Page MenuHomeIsabelle/Phabricator

No OneTemporary

This file is larger than 256 KB, so syntax highlighting was skipped.
diff --git a/thys/Conditional_Transfer_Rule/UD/UD_Reference.thy b/thys/Conditional_Transfer_Rule/UD/UD_Reference.thy
--- a/thys/Conditional_Transfer_Rule/UD/UD_Reference.thy
+++ b/thys/Conditional_Transfer_Rule/UD/UD_Reference.thy
@@ -1,346 +1,346 @@
-(* Title: UD/UD_Reference.thy
- Author: Mihails Milehins
- Copyright 2021 (C) Mihails Milehins
-
-Reference manual for the UD.
-*)
-
-section\<open>UD\<close>
-theory UD_Reference
- imports
- UD
- "../Reference_Prerequisites"
-begin
-
-
-
-subsection\<open>Introduction\<close>
-
-
-subsubsection\<open>Background\<close>
-
-text\<open>
-This section presents a reference manual for the sub-framework UD.
-The UD can be used for the elimination of
-\textit{sort constraints} (e.g., see \cite{altenkirch_constructive_2007})
-and unoverloading of definitions in the object logic Isabelle/HOL of the
-formal proof assistant Isabelle.
-The UD evolved from the author's work on an extension of the
-framework \textit{Types-To-Sets}
-(see
-\cite{blanchette_types_2016, kuncar_types_2019, immler_smooth_2019, immler_automation_2019},
-for a description of the framework Types-To-Sets
-and \cite{milehins_extension_2021} for a description of the author's extension)
-and builds upon certain ideas expressed in \cite{kaufmann_mechanized_2010}.
-\<close>
-
-
-subsubsection\<open>Purpose and scope\<close>
-
-text\<open>
-The primary functionality of the framework is available via the Isabelle/Isar
-command @{command ud}. This command automates the processes of the
-elimination of sort constraints and unoverloading of definitions.
-Thus, the command @{command ud} allows for the synthesis
-of the convenience constants and theorems that are usually needed for the
-application of the derivation step 2 of the original relativization algorithm
-of Types-To-Sets (see subsection 5.4 in \cite{blanchette_types_2016}). However,
-it is expected that the command can be useful for other purposes.
-\<close>
-
-
-subsubsection\<open>Related and previous work\<close>
-
-text\<open>
-The functionality provided by the command @{command ud} shares similarities
-with the functionality provided by the algorithms for the elimination of
-sort constraints and elimination of overloading that were
-presented in \cite{kaufmann_mechanized_2010}
-and with the algorithm associated with the command
-\mbox{\textbf{unoverload\_definition}} that was proposed
-in \cite{immler_automation_2019}.
-Nonetheless, technically, unlike \mbox{\textbf{unoverload\_definition}},
-the command @{command ud} does
-not require the additional axiom UO associated with Types-To-Sets for
-its operation (see \cite{blanchette_types_2016},
-\cite{immler_automation_2019}), it uses
-the \textit{definitional axioms} (e.g., see \cite{kaufmann_mechanized_2010})
-instead of arbitrary theorems supplied by the user
-and it is independent of the infrastructure associated with
-the \textit{axiomatic type classes}
-\cite{nipkow_type_1991,wenzel_type_1997,altenkirch_constructive_2007}.
-
-It should also be mentioned that the Isabelle/ML code from the main
-distribution of Isabelle was frequently reused during the
-development of the UD.
-\<close>
-
-
-
-subsection\<open>Theory\label{sec:ud_theory}\<close>
-
-
-text\<open>
-The general references for this subsection are
-\cite{kaufmann_mechanized_2010} and \cite{yang_comprehending_2017}.
-The command @{command ud} relies
-on a restricted (non-recursive) variant of the
-\textit{classical overloading elimination algorithm}
-that was originally proposed in \cite{kaufmann_mechanized_2010}.
-It is assumed that there exists
-a variable $ud_{\mathsf{with}}$ that stores theorems of the
-form $c_{\tau} = c_{\mathsf{with}}\ \bar{*}$, where $c_{\tau}$ and
-$c_{\mathsf{with}}$ are distinct \textit{constant-instances}
-and $\bar{*}$ is a finite sequence of \textit{uninterpreted constant-instances},
-such that, if $c_{\tau}$ depends on a type variable $\alpha_{\Upsilon}$,
-with $\Upsilon$ being a \textit{type class}
-\cite{nipkow_type_1991,wenzel_type_1997,altenkirch_constructive_2007}
-that depends on the overloaded
-constants $\bar{*'}$, then $\bar{*}$ contains $\bar{*'}$ as a subsequence.
-Lastly, the binary operation $\cup$ is defined in a manner such that
-for any sequences $\bar{*}$ and $\bar{*'}$, $\bar{*} \cup \bar{*'}$
-is a sequence that consists of all elements of the union of the
-elements of $\bar{*}$ and $\bar{*'}$ without duplication.
-Assuming an underlying
-\textit{well-formed definitional theory} $D$,
-the input to the algorithm is a constant-instance $c_{\sigma}$.
-Given the constant-instance $c_{\sigma}$,
-there exists at most one definitional axiom
-$c_{\tau} = \phi_{\tau}\left[\bar{*}\right]$
-in $D$ such that $c_{\sigma} \leq c_{\tau}$: otherwise
-the \textit{orthogonality} of $D$ and,
-therefore, the \textit{well-formedness}
-of $D$ are violated ($\phi$ is assumed to be parameterized by
-the types that it can have with respect to the
-type substitution operation,
-and $\bar{*}$ in $c_{\tau} = \phi_{\tau}\left[\bar{*}\right]$
-is a list of all uninterpreted constant-instances that
-occur in $\phi_{\tau}\left[\bar{*}\right]$).
-
-If a definitional axiom $c_{\tau}=\phi_{\tau}\left[\bar{*}\right]$
-such that $c_{\sigma} \leq c_{\tau}$
-exists for the constant-instance $c_{\sigma}$,
-then the following derivation is applied to it by the algorithm
-\[
-\infer[(6)]
-{\vdash c_{\sigma} = c_{\mathsf{with}}\ \left(\bar{*} \cup \bar{*'}\right)}
-{
-\infer[(5)]
-{
-\vdash c_{\mathsf{with}}\ \left(\bar{*} \cup \bar{*'}\right) =
-\phi_{\mathsf{with}}\left[\bar{*} \cup \bar{*'}\right]
-}
-{
-\infer[(4)]
-{\vdash c_{\mathsf{with}}\ ?\bar{f} = \phi_{\mathsf{with}}\left[?\bar{f}\right]}
-{
-\infer[(3)]
-{\vdash c_{\mathsf{with}} = (\lambda \bar{f}.\ \phi_{\mathsf{with}}\left[\bar{f}\right])}
-{
-\infer[(2)]
-{\vdash c_{\sigma}=\phi_{\mathsf{with}}\left[\bar{*} \cup \bar{*'}\right]}
-{
-\infer[(1)]
-{\vdash c_{\sigma}=\phi_{\sigma}\left[\bar{*}\right]}
-{\vdash c_{\tau}=\phi_{\tau}\left[\bar{*}\right]}
-}
-}
-}
-}
-}
-\]
-In step 1, the previously established
-property $c_{\sigma} \leq c_{\tau}$ is used to create the
-(extended variant of the) type substitution
-map $\rho$ such that $\sigma = \rho \left( \tau \right)$
-(see \cite{kuncar_types_2015}) and perform the type
-substitution in $c_{\tau}=\phi_{\tau}\left[\bar{*}\right]$
-to obtain $c_{\sigma}=\phi_{\sigma}\left[\bar{*}\right]$;
-in step 2, the collection of theorems $ud_{\mathsf{with}}$ is unfolded,
-using it as a term rewriting system, possibly introducing further uninterpreted
-constants $\bar{*'}$; in step 3, the term on the right-hand side of the
-theorem is processed by removing the sort constraints from all type
-variables that occur in it, replacing every uninterpreted constant-instance
-(this excludes all built-in constants of Isabelle/HOL) that occurs in it by a
-fresh term variable, and applying the abstraction until the resulting term
-is closed: this term forms the right-hand side of a new definitional axiom
-of a fresh constant $c_{\mathsf{with}}$ (if the conditions associated with
-the definitional principles of Isabelle/HOL \cite{yang_comprehending_2017}
-are satisfied); step 4 is justified by the beta-contraction;
-step 5 is a substitution of the uninterpreted constants $\bar{*} \cup \bar{*'}$;
-step 6 follows trivially from the results of the application of steps 2 and 5.
-
-The implementation of the command @{command ud} closely follows the steps of
-the algorithm outlined above. Thus, at the end of the successful
-execution, the command declares the constant $c_{\mathsf{with}}$ and stores the
-constant-instance definition that is obtained at the end of step 3 of
-the algorithm UD; furthermore, the command adds the theorem that is
-obtained after the execution of step 6 of the algorithm
-to $ud_{\mathsf{with}}$.
-
-Unlike the classical overloading elimination algorithm,
-the algorithm employed in the implementation
-of the command @{command ud} is not recursive. Thus, the users are responsible
-for maintaining an adequate collection of theorems $ud_{\mathsf{with}}$.
-Nonetheless, in this case, the users can provide their own
-unoverloaded constants $c_{\mathsf{with}}$ and the associated theorems
-$c_{\sigma} = c_{\mathsf{with}}\ \bar{*}$ for any constant-instance $c_{\sigma}$.
-From the perspective of the relativization algorithm associated with
-Types-To-Sets this can be useful because there is no
-guarantee that the automatically synthesized constants $c_{\mathsf{with}}$
-will possess desirable parametricity characteristics
-(e.g., see \cite{kuncar_types_2015} and \cite{immler_smooth_2019}).
-Unfortunately, the implemented algorithm still suffers from the fundamental
-limitation that was already outlined in \cite{kaufmann_mechanized_2010},
-\cite{blanchette_types_2016} and \cite{kuncar_types_2019}:
-it does not offer a solution for handling the
-constants whose types contain occurrences of the type constructors whose
-type definitions contain occurrences of unresolvable overloading.
-\<close>
-
-
-
-subsection\<open>Syntax\<close>
-
-text\<open>
-This subsection presents the syntactic categories that are associated with the
-command @{command ud}. It is important to note that the presentation is
-only approximate.
-\<close>
-
-text\<open>
-
-\begin{matharray}{rcl}
- @{command_def "ud"} & : & \<open>theory \<rightarrow> theory\<close>\\
-\end{matharray}
-
- \<^medskip>
-
- \<^rail>\<open>@@{command ud} binding? const mixfix?\<close>
-
- \<^descr> \<^theory_text>\<open>ud\<close> (\<open>b\<close>) \<open>const\<close> (\<open>mixfix\<close>) provides access to the algorithm for
-the elimination of sort constraints and unoverloading of definitions
-that was described in subsection \ref{sec:ud_theory}.
-The optional binding \<open>b\<close> is used for the specification
-of the names of the entities added by the command to the theory and the
-optional argument \<open>mixfix\<close> is used for the specification
-of the concrete inner syntax for the constant in the usual manner
-(e.g., see \cite{wenzel_isabelle/isar_2019-1}).
-If either \<open>b\<close> or \<open>mixfix\<close> are not specified by the user, then the command
-introduces sensible defaults. Following the specification of the
-definition of the constant, an additional theorem that establishes
-the relationship between the newly introduced constant and the
-constant provided by the user as an input is established and added
-to the dynamic fact @{thm [source] ud_with}.
-\<close>
-
-
-
-subsection\<open>Examples\label{sec:ud_ex}\<close>
-
-text\<open>
-In this subsection, some of the capabilities of the UD are
-demonstrated by example. The examples that are presented in this subsection are
-expected to be sufficient for beginning an independent exploration of the
-framework, but do not cover the entire spectrum of the functionality
-and the problems that one may encounter while using it.
-\<close>
-
-
-subsubsection\<open>Type classes\<close>
-
-text\<open>
-We begin the exploration of the capabilities of the framework by considering
-the constant @{const mono}.
-The constant @{const mono} is an overloaded constant that is defined
-in the standard library of Isabelle/HOL \cite{noauthor_isabellehol_2020} as
-\begin{center}
-@{thm [names_short = true] mono_def[no_vars]}
-\end{center}
-for any @{term [show_sorts] "f::'a::order\<Rightarrow>'b::order"}.
-Technically, there exist two distinct constants that are associated with
-the definition of the constant @{const mono}
-(see, e.g., \cite{berardi_local_2009}
-for further details): @{const order.mono} that is provided by the
-Isabelle system automatically and the constant @{const mono} itself.
-The constants are unoverloaded successively using the command @{command ud}:
-\<close>
-ud \<open>order.mono\<close>
-ud mono' \<open>mono\<close>
-text\<open>
-The invocation of the commands above declares the constant @{const mono.with}
-that is defined as
-\begin{center}
-@{thm mono.with_def[no_vars]}
-\end{center}
-and provides the theorem @{thm [source] mono.with} given by
-\begin{center}
-@{thm mono.with[no_vars]}
-\end{center}
-and the theorem @{thm [source] mono'.with} given by
-\begin{center}
-@{thm mono'.with[no_vars]}.
-\end{center}
-The theorems establish the relationship between the unoverloaded constant
-@{const mono.with} and the overloaded constant @{const mono}:
-both theorems are automatically added to the dynamic fact
-@{thm [source] ud_with}.
-\<close>
-
-
-subsubsection\<open>Low-level overloading\<close>
-
-text\<open>
-The following example closely follows Example 5 in section 5.2. in
-\cite{kaufmann_mechanized_2010}.
-\<close>
-
-consts pls :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"
-
-overloading
-pls_nat \<equiv> "pls::nat \<Rightarrow> nat \<Rightarrow> nat"
-pls_times \<equiv> "pls::'a \<times> 'b \<Rightarrow> 'a \<times> 'b \<Rightarrow> 'a \<times> 'b"
-begin
-definition pls_nat :: "nat \<Rightarrow> nat \<Rightarrow> nat" where "pls_nat a b = a + b"
-definition pls_times :: "'a \<times> 'b \<Rightarrow> 'a \<times> 'b \<Rightarrow> 'a \<times> 'b"
- where "pls_times \<equiv> \<lambda>x y. (pls (fst x) (fst y), pls (snd x) (snd y))"
-end
-
-ud pls_nat \<open>pls::nat \<Rightarrow> nat \<Rightarrow> nat\<close>
-ud pls_times \<open>pls::'a \<times> 'b \<Rightarrow> 'a \<times> 'b \<Rightarrow> 'a \<times> 'b\<close>
-
-text\<open>
-As expected, two new unoverloaded constants are produced via
-the invocations of the command @{command ud} above. The first constant,
-\<^const>\<open>pls_nat.with\<close>, corresponds to \<open>pls::nat \<Rightarrow> nat \<Rightarrow> nat\<close> and is given by
-\begin{center}
-@{thm pls_nat.with_def[no_vars]},
-\end{center}
-the second constant, \<^const>\<open>pls_times.with\<close>, corresponds to
-\begin{center}
-\<open>pls::'a \<times> 'b \<Rightarrow> 'a \<times> 'b \<Rightarrow> 'a \<times> 'b\<close>
-\end{center}
-and is given by
-\begin{center}
-@{thm pls_times.with_def[no_vars]}.
-\end{center}
-The theorems that establish the relationship between the overloaded and
-the unoverloaded constants are given by
-\begin{center}
-@{thm pls_nat.with}
-\end{center}
-and
-\begin{center}
-@{thm pls_times.with}.
-\end{center}
-The definitions of the constants \<^const>\<open>pls_nat.with\<close> and
-\<^const>\<open>pls_times.with\<close> are consistent with the ones suggested in
-\cite{kaufmann_mechanized_2010}. Nonetheless, of course, it is
-important to keep in mind that the command @{command ud}
-has a more restricted scope of applicability than the
-algorithm suggested in \cite{kaufmann_mechanized_2010}.
-\<close>
-
-text\<open>\newpage\<close>
-
+(* Title: UD/UD_Reference.thy
+ Author: Mihails Milehins
+ Copyright 2021 (C) Mihails Milehins
+
+Reference manual for the UD.
+*)
+
+section\<open>UD\<close>
+theory UD_Reference
+ imports
+ UD
+ "../Reference_Prerequisites"
+begin
+
+
+
+subsection\<open>Introduction\<close>
+
+
+subsubsection\<open>Background\<close>
+
+text\<open>
+This section presents a reference manual for the sub-framework UD.
+The UD can be used for the elimination of
+\textit{sort constraints} (e.g., see \cite{altenkirch_constructive_2007})
+and unoverloading of definitions in the object logic Isabelle/HOL of the
+formal proof assistant Isabelle.
+The UD evolved from the author's work on an extension of the
+framework \textit{Types-To-Sets}
+(see
+\cite{blanchette_types_2016, kuncar_types_2019, immler_smooth_2019, immler_automation_2019},
+for a description of the framework Types-To-Sets
+and \cite{milehins_extension_2021} for a description of the author's extension)
+and builds upon certain ideas expressed in \cite{kaufmann_mechanized_2010}.
+\<close>
+
+
+subsubsection\<open>Purpose and scope\<close>
+
+text\<open>
+The primary functionality of the framework is available via the Isabelle/Isar
+command @{command ud}. This command automates the processes of the
+elimination of sort constraints and unoverloading of definitions.
+Thus, the command @{command ud} allows for the synthesis
+of the convenience constants and theorems that are usually needed for the
+application of the derivation step 2 of the original relativization algorithm
+of Types-To-Sets (see subsection 5.4 in \cite{blanchette_types_2016}). However,
+it is expected that the command can be useful for other purposes.
+\<close>
+
+
+subsubsection\<open>Related and previous work\<close>
+
+text\<open>
+The functionality provided by the command @{command ud} shares similarities
+with the functionality provided by the algorithms for the elimination of
+sort constraints and elimination of overloading that were
+presented in \cite{kaufmann_mechanized_2010}
+and with the algorithm associated with the command
+\mbox{\textbf{unoverload\_definition}} that was proposed
+in \cite{immler_automation_2019}.
+Nonetheless, technically, unlike \mbox{\textbf{unoverload\_definition}},
+the command @{command ud} does
+not require the additional axiom UO associated with Types-To-Sets for
+its operation (see \cite{blanchette_types_2016},
+\cite{immler_automation_2019}), it uses
+the \textit{definitional axioms} (e.g., see \cite{kaufmann_mechanized_2010})
+instead of arbitrary theorems supplied by the user
+and it is independent of the infrastructure associated with
+the \textit{axiomatic type classes}
+\cite{nipkow_type_1991,wenzel_type_1997,altenkirch_constructive_2007}.
+
+It should also be mentioned that the Isabelle/ML code from the main
+distribution of Isabelle was frequently reused during the
+development of the UD.
+\<close>
+
+
+
+subsection\<open>Theory\label{sec:ud_theory}\<close>
+
+
+text\<open>
+The general references for this subsection are
+\cite{kaufmann_mechanized_2010} and \cite{yang_comprehending_2017}.
+The command @{command ud} relies
+on a restricted (non-recursive) variant of the
+\textit{classical overloading elimination algorithm}
+that was originally proposed in \cite{kaufmann_mechanized_2010}.
+It is assumed that there exists
+a variable $ud_{\mathsf{with}}$ that stores theorems of the
+form $c_{\tau} = c_{\mathsf{with}}\ \bar{*}$, where $c_{\tau}$ and
+$c_{\mathsf{with}}$ are distinct \textit{constant-instances}
+and $\bar{*}$ is a finite sequence of \textit{uninterpreted constant-instances},
+such that, if $c_{\tau}$ depends on a type variable $\alpha_{\Upsilon}$,
+with $\Upsilon$ being a \textit{type class}
+\cite{nipkow_type_1991,wenzel_type_1997,altenkirch_constructive_2007}
+that depends on the overloaded
+constants $\bar{*'}$, then $\bar{*}$ contains $\bar{*'}$ as a subsequence.
+Lastly, the binary operation $\cup$ is defined in a manner such that
+for any sequences $\bar{*}$ and $\bar{*'}$, $\bar{*} \cup \bar{*'}$
+is a sequence that consists of all elements of the union of the
+elements of $\bar{*}$ and $\bar{*'}$ without duplication.
+Assuming an underlying
+\textit{well-formed definitional theory} $D$,
+the input to the algorithm is a constant-instance $c_{\sigma}$.
+Given the constant-instance $c_{\sigma}$,
+there exists at most one definitional axiom
+$c_{\tau} = \phi_{\tau}\left[\bar{*}\right]$
+in $D$ such that $c_{\sigma} \leq c_{\tau}$: otherwise
+the \textit{orthogonality} of $D$ and,
+therefore, the \textit{well-formedness}
+of $D$ are violated ($\phi$ is assumed to be parameterized by
+the types that it can have with respect to the
+type substitution operation,
+and $\bar{*}$ in $c_{\tau} = \phi_{\tau}\left[\bar{*}\right]$
+is a list of all uninterpreted constant-instances that
+occur in $\phi_{\tau}\left[\bar{*}\right]$).
+
+If a definitional axiom $c_{\tau}=\phi_{\tau}\left[\bar{*}\right]$
+such that $c_{\sigma} \leq c_{\tau}$
+exists for the constant-instance $c_{\sigma}$,
+then the following derivation is applied to it by the algorithm
+\[
+\infer[(6)]
+{\vdash c_{\sigma} = c_{\mathsf{with}}\ \left(\bar{*} \cup \bar{*'}\right)}
+{
+\infer[(5)]
+{
+\vdash c_{\mathsf{with}}\ \left(\bar{*} \cup \bar{*'}\right) =
+\phi_{\mathsf{with}}\left[\bar{*} \cup \bar{*'}\right]
+}
+{
+\infer[(4)]
+{\vdash c_{\mathsf{with}}\ ?\bar{f} = \phi_{\mathsf{with}}\left[?\bar{f}\right]}
+{
+\infer[(3)]
+{\vdash c_{\mathsf{with}} = (\lambda \bar{f}.\ \phi_{\mathsf{with}}\left[\bar{f}\right])}
+{
+\infer[(2)]
+{\vdash c_{\sigma}=\phi_{\mathsf{with}}\left[\bar{*} \cup \bar{*'}\right]}
+{
+\infer[(1)]
+{\vdash c_{\sigma}=\phi_{\sigma}\left[\bar{*}\right]}
+{\vdash c_{\tau}=\phi_{\tau}\left[\bar{*}\right]}
+}
+}
+}
+}
+}
+\]
+In step 1, the previously established
+property $c_{\sigma} \leq c_{\tau}$ is used to create the
+(extended variant of the) type substitution
+map $\rho$ such that $\sigma = \rho \left( \tau \right)$
+(see \cite{kuncar_types_2015}) and perform the type
+substitution in $c_{\tau}=\phi_{\tau}\left[\bar{*}\right]$
+to obtain $c_{\sigma}=\phi_{\sigma}\left[\bar{*}\right]$;
+in step 2, the collection of theorems $ud_{\mathsf{with}}$ is unfolded,
+using it as a term rewriting system, possibly introducing further uninterpreted
+constants $\bar{*'}$; in step 3, the term on the right-hand side of the
+theorem is processed by removing the sort constraints from all type
+variables that occur in it, replacing every uninterpreted constant-instance
+(this excludes all built-in constants of Isabelle/HOL) that occurs in it by a
+fresh term variable, and applying the abstraction until the resulting term
+is closed: this term forms the right-hand side of a new definitional axiom
+of a fresh constant $c_{\mathsf{with}}$ (if the conditions associated with
+the definitional principles of Isabelle/HOL \cite{yang_comprehending_2017}
+are satisfied); step 4 is justified by the beta-contraction;
+step 5 is a substitution of the uninterpreted constants $\bar{*} \cup \bar{*'}$;
+step 6 follows trivially from the results of the application of steps 2 and 5.
+
+The implementation of the command @{command ud} closely follows the steps of
+the algorithm outlined above. Thus, at the end of the successful
+execution, the command declares the constant $c_{\mathsf{with}}$ and stores the
+constant-instance definition that is obtained at the end of step 3 of
+the algorithm UD; furthermore, the command adds the theorem that is
+obtained after the execution of step 6 of the algorithm
+to $ud_{\mathsf{with}}$.
+
+Unlike the classical overloading elimination algorithm,
+the algorithm employed in the implementation
+of the command @{command ud} is not recursive. Thus, the users are responsible
+for maintaining an adequate collection of theorems $ud_{\mathsf{with}}$.
+Nonetheless, in this case, the users can provide their own
+unoverloaded constants $c_{\mathsf{with}}$ and the associated theorems
+$c_{\sigma} = c_{\mathsf{with}}\ \bar{*}$ for any constant-instance $c_{\sigma}$.
+From the perspective of the relativization algorithm associated with
+Types-To-Sets this can be useful because there is no
+guarantee that the automatically synthesized constants $c_{\mathsf{with}}$
+will possess desirable parametricity characteristics
+(e.g., see \cite{kuncar_types_2015} and \cite{immler_smooth_2019}).
+Unfortunately, the implemented algorithm still suffers from the fundamental
+limitation that was already outlined in \cite{kaufmann_mechanized_2010},
+\cite{blanchette_types_2016} and \cite{kuncar_types_2019}:
+it does not offer a solution for handling the
+constants whose types contain occurrences of the type constructors whose
+type definitions contain occurrences of unresolvable overloading.
+\<close>
+
+
+
+subsection\<open>Syntax\<close>
+
+text\<open>
+This subsection presents the syntactic categories that are associated with the
+command @{command ud}. It is important to note that the presentation is
+only approximate.
+\<close>
+
+text\<open>
+
+\begin{matharray}{rcl}
+ @{command_def "ud"} & : & \<open>theory \<rightarrow> theory\<close>\\
+\end{matharray}
+
+ \<^medskip>
+
+ \<^rail>\<open>@@{command ud} binding? const mixfix?\<close>
+
+ \<^descr> \<^theory_text>\<open>ud\<close> (\<open>b\<close>) \<open>const\<close> (\<open>mixfix\<close>) provides access to the algorithm for
+the elimination of sort constraints and unoverloading of definitions
+that was described in subsection \ref{sec:ud_theory}.
+The optional binding \<open>b\<close> is used for the specification
+of the names of the entities added by the command to the theory and the
+optional argument \<open>mixfix\<close> is used for the specification
+of the concrete inner syntax for the constant in the usual manner
+(e.g., see \cite{wenzel_isabelle/isar_2019-1}).
+If either \<open>b\<close> or \<open>mixfix\<close> are not specified by the user, then the command
+introduces sensible defaults. Following the specification of the
+definition of the constant, an additional theorem that establishes
+the relationship between the newly introduced constant and the
+constant provided by the user as an input is established and added
+to the dynamic fact @{thm [source] ud_with}.
+\<close>
+
+
+
+subsection\<open>Examples\label{sec:ud_ex}\<close>
+
+text\<open>
+In this subsection, some of the capabilities of the UD are
+demonstrated by example. The examples that are presented in this subsection are
+expected to be sufficient for beginning an independent exploration of the
+framework, but do not cover the entire spectrum of the functionality
+and the problems that one may encounter while using it.
+\<close>
+
+
+subsubsection\<open>Type classes\<close>
+
+text\<open>
+We begin the exploration of the capabilities of the framework by considering
+the constant @{const mono}.
+The constant @{const mono} is an overloaded constant that is defined
+in the standard library of Isabelle/HOL \cite{noauthor_isabellehol_2020} as
+\begin{center}
+@{thm [names_short = true] mono_def[no_vars]}
+\end{center}
+for any @{term [show_sorts] "f::'a::order\<Rightarrow>'b::order"}.
+Technically, there exist two distinct constants that are associated with
+the definition of the constant @{const mono}
+(see, e.g., \cite{berardi_local_2009}
+for further details): @{const order.mono} that is provided by the
+Isabelle system automatically and the constant @{const mono} itself.
+The constants are unoverloaded successively using the command @{command ud}:
+\<close>
+ud \<open>order.mono\<close>
+ud mono' \<open>mono\<close>
+text\<open>
+The invocation of the commands above declares the constant @{const mono.with}
+that is defined as
+\begin{center}
+@{thm mono.with_def[no_vars]}
+\end{center}
+and provides the theorem @{thm [source] mono.with} given by
+\begin{center}
+@{thm mono.with[no_vars]}
+\end{center}
+and the theorem @{thm [source] mono'.with} given by
+\begin{center}
+@{thm mono'.with[no_vars]}.
+\end{center}
+The theorems establish the relationship between the unoverloaded constant
+@{const mono.with} and the overloaded constant @{const mono}:
+both theorems are automatically added to the dynamic fact
+@{thm [source] ud_with}.
+\<close>
+
+
+subsubsection\<open>Low-level overloading\<close>
+
+text\<open>
+The following example closely follows Example 5 in section 5.2. in
+\cite{kaufmann_mechanized_2010}.
+\<close>
+
+consts pls :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"
+
+overloading
+pls_nat \<equiv> "pls::nat \<Rightarrow> nat \<Rightarrow> nat"
+pls_times \<equiv> "pls::'a \<times> 'b \<Rightarrow> 'a \<times> 'b \<Rightarrow> 'a \<times> 'b"
+begin
+definition pls_nat :: "nat \<Rightarrow> nat \<Rightarrow> nat" where "pls_nat a b = a + b"
+definition pls_times :: "'a \<times> 'b \<Rightarrow> 'a \<times> 'b \<Rightarrow> 'a \<times> 'b"
+ where "pls_times \<equiv> \<lambda>x y. (pls (fst x) (fst y), pls (snd x) (snd y))"
+end
+
+ud pls_nat \<open>pls::nat \<Rightarrow> nat \<Rightarrow> nat\<close>
+ud pls_times \<open>pls::'a \<times> 'b \<Rightarrow> 'a \<times> 'b \<Rightarrow> 'a \<times> 'b\<close>
+
+text\<open>
+As expected, two new unoverloaded constants are produced via
+the invocations of the command @{command ud} above. The first constant,
+\<^const>\<open>pls_nat.with\<close>, corresponds to \<open>pls::nat \<Rightarrow> nat \<Rightarrow> nat\<close> and is given by
+\begin{center}
+@{thm pls_nat.with_def[no_vars]},
+\end{center}
+the second constant, \<^const>\<open>pls_times.with\<close>, corresponds to
+\begin{center}
+\<open>pls::'a \<times> 'b \<Rightarrow> 'a \<times> 'b \<Rightarrow> 'a \<times> 'b\<close>
+\end{center}
+and is given by
+\begin{center}
+@{thm pls_times.with_def[no_vars]}.
+\end{center}
+The theorems that establish the relationship between the overloaded and
+the unoverloaded constants are given by
+\begin{center}
+@{thm pls_nat.with}
+\end{center}
+and
+\begin{center}
+@{thm pls_times.with}.
+\end{center}
+The definitions of the constants \<^const>\<open>pls_nat.with\<close> and
+\<^const>\<open>pls_times.with\<close> are consistent with the ones suggested in
+\cite{kaufmann_mechanized_2010}. Nonetheless, of course, it is
+important to keep in mind that the command @{command ud}
+has a more restricted scope of applicability than the
+algorithm suggested in \cite{kaufmann_mechanized_2010}.
+\<close>
+
+text\<open>\newpage\<close>
+
end
\ No newline at end of file
diff --git a/thys/Design_Theory/BIBD.thy b/thys/Design_Theory/BIBD.thy
--- a/thys/Design_Theory/BIBD.thy
+++ b/thys/Design_Theory/BIBD.thy
@@ -1,943 +1,943 @@
-(* Title: BIBD
- Author: Chelsea Edmonds
-*)
-
-theory BIBD imports Block_Designs
-begin
-
-section \<open>BIBD's\<close>
-text \<open>BIBD's are perhaps the most commonly studied type of design in combinatorial design theory,
-and usually the first type of design explored in a design theory course.
-These designs are a type of t-design, where $t = 2$\<close>
-
-subsection \<open>BIBD Basics\<close>
-locale bibd = t_design \<V> \<B> \<k> 2 \<Lambda>
- for point_set ("\<V>") and block_collection ("\<B>")
- and u_block_size ("\<k>") and index ("\<Lambda>")
-
-begin
-
-lemma min_block_size_2: "\<k> \<ge> 2"
- using block_size_t by simp
-
-lemma points_index_pair: "y \<in> \<V> \<Longrightarrow> x \<in> \<V> \<Longrightarrow> x \<noteq> y \<Longrightarrow> size ({# bl \<in># \<B> . {x, y} \<subseteq> bl#}) = \<Lambda>"
- using balanced card_2_iff empty_subsetI insert_subset points_index_def
- by (metis of_nat_numeral)
-
-lemma index_one_empty_rm_blv [simp]:
- assumes "\<Lambda> = 1" and " blv \<in># \<B>" and "p \<subseteq> blv" and "card p = 2"
- shows "{#bl \<in># remove1_mset blv \<B> . p \<subseteq> bl#} = {#}"
-proof -
- have blv_in: "blv \<in># filter_mset ((\<subseteq>) p) \<B>"
- using assms by simp
- have "p \<subseteq> \<V>" using assms wellformed by auto
- then have "size (filter_mset ((\<subseteq>) p) \<B>) = 1"
- using balanced assms by (simp add: points_index_def)
- then show ?thesis using blv_in filter_diff_mset filter_single_mset
- by (metis (no_types, lifting) add_mset_eq_single assms(3) insert_DiffM size_1_singleton_mset)
-qed
-
-lemma index_one_alt_bl_not_exist:
- assumes "\<Lambda> = 1" and " blv \<in># \<B>" and "p \<subseteq> blv" and "card p = 2"
- shows" \<And> bl. bl \<in># remove1_mset blv \<B> \<Longrightarrow> \<not> (p \<subseteq> bl) "
- using index_one_empty_rm_blv
- by (metis assms(1) assms(2) assms(3) assms(4) filter_mset_empty_conv)
-
-subsection \<open>Necessary Conditions for Existence\<close>
-
-text \<open>The necessary conditions on the existence of a $(v, k, \lambda)$-bibd are one of the
-fundamental first theorems on designs. Proofs based off MATH3301 lecture notes \cite{HerkeLectureNotes2016}
- and Stinson \cite{stinsonCombinatorialDesignsConstructions2004}\<close>
-
-lemma necess_cond_1_rhs:
- assumes "x \<in> \<V>"
- shows "size ({# p \<in># (mset_set (\<V> - {x}) \<times># {# bl \<in># \<B> . x \<in> bl #}). fst p \<in> snd p#}) = \<Lambda> * (\<v>- 1)"
-proof -
- let ?M = "mset_set (\<V> - {x})"
- let ?B = "{# bl \<in># \<B> . x \<in> bl #}"
- have m_distinct: "distinct_mset ?M" using assms mset_points_distinct_diff_one by simp
- have y_point: "\<And> y . y \<in># ?M \<Longrightarrow> y \<in> \<V>" using assms
- by (simp add: finite_sets)
- have b_contents: "\<And> bl. bl \<in># ?B \<Longrightarrow> x \<in> bl" using assms by auto
- have "\<And> y. y \<in># ?M \<Longrightarrow> y \<noteq> x" using assms
- by (simp add: finite_sets)
- then have "\<And> y .y \<in># ?M \<Longrightarrow> size ({# bl \<in># ?B . {x, y} \<subseteq> bl#}) = nat \<Lambda>"
- using points_index_pair filter_filter_mset_ss_member y_point assms finite_sets index_ge_zero
- by (metis nat_0_le nat_int_comparison(1))
- then have "\<And> y .y \<in># ?M \<Longrightarrow> size ({# bl \<in># ?B . x \<in> bl \<and> y \<in> bl#}) = nat \<Lambda>"
- by auto
- then have bl_set_size: "\<And> y . y \<in># ?M \<Longrightarrow> size ({# bl \<in># ?B . y \<in> bl#}) = nat \<Lambda>"
- using b_contents by (metis (no_types, lifting) filter_mset_cong)
- then have final_size: "size (\<Sum>p\<in>#?M . ({#p#} \<times># {#bl \<in># ?B. p \<in> bl#})) = size (?M) * (nat \<Lambda>)"
- using m_distinct size_Union_distinct_cart_prod_filter bl_set_size index_ge_zero by blast
- have "size ?M = \<v> - 1" using v_non_zero
- by (simp add: assms(1) finite_sets)
- thus ?thesis using final_size
- by (simp add: set_break_down_left index_ge_zero)
-qed
-
-lemma necess_cond_1_lhs:
- assumes "x \<in> \<V>"
- shows "size ({# p \<in># (mset_set (\<V> - {x}) \<times># {# bl \<in># \<B> . x \<in> bl #}). fst p \<in> snd p#})
- = (\<B> rep x) * (\<k> - 1)"
- (is "size ({# p \<in># (?M \<times># ?B). fst p \<in> snd p#}) = (\<B> rep x) * (\<k> - 1) ")
-proof -
- have "\<And> y. y \<in># ?M \<Longrightarrow> y \<noteq> x" using assms
- by (simp add: finite_sets)
- have distinct_m: "distinct_mset ?M" using assms mset_points_distinct_diff_one by simp
- have finite_M: "finite (\<V> - {x})" using finite_sets by auto
- have block_choices: "size ?B = \<B> rep x"
- by (simp add: assms(1) point_replication_number_def)
- have bl_size: "\<forall> bl \<in># ?B. card {p \<in> \<V> . p \<in> bl } = \<k> " using uniform_unfold_point_set by simp
- have x_in_set: "\<forall> bl \<in># ?B . {x} \<subseteq> {p \<in> \<V>. p \<in> bl}" using assms by auto
- then have "\<forall> bl \<in># ?B. card {p \<in> (\<V> - {x}) . p \<in> bl } = card ({p \<in> \<V> . p \<in> bl } - {x})"
- by (simp add: set_filter_diff_card)
- then have "\<forall> bl \<in># ?B. card {p \<in> (\<V> - {x}) . p \<in> bl } = \<k> - 1"
- using bl_size x_in_set card_Diff_subset finite_sets k_non_zero by auto
- then have "\<And> bl . bl \<in># ?B \<Longrightarrow> size {#p \<in># ?M . p \<in> bl#} = nat (\<k> - 1)"
- using assms finite_M card_size_filter_eq by auto
- then have "size (\<Sum>bl\<in>#?B. ( {# p \<in># ?M . p \<in> bl #} \<times># {#bl#})) = size (?B) * nat (\<k> - 1)"
- using distinct_m size_Union_distinct_cart_prod_filter2 by blast
- thus ?thesis using block_choices k_non_zero by (simp add: set_break_down_right)
-qed
-
-lemma r_constant: "x \<in> \<V> \<Longrightarrow> (\<B> rep x) * (\<k> -1) = \<Lambda> * (\<v> - 1)"
- using necess_cond_1_rhs necess_cond_1_lhs design_points_nempty by force
-
-lemma replication_number_value:
- assumes "x \<in> \<V>"
- shows "(\<B> rep x) = \<Lambda> * (\<v> - 1) div (\<k> - 1)"
- using min_block_size_2 r_constant assms diff_gt_0_iff_gt diff_self zle_diff1_eq numeral_le_one_iff
- by (metis less_int_code(1) linorder_neqE_linordered_idom nonzero_mult_div_cancel_right semiring_norm(69))
-
-lemma r_constant_alt: "\<forall> x \<in> \<V>. \<B> rep x = \<Lambda> * (\<v> - 1) div (\<k> - 1)"
- using r_constant replication_number_value by blast
-
-end
-
-text \<open>Using the first necessary condition, it is possible to show that a bibd has
-a constant replication number\<close>
-
-sublocale bibd \<subseteq> constant_rep_design \<V> \<B> "(\<Lambda> * (\<v> - 1) div (\<k> - 1))"
- using r_constant_alt by (unfold_locales) simp_all
-
-lemma (in t_design) bibdI [intro]: "\<t> = 2 \<Longrightarrow> bibd \<V> \<B> \<k> \<Lambda>\<^sub>t"
- using t_lt_order block_size_t by (unfold_locales) (simp_all)
-
-context bibd
-begin
-
-abbreviation "\<r> \<equiv> (\<Lambda> * (\<v> - 1) div (\<k> - 1))"
-
-lemma necessary_condition_one:
- shows "\<r> * (\<k> - 1) = \<Lambda> * (\<v> - 1)"
- using necess_cond_1_rhs necess_cond_1_lhs design_points_nempty rep_number by auto
-
-lemma bibd_point_occ_rep:
- assumes "x \<in> bl"
- assumes "bl \<in># \<B>"
- shows "(\<B> - {#bl#}) rep x = \<r> - 1"
-proof -
- have xin: "x \<in> \<V>" using assms wf_invalid_point by blast
- then have rep: "size {# blk \<in># \<B>. x \<in> blk #} = \<r>" using rep_number_unfold_set by simp
- have "(\<B> - {#bl#}) rep x = size {# blk \<in># (\<B> - {#bl#}). x \<in> blk #}"
- by (simp add: point_replication_number_def)
- then have "(\<B> - {#bl#}) rep x = size {# blk \<in># \<B>. x \<in> blk #} - 1"
- by (simp add: assms size_Diff_singleton)
- then show ?thesis using assms rep r_gzero by simp
-qed
-
-lemma necess_cond_2_lhs: "size {# x \<in># (mset_set \<V> \<times># \<B>) . (fst x) \<in> (snd x) #} = \<v> * \<r>"
-proof -
- let ?M = "mset_set \<V>"
- have "\<And> p . p \<in># ?M \<Longrightarrow> size ({# bl \<in># \<B> . p \<in> bl #}) = nat (\<r>)"
- using finite_sets rep_number_unfold_set r_gzero nat_eq_iff2 by auto
- then have "size (\<Sum>p\<in>#?M. ({#p#} \<times># {#bl \<in># \<B>. p \<in> bl#})) = size ?M * nat (\<r>)"
- using mset_points_distinct size_Union_distinct_cart_prod_filter by blast
- thus ?thesis using r_gzero
- by (simp add: set_break_down_left)
-qed
-
-lemma necess_cond_2_rhs: "size {# x \<in># (mset_set \<V> \<times># \<B>) . (fst x) \<in> (snd x) #} = \<b>*\<k>"
- (is "size {# x \<in># (?M \<times># ?B). (fst x) \<in> (snd x) #} = \<b>*\<k>")
-proof -
- have "\<And> bl . bl \<in># ?B \<Longrightarrow> size ({# p \<in># ?M . p \<in> bl #}) = nat \<k>"
- using uniform k_non_zero uniform_unfold_point_set_mset by fastforce
- then have "size (\<Sum>bl\<in>#?B. ( {# p \<in># ?M . p \<in> bl #} \<times># {#bl#})) = size (?B) * (nat \<k>)"
- using mset_points_distinct size_Union_distinct_cart_prod_filter2 by blast
- thus ?thesis using k_non_zero by (simp add: set_break_down_right)
-qed
-
-lemma necessary_condition_two:
- shows "\<v> * \<r> = \<b> * \<k>"
- using necess_cond_2_lhs necess_cond_2_rhs by simp
-
-theorem admissability_conditions:
-"\<r> * (\<k> - 1) = \<Lambda> * (\<v> - 1)"
-"\<v> * \<r> = \<b> * \<k>"
- using necessary_condition_one necessary_condition_two by auto
-
-subsubsection \<open>BIBD Param Relationships\<close>
-
-lemma bibd_block_number: "\<b> = \<Lambda> * \<v> * (\<v> - 1) div (\<k> * (\<k>-1))"
-proof -
- have "\<b> * \<k> = (\<v> * \<r>)" using necessary_condition_two by simp
- then have k_dvd: "\<k> dvd (\<v> * \<r>)" by (metis dvd_triv_right)
- then have "\<b> = (\<v> * \<r>) div \<k>" using necessary_condition_two min_block_size_2 by auto
- then have "\<b> = (\<v> * ((\<Lambda> * (\<v> - 1) div (\<k> - 1)))) div \<k>" by simp
- then have "\<b> = (\<v> * \<Lambda> * (\<v> - 1)) div ((\<k> - 1)* \<k>)" using necessary_condition_one
- necessary_condition_two dvd_div_div_eq_mult dvd_div_eq_0_iff
- by (smt (z3) dvd_triv_right mult.assoc mult.commute mult.left_commute mult_eq_0_iff )
- then show ?thesis by (simp add: mult.commute)
-qed
-
-lemma symmetric_condition_1: "\<Lambda> * (\<v> - 1) = \<k> * (\<k> - 1) \<Longrightarrow> \<b> = \<v> \<and> \<r> = \<k>"
- using b_non_zero bibd_block_number mult_eq_0_iff necessary_condition_two necessary_condition_one
- by auto
-
-lemma index_lt_replication: "\<Lambda> < \<r>"
-proof -
- have 1: "\<r> * (\<k> - 1) = \<Lambda> * (\<v> - 1)" using admissability_conditions by simp
- have lhsnot0: "\<r> * (\<k> - 1) \<noteq> 0"
- using no_zero_divisors rep_not_zero zdiv_eq_0_iff by blast
- then have rhsnot0: "\<Lambda> * (\<v> - 1) \<noteq> 0" using 1 by simp
- have "\<k> - 1 < \<v> - 1" using incomplete b_non_zero bibd_block_number not_less_eq by fastforce
- thus ?thesis using 1 lhsnot0 rhsnot0
- by (smt (verit, best) k_non_zero mult_le_less_imp_less r_gzero)
-qed
-
-lemma index_not_zero: "\<Lambda> \<ge> 1"
- using index_ge_zero index_lt_replication int_one_le_iff_zero_less by fastforce
-
-lemma r_ge_two: "\<r> \<ge> 2"
- using index_lt_replication index_not_zero by linarith
-
-lemma block_num_gt_rep: "\<b> > \<r>"
-proof -
- have fact: "\<b> * \<k> = \<v> * \<r>" using admissability_conditions by auto
- have lhsnot0: "\<b> * \<k> \<noteq> 0" using k_non_zero b_non_zero by auto
- then have rhsnot0: "\<v> * \<r> \<noteq> 0" using fact by simp
- then show ?thesis using incomplete lhsnot0
- using complement_rep_number constant_rep_design.r_gzero incomplete_imp_incomp_block by fastforce
-qed
-
-lemma bibd_subset_occ:
- assumes "x \<subseteq> bl" and "bl \<in># \<B>" and "card x = 2"
- shows "size {# blk \<in># (\<B> - {#bl#}). x \<subseteq> blk #} = \<Lambda> - 1"
-proof -
- have index: "size {# blk \<in># \<B>. x \<subseteq> blk #} = \<Lambda>" using points_index_def balanced assms
- by (metis (full_types) of_nat_numeral subset_eq wf_invalid_point)
- then have "size {# blk \<in># (\<B> - {#bl#}). x \<subseteq> blk #} = size {# blk \<in># \<B>. x \<subseteq> blk #} - 1"
- by (simp add: assms size_Diff_singleton)
- then show ?thesis using assms index_not_zero index by simp
-qed
-
-lemma necess_cond_one_param_balance: "\<b> > \<v> \<Longrightarrow> \<r> > \<k>"
- using necessary_condition_two
- by (smt mult_nonneg_nonneg nonzero_mult_div_cancel_right of_nat_0_le_iff r_gzero zdiv_mono2)
-
-subsection \<open>Constructing New bibd's\<close>
-text \<open>There are many constructions on bibd's to establish new bibds (or other types of designs).
-This section demonstrates this using both existing constructions, and by defining new constructions.\<close>
-subsubsection \<open>BIBD Complement, Multiple, Combine\<close>
-
-lemma comp_params_index_pair:
- assumes "{x, y} \<subseteq> \<V>"
- assumes "x \<noteq> y"
- shows "\<B>\<^sup>C index {x, y} = \<b> + \<Lambda> - 2*\<r>"
-proof -
- have xin: "x \<in> \<V>" and yin: "y \<in> \<V>" using assms by auto
- have ge: "2*\<r> \<ge> \<Lambda>" using index_lt_replication
- using r_gzero by linarith
- have "size {# b \<in># \<B> . x \<in> b \<and> y \<in> b#} = \<Lambda>" using points_index_pair assms by simp
- then have lambda: "size {# b \<in># \<B> . x \<in> b \<and> y \<in> b#} = nat \<Lambda>"
- using index_ge_zero by auto
- have "\<B>\<^sup>C index {x, y} = size {# b \<in># \<B> . x \<notin> b \<and> y \<notin> b #}"
- using complement_index_2 assms by simp
- also have "\<dots> = size \<B> - (size {# b \<in># \<B> . \<not> (x \<notin> b \<and> y \<notin> b) #})"
- using size_filter_neg by blast
- also have "... = size \<B> - (size {# b \<in># \<B> . x \<in> b \<or> y \<in> b#})" by auto
- also have "... = \<b> - (size {# b \<in># \<B> . x \<in> b \<or> y \<in> b#})" by (simp add: of_nat_diff)
- also have "... = \<b> - (size {# b \<in># \<B> . x \<in> b#} +
- size {# b \<in># \<B> . y \<in> b#} - size {# b \<in># \<B> . x \<in> b \<and> y \<in> b#})"
- by (simp add: mset_size_partition_dep)
- also have "... = \<b> - (nat \<r> + nat \<r> - nat (\<Lambda>))" using rep_number_unfold_set lambda xin yin
- by (metis (no_types, lifting) nat_int)
- finally have "\<B>\<^sup>C index {x, y} = \<b> - (2*\<r> - \<Lambda>)"
- using index_ge_zero index_lt_replication by linarith
- thus ?thesis using ge diff_diff_right by simp
-qed
-
-lemma complement_bibd_index:
- assumes "ps \<subseteq> \<V>"
- assumes "card ps = 2"
- shows "\<B>\<^sup>C index ps = \<b> + \<Lambda> - 2*\<r>"
-proof -
- obtain x y where set: "ps = {x, y}" using b_non_zero bibd_block_number diff_is_0_eq incomplete
- mult_0_right nat_less_le design_points_nempty assms by (metis card_2_iff)
- then have "x \<noteq> y" using assms by auto
- thus ?thesis using comp_params_index_pair assms
- by (simp add: set)
-qed
-
-lemma complement_bibd:
- assumes "\<k> \<le> \<v> - 2"
- shows "bibd \<V> \<B>\<^sup>C (\<v> - \<k>) (\<b> + \<Lambda> - 2*\<r>)"
-proof -
- interpret des: incomplete_design \<V> "\<B>\<^sup>C" "(\<v> - \<k>)"
- using assms complement_incomplete by blast
- show ?thesis proof (unfold_locales, simp_all)
- show "2 \<le> des.\<v>" using assms block_size_t by linarith
- show "\<And>ps. ps \<subseteq> \<V> \<Longrightarrow> card ps = 2 \<Longrightarrow>
- \<B>\<^sup>C index ps = \<b> + \<Lambda> - 2 * (\<Lambda> * (des.\<v> - 1) div (\<k> - 1))"
- using complement_bibd_index by simp
- show "2 \<le> des.\<v> - \<k>" using assms block_size_t by linarith
- qed
-qed
-
-lemma multiple_bibd: "n > 0 \<Longrightarrow> bibd \<V> (multiple_blocks n) \<k> (\<Lambda> * n)"
- using multiple_t_design by (simp add: bibd_def)
-
-end
-
-locale two_bibd_eq_points = two_t_designs_eq_points \<V> \<B> \<k> \<B>' 2 \<Lambda> \<Lambda>'
- + des1: bibd \<V> \<B> \<k> \<Lambda> + des2: bibd \<V> \<B>' \<k> \<Lambda>' for \<V> \<B> \<k> \<B>' \<Lambda> \<Lambda>'
-begin
-
-lemma combine_is_bibd: "bibd \<V>\<^sup>+ \<B>\<^sup>+ \<k> (\<Lambda> + \<Lambda>')"
- by (unfold_locales)
-
-sublocale combine_bibd: bibd "\<V>\<^sup>+" "\<B>\<^sup>+" "\<k>" "(\<Lambda> + \<Lambda>')"
- by (unfold_locales)
-
-end
-
-subsubsection \<open>Derived Designs\<close>
-text \<open>A derived bibd takes a block from a valid bibd as the new point sets, and the intersection
-of that block with other blocks as it's block set\<close>
-
-locale bibd_block_transformations = bibd +
- fixes block :: "'a set" ("bl")
- assumes valid_block: "bl \<in># \<B>"
-begin
-
-definition derived_blocks :: "'a set multiset" ("(\<B>\<^sup>D)") where
-"\<B>\<^sup>D \<equiv> {# bl \<inter> b . b \<in># (\<B> - {#bl#}) #}"
-
-lemma derive_define_flip: "{# b \<inter> bl . b \<in># (\<B> - {#bl#}) #} = \<B>\<^sup>D"
- by (simp add: derived_blocks_def inf_sup_aci(1))
-
-lemma derived_points_order: "card bl = \<k>"
- using uniform valid_block by simp
-
-lemma derived_block_num: "bl \<in># \<B> \<Longrightarrow> size \<B>\<^sup>D = \<b> - 1"
- apply (simp add: derived_blocks_def size_remove1_mset_If valid_block)
- using valid_block int_ops(6) by fastforce
-
-lemma derived_is_wellformed: "b \<in># \<B>\<^sup>D \<Longrightarrow> b \<subseteq> bl"
- by (simp add: derived_blocks_def valid_block) (auto)
-
-lemma derived_point_subset_orig: "ps \<subseteq> bl \<Longrightarrow> ps \<subset> \<V>"
- by (simp add: valid_block incomplete_imp_proper_subset subset_psubset_trans)
-
-lemma derived_obtain_orig_block:
- assumes "b \<in># \<B>\<^sup>D"
- obtains b2 where "b = b2 \<inter> bl" and "b2 \<in># remove1_mset bl \<B>"
- using assms derived_blocks_def by auto
-
-sublocale derived_incidence_sys: incidence_system "bl" "\<B>\<^sup>D"
- using derived_is_wellformed valid_block by (unfold_locales) (auto)
-
-sublocale derived_fin_incidence_system: finite_incidence_system "bl" "\<B>\<^sup>D"
- using valid_block finite_blocks by (unfold_locales) simp_all
-
-lemma derived_blocks_nempty:
- assumes "\<And> b .b \<in># remove1_mset bl \<B> \<Longrightarrow> bl |\<inter>| b > 0"
- assumes "bld \<in># \<B>\<^sup>D"
- shows "bld \<noteq> {}"
-proof -
- obtain bl2 where inter: "bld = bl2 \<inter> bl" and member: "bl2 \<in># remove1_mset bl \<B>"
- using assms derived_obtain_orig_block by blast
- then have "bl |\<inter>| bl2 > 0" using assms(1) by blast
- thus ?thesis using intersection_number_empty_iff finite_blocks valid_block
- by (metis Int_commute dual_order.irrefl inter)
-qed
-
-lemma derived_is_design:
- assumes "\<And> b. b \<in># remove1_mset bl \<B> \<Longrightarrow> bl |\<inter>| b > 0"
- shows "design bl \<B>\<^sup>D"
-proof -
- interpret fin: finite_incidence_system "bl" "\<B>\<^sup>D"
- by (unfold_locales)
- show ?thesis using assms derived_blocks_nempty by (unfold_locales) simp
-qed
-
-lemma derived_is_proper:
- assumes "\<And> b. b \<in># remove1_mset bl \<B> \<Longrightarrow> bl |\<inter>| b > 0"
- shows "proper_design bl \<B>\<^sup>D"
-proof -
- interpret des: design "bl" "\<B>\<^sup>D"
- using derived_is_design assms by fastforce
- have "\<b> - 1 > 1" using block_num_gt_rep r_ge_two by linarith
- then show ?thesis by (unfold_locales) (simp add: derived_block_num valid_block)
-qed
-
-
-subsubsection \<open>Residual Designs\<close>
-text \<open>Similar to derived designs, a residual design takes the complement of a block bl as it's new
-point set, and the complement of all other blocks with respect to bl.\<close>
-
-definition residual_blocks :: "'a set multiset" ("(\<B>\<^sup>R)") where
-"\<B>\<^sup>R \<equiv> {# b - bl . b \<in># (\<B> - {#bl#}) #}"
-
-lemma residual_order: "card (bl\<^sup>c) = \<v> - \<k>"
- apply (simp add: valid_block wellformed block_complement_size)
- using block_size_lt_v derived_points_order by auto
-
-lemma residual_block_num: "size (\<B>\<^sup>R) = \<b> - 1"
- using b_positive by (simp add: residual_blocks_def size_remove1_mset_If valid_block int_ops(6))
-
-lemma residual_obtain_orig_block:
- assumes "b \<in># \<B>\<^sup>R"
- obtains bl2 where "b = bl2 - bl" and "bl2 \<in># remove1_mset bl \<B>"
- using assms residual_blocks_def by auto
-
-lemma residual_blocks_ss: assumes "b \<in># \<B>\<^sup>R" shows "b \<subseteq> \<V>"
-proof -
- have "b \<subseteq> (bl\<^sup>c)" using residual_obtain_orig_block
- by (metis Diff_mono assms block_complement_def in_diffD order_refl wellformed)
- thus ?thesis
- using block_complement_subset_points by auto
-qed
-
-lemma residual_blocks_exclude: "b \<in># \<B>\<^sup>R \<Longrightarrow> x \<in> b \<Longrightarrow> x \<notin> bl"
- using residual_obtain_orig_block by auto
-
-lemma residual_is_wellformed: "b \<in># \<B>\<^sup>R \<Longrightarrow> b \<subseteq> (bl\<^sup>c)"
- apply (auto simp add: residual_blocks_def)
- by (metis DiffI block_complement_def in_diffD wf_invalid_point)
-
-sublocale residual_incidence_sys: incidence_system "bl\<^sup>c" "\<B>\<^sup>R"
- using residual_is_wellformed by (unfold_locales)
-
-lemma residual_is_finite: "finite (bl\<^sup>c)"
- by (simp add: block_complement_def finite_sets)
-
-sublocale residual_fin_incidence_sys: finite_incidence_system "bl\<^sup>c" "\<B>\<^sup>R"
- using residual_is_finite by (unfold_locales)
-
-lemma residual_blocks_nempty:
- assumes "bld \<in># \<B>\<^sup>R"
- assumes "multiplicity bl = 1"
- shows "bld \<noteq> {}"
-proof -
- obtain bl2 where inter: "bld = bl2 - bl" and member: "bl2 \<in># remove1_mset bl \<B>"
- using assms residual_blocks_def by auto
- then have ne: "bl2 \<noteq> bl" using assms
- by (metis count_eq_zero_iff in_diff_count less_one union_single_eq_member)
- have "card bl2 = card bl" using uniform valid_block member
- by (metis in_diffD of_nat_eq_iff)
- then have "card (bl2 - bl) > 0"
- using finite_blocks member uniform set_card_diff_ge_zero valid_block by (metis in_diffD ne)
- thus ?thesis using inter by fastforce
-qed
-
-lemma residual_is_design: "multiplicity bl = 1 \<Longrightarrow> design (bl\<^sup>c) \<B>\<^sup>R"
- using residual_blocks_nempty by (unfold_locales)
-
-lemma residual_is_proper:
- assumes "multiplicity bl = 1"
- shows "proper_design (bl\<^sup>c) \<B>\<^sup>R"
-proof -
- interpret des: design "bl\<^sup>c" "\<B>\<^sup>R" using residual_is_design assms by blast
- have "\<b> - 1 > 1" using r_ge_two block_num_gt_rep by linarith
- then show ?thesis using residual_block_num by (unfold_locales) auto
-qed
-
-end
-
-subsection \<open>Symmetric BIBD's\<close>
-text \<open>Symmetric bibd's are those where the order of the design equals the number of blocks\<close>
-
-locale symmetric_bibd = bibd +
- assumes symmetric: "\<b> = \<v>"
-begin
-
-lemma rep_value_sym: "\<r> = \<k>"
- using b_non_zero local.symmetric necessary_condition_two by auto
-
-lemma symmetric_condition_2: "\<Lambda> * (\<v> - 1) = \<k> * (\<k> - 1)"
- using necessary_condition_one rep_value_sym by auto
-
-lemma sym_design_vk_gt_kl:
- assumes "\<k> \<ge> \<Lambda> + 2"
- shows "\<v> - \<k> > \<k> - \<Lambda>"
-proof (rule ccontr)
- assume "\<not> (\<v> - \<k> > \<k> - \<Lambda>)"
- then have "\<v> \<le> 2 * \<k> - \<Lambda>"
- by linarith
- then have "\<v> - 1 \<le> 2 * \<k> - \<Lambda> - 1" by linarith
- then have "\<Lambda>* (\<v> - 1) \<le> \<Lambda>*( 2 * \<k> - \<Lambda> - 1)"
- using index_ge_zero mult_le_cancel_left by fastforce
- then have "\<k> * (\<k> - 1) \<le> \<Lambda>*( 2 * \<k> - \<Lambda> - 1)"
- by (simp add: symmetric_condition_2)
- then have "\<k> * (\<k> - 1) - \<Lambda>*( 2 * \<k> - \<Lambda> - 1) \<le> 0" by linarith
- then have "(\<k> - \<Lambda>)*(\<k> - \<Lambda> - 1) \<le> 0"
- by (simp add: mult.commute right_diff_distrib')
- then have "\<k> = \<Lambda> \<or> \<k> = \<Lambda> + 1"
- using mult_le_0_iff by force
- thus False using assms
- by simp
-qed
-
-end
-
-context bibd
-begin
-
-lemma symmetric_bibdI: "\<b> = \<v> \<Longrightarrow> symmetric_bibd \<V> \<B> \<k> \<Lambda>"
- by unfold_locales simp
-
-lemma symmetric_bibdII: "\<Lambda> * (\<v> - 1) = \<k> * (\<k> - 1) \<Longrightarrow> symmetric_bibd \<V> \<B> \<k> \<Lambda>"
- using symmetric_condition_1 by unfold_locales blast
-
-lemma symmetric_not_admissable: "\<Lambda> * (\<v> - 1) \<noteq> \<k> * (\<k> - 1) \<Longrightarrow> \<not> symmetric_bibd \<V> \<B> \<k> \<Lambda>"
- using symmetric_bibd.symmetric_condition_2 by blast
-end
-
-context symmetric_bibd
-begin
-
-subsubsection \<open>Intersection Property on Symmetric BIBDs\<close>
-text \<open>Below is a proof of an important property on symmetric BIBD's regarding the equivalence
-of intersection numbers and the design index. This is an intuitive counting proof, and involved
-significantly more work in a formal environment. Based of Lecture Note \cite{HerkeLectureNotes2016}\<close>
-
-lemma intersect_mult_set_eq_block:
- assumes "blv \<in># \<B>"
- shows "p \<in># \<Sum>\<^sub>#{# mset_set (bl \<inter> blv) .bl \<in># (\<B> - {#blv#})#} \<longleftrightarrow> p \<in> blv"
-proof (auto, simp add: assms finite_blocks)
- assume assm: "p \<in> blv"
- then have "(\<B> - {#blv#}) rep p > 0" using bibd_point_occ_rep r_ge_two assms by auto
- then obtain bl where "bl \<in># (\<B> - {#blv#}) \<and> p \<in> bl" using assms rep_number_g0_exists by metis
- then show "\<exists>x\<in>#remove1_mset blv \<B>. p \<in># mset_set (x \<inter> blv)"
- using assms assm finite_blocks by auto
-qed
-
-lemma intersect_mult_set_block_subset_iff:
- assumes "blv \<in># \<B>"
- assumes "p \<in># \<Sum>\<^sub>#{# mset_set {y .y \<subseteq> blv \<inter> b2 \<and> card y = 2} .b2 \<in># (\<B> - {#blv#})#}"
- shows "p \<subseteq> blv"
-proof (rule subsetI)
- fix x
- assume asm: "x \<in> p"
- obtain b2 where "p \<in># mset_set {y . y \<subseteq> blv \<inter> b2 \<and> card y = 2} \<and> b2 \<in>#(\<B> - {#blv#})"
- using assms by blast
- then have "p \<subseteq> blv \<inter> b2"
- by (metis (no_types, lifting) elem_mset_set equals0D infinite_set_mset_mset_set mem_Collect_eq)
- thus "x \<in> blv" using asm by auto
-qed
-
-lemma intersect_mult_set_block_subset_card:
- assumes "blv \<in># \<B>"
- assumes "p \<in># \<Sum>\<^sub>#{# mset_set {y .y \<subseteq> blv \<inter> b2 \<and> card y = 2} .b2 \<in># (\<B> - {#blv#})#}"
- shows "card p = 2"
-proof -
- obtain b2 where "p \<in># mset_set {y . y \<subseteq> blv \<inter> b2 \<and> card y = 2} \<and> b2 \<in>#(\<B> - {#blv#})"
- using assms by blast
- thus ?thesis
- by (metis (mono_tags, lifting) elem_mset_set equals0D infinite_set_mset_mset_set mem_Collect_eq)
-qed
-
-lemma intersect_mult_set_block_with_point_exists:
- assumes "blv \<in># \<B>" and "p \<subseteq> blv" and "\<Lambda> \<ge> 2" and "card p = 2"
- shows "\<exists>x\<in>#remove1_mset blv \<B>. p \<in># mset_set {y. y \<subseteq> blv \<and> y \<subseteq> x \<and> card y = 2}"
-proof -
- have "size {#b \<in># \<B> . p \<subseteq> b#} = \<Lambda>" using points_index_def assms
- by (metis balanced_alt_def_all dual_order.trans of_nat_numeral wellformed)
- then have "size {#bl \<in># (\<B> - {#blv#}) . p \<subseteq> bl#} \<ge> 1"
- using assms by (simp add: size_Diff_singleton)
- then obtain bl where "bl \<in># (\<B> - {#blv#}) \<and> p \<subseteq> bl" using assms filter_mset_empty_conv
- by (metis diff_diff_cancel diff_is_0_eq' le_numeral_extra(4) size_empty zero_neq_one)
- thus ?thesis
- using assms finite_blocks by auto
-qed
-
-lemma intersect_mult_set_block_subset_iff_2:
- assumes "blv \<in># \<B>" and "p \<subseteq> blv" and "\<Lambda> \<ge> 2" and "card p = 2"
- shows "p \<in># \<Sum>\<^sub>#{# mset_set {y .y \<subseteq> blv \<inter> b2 \<and> card y = 2} .b2 \<in># (\<B> - {#blv#})#}"
- by (auto simp add: intersect_mult_set_block_with_point_exists assms)
-
-lemma sym_sum_mset_inter_sets_count:
- assumes "blv \<in># \<B>"
- assumes "p \<in> blv"
- shows "count (\<Sum>\<^sub>#{# mset_set (bl \<inter> blv) .bl \<in># (\<B> - {#blv#})#}) p = \<r> - 1"
- (is "count (\<Sum>\<^sub>#?M) p = \<r> - 1")
-proof -
- have size_inter: "size {# mset_set (bl \<inter> blv) | bl \<in># (\<B> - {#blv#}) . p \<in> bl#} = \<r> - 1"
- using bibd_point_occ_rep point_replication_number_def
- by (metis assms(1) assms(2) size_image_mset)
- have inter_finite: "\<forall> bl \<in># (\<B> - {#blv#}) . finite (bl \<inter> blv)"
- by (simp add: assms(1) finite_blocks)
- have "\<And> bl . bl \<in># (\<B> - {#blv#}) \<Longrightarrow> p \<in> bl \<longrightarrow> count (mset_set (bl \<inter> blv)) p = 1"
- using assms count_mset_set(1) inter_finite by simp
- then have "\<And> bl . bl \<in># {#b1 \<in>#(\<B> - {#blv#}) . p \<in> b1#} \<Longrightarrow> count (mset_set (bl \<inter> blv)) p = 1"
- by (metis (full_types) count_eq_zero_iff count_filter_mset)
- then have pin: "\<And> P. P \<in># {# mset_set (bl \<inter> blv) | bl \<in># (\<B> - {#blv#}) . p \<in> bl#}
- \<Longrightarrow> count P p = 1" by blast
- have "?M = {# mset_set (bl \<inter> blv) | bl \<in># (\<B> - {#blv#}) . p \<in> bl#}
- + {# mset_set (bl \<inter> blv) | bl \<in># (\<B> - {#blv#}) . p \<notin> bl#}"
- by (metis image_mset_union multiset_partition)
- then have "count (\<Sum>\<^sub>#?M) p = size {# mset_set (bl \<inter> blv) | bl \<in># (\<B> - {#blv#}) . p \<in> bl#} "
- using pin by (auto simp add: count_sum_mset)
- then show ?thesis using size_inter by linarith
-qed
-
-lemma sym_sum_mset_inter_sets_size:
- assumes "blv \<in># \<B>"
- shows "size (\<Sum>\<^sub>#{# mset_set (bl \<inter> blv) .bl \<in># (\<B> - {#blv#})#}) = \<k> * (\<r> - 1)"
- (is "size (\<Sum>\<^sub>#?M) = \<k>* (\<r> - 1)")
-proof -
- have eq: "set_mset (\<Sum>\<^sub>#{# mset_set (bl \<inter> blv) .bl \<in># (\<B> - {#blv#})#}) = blv"
- using intersect_mult_set_eq_block assms by auto
- then have k: "card (set_mset (\<Sum>\<^sub>#?M)) = \<k>"
- by (simp add: assms)
- have "\<And> p. p \<in># (\<Sum>\<^sub>#?M) \<Longrightarrow> count (\<Sum>\<^sub>#?M) p = \<r> - 1"
- using sym_sum_mset_inter_sets_count assms eq by blast
- thus ?thesis using k size_multiset_int_count by metis
-qed
-
-lemma sym_sum_inter_num:
- assumes "b1 \<in># \<B>"
- shows "(\<Sum>b2 \<in>#(\<B> - {#b1#}). b1 |\<inter>| b2) = \<k>* (\<r> - 1)"
-proof -
- have "(\<Sum>b2 \<in>#(\<B> - {#b1#}). b1 |\<inter>| b2) = (\<Sum>b2 \<in>#(\<B> - {#b1#}). size (mset_set (b1 \<inter> b2)))"
- by (simp add: intersection_number_def)
- also have "... = size (\<Sum>\<^sub>#{#mset_set (b1 \<inter> bl). bl \<in># (\<B> - {#b1#})#})"
- by (auto simp add: size_big_union_sum)
- also have "... = size (\<Sum>\<^sub>#{#mset_set (bl \<inter> b1). bl \<in># (\<B> - {#b1#})#})"
- by (metis Int_commute)
- finally have "(\<Sum>b2 \<in>#(\<B> - {#b1#}). b1 |\<inter>| b2) = \<k> * (\<r> - 1)"
- using sym_sum_mset_inter_sets_size assms by auto
- then show ?thesis by simp
-qed
-
-lemma choose_two_int:
- assumes " x \<ge> 0"
- shows "nat (x :: int) choose 2 = ((x ::int) * ( x - 1)) div 2 "
- using choose_two assms dvd_div_mult_self even_numeral int_nat_eq mult_cancel_right2 mult_eq_0_iff
- mult_nonneg_nonneg nat_diff_distrib' nat_mult_distrib nat_one_as_int
- numeral_Bit0_div_2 numerals(1) of_nat_numeral zdiv_int by (smt (verit)) (* Slow *)
-
-lemma sym_sum_mset_inter2_sets_count:
- assumes "blv \<in># \<B>"
- assumes "p \<subseteq> blv"
- assumes "card p = 2"
- shows "count (\<Sum>\<^sub>#{#mset_set {y .y \<subseteq> blv \<inter> b2 \<and> card y = 2}. b2 \<in># (\<B> - {#blv#})#}) p = \<Lambda> - 1"
- (is "count (\<Sum>\<^sub>#?M) p = \<Lambda> - 1")
-proof -
- have size_inter: "size {# mset_set {y .y \<subseteq> blv \<inter> b2 \<and> card y = 2} | b2 \<in># (\<B> - {#blv#}) . p \<subseteq> b2#}
- = \<Lambda> - 1"
- using bibd_subset_occ assms by simp
- have "\<forall> b2 \<in># (\<B> - {#blv#}) . p \<subseteq> b2 \<longrightarrow> count (mset_set{y .y \<subseteq> blv \<inter> b2 \<and> card y = 2}) p = 1"
- using assms(2) count_mset_set(1) assms(3) by (auto simp add: assms(1) finite_blocks)
- then have "\<forall> bl \<in># {#b1 \<in>#(\<B> - {#blv#}) . p \<subseteq> b1#}.
- count (mset_set {y .y \<subseteq> blv \<inter> bl \<and> card y = 2}) p = 1"
- using count_eq_zero_iff count_filter_mset by (metis (no_types, lifting))
- then have pin: "\<forall> P \<in># {# mset_set {y .y \<subseteq> blv \<inter> b2 \<and> card y = 2} | b2 \<in># (\<B> - {#blv#}) . p \<subseteq> b2#}.
- count P p = 1"
- using count_eq_zero_iff count_filter_mset by blast
- have "?M = {# mset_set {y .y \<subseteq> blv \<inter> b2 \<and> card y = 2} | b2 \<in># (\<B> - {#blv#}) . p \<subseteq> b2#} +
- {# mset_set {y .y \<subseteq> blv \<inter> b2 \<and> card y = 2} | b2 \<in># (\<B> - {#blv#}) . \<not> (p \<subseteq> b2)#}"
- by (metis image_mset_union multiset_partition)
- then have "count (\<Sum>\<^sub>#?M) p =
- size {# mset_set {y .y \<subseteq> blv \<inter> b2 \<and> card y = 2} | b2 \<in># (\<B> - {#blv#}) . p \<subseteq> b2#}"
- using pin by (auto simp add: count_sum_mset)
- then show ?thesis using size_inter by linarith
-qed
-
-lemma sym_sum_mset_inter2_sets_size:
- assumes "blv \<in># \<B>"
- shows "size (\<Sum>\<^sub>#{#mset_set {y .y \<subseteq> blv \<inter> b2 \<and> card y = 2}. b2 \<in># (\<B> - {#blv#})#}) =
- ((nat \<k>) choose 2) * (\<Lambda> -1)"
- (is "size (\<Sum>\<^sub>#?M) = ((nat \<k>) choose 2) * (\<Lambda> -1)")
-proof (cases "\<Lambda> = 1")
- case True
- have empty: "\<And> b2 . b2 \<in># remove1_mset blv \<B> \<Longrightarrow> {y .y \<subseteq> blv \<and> y \<subseteq> b2 \<and> card y = 2} = {}"
- using index_one_alt_bl_not_exist assms True by blast
- then show ?thesis using sum_mset.neutral True by (simp add: empty)
-next
- case False
- then have index_min: "\<Lambda> \<ge> 2" using index_not_zero by linarith
- have subset_card: "\<And> x . x \<in># (\<Sum>\<^sub>#?M) \<Longrightarrow> card x = 2"
- proof -
- fix x
- assume a: "x \<in># (\<Sum>\<^sub>#?M)"
- then obtain b2 where "x \<in># mset_set {y . y \<subseteq> blv \<inter> b2 \<and> card y = 2} \<and> b2 \<in>#(\<B> - {#blv#})"
- by blast
- thus "card x = 2" using mem_Collect_eq
- by (metis (mono_tags, lifting) elem_mset_set equals0D infinite_set_mset_mset_set)
- qed
- have eq: "set_mset (\<Sum>\<^sub>#?M) = {bl . bl \<subseteq> blv \<and> card bl = 2}"
- proof
- show "set_mset (\<Sum>\<^sub>#?M) \<subseteq> {bl . bl \<subseteq> blv \<and> card bl = 2}"
- using subset_card intersect_mult_set_block_subset_iff assms by blast
- show "{bl . bl \<subseteq> blv \<and> card bl = 2} \<subseteq> set_mset (\<Sum>\<^sub>#?M)"
- using intersect_mult_set_block_subset_iff_2 assms index_min by blast
- qed
- have "card blv = (nat \<k>)" using uniform assms by (metis nat_int)
- then have k: "card (set_mset (\<Sum>\<^sub>#?M)) = ((nat \<k>) choose 2)" using eq n_subsets
- by (simp add: n_subsets assms finite_blocks)
- thus ?thesis using k size_multiset_int_count sym_sum_mset_inter2_sets_count assms eq subset_card
- by (metis (no_types, lifting) intersect_mult_set_block_subset_iff)
-qed
-
-lemma sum_choose_two_inter_num:
- assumes "b1 \<in># \<B>"
- shows "(\<Sum>b2 \<in># (\<B> - {#b1#}). (nat (b1 |\<inter>| b2) choose 2)) = ((\<Lambda> * (\<Lambda> - 1) div 2)) * (\<v> -1)"
-proof -
- have div_fact: "2 dvd (\<Lambda> * (\<Lambda> - 1))"by simp
- have div_fact_2: "2 dvd (\<Lambda> * (\<v> - 1))" using symmetric_condition_2 by simp
- have "(\<Sum>b2 \<in># (\<B> - {#b1#}). (nat (b1 |\<inter>| b2) choose 2)) = (\<Sum>b2 \<in># (\<B> - {#b1#}). nat (b1 |\<inter>|\<^sub>2 b2 ))"
- using n_inter_num_choose_design_inter assms by (simp add: in_diffD)
- then have sum_fact: "(\<Sum>b2 \<in># (\<B> - {#b1#}).(nat (b1 |\<inter>| b2) choose 2))
- = ((nat \<k>) choose 2) * (\<Lambda> -1)"
- using assms sym_sum_mset_inter2_sets_size
- by (auto simp add: size_big_union_sum n_intersect_num_subset_def)
- have "((nat \<k>) choose 2) * (\<Lambda> -1) = ((\<Lambda> * (\<v> - 1) div 2)) * (\<Lambda> -1)"
- using choose_two_int symmetric_condition_2 k_non_zero by auto
- then have "((nat \<k>) choose 2) * (\<Lambda> -1) = ((\<Lambda> * (\<Lambda> - 1) div 2)) * (\<v> -1)"
- using div_fact div_fact_2 by (smt div_mult_swap mult.assoc mult.commute)
- then show ?thesis using sum_fact by simp
-qed
-
-lemma sym_sum_inter_num_sq:
- assumes "b1 \<in># \<B>"
- shows "(\<Sum>bl \<in># (remove1_mset b1 \<B>). (b1 |\<inter>| bl)^2) = \<Lambda>^2 * ( \<v> - 1)"
-proof -
- have dvd: "2 dvd (( \<v> - 1) * (\<Lambda> * (\<Lambda> - 1)))" by simp
- have a: "(\<Sum>b2 \<in>#(\<B> - {#b1#}). int (nat (b1 |\<inter>| b2) choose 2)) =
- (\<Sum>bl \<in># (remove1_mset b1 \<B>). ((b1 |\<inter>| bl) * ((b1 |\<inter>| bl) - 1)) div 2)"
- using choose_two_int by (simp add: intersection_num_non_neg)
- have b: "(\<Sum>b2 \<in>#(\<B> - {#b1#}). int (nat (b1 |\<inter>| b2) choose 2)) =
- (\<Sum>b2 \<in>#(\<B> - {#b1#}). (nat (b1 |\<inter>| b2) choose 2))" by simp
- have "(\<Sum>b2 \<in>#(\<B> - {#b1#}). (nat (b1 |\<inter>| b2) choose 2)) = ((\<Lambda> * (\<Lambda> - 1)) div 2) * ( \<v> - 1)"
- using sum_choose_two_inter_num assms by blast
- then have start: "(\<Sum>bl \<in># (remove1_mset b1 \<B>). ((b1 |\<inter>| bl) * ((b1 |\<inter>| bl) - 1)) div 2)
- = ((\<Lambda> * (\<Lambda> - 1)) div 2) * (\<v> - 1)"
- using a b by linarith
- have sum_dvd: "2 dvd (\<Sum>bl \<in># (remove1_mset b1 \<B>). (b1 |\<inter>| bl) * ((b1 |\<inter>| bl) - 1))"
- by (simp add: sum_mset_dvd)
- then have "(\<Sum>bl \<in># (remove1_mset b1 \<B>). (b1 |\<inter>| bl) * ((b1 |\<inter>| bl) - 1)) div 2
- = (\<v> - 1) * ((\<Lambda> * (\<Lambda> - 1)) div 2)"
- using start by (simp add: sum_mset_distrib_div_if_dvd)
- then have "(\<Sum>bl \<in># (remove1_mset b1 \<B>). (b1 |\<inter>| bl)^2)
- - (\<Sum>bl \<in># (remove1_mset b1 \<B>). (b1 |\<inter>| bl)) = ((\<v> - 1) * (\<Lambda> * (\<Lambda> - 1)))"
- using sum_dvd dvd
- by (simp add: dvd_div_eq_iff div_mult_swap int_distrib(4) power2_eq_square sum_mset_add_diff)
- then have "(\<Sum>bl \<in># (remove1_mset b1 \<B>). (b1 |\<inter>| bl)^2) - (\<Lambda> * (\<v> - 1)) = ((\<v> - 1) * (\<Lambda> * (\<Lambda> - 1)))"
- using sym_sum_inter_num assms rep_value_sym symmetric_condition_2 by auto
- then have "(\<Sum>bl \<in># (remove1_mset b1 \<B>). (b1 |\<inter>| bl)^2) = (\<Lambda> * (\<v> - 1)) * (\<Lambda> - 1) + (\<Lambda> * (\<v> - 1))"
- using diff_eq_eq by fastforce
- then have "(\<Sum>bl \<in># (remove1_mset b1 \<B>). (b1 |\<inter>| bl)^2) = (\<Lambda> * (\<v> - 1)) * (\<Lambda> - 1 + 1)"
- using int_distrib(2) by (metis mult_numeral_1_right numeral_One)
- thus ?thesis by (simp add: power2_eq_square)
-qed
-
-lemma sym_sum_inter_num_to_zero:
- assumes "b1 \<in># \<B>"
- shows "(\<Sum>bl \<in># (remove1_mset b1 \<B>). ((b1 |\<inter>| bl) - \<Lambda>)^2) = 0"
-proof -
- have rm1_size: "size (remove1_mset b1 \<B>) = \<v> - 1" using assms b_non_zero int_ops(6)
- by (auto simp add: symmetric size_remove1_mset_If)
- have "\<And> bl . bl \<in># (remove1_mset b1 \<B>) \<Longrightarrow> ((b1 |\<inter>| bl) - \<Lambda>)^2 =
- (((b1 |\<inter>| bl)^2) - (2 * \<Lambda> * (b1 |\<inter>| bl)) + (\<Lambda>^2))"
- by (simp add: power2_diff)
- then have "(\<Sum>bl \<in># (remove1_mset b1 \<B>). ((b1 |\<inter>| bl) - \<Lambda>)^2) =
- (\<Sum>bl \<in># (remove1_mset b1 \<B>). (((b1 |\<inter>| bl)^2) - (2 * \<Lambda> * (b1 |\<inter>| bl)) + (\<Lambda>^2)))"
- using sum_over_fun_eq by auto
- also have "... = \<Lambda>^2 * (\<v> - 1) - 2 * \<Lambda> * (\<Sum>bl \<in># (remove1_mset b1 \<B>). ((b1 |\<inter>| bl)))
- + (\<v> - 1) * (\<Lambda>^2)"
- using sym_sum_inter_num_sq rm1_size
- by (simp add: assms sum_mset.distrib sum_mset_add_diff sum_mset_distrib_left)
- finally have "(\<Sum>bl \<in># (remove1_mset b1 \<B>). ((b1 |\<inter>| bl) - \<Lambda>)^2) = 0"
- using rep_value_sym symmetric_condition_2 sym_sum_inter_num assms
- by (auto simp add: power2_eq_square)
- thus ?thesis by simp
-qed
-
-theorem sym_block_intersections_index [simp]:
- assumes "b1 \<in># \<B>"
- assumes "b2 \<in># (\<B> - {#b1#})"
- shows "b1 |\<inter>| b2 = \<Lambda>"
-proof -
- have pos: "\<And> bl . ((b1 |\<inter>| bl) - \<Lambda>)^2 \<ge> 0" by simp
- have "(\<Sum>bl \<in># (remove1_mset b1 \<B>). ((b1 |\<inter>| bl) - \<Lambda>)^2) = 0"
- using sym_sum_inter_num_to_zero assms by simp
- then have "\<And> bl. bl \<in> set_mset (remove1_mset b1 \<B>) \<Longrightarrow> ((b1 |\<inter>| bl) - \<Lambda>)^2 = 0"
- using sum_mset_0_iff_ge_0 pos by (metis (no_types, lifting))
- thus ?thesis
- using assms(2) by auto
-qed
-
-subsubsection \<open>Symmetric BIBD is Simple\<close>
-
-lemma sym_block_mult_one [simp]:
- assumes "bl \<in># \<B>"
- shows "multiplicity bl = 1"
-proof (rule ccontr)
- assume "\<not> (multiplicity bl = 1)"
- then have not: "multiplicity bl \<noteq> 1" by simp
- have "multiplicity bl \<noteq> 0" using assms
- by simp
- then have m: "multiplicity bl \<ge> 2" using not by linarith
- then have blleft: "bl \<in># (\<B> - {#bl#})"
- using in_diff_count by fastforce
- have "bl |\<inter>| bl = \<k>" using k_non_zero assms
- by (simp add: intersection_number_def)
- then have keql: "\<k> = \<Lambda>" using sym_block_intersections_index blleft assms by simp
- then have "\<v> = \<k>"
- using keql index_lt_replication rep_value_sym block_size_lt_v diffs0_imp_equal k_non_zero zero_diff by linarith
- then show False using incomplete
- by simp
-qed
-
-end
-
-sublocale symmetric_bibd \<subseteq> simple_design
- by unfold_locales simp
-
-subsubsection \<open>Residual/Derived Sym BIBD Constructions\<close>
-text \<open>Using the intersect result, we can reason further on residual and derived designs.
-Proofs based off lecture notes \cite{HerkeLectureNotes2016}\<close>
-
-locale symmetric_bibd_block_transformations = symmetric_bibd + bibd_block_transformations
-begin
-
-lemma derived_block_size [simp]:
- assumes "b \<in># \<B>\<^sup>D"
- shows "card b = \<Lambda>"
-proof -
- obtain bl2 where set: "bl2 \<in># remove1_mset bl \<B>" and inter: "b = bl2 \<inter> bl"
- using derived_blocks_def assms by (meson derived_obtain_orig_block)
- then have "card b = bl2 |\<inter>| bl"
- by (simp add: intersection_number_def)
- thus ?thesis using sym_block_intersections_index
- using set intersect_num_commute valid_block by fastforce
-qed
-
-lemma derived_points_index [simp]:
- assumes "ps \<subseteq> bl"
- assumes "card ps = 2"
- shows "\<B>\<^sup>D index ps = \<Lambda> - 1"
-proof -
- have b_in: "\<And> b . b \<in># (remove1_mset bl \<B>) \<Longrightarrow> ps \<subseteq> b \<Longrightarrow> ps \<subseteq> b \<inter> bl"
- using assms by blast
- then have orig: "ps \<subseteq> \<V>"
- using valid_block assms wellformed by blast
- then have lam: "size {# b \<in># \<B> . ps \<subseteq> b #} = \<Lambda>" using balanced
- by (simp add: assms(2) points_index_def)
- then have "size {# b \<in># remove1_mset bl \<B> . ps \<subseteq> b #} = size {# b \<in># \<B> . ps \<subseteq> b #} - 1"
- using assms valid_block by (simp add: size_Diff_submset)
- then have "size {# b \<in># remove1_mset bl \<B> . ps \<subseteq> b #} = \<Lambda> - 1"
- using lam index_not_zero by linarith
- then have "size {# bl \<inter> b | b \<in># (remove1_mset bl \<B>) . ps \<subseteq> bl \<inter> b #} = \<Lambda> - 1"
- using b_in by (metis (no_types, lifting) Int_subset_iff filter_mset_cong size_image_mset)
- then have "size {# x \<in># {# bl \<inter> b . b \<in># (remove1_mset bl \<B>) #} . ps \<subseteq> x #} = \<Lambda> - 1"
- by (metis image_mset_filter_swap)
- then have "size {# x \<in># \<B>\<^sup>D . ps \<subseteq> x #} = \<Lambda> - 1" by (simp add: derived_blocks_def)
- thus ?thesis by (simp add: points_index_def)
-qed
-
-lemma sym_derive_design_bibd:
- assumes "\<Lambda> > 1"
- shows "bibd bl \<B>\<^sup>D \<Lambda> (\<Lambda> - 1)"
-proof -
- interpret des: proper_design bl "\<B>\<^sup>D" using derived_is_proper assms valid_block by auto
- have "\<Lambda> < \<k>" using index_lt_replication rep_value_sym by linarith
- then show ?thesis using derived_block_size assms derived_points_index derived_points_order
- by (unfold_locales) (simp_all)
-qed
-
-lemma residual_block_size [simp]:
- assumes "b \<in># \<B>\<^sup>R"
- shows "card b = \<k> - \<Lambda>"
-proof -
- obtain bl2 where sub: "b = bl2 - bl" and mem: "bl2 \<in># remove1_mset bl \<B>"
- using assms residual_blocks_def by auto
- then have "card b = card bl2 - card (bl2 \<inter> bl)"
- using card_Diff_subset_Int valid_block finite_blocks
- by (simp add: card_Diff_subset_Int)
- then have "card b = card bl2 - bl2 |\<inter>| bl"
- using intersection_number_def finite_blocks card_inter_lt_single
- by (metis assms derived_fin_incidence_system.finite_sets finite_Diff2 of_nat_diff
- residual_fin_incidence_sys.finite_blocks sub)
- thus ?thesis using sym_block_intersections_index uniform
- by (metis valid_block in_diffD intersect_num_commute mem)
-qed
-
-lemma residual_index [simp]:
- assumes "ps \<subseteq> bl\<^sup>c"
- assumes "card ps = 2"
- shows "(\<B>\<^sup>R) index ps = \<Lambda>"
-proof -
- have a: "\<And> b . (b \<in># remove1_mset bl \<B> \<Longrightarrow> ps \<subseteq> b \<Longrightarrow> ps \<subseteq> (b - bl))" using assms
- by (smt DiffI block_comp_elem_alt_left in_diffD subset_eq wellformed)
- have b: "\<And> b . (b \<in># remove1_mset bl \<B> \<Longrightarrow> ps \<subseteq> (b - bl) \<Longrightarrow> ps \<subseteq> b)"
- by auto
- have not_ss: "\<not> (ps \<subseteq> bl)" using set_diff_non_empty_not_subset blocks_nempty t_non_zero assms
- block_complement_def by fastforce
- have "\<B>\<^sup>R index ps = size {# x \<in># {# b - bl . b \<in># (remove1_mset bl \<B>) #} . ps \<subseteq> x #}"
- using assms valid_block by (simp add: points_index_def residual_blocks_def)
- also have "... = size {# b - bl | b \<in># (remove1_mset bl \<B>) . ps \<subseteq> b - bl #} "
- by (metis image_mset_filter_swap)
- finally have "\<B>\<^sup>R index ps = size {# b \<in># (remove1_mset bl \<B>) . ps \<subseteq> b #} " using a b
- by (metis (no_types, lifting) filter_mset_cong size_image_mset)
- thus ?thesis
- using balanced not_ss assms points_index_alt_def block_complement_subset_points by auto
-qed
-
-lemma sym_residual_design_bibd:
- assumes "\<k> \<ge> \<Lambda> + 2"
- shows "bibd (bl\<^sup>c) \<B>\<^sup>R (\<k> - \<Lambda>) \<Lambda>"
-proof -
- interpret des: proper_design "bl\<^sup>c" "\<B>\<^sup>R"
- using residual_is_proper assms(1) valid_block sym_block_mult_one by fastforce
- show ?thesis using residual_block_size assms sym_design_vk_gt_kl residual_order residual_index
- by(unfold_locales) simp_all
-qed
-
-end
-
-subsection \<open>BIBD's and Other Block Designs\<close>
-text \<open>BIBD's are closely related to other block designs by indirect inheritance\<close>
-
-sublocale bibd \<subseteq> k_\<Lambda>_PBD \<V> \<B> \<Lambda> \<k>
- using block_size_gt_t by (unfold_locales) simp_all
-
-lemma incomplete_PBD_is_bibd:
- assumes "k < card V" and "k_\<Lambda>_PBD V B \<Lambda> k"
- shows "bibd V B k \<Lambda>"
-proof -
- interpret inc: incomplete_design V B k using assms
- by (auto simp add: block_design.incomplete_designI k_\<Lambda>_PBD.axioms(2))
- interpret pairwise_balance: pairwise_balance V B \<Lambda> using assms
- by (auto simp add: k_\<Lambda>_PBD.axioms(1))
- show ?thesis using assms k_\<Lambda>_PBD.block_size_t by (unfold_locales) (simp_all)
-qed
-
-lemma (in bibd) bibd_to_pbdI[intro]:
- assumes "\<Lambda> = 1"
- shows "k_PBD \<V> \<B> \<k>"
-proof -
- interpret pbd: k_\<Lambda>_PBD \<V> \<B> \<Lambda> \<k>
- by (simp add: k_\<Lambda>_PBD_axioms)
- show ?thesis using assms by (unfold_locales) (simp_all add: t_lt_order min_block_size_2)
-qed
-
-locale incomplete_PBD = incomplete_design + k_\<Lambda>_PBD
-
-sublocale incomplete_PBD \<subseteq> bibd
- using block_size_t by (unfold_locales) simp
-
+(* Title: BIBD
+ Author: Chelsea Edmonds
+*)
+
+theory BIBD imports Block_Designs
+begin
+
+section \<open>BIBD's\<close>
+text \<open>BIBD's are perhaps the most commonly studied type of design in combinatorial design theory,
+and usually the first type of design explored in a design theory course.
+These designs are a type of t-design, where $t = 2$\<close>
+
+subsection \<open>BIBD Basics\<close>
+locale bibd = t_design \<V> \<B> \<k> 2 \<Lambda>
+ for point_set ("\<V>") and block_collection ("\<B>")
+ and u_block_size ("\<k>") and index ("\<Lambda>")
+
+begin
+
+lemma min_block_size_2: "\<k> \<ge> 2"
+ using block_size_t by simp
+
+lemma points_index_pair: "y \<in> \<V> \<Longrightarrow> x \<in> \<V> \<Longrightarrow> x \<noteq> y \<Longrightarrow> size ({# bl \<in># \<B> . {x, y} \<subseteq> bl#}) = \<Lambda>"
+ using balanced card_2_iff empty_subsetI insert_subset points_index_def
+ by (metis of_nat_numeral)
+
+lemma index_one_empty_rm_blv [simp]:
+ assumes "\<Lambda> = 1" and " blv \<in># \<B>" and "p \<subseteq> blv" and "card p = 2"
+ shows "{#bl \<in># remove1_mset blv \<B> . p \<subseteq> bl#} = {#}"
+proof -
+ have blv_in: "blv \<in># filter_mset ((\<subseteq>) p) \<B>"
+ using assms by simp
+ have "p \<subseteq> \<V>" using assms wellformed by auto
+ then have "size (filter_mset ((\<subseteq>) p) \<B>) = 1"
+ using balanced assms by (simp add: points_index_def)
+ then show ?thesis using blv_in filter_diff_mset filter_single_mset
+ by (metis (no_types, lifting) add_mset_eq_single assms(3) insert_DiffM size_1_singleton_mset)
+qed
+
+lemma index_one_alt_bl_not_exist:
+ assumes "\<Lambda> = 1" and " blv \<in># \<B>" and "p \<subseteq> blv" and "card p = 2"
+ shows" \<And> bl. bl \<in># remove1_mset blv \<B> \<Longrightarrow> \<not> (p \<subseteq> bl) "
+ using index_one_empty_rm_blv
+ by (metis assms(1) assms(2) assms(3) assms(4) filter_mset_empty_conv)
+
+subsection \<open>Necessary Conditions for Existence\<close>
+
+text \<open>The necessary conditions on the existence of a $(v, k, \lambda)$-bibd are one of the
+fundamental first theorems on designs. Proofs based off MATH3301 lecture notes \cite{HerkeLectureNotes2016}
+ and Stinson \cite{stinsonCombinatorialDesignsConstructions2004}\<close>
+
+lemma necess_cond_1_rhs:
+ assumes "x \<in> \<V>"
+ shows "size ({# p \<in># (mset_set (\<V> - {x}) \<times># {# bl \<in># \<B> . x \<in> bl #}). fst p \<in> snd p#}) = \<Lambda> * (\<v>- 1)"
+proof -
+ let ?M = "mset_set (\<V> - {x})"
+ let ?B = "{# bl \<in># \<B> . x \<in> bl #}"
+ have m_distinct: "distinct_mset ?M" using assms mset_points_distinct_diff_one by simp
+ have y_point: "\<And> y . y \<in># ?M \<Longrightarrow> y \<in> \<V>" using assms
+ by (simp add: finite_sets)
+ have b_contents: "\<And> bl. bl \<in># ?B \<Longrightarrow> x \<in> bl" using assms by auto
+ have "\<And> y. y \<in># ?M \<Longrightarrow> y \<noteq> x" using assms
+ by (simp add: finite_sets)
+ then have "\<And> y .y \<in># ?M \<Longrightarrow> size ({# bl \<in># ?B . {x, y} \<subseteq> bl#}) = nat \<Lambda>"
+ using points_index_pair filter_filter_mset_ss_member y_point assms finite_sets index_ge_zero
+ by (metis nat_0_le nat_int_comparison(1))
+ then have "\<And> y .y \<in># ?M \<Longrightarrow> size ({# bl \<in># ?B . x \<in> bl \<and> y \<in> bl#}) = nat \<Lambda>"
+ by auto
+ then have bl_set_size: "\<And> y . y \<in># ?M \<Longrightarrow> size ({# bl \<in># ?B . y \<in> bl#}) = nat \<Lambda>"
+ using b_contents by (metis (no_types, lifting) filter_mset_cong)
+ then have final_size: "size (\<Sum>p\<in>#?M . ({#p#} \<times># {#bl \<in># ?B. p \<in> bl#})) = size (?M) * (nat \<Lambda>)"
+ using m_distinct size_Union_distinct_cart_prod_filter bl_set_size index_ge_zero by blast
+ have "size ?M = \<v> - 1" using v_non_zero
+ by (simp add: assms(1) finite_sets)
+ thus ?thesis using final_size
+ by (simp add: set_break_down_left index_ge_zero)
+qed
+
+lemma necess_cond_1_lhs:
+ assumes "x \<in> \<V>"
+ shows "size ({# p \<in># (mset_set (\<V> - {x}) \<times># {# bl \<in># \<B> . x \<in> bl #}). fst p \<in> snd p#})
+ = (\<B> rep x) * (\<k> - 1)"
+ (is "size ({# p \<in># (?M \<times># ?B). fst p \<in> snd p#}) = (\<B> rep x) * (\<k> - 1) ")
+proof -
+ have "\<And> y. y \<in># ?M \<Longrightarrow> y \<noteq> x" using assms
+ by (simp add: finite_sets)
+ have distinct_m: "distinct_mset ?M" using assms mset_points_distinct_diff_one by simp
+ have finite_M: "finite (\<V> - {x})" using finite_sets by auto
+ have block_choices: "size ?B = \<B> rep x"
+ by (simp add: assms(1) point_replication_number_def)
+ have bl_size: "\<forall> bl \<in># ?B. card {p \<in> \<V> . p \<in> bl } = \<k> " using uniform_unfold_point_set by simp
+ have x_in_set: "\<forall> bl \<in># ?B . {x} \<subseteq> {p \<in> \<V>. p \<in> bl}" using assms by auto
+ then have "\<forall> bl \<in># ?B. card {p \<in> (\<V> - {x}) . p \<in> bl } = card ({p \<in> \<V> . p \<in> bl } - {x})"
+ by (simp add: set_filter_diff_card)
+ then have "\<forall> bl \<in># ?B. card {p \<in> (\<V> - {x}) . p \<in> bl } = \<k> - 1"
+ using bl_size x_in_set card_Diff_subset finite_sets k_non_zero by auto
+ then have "\<And> bl . bl \<in># ?B \<Longrightarrow> size {#p \<in># ?M . p \<in> bl#} = nat (\<k> - 1)"
+ using assms finite_M card_size_filter_eq by auto
+ then have "size (\<Sum>bl\<in>#?B. ( {# p \<in># ?M . p \<in> bl #} \<times># {#bl#})) = size (?B) * nat (\<k> - 1)"
+ using distinct_m size_Union_distinct_cart_prod_filter2 by blast
+ thus ?thesis using block_choices k_non_zero by (simp add: set_break_down_right)
+qed
+
+lemma r_constant: "x \<in> \<V> \<Longrightarrow> (\<B> rep x) * (\<k> -1) = \<Lambda> * (\<v> - 1)"
+ using necess_cond_1_rhs necess_cond_1_lhs design_points_nempty by force
+
+lemma replication_number_value:
+ assumes "x \<in> \<V>"
+ shows "(\<B> rep x) = \<Lambda> * (\<v> - 1) div (\<k> - 1)"
+ using min_block_size_2 r_constant assms diff_gt_0_iff_gt diff_self zle_diff1_eq numeral_le_one_iff
+ by (metis less_int_code(1) linorder_neqE_linordered_idom nonzero_mult_div_cancel_right semiring_norm(69))
+
+lemma r_constant_alt: "\<forall> x \<in> \<V>. \<B> rep x = \<Lambda> * (\<v> - 1) div (\<k> - 1)"
+ using r_constant replication_number_value by blast
+
+end
+
+text \<open>Using the first necessary condition, it is possible to show that a bibd has
+a constant replication number\<close>
+
+sublocale bibd \<subseteq> constant_rep_design \<V> \<B> "(\<Lambda> * (\<v> - 1) div (\<k> - 1))"
+ using r_constant_alt by (unfold_locales) simp_all
+
+lemma (in t_design) bibdI [intro]: "\<t> = 2 \<Longrightarrow> bibd \<V> \<B> \<k> \<Lambda>\<^sub>t"
+ using t_lt_order block_size_t by (unfold_locales) (simp_all)
+
+context bibd
+begin
+
+abbreviation "\<r> \<equiv> (\<Lambda> * (\<v> - 1) div (\<k> - 1))"
+
+lemma necessary_condition_one:
+ shows "\<r> * (\<k> - 1) = \<Lambda> * (\<v> - 1)"
+ using necess_cond_1_rhs necess_cond_1_lhs design_points_nempty rep_number by auto
+
+lemma bibd_point_occ_rep:
+ assumes "x \<in> bl"
+ assumes "bl \<in># \<B>"
+ shows "(\<B> - {#bl#}) rep x = \<r> - 1"
+proof -
+ have xin: "x \<in> \<V>" using assms wf_invalid_point by blast
+ then have rep: "size {# blk \<in># \<B>. x \<in> blk #} = \<r>" using rep_number_unfold_set by simp
+ have "(\<B> - {#bl#}) rep x = size {# blk \<in># (\<B> - {#bl#}). x \<in> blk #}"
+ by (simp add: point_replication_number_def)
+ then have "(\<B> - {#bl#}) rep x = size {# blk \<in># \<B>. x \<in> blk #} - 1"
+ by (simp add: assms size_Diff_singleton)
+ then show ?thesis using assms rep r_gzero by simp
+qed
+
+lemma necess_cond_2_lhs: "size {# x \<in># (mset_set \<V> \<times># \<B>) . (fst x) \<in> (snd x) #} = \<v> * \<r>"
+proof -
+ let ?M = "mset_set \<V>"
+ have "\<And> p . p \<in># ?M \<Longrightarrow> size ({# bl \<in># \<B> . p \<in> bl #}) = nat (\<r>)"
+ using finite_sets rep_number_unfold_set r_gzero nat_eq_iff2 by auto
+ then have "size (\<Sum>p\<in>#?M. ({#p#} \<times># {#bl \<in># \<B>. p \<in> bl#})) = size ?M * nat (\<r>)"
+ using mset_points_distinct size_Union_distinct_cart_prod_filter by blast
+ thus ?thesis using r_gzero
+ by (simp add: set_break_down_left)
+qed
+
+lemma necess_cond_2_rhs: "size {# x \<in># (mset_set \<V> \<times># \<B>) . (fst x) \<in> (snd x) #} = \<b>*\<k>"
+ (is "size {# x \<in># (?M \<times># ?B). (fst x) \<in> (snd x) #} = \<b>*\<k>")
+proof -
+ have "\<And> bl . bl \<in># ?B \<Longrightarrow> size ({# p \<in># ?M . p \<in> bl #}) = nat \<k>"
+ using uniform k_non_zero uniform_unfold_point_set_mset by fastforce
+ then have "size (\<Sum>bl\<in>#?B. ( {# p \<in># ?M . p \<in> bl #} \<times># {#bl#})) = size (?B) * (nat \<k>)"
+ using mset_points_distinct size_Union_distinct_cart_prod_filter2 by blast
+ thus ?thesis using k_non_zero by (simp add: set_break_down_right)
+qed
+
+lemma necessary_condition_two:
+ shows "\<v> * \<r> = \<b> * \<k>"
+ using necess_cond_2_lhs necess_cond_2_rhs by simp
+
+theorem admissability_conditions:
+"\<r> * (\<k> - 1) = \<Lambda> * (\<v> - 1)"
+"\<v> * \<r> = \<b> * \<k>"
+ using necessary_condition_one necessary_condition_two by auto
+
+subsubsection \<open>BIBD Param Relationships\<close>
+
+lemma bibd_block_number: "\<b> = \<Lambda> * \<v> * (\<v> - 1) div (\<k> * (\<k>-1))"
+proof -
+ have "\<b> * \<k> = (\<v> * \<r>)" using necessary_condition_two by simp
+ then have k_dvd: "\<k> dvd (\<v> * \<r>)" by (metis dvd_triv_right)
+ then have "\<b> = (\<v> * \<r>) div \<k>" using necessary_condition_two min_block_size_2 by auto
+ then have "\<b> = (\<v> * ((\<Lambda> * (\<v> - 1) div (\<k> - 1)))) div \<k>" by simp
+ then have "\<b> = (\<v> * \<Lambda> * (\<v> - 1)) div ((\<k> - 1)* \<k>)" using necessary_condition_one
+ necessary_condition_two dvd_div_div_eq_mult dvd_div_eq_0_iff
+ by (smt (z3) dvd_triv_right mult.assoc mult.commute mult.left_commute mult_eq_0_iff )
+ then show ?thesis by (simp add: mult.commute)
+qed
+
+lemma symmetric_condition_1: "\<Lambda> * (\<v> - 1) = \<k> * (\<k> - 1) \<Longrightarrow> \<b> = \<v> \<and> \<r> = \<k>"
+ using b_non_zero bibd_block_number mult_eq_0_iff necessary_condition_two necessary_condition_one
+ by auto
+
+lemma index_lt_replication: "\<Lambda> < \<r>"
+proof -
+ have 1: "\<r> * (\<k> - 1) = \<Lambda> * (\<v> - 1)" using admissability_conditions by simp
+ have lhsnot0: "\<r> * (\<k> - 1) \<noteq> 0"
+ using no_zero_divisors rep_not_zero zdiv_eq_0_iff by blast
+ then have rhsnot0: "\<Lambda> * (\<v> - 1) \<noteq> 0" using 1 by simp
+ have "\<k> - 1 < \<v> - 1" using incomplete b_non_zero bibd_block_number not_less_eq by fastforce
+ thus ?thesis using 1 lhsnot0 rhsnot0
+ by (smt (verit, best) k_non_zero mult_le_less_imp_less r_gzero)
+qed
+
+lemma index_not_zero: "\<Lambda> \<ge> 1"
+ using index_ge_zero index_lt_replication int_one_le_iff_zero_less by fastforce
+
+lemma r_ge_two: "\<r> \<ge> 2"
+ using index_lt_replication index_not_zero by linarith
+
+lemma block_num_gt_rep: "\<b> > \<r>"
+proof -
+ have fact: "\<b> * \<k> = \<v> * \<r>" using admissability_conditions by auto
+ have lhsnot0: "\<b> * \<k> \<noteq> 0" using k_non_zero b_non_zero by auto
+ then have rhsnot0: "\<v> * \<r> \<noteq> 0" using fact by simp
+ then show ?thesis using incomplete lhsnot0
+ using complement_rep_number constant_rep_design.r_gzero incomplete_imp_incomp_block by fastforce
+qed
+
+lemma bibd_subset_occ:
+ assumes "x \<subseteq> bl" and "bl \<in># \<B>" and "card x = 2"
+ shows "size {# blk \<in># (\<B> - {#bl#}). x \<subseteq> blk #} = \<Lambda> - 1"
+proof -
+ have index: "size {# blk \<in># \<B>. x \<subseteq> blk #} = \<Lambda>" using points_index_def balanced assms
+ by (metis (full_types) of_nat_numeral subset_eq wf_invalid_point)
+ then have "size {# blk \<in># (\<B> - {#bl#}). x \<subseteq> blk #} = size {# blk \<in># \<B>. x \<subseteq> blk #} - 1"
+ by (simp add: assms size_Diff_singleton)
+ then show ?thesis using assms index_not_zero index by simp
+qed
+
+lemma necess_cond_one_param_balance: "\<b> > \<v> \<Longrightarrow> \<r> > \<k>"
+ using necessary_condition_two
+ by (smt mult_nonneg_nonneg nonzero_mult_div_cancel_right of_nat_0_le_iff r_gzero zdiv_mono2)
+
+subsection \<open>Constructing New bibd's\<close>
+text \<open>There are many constructions on bibd's to establish new bibds (or other types of designs).
+This section demonstrates this using both existing constructions, and by defining new constructions.\<close>
+subsubsection \<open>BIBD Complement, Multiple, Combine\<close>
+
+lemma comp_params_index_pair:
+ assumes "{x, y} \<subseteq> \<V>"
+ assumes "x \<noteq> y"
+ shows "\<B>\<^sup>C index {x, y} = \<b> + \<Lambda> - 2*\<r>"
+proof -
+ have xin: "x \<in> \<V>" and yin: "y \<in> \<V>" using assms by auto
+ have ge: "2*\<r> \<ge> \<Lambda>" using index_lt_replication
+ using r_gzero by linarith
+ have "size {# b \<in># \<B> . x \<in> b \<and> y \<in> b#} = \<Lambda>" using points_index_pair assms by simp
+ then have lambda: "size {# b \<in># \<B> . x \<in> b \<and> y \<in> b#} = nat \<Lambda>"
+ using index_ge_zero by auto
+ have "\<B>\<^sup>C index {x, y} = size {# b \<in># \<B> . x \<notin> b \<and> y \<notin> b #}"
+ using complement_index_2 assms by simp
+ also have "\<dots> = size \<B> - (size {# b \<in># \<B> . \<not> (x \<notin> b \<and> y \<notin> b) #})"
+ using size_filter_neg by blast
+ also have "... = size \<B> - (size {# b \<in># \<B> . x \<in> b \<or> y \<in> b#})" by auto
+ also have "... = \<b> - (size {# b \<in># \<B> . x \<in> b \<or> y \<in> b#})" by (simp add: of_nat_diff)
+ also have "... = \<b> - (size {# b \<in># \<B> . x \<in> b#} +
+ size {# b \<in># \<B> . y \<in> b#} - size {# b \<in># \<B> . x \<in> b \<and> y \<in> b#})"
+ by (simp add: mset_size_partition_dep)
+ also have "... = \<b> - (nat \<r> + nat \<r> - nat (\<Lambda>))" using rep_number_unfold_set lambda xin yin
+ by (metis (no_types, lifting) nat_int)
+ finally have "\<B>\<^sup>C index {x, y} = \<b> - (2*\<r> - \<Lambda>)"
+ using index_ge_zero index_lt_replication by linarith
+ thus ?thesis using ge diff_diff_right by simp
+qed
+
+lemma complement_bibd_index:
+ assumes "ps \<subseteq> \<V>"
+ assumes "card ps = 2"
+ shows "\<B>\<^sup>C index ps = \<b> + \<Lambda> - 2*\<r>"
+proof -
+ obtain x y where set: "ps = {x, y}" using b_non_zero bibd_block_number diff_is_0_eq incomplete
+ mult_0_right nat_less_le design_points_nempty assms by (metis card_2_iff)
+ then have "x \<noteq> y" using assms by auto
+ thus ?thesis using comp_params_index_pair assms
+ by (simp add: set)
+qed
+
+lemma complement_bibd:
+ assumes "\<k> \<le> \<v> - 2"
+ shows "bibd \<V> \<B>\<^sup>C (\<v> - \<k>) (\<b> + \<Lambda> - 2*\<r>)"
+proof -
+ interpret des: incomplete_design \<V> "\<B>\<^sup>C" "(\<v> - \<k>)"
+ using assms complement_incomplete by blast
+ show ?thesis proof (unfold_locales, simp_all)
+ show "2 \<le> des.\<v>" using assms block_size_t by linarith
+ show "\<And>ps. ps \<subseteq> \<V> \<Longrightarrow> card ps = 2 \<Longrightarrow>
+ \<B>\<^sup>C index ps = \<b> + \<Lambda> - 2 * (\<Lambda> * (des.\<v> - 1) div (\<k> - 1))"
+ using complement_bibd_index by simp
+ show "2 \<le> des.\<v> - \<k>" using assms block_size_t by linarith
+ qed
+qed
+
+lemma multiple_bibd: "n > 0 \<Longrightarrow> bibd \<V> (multiple_blocks n) \<k> (\<Lambda> * n)"
+ using multiple_t_design by (simp add: bibd_def)
+
+end
+
+locale two_bibd_eq_points = two_t_designs_eq_points \<V> \<B> \<k> \<B>' 2 \<Lambda> \<Lambda>'
+ + des1: bibd \<V> \<B> \<k> \<Lambda> + des2: bibd \<V> \<B>' \<k> \<Lambda>' for \<V> \<B> \<k> \<B>' \<Lambda> \<Lambda>'
+begin
+
+lemma combine_is_bibd: "bibd \<V>\<^sup>+ \<B>\<^sup>+ \<k> (\<Lambda> + \<Lambda>')"
+ by (unfold_locales)
+
+sublocale combine_bibd: bibd "\<V>\<^sup>+" "\<B>\<^sup>+" "\<k>" "(\<Lambda> + \<Lambda>')"
+ by (unfold_locales)
+
+end
+
+subsubsection \<open>Derived Designs\<close>
+text \<open>A derived bibd takes a block from a valid bibd as the new point sets, and the intersection
+of that block with other blocks as it's block set\<close>
+
+locale bibd_block_transformations = bibd +
+ fixes block :: "'a set" ("bl")
+ assumes valid_block: "bl \<in># \<B>"
+begin
+
+definition derived_blocks :: "'a set multiset" ("(\<B>\<^sup>D)") where
+"\<B>\<^sup>D \<equiv> {# bl \<inter> b . b \<in># (\<B> - {#bl#}) #}"
+
+lemma derive_define_flip: "{# b \<inter> bl . b \<in># (\<B> - {#bl#}) #} = \<B>\<^sup>D"
+ by (simp add: derived_blocks_def inf_sup_aci(1))
+
+lemma derived_points_order: "card bl = \<k>"
+ using uniform valid_block by simp
+
+lemma derived_block_num: "bl \<in># \<B> \<Longrightarrow> size \<B>\<^sup>D = \<b> - 1"
+ apply (simp add: derived_blocks_def size_remove1_mset_If valid_block)
+ using valid_block int_ops(6) by fastforce
+
+lemma derived_is_wellformed: "b \<in># \<B>\<^sup>D \<Longrightarrow> b \<subseteq> bl"
+ by (simp add: derived_blocks_def valid_block) (auto)
+
+lemma derived_point_subset_orig: "ps \<subseteq> bl \<Longrightarrow> ps \<subset> \<V>"
+ by (simp add: valid_block incomplete_imp_proper_subset subset_psubset_trans)
+
+lemma derived_obtain_orig_block:
+ assumes "b \<in># \<B>\<^sup>D"
+ obtains b2 where "b = b2 \<inter> bl" and "b2 \<in># remove1_mset bl \<B>"
+ using assms derived_blocks_def by auto
+
+sublocale derived_incidence_sys: incidence_system "bl" "\<B>\<^sup>D"
+ using derived_is_wellformed valid_block by (unfold_locales) (auto)
+
+sublocale derived_fin_incidence_system: finite_incidence_system "bl" "\<B>\<^sup>D"
+ using valid_block finite_blocks by (unfold_locales) simp_all
+
+lemma derived_blocks_nempty:
+ assumes "\<And> b .b \<in># remove1_mset bl \<B> \<Longrightarrow> bl |\<inter>| b > 0"
+ assumes "bld \<in># \<B>\<^sup>D"
+ shows "bld \<noteq> {}"
+proof -
+ obtain bl2 where inter: "bld = bl2 \<inter> bl" and member: "bl2 \<in># remove1_mset bl \<B>"
+ using assms derived_obtain_orig_block by blast
+ then have "bl |\<inter>| bl2 > 0" using assms(1) by blast
+ thus ?thesis using intersection_number_empty_iff finite_blocks valid_block
+ by (metis Int_commute dual_order.irrefl inter)
+qed
+
+lemma derived_is_design:
+ assumes "\<And> b. b \<in># remove1_mset bl \<B> \<Longrightarrow> bl |\<inter>| b > 0"
+ shows "design bl \<B>\<^sup>D"
+proof -
+ interpret fin: finite_incidence_system "bl" "\<B>\<^sup>D"
+ by (unfold_locales)
+ show ?thesis using assms derived_blocks_nempty by (unfold_locales) simp
+qed
+
+lemma derived_is_proper:
+ assumes "\<And> b. b \<in># remove1_mset bl \<B> \<Longrightarrow> bl |\<inter>| b > 0"
+ shows "proper_design bl \<B>\<^sup>D"
+proof -
+ interpret des: design "bl" "\<B>\<^sup>D"
+ using derived_is_design assms by fastforce
+ have "\<b> - 1 > 1" using block_num_gt_rep r_ge_two by linarith
+ then show ?thesis by (unfold_locales) (simp add: derived_block_num valid_block)
+qed
+
+
+subsubsection \<open>Residual Designs\<close>
+text \<open>Similar to derived designs, a residual design takes the complement of a block bl as it's new
+point set, and the complement of all other blocks with respect to bl.\<close>
+
+definition residual_blocks :: "'a set multiset" ("(\<B>\<^sup>R)") where
+"\<B>\<^sup>R \<equiv> {# b - bl . b \<in># (\<B> - {#bl#}) #}"
+
+lemma residual_order: "card (bl\<^sup>c) = \<v> - \<k>"
+ apply (simp add: valid_block wellformed block_complement_size)
+ using block_size_lt_v derived_points_order by auto
+
+lemma residual_block_num: "size (\<B>\<^sup>R) = \<b> - 1"
+ using b_positive by (simp add: residual_blocks_def size_remove1_mset_If valid_block int_ops(6))
+
+lemma residual_obtain_orig_block:
+ assumes "b \<in># \<B>\<^sup>R"
+ obtains bl2 where "b = bl2 - bl" and "bl2 \<in># remove1_mset bl \<B>"
+ using assms residual_blocks_def by auto
+
+lemma residual_blocks_ss: assumes "b \<in># \<B>\<^sup>R" shows "b \<subseteq> \<V>"
+proof -
+ have "b \<subseteq> (bl\<^sup>c)" using residual_obtain_orig_block
+ by (metis Diff_mono assms block_complement_def in_diffD order_refl wellformed)
+ thus ?thesis
+ using block_complement_subset_points by auto
+qed
+
+lemma residual_blocks_exclude: "b \<in># \<B>\<^sup>R \<Longrightarrow> x \<in> b \<Longrightarrow> x \<notin> bl"
+ using residual_obtain_orig_block by auto
+
+lemma residual_is_wellformed: "b \<in># \<B>\<^sup>R \<Longrightarrow> b \<subseteq> (bl\<^sup>c)"
+ apply (auto simp add: residual_blocks_def)
+ by (metis DiffI block_complement_def in_diffD wf_invalid_point)
+
+sublocale residual_incidence_sys: incidence_system "bl\<^sup>c" "\<B>\<^sup>R"
+ using residual_is_wellformed by (unfold_locales)
+
+lemma residual_is_finite: "finite (bl\<^sup>c)"
+ by (simp add: block_complement_def finite_sets)
+
+sublocale residual_fin_incidence_sys: finite_incidence_system "bl\<^sup>c" "\<B>\<^sup>R"
+ using residual_is_finite by (unfold_locales)
+
+lemma residual_blocks_nempty:
+ assumes "bld \<in># \<B>\<^sup>R"
+ assumes "multiplicity bl = 1"
+ shows "bld \<noteq> {}"
+proof -
+ obtain bl2 where inter: "bld = bl2 - bl" and member: "bl2 \<in># remove1_mset bl \<B>"
+ using assms residual_blocks_def by auto
+ then have ne: "bl2 \<noteq> bl" using assms
+ by (metis count_eq_zero_iff in_diff_count less_one union_single_eq_member)
+ have "card bl2 = card bl" using uniform valid_block member
+ by (metis in_diffD of_nat_eq_iff)
+ then have "card (bl2 - bl) > 0"
+ using finite_blocks member uniform set_card_diff_ge_zero valid_block by (metis in_diffD ne)
+ thus ?thesis using inter by fastforce
+qed
+
+lemma residual_is_design: "multiplicity bl = 1 \<Longrightarrow> design (bl\<^sup>c) \<B>\<^sup>R"
+ using residual_blocks_nempty by (unfold_locales)
+
+lemma residual_is_proper:
+ assumes "multiplicity bl = 1"
+ shows "proper_design (bl\<^sup>c) \<B>\<^sup>R"
+proof -
+ interpret des: design "bl\<^sup>c" "\<B>\<^sup>R" using residual_is_design assms by blast
+ have "\<b> - 1 > 1" using r_ge_two block_num_gt_rep by linarith
+ then show ?thesis using residual_block_num by (unfold_locales) auto
+qed
+
+end
+
+subsection \<open>Symmetric BIBD's\<close>
+text \<open>Symmetric bibd's are those where the order of the design equals the number of blocks\<close>
+
+locale symmetric_bibd = bibd +
+ assumes symmetric: "\<b> = \<v>"
+begin
+
+lemma rep_value_sym: "\<r> = \<k>"
+ using b_non_zero local.symmetric necessary_condition_two by auto
+
+lemma symmetric_condition_2: "\<Lambda> * (\<v> - 1) = \<k> * (\<k> - 1)"
+ using necessary_condition_one rep_value_sym by auto
+
+lemma sym_design_vk_gt_kl:
+ assumes "\<k> \<ge> \<Lambda> + 2"
+ shows "\<v> - \<k> > \<k> - \<Lambda>"
+proof (rule ccontr)
+ assume "\<not> (\<v> - \<k> > \<k> - \<Lambda>)"
+ then have "\<v> \<le> 2 * \<k> - \<Lambda>"
+ by linarith
+ then have "\<v> - 1 \<le> 2 * \<k> - \<Lambda> - 1" by linarith
+ then have "\<Lambda>* (\<v> - 1) \<le> \<Lambda>*( 2 * \<k> - \<Lambda> - 1)"
+ using index_ge_zero mult_le_cancel_left by fastforce
+ then have "\<k> * (\<k> - 1) \<le> \<Lambda>*( 2 * \<k> - \<Lambda> - 1)"
+ by (simp add: symmetric_condition_2)
+ then have "\<k> * (\<k> - 1) - \<Lambda>*( 2 * \<k> - \<Lambda> - 1) \<le> 0" by linarith
+ then have "(\<k> - \<Lambda>)*(\<k> - \<Lambda> - 1) \<le> 0"
+ by (simp add: mult.commute right_diff_distrib')
+ then have "\<k> = \<Lambda> \<or> \<k> = \<Lambda> + 1"
+ using mult_le_0_iff by force
+ thus False using assms
+ by simp
+qed
+
+end
+
+context bibd
+begin
+
+lemma symmetric_bibdI: "\<b> = \<v> \<Longrightarrow> symmetric_bibd \<V> \<B> \<k> \<Lambda>"
+ by unfold_locales simp
+
+lemma symmetric_bibdII: "\<Lambda> * (\<v> - 1) = \<k> * (\<k> - 1) \<Longrightarrow> symmetric_bibd \<V> \<B> \<k> \<Lambda>"
+ using symmetric_condition_1 by unfold_locales blast
+
+lemma symmetric_not_admissable: "\<Lambda> * (\<v> - 1) \<noteq> \<k> * (\<k> - 1) \<Longrightarrow> \<not> symmetric_bibd \<V> \<B> \<k> \<Lambda>"
+ using symmetric_bibd.symmetric_condition_2 by blast
+end
+
+context symmetric_bibd
+begin
+
+subsubsection \<open>Intersection Property on Symmetric BIBDs\<close>
+text \<open>Below is a proof of an important property on symmetric BIBD's regarding the equivalence
+of intersection numbers and the design index. This is an intuitive counting proof, and involved
+significantly more work in a formal environment. Based of Lecture Note \cite{HerkeLectureNotes2016}\<close>
+
+lemma intersect_mult_set_eq_block:
+ assumes "blv \<in># \<B>"
+ shows "p \<in># \<Sum>\<^sub>#{# mset_set (bl \<inter> blv) .bl \<in># (\<B> - {#blv#})#} \<longleftrightarrow> p \<in> blv"
+proof (auto, simp add: assms finite_blocks)
+ assume assm: "p \<in> blv"
+ then have "(\<B> - {#blv#}) rep p > 0" using bibd_point_occ_rep r_ge_two assms by auto
+ then obtain bl where "bl \<in># (\<B> - {#blv#}) \<and> p \<in> bl" using assms rep_number_g0_exists by metis
+ then show "\<exists>x\<in>#remove1_mset blv \<B>. p \<in># mset_set (x \<inter> blv)"
+ using assms assm finite_blocks by auto
+qed
+
+lemma intersect_mult_set_block_subset_iff:
+ assumes "blv \<in># \<B>"
+ assumes "p \<in># \<Sum>\<^sub>#{# mset_set {y .y \<subseteq> blv \<inter> b2 \<and> card y = 2} .b2 \<in># (\<B> - {#blv#})#}"
+ shows "p \<subseteq> blv"
+proof (rule subsetI)
+ fix x
+ assume asm: "x \<in> p"
+ obtain b2 where "p \<in># mset_set {y . y \<subseteq> blv \<inter> b2 \<and> card y = 2} \<and> b2 \<in>#(\<B> - {#blv#})"
+ using assms by blast
+ then have "p \<subseteq> blv \<inter> b2"
+ by (metis (no_types, lifting) elem_mset_set equals0D infinite_set_mset_mset_set mem_Collect_eq)
+ thus "x \<in> blv" using asm by auto
+qed
+
+lemma intersect_mult_set_block_subset_card:
+ assumes "blv \<in># \<B>"
+ assumes "p \<in># \<Sum>\<^sub>#{# mset_set {y .y \<subseteq> blv \<inter> b2 \<and> card y = 2} .b2 \<in># (\<B> - {#blv#})#}"
+ shows "card p = 2"
+proof -
+ obtain b2 where "p \<in># mset_set {y . y \<subseteq> blv \<inter> b2 \<and> card y = 2} \<and> b2 \<in>#(\<B> - {#blv#})"
+ using assms by blast
+ thus ?thesis
+ by (metis (mono_tags, lifting) elem_mset_set equals0D infinite_set_mset_mset_set mem_Collect_eq)
+qed
+
+lemma intersect_mult_set_block_with_point_exists:
+ assumes "blv \<in># \<B>" and "p \<subseteq> blv" and "\<Lambda> \<ge> 2" and "card p = 2"
+ shows "\<exists>x\<in>#remove1_mset blv \<B>. p \<in># mset_set {y. y \<subseteq> blv \<and> y \<subseteq> x \<and> card y = 2}"
+proof -
+ have "size {#b \<in># \<B> . p \<subseteq> b#} = \<Lambda>" using points_index_def assms
+ by (metis balanced_alt_def_all dual_order.trans of_nat_numeral wellformed)
+ then have "size {#bl \<in># (\<B> - {#blv#}) . p \<subseteq> bl#} \<ge> 1"
+ using assms by (simp add: size_Diff_singleton)
+ then obtain bl where "bl \<in># (\<B> - {#blv#}) \<and> p \<subseteq> bl" using assms filter_mset_empty_conv
+ by (metis diff_diff_cancel diff_is_0_eq' le_numeral_extra(4) size_empty zero_neq_one)
+ thus ?thesis
+ using assms finite_blocks by auto
+qed
+
+lemma intersect_mult_set_block_subset_iff_2:
+ assumes "blv \<in># \<B>" and "p \<subseteq> blv" and "\<Lambda> \<ge> 2" and "card p = 2"
+ shows "p \<in># \<Sum>\<^sub>#{# mset_set {y .y \<subseteq> blv \<inter> b2 \<and> card y = 2} .b2 \<in># (\<B> - {#blv#})#}"
+ by (auto simp add: intersect_mult_set_block_with_point_exists assms)
+
+lemma sym_sum_mset_inter_sets_count:
+ assumes "blv \<in># \<B>"
+ assumes "p \<in> blv"
+ shows "count (\<Sum>\<^sub>#{# mset_set (bl \<inter> blv) .bl \<in># (\<B> - {#blv#})#}) p = \<r> - 1"
+ (is "count (\<Sum>\<^sub>#?M) p = \<r> - 1")
+proof -
+ have size_inter: "size {# mset_set (bl \<inter> blv) | bl \<in># (\<B> - {#blv#}) . p \<in> bl#} = \<r> - 1"
+ using bibd_point_occ_rep point_replication_number_def
+ by (metis assms(1) assms(2) size_image_mset)
+ have inter_finite: "\<forall> bl \<in># (\<B> - {#blv#}) . finite (bl \<inter> blv)"
+ by (simp add: assms(1) finite_blocks)
+ have "\<And> bl . bl \<in># (\<B> - {#blv#}) \<Longrightarrow> p \<in> bl \<longrightarrow> count (mset_set (bl \<inter> blv)) p = 1"
+ using assms count_mset_set(1) inter_finite by simp
+ then have "\<And> bl . bl \<in># {#b1 \<in>#(\<B> - {#blv#}) . p \<in> b1#} \<Longrightarrow> count (mset_set (bl \<inter> blv)) p = 1"
+ by (metis (full_types) count_eq_zero_iff count_filter_mset)
+ then have pin: "\<And> P. P \<in># {# mset_set (bl \<inter> blv) | bl \<in># (\<B> - {#blv#}) . p \<in> bl#}
+ \<Longrightarrow> count P p = 1" by blast
+ have "?M = {# mset_set (bl \<inter> blv) | bl \<in># (\<B> - {#blv#}) . p \<in> bl#}
+ + {# mset_set (bl \<inter> blv) | bl \<in># (\<B> - {#blv#}) . p \<notin> bl#}"
+ by (metis image_mset_union multiset_partition)
+ then have "count (\<Sum>\<^sub>#?M) p = size {# mset_set (bl \<inter> blv) | bl \<in># (\<B> - {#blv#}) . p \<in> bl#} "
+ using pin by (auto simp add: count_sum_mset)
+ then show ?thesis using size_inter by linarith
+qed
+
+lemma sym_sum_mset_inter_sets_size:
+ assumes "blv \<in># \<B>"
+ shows "size (\<Sum>\<^sub>#{# mset_set (bl \<inter> blv) .bl \<in># (\<B> - {#blv#})#}) = \<k> * (\<r> - 1)"
+ (is "size (\<Sum>\<^sub>#?M) = \<k>* (\<r> - 1)")
+proof -
+ have eq: "set_mset (\<Sum>\<^sub>#{# mset_set (bl \<inter> blv) .bl \<in># (\<B> - {#blv#})#}) = blv"
+ using intersect_mult_set_eq_block assms by auto
+ then have k: "card (set_mset (\<Sum>\<^sub>#?M)) = \<k>"
+ by (simp add: assms)
+ have "\<And> p. p \<in># (\<Sum>\<^sub>#?M) \<Longrightarrow> count (\<Sum>\<^sub>#?M) p = \<r> - 1"
+ using sym_sum_mset_inter_sets_count assms eq by blast
+ thus ?thesis using k size_multiset_int_count by metis
+qed
+
+lemma sym_sum_inter_num:
+ assumes "b1 \<in># \<B>"
+ shows "(\<Sum>b2 \<in>#(\<B> - {#b1#}). b1 |\<inter>| b2) = \<k>* (\<r> - 1)"
+proof -
+ have "(\<Sum>b2 \<in>#(\<B> - {#b1#}). b1 |\<inter>| b2) = (\<Sum>b2 \<in>#(\<B> - {#b1#}). size (mset_set (b1 \<inter> b2)))"
+ by (simp add: intersection_number_def)
+ also have "... = size (\<Sum>\<^sub>#{#mset_set (b1 \<inter> bl). bl \<in># (\<B> - {#b1#})#})"
+ by (auto simp add: size_big_union_sum)
+ also have "... = size (\<Sum>\<^sub>#{#mset_set (bl \<inter> b1). bl \<in># (\<B> - {#b1#})#})"
+ by (metis Int_commute)
+ finally have "(\<Sum>b2 \<in>#(\<B> - {#b1#}). b1 |\<inter>| b2) = \<k> * (\<r> - 1)"
+ using sym_sum_mset_inter_sets_size assms by auto
+ then show ?thesis by simp
+qed
+
+lemma choose_two_int:
+ assumes " x \<ge> 0"
+ shows "nat (x :: int) choose 2 = ((x ::int) * ( x - 1)) div 2 "
+ using choose_two assms dvd_div_mult_self even_numeral int_nat_eq mult_cancel_right2 mult_eq_0_iff
+ mult_nonneg_nonneg nat_diff_distrib' nat_mult_distrib nat_one_as_int
+ numeral_Bit0_div_2 numerals(1) of_nat_numeral zdiv_int by (smt (verit)) (* Slow *)
+
+lemma sym_sum_mset_inter2_sets_count:
+ assumes "blv \<in># \<B>"
+ assumes "p \<subseteq> blv"
+ assumes "card p = 2"
+ shows "count (\<Sum>\<^sub>#{#mset_set {y .y \<subseteq> blv \<inter> b2 \<and> card y = 2}. b2 \<in># (\<B> - {#blv#})#}) p = \<Lambda> - 1"
+ (is "count (\<Sum>\<^sub>#?M) p = \<Lambda> - 1")
+proof -
+ have size_inter: "size {# mset_set {y .y \<subseteq> blv \<inter> b2 \<and> card y = 2} | b2 \<in># (\<B> - {#blv#}) . p \<subseteq> b2#}
+ = \<Lambda> - 1"
+ using bibd_subset_occ assms by simp
+ have "\<forall> b2 \<in># (\<B> - {#blv#}) . p \<subseteq> b2 \<longrightarrow> count (mset_set{y .y \<subseteq> blv \<inter> b2 \<and> card y = 2}) p = 1"
+ using assms(2) count_mset_set(1) assms(3) by (auto simp add: assms(1) finite_blocks)
+ then have "\<forall> bl \<in># {#b1 \<in>#(\<B> - {#blv#}) . p \<subseteq> b1#}.
+ count (mset_set {y .y \<subseteq> blv \<inter> bl \<and> card y = 2}) p = 1"
+ using count_eq_zero_iff count_filter_mset by (metis (no_types, lifting))
+ then have pin: "\<forall> P \<in># {# mset_set {y .y \<subseteq> blv \<inter> b2 \<and> card y = 2} | b2 \<in># (\<B> - {#blv#}) . p \<subseteq> b2#}.
+ count P p = 1"
+ using count_eq_zero_iff count_filter_mset by blast
+ have "?M = {# mset_set {y .y \<subseteq> blv \<inter> b2 \<and> card y = 2} | b2 \<in># (\<B> - {#blv#}) . p \<subseteq> b2#} +
+ {# mset_set {y .y \<subseteq> blv \<inter> b2 \<and> card y = 2} | b2 \<in># (\<B> - {#blv#}) . \<not> (p \<subseteq> b2)#}"
+ by (metis image_mset_union multiset_partition)
+ then have "count (\<Sum>\<^sub>#?M) p =
+ size {# mset_set {y .y \<subseteq> blv \<inter> b2 \<and> card y = 2} | b2 \<in># (\<B> - {#blv#}) . p \<subseteq> b2#}"
+ using pin by (auto simp add: count_sum_mset)
+ then show ?thesis using size_inter by linarith
+qed
+
+lemma sym_sum_mset_inter2_sets_size:
+ assumes "blv \<in># \<B>"
+ shows "size (\<Sum>\<^sub>#{#mset_set {y .y \<subseteq> blv \<inter> b2 \<and> card y = 2}. b2 \<in># (\<B> - {#blv#})#}) =
+ ((nat \<k>) choose 2) * (\<Lambda> -1)"
+ (is "size (\<Sum>\<^sub>#?M) = ((nat \<k>) choose 2) * (\<Lambda> -1)")
+proof (cases "\<Lambda> = 1")
+ case True
+ have empty: "\<And> b2 . b2 \<in># remove1_mset blv \<B> \<Longrightarrow> {y .y \<subseteq> blv \<and> y \<subseteq> b2 \<and> card y = 2} = {}"
+ using index_one_alt_bl_not_exist assms True by blast
+ then show ?thesis using sum_mset.neutral True by (simp add: empty)
+next
+ case False
+ then have index_min: "\<Lambda> \<ge> 2" using index_not_zero by linarith
+ have subset_card: "\<And> x . x \<in># (\<Sum>\<^sub>#?M) \<Longrightarrow> card x = 2"
+ proof -
+ fix x
+ assume a: "x \<in># (\<Sum>\<^sub>#?M)"
+ then obtain b2 where "x \<in># mset_set {y . y \<subseteq> blv \<inter> b2 \<and> card y = 2} \<and> b2 \<in>#(\<B> - {#blv#})"
+ by blast
+ thus "card x = 2" using mem_Collect_eq
+ by (metis (mono_tags, lifting) elem_mset_set equals0D infinite_set_mset_mset_set)
+ qed
+ have eq: "set_mset (\<Sum>\<^sub>#?M) = {bl . bl \<subseteq> blv \<and> card bl = 2}"
+ proof
+ show "set_mset (\<Sum>\<^sub>#?M) \<subseteq> {bl . bl \<subseteq> blv \<and> card bl = 2}"
+ using subset_card intersect_mult_set_block_subset_iff assms by blast
+ show "{bl . bl \<subseteq> blv \<and> card bl = 2} \<subseteq> set_mset (\<Sum>\<^sub>#?M)"
+ using intersect_mult_set_block_subset_iff_2 assms index_min by blast
+ qed
+ have "card blv = (nat \<k>)" using uniform assms by (metis nat_int)
+ then have k: "card (set_mset (\<Sum>\<^sub>#?M)) = ((nat \<k>) choose 2)" using eq n_subsets
+ by (simp add: n_subsets assms finite_blocks)
+ thus ?thesis using k size_multiset_int_count sym_sum_mset_inter2_sets_count assms eq subset_card
+ by (metis (no_types, lifting) intersect_mult_set_block_subset_iff)
+qed
+
+lemma sum_choose_two_inter_num:
+ assumes "b1 \<in># \<B>"
+ shows "(\<Sum>b2 \<in># (\<B> - {#b1#}). (nat (b1 |\<inter>| b2) choose 2)) = ((\<Lambda> * (\<Lambda> - 1) div 2)) * (\<v> -1)"
+proof -
+ have div_fact: "2 dvd (\<Lambda> * (\<Lambda> - 1))"by simp
+ have div_fact_2: "2 dvd (\<Lambda> * (\<v> - 1))" using symmetric_condition_2 by simp
+ have "(\<Sum>b2 \<in># (\<B> - {#b1#}). (nat (b1 |\<inter>| b2) choose 2)) = (\<Sum>b2 \<in># (\<B> - {#b1#}). nat (b1 |\<inter>|\<^sub>2 b2 ))"
+ using n_inter_num_choose_design_inter assms by (simp add: in_diffD)
+ then have sum_fact: "(\<Sum>b2 \<in># (\<B> - {#b1#}).(nat (b1 |\<inter>| b2) choose 2))
+ = ((nat \<k>) choose 2) * (\<Lambda> -1)"
+ using assms sym_sum_mset_inter2_sets_size
+ by (auto simp add: size_big_union_sum n_intersect_num_subset_def)
+ have "((nat \<k>) choose 2) * (\<Lambda> -1) = ((\<Lambda> * (\<v> - 1) div 2)) * (\<Lambda> -1)"
+ using choose_two_int symmetric_condition_2 k_non_zero by auto
+ then have "((nat \<k>) choose 2) * (\<Lambda> -1) = ((\<Lambda> * (\<Lambda> - 1) div 2)) * (\<v> -1)"
+ using div_fact div_fact_2 by (smt div_mult_swap mult.assoc mult.commute)
+ then show ?thesis using sum_fact by simp
+qed
+
+lemma sym_sum_inter_num_sq:
+ assumes "b1 \<in># \<B>"
+ shows "(\<Sum>bl \<in># (remove1_mset b1 \<B>). (b1 |\<inter>| bl)^2) = \<Lambda>^2 * ( \<v> - 1)"
+proof -
+ have dvd: "2 dvd (( \<v> - 1) * (\<Lambda> * (\<Lambda> - 1)))" by simp
+ have a: "(\<Sum>b2 \<in>#(\<B> - {#b1#}). int (nat (b1 |\<inter>| b2) choose 2)) =
+ (\<Sum>bl \<in># (remove1_mset b1 \<B>). ((b1 |\<inter>| bl) * ((b1 |\<inter>| bl) - 1)) div 2)"
+ using choose_two_int by (simp add: intersection_num_non_neg)
+ have b: "(\<Sum>b2 \<in>#(\<B> - {#b1#}). int (nat (b1 |\<inter>| b2) choose 2)) =
+ (\<Sum>b2 \<in>#(\<B> - {#b1#}). (nat (b1 |\<inter>| b2) choose 2))" by simp
+ have "(\<Sum>b2 \<in>#(\<B> - {#b1#}). (nat (b1 |\<inter>| b2) choose 2)) = ((\<Lambda> * (\<Lambda> - 1)) div 2) * ( \<v> - 1)"
+ using sum_choose_two_inter_num assms by blast
+ then have start: "(\<Sum>bl \<in># (remove1_mset b1 \<B>). ((b1 |\<inter>| bl) * ((b1 |\<inter>| bl) - 1)) div 2)
+ = ((\<Lambda> * (\<Lambda> - 1)) div 2) * (\<v> - 1)"
+ using a b by linarith
+ have sum_dvd: "2 dvd (\<Sum>bl \<in># (remove1_mset b1 \<B>). (b1 |\<inter>| bl) * ((b1 |\<inter>| bl) - 1))"
+ by (simp add: sum_mset_dvd)
+ then have "(\<Sum>bl \<in># (remove1_mset b1 \<B>). (b1 |\<inter>| bl) * ((b1 |\<inter>| bl) - 1)) div 2
+ = (\<v> - 1) * ((\<Lambda> * (\<Lambda> - 1)) div 2)"
+ using start by (simp add: sum_mset_distrib_div_if_dvd)
+ then have "(\<Sum>bl \<in># (remove1_mset b1 \<B>). (b1 |\<inter>| bl)^2)
+ - (\<Sum>bl \<in># (remove1_mset b1 \<B>). (b1 |\<inter>| bl)) = ((\<v> - 1) * (\<Lambda> * (\<Lambda> - 1)))"
+ using sum_dvd dvd
+ by (simp add: dvd_div_eq_iff div_mult_swap int_distrib(4) power2_eq_square sum_mset_add_diff)
+ then have "(\<Sum>bl \<in># (remove1_mset b1 \<B>). (b1 |\<inter>| bl)^2) - (\<Lambda> * (\<v> - 1)) = ((\<v> - 1) * (\<Lambda> * (\<Lambda> - 1)))"
+ using sym_sum_inter_num assms rep_value_sym symmetric_condition_2 by auto
+ then have "(\<Sum>bl \<in># (remove1_mset b1 \<B>). (b1 |\<inter>| bl)^2) = (\<Lambda> * (\<v> - 1)) * (\<Lambda> - 1) + (\<Lambda> * (\<v> - 1))"
+ using diff_eq_eq by fastforce
+ then have "(\<Sum>bl \<in># (remove1_mset b1 \<B>). (b1 |\<inter>| bl)^2) = (\<Lambda> * (\<v> - 1)) * (\<Lambda> - 1 + 1)"
+ using int_distrib(2) by (metis mult_numeral_1_right numeral_One)
+ thus ?thesis by (simp add: power2_eq_square)
+qed
+
+lemma sym_sum_inter_num_to_zero:
+ assumes "b1 \<in># \<B>"
+ shows "(\<Sum>bl \<in># (remove1_mset b1 \<B>). ((b1 |\<inter>| bl) - \<Lambda>)^2) = 0"
+proof -
+ have rm1_size: "size (remove1_mset b1 \<B>) = \<v> - 1" using assms b_non_zero int_ops(6)
+ by (auto simp add: symmetric size_remove1_mset_If)
+ have "\<And> bl . bl \<in># (remove1_mset b1 \<B>) \<Longrightarrow> ((b1 |\<inter>| bl) - \<Lambda>)^2 =
+ (((b1 |\<inter>| bl)^2) - (2 * \<Lambda> * (b1 |\<inter>| bl)) + (\<Lambda>^2))"
+ by (simp add: power2_diff)
+ then have "(\<Sum>bl \<in># (remove1_mset b1 \<B>). ((b1 |\<inter>| bl) - \<Lambda>)^2) =
+ (\<Sum>bl \<in># (remove1_mset b1 \<B>). (((b1 |\<inter>| bl)^2) - (2 * \<Lambda> * (b1 |\<inter>| bl)) + (\<Lambda>^2)))"
+ using sum_over_fun_eq by auto
+ also have "... = \<Lambda>^2 * (\<v> - 1) - 2 * \<Lambda> * (\<Sum>bl \<in># (remove1_mset b1 \<B>). ((b1 |\<inter>| bl)))
+ + (\<v> - 1) * (\<Lambda>^2)"
+ using sym_sum_inter_num_sq rm1_size
+ by (simp add: assms sum_mset.distrib sum_mset_add_diff sum_mset_distrib_left)
+ finally have "(\<Sum>bl \<in># (remove1_mset b1 \<B>). ((b1 |\<inter>| bl) - \<Lambda>)^2) = 0"
+ using rep_value_sym symmetric_condition_2 sym_sum_inter_num assms
+ by (auto simp add: power2_eq_square)
+ thus ?thesis by simp
+qed
+
+theorem sym_block_intersections_index [simp]:
+ assumes "b1 \<in># \<B>"
+ assumes "b2 \<in># (\<B> - {#b1#})"
+ shows "b1 |\<inter>| b2 = \<Lambda>"
+proof -
+ have pos: "\<And> bl . ((b1 |\<inter>| bl) - \<Lambda>)^2 \<ge> 0" by simp
+ have "(\<Sum>bl \<in># (remove1_mset b1 \<B>). ((b1 |\<inter>| bl) - \<Lambda>)^2) = 0"
+ using sym_sum_inter_num_to_zero assms by simp
+ then have "\<And> bl. bl \<in> set_mset (remove1_mset b1 \<B>) \<Longrightarrow> ((b1 |\<inter>| bl) - \<Lambda>)^2 = 0"
+ using sum_mset_0_iff_ge_0 pos by (metis (no_types, lifting))
+ thus ?thesis
+ using assms(2) by auto
+qed
+
+subsubsection \<open>Symmetric BIBD is Simple\<close>
+
+lemma sym_block_mult_one [simp]:
+ assumes "bl \<in># \<B>"
+ shows "multiplicity bl = 1"
+proof (rule ccontr)
+ assume "\<not> (multiplicity bl = 1)"
+ then have not: "multiplicity bl \<noteq> 1" by simp
+ have "multiplicity bl \<noteq> 0" using assms
+ by simp
+ then have m: "multiplicity bl \<ge> 2" using not by linarith
+ then have blleft: "bl \<in># (\<B> - {#bl#})"
+ using in_diff_count by fastforce
+ have "bl |\<inter>| bl = \<k>" using k_non_zero assms
+ by (simp add: intersection_number_def)
+ then have keql: "\<k> = \<Lambda>" using sym_block_intersections_index blleft assms by simp
+ then have "\<v> = \<k>"
+ using keql index_lt_replication rep_value_sym block_size_lt_v diffs0_imp_equal k_non_zero zero_diff by linarith
+ then show False using incomplete
+ by simp
+qed
+
+end
+
+sublocale symmetric_bibd \<subseteq> simple_design
+ by unfold_locales simp
+
+subsubsection \<open>Residual/Derived Sym BIBD Constructions\<close>
+text \<open>Using the intersect result, we can reason further on residual and derived designs.
+Proofs based off lecture notes \cite{HerkeLectureNotes2016}\<close>
+
+locale symmetric_bibd_block_transformations = symmetric_bibd + bibd_block_transformations
+begin
+
+lemma derived_block_size [simp]:
+ assumes "b \<in># \<B>\<^sup>D"
+ shows "card b = \<Lambda>"
+proof -
+ obtain bl2 where set: "bl2 \<in># remove1_mset bl \<B>" and inter: "b = bl2 \<inter> bl"
+ using derived_blocks_def assms by (meson derived_obtain_orig_block)
+ then have "card b = bl2 |\<inter>| bl"
+ by (simp add: intersection_number_def)
+ thus ?thesis using sym_block_intersections_index
+ using set intersect_num_commute valid_block by fastforce
+qed
+
+lemma derived_points_index [simp]:
+ assumes "ps \<subseteq> bl"
+ assumes "card ps = 2"
+ shows "\<B>\<^sup>D index ps = \<Lambda> - 1"
+proof -
+ have b_in: "\<And> b . b \<in># (remove1_mset bl \<B>) \<Longrightarrow> ps \<subseteq> b \<Longrightarrow> ps \<subseteq> b \<inter> bl"
+ using assms by blast
+ then have orig: "ps \<subseteq> \<V>"
+ using valid_block assms wellformed by blast
+ then have lam: "size {# b \<in># \<B> . ps \<subseteq> b #} = \<Lambda>" using balanced
+ by (simp add: assms(2) points_index_def)
+ then have "size {# b \<in># remove1_mset bl \<B> . ps \<subseteq> b #} = size {# b \<in># \<B> . ps \<subseteq> b #} - 1"
+ using assms valid_block by (simp add: size_Diff_submset)
+ then have "size {# b \<in># remove1_mset bl \<B> . ps \<subseteq> b #} = \<Lambda> - 1"
+ using lam index_not_zero by linarith
+ then have "size {# bl \<inter> b | b \<in># (remove1_mset bl \<B>) . ps \<subseteq> bl \<inter> b #} = \<Lambda> - 1"
+ using b_in by (metis (no_types, lifting) Int_subset_iff filter_mset_cong size_image_mset)
+ then have "size {# x \<in># {# bl \<inter> b . b \<in># (remove1_mset bl \<B>) #} . ps \<subseteq> x #} = \<Lambda> - 1"
+ by (metis image_mset_filter_swap)
+ then have "size {# x \<in># \<B>\<^sup>D . ps \<subseteq> x #} = \<Lambda> - 1" by (simp add: derived_blocks_def)
+ thus ?thesis by (simp add: points_index_def)
+qed
+
+lemma sym_derive_design_bibd:
+ assumes "\<Lambda> > 1"
+ shows "bibd bl \<B>\<^sup>D \<Lambda> (\<Lambda> - 1)"
+proof -
+ interpret des: proper_design bl "\<B>\<^sup>D" using derived_is_proper assms valid_block by auto
+ have "\<Lambda> < \<k>" using index_lt_replication rep_value_sym by linarith
+ then show ?thesis using derived_block_size assms derived_points_index derived_points_order
+ by (unfold_locales) (simp_all)
+qed
+
+lemma residual_block_size [simp]:
+ assumes "b \<in># \<B>\<^sup>R"
+ shows "card b = \<k> - \<Lambda>"
+proof -
+ obtain bl2 where sub: "b = bl2 - bl" and mem: "bl2 \<in># remove1_mset bl \<B>"
+ using assms residual_blocks_def by auto
+ then have "card b = card bl2 - card (bl2 \<inter> bl)"
+ using card_Diff_subset_Int valid_block finite_blocks
+ by (simp add: card_Diff_subset_Int)
+ then have "card b = card bl2 - bl2 |\<inter>| bl"
+ using intersection_number_def finite_blocks card_inter_lt_single
+ by (metis assms derived_fin_incidence_system.finite_sets finite_Diff2 of_nat_diff
+ residual_fin_incidence_sys.finite_blocks sub)
+ thus ?thesis using sym_block_intersections_index uniform
+ by (metis valid_block in_diffD intersect_num_commute mem)
+qed
+
+lemma residual_index [simp]:
+ assumes "ps \<subseteq> bl\<^sup>c"
+ assumes "card ps = 2"
+ shows "(\<B>\<^sup>R) index ps = \<Lambda>"
+proof -
+ have a: "\<And> b . (b \<in># remove1_mset bl \<B> \<Longrightarrow> ps \<subseteq> b \<Longrightarrow> ps \<subseteq> (b - bl))" using assms
+ by (smt DiffI block_comp_elem_alt_left in_diffD subset_eq wellformed)
+ have b: "\<And> b . (b \<in># remove1_mset bl \<B> \<Longrightarrow> ps \<subseteq> (b - bl) \<Longrightarrow> ps \<subseteq> b)"
+ by auto
+ have not_ss: "\<not> (ps \<subseteq> bl)" using set_diff_non_empty_not_subset blocks_nempty t_non_zero assms
+ block_complement_def by fastforce
+ have "\<B>\<^sup>R index ps = size {# x \<in># {# b - bl . b \<in># (remove1_mset bl \<B>) #} . ps \<subseteq> x #}"
+ using assms valid_block by (simp add: points_index_def residual_blocks_def)
+ also have "... = size {# b - bl | b \<in># (remove1_mset bl \<B>) . ps \<subseteq> b - bl #} "
+ by (metis image_mset_filter_swap)
+ finally have "\<B>\<^sup>R index ps = size {# b \<in># (remove1_mset bl \<B>) . ps \<subseteq> b #} " using a b
+ by (metis (no_types, lifting) filter_mset_cong size_image_mset)
+ thus ?thesis
+ using balanced not_ss assms points_index_alt_def block_complement_subset_points by auto
+qed
+
+lemma sym_residual_design_bibd:
+ assumes "\<k> \<ge> \<Lambda> + 2"
+ shows "bibd (bl\<^sup>c) \<B>\<^sup>R (\<k> - \<Lambda>) \<Lambda>"
+proof -
+ interpret des: proper_design "bl\<^sup>c" "\<B>\<^sup>R"
+ using residual_is_proper assms(1) valid_block sym_block_mult_one by fastforce
+ show ?thesis using residual_block_size assms sym_design_vk_gt_kl residual_order residual_index
+ by(unfold_locales) simp_all
+qed
+
+end
+
+subsection \<open>BIBD's and Other Block Designs\<close>
+text \<open>BIBD's are closely related to other block designs by indirect inheritance\<close>
+
+sublocale bibd \<subseteq> k_\<Lambda>_PBD \<V> \<B> \<Lambda> \<k>
+ using block_size_gt_t by (unfold_locales) simp_all
+
+lemma incomplete_PBD_is_bibd:
+ assumes "k < card V" and "k_\<Lambda>_PBD V B \<Lambda> k"
+ shows "bibd V B k \<Lambda>"
+proof -
+ interpret inc: incomplete_design V B k using assms
+ by (auto simp add: block_design.incomplete_designI k_\<Lambda>_PBD.axioms(2))
+ interpret pairwise_balance: pairwise_balance V B \<Lambda> using assms
+ by (auto simp add: k_\<Lambda>_PBD.axioms(1))
+ show ?thesis using assms k_\<Lambda>_PBD.block_size_t by (unfold_locales) (simp_all)
+qed
+
+lemma (in bibd) bibd_to_pbdI[intro]:
+ assumes "\<Lambda> = 1"
+ shows "k_PBD \<V> \<B> \<k>"
+proof -
+ interpret pbd: k_\<Lambda>_PBD \<V> \<B> \<Lambda> \<k>
+ by (simp add: k_\<Lambda>_PBD_axioms)
+ show ?thesis using assms by (unfold_locales) (simp_all add: t_lt_order min_block_size_2)
+qed
+
+locale incomplete_PBD = incomplete_design + k_\<Lambda>_PBD
+
+sublocale incomplete_PBD \<subseteq> bibd
+ using block_size_t by (unfold_locales) simp
+
end
\ No newline at end of file
diff --git a/thys/Design_Theory/Block_Designs.thy b/thys/Design_Theory/Block_Designs.thy
--- a/thys/Design_Theory/Block_Designs.thy
+++ b/thys/Design_Theory/Block_Designs.thy
@@ -1,558 +1,558 @@
-(* Title: Block_Designs.thy
- Author: Chelsea Edmonds
-*)
-
-section \<open>Block and Balanced Designs\<close>
-text \<open>We define a selection of the many different types of block and balanced designs, building up
-to properties required for defining a BIBD, in addition to several base generalisations\<close>
-
-theory Block_Designs imports Design_Operations
-begin
-
-subsection \<open>Block Designs\<close>
-text \<open>A block design is a design where all blocks have the same size.\<close>
-
-subsubsection \<open>K Block Designs\<close>
-text \<open>An important generalisation of a typical block design is the $\mathcal{K}$ block design,
-where all blocks must have a size $x$ where $x \in \mathcal{K}$\<close>
-locale K_block_design = proper_design +
- fixes sizes :: "int set" ("\<K>")
- assumes block_sizes: "bl \<in># \<B> \<Longrightarrow> (int (card bl)) \<in> \<K>"
- assumes positive_ints: "x \<in> \<K> \<Longrightarrow> x > 0"
-begin
-
-lemma sys_block_size_subset: "sys_block_sizes \<subseteq> \<K>"
- using block_sizes sys_block_sizes_obtain_bl by blast
-
-end
-
-subsubsection\<open>Uniform Block Design\<close>
-text \<open>The typical uniform block design is defined below\<close>
-locale block_design = proper_design +
- fixes u_block_size :: int ("\<k>")
- assumes uniform [simp]: "bl \<in># \<B> \<Longrightarrow> card bl = \<k>"
-begin
-
-lemma k_non_zero: "\<k> \<ge> 1"
-proof -
- obtain bl where bl_in: "bl \<in># \<B>"
- using design_blocks_nempty by auto
- then have "int (card bl) \<ge> 1" using block_size_gt_0
- by (metis less_not_refl less_one not_le_imp_less of_nat_1 of_nat_less_iff)
- thus ?thesis by (simp add: bl_in)
-qed
-
-lemma uniform_alt_def_all: "\<forall> bl \<in># \<B> .card bl = \<k>"
- using uniform by auto
-
-lemma uniform_unfold_point_set: "bl \<in># \<B> \<Longrightarrow> card {p \<in> \<V>. p \<in> bl} = \<k>"
- using uniform wellformed by (simp add: Collect_conj_eq inf.absorb_iff2)
-
-lemma uniform_unfold_point_set_mset: "bl \<in># \<B> \<Longrightarrow> size {#p \<in># mset_set \<V>. p \<in> bl #} = \<k>"
- using uniform_unfold_point_set by (simp add: finite_sets)
-
-lemma sys_block_sizes_uniform [simp]: "sys_block_sizes = {\<k>}"
-proof -
- have "sys_block_sizes = {bs . \<exists> bl . bs = card bl \<and> bl\<in># \<B>}" by (simp add: sys_block_sizes_def)
- then have "sys_block_sizes = {bs . bs = \<k>}" using uniform uniform_unfold_point_set
- b_positive block_set_nempty_imp_block_ex
- by (smt (verit, best) Collect_cong design_blocks_nempty)
- thus ?thesis by auto
-qed
-
-lemma sys_block_sizes_uniform_single: "is_singleton (sys_block_sizes)"
- by simp
-
-lemma uniform_size_incomp: "\<k> \<le> \<v> - 1 \<Longrightarrow> bl \<in># \<B> \<Longrightarrow> incomplete_block bl"
- using uniform k_non_zero of_nat_less_iff zle_diff1_eq by metis
-
-lemma uniform_complement_block_size:
- assumes "bl \<in># \<B>\<^sup>C"
- shows "card bl = \<v> - \<k>"
-proof -
- obtain bl' where bl_assm: "bl = bl'\<^sup>c \<and> bl' \<in># \<B>"
- using wellformed assms by (auto simp add: complement_blocks_def)
- then have "int (card bl') = \<k>" by simp
- thus ?thesis using bl_assm block_complement_size wellformed
- by (simp add: block_size_lt_order of_nat_diff)
-qed
-
-lemma uniform_complement[intro]:
- assumes "\<k> \<le> \<v> - 1"
- shows "block_design \<V> \<B>\<^sup>C (\<v> - \<k>)"
-proof -
- interpret des: proper_design \<V> "\<B>\<^sup>C"
- using uniform_size_incomp assms complement_proper_design by auto
- show ?thesis using assms uniform_complement_block_size by (unfold_locales) (simp)
-qed
-
-lemma block_size_lt_v: "\<k> \<le> \<v>"
- using v_non_zero block_size_lt_v design_blocks_nempty uniform by auto
-
-end
-
-lemma (in proper_design) block_designI[intro]: "(\<And> bl . bl \<in># \<B> \<Longrightarrow> card bl = k)
- \<Longrightarrow> block_design \<V> \<B> k"
- by (unfold_locales) (auto)
-
-context block_design
-begin
-
-lemma block_design_multiple: "n > 0 \<Longrightarrow> block_design \<V> (multiple_blocks n) \<k>"
- using elem_in_repeat_in_original multiple_proper_design proper_design.block_designI
- by (metis block_set_nempty_imp_block_ex design_blocks_nempty int_int_eq uniform_alt_def_all)
-
-end
-text \<open>A uniform block design is clearly a type of $K$\_block\_design with a singleton $K$ set\<close>
-sublocale block_design \<subseteq> K_block_design \<V> \<B> "{\<k>}"
- using k_non_zero uniform by unfold_locales simp_all
-
-subsubsection \<open>Incomplete Designs\<close>
-text \<open>An incomplete design is a design where $k < v$, i.e. no block is equal to the point set\<close>
-locale incomplete_design = block_design +
- assumes incomplete: "\<k> < \<v>"
-
-begin
-
-lemma incomplete_imp_incomp_block: "bl \<in># \<B> \<Longrightarrow> incomplete_block bl"
- using incomplete uniform uniform_size_incomp by fastforce
-
-lemma incomplete_imp_proper_subset: "bl \<in># \<B> \<Longrightarrow> bl \<subset> \<V>"
- by (simp add: incomplete_block_proper_subset incomplete_imp_incomp_block wellformed)
-end
-
-lemma (in block_design) incomplete_designI[intro]: "\<k> < \<v> \<Longrightarrow> incomplete_design \<V> \<B> \<k>"
- by unfold_locales auto
-
-context incomplete_design
-begin
-
-lemma multiple_incomplete: "n > 0 \<Longrightarrow> incomplete_design \<V> (multiple_blocks n) \<k>"
- using block_design_multiple incomplete by (simp add: block_design.incomplete_designI)
-
-lemma complement_incomplete: "incomplete_design \<V> (\<B>\<^sup>C) (\<v> - \<k>)"
-proof -
- have "\<v> - \<k> < \<v>" using v_non_zero k_non_zero by linarith
- thus ?thesis using uniform_complement incomplete incomplete_designI
- by (simp add: block_design.incomplete_designI)
-qed
-
-end
-
-subsection \<open>Balanced Designs\<close>
-text \<open>t-wise balance is a design with the property that all point subsets of size $t$ occur in
-$\lambda_t$ blocks\<close>
-
-locale t_wise_balance = proper_design +
- fixes grouping :: int ("\<t>") and index :: int ("\<Lambda>\<^sub>t")
- assumes t_non_zero: "\<t> \<ge> 1"
- assumes t_lt_order: "\<t> \<le> \<v>"
- assumes balanced [simp]: "ps \<subseteq> \<V> \<Longrightarrow> card ps = \<t> \<Longrightarrow> \<B> index ps = \<Lambda>\<^sub>t"
-begin
-
-lemma balanced_alt_def_all: "\<forall> ps \<subseteq> \<V> . card ps = \<t> \<longrightarrow> \<B> index ps = \<Lambda>\<^sub>t"
- using balanced by auto
-
-end
-
-lemma (in proper_design) t_wise_balanceI[intro]: "\<t> \<le> \<v> \<Longrightarrow> \<t> \<ge> 1 \<Longrightarrow>
- (\<And> ps . ps \<subseteq> \<V> \<Longrightarrow> card ps = \<t> \<Longrightarrow> \<B> index ps = \<Lambda>\<^sub>t) \<Longrightarrow> t_wise_balance \<V> \<B> \<t> \<Lambda>\<^sub>t"
- by (unfold_locales) auto
-
-context t_wise_balance
-begin
-
-lemma obtain_t_subset_points:
- obtains T where "T \<subseteq> \<V>" "card T = \<t>" "finite T"
- using obtain_subset_with_card_int_n design_points_nempty t_lt_order t_non_zero finite_sets
- by (metis int_one_le_iff_zero_less less_eq_int_code(1) linorder_not_le verit_la_generic)
-
-lemma multiple_t_wise_balance_index [simp]:
- assumes "ps \<subseteq> \<V>"
- assumes "card ps = \<t>"
- shows "(multiple_blocks n) index ps = \<Lambda>\<^sub>t * n"
- using multiple_point_index balanced assms by fastforce
-
-lemma multiple_t_wise_balance:
- assumes "n > 0"
- shows "t_wise_balance \<V> (multiple_blocks n) \<t> (\<Lambda>\<^sub>t * n)"
-proof -
- interpret des: proper_design \<V> "(multiple_blocks n)" by (simp add: assms multiple_proper_design)
- show ?thesis using t_non_zero t_lt_order multiple_t_wise_balance_index
- by (unfold_locales) (simp_all)
-qed
-
-lemma twise_set_pair_index: "ps \<subseteq> \<V> \<Longrightarrow> ps2 \<subseteq> \<V> \<Longrightarrow> ps \<noteq> ps2 \<Longrightarrow> card ps = \<t> \<Longrightarrow> card ps2 = \<t>
- \<Longrightarrow> \<B> index ps = \<B> index ps2"
- using balanced by (metis of_nat_eq_iff)
-
-lemma t_wise_balance_alt: "ps \<subseteq> \<V> \<Longrightarrow> card ps = \<t> \<Longrightarrow> \<B> index ps = l2
- \<Longrightarrow> (\<And> ps . ps \<subseteq> \<V> \<Longrightarrow> card ps = \<t> \<Longrightarrow> \<B> index ps = l2)"
- using twise_set_pair_index by blast
-
-lemma index_ge_zero: "\<Lambda>\<^sub>t \<ge> 0"
-proof -
- obtain ps where "ps \<subseteq> \<V> \<and> card ps = \<t>" using t_non_zero t_lt_order obtain_subset_with_card_n
- by (metis dual_order.trans of_nat_le_iff zero_le_imp_eq_int zero_le_one)
- thus ?thesis
- using balanced_alt_def_all of_nat_0_le_iff by blast
-qed
-
-lemma index_1_imp_mult_1 [simp]:
- assumes "\<Lambda>\<^sub>t = 1"
- assumes "bl \<in># \<B>"
- assumes "card bl \<ge> \<t>"
- shows "multiplicity bl = 1"
-proof (rule ccontr)
- assume "\<not> (multiplicity bl = 1)"
- then have not: "multiplicity bl \<noteq> 1" by simp
- have "multiplicity bl \<noteq> 0" using assms by simp
- then have m: "multiplicity bl \<ge> 2" using not by linarith
- obtain ps where ps: "ps \<subseteq> bl \<and> card ps = \<t>"
- using assms obtain_t_subset_points
- by (metis obtain_subset_with_card_int_n of_nat_0_le_iff)
- then have "\<B> index ps \<ge> 2"
- using m points_index_count_min ps by blast
- then show False using balanced ps antisym_conv2 not_numeral_less_zero numeral_le_one_iff
- points_index_ps_nin semiring_norm(69) zero_neq_numeral
- by (metis assms(1) int_int_eq int_ops(2))
-qed
-
-end
-
-subsubsection \<open>Sub-types of t-wise balance\<close>
-
-text \<open>Pairwise balance is when $t = 2$. These are commonly of interest\<close>
-locale pairwise_balance = t_wise_balance \<V> \<B> 2 \<Lambda>
- for point_set ("\<V>") and block_collection ("\<B>") and index ("\<Lambda>")
-
-text \<open>We can combine the balance properties with $K$\_block design to define tBD's
-(t-wise balanced designs), and PBD's (pairwise balanced designs)\<close>
-
-locale tBD = t_wise_balance + K_block_design +
- assumes block_size_gt_t: "k \<in> \<K> \<Longrightarrow> k \<ge> \<t>"
-
-locale \<Lambda>_PBD = pairwise_balance + K_block_design +
- assumes block_size_gt_t: "k \<in> \<K> \<Longrightarrow> k \<ge> 2"
-
-sublocale \<Lambda>_PBD \<subseteq> tBD \<V> \<B> 2 \<Lambda> \<K>
- using t_lt_order block_size_gt_t by (unfold_locales) (simp_all)
-
-locale PBD = \<Lambda>_PBD \<V> \<B> 1 \<K> for point_set ("\<V>") and block_collection ("\<B>") and sizes ("\<K>")
-begin
-lemma multiplicity_is_1:
- assumes "bl \<in># \<B>"
- shows "multiplicity bl = 1"
- using block_size_gt_t index_1_imp_mult_1 by (simp add: assms block_sizes)
-
-end
-
-sublocale PBD \<subseteq> simple_design
- using multiplicity_is_1 by (unfold_locales)
-
-text \<open>PBD's are often only used in the case where $k$ is uniform, defined here.\<close>
-locale k_\<Lambda>_PBD = pairwise_balance + block_design +
- assumes block_size_t: "2 \<le> \<k>"
-
-sublocale k_\<Lambda>_PBD \<subseteq> \<Lambda>_PBD \<V> \<B> \<Lambda> "{\<k>}"
- using k_non_zero uniform block_size_t by(unfold_locales) (simp_all)
-
-locale k_PBD = k_\<Lambda>_PBD \<V> \<B> 1 \<k> for point_set ("\<V>") and block_collection ("\<B>") and u_block_size ("\<k>")
-
-sublocale k_PBD \<subseteq> PBD \<V> \<B> "{\<k>}"
- using block_size_t by (unfold_locales, simp_all)
-
-subsubsection \<open>Covering and Packing Designs\<close>
-text \<open>Covering and packing designs involve a looser balance restriction. Upper/lower bounds
-are placed on the points index, instead of a strict equality\<close>
-
-text \<open>A t-covering design is a relaxed version of a tBD, where, for all point subsets of size t,
-a lower bound is put on the points index\<close>
-locale t_covering_design = block_design +
- fixes grouping :: int ("\<t>")
- fixes min_index :: int ("\<Lambda>\<^sub>t")
- assumes covering: "ps \<subseteq> \<V> \<Longrightarrow> card ps = \<t> \<Longrightarrow> \<B> index ps \<ge> \<Lambda>\<^sub>t"
- assumes block_size_t: "\<t> \<le> \<k>"
- assumes t_non_zero: "\<t> \<ge> 1"
-begin
-
-lemma covering_alt_def_all: "\<forall> ps \<subseteq> \<V> . card ps = \<t> \<longrightarrow> \<B> index ps \<ge> \<Lambda>\<^sub>t"
- using covering by auto
-
-end
-
-lemma (in block_design) t_covering_designI [intro]: "t \<le> \<k> \<Longrightarrow> t \<ge> 1 \<Longrightarrow>
- (\<And> ps. ps \<subseteq> \<V> \<Longrightarrow> card ps = t \<Longrightarrow> \<B> index ps \<ge> \<Lambda>\<^sub>t) \<Longrightarrow> t_covering_design \<V> \<B> \<k> t \<Lambda>\<^sub>t"
- by (unfold_locales) simp_all
-
-text \<open>A t-packing design is a relaxed version of a tBD, where, for all point subsets of size t,
-an upper bound is put on the points index\<close>
-locale t_packing_design = block_design +
- fixes grouping :: int ("\<t>")
- fixes min_index :: int ("\<Lambda>\<^sub>t")
- assumes packing: "ps \<subseteq> \<V> \<Longrightarrow> card ps = \<t> \<Longrightarrow> \<B> index ps \<le> \<Lambda>\<^sub>t"
- assumes block_size_t: "\<t> \<le> \<k>"
- assumes t_non_zero: "\<t> \<ge> 1"
-begin
-
-lemma packing_alt_def_all: "\<forall> ps \<subseteq> \<V> . card ps = \<t> \<longrightarrow> \<B> index ps \<le> \<Lambda>\<^sub>t"
- using packing by auto
-
-end
-
-lemma (in block_design) t_packing_designI [intro]: "t \<le> \<k> \<Longrightarrow> t \<ge> 1 \<Longrightarrow>
- (\<And> ps . ps \<subseteq> \<V> \<Longrightarrow> card ps = t \<Longrightarrow> \<B> index ps \<le> \<Lambda>\<^sub>t) \<Longrightarrow> t_packing_design \<V> \<B> \<k> t \<Lambda>\<^sub>t"
- by (unfold_locales) simp_all
-
-lemma packing_covering_imp_balance:
- assumes "t_packing_design V B k t \<Lambda>\<^sub>t"
- assumes "t_covering_design V B k t \<Lambda>\<^sub>t"
- shows "t_wise_balance V B t \<Lambda>\<^sub>t"
-proof -
- from assms interpret des: proper_design V B
- using block_design.axioms(1) t_covering_design.axioms(1) by blast
- show ?thesis
- proof (unfold_locales)
- show "1 \<le> t" using assms by (simp add: t_packing_design.t_non_zero)
- show "t \<le> des.\<v>" using block_design.block_size_lt_v t_packing_design.axioms(1)
- by (metis assms(1) dual_order.trans t_packing_design.block_size_t)
- show "\<And>ps. ps \<subseteq> V \<Longrightarrow> card ps = t \<Longrightarrow> B index ps = \<Lambda>\<^sub>t"
- using t_packing_design.packing t_covering_design.covering by (metis assms dual_order.antisym)
- qed
-qed
-
-subsection \<open>Constant Replication Design\<close>
-text \<open>When the replication number for all points in a design is constant, it is the
-design replication number.\<close>
-locale constant_rep_design = proper_design +
- fixes design_rep_number :: int ("\<r>")
- assumes rep_number [simp]: "x \<in> \<V> \<Longrightarrow> \<B> rep x = \<r>"
-
-begin
-
-lemma rep_number_alt_def_all: "\<forall> x \<in> \<V>. \<B> rep x = \<r>"
- by (simp)
-
-lemma rep_number_unfold_set: "x \<in> \<V> \<Longrightarrow> size {#bl \<in># \<B> . x \<in> bl#} = \<r>"
- using rep_number by (simp add: point_replication_number_def)
-
-lemma rep_numbers_constant [simp]: "replication_numbers = {\<r>}"
- unfolding replication_numbers_def using rep_number design_points_nempty Collect_cong finite.cases
- finite_sets insertCI singleton_conv
- by (smt (verit, ccfv_threshold) fst_conv snd_conv)
-
-lemma replication_number_single: "is_singleton (replication_numbers)"
- using is_singleton_the_elem by simp
-
-lemma constant_rep_point_pair: "x1 \<in> \<V> \<Longrightarrow> x2 \<in> \<V> \<Longrightarrow> x1 \<noteq> x2 \<Longrightarrow> \<B> rep x1 = \<B> rep x2"
- using rep_number by auto
-
-lemma constant_rep_alt: "x1 \<in> \<V> \<Longrightarrow> \<B> rep x1 = r2 \<Longrightarrow> (\<And> x . x \<in> \<V> \<Longrightarrow> \<B> rep x = r2)"
- by (simp)
-
-lemma constant_rep_point_not_0:
- assumes "x \<in> \<V>"
- shows "\<B> rep x \<noteq> 0"
-proof (rule ccontr)
- assume "\<not> \<B> rep x \<noteq> 0"
- then have "\<And> x . x \<in> \<V> \<Longrightarrow> \<B> rep x = 0" using rep_number assms by auto
- then have "\<And> x . x \<in> \<V> \<Longrightarrow> size {#bl \<in># \<B> . x \<in> bl#} = 0"
- by (simp add: point_replication_number_def)
- then show False using design_blocks_nempty wf_design wf_design_iff wf_invalid_point
- by (metis ex_in_conv filter_mset_empty_conv multiset_nonemptyE size_eq_0_iff_empty)
-qed
-
-lemma rep_not_zero: "\<r> \<noteq> 0"
- using rep_number constant_rep_point_not_0 design_points_nempty by auto
-
-lemma r_gzero: "\<r> > 0"
- using point_replication_number_def rep_number constant_rep_design.rep_not_zero
- by (metis constant_rep_design.intro constant_rep_design_axioms.intro leI of_nat_less_0_iff
- proper_design_axioms verit_la_disequality)
-
-lemma r_lt_eq_b: "\<r> \<le> \<b>"
- using rep_number max_point_rep
- by (metis all_not_in_conv design_points_nempty)
-
-lemma complement_rep_number:
- assumes "\<And> bl . bl \<in># \<B> \<Longrightarrow> incomplete_block bl"
- shows "constant_rep_design \<V> \<B>\<^sup>C (\<b> - \<r>)"
-proof -
- interpret d: proper_design \<V> "(\<B>\<^sup>C)" using complement_proper_design
- by (simp add: assms)
- show ?thesis using complement_rep_number rep_number by (unfold_locales) simp
-qed
-
-lemma multiple_rep_number:
- assumes "n > 0"
- shows "constant_rep_design \<V> (multiple_blocks n) (\<r> * n)"
-proof -
- interpret d: proper_design \<V> "(multiple_blocks n)" using multiple_proper_design
- by (simp add: assms)
- show ?thesis using multiple_point_rep_num by (unfold_locales) (simp_all)
-qed
-end
-
-lemma (in proper_design) constant_rep_designI [intro]: "(\<And> x . x \<in> \<V> \<Longrightarrow> \<B> rep x = \<r>)
- \<Longrightarrow> constant_rep_design \<V> \<B> \<r>"
- by unfold_locales auto
-
-subsection \<open>T-designs\<close>
-text \<open>All the before mentioned designs build up to the concept of a t-design, which has uniform
-block size and is t-wise balanced. We limit $t$ to be less than $k$, so the balance condition has
-relevance\<close>
-locale t_design = incomplete_design + t_wise_balance +
- assumes block_size_t: "\<t> \<le> \<k>"
-begin
-
-lemma point_indices_balanced: "point_indices \<t> = {\<Lambda>\<^sub>t}"
-proof -
- have "point_indices \<t> = {i . \<exists> ps . i = \<B> index ps \<and> int (card ps) = \<t> \<and> ps \<subseteq> \<V>}"
- by (simp add: point_indices_def)
- then have "point_indices \<t> = {i . i = \<Lambda>\<^sub>t}" using balanced Collect_cong obtain_t_subset_points
- by smt
- thus ?thesis by auto
-qed
-
-lemma point_indices_singleton: "is_singleton (point_indices \<t>)"
- using point_indices_balanced is_singleton_the_elem by simp
-
-end
-
-lemma t_designI [intro]:
- assumes "incomplete_design V B k"
- assumes "t_wise_balance V B t \<Lambda>\<^sub>t"
- assumes "t \<le> k"
- shows "t_design V B k t \<Lambda>\<^sub>t"
- by (simp add: assms(1) assms(2) assms(3) t_design.intro t_design_axioms.intro)
-
-sublocale t_design \<subseteq> t_covering_design \<V> \<B> \<k> \<t> \<Lambda>\<^sub>t
- using t_non_zero by (unfold_locales) (auto simp add: block_size_t)
-
-sublocale t_design \<subseteq> t_packing_design \<V> \<B> \<k> \<t> \<Lambda>\<^sub>t
- using t_non_zero by (unfold_locales) (auto simp add: block_size_t)
-
-lemma t_design_pack_cov [intro]:
- assumes "k < card V"
- assumes "t_covering_design V B k t \<Lambda>\<^sub>t"
- assumes "t_packing_design V B k t \<Lambda>\<^sub>t"
- shows "t_design V B k t \<Lambda>\<^sub>t"
-proof -
- from assms interpret id: incomplete_design V B k
- using block_design.incomplete_designI t_packing_design.axioms(1)
- by (metis of_nat_less_iff)
- from assms interpret balance: t_wise_balance V B t \<Lambda>\<^sub>t
- using packing_covering_imp_balance by blast
- show ?thesis using assms(3)
- by (unfold_locales) (simp_all add: t_packing_design.block_size_t)
-qed
-
-sublocale t_design \<subseteq> tBD \<V> \<B> \<t> \<Lambda>\<^sub>t "{\<k>}"
- using uniform k_non_zero block_size_t by (unfold_locales) simp_all
-
-context t_design
-begin
-
-lemma multiple_t_design: "n > 0 \<Longrightarrow> t_design \<V> (multiple_blocks n) \<k> \<t> (\<Lambda>\<^sub>t * n)"
- using multiple_t_wise_balance multiple_incomplete block_size_t by (simp add: t_designI)
-
-lemma t_design_min_v: "\<v> > 1"
- using k_non_zero incomplete by simp
-
-end
-
-subsection \<open>Steiner Systems\<close>
-
-text \<open>Steiner systems are a special type of t-design where $\Lambda_t = 1$\<close>
-locale steiner_system = t_design \<V> \<B> \<k> \<t> 1
- for point_set ("\<V>") and block_collection ("\<B>") and u_block_size ("\<k>") and grouping ("\<t>")
-
-begin
-
-lemma block_multiplicity [simp]:
- assumes "bl \<in># \<B>"
- shows "multiplicity bl = 1"
- by (simp add: assms block_size_t)
-
-end
-
-sublocale steiner_system \<subseteq> simple_design
- by unfold_locales (simp)
-
-lemma (in t_design) steiner_systemI[intro]: "\<Lambda>\<^sub>t = 1 \<Longrightarrow> steiner_system \<V> \<B> \<k> \<t>"
- using t_non_zero t_lt_order block_size_t
- by unfold_locales auto
-
-subsection \<open>Combining block designs\<close>
-text \<open>We define some closure properties for various block designs under the combine operator.
-This is done using locales to reason on multiple instances of the same type of design, building
-on what was presented in the design operations theory\<close>
-
-locale two_t_wise_eq_points = two_designs_proper \<V> \<B> \<V> \<B>' + des1: t_wise_balance \<V> \<B> \<t> \<Lambda>\<^sub>t +
- des2: t_wise_balance \<V> \<B>' \<t> \<Lambda>\<^sub>t' for \<V> \<B> \<t> \<Lambda>\<^sub>t \<B>' \<Lambda>\<^sub>t'
-begin
-
-lemma combine_t_wise_balance_index: "ps \<subseteq> \<V> \<Longrightarrow> card ps = \<t> \<Longrightarrow> \<B>\<^sup>+ index ps = (\<Lambda>\<^sub>t + \<Lambda>\<^sub>t')"
- using des1.balanced des2.balanced by (simp add: combine_points_index)
-
-lemma combine_t_wise_balance: "t_wise_balance \<V>\<^sup>+ \<B>\<^sup>+ \<t> (\<Lambda>\<^sub>t + \<Lambda>\<^sub>t')"
-proof (unfold_locales, simp add: des1.t_non_zero)
- have "card \<V>\<^sup>+ \<ge> card \<V>" by simp
- then show "\<t> \<le> card (\<V>\<^sup>+)" using des1.t_lt_order by linarith
- show "\<And>ps. ps \<subseteq> \<V>\<^sup>+ \<Longrightarrow> card ps = \<t> \<Longrightarrow> (\<B>\<^sup>+ index ps) = \<Lambda>\<^sub>t + \<Lambda>\<^sub>t'"
- using combine_t_wise_balance_index by blast
-qed
-
-sublocale combine_t_wise_des: t_wise_balance "\<V>\<^sup>+" "\<B>\<^sup>+" "\<t>" "(\<Lambda>\<^sub>t + \<Lambda>\<^sub>t')"
- using combine_t_wise_balance by auto
-
-end
-
-locale two_k_block_designs = two_designs_proper \<V> \<B> \<V>' \<B>' + des1: block_design \<V> \<B> \<k> +
- des2: block_design \<V>' \<B>' \<k> for \<V> \<B> \<k> \<V>' \<B>'
-begin
-
-lemma block_design_combine: "block_design \<V>\<^sup>+ \<B>\<^sup>+ \<k>"
- using des1.uniform des2.uniform by (unfold_locales) (auto)
-
-sublocale combine_block_des: block_design "\<V>\<^sup>+" "\<B>\<^sup>+" "\<k>"
- using block_design_combine by simp
-
-end
-
-locale two_rep_designs_eq_points = two_designs_proper \<V> \<B> \<V> \<B>' + des1: constant_rep_design \<V> \<B> \<r> +
- des2: constant_rep_design \<V> \<B>' \<r>' for \<V> \<B> \<r> \<B>' \<r>'
-begin
-
-lemma combine_rep_number: "constant_rep_design \<V>\<^sup>+ \<B>\<^sup>+ (\<r> + \<r>')"
- using combine_rep_number des1.rep_number des2.rep_number by (unfold_locales) (simp)
-
-sublocale combine_const_rep: constant_rep_design "\<V>\<^sup>+" "\<B>\<^sup>+" "(\<r> + \<r>')"
- using combine_rep_number by simp
-
-end
-
-locale two_incomplete_designs = two_k_block_designs \<V> \<B> \<k> \<V>' \<B>' + des1: incomplete_design \<V> \<B> \<k> +
- des2: incomplete_design \<V>' \<B>' \<k> for \<V> \<B> \<k> \<V>' \<B>'
-begin
-
-lemma combine_is_incomplete: "incomplete_design \<V>\<^sup>+ \<B>\<^sup>+ \<k>"
- using combine_order des1.incomplete des2.incomplete by (unfold_locales) (simp)
-
-sublocale combine_incomplete: incomplete_design "\<V>\<^sup>+" "\<B>\<^sup>+" "\<k>"
- using combine_is_incomplete by simp
-end
-
-locale two_t_designs_eq_points = two_incomplete_designs \<V> \<B> \<k> \<V> \<B>'
- + two_t_wise_eq_points \<V> \<B> \<t> \<Lambda>\<^sub>t \<B>' \<Lambda>\<^sub>t' + des1: t_design \<V> \<B> \<k> \<t> \<Lambda>\<^sub>t +
- des2: t_design \<V> \<B>' \<k> \<t> \<Lambda>\<^sub>t' for \<V> \<B> \<k> \<B>' \<t> \<Lambda>\<^sub>t \<Lambda>\<^sub>t'
-begin
-
-lemma combine_is_t_des: "t_design \<V>\<^sup>+ \<B>\<^sup>+ \<k> \<t> (\<Lambda>\<^sub>t + \<Lambda>\<^sub>t')"
- using des1.block_size_t des2.block_size_t by (unfold_locales)
-
-sublocale combine_t_des: t_design "\<V>\<^sup>+" "\<B>\<^sup>+" "\<k>" "\<t>" "(\<Lambda>\<^sub>t + \<Lambda>\<^sub>t')"
- using combine_is_t_des by blast
-
-end
+(* Title: Block_Designs.thy
+ Author: Chelsea Edmonds
+*)
+
+section \<open>Block and Balanced Designs\<close>
+text \<open>We define a selection of the many different types of block and balanced designs, building up
+to properties required for defining a BIBD, in addition to several base generalisations\<close>
+
+theory Block_Designs imports Design_Operations
+begin
+
+subsection \<open>Block Designs\<close>
+text \<open>A block design is a design where all blocks have the same size.\<close>
+
+subsubsection \<open>K Block Designs\<close>
+text \<open>An important generalisation of a typical block design is the $\mathcal{K}$ block design,
+where all blocks must have a size $x$ where $x \in \mathcal{K}$\<close>
+locale K_block_design = proper_design +
+ fixes sizes :: "int set" ("\<K>")
+ assumes block_sizes: "bl \<in># \<B> \<Longrightarrow> (int (card bl)) \<in> \<K>"
+ assumes positive_ints: "x \<in> \<K> \<Longrightarrow> x > 0"
+begin
+
+lemma sys_block_size_subset: "sys_block_sizes \<subseteq> \<K>"
+ using block_sizes sys_block_sizes_obtain_bl by blast
+
+end
+
+subsubsection\<open>Uniform Block Design\<close>
+text \<open>The typical uniform block design is defined below\<close>
+locale block_design = proper_design +
+ fixes u_block_size :: int ("\<k>")
+ assumes uniform [simp]: "bl \<in># \<B> \<Longrightarrow> card bl = \<k>"
+begin
+
+lemma k_non_zero: "\<k> \<ge> 1"
+proof -
+ obtain bl where bl_in: "bl \<in># \<B>"
+ using design_blocks_nempty by auto
+ then have "int (card bl) \<ge> 1" using block_size_gt_0
+ by (metis less_not_refl less_one not_le_imp_less of_nat_1 of_nat_less_iff)
+ thus ?thesis by (simp add: bl_in)
+qed
+
+lemma uniform_alt_def_all: "\<forall> bl \<in># \<B> .card bl = \<k>"
+ using uniform by auto
+
+lemma uniform_unfold_point_set: "bl \<in># \<B> \<Longrightarrow> card {p \<in> \<V>. p \<in> bl} = \<k>"
+ using uniform wellformed by (simp add: Collect_conj_eq inf.absorb_iff2)
+
+lemma uniform_unfold_point_set_mset: "bl \<in># \<B> \<Longrightarrow> size {#p \<in># mset_set \<V>. p \<in> bl #} = \<k>"
+ using uniform_unfold_point_set by (simp add: finite_sets)
+
+lemma sys_block_sizes_uniform [simp]: "sys_block_sizes = {\<k>}"
+proof -
+ have "sys_block_sizes = {bs . \<exists> bl . bs = card bl \<and> bl\<in># \<B>}" by (simp add: sys_block_sizes_def)
+ then have "sys_block_sizes = {bs . bs = \<k>}" using uniform uniform_unfold_point_set
+ b_positive block_set_nempty_imp_block_ex
+ by (smt (verit, best) Collect_cong design_blocks_nempty)
+ thus ?thesis by auto
+qed
+
+lemma sys_block_sizes_uniform_single: "is_singleton (sys_block_sizes)"
+ by simp
+
+lemma uniform_size_incomp: "\<k> \<le> \<v> - 1 \<Longrightarrow> bl \<in># \<B> \<Longrightarrow> incomplete_block bl"
+ using uniform k_non_zero of_nat_less_iff zle_diff1_eq by metis
+
+lemma uniform_complement_block_size:
+ assumes "bl \<in># \<B>\<^sup>C"
+ shows "card bl = \<v> - \<k>"
+proof -
+ obtain bl' where bl_assm: "bl = bl'\<^sup>c \<and> bl' \<in># \<B>"
+ using wellformed assms by (auto simp add: complement_blocks_def)
+ then have "int (card bl') = \<k>" by simp
+ thus ?thesis using bl_assm block_complement_size wellformed
+ by (simp add: block_size_lt_order of_nat_diff)
+qed
+
+lemma uniform_complement[intro]:
+ assumes "\<k> \<le> \<v> - 1"
+ shows "block_design \<V> \<B>\<^sup>C (\<v> - \<k>)"
+proof -
+ interpret des: proper_design \<V> "\<B>\<^sup>C"
+ using uniform_size_incomp assms complement_proper_design by auto
+ show ?thesis using assms uniform_complement_block_size by (unfold_locales) (simp)
+qed
+
+lemma block_size_lt_v: "\<k> \<le> \<v>"
+ using v_non_zero block_size_lt_v design_blocks_nempty uniform by auto
+
+end
+
+lemma (in proper_design) block_designI[intro]: "(\<And> bl . bl \<in># \<B> \<Longrightarrow> card bl = k)
+ \<Longrightarrow> block_design \<V> \<B> k"
+ by (unfold_locales) (auto)
+
+context block_design
+begin
+
+lemma block_design_multiple: "n > 0 \<Longrightarrow> block_design \<V> (multiple_blocks n) \<k>"
+ using elem_in_repeat_in_original multiple_proper_design proper_design.block_designI
+ by (metis block_set_nempty_imp_block_ex design_blocks_nempty int_int_eq uniform_alt_def_all)
+
+end
+text \<open>A uniform block design is clearly a type of $K$\_block\_design with a singleton $K$ set\<close>
+sublocale block_design \<subseteq> K_block_design \<V> \<B> "{\<k>}"
+ using k_non_zero uniform by unfold_locales simp_all
+
+subsubsection \<open>Incomplete Designs\<close>
+text \<open>An incomplete design is a design where $k < v$, i.e. no block is equal to the point set\<close>
+locale incomplete_design = block_design +
+ assumes incomplete: "\<k> < \<v>"
+
+begin
+
+lemma incomplete_imp_incomp_block: "bl \<in># \<B> \<Longrightarrow> incomplete_block bl"
+ using incomplete uniform uniform_size_incomp by fastforce
+
+lemma incomplete_imp_proper_subset: "bl \<in># \<B> \<Longrightarrow> bl \<subset> \<V>"
+ by (simp add: incomplete_block_proper_subset incomplete_imp_incomp_block wellformed)
+end
+
+lemma (in block_design) incomplete_designI[intro]: "\<k> < \<v> \<Longrightarrow> incomplete_design \<V> \<B> \<k>"
+ by unfold_locales auto
+
+context incomplete_design
+begin
+
+lemma multiple_incomplete: "n > 0 \<Longrightarrow> incomplete_design \<V> (multiple_blocks n) \<k>"
+ using block_design_multiple incomplete by (simp add: block_design.incomplete_designI)
+
+lemma complement_incomplete: "incomplete_design \<V> (\<B>\<^sup>C) (\<v> - \<k>)"
+proof -
+ have "\<v> - \<k> < \<v>" using v_non_zero k_non_zero by linarith
+ thus ?thesis using uniform_complement incomplete incomplete_designI
+ by (simp add: block_design.incomplete_designI)
+qed
+
+end
+
+subsection \<open>Balanced Designs\<close>
+text \<open>t-wise balance is a design with the property that all point subsets of size $t$ occur in
+$\lambda_t$ blocks\<close>
+
+locale t_wise_balance = proper_design +
+ fixes grouping :: int ("\<t>") and index :: int ("\<Lambda>\<^sub>t")
+ assumes t_non_zero: "\<t> \<ge> 1"
+ assumes t_lt_order: "\<t> \<le> \<v>"
+ assumes balanced [simp]: "ps \<subseteq> \<V> \<Longrightarrow> card ps = \<t> \<Longrightarrow> \<B> index ps = \<Lambda>\<^sub>t"
+begin
+
+lemma balanced_alt_def_all: "\<forall> ps \<subseteq> \<V> . card ps = \<t> \<longrightarrow> \<B> index ps = \<Lambda>\<^sub>t"
+ using balanced by auto
+
+end
+
+lemma (in proper_design) t_wise_balanceI[intro]: "\<t> \<le> \<v> \<Longrightarrow> \<t> \<ge> 1 \<Longrightarrow>
+ (\<And> ps . ps \<subseteq> \<V> \<Longrightarrow> card ps = \<t> \<Longrightarrow> \<B> index ps = \<Lambda>\<^sub>t) \<Longrightarrow> t_wise_balance \<V> \<B> \<t> \<Lambda>\<^sub>t"
+ by (unfold_locales) auto
+
+context t_wise_balance
+begin
+
+lemma obtain_t_subset_points:
+ obtains T where "T \<subseteq> \<V>" "card T = \<t>" "finite T"
+ using obtain_subset_with_card_int_n design_points_nempty t_lt_order t_non_zero finite_sets
+ by (metis int_one_le_iff_zero_less less_eq_int_code(1) linorder_not_le verit_la_generic)
+
+lemma multiple_t_wise_balance_index [simp]:
+ assumes "ps \<subseteq> \<V>"
+ assumes "card ps = \<t>"
+ shows "(multiple_blocks n) index ps = \<Lambda>\<^sub>t * n"
+ using multiple_point_index balanced assms by fastforce
+
+lemma multiple_t_wise_balance:
+ assumes "n > 0"
+ shows "t_wise_balance \<V> (multiple_blocks n) \<t> (\<Lambda>\<^sub>t * n)"
+proof -
+ interpret des: proper_design \<V> "(multiple_blocks n)" by (simp add: assms multiple_proper_design)
+ show ?thesis using t_non_zero t_lt_order multiple_t_wise_balance_index
+ by (unfold_locales) (simp_all)
+qed
+
+lemma twise_set_pair_index: "ps \<subseteq> \<V> \<Longrightarrow> ps2 \<subseteq> \<V> \<Longrightarrow> ps \<noteq> ps2 \<Longrightarrow> card ps = \<t> \<Longrightarrow> card ps2 = \<t>
+ \<Longrightarrow> \<B> index ps = \<B> index ps2"
+ using balanced by (metis of_nat_eq_iff)
+
+lemma t_wise_balance_alt: "ps \<subseteq> \<V> \<Longrightarrow> card ps = \<t> \<Longrightarrow> \<B> index ps = l2
+ \<Longrightarrow> (\<And> ps . ps \<subseteq> \<V> \<Longrightarrow> card ps = \<t> \<Longrightarrow> \<B> index ps = l2)"
+ using twise_set_pair_index by blast
+
+lemma index_ge_zero: "\<Lambda>\<^sub>t \<ge> 0"
+proof -
+ obtain ps where "ps \<subseteq> \<V> \<and> card ps = \<t>" using t_non_zero t_lt_order obtain_subset_with_card_n
+ by (metis dual_order.trans of_nat_le_iff zero_le_imp_eq_int zero_le_one)
+ thus ?thesis
+ using balanced_alt_def_all of_nat_0_le_iff by blast
+qed
+
+lemma index_1_imp_mult_1 [simp]:
+ assumes "\<Lambda>\<^sub>t = 1"
+ assumes "bl \<in># \<B>"
+ assumes "card bl \<ge> \<t>"
+ shows "multiplicity bl = 1"
+proof (rule ccontr)
+ assume "\<not> (multiplicity bl = 1)"
+ then have not: "multiplicity bl \<noteq> 1" by simp
+ have "multiplicity bl \<noteq> 0" using assms by simp
+ then have m: "multiplicity bl \<ge> 2" using not by linarith
+ obtain ps where ps: "ps \<subseteq> bl \<and> card ps = \<t>"
+ using assms obtain_t_subset_points
+ by (metis obtain_subset_with_card_int_n of_nat_0_le_iff)
+ then have "\<B> index ps \<ge> 2"
+ using m points_index_count_min ps by blast
+ then show False using balanced ps antisym_conv2 not_numeral_less_zero numeral_le_one_iff
+ points_index_ps_nin semiring_norm(69) zero_neq_numeral
+ by (metis assms(1) int_int_eq int_ops(2))
+qed
+
+end
+
+subsubsection \<open>Sub-types of t-wise balance\<close>
+
+text \<open>Pairwise balance is when $t = 2$. These are commonly of interest\<close>
+locale pairwise_balance = t_wise_balance \<V> \<B> 2 \<Lambda>
+ for point_set ("\<V>") and block_collection ("\<B>") and index ("\<Lambda>")
+
+text \<open>We can combine the balance properties with $K$\_block design to define tBD's
+(t-wise balanced designs), and PBD's (pairwise balanced designs)\<close>
+
+locale tBD = t_wise_balance + K_block_design +
+ assumes block_size_gt_t: "k \<in> \<K> \<Longrightarrow> k \<ge> \<t>"
+
+locale \<Lambda>_PBD = pairwise_balance + K_block_design +
+ assumes block_size_gt_t: "k \<in> \<K> \<Longrightarrow> k \<ge> 2"
+
+sublocale \<Lambda>_PBD \<subseteq> tBD \<V> \<B> 2 \<Lambda> \<K>
+ using t_lt_order block_size_gt_t by (unfold_locales) (simp_all)
+
+locale PBD = \<Lambda>_PBD \<V> \<B> 1 \<K> for point_set ("\<V>") and block_collection ("\<B>") and sizes ("\<K>")
+begin
+lemma multiplicity_is_1:
+ assumes "bl \<in># \<B>"
+ shows "multiplicity bl = 1"
+ using block_size_gt_t index_1_imp_mult_1 by (simp add: assms block_sizes)
+
+end
+
+sublocale PBD \<subseteq> simple_design
+ using multiplicity_is_1 by (unfold_locales)
+
+text \<open>PBD's are often only used in the case where $k$ is uniform, defined here.\<close>
+locale k_\<Lambda>_PBD = pairwise_balance + block_design +
+ assumes block_size_t: "2 \<le> \<k>"
+
+sublocale k_\<Lambda>_PBD \<subseteq> \<Lambda>_PBD \<V> \<B> \<Lambda> "{\<k>}"
+ using k_non_zero uniform block_size_t by(unfold_locales) (simp_all)
+
+locale k_PBD = k_\<Lambda>_PBD \<V> \<B> 1 \<k> for point_set ("\<V>") and block_collection ("\<B>") and u_block_size ("\<k>")
+
+sublocale k_PBD \<subseteq> PBD \<V> \<B> "{\<k>}"
+ using block_size_t by (unfold_locales, simp_all)
+
+subsubsection \<open>Covering and Packing Designs\<close>
+text \<open>Covering and packing designs involve a looser balance restriction. Upper/lower bounds
+are placed on the points index, instead of a strict equality\<close>
+
+text \<open>A t-covering design is a relaxed version of a tBD, where, for all point subsets of size t,
+a lower bound is put on the points index\<close>
+locale t_covering_design = block_design +
+ fixes grouping :: int ("\<t>")
+ fixes min_index :: int ("\<Lambda>\<^sub>t")
+ assumes covering: "ps \<subseteq> \<V> \<Longrightarrow> card ps = \<t> \<Longrightarrow> \<B> index ps \<ge> \<Lambda>\<^sub>t"
+ assumes block_size_t: "\<t> \<le> \<k>"
+ assumes t_non_zero: "\<t> \<ge> 1"
+begin
+
+lemma covering_alt_def_all: "\<forall> ps \<subseteq> \<V> . card ps = \<t> \<longrightarrow> \<B> index ps \<ge> \<Lambda>\<^sub>t"
+ using covering by auto
+
+end
+
+lemma (in block_design) t_covering_designI [intro]: "t \<le> \<k> \<Longrightarrow> t \<ge> 1 \<Longrightarrow>
+ (\<And> ps. ps \<subseteq> \<V> \<Longrightarrow> card ps = t \<Longrightarrow> \<B> index ps \<ge> \<Lambda>\<^sub>t) \<Longrightarrow> t_covering_design \<V> \<B> \<k> t \<Lambda>\<^sub>t"
+ by (unfold_locales) simp_all
+
+text \<open>A t-packing design is a relaxed version of a tBD, where, for all point subsets of size t,
+an upper bound is put on the points index\<close>
+locale t_packing_design = block_design +
+ fixes grouping :: int ("\<t>")
+ fixes min_index :: int ("\<Lambda>\<^sub>t")
+ assumes packing: "ps \<subseteq> \<V> \<Longrightarrow> card ps = \<t> \<Longrightarrow> \<B> index ps \<le> \<Lambda>\<^sub>t"
+ assumes block_size_t: "\<t> \<le> \<k>"
+ assumes t_non_zero: "\<t> \<ge> 1"
+begin
+
+lemma packing_alt_def_all: "\<forall> ps \<subseteq> \<V> . card ps = \<t> \<longrightarrow> \<B> index ps \<le> \<Lambda>\<^sub>t"
+ using packing by auto
+
+end
+
+lemma (in block_design) t_packing_designI [intro]: "t \<le> \<k> \<Longrightarrow> t \<ge> 1 \<Longrightarrow>
+ (\<And> ps . ps \<subseteq> \<V> \<Longrightarrow> card ps = t \<Longrightarrow> \<B> index ps \<le> \<Lambda>\<^sub>t) \<Longrightarrow> t_packing_design \<V> \<B> \<k> t \<Lambda>\<^sub>t"
+ by (unfold_locales) simp_all
+
+lemma packing_covering_imp_balance:
+ assumes "t_packing_design V B k t \<Lambda>\<^sub>t"
+ assumes "t_covering_design V B k t \<Lambda>\<^sub>t"
+ shows "t_wise_balance V B t \<Lambda>\<^sub>t"
+proof -
+ from assms interpret des: proper_design V B
+ using block_design.axioms(1) t_covering_design.axioms(1) by blast
+ show ?thesis
+ proof (unfold_locales)
+ show "1 \<le> t" using assms by (simp add: t_packing_design.t_non_zero)
+ show "t \<le> des.\<v>" using block_design.block_size_lt_v t_packing_design.axioms(1)
+ by (metis assms(1) dual_order.trans t_packing_design.block_size_t)
+ show "\<And>ps. ps \<subseteq> V \<Longrightarrow> card ps = t \<Longrightarrow> B index ps = \<Lambda>\<^sub>t"
+ using t_packing_design.packing t_covering_design.covering by (metis assms dual_order.antisym)
+ qed
+qed
+
+subsection \<open>Constant Replication Design\<close>
+text \<open>When the replication number for all points in a design is constant, it is the
+design replication number.\<close>
+locale constant_rep_design = proper_design +
+ fixes design_rep_number :: int ("\<r>")
+ assumes rep_number [simp]: "x \<in> \<V> \<Longrightarrow> \<B> rep x = \<r>"
+
+begin
+
+lemma rep_number_alt_def_all: "\<forall> x \<in> \<V>. \<B> rep x = \<r>"
+ by (simp)
+
+lemma rep_number_unfold_set: "x \<in> \<V> \<Longrightarrow> size {#bl \<in># \<B> . x \<in> bl#} = \<r>"
+ using rep_number by (simp add: point_replication_number_def)
+
+lemma rep_numbers_constant [simp]: "replication_numbers = {\<r>}"
+ unfolding replication_numbers_def using rep_number design_points_nempty Collect_cong finite.cases
+ finite_sets insertCI singleton_conv
+ by (smt (verit, ccfv_threshold) fst_conv snd_conv)
+
+lemma replication_number_single: "is_singleton (replication_numbers)"
+ using is_singleton_the_elem by simp
+
+lemma constant_rep_point_pair: "x1 \<in> \<V> \<Longrightarrow> x2 \<in> \<V> \<Longrightarrow> x1 \<noteq> x2 \<Longrightarrow> \<B> rep x1 = \<B> rep x2"
+ using rep_number by auto
+
+lemma constant_rep_alt: "x1 \<in> \<V> \<Longrightarrow> \<B> rep x1 = r2 \<Longrightarrow> (\<And> x . x \<in> \<V> \<Longrightarrow> \<B> rep x = r2)"
+ by (simp)
+
+lemma constant_rep_point_not_0:
+ assumes "x \<in> \<V>"
+ shows "\<B> rep x \<noteq> 0"
+proof (rule ccontr)
+ assume "\<not> \<B> rep x \<noteq> 0"
+ then have "\<And> x . x \<in> \<V> \<Longrightarrow> \<B> rep x = 0" using rep_number assms by auto
+ then have "\<And> x . x \<in> \<V> \<Longrightarrow> size {#bl \<in># \<B> . x \<in> bl#} = 0"
+ by (simp add: point_replication_number_def)
+ then show False using design_blocks_nempty wf_design wf_design_iff wf_invalid_point
+ by (metis ex_in_conv filter_mset_empty_conv multiset_nonemptyE size_eq_0_iff_empty)
+qed
+
+lemma rep_not_zero: "\<r> \<noteq> 0"
+ using rep_number constant_rep_point_not_0 design_points_nempty by auto
+
+lemma r_gzero: "\<r> > 0"
+ using point_replication_number_def rep_number constant_rep_design.rep_not_zero
+ by (metis constant_rep_design.intro constant_rep_design_axioms.intro leI of_nat_less_0_iff
+ proper_design_axioms verit_la_disequality)
+
+lemma r_lt_eq_b: "\<r> \<le> \<b>"
+ using rep_number max_point_rep
+ by (metis all_not_in_conv design_points_nempty)
+
+lemma complement_rep_number:
+ assumes "\<And> bl . bl \<in># \<B> \<Longrightarrow> incomplete_block bl"
+ shows "constant_rep_design \<V> \<B>\<^sup>C (\<b> - \<r>)"
+proof -
+ interpret d: proper_design \<V> "(\<B>\<^sup>C)" using complement_proper_design
+ by (simp add: assms)
+ show ?thesis using complement_rep_number rep_number by (unfold_locales) simp
+qed
+
+lemma multiple_rep_number:
+ assumes "n > 0"
+ shows "constant_rep_design \<V> (multiple_blocks n) (\<r> * n)"
+proof -
+ interpret d: proper_design \<V> "(multiple_blocks n)" using multiple_proper_design
+ by (simp add: assms)
+ show ?thesis using multiple_point_rep_num by (unfold_locales) (simp_all)
+qed
+end
+
+lemma (in proper_design) constant_rep_designI [intro]: "(\<And> x . x \<in> \<V> \<Longrightarrow> \<B> rep x = \<r>)
+ \<Longrightarrow> constant_rep_design \<V> \<B> \<r>"
+ by unfold_locales auto
+
+subsection \<open>T-designs\<close>
+text \<open>All the before mentioned designs build up to the concept of a t-design, which has uniform
+block size and is t-wise balanced. We limit $t$ to be less than $k$, so the balance condition has
+relevance\<close>
+locale t_design = incomplete_design + t_wise_balance +
+ assumes block_size_t: "\<t> \<le> \<k>"
+begin
+
+lemma point_indices_balanced: "point_indices \<t> = {\<Lambda>\<^sub>t}"
+proof -
+ have "point_indices \<t> = {i . \<exists> ps . i = \<B> index ps \<and> int (card ps) = \<t> \<and> ps \<subseteq> \<V>}"
+ by (simp add: point_indices_def)
+ then have "point_indices \<t> = {i . i = \<Lambda>\<^sub>t}" using balanced Collect_cong obtain_t_subset_points
+ by smt
+ thus ?thesis by auto
+qed
+
+lemma point_indices_singleton: "is_singleton (point_indices \<t>)"
+ using point_indices_balanced is_singleton_the_elem by simp
+
+end
+
+lemma t_designI [intro]:
+ assumes "incomplete_design V B k"
+ assumes "t_wise_balance V B t \<Lambda>\<^sub>t"
+ assumes "t \<le> k"
+ shows "t_design V B k t \<Lambda>\<^sub>t"
+ by (simp add: assms(1) assms(2) assms(3) t_design.intro t_design_axioms.intro)
+
+sublocale t_design \<subseteq> t_covering_design \<V> \<B> \<k> \<t> \<Lambda>\<^sub>t
+ using t_non_zero by (unfold_locales) (auto simp add: block_size_t)
+
+sublocale t_design \<subseteq> t_packing_design \<V> \<B> \<k> \<t> \<Lambda>\<^sub>t
+ using t_non_zero by (unfold_locales) (auto simp add: block_size_t)
+
+lemma t_design_pack_cov [intro]:
+ assumes "k < card V"
+ assumes "t_covering_design V B k t \<Lambda>\<^sub>t"
+ assumes "t_packing_design V B k t \<Lambda>\<^sub>t"
+ shows "t_design V B k t \<Lambda>\<^sub>t"
+proof -
+ from assms interpret id: incomplete_design V B k
+ using block_design.incomplete_designI t_packing_design.axioms(1)
+ by (metis of_nat_less_iff)
+ from assms interpret balance: t_wise_balance V B t \<Lambda>\<^sub>t
+ using packing_covering_imp_balance by blast
+ show ?thesis using assms(3)
+ by (unfold_locales) (simp_all add: t_packing_design.block_size_t)
+qed
+
+sublocale t_design \<subseteq> tBD \<V> \<B> \<t> \<Lambda>\<^sub>t "{\<k>}"
+ using uniform k_non_zero block_size_t by (unfold_locales) simp_all
+
+context t_design
+begin
+
+lemma multiple_t_design: "n > 0 \<Longrightarrow> t_design \<V> (multiple_blocks n) \<k> \<t> (\<Lambda>\<^sub>t * n)"
+ using multiple_t_wise_balance multiple_incomplete block_size_t by (simp add: t_designI)
+
+lemma t_design_min_v: "\<v> > 1"
+ using k_non_zero incomplete by simp
+
+end
+
+subsection \<open>Steiner Systems\<close>
+
+text \<open>Steiner systems are a special type of t-design where $\Lambda_t = 1$\<close>
+locale steiner_system = t_design \<V> \<B> \<k> \<t> 1
+ for point_set ("\<V>") and block_collection ("\<B>") and u_block_size ("\<k>") and grouping ("\<t>")
+
+begin
+
+lemma block_multiplicity [simp]:
+ assumes "bl \<in># \<B>"
+ shows "multiplicity bl = 1"
+ by (simp add: assms block_size_t)
+
+end
+
+sublocale steiner_system \<subseteq> simple_design
+ by unfold_locales (simp)
+
+lemma (in t_design) steiner_systemI[intro]: "\<Lambda>\<^sub>t = 1 \<Longrightarrow> steiner_system \<V> \<B> \<k> \<t>"
+ using t_non_zero t_lt_order block_size_t
+ by unfold_locales auto
+
+subsection \<open>Combining block designs\<close>
+text \<open>We define some closure properties for various block designs under the combine operator.
+This is done using locales to reason on multiple instances of the same type of design, building
+on what was presented in the design operations theory\<close>
+
+locale two_t_wise_eq_points = two_designs_proper \<V> \<B> \<V> \<B>' + des1: t_wise_balance \<V> \<B> \<t> \<Lambda>\<^sub>t +
+ des2: t_wise_balance \<V> \<B>' \<t> \<Lambda>\<^sub>t' for \<V> \<B> \<t> \<Lambda>\<^sub>t \<B>' \<Lambda>\<^sub>t'
+begin
+
+lemma combine_t_wise_balance_index: "ps \<subseteq> \<V> \<Longrightarrow> card ps = \<t> \<Longrightarrow> \<B>\<^sup>+ index ps = (\<Lambda>\<^sub>t + \<Lambda>\<^sub>t')"
+ using des1.balanced des2.balanced by (simp add: combine_points_index)
+
+lemma combine_t_wise_balance: "t_wise_balance \<V>\<^sup>+ \<B>\<^sup>+ \<t> (\<Lambda>\<^sub>t + \<Lambda>\<^sub>t')"
+proof (unfold_locales, simp add: des1.t_non_zero)
+ have "card \<V>\<^sup>+ \<ge> card \<V>" by simp
+ then show "\<t> \<le> card (\<V>\<^sup>+)" using des1.t_lt_order by linarith
+ show "\<And>ps. ps \<subseteq> \<V>\<^sup>+ \<Longrightarrow> card ps = \<t> \<Longrightarrow> (\<B>\<^sup>+ index ps) = \<Lambda>\<^sub>t + \<Lambda>\<^sub>t'"
+ using combine_t_wise_balance_index by blast
+qed
+
+sublocale combine_t_wise_des: t_wise_balance "\<V>\<^sup>+" "\<B>\<^sup>+" "\<t>" "(\<Lambda>\<^sub>t + \<Lambda>\<^sub>t')"
+ using combine_t_wise_balance by auto
+
+end
+
+locale two_k_block_designs = two_designs_proper \<V> \<B> \<V>' \<B>' + des1: block_design \<V> \<B> \<k> +
+ des2: block_design \<V>' \<B>' \<k> for \<V> \<B> \<k> \<V>' \<B>'
+begin
+
+lemma block_design_combine: "block_design \<V>\<^sup>+ \<B>\<^sup>+ \<k>"
+ using des1.uniform des2.uniform by (unfold_locales) (auto)
+
+sublocale combine_block_des: block_design "\<V>\<^sup>+" "\<B>\<^sup>+" "\<k>"
+ using block_design_combine by simp
+
+end
+
+locale two_rep_designs_eq_points = two_designs_proper \<V> \<B> \<V> \<B>' + des1: constant_rep_design \<V> \<B> \<r> +
+ des2: constant_rep_design \<V> \<B>' \<r>' for \<V> \<B> \<r> \<B>' \<r>'
+begin
+
+lemma combine_rep_number: "constant_rep_design \<V>\<^sup>+ \<B>\<^sup>+ (\<r> + \<r>')"
+ using combine_rep_number des1.rep_number des2.rep_number by (unfold_locales) (simp)
+
+sublocale combine_const_rep: constant_rep_design "\<V>\<^sup>+" "\<B>\<^sup>+" "(\<r> + \<r>')"
+ using combine_rep_number by simp
+
+end
+
+locale two_incomplete_designs = two_k_block_designs \<V> \<B> \<k> \<V>' \<B>' + des1: incomplete_design \<V> \<B> \<k> +
+ des2: incomplete_design \<V>' \<B>' \<k> for \<V> \<B> \<k> \<V>' \<B>'
+begin
+
+lemma combine_is_incomplete: "incomplete_design \<V>\<^sup>+ \<B>\<^sup>+ \<k>"
+ using combine_order des1.incomplete des2.incomplete by (unfold_locales) (simp)
+
+sublocale combine_incomplete: incomplete_design "\<V>\<^sup>+" "\<B>\<^sup>+" "\<k>"
+ using combine_is_incomplete by simp
+end
+
+locale two_t_designs_eq_points = two_incomplete_designs \<V> \<B> \<k> \<V> \<B>'
+ + two_t_wise_eq_points \<V> \<B> \<t> \<Lambda>\<^sub>t \<B>' \<Lambda>\<^sub>t' + des1: t_design \<V> \<B> \<k> \<t> \<Lambda>\<^sub>t +
+ des2: t_design \<V> \<B>' \<k> \<t> \<Lambda>\<^sub>t' for \<V> \<B> \<k> \<B>' \<t> \<Lambda>\<^sub>t \<Lambda>\<^sub>t'
+begin
+
+lemma combine_is_t_des: "t_design \<V>\<^sup>+ \<B>\<^sup>+ \<k> \<t> (\<Lambda>\<^sub>t + \<Lambda>\<^sub>t')"
+ using des1.block_size_t des2.block_size_t by (unfold_locales)
+
+sublocale combine_t_des: t_design "\<V>\<^sup>+" "\<B>\<^sup>+" "\<k>" "\<t>" "(\<Lambda>\<^sub>t + \<Lambda>\<^sub>t')"
+ using combine_is_t_des by blast
+
+end
end
\ No newline at end of file
diff --git a/thys/Design_Theory/Design_Basics.thy b/thys/Design_Theory/Design_Basics.thy
--- a/thys/Design_Theory/Design_Basics.thy
+++ b/thys/Design_Theory/Design_Basics.thy
@@ -1,824 +1,824 @@
-theory Design_Basics imports Main Multisets_Extras "HOL-Library.Disjoint_Sets"
-begin
-
-section \<open>Design Theory Basics\<close>
-text \<open>All definitions in this section reference the handbook of combinatorial designs
- \cite{colbournHandbookCombinatorialDesigns2007}\<close>
-
-subsection \<open>Initial setup\<close>
-
-text \<open>Enable coercion of nats to ints to aid with reasoning on design properties\<close>
-declare [[coercion_enabled]]
-declare [[coercion "of_nat :: nat \<Rightarrow> int"]]
-
-subsection \<open>Incidence System\<close>
-
-text \<open>An incidence system is defined to be a wellformed set system. i.e. each block is a subset
-of the base point set. Alternatively, an incidence system can be looked at as the point set
-and an incidence relation which indicates if they are in the same block\<close>
-
-locale incidence_system =
- fixes point_set :: "'a set" ("\<V>")
- fixes block_collection :: "'a set multiset" ("\<B>")
- assumes wellformed: "b \<in># \<B> \<Longrightarrow> b \<subseteq> \<V>"
-begin
-
-definition "\<I> \<equiv> { (x, b) . b \<in># \<B> \<and> x \<in> b}" (* incidence relation *)
-
-definition incident :: "'a \<Rightarrow> 'a set \<Rightarrow> bool" where
-"incident p b \<equiv> (p, b) \<in> \<I>"
-
-text \<open>Defines common notation used to indicate number of points ($v$) and number of blocks ($b$)\<close>
-abbreviation "\<v> \<equiv> int (card \<V>)"
-
-abbreviation "\<b> \<equiv> int (size \<B>)"
-
-text \<open>Basic incidence lemmas\<close>
-
-lemma incidence_alt_def:
- assumes "p \<in> \<V>"
- assumes "b \<in># \<B>"
- shows "incident p b \<longleftrightarrow> p \<in> b"
- by (auto simp add: incident_def \<I>_def assms)
-
-lemma wf_invalid_point: "x \<notin> \<V> \<Longrightarrow> b \<in># \<B> \<Longrightarrow> x \<notin> b"
- using wellformed by auto
-
-lemma block_set_nempty_imp_block_ex: "\<B> \<noteq> {#} \<Longrightarrow> \<exists> bl . bl \<in># \<B>"
- by auto
-
-text \<open>Abbreviations for all incidence systems\<close>
-abbreviation multiplicity :: "'a set \<Rightarrow> nat" where
-"multiplicity b \<equiv> count \<B> b"
-
-abbreviation incomplete_block :: "'a set \<Rightarrow> bool" where
-"incomplete_block bl \<equiv> card bl < card \<V> \<and> bl \<in># \<B>"
-
-lemma incomplete_alt_size: "incomplete_block bl \<Longrightarrow> card bl < \<v>"
- by simp
-
-lemma incomplete_alt_in: "incomplete_block bl \<Longrightarrow> bl \<in># \<B>"
- by simp
-
-lemma incomplete_alt_imp[intro]: "card bl < \<v> \<Longrightarrow> bl \<in># \<B> \<Longrightarrow> incomplete_block bl"
- by simp
-
-definition design_support :: "'a set set" where
-"design_support \<equiv> set_mset \<B>"
-
-end
-
-subsection \<open>Finite Incidence Systems\<close>
-
-text \<open>These simply require the point set to be finite.
-As multisets are only defined to be finite, it is implied that the block set must be finite already\<close>
-
-locale finite_incidence_system = incidence_system +
- assumes finite_sets: "finite \<V>"
-begin
-
-lemma finite_blocks: "b \<in># \<B> \<Longrightarrow> finite b"
- using wellformed finite_sets finite_subset by blast
-
-lemma mset_points_distinct: "distinct_mset (mset_set \<V>)"
- using finite_sets by (simp add: distinct_mset_def)
-
-lemma mset_points_distinct_diff_one: "distinct_mset (mset_set (\<V> - {x}))"
- by (meson count_mset_set_le_one distinct_mset_count_less_1)
-
-lemma finite_design_support: "finite (design_support)"
- using design_support_def by auto
-
-lemma block_size_lt_order: "bl \<in># \<B> \<Longrightarrow> card bl \<le> card \<V>"
- using wellformed by (simp add: card_mono finite_sets)
-
-end
-
-subsection \<open>Designs\<close>
-
-text \<open>There are many varied definitions of a design in literature. However, the most
-commonly accepted definition is a finite point set, $V$ and collection of blocks $B$, where
-no block in $B$ can be empty\<close>
-locale design = finite_incidence_system +
- assumes blocks_nempty: "bl \<in># \<B> \<Longrightarrow> bl \<noteq> {}"
-begin
-
-lemma wf_design: "design \<V> \<B>" by intro_locales
-
-lemma wf_design_iff: "bl \<in># \<B> \<Longrightarrow> design \<V> \<B> \<longleftrightarrow> (bl \<subseteq> \<V> \<and> finite \<V> \<and> bl \<noteq> {})"
- using blocks_nempty wellformed finite_sets
- by (simp add: wf_design)
-
-text \<open>Reasoning on non empty properties and non zero parameters\<close>
-lemma blocks_nempty_alt: "\<forall> bl \<in># \<B>. bl \<noteq> {}"
- using blocks_nempty by auto
-
-lemma block_set_nempty_imp_points: "\<B> \<noteq> {#} \<Longrightarrow> \<V> \<noteq> {}"
- using wf_design wf_design_iff by auto
-
-lemma b_non_zero_imp_v_non_zero: "\<b> > 0 \<Longrightarrow> \<v> > 0"
- using block_set_nempty_imp_points finite_sets by fastforce
-
-lemma v_eq0_imp_b_eq_0: "\<v> = 0 \<Longrightarrow> \<b> = 0"
- using b_non_zero_imp_v_non_zero by auto
-
-text \<open>Size lemmas\<close>
-lemma block_size_lt_v: "bl \<in># \<B> \<Longrightarrow> card bl \<le> \<v>"
- by (simp add: card_mono finite_sets wellformed)
-
-lemma block_size_gt_0: "bl \<in># \<B> \<Longrightarrow> card bl > 0"
- using finite_sets blocks_nempty finite_blocks by fastforce
-
-lemma design_cart_product_size: "size ((mset_set \<V>) \<times># \<B>) = \<v> * \<b>"
- by (simp add: size_cartesian_product)
-
-end
-
-text \<open>Intro rules for design locale\<close>
-
-lemma wf_design_implies:
- assumes "(\<And> b . b \<in># \<B> \<Longrightarrow> b \<subseteq> V)"
- assumes "\<And> b . b \<in># \<B> \<Longrightarrow> b \<noteq> {}"
- assumes "finite V"
- assumes "\<B> \<noteq> {#}"
- assumes "V \<noteq> {}"
- shows "design V \<B>"
- using assms by (unfold_locales) simp_all
-
-lemma (in incidence_system) finite_sysI[intro]: "finite \<V> \<Longrightarrow> finite_incidence_system \<V> \<B>"
- by (unfold_locales) simp_all
-
-lemma (in finite_incidence_system) designI[intro]: "(\<And> b. b \<in># \<B> \<Longrightarrow> b \<noteq> {}) \<Longrightarrow> \<B> \<noteq> {#}
- \<Longrightarrow> \<V> \<noteq> {} \<Longrightarrow> design \<V> \<B>"
- by (unfold_locales) simp_all
-
-subsection \<open>Core Property Definitions\<close>
-
-subsubsection \<open>Replication Number\<close>
-
-text \<open>The replication number for a point is the number of blocks that point is incident with\<close>
-
-definition point_replication_number :: "'a set multiset \<Rightarrow> 'a \<Rightarrow> int" (infix "rep" 75) where
-"B rep x \<equiv> size {#b \<in># B . x \<in> b#}"
-
-lemma max_point_rep: "B rep x \<le> size B"
- using size_filter_mset_lesseq by (simp add: point_replication_number_def)
-
-lemma rep_number_g0_exists:
- assumes "B rep x > 0"
- obtains b where "b \<in># B" and "x \<in> b"
-proof -
- have "size {#b \<in># B . x \<in> b#} > 0" using assms point_replication_number_def
- by (metis of_nat_0_less_iff)
- thus ?thesis
- by (metis filter_mset_empty_conv nonempty_has_size that)
-qed
-
-lemma rep_number_on_set_def: "finite B \<Longrightarrow> (mset_set B) rep x = card {b \<in> B . x \<in> b}"
- by (simp add: point_replication_number_def)
-
-lemma point_rep_number_split[simp]: "(A + B) rep x = A rep x + B rep x"
- by (simp add: point_replication_number_def)
-
-lemma point_rep_singleton_val [simp]: "x \<in> b \<Longrightarrow> {#b#} rep x = 1"
- by (simp add: point_replication_number_def)
-
-lemma point_rep_singleton_inval [simp]: "x \<notin> b \<Longrightarrow> {#b#} rep x = 0"
- by (simp add: point_replication_number_def)
-
-context incidence_system
-begin
-
-lemma point_rep_number_alt_def: "\<B> rep x = size {# b \<in># \<B> . x \<in> b#}"
- by (simp add: point_replication_number_def)
-
-lemma rep_number_non_zero_system_point: " \<B> rep x > 0 \<Longrightarrow> x \<in> \<V>"
- using rep_number_g0_exists wellformed
- by (metis wf_invalid_point)
-
-lemma point_rep_non_existance [simp]: "x \<notin> \<V> \<Longrightarrow> \<B> rep x = 0"
- using wf_invalid_point by (simp add: point_replication_number_def filter_mset_empty_conv)
-
-lemma point_rep_number_inv: "size {# b \<in># \<B> . x \<notin> b #} = \<b> - (\<B> rep x)"
-proof -
- have "\<b> = size {# b \<in># \<B> . x \<notin> b #} + size {# b \<in># \<B> . x \<in> b #}"
- using multiset_partition by (metis add.commute size_union)
- thus ?thesis by (simp add: point_replication_number_def)
-qed
-
-lemma point_rep_num_inv_non_empty: "(\<B> rep x) < \<b> \<Longrightarrow> \<B> \<noteq> {#} \<Longrightarrow> {# b \<in># \<B> . x \<notin> b #} \<noteq> {#}"
- by (metis diff_zero point_replication_number_def size_empty size_filter_neg verit_comp_simplify1(1))
-
-end
-
-subsubsection \<open>Point Index\<close>
-
-text \<open>The point index of a subset of points in a design, is the number of times those points
-occur together in a block of the design\<close>
-definition points_index :: "'a set multiset \<Rightarrow> 'a set \<Rightarrow> nat" (infix "index" 75) where
-"B index ps \<equiv> size {#b \<in># B . ps \<subseteq> b#}"
-
-lemma points_index_empty [simp]: "{#} index ps = 0"
- by (simp add: points_index_def)
-
-lemma point_index_distrib: "(B1 + B2) index ps = B1 index ps + B2 index ps"
- by (simp add: points_index_def)
-
-lemma point_index_diff: "B1 index ps = (B1 + B2) index ps - B2 index ps"
- by (simp add: points_index_def)
-
-lemma points_index_singleton: "{#b#} index ps = 1 \<longleftrightarrow> ps \<subseteq> b"
- by (simp add: points_index_def)
-
-lemma points_index_singleton_zero: "\<not> (ps \<subseteq> b) \<Longrightarrow> {#b#} index ps = 0"
- by (simp add: points_index_def)
-
-lemma points_index_sum: "(\<Sum>\<^sub># B ) index ps = (\<Sum>b \<in># B . (b index ps))"
- using points_index_empty by (induction B) (auto simp add: point_index_distrib)
-
-lemma points_index_block_image_add_eq:
- assumes "x \<notin> ps"
- assumes "B index ps = l"
- shows "{# insert x b . b \<in># B#} index ps = l"
- using points_index_def by (metis (no_types, lifting) assms filter_mset_cong
- image_mset_filter_swap2 points_index_def size_image_mset subset_insert)
-
-lemma points_index_on_set_def [simp]:
- assumes "finite B"
- shows "(mset_set B) index ps = card {b \<in> B. ps \<subseteq> b}"
- by (simp add: points_index_def assms)
-
-lemma points_index_single_rep_num: "B index {x} = B rep x"
- by (simp add: points_index_def point_replication_number_def)
-
-lemma points_index_pair_rep_num:
- assumes "\<And> b. b \<in># B \<Longrightarrow> x \<in> b"
- shows "B index {x, y} = B rep y"
- using point_replication_number_def points_index_def
- by (metis assms empty_subsetI filter_mset_cong insert_subset)
-
-lemma points_index_0_left_imp:
- assumes "B index ps = 0"
- assumes "b \<in># B"
- shows "\<not> (ps \<subseteq> b)"
-proof (rule ccontr)
- assume "\<not> \<not> ps \<subseteq> b"
- then have a: "ps \<subseteq> b" by auto
- then have "b \<in># {#bl \<in># B . ps \<subseteq> bl#}" by (simp add: assms(2))
- thus False by (metis assms(1) count_greater_eq_Suc_zero_iff count_size_set_repr not_less_eq_eq
- points_index_def size_filter_mset_lesseq)
-qed
-
-lemma points_index_0_right_imp:
- assumes "\<And> b . b \<in># B \<Longrightarrow> (\<not> ps \<subseteq> b)"
- shows "B index ps = 0"
- using assms by (simp add: filter_mset_empty_conv points_index_def)
-
-lemma points_index_0_iff: "B index ps = 0 \<longleftrightarrow> (\<forall> b. b \<in># B \<longrightarrow> (\<not> ps \<subseteq> b))"
- using points_index_0_left_imp points_index_0_right_imp by metis
-
-lemma points_index_gt0_impl_existance:
- assumes "B index ps > 0"
- shows "(\<exists> bl . (bl \<in># B \<and> ps \<subseteq> bl))"
-proof -
- have "size {#bl \<in># B . ps \<subseteq> bl#} > 0"
- by (metis assms points_index_def)
- then obtain bl where "bl \<in># B" and "ps \<subseteq> bl"
- by (metis filter_mset_empty_conv nonempty_has_size)
- thus ?thesis by auto
-qed
-
-lemma points_index_one_unique:
- assumes "B index ps = 1"
- assumes "bl \<in># B" and "ps \<subseteq> bl" and "bl' \<in># B" and "ps \<subseteq> bl'"
- shows "bl = bl'"
-proof (rule ccontr)
- assume assm: "bl \<noteq> bl'"
- then have bl1: "bl \<in># {#bl \<in># B . ps \<subseteq> bl#}" using assms by simp
- then have bl2: "bl'\<in># {#bl \<in># B . ps \<subseteq> bl#}" using assms by simp
- then have "{#bl, bl'#} \<subseteq># {#bl \<in># B . ps \<subseteq> bl#}" using assms by (metis bl1 bl2 points_index_def
- add_mset_subseteq_single_iff assm mset_subset_eq_single size_single subseteq_mset_size_eql)
- then have "size {#bl \<in># B . ps \<subseteq> bl#} \<ge> 2" using size_mset_mono by fastforce
- thus False using assms by (metis numeral_le_one_iff points_index_def semiring_norm(69))
-qed
-
-lemma points_index_one_unique_block:
- assumes "B index ps = 1"
- shows "\<exists>! bl . (bl \<in># B \<and> ps \<subseteq> bl)"
- using assms points_index_gt0_impl_existance points_index_one_unique
- by (metis zero_less_one)
-
-lemma points_index_one_not_unique_block:
- assumes "B index ps = 1"
- assumes "ps \<subseteq> bl"
- assumes "bl \<in># B"
- assumes "bl' \<in># B - {#bl#}"
- shows "\<not> ps \<subseteq> bl'"
-proof -
- have "B = (B - {#bl#}) + {#bl#}" by (simp add: assms(3))
- then have "(B - {#bl#}) index ps = B index ps - {#bl#} index ps"
- by (metis point_index_diff)
- then have "(B - {#bl#}) index ps = 0" using assms points_index_singleton
- by (metis diff_self_eq_0)
- thus ?thesis using assms(4) points_index_0_left_imp by auto
-qed
-
-lemma (in incidence_system) points_index_alt_def: "\<B> index ps = size {#b \<in># \<B> . ps \<subseteq> b#}"
- by (simp add: points_index_def)
-
-lemma (in incidence_system) points_index_ps_nin: "\<not> (ps \<subseteq> \<V>) \<Longrightarrow> \<B> index ps = 0"
- using points_index_alt_def filter_mset_empty_conv in_mono size_empty subsetI wf_invalid_point
- by metis
-
-lemma (in incidence_system) points_index_count_bl:
- "multiplicity bl \<ge> n \<Longrightarrow> ps \<subseteq> bl \<Longrightarrow> count {#bl \<in># \<B> . ps \<subseteq> bl#} bl \<ge> n"
- by simp
-
-lemma (in finite_incidence_system) points_index_zero:
- assumes "card ps > card \<V>"
- shows "\<B> index ps = 0"
-proof -
- have "\<And> b. b \<in># \<B> \<Longrightarrow> card ps > card b"
- using block_size_lt_order card_subset_not_gt_card finite_sets assms by fastforce
- then have "{#b \<in># \<B> . ps \<subseteq> b#} = {#}"
- by (simp add: card_subset_not_gt_card filter_mset_empty_conv finite_blocks)
- thus ?thesis using points_index_alt_def by simp
-qed
-
-lemma (in design) points_index_subset:
- "x \<subseteq># {#bl \<in># \<B> . ps \<subseteq> bl#} \<Longrightarrow> ps \<subseteq> \<V> \<Longrightarrow> (\<B> index ps) \<ge> (size x)"
- by (simp add: points_index_def size_mset_mono)
-
-lemma (in design) points_index_count_min: "multiplicity bl \<ge> n \<Longrightarrow> ps \<subseteq> bl \<Longrightarrow> \<B> index ps \<ge> n"
- using points_index_alt_def set_count_size_min by (metis filter_mset.rep_eq)
-
-subsubsection \<open>Intersection Number\<close>
-
-text \<open>The intersection number of two blocks is the size of the intersection of those blocks. i.e.
-the number of points which occur in both blocks\<close>
-definition intersection_number :: "'a set \<Rightarrow> 'a set \<Rightarrow> int" (infix "|\<inter>|" 70) where
-"b1 |\<inter>| b2 \<equiv> card (b1 \<inter> b2)"
-
-lemma intersection_num_non_neg: "b1 |\<inter>| b2 \<ge> 0"
- by (simp add: intersection_number_def)
-
-lemma intersection_number_empty_iff:
- assumes "finite b1"
- shows "b1 \<inter> b2 = {} \<longleftrightarrow> b1 |\<inter>| b2 = 0"
- by (simp add: intersection_number_def assms)
-
-lemma intersect_num_commute: "b1 |\<inter>| b2 = b2 |\<inter>| b1"
- by (simp add: inf_commute intersection_number_def)
-
-definition n_intersect_number :: "'a set \<Rightarrow> nat\<Rightarrow> 'a set \<Rightarrow> int" where
-"n_intersect_number b1 n b2 \<equiv> card { x \<in> Pow (b1 \<inter> b2) . card x = n}"
-
-notation n_intersect_number ("(_ |\<inter>|\<^sub>_ _)" [52, 51, 52] 50)
-
-lemma n_intersect_num_subset_def: "b1 |\<inter>|\<^sub>n b2 = card {x . x \<subseteq> b1 \<inter> b2 \<and> card x = n}"
- using n_intersect_number_def by auto
-
-lemma n_inter_num_one: "finite b1 \<Longrightarrow> finite b2 \<Longrightarrow> b1 |\<inter>|\<^sub>1 b2 = b1 |\<inter>| b2"
- using n_intersect_number_def intersection_number_def card_Pow_filter_one
- by (metis (full_types) finite_Int)
-
-lemma n_inter_num_choose: "finite b1 \<Longrightarrow> finite b2 \<Longrightarrow> b1 |\<inter>|\<^sub>n b2 = (card (b1 \<inter> b2) choose n)"
- using n_subsets n_intersect_num_subset_def
- by (metis (full_types) finite_Int)
-
-lemma set_filter_single: "x \<in> A \<Longrightarrow> {a \<in> A . a = x} = {x}"
- by auto
-
-lemma (in design) n_inter_num_zero:
- assumes "b1 \<in># \<B>" and "b2 \<in># \<B>"
- shows "b1 |\<inter>|\<^sub>0 b2 = 1"
-proof -
- have empty: "\<And>x . finite x \<Longrightarrow> card x = 0 \<Longrightarrow> x = {}"
- by simp
- have empt_in: "{} \<in> Pow (b1 \<inter> b2)" by simp
- have "finite (b1 \<inter> b2)" using finite_blocks assms by simp
- then have "\<And> x . x \<in> Pow (b1 \<inter> b2) \<Longrightarrow> finite x" by (meson PowD finite_subset)
- then have "{x \<in> Pow (b1 \<inter> b2) . card x = 0} = {x \<in> Pow (b1 \<inter> b2) . x = {}}"
- using empty by (metis card.empty)
- then have "{x \<in> Pow (b1 \<inter> b2) . card x = 0} = {{}}"
- by (simp add: empt_in set_filter_single Collect_conv_if)
- thus ?thesis by (simp add: n_intersect_number_def)
-qed
-
-lemma (in design) n_inter_num_choose_design: "b1 \<in># \<B> \<Longrightarrow> b2 \<in># \<B>
- \<Longrightarrow> b1 |\<inter>|\<^sub>n b2 = (card (b1 \<inter> b2) choose n) "
- using finite_blocks by (simp add: n_inter_num_choose)
-
-lemma (in design) n_inter_num_choose_design_inter: "b1 \<in># \<B> \<Longrightarrow> b2 \<in># \<B>
- \<Longrightarrow> b1 |\<inter>|\<^sub>n b2 = (nat (b1 |\<inter>| b2) choose n) "
- using finite_blocks by (simp add: n_inter_num_choose intersection_number_def)
-
-subsection \<open>Incidence System Set Property Definitions\<close>
-context incidence_system
-begin
-
-text \<open>The set of replication numbers for all points of design\<close>
-definition replication_numbers :: "int set" where
-"replication_numbers \<equiv> {\<B> rep x | x . x \<in> \<V>}"
-
-lemma replication_numbers_non_empty:
- assumes "\<V> \<noteq> {}"
- shows "replication_numbers \<noteq> {}"
- by (simp add: assms replication_numbers_def)
-
-lemma obtain_point_with_rep: "r \<in> replication_numbers \<Longrightarrow> \<exists> x. x \<in> \<V> \<and> \<B> rep x = r"
- using replication_numbers_def by auto
-
-lemma point_rep_number_in_set: "x \<in> \<V> \<Longrightarrow> (\<B> rep x) \<in> replication_numbers"
- by (auto simp add: replication_numbers_def)
-
-lemma (in finite_incidence_system) replication_numbers_finite: "finite replication_numbers"
- using finite_sets by (simp add: replication_numbers_def)
-
-text \<open>The set of all block sizes in a system\<close>
-
-definition sys_block_sizes :: "int set" where
-"sys_block_sizes \<equiv> { (int (card bl)) | bl. bl \<in># \<B>}"
-
-lemma block_sizes_non_empty_set:
- assumes "\<B> \<noteq> {#}"
- shows "sys_block_sizes \<noteq> {}"
-by (simp add: sys_block_sizes_def assms)
-
-lemma finite_block_sizes: "finite (sys_block_sizes)"
- by (simp add: sys_block_sizes_def)
-
-lemma block_sizes_non_empty:
- assumes "\<B> \<noteq> {#}"
- shows "card (sys_block_sizes) > 0"
- using finite_block_sizes block_sizes_non_empty_set
- by (simp add: assms card_gt_0_iff)
-
-lemma sys_block_sizes_in: "bl \<in># \<B> \<Longrightarrow> card bl \<in> sys_block_sizes"
- unfolding sys_block_sizes_def by auto
-
-lemma sys_block_sizes_obtain_bl: "x \<in> sys_block_sizes \<Longrightarrow> (\<exists> bl \<in># \<B>. int (card bl) = x)"
- by (auto simp add: sys_block_sizes_def)
-
-text \<open>The set of all possible intersection numbers in a system.\<close>
-
-definition intersection_numbers :: "int set" where
-"intersection_numbers \<equiv> { b1 |\<inter>| b2 | b1 b2 . b1 \<in># \<B> \<and> b2 \<in># (\<B> - {#b1#})}"
-
-lemma obtain_blocks_intersect_num: "n \<in> intersection_numbers \<Longrightarrow>
- \<exists> b1 b2. b1 \<in># \<B> \<and> b2 \<in># (\<B> - {#b1#}) \<and> b1 |\<inter>| b2 = n"
- by (auto simp add: intersection_numbers_def)
-
-lemma intersect_num_in_set: "b1 \<in># \<B> \<Longrightarrow> b2 \<in># (\<B> - {#b1#}) \<Longrightarrow> b1 |\<inter>| b2 \<in> intersection_numbers"
- by (auto simp add: intersection_numbers_def)
-
-text \<open>The set of all possible point indices\<close>
-definition point_indices :: "int \<Rightarrow> int set" where
-"point_indices t \<equiv> {\<B> index ps | ps. int (card ps) = t \<and> ps \<subseteq> \<V>}"
-
-lemma point_indices_elem_in: "ps \<subseteq> \<V> \<Longrightarrow> card ps = t \<Longrightarrow> \<B> index ps \<in> point_indices t"
- by (auto simp add: point_indices_def)
-
-lemma point_indices_alt_def: "point_indices t = { \<B> index ps | ps. int (card ps) = t \<and> ps \<subseteq> \<V>}"
- by (simp add: point_indices_def)
-
-end
-
-subsection \<open>Basic Constructions on designs\<close>
-
-text \<open>This section defines some of the most common universal constructions found in design theory
-involving only a single design\<close>
-
-subsubsection \<open>Design Complements\<close>
-
-context incidence_system
-begin
-
-text \<open>The complement of a block are all the points in the design not in that block.
-The complement of a design is therefore the original point sets, and set of all block complements\<close>
-definition block_complement:: "'a set \<Rightarrow> 'a set" ("_\<^sup>c" [56] 55) where
-"block_complement b \<equiv> \<V> - b"
-
-definition complement_blocks :: "'a set multiset" ("(\<B>\<^sup>C)")where
-"complement_blocks \<equiv> {# bl\<^sup>c . bl \<in># \<B> #}"
-
-lemma block_complement_elem_iff:
- assumes "ps \<subseteq> \<V>"
- shows "ps \<subseteq> bl\<^sup>c \<longleftrightarrow> (\<forall> x \<in> ps. x \<notin> bl)"
- using assms block_complement_def by (auto)
-
-lemma block_complement_inter_empty: "bl1\<^sup>c = bl2 \<Longrightarrow> bl1 \<inter> bl2 = {}"
- using block_complement_def by auto
-
-lemma block_complement_inv:
- assumes "bl \<in># \<B>"
- assumes "bl\<^sup>c = bl2"
- shows "bl2\<^sup>c = bl"
- by (metis Diff_Diff_Int assms(1) assms(2) block_complement_def inf.absorb_iff2 wellformed)
-
-lemma block_complement_subset_points: "ps \<subseteq> (bl\<^sup>c) \<Longrightarrow> ps \<subseteq> \<V>"
- using block_complement_def by blast
-
-lemma obtain_comp_block_orig:
- assumes "bl1 \<in># \<B>\<^sup>C"
- obtains bl2 where "bl2 \<in># \<B>" and "bl1 = bl2\<^sup>c"
- using wellformed assms by (auto simp add: complement_blocks_def)
-
-lemma complement_same_b [simp]: "size \<B>\<^sup>C = size \<B>"
- by (simp add: complement_blocks_def)
-
-lemma block_comp_elem_alt_left: "x \<in> bl \<Longrightarrow> ps \<subseteq> bl\<^sup>c \<Longrightarrow> x \<notin> ps"
- by (auto simp add: block_complement_def block_complement_elem_iff)
-
-lemma block_comp_elem_alt_right: "ps \<subseteq> \<V> \<Longrightarrow> (\<And> x . x \<in> ps \<Longrightarrow> x \<notin> bl) \<Longrightarrow> ps \<subseteq> bl\<^sup>c"
- by (auto simp add: block_complement_elem_iff)
-
-lemma complement_index:
- assumes "ps \<subseteq> \<V>"
- shows "\<B>\<^sup>C index ps = size {# b \<in># \<B> . (\<forall> x \<in> ps . x \<notin> b) #}"
-proof -
- have "\<B>\<^sup>C index ps = size {# b \<in># {# bl\<^sup>c . bl \<in># \<B>#}. ps \<subseteq> b #}"
- by (simp add: complement_blocks_def points_index_def)
- then have "\<B>\<^sup>C index ps = size {# bl\<^sup>c | bl \<in># \<B> . ps \<subseteq> bl\<^sup>c #}"
- by (metis image_mset_filter_swap)
- thus ?thesis using assms by (simp add: block_complement_elem_iff)
-qed
-
-lemma complement_index_2:
- assumes "{x, y} \<subseteq> \<V>"
- shows "\<B>\<^sup>C index {x, y} = size {# b \<in># \<B> . x \<notin> b \<and> y \<notin> b #}"
-proof -
- have a: "\<And> b. b \<in># \<B> \<Longrightarrow> \<forall> x' \<in> {x, y} . x' \<notin> b \<Longrightarrow> x \<notin> b \<and> y \<notin> b"
- by simp
- have "\<And> b. b \<in># \<B> \<Longrightarrow> x \<notin> b \<and> y \<notin> b \<Longrightarrow> \<forall> x' \<in> {x, y} . x' \<notin> b "
- by simp
- thus ?thesis using assms a complement_index
- by (smt (verit) filter_mset_cong)
-qed
-
-lemma complement_rep_number:
- assumes "x \<in> \<V>" and "\<B> rep x = r"
- shows "\<B>\<^sup>C rep x = \<b> - r"
-proof -
- have r: "size {#b \<in># \<B> . x \<in> b#} = r" using assms by (simp add: point_replication_number_def)
- then have a: "\<And> b . b \<in># \<B> \<Longrightarrow> x \<in> b \<Longrightarrow> x \<notin> b\<^sup>c"
- by (simp add: block_complement_def)
- have "\<And> b . b \<in># \<B> \<Longrightarrow> x \<notin> b \<Longrightarrow> x \<in> b\<^sup>c"
- by (simp add: assms(1) block_complement_def)
- then have alt: "(image_mset block_complement \<B>) rep x = size {#b \<in># \<B> . x \<notin> b#}"
- using a filter_mset_cong image_mset_filter_swap2 point_replication_number_def
- by (smt (verit, ccfv_SIG) size_image_mset)
- have "\<b> = size {#b \<in># \<B> . x \<in> b#} + size {#b \<in># \<B> . x \<notin> b#}"
- by (metis multiset_partition size_union)
- thus ?thesis using alt
- by (simp add: r complement_blocks_def)
-qed
-
-lemma complement_blocks_wf: "bl \<in># \<B>\<^sup>C \<Longrightarrow> bl \<subseteq> \<V>"
- by (auto simp add: complement_blocks_def block_complement_def)
-
-lemma complement_wf [intro]: "incidence_system \<V> \<B>\<^sup>C"
- using complement_blocks_wf by (unfold_locales)
-
-interpretation sys_complement: incidence_system "\<V>" "\<B>\<^sup>C"
- using complement_wf by simp
-end
-
-context finite_incidence_system
-begin
-lemma block_complement_size: "b \<subseteq> \<V> \<Longrightarrow> card (b\<^sup>c) = card \<V> - card b"
- by (simp add: block_complement_def card_Diff_subset finite_subset card_mono of_nat_diff finite_sets)
-
-lemma block_comp_incomplete: "incomplete_block bl \<Longrightarrow> card (bl\<^sup>c) > 0"
- using block_complement_size by (simp add: wellformed)
-
-lemma block_comp_incomplete_nempty: "incomplete_block bl \<Longrightarrow> bl\<^sup>c \<noteq> {}"
- using wellformed block_complement_def finite_blocks
- by (auto simp add: block_complement_size block_comp_incomplete card_subset_not_gt_card)
-
-lemma incomplete_block_proper_subset: "incomplete_block bl \<Longrightarrow> bl \<subset> \<V>"
- using wellformed by fastforce
-
-lemma complement_finite: "finite_incidence_system \<V> \<B>\<^sup>C"
- using complement_wf finite_sets by (simp add: incidence_system.finite_sysI)
-
-interpretation comp_fin: finite_incidence_system \<V> "\<B>\<^sup>C"
- using complement_finite by simp
-
-end
-
-context design
-begin
-lemma (in design) complement_design:
- assumes "\<And> bl . bl \<in># \<B> \<Longrightarrow> incomplete_block bl"
- shows "design \<V> (\<B>\<^sup>C)"
-proof -
- interpret fin: finite_incidence_system \<V> "\<B>\<^sup>C" using complement_finite by simp
- show ?thesis using assms block_comp_incomplete_nempty wellformed
- by (unfold_locales) (auto simp add: complement_blocks_def)
-qed
-
-end
-subsubsection \<open>Multiples\<close>
-text \<open>An easy way to construct new set systems is to simply multiply the block collection by some
-constant\<close>
-
-context incidence_system
-begin
-
-abbreviation multiple_blocks :: "nat \<Rightarrow> 'a set multiset" where
-"multiple_blocks n \<equiv> repeat_mset n \<B>"
-
-lemma multiple_block_in_original: "b \<in># multiple_blocks n \<Longrightarrow> b \<in># \<B>"
- by (simp add: elem_in_repeat_in_original)
-
-lemma multiple_block_in: "n > 0 \<Longrightarrow> b \<in># \<B> \<Longrightarrow> b \<in># multiple_blocks n"
- by (simp add: elem_in_original_in_repeat)
-
-lemma multiple_blocks_gt: "n > 0 \<Longrightarrow> size (multiple_blocks n) \<ge> size \<B>"
- by (simp)
-
-lemma block_original_count_le: "n > 0 \<Longrightarrow> count \<B> b \<le> count (multiple_blocks n) b"
- using count_repeat_mset by simp
-
-lemma multiple_blocks_sub: "n > 0 \<Longrightarrow> \<B> \<subseteq># (multiple_blocks n)"
- by (simp add: mset_subset_eqI block_original_count_le)
-
-lemma multiple_1_same: "multiple_blocks 1 = \<B>"
- by simp
-
-lemma multiple_unfold_1: "multiple_blocks (Suc n) = (multiple_blocks n) + \<B>"
- by simp
-
-lemma multiple_point_rep_num: "(multiple_blocks n) rep x = (\<B> rep x) * n"
-proof (induction n)
- case 0
- then show ?case by (simp add: point_replication_number_def)
-next
- case (Suc n)
- then have "multiple_blocks (Suc n) rep x = \<B> rep x * n + (\<B> rep x)"
- using Suc.IH Suc.prems by (simp add: union_commute point_replication_number_def)
- then show ?case
- by (simp add: int_distrib(2))
-qed
-
-lemma multiple_point_index: "(multiple_blocks n) index ps = (\<B> index ps) * n"
- by (induction n) (auto simp add: points_index_def)
-
-lemma repeat_mset_block_point_rel: "\<And>b x. b \<in># multiple_blocks n \<Longrightarrow> x \<in> b \<Longrightarrow> x \<in> \<V>"
- by (induction n) (auto, meson subset_iff wellformed)
-
-lemma multiple_is_wellformed: "incidence_system \<V> (multiple_blocks n)"
- using repeat_mset_subset_in wellformed repeat_mset_block_point_rel by (unfold_locales) (auto)
-
-lemma multiple_blocks_num [simp]: "size (multiple_blocks n) = n*\<b>"
- by simp
-
-interpretation mult_sys: incidence_system \<V> "(multiple_blocks n)"
- by (simp add: multiple_is_wellformed)
-
-lemma multiple_block_multiplicity [simp]: "mult_sys.multiplicity n bl = (multiplicity bl) * n"
- by (simp)
-
-lemma multiple_block_sizes_same:
- assumes "n > 0"
- shows "sys_block_sizes = mult_sys.sys_block_sizes n"
-proof -
- have def: "mult_sys.sys_block_sizes n = {card bl | bl. bl \<in># (multiple_blocks n)}"
- by (simp add: mult_sys.sys_block_sizes_def)
- then have eq: "\<And> bl. bl \<in># (multiple_blocks n) \<longleftrightarrow> bl \<in># \<B>"
- using assms multiple_block_in multiple_block_in_original by blast
- thus ?thesis using def by (simp add: sys_block_sizes_def eq)
-qed
-
-end
-
-context finite_incidence_system
-begin
-
-lemma multiple_is_finite: "finite_incidence_system \<V> (multiple_blocks n)"
- using multiple_is_wellformed finite_sets by (unfold_locales) (auto simp add: incidence_system_def)
-
-end
-
-context design
-begin
-
-lemma multiple_is_design: "design \<V> (multiple_blocks n)"
-proof -
- interpret fis: finite_incidence_system \<V> "multiple_blocks n" using multiple_is_finite by simp
- show ?thesis using blocks_nempty
- by (unfold_locales) (auto simp add: elem_in_repeat_in_original repeat_mset_not_empty)
-qed
-
-end
-
-subsection \<open>Simple Designs\<close>
-
-text \<open>Simple designs are those in which the multiplicity of each block is at most one.
-In other words, the block collection is a set. This can significantly ease reasoning.\<close>
-
-locale simple_incidence_system = incidence_system +
- assumes simple [simp]: "bl \<in># \<B> \<Longrightarrow> multiplicity bl = 1"
-
-begin
-
-lemma simple_alt_def_all: "\<forall> bl \<in># \<B> . multiplicity bl = 1"
- using simple by auto
-
-lemma simple_blocks_eq_sup: "mset_set (design_support) = \<B>"
- using distinct_mset_def simple design_support_def by (metis distinct_mset_set_mset_ident)
-
-lemma simple_block_size_eq_card: "\<b> = card (design_support)"
- by (metis simple_blocks_eq_sup size_mset_set)
-
-lemma points_index_simple_def: "\<B> index ps = card {b \<in> design_support . ps \<subseteq> b}"
- using design_support_def points_index_def card_size_filter_eq simple_blocks_eq_sup
- by (metis finite_set_mset)
-
-lemma replication_num_simple_def: "\<B> rep x = card {b \<in> design_support . x \<in> b}"
- using design_support_def point_replication_number_def card_size_filter_eq simple_blocks_eq_sup
- by (metis finite_set_mset)
-
-end
-
-locale simple_design = design + simple_incidence_system
-
-text \<open>Additional reasoning about when something is not simple\<close>
-context incidence_system
-begin
-lemma simple_not_multiplicity: "b \<in># \<B> \<Longrightarrow> multiplicity b > 1 \<Longrightarrow> \<not> simple_incidence_system \<V> \<B>"
- using simple_incidence_system_def simple_incidence_system_axioms_def by (metis nat_neq_iff)
-
-lemma multiple_not_simple:
- assumes "n > 1"
- assumes "\<B> \<noteq> {#}"
- shows "\<not> simple_incidence_system \<V> (multiple_blocks n)"
-proof (rule ccontr, simp)
- assume "simple_incidence_system \<V> (multiple_blocks n)"
- then have "\<And> bl. bl \<in># \<B> \<Longrightarrow> count (multiple_blocks n) bl = 1"
- using assms(1) elem_in_original_in_repeat
- by (metis not_gr_zero not_less_zero simple_incidence_system.simple)
- thus False using assms by auto
-qed
-
-end
-
-subsection \<open>Proper Designs\<close>
-text \<open>Many types of designs rely on parameter conditions that only make sense for non-empty designs.
-i.e. designs with at least one block, and therefore given well-formed condition, at least one point.
-To this end we define the notion of a "proper" design\<close>
-
-locale proper_design = design +
- assumes b_non_zero: "\<b> \<noteq> 0"
-begin
-
-lemma is_proper: "proper_design \<V> \<B>" by intro_locales
-
-lemma v_non_zero: "\<v> > 0"
- using b_non_zero v_eq0_imp_b_eq_0 by auto
-
-lemma b_positive: "\<b> > 0" using b_non_zero
- by (simp add: nonempty_has_size)
-
-lemma design_points_nempty: "\<V> \<noteq> {}"
- using v_non_zero by auto
-
-lemma design_blocks_nempty: "\<B> \<noteq> {#}"
- using b_non_zero by auto
-
-end
-
-text \<open>Intro rules for a proper design\<close>
-lemma (in design) proper_designI[intro]: "\<b> \<noteq> 0 \<Longrightarrow> proper_design \<V> \<B>"
- by (unfold_locales) simp
-
-lemma proper_designII[intro]:
- assumes "design V B" and "B \<noteq> {#}"
- shows "proper_design V B"
-proof -
- interpret des: design V B using assms by simp
- show ?thesis using assms by unfold_locales simp
-qed
-
-text \<open>Reasoning on construction closure for proper designs\<close>
-context proper_design
-begin
-
-lemma multiple_proper_design:
- assumes "n > 0"
- shows "proper_design \<V> (multiple_blocks n)"
- using multiple_is_design assms design_blocks_nempty multiple_block_in
- by (metis block_set_nempty_imp_block_ex empty_iff proper_designII set_mset_empty)
-
-lemma complement_proper_design:
- assumes "\<And> bl . bl \<in># \<B> \<Longrightarrow> incomplete_block bl"
- shows "proper_design \<V> \<B>\<^sup>C"
-proof -
- interpret des: design \<V> "\<B>\<^sup>C"
- by (simp add: assms complement_design)
- show ?thesis using b_non_zero by (unfold_locales) auto
-qed
-
-end
+theory Design_Basics imports Main Multisets_Extras "HOL-Library.Disjoint_Sets"
+begin
+
+section \<open>Design Theory Basics\<close>
+text \<open>All definitions in this section reference the handbook of combinatorial designs
+ \cite{colbournHandbookCombinatorialDesigns2007}\<close>
+
+subsection \<open>Initial setup\<close>
+
+text \<open>Enable coercion of nats to ints to aid with reasoning on design properties\<close>
+declare [[coercion_enabled]]
+declare [[coercion "of_nat :: nat \<Rightarrow> int"]]
+
+subsection \<open>Incidence System\<close>
+
+text \<open>An incidence system is defined to be a wellformed set system. i.e. each block is a subset
+of the base point set. Alternatively, an incidence system can be looked at as the point set
+and an incidence relation which indicates if they are in the same block\<close>
+
+locale incidence_system =
+ fixes point_set :: "'a set" ("\<V>")
+ fixes block_collection :: "'a set multiset" ("\<B>")
+ assumes wellformed: "b \<in># \<B> \<Longrightarrow> b \<subseteq> \<V>"
+begin
+
+definition "\<I> \<equiv> { (x, b) . b \<in># \<B> \<and> x \<in> b}" (* incidence relation *)
+
+definition incident :: "'a \<Rightarrow> 'a set \<Rightarrow> bool" where
+"incident p b \<equiv> (p, b) \<in> \<I>"
+
+text \<open>Defines common notation used to indicate number of points ($v$) and number of blocks ($b$)\<close>
+abbreviation "\<v> \<equiv> int (card \<V>)"
+
+abbreviation "\<b> \<equiv> int (size \<B>)"
+
+text \<open>Basic incidence lemmas\<close>
+
+lemma incidence_alt_def:
+ assumes "p \<in> \<V>"
+ assumes "b \<in># \<B>"
+ shows "incident p b \<longleftrightarrow> p \<in> b"
+ by (auto simp add: incident_def \<I>_def assms)
+
+lemma wf_invalid_point: "x \<notin> \<V> \<Longrightarrow> b \<in># \<B> \<Longrightarrow> x \<notin> b"
+ using wellformed by auto
+
+lemma block_set_nempty_imp_block_ex: "\<B> \<noteq> {#} \<Longrightarrow> \<exists> bl . bl \<in># \<B>"
+ by auto
+
+text \<open>Abbreviations for all incidence systems\<close>
+abbreviation multiplicity :: "'a set \<Rightarrow> nat" where
+"multiplicity b \<equiv> count \<B> b"
+
+abbreviation incomplete_block :: "'a set \<Rightarrow> bool" where
+"incomplete_block bl \<equiv> card bl < card \<V> \<and> bl \<in># \<B>"
+
+lemma incomplete_alt_size: "incomplete_block bl \<Longrightarrow> card bl < \<v>"
+ by simp
+
+lemma incomplete_alt_in: "incomplete_block bl \<Longrightarrow> bl \<in># \<B>"
+ by simp
+
+lemma incomplete_alt_imp[intro]: "card bl < \<v> \<Longrightarrow> bl \<in># \<B> \<Longrightarrow> incomplete_block bl"
+ by simp
+
+definition design_support :: "'a set set" where
+"design_support \<equiv> set_mset \<B>"
+
+end
+
+subsection \<open>Finite Incidence Systems\<close>
+
+text \<open>These simply require the point set to be finite.
+As multisets are only defined to be finite, it is implied that the block set must be finite already\<close>
+
+locale finite_incidence_system = incidence_system +
+ assumes finite_sets: "finite \<V>"
+begin
+
+lemma finite_blocks: "b \<in># \<B> \<Longrightarrow> finite b"
+ using wellformed finite_sets finite_subset by blast
+
+lemma mset_points_distinct: "distinct_mset (mset_set \<V>)"
+ using finite_sets by (simp add: distinct_mset_def)
+
+lemma mset_points_distinct_diff_one: "distinct_mset (mset_set (\<V> - {x}))"
+ by (meson count_mset_set_le_one distinct_mset_count_less_1)
+
+lemma finite_design_support: "finite (design_support)"
+ using design_support_def by auto
+
+lemma block_size_lt_order: "bl \<in># \<B> \<Longrightarrow> card bl \<le> card \<V>"
+ using wellformed by (simp add: card_mono finite_sets)
+
+end
+
+subsection \<open>Designs\<close>
+
+text \<open>There are many varied definitions of a design in literature. However, the most
+commonly accepted definition is a finite point set, $V$ and collection of blocks $B$, where
+no block in $B$ can be empty\<close>
+locale design = finite_incidence_system +
+ assumes blocks_nempty: "bl \<in># \<B> \<Longrightarrow> bl \<noteq> {}"
+begin
+
+lemma wf_design: "design \<V> \<B>" by intro_locales
+
+lemma wf_design_iff: "bl \<in># \<B> \<Longrightarrow> design \<V> \<B> \<longleftrightarrow> (bl \<subseteq> \<V> \<and> finite \<V> \<and> bl \<noteq> {})"
+ using blocks_nempty wellformed finite_sets
+ by (simp add: wf_design)
+
+text \<open>Reasoning on non empty properties and non zero parameters\<close>
+lemma blocks_nempty_alt: "\<forall> bl \<in># \<B>. bl \<noteq> {}"
+ using blocks_nempty by auto
+
+lemma block_set_nempty_imp_points: "\<B> \<noteq> {#} \<Longrightarrow> \<V> \<noteq> {}"
+ using wf_design wf_design_iff by auto
+
+lemma b_non_zero_imp_v_non_zero: "\<b> > 0 \<Longrightarrow> \<v> > 0"
+ using block_set_nempty_imp_points finite_sets by fastforce
+
+lemma v_eq0_imp_b_eq_0: "\<v> = 0 \<Longrightarrow> \<b> = 0"
+ using b_non_zero_imp_v_non_zero by auto
+
+text \<open>Size lemmas\<close>
+lemma block_size_lt_v: "bl \<in># \<B> \<Longrightarrow> card bl \<le> \<v>"
+ by (simp add: card_mono finite_sets wellformed)
+
+lemma block_size_gt_0: "bl \<in># \<B> \<Longrightarrow> card bl > 0"
+ using finite_sets blocks_nempty finite_blocks by fastforce
+
+lemma design_cart_product_size: "size ((mset_set \<V>) \<times># \<B>) = \<v> * \<b>"
+ by (simp add: size_cartesian_product)
+
+end
+
+text \<open>Intro rules for design locale\<close>
+
+lemma wf_design_implies:
+ assumes "(\<And> b . b \<in># \<B> \<Longrightarrow> b \<subseteq> V)"
+ assumes "\<And> b . b \<in># \<B> \<Longrightarrow> b \<noteq> {}"
+ assumes "finite V"
+ assumes "\<B> \<noteq> {#}"
+ assumes "V \<noteq> {}"
+ shows "design V \<B>"
+ using assms by (unfold_locales) simp_all
+
+lemma (in incidence_system) finite_sysI[intro]: "finite \<V> \<Longrightarrow> finite_incidence_system \<V> \<B>"
+ by (unfold_locales) simp_all
+
+lemma (in finite_incidence_system) designI[intro]: "(\<And> b. b \<in># \<B> \<Longrightarrow> b \<noteq> {}) \<Longrightarrow> \<B> \<noteq> {#}
+ \<Longrightarrow> \<V> \<noteq> {} \<Longrightarrow> design \<V> \<B>"
+ by (unfold_locales) simp_all
+
+subsection \<open>Core Property Definitions\<close>
+
+subsubsection \<open>Replication Number\<close>
+
+text \<open>The replication number for a point is the number of blocks that point is incident with\<close>
+
+definition point_replication_number :: "'a set multiset \<Rightarrow> 'a \<Rightarrow> int" (infix "rep" 75) where
+"B rep x \<equiv> size {#b \<in># B . x \<in> b#}"
+
+lemma max_point_rep: "B rep x \<le> size B"
+ using size_filter_mset_lesseq by (simp add: point_replication_number_def)
+
+lemma rep_number_g0_exists:
+ assumes "B rep x > 0"
+ obtains b where "b \<in># B" and "x \<in> b"
+proof -
+ have "size {#b \<in># B . x \<in> b#} > 0" using assms point_replication_number_def
+ by (metis of_nat_0_less_iff)
+ thus ?thesis
+ by (metis filter_mset_empty_conv nonempty_has_size that)
+qed
+
+lemma rep_number_on_set_def: "finite B \<Longrightarrow> (mset_set B) rep x = card {b \<in> B . x \<in> b}"
+ by (simp add: point_replication_number_def)
+
+lemma point_rep_number_split[simp]: "(A + B) rep x = A rep x + B rep x"
+ by (simp add: point_replication_number_def)
+
+lemma point_rep_singleton_val [simp]: "x \<in> b \<Longrightarrow> {#b#} rep x = 1"
+ by (simp add: point_replication_number_def)
+
+lemma point_rep_singleton_inval [simp]: "x \<notin> b \<Longrightarrow> {#b#} rep x = 0"
+ by (simp add: point_replication_number_def)
+
+context incidence_system
+begin
+
+lemma point_rep_number_alt_def: "\<B> rep x = size {# b \<in># \<B> . x \<in> b#}"
+ by (simp add: point_replication_number_def)
+
+lemma rep_number_non_zero_system_point: " \<B> rep x > 0 \<Longrightarrow> x \<in> \<V>"
+ using rep_number_g0_exists wellformed
+ by (metis wf_invalid_point)
+
+lemma point_rep_non_existance [simp]: "x \<notin> \<V> \<Longrightarrow> \<B> rep x = 0"
+ using wf_invalid_point by (simp add: point_replication_number_def filter_mset_empty_conv)
+
+lemma point_rep_number_inv: "size {# b \<in># \<B> . x \<notin> b #} = \<b> - (\<B> rep x)"
+proof -
+ have "\<b> = size {# b \<in># \<B> . x \<notin> b #} + size {# b \<in># \<B> . x \<in> b #}"
+ using multiset_partition by (metis add.commute size_union)
+ thus ?thesis by (simp add: point_replication_number_def)
+qed
+
+lemma point_rep_num_inv_non_empty: "(\<B> rep x) < \<b> \<Longrightarrow> \<B> \<noteq> {#} \<Longrightarrow> {# b \<in># \<B> . x \<notin> b #} \<noteq> {#}"
+ by (metis diff_zero point_replication_number_def size_empty size_filter_neg verit_comp_simplify1(1))
+
+end
+
+subsubsection \<open>Point Index\<close>
+
+text \<open>The point index of a subset of points in a design, is the number of times those points
+occur together in a block of the design\<close>
+definition points_index :: "'a set multiset \<Rightarrow> 'a set \<Rightarrow> nat" (infix "index" 75) where
+"B index ps \<equiv> size {#b \<in># B . ps \<subseteq> b#}"
+
+lemma points_index_empty [simp]: "{#} index ps = 0"
+ by (simp add: points_index_def)
+
+lemma point_index_distrib: "(B1 + B2) index ps = B1 index ps + B2 index ps"
+ by (simp add: points_index_def)
+
+lemma point_index_diff: "B1 index ps = (B1 + B2) index ps - B2 index ps"
+ by (simp add: points_index_def)
+
+lemma points_index_singleton: "{#b#} index ps = 1 \<longleftrightarrow> ps \<subseteq> b"
+ by (simp add: points_index_def)
+
+lemma points_index_singleton_zero: "\<not> (ps \<subseteq> b) \<Longrightarrow> {#b#} index ps = 0"
+ by (simp add: points_index_def)
+
+lemma points_index_sum: "(\<Sum>\<^sub># B ) index ps = (\<Sum>b \<in># B . (b index ps))"
+ using points_index_empty by (induction B) (auto simp add: point_index_distrib)
+
+lemma points_index_block_image_add_eq:
+ assumes "x \<notin> ps"
+ assumes "B index ps = l"
+ shows "{# insert x b . b \<in># B#} index ps = l"
+ using points_index_def by (metis (no_types, lifting) assms filter_mset_cong
+ image_mset_filter_swap2 points_index_def size_image_mset subset_insert)
+
+lemma points_index_on_set_def [simp]:
+ assumes "finite B"
+ shows "(mset_set B) index ps = card {b \<in> B. ps \<subseteq> b}"
+ by (simp add: points_index_def assms)
+
+lemma points_index_single_rep_num: "B index {x} = B rep x"
+ by (simp add: points_index_def point_replication_number_def)
+
+lemma points_index_pair_rep_num:
+ assumes "\<And> b. b \<in># B \<Longrightarrow> x \<in> b"
+ shows "B index {x, y} = B rep y"
+ using point_replication_number_def points_index_def
+ by (metis assms empty_subsetI filter_mset_cong insert_subset)
+
+lemma points_index_0_left_imp:
+ assumes "B index ps = 0"
+ assumes "b \<in># B"
+ shows "\<not> (ps \<subseteq> b)"
+proof (rule ccontr)
+ assume "\<not> \<not> ps \<subseteq> b"
+ then have a: "ps \<subseteq> b" by auto
+ then have "b \<in># {#bl \<in># B . ps \<subseteq> bl#}" by (simp add: assms(2))
+ thus False by (metis assms(1) count_greater_eq_Suc_zero_iff count_size_set_repr not_less_eq_eq
+ points_index_def size_filter_mset_lesseq)
+qed
+
+lemma points_index_0_right_imp:
+ assumes "\<And> b . b \<in># B \<Longrightarrow> (\<not> ps \<subseteq> b)"
+ shows "B index ps = 0"
+ using assms by (simp add: filter_mset_empty_conv points_index_def)
+
+lemma points_index_0_iff: "B index ps = 0 \<longleftrightarrow> (\<forall> b. b \<in># B \<longrightarrow> (\<not> ps \<subseteq> b))"
+ using points_index_0_left_imp points_index_0_right_imp by metis
+
+lemma points_index_gt0_impl_existance:
+ assumes "B index ps > 0"
+ shows "(\<exists> bl . (bl \<in># B \<and> ps \<subseteq> bl))"
+proof -
+ have "size {#bl \<in># B . ps \<subseteq> bl#} > 0"
+ by (metis assms points_index_def)
+ then obtain bl where "bl \<in># B" and "ps \<subseteq> bl"
+ by (metis filter_mset_empty_conv nonempty_has_size)
+ thus ?thesis by auto
+qed
+
+lemma points_index_one_unique:
+ assumes "B index ps = 1"
+ assumes "bl \<in># B" and "ps \<subseteq> bl" and "bl' \<in># B" and "ps \<subseteq> bl'"
+ shows "bl = bl'"
+proof (rule ccontr)
+ assume assm: "bl \<noteq> bl'"
+ then have bl1: "bl \<in># {#bl \<in># B . ps \<subseteq> bl#}" using assms by simp
+ then have bl2: "bl'\<in># {#bl \<in># B . ps \<subseteq> bl#}" using assms by simp
+ then have "{#bl, bl'#} \<subseteq># {#bl \<in># B . ps \<subseteq> bl#}" using assms by (metis bl1 bl2 points_index_def
+ add_mset_subseteq_single_iff assm mset_subset_eq_single size_single subseteq_mset_size_eql)
+ then have "size {#bl \<in># B . ps \<subseteq> bl#} \<ge> 2" using size_mset_mono by fastforce
+ thus False using assms by (metis numeral_le_one_iff points_index_def semiring_norm(69))
+qed
+
+lemma points_index_one_unique_block:
+ assumes "B index ps = 1"
+ shows "\<exists>! bl . (bl \<in># B \<and> ps \<subseteq> bl)"
+ using assms points_index_gt0_impl_existance points_index_one_unique
+ by (metis zero_less_one)
+
+lemma points_index_one_not_unique_block:
+ assumes "B index ps = 1"
+ assumes "ps \<subseteq> bl"
+ assumes "bl \<in># B"
+ assumes "bl' \<in># B - {#bl#}"
+ shows "\<not> ps \<subseteq> bl'"
+proof -
+ have "B = (B - {#bl#}) + {#bl#}" by (simp add: assms(3))
+ then have "(B - {#bl#}) index ps = B index ps - {#bl#} index ps"
+ by (metis point_index_diff)
+ then have "(B - {#bl#}) index ps = 0" using assms points_index_singleton
+ by (metis diff_self_eq_0)
+ thus ?thesis using assms(4) points_index_0_left_imp by auto
+qed
+
+lemma (in incidence_system) points_index_alt_def: "\<B> index ps = size {#b \<in># \<B> . ps \<subseteq> b#}"
+ by (simp add: points_index_def)
+
+lemma (in incidence_system) points_index_ps_nin: "\<not> (ps \<subseteq> \<V>) \<Longrightarrow> \<B> index ps = 0"
+ using points_index_alt_def filter_mset_empty_conv in_mono size_empty subsetI wf_invalid_point
+ by metis
+
+lemma (in incidence_system) points_index_count_bl:
+ "multiplicity bl \<ge> n \<Longrightarrow> ps \<subseteq> bl \<Longrightarrow> count {#bl \<in># \<B> . ps \<subseteq> bl#} bl \<ge> n"
+ by simp
+
+lemma (in finite_incidence_system) points_index_zero:
+ assumes "card ps > card \<V>"
+ shows "\<B> index ps = 0"
+proof -
+ have "\<And> b. b \<in># \<B> \<Longrightarrow> card ps > card b"
+ using block_size_lt_order card_subset_not_gt_card finite_sets assms by fastforce
+ then have "{#b \<in># \<B> . ps \<subseteq> b#} = {#}"
+ by (simp add: card_subset_not_gt_card filter_mset_empty_conv finite_blocks)
+ thus ?thesis using points_index_alt_def by simp
+qed
+
+lemma (in design) points_index_subset:
+ "x \<subseteq># {#bl \<in># \<B> . ps \<subseteq> bl#} \<Longrightarrow> ps \<subseteq> \<V> \<Longrightarrow> (\<B> index ps) \<ge> (size x)"
+ by (simp add: points_index_def size_mset_mono)
+
+lemma (in design) points_index_count_min: "multiplicity bl \<ge> n \<Longrightarrow> ps \<subseteq> bl \<Longrightarrow> \<B> index ps \<ge> n"
+ using points_index_alt_def set_count_size_min by (metis filter_mset.rep_eq)
+
+subsubsection \<open>Intersection Number\<close>
+
+text \<open>The intersection number of two blocks is the size of the intersection of those blocks. i.e.
+the number of points which occur in both blocks\<close>
+definition intersection_number :: "'a set \<Rightarrow> 'a set \<Rightarrow> int" (infix "|\<inter>|" 70) where
+"b1 |\<inter>| b2 \<equiv> card (b1 \<inter> b2)"
+
+lemma intersection_num_non_neg: "b1 |\<inter>| b2 \<ge> 0"
+ by (simp add: intersection_number_def)
+
+lemma intersection_number_empty_iff:
+ assumes "finite b1"
+ shows "b1 \<inter> b2 = {} \<longleftrightarrow> b1 |\<inter>| b2 = 0"
+ by (simp add: intersection_number_def assms)
+
+lemma intersect_num_commute: "b1 |\<inter>| b2 = b2 |\<inter>| b1"
+ by (simp add: inf_commute intersection_number_def)
+
+definition n_intersect_number :: "'a set \<Rightarrow> nat\<Rightarrow> 'a set \<Rightarrow> int" where
+"n_intersect_number b1 n b2 \<equiv> card { x \<in> Pow (b1 \<inter> b2) . card x = n}"
+
+notation n_intersect_number ("(_ |\<inter>|\<^sub>_ _)" [52, 51, 52] 50)
+
+lemma n_intersect_num_subset_def: "b1 |\<inter>|\<^sub>n b2 = card {x . x \<subseteq> b1 \<inter> b2 \<and> card x = n}"
+ using n_intersect_number_def by auto
+
+lemma n_inter_num_one: "finite b1 \<Longrightarrow> finite b2 \<Longrightarrow> b1 |\<inter>|\<^sub>1 b2 = b1 |\<inter>| b2"
+ using n_intersect_number_def intersection_number_def card_Pow_filter_one
+ by (metis (full_types) finite_Int)
+
+lemma n_inter_num_choose: "finite b1 \<Longrightarrow> finite b2 \<Longrightarrow> b1 |\<inter>|\<^sub>n b2 = (card (b1 \<inter> b2) choose n)"
+ using n_subsets n_intersect_num_subset_def
+ by (metis (full_types) finite_Int)
+
+lemma set_filter_single: "x \<in> A \<Longrightarrow> {a \<in> A . a = x} = {x}"
+ by auto
+
+lemma (in design) n_inter_num_zero:
+ assumes "b1 \<in># \<B>" and "b2 \<in># \<B>"
+ shows "b1 |\<inter>|\<^sub>0 b2 = 1"
+proof -
+ have empty: "\<And>x . finite x \<Longrightarrow> card x = 0 \<Longrightarrow> x = {}"
+ by simp
+ have empt_in: "{} \<in> Pow (b1 \<inter> b2)" by simp
+ have "finite (b1 \<inter> b2)" using finite_blocks assms by simp
+ then have "\<And> x . x \<in> Pow (b1 \<inter> b2) \<Longrightarrow> finite x" by (meson PowD finite_subset)
+ then have "{x \<in> Pow (b1 \<inter> b2) . card x = 0} = {x \<in> Pow (b1 \<inter> b2) . x = {}}"
+ using empty by (metis card.empty)
+ then have "{x \<in> Pow (b1 \<inter> b2) . card x = 0} = {{}}"
+ by (simp add: empt_in set_filter_single Collect_conv_if)
+ thus ?thesis by (simp add: n_intersect_number_def)
+qed
+
+lemma (in design) n_inter_num_choose_design: "b1 \<in># \<B> \<Longrightarrow> b2 \<in># \<B>
+ \<Longrightarrow> b1 |\<inter>|\<^sub>n b2 = (card (b1 \<inter> b2) choose n) "
+ using finite_blocks by (simp add: n_inter_num_choose)
+
+lemma (in design) n_inter_num_choose_design_inter: "b1 \<in># \<B> \<Longrightarrow> b2 \<in># \<B>
+ \<Longrightarrow> b1 |\<inter>|\<^sub>n b2 = (nat (b1 |\<inter>| b2) choose n) "
+ using finite_blocks by (simp add: n_inter_num_choose intersection_number_def)
+
+subsection \<open>Incidence System Set Property Definitions\<close>
+context incidence_system
+begin
+
+text \<open>The set of replication numbers for all points of design\<close>
+definition replication_numbers :: "int set" where
+"replication_numbers \<equiv> {\<B> rep x | x . x \<in> \<V>}"
+
+lemma replication_numbers_non_empty:
+ assumes "\<V> \<noteq> {}"
+ shows "replication_numbers \<noteq> {}"
+ by (simp add: assms replication_numbers_def)
+
+lemma obtain_point_with_rep: "r \<in> replication_numbers \<Longrightarrow> \<exists> x. x \<in> \<V> \<and> \<B> rep x = r"
+ using replication_numbers_def by auto
+
+lemma point_rep_number_in_set: "x \<in> \<V> \<Longrightarrow> (\<B> rep x) \<in> replication_numbers"
+ by (auto simp add: replication_numbers_def)
+
+lemma (in finite_incidence_system) replication_numbers_finite: "finite replication_numbers"
+ using finite_sets by (simp add: replication_numbers_def)
+
+text \<open>The set of all block sizes in a system\<close>
+
+definition sys_block_sizes :: "int set" where
+"sys_block_sizes \<equiv> { (int (card bl)) | bl. bl \<in># \<B>}"
+
+lemma block_sizes_non_empty_set:
+ assumes "\<B> \<noteq> {#}"
+ shows "sys_block_sizes \<noteq> {}"
+by (simp add: sys_block_sizes_def assms)
+
+lemma finite_block_sizes: "finite (sys_block_sizes)"
+ by (simp add: sys_block_sizes_def)
+
+lemma block_sizes_non_empty:
+ assumes "\<B> \<noteq> {#}"
+ shows "card (sys_block_sizes) > 0"
+ using finite_block_sizes block_sizes_non_empty_set
+ by (simp add: assms card_gt_0_iff)
+
+lemma sys_block_sizes_in: "bl \<in># \<B> \<Longrightarrow> card bl \<in> sys_block_sizes"
+ unfolding sys_block_sizes_def by auto
+
+lemma sys_block_sizes_obtain_bl: "x \<in> sys_block_sizes \<Longrightarrow> (\<exists> bl \<in># \<B>. int (card bl) = x)"
+ by (auto simp add: sys_block_sizes_def)
+
+text \<open>The set of all possible intersection numbers in a system.\<close>
+
+definition intersection_numbers :: "int set" where
+"intersection_numbers \<equiv> { b1 |\<inter>| b2 | b1 b2 . b1 \<in># \<B> \<and> b2 \<in># (\<B> - {#b1#})}"
+
+lemma obtain_blocks_intersect_num: "n \<in> intersection_numbers \<Longrightarrow>
+ \<exists> b1 b2. b1 \<in># \<B> \<and> b2 \<in># (\<B> - {#b1#}) \<and> b1 |\<inter>| b2 = n"
+ by (auto simp add: intersection_numbers_def)
+
+lemma intersect_num_in_set: "b1 \<in># \<B> \<Longrightarrow> b2 \<in># (\<B> - {#b1#}) \<Longrightarrow> b1 |\<inter>| b2 \<in> intersection_numbers"
+ by (auto simp add: intersection_numbers_def)
+
+text \<open>The set of all possible point indices\<close>
+definition point_indices :: "int \<Rightarrow> int set" where
+"point_indices t \<equiv> {\<B> index ps | ps. int (card ps) = t \<and> ps \<subseteq> \<V>}"
+
+lemma point_indices_elem_in: "ps \<subseteq> \<V> \<Longrightarrow> card ps = t \<Longrightarrow> \<B> index ps \<in> point_indices t"
+ by (auto simp add: point_indices_def)
+
+lemma point_indices_alt_def: "point_indices t = { \<B> index ps | ps. int (card ps) = t \<and> ps \<subseteq> \<V>}"
+ by (simp add: point_indices_def)
+
+end
+
+subsection \<open>Basic Constructions on designs\<close>
+
+text \<open>This section defines some of the most common universal constructions found in design theory
+involving only a single design\<close>
+
+subsubsection \<open>Design Complements\<close>
+
+context incidence_system
+begin
+
+text \<open>The complement of a block are all the points in the design not in that block.
+The complement of a design is therefore the original point sets, and set of all block complements\<close>
+definition block_complement:: "'a set \<Rightarrow> 'a set" ("_\<^sup>c" [56] 55) where
+"block_complement b \<equiv> \<V> - b"
+
+definition complement_blocks :: "'a set multiset" ("(\<B>\<^sup>C)")where
+"complement_blocks \<equiv> {# bl\<^sup>c . bl \<in># \<B> #}"
+
+lemma block_complement_elem_iff:
+ assumes "ps \<subseteq> \<V>"
+ shows "ps \<subseteq> bl\<^sup>c \<longleftrightarrow> (\<forall> x \<in> ps. x \<notin> bl)"
+ using assms block_complement_def by (auto)
+
+lemma block_complement_inter_empty: "bl1\<^sup>c = bl2 \<Longrightarrow> bl1 \<inter> bl2 = {}"
+ using block_complement_def by auto
+
+lemma block_complement_inv:
+ assumes "bl \<in># \<B>"
+ assumes "bl\<^sup>c = bl2"
+ shows "bl2\<^sup>c = bl"
+ by (metis Diff_Diff_Int assms(1) assms(2) block_complement_def inf.absorb_iff2 wellformed)
+
+lemma block_complement_subset_points: "ps \<subseteq> (bl\<^sup>c) \<Longrightarrow> ps \<subseteq> \<V>"
+ using block_complement_def by blast
+
+lemma obtain_comp_block_orig:
+ assumes "bl1 \<in># \<B>\<^sup>C"
+ obtains bl2 where "bl2 \<in># \<B>" and "bl1 = bl2\<^sup>c"
+ using wellformed assms by (auto simp add: complement_blocks_def)
+
+lemma complement_same_b [simp]: "size \<B>\<^sup>C = size \<B>"
+ by (simp add: complement_blocks_def)
+
+lemma block_comp_elem_alt_left: "x \<in> bl \<Longrightarrow> ps \<subseteq> bl\<^sup>c \<Longrightarrow> x \<notin> ps"
+ by (auto simp add: block_complement_def block_complement_elem_iff)
+
+lemma block_comp_elem_alt_right: "ps \<subseteq> \<V> \<Longrightarrow> (\<And> x . x \<in> ps \<Longrightarrow> x \<notin> bl) \<Longrightarrow> ps \<subseteq> bl\<^sup>c"
+ by (auto simp add: block_complement_elem_iff)
+
+lemma complement_index:
+ assumes "ps \<subseteq> \<V>"
+ shows "\<B>\<^sup>C index ps = size {# b \<in># \<B> . (\<forall> x \<in> ps . x \<notin> b) #}"
+proof -
+ have "\<B>\<^sup>C index ps = size {# b \<in># {# bl\<^sup>c . bl \<in># \<B>#}. ps \<subseteq> b #}"
+ by (simp add: complement_blocks_def points_index_def)
+ then have "\<B>\<^sup>C index ps = size {# bl\<^sup>c | bl \<in># \<B> . ps \<subseteq> bl\<^sup>c #}"
+ by (metis image_mset_filter_swap)
+ thus ?thesis using assms by (simp add: block_complement_elem_iff)
+qed
+
+lemma complement_index_2:
+ assumes "{x, y} \<subseteq> \<V>"
+ shows "\<B>\<^sup>C index {x, y} = size {# b \<in># \<B> . x \<notin> b \<and> y \<notin> b #}"
+proof -
+ have a: "\<And> b. b \<in># \<B> \<Longrightarrow> \<forall> x' \<in> {x, y} . x' \<notin> b \<Longrightarrow> x \<notin> b \<and> y \<notin> b"
+ by simp
+ have "\<And> b. b \<in># \<B> \<Longrightarrow> x \<notin> b \<and> y \<notin> b \<Longrightarrow> \<forall> x' \<in> {x, y} . x' \<notin> b "
+ by simp
+ thus ?thesis using assms a complement_index
+ by (smt (verit) filter_mset_cong)
+qed
+
+lemma complement_rep_number:
+ assumes "x \<in> \<V>" and "\<B> rep x = r"
+ shows "\<B>\<^sup>C rep x = \<b> - r"
+proof -
+ have r: "size {#b \<in># \<B> . x \<in> b#} = r" using assms by (simp add: point_replication_number_def)
+ then have a: "\<And> b . b \<in># \<B> \<Longrightarrow> x \<in> b \<Longrightarrow> x \<notin> b\<^sup>c"
+ by (simp add: block_complement_def)
+ have "\<And> b . b \<in># \<B> \<Longrightarrow> x \<notin> b \<Longrightarrow> x \<in> b\<^sup>c"
+ by (simp add: assms(1) block_complement_def)
+ then have alt: "(image_mset block_complement \<B>) rep x = size {#b \<in># \<B> . x \<notin> b#}"
+ using a filter_mset_cong image_mset_filter_swap2 point_replication_number_def
+ by (smt (verit, ccfv_SIG) size_image_mset)
+ have "\<b> = size {#b \<in># \<B> . x \<in> b#} + size {#b \<in># \<B> . x \<notin> b#}"
+ by (metis multiset_partition size_union)
+ thus ?thesis using alt
+ by (simp add: r complement_blocks_def)
+qed
+
+lemma complement_blocks_wf: "bl \<in># \<B>\<^sup>C \<Longrightarrow> bl \<subseteq> \<V>"
+ by (auto simp add: complement_blocks_def block_complement_def)
+
+lemma complement_wf [intro]: "incidence_system \<V> \<B>\<^sup>C"
+ using complement_blocks_wf by (unfold_locales)
+
+interpretation sys_complement: incidence_system "\<V>" "\<B>\<^sup>C"
+ using complement_wf by simp
+end
+
+context finite_incidence_system
+begin
+lemma block_complement_size: "b \<subseteq> \<V> \<Longrightarrow> card (b\<^sup>c) = card \<V> - card b"
+ by (simp add: block_complement_def card_Diff_subset finite_subset card_mono of_nat_diff finite_sets)
+
+lemma block_comp_incomplete: "incomplete_block bl \<Longrightarrow> card (bl\<^sup>c) > 0"
+ using block_complement_size by (simp add: wellformed)
+
+lemma block_comp_incomplete_nempty: "incomplete_block bl \<Longrightarrow> bl\<^sup>c \<noteq> {}"
+ using wellformed block_complement_def finite_blocks
+ by (auto simp add: block_complement_size block_comp_incomplete card_subset_not_gt_card)
+
+lemma incomplete_block_proper_subset: "incomplete_block bl \<Longrightarrow> bl \<subset> \<V>"
+ using wellformed by fastforce
+
+lemma complement_finite: "finite_incidence_system \<V> \<B>\<^sup>C"
+ using complement_wf finite_sets by (simp add: incidence_system.finite_sysI)
+
+interpretation comp_fin: finite_incidence_system \<V> "\<B>\<^sup>C"
+ using complement_finite by simp
+
+end
+
+context design
+begin
+lemma (in design) complement_design:
+ assumes "\<And> bl . bl \<in># \<B> \<Longrightarrow> incomplete_block bl"
+ shows "design \<V> (\<B>\<^sup>C)"
+proof -
+ interpret fin: finite_incidence_system \<V> "\<B>\<^sup>C" using complement_finite by simp
+ show ?thesis using assms block_comp_incomplete_nempty wellformed
+ by (unfold_locales) (auto simp add: complement_blocks_def)
+qed
+
+end
+subsubsection \<open>Multiples\<close>
+text \<open>An easy way to construct new set systems is to simply multiply the block collection by some
+constant\<close>
+
+context incidence_system
+begin
+
+abbreviation multiple_blocks :: "nat \<Rightarrow> 'a set multiset" where
+"multiple_blocks n \<equiv> repeat_mset n \<B>"
+
+lemma multiple_block_in_original: "b \<in># multiple_blocks n \<Longrightarrow> b \<in># \<B>"
+ by (simp add: elem_in_repeat_in_original)
+
+lemma multiple_block_in: "n > 0 \<Longrightarrow> b \<in># \<B> \<Longrightarrow> b \<in># multiple_blocks n"
+ by (simp add: elem_in_original_in_repeat)
+
+lemma multiple_blocks_gt: "n > 0 \<Longrightarrow> size (multiple_blocks n) \<ge> size \<B>"
+ by (simp)
+
+lemma block_original_count_le: "n > 0 \<Longrightarrow> count \<B> b \<le> count (multiple_blocks n) b"
+ using count_repeat_mset by simp
+
+lemma multiple_blocks_sub: "n > 0 \<Longrightarrow> \<B> \<subseteq># (multiple_blocks n)"
+ by (simp add: mset_subset_eqI block_original_count_le)
+
+lemma multiple_1_same: "multiple_blocks 1 = \<B>"
+ by simp
+
+lemma multiple_unfold_1: "multiple_blocks (Suc n) = (multiple_blocks n) + \<B>"
+ by simp
+
+lemma multiple_point_rep_num: "(multiple_blocks n) rep x = (\<B> rep x) * n"
+proof (induction n)
+ case 0
+ then show ?case by (simp add: point_replication_number_def)
+next
+ case (Suc n)
+ then have "multiple_blocks (Suc n) rep x = \<B> rep x * n + (\<B> rep x)"
+ using Suc.IH Suc.prems by (simp add: union_commute point_replication_number_def)
+ then show ?case
+ by (simp add: int_distrib(2))
+qed
+
+lemma multiple_point_index: "(multiple_blocks n) index ps = (\<B> index ps) * n"
+ by (induction n) (auto simp add: points_index_def)
+
+lemma repeat_mset_block_point_rel: "\<And>b x. b \<in># multiple_blocks n \<Longrightarrow> x \<in> b \<Longrightarrow> x \<in> \<V>"
+ by (induction n) (auto, meson subset_iff wellformed)
+
+lemma multiple_is_wellformed: "incidence_system \<V> (multiple_blocks n)"
+ using repeat_mset_subset_in wellformed repeat_mset_block_point_rel by (unfold_locales) (auto)
+
+lemma multiple_blocks_num [simp]: "size (multiple_blocks n) = n*\<b>"
+ by simp
+
+interpretation mult_sys: incidence_system \<V> "(multiple_blocks n)"
+ by (simp add: multiple_is_wellformed)
+
+lemma multiple_block_multiplicity [simp]: "mult_sys.multiplicity n bl = (multiplicity bl) * n"
+ by (simp)
+
+lemma multiple_block_sizes_same:
+ assumes "n > 0"
+ shows "sys_block_sizes = mult_sys.sys_block_sizes n"
+proof -
+ have def: "mult_sys.sys_block_sizes n = {card bl | bl. bl \<in># (multiple_blocks n)}"
+ by (simp add: mult_sys.sys_block_sizes_def)
+ then have eq: "\<And> bl. bl \<in># (multiple_blocks n) \<longleftrightarrow> bl \<in># \<B>"
+ using assms multiple_block_in multiple_block_in_original by blast
+ thus ?thesis using def by (simp add: sys_block_sizes_def eq)
+qed
+
+end
+
+context finite_incidence_system
+begin
+
+lemma multiple_is_finite: "finite_incidence_system \<V> (multiple_blocks n)"
+ using multiple_is_wellformed finite_sets by (unfold_locales) (auto simp add: incidence_system_def)
+
+end
+
+context design
+begin
+
+lemma multiple_is_design: "design \<V> (multiple_blocks n)"
+proof -
+ interpret fis: finite_incidence_system \<V> "multiple_blocks n" using multiple_is_finite by simp
+ show ?thesis using blocks_nempty
+ by (unfold_locales) (auto simp add: elem_in_repeat_in_original repeat_mset_not_empty)
+qed
+
+end
+
+subsection \<open>Simple Designs\<close>
+
+text \<open>Simple designs are those in which the multiplicity of each block is at most one.
+In other words, the block collection is a set. This can significantly ease reasoning.\<close>
+
+locale simple_incidence_system = incidence_system +
+ assumes simple [simp]: "bl \<in># \<B> \<Longrightarrow> multiplicity bl = 1"
+
+begin
+
+lemma simple_alt_def_all: "\<forall> bl \<in># \<B> . multiplicity bl = 1"
+ using simple by auto
+
+lemma simple_blocks_eq_sup: "mset_set (design_support) = \<B>"
+ using distinct_mset_def simple design_support_def by (metis distinct_mset_set_mset_ident)
+
+lemma simple_block_size_eq_card: "\<b> = card (design_support)"
+ by (metis simple_blocks_eq_sup size_mset_set)
+
+lemma points_index_simple_def: "\<B> index ps = card {b \<in> design_support . ps \<subseteq> b}"
+ using design_support_def points_index_def card_size_filter_eq simple_blocks_eq_sup
+ by (metis finite_set_mset)
+
+lemma replication_num_simple_def: "\<B> rep x = card {b \<in> design_support . x \<in> b}"
+ using design_support_def point_replication_number_def card_size_filter_eq simple_blocks_eq_sup
+ by (metis finite_set_mset)
+
+end
+
+locale simple_design = design + simple_incidence_system
+
+text \<open>Additional reasoning about when something is not simple\<close>
+context incidence_system
+begin
+lemma simple_not_multiplicity: "b \<in># \<B> \<Longrightarrow> multiplicity b > 1 \<Longrightarrow> \<not> simple_incidence_system \<V> \<B>"
+ using simple_incidence_system_def simple_incidence_system_axioms_def by (metis nat_neq_iff)
+
+lemma multiple_not_simple:
+ assumes "n > 1"
+ assumes "\<B> \<noteq> {#}"
+ shows "\<not> simple_incidence_system \<V> (multiple_blocks n)"
+proof (rule ccontr, simp)
+ assume "simple_incidence_system \<V> (multiple_blocks n)"
+ then have "\<And> bl. bl \<in># \<B> \<Longrightarrow> count (multiple_blocks n) bl = 1"
+ using assms(1) elem_in_original_in_repeat
+ by (metis not_gr_zero not_less_zero simple_incidence_system.simple)
+ thus False using assms by auto
+qed
+
+end
+
+subsection \<open>Proper Designs\<close>
+text \<open>Many types of designs rely on parameter conditions that only make sense for non-empty designs.
+i.e. designs with at least one block, and therefore given well-formed condition, at least one point.
+To this end we define the notion of a "proper" design\<close>
+
+locale proper_design = design +
+ assumes b_non_zero: "\<b> \<noteq> 0"
+begin
+
+lemma is_proper: "proper_design \<V> \<B>" by intro_locales
+
+lemma v_non_zero: "\<v> > 0"
+ using b_non_zero v_eq0_imp_b_eq_0 by auto
+
+lemma b_positive: "\<b> > 0" using b_non_zero
+ by (simp add: nonempty_has_size)
+
+lemma design_points_nempty: "\<V> \<noteq> {}"
+ using v_non_zero by auto
+
+lemma design_blocks_nempty: "\<B> \<noteq> {#}"
+ using b_non_zero by auto
+
+end
+
+text \<open>Intro rules for a proper design\<close>
+lemma (in design) proper_designI[intro]: "\<b> \<noteq> 0 \<Longrightarrow> proper_design \<V> \<B>"
+ by (unfold_locales) simp
+
+lemma proper_designII[intro]:
+ assumes "design V B" and "B \<noteq> {#}"
+ shows "proper_design V B"
+proof -
+ interpret des: design V B using assms by simp
+ show ?thesis using assms by unfold_locales simp
+qed
+
+text \<open>Reasoning on construction closure for proper designs\<close>
+context proper_design
+begin
+
+lemma multiple_proper_design:
+ assumes "n > 0"
+ shows "proper_design \<V> (multiple_blocks n)"
+ using multiple_is_design assms design_blocks_nempty multiple_block_in
+ by (metis block_set_nempty_imp_block_ex empty_iff proper_designII set_mset_empty)
+
+lemma complement_proper_design:
+ assumes "\<And> bl . bl \<in># \<B> \<Longrightarrow> incomplete_block bl"
+ shows "proper_design \<V> \<B>\<^sup>C"
+proof -
+ interpret des: design \<V> "\<B>\<^sup>C"
+ by (simp add: assms complement_design)
+ show ?thesis using b_non_zero by (unfold_locales) auto
+qed
+
+end
end
\ No newline at end of file
diff --git a/thys/Design_Theory/Design_Isomorphisms.thy b/thys/Design_Theory/Design_Isomorphisms.thy
--- a/thys/Design_Theory/Design_Isomorphisms.thy
+++ b/thys/Design_Theory/Design_Isomorphisms.thy
@@ -1,398 +1,398 @@
-(* Title: Design_Isomorphisms
- Author: Chelsea Edmonds
-*)
-
-section \<open>Design Isomorphisms\<close>
-
-theory Design_Isomorphisms imports Design_Basics Sub_Designs
-begin
-
-subsection \<open>Images of Set Systems\<close>
-
-text \<open>We loosely define the concept of taking the "image" of a set system, as done in isomorphisms.
-Note that this is not based off mathematical theory, but is for ease of notation\<close>
-definition blocks_image :: "'a set multiset \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'b set multiset" where
-"blocks_image B f \<equiv> image_mset ((`) f) B"
-
-lemma image_block_set_constant_size: "size (B) = size (blocks_image B f)"
- by (simp add: blocks_image_def)
-
-lemma (in incidence_system) image_set_system_wellformed:
- "incidence_system (f ` \<V>) (blocks_image \<B> f)"
- by (unfold_locales, auto simp add: blocks_image_def) (meson image_eqI wf_invalid_point)
-
-lemma (in finite_incidence_system) image_set_system_finite:
- "finite_incidence_system (f ` \<V>) (blocks_image \<B> f)"
- using image_set_system_wellformed finite_sets
- by (intro_locales) (simp_all add: blocks_image_def finite_incidence_system_axioms.intro)
-
-subsection \<open>Incidence System Isomorphisms\<close>
-
-text \<open>Isomorphism's are defined by the Handbook of Combinatorial Designs
-\cite{colbournHandbookCombinatorialDesigns2007}\<close>
-
-locale incidence_system_isomorphism = source: incidence_system \<V> \<B> + target: incidence_system \<V>' \<B>'
- for "\<V>" and "\<B>" and "\<V>'" and "\<B>'" + fixes bij_map ("\<pi>")
- assumes bij: "bij_betw \<pi> \<V> \<V>'"
- assumes block_img: "image_mset ((`) \<pi>) \<B> = \<B>'"
-begin
-
-lemma iso_eq_order: "card \<V> = card \<V>'"
- using bij bij_betw_same_card by auto
-
-lemma iso_eq_block_num: "size \<B> = size \<B>'"
- using block_img by (metis size_image_mset)
-
-lemma iso_block_img_alt_rep: "{# \<pi> ` bl . bl \<in># \<B>#} = \<B>'"
- using block_img by simp
-
-lemma inv_iso_block_img: "image_mset ((`) (inv_into \<V> \<pi>)) \<B>' = \<B>"
-proof -
- have "\<And> x. x \<in> \<V> \<Longrightarrow> ((inv_into \<V> \<pi>) \<circ> \<pi>) x = x"
- using bij bij_betw_inv_into_left comp_apply by fastforce
- then have "\<And> bl x . bl \<in># \<B> \<Longrightarrow> x \<in> bl \<Longrightarrow> ((inv_into \<V> \<pi>) \<circ> \<pi>) x = x"
- using source.wellformed by blast
- then have img: "\<And> bl . bl \<in># \<B> \<Longrightarrow> image ((inv_into \<V> \<pi>) \<circ> \<pi>) bl = bl"
- by simp
- have "image_mset ((`) (inv_into \<V> \<pi>)) \<B>' = image_mset ((`) (inv_into \<V> \<pi>)) (image_mset ((`) \<pi>) \<B>)"
- using block_img by simp
- then have "image_mset ((`) (inv_into \<V> \<pi>)) \<B>' = image_mset ((`) ((inv_into \<V> \<pi>) \<circ> \<pi>)) \<B>"
- by (metis (no_types, opaque_lifting) block_img comp_apply image_comp multiset.map_comp multiset.map_cong0)
- thus ?thesis using img by simp
-qed
-
-lemma inverse_incidence_sys_iso: "incidence_system_isomorphism \<V>' \<B>' \<V> \<B> (inv_into \<V> \<pi>)"
- using bij bij_betw_inv_into inv_iso_block_img by (unfold_locales) simp
-
-lemma iso_points_map: "\<pi> ` \<V> = \<V>'"
- using bij by (simp add: bij_betw_imp_surj_on)
-
-lemma iso_points_inv_map: "(inv_into \<V> \<pi>) ` \<V>' = \<V>"
- using incidence_system_isomorphism.iso_points_map inverse_incidence_sys_iso by blast
-
-lemma iso_points_ss_card:
- assumes "ps \<subseteq> \<V>"
- shows "card ps = card (\<pi> ` ps)"
- using assms bij bij_betw_same_card bij_betw_subset by blast
-
-lemma iso_block_in: "bl \<in># \<B> \<Longrightarrow> (\<pi> ` bl) \<in># \<B>'"
- using iso_block_img_alt_rep
- by (metis image_eqI in_image_mset)
-
-lemma iso_inv_block_in: "x \<in># \<B>' \<Longrightarrow> x \<in> (`) \<pi> ` set_mset \<B>"
- by (metis block_img in_image_mset)
-
-lemma iso_img_block_orig_exists: "x \<in># \<B>' \<Longrightarrow> \<exists> bl . bl \<in># \<B> \<and> x = \<pi> ` bl"
- using iso_inv_block_in by blast
-
-lemma iso_blocks_map_inj: "x \<in># \<B> \<Longrightarrow> y \<in># \<B> \<Longrightarrow> \<pi> ` x = \<pi> ` y \<Longrightarrow> x = y"
- using image_inv_into_cancel incidence_system.wellformed iso_points_inv_map iso_points_map
- by (metis (no_types, lifting) source.incidence_system_axioms subset_image_iff)
-
-lemma iso_bij_betwn_block_sets: "bij_betw ((`) \<pi>) (set_mset \<B>) (set_mset \<B>')"
- apply ( simp add: bij_betw_def inj_on_def)
- using iso_block_in iso_inv_block_in iso_blocks_map_inj by auto
-
-lemma iso_bij_betwn_block_sets_inv: "bij_betw ((`) (inv_into \<V> \<pi>)) (set_mset \<B>') (set_mset \<B>)"
- using incidence_system_isomorphism.iso_bij_betwn_block_sets inverse_incidence_sys_iso by blast
-
-lemma iso_bij_betw_individual_blocks: "bl \<in># \<B> \<Longrightarrow> bij_betw \<pi> bl (\<pi> ` bl)"
- using bij bij_betw_subset source.wellformed by blast
-
-lemma iso_bij_betw_individual_blocks_inv: "bl \<in># \<B> \<Longrightarrow> bij_betw (inv_into \<V> \<pi>) (\<pi> ` bl) bl"
- using bij bij_betw_subset source.wellformed bij_betw_inv_into_subset by fastforce
-
-lemma iso_bij_betw_individual_blocks_inv_alt:
- "bl \<in># \<B>' \<Longrightarrow> bij_betw (inv_into \<V> \<pi>) bl ((inv_into \<V> \<pi>) ` bl)"
- using incidence_system_isomorphism.iso_bij_betw_individual_blocks inverse_incidence_sys_iso
- by blast
-
-lemma iso_inv_block_in_alt: "(\<pi> ` bl) \<in># \<B>' \<Longrightarrow> bl \<subseteq> \<V> \<Longrightarrow> bl \<in># \<B>"
- using image_eqI image_inv_into_cancel inv_iso_block_img iso_points_inv_map
- by (metis (no_types, lifting) iso_points_map multiset.set_map subset_image_iff)
-
-lemma iso_img_block_not_in:
- assumes "x \<notin># \<B>"
- assumes "x \<subseteq> \<V>"
- shows "(\<pi> ` x) \<notin># \<B>'"
-proof (rule ccontr)
- assume a: "\<not> \<pi> ` x \<notin># \<B>'"
- then have a: "\<pi> ` x \<in># \<B>'" by simp
- then have "\<And> y . y \<in> (\<pi> ` x) \<Longrightarrow> (inv_into \<V> \<pi>) y \<in> \<V>"
- using target.wf_invalid_point iso_points_inv_map by auto
- then have "((`) (inv_into \<V> \<pi>)) (\<pi> ` x) \<in># \<B>"
- using iso_bij_betwn_block_sets_inv by (meson a bij_betw_apply)
- thus False
- using a assms(1) assms(2) iso_inv_block_in_alt by blast
-qed
-
-lemma iso_block_multiplicity:
- assumes "bl \<subseteq> \<V>"
- shows "source.multiplicity bl = target.multiplicity (\<pi> ` bl)"
-proof (cases "bl \<in># \<B>")
- case True
- have "inj_on ((`) \<pi>) (set_mset \<B>)"
- using bij_betw_imp_inj_on iso_bij_betwn_block_sets by auto
- then have "count \<B> bl = count \<B>' (\<pi> ` bl)"
- using count_image_mset_le_count_inj_on count_image_mset_ge_count True block_img inv_into_f_f
- less_le_not_le order.not_eq_order_implies_strict by metis
- thus ?thesis by simp
-next
- case False
- have s_mult: "source.multiplicity bl = 0"
- by (simp add: False count_eq_zero_iff)
- then have "target.multiplicity (\<pi> ` bl) = 0"
- using False count_inI iso_inv_block_in_alt
- by (metis assms)
- thus ?thesis
- using s_mult by simp
-qed
-
-lemma iso_point_in_block_img_iff: "p \<in> \<V> \<Longrightarrow> bl \<in># \<B> \<Longrightarrow> p \<in> bl \<longleftrightarrow> (\<pi> p) \<in> (\<pi> ` bl)"
- by (metis bij bij_betw_imp_surj_on iso_bij_betw_individual_blocks_inv bij_betw_inv_into_left imageI)
-
-lemma iso_point_subset_block_iff: "p \<subseteq> \<V> \<Longrightarrow> bl \<in># \<B> \<Longrightarrow> p \<subseteq> bl \<longleftrightarrow> (\<pi> ` p) \<subseteq> (\<pi> ` bl)"
- apply auto
- using image_subset_iff iso_point_in_block_img_iff subset_iff by metis
-
-lemma iso_is_image_block: "\<B>' = blocks_image \<B> \<pi>"
- unfolding blocks_image_def by (simp add: block_img iso_points_map)
-
-end
-
-subsection \<open>Design Isomorphisms\<close>
-text \<open>Apply the concept of isomorphisms to designs only\<close>
-
-locale design_isomorphism = incidence_system_isomorphism \<V> \<B> \<V>' \<B>' \<pi> + source: design \<V> \<B> +
- target: design \<V>' \<B>' for \<V> and \<B> and \<V>' and \<B>' and bij_map ("\<pi>")
-
-context design_isomorphism
-begin
-
-lemma inverse_design_isomorphism: "design_isomorphism \<V>' \<B>' \<V> \<B> (inv_into \<V> \<pi>)"
- using inverse_incidence_sys_iso source.wf_design target.wf_design
- by (simp add: design_isomorphism.intro)
-
-end
-
-subsubsection \<open>Isomorphism Operation\<close>
-text \<open>Define the concept of isomorphic designs outside the scope of locale\<close>
-
-definition isomorphic_designs (infixl "\<cong>\<^sub>D" 50) where
-"\<D> \<cong>\<^sub>D \<D>' \<longleftrightarrow> (\<exists> \<pi> . design_isomorphism (fst \<D>) (snd \<D>) (fst \<D>') (snd \<D>') \<pi>)"
-
-lemma isomorphic_designs_symmetric: "(\<V>, \<B>) \<cong>\<^sub>D (\<V>', \<B>') \<Longrightarrow> (\<V>', \<B>') \<cong>\<^sub>D (\<V>, \<B>)"
- using isomorphic_designs_def design_isomorphism.inverse_design_isomorphism
- by metis
-
-lemma isomorphic_designs_implies_bij: "(\<V>, \<B>) \<cong>\<^sub>D (\<V>', \<B>') \<Longrightarrow> \<exists> \<pi> . bij_betw \<pi> \<V> \<V>'"
- using incidence_system_isomorphism.bij isomorphic_designs_def
- by (metis design_isomorphism.axioms(1) fst_conv)
-
-lemma isomorphic_designs_implies_block_map: "(\<V>, \<B>) \<cong>\<^sub>D (\<V>', \<B>') \<Longrightarrow> \<exists> \<pi> . image_mset ((`) \<pi>) \<B> = \<B>'"
- using incidence_system_isomorphism.block_img isomorphic_designs_def
- using design_isomorphism.axioms(1) by fastforce
-
-context design
-begin
-
-lemma isomorphic_designsI [intro]: "design \<V>' \<B>' \<Longrightarrow> bij_betw \<pi> \<V> \<V>' \<Longrightarrow> image_mset ((`) \<pi>) \<B> = \<B>'
- \<Longrightarrow> (\<V>, \<B>) \<cong>\<^sub>D (\<V>', \<B>')"
- using design_isomorphism.intro isomorphic_designs_def wf_design image_set_system_wellformed
- by (metis bij_betw_imp_surj_on blocks_image_def fst_conv incidence_system_axioms
- incidence_system_isomorphism.intro incidence_system_isomorphism_axioms_def snd_conv)
-
-lemma eq_designs_isomorphic:
- assumes "\<V> = \<V>'"
- assumes "\<B> = \<B>'"
- shows "(\<V>, \<B>) \<cong>\<^sub>D (\<V>', \<B>')"
-proof -
- interpret d1: design \<V> \<B> using assms
- using wf_design by auto
- interpret d2: design \<V>' \<B>' using assms
- using wf_design by blast
- have "design_isomorphism \<V> \<B> \<V>' \<B>' id" using assms by (unfold_locales) simp_all
- thus ?thesis unfolding isomorphic_designs_def by auto
-qed
-
-end
-
-context design_isomorphism
-begin
-
-subsubsection \<open>Design Properties/Operations under Isomorphism\<close>
-
-lemma design_iso_point_rep_num_eq:
- assumes "p \<in> \<V>"
- shows "\<B> rep p = \<B>' rep (\<pi> p)"
-proof -
- have "{#b \<in># \<B> . p \<in> b#} = {#b \<in># \<B> . \<pi> p \<in> \<pi> ` b#}"
- using assms filter_mset_cong iso_point_in_block_img_iff assms by force
- then have "{#b \<in># \<B>' . \<pi> p \<in> b#} = image_mset ((`) \<pi>) {#b \<in># \<B> . p \<in> b#}"
- by (simp add: image_mset_filter_swap block_img)
- thus ?thesis
- by (simp add: point_replication_number_def)
-qed
-
-lemma design_iso_rep_numbers_eq: "source.replication_numbers = target.replication_numbers"
- apply (simp add: source.replication_numbers_def target.replication_numbers_def)
- using design_iso_point_rep_num_eq design_isomorphism.design_iso_point_rep_num_eq iso_points_map
- by (metis (no_types, lifting) inverse_design_isomorphism iso_points_inv_map rev_image_eqI)
-
-lemma design_iso_block_size_eq: "bl \<in># \<B> \<Longrightarrow> card bl = card (\<pi> ` bl)"
- using card_image_le finite_subset_image image_inv_into_cancel
- by (metis iso_points_inv_map iso_points_map le_antisym source.finite_blocks source.wellformed)
-
-lemma design_iso_block_sizes_eq: "source.sys_block_sizes = target.sys_block_sizes"
- apply (simp add: source.sys_block_sizes_def target.sys_block_sizes_def)
- by (metis (no_types, lifting) design_iso_block_size_eq iso_block_in iso_img_block_orig_exists)
-
-lemma design_iso_points_index_eq:
- assumes "ps \<subseteq> \<V>"
- shows "\<B> index ps = \<B>' index (\<pi> ` ps)"
-proof -
- have "\<And> b . b \<in># \<B> \<Longrightarrow> ((ps \<subseteq> b) = ((\<pi> ` ps) \<subseteq> \<pi> ` b))"
- using iso_point_subset_block_iff assms by blast
- then have "{#b \<in># \<B> . ps \<subseteq> b#} = {#b \<in># \<B> . (\<pi> ` ps) \<subseteq> (\<pi> ` b)#}"
- using assms filter_mset_cong by force
- then have "{#b \<in># \<B>' . \<pi> ` ps \<subseteq> b#} = image_mset ((`) \<pi>) {#b \<in># \<B> . ps \<subseteq> b#}"
- by (simp add: image_mset_filter_swap block_img)
- thus ?thesis
- by (simp add: points_index_def)
-qed
-
-lemma design_iso_points_indices_imp:
- assumes "x \<in> source.point_indices t"
- shows "x \<in> target.point_indices t"
-proof -
- obtain ps where t: "card ps = t" and ss: "ps \<subseteq> \<V>" and x: "\<B> index ps = x" using assms
- by (auto simp add: source.point_indices_def)
- then have x_val: "x = \<B>' index (\<pi> ` ps)" using design_iso_points_index_eq by auto
- have x_img: " (\<pi> ` ps) \<subseteq> \<V>'"
- using ss bij iso_points_map by fastforce
- then have "card (\<pi> ` ps) = t" using t ss iso_points_ss_card by auto
- then show ?thesis using target.point_indices_elem_in x_img x_val by blast
-qed
-
-lemma design_iso_points_indices_eq: "source.point_indices t = target.point_indices t"
- using inverse_design_isomorphism design_isomorphism.design_iso_points_indices_imp
- design_iso_points_indices_imp by blast
-
-lemma design_iso_block_intersect_num_eq:
- assumes "b1 \<in># \<B>"
- assumes "b2 \<in># \<B>"
- shows "b1 |\<inter>| b2 = (\<pi> ` b1) |\<inter>| (\<pi> ` b2)"
-proof -
- have split: "\<pi> ` (b1 \<inter> b2) = (\<pi> ` b1) \<inter> (\<pi> ` b2)" using assms bij bij_betw_inter_subsets
- by (metis source.wellformed)
- thus ?thesis using source.wellformed
- by (simp add: intersection_number_def iso_points_ss_card split assms(2) inf.coboundedI2)
-qed
-
-lemma design_iso_inter_numbers_imp:
- assumes "x \<in> source.intersection_numbers"
- shows "x \<in> target.intersection_numbers"
-proof -
- obtain b1 b2 where 1: "b1 \<in># \<B>" and 2: "b2 \<in># (remove1_mset b1 \<B>)" and xval: "x = b1 |\<inter>| b2"
- using assms by (auto simp add: source.intersection_numbers_def)
- then have pi1: "\<pi> ` b1 \<in># \<B>'" by (simp add: iso_block_in)
- have pi2: "\<pi> ` b2 \<in># (remove1_mset (\<pi> ` b1) \<B>')" using iso_block_in 2
- by (metis (no_types, lifting) "1" block_img image_mset_remove1_mset_if in_remove1_mset_neq
- iso_blocks_map_inj more_than_one_mset_mset_diff multiset.set_map)
- have "x = (\<pi> ` b1) |\<inter>| (\<pi> ` b2)" using 1 2 design_iso_block_intersect_num_eq
- by (metis in_diffD xval)
- then have "x \<in> {b1 |\<inter>| b2 | b1 b2 . b1 \<in># \<B>' \<and> b2 \<in># (\<B>' - {#b1#})}"
- using pi1 pi2 by blast
- then show ?thesis by (simp add: target.intersection_numbers_def)
-qed
-
-lemma design_iso_intersection_numbers: "source.intersection_numbers = target.intersection_numbers"
- using inverse_design_isomorphism design_isomorphism.design_iso_inter_numbers_imp
- design_iso_inter_numbers_imp by blast
-
-lemma design_iso_n_intersect_num:
- assumes "b1 \<in># \<B>"
- assumes "b2 \<in># \<B>"
- shows "b1 |\<inter>|\<^sub>n b2 = ((\<pi> ` b1) |\<inter>|\<^sub>n (\<pi> ` b2))"
-proof -
- let ?A = "{x . x \<subseteq> b1 \<and> x \<subseteq> b2 \<and> card x = n}"
- let ?B = "{y . y \<subseteq> (\<pi> ` b1) \<and> y \<subseteq> (\<pi> ` b2) \<and> card y = n}"
- have b1v: "b1 \<subseteq> \<V>" by (simp add: assms(1) source.wellformed)
- have b2v: "b2 \<subseteq> \<V>" by (simp add: assms(2) source.wellformed)
- then have "\<And>x y . x \<subseteq> b1 \<Longrightarrow> x \<subseteq> b2 \<Longrightarrow> y \<subseteq> b1 \<Longrightarrow> y \<subseteq> b2 \<Longrightarrow> \<pi> ` x = \<pi> ` y \<Longrightarrow> x = y"
- using b1v bij by (metis bij_betw_imp_surj_on bij_betw_inv_into_subset dual_order.trans)
- then have inj: "inj_on ((`) \<pi>) ?A" by (simp add: inj_on_def)
- have eqcard: "\<And>xa. xa \<subseteq> b1 \<Longrightarrow> xa \<subseteq> b2 \<Longrightarrow> card (\<pi> ` xa) = card xa" using b1v b2v bij
- using iso_points_ss_card by auto
- have surj: "\<And>x. x \<subseteq> \<pi> ` b1 \<Longrightarrow> x \<subseteq> \<pi> ` b2 \<Longrightarrow>
- x \<in> {(\<pi> ` xa) | xa . xa \<subseteq> b1 \<and> xa \<subseteq> b2 \<and> card xa = card x}"
- proof -
- fix x
- assume x1: "x \<subseteq> \<pi> ` b1" and x2: "x \<subseteq> \<pi> ` b2"
- then obtain xa where eq_x: "\<pi> ` xa = x" and ss: "xa \<subseteq> \<V>"
- by (metis b1v dual_order.trans subset_imageE)
- then have f1: "xa \<subseteq> b1" by (simp add: x1 assms(1) iso_point_subset_block_iff)
- then have f2: "xa \<subseteq> b2" by (simp add: eq_x ss assms(2) iso_point_subset_block_iff x2)
- then have f3: "card xa = card x" using bij by (simp add: eq_x ss iso_points_ss_card)
- then show "x \<in> {(\<pi> ` xa) | xa . xa \<subseteq> b1 \<and> xa \<subseteq> b2 \<and> card xa = card x}"
- using f1 f2 f3 \<open>\<pi> ` xa = x\<close> by auto
- qed
- have "bij_betw ( (`) \<pi>) ?A ?B"
- proof (auto simp add: bij_betw_def)
- show "inj_on ((`) \<pi>) {x. x \<subseteq> b1 \<and> x \<subseteq> b2 \<and> card x = n}" using inj by simp
- show "\<And>xa. xa \<subseteq> b1 \<Longrightarrow> xa \<subseteq> b2 \<Longrightarrow> n = card xa \<Longrightarrow> card (\<pi> ` xa) = card xa"
- using eqcard by simp
- show "\<And>x. x \<subseteq> \<pi> ` b1 \<Longrightarrow> x \<subseteq> \<pi> ` b2 \<Longrightarrow> n = card x \<Longrightarrow>
- x \<in> (`) \<pi> ` {xa. xa \<subseteq> b1 \<and> xa \<subseteq> b2 \<and> card xa = card x}"
- using surj by (simp add: setcompr_eq_image)
- qed
- thus ?thesis
- using bij_betw_same_card by (auto simp add: n_intersect_number_def)
-qed
-
-lemma subdesign_iso_implies:
- assumes "sub_set_system V B \<V> \<B>"
- shows "sub_set_system (\<pi> ` V) (blocks_image B \<pi>) \<V>' \<B>'"
-proof (unfold_locales)
- show "\<pi> ` V \<subseteq> \<V>'"
- by (metis assms image_mono iso_points_map sub_set_system.points_subset)
- show "blocks_image B \<pi> \<subseteq># \<B>'"
- by (metis assms block_img blocks_image_def image_mset_subseteq_mono sub_set_system.blocks_subset)
-qed
-
-lemma subdesign_image_is_design:
- assumes "sub_set_system V B \<V> \<B>"
- assumes "design V B"
- shows "design (\<pi> ` V) (blocks_image B \<pi>)"
-proof -
- interpret fin: finite_incidence_system "(\<pi> ` V)" "(blocks_image B \<pi>)" using assms(2)
- by (simp add: design.axioms(1) finite_incidence_system.image_set_system_finite)
- interpret des: sub_design V B \<V> \<B> using assms design.wf_design_iff
- by (unfold_locales, auto simp add: sub_set_system.points_subset sub_set_system.blocks_subset)
- have bl_img: "blocks_image B \<pi> \<subseteq># \<B>'"
- by (simp add: blocks_image_def des.blocks_subset image_mset_subseteq_mono iso_is_image_block)
- then show ?thesis
- proof (unfold_locales, auto)
- show "{} \<in># blocks_image B \<pi> \<Longrightarrow> False"
- using assms subdesign_iso_implies target.blocks_nempty bl_img by auto
- qed
-qed
-
-lemma sub_design_isomorphism:
- assumes "sub_set_system V B \<V> \<B>"
- assumes "design V B"
- shows "design_isomorphism V B (\<pi> ` V) (blocks_image B \<pi>) \<pi>"
-proof -
- interpret design "(\<pi> ` V)" "(blocks_image B \<pi>)"
- by (simp add: assms(1) assms(2) subdesign_image_is_design)
- interpret des: design V B by fact
- show ?thesis
- proof (unfold_locales)
- show "bij_betw \<pi> V (\<pi> ` V)" using bij
- by (metis assms(1) bij_betw_subset sub_set_system.points_subset)
- show "image_mset ((`) \<pi>) B = blocks_image B \<pi>" by (simp add: blocks_image_def)
- qed
-qed
-
-end
+(* Title: Design_Isomorphisms
+ Author: Chelsea Edmonds
+*)
+
+section \<open>Design Isomorphisms\<close>
+
+theory Design_Isomorphisms imports Design_Basics Sub_Designs
+begin
+
+subsection \<open>Images of Set Systems\<close>
+
+text \<open>We loosely define the concept of taking the "image" of a set system, as done in isomorphisms.
+Note that this is not based off mathematical theory, but is for ease of notation\<close>
+definition blocks_image :: "'a set multiset \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'b set multiset" where
+"blocks_image B f \<equiv> image_mset ((`) f) B"
+
+lemma image_block_set_constant_size: "size (B) = size (blocks_image B f)"
+ by (simp add: blocks_image_def)
+
+lemma (in incidence_system) image_set_system_wellformed:
+ "incidence_system (f ` \<V>) (blocks_image \<B> f)"
+ by (unfold_locales, auto simp add: blocks_image_def) (meson image_eqI wf_invalid_point)
+
+lemma (in finite_incidence_system) image_set_system_finite:
+ "finite_incidence_system (f ` \<V>) (blocks_image \<B> f)"
+ using image_set_system_wellformed finite_sets
+ by (intro_locales) (simp_all add: blocks_image_def finite_incidence_system_axioms.intro)
+
+subsection \<open>Incidence System Isomorphisms\<close>
+
+text \<open>Isomorphism's are defined by the Handbook of Combinatorial Designs
+\cite{colbournHandbookCombinatorialDesigns2007}\<close>
+
+locale incidence_system_isomorphism = source: incidence_system \<V> \<B> + target: incidence_system \<V>' \<B>'
+ for "\<V>" and "\<B>" and "\<V>'" and "\<B>'" + fixes bij_map ("\<pi>")
+ assumes bij: "bij_betw \<pi> \<V> \<V>'"
+ assumes block_img: "image_mset ((`) \<pi>) \<B> = \<B>'"
+begin
+
+lemma iso_eq_order: "card \<V> = card \<V>'"
+ using bij bij_betw_same_card by auto
+
+lemma iso_eq_block_num: "size \<B> = size \<B>'"
+ using block_img by (metis size_image_mset)
+
+lemma iso_block_img_alt_rep: "{# \<pi> ` bl . bl \<in># \<B>#} = \<B>'"
+ using block_img by simp
+
+lemma inv_iso_block_img: "image_mset ((`) (inv_into \<V> \<pi>)) \<B>' = \<B>"
+proof -
+ have "\<And> x. x \<in> \<V> \<Longrightarrow> ((inv_into \<V> \<pi>) \<circ> \<pi>) x = x"
+ using bij bij_betw_inv_into_left comp_apply by fastforce
+ then have "\<And> bl x . bl \<in># \<B> \<Longrightarrow> x \<in> bl \<Longrightarrow> ((inv_into \<V> \<pi>) \<circ> \<pi>) x = x"
+ using source.wellformed by blast
+ then have img: "\<And> bl . bl \<in># \<B> \<Longrightarrow> image ((inv_into \<V> \<pi>) \<circ> \<pi>) bl = bl"
+ by simp
+ have "image_mset ((`) (inv_into \<V> \<pi>)) \<B>' = image_mset ((`) (inv_into \<V> \<pi>)) (image_mset ((`) \<pi>) \<B>)"
+ using block_img by simp
+ then have "image_mset ((`) (inv_into \<V> \<pi>)) \<B>' = image_mset ((`) ((inv_into \<V> \<pi>) \<circ> \<pi>)) \<B>"
+ by (metis (no_types, opaque_lifting) block_img comp_apply image_comp multiset.map_comp multiset.map_cong0)
+ thus ?thesis using img by simp
+qed
+
+lemma inverse_incidence_sys_iso: "incidence_system_isomorphism \<V>' \<B>' \<V> \<B> (inv_into \<V> \<pi>)"
+ using bij bij_betw_inv_into inv_iso_block_img by (unfold_locales) simp
+
+lemma iso_points_map: "\<pi> ` \<V> = \<V>'"
+ using bij by (simp add: bij_betw_imp_surj_on)
+
+lemma iso_points_inv_map: "(inv_into \<V> \<pi>) ` \<V>' = \<V>"
+ using incidence_system_isomorphism.iso_points_map inverse_incidence_sys_iso by blast
+
+lemma iso_points_ss_card:
+ assumes "ps \<subseteq> \<V>"
+ shows "card ps = card (\<pi> ` ps)"
+ using assms bij bij_betw_same_card bij_betw_subset by blast
+
+lemma iso_block_in: "bl \<in># \<B> \<Longrightarrow> (\<pi> ` bl) \<in># \<B>'"
+ using iso_block_img_alt_rep
+ by (metis image_eqI in_image_mset)
+
+lemma iso_inv_block_in: "x \<in># \<B>' \<Longrightarrow> x \<in> (`) \<pi> ` set_mset \<B>"
+ by (metis block_img in_image_mset)
+
+lemma iso_img_block_orig_exists: "x \<in># \<B>' \<Longrightarrow> \<exists> bl . bl \<in># \<B> \<and> x = \<pi> ` bl"
+ using iso_inv_block_in by blast
+
+lemma iso_blocks_map_inj: "x \<in># \<B> \<Longrightarrow> y \<in># \<B> \<Longrightarrow> \<pi> ` x = \<pi> ` y \<Longrightarrow> x = y"
+ using image_inv_into_cancel incidence_system.wellformed iso_points_inv_map iso_points_map
+ by (metis (no_types, lifting) source.incidence_system_axioms subset_image_iff)
+
+lemma iso_bij_betwn_block_sets: "bij_betw ((`) \<pi>) (set_mset \<B>) (set_mset \<B>')"
+ apply ( simp add: bij_betw_def inj_on_def)
+ using iso_block_in iso_inv_block_in iso_blocks_map_inj by auto
+
+lemma iso_bij_betwn_block_sets_inv: "bij_betw ((`) (inv_into \<V> \<pi>)) (set_mset \<B>') (set_mset \<B>)"
+ using incidence_system_isomorphism.iso_bij_betwn_block_sets inverse_incidence_sys_iso by blast
+
+lemma iso_bij_betw_individual_blocks: "bl \<in># \<B> \<Longrightarrow> bij_betw \<pi> bl (\<pi> ` bl)"
+ using bij bij_betw_subset source.wellformed by blast
+
+lemma iso_bij_betw_individual_blocks_inv: "bl \<in># \<B> \<Longrightarrow> bij_betw (inv_into \<V> \<pi>) (\<pi> ` bl) bl"
+ using bij bij_betw_subset source.wellformed bij_betw_inv_into_subset by fastforce
+
+lemma iso_bij_betw_individual_blocks_inv_alt:
+ "bl \<in># \<B>' \<Longrightarrow> bij_betw (inv_into \<V> \<pi>) bl ((inv_into \<V> \<pi>) ` bl)"
+ using incidence_system_isomorphism.iso_bij_betw_individual_blocks inverse_incidence_sys_iso
+ by blast
+
+lemma iso_inv_block_in_alt: "(\<pi> ` bl) \<in># \<B>' \<Longrightarrow> bl \<subseteq> \<V> \<Longrightarrow> bl \<in># \<B>"
+ using image_eqI image_inv_into_cancel inv_iso_block_img iso_points_inv_map
+ by (metis (no_types, lifting) iso_points_map multiset.set_map subset_image_iff)
+
+lemma iso_img_block_not_in:
+ assumes "x \<notin># \<B>"
+ assumes "x \<subseteq> \<V>"
+ shows "(\<pi> ` x) \<notin># \<B>'"
+proof (rule ccontr)
+ assume a: "\<not> \<pi> ` x \<notin># \<B>'"
+ then have a: "\<pi> ` x \<in># \<B>'" by simp
+ then have "\<And> y . y \<in> (\<pi> ` x) \<Longrightarrow> (inv_into \<V> \<pi>) y \<in> \<V>"
+ using target.wf_invalid_point iso_points_inv_map by auto
+ then have "((`) (inv_into \<V> \<pi>)) (\<pi> ` x) \<in># \<B>"
+ using iso_bij_betwn_block_sets_inv by (meson a bij_betw_apply)
+ thus False
+ using a assms(1) assms(2) iso_inv_block_in_alt by blast
+qed
+
+lemma iso_block_multiplicity:
+ assumes "bl \<subseteq> \<V>"
+ shows "source.multiplicity bl = target.multiplicity (\<pi> ` bl)"
+proof (cases "bl \<in># \<B>")
+ case True
+ have "inj_on ((`) \<pi>) (set_mset \<B>)"
+ using bij_betw_imp_inj_on iso_bij_betwn_block_sets by auto
+ then have "count \<B> bl = count \<B>' (\<pi> ` bl)"
+ using count_image_mset_le_count_inj_on count_image_mset_ge_count True block_img inv_into_f_f
+ less_le_not_le order.not_eq_order_implies_strict by metis
+ thus ?thesis by simp
+next
+ case False
+ have s_mult: "source.multiplicity bl = 0"
+ by (simp add: False count_eq_zero_iff)
+ then have "target.multiplicity (\<pi> ` bl) = 0"
+ using False count_inI iso_inv_block_in_alt
+ by (metis assms)
+ thus ?thesis
+ using s_mult by simp
+qed
+
+lemma iso_point_in_block_img_iff: "p \<in> \<V> \<Longrightarrow> bl \<in># \<B> \<Longrightarrow> p \<in> bl \<longleftrightarrow> (\<pi> p) \<in> (\<pi> ` bl)"
+ by (metis bij bij_betw_imp_surj_on iso_bij_betw_individual_blocks_inv bij_betw_inv_into_left imageI)
+
+lemma iso_point_subset_block_iff: "p \<subseteq> \<V> \<Longrightarrow> bl \<in># \<B> \<Longrightarrow> p \<subseteq> bl \<longleftrightarrow> (\<pi> ` p) \<subseteq> (\<pi> ` bl)"
+ apply auto
+ using image_subset_iff iso_point_in_block_img_iff subset_iff by metis
+
+lemma iso_is_image_block: "\<B>' = blocks_image \<B> \<pi>"
+ unfolding blocks_image_def by (simp add: block_img iso_points_map)
+
+end
+
+subsection \<open>Design Isomorphisms\<close>
+text \<open>Apply the concept of isomorphisms to designs only\<close>
+
+locale design_isomorphism = incidence_system_isomorphism \<V> \<B> \<V>' \<B>' \<pi> + source: design \<V> \<B> +
+ target: design \<V>' \<B>' for \<V> and \<B> and \<V>' and \<B>' and bij_map ("\<pi>")
+
+context design_isomorphism
+begin
+
+lemma inverse_design_isomorphism: "design_isomorphism \<V>' \<B>' \<V> \<B> (inv_into \<V> \<pi>)"
+ using inverse_incidence_sys_iso source.wf_design target.wf_design
+ by (simp add: design_isomorphism.intro)
+
+end
+
+subsubsection \<open>Isomorphism Operation\<close>
+text \<open>Define the concept of isomorphic designs outside the scope of locale\<close>
+
+definition isomorphic_designs (infixl "\<cong>\<^sub>D" 50) where
+"\<D> \<cong>\<^sub>D \<D>' \<longleftrightarrow> (\<exists> \<pi> . design_isomorphism (fst \<D>) (snd \<D>) (fst \<D>') (snd \<D>') \<pi>)"
+
+lemma isomorphic_designs_symmetric: "(\<V>, \<B>) \<cong>\<^sub>D (\<V>', \<B>') \<Longrightarrow> (\<V>', \<B>') \<cong>\<^sub>D (\<V>, \<B>)"
+ using isomorphic_designs_def design_isomorphism.inverse_design_isomorphism
+ by metis
+
+lemma isomorphic_designs_implies_bij: "(\<V>, \<B>) \<cong>\<^sub>D (\<V>', \<B>') \<Longrightarrow> \<exists> \<pi> . bij_betw \<pi> \<V> \<V>'"
+ using incidence_system_isomorphism.bij isomorphic_designs_def
+ by (metis design_isomorphism.axioms(1) fst_conv)
+
+lemma isomorphic_designs_implies_block_map: "(\<V>, \<B>) \<cong>\<^sub>D (\<V>', \<B>') \<Longrightarrow> \<exists> \<pi> . image_mset ((`) \<pi>) \<B> = \<B>'"
+ using incidence_system_isomorphism.block_img isomorphic_designs_def
+ using design_isomorphism.axioms(1) by fastforce
+
+context design
+begin
+
+lemma isomorphic_designsI [intro]: "design \<V>' \<B>' \<Longrightarrow> bij_betw \<pi> \<V> \<V>' \<Longrightarrow> image_mset ((`) \<pi>) \<B> = \<B>'
+ \<Longrightarrow> (\<V>, \<B>) \<cong>\<^sub>D (\<V>', \<B>')"
+ using design_isomorphism.intro isomorphic_designs_def wf_design image_set_system_wellformed
+ by (metis bij_betw_imp_surj_on blocks_image_def fst_conv incidence_system_axioms
+ incidence_system_isomorphism.intro incidence_system_isomorphism_axioms_def snd_conv)
+
+lemma eq_designs_isomorphic:
+ assumes "\<V> = \<V>'"
+ assumes "\<B> = \<B>'"
+ shows "(\<V>, \<B>) \<cong>\<^sub>D (\<V>', \<B>')"
+proof -
+ interpret d1: design \<V> \<B> using assms
+ using wf_design by auto
+ interpret d2: design \<V>' \<B>' using assms
+ using wf_design by blast
+ have "design_isomorphism \<V> \<B> \<V>' \<B>' id" using assms by (unfold_locales) simp_all
+ thus ?thesis unfolding isomorphic_designs_def by auto
+qed
+
+end
+
+context design_isomorphism
+begin
+
+subsubsection \<open>Design Properties/Operations under Isomorphism\<close>
+
+lemma design_iso_point_rep_num_eq:
+ assumes "p \<in> \<V>"
+ shows "\<B> rep p = \<B>' rep (\<pi> p)"
+proof -
+ have "{#b \<in># \<B> . p \<in> b#} = {#b \<in># \<B> . \<pi> p \<in> \<pi> ` b#}"
+ using assms filter_mset_cong iso_point_in_block_img_iff assms by force
+ then have "{#b \<in># \<B>' . \<pi> p \<in> b#} = image_mset ((`) \<pi>) {#b \<in># \<B> . p \<in> b#}"
+ by (simp add: image_mset_filter_swap block_img)
+ thus ?thesis
+ by (simp add: point_replication_number_def)
+qed
+
+lemma design_iso_rep_numbers_eq: "source.replication_numbers = target.replication_numbers"
+ apply (simp add: source.replication_numbers_def target.replication_numbers_def)
+ using design_iso_point_rep_num_eq design_isomorphism.design_iso_point_rep_num_eq iso_points_map
+ by (metis (no_types, lifting) inverse_design_isomorphism iso_points_inv_map rev_image_eqI)
+
+lemma design_iso_block_size_eq: "bl \<in># \<B> \<Longrightarrow> card bl = card (\<pi> ` bl)"
+ using card_image_le finite_subset_image image_inv_into_cancel
+ by (metis iso_points_inv_map iso_points_map le_antisym source.finite_blocks source.wellformed)
+
+lemma design_iso_block_sizes_eq: "source.sys_block_sizes = target.sys_block_sizes"
+ apply (simp add: source.sys_block_sizes_def target.sys_block_sizes_def)
+ by (metis (no_types, lifting) design_iso_block_size_eq iso_block_in iso_img_block_orig_exists)
+
+lemma design_iso_points_index_eq:
+ assumes "ps \<subseteq> \<V>"
+ shows "\<B> index ps = \<B>' index (\<pi> ` ps)"
+proof -
+ have "\<And> b . b \<in># \<B> \<Longrightarrow> ((ps \<subseteq> b) = ((\<pi> ` ps) \<subseteq> \<pi> ` b))"
+ using iso_point_subset_block_iff assms by blast
+ then have "{#b \<in># \<B> . ps \<subseteq> b#} = {#b \<in># \<B> . (\<pi> ` ps) \<subseteq> (\<pi> ` b)#}"
+ using assms filter_mset_cong by force
+ then have "{#b \<in># \<B>' . \<pi> ` ps \<subseteq> b#} = image_mset ((`) \<pi>) {#b \<in># \<B> . ps \<subseteq> b#}"
+ by (simp add: image_mset_filter_swap block_img)
+ thus ?thesis
+ by (simp add: points_index_def)
+qed
+
+lemma design_iso_points_indices_imp:
+ assumes "x \<in> source.point_indices t"
+ shows "x \<in> target.point_indices t"
+proof -
+ obtain ps where t: "card ps = t" and ss: "ps \<subseteq> \<V>" and x: "\<B> index ps = x" using assms
+ by (auto simp add: source.point_indices_def)
+ then have x_val: "x = \<B>' index (\<pi> ` ps)" using design_iso_points_index_eq by auto
+ have x_img: " (\<pi> ` ps) \<subseteq> \<V>'"
+ using ss bij iso_points_map by fastforce
+ then have "card (\<pi> ` ps) = t" using t ss iso_points_ss_card by auto
+ then show ?thesis using target.point_indices_elem_in x_img x_val by blast
+qed
+
+lemma design_iso_points_indices_eq: "source.point_indices t = target.point_indices t"
+ using inverse_design_isomorphism design_isomorphism.design_iso_points_indices_imp
+ design_iso_points_indices_imp by blast
+
+lemma design_iso_block_intersect_num_eq:
+ assumes "b1 \<in># \<B>"
+ assumes "b2 \<in># \<B>"
+ shows "b1 |\<inter>| b2 = (\<pi> ` b1) |\<inter>| (\<pi> ` b2)"
+proof -
+ have split: "\<pi> ` (b1 \<inter> b2) = (\<pi> ` b1) \<inter> (\<pi> ` b2)" using assms bij bij_betw_inter_subsets
+ by (metis source.wellformed)
+ thus ?thesis using source.wellformed
+ by (simp add: intersection_number_def iso_points_ss_card split assms(2) inf.coboundedI2)
+qed
+
+lemma design_iso_inter_numbers_imp:
+ assumes "x \<in> source.intersection_numbers"
+ shows "x \<in> target.intersection_numbers"
+proof -
+ obtain b1 b2 where 1: "b1 \<in># \<B>" and 2: "b2 \<in># (remove1_mset b1 \<B>)" and xval: "x = b1 |\<inter>| b2"
+ using assms by (auto simp add: source.intersection_numbers_def)
+ then have pi1: "\<pi> ` b1 \<in># \<B>'" by (simp add: iso_block_in)
+ have pi2: "\<pi> ` b2 \<in># (remove1_mset (\<pi> ` b1) \<B>')" using iso_block_in 2
+ by (metis (no_types, lifting) "1" block_img image_mset_remove1_mset_if in_remove1_mset_neq
+ iso_blocks_map_inj more_than_one_mset_mset_diff multiset.set_map)
+ have "x = (\<pi> ` b1) |\<inter>| (\<pi> ` b2)" using 1 2 design_iso_block_intersect_num_eq
+ by (metis in_diffD xval)
+ then have "x \<in> {b1 |\<inter>| b2 | b1 b2 . b1 \<in># \<B>' \<and> b2 \<in># (\<B>' - {#b1#})}"
+ using pi1 pi2 by blast
+ then show ?thesis by (simp add: target.intersection_numbers_def)
+qed
+
+lemma design_iso_intersection_numbers: "source.intersection_numbers = target.intersection_numbers"
+ using inverse_design_isomorphism design_isomorphism.design_iso_inter_numbers_imp
+ design_iso_inter_numbers_imp by blast
+
+lemma design_iso_n_intersect_num:
+ assumes "b1 \<in># \<B>"
+ assumes "b2 \<in># \<B>"
+ shows "b1 |\<inter>|\<^sub>n b2 = ((\<pi> ` b1) |\<inter>|\<^sub>n (\<pi> ` b2))"
+proof -
+ let ?A = "{x . x \<subseteq> b1 \<and> x \<subseteq> b2 \<and> card x = n}"
+ let ?B = "{y . y \<subseteq> (\<pi> ` b1) \<and> y \<subseteq> (\<pi> ` b2) \<and> card y = n}"
+ have b1v: "b1 \<subseteq> \<V>" by (simp add: assms(1) source.wellformed)
+ have b2v: "b2 \<subseteq> \<V>" by (simp add: assms(2) source.wellformed)
+ then have "\<And>x y . x \<subseteq> b1 \<Longrightarrow> x \<subseteq> b2 \<Longrightarrow> y \<subseteq> b1 \<Longrightarrow> y \<subseteq> b2 \<Longrightarrow> \<pi> ` x = \<pi> ` y \<Longrightarrow> x = y"
+ using b1v bij by (metis bij_betw_imp_surj_on bij_betw_inv_into_subset dual_order.trans)
+ then have inj: "inj_on ((`) \<pi>) ?A" by (simp add: inj_on_def)
+ have eqcard: "\<And>xa. xa \<subseteq> b1 \<Longrightarrow> xa \<subseteq> b2 \<Longrightarrow> card (\<pi> ` xa) = card xa" using b1v b2v bij
+ using iso_points_ss_card by auto
+ have surj: "\<And>x. x \<subseteq> \<pi> ` b1 \<Longrightarrow> x \<subseteq> \<pi> ` b2 \<Longrightarrow>
+ x \<in> {(\<pi> ` xa) | xa . xa \<subseteq> b1 \<and> xa \<subseteq> b2 \<and> card xa = card x}"
+ proof -
+ fix x
+ assume x1: "x \<subseteq> \<pi> ` b1" and x2: "x \<subseteq> \<pi> ` b2"
+ then obtain xa where eq_x: "\<pi> ` xa = x" and ss: "xa \<subseteq> \<V>"
+ by (metis b1v dual_order.trans subset_imageE)
+ then have f1: "xa \<subseteq> b1" by (simp add: x1 assms(1) iso_point_subset_block_iff)
+ then have f2: "xa \<subseteq> b2" by (simp add: eq_x ss assms(2) iso_point_subset_block_iff x2)
+ then have f3: "card xa = card x" using bij by (simp add: eq_x ss iso_points_ss_card)
+ then show "x \<in> {(\<pi> ` xa) | xa . xa \<subseteq> b1 \<and> xa \<subseteq> b2 \<and> card xa = card x}"
+ using f1 f2 f3 \<open>\<pi> ` xa = x\<close> by auto
+ qed
+ have "bij_betw ( (`) \<pi>) ?A ?B"
+ proof (auto simp add: bij_betw_def)
+ show "inj_on ((`) \<pi>) {x. x \<subseteq> b1 \<and> x \<subseteq> b2 \<and> card x = n}" using inj by simp
+ show "\<And>xa. xa \<subseteq> b1 \<Longrightarrow> xa \<subseteq> b2 \<Longrightarrow> n = card xa \<Longrightarrow> card (\<pi> ` xa) = card xa"
+ using eqcard by simp
+ show "\<And>x. x \<subseteq> \<pi> ` b1 \<Longrightarrow> x \<subseteq> \<pi> ` b2 \<Longrightarrow> n = card x \<Longrightarrow>
+ x \<in> (`) \<pi> ` {xa. xa \<subseteq> b1 \<and> xa \<subseteq> b2 \<and> card xa = card x}"
+ using surj by (simp add: setcompr_eq_image)
+ qed
+ thus ?thesis
+ using bij_betw_same_card by (auto simp add: n_intersect_number_def)
+qed
+
+lemma subdesign_iso_implies:
+ assumes "sub_set_system V B \<V> \<B>"
+ shows "sub_set_system (\<pi> ` V) (blocks_image B \<pi>) \<V>' \<B>'"
+proof (unfold_locales)
+ show "\<pi> ` V \<subseteq> \<V>'"
+ by (metis assms image_mono iso_points_map sub_set_system.points_subset)
+ show "blocks_image B \<pi> \<subseteq># \<B>'"
+ by (metis assms block_img blocks_image_def image_mset_subseteq_mono sub_set_system.blocks_subset)
+qed
+
+lemma subdesign_image_is_design:
+ assumes "sub_set_system V B \<V> \<B>"
+ assumes "design V B"
+ shows "design (\<pi> ` V) (blocks_image B \<pi>)"
+proof -
+ interpret fin: finite_incidence_system "(\<pi> ` V)" "(blocks_image B \<pi>)" using assms(2)
+ by (simp add: design.axioms(1) finite_incidence_system.image_set_system_finite)
+ interpret des: sub_design V B \<V> \<B> using assms design.wf_design_iff
+ by (unfold_locales, auto simp add: sub_set_system.points_subset sub_set_system.blocks_subset)
+ have bl_img: "blocks_image B \<pi> \<subseteq># \<B>'"
+ by (simp add: blocks_image_def des.blocks_subset image_mset_subseteq_mono iso_is_image_block)
+ then show ?thesis
+ proof (unfold_locales, auto)
+ show "{} \<in># blocks_image B \<pi> \<Longrightarrow> False"
+ using assms subdesign_iso_implies target.blocks_nempty bl_img by auto
+ qed
+qed
+
+lemma sub_design_isomorphism:
+ assumes "sub_set_system V B \<V> \<B>"
+ assumes "design V B"
+ shows "design_isomorphism V B (\<pi> ` V) (blocks_image B \<pi>) \<pi>"
+proof -
+ interpret design "(\<pi> ` V)" "(blocks_image B \<pi>)"
+ by (simp add: assms(1) assms(2) subdesign_image_is_design)
+ interpret des: design V B by fact
+ show ?thesis
+ proof (unfold_locales)
+ show "bij_betw \<pi> V (\<pi> ` V)" using bij
+ by (metis assms(1) bij_betw_subset sub_set_system.points_subset)
+ show "image_mset ((`) \<pi>) B = blocks_image B \<pi>" by (simp add: blocks_image_def)
+ qed
+qed
+
+end
end
\ No newline at end of file
diff --git a/thys/Design_Theory/Design_Theory_Root.thy b/thys/Design_Theory/Design_Theory_Root.thy
--- a/thys/Design_Theory/Design_Theory_Root.thy
+++ b/thys/Design_Theory/Design_Theory_Root.thy
@@ -1,21 +1,21 @@
-(* Title: Design_Theory_Root.thy
- Author: Chelsea Edmonds
-*)
-
-theory Design_Theory_Root
-imports
- Multisets_Extras
-
- Design_Basics
- Design_Operations
- Block_Designs
- BIBD
-
- Resolvable_Designs
- Group_Divisible_Designs
- Designs_And_Graphs
- Design_Isomorphisms
- Sub_Designs
-begin
-
+(* Title: Design_Theory_Root.thy
+ Author: Chelsea Edmonds
+*)
+
+theory Design_Theory_Root
+imports
+ Multisets_Extras
+
+ Design_Basics
+ Design_Operations
+ Block_Designs
+ BIBD
+
+ Resolvable_Designs
+ Group_Divisible_Designs
+ Designs_And_Graphs
+ Design_Isomorphisms
+ Sub_Designs
+begin
+
end
\ No newline at end of file
diff --git a/thys/Design_Theory/Designs_And_Graphs.thy b/thys/Design_Theory/Designs_And_Graphs.thy
--- a/thys/Design_Theory/Designs_And_Graphs.thy
+++ b/thys/Design_Theory/Designs_And_Graphs.thy
@@ -1,336 +1,336 @@
-(* Title: Designs_And_Graphs.thy
- Author: Chelsea Edmonds
-*)
-
-section \<open>Graphs and Designs\<close>
-text \<open>There are many links between graphs and design - most fundamentally that graphs are designs\<close>
-
-theory Designs_And_Graphs imports Block_Designs Graph_Theory.Digraph Graph_Theory.Digraph_Component
-begin
-
-subsection \<open>Non-empty digraphs\<close>
-text \<open>First, we define the concept of a non-empty digraph. This mirrors the idea of a "proper design"
-in the design theory library\<close>
-locale non_empty_digraph = wf_digraph +
- assumes arcs_not_empty: "arcs G \<noteq> {}"
-
-begin
-
-lemma verts_not_empty: "verts G \<noteq> {}"
- using wf arcs_not_empty head_in_verts by fastforce
-
-end
-
-subsection \<open>Arcs to Blocks\<close>
-text \<open>A digraph uses a pair of points to define an ordered edge. In the case of simple graphs,
-both possible orderings will be in the arcs set. Blocks are inherently unordered, and as such
-a method is required to convert between the two representations\<close>
-context graph
-begin
-
-definition arc_to_block :: "'b \<Rightarrow> 'a set" where
- "arc_to_block e \<equiv> {tail G e, head G e}"
-
-lemma arc_to_block_to_ends: "{fst (arc_to_ends G e), snd (arc_to_ends G e)} = arc_to_block e"
- by (simp add: arc_to_ends_def arc_to_block_def)
-
-lemma arc_to_block_to_ends_swap: "{snd (arc_to_ends G e), fst (arc_to_ends G e)} = arc_to_block e"
- using arc_to_block_to_ends
- by (simp add: arc_to_block_to_ends insert_commute)
-
-lemma arc_to_ends_to_block: "arc_to_block e = {x, y} \<Longrightarrow>
- arc_to_ends G e = (x, y) \<or> arc_to_ends G e = (y, x)"
- by (metis arc_to_block_def arc_to_ends_def doubleton_eq_iff)
-
-lemma arc_to_block_sym: "arc_to_ends G e1 = (u, v) \<Longrightarrow> arc_to_ends G e2 = (v, u) \<Longrightarrow>
- arc_to_block e1 = arc_to_block e2"
- by (simp add: arc_to_block_def arc_to_ends_def insert_commute)
-
-definition arcs_blocks :: "'a set multiset" where
-"arcs_blocks \<equiv> mset_set (arc_to_block ` (arcs G))"
-
-lemma arcs_blocks_ends: "(x, y) \<in> arcs_ends G \<Longrightarrow> {x, y} \<in># arcs_blocks"
-proof (auto simp add: arcs_ends_def arcs_blocks_def )
- fix xa
- assume assm1: "(x, y) = arc_to_ends G xa" and assm2: "xa \<in> arcs G"
- obtain z where zin: "z \<in> (arc_to_block ` (arcs G))" and "z = arc_to_block xa"
- using assm2 by blast
- thus "{x, y} \<in> arc_to_block ` (arcs G)" using assm1 arc_to_block_to_ends
- by (metis fst_conv snd_conv)
-qed
-
-lemma arc_ends_blocks_subset: "E \<subseteq> arcs G \<Longrightarrow> (x, y) \<in> ((arc_to_ends G) ` E) \<Longrightarrow>
- {x, y} \<in> (arc_to_block ` E)"
- by (auto simp add: arc_to_ends_def arc_to_block_def )
-
-lemma arc_blocks_end_subset: assumes "E \<subseteq> arcs G" and "{x, y} \<in> (arc_to_block ` E)"
- shows "(x, y) \<in> ((arc_to_ends G) ` E) \<or> (y, x) \<in> ((arc_to_ends G) ` E)"
-proof -
- obtain e where "e \<in> E" and "arc_to_block e = {x,y}" using assms
- by fastforce
- then have "arc_to_ends G e = (x, y) \<or> arc_to_ends G e = (y, x)"
- using arc_to_ends_to_block by simp
- thus ?thesis
- by (metis \<open>e \<in> E\<close> image_iff)
-qed
-
-lemma arcs_ends_blocks: "{x, y} \<in># arcs_blocks \<Longrightarrow> (x, y) \<in> arcs_ends G \<and> (y, x) \<in> arcs_ends G"
-proof (auto simp add: arcs_ends_def arcs_blocks_def )
- fix xa
- assume assm1: "{x, y} = arc_to_block xa" and assm2: "xa \<in> (arcs G)"
- obtain z where zin: "z \<in> (arc_to_ends G ` (arcs G))" and "z = arc_to_ends G xa"
- using assm2 by blast
- then have "z = (x, y) \<or> z = (y, x)" using arc_to_block_to_ends assm1
- by (metis arc_to_ends_def doubleton_eq_iff fst_conv snd_conv) (* Slow *)
- thus "(x, y) \<in> arc_to_ends G ` (arcs G)" using assm2
- by (metis arcs_ends_def arcs_ends_symmetric sym_arcs zin)
-next
- fix xa
- assume assm1: "{x, y} = arc_to_block xa" and assm2: "xa \<in> (arcs G)"
- thus "(y, x) \<in> arc_to_ends G ` arcs G" using arcs_ends_def
- by (metis dual_order.refl graph.arc_blocks_end_subset graph_axioms graph_symmetric imageI)
-qed
-
-lemma arcs_blocks_iff: "{x, y} \<in># arcs_blocks \<longleftrightarrow> (x, y) \<in> arcs_ends G \<and> (y, x) \<in> arcs_ends G"
- using arcs_ends_blocks arcs_blocks_ends by blast
-
-lemma arcs_ends_wf: "(x, y) \<in> arcs_ends G \<Longrightarrow> x \<in> verts G \<and> y \<in> verts G"
- by auto
-
-lemma arcs_blocks_elem: "bl \<in># arcs_blocks \<Longrightarrow> \<exists> x y . bl = {x, y}"
- apply (auto simp add: arcs_blocks_def)
- by (meson arc_to_block_def)
-
-lemma arcs_ends_blocks_wf:
- assumes "bl \<in># arcs_blocks"
- shows "bl \<subseteq> verts G"
-proof -
- obtain x y where blpair: "bl = {x, y}" using arcs_blocks_elem assms
- by fastforce
- then have "(x, y) \<in> arcs_ends G" using arcs_ends_blocks assms by simp
- thus ?thesis using arcs_ends_wf blpair by auto
-qed
-
-lemma arcs_blocks_simple: "bl \<in># arcs_blocks \<Longrightarrow> count (arcs_blocks) bl = 1"
- by (simp add: arcs_blocks_def)
-
-lemma arcs_blocks_ne: "arcs G \<noteq> {} \<Longrightarrow> arcs_blocks \<noteq> {#}"
- by (simp add: arcs_blocks_iff arcs_blocks_def mset_set_empty_iff)
-
-end
-
-subsection \<open>Graphs are designs\<close>
-
-text \<open>Prove that a graph is a number of different types of designs\<close>
-sublocale graph \<subseteq> design "verts G" "arcs_blocks"
- using arcs_ends_blocks_wf arcs_blocks_elem by (unfold_locales) (auto)
-
-sublocale graph \<subseteq> simple_design "verts G" "arcs_blocks"
- using arcs_ends_blocks_wf arcs_blocks_elem arcs_blocks_simple by (unfold_locales) (auto)
-
-locale non_empty_graph = graph + non_empty_digraph
-
-sublocale non_empty_graph \<subseteq> proper_design "verts G" "arcs_blocks"
- using arcs_blocks_ne arcs_not_empty by (unfold_locales) simp
-
-lemma (in graph) graph_block_size: assumes "bl \<in># arcs_blocks" shows "card bl = 2"
-proof -
- obtain x y where blrep: "bl = {x, y}" using assms arcs_blocks_elem
- by fastforce
- then have "(x, y) \<in> arcs_ends G" using arcs_ends_blocks assms by simp
- then have "x \<noteq> y" using no_loops using adj_not_same by blast
- thus ?thesis using blrep by simp
-qed
-
-sublocale non_empty_graph \<subseteq> block_design "verts G" "arcs_blocks" 2
- using arcs_not_empty graph_block_size arcs_blocks_ne by (unfold_locales) simp_all
-
-subsection \<open>R-regular graphs\<close>
-text \<open>To reason on r-regular graphs and their link to designs, we require a number of extensions to
-lemmas reasoning around the degrees of vertices\<close>
-context sym_digraph
-begin
-
-lemma in_out_arcs_reflexive: "v \<in> verts G \<Longrightarrow> (e \<in> (in_arcs G v) \<Longrightarrow>
- \<exists> e' . (e' \<in> (out_arcs G v) \<and> head G e' = tail G e))"
- using symmetric_conv sym_arcs by fastforce
-
-lemma out_in_arcs_reflexive: "v \<in> verts G \<Longrightarrow> (e \<in> (out_arcs G v) \<Longrightarrow>
- \<exists> e' . (e' \<in> (in_arcs G v) \<and> tail G e' = head G e))"
- using symmetric_conv sym_arcs by fastforce
-
-end
-
-context nomulti_digraph
-begin
-
-lemma in_arcs_single_per_vert:
- assumes "v \<in> verts G" and "u \<in> verts G"
- assumes "e1 \<in> in_arcs G v" and " e2 \<in> in_arcs G v"
- assumes "tail G e1 = u" and "tail G e2 = u"
- shows "e1 = e2"
-proof -
- have in_arcs1: "e1 \<in> arcs G" and in_arcs2: "e2 \<in> arcs G" using assms by auto
- have "arc_to_ends G e1 = arc_to_ends G e2" using assms arc_to_ends_def
- by (metis in_in_arcs_conv)
- thus ?thesis using in_arcs1 in_arcs2 no_multi_arcs by simp
-qed
-
-lemma out_arcs_single_per_vert:
- assumes "v \<in> verts G" and "u \<in> verts G"
- assumes "e1 \<in> out_arcs G v" and " e2 \<in> out_arcs G v"
- assumes "head G e1 = u" and "head G e2 = u"
- shows "e1 = e2"
-proof -
- have in_arcs1: "e1 \<in> arcs G" and in_arcs2: "e2 \<in> arcs G" using assms by auto
- have "arc_to_ends G e1 = arc_to_ends G e2" using assms arc_to_ends_def
- by (metis in_out_arcs_conv)
- thus ?thesis using in_arcs1 in_arcs2 no_multi_arcs by simp
-qed
-
-end
-
-text \<open>Some helpers on the transformation arc definition\<close>
-context graph
-begin
-
-lemma arc_to_block_is_inj_in_arcs: "inj_on arc_to_block (in_arcs G v)"
- apply (auto simp add: arc_to_block_def inj_on_def)
- by (metis arc_to_ends_def doubleton_eq_iff no_multi_arcs)
-
-lemma arc_to_block_is_inj_out_arcs: "inj_on arc_to_block (out_arcs G v)"
- apply (auto simp add: arc_to_block_def inj_on_def)
- by (metis arc_to_ends_def doubleton_eq_iff no_multi_arcs)
-
-lemma in_out_arcs_reflexive_uniq: "v \<in> verts G \<Longrightarrow> (e \<in> (in_arcs G v) \<Longrightarrow>
- \<exists>! e' . (e' \<in> (out_arcs G v) \<and> head G e' = tail G e))"
- apply auto
- using symmetric_conv apply fastforce
- using out_arcs_single_per_vert by (metis head_in_verts in_out_arcs_conv)
-
-lemma out_in_arcs_reflexive_uniq: "v \<in> verts G \<Longrightarrow> e \<in> (out_arcs G v) \<Longrightarrow>
- \<exists>! e' . (e' \<in> (in_arcs G v) \<and> tail G e' = head G e)"
- apply auto
- using symmetric_conv apply fastforce
- using in_arcs_single_per_vert by (metis tail_in_verts in_in_arcs_conv)
-
-lemma in_eq_out_arc_ends: "(u, v) \<in> ((arc_to_ends G) ` (in_arcs G v)) \<longleftrightarrow>
- (v, u) \<in> ((arc_to_ends G) ` (out_arcs G v))"
- using arc_to_ends_def in_in_arcs_conv in_out_arcs_conv
- by (smt (z3) Pair_inject adj_in_verts(1) dominatesI image_iff out_in_arcs_reflexive_uniq)
-
-lemma in_degree_eq_card_arc_ends: "in_degree G v = card ((arc_to_ends G) ` (in_arcs G v))"
- apply (simp add: in_degree_def)
- using no_multi_arcs by (metis card_image in_arcs_in_arcs inj_onI)
-
-lemma in_degree_eq_card_arc_blocks: "in_degree G v = card (arc_to_block ` (in_arcs G v))"
- apply (simp add: in_degree_def)
- using no_multi_arcs arc_to_block_is_inj_in_arcs by (simp add: card_image)
-
-lemma out_degree_eq_card_arc_blocks: "out_degree G v = card (arc_to_block ` (out_arcs G v))"
- apply (simp add: out_degree_def)
- using no_multi_arcs arc_to_block_is_inj_out_arcs by (simp add: card_image)
-
-lemma out_degree_eq_card_arc_ends: "out_degree G v = card ((arc_to_ends G) ` (out_arcs G v))"
- apply (simp add: out_degree_def)
- using no_multi_arcs by (metis card_image out_arcs_in_arcs inj_onI)
-
-lemma bij_betw_in_out_arcs: "bij_betw (\<lambda> (u, v) . (v, u)) ((arc_to_ends G) ` (in_arcs G v))
- ((arc_to_ends G) ` (out_arcs G v))"
- apply (auto simp add: bij_betw_def)
- apply (simp add: swap_inj_on)
- apply (metis Pair_inject arc_to_ends_def image_eqI in_eq_out_arc_ends in_in_arcs_conv)
- by (metis arc_to_ends_def imageI in_eq_out_arc_ends in_out_arcs_conv pair_imageI)
-
-lemma in_eq_out_degree: "in_degree G v = out_degree G v"
- using bij_betw_in_out_arcs bij_betw_same_card in_degree_eq_card_arc_ends
- out_degree_eq_card_arc_ends by auto
-
-lemma in_out_arcs_blocks: "arc_to_block ` (in_arcs G v) = arc_to_block ` (out_arcs G v)"
-proof (auto)
- fix xa
- assume a1: "xa \<in> arcs G" and a2: "v = head G xa"
- then have "xa \<in> in_arcs G v" by simp
- then obtain e where out: "e \<in> out_arcs G v" and "head G e = tail G xa"
- using out_in_arcs_reflexive_uniq by force
- then have "arc_to_ends G e = (v, tail G xa)"
- by (simp add: arc_to_ends_def)
- then have "arc_to_block xa = arc_to_block e"
- using arc_to_block_sym by (metis a2 arc_to_ends_def)
- then show "arc_to_block xa \<in> arc_to_block ` out_arcs G (head G xa)"
- using out a2 by blast
-next
- fix xa
- assume a1: "xa \<in> arcs G" and a2: "v = tail G xa"
- then have "xa \<in> out_arcs G v" by simp
- then obtain e where ina: "e \<in> in_arcs G v" and "tail G e = head G xa"
- using out_in_arcs_reflexive_uniq by force
- then have "arc_to_ends G e = (head G xa, v)"
- by (simp add: arc_to_ends_def)
- then have "arc_to_block xa = arc_to_block e"
- using arc_to_block_sym by (metis a2 arc_to_ends_def)
- then show "arc_to_block xa \<in> arc_to_block ` in_arcs G (tail G xa)"
- using ina a2 by blast
-qed
-
-end
-
-text \<open>A regular digraph is defined as one where the in degree equals the out degree which in turn
-equals some fixed integer $\mathrm{r}$\<close>
-locale regular_digraph = wf_digraph +
- fixes \<r> :: int
- assumes in_deg_r: "v \<in> verts G \<Longrightarrow> in_degree G v = \<r>"
- assumes out_deg_r: "v \<in> verts G \<Longrightarrow> out_degree G v = \<r>"
-
-locale regular_graph = graph + regular_digraph
-begin
-
-lemma rep_vertices_in_blocks [simp]:
- assumes "x \<in> verts G"
- shows "size {# e \<in># arcs_blocks . x \<in> e #} = \<r>"
-proof -
- have "\<And> e . e \<in> (arc_to_block ` (arcs G)) \<Longrightarrow> x \<in> e \<Longrightarrow> e \<in> (arc_to_block ` in_arcs G x)"
- using arc_to_block_def in_in_arcs_conv insert_commute insert_iff singleton_iff sym_arcs
- symmetric_conv by fastforce
- then have "{ e \<in> (arc_to_block ` (arcs G)) . x \<in> e} = (arc_to_block ` (in_arcs G x))"
- using arc_to_block_def by auto
- then have "card { e \<in> (arc_to_block ` (arcs G)) . x \<in> e} = \<r>"
- using in_deg_r in_degree_eq_card_arc_blocks assms by auto
- thus ?thesis
- using arcs_blocks_def finite_arcs by force
-qed
-
-end
-
-text \<open>Intro rules for regular graphs\<close>
-lemma graph_in_degree_r_imp_reg[intro]: assumes "graph G"
- assumes "(\<And> v . v \<in> (verts G) \<Longrightarrow> in_degree G v = \<r>)"
- shows "regular_graph G \<r>"
-proof -
- interpret g: graph G using assms by simp
- interpret wf: wf_digraph G by (simp add: g.wf_digraph_axioms)
- show ?thesis
- using assms(2) g.in_eq_out_degree by (unfold_locales) simp_all
-qed
-
-lemma graph_out_degree_r_imp_reg[intro]: assumes "graph G"
- assumes "(\<And> v . v \<in> (verts G) \<Longrightarrow> out_degree G v = \<r>)"
- shows "regular_graph G \<r>"
-proof -
- interpret g: graph G using assms by simp
- interpret wf: wf_digraph G by (simp add: g.wf_digraph_axioms)
- show ?thesis
- using assms(2) g.in_eq_out_degree by (unfold_locales) simp_all
-qed
-
-text \<open>Regular graphs (non-empty) can be shown to be a constant rep design\<close>
-locale non_empty_regular_graph = regular_graph + non_empty_digraph
-
-sublocale non_empty_regular_graph \<subseteq> non_empty_graph
- by unfold_locales
-
-sublocale non_empty_regular_graph \<subseteq> constant_rep_design "verts G" "arcs_blocks" \<r>
- using arcs_blocks_ne arcs_not_empty
- by (unfold_locales)(simp_all add: point_replication_number_def)
-
+(* Title: Designs_And_Graphs.thy
+ Author: Chelsea Edmonds
+*)
+
+section \<open>Graphs and Designs\<close>
+text \<open>There are many links between graphs and design - most fundamentally that graphs are designs\<close>
+
+theory Designs_And_Graphs imports Block_Designs Graph_Theory.Digraph Graph_Theory.Digraph_Component
+begin
+
+subsection \<open>Non-empty digraphs\<close>
+text \<open>First, we define the concept of a non-empty digraph. This mirrors the idea of a "proper design"
+in the design theory library\<close>
+locale non_empty_digraph = wf_digraph +
+ assumes arcs_not_empty: "arcs G \<noteq> {}"
+
+begin
+
+lemma verts_not_empty: "verts G \<noteq> {}"
+ using wf arcs_not_empty head_in_verts by fastforce
+
+end
+
+subsection \<open>Arcs to Blocks\<close>
+text \<open>A digraph uses a pair of points to define an ordered edge. In the case of simple graphs,
+both possible orderings will be in the arcs set. Blocks are inherently unordered, and as such
+a method is required to convert between the two representations\<close>
+context graph
+begin
+
+definition arc_to_block :: "'b \<Rightarrow> 'a set" where
+ "arc_to_block e \<equiv> {tail G e, head G e}"
+
+lemma arc_to_block_to_ends: "{fst (arc_to_ends G e), snd (arc_to_ends G e)} = arc_to_block e"
+ by (simp add: arc_to_ends_def arc_to_block_def)
+
+lemma arc_to_block_to_ends_swap: "{snd (arc_to_ends G e), fst (arc_to_ends G e)} = arc_to_block e"
+ using arc_to_block_to_ends
+ by (simp add: arc_to_block_to_ends insert_commute)
+
+lemma arc_to_ends_to_block: "arc_to_block e = {x, y} \<Longrightarrow>
+ arc_to_ends G e = (x, y) \<or> arc_to_ends G e = (y, x)"
+ by (metis arc_to_block_def arc_to_ends_def doubleton_eq_iff)
+
+lemma arc_to_block_sym: "arc_to_ends G e1 = (u, v) \<Longrightarrow> arc_to_ends G e2 = (v, u) \<Longrightarrow>
+ arc_to_block e1 = arc_to_block e2"
+ by (simp add: arc_to_block_def arc_to_ends_def insert_commute)
+
+definition arcs_blocks :: "'a set multiset" where
+"arcs_blocks \<equiv> mset_set (arc_to_block ` (arcs G))"
+
+lemma arcs_blocks_ends: "(x, y) \<in> arcs_ends G \<Longrightarrow> {x, y} \<in># arcs_blocks"
+proof (auto simp add: arcs_ends_def arcs_blocks_def )
+ fix xa
+ assume assm1: "(x, y) = arc_to_ends G xa" and assm2: "xa \<in> arcs G"
+ obtain z where zin: "z \<in> (arc_to_block ` (arcs G))" and "z = arc_to_block xa"
+ using assm2 by blast
+ thus "{x, y} \<in> arc_to_block ` (arcs G)" using assm1 arc_to_block_to_ends
+ by (metis fst_conv snd_conv)
+qed
+
+lemma arc_ends_blocks_subset: "E \<subseteq> arcs G \<Longrightarrow> (x, y) \<in> ((arc_to_ends G) ` E) \<Longrightarrow>
+ {x, y} \<in> (arc_to_block ` E)"
+ by (auto simp add: arc_to_ends_def arc_to_block_def )
+
+lemma arc_blocks_end_subset: assumes "E \<subseteq> arcs G" and "{x, y} \<in> (arc_to_block ` E)"
+ shows "(x, y) \<in> ((arc_to_ends G) ` E) \<or> (y, x) \<in> ((arc_to_ends G) ` E)"
+proof -
+ obtain e where "e \<in> E" and "arc_to_block e = {x,y}" using assms
+ by fastforce
+ then have "arc_to_ends G e = (x, y) \<or> arc_to_ends G e = (y, x)"
+ using arc_to_ends_to_block by simp
+ thus ?thesis
+ by (metis \<open>e \<in> E\<close> image_iff)
+qed
+
+lemma arcs_ends_blocks: "{x, y} \<in># arcs_blocks \<Longrightarrow> (x, y) \<in> arcs_ends G \<and> (y, x) \<in> arcs_ends G"
+proof (auto simp add: arcs_ends_def arcs_blocks_def )
+ fix xa
+ assume assm1: "{x, y} = arc_to_block xa" and assm2: "xa \<in> (arcs G)"
+ obtain z where zin: "z \<in> (arc_to_ends G ` (arcs G))" and "z = arc_to_ends G xa"
+ using assm2 by blast
+ then have "z = (x, y) \<or> z = (y, x)" using arc_to_block_to_ends assm1
+ by (metis arc_to_ends_def doubleton_eq_iff fst_conv snd_conv) (* Slow *)
+ thus "(x, y) \<in> arc_to_ends G ` (arcs G)" using assm2
+ by (metis arcs_ends_def arcs_ends_symmetric sym_arcs zin)
+next
+ fix xa
+ assume assm1: "{x, y} = arc_to_block xa" and assm2: "xa \<in> (arcs G)"
+ thus "(y, x) \<in> arc_to_ends G ` arcs G" using arcs_ends_def
+ by (metis dual_order.refl graph.arc_blocks_end_subset graph_axioms graph_symmetric imageI)
+qed
+
+lemma arcs_blocks_iff: "{x, y} \<in># arcs_blocks \<longleftrightarrow> (x, y) \<in> arcs_ends G \<and> (y, x) \<in> arcs_ends G"
+ using arcs_ends_blocks arcs_blocks_ends by blast
+
+lemma arcs_ends_wf: "(x, y) \<in> arcs_ends G \<Longrightarrow> x \<in> verts G \<and> y \<in> verts G"
+ by auto
+
+lemma arcs_blocks_elem: "bl \<in># arcs_blocks \<Longrightarrow> \<exists> x y . bl = {x, y}"
+ apply (auto simp add: arcs_blocks_def)
+ by (meson arc_to_block_def)
+
+lemma arcs_ends_blocks_wf:
+ assumes "bl \<in># arcs_blocks"
+ shows "bl \<subseteq> verts G"
+proof -
+ obtain x y where blpair: "bl = {x, y}" using arcs_blocks_elem assms
+ by fastforce
+ then have "(x, y) \<in> arcs_ends G" using arcs_ends_blocks assms by simp
+ thus ?thesis using arcs_ends_wf blpair by auto
+qed
+
+lemma arcs_blocks_simple: "bl \<in># arcs_blocks \<Longrightarrow> count (arcs_blocks) bl = 1"
+ by (simp add: arcs_blocks_def)
+
+lemma arcs_blocks_ne: "arcs G \<noteq> {} \<Longrightarrow> arcs_blocks \<noteq> {#}"
+ by (simp add: arcs_blocks_iff arcs_blocks_def mset_set_empty_iff)
+
+end
+
+subsection \<open>Graphs are designs\<close>
+
+text \<open>Prove that a graph is a number of different types of designs\<close>
+sublocale graph \<subseteq> design "verts G" "arcs_blocks"
+ using arcs_ends_blocks_wf arcs_blocks_elem by (unfold_locales) (auto)
+
+sublocale graph \<subseteq> simple_design "verts G" "arcs_blocks"
+ using arcs_ends_blocks_wf arcs_blocks_elem arcs_blocks_simple by (unfold_locales) (auto)
+
+locale non_empty_graph = graph + non_empty_digraph
+
+sublocale non_empty_graph \<subseteq> proper_design "verts G" "arcs_blocks"
+ using arcs_blocks_ne arcs_not_empty by (unfold_locales) simp
+
+lemma (in graph) graph_block_size: assumes "bl \<in># arcs_blocks" shows "card bl = 2"
+proof -
+ obtain x y where blrep: "bl = {x, y}" using assms arcs_blocks_elem
+ by fastforce
+ then have "(x, y) \<in> arcs_ends G" using arcs_ends_blocks assms by simp
+ then have "x \<noteq> y" using no_loops using adj_not_same by blast
+ thus ?thesis using blrep by simp
+qed
+
+sublocale non_empty_graph \<subseteq> block_design "verts G" "arcs_blocks" 2
+ using arcs_not_empty graph_block_size arcs_blocks_ne by (unfold_locales) simp_all
+
+subsection \<open>R-regular graphs\<close>
+text \<open>To reason on r-regular graphs and their link to designs, we require a number of extensions to
+lemmas reasoning around the degrees of vertices\<close>
+context sym_digraph
+begin
+
+lemma in_out_arcs_reflexive: "v \<in> verts G \<Longrightarrow> (e \<in> (in_arcs G v) \<Longrightarrow>
+ \<exists> e' . (e' \<in> (out_arcs G v) \<and> head G e' = tail G e))"
+ using symmetric_conv sym_arcs by fastforce
+
+lemma out_in_arcs_reflexive: "v \<in> verts G \<Longrightarrow> (e \<in> (out_arcs G v) \<Longrightarrow>
+ \<exists> e' . (e' \<in> (in_arcs G v) \<and> tail G e' = head G e))"
+ using symmetric_conv sym_arcs by fastforce
+
+end
+
+context nomulti_digraph
+begin
+
+lemma in_arcs_single_per_vert:
+ assumes "v \<in> verts G" and "u \<in> verts G"
+ assumes "e1 \<in> in_arcs G v" and " e2 \<in> in_arcs G v"
+ assumes "tail G e1 = u" and "tail G e2 = u"
+ shows "e1 = e2"
+proof -
+ have in_arcs1: "e1 \<in> arcs G" and in_arcs2: "e2 \<in> arcs G" using assms by auto
+ have "arc_to_ends G e1 = arc_to_ends G e2" using assms arc_to_ends_def
+ by (metis in_in_arcs_conv)
+ thus ?thesis using in_arcs1 in_arcs2 no_multi_arcs by simp
+qed
+
+lemma out_arcs_single_per_vert:
+ assumes "v \<in> verts G" and "u \<in> verts G"
+ assumes "e1 \<in> out_arcs G v" and " e2 \<in> out_arcs G v"
+ assumes "head G e1 = u" and "head G e2 = u"
+ shows "e1 = e2"
+proof -
+ have in_arcs1: "e1 \<in> arcs G" and in_arcs2: "e2 \<in> arcs G" using assms by auto
+ have "arc_to_ends G e1 = arc_to_ends G e2" using assms arc_to_ends_def
+ by (metis in_out_arcs_conv)
+ thus ?thesis using in_arcs1 in_arcs2 no_multi_arcs by simp
+qed
+
+end
+
+text \<open>Some helpers on the transformation arc definition\<close>
+context graph
+begin
+
+lemma arc_to_block_is_inj_in_arcs: "inj_on arc_to_block (in_arcs G v)"
+ apply (auto simp add: arc_to_block_def inj_on_def)
+ by (metis arc_to_ends_def doubleton_eq_iff no_multi_arcs)
+
+lemma arc_to_block_is_inj_out_arcs: "inj_on arc_to_block (out_arcs G v)"
+ apply (auto simp add: arc_to_block_def inj_on_def)
+ by (metis arc_to_ends_def doubleton_eq_iff no_multi_arcs)
+
+lemma in_out_arcs_reflexive_uniq: "v \<in> verts G \<Longrightarrow> (e \<in> (in_arcs G v) \<Longrightarrow>
+ \<exists>! e' . (e' \<in> (out_arcs G v) \<and> head G e' = tail G e))"
+ apply auto
+ using symmetric_conv apply fastforce
+ using out_arcs_single_per_vert by (metis head_in_verts in_out_arcs_conv)
+
+lemma out_in_arcs_reflexive_uniq: "v \<in> verts G \<Longrightarrow> e \<in> (out_arcs G v) \<Longrightarrow>
+ \<exists>! e' . (e' \<in> (in_arcs G v) \<and> tail G e' = head G e)"
+ apply auto
+ using symmetric_conv apply fastforce
+ using in_arcs_single_per_vert by (metis tail_in_verts in_in_arcs_conv)
+
+lemma in_eq_out_arc_ends: "(u, v) \<in> ((arc_to_ends G) ` (in_arcs G v)) \<longleftrightarrow>
+ (v, u) \<in> ((arc_to_ends G) ` (out_arcs G v))"
+ using arc_to_ends_def in_in_arcs_conv in_out_arcs_conv
+ by (smt (z3) Pair_inject adj_in_verts(1) dominatesI image_iff out_in_arcs_reflexive_uniq)
+
+lemma in_degree_eq_card_arc_ends: "in_degree G v = card ((arc_to_ends G) ` (in_arcs G v))"
+ apply (simp add: in_degree_def)
+ using no_multi_arcs by (metis card_image in_arcs_in_arcs inj_onI)
+
+lemma in_degree_eq_card_arc_blocks: "in_degree G v = card (arc_to_block ` (in_arcs G v))"
+ apply (simp add: in_degree_def)
+ using no_multi_arcs arc_to_block_is_inj_in_arcs by (simp add: card_image)
+
+lemma out_degree_eq_card_arc_blocks: "out_degree G v = card (arc_to_block ` (out_arcs G v))"
+ apply (simp add: out_degree_def)
+ using no_multi_arcs arc_to_block_is_inj_out_arcs by (simp add: card_image)
+
+lemma out_degree_eq_card_arc_ends: "out_degree G v = card ((arc_to_ends G) ` (out_arcs G v))"
+ apply (simp add: out_degree_def)
+ using no_multi_arcs by (metis card_image out_arcs_in_arcs inj_onI)
+
+lemma bij_betw_in_out_arcs: "bij_betw (\<lambda> (u, v) . (v, u)) ((arc_to_ends G) ` (in_arcs G v))
+ ((arc_to_ends G) ` (out_arcs G v))"
+ apply (auto simp add: bij_betw_def)
+ apply (simp add: swap_inj_on)
+ apply (metis Pair_inject arc_to_ends_def image_eqI in_eq_out_arc_ends in_in_arcs_conv)
+ by (metis arc_to_ends_def imageI in_eq_out_arc_ends in_out_arcs_conv pair_imageI)
+
+lemma in_eq_out_degree: "in_degree G v = out_degree G v"
+ using bij_betw_in_out_arcs bij_betw_same_card in_degree_eq_card_arc_ends
+ out_degree_eq_card_arc_ends by auto
+
+lemma in_out_arcs_blocks: "arc_to_block ` (in_arcs G v) = arc_to_block ` (out_arcs G v)"
+proof (auto)
+ fix xa
+ assume a1: "xa \<in> arcs G" and a2: "v = head G xa"
+ then have "xa \<in> in_arcs G v" by simp
+ then obtain e where out: "e \<in> out_arcs G v" and "head G e = tail G xa"
+ using out_in_arcs_reflexive_uniq by force
+ then have "arc_to_ends G e = (v, tail G xa)"
+ by (simp add: arc_to_ends_def)
+ then have "arc_to_block xa = arc_to_block e"
+ using arc_to_block_sym by (metis a2 arc_to_ends_def)
+ then show "arc_to_block xa \<in> arc_to_block ` out_arcs G (head G xa)"
+ using out a2 by blast
+next
+ fix xa
+ assume a1: "xa \<in> arcs G" and a2: "v = tail G xa"
+ then have "xa \<in> out_arcs G v" by simp
+ then obtain e where ina: "e \<in> in_arcs G v" and "tail G e = head G xa"
+ using out_in_arcs_reflexive_uniq by force
+ then have "arc_to_ends G e = (head G xa, v)"
+ by (simp add: arc_to_ends_def)
+ then have "arc_to_block xa = arc_to_block e"
+ using arc_to_block_sym by (metis a2 arc_to_ends_def)
+ then show "arc_to_block xa \<in> arc_to_block ` in_arcs G (tail G xa)"
+ using ina a2 by blast
+qed
+
+end
+
+text \<open>A regular digraph is defined as one where the in degree equals the out degree which in turn
+equals some fixed integer $\mathrm{r}$\<close>
+locale regular_digraph = wf_digraph +
+ fixes \<r> :: int
+ assumes in_deg_r: "v \<in> verts G \<Longrightarrow> in_degree G v = \<r>"
+ assumes out_deg_r: "v \<in> verts G \<Longrightarrow> out_degree G v = \<r>"
+
+locale regular_graph = graph + regular_digraph
+begin
+
+lemma rep_vertices_in_blocks [simp]:
+ assumes "x \<in> verts G"
+ shows "size {# e \<in># arcs_blocks . x \<in> e #} = \<r>"
+proof -
+ have "\<And> e . e \<in> (arc_to_block ` (arcs G)) \<Longrightarrow> x \<in> e \<Longrightarrow> e \<in> (arc_to_block ` in_arcs G x)"
+ using arc_to_block_def in_in_arcs_conv insert_commute insert_iff singleton_iff sym_arcs
+ symmetric_conv by fastforce
+ then have "{ e \<in> (arc_to_block ` (arcs G)) . x \<in> e} = (arc_to_block ` (in_arcs G x))"
+ using arc_to_block_def by auto
+ then have "card { e \<in> (arc_to_block ` (arcs G)) . x \<in> e} = \<r>"
+ using in_deg_r in_degree_eq_card_arc_blocks assms by auto
+ thus ?thesis
+ using arcs_blocks_def finite_arcs by force
+qed
+
+end
+
+text \<open>Intro rules for regular graphs\<close>
+lemma graph_in_degree_r_imp_reg[intro]: assumes "graph G"
+ assumes "(\<And> v . v \<in> (verts G) \<Longrightarrow> in_degree G v = \<r>)"
+ shows "regular_graph G \<r>"
+proof -
+ interpret g: graph G using assms by simp
+ interpret wf: wf_digraph G by (simp add: g.wf_digraph_axioms)
+ show ?thesis
+ using assms(2) g.in_eq_out_degree by (unfold_locales) simp_all
+qed
+
+lemma graph_out_degree_r_imp_reg[intro]: assumes "graph G"
+ assumes "(\<And> v . v \<in> (verts G) \<Longrightarrow> out_degree G v = \<r>)"
+ shows "regular_graph G \<r>"
+proof -
+ interpret g: graph G using assms by simp
+ interpret wf: wf_digraph G by (simp add: g.wf_digraph_axioms)
+ show ?thesis
+ using assms(2) g.in_eq_out_degree by (unfold_locales) simp_all
+qed
+
+text \<open>Regular graphs (non-empty) can be shown to be a constant rep design\<close>
+locale non_empty_regular_graph = regular_graph + non_empty_digraph
+
+sublocale non_empty_regular_graph \<subseteq> non_empty_graph
+ by unfold_locales
+
+sublocale non_empty_regular_graph \<subseteq> constant_rep_design "verts G" "arcs_blocks" \<r>
+ using arcs_blocks_ne arcs_not_empty
+ by (unfold_locales)(simp_all add: point_replication_number_def)
+
end
\ No newline at end of file
diff --git a/thys/Design_Theory/Group_Divisible_Designs.thy b/thys/Design_Theory/Group_Divisible_Designs.thy
--- a/thys/Design_Theory/Group_Divisible_Designs.thy
+++ b/thys/Design_Theory/Group_Divisible_Designs.thy
@@ -1,1091 +1,1091 @@
-(* Title: Group_Divisible_Designs.thy
- Author: Chelsea Edmonds
-*)
-
-section \<open>Group Divisible Designs\<close>
-text \<open>Definitions in this section taken from the handbook \cite{colbournHandbookCombinatorialDesigns2007}
-and Stinson \cite{stinsonCombinatorialDesignsConstructions2004}\<close>
-theory Group_Divisible_Designs imports Resolvable_Designs
-begin
-
-subsection \<open>Group design\<close>
-text \<open>We define a group design to have an additional paramater $G$ which is a partition on the point
-set $V$. This is not defined in the handbook, but is a precursor to GDD's without index constraints\<close>
-
-locale group_design = proper_design +
- fixes groups :: "'a set set" ("\<G>")
- assumes group_partitions: "partition_on \<V> \<G>"
- assumes groups_size: "card \<G> > 1"
-begin
-
-lemma groups_not_empty: "\<G> \<noteq> {}"
- using groups_size by auto
-
-lemma num_groups_lt_points: "card \<G> \<le> \<v>"
- by (simp add: partition_on_le_set_elements finite_sets group_partitions)
-
-lemma groups_disjoint: "disjoint \<G>"
- using group_partitions partition_onD2 by auto
-
-lemma groups_disjoint_pairwise: "G1 \<in> \<G> \<Longrightarrow> G2 \<in> \<G> \<Longrightarrow> G1 \<noteq> G2 \<Longrightarrow> disjnt G1 G2"
- using group_partitions partition_onD2 pairwiseD by fastforce
-
-lemma point_in_one_group: "x \<in> G1 \<Longrightarrow> G1 \<in> \<G> \<Longrightarrow> G2 \<in> \<G> \<Longrightarrow> G1 \<noteq> G2 \<Longrightarrow> x \<notin> G2"
- using groups_disjoint_pairwise by (simp add: disjnt_iff)
-
-lemma point_has_unique_group: "x \<in> \<V> \<Longrightarrow> \<exists>!G. x \<in> G \<and> G \<in> \<G>"
- using partition_on_partition_on_unique group_partitions
- by fastforce
-
-lemma rep_number_point_group_one:
- assumes "x \<in> \<V>"
- shows "card {g \<in> \<G> . x \<in> g} = 1"
-proof -
- obtain g' where "g' \<in> \<G>" and "x \<in> g'"
- using assms point_has_unique_group by blast
- then have "{g \<in> \<G> . x \<in> g} = {g'}"
- using group_partitions partition_onD4 by force
- thus ?thesis
- by simp
-qed
-
-lemma point_in_group: "G \<in> \<G> \<Longrightarrow> x \<in> G \<Longrightarrow> x \<in> \<V>"
- using group_partitions partition_onD1 by auto
-
-lemma point_subset_in_group: "G \<in> \<G> \<Longrightarrow> ps \<subseteq> G \<Longrightarrow> ps \<subseteq> \<V>"
- using point_in_group by auto
-
-lemma group_subset_point_subset: "G \<in> \<G> \<Longrightarrow> G' \<subseteq> G \<Longrightarrow> ps \<subseteq> G' \<Longrightarrow> ps \<subseteq> \<V>"
- using point_subset_in_group by auto
-
-lemma groups_finite: "finite \<G>"
- using finite_elements finite_sets group_partitions by auto
-
-lemma group_elements_finite: "G \<in> \<G> \<Longrightarrow> finite G"
- using groups_finite finite_sets group_partitions
- by (meson finite_subset point_in_group subset_iff)
-
-lemma v_equals_sum_group_sizes: "\<v> = (\<Sum>G \<in> \<G>. card G)"
- using group_partitions groups_disjoint partition_onD1 card_Union_disjoint group_elements_finite
- by fastforce
-
-lemma gdd_min_v: "\<v> \<ge> 2"
-proof -
- have assm: "card \<G> \<ge> 2" using groups_size by simp
- then have "\<And> G . G \<in> \<G> \<Longrightarrow> G \<noteq> {}" using partition_onD3 group_partitions by auto
- then have "\<And> G . G \<in> \<G> \<Longrightarrow> card G \<ge> 1"
- using group_elements_finite card_0_eq by fastforce
- then have " (\<Sum>G \<in> \<G>. card G) \<ge> 2" using assm
- using sum_mono by force
- thus ?thesis using v_equals_sum_group_sizes
- by linarith
-qed
-
-lemma min_group_size: "G \<in> \<G> \<Longrightarrow> card G \<ge> 1"
- using partition_onD3 group_partitions
- using group_elements_finite not_le_imp_less by fastforce
-
-lemma group_size_lt_v:
- assumes "G \<in> \<G>"
- shows "card G < \<v>"
-proof -
- have "(\<Sum>G' \<in> \<G>. card G') = \<v>" using gdd_min_v v_equals_sum_group_sizes
- by linarith
- then have split_sum: "card G + (\<Sum>G' \<in> (\<G> - {G}). card G') = \<v>" using assms sum.remove
- by (metis groups_finite v_equals_sum_group_sizes)
- have "card (\<G> - {G}) \<ge> 1" using groups_size
- by (simp add: assms groups_finite)
- then obtain G' where gin: "G' \<in> (\<G> - {G})"
- by (meson elem_exists_non_empty_set less_le_trans less_numeral_extra(1))
- then have "card G' \<ge> 1" using min_group_size by auto
- then have "(\<Sum>G' \<in> (\<G> - {G}). card G') \<ge> 1"
- by (metis gin finite_Diff groups_finite leI less_one sum_eq_0_iff)
- thus ?thesis using split_sum
- by linarith
-qed
-
-subsubsection \<open>Group Type\<close>
-
-text \<open>GDD's have a "type", which is defined by a sequence of group sizes $g_i$, and the number
-of groups of that size $a_i$: $g_1^{a_1}g2^{a_2}...g_n^{a_n}$\<close>
-definition group_sizes :: "nat set" where
-"group_sizes \<equiv> {card G | G . G \<in> \<G>}"
-
-definition groups_of_size :: "nat \<Rightarrow> nat" where
-"groups_of_size g \<equiv> card { G \<in> \<G> . card G = g }"
-
-definition group_type :: "(nat \<times> nat) set" where
-"group_type \<equiv> {(g, groups_of_size g) | g . g \<in> group_sizes }"
-
-lemma group_sizes_min: "x \<in> group_sizes \<Longrightarrow> x \<ge> 1 "
- unfolding group_sizes_def using min_group_size group_size_lt_v by auto
-
-lemma group_sizes_max: "x \<in> group_sizes \<Longrightarrow> x < \<v> "
- unfolding group_sizes_def using min_group_size group_size_lt_v by auto
-
-lemma group_size_implies_group_existance: "x \<in> group_sizes \<Longrightarrow> \<exists>G. G \<in> \<G> \<and> card G = x"
- unfolding group_sizes_def by auto
-
-lemma groups_of_size_zero: "groups_of_size 0 = 0"
-proof -
- have empty: "{G \<in> \<G> . card G = 0} = {}" using min_group_size
- by fastforce
- thus ?thesis unfolding groups_of_size_def
- by (simp add: empty)
-qed
-
-lemma groups_of_size_max:
- assumes "g \<ge> \<v>"
- shows "groups_of_size g = 0"
-proof -
- have "{G \<in> \<G> . card G = g} = {}" using group_size_lt_v assms by fastforce
- thus ?thesis unfolding groups_of_size_def
- by (simp add: \<open>{G \<in> \<G>. card G = g} = {}\<close>)
-qed
-
-lemma group_type_contained_sizes: "(g, a) \<in> group_type \<Longrightarrow> g \<in> group_sizes"
- unfolding group_type_def by simp
-
-lemma group_type_contained_count: "(g, a) \<in> group_type \<Longrightarrow> card {G \<in> \<G> . card G = g} = a"
- unfolding group_type_def groups_of_size_def by simp
-
-lemma group_card_in_sizes: "g \<in> \<G> \<Longrightarrow> card g \<in> group_sizes"
- unfolding group_sizes_def by auto
-
-lemma group_card_non_zero_groups_of_size_min:
- assumes "g \<in> \<G>"
- assumes "card g = a"
- shows "groups_of_size a \<ge> 1"
-proof -
- have "g \<in> {G \<in> \<G> . card G = a}" using assms by simp
- then have "{G \<in> \<G> . card G = a} \<noteq> {}" by auto
- then have "card {G \<in> \<G> . card G = a} \<noteq> 0"
- by (simp add: groups_finite)
- thus ?thesis unfolding groups_of_size_def by simp
-qed
-
-lemma elem_in_group_sizes_min_of_size:
- assumes "a \<in> group_sizes"
- shows "groups_of_size a \<ge> 1"
- using assms group_card_non_zero_groups_of_size_min group_size_implies_group_existance by blast
-
-lemma group_card_non_zero_groups_of_size_max:
- shows "groups_of_size a \<le> \<v>"
-proof -
- have "{G \<in> \<G> . card G = a} \<subseteq> \<G>" by simp
- then have "card {G \<in> \<G> . card G = a} \<le> card \<G>"
- by (simp add: card_mono groups_finite)
- thus ?thesis
- using groups_of_size_def num_groups_lt_points by auto
-qed
-
-lemma group_card_in_type: "g \<in> \<G> \<Longrightarrow> \<exists> x . (card g, x) \<in> group_type \<and> x \<ge> 1"
- unfolding group_type_def using group_card_non_zero_groups_of_size_min
- by (simp add: group_card_in_sizes)
-
-lemma partition_groups_on_size: "partition_on \<G> {{ G \<in> \<G> . card G = g } | g . g \<in> group_sizes}"
-proof (intro partition_onI, auto)
- fix g
- assume a1: "g \<in> group_sizes"
- assume " \<forall>x. x \<in> \<G> \<longrightarrow> card x \<noteq> g"
- then show False using a1 group_size_implies_group_existance by auto
-next
- fix x
- assume "x \<in> \<G>"
- then show "\<exists>xa. (\<exists>g. xa = {G \<in> \<G>. card G = g} \<and> g \<in> group_sizes) \<and> x \<in> xa"
- using group_card_in_sizes by auto
-qed
-
-lemma group_size_partition_covers_points: "\<Union>(\<Union>{{ G \<in> \<G> . card G = g } | g . g \<in> group_sizes}) = \<V>"
- by (metis (no_types, lifting) group_partitions partition_groups_on_size partition_onD1)
-
-lemma groups_of_size_alt_def_count: "groups_of_size g = count {# card G . G \<in># mset_set \<G> #} g"
-proof -
- have a: "groups_of_size g = card { G \<in> \<G> . card G = g }" unfolding groups_of_size_def by simp
- then have "groups_of_size g = size {# G \<in># (mset_set \<G>) . card G = g #}"
- using groups_finite by auto
- then have size_repr: "groups_of_size g = size {# x \<in># {# card G . G \<in># mset_set \<G> #} . x = g #}"
- using groups_finite by (simp add: filter_mset_image_mset)
- have "group_sizes = set_mset ({# card G . G \<in># mset_set \<G> #})"
- using group_sizes_def groups_finite by auto
- thus ?thesis using size_repr by (simp add: count_size_set_repr)
-qed
-
-lemma v_sum_type_rep: "\<v> = (\<Sum> g \<in> group_sizes . g * (groups_of_size g))"
-proof -
- have gs: "set_mset {# card G . G \<in># mset_set \<G> #} = group_sizes"
- unfolding group_sizes_def using groups_finite by auto
- have "\<v> = card (\<Union>(\<Union>{{ G \<in> \<G> . card G = g } | g . g \<in> group_sizes}))"
- using group_size_partition_covers_points by simp
- have v1: "\<v> = (\<Sum>x \<in># {# card G . G \<in># mset_set \<G> #}. x)"
- by (simp add: sum_unfold_sum_mset v_equals_sum_group_sizes)
- then have "\<v> = (\<Sum>x \<in> set_mset {# card G . G \<in># mset_set \<G> #} . x * (count {# card G . G \<in># mset_set \<G> #} x))"
- using mset_set_size_card_count by (simp add: v1)
- thus ?thesis using gs groups_of_size_alt_def_count by auto
-qed
-
-end
-
-subsubsection \<open>Uniform Group designs\<close>
-text \<open>A group design requiring all groups are the same size\<close>
-locale uniform_group_design = group_design +
- fixes u_group_size :: nat ("\<m>")
- assumes uniform_groups: "G \<in> \<G> \<Longrightarrow> card G = \<m>"
-
-begin
-
-lemma m_positive: "\<m> \<ge> 1"
-proof -
- obtain G where "G \<in> \<G>" using groups_size elem_exists_non_empty_set gr_implies_not_zero by blast
- thus ?thesis using uniform_groups min_group_size by fastforce
-qed
-
-lemma uniform_groups_alt: " \<forall> G \<in> \<G> . card G = \<m>"
- using uniform_groups by blast
-
-lemma uniform_groups_group_sizes: "group_sizes = {\<m>}"
- using design_points_nempty group_card_in_sizes group_size_implies_group_existance
- point_has_unique_group uniform_groups_alt by force
-
-lemma uniform_groups_group_size_singleton: "is_singleton (group_sizes)"
- using uniform_groups_group_sizes by auto
-
-lemma set_filter_eq_P_forall:"\<forall> x \<in> X . P x \<Longrightarrow> Set.filter P X = X"
- by (simp add: Collect_conj_eq Int_absorb2 Set.filter_def subsetI)
-
-lemma uniform_groups_groups_of_size_m: "groups_of_size \<m> = card \<G>"
-proof(simp add: groups_of_size_def)
- have "{G \<in> \<G>. card G = \<m>} = \<G>" using uniform_groups_alt set_filter_eq_P_forall by auto
- thus "card {G \<in> \<G>. card G = \<m>} = card \<G>" by simp
-qed
-
-lemma uniform_groups_of_size_not_m: "x \<noteq> \<m> \<Longrightarrow> groups_of_size x = 0"
- by (simp add: groups_of_size_def card_eq_0_iff uniform_groups)
-
-end
-
-subsection \<open>GDD\<close>
-text \<open>A GDD extends a group design with an additional index parameter.
-Each pair of elements must occur either \Lambda times if in diff groups, or 0 times if in the same
-group\<close>
-
-locale GDD = group_design +
- fixes index :: int ("\<Lambda>")
- assumes index_ge_1: "\<Lambda> \<ge> 1"
- assumes index_together: "G \<in> \<G> \<Longrightarrow> x \<in> G \<Longrightarrow> y \<in> G \<Longrightarrow> x \<noteq> y \<Longrightarrow> \<B> index {x, y} = 0"
- assumes index_distinct: "G1 \<in> \<G> \<Longrightarrow> G2 \<in> \<G> \<Longrightarrow> G1 \<noteq> G2 \<Longrightarrow> x \<in> G1 \<Longrightarrow> y \<in> G2 \<Longrightarrow>
- \<B> index {x, y} = \<Lambda>"
-begin
-
-lemma points_sep_groups_ne: "G1 \<in> \<G> \<Longrightarrow> G2 \<in> \<G> \<Longrightarrow> G1 \<noteq> G2 \<Longrightarrow> x \<in> G1 \<Longrightarrow> y \<in> G2 \<Longrightarrow> x \<noteq> y"
- by (meson point_in_one_group)
-
-lemma index_together_alt_ss: "ps \<subseteq> G \<Longrightarrow> G \<in> \<G> \<Longrightarrow> card ps = 2 \<Longrightarrow> \<B> index ps = 0"
- using index_together by (metis card_2_iff insert_subset)
-
-lemma index_distinct_alt_ss: "ps \<subseteq> \<V> \<Longrightarrow> card ps = 2 \<Longrightarrow> (\<And> G . G \<in> \<G> \<Longrightarrow> \<not> ps \<subseteq> G) \<Longrightarrow>
- \<B> index ps = \<Lambda>"
- using index_distinct by (metis card_2_iff empty_subsetI insert_subset point_has_unique_group)
-
-lemma gdd_index_options: "ps \<subseteq> \<V> \<Longrightarrow> card ps = 2 \<Longrightarrow> \<B> index ps = 0 \<or> \<B> index ps = \<Lambda>"
- using index_distinct_alt_ss index_together_alt_ss by blast
-
-lemma index_zero_implies_same_group: "ps \<subseteq> \<V> \<Longrightarrow> card ps = 2 \<Longrightarrow> \<B> index ps = 0 \<Longrightarrow>
- \<exists> G \<in> \<G> . ps \<subseteq> G" using index_distinct_alt_ss gr_implies_not_zero
- by (metis index_ge_1 less_one of_nat_0 of_nat_1 of_nat_le_0_iff)
-
-lemma index_zero_implies_same_group_unique: "ps \<subseteq> \<V> \<Longrightarrow> card ps = 2 \<Longrightarrow> \<B> index ps = 0 \<Longrightarrow>
- \<exists>! G \<in> \<G> . ps \<subseteq> G"
- by (meson GDD.index_zero_implies_same_group GDD_axioms card_2_iff' group_design.point_in_one_group
- group_design_axioms in_mono)
-
-lemma index_not_zero_impl_diff_group: "ps \<subseteq> \<V> \<Longrightarrow> card ps = 2 \<Longrightarrow> \<B> index ps = \<Lambda> \<Longrightarrow>
- (\<And> G . G \<in> \<G> \<Longrightarrow> \<not> ps \<subseteq> G)"
- using index_ge_1 index_together_alt_ss by auto
-
-lemma index_zero_implies_one_group:
- assumes "ps \<subseteq> \<V>"
- and "card ps = 2"
- and "\<B> index ps = 0"
- shows "size {#b \<in># mset_set \<G> . ps \<subseteq> b#} = 1"
-proof -
- obtain G where ging: "G \<in> \<G>" and psin: "ps \<subseteq> G"
- using index_zero_implies_same_group groups_size assms by blast
- then have unique: "\<And> G2 . G2 \<in> \<G> \<Longrightarrow> G \<noteq> G2 \<Longrightarrow> \<not> ps \<subseteq> G2"
- using index_zero_implies_same_group_unique by (metis assms)
- have "\<And> G'. G' \<in> \<G> \<longleftrightarrow> G' \<in># mset_set \<G>"
- by (simp add: groups_finite)
- then have eq_mset: "{#b \<in># mset_set \<G> . ps \<subseteq> b#} = mset_set {b \<in> \<G> . ps \<subseteq> b}"
- using filter_mset_mset_set groups_finite by blast
- then have "{b \<in> \<G> . ps \<subseteq> b} = {G}" using unique psin
- by (smt Collect_cong ging singleton_conv)
- thus ?thesis by (simp add: eq_mset)
-qed
-
-lemma index_distinct_group_num_alt_def: "ps \<subseteq> \<V> \<Longrightarrow> card ps = 2 \<Longrightarrow>
- size {#b \<in># mset_set \<G> . ps \<subseteq> b#} = 0 \<Longrightarrow> \<B> index ps = \<Lambda>"
- by (metis gdd_index_options index_zero_implies_one_group numeral_One zero_neq_numeral)
-
-lemma index_non_zero_implies_no_group:
- assumes "ps \<subseteq> \<V>"
- and "card ps = 2"
- and "\<B> index ps = \<Lambda>"
- shows "size {#b \<in># mset_set \<G> . ps \<subseteq> b#} = 0"
-proof -
- have "\<And> G . G \<in> \<G> \<Longrightarrow> \<not> ps \<subseteq> G" using index_not_zero_impl_diff_group assms by simp
- then have "{#b \<in># mset_set \<G> . ps \<subseteq> b#} = {#}"
- using filter_mset_empty_if_finite_and_filter_set_empty by force
- thus ?thesis by simp
-qed
-
-lemma gdd_index_non_zero_iff: "ps \<subseteq> \<V> \<Longrightarrow> card ps = 2 \<Longrightarrow>
- \<B> index ps = \<Lambda> \<longleftrightarrow> size {#b \<in># mset_set \<G> . ps \<subseteq> b#} = 0"
- using index_non_zero_implies_no_group index_distinct_group_num_alt_def by auto
-
-lemma gdd_index_zero_iff: "ps \<subseteq> \<V> \<Longrightarrow> card ps = 2 \<Longrightarrow>
- \<B> index ps = 0 \<longleftrightarrow> size {#b \<in># mset_set \<G> . ps \<subseteq> b#} = 1"
- apply (auto simp add: index_zero_implies_one_group)
- by (metis GDD.gdd_index_options GDD_axioms index_non_zero_implies_no_group old.nat.distinct(2))
-
-lemma points_index_upper_bound: "ps \<subseteq> \<V> \<Longrightarrow> card ps = 2 \<Longrightarrow> \<B> index ps \<le> \<Lambda>"
- using gdd_index_options index_ge_1
- by (metis int_one_le_iff_zero_less le_refl of_nat_0 of_nat_0_le_iff of_nat_le_iff zero_less_imp_eq_int)
-
-lemma index_1_imp_mult_1:
- assumes "\<Lambda> = 1"
- assumes "bl \<in># \<B>"
- assumes "card bl \<ge> 2"
- shows "multiplicity bl = 1"
-proof (rule ccontr)
- assume "\<not> (multiplicity bl = 1)"
- then have "multiplicity bl \<noteq> 1" and "multiplicity bl \<noteq> 0" using assms by simp_all
- then have m: "multiplicity bl \<ge> 2" by linarith
- obtain ps where ps: "ps \<subseteq> bl \<and> card ps = 2"
- using nat_int_comparison(3) obtain_subset_with_card_n by (metis assms(3))
- then have "\<B> index ps \<ge> 2"
- using m points_index_count_min ps by blast
- then show False using assms index_distinct ps antisym_conv2 not_numeral_less_zero
- numeral_le_one_iff points_index_ps_nin semiring_norm(69) zero_neq_numeral
- by (metis gdd_index_options int_int_eq int_ops(2))
-qed
-
-lemma simple_if_block_size_gt_2:
- assumes "\<And> bl . card bl \<ge> 2"
- assumes "\<Lambda> = 1"
- shows "simple_design \<V> \<B>"
- using index_1_imp_mult_1 assms apply (unfold_locales)
- by (metis card.empty not_numeral_le_zero)
-
-end
-
-subsubsection \<open>Sub types of GDD's\<close>
-
-text \<open>In literature, a GDD is usually defined in a number of different ways,
-including factors such as block size limitations\<close>
-locale K_\<Lambda>_GDD = K_block_design + GDD
-
-locale k_\<Lambda>_GDD = block_design + GDD
-
-sublocale k_\<Lambda>_GDD \<subseteq> K_\<Lambda>_GDD \<V> \<B> "{\<k>}" \<G> \<Lambda>
- by (unfold_locales)
-
-locale K_GDD = K_\<Lambda>_GDD \<V> \<B> \<K> \<G> 1
- for point_set ("\<V>") and block_collection ("\<B>") and sizes ("\<K>") and groups ("\<G>")
-
-locale k_GDD = k_\<Lambda>_GDD \<V> \<B> \<k> \<G> 1
- for point_set ("\<V>") and block_collection ("\<B>") and u_block_size ("\<k>") and groups ("\<G>")
-
-sublocale k_GDD \<subseteq> K_GDD \<V> \<B> "{\<k>}" \<G>
- by (unfold_locales)
-
-lemma (in K_GDD) multiplicity_1: "bl \<in># \<B> \<Longrightarrow> card bl \<ge> 2 \<Longrightarrow> multiplicity bl = 1"
- using index_1_imp_mult_1 by simp
-
-locale RGDD = GDD + resolvable_design
-
-subsection \<open>GDD and PBD Constructions\<close>
-text \<open>GDD's are commonly studied alongside PBD's (pairwise balanced designs). Many constructions
-have been developed for designs to create a GDD from a PBD and vice versa. In particular,
-Wilsons Construction is a well known construction, which is formalised in this section. It
-should be noted that many of the more basic constructions in this section are often stated without
-proof/all the necessary assumptions in textbooks/course notes.\<close>
-
-context GDD
-begin
-
-subsubsection \<open>GDD Delete Point construction\<close>
-lemma delete_point_index_zero:
- assumes "G \<in> {g - {x} |g. g \<in> \<G> \<and> g \<noteq> {x}}"
- and "y \<in> G" and "z \<in> G" and "z\<noteq> y"
-shows "(del_point_blocks x) index {y, z} = 0"
-proof -
- have "y \<noteq> x" using assms(1) assms(2) by blast
- have "z \<noteq> x" using assms(1) assms(3) by blast
- obtain G' where ing: "G' \<in> \<G>" and ss: "G \<subseteq> G'"
- using assms(1) by auto
- have "{y, z} \<subseteq> G" by (simp add: assms(2) assms(3))
- then have "{y, z} \<subseteq> \<V>"
- by (meson ss ing group_subset_point_subset)
- then have "{y, z} \<subseteq> (del_point x)"
- using \<open>y \<noteq> x\<close> \<open>z \<noteq> x\<close> del_point_def by fastforce
- thus ?thesis using delete_point_index_eq index_together
- by (metis assms(2) assms(3) assms(4) in_mono ing ss)
-qed
-
-lemma delete_point_index:
- assumes "G1 \<in> {g - {x} |g. g \<in> \<G> \<and> g \<noteq> {x}}"
- assumes "G2 \<in> {g - {x} |g. g \<in> \<G> \<and> g \<noteq> {x}}"
- assumes "G1 \<noteq> G2" and "y \<in> G1" and "z \<in> G2"
- shows "del_point_blocks x index {y, z} = \<Lambda>"
-proof -
- have "y \<noteq> x" using assms by blast
- have "z \<noteq> x" using assms by blast
- obtain G1' where ing1: "G1' \<in> \<G>" and t1: "G1 = G1' - {x}"
- using assms(1) by auto
- obtain G2' where ing2: "G2' \<in> \<G>" and t2: "G2 = G2' - {x}"
- using assms(2) by auto
- then have ss1: "G1 \<subseteq> G1'" and ss2: "G2 \<subseteq> G2'" using t1 by auto
- then have "{y, z} \<subseteq> \<V>" using ing1 ing2 ss1 ss2 assms(4) assms(5)
- by (metis empty_subsetI insert_absorb insert_subset point_in_group)
- then have "{y, z} \<subseteq> del_point x"
- using \<open>y \<noteq> x\<close> \<open>z \<noteq> x\<close> del_point_def by auto
- then have indx: "del_point_blocks x index {y, z} = \<B> index {y, z}"
- using delete_point_index_eq by auto
- have "G1' \<noteq> G2'" using assms t1 t2 by fastforce
- thus ?thesis using index_distinct
- using indx assms(4) assms(5) ing1 ing2 t1 t2 by auto
-qed
-
-lemma delete_point_group_size:
- assumes "{x} \<in> \<G> \<Longrightarrow> card \<G> > 2"
- shows "1 < card {g - {x} |g. g \<in> \<G> \<and> g \<noteq> {x}}"
-proof (cases "{x} \<in> \<G>")
- case True
- then have "\<And> g . g \<in> (\<G> - {{x}}) \<Longrightarrow> x \<notin> g"
- by (meson disjnt_insert1 groups_disjoint pairwise_alt)
- then have simpg: "\<And> g . g \<in> (\<G> - {{x}}) \<Longrightarrow> g - {x} = g"
- by simp
- have "{g - {x} |g. g \<in> \<G> \<and> g \<noteq> {x}} = {g - {x} |g. (g \<in> \<G> - {{x}})}" using True
- by force
- then have "{g - {x} |g. g \<in> \<G> \<and> g \<noteq> {x}} = {g |g. (g \<in> \<G> - {{x}})}" using simpg
- by (smt (verit) Collect_cong)
- then have eq: "{g - {x} |g. g \<in> \<G> \<and> g \<noteq> {x}} = \<G> - {{x}}" using set_self_img_compr by blast
- have "card (\<G> - {{x}}) = card \<G> - 1" using True
- by (simp add: groups_finite)
- then show ?thesis using True assms eq diff_is_0_eq' by force
-next
- case False
- then have "\<And>g' y. {x} \<notin> \<G> \<Longrightarrow> g' \<in> \<G> \<Longrightarrow> y \<in> \<G> \<Longrightarrow> g' - {x} = y - {x} \<Longrightarrow> g' = y"
- by (metis all_not_in_conv insert_Diff_single insert_absorb insert_iff points_sep_groups_ne)
- then have inj: "inj_on (\<lambda> g . g - {x}) \<G>" by (simp add: inj_onI False)
- have "{g - {x} |g. g \<in> \<G> \<and> g \<noteq> {x}} = {g - {x} |g. g \<in> \<G>}" using False by auto
- then have "card {g - {x} |g. g \<in> \<G> \<and> g \<noteq> {x}} = card \<G>" using inj groups_finite card_image
- by (auto simp add: card_image setcompr_eq_image)
- then show ?thesis using groups_size by presburger
-qed
-
-lemma GDD_by_deleting_point:
- assumes "\<And>bl. bl \<in># \<B> \<Longrightarrow> x \<in> bl \<Longrightarrow> 2 \<le> card bl"
- assumes "{x} \<in> \<G> \<Longrightarrow> card \<G> > 2"
- shows "GDD (del_point x) (del_point_blocks x) {g - {x} | g . g \<in> \<G> \<and> g \<noteq> {x}} \<Lambda>"
-proof -
- interpret pd: proper_design "del_point x" "del_point_blocks x"
- using delete_point_proper assms by blast
- show ?thesis using delete_point_index_zero delete_point_index assms delete_point_group_size
- by(unfold_locales) (simp_all add: partition_on_remove_pt group_partitions index_ge_1 del_point_def)
-qed
-
-end
-
-context K_GDD begin
-
-subsubsection \<open>PBD construction from GDD\<close>
-text \<open>Two well known PBD constructions involve taking a GDD and either combining the groups and
-blocks to form a new block collection, or by adjoining a point\<close>
-
-text \<open>First prove that combining the groups and block set results in a constant index\<close>
-lemma kgdd1_points_index_group_block:
- assumes "ps \<subseteq> \<V>"
- and "card ps = 2"
- shows "(\<B> + mset_set \<G>) index ps = 1"
-proof -
- have index1: "(\<And> G . G \<in> \<G> \<Longrightarrow> \<not> ps \<subseteq> G) \<Longrightarrow> \<B> index ps = 1"
- using index_distinct_alt_ss assms by fastforce
- have groups1: "\<B> index ps = 0 \<Longrightarrow> size {#b \<in># mset_set \<G> . ps \<subseteq> b#} = 1"
- using index_zero_implies_one_group assms by simp
- then have "(\<B> + mset_set \<G>) index ps = size (filter_mset ((\<subseteq>) ps) (\<B> + mset_set \<G>))"
- by (simp add: points_index_def)
- thus ?thesis using index1 groups1 gdd_index_non_zero_iff gdd_index_zero_iff assms
- gdd_index_options points_index_def filter_union_mset union_commute
- by (smt (z3) empty_neutral(1) less_irrefl_nat nonempty_has_size of_nat_1_eq_iff)
-qed
-
-text \<open>Combining blocks and the group set forms a PBD\<close>
-lemma combine_block_groups_pairwise: "pairwise_balance \<V> (\<B> + mset_set \<G>) 1"
-proof -
- let ?B = "\<B> + mset_set \<G>"
- have ss: "\<And> G. G \<in> \<G> \<Longrightarrow> G \<subseteq> \<V>"
- by (simp add: point_in_group subsetI)
- have "\<And> G. G \<in> \<G> \<Longrightarrow> G \<noteq> {}" using group_partitions
- using partition_onD3 by auto
- then interpret inc: design \<V> ?B
- proof (unfold_locales)
- show "\<And>b. (\<And>G. G \<in> \<G> \<Longrightarrow> G \<noteq> {}) \<Longrightarrow> b \<in># \<B> + mset_set \<G> \<Longrightarrow> b \<subseteq> \<V>"
- by (metis finite_set_mset_mset_set groups_finite ss union_iff wellformed)
- show "(\<And>G. G \<in> \<G> \<Longrightarrow> G \<noteq> {}) \<Longrightarrow> finite \<V>" by (simp add: finite_sets)
- show "\<And>bl. (\<And>G. G \<in> \<G> \<Longrightarrow> G \<noteq> {}) \<Longrightarrow> bl \<in># \<B> + mset_set \<G> \<Longrightarrow> bl \<noteq> {}"
- using blocks_nempty groups_finite by auto
- qed
- show ?thesis proof (unfold_locales)
- show "inc.\<b> \<noteq> 0" using b_positive by auto
- show "(1 ::int) \<le> 2" by simp
- show "2 \<le> inc.\<v>" by (simp add: gdd_min_v)
- then show "\<And>ps. ps \<subseteq> \<V> \<Longrightarrow> int (card ps) = 2 \<Longrightarrow> int ((\<B> + mset_set \<G>) index ps) = 1"
- using kgdd1_points_index_group_block by simp
- qed
-qed
-
-lemma combine_block_groups_PBD:
- assumes "\<And> G. G \<in> \<G> \<Longrightarrow> card G \<in> \<K>"
- assumes "\<And> k . k \<in> \<K> \<Longrightarrow> k \<ge> 2"
- shows "PBD \<V> (\<B> + mset_set \<G>) \<K>"
-proof -
- let ?B = "\<B> + mset_set \<G>"
- interpret inc: pairwise_balance \<V> ?B 1 using combine_block_groups_pairwise by simp
- show ?thesis using assms block_sizes groups_finite positive_ints
- by (unfold_locales) auto
-qed
-
-text \<open>Prove adjoining a point to each group set results in a constant points index\<close>
-lemma kgdd1_index_adjoin_group_block:
- assumes "x \<notin> \<V>"
- assumes "ps \<subseteq> insert x \<V>"
- assumes "card ps = 2"
- shows "(\<B> + mset_set {insert x g |g. g \<in> \<G>}) index ps = 1"
-proof -
- have "inj_on ((insert) x) \<G>"
- by (meson assms(1) inj_onI insert_ident point_in_group)
- then have eq: "mset_set {insert x g |g. g \<in> \<G>} = {# insert x g . g \<in># mset_set \<G>#}"
- by (simp add: image_mset_mset_set setcompr_eq_image)
- thus ?thesis
- proof (cases "x \<in> ps")
- case True
- then obtain y where y_ps: "ps = {x, y}" using assms(3)
- by (metis card_2_iff doubleton_eq_iff insertE singletonD)
- then have ynex: "y \<noteq> x" using assms by fastforce
- have yinv: "y \<in> \<V>"
- using assms(2) y_ps ynex by auto
- have all_g: "\<And> g. g \<in># (mset_set {insert x g |g. g \<in> \<G>}) \<Longrightarrow> x \<in> g"
- using eq by force
- have iff: "\<And> g . g \<in> \<G> \<Longrightarrow> y \<in> (insert x g) \<longleftrightarrow> y \<in> g" using ynex by simp
- have b: "\<B> index ps = 0"
- using True assms(1) points_index_ps_nin by fastforce
- then have "(\<B> + mset_set {insert x g |g. g \<in> \<G>}) index ps =
- (mset_set {insert x g |g. g \<in> \<G>}) index ps"
- using eq by (simp add: point_index_distrib)
- also have "... = (mset_set {insert x g |g. g \<in> \<G>}) rep y" using points_index_pair_rep_num
- by (metis (no_types, lifting) all_g y_ps)
- also have 0: "... = card {b \<in> {insert x g |g. g \<in> \<G>} . y \<in> b}"
- by (simp add: groups_finite rep_number_on_set_def)
- also have 1: "... = card {insert x g |g. g \<in> \<G> \<and> y \<in> insert x g}"
- by (smt (verit) Collect_cong mem_Collect_eq)
- also have 2: " ... = card {insert x g |g. g \<in> \<G> \<and> y \<in> g}"
- using iff by metis
- also have "... = card {g \<in> \<G> . y \<in> g}" using 1 2 0 empty_iff eq groups_finite ynex insert_iff
- by (metis points_index_block_image_add_eq points_index_single_rep_num rep_number_on_set_def)
- finally have "(\<B> + mset_set {insert x g |g. g \<in> \<G>}) index ps = 1"
- using rep_number_point_group_one yinv by simp
- then show ?thesis
- by simp
- next
- case False
- then have v: "ps \<subseteq> \<V>" using assms(2) by auto
- then have "(\<B> + mset_set {insert x g |g. g \<in> \<G>}) index ps = (\<B> + mset_set \<G>) index ps"
- using eq by (simp add: points_index_block_image_add_eq False point_index_distrib)
- then show ?thesis using v assms kgdd1_points_index_group_block by simp
- qed
-qed
-
-lemma pairwise_by_adjoining_point:
- assumes "x \<notin> \<V>"
- shows "pairwise_balance (add_point x) (\<B> + mset_set { insert x g | g. g \<in> \<G>}) 1"
-proof -
- let ?B = "\<B> + mset_set { insert x g | g. g \<in> \<G>}"
- let ?V = "add_point x"
- have vdef: "?V = \<V> \<union> {x}" using add_point_def by simp
- show ?thesis unfolding add_point_def using finite_sets b_positive
- proof (unfold_locales, simp_all)
- have "\<And> G. G \<in> \<G> \<Longrightarrow> insert x G \<subseteq> ?V"
- by (simp add: point_in_group subsetI vdef)
- then have "\<And> G. G \<in># (mset_set { insert x g | g. g \<in> \<G>}) \<Longrightarrow> G \<subseteq> ?V"
- by (smt (verit, del_insts) elem_mset_set empty_iff infinite_set_mset_mset_set mem_Collect_eq)
- then show "\<And>b. b \<in># \<B> \<or> b \<in># mset_set {insert x g |g. g \<in> \<G>} \<Longrightarrow> b \<subseteq> insert x \<V>"
- using wellformed add_point_def by fastforce
- next
- have "\<And> G. G \<in> \<G> \<Longrightarrow> insert x G \<noteq> {}" using group_partitions
- using partition_onD3 by auto
- then have gnempty: "\<And> G. G \<in># (mset_set { insert x g | g. g \<in> \<G>}) \<Longrightarrow> G \<noteq> {}"
- by (smt (verit, del_insts) elem_mset_set empty_iff infinite_set_mset_mset_set mem_Collect_eq)
- then show "\<And>bl. bl \<in># \<B> \<or> bl \<in># mset_set {insert x g |g. g \<in> \<G>} \<Longrightarrow> bl \<noteq> {}"
- using blocks_nempty by auto
- next
- have "card \<V> \<ge> 2" using gdd_min_v by simp
- then have "card (insert x \<V>) \<ge> 2"
- by (meson card_insert_le dual_order.trans finite_sets)
- then show "2 \<le> int (card (insert x \<V>))" by auto
- next
- show "\<And>ps. ps \<subseteq> insert x \<V> \<Longrightarrow>
- card ps = 2 \<Longrightarrow> (\<B> + mset_set {insert x g |g. g \<in> \<G>}) index ps = Suc 0"
- using kgdd1_index_adjoin_group_block by (simp add: assms)
- qed
-qed
-
-lemma PBD_by_adjoining_point:
- assumes "x \<notin> \<V>"
- assumes "\<And> k . k \<in> \<K> \<Longrightarrow> k \<ge> 2"
- shows "PBD (add_point x) (\<B> + mset_set { insert x g | g. g \<in> \<G>}) (\<K> \<union> {(card g) + 1 | g . g \<in> \<G>})"
-proof -
- let ?B = "\<B> + mset_set { insert x g | g. g \<in> \<G>}"
- let ?V = "(add_point x)"
- interpret inc: pairwise_balance ?V ?B 1 using pairwise_by_adjoining_point assms by auto
- show ?thesis using block_sizes positive_ints proof (unfold_locales)
- have xg: "\<And> g. g \<in> \<G> \<Longrightarrow> x \<notin> g"
- using assms point_in_group by auto
- have "\<And> bl . bl \<in># \<B> \<Longrightarrow> card bl \<in> \<K>" by (simp add: block_sizes)
- have "\<And> bl . bl \<in># mset_set {insert x g |g. g \<in> \<G>} \<Longrightarrow> bl \<in> {insert x g | g . g \<in> \<G>}"
- by (simp add: groups_finite)
- then have "\<And> bl . bl \<in># mset_set {insert x g |g. g \<in> \<G>} \<Longrightarrow>
- card bl \<in> {int (card g + 1) |g. g \<in> \<G>}"
- proof -
- fix bl
- assume "bl \<in># mset_set {insert x g |g. g \<in> \<G>}"
- then have "bl \<in> {insert x g | g . g \<in> \<G>}" by (simp add: groups_finite)
- then obtain g where gin: "g \<in> \<G>" and i: "bl = insert x g" by auto
- thus "card bl \<in> {int (card g + 1) |g. g \<in> \<G>}"
- using gin group_elements_finite i xg by auto
- qed
- then show "\<And>bl. bl \<in># \<B> + mset_set {insert x g |g. g \<in> \<G>} \<Longrightarrow>
- int (card bl) \<in> \<K> \<union> {int (card g + 1) |g. g \<in> \<G>}"
- using UnI1 UnI2 block_sizes union_iff by (smt (z3) mem_Collect_eq)
- show "\<And>x. x \<in> \<K> \<union> {int (card g + 1) |g. g \<in> \<G>} \<Longrightarrow> 0 < x"
- using min_group_size positive_ints by auto
- show "\<And>k. k \<in> \<K> \<union> {int (card g + 1) |g. g \<in> \<G>} \<Longrightarrow> 2 \<le> k"
- using min_group_size positive_ints assms by fastforce
- qed
-qed
-
-subsubsection \<open>Wilson's Construction\<close>
-text \<open>Wilson's construction involves the combination of multiple k-GDD's. This proof was
-based of Stinson \cite{stinsonCombinatorialDesignsConstructions2004}\<close>
-
-lemma wilsons_construction_proper:
- assumes "card I = w"
- assumes "w > 0"
- assumes "\<And> n. n \<in> \<K>' \<Longrightarrow> n \<ge> 2"
- assumes "\<And> B . B \<in># \<B> \<Longrightarrow> K_GDD (B \<times> I) (f B) \<K>' {{x} \<times> I |x . x \<in> B }"
- shows "proper_design (\<V> \<times> I) (\<Sum>B \<in># \<B>. (f B))" (is "proper_design ?Y ?B")
-proof (unfold_locales, simp_all)
- show "\<And>b. \<exists>x\<in>#\<B>. b \<in># f x \<Longrightarrow> b \<subseteq> \<V> \<times> I"
- proof -
- fix b
- assume "\<exists>x\<in>#\<B>. b \<in># f x"
- then obtain B where "B \<in># \<B>" and "b \<in># (f B)" by auto
- then interpret kgdd: K_GDD "(B \<times> I)" "(f B)" \<K>' "{{x} \<times> I |x . x \<in> B }" using assms by auto
- show "b \<subseteq> \<V> \<times> I" using kgdd.wellformed
- using \<open>B \<in># \<B>\<close> \<open>b \<in># f B\<close> wellformed by fastforce
- qed
- show "finite (\<V> \<times> I)" using finite_sets assms bot_nat_0.not_eq_extremum card.infinite by blast
- show "\<And>bl. \<exists>x\<in>#\<B>. bl \<in># f x \<Longrightarrow> bl \<noteq> {}"
- proof -
- fix bl
- assume "\<exists>x\<in>#\<B>. bl \<in># f x"
- then obtain B where "B \<in># \<B>" and "bl \<in># (f B)" by auto
- then interpret kgdd: K_GDD "(B \<times> I)" "(f B)" \<K>' "{{x} \<times> I |x . x \<in> B }" using assms by auto
- show "bl \<noteq> {}" using kgdd.blocks_nempty by (simp add: \<open>bl \<in># f B\<close>)
- qed
- show "\<exists>i\<in>#\<B>. f i \<noteq> {#}"
- proof -
- obtain B where "B \<in># \<B>"
- using design_blocks_nempty by auto
- then interpret kgdd: K_GDD "(B \<times> I)" "(f B)" \<K>' "{{x} \<times> I |x . x \<in> B }" using assms by auto
- have "f B \<noteq> {#}" using kgdd.design_blocks_nempty by simp
- then show "\<exists>i\<in>#\<B>. f i \<noteq> {#}" using \<open>B \<in># \<B>\<close> by auto
- qed
-qed
-
-lemma pair_construction_block_sizes:
- assumes "K_GDD (B \<times> I) (f B) \<K>' {{x} \<times> I |x . x \<in> B }"
- assumes "B \<in># \<B>"
- assumes "b \<in># (f B)"
- shows "card b \<in> \<K>'"
-proof -
- interpret bkgdd: K_GDD "(B \<times> I)" "(f B)" \<K>' "{{x} \<times> I |x . x \<in> B }"
- using assms by simp
- show "card b \<in> \<K>'" using bkgdd.block_sizes by (simp add:assms)
-qed
-
-lemma wilsons_construction_index_0:
- assumes "\<And> B . B \<in># \<B> \<Longrightarrow> K_GDD (B \<times> I) (f B) \<K>' {{x} \<times> I |x . x \<in> B }"
- assumes "G \<in> {GG \<times> I |GG. GG \<in> \<G>}"
- assumes "X \<in> G"
- assumes "Y \<in> G"
- assumes "X \<noteq> Y"
- shows "(\<Sum>\<^sub># (image_mset f \<B>)) index {X, Y} = 0"
-proof -
- obtain G' where gi: "G = G' \<times> I" and ging: "G' \<in> \<G>" using assms by auto
- obtain x y ix iy where xpair: "X = (x, ix)" and ypair: "Y = (y, iy)" using assms by auto
- then have ixin: "ix \<in> I" and xing: "x \<in> G'" using assms gi by auto
- have iyin: "iy \<in> I" and ying: "y \<in> G'" using assms ypair gi by auto
- have ne_index_0: "x \<noteq> y \<Longrightarrow> \<B> index {x, y} = 0"
- using ying xing index_together ging by simp
- have "\<And> B. B \<in># \<B> \<Longrightarrow> (f B) index {(x, ix), (y, iy)} = 0"
- proof -
- fix B
- assume assm: "B \<in># \<B>"
- then interpret kgdd: K_GDD "(B \<times> I)" "(f B)" \<K>' "{{x} \<times> I |x . x \<in> B }" using assms by simp
- have not_ss_0: "\<not> ({(x, ix), (y, iy)} \<subseteq> (B \<times> I)) \<Longrightarrow> (f B) index {(x, ix), (y, iy)} = 0"
- by (metis kgdd.points_index_ps_nin)
- have "x \<noteq> y \<Longrightarrow> \<not> {x, y} \<subseteq> B" using ne_index_0 assm points_index_0_left_imp by auto
- then have "x \<noteq> y \<Longrightarrow> \<not> ({(x, ix), (y, iy)} \<subseteq> (B \<times> I))" using assms
- by (meson empty_subsetI insert_subset mem_Sigma_iff)
- then have nexy: "x \<noteq> y \<Longrightarrow> (f B) index {(x, ix), (y, iy)} = 0" using not_ss_0 by simp
- have "x = y \<Longrightarrow> ({(x, ix), (y, iy)} \<subseteq> (B \<times> I)) \<Longrightarrow> (f B) index {(x, ix), (y, iy)} = 0"
- proof -
- assume eq: "x = y"
- assume "({(x, ix), (y, iy)} \<subseteq> (B \<times> I))"
- then obtain g where "g \<in> {{x} \<times> I |x . x \<in> B }" and "(x, ix) \<in> g" and "(y, ix) \<in> g"
- using eq by auto
- then show ?thesis using kgdd.index_together
- by (smt (verit, best) SigmaD1 SigmaD2 SigmaI assms(4) assms(5) gi mem_Collect_eq xpair ypair)
- qed
- then show "(f B) index {(x, ix), (y, iy)} = 0" using not_ss_0 nexy by auto
- qed
- then have "\<And> B. B \<in># (image_mset f \<B>) \<Longrightarrow> B index {(x, ix), (y, iy)} = 0" by auto
- then show "(\<Sum>\<^sub># (image_mset f \<B>)) index {X, Y} = 0"
- by (simp add: points_index_sum xpair ypair)
-qed
-
-lemma wilsons_construction_index_1:
- assumes "\<And> B . B \<in># \<B> \<Longrightarrow> K_GDD (B \<times> I) (f B) \<K>' {{x} \<times> I |x . x \<in> B }"
- assumes "G1 \<in> {G \<times> I |G. G \<in> \<G>}"
- assumes "G2 \<in> {G \<times> I |G. G \<in> \<G>}"
- assumes "G1 \<noteq> G2"
- and "(x, ix) \<in> G1" and "(y, iy) \<in> G2"
- shows "(\<Sum>\<^sub># (image_mset f \<B>)) index {(x, ix), (y, iy)} = (1 ::int)"
-proof -
- obtain G1' where gi1: "G1 = G1' \<times> I" and ging1: "G1' \<in> \<G>" using assms by auto
- obtain G2' where gi2: "G2 = G2' \<times> I" and ging2: "G2' \<in> \<G>" using assms by auto
- have xing: "x \<in> G1'" using assms gi1 by simp
- have ying: "y \<in> G2'" using assms gi2 by simp
- have gne: "G1' \<noteq> G2'" using assms gi1 gi2 by auto
- then have xyne: "x \<noteq> y" using xing ying ging1 ging2 point_in_one_group by blast
- have "\<exists>! bl . bl \<in># \<B> \<and> {x, y} \<subseteq> bl" using index_distinct points_index_one_unique_block
- by (metis ging1 ging2 gne of_nat_1_eq_iff xing ying)
- then obtain bl where blinb:"bl \<in># \<B>" and xyblss: "{x, y} \<subseteq> bl" by auto
- then have "\<And> b . b \<in># \<B> - {#bl#} \<Longrightarrow> \<not> {x, y} \<subseteq> b" using points_index_one_not_unique_block
- by (metis ging1 ging2 gne index_distinct int_ops(2) nat_int_comparison(1) xing ying)
- then have not_ss: "\<And> b . b \<in># \<B> - {#bl#} \<Longrightarrow> \<not> ({(x, ix), (y, iy)} \<subseteq> (b \<times> I))" using assms
- by (meson SigmaD1 empty_subsetI insert_subset)
- then have pi0: "\<And> b . b \<in># \<B> - {#bl#} \<Longrightarrow> (f b) index {(x, ix), (y, iy)} = 0"
- proof -
- fix b
- assume assm: "b \<in># \<B> - {#bl#}"
- then have "b \<in># \<B>" by (meson in_diffD)
- then interpret kgdd: K_GDD "(b \<times> I)" "(f b)" \<K>' "{{x} \<times> I |x . x \<in> b }" using assms by simp
- show "(f b) index {(x, ix), (y, iy)} = 0"
- using assm not_ss by (metis kgdd.points_index_ps_nin)
- qed
- let ?G = "{{x} \<times> I |x . x \<in> bl }"
- interpret bkgdd: K_GDD "(bl \<times> I)" "(f bl)" \<K>' ?G using assms blinb by simp
- obtain g1 g2 where xing1: "(x, ix) \<in> g1" and ying2: "(y, iy) \<in> g2" and g1g: "g1 \<in> ?G"
- and g2g: "g2 \<in> ?G" using assms(5) assms(6) gi1 gi2
- by (metis (no_types, lifting) bkgdd.point_has_unique_group insert_subset mem_Sigma_iff xyblss)
- then have "g1 \<noteq> g2" using xyne by blast
- then have pi1: "(f bl) index {(x, ix), (y, iy)} = 1"
- using bkgdd.index_distinct xing1 ying2 g1g g2g by simp
- have "(\<Sum>\<^sub># (image_mset f \<B>)) index {(x, ix), (y, iy)} =
- (\<Sum>B \<in># \<B>. (f B) index {(x, ix), (y, iy)} )"
- by (simp add: points_index_sum)
- then have "(\<Sum>\<^sub># (image_mset f \<B>)) index {(x, ix), (y, iy)} =
- (\<Sum>B \<in># (\<B> - {#bl#}). (f B) index {(x, ix), (y, iy)}) + (f bl) index {(x, ix), (y, iy)}"
- by (metis (no_types, lifting) add.commute blinb insert_DiffM sum_mset.insert)
- thus ?thesis using pi0 pi1 by simp
-qed
-
-theorem Wilsons_Construction:
- assumes "card I = w"
- assumes "w > 0"
- assumes "\<And> n. n \<in> \<K>' \<Longrightarrow> n \<ge> 2"
- assumes "\<And> B . B \<in># \<B> \<Longrightarrow> K_GDD (B \<times> I) (f B) \<K>' {{x} \<times> I |x . x \<in> B }"
- shows "K_GDD (\<V> \<times> I) (\<Sum>B \<in># \<B>. (f B)) \<K>' {G \<times> I | G . G \<in> \<G>}"
-proof -
- let ?Y = "\<V> \<times> I" and ?H = "{G \<times> I | G . G \<in> \<G>}" and ?B = "\<Sum>B \<in># \<B>. (f B)"
- interpret pd: proper_design ?Y ?B using wilsons_construction_proper assms by auto
- have "\<And> bl . bl \<in># (\<Sum>B \<in># \<B>. (f B)) \<Longrightarrow> card bl \<in> \<K>'"
- using assms pair_construction_block_sizes by blast
- then interpret kdes: K_block_design ?Y ?B \<K>'
- using assms(3) by (unfold_locales) (simp_all,fastforce)
- interpret gdd: GDD ?Y ?B ?H "1:: int"
- proof (unfold_locales)
- show "partition_on (\<V> \<times> I) {G \<times> I |G. G \<in> \<G>}"
- using assms groups_not_empty design_points_nempty group_partitions
- by (simp add: partition_on_cart_prod)
- have "inj_on (\<lambda> G. G \<times> I) \<G>"
- using inj_on_def pd.design_points_nempty by auto
- then have "card {G \<times> I |G. G \<in> \<G>} = card \<G>" using card_image by (simp add: Setcompr_eq_image)
- then show "1 < card {G \<times> I |G. G \<in> \<G>}" using groups_size by linarith
- show "(1::int) \<le> 1" by simp
- have gdd_fact: "\<And> B . B \<in># \<B> \<Longrightarrow> K_GDD (B \<times> I) (f B) \<K>' {{x} \<times> I |x . x \<in> B }"
- using assms by simp
- show "\<And>G X Y. G \<in> {GG \<times> I |GG. GG \<in> \<G>} \<Longrightarrow> X \<in> G \<Longrightarrow> Y \<in> G \<Longrightarrow> X \<noteq> Y
- \<Longrightarrow> (\<Sum>\<^sub># (image_mset f \<B>)) index {X, Y} = 0"
- using wilsons_construction_index_0[OF assms(4)] by auto
- show "\<And>G1 G2 X Y. G1 \<in> {G \<times> I |G. G \<in> \<G>} \<Longrightarrow> G2 \<in> {G \<times> I |G. G \<in> \<G>}
- \<Longrightarrow> G1 \<noteq> G2 \<Longrightarrow> X \<in> G1 \<Longrightarrow> Y \<in> G2 \<Longrightarrow> ((\<Sum>\<^sub># (image_mset f \<B>)) index {X, Y}) = (1 ::int)"
- using wilsons_construction_index_1[OF assms(4)] by blast
- qed
- show ?thesis by (unfold_locales)
-qed
-
-end
-
-context pairwise_balance
-begin
-
-lemma PBD_by_deleting_point:
- assumes "\<v> > 2"
- assumes "\<And> bl . bl \<in># \<B> \<Longrightarrow> card bl \<ge> 2"
- shows "pairwise_balance (del_point x) (del_point_blocks x) \<Lambda>"
-proof (cases "x \<in> \<V>")
- case True
- interpret des: design "del_point x" "del_point_blocks x"
- using delete_point_design assms by blast
- show ?thesis using assms design_blocks_nempty del_point_def del_point_blocks_def
- proof (unfold_locales, simp_all)
- show "2 < \<v> \<Longrightarrow> (\<And>bl. bl \<in># \<B> \<Longrightarrow> 2 \<le> card bl) \<Longrightarrow> 2 \<le> int (card (\<V> - {x}))"
- using card_Diff_singleton_if diff_diff_cancel diff_le_mono2 finite_sets less_one
- by (metis int_ops(3) less_nat_zero_code nat_le_linear verit_comp_simplify1(3) zle_int)
- have "\<And> ps . ps \<subseteq> \<V> - {x} \<Longrightarrow> ps \<subseteq> \<V>" by auto
- then show "\<And>ps. ps \<subseteq> \<V> - {x} \<Longrightarrow> card ps = 2 \<Longrightarrow> {#bl - {x}. bl \<in># \<B>#} index ps = \<Lambda>"
- using delete_point_index_eq del_point_def del_point_blocks_def by simp
- qed
-next
- case False
- then show ?thesis
- by (simp add: del_invalid_point del_invalid_point_blocks pairwise_balance_axioms)
-qed
-end
-
-context k_GDD
-begin
-
-lemma bibd_from_kGDD:
- assumes "\<k> > 1"
- assumes "\<And> g. g \<in> \<G> \<Longrightarrow> card g = \<k> - 1"
- assumes " x \<notin> \<V>"
- shows "bibd (add_point x) (\<B> + mset_set { insert x g | g. g \<in> \<G>}) (\<k>) 1"
-proof -
- have "({\<k>} \<union> {(card g) + 1 | g . g \<in> \<G>}) = {\<k>}"
- using assms(2) by fastforce
- then interpret pbd: PBD "(add_point x)" "\<B> + mset_set { insert x g | g. g \<in> \<G>}" "{\<k>}"
- using PBD_by_adjoining_point assms sys_block_sizes_obtain_bl add_point_def
- by (smt (verit, best) Collect_cong sys_block_sizes_uniform uniform_alt_def_all)
- show ?thesis using assms pbd.block_sizes block_size_lt_v finite_sets add_point_def
- by (unfold_locales) (simp_all)
-qed
-
-end
-
-context PBD
-begin
-
-lemma pbd_points_index1: "ps \<subseteq> \<V> \<Longrightarrow> card ps = 2 \<Longrightarrow> \<B> index ps = 1"
- using balanced by (metis int_eq_iff_numeral of_nat_1_eq_iff)
-
-lemma pbd_index1_points_imply_unique_block:
- assumes "b1 \<in># \<B>" and "b2 \<in># \<B>" and "b1 \<noteq> b2"
- assumes "x \<noteq> y" and "{x, y} \<subseteq> b1" and "x \<in> b2"
- shows "y \<notin> b2"
-proof (rule ccontr)
- let ?ps = "{# b \<in># \<B> . {x, y} \<subseteq> b#}"
- assume "\<not> y \<notin> b2"
- then have a: "y \<in> b2" by linarith
- then have "{x, y} \<subseteq> b2"
- by (simp add: assms(6))
- then have "b1 \<in># ?ps" and "b2 \<in># ?ps" using assms by auto
- then have ss: "{#b1, b2#} \<subseteq># ?ps" using assms
- by (metis insert_noteq_member mset_add mset_subset_eq_add_mset_cancel single_subset_iff)
- have "size {#b1, b2#} = 2" using assms by auto
- then have ge2: "size ?ps \<ge> 2" using assms ss by (metis size_mset_mono)
- have pair: "card {x, y} = 2" using assms by auto
- have "{x, y} \<subseteq> \<V>" using assms wellformed by auto
- then have "\<B> index {x, y} = 1" using pbd_points_index1 pair by simp
- then show False using points_index_def ge2
- by (metis numeral_le_one_iff semiring_norm(69))
-qed
-
-lemma strong_delete_point_groups_index_zero:
- assumes "G \<in> {b - {x} |b. b \<in># \<B> \<and> x \<in> b}"
- assumes "xa \<in> G" and "y \<in> G" and "xa \<noteq> y"
- shows "(str_del_point_blocks x) index {xa, y} = 0"
-proof (auto simp add: points_index_0_iff str_del_point_blocks_def)
- fix b
- assume a1: "b \<in># \<B>" and a2: "x \<notin> b" and a3: "xa \<in> b" and a4: "y \<in> b"
- obtain b' where "G = b' - {x}" and "b' \<in># \<B>" and "x \<in> b'" using assms by blast
- then show False using a1 a2 a3 a4 assms pbd_index1_points_imply_unique_block
- by fastforce
-qed
-
-lemma strong_delete_point_groups_index_one:
- assumes "G1 \<in> {b - {x} |b. b \<in># \<B> \<and> x \<in> b}"
- assumes "G2 \<in> {b - {x} |b. b \<in># \<B> \<and> x \<in> b}"
- assumes "G1 \<noteq> G2" and "xa \<in> G1" and "y \<in> G2"
- shows "(str_del_point_blocks x) index {xa, y} = 1"
-proof -
- obtain b1 where gb1: "G1 = b1 - {x}" and b1in: "b1 \<in># \<B>" and xin1: "x \<in> b1" using assms by blast
- obtain b2 where gb2: "G2 = b2 - {x}" and b2in: "b2 \<in># \<B>" and xin2:"x \<in> b2" using assms by blast
- have bneq: "b1 \<noteq> b2 " using assms(3) gb1 gb2 by auto
- have "xa \<noteq> y" using gb1 b1in xin1 gb2 b2in xin2 assms(3) assms(4) assms(5) insert_subset
- by (smt (verit, best) Diff_eq_empty_iff Diff_iff empty_Diff insertCI pbd_index1_points_imply_unique_block)
- then have pair: "card {xa, y} = 2" by simp
- have inv: "{xa, y} \<subseteq> \<V>" using gb1 b1in gb2 b2in assms(4) assms(5)
- by (metis Diff_cancel Diff_subset insert_Diff insert_subset wellformed)
- have "{# bl \<in># \<B> . x \<in> bl#} index {xa, y} = 0"
- proof (auto simp add: points_index_0_iff)
- fix b assume a1: "b \<in># \<B>" and a2: "x \<in> b" and a3: "xa \<in> b" and a4: "y \<in> b"
- then have yxss: "{y, x} \<subseteq> b2"
- using assms(5) gb2 xin2 by blast
- have "{xa, x} \<subseteq> b1"
- using assms(4) gb1 xin1 by auto
- then have "xa \<notin> b2" using pbd_index1_points_imply_unique_block
- by (metis DiffE assms(4) b1in b2in bneq gb1 singletonI xin2)
- then have "b2 \<noteq> b" using a3 by auto
- then show False using pbd_index1_points_imply_unique_block
- by (metis DiffD2 yxss a1 a2 a4 assms(5) b2in gb2 insertI1)
- qed
- then have "(str_del_point_blocks x) index {xa, y} = \<B> index {xa, y}"
- by (metis multiset_partition plus_nat.add_0 point_index_distrib str_del_point_blocks_def)
- thus ?thesis using pbd_points_index1 pair inv by fastforce
-qed
-
-lemma blocks_with_x_partition:
- assumes "x \<in> \<V>"
- shows "partition_on (\<V> - {x}) {b - {x} |b. b \<in># \<B> \<and> x \<in> b}"
-proof (intro partition_onI )
- have gtt: "\<And> bl. bl \<in># \<B> \<Longrightarrow> card bl \<ge> 2" using block_size_gt_t
- by (simp add: block_sizes nat_int_comparison(3))
- show "\<And>p. p \<in> {b - {x} |b. b \<in># \<B> \<and> x \<in> b} \<Longrightarrow> p \<noteq> {}"
- proof -
- fix p assume "p \<in> {b - {x} |b. b \<in># \<B> \<and> x \<in> b}"
- then obtain b where ptx: "p = b - {x}" and "b \<in># \<B>" and xinb: "x \<in> b" by blast
- then have ge2: "card b \<ge> 2" using gtt by (simp add: nat_int_comparison(3))
- then have "finite b" by (metis card.infinite not_numeral_le_zero)
- then have "card p = card b - 1" using xinb ptx by simp
- then have "card p \<ge> 1" using ge2 by linarith
- thus "p \<noteq> {}" by auto
- qed
- show "\<Union> {b - {x} |b. b \<in># \<B> \<and> x \<in> b} = \<V> - {x}"
- proof (intro subset_antisym subsetI)
- fix xa
- assume "xa \<in> \<Union> {b - {x} |b. b \<in># \<B> \<and> x \<in> b}"
- then obtain b where "xa \<in> b" and "b \<in># \<B>" and "x \<in> b" and "xa \<noteq> x" by auto
- then show "xa \<in> \<V> - {x}" using wf_invalid_point by blast
- next
- fix xa
- assume a: "xa \<in> \<V> - {x}"
- then have nex: "xa \<noteq> x" by simp
- then have pair: "card {xa, x} = 2" by simp
- have "{xa, x} \<subseteq> \<V>" using a assms by auto
- then have "card {b \<in> design_support . {xa, x} \<subseteq> b} = 1"
- using balanced points_index_simple_def pbd_points_index1 assms by (metis pair)
- then obtain b where des: "b \<in> design_support" and ss: "{xa, x} \<subseteq> b"
- by (metis (no_types, lifting) card_1_singletonE mem_Collect_eq singletonI)
- then show "xa \<in> \<Union> {b - {x} |b. b \<in># \<B> \<and> x \<in> b}"
- using des ss nex design_support_def by auto
- qed
- show "\<And>p p'. p \<in> {b - {x} |b. b \<in># \<B> \<and> x \<in> b} \<Longrightarrow> p' \<in> {b - {x} |b. b \<in># \<B> \<and> x \<in> b} \<Longrightarrow>
- p \<noteq> p' \<Longrightarrow> p \<inter> p' = {}"
- proof -
- fix p p'
- assume p1: "p \<in> {b - {x} |b. b \<in># \<B> \<and> x \<in> b}" and p2: "p' \<in> {b - {x} |b. b \<in># \<B> \<and> x \<in> b}"
- and pne: "p \<noteq> p'"
- then obtain b where b1: "p = b - {x}" and b1in:"b \<in># \<B>" and xinb1:"x \<in> b" by blast
- then obtain b' where b2: "p' = b' - {x}" and b2in: "b' \<in># \<B>" and xinb2: "x \<in> b'"
- using p2 by blast
- then have "b \<noteq> b'" using pne b1 by auto
- then have "\<And> y. y \<in> b \<Longrightarrow> y \<noteq> x \<Longrightarrow> y \<notin> b'"
- using b1in b2in xinb1 xinb2 pbd_index1_points_imply_unique_block
- by (meson empty_subsetI insert_subset)
- then have "\<And> y. y \<in> p \<Longrightarrow> y \<notin> p'"
- by (metis Diff_iff b1 b2 insertI1)
- then show "p \<inter> p' = {}" using disjoint_iff by auto
- qed
-qed
-
-lemma KGDD_by_deleting_point:
- assumes "x \<in> \<V>"
- assumes "\<B> rep x < \<b>"
- assumes "\<B> rep x > 1"
- shows "K_GDD (del_point x) (str_del_point_blocks x) \<K> { b - {x} | b . b \<in># \<B> \<and> x \<in> b}"
-proof -
- have "\<And> bl. bl \<in># \<B> \<Longrightarrow> card bl \<ge> 2" using block_size_gt_t
- by (simp add: block_sizes nat_int_comparison(3))
- then interpret des: proper_design "(del_point x)" "(str_del_point_blocks x)"
- using strong_delete_point_proper assms by blast
- show ?thesis using blocks_with_x_partition strong_delete_point_groups_index_zero
- strong_delete_point_groups_index_one str_del_point_blocks_def del_point_def
- proof (unfold_locales, simp_all add: block_sizes positive_ints assms)
- have ge1: "card {b . b \<in># \<B> \<and> x \<in> b} > 1"
- using assms(3) replication_num_simple_def design_support_def by auto
- have fin: "finite {b . b \<in># \<B> \<and> x \<in> b}" by simp
- have inj: "inj_on (\<lambda> b . b - {x}) {b . b \<in># \<B> \<and> x \<in> b}"
- using assms(2) inj_on_def mem_Collect_eq by auto
- then have "card {b - {x} |b. b \<in># \<B> \<and> x \<in> b} = card {b . b \<in># \<B> \<and> x \<in> b}"
- using card_image fin by (simp add: inj card_image setcompr_eq_image)
- then show "Suc 0 < card {b - {x} |b. b \<in># \<B> \<and> x \<in> b}" using ge1
- by presburger
- qed
-qed
-
-lemma card_singletons_eq: "card {{a} | a . a \<in> A} = card A"
- by (simp add: card_image Setcompr_eq_image)
-
-lemma KGDD_from_PBD: "K_GDD \<V> \<B> \<K> {{x} | x . x \<in> \<V>}"
-proof (unfold_locales,auto simp add: Setcompr_eq_image partition_on_singletons)
- have "card ((\<lambda>x. {x}) ` \<V>) \<ge> 2" using t_lt_order card_singletons_eq
- by (metis Collect_mem_eq nat_leq_as_int of_nat_numeral setcompr_eq_image)
- then show "Suc 0 < card ((\<lambda>x. {x}) ` \<V>)" by linarith
- show "\<And>xa xb. xa \<in> \<V> \<Longrightarrow> xb \<in> \<V> \<Longrightarrow> \<B> index {xa, xb} \<noteq> Suc 0 \<Longrightarrow> xa = xb"
- proof (rule ccontr)
- fix xa xb
- assume ain: "xa \<in> \<V>" and bin: "xb \<in> \<V>" and ne1: "\<B> index {xa, xb} \<noteq> Suc 0"
- assume "xa \<noteq> xb"
- then have "card {xa, xb} = 2" by auto
- then have "\<B> index {xa, xb} = 1"
- by (simp add: ain bin pbd_points_index1)
- thus False using ne1 by linarith
- qed
-qed
-
-end
-
-context bibd
-begin
-lemma kGDD_from_bibd:
- assumes "\<Lambda> = 1"
- assumes "x \<in> \<V>"
- shows "k_GDD (del_point x) (str_del_point_blocks x) \<k> { b - {x} | b . b \<in># \<B> \<and> x \<in> b}"
-proof -
- interpret pbd: PBD \<V> \<B> "{\<k>}" using assms
- using PBD.intro \<Lambda>_PBD_axioms by auto
- have lt: "\<B> rep x < \<b>" using block_num_gt_rep
- by (simp add: assms(2))
- have "\<B> rep x > 1" using r_ge_two assms by simp
- then interpret kgdd: K_GDD "(del_point x)" "str_del_point_blocks x"
- "{\<k>}" "{ b - {x} | b . b \<in># \<B> \<and> x \<in> b}"
- using pbd.KGDD_by_deleting_point lt assms by blast
- show ?thesis using del_point_def str_del_point_blocks_def by (unfold_locales) (simp_all)
-qed
-
-end
+(* Title: Group_Divisible_Designs.thy
+ Author: Chelsea Edmonds
+*)
+
+section \<open>Group Divisible Designs\<close>
+text \<open>Definitions in this section taken from the handbook \cite{colbournHandbookCombinatorialDesigns2007}
+and Stinson \cite{stinsonCombinatorialDesignsConstructions2004}\<close>
+theory Group_Divisible_Designs imports Resolvable_Designs
+begin
+
+subsection \<open>Group design\<close>
+text \<open>We define a group design to have an additional paramater $G$ which is a partition on the point
+set $V$. This is not defined in the handbook, but is a precursor to GDD's without index constraints\<close>
+
+locale group_design = proper_design +
+ fixes groups :: "'a set set" ("\<G>")
+ assumes group_partitions: "partition_on \<V> \<G>"
+ assumes groups_size: "card \<G> > 1"
+begin
+
+lemma groups_not_empty: "\<G> \<noteq> {}"
+ using groups_size by auto
+
+lemma num_groups_lt_points: "card \<G> \<le> \<v>"
+ by (simp add: partition_on_le_set_elements finite_sets group_partitions)
+
+lemma groups_disjoint: "disjoint \<G>"
+ using group_partitions partition_onD2 by auto
+
+lemma groups_disjoint_pairwise: "G1 \<in> \<G> \<Longrightarrow> G2 \<in> \<G> \<Longrightarrow> G1 \<noteq> G2 \<Longrightarrow> disjnt G1 G2"
+ using group_partitions partition_onD2 pairwiseD by fastforce
+
+lemma point_in_one_group: "x \<in> G1 \<Longrightarrow> G1 \<in> \<G> \<Longrightarrow> G2 \<in> \<G> \<Longrightarrow> G1 \<noteq> G2 \<Longrightarrow> x \<notin> G2"
+ using groups_disjoint_pairwise by (simp add: disjnt_iff)
+
+lemma point_has_unique_group: "x \<in> \<V> \<Longrightarrow> \<exists>!G. x \<in> G \<and> G \<in> \<G>"
+ using partition_on_partition_on_unique group_partitions
+ by fastforce
+
+lemma rep_number_point_group_one:
+ assumes "x \<in> \<V>"
+ shows "card {g \<in> \<G> . x \<in> g} = 1"
+proof -
+ obtain g' where "g' \<in> \<G>" and "x \<in> g'"
+ using assms point_has_unique_group by blast
+ then have "{g \<in> \<G> . x \<in> g} = {g'}"
+ using group_partitions partition_onD4 by force
+ thus ?thesis
+ by simp
+qed
+
+lemma point_in_group: "G \<in> \<G> \<Longrightarrow> x \<in> G \<Longrightarrow> x \<in> \<V>"
+ using group_partitions partition_onD1 by auto
+
+lemma point_subset_in_group: "G \<in> \<G> \<Longrightarrow> ps \<subseteq> G \<Longrightarrow> ps \<subseteq> \<V>"
+ using point_in_group by auto
+
+lemma group_subset_point_subset: "G \<in> \<G> \<Longrightarrow> G' \<subseteq> G \<Longrightarrow> ps \<subseteq> G' \<Longrightarrow> ps \<subseteq> \<V>"
+ using point_subset_in_group by auto
+
+lemma groups_finite: "finite \<G>"
+ using finite_elements finite_sets group_partitions by auto
+
+lemma group_elements_finite: "G \<in> \<G> \<Longrightarrow> finite G"
+ using groups_finite finite_sets group_partitions
+ by (meson finite_subset point_in_group subset_iff)
+
+lemma v_equals_sum_group_sizes: "\<v> = (\<Sum>G \<in> \<G>. card G)"
+ using group_partitions groups_disjoint partition_onD1 card_Union_disjoint group_elements_finite
+ by fastforce
+
+lemma gdd_min_v: "\<v> \<ge> 2"
+proof -
+ have assm: "card \<G> \<ge> 2" using groups_size by simp
+ then have "\<And> G . G \<in> \<G> \<Longrightarrow> G \<noteq> {}" using partition_onD3 group_partitions by auto
+ then have "\<And> G . G \<in> \<G> \<Longrightarrow> card G \<ge> 1"
+ using group_elements_finite card_0_eq by fastforce
+ then have " (\<Sum>G \<in> \<G>. card G) \<ge> 2" using assm
+ using sum_mono by force
+ thus ?thesis using v_equals_sum_group_sizes
+ by linarith
+qed
+
+lemma min_group_size: "G \<in> \<G> \<Longrightarrow> card G \<ge> 1"
+ using partition_onD3 group_partitions
+ using group_elements_finite not_le_imp_less by fastforce
+
+lemma group_size_lt_v:
+ assumes "G \<in> \<G>"
+ shows "card G < \<v>"
+proof -
+ have "(\<Sum>G' \<in> \<G>. card G') = \<v>" using gdd_min_v v_equals_sum_group_sizes
+ by linarith
+ then have split_sum: "card G + (\<Sum>G' \<in> (\<G> - {G}). card G') = \<v>" using assms sum.remove
+ by (metis groups_finite v_equals_sum_group_sizes)
+ have "card (\<G> - {G}) \<ge> 1" using groups_size
+ by (simp add: assms groups_finite)
+ then obtain G' where gin: "G' \<in> (\<G> - {G})"
+ by (meson elem_exists_non_empty_set less_le_trans less_numeral_extra(1))
+ then have "card G' \<ge> 1" using min_group_size by auto
+ then have "(\<Sum>G' \<in> (\<G> - {G}). card G') \<ge> 1"
+ by (metis gin finite_Diff groups_finite leI less_one sum_eq_0_iff)
+ thus ?thesis using split_sum
+ by linarith
+qed
+
+subsubsection \<open>Group Type\<close>
+
+text \<open>GDD's have a "type", which is defined by a sequence of group sizes $g_i$, and the number
+of groups of that size $a_i$: $g_1^{a_1}g2^{a_2}...g_n^{a_n}$\<close>
+definition group_sizes :: "nat set" where
+"group_sizes \<equiv> {card G | G . G \<in> \<G>}"
+
+definition groups_of_size :: "nat \<Rightarrow> nat" where
+"groups_of_size g \<equiv> card { G \<in> \<G> . card G = g }"
+
+definition group_type :: "(nat \<times> nat) set" where
+"group_type \<equiv> {(g, groups_of_size g) | g . g \<in> group_sizes }"
+
+lemma group_sizes_min: "x \<in> group_sizes \<Longrightarrow> x \<ge> 1 "
+ unfolding group_sizes_def using min_group_size group_size_lt_v by auto
+
+lemma group_sizes_max: "x \<in> group_sizes \<Longrightarrow> x < \<v> "
+ unfolding group_sizes_def using min_group_size group_size_lt_v by auto
+
+lemma group_size_implies_group_existance: "x \<in> group_sizes \<Longrightarrow> \<exists>G. G \<in> \<G> \<and> card G = x"
+ unfolding group_sizes_def by auto
+
+lemma groups_of_size_zero: "groups_of_size 0 = 0"
+proof -
+ have empty: "{G \<in> \<G> . card G = 0} = {}" using min_group_size
+ by fastforce
+ thus ?thesis unfolding groups_of_size_def
+ by (simp add: empty)
+qed
+
+lemma groups_of_size_max:
+ assumes "g \<ge> \<v>"
+ shows "groups_of_size g = 0"
+proof -
+ have "{G \<in> \<G> . card G = g} = {}" using group_size_lt_v assms by fastforce
+ thus ?thesis unfolding groups_of_size_def
+ by (simp add: \<open>{G \<in> \<G>. card G = g} = {}\<close>)
+qed
+
+lemma group_type_contained_sizes: "(g, a) \<in> group_type \<Longrightarrow> g \<in> group_sizes"
+ unfolding group_type_def by simp
+
+lemma group_type_contained_count: "(g, a) \<in> group_type \<Longrightarrow> card {G \<in> \<G> . card G = g} = a"
+ unfolding group_type_def groups_of_size_def by simp
+
+lemma group_card_in_sizes: "g \<in> \<G> \<Longrightarrow> card g \<in> group_sizes"
+ unfolding group_sizes_def by auto
+
+lemma group_card_non_zero_groups_of_size_min:
+ assumes "g \<in> \<G>"
+ assumes "card g = a"
+ shows "groups_of_size a \<ge> 1"
+proof -
+ have "g \<in> {G \<in> \<G> . card G = a}" using assms by simp
+ then have "{G \<in> \<G> . card G = a} \<noteq> {}" by auto
+ then have "card {G \<in> \<G> . card G = a} \<noteq> 0"
+ by (simp add: groups_finite)
+ thus ?thesis unfolding groups_of_size_def by simp
+qed
+
+lemma elem_in_group_sizes_min_of_size:
+ assumes "a \<in> group_sizes"
+ shows "groups_of_size a \<ge> 1"
+ using assms group_card_non_zero_groups_of_size_min group_size_implies_group_existance by blast
+
+lemma group_card_non_zero_groups_of_size_max:
+ shows "groups_of_size a \<le> \<v>"
+proof -
+ have "{G \<in> \<G> . card G = a} \<subseteq> \<G>" by simp
+ then have "card {G \<in> \<G> . card G = a} \<le> card \<G>"
+ by (simp add: card_mono groups_finite)
+ thus ?thesis
+ using groups_of_size_def num_groups_lt_points by auto
+qed
+
+lemma group_card_in_type: "g \<in> \<G> \<Longrightarrow> \<exists> x . (card g, x) \<in> group_type \<and> x \<ge> 1"
+ unfolding group_type_def using group_card_non_zero_groups_of_size_min
+ by (simp add: group_card_in_sizes)
+
+lemma partition_groups_on_size: "partition_on \<G> {{ G \<in> \<G> . card G = g } | g . g \<in> group_sizes}"
+proof (intro partition_onI, auto)
+ fix g
+ assume a1: "g \<in> group_sizes"
+ assume " \<forall>x. x \<in> \<G> \<longrightarrow> card x \<noteq> g"
+ then show False using a1 group_size_implies_group_existance by auto
+next
+ fix x
+ assume "x \<in> \<G>"
+ then show "\<exists>xa. (\<exists>g. xa = {G \<in> \<G>. card G = g} \<and> g \<in> group_sizes) \<and> x \<in> xa"
+ using group_card_in_sizes by auto
+qed
+
+lemma group_size_partition_covers_points: "\<Union>(\<Union>{{ G \<in> \<G> . card G = g } | g . g \<in> group_sizes}) = \<V>"
+ by (metis (no_types, lifting) group_partitions partition_groups_on_size partition_onD1)
+
+lemma groups_of_size_alt_def_count: "groups_of_size g = count {# card G . G \<in># mset_set \<G> #} g"
+proof -
+ have a: "groups_of_size g = card { G \<in> \<G> . card G = g }" unfolding groups_of_size_def by simp
+ then have "groups_of_size g = size {# G \<in># (mset_set \<G>) . card G = g #}"
+ using groups_finite by auto
+ then have size_repr: "groups_of_size g = size {# x \<in># {# card G . G \<in># mset_set \<G> #} . x = g #}"
+ using groups_finite by (simp add: filter_mset_image_mset)
+ have "group_sizes = set_mset ({# card G . G \<in># mset_set \<G> #})"
+ using group_sizes_def groups_finite by auto
+ thus ?thesis using size_repr by (simp add: count_size_set_repr)
+qed
+
+lemma v_sum_type_rep: "\<v> = (\<Sum> g \<in> group_sizes . g * (groups_of_size g))"
+proof -
+ have gs: "set_mset {# card G . G \<in># mset_set \<G> #} = group_sizes"
+ unfolding group_sizes_def using groups_finite by auto
+ have "\<v> = card (\<Union>(\<Union>{{ G \<in> \<G> . card G = g } | g . g \<in> group_sizes}))"
+ using group_size_partition_covers_points by simp
+ have v1: "\<v> = (\<Sum>x \<in># {# card G . G \<in># mset_set \<G> #}. x)"
+ by (simp add: sum_unfold_sum_mset v_equals_sum_group_sizes)
+ then have "\<v> = (\<Sum>x \<in> set_mset {# card G . G \<in># mset_set \<G> #} . x * (count {# card G . G \<in># mset_set \<G> #} x))"
+ using mset_set_size_card_count by (simp add: v1)
+ thus ?thesis using gs groups_of_size_alt_def_count by auto
+qed
+
+end
+
+subsubsection \<open>Uniform Group designs\<close>
+text \<open>A group design requiring all groups are the same size\<close>
+locale uniform_group_design = group_design +
+ fixes u_group_size :: nat ("\<m>")
+ assumes uniform_groups: "G \<in> \<G> \<Longrightarrow> card G = \<m>"
+
+begin
+
+lemma m_positive: "\<m> \<ge> 1"
+proof -
+ obtain G where "G \<in> \<G>" using groups_size elem_exists_non_empty_set gr_implies_not_zero by blast
+ thus ?thesis using uniform_groups min_group_size by fastforce
+qed
+
+lemma uniform_groups_alt: " \<forall> G \<in> \<G> . card G = \<m>"
+ using uniform_groups by blast
+
+lemma uniform_groups_group_sizes: "group_sizes = {\<m>}"
+ using design_points_nempty group_card_in_sizes group_size_implies_group_existance
+ point_has_unique_group uniform_groups_alt by force
+
+lemma uniform_groups_group_size_singleton: "is_singleton (group_sizes)"
+ using uniform_groups_group_sizes by auto
+
+lemma set_filter_eq_P_forall:"\<forall> x \<in> X . P x \<Longrightarrow> Set.filter P X = X"
+ by (simp add: Collect_conj_eq Int_absorb2 Set.filter_def subsetI)
+
+lemma uniform_groups_groups_of_size_m: "groups_of_size \<m> = card \<G>"
+proof(simp add: groups_of_size_def)
+ have "{G \<in> \<G>. card G = \<m>} = \<G>" using uniform_groups_alt set_filter_eq_P_forall by auto
+ thus "card {G \<in> \<G>. card G = \<m>} = card \<G>" by simp
+qed
+
+lemma uniform_groups_of_size_not_m: "x \<noteq> \<m> \<Longrightarrow> groups_of_size x = 0"
+ by (simp add: groups_of_size_def card_eq_0_iff uniform_groups)
+
+end
+
+subsection \<open>GDD\<close>
+text \<open>A GDD extends a group design with an additional index parameter.
+Each pair of elements must occur either \Lambda times if in diff groups, or 0 times if in the same
+group\<close>
+
+locale GDD = group_design +
+ fixes index :: int ("\<Lambda>")
+ assumes index_ge_1: "\<Lambda> \<ge> 1"
+ assumes index_together: "G \<in> \<G> \<Longrightarrow> x \<in> G \<Longrightarrow> y \<in> G \<Longrightarrow> x \<noteq> y \<Longrightarrow> \<B> index {x, y} = 0"
+ assumes index_distinct: "G1 \<in> \<G> \<Longrightarrow> G2 \<in> \<G> \<Longrightarrow> G1 \<noteq> G2 \<Longrightarrow> x \<in> G1 \<Longrightarrow> y \<in> G2 \<Longrightarrow>
+ \<B> index {x, y} = \<Lambda>"
+begin
+
+lemma points_sep_groups_ne: "G1 \<in> \<G> \<Longrightarrow> G2 \<in> \<G> \<Longrightarrow> G1 \<noteq> G2 \<Longrightarrow> x \<in> G1 \<Longrightarrow> y \<in> G2 \<Longrightarrow> x \<noteq> y"
+ by (meson point_in_one_group)
+
+lemma index_together_alt_ss: "ps \<subseteq> G \<Longrightarrow> G \<in> \<G> \<Longrightarrow> card ps = 2 \<Longrightarrow> \<B> index ps = 0"
+ using index_together by (metis card_2_iff insert_subset)
+
+lemma index_distinct_alt_ss: "ps \<subseteq> \<V> \<Longrightarrow> card ps = 2 \<Longrightarrow> (\<And> G . G \<in> \<G> \<Longrightarrow> \<not> ps \<subseteq> G) \<Longrightarrow>
+ \<B> index ps = \<Lambda>"
+ using index_distinct by (metis card_2_iff empty_subsetI insert_subset point_has_unique_group)
+
+lemma gdd_index_options: "ps \<subseteq> \<V> \<Longrightarrow> card ps = 2 \<Longrightarrow> \<B> index ps = 0 \<or> \<B> index ps = \<Lambda>"
+ using index_distinct_alt_ss index_together_alt_ss by blast
+
+lemma index_zero_implies_same_group: "ps \<subseteq> \<V> \<Longrightarrow> card ps = 2 \<Longrightarrow> \<B> index ps = 0 \<Longrightarrow>
+ \<exists> G \<in> \<G> . ps \<subseteq> G" using index_distinct_alt_ss gr_implies_not_zero
+ by (metis index_ge_1 less_one of_nat_0 of_nat_1 of_nat_le_0_iff)
+
+lemma index_zero_implies_same_group_unique: "ps \<subseteq> \<V> \<Longrightarrow> card ps = 2 \<Longrightarrow> \<B> index ps = 0 \<Longrightarrow>
+ \<exists>! G \<in> \<G> . ps \<subseteq> G"
+ by (meson GDD.index_zero_implies_same_group GDD_axioms card_2_iff' group_design.point_in_one_group
+ group_design_axioms in_mono)
+
+lemma index_not_zero_impl_diff_group: "ps \<subseteq> \<V> \<Longrightarrow> card ps = 2 \<Longrightarrow> \<B> index ps = \<Lambda> \<Longrightarrow>
+ (\<And> G . G \<in> \<G> \<Longrightarrow> \<not> ps \<subseteq> G)"
+ using index_ge_1 index_together_alt_ss by auto
+
+lemma index_zero_implies_one_group:
+ assumes "ps \<subseteq> \<V>"
+ and "card ps = 2"
+ and "\<B> index ps = 0"
+ shows "size {#b \<in># mset_set \<G> . ps \<subseteq> b#} = 1"
+proof -
+ obtain G where ging: "G \<in> \<G>" and psin: "ps \<subseteq> G"
+ using index_zero_implies_same_group groups_size assms by blast
+ then have unique: "\<And> G2 . G2 \<in> \<G> \<Longrightarrow> G \<noteq> G2 \<Longrightarrow> \<not> ps \<subseteq> G2"
+ using index_zero_implies_same_group_unique by (metis assms)
+ have "\<And> G'. G' \<in> \<G> \<longleftrightarrow> G' \<in># mset_set \<G>"
+ by (simp add: groups_finite)
+ then have eq_mset: "{#b \<in># mset_set \<G> . ps \<subseteq> b#} = mset_set {b \<in> \<G> . ps \<subseteq> b}"
+ using filter_mset_mset_set groups_finite by blast
+ then have "{b \<in> \<G> . ps \<subseteq> b} = {G}" using unique psin
+ by (smt Collect_cong ging singleton_conv)
+ thus ?thesis by (simp add: eq_mset)
+qed
+
+lemma index_distinct_group_num_alt_def: "ps \<subseteq> \<V> \<Longrightarrow> card ps = 2 \<Longrightarrow>
+ size {#b \<in># mset_set \<G> . ps \<subseteq> b#} = 0 \<Longrightarrow> \<B> index ps = \<Lambda>"
+ by (metis gdd_index_options index_zero_implies_one_group numeral_One zero_neq_numeral)
+
+lemma index_non_zero_implies_no_group:
+ assumes "ps \<subseteq> \<V>"
+ and "card ps = 2"
+ and "\<B> index ps = \<Lambda>"
+ shows "size {#b \<in># mset_set \<G> . ps \<subseteq> b#} = 0"
+proof -
+ have "\<And> G . G \<in> \<G> \<Longrightarrow> \<not> ps \<subseteq> G" using index_not_zero_impl_diff_group assms by simp
+ then have "{#b \<in># mset_set \<G> . ps \<subseteq> b#} = {#}"
+ using filter_mset_empty_if_finite_and_filter_set_empty by force
+ thus ?thesis by simp
+qed
+
+lemma gdd_index_non_zero_iff: "ps \<subseteq> \<V> \<Longrightarrow> card ps = 2 \<Longrightarrow>
+ \<B> index ps = \<Lambda> \<longleftrightarrow> size {#b \<in># mset_set \<G> . ps \<subseteq> b#} = 0"
+ using index_non_zero_implies_no_group index_distinct_group_num_alt_def by auto
+
+lemma gdd_index_zero_iff: "ps \<subseteq> \<V> \<Longrightarrow> card ps = 2 \<Longrightarrow>
+ \<B> index ps = 0 \<longleftrightarrow> size {#b \<in># mset_set \<G> . ps \<subseteq> b#} = 1"
+ apply (auto simp add: index_zero_implies_one_group)
+ by (metis GDD.gdd_index_options GDD_axioms index_non_zero_implies_no_group old.nat.distinct(2))
+
+lemma points_index_upper_bound: "ps \<subseteq> \<V> \<Longrightarrow> card ps = 2 \<Longrightarrow> \<B> index ps \<le> \<Lambda>"
+ using gdd_index_options index_ge_1
+ by (metis int_one_le_iff_zero_less le_refl of_nat_0 of_nat_0_le_iff of_nat_le_iff zero_less_imp_eq_int)
+
+lemma index_1_imp_mult_1:
+ assumes "\<Lambda> = 1"
+ assumes "bl \<in># \<B>"
+ assumes "card bl \<ge> 2"
+ shows "multiplicity bl = 1"
+proof (rule ccontr)
+ assume "\<not> (multiplicity bl = 1)"
+ then have "multiplicity bl \<noteq> 1" and "multiplicity bl \<noteq> 0" using assms by simp_all
+ then have m: "multiplicity bl \<ge> 2" by linarith
+ obtain ps where ps: "ps \<subseteq> bl \<and> card ps = 2"
+ using nat_int_comparison(3) obtain_subset_with_card_n by (metis assms(3))
+ then have "\<B> index ps \<ge> 2"
+ using m points_index_count_min ps by blast
+ then show False using assms index_distinct ps antisym_conv2 not_numeral_less_zero
+ numeral_le_one_iff points_index_ps_nin semiring_norm(69) zero_neq_numeral
+ by (metis gdd_index_options int_int_eq int_ops(2))
+qed
+
+lemma simple_if_block_size_gt_2:
+ assumes "\<And> bl . card bl \<ge> 2"
+ assumes "\<Lambda> = 1"
+ shows "simple_design \<V> \<B>"
+ using index_1_imp_mult_1 assms apply (unfold_locales)
+ by (metis card.empty not_numeral_le_zero)
+
+end
+
+subsubsection \<open>Sub types of GDD's\<close>
+
+text \<open>In literature, a GDD is usually defined in a number of different ways,
+including factors such as block size limitations\<close>
+locale K_\<Lambda>_GDD = K_block_design + GDD
+
+locale k_\<Lambda>_GDD = block_design + GDD
+
+sublocale k_\<Lambda>_GDD \<subseteq> K_\<Lambda>_GDD \<V> \<B> "{\<k>}" \<G> \<Lambda>
+ by (unfold_locales)
+
+locale K_GDD = K_\<Lambda>_GDD \<V> \<B> \<K> \<G> 1
+ for point_set ("\<V>") and block_collection ("\<B>") and sizes ("\<K>") and groups ("\<G>")
+
+locale k_GDD = k_\<Lambda>_GDD \<V> \<B> \<k> \<G> 1
+ for point_set ("\<V>") and block_collection ("\<B>") and u_block_size ("\<k>") and groups ("\<G>")
+
+sublocale k_GDD \<subseteq> K_GDD \<V> \<B> "{\<k>}" \<G>
+ by (unfold_locales)
+
+lemma (in K_GDD) multiplicity_1: "bl \<in># \<B> \<Longrightarrow> card bl \<ge> 2 \<Longrightarrow> multiplicity bl = 1"
+ using index_1_imp_mult_1 by simp
+
+locale RGDD = GDD + resolvable_design
+
+subsection \<open>GDD and PBD Constructions\<close>
+text \<open>GDD's are commonly studied alongside PBD's (pairwise balanced designs). Many constructions
+have been developed for designs to create a GDD from a PBD and vice versa. In particular,
+Wilsons Construction is a well known construction, which is formalised in this section. It
+should be noted that many of the more basic constructions in this section are often stated without
+proof/all the necessary assumptions in textbooks/course notes.\<close>
+
+context GDD
+begin
+
+subsubsection \<open>GDD Delete Point construction\<close>
+lemma delete_point_index_zero:
+ assumes "G \<in> {g - {x} |g. g \<in> \<G> \<and> g \<noteq> {x}}"
+ and "y \<in> G" and "z \<in> G" and "z\<noteq> y"
+shows "(del_point_blocks x) index {y, z} = 0"
+proof -
+ have "y \<noteq> x" using assms(1) assms(2) by blast
+ have "z \<noteq> x" using assms(1) assms(3) by blast
+ obtain G' where ing: "G' \<in> \<G>" and ss: "G \<subseteq> G'"
+ using assms(1) by auto
+ have "{y, z} \<subseteq> G" by (simp add: assms(2) assms(3))
+ then have "{y, z} \<subseteq> \<V>"
+ by (meson ss ing group_subset_point_subset)
+ then have "{y, z} \<subseteq> (del_point x)"
+ using \<open>y \<noteq> x\<close> \<open>z \<noteq> x\<close> del_point_def by fastforce
+ thus ?thesis using delete_point_index_eq index_together
+ by (metis assms(2) assms(3) assms(4) in_mono ing ss)
+qed
+
+lemma delete_point_index:
+ assumes "G1 \<in> {g - {x} |g. g \<in> \<G> \<and> g \<noteq> {x}}"
+ assumes "G2 \<in> {g - {x} |g. g \<in> \<G> \<and> g \<noteq> {x}}"
+ assumes "G1 \<noteq> G2" and "y \<in> G1" and "z \<in> G2"
+ shows "del_point_blocks x index {y, z} = \<Lambda>"
+proof -
+ have "y \<noteq> x" using assms by blast
+ have "z \<noteq> x" using assms by blast
+ obtain G1' where ing1: "G1' \<in> \<G>" and t1: "G1 = G1' - {x}"
+ using assms(1) by auto
+ obtain G2' where ing2: "G2' \<in> \<G>" and t2: "G2 = G2' - {x}"
+ using assms(2) by auto
+ then have ss1: "G1 \<subseteq> G1'" and ss2: "G2 \<subseteq> G2'" using t1 by auto
+ then have "{y, z} \<subseteq> \<V>" using ing1 ing2 ss1 ss2 assms(4) assms(5)
+ by (metis empty_subsetI insert_absorb insert_subset point_in_group)
+ then have "{y, z} \<subseteq> del_point x"
+ using \<open>y \<noteq> x\<close> \<open>z \<noteq> x\<close> del_point_def by auto
+ then have indx: "del_point_blocks x index {y, z} = \<B> index {y, z}"
+ using delete_point_index_eq by auto
+ have "G1' \<noteq> G2'" using assms t1 t2 by fastforce
+ thus ?thesis using index_distinct
+ using indx assms(4) assms(5) ing1 ing2 t1 t2 by auto
+qed
+
+lemma delete_point_group_size:
+ assumes "{x} \<in> \<G> \<Longrightarrow> card \<G> > 2"
+ shows "1 < card {g - {x} |g. g \<in> \<G> \<and> g \<noteq> {x}}"
+proof (cases "{x} \<in> \<G>")
+ case True
+ then have "\<And> g . g \<in> (\<G> - {{x}}) \<Longrightarrow> x \<notin> g"
+ by (meson disjnt_insert1 groups_disjoint pairwise_alt)
+ then have simpg: "\<And> g . g \<in> (\<G> - {{x}}) \<Longrightarrow> g - {x} = g"
+ by simp
+ have "{g - {x} |g. g \<in> \<G> \<and> g \<noteq> {x}} = {g - {x} |g. (g \<in> \<G> - {{x}})}" using True
+ by force
+ then have "{g - {x} |g. g \<in> \<G> \<and> g \<noteq> {x}} = {g |g. (g \<in> \<G> - {{x}})}" using simpg
+ by (smt (verit) Collect_cong)
+ then have eq: "{g - {x} |g. g \<in> \<G> \<and> g \<noteq> {x}} = \<G> - {{x}}" using set_self_img_compr by blast
+ have "card (\<G> - {{x}}) = card \<G> - 1" using True
+ by (simp add: groups_finite)
+ then show ?thesis using True assms eq diff_is_0_eq' by force
+next
+ case False
+ then have "\<And>g' y. {x} \<notin> \<G> \<Longrightarrow> g' \<in> \<G> \<Longrightarrow> y \<in> \<G> \<Longrightarrow> g' - {x} = y - {x} \<Longrightarrow> g' = y"
+ by (metis all_not_in_conv insert_Diff_single insert_absorb insert_iff points_sep_groups_ne)
+ then have inj: "inj_on (\<lambda> g . g - {x}) \<G>" by (simp add: inj_onI False)
+ have "{g - {x} |g. g \<in> \<G> \<and> g \<noteq> {x}} = {g - {x} |g. g \<in> \<G>}" using False by auto
+ then have "card {g - {x} |g. g \<in> \<G> \<and> g \<noteq> {x}} = card \<G>" using inj groups_finite card_image
+ by (auto simp add: card_image setcompr_eq_image)
+ then show ?thesis using groups_size by presburger
+qed
+
+lemma GDD_by_deleting_point:
+ assumes "\<And>bl. bl \<in># \<B> \<Longrightarrow> x \<in> bl \<Longrightarrow> 2 \<le> card bl"
+ assumes "{x} \<in> \<G> \<Longrightarrow> card \<G> > 2"
+ shows "GDD (del_point x) (del_point_blocks x) {g - {x} | g . g \<in> \<G> \<and> g \<noteq> {x}} \<Lambda>"
+proof -
+ interpret pd: proper_design "del_point x" "del_point_blocks x"
+ using delete_point_proper assms by blast
+ show ?thesis using delete_point_index_zero delete_point_index assms delete_point_group_size
+ by(unfold_locales) (simp_all add: partition_on_remove_pt group_partitions index_ge_1 del_point_def)
+qed
+
+end
+
+context K_GDD begin
+
+subsubsection \<open>PBD construction from GDD\<close>
+text \<open>Two well known PBD constructions involve taking a GDD and either combining the groups and
+blocks to form a new block collection, or by adjoining a point\<close>
+
+text \<open>First prove that combining the groups and block set results in a constant index\<close>
+lemma kgdd1_points_index_group_block:
+ assumes "ps \<subseteq> \<V>"
+ and "card ps = 2"
+ shows "(\<B> + mset_set \<G>) index ps = 1"
+proof -
+ have index1: "(\<And> G . G \<in> \<G> \<Longrightarrow> \<not> ps \<subseteq> G) \<Longrightarrow> \<B> index ps = 1"
+ using index_distinct_alt_ss assms by fastforce
+ have groups1: "\<B> index ps = 0 \<Longrightarrow> size {#b \<in># mset_set \<G> . ps \<subseteq> b#} = 1"
+ using index_zero_implies_one_group assms by simp
+ then have "(\<B> + mset_set \<G>) index ps = size (filter_mset ((\<subseteq>) ps) (\<B> + mset_set \<G>))"
+ by (simp add: points_index_def)
+ thus ?thesis using index1 groups1 gdd_index_non_zero_iff gdd_index_zero_iff assms
+ gdd_index_options points_index_def filter_union_mset union_commute
+ by (smt (z3) empty_neutral(1) less_irrefl_nat nonempty_has_size of_nat_1_eq_iff)
+qed
+
+text \<open>Combining blocks and the group set forms a PBD\<close>
+lemma combine_block_groups_pairwise: "pairwise_balance \<V> (\<B> + mset_set \<G>) 1"
+proof -
+ let ?B = "\<B> + mset_set \<G>"
+ have ss: "\<And> G. G \<in> \<G> \<Longrightarrow> G \<subseteq> \<V>"
+ by (simp add: point_in_group subsetI)
+ have "\<And> G. G \<in> \<G> \<Longrightarrow> G \<noteq> {}" using group_partitions
+ using partition_onD3 by auto
+ then interpret inc: design \<V> ?B
+ proof (unfold_locales)
+ show "\<And>b. (\<And>G. G \<in> \<G> \<Longrightarrow> G \<noteq> {}) \<Longrightarrow> b \<in># \<B> + mset_set \<G> \<Longrightarrow> b \<subseteq> \<V>"
+ by (metis finite_set_mset_mset_set groups_finite ss union_iff wellformed)
+ show "(\<And>G. G \<in> \<G> \<Longrightarrow> G \<noteq> {}) \<Longrightarrow> finite \<V>" by (simp add: finite_sets)
+ show "\<And>bl. (\<And>G. G \<in> \<G> \<Longrightarrow> G \<noteq> {}) \<Longrightarrow> bl \<in># \<B> + mset_set \<G> \<Longrightarrow> bl \<noteq> {}"
+ using blocks_nempty groups_finite by auto
+ qed
+ show ?thesis proof (unfold_locales)
+ show "inc.\<b> \<noteq> 0" using b_positive by auto
+ show "(1 ::int) \<le> 2" by simp
+ show "2 \<le> inc.\<v>" by (simp add: gdd_min_v)
+ then show "\<And>ps. ps \<subseteq> \<V> \<Longrightarrow> int (card ps) = 2 \<Longrightarrow> int ((\<B> + mset_set \<G>) index ps) = 1"
+ using kgdd1_points_index_group_block by simp
+ qed
+qed
+
+lemma combine_block_groups_PBD:
+ assumes "\<And> G. G \<in> \<G> \<Longrightarrow> card G \<in> \<K>"
+ assumes "\<And> k . k \<in> \<K> \<Longrightarrow> k \<ge> 2"
+ shows "PBD \<V> (\<B> + mset_set \<G>) \<K>"
+proof -
+ let ?B = "\<B> + mset_set \<G>"
+ interpret inc: pairwise_balance \<V> ?B 1 using combine_block_groups_pairwise by simp
+ show ?thesis using assms block_sizes groups_finite positive_ints
+ by (unfold_locales) auto
+qed
+
+text \<open>Prove adjoining a point to each group set results in a constant points index\<close>
+lemma kgdd1_index_adjoin_group_block:
+ assumes "x \<notin> \<V>"
+ assumes "ps \<subseteq> insert x \<V>"
+ assumes "card ps = 2"
+ shows "(\<B> + mset_set {insert x g |g. g \<in> \<G>}) index ps = 1"
+proof -
+ have "inj_on ((insert) x) \<G>"
+ by (meson assms(1) inj_onI insert_ident point_in_group)
+ then have eq: "mset_set {insert x g |g. g \<in> \<G>} = {# insert x g . g \<in># mset_set \<G>#}"
+ by (simp add: image_mset_mset_set setcompr_eq_image)
+ thus ?thesis
+ proof (cases "x \<in> ps")
+ case True
+ then obtain y where y_ps: "ps = {x, y}" using assms(3)
+ by (metis card_2_iff doubleton_eq_iff insertE singletonD)
+ then have ynex: "y \<noteq> x" using assms by fastforce
+ have yinv: "y \<in> \<V>"
+ using assms(2) y_ps ynex by auto
+ have all_g: "\<And> g. g \<in># (mset_set {insert x g |g. g \<in> \<G>}) \<Longrightarrow> x \<in> g"
+ using eq by force
+ have iff: "\<And> g . g \<in> \<G> \<Longrightarrow> y \<in> (insert x g) \<longleftrightarrow> y \<in> g" using ynex by simp
+ have b: "\<B> index ps = 0"
+ using True assms(1) points_index_ps_nin by fastforce
+ then have "(\<B> + mset_set {insert x g |g. g \<in> \<G>}) index ps =
+ (mset_set {insert x g |g. g \<in> \<G>}) index ps"
+ using eq by (simp add: point_index_distrib)
+ also have "... = (mset_set {insert x g |g. g \<in> \<G>}) rep y" using points_index_pair_rep_num
+ by (metis (no_types, lifting) all_g y_ps)
+ also have 0: "... = card {b \<in> {insert x g |g. g \<in> \<G>} . y \<in> b}"
+ by (simp add: groups_finite rep_number_on_set_def)
+ also have 1: "... = card {insert x g |g. g \<in> \<G> \<and> y \<in> insert x g}"
+ by (smt (verit) Collect_cong mem_Collect_eq)
+ also have 2: " ... = card {insert x g |g. g \<in> \<G> \<and> y \<in> g}"
+ using iff by metis
+ also have "... = card {g \<in> \<G> . y \<in> g}" using 1 2 0 empty_iff eq groups_finite ynex insert_iff
+ by (metis points_index_block_image_add_eq points_index_single_rep_num rep_number_on_set_def)
+ finally have "(\<B> + mset_set {insert x g |g. g \<in> \<G>}) index ps = 1"
+ using rep_number_point_group_one yinv by simp
+ then show ?thesis
+ by simp
+ next
+ case False
+ then have v: "ps \<subseteq> \<V>" using assms(2) by auto
+ then have "(\<B> + mset_set {insert x g |g. g \<in> \<G>}) index ps = (\<B> + mset_set \<G>) index ps"
+ using eq by (simp add: points_index_block_image_add_eq False point_index_distrib)
+ then show ?thesis using v assms kgdd1_points_index_group_block by simp
+ qed
+qed
+
+lemma pairwise_by_adjoining_point:
+ assumes "x \<notin> \<V>"
+ shows "pairwise_balance (add_point x) (\<B> + mset_set { insert x g | g. g \<in> \<G>}) 1"
+proof -
+ let ?B = "\<B> + mset_set { insert x g | g. g \<in> \<G>}"
+ let ?V = "add_point x"
+ have vdef: "?V = \<V> \<union> {x}" using add_point_def by simp
+ show ?thesis unfolding add_point_def using finite_sets b_positive
+ proof (unfold_locales, simp_all)
+ have "\<And> G. G \<in> \<G> \<Longrightarrow> insert x G \<subseteq> ?V"
+ by (simp add: point_in_group subsetI vdef)
+ then have "\<And> G. G \<in># (mset_set { insert x g | g. g \<in> \<G>}) \<Longrightarrow> G \<subseteq> ?V"
+ by (smt (verit, del_insts) elem_mset_set empty_iff infinite_set_mset_mset_set mem_Collect_eq)
+ then show "\<And>b. b \<in># \<B> \<or> b \<in># mset_set {insert x g |g. g \<in> \<G>} \<Longrightarrow> b \<subseteq> insert x \<V>"
+ using wellformed add_point_def by fastforce
+ next
+ have "\<And> G. G \<in> \<G> \<Longrightarrow> insert x G \<noteq> {}" using group_partitions
+ using partition_onD3 by auto
+ then have gnempty: "\<And> G. G \<in># (mset_set { insert x g | g. g \<in> \<G>}) \<Longrightarrow> G \<noteq> {}"
+ by (smt (verit, del_insts) elem_mset_set empty_iff infinite_set_mset_mset_set mem_Collect_eq)
+ then show "\<And>bl. bl \<in># \<B> \<or> bl \<in># mset_set {insert x g |g. g \<in> \<G>} \<Longrightarrow> bl \<noteq> {}"
+ using blocks_nempty by auto
+ next
+ have "card \<V> \<ge> 2" using gdd_min_v by simp
+ then have "card (insert x \<V>) \<ge> 2"
+ by (meson card_insert_le dual_order.trans finite_sets)
+ then show "2 \<le> int (card (insert x \<V>))" by auto
+ next
+ show "\<And>ps. ps \<subseteq> insert x \<V> \<Longrightarrow>
+ card ps = 2 \<Longrightarrow> (\<B> + mset_set {insert x g |g. g \<in> \<G>}) index ps = Suc 0"
+ using kgdd1_index_adjoin_group_block by (simp add: assms)
+ qed
+qed
+
+lemma PBD_by_adjoining_point:
+ assumes "x \<notin> \<V>"
+ assumes "\<And> k . k \<in> \<K> \<Longrightarrow> k \<ge> 2"
+ shows "PBD (add_point x) (\<B> + mset_set { insert x g | g. g \<in> \<G>}) (\<K> \<union> {(card g) + 1 | g . g \<in> \<G>})"
+proof -
+ let ?B = "\<B> + mset_set { insert x g | g. g \<in> \<G>}"
+ let ?V = "(add_point x)"
+ interpret inc: pairwise_balance ?V ?B 1 using pairwise_by_adjoining_point assms by auto
+ show ?thesis using block_sizes positive_ints proof (unfold_locales)
+ have xg: "\<And> g. g \<in> \<G> \<Longrightarrow> x \<notin> g"
+ using assms point_in_group by auto
+ have "\<And> bl . bl \<in># \<B> \<Longrightarrow> card bl \<in> \<K>" by (simp add: block_sizes)
+ have "\<And> bl . bl \<in># mset_set {insert x g |g. g \<in> \<G>} \<Longrightarrow> bl \<in> {insert x g | g . g \<in> \<G>}"
+ by (simp add: groups_finite)
+ then have "\<And> bl . bl \<in># mset_set {insert x g |g. g \<in> \<G>} \<Longrightarrow>
+ card bl \<in> {int (card g + 1) |g. g \<in> \<G>}"
+ proof -
+ fix bl
+ assume "bl \<in># mset_set {insert x g |g. g \<in> \<G>}"
+ then have "bl \<in> {insert x g | g . g \<in> \<G>}" by (simp add: groups_finite)
+ then obtain g where gin: "g \<in> \<G>" and i: "bl = insert x g" by auto
+ thus "card bl \<in> {int (card g + 1) |g. g \<in> \<G>}"
+ using gin group_elements_finite i xg by auto
+ qed
+ then show "\<And>bl. bl \<in># \<B> + mset_set {insert x g |g. g \<in> \<G>} \<Longrightarrow>
+ int (card bl) \<in> \<K> \<union> {int (card g + 1) |g. g \<in> \<G>}"
+ using UnI1 UnI2 block_sizes union_iff by (smt (z3) mem_Collect_eq)
+ show "\<And>x. x \<in> \<K> \<union> {int (card g + 1) |g. g \<in> \<G>} \<Longrightarrow> 0 < x"
+ using min_group_size positive_ints by auto
+ show "\<And>k. k \<in> \<K> \<union> {int (card g + 1) |g. g \<in> \<G>} \<Longrightarrow> 2 \<le> k"
+ using min_group_size positive_ints assms by fastforce
+ qed
+qed
+
+subsubsection \<open>Wilson's Construction\<close>
+text \<open>Wilson's construction involves the combination of multiple k-GDD's. This proof was
+based of Stinson \cite{stinsonCombinatorialDesignsConstructions2004}\<close>
+
+lemma wilsons_construction_proper:
+ assumes "card I = w"
+ assumes "w > 0"
+ assumes "\<And> n. n \<in> \<K>' \<Longrightarrow> n \<ge> 2"
+ assumes "\<And> B . B \<in># \<B> \<Longrightarrow> K_GDD (B \<times> I) (f B) \<K>' {{x} \<times> I |x . x \<in> B }"
+ shows "proper_design (\<V> \<times> I) (\<Sum>B \<in># \<B>. (f B))" (is "proper_design ?Y ?B")
+proof (unfold_locales, simp_all)
+ show "\<And>b. \<exists>x\<in>#\<B>. b \<in># f x \<Longrightarrow> b \<subseteq> \<V> \<times> I"
+ proof -
+ fix b
+ assume "\<exists>x\<in>#\<B>. b \<in># f x"
+ then obtain B where "B \<in># \<B>" and "b \<in># (f B)" by auto
+ then interpret kgdd: K_GDD "(B \<times> I)" "(f B)" \<K>' "{{x} \<times> I |x . x \<in> B }" using assms by auto
+ show "b \<subseteq> \<V> \<times> I" using kgdd.wellformed
+ using \<open>B \<in># \<B>\<close> \<open>b \<in># f B\<close> wellformed by fastforce
+ qed
+ show "finite (\<V> \<times> I)" using finite_sets assms bot_nat_0.not_eq_extremum card.infinite by blast
+ show "\<And>bl. \<exists>x\<in>#\<B>. bl \<in># f x \<Longrightarrow> bl \<noteq> {}"
+ proof -
+ fix bl
+ assume "\<exists>x\<in>#\<B>. bl \<in># f x"
+ then obtain B where "B \<in># \<B>" and "bl \<in># (f B)" by auto
+ then interpret kgdd: K_GDD "(B \<times> I)" "(f B)" \<K>' "{{x} \<times> I |x . x \<in> B }" using assms by auto
+ show "bl \<noteq> {}" using kgdd.blocks_nempty by (simp add: \<open>bl \<in># f B\<close>)
+ qed
+ show "\<exists>i\<in>#\<B>. f i \<noteq> {#}"
+ proof -
+ obtain B where "B \<in># \<B>"
+ using design_blocks_nempty by auto
+ then interpret kgdd: K_GDD "(B \<times> I)" "(f B)" \<K>' "{{x} \<times> I |x . x \<in> B }" using assms by auto
+ have "f B \<noteq> {#}" using kgdd.design_blocks_nempty by simp
+ then show "\<exists>i\<in>#\<B>. f i \<noteq> {#}" using \<open>B \<in># \<B>\<close> by auto
+ qed
+qed
+
+lemma pair_construction_block_sizes:
+ assumes "K_GDD (B \<times> I) (f B) \<K>' {{x} \<times> I |x . x \<in> B }"
+ assumes "B \<in># \<B>"
+ assumes "b \<in># (f B)"
+ shows "card b \<in> \<K>'"
+proof -
+ interpret bkgdd: K_GDD "(B \<times> I)" "(f B)" \<K>' "{{x} \<times> I |x . x \<in> B }"
+ using assms by simp
+ show "card b \<in> \<K>'" using bkgdd.block_sizes by (simp add:assms)
+qed
+
+lemma wilsons_construction_index_0:
+ assumes "\<And> B . B \<in># \<B> \<Longrightarrow> K_GDD (B \<times> I) (f B) \<K>' {{x} \<times> I |x . x \<in> B }"
+ assumes "G \<in> {GG \<times> I |GG. GG \<in> \<G>}"
+ assumes "X \<in> G"
+ assumes "Y \<in> G"
+ assumes "X \<noteq> Y"
+ shows "(\<Sum>\<^sub># (image_mset f \<B>)) index {X, Y} = 0"
+proof -
+ obtain G' where gi: "G = G' \<times> I" and ging: "G' \<in> \<G>" using assms by auto
+ obtain x y ix iy where xpair: "X = (x, ix)" and ypair: "Y = (y, iy)" using assms by auto
+ then have ixin: "ix \<in> I" and xing: "x \<in> G'" using assms gi by auto
+ have iyin: "iy \<in> I" and ying: "y \<in> G'" using assms ypair gi by auto
+ have ne_index_0: "x \<noteq> y \<Longrightarrow> \<B> index {x, y} = 0"
+ using ying xing index_together ging by simp
+ have "\<And> B. B \<in># \<B> \<Longrightarrow> (f B) index {(x, ix), (y, iy)} = 0"
+ proof -
+ fix B
+ assume assm: "B \<in># \<B>"
+ then interpret kgdd: K_GDD "(B \<times> I)" "(f B)" \<K>' "{{x} \<times> I |x . x \<in> B }" using assms by simp
+ have not_ss_0: "\<not> ({(x, ix), (y, iy)} \<subseteq> (B \<times> I)) \<Longrightarrow> (f B) index {(x, ix), (y, iy)} = 0"
+ by (metis kgdd.points_index_ps_nin)
+ have "x \<noteq> y \<Longrightarrow> \<not> {x, y} \<subseteq> B" using ne_index_0 assm points_index_0_left_imp by auto
+ then have "x \<noteq> y \<Longrightarrow> \<not> ({(x, ix), (y, iy)} \<subseteq> (B \<times> I))" using assms
+ by (meson empty_subsetI insert_subset mem_Sigma_iff)
+ then have nexy: "x \<noteq> y \<Longrightarrow> (f B) index {(x, ix), (y, iy)} = 0" using not_ss_0 by simp
+ have "x = y \<Longrightarrow> ({(x, ix), (y, iy)} \<subseteq> (B \<times> I)) \<Longrightarrow> (f B) index {(x, ix), (y, iy)} = 0"
+ proof -
+ assume eq: "x = y"
+ assume "({(x, ix), (y, iy)} \<subseteq> (B \<times> I))"
+ then obtain g where "g \<in> {{x} \<times> I |x . x \<in> B }" and "(x, ix) \<in> g" and "(y, ix) \<in> g"
+ using eq by auto
+ then show ?thesis using kgdd.index_together
+ by (smt (verit, best) SigmaD1 SigmaD2 SigmaI assms(4) assms(5) gi mem_Collect_eq xpair ypair)
+ qed
+ then show "(f B) index {(x, ix), (y, iy)} = 0" using not_ss_0 nexy by auto
+ qed
+ then have "\<And> B. B \<in># (image_mset f \<B>) \<Longrightarrow> B index {(x, ix), (y, iy)} = 0" by auto
+ then show "(\<Sum>\<^sub># (image_mset f \<B>)) index {X, Y} = 0"
+ by (simp add: points_index_sum xpair ypair)
+qed
+
+lemma wilsons_construction_index_1:
+ assumes "\<And> B . B \<in># \<B> \<Longrightarrow> K_GDD (B \<times> I) (f B) \<K>' {{x} \<times> I |x . x \<in> B }"
+ assumes "G1 \<in> {G \<times> I |G. G \<in> \<G>}"
+ assumes "G2 \<in> {G \<times> I |G. G \<in> \<G>}"
+ assumes "G1 \<noteq> G2"
+ and "(x, ix) \<in> G1" and "(y, iy) \<in> G2"
+ shows "(\<Sum>\<^sub># (image_mset f \<B>)) index {(x, ix), (y, iy)} = (1 ::int)"
+proof -
+ obtain G1' where gi1: "G1 = G1' \<times> I" and ging1: "G1' \<in> \<G>" using assms by auto
+ obtain G2' where gi2: "G2 = G2' \<times> I" and ging2: "G2' \<in> \<G>" using assms by auto
+ have xing: "x \<in> G1'" using assms gi1 by simp
+ have ying: "y \<in> G2'" using assms gi2 by simp
+ have gne: "G1' \<noteq> G2'" using assms gi1 gi2 by auto
+ then have xyne: "x \<noteq> y" using xing ying ging1 ging2 point_in_one_group by blast
+ have "\<exists>! bl . bl \<in># \<B> \<and> {x, y} \<subseteq> bl" using index_distinct points_index_one_unique_block
+ by (metis ging1 ging2 gne of_nat_1_eq_iff xing ying)
+ then obtain bl where blinb:"bl \<in># \<B>" and xyblss: "{x, y} \<subseteq> bl" by auto
+ then have "\<And> b . b \<in># \<B> - {#bl#} \<Longrightarrow> \<not> {x, y} \<subseteq> b" using points_index_one_not_unique_block
+ by (metis ging1 ging2 gne index_distinct int_ops(2) nat_int_comparison(1) xing ying)
+ then have not_ss: "\<And> b . b \<in># \<B> - {#bl#} \<Longrightarrow> \<not> ({(x, ix), (y, iy)} \<subseteq> (b \<times> I))" using assms
+ by (meson SigmaD1 empty_subsetI insert_subset)
+ then have pi0: "\<And> b . b \<in># \<B> - {#bl#} \<Longrightarrow> (f b) index {(x, ix), (y, iy)} = 0"
+ proof -
+ fix b
+ assume assm: "b \<in># \<B> - {#bl#}"
+ then have "b \<in># \<B>" by (meson in_diffD)
+ then interpret kgdd: K_GDD "(b \<times> I)" "(f b)" \<K>' "{{x} \<times> I |x . x \<in> b }" using assms by simp
+ show "(f b) index {(x, ix), (y, iy)} = 0"
+ using assm not_ss by (metis kgdd.points_index_ps_nin)
+ qed
+ let ?G = "{{x} \<times> I |x . x \<in> bl }"
+ interpret bkgdd: K_GDD "(bl \<times> I)" "(f bl)" \<K>' ?G using assms blinb by simp
+ obtain g1 g2 where xing1: "(x, ix) \<in> g1" and ying2: "(y, iy) \<in> g2" and g1g: "g1 \<in> ?G"
+ and g2g: "g2 \<in> ?G" using assms(5) assms(6) gi1 gi2
+ by (metis (no_types, lifting) bkgdd.point_has_unique_group insert_subset mem_Sigma_iff xyblss)
+ then have "g1 \<noteq> g2" using xyne by blast
+ then have pi1: "(f bl) index {(x, ix), (y, iy)} = 1"
+ using bkgdd.index_distinct xing1 ying2 g1g g2g by simp
+ have "(\<Sum>\<^sub># (image_mset f \<B>)) index {(x, ix), (y, iy)} =
+ (\<Sum>B \<in># \<B>. (f B) index {(x, ix), (y, iy)} )"
+ by (simp add: points_index_sum)
+ then have "(\<Sum>\<^sub># (image_mset f \<B>)) index {(x, ix), (y, iy)} =
+ (\<Sum>B \<in># (\<B> - {#bl#}). (f B) index {(x, ix), (y, iy)}) + (f bl) index {(x, ix), (y, iy)}"
+ by (metis (no_types, lifting) add.commute blinb insert_DiffM sum_mset.insert)
+ thus ?thesis using pi0 pi1 by simp
+qed
+
+theorem Wilsons_Construction:
+ assumes "card I = w"
+ assumes "w > 0"
+ assumes "\<And> n. n \<in> \<K>' \<Longrightarrow> n \<ge> 2"
+ assumes "\<And> B . B \<in># \<B> \<Longrightarrow> K_GDD (B \<times> I) (f B) \<K>' {{x} \<times> I |x . x \<in> B }"
+ shows "K_GDD (\<V> \<times> I) (\<Sum>B \<in># \<B>. (f B)) \<K>' {G \<times> I | G . G \<in> \<G>}"
+proof -
+ let ?Y = "\<V> \<times> I" and ?H = "{G \<times> I | G . G \<in> \<G>}" and ?B = "\<Sum>B \<in># \<B>. (f B)"
+ interpret pd: proper_design ?Y ?B using wilsons_construction_proper assms by auto
+ have "\<And> bl . bl \<in># (\<Sum>B \<in># \<B>. (f B)) \<Longrightarrow> card bl \<in> \<K>'"
+ using assms pair_construction_block_sizes by blast
+ then interpret kdes: K_block_design ?Y ?B \<K>'
+ using assms(3) by (unfold_locales) (simp_all,fastforce)
+ interpret gdd: GDD ?Y ?B ?H "1:: int"
+ proof (unfold_locales)
+ show "partition_on (\<V> \<times> I) {G \<times> I |G. G \<in> \<G>}"
+ using assms groups_not_empty design_points_nempty group_partitions
+ by (simp add: partition_on_cart_prod)
+ have "inj_on (\<lambda> G. G \<times> I) \<G>"
+ using inj_on_def pd.design_points_nempty by auto
+ then have "card {G \<times> I |G. G \<in> \<G>} = card \<G>" using card_image by (simp add: Setcompr_eq_image)
+ then show "1 < card {G \<times> I |G. G \<in> \<G>}" using groups_size by linarith
+ show "(1::int) \<le> 1" by simp
+ have gdd_fact: "\<And> B . B \<in># \<B> \<Longrightarrow> K_GDD (B \<times> I) (f B) \<K>' {{x} \<times> I |x . x \<in> B }"
+ using assms by simp
+ show "\<And>G X Y. G \<in> {GG \<times> I |GG. GG \<in> \<G>} \<Longrightarrow> X \<in> G \<Longrightarrow> Y \<in> G \<Longrightarrow> X \<noteq> Y
+ \<Longrightarrow> (\<Sum>\<^sub># (image_mset f \<B>)) index {X, Y} = 0"
+ using wilsons_construction_index_0[OF assms(4)] by auto
+ show "\<And>G1 G2 X Y. G1 \<in> {G \<times> I |G. G \<in> \<G>} \<Longrightarrow> G2 \<in> {G \<times> I |G. G \<in> \<G>}
+ \<Longrightarrow> G1 \<noteq> G2 \<Longrightarrow> X \<in> G1 \<Longrightarrow> Y \<in> G2 \<Longrightarrow> ((\<Sum>\<^sub># (image_mset f \<B>)) index {X, Y}) = (1 ::int)"
+ using wilsons_construction_index_1[OF assms(4)] by blast
+ qed
+ show ?thesis by (unfold_locales)
+qed
+
+end
+
+context pairwise_balance
+begin
+
+lemma PBD_by_deleting_point:
+ assumes "\<v> > 2"
+ assumes "\<And> bl . bl \<in># \<B> \<Longrightarrow> card bl \<ge> 2"
+ shows "pairwise_balance (del_point x) (del_point_blocks x) \<Lambda>"
+proof (cases "x \<in> \<V>")
+ case True
+ interpret des: design "del_point x" "del_point_blocks x"
+ using delete_point_design assms by blast
+ show ?thesis using assms design_blocks_nempty del_point_def del_point_blocks_def
+ proof (unfold_locales, simp_all)
+ show "2 < \<v> \<Longrightarrow> (\<And>bl. bl \<in># \<B> \<Longrightarrow> 2 \<le> card bl) \<Longrightarrow> 2 \<le> int (card (\<V> - {x}))"
+ using card_Diff_singleton_if diff_diff_cancel diff_le_mono2 finite_sets less_one
+ by (metis int_ops(3) less_nat_zero_code nat_le_linear verit_comp_simplify1(3) zle_int)
+ have "\<And> ps . ps \<subseteq> \<V> - {x} \<Longrightarrow> ps \<subseteq> \<V>" by auto
+ then show "\<And>ps. ps \<subseteq> \<V> - {x} \<Longrightarrow> card ps = 2 \<Longrightarrow> {#bl - {x}. bl \<in># \<B>#} index ps = \<Lambda>"
+ using delete_point_index_eq del_point_def del_point_blocks_def by simp
+ qed
+next
+ case False
+ then show ?thesis
+ by (simp add: del_invalid_point del_invalid_point_blocks pairwise_balance_axioms)
+qed
+end
+
+context k_GDD
+begin
+
+lemma bibd_from_kGDD:
+ assumes "\<k> > 1"
+ assumes "\<And> g. g \<in> \<G> \<Longrightarrow> card g = \<k> - 1"
+ assumes " x \<notin> \<V>"
+ shows "bibd (add_point x) (\<B> + mset_set { insert x g | g. g \<in> \<G>}) (\<k>) 1"
+proof -
+ have "({\<k>} \<union> {(card g) + 1 | g . g \<in> \<G>}) = {\<k>}"
+ using assms(2) by fastforce
+ then interpret pbd: PBD "(add_point x)" "\<B> + mset_set { insert x g | g. g \<in> \<G>}" "{\<k>}"
+ using PBD_by_adjoining_point assms sys_block_sizes_obtain_bl add_point_def
+ by (smt (verit, best) Collect_cong sys_block_sizes_uniform uniform_alt_def_all)
+ show ?thesis using assms pbd.block_sizes block_size_lt_v finite_sets add_point_def
+ by (unfold_locales) (simp_all)
+qed
+
+end
+
+context PBD
+begin
+
+lemma pbd_points_index1: "ps \<subseteq> \<V> \<Longrightarrow> card ps = 2 \<Longrightarrow> \<B> index ps = 1"
+ using balanced by (metis int_eq_iff_numeral of_nat_1_eq_iff)
+
+lemma pbd_index1_points_imply_unique_block:
+ assumes "b1 \<in># \<B>" and "b2 \<in># \<B>" and "b1 \<noteq> b2"
+ assumes "x \<noteq> y" and "{x, y} \<subseteq> b1" and "x \<in> b2"
+ shows "y \<notin> b2"
+proof (rule ccontr)
+ let ?ps = "{# b \<in># \<B> . {x, y} \<subseteq> b#}"
+ assume "\<not> y \<notin> b2"
+ then have a: "y \<in> b2" by linarith
+ then have "{x, y} \<subseteq> b2"
+ by (simp add: assms(6))
+ then have "b1 \<in># ?ps" and "b2 \<in># ?ps" using assms by auto
+ then have ss: "{#b1, b2#} \<subseteq># ?ps" using assms
+ by (metis insert_noteq_member mset_add mset_subset_eq_add_mset_cancel single_subset_iff)
+ have "size {#b1, b2#} = 2" using assms by auto
+ then have ge2: "size ?ps \<ge> 2" using assms ss by (metis size_mset_mono)
+ have pair: "card {x, y} = 2" using assms by auto
+ have "{x, y} \<subseteq> \<V>" using assms wellformed by auto
+ then have "\<B> index {x, y} = 1" using pbd_points_index1 pair by simp
+ then show False using points_index_def ge2
+ by (metis numeral_le_one_iff semiring_norm(69))
+qed
+
+lemma strong_delete_point_groups_index_zero:
+ assumes "G \<in> {b - {x} |b. b \<in># \<B> \<and> x \<in> b}"
+ assumes "xa \<in> G" and "y \<in> G" and "xa \<noteq> y"
+ shows "(str_del_point_blocks x) index {xa, y} = 0"
+proof (auto simp add: points_index_0_iff str_del_point_blocks_def)
+ fix b
+ assume a1: "b \<in># \<B>" and a2: "x \<notin> b" and a3: "xa \<in> b" and a4: "y \<in> b"
+ obtain b' where "G = b' - {x}" and "b' \<in># \<B>" and "x \<in> b'" using assms by blast
+ then show False using a1 a2 a3 a4 assms pbd_index1_points_imply_unique_block
+ by fastforce
+qed
+
+lemma strong_delete_point_groups_index_one:
+ assumes "G1 \<in> {b - {x} |b. b \<in># \<B> \<and> x \<in> b}"
+ assumes "G2 \<in> {b - {x} |b. b \<in># \<B> \<and> x \<in> b}"
+ assumes "G1 \<noteq> G2" and "xa \<in> G1" and "y \<in> G2"
+ shows "(str_del_point_blocks x) index {xa, y} = 1"
+proof -
+ obtain b1 where gb1: "G1 = b1 - {x}" and b1in: "b1 \<in># \<B>" and xin1: "x \<in> b1" using assms by blast
+ obtain b2 where gb2: "G2 = b2 - {x}" and b2in: "b2 \<in># \<B>" and xin2:"x \<in> b2" using assms by blast
+ have bneq: "b1 \<noteq> b2 " using assms(3) gb1 gb2 by auto
+ have "xa \<noteq> y" using gb1 b1in xin1 gb2 b2in xin2 assms(3) assms(4) assms(5) insert_subset
+ by (smt (verit, best) Diff_eq_empty_iff Diff_iff empty_Diff insertCI pbd_index1_points_imply_unique_block)
+ then have pair: "card {xa, y} = 2" by simp
+ have inv: "{xa, y} \<subseteq> \<V>" using gb1 b1in gb2 b2in assms(4) assms(5)
+ by (metis Diff_cancel Diff_subset insert_Diff insert_subset wellformed)
+ have "{# bl \<in># \<B> . x \<in> bl#} index {xa, y} = 0"
+ proof (auto simp add: points_index_0_iff)
+ fix b assume a1: "b \<in># \<B>" and a2: "x \<in> b" and a3: "xa \<in> b" and a4: "y \<in> b"
+ then have yxss: "{y, x} \<subseteq> b2"
+ using assms(5) gb2 xin2 by blast
+ have "{xa, x} \<subseteq> b1"
+ using assms(4) gb1 xin1 by auto
+ then have "xa \<notin> b2" using pbd_index1_points_imply_unique_block
+ by (metis DiffE assms(4) b1in b2in bneq gb1 singletonI xin2)
+ then have "b2 \<noteq> b" using a3 by auto
+ then show False using pbd_index1_points_imply_unique_block
+ by (metis DiffD2 yxss a1 a2 a4 assms(5) b2in gb2 insertI1)
+ qed
+ then have "(str_del_point_blocks x) index {xa, y} = \<B> index {xa, y}"
+ by (metis multiset_partition plus_nat.add_0 point_index_distrib str_del_point_blocks_def)
+ thus ?thesis using pbd_points_index1 pair inv by fastforce
+qed
+
+lemma blocks_with_x_partition:
+ assumes "x \<in> \<V>"
+ shows "partition_on (\<V> - {x}) {b - {x} |b. b \<in># \<B> \<and> x \<in> b}"
+proof (intro partition_onI )
+ have gtt: "\<And> bl. bl \<in># \<B> \<Longrightarrow> card bl \<ge> 2" using block_size_gt_t
+ by (simp add: block_sizes nat_int_comparison(3))
+ show "\<And>p. p \<in> {b - {x} |b. b \<in># \<B> \<and> x \<in> b} \<Longrightarrow> p \<noteq> {}"
+ proof -
+ fix p assume "p \<in> {b - {x} |b. b \<in># \<B> \<and> x \<in> b}"
+ then obtain b where ptx: "p = b - {x}" and "b \<in># \<B>" and xinb: "x \<in> b" by blast
+ then have ge2: "card b \<ge> 2" using gtt by (simp add: nat_int_comparison(3))
+ then have "finite b" by (metis card.infinite not_numeral_le_zero)
+ then have "card p = card b - 1" using xinb ptx by simp
+ then have "card p \<ge> 1" using ge2 by linarith
+ thus "p \<noteq> {}" by auto
+ qed
+ show "\<Union> {b - {x} |b. b \<in># \<B> \<and> x \<in> b} = \<V> - {x}"
+ proof (intro subset_antisym subsetI)
+ fix xa
+ assume "xa \<in> \<Union> {b - {x} |b. b \<in># \<B> \<and> x \<in> b}"
+ then obtain b where "xa \<in> b" and "b \<in># \<B>" and "x \<in> b" and "xa \<noteq> x" by auto
+ then show "xa \<in> \<V> - {x}" using wf_invalid_point by blast
+ next
+ fix xa
+ assume a: "xa \<in> \<V> - {x}"
+ then have nex: "xa \<noteq> x" by simp
+ then have pair: "card {xa, x} = 2" by simp
+ have "{xa, x} \<subseteq> \<V>" using a assms by auto
+ then have "card {b \<in> design_support . {xa, x} \<subseteq> b} = 1"
+ using balanced points_index_simple_def pbd_points_index1 assms by (metis pair)
+ then obtain b where des: "b \<in> design_support" and ss: "{xa, x} \<subseteq> b"
+ by (metis (no_types, lifting) card_1_singletonE mem_Collect_eq singletonI)
+ then show "xa \<in> \<Union> {b - {x} |b. b \<in># \<B> \<and> x \<in> b}"
+ using des ss nex design_support_def by auto
+ qed
+ show "\<And>p p'. p \<in> {b - {x} |b. b \<in># \<B> \<and> x \<in> b} \<Longrightarrow> p' \<in> {b - {x} |b. b \<in># \<B> \<and> x \<in> b} \<Longrightarrow>
+ p \<noteq> p' \<Longrightarrow> p \<inter> p' = {}"
+ proof -
+ fix p p'
+ assume p1: "p \<in> {b - {x} |b. b \<in># \<B> \<and> x \<in> b}" and p2: "p' \<in> {b - {x} |b. b \<in># \<B> \<and> x \<in> b}"
+ and pne: "p \<noteq> p'"
+ then obtain b where b1: "p = b - {x}" and b1in:"b \<in># \<B>" and xinb1:"x \<in> b" by blast
+ then obtain b' where b2: "p' = b' - {x}" and b2in: "b' \<in># \<B>" and xinb2: "x \<in> b'"
+ using p2 by blast
+ then have "b \<noteq> b'" using pne b1 by auto
+ then have "\<And> y. y \<in> b \<Longrightarrow> y \<noteq> x \<Longrightarrow> y \<notin> b'"
+ using b1in b2in xinb1 xinb2 pbd_index1_points_imply_unique_block
+ by (meson empty_subsetI insert_subset)
+ then have "\<And> y. y \<in> p \<Longrightarrow> y \<notin> p'"
+ by (metis Diff_iff b1 b2 insertI1)
+ then show "p \<inter> p' = {}" using disjoint_iff by auto
+ qed
+qed
+
+lemma KGDD_by_deleting_point:
+ assumes "x \<in> \<V>"
+ assumes "\<B> rep x < \<b>"
+ assumes "\<B> rep x > 1"
+ shows "K_GDD (del_point x) (str_del_point_blocks x) \<K> { b - {x} | b . b \<in># \<B> \<and> x \<in> b}"
+proof -
+ have "\<And> bl. bl \<in># \<B> \<Longrightarrow> card bl \<ge> 2" using block_size_gt_t
+ by (simp add: block_sizes nat_int_comparison(3))
+ then interpret des: proper_design "(del_point x)" "(str_del_point_blocks x)"
+ using strong_delete_point_proper assms by blast
+ show ?thesis using blocks_with_x_partition strong_delete_point_groups_index_zero
+ strong_delete_point_groups_index_one str_del_point_blocks_def del_point_def
+ proof (unfold_locales, simp_all add: block_sizes positive_ints assms)
+ have ge1: "card {b . b \<in># \<B> \<and> x \<in> b} > 1"
+ using assms(3) replication_num_simple_def design_support_def by auto
+ have fin: "finite {b . b \<in># \<B> \<and> x \<in> b}" by simp
+ have inj: "inj_on (\<lambda> b . b - {x}) {b . b \<in># \<B> \<and> x \<in> b}"
+ using assms(2) inj_on_def mem_Collect_eq by auto
+ then have "card {b - {x} |b. b \<in># \<B> \<and> x \<in> b} = card {b . b \<in># \<B> \<and> x \<in> b}"
+ using card_image fin by (simp add: inj card_image setcompr_eq_image)
+ then show "Suc 0 < card {b - {x} |b. b \<in># \<B> \<and> x \<in> b}" using ge1
+ by presburger
+ qed
+qed
+
+lemma card_singletons_eq: "card {{a} | a . a \<in> A} = card A"
+ by (simp add: card_image Setcompr_eq_image)
+
+lemma KGDD_from_PBD: "K_GDD \<V> \<B> \<K> {{x} | x . x \<in> \<V>}"
+proof (unfold_locales,auto simp add: Setcompr_eq_image partition_on_singletons)
+ have "card ((\<lambda>x. {x}) ` \<V>) \<ge> 2" using t_lt_order card_singletons_eq
+ by (metis Collect_mem_eq nat_leq_as_int of_nat_numeral setcompr_eq_image)
+ then show "Suc 0 < card ((\<lambda>x. {x}) ` \<V>)" by linarith
+ show "\<And>xa xb. xa \<in> \<V> \<Longrightarrow> xb \<in> \<V> \<Longrightarrow> \<B> index {xa, xb} \<noteq> Suc 0 \<Longrightarrow> xa = xb"
+ proof (rule ccontr)
+ fix xa xb
+ assume ain: "xa \<in> \<V>" and bin: "xb \<in> \<V>" and ne1: "\<B> index {xa, xb} \<noteq> Suc 0"
+ assume "xa \<noteq> xb"
+ then have "card {xa, xb} = 2" by auto
+ then have "\<B> index {xa, xb} = 1"
+ by (simp add: ain bin pbd_points_index1)
+ thus False using ne1 by linarith
+ qed
+qed
+
+end
+
+context bibd
+begin
+lemma kGDD_from_bibd:
+ assumes "\<Lambda> = 1"
+ assumes "x \<in> \<V>"
+ shows "k_GDD (del_point x) (str_del_point_blocks x) \<k> { b - {x} | b . b \<in># \<B> \<and> x \<in> b}"
+proof -
+ interpret pbd: PBD \<V> \<B> "{\<k>}" using assms
+ using PBD.intro \<Lambda>_PBD_axioms by auto
+ have lt: "\<B> rep x < \<b>" using block_num_gt_rep
+ by (simp add: assms(2))
+ have "\<B> rep x > 1" using r_ge_two assms by simp
+ then interpret kgdd: K_GDD "(del_point x)" "str_del_point_blocks x"
+ "{\<k>}" "{ b - {x} | b . b \<in># \<B> \<and> x \<in> b}"
+ using pbd.KGDD_by_deleting_point lt assms by blast
+ show ?thesis using del_point_def str_del_point_blocks_def by (unfold_locales) (simp_all)
+qed
+
+end
end
\ No newline at end of file
diff --git a/thys/Design_Theory/Multisets_Extras.thy b/thys/Design_Theory/Multisets_Extras.thy
--- a/thys/Design_Theory/Multisets_Extras.thy
+++ b/thys/Design_Theory/Multisets_Extras.thy
@@ -1,796 +1,796 @@
-(* Title: Multisets_Extras
- Author: Chelsea Edmonds
-*)
-
-section \<open>Micellanious Helper Functions on Sets and Multisets\<close>
-
-theory Multisets_Extras imports
- "HOL-Library.Multiset"
- Card_Partitions.Set_Partition
- Nested_Multisets_Ordinals.Multiset_More
- Nested_Multisets_Ordinals.Duplicate_Free_Multiset
- "HOL-Library.Disjoint_Sets"
-begin
-
-subsection \<open>Set Theory Extras\<close>
-
-text \<open>A number of extra helper lemmas for reasoning on sets (finite) required for Design Theory
-proofs\<close>
-
-lemma card_Pow_filter_one:
- assumes "finite A"
- shows "card {x \<in> Pow A . card x = 1} = card (A)"
- using assms
-proof (induct rule: finite_induct)
- case empty
- then show ?case by auto
-next
- case (insert x F)
- have "Pow (insert x F) = Pow F \<union> insert x ` Pow F"
- by (simp add: Pow_insert)
- then have split: "{y \<in> Pow (insert x F) . card y = 1} =
- {y \<in> (Pow F) . card y = 1} \<union> {y \<in> (insert x ` Pow F) . card y = 1}"
- by blast
- have "\<And> y . y \<in> (insert x ` Pow F) \<Longrightarrow> finite y"
- using finite_subset insert.hyps(1) by fastforce
- then have single: "\<And> y . y \<in> (insert x ` Pow F) \<Longrightarrow> card y = 1 \<Longrightarrow> y = {x}"
- by (metis card_1_singletonE empty_iff image_iff insertCI insertE)
- then have "card {y \<in> (insert x ` Pow F) . card y = 1} = 1"
- using empty_iff imageI is_singletonI is_singletonI' is_singleton_altdef (* LONG *)
- by (metis (full_types, lifting) Collect_empty_eq_bot Pow_bottom bot_empty_eq mem_Collect_eq)
- then have " {y \<in> (insert x ` Pow F) . card y = 1} = {{x}}"
- using single card_1_singletonE card_eq_0_iff
- by (smt empty_Collect_eq mem_Collect_eq singletonD zero_neq_one)
- then have split2:"{y \<in> Pow (insert x F) . card y = 1} = {y \<in> (Pow F) . card y = 1} \<union> {{x}}"
- using split by simp
- then show ?case
- proof (cases "x \<in> F")
- case True
- then show ?thesis using insert.hyps(2) by auto
- next
- case False
- then have "{y \<in> (Pow F) . card y = 1} \<inter> {{x}} = {}" by blast
- then have fact:"card {y \<in> Pow (insert x F) . card y = 1} =
- card {y \<in> (Pow F) . card y = 1} + card {{x}}"
- using split2 card_Un_disjoint insert.hyps(1) by auto
- have "card (insert x F) = card F + 1"
- using False card_insert_disjoint by (metis Suc_eq_plus1 insert.hyps(1))
- then show ?thesis using fact insert.hyps(3) by auto
- qed
-qed
-
-lemma elem_exists_non_empty_set:
- assumes "card A > 0"
- obtains x where "x \<in> A"
- using assms card_gt_0_iff by fastforce
-
-lemma set_self_img_compr: "{a | a . a \<in> A} = A"
- by blast
-
-lemma card_subset_not_gt_card: "finite A \<Longrightarrow> card ps > card A \<Longrightarrow> \<not> (ps \<subseteq> A)"
- using card_mono leD by auto
-
-lemma card_inter_lt_single: "finite A \<Longrightarrow> finite B \<Longrightarrow> card (A \<inter> B) \<le> card A"
- by (simp add: card_mono)
-
-lemma set_diff_non_empty_not_subset:
- assumes "A \<subseteq> (B - C)"
- assumes "C \<noteq> {}"
- assumes "A \<noteq> {}"
- assumes "B \<noteq> {}"
- shows " \<not> (A \<subseteq> C)"
-proof (rule ccontr)
- assume " \<not> \<not> (A \<subseteq> C)"
- then have a: "\<And> x . x \<in> A \<Longrightarrow> x \<in> C" by blast
- thus False using a assms by blast
-qed
-
-lemma set_card_diff_ge_zero: "finite A \<Longrightarrow> finite B \<Longrightarrow> A \<noteq> B \<Longrightarrow> card A = card B \<Longrightarrow>
- card (A - B) > 0"
- by (meson Diff_eq_empty_iff card_0_eq card_subset_eq finite_Diff neq0_conv)
-
-lemma set_filter_diff: "{a \<in> A . P a } - {x} = {a \<in> (A - {x}) . (P a )}"
- by (auto)
-
-lemma set_filter_diff_card: "card ({a \<in> A . P a } - {x}) = card {a \<in> (A - {x}) . (P a )}"
- by (simp add: set_filter_diff)
-
-lemma obtain_subset_with_card_int_n:
- assumes "(n ::int) \<le> of_nat (card S)"
- assumes "(n ::int) \<ge> 0"
- obtains T where "T \<subseteq> S" "of_nat (card T) = (n ::int)" "finite T"
- using obtain_subset_with_card_n assms
- by (metis nonneg_int_cases of_nat_le_iff)
-
-lemma transform_filter_img_empty_rm:
- assumes "\<And> g . g \<in> G \<Longrightarrow> g \<noteq> {}"
- shows "{g - {x} | g. g \<in> G \<and> g \<noteq> {x}} = {g - {x} | g. g \<in> G } - {{}}"
-proof -
- let ?f = "\<lambda> g . g - {x}"
- have "\<And> g . g \<in> G \<Longrightarrow> g \<noteq> {x} \<longleftrightarrow> ?f g \<noteq> {}" using assms
- by (metis Diff_cancel Diff_empty Diff_insert0 insert_Diff)
- thus ?thesis by auto
-qed
-
-lemma bij_betw_inter_subsets: "bij_betw f A B \<Longrightarrow> a1 \<subseteq> A \<Longrightarrow> a2 \<subseteq> A
- \<Longrightarrow> f ` (a1 \<inter> a2) = (f ` a1) \<inter> (f ` a2)"
- by (meson bij_betw_imp_inj_on inj_on_image_Int)
-
-text\<open>Partition related set theory lemmas\<close>
-
-lemma partition_on_remove_pt:
- assumes "partition_on A G"
- shows "partition_on (A - {x}) {g - {x} | g. g \<in> G \<and> g \<noteq> {x}}"
-proof (intro partition_onI)
- show "\<And>p. p \<in> {g - {x} |g. g \<in> G \<and> g \<noteq> {x}} \<Longrightarrow> p \<noteq> {}"
- using assms partition_onD3 subset_singletonD by force
- let ?f = "(\<lambda> g . g - {x})"
- have un_img: "\<Union>({?f g | g. g \<in> G }) = ?f (\<Union> G)" by blast
- have empty: "\<Union> {g - {x} |g. g \<in> G \<and> g \<noteq> {x}} = \<Union>({g - {x} | g. g \<in> G } - {{}})"
- by blast
- then have "\<Union>({g - {x} | g. g \<in> G } - {{}}) = \<Union>({g - {x} | g. g \<in> G })" by blast
- then show " \<Union> {g - {x} |g. g \<in> G \<and> g \<noteq> {x}} = A - {x}" using partition_onD1 assms un_img
- by (metis empty)
- then show "\<And>p p'.
- p \<in> {g - {x} |g. g \<in> G \<and> g \<noteq> {x}} \<Longrightarrow>
- p' \<in> {g - {x} |g. g \<in> G \<and> g \<noteq> {x}} \<Longrightarrow> p \<noteq> p' \<Longrightarrow> p \<inter> p' = {}"
- proof -
- fix p1 p2
- assume p1: "p1 \<in> {g - {x} |g. g \<in> G \<and> g \<noteq> {x}}"
- and p2: "p2 \<in> {g - {x} |g. g \<in> G \<and> g \<noteq> {x}}"
- and ne: "p1 \<noteq> p2"
- obtain p1' p2' where orig1: "p1 = p1' - {x}" and orig2: "p2 = p2' - {x}"
- and origne: "p1' \<noteq> p2'" and ne1: "p1' \<noteq> {x}" and ne2:"p2' \<noteq> {x}" and ing1: "p1' \<in> G"
- and ing2: "p2' \<in> G"
- using p1 p2 using mem_Collect_eq ne by blast
- then have "p1' \<inter> p2' = {}" using assms partition_onD2 ing1 ing2 origne disjointD by blast
- thus "p1 \<inter> p2 = {}" using orig1 orig2 by blast
- qed
-qed
-
-lemma partition_on_cart_prod:
- assumes "card I > 0"
- assumes "A \<noteq> {}"
- assumes "G \<noteq> {}"
- assumes "partition_on A G"
- shows "partition_on (A \<times> I) {g \<times> I |g. g \<in> G}"
-proof (intro partition_onI)
- show "\<And>p. p \<in> {g \<times> I |g. g \<in> G} \<Longrightarrow> p \<noteq> {}"
- using assms(1) assms(4) partition_onD3 by fastforce
- show "\<Union> {g \<times> I |g. g \<in> G} = A \<times> I"
- by (metis Setcompr_eq_image Sigma_Union assms(4) partition_onD1)
- show "\<And>p p'. p \<in> {g \<times> I |g. g \<in> G} \<Longrightarrow> p' \<in> {g \<times> I |g. g \<in> G} \<Longrightarrow> p \<noteq> p' \<Longrightarrow> p \<inter> p' = {}"
- by (smt (verit, best) Sigma_Int_distrib1 Sigma_empty1 assms(4) mem_Collect_eq partition_onE)
-qed
-
-subsection \<open>Multiset Helpers\<close>
-
-text \<open>Generic Size, count and card helpers\<close>
-
-lemma count_size_set_repr: "size {# x \<in># A . x = g#} = count A g"
- by (simp add: filter_eq_replicate_mset)
-
-lemma mset_nempty_set_nempty: "A \<noteq> {#} \<longleftrightarrow> (set_mset A) \<noteq> {}"
- by simp
-
-lemma mset_size_ne0_set_card: "size A > 0 \<Longrightarrow> card (set_mset A) > 0"
- using mset_nempty_set_nempty by fastforce
-
-lemma set_count_size_min: "count A a \<ge> n \<Longrightarrow> size A \<ge> n"
- by (metis (full_types) count_le_replicate_mset_subset_eq size_mset_mono size_replicate_mset)
-
-lemma card_size_filter_eq: "finite A \<Longrightarrow> card {a \<in> A . P a} = size {#a \<in># mset_set A . P a#}"
- by simp
-
-lemma size_multiset_int_count:
- assumes "of_nat (card (set_mset A)) = (ca :: int)"
- assumes "\<And>p. p \<in># A \<Longrightarrow> of_nat (count A p) = (ca2 :: int)"
- shows "of_nat (size A) = ((ca :: int) * ca2)"
-proof -
- have "size A = (\<Sum> p \<in> (set_mset A) . count A p)" using size_multiset_overloaded_eq by auto
- then have "of_nat (size A) = (\<Sum> p \<in> (set_mset A) . ca2)" using assms by simp
- thus ?thesis using assms(1) by auto
-qed
-
-lemma mset_union_size: "size (A \<union># B) = size (A) + size (B - A)"
- by (simp add: union_mset_def)
-
-lemma mset_union_size_inter: "size (A \<union># B) = size (A) + size B - size (A \<inter># B)"
- by (metis diff_add_inverse2 size_Un_Int)
-
-text \<open>Lemmas for repeat\_mset\<close>
-
-lemma repeat_mset_size [simp]: "size (repeat_mset n A) = n * size A"
- by (induction n) auto
-
-lemma repeat_mset_subset_in:
- assumes "\<And> a . a \<in># A \<Longrightarrow> a \<subseteq> B"
- assumes "X \<in># repeat_mset n A"
- assumes "x \<in> X"
- shows " x \<in> B"
- using assms by (induction n) auto
-
-lemma repeat_mset_not_empty: "n > 0 \<Longrightarrow> A \<noteq> {#} \<Longrightarrow> repeat_mset n A \<noteq> {#}"
- by (induction n) auto
-
-lemma elem_in_repeat_in_original: "a \<in># repeat_mset n A \<Longrightarrow> a \<in># A"
- by (metis count_inI count_repeat_mset in_countE mult.commute mult_zero_left nat.distinct(1))
-
-lemma elem_in_original_in_repeat: "n > 0 \<Longrightarrow> a \<in># A \<Longrightarrow> a \<in># repeat_mset n A"
- by (metis count_greater_zero_iff count_repeat_mset nat_0_less_mult_iff)
-
-text \<open>Lemmas on image and filter for multisets\<close>
-
-lemma multiset_add_filter_size: "size {# a \<in># (A1 + A2) . P a #} = size {# a \<in># A1 . P a #} +
- size {# a \<in># A2 . P a #}"
- by simp
-
-lemma size_filter_neg: "size {#a \<in># A . P a #} = size A - size {# a \<in># A . \<not> P a #}"
- using size_filter_mset_lesseq size_union union_filter_mset_complement
- by (metis ordered_cancel_comm_monoid_diff_class.le_imp_diff_is_add)
-
-lemma filter_filter_mset_cond_simp:
- assumes "\<And> a . P a \<Longrightarrow> Q a"
- shows "filter_mset P A = filter_mset P (filter_mset Q A)"
-proof -
- have "filter_mset P (filter_mset Q A) = filter_mset (\<lambda> a. Q a \<and> P a) A"
- by (simp add: filter_filter_mset)
- thus ?thesis using assms
- by (metis (mono_tags, lifting) filter_mset_cong)
-qed
-
-lemma filter_filter_mset_ss_member: "filter_mset (\<lambda> a . {x, y} \<subseteq> a) A =
- filter_mset (\<lambda> a . {x, y} \<subseteq> a) (filter_mset (\<lambda> a . x \<in> a) A)"
-proof -
- have filter: "filter_mset (\<lambda> a . {x, y} \<subseteq> a) (filter_mset (\<lambda> a . x \<in> a) A) =
- filter_mset (\<lambda> a . x \<in> a \<and> {x, y} \<subseteq> a) A" by (simp add: filter_filter_mset)
- have "\<And> a. {x, y} \<subseteq> a \<Longrightarrow> x \<in> a" by simp
- thus ?thesis using filter by auto
-qed
-
-lemma multiset_image_do_nothing: "(\<And> x .x \<in># A \<Longrightarrow> f x = x) \<Longrightarrow> image_mset f A = A"
- by (induct A) auto
-
-lemma set_mset_filter: "set_mset {# f a . a \<in># A #} = {f a | a. a \<in># A}"
- by (simp add: Setcompr_eq_image)
-
-lemma mset_exists_imply: "x \<in># {# f a . a \<in># A #} \<Longrightarrow> \<exists> y \<in># A . x = f y"
- by auto
-
-lemma filter_mset_image_mset:
- "filter_mset P (image_mset f A) = image_mset f (filter_mset (\<lambda>x. P (f x)) A)"
- by (induction A) auto
-
-lemma mset_bunion_filter: "{# a \<in># A . P a \<or> Q a #} = {# a \<in># A . P a #} \<union># {# a \<in># A . Q a #}"
- by (rule multiset_eqI) simp
-
-lemma mset_inter_filter: "{# a \<in># A . P a \<and> Q a #} = {# a \<in># A . P a #} \<inter># {# a \<in># A . Q a #}"
- by (rule multiset_eqI) simp
-
-lemma image_image_mset: "image_mset (\<lambda> x . f x) (image_mset (\<lambda> y . g y) A) =
- image_mset (\<lambda> x. f (g x)) A"
- by simp
-
-text \<open>Big Union over multiset helpers\<close>
-
-lemma mset_big_union_obtain:
- assumes "x \<in># \<Sum>\<^sub># A"
- obtains a where "a \<in># A" and "x \<in># a"
- using assms by blast
-
-lemma size_big_union_sum: "size (\<Sum>\<^sub># (M :: 'a multiset multiset)) = (\<Sum>x \<in>#M . size x)"
- by (induct M) auto
-
-text \<open>Cartesian Product on Multisets\<close>
-
-lemma size_cartesian_product_singleton [simp]: "size ({#a#} \<times># B) = size B"
- by (simp add: Times_mset_single_left)
-
-lemma size_cartesian_product_singleton_right [simp]: "size (A \<times># {#b#}) = size A"
- by (simp add: Times_mset_single_right)
-
-lemma size_cartesian_product_empty [simp]: "size (A \<times># {#}) = 0"
- by simp
-
-lemma size_add_elem_step_eq:
- assumes "size (A \<times># B) = size A * size B"
- shows "size (add_mset x A \<times># B) = size (add_mset x A) * size B"
-proof -
- have "(add_mset x A \<times># B) = A \<times># B + {#x#} \<times># B"
- by (metis Sigma_mset_plus_distrib1 add_mset_add_single)
- then have "size (add_mset x A \<times># B) = size (A \<times># B) + size B" by auto
- also have "... = size A * size B + size B"
- by (simp add: assms)
- finally have "size (add_mset x A \<times># B) = (size A + 1) * size B"
- by auto
- thus ?thesis by simp
-qed
-
-lemma size_cartesian_product: "size (A \<times># B) = size A * size B"
- by (induct A) (simp_all add: size_add_elem_step_eq)
-
-lemma cart_prod_distinct_mset:
- assumes assm1: "distinct_mset A"
- assumes assm2: "distinct_mset B"
- shows "distinct_mset (A \<times># B)"
- unfolding distinct_mset_count_less_1
-proof (rule allI)
- fix x
- have count_mult: "count (A \<times># B) x = count A (fst x) * count B (snd x)"
- using count_Sigma_mset by (metis prod.exhaust_sel)
- then have "count A (fst x) * count B (snd x) \<le> 1" using assm1 assm2
- unfolding distinct_mset_count_less_1 using mult_le_one by blast
- thus "count (A \<times># B) x \<le> 1" using count_mult by simp
-qed
-
-lemma cart_product_single_intersect: "x1 \<noteq> x2 \<Longrightarrow> ({#x1#} \<times># A) \<inter># ({#x2#} \<times># B) = {#}"
- using multiset_inter_single by fastforce
-
-lemma size_union_distinct_cart_prod: "x1 \<noteq> x2 \<Longrightarrow> size (({#x1#} \<times># A) \<union># ({#x2#} \<times># B)) =
- size ({#x1#} \<times># A) + size ({#x2#} \<times># B)"
- by (simp add: cart_product_single_intersect size_Un_disjoint)
-
-lemma size_Union_distinct_cart_prod: "distinct_mset M \<Longrightarrow>
- size (\<Sum>p\<in>#M. ({#p#} \<times># B)) = size (M) * size (B)"
- by (induction M) auto
-
-lemma size_Union_distinct_cart_prod_filter: "distinct_mset M \<Longrightarrow>
- (\<And> p . p \<in># M \<Longrightarrow> size ({# b \<in># B . P p b #}) = c) \<Longrightarrow>
- size (\<Sum>p\<in>#M. ({#p#} \<times># {# b \<in># B . P p b #})) = size (M) * c"
- by (induction M) auto
-
-lemma size_Union_distinct_cart_prod_filter2: "distinct_mset V \<Longrightarrow>
- (\<And> b . b \<in># B \<Longrightarrow> size ({# v \<in># V . P v b #}) = c) \<Longrightarrow>
- size (\<Sum>b\<in>#B. ( {# v \<in># V . P v b #} \<times># {#b#})) = size (B) * c"
- by (induction B) auto
-
-lemma cart_product_add_1: "(add_mset a A) \<times># B = ({#a#} \<times># B) + (A \<times># B)"
- by (metis Sigma_mset_plus_distrib1 add_mset_add_single union_commute)
-
-lemma cart_product_add_1_filter: "{#m \<in># ((add_mset a M) \<times># N) . P m #} =
- {#m \<in># (M \<times># N) . P m #} + {#m \<in># ({#a#} \<times># N) . P m #}"
- unfolding add_mset_add_single [of a M] Sigma_mset_plus_distrib1
- by (simp add: Times_mset_single_left)
-
-lemma cart_product_add_1_filter2: "{#m \<in># (M \<times># (add_mset b N)) . P m #} =
- {#m \<in># (M \<times># N) . P m #} + {#m \<in># (M \<times># {#b#}) . P m #}"
- unfolding add_mset_add_single [of b N] Sigma_mset_plus_distrib1
- by (metis Times_insert_left Times_mset_single_right add_mset_add_single filter_union_mset)
-
-lemma cart_prod_singleton_right_gen:
- assumes "\<And> x . x \<in># (A \<times># {#b#}) \<Longrightarrow> P x \<longleftrightarrow> Q (fst x)"
- shows "{#x \<in># (A \<times># {#b#}). P x#} = {# a \<in># A . Q a#} \<times># {#b#}"
- using assms
-proof (induction A)
- case empty
- then show ?case by simp
-next
- case (add x A)
- have "add_mset x A \<times># {#b#} = add_mset (x, b) (A \<times># {#b#})"
- by (simp add: Times_mset_single_right)
- then have lhs: "filter_mset P (add_mset x A \<times># {#b#}) = filter_mset P (A \<times># {#b#}) +
- filter_mset P {#(x, b)#}" by simp
- have rhs: "filter_mset Q (add_mset x A) \<times># {#b#} = filter_mset Q A \<times># {#b#} +
- filter_mset Q {#x#} \<times># {#b#}"
- by (metis Sigma_mset_plus_distrib1 add_mset_add_single filter_union_mset)
- have "filter_mset P {#(x, b)#} = filter_mset Q {#x#} \<times># {#b#}"
- using add.prems by fastforce
- then show ?case using lhs rhs add.IH add.prems by force
-qed
-
-lemma cart_prod_singleton_left_gen:
- assumes "\<And> x . x \<in># ({#a#} \<times># B) \<Longrightarrow> P x \<longleftrightarrow> Q (snd x)"
- shows "{#x \<in># ({#a#} \<times># B). P x#} = {#a#} \<times># {#b \<in># B . Q b#}"
- using assms
-proof (induction B)
- case empty
- then show ?case by simp
-next
- case (add x B)
- have lhs: "filter_mset P ({#a#} \<times># add_mset x B) = filter_mset P ({#a#} \<times># B) +
- filter_mset P {#(a, x)#}"
- by (simp add: cart_product_add_1_filter2)
- have rhs: "{#a#} \<times># filter_mset Q (add_mset x B) = {#a#} \<times># filter_mset Q B +
- {#a#} \<times># filter_mset Q {#x#}"
- using add_mset_add_single filter_union_mset by (metis Times_mset_single_left image_mset_union)
- have "filter_mset P {#(a, x)#} = {#a#} \<times># filter_mset Q {#x#}"
- using add.prems by fastforce
- then show ?case using lhs rhs add.IH add.prems by force
-qed
-
-lemma cart_product_singleton_left: "{#m \<in># ({#a#} \<times># N) . fst m \<in> snd m #} =
- ({#a#} \<times># {# n \<in># N . a \<in> n #})" (is "?A = ?B")
-proof -
- have stmt: "\<And>m. m \<in># ({#a#} \<times># N) \<Longrightarrow> fst m \<in> snd m \<longleftrightarrow> a \<in> snd m"
- by (simp add: mem_Times_mset_iff)
- thus ?thesis by (metis (no_types, lifting) Sigma_mset_cong stmt cart_prod_singleton_left_gen)
-qed
-
-lemma cart_product_singleton_right: "{#m \<in># (N \<times># {#b#}) . fst m \<in> snd m #} =
- ({# n \<in># N . n \<in> b #} \<times># {# b #})" (is "?A = ?B")
-proof -
- have stmt: "\<And>m. m \<in># (N \<times># {#b#}) \<Longrightarrow> fst m \<in> snd m \<longleftrightarrow> fst m \<in>b"
- by (simp add: mem_Times_mset_iff)
- thus ?thesis by (metis (no_types, lifting) Sigma_mset_cong stmt cart_prod_singleton_right_gen)
-qed
-
-lemma cart_product_add_1_filter_eq: "{#m \<in># ((add_mset a M) \<times># N) . (fst m \<in> snd m) #} =
- {#m \<in># (M \<times># N) . (fst m \<in> snd m) #} + ({#a#} \<times># {# n \<in># N . a \<in> n #})"
- unfolding add_mset_add_single [of a M] Sigma_mset_plus_distrib1
- using cart_product_singleton_left cart_product_add_1_filter by fastforce
-
-lemma cart_product_add_1_filter_eq_mirror: "{#m \<in># M \<times># (add_mset b N) . (fst m \<in> snd m) #} =
- {#m \<in># (M \<times># N) . (fst m \<in> snd m) #} + ({# n \<in># M . n \<in> b #} \<times># {#b#})"
- unfolding add_mset_add_single [of b N] Sigma_mset_plus_distrib1 (* longish *)
- by (metis (no_types) add_mset_add_single cart_product_add_1_filter2 cart_product_singleton_right)
-
-lemma set_break_down_left:
- shows "{# m \<in># (M \<times># N) . (fst m) \<in> (snd m) #} = (\<Sum>m\<in>#M. ({#m#} \<times># {#n \<in># N. m \<in> n#}))"
- by (induction M) (auto simp add: cart_product_add_1_filter_eq)
-
-lemma set_break_down_right:
- shows "{# x \<in># M \<times># N . (fst x) \<in> (snd x) #} = (\<Sum>n\<in>#N. ({#m \<in># M. m \<in> n#} \<times># {#n#}))"
- by (induction N) (auto simp add: cart_product_add_1_filter_eq_mirror)
-
-text \<open>Reasoning on sums of elements over multisets\<close>
-
-lemma sum_over_fun_eq:
- assumes "\<And> x . x \<in># A \<Longrightarrow> f x = g x"
- shows "(\<Sum>x \<in># A . f(x)) = (\<Sum> x \<in># A . g (x))"
- using assms by auto
-
-context ring_1
-begin
-
-lemma sum_mset_add_diff: "(\<Sum> x \<in># A. f x - g x) = (\<Sum> x \<in># A . f x) - (\<Sum> x \<in># A . g x)"
- by (induction A) (auto simp add: algebra_simps)
-
-end
-
-context ordered_ring
-begin
-
-lemma sum_mset_ge0:"(\<And> x . f x \<ge> 0) \<Longrightarrow> (\<Sum> x \<in># A. f x ) \<ge> 0"
-proof (induction A)
- case empty
- then show ?case by simp
-next
- case (add x A)
- then have hyp2: "0 \<le> sum_mset (image_mset f A)" by blast
- then have " sum_mset (image_mset f (add_mset x A)) = sum_mset (image_mset f A) + f x"
- by (simp add: add_commute)
- then show ?case
- by (simp add: add.IH add.prems)
-qed
-
-lemma sum_order_add_mset: "(\<And> x . f x \<ge> 0) \<Longrightarrow> (\<Sum> x \<in># A. f x ) \<le> (\<Sum> x \<in># add_mset a A. f x )"
- by simp
-
-lemma sum_mset_0_left: "(\<And> x . f x \<ge> 0) \<Longrightarrow> (\<Sum> x \<in># A. f x ) = 0 \<Longrightarrow> (\<forall> x \<in># A .f x = 0)"
- apply (induction A)
- apply auto
- using local.add_nonneg_eq_0_iff sum_mset_ge0 apply blast
- using local.add_nonneg_eq_0_iff sum_mset_ge0 by blast
-
-lemma sum_mset_0_iff_ge_0:
- assumes "(\<And> x . f x \<ge> 0)"
- shows "(\<Sum> x \<in># A. f x ) = 0 \<longleftrightarrow> (\<forall> x \<in> set_mset A .f x = 0)"
- using sum_mset_0_left assms by auto
-
-end
-
-lemma mset_set_size_card_count: "(\<Sum>x \<in># A. x) = (\<Sum>x \<in> set_mset A . x * (count A x))"
-proof (induction A)
- case empty
- then show ?case by simp
-next
- case (add y A)
- have lhs: "(\<Sum>x\<in>#add_mset y A. x) = (\<Sum>x\<in># A. x) + y" by simp
- have rhs: "(\<Sum>x\<in>set_mset (add_mset y A). x * count (add_mset y A) x) =
- (\<Sum>x\<in>(insert y (set_mset A)) . x * count (add_mset y A) x)"
- by simp
- then show ?case
- proof (cases "y \<in># A")
- case True
- have x_val: "\<And> x . x \<in> (insert y (set_mset A)) \<Longrightarrow> x \<noteq> y \<Longrightarrow>
- x* count (add_mset y A) x = x * (count A x)"
- by auto
- have y_count: "count (add_mset y A) y = 1 + count A y"
- using True count_inI by fastforce
- then have "(\<Sum>x\<in>set_mset (add_mset y A). x * count (add_mset y A) x) =
- (y * (count (add_mset y A) y)) + (\<Sum>x\<in>(set_mset A) - {y}. x * count A x)"
- using x_val finite_set_mset sum.cong sum.insert rhs
- by (smt DiffD1 Diff_insert_absorb insert_absorb mk_disjoint_insert sum.insert_remove)
- then have s1: "(\<Sum>x\<in>set_mset (add_mset y A). x * count (add_mset y A) x) =
- y + y * (count A y) + (\<Sum>x\<in>(set_mset A) - {y}. x * count A x)"
- using y_count by simp
- then have "(\<Sum>x\<in>set_mset (add_mset y A). x * count (add_mset y A) x) =
- y + (\<Sum>x\<in>insert y ((set_mset A) - {y} ) . x * count A x)"
- by (simp add: sum.insert_remove)
- then have "(\<Sum>x\<in>set_mset (add_mset y A). x * count (add_mset y A) x) =
- y + (\<Sum>x\<in>(set_mset A) . x * count A x)"
- by (simp add: True insert_absorb)
- then show ?thesis using lhs add.IH
- by linarith
- next
- case False
- have x_val: "\<And> x . x \<in> set_mset A \<Longrightarrow> x* count (add_mset y A) x = x * (count A x)"
- using False by auto
- have y_count: "count (add_mset y A) y = 1" using False count_inI by fastforce
- have lhs: "(\<Sum>x\<in>#add_mset y A. x) = (\<Sum>x\<in># A. x) + y" by simp
- have "(\<Sum>x\<in>set_mset (add_mset y A). x * count (add_mset y A) x) =
- (y * count (add_mset y A) y) + (\<Sum>x\<in>set_mset A. x * count A x)"
- using x_val rhs by (metis (no_types, lifting) False finite_set_mset sum.cong sum.insert)
- then have "(\<Sum>x\<in>set_mset (add_mset y A). x * count (add_mset y A) x) =
- y + (\<Sum>x\<in>set_mset A. x * count A x)"
- using y_count by simp
- then show ?thesis using lhs add.IH by linarith
- qed
-qed
-
-subsection \<open>Partitions on Multisets\<close>
-
-text \<open>A partition on a multiset A is a multiset of multisets, where the sum over P equals A and the
-empty multiset is not in the partition. Based off set partition definition.
-We note that unlike set partitions, there is no requirement for elements in the multisets to be
-distinct due to the definition of union on multisets \cite{benderPartitionsMultisets1974}\<close>
-
-lemma mset_size_partition_dep: "size {# a \<in># A . P a \<or> Q a #} =
- size {# a \<in># A . P a #} + size {# a \<in># A . Q a #} - size {# a \<in># A . P a \<and> Q a #}"
- by (simp add: mset_bunion_filter mset_inter_filter mset_union_size_inter)
-
-definition partition_on_mset :: "'a multiset \<Rightarrow> 'a multiset multiset \<Rightarrow> bool" where
-"partition_on_mset A P \<longleftrightarrow> \<Sum>\<^sub>#P = A \<and> {#} \<notin># P"
-
-lemma partition_on_msetI [intro]: "\<Sum>\<^sub>#P = A \<Longrightarrow> {#} \<notin># P \<Longrightarrow> partition_on_mset A P"
- by (simp add: partition_on_mset_def)
-
-lemma partition_on_msetD1: "partition_on_mset A P \<Longrightarrow> \<Sum>\<^sub>#P = A"
- by (simp add: partition_on_mset_def)
-
-lemma partition_on_msetD2: "partition_on_mset A P \<Longrightarrow> {#} \<notin># P"
- by (simp add: partition_on_mset_def)
-
-lemma partition_on_mset_empty: "partition_on_mset {#} P \<longleftrightarrow> P = {#}"
- unfolding partition_on_mset_def
- using multiset_nonemptyE by fastforce
-
-lemma partition_on_mset_all: "A \<noteq> {#} \<Longrightarrow> partition_on_mset A {#A #}"
- by (simp add: partition_on_mset_def)
-
-lemma partition_on_mset_singletons: "partition_on_mset A (image_mset (\<lambda> x . {#x#}) A)"
- by (auto simp: partition_on_mset_def)
-
-lemma partition_on_mset_not_empty: "A \<noteq> {#} \<Longrightarrow> partition_on_mset A P \<Longrightarrow> P \<noteq> {#}"
- by (auto simp: partition_on_mset_def)
-
-lemma partition_on_msetI2: "\<Sum>\<^sub>#P = A \<Longrightarrow> (\<And> p . p \<in># P \<Longrightarrow> p \<noteq> {#}) \<Longrightarrow> partition_on_mset A P"
- by (auto simp: partition_on_mset_def)
-
-lemma partition_on_mset_elems: "partition_on_mset A P \<Longrightarrow> p1 \<in># P \<Longrightarrow> x \<in># p1 \<Longrightarrow> x \<in># A"
- by (auto simp: partition_on_mset_def)
-
-lemma partition_on_mset_sum_size_eq: "partition_on_mset A P \<Longrightarrow> (\<Sum>x \<in># P. size x) = size A"
- by (metis partition_on_msetD1 size_big_union_sum)
-
-lemma partition_on_mset_card: assumes "partition_on_mset A P" shows " size P \<le> size A"
-proof (rule ccontr)
- assume "\<not> size P \<le> size A"
- then have a: "size P > size A" by simp
- have "\<And> x . x \<in># P \<Longrightarrow> size x > 0" using partition_on_msetD2
- using assms nonempty_has_size by auto
- then have " (\<Sum>x \<in># P. size x) \<ge> size P"
- by (metis leI less_one not_less_zero size_eq_sum_mset sum_mset_mono)
- thus False using a partition_on_mset_sum_size_eq
- using assms by fastforce
-qed
-
-lemma partition_on_mset_count_eq: "partition_on_mset A P \<Longrightarrow> a \<in># A \<Longrightarrow>
- (\<Sum>x \<in># P. count x a) = count A a"
- by (metis count_sum_mset partition_on_msetD1)
-
-lemma partition_on_mset_subsets: "partition_on_mset A P \<Longrightarrow> x \<in># P \<Longrightarrow> x \<subseteq># A"
- by (auto simp add: partition_on_mset_def)
-
-lemma partition_on_mset_distinct:
- assumes "partition_on_mset A P"
- assumes "distinct_mset A"
- shows "distinct_mset P"
-proof (rule ccontr)
- assume "\<not> distinct_mset P"
- then obtain p1 where count: "count P p1 \<ge> 2"
- by (metis Suc_1 distinct_mset_count_less_1 less_Suc_eq_le not_less_eq)
- then have cge: "\<And> x . x \<in># p1 \<Longrightarrow> (\<Sum>p \<in># P. count p x ) \<ge> 2"
- by (smt count_greater_eq_one_iff count_sum_mset_if_1_0 dual_order.trans sum_mset_mono zero_le)
- have elem_in: "\<And> x . x \<in># p1 \<Longrightarrow> x \<in># A" using partition_on_mset_elems
- by (metis count assms(1) count_eq_zero_iff not_numeral_le_zero)
- have "\<And> x . x \<in># A \<Longrightarrow> count A x = 1" using assms
- by (simp add: distinct_mset_def)
- thus False
- using assms partition_on_mset_count_eq cge elem_in count_inI local.count multiset_nonemptyE
- by (metis (mono_tags) not_numeral_le_zero numeral_One numeral_le_iff partition_on_mset_def semiring_norm(69))
-qed
-
-lemma partition_on_mset_distinct_disjoint:
- assumes "partition_on_mset A P"
- assumes "distinct_mset A"
- assumes "p1 \<in># P"
- assumes "p2 \<in># P - {#p1#}"
- shows "p1 \<inter># p2 = {#}"
- using Diff_eq_empty_iff_mset assms diff_add_zero distinct_mset_add multiset_inter_assoc sum_mset.remove
- by (smt partition_on_msetD1 subset_mset.inf.absorb_iff2 subset_mset.le_add_same_cancel1 subset_mset.le_iff_inf)
-
-lemma partition_on_mset_diff:
- assumes "partition_on_mset A P"
- assumes "Q \<subseteq>#P"
- shows "partition_on_mset (A - \<Sum>\<^sub>#Q) (P - Q)"
- using assms partition_on_mset_def
- by (smt diff_union_cancelL subset_mset.add_diff_inverse sum_mset.union union_iff)
-
-lemma sigma_over_set_partition_count:
- assumes "finite A"
- assumes "partition_on A P"
- assumes "x \<in># \<Sum>\<^sub># (mset_set (mset_set ` P))"
- shows "count (\<Sum>\<^sub># (mset_set (mset_set ` P))) x = 1"
-proof -
- have disj: "disjoint P" using assms partition_onD2 by auto
- then obtain p where pin: "p \<in># mset_set (mset_set ` P)" and xin: "x \<in># p"
- using assms by blast
- then have "count (mset_set (mset_set ` P)) p = 1"
- by (meson count_eq_zero_iff count_mset_set')
- then have filter: "\<And> p' . p' \<in># ((mset_set (mset_set` P)) - {#p#}) \<Longrightarrow> p \<noteq> p'"
- using count_eq_zero_iff count_single by fastforce
- have zero: "\<And> p'. p' \<in># mset_set (mset_set ` P) \<Longrightarrow> p' \<noteq> p \<Longrightarrow> count p' x = 0"
- proof (rule ccontr)
- fix p'
- assume assm: "p' \<in># mset_set (mset_set ` P)" and ne: "p' \<noteq> p" and n0: "count p' x \<noteq> 0"
- then have xin2: "x \<in># p'" by auto
- obtain p1 p2 where p1in: "p1 \<in> P" and p2in: "p2 \<in> P" and p1eq: "mset_set p1 = p"
- and p2eq: "mset_set p2 = p'" using assm assms(1) assms(2) pin
- by (metis (no_types, lifting) elem_mset_set finite_elements finite_imageI image_iff)
- have origne: "p1 \<noteq> p2" using ne p1eq p2eq by auto
- have "p1 = p2" using partition_onD4 xin xin2
- by (metis assms(2) count_eq_zero_iff count_mset_set' p1eq p1in p2eq p2in)
- then show False using origne by simp
- qed
- have one: "count p x = 1" using pin xin assms count_eq_zero_iff count_greater_eq_one_iff
- by (metis count_mset_set(3) count_mset_set_le_one image_iff le_antisym)
- then have "count (\<Sum>\<^sub># (mset_set (mset_set ` P))) x =
- (\<Sum>p' \<in># (mset_set (mset_set ` P)) . count p' x)"
- using count_sum_mset by auto
- also have "... = (count p x) + (\<Sum>p' \<in># ((mset_set (mset_set ` P)) - {#p#}) . count p' x)"
- by (metis (mono_tags, lifting) insert_DiffM pin sum_mset.insert)
- also have "... = 1 + (\<Sum>p' \<in># ((mset_set (mset_set ` P)) - {#p#}) . count p' x)"
- using one by presburger
- finally have "count (\<Sum>\<^sub># (mset_set (mset_set ` P))) x =
- 1 + (\<Sum>p' \<in># ((mset_set (mset_set ` P)) - {#p#}) . 0)"
- using zero filter by (metis (mono_tags, lifting) in_diffD sum_over_fun_eq)
- then show "count (\<Sum>\<^sub># (mset_set (mset_set ` P))) x = 1" by simp
-qed
-
-lemma partition_on_mset_set:
- assumes "finite A"
- assumes "partition_on A P"
- shows "partition_on_mset (mset_set A) (mset_set (image (\<lambda> x. mset_set x) P))"
-proof (intro partition_on_msetI)
- have partd1: "\<Union>P = A" using assms partition_onD1 by auto
- have imp: "\<And>x. x \<in># \<Sum>\<^sub># (mset_set (mset_set ` P)) \<Longrightarrow> x \<in># mset_set A"
- proof -
- fix x
- assume "x \<in># \<Sum>\<^sub># (mset_set (mset_set ` P))"
- then obtain p where "p \<in> (mset_set ` P)" and xin: "x \<in># p"
- by (metis elem_mset_set equals0D infinite_set_mset_mset_set mset_big_union_obtain)
- then have "set_mset p \<in> P"
- by (metis empty_iff finite_set_mset_mset_set image_iff infinite_set_mset_mset_set)
- then show "x \<in># mset_set A"
- using partd1 xin assms(1) by auto
- qed
- have imp2: "\<And>x . x \<in># mset_set A \<Longrightarrow> x \<in># \<Sum>\<^sub># (mset_set (mset_set ` P))"
- proof -
- fix x
- assume "x \<in># mset_set A"
- then have "x \<in> A" by (simp add: assms(1))
- then obtain p where "p \<in> P" and "x \<in> p" using assms(2) using partd1 by blast
- then obtain p' where "p' \<in> (mset_set ` P)" and "p' = mset_set p" by blast
- thus "x \<in># \<Sum>\<^sub># (mset_set (mset_set ` P))" using assms \<open>p \<in> P\<close> \<open>x \<in> p\<close> finite_elements partd1
- by (metis Sup_upper finite_imageI finite_set_mset_mset_set in_Union_mset_iff rev_finite_subset)
- qed
- have a1: "\<And> x . x \<in># mset_set A \<Longrightarrow> count (mset_set A) x = 1"
- using assms(1) by fastforce
- then show "\<Sum>\<^sub># (mset_set (mset_set ` P)) = mset_set A" using imp imp2 a1
- by (metis assms(1) assms(2) count_eq_zero_iff multiset_eqI sigma_over_set_partition_count)
- have "\<And> p. p \<in> P \<Longrightarrow> p \<noteq> {} " using assms partition_onD3 by auto
- then have "\<And> p. p \<in> P \<Longrightarrow> mset_set p \<noteq> {#}" using mset_set_empty_iff
- by (metis Union_upper assms(1) partd1 rev_finite_subset)
- then show "{#} \<notin># mset_set (mset_set ` P)"
- by (metis elem_mset_set equals0D image_iff infinite_set_mset_mset_set)
-qed
-
-lemma partition_on_mset_distinct_inter:
- assumes "partition_on_mset A P"
- assumes "distinct_mset A"
- assumes "p1 \<in># P" and "p2 \<in># P" and "p1 \<noteq> p2"
- shows "p1 \<inter># p2 = {#}"
- by (metis assms in_remove1_mset_neq partition_on_mset_distinct_disjoint)
-
-lemma partition_on_set_mset_distinct:
- assumes "partition_on_mset A P"
- assumes "distinct_mset A"
- assumes "p \<in># image_mset set_mset P"
- assumes "p' \<in># image_mset set_mset P"
- assumes "p \<noteq> p'"
- shows "p \<inter> p' = {}"
-proof -
- obtain p1 where p1in: "p1 \<in># P" and p1eq: "set_mset p1 = p" using assms(3)
- by blast
- obtain p2 where p2in: "p2 \<in># P" and p2eq: "set_mset p2 = p'" using assms(4) by blast
- have "distinct_mset P" using assms partition_on_mset_distinct by blast
- then have "p1 \<noteq> p2" using assms using p1eq p2eq by fastforce
- then have "p1 \<inter># p2 = {#}" using partition_on_mset_distinct_inter
- using assms(1) assms(2) p1in p2in by auto
- thus ?thesis using p1eq p2eq
- by (metis set_mset_empty set_mset_inter)
-qed
-
-lemma partition_on_set_mset:
- assumes "partition_on_mset A P"
- assumes "distinct_mset A"
- shows "partition_on (set_mset A) (set_mset (image_mset set_mset P))"
-proof (intro partition_onI)
- show "\<And>p. p \<in># image_mset set_mset P \<Longrightarrow> p \<noteq> {}"
- using assms(1) partition_on_msetD2 by fastforce
-next
- have "\<And> x . x \<in> set_mset A \<Longrightarrow> x \<in> \<Union> (set_mset (image_mset set_mset P))"
- by (metis Union_iff assms(1) image_eqI mset_big_union_obtain partition_on_msetD1 set_image_mset)
- then show "\<Union> (set_mset (image_mset set_mset P)) = set_mset A"
- using set_eqI' partition_on_mset_elems assms by auto
- show "\<And>p p'. p \<in># image_mset set_mset P \<Longrightarrow> p' \<in># image_mset set_mset P \<Longrightarrow>
- p \<noteq> p' \<Longrightarrow> p \<inter> p' = {}"
- using partition_on_set_mset_distinct assms by blast
-qed
-
-lemma partition_on_mset_eq_imp_eq_carrier:
- assumes "partition_on_mset A P"
- assumes "partition_on_mset B P"
- shows "A = B"
- using assms partition_on_msetD1 by auto
-
-lemma partition_on_mset_add_single:
- assumes "partition_on_mset A P"
- shows "partition_on_mset (add_mset a A) (add_mset {#a#} P)"
- using assms by (auto simp: partition_on_mset_def)
-
-lemma partition_on_mset_add_part:
- assumes "partition_on_mset A P"
- assumes "X \<noteq> {#}"
- assumes "A + X = A'"
- shows "partition_on_mset A' (add_mset X P)"
- using assms by (auto simp: partition_on_mset_def)
-
-lemma partition_on_mset_add:
- assumes "partition_on_mset A P"
- assumes "X \<in># P"
- assumes "add_mset a X = X'"
- shows "partition_on_mset (add_mset a A) (add_mset X' (P - {#X#}))"
- using add_mset_add_single assms empty_not_add_mset mset_subset_eq_single partition_on_mset_all
- by (smt partition_on_mset_def subset_mset.add_diff_inverse sum_mset.add_mset sum_mset.remove union_iff union_mset_add_mset_left)
-
-lemma partition_on_mset_elem_exists_part:
- assumes "partition_on_mset A P"
- assumes "x \<in># A"
- obtains p where "p \<in># P" and "x \<in># p"
- using assms in_Union_mset_iff partition_on_msetD2 partition_on_msetI
- by (metis partition_on_mset_eq_imp_eq_carrier)
-
-lemma partition_on_mset_combine:
- assumes "partition_on_mset A P"
- assumes "partition_on_mset B Q"
- shows "partition_on_mset (A + B) (P + Q)"
- unfolding partition_on_mset_def
- using assms partition_on_msetD1 partition_on_msetD2 by auto
-
-lemma partition_on_mset_split:
- assumes "partition_on_mset A (P + Q)"
- shows "partition_on_mset (\<Sum>\<^sub>#P) P"
- using partition_on_mset_def partition_on_msetD2 assms by fastforce
+(* Title: Multisets_Extras
+ Author: Chelsea Edmonds
+*)
+
+section \<open>Micellanious Helper Functions on Sets and Multisets\<close>
+
+theory Multisets_Extras imports
+ "HOL-Library.Multiset"
+ Card_Partitions.Set_Partition
+ Nested_Multisets_Ordinals.Multiset_More
+ Nested_Multisets_Ordinals.Duplicate_Free_Multiset
+ "HOL-Library.Disjoint_Sets"
+begin
+
+subsection \<open>Set Theory Extras\<close>
+
+text \<open>A number of extra helper lemmas for reasoning on sets (finite) required for Design Theory
+proofs\<close>
+
+lemma card_Pow_filter_one:
+ assumes "finite A"
+ shows "card {x \<in> Pow A . card x = 1} = card (A)"
+ using assms
+proof (induct rule: finite_induct)
+ case empty
+ then show ?case by auto
+next
+ case (insert x F)
+ have "Pow (insert x F) = Pow F \<union> insert x ` Pow F"
+ by (simp add: Pow_insert)
+ then have split: "{y \<in> Pow (insert x F) . card y = 1} =
+ {y \<in> (Pow F) . card y = 1} \<union> {y \<in> (insert x ` Pow F) . card y = 1}"
+ by blast
+ have "\<And> y . y \<in> (insert x ` Pow F) \<Longrightarrow> finite y"
+ using finite_subset insert.hyps(1) by fastforce
+ then have single: "\<And> y . y \<in> (insert x ` Pow F) \<Longrightarrow> card y = 1 \<Longrightarrow> y = {x}"
+ by (metis card_1_singletonE empty_iff image_iff insertCI insertE)
+ then have "card {y \<in> (insert x ` Pow F) . card y = 1} = 1"
+ using empty_iff imageI is_singletonI is_singletonI' is_singleton_altdef (* LONG *)
+ by (metis (full_types, lifting) Collect_empty_eq_bot Pow_bottom bot_empty_eq mem_Collect_eq)
+ then have " {y \<in> (insert x ` Pow F) . card y = 1} = {{x}}"
+ using single card_1_singletonE card_eq_0_iff
+ by (smt empty_Collect_eq mem_Collect_eq singletonD zero_neq_one)
+ then have split2:"{y \<in> Pow (insert x F) . card y = 1} = {y \<in> (Pow F) . card y = 1} \<union> {{x}}"
+ using split by simp
+ then show ?case
+ proof (cases "x \<in> F")
+ case True
+ then show ?thesis using insert.hyps(2) by auto
+ next
+ case False
+ then have "{y \<in> (Pow F) . card y = 1} \<inter> {{x}} = {}" by blast
+ then have fact:"card {y \<in> Pow (insert x F) . card y = 1} =
+ card {y \<in> (Pow F) . card y = 1} + card {{x}}"
+ using split2 card_Un_disjoint insert.hyps(1) by auto
+ have "card (insert x F) = card F + 1"
+ using False card_insert_disjoint by (metis Suc_eq_plus1 insert.hyps(1))
+ then show ?thesis using fact insert.hyps(3) by auto
+ qed
+qed
+
+lemma elem_exists_non_empty_set:
+ assumes "card A > 0"
+ obtains x where "x \<in> A"
+ using assms card_gt_0_iff by fastforce
+
+lemma set_self_img_compr: "{a | a . a \<in> A} = A"
+ by blast
+
+lemma card_subset_not_gt_card: "finite A \<Longrightarrow> card ps > card A \<Longrightarrow> \<not> (ps \<subseteq> A)"
+ using card_mono leD by auto
+
+lemma card_inter_lt_single: "finite A \<Longrightarrow> finite B \<Longrightarrow> card (A \<inter> B) \<le> card A"
+ by (simp add: card_mono)
+
+lemma set_diff_non_empty_not_subset:
+ assumes "A \<subseteq> (B - C)"
+ assumes "C \<noteq> {}"
+ assumes "A \<noteq> {}"
+ assumes "B \<noteq> {}"
+ shows " \<not> (A \<subseteq> C)"
+proof (rule ccontr)
+ assume " \<not> \<not> (A \<subseteq> C)"
+ then have a: "\<And> x . x \<in> A \<Longrightarrow> x \<in> C" by blast
+ thus False using a assms by blast
+qed
+
+lemma set_card_diff_ge_zero: "finite A \<Longrightarrow> finite B \<Longrightarrow> A \<noteq> B \<Longrightarrow> card A = card B \<Longrightarrow>
+ card (A - B) > 0"
+ by (meson Diff_eq_empty_iff card_0_eq card_subset_eq finite_Diff neq0_conv)
+
+lemma set_filter_diff: "{a \<in> A . P a } - {x} = {a \<in> (A - {x}) . (P a )}"
+ by (auto)
+
+lemma set_filter_diff_card: "card ({a \<in> A . P a } - {x}) = card {a \<in> (A - {x}) . (P a )}"
+ by (simp add: set_filter_diff)
+
+lemma obtain_subset_with_card_int_n:
+ assumes "(n ::int) \<le> of_nat (card S)"
+ assumes "(n ::int) \<ge> 0"
+ obtains T where "T \<subseteq> S" "of_nat (card T) = (n ::int)" "finite T"
+ using obtain_subset_with_card_n assms
+ by (metis nonneg_int_cases of_nat_le_iff)
+
+lemma transform_filter_img_empty_rm:
+ assumes "\<And> g . g \<in> G \<Longrightarrow> g \<noteq> {}"
+ shows "{g - {x} | g. g \<in> G \<and> g \<noteq> {x}} = {g - {x} | g. g \<in> G } - {{}}"
+proof -
+ let ?f = "\<lambda> g . g - {x}"
+ have "\<And> g . g \<in> G \<Longrightarrow> g \<noteq> {x} \<longleftrightarrow> ?f g \<noteq> {}" using assms
+ by (metis Diff_cancel Diff_empty Diff_insert0 insert_Diff)
+ thus ?thesis by auto
+qed
+
+lemma bij_betw_inter_subsets: "bij_betw f A B \<Longrightarrow> a1 \<subseteq> A \<Longrightarrow> a2 \<subseteq> A
+ \<Longrightarrow> f ` (a1 \<inter> a2) = (f ` a1) \<inter> (f ` a2)"
+ by (meson bij_betw_imp_inj_on inj_on_image_Int)
+
+text\<open>Partition related set theory lemmas\<close>
+
+lemma partition_on_remove_pt:
+ assumes "partition_on A G"
+ shows "partition_on (A - {x}) {g - {x} | g. g \<in> G \<and> g \<noteq> {x}}"
+proof (intro partition_onI)
+ show "\<And>p. p \<in> {g - {x} |g. g \<in> G \<and> g \<noteq> {x}} \<Longrightarrow> p \<noteq> {}"
+ using assms partition_onD3 subset_singletonD by force
+ let ?f = "(\<lambda> g . g - {x})"
+ have un_img: "\<Union>({?f g | g. g \<in> G }) = ?f (\<Union> G)" by blast
+ have empty: "\<Union> {g - {x} |g. g \<in> G \<and> g \<noteq> {x}} = \<Union>({g - {x} | g. g \<in> G } - {{}})"
+ by blast
+ then have "\<Union>({g - {x} | g. g \<in> G } - {{}}) = \<Union>({g - {x} | g. g \<in> G })" by blast
+ then show " \<Union> {g - {x} |g. g \<in> G \<and> g \<noteq> {x}} = A - {x}" using partition_onD1 assms un_img
+ by (metis empty)
+ then show "\<And>p p'.
+ p \<in> {g - {x} |g. g \<in> G \<and> g \<noteq> {x}} \<Longrightarrow>
+ p' \<in> {g - {x} |g. g \<in> G \<and> g \<noteq> {x}} \<Longrightarrow> p \<noteq> p' \<Longrightarrow> p \<inter> p' = {}"
+ proof -
+ fix p1 p2
+ assume p1: "p1 \<in> {g - {x} |g. g \<in> G \<and> g \<noteq> {x}}"
+ and p2: "p2 \<in> {g - {x} |g. g \<in> G \<and> g \<noteq> {x}}"
+ and ne: "p1 \<noteq> p2"
+ obtain p1' p2' where orig1: "p1 = p1' - {x}" and orig2: "p2 = p2' - {x}"
+ and origne: "p1' \<noteq> p2'" and ne1: "p1' \<noteq> {x}" and ne2:"p2' \<noteq> {x}" and ing1: "p1' \<in> G"
+ and ing2: "p2' \<in> G"
+ using p1 p2 using mem_Collect_eq ne by blast
+ then have "p1' \<inter> p2' = {}" using assms partition_onD2 ing1 ing2 origne disjointD by blast
+ thus "p1 \<inter> p2 = {}" using orig1 orig2 by blast
+ qed
+qed
+
+lemma partition_on_cart_prod:
+ assumes "card I > 0"
+ assumes "A \<noteq> {}"
+ assumes "G \<noteq> {}"
+ assumes "partition_on A G"
+ shows "partition_on (A \<times> I) {g \<times> I |g. g \<in> G}"
+proof (intro partition_onI)
+ show "\<And>p. p \<in> {g \<times> I |g. g \<in> G} \<Longrightarrow> p \<noteq> {}"
+ using assms(1) assms(4) partition_onD3 by fastforce
+ show "\<Union> {g \<times> I |g. g \<in> G} = A \<times> I"
+ by (metis Setcompr_eq_image Sigma_Union assms(4) partition_onD1)
+ show "\<And>p p'. p \<in> {g \<times> I |g. g \<in> G} \<Longrightarrow> p' \<in> {g \<times> I |g. g \<in> G} \<Longrightarrow> p \<noteq> p' \<Longrightarrow> p \<inter> p' = {}"
+ by (smt (verit, best) Sigma_Int_distrib1 Sigma_empty1 assms(4) mem_Collect_eq partition_onE)
+qed
+
+subsection \<open>Multiset Helpers\<close>
+
+text \<open>Generic Size, count and card helpers\<close>
+
+lemma count_size_set_repr: "size {# x \<in># A . x = g#} = count A g"
+ by (simp add: filter_eq_replicate_mset)
+
+lemma mset_nempty_set_nempty: "A \<noteq> {#} \<longleftrightarrow> (set_mset A) \<noteq> {}"
+ by simp
+
+lemma mset_size_ne0_set_card: "size A > 0 \<Longrightarrow> card (set_mset A) > 0"
+ using mset_nempty_set_nempty by fastforce
+
+lemma set_count_size_min: "count A a \<ge> n \<Longrightarrow> size A \<ge> n"
+ by (metis (full_types) count_le_replicate_mset_subset_eq size_mset_mono size_replicate_mset)
+
+lemma card_size_filter_eq: "finite A \<Longrightarrow> card {a \<in> A . P a} = size {#a \<in># mset_set A . P a#}"
+ by simp
+
+lemma size_multiset_int_count:
+ assumes "of_nat (card (set_mset A)) = (ca :: int)"
+ assumes "\<And>p. p \<in># A \<Longrightarrow> of_nat (count A p) = (ca2 :: int)"
+ shows "of_nat (size A) = ((ca :: int) * ca2)"
+proof -
+ have "size A = (\<Sum> p \<in> (set_mset A) . count A p)" using size_multiset_overloaded_eq by auto
+ then have "of_nat (size A) = (\<Sum> p \<in> (set_mset A) . ca2)" using assms by simp
+ thus ?thesis using assms(1) by auto
+qed
+
+lemma mset_union_size: "size (A \<union># B) = size (A) + size (B - A)"
+ by (simp add: union_mset_def)
+
+lemma mset_union_size_inter: "size (A \<union># B) = size (A) + size B - size (A \<inter># B)"
+ by (metis diff_add_inverse2 size_Un_Int)
+
+text \<open>Lemmas for repeat\_mset\<close>
+
+lemma repeat_mset_size [simp]: "size (repeat_mset n A) = n * size A"
+ by (induction n) auto
+
+lemma repeat_mset_subset_in:
+ assumes "\<And> a . a \<in># A \<Longrightarrow> a \<subseteq> B"
+ assumes "X \<in># repeat_mset n A"
+ assumes "x \<in> X"
+ shows " x \<in> B"
+ using assms by (induction n) auto
+
+lemma repeat_mset_not_empty: "n > 0 \<Longrightarrow> A \<noteq> {#} \<Longrightarrow> repeat_mset n A \<noteq> {#}"
+ by (induction n) auto
+
+lemma elem_in_repeat_in_original: "a \<in># repeat_mset n A \<Longrightarrow> a \<in># A"
+ by (metis count_inI count_repeat_mset in_countE mult.commute mult_zero_left nat.distinct(1))
+
+lemma elem_in_original_in_repeat: "n > 0 \<Longrightarrow> a \<in># A \<Longrightarrow> a \<in># repeat_mset n A"
+ by (metis count_greater_zero_iff count_repeat_mset nat_0_less_mult_iff)
+
+text \<open>Lemmas on image and filter for multisets\<close>
+
+lemma multiset_add_filter_size: "size {# a \<in># (A1 + A2) . P a #} = size {# a \<in># A1 . P a #} +
+ size {# a \<in># A2 . P a #}"
+ by simp
+
+lemma size_filter_neg: "size {#a \<in># A . P a #} = size A - size {# a \<in># A . \<not> P a #}"
+ using size_filter_mset_lesseq size_union union_filter_mset_complement
+ by (metis ordered_cancel_comm_monoid_diff_class.le_imp_diff_is_add)
+
+lemma filter_filter_mset_cond_simp:
+ assumes "\<And> a . P a \<Longrightarrow> Q a"
+ shows "filter_mset P A = filter_mset P (filter_mset Q A)"
+proof -
+ have "filter_mset P (filter_mset Q A) = filter_mset (\<lambda> a. Q a \<and> P a) A"
+ by (simp add: filter_filter_mset)
+ thus ?thesis using assms
+ by (metis (mono_tags, lifting) filter_mset_cong)
+qed
+
+lemma filter_filter_mset_ss_member: "filter_mset (\<lambda> a . {x, y} \<subseteq> a) A =
+ filter_mset (\<lambda> a . {x, y} \<subseteq> a) (filter_mset (\<lambda> a . x \<in> a) A)"
+proof -
+ have filter: "filter_mset (\<lambda> a . {x, y} \<subseteq> a) (filter_mset (\<lambda> a . x \<in> a) A) =
+ filter_mset (\<lambda> a . x \<in> a \<and> {x, y} \<subseteq> a) A" by (simp add: filter_filter_mset)
+ have "\<And> a. {x, y} \<subseteq> a \<Longrightarrow> x \<in> a" by simp
+ thus ?thesis using filter by auto
+qed
+
+lemma multiset_image_do_nothing: "(\<And> x .x \<in># A \<Longrightarrow> f x = x) \<Longrightarrow> image_mset f A = A"
+ by (induct A) auto
+
+lemma set_mset_filter: "set_mset {# f a . a \<in># A #} = {f a | a. a \<in># A}"
+ by (simp add: Setcompr_eq_image)
+
+lemma mset_exists_imply: "x \<in># {# f a . a \<in># A #} \<Longrightarrow> \<exists> y \<in># A . x = f y"
+ by auto
+
+lemma filter_mset_image_mset:
+ "filter_mset P (image_mset f A) = image_mset f (filter_mset (\<lambda>x. P (f x)) A)"
+ by (induction A) auto
+
+lemma mset_bunion_filter: "{# a \<in># A . P a \<or> Q a #} = {# a \<in># A . P a #} \<union># {# a \<in># A . Q a #}"
+ by (rule multiset_eqI) simp
+
+lemma mset_inter_filter: "{# a \<in># A . P a \<and> Q a #} = {# a \<in># A . P a #} \<inter># {# a \<in># A . Q a #}"
+ by (rule multiset_eqI) simp
+
+lemma image_image_mset: "image_mset (\<lambda> x . f x) (image_mset (\<lambda> y . g y) A) =
+ image_mset (\<lambda> x. f (g x)) A"
+ by simp
+
+text \<open>Big Union over multiset helpers\<close>
+
+lemma mset_big_union_obtain:
+ assumes "x \<in># \<Sum>\<^sub># A"
+ obtains a where "a \<in># A" and "x \<in># a"
+ using assms by blast
+
+lemma size_big_union_sum: "size (\<Sum>\<^sub># (M :: 'a multiset multiset)) = (\<Sum>x \<in>#M . size x)"
+ by (induct M) auto
+
+text \<open>Cartesian Product on Multisets\<close>
+
+lemma size_cartesian_product_singleton [simp]: "size ({#a#} \<times># B) = size B"
+ by (simp add: Times_mset_single_left)
+
+lemma size_cartesian_product_singleton_right [simp]: "size (A \<times># {#b#}) = size A"
+ by (simp add: Times_mset_single_right)
+
+lemma size_cartesian_product_empty [simp]: "size (A \<times># {#}) = 0"
+ by simp
+
+lemma size_add_elem_step_eq:
+ assumes "size (A \<times># B) = size A * size B"
+ shows "size (add_mset x A \<times># B) = size (add_mset x A) * size B"
+proof -
+ have "(add_mset x A \<times># B) = A \<times># B + {#x#} \<times># B"
+ by (metis Sigma_mset_plus_distrib1 add_mset_add_single)
+ then have "size (add_mset x A \<times># B) = size (A \<times># B) + size B" by auto
+ also have "... = size A * size B + size B"
+ by (simp add: assms)
+ finally have "size (add_mset x A \<times># B) = (size A + 1) * size B"
+ by auto
+ thus ?thesis by simp
+qed
+
+lemma size_cartesian_product: "size (A \<times># B) = size A * size B"
+ by (induct A) (simp_all add: size_add_elem_step_eq)
+
+lemma cart_prod_distinct_mset:
+ assumes assm1: "distinct_mset A"
+ assumes assm2: "distinct_mset B"
+ shows "distinct_mset (A \<times># B)"
+ unfolding distinct_mset_count_less_1
+proof (rule allI)
+ fix x
+ have count_mult: "count (A \<times># B) x = count A (fst x) * count B (snd x)"
+ using count_Sigma_mset by (metis prod.exhaust_sel)
+ then have "count A (fst x) * count B (snd x) \<le> 1" using assm1 assm2
+ unfolding distinct_mset_count_less_1 using mult_le_one by blast
+ thus "count (A \<times># B) x \<le> 1" using count_mult by simp
+qed
+
+lemma cart_product_single_intersect: "x1 \<noteq> x2 \<Longrightarrow> ({#x1#} \<times># A) \<inter># ({#x2#} \<times># B) = {#}"
+ using multiset_inter_single by fastforce
+
+lemma size_union_distinct_cart_prod: "x1 \<noteq> x2 \<Longrightarrow> size (({#x1#} \<times># A) \<union># ({#x2#} \<times># B)) =
+ size ({#x1#} \<times># A) + size ({#x2#} \<times># B)"
+ by (simp add: cart_product_single_intersect size_Un_disjoint)
+
+lemma size_Union_distinct_cart_prod: "distinct_mset M \<Longrightarrow>
+ size (\<Sum>p\<in>#M. ({#p#} \<times># B)) = size (M) * size (B)"
+ by (induction M) auto
+
+lemma size_Union_distinct_cart_prod_filter: "distinct_mset M \<Longrightarrow>
+ (\<And> p . p \<in># M \<Longrightarrow> size ({# b \<in># B . P p b #}) = c) \<Longrightarrow>
+ size (\<Sum>p\<in>#M. ({#p#} \<times># {# b \<in># B . P p b #})) = size (M) * c"
+ by (induction M) auto
+
+lemma size_Union_distinct_cart_prod_filter2: "distinct_mset V \<Longrightarrow>
+ (\<And> b . b \<in># B \<Longrightarrow> size ({# v \<in># V . P v b #}) = c) \<Longrightarrow>
+ size (\<Sum>b\<in>#B. ( {# v \<in># V . P v b #} \<times># {#b#})) = size (B) * c"
+ by (induction B) auto
+
+lemma cart_product_add_1: "(add_mset a A) \<times># B = ({#a#} \<times># B) + (A \<times># B)"
+ by (metis Sigma_mset_plus_distrib1 add_mset_add_single union_commute)
+
+lemma cart_product_add_1_filter: "{#m \<in># ((add_mset a M) \<times># N) . P m #} =
+ {#m \<in># (M \<times># N) . P m #} + {#m \<in># ({#a#} \<times># N) . P m #}"
+ unfolding add_mset_add_single [of a M] Sigma_mset_plus_distrib1
+ by (simp add: Times_mset_single_left)
+
+lemma cart_product_add_1_filter2: "{#m \<in># (M \<times># (add_mset b N)) . P m #} =
+ {#m \<in># (M \<times># N) . P m #} + {#m \<in># (M \<times># {#b#}) . P m #}"
+ unfolding add_mset_add_single [of b N] Sigma_mset_plus_distrib1
+ by (metis Times_insert_left Times_mset_single_right add_mset_add_single filter_union_mset)
+
+lemma cart_prod_singleton_right_gen:
+ assumes "\<And> x . x \<in># (A \<times># {#b#}) \<Longrightarrow> P x \<longleftrightarrow> Q (fst x)"
+ shows "{#x \<in># (A \<times># {#b#}). P x#} = {# a \<in># A . Q a#} \<times># {#b#}"
+ using assms
+proof (induction A)
+ case empty
+ then show ?case by simp
+next
+ case (add x A)
+ have "add_mset x A \<times># {#b#} = add_mset (x, b) (A \<times># {#b#})"
+ by (simp add: Times_mset_single_right)
+ then have lhs: "filter_mset P (add_mset x A \<times># {#b#}) = filter_mset P (A \<times># {#b#}) +
+ filter_mset P {#(x, b)#}" by simp
+ have rhs: "filter_mset Q (add_mset x A) \<times># {#b#} = filter_mset Q A \<times># {#b#} +
+ filter_mset Q {#x#} \<times># {#b#}"
+ by (metis Sigma_mset_plus_distrib1 add_mset_add_single filter_union_mset)
+ have "filter_mset P {#(x, b)#} = filter_mset Q {#x#} \<times># {#b#}"
+ using add.prems by fastforce
+ then show ?case using lhs rhs add.IH add.prems by force
+qed
+
+lemma cart_prod_singleton_left_gen:
+ assumes "\<And> x . x \<in># ({#a#} \<times># B) \<Longrightarrow> P x \<longleftrightarrow> Q (snd x)"
+ shows "{#x \<in># ({#a#} \<times># B). P x#} = {#a#} \<times># {#b \<in># B . Q b#}"
+ using assms
+proof (induction B)
+ case empty
+ then show ?case by simp
+next
+ case (add x B)
+ have lhs: "filter_mset P ({#a#} \<times># add_mset x B) = filter_mset P ({#a#} \<times># B) +
+ filter_mset P {#(a, x)#}"
+ by (simp add: cart_product_add_1_filter2)
+ have rhs: "{#a#} \<times># filter_mset Q (add_mset x B) = {#a#} \<times># filter_mset Q B +
+ {#a#} \<times># filter_mset Q {#x#}"
+ using add_mset_add_single filter_union_mset by (metis Times_mset_single_left image_mset_union)
+ have "filter_mset P {#(a, x)#} = {#a#} \<times># filter_mset Q {#x#}"
+ using add.prems by fastforce
+ then show ?case using lhs rhs add.IH add.prems by force
+qed
+
+lemma cart_product_singleton_left: "{#m \<in># ({#a#} \<times># N) . fst m \<in> snd m #} =
+ ({#a#} \<times># {# n \<in># N . a \<in> n #})" (is "?A = ?B")
+proof -
+ have stmt: "\<And>m. m \<in># ({#a#} \<times># N) \<Longrightarrow> fst m \<in> snd m \<longleftrightarrow> a \<in> snd m"
+ by (simp add: mem_Times_mset_iff)
+ thus ?thesis by (metis (no_types, lifting) Sigma_mset_cong stmt cart_prod_singleton_left_gen)
+qed
+
+lemma cart_product_singleton_right: "{#m \<in># (N \<times># {#b#}) . fst m \<in> snd m #} =
+ ({# n \<in># N . n \<in> b #} \<times># {# b #})" (is "?A = ?B")
+proof -
+ have stmt: "\<And>m. m \<in># (N \<times># {#b#}) \<Longrightarrow> fst m \<in> snd m \<longleftrightarrow> fst m \<in>b"
+ by (simp add: mem_Times_mset_iff)
+ thus ?thesis by (metis (no_types, lifting) Sigma_mset_cong stmt cart_prod_singleton_right_gen)
+qed
+
+lemma cart_product_add_1_filter_eq: "{#m \<in># ((add_mset a M) \<times># N) . (fst m \<in> snd m) #} =
+ {#m \<in># (M \<times># N) . (fst m \<in> snd m) #} + ({#a#} \<times># {# n \<in># N . a \<in> n #})"
+ unfolding add_mset_add_single [of a M] Sigma_mset_plus_distrib1
+ using cart_product_singleton_left cart_product_add_1_filter by fastforce
+
+lemma cart_product_add_1_filter_eq_mirror: "{#m \<in># M \<times># (add_mset b N) . (fst m \<in> snd m) #} =
+ {#m \<in># (M \<times># N) . (fst m \<in> snd m) #} + ({# n \<in># M . n \<in> b #} \<times># {#b#})"
+ unfolding add_mset_add_single [of b N] Sigma_mset_plus_distrib1 (* longish *)
+ by (metis (no_types) add_mset_add_single cart_product_add_1_filter2 cart_product_singleton_right)
+
+lemma set_break_down_left:
+ shows "{# m \<in># (M \<times># N) . (fst m) \<in> (snd m) #} = (\<Sum>m\<in>#M. ({#m#} \<times># {#n \<in># N. m \<in> n#}))"
+ by (induction M) (auto simp add: cart_product_add_1_filter_eq)
+
+lemma set_break_down_right:
+ shows "{# x \<in># M \<times># N . (fst x) \<in> (snd x) #} = (\<Sum>n\<in>#N. ({#m \<in># M. m \<in> n#} \<times># {#n#}))"
+ by (induction N) (auto simp add: cart_product_add_1_filter_eq_mirror)
+
+text \<open>Reasoning on sums of elements over multisets\<close>
+
+lemma sum_over_fun_eq:
+ assumes "\<And> x . x \<in># A \<Longrightarrow> f x = g x"
+ shows "(\<Sum>x \<in># A . f(x)) = (\<Sum> x \<in># A . g (x))"
+ using assms by auto
+
+context ring_1
+begin
+
+lemma sum_mset_add_diff: "(\<Sum> x \<in># A. f x - g x) = (\<Sum> x \<in># A . f x) - (\<Sum> x \<in># A . g x)"
+ by (induction A) (auto simp add: algebra_simps)
+
+end
+
+context ordered_ring
+begin
+
+lemma sum_mset_ge0:"(\<And> x . f x \<ge> 0) \<Longrightarrow> (\<Sum> x \<in># A. f x ) \<ge> 0"
+proof (induction A)
+ case empty
+ then show ?case by simp
+next
+ case (add x A)
+ then have hyp2: "0 \<le> sum_mset (image_mset f A)" by blast
+ then have " sum_mset (image_mset f (add_mset x A)) = sum_mset (image_mset f A) + f x"
+ by (simp add: add_commute)
+ then show ?case
+ by (simp add: add.IH add.prems)
+qed
+
+lemma sum_order_add_mset: "(\<And> x . f x \<ge> 0) \<Longrightarrow> (\<Sum> x \<in># A. f x ) \<le> (\<Sum> x \<in># add_mset a A. f x )"
+ by simp
+
+lemma sum_mset_0_left: "(\<And> x . f x \<ge> 0) \<Longrightarrow> (\<Sum> x \<in># A. f x ) = 0 \<Longrightarrow> (\<forall> x \<in># A .f x = 0)"
+ apply (induction A)
+ apply auto
+ using local.add_nonneg_eq_0_iff sum_mset_ge0 apply blast
+ using local.add_nonneg_eq_0_iff sum_mset_ge0 by blast
+
+lemma sum_mset_0_iff_ge_0:
+ assumes "(\<And> x . f x \<ge> 0)"
+ shows "(\<Sum> x \<in># A. f x ) = 0 \<longleftrightarrow> (\<forall> x \<in> set_mset A .f x = 0)"
+ using sum_mset_0_left assms by auto
+
+end
+
+lemma mset_set_size_card_count: "(\<Sum>x \<in># A. x) = (\<Sum>x \<in> set_mset A . x * (count A x))"
+proof (induction A)
+ case empty
+ then show ?case by simp
+next
+ case (add y A)
+ have lhs: "(\<Sum>x\<in>#add_mset y A. x) = (\<Sum>x\<in># A. x) + y" by simp
+ have rhs: "(\<Sum>x\<in>set_mset (add_mset y A). x * count (add_mset y A) x) =
+ (\<Sum>x\<in>(insert y (set_mset A)) . x * count (add_mset y A) x)"
+ by simp
+ then show ?case
+ proof (cases "y \<in># A")
+ case True
+ have x_val: "\<And> x . x \<in> (insert y (set_mset A)) \<Longrightarrow> x \<noteq> y \<Longrightarrow>
+ x* count (add_mset y A) x = x * (count A x)"
+ by auto
+ have y_count: "count (add_mset y A) y = 1 + count A y"
+ using True count_inI by fastforce
+ then have "(\<Sum>x\<in>set_mset (add_mset y A). x * count (add_mset y A) x) =
+ (y * (count (add_mset y A) y)) + (\<Sum>x\<in>(set_mset A) - {y}. x * count A x)"
+ using x_val finite_set_mset sum.cong sum.insert rhs
+ by (smt DiffD1 Diff_insert_absorb insert_absorb mk_disjoint_insert sum.insert_remove)
+ then have s1: "(\<Sum>x\<in>set_mset (add_mset y A). x * count (add_mset y A) x) =
+ y + y * (count A y) + (\<Sum>x\<in>(set_mset A) - {y}. x * count A x)"
+ using y_count by simp
+ then have "(\<Sum>x\<in>set_mset (add_mset y A). x * count (add_mset y A) x) =
+ y + (\<Sum>x\<in>insert y ((set_mset A) - {y} ) . x * count A x)"
+ by (simp add: sum.insert_remove)
+ then have "(\<Sum>x\<in>set_mset (add_mset y A). x * count (add_mset y A) x) =
+ y + (\<Sum>x\<in>(set_mset A) . x * count A x)"
+ by (simp add: True insert_absorb)
+ then show ?thesis using lhs add.IH
+ by linarith
+ next
+ case False
+ have x_val: "\<And> x . x \<in> set_mset A \<Longrightarrow> x* count (add_mset y A) x = x * (count A x)"
+ using False by auto
+ have y_count: "count (add_mset y A) y = 1" using False count_inI by fastforce
+ have lhs: "(\<Sum>x\<in>#add_mset y A. x) = (\<Sum>x\<in># A. x) + y" by simp
+ have "(\<Sum>x\<in>set_mset (add_mset y A). x * count (add_mset y A) x) =
+ (y * count (add_mset y A) y) + (\<Sum>x\<in>set_mset A. x * count A x)"
+ using x_val rhs by (metis (no_types, lifting) False finite_set_mset sum.cong sum.insert)
+ then have "(\<Sum>x\<in>set_mset (add_mset y A). x * count (add_mset y A) x) =
+ y + (\<Sum>x\<in>set_mset A. x * count A x)"
+ using y_count by simp
+ then show ?thesis using lhs add.IH by linarith
+ qed
+qed
+
+subsection \<open>Partitions on Multisets\<close>
+
+text \<open>A partition on a multiset A is a multiset of multisets, where the sum over P equals A and the
+empty multiset is not in the partition. Based off set partition definition.
+We note that unlike set partitions, there is no requirement for elements in the multisets to be
+distinct due to the definition of union on multisets \cite{benderPartitionsMultisets1974}\<close>
+
+lemma mset_size_partition_dep: "size {# a \<in># A . P a \<or> Q a #} =
+ size {# a \<in># A . P a #} + size {# a \<in># A . Q a #} - size {# a \<in># A . P a \<and> Q a #}"
+ by (simp add: mset_bunion_filter mset_inter_filter mset_union_size_inter)
+
+definition partition_on_mset :: "'a multiset \<Rightarrow> 'a multiset multiset \<Rightarrow> bool" where
+"partition_on_mset A P \<longleftrightarrow> \<Sum>\<^sub>#P = A \<and> {#} \<notin># P"
+
+lemma partition_on_msetI [intro]: "\<Sum>\<^sub>#P = A \<Longrightarrow> {#} \<notin># P \<Longrightarrow> partition_on_mset A P"
+ by (simp add: partition_on_mset_def)
+
+lemma partition_on_msetD1: "partition_on_mset A P \<Longrightarrow> \<Sum>\<^sub>#P = A"
+ by (simp add: partition_on_mset_def)
+
+lemma partition_on_msetD2: "partition_on_mset A P \<Longrightarrow> {#} \<notin># P"
+ by (simp add: partition_on_mset_def)
+
+lemma partition_on_mset_empty: "partition_on_mset {#} P \<longleftrightarrow> P = {#}"
+ unfolding partition_on_mset_def
+ using multiset_nonemptyE by fastforce
+
+lemma partition_on_mset_all: "A \<noteq> {#} \<Longrightarrow> partition_on_mset A {#A #}"
+ by (simp add: partition_on_mset_def)
+
+lemma partition_on_mset_singletons: "partition_on_mset A (image_mset (\<lambda> x . {#x#}) A)"
+ by (auto simp: partition_on_mset_def)
+
+lemma partition_on_mset_not_empty: "A \<noteq> {#} \<Longrightarrow> partition_on_mset A P \<Longrightarrow> P \<noteq> {#}"
+ by (auto simp: partition_on_mset_def)
+
+lemma partition_on_msetI2: "\<Sum>\<^sub>#P = A \<Longrightarrow> (\<And> p . p \<in># P \<Longrightarrow> p \<noteq> {#}) \<Longrightarrow> partition_on_mset A P"
+ by (auto simp: partition_on_mset_def)
+
+lemma partition_on_mset_elems: "partition_on_mset A P \<Longrightarrow> p1 \<in># P \<Longrightarrow> x \<in># p1 \<Longrightarrow> x \<in># A"
+ by (auto simp: partition_on_mset_def)
+
+lemma partition_on_mset_sum_size_eq: "partition_on_mset A P \<Longrightarrow> (\<Sum>x \<in># P. size x) = size A"
+ by (metis partition_on_msetD1 size_big_union_sum)
+
+lemma partition_on_mset_card: assumes "partition_on_mset A P" shows " size P \<le> size A"
+proof (rule ccontr)
+ assume "\<not> size P \<le> size A"
+ then have a: "size P > size A" by simp
+ have "\<And> x . x \<in># P \<Longrightarrow> size x > 0" using partition_on_msetD2
+ using assms nonempty_has_size by auto
+ then have " (\<Sum>x \<in># P. size x) \<ge> size P"
+ by (metis leI less_one not_less_zero size_eq_sum_mset sum_mset_mono)
+ thus False using a partition_on_mset_sum_size_eq
+ using assms by fastforce
+qed
+
+lemma partition_on_mset_count_eq: "partition_on_mset A P \<Longrightarrow> a \<in># A \<Longrightarrow>
+ (\<Sum>x \<in># P. count x a) = count A a"
+ by (metis count_sum_mset partition_on_msetD1)
+
+lemma partition_on_mset_subsets: "partition_on_mset A P \<Longrightarrow> x \<in># P \<Longrightarrow> x \<subseteq># A"
+ by (auto simp add: partition_on_mset_def)
+
+lemma partition_on_mset_distinct:
+ assumes "partition_on_mset A P"
+ assumes "distinct_mset A"
+ shows "distinct_mset P"
+proof (rule ccontr)
+ assume "\<not> distinct_mset P"
+ then obtain p1 where count: "count P p1 \<ge> 2"
+ by (metis Suc_1 distinct_mset_count_less_1 less_Suc_eq_le not_less_eq)
+ then have cge: "\<And> x . x \<in># p1 \<Longrightarrow> (\<Sum>p \<in># P. count p x ) \<ge> 2"
+ by (smt count_greater_eq_one_iff count_sum_mset_if_1_0 dual_order.trans sum_mset_mono zero_le)
+ have elem_in: "\<And> x . x \<in># p1 \<Longrightarrow> x \<in># A" using partition_on_mset_elems
+ by (metis count assms(1) count_eq_zero_iff not_numeral_le_zero)
+ have "\<And> x . x \<in># A \<Longrightarrow> count A x = 1" using assms
+ by (simp add: distinct_mset_def)
+ thus False
+ using assms partition_on_mset_count_eq cge elem_in count_inI local.count multiset_nonemptyE
+ by (metis (mono_tags) not_numeral_le_zero numeral_One numeral_le_iff partition_on_mset_def semiring_norm(69))
+qed
+
+lemma partition_on_mset_distinct_disjoint:
+ assumes "partition_on_mset A P"
+ assumes "distinct_mset A"
+ assumes "p1 \<in># P"
+ assumes "p2 \<in># P - {#p1#}"
+ shows "p1 \<inter># p2 = {#}"
+ using Diff_eq_empty_iff_mset assms diff_add_zero distinct_mset_add multiset_inter_assoc sum_mset.remove
+ by (smt partition_on_msetD1 subset_mset.inf.absorb_iff2 subset_mset.le_add_same_cancel1 subset_mset.le_iff_inf)
+
+lemma partition_on_mset_diff:
+ assumes "partition_on_mset A P"
+ assumes "Q \<subseteq>#P"
+ shows "partition_on_mset (A - \<Sum>\<^sub>#Q) (P - Q)"
+ using assms partition_on_mset_def
+ by (smt diff_union_cancelL subset_mset.add_diff_inverse sum_mset.union union_iff)
+
+lemma sigma_over_set_partition_count:
+ assumes "finite A"
+ assumes "partition_on A P"
+ assumes "x \<in># \<Sum>\<^sub># (mset_set (mset_set ` P))"
+ shows "count (\<Sum>\<^sub># (mset_set (mset_set ` P))) x = 1"
+proof -
+ have disj: "disjoint P" using assms partition_onD2 by auto
+ then obtain p where pin: "p \<in># mset_set (mset_set ` P)" and xin: "x \<in># p"
+ using assms by blast
+ then have "count (mset_set (mset_set ` P)) p = 1"
+ by (meson count_eq_zero_iff count_mset_set')
+ then have filter: "\<And> p' . p' \<in># ((mset_set (mset_set` P)) - {#p#}) \<Longrightarrow> p \<noteq> p'"
+ using count_eq_zero_iff count_single by fastforce
+ have zero: "\<And> p'. p' \<in># mset_set (mset_set ` P) \<Longrightarrow> p' \<noteq> p \<Longrightarrow> count p' x = 0"
+ proof (rule ccontr)
+ fix p'
+ assume assm: "p' \<in># mset_set (mset_set ` P)" and ne: "p' \<noteq> p" and n0: "count p' x \<noteq> 0"
+ then have xin2: "x \<in># p'" by auto
+ obtain p1 p2 where p1in: "p1 \<in> P" and p2in: "p2 \<in> P" and p1eq: "mset_set p1 = p"
+ and p2eq: "mset_set p2 = p'" using assm assms(1) assms(2) pin
+ by (metis (no_types, lifting) elem_mset_set finite_elements finite_imageI image_iff)
+ have origne: "p1 \<noteq> p2" using ne p1eq p2eq by auto
+ have "p1 = p2" using partition_onD4 xin xin2
+ by (metis assms(2) count_eq_zero_iff count_mset_set' p1eq p1in p2eq p2in)
+ then show False using origne by simp
+ qed
+ have one: "count p x = 1" using pin xin assms count_eq_zero_iff count_greater_eq_one_iff
+ by (metis count_mset_set(3) count_mset_set_le_one image_iff le_antisym)
+ then have "count (\<Sum>\<^sub># (mset_set (mset_set ` P))) x =
+ (\<Sum>p' \<in># (mset_set (mset_set ` P)) . count p' x)"
+ using count_sum_mset by auto
+ also have "... = (count p x) + (\<Sum>p' \<in># ((mset_set (mset_set ` P)) - {#p#}) . count p' x)"
+ by (metis (mono_tags, lifting) insert_DiffM pin sum_mset.insert)
+ also have "... = 1 + (\<Sum>p' \<in># ((mset_set (mset_set ` P)) - {#p#}) . count p' x)"
+ using one by presburger
+ finally have "count (\<Sum>\<^sub># (mset_set (mset_set ` P))) x =
+ 1 + (\<Sum>p' \<in># ((mset_set (mset_set ` P)) - {#p#}) . 0)"
+ using zero filter by (metis (mono_tags, lifting) in_diffD sum_over_fun_eq)
+ then show "count (\<Sum>\<^sub># (mset_set (mset_set ` P))) x = 1" by simp
+qed
+
+lemma partition_on_mset_set:
+ assumes "finite A"
+ assumes "partition_on A P"
+ shows "partition_on_mset (mset_set A) (mset_set (image (\<lambda> x. mset_set x) P))"
+proof (intro partition_on_msetI)
+ have partd1: "\<Union>P = A" using assms partition_onD1 by auto
+ have imp: "\<And>x. x \<in># \<Sum>\<^sub># (mset_set (mset_set ` P)) \<Longrightarrow> x \<in># mset_set A"
+ proof -
+ fix x
+ assume "x \<in># \<Sum>\<^sub># (mset_set (mset_set ` P))"
+ then obtain p where "p \<in> (mset_set ` P)" and xin: "x \<in># p"
+ by (metis elem_mset_set equals0D infinite_set_mset_mset_set mset_big_union_obtain)
+ then have "set_mset p \<in> P"
+ by (metis empty_iff finite_set_mset_mset_set image_iff infinite_set_mset_mset_set)
+ then show "x \<in># mset_set A"
+ using partd1 xin assms(1) by auto
+ qed
+ have imp2: "\<And>x . x \<in># mset_set A \<Longrightarrow> x \<in># \<Sum>\<^sub># (mset_set (mset_set ` P))"
+ proof -
+ fix x
+ assume "x \<in># mset_set A"
+ then have "x \<in> A" by (simp add: assms(1))
+ then obtain p where "p \<in> P" and "x \<in> p" using assms(2) using partd1 by blast
+ then obtain p' where "p' \<in> (mset_set ` P)" and "p' = mset_set p" by blast
+ thus "x \<in># \<Sum>\<^sub># (mset_set (mset_set ` P))" using assms \<open>p \<in> P\<close> \<open>x \<in> p\<close> finite_elements partd1
+ by (metis Sup_upper finite_imageI finite_set_mset_mset_set in_Union_mset_iff rev_finite_subset)
+ qed
+ have a1: "\<And> x . x \<in># mset_set A \<Longrightarrow> count (mset_set A) x = 1"
+ using assms(1) by fastforce
+ then show "\<Sum>\<^sub># (mset_set (mset_set ` P)) = mset_set A" using imp imp2 a1
+ by (metis assms(1) assms(2) count_eq_zero_iff multiset_eqI sigma_over_set_partition_count)
+ have "\<And> p. p \<in> P \<Longrightarrow> p \<noteq> {} " using assms partition_onD3 by auto
+ then have "\<And> p. p \<in> P \<Longrightarrow> mset_set p \<noteq> {#}" using mset_set_empty_iff
+ by (metis Union_upper assms(1) partd1 rev_finite_subset)
+ then show "{#} \<notin># mset_set (mset_set ` P)"
+ by (metis elem_mset_set equals0D image_iff infinite_set_mset_mset_set)
+qed
+
+lemma partition_on_mset_distinct_inter:
+ assumes "partition_on_mset A P"
+ assumes "distinct_mset A"
+ assumes "p1 \<in># P" and "p2 \<in># P" and "p1 \<noteq> p2"
+ shows "p1 \<inter># p2 = {#}"
+ by (metis assms in_remove1_mset_neq partition_on_mset_distinct_disjoint)
+
+lemma partition_on_set_mset_distinct:
+ assumes "partition_on_mset A P"
+ assumes "distinct_mset A"
+ assumes "p \<in># image_mset set_mset P"
+ assumes "p' \<in># image_mset set_mset P"
+ assumes "p \<noteq> p'"
+ shows "p \<inter> p' = {}"
+proof -
+ obtain p1 where p1in: "p1 \<in># P" and p1eq: "set_mset p1 = p" using assms(3)
+ by blast
+ obtain p2 where p2in: "p2 \<in># P" and p2eq: "set_mset p2 = p'" using assms(4) by blast
+ have "distinct_mset P" using assms partition_on_mset_distinct by blast
+ then have "p1 \<noteq> p2" using assms using p1eq p2eq by fastforce
+ then have "p1 \<inter># p2 = {#}" using partition_on_mset_distinct_inter
+ using assms(1) assms(2) p1in p2in by auto
+ thus ?thesis using p1eq p2eq
+ by (metis set_mset_empty set_mset_inter)
+qed
+
+lemma partition_on_set_mset:
+ assumes "partition_on_mset A P"
+ assumes "distinct_mset A"
+ shows "partition_on (set_mset A) (set_mset (image_mset set_mset P))"
+proof (intro partition_onI)
+ show "\<And>p. p \<in># image_mset set_mset P \<Longrightarrow> p \<noteq> {}"
+ using assms(1) partition_on_msetD2 by fastforce
+next
+ have "\<And> x . x \<in> set_mset A \<Longrightarrow> x \<in> \<Union> (set_mset (image_mset set_mset P))"
+ by (metis Union_iff assms(1) image_eqI mset_big_union_obtain partition_on_msetD1 set_image_mset)
+ then show "\<Union> (set_mset (image_mset set_mset P)) = set_mset A"
+ using set_eqI' partition_on_mset_elems assms by auto
+ show "\<And>p p'. p \<in># image_mset set_mset P \<Longrightarrow> p' \<in># image_mset set_mset P \<Longrightarrow>
+ p \<noteq> p' \<Longrightarrow> p \<inter> p' = {}"
+ using partition_on_set_mset_distinct assms by blast
+qed
+
+lemma partition_on_mset_eq_imp_eq_carrier:
+ assumes "partition_on_mset A P"
+ assumes "partition_on_mset B P"
+ shows "A = B"
+ using assms partition_on_msetD1 by auto
+
+lemma partition_on_mset_add_single:
+ assumes "partition_on_mset A P"
+ shows "partition_on_mset (add_mset a A) (add_mset {#a#} P)"
+ using assms by (auto simp: partition_on_mset_def)
+
+lemma partition_on_mset_add_part:
+ assumes "partition_on_mset A P"
+ assumes "X \<noteq> {#}"
+ assumes "A + X = A'"
+ shows "partition_on_mset A' (add_mset X P)"
+ using assms by (auto simp: partition_on_mset_def)
+
+lemma partition_on_mset_add:
+ assumes "partition_on_mset A P"
+ assumes "X \<in># P"
+ assumes "add_mset a X = X'"
+ shows "partition_on_mset (add_mset a A) (add_mset X' (P - {#X#}))"
+ using add_mset_add_single assms empty_not_add_mset mset_subset_eq_single partition_on_mset_all
+ by (smt partition_on_mset_def subset_mset.add_diff_inverse sum_mset.add_mset sum_mset.remove union_iff union_mset_add_mset_left)
+
+lemma partition_on_mset_elem_exists_part:
+ assumes "partition_on_mset A P"
+ assumes "x \<in># A"
+ obtains p where "p \<in># P" and "x \<in># p"
+ using assms in_Union_mset_iff partition_on_msetD2 partition_on_msetI
+ by (metis partition_on_mset_eq_imp_eq_carrier)
+
+lemma partition_on_mset_combine:
+ assumes "partition_on_mset A P"
+ assumes "partition_on_mset B Q"
+ shows "partition_on_mset (A + B) (P + Q)"
+ unfolding partition_on_mset_def
+ using assms partition_on_msetD1 partition_on_msetD2 by auto
+
+lemma partition_on_mset_split:
+ assumes "partition_on_mset A (P + Q)"
+ shows "partition_on_mset (\<Sum>\<^sub>#P) P"
+ using partition_on_mset_def partition_on_msetD2 assms by fastforce
end
\ No newline at end of file
diff --git a/thys/Design_Theory/Resolvable_Designs.thy b/thys/Design_Theory/Resolvable_Designs.thy
--- a/thys/Design_Theory/Resolvable_Designs.thy
+++ b/thys/Design_Theory/Resolvable_Designs.thy
@@ -1,204 +1,204 @@
-(* Title: Resolvable_Designs.thy
- Author: Chelsea Edmonds
-*)
-
-section \<open>Resolvable Designs\<close>
-text \<open>Resolvable designs have further structure, and can be "resolved" into a set of resolution
-classes. A resolution class is a subset of blocks which exactly partitions the point set.
-Definitions based off the handbook \cite{colbournHandbookCombinatorialDesigns2007}
- and Stinson \cite{stinsonCombinatorialDesignsConstructions2004}.
-This theory includes a proof of an alternate statement of Bose's theorem\<close>
-
-theory Resolvable_Designs imports BIBD
-begin
-
-subsection \<open>Resolutions and Resolution Classes\<close>
-text \<open>A resolution class is a partition of the point set using a set of blocks from the design
-A resolution is a group of resolution classes partitioning the block collection\<close>
-
-context incidence_system
-begin
-
-definition resolution_class :: "'a set set \<Rightarrow> bool" where
-"resolution_class S \<longleftrightarrow> partition_on \<V> S \<and> (\<forall> bl \<in> S . bl \<in># \<B>)"
-
-lemma resolution_classI [intro]: "partition_on \<V> S \<Longrightarrow> (\<And> bl . bl \<in> S \<Longrightarrow> bl \<in># \<B>)
- \<Longrightarrow> resolution_class S"
- by (simp add: resolution_class_def)
-
-lemma resolution_classD1: "resolution_class S \<Longrightarrow> partition_on \<V> S"
- by (simp add: resolution_class_def)
-
-lemma resolution_classD2: "resolution_class S \<Longrightarrow> bl \<in> S \<Longrightarrow> bl \<in># \<B>"
- by (simp add: resolution_class_def)
-
-lemma resolution_class_empty_iff: "resolution_class {} \<longleftrightarrow> \<V> = {}"
- by (auto simp add: resolution_class_def partition_on_def)
-
-lemma resolution_class_complete: "\<V> \<noteq> {} \<Longrightarrow> \<V> \<in># \<B> \<Longrightarrow> resolution_class {\<V>}"
- by (auto simp add: resolution_class_def partition_on_space)
-
-lemma resolution_class_union: "resolution_class S \<Longrightarrow> \<Union>S = \<V> "
- by (simp add: resolution_class_def partition_on_def)
-
-lemma (in finite_incidence_system) resolution_class_finite: "resolution_class S \<Longrightarrow> finite S"
- using finite_elements finite_sets by (auto simp add: resolution_class_def)
-
-lemma (in design) resolution_class_sum_card: "resolution_class S \<Longrightarrow> (\<Sum>bl \<in> S . card bl) = \<v>"
- using resolution_class_union finite_blocks
- by (auto simp add: resolution_class_def partition_on_def card_Union_disjoint)
-
-definition resolution:: "'a set multiset multiset \<Rightarrow> bool" where
-"resolution P \<longleftrightarrow> partition_on_mset \<B> P \<and> (\<forall> S \<in># P . distinct_mset S \<and> resolution_class (set_mset S))"
-
-lemma resolutionI : "partition_on_mset \<B> P \<Longrightarrow> (\<And> S . S \<in>#P \<Longrightarrow> distinct_mset S) \<Longrightarrow>
- (\<And> S . S\<in># P \<Longrightarrow> resolution_class (set_mset S)) \<Longrightarrow> resolution P"
- by (simp add: resolution_def)
-
-lemma (in proper_design) resolution_blocks: "distinct_mset \<B> \<Longrightarrow> disjoint (set_mset \<B>) \<Longrightarrow>
- \<Union>(set_mset \<B>) = \<V> \<Longrightarrow> resolution {#\<B>#}"
- unfolding resolution_def resolution_class_def partition_on_mset_def partition_on_def
- using design_blocks_nempty blocks_nempty by auto
-
-end
-
-subsection \<open>Resolvable Design Locale\<close>
-text \<open>A resolvable design is one with a resolution P\<close>
-locale resolvable_design = design +
- fixes partition :: "'a set multiset multiset" ("\<P>")
- assumes resolvable: "resolution \<P>"
-begin
-
-lemma resolutionD1: "partition_on_mset \<B> \<P>"
- using resolvable by (simp add: resolution_def)
-
-lemma resolutionD2: "S \<in>#\<P> \<Longrightarrow> distinct_mset S"
- using resolvable by (simp add: resolution_def)
-
-lemma resolutionD3: " S\<in># \<P> \<Longrightarrow> resolution_class (set_mset S)"
- using resolvable by (simp add: resolution_def)
-
-lemma resolution_class_blocks_disjoint: "S \<in># \<P> \<Longrightarrow> disjoint (set_mset S)"
- using resolutionD3 by (simp add: partition_on_def resolution_class_def)
-
-lemma resolution_not_empty: "\<B> \<noteq> {#} \<Longrightarrow> \<P> \<noteq> {#}"
- using partition_on_mset_not_empty resolutionD1 by auto
-
-lemma resolution_blocks_subset: "S \<in># \<P> \<Longrightarrow> S \<subseteq># \<B>"
- using partition_on_mset_subsets resolutionD1 by auto
-
-end
-
-lemma (in incidence_system) resolvable_designI [intro]: "resolution \<P> \<Longrightarrow> design \<V> \<B> \<Longrightarrow>
- resolvable_design \<V> \<B> \<P>"
- by (simp add: resolvable_design.intro resolvable_design_axioms.intro)
-
-subsection \<open>Resolvable Block Designs\<close>
-text \<open>An RBIBD is a resolvable BIBD - a common subclass of interest for block designs\<close>
-locale r_block_design = resolvable_design + block_design
-begin
-lemma resolution_class_blocks_constant_size: "S \<in># \<P> \<Longrightarrow> bl \<in># S \<Longrightarrow> card bl = \<k>"
- by (metis resolutionD3 resolution_classD2 uniform_alt_def_all)
-
-lemma resolution_class_size1:
- assumes "S \<in># \<P>"
- shows "\<v> = \<k> * size S"
-proof -
- have "(\<Sum>bl \<in># S . card bl) = (\<Sum>bl \<in> (set_mset S) . card bl)" using resolutionD2 assms
- by (simp add: sum_unfold_sum_mset)
- then have eqv: "(\<Sum>bl \<in># S . card bl) = \<v>" using resolutionD3 assms resolution_class_sum_card
- by presburger
- have "(\<Sum>bl \<in># S . card bl) = (\<Sum>bl \<in># S . \<k>)" using resolution_class_blocks_constant_size assms
- by auto
- thus ?thesis using eqv by (metis mult.commute sum_mset_constant)
-qed
-
-lemma resolution_class_size2:
- assumes "S \<in># \<P>"
- shows "size S = \<v> div \<k>"
- using resolution_class_size1 assms
- by (metis nonzero_mult_div_cancel_left not_one_le_zero k_non_zero)
-
-lemma resolvable_necessary_cond_v: "\<k> dvd \<v>"
-proof -
- obtain S where s_in: "S \<in>#\<P>" using resolution_not_empty design_blocks_nempty by blast
- then have "\<k> * size S = \<v>" using resolution_class_size1 by simp
- thus ?thesis by (metis dvd_triv_left)
-qed
-
-end
-
-locale rbibd = r_block_design + bibd
-
-begin
-
-lemma resolvable_design_num_res_classes: "size \<P> = \<r>"
-proof -
- have k_ne0: "\<k> \<noteq> 0" using k_non_zero by auto
- have f1: "\<b> = (\<Sum>S \<in># \<P> . size S)"
- by (metis partition_on_msetD1 resolutionD1 size_big_union_sum)
- then have "\<b> = (\<Sum>S \<in># \<P> . \<v> div \<k>)" using resolution_class_size2 f1 by auto
- then have f2: "\<b> = (size \<P>) * (\<v> div \<k>)" by simp
- then have "size \<P> = \<b> div (\<v> div \<k>)"
- using b_non_zero by auto
- then have "size \<P> = (\<b> * \<k>) div \<v>" using f2 resolvable_necessary_cond_v
- by (metis div_div_div_same div_dvd_div dvd_triv_right k_ne0 nonzero_mult_div_cancel_right)
- thus ?thesis using necessary_condition_two
- by (metis nonzero_mult_div_cancel_left not_one_less_zero t_design_min_v)
-qed
-
-lemma resolvable_necessary_cond_b: "\<r> dvd \<b>"
-proof -
- have f1: "\<b> = (\<Sum>S \<in># \<P> . size S)"
- by (metis partition_on_msetD1 resolutionD1 size_big_union_sum)
- then have "\<b> = (\<Sum>S \<in># \<P> . \<v> div \<k>)" using resolution_class_size2 f1 by auto
- thus ?thesis using resolvable_design_num_res_classes by simp
-qed
-
-subsubsection \<open>Bose's Inequality\<close>
-text \<open>Boses inequality is an important theorem on RBIBD's. This is a proof
-of an alternate statement of the thm, which does not require a linear algebraic approach,
-taken directly from Stinson \cite{stinsonCombinatorialDesignsConstructions2004}\<close>
-theorem bose_inequality_alternate: "\<b> \<ge> \<v> + \<r> - 1 \<longleftrightarrow> \<r> \<ge> \<k> + \<Lambda>"
-proof -
- have kdvd: "\<k> dvd (\<v> * (\<r> - \<k>))"
- using necessary_condition_two by (simp add: right_diff_distrib')
- have v_eq: "\<v> = (\<r> * (\<k> - 1) + \<Lambda> ) div \<Lambda>"
- using necessary_condition_one index_not_zero by auto
- have ldvd: " \<And> x. \<Lambda> dvd (x * (\<r> * (\<k> - 1) + \<Lambda>))"
- using necessary_condition_one by auto
- have "(\<b> \<ge> \<v> + \<r> - 1) \<longleftrightarrow> ((\<v> * \<r>) div \<k> \<ge> \<v> + \<r> - 1)"
- using necessary_condition_two k_non_zero by auto
- also have "... \<longleftrightarrow> (((\<v> * \<r>) - (\<v> * \<k>)) div \<k> \<ge> \<r> - 1)"
- using k_non_zero div_mult_self3 k_non_zero necessary_condition_two by auto
- also have f2: " ... \<longleftrightarrow> ((\<v> * ( \<r> - \<k>)) \<ge> \<k> * ( \<r> - 1))"
- using k_non_zero kdvd by (auto simp add: int_distrib(3) mult_of_nat_commute)
- also have "... \<longleftrightarrow> ((((\<r> * (\<k> - 1) + \<Lambda> ) div \<Lambda>) * (\<r> - \<k>)) \<ge> \<k> * (\<r> - 1))"
- using v_eq by presburger
- also have "... \<longleftrightarrow> ( (\<r> - \<k>) * ((\<r> * (\<k> - 1) + \<Lambda> ) div \<Lambda>) \<ge> (\<k> * (\<r> - 1)))"
- by (simp add: mult.commute)
- also have " ... \<longleftrightarrow> ( ((\<r> - \<k>) * (\<r> * (\<k> - 1) + \<Lambda> )) div \<Lambda> \<ge> (\<k> * (\<r> - 1)))"
- by (metis div_mult_swap dvd_add_triv_right_iff dvd_triv_left necessary_condition_one)
- also have " ... \<longleftrightarrow> (((\<r> - \<k>) * (\<r> * (\<k> - 1) + \<Lambda> )) \<ge> \<Lambda> * (\<k> * (\<r> - 1)))"
- using ldvd by (smt dvd_mult_div_cancel index_not_zero mult_strict_left_mono)
- also have "... \<longleftrightarrow> (((\<r> - \<k>) * (\<r> * (\<k> - 1))) + ((\<r> - \<k>) * \<Lambda> ) \<ge> \<Lambda> * (\<k> * (\<r> - 1)))"
- by (simp add: distrib_left)
- also have "... \<longleftrightarrow> (((\<r> - \<k>) * \<r> * (\<k> - 1)) \<ge> \<Lambda> * \<k> * (\<r> - 1) - ((\<r> - \<k>) * \<Lambda> ))"
- using mult.assoc by linarith
- also have "... \<longleftrightarrow> (((\<r> - \<k>) * \<r> * (\<k> - 1)) \<ge> (\<Lambda> * \<k> * \<r>) - (\<Lambda> * \<k>) - ((\<r> * \<Lambda>) -(\<k> * \<Lambda> )))"
- using distrib_right by (simp add: distrib_left right_diff_distrib' left_diff_distrib')
- also have "... \<longleftrightarrow> (((\<r> - \<k>) * \<r> * (\<k> - 1)) \<ge> (\<Lambda> * \<k> * \<r>) - ( \<Lambda> * \<r>))"
- by (simp add: mult.commute)
- also have "... \<longleftrightarrow> (((\<r> - \<k>) * \<r> * (\<k> - 1)) \<ge> (\<Lambda> * (\<k> * \<r>)) - ( \<Lambda> * \<r>))"
- by linarith
- also have "... \<longleftrightarrow> (((\<r> - \<k>) * \<r> * (\<k> - 1)) \<ge> (\<Lambda> * (\<r> * \<k>)) - ( \<Lambda> * \<r>))"
- by (simp add: mult.commute)
- also have "... \<longleftrightarrow> (((\<r> - \<k>) * \<r> * (\<k> - 1)) \<ge> \<Lambda> * \<r> * ( \<k> - 1))"
- by (simp add: mult.assoc int_distrib(4))
- finally have "(\<b> \<ge> \<v> + \<r> - 1) \<longleftrightarrow> (\<r> \<ge> \<k> + \<Lambda>)"
- using index_lt_replication mult_right_le_imp_le r_gzero
- by (smt mult_cancel_right k_non_zero)
- thus ?thesis by simp
-qed
-end
+(* Title: Resolvable_Designs.thy
+ Author: Chelsea Edmonds
+*)
+
+section \<open>Resolvable Designs\<close>
+text \<open>Resolvable designs have further structure, and can be "resolved" into a set of resolution
+classes. A resolution class is a subset of blocks which exactly partitions the point set.
+Definitions based off the handbook \cite{colbournHandbookCombinatorialDesigns2007}
+ and Stinson \cite{stinsonCombinatorialDesignsConstructions2004}.
+This theory includes a proof of an alternate statement of Bose's theorem\<close>
+
+theory Resolvable_Designs imports BIBD
+begin
+
+subsection \<open>Resolutions and Resolution Classes\<close>
+text \<open>A resolution class is a partition of the point set using a set of blocks from the design
+A resolution is a group of resolution classes partitioning the block collection\<close>
+
+context incidence_system
+begin
+
+definition resolution_class :: "'a set set \<Rightarrow> bool" where
+"resolution_class S \<longleftrightarrow> partition_on \<V> S \<and> (\<forall> bl \<in> S . bl \<in># \<B>)"
+
+lemma resolution_classI [intro]: "partition_on \<V> S \<Longrightarrow> (\<And> bl . bl \<in> S \<Longrightarrow> bl \<in># \<B>)
+ \<Longrightarrow> resolution_class S"
+ by (simp add: resolution_class_def)
+
+lemma resolution_classD1: "resolution_class S \<Longrightarrow> partition_on \<V> S"
+ by (simp add: resolution_class_def)
+
+lemma resolution_classD2: "resolution_class S \<Longrightarrow> bl \<in> S \<Longrightarrow> bl \<in># \<B>"
+ by (simp add: resolution_class_def)
+
+lemma resolution_class_empty_iff: "resolution_class {} \<longleftrightarrow> \<V> = {}"
+ by (auto simp add: resolution_class_def partition_on_def)
+
+lemma resolution_class_complete: "\<V> \<noteq> {} \<Longrightarrow> \<V> \<in># \<B> \<Longrightarrow> resolution_class {\<V>}"
+ by (auto simp add: resolution_class_def partition_on_space)
+
+lemma resolution_class_union: "resolution_class S \<Longrightarrow> \<Union>S = \<V> "
+ by (simp add: resolution_class_def partition_on_def)
+
+lemma (in finite_incidence_system) resolution_class_finite: "resolution_class S \<Longrightarrow> finite S"
+ using finite_elements finite_sets by (auto simp add: resolution_class_def)
+
+lemma (in design) resolution_class_sum_card: "resolution_class S \<Longrightarrow> (\<Sum>bl \<in> S . card bl) = \<v>"
+ using resolution_class_union finite_blocks
+ by (auto simp add: resolution_class_def partition_on_def card_Union_disjoint)
+
+definition resolution:: "'a set multiset multiset \<Rightarrow> bool" where
+"resolution P \<longleftrightarrow> partition_on_mset \<B> P \<and> (\<forall> S \<in># P . distinct_mset S \<and> resolution_class (set_mset S))"
+
+lemma resolutionI : "partition_on_mset \<B> P \<Longrightarrow> (\<And> S . S \<in>#P \<Longrightarrow> distinct_mset S) \<Longrightarrow>
+ (\<And> S . S\<in># P \<Longrightarrow> resolution_class (set_mset S)) \<Longrightarrow> resolution P"
+ by (simp add: resolution_def)
+
+lemma (in proper_design) resolution_blocks: "distinct_mset \<B> \<Longrightarrow> disjoint (set_mset \<B>) \<Longrightarrow>
+ \<Union>(set_mset \<B>) = \<V> \<Longrightarrow> resolution {#\<B>#}"
+ unfolding resolution_def resolution_class_def partition_on_mset_def partition_on_def
+ using design_blocks_nempty blocks_nempty by auto
+
+end
+
+subsection \<open>Resolvable Design Locale\<close>
+text \<open>A resolvable design is one with a resolution P\<close>
+locale resolvable_design = design +
+ fixes partition :: "'a set multiset multiset" ("\<P>")
+ assumes resolvable: "resolution \<P>"
+begin
+
+lemma resolutionD1: "partition_on_mset \<B> \<P>"
+ using resolvable by (simp add: resolution_def)
+
+lemma resolutionD2: "S \<in>#\<P> \<Longrightarrow> distinct_mset S"
+ using resolvable by (simp add: resolution_def)
+
+lemma resolutionD3: " S\<in># \<P> \<Longrightarrow> resolution_class (set_mset S)"
+ using resolvable by (simp add: resolution_def)
+
+lemma resolution_class_blocks_disjoint: "S \<in># \<P> \<Longrightarrow> disjoint (set_mset S)"
+ using resolutionD3 by (simp add: partition_on_def resolution_class_def)
+
+lemma resolution_not_empty: "\<B> \<noteq> {#} \<Longrightarrow> \<P> \<noteq> {#}"
+ using partition_on_mset_not_empty resolutionD1 by auto
+
+lemma resolution_blocks_subset: "S \<in># \<P> \<Longrightarrow> S \<subseteq># \<B>"
+ using partition_on_mset_subsets resolutionD1 by auto
+
+end
+
+lemma (in incidence_system) resolvable_designI [intro]: "resolution \<P> \<Longrightarrow> design \<V> \<B> \<Longrightarrow>
+ resolvable_design \<V> \<B> \<P>"
+ by (simp add: resolvable_design.intro resolvable_design_axioms.intro)
+
+subsection \<open>Resolvable Block Designs\<close>
+text \<open>An RBIBD is a resolvable BIBD - a common subclass of interest for block designs\<close>
+locale r_block_design = resolvable_design + block_design
+begin
+lemma resolution_class_blocks_constant_size: "S \<in># \<P> \<Longrightarrow> bl \<in># S \<Longrightarrow> card bl = \<k>"
+ by (metis resolutionD3 resolution_classD2 uniform_alt_def_all)
+
+lemma resolution_class_size1:
+ assumes "S \<in># \<P>"
+ shows "\<v> = \<k> * size S"
+proof -
+ have "(\<Sum>bl \<in># S . card bl) = (\<Sum>bl \<in> (set_mset S) . card bl)" using resolutionD2 assms
+ by (simp add: sum_unfold_sum_mset)
+ then have eqv: "(\<Sum>bl \<in># S . card bl) = \<v>" using resolutionD3 assms resolution_class_sum_card
+ by presburger
+ have "(\<Sum>bl \<in># S . card bl) = (\<Sum>bl \<in># S . \<k>)" using resolution_class_blocks_constant_size assms
+ by auto
+ thus ?thesis using eqv by (metis mult.commute sum_mset_constant)
+qed
+
+lemma resolution_class_size2:
+ assumes "S \<in># \<P>"
+ shows "size S = \<v> div \<k>"
+ using resolution_class_size1 assms
+ by (metis nonzero_mult_div_cancel_left not_one_le_zero k_non_zero)
+
+lemma resolvable_necessary_cond_v: "\<k> dvd \<v>"
+proof -
+ obtain S where s_in: "S \<in>#\<P>" using resolution_not_empty design_blocks_nempty by blast
+ then have "\<k> * size S = \<v>" using resolution_class_size1 by simp
+ thus ?thesis by (metis dvd_triv_left)
+qed
+
+end
+
+locale rbibd = r_block_design + bibd
+
+begin
+
+lemma resolvable_design_num_res_classes: "size \<P> = \<r>"
+proof -
+ have k_ne0: "\<k> \<noteq> 0" using k_non_zero by auto
+ have f1: "\<b> = (\<Sum>S \<in># \<P> . size S)"
+ by (metis partition_on_msetD1 resolutionD1 size_big_union_sum)
+ then have "\<b> = (\<Sum>S \<in># \<P> . \<v> div \<k>)" using resolution_class_size2 f1 by auto
+ then have f2: "\<b> = (size \<P>) * (\<v> div \<k>)" by simp
+ then have "size \<P> = \<b> div (\<v> div \<k>)"
+ using b_non_zero by auto
+ then have "size \<P> = (\<b> * \<k>) div \<v>" using f2 resolvable_necessary_cond_v
+ by (metis div_div_div_same div_dvd_div dvd_triv_right k_ne0 nonzero_mult_div_cancel_right)
+ thus ?thesis using necessary_condition_two
+ by (metis nonzero_mult_div_cancel_left not_one_less_zero t_design_min_v)
+qed
+
+lemma resolvable_necessary_cond_b: "\<r> dvd \<b>"
+proof -
+ have f1: "\<b> = (\<Sum>S \<in># \<P> . size S)"
+ by (metis partition_on_msetD1 resolutionD1 size_big_union_sum)
+ then have "\<b> = (\<Sum>S \<in># \<P> . \<v> div \<k>)" using resolution_class_size2 f1 by auto
+ thus ?thesis using resolvable_design_num_res_classes by simp
+qed
+
+subsubsection \<open>Bose's Inequality\<close>
+text \<open>Boses inequality is an important theorem on RBIBD's. This is a proof
+of an alternate statement of the thm, which does not require a linear algebraic approach,
+taken directly from Stinson \cite{stinsonCombinatorialDesignsConstructions2004}\<close>
+theorem bose_inequality_alternate: "\<b> \<ge> \<v> + \<r> - 1 \<longleftrightarrow> \<r> \<ge> \<k> + \<Lambda>"
+proof -
+ have kdvd: "\<k> dvd (\<v> * (\<r> - \<k>))"
+ using necessary_condition_two by (simp add: right_diff_distrib')
+ have v_eq: "\<v> = (\<r> * (\<k> - 1) + \<Lambda> ) div \<Lambda>"
+ using necessary_condition_one index_not_zero by auto
+ have ldvd: " \<And> x. \<Lambda> dvd (x * (\<r> * (\<k> - 1) + \<Lambda>))"
+ using necessary_condition_one by auto
+ have "(\<b> \<ge> \<v> + \<r> - 1) \<longleftrightarrow> ((\<v> * \<r>) div \<k> \<ge> \<v> + \<r> - 1)"
+ using necessary_condition_two k_non_zero by auto
+ also have "... \<longleftrightarrow> (((\<v> * \<r>) - (\<v> * \<k>)) div \<k> \<ge> \<r> - 1)"
+ using k_non_zero div_mult_self3 k_non_zero necessary_condition_two by auto
+ also have f2: " ... \<longleftrightarrow> ((\<v> * ( \<r> - \<k>)) \<ge> \<k> * ( \<r> - 1))"
+ using k_non_zero kdvd by (auto simp add: int_distrib(3) mult_of_nat_commute)
+ also have "... \<longleftrightarrow> ((((\<r> * (\<k> - 1) + \<Lambda> ) div \<Lambda>) * (\<r> - \<k>)) \<ge> \<k> * (\<r> - 1))"
+ using v_eq by presburger
+ also have "... \<longleftrightarrow> ( (\<r> - \<k>) * ((\<r> * (\<k> - 1) + \<Lambda> ) div \<Lambda>) \<ge> (\<k> * (\<r> - 1)))"
+ by (simp add: mult.commute)
+ also have " ... \<longleftrightarrow> ( ((\<r> - \<k>) * (\<r> * (\<k> - 1) + \<Lambda> )) div \<Lambda> \<ge> (\<k> * (\<r> - 1)))"
+ by (metis div_mult_swap dvd_add_triv_right_iff dvd_triv_left necessary_condition_one)
+ also have " ... \<longleftrightarrow> (((\<r> - \<k>) * (\<r> * (\<k> - 1) + \<Lambda> )) \<ge> \<Lambda> * (\<k> * (\<r> - 1)))"
+ using ldvd by (smt dvd_mult_div_cancel index_not_zero mult_strict_left_mono)
+ also have "... \<longleftrightarrow> (((\<r> - \<k>) * (\<r> * (\<k> - 1))) + ((\<r> - \<k>) * \<Lambda> ) \<ge> \<Lambda> * (\<k> * (\<r> - 1)))"
+ by (simp add: distrib_left)
+ also have "... \<longleftrightarrow> (((\<r> - \<k>) * \<r> * (\<k> - 1)) \<ge> \<Lambda> * \<k> * (\<r> - 1) - ((\<r> - \<k>) * \<Lambda> ))"
+ using mult.assoc by linarith
+ also have "... \<longleftrightarrow> (((\<r> - \<k>) * \<r> * (\<k> - 1)) \<ge> (\<Lambda> * \<k> * \<r>) - (\<Lambda> * \<k>) - ((\<r> * \<Lambda>) -(\<k> * \<Lambda> )))"
+ using distrib_right by (simp add: distrib_left right_diff_distrib' left_diff_distrib')
+ also have "... \<longleftrightarrow> (((\<r> - \<k>) * \<r> * (\<k> - 1)) \<ge> (\<Lambda> * \<k> * \<r>) - ( \<Lambda> * \<r>))"
+ by (simp add: mult.commute)
+ also have "... \<longleftrightarrow> (((\<r> - \<k>) * \<r> * (\<k> - 1)) \<ge> (\<Lambda> * (\<k> * \<r>)) - ( \<Lambda> * \<r>))"
+ by linarith
+ also have "... \<longleftrightarrow> (((\<r> - \<k>) * \<r> * (\<k> - 1)) \<ge> (\<Lambda> * (\<r> * \<k>)) - ( \<Lambda> * \<r>))"
+ by (simp add: mult.commute)
+ also have "... \<longleftrightarrow> (((\<r> - \<k>) * \<r> * (\<k> - 1)) \<ge> \<Lambda> * \<r> * ( \<k> - 1))"
+ by (simp add: mult.assoc int_distrib(4))
+ finally have "(\<b> \<ge> \<v> + \<r> - 1) \<longleftrightarrow> (\<r> \<ge> \<k> + \<Lambda>)"
+ using index_lt_replication mult_right_le_imp_le r_gzero
+ by (smt mult_cancel_right k_non_zero)
+ thus ?thesis by simp
+qed
+end
end
\ No newline at end of file
diff --git a/thys/Design_Theory/Sub_Designs.thy b/thys/Design_Theory/Sub_Designs.thy
--- a/thys/Design_Theory/Sub_Designs.thy
+++ b/thys/Design_Theory/Sub_Designs.thy
@@ -1,202 +1,202 @@
-(* Title: Sub_Designs.thy
- Author: Chelsea Edmonds
-*)
-
-section \<open>Sub-designs\<close>
-text \<open>Sub designs are a relationship between two designs using the subset and submultiset relations
-This theory defines the concept at the incidence system level, before extending to defining on
-well defined designs\<close>
-
-theory Sub_Designs imports Design_Operations
-begin
-
-subsection \<open>Sub-system and Sub-design Locales\<close>
-locale sub_set_system = incidence_system \<V> \<B>
- for "\<U>" and "\<A>" and "\<V>" and "\<B>" +
- assumes points_subset: "\<U> \<subseteq> \<V>"
- assumes blocks_subset: "\<A> \<subseteq># \<B>"
-begin
-
-lemma sub_points: "x \<in> \<U> \<Longrightarrow> x \<in> \<V>"
- using points_subset by auto
-
-lemma sub_blocks: "bl \<in># \<A> \<Longrightarrow> bl \<in># \<B>"
- using blocks_subset by auto
-
-lemma sub_blocks_count: "count \<A> b \<le> count \<B> b"
- by (simp add: mset_subset_eq_count blocks_subset)
-
-end
-
-locale sub_incidence_system = sub_set_system + ins: incidence_system \<U> \<A>
-
-locale sub_design = sub_incidence_system + des: design \<V> \<B>
-begin
-
-lemma sub_non_empty_blocks: "A \<in># \<A> \<Longrightarrow> A \<noteq> {}"
- using des.blocks_nempty sub_blocks by simp
-
-sublocale sub_des: design \<U> \<A>
- using des.finite_sets finite_subset points_subset sub_non_empty_blocks
- by (unfold_locales) (auto)
-
-end
-
-locale proper_sub_set_system = incidence_system \<V> \<B>
- for "\<U>" and "\<A>" and "\<V>" and "\<B>" +
- assumes points_psubset: "\<U> \<subset> \<V>"
- assumes blocks_subset: "\<A> \<subseteq># \<B>"
-begin
-
-lemma point_sets_ne: "\<U> \<noteq> \<V>"
- using points_psubset by auto
-
-end
-
-sublocale proper_sub_set_system \<subseteq> sub_set_system
- using points_psubset blocks_subset by (unfold_locales) simp_all
-
-context sub_set_system
-begin
-
-lemma sub_is_proper: "\<U> \<noteq> \<V> \<Longrightarrow> proper_sub_set_system \<U> \<A> \<V> \<B>"
- using blocks_subset incidence_system_axioms
- by (simp add: points_subset proper_sub_set_system.intro proper_sub_set_system_axioms_def psubsetI)
-
-end
-
-locale proper_sub_incidence_system = proper_sub_set_system + ins: incidence_system \<U> \<A>
-
-sublocale proper_sub_incidence_system \<subseteq> sub_incidence_system
- by (unfold_locales)
-
-context sub_incidence_system
-begin
-lemma sub_is_proper: "\<U> \<noteq> \<V> \<Longrightarrow> proper_sub_incidence_system \<U> \<A> \<V> \<B>"
- by (simp add: ins.incidence_system_axioms proper_sub_incidence_system_def sub_is_proper)
-
-end
-
-locale proper_sub_design = proper_sub_incidence_system + des: design \<V> \<B>
-
-sublocale proper_sub_design \<subseteq> sub_design
- by (unfold_locales)
-
-context sub_design
-begin
-lemma sub_is_proper: "\<U> \<noteq> \<V> \<Longrightarrow> proper_sub_design \<U> \<A> \<V> \<B>"
- by (simp add: des.wf_design proper_sub_design.intro sub_is_proper)
-
-end
-
-lemma ss_proper_implies_sub [intro]: "proper_sub_set_system \<U> \<A> \<V> \<B> \<Longrightarrow> sub_set_system \<U> \<A> \<V> \<B>"
- using proper_sub_set_system.axioms(1) proper_sub_set_system.blocks_subset psubsetE
- by (metis proper_sub_set_system.points_psubset sub_set_system.intro sub_set_system_axioms_def)
-
-lemma sub_ssI [intro!]: "incidence_system \<V> \<B> \<Longrightarrow> \<U> \<subseteq> \<V> \<Longrightarrow> \<A> \<subseteq># \<B> \<Longrightarrow> sub_set_system \<U> \<A> \<V> \<B>"
- using incidence_system_def subset_iff
- by (unfold_locales) (simp_all add: incidence_system.wellformed)
-
-lemma sub_ss_equality:
- assumes "sub_set_system \<U> \<A> \<V> \<B>"
- and "sub_set_system \<V> \<B> \<U> \<A>"
- shows "\<U> = \<V>" and "\<A> = \<B>"
- using assms(1) assms(2) sub_set_system.points_subset apply blast
- by (meson assms(1) assms(2) sub_set_system.blocks_subset subset_mset.eq_iff)
-
-subsection \<open>Reasoning on Sub-designs\<close>
-
-subsubsection \<open>Reasoning on Incidence Sys property relationships\<close>
-
-context sub_incidence_system
-begin
-
-lemma sub_sys_block_sizes: "ins.sys_block_sizes \<subseteq> sys_block_sizes"
- by (auto simp add: sys_block_sizes_def ins.sys_block_sizes_def blocks_subset sub_blocks)
-
-lemma sub_point_rep_number_le: "x \<in> \<U> \<Longrightarrow> \<A> rep x \<le> \<B> rep x"
- by (simp add: point_replication_number_def blocks_subset multiset_filter_mono size_mset_mono)
-
-lemma sub_point_index_le: "ps \<subseteq> \<U> \<Longrightarrow> \<A> index ps \<le> \<B> index ps"
- by (simp add: points_index_def blocks_subset multiset_filter_mono size_mset_mono)
-
-lemma sub_sys_intersection_numbers: "ins.intersection_numbers \<subseteq> intersection_numbers"
- apply (auto simp add: intersection_numbers_def ins.intersection_numbers_def)
- by (metis blocks_subset insert_DiffM insert_subset_eq_iff)
-
-end
-
-subsubsection \<open>Reasoning on Incidence Sys/Design operations\<close>
-context incidence_system
-begin
-
-lemma sub_set_sysI[intro]: "\<U> \<subseteq> \<V> \<Longrightarrow> \<A> \<subseteq># \<B> \<Longrightarrow> sub_set_system \<U> \<A> \<V> \<B>"
- by (simp add: sub_ssI incidence_system_axioms)
-
-lemma sub_inc_sysI[intro]: "incidence_system \<U> \<A> \<Longrightarrow> \<U> \<subseteq> \<V> \<Longrightarrow> \<A> \<subseteq># \<B> \<Longrightarrow>
- sub_incidence_system \<U> \<A> \<V> \<B>"
- by (simp add: sub_incidence_system.intro sub_set_sysI)
-
-lemma multiple_orig_sub_system:
- assumes "n > 0"
- shows "sub_incidence_system \<V> \<B> \<V> (multiple_blocks n)"
- using multiple_block_in_original wellformed multiple_blocks_sub assms
- by (unfold_locales) simp_all
-
-lemma add_point_sub_sys: "sub_incidence_system \<V> \<B> (add_point p) \<B>"
- using add_point_wf add_point_def
- by (simp add: sub_ssI subset_insertI incidence_system_axioms sub_incidence_system.intro)
-
-lemma strong_del_point_sub_sys: "sub_incidence_system (del_point p) (str_del_point_blocks p) \<V> \<B> "
- using strong_del_point_incidence_wf wf_invalid_point del_point_def str_del_point_blocks_def
- by (unfold_locales) (auto)
-
-lemma add_block_sub_sys: "sub_incidence_system \<V> \<B> (\<V> \<union> b) (add_block b)"
- using add_block_wf wf_invalid_point add_block_def by (unfold_locales) (auto)
-
-lemma delete_block_sub_sys: "sub_incidence_system \<V> (del_block b) \<V> \<B> "
- using delete_block_wf delete_block_subset incidence_system_def by (unfold_locales, auto)
-
-end
-
-lemma (in two_set_systems) combine_sub_sys: "sub_incidence_system \<V> \<B> \<V>\<^sup>+ \<B>\<^sup>+"
- by (unfold_locales) (simp_all)
-
-lemma (in two_set_systems) combine_sub_sys_alt: "sub_incidence_system \<V>' \<B>' \<V>\<^sup>+ \<B>\<^sup>+"
- by (unfold_locales) (simp_all)
-
-context design
-begin
-
-lemma sub_designI [intro]: "design \<U> \<A> \<Longrightarrow> sub_incidence_system \<U> \<A> \<V> \<B> \<Longrightarrow> sub_design \<U> \<A> \<V> \<B>"
- by (simp add: sub_design.intro wf_design)
-
-lemma sub_designII [intro]: "design \<U> \<A> \<Longrightarrow> sub_incidence_system \<V> \<B> \<U> \<A> \<Longrightarrow> sub_design \<V> \<B> \<U> \<A>"
- by (simp add: sub_design_def)
-
-lemma multiple_orig_sub_des:
- assumes "n > 0"
- shows "sub_design \<V> \<B> \<V> (multiple_blocks n)"
- using multiple_is_design assms multiple_orig_sub_system by (simp add: sub_design.intro)
-
-lemma add_point_sub_des: "sub_design \<V> \<B> (add_point p) \<B>"
- using add_point_design add_point_sub_sys sub_design.intro by fastforce
-
-lemma strong_del_point_sub_des: "sub_design (del_point p) (str_del_point_blocks p) \<V> \<B> "
- using strong_del_point_sub_sys sub_design.intro wf_design by blast
-
-lemma add_block_sub_des: "finite b \<Longrightarrow> b \<noteq> {} \<Longrightarrow> sub_design \<V> \<B> (\<V> \<union> b) (add_block b)"
- using add_block_sub_sys add_block_design sub_designII by fastforce
-
-lemma delete_block_sub_des: "sub_design \<V> (del_block b) \<V> \<B> "
- using delete_block_design delete_block_sub_sys sub_designI by auto
-
-end
-
-lemma (in two_designs) combine_sub_des: "sub_design \<V> \<B> \<V>\<^sup>+ \<B>\<^sup>+"
- by (unfold_locales) (simp_all)
-
-lemma (in two_designs) combine_sub_des_alt: "sub_design \<V>' \<B>' \<V>\<^sup>+ \<B>\<^sup>+"
- by (unfold_locales) (simp_all)
-
+(* Title: Sub_Designs.thy
+ Author: Chelsea Edmonds
+*)
+
+section \<open>Sub-designs\<close>
+text \<open>Sub designs are a relationship between two designs using the subset and submultiset relations
+This theory defines the concept at the incidence system level, before extending to defining on
+well defined designs\<close>
+
+theory Sub_Designs imports Design_Operations
+begin
+
+subsection \<open>Sub-system and Sub-design Locales\<close>
+locale sub_set_system = incidence_system \<V> \<B>
+ for "\<U>" and "\<A>" and "\<V>" and "\<B>" +
+ assumes points_subset: "\<U> \<subseteq> \<V>"
+ assumes blocks_subset: "\<A> \<subseteq># \<B>"
+begin
+
+lemma sub_points: "x \<in> \<U> \<Longrightarrow> x \<in> \<V>"
+ using points_subset by auto
+
+lemma sub_blocks: "bl \<in># \<A> \<Longrightarrow> bl \<in># \<B>"
+ using blocks_subset by auto
+
+lemma sub_blocks_count: "count \<A> b \<le> count \<B> b"
+ by (simp add: mset_subset_eq_count blocks_subset)
+
+end
+
+locale sub_incidence_system = sub_set_system + ins: incidence_system \<U> \<A>
+
+locale sub_design = sub_incidence_system + des: design \<V> \<B>
+begin
+
+lemma sub_non_empty_blocks: "A \<in># \<A> \<Longrightarrow> A \<noteq> {}"
+ using des.blocks_nempty sub_blocks by simp
+
+sublocale sub_des: design \<U> \<A>
+ using des.finite_sets finite_subset points_subset sub_non_empty_blocks
+ by (unfold_locales) (auto)
+
+end
+
+locale proper_sub_set_system = incidence_system \<V> \<B>
+ for "\<U>" and "\<A>" and "\<V>" and "\<B>" +
+ assumes points_psubset: "\<U> \<subset> \<V>"
+ assumes blocks_subset: "\<A> \<subseteq># \<B>"
+begin
+
+lemma point_sets_ne: "\<U> \<noteq> \<V>"
+ using points_psubset by auto
+
+end
+
+sublocale proper_sub_set_system \<subseteq> sub_set_system
+ using points_psubset blocks_subset by (unfold_locales) simp_all
+
+context sub_set_system
+begin
+
+lemma sub_is_proper: "\<U> \<noteq> \<V> \<Longrightarrow> proper_sub_set_system \<U> \<A> \<V> \<B>"
+ using blocks_subset incidence_system_axioms
+ by (simp add: points_subset proper_sub_set_system.intro proper_sub_set_system_axioms_def psubsetI)
+
+end
+
+locale proper_sub_incidence_system = proper_sub_set_system + ins: incidence_system \<U> \<A>
+
+sublocale proper_sub_incidence_system \<subseteq> sub_incidence_system
+ by (unfold_locales)
+
+context sub_incidence_system
+begin
+lemma sub_is_proper: "\<U> \<noteq> \<V> \<Longrightarrow> proper_sub_incidence_system \<U> \<A> \<V> \<B>"
+ by (simp add: ins.incidence_system_axioms proper_sub_incidence_system_def sub_is_proper)
+
+end
+
+locale proper_sub_design = proper_sub_incidence_system + des: design \<V> \<B>
+
+sublocale proper_sub_design \<subseteq> sub_design
+ by (unfold_locales)
+
+context sub_design
+begin
+lemma sub_is_proper: "\<U> \<noteq> \<V> \<Longrightarrow> proper_sub_design \<U> \<A> \<V> \<B>"
+ by (simp add: des.wf_design proper_sub_design.intro sub_is_proper)
+
+end
+
+lemma ss_proper_implies_sub [intro]: "proper_sub_set_system \<U> \<A> \<V> \<B> \<Longrightarrow> sub_set_system \<U> \<A> \<V> \<B>"
+ using proper_sub_set_system.axioms(1) proper_sub_set_system.blocks_subset psubsetE
+ by (metis proper_sub_set_system.points_psubset sub_set_system.intro sub_set_system_axioms_def)
+
+lemma sub_ssI [intro!]: "incidence_system \<V> \<B> \<Longrightarrow> \<U> \<subseteq> \<V> \<Longrightarrow> \<A> \<subseteq># \<B> \<Longrightarrow> sub_set_system \<U> \<A> \<V> \<B>"
+ using incidence_system_def subset_iff
+ by (unfold_locales) (simp_all add: incidence_system.wellformed)
+
+lemma sub_ss_equality:
+ assumes "sub_set_system \<U> \<A> \<V> \<B>"
+ and "sub_set_system \<V> \<B> \<U> \<A>"
+ shows "\<U> = \<V>" and "\<A> = \<B>"
+ using assms(1) assms(2) sub_set_system.points_subset apply blast
+ by (meson assms(1) assms(2) sub_set_system.blocks_subset subset_mset.eq_iff)
+
+subsection \<open>Reasoning on Sub-designs\<close>
+
+subsubsection \<open>Reasoning on Incidence Sys property relationships\<close>
+
+context sub_incidence_system
+begin
+
+lemma sub_sys_block_sizes: "ins.sys_block_sizes \<subseteq> sys_block_sizes"
+ by (auto simp add: sys_block_sizes_def ins.sys_block_sizes_def blocks_subset sub_blocks)
+
+lemma sub_point_rep_number_le: "x \<in> \<U> \<Longrightarrow> \<A> rep x \<le> \<B> rep x"
+ by (simp add: point_replication_number_def blocks_subset multiset_filter_mono size_mset_mono)
+
+lemma sub_point_index_le: "ps \<subseteq> \<U> \<Longrightarrow> \<A> index ps \<le> \<B> index ps"
+ by (simp add: points_index_def blocks_subset multiset_filter_mono size_mset_mono)
+
+lemma sub_sys_intersection_numbers: "ins.intersection_numbers \<subseteq> intersection_numbers"
+ apply (auto simp add: intersection_numbers_def ins.intersection_numbers_def)
+ by (metis blocks_subset insert_DiffM insert_subset_eq_iff)
+
+end
+
+subsubsection \<open>Reasoning on Incidence Sys/Design operations\<close>
+context incidence_system
+begin
+
+lemma sub_set_sysI[intro]: "\<U> \<subseteq> \<V> \<Longrightarrow> \<A> \<subseteq># \<B> \<Longrightarrow> sub_set_system \<U> \<A> \<V> \<B>"
+ by (simp add: sub_ssI incidence_system_axioms)
+
+lemma sub_inc_sysI[intro]: "incidence_system \<U> \<A> \<Longrightarrow> \<U> \<subseteq> \<V> \<Longrightarrow> \<A> \<subseteq># \<B> \<Longrightarrow>
+ sub_incidence_system \<U> \<A> \<V> \<B>"
+ by (simp add: sub_incidence_system.intro sub_set_sysI)
+
+lemma multiple_orig_sub_system:
+ assumes "n > 0"
+ shows "sub_incidence_system \<V> \<B> \<V> (multiple_blocks n)"
+ using multiple_block_in_original wellformed multiple_blocks_sub assms
+ by (unfold_locales) simp_all
+
+lemma add_point_sub_sys: "sub_incidence_system \<V> \<B> (add_point p) \<B>"
+ using add_point_wf add_point_def
+ by (simp add: sub_ssI subset_insertI incidence_system_axioms sub_incidence_system.intro)
+
+lemma strong_del_point_sub_sys: "sub_incidence_system (del_point p) (str_del_point_blocks p) \<V> \<B> "
+ using strong_del_point_incidence_wf wf_invalid_point del_point_def str_del_point_blocks_def
+ by (unfold_locales) (auto)
+
+lemma add_block_sub_sys: "sub_incidence_system \<V> \<B> (\<V> \<union> b) (add_block b)"
+ using add_block_wf wf_invalid_point add_block_def by (unfold_locales) (auto)
+
+lemma delete_block_sub_sys: "sub_incidence_system \<V> (del_block b) \<V> \<B> "
+ using delete_block_wf delete_block_subset incidence_system_def by (unfold_locales, auto)
+
+end
+
+lemma (in two_set_systems) combine_sub_sys: "sub_incidence_system \<V> \<B> \<V>\<^sup>+ \<B>\<^sup>+"
+ by (unfold_locales) (simp_all)
+
+lemma (in two_set_systems) combine_sub_sys_alt: "sub_incidence_system \<V>' \<B>' \<V>\<^sup>+ \<B>\<^sup>+"
+ by (unfold_locales) (simp_all)
+
+context design
+begin
+
+lemma sub_designI [intro]: "design \<U> \<A> \<Longrightarrow> sub_incidence_system \<U> \<A> \<V> \<B> \<Longrightarrow> sub_design \<U> \<A> \<V> \<B>"
+ by (simp add: sub_design.intro wf_design)
+
+lemma sub_designII [intro]: "design \<U> \<A> \<Longrightarrow> sub_incidence_system \<V> \<B> \<U> \<A> \<Longrightarrow> sub_design \<V> \<B> \<U> \<A>"
+ by (simp add: sub_design_def)
+
+lemma multiple_orig_sub_des:
+ assumes "n > 0"
+ shows "sub_design \<V> \<B> \<V> (multiple_blocks n)"
+ using multiple_is_design assms multiple_orig_sub_system by (simp add: sub_design.intro)
+
+lemma add_point_sub_des: "sub_design \<V> \<B> (add_point p) \<B>"
+ using add_point_design add_point_sub_sys sub_design.intro by fastforce
+
+lemma strong_del_point_sub_des: "sub_design (del_point p) (str_del_point_blocks p) \<V> \<B> "
+ using strong_del_point_sub_sys sub_design.intro wf_design by blast
+
+lemma add_block_sub_des: "finite b \<Longrightarrow> b \<noteq> {} \<Longrightarrow> sub_design \<V> \<B> (\<V> \<union> b) (add_block b)"
+ using add_block_sub_sys add_block_design sub_designII by fastforce
+
+lemma delete_block_sub_des: "sub_design \<V> (del_block b) \<V> \<B> "
+ using delete_block_design delete_block_sub_sys sub_designI by auto
+
+end
+
+lemma (in two_designs) combine_sub_des: "sub_design \<V> \<B> \<V>\<^sup>+ \<B>\<^sup>+"
+ by (unfold_locales) (simp_all)
+
+lemma (in two_designs) combine_sub_des_alt: "sub_design \<V>' \<B>' \<V>\<^sup>+ \<B>\<^sup>+"
+ by (unfold_locales) (simp_all)
+
end
\ No newline at end of file
diff --git a/thys/Isabelle_Marries_Dirac/Deutsch.thy b/thys/Isabelle_Marries_Dirac/Deutsch.thy
--- a/thys/Isabelle_Marries_Dirac/Deutsch.thy
+++ b/thys/Isabelle_Marries_Dirac/Deutsch.thy
@@ -1,526 +1,526 @@
-(*
-Authors:
-
- Hanna Lachnitt, TU Wien, lachnitt@student.tuwien.ac.at
- Anthony Bordg, University of Cambridge, apdb3@cam.ac.uk
-*)
-
-section \<open>The Deutsch Algorithm\<close>
-
-theory Deutsch
-imports
- More_Tensor
- Measurement
-begin
-
-
-text \<open>
-Given a function $f:{0,1}\mapsto {0,1}$, Deutsch's algorithm decides if this function is constant
-or balanced with a single $f(x)$ circuit to evaluate the function for multiple values of $x$
-simultaneously. The algorithm makes use of quantum parallelism and quantum interference.
-\<close>
-
-text \<open>
-A constant function with values in {0,1} returns either always 0 or always 1.
-A balanced function is 0 for half of the inputs and 1 for the other half.
-\<close>
-
-locale deutsch =
- fixes f:: "nat \<Rightarrow> nat"
- assumes dom: "f \<in> ({0,1} \<rightarrow>\<^sub>E {0,1})"
-
-context deutsch
-begin
-
-definition is_swap:: bool where
-"is_swap = (\<forall>x \<in> {0,1}. f x = 1 - x)"
-
-lemma is_swap_values:
- assumes "is_swap"
- shows "f 0 = 1" and "f 1 = 0"
- using assms is_swap_def by auto
-
-lemma is_swap_sum_mod_2:
- assumes "is_swap"
- shows "(f 0 + f 1) mod 2 = 1"
- using assms is_swap_def by simp
-
-definition const:: "nat \<Rightarrow> bool" where
-"const n = (\<forall>x \<in> {0,1}.(f x = n))"
-
-definition is_const:: "bool" where
-"is_const \<equiv> const 0 \<or> const 1"
-
-definition is_balanced:: "bool" where
-"is_balanced \<equiv> (\<forall>x \<in> {0,1}.(f x = x)) \<or> is_swap"
-
-lemma f_values: "(f 0 = 0 \<or> f 0 = 1) \<and> (f 1 = 0 \<or> f 1 = 1)"
- using dom by auto
-
-lemma f_cases:
- shows "is_const \<or> is_balanced"
- using dom is_balanced_def const_def is_const_def is_swap_def f_values by auto
-
-lemma const_0_sum_mod_2:
- assumes "const 0"
- shows "(f 0 + f 1) mod 2 = 0"
- using assms const_def by simp
-
-lemma const_1_sum_mod_2:
- assumes "const 1"
- shows "(f 0 + f 1) mod 2 = 0"
- using assms const_def by simp
-
-lemma is_const_sum_mod_2:
- assumes "is_const"
- shows "(f 0 + f 1) mod 2 = 0"
- using assms is_const_def const_0_sum_mod_2 const_1_sum_mod_2 by auto
-
-lemma id_sum_mod_2:
- assumes "f = id"
- shows "(f 0 + f 1) mod 2 = 1"
- using assms by simp
-
-lemma is_balanced_sum_mod_2:
- assumes "is_balanced"
- shows "(f 0 + f 1) mod 2 = 1"
- using assms is_balanced_def id_sum_mod_2 is_swap_sum_mod_2 by auto
-
-lemma f_ge_0: "\<forall> x. (f x \<ge> 0)" by simp
-
-end (* context deutsch *)
-
-text \<open>The Deutsch's Transform @{text U\<^sub>f}.\<close>
-
-definition (in deutsch) deutsch_transform:: "complex Matrix.mat" ("U\<^sub>f") where
-"U\<^sub>f \<equiv> mat_of_cols_list 4 [[1 - f(0), f(0), 0, 0],
- [f(0), 1 - f(0), 0, 0],
- [0, 0, 1 - f(1), f(1)],
- [0, 0, f(1), 1 - f(1)]]"
-
-lemma (in deutsch) deutsch_transform_dim [simp]:
- shows "dim_row U\<^sub>f = 4" and "dim_col U\<^sub>f = 4"
- by (auto simp add: deutsch_transform_def mat_of_cols_list_def)
-
-lemma (in deutsch) deutsch_transform_coeff_is_zero [simp]:
- shows "U\<^sub>f $$ (0,2) = 0" and "U\<^sub>f $$ (0,3) = 0"
- and "U\<^sub>f $$ (1,2) = 0" and "U\<^sub>f $$(1,3) = 0"
- and "U\<^sub>f $$ (2,0) = 0" and "U\<^sub>f $$(2,1) = 0"
- and "U\<^sub>f $$ (3,0) = 0" and "U\<^sub>f $$ (3,1) = 0"
- using deutsch_transform_def by auto
-
-lemma (in deutsch) deutsch_transform_coeff [simp]:
- shows "U\<^sub>f $$ (0,1) = f(0)" and "U\<^sub>f $$ (1,0) = f(0)"
- and "U\<^sub>f $$(2,3) = f(1)" and "U\<^sub>f $$ (3,2) = f(1)"
- and "U\<^sub>f $$ (0,0) = 1 - f(0)" and "U\<^sub>f $$(1,1) = 1 - f(0)"
- and "U\<^sub>f $$ (2,2) = 1 - f(1)" and "U\<^sub>f $$ (3,3) = 1 - f(1)"
- using deutsch_transform_def by auto
-
-abbreviation (in deutsch) V\<^sub>f:: "complex Matrix.mat" where
-"V\<^sub>f \<equiv> Matrix.mat 4 4 (\<lambda>(i,j).
- if i=0 \<and> j=0 then 1 - f(0) else
- (if i=0 \<and> j=1 then f(0) else
- (if i=1 \<and> j=0 then f(0) else
- (if i=1 \<and> j=1 then 1 - f(0) else
- (if i=2 \<and> j=2 then 1 - f(1) else
- (if i=2 \<and> j=3 then f(1) else
- (if i=3 \<and> j=2 then f(1) else
- (if i=3 \<and> j=3 then 1 - f(1) else 0))))))))"
-
-lemma (in deutsch) deutsch_transform_alt_rep_coeff_is_zero [simp]:
- shows "V\<^sub>f $$ (0,2) = 0" and "V\<^sub>f $$ (0,3) = 0"
- and "V\<^sub>f $$ (1,2) = 0" and "V\<^sub>f $$(1,3) = 0"
- and "V\<^sub>f $$ (2,0) = 0" and "V\<^sub>f $$(2,1) = 0"
- and "V\<^sub>f $$ (3,0) = 0" and "V\<^sub>f $$ (3,1) = 0"
- by auto
-
-lemma (in deutsch) deutsch_transform_alt_rep_coeff [simp]:
- shows "V\<^sub>f $$ (0,1) = f(0)" and "V\<^sub>f $$ (1,0) = f(0)"
- and "V\<^sub>f $$(2,3) = f(1)" and "V\<^sub>f $$ (3,2) = f(1)"
- and "V\<^sub>f $$ (0,0) = 1 - f(0)" and "V\<^sub>f $$(1,1) = 1 - f(0)"
- and "V\<^sub>f $$ (2,2) = 1 - f(1)" and "V\<^sub>f $$ (3,3) = 1 - f(1)"
- by auto
-
-lemma (in deutsch) deutsch_transform_alt_rep:
- shows "U\<^sub>f = V\<^sub>f"
-proof
- show c0:"dim_row U\<^sub>f = dim_row V\<^sub>f" by simp
- show c1:"dim_col U\<^sub>f = dim_col V\<^sub>f" by simp
- fix i j:: nat
- assume "i < dim_row V\<^sub>f" and "j < dim_col V\<^sub>f"
- then have "i < 4" and "j < 4" by auto
- thus "U\<^sub>f $$ (i,j) = V\<^sub>f $$ (i,j)"
- by (smt deutsch_transform_alt_rep_coeff deutsch_transform_alt_rep_coeff_is_zero deutsch_transform_coeff
- deutsch_transform_coeff_is_zero set_4_disj)
-qed
-
-text \<open>@{text U\<^sub>f} is a gate.\<close>
-
-lemma (in deutsch) transpose_of_deutsch_transform:
- shows "(U\<^sub>f)\<^sup>t = U\<^sub>f"
-proof
- show "dim_row (U\<^sub>f\<^sup>t) = dim_row U\<^sub>f" by simp
- show "dim_col (U\<^sub>f\<^sup>t) = dim_col U\<^sub>f" by simp
- fix i j:: nat
- assume "i < dim_row U\<^sub>f" and "j < dim_col U\<^sub>f"
- thus "U\<^sub>f\<^sup>t $$ (i, j) = U\<^sub>f $$ (i, j)"
- by (auto simp add: transpose_mat_def)
- (metis deutsch_transform_coeff(1-4) deutsch_transform_coeff_is_zero set_4_disj)
-qed
-
-lemma (in deutsch) adjoint_of_deutsch_transform:
- shows "(U\<^sub>f)\<^sup>\<dagger> = U\<^sub>f"
-proof
- show "dim_row (U\<^sub>f\<^sup>\<dagger>) = dim_row U\<^sub>f" by simp
- show "dim_col (U\<^sub>f\<^sup>\<dagger>) = dim_col U\<^sub>f" by simp
- fix i j:: nat
- assume "i < dim_row U\<^sub>f" and "j < dim_col U\<^sub>f"
- thus "U\<^sub>f\<^sup>\<dagger> $$ (i, j) = U\<^sub>f $$ (i, j)"
- by (auto simp add: dagger_def)
- (metis complex_cnj_of_nat complex_cnj_zero deutsch_transform_coeff
- deutsch_transform_coeff_is_zero set_4_disj)
-qed
-
-lemma (in deutsch) deutsch_transform_is_gate:
- shows "gate 2 U\<^sub>f"
-proof
- show "dim_row U\<^sub>f = 2\<^sup>2" by simp
- show "square_mat U\<^sub>f" by simp
- show "unitary U\<^sub>f"
- proof-
- have "U\<^sub>f * U\<^sub>f = 1\<^sub>m (dim_col U\<^sub>f)"
- proof
- show "dim_row (U\<^sub>f * U\<^sub>f) = dim_row (1\<^sub>m (dim_col U\<^sub>f))" by simp
- next
- show "dim_col (U\<^sub>f * U\<^sub>f) = dim_col (1\<^sub>m (dim_col U\<^sub>f))" by simp
- next
- fix i j:: nat
- assume "i < dim_row (1\<^sub>m (dim_col U\<^sub>f))" and "j < dim_col (1\<^sub>m (dim_col U\<^sub>f))"
- then show "(U\<^sub>f * U\<^sub>f) $$ (i,j) = 1\<^sub>m (dim_col U\<^sub>f) $$ (i, j)"
- apply (auto simp add: deutsch_transform_alt_rep one_mat_def times_mat_def)
- apply (auto simp: scalar_prod_def)
- using f_values by auto
- qed
- thus ?thesis by (simp add: adjoint_of_deutsch_transform unitary_def)
- qed
-qed
-
-text \<open>
-Two qubits are prepared.
-The first one in the state $|0\rangle$, the second one in the state $|1\rangle$.
-\<close>
-
-abbreviation zero where "zero \<equiv> unit_vec 2 0"
-abbreviation one where "one \<equiv> unit_vec 2 1"
-
-lemma ket_zero_is_state:
- shows "state 1 |zero\<rangle>"
- by (simp add: state_def ket_vec_def cpx_vec_length_def numerals(2))
-
-lemma ket_one_is_state:
- shows "state 1 |one\<rangle>"
- by (simp add: state_def ket_vec_def cpx_vec_length_def numerals(2))
-
-lemma ket_zero_to_mat_of_cols_list [simp]: "|zero\<rangle> = mat_of_cols_list 2 [[1, 0]]"
- by (auto simp add: ket_vec_def mat_of_cols_list_def)
-
-lemma ket_one_to_mat_of_cols_list [simp]: "|one\<rangle> = mat_of_cols_list 2 [[0, 1]]"
- apply (auto simp add: ket_vec_def unit_vec_def mat_of_cols_list_def)
- using less_2_cases by fastforce
-
-text \<open>
-Applying the Hadamard gate to the state $|0\rangle$ results in the new state
-@{term "\<psi>\<^sub>0\<^sub>0"} = $\dfrac {(|0\rangle + |1\rangle)} {\sqrt 2 }$
-\<close>
-
-abbreviation \<psi>\<^sub>0\<^sub>0:: "complex Matrix.mat" where
-"\<psi>\<^sub>0\<^sub>0 \<equiv> mat_of_cols_list 2 [[1/sqrt(2), 1/sqrt(2)]]"
-
-lemma H_on_ket_zero:
- shows "(H * |zero\<rangle>) = \<psi>\<^sub>0\<^sub>0"
-proof
- fix i j:: nat
- assume "i < dim_row \<psi>\<^sub>0\<^sub>0" and "j < dim_col \<psi>\<^sub>0\<^sub>0"
- then have "i \<in> {0,1} \<and> j = 0" by (simp add: mat_of_cols_list_def less_2_cases)
- then show "(H * |zero\<rangle>) $$ (i,j) = \<psi>\<^sub>0\<^sub>0 $$ (i,j)"
- by (auto simp add: mat_of_cols_list_def times_mat_def scalar_prod_def H_def)
-next
- show "dim_row (H * |zero\<rangle>) = dim_row \<psi>\<^sub>0\<^sub>0" by (simp add: H_def mat_of_cols_list_def)
- show "dim_col (H * |zero\<rangle>) = dim_col \<psi>\<^sub>0\<^sub>0" by (simp add: H_def mat_of_cols_list_def)
-qed
-
-lemma H_on_ket_zero_is_state:
- shows "state 1 (H * |zero\<rangle>)"
-proof
- show "gate 1 H"
- using H_is_gate by simp
- show "state 1 |zero\<rangle>"
- using ket_zero_is_state by simp
-qed
-
-text \<open>
-Applying the Hadamard gate to the state $|0\rangle$ results in the new state
-@{text \<psi>\<^sub>0\<^sub>1} = $\dfrac {(|0\rangle - |1\rangle)} {\sqrt 2}$.
-\<close>
-
-abbreviation \<psi>\<^sub>0\<^sub>1:: "complex Matrix.mat" where
-"\<psi>\<^sub>0\<^sub>1 \<equiv> mat_of_cols_list 2 [[1/sqrt(2), -1/sqrt(2)]]"
-
-lemma H_on_ket_one:
- shows "(H * |one\<rangle>) = \<psi>\<^sub>0\<^sub>1"
-proof
- fix i j:: nat
- assume "i < dim_row \<psi>\<^sub>0\<^sub>1" and "j < dim_col \<psi>\<^sub>0\<^sub>1"
- then have "i \<in> {0,1} \<and> j = 0" by (simp add: mat_of_cols_list_def less_2_cases)
- then show "(H * |one\<rangle>) $$ (i,j) = \<psi>\<^sub>0\<^sub>1 $$ (i,j)"
- by (auto simp add: mat_of_cols_list_def times_mat_def scalar_prod_def H_def ket_vec_def)
-next
- show "dim_row (H * |one\<rangle>) = dim_row \<psi>\<^sub>0\<^sub>1" by (simp add: H_def mat_of_cols_list_def)
- show "dim_col (H * |one\<rangle>) = dim_col \<psi>\<^sub>0\<^sub>1" by (simp add: H_def mat_of_cols_list_def ket_vec_def)
-qed
-
-lemma H_on_ket_one_is_state:
- shows "state 1 (H * |one\<rangle>)"
- using H_is_gate ket_one_is_state by simp
-
-text\<open>
-Then, the state @{text \<psi>\<^sub>1} = $\dfrac {(|00\rangle - |01\rangle + |10\rangle - |11\rangle)} {2} $
-is obtained by taking the tensor product of the states
-@{text \<psi>\<^sub>0\<^sub>0} = $\dfrac {(|0\rangle + |1\rangle)} {\sqrt 2} $ and
-@{text \<psi>\<^sub>0\<^sub>1} = $\dfrac {(|0\rangle - |1\rangle)} {\sqrt 2} $.
-\<close>
-
-abbreviation \<psi>\<^sub>1:: "complex Matrix.mat" where
-"\<psi>\<^sub>1 \<equiv> mat_of_cols_list 4 [[1/2, -1/2, 1/2, -1/2]]"
-
-lemma \<psi>\<^sub>0_to_\<psi>\<^sub>1:
- shows "(\<psi>\<^sub>0\<^sub>0 \<Otimes> \<psi>\<^sub>0\<^sub>1) = \<psi>\<^sub>1"
-proof
- fix i j:: nat
- assume "i < dim_row \<psi>\<^sub>1" and "j < dim_col \<psi>\<^sub>1"
- then have "i \<in> {0,1,2,3}" and "j = 0" using mat_of_cols_list_def by auto
- moreover have "complex_of_real (sqrt 2) * complex_of_real (sqrt 2) = 2"
- by (metis mult_2_right numeral_Bit0 of_real_mult of_real_numeral real_sqrt_four real_sqrt_mult)
- ultimately show "(\<psi>\<^sub>0\<^sub>0 \<Otimes> \<psi>\<^sub>0\<^sub>1) $$ (i,j) = \<psi>\<^sub>1 $$ (i,j)" using mat_of_cols_list_def by auto
-next
- show "dim_row (\<psi>\<^sub>0\<^sub>0 \<Otimes> \<psi>\<^sub>0\<^sub>1) = dim_row \<psi>\<^sub>1" by (simp add: mat_of_cols_list_def)
- show "dim_col (\<psi>\<^sub>0\<^sub>0 \<Otimes> \<psi>\<^sub>0\<^sub>1) = dim_col \<psi>\<^sub>1" by (simp add: mat_of_cols_list_def)
-qed
-
-lemma \<psi>\<^sub>1_is_state:
- shows "state 2 \<psi>\<^sub>1"
-proof
- show "dim_col \<psi>\<^sub>1 = 1"
- by (simp add: Tensor.mat_of_cols_list_def)
- show "dim_row \<psi>\<^sub>1 = 2\<^sup>2"
- by (simp add: Tensor.mat_of_cols_list_def)
- show "\<parallel>Matrix.col \<psi>\<^sub>1 0\<parallel> = 1"
- using H_on_ket_one_is_state H_on_ket_zero_is_state state.is_normal tensor_state2 \<psi>\<^sub>0_to_\<psi>\<^sub>1
- H_on_ket_one H_on_ket_zero by force
-qed
-
-text \<open>
-Next, the gate @{text U\<^sub>f} is applied to the state
-@{text \<psi>\<^sub>1} = $\dfrac {(|00\rangle - |01\rangle + |10\rangle - |11\rangle)} {2}$ and
-@{text \<psi>\<^sub>2}= $\dfrac {(|0f(0)\oplus 0\rangle - |0 f(0) \oplus 1\rangle + |1 f(1)\oplus 0\rangle - |1f(1)\oplus 1\rangle)} {2}$
-is obtained. This simplifies to
-@{text \<psi>\<^sub>2}= $\dfrac {(|0f(0)\rangle - |0 \overline{f(0)} \rangle + |1 f(1)\rangle - |1\overline{f(1)}\rangle)} {2}$
-\<close>
-
-abbreviation (in deutsch) \<psi>\<^sub>2:: "complex Matrix.mat" where
-"\<psi>\<^sub>2 \<equiv> mat_of_cols_list 4 [[(1 - f(0))/2 - f(0)/2,
- f(0)/2 - (1 - f(0))/2,
- (1 - f(1))/2 - f(1)/2,
- f(1)/2 - (1- f(1))/2]]"
-
-lemma (in deutsch) \<psi>\<^sub>1_to_\<psi>\<^sub>2:
- shows "U\<^sub>f * \<psi>\<^sub>1 = \<psi>\<^sub>2"
-proof
- fix i j:: nat
- assume "i < dim_row \<psi>\<^sub>2 " and "j < dim_col \<psi>\<^sub>2"
- then have asm:"i \<in> {0,1,2,3} \<and> j = 0 " by (auto simp add: mat_of_cols_list_def)
- then have "i < dim_row U\<^sub>f \<and> j < dim_col \<psi>\<^sub>1"
- using deutsch_transform_def mat_of_cols_list_def by auto
- then have "(U\<^sub>f * \<psi>\<^sub>1) $$ (i, j)
- = (\<Sum> k \<in> {0 ..< dim_vec \<psi>\<^sub>1}. (Matrix.row U\<^sub>f i) $ k * (Matrix.col \<psi>\<^sub>1 j) $ k)"
- apply (auto simp add: times_mat_def scalar_prod_def).
- thus "(U\<^sub>f * \<psi>\<^sub>1) $$ (i, j) = \<psi>\<^sub>2 $$ (i, j)"
- using mat_of_cols_list_def deutsch_transform_def asm by auto
-next
- show "dim_row (U\<^sub>f * \<psi>\<^sub>1) = dim_row \<psi>\<^sub>2" by (simp add: mat_of_cols_list_def)
- show "dim_col (U\<^sub>f * \<psi>\<^sub>1) = dim_col \<psi>\<^sub>2" by (simp add: mat_of_cols_list_def)
-qed
-
-lemma (in deutsch) \<psi>\<^sub>2_is_state:
- shows "state 2 \<psi>\<^sub>2"
-proof
- show "dim_col \<psi>\<^sub>2 = 1"
- by (simp add: Tensor.mat_of_cols_list_def)
- show "dim_row \<psi>\<^sub>2 = 2\<^sup>2"
- by (simp add: Tensor.mat_of_cols_list_def)
- show "\<parallel>Matrix.col \<psi>\<^sub>2 0\<parallel> = 1"
- using gate_on_state_is_state \<psi>\<^sub>1_is_state deutsch_transform_is_gate \<psi>\<^sub>1_to_\<psi>\<^sub>2 state_def
- by (metis (no_types, lifting))
-qed
-
-lemma H_tensor_Id_1:
- defines d:"v \<equiv> mat_of_cols_list 4 [[1/sqrt(2), 0, 1/sqrt(2), 0],
- [0, 1/sqrt(2), 0, 1/sqrt(2)],
- [1/sqrt(2), 0, -1/sqrt(2), 0],
- [0, 1/sqrt(2), 0, -1/sqrt(2)]]"
- shows "(H \<Otimes> Id 1) = v"
-proof
- show "dim_col (H \<Otimes> Id 1) = dim_col v"
- by (simp add: d H_def Id_def mat_of_cols_list_def)
- show "dim_row (H \<Otimes> Id 1) = dim_row v"
- by (simp add: d H_def Id_def mat_of_cols_list_def)
- fix i j:: nat assume "i < dim_row v" and "j < dim_col v"
- then have "i \<in> {0..<4} \<and> j \<in> {0..<4}"
- by (auto simp add: d mat_of_cols_list_def)
- thus "(H \<Otimes> Id 1) $$ (i, j) = v $$ (i, j)"
- by (auto simp add: d Id_def H_def mat_of_cols_list_def)
-qed
-
-lemma H_tensor_Id_1_is_gate:
- shows "gate 2 (H \<Otimes> Id 1)"
-proof
- show "dim_row (H \<Otimes> Quantum.Id 1) = 2\<^sup>2"
- using H_tensor_Id_1 by (simp add: mat_of_cols_list_def)
- show "square_mat (H \<Otimes> Quantum.Id 1)"
- using H_is_gate id_is_gate tensor_gate_sqr_mat by blast
- show "unitary (H \<Otimes> Quantum.Id 1)"
- using H_is_gate gate_def id_is_gate tensor_gate by blast
-qed
-
-text \<open>
-Applying the Hadamard gate to the first qubit of @{text \<psi>\<^sub>2} results in @{text \<psi>\<^sub>3} =
-$\pm |f(0)\oplus f(1)\rangle \left[ \dfrac {(|0\rangle - |1\rangle)} {\sqrt 2}\right]$
-\<close>
-
-abbreviation (in deutsch) \<psi>\<^sub>3:: "complex Matrix.mat" where
-"\<psi>\<^sub>3 \<equiv> mat_of_cols_list 4
-[[(1-f(0))/(2*sqrt(2)) - f(0)/(2*sqrt(2)) + (1-f(1))/(2*sqrt(2)) - f(1)/(2*sqrt(2)),
- f(0)/(2*sqrt(2)) - (1-f(0))/(2*sqrt(2)) + (f(1)/(2*sqrt(2)) - (1-f(1))/(2*sqrt(2))),
- (1-f(0))/(2*sqrt(2)) - f(0)/(2*sqrt(2)) - (1-f(1))/(2*sqrt(2)) + f(1)/(2*sqrt(2)),
- f(0)/(2*sqrt(2)) - (1-f(0))/(2*sqrt(2)) - f(1)/(2*sqrt(2)) + (1-f(1))/(2*sqrt(2))]]"
-
-lemma (in deutsch) \<psi>\<^sub>2_to_\<psi>\<^sub>3:
- shows "(H \<Otimes> Id 1) * \<psi>\<^sub>2 = \<psi>\<^sub>3"
-proof
- fix i j:: nat
- assume "i < dim_row \<psi>\<^sub>3" and "j < dim_col \<psi>\<^sub>3"
- then have a0:"i \<in> {0,1,2,3} \<and> j = 0" by (auto simp add: mat_of_cols_list_def)
- then have "i < dim_row (H \<Otimes> Id 1) \<and> j < dim_col \<psi>\<^sub>2"
- using mat_of_cols_list_def H_tensor_Id_1 by auto
- then have "((H \<Otimes> Id 1)*\<psi>\<^sub>2) $$ (i,j)
- = (\<Sum> k \<in> {0 ..< dim_vec \<psi>\<^sub>2}. (Matrix.row (H \<Otimes> Id 1) i) $ k * (Matrix.col \<psi>\<^sub>2 j) $ k)"
- by (auto simp: times_mat_def scalar_prod_def)
- thus "((H \<Otimes> Id 1) * \<psi>\<^sub>2) $$ (i, j) = \<psi>\<^sub>3 $$ (i, j)"
- using mat_of_cols_list_def H_tensor_Id_1 a0 f_ge_0
- by (auto simp: diff_divide_distrib)
-next
- show "dim_row ((H \<Otimes> Id 1) * \<psi>\<^sub>2) = dim_row \<psi>\<^sub>3"
- using H_tensor_Id_1 mat_of_cols_list_def by simp
- show "dim_col ((H \<Otimes> Id 1) * \<psi>\<^sub>2) = dim_col \<psi>\<^sub>3"
- using H_tensor_Id_1 mat_of_cols_list_def by simp
-qed
-
-lemma (in deutsch) \<psi>\<^sub>3_is_state:
- shows "state 2 \<psi>\<^sub>3"
-proof-
- have "gate 2 (H \<Otimes> Id 1)"
- using H_tensor_Id_1_is_gate by simp
- thus "state 2 \<psi>\<^sub>3"
- using \<psi>\<^sub>2_is_state \<psi>\<^sub>2_to_\<psi>\<^sub>3 by (metis gate_on_state_is_state)
-qed
-
-text \<open>
-Finally, all steps are put together. The result depends on the function f. If f is constant
-the first qubit of $\pm |f(0)\oplus f(1)\rangle \left[ \dfrac {(|0\rangle - |1\rangle)} {\sqrt 2}\right]$
-is 0, if it is is\_balanced it is 1.
-The algorithm only uses one evaluation of f(x) and will always succeed.
-\<close>
-
-definition (in deutsch) deutsch_algo:: "complex Matrix.mat" where
-"deutsch_algo \<equiv> (H \<Otimes> Id 1) * (U\<^sub>f * ((H * |zero\<rangle>) \<Otimes> (H * |one\<rangle>)))"
-
-lemma (in deutsch) deutsch_algo_result [simp]:
- shows "deutsch_algo = \<psi>\<^sub>3"
- using deutsch_algo_def H_on_ket_zero H_on_ket_one \<psi>\<^sub>0_to_\<psi>\<^sub>1 \<psi>\<^sub>1_to_\<psi>\<^sub>2 \<psi>\<^sub>2_to_\<psi>\<^sub>3 by simp
-
-lemma (in deutsch) deutsch_algo_result_is_state:
- shows "state 2 deutsch_algo"
- using \<psi>\<^sub>3_is_state by simp
-
-
-text \<open>
-If the function is constant then the measurement of the first qubit should result in the state
-$|0\rangle$ with probability 1.
-\<close>
-
-lemma (in deutsch) prob0_deutsch_algo_const:
- assumes "is_const"
- shows "prob0 2 deutsch_algo 0 = 1"
-proof -
- have "{k| k::nat. (k < 4) \<and> \<not> select_index 2 0 k} = {0,1}"
- using select_index_def by auto
- then have "prob0 2 deutsch_algo 0 = (\<Sum>j\<in>{0,1}. (cmod(deutsch_algo $$ (j,0)))\<^sup>2)"
- using deutsch_algo_result_is_state prob0_def by simp
- thus "prob0 2 deutsch_algo 0 = 1"
- using assms is_const_def const_def by auto
-qed
-
-lemma (in deutsch) prob1_deutsch_algo_const:
- assumes "is_const"
- shows "prob1 2 deutsch_algo 0 = 0"
- using assms prob0_deutsch_algo_const prob_sum_is_one[of "2" "deutsch_algo" "0"]
-deutsch_algo_result_is_state by simp
-
-text \<open>
-If the function is balanced the measurement of the first qubit should result in the state $|1\rangle$
-with probability 1.
-\<close>
-
-lemma (in deutsch) prob0_deutsch_algo_balanced:
- assumes "is_balanced"
- shows "prob0 2 deutsch_algo 0 = 0"
-proof-
- have "{k| k::nat. (k < 4) \<and> \<not> select_index 2 0 k} = {0,1}"
- using select_index_def by auto
- then have "prob0 2 deutsch_algo 0 = (\<Sum>j \<in> {0,1}. (cmod(deutsch_algo $$ (j,0)))\<^sup>2)"
- using deutsch_algo_result_is_state prob0_def by simp
- thus "prob0 2 deutsch_algo 0 = 0"
- using is_swap_values assms is_balanced_def by auto
-qed
-
-lemma (in deutsch) prob1_deutsch_algo_balanced:
- assumes "is_balanced"
- shows "prob1 2 deutsch_algo 0 = 1"
-using assms prob0_deutsch_algo_balanced prob_sum_is_one[of "2" "deutsch_algo" "0"]
-deutsch_algo_result_is_state by simp
-
-text \<open>Eventually, the measurement of the first qubit results in $f(0)\oplus f(1)$. \<close>
-
-definition (in deutsch) deutsch_algo_eval:: "real" where
-"deutsch_algo_eval \<equiv> prob1 2 deutsch_algo 0"
-
-lemma (in deutsch) sum_mod_2_cases:
- shows "(f 0 + f 1) mod 2 = 0 \<longrightarrow> is_const"
- and "(f 0 + f 1) mod 2 = 1 \<longrightarrow> is_balanced"
- using f_cases is_balanced_sum_mod_2 is_const_sum_mod_2 by auto
-
-lemma (in deutsch) deutsch_algo_eval_is_sum_mod_2:
- shows "deutsch_algo_eval = (f 0 + f 1) mod 2"
- using deutsch_algo_eval_def f_cases is_const_sum_mod_2 is_balanced_sum_mod_2
-prob1_deutsch_algo_const prob1_deutsch_algo_balanced by auto
-
-text \<open>
-If the algorithm returns 0 then one concludes that the input function is constant and
-if it returns 1 then the function is balanced.
-\<close>
-
-theorem (in deutsch) deutsch_algo_is_correct:
- shows "deutsch_algo_eval = 0 \<longrightarrow> is_const" and "deutsch_algo_eval = 1 \<longrightarrow> is_balanced"
- using deutsch_algo_eval_is_sum_mod_2 sum_mod_2_cases by auto
-
+(*
+Authors:
+
+ Hanna Lachnitt, TU Wien, lachnitt@student.tuwien.ac.at
+ Anthony Bordg, University of Cambridge, apdb3@cam.ac.uk
+*)
+
+section \<open>The Deutsch Algorithm\<close>
+
+theory Deutsch
+imports
+ More_Tensor
+ Measurement
+begin
+
+
+text \<open>
+Given a function $f:{0,1}\mapsto {0,1}$, Deutsch's algorithm decides if this function is constant
+or balanced with a single $f(x)$ circuit to evaluate the function for multiple values of $x$
+simultaneously. The algorithm makes use of quantum parallelism and quantum interference.
+\<close>
+
+text \<open>
+A constant function with values in {0,1} returns either always 0 or always 1.
+A balanced function is 0 for half of the inputs and 1 for the other half.
+\<close>
+
+locale deutsch =
+ fixes f:: "nat \<Rightarrow> nat"
+ assumes dom: "f \<in> ({0,1} \<rightarrow>\<^sub>E {0,1})"
+
+context deutsch
+begin
+
+definition is_swap:: bool where
+"is_swap = (\<forall>x \<in> {0,1}. f x = 1 - x)"
+
+lemma is_swap_values:
+ assumes "is_swap"
+ shows "f 0 = 1" and "f 1 = 0"
+ using assms is_swap_def by auto
+
+lemma is_swap_sum_mod_2:
+ assumes "is_swap"
+ shows "(f 0 + f 1) mod 2 = 1"
+ using assms is_swap_def by simp
+
+definition const:: "nat \<Rightarrow> bool" where
+"const n = (\<forall>x \<in> {0,1}.(f x = n))"
+
+definition is_const:: "bool" where
+"is_const \<equiv> const 0 \<or> const 1"
+
+definition is_balanced:: "bool" where
+"is_balanced \<equiv> (\<forall>x \<in> {0,1}.(f x = x)) \<or> is_swap"
+
+lemma f_values: "(f 0 = 0 \<or> f 0 = 1) \<and> (f 1 = 0 \<or> f 1 = 1)"
+ using dom by auto
+
+lemma f_cases:
+ shows "is_const \<or> is_balanced"
+ using dom is_balanced_def const_def is_const_def is_swap_def f_values by auto
+
+lemma const_0_sum_mod_2:
+ assumes "const 0"
+ shows "(f 0 + f 1) mod 2 = 0"
+ using assms const_def by simp
+
+lemma const_1_sum_mod_2:
+ assumes "const 1"
+ shows "(f 0 + f 1) mod 2 = 0"
+ using assms const_def by simp
+
+lemma is_const_sum_mod_2:
+ assumes "is_const"
+ shows "(f 0 + f 1) mod 2 = 0"
+ using assms is_const_def const_0_sum_mod_2 const_1_sum_mod_2 by auto
+
+lemma id_sum_mod_2:
+ assumes "f = id"
+ shows "(f 0 + f 1) mod 2 = 1"
+ using assms by simp
+
+lemma is_balanced_sum_mod_2:
+ assumes "is_balanced"
+ shows "(f 0 + f 1) mod 2 = 1"
+ using assms is_balanced_def id_sum_mod_2 is_swap_sum_mod_2 by auto
+
+lemma f_ge_0: "\<forall> x. (f x \<ge> 0)" by simp
+
+end (* context deutsch *)
+
+text \<open>The Deutsch's Transform @{text U\<^sub>f}.\<close>
+
+definition (in deutsch) deutsch_transform:: "complex Matrix.mat" ("U\<^sub>f") where
+"U\<^sub>f \<equiv> mat_of_cols_list 4 [[1 - f(0), f(0), 0, 0],
+ [f(0), 1 - f(0), 0, 0],
+ [0, 0, 1 - f(1), f(1)],
+ [0, 0, f(1), 1 - f(1)]]"
+
+lemma (in deutsch) deutsch_transform_dim [simp]:
+ shows "dim_row U\<^sub>f = 4" and "dim_col U\<^sub>f = 4"
+ by (auto simp add: deutsch_transform_def mat_of_cols_list_def)
+
+lemma (in deutsch) deutsch_transform_coeff_is_zero [simp]:
+ shows "U\<^sub>f $$ (0,2) = 0" and "U\<^sub>f $$ (0,3) = 0"
+ and "U\<^sub>f $$ (1,2) = 0" and "U\<^sub>f $$(1,3) = 0"
+ and "U\<^sub>f $$ (2,0) = 0" and "U\<^sub>f $$(2,1) = 0"
+ and "U\<^sub>f $$ (3,0) = 0" and "U\<^sub>f $$ (3,1) = 0"
+ using deutsch_transform_def by auto
+
+lemma (in deutsch) deutsch_transform_coeff [simp]:
+ shows "U\<^sub>f $$ (0,1) = f(0)" and "U\<^sub>f $$ (1,0) = f(0)"
+ and "U\<^sub>f $$(2,3) = f(1)" and "U\<^sub>f $$ (3,2) = f(1)"
+ and "U\<^sub>f $$ (0,0) = 1 - f(0)" and "U\<^sub>f $$(1,1) = 1 - f(0)"
+ and "U\<^sub>f $$ (2,2) = 1 - f(1)" and "U\<^sub>f $$ (3,3) = 1 - f(1)"
+ using deutsch_transform_def by auto
+
+abbreviation (in deutsch) V\<^sub>f:: "complex Matrix.mat" where
+"V\<^sub>f \<equiv> Matrix.mat 4 4 (\<lambda>(i,j).
+ if i=0 \<and> j=0 then 1 - f(0) else
+ (if i=0 \<and> j=1 then f(0) else
+ (if i=1 \<and> j=0 then f(0) else
+ (if i=1 \<and> j=1 then 1 - f(0) else
+ (if i=2 \<and> j=2 then 1 - f(1) else
+ (if i=2 \<and> j=3 then f(1) else
+ (if i=3 \<and> j=2 then f(1) else
+ (if i=3 \<and> j=3 then 1 - f(1) else 0))))))))"
+
+lemma (in deutsch) deutsch_transform_alt_rep_coeff_is_zero [simp]:
+ shows "V\<^sub>f $$ (0,2) = 0" and "V\<^sub>f $$ (0,3) = 0"
+ and "V\<^sub>f $$ (1,2) = 0" and "V\<^sub>f $$(1,3) = 0"
+ and "V\<^sub>f $$ (2,0) = 0" and "V\<^sub>f $$(2,1) = 0"
+ and "V\<^sub>f $$ (3,0) = 0" and "V\<^sub>f $$ (3,1) = 0"
+ by auto
+
+lemma (in deutsch) deutsch_transform_alt_rep_coeff [simp]:
+ shows "V\<^sub>f $$ (0,1) = f(0)" and "V\<^sub>f $$ (1,0) = f(0)"
+ and "V\<^sub>f $$(2,3) = f(1)" and "V\<^sub>f $$ (3,2) = f(1)"
+ and "V\<^sub>f $$ (0,0) = 1 - f(0)" and "V\<^sub>f $$(1,1) = 1 - f(0)"
+ and "V\<^sub>f $$ (2,2) = 1 - f(1)" and "V\<^sub>f $$ (3,3) = 1 - f(1)"
+ by auto
+
+lemma (in deutsch) deutsch_transform_alt_rep:
+ shows "U\<^sub>f = V\<^sub>f"
+proof
+ show c0:"dim_row U\<^sub>f = dim_row V\<^sub>f" by simp
+ show c1:"dim_col U\<^sub>f = dim_col V\<^sub>f" by simp
+ fix i j:: nat
+ assume "i < dim_row V\<^sub>f" and "j < dim_col V\<^sub>f"
+ then have "i < 4" and "j < 4" by auto
+ thus "U\<^sub>f $$ (i,j) = V\<^sub>f $$ (i,j)"
+ by (smt deutsch_transform_alt_rep_coeff deutsch_transform_alt_rep_coeff_is_zero deutsch_transform_coeff
+ deutsch_transform_coeff_is_zero set_4_disj)
+qed
+
+text \<open>@{text U\<^sub>f} is a gate.\<close>
+
+lemma (in deutsch) transpose_of_deutsch_transform:
+ shows "(U\<^sub>f)\<^sup>t = U\<^sub>f"
+proof
+ show "dim_row (U\<^sub>f\<^sup>t) = dim_row U\<^sub>f" by simp
+ show "dim_col (U\<^sub>f\<^sup>t) = dim_col U\<^sub>f" by simp
+ fix i j:: nat
+ assume "i < dim_row U\<^sub>f" and "j < dim_col U\<^sub>f"
+ thus "U\<^sub>f\<^sup>t $$ (i, j) = U\<^sub>f $$ (i, j)"
+ by (auto simp add: transpose_mat_def)
+ (metis deutsch_transform_coeff(1-4) deutsch_transform_coeff_is_zero set_4_disj)
+qed
+
+lemma (in deutsch) adjoint_of_deutsch_transform:
+ shows "(U\<^sub>f)\<^sup>\<dagger> = U\<^sub>f"
+proof
+ show "dim_row (U\<^sub>f\<^sup>\<dagger>) = dim_row U\<^sub>f" by simp
+ show "dim_col (U\<^sub>f\<^sup>\<dagger>) = dim_col U\<^sub>f" by simp
+ fix i j:: nat
+ assume "i < dim_row U\<^sub>f" and "j < dim_col U\<^sub>f"
+ thus "U\<^sub>f\<^sup>\<dagger> $$ (i, j) = U\<^sub>f $$ (i, j)"
+ by (auto simp add: dagger_def)
+ (metis complex_cnj_of_nat complex_cnj_zero deutsch_transform_coeff
+ deutsch_transform_coeff_is_zero set_4_disj)
+qed
+
+lemma (in deutsch) deutsch_transform_is_gate:
+ shows "gate 2 U\<^sub>f"
+proof
+ show "dim_row U\<^sub>f = 2\<^sup>2" by simp
+ show "square_mat U\<^sub>f" by simp
+ show "unitary U\<^sub>f"
+ proof-
+ have "U\<^sub>f * U\<^sub>f = 1\<^sub>m (dim_col U\<^sub>f)"
+ proof
+ show "dim_row (U\<^sub>f * U\<^sub>f) = dim_row (1\<^sub>m (dim_col U\<^sub>f))" by simp
+ next
+ show "dim_col (U\<^sub>f * U\<^sub>f) = dim_col (1\<^sub>m (dim_col U\<^sub>f))" by simp
+ next
+ fix i j:: nat
+ assume "i < dim_row (1\<^sub>m (dim_col U\<^sub>f))" and "j < dim_col (1\<^sub>m (dim_col U\<^sub>f))"
+ then show "(U\<^sub>f * U\<^sub>f) $$ (i,j) = 1\<^sub>m (dim_col U\<^sub>f) $$ (i, j)"
+ apply (auto simp add: deutsch_transform_alt_rep one_mat_def times_mat_def)
+ apply (auto simp: scalar_prod_def)
+ using f_values by auto
+ qed
+ thus ?thesis by (simp add: adjoint_of_deutsch_transform unitary_def)
+ qed
+qed
+
+text \<open>
+Two qubits are prepared.
+The first one in the state $|0\rangle$, the second one in the state $|1\rangle$.
+\<close>
+
+abbreviation zero where "zero \<equiv> unit_vec 2 0"
+abbreviation one where "one \<equiv> unit_vec 2 1"
+
+lemma ket_zero_is_state:
+ shows "state 1 |zero\<rangle>"
+ by (simp add: state_def ket_vec_def cpx_vec_length_def numerals(2))
+
+lemma ket_one_is_state:
+ shows "state 1 |one\<rangle>"
+ by (simp add: state_def ket_vec_def cpx_vec_length_def numerals(2))
+
+lemma ket_zero_to_mat_of_cols_list [simp]: "|zero\<rangle> = mat_of_cols_list 2 [[1, 0]]"
+ by (auto simp add: ket_vec_def mat_of_cols_list_def)
+
+lemma ket_one_to_mat_of_cols_list [simp]: "|one\<rangle> = mat_of_cols_list 2 [[0, 1]]"
+ apply (auto simp add: ket_vec_def unit_vec_def mat_of_cols_list_def)
+ using less_2_cases by fastforce
+
+text \<open>
+Applying the Hadamard gate to the state $|0\rangle$ results in the new state
+@{term "\<psi>\<^sub>0\<^sub>0"} = $\dfrac {(|0\rangle + |1\rangle)} {\sqrt 2 }$
+\<close>
+
+abbreviation \<psi>\<^sub>0\<^sub>0:: "complex Matrix.mat" where
+"\<psi>\<^sub>0\<^sub>0 \<equiv> mat_of_cols_list 2 [[1/sqrt(2), 1/sqrt(2)]]"
+
+lemma H_on_ket_zero:
+ shows "(H * |zero\<rangle>) = \<psi>\<^sub>0\<^sub>0"
+proof
+ fix i j:: nat
+ assume "i < dim_row \<psi>\<^sub>0\<^sub>0" and "j < dim_col \<psi>\<^sub>0\<^sub>0"
+ then have "i \<in> {0,1} \<and> j = 0" by (simp add: mat_of_cols_list_def less_2_cases)
+ then show "(H * |zero\<rangle>) $$ (i,j) = \<psi>\<^sub>0\<^sub>0 $$ (i,j)"
+ by (auto simp add: mat_of_cols_list_def times_mat_def scalar_prod_def H_def)
+next
+ show "dim_row (H * |zero\<rangle>) = dim_row \<psi>\<^sub>0\<^sub>0" by (simp add: H_def mat_of_cols_list_def)
+ show "dim_col (H * |zero\<rangle>) = dim_col \<psi>\<^sub>0\<^sub>0" by (simp add: H_def mat_of_cols_list_def)
+qed
+
+lemma H_on_ket_zero_is_state:
+ shows "state 1 (H * |zero\<rangle>)"
+proof
+ show "gate 1 H"
+ using H_is_gate by simp
+ show "state 1 |zero\<rangle>"
+ using ket_zero_is_state by simp
+qed
+
+text \<open>
+Applying the Hadamard gate to the state $|0\rangle$ results in the new state
+@{text \<psi>\<^sub>0\<^sub>1} = $\dfrac {(|0\rangle - |1\rangle)} {\sqrt 2}$.
+\<close>
+
+abbreviation \<psi>\<^sub>0\<^sub>1:: "complex Matrix.mat" where
+"\<psi>\<^sub>0\<^sub>1 \<equiv> mat_of_cols_list 2 [[1/sqrt(2), -1/sqrt(2)]]"
+
+lemma H_on_ket_one:
+ shows "(H * |one\<rangle>) = \<psi>\<^sub>0\<^sub>1"
+proof
+ fix i j:: nat
+ assume "i < dim_row \<psi>\<^sub>0\<^sub>1" and "j < dim_col \<psi>\<^sub>0\<^sub>1"
+ then have "i \<in> {0,1} \<and> j = 0" by (simp add: mat_of_cols_list_def less_2_cases)
+ then show "(H * |one\<rangle>) $$ (i,j) = \<psi>\<^sub>0\<^sub>1 $$ (i,j)"
+ by (auto simp add: mat_of_cols_list_def times_mat_def scalar_prod_def H_def ket_vec_def)
+next
+ show "dim_row (H * |one\<rangle>) = dim_row \<psi>\<^sub>0\<^sub>1" by (simp add: H_def mat_of_cols_list_def)
+ show "dim_col (H * |one\<rangle>) = dim_col \<psi>\<^sub>0\<^sub>1" by (simp add: H_def mat_of_cols_list_def ket_vec_def)
+qed
+
+lemma H_on_ket_one_is_state:
+ shows "state 1 (H * |one\<rangle>)"
+ using H_is_gate ket_one_is_state by simp
+
+text\<open>
+Then, the state @{text \<psi>\<^sub>1} = $\dfrac {(|00\rangle - |01\rangle + |10\rangle - |11\rangle)} {2} $
+is obtained by taking the tensor product of the states
+@{text \<psi>\<^sub>0\<^sub>0} = $\dfrac {(|0\rangle + |1\rangle)} {\sqrt 2} $ and
+@{text \<psi>\<^sub>0\<^sub>1} = $\dfrac {(|0\rangle - |1\rangle)} {\sqrt 2} $.
+\<close>
+
+abbreviation \<psi>\<^sub>1:: "complex Matrix.mat" where
+"\<psi>\<^sub>1 \<equiv> mat_of_cols_list 4 [[1/2, -1/2, 1/2, -1/2]]"
+
+lemma \<psi>\<^sub>0_to_\<psi>\<^sub>1:
+ shows "(\<psi>\<^sub>0\<^sub>0 \<Otimes> \<psi>\<^sub>0\<^sub>1) = \<psi>\<^sub>1"
+proof
+ fix i j:: nat
+ assume "i < dim_row \<psi>\<^sub>1" and "j < dim_col \<psi>\<^sub>1"
+ then have "i \<in> {0,1,2,3}" and "j = 0" using mat_of_cols_list_def by auto
+ moreover have "complex_of_real (sqrt 2) * complex_of_real (sqrt 2) = 2"
+ by (metis mult_2_right numeral_Bit0 of_real_mult of_real_numeral real_sqrt_four real_sqrt_mult)
+ ultimately show "(\<psi>\<^sub>0\<^sub>0 \<Otimes> \<psi>\<^sub>0\<^sub>1) $$ (i,j) = \<psi>\<^sub>1 $$ (i,j)" using mat_of_cols_list_def by auto
+next
+ show "dim_row (\<psi>\<^sub>0\<^sub>0 \<Otimes> \<psi>\<^sub>0\<^sub>1) = dim_row \<psi>\<^sub>1" by (simp add: mat_of_cols_list_def)
+ show "dim_col (\<psi>\<^sub>0\<^sub>0 \<Otimes> \<psi>\<^sub>0\<^sub>1) = dim_col \<psi>\<^sub>1" by (simp add: mat_of_cols_list_def)
+qed
+
+lemma \<psi>\<^sub>1_is_state:
+ shows "state 2 \<psi>\<^sub>1"
+proof
+ show "dim_col \<psi>\<^sub>1 = 1"
+ by (simp add: Tensor.mat_of_cols_list_def)
+ show "dim_row \<psi>\<^sub>1 = 2\<^sup>2"
+ by (simp add: Tensor.mat_of_cols_list_def)
+ show "\<parallel>Matrix.col \<psi>\<^sub>1 0\<parallel> = 1"
+ using H_on_ket_one_is_state H_on_ket_zero_is_state state.is_normal tensor_state2 \<psi>\<^sub>0_to_\<psi>\<^sub>1
+ H_on_ket_one H_on_ket_zero by force
+qed
+
+text \<open>
+Next, the gate @{text U\<^sub>f} is applied to the state
+@{text \<psi>\<^sub>1} = $\dfrac {(|00\rangle - |01\rangle + |10\rangle - |11\rangle)} {2}$ and
+@{text \<psi>\<^sub>2}= $\dfrac {(|0f(0)\oplus 0\rangle - |0 f(0) \oplus 1\rangle + |1 f(1)\oplus 0\rangle - |1f(1)\oplus 1\rangle)} {2}$
+is obtained. This simplifies to
+@{text \<psi>\<^sub>2}= $\dfrac {(|0f(0)\rangle - |0 \overline{f(0)} \rangle + |1 f(1)\rangle - |1\overline{f(1)}\rangle)} {2}$
+\<close>
+
+abbreviation (in deutsch) \<psi>\<^sub>2:: "complex Matrix.mat" where
+"\<psi>\<^sub>2 \<equiv> mat_of_cols_list 4 [[(1 - f(0))/2 - f(0)/2,
+ f(0)/2 - (1 - f(0))/2,
+ (1 - f(1))/2 - f(1)/2,
+ f(1)/2 - (1- f(1))/2]]"
+
+lemma (in deutsch) \<psi>\<^sub>1_to_\<psi>\<^sub>2:
+ shows "U\<^sub>f * \<psi>\<^sub>1 = \<psi>\<^sub>2"
+proof
+ fix i j:: nat
+ assume "i < dim_row \<psi>\<^sub>2 " and "j < dim_col \<psi>\<^sub>2"
+ then have asm:"i \<in> {0,1,2,3} \<and> j = 0 " by (auto simp add: mat_of_cols_list_def)
+ then have "i < dim_row U\<^sub>f \<and> j < dim_col \<psi>\<^sub>1"
+ using deutsch_transform_def mat_of_cols_list_def by auto
+ then have "(U\<^sub>f * \<psi>\<^sub>1) $$ (i, j)
+ = (\<Sum> k \<in> {0 ..< dim_vec \<psi>\<^sub>1}. (Matrix.row U\<^sub>f i) $ k * (Matrix.col \<psi>\<^sub>1 j) $ k)"
+ apply (auto simp add: times_mat_def scalar_prod_def).
+ thus "(U\<^sub>f * \<psi>\<^sub>1) $$ (i, j) = \<psi>\<^sub>2 $$ (i, j)"
+ using mat_of_cols_list_def deutsch_transform_def asm by auto
+next
+ show "dim_row (U\<^sub>f * \<psi>\<^sub>1) = dim_row \<psi>\<^sub>2" by (simp add: mat_of_cols_list_def)
+ show "dim_col (U\<^sub>f * \<psi>\<^sub>1) = dim_col \<psi>\<^sub>2" by (simp add: mat_of_cols_list_def)
+qed
+
+lemma (in deutsch) \<psi>\<^sub>2_is_state:
+ shows "state 2 \<psi>\<^sub>2"
+proof
+ show "dim_col \<psi>\<^sub>2 = 1"
+ by (simp add: Tensor.mat_of_cols_list_def)
+ show "dim_row \<psi>\<^sub>2 = 2\<^sup>2"
+ by (simp add: Tensor.mat_of_cols_list_def)
+ show "\<parallel>Matrix.col \<psi>\<^sub>2 0\<parallel> = 1"
+ using gate_on_state_is_state \<psi>\<^sub>1_is_state deutsch_transform_is_gate \<psi>\<^sub>1_to_\<psi>\<^sub>2 state_def
+ by (metis (no_types, lifting))
+qed
+
+lemma H_tensor_Id_1:
+ defines d:"v \<equiv> mat_of_cols_list 4 [[1/sqrt(2), 0, 1/sqrt(2), 0],
+ [0, 1/sqrt(2), 0, 1/sqrt(2)],
+ [1/sqrt(2), 0, -1/sqrt(2), 0],
+ [0, 1/sqrt(2), 0, -1/sqrt(2)]]"
+ shows "(H \<Otimes> Id 1) = v"
+proof
+ show "dim_col (H \<Otimes> Id 1) = dim_col v"
+ by (simp add: d H_def Id_def mat_of_cols_list_def)
+ show "dim_row (H \<Otimes> Id 1) = dim_row v"
+ by (simp add: d H_def Id_def mat_of_cols_list_def)
+ fix i j:: nat assume "i < dim_row v" and "j < dim_col v"
+ then have "i \<in> {0..<4} \<and> j \<in> {0..<4}"
+ by (auto simp add: d mat_of_cols_list_def)
+ thus "(H \<Otimes> Id 1) $$ (i, j) = v $$ (i, j)"
+ by (auto simp add: d Id_def H_def mat_of_cols_list_def)
+qed
+
+lemma H_tensor_Id_1_is_gate:
+ shows "gate 2 (H \<Otimes> Id 1)"
+proof
+ show "dim_row (H \<Otimes> Quantum.Id 1) = 2\<^sup>2"
+ using H_tensor_Id_1 by (simp add: mat_of_cols_list_def)
+ show "square_mat (H \<Otimes> Quantum.Id 1)"
+ using H_is_gate id_is_gate tensor_gate_sqr_mat by blast
+ show "unitary (H \<Otimes> Quantum.Id 1)"
+ using H_is_gate gate_def id_is_gate tensor_gate by blast
+qed
+
+text \<open>
+Applying the Hadamard gate to the first qubit of @{text \<psi>\<^sub>2} results in @{text \<psi>\<^sub>3} =
+$\pm |f(0)\oplus f(1)\rangle \left[ \dfrac {(|0\rangle - |1\rangle)} {\sqrt 2}\right]$
+\<close>
+
+abbreviation (in deutsch) \<psi>\<^sub>3:: "complex Matrix.mat" where
+"\<psi>\<^sub>3 \<equiv> mat_of_cols_list 4
+[[(1-f(0))/(2*sqrt(2)) - f(0)/(2*sqrt(2)) + (1-f(1))/(2*sqrt(2)) - f(1)/(2*sqrt(2)),
+ f(0)/(2*sqrt(2)) - (1-f(0))/(2*sqrt(2)) + (f(1)/(2*sqrt(2)) - (1-f(1))/(2*sqrt(2))),
+ (1-f(0))/(2*sqrt(2)) - f(0)/(2*sqrt(2)) - (1-f(1))/(2*sqrt(2)) + f(1)/(2*sqrt(2)),
+ f(0)/(2*sqrt(2)) - (1-f(0))/(2*sqrt(2)) - f(1)/(2*sqrt(2)) + (1-f(1))/(2*sqrt(2))]]"
+
+lemma (in deutsch) \<psi>\<^sub>2_to_\<psi>\<^sub>3:
+ shows "(H \<Otimes> Id 1) * \<psi>\<^sub>2 = \<psi>\<^sub>3"
+proof
+ fix i j:: nat
+ assume "i < dim_row \<psi>\<^sub>3" and "j < dim_col \<psi>\<^sub>3"
+ then have a0:"i \<in> {0,1,2,3} \<and> j = 0" by (auto simp add: mat_of_cols_list_def)
+ then have "i < dim_row (H \<Otimes> Id 1) \<and> j < dim_col \<psi>\<^sub>2"
+ using mat_of_cols_list_def H_tensor_Id_1 by auto
+ then have "((H \<Otimes> Id 1)*\<psi>\<^sub>2) $$ (i,j)
+ = (\<Sum> k \<in> {0 ..< dim_vec \<psi>\<^sub>2}. (Matrix.row (H \<Otimes> Id 1) i) $ k * (Matrix.col \<psi>\<^sub>2 j) $ k)"
+ by (auto simp: times_mat_def scalar_prod_def)
+ thus "((H \<Otimes> Id 1) * \<psi>\<^sub>2) $$ (i, j) = \<psi>\<^sub>3 $$ (i, j)"
+ using mat_of_cols_list_def H_tensor_Id_1 a0 f_ge_0
+ by (auto simp: diff_divide_distrib)
+next
+ show "dim_row ((H \<Otimes> Id 1) * \<psi>\<^sub>2) = dim_row \<psi>\<^sub>3"
+ using H_tensor_Id_1 mat_of_cols_list_def by simp
+ show "dim_col ((H \<Otimes> Id 1) * \<psi>\<^sub>2) = dim_col \<psi>\<^sub>3"
+ using H_tensor_Id_1 mat_of_cols_list_def by simp
+qed
+
+lemma (in deutsch) \<psi>\<^sub>3_is_state:
+ shows "state 2 \<psi>\<^sub>3"
+proof-
+ have "gate 2 (H \<Otimes> Id 1)"
+ using H_tensor_Id_1_is_gate by simp
+ thus "state 2 \<psi>\<^sub>3"
+ using \<psi>\<^sub>2_is_state \<psi>\<^sub>2_to_\<psi>\<^sub>3 by (metis gate_on_state_is_state)
+qed
+
+text \<open>
+Finally, all steps are put together. The result depends on the function f. If f is constant
+the first qubit of $\pm |f(0)\oplus f(1)\rangle \left[ \dfrac {(|0\rangle - |1\rangle)} {\sqrt 2}\right]$
+is 0, if it is is\_balanced it is 1.
+The algorithm only uses one evaluation of f(x) and will always succeed.
+\<close>
+
+definition (in deutsch) deutsch_algo:: "complex Matrix.mat" where
+"deutsch_algo \<equiv> (H \<Otimes> Id 1) * (U\<^sub>f * ((H * |zero\<rangle>) \<Otimes> (H * |one\<rangle>)))"
+
+lemma (in deutsch) deutsch_algo_result [simp]:
+ shows "deutsch_algo = \<psi>\<^sub>3"
+ using deutsch_algo_def H_on_ket_zero H_on_ket_one \<psi>\<^sub>0_to_\<psi>\<^sub>1 \<psi>\<^sub>1_to_\<psi>\<^sub>2 \<psi>\<^sub>2_to_\<psi>\<^sub>3 by simp
+
+lemma (in deutsch) deutsch_algo_result_is_state:
+ shows "state 2 deutsch_algo"
+ using \<psi>\<^sub>3_is_state by simp
+
+
+text \<open>
+If the function is constant then the measurement of the first qubit should result in the state
+$|0\rangle$ with probability 1.
+\<close>
+
+lemma (in deutsch) prob0_deutsch_algo_const:
+ assumes "is_const"
+ shows "prob0 2 deutsch_algo 0 = 1"
+proof -
+ have "{k| k::nat. (k < 4) \<and> \<not> select_index 2 0 k} = {0,1}"
+ using select_index_def by auto
+ then have "prob0 2 deutsch_algo 0 = (\<Sum>j\<in>{0,1}. (cmod(deutsch_algo $$ (j,0)))\<^sup>2)"
+ using deutsch_algo_result_is_state prob0_def by simp
+ thus "prob0 2 deutsch_algo 0 = 1"
+ using assms is_const_def const_def by auto
+qed
+
+lemma (in deutsch) prob1_deutsch_algo_const:
+ assumes "is_const"
+ shows "prob1 2 deutsch_algo 0 = 0"
+ using assms prob0_deutsch_algo_const prob_sum_is_one[of "2" "deutsch_algo" "0"]
+deutsch_algo_result_is_state by simp
+
+text \<open>
+If the function is balanced the measurement of the first qubit should result in the state $|1\rangle$
+with probability 1.
+\<close>
+
+lemma (in deutsch) prob0_deutsch_algo_balanced:
+ assumes "is_balanced"
+ shows "prob0 2 deutsch_algo 0 = 0"
+proof-
+ have "{k| k::nat. (k < 4) \<and> \<not> select_index 2 0 k} = {0,1}"
+ using select_index_def by auto
+ then have "prob0 2 deutsch_algo 0 = (\<Sum>j \<in> {0,1}. (cmod(deutsch_algo $$ (j,0)))\<^sup>2)"
+ using deutsch_algo_result_is_state prob0_def by simp
+ thus "prob0 2 deutsch_algo 0 = 0"
+ using is_swap_values assms is_balanced_def by auto
+qed
+
+lemma (in deutsch) prob1_deutsch_algo_balanced:
+ assumes "is_balanced"
+ shows "prob1 2 deutsch_algo 0 = 1"
+using assms prob0_deutsch_algo_balanced prob_sum_is_one[of "2" "deutsch_algo" "0"]
+deutsch_algo_result_is_state by simp
+
+text \<open>Eventually, the measurement of the first qubit results in $f(0)\oplus f(1)$. \<close>
+
+definition (in deutsch) deutsch_algo_eval:: "real" where
+"deutsch_algo_eval \<equiv> prob1 2 deutsch_algo 0"
+
+lemma (in deutsch) sum_mod_2_cases:
+ shows "(f 0 + f 1) mod 2 = 0 \<longrightarrow> is_const"
+ and "(f 0 + f 1) mod 2 = 1 \<longrightarrow> is_balanced"
+ using f_cases is_balanced_sum_mod_2 is_const_sum_mod_2 by auto
+
+lemma (in deutsch) deutsch_algo_eval_is_sum_mod_2:
+ shows "deutsch_algo_eval = (f 0 + f 1) mod 2"
+ using deutsch_algo_eval_def f_cases is_const_sum_mod_2 is_balanced_sum_mod_2
+prob1_deutsch_algo_const prob1_deutsch_algo_balanced by auto
+
+text \<open>
+If the algorithm returns 0 then one concludes that the input function is constant and
+if it returns 1 then the function is balanced.
+\<close>
+
+theorem (in deutsch) deutsch_algo_is_correct:
+ shows "deutsch_algo_eval = 0 \<longrightarrow> is_const" and "deutsch_algo_eval = 1 \<longrightarrow> is_balanced"
+ using deutsch_algo_eval_is_sum_mod_2 sum_mod_2_cases by auto
+
end
\ No newline at end of file
diff --git a/thys/JinjaDCI/BV/BVConform.thy b/thys/JinjaDCI/BV/BVConform.thy
--- a/thys/JinjaDCI/BV/BVConform.thy
+++ b/thys/JinjaDCI/BV/BVConform.thy
@@ -1,433 +1,433 @@
-(* Title: JinjaDCI/BV/BVConform.thy
-
- Author: Cornelia Pusch, Gerwin Klein, Susannah Mansky
- Copyright 1999 Technische Universitaet Muenchen, 2019-20 UIUC
-
- Based on the Jinja theory BV/BVConform.thy by Cornelia Pusch and Gerwin Klein
-
-The invariant for the type safety proof.
-*)
-
-section \<open> BV Type Safety Invariant \<close>
-
-theory BVConform
-imports BVSpec "../JVM/JVMExec" "../Common/Conform"
-begin
-
-subsection \<open> @{text "correct_state"} definitions \<close>
-
-definition confT :: "'c prog \<Rightarrow> heap \<Rightarrow> val \<Rightarrow> ty err \<Rightarrow> bool"
- ("_,_ \<turnstile> _ :\<le>\<^sub>\<top> _" [51,51,51,51] 50)
-where
- "P,h \<turnstile> v :\<le>\<^sub>\<top> E \<equiv> case E of Err \<Rightarrow> True | OK T \<Rightarrow> P,h \<turnstile> v :\<le> T"
-
-notation (ASCII)
- confT ("_,_ |- _ :<=T _" [51,51,51,51] 50)
-
-abbreviation
- confTs :: "'c prog \<Rightarrow> heap \<Rightarrow> val list \<Rightarrow> ty\<^sub>l \<Rightarrow> bool"
- ("_,_ \<turnstile> _ [:\<le>\<^sub>\<top>] _" [51,51,51,51] 50) where
- "P,h \<turnstile> vs [:\<le>\<^sub>\<top>] Ts \<equiv> list_all2 (confT P h) vs Ts"
-
-notation (ASCII)
- confTs ("_,_ |- _ [:<=T] _" [51,51,51,51] 50)
-
-fun Called_context :: "jvm_prog \<Rightarrow> cname \<Rightarrow> instr \<Rightarrow> bool" where
-"Called_context P C\<^sub>0 (New C') = (C\<^sub>0=C')" |
-"Called_context P C\<^sub>0 (Getstatic C F D) = ((C\<^sub>0=D) \<and> (\<exists>t. P \<turnstile> C has F,Static:t in D))" |
-"Called_context P C\<^sub>0 (Putstatic C F D) = ((C\<^sub>0=D) \<and> (\<exists>t. P \<turnstile> C has F,Static:t in D))" |
-"Called_context P C\<^sub>0 (Invokestatic C M n)
- = (\<exists>Ts T m D. (C\<^sub>0=D) \<and> P \<turnstile> C sees M,Static:Ts \<rightarrow> T = m in D)" |
-"Called_context P _ _ = False"
-
-abbreviation Called_set :: "instr set" where
-"Called_set \<equiv> {i. \<exists>C. i = New C} \<union> {i. \<exists>C M n. i = Invokestatic C M n}
- \<union> {i. \<exists>C F D. i = Getstatic C F D} \<union> {i. \<exists>C F D. i = Putstatic C F D}"
-
-lemma Called_context_Called_set:
- "Called_context P D i \<Longrightarrow> i \<in> Called_set" by(cases i, auto)
-
-fun valid_ics :: "jvm_prog \<Rightarrow> heap \<Rightarrow> sheap \<Rightarrow> cname \<times> mname \<times> pc \<times> init_call_status \<Rightarrow> bool"
- ("_,_,_ \<turnstile>\<^sub>i _" [51,51,51,51] 50) where
-"valid_ics P h sh (C,M,pc,Calling C' Cs)
- = (let ins = instrs_of P C M in Called_context P (last (C'#Cs)) (ins!pc)
- \<and> is_class P C')" |
-"valid_ics P h sh (C,M,pc,Throwing Cs a)
- =(let ins = instrs_of P C M in \<exists>C1. Called_context P C1 (ins!pc)
- \<and> (\<exists>obj. h a = Some obj))" |
-"valid_ics P h sh (C,M,pc,Called Cs)
- = (let ins = instrs_of P C M
- in \<exists>C1 sobj. Called_context P C1 (ins!pc) \<and> sh C1 = Some sobj)" |
-"valid_ics P _ _ _ = True"
-
-definition conf_f :: "jvm_prog \<Rightarrow> heap \<Rightarrow> sheap \<Rightarrow> ty\<^sub>i \<Rightarrow> bytecode \<Rightarrow> frame \<Rightarrow> bool"
-where
- "conf_f P h sh \<equiv> \<lambda>(ST,LT) is (stk,loc,C,M,pc,ics).
- P,h \<turnstile> stk [:\<le>] ST \<and> P,h \<turnstile> loc [:\<le>\<^sub>\<top>] LT \<and> pc < size is \<and> P,h,sh \<turnstile>\<^sub>i (C,M,pc,ics)"
-
-lemma conf_f_def2:
- "conf_f P h sh (ST,LT) is (stk,loc,C,M,pc,ics) \<equiv>
- P,h \<turnstile> stk [:\<le>] ST \<and> P,h \<turnstile> loc [:\<le>\<^sub>\<top>] LT \<and> pc < size is \<and> P,h,sh \<turnstile>\<^sub>i (C,M,pc,ics)"
- by (simp add: conf_f_def)
-
-primrec conf_fs :: "[jvm_prog,heap,sheap,ty\<^sub>P,cname,mname,nat,ty,frame list] \<Rightarrow> bool"
-where
- "conf_fs P h sh \<Phi> C\<^sub>0 M\<^sub>0 n\<^sub>0 T\<^sub>0 [] = True"
-| "conf_fs P h sh \<Phi> C\<^sub>0 M\<^sub>0 n\<^sub>0 T\<^sub>0 (f#frs) =
- (let (stk,loc,C,M,pc,ics) = f in
- (\<exists>ST LT b Ts T mxs mxl\<^sub>0 is xt.
- \<Phi> C M ! pc = Some (ST,LT) \<and>
- (P \<turnstile> C sees M,b:Ts \<rightarrow> T = (mxs,mxl\<^sub>0,is,xt) in C) \<and>
- ((\<exists>D Ts' T' m D'. M\<^sub>0 \<noteq> clinit \<and> ics = No_ics \<and>
- is!pc = Invoke M\<^sub>0 n\<^sub>0 \<and> ST!n\<^sub>0 = Class D \<and>
- P \<turnstile> D sees M\<^sub>0,NonStatic:Ts' \<rightarrow> T' = m in D' \<and> P \<turnstile> C\<^sub>0 \<preceq>\<^sup>* D' \<and> P \<turnstile> T\<^sub>0 \<le> T') \<or>
- (\<exists>D Ts' T' m. M\<^sub>0 \<noteq> clinit \<and> ics = No_ics \<and>
- is!pc = Invokestatic D M\<^sub>0 n\<^sub>0 \<and>
- P \<turnstile> D sees M\<^sub>0,Static:Ts' \<rightarrow> T' = m in C\<^sub>0 \<and> P \<turnstile> T\<^sub>0 \<le> T') \<or>
- (M\<^sub>0 = clinit \<and> (\<exists>Cs. ics = Called Cs))) \<and>
- conf_f P h sh (ST, LT) is f \<and> conf_fs P h sh \<Phi> C M (size Ts) T frs))"
-
-fun ics_classes :: "init_call_status \<Rightarrow> cname list" where
-"ics_classes (Calling C Cs) = Cs" |
-"ics_classes (Throwing Cs a) = Cs" |
-"ics_classes (Called Cs) = Cs" |
-"ics_classes _ = []"
-
-fun frame_clinit_classes :: "frame \<Rightarrow> cname list" where
-"frame_clinit_classes (stk,loc,C,M,pc,ics) = (if M=clinit then [C] else []) @ ics_classes ics"
-
-abbreviation clinit_classes :: "frame list \<Rightarrow> cname list" where
-"clinit_classes frs \<equiv> concat (map frame_clinit_classes frs)"
-
-definition distinct_clinit :: "frame list \<Rightarrow> bool" where
-"distinct_clinit frs \<equiv> distinct (clinit_classes frs)"
-
-definition conf_clinit :: "jvm_prog \<Rightarrow> sheap \<Rightarrow> frame list \<Rightarrow> bool" where
-"conf_clinit P sh frs
- \<equiv> distinct_clinit frs \<and>
- (\<forall>C \<in> set(clinit_classes frs). is_class P C \<and> (\<exists>sfs. sh C = Some(sfs, Processing)))"
-
-(*************************)
-
-definition correct_state :: "[jvm_prog,ty\<^sub>P,jvm_state] \<Rightarrow> bool" ("_,_ \<turnstile> _ \<surd>" [61,0,0] 61)
-where
- "correct_state P \<Phi> \<equiv> \<lambda>(xp,h,frs,sh).
- case xp of
- None \<Rightarrow> (case frs of
- [] \<Rightarrow> True
- | (f#fs) \<Rightarrow> P\<turnstile> h\<surd> \<and> P,h\<turnstile>\<^sub>s sh\<surd> \<and> conf_clinit P sh frs \<and>
- (let (stk,loc,C,M,pc,ics) = f
- in \<exists>b Ts T mxs mxl\<^sub>0 is xt \<tau>.
- (P \<turnstile> C sees M,b:Ts\<rightarrow>T = (mxs,mxl\<^sub>0,is,xt) in C) \<and>
- \<Phi> C M ! pc = Some \<tau> \<and>
- conf_f P h sh \<tau> is f \<and> conf_fs P h sh \<Phi> C M (size Ts) T fs))
- | Some x \<Rightarrow> frs = []"
-
-notation
- correct_state ("_,_ |- _ [ok]" [61,0,0] 61)
-
-subsection \<open> Values and @{text "\<top>"} \<close>
-
-lemma confT_Err [iff]: "P,h \<turnstile> x :\<le>\<^sub>\<top> Err"
- by (simp add: confT_def)
-
-lemma confT_OK [iff]: "P,h \<turnstile> x :\<le>\<^sub>\<top> OK T = (P,h \<turnstile> x :\<le> T)"
- by (simp add: confT_def)
-
-lemma confT_cases:
- "P,h \<turnstile> x :\<le>\<^sub>\<top> X = (X = Err \<or> (\<exists>T. X = OK T \<and> P,h \<turnstile> x :\<le> T))"
- by (cases X) auto
-
-lemma confT_hext [intro?, trans]:
- "\<lbrakk> P,h \<turnstile> x :\<le>\<^sub>\<top> T; h \<unlhd> h' \<rbrakk> \<Longrightarrow> P,h' \<turnstile> x :\<le>\<^sub>\<top> T"
- by (cases T) (blast intro: conf_hext)+
-
-lemma confT_widen [intro?, trans]:
- "\<lbrakk> P,h \<turnstile> x :\<le>\<^sub>\<top> T; P \<turnstile> T \<le>\<^sub>\<top> T' \<rbrakk> \<Longrightarrow> P,h \<turnstile> x :\<le>\<^sub>\<top> T'"
- by (cases T', auto intro: conf_widen)
-
-
-subsection \<open> Stack and Registers \<close>
-
-lemmas confTs_Cons1 [iff] = list_all2_Cons1 [of "confT P h"] for P h
-
-lemma confTs_confT_sup:
-assumes confTs: "P,h \<turnstile> loc [:\<le>\<^sub>\<top>] LT" and n: "n < size LT" and
- LTn: "LT!n = OK T" and subtype: "P \<turnstile> T \<le> T'"
-shows "P,h \<turnstile> (loc!n) :\<le> T'"
-(*<*)
-proof -
- have len: "n < length loc" using list_all2_lengthD[OF confTs] n
- by simp
- show ?thesis
- using list_all2_nthD[OF confTs len] conf_widen[OF _ subtype] LTn
- by simp
-qed
-(*>*)
-
-lemma confTs_hext [intro?]:
- "P,h \<turnstile> loc [:\<le>\<^sub>\<top>] LT \<Longrightarrow> h \<unlhd> h' \<Longrightarrow> P,h' \<turnstile> loc [:\<le>\<^sub>\<top>] LT"
- by (fast elim: list_all2_mono confT_hext)
-
-lemma confTs_widen [intro?, trans]:
- "P,h \<turnstile> loc [:\<le>\<^sub>\<top>] LT \<Longrightarrow> P \<turnstile> LT [\<le>\<^sub>\<top>] LT' \<Longrightarrow> P,h \<turnstile> loc [:\<le>\<^sub>\<top>] LT'"
- by (rule list_all2_trans, rule confT_widen)
-
-lemma confTs_map [iff]:
- "\<And>vs. (P,h \<turnstile> vs [:\<le>\<^sub>\<top>] map OK Ts) = (P,h \<turnstile> vs [:\<le>] Ts)"
- by (induct Ts) (auto simp: list_all2_Cons2)
-
-lemma reg_widen_Err [iff]:
- "\<And>LT. (P \<turnstile> replicate n Err [\<le>\<^sub>\<top>] LT) = (LT = replicate n Err)"
- by (induct n) (auto simp: list_all2_Cons1)
-
-lemma confTs_Err [iff]:
- "P,h \<turnstile> replicate n v [:\<le>\<^sub>\<top>] replicate n Err"
- by (induct n) auto
-
-subsection \<open> valid @{text "init_call_status"} \<close>
-
-lemma valid_ics_shupd:
-assumes "P,h,sh \<turnstile>\<^sub>i (C, M, pc, ics)" and "distinct (C'#ics_classes ics)"
-shows "P,h,sh(C' \<mapsto> (sfs, i')) \<turnstile>\<^sub>i (C, M, pc, ics)"
-using assms by(cases ics; clarsimp simp: fun_upd_apply) fastforce
-
-subsection \<open> correct-frame \<close>
-
-lemma conf_f_Throwing:
-assumes "conf_f P h sh (ST, LT) is (stk, loc, C, M, pc, Called Cs)"
- and "is_class P C'" and "h xcp = Some obj" and "sh C' = Some(sfs,Processing)"
-shows "conf_f P h sh (ST, LT) is (stk, loc, C, M, pc, Throwing (C' # Cs) xcp)"
-using assms by(auto simp: conf_f_def2)
-
-lemma conf_f_shupd:
-assumes "conf_f P h sh (ST,LT) ins f"
- and "i = Processing
- \<or> (distinct (C#ics_classes (ics_of f)) \<and> (curr_method f = clinit \<longrightarrow> C \<noteq> curr_class f))"
-shows "conf_f P h (sh(C \<mapsto> (sfs, i))) (ST,LT) ins f"
-using assms
- by(cases f, cases "ics_of f"; clarsimp simp: conf_f_def2 fun_upd_apply) fastforce+
-
-lemma conf_f_shupd':
-assumes "conf_f P h sh (ST,LT) ins f"
- and "sh C = Some(sfs,i)"
-shows "conf_f P h (sh(C \<mapsto> (sfs', i))) (ST,LT) ins f"
-using assms
- by(cases f, cases "ics_of f"; clarsimp simp: conf_f_def2 fun_upd_apply) fastforce+
-
-subsection \<open> correct-frames \<close>
-
-lemmas [simp del] = fun_upd_apply
-
-lemma conf_fs_hext:
- "\<And>C M n T\<^sub>r.
- \<lbrakk> conf_fs P h sh \<Phi> C M n T\<^sub>r frs; h \<unlhd> h' \<rbrakk> \<Longrightarrow> conf_fs P h' sh \<Phi> C M n T\<^sub>r frs"
-(*<*)
-proof(induct frs)
- case (Cons fr frs)
- obtain stk ls C M pc ics where fr: "fr = (stk, ls, C, M, pc, ics)" by(cases fr) simp
- moreover obtain ST LT where \<Phi>: "\<Phi> C M ! pc = \<lfloor>(ST, LT)\<rfloor>" and
- ST: "P,h \<turnstile> stk [:\<le>] ST" and LT: "P,h \<turnstile> ls [:\<le>\<^sub>\<top>] LT"
- using Cons.prems(1) fr by(auto simp: conf_f_def)
- ultimately show ?case using Cons confs_hext[OF ST Cons(3)] confTs_hext[OF LT Cons(3)]
- by (fastforce simp: conf_f_def)
-qed simp
-(*>*)
-
-
-lemma conf_fs_shupd:
-assumes "conf_fs P h sh \<Phi> C\<^sub>0 M n T frs"
- and dist: "distinct (C#clinit_classes frs)"
-shows "conf_fs P h (sh(C \<mapsto> (sfs, i))) \<Phi> C\<^sub>0 M n T frs"
-using assms proof(induct frs arbitrary: C\<^sub>0 C M n T)
- case (Cons f' frs')
- then obtain stk' loc' C' M' pc' ics' where f': "f' = (stk',loc',C',M',pc',ics')" by(cases f')
- with assms Cons obtain ST LT b Ts T1 mxs mxl\<^sub>0 ins xt where
- ty: "\<Phi> C' M' ! pc' = Some (ST,LT)" and
- meth: "P \<turnstile> C' sees M',b:Ts \<rightarrow> T1 = (mxs,mxl\<^sub>0,ins,xt) in C'" and
- conf: "conf_f P h sh (ST, LT) ins f'" and
- confs: "conf_fs P h sh \<Phi> C' M' (size Ts) T1 frs'" by clarsimp
-
- from f' Cons.prems(2) have
- "distinct (C#ics_classes (ics_of f')) \<and> (curr_method f' = clinit \<longrightarrow> C \<noteq> curr_class f')"
- by fastforce
- with conf_f_shupd[where C=C, OF conf] have
- conf': "conf_f P h (sh(C \<mapsto> (sfs, i))) (ST, LT) ins f'" by simp
-
- from Cons.prems(2) have dist': "distinct (C # clinit_classes frs')"
- by(auto simp: distinct_length_2_or_more)
- from Cons.hyps[OF confs dist'] have
- confs': "conf_fs P h (sh(C \<mapsto> (sfs, i))) \<Phi> C' M' (length Ts) T1 frs'" by simp
-
- from conf' confs' ty meth f' Cons.prems show ?case by(fastforce dest: sees_method_fun)
-qed(simp)
-
-lemma conf_fs_shupd':
-assumes "conf_fs P h sh \<Phi> C\<^sub>0 M n T frs"
- and shC: "sh C = Some(sfs,i)"
-shows "conf_fs P h (sh(C \<mapsto> (sfs', i))) \<Phi> C\<^sub>0 M n T frs"
-using assms proof(induct frs arbitrary: C\<^sub>0 C M n T sfs i sfs')
- case (Cons f' frs')
- then obtain stk' loc' C' M' pc' ics' where f': "f' = (stk',loc',C',M',pc',ics')" by(cases f')
- with assms Cons obtain ST LT b Ts T1 mxs mxl\<^sub>0 ins xt where
- ty: "\<Phi> C' M' ! pc' = Some (ST,LT)" and
- meth: "P \<turnstile> C' sees M',b:Ts \<rightarrow> T1 = (mxs,mxl\<^sub>0,ins,xt) in C'" and
- conf: "conf_f P h sh (ST, LT) ins f'" and
- confs: "conf_fs P h sh \<Phi> C' M' (size Ts) T1 frs'" and
- shC': "sh C = Some(sfs,i)" by clarsimp
-
- have conf': "conf_f P h (sh(C \<mapsto> (sfs', i))) (ST, LT) ins f'" by(rule conf_f_shupd'[OF conf shC'])
-
- from Cons.hyps[OF confs shC'] have
- confs': "conf_fs P h (sh(C \<mapsto> (sfs', i))) \<Phi> C' M' (length Ts) T1 frs'" by simp
-
- from conf' confs' ty meth f' Cons.prems show ?case by(fastforce dest: sees_method_fun)
-qed(simp)
-
-subsection \<open> correctness wrt @{term clinit} use \<close>
-
-lemma conf_clinit_Cons:
-assumes "conf_clinit P sh (f#frs)"
-shows "conf_clinit P sh frs"
-proof -
- from assms have dist: "distinct_clinit (f#frs)"
- by(cases "curr_method f = clinit", auto simp: conf_clinit_def)
- then have dist': "distinct_clinit frs" by(simp add: distinct_clinit_def)
-
- with assms show ?thesis by(cases frs; fastforce simp: conf_clinit_def)
-qed
-
-lemma conf_clinit_Cons_Cons:
- "conf_clinit P sh (f'#f#frs) \<Longrightarrow> conf_clinit P sh (f'#frs)"
- by(auto simp: conf_clinit_def distinct_clinit_def)
-
-lemma conf_clinit_diff:
-assumes "conf_clinit P sh ((stk,loc,C,M,pc,ics)#frs)"
-shows "conf_clinit P sh ((stk',loc',C,M,pc',ics)#frs)"
-using assms by(cases "M = clinit", simp_all add: conf_clinit_def distinct_clinit_def)
-
-lemma conf_clinit_diff':
-assumes "conf_clinit P sh ((stk,loc,C,M,pc,ics)#frs)"
-shows "conf_clinit P sh ((stk',loc',C,M,pc',No_ics)#frs)"
-using assms by(cases "M = clinit", simp_all add: conf_clinit_def distinct_clinit_def)
-
-lemma conf_clinit_Called_Throwing:
- "conf_clinit P sh ((stk', loc', C', clinit, pc', ics') # (stk, loc, C, M, pc, Called Cs) # fs)
- \<Longrightarrow> conf_clinit P sh ((stk, loc, C, M, pc, Throwing (C' # Cs) xcp) # fs)"
- by(simp add: conf_clinit_def distinct_clinit_def)
-
-lemma conf_clinit_Throwing:
- "conf_clinit P sh ((stk, loc, C, M, pc, Throwing (C'#Cs) xcp) # fs)
- \<Longrightarrow> conf_clinit P sh ((stk, loc, C, M, pc, Throwing Cs xcp) # fs)"
- by(simp add: conf_clinit_def distinct_clinit_def)
-
-lemma conf_clinit_Called:
- "\<lbrakk> conf_clinit P sh ((stk, loc, C, M, pc, Called (C'#Cs)) # frs);
- P \<turnstile> C' sees clinit,Static: [] \<rightarrow> Void=(mxs',mxl',ins',xt') in C' \<rbrakk>
- \<Longrightarrow> conf_clinit P sh (create_init_frame P C' # (stk, loc, C, M, pc, Called Cs) # frs)"
- by(simp add: conf_clinit_def distinct_clinit_def)
-
-lemma conf_clinit_Cons_nclinit:
-assumes "conf_clinit P sh frs" and nclinit: "M \<noteq> clinit"
-shows "conf_clinit P sh ((stk, loc, C, M, pc, No_ics) # frs)"
-proof -
- from nclinit
- have "clinit_classes ((stk, loc, C, M, pc, No_ics) # frs) = clinit_classes frs" by simp
- with assms show ?thesis by(simp add: conf_clinit_def distinct_clinit_def)
-qed
-
-lemma conf_clinit_Invoke:
-assumes "conf_clinit P sh ((stk, loc, C, M, pc, ics) # frs)" and "M' \<noteq> clinit"
-shows "conf_clinit P sh ((stk', loc', C', M', pc', No_ics) # (stk, loc, C, M, pc, No_ics) # frs)"
- using assms conf_clinit_Cons_nclinit conf_clinit_diff' by auto
-
-lemma conf_clinit_nProc_dist:
-assumes "conf_clinit P sh frs"
- and "\<forall>sfs. sh C \<noteq> Some(sfs,Processing)"
-shows "distinct (C # clinit_classes frs)"
-using assms by(auto simp: conf_clinit_def distinct_clinit_def)
-
-
-lemma conf_clinit_shupd:
-assumes "conf_clinit P sh frs"
- and dist: "distinct (C#clinit_classes frs)"
-shows "conf_clinit P (sh(C \<mapsto> (sfs, i))) frs"
-using assms by(simp add: conf_clinit_def fun_upd_apply)
-
-lemma conf_clinit_shupd':
-assumes "conf_clinit P sh frs"
- and "sh C = Some(sfs,i)"
-shows "conf_clinit P (sh(C \<mapsto> (sfs', i))) frs"
-using assms by(fastforce simp: conf_clinit_def fun_upd_apply)
-
-lemma conf_clinit_shupd_Called:
-assumes "conf_clinit P sh ((stk,loc,C,M,pc,Calling C' Cs)#frs)"
- and dist: "distinct (C'#clinit_classes ((stk,loc,C,M,pc,Calling C' Cs)#frs))"
- and cls: "is_class P C'"
-shows "conf_clinit P (sh(C' \<mapsto> (sfs, Processing))) ((stk,loc,C,M,pc,Called (C'#Cs))#frs)"
-using assms by(clarsimp simp: conf_clinit_def fun_upd_apply distinct_clinit_def)
-
-lemma conf_clinit_shupd_Calling:
-assumes "conf_clinit P sh ((stk,loc,C,M,pc,Calling C' Cs)#frs)"
- and dist: "distinct (C'#clinit_classes ((stk,loc,C,M,pc,Calling C' Cs)#frs))"
- and cls: "is_class P C'"
-shows "conf_clinit P (sh(C' \<mapsto> (sfs, Processing)))
- ((stk,loc,C,M,pc,Calling (fst(the(class P C'))) (C'#Cs))#frs)"
-using assms by(clarsimp simp: conf_clinit_def fun_upd_apply distinct_clinit_def)
-
-subsection \<open> correct state \<close>
-
-lemma correct_state_Cons:
-assumes cr: "P,\<Phi> |- (xp,h,f#frs,sh) [ok]"
-shows "P,\<Phi> |- (xp,h,frs,sh) [ok]"
-proof -
- from cr have dist: "conf_clinit P sh (f#frs)"
- by(simp add: correct_state_def)
- then have "conf_clinit P sh frs" by(rule conf_clinit_Cons)
-
- with cr show ?thesis by(cases frs; fastforce simp: correct_state_def)
-qed
-
-lemma correct_state_shupd:
-assumes cs: "P,\<Phi> |- (xp,h,frs,sh) [ok]" and shC: "sh C = Some(sfs,i)"
- and dist: "distinct (C#clinit_classes frs)"
-shows "P,\<Phi> |- (xp,h,frs,sh(C \<mapsto> (sfs, i'))) [ok]"
-using assms
-proof(cases xp)
- case None with assms show ?thesis
- proof(cases frs)
- case (Cons f' frs')
- let ?sh = "sh(C \<mapsto> (sfs, i'))"
-
- obtain stk' loc' C' M' pc' ics' where f': "f' = (stk',loc',C',M',pc',ics')" by(cases f')
- with cs Cons None obtain b Ts T mxs mxl\<^sub>0 ins xt ST LT where
- meth: "P \<turnstile> C' sees M',b:Ts\<rightarrow>T = (mxs,mxl\<^sub>0,ins,xt) in C'"
- and ty: "\<Phi> C' M' ! pc' = Some (ST,LT)" and conf: "conf_f P h sh (ST,LT) ins f'"
- and confs: "conf_fs P h sh \<Phi> C' M' (size Ts) T frs'"
- and confc: "conf_clinit P sh frs"
- and h_ok: "P\<turnstile> h\<surd>" and sh_ok: "P,h \<turnstile>\<^sub>s sh \<surd>"
- by(auto simp: correct_state_def)
-
- from Cons dist have dist': "distinct (C#clinit_classes frs')"
- by(auto simp: distinct_length_2_or_more)
-
- from shconf_upd_obj[OF sh_ok shconfD[OF sh_ok shC]] have sh_ok': "P,h \<turnstile>\<^sub>s ?sh \<surd>"
- by simp
-
- from conf f' valid_ics_shupd Cons dist have conf': "conf_f P h ?sh (ST,LT) ins f'"
- by(auto simp: conf_f_def2 fun_upd_apply)
- have confs': "conf_fs P h ?sh \<Phi> C' M' (size Ts) T frs'" by(rule conf_fs_shupd[OF confs dist'])
-
- have confc': "conf_clinit P ?sh frs" by(rule conf_clinit_shupd[OF confc dist])
-
- with h_ok sh_ok' meth ty conf' confs' f' Cons None show ?thesis
- by(fastforce simp: correct_state_def)
- qed(simp add: correct_state_def)
-qed(simp add: correct_state_def)
-
-lemma correct_state_Throwing_ex:
-assumes correct: "P,\<Phi> \<turnstile> (xp,h,(stk,loc,C,M,pc,ics)#frs,sh)\<surd>"
-shows "\<And>Cs a. ics = Throwing Cs a \<Longrightarrow> \<exists>obj. h a = Some obj"
-using correct by(clarsimp simp: correct_state_def conf_f_def)
-
-end
+(* Title: JinjaDCI/BV/BVConform.thy
+
+ Author: Cornelia Pusch, Gerwin Klein, Susannah Mansky
+ Copyright 1999 Technische Universitaet Muenchen, 2019-20 UIUC
+
+ Based on the Jinja theory BV/BVConform.thy by Cornelia Pusch and Gerwin Klein
+
+The invariant for the type safety proof.
+*)
+
+section \<open> BV Type Safety Invariant \<close>
+
+theory BVConform
+imports BVSpec "../JVM/JVMExec" "../Common/Conform"
+begin
+
+subsection \<open> @{text "correct_state"} definitions \<close>
+
+definition confT :: "'c prog \<Rightarrow> heap \<Rightarrow> val \<Rightarrow> ty err \<Rightarrow> bool"
+ ("_,_ \<turnstile> _ :\<le>\<^sub>\<top> _" [51,51,51,51] 50)
+where
+ "P,h \<turnstile> v :\<le>\<^sub>\<top> E \<equiv> case E of Err \<Rightarrow> True | OK T \<Rightarrow> P,h \<turnstile> v :\<le> T"
+
+notation (ASCII)
+ confT ("_,_ |- _ :<=T _" [51,51,51,51] 50)
+
+abbreviation
+ confTs :: "'c prog \<Rightarrow> heap \<Rightarrow> val list \<Rightarrow> ty\<^sub>l \<Rightarrow> bool"
+ ("_,_ \<turnstile> _ [:\<le>\<^sub>\<top>] _" [51,51,51,51] 50) where
+ "P,h \<turnstile> vs [:\<le>\<^sub>\<top>] Ts \<equiv> list_all2 (confT P h) vs Ts"
+
+notation (ASCII)
+ confTs ("_,_ |- _ [:<=T] _" [51,51,51,51] 50)
+
+fun Called_context :: "jvm_prog \<Rightarrow> cname \<Rightarrow> instr \<Rightarrow> bool" where
+"Called_context P C\<^sub>0 (New C') = (C\<^sub>0=C')" |
+"Called_context P C\<^sub>0 (Getstatic C F D) = ((C\<^sub>0=D) \<and> (\<exists>t. P \<turnstile> C has F,Static:t in D))" |
+"Called_context P C\<^sub>0 (Putstatic C F D) = ((C\<^sub>0=D) \<and> (\<exists>t. P \<turnstile> C has F,Static:t in D))" |
+"Called_context P C\<^sub>0 (Invokestatic C M n)
+ = (\<exists>Ts T m D. (C\<^sub>0=D) \<and> P \<turnstile> C sees M,Static:Ts \<rightarrow> T = m in D)" |
+"Called_context P _ _ = False"
+
+abbreviation Called_set :: "instr set" where
+"Called_set \<equiv> {i. \<exists>C. i = New C} \<union> {i. \<exists>C M n. i = Invokestatic C M n}
+ \<union> {i. \<exists>C F D. i = Getstatic C F D} \<union> {i. \<exists>C F D. i = Putstatic C F D}"
+
+lemma Called_context_Called_set:
+ "Called_context P D i \<Longrightarrow> i \<in> Called_set" by(cases i, auto)
+
+fun valid_ics :: "jvm_prog \<Rightarrow> heap \<Rightarrow> sheap \<Rightarrow> cname \<times> mname \<times> pc \<times> init_call_status \<Rightarrow> bool"
+ ("_,_,_ \<turnstile>\<^sub>i _" [51,51,51,51] 50) where
+"valid_ics P h sh (C,M,pc,Calling C' Cs)
+ = (let ins = instrs_of P C M in Called_context P (last (C'#Cs)) (ins!pc)
+ \<and> is_class P C')" |
+"valid_ics P h sh (C,M,pc,Throwing Cs a)
+ =(let ins = instrs_of P C M in \<exists>C1. Called_context P C1 (ins!pc)
+ \<and> (\<exists>obj. h a = Some obj))" |
+"valid_ics P h sh (C,M,pc,Called Cs)
+ = (let ins = instrs_of P C M
+ in \<exists>C1 sobj. Called_context P C1 (ins!pc) \<and> sh C1 = Some sobj)" |
+"valid_ics P _ _ _ = True"
+
+definition conf_f :: "jvm_prog \<Rightarrow> heap \<Rightarrow> sheap \<Rightarrow> ty\<^sub>i \<Rightarrow> bytecode \<Rightarrow> frame \<Rightarrow> bool"
+where
+ "conf_f P h sh \<equiv> \<lambda>(ST,LT) is (stk,loc,C,M,pc,ics).
+ P,h \<turnstile> stk [:\<le>] ST \<and> P,h \<turnstile> loc [:\<le>\<^sub>\<top>] LT \<and> pc < size is \<and> P,h,sh \<turnstile>\<^sub>i (C,M,pc,ics)"
+
+lemma conf_f_def2:
+ "conf_f P h sh (ST,LT) is (stk,loc,C,M,pc,ics) \<equiv>
+ P,h \<turnstile> stk [:\<le>] ST \<and> P,h \<turnstile> loc [:\<le>\<^sub>\<top>] LT \<and> pc < size is \<and> P,h,sh \<turnstile>\<^sub>i (C,M,pc,ics)"
+ by (simp add: conf_f_def)
+
+primrec conf_fs :: "[jvm_prog,heap,sheap,ty\<^sub>P,cname,mname,nat,ty,frame list] \<Rightarrow> bool"
+where
+ "conf_fs P h sh \<Phi> C\<^sub>0 M\<^sub>0 n\<^sub>0 T\<^sub>0 [] = True"
+| "conf_fs P h sh \<Phi> C\<^sub>0 M\<^sub>0 n\<^sub>0 T\<^sub>0 (f#frs) =
+ (let (stk,loc,C,M,pc,ics) = f in
+ (\<exists>ST LT b Ts T mxs mxl\<^sub>0 is xt.
+ \<Phi> C M ! pc = Some (ST,LT) \<and>
+ (P \<turnstile> C sees M,b:Ts \<rightarrow> T = (mxs,mxl\<^sub>0,is,xt) in C) \<and>
+ ((\<exists>D Ts' T' m D'. M\<^sub>0 \<noteq> clinit \<and> ics = No_ics \<and>
+ is!pc = Invoke M\<^sub>0 n\<^sub>0 \<and> ST!n\<^sub>0 = Class D \<and>
+ P \<turnstile> D sees M\<^sub>0,NonStatic:Ts' \<rightarrow> T' = m in D' \<and> P \<turnstile> C\<^sub>0 \<preceq>\<^sup>* D' \<and> P \<turnstile> T\<^sub>0 \<le> T') \<or>
+ (\<exists>D Ts' T' m. M\<^sub>0 \<noteq> clinit \<and> ics = No_ics \<and>
+ is!pc = Invokestatic D M\<^sub>0 n\<^sub>0 \<and>
+ P \<turnstile> D sees M\<^sub>0,Static:Ts' \<rightarrow> T' = m in C\<^sub>0 \<and> P \<turnstile> T\<^sub>0 \<le> T') \<or>
+ (M\<^sub>0 = clinit \<and> (\<exists>Cs. ics = Called Cs))) \<and>
+ conf_f P h sh (ST, LT) is f \<and> conf_fs P h sh \<Phi> C M (size Ts) T frs))"
+
+fun ics_classes :: "init_call_status \<Rightarrow> cname list" where
+"ics_classes (Calling C Cs) = Cs" |
+"ics_classes (Throwing Cs a) = Cs" |
+"ics_classes (Called Cs) = Cs" |
+"ics_classes _ = []"
+
+fun frame_clinit_classes :: "frame \<Rightarrow> cname list" where
+"frame_clinit_classes (stk,loc,C,M,pc,ics) = (if M=clinit then [C] else []) @ ics_classes ics"
+
+abbreviation clinit_classes :: "frame list \<Rightarrow> cname list" where
+"clinit_classes frs \<equiv> concat (map frame_clinit_classes frs)"
+
+definition distinct_clinit :: "frame list \<Rightarrow> bool" where
+"distinct_clinit frs \<equiv> distinct (clinit_classes frs)"
+
+definition conf_clinit :: "jvm_prog \<Rightarrow> sheap \<Rightarrow> frame list \<Rightarrow> bool" where
+"conf_clinit P sh frs
+ \<equiv> distinct_clinit frs \<and>
+ (\<forall>C \<in> set(clinit_classes frs). is_class P C \<and> (\<exists>sfs. sh C = Some(sfs, Processing)))"
+
+(*************************)
+
+definition correct_state :: "[jvm_prog,ty\<^sub>P,jvm_state] \<Rightarrow> bool" ("_,_ \<turnstile> _ \<surd>" [61,0,0] 61)
+where
+ "correct_state P \<Phi> \<equiv> \<lambda>(xp,h,frs,sh).
+ case xp of
+ None \<Rightarrow> (case frs of
+ [] \<Rightarrow> True
+ | (f#fs) \<Rightarrow> P\<turnstile> h\<surd> \<and> P,h\<turnstile>\<^sub>s sh\<surd> \<and> conf_clinit P sh frs \<and>
+ (let (stk,loc,C,M,pc,ics) = f
+ in \<exists>b Ts T mxs mxl\<^sub>0 is xt \<tau>.
+ (P \<turnstile> C sees M,b:Ts\<rightarrow>T = (mxs,mxl\<^sub>0,is,xt) in C) \<and>
+ \<Phi> C M ! pc = Some \<tau> \<and>
+ conf_f P h sh \<tau> is f \<and> conf_fs P h sh \<Phi> C M (size Ts) T fs))
+ | Some x \<Rightarrow> frs = []"
+
+notation
+ correct_state ("_,_ |- _ [ok]" [61,0,0] 61)
+
+subsection \<open> Values and @{text "\<top>"} \<close>
+
+lemma confT_Err [iff]: "P,h \<turnstile> x :\<le>\<^sub>\<top> Err"
+ by (simp add: confT_def)
+
+lemma confT_OK [iff]: "P,h \<turnstile> x :\<le>\<^sub>\<top> OK T = (P,h \<turnstile> x :\<le> T)"
+ by (simp add: confT_def)
+
+lemma confT_cases:
+ "P,h \<turnstile> x :\<le>\<^sub>\<top> X = (X = Err \<or> (\<exists>T. X = OK T \<and> P,h \<turnstile> x :\<le> T))"
+ by (cases X) auto
+
+lemma confT_hext [intro?, trans]:
+ "\<lbrakk> P,h \<turnstile> x :\<le>\<^sub>\<top> T; h \<unlhd> h' \<rbrakk> \<Longrightarrow> P,h' \<turnstile> x :\<le>\<^sub>\<top> T"
+ by (cases T) (blast intro: conf_hext)+
+
+lemma confT_widen [intro?, trans]:
+ "\<lbrakk> P,h \<turnstile> x :\<le>\<^sub>\<top> T; P \<turnstile> T \<le>\<^sub>\<top> T' \<rbrakk> \<Longrightarrow> P,h \<turnstile> x :\<le>\<^sub>\<top> T'"
+ by (cases T', auto intro: conf_widen)
+
+
+subsection \<open> Stack and Registers \<close>
+
+lemmas confTs_Cons1 [iff] = list_all2_Cons1 [of "confT P h"] for P h
+
+lemma confTs_confT_sup:
+assumes confTs: "P,h \<turnstile> loc [:\<le>\<^sub>\<top>] LT" and n: "n < size LT" and
+ LTn: "LT!n = OK T" and subtype: "P \<turnstile> T \<le> T'"
+shows "P,h \<turnstile> (loc!n) :\<le> T'"
+(*<*)
+proof -
+ have len: "n < length loc" using list_all2_lengthD[OF confTs] n
+ by simp
+ show ?thesis
+ using list_all2_nthD[OF confTs len] conf_widen[OF _ subtype] LTn
+ by simp
+qed
+(*>*)
+
+lemma confTs_hext [intro?]:
+ "P,h \<turnstile> loc [:\<le>\<^sub>\<top>] LT \<Longrightarrow> h \<unlhd> h' \<Longrightarrow> P,h' \<turnstile> loc [:\<le>\<^sub>\<top>] LT"
+ by (fast elim: list_all2_mono confT_hext)
+
+lemma confTs_widen [intro?, trans]:
+ "P,h \<turnstile> loc [:\<le>\<^sub>\<top>] LT \<Longrightarrow> P \<turnstile> LT [\<le>\<^sub>\<top>] LT' \<Longrightarrow> P,h \<turnstile> loc [:\<le>\<^sub>\<top>] LT'"
+ by (rule list_all2_trans, rule confT_widen)
+
+lemma confTs_map [iff]:
+ "\<And>vs. (P,h \<turnstile> vs [:\<le>\<^sub>\<top>] map OK Ts) = (P,h \<turnstile> vs [:\<le>] Ts)"
+ by (induct Ts) (auto simp: list_all2_Cons2)
+
+lemma reg_widen_Err [iff]:
+ "\<And>LT. (P \<turnstile> replicate n Err [\<le>\<^sub>\<top>] LT) = (LT = replicate n Err)"
+ by (induct n) (auto simp: list_all2_Cons1)
+
+lemma confTs_Err [iff]:
+ "P,h \<turnstile> replicate n v [:\<le>\<^sub>\<top>] replicate n Err"
+ by (induct n) auto
+
+subsection \<open> valid @{text "init_call_status"} \<close>
+
+lemma valid_ics_shupd:
+assumes "P,h,sh \<turnstile>\<^sub>i (C, M, pc, ics)" and "distinct (C'#ics_classes ics)"
+shows "P,h,sh(C' \<mapsto> (sfs, i')) \<turnstile>\<^sub>i (C, M, pc, ics)"
+using assms by(cases ics; clarsimp simp: fun_upd_apply) fastforce
+
+subsection \<open> correct-frame \<close>
+
+lemma conf_f_Throwing:
+assumes "conf_f P h sh (ST, LT) is (stk, loc, C, M, pc, Called Cs)"
+ and "is_class P C'" and "h xcp = Some obj" and "sh C' = Some(sfs,Processing)"
+shows "conf_f P h sh (ST, LT) is (stk, loc, C, M, pc, Throwing (C' # Cs) xcp)"
+using assms by(auto simp: conf_f_def2)
+
+lemma conf_f_shupd:
+assumes "conf_f P h sh (ST,LT) ins f"
+ and "i = Processing
+ \<or> (distinct (C#ics_classes (ics_of f)) \<and> (curr_method f = clinit \<longrightarrow> C \<noteq> curr_class f))"
+shows "conf_f P h (sh(C \<mapsto> (sfs, i))) (ST,LT) ins f"
+using assms
+ by(cases f, cases "ics_of f"; clarsimp simp: conf_f_def2 fun_upd_apply) fastforce+
+
+lemma conf_f_shupd':
+assumes "conf_f P h sh (ST,LT) ins f"
+ and "sh C = Some(sfs,i)"
+shows "conf_f P h (sh(C \<mapsto> (sfs', i))) (ST,LT) ins f"
+using assms
+ by(cases f, cases "ics_of f"; clarsimp simp: conf_f_def2 fun_upd_apply) fastforce+
+
+subsection \<open> correct-frames \<close>
+
+lemmas [simp del] = fun_upd_apply
+
+lemma conf_fs_hext:
+ "\<And>C M n T\<^sub>r.
+ \<lbrakk> conf_fs P h sh \<Phi> C M n T\<^sub>r frs; h \<unlhd> h' \<rbrakk> \<Longrightarrow> conf_fs P h' sh \<Phi> C M n T\<^sub>r frs"
+(*<*)
+proof(induct frs)
+ case (Cons fr frs)
+ obtain stk ls C M pc ics where fr: "fr = (stk, ls, C, M, pc, ics)" by(cases fr) simp
+ moreover obtain ST LT where \<Phi>: "\<Phi> C M ! pc = \<lfloor>(ST, LT)\<rfloor>" and
+ ST: "P,h \<turnstile> stk [:\<le>] ST" and LT: "P,h \<turnstile> ls [:\<le>\<^sub>\<top>] LT"
+ using Cons.prems(1) fr by(auto simp: conf_f_def)
+ ultimately show ?case using Cons confs_hext[OF ST Cons(3)] confTs_hext[OF LT Cons(3)]
+ by (fastforce simp: conf_f_def)
+qed simp
+(*>*)
+
+
+lemma conf_fs_shupd:
+assumes "conf_fs P h sh \<Phi> C\<^sub>0 M n T frs"
+ and dist: "distinct (C#clinit_classes frs)"
+shows "conf_fs P h (sh(C \<mapsto> (sfs, i))) \<Phi> C\<^sub>0 M n T frs"
+using assms proof(induct frs arbitrary: C\<^sub>0 C M n T)
+ case (Cons f' frs')
+ then obtain stk' loc' C' M' pc' ics' where f': "f' = (stk',loc',C',M',pc',ics')" by(cases f')
+ with assms Cons obtain ST LT b Ts T1 mxs mxl\<^sub>0 ins xt where
+ ty: "\<Phi> C' M' ! pc' = Some (ST,LT)" and
+ meth: "P \<turnstile> C' sees M',b:Ts \<rightarrow> T1 = (mxs,mxl\<^sub>0,ins,xt) in C'" and
+ conf: "conf_f P h sh (ST, LT) ins f'" and
+ confs: "conf_fs P h sh \<Phi> C' M' (size Ts) T1 frs'" by clarsimp
+
+ from f' Cons.prems(2) have
+ "distinct (C#ics_classes (ics_of f')) \<and> (curr_method f' = clinit \<longrightarrow> C \<noteq> curr_class f')"
+ by fastforce
+ with conf_f_shupd[where C=C, OF conf] have
+ conf': "conf_f P h (sh(C \<mapsto> (sfs, i))) (ST, LT) ins f'" by simp
+
+ from Cons.prems(2) have dist': "distinct (C # clinit_classes frs')"
+ by(auto simp: distinct_length_2_or_more)
+ from Cons.hyps[OF confs dist'] have
+ confs': "conf_fs P h (sh(C \<mapsto> (sfs, i))) \<Phi> C' M' (length Ts) T1 frs'" by simp
+
+ from conf' confs' ty meth f' Cons.prems show ?case by(fastforce dest: sees_method_fun)
+qed(simp)
+
+lemma conf_fs_shupd':
+assumes "conf_fs P h sh \<Phi> C\<^sub>0 M n T frs"
+ and shC: "sh C = Some(sfs,i)"
+shows "conf_fs P h (sh(C \<mapsto> (sfs', i))) \<Phi> C\<^sub>0 M n T frs"
+using assms proof(induct frs arbitrary: C\<^sub>0 C M n T sfs i sfs')
+ case (Cons f' frs')
+ then obtain stk' loc' C' M' pc' ics' where f': "f' = (stk',loc',C',M',pc',ics')" by(cases f')
+ with assms Cons obtain ST LT b Ts T1 mxs mxl\<^sub>0 ins xt where
+ ty: "\<Phi> C' M' ! pc' = Some (ST,LT)" and
+ meth: "P \<turnstile> C' sees M',b:Ts \<rightarrow> T1 = (mxs,mxl\<^sub>0,ins,xt) in C'" and
+ conf: "conf_f P h sh (ST, LT) ins f'" and
+ confs: "conf_fs P h sh \<Phi> C' M' (size Ts) T1 frs'" and
+ shC': "sh C = Some(sfs,i)" by clarsimp
+
+ have conf': "conf_f P h (sh(C \<mapsto> (sfs', i))) (ST, LT) ins f'" by(rule conf_f_shupd'[OF conf shC'])
+
+ from Cons.hyps[OF confs shC'] have
+ confs': "conf_fs P h (sh(C \<mapsto> (sfs', i))) \<Phi> C' M' (length Ts) T1 frs'" by simp
+
+ from conf' confs' ty meth f' Cons.prems show ?case by(fastforce dest: sees_method_fun)
+qed(simp)
+
+subsection \<open> correctness wrt @{term clinit} use \<close>
+
+lemma conf_clinit_Cons:
+assumes "conf_clinit P sh (f#frs)"
+shows "conf_clinit P sh frs"
+proof -
+ from assms have dist: "distinct_clinit (f#frs)"
+ by(cases "curr_method f = clinit", auto simp: conf_clinit_def)
+ then have dist': "distinct_clinit frs" by(simp add: distinct_clinit_def)
+
+ with assms show ?thesis by(cases frs; fastforce simp: conf_clinit_def)
+qed
+
+lemma conf_clinit_Cons_Cons:
+ "conf_clinit P sh (f'#f#frs) \<Longrightarrow> conf_clinit P sh (f'#frs)"
+ by(auto simp: conf_clinit_def distinct_clinit_def)
+
+lemma conf_clinit_diff:
+assumes "conf_clinit P sh ((stk,loc,C,M,pc,ics)#frs)"
+shows "conf_clinit P sh ((stk',loc',C,M,pc',ics)#frs)"
+using assms by(cases "M = clinit", simp_all add: conf_clinit_def distinct_clinit_def)
+
+lemma conf_clinit_diff':
+assumes "conf_clinit P sh ((stk,loc,C,M,pc,ics)#frs)"
+shows "conf_clinit P sh ((stk',loc',C,M,pc',No_ics)#frs)"
+using assms by(cases "M = clinit", simp_all add: conf_clinit_def distinct_clinit_def)
+
+lemma conf_clinit_Called_Throwing:
+ "conf_clinit P sh ((stk', loc', C', clinit, pc', ics') # (stk, loc, C, M, pc, Called Cs) # fs)
+ \<Longrightarrow> conf_clinit P sh ((stk, loc, C, M, pc, Throwing (C' # Cs) xcp) # fs)"
+ by(simp add: conf_clinit_def distinct_clinit_def)
+
+lemma conf_clinit_Throwing:
+ "conf_clinit P sh ((stk, loc, C, M, pc, Throwing (C'#Cs) xcp) # fs)
+ \<Longrightarrow> conf_clinit P sh ((stk, loc, C, M, pc, Throwing Cs xcp) # fs)"
+ by(simp add: conf_clinit_def distinct_clinit_def)
+
+lemma conf_clinit_Called:
+ "\<lbrakk> conf_clinit P sh ((stk, loc, C, M, pc, Called (C'#Cs)) # frs);
+ P \<turnstile> C' sees clinit,Static: [] \<rightarrow> Void=(mxs',mxl',ins',xt') in C' \<rbrakk>
+ \<Longrightarrow> conf_clinit P sh (create_init_frame P C' # (stk, loc, C, M, pc, Called Cs) # frs)"
+ by(simp add: conf_clinit_def distinct_clinit_def)
+
+lemma conf_clinit_Cons_nclinit:
+assumes "conf_clinit P sh frs" and nclinit: "M \<noteq> clinit"
+shows "conf_clinit P sh ((stk, loc, C, M, pc, No_ics) # frs)"
+proof -
+ from nclinit
+ have "clinit_classes ((stk, loc, C, M, pc, No_ics) # frs) = clinit_classes frs" by simp
+ with assms show ?thesis by(simp add: conf_clinit_def distinct_clinit_def)
+qed
+
+lemma conf_clinit_Invoke:
+assumes "conf_clinit P sh ((stk, loc, C, M, pc, ics) # frs)" and "M' \<noteq> clinit"
+shows "conf_clinit P sh ((stk', loc', C', M', pc', No_ics) # (stk, loc, C, M, pc, No_ics) # frs)"
+ using assms conf_clinit_Cons_nclinit conf_clinit_diff' by auto
+
+lemma conf_clinit_nProc_dist:
+assumes "conf_clinit P sh frs"
+ and "\<forall>sfs. sh C \<noteq> Some(sfs,Processing)"
+shows "distinct (C # clinit_classes frs)"
+using assms by(auto simp: conf_clinit_def distinct_clinit_def)
+
+
+lemma conf_clinit_shupd:
+assumes "conf_clinit P sh frs"
+ and dist: "distinct (C#clinit_classes frs)"
+shows "conf_clinit P (sh(C \<mapsto> (sfs, i))) frs"
+using assms by(simp add: conf_clinit_def fun_upd_apply)
+
+lemma conf_clinit_shupd':
+assumes "conf_clinit P sh frs"
+ and "sh C = Some(sfs,i)"
+shows "conf_clinit P (sh(C \<mapsto> (sfs', i))) frs"
+using assms by(fastforce simp: conf_clinit_def fun_upd_apply)
+
+lemma conf_clinit_shupd_Called:
+assumes "conf_clinit P sh ((stk,loc,C,M,pc,Calling C' Cs)#frs)"
+ and dist: "distinct (C'#clinit_classes ((stk,loc,C,M,pc,Calling C' Cs)#frs))"
+ and cls: "is_class P C'"
+shows "conf_clinit P (sh(C' \<mapsto> (sfs, Processing))) ((stk,loc,C,M,pc,Called (C'#Cs))#frs)"
+using assms by(clarsimp simp: conf_clinit_def fun_upd_apply distinct_clinit_def)
+
+lemma conf_clinit_shupd_Calling:
+assumes "conf_clinit P sh ((stk,loc,C,M,pc,Calling C' Cs)#frs)"
+ and dist: "distinct (C'#clinit_classes ((stk,loc,C,M,pc,Calling C' Cs)#frs))"
+ and cls: "is_class P C'"
+shows "conf_clinit P (sh(C' \<mapsto> (sfs, Processing)))
+ ((stk,loc,C,M,pc,Calling (fst(the(class P C'))) (C'#Cs))#frs)"
+using assms by(clarsimp simp: conf_clinit_def fun_upd_apply distinct_clinit_def)
+
+subsection \<open> correct state \<close>
+
+lemma correct_state_Cons:
+assumes cr: "P,\<Phi> |- (xp,h,f#frs,sh) [ok]"
+shows "P,\<Phi> |- (xp,h,frs,sh) [ok]"
+proof -
+ from cr have dist: "conf_clinit P sh (f#frs)"
+ by(simp add: correct_state_def)
+ then have "conf_clinit P sh frs" by(rule conf_clinit_Cons)
+
+ with cr show ?thesis by(cases frs; fastforce simp: correct_state_def)
+qed
+
+lemma correct_state_shupd:
+assumes cs: "P,\<Phi> |- (xp,h,frs,sh) [ok]" and shC: "sh C = Some(sfs,i)"
+ and dist: "distinct (C#clinit_classes frs)"
+shows "P,\<Phi> |- (xp,h,frs,sh(C \<mapsto> (sfs, i'))) [ok]"
+using assms
+proof(cases xp)
+ case None with assms show ?thesis
+ proof(cases frs)
+ case (Cons f' frs')
+ let ?sh = "sh(C \<mapsto> (sfs, i'))"
+
+ obtain stk' loc' C' M' pc' ics' where f': "f' = (stk',loc',C',M',pc',ics')" by(cases f')
+ with cs Cons None obtain b Ts T mxs mxl\<^sub>0 ins xt ST LT where
+ meth: "P \<turnstile> C' sees M',b:Ts\<rightarrow>T = (mxs,mxl\<^sub>0,ins,xt) in C'"
+ and ty: "\<Phi> C' M' ! pc' = Some (ST,LT)" and conf: "conf_f P h sh (ST,LT) ins f'"
+ and confs: "conf_fs P h sh \<Phi> C' M' (size Ts) T frs'"
+ and confc: "conf_clinit P sh frs"
+ and h_ok: "P\<turnstile> h\<surd>" and sh_ok: "P,h \<turnstile>\<^sub>s sh \<surd>"
+ by(auto simp: correct_state_def)
+
+ from Cons dist have dist': "distinct (C#clinit_classes frs')"
+ by(auto simp: distinct_length_2_or_more)
+
+ from shconf_upd_obj[OF sh_ok shconfD[OF sh_ok shC]] have sh_ok': "P,h \<turnstile>\<^sub>s ?sh \<surd>"
+ by simp
+
+ from conf f' valid_ics_shupd Cons dist have conf': "conf_f P h ?sh (ST,LT) ins f'"
+ by(auto simp: conf_f_def2 fun_upd_apply)
+ have confs': "conf_fs P h ?sh \<Phi> C' M' (size Ts) T frs'" by(rule conf_fs_shupd[OF confs dist'])
+
+ have confc': "conf_clinit P ?sh frs" by(rule conf_clinit_shupd[OF confc dist])
+
+ with h_ok sh_ok' meth ty conf' confs' f' Cons None show ?thesis
+ by(fastforce simp: correct_state_def)
+ qed(simp add: correct_state_def)
+qed(simp add: correct_state_def)
+
+lemma correct_state_Throwing_ex:
+assumes correct: "P,\<Phi> \<turnstile> (xp,h,(stk,loc,C,M,pc,ics)#frs,sh)\<surd>"
+shows "\<And>Cs a. ics = Throwing Cs a \<Longrightarrow> \<exists>obj. h a = Some obj"
+using correct by(clarsimp simp: correct_state_def conf_f_def)
+
+end
diff --git a/thys/JinjaDCI/BV/BVExec.thy b/thys/JinjaDCI/BV/BVExec.thy
--- a/thys/JinjaDCI/BV/BVExec.thy
+++ b/thys/JinjaDCI/BV/BVExec.thy
@@ -1,506 +1,506 @@
-(* Title: JinjaDCI/BV/BVExec.thy
-
- Author: Tobias Nipkow, Gerwin Klein, Susannah Mansky
- Copyright 2000 TUM, 2020 UIUC
-
- Based on the Jinja theory BV/BVExec.thy by Tobias Nipkow and Gerwin Klein
-*)
-
-section \<open> Kildall for the JVM \label{sec:JVM} \<close>
-
-theory BVExec
-imports Jinja.Abstract_BV TF_JVM Jinja.Typing_Framework_2
-begin
-
-definition kiljvm :: "jvm_prog \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> ty \<Rightarrow>
- instr list \<Rightarrow> ex_table \<Rightarrow> ty\<^sub>i' err list \<Rightarrow> ty\<^sub>i' err list"
-where
- "kiljvm P mxs mxl T\<^sub>r is xt \<equiv>
- kildall (JVM_SemiType.le P mxs mxl) (JVM_SemiType.sup P mxs mxl)
- (exec P mxs T\<^sub>r xt is)"
-
-definition wt_kildall :: "jvm_prog \<Rightarrow> cname \<Rightarrow> staticb \<Rightarrow> ty list \<Rightarrow> ty \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow>
- instr list \<Rightarrow> ex_table \<Rightarrow> bool"
-where
- "wt_kildall P C' b Ts T\<^sub>r mxs mxl\<^sub>0 is xt \<equiv> (b = Static \<or> b = NonStatic) \<and>
- 0 < size is \<and>
- (let first = Some ([],(case b of Static \<Rightarrow> [] | NonStatic \<Rightarrow> [OK (Class C')])
- @(map OK Ts)@(replicate mxl\<^sub>0 Err));
- start = (OK first)#(replicate (size is - 1) (OK None));
- result = kiljvm P mxs
- ((case b of Static \<Rightarrow> 0 | NonStatic \<Rightarrow> 1)+size Ts+mxl\<^sub>0)
- T\<^sub>r is xt start
- in \<forall>n < size is. result!n \<noteq> Err)"
-
-definition wf_jvm_prog\<^sub>k :: "jvm_prog \<Rightarrow> bool"
-where
- "wf_jvm_prog\<^sub>k P \<equiv>
- wf_prog (\<lambda>P C' (M,b,Ts,T\<^sub>r,(mxs,mxl\<^sub>0,is,xt)). wt_kildall P C' b Ts T\<^sub>r mxs mxl\<^sub>0 is xt) P"
-
-
-context start_context
-begin
-
-lemma Cons_less_Conss3 [simp]:
- "x#xs [\<sqsubset>\<^bsub>r\<^esub>] y#ys = (x \<sqsubset>\<^bsub>r\<^esub> y \<and> xs [\<sqsubseteq>\<^bsub>r\<^esub>] ys \<or> x = y \<and> xs [\<sqsubset>\<^bsub>r\<^esub>] ys)"
- apply (unfold lesssub_def )
- apply auto
- apply (insert sup_state_opt_err)
- apply (unfold lesssub_def lesub_def sup_state_opt_def sup_state_def sup_ty_opt_def)
- apply (simp only: JVM_le_unfold )
- apply fastforce
- done
-
-lemma acc_le_listI3 [intro!]:
- " acc r \<Longrightarrow> acc (Listn.le r)"
-apply (unfold acc_def)
-apply (subgoal_tac
- "wf(UN n. {(ys,xs). size xs = n \<and> size ys = n \<and> xs <_(Listn.le r) ys})")
- apply (erule wf_subset)
- apply (blast intro: lesssub_lengthD)
-apply (rule wf_UN)
- prefer 2
- apply (rename_tac m n)
- apply (case_tac "m=n")
- apply simp
- apply (fast intro!: equals0I dest: not_sym)
-apply (rename_tac n)
-apply (induct_tac n)
- apply (simp add: lesssub_def cong: conj_cong)
-apply (rename_tac k)
-apply (simp add: wf_eq_minimal del: r_def f_def step_def A_def)
-apply (simp (no_asm) add: length_Suc_conv cong: conj_cong del: r_def f_def step_def A_def)
-apply clarify
-apply (rename_tac M m)
-apply (case_tac "\<exists>x xs. size xs = k \<and> x#xs \<in> M")
- prefer 2
- apply (erule thin_rl)
- apply (erule thin_rl)
- apply blast
-apply (erule_tac x = "{a. \<exists>xs. size xs = k \<and> a#xs:M}" in allE)
-apply (erule impE)
- apply blast
-apply (thin_tac "\<exists>x xs. P x xs" for P)
-apply clarify
-apply (rename_tac maxA xs)
-apply (erule_tac x = "{ys. size ys = size xs \<and> maxA#ys \<in> M}" in allE)
-apply (erule impE)
- apply blast
-apply clarify
-apply (thin_tac "m \<in> M")
-apply (thin_tac "maxA#xs \<in> M")
-apply (rule bexI)
- prefer 2
- apply assumption
-apply clarify
-apply (simp del: r_def f_def step_def A_def)
-apply blast
- done
-
-
-lemma wf_jvm: " wf {(ss', ss). ss [\<sqsubset>\<^bsub>r\<^esub>] ss'}"
- apply (insert acc_le_listI3 acc_JVM [OF wf])
- by (simp add: acc_def r_def)
-
-lemma iter_properties_bv[rule_format]:
-shows "\<lbrakk> \<forall>p\<in>w0. p < n; ss0 \<in> list n A; \<forall>p<n. p \<notin> w0 \<longrightarrow> stable r step ss0 p \<rbrakk> \<Longrightarrow>
- iter f step ss0 w0 = (ss',w') \<longrightarrow>
- ss' \<in> list n A \<and> stables r step ss' \<and> ss0 [\<sqsubseteq>\<^sub>r] ss' \<and>
- (\<forall>ts\<in>list n A. ss0 [\<sqsubseteq>\<^sub>r] ts \<and> stables r step ts \<longrightarrow> ss' [\<sqsubseteq>\<^sub>r] ts)"
-(*<*) (is "PROP ?P")
-
-proof -
- show "PROP ?P"
- apply (insert semi bounded_step exec_pres_type step_mono[OF wf])
- apply (unfold iter_def stables_def)
-
- apply (rule_tac P = "\<lambda>(ss,w).
- ss \<in> list n A \<and> (\<forall>p<n. p \<notin> w \<longrightarrow> stable r step ss p) \<and> ss0 [\<sqsubseteq>\<^sub>r] ss \<and>
- (\<forall>ts\<in>list n A. ss0 [\<sqsubseteq>\<^sub>r] ts \<and> stables r step ts \<longrightarrow> ss [\<sqsubseteq>\<^sub>r] ts) \<and>
- (\<forall>p\<in>w. p < n)" and
- r = "{(ss',ss) . ss [\<sqsubset>\<^sub>r] ss'} <*lex*> finite_psubset"
- in while_rule)
-
- \<comment> \<open>Invariant holds initially:\<close>
- apply (simp add:stables_def semilat_Def del: n_def A_def r_def f_def step_def)
- apply (blast intro:le_list_refl')
-
- \<comment> \<open>Invariant is preserved:\<close>
- apply(simp add: stables_def split_paired_all del: A_def r_def f_def step_def n_def)
- apply(rename_tac ss w)
- apply(subgoal_tac "(SOME p. p \<in> w) \<in> w")
- prefer 2 apply (fast intro: someI)
- apply(subgoal_tac "\<forall>(q,t) \<in> set (step (SOME p. p \<in> w) (ss ! (SOME p. p \<in> w))). q < length ss \<and> t \<in> A")
- prefer 2
- apply clarify
- apply (rule conjI)
- apply(clarsimp, blast dest!: boundedD)
- apply (subgoal_tac "(SOME p. p \<in> w) < n")
- apply (subgoal_tac "(ss ! (SOME p. p \<in> w)) \<in> A")
- apply (fastforce simp only:n_def dest:pres_typeD )
- apply simp
- apply simp
- apply (subst decomp_propa)
- apply blast
- apply (simp del:A_def r_def f_def step_def n_def)
- apply (rule conjI)
- apply (rule Semilat.merges_preserves_type[OF Semilat.intro, OF semi])
- apply blast
- apply clarify
- apply (rule conjI)
- apply(clarsimp, blast dest!: boundedD)
- apply (erule pres_typeD)
- prefer 3
- apply assumption
- apply (erule listE_nth_in)
- apply blast
- apply (simp only:n_def)
- apply (rule conjI)
- apply clarify
- apply (subgoal_tac "ss \<in> list (length is) A" "\<forall>p\<in>w. p < (length is) " "\<forall>p<(length is). p \<notin> w \<longrightarrow> stable r step ss p "
- "p < length is")
- apply (blast intro!: Semilat.stable_pres_lemma[OF Semilat.intro, OF semi])
- apply (simp only:n_def)
- apply (simp only:n_def)
- apply (simp only:n_def)
- apply (simp only:n_def)
- apply (rule conjI)
- apply (subgoal_tac "ss \<in> list (length is) A"
- "\<forall>(q,t)\<in>set (step (SOME p. p \<in> w) (ss ! (SOME p. p \<in> w))). q < length is \<and> t \<in> A"
- "ss [\<sqsubseteq>\<^bsub>r\<^esub>] merges f (step (SOME p. p \<in> w) (ss ! (SOME p. p \<in> w))) ss" "ss0\<in> list (size is) A"
- "merges f (step (SOME p. p \<in> w) (ss ! (SOME p. p \<in> w))) ss \<in> list (size is) A"
- "ss \<in>list (size is) A" "order r A" "ss0 [\<sqsubseteq>\<^bsub>r\<^esub>] ss ")
- apply (blast dest: le_list_trans)
- apply simp
- apply (simp only:semilat_Def)
- apply (simp only:n_def)
- apply (fastforce simp only: n_def dest:Semilat.merges_preserves_type[OF Semilat.intro, OF semi] )
- apply (simp only:n_def)
- apply (blast intro:Semilat.merges_incr[OF Semilat.intro, OF semi])
- apply (subgoal_tac "length ss = n")
- apply (simp only:n_def)
- apply (subgoal_tac "ss \<in>list n A")
- defer
- apply simp
- apply (simp only:n_def)
- prefer 5
- apply (simp only:listE_length n_def)
- apply(rule conjI)
- apply (clarsimp simp del: A_def r_def f_def step_def)
- apply (blast intro!: Semilat.merges_bounded_lemma[OF Semilat.intro, OF semi])
- apply (subgoal_tac "bounded step n")
- apply (blast dest!: boundedD)
- apply (simp only:n_def)
-
- \<comment> \<open>Postcondition holds upon termination:\<close>
- apply(clarsimp simp add: stables_def split_paired_all)
-
- \<comment> \<open>Well-foundedness of the termination relation:\<close>
- apply (rule wf_lex_prod)
- apply (simp only:wf_jvm)
- apply (rule wf_finite_psubset)
-
- \<comment> \<open>Loop decreases along termination relation:\<close>
- apply(simp add: stables_def split_paired_all del: A_def r_def f_def step_def)
- apply(rename_tac ss w)
- apply(subgoal_tac "(SOME p. p \<in> w) \<in> w")
- prefer 2 apply (fast intro: someI)
- apply(subgoal_tac "\<forall>(q,t) \<in> set (step (SOME p. p \<in> w) (ss ! (SOME p. p \<in> w))). q < length ss \<and> t \<in> A")
- prefer 2
- apply clarify
- apply (rule conjI)
- apply(clarsimp, blast dest!: boundedD)
- apply (erule pres_typeD)
- prefer 3
- apply assumption
- apply (erule listE_nth_in)
- apply blast
- apply blast
- apply (subst decomp_propa)
- apply blast
- apply clarify
- apply (simp del: listE_length A_def r_def f_def step_def
- add: lex_prod_def finite_psubset_def
- bounded_nat_set_is_finite)
- apply (rule termination_lemma)
- apply (insert Semilat.intro)
- apply assumption+
- defer
- apply assumption
- defer
- apply clarsimp
- done
-qed
-
-(*>*)
-
-
-lemma kildall_properties_bv:
-shows "\<lbrakk> ss0 \<in> list n A \<rbrakk> \<Longrightarrow>
- kildall r f step ss0 \<in> list n A \<and>
- stables r step (kildall r f step ss0) \<and>
- ss0 [\<sqsubseteq>\<^sub>r] kildall r f step ss0 \<and>
- (\<forall>ts\<in>list n A. ss0 [\<sqsubseteq>\<^sub>r] ts \<and> stables r step ts \<longrightarrow>
- kildall r f step ss0 [\<sqsubseteq>\<^sub>r] ts)"
-(*<*) (is "PROP ?P")
-proof -
- show "PROP ?P"
- apply (unfold kildall_def)
- apply(case_tac "iter f step ss0 (unstables r step ss0)")
- apply (simp del:r_def f_def n_def step_def A_def)
- apply (rule iter_properties_bv)
- apply (simp_all add: unstables_def stable_def)
- done
-qed
-
-end
-
-theorem (in start_context) is_bcv_kiljvm:
- "is_bcv r Err step (size is) A (kiljvm P mxs mxl T\<^sub>r is xt)"
-(*<*)
- apply (insert wf)
- apply (unfold kiljvm_def)
- apply (fold r_def f_def step_def_exec n_def)
- apply(unfold is_bcv_def wt_step_def)
- apply(insert semi kildall_properties_bv)
- apply(simp only:stables_def)
- apply clarify
- apply(subgoal_tac "kildall r f step \<tau>s\<^sub>0 \<in> list n A")
- prefer 2
- apply (fastforce intro: kildall_properties_bv)
- apply (rule iffI)
- apply (rule_tac x = "kildall r f step \<tau>s\<^sub>0" in bexI)
- apply (rule conjI)
- apply (fastforce intro: kildall_properties_bv)
- apply (force intro: kildall_properties_bv)
- apply simp
- apply clarify
- apply(subgoal_tac "kildall r f step \<tau>s\<^sub>0!pa <=_r \<tau>s!pa")
- defer
- apply (blast intro!: le_listD less_lengthI)
- apply (subgoal_tac "\<tau>s!pa \<noteq> Err")
- defer
- apply simp
- apply (rule ccontr)
- apply (simp only:top_def r_def JVM_le_unfold)
- apply fastforce
- done
-(*
-proof -
- let ?n = "length is"
- have "Semilat A r f" using semilat_JVM[OF wf]
- by (simp add: Semilat.intro sl_def2)
- moreover have "acc r" using wf by simp blast
- moreover have "top r Err" by (simp add: JVM_le_unfold)
- moreover have "pres_type step ?n A" by (rule exec_pres_type)
- moreover have "bounded step ?n" by (rule bounded_step)
- moreover have "mono r step ?n A" using step_mono[OF wf] by simp
- ultimately have "is_bcv r Err step ?n A (kildall r f step)"
- by(rule is_bcv_kildall)
- moreover have kileq: "kiljvm P mxs mxl T\<^sub>r is xt = kildall r f step"
- using f_def kiljvm_def r_def step_def_exec by blast
- ultimately show ?thesis by simp
-qed
-*)
-(*>*)
-
-(* FIXME: move? *)
-lemma subset_replicate [intro?]: "set (replicate n x) \<subseteq> {x}"
- by (induct n) auto
-
-lemma in_set_replicate:
- assumes "x \<in> set (replicate n y)"
- shows "x = y"
-(*<*)
-proof -
- note assms
- also have "set (replicate n y) \<subseteq> {y}" ..
- finally show ?thesis by simp
-qed
-(*>*)
-
-lemma (in start_context) start_in_A [intro?]:
- "0 < size is \<Longrightarrow> start \<in> list (size is) A"
- using Ts C
-(*<*)
-proof(cases b)
-qed (force simp: JVM_states_unfold intro!: listI list_appendI
- dest!: in_set_replicate)+
-(*>*)
-
-
-theorem (in start_context) wt_kil_correct:
- assumes wtk: "wt_kildall P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt"
- shows "\<exists>\<tau>s. wt_method P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt \<tau>s"
-(*<*)
-proof -
- from wtk obtain res where
- result: "res = kiljvm P mxs mxl T\<^sub>r is xt start" and
- success: "\<forall>n < size is. res!n \<noteq> Err" and
- instrs: "0 < size is" and
- stab: "b = Static \<or> b = NonStatic"
- by (unfold wt_kildall_def) simp
-
- have bcv: "is_bcv r Err step (size is) A (kiljvm P mxs mxl T\<^sub>r is xt)"
- by (rule is_bcv_kiljvm)
-
- from instrs have "start \<in> list (size is) A" ..
- with bcv success result have
- "\<exists>ts\<in>list (size is) A. start [\<sqsubseteq>\<^sub>r] ts \<and> wt_step r Err step ts"
- by (unfold is_bcv_def) blast
- then obtain \<tau>s' where
- in_A: "\<tau>s' \<in> list (size is) A" and
- s: "start [\<sqsubseteq>\<^sub>r] \<tau>s'" and
- w: "wt_step r Err step \<tau>s'"
- by blast
- hence wt_err_step: "wt_err_step (sup_state_opt P) step \<tau>s'"
- by (simp add: wt_err_step_def JVM_le_Err_conv)
-
- from in_A have l: "size \<tau>s' = size is" by simp
- moreover {
- from in_A have "check_types P mxs mxl \<tau>s'" by (simp add: check_types_def)
- also from w have "\<forall>x \<in> set \<tau>s'. x \<noteq> Err"
- by (auto simp add: wt_step_def all_set_conv_all_nth)
- hence [symmetric]: "map OK (map ok_val \<tau>s') = \<tau>s'"
- by (auto intro!: map_idI simp add: wt_step_def)
- finally have "check_types P mxs mxl (map OK (map ok_val \<tau>s'))" .
- }
- moreover {
- from s have "start!0 \<sqsubseteq>\<^sub>r \<tau>s'!0" by (rule le_listD) simp
- moreover
- from instrs w l
- have "\<tau>s'!0 \<noteq> Err" by (unfold wt_step_def) simp
- then obtain \<tau>s0 where "\<tau>s'!0 = OK \<tau>s0" by auto
- ultimately
- have "wt_start P C b Ts mxl\<^sub>0 (map ok_val \<tau>s')" using l instrs
- by (unfold wt_start_def)
- (cases b; simp add: lesub_def JVM_le_Err_conv Err.le_def)
- }
- moreover
- from in_A have "set \<tau>s' \<subseteq> A" by simp
- with wt_err_step bounded_step
- have "wt_app_eff (sup_state_opt P) app eff (map ok_val \<tau>s')"
- by (auto intro: wt_err_imp_wt_app_eff simp add: l)
- ultimately
- have "wt_method P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt (map ok_val \<tau>s')"
- using instrs stab by (simp add: wt_method_def2 check_types_def del: map_map)
- thus ?thesis by blast
-qed
-(*>*)
-
-
-theorem (in start_context) wt_kil_complete:
- assumes wtm: "wt_method P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt \<tau>s"
- shows "wt_kildall P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt"
-(*<*)
-proof -
- from wtm obtain
- instrs: "0 < size is" and
- length: "length \<tau>s = length is" and
- ck_type: "check_types P mxs mxl (map OK \<tau>s)" and
- wt_start: "wt_start P C b Ts mxl\<^sub>0 \<tau>s" and
- app_eff: "wt_app_eff (sup_state_opt P) app eff \<tau>s" and
- stab: "b = Static \<or> b = NonStatic"
- by (simp add: wt_method_def2 check_types_def )
-
- from ck_type
- have in_A: "set (map OK \<tau>s) \<subseteq> A"
- by (simp add: check_types_def)
- with app_eff in_A bounded_step
- have "wt_err_step (sup_state_opt P) (err_step (size \<tau>s) app eff) (map OK \<tau>s)"
- by - (erule wt_app_eff_imp_wt_err,
- auto simp add: exec_def length states_def)
- hence wt_err: "wt_err_step (sup_state_opt P) step (map OK \<tau>s)"
- by (simp add: length)
- have is_bcv: "is_bcv r Err step (size is) A (kiljvm P mxs mxl T\<^sub>r is xt)"
- by (rule is_bcv_kiljvm)
- moreover from instrs have "start \<in> list (size is) A" ..
- moreover
- let ?\<tau>s = "map OK \<tau>s"
- have less_\<tau>s: "start [\<sqsubseteq>\<^sub>r] ?\<tau>s"
- proof (rule le_listI)
- from length instrs
- show "length start = length (map OK \<tau>s)" by simp
- next
- fix n
- from wt_start have "P \<turnstile> ok_val (start!0) \<le>' \<tau>s!0"
- by (cases b; simp add: wt_start_def)
- moreover from instrs length have "0 < length \<tau>s" by simp
- ultimately have "start!0 \<sqsubseteq>\<^sub>r ?\<tau>s!0"
- by (simp add: JVM_le_Err_conv lesub_def)
- moreover {
- fix n'
- have "OK None \<sqsubseteq>\<^sub>r ?\<tau>s!n"
- by (auto simp add: JVM_le_Err_conv Err.le_def lesub_def
- split: err.splits)
- hence "\<lbrakk>n = Suc n'; n < size start\<rbrakk> \<Longrightarrow> start!n \<sqsubseteq>\<^sub>r ?\<tau>s!n" by simp
- }
- ultimately
- show "n < size start \<Longrightarrow> start!n \<sqsubseteq>\<^sub>r ?\<tau>s!n" by (cases n, blast+)
- qed
- moreover
- from ck_type length
- have "?\<tau>s \<in> list (size is) A"
- by (auto intro!: listI simp add: check_types_def)
- moreover
- from wt_err have "wt_step r Err step ?\<tau>s"
- by (simp add: wt_err_step_def JVM_le_Err_conv)
- ultimately
- have "\<forall>p. p < size is \<longrightarrow> kiljvm P mxs mxl T\<^sub>r is xt start ! p \<noteq> Err"
- by (unfold is_bcv_def) blast
- with instrs
- show "wt_kildall P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt"
- using start_context_intro_auxi[OF staticb] using stab by (unfold wt_kildall_def) simp
-qed
-(*>*)
-
-
-theorem jvm_kildall_correct:
- "wf_jvm_prog\<^sub>k P = wf_jvm_prog P"
-(*<*)
-proof -
- let ?\<Phi> = "\<lambda>C M. let (C,b,Ts,T\<^sub>r,(mxs,mxl\<^sub>0,is,xt)) = method P C M in
- SOME \<tau>s. wt_method P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt \<tau>s"
- let ?A = "\<lambda>P C' (M, b, Ts, T\<^sub>r, mxs, mxl\<^sub>0, is, xt). wt_kildall P C' b Ts T\<^sub>r mxs mxl\<^sub>0 is xt"
- let ?B\<^sub>\<Phi> = "\<lambda>\<Phi>. (\<lambda>P C (M,b,Ts,T\<^sub>r,(mxs,mxl\<^sub>0,is,xt)).
- wt_method P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt (\<Phi> C M))"
-
- show ?thesis proof(rule iffI)
- \<comment> \<open>soundness\<close>
- assume wt: "wf_jvm_prog\<^sub>k P"
- then have wt': "wf_prog ?A P" by(simp add: wf_jvm_prog\<^sub>k_def)
- moreover {
- fix wf_md C M b Ts Ca T m bd
-
- assume ass1: "wf_prog wf_md P" and sees: "P \<turnstile> Ca sees M, b : Ts\<rightarrow>T = m in Ca" and
- ass2: "set Ts \<subseteq> types P" and ass3: "bd = (M, b, Ts, T, m)" and
- ass4: "?A P Ca bd"
- from ass3 ass4 have stab: "b = Static \<or> b = NonStatic" by (simp add:wt_kildall_def)
- from ass1 sees ass2 ass3 ass4 have "(?B\<^sub>\<Phi> ?\<Phi>) P Ca bd" using sees_method_is_class[OF sees]
- by (auto dest!: start_context.wt_kil_correct[OF start_context_intro_auxi[OF stab]]
- intro: someI)
- }
- ultimately have "wf_prog (?B\<^sub>\<Phi> ?\<Phi>) P" by(rule wf_prog_lift)
- then have "wf_jvm_prog\<^bsub>?\<Phi>\<^esub> P" by (simp add: wf_jvm_prog_phi_def)
- thus "wf_jvm_prog P" by (unfold wf_jvm_prog_def) fast
- next
- \<comment> \<open>completeness\<close>
-
- assume wt: "wf_jvm_prog P"
- then obtain \<Phi> where "wf_prog (?B\<^sub>\<Phi> \<Phi>) P" by(clarsimp simp: wf_jvm_prog_def wf_jvm_prog_phi_def)
- moreover {
- fix wf_md C M b Ts Ca T m bd
- assume ass1: "wf_prog wf_md P" and sees: "P \<turnstile> Ca sees M, b : Ts\<rightarrow>T = m in Ca" and
- ass2: "set Ts \<subseteq> types P" and ass3: "bd = (M, b, Ts, T, m)" and
- ass4: "(?B\<^sub>\<Phi> \<Phi>) P Ca bd"
- from ass3 ass4 have stab: "b = Static \<or> b = NonStatic" by (simp add:wt_method_def)
- from ass1 sees ass2 ass3 ass4 have "?A P Ca bd" using sees_method_is_class[OF sees] using JVM_sl.staticb
- by (auto intro!: start_context.wt_kil_complete start_context_intro_auxi[OF stab])
- }
- ultimately have "wf_prog ?A P" by(rule wf_prog_lift)
- thus "wf_jvm_prog\<^sub>k P" by (simp add: wf_jvm_prog\<^sub>k_def)
- qed
-qed
-(*>*)
-
-end
+(* Title: JinjaDCI/BV/BVExec.thy
+
+ Author: Tobias Nipkow, Gerwin Klein, Susannah Mansky
+ Copyright 2000 TUM, 2020 UIUC
+
+ Based on the Jinja theory BV/BVExec.thy by Tobias Nipkow and Gerwin Klein
+*)
+
+section \<open> Kildall for the JVM \label{sec:JVM} \<close>
+
+theory BVExec
+imports Jinja.Abstract_BV TF_JVM Jinja.Typing_Framework_2
+begin
+
+definition kiljvm :: "jvm_prog \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> ty \<Rightarrow>
+ instr list \<Rightarrow> ex_table \<Rightarrow> ty\<^sub>i' err list \<Rightarrow> ty\<^sub>i' err list"
+where
+ "kiljvm P mxs mxl T\<^sub>r is xt \<equiv>
+ kildall (JVM_SemiType.le P mxs mxl) (JVM_SemiType.sup P mxs mxl)
+ (exec P mxs T\<^sub>r xt is)"
+
+definition wt_kildall :: "jvm_prog \<Rightarrow> cname \<Rightarrow> staticb \<Rightarrow> ty list \<Rightarrow> ty \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow>
+ instr list \<Rightarrow> ex_table \<Rightarrow> bool"
+where
+ "wt_kildall P C' b Ts T\<^sub>r mxs mxl\<^sub>0 is xt \<equiv> (b = Static \<or> b = NonStatic) \<and>
+ 0 < size is \<and>
+ (let first = Some ([],(case b of Static \<Rightarrow> [] | NonStatic \<Rightarrow> [OK (Class C')])
+ @(map OK Ts)@(replicate mxl\<^sub>0 Err));
+ start = (OK first)#(replicate (size is - 1) (OK None));
+ result = kiljvm P mxs
+ ((case b of Static \<Rightarrow> 0 | NonStatic \<Rightarrow> 1)+size Ts+mxl\<^sub>0)
+ T\<^sub>r is xt start
+ in \<forall>n < size is. result!n \<noteq> Err)"
+
+definition wf_jvm_prog\<^sub>k :: "jvm_prog \<Rightarrow> bool"
+where
+ "wf_jvm_prog\<^sub>k P \<equiv>
+ wf_prog (\<lambda>P C' (M,b,Ts,T\<^sub>r,(mxs,mxl\<^sub>0,is,xt)). wt_kildall P C' b Ts T\<^sub>r mxs mxl\<^sub>0 is xt) P"
+
+
+context start_context
+begin
+
+lemma Cons_less_Conss3 [simp]:
+ "x#xs [\<sqsubset>\<^bsub>r\<^esub>] y#ys = (x \<sqsubset>\<^bsub>r\<^esub> y \<and> xs [\<sqsubseteq>\<^bsub>r\<^esub>] ys \<or> x = y \<and> xs [\<sqsubset>\<^bsub>r\<^esub>] ys)"
+ apply (unfold lesssub_def )
+ apply auto
+ apply (insert sup_state_opt_err)
+ apply (unfold lesssub_def lesub_def sup_state_opt_def sup_state_def sup_ty_opt_def)
+ apply (simp only: JVM_le_unfold )
+ apply fastforce
+ done
+
+lemma acc_le_listI3 [intro!]:
+ " acc r \<Longrightarrow> acc (Listn.le r)"
+apply (unfold acc_def)
+apply (subgoal_tac
+ "wf(UN n. {(ys,xs). size xs = n \<and> size ys = n \<and> xs <_(Listn.le r) ys})")
+ apply (erule wf_subset)
+ apply (blast intro: lesssub_lengthD)
+apply (rule wf_UN)
+ prefer 2
+ apply (rename_tac m n)
+ apply (case_tac "m=n")
+ apply simp
+ apply (fast intro!: equals0I dest: not_sym)
+apply (rename_tac n)
+apply (induct_tac n)
+ apply (simp add: lesssub_def cong: conj_cong)
+apply (rename_tac k)
+apply (simp add: wf_eq_minimal del: r_def f_def step_def A_def)
+apply (simp (no_asm) add: length_Suc_conv cong: conj_cong del: r_def f_def step_def A_def)
+apply clarify
+apply (rename_tac M m)
+apply (case_tac "\<exists>x xs. size xs = k \<and> x#xs \<in> M")
+ prefer 2
+ apply (erule thin_rl)
+ apply (erule thin_rl)
+ apply blast
+apply (erule_tac x = "{a. \<exists>xs. size xs = k \<and> a#xs:M}" in allE)
+apply (erule impE)
+ apply blast
+apply (thin_tac "\<exists>x xs. P x xs" for P)
+apply clarify
+apply (rename_tac maxA xs)
+apply (erule_tac x = "{ys. size ys = size xs \<and> maxA#ys \<in> M}" in allE)
+apply (erule impE)
+ apply blast
+apply clarify
+apply (thin_tac "m \<in> M")
+apply (thin_tac "maxA#xs \<in> M")
+apply (rule bexI)
+ prefer 2
+ apply assumption
+apply clarify
+apply (simp del: r_def f_def step_def A_def)
+apply blast
+ done
+
+
+lemma wf_jvm: " wf {(ss', ss). ss [\<sqsubset>\<^bsub>r\<^esub>] ss'}"
+ apply (insert acc_le_listI3 acc_JVM [OF wf])
+ by (simp add: acc_def r_def)
+
+lemma iter_properties_bv[rule_format]:
+shows "\<lbrakk> \<forall>p\<in>w0. p < n; ss0 \<in> list n A; \<forall>p<n. p \<notin> w0 \<longrightarrow> stable r step ss0 p \<rbrakk> \<Longrightarrow>
+ iter f step ss0 w0 = (ss',w') \<longrightarrow>
+ ss' \<in> list n A \<and> stables r step ss' \<and> ss0 [\<sqsubseteq>\<^sub>r] ss' \<and>
+ (\<forall>ts\<in>list n A. ss0 [\<sqsubseteq>\<^sub>r] ts \<and> stables r step ts \<longrightarrow> ss' [\<sqsubseteq>\<^sub>r] ts)"
+(*<*) (is "PROP ?P")
+
+proof -
+ show "PROP ?P"
+ apply (insert semi bounded_step exec_pres_type step_mono[OF wf])
+ apply (unfold iter_def stables_def)
+
+ apply (rule_tac P = "\<lambda>(ss,w).
+ ss \<in> list n A \<and> (\<forall>p<n. p \<notin> w \<longrightarrow> stable r step ss p) \<and> ss0 [\<sqsubseteq>\<^sub>r] ss \<and>
+ (\<forall>ts\<in>list n A. ss0 [\<sqsubseteq>\<^sub>r] ts \<and> stables r step ts \<longrightarrow> ss [\<sqsubseteq>\<^sub>r] ts) \<and>
+ (\<forall>p\<in>w. p < n)" and
+ r = "{(ss',ss) . ss [\<sqsubset>\<^sub>r] ss'} <*lex*> finite_psubset"
+ in while_rule)
+
+ \<comment> \<open>Invariant holds initially:\<close>
+ apply (simp add:stables_def semilat_Def del: n_def A_def r_def f_def step_def)
+ apply (blast intro:le_list_refl')
+
+ \<comment> \<open>Invariant is preserved:\<close>
+ apply(simp add: stables_def split_paired_all del: A_def r_def f_def step_def n_def)
+ apply(rename_tac ss w)
+ apply(subgoal_tac "(SOME p. p \<in> w) \<in> w")
+ prefer 2 apply (fast intro: someI)
+ apply(subgoal_tac "\<forall>(q,t) \<in> set (step (SOME p. p \<in> w) (ss ! (SOME p. p \<in> w))). q < length ss \<and> t \<in> A")
+ prefer 2
+ apply clarify
+ apply (rule conjI)
+ apply(clarsimp, blast dest!: boundedD)
+ apply (subgoal_tac "(SOME p. p \<in> w) < n")
+ apply (subgoal_tac "(ss ! (SOME p. p \<in> w)) \<in> A")
+ apply (fastforce simp only:n_def dest:pres_typeD )
+ apply simp
+ apply simp
+ apply (subst decomp_propa)
+ apply blast
+ apply (simp del:A_def r_def f_def step_def n_def)
+ apply (rule conjI)
+ apply (rule Semilat.merges_preserves_type[OF Semilat.intro, OF semi])
+ apply blast
+ apply clarify
+ apply (rule conjI)
+ apply(clarsimp, blast dest!: boundedD)
+ apply (erule pres_typeD)
+ prefer 3
+ apply assumption
+ apply (erule listE_nth_in)
+ apply blast
+ apply (simp only:n_def)
+ apply (rule conjI)
+ apply clarify
+ apply (subgoal_tac "ss \<in> list (length is) A" "\<forall>p\<in>w. p < (length is) " "\<forall>p<(length is). p \<notin> w \<longrightarrow> stable r step ss p "
+ "p < length is")
+ apply (blast intro!: Semilat.stable_pres_lemma[OF Semilat.intro, OF semi])
+ apply (simp only:n_def)
+ apply (simp only:n_def)
+ apply (simp only:n_def)
+ apply (simp only:n_def)
+ apply (rule conjI)
+ apply (subgoal_tac "ss \<in> list (length is) A"
+ "\<forall>(q,t)\<in>set (step (SOME p. p \<in> w) (ss ! (SOME p. p \<in> w))). q < length is \<and> t \<in> A"
+ "ss [\<sqsubseteq>\<^bsub>r\<^esub>] merges f (step (SOME p. p \<in> w) (ss ! (SOME p. p \<in> w))) ss" "ss0\<in> list (size is) A"
+ "merges f (step (SOME p. p \<in> w) (ss ! (SOME p. p \<in> w))) ss \<in> list (size is) A"
+ "ss \<in>list (size is) A" "order r A" "ss0 [\<sqsubseteq>\<^bsub>r\<^esub>] ss ")
+ apply (blast dest: le_list_trans)
+ apply simp
+ apply (simp only:semilat_Def)
+ apply (simp only:n_def)
+ apply (fastforce simp only: n_def dest:Semilat.merges_preserves_type[OF Semilat.intro, OF semi] )
+ apply (simp only:n_def)
+ apply (blast intro:Semilat.merges_incr[OF Semilat.intro, OF semi])
+ apply (subgoal_tac "length ss = n")
+ apply (simp only:n_def)
+ apply (subgoal_tac "ss \<in>list n A")
+ defer
+ apply simp
+ apply (simp only:n_def)
+ prefer 5
+ apply (simp only:listE_length n_def)
+ apply(rule conjI)
+ apply (clarsimp simp del: A_def r_def f_def step_def)
+ apply (blast intro!: Semilat.merges_bounded_lemma[OF Semilat.intro, OF semi])
+ apply (subgoal_tac "bounded step n")
+ apply (blast dest!: boundedD)
+ apply (simp only:n_def)
+
+ \<comment> \<open>Postcondition holds upon termination:\<close>
+ apply(clarsimp simp add: stables_def split_paired_all)
+
+ \<comment> \<open>Well-foundedness of the termination relation:\<close>
+ apply (rule wf_lex_prod)
+ apply (simp only:wf_jvm)
+ apply (rule wf_finite_psubset)
+
+ \<comment> \<open>Loop decreases along termination relation:\<close>
+ apply(simp add: stables_def split_paired_all del: A_def r_def f_def step_def)
+ apply(rename_tac ss w)
+ apply(subgoal_tac "(SOME p. p \<in> w) \<in> w")
+ prefer 2 apply (fast intro: someI)
+ apply(subgoal_tac "\<forall>(q,t) \<in> set (step (SOME p. p \<in> w) (ss ! (SOME p. p \<in> w))). q < length ss \<and> t \<in> A")
+ prefer 2
+ apply clarify
+ apply (rule conjI)
+ apply(clarsimp, blast dest!: boundedD)
+ apply (erule pres_typeD)
+ prefer 3
+ apply assumption
+ apply (erule listE_nth_in)
+ apply blast
+ apply blast
+ apply (subst decomp_propa)
+ apply blast
+ apply clarify
+ apply (simp del: listE_length A_def r_def f_def step_def
+ add: lex_prod_def finite_psubset_def
+ bounded_nat_set_is_finite)
+ apply (rule termination_lemma)
+ apply (insert Semilat.intro)
+ apply assumption+
+ defer
+ apply assumption
+ defer
+ apply clarsimp
+ done
+qed
+
+(*>*)
+
+
+lemma kildall_properties_bv:
+shows "\<lbrakk> ss0 \<in> list n A \<rbrakk> \<Longrightarrow>
+ kildall r f step ss0 \<in> list n A \<and>
+ stables r step (kildall r f step ss0) \<and>
+ ss0 [\<sqsubseteq>\<^sub>r] kildall r f step ss0 \<and>
+ (\<forall>ts\<in>list n A. ss0 [\<sqsubseteq>\<^sub>r] ts \<and> stables r step ts \<longrightarrow>
+ kildall r f step ss0 [\<sqsubseteq>\<^sub>r] ts)"
+(*<*) (is "PROP ?P")
+proof -
+ show "PROP ?P"
+ apply (unfold kildall_def)
+ apply(case_tac "iter f step ss0 (unstables r step ss0)")
+ apply (simp del:r_def f_def n_def step_def A_def)
+ apply (rule iter_properties_bv)
+ apply (simp_all add: unstables_def stable_def)
+ done
+qed
+
+end
+
+theorem (in start_context) is_bcv_kiljvm:
+ "is_bcv r Err step (size is) A (kiljvm P mxs mxl T\<^sub>r is xt)"
+(*<*)
+ apply (insert wf)
+ apply (unfold kiljvm_def)
+ apply (fold r_def f_def step_def_exec n_def)
+ apply(unfold is_bcv_def wt_step_def)
+ apply(insert semi kildall_properties_bv)
+ apply(simp only:stables_def)
+ apply clarify
+ apply(subgoal_tac "kildall r f step \<tau>s\<^sub>0 \<in> list n A")
+ prefer 2
+ apply (fastforce intro: kildall_properties_bv)
+ apply (rule iffI)
+ apply (rule_tac x = "kildall r f step \<tau>s\<^sub>0" in bexI)
+ apply (rule conjI)
+ apply (fastforce intro: kildall_properties_bv)
+ apply (force intro: kildall_properties_bv)
+ apply simp
+ apply clarify
+ apply(subgoal_tac "kildall r f step \<tau>s\<^sub>0!pa <=_r \<tau>s!pa")
+ defer
+ apply (blast intro!: le_listD less_lengthI)
+ apply (subgoal_tac "\<tau>s!pa \<noteq> Err")
+ defer
+ apply simp
+ apply (rule ccontr)
+ apply (simp only:top_def r_def JVM_le_unfold)
+ apply fastforce
+ done
+(*
+proof -
+ let ?n = "length is"
+ have "Semilat A r f" using semilat_JVM[OF wf]
+ by (simp add: Semilat.intro sl_def2)
+ moreover have "acc r" using wf by simp blast
+ moreover have "top r Err" by (simp add: JVM_le_unfold)
+ moreover have "pres_type step ?n A" by (rule exec_pres_type)
+ moreover have "bounded step ?n" by (rule bounded_step)
+ moreover have "mono r step ?n A" using step_mono[OF wf] by simp
+ ultimately have "is_bcv r Err step ?n A (kildall r f step)"
+ by(rule is_bcv_kildall)
+ moreover have kileq: "kiljvm P mxs mxl T\<^sub>r is xt = kildall r f step"
+ using f_def kiljvm_def r_def step_def_exec by blast
+ ultimately show ?thesis by simp
+qed
+*)
+(*>*)
+
+(* FIXME: move? *)
+lemma subset_replicate [intro?]: "set (replicate n x) \<subseteq> {x}"
+ by (induct n) auto
+
+lemma in_set_replicate:
+ assumes "x \<in> set (replicate n y)"
+ shows "x = y"
+(*<*)
+proof -
+ note assms
+ also have "set (replicate n y) \<subseteq> {y}" ..
+ finally show ?thesis by simp
+qed
+(*>*)
+
+lemma (in start_context) start_in_A [intro?]:
+ "0 < size is \<Longrightarrow> start \<in> list (size is) A"
+ using Ts C
+(*<*)
+proof(cases b)
+qed (force simp: JVM_states_unfold intro!: listI list_appendI
+ dest!: in_set_replicate)+
+(*>*)
+
+
+theorem (in start_context) wt_kil_correct:
+ assumes wtk: "wt_kildall P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt"
+ shows "\<exists>\<tau>s. wt_method P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt \<tau>s"
+(*<*)
+proof -
+ from wtk obtain res where
+ result: "res = kiljvm P mxs mxl T\<^sub>r is xt start" and
+ success: "\<forall>n < size is. res!n \<noteq> Err" and
+ instrs: "0 < size is" and
+ stab: "b = Static \<or> b = NonStatic"
+ by (unfold wt_kildall_def) simp
+
+ have bcv: "is_bcv r Err step (size is) A (kiljvm P mxs mxl T\<^sub>r is xt)"
+ by (rule is_bcv_kiljvm)
+
+ from instrs have "start \<in> list (size is) A" ..
+ with bcv success result have
+ "\<exists>ts\<in>list (size is) A. start [\<sqsubseteq>\<^sub>r] ts \<and> wt_step r Err step ts"
+ by (unfold is_bcv_def) blast
+ then obtain \<tau>s' where
+ in_A: "\<tau>s' \<in> list (size is) A" and
+ s: "start [\<sqsubseteq>\<^sub>r] \<tau>s'" and
+ w: "wt_step r Err step \<tau>s'"
+ by blast
+ hence wt_err_step: "wt_err_step (sup_state_opt P) step \<tau>s'"
+ by (simp add: wt_err_step_def JVM_le_Err_conv)
+
+ from in_A have l: "size \<tau>s' = size is" by simp
+ moreover {
+ from in_A have "check_types P mxs mxl \<tau>s'" by (simp add: check_types_def)
+ also from w have "\<forall>x \<in> set \<tau>s'. x \<noteq> Err"
+ by (auto simp add: wt_step_def all_set_conv_all_nth)
+ hence [symmetric]: "map OK (map ok_val \<tau>s') = \<tau>s'"
+ by (auto intro!: map_idI simp add: wt_step_def)
+ finally have "check_types P mxs mxl (map OK (map ok_val \<tau>s'))" .
+ }
+ moreover {
+ from s have "start!0 \<sqsubseteq>\<^sub>r \<tau>s'!0" by (rule le_listD) simp
+ moreover
+ from instrs w l
+ have "\<tau>s'!0 \<noteq> Err" by (unfold wt_step_def) simp
+ then obtain \<tau>s0 where "\<tau>s'!0 = OK \<tau>s0" by auto
+ ultimately
+ have "wt_start P C b Ts mxl\<^sub>0 (map ok_val \<tau>s')" using l instrs
+ by (unfold wt_start_def)
+ (cases b; simp add: lesub_def JVM_le_Err_conv Err.le_def)
+ }
+ moreover
+ from in_A have "set \<tau>s' \<subseteq> A" by simp
+ with wt_err_step bounded_step
+ have "wt_app_eff (sup_state_opt P) app eff (map ok_val \<tau>s')"
+ by (auto intro: wt_err_imp_wt_app_eff simp add: l)
+ ultimately
+ have "wt_method P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt (map ok_val \<tau>s')"
+ using instrs stab by (simp add: wt_method_def2 check_types_def del: map_map)
+ thus ?thesis by blast
+qed
+(*>*)
+
+
+theorem (in start_context) wt_kil_complete:
+ assumes wtm: "wt_method P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt \<tau>s"
+ shows "wt_kildall P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt"
+(*<*)
+proof -
+ from wtm obtain
+ instrs: "0 < size is" and
+ length: "length \<tau>s = length is" and
+ ck_type: "check_types P mxs mxl (map OK \<tau>s)" and
+ wt_start: "wt_start P C b Ts mxl\<^sub>0 \<tau>s" and
+ app_eff: "wt_app_eff (sup_state_opt P) app eff \<tau>s" and
+ stab: "b = Static \<or> b = NonStatic"
+ by (simp add: wt_method_def2 check_types_def )
+
+ from ck_type
+ have in_A: "set (map OK \<tau>s) \<subseteq> A"
+ by (simp add: check_types_def)
+ with app_eff in_A bounded_step
+ have "wt_err_step (sup_state_opt P) (err_step (size \<tau>s) app eff) (map OK \<tau>s)"
+ by - (erule wt_app_eff_imp_wt_err,
+ auto simp add: exec_def length states_def)
+ hence wt_err: "wt_err_step (sup_state_opt P) step (map OK \<tau>s)"
+ by (simp add: length)
+ have is_bcv: "is_bcv r Err step (size is) A (kiljvm P mxs mxl T\<^sub>r is xt)"
+ by (rule is_bcv_kiljvm)
+ moreover from instrs have "start \<in> list (size is) A" ..
+ moreover
+ let ?\<tau>s = "map OK \<tau>s"
+ have less_\<tau>s: "start [\<sqsubseteq>\<^sub>r] ?\<tau>s"
+ proof (rule le_listI)
+ from length instrs
+ show "length start = length (map OK \<tau>s)" by simp
+ next
+ fix n
+ from wt_start have "P \<turnstile> ok_val (start!0) \<le>' \<tau>s!0"
+ by (cases b; simp add: wt_start_def)
+ moreover from instrs length have "0 < length \<tau>s" by simp
+ ultimately have "start!0 \<sqsubseteq>\<^sub>r ?\<tau>s!0"
+ by (simp add: JVM_le_Err_conv lesub_def)
+ moreover {
+ fix n'
+ have "OK None \<sqsubseteq>\<^sub>r ?\<tau>s!n"
+ by (auto simp add: JVM_le_Err_conv Err.le_def lesub_def
+ split: err.splits)
+ hence "\<lbrakk>n = Suc n'; n < size start\<rbrakk> \<Longrightarrow> start!n \<sqsubseteq>\<^sub>r ?\<tau>s!n" by simp
+ }
+ ultimately
+ show "n < size start \<Longrightarrow> start!n \<sqsubseteq>\<^sub>r ?\<tau>s!n" by (cases n, blast+)
+ qed
+ moreover
+ from ck_type length
+ have "?\<tau>s \<in> list (size is) A"
+ by (auto intro!: listI simp add: check_types_def)
+ moreover
+ from wt_err have "wt_step r Err step ?\<tau>s"
+ by (simp add: wt_err_step_def JVM_le_Err_conv)
+ ultimately
+ have "\<forall>p. p < size is \<longrightarrow> kiljvm P mxs mxl T\<^sub>r is xt start ! p \<noteq> Err"
+ by (unfold is_bcv_def) blast
+ with instrs
+ show "wt_kildall P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt"
+ using start_context_intro_auxi[OF staticb] using stab by (unfold wt_kildall_def) simp
+qed
+(*>*)
+
+
+theorem jvm_kildall_correct:
+ "wf_jvm_prog\<^sub>k P = wf_jvm_prog P"
+(*<*)
+proof -
+ let ?\<Phi> = "\<lambda>C M. let (C,b,Ts,T\<^sub>r,(mxs,mxl\<^sub>0,is,xt)) = method P C M in
+ SOME \<tau>s. wt_method P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt \<tau>s"
+ let ?A = "\<lambda>P C' (M, b, Ts, T\<^sub>r, mxs, mxl\<^sub>0, is, xt). wt_kildall P C' b Ts T\<^sub>r mxs mxl\<^sub>0 is xt"
+ let ?B\<^sub>\<Phi> = "\<lambda>\<Phi>. (\<lambda>P C (M,b,Ts,T\<^sub>r,(mxs,mxl\<^sub>0,is,xt)).
+ wt_method P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt (\<Phi> C M))"
+
+ show ?thesis proof(rule iffI)
+ \<comment> \<open>soundness\<close>
+ assume wt: "wf_jvm_prog\<^sub>k P"
+ then have wt': "wf_prog ?A P" by(simp add: wf_jvm_prog\<^sub>k_def)
+ moreover {
+ fix wf_md C M b Ts Ca T m bd
+
+ assume ass1: "wf_prog wf_md P" and sees: "P \<turnstile> Ca sees M, b : Ts\<rightarrow>T = m in Ca" and
+ ass2: "set Ts \<subseteq> types P" and ass3: "bd = (M, b, Ts, T, m)" and
+ ass4: "?A P Ca bd"
+ from ass3 ass4 have stab: "b = Static \<or> b = NonStatic" by (simp add:wt_kildall_def)
+ from ass1 sees ass2 ass3 ass4 have "(?B\<^sub>\<Phi> ?\<Phi>) P Ca bd" using sees_method_is_class[OF sees]
+ by (auto dest!: start_context.wt_kil_correct[OF start_context_intro_auxi[OF stab]]
+ intro: someI)
+ }
+ ultimately have "wf_prog (?B\<^sub>\<Phi> ?\<Phi>) P" by(rule wf_prog_lift)
+ then have "wf_jvm_prog\<^bsub>?\<Phi>\<^esub> P" by (simp add: wf_jvm_prog_phi_def)
+ thus "wf_jvm_prog P" by (unfold wf_jvm_prog_def) fast
+ next
+ \<comment> \<open>completeness\<close>
+
+ assume wt: "wf_jvm_prog P"
+ then obtain \<Phi> where "wf_prog (?B\<^sub>\<Phi> \<Phi>) P" by(clarsimp simp: wf_jvm_prog_def wf_jvm_prog_phi_def)
+ moreover {
+ fix wf_md C M b Ts Ca T m bd
+ assume ass1: "wf_prog wf_md P" and sees: "P \<turnstile> Ca sees M, b : Ts\<rightarrow>T = m in Ca" and
+ ass2: "set Ts \<subseteq> types P" and ass3: "bd = (M, b, Ts, T, m)" and
+ ass4: "(?B\<^sub>\<Phi> \<Phi>) P Ca bd"
+ from ass3 ass4 have stab: "b = Static \<or> b = NonStatic" by (simp add:wt_method_def)
+ from ass1 sees ass2 ass3 ass4 have "?A P Ca bd" using sees_method_is_class[OF sees] using JVM_sl.staticb
+ by (auto intro!: start_context.wt_kil_complete start_context_intro_auxi[OF stab])
+ }
+ ultimately have "wf_prog ?A P" by(rule wf_prog_lift)
+ thus "wf_jvm_prog\<^sub>k P" by (simp add: wf_jvm_prog\<^sub>k_def)
+ qed
+qed
+(*>*)
+
+end
diff --git a/thys/JinjaDCI/BV/BVNoTypeError.thy b/thys/JinjaDCI/BV/BVNoTypeError.thy
--- a/thys/JinjaDCI/BV/BVNoTypeError.thy
+++ b/thys/JinjaDCI/BV/BVNoTypeError.thy
@@ -1,389 +1,389 @@
-(* Title: JinjaDCI/BV/BVNoTypeErrors.thy
-
- Author: Gerwin Klein, Susannah Mansky
- Copyright GPL
-
- Based on the Jinja theory BV/BVNoTypeErrors.thy by Gerwin Klein
-*)
-
-section \<open> Welltyped Programs produce no Type Errors \<close>
-
-theory BVNoTypeError
-imports "../JVM/JVMDefensive" BVSpecTypeSafe
-begin
-
-lemma has_methodI:
- "P \<turnstile> C sees M,b:Ts\<rightarrow>T = m in D \<Longrightarrow> P \<turnstile> C has M,b"
- by (unfold has_method_def) blast
-
-text \<open>
- Some simple lemmas about the type testing functions of the
- defensive JVM:
-\<close>
-lemma typeof_NoneD [simp,dest]: "typeof v = Some x \<Longrightarrow> \<not>is_Addr v"
- by (cases v) auto
-
-lemma is_Ref_def2:
- "is_Ref v = (v = Null \<or> (\<exists>a. v = Addr a))"
- by (cases v) (auto simp add: is_Ref_def)
-
-lemma [iff]: "is_Ref Null" by (simp add: is_Ref_def2)
-
-lemma is_RefI [intro, simp]: "P,h \<turnstile> v :\<le> T \<Longrightarrow> is_refT T \<Longrightarrow> is_Ref v"
-(*<*)
-proof(cases T)
-qed (auto simp add: is_refT_def is_Ref_def dest: conf_ClassD)
-(*>*)
-
-lemma is_IntgI [intro, simp]: "P,h \<turnstile> v :\<le> Integer \<Longrightarrow> is_Intg v"
-(*<*)by (unfold conf_def) auto(*>*)
-
-lemma is_BoolI [intro, simp]: "P,h \<turnstile> v :\<le> Boolean \<Longrightarrow> is_Bool v"
-(*<*)by (unfold conf_def) auto(*>*)
-
-declare defs1 [simp del]
-
-lemma wt_jvm_prog_states_NonStatic:
-assumes wf: "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P"
- and mC: "P \<turnstile> C sees M,NonStatic: Ts\<rightarrow>T = (mxs, mxl, ins, et) in C"
- and \<Phi>: "\<Phi> C M ! pc = \<tau>" and pc: "pc < size ins"
-shows "OK \<tau> \<in> states P mxs (1+size Ts+mxl)"
-(*<*)
-proof -
- let ?wf_md = "(\<lambda>P C (M, b, Ts, T\<^sub>r, mxs, mxl\<^sub>0, is, xt).
- wt_method P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt (\<Phi> C M))"
- have wfmd: "wf_prog ?wf_md P" using wf
- by (unfold wf_jvm_prog_phi_def) assumption
- show ?thesis using sees_wf_mdecl[OF wfmd mC] \<Phi> pc
- by (simp add: wf_mdecl_def wt_method_def check_types_def)
- (blast intro: nth_in)
-qed
-(*>*)
-
-lemma wt_jvm_prog_states_Static:
-assumes wf: "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P"
- and mC: "P \<turnstile> C sees M,Static: Ts\<rightarrow>T = (mxs, mxl, ins, et) in C"
- and \<Phi>: "\<Phi> C M ! pc = \<tau>" and pc: "pc < size ins"
-shows "OK \<tau> \<in> states P mxs (size Ts+mxl)"
-(*<*)
-proof -
- let ?wf_md = "(\<lambda>P C (M, b, Ts, T\<^sub>r, mxs, mxl\<^sub>0, is, xt).
- wt_method P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt (\<Phi> C M))"
- have wfmd: "wf_prog ?wf_md P" using wf
- by (unfold wf_jvm_prog_phi_def) assumption
- show ?thesis using sees_wf_mdecl[OF wfmd mC] \<Phi> pc
- by (simp add: wf_mdecl_def wt_method_def check_types_def)
- (blast intro: nth_in)
-qed
-(*>*)
-
-text \<open>
- The main theorem: welltyped programs do not produce type errors if they
- are started in a conformant state.
-\<close>
-theorem no_type_error:
- fixes \<sigma> :: jvm_state
- assumes welltyped: "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P" and conforms: "P,\<Phi> \<turnstile> \<sigma> \<surd>"
- shows "exec_d P \<sigma> \<noteq> TypeError"
-(*<*)
-proof -
- from welltyped obtain mb where wf: "wf_prog mb P" by (fast dest: wt_jvm_progD)
-
- obtain xcp h frs sh where s [simp]: "\<sigma> = (xcp, h, frs, sh)" by (cases \<sigma>)
-
- from conforms have "xcp \<noteq> None \<or> frs = [] \<Longrightarrow> check P \<sigma>"
- by (unfold correct_state_def check_def) auto
- moreover {
- assume "\<not>(xcp \<noteq> None \<or> frs = [])"
- then obtain stk reg C M pc ics frs' where
- xcp [simp]: "xcp = None" and
- frs [simp]: "frs = (stk,reg,C,M,pc,ics)#frs'"
- by (clarsimp simp add: neq_Nil_conv)
-
- from conforms obtain ST LT b Ts T mxs mxl ins xt where
- hconf: "P \<turnstile> h \<surd>" and
- shconf: "P,h \<turnstile>\<^sub>s sh \<surd>" and
- meth: "P \<turnstile> C sees M,b:Ts\<rightarrow>T = (mxs, mxl, ins, xt) in C" and
- \<Phi>: "\<Phi> C M ! pc = Some (ST,LT)" and
- frame: "conf_f P h sh (ST,LT) ins (stk,reg,C,M,pc,ics)" and
- frames: "conf_fs P h sh \<Phi> C M (size Ts) T frs'"
- by (fastforce simp add: correct_state_def dest: sees_method_fun)
-
- from frame obtain
- stk: "P,h \<turnstile> stk [:\<le>] ST" and
- reg: "P,h \<turnstile> reg [:\<le>\<^sub>\<top>] LT" and
- pc: "pc < size ins"
- by (simp add: conf_f_def)
-
- from welltyped meth \<Phi> pc
- have "OK (Some (ST, LT)) \<in> states P mxs (1+size Ts+mxl)
- \<or> OK (Some (ST, LT)) \<in> states P mxs (size Ts+mxl)"
- by (cases b, auto dest: wt_jvm_prog_states_NonStatic wt_jvm_prog_states_Static)
- hence "size ST \<le> mxs" by (auto simp add: JVM_states_unfold)
- with stk have mxs: "size stk \<le> mxs"
- by (auto dest: list_all2_lengthD)
-
- from welltyped meth pc
- have "P,T,mxs,size ins,xt \<turnstile> ins!pc,pc :: \<Phi> C M"
- by (rule wt_jvm_prog_impl_wt_instr)
- hence app\<^sub>0: "app (ins!pc) P mxs T pc (size ins) xt (\<Phi> C M!pc) "
- by (simp add: wt_instr_def)
- with \<Phi> have eff:
- "\<forall>(pc',s')\<in>set (eff (ins ! pc) P pc xt (\<Phi> C M ! pc)). pc' < size ins"
- by (unfold app_def) simp
-
- from app\<^sub>0 \<Phi> have app:
- "xcpt_app (ins!pc) P pc mxs xt (ST,LT) \<and> app\<^sub>i (ins!pc, P, pc, mxs, T, (ST,LT))"
- by (clarsimp simp add: app_def)
-
- with eff stk reg
- have "check_instr (ins!pc) P h stk reg C M pc frs' sh"
- proof (cases "ins!pc")
- case (Getfield F C)
- with app stk reg \<Phi> obtain v vT stk' where
- field: "P \<turnstile> C sees F,NonStatic:vT in C" and
- stk: "stk = v # stk'" and
- conf: "P,h \<turnstile> v :\<le> Class C"
- by auto
- from conf have is_Ref: "is_Ref v" by auto
- moreover {
- assume "v \<noteq> Null"
- with conf field is_Ref wf
- have "\<exists>D vs. h (the_Addr v) = Some (D,vs) \<and> P \<turnstile> D \<preceq>\<^sup>* C"
- by (auto dest!: non_npD)
- }
- ultimately show ?thesis using Getfield field stk
- has_field_mono[OF has_visible_field[OF field]] hconfD[OF hconf]
- by (unfold oconf_def has_field_def) (fastforce dest: has_fields_fun)
- next
- case (Getstatic C F D)
- with app stk reg \<Phi> obtain vT where
- field: "P \<turnstile> C sees F,Static:vT in D"
- by auto
-
- then show ?thesis using Getstatic stk
- has_field_idemp[OF has_visible_field[OF field]] shconfD[OF shconf]
- by (unfold soconf_def has_field_def) (fastforce dest: has_fields_fun)
- next
- case (Putfield F C)
- with app stk reg \<Phi> obtain v ref vT stk' where
- field: "P \<turnstile> C sees F,NonStatic:vT in C" and
- stk: "stk = v # ref # stk'" and
- confv: "P,h \<turnstile> v :\<le> vT" and
- confr: "P,h \<turnstile> ref :\<le> Class C"
- by fastforce
- from confr have is_Ref: "is_Ref ref" by simp
- moreover {
- assume "ref \<noteq> Null"
- with confr field is_Ref wf
- have "\<exists>D vs. h (the_Addr ref) = Some (D,vs) \<and> P \<turnstile> D \<preceq>\<^sup>* C"
- by (auto dest: non_npD)
- }
- ultimately show ?thesis using Putfield field stk confv by fastforce
- next
- case (Invoke M' n)
- with app have n: "n < size ST" by simp
-
- from stk have [simp]: "size stk = size ST" by (rule list_all2_lengthD)
-
- { assume "stk!n = Null" with n Invoke have ?thesis by simp }
- moreover {
- assume "ST!n = NT"
- with n stk have "stk!n = Null" by (auto simp: list_all2_conv_all_nth)
- with n Invoke have ?thesis by simp
- }
- moreover {
- assume Null: "stk!n \<noteq> Null" and NT: "ST!n \<noteq> NT"
-
- from NT app Invoke
- obtain D D' Ts T m where
- D: "ST!n = Class D" and
- M': "P \<turnstile> D sees M',NonStatic: Ts\<rightarrow>T = m in D'" and
- Ts: "P \<turnstile> rev (take n ST) [\<le>] Ts"
- by auto
-
- from D stk n have "P,h \<turnstile> stk!n :\<le> Class D"
- by (auto simp: list_all2_conv_all_nth)
- with Null obtain a C' fs where
- [simp]: "stk!n = Addr a" "h a = Some (C',fs)" and
- "P \<turnstile> C' \<preceq>\<^sup>* D"
- by (fastforce dest!: conf_ClassD)
-
- with M' wf obtain m' Ts' T' D'' where
- C': "P \<turnstile> C' sees M',NonStatic: Ts'\<rightarrow>T' = m' in D''" and
- Ts': "P \<turnstile> Ts [\<le>] Ts'"
- by (auto dest!: sees_method_mono)
-
- from stk have "P,h \<turnstile> take n stk [:\<le>] take n ST" ..
- hence "P,h \<turnstile> rev (take n stk) [:\<le>] rev (take n ST)" ..
- also note Ts also note Ts'
- finally have "P,h \<turnstile> rev (take n stk) [:\<le>] Ts'" .
-
- with Invoke Null n C'
- have ?thesis by (auto simp add: is_Ref_def2 has_methodI)
- }
- ultimately show ?thesis by blast
- next
- case (Invokestatic C M' n)
- with app have n: "n \<le> size ST" by simp
-
- from stk have [simp]: "size stk = size ST" by (rule list_all2_lengthD)
-
- from app Invokestatic
- obtain D Ts T m where
- M': "P \<turnstile> C sees M',Static: Ts\<rightarrow>T = m in D" and
- Ts: "P \<turnstile> rev (take n ST) [\<le>] Ts"
- by auto
-
- from stk have "P,h \<turnstile> take n stk [:\<le>] take n ST" ..
- hence "P,h \<turnstile> rev (take n stk) [:\<le>] rev (take n ST)" ..
- also note Ts
- finally have "P,h \<turnstile> rev (take n stk) [:\<le>] Ts" .
-
- with Invokestatic n M'
- show ?thesis by (auto simp add: is_Ref_def2 has_methodI)
- next
- case Return
- show ?thesis
- proof(cases "M = clinit")
- case True
- have "is_class P C" by(rule sees_method_is_class[OF meth])
- with wf_sees_clinit[OF wf]
- obtain m where "P \<turnstile> C sees clinit,Static: [] \<rightarrow> Void = m in C"
- by(fastforce simp: is_class_def)
-
- with stk app \<Phi> meth frames True Return
- show ?thesis by (auto simp add: has_methodI)
- next
- case False with stk app \<Phi> meth frames Return
- show ?thesis by (auto intro: has_methodI)
- qed
- qed (auto simp add: list_all2_lengthD)
- hence "check P \<sigma>" using meth pc mxs by (auto simp: check_def intro: has_methodI)
- } ultimately
- have "check P \<sigma>" by blast
- thus "exec_d P \<sigma> \<noteq> TypeError" ..
-qed
-(*>*)
-
-
-text \<open>
- The theorem above tells us that, in welltyped programs, the
- defensive machine reaches the same result as the aggressive
- one (after arbitrarily many steps).
-\<close>
-theorem welltyped_aggressive_imp_defensive:
- "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P \<Longrightarrow> P,\<Phi> \<turnstile> \<sigma> \<surd> \<Longrightarrow> P \<turnstile> \<sigma> -jvm\<rightarrow> \<sigma>'
- \<Longrightarrow> P \<turnstile> (Normal \<sigma>) -jvmd\<rightarrow> (Normal \<sigma>')"
-(*<*)
-proof -
- assume wf: "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P" and cf: "P,\<Phi> \<turnstile> \<sigma> \<surd>" and exec: "P \<turnstile> \<sigma> -jvm\<rightarrow> \<sigma>'"
- then have "(\<sigma>, \<sigma>') \<in> {(\<sigma>, \<sigma>'). exec (P, \<sigma>) = \<lfloor>\<sigma>'\<rfloor>}\<^sup>*" by(simp only: exec_all_def)
- then show ?thesis proof(induct rule: rtrancl_induct)
- case base
- then show ?case by (simp add: exec_all_d_def1)
- next
- case (step y z)
- then have \<sigma>y: "P \<turnstile> \<sigma> -jvm\<rightarrow> y" by (simp only: exec_all_def [symmetric])
- have exec_d: "exec_d P y = Normal \<lfloor>z\<rfloor>" using step
- no_type_error_commutes[OF no_type_error[OF wf BV_correct[OF wf \<sigma>y cf]]]
- by (simp add: exec_all_d_def1)
- show ?case using step.hyps(3) exec_1_d_NormalI[OF exec_d]
- by (simp add: exec_all_d_def1)
- qed
-qed
-(*>*)
-
-
-text \<open>
- As corollary we get that the aggressive and the defensive machine
- are equivalent for welltyped programs (if started in a conformant
- state or in the canonical start state)
-\<close>
-corollary welltyped_commutes:
- fixes \<sigma> :: jvm_state
- assumes wf: "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P" and conforms: "P,\<Phi> \<turnstile> \<sigma> \<surd>"
- shows "P \<turnstile> (Normal \<sigma>) -jvmd\<rightarrow> (Normal \<sigma>') = P \<turnstile> \<sigma> -jvm\<rightarrow> \<sigma>'"
-proof(rule iffI)
- assume "P \<turnstile> Normal \<sigma> -jvmd\<rightarrow> Normal \<sigma>'" then show "P \<turnstile> \<sigma> -jvm\<rightarrow> \<sigma>'"
- by (rule defensive_imp_aggressive)
-next
- assume "P \<turnstile> \<sigma> -jvm\<rightarrow> \<sigma>'" then show "P \<turnstile> Normal \<sigma> -jvmd\<rightarrow> Normal \<sigma>'"
- by (rule welltyped_aggressive_imp_defensive [OF wf conforms])
-qed
-
-corollary welltyped_initial_commutes:
- assumes wf: "wf_jvm_prog P"
- assumes nstart: "\<not> is_class P Start"
- assumes meth: "P \<turnstile> C sees M,Static:[]\<rightarrow>Void = b in C"
- assumes nclinit: "M \<noteq> clinit"
- assumes Obj_start_m:
- "(\<And>b' Ts' T' m' D'. P \<turnstile> Object sees start_m, b' : Ts'\<rightarrow>T' = m' in D'
- \<Longrightarrow> b' = Static \<and> Ts' = [] \<and> T' = Void)"
- defines start: "\<sigma> \<equiv> start_state P"
- shows "start_prog P C M \<turnstile> (Normal \<sigma>) -jvmd\<rightarrow> (Normal \<sigma>') = start_prog P C M \<turnstile> \<sigma> -jvm\<rightarrow> \<sigma>'"
-proof -
- from wf obtain \<Phi> where wf': "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P" by (auto simp: wf_jvm_prog_def)
- let ?\<Phi> = "\<Phi>_start \<Phi>"
- from start_prog_wf_jvm_prog_phi[where \<Phi>'="?\<Phi>", OF wf' nstart meth nclinit \<Phi>_start Obj_start_m]
- have "wf_jvm_prog\<^bsub>?\<Phi>\<^esub>(start_prog P C M)" by simp
- moreover
- from wf' nstart meth nclinit \<Phi>_start(2) have "start_prog P C M,?\<Phi> \<turnstile> \<sigma> \<surd>"
- unfolding start by (rule BV_correct_initial)
- ultimately show ?thesis by (rule welltyped_commutes)
-qed
-
-
-lemma not_TypeError_eq [iff]:
- "x \<noteq> TypeError = (\<exists>t. x = Normal t)"
- by (cases x) auto
-
-locale cnf =
- fixes P and \<Phi> and \<sigma>
- assumes wf: "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P"
- assumes cnf: "correct_state P \<Phi> \<sigma>"
-
-theorem (in cnf) no_type_errors:
- "P \<turnstile> (Normal \<sigma>) -jvmd\<rightarrow> \<sigma>' \<Longrightarrow> \<sigma>' \<noteq> TypeError"
-proof -
- assume "P \<turnstile> (Normal \<sigma>) -jvmd\<rightarrow> \<sigma>'"
- then have "(Normal \<sigma>, \<sigma>') \<in> (exec_1_d P)\<^sup>*" by (unfold exec_all_d_def1) simp
- then show ?thesis proof(induct rule: rtrancl_induct)
- case (step y z)
- then obtain y\<^sub>n where [simp]: "y = Normal y\<^sub>n" by clarsimp
- have n\<sigma>y: "P \<turnstile> Normal \<sigma> -jvmd\<rightarrow> Normal y\<^sub>n" using step.hyps(1)
- by (fold exec_all_d_def1) simp
- have \<sigma>y: "P \<turnstile> \<sigma> -jvm\<rightarrow> y\<^sub>n" using defensive_imp_aggressive[OF n\<sigma>y] by simp
- show ?case using step no_type_error[OF wf BV_correct[OF wf \<sigma>y cnf]]
- by (auto simp add: exec_1_d_eq)
- qed simp
-qed
-
-locale start =
- fixes P and C and M and \<sigma> and T and b and P\<^sub>0
- assumes wf: "wf_jvm_prog P"
- assumes nstart: "\<not> is_class P Start"
- assumes sees: "P \<turnstile> C sees M,Static:[]\<rightarrow>Void = b in C"
- assumes nclinit: "M \<noteq> clinit"
- assumes Obj_start_m: "(\<And>b' Ts' T' m' D'. P \<turnstile> Object sees start_m, b' : Ts'\<rightarrow>T' = m' in D'
- \<Longrightarrow> b' = Static \<and> Ts' = [] \<and> T' = Void)"
- defines "\<sigma> \<equiv> Normal (start_state P)"
- defines [simp]: "P\<^sub>0 \<equiv> start_prog P C M"
-
-corollary (in start) bv_no_type_error:
- shows "P\<^sub>0 \<turnstile> \<sigma> -jvmd\<rightarrow> \<sigma>' \<Longrightarrow> \<sigma>' \<noteq> TypeError"
-proof -
- from wf obtain \<Phi> where wf': "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P" by (auto simp: wf_jvm_prog_def)
- let ?\<Phi> = "\<Phi>_start \<Phi>"
- from start_prog_wf_jvm_prog_phi[where \<Phi>'="?\<Phi>", OF wf' nstart sees nclinit \<Phi>_start Obj_start_m]
- have "wf_jvm_prog\<^bsub>?\<Phi>\<^esub>P\<^sub>0" by simp
- moreover
- from BV_correct_initial[where \<Phi>'="?\<Phi>", OF wf' nstart sees nclinit \<Phi>_start(2)]
- have "correct_state P\<^sub>0 ?\<Phi> (start_state P)" by simp
- ultimately have "cnf P\<^sub>0 ?\<Phi> (start_state P)" by (rule cnf.intro)
- moreover assume "P\<^sub>0 \<turnstile> \<sigma> -jvmd\<rightarrow> \<sigma>'"
- ultimately show ?thesis by (unfold \<sigma>_def) (rule cnf.no_type_errors)
-qed
-
-
-end
+(* Title: JinjaDCI/BV/BVNoTypeErrors.thy
+
+ Author: Gerwin Klein, Susannah Mansky
+ Copyright GPL
+
+ Based on the Jinja theory BV/BVNoTypeErrors.thy by Gerwin Klein
+*)
+
+section \<open> Welltyped Programs produce no Type Errors \<close>
+
+theory BVNoTypeError
+imports "../JVM/JVMDefensive" BVSpecTypeSafe
+begin
+
+lemma has_methodI:
+ "P \<turnstile> C sees M,b:Ts\<rightarrow>T = m in D \<Longrightarrow> P \<turnstile> C has M,b"
+ by (unfold has_method_def) blast
+
+text \<open>
+ Some simple lemmas about the type testing functions of the
+ defensive JVM:
+\<close>
+lemma typeof_NoneD [simp,dest]: "typeof v = Some x \<Longrightarrow> \<not>is_Addr v"
+ by (cases v) auto
+
+lemma is_Ref_def2:
+ "is_Ref v = (v = Null \<or> (\<exists>a. v = Addr a))"
+ by (cases v) (auto simp add: is_Ref_def)
+
+lemma [iff]: "is_Ref Null" by (simp add: is_Ref_def2)
+
+lemma is_RefI [intro, simp]: "P,h \<turnstile> v :\<le> T \<Longrightarrow> is_refT T \<Longrightarrow> is_Ref v"
+(*<*)
+proof(cases T)
+qed (auto simp add: is_refT_def is_Ref_def dest: conf_ClassD)
+(*>*)
+
+lemma is_IntgI [intro, simp]: "P,h \<turnstile> v :\<le> Integer \<Longrightarrow> is_Intg v"
+(*<*)by (unfold conf_def) auto(*>*)
+
+lemma is_BoolI [intro, simp]: "P,h \<turnstile> v :\<le> Boolean \<Longrightarrow> is_Bool v"
+(*<*)by (unfold conf_def) auto(*>*)
+
+declare defs1 [simp del]
+
+lemma wt_jvm_prog_states_NonStatic:
+assumes wf: "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P"
+ and mC: "P \<turnstile> C sees M,NonStatic: Ts\<rightarrow>T = (mxs, mxl, ins, et) in C"
+ and \<Phi>: "\<Phi> C M ! pc = \<tau>" and pc: "pc < size ins"
+shows "OK \<tau> \<in> states P mxs (1+size Ts+mxl)"
+(*<*)
+proof -
+ let ?wf_md = "(\<lambda>P C (M, b, Ts, T\<^sub>r, mxs, mxl\<^sub>0, is, xt).
+ wt_method P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt (\<Phi> C M))"
+ have wfmd: "wf_prog ?wf_md P" using wf
+ by (unfold wf_jvm_prog_phi_def) assumption
+ show ?thesis using sees_wf_mdecl[OF wfmd mC] \<Phi> pc
+ by (simp add: wf_mdecl_def wt_method_def check_types_def)
+ (blast intro: nth_in)
+qed
+(*>*)
+
+lemma wt_jvm_prog_states_Static:
+assumes wf: "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P"
+ and mC: "P \<turnstile> C sees M,Static: Ts\<rightarrow>T = (mxs, mxl, ins, et) in C"
+ and \<Phi>: "\<Phi> C M ! pc = \<tau>" and pc: "pc < size ins"
+shows "OK \<tau> \<in> states P mxs (size Ts+mxl)"
+(*<*)
+proof -
+ let ?wf_md = "(\<lambda>P C (M, b, Ts, T\<^sub>r, mxs, mxl\<^sub>0, is, xt).
+ wt_method P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt (\<Phi> C M))"
+ have wfmd: "wf_prog ?wf_md P" using wf
+ by (unfold wf_jvm_prog_phi_def) assumption
+ show ?thesis using sees_wf_mdecl[OF wfmd mC] \<Phi> pc
+ by (simp add: wf_mdecl_def wt_method_def check_types_def)
+ (blast intro: nth_in)
+qed
+(*>*)
+
+text \<open>
+ The main theorem: welltyped programs do not produce type errors if they
+ are started in a conformant state.
+\<close>
+theorem no_type_error:
+ fixes \<sigma> :: jvm_state
+ assumes welltyped: "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P" and conforms: "P,\<Phi> \<turnstile> \<sigma> \<surd>"
+ shows "exec_d P \<sigma> \<noteq> TypeError"
+(*<*)
+proof -
+ from welltyped obtain mb where wf: "wf_prog mb P" by (fast dest: wt_jvm_progD)
+
+ obtain xcp h frs sh where s [simp]: "\<sigma> = (xcp, h, frs, sh)" by (cases \<sigma>)
+
+ from conforms have "xcp \<noteq> None \<or> frs = [] \<Longrightarrow> check P \<sigma>"
+ by (unfold correct_state_def check_def) auto
+ moreover {
+ assume "\<not>(xcp \<noteq> None \<or> frs = [])"
+ then obtain stk reg C M pc ics frs' where
+ xcp [simp]: "xcp = None" and
+ frs [simp]: "frs = (stk,reg,C,M,pc,ics)#frs'"
+ by (clarsimp simp add: neq_Nil_conv)
+
+ from conforms obtain ST LT b Ts T mxs mxl ins xt where
+ hconf: "P \<turnstile> h \<surd>" and
+ shconf: "P,h \<turnstile>\<^sub>s sh \<surd>" and
+ meth: "P \<turnstile> C sees M,b:Ts\<rightarrow>T = (mxs, mxl, ins, xt) in C" and
+ \<Phi>: "\<Phi> C M ! pc = Some (ST,LT)" and
+ frame: "conf_f P h sh (ST,LT) ins (stk,reg,C,M,pc,ics)" and
+ frames: "conf_fs P h sh \<Phi> C M (size Ts) T frs'"
+ by (fastforce simp add: correct_state_def dest: sees_method_fun)
+
+ from frame obtain
+ stk: "P,h \<turnstile> stk [:\<le>] ST" and
+ reg: "P,h \<turnstile> reg [:\<le>\<^sub>\<top>] LT" and
+ pc: "pc < size ins"
+ by (simp add: conf_f_def)
+
+ from welltyped meth \<Phi> pc
+ have "OK (Some (ST, LT)) \<in> states P mxs (1+size Ts+mxl)
+ \<or> OK (Some (ST, LT)) \<in> states P mxs (size Ts+mxl)"
+ by (cases b, auto dest: wt_jvm_prog_states_NonStatic wt_jvm_prog_states_Static)
+ hence "size ST \<le> mxs" by (auto simp add: JVM_states_unfold)
+ with stk have mxs: "size stk \<le> mxs"
+ by (auto dest: list_all2_lengthD)
+
+ from welltyped meth pc
+ have "P,T,mxs,size ins,xt \<turnstile> ins!pc,pc :: \<Phi> C M"
+ by (rule wt_jvm_prog_impl_wt_instr)
+ hence app\<^sub>0: "app (ins!pc) P mxs T pc (size ins) xt (\<Phi> C M!pc) "
+ by (simp add: wt_instr_def)
+ with \<Phi> have eff:
+ "\<forall>(pc',s')\<in>set (eff (ins ! pc) P pc xt (\<Phi> C M ! pc)). pc' < size ins"
+ by (unfold app_def) simp
+
+ from app\<^sub>0 \<Phi> have app:
+ "xcpt_app (ins!pc) P pc mxs xt (ST,LT) \<and> app\<^sub>i (ins!pc, P, pc, mxs, T, (ST,LT))"
+ by (clarsimp simp add: app_def)
+
+ with eff stk reg
+ have "check_instr (ins!pc) P h stk reg C M pc frs' sh"
+ proof (cases "ins!pc")
+ case (Getfield F C)
+ with app stk reg \<Phi> obtain v vT stk' where
+ field: "P \<turnstile> C sees F,NonStatic:vT in C" and
+ stk: "stk = v # stk'" and
+ conf: "P,h \<turnstile> v :\<le> Class C"
+ by auto
+ from conf have is_Ref: "is_Ref v" by auto
+ moreover {
+ assume "v \<noteq> Null"
+ with conf field is_Ref wf
+ have "\<exists>D vs. h (the_Addr v) = Some (D,vs) \<and> P \<turnstile> D \<preceq>\<^sup>* C"
+ by (auto dest!: non_npD)
+ }
+ ultimately show ?thesis using Getfield field stk
+ has_field_mono[OF has_visible_field[OF field]] hconfD[OF hconf]
+ by (unfold oconf_def has_field_def) (fastforce dest: has_fields_fun)
+ next
+ case (Getstatic C F D)
+ with app stk reg \<Phi> obtain vT where
+ field: "P \<turnstile> C sees F,Static:vT in D"
+ by auto
+
+ then show ?thesis using Getstatic stk
+ has_field_idemp[OF has_visible_field[OF field]] shconfD[OF shconf]
+ by (unfold soconf_def has_field_def) (fastforce dest: has_fields_fun)
+ next
+ case (Putfield F C)
+ with app stk reg \<Phi> obtain v ref vT stk' where
+ field: "P \<turnstile> C sees F,NonStatic:vT in C" and
+ stk: "stk = v # ref # stk'" and
+ confv: "P,h \<turnstile> v :\<le> vT" and
+ confr: "P,h \<turnstile> ref :\<le> Class C"
+ by fastforce
+ from confr have is_Ref: "is_Ref ref" by simp
+ moreover {
+ assume "ref \<noteq> Null"
+ with confr field is_Ref wf
+ have "\<exists>D vs. h (the_Addr ref) = Some (D,vs) \<and> P \<turnstile> D \<preceq>\<^sup>* C"
+ by (auto dest: non_npD)
+ }
+ ultimately show ?thesis using Putfield field stk confv by fastforce
+ next
+ case (Invoke M' n)
+ with app have n: "n < size ST" by simp
+
+ from stk have [simp]: "size stk = size ST" by (rule list_all2_lengthD)
+
+ { assume "stk!n = Null" with n Invoke have ?thesis by simp }
+ moreover {
+ assume "ST!n = NT"
+ with n stk have "stk!n = Null" by (auto simp: list_all2_conv_all_nth)
+ with n Invoke have ?thesis by simp
+ }
+ moreover {
+ assume Null: "stk!n \<noteq> Null" and NT: "ST!n \<noteq> NT"
+
+ from NT app Invoke
+ obtain D D' Ts T m where
+ D: "ST!n = Class D" and
+ M': "P \<turnstile> D sees M',NonStatic: Ts\<rightarrow>T = m in D'" and
+ Ts: "P \<turnstile> rev (take n ST) [\<le>] Ts"
+ by auto
+
+ from D stk n have "P,h \<turnstile> stk!n :\<le> Class D"
+ by (auto simp: list_all2_conv_all_nth)
+ with Null obtain a C' fs where
+ [simp]: "stk!n = Addr a" "h a = Some (C',fs)" and
+ "P \<turnstile> C' \<preceq>\<^sup>* D"
+ by (fastforce dest!: conf_ClassD)
+
+ with M' wf obtain m' Ts' T' D'' where
+ C': "P \<turnstile> C' sees M',NonStatic: Ts'\<rightarrow>T' = m' in D''" and
+ Ts': "P \<turnstile> Ts [\<le>] Ts'"
+ by (auto dest!: sees_method_mono)
+
+ from stk have "P,h \<turnstile> take n stk [:\<le>] take n ST" ..
+ hence "P,h \<turnstile> rev (take n stk) [:\<le>] rev (take n ST)" ..
+ also note Ts also note Ts'
+ finally have "P,h \<turnstile> rev (take n stk) [:\<le>] Ts'" .
+
+ with Invoke Null n C'
+ have ?thesis by (auto simp add: is_Ref_def2 has_methodI)
+ }
+ ultimately show ?thesis by blast
+ next
+ case (Invokestatic C M' n)
+ with app have n: "n \<le> size ST" by simp
+
+ from stk have [simp]: "size stk = size ST" by (rule list_all2_lengthD)
+
+ from app Invokestatic
+ obtain D Ts T m where
+ M': "P \<turnstile> C sees M',Static: Ts\<rightarrow>T = m in D" and
+ Ts: "P \<turnstile> rev (take n ST) [\<le>] Ts"
+ by auto
+
+ from stk have "P,h \<turnstile> take n stk [:\<le>] take n ST" ..
+ hence "P,h \<turnstile> rev (take n stk) [:\<le>] rev (take n ST)" ..
+ also note Ts
+ finally have "P,h \<turnstile> rev (take n stk) [:\<le>] Ts" .
+
+ with Invokestatic n M'
+ show ?thesis by (auto simp add: is_Ref_def2 has_methodI)
+ next
+ case Return
+ show ?thesis
+ proof(cases "M = clinit")
+ case True
+ have "is_class P C" by(rule sees_method_is_class[OF meth])
+ with wf_sees_clinit[OF wf]
+ obtain m where "P \<turnstile> C sees clinit,Static: [] \<rightarrow> Void = m in C"
+ by(fastforce simp: is_class_def)
+
+ with stk app \<Phi> meth frames True Return
+ show ?thesis by (auto simp add: has_methodI)
+ next
+ case False with stk app \<Phi> meth frames Return
+ show ?thesis by (auto intro: has_methodI)
+ qed
+ qed (auto simp add: list_all2_lengthD)
+ hence "check P \<sigma>" using meth pc mxs by (auto simp: check_def intro: has_methodI)
+ } ultimately
+ have "check P \<sigma>" by blast
+ thus "exec_d P \<sigma> \<noteq> TypeError" ..
+qed
+(*>*)
+
+
+text \<open>
+ The theorem above tells us that, in welltyped programs, the
+ defensive machine reaches the same result as the aggressive
+ one (after arbitrarily many steps).
+\<close>
+theorem welltyped_aggressive_imp_defensive:
+ "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P \<Longrightarrow> P,\<Phi> \<turnstile> \<sigma> \<surd> \<Longrightarrow> P \<turnstile> \<sigma> -jvm\<rightarrow> \<sigma>'
+ \<Longrightarrow> P \<turnstile> (Normal \<sigma>) -jvmd\<rightarrow> (Normal \<sigma>')"
+(*<*)
+proof -
+ assume wf: "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P" and cf: "P,\<Phi> \<turnstile> \<sigma> \<surd>" and exec: "P \<turnstile> \<sigma> -jvm\<rightarrow> \<sigma>'"
+ then have "(\<sigma>, \<sigma>') \<in> {(\<sigma>, \<sigma>'). exec (P, \<sigma>) = \<lfloor>\<sigma>'\<rfloor>}\<^sup>*" by(simp only: exec_all_def)
+ then show ?thesis proof(induct rule: rtrancl_induct)
+ case base
+ then show ?case by (simp add: exec_all_d_def1)
+ next
+ case (step y z)
+ then have \<sigma>y: "P \<turnstile> \<sigma> -jvm\<rightarrow> y" by (simp only: exec_all_def [symmetric])
+ have exec_d: "exec_d P y = Normal \<lfloor>z\<rfloor>" using step
+ no_type_error_commutes[OF no_type_error[OF wf BV_correct[OF wf \<sigma>y cf]]]
+ by (simp add: exec_all_d_def1)
+ show ?case using step.hyps(3) exec_1_d_NormalI[OF exec_d]
+ by (simp add: exec_all_d_def1)
+ qed
+qed
+(*>*)
+
+
+text \<open>
+ As corollary we get that the aggressive and the defensive machine
+ are equivalent for welltyped programs (if started in a conformant
+ state or in the canonical start state)
+\<close>
+corollary welltyped_commutes:
+ fixes \<sigma> :: jvm_state
+ assumes wf: "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P" and conforms: "P,\<Phi> \<turnstile> \<sigma> \<surd>"
+ shows "P \<turnstile> (Normal \<sigma>) -jvmd\<rightarrow> (Normal \<sigma>') = P \<turnstile> \<sigma> -jvm\<rightarrow> \<sigma>'"
+proof(rule iffI)
+ assume "P \<turnstile> Normal \<sigma> -jvmd\<rightarrow> Normal \<sigma>'" then show "P \<turnstile> \<sigma> -jvm\<rightarrow> \<sigma>'"
+ by (rule defensive_imp_aggressive)
+next
+ assume "P \<turnstile> \<sigma> -jvm\<rightarrow> \<sigma>'" then show "P \<turnstile> Normal \<sigma> -jvmd\<rightarrow> Normal \<sigma>'"
+ by (rule welltyped_aggressive_imp_defensive [OF wf conforms])
+qed
+
+corollary welltyped_initial_commutes:
+ assumes wf: "wf_jvm_prog P"
+ assumes nstart: "\<not> is_class P Start"
+ assumes meth: "P \<turnstile> C sees M,Static:[]\<rightarrow>Void = b in C"
+ assumes nclinit: "M \<noteq> clinit"
+ assumes Obj_start_m:
+ "(\<And>b' Ts' T' m' D'. P \<turnstile> Object sees start_m, b' : Ts'\<rightarrow>T' = m' in D'
+ \<Longrightarrow> b' = Static \<and> Ts' = [] \<and> T' = Void)"
+ defines start: "\<sigma> \<equiv> start_state P"
+ shows "start_prog P C M \<turnstile> (Normal \<sigma>) -jvmd\<rightarrow> (Normal \<sigma>') = start_prog P C M \<turnstile> \<sigma> -jvm\<rightarrow> \<sigma>'"
+proof -
+ from wf obtain \<Phi> where wf': "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P" by (auto simp: wf_jvm_prog_def)
+ let ?\<Phi> = "\<Phi>_start \<Phi>"
+ from start_prog_wf_jvm_prog_phi[where \<Phi>'="?\<Phi>", OF wf' nstart meth nclinit \<Phi>_start Obj_start_m]
+ have "wf_jvm_prog\<^bsub>?\<Phi>\<^esub>(start_prog P C M)" by simp
+ moreover
+ from wf' nstart meth nclinit \<Phi>_start(2) have "start_prog P C M,?\<Phi> \<turnstile> \<sigma> \<surd>"
+ unfolding start by (rule BV_correct_initial)
+ ultimately show ?thesis by (rule welltyped_commutes)
+qed
+
+
+lemma not_TypeError_eq [iff]:
+ "x \<noteq> TypeError = (\<exists>t. x = Normal t)"
+ by (cases x) auto
+
+locale cnf =
+ fixes P and \<Phi> and \<sigma>
+ assumes wf: "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P"
+ assumes cnf: "correct_state P \<Phi> \<sigma>"
+
+theorem (in cnf) no_type_errors:
+ "P \<turnstile> (Normal \<sigma>) -jvmd\<rightarrow> \<sigma>' \<Longrightarrow> \<sigma>' \<noteq> TypeError"
+proof -
+ assume "P \<turnstile> (Normal \<sigma>) -jvmd\<rightarrow> \<sigma>'"
+ then have "(Normal \<sigma>, \<sigma>') \<in> (exec_1_d P)\<^sup>*" by (unfold exec_all_d_def1) simp
+ then show ?thesis proof(induct rule: rtrancl_induct)
+ case (step y z)
+ then obtain y\<^sub>n where [simp]: "y = Normal y\<^sub>n" by clarsimp
+ have n\<sigma>y: "P \<turnstile> Normal \<sigma> -jvmd\<rightarrow> Normal y\<^sub>n" using step.hyps(1)
+ by (fold exec_all_d_def1) simp
+ have \<sigma>y: "P \<turnstile> \<sigma> -jvm\<rightarrow> y\<^sub>n" using defensive_imp_aggressive[OF n\<sigma>y] by simp
+ show ?case using step no_type_error[OF wf BV_correct[OF wf \<sigma>y cnf]]
+ by (auto simp add: exec_1_d_eq)
+ qed simp
+qed
+
+locale start =
+ fixes P and C and M and \<sigma> and T and b and P\<^sub>0
+ assumes wf: "wf_jvm_prog P"
+ assumes nstart: "\<not> is_class P Start"
+ assumes sees: "P \<turnstile> C sees M,Static:[]\<rightarrow>Void = b in C"
+ assumes nclinit: "M \<noteq> clinit"
+ assumes Obj_start_m: "(\<And>b' Ts' T' m' D'. P \<turnstile> Object sees start_m, b' : Ts'\<rightarrow>T' = m' in D'
+ \<Longrightarrow> b' = Static \<and> Ts' = [] \<and> T' = Void)"
+ defines "\<sigma> \<equiv> Normal (start_state P)"
+ defines [simp]: "P\<^sub>0 \<equiv> start_prog P C M"
+
+corollary (in start) bv_no_type_error:
+ shows "P\<^sub>0 \<turnstile> \<sigma> -jvmd\<rightarrow> \<sigma>' \<Longrightarrow> \<sigma>' \<noteq> TypeError"
+proof -
+ from wf obtain \<Phi> where wf': "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P" by (auto simp: wf_jvm_prog_def)
+ let ?\<Phi> = "\<Phi>_start \<Phi>"
+ from start_prog_wf_jvm_prog_phi[where \<Phi>'="?\<Phi>", OF wf' nstart sees nclinit \<Phi>_start Obj_start_m]
+ have "wf_jvm_prog\<^bsub>?\<Phi>\<^esub>P\<^sub>0" by simp
+ moreover
+ from BV_correct_initial[where \<Phi>'="?\<Phi>", OF wf' nstart sees nclinit \<Phi>_start(2)]
+ have "correct_state P\<^sub>0 ?\<Phi> (start_state P)" by simp
+ ultimately have "cnf P\<^sub>0 ?\<Phi> (start_state P)" by (rule cnf.intro)
+ moreover assume "P\<^sub>0 \<turnstile> \<sigma> -jvmd\<rightarrow> \<sigma>'"
+ ultimately show ?thesis by (unfold \<sigma>_def) (rule cnf.no_type_errors)
+qed
+
+
+end
diff --git a/thys/JinjaDCI/BV/BVSpec.thy b/thys/JinjaDCI/BV/BVSpec.thy
--- a/thys/JinjaDCI/BV/BVSpec.thy
+++ b/thys/JinjaDCI/BV/BVSpec.thy
@@ -1,113 +1,113 @@
-(* Title: JinjaDCI/BV/BVSpec.thy
-
- Author: Cornelia Pusch, Gerwin Klein, Susannah Mansky
- Copyright 1999 Technische Universitaet Muenchen, 2019-20 UIUC
-
- Based on the Jinja theory BV/BVSpec.thy by Tobias Nipkow
-*)
-
-section \<open> The Bytecode Verifier \label{sec:BVSpec} \<close>
-
-theory BVSpec
-imports Effect
-begin
-
-text \<open>
- This theory contains a specification of the BV. The specification
- describes correct typings of method bodies; it corresponds
- to type \emph{checking}.
-\<close>
-
-
-definition
- \<comment> \<open>The method type only contains declared classes:\<close>
- check_types :: "'m prog \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> ty\<^sub>i' err list \<Rightarrow> bool"
-where
- "check_types P mxs mxl \<tau>s \<equiv> set \<tau>s \<subseteq> states P mxs mxl"
-
- \<comment> \<open>An instruction is welltyped if it is applicable and its effect\<close>
- \<comment> \<open>is compatible with the type at all successor instructions:\<close>
-definition
- wt_instr :: "['m prog,ty,nat,pc,ex_table,instr,pc,ty\<^sub>m] \<Rightarrow> bool"
- ("_,_,_,_,_ \<turnstile> _,_ :: _" [60,0,0,0,0,0,0,61] 60)
-where
- "P,T,mxs,mpc,xt \<turnstile> i,pc :: \<tau>s \<equiv>
- app i P mxs T pc mpc xt (\<tau>s!pc) \<and>
- (\<forall>(pc',\<tau>') \<in> set (eff i P pc xt (\<tau>s!pc)). P \<turnstile> \<tau>' \<le>' \<tau>s!pc')"
-
- \<comment> \<open>The type at @{text "pc=0"} conforms to the method calling convention:\<close>
-definition wt_start :: "['m prog,cname,staticb,ty list,nat,ty\<^sub>m] \<Rightarrow> bool"
-where
- "wt_start P C b Ts mxl\<^sub>0 \<tau>s \<equiv>
-case b of NonStatic \<Rightarrow> P \<turnstile> Some ([],OK (Class C)#map OK Ts@replicate mxl\<^sub>0 Err) \<le>' \<tau>s!0
- | Static \<Rightarrow> P \<turnstile> Some ([],map OK Ts@replicate mxl\<^sub>0 Err) \<le>' \<tau>s!0"
-
- \<comment> \<open>A method is welltyped if the body is not empty,\<close>
- \<comment> \<open>if the method type covers all instructions and mentions\<close>
- \<comment> \<open>declared classes only, if the method calling convention is respected, and\<close>
- \<comment> \<open>if all instructions are welltyped.\<close>
-definition wt_method :: "['m prog,cname,staticb,ty list,ty,nat,nat,instr list,
- ex_table,ty\<^sub>m] \<Rightarrow> bool"
-where
- "wt_method P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt \<tau>s \<equiv> (b = Static \<or> b = NonStatic) \<and>
- 0 < size is \<and> size \<tau>s = size is \<and>
- check_types P mxs ((case b of Static \<Rightarrow> 0 | NonStatic \<Rightarrow> 1)+size Ts+mxl\<^sub>0) (map OK \<tau>s) \<and>
- wt_start P C b Ts mxl\<^sub>0 \<tau>s \<and>
- (\<forall>pc < size is. P,T\<^sub>r,mxs,size is,xt \<turnstile> is!pc,pc :: \<tau>s)"
-
- \<comment> \<open>A program is welltyped if it is wellformed and all methods are welltyped\<close>
-definition wf_jvm_prog_phi :: "ty\<^sub>P \<Rightarrow> jvm_prog \<Rightarrow> bool" ("wf'_jvm'_prog\<^bsub>_\<^esub>")
-where
- "wf_jvm_prog\<^bsub>\<Phi>\<^esub> \<equiv>
- wf_prog (\<lambda>P C (M,b,Ts,T\<^sub>r,(mxs,mxl\<^sub>0,is,xt)).
- wt_method P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt (\<Phi> C M))"
-
-definition wf_jvm_prog :: "jvm_prog \<Rightarrow> bool"
-where
- "wf_jvm_prog P \<equiv> \<exists>\<Phi>. wf_jvm_prog\<^bsub>\<Phi>\<^esub> P"
-
-lemma wt_jvm_progD:
- "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P \<Longrightarrow> \<exists>wt. wf_prog wt P"
-(*<*) by (unfold wf_jvm_prog_phi_def, blast) (*>*)
-
-lemma wt_jvm_prog_impl_wt_instr:
-assumes wf: "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P" and
- sees: "P \<turnstile> C sees M,b:Ts \<rightarrow> T = (mxs,mxl\<^sub>0,ins,xt) in C" and
- pc: "pc < size ins"
-shows "P,T,mxs,size ins,xt \<turnstile> ins!pc,pc :: \<Phi> C M"
-(*<*)
-proof -
- have wfm: "wf_prog
- (\<lambda>P C (M, b, Ts, T\<^sub>r, mxs, mxl\<^sub>0, is, xt).
- wt_method P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt (\<Phi> C M)) P" using wf
- by (unfold wf_jvm_prog_phi_def)
- show ?thesis using sees_wf_mdecl[OF wfm sees] pc
- by (simp add: wf_mdecl_def wt_method_def)
-qed
-(*>*)
-
-lemma wt_jvm_prog_impl_wt_start:
-assumes wf: "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P" and
- sees: "P \<turnstile> C sees M,b:Ts \<rightarrow> T = (mxs,mxl\<^sub>0,ins,xt) in C"
-shows "0 < size ins \<and> wt_start P C b Ts mxl\<^sub>0 (\<Phi> C M)"
-(*<*)
-proof -
- have wfm: "wf_prog
- (\<lambda>P C (M, b, Ts, T\<^sub>r, mxs, mxl\<^sub>0, is, xt).
- wt_method P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt (\<Phi> C M)) P" using wf
- by (unfold wf_jvm_prog_phi_def)
- show ?thesis using sees_wf_mdecl[OF wfm sees]
- by (simp add: wf_mdecl_def wt_method_def)
-qed
-(*>*)
-
-lemma wf_jvm_prog_nclinit:
-assumes wtp: "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P"
- and meth: "P \<turnstile> C sees M, b : Ts\<rightarrow>T = (mxs, mxl\<^sub>0, ins, xt) in D"
- and wt: "P,T,mxs,size ins,xt \<turnstile> ins!pc,pc :: \<Phi> C M"
- and pc: "pc < length ins" and \<Phi>: "\<Phi> C M ! pc = Some(ST,LT)"
- and ins: "ins ! pc = Invokestatic C\<^sub>0 M\<^sub>0 n"
-shows "M\<^sub>0 \<noteq> clinit"
- using assms by(simp add: wf_jvm_prog_phi_def wt_instr_def app_def)
-
-end
+(* Title: JinjaDCI/BV/BVSpec.thy
+
+ Author: Cornelia Pusch, Gerwin Klein, Susannah Mansky
+ Copyright 1999 Technische Universitaet Muenchen, 2019-20 UIUC
+
+ Based on the Jinja theory BV/BVSpec.thy by Tobias Nipkow
+*)
+
+section \<open> The Bytecode Verifier \label{sec:BVSpec} \<close>
+
+theory BVSpec
+imports Effect
+begin
+
+text \<open>
+ This theory contains a specification of the BV. The specification
+ describes correct typings of method bodies; it corresponds
+ to type \emph{checking}.
+\<close>
+
+
+definition
+ \<comment> \<open>The method type only contains declared classes:\<close>
+ check_types :: "'m prog \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> ty\<^sub>i' err list \<Rightarrow> bool"
+where
+ "check_types P mxs mxl \<tau>s \<equiv> set \<tau>s \<subseteq> states P mxs mxl"
+
+ \<comment> \<open>An instruction is welltyped if it is applicable and its effect\<close>
+ \<comment> \<open>is compatible with the type at all successor instructions:\<close>
+definition
+ wt_instr :: "['m prog,ty,nat,pc,ex_table,instr,pc,ty\<^sub>m] \<Rightarrow> bool"
+ ("_,_,_,_,_ \<turnstile> _,_ :: _" [60,0,0,0,0,0,0,61] 60)
+where
+ "P,T,mxs,mpc,xt \<turnstile> i,pc :: \<tau>s \<equiv>
+ app i P mxs T pc mpc xt (\<tau>s!pc) \<and>
+ (\<forall>(pc',\<tau>') \<in> set (eff i P pc xt (\<tau>s!pc)). P \<turnstile> \<tau>' \<le>' \<tau>s!pc')"
+
+ \<comment> \<open>The type at @{text "pc=0"} conforms to the method calling convention:\<close>
+definition wt_start :: "['m prog,cname,staticb,ty list,nat,ty\<^sub>m] \<Rightarrow> bool"
+where
+ "wt_start P C b Ts mxl\<^sub>0 \<tau>s \<equiv>
+case b of NonStatic \<Rightarrow> P \<turnstile> Some ([],OK (Class C)#map OK Ts@replicate mxl\<^sub>0 Err) \<le>' \<tau>s!0
+ | Static \<Rightarrow> P \<turnstile> Some ([],map OK Ts@replicate mxl\<^sub>0 Err) \<le>' \<tau>s!0"
+
+ \<comment> \<open>A method is welltyped if the body is not empty,\<close>
+ \<comment> \<open>if the method type covers all instructions and mentions\<close>
+ \<comment> \<open>declared classes only, if the method calling convention is respected, and\<close>
+ \<comment> \<open>if all instructions are welltyped.\<close>
+definition wt_method :: "['m prog,cname,staticb,ty list,ty,nat,nat,instr list,
+ ex_table,ty\<^sub>m] \<Rightarrow> bool"
+where
+ "wt_method P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt \<tau>s \<equiv> (b = Static \<or> b = NonStatic) \<and>
+ 0 < size is \<and> size \<tau>s = size is \<and>
+ check_types P mxs ((case b of Static \<Rightarrow> 0 | NonStatic \<Rightarrow> 1)+size Ts+mxl\<^sub>0) (map OK \<tau>s) \<and>
+ wt_start P C b Ts mxl\<^sub>0 \<tau>s \<and>
+ (\<forall>pc < size is. P,T\<^sub>r,mxs,size is,xt \<turnstile> is!pc,pc :: \<tau>s)"
+
+ \<comment> \<open>A program is welltyped if it is wellformed and all methods are welltyped\<close>
+definition wf_jvm_prog_phi :: "ty\<^sub>P \<Rightarrow> jvm_prog \<Rightarrow> bool" ("wf'_jvm'_prog\<^bsub>_\<^esub>")
+where
+ "wf_jvm_prog\<^bsub>\<Phi>\<^esub> \<equiv>
+ wf_prog (\<lambda>P C (M,b,Ts,T\<^sub>r,(mxs,mxl\<^sub>0,is,xt)).
+ wt_method P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt (\<Phi> C M))"
+
+definition wf_jvm_prog :: "jvm_prog \<Rightarrow> bool"
+where
+ "wf_jvm_prog P \<equiv> \<exists>\<Phi>. wf_jvm_prog\<^bsub>\<Phi>\<^esub> P"
+
+lemma wt_jvm_progD:
+ "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P \<Longrightarrow> \<exists>wt. wf_prog wt P"
+(*<*) by (unfold wf_jvm_prog_phi_def, blast) (*>*)
+
+lemma wt_jvm_prog_impl_wt_instr:
+assumes wf: "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P" and
+ sees: "P \<turnstile> C sees M,b:Ts \<rightarrow> T = (mxs,mxl\<^sub>0,ins,xt) in C" and
+ pc: "pc < size ins"
+shows "P,T,mxs,size ins,xt \<turnstile> ins!pc,pc :: \<Phi> C M"
+(*<*)
+proof -
+ have wfm: "wf_prog
+ (\<lambda>P C (M, b, Ts, T\<^sub>r, mxs, mxl\<^sub>0, is, xt).
+ wt_method P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt (\<Phi> C M)) P" using wf
+ by (unfold wf_jvm_prog_phi_def)
+ show ?thesis using sees_wf_mdecl[OF wfm sees] pc
+ by (simp add: wf_mdecl_def wt_method_def)
+qed
+(*>*)
+
+lemma wt_jvm_prog_impl_wt_start:
+assumes wf: "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P" and
+ sees: "P \<turnstile> C sees M,b:Ts \<rightarrow> T = (mxs,mxl\<^sub>0,ins,xt) in C"
+shows "0 < size ins \<and> wt_start P C b Ts mxl\<^sub>0 (\<Phi> C M)"
+(*<*)
+proof -
+ have wfm: "wf_prog
+ (\<lambda>P C (M, b, Ts, T\<^sub>r, mxs, mxl\<^sub>0, is, xt).
+ wt_method P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt (\<Phi> C M)) P" using wf
+ by (unfold wf_jvm_prog_phi_def)
+ show ?thesis using sees_wf_mdecl[OF wfm sees]
+ by (simp add: wf_mdecl_def wt_method_def)
+qed
+(*>*)
+
+lemma wf_jvm_prog_nclinit:
+assumes wtp: "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P"
+ and meth: "P \<turnstile> C sees M, b : Ts\<rightarrow>T = (mxs, mxl\<^sub>0, ins, xt) in D"
+ and wt: "P,T,mxs,size ins,xt \<turnstile> ins!pc,pc :: \<Phi> C M"
+ and pc: "pc < length ins" and \<Phi>: "\<Phi> C M ! pc = Some(ST,LT)"
+ and ins: "ins ! pc = Invokestatic C\<^sub>0 M\<^sub>0 n"
+shows "M\<^sub>0 \<noteq> clinit"
+ using assms by(simp add: wf_jvm_prog_phi_def wt_instr_def app_def)
+
+end
diff --git a/thys/JinjaDCI/BV/BVSpecTypeSafe.thy b/thys/JinjaDCI/BV/BVSpecTypeSafe.thy
--- a/thys/JinjaDCI/BV/BVSpecTypeSafe.thy
+++ b/thys/JinjaDCI/BV/BVSpecTypeSafe.thy
@@ -1,2345 +1,2345 @@
-(* Title: JinjaDCI/BV/BVSpecTypeSafe.thy
-
- Author: Cornelia Pusch, Gerwin Klein, Susannah Mansky
- Copyright 1999 Technische Universitaet Muenchen, 2019-20 UIUC
-
- Based on the Jinja theory BV/BVSpecTypeSafe.thy by Cornelia Pusch and Gerwin Klein
-*)
-
-
-section \<open> BV Type Safety Proof \label{sec:BVSpecTypeSafe} \<close>
-
-theory BVSpecTypeSafe
-imports BVConform StartProg
-begin
-
-text \<open>
- This theory contains proof that the specification of the bytecode
- verifier only admits type safe programs.
-\<close>
-
-subsection \<open> Preliminaries \<close>
-
-text \<open>
- Simp and intro setup for the type safety proof:
-\<close>
-lemmas defs1 = correct_state_def conf_f_def wt_instr_def eff_def norm_eff_def app_def xcpt_app_def
-
-lemmas widen_rules [intro] = conf_widen confT_widen confs_widens confTs_widen
-
-
-subsection \<open> Exception Handling \<close>
-
-
-text \<open>
- For the @{text Invoke} instruction the BV has checked all handlers
- that guard the current @{text pc}.
-\<close>
-lemma Invoke_handlers:
- "match_ex_table P C pc xt = Some (pc',d') \<Longrightarrow>
- \<exists>(f,t,D,h,d) \<in> set (relevant_entries P (Invoke n M) pc xt).
- P \<turnstile> C \<preceq>\<^sup>* D \<and> pc \<in> {f..<t} \<and> pc' = h \<and> d' = d"
- by (induct xt) (auto simp: relevant_entries_def matches_ex_entry_def
- is_relevant_entry_def split: if_split_asm)
-
-text \<open>
- For the @{text Invokestatic} instruction the BV has checked all handlers
- that guard the current @{text pc}.
-\<close>
-lemma Invokestatic_handlers:
- "match_ex_table P C pc xt = Some (pc',d') \<Longrightarrow>
- \<exists>(f,t,D,h,d) \<in> set (relevant_entries P (Invokestatic C\<^sub>0 n M) pc xt).
- P \<turnstile> C \<preceq>\<^sup>* D \<and> pc \<in> {f..<t} \<and> pc' = h \<and> d' = d"
- by (induct xt) (auto simp: relevant_entries_def matches_ex_entry_def
- is_relevant_entry_def split: if_split_asm)
-
-text \<open>
- For the instrs in @{text Called_set} the BV has checked all handlers
- that guard the current @{text pc}.
-\<close>
-lemma Called_set_handlers:
- "match_ex_table P C pc xt = Some (pc',d') \<Longrightarrow> i \<in> Called_set \<Longrightarrow>
- \<exists>(f,t,D,h,d) \<in> set (relevant_entries P i pc xt).
- P \<turnstile> C \<preceq>\<^sup>* D \<and> pc \<in> {f..<t} \<and> pc' = h \<and> d' = d"
- by (induct xt) (auto simp: relevant_entries_def matches_ex_entry_def
- is_relevant_entry_def split: if_split_asm)
-
-text \<open>
- We can prove separately that the recursive search for exception
- handlers (@{text find_handler}) in the frame stack results in
- a conforming state (if there was no matching exception handler
- in the current frame). We require that the exception is a valid
- heap address, and that the state before the exception occurred
- conforms.
-\<close>
-lemma uncaught_xcpt_correct:
- assumes wt: "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P"
- assumes h: "h xcp = Some obj"
- shows "\<And>f. P,\<Phi> \<turnstile> (None, h, f#frs, sh)\<surd>
- \<Longrightarrow> curr_method f \<noteq> clinit \<Longrightarrow> P,\<Phi> \<turnstile> find_handler P xcp h frs sh \<surd>"
- (is "\<And>f. ?correct (None, h, f#frs, sh) \<Longrightarrow> ?prem f \<Longrightarrow> ?correct (?find frs)")
-(*<*)
-proof (induct frs)
- \<comment> \<open>the base
- case is trivial as it should be\<close>
- show "?correct (?find [])" by (simp add: correct_state_def)
-next
- \<comment> \<open>we will need both forms @{text wf_jvm_prog} and @{text wf_prog} later\<close>
- from wt obtain mb where wf: "wf_prog mb P" by (simp add: wf_jvm_prog_phi_def)
-
- \<comment> \<open>the assumptions for the cons case:\<close>
- fix f f' frs' assume cr: "?correct (None, h, f#f'#frs', sh)"
- assume pr: "?prem f"
-
- \<comment> \<open>the induction hypothesis:\<close>
- assume IH: "\<And>f. ?correct (None, h, f#frs', sh) \<Longrightarrow> ?prem f \<Longrightarrow> ?correct (?find frs')"
-
- from cr pr conf_clinit_Cons[where frs="f'#frs'" and f=f] obtain
- confc: "conf_clinit P sh (f'#frs')"
- and cr': "?correct (None, h, f'#frs', sh)" by(fastforce simp: correct_state_def)
-
- obtain stk loc C M pc ics where [simp]: "f' = (stk,loc,C,M,pc,ics)" by (cases f')
-
- from cr' obtain b Ts T mxs mxl\<^sub>0 ins xt where
- meth: "P \<turnstile> C sees M,b:Ts \<rightarrow> T = (mxs,mxl\<^sub>0,ins,xt) in C"
- by (simp add: correct_state_def, blast)
-
- hence xt[simp]: "ex_table_of P C M = xt" by simp
-
- have cls: "is_class P C" by(rule sees_method_is_class'[OF meth])
- from cr' obtain sfs where
- sfs: "M = clinit \<Longrightarrow> sh C = Some(sfs,Processing)" by(fastforce simp: defs1 conf_clinit_def)
-
- show "?correct (?find (f'#frs'))"
- proof (cases "match_ex_table P (cname_of h xcp) pc xt")
- case None with cr' IH[of f'] show ?thesis
- proof(cases "M=clinit")
- case True then show ?thesis using xt cr' IH[of f'] None h conf_clinit_Called_Throwing
- conf_f_Throwing[where h=h and sh=sh, OF _ cls h sfs]
- by(cases frs', auto simp: correct_state_def image_iff) fastforce
- qed(auto)
- next
- fix pc_d
- assume "match_ex_table P (cname_of h xcp) pc xt = Some pc_d"
- then obtain pc' d' where
- match: "match_ex_table P (cname_of h xcp) pc xt = Some (pc',d')"
- (is "?match (cname_of h xcp) = _")
- by (cases pc_d) auto
-
- from wt meth cr' [simplified]
- have wti: "P,T,mxs,size ins,xt \<turnstile> ins!pc,pc :: \<Phi> C M"
- by (fastforce simp: correct_state_def conf_f_def
- dest: sees_method_fun
- elim!: wt_jvm_prog_impl_wt_instr)
-
- from cr' obtain ST LT where \<Phi>: "\<Phi> C M ! pc = Some (ST, LT)"
- by(fastforce dest: sees_method_fun simp: correct_state_def)
-
- from cr' \<Phi> meth have conf': "conf_f P h sh (ST, LT) ins f'"
- by (unfold correct_state_def) (fastforce dest: sees_method_fun)
- hence loc: "P,h \<turnstile> loc [:\<le>\<^sub>\<top>] LT" and
- stk: "P,h \<turnstile> stk [:\<le>] ST" by (unfold conf_f_def) auto
- hence [simp]: "size stk = size ST" by (simp add: list_all2_lengthD)
-
- from cr meth pr
- obtain D n M' where
- ins: "ins!pc = Invoke n M' \<or> ins!pc = Invokestatic D n M'" (is "_ = ?i \<or> _ = ?i'")
- by(fastforce dest: sees_method_fun simp: correct_state_def)
-
- with match obtain f1 t D where
- rel: "(f1,t,D,pc',d') \<in> set (relevant_entries P (ins!pc) pc xt)" and
- D: "P \<turnstile> cname_of h xcp \<preceq>\<^sup>* D"
- by(fastforce dest: Invoke_handlers Invokestatic_handlers)
-
- from rel have
- "(pc', Some (Class D # drop (size ST - d') ST, LT)) \<in> set (xcpt_eff (ins!pc) P pc (ST,LT) xt)"
- (is "(_, Some (?ST',_)) \<in> _")
- by (force simp: xcpt_eff_def image_def)
- with wti \<Phi> obtain
- pc: "pc' < size ins" and
- "P \<turnstile> Some (?ST', LT) \<le>' \<Phi> C M ! pc'"
- by (clarsimp simp: defs1) blast
- then obtain ST' LT' where
- \<Phi>': "\<Phi> C M ! pc' = Some (ST',LT')" and
- less: "P \<turnstile> (?ST', LT) \<le>\<^sub>i (ST',LT')"
- by (auto simp: sup_state_opt_any_Some)
-
- let ?f = "(Addr xcp # drop (length stk - d') stk, loc, C, M, pc',No_ics)"
- have "conf_f P h sh (ST',LT') ins ?f"
- proof -
- from wf less loc have "P,h \<turnstile> loc [:\<le>\<^sub>\<top>] LT'" by simp blast
- moreover from D h have "P,h \<turnstile> Addr xcp :\<le> Class D"
- by (simp add: conf_def obj_ty_def case_prod_unfold)
- with less stk
- have "P,h \<turnstile> Addr xcp # drop (length stk - d') stk [:\<le>] ST'"
- by (auto intro!: list_all2_dropI)
- ultimately show ?thesis using pc conf' by(auto simp: conf_f_def)
- qed
-
- with cr' match \<Phi>' meth pc
- show ?thesis by (unfold correct_state_def)
- (cases "M=clinit"; fastforce dest: sees_method_fun simp: conf_clinit_def distinct_clinit_def)
- qed
-qed
-(*>*)
-
-text \<open>
- The requirement of lemma @{text uncaught_xcpt_correct} (that
- the exception is a valid reference on the heap) is always met
- for welltyped instructions and conformant states:
-\<close>
-lemma exec_instr_xcpt_h:
- "\<lbrakk> fst (exec_instr (ins!pc) P h stk vars C M pc ics frs sh) = Some xcp;
- P,T,mxs,size ins,xt \<turnstile> ins!pc,pc :: \<Phi> C M;
- P,\<Phi> \<turnstile> (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\<surd> \<rbrakk>
- \<Longrightarrow> \<exists>obj. h xcp = Some obj"
- (is "\<lbrakk> ?xcpt; ?wt; ?correct \<rbrakk> \<Longrightarrow> ?thesis")
-(*<*)
-proof -
- note [simp] = split_beta
- note [split] = if_split_asm option.split_asm
-
- assume wt: ?wt ?correct
- hence pre: "preallocated h" by (simp add: correct_state_def hconf_def)
-
- assume xcpt: ?xcpt
- with exec_instr_xcpts have
- opt: "ins!pc = Throw \<or> xcp \<in> {a. \<exists>x \<in> sys_xcpts. a = addr_of_sys_xcpt x}" by simp
-
- with pre show ?thesis
- proof (cases "ins!pc")
- case Throw with xcpt wt pre show ?thesis
- by (clarsimp iff: list_all2_Cons2 simp: defs1)
- (auto dest: non_npD simp: is_refT_def elim: preallocatedE)
- qed (auto elim: preallocatedE)
-qed
-(*>*)
-
-lemma exec_step_xcpt_h:
-assumes xcpt: "fst (exec_step P h stk vars C M pc ics frs sh) = Some xcp"
- and ins: "instrs_of P C M = ins"
- and wti: "P,T,mxs,size ins,xt \<turnstile> ins!pc,pc :: \<Phi> C M"
- and correct: "P,\<Phi> \<turnstile> (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\<surd>"
-shows "\<exists>obj. h xcp = Some obj"
-proof -
- from correct have pre: "preallocated h" by(simp add: defs1 hconf_def)
- { fix C' Cs assume ics[simp]: "ics = Calling C' Cs"
- with xcpt have "xcp = addr_of_sys_xcpt NoClassDefFoundError"
- by(cases ics, auto simp: split_beta split: init_state.splits if_split_asm)
- with pre have ?thesis using preallocated_def by force
- }
- moreover
- { fix Cs a assume [simp]: "ics = Throwing Cs a"
- with xcpt have eq: "a = xcp" by(cases Cs; simp)
-
- from correct have "P,h,sh \<turnstile>\<^sub>i (C,M,pc,ics)" by(auto simp: defs1)
- with eq have ?thesis by simp
- }
- moreover
- { fix Cs assume ics: "ics = No_ics \<or> ics = Called Cs"
- with exec_instr_xcpt_h[OF _ wti correct] xcpt ins have ?thesis by(cases Cs, auto)
- }
- ultimately show ?thesis by(cases ics, auto)
-qed
-
-lemma conf_sys_xcpt:
- "\<lbrakk>preallocated h; C \<in> sys_xcpts\<rbrakk> \<Longrightarrow> P,h \<turnstile> Addr (addr_of_sys_xcpt C) :\<le> Class C"
- by (auto elim: preallocatedE)
-
-lemma match_ex_table_SomeD:
- "match_ex_table P C pc xt = Some (pc',d') \<Longrightarrow>
- \<exists>(f,t,D,h,d) \<in> set xt. matches_ex_entry P C pc (f,t,D,h,d) \<and> h = pc' \<and> d=d'"
- by (induct xt) (auto split: if_split_asm)
-
-text \<open>
- Finally we can state that, whenever an exception occurs, the
- next state always conforms:
-\<close>
-lemma xcpt_correct:
- fixes \<sigma>' :: jvm_state
- assumes wtp: "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P"
- assumes meth: "P \<turnstile> C sees M,b:Ts\<rightarrow>T=(mxs,mxl\<^sub>0,ins,xt) in C"
- assumes wt: "P,T,mxs,size ins,xt \<turnstile> ins!pc,pc :: \<Phi> C M"
- assumes xp: "fst (exec_step P h stk loc C M pc ics frs sh) = Some xcp"
- assumes s': "Some \<sigma>' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)"
- assumes correct: "P,\<Phi> \<turnstile> (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\<surd>"
- shows "P,\<Phi> \<turnstile> \<sigma>'\<surd>"
-(*<*)
-proof -
- from wtp obtain wfmb where wf: "wf_prog wfmb P"
- by (simp add: wf_jvm_prog_phi_def)
-
- from meth have ins[simp]: "instrs_of P C M = ins" by simp
- have cls: "is_class P C" by(rule sees_method_is_class[OF meth])
- from correct obtain sfs where
- sfs: "M = clinit \<Longrightarrow> sh C = Some(sfs,Processing)"
- by(auto simp: correct_state_def conf_clinit_def conf_f_def2)
-
- note conf_sys_xcpt [elim!]
- note xp' = meth s' xp
-
- from correct meth
- obtain ST LT where
- h_ok: "P \<turnstile> h \<surd>" and
- sh_ok: "P,h \<turnstile>\<^sub>s sh \<surd>" and
- \<Phi>_pc: "\<Phi> C M ! pc = Some (ST, LT)" and
- frame: "conf_f P h sh (ST,LT) ins (stk,loc,C,M,pc,ics)" and
- frames: "conf_fs P h sh \<Phi> C M (size Ts) T frs" and
- confc: "conf_clinit P sh ((stk,loc,C,M,pc,ics)#frs)" and
- vics: "P,h,sh \<turnstile>\<^sub>i (C,M,pc,ics)"
- by(auto simp: defs1 dest: sees_method_fun)
-
- from frame obtain
- stk: "P,h \<turnstile> stk [:\<le>] ST" and
- loc: "P,h \<turnstile> loc [:\<le>\<^sub>\<top>] LT" and
- pc: "pc < size ins"
- by (unfold conf_f_def) auto
-
- from h_ok have preh: "preallocated h" by (simp add: hconf_def)
-
- note wtp
- moreover
- from exec_step_xcpt_h[OF xp ins wt correct]
- obtain obj where h: "h xcp = Some obj" by clarify
- moreover note correct
- ultimately
- have fh: "curr_method (stk,loc,C,M,pc,ics) \<noteq> clinit
- \<Longrightarrow> P,\<Phi> \<turnstile> find_handler P xcp h frs sh \<surd>" by (rule uncaught_xcpt_correct)
- with xp'
- have "M \<noteq> clinit \<Longrightarrow> \<forall>Cs a. ics \<noteq> Throwing Cs a
- \<Longrightarrow> match_ex_table P (cname_of h xcp) pc xt = None \<Longrightarrow> ?thesis"
- (is "?nc \<Longrightarrow> ?t \<Longrightarrow> ?m (cname_of h xcp) = _ \<Longrightarrow> _" is "?nc \<Longrightarrow> ?t \<Longrightarrow> ?match = _ \<Longrightarrow> _")
- by(cases ics; simp add: split_beta)
- moreover
- from correct xp' conf_clinit_Called_Throwing conf_f_Throwing[where h=h and sh=sh, OF _ cls h sfs]
- have "M = clinit \<Longrightarrow> \<forall>Cs a. ics \<noteq> Throwing Cs a
- \<Longrightarrow> match_ex_table P (cname_of h xcp) pc xt = None \<Longrightarrow> ?thesis"
- by(cases frs, auto simp: correct_state_def image_iff split_beta) fastforce
- moreover
- { fix pc_d assume "?match = Some pc_d"
- then obtain pc' d' where some_handler: "?match = Some (pc',d')"
- by (cases pc_d) auto
-
- from stk have [simp]: "size stk = size ST" ..
-
- from wt \<Phi>_pc have
- eff: "\<forall>(pc', s')\<in>set (xcpt_eff (ins!pc) P pc (ST,LT) xt).
- pc' < size ins \<and> P \<turnstile> s' \<le>' \<Phi> C M!pc'"
- by (auto simp: defs1)
-
- from some_handler obtain f t D where
- xt: "(f,t,D,pc',d') \<in> set xt" and
- "matches_ex_entry P (cname_of h xcp) pc (f,t,D,pc',d')"
- by (auto dest: match_ex_table_SomeD)
-
- hence match: "P \<turnstile> cname_of h xcp \<preceq>\<^sup>* D" "pc \<in> {f..<t}"
- by (auto simp: matches_ex_entry_def)
-
- { fix C' Cs assume ics: "ics = Calling C' Cs \<or> ics = Called (C'#Cs)"
-
- let ?stk' = "Addr xcp # drop (length stk - d') stk"
- let ?f = "(?stk', loc, C, M, pc', No_ics)"
- from some_handler xp' ics
- have \<sigma>': "\<sigma>' = (None, h, ?f#frs, sh)"
- by (cases ics; simp add: split_beta)
-
- from xp ics have "xcp = addr_of_sys_xcpt NoClassDefFoundError"
- by(cases ics, auto simp: split_beta split: init_state.splits if_split_asm)
- with match preh have conf: "P,h \<turnstile> Addr xcp :\<le> Class D" by fastforce
-
- from correct ics obtain C1 where "Called_context P C1 (ins!pc)"
- by(fastforce simp: correct_state_def conf_f_def2)
- then have "ins!pc \<in> Called_set" by(rule Called_context_Called_set)
- with xt match have "(f,t,D,pc',d') \<in> set (relevant_entries P (ins!pc) pc xt)"
- by(auto simp: relevant_entries_def is_relevant_entry_def)
-
- with eff obtain ST' LT' where
- \<Phi>_pc': "\<Phi> C M ! pc' = Some (ST', LT')" and
- pc': "pc' < size ins" and
- less: "P \<turnstile> (Class D # drop (size ST - d') ST, LT) \<le>\<^sub>i (ST', LT')"
- by (fastforce simp: xcpt_eff_def sup_state_opt_any_Some)
-
- with conf loc stk conf_f_def2 frame ics have "conf_f P h sh (ST',LT') ins ?f"
- by (auto simp: defs1 intro: list_all2_dropI)
- with meth h_ok frames \<Phi>_pc' \<sigma>' sh_ok confc ics
- have ?thesis
- by (unfold correct_state_def)
- (auto dest: sees_method_fun conf_clinit_diff' sees_method_is_class; fastforce)
- }
- moreover
- { assume ics: "ics = No_ics \<or> ics = Called []"
-
- let ?stk' = "Addr xcp # drop (length stk - d') stk"
- let ?f = "(?stk', loc, C, M, pc', No_ics)"
- from some_handler xp' ics
- have \<sigma>': "\<sigma>' = (None, h, ?f#frs, sh)"
- by (cases ics; simp add: split_beta)
-
- from xp ics obtain
- "(f,t,D,pc',d') \<in> set (relevant_entries P (ins!pc) pc xt)" and
- conf: "P,h \<turnstile> Addr xcp :\<le> Class D"
- proof (cases "ins!pc")
- case Return
- with xp ics have False by(cases ics; cases frs, auto simp: split_beta split: if_split_asm)
- then show ?thesis by simp
- next
- case New with xp match
- have "is_relevant_entry P (ins!pc) pc (f,t,D,pc',d')"
- by (simp add: is_relevant_entry_def)
- moreover
- from xp wt correct obtain obj where xcp: "h xcp = Some obj"
- by (blast dest: exec_step_xcpt_h[OF _ ins])
- ultimately
- show ?thesis using xt match
- by (auto simp: relevant_entries_def conf_def case_prod_unfold intro: that)
- next
- case Getfield with xp ics
- have xcp: "xcp = addr_of_sys_xcpt NullPointer \<or> xcp = addr_of_sys_xcpt NoSuchFieldError
- \<or> xcp = addr_of_sys_xcpt IncompatibleClassChangeError"
- by (cases ics; simp add: split_beta split: if_split_asm staticb.splits)
- with Getfield match preh have "is_relevant_entry P (ins!pc) pc (f,t,D,pc',d')"
- by (fastforce simp: is_relevant_entry_def)
- with match preh xt xcp
- show ?thesis by(fastforce simp: relevant_entries_def intro: that)
- next
- case Getstatic with xp match
- have "is_relevant_entry P (ins!pc) pc (f,t,D,pc',d')"
- by (simp add: is_relevant_entry_def)
- moreover
- from xp wt correct obtain obj where xcp: "h xcp = Some obj"
- by (blast dest: exec_step_xcpt_h[OF _ ins])
- ultimately
- show ?thesis using xt match
- by (auto simp: relevant_entries_def conf_def case_prod_unfold intro: that)
- next
- case Putfield with xp ics
- have xcp: "xcp = addr_of_sys_xcpt NullPointer \<or> xcp = addr_of_sys_xcpt NoSuchFieldError
- \<or> xcp = addr_of_sys_xcpt IncompatibleClassChangeError"
- by (cases ics; simp add: split_beta split: if_split_asm staticb.splits)
- with Putfield match preh have "is_relevant_entry P (ins!pc) pc (f,t,D,pc',d')"
- by (fastforce simp: is_relevant_entry_def)
- with match preh xt xcp
- show ?thesis by (fastforce simp: relevant_entries_def intro: that)
- next
- case Putstatic with xp match
- have "is_relevant_entry P (ins!pc) pc (f,t,D,pc',d')"
- by (simp add: is_relevant_entry_def)
- moreover
- from xp wt correct obtain obj where xcp: "h xcp = Some obj"
- by (blast dest: exec_step_xcpt_h[OF _ ins])
- ultimately
- show ?thesis using xt match
- by (auto simp: relevant_entries_def conf_def case_prod_unfold intro: that)
- next
- case Checkcast with xp ics
- have [simp]: "xcp = addr_of_sys_xcpt ClassCast"
- by (cases ics; simp add: split_beta split: if_split_asm)
- with Checkcast match preh have "is_relevant_entry P (ins!pc) pc (f,t,D,pc',d')"
- by (simp add: is_relevant_entry_def)
- with match preh xt
- show ?thesis by (fastforce simp: relevant_entries_def intro: that)
- next
- case Invoke with xp match
- have "is_relevant_entry P (ins!pc) pc (f,t,D,pc',d')"
- by (simp add: is_relevant_entry_def)
- moreover
- from xp wt correct obtain obj where xcp: "h xcp = Some obj"
- by (blast dest: exec_step_xcpt_h[OF _ ins])
- ultimately
- show ?thesis using xt match
- by (auto simp: relevant_entries_def conf_def case_prod_unfold intro: that)
- next
- case Invokestatic with xp match
- have "is_relevant_entry P (ins!pc) pc (f,t,D,pc',d')"
- by (simp add: is_relevant_entry_def)
- moreover
- from xp wt correct obtain obj where xcp: "h xcp = Some obj"
- by (blast dest: exec_step_xcpt_h[OF _ ins])
- ultimately
- show ?thesis using xt match
- by (auto simp: relevant_entries_def conf_def case_prod_unfold intro: that)
- next
- case Throw with xp match preh
- have "is_relevant_entry P (ins!pc) pc (f,t,D,pc',d')"
- by (simp add: is_relevant_entry_def)
- moreover
- from xp wt correct obtain obj where xcp: "h xcp = Some obj"
- by (blast dest: exec_step_xcpt_h[OF _ ins])
- ultimately
- show ?thesis using xt match
- by (auto simp: relevant_entries_def conf_def case_prod_unfold intro: that)
- qed(cases ics, (auto)[5])+
-
- with eff obtain ST' LT' where
- \<Phi>_pc': "\<Phi> C M ! pc' = Some (ST', LT')" and
- pc': "pc' < size ins" and
- less: "P \<turnstile> (Class D # drop (size ST - d') ST, LT) \<le>\<^sub>i (ST', LT')"
- by (fastforce simp: xcpt_eff_def sup_state_opt_any_Some)
-
- with conf loc stk conf_f_def2 frame ics have "conf_f P h sh (ST',LT') ins ?f"
- by (auto simp: defs1 intro: list_all2_dropI)
- with meth h_ok frames \<Phi>_pc' \<sigma>' sh_ok confc ics
- have ?thesis
- by (unfold correct_state_def) (auto dest: sees_method_fun conf_clinit_diff'; fastforce)
- }
- ultimately
- have "\<forall>Cs a. ics \<noteq> Throwing Cs a \<Longrightarrow> ?thesis" by(cases ics; metis list.exhaust)
- }
- moreover
- { fix Cs a assume "ics = Throwing Cs a"
- with xp' have ics: "ics = Throwing [] xcp" by(cases Cs; clarsimp)
-
- let ?frs = "(stk,loc,C,M,pc,No_ics)#frs"
-
- have eT: "exec_step P h stk loc C M pc (Throwing [] xcp) frs sh = (Some xcp, h, ?frs, sh)"
- by auto
- with xp' ics have \<sigma>'_fh: "\<sigma>' = find_handler P xcp h ?frs sh" by simp
-
- from meth have [simp]: "xt = ex_table_of P C M" by simp
-
- let ?match = "match_ex_table P (cname_of h xcp) pc xt"
-
- { assume clinit: "M = clinit" and None: "?match = None"
- note asms = clinit None
-
- have "P,\<Phi> |- find_handler P xcp h ?frs sh [ok]"
- proof(cases frs)
- case Nil
- with h_ok sh_ok asms show "P,\<Phi> |- find_handler P xcp h ?frs sh [ok]"
- by(simp add: correct_state_def)
- next
- case [simp]: (Cons f' frs')
- obtain stk' loc' C' M' pc' ics' where
- [simp]: "f' = (stk',loc',C',M',pc',ics')" by(cases f')
-
- have cls: "is_class P C" by(rule sees_method_is_class[OF meth])
- have shC: "sh C = Some(sfs,Processing)" by(rule sfs[OF clinit])
-
- from correct obtain b Ts T mxs' mxl\<^sub>0' ins' xt' ST' LT' where
- meth': "P \<turnstile> C' sees M', b : Ts\<rightarrow>T = (mxs', mxl\<^sub>0', ins', xt') in C'" and
- \<Phi>_pc': "\<Phi> C' M' ! pc' = \<lfloor>(ST', LT')\<rfloor>" and
- frame': "conf_f P h sh (ST',LT') ins' (stk', loc', C', M', pc', ics')" and
- frames': "conf_fs P h sh \<Phi> C' M' (length Ts) T frs'" and
- confc': "conf_clinit P sh ((stk',loc',C',M',pc',ics')#frs')"
- by(auto dest: conf_clinit_Cons simp: correct_state_def)
-
- from meth' have
- ins'[simp]: "instrs_of P C' M' = ins'"
- and [simp]: "xt' = ex_table_of P C' M'" by simp+
-
- let ?f' = "case ics' of Called Cs' \<Rightarrow> (stk',loc',C',M',pc',Throwing (C#Cs') xcp)
- | _ \<Rightarrow> (stk',loc',C',M',pc',ics')"
-
- from asms confc have confc_T: "conf_clinit P sh (?f'#frs')"
- by(cases ics', auto simp: conf_clinit_def distinct_clinit_def)
-
- from asms conf_f_Throwing[where h=h and sh=sh, OF _ cls h shC] frame' have
- frame_T: "conf_f P h sh (ST', LT') ins' ?f'" by(cases ics'; simp)
- with h_ok sh_ok meth' \<Phi>_pc' confc_T frames'
- have "P,\<Phi> |- (None, h, ?f'#frs', sh) [ok]"
- by(cases ics') (fastforce simp: correct_state_def)+
-
- with asms show ?thesis by(cases ics'; simp)
- qed
- }
- moreover
- { assume asms: "M \<noteq> clinit" "?match = None"
-
- from asms uncaught_xcpt_correct[OF wtp h correct]
- have "P,\<Phi> |- find_handler P xcp h frs sh [ok]" by simp
- with asms have "P,\<Phi> |- find_handler P xcp h ?frs sh [ok]" by auto
- }
- moreover
- { fix pc_d assume some_handler: "?match = \<lfloor>pc_d\<rfloor>"
- (is "?match = \<lfloor>pc_d\<rfloor>")
- then obtain pc1 d1 where sh': "?match = Some(pc1,d1)" by(cases pc_d, simp)
-
- let ?stk' = "Addr xcp # drop (length stk - d1) stk"
- let ?f = "(?stk', loc, C, M, pc1, No_ics)"
-
- from stk have [simp]: "size stk = size ST" ..
-
- from wt \<Phi>_pc have
- eff: "\<forall>(pc1, s')\<in>set (xcpt_eff (ins!pc) P pc (ST,LT) xt).
- pc1 < size ins \<and> P \<turnstile> s' \<le>' \<Phi> C M!pc1"
- by (auto simp: defs1)
-
- from match_ex_table_SomeD[OF sh'] obtain f t D where
- xt: "(f,t,D,pc1,d1) \<in> set xt" and
- "matches_ex_entry P (cname_of h xcp) pc (f,t,D,pc1,d1)" by auto
-
- hence match: "P \<turnstile> cname_of h xcp \<preceq>\<^sup>* D" "pc \<in> {f..<t}"
- by (auto simp: matches_ex_entry_def)
-
- from ics vics obtain C1 where "Called_context P C1 (ins ! pc)" by auto
- then have "ins!pc \<in> Called_set" by(rule Called_context_Called_set)
- with match xt xp ics obtain
- res: "(f,t,D,pc1,d1) \<in> set (relevant_entries P (ins!pc) pc xt)"
- by(auto simp: relevant_entries_def is_relevant_entry_def)
-
- with h match xt xp ics have conf: "P,h \<turnstile> Addr xcp :\<le> Class D"
- by (auto simp: relevant_entries_def conf_def case_prod_unfold)
-
- with eff res obtain ST1 LT1 where
- \<Phi>_pc1: "\<Phi> C M ! pc1 = Some (ST1, LT1)" and
- pc1: "pc1 < size ins" and
- less1: "P \<turnstile> (Class D # drop (size ST - d1) ST, LT) \<le>\<^sub>i (ST1, LT1)"
- by (fastforce simp: xcpt_eff_def sup_state_opt_any_Some)
-
- with conf loc stk conf_f_def2 frame ics have frame1: "conf_f P h sh (ST1,LT1) ins ?f"
- by (auto simp: defs1 intro: list_all2_dropI)
-
- from \<Phi>_pc1 h_ok sh_ok meth frame1 frames conf_clinit_diff'[OF confc] have
- "P,\<Phi> |- (None, h, ?f # frs, sh) [ok]" by(fastforce simp: correct_state_def)
- with sh' have "P,\<Phi> |- find_handler P xcp h ?frs sh [ok]" by auto
- }
- ultimately
- have cr': "P,\<Phi> |- find_handler P xcp h ?frs sh [ok]" by(cases "?match") blast+
-
- with \<sigma>'_fh have ?thesis by simp
- }
- ultimately
- show ?thesis by (cases "?match") blast+
-qed
-(*>*)
-
-(**********Non-exception Single-step correctness*************************)
-declare defs1 [simp]
-
-subsection \<open> Initialization procedure steps \<close>
-
-text \<open>
- In this section we prove that, for states that result in a step of the
- initialization procedure rather than an instruction execution, the state
- after execution of the step still conforms.
-\<close>
-
-lemma Calling_correct:
- fixes \<sigma>' :: jvm_state
- assumes wtprog: "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P"
- assumes mC: "P \<turnstile> C sees M,b:Ts\<rightarrow>T=(mxs,mxl\<^sub>0,ins,xt) in C"
- assumes s': "Some \<sigma>' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)"
- assumes cf: "P,\<Phi> \<turnstile> (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\<surd>"
- assumes xc: "fst (exec_step P h stk loc C M pc ics frs sh) = None"
- assumes ics: "ics = Calling C' Cs"
-
- shows "P,\<Phi> \<turnstile> \<sigma>'\<surd>"
-proof -
- from wtprog obtain wfmb where wf: "wf_prog wfmb P"
- by (simp add: wf_jvm_prog_phi_def)
-
- from mC cf obtain ST LT where
- h_ok: "P \<turnstile> h \<surd>" and
- sh_ok: "P,h \<turnstile>\<^sub>s sh \<surd>" and
- \<Phi>: "\<Phi> C M ! pc = Some (ST,LT)" and
- stk: "P,h \<turnstile> stk [:\<le>] ST" and loc: "P,h \<turnstile> loc [:\<le>\<^sub>\<top>] LT" and
- pc: "pc < size ins" and
- frame: "conf_f P h sh (ST, LT) ins (stk,loc,C,M,pc,ics)" and
- fs: "conf_fs P h sh \<Phi> C M (size Ts) T frs" and
- confc: "conf_clinit P sh ((stk,loc,C,M,pc,ics)#frs)" and
- vics: "P,h,sh \<turnstile>\<^sub>i (C,M,pc,ics)"
- by (fastforce dest: sees_method_fun)
-
- with ics have confc\<^sub>0: "conf_clinit P sh ((stk,loc,C,M,pc,Calling C' Cs)#frs)" by simp
-
- from vics ics have cls': "is_class P C'" by auto
-
- { assume None: "sh C' = None"
-
- let ?sh = "sh(C' \<mapsto> (sblank P C', Prepared))"
-
- obtain FDTs where
- flds: "P \<turnstile> C' has_fields FDTs" using wf_Fields_Ex[OF wf cls'] by clarsimp
-
- from shconf_upd_obj[where C=C', OF sh_ok soconf_sblank[OF flds]]
- have sh_ok': "P,h \<turnstile>\<^sub>s ?sh \<surd>" by simp
-
- from None have "\<forall>sfs. sh C' \<noteq> Some(sfs,Processing)" by simp
- with conf_clinit_nProc_dist[OF confc] have
- dist': "distinct (C' # clinit_classes ((stk, loc, C, M, pc, ics) # frs))" by simp
- then have dist'': "distinct (C' # clinit_classes frs)" by simp
-
- have confc': "conf_clinit P ?sh ((stk, loc, C, M, pc, ics) # frs)"
- by(rule conf_clinit_shupd[OF confc dist'])
- have fs': "conf_fs P h ?sh \<Phi> C M (size Ts) T frs" by(rule conf_fs_shupd[OF fs dist''])
- from vics ics have vics': "P,h,?sh \<turnstile>\<^sub>i (C, M, pc, ics)" by auto
-
- from s' ics None have "\<sigma>' = (None, h, (stk, loc, C, M, pc, ics)#frs, ?sh)" by auto
-
- with mC h_ok sh_ok' \<Phi> stk loc pc fs' confc vics' confc' frame None
- have ?thesis by fastforce
- }
- moreover
- { fix a assume "sh C' = Some a"
- then obtain sfs i where shC'[simp]: "sh C' = Some(sfs,i)" by(cases a, simp)
-
- from confc ics have last: "\<exists>sobj. sh (last(C'#Cs)) = Some sobj"
- by(fastforce simp: conf_clinit_def)
-
- let "?f" = "\<lambda>ics'. (stk, loc, C, M, pc, ics'::init_call_status)"
-
- { assume i: "i = Done \<or> i = Processing"
- let ?ics = "Called Cs"
-
- from last vics ics have vics': "P,h,sh \<turnstile>\<^sub>i (C, M, pc, ?ics)" by auto
- from confc ics have confc': "conf_clinit P sh (?f ?ics#frs)"
- by(cases "M=clinit"; clarsimp simp: conf_clinit_def distinct_clinit_def)
-
- from i s' ics have "\<sigma>' = (None, h, ?f ?ics#frs, sh)" by auto
-
- with mC h_ok sh_ok \<Phi> stk loc pc fs confc' vics' frame ics
- have ?thesis by fastforce
- }
- moreover
- { assume i[simp]: "i = Error"
- let ?a = "addr_of_sys_xcpt NoClassDefFoundError"
- let ?ics = "Throwing Cs ?a"
-
- from h_ok have preh: "preallocated h" by (simp add: hconf_def)
- then obtain obj where ha: "h ?a = Some obj" by(clarsimp simp: preallocated_def sys_xcpts_def)
- with vics ics have vics': "P,h,sh \<turnstile>\<^sub>i (C, M, pc, ?ics)" by auto
-
- from confc ics have confc'': "conf_clinit P sh (?f ?ics#frs)"
- by(cases "M=clinit"; clarsimp simp: conf_clinit_def distinct_clinit_def)
-
- from s' ics have \<sigma>': "\<sigma>' = (None, h, ?f ?ics#frs, sh)" by auto
-
- from mC h_ok sh_ok \<Phi> stk loc pc fs confc'' vics \<sigma>' ics ha
- have ?thesis by fastforce
- }
- moreover
- { assume i[simp]: "i = Prepared"
- let ?sh = "sh(C' \<mapsto> (sfs,Processing))"
- let ?D = "fst(the(class P C'))"
- let ?ics = "if C' = Object then Called (C'#Cs) else Calling ?D (C'#Cs)"
-
- from shconf_upd_obj[where C=C', OF sh_ok shconfD[OF sh_ok shC']]
- have sh_ok': "P,h \<turnstile>\<^sub>s ?sh \<surd>" by simp
-
- from cls' have "C' \<noteq> Object \<Longrightarrow> P \<turnstile> C' \<preceq>\<^sup>* ?D" by(auto simp: is_class_def intro!: subcls1I)
- with is_class_supclass[OF wf _ cls'] have D: "C' \<noteq> Object \<Longrightarrow> is_class P ?D" by simp
-
- from i have "\<forall>sfs. sh C' \<noteq> Some(sfs,Processing)" by simp
- with conf_clinit_nProc_dist[OF confc\<^sub>0] have
- dist': "distinct (C' # clinit_classes ((stk, loc, C, M, pc, Calling C' Cs) # frs))" by fast
- then have dist'': "distinct (C' # clinit_classes frs)" by simp
-
- from conf_clinit_shupd_Calling[OF confc\<^sub>0 dist' cls']
- conf_clinit_shupd_Called[OF confc\<^sub>0 dist' cls']
- have confc': "conf_clinit P ?sh (?f ?ics#frs)" by clarsimp
- with last ics have "\<exists>sobj. ?sh (last(C'#Cs)) = Some sobj"
- by(auto simp: conf_clinit_def fun_upd_apply)
- with D vics ics have vics': "P,h,?sh \<turnstile>\<^sub>i (C, M, pc, ?ics)" by auto
-
- have fs': "conf_fs P h ?sh \<Phi> C M (size Ts) T frs" by(rule conf_fs_shupd[OF fs dist''])
-
- from frame vics' have frame': "conf_f P h ?sh (ST, LT) ins (?f ?ics)" by simp
-
- from i s' ics have "\<sigma>' = (None, h, ?f ?ics#frs, ?sh)" by(auto simp: if_split_asm)
-
- with mC h_ok sh_ok' \<Phi> stk loc pc fs' confc' frame' ics
- have ?thesis by fastforce
- }
- ultimately have ?thesis by(cases i, auto)
- }
- ultimately show ?thesis by(cases "sh C'", auto)
-qed
-
-lemma Throwing_correct:
- fixes \<sigma>' :: jvm_state
- assumes wtprog: "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P"
- assumes mC: "P \<turnstile> C sees M,b:Ts\<rightarrow>T=(mxs,mxl\<^sub>0,ins,xt) in C"
- assumes s': "Some \<sigma>' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)"
- assumes cf: "P,\<Phi> \<turnstile> (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\<surd>"
- assumes xc: "fst (exec_step P h stk loc C M pc ics frs sh) = None"
- assumes ics: "ics = Throwing (C'#Cs) a"
-
- shows "P,\<Phi> \<turnstile> \<sigma>'\<surd>"
-proof -
- from wtprog obtain wfmb where wf: "wf_prog wfmb P"
- by (simp add: wf_jvm_prog_phi_def)
-
- from mC cf obtain ST LT where
- h_ok: "P \<turnstile> h \<surd>" and
- sh_ok: "P,h \<turnstile>\<^sub>s sh \<surd>" and
- \<Phi>: "\<Phi> C M ! pc = Some (ST,LT)" and
- stk: "P,h \<turnstile> stk [:\<le>] ST" and loc: "P,h \<turnstile> loc [:\<le>\<^sub>\<top>] LT" and
- pc: "pc < size ins" and
- frame: "conf_f P h sh (ST, LT) ins (stk,loc,C,M,pc,ics)" and
- fs: "conf_fs P h sh \<Phi> C M (size Ts) T frs" and
- confc: "conf_clinit P sh ((stk,loc,C,M,pc,ics)#frs)" and
- vics: "P,h,sh \<turnstile>\<^sub>i (C,M,pc,ics)"
- by (fastforce dest: sees_method_fun)
-
- with ics have confc\<^sub>0: "conf_clinit P sh ((stk,loc,C,M,pc,Throwing (C'#Cs) a)#frs)" by simp
-
- from frame ics mC have
- cc: "\<exists>C1. Called_context P C1 (ins ! pc)" by(clarsimp simp: conf_f_def2)
-
- from frame ics obtain obj where ha: "h a = Some obj" by(auto simp: conf_f_def2)
-
- from confc ics obtain sfs i where shC': "sh C' = Some(sfs,i)" by(clarsimp simp: conf_clinit_def)
- then have sfs: "P,h,C' \<turnstile>\<^sub>s sfs \<surd>" by(rule shconfD[OF sh_ok])
-
- from s' ics
- have \<sigma>': "\<sigma>' = (None, h, (stk,loc,C,M,pc,Throwing Cs a)#frs, sh(C' \<mapsto> (fst(the(sh C')), Error)))"
- (is "\<sigma>' = (None, h, ?f'#frs, ?sh')")
- by simp
-
- from confc ics have dist: "distinct (C' # clinit_classes (?f' # frs))"
- by (simp add: conf_clinit_def distinct_clinit_def)
- then have dist': "distinct (C' # clinit_classes frs)" by simp
-
- from conf_clinit_Throwing confc ics have confc': "conf_clinit P sh (?f' # frs)" by simp
-
- from shconf_upd_obj[OF sh_ok sfs] shC' have "P,h \<turnstile>\<^sub>s ?sh' \<surd>" by simp
- moreover
- have "conf_fs P h ?sh' \<Phi> C M (length Ts) T frs" by(rule conf_fs_shupd[OF fs dist'])
- moreover
- have "conf_clinit P ?sh' (?f' # frs)" by(rule conf_clinit_shupd[OF confc' dist])
- moreover note \<sigma>' h_ok mC \<Phi> pc stk loc ha cc
- ultimately show "P,\<Phi> \<turnstile> \<sigma>' \<surd>" by fastforce
-qed
-
-lemma Called_correct:
- fixes \<sigma>' :: jvm_state
- assumes wtprog: "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P"
- assumes mC: "P \<turnstile> C sees M,b:Ts\<rightarrow>T=(mxs,mxl\<^sub>0,ins,xt) in C"
- assumes s': "Some \<sigma>' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)"
- assumes cf: "P,\<Phi> \<turnstile> (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\<surd>"
- assumes xc: "fst (exec_step P h stk loc C M pc ics frs sh) = None"
- assumes ics[simp]: "ics = Called (C'#Cs)"
-
- shows "P,\<Phi> \<turnstile> \<sigma>'\<surd>"
-proof -
- from wtprog obtain wfmb where wf: "wf_prog wfmb P"
- by (simp add: wf_jvm_prog_phi_def)
-
- from mC cf obtain ST LT where
- h_ok: "P \<turnstile> h \<surd>" and
- sh_ok: "P,h \<turnstile>\<^sub>s sh \<surd>" and
- \<Phi>: "\<Phi> C M ! pc = Some (ST,LT)" and
- stk: "P,h \<turnstile> stk [:\<le>] ST" and loc: "P,h \<turnstile> loc [:\<le>\<^sub>\<top>] LT" and
- pc: "pc < size ins" and
- frame: "conf_f P h sh (ST, LT) ins (stk,loc,C,M,pc,ics)" and
- fs: "conf_fs P h sh \<Phi> C M (size Ts) T frs" and
- confc: "conf_clinit P sh ((stk,loc,C,M,pc,ics)#frs)" and
- vics: "P,h,sh \<turnstile>\<^sub>i (C,M,pc,ics)"
- by (fastforce dest: sees_method_fun)
-
- then have confc\<^sub>0: "conf_clinit P sh ((stk,loc,C,M,pc,Called (C'#Cs))#frs)" by simp
-
- from frame mC obtain C1 sobj where
- ss: "Called_context P C1 (ins ! pc)" and
- shC1: "sh C1 = Some sobj" by(clarsimp simp: conf_f_def2)
-
- from confc wf_sees_clinit[OF wf] obtain mxs' mxl' ins' xt' where
- clinit: "P \<turnstile> C' sees clinit,Static: [] \<rightarrow> Void=(mxs',mxl',ins',xt') in C'"
- by(fastforce simp: conf_clinit_def is_class_def)
-
- let ?loc' = "replicate mxl' undefined"
-
- from s' clinit
- have \<sigma>': "\<sigma>' = (None, h, ([],?loc',C',clinit,0,No_ics)#(stk,loc,C,M,pc,Called Cs)#frs, sh)"
- (is "\<sigma>' = (None, h, ?if#?f'#frs, sh)")
- by simp
-
- with wtprog clinit
- obtain start: "wt_start P C' Static [] mxl' (\<Phi> C' clinit)" and ins': "ins' \<noteq> []"
- by (auto dest: wt_jvm_prog_impl_wt_start)
- then obtain LT\<^sub>0 where LT\<^sub>0: "\<Phi> C' clinit ! 0 = Some ([], LT\<^sub>0)"
- by (clarsimp simp: wt_start_def defs1 sup_state_opt_any_Some split: staticb.splits)
- moreover
- have "conf_f P h sh ([], LT\<^sub>0) ins' ?if"
- proof -
- let ?LT = "replicate mxl' Err"
- have "P,h \<turnstile> ?loc' [:\<le>\<^sub>\<top>] ?LT" by simp
- also from start LT\<^sub>0 have "P \<turnstile> \<dots> [\<le>\<^sub>\<top>] LT\<^sub>0" by (simp add: wt_start_def)
- finally have "P,h \<turnstile> ?loc' [:\<le>\<^sub>\<top>] LT\<^sub>0" .
- thus ?thesis using ins' by simp
- qed
- moreover
- from conf_clinit_Called confc clinit have "conf_clinit P sh (?if # ?f' # frs)" by simp
- moreover note \<sigma>' h_ok sh_ok mC \<Phi> pc stk loc clinit ss shC1 fs
- ultimately show "P,\<Phi> \<turnstile> \<sigma>' \<surd>" by fastforce
-qed
-
-subsection \<open> Single Instructions \<close>
-
-text \<open>
- In this section we prove for each single (welltyped) instruction
- that the state after execution of the instruction still conforms.
- Since we have already handled exceptions above, we can now assume that
- no exception occurs in this step. For instructions that may call
- the initialization procedure, we cover the calling and non-calling
- cases separately.
-\<close>
-
-lemma Invoke_correct:
- fixes \<sigma>' :: jvm_state
- assumes wtprog: "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P"
- assumes meth_C: "P \<turnstile> C sees M,b:Ts\<rightarrow>T=(mxs,mxl\<^sub>0,ins,xt) in C"
- assumes ins: "ins ! pc = Invoke M' n"
- assumes wti: "P,T,mxs,size ins,xt \<turnstile> ins!pc,pc :: \<Phi> C M"
- assumes \<sigma>': "Some \<sigma>' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)"
- assumes approx: "P,\<Phi> \<turnstile> (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\<surd>"
- assumes no_xcp: "fst (exec_step P h stk loc C M pc ics frs sh) = None"
- shows "P,\<Phi> \<turnstile> \<sigma>'\<surd>"
-(*<*)
-proof -
- from meth_C approx ins have [simp]: "ics = No_ics" by(cases ics, auto)
-
- note split_paired_Ex [simp del]
-
- from wtprog obtain wfmb where wfprog: "wf_prog wfmb P"
- by (simp add: wf_jvm_prog_phi_def)
-
- from ins meth_C approx obtain ST LT where
- heap_ok: "P\<turnstile> h\<surd>" and
- \<Phi>_pc: "\<Phi> C M!pc = Some (ST,LT)" and
- frame: "conf_f P h sh (ST,LT) ins (stk,loc,C,M,pc,ics)" and
- frames: "conf_fs P h sh \<Phi> C M (size Ts) T frs" and
- confc: "conf_clinit P sh ((stk,loc,C,M,pc,ics)#frs)"
- by (fastforce dest: sees_method_fun)
-
- from ins wti \<Phi>_pc
- have n: "n < size ST" by simp
-
- { assume "stk!n = Null"
- with ins no_xcp meth_C have False by (simp add: split_beta)
- hence ?thesis ..
- }
- moreover
- { assume "ST!n = NT"
- moreover
- from frame have "P,h \<turnstile> stk [:\<le>] ST" by simp
- with n have "P,h \<turnstile> stk!n :\<le> ST!n" by (simp add: list_all2_conv_all_nth)
- ultimately
- have "stk!n = Null" by simp
- with ins no_xcp meth_C have False by (simp add: split_beta)
- hence ?thesis ..
- }
- moreover {
- assume NT: "ST!n \<noteq> NT" and Null: "stk!n \<noteq> Null"
-
- from NT ins wti \<Phi>_pc obtain D D' b Ts T m ST' LT' where
- D: "ST!n = Class D" and
- pc': "pc+1 < size ins" and
- m_D: "P \<turnstile> D sees M',b: Ts\<rightarrow>T = m in D'" and
- Ts: "P \<turnstile> rev (take n ST) [\<le>] Ts" and
- \<Phi>': "\<Phi> C M ! (pc+1) = Some (ST', LT')" and
- LT': "P \<turnstile> LT [\<le>\<^sub>\<top>] LT'" and
- ST': "P \<turnstile> (T # drop (n+1) ST) [\<le>] ST'" and
- b[simp]: "b = NonStatic"
- by (clarsimp simp: sup_state_opt_any_Some)
-
- from frame obtain
- stk: "P,h \<turnstile> stk [:\<le>] ST" and
- loc: "P,h \<turnstile> loc [:\<le>\<^sub>\<top>] LT" by simp
-
- from n stk D have "P,h \<turnstile> stk!n :\<le> Class D"
- by (auto simp: list_all2_conv_all_nth)
- with Null obtain a C' fs where
- Addr: "stk!n = Addr a" and
- obj: "h a = Some (C',fs)" and
- C'subD: "P \<turnstile> C' \<preceq>\<^sup>* D"
- by (fastforce dest!: conf_ClassD)
-
- with wfprog m_D no_xcp
- obtain Ts' T' D'' mxs' mxl' ins' xt' where
- m_C': "P \<turnstile> C' sees M',NonStatic: Ts'\<rightarrow>T' = (mxs',mxl',ins',xt') in D''" and
- T': "P \<turnstile> T' \<le> T" and
- Ts': "P \<turnstile> Ts [\<le>] Ts'"
- by (auto dest: sees_method_mono)
- with wf_NonStatic_nclinit wtprog have nclinit: "M' \<noteq> clinit" by(simp add: wf_jvm_prog_phi_def)
-
- have D''subD': "P \<turnstile> D'' \<preceq>\<^sup>* D'" by(rule sees_method_decl_mono[OF C'subD m_D m_C'])
-
- let ?loc' = "Addr a # rev (take n stk) @ replicate mxl' undefined"
- let ?f' = "([], ?loc', D'', M', 0, No_ics)"
- let ?f = "(stk, loc, C, M, pc, ics)"
-
- from Addr obj m_C' ins \<sigma>' meth_C no_xcp
- have s': "\<sigma>' = (None, h, ?f' # ?f # frs, sh)" by simp
-
- from Ts n have [simp]: "size Ts = n"
- by (auto dest: list_all2_lengthD simp: min_def)
- with Ts' have [simp]: "size Ts' = n"
- by (auto dest: list_all2_lengthD)
-
- from m_C' wfprog
- obtain mD'': "P \<turnstile> D'' sees M',NonStatic:Ts'\<rightarrow>T'=(mxs',mxl',ins',xt') in D''"
- by (fast dest: sees_method_idemp)
- moreover
- with wtprog
- obtain start: "wt_start P D'' NonStatic Ts' mxl' (\<Phi> D'' M')" and ins': "ins' \<noteq> []"
- by (auto dest: wt_jvm_prog_impl_wt_start)
- then obtain LT\<^sub>0 where LT\<^sub>0: "\<Phi> D'' M' ! 0 = Some ([], LT\<^sub>0)"
- by (clarsimp simp: wt_start_def defs1 sup_state_opt_any_Some split: staticb.splits)
- moreover
- have "conf_f P h sh ([], LT\<^sub>0) ins' ?f'"
- proof -
- let ?LT = "OK (Class D'') # (map OK Ts') @ (replicate mxl' Err)"
-
- from stk have "P,h \<turnstile> take n stk [:\<le>] take n ST" ..
- hence "P,h \<turnstile> rev (take n stk) [:\<le>] rev (take n ST)" by simp
- also note Ts also note Ts' finally
- have "P,h \<turnstile> rev (take n stk) [:\<le>\<^sub>\<top>] map OK Ts'" by simp
- also
- have "P,h \<turnstile> replicate mxl' undefined [:\<le>\<^sub>\<top>] replicate mxl' Err"
- by simp
- also from m_C' have "P \<turnstile> C' \<preceq>\<^sup>* D''" by (rule sees_method_decl_above)
- with obj have "P,h \<turnstile> Addr a :\<le> Class D''" by (simp add: conf_def)
- ultimately
- have "P,h \<turnstile> ?loc' [:\<le>\<^sub>\<top>] ?LT" by simp
- also from start LT\<^sub>0 have "P \<turnstile> \<dots> [\<le>\<^sub>\<top>] LT\<^sub>0" by (simp add: wt_start_def)
- finally have "P,h \<turnstile> ?loc' [:\<le>\<^sub>\<top>] LT\<^sub>0" .
- thus ?thesis using ins' nclinit by simp
- qed
- moreover
- have "conf_clinit P sh (?f'#?f#frs)" using conf_clinit_Invoke[OF confc nclinit] by simp
- ultimately
- have ?thesis using s' \<Phi>_pc approx meth_C m_D T' ins D nclinit D''subD'
- by(fastforce dest: sees_method_fun [of _ C])
- }
- ultimately show ?thesis by blast
-qed
-(*>*)
-
-lemma Invokestatic_nInit_correct:
- fixes \<sigma>' :: jvm_state
- assumes wtprog: "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P"
- assumes meth_C: "P \<turnstile> C sees M,b:Ts\<rightarrow>T=(mxs,mxl\<^sub>0,ins,xt) in C"
- assumes ins: "ins ! pc = Invokestatic D M' n" and nclinit: "M' \<noteq> clinit"
- assumes wti: "P,T,mxs,size ins,xt \<turnstile> ins!pc,pc :: \<Phi> C M"
- assumes \<sigma>': "Some \<sigma>' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)"
- assumes approx: "P,\<Phi> \<turnstile> (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\<surd>"
- assumes no_xcp: "fst (exec_step P h stk loc C M pc ics frs sh) = None"
- assumes cs: "ics = Called [] \<or> (ics = No_ics \<and> (\<exists>sfs. sh (fst(method P D M')) = Some(sfs, Done)))"
- shows "P,\<Phi> \<turnstile> \<sigma>'\<surd>"
-(*<*)
-proof -
- note split_paired_Ex [simp del]
-
- from wtprog obtain wfmb where wfprog: "wf_prog wfmb P"
- by (simp add: wf_jvm_prog_phi_def)
-
- from ins meth_C approx obtain ST LT where
- heap_ok: "P\<turnstile> h\<surd>" and
- \<Phi>_pc: "\<Phi> C M!pc = Some (ST,LT)" and
- frame: "conf_f P h sh (ST,LT) ins (stk,loc,C,M,pc,ics)" and
- frames: "conf_fs P h sh \<Phi> C M (size Ts) T frs" and
- confc: "conf_clinit P sh ((stk,loc,C,M,pc,ics)#frs)"
- by (fastforce dest: sees_method_fun)
-
- from ins wti \<Phi>_pc have n: "n \<le> size ST" by simp
-
- from ins wti \<Phi>_pc obtain D' b Ts T mxs' mxl' ins' xt' ST' LT' where
- pc': "pc+1 < size ins" and
- m_D: "P \<turnstile> D sees M',b: Ts\<rightarrow>T = (mxs',mxl',ins',xt') in D'" and
- Ts: "P \<turnstile> rev (take n ST) [\<le>] Ts" and
- \<Phi>': "\<Phi> C M ! (pc+1) = Some (ST', LT')" and
- LT': "P \<turnstile> LT [\<le>\<^sub>\<top>] LT'" and
- ST': "P \<turnstile> (T # drop n ST) [\<le>] ST'" and
- b[simp]: "b = Static"
- by (clarsimp simp: sup_state_opt_any_Some)
-
- from frame obtain
- stk: "P,h \<turnstile> stk [:\<le>] ST" and
- loc: "P,h \<turnstile> loc [:\<le>\<^sub>\<top>] LT" by simp
-
- let ?loc' = "rev (take n stk) @ replicate mxl' undefined"
- let ?f' = "([], ?loc', D', M', 0, No_ics)"
- let ?f = "(stk, loc, C, M, pc, No_ics)"
-
- from m_D ins \<sigma>' meth_C no_xcp cs
- have s': "\<sigma>' = (None, h, ?f' # ?f # frs, sh)" by auto
-
- from Ts n have [simp]: "size Ts = n"
- by (auto dest: list_all2_lengthD)
-
- from m_D wfprog b
- obtain mD': "P \<turnstile> D' sees M',Static:Ts\<rightarrow>T=(mxs',mxl',ins',xt') in D'"
- by (fast dest: sees_method_idemp)
- moreover
- with wtprog
- obtain start: "wt_start P D' Static Ts mxl' (\<Phi> D' M')" and ins': "ins' \<noteq> []"
- by (auto dest: wt_jvm_prog_impl_wt_start)
- then obtain LT\<^sub>0 where LT\<^sub>0: "\<Phi> D' M' ! 0 = Some ([], LT\<^sub>0)"
- by (clarsimp simp: wt_start_def defs1 sup_state_opt_any_Some split: staticb.splits)
- moreover
- have "conf_f P h sh ([], LT\<^sub>0) ins' ?f'"
- proof -
- let ?LT = "(map OK Ts) @ (replicate mxl' Err)"
-
- from stk have "P,h \<turnstile> take n stk [:\<le>] take n ST" ..
- hence "P,h \<turnstile> rev (take n stk) [:\<le>] rev (take n ST)" by simp
- also note Ts finally
- have "P,h \<turnstile> rev (take n stk) [:\<le>\<^sub>\<top>] map OK Ts" by simp
- also
- have "P,h \<turnstile> replicate mxl' undefined [:\<le>\<^sub>\<top>] replicate mxl' Err"
- by simp
- also from m_D have "P \<turnstile> D \<preceq>\<^sup>* D'" by (rule sees_method_decl_above)
- ultimately
- have "P,h \<turnstile> ?loc' [:\<le>\<^sub>\<top>] ?LT" by simp
- also from start LT\<^sub>0 have "P \<turnstile> \<dots> [\<le>\<^sub>\<top>] LT\<^sub>0" by (simp add: wt_start_def)
- finally have "P,h \<turnstile> ?loc' [:\<le>\<^sub>\<top>] LT\<^sub>0" .
- thus ?thesis using ins' by simp
- qed
- moreover
- have "conf_clinit P sh (?f'#?f#frs)" by(rule conf_clinit_Invoke[OF confc nclinit])
- ultimately
- show ?thesis using s' \<Phi>_pc approx meth_C m_D ins nclinit
- by (fastforce dest: sees_method_fun [of _ C])
-qed
-(*>*)
-
-lemma Invokestatic_Init_correct:
- fixes \<sigma>' :: jvm_state
- assumes wtprog: "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P"
- assumes meth_C: "P \<turnstile> C sees M,b:Ts\<rightarrow>T=(mxs,mxl\<^sub>0,ins,xt) in C"
- assumes ins: "ins ! pc = Invokestatic D M' n" and nclinit: "M' \<noteq> clinit"
- assumes wti: "P,T,mxs,size ins,xt \<turnstile> ins!pc,pc :: \<Phi> C M"
- assumes \<sigma>': "Some \<sigma>' = exec (P, None, h, (stk,loc,C,M,pc,No_ics)#frs, sh)"
- assumes approx: "P,\<Phi> \<turnstile> (None, h, (stk,loc,C,M,pc,No_ics)#frs, sh)\<surd>"
- assumes no_xcp: "fst (exec_step P h stk loc C M pc No_ics frs sh) = None"
- assumes nDone: "\<forall>sfs. sh (fst(method P D M')) \<noteq> Some(sfs, Done)"
- shows "P,\<Phi> \<turnstile> \<sigma>'\<surd>"
-(*<*)
-proof -
- note split_paired_Ex [simp del]
-
- from wtprog obtain wfmb where wfprog: "wf_prog wfmb P"
- by (simp add: wf_jvm_prog_phi_def)
-
- from ins meth_C approx obtain ST LT where
- heap_ok: "P\<turnstile> h\<surd>" and
- \<Phi>_pc: "\<Phi> C M!pc = Some (ST,LT)" and
- stk: "P,h \<turnstile> stk [:\<le>] ST" and loc: "P,h \<turnstile> loc [:\<le>\<^sub>\<top>] LT" and
- pc: "pc < size ins" and
- frames: "conf_fs P h sh \<Phi> C M (size Ts) T frs" and
- confc: "conf_clinit P sh ((stk,loc,C,M,pc,No_ics)#frs)" and
- pc: "pc < size ins"
- by (fastforce dest: sees_method_fun)
-
- from ins wti \<Phi>_pc obtain D' b Ts T mxs' mxl' ins' xt' where
- m_D: "P \<turnstile> D sees M',b: Ts\<rightarrow>T = (mxs',mxl',ins',xt') in D'" and
- b[simp]: "b = Static"
- by clarsimp
-
- let ?f = "(stk, loc, C, M, pc, Calling D' [])"
-
- from m_D ins \<sigma>' meth_C no_xcp nDone
- have s': "\<sigma>' = (None, h, ?f # frs, sh)" by(auto split: init_state.splits)
-
- have cls: "is_class P D'" by(rule sees_method_is_class'[OF m_D])
-
- from confc have confc': "conf_clinit P sh (?f#frs)"
- by(auto simp: conf_clinit_def distinct_clinit_def split: if_split_asm)
- with s' \<Phi>_pc approx meth_C m_D ins nclinit stk loc pc cls frames
- show ?thesis by(fastforce dest: sees_method_fun [of _ C])
-qed
-(*>*)
-
-declare list_all2_Cons2 [iff]
-
-lemma Return_correct:
- fixes \<sigma>' :: jvm_state
- assumes wt_prog: "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P"
- assumes meth: "P \<turnstile> C sees M,b:Ts\<rightarrow>T=(mxs,mxl\<^sub>0,ins,xt) in C"
- assumes ins: "ins ! pc = Return"
- assumes wt: "P,T,mxs,size ins,xt \<turnstile> ins!pc,pc :: \<Phi> C M"
- assumes s': "Some \<sigma>' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)"
- assumes correct: "P,\<Phi> \<turnstile> (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\<surd>"
-
- shows "P,\<Phi> \<turnstile> \<sigma>'\<surd>"
-(*<*)
-proof -
- from meth correct ins have [simp]: "ics = No_ics" by(cases ics, auto)
-
- from wt_prog
- obtain wfmb where wf: "wf_prog wfmb P" by (simp add: wf_jvm_prog_phi_def)
-
- from meth ins s'
- have "frs = [] \<Longrightarrow> ?thesis" by (simp add: correct_state_def)
- moreover
- { fix f frs' assume frs': "frs = f#frs'"
- then obtain stk' loc' C' M' pc' ics' where
- f: "f = (stk',loc',C',M',pc',ics')" by (cases f)
-
- from correct meth
- obtain ST LT where
- h_ok: "P \<turnstile> h \<surd>" and
- sh_ok: "P,h \<turnstile>\<^sub>s sh \<surd>" and
- \<Phi>_pc: "\<Phi> C M ! pc = Some (ST, LT)" and
- frame: "conf_f P h sh (ST, LT) ins (stk,loc,C,M,pc,ics)" and
- frames: "conf_fs P h sh \<Phi> C M (size Ts) T frs" and
- confc: "conf_clinit P sh frs"
- by (auto dest: sees_method_fun conf_clinit_Cons simp: correct_state_def)
-
- from \<Phi>_pc ins wt
- obtain U ST\<^sub>0 where "ST = U # ST\<^sub>0" "P \<turnstile> U \<le> T"
- by (simp add: wt_instr_def app_def) blast
- with wf frame
- have hd_stk: "P,h \<turnstile> hd stk :\<le> T" by (auto simp: conf_f_def)
-
- from f frs' frames meth
- obtain ST' LT' b' Ts'' T'' mxs' mxl\<^sub>0' ins' xt' where
- \<Phi>': "\<Phi> C' M' ! pc' = Some (ST', LT')" and
- meth_C': "P \<turnstile> C' sees M',b':Ts''\<rightarrow>T''=(mxs',mxl\<^sub>0',ins',xt') in C'" and
- frame': "conf_f P h sh (ST',LT') ins' f" and
- conf_fs: "conf_fs P h sh \<Phi> C' M' (size Ts'') T'' frs'"
- by clarsimp
-
- from f frame' obtain
- stk': "P,h \<turnstile> stk' [:\<le>] ST'" and
- loc': "P,h \<turnstile> loc' [:\<le>\<^sub>\<top>] LT'" and
- pc': "pc' < size ins'"
- by (simp add: conf_f_def)
-
- { assume b[simp]: "b = NonStatic"
-
- from wf_NonStatic_nclinit[OF wf] meth have nclinit[simp]: "M \<noteq> clinit" by simp
-
- from f frs' meth ins s'
- have \<sigma>':
- "\<sigma>' = (None,h,(hd stk#(drop (1+size Ts) stk'),loc',C',M',pc'+1,ics')#frs',sh)"
- (is "\<sigma>' = (None,h,?f'#frs',sh)")
- by simp
- from f frs' confc conf_clinit_diff have confc'': "conf_clinit P sh (?f'#frs')" by blast
-
- with \<Phi>' meth_C' f frs' frames meth
- obtain D Ts' T' m D' where
- ins': "ins' ! pc' = Invoke M (size Ts)" and
- D: "ST' ! (size Ts) = Class D" and
- meth_D: "P \<turnstile> D sees M,b: Ts'\<rightarrow>T' = m in D'" and
- T': "P \<turnstile> T \<le> T'" and
- CsubD': "P \<turnstile> C \<preceq>\<^sup>* D'"
- by(auto dest: sees_method_fun sees_method_fun[OF sees_method_idemp])
-
- from wt_prog meth_C' pc'
- have "P,T'',mxs',size ins',xt' \<turnstile> ins'!pc',pc' :: \<Phi> C' M'"
- by (rule wt_jvm_prog_impl_wt_instr)
- with ins' \<Phi>' D meth_D
- obtain ST'' LT'' where
- \<Phi>_suc: "\<Phi> C' M' ! Suc pc' = Some (ST'', LT'')" and
- less: "P \<turnstile> (T' # drop (size Ts+1) ST', LT') \<le>\<^sub>i (ST'', LT'')" and
- suc_pc': "Suc pc' < size ins'"
- by (clarsimp simp: sup_state_opt_any_Some)
-
- from hd_stk T' have hd_stk': "P,h \<turnstile> hd stk :\<le> T'" ..
-
- have frame'':
- "conf_f P h sh (ST'',LT'') ins' ?f'"
- proof -
- from stk'
- have "P,h \<turnstile> drop (1+size Ts) stk' [:\<le>] drop (1+size Ts) ST'" ..
- moreover
- with hd_stk' less
- have "P,h \<turnstile> hd stk # drop (1+size Ts) stk' [:\<le>] ST''" by auto
- moreover
- from wf loc' less have "P,h \<turnstile> loc' [:\<le>\<^sub>\<top>] LT''" by auto
- moreover note suc_pc'
- moreover
- from f frs' frames (* ics' = No_ics *)
- have "P,h,sh \<turnstile>\<^sub>i (C', M', Suc pc', ics')" by auto
- ultimately show ?thesis by (simp add: conf_f_def)
- qed
-
- with \<sigma>' frs' f meth h_ok sh_ok hd_stk \<Phi>_suc frames confc'' meth_C' \<Phi>'
- have ?thesis by(fastforce dest: sees_method_fun [of _ C'])
- }
- moreover
- { assume b[simp]: "b = Static" and nclinit[simp]: "M \<noteq> clinit"
-
- from f frs' meth ins s'
- have \<sigma>':
- "\<sigma>' = (None,h,(hd stk#(drop (size Ts) stk'),loc',C',M',pc'+1,ics')#frs',sh)"
- (is "\<sigma>' = (None,h,?f'#frs',sh)")
- by simp
- from f frs' confc conf_clinit_diff have confc'': "conf_clinit P sh (?f'#frs')" by blast
-
- with \<Phi>' meth_C' f frs' frames meth
- obtain D Ts' T' m where
- ins': "ins' ! pc' = Invokestatic D M (size Ts)" and
- meth_D: "P \<turnstile> D sees M,b: Ts'\<rightarrow>T' = m in C" and
- T': "P \<turnstile> T \<le> T'"
- by(auto dest: sees_method_fun sees_method_mono2[OF _ wf sees_method_idemp])
-
- from wt_prog meth_C' pc'
- have "P,T'',mxs',size ins',xt' \<turnstile> ins'!pc',pc' :: \<Phi> C' M'"
- by (rule wt_jvm_prog_impl_wt_instr)
- with ins' \<Phi>' meth_D
- obtain ST'' LT'' where
- \<Phi>_suc: "\<Phi> C' M' ! Suc pc' = Some (ST'', LT'')" and
- less: "P \<turnstile> (T' # drop (size Ts) ST', LT') \<le>\<^sub>i (ST'', LT'')" and
- suc_pc': "Suc pc' < size ins'"
- by (clarsimp simp: sup_state_opt_any_Some)
-
- from hd_stk T' have hd_stk': "P,h \<turnstile> hd stk :\<le> T'" ..
-
- have frame'':
- "conf_f P h sh (ST'',LT'') ins' ?f'"
- proof -
- from stk'
- have "P,h \<turnstile> drop (size Ts) stk' [:\<le>] drop (size Ts) ST'" ..
- moreover
- with hd_stk' less
- have "P,h \<turnstile> hd stk # drop (size Ts) stk' [:\<le>] ST''" by auto
- moreover
- from wf loc' less have "P,h \<turnstile> loc' [:\<le>\<^sub>\<top>] LT''" by auto
- moreover note suc_pc'
- moreover
- from f frs' frames (* ics' = No_ics *)
- have "P,h,sh \<turnstile>\<^sub>i (C', M', Suc pc', ics')" by auto
- ultimately show ?thesis by (simp add: conf_f_def)
- qed
-
- with \<sigma>' frs' f meth h_ok sh_ok hd_stk \<Phi>_suc frames confc'' meth_C' \<Phi>'
- have ?thesis by(fastforce dest: sees_method_fun [of _ C'])
- }
- moreover
- { assume b[simp]: "b = Static" and clinit[simp]: "M = clinit"
-
- from frs' meth ins s'
- have \<sigma>':
- "\<sigma>' = (None,h,frs,sh(C\<mapsto>(fst(the(sh C)), Done)))" (is "\<sigma>' = (None,h,frs,?sh)")
- by simp
-
- from correct have dist': "distinct (C # clinit_classes frs)"
- by(simp add: conf_clinit_def distinct_clinit_def)
-
- from f frs' correct have confc1:
- "conf_clinit P sh ((stk, loc, C, clinit, pc, No_ics) # (stk',loc',C',M',pc',ics') # frs')"
- by simp
- then have ics_dist: "distinct (C # ics_classes ics')"
- by(simp add: conf_clinit_def distinct_clinit_def)
-
- from conf_clinit_Cons_Cons[OF confc1] have dist'': "distinct (C # clinit_classes frs')"
- by(simp add: conf_clinit_def distinct_clinit_def)
-
- from correct shconf_upd_obj[OF sh_ok _ [OF shconfD[OF sh_ok]]]
- have sh'_ok: "P,h \<turnstile>\<^sub>s ?sh \<surd>" by(clarsimp simp: conf_clinit_def)
-
- have frame'':
- "conf_f P h ?sh (ST',LT') ins' f"
- proof -
- note stk' loc' pc' f valid_ics_shupd[OF _ ics_dist]
- moreover
- from f frs' frames
- have "P,h,sh \<turnstile>\<^sub>i (C', M', pc', ics')" by auto
- ultimately show ?thesis by (simp add: conf_f_def2)
- qed
- have conf_fs': "conf_fs P h ?sh \<Phi> C' M' (length Ts'') T'' frs'"
- by(rule conf_fs_shupd[OF conf_fs dist''])
-
- have confc'': "conf_clinit P ?sh frs" by(rule conf_clinit_shupd[OF confc dist'])
-
- from \<sigma>' f frs' h_ok sh'_ok conf_fs' frame'' \<Phi>' stk' loc' pc' meth_C' confc''
- have ?thesis by(fastforce dest: sees_method_fun)
- }
- ultimately have ?thesis by (cases b) blast+
- }
- ultimately
- show ?thesis by (cases frs) blast+
-qed
-(*>*)
-
-declare sup_state_opt_any_Some [iff]
-declare not_Err_eq [iff]
-
-lemma Load_correct:
-assumes "wf_prog wt P" and
- mC: "P \<turnstile> C sees M,b:Ts\<rightarrow>T=(mxs,mxl\<^sub>0,ins,xt) in C" and
- i: "ins!pc = Load idx" and
- "P,T,mxs,size ins,xt \<turnstile> ins!pc,pc :: \<Phi> C M" and
- "Some \<sigma>' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)" and
- cf: "P,\<Phi> \<turnstile> (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\<surd>"
-shows "P,\<Phi> |- \<sigma>' [ok]"
-(*<*)
-proof -
- have "ics = No_ics" using mC i cf by(cases ics) auto
- then show ?thesis using assms(1,3-6) sees_method_fun[OF mC]
- by(fastforce elim!: confTs_confT_sup conf_clinit_diff)
-qed
-(*>*)
-
-declare [[simproc del: list_to_set_comprehension]]
-
-lemma Store_correct:
-assumes "wf_prog wt P" and
- mC: "P \<turnstile> C sees M,b:Ts\<rightarrow>T=(mxs,mxl\<^sub>0,ins,xt) in C" and
- i: "ins!pc = Store idx" and
- "P,T,mxs,size ins,xt \<turnstile> ins!pc,pc :: \<Phi> C M" and
- "Some \<sigma>' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)" and
- cf: "P,\<Phi> \<turnstile> (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\<surd>"
-shows "P,\<Phi> |- \<sigma>' [ok]"
-(*<*)
-proof -
- have "ics = No_ics" using mC i cf by(cases ics) auto
- then show ?thesis using assms(1,3-6) sees_method_fun[OF mC]
- by clarsimp (blast intro!: list_all2_update_cong conf_clinit_diff)
-qed
-(*>*)
-
-
-lemma Push_correct:
-assumes "wf_prog wt P" and
- mC: "P \<turnstile> C sees M,b:Ts\<rightarrow>T=(mxs,mxl\<^sub>0,ins,xt) in C" and
- i: "ins!pc = Push v" and
- "P,T,mxs,size ins,xt \<turnstile> ins!pc,pc :: \<Phi> C M" and
- "Some \<sigma>' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)" and
- cf: "P,\<Phi> \<turnstile> (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\<surd>"
-shows "P,\<Phi> |- \<sigma>' [ok]"
-(*<*)
-proof -
- have "ics = No_ics" using mC i cf by(cases ics) auto
- then show ?thesis using assms(1,3-6) sees_method_fun[OF mC]
- by clarsimp (blast dest: typeof_lit_conf conf_clinit_diff)
-qed
-
-
-lemma Cast_conf2:
- "\<lbrakk> wf_prog ok P; P,h \<turnstile> v :\<le> T; is_refT T; cast_ok P C h v;
- P \<turnstile> Class C \<le> T'; is_class P C\<rbrakk>
- \<Longrightarrow> P,h \<turnstile> v :\<le> T'"
-(*<*)
-proof -
- assume "wf_prog ok P" and "P,h \<turnstile> v :\<le> T" and "is_refT T" and
- "cast_ok P C h v" and wid: "P \<turnstile> Class C \<le> T'" and "is_class P C"
- then show "P,h \<turnstile> v :\<le> T'" using Class_widen[OF wid]
- by(cases v)
- (auto simp: cast_ok_def is_refT_def conf_def obj_ty_def
- intro: rtrancl_trans)
-qed
-(*>*)
-
-
-lemma Checkcast_correct:
-assumes "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P" and
- mC: "P \<turnstile> C sees M,b:Ts\<rightarrow>T=(mxs,mxl\<^sub>0,ins,xt) in C" and
- i: "ins!pc = Checkcast D" and
- "P,T,mxs,size ins,xt \<turnstile> ins!pc,pc :: \<Phi> C M" and
- "Some \<sigma>' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)" and
- cf: "P,\<Phi> \<turnstile> (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\<surd>" and
- "fst (exec_step P h stk loc C M pc ics frs sh) = None"
-shows "P,\<Phi> |- \<sigma>' [ok]"
-(*<*)
-proof -
- have "ics = No_ics" using mC i cf by(cases ics) auto
- then show ?thesis using assms
- by (clarsimp simp: wf_jvm_prog_phi_def split: if_split_asm)
- (blast intro: Cast_conf2 dest: sees_method_fun conf_clinit_diff)
-qed
-(*>*)
-
-declare split_paired_All [simp del]
-
-lemmas widens_Cons [iff] = list_all2_Cons1 [of "widen P"] for P
-
-lemma Getfield_correct:
- fixes \<sigma>' :: jvm_state
- assumes wf: "wf_prog wt P"
- assumes mC: "P \<turnstile> C sees M,b:Ts\<rightarrow>T=(mxs,mxl\<^sub>0,ins,xt) in C"
- assumes i: "ins!pc = Getfield F D"
- assumes wt: "P,T,mxs,size ins,xt \<turnstile> ins!pc,pc :: \<Phi> C M"
- assumes s': "Some \<sigma>' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)"
- assumes cf: "P,\<Phi> \<turnstile> (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\<surd>"
- assumes xc: "fst (exec_step P h stk loc C M pc ics frs sh) = None"
-
- shows "P,\<Phi> \<turnstile> \<sigma>'\<surd>"
-(*<*)
-proof -
- from mC cf i have [simp]: "ics = No_ics" by(cases ics, auto)
-
- from mC cf obtain ST LT where
- "h\<surd>": "P \<turnstile> h \<surd>" and
- "sh\<surd>": "P,h \<turnstile>\<^sub>s sh \<surd>" and
- \<Phi>: "\<Phi> C M ! pc = Some (ST,LT)" and
- stk: "P,h \<turnstile> stk [:\<le>] ST" and loc: "P,h \<turnstile> loc [:\<le>\<^sub>\<top>] LT" and
- pc: "pc < size ins" and
- fs: "conf_fs P h sh \<Phi> C M (size Ts) T frs" and
- confc: "conf_clinit P sh ((stk,loc,C,M,pc,ics)#frs)"
- by (fastforce dest: sees_method_fun)
-
- from i \<Phi> wt obtain oT ST'' vT ST' LT' vT' where
- oT: "P \<turnstile> oT \<le> Class D" and
- ST: "ST = oT # ST''" and
- F: "P \<turnstile> D sees F,NonStatic:vT in D" and
- pc': "pc+1 < size ins" and
- \<Phi>': "\<Phi> C M ! (pc+1) = Some (vT'#ST', LT')" and
- ST': "P \<turnstile> ST'' [\<le>] ST'" and LT': "P \<turnstile> LT [\<le>\<^sub>\<top>] LT'" and
- vT': "P \<turnstile> vT \<le> vT'"
- by fastforce
-
- from stk ST obtain ref stk' where
- stk': "stk = ref#stk'" and
- ref: "P,h \<turnstile> ref :\<le> oT" and
- ST'': "P,h \<turnstile> stk' [:\<le>] ST''"
- by auto
-
- from stk' i mC s' xc have "ref \<noteq> Null"
- by (simp add: split_beta split:if_split_asm)
- moreover from ref oT have "P,h \<turnstile> ref :\<le> Class D" ..
- ultimately obtain a D' fs where
- a: "ref = Addr a" and h: "h a = Some (D', fs)" and D': "P \<turnstile> D' \<preceq>\<^sup>* D"
- by (blast dest: non_npD)
-
- from D' F have has_field: "P \<turnstile> D' has F,NonStatic:vT in D"
- by (blast intro: has_field_mono has_visible_field)
- moreover from "h\<surd>" h have "P,h \<turnstile> (D', fs) \<surd>" by (rule hconfD)
- ultimately obtain v where v: "fs (F, D) = Some v" "P,h \<turnstile> v :\<le> vT"
- by (clarsimp simp: oconf_def has_field_def)
- (blast dest: has_fields_fun)
-
- from conf_clinit_diff[OF confc]
- have confc': "conf_clinit P sh ((v#stk',loc,C,M,pc+1,ics)#frs)" by simp
-
- from a h i mC s' stk' v xc has_field
- have "\<sigma>' = (None, h, (v#stk',loc,C,M,pc+1,ics)#frs, sh)"
- by(simp add: split_beta split: if_split_asm)
- moreover
- from ST'' ST' have "P,h \<turnstile> stk' [:\<le>] ST'" ..
- moreover
- from v vT' have "P,h \<turnstile> v :\<le> vT'" by blast
- moreover
- from loc LT' have "P,h \<turnstile> loc [:\<le>\<^sub>\<top>] LT'" ..
- moreover
- note "h\<surd>" "sh\<surd>" mC \<Phi>' pc' v fs confc'
- ultimately
- show "P,\<Phi> \<turnstile> \<sigma>' \<surd>" by fastforce
-qed
-(*>*)
-
-lemma Getstatic_nInit_correct:
- fixes \<sigma>' :: jvm_state
- assumes wf: "wf_prog wt P"
- assumes mC: "P \<turnstile> C sees M,b:Ts\<rightarrow>T=(mxs,mxl\<^sub>0,ins,xt) in C"
- assumes i: "ins!pc = Getstatic C' F D"
- assumes wt: "P,T,mxs,size ins,xt \<turnstile> ins!pc,pc :: \<Phi> C M"
- assumes s': "Some \<sigma>' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)"
- assumes cf: "P,\<Phi> \<turnstile> (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\<surd>"
- assumes xc: "fst (exec_step P h stk loc C M pc ics frs sh) = None"
- assumes cs: "ics = Called [] \<or> (ics = No_ics \<and> (\<exists>sfs. sh (fst(field P D F)) = Some(sfs, Done)))"
-
- shows "P,\<Phi> \<turnstile> \<sigma>'\<surd>"
-(*<*)
-proof -
- from mC cf obtain ST LT where
- "h\<surd>": "P \<turnstile> h \<surd>" and
- "sh\<surd>": "P,h \<turnstile>\<^sub>s sh \<surd>" and
- \<Phi>: "\<Phi> C M ! pc = Some (ST,LT)" and
- stk: "P,h \<turnstile> stk [:\<le>] ST" and loc: "P,h \<turnstile> loc [:\<le>\<^sub>\<top>] LT" and
- pc: "pc < size ins" and
- fs: "conf_fs P h sh \<Phi> C M (size Ts) T frs" and
- confc: "conf_clinit P sh ((stk,loc,C,M,pc,ics)#frs)" and
- vics: "P,h,sh \<turnstile>\<^sub>i (C,M,pc,ics)"
- by (fastforce dest: sees_method_fun)
-
- from i \<Phi> wt cs obtain vT ST' LT' vT' where
- F: "P \<turnstile> C' sees F,Static:vT in D" and
- pc': "pc+1 < size ins" and
- \<Phi>': "\<Phi> C M ! (pc+1) = Some (vT'#ST', LT')" and
- ST': "P \<turnstile> ST [\<le>] ST'" and LT': "P \<turnstile> LT [\<le>\<^sub>\<top>] LT'" and
- vT': "P \<turnstile> vT \<le> vT'"
- by fastforce
-
- with mC i vics obtain sobj where
- cc': "ics = Called [] \<Longrightarrow> Called_context P D (ins!pc) \<and> sh D = Some sobj"
- by(fastforce dest: has_visible_field)
-
- from field_def2[OF sees_field_idemp[OF F]] have D[simp]: "fst(field P D F) = D" by simp
- from cs cc' obtain sfs i where shD: "sh D = Some(sfs,i)" by(cases sobj, auto)
-
- note has_field_idemp[OF has_visible_field[OF F]]
- moreover from "sh\<surd>" shD have "P,h,D \<turnstile>\<^sub>s sfs \<surd>" by (rule shconfD)
- ultimately obtain v where v: "sfs F = Some v" "P,h \<turnstile> v :\<le> vT"
- by (clarsimp simp: soconf_def has_field_def) blast
-
- from i mC s' v xc F cs cc' shD
- have "\<sigma>' = (None, h, (v#stk,loc,C,M,pc+1,No_ics)#frs, sh)"
- by(fastforce simp: split_beta split: if_split_asm init_call_status.splits)
- moreover
- from stk ST' have "P,h \<turnstile> stk [:\<le>] ST'" ..
- moreover
- from v vT' have "P,h \<turnstile> v :\<le> vT'" by blast
- moreover
- from loc LT' have "P,h \<turnstile> loc [:\<le>\<^sub>\<top>] LT'" ..
- moreover
- have "conf_clinit P sh ((v#stk,loc,C,M,pc+1,No_ics)#frs)" by(rule conf_clinit_diff'[OF confc])
- moreover
- note "h\<surd>" "sh\<surd>" mC \<Phi>' pc' v fs
- ultimately
- show "P,\<Phi> \<turnstile> \<sigma>' \<surd>" by fastforce
-qed
-(*>*)
-
-lemma Getstatic_Init_correct:
- fixes \<sigma>' :: jvm_state
- assumes wf: "wf_prog wt P"
- assumes mC: "P \<turnstile> C sees M,b:Ts\<rightarrow>T=(mxs,mxl\<^sub>0,ins,xt) in C"
- assumes i: "ins!pc = Getstatic C' F D"
- assumes wt: "P,T,mxs,size ins,xt \<turnstile> ins!pc,pc :: \<Phi> C M"
- assumes s': "Some \<sigma>' = exec (P, None, h, (stk,loc,C,M,pc,No_ics)#frs, sh)"
- assumes cf: "P,\<Phi> \<turnstile> (None, h, (stk,loc,C,M,pc,No_ics)#frs, sh)\<surd>"
- assumes xc: "fst (exec_step P h stk loc C M pc No_ics frs sh) = None"
- assumes nDone: "\<forall>sfs. sh (fst(field P D F)) \<noteq> Some(sfs, Done)"
-
- shows "P,\<Phi> \<turnstile> \<sigma>'\<surd>"
-(*<*)
-proof -
- from mC cf obtain ST LT where
- "h\<surd>": "P \<turnstile> h \<surd>" and
- "sh\<surd>": "P,h \<turnstile>\<^sub>s sh \<surd>" and
- \<Phi>: "\<Phi> C M ! pc = Some (ST,LT)" and
- stk: "P,h \<turnstile> stk [:\<le>] ST" and loc: "P,h \<turnstile> loc [:\<le>\<^sub>\<top>] LT" and
- pc: "pc < size ins" and
- fs: "conf_fs P h sh \<Phi> C M (size Ts) T frs" and
- confc: "conf_clinit P sh ((stk,loc,C,M,pc,No_ics)#frs)"
- by (fastforce dest: sees_method_fun)
-
- from i \<Phi> wt nDone obtain vT where
- F: "P \<turnstile> C' sees F,Static:vT in D"
- by fastforce
- then have has_field: "P \<turnstile> C' has F,Static:vT in D" by(rule has_visible_field)
-
- from field_def2[OF sees_field_idemp[OF F]] has_field_is_class'[OF has_field] obtain
- D[simp]: "fst(field P D F) = D" and
- cls: "is_class P D" by simp
-
- from i mC s' xc F nDone
- have "\<sigma>' = (None, h, (stk,loc,C,M,pc,Calling D [])#frs, sh)"
- by(auto simp: split_beta split: if_split_asm init_state.splits)
- moreover
- from confc have "conf_clinit P sh ((stk,loc,C,M,pc,Calling D [])#frs)"
- by(auto simp: conf_clinit_def distinct_clinit_def split: if_split_asm)
- moreover
- note loc stk "h\<surd>" "sh\<surd>" mC \<Phi> pc fs i has_field cls
- ultimately
- show "P,\<Phi> \<turnstile> \<sigma>' \<surd>" by fastforce
-qed
-(*>*)
-
-lemma Putfield_correct:
- fixes \<sigma>' :: jvm_state
- assumes wf: "wf_prog wt P"
- assumes mC: "P \<turnstile> C sees M,b:Ts\<rightarrow>T=(mxs,mxl\<^sub>0,ins,xt) in C"
- assumes i: "ins!pc = Putfield F D"
- assumes wt: "P,T,mxs,size ins,xt \<turnstile> ins!pc,pc :: \<Phi> C M"
- assumes s': "Some \<sigma>' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)"
- assumes cf: "P,\<Phi> \<turnstile> (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\<surd>"
- assumes xc: "fst (exec_step P h stk loc C M pc ics frs sh) = None"
-
- shows "P,\<Phi> \<turnstile> \<sigma>'\<surd>"
-(*<*)
-proof -
- from mC cf i have [simp]: "ics = No_ics" by(cases ics, auto)
-
- from mC cf obtain ST LT where
- "h\<surd>": "P \<turnstile> h \<surd>" and
- "sh\<surd>": "P,h \<turnstile>\<^sub>s sh \<surd>" and
- \<Phi>: "\<Phi> C M ! pc = Some (ST,LT)" and
- stk: "P,h \<turnstile> stk [:\<le>] ST" and loc: "P,h \<turnstile> loc [:\<le>\<^sub>\<top>] LT" and
- pc: "pc < size ins" and
- fs: "conf_fs P h sh \<Phi> C M (size Ts) T frs" and
- confc: "conf_clinit P sh ((stk,loc,C,M,pc,ics) # frs)"
- by (fastforce dest: sees_method_fun)
-
- from i \<Phi> wt obtain vT vT' oT ST'' ST' LT' where
- ST: "ST = vT # oT # ST''" and
- field: "P \<turnstile> D sees F,NonStatic:vT' in D" and
- oT: "P \<turnstile> oT \<le> Class D" and vT: "P \<turnstile> vT \<le> vT'" and
- pc': "pc+1 < size ins" and
- \<Phi>': "\<Phi> C M!(pc+1) = Some (ST',LT')" and
- ST': "P \<turnstile> ST'' [\<le>] ST'" and LT': "P \<turnstile> LT [\<le>\<^sub>\<top>] LT'"
- by clarsimp
-
- from stk ST obtain v ref stk' where
- stk': "stk = v#ref#stk'" and
- v: "P,h \<turnstile> v :\<le> vT" and
- ref: "P,h \<turnstile> ref :\<le> oT" and
- ST'': "P,h \<turnstile> stk' [:\<le>] ST''"
- by auto
-
- from stk' i mC s' xc have "ref \<noteq> Null" by (auto simp: split_beta)
- moreover from ref oT have "P,h \<turnstile> ref :\<le> Class D" ..
- ultimately obtain a D' fs where
- a: "ref = Addr a" and h: "h a = Some (D', fs)" and D': "P \<turnstile> D' \<preceq>\<^sup>* D"
- by (blast dest: non_npD)
-
- from v vT have vT': "P,h \<turnstile> v :\<le> vT'" ..
-
- from field D' have has_field: "P \<turnstile> D' has F,NonStatic:vT' in D"
- by (blast intro: has_field_mono has_visible_field)
-
- let ?h' = "h(a\<mapsto>(D', fs((F, D)\<mapsto>v)))" and ?f' = "(stk',loc,C,M,pc+1,ics)"
- from h have hext: "h \<unlhd> ?h'" by (rule hext_upd_obj)
-
- have "sh\<surd>'": "P,?h' \<turnstile>\<^sub>s sh \<surd>" by(rule shconf_hupd_obj[OF "sh\<surd>" h])
-
- from a h i mC s' stk' has_field field
- have "\<sigma>' = (None, ?h', ?f'#frs, sh)" by(simp split: if_split_asm)
- moreover
- from "h\<surd>" h have "P,h \<turnstile> (D',fs)\<surd>" by (rule hconfD)
- with has_field vT' have "P,h \<turnstile> (D',fs((F, D)\<mapsto>v))\<surd>" ..
- with "h\<surd>" h have "P \<turnstile> ?h'\<surd>" by (rule hconf_upd_obj)
- moreover
- from ST'' ST' have "P,h \<turnstile> stk' [:\<le>] ST'" ..
- from this hext have "P,?h' \<turnstile> stk' [:\<le>] ST'" by (rule confs_hext)
- moreover
- from loc LT' have "P,h \<turnstile> loc [:\<le>\<^sub>\<top>] LT'" ..
- from this hext have "P,?h' \<turnstile> loc [:\<le>\<^sub>\<top>] LT'" by (rule confTs_hext)
- moreover
- from fs hext
- have "conf_fs P ?h' sh \<Phi> C M (size Ts) T frs" by (rule conf_fs_hext)
- moreover
- have "conf_clinit P sh (?f' # frs)" by(rule conf_clinit_diff[OF confc])
- moreover
- note mC \<Phi>' pc' "sh\<surd>'"
- ultimately
- show "P,\<Phi> \<turnstile> \<sigma>' \<surd>" by fastforce
-qed
-(*>*)
-
-lemma Putstatic_nInit_correct:
- fixes \<sigma>' :: jvm_state
- assumes wf: "wf_prog wt P"
- assumes mC: "P \<turnstile> C sees M,b:Ts\<rightarrow>T=(mxs,mxl\<^sub>0,ins,xt) in C"
- assumes i: "ins!pc = Putstatic C' F D"
- assumes wt: "P,T,mxs,size ins,xt \<turnstile> ins!pc,pc :: \<Phi> C M"
- assumes s': "Some \<sigma>' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)"
- assumes cf: "P,\<Phi> \<turnstile> (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\<surd>"
- assumes xc: "fst (exec_step P h stk loc C M pc ics frs sh) = None"
- assumes cs: "ics = Called [] \<or> (ics = No_ics \<and> (\<exists>sfs. sh (fst(field P D F)) = Some(sfs, Done)))"
-
- shows "P,\<Phi> \<turnstile> \<sigma>'\<surd>"
-(*<*)
-proof -
- from mC cf obtain ST LT where
- "h\<surd>": "P \<turnstile> h \<surd>" and
- "sh\<surd>": "P,h \<turnstile>\<^sub>s sh \<surd>" and
- \<Phi>: "\<Phi> C M ! pc = Some (ST,LT)" and
- stk: "P,h \<turnstile> stk [:\<le>] ST" and loc: "P,h \<turnstile> loc [:\<le>\<^sub>\<top>] LT" and
- pc: "pc < size ins" and
- fs: "conf_fs P h sh \<Phi> C M (size Ts) T frs" and
- confc: "conf_clinit P sh ((stk,loc,C,M,pc,ics)#frs)" and
- vics: "P,h,sh \<turnstile>\<^sub>i (C,M,pc,ics)"
- by (fastforce dest: sees_method_fun)
-
- from i \<Phi> wt cs obtain vT vT' ST'' ST' LT' where
- ST: "ST = vT # ST''" and
- F: "P \<turnstile> C' sees F,Static:vT' in D" and
- vT: "P \<turnstile> vT \<le> vT'" and
- pc': "pc+1 < size ins" and
- \<Phi>': "\<Phi> C M ! (pc+1) = Some (ST', LT')" and
- ST': "P \<turnstile> ST'' [\<le>] ST'" and LT': "P \<turnstile> LT [\<le>\<^sub>\<top>] LT'"
- by fastforce
-
- from stk ST obtain v stk' where
- stk': "stk = v#stk'" and
- v: "P,h \<turnstile> v :\<le> vT" and
- ST'': "P,h \<turnstile> stk' [:\<le>] ST''"
- by auto
-
- from v vT have vT': "P,h \<turnstile> v :\<le> vT'" ..
-
- with mC i vics obtain sobj where
- cc': "ics = Called [] \<Longrightarrow> Called_context P D (ins!pc) \<and> sh D = Some sobj"
- by(fastforce dest: has_visible_field)
-
- from field_def2[OF sees_field_idemp[OF F]] have D[simp]: "fst(field P D F) = D" by simp
- from cs cc' obtain sfs i where shD: "sh D = Some(sfs,i)" by(cases sobj, auto)
-
- let ?sh' = "sh(D\<mapsto>(sfs(F\<mapsto>v),i))" and ?f' = "(stk',loc,C,M,pc+1,No_ics)"
-
- have m_D: "P \<turnstile> D has F,Static:vT' in D" by (rule has_field_idemp[OF has_visible_field[OF F]])
- from "sh\<surd>" shD have sfs: "P,h,D \<turnstile>\<^sub>s sfs \<surd>" by (rule shconfD)
-
- have "sh'\<surd>": "P,h \<turnstile>\<^sub>s ?sh' \<surd>" by (rule shconf_upd_obj[OF "sh\<surd>" soconf_fupd[OF m_D vT' sfs]])
-
- from i mC s' v xc F cs cc' shD stk'
- have "\<sigma>' = (None, h, (stk',loc,C,M,pc+1,No_ics)#frs, ?sh')"
- by(fastforce simp: split_beta split: if_split_asm init_call_status.splits)
- moreover
- from ST'' ST' have "P,h \<turnstile> stk' [:\<le>] ST'" ..
- moreover
- from loc LT' have "P,h \<turnstile> loc [:\<le>\<^sub>\<top>] LT'" ..
- moreover
- have "conf_fs P h ?sh' \<Phi> C M (size Ts) T frs" by (rule conf_fs_shupd'[OF fs shD])
- moreover
- have "conf_clinit P ?sh' ((stk',loc,C,M,pc+1,No_ics)#frs)"
- by(rule conf_clinit_diff'[OF conf_clinit_shupd'[OF confc shD]])
- moreover
- note "h\<surd>" "sh'\<surd>" mC \<Phi>' pc' v vT'
- ultimately
- show "P,\<Phi> \<turnstile> \<sigma>' \<surd>" by fastforce
-qed
-(*>*)
-
-lemma Putstatic_Init_correct:
- fixes \<sigma>' :: jvm_state
- assumes wf: "wf_prog wt P"
- assumes mC: "P \<turnstile> C sees M,b:Ts\<rightarrow>T=(mxs,mxl\<^sub>0,ins,xt) in C"
- assumes i: "ins!pc = Putstatic C' F D"
- assumes wt: "P,T,mxs,size ins,xt \<turnstile> ins!pc,pc :: \<Phi> C M"
- assumes s': "Some \<sigma>' = exec (P, None, h, (stk,loc,C,M,pc,No_ics)#frs, sh)"
- assumes cf: "P,\<Phi> \<turnstile> (None, h, (stk,loc,C,M,pc,No_ics)#frs, sh)\<surd>"
- assumes xc: "fst (exec_step P h stk loc C M pc No_ics frs sh) = None"
- assumes nDone: "\<forall>sfs. sh (fst(field P D F)) \<noteq> Some(sfs, Done)"
-
- shows "P,\<Phi> \<turnstile> \<sigma>'\<surd>"
-(*<*)
-proof -
- from mC cf obtain ST LT where
- "h\<surd>": "P \<turnstile> h \<surd>" and
- "sh\<surd>": "P,h \<turnstile>\<^sub>s sh \<surd>" and
- \<Phi>: "\<Phi> C M ! pc = Some (ST,LT)" and
- stk: "P,h \<turnstile> stk [:\<le>] ST" and loc: "P,h \<turnstile> loc [:\<le>\<^sub>\<top>] LT" and
- pc: "pc < size ins" and
- fs: "conf_fs P h sh \<Phi> C M (size Ts) T frs" and
- confc: "conf_clinit P sh ((stk,loc,C,M,pc,No_ics)#frs)"
- by (fastforce dest: sees_method_fun)
-
- from i \<Phi> wt nDone obtain vT where
- F: "P \<turnstile> C' sees F,Static:vT in D"
- by fastforce
- then have has_field: "P \<turnstile> C' has F,Static:vT in D" by(rule has_visible_field)
-
- from field_def2[OF sees_field_idemp[OF F]] has_field_is_class'[OF has_field] obtain
- D[simp]: "fst(field P D F) = D" and
- cls: "is_class P D" by simp
-
- from i mC s' xc F nDone
- have "\<sigma>' = (None, h, (stk,loc,C,M,pc,Calling D [])#frs, sh)"
- by(auto simp: split_beta split: if_split_asm init_state.splits)
- moreover
- from confc have "conf_clinit P sh ((stk,loc,C,M,pc,Calling D [])#frs)"
- by(auto simp: conf_clinit_def distinct_clinit_def split: if_split_asm)
- moreover
- note loc stk "h\<surd>" "sh\<surd>" mC \<Phi> pc fs i has_field cls
- ultimately
- show "P,\<Phi> \<turnstile> \<sigma>' \<surd>" by fastforce
-qed
-(*>*)
-
-(* FIXME: move *)
-lemma oconf_blank2 [intro, simp]:
- "\<lbrakk>is_class P C; wf_prog wt P\<rbrakk> \<Longrightarrow> P,h \<turnstile> blank P C \<surd>"
-(*<*)
- by (fastforce simp: oconf_blank dest: wf_Fields_Ex)
-(*>*)
-
-lemma obj_ty_blank [iff]: "obj_ty (blank P C) = Class C"
- by (simp add: blank_def)
-
-lemma New_nInit_correct:
- fixes \<sigma>' :: jvm_state
- assumes wf: "wf_prog wt P"
- assumes meth: "P \<turnstile> C sees M,b:Ts\<rightarrow>T=(mxs,mxl\<^sub>0,ins,xt) in C"
- assumes ins: "ins!pc = New X"
- assumes wt: "P,T,mxs,size ins,xt \<turnstile> ins!pc,pc :: \<Phi> C M"
- assumes exec: "Some \<sigma>' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)"
- assumes conf: "P,\<Phi> \<turnstile> (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\<surd>"
- assumes no_x: "fst (exec_step P h stk loc C M pc ics frs sh) = None"
- assumes cs: "ics = Called [] \<or> (ics = No_ics \<and> (\<exists>sfs. sh X = Some(sfs, Done)))"
- shows "P,\<Phi> \<turnstile> \<sigma>'\<surd>"
-(*<*)
-proof -
- from ins conf meth
- obtain ST LT where
- heap_ok: "P\<turnstile> h\<surd>" and
- sheap_ok: "P,h \<turnstile>\<^sub>s sh \<surd>" and
- \<Phi>_pc: "\<Phi> C M!pc = Some (ST,LT)" and
- frame: "conf_f P h sh (ST,LT) ins (stk,loc,C,M,pc,ics)" and
- frames: "conf_fs P h sh \<Phi> C M (size Ts) T frs" and
- confc: "conf_clinit P sh ((stk,loc,C,M,pc,ics) # frs)"
- by (auto dest: sees_method_fun)
-
- from \<Phi>_pc ins wt
- obtain ST' LT' where
- is_class_X: "is_class P X" and
- mxs: "size ST < mxs" and
- suc_pc: "pc+1 < size ins" and
- \<Phi>_suc: "\<Phi> C M!(pc+1) = Some (ST', LT')" and
- less: "P \<turnstile> (Class X # ST, LT) \<le>\<^sub>i (ST', LT')"
- by auto
-
- from ins no_x cs meth obtain oref where new_Addr: "new_Addr h = Some oref" by auto
- hence h: "h oref = None" by (rule new_Addr_SomeD)
-
- with exec ins meth new_Addr cs have \<sigma>':
- "\<sigma>' = (None, h(oref \<mapsto> blank P X), (Addr oref#stk,loc,C,M,pc+1,No_ics)#frs, sh)"
- (is "\<sigma>' = (None, ?h', ?f # frs, sh)")
- by auto
- moreover
- from wf h heap_ok is_class_X have h': "P \<turnstile> ?h' \<surd>"
- by (auto intro: hconf_new)
- moreover
- from h frame less suc_pc wf
- have "conf_f P ?h' sh (ST', LT') ins ?f"
- by (clarsimp simp: fun_upd_apply conf_def blank_def split_beta)
- (auto intro: confs_hext confTs_hext)
- moreover
- from h have "h \<unlhd> ?h'" by simp
- with frames have "conf_fs P ?h' sh \<Phi> C M (size Ts) T frs" by (rule conf_fs_hext)
- moreover
- have "P,?h' \<turnstile>\<^sub>s sh \<surd>" by(rule shconf_hnew[OF sheap_ok h])
- moreover
- have "conf_clinit P sh (?f # frs)" by(rule conf_clinit_diff'[OF confc])
- ultimately
- show ?thesis using meth \<Phi>_suc by fastforce
-qed
-(*>*)
-
-lemma New_Init_correct:
- fixes \<sigma>' :: jvm_state
- assumes wf: "wf_prog wt P"
- assumes meth: "P \<turnstile> C sees M,b:Ts\<rightarrow>T=(mxs,mxl\<^sub>0,ins,xt) in C"
- assumes ins: "ins!pc = New X"
- assumes wt: "P,T,mxs,size ins,xt \<turnstile> ins!pc,pc :: \<Phi> C M"
- assumes exec: "Some \<sigma>' = exec (P, None, h, (stk,loc,C,M,pc,No_ics)#frs, sh)"
- assumes conf: "P,\<Phi> \<turnstile> (None, h, (stk,loc,C,M,pc,No_ics)#frs, sh)\<surd>"
- assumes no_x: "fst (exec_step P h stk loc C M pc No_ics frs sh) = None"
- assumes nDone: "\<forall>sfs. sh X \<noteq> Some(sfs, Done)"
- shows "P,\<Phi> \<turnstile> \<sigma>'\<surd>"
-(*<*)
-proof -
- from ins conf meth
- obtain ST LT where
- heap_ok: "P\<turnstile> h\<surd>" and
- sheap_ok: "P,h \<turnstile>\<^sub>s sh \<surd>" and
- \<Phi>_pc: "\<Phi> C M!pc = Some (ST,LT)" and
- frame: "conf_f P h sh (ST,LT) ins (stk,loc,C,M,pc,No_ics)" and
- frames: "conf_fs P h sh \<Phi> C M (size Ts) T frs" and
- confc: "conf_clinit P sh ((stk,loc,C,M,pc,No_ics) # frs)"
- by (auto dest: sees_method_fun)
-
- from \<Phi>_pc ins wt
- obtain ST' LT' where
- is_class_X: "is_class P X" and
- mxs: "size ST < mxs" and
- suc_pc: "pc+1 < size ins" and
- \<Phi>_suc: "\<Phi> C M!(pc+1) = Some (ST', LT')" and
- less: "P \<turnstile> (Class X # ST, LT) \<le>\<^sub>i (ST', LT')"
- by auto
-
- with exec ins meth nDone have \<sigma>':
- "\<sigma>' = (None, h, (stk,loc,C,M,pc,Calling X [])#frs, sh)"
- (is "\<sigma>' = (None, h, ?f # frs, sh)")
- by(auto split: init_state.splits)
- moreover
- from meth frame is_class_X ins
- have "conf_f P h sh (ST, LT) ins ?f" by auto
- moreover note heap_ok sheap_ok frames
- moreover
- from confc have "conf_clinit P sh (?f # frs)"
- by(auto simp: conf_clinit_def distinct_clinit_def split: if_split_asm)
- ultimately
- show ?thesis using meth \<Phi>_pc by fastforce
-qed
-(*>*)
-
-
-lemma Goto_correct:
-assumes "wf_prog wt P" and
- mC: "P \<turnstile> C sees M,b:Ts\<rightarrow>T=(mxs,mxl\<^sub>0,ins,xt) in C" and
- i: "ins!pc = Goto branch" and
- "P,T,mxs,size ins,xt \<turnstile> ins!pc,pc :: \<Phi> C M" and
- "Some \<sigma>' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)" and
- cf: "P,\<Phi> \<turnstile> (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\<surd>"
-shows "P,\<Phi> |- \<sigma>' [ok]"
-(*<*)
-proof -
- have "ics = No_ics" using mC i cf by(cases ics) auto
- then show ?thesis using assms(1,3-6) sees_method_fun[OF mC]
- by(fastforce elim!: conf_clinit_diff)
-qed
-(*>*)
-
-
-lemma IfFalse_correct:
-assumes "wf_prog wt P" and
- mC: "P \<turnstile> C sees M,b:Ts\<rightarrow>T=(mxs,mxl\<^sub>0,ins,xt) in C" and
- i: "ins!pc = IfFalse branch" and
- "P,T,mxs,size ins,xt \<turnstile> ins!pc,pc :: \<Phi> C M" and
- "Some \<sigma>' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)" and
- cf: "P,\<Phi> \<turnstile> (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\<surd>"
-shows "P,\<Phi> |- \<sigma>' [ok]"
-(*<*)
-proof -
- have "ics = No_ics" using mC i cf by(cases ics) auto
- then show ?thesis using assms(1,3-6) sees_method_fun[OF mC]
- by(fastforce elim!: conf_clinit_diff)
-qed
-(*>*)
-
-lemma CmpEq_correct:
-assumes "wf_prog wt P" and
- mC: "P \<turnstile> C sees M,b:Ts\<rightarrow>T=(mxs,mxl\<^sub>0,ins,xt) in C" and
- i: "ins!pc = CmpEq" and
- "P,T,mxs,size ins,xt \<turnstile> ins!pc,pc :: \<Phi> C M" and
- "Some \<sigma>' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)" and
- cf: "P,\<Phi> \<turnstile> (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\<surd>"
-shows "P,\<Phi> |- \<sigma>' [ok]"
-(*<*)
-proof -
- have "ics = No_ics" using mC i cf by(cases ics) auto
- then show ?thesis using assms(1,3-6) sees_method_fun[OF mC]
- by(fastforce elim!: conf_clinit_diff)
-qed
-(*>*)
-
-lemma Pop_correct:
-assumes "wf_prog wt P" and
- mC: "P \<turnstile> C sees M,b:Ts\<rightarrow>T=(mxs,mxl\<^sub>0,ins,xt) in C" and
- i: "ins!pc = Pop" and
- "P,T,mxs,size ins,xt \<turnstile> ins!pc,pc :: \<Phi> C M" and
- "Some \<sigma>' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)" and
- cf: "P,\<Phi> \<turnstile> (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\<surd>"
-shows "P,\<Phi> |- \<sigma>' [ok]"
-(*<*)
-proof -
- have "ics = No_ics" using mC i cf by(cases ics) auto
- then show ?thesis using assms(1,3-6) sees_method_fun[OF mC]
- by(fastforce elim!: conf_clinit_diff)
-qed
-(*>*)
-
-
-lemma IAdd_correct:
-assumes "wf_prog wt P" and
- mC: "P \<turnstile> C sees M,b:Ts\<rightarrow>T=(mxs,mxl\<^sub>0,ins,xt) in C" and
- i: "ins!pc = IAdd" and
- "P,T,mxs,size ins,xt \<turnstile> ins!pc,pc :: \<Phi> C M" and
- "Some \<sigma>' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)" and
- cf: "P,\<Phi> \<turnstile> (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\<surd>"
-shows "P,\<Phi> |- \<sigma>' [ok]"
-(*<*)
-proof -
- have "ics = No_ics" using mC i cf by(cases ics) auto
- then show ?thesis using assms(1,3-6) sees_method_fun[OF mC]
- by(fastforce elim!: conf_clinit_diff)
-qed
-(*>*)
-
-
-lemma Throw_correct:
-assumes "wf_prog wt P" and
- mC: "P \<turnstile> C sees M,b:Ts\<rightarrow>T=(mxs,mxl\<^sub>0,ins,xt) in C" and
- i: "ins!pc = Throw" and
- "Some \<sigma>' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)" and
- cf: "P,\<Phi> \<turnstile> (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\<surd>" and
- "fst (exec_step P h stk loc C M pc ics frs sh) = None"
-shows "P,\<Phi> |- \<sigma>' [ok]"
-(*<*)
-proof -
- have "ics = No_ics" using mC i cf by(cases ics) auto
- then show ?thesis using assms by simp
-qed
-(*>*)
-
-text \<open>
- The next theorem collects the results of the sections above,
- i.e.~exception handling, initialization procedure steps, and
- the execution step for each instruction. It states type safety
- for single step execution: in welltyped programs, a conforming
- state is transformed into another conforming state when one
- step of execution is performed.
-\<close>
-lemma step_correct:
-fixes \<sigma>' :: jvm_state
-assumes wtp: "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P"
- and meth: "P \<turnstile> C sees M,b:Ts\<rightarrow>T=(mxs,mxl\<^sub>0,ins,xt) in C"
- and exec: "Some \<sigma>' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)"
- and conf: "P,\<Phi> \<turnstile> (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\<surd>"
-shows "P,\<Phi> \<turnstile> \<sigma>'\<surd>"
-(*<*)
-proof -
- from assms have pc: "pc < length ins" by(auto dest: sees_method_fun)
- with wt_jvm_prog_impl_wt_instr[OF wtp meth] have wt: "P,T,mxs,size ins,xt \<turnstile> ins!pc,pc :: \<Phi> C M"
- by simp
-
- from conf obtain ST LT where \<Phi>: "\<Phi> C M ! pc = Some(ST,LT)" by clarsimp
-
- show ?thesis
- proof(cases "fst (exec_step P h stk loc C M pc ics frs sh)")
- case Some show ?thesis by(rule xcpt_correct[OF wtp meth wt Some exec conf])
- next
- case None
- from wt_jvm_progD[OF wtp] obtain wf_md where wf: "wf_prog wf_md P" by clarify
-
- { assume [simp]: "ics = No_ics"
-
- from exec conf None obtain
- exec': "Some \<sigma>' = exec (P, None, h, (stk,loc,C,M,pc,No_ics)#frs, sh)"
- and conf': "P,\<Phi> \<turnstile> (None, h, (stk,loc,C,M,pc,No_ics)#frs, sh)\<surd>"
- and None': "fst (exec_step P h stk loc C M pc No_ics frs sh) = None" by simp
-
- have ?thesis
- proof(cases "ins!pc")
- case Load from Load_correct[OF wf meth this wt exec conf] show ?thesis by simp
- next
- case Store from Store_correct[OF wf meth this wt exec conf] show ?thesis by simp
- next
- case Push from Push_correct[OF wf meth this wt exec conf] show ?thesis by simp
- next
- case (New C) show ?thesis
- proof(cases "\<exists>sfs. sh C = Some(sfs, Done)")
- case True
- with New_nInit_correct[OF wf meth New wt exec conf None] show ?thesis by simp
- next
- case False
- with New_Init_correct[OF wf meth New wt exec' conf' None'] show ?thesis by simp
- qed
- next
- case Getfield from Getfield_correct[OF wf meth this wt exec conf None]
- show ?thesis by simp
- next
- case (Getstatic C F D) show ?thesis
- proof(cases "\<exists>sfs. sh (fst (field P D F)) = Some(sfs, Done)")
- case True
- with Getstatic_nInit_correct[OF wf meth Getstatic wt exec conf None] show ?thesis by simp
- next
- case False
- with Getstatic_Init_correct[OF wf meth Getstatic wt exec' conf' None']
- show ?thesis by simp
- qed
- next
- case Putfield from Putfield_correct[OF wf meth this wt exec conf None]
- show ?thesis by simp
- next
- case (Putstatic C F D) show ?thesis
- proof(cases "\<exists>sfs. sh (fst (field P D F)) = Some(sfs, Done)")
- case True
- with Putstatic_nInit_correct[OF wf meth Putstatic wt exec conf None] show ?thesis by simp
- next
- case False
- with Putstatic_Init_correct[OF wf meth Putstatic wt exec' conf' None']
- show ?thesis by simp
- qed
- next
- case Checkcast from Checkcast_correct[OF wtp meth this wt exec conf None]
- show ?thesis by simp
- next
- case Invoke with Invoke_correct[OF wtp meth this wt exec conf None] show ?thesis by simp
- next
- case (Invokestatic C M n)
- from wf_jvm_prog_nclinit[OF wtp meth wt pc \<Phi> this] have ncl: "M \<noteq> clinit" by simp
- show ?thesis
- proof(cases "\<exists>sfs. sh (fst (method P C M)) = Some(sfs, Done)")
- case True
- with Invokestatic_nInit_correct[OF wtp meth Invokestatic ncl wt exec conf None]
- show ?thesis by simp
- next
- case False
- with Invokestatic_Init_correct[OF wtp meth Invokestatic ncl wt exec' conf' None']
- show ?thesis by simp
- qed
- next
- case Return with Return_correct[OF wtp meth this wt exec conf] show ?thesis by simp
- next
- case Pop with Pop_correct[OF wf meth this wt exec conf] show ?thesis by simp
- next
- case IAdd with IAdd_correct[OF wf meth this wt exec conf] show ?thesis by simp
- next
- case Goto with Goto_correct[OF wf meth this wt exec conf] show ?thesis by simp
- next
- case CmpEq with CmpEq_correct[OF wf meth this wt exec conf] show ?thesis by simp
- next
- case IfFalse with IfFalse_correct[OF wf meth this wt exec conf] show ?thesis by simp
- next
- case Throw with Throw_correct[OF wf meth this exec conf None] show ?thesis by simp
- qed
- }
- moreover
- { fix Cs assume [simp]: "ics = Called Cs"
- have ?thesis
- proof(cases Cs)
- case [simp]: Nil
- from conf meth obtain C1 where "Called_context P C1 (ins ! pc)"
- by(clarsimp simp: conf_f_def2 intro!: Called_context_Called_set)
- then have "ins!pc \<in> Called_set" by(rule Called_context_Called_set)
- then show ?thesis
- proof(cases "ins!pc")
- case (New C)
- from New_nInit_correct[OF wf meth this wt exec conf None] show ?thesis by simp
- next
- case (Getstatic C F D)
- from Getstatic_nInit_correct[OF wf meth this wt exec conf None] show ?thesis by simp
- next
- case (Putstatic C F D)
- from Putstatic_nInit_correct[OF wf meth this wt exec conf None] show ?thesis by simp
- next
- case (Invokestatic C M n)
- from wf_jvm_prog_nclinit[OF wtp meth wt pc \<Phi> this] have ncl: "M \<noteq> clinit" by simp
- with Invokestatic_nInit_correct[OF wtp meth Invokestatic ncl wt exec conf None]
- show ?thesis by simp
- qed(simp_all)
- next
- case (Cons C' Cs') with Called_correct[OF wtp meth exec conf None] show ?thesis by simp
- qed
- }
- moreover
- { fix C' Cs assume [simp]: "ics = Calling C' Cs"
- with Calling_correct[OF wtp meth exec conf None] have ?thesis by simp
- }
- moreover
- { fix Cs a assume [simp]: "ics = Throwing Cs a"
- have ?thesis
- proof(cases Cs)
- case Nil with exec None show ?thesis by simp
- next
- case (Cons C' Cs')
- with Throwing_correct[OF wtp meth exec conf None] show ?thesis by simp
- qed
- }
- ultimately show ?thesis by(cases ics) auto
- qed
-qed
-(*>*)
-
-subsection \<open> Main \<close>
-
-lemma correct_state_impl_Some_method:
- "P,\<Phi> \<turnstile> (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\<surd>
- \<Longrightarrow> \<exists>b m Ts T. P \<turnstile> C sees M,b:Ts\<rightarrow>T = m in C"
- by fastforce
-
-lemma BV_correct_1 [rule_format]:
-"\<And>\<sigma>. \<lbrakk> wf_jvm_prog\<^bsub>\<Phi>\<^esub> P; P,\<Phi> \<turnstile> \<sigma>\<surd>\<rbrakk> \<Longrightarrow> P \<turnstile> \<sigma> -jvm\<rightarrow>\<^sub>1 \<sigma>' \<longrightarrow> P,\<Phi> \<turnstile> \<sigma>'\<surd>"
-(*<*)
-proof -
- fix \<sigma> assume wf: "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P" and cf: "P,\<Phi> \<turnstile> \<sigma>\<surd>"
- obtain xp h frs sh where \<sigma>[simp]: "\<sigma> = (xp, h, frs, sh)" by(cases \<sigma>) simp
- have "exec (P, xp, h, frs, sh) = \<lfloor>\<sigma>'\<rfloor> \<longrightarrow> P,\<Phi> |- \<sigma>' [ok]"
- proof(cases xp)
- case None
- with wf cf show ?thesis
- proof(cases frs)
- case (Cons fr frs')
- obtain stk loc C M pc ics where [simp]: "fr = (stk,loc,C,M,pc,ics)" by(cases fr) simp
- then have cf': "P,\<Phi> |- (None, h, (stk,loc,C,M,pc,ics) # frs', sh) [ok]"
- using Cons None cf by simp
- then obtain b mxs mxl\<^sub>0 ins xt Ts T
- where mC: "P \<turnstile> C sees M, b : Ts\<rightarrow>T = (mxs, mxl\<^sub>0, ins, xt) in C"
- using correct_state_impl_Some_method[OF cf'] by clarify
- show ?thesis using Cons None step_correct[OF wf mC _ cf'] by simp
- qed simp
- next
- case (Some a)
- then show ?thesis using wf cf by (case_tac frs) simp_all
- qed
- then show "P \<turnstile> \<sigma> -jvm\<rightarrow>\<^sub>1 \<sigma>' \<longrightarrow> P,\<Phi> \<turnstile> \<sigma>'\<surd>" by(simp add: exec_1_iff)
-qed
-(*>*)
-
-
-theorem progress:
- "\<lbrakk> xp=None; frs\<noteq>[] \<rbrakk> \<Longrightarrow> \<exists>\<sigma>'. P \<turnstile> (xp,h,frs,sh) -jvm\<rightarrow>\<^sub>1 \<sigma>'"
- by (clarsimp simp: exec_1_iff neq_Nil_conv split_beta
- simp del: split_paired_Ex)
-
-lemma progress_conform:
- "\<lbrakk>wf_jvm_prog\<^bsub>\<Phi>\<^esub> P; P,\<Phi> \<turnstile> (xp,h,frs,sh)\<surd>; xp=None; frs\<noteq>[]\<rbrakk>
- \<Longrightarrow> \<exists>\<sigma>'. P \<turnstile> (xp,h,frs,sh) -jvm\<rightarrow>\<^sub>1 \<sigma>' \<and> P,\<Phi> \<turnstile> \<sigma>'\<surd>"
-(*<*)by (drule (1) progress) (fast intro: BV_correct_1)(*>*)
-
-theorem BV_correct [rule_format]:
-"\<lbrakk> wf_jvm_prog\<^bsub>\<Phi>\<^esub> P; P \<turnstile> \<sigma> -jvm\<rightarrow> \<sigma>' \<rbrakk> \<Longrightarrow> P,\<Phi> \<turnstile> \<sigma>\<surd> \<longrightarrow> P,\<Phi> \<turnstile> \<sigma>'\<surd>"
-(*<*)
-proof -
- assume wf: "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P" and "P \<turnstile> \<sigma> -jvm\<rightarrow> \<sigma>'"
- then have "(\<sigma>, \<sigma>') \<in> (exec_1 P)\<^sup>*" by (simp only: exec_all_def1)
- then show ?thesis proof(induct rule: rtrancl_induct)
- case (step y z)
- then show ?case by clarify (erule (1) BV_correct_1[OF wf])
- qed simp
-qed
-(*>*)
-
-lemma hconf_start:
- assumes wf: "wf_prog wf_mb P"
- shows "P \<turnstile> (start_heap P) \<surd>"
-(*<*)
-proof -
- { fix a obj assume assm: "start_heap P a = \<lfloor>obj\<rfloor>"
- have "P,start_heap P \<turnstile> obj \<surd>" using wf assm[THEN sym]
- by (unfold start_heap_def)
- (auto simp: fun_upd_apply is_class_xcpt split: if_split_asm)
- }
- then show ?thesis using preallocated_start[of P]
- by (unfold hconf_def) simp
-qed
-(*>*)
-
-lemma shconf_start:
- "\<not> is_class P Start \<Longrightarrow> P,start_heap P \<turnstile>\<^sub>s start_sheap \<surd>"
-(*<*)
- by (unfold shconf_def)
- (clarsimp simp: preallocated_start fun_upd_apply soconf_def has_field_is_class)
-(*>*)
-
-lemma BV_correct_initial:
-assumes wf: "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P"
- and nStart: "\<not>is_class P Start"
- and mC: "P \<turnstile> C sees M,Static:[]\<rightarrow>Void = m in C"
- and nclinit: "M \<noteq> clinit"
- and \<Phi>: "\<Phi>' Start start_m = start_\<phi>\<^sub>m"
-shows "start_prog P C M,\<Phi>' \<turnstile> start_state P \<surd>"
-(*<*)
-proof -
- let ?P = "class_add P (start_class C M)"
- let ?h = "start_heap P" and ?sh = "[Start \<mapsto> (Map.empty, Done)]"
- let ?C = Start and ?M = start_m and ?pc = 0 and ?ics = No_ics
- let ?f = "([], [], ?C, ?M, ?pc, ?ics)" and ?fs = "[]"
- let ?frs = "?f#?fs"
- have "is_class P Object" using wf by(simp add: wf_jvm_prog_phi_def)
- then obtain Mm where Mm: "P \<turnstile> Object sees_methods Mm"
- by(fastforce simp: is_class_def dest: sees_methods_Object)
- obtain mxs mxl\<^sub>0 ins xt where
- mC': "P \<turnstile> C sees M,Static:[]\<rightarrow>Void = (mxs,mxl\<^sub>0,ins,xt) in C"
- using mC by (cases m) simp
- have "?P\<turnstile> ?h\<surd>" using wf nStart class_add_hconf_wf[OF _ hconf_start]
- by (simp add: wf_jvm_prog_phi_def start_heap_nStart)
- moreover have "?P,?h\<turnstile>\<^sub>s ?sh\<surd>" by(rule start_prog_start_shconf)
- moreover have "conf_clinit ?P ?sh ?frs"
- by(simp add: conf_clinit_def distinct_clinit_def)
- moreover have "\<exists>b Ts T mxs mxl\<^sub>0 is xt \<tau>.
- (?P \<turnstile> ?C sees ?M,b:Ts\<rightarrow>T = (mxs,mxl\<^sub>0,is,xt) in ?C)
- \<and> \<Phi>' ?C ?M ! ?pc = Some \<tau>
- \<and> conf_f ?P ?h ?sh \<tau> is ?f \<and> conf_fs ?P ?h ?sh \<Phi>' ?C ?M (size Ts) T ?fs"
- using \<Phi> start_prog_Start_sees_start_method[OF Mm]
- by (unfold conf_f_def wt_start_def) fastforce
- ultimately show ?thesis
- by (simp del: defs1 add: start_state_def correct_state_def)
-qed
-
-declare [[simproc add: list_to_set_comprehension]]
-(*>*)
-
-theorem typesafe:
- assumes welltyped: "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P"
- assumes nstart: "\<not> is_class P Start"
- assumes main_method: "P \<turnstile> C sees M,Static:[]\<rightarrow>Void = m in C"
- assumes nclinit: "M \<noteq> clinit"
- assumes \<Phi>: "\<And>C. C \<noteq> Start \<Longrightarrow> \<Phi>' C = \<Phi> C"
- assumes \<Phi>': "\<Phi>' Start start_m = start_\<phi>\<^sub>m" "\<Phi>' Start clinit = start_\<phi>\<^sub>m"
- assumes Obj_start_m:
- "(\<And>b' Ts' T' m' D'. P \<turnstile> Object sees start_m, b' : Ts'\<rightarrow>T' = m' in D'
- \<Longrightarrow> b' = Static \<and> Ts' = [] \<and> T' = Void)"
- shows "start_prog P C M \<turnstile> start_state P -jvm\<rightarrow> \<sigma> \<Longrightarrow> start_prog P C M,\<Phi>' \<turnstile> \<sigma> \<surd>"
-(*<*)
-proof -
- from welltyped nstart main_method nclinit \<Phi>'(1)
- have "start_prog P C M,\<Phi>' \<turnstile> start_state P \<surd>" by (rule BV_correct_initial)
- moreover
- assume "start_prog P C M \<turnstile> start_state P -jvm\<rightarrow> \<sigma>"
- moreover
- from start_prog_wf_jvm_prog_phi[OF welltyped nstart main_method nclinit \<Phi> \<Phi>' Obj_start_m]
- have "wf_jvm_prog\<^bsub>\<Phi>'\<^esub>(start_prog P C M)" by simp
- ultimately
- show "start_prog P C M,\<Phi>' \<turnstile> \<sigma> \<surd>" using welltyped by - (rule BV_correct)
-qed
-(*>*)
-
-end
+(* Title: JinjaDCI/BV/BVSpecTypeSafe.thy
+
+ Author: Cornelia Pusch, Gerwin Klein, Susannah Mansky
+ Copyright 1999 Technische Universitaet Muenchen, 2019-20 UIUC
+
+ Based on the Jinja theory BV/BVSpecTypeSafe.thy by Cornelia Pusch and Gerwin Klein
+*)
+
+
+section \<open> BV Type Safety Proof \label{sec:BVSpecTypeSafe} \<close>
+
+theory BVSpecTypeSafe
+imports BVConform StartProg
+begin
+
+text \<open>
+ This theory contains proof that the specification of the bytecode
+ verifier only admits type safe programs.
+\<close>
+
+subsection \<open> Preliminaries \<close>
+
+text \<open>
+ Simp and intro setup for the type safety proof:
+\<close>
+lemmas defs1 = correct_state_def conf_f_def wt_instr_def eff_def norm_eff_def app_def xcpt_app_def
+
+lemmas widen_rules [intro] = conf_widen confT_widen confs_widens confTs_widen
+
+
+subsection \<open> Exception Handling \<close>
+
+
+text \<open>
+ For the @{text Invoke} instruction the BV has checked all handlers
+ that guard the current @{text pc}.
+\<close>
+lemma Invoke_handlers:
+ "match_ex_table P C pc xt = Some (pc',d') \<Longrightarrow>
+ \<exists>(f,t,D,h,d) \<in> set (relevant_entries P (Invoke n M) pc xt).
+ P \<turnstile> C \<preceq>\<^sup>* D \<and> pc \<in> {f..<t} \<and> pc' = h \<and> d' = d"
+ by (induct xt) (auto simp: relevant_entries_def matches_ex_entry_def
+ is_relevant_entry_def split: if_split_asm)
+
+text \<open>
+ For the @{text Invokestatic} instruction the BV has checked all handlers
+ that guard the current @{text pc}.
+\<close>
+lemma Invokestatic_handlers:
+ "match_ex_table P C pc xt = Some (pc',d') \<Longrightarrow>
+ \<exists>(f,t,D,h,d) \<in> set (relevant_entries P (Invokestatic C\<^sub>0 n M) pc xt).
+ P \<turnstile> C \<preceq>\<^sup>* D \<and> pc \<in> {f..<t} \<and> pc' = h \<and> d' = d"
+ by (induct xt) (auto simp: relevant_entries_def matches_ex_entry_def
+ is_relevant_entry_def split: if_split_asm)
+
+text \<open>
+ For the instrs in @{text Called_set} the BV has checked all handlers
+ that guard the current @{text pc}.
+\<close>
+lemma Called_set_handlers:
+ "match_ex_table P C pc xt = Some (pc',d') \<Longrightarrow> i \<in> Called_set \<Longrightarrow>
+ \<exists>(f,t,D,h,d) \<in> set (relevant_entries P i pc xt).
+ P \<turnstile> C \<preceq>\<^sup>* D \<and> pc \<in> {f..<t} \<and> pc' = h \<and> d' = d"
+ by (induct xt) (auto simp: relevant_entries_def matches_ex_entry_def
+ is_relevant_entry_def split: if_split_asm)
+
+text \<open>
+ We can prove separately that the recursive search for exception
+ handlers (@{text find_handler}) in the frame stack results in
+ a conforming state (if there was no matching exception handler
+ in the current frame). We require that the exception is a valid
+ heap address, and that the state before the exception occurred
+ conforms.
+\<close>
+lemma uncaught_xcpt_correct:
+ assumes wt: "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P"
+ assumes h: "h xcp = Some obj"
+ shows "\<And>f. P,\<Phi> \<turnstile> (None, h, f#frs, sh)\<surd>
+ \<Longrightarrow> curr_method f \<noteq> clinit \<Longrightarrow> P,\<Phi> \<turnstile> find_handler P xcp h frs sh \<surd>"
+ (is "\<And>f. ?correct (None, h, f#frs, sh) \<Longrightarrow> ?prem f \<Longrightarrow> ?correct (?find frs)")
+(*<*)
+proof (induct frs)
+ \<comment> \<open>the base
+ case is trivial as it should be\<close>
+ show "?correct (?find [])" by (simp add: correct_state_def)
+next
+ \<comment> \<open>we will need both forms @{text wf_jvm_prog} and @{text wf_prog} later\<close>
+ from wt obtain mb where wf: "wf_prog mb P" by (simp add: wf_jvm_prog_phi_def)
+
+ \<comment> \<open>the assumptions for the cons case:\<close>
+ fix f f' frs' assume cr: "?correct (None, h, f#f'#frs', sh)"
+ assume pr: "?prem f"
+
+ \<comment> \<open>the induction hypothesis:\<close>
+ assume IH: "\<And>f. ?correct (None, h, f#frs', sh) \<Longrightarrow> ?prem f \<Longrightarrow> ?correct (?find frs')"
+
+ from cr pr conf_clinit_Cons[where frs="f'#frs'" and f=f] obtain
+ confc: "conf_clinit P sh (f'#frs')"
+ and cr': "?correct (None, h, f'#frs', sh)" by(fastforce simp: correct_state_def)
+
+ obtain stk loc C M pc ics where [simp]: "f' = (stk,loc,C,M,pc,ics)" by (cases f')
+
+ from cr' obtain b Ts T mxs mxl\<^sub>0 ins xt where
+ meth: "P \<turnstile> C sees M,b:Ts \<rightarrow> T = (mxs,mxl\<^sub>0,ins,xt) in C"
+ by (simp add: correct_state_def, blast)
+
+ hence xt[simp]: "ex_table_of P C M = xt" by simp
+
+ have cls: "is_class P C" by(rule sees_method_is_class'[OF meth])
+ from cr' obtain sfs where
+ sfs: "M = clinit \<Longrightarrow> sh C = Some(sfs,Processing)" by(fastforce simp: defs1 conf_clinit_def)
+
+ show "?correct (?find (f'#frs'))"
+ proof (cases "match_ex_table P (cname_of h xcp) pc xt")
+ case None with cr' IH[of f'] show ?thesis
+ proof(cases "M=clinit")
+ case True then show ?thesis using xt cr' IH[of f'] None h conf_clinit_Called_Throwing
+ conf_f_Throwing[where h=h and sh=sh, OF _ cls h sfs]
+ by(cases frs', auto simp: correct_state_def image_iff) fastforce
+ qed(auto)
+ next
+ fix pc_d
+ assume "match_ex_table P (cname_of h xcp) pc xt = Some pc_d"
+ then obtain pc' d' where
+ match: "match_ex_table P (cname_of h xcp) pc xt = Some (pc',d')"
+ (is "?match (cname_of h xcp) = _")
+ by (cases pc_d) auto
+
+ from wt meth cr' [simplified]
+ have wti: "P,T,mxs,size ins,xt \<turnstile> ins!pc,pc :: \<Phi> C M"
+ by (fastforce simp: correct_state_def conf_f_def
+ dest: sees_method_fun
+ elim!: wt_jvm_prog_impl_wt_instr)
+
+ from cr' obtain ST LT where \<Phi>: "\<Phi> C M ! pc = Some (ST, LT)"
+ by(fastforce dest: sees_method_fun simp: correct_state_def)
+
+ from cr' \<Phi> meth have conf': "conf_f P h sh (ST, LT) ins f'"
+ by (unfold correct_state_def) (fastforce dest: sees_method_fun)
+ hence loc: "P,h \<turnstile> loc [:\<le>\<^sub>\<top>] LT" and
+ stk: "P,h \<turnstile> stk [:\<le>] ST" by (unfold conf_f_def) auto
+ hence [simp]: "size stk = size ST" by (simp add: list_all2_lengthD)
+
+ from cr meth pr
+ obtain D n M' where
+ ins: "ins!pc = Invoke n M' \<or> ins!pc = Invokestatic D n M'" (is "_ = ?i \<or> _ = ?i'")
+ by(fastforce dest: sees_method_fun simp: correct_state_def)
+
+ with match obtain f1 t D where
+ rel: "(f1,t,D,pc',d') \<in> set (relevant_entries P (ins!pc) pc xt)" and
+ D: "P \<turnstile> cname_of h xcp \<preceq>\<^sup>* D"
+ by(fastforce dest: Invoke_handlers Invokestatic_handlers)
+
+ from rel have
+ "(pc', Some (Class D # drop (size ST - d') ST, LT)) \<in> set (xcpt_eff (ins!pc) P pc (ST,LT) xt)"
+ (is "(_, Some (?ST',_)) \<in> _")
+ by (force simp: xcpt_eff_def image_def)
+ with wti \<Phi> obtain
+ pc: "pc' < size ins" and
+ "P \<turnstile> Some (?ST', LT) \<le>' \<Phi> C M ! pc'"
+ by (clarsimp simp: defs1) blast
+ then obtain ST' LT' where
+ \<Phi>': "\<Phi> C M ! pc' = Some (ST',LT')" and
+ less: "P \<turnstile> (?ST', LT) \<le>\<^sub>i (ST',LT')"
+ by (auto simp: sup_state_opt_any_Some)
+
+ let ?f = "(Addr xcp # drop (length stk - d') stk, loc, C, M, pc',No_ics)"
+ have "conf_f P h sh (ST',LT') ins ?f"
+ proof -
+ from wf less loc have "P,h \<turnstile> loc [:\<le>\<^sub>\<top>] LT'" by simp blast
+ moreover from D h have "P,h \<turnstile> Addr xcp :\<le> Class D"
+ by (simp add: conf_def obj_ty_def case_prod_unfold)
+ with less stk
+ have "P,h \<turnstile> Addr xcp # drop (length stk - d') stk [:\<le>] ST'"
+ by (auto intro!: list_all2_dropI)
+ ultimately show ?thesis using pc conf' by(auto simp: conf_f_def)
+ qed
+
+ with cr' match \<Phi>' meth pc
+ show ?thesis by (unfold correct_state_def)
+ (cases "M=clinit"; fastforce dest: sees_method_fun simp: conf_clinit_def distinct_clinit_def)
+ qed
+qed
+(*>*)
+
+text \<open>
+ The requirement of lemma @{text uncaught_xcpt_correct} (that
+ the exception is a valid reference on the heap) is always met
+ for welltyped instructions and conformant states:
+\<close>
+lemma exec_instr_xcpt_h:
+ "\<lbrakk> fst (exec_instr (ins!pc) P h stk vars C M pc ics frs sh) = Some xcp;
+ P,T,mxs,size ins,xt \<turnstile> ins!pc,pc :: \<Phi> C M;
+ P,\<Phi> \<turnstile> (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\<surd> \<rbrakk>
+ \<Longrightarrow> \<exists>obj. h xcp = Some obj"
+ (is "\<lbrakk> ?xcpt; ?wt; ?correct \<rbrakk> \<Longrightarrow> ?thesis")
+(*<*)
+proof -
+ note [simp] = split_beta
+ note [split] = if_split_asm option.split_asm
+
+ assume wt: ?wt ?correct
+ hence pre: "preallocated h" by (simp add: correct_state_def hconf_def)
+
+ assume xcpt: ?xcpt
+ with exec_instr_xcpts have
+ opt: "ins!pc = Throw \<or> xcp \<in> {a. \<exists>x \<in> sys_xcpts. a = addr_of_sys_xcpt x}" by simp
+
+ with pre show ?thesis
+ proof (cases "ins!pc")
+ case Throw with xcpt wt pre show ?thesis
+ by (clarsimp iff: list_all2_Cons2 simp: defs1)
+ (auto dest: non_npD simp: is_refT_def elim: preallocatedE)
+ qed (auto elim: preallocatedE)
+qed
+(*>*)
+
+lemma exec_step_xcpt_h:
+assumes xcpt: "fst (exec_step P h stk vars C M pc ics frs sh) = Some xcp"
+ and ins: "instrs_of P C M = ins"
+ and wti: "P,T,mxs,size ins,xt \<turnstile> ins!pc,pc :: \<Phi> C M"
+ and correct: "P,\<Phi> \<turnstile> (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\<surd>"
+shows "\<exists>obj. h xcp = Some obj"
+proof -
+ from correct have pre: "preallocated h" by(simp add: defs1 hconf_def)
+ { fix C' Cs assume ics[simp]: "ics = Calling C' Cs"
+ with xcpt have "xcp = addr_of_sys_xcpt NoClassDefFoundError"
+ by(cases ics, auto simp: split_beta split: init_state.splits if_split_asm)
+ with pre have ?thesis using preallocated_def by force
+ }
+ moreover
+ { fix Cs a assume [simp]: "ics = Throwing Cs a"
+ with xcpt have eq: "a = xcp" by(cases Cs; simp)
+
+ from correct have "P,h,sh \<turnstile>\<^sub>i (C,M,pc,ics)" by(auto simp: defs1)
+ with eq have ?thesis by simp
+ }
+ moreover
+ { fix Cs assume ics: "ics = No_ics \<or> ics = Called Cs"
+ with exec_instr_xcpt_h[OF _ wti correct] xcpt ins have ?thesis by(cases Cs, auto)
+ }
+ ultimately show ?thesis by(cases ics, auto)
+qed
+
+lemma conf_sys_xcpt:
+ "\<lbrakk>preallocated h; C \<in> sys_xcpts\<rbrakk> \<Longrightarrow> P,h \<turnstile> Addr (addr_of_sys_xcpt C) :\<le> Class C"
+ by (auto elim: preallocatedE)
+
+lemma match_ex_table_SomeD:
+ "match_ex_table P C pc xt = Some (pc',d') \<Longrightarrow>
+ \<exists>(f,t,D,h,d) \<in> set xt. matches_ex_entry P C pc (f,t,D,h,d) \<and> h = pc' \<and> d=d'"
+ by (induct xt) (auto split: if_split_asm)
+
+text \<open>
+ Finally we can state that, whenever an exception occurs, the
+ next state always conforms:
+\<close>
+lemma xcpt_correct:
+ fixes \<sigma>' :: jvm_state
+ assumes wtp: "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P"
+ assumes meth: "P \<turnstile> C sees M,b:Ts\<rightarrow>T=(mxs,mxl\<^sub>0,ins,xt) in C"
+ assumes wt: "P,T,mxs,size ins,xt \<turnstile> ins!pc,pc :: \<Phi> C M"
+ assumes xp: "fst (exec_step P h stk loc C M pc ics frs sh) = Some xcp"
+ assumes s': "Some \<sigma>' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)"
+ assumes correct: "P,\<Phi> \<turnstile> (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\<surd>"
+ shows "P,\<Phi> \<turnstile> \<sigma>'\<surd>"
+(*<*)
+proof -
+ from wtp obtain wfmb where wf: "wf_prog wfmb P"
+ by (simp add: wf_jvm_prog_phi_def)
+
+ from meth have ins[simp]: "instrs_of P C M = ins" by simp
+ have cls: "is_class P C" by(rule sees_method_is_class[OF meth])
+ from correct obtain sfs where
+ sfs: "M = clinit \<Longrightarrow> sh C = Some(sfs,Processing)"
+ by(auto simp: correct_state_def conf_clinit_def conf_f_def2)
+
+ note conf_sys_xcpt [elim!]
+ note xp' = meth s' xp
+
+ from correct meth
+ obtain ST LT where
+ h_ok: "P \<turnstile> h \<surd>" and
+ sh_ok: "P,h \<turnstile>\<^sub>s sh \<surd>" and
+ \<Phi>_pc: "\<Phi> C M ! pc = Some (ST, LT)" and
+ frame: "conf_f P h sh (ST,LT) ins (stk,loc,C,M,pc,ics)" and
+ frames: "conf_fs P h sh \<Phi> C M (size Ts) T frs" and
+ confc: "conf_clinit P sh ((stk,loc,C,M,pc,ics)#frs)" and
+ vics: "P,h,sh \<turnstile>\<^sub>i (C,M,pc,ics)"
+ by(auto simp: defs1 dest: sees_method_fun)
+
+ from frame obtain
+ stk: "P,h \<turnstile> stk [:\<le>] ST" and
+ loc: "P,h \<turnstile> loc [:\<le>\<^sub>\<top>] LT" and
+ pc: "pc < size ins"
+ by (unfold conf_f_def) auto
+
+ from h_ok have preh: "preallocated h" by (simp add: hconf_def)
+
+ note wtp
+ moreover
+ from exec_step_xcpt_h[OF xp ins wt correct]
+ obtain obj where h: "h xcp = Some obj" by clarify
+ moreover note correct
+ ultimately
+ have fh: "curr_method (stk,loc,C,M,pc,ics) \<noteq> clinit
+ \<Longrightarrow> P,\<Phi> \<turnstile> find_handler P xcp h frs sh \<surd>" by (rule uncaught_xcpt_correct)
+ with xp'
+ have "M \<noteq> clinit \<Longrightarrow> \<forall>Cs a. ics \<noteq> Throwing Cs a
+ \<Longrightarrow> match_ex_table P (cname_of h xcp) pc xt = None \<Longrightarrow> ?thesis"
+ (is "?nc \<Longrightarrow> ?t \<Longrightarrow> ?m (cname_of h xcp) = _ \<Longrightarrow> _" is "?nc \<Longrightarrow> ?t \<Longrightarrow> ?match = _ \<Longrightarrow> _")
+ by(cases ics; simp add: split_beta)
+ moreover
+ from correct xp' conf_clinit_Called_Throwing conf_f_Throwing[where h=h and sh=sh, OF _ cls h sfs]
+ have "M = clinit \<Longrightarrow> \<forall>Cs a. ics \<noteq> Throwing Cs a
+ \<Longrightarrow> match_ex_table P (cname_of h xcp) pc xt = None \<Longrightarrow> ?thesis"
+ by(cases frs, auto simp: correct_state_def image_iff split_beta) fastforce
+ moreover
+ { fix pc_d assume "?match = Some pc_d"
+ then obtain pc' d' where some_handler: "?match = Some (pc',d')"
+ by (cases pc_d) auto
+
+ from stk have [simp]: "size stk = size ST" ..
+
+ from wt \<Phi>_pc have
+ eff: "\<forall>(pc', s')\<in>set (xcpt_eff (ins!pc) P pc (ST,LT) xt).
+ pc' < size ins \<and> P \<turnstile> s' \<le>' \<Phi> C M!pc'"
+ by (auto simp: defs1)
+
+ from some_handler obtain f t D where
+ xt: "(f,t,D,pc',d') \<in> set xt" and
+ "matches_ex_entry P (cname_of h xcp) pc (f,t,D,pc',d')"
+ by (auto dest: match_ex_table_SomeD)
+
+ hence match: "P \<turnstile> cname_of h xcp \<preceq>\<^sup>* D" "pc \<in> {f..<t}"
+ by (auto simp: matches_ex_entry_def)
+
+ { fix C' Cs assume ics: "ics = Calling C' Cs \<or> ics = Called (C'#Cs)"
+
+ let ?stk' = "Addr xcp # drop (length stk - d') stk"
+ let ?f = "(?stk', loc, C, M, pc', No_ics)"
+ from some_handler xp' ics
+ have \<sigma>': "\<sigma>' = (None, h, ?f#frs, sh)"
+ by (cases ics; simp add: split_beta)
+
+ from xp ics have "xcp = addr_of_sys_xcpt NoClassDefFoundError"
+ by(cases ics, auto simp: split_beta split: init_state.splits if_split_asm)
+ with match preh have conf: "P,h \<turnstile> Addr xcp :\<le> Class D" by fastforce
+
+ from correct ics obtain C1 where "Called_context P C1 (ins!pc)"
+ by(fastforce simp: correct_state_def conf_f_def2)
+ then have "ins!pc \<in> Called_set" by(rule Called_context_Called_set)
+ with xt match have "(f,t,D,pc',d') \<in> set (relevant_entries P (ins!pc) pc xt)"
+ by(auto simp: relevant_entries_def is_relevant_entry_def)
+
+ with eff obtain ST' LT' where
+ \<Phi>_pc': "\<Phi> C M ! pc' = Some (ST', LT')" and
+ pc': "pc' < size ins" and
+ less: "P \<turnstile> (Class D # drop (size ST - d') ST, LT) \<le>\<^sub>i (ST', LT')"
+ by (fastforce simp: xcpt_eff_def sup_state_opt_any_Some)
+
+ with conf loc stk conf_f_def2 frame ics have "conf_f P h sh (ST',LT') ins ?f"
+ by (auto simp: defs1 intro: list_all2_dropI)
+ with meth h_ok frames \<Phi>_pc' \<sigma>' sh_ok confc ics
+ have ?thesis
+ by (unfold correct_state_def)
+ (auto dest: sees_method_fun conf_clinit_diff' sees_method_is_class; fastforce)
+ }
+ moreover
+ { assume ics: "ics = No_ics \<or> ics = Called []"
+
+ let ?stk' = "Addr xcp # drop (length stk - d') stk"
+ let ?f = "(?stk', loc, C, M, pc', No_ics)"
+ from some_handler xp' ics
+ have \<sigma>': "\<sigma>' = (None, h, ?f#frs, sh)"
+ by (cases ics; simp add: split_beta)
+
+ from xp ics obtain
+ "(f,t,D,pc',d') \<in> set (relevant_entries P (ins!pc) pc xt)" and
+ conf: "P,h \<turnstile> Addr xcp :\<le> Class D"
+ proof (cases "ins!pc")
+ case Return
+ with xp ics have False by(cases ics; cases frs, auto simp: split_beta split: if_split_asm)
+ then show ?thesis by simp
+ next
+ case New with xp match
+ have "is_relevant_entry P (ins!pc) pc (f,t,D,pc',d')"
+ by (simp add: is_relevant_entry_def)
+ moreover
+ from xp wt correct obtain obj where xcp: "h xcp = Some obj"
+ by (blast dest: exec_step_xcpt_h[OF _ ins])
+ ultimately
+ show ?thesis using xt match
+ by (auto simp: relevant_entries_def conf_def case_prod_unfold intro: that)
+ next
+ case Getfield with xp ics
+ have xcp: "xcp = addr_of_sys_xcpt NullPointer \<or> xcp = addr_of_sys_xcpt NoSuchFieldError
+ \<or> xcp = addr_of_sys_xcpt IncompatibleClassChangeError"
+ by (cases ics; simp add: split_beta split: if_split_asm staticb.splits)
+ with Getfield match preh have "is_relevant_entry P (ins!pc) pc (f,t,D,pc',d')"
+ by (fastforce simp: is_relevant_entry_def)
+ with match preh xt xcp
+ show ?thesis by(fastforce simp: relevant_entries_def intro: that)
+ next
+ case Getstatic with xp match
+ have "is_relevant_entry P (ins!pc) pc (f,t,D,pc',d')"
+ by (simp add: is_relevant_entry_def)
+ moreover
+ from xp wt correct obtain obj where xcp: "h xcp = Some obj"
+ by (blast dest: exec_step_xcpt_h[OF _ ins])
+ ultimately
+ show ?thesis using xt match
+ by (auto simp: relevant_entries_def conf_def case_prod_unfold intro: that)
+ next
+ case Putfield with xp ics
+ have xcp: "xcp = addr_of_sys_xcpt NullPointer \<or> xcp = addr_of_sys_xcpt NoSuchFieldError
+ \<or> xcp = addr_of_sys_xcpt IncompatibleClassChangeError"
+ by (cases ics; simp add: split_beta split: if_split_asm staticb.splits)
+ with Putfield match preh have "is_relevant_entry P (ins!pc) pc (f,t,D,pc',d')"
+ by (fastforce simp: is_relevant_entry_def)
+ with match preh xt xcp
+ show ?thesis by (fastforce simp: relevant_entries_def intro: that)
+ next
+ case Putstatic with xp match
+ have "is_relevant_entry P (ins!pc) pc (f,t,D,pc',d')"
+ by (simp add: is_relevant_entry_def)
+ moreover
+ from xp wt correct obtain obj where xcp: "h xcp = Some obj"
+ by (blast dest: exec_step_xcpt_h[OF _ ins])
+ ultimately
+ show ?thesis using xt match
+ by (auto simp: relevant_entries_def conf_def case_prod_unfold intro: that)
+ next
+ case Checkcast with xp ics
+ have [simp]: "xcp = addr_of_sys_xcpt ClassCast"
+ by (cases ics; simp add: split_beta split: if_split_asm)
+ with Checkcast match preh have "is_relevant_entry P (ins!pc) pc (f,t,D,pc',d')"
+ by (simp add: is_relevant_entry_def)
+ with match preh xt
+ show ?thesis by (fastforce simp: relevant_entries_def intro: that)
+ next
+ case Invoke with xp match
+ have "is_relevant_entry P (ins!pc) pc (f,t,D,pc',d')"
+ by (simp add: is_relevant_entry_def)
+ moreover
+ from xp wt correct obtain obj where xcp: "h xcp = Some obj"
+ by (blast dest: exec_step_xcpt_h[OF _ ins])
+ ultimately
+ show ?thesis using xt match
+ by (auto simp: relevant_entries_def conf_def case_prod_unfold intro: that)
+ next
+ case Invokestatic with xp match
+ have "is_relevant_entry P (ins!pc) pc (f,t,D,pc',d')"
+ by (simp add: is_relevant_entry_def)
+ moreover
+ from xp wt correct obtain obj where xcp: "h xcp = Some obj"
+ by (blast dest: exec_step_xcpt_h[OF _ ins])
+ ultimately
+ show ?thesis using xt match
+ by (auto simp: relevant_entries_def conf_def case_prod_unfold intro: that)
+ next
+ case Throw with xp match preh
+ have "is_relevant_entry P (ins!pc) pc (f,t,D,pc',d')"
+ by (simp add: is_relevant_entry_def)
+ moreover
+ from xp wt correct obtain obj where xcp: "h xcp = Some obj"
+ by (blast dest: exec_step_xcpt_h[OF _ ins])
+ ultimately
+ show ?thesis using xt match
+ by (auto simp: relevant_entries_def conf_def case_prod_unfold intro: that)
+ qed(cases ics, (auto)[5])+
+
+ with eff obtain ST' LT' where
+ \<Phi>_pc': "\<Phi> C M ! pc' = Some (ST', LT')" and
+ pc': "pc' < size ins" and
+ less: "P \<turnstile> (Class D # drop (size ST - d') ST, LT) \<le>\<^sub>i (ST', LT')"
+ by (fastforce simp: xcpt_eff_def sup_state_opt_any_Some)
+
+ with conf loc stk conf_f_def2 frame ics have "conf_f P h sh (ST',LT') ins ?f"
+ by (auto simp: defs1 intro: list_all2_dropI)
+ with meth h_ok frames \<Phi>_pc' \<sigma>' sh_ok confc ics
+ have ?thesis
+ by (unfold correct_state_def) (auto dest: sees_method_fun conf_clinit_diff'; fastforce)
+ }
+ ultimately
+ have "\<forall>Cs a. ics \<noteq> Throwing Cs a \<Longrightarrow> ?thesis" by(cases ics; metis list.exhaust)
+ }
+ moreover
+ { fix Cs a assume "ics = Throwing Cs a"
+ with xp' have ics: "ics = Throwing [] xcp" by(cases Cs; clarsimp)
+
+ let ?frs = "(stk,loc,C,M,pc,No_ics)#frs"
+
+ have eT: "exec_step P h stk loc C M pc (Throwing [] xcp) frs sh = (Some xcp, h, ?frs, sh)"
+ by auto
+ with xp' ics have \<sigma>'_fh: "\<sigma>' = find_handler P xcp h ?frs sh" by simp
+
+ from meth have [simp]: "xt = ex_table_of P C M" by simp
+
+ let ?match = "match_ex_table P (cname_of h xcp) pc xt"
+
+ { assume clinit: "M = clinit" and None: "?match = None"
+ note asms = clinit None
+
+ have "P,\<Phi> |- find_handler P xcp h ?frs sh [ok]"
+ proof(cases frs)
+ case Nil
+ with h_ok sh_ok asms show "P,\<Phi> |- find_handler P xcp h ?frs sh [ok]"
+ by(simp add: correct_state_def)
+ next
+ case [simp]: (Cons f' frs')
+ obtain stk' loc' C' M' pc' ics' where
+ [simp]: "f' = (stk',loc',C',M',pc',ics')" by(cases f')
+
+ have cls: "is_class P C" by(rule sees_method_is_class[OF meth])
+ have shC: "sh C = Some(sfs,Processing)" by(rule sfs[OF clinit])
+
+ from correct obtain b Ts T mxs' mxl\<^sub>0' ins' xt' ST' LT' where
+ meth': "P \<turnstile> C' sees M', b : Ts\<rightarrow>T = (mxs', mxl\<^sub>0', ins', xt') in C'" and
+ \<Phi>_pc': "\<Phi> C' M' ! pc' = \<lfloor>(ST', LT')\<rfloor>" and
+ frame': "conf_f P h sh (ST',LT') ins' (stk', loc', C', M', pc', ics')" and
+ frames': "conf_fs P h sh \<Phi> C' M' (length Ts) T frs'" and
+ confc': "conf_clinit P sh ((stk',loc',C',M',pc',ics')#frs')"
+ by(auto dest: conf_clinit_Cons simp: correct_state_def)
+
+ from meth' have
+ ins'[simp]: "instrs_of P C' M' = ins'"
+ and [simp]: "xt' = ex_table_of P C' M'" by simp+
+
+ let ?f' = "case ics' of Called Cs' \<Rightarrow> (stk',loc',C',M',pc',Throwing (C#Cs') xcp)
+ | _ \<Rightarrow> (stk',loc',C',M',pc',ics')"
+
+ from asms confc have confc_T: "conf_clinit P sh (?f'#frs')"
+ by(cases ics', auto simp: conf_clinit_def distinct_clinit_def)
+
+ from asms conf_f_Throwing[where h=h and sh=sh, OF _ cls h shC] frame' have
+ frame_T: "conf_f P h sh (ST', LT') ins' ?f'" by(cases ics'; simp)
+ with h_ok sh_ok meth' \<Phi>_pc' confc_T frames'
+ have "P,\<Phi> |- (None, h, ?f'#frs', sh) [ok]"
+ by(cases ics') (fastforce simp: correct_state_def)+
+
+ with asms show ?thesis by(cases ics'; simp)
+ qed
+ }
+ moreover
+ { assume asms: "M \<noteq> clinit" "?match = None"
+
+ from asms uncaught_xcpt_correct[OF wtp h correct]
+ have "P,\<Phi> |- find_handler P xcp h frs sh [ok]" by simp
+ with asms have "P,\<Phi> |- find_handler P xcp h ?frs sh [ok]" by auto
+ }
+ moreover
+ { fix pc_d assume some_handler: "?match = \<lfloor>pc_d\<rfloor>"
+ (is "?match = \<lfloor>pc_d\<rfloor>")
+ then obtain pc1 d1 where sh': "?match = Some(pc1,d1)" by(cases pc_d, simp)
+
+ let ?stk' = "Addr xcp # drop (length stk - d1) stk"
+ let ?f = "(?stk', loc, C, M, pc1, No_ics)"
+
+ from stk have [simp]: "size stk = size ST" ..
+
+ from wt \<Phi>_pc have
+ eff: "\<forall>(pc1, s')\<in>set (xcpt_eff (ins!pc) P pc (ST,LT) xt).
+ pc1 < size ins \<and> P \<turnstile> s' \<le>' \<Phi> C M!pc1"
+ by (auto simp: defs1)
+
+ from match_ex_table_SomeD[OF sh'] obtain f t D where
+ xt: "(f,t,D,pc1,d1) \<in> set xt" and
+ "matches_ex_entry P (cname_of h xcp) pc (f,t,D,pc1,d1)" by auto
+
+ hence match: "P \<turnstile> cname_of h xcp \<preceq>\<^sup>* D" "pc \<in> {f..<t}"
+ by (auto simp: matches_ex_entry_def)
+
+ from ics vics obtain C1 where "Called_context P C1 (ins ! pc)" by auto
+ then have "ins!pc \<in> Called_set" by(rule Called_context_Called_set)
+ with match xt xp ics obtain
+ res: "(f,t,D,pc1,d1) \<in> set (relevant_entries P (ins!pc) pc xt)"
+ by(auto simp: relevant_entries_def is_relevant_entry_def)
+
+ with h match xt xp ics have conf: "P,h \<turnstile> Addr xcp :\<le> Class D"
+ by (auto simp: relevant_entries_def conf_def case_prod_unfold)
+
+ with eff res obtain ST1 LT1 where
+ \<Phi>_pc1: "\<Phi> C M ! pc1 = Some (ST1, LT1)" and
+ pc1: "pc1 < size ins" and
+ less1: "P \<turnstile> (Class D # drop (size ST - d1) ST, LT) \<le>\<^sub>i (ST1, LT1)"
+ by (fastforce simp: xcpt_eff_def sup_state_opt_any_Some)
+
+ with conf loc stk conf_f_def2 frame ics have frame1: "conf_f P h sh (ST1,LT1) ins ?f"
+ by (auto simp: defs1 intro: list_all2_dropI)
+
+ from \<Phi>_pc1 h_ok sh_ok meth frame1 frames conf_clinit_diff'[OF confc] have
+ "P,\<Phi> |- (None, h, ?f # frs, sh) [ok]" by(fastforce simp: correct_state_def)
+ with sh' have "P,\<Phi> |- find_handler P xcp h ?frs sh [ok]" by auto
+ }
+ ultimately
+ have cr': "P,\<Phi> |- find_handler P xcp h ?frs sh [ok]" by(cases "?match") blast+
+
+ with \<sigma>'_fh have ?thesis by simp
+ }
+ ultimately
+ show ?thesis by (cases "?match") blast+
+qed
+(*>*)
+
+(**********Non-exception Single-step correctness*************************)
+declare defs1 [simp]
+
+subsection \<open> Initialization procedure steps \<close>
+
+text \<open>
+ In this section we prove that, for states that result in a step of the
+ initialization procedure rather than an instruction execution, the state
+ after execution of the step still conforms.
+\<close>
+
+lemma Calling_correct:
+ fixes \<sigma>' :: jvm_state
+ assumes wtprog: "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P"
+ assumes mC: "P \<turnstile> C sees M,b:Ts\<rightarrow>T=(mxs,mxl\<^sub>0,ins,xt) in C"
+ assumes s': "Some \<sigma>' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)"
+ assumes cf: "P,\<Phi> \<turnstile> (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\<surd>"
+ assumes xc: "fst (exec_step P h stk loc C M pc ics frs sh) = None"
+ assumes ics: "ics = Calling C' Cs"
+
+ shows "P,\<Phi> \<turnstile> \<sigma>'\<surd>"
+proof -
+ from wtprog obtain wfmb where wf: "wf_prog wfmb P"
+ by (simp add: wf_jvm_prog_phi_def)
+
+ from mC cf obtain ST LT where
+ h_ok: "P \<turnstile> h \<surd>" and
+ sh_ok: "P,h \<turnstile>\<^sub>s sh \<surd>" and
+ \<Phi>: "\<Phi> C M ! pc = Some (ST,LT)" and
+ stk: "P,h \<turnstile> stk [:\<le>] ST" and loc: "P,h \<turnstile> loc [:\<le>\<^sub>\<top>] LT" and
+ pc: "pc < size ins" and
+ frame: "conf_f P h sh (ST, LT) ins (stk,loc,C,M,pc,ics)" and
+ fs: "conf_fs P h sh \<Phi> C M (size Ts) T frs" and
+ confc: "conf_clinit P sh ((stk,loc,C,M,pc,ics)#frs)" and
+ vics: "P,h,sh \<turnstile>\<^sub>i (C,M,pc,ics)"
+ by (fastforce dest: sees_method_fun)
+
+ with ics have confc\<^sub>0: "conf_clinit P sh ((stk,loc,C,M,pc,Calling C' Cs)#frs)" by simp
+
+ from vics ics have cls': "is_class P C'" by auto
+
+ { assume None: "sh C' = None"
+
+ let ?sh = "sh(C' \<mapsto> (sblank P C', Prepared))"
+
+ obtain FDTs where
+ flds: "P \<turnstile> C' has_fields FDTs" using wf_Fields_Ex[OF wf cls'] by clarsimp
+
+ from shconf_upd_obj[where C=C', OF sh_ok soconf_sblank[OF flds]]
+ have sh_ok': "P,h \<turnstile>\<^sub>s ?sh \<surd>" by simp
+
+ from None have "\<forall>sfs. sh C' \<noteq> Some(sfs,Processing)" by simp
+ with conf_clinit_nProc_dist[OF confc] have
+ dist': "distinct (C' # clinit_classes ((stk, loc, C, M, pc, ics) # frs))" by simp
+ then have dist'': "distinct (C' # clinit_classes frs)" by simp
+
+ have confc': "conf_clinit P ?sh ((stk, loc, C, M, pc, ics) # frs)"
+ by(rule conf_clinit_shupd[OF confc dist'])
+ have fs': "conf_fs P h ?sh \<Phi> C M (size Ts) T frs" by(rule conf_fs_shupd[OF fs dist''])
+ from vics ics have vics': "P,h,?sh \<turnstile>\<^sub>i (C, M, pc, ics)" by auto
+
+ from s' ics None have "\<sigma>' = (None, h, (stk, loc, C, M, pc, ics)#frs, ?sh)" by auto
+
+ with mC h_ok sh_ok' \<Phi> stk loc pc fs' confc vics' confc' frame None
+ have ?thesis by fastforce
+ }
+ moreover
+ { fix a assume "sh C' = Some a"
+ then obtain sfs i where shC'[simp]: "sh C' = Some(sfs,i)" by(cases a, simp)
+
+ from confc ics have last: "\<exists>sobj. sh (last(C'#Cs)) = Some sobj"
+ by(fastforce simp: conf_clinit_def)
+
+ let "?f" = "\<lambda>ics'. (stk, loc, C, M, pc, ics'::init_call_status)"
+
+ { assume i: "i = Done \<or> i = Processing"
+ let ?ics = "Called Cs"
+
+ from last vics ics have vics': "P,h,sh \<turnstile>\<^sub>i (C, M, pc, ?ics)" by auto
+ from confc ics have confc': "conf_clinit P sh (?f ?ics#frs)"
+ by(cases "M=clinit"; clarsimp simp: conf_clinit_def distinct_clinit_def)
+
+ from i s' ics have "\<sigma>' = (None, h, ?f ?ics#frs, sh)" by auto
+
+ with mC h_ok sh_ok \<Phi> stk loc pc fs confc' vics' frame ics
+ have ?thesis by fastforce
+ }
+ moreover
+ { assume i[simp]: "i = Error"
+ let ?a = "addr_of_sys_xcpt NoClassDefFoundError"
+ let ?ics = "Throwing Cs ?a"
+
+ from h_ok have preh: "preallocated h" by (simp add: hconf_def)
+ then obtain obj where ha: "h ?a = Some obj" by(clarsimp simp: preallocated_def sys_xcpts_def)
+ with vics ics have vics': "P,h,sh \<turnstile>\<^sub>i (C, M, pc, ?ics)" by auto
+
+ from confc ics have confc'': "conf_clinit P sh (?f ?ics#frs)"
+ by(cases "M=clinit"; clarsimp simp: conf_clinit_def distinct_clinit_def)
+
+ from s' ics have \<sigma>': "\<sigma>' = (None, h, ?f ?ics#frs, sh)" by auto
+
+ from mC h_ok sh_ok \<Phi> stk loc pc fs confc'' vics \<sigma>' ics ha
+ have ?thesis by fastforce
+ }
+ moreover
+ { assume i[simp]: "i = Prepared"
+ let ?sh = "sh(C' \<mapsto> (sfs,Processing))"
+ let ?D = "fst(the(class P C'))"
+ let ?ics = "if C' = Object then Called (C'#Cs) else Calling ?D (C'#Cs)"
+
+ from shconf_upd_obj[where C=C', OF sh_ok shconfD[OF sh_ok shC']]
+ have sh_ok': "P,h \<turnstile>\<^sub>s ?sh \<surd>" by simp
+
+ from cls' have "C' \<noteq> Object \<Longrightarrow> P \<turnstile> C' \<preceq>\<^sup>* ?D" by(auto simp: is_class_def intro!: subcls1I)
+ with is_class_supclass[OF wf _ cls'] have D: "C' \<noteq> Object \<Longrightarrow> is_class P ?D" by simp
+
+ from i have "\<forall>sfs. sh C' \<noteq> Some(sfs,Processing)" by simp
+ with conf_clinit_nProc_dist[OF confc\<^sub>0] have
+ dist': "distinct (C' # clinit_classes ((stk, loc, C, M, pc, Calling C' Cs) # frs))" by fast
+ then have dist'': "distinct (C' # clinit_classes frs)" by simp
+
+ from conf_clinit_shupd_Calling[OF confc\<^sub>0 dist' cls']
+ conf_clinit_shupd_Called[OF confc\<^sub>0 dist' cls']
+ have confc': "conf_clinit P ?sh (?f ?ics#frs)" by clarsimp
+ with last ics have "\<exists>sobj. ?sh (last(C'#Cs)) = Some sobj"
+ by(auto simp: conf_clinit_def fun_upd_apply)
+ with D vics ics have vics': "P,h,?sh \<turnstile>\<^sub>i (C, M, pc, ?ics)" by auto
+
+ have fs': "conf_fs P h ?sh \<Phi> C M (size Ts) T frs" by(rule conf_fs_shupd[OF fs dist''])
+
+ from frame vics' have frame': "conf_f P h ?sh (ST, LT) ins (?f ?ics)" by simp
+
+ from i s' ics have "\<sigma>' = (None, h, ?f ?ics#frs, ?sh)" by(auto simp: if_split_asm)
+
+ with mC h_ok sh_ok' \<Phi> stk loc pc fs' confc' frame' ics
+ have ?thesis by fastforce
+ }
+ ultimately have ?thesis by(cases i, auto)
+ }
+ ultimately show ?thesis by(cases "sh C'", auto)
+qed
+
+lemma Throwing_correct:
+ fixes \<sigma>' :: jvm_state
+ assumes wtprog: "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P"
+ assumes mC: "P \<turnstile> C sees M,b:Ts\<rightarrow>T=(mxs,mxl\<^sub>0,ins,xt) in C"
+ assumes s': "Some \<sigma>' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)"
+ assumes cf: "P,\<Phi> \<turnstile> (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\<surd>"
+ assumes xc: "fst (exec_step P h stk loc C M pc ics frs sh) = None"
+ assumes ics: "ics = Throwing (C'#Cs) a"
+
+ shows "P,\<Phi> \<turnstile> \<sigma>'\<surd>"
+proof -
+ from wtprog obtain wfmb where wf: "wf_prog wfmb P"
+ by (simp add: wf_jvm_prog_phi_def)
+
+ from mC cf obtain ST LT where
+ h_ok: "P \<turnstile> h \<surd>" and
+ sh_ok: "P,h \<turnstile>\<^sub>s sh \<surd>" and
+ \<Phi>: "\<Phi> C M ! pc = Some (ST,LT)" and
+ stk: "P,h \<turnstile> stk [:\<le>] ST" and loc: "P,h \<turnstile> loc [:\<le>\<^sub>\<top>] LT" and
+ pc: "pc < size ins" and
+ frame: "conf_f P h sh (ST, LT) ins (stk,loc,C,M,pc,ics)" and
+ fs: "conf_fs P h sh \<Phi> C M (size Ts) T frs" and
+ confc: "conf_clinit P sh ((stk,loc,C,M,pc,ics)#frs)" and
+ vics: "P,h,sh \<turnstile>\<^sub>i (C,M,pc,ics)"
+ by (fastforce dest: sees_method_fun)
+
+ with ics have confc\<^sub>0: "conf_clinit P sh ((stk,loc,C,M,pc,Throwing (C'#Cs) a)#frs)" by simp
+
+ from frame ics mC have
+ cc: "\<exists>C1. Called_context P C1 (ins ! pc)" by(clarsimp simp: conf_f_def2)
+
+ from frame ics obtain obj where ha: "h a = Some obj" by(auto simp: conf_f_def2)
+
+ from confc ics obtain sfs i where shC': "sh C' = Some(sfs,i)" by(clarsimp simp: conf_clinit_def)
+ then have sfs: "P,h,C' \<turnstile>\<^sub>s sfs \<surd>" by(rule shconfD[OF sh_ok])
+
+ from s' ics
+ have \<sigma>': "\<sigma>' = (None, h, (stk,loc,C,M,pc,Throwing Cs a)#frs, sh(C' \<mapsto> (fst(the(sh C')), Error)))"
+ (is "\<sigma>' = (None, h, ?f'#frs, ?sh')")
+ by simp
+
+ from confc ics have dist: "distinct (C' # clinit_classes (?f' # frs))"
+ by (simp add: conf_clinit_def distinct_clinit_def)
+ then have dist': "distinct (C' # clinit_classes frs)" by simp
+
+ from conf_clinit_Throwing confc ics have confc': "conf_clinit P sh (?f' # frs)" by simp
+
+ from shconf_upd_obj[OF sh_ok sfs] shC' have "P,h \<turnstile>\<^sub>s ?sh' \<surd>" by simp
+ moreover
+ have "conf_fs P h ?sh' \<Phi> C M (length Ts) T frs" by(rule conf_fs_shupd[OF fs dist'])
+ moreover
+ have "conf_clinit P ?sh' (?f' # frs)" by(rule conf_clinit_shupd[OF confc' dist])
+ moreover note \<sigma>' h_ok mC \<Phi> pc stk loc ha cc
+ ultimately show "P,\<Phi> \<turnstile> \<sigma>' \<surd>" by fastforce
+qed
+
+lemma Called_correct:
+ fixes \<sigma>' :: jvm_state
+ assumes wtprog: "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P"
+ assumes mC: "P \<turnstile> C sees M,b:Ts\<rightarrow>T=(mxs,mxl\<^sub>0,ins,xt) in C"
+ assumes s': "Some \<sigma>' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)"
+ assumes cf: "P,\<Phi> \<turnstile> (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\<surd>"
+ assumes xc: "fst (exec_step P h stk loc C M pc ics frs sh) = None"
+ assumes ics[simp]: "ics = Called (C'#Cs)"
+
+ shows "P,\<Phi> \<turnstile> \<sigma>'\<surd>"
+proof -
+ from wtprog obtain wfmb where wf: "wf_prog wfmb P"
+ by (simp add: wf_jvm_prog_phi_def)
+
+ from mC cf obtain ST LT where
+ h_ok: "P \<turnstile> h \<surd>" and
+ sh_ok: "P,h \<turnstile>\<^sub>s sh \<surd>" and
+ \<Phi>: "\<Phi> C M ! pc = Some (ST,LT)" and
+ stk: "P,h \<turnstile> stk [:\<le>] ST" and loc: "P,h \<turnstile> loc [:\<le>\<^sub>\<top>] LT" and
+ pc: "pc < size ins" and
+ frame: "conf_f P h sh (ST, LT) ins (stk,loc,C,M,pc,ics)" and
+ fs: "conf_fs P h sh \<Phi> C M (size Ts) T frs" and
+ confc: "conf_clinit P sh ((stk,loc,C,M,pc,ics)#frs)" and
+ vics: "P,h,sh \<turnstile>\<^sub>i (C,M,pc,ics)"
+ by (fastforce dest: sees_method_fun)
+
+ then have confc\<^sub>0: "conf_clinit P sh ((stk,loc,C,M,pc,Called (C'#Cs))#frs)" by simp
+
+ from frame mC obtain C1 sobj where
+ ss: "Called_context P C1 (ins ! pc)" and
+ shC1: "sh C1 = Some sobj" by(clarsimp simp: conf_f_def2)
+
+ from confc wf_sees_clinit[OF wf] obtain mxs' mxl' ins' xt' where
+ clinit: "P \<turnstile> C' sees clinit,Static: [] \<rightarrow> Void=(mxs',mxl',ins',xt') in C'"
+ by(fastforce simp: conf_clinit_def is_class_def)
+
+ let ?loc' = "replicate mxl' undefined"
+
+ from s' clinit
+ have \<sigma>': "\<sigma>' = (None, h, ([],?loc',C',clinit,0,No_ics)#(stk,loc,C,M,pc,Called Cs)#frs, sh)"
+ (is "\<sigma>' = (None, h, ?if#?f'#frs, sh)")
+ by simp
+
+ with wtprog clinit
+ obtain start: "wt_start P C' Static [] mxl' (\<Phi> C' clinit)" and ins': "ins' \<noteq> []"
+ by (auto dest: wt_jvm_prog_impl_wt_start)
+ then obtain LT\<^sub>0 where LT\<^sub>0: "\<Phi> C' clinit ! 0 = Some ([], LT\<^sub>0)"
+ by (clarsimp simp: wt_start_def defs1 sup_state_opt_any_Some split: staticb.splits)
+ moreover
+ have "conf_f P h sh ([], LT\<^sub>0) ins' ?if"
+ proof -
+ let ?LT = "replicate mxl' Err"
+ have "P,h \<turnstile> ?loc' [:\<le>\<^sub>\<top>] ?LT" by simp
+ also from start LT\<^sub>0 have "P \<turnstile> \<dots> [\<le>\<^sub>\<top>] LT\<^sub>0" by (simp add: wt_start_def)
+ finally have "P,h \<turnstile> ?loc' [:\<le>\<^sub>\<top>] LT\<^sub>0" .
+ thus ?thesis using ins' by simp
+ qed
+ moreover
+ from conf_clinit_Called confc clinit have "conf_clinit P sh (?if # ?f' # frs)" by simp
+ moreover note \<sigma>' h_ok sh_ok mC \<Phi> pc stk loc clinit ss shC1 fs
+ ultimately show "P,\<Phi> \<turnstile> \<sigma>' \<surd>" by fastforce
+qed
+
+subsection \<open> Single Instructions \<close>
+
+text \<open>
+ In this section we prove for each single (welltyped) instruction
+ that the state after execution of the instruction still conforms.
+ Since we have already handled exceptions above, we can now assume that
+ no exception occurs in this step. For instructions that may call
+ the initialization procedure, we cover the calling and non-calling
+ cases separately.
+\<close>
+
+lemma Invoke_correct:
+ fixes \<sigma>' :: jvm_state
+ assumes wtprog: "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P"
+ assumes meth_C: "P \<turnstile> C sees M,b:Ts\<rightarrow>T=(mxs,mxl\<^sub>0,ins,xt) in C"
+ assumes ins: "ins ! pc = Invoke M' n"
+ assumes wti: "P,T,mxs,size ins,xt \<turnstile> ins!pc,pc :: \<Phi> C M"
+ assumes \<sigma>': "Some \<sigma>' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)"
+ assumes approx: "P,\<Phi> \<turnstile> (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\<surd>"
+ assumes no_xcp: "fst (exec_step P h stk loc C M pc ics frs sh) = None"
+ shows "P,\<Phi> \<turnstile> \<sigma>'\<surd>"
+(*<*)
+proof -
+ from meth_C approx ins have [simp]: "ics = No_ics" by(cases ics, auto)
+
+ note split_paired_Ex [simp del]
+
+ from wtprog obtain wfmb where wfprog: "wf_prog wfmb P"
+ by (simp add: wf_jvm_prog_phi_def)
+
+ from ins meth_C approx obtain ST LT where
+ heap_ok: "P\<turnstile> h\<surd>" and
+ \<Phi>_pc: "\<Phi> C M!pc = Some (ST,LT)" and
+ frame: "conf_f P h sh (ST,LT) ins (stk,loc,C,M,pc,ics)" and
+ frames: "conf_fs P h sh \<Phi> C M (size Ts) T frs" and
+ confc: "conf_clinit P sh ((stk,loc,C,M,pc,ics)#frs)"
+ by (fastforce dest: sees_method_fun)
+
+ from ins wti \<Phi>_pc
+ have n: "n < size ST" by simp
+
+ { assume "stk!n = Null"
+ with ins no_xcp meth_C have False by (simp add: split_beta)
+ hence ?thesis ..
+ }
+ moreover
+ { assume "ST!n = NT"
+ moreover
+ from frame have "P,h \<turnstile> stk [:\<le>] ST" by simp
+ with n have "P,h \<turnstile> stk!n :\<le> ST!n" by (simp add: list_all2_conv_all_nth)
+ ultimately
+ have "stk!n = Null" by simp
+ with ins no_xcp meth_C have False by (simp add: split_beta)
+ hence ?thesis ..
+ }
+ moreover {
+ assume NT: "ST!n \<noteq> NT" and Null: "stk!n \<noteq> Null"
+
+ from NT ins wti \<Phi>_pc obtain D D' b Ts T m ST' LT' where
+ D: "ST!n = Class D" and
+ pc': "pc+1 < size ins" and
+ m_D: "P \<turnstile> D sees M',b: Ts\<rightarrow>T = m in D'" and
+ Ts: "P \<turnstile> rev (take n ST) [\<le>] Ts" and
+ \<Phi>': "\<Phi> C M ! (pc+1) = Some (ST', LT')" and
+ LT': "P \<turnstile> LT [\<le>\<^sub>\<top>] LT'" and
+ ST': "P \<turnstile> (T # drop (n+1) ST) [\<le>] ST'" and
+ b[simp]: "b = NonStatic"
+ by (clarsimp simp: sup_state_opt_any_Some)
+
+ from frame obtain
+ stk: "P,h \<turnstile> stk [:\<le>] ST" and
+ loc: "P,h \<turnstile> loc [:\<le>\<^sub>\<top>] LT" by simp
+
+ from n stk D have "P,h \<turnstile> stk!n :\<le> Class D"
+ by (auto simp: list_all2_conv_all_nth)
+ with Null obtain a C' fs where
+ Addr: "stk!n = Addr a" and
+ obj: "h a = Some (C',fs)" and
+ C'subD: "P \<turnstile> C' \<preceq>\<^sup>* D"
+ by (fastforce dest!: conf_ClassD)
+
+ with wfprog m_D no_xcp
+ obtain Ts' T' D'' mxs' mxl' ins' xt' where
+ m_C': "P \<turnstile> C' sees M',NonStatic: Ts'\<rightarrow>T' = (mxs',mxl',ins',xt') in D''" and
+ T': "P \<turnstile> T' \<le> T" and
+ Ts': "P \<turnstile> Ts [\<le>] Ts'"
+ by (auto dest: sees_method_mono)
+ with wf_NonStatic_nclinit wtprog have nclinit: "M' \<noteq> clinit" by(simp add: wf_jvm_prog_phi_def)
+
+ have D''subD': "P \<turnstile> D'' \<preceq>\<^sup>* D'" by(rule sees_method_decl_mono[OF C'subD m_D m_C'])
+
+ let ?loc' = "Addr a # rev (take n stk) @ replicate mxl' undefined"
+ let ?f' = "([], ?loc', D'', M', 0, No_ics)"
+ let ?f = "(stk, loc, C, M, pc, ics)"
+
+ from Addr obj m_C' ins \<sigma>' meth_C no_xcp
+ have s': "\<sigma>' = (None, h, ?f' # ?f # frs, sh)" by simp
+
+ from Ts n have [simp]: "size Ts = n"
+ by (auto dest: list_all2_lengthD simp: min_def)
+ with Ts' have [simp]: "size Ts' = n"
+ by (auto dest: list_all2_lengthD)
+
+ from m_C' wfprog
+ obtain mD'': "P \<turnstile> D'' sees M',NonStatic:Ts'\<rightarrow>T'=(mxs',mxl',ins',xt') in D''"
+ by (fast dest: sees_method_idemp)
+ moreover
+ with wtprog
+ obtain start: "wt_start P D'' NonStatic Ts' mxl' (\<Phi> D'' M')" and ins': "ins' \<noteq> []"
+ by (auto dest: wt_jvm_prog_impl_wt_start)
+ then obtain LT\<^sub>0 where LT\<^sub>0: "\<Phi> D'' M' ! 0 = Some ([], LT\<^sub>0)"
+ by (clarsimp simp: wt_start_def defs1 sup_state_opt_any_Some split: staticb.splits)
+ moreover
+ have "conf_f P h sh ([], LT\<^sub>0) ins' ?f'"
+ proof -
+ let ?LT = "OK (Class D'') # (map OK Ts') @ (replicate mxl' Err)"
+
+ from stk have "P,h \<turnstile> take n stk [:\<le>] take n ST" ..
+ hence "P,h \<turnstile> rev (take n stk) [:\<le>] rev (take n ST)" by simp
+ also note Ts also note Ts' finally
+ have "P,h \<turnstile> rev (take n stk) [:\<le>\<^sub>\<top>] map OK Ts'" by simp
+ also
+ have "P,h \<turnstile> replicate mxl' undefined [:\<le>\<^sub>\<top>] replicate mxl' Err"
+ by simp
+ also from m_C' have "P \<turnstile> C' \<preceq>\<^sup>* D''" by (rule sees_method_decl_above)
+ with obj have "P,h \<turnstile> Addr a :\<le> Class D''" by (simp add: conf_def)
+ ultimately
+ have "P,h \<turnstile> ?loc' [:\<le>\<^sub>\<top>] ?LT" by simp
+ also from start LT\<^sub>0 have "P \<turnstile> \<dots> [\<le>\<^sub>\<top>] LT\<^sub>0" by (simp add: wt_start_def)
+ finally have "P,h \<turnstile> ?loc' [:\<le>\<^sub>\<top>] LT\<^sub>0" .
+ thus ?thesis using ins' nclinit by simp
+ qed
+ moreover
+ have "conf_clinit P sh (?f'#?f#frs)" using conf_clinit_Invoke[OF confc nclinit] by simp
+ ultimately
+ have ?thesis using s' \<Phi>_pc approx meth_C m_D T' ins D nclinit D''subD'
+ by(fastforce dest: sees_method_fun [of _ C])
+ }
+ ultimately show ?thesis by blast
+qed
+(*>*)
+
+lemma Invokestatic_nInit_correct:
+ fixes \<sigma>' :: jvm_state
+ assumes wtprog: "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P"
+ assumes meth_C: "P \<turnstile> C sees M,b:Ts\<rightarrow>T=(mxs,mxl\<^sub>0,ins,xt) in C"
+ assumes ins: "ins ! pc = Invokestatic D M' n" and nclinit: "M' \<noteq> clinit"
+ assumes wti: "P,T,mxs,size ins,xt \<turnstile> ins!pc,pc :: \<Phi> C M"
+ assumes \<sigma>': "Some \<sigma>' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)"
+ assumes approx: "P,\<Phi> \<turnstile> (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\<surd>"
+ assumes no_xcp: "fst (exec_step P h stk loc C M pc ics frs sh) = None"
+ assumes cs: "ics = Called [] \<or> (ics = No_ics \<and> (\<exists>sfs. sh (fst(method P D M')) = Some(sfs, Done)))"
+ shows "P,\<Phi> \<turnstile> \<sigma>'\<surd>"
+(*<*)
+proof -
+ note split_paired_Ex [simp del]
+
+ from wtprog obtain wfmb where wfprog: "wf_prog wfmb P"
+ by (simp add: wf_jvm_prog_phi_def)
+
+ from ins meth_C approx obtain ST LT where
+ heap_ok: "P\<turnstile> h\<surd>" and
+ \<Phi>_pc: "\<Phi> C M!pc = Some (ST,LT)" and
+ frame: "conf_f P h sh (ST,LT) ins (stk,loc,C,M,pc,ics)" and
+ frames: "conf_fs P h sh \<Phi> C M (size Ts) T frs" and
+ confc: "conf_clinit P sh ((stk,loc,C,M,pc,ics)#frs)"
+ by (fastforce dest: sees_method_fun)
+
+ from ins wti \<Phi>_pc have n: "n \<le> size ST" by simp
+
+ from ins wti \<Phi>_pc obtain D' b Ts T mxs' mxl' ins' xt' ST' LT' where
+ pc': "pc+1 < size ins" and
+ m_D: "P \<turnstile> D sees M',b: Ts\<rightarrow>T = (mxs',mxl',ins',xt') in D'" and
+ Ts: "P \<turnstile> rev (take n ST) [\<le>] Ts" and
+ \<Phi>': "\<Phi> C M ! (pc+1) = Some (ST', LT')" and
+ LT': "P \<turnstile> LT [\<le>\<^sub>\<top>] LT'" and
+ ST': "P \<turnstile> (T # drop n ST) [\<le>] ST'" and
+ b[simp]: "b = Static"
+ by (clarsimp simp: sup_state_opt_any_Some)
+
+ from frame obtain
+ stk: "P,h \<turnstile> stk [:\<le>] ST" and
+ loc: "P,h \<turnstile> loc [:\<le>\<^sub>\<top>] LT" by simp
+
+ let ?loc' = "rev (take n stk) @ replicate mxl' undefined"
+ let ?f' = "([], ?loc', D', M', 0, No_ics)"
+ let ?f = "(stk, loc, C, M, pc, No_ics)"
+
+ from m_D ins \<sigma>' meth_C no_xcp cs
+ have s': "\<sigma>' = (None, h, ?f' # ?f # frs, sh)" by auto
+
+ from Ts n have [simp]: "size Ts = n"
+ by (auto dest: list_all2_lengthD)
+
+ from m_D wfprog b
+ obtain mD': "P \<turnstile> D' sees M',Static:Ts\<rightarrow>T=(mxs',mxl',ins',xt') in D'"
+ by (fast dest: sees_method_idemp)
+ moreover
+ with wtprog
+ obtain start: "wt_start P D' Static Ts mxl' (\<Phi> D' M')" and ins': "ins' \<noteq> []"
+ by (auto dest: wt_jvm_prog_impl_wt_start)
+ then obtain LT\<^sub>0 where LT\<^sub>0: "\<Phi> D' M' ! 0 = Some ([], LT\<^sub>0)"
+ by (clarsimp simp: wt_start_def defs1 sup_state_opt_any_Some split: staticb.splits)
+ moreover
+ have "conf_f P h sh ([], LT\<^sub>0) ins' ?f'"
+ proof -
+ let ?LT = "(map OK Ts) @ (replicate mxl' Err)"
+
+ from stk have "P,h \<turnstile> take n stk [:\<le>] take n ST" ..
+ hence "P,h \<turnstile> rev (take n stk) [:\<le>] rev (take n ST)" by simp
+ also note Ts finally
+ have "P,h \<turnstile> rev (take n stk) [:\<le>\<^sub>\<top>] map OK Ts" by simp
+ also
+ have "P,h \<turnstile> replicate mxl' undefined [:\<le>\<^sub>\<top>] replicate mxl' Err"
+ by simp
+ also from m_D have "P \<turnstile> D \<preceq>\<^sup>* D'" by (rule sees_method_decl_above)
+ ultimately
+ have "P,h \<turnstile> ?loc' [:\<le>\<^sub>\<top>] ?LT" by simp
+ also from start LT\<^sub>0 have "P \<turnstile> \<dots> [\<le>\<^sub>\<top>] LT\<^sub>0" by (simp add: wt_start_def)
+ finally have "P,h \<turnstile> ?loc' [:\<le>\<^sub>\<top>] LT\<^sub>0" .
+ thus ?thesis using ins' by simp
+ qed
+ moreover
+ have "conf_clinit P sh (?f'#?f#frs)" by(rule conf_clinit_Invoke[OF confc nclinit])
+ ultimately
+ show ?thesis using s' \<Phi>_pc approx meth_C m_D ins nclinit
+ by (fastforce dest: sees_method_fun [of _ C])
+qed
+(*>*)
+
+lemma Invokestatic_Init_correct:
+ fixes \<sigma>' :: jvm_state
+ assumes wtprog: "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P"
+ assumes meth_C: "P \<turnstile> C sees M,b:Ts\<rightarrow>T=(mxs,mxl\<^sub>0,ins,xt) in C"
+ assumes ins: "ins ! pc = Invokestatic D M' n" and nclinit: "M' \<noteq> clinit"
+ assumes wti: "P,T,mxs,size ins,xt \<turnstile> ins!pc,pc :: \<Phi> C M"
+ assumes \<sigma>': "Some \<sigma>' = exec (P, None, h, (stk,loc,C,M,pc,No_ics)#frs, sh)"
+ assumes approx: "P,\<Phi> \<turnstile> (None, h, (stk,loc,C,M,pc,No_ics)#frs, sh)\<surd>"
+ assumes no_xcp: "fst (exec_step P h stk loc C M pc No_ics frs sh) = None"
+ assumes nDone: "\<forall>sfs. sh (fst(method P D M')) \<noteq> Some(sfs, Done)"
+ shows "P,\<Phi> \<turnstile> \<sigma>'\<surd>"
+(*<*)
+proof -
+ note split_paired_Ex [simp del]
+
+ from wtprog obtain wfmb where wfprog: "wf_prog wfmb P"
+ by (simp add: wf_jvm_prog_phi_def)
+
+ from ins meth_C approx obtain ST LT where
+ heap_ok: "P\<turnstile> h\<surd>" and
+ \<Phi>_pc: "\<Phi> C M!pc = Some (ST,LT)" and
+ stk: "P,h \<turnstile> stk [:\<le>] ST" and loc: "P,h \<turnstile> loc [:\<le>\<^sub>\<top>] LT" and
+ pc: "pc < size ins" and
+ frames: "conf_fs P h sh \<Phi> C M (size Ts) T frs" and
+ confc: "conf_clinit P sh ((stk,loc,C,M,pc,No_ics)#frs)" and
+ pc: "pc < size ins"
+ by (fastforce dest: sees_method_fun)
+
+ from ins wti \<Phi>_pc obtain D' b Ts T mxs' mxl' ins' xt' where
+ m_D: "P \<turnstile> D sees M',b: Ts\<rightarrow>T = (mxs',mxl',ins',xt') in D'" and
+ b[simp]: "b = Static"
+ by clarsimp
+
+ let ?f = "(stk, loc, C, M, pc, Calling D' [])"
+
+ from m_D ins \<sigma>' meth_C no_xcp nDone
+ have s': "\<sigma>' = (None, h, ?f # frs, sh)" by(auto split: init_state.splits)
+
+ have cls: "is_class P D'" by(rule sees_method_is_class'[OF m_D])
+
+ from confc have confc': "conf_clinit P sh (?f#frs)"
+ by(auto simp: conf_clinit_def distinct_clinit_def split: if_split_asm)
+ with s' \<Phi>_pc approx meth_C m_D ins nclinit stk loc pc cls frames
+ show ?thesis by(fastforce dest: sees_method_fun [of _ C])
+qed
+(*>*)
+
+declare list_all2_Cons2 [iff]
+
+lemma Return_correct:
+ fixes \<sigma>' :: jvm_state
+ assumes wt_prog: "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P"
+ assumes meth: "P \<turnstile> C sees M,b:Ts\<rightarrow>T=(mxs,mxl\<^sub>0,ins,xt) in C"
+ assumes ins: "ins ! pc = Return"
+ assumes wt: "P,T,mxs,size ins,xt \<turnstile> ins!pc,pc :: \<Phi> C M"
+ assumes s': "Some \<sigma>' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)"
+ assumes correct: "P,\<Phi> \<turnstile> (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\<surd>"
+
+ shows "P,\<Phi> \<turnstile> \<sigma>'\<surd>"
+(*<*)
+proof -
+ from meth correct ins have [simp]: "ics = No_ics" by(cases ics, auto)
+
+ from wt_prog
+ obtain wfmb where wf: "wf_prog wfmb P" by (simp add: wf_jvm_prog_phi_def)
+
+ from meth ins s'
+ have "frs = [] \<Longrightarrow> ?thesis" by (simp add: correct_state_def)
+ moreover
+ { fix f frs' assume frs': "frs = f#frs'"
+ then obtain stk' loc' C' M' pc' ics' where
+ f: "f = (stk',loc',C',M',pc',ics')" by (cases f)
+
+ from correct meth
+ obtain ST LT where
+ h_ok: "P \<turnstile> h \<surd>" and
+ sh_ok: "P,h \<turnstile>\<^sub>s sh \<surd>" and
+ \<Phi>_pc: "\<Phi> C M ! pc = Some (ST, LT)" and
+ frame: "conf_f P h sh (ST, LT) ins (stk,loc,C,M,pc,ics)" and
+ frames: "conf_fs P h sh \<Phi> C M (size Ts) T frs" and
+ confc: "conf_clinit P sh frs"
+ by (auto dest: sees_method_fun conf_clinit_Cons simp: correct_state_def)
+
+ from \<Phi>_pc ins wt
+ obtain U ST\<^sub>0 where "ST = U # ST\<^sub>0" "P \<turnstile> U \<le> T"
+ by (simp add: wt_instr_def app_def) blast
+ with wf frame
+ have hd_stk: "P,h \<turnstile> hd stk :\<le> T" by (auto simp: conf_f_def)
+
+ from f frs' frames meth
+ obtain ST' LT' b' Ts'' T'' mxs' mxl\<^sub>0' ins' xt' where
+ \<Phi>': "\<Phi> C' M' ! pc' = Some (ST', LT')" and
+ meth_C': "P \<turnstile> C' sees M',b':Ts''\<rightarrow>T''=(mxs',mxl\<^sub>0',ins',xt') in C'" and
+ frame': "conf_f P h sh (ST',LT') ins' f" and
+ conf_fs: "conf_fs P h sh \<Phi> C' M' (size Ts'') T'' frs'"
+ by clarsimp
+
+ from f frame' obtain
+ stk': "P,h \<turnstile> stk' [:\<le>] ST'" and
+ loc': "P,h \<turnstile> loc' [:\<le>\<^sub>\<top>] LT'" and
+ pc': "pc' < size ins'"
+ by (simp add: conf_f_def)
+
+ { assume b[simp]: "b = NonStatic"
+
+ from wf_NonStatic_nclinit[OF wf] meth have nclinit[simp]: "M \<noteq> clinit" by simp
+
+ from f frs' meth ins s'
+ have \<sigma>':
+ "\<sigma>' = (None,h,(hd stk#(drop (1+size Ts) stk'),loc',C',M',pc'+1,ics')#frs',sh)"
+ (is "\<sigma>' = (None,h,?f'#frs',sh)")
+ by simp
+ from f frs' confc conf_clinit_diff have confc'': "conf_clinit P sh (?f'#frs')" by blast
+
+ with \<Phi>' meth_C' f frs' frames meth
+ obtain D Ts' T' m D' where
+ ins': "ins' ! pc' = Invoke M (size Ts)" and
+ D: "ST' ! (size Ts) = Class D" and
+ meth_D: "P \<turnstile> D sees M,b: Ts'\<rightarrow>T' = m in D'" and
+ T': "P \<turnstile> T \<le> T'" and
+ CsubD': "P \<turnstile> C \<preceq>\<^sup>* D'"
+ by(auto dest: sees_method_fun sees_method_fun[OF sees_method_idemp])
+
+ from wt_prog meth_C' pc'
+ have "P,T'',mxs',size ins',xt' \<turnstile> ins'!pc',pc' :: \<Phi> C' M'"
+ by (rule wt_jvm_prog_impl_wt_instr)
+ with ins' \<Phi>' D meth_D
+ obtain ST'' LT'' where
+ \<Phi>_suc: "\<Phi> C' M' ! Suc pc' = Some (ST'', LT'')" and
+ less: "P \<turnstile> (T' # drop (size Ts+1) ST', LT') \<le>\<^sub>i (ST'', LT'')" and
+ suc_pc': "Suc pc' < size ins'"
+ by (clarsimp simp: sup_state_opt_any_Some)
+
+ from hd_stk T' have hd_stk': "P,h \<turnstile> hd stk :\<le> T'" ..
+
+ have frame'':
+ "conf_f P h sh (ST'',LT'') ins' ?f'"
+ proof -
+ from stk'
+ have "P,h \<turnstile> drop (1+size Ts) stk' [:\<le>] drop (1+size Ts) ST'" ..
+ moreover
+ with hd_stk' less
+ have "P,h \<turnstile> hd stk # drop (1+size Ts) stk' [:\<le>] ST''" by auto
+ moreover
+ from wf loc' less have "P,h \<turnstile> loc' [:\<le>\<^sub>\<top>] LT''" by auto
+ moreover note suc_pc'
+ moreover
+ from f frs' frames (* ics' = No_ics *)
+ have "P,h,sh \<turnstile>\<^sub>i (C', M', Suc pc', ics')" by auto
+ ultimately show ?thesis by (simp add: conf_f_def)
+ qed
+
+ with \<sigma>' frs' f meth h_ok sh_ok hd_stk \<Phi>_suc frames confc'' meth_C' \<Phi>'
+ have ?thesis by(fastforce dest: sees_method_fun [of _ C'])
+ }
+ moreover
+ { assume b[simp]: "b = Static" and nclinit[simp]: "M \<noteq> clinit"
+
+ from f frs' meth ins s'
+ have \<sigma>':
+ "\<sigma>' = (None,h,(hd stk#(drop (size Ts) stk'),loc',C',M',pc'+1,ics')#frs',sh)"
+ (is "\<sigma>' = (None,h,?f'#frs',sh)")
+ by simp
+ from f frs' confc conf_clinit_diff have confc'': "conf_clinit P sh (?f'#frs')" by blast
+
+ with \<Phi>' meth_C' f frs' frames meth
+ obtain D Ts' T' m where
+ ins': "ins' ! pc' = Invokestatic D M (size Ts)" and
+ meth_D: "P \<turnstile> D sees M,b: Ts'\<rightarrow>T' = m in C" and
+ T': "P \<turnstile> T \<le> T'"
+ by(auto dest: sees_method_fun sees_method_mono2[OF _ wf sees_method_idemp])
+
+ from wt_prog meth_C' pc'
+ have "P,T'',mxs',size ins',xt' \<turnstile> ins'!pc',pc' :: \<Phi> C' M'"
+ by (rule wt_jvm_prog_impl_wt_instr)
+ with ins' \<Phi>' meth_D
+ obtain ST'' LT'' where
+ \<Phi>_suc: "\<Phi> C' M' ! Suc pc' = Some (ST'', LT'')" and
+ less: "P \<turnstile> (T' # drop (size Ts) ST', LT') \<le>\<^sub>i (ST'', LT'')" and
+ suc_pc': "Suc pc' < size ins'"
+ by (clarsimp simp: sup_state_opt_any_Some)
+
+ from hd_stk T' have hd_stk': "P,h \<turnstile> hd stk :\<le> T'" ..
+
+ have frame'':
+ "conf_f P h sh (ST'',LT'') ins' ?f'"
+ proof -
+ from stk'
+ have "P,h \<turnstile> drop (size Ts) stk' [:\<le>] drop (size Ts) ST'" ..
+ moreover
+ with hd_stk' less
+ have "P,h \<turnstile> hd stk # drop (size Ts) stk' [:\<le>] ST''" by auto
+ moreover
+ from wf loc' less have "P,h \<turnstile> loc' [:\<le>\<^sub>\<top>] LT''" by auto
+ moreover note suc_pc'
+ moreover
+ from f frs' frames (* ics' = No_ics *)
+ have "P,h,sh \<turnstile>\<^sub>i (C', M', Suc pc', ics')" by auto
+ ultimately show ?thesis by (simp add: conf_f_def)
+ qed
+
+ with \<sigma>' frs' f meth h_ok sh_ok hd_stk \<Phi>_suc frames confc'' meth_C' \<Phi>'
+ have ?thesis by(fastforce dest: sees_method_fun [of _ C'])
+ }
+ moreover
+ { assume b[simp]: "b = Static" and clinit[simp]: "M = clinit"
+
+ from frs' meth ins s'
+ have \<sigma>':
+ "\<sigma>' = (None,h,frs,sh(C\<mapsto>(fst(the(sh C)), Done)))" (is "\<sigma>' = (None,h,frs,?sh)")
+ by simp
+
+ from correct have dist': "distinct (C # clinit_classes frs)"
+ by(simp add: conf_clinit_def distinct_clinit_def)
+
+ from f frs' correct have confc1:
+ "conf_clinit P sh ((stk, loc, C, clinit, pc, No_ics) # (stk',loc',C',M',pc',ics') # frs')"
+ by simp
+ then have ics_dist: "distinct (C # ics_classes ics')"
+ by(simp add: conf_clinit_def distinct_clinit_def)
+
+ from conf_clinit_Cons_Cons[OF confc1] have dist'': "distinct (C # clinit_classes frs')"
+ by(simp add: conf_clinit_def distinct_clinit_def)
+
+ from correct shconf_upd_obj[OF sh_ok _ [OF shconfD[OF sh_ok]]]
+ have sh'_ok: "P,h \<turnstile>\<^sub>s ?sh \<surd>" by(clarsimp simp: conf_clinit_def)
+
+ have frame'':
+ "conf_f P h ?sh (ST',LT') ins' f"
+ proof -
+ note stk' loc' pc' f valid_ics_shupd[OF _ ics_dist]
+ moreover
+ from f frs' frames
+ have "P,h,sh \<turnstile>\<^sub>i (C', M', pc', ics')" by auto
+ ultimately show ?thesis by (simp add: conf_f_def2)
+ qed
+ have conf_fs': "conf_fs P h ?sh \<Phi> C' M' (length Ts'') T'' frs'"
+ by(rule conf_fs_shupd[OF conf_fs dist''])
+
+ have confc'': "conf_clinit P ?sh frs" by(rule conf_clinit_shupd[OF confc dist'])
+
+ from \<sigma>' f frs' h_ok sh'_ok conf_fs' frame'' \<Phi>' stk' loc' pc' meth_C' confc''
+ have ?thesis by(fastforce dest: sees_method_fun)
+ }
+ ultimately have ?thesis by (cases b) blast+
+ }
+ ultimately
+ show ?thesis by (cases frs) blast+
+qed
+(*>*)
+
+declare sup_state_opt_any_Some [iff]
+declare not_Err_eq [iff]
+
+lemma Load_correct:
+assumes "wf_prog wt P" and
+ mC: "P \<turnstile> C sees M,b:Ts\<rightarrow>T=(mxs,mxl\<^sub>0,ins,xt) in C" and
+ i: "ins!pc = Load idx" and
+ "P,T,mxs,size ins,xt \<turnstile> ins!pc,pc :: \<Phi> C M" and
+ "Some \<sigma>' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)" and
+ cf: "P,\<Phi> \<turnstile> (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\<surd>"
+shows "P,\<Phi> |- \<sigma>' [ok]"
+(*<*)
+proof -
+ have "ics = No_ics" using mC i cf by(cases ics) auto
+ then show ?thesis using assms(1,3-6) sees_method_fun[OF mC]
+ by(fastforce elim!: confTs_confT_sup conf_clinit_diff)
+qed
+(*>*)
+
+declare [[simproc del: list_to_set_comprehension]]
+
+lemma Store_correct:
+assumes "wf_prog wt P" and
+ mC: "P \<turnstile> C sees M,b:Ts\<rightarrow>T=(mxs,mxl\<^sub>0,ins,xt) in C" and
+ i: "ins!pc = Store idx" and
+ "P,T,mxs,size ins,xt \<turnstile> ins!pc,pc :: \<Phi> C M" and
+ "Some \<sigma>' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)" and
+ cf: "P,\<Phi> \<turnstile> (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\<surd>"
+shows "P,\<Phi> |- \<sigma>' [ok]"
+(*<*)
+proof -
+ have "ics = No_ics" using mC i cf by(cases ics) auto
+ then show ?thesis using assms(1,3-6) sees_method_fun[OF mC]
+ by clarsimp (blast intro!: list_all2_update_cong conf_clinit_diff)
+qed
+(*>*)
+
+
+lemma Push_correct:
+assumes "wf_prog wt P" and
+ mC: "P \<turnstile> C sees M,b:Ts\<rightarrow>T=(mxs,mxl\<^sub>0,ins,xt) in C" and
+ i: "ins!pc = Push v" and
+ "P,T,mxs,size ins,xt \<turnstile> ins!pc,pc :: \<Phi> C M" and
+ "Some \<sigma>' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)" and
+ cf: "P,\<Phi> \<turnstile> (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\<surd>"
+shows "P,\<Phi> |- \<sigma>' [ok]"
+(*<*)
+proof -
+ have "ics = No_ics" using mC i cf by(cases ics) auto
+ then show ?thesis using assms(1,3-6) sees_method_fun[OF mC]
+ by clarsimp (blast dest: typeof_lit_conf conf_clinit_diff)
+qed
+
+
+lemma Cast_conf2:
+ "\<lbrakk> wf_prog ok P; P,h \<turnstile> v :\<le> T; is_refT T; cast_ok P C h v;
+ P \<turnstile> Class C \<le> T'; is_class P C\<rbrakk>
+ \<Longrightarrow> P,h \<turnstile> v :\<le> T'"
+(*<*)
+proof -
+ assume "wf_prog ok P" and "P,h \<turnstile> v :\<le> T" and "is_refT T" and
+ "cast_ok P C h v" and wid: "P \<turnstile> Class C \<le> T'" and "is_class P C"
+ then show "P,h \<turnstile> v :\<le> T'" using Class_widen[OF wid]
+ by(cases v)
+ (auto simp: cast_ok_def is_refT_def conf_def obj_ty_def
+ intro: rtrancl_trans)
+qed
+(*>*)
+
+
+lemma Checkcast_correct:
+assumes "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P" and
+ mC: "P \<turnstile> C sees M,b:Ts\<rightarrow>T=(mxs,mxl\<^sub>0,ins,xt) in C" and
+ i: "ins!pc = Checkcast D" and
+ "P,T,mxs,size ins,xt \<turnstile> ins!pc,pc :: \<Phi> C M" and
+ "Some \<sigma>' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)" and
+ cf: "P,\<Phi> \<turnstile> (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\<surd>" and
+ "fst (exec_step P h stk loc C M pc ics frs sh) = None"
+shows "P,\<Phi> |- \<sigma>' [ok]"
+(*<*)
+proof -
+ have "ics = No_ics" using mC i cf by(cases ics) auto
+ then show ?thesis using assms
+ by (clarsimp simp: wf_jvm_prog_phi_def split: if_split_asm)
+ (blast intro: Cast_conf2 dest: sees_method_fun conf_clinit_diff)
+qed
+(*>*)
+
+declare split_paired_All [simp del]
+
+lemmas widens_Cons [iff] = list_all2_Cons1 [of "widen P"] for P
+
+lemma Getfield_correct:
+ fixes \<sigma>' :: jvm_state
+ assumes wf: "wf_prog wt P"
+ assumes mC: "P \<turnstile> C sees M,b:Ts\<rightarrow>T=(mxs,mxl\<^sub>0,ins,xt) in C"
+ assumes i: "ins!pc = Getfield F D"
+ assumes wt: "P,T,mxs,size ins,xt \<turnstile> ins!pc,pc :: \<Phi> C M"
+ assumes s': "Some \<sigma>' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)"
+ assumes cf: "P,\<Phi> \<turnstile> (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\<surd>"
+ assumes xc: "fst (exec_step P h stk loc C M pc ics frs sh) = None"
+
+ shows "P,\<Phi> \<turnstile> \<sigma>'\<surd>"
+(*<*)
+proof -
+ from mC cf i have [simp]: "ics = No_ics" by(cases ics, auto)
+
+ from mC cf obtain ST LT where
+ "h\<surd>": "P \<turnstile> h \<surd>" and
+ "sh\<surd>": "P,h \<turnstile>\<^sub>s sh \<surd>" and
+ \<Phi>: "\<Phi> C M ! pc = Some (ST,LT)" and
+ stk: "P,h \<turnstile> stk [:\<le>] ST" and loc: "P,h \<turnstile> loc [:\<le>\<^sub>\<top>] LT" and
+ pc: "pc < size ins" and
+ fs: "conf_fs P h sh \<Phi> C M (size Ts) T frs" and
+ confc: "conf_clinit P sh ((stk,loc,C,M,pc,ics)#frs)"
+ by (fastforce dest: sees_method_fun)
+
+ from i \<Phi> wt obtain oT ST'' vT ST' LT' vT' where
+ oT: "P \<turnstile> oT \<le> Class D" and
+ ST: "ST = oT # ST''" and
+ F: "P \<turnstile> D sees F,NonStatic:vT in D" and
+ pc': "pc+1 < size ins" and
+ \<Phi>': "\<Phi> C M ! (pc+1) = Some (vT'#ST', LT')" and
+ ST': "P \<turnstile> ST'' [\<le>] ST'" and LT': "P \<turnstile> LT [\<le>\<^sub>\<top>] LT'" and
+ vT': "P \<turnstile> vT \<le> vT'"
+ by fastforce
+
+ from stk ST obtain ref stk' where
+ stk': "stk = ref#stk'" and
+ ref: "P,h \<turnstile> ref :\<le> oT" and
+ ST'': "P,h \<turnstile> stk' [:\<le>] ST''"
+ by auto
+
+ from stk' i mC s' xc have "ref \<noteq> Null"
+ by (simp add: split_beta split:if_split_asm)
+ moreover from ref oT have "P,h \<turnstile> ref :\<le> Class D" ..
+ ultimately obtain a D' fs where
+ a: "ref = Addr a" and h: "h a = Some (D', fs)" and D': "P \<turnstile> D' \<preceq>\<^sup>* D"
+ by (blast dest: non_npD)
+
+ from D' F have has_field: "P \<turnstile> D' has F,NonStatic:vT in D"
+ by (blast intro: has_field_mono has_visible_field)
+ moreover from "h\<surd>" h have "P,h \<turnstile> (D', fs) \<surd>" by (rule hconfD)
+ ultimately obtain v where v: "fs (F, D) = Some v" "P,h \<turnstile> v :\<le> vT"
+ by (clarsimp simp: oconf_def has_field_def)
+ (blast dest: has_fields_fun)
+
+ from conf_clinit_diff[OF confc]
+ have confc': "conf_clinit P sh ((v#stk',loc,C,M,pc+1,ics)#frs)" by simp
+
+ from a h i mC s' stk' v xc has_field
+ have "\<sigma>' = (None, h, (v#stk',loc,C,M,pc+1,ics)#frs, sh)"
+ by(simp add: split_beta split: if_split_asm)
+ moreover
+ from ST'' ST' have "P,h \<turnstile> stk' [:\<le>] ST'" ..
+ moreover
+ from v vT' have "P,h \<turnstile> v :\<le> vT'" by blast
+ moreover
+ from loc LT' have "P,h \<turnstile> loc [:\<le>\<^sub>\<top>] LT'" ..
+ moreover
+ note "h\<surd>" "sh\<surd>" mC \<Phi>' pc' v fs confc'
+ ultimately
+ show "P,\<Phi> \<turnstile> \<sigma>' \<surd>" by fastforce
+qed
+(*>*)
+
+lemma Getstatic_nInit_correct:
+ fixes \<sigma>' :: jvm_state
+ assumes wf: "wf_prog wt P"
+ assumes mC: "P \<turnstile> C sees M,b:Ts\<rightarrow>T=(mxs,mxl\<^sub>0,ins,xt) in C"
+ assumes i: "ins!pc = Getstatic C' F D"
+ assumes wt: "P,T,mxs,size ins,xt \<turnstile> ins!pc,pc :: \<Phi> C M"
+ assumes s': "Some \<sigma>' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)"
+ assumes cf: "P,\<Phi> \<turnstile> (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\<surd>"
+ assumes xc: "fst (exec_step P h stk loc C M pc ics frs sh) = None"
+ assumes cs: "ics = Called [] \<or> (ics = No_ics \<and> (\<exists>sfs. sh (fst(field P D F)) = Some(sfs, Done)))"
+
+ shows "P,\<Phi> \<turnstile> \<sigma>'\<surd>"
+(*<*)
+proof -
+ from mC cf obtain ST LT where
+ "h\<surd>": "P \<turnstile> h \<surd>" and
+ "sh\<surd>": "P,h \<turnstile>\<^sub>s sh \<surd>" and
+ \<Phi>: "\<Phi> C M ! pc = Some (ST,LT)" and
+ stk: "P,h \<turnstile> stk [:\<le>] ST" and loc: "P,h \<turnstile> loc [:\<le>\<^sub>\<top>] LT" and
+ pc: "pc < size ins" and
+ fs: "conf_fs P h sh \<Phi> C M (size Ts) T frs" and
+ confc: "conf_clinit P sh ((stk,loc,C,M,pc,ics)#frs)" and
+ vics: "P,h,sh \<turnstile>\<^sub>i (C,M,pc,ics)"
+ by (fastforce dest: sees_method_fun)
+
+ from i \<Phi> wt cs obtain vT ST' LT' vT' where
+ F: "P \<turnstile> C' sees F,Static:vT in D" and
+ pc': "pc+1 < size ins" and
+ \<Phi>': "\<Phi> C M ! (pc+1) = Some (vT'#ST', LT')" and
+ ST': "P \<turnstile> ST [\<le>] ST'" and LT': "P \<turnstile> LT [\<le>\<^sub>\<top>] LT'" and
+ vT': "P \<turnstile> vT \<le> vT'"
+ by fastforce
+
+ with mC i vics obtain sobj where
+ cc': "ics = Called [] \<Longrightarrow> Called_context P D (ins!pc) \<and> sh D = Some sobj"
+ by(fastforce dest: has_visible_field)
+
+ from field_def2[OF sees_field_idemp[OF F]] have D[simp]: "fst(field P D F) = D" by simp
+ from cs cc' obtain sfs i where shD: "sh D = Some(sfs,i)" by(cases sobj, auto)
+
+ note has_field_idemp[OF has_visible_field[OF F]]
+ moreover from "sh\<surd>" shD have "P,h,D \<turnstile>\<^sub>s sfs \<surd>" by (rule shconfD)
+ ultimately obtain v where v: "sfs F = Some v" "P,h \<turnstile> v :\<le> vT"
+ by (clarsimp simp: soconf_def has_field_def) blast
+
+ from i mC s' v xc F cs cc' shD
+ have "\<sigma>' = (None, h, (v#stk,loc,C,M,pc+1,No_ics)#frs, sh)"
+ by(fastforce simp: split_beta split: if_split_asm init_call_status.splits)
+ moreover
+ from stk ST' have "P,h \<turnstile> stk [:\<le>] ST'" ..
+ moreover
+ from v vT' have "P,h \<turnstile> v :\<le> vT'" by blast
+ moreover
+ from loc LT' have "P,h \<turnstile> loc [:\<le>\<^sub>\<top>] LT'" ..
+ moreover
+ have "conf_clinit P sh ((v#stk,loc,C,M,pc+1,No_ics)#frs)" by(rule conf_clinit_diff'[OF confc])
+ moreover
+ note "h\<surd>" "sh\<surd>" mC \<Phi>' pc' v fs
+ ultimately
+ show "P,\<Phi> \<turnstile> \<sigma>' \<surd>" by fastforce
+qed
+(*>*)
+
+lemma Getstatic_Init_correct:
+ fixes \<sigma>' :: jvm_state
+ assumes wf: "wf_prog wt P"
+ assumes mC: "P \<turnstile> C sees M,b:Ts\<rightarrow>T=(mxs,mxl\<^sub>0,ins,xt) in C"
+ assumes i: "ins!pc = Getstatic C' F D"
+ assumes wt: "P,T,mxs,size ins,xt \<turnstile> ins!pc,pc :: \<Phi> C M"
+ assumes s': "Some \<sigma>' = exec (P, None, h, (stk,loc,C,M,pc,No_ics)#frs, sh)"
+ assumes cf: "P,\<Phi> \<turnstile> (None, h, (stk,loc,C,M,pc,No_ics)#frs, sh)\<surd>"
+ assumes xc: "fst (exec_step P h stk loc C M pc No_ics frs sh) = None"
+ assumes nDone: "\<forall>sfs. sh (fst(field P D F)) \<noteq> Some(sfs, Done)"
+
+ shows "P,\<Phi> \<turnstile> \<sigma>'\<surd>"
+(*<*)
+proof -
+ from mC cf obtain ST LT where
+ "h\<surd>": "P \<turnstile> h \<surd>" and
+ "sh\<surd>": "P,h \<turnstile>\<^sub>s sh \<surd>" and
+ \<Phi>: "\<Phi> C M ! pc = Some (ST,LT)" and
+ stk: "P,h \<turnstile> stk [:\<le>] ST" and loc: "P,h \<turnstile> loc [:\<le>\<^sub>\<top>] LT" and
+ pc: "pc < size ins" and
+ fs: "conf_fs P h sh \<Phi> C M (size Ts) T frs" and
+ confc: "conf_clinit P sh ((stk,loc,C,M,pc,No_ics)#frs)"
+ by (fastforce dest: sees_method_fun)
+
+ from i \<Phi> wt nDone obtain vT where
+ F: "P \<turnstile> C' sees F,Static:vT in D"
+ by fastforce
+ then have has_field: "P \<turnstile> C' has F,Static:vT in D" by(rule has_visible_field)
+
+ from field_def2[OF sees_field_idemp[OF F]] has_field_is_class'[OF has_field] obtain
+ D[simp]: "fst(field P D F) = D" and
+ cls: "is_class P D" by simp
+
+ from i mC s' xc F nDone
+ have "\<sigma>' = (None, h, (stk,loc,C,M,pc,Calling D [])#frs, sh)"
+ by(auto simp: split_beta split: if_split_asm init_state.splits)
+ moreover
+ from confc have "conf_clinit P sh ((stk,loc,C,M,pc,Calling D [])#frs)"
+ by(auto simp: conf_clinit_def distinct_clinit_def split: if_split_asm)
+ moreover
+ note loc stk "h\<surd>" "sh\<surd>" mC \<Phi> pc fs i has_field cls
+ ultimately
+ show "P,\<Phi> \<turnstile> \<sigma>' \<surd>" by fastforce
+qed
+(*>*)
+
+lemma Putfield_correct:
+ fixes \<sigma>' :: jvm_state
+ assumes wf: "wf_prog wt P"
+ assumes mC: "P \<turnstile> C sees M,b:Ts\<rightarrow>T=(mxs,mxl\<^sub>0,ins,xt) in C"
+ assumes i: "ins!pc = Putfield F D"
+ assumes wt: "P,T,mxs,size ins,xt \<turnstile> ins!pc,pc :: \<Phi> C M"
+ assumes s': "Some \<sigma>' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)"
+ assumes cf: "P,\<Phi> \<turnstile> (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\<surd>"
+ assumes xc: "fst (exec_step P h stk loc C M pc ics frs sh) = None"
+
+ shows "P,\<Phi> \<turnstile> \<sigma>'\<surd>"
+(*<*)
+proof -
+ from mC cf i have [simp]: "ics = No_ics" by(cases ics, auto)
+
+ from mC cf obtain ST LT where
+ "h\<surd>": "P \<turnstile> h \<surd>" and
+ "sh\<surd>": "P,h \<turnstile>\<^sub>s sh \<surd>" and
+ \<Phi>: "\<Phi> C M ! pc = Some (ST,LT)" and
+ stk: "P,h \<turnstile> stk [:\<le>] ST" and loc: "P,h \<turnstile> loc [:\<le>\<^sub>\<top>] LT" and
+ pc: "pc < size ins" and
+ fs: "conf_fs P h sh \<Phi> C M (size Ts) T frs" and
+ confc: "conf_clinit P sh ((stk,loc,C,M,pc,ics) # frs)"
+ by (fastforce dest: sees_method_fun)
+
+ from i \<Phi> wt obtain vT vT' oT ST'' ST' LT' where
+ ST: "ST = vT # oT # ST''" and
+ field: "P \<turnstile> D sees F,NonStatic:vT' in D" and
+ oT: "P \<turnstile> oT \<le> Class D" and vT: "P \<turnstile> vT \<le> vT'" and
+ pc': "pc+1 < size ins" and
+ \<Phi>': "\<Phi> C M!(pc+1) = Some (ST',LT')" and
+ ST': "P \<turnstile> ST'' [\<le>] ST'" and LT': "P \<turnstile> LT [\<le>\<^sub>\<top>] LT'"
+ by clarsimp
+
+ from stk ST obtain v ref stk' where
+ stk': "stk = v#ref#stk'" and
+ v: "P,h \<turnstile> v :\<le> vT" and
+ ref: "P,h \<turnstile> ref :\<le> oT" and
+ ST'': "P,h \<turnstile> stk' [:\<le>] ST''"
+ by auto
+
+ from stk' i mC s' xc have "ref \<noteq> Null" by (auto simp: split_beta)
+ moreover from ref oT have "P,h \<turnstile> ref :\<le> Class D" ..
+ ultimately obtain a D' fs where
+ a: "ref = Addr a" and h: "h a = Some (D', fs)" and D': "P \<turnstile> D' \<preceq>\<^sup>* D"
+ by (blast dest: non_npD)
+
+ from v vT have vT': "P,h \<turnstile> v :\<le> vT'" ..
+
+ from field D' have has_field: "P \<turnstile> D' has F,NonStatic:vT' in D"
+ by (blast intro: has_field_mono has_visible_field)
+
+ let ?h' = "h(a\<mapsto>(D', fs((F, D)\<mapsto>v)))" and ?f' = "(stk',loc,C,M,pc+1,ics)"
+ from h have hext: "h \<unlhd> ?h'" by (rule hext_upd_obj)
+
+ have "sh\<surd>'": "P,?h' \<turnstile>\<^sub>s sh \<surd>" by(rule shconf_hupd_obj[OF "sh\<surd>" h])
+
+ from a h i mC s' stk' has_field field
+ have "\<sigma>' = (None, ?h', ?f'#frs, sh)" by(simp split: if_split_asm)
+ moreover
+ from "h\<surd>" h have "P,h \<turnstile> (D',fs)\<surd>" by (rule hconfD)
+ with has_field vT' have "P,h \<turnstile> (D',fs((F, D)\<mapsto>v))\<surd>" ..
+ with "h\<surd>" h have "P \<turnstile> ?h'\<surd>" by (rule hconf_upd_obj)
+ moreover
+ from ST'' ST' have "P,h \<turnstile> stk' [:\<le>] ST'" ..
+ from this hext have "P,?h' \<turnstile> stk' [:\<le>] ST'" by (rule confs_hext)
+ moreover
+ from loc LT' have "P,h \<turnstile> loc [:\<le>\<^sub>\<top>] LT'" ..
+ from this hext have "P,?h' \<turnstile> loc [:\<le>\<^sub>\<top>] LT'" by (rule confTs_hext)
+ moreover
+ from fs hext
+ have "conf_fs P ?h' sh \<Phi> C M (size Ts) T frs" by (rule conf_fs_hext)
+ moreover
+ have "conf_clinit P sh (?f' # frs)" by(rule conf_clinit_diff[OF confc])
+ moreover
+ note mC \<Phi>' pc' "sh\<surd>'"
+ ultimately
+ show "P,\<Phi> \<turnstile> \<sigma>' \<surd>" by fastforce
+qed
+(*>*)
+
+lemma Putstatic_nInit_correct:
+ fixes \<sigma>' :: jvm_state
+ assumes wf: "wf_prog wt P"
+ assumes mC: "P \<turnstile> C sees M,b:Ts\<rightarrow>T=(mxs,mxl\<^sub>0,ins,xt) in C"
+ assumes i: "ins!pc = Putstatic C' F D"
+ assumes wt: "P,T,mxs,size ins,xt \<turnstile> ins!pc,pc :: \<Phi> C M"
+ assumes s': "Some \<sigma>' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)"
+ assumes cf: "P,\<Phi> \<turnstile> (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\<surd>"
+ assumes xc: "fst (exec_step P h stk loc C M pc ics frs sh) = None"
+ assumes cs: "ics = Called [] \<or> (ics = No_ics \<and> (\<exists>sfs. sh (fst(field P D F)) = Some(sfs, Done)))"
+
+ shows "P,\<Phi> \<turnstile> \<sigma>'\<surd>"
+(*<*)
+proof -
+ from mC cf obtain ST LT where
+ "h\<surd>": "P \<turnstile> h \<surd>" and
+ "sh\<surd>": "P,h \<turnstile>\<^sub>s sh \<surd>" and
+ \<Phi>: "\<Phi> C M ! pc = Some (ST,LT)" and
+ stk: "P,h \<turnstile> stk [:\<le>] ST" and loc: "P,h \<turnstile> loc [:\<le>\<^sub>\<top>] LT" and
+ pc: "pc < size ins" and
+ fs: "conf_fs P h sh \<Phi> C M (size Ts) T frs" and
+ confc: "conf_clinit P sh ((stk,loc,C,M,pc,ics)#frs)" and
+ vics: "P,h,sh \<turnstile>\<^sub>i (C,M,pc,ics)"
+ by (fastforce dest: sees_method_fun)
+
+ from i \<Phi> wt cs obtain vT vT' ST'' ST' LT' where
+ ST: "ST = vT # ST''" and
+ F: "P \<turnstile> C' sees F,Static:vT' in D" and
+ vT: "P \<turnstile> vT \<le> vT'" and
+ pc': "pc+1 < size ins" and
+ \<Phi>': "\<Phi> C M ! (pc+1) = Some (ST', LT')" and
+ ST': "P \<turnstile> ST'' [\<le>] ST'" and LT': "P \<turnstile> LT [\<le>\<^sub>\<top>] LT'"
+ by fastforce
+
+ from stk ST obtain v stk' where
+ stk': "stk = v#stk'" and
+ v: "P,h \<turnstile> v :\<le> vT" and
+ ST'': "P,h \<turnstile> stk' [:\<le>] ST''"
+ by auto
+
+ from v vT have vT': "P,h \<turnstile> v :\<le> vT'" ..
+
+ with mC i vics obtain sobj where
+ cc': "ics = Called [] \<Longrightarrow> Called_context P D (ins!pc) \<and> sh D = Some sobj"
+ by(fastforce dest: has_visible_field)
+
+ from field_def2[OF sees_field_idemp[OF F]] have D[simp]: "fst(field P D F) = D" by simp
+ from cs cc' obtain sfs i where shD: "sh D = Some(sfs,i)" by(cases sobj, auto)
+
+ let ?sh' = "sh(D\<mapsto>(sfs(F\<mapsto>v),i))" and ?f' = "(stk',loc,C,M,pc+1,No_ics)"
+
+ have m_D: "P \<turnstile> D has F,Static:vT' in D" by (rule has_field_idemp[OF has_visible_field[OF F]])
+ from "sh\<surd>" shD have sfs: "P,h,D \<turnstile>\<^sub>s sfs \<surd>" by (rule shconfD)
+
+ have "sh'\<surd>": "P,h \<turnstile>\<^sub>s ?sh' \<surd>" by (rule shconf_upd_obj[OF "sh\<surd>" soconf_fupd[OF m_D vT' sfs]])
+
+ from i mC s' v xc F cs cc' shD stk'
+ have "\<sigma>' = (None, h, (stk',loc,C,M,pc+1,No_ics)#frs, ?sh')"
+ by(fastforce simp: split_beta split: if_split_asm init_call_status.splits)
+ moreover
+ from ST'' ST' have "P,h \<turnstile> stk' [:\<le>] ST'" ..
+ moreover
+ from loc LT' have "P,h \<turnstile> loc [:\<le>\<^sub>\<top>] LT'" ..
+ moreover
+ have "conf_fs P h ?sh' \<Phi> C M (size Ts) T frs" by (rule conf_fs_shupd'[OF fs shD])
+ moreover
+ have "conf_clinit P ?sh' ((stk',loc,C,M,pc+1,No_ics)#frs)"
+ by(rule conf_clinit_diff'[OF conf_clinit_shupd'[OF confc shD]])
+ moreover
+ note "h\<surd>" "sh'\<surd>" mC \<Phi>' pc' v vT'
+ ultimately
+ show "P,\<Phi> \<turnstile> \<sigma>' \<surd>" by fastforce
+qed
+(*>*)
+
+lemma Putstatic_Init_correct:
+ fixes \<sigma>' :: jvm_state
+ assumes wf: "wf_prog wt P"
+ assumes mC: "P \<turnstile> C sees M,b:Ts\<rightarrow>T=(mxs,mxl\<^sub>0,ins,xt) in C"
+ assumes i: "ins!pc = Putstatic C' F D"
+ assumes wt: "P,T,mxs,size ins,xt \<turnstile> ins!pc,pc :: \<Phi> C M"
+ assumes s': "Some \<sigma>' = exec (P, None, h, (stk,loc,C,M,pc,No_ics)#frs, sh)"
+ assumes cf: "P,\<Phi> \<turnstile> (None, h, (stk,loc,C,M,pc,No_ics)#frs, sh)\<surd>"
+ assumes xc: "fst (exec_step P h stk loc C M pc No_ics frs sh) = None"
+ assumes nDone: "\<forall>sfs. sh (fst(field P D F)) \<noteq> Some(sfs, Done)"
+
+ shows "P,\<Phi> \<turnstile> \<sigma>'\<surd>"
+(*<*)
+proof -
+ from mC cf obtain ST LT where
+ "h\<surd>": "P \<turnstile> h \<surd>" and
+ "sh\<surd>": "P,h \<turnstile>\<^sub>s sh \<surd>" and
+ \<Phi>: "\<Phi> C M ! pc = Some (ST,LT)" and
+ stk: "P,h \<turnstile> stk [:\<le>] ST" and loc: "P,h \<turnstile> loc [:\<le>\<^sub>\<top>] LT" and
+ pc: "pc < size ins" and
+ fs: "conf_fs P h sh \<Phi> C M (size Ts) T frs" and
+ confc: "conf_clinit P sh ((stk,loc,C,M,pc,No_ics)#frs)"
+ by (fastforce dest: sees_method_fun)
+
+ from i \<Phi> wt nDone obtain vT where
+ F: "P \<turnstile> C' sees F,Static:vT in D"
+ by fastforce
+ then have has_field: "P \<turnstile> C' has F,Static:vT in D" by(rule has_visible_field)
+
+ from field_def2[OF sees_field_idemp[OF F]] has_field_is_class'[OF has_field] obtain
+ D[simp]: "fst(field P D F) = D" and
+ cls: "is_class P D" by simp
+
+ from i mC s' xc F nDone
+ have "\<sigma>' = (None, h, (stk,loc,C,M,pc,Calling D [])#frs, sh)"
+ by(auto simp: split_beta split: if_split_asm init_state.splits)
+ moreover
+ from confc have "conf_clinit P sh ((stk,loc,C,M,pc,Calling D [])#frs)"
+ by(auto simp: conf_clinit_def distinct_clinit_def split: if_split_asm)
+ moreover
+ note loc stk "h\<surd>" "sh\<surd>" mC \<Phi> pc fs i has_field cls
+ ultimately
+ show "P,\<Phi> \<turnstile> \<sigma>' \<surd>" by fastforce
+qed
+(*>*)
+
+(* FIXME: move *)
+lemma oconf_blank2 [intro, simp]:
+ "\<lbrakk>is_class P C; wf_prog wt P\<rbrakk> \<Longrightarrow> P,h \<turnstile> blank P C \<surd>"
+(*<*)
+ by (fastforce simp: oconf_blank dest: wf_Fields_Ex)
+(*>*)
+
+lemma obj_ty_blank [iff]: "obj_ty (blank P C) = Class C"
+ by (simp add: blank_def)
+
+lemma New_nInit_correct:
+ fixes \<sigma>' :: jvm_state
+ assumes wf: "wf_prog wt P"
+ assumes meth: "P \<turnstile> C sees M,b:Ts\<rightarrow>T=(mxs,mxl\<^sub>0,ins,xt) in C"
+ assumes ins: "ins!pc = New X"
+ assumes wt: "P,T,mxs,size ins,xt \<turnstile> ins!pc,pc :: \<Phi> C M"
+ assumes exec: "Some \<sigma>' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)"
+ assumes conf: "P,\<Phi> \<turnstile> (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\<surd>"
+ assumes no_x: "fst (exec_step P h stk loc C M pc ics frs sh) = None"
+ assumes cs: "ics = Called [] \<or> (ics = No_ics \<and> (\<exists>sfs. sh X = Some(sfs, Done)))"
+ shows "P,\<Phi> \<turnstile> \<sigma>'\<surd>"
+(*<*)
+proof -
+ from ins conf meth
+ obtain ST LT where
+ heap_ok: "P\<turnstile> h\<surd>" and
+ sheap_ok: "P,h \<turnstile>\<^sub>s sh \<surd>" and
+ \<Phi>_pc: "\<Phi> C M!pc = Some (ST,LT)" and
+ frame: "conf_f P h sh (ST,LT) ins (stk,loc,C,M,pc,ics)" and
+ frames: "conf_fs P h sh \<Phi> C M (size Ts) T frs" and
+ confc: "conf_clinit P sh ((stk,loc,C,M,pc,ics) # frs)"
+ by (auto dest: sees_method_fun)
+
+ from \<Phi>_pc ins wt
+ obtain ST' LT' where
+ is_class_X: "is_class P X" and
+ mxs: "size ST < mxs" and
+ suc_pc: "pc+1 < size ins" and
+ \<Phi>_suc: "\<Phi> C M!(pc+1) = Some (ST', LT')" and
+ less: "P \<turnstile> (Class X # ST, LT) \<le>\<^sub>i (ST', LT')"
+ by auto
+
+ from ins no_x cs meth obtain oref where new_Addr: "new_Addr h = Some oref" by auto
+ hence h: "h oref = None" by (rule new_Addr_SomeD)
+
+ with exec ins meth new_Addr cs have \<sigma>':
+ "\<sigma>' = (None, h(oref \<mapsto> blank P X), (Addr oref#stk,loc,C,M,pc+1,No_ics)#frs, sh)"
+ (is "\<sigma>' = (None, ?h', ?f # frs, sh)")
+ by auto
+ moreover
+ from wf h heap_ok is_class_X have h': "P \<turnstile> ?h' \<surd>"
+ by (auto intro: hconf_new)
+ moreover
+ from h frame less suc_pc wf
+ have "conf_f P ?h' sh (ST', LT') ins ?f"
+ by (clarsimp simp: fun_upd_apply conf_def blank_def split_beta)
+ (auto intro: confs_hext confTs_hext)
+ moreover
+ from h have "h \<unlhd> ?h'" by simp
+ with frames have "conf_fs P ?h' sh \<Phi> C M (size Ts) T frs" by (rule conf_fs_hext)
+ moreover
+ have "P,?h' \<turnstile>\<^sub>s sh \<surd>" by(rule shconf_hnew[OF sheap_ok h])
+ moreover
+ have "conf_clinit P sh (?f # frs)" by(rule conf_clinit_diff'[OF confc])
+ ultimately
+ show ?thesis using meth \<Phi>_suc by fastforce
+qed
+(*>*)
+
+lemma New_Init_correct:
+ fixes \<sigma>' :: jvm_state
+ assumes wf: "wf_prog wt P"
+ assumes meth: "P \<turnstile> C sees M,b:Ts\<rightarrow>T=(mxs,mxl\<^sub>0,ins,xt) in C"
+ assumes ins: "ins!pc = New X"
+ assumes wt: "P,T,mxs,size ins,xt \<turnstile> ins!pc,pc :: \<Phi> C M"
+ assumes exec: "Some \<sigma>' = exec (P, None, h, (stk,loc,C,M,pc,No_ics)#frs, sh)"
+ assumes conf: "P,\<Phi> \<turnstile> (None, h, (stk,loc,C,M,pc,No_ics)#frs, sh)\<surd>"
+ assumes no_x: "fst (exec_step P h stk loc C M pc No_ics frs sh) = None"
+ assumes nDone: "\<forall>sfs. sh X \<noteq> Some(sfs, Done)"
+ shows "P,\<Phi> \<turnstile> \<sigma>'\<surd>"
+(*<*)
+proof -
+ from ins conf meth
+ obtain ST LT where
+ heap_ok: "P\<turnstile> h\<surd>" and
+ sheap_ok: "P,h \<turnstile>\<^sub>s sh \<surd>" and
+ \<Phi>_pc: "\<Phi> C M!pc = Some (ST,LT)" and
+ frame: "conf_f P h sh (ST,LT) ins (stk,loc,C,M,pc,No_ics)" and
+ frames: "conf_fs P h sh \<Phi> C M (size Ts) T frs" and
+ confc: "conf_clinit P sh ((stk,loc,C,M,pc,No_ics) # frs)"
+ by (auto dest: sees_method_fun)
+
+ from \<Phi>_pc ins wt
+ obtain ST' LT' where
+ is_class_X: "is_class P X" and
+ mxs: "size ST < mxs" and
+ suc_pc: "pc+1 < size ins" and
+ \<Phi>_suc: "\<Phi> C M!(pc+1) = Some (ST', LT')" and
+ less: "P \<turnstile> (Class X # ST, LT) \<le>\<^sub>i (ST', LT')"
+ by auto
+
+ with exec ins meth nDone have \<sigma>':
+ "\<sigma>' = (None, h, (stk,loc,C,M,pc,Calling X [])#frs, sh)"
+ (is "\<sigma>' = (None, h, ?f # frs, sh)")
+ by(auto split: init_state.splits)
+ moreover
+ from meth frame is_class_X ins
+ have "conf_f P h sh (ST, LT) ins ?f" by auto
+ moreover note heap_ok sheap_ok frames
+ moreover
+ from confc have "conf_clinit P sh (?f # frs)"
+ by(auto simp: conf_clinit_def distinct_clinit_def split: if_split_asm)
+ ultimately
+ show ?thesis using meth \<Phi>_pc by fastforce
+qed
+(*>*)
+
+
+lemma Goto_correct:
+assumes "wf_prog wt P" and
+ mC: "P \<turnstile> C sees M,b:Ts\<rightarrow>T=(mxs,mxl\<^sub>0,ins,xt) in C" and
+ i: "ins!pc = Goto branch" and
+ "P,T,mxs,size ins,xt \<turnstile> ins!pc,pc :: \<Phi> C M" and
+ "Some \<sigma>' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)" and
+ cf: "P,\<Phi> \<turnstile> (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\<surd>"
+shows "P,\<Phi> |- \<sigma>' [ok]"
+(*<*)
+proof -
+ have "ics = No_ics" using mC i cf by(cases ics) auto
+ then show ?thesis using assms(1,3-6) sees_method_fun[OF mC]
+ by(fastforce elim!: conf_clinit_diff)
+qed
+(*>*)
+
+
+lemma IfFalse_correct:
+assumes "wf_prog wt P" and
+ mC: "P \<turnstile> C sees M,b:Ts\<rightarrow>T=(mxs,mxl\<^sub>0,ins,xt) in C" and
+ i: "ins!pc = IfFalse branch" and
+ "P,T,mxs,size ins,xt \<turnstile> ins!pc,pc :: \<Phi> C M" and
+ "Some \<sigma>' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)" and
+ cf: "P,\<Phi> \<turnstile> (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\<surd>"
+shows "P,\<Phi> |- \<sigma>' [ok]"
+(*<*)
+proof -
+ have "ics = No_ics" using mC i cf by(cases ics) auto
+ then show ?thesis using assms(1,3-6) sees_method_fun[OF mC]
+ by(fastforce elim!: conf_clinit_diff)
+qed
+(*>*)
+
+lemma CmpEq_correct:
+assumes "wf_prog wt P" and
+ mC: "P \<turnstile> C sees M,b:Ts\<rightarrow>T=(mxs,mxl\<^sub>0,ins,xt) in C" and
+ i: "ins!pc = CmpEq" and
+ "P,T,mxs,size ins,xt \<turnstile> ins!pc,pc :: \<Phi> C M" and
+ "Some \<sigma>' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)" and
+ cf: "P,\<Phi> \<turnstile> (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\<surd>"
+shows "P,\<Phi> |- \<sigma>' [ok]"
+(*<*)
+proof -
+ have "ics = No_ics" using mC i cf by(cases ics) auto
+ then show ?thesis using assms(1,3-6) sees_method_fun[OF mC]
+ by(fastforce elim!: conf_clinit_diff)
+qed
+(*>*)
+
+lemma Pop_correct:
+assumes "wf_prog wt P" and
+ mC: "P \<turnstile> C sees M,b:Ts\<rightarrow>T=(mxs,mxl\<^sub>0,ins,xt) in C" and
+ i: "ins!pc = Pop" and
+ "P,T,mxs,size ins,xt \<turnstile> ins!pc,pc :: \<Phi> C M" and
+ "Some \<sigma>' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)" and
+ cf: "P,\<Phi> \<turnstile> (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\<surd>"
+shows "P,\<Phi> |- \<sigma>' [ok]"
+(*<*)
+proof -
+ have "ics = No_ics" using mC i cf by(cases ics) auto
+ then show ?thesis using assms(1,3-6) sees_method_fun[OF mC]
+ by(fastforce elim!: conf_clinit_diff)
+qed
+(*>*)
+
+
+lemma IAdd_correct:
+assumes "wf_prog wt P" and
+ mC: "P \<turnstile> C sees M,b:Ts\<rightarrow>T=(mxs,mxl\<^sub>0,ins,xt) in C" and
+ i: "ins!pc = IAdd" and
+ "P,T,mxs,size ins,xt \<turnstile> ins!pc,pc :: \<Phi> C M" and
+ "Some \<sigma>' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)" and
+ cf: "P,\<Phi> \<turnstile> (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\<surd>"
+shows "P,\<Phi> |- \<sigma>' [ok]"
+(*<*)
+proof -
+ have "ics = No_ics" using mC i cf by(cases ics) auto
+ then show ?thesis using assms(1,3-6) sees_method_fun[OF mC]
+ by(fastforce elim!: conf_clinit_diff)
+qed
+(*>*)
+
+
+lemma Throw_correct:
+assumes "wf_prog wt P" and
+ mC: "P \<turnstile> C sees M,b:Ts\<rightarrow>T=(mxs,mxl\<^sub>0,ins,xt) in C" and
+ i: "ins!pc = Throw" and
+ "Some \<sigma>' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)" and
+ cf: "P,\<Phi> \<turnstile> (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\<surd>" and
+ "fst (exec_step P h stk loc C M pc ics frs sh) = None"
+shows "P,\<Phi> |- \<sigma>' [ok]"
+(*<*)
+proof -
+ have "ics = No_ics" using mC i cf by(cases ics) auto
+ then show ?thesis using assms by simp
+qed
+(*>*)
+
+text \<open>
+ The next theorem collects the results of the sections above,
+ i.e.~exception handling, initialization procedure steps, and
+ the execution step for each instruction. It states type safety
+ for single step execution: in welltyped programs, a conforming
+ state is transformed into another conforming state when one
+ step of execution is performed.
+\<close>
+lemma step_correct:
+fixes \<sigma>' :: jvm_state
+assumes wtp: "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P"
+ and meth: "P \<turnstile> C sees M,b:Ts\<rightarrow>T=(mxs,mxl\<^sub>0,ins,xt) in C"
+ and exec: "Some \<sigma>' = exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh)"
+ and conf: "P,\<Phi> \<turnstile> (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\<surd>"
+shows "P,\<Phi> \<turnstile> \<sigma>'\<surd>"
+(*<*)
+proof -
+ from assms have pc: "pc < length ins" by(auto dest: sees_method_fun)
+ with wt_jvm_prog_impl_wt_instr[OF wtp meth] have wt: "P,T,mxs,size ins,xt \<turnstile> ins!pc,pc :: \<Phi> C M"
+ by simp
+
+ from conf obtain ST LT where \<Phi>: "\<Phi> C M ! pc = Some(ST,LT)" by clarsimp
+
+ show ?thesis
+ proof(cases "fst (exec_step P h stk loc C M pc ics frs sh)")
+ case Some show ?thesis by(rule xcpt_correct[OF wtp meth wt Some exec conf])
+ next
+ case None
+ from wt_jvm_progD[OF wtp] obtain wf_md where wf: "wf_prog wf_md P" by clarify
+
+ { assume [simp]: "ics = No_ics"
+
+ from exec conf None obtain
+ exec': "Some \<sigma>' = exec (P, None, h, (stk,loc,C,M,pc,No_ics)#frs, sh)"
+ and conf': "P,\<Phi> \<turnstile> (None, h, (stk,loc,C,M,pc,No_ics)#frs, sh)\<surd>"
+ and None': "fst (exec_step P h stk loc C M pc No_ics frs sh) = None" by simp
+
+ have ?thesis
+ proof(cases "ins!pc")
+ case Load from Load_correct[OF wf meth this wt exec conf] show ?thesis by simp
+ next
+ case Store from Store_correct[OF wf meth this wt exec conf] show ?thesis by simp
+ next
+ case Push from Push_correct[OF wf meth this wt exec conf] show ?thesis by simp
+ next
+ case (New C) show ?thesis
+ proof(cases "\<exists>sfs. sh C = Some(sfs, Done)")
+ case True
+ with New_nInit_correct[OF wf meth New wt exec conf None] show ?thesis by simp
+ next
+ case False
+ with New_Init_correct[OF wf meth New wt exec' conf' None'] show ?thesis by simp
+ qed
+ next
+ case Getfield from Getfield_correct[OF wf meth this wt exec conf None]
+ show ?thesis by simp
+ next
+ case (Getstatic C F D) show ?thesis
+ proof(cases "\<exists>sfs. sh (fst (field P D F)) = Some(sfs, Done)")
+ case True
+ with Getstatic_nInit_correct[OF wf meth Getstatic wt exec conf None] show ?thesis by simp
+ next
+ case False
+ with Getstatic_Init_correct[OF wf meth Getstatic wt exec' conf' None']
+ show ?thesis by simp
+ qed
+ next
+ case Putfield from Putfield_correct[OF wf meth this wt exec conf None]
+ show ?thesis by simp
+ next
+ case (Putstatic C F D) show ?thesis
+ proof(cases "\<exists>sfs. sh (fst (field P D F)) = Some(sfs, Done)")
+ case True
+ with Putstatic_nInit_correct[OF wf meth Putstatic wt exec conf None] show ?thesis by simp
+ next
+ case False
+ with Putstatic_Init_correct[OF wf meth Putstatic wt exec' conf' None']
+ show ?thesis by simp
+ qed
+ next
+ case Checkcast from Checkcast_correct[OF wtp meth this wt exec conf None]
+ show ?thesis by simp
+ next
+ case Invoke with Invoke_correct[OF wtp meth this wt exec conf None] show ?thesis by simp
+ next
+ case (Invokestatic C M n)
+ from wf_jvm_prog_nclinit[OF wtp meth wt pc \<Phi> this] have ncl: "M \<noteq> clinit" by simp
+ show ?thesis
+ proof(cases "\<exists>sfs. sh (fst (method P C M)) = Some(sfs, Done)")
+ case True
+ with Invokestatic_nInit_correct[OF wtp meth Invokestatic ncl wt exec conf None]
+ show ?thesis by simp
+ next
+ case False
+ with Invokestatic_Init_correct[OF wtp meth Invokestatic ncl wt exec' conf' None']
+ show ?thesis by simp
+ qed
+ next
+ case Return with Return_correct[OF wtp meth this wt exec conf] show ?thesis by simp
+ next
+ case Pop with Pop_correct[OF wf meth this wt exec conf] show ?thesis by simp
+ next
+ case IAdd with IAdd_correct[OF wf meth this wt exec conf] show ?thesis by simp
+ next
+ case Goto with Goto_correct[OF wf meth this wt exec conf] show ?thesis by simp
+ next
+ case CmpEq with CmpEq_correct[OF wf meth this wt exec conf] show ?thesis by simp
+ next
+ case IfFalse with IfFalse_correct[OF wf meth this wt exec conf] show ?thesis by simp
+ next
+ case Throw with Throw_correct[OF wf meth this exec conf None] show ?thesis by simp
+ qed
+ }
+ moreover
+ { fix Cs assume [simp]: "ics = Called Cs"
+ have ?thesis
+ proof(cases Cs)
+ case [simp]: Nil
+ from conf meth obtain C1 where "Called_context P C1 (ins ! pc)"
+ by(clarsimp simp: conf_f_def2 intro!: Called_context_Called_set)
+ then have "ins!pc \<in> Called_set" by(rule Called_context_Called_set)
+ then show ?thesis
+ proof(cases "ins!pc")
+ case (New C)
+ from New_nInit_correct[OF wf meth this wt exec conf None] show ?thesis by simp
+ next
+ case (Getstatic C F D)
+ from Getstatic_nInit_correct[OF wf meth this wt exec conf None] show ?thesis by simp
+ next
+ case (Putstatic C F D)
+ from Putstatic_nInit_correct[OF wf meth this wt exec conf None] show ?thesis by simp
+ next
+ case (Invokestatic C M n)
+ from wf_jvm_prog_nclinit[OF wtp meth wt pc \<Phi> this] have ncl: "M \<noteq> clinit" by simp
+ with Invokestatic_nInit_correct[OF wtp meth Invokestatic ncl wt exec conf None]
+ show ?thesis by simp
+ qed(simp_all)
+ next
+ case (Cons C' Cs') with Called_correct[OF wtp meth exec conf None] show ?thesis by simp
+ qed
+ }
+ moreover
+ { fix C' Cs assume [simp]: "ics = Calling C' Cs"
+ with Calling_correct[OF wtp meth exec conf None] have ?thesis by simp
+ }
+ moreover
+ { fix Cs a assume [simp]: "ics = Throwing Cs a"
+ have ?thesis
+ proof(cases Cs)
+ case Nil with exec None show ?thesis by simp
+ next
+ case (Cons C' Cs')
+ with Throwing_correct[OF wtp meth exec conf None] show ?thesis by simp
+ qed
+ }
+ ultimately show ?thesis by(cases ics) auto
+ qed
+qed
+(*>*)
+
+subsection \<open> Main \<close>
+
+lemma correct_state_impl_Some_method:
+ "P,\<Phi> \<turnstile> (None, h, (stk,loc,C,M,pc,ics)#frs, sh)\<surd>
+ \<Longrightarrow> \<exists>b m Ts T. P \<turnstile> C sees M,b:Ts\<rightarrow>T = m in C"
+ by fastforce
+
+lemma BV_correct_1 [rule_format]:
+"\<And>\<sigma>. \<lbrakk> wf_jvm_prog\<^bsub>\<Phi>\<^esub> P; P,\<Phi> \<turnstile> \<sigma>\<surd>\<rbrakk> \<Longrightarrow> P \<turnstile> \<sigma> -jvm\<rightarrow>\<^sub>1 \<sigma>' \<longrightarrow> P,\<Phi> \<turnstile> \<sigma>'\<surd>"
+(*<*)
+proof -
+ fix \<sigma> assume wf: "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P" and cf: "P,\<Phi> \<turnstile> \<sigma>\<surd>"
+ obtain xp h frs sh where \<sigma>[simp]: "\<sigma> = (xp, h, frs, sh)" by(cases \<sigma>) simp
+ have "exec (P, xp, h, frs, sh) = \<lfloor>\<sigma>'\<rfloor> \<longrightarrow> P,\<Phi> |- \<sigma>' [ok]"
+ proof(cases xp)
+ case None
+ with wf cf show ?thesis
+ proof(cases frs)
+ case (Cons fr frs')
+ obtain stk loc C M pc ics where [simp]: "fr = (stk,loc,C,M,pc,ics)" by(cases fr) simp
+ then have cf': "P,\<Phi> |- (None, h, (stk,loc,C,M,pc,ics) # frs', sh) [ok]"
+ using Cons None cf by simp
+ then obtain b mxs mxl\<^sub>0 ins xt Ts T
+ where mC: "P \<turnstile> C sees M, b : Ts\<rightarrow>T = (mxs, mxl\<^sub>0, ins, xt) in C"
+ using correct_state_impl_Some_method[OF cf'] by clarify
+ show ?thesis using Cons None step_correct[OF wf mC _ cf'] by simp
+ qed simp
+ next
+ case (Some a)
+ then show ?thesis using wf cf by (case_tac frs) simp_all
+ qed
+ then show "P \<turnstile> \<sigma> -jvm\<rightarrow>\<^sub>1 \<sigma>' \<longrightarrow> P,\<Phi> \<turnstile> \<sigma>'\<surd>" by(simp add: exec_1_iff)
+qed
+(*>*)
+
+
+theorem progress:
+ "\<lbrakk> xp=None; frs\<noteq>[] \<rbrakk> \<Longrightarrow> \<exists>\<sigma>'. P \<turnstile> (xp,h,frs,sh) -jvm\<rightarrow>\<^sub>1 \<sigma>'"
+ by (clarsimp simp: exec_1_iff neq_Nil_conv split_beta
+ simp del: split_paired_Ex)
+
+lemma progress_conform:
+ "\<lbrakk>wf_jvm_prog\<^bsub>\<Phi>\<^esub> P; P,\<Phi> \<turnstile> (xp,h,frs,sh)\<surd>; xp=None; frs\<noteq>[]\<rbrakk>
+ \<Longrightarrow> \<exists>\<sigma>'. P \<turnstile> (xp,h,frs,sh) -jvm\<rightarrow>\<^sub>1 \<sigma>' \<and> P,\<Phi> \<turnstile> \<sigma>'\<surd>"
+(*<*)by (drule (1) progress) (fast intro: BV_correct_1)(*>*)
+
+theorem BV_correct [rule_format]:
+"\<lbrakk> wf_jvm_prog\<^bsub>\<Phi>\<^esub> P; P \<turnstile> \<sigma> -jvm\<rightarrow> \<sigma>' \<rbrakk> \<Longrightarrow> P,\<Phi> \<turnstile> \<sigma>\<surd> \<longrightarrow> P,\<Phi> \<turnstile> \<sigma>'\<surd>"
+(*<*)
+proof -
+ assume wf: "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P" and "P \<turnstile> \<sigma> -jvm\<rightarrow> \<sigma>'"
+ then have "(\<sigma>, \<sigma>') \<in> (exec_1 P)\<^sup>*" by (simp only: exec_all_def1)
+ then show ?thesis proof(induct rule: rtrancl_induct)
+ case (step y z)
+ then show ?case by clarify (erule (1) BV_correct_1[OF wf])
+ qed simp
+qed
+(*>*)
+
+lemma hconf_start:
+ assumes wf: "wf_prog wf_mb P"
+ shows "P \<turnstile> (start_heap P) \<surd>"
+(*<*)
+proof -
+ { fix a obj assume assm: "start_heap P a = \<lfloor>obj\<rfloor>"
+ have "P,start_heap P \<turnstile> obj \<surd>" using wf assm[THEN sym]
+ by (unfold start_heap_def)
+ (auto simp: fun_upd_apply is_class_xcpt split: if_split_asm)
+ }
+ then show ?thesis using preallocated_start[of P]
+ by (unfold hconf_def) simp
+qed
+(*>*)
+
+lemma shconf_start:
+ "\<not> is_class P Start \<Longrightarrow> P,start_heap P \<turnstile>\<^sub>s start_sheap \<surd>"
+(*<*)
+ by (unfold shconf_def)
+ (clarsimp simp: preallocated_start fun_upd_apply soconf_def has_field_is_class)
+(*>*)
+
+lemma BV_correct_initial:
+assumes wf: "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P"
+ and nStart: "\<not>is_class P Start"
+ and mC: "P \<turnstile> C sees M,Static:[]\<rightarrow>Void = m in C"
+ and nclinit: "M \<noteq> clinit"
+ and \<Phi>: "\<Phi>' Start start_m = start_\<phi>\<^sub>m"
+shows "start_prog P C M,\<Phi>' \<turnstile> start_state P \<surd>"
+(*<*)
+proof -
+ let ?P = "class_add P (start_class C M)"
+ let ?h = "start_heap P" and ?sh = "[Start \<mapsto> (Map.empty, Done)]"
+ let ?C = Start and ?M = start_m and ?pc = 0 and ?ics = No_ics
+ let ?f = "([], [], ?C, ?M, ?pc, ?ics)" and ?fs = "[]"
+ let ?frs = "?f#?fs"
+ have "is_class P Object" using wf by(simp add: wf_jvm_prog_phi_def)
+ then obtain Mm where Mm: "P \<turnstile> Object sees_methods Mm"
+ by(fastforce simp: is_class_def dest: sees_methods_Object)
+ obtain mxs mxl\<^sub>0 ins xt where
+ mC': "P \<turnstile> C sees M,Static:[]\<rightarrow>Void = (mxs,mxl\<^sub>0,ins,xt) in C"
+ using mC by (cases m) simp
+ have "?P\<turnstile> ?h\<surd>" using wf nStart class_add_hconf_wf[OF _ hconf_start]
+ by (simp add: wf_jvm_prog_phi_def start_heap_nStart)
+ moreover have "?P,?h\<turnstile>\<^sub>s ?sh\<surd>" by(rule start_prog_start_shconf)
+ moreover have "conf_clinit ?P ?sh ?frs"
+ by(simp add: conf_clinit_def distinct_clinit_def)
+ moreover have "\<exists>b Ts T mxs mxl\<^sub>0 is xt \<tau>.
+ (?P \<turnstile> ?C sees ?M,b:Ts\<rightarrow>T = (mxs,mxl\<^sub>0,is,xt) in ?C)
+ \<and> \<Phi>' ?C ?M ! ?pc = Some \<tau>
+ \<and> conf_f ?P ?h ?sh \<tau> is ?f \<and> conf_fs ?P ?h ?sh \<Phi>' ?C ?M (size Ts) T ?fs"
+ using \<Phi> start_prog_Start_sees_start_method[OF Mm]
+ by (unfold conf_f_def wt_start_def) fastforce
+ ultimately show ?thesis
+ by (simp del: defs1 add: start_state_def correct_state_def)
+qed
+
+declare [[simproc add: list_to_set_comprehension]]
+(*>*)
+
+theorem typesafe:
+ assumes welltyped: "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P"
+ assumes nstart: "\<not> is_class P Start"
+ assumes main_method: "P \<turnstile> C sees M,Static:[]\<rightarrow>Void = m in C"
+ assumes nclinit: "M \<noteq> clinit"
+ assumes \<Phi>: "\<And>C. C \<noteq> Start \<Longrightarrow> \<Phi>' C = \<Phi> C"
+ assumes \<Phi>': "\<Phi>' Start start_m = start_\<phi>\<^sub>m" "\<Phi>' Start clinit = start_\<phi>\<^sub>m"
+ assumes Obj_start_m:
+ "(\<And>b' Ts' T' m' D'. P \<turnstile> Object sees start_m, b' : Ts'\<rightarrow>T' = m' in D'
+ \<Longrightarrow> b' = Static \<and> Ts' = [] \<and> T' = Void)"
+ shows "start_prog P C M \<turnstile> start_state P -jvm\<rightarrow> \<sigma> \<Longrightarrow> start_prog P C M,\<Phi>' \<turnstile> \<sigma> \<surd>"
+(*<*)
+proof -
+ from welltyped nstart main_method nclinit \<Phi>'(1)
+ have "start_prog P C M,\<Phi>' \<turnstile> start_state P \<surd>" by (rule BV_correct_initial)
+ moreover
+ assume "start_prog P C M \<turnstile> start_state P -jvm\<rightarrow> \<sigma>"
+ moreover
+ from start_prog_wf_jvm_prog_phi[OF welltyped nstart main_method nclinit \<Phi> \<Phi>' Obj_start_m]
+ have "wf_jvm_prog\<^bsub>\<Phi>'\<^esub>(start_prog P C M)" by simp
+ ultimately
+ show "start_prog P C M,\<Phi>' \<turnstile> \<sigma> \<surd>" using welltyped by - (rule BV_correct)
+qed
+(*>*)
+
+end
diff --git a/thys/JinjaDCI/BV/ClassAdd.thy b/thys/JinjaDCI/BV/ClassAdd.thy
--- a/thys/JinjaDCI/BV/ClassAdd.thy
+++ b/thys/JinjaDCI/BV/ClassAdd.thy
@@ -1,685 +1,685 @@
-(* Title: JinjaDCI/BV/ClassAdd.thy
- Author: Susannah Mansky
- 2019-20, UIUC
-*)
-
-section \<open> Property preservation under @{text "class_add"} \<close>
-
-theory ClassAdd
-imports BVConform
-begin
-
-
-lemma err_mono: "A \<subseteq> B \<Longrightarrow> err A \<subseteq> err B"
- by(unfold err_def) auto
-
-lemma opt_mono: "A \<subseteq> B \<Longrightarrow> opt A \<subseteq> opt B"
- by(unfold opt_def) auto
-
-lemma list_mono:
-assumes "A \<subseteq> B" shows "list n A \<subseteq> list n B"
-proof(rule)
- fix xs assume "xs \<in> list n A"
- then obtain size: "size xs = n" and inA: "set xs \<subseteq> A" by simp
- with assms have "set xs \<subseteq> B" by simp
- with size show "xs \<in> list n B" by(clarsimp intro!: listI)
-qed
-
-(****************************************************************)
-
-\<comment> \<open> adding a class in the simplest way \<close>
-abbreviation class_add :: "jvm_prog \<Rightarrow> jvm_method cdecl \<Rightarrow> jvm_prog" where
-"class_add P cd \<equiv> cd#P"
-
-
-subsection "Fields"
-
-lemma class_add_has_fields:
-assumes fs: "P \<turnstile> D has_fields FDTs" and nc: "\<not>is_class P C"
-shows "class_add P (C, cdec) \<turnstile> D has_fields FDTs"
-using assms
-proof(induct rule: Fields.induct)
- case (has_fields_Object D fs ms FDTs)
- from has_fields_is_class_Object[OF fs] nc have "C \<noteq> Object" by fast
- with has_fields_Object show ?case
- by(auto simp: class_def fun_upd_apply intro!: TypeRel.has_fields_Object)
-next
- case rec: (has_fields_rec C1 D fs ms FDTs FDTs')
- with has_fields_is_class have [simp]: "D \<noteq> C" by auto
- with rec have "C1 \<noteq> C" by(clarsimp simp: is_class_def)
- with rec show ?case
- by(auto simp: class_def fun_upd_apply intro: TypeRel.has_fields_rec)
-qed
-
-lemma class_add_has_fields_rev:
- "\<lbrakk> class_add P (C, cdec) \<turnstile> D has_fields FDTs; \<not>P \<turnstile> D \<preceq>\<^sup>* C \<rbrakk>
- \<Longrightarrow> P \<turnstile> D has_fields FDTs"
-proof(induct rule: Fields.induct)
- case (has_fields_Object D fs ms FDTs)
- then show ?case by(auto simp: class_def fun_upd_apply intro!: TypeRel.has_fields_Object)
-next
- case rec: (has_fields_rec C1 D fs ms FDTs FDTs')
- then have sub1: "P \<turnstile> C1 \<prec>\<^sup>1 D"
- by(auto simp: class_def fun_upd_apply intro!: subcls1I split: if_split_asm)
- with rec.prems have cls: "\<not> P \<turnstile> D \<preceq>\<^sup>* C" by (meson converse_rtrancl_into_rtrancl)
- with cls rec show ?case
- by(auto simp: class_def fun_upd_apply
- intro: TypeRel.has_fields_rec split: if_split_asm)
-qed
-
-lemma class_add_has_field:
-assumes "P \<turnstile> C\<^sub>0 has F,b:T in D" and "\<not> is_class P C"
-shows "class_add P (C, cdec) \<turnstile> C\<^sub>0 has F,b:T in D"
-using assms by(auto simp: has_field_def dest!: class_add_has_fields[of P C\<^sub>0])
-
-lemma class_add_has_field_rev:
-assumes has: "class_add P (C, cdec) \<turnstile> C\<^sub>0 has F,b:T in D"
- and ncp: "\<And>D'. P \<turnstile> C\<^sub>0 \<preceq>\<^sup>* D' \<Longrightarrow> D' \<noteq> C"
-shows "P \<turnstile> C\<^sub>0 has F,b:T in D"
-using assms by(auto simp: has_field_def dest!: class_add_has_fields_rev)
-
-lemma class_add_sees_field:
-assumes "P \<turnstile> C\<^sub>0 sees F,b:T in D" and "\<not> is_class P C"
-shows "class_add P (C, cdec) \<turnstile> C\<^sub>0 sees F,b:T in D"
-using assms by(auto simp: sees_field_def dest!: class_add_has_fields[of P C\<^sub>0])
-
-lemma class_add_sees_field_rev:
-assumes has: "class_add P (C, cdec) \<turnstile> C\<^sub>0 sees F,b:T in D"
- and ncp: "\<And>D'. P \<turnstile> C\<^sub>0 \<preceq>\<^sup>* D' \<Longrightarrow> D' \<noteq> C"
-shows "P \<turnstile> C\<^sub>0 sees F,b:T in D"
-using assms by(auto simp: sees_field_def dest!: class_add_has_fields_rev)
-
-lemma class_add_field:
-assumes fd: "P \<turnstile> C\<^sub>0 sees F,b:T in D" and "\<not> is_class P C"
-shows "field P C\<^sub>0 F = field (class_add P (C, cdec)) C\<^sub>0 F"
-using class_add_sees_field[OF assms, of cdec] fd by simp
-
-subsection "Methods"
-
-lemma class_add_sees_methods:
-assumes ms: "P \<turnstile> D sees_methods Mm" and nc: "\<not>is_class P C"
-shows "class_add P (C, cdec) \<turnstile> D sees_methods Mm"
-using assms
-proof(induct rule: Methods.induct)
- case (sees_methods_Object D fs ms Mm)
- from sees_methods_is_class_Object[OF ms] nc have "C \<noteq> Object" by fast
- with sees_methods_Object show ?case
- by(auto simp: class_def fun_upd_apply intro!: TypeRel.sees_methods_Object)
-next
- case rec: (sees_methods_rec C1 D fs ms Mm Mm')
- with sees_methods_is_class have [simp]: "D \<noteq> C" by auto
- with rec have "C1 \<noteq> C" by(clarsimp simp: is_class_def)
- with rec show ?case
- by(auto simp: class_def fun_upd_apply intro: TypeRel.sees_methods_rec)
-qed
-
-lemma class_add_sees_methods_rev:
- "\<lbrakk> class_add P (C, cdec) \<turnstile> D sees_methods Mm;
- \<And>D'. P \<turnstile> D \<preceq>\<^sup>* D' \<Longrightarrow> D' \<noteq> C \<rbrakk>
- \<Longrightarrow> P \<turnstile> D sees_methods Mm"
-proof(induct rule: Methods.induct)
- case (sees_methods_Object D fs ms Mm)
- then show ?case
- by(auto simp: class_def fun_upd_apply intro!: TypeRel.sees_methods_Object)
-next
- case rec: (sees_methods_rec C1 D fs ms Mm Mm')
- then have sub1: "P \<turnstile> C1 \<prec>\<^sup>1 D"
- by(auto simp: class_def fun_upd_apply intro!: subcls1I)
- have cls: "\<And>D'. P \<turnstile> D \<preceq>\<^sup>* D' \<Longrightarrow> D' \<noteq> C"
- proof -
- fix D' assume "P \<turnstile> D \<preceq>\<^sup>* D'"
- with sub1 have "P \<turnstile> C1 \<preceq>\<^sup>* D'" by simp
- with rec.prems show "D' \<noteq> C" by simp
- qed
- with cls rec show ?case
- by(auto simp: class_def fun_upd_apply intro: TypeRel.sees_methods_rec)
-qed
-
-lemma class_add_sees_methods_Obj:
-assumes "P \<turnstile> Object sees_methods Mm" and nObj: "C \<noteq> Object"
-shows "class_add P (C, cdec) \<turnstile> Object sees_methods Mm"
-proof -
- from assms obtain C' fs ms where cls: "class P Object = Some(C',fs,ms)"
- by(auto elim!: Methods.cases)
- with nObj have cls': "class (class_add P (C, cdec)) Object = Some(C',fs,ms)"
- by(simp add: class_def fun_upd_apply)
- from assms cls have "Mm = map_option (\<lambda>m. (m, Object)) \<circ> map_of ms" by(auto elim!: Methods.cases)
- with assms cls' show ?thesis
- by(auto simp: is_class_def fun_upd_apply intro!: sees_methods_Object)
-qed
-
-lemma class_add_sees_methods_rev_Obj:
-assumes "class_add P (C, cdec) \<turnstile> Object sees_methods Mm" and nObj: "C \<noteq> Object"
-shows "P \<turnstile> Object sees_methods Mm"
-proof -
- from assms obtain C' fs ms where cls: "class (class_add P (C, cdec)) Object = Some(C',fs,ms)"
- by(auto elim!: Methods.cases)
- with nObj have cls': "class P Object = Some(C',fs,ms)"
- by(simp add: class_def fun_upd_apply)
- from assms cls have "Mm = map_option (\<lambda>m. (m, Object)) \<circ> map_of ms" by(auto elim!: Methods.cases)
- with assms cls' show ?thesis
- by(auto simp: is_class_def fun_upd_apply intro!: sees_methods_Object)
-qed
-
-lemma class_add_sees_method:
-assumes "P \<turnstile> C\<^sub>0 sees M\<^sub>0, b : Ts\<rightarrow>T = m in D" and "\<not> is_class P C"
-shows "class_add P (C, cdec) \<turnstile> C\<^sub>0 sees M\<^sub>0, b : Ts\<rightarrow>T = m in D"
-using assms by(auto simp: Method_def dest!: class_add_sees_methods[of P C\<^sub>0])
-
-lemma class_add_method:
-assumes md: "P \<turnstile> C\<^sub>0 sees M\<^sub>0, b : Ts\<rightarrow>T = m in D" and "\<not> is_class P C"
-shows "method P C\<^sub>0 M\<^sub>0 = method (class_add P (C, cdec)) C\<^sub>0 M\<^sub>0"
-using class_add_sees_method[OF assms, of cdec] md by simp
-
-lemma class_add_sees_method_rev:
- "\<lbrakk> class_add P (C, cdec) \<turnstile> C\<^sub>0 sees M\<^sub>0, b : Ts\<rightarrow>T = m in D;
- \<not> P \<turnstile> C\<^sub>0 \<preceq>\<^sup>* C \<rbrakk>
- \<Longrightarrow> P \<turnstile> C\<^sub>0 sees M\<^sub>0, b : Ts\<rightarrow>T = m in D"
- by(auto simp: Method_def dest!: class_add_sees_methods_rev)
-
-lemma class_add_sees_method_Obj:
- "\<lbrakk> P \<turnstile> Object sees M\<^sub>0, b : Ts\<rightarrow>T = m in D; C \<noteq> Object \<rbrakk>
- \<Longrightarrow> class_add P (C, cdec) \<turnstile> Object sees M\<^sub>0, b : Ts\<rightarrow>T = m in D"
- by(auto simp: Method_def dest!: class_add_sees_methods_Obj[where P=P])
-
-lemma class_add_sees_method_rev_Obj:
- "\<lbrakk> class_add P (C, cdec) \<turnstile> Object sees M\<^sub>0, b : Ts\<rightarrow>T = m in D; C \<noteq> Object \<rbrakk>
- \<Longrightarrow> P \<turnstile> Object sees M\<^sub>0, b : Ts\<rightarrow>T = m in D"
- by(auto simp: Method_def dest!: class_add_sees_methods_rev_Obj[where P=P])
-
-subsection "Types and states"
-
-lemma class_add_is_type:
- "is_type P T \<Longrightarrow> is_type (class_add P (C, cdec)) T"
- by(cases cdec, simp add: is_type_def is_class_def class_def fun_upd_apply split: ty.splits)
-
-lemma class_add_types:
- "types P \<subseteq> types (class_add P (C, cdec))"
-using class_add_is_type by(cases cdec, clarsimp)
-
-lemma class_add_states:
- "states P mxs mxl \<subseteq> states (class_add P (C, cdec)) mxs mxl"
-proof -
- let ?A = "types P" and ?B = "types (class_add P (C, cdec))"
- have ab: "?A \<subseteq> ?B" by(rule class_add_types)
- moreover have "\<And>n. list n ?A \<subseteq> list n ?B" using ab by(rule list_mono)
- moreover have "list mxl (err ?A) \<subseteq> list mxl (err ?B)" using err_mono[OF ab] by(rule list_mono)
- ultimately show ?thesis by(auto simp: JVM_states_unfold intro!: err_mono opt_mono)
-qed
-
-lemma class_add_check_types:
- "check_types P mxs mxl \<tau>s \<Longrightarrow> check_types (class_add P (C, cdec)) mxs mxl \<tau>s"
-using class_add_states by(fastforce simp: check_types_def)
-
-subsection "Subclasses and subtypes"
-
-lemma class_add_subcls:
- "\<lbrakk> P \<turnstile> D \<preceq>\<^sup>* D'; \<not> is_class P C \<rbrakk>
- \<Longrightarrow> class_add P (C, cdec) \<turnstile> D \<preceq>\<^sup>* D'"
-proof(induct rule: rtrancl.induct)
- case (rtrancl_into_rtrancl a b c)
- then have "b \<noteq> C" by(clarsimp simp: is_class_def dest!: subcls1D)
- with rtrancl_into_rtrancl show ?case
- by(fastforce dest!: subcls1D simp: class_def fun_upd_apply
- intro!: rtrancl_trans[of a b] subcls1I)
-qed(simp)
-
-lemma class_add_subcls_rev:
- "\<lbrakk> class_add P (C, cdec) \<turnstile> D \<preceq>\<^sup>* D'; \<not>P \<turnstile> D \<preceq>\<^sup>* C \<rbrakk>
- \<Longrightarrow> P \<turnstile> D \<preceq>\<^sup>* D'"
-proof(induct rule: rtrancl.induct)
- case (rtrancl_into_rtrancl a b c)
- then have "b \<noteq> C" by(clarsimp simp: is_class_def dest!: subcls1D)
- with rtrancl_into_rtrancl show ?case
- by(fastforce dest!: subcls1D simp: class_def fun_upd_apply
- intro!: rtrancl_trans[of a b] subcls1I)
-qed(simp)
-
-lemma class_add_subtype:
- "\<lbrakk> subtype P x y; \<not> is_class P C \<rbrakk>
- \<Longrightarrow> subtype (class_add P (C, cdec)) x y"
-proof(induct rule: widen.induct)
- case (widen_subcls C D)
- then show ?case using class_add_subcls by simp
-qed(simp+)
-
-lemma class_add_widens:
- "\<lbrakk> P \<turnstile> Ts [\<le>] Ts'; \<not> is_class P C \<rbrakk>
- \<Longrightarrow> (class_add P (C, cdec)) \<turnstile> Ts [\<le>] Ts'"
-using class_add_subtype by (metis (no_types) list_all2_mono)
-
-lemma class_add_sup_ty_opt:
- "\<lbrakk> P \<turnstile> l1 \<le>\<^sub>\<top> l2; \<not> is_class P C \<rbrakk>
- \<Longrightarrow> class_add P (C, cdec) \<turnstile> l1 \<le>\<^sub>\<top> l2"
-using class_add_subtype by(auto simp: sup_ty_opt_def Err.le_def lesub_def split: err.splits)
-
-lemma class_add_sup_loc:
-"\<lbrakk> P \<turnstile> LT [\<le>\<^sub>\<top>] LT'; \<not> is_class P C \<rbrakk>
- \<Longrightarrow> class_add P (C, cdec) \<turnstile> LT [\<le>\<^sub>\<top>] LT'"
-using class_add_sup_ty_opt[where P=P and C=C] by (simp add: list.rel_mono_strong)
-
-lemma class_add_sup_state:
- "\<lbrakk> P \<turnstile> \<tau> \<le>\<^sub>i \<tau>'; \<not> is_class P C \<rbrakk>
- \<Longrightarrow> class_add P (C, cdec) \<turnstile> \<tau> \<le>\<^sub>i \<tau>'"
-using class_add_subtype class_add_sup_ty_opt
- by(auto simp: sup_state_def Listn.le_def Product.le_def lesub_def class_add_widens
- class_add_sup_ty_opt list_all2_mono)
-
-lemma class_add_sup_state_opt:
- "\<lbrakk> P \<turnstile> \<tau> \<le>' \<tau>'; \<not> is_class P C \<rbrakk>
- \<Longrightarrow> class_add P (C, cdec) \<turnstile> \<tau> \<le>' \<tau>'"
- by(auto simp: sup_state_opt_def Opt.le_def lesub_def class_add_widens
- class_add_sup_ty_opt list_all2_mono)
-
-subsection "Effect"
-
-lemma class_add_is_relevant_class:
- "\<lbrakk> is_relevant_class i P C\<^sub>0; \<not> is_class P C \<rbrakk>
- \<Longrightarrow> is_relevant_class i (class_add P (C, cdec)) C\<^sub>0"
- by(cases i, auto simp: class_add_subcls)
-
-lemma class_add_is_relevant_class_rev:
-assumes irc: "is_relevant_class i (class_add P (C, cdec)) C\<^sub>0"
- and ncp: "\<And>cd D'. cd \<in> set P \<Longrightarrow> \<not>P \<turnstile> fst cd \<preceq>\<^sup>* C"
- and wfxp: "wf_syscls P"
-shows "is_relevant_class i P C\<^sub>0"
-using assms
-proof(cases i)
- case (Getfield F D) with assms
- show ?thesis by(fastforce simp: wf_syscls_def sys_xcpts_def dest!: class_add_subcls_rev)
-next
- case (Putfield F D) with assms
- show ?thesis by(fastforce simp: wf_syscls_def sys_xcpts_def dest!: class_add_subcls_rev)
-next
- case (Checkcast D) with assms
- show ?thesis by(fastforce simp: wf_syscls_def sys_xcpts_def dest!: class_add_subcls_rev)
-qed(simp_all)
-
-lemma class_add_is_relevant_entry:
- "\<lbrakk> is_relevant_entry P i pc e; \<not> is_class P C \<rbrakk>
- \<Longrightarrow> is_relevant_entry (class_add P (C, cdec)) i pc e"
- by(clarsimp simp: is_relevant_entry_def class_add_is_relevant_class)
-
-lemma class_add_is_relevant_entry_rev:
- "\<lbrakk> is_relevant_entry (class_add P (C, cdec)) i pc e;
- \<And>cd D'. cd \<in> set P \<Longrightarrow> \<not>P \<turnstile> fst cd \<preceq>\<^sup>* C;
- wf_syscls P \<rbrakk>
- \<Longrightarrow> is_relevant_entry P i pc e"
- by(auto simp: is_relevant_entry_def dest!: class_add_is_relevant_class_rev)
-
-lemma class_add_relevant_entries:
- "\<not> is_class P C
- \<Longrightarrow> set (relevant_entries P i pc xt) \<subseteq> set (relevant_entries (class_add P (C, cdec)) i pc xt)"
- by(clarsimp simp: relevant_entries_def class_add_is_relevant_entry)
-
-lemma class_add_relevant_entries_eq:
-assumes wf: "wf_prog wf_md P" and nclass: "\<not> is_class P C"
-shows "relevant_entries P i pc xt = relevant_entries (class_add P (C, cdec)) i pc xt"
-proof -
- have ncp: "\<And>cd D'. cd \<in> set P \<Longrightarrow> \<not>P \<turnstile> fst cd \<preceq>\<^sup>* C"
- by(rule wf_subcls_nCls'[OF assms])
- moreover from wf have wfsys: "wf_syscls P" by(simp add: wf_prog_def)
- moreover
- note class_add_is_relevant_entry[OF _ nclass, of i pc _ cdec]
- class_add_is_relevant_entry_rev[OF _ ncp wfsys, of cdec i pc]
- ultimately show ?thesis by (metis filter_cong relevant_entries_def)
-qed
-
-lemma class_add_norm_eff_pc:
-assumes ne: "\<forall>(pc',\<tau>') \<in> set (norm_eff i P pc \<tau>). pc' < mpc"
-shows "\<forall>(pc',\<tau>') \<in> set (norm_eff i (class_add P (C, cdec)) pc \<tau>). pc' < mpc"
-using assms by(cases i, auto simp: norm_eff_def)
-
-lemma class_add_norm_eff_sup_state_opt:
-assumes ne: "\<forall>(pc',\<tau>') \<in> set (norm_eff i P pc \<tau>). P \<turnstile> \<tau>' \<le>' \<tau>s!pc'"
- and nclass: "\<not> is_class P C" and app: "app\<^sub>i (i, P, pc, mxs, T, \<tau>)"
-shows "\<forall>(pc',\<tau>') \<in> set (norm_eff i (class_add P (C, cdec)) pc \<tau>). (class_add P (C, cdec)) \<turnstile> \<tau>' \<le>' \<tau>s!pc'"
-proof -
- obtain ST LT where "\<tau> = (ST,LT)" by(cases \<tau>)
- with assms show ?thesis proof(cases i)
- qed(fastforce simp: norm_eff_def
- dest!: class_add_field[where cdec=cdec] class_add_method[where cdec=cdec]
- class_add_sup_loc[OF _ nclass] class_add_subtype[OF _ nclass]
- class_add_widens[OF _ nclass] class_add_sup_state_opt[OF _ nclass])+
-qed
-
-lemma class_add_xcpt_eff_eq:
-assumes wf: "wf_prog wf_md P" and nclass: "\<not> is_class P C"
-shows "xcpt_eff i P pc \<tau> xt = xcpt_eff i (class_add P (C, cdec)) pc \<tau> xt"
-using class_add_relevant_entries_eq[OF assms, of i pc xt cdec] by(cases \<tau>, simp add: xcpt_eff_def)
-
-lemma class_add_eff_pc:
-assumes eff: "\<forall>(pc',\<tau>') \<in> set (eff i P pc xt (Some \<tau>)). pc' < mpc"
- and wf: "wf_prog wf_md P" and nclass: "\<not> is_class P C"
-shows "\<forall>(pc',\<tau>') \<in> set (eff i (class_add P (C, cdec)) pc xt (Some \<tau>)). pc' < mpc"
-using eff class_add_norm_eff_pc class_add_xcpt_eff_eq[OF wf nclass]
- by(auto simp: norm_eff_def eff_def)
-
-lemma class_add_eff_sup_state_opt:
-assumes eff: "\<forall>(pc',\<tau>') \<in> set (eff i P pc xt (Some \<tau>)). P \<turnstile> \<tau>' \<le>' \<tau>s!pc'"
- and wf: "wf_prog wf_md P"and nclass: "\<not> is_class P C"
- and app: "app\<^sub>i (i, P, pc, mxs, T, \<tau>)"
-shows "\<forall>(pc',\<tau>') \<in> set (eff i (class_add P (C, cdec)) pc xt (Some \<tau>)).
- (class_add P (C, cdec)) \<turnstile> \<tau>' \<le>' \<tau>s!pc'"
-proof -
- from eff have ne: "\<forall>(pc', \<tau>')\<in>set (norm_eff i P pc \<tau>). P \<turnstile> \<tau>' \<le>' \<tau>s ! pc'"
- by(simp add: norm_eff_def eff_def)
- from eff have "\<forall>(pc', \<tau>')\<in>set (xcpt_eff i P pc \<tau> xt). P \<turnstile> \<tau>' \<le>' \<tau>s ! pc'"
- by(simp add: xcpt_eff_def eff_def)
- with class_add_norm_eff_sup_state_opt[OF ne nclass app]
- class_add_xcpt_eff_eq[OF wf nclass]class_add_sup_state_opt[OF _ nclass]
- show ?thesis by(cases cdec, auto simp: eff_def norm_eff_def xcpt_app_def)
-qed
-
-
-lemma class_add_app\<^sub>i:
-assumes "app\<^sub>i (i, P, pc, mxs, T\<^sub>r, ST, LT)" and "\<not> is_class P C"
-shows "app\<^sub>i (i, class_add P (C, cdec), pc, mxs, T\<^sub>r, ST, LT)"
-using assms
-proof(cases i)
- case New then show ?thesis using assms by(fastforce simp: is_class_def class_def fun_upd_apply)
-next
- case Getfield then show ?thesis using assms
- by(auto simp: class_add_subtype dest!: class_add_sees_field[where P=P])
-next
- case Getstatic then show ?thesis using assms by(auto dest!: class_add_sees_field[where P=P])
-next
- case Putfield then show ?thesis using assms
- by(auto dest!: class_add_subtype[where P=P] class_add_sees_field[where P=P])
-next
- case Putstatic then show ?thesis using assms
- by(auto dest!: class_add_subtype[where P=P] class_add_sees_field[where P=P])
-next
- case Checkcast then show ?thesis using assms
- by(clarsimp simp: is_class_def class_def fun_upd_apply)
-next
- case Invoke then show ?thesis using assms
- by(fastforce dest!: class_add_widens[where P=P] class_add_sees_method[where P=P])
-next
- case Invokestatic then show ?thesis using assms
- by(fastforce dest!: class_add_widens[where P=P] class_add_sees_method[where P=P])
-next
- case Return then show ?thesis using assms by(clarsimp simp: class_add_subtype)
-qed(simp+)
-
-lemma class_add_xcpt_app:
-assumes xa: "xcpt_app i P pc mxs xt \<tau>"
- and wf: "wf_prog wf_md P" and nclass: "\<not> is_class P C"
-shows "xcpt_app i (class_add P (C, cdec)) pc mxs xt \<tau>"
-using xa class_add_relevant_entries_eq[OF wf nclass] nclass
- by(auto simp: xcpt_app_def is_class_def class_def fun_upd_apply) auto
-
-lemma class_add_app:
-assumes app: "app i P mxs T pc mpc xt t"
- and wf: "wf_prog wf_md P" and nclass: "\<not> is_class P C"
-shows "app i (class_add P (C, cdec)) mxs T pc mpc xt t"
-proof(cases t)
- case (Some \<tau>)
- let ?P = "class_add P (C, cdec)"
- from assms Some have eff: "\<forall>(pc', \<tau>')\<in>set (eff i P pc xt \<lfloor>\<tau>\<rfloor>). pc' < mpc" by(simp add: app_def)
- from assms Some have app\<^sub>i: "app\<^sub>i (i,P,pc,mxs,T,\<tau>)" by(simp add: app_def)
- with class_add_app\<^sub>i[OF _ nclass] Some have "app\<^sub>i (i,?P,pc,mxs,T,\<tau>)" by(cases \<tau>,simp)
- moreover
- from app class_add_xcpt_app[OF _ wf nclass] Some
- have "xcpt_app i ?P pc mxs xt \<tau>" by(simp add: app_def del: xcpt_app_def)
- moreover
- from app class_add_eff_pc[OF eff wf nclass] Some
- have "\<forall>(pc',\<tau>') \<in> set (eff i ?P pc xt t). pc' < mpc" by auto
- moreover note app Some
- ultimately show ?thesis by(simp add: app_def)
-qed(simp)
-
-subsection "Well-formedness and well-typedness"
-
-lemma class_add_wf_mdecl:
- "\<lbrakk> wf_mdecl wf_md P C\<^sub>0 md;
- \<And>C\<^sub>0 md. wf_md P C\<^sub>0 md \<Longrightarrow> wf_md (class_add P (C, cdec)) C\<^sub>0 md \<rbrakk>
- \<Longrightarrow> wf_mdecl wf_md (class_add P (C, cdec)) C\<^sub>0 md"
- by(clarsimp simp: wf_mdecl_def class_add_is_type)
-
-lemma class_add_wf_mdecl':
-assumes wfd: "wf_mdecl wf_md P C\<^sub>0 md"
- and ms: "(C\<^sub>0,S,fs,ms) \<in> set P" and md: "md \<in> set ms"
- and wf_md': "\<And>C\<^sub>0 S fs ms m.\<lbrakk>(C\<^sub>0,S,fs,ms) \<in> set P; m \<in> set ms\<rbrakk> \<Longrightarrow> wf_md' (class_add P (C, cdec)) C\<^sub>0 m"
-shows "wf_mdecl wf_md' (class_add P (C, cdec)) C\<^sub>0 md"
-using assms by(clarsimp simp: wf_mdecl_def class_add_is_type)
-
-lemma class_add_wf_cdecl:
-assumes wfcd: "wf_cdecl wf_md P cd" and cdP: "cd \<in> set P"
- and ncp: "\<not> P \<turnstile> fst cd \<preceq>\<^sup>* C" and dist: "distinct_fst P"
- and wfmd: "\<And>C\<^sub>0 md. wf_md P C\<^sub>0 md \<Longrightarrow> wf_md (class_add P (C, cdec)) C\<^sub>0 md"
- and nclass: "\<not> is_class P C"
-shows "wf_cdecl wf_md (class_add P (C, cdec)) cd"
-proof -
- let ?P = "class_add P (C, cdec)"
- obtain C1 D fs ms where [simp]: "cd = (C1,(D,fs,ms))" by(cases cd)
- from wfcd
- have "\<forall>f\<in>set fs. wf_fdecl ?P f" by(auto simp: wf_cdecl_def wf_fdecl_def class_add_is_type)
- moreover
- from wfcd wfmd class_add_wf_mdecl
- have "\<forall>m\<in>set ms. wf_mdecl wf_md ?P C1 m" by(auto simp: wf_cdecl_def)
- moreover
- have "C1 \<noteq> Object \<Longrightarrow> is_class ?P D \<and> \<not> ?P \<turnstile> D \<preceq>\<^sup>* C1
- \<and> (\<forall>(M,b,Ts,T,m)\<in>set ms.
- \<forall>D' b' Ts' T' m'. ?P \<turnstile> D sees M,b':Ts' \<rightarrow> T' = m' in D' \<longrightarrow>
- b = b' \<and> ?P \<turnstile> Ts' [\<le>] Ts \<and> ?P \<turnstile> T \<le> T')"
- proof -
- assume nObj[simp]: "C1 \<noteq> Object"
- with cdP dist have sub1: "P \<turnstile> C1 \<prec>\<^sup>1 D" by(auto simp: class_def intro!: subcls1I map_of_SomeI)
- with ncp have ncp': "\<not> P \<turnstile> D \<preceq>\<^sup>* C" by(auto simp: converse_rtrancl_into_rtrancl)
- with wfcd
- have clsD: "is_class ?P D"
- by(auto simp: wf_cdecl_def is_class_def class_def fun_upd_apply)
- moreover
- from wfcd sub1
- have "\<not> ?P \<turnstile> D \<preceq>\<^sup>* C1" by(auto simp: wf_cdecl_def dest!: class_add_subcls_rev[OF _ ncp'])
- moreover
- have "\<And>M b Ts T m D' b' Ts' T' m'. (M,b,Ts,T,m) \<in> set ms
- \<Longrightarrow> ?P \<turnstile> D sees M,b':Ts' \<rightarrow> T' = m' in D'
- \<Longrightarrow> b = b' \<and> ?P \<turnstile> Ts' [\<le>] Ts \<and> ?P \<turnstile> T \<le> T'"
- proof -
- fix M b Ts T m D' b' Ts' T' m'
- assume ms: "(M,b,Ts,T,m) \<in> set ms" and meth': "?P \<turnstile> D sees M,b':Ts' \<rightarrow> T' = m' in D'"
- with sub1
- have "P \<turnstile> D sees M,b':Ts' \<rightarrow> T' = m' in D'"
- by(fastforce dest!: class_add_sees_method_rev[OF _ ncp'])
- moreover
- with wfcd ms meth'
- have "b = b' \<and> P \<turnstile> Ts' [\<le>] Ts \<and> P \<turnstile> T \<le> T'"
- by(cases m', fastforce simp: wf_cdecl_def elim!: ballE[where x="(M,b,Ts,T,m)"])
- ultimately show "b = b' \<and> ?P \<turnstile> Ts' [\<le>] Ts \<and> ?P \<turnstile> T \<le> T'"
- by(auto dest!: class_add_subtype[OF _ nclass] class_add_widens[OF _ nclass])
- qed
- ultimately show ?thesis by clarsimp
- qed
- moreover note wfcd
- ultimately show ?thesis by(simp add: wf_cdecl_def)
-qed
-
-lemma class_add_wf_cdecl':
-assumes wfcd: "wf_cdecl wf_md P cd" and cdP: "cd \<in> set P"
- and ncp: "\<not>P \<turnstile> fst cd \<preceq>\<^sup>* C" and dist: "distinct_fst P"
- and wfmd: "\<And>C\<^sub>0 S fs ms m.\<lbrakk>(C\<^sub>0,S,fs,ms) \<in> set P; m \<in> set ms\<rbrakk> \<Longrightarrow> wf_md' (class_add P (C, cdec)) C\<^sub>0 m"
- and nclass: "\<not> is_class P C"
-shows "wf_cdecl wf_md' (class_add P (C, cdec)) cd"
-proof -
- let ?P = "class_add P (C, cdec)"
- obtain C1 D fs ms where [simp]: "cd = (C1,(D,fs,ms))" by(cases cd)
- from wfcd
- have "\<forall>f\<in>set fs. wf_fdecl ?P f" by(auto simp: wf_cdecl_def wf_fdecl_def class_add_is_type)
- moreover
- from cdP wfcd wfmd
- have "\<forall>m\<in>set ms. wf_mdecl wf_md' ?P C1 m"
- by(auto simp: wf_cdecl_def wf_mdecl_def class_add_is_type)
- moreover
- have "C1 \<noteq> Object \<Longrightarrow> is_class ?P D \<and> \<not> ?P \<turnstile> D \<preceq>\<^sup>* C1
- \<and> (\<forall>(M,b,Ts,T,m)\<in>set ms.
- \<forall>D' b' Ts' T' m'. ?P \<turnstile> D sees M,b':Ts' \<rightarrow> T' = m' in D' \<longrightarrow>
- b = b' \<and> ?P \<turnstile> Ts' [\<le>] Ts \<and> ?P \<turnstile> T \<le> T')"
- proof -
- assume nObj[simp]: "C1 \<noteq> Object"
- with cdP dist have sub1: "P \<turnstile> C1 \<prec>\<^sup>1 D" by(auto simp: class_def intro!: subcls1I map_of_SomeI)
- with ncp have ncp': "\<not> P \<turnstile> D \<preceq>\<^sup>* C" by(auto simp: converse_rtrancl_into_rtrancl)
- with wfcd
- have clsD: "is_class ?P D"
- by(auto simp: wf_cdecl_def is_class_def class_def fun_upd_apply)
- moreover
- from wfcd sub1
- have "\<not> ?P \<turnstile> D \<preceq>\<^sup>* C1" by(auto simp: wf_cdecl_def dest!: class_add_subcls_rev[OF _ ncp'])
- moreover
- have "\<And>M b Ts T m D' b' Ts' T' m'. (M,b,Ts,T,m) \<in> set ms
- \<Longrightarrow> ?P \<turnstile> D sees M,b':Ts' \<rightarrow> T' = m' in D'
- \<Longrightarrow> b = b' \<and> ?P \<turnstile> Ts' [\<le>] Ts \<and> ?P \<turnstile> T \<le> T'"
- proof -
- fix M b Ts T m D' b' Ts' T' m'
- assume ms: "(M,b,Ts,T,m) \<in> set ms" and meth': "?P \<turnstile> D sees M,b':Ts' \<rightarrow> T' = m' in D'"
- with sub1
- have "P \<turnstile> D sees M,b':Ts' \<rightarrow> T' = m' in D'"
- by(fastforce dest!: class_add_sees_method_rev[OF _ ncp'])
- moreover
- with wfcd ms meth'
- have "b = b' \<and> P \<turnstile> Ts' [\<le>] Ts \<and> P \<turnstile> T \<le> T'"
- by(cases m', fastforce simp: wf_cdecl_def elim!: ballE[where x="(M,b,Ts,T,m)"])
- ultimately show "b = b' \<and> ?P \<turnstile> Ts' [\<le>] Ts \<and> ?P \<turnstile> T \<le> T'"
- by(auto dest!: class_add_subtype[OF _ nclass] class_add_widens[OF _ nclass])
- qed
- ultimately show ?thesis by clarsimp
- qed
- moreover note wfcd
- ultimately show ?thesis by(simp add: wf_cdecl_def)
-qed
-
-lemma class_add_wt_start:
- "\<lbrakk> wt_start P C\<^sub>0 b Ts mxl \<tau>s; \<not> is_class P C \<rbrakk>
- \<Longrightarrow> wt_start (class_add P (C, cdec)) C\<^sub>0 b Ts mxl \<tau>s"
-using class_add_sup_state_opt by(clarsimp simp: wt_start_def split: staticb.splits)
-
-lemma class_add_wt_instr:
-assumes wti: "P,T,mxs,mpc,xt \<turnstile> i,pc :: \<tau>s"
- and wf: "wf_prog wf_md P" and nclass: "\<not> is_class P C"
-shows "class_add P (C, cdec),T,mxs,mpc,xt \<turnstile> i,pc :: \<tau>s"
-proof -
- let ?P = "class_add P (C, cdec)"
- from wti have eff: "\<forall>(pc', \<tau>')\<in>set (eff i P pc xt (\<tau>s ! pc)). P \<turnstile> \<tau>' \<le>' \<tau>s ! pc'"
- by(simp add: wt_instr_def)
- from wti have app\<^sub>i: "\<tau>s!pc \<noteq> None \<Longrightarrow> app\<^sub>i (i,P,pc,mxs,T,the (\<tau>s!pc))"
- by(simp add: wt_instr_def app_def)
- from wti class_add_app[OF _ wf nclass]
- have "app i ?P mxs T pc mpc xt (\<tau>s!pc)" by(simp add: wt_instr_def)
- moreover
- have "\<forall>(pc',\<tau>') \<in> set (eff i ?P pc xt (\<tau>s!pc)). ?P \<turnstile> \<tau>' \<le>' \<tau>s!pc'"
- proof(cases "\<tau>s!pc")
- case Some with eff class_add_eff_sup_state_opt[OF _ wf nclass app\<^sub>i] show ?thesis by auto
- qed(simp add: eff_def)
- moreover note wti
- ultimately show ?thesis by(clarsimp simp: wt_instr_def)
-qed
-
-lemma class_add_wt_method:
-assumes wtm: "wt_method P C\<^sub>0 b Ts T\<^sub>r mxs mxl\<^sub>0 is xt (\<Phi> C\<^sub>0 M\<^sub>0)"
- and wf: "wf_prog wf_md P" and nclass: "\<not> is_class P C"
-shows "wt_method (class_add P (C, cdec)) C\<^sub>0 b Ts T\<^sub>r mxs mxl\<^sub>0 is xt (\<Phi> C\<^sub>0 M\<^sub>0)"
-proof -
- let ?P = "class_add P (C, cdec)"
- let ?\<tau>s = "\<Phi> C\<^sub>0 M\<^sub>0"
- from wtm class_add_check_types
- have "check_types ?P mxs ((case b of Static \<Rightarrow> 0 | NonStatic \<Rightarrow> 1)+size Ts+mxl\<^sub>0) (map OK ?\<tau>s)"
- by(simp add: wt_method_def)
- moreover
- from wtm class_add_wt_start nclass
- have "wt_start ?P C\<^sub>0 b Ts mxl\<^sub>0 ?\<tau>s" by(simp add: wt_method_def)
- moreover
- from wtm class_add_wt_instr[OF _ wf nclass]
- have "\<forall>pc < size is. ?P,T\<^sub>r,mxs,size is,xt \<turnstile> is!pc,pc :: ?\<tau>s" by(clarsimp simp: wt_method_def)
- moreover note wtm
- ultimately
- show ?thesis by(clarsimp simp: wt_method_def)
-qed
-
-lemma class_add_wt_method':
- "\<lbrakk> (\<lambda>P C (M,b,Ts,T\<^sub>r,(mxs,mxl\<^sub>0,is,xt)). wt_method P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt (\<Phi> C M)) P C\<^sub>0 md;
- wf_prog wf_md P; \<not> is_class P C \<rbrakk>
- \<Longrightarrow> (\<lambda>P C (M,b,Ts,T\<^sub>r,(mxs,mxl\<^sub>0,is,xt)). wt_method P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt (\<Phi> C M))
- (class_add P (C, cdec)) C\<^sub>0 md"
- by(clarsimp simp: class_add_wt_method)
-
-subsection \<open> @{text "distinct_fst"} \<close>
-
-lemma class_add_distinct_fst:
-"\<lbrakk> distinct_fst P; \<not> is_class P C \<rbrakk>
- \<Longrightarrow> distinct_fst (class_add P (C, cdec))"
- by(clarsimp simp: distinct_fst_def is_class_def class_def)
-
-subsection "Conformance"
-
-lemma class_add_conf:
- "\<lbrakk> P,h \<turnstile> v :\<le> T; \<not> is_class P C \<rbrakk>
- \<Longrightarrow> class_add P (C, cdec),h \<turnstile> v :\<le> T"
- by(clarsimp simp: conf_def class_add_subtype)
-
-lemma class_add_oconf:
-fixes obj::obj
-assumes oc: "P,h \<turnstile> obj \<surd>" and ns: "\<not> is_class P C"
- and ncp: "\<And>D'. P \<turnstile> fst(obj) \<preceq>\<^sup>* D' \<Longrightarrow> D' \<noteq> C"
-shows "(class_add P (C, cdec)),h \<turnstile> obj \<surd>"
-proof -
- obtain C\<^sub>0 fs where [simp]: "obj=(C\<^sub>0,fs)" by(cases obj)
- from oc have
- oc': "\<And>F D T. P \<turnstile> C\<^sub>0 has F,NonStatic:T in D \<Longrightarrow> (\<exists>v. fs (F, D) = \<lfloor>v\<rfloor> \<and> P,h \<turnstile> v :\<le> T)"
- by(simp add: oconf_def)
- have "\<And>F D T. class_add P (C, cdec) \<turnstile> C\<^sub>0 has F,NonStatic:T in D
- \<Longrightarrow> \<exists>v. fs(F,D) = Some v \<and> class_add P (C, cdec),h \<turnstile> v :\<le> T"
- proof -
- fix F D T assume "class_add P (C, cdec) \<turnstile> C\<^sub>0 has F,NonStatic:T in D"
- with class_add_has_field_rev[OF _ ncp] have meth: "P \<turnstile> C\<^sub>0 has F,NonStatic:T in D" by simp
- then show "\<exists>v. fs(F,D) = Some v \<and> class_add P (C, cdec),h \<turnstile> v :\<le> T"
- using oc'[OF meth] class_add_conf[OF _ ns] by(fastforce simp: oconf_def)
- qed
- then show ?thesis by(simp add: oconf_def)
-qed
-
-lemma class_add_soconf:
-assumes soc: "P,h,C\<^sub>0 \<turnstile>\<^sub>s sfs \<surd>" and ns: "\<not> is_class P C"
- and ncp: "\<And>D'. P \<turnstile> C\<^sub>0 \<preceq>\<^sup>* D' \<Longrightarrow> D' \<noteq> C"
-shows "(class_add P (C, cdec)),h,C\<^sub>0 \<turnstile>\<^sub>s sfs \<surd>"
-proof -
- from soc have
- oc': "\<And>F T. P \<turnstile> C\<^sub>0 has F,Static:T in C\<^sub>0 \<Longrightarrow> (\<exists>v. sfs F = \<lfloor>v\<rfloor> \<and> P,h \<turnstile> v :\<le> T)"
- by(simp add: soconf_def)
- have "\<And>F T. class_add P (C, cdec) \<turnstile> C\<^sub>0 has F,Static:T in C\<^sub>0
- \<Longrightarrow> \<exists>v. sfs F = Some v \<and> class_add P (C, cdec),h \<turnstile> v :\<le> T"
- proof -
- fix F T assume "class_add P (C, cdec) \<turnstile> C\<^sub>0 has F,Static:T in C\<^sub>0"
- with class_add_has_field_rev[OF _ ncp] have meth: "P \<turnstile> C\<^sub>0 has F,Static:T in C\<^sub>0" by simp
- then show "\<exists>v. sfs F = Some v \<and> class_add P (C, cdec),h \<turnstile> v :\<le> T"
- using oc'[OF meth] class_add_conf[OF _ ns] by(fastforce simp: soconf_def)
- qed
- then show ?thesis by(simp add: soconf_def)
-qed
-
-lemma class_add_hconf:
-assumes "P \<turnstile> h \<surd>" and "\<not> is_class P C"
- and "\<And>a obj D'. h a = Some obj \<Longrightarrow> P \<turnstile> fst(obj) \<preceq>\<^sup>* D' \<Longrightarrow> D' \<noteq> C"
-shows "class_add P (C, cdec) \<turnstile> h \<surd>"
-using assms by(auto simp: hconf_def intro!: class_add_oconf)
-
-lemma class_add_hconf_wf:
-assumes wf: "wf_prog wf_md P" and "P \<turnstile> h \<surd>" and "\<not> is_class P C"
- and "\<And>a obj. h a = Some obj \<Longrightarrow> fst(obj) \<noteq> C"
-shows "class_add P (C, cdec) \<turnstile> h \<surd>"
-using wf_subcls_nCls[OF wf] assms by(fastforce simp: hconf_def intro!: class_add_oconf)
-
-lemma class_add_shconf:
-assumes "P,h \<turnstile>\<^sub>s sh \<surd>" and ns: "\<not> is_class P C"
- and "\<And>C sobj D'. sh C = Some sobj \<Longrightarrow> P \<turnstile> C \<preceq>\<^sup>* D' \<Longrightarrow> D' \<noteq> C"
-shows "class_add P (C, cdec),h \<turnstile>\<^sub>s sh \<surd>"
-using assms by(fastforce simp: shconf_def)
-
-lemma class_add_shconf_wf:
-assumes wf: "wf_prog wf_md P" and "P,h \<turnstile>\<^sub>s sh \<surd>" and "\<not> is_class P C"
- and "\<And>C sobj. sh C = Some sobj \<Longrightarrow> C \<noteq> C"
-shows "class_add P (C, cdec),h \<turnstile>\<^sub>s sh \<surd>"
-using wf_subcls_nCls[OF wf] assms by(fastforce simp: shconf_def)
-
-
+(* Title: JinjaDCI/BV/ClassAdd.thy
+ Author: Susannah Mansky
+ 2019-20, UIUC
+*)
+
+section \<open> Property preservation under @{text "class_add"} \<close>
+
+theory ClassAdd
+imports BVConform
+begin
+
+
+lemma err_mono: "A \<subseteq> B \<Longrightarrow> err A \<subseteq> err B"
+ by(unfold err_def) auto
+
+lemma opt_mono: "A \<subseteq> B \<Longrightarrow> opt A \<subseteq> opt B"
+ by(unfold opt_def) auto
+
+lemma list_mono:
+assumes "A \<subseteq> B" shows "list n A \<subseteq> list n B"
+proof(rule)
+ fix xs assume "xs \<in> list n A"
+ then obtain size: "size xs = n" and inA: "set xs \<subseteq> A" by simp
+ with assms have "set xs \<subseteq> B" by simp
+ with size show "xs \<in> list n B" by(clarsimp intro!: listI)
+qed
+
+(****************************************************************)
+
+\<comment> \<open> adding a class in the simplest way \<close>
+abbreviation class_add :: "jvm_prog \<Rightarrow> jvm_method cdecl \<Rightarrow> jvm_prog" where
+"class_add P cd \<equiv> cd#P"
+
+
+subsection "Fields"
+
+lemma class_add_has_fields:
+assumes fs: "P \<turnstile> D has_fields FDTs" and nc: "\<not>is_class P C"
+shows "class_add P (C, cdec) \<turnstile> D has_fields FDTs"
+using assms
+proof(induct rule: Fields.induct)
+ case (has_fields_Object D fs ms FDTs)
+ from has_fields_is_class_Object[OF fs] nc have "C \<noteq> Object" by fast
+ with has_fields_Object show ?case
+ by(auto simp: class_def fun_upd_apply intro!: TypeRel.has_fields_Object)
+next
+ case rec: (has_fields_rec C1 D fs ms FDTs FDTs')
+ with has_fields_is_class have [simp]: "D \<noteq> C" by auto
+ with rec have "C1 \<noteq> C" by(clarsimp simp: is_class_def)
+ with rec show ?case
+ by(auto simp: class_def fun_upd_apply intro: TypeRel.has_fields_rec)
+qed
+
+lemma class_add_has_fields_rev:
+ "\<lbrakk> class_add P (C, cdec) \<turnstile> D has_fields FDTs; \<not>P \<turnstile> D \<preceq>\<^sup>* C \<rbrakk>
+ \<Longrightarrow> P \<turnstile> D has_fields FDTs"
+proof(induct rule: Fields.induct)
+ case (has_fields_Object D fs ms FDTs)
+ then show ?case by(auto simp: class_def fun_upd_apply intro!: TypeRel.has_fields_Object)
+next
+ case rec: (has_fields_rec C1 D fs ms FDTs FDTs')
+ then have sub1: "P \<turnstile> C1 \<prec>\<^sup>1 D"
+ by(auto simp: class_def fun_upd_apply intro!: subcls1I split: if_split_asm)
+ with rec.prems have cls: "\<not> P \<turnstile> D \<preceq>\<^sup>* C" by (meson converse_rtrancl_into_rtrancl)
+ with cls rec show ?case
+ by(auto simp: class_def fun_upd_apply
+ intro: TypeRel.has_fields_rec split: if_split_asm)
+qed
+
+lemma class_add_has_field:
+assumes "P \<turnstile> C\<^sub>0 has F,b:T in D" and "\<not> is_class P C"
+shows "class_add P (C, cdec) \<turnstile> C\<^sub>0 has F,b:T in D"
+using assms by(auto simp: has_field_def dest!: class_add_has_fields[of P C\<^sub>0])
+
+lemma class_add_has_field_rev:
+assumes has: "class_add P (C, cdec) \<turnstile> C\<^sub>0 has F,b:T in D"
+ and ncp: "\<And>D'. P \<turnstile> C\<^sub>0 \<preceq>\<^sup>* D' \<Longrightarrow> D' \<noteq> C"
+shows "P \<turnstile> C\<^sub>0 has F,b:T in D"
+using assms by(auto simp: has_field_def dest!: class_add_has_fields_rev)
+
+lemma class_add_sees_field:
+assumes "P \<turnstile> C\<^sub>0 sees F,b:T in D" and "\<not> is_class P C"
+shows "class_add P (C, cdec) \<turnstile> C\<^sub>0 sees F,b:T in D"
+using assms by(auto simp: sees_field_def dest!: class_add_has_fields[of P C\<^sub>0])
+
+lemma class_add_sees_field_rev:
+assumes has: "class_add P (C, cdec) \<turnstile> C\<^sub>0 sees F,b:T in D"
+ and ncp: "\<And>D'. P \<turnstile> C\<^sub>0 \<preceq>\<^sup>* D' \<Longrightarrow> D' \<noteq> C"
+shows "P \<turnstile> C\<^sub>0 sees F,b:T in D"
+using assms by(auto simp: sees_field_def dest!: class_add_has_fields_rev)
+
+lemma class_add_field:
+assumes fd: "P \<turnstile> C\<^sub>0 sees F,b:T in D" and "\<not> is_class P C"
+shows "field P C\<^sub>0 F = field (class_add P (C, cdec)) C\<^sub>0 F"
+using class_add_sees_field[OF assms, of cdec] fd by simp
+
+subsection "Methods"
+
+lemma class_add_sees_methods:
+assumes ms: "P \<turnstile> D sees_methods Mm" and nc: "\<not>is_class P C"
+shows "class_add P (C, cdec) \<turnstile> D sees_methods Mm"
+using assms
+proof(induct rule: Methods.induct)
+ case (sees_methods_Object D fs ms Mm)
+ from sees_methods_is_class_Object[OF ms] nc have "C \<noteq> Object" by fast
+ with sees_methods_Object show ?case
+ by(auto simp: class_def fun_upd_apply intro!: TypeRel.sees_methods_Object)
+next
+ case rec: (sees_methods_rec C1 D fs ms Mm Mm')
+ with sees_methods_is_class have [simp]: "D \<noteq> C" by auto
+ with rec have "C1 \<noteq> C" by(clarsimp simp: is_class_def)
+ with rec show ?case
+ by(auto simp: class_def fun_upd_apply intro: TypeRel.sees_methods_rec)
+qed
+
+lemma class_add_sees_methods_rev:
+ "\<lbrakk> class_add P (C, cdec) \<turnstile> D sees_methods Mm;
+ \<And>D'. P \<turnstile> D \<preceq>\<^sup>* D' \<Longrightarrow> D' \<noteq> C \<rbrakk>
+ \<Longrightarrow> P \<turnstile> D sees_methods Mm"
+proof(induct rule: Methods.induct)
+ case (sees_methods_Object D fs ms Mm)
+ then show ?case
+ by(auto simp: class_def fun_upd_apply intro!: TypeRel.sees_methods_Object)
+next
+ case rec: (sees_methods_rec C1 D fs ms Mm Mm')
+ then have sub1: "P \<turnstile> C1 \<prec>\<^sup>1 D"
+ by(auto simp: class_def fun_upd_apply intro!: subcls1I)
+ have cls: "\<And>D'. P \<turnstile> D \<preceq>\<^sup>* D' \<Longrightarrow> D' \<noteq> C"
+ proof -
+ fix D' assume "P \<turnstile> D \<preceq>\<^sup>* D'"
+ with sub1 have "P \<turnstile> C1 \<preceq>\<^sup>* D'" by simp
+ with rec.prems show "D' \<noteq> C" by simp
+ qed
+ with cls rec show ?case
+ by(auto simp: class_def fun_upd_apply intro: TypeRel.sees_methods_rec)
+qed
+
+lemma class_add_sees_methods_Obj:
+assumes "P \<turnstile> Object sees_methods Mm" and nObj: "C \<noteq> Object"
+shows "class_add P (C, cdec) \<turnstile> Object sees_methods Mm"
+proof -
+ from assms obtain C' fs ms where cls: "class P Object = Some(C',fs,ms)"
+ by(auto elim!: Methods.cases)
+ with nObj have cls': "class (class_add P (C, cdec)) Object = Some(C',fs,ms)"
+ by(simp add: class_def fun_upd_apply)
+ from assms cls have "Mm = map_option (\<lambda>m. (m, Object)) \<circ> map_of ms" by(auto elim!: Methods.cases)
+ with assms cls' show ?thesis
+ by(auto simp: is_class_def fun_upd_apply intro!: sees_methods_Object)
+qed
+
+lemma class_add_sees_methods_rev_Obj:
+assumes "class_add P (C, cdec) \<turnstile> Object sees_methods Mm" and nObj: "C \<noteq> Object"
+shows "P \<turnstile> Object sees_methods Mm"
+proof -
+ from assms obtain C' fs ms where cls: "class (class_add P (C, cdec)) Object = Some(C',fs,ms)"
+ by(auto elim!: Methods.cases)
+ with nObj have cls': "class P Object = Some(C',fs,ms)"
+ by(simp add: class_def fun_upd_apply)
+ from assms cls have "Mm = map_option (\<lambda>m. (m, Object)) \<circ> map_of ms" by(auto elim!: Methods.cases)
+ with assms cls' show ?thesis
+ by(auto simp: is_class_def fun_upd_apply intro!: sees_methods_Object)
+qed
+
+lemma class_add_sees_method:
+assumes "P \<turnstile> C\<^sub>0 sees M\<^sub>0, b : Ts\<rightarrow>T = m in D" and "\<not> is_class P C"
+shows "class_add P (C, cdec) \<turnstile> C\<^sub>0 sees M\<^sub>0, b : Ts\<rightarrow>T = m in D"
+using assms by(auto simp: Method_def dest!: class_add_sees_methods[of P C\<^sub>0])
+
+lemma class_add_method:
+assumes md: "P \<turnstile> C\<^sub>0 sees M\<^sub>0, b : Ts\<rightarrow>T = m in D" and "\<not> is_class P C"
+shows "method P C\<^sub>0 M\<^sub>0 = method (class_add P (C, cdec)) C\<^sub>0 M\<^sub>0"
+using class_add_sees_method[OF assms, of cdec] md by simp
+
+lemma class_add_sees_method_rev:
+ "\<lbrakk> class_add P (C, cdec) \<turnstile> C\<^sub>0 sees M\<^sub>0, b : Ts\<rightarrow>T = m in D;
+ \<not> P \<turnstile> C\<^sub>0 \<preceq>\<^sup>* C \<rbrakk>
+ \<Longrightarrow> P \<turnstile> C\<^sub>0 sees M\<^sub>0, b : Ts\<rightarrow>T = m in D"
+ by(auto simp: Method_def dest!: class_add_sees_methods_rev)
+
+lemma class_add_sees_method_Obj:
+ "\<lbrakk> P \<turnstile> Object sees M\<^sub>0, b : Ts\<rightarrow>T = m in D; C \<noteq> Object \<rbrakk>
+ \<Longrightarrow> class_add P (C, cdec) \<turnstile> Object sees M\<^sub>0, b : Ts\<rightarrow>T = m in D"
+ by(auto simp: Method_def dest!: class_add_sees_methods_Obj[where P=P])
+
+lemma class_add_sees_method_rev_Obj:
+ "\<lbrakk> class_add P (C, cdec) \<turnstile> Object sees M\<^sub>0, b : Ts\<rightarrow>T = m in D; C \<noteq> Object \<rbrakk>
+ \<Longrightarrow> P \<turnstile> Object sees M\<^sub>0, b : Ts\<rightarrow>T = m in D"
+ by(auto simp: Method_def dest!: class_add_sees_methods_rev_Obj[where P=P])
+
+subsection "Types and states"
+
+lemma class_add_is_type:
+ "is_type P T \<Longrightarrow> is_type (class_add P (C, cdec)) T"
+ by(cases cdec, simp add: is_type_def is_class_def class_def fun_upd_apply split: ty.splits)
+
+lemma class_add_types:
+ "types P \<subseteq> types (class_add P (C, cdec))"
+using class_add_is_type by(cases cdec, clarsimp)
+
+lemma class_add_states:
+ "states P mxs mxl \<subseteq> states (class_add P (C, cdec)) mxs mxl"
+proof -
+ let ?A = "types P" and ?B = "types (class_add P (C, cdec))"
+ have ab: "?A \<subseteq> ?B" by(rule class_add_types)
+ moreover have "\<And>n. list n ?A \<subseteq> list n ?B" using ab by(rule list_mono)
+ moreover have "list mxl (err ?A) \<subseteq> list mxl (err ?B)" using err_mono[OF ab] by(rule list_mono)
+ ultimately show ?thesis by(auto simp: JVM_states_unfold intro!: err_mono opt_mono)
+qed
+
+lemma class_add_check_types:
+ "check_types P mxs mxl \<tau>s \<Longrightarrow> check_types (class_add P (C, cdec)) mxs mxl \<tau>s"
+using class_add_states by(fastforce simp: check_types_def)
+
+subsection "Subclasses and subtypes"
+
+lemma class_add_subcls:
+ "\<lbrakk> P \<turnstile> D \<preceq>\<^sup>* D'; \<not> is_class P C \<rbrakk>
+ \<Longrightarrow> class_add P (C, cdec) \<turnstile> D \<preceq>\<^sup>* D'"
+proof(induct rule: rtrancl.induct)
+ case (rtrancl_into_rtrancl a b c)
+ then have "b \<noteq> C" by(clarsimp simp: is_class_def dest!: subcls1D)
+ with rtrancl_into_rtrancl show ?case
+ by(fastforce dest!: subcls1D simp: class_def fun_upd_apply
+ intro!: rtrancl_trans[of a b] subcls1I)
+qed(simp)
+
+lemma class_add_subcls_rev:
+ "\<lbrakk> class_add P (C, cdec) \<turnstile> D \<preceq>\<^sup>* D'; \<not>P \<turnstile> D \<preceq>\<^sup>* C \<rbrakk>
+ \<Longrightarrow> P \<turnstile> D \<preceq>\<^sup>* D'"
+proof(induct rule: rtrancl.induct)
+ case (rtrancl_into_rtrancl a b c)
+ then have "b \<noteq> C" by(clarsimp simp: is_class_def dest!: subcls1D)
+ with rtrancl_into_rtrancl show ?case
+ by(fastforce dest!: subcls1D simp: class_def fun_upd_apply
+ intro!: rtrancl_trans[of a b] subcls1I)
+qed(simp)
+
+lemma class_add_subtype:
+ "\<lbrakk> subtype P x y; \<not> is_class P C \<rbrakk>
+ \<Longrightarrow> subtype (class_add P (C, cdec)) x y"
+proof(induct rule: widen.induct)
+ case (widen_subcls C D)
+ then show ?case using class_add_subcls by simp
+qed(simp+)
+
+lemma class_add_widens:
+ "\<lbrakk> P \<turnstile> Ts [\<le>] Ts'; \<not> is_class P C \<rbrakk>
+ \<Longrightarrow> (class_add P (C, cdec)) \<turnstile> Ts [\<le>] Ts'"
+using class_add_subtype by (metis (no_types) list_all2_mono)
+
+lemma class_add_sup_ty_opt:
+ "\<lbrakk> P \<turnstile> l1 \<le>\<^sub>\<top> l2; \<not> is_class P C \<rbrakk>
+ \<Longrightarrow> class_add P (C, cdec) \<turnstile> l1 \<le>\<^sub>\<top> l2"
+using class_add_subtype by(auto simp: sup_ty_opt_def Err.le_def lesub_def split: err.splits)
+
+lemma class_add_sup_loc:
+"\<lbrakk> P \<turnstile> LT [\<le>\<^sub>\<top>] LT'; \<not> is_class P C \<rbrakk>
+ \<Longrightarrow> class_add P (C, cdec) \<turnstile> LT [\<le>\<^sub>\<top>] LT'"
+using class_add_sup_ty_opt[where P=P and C=C] by (simp add: list.rel_mono_strong)
+
+lemma class_add_sup_state:
+ "\<lbrakk> P \<turnstile> \<tau> \<le>\<^sub>i \<tau>'; \<not> is_class P C \<rbrakk>
+ \<Longrightarrow> class_add P (C, cdec) \<turnstile> \<tau> \<le>\<^sub>i \<tau>'"
+using class_add_subtype class_add_sup_ty_opt
+ by(auto simp: sup_state_def Listn.le_def Product.le_def lesub_def class_add_widens
+ class_add_sup_ty_opt list_all2_mono)
+
+lemma class_add_sup_state_opt:
+ "\<lbrakk> P \<turnstile> \<tau> \<le>' \<tau>'; \<not> is_class P C \<rbrakk>
+ \<Longrightarrow> class_add P (C, cdec) \<turnstile> \<tau> \<le>' \<tau>'"
+ by(auto simp: sup_state_opt_def Opt.le_def lesub_def class_add_widens
+ class_add_sup_ty_opt list_all2_mono)
+
+subsection "Effect"
+
+lemma class_add_is_relevant_class:
+ "\<lbrakk> is_relevant_class i P C\<^sub>0; \<not> is_class P C \<rbrakk>
+ \<Longrightarrow> is_relevant_class i (class_add P (C, cdec)) C\<^sub>0"
+ by(cases i, auto simp: class_add_subcls)
+
+lemma class_add_is_relevant_class_rev:
+assumes irc: "is_relevant_class i (class_add P (C, cdec)) C\<^sub>0"
+ and ncp: "\<And>cd D'. cd \<in> set P \<Longrightarrow> \<not>P \<turnstile> fst cd \<preceq>\<^sup>* C"
+ and wfxp: "wf_syscls P"
+shows "is_relevant_class i P C\<^sub>0"
+using assms
+proof(cases i)
+ case (Getfield F D) with assms
+ show ?thesis by(fastforce simp: wf_syscls_def sys_xcpts_def dest!: class_add_subcls_rev)
+next
+ case (Putfield F D) with assms
+ show ?thesis by(fastforce simp: wf_syscls_def sys_xcpts_def dest!: class_add_subcls_rev)
+next
+ case (Checkcast D) with assms
+ show ?thesis by(fastforce simp: wf_syscls_def sys_xcpts_def dest!: class_add_subcls_rev)
+qed(simp_all)
+
+lemma class_add_is_relevant_entry:
+ "\<lbrakk> is_relevant_entry P i pc e; \<not> is_class P C \<rbrakk>
+ \<Longrightarrow> is_relevant_entry (class_add P (C, cdec)) i pc e"
+ by(clarsimp simp: is_relevant_entry_def class_add_is_relevant_class)
+
+lemma class_add_is_relevant_entry_rev:
+ "\<lbrakk> is_relevant_entry (class_add P (C, cdec)) i pc e;
+ \<And>cd D'. cd \<in> set P \<Longrightarrow> \<not>P \<turnstile> fst cd \<preceq>\<^sup>* C;
+ wf_syscls P \<rbrakk>
+ \<Longrightarrow> is_relevant_entry P i pc e"
+ by(auto simp: is_relevant_entry_def dest!: class_add_is_relevant_class_rev)
+
+lemma class_add_relevant_entries:
+ "\<not> is_class P C
+ \<Longrightarrow> set (relevant_entries P i pc xt) \<subseteq> set (relevant_entries (class_add P (C, cdec)) i pc xt)"
+ by(clarsimp simp: relevant_entries_def class_add_is_relevant_entry)
+
+lemma class_add_relevant_entries_eq:
+assumes wf: "wf_prog wf_md P" and nclass: "\<not> is_class P C"
+shows "relevant_entries P i pc xt = relevant_entries (class_add P (C, cdec)) i pc xt"
+proof -
+ have ncp: "\<And>cd D'. cd \<in> set P \<Longrightarrow> \<not>P \<turnstile> fst cd \<preceq>\<^sup>* C"
+ by(rule wf_subcls_nCls'[OF assms])
+ moreover from wf have wfsys: "wf_syscls P" by(simp add: wf_prog_def)
+ moreover
+ note class_add_is_relevant_entry[OF _ nclass, of i pc _ cdec]
+ class_add_is_relevant_entry_rev[OF _ ncp wfsys, of cdec i pc]
+ ultimately show ?thesis by (metis filter_cong relevant_entries_def)
+qed
+
+lemma class_add_norm_eff_pc:
+assumes ne: "\<forall>(pc',\<tau>') \<in> set (norm_eff i P pc \<tau>). pc' < mpc"
+shows "\<forall>(pc',\<tau>') \<in> set (norm_eff i (class_add P (C, cdec)) pc \<tau>). pc' < mpc"
+using assms by(cases i, auto simp: norm_eff_def)
+
+lemma class_add_norm_eff_sup_state_opt:
+assumes ne: "\<forall>(pc',\<tau>') \<in> set (norm_eff i P pc \<tau>). P \<turnstile> \<tau>' \<le>' \<tau>s!pc'"
+ and nclass: "\<not> is_class P C" and app: "app\<^sub>i (i, P, pc, mxs, T, \<tau>)"
+shows "\<forall>(pc',\<tau>') \<in> set (norm_eff i (class_add P (C, cdec)) pc \<tau>). (class_add P (C, cdec)) \<turnstile> \<tau>' \<le>' \<tau>s!pc'"
+proof -
+ obtain ST LT where "\<tau> = (ST,LT)" by(cases \<tau>)
+ with assms show ?thesis proof(cases i)
+ qed(fastforce simp: norm_eff_def
+ dest!: class_add_field[where cdec=cdec] class_add_method[where cdec=cdec]
+ class_add_sup_loc[OF _ nclass] class_add_subtype[OF _ nclass]
+ class_add_widens[OF _ nclass] class_add_sup_state_opt[OF _ nclass])+
+qed
+
+lemma class_add_xcpt_eff_eq:
+assumes wf: "wf_prog wf_md P" and nclass: "\<not> is_class P C"
+shows "xcpt_eff i P pc \<tau> xt = xcpt_eff i (class_add P (C, cdec)) pc \<tau> xt"
+using class_add_relevant_entries_eq[OF assms, of i pc xt cdec] by(cases \<tau>, simp add: xcpt_eff_def)
+
+lemma class_add_eff_pc:
+assumes eff: "\<forall>(pc',\<tau>') \<in> set (eff i P pc xt (Some \<tau>)). pc' < mpc"
+ and wf: "wf_prog wf_md P" and nclass: "\<not> is_class P C"
+shows "\<forall>(pc',\<tau>') \<in> set (eff i (class_add P (C, cdec)) pc xt (Some \<tau>)). pc' < mpc"
+using eff class_add_norm_eff_pc class_add_xcpt_eff_eq[OF wf nclass]
+ by(auto simp: norm_eff_def eff_def)
+
+lemma class_add_eff_sup_state_opt:
+assumes eff: "\<forall>(pc',\<tau>') \<in> set (eff i P pc xt (Some \<tau>)). P \<turnstile> \<tau>' \<le>' \<tau>s!pc'"
+ and wf: "wf_prog wf_md P"and nclass: "\<not> is_class P C"
+ and app: "app\<^sub>i (i, P, pc, mxs, T, \<tau>)"
+shows "\<forall>(pc',\<tau>') \<in> set (eff i (class_add P (C, cdec)) pc xt (Some \<tau>)).
+ (class_add P (C, cdec)) \<turnstile> \<tau>' \<le>' \<tau>s!pc'"
+proof -
+ from eff have ne: "\<forall>(pc', \<tau>')\<in>set (norm_eff i P pc \<tau>). P \<turnstile> \<tau>' \<le>' \<tau>s ! pc'"
+ by(simp add: norm_eff_def eff_def)
+ from eff have "\<forall>(pc', \<tau>')\<in>set (xcpt_eff i P pc \<tau> xt). P \<turnstile> \<tau>' \<le>' \<tau>s ! pc'"
+ by(simp add: xcpt_eff_def eff_def)
+ with class_add_norm_eff_sup_state_opt[OF ne nclass app]
+ class_add_xcpt_eff_eq[OF wf nclass]class_add_sup_state_opt[OF _ nclass]
+ show ?thesis by(cases cdec, auto simp: eff_def norm_eff_def xcpt_app_def)
+qed
+
+
+lemma class_add_app\<^sub>i:
+assumes "app\<^sub>i (i, P, pc, mxs, T\<^sub>r, ST, LT)" and "\<not> is_class P C"
+shows "app\<^sub>i (i, class_add P (C, cdec), pc, mxs, T\<^sub>r, ST, LT)"
+using assms
+proof(cases i)
+ case New then show ?thesis using assms by(fastforce simp: is_class_def class_def fun_upd_apply)
+next
+ case Getfield then show ?thesis using assms
+ by(auto simp: class_add_subtype dest!: class_add_sees_field[where P=P])
+next
+ case Getstatic then show ?thesis using assms by(auto dest!: class_add_sees_field[where P=P])
+next
+ case Putfield then show ?thesis using assms
+ by(auto dest!: class_add_subtype[where P=P] class_add_sees_field[where P=P])
+next
+ case Putstatic then show ?thesis using assms
+ by(auto dest!: class_add_subtype[where P=P] class_add_sees_field[where P=P])
+next
+ case Checkcast then show ?thesis using assms
+ by(clarsimp simp: is_class_def class_def fun_upd_apply)
+next
+ case Invoke then show ?thesis using assms
+ by(fastforce dest!: class_add_widens[where P=P] class_add_sees_method[where P=P])
+next
+ case Invokestatic then show ?thesis using assms
+ by(fastforce dest!: class_add_widens[where P=P] class_add_sees_method[where P=P])
+next
+ case Return then show ?thesis using assms by(clarsimp simp: class_add_subtype)
+qed(simp+)
+
+lemma class_add_xcpt_app:
+assumes xa: "xcpt_app i P pc mxs xt \<tau>"
+ and wf: "wf_prog wf_md P" and nclass: "\<not> is_class P C"
+shows "xcpt_app i (class_add P (C, cdec)) pc mxs xt \<tau>"
+using xa class_add_relevant_entries_eq[OF wf nclass] nclass
+ by(auto simp: xcpt_app_def is_class_def class_def fun_upd_apply) auto
+
+lemma class_add_app:
+assumes app: "app i P mxs T pc mpc xt t"
+ and wf: "wf_prog wf_md P" and nclass: "\<not> is_class P C"
+shows "app i (class_add P (C, cdec)) mxs T pc mpc xt t"
+proof(cases t)
+ case (Some \<tau>)
+ let ?P = "class_add P (C, cdec)"
+ from assms Some have eff: "\<forall>(pc', \<tau>')\<in>set (eff i P pc xt \<lfloor>\<tau>\<rfloor>). pc' < mpc" by(simp add: app_def)
+ from assms Some have app\<^sub>i: "app\<^sub>i (i,P,pc,mxs,T,\<tau>)" by(simp add: app_def)
+ with class_add_app\<^sub>i[OF _ nclass] Some have "app\<^sub>i (i,?P,pc,mxs,T,\<tau>)" by(cases \<tau>,simp)
+ moreover
+ from app class_add_xcpt_app[OF _ wf nclass] Some
+ have "xcpt_app i ?P pc mxs xt \<tau>" by(simp add: app_def del: xcpt_app_def)
+ moreover
+ from app class_add_eff_pc[OF eff wf nclass] Some
+ have "\<forall>(pc',\<tau>') \<in> set (eff i ?P pc xt t). pc' < mpc" by auto
+ moreover note app Some
+ ultimately show ?thesis by(simp add: app_def)
+qed(simp)
+
+subsection "Well-formedness and well-typedness"
+
+lemma class_add_wf_mdecl:
+ "\<lbrakk> wf_mdecl wf_md P C\<^sub>0 md;
+ \<And>C\<^sub>0 md. wf_md P C\<^sub>0 md \<Longrightarrow> wf_md (class_add P (C, cdec)) C\<^sub>0 md \<rbrakk>
+ \<Longrightarrow> wf_mdecl wf_md (class_add P (C, cdec)) C\<^sub>0 md"
+ by(clarsimp simp: wf_mdecl_def class_add_is_type)
+
+lemma class_add_wf_mdecl':
+assumes wfd: "wf_mdecl wf_md P C\<^sub>0 md"
+ and ms: "(C\<^sub>0,S,fs,ms) \<in> set P" and md: "md \<in> set ms"
+ and wf_md': "\<And>C\<^sub>0 S fs ms m.\<lbrakk>(C\<^sub>0,S,fs,ms) \<in> set P; m \<in> set ms\<rbrakk> \<Longrightarrow> wf_md' (class_add P (C, cdec)) C\<^sub>0 m"
+shows "wf_mdecl wf_md' (class_add P (C, cdec)) C\<^sub>0 md"
+using assms by(clarsimp simp: wf_mdecl_def class_add_is_type)
+
+lemma class_add_wf_cdecl:
+assumes wfcd: "wf_cdecl wf_md P cd" and cdP: "cd \<in> set P"
+ and ncp: "\<not> P \<turnstile> fst cd \<preceq>\<^sup>* C" and dist: "distinct_fst P"
+ and wfmd: "\<And>C\<^sub>0 md. wf_md P C\<^sub>0 md \<Longrightarrow> wf_md (class_add P (C, cdec)) C\<^sub>0 md"
+ and nclass: "\<not> is_class P C"
+shows "wf_cdecl wf_md (class_add P (C, cdec)) cd"
+proof -
+ let ?P = "class_add P (C, cdec)"
+ obtain C1 D fs ms where [simp]: "cd = (C1,(D,fs,ms))" by(cases cd)
+ from wfcd
+ have "\<forall>f\<in>set fs. wf_fdecl ?P f" by(auto simp: wf_cdecl_def wf_fdecl_def class_add_is_type)
+ moreover
+ from wfcd wfmd class_add_wf_mdecl
+ have "\<forall>m\<in>set ms. wf_mdecl wf_md ?P C1 m" by(auto simp: wf_cdecl_def)
+ moreover
+ have "C1 \<noteq> Object \<Longrightarrow> is_class ?P D \<and> \<not> ?P \<turnstile> D \<preceq>\<^sup>* C1
+ \<and> (\<forall>(M,b,Ts,T,m)\<in>set ms.
+ \<forall>D' b' Ts' T' m'. ?P \<turnstile> D sees M,b':Ts' \<rightarrow> T' = m' in D' \<longrightarrow>
+ b = b' \<and> ?P \<turnstile> Ts' [\<le>] Ts \<and> ?P \<turnstile> T \<le> T')"
+ proof -
+ assume nObj[simp]: "C1 \<noteq> Object"
+ with cdP dist have sub1: "P \<turnstile> C1 \<prec>\<^sup>1 D" by(auto simp: class_def intro!: subcls1I map_of_SomeI)
+ with ncp have ncp': "\<not> P \<turnstile> D \<preceq>\<^sup>* C" by(auto simp: converse_rtrancl_into_rtrancl)
+ with wfcd
+ have clsD: "is_class ?P D"
+ by(auto simp: wf_cdecl_def is_class_def class_def fun_upd_apply)
+ moreover
+ from wfcd sub1
+ have "\<not> ?P \<turnstile> D \<preceq>\<^sup>* C1" by(auto simp: wf_cdecl_def dest!: class_add_subcls_rev[OF _ ncp'])
+ moreover
+ have "\<And>M b Ts T m D' b' Ts' T' m'. (M,b,Ts,T,m) \<in> set ms
+ \<Longrightarrow> ?P \<turnstile> D sees M,b':Ts' \<rightarrow> T' = m' in D'
+ \<Longrightarrow> b = b' \<and> ?P \<turnstile> Ts' [\<le>] Ts \<and> ?P \<turnstile> T \<le> T'"
+ proof -
+ fix M b Ts T m D' b' Ts' T' m'
+ assume ms: "(M,b,Ts,T,m) \<in> set ms" and meth': "?P \<turnstile> D sees M,b':Ts' \<rightarrow> T' = m' in D'"
+ with sub1
+ have "P \<turnstile> D sees M,b':Ts' \<rightarrow> T' = m' in D'"
+ by(fastforce dest!: class_add_sees_method_rev[OF _ ncp'])
+ moreover
+ with wfcd ms meth'
+ have "b = b' \<and> P \<turnstile> Ts' [\<le>] Ts \<and> P \<turnstile> T \<le> T'"
+ by(cases m', fastforce simp: wf_cdecl_def elim!: ballE[where x="(M,b,Ts,T,m)"])
+ ultimately show "b = b' \<and> ?P \<turnstile> Ts' [\<le>] Ts \<and> ?P \<turnstile> T \<le> T'"
+ by(auto dest!: class_add_subtype[OF _ nclass] class_add_widens[OF _ nclass])
+ qed
+ ultimately show ?thesis by clarsimp
+ qed
+ moreover note wfcd
+ ultimately show ?thesis by(simp add: wf_cdecl_def)
+qed
+
+lemma class_add_wf_cdecl':
+assumes wfcd: "wf_cdecl wf_md P cd" and cdP: "cd \<in> set P"
+ and ncp: "\<not>P \<turnstile> fst cd \<preceq>\<^sup>* C" and dist: "distinct_fst P"
+ and wfmd: "\<And>C\<^sub>0 S fs ms m.\<lbrakk>(C\<^sub>0,S,fs,ms) \<in> set P; m \<in> set ms\<rbrakk> \<Longrightarrow> wf_md' (class_add P (C, cdec)) C\<^sub>0 m"
+ and nclass: "\<not> is_class P C"
+shows "wf_cdecl wf_md' (class_add P (C, cdec)) cd"
+proof -
+ let ?P = "class_add P (C, cdec)"
+ obtain C1 D fs ms where [simp]: "cd = (C1,(D,fs,ms))" by(cases cd)
+ from wfcd
+ have "\<forall>f\<in>set fs. wf_fdecl ?P f" by(auto simp: wf_cdecl_def wf_fdecl_def class_add_is_type)
+ moreover
+ from cdP wfcd wfmd
+ have "\<forall>m\<in>set ms. wf_mdecl wf_md' ?P C1 m"
+ by(auto simp: wf_cdecl_def wf_mdecl_def class_add_is_type)
+ moreover
+ have "C1 \<noteq> Object \<Longrightarrow> is_class ?P D \<and> \<not> ?P \<turnstile> D \<preceq>\<^sup>* C1
+ \<and> (\<forall>(M,b,Ts,T,m)\<in>set ms.
+ \<forall>D' b' Ts' T' m'. ?P \<turnstile> D sees M,b':Ts' \<rightarrow> T' = m' in D' \<longrightarrow>
+ b = b' \<and> ?P \<turnstile> Ts' [\<le>] Ts \<and> ?P \<turnstile> T \<le> T')"
+ proof -
+ assume nObj[simp]: "C1 \<noteq> Object"
+ with cdP dist have sub1: "P \<turnstile> C1 \<prec>\<^sup>1 D" by(auto simp: class_def intro!: subcls1I map_of_SomeI)
+ with ncp have ncp': "\<not> P \<turnstile> D \<preceq>\<^sup>* C" by(auto simp: converse_rtrancl_into_rtrancl)
+ with wfcd
+ have clsD: "is_class ?P D"
+ by(auto simp: wf_cdecl_def is_class_def class_def fun_upd_apply)
+ moreover
+ from wfcd sub1
+ have "\<not> ?P \<turnstile> D \<preceq>\<^sup>* C1" by(auto simp: wf_cdecl_def dest!: class_add_subcls_rev[OF _ ncp'])
+ moreover
+ have "\<And>M b Ts T m D' b' Ts' T' m'. (M,b,Ts,T,m) \<in> set ms
+ \<Longrightarrow> ?P \<turnstile> D sees M,b':Ts' \<rightarrow> T' = m' in D'
+ \<Longrightarrow> b = b' \<and> ?P \<turnstile> Ts' [\<le>] Ts \<and> ?P \<turnstile> T \<le> T'"
+ proof -
+ fix M b Ts T m D' b' Ts' T' m'
+ assume ms: "(M,b,Ts,T,m) \<in> set ms" and meth': "?P \<turnstile> D sees M,b':Ts' \<rightarrow> T' = m' in D'"
+ with sub1
+ have "P \<turnstile> D sees M,b':Ts' \<rightarrow> T' = m' in D'"
+ by(fastforce dest!: class_add_sees_method_rev[OF _ ncp'])
+ moreover
+ with wfcd ms meth'
+ have "b = b' \<and> P \<turnstile> Ts' [\<le>] Ts \<and> P \<turnstile> T \<le> T'"
+ by(cases m', fastforce simp: wf_cdecl_def elim!: ballE[where x="(M,b,Ts,T,m)"])
+ ultimately show "b = b' \<and> ?P \<turnstile> Ts' [\<le>] Ts \<and> ?P \<turnstile> T \<le> T'"
+ by(auto dest!: class_add_subtype[OF _ nclass] class_add_widens[OF _ nclass])
+ qed
+ ultimately show ?thesis by clarsimp
+ qed
+ moreover note wfcd
+ ultimately show ?thesis by(simp add: wf_cdecl_def)
+qed
+
+lemma class_add_wt_start:
+ "\<lbrakk> wt_start P C\<^sub>0 b Ts mxl \<tau>s; \<not> is_class P C \<rbrakk>
+ \<Longrightarrow> wt_start (class_add P (C, cdec)) C\<^sub>0 b Ts mxl \<tau>s"
+using class_add_sup_state_opt by(clarsimp simp: wt_start_def split: staticb.splits)
+
+lemma class_add_wt_instr:
+assumes wti: "P,T,mxs,mpc,xt \<turnstile> i,pc :: \<tau>s"
+ and wf: "wf_prog wf_md P" and nclass: "\<not> is_class P C"
+shows "class_add P (C, cdec),T,mxs,mpc,xt \<turnstile> i,pc :: \<tau>s"
+proof -
+ let ?P = "class_add P (C, cdec)"
+ from wti have eff: "\<forall>(pc', \<tau>')\<in>set (eff i P pc xt (\<tau>s ! pc)). P \<turnstile> \<tau>' \<le>' \<tau>s ! pc'"
+ by(simp add: wt_instr_def)
+ from wti have app\<^sub>i: "\<tau>s!pc \<noteq> None \<Longrightarrow> app\<^sub>i (i,P,pc,mxs,T,the (\<tau>s!pc))"
+ by(simp add: wt_instr_def app_def)
+ from wti class_add_app[OF _ wf nclass]
+ have "app i ?P mxs T pc mpc xt (\<tau>s!pc)" by(simp add: wt_instr_def)
+ moreover
+ have "\<forall>(pc',\<tau>') \<in> set (eff i ?P pc xt (\<tau>s!pc)). ?P \<turnstile> \<tau>' \<le>' \<tau>s!pc'"
+ proof(cases "\<tau>s!pc")
+ case Some with eff class_add_eff_sup_state_opt[OF _ wf nclass app\<^sub>i] show ?thesis by auto
+ qed(simp add: eff_def)
+ moreover note wti
+ ultimately show ?thesis by(clarsimp simp: wt_instr_def)
+qed
+
+lemma class_add_wt_method:
+assumes wtm: "wt_method P C\<^sub>0 b Ts T\<^sub>r mxs mxl\<^sub>0 is xt (\<Phi> C\<^sub>0 M\<^sub>0)"
+ and wf: "wf_prog wf_md P" and nclass: "\<not> is_class P C"
+shows "wt_method (class_add P (C, cdec)) C\<^sub>0 b Ts T\<^sub>r mxs mxl\<^sub>0 is xt (\<Phi> C\<^sub>0 M\<^sub>0)"
+proof -
+ let ?P = "class_add P (C, cdec)"
+ let ?\<tau>s = "\<Phi> C\<^sub>0 M\<^sub>0"
+ from wtm class_add_check_types
+ have "check_types ?P mxs ((case b of Static \<Rightarrow> 0 | NonStatic \<Rightarrow> 1)+size Ts+mxl\<^sub>0) (map OK ?\<tau>s)"
+ by(simp add: wt_method_def)
+ moreover
+ from wtm class_add_wt_start nclass
+ have "wt_start ?P C\<^sub>0 b Ts mxl\<^sub>0 ?\<tau>s" by(simp add: wt_method_def)
+ moreover
+ from wtm class_add_wt_instr[OF _ wf nclass]
+ have "\<forall>pc < size is. ?P,T\<^sub>r,mxs,size is,xt \<turnstile> is!pc,pc :: ?\<tau>s" by(clarsimp simp: wt_method_def)
+ moreover note wtm
+ ultimately
+ show ?thesis by(clarsimp simp: wt_method_def)
+qed
+
+lemma class_add_wt_method':
+ "\<lbrakk> (\<lambda>P C (M,b,Ts,T\<^sub>r,(mxs,mxl\<^sub>0,is,xt)). wt_method P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt (\<Phi> C M)) P C\<^sub>0 md;
+ wf_prog wf_md P; \<not> is_class P C \<rbrakk>
+ \<Longrightarrow> (\<lambda>P C (M,b,Ts,T\<^sub>r,(mxs,mxl\<^sub>0,is,xt)). wt_method P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt (\<Phi> C M))
+ (class_add P (C, cdec)) C\<^sub>0 md"
+ by(clarsimp simp: class_add_wt_method)
+
+subsection \<open> @{text "distinct_fst"} \<close>
+
+lemma class_add_distinct_fst:
+"\<lbrakk> distinct_fst P; \<not> is_class P C \<rbrakk>
+ \<Longrightarrow> distinct_fst (class_add P (C, cdec))"
+ by(clarsimp simp: distinct_fst_def is_class_def class_def)
+
+subsection "Conformance"
+
+lemma class_add_conf:
+ "\<lbrakk> P,h \<turnstile> v :\<le> T; \<not> is_class P C \<rbrakk>
+ \<Longrightarrow> class_add P (C, cdec),h \<turnstile> v :\<le> T"
+ by(clarsimp simp: conf_def class_add_subtype)
+
+lemma class_add_oconf:
+fixes obj::obj
+assumes oc: "P,h \<turnstile> obj \<surd>" and ns: "\<not> is_class P C"
+ and ncp: "\<And>D'. P \<turnstile> fst(obj) \<preceq>\<^sup>* D' \<Longrightarrow> D' \<noteq> C"
+shows "(class_add P (C, cdec)),h \<turnstile> obj \<surd>"
+proof -
+ obtain C\<^sub>0 fs where [simp]: "obj=(C\<^sub>0,fs)" by(cases obj)
+ from oc have
+ oc': "\<And>F D T. P \<turnstile> C\<^sub>0 has F,NonStatic:T in D \<Longrightarrow> (\<exists>v. fs (F, D) = \<lfloor>v\<rfloor> \<and> P,h \<turnstile> v :\<le> T)"
+ by(simp add: oconf_def)
+ have "\<And>F D T. class_add P (C, cdec) \<turnstile> C\<^sub>0 has F,NonStatic:T in D
+ \<Longrightarrow> \<exists>v. fs(F,D) = Some v \<and> class_add P (C, cdec),h \<turnstile> v :\<le> T"
+ proof -
+ fix F D T assume "class_add P (C, cdec) \<turnstile> C\<^sub>0 has F,NonStatic:T in D"
+ with class_add_has_field_rev[OF _ ncp] have meth: "P \<turnstile> C\<^sub>0 has F,NonStatic:T in D" by simp
+ then show "\<exists>v. fs(F,D) = Some v \<and> class_add P (C, cdec),h \<turnstile> v :\<le> T"
+ using oc'[OF meth] class_add_conf[OF _ ns] by(fastforce simp: oconf_def)
+ qed
+ then show ?thesis by(simp add: oconf_def)
+qed
+
+lemma class_add_soconf:
+assumes soc: "P,h,C\<^sub>0 \<turnstile>\<^sub>s sfs \<surd>" and ns: "\<not> is_class P C"
+ and ncp: "\<And>D'. P \<turnstile> C\<^sub>0 \<preceq>\<^sup>* D' \<Longrightarrow> D' \<noteq> C"
+shows "(class_add P (C, cdec)),h,C\<^sub>0 \<turnstile>\<^sub>s sfs \<surd>"
+proof -
+ from soc have
+ oc': "\<And>F T. P \<turnstile> C\<^sub>0 has F,Static:T in C\<^sub>0 \<Longrightarrow> (\<exists>v. sfs F = \<lfloor>v\<rfloor> \<and> P,h \<turnstile> v :\<le> T)"
+ by(simp add: soconf_def)
+ have "\<And>F T. class_add P (C, cdec) \<turnstile> C\<^sub>0 has F,Static:T in C\<^sub>0
+ \<Longrightarrow> \<exists>v. sfs F = Some v \<and> class_add P (C, cdec),h \<turnstile> v :\<le> T"
+ proof -
+ fix F T assume "class_add P (C, cdec) \<turnstile> C\<^sub>0 has F,Static:T in C\<^sub>0"
+ with class_add_has_field_rev[OF _ ncp] have meth: "P \<turnstile> C\<^sub>0 has F,Static:T in C\<^sub>0" by simp
+ then show "\<exists>v. sfs F = Some v \<and> class_add P (C, cdec),h \<turnstile> v :\<le> T"
+ using oc'[OF meth] class_add_conf[OF _ ns] by(fastforce simp: soconf_def)
+ qed
+ then show ?thesis by(simp add: soconf_def)
+qed
+
+lemma class_add_hconf:
+assumes "P \<turnstile> h \<surd>" and "\<not> is_class P C"
+ and "\<And>a obj D'. h a = Some obj \<Longrightarrow> P \<turnstile> fst(obj) \<preceq>\<^sup>* D' \<Longrightarrow> D' \<noteq> C"
+shows "class_add P (C, cdec) \<turnstile> h \<surd>"
+using assms by(auto simp: hconf_def intro!: class_add_oconf)
+
+lemma class_add_hconf_wf:
+assumes wf: "wf_prog wf_md P" and "P \<turnstile> h \<surd>" and "\<not> is_class P C"
+ and "\<And>a obj. h a = Some obj \<Longrightarrow> fst(obj) \<noteq> C"
+shows "class_add P (C, cdec) \<turnstile> h \<surd>"
+using wf_subcls_nCls[OF wf] assms by(fastforce simp: hconf_def intro!: class_add_oconf)
+
+lemma class_add_shconf:
+assumes "P,h \<turnstile>\<^sub>s sh \<surd>" and ns: "\<not> is_class P C"
+ and "\<And>C sobj D'. sh C = Some sobj \<Longrightarrow> P \<turnstile> C \<preceq>\<^sup>* D' \<Longrightarrow> D' \<noteq> C"
+shows "class_add P (C, cdec),h \<turnstile>\<^sub>s sh \<surd>"
+using assms by(fastforce simp: shconf_def)
+
+lemma class_add_shconf_wf:
+assumes wf: "wf_prog wf_md P" and "P,h \<turnstile>\<^sub>s sh \<surd>" and "\<not> is_class P C"
+ and "\<And>C sobj. sh C = Some sobj \<Longrightarrow> C \<noteq> C"
+shows "class_add P (C, cdec),h \<turnstile>\<^sub>s sh \<surd>"
+using wf_subcls_nCls[OF wf] assms by(fastforce simp: shconf_def)
+
+
end
\ No newline at end of file
diff --git a/thys/JinjaDCI/BV/Effect.thy b/thys/JinjaDCI/BV/Effect.thy
--- a/thys/JinjaDCI/BV/Effect.thy
+++ b/thys/JinjaDCI/BV/Effect.thy
@@ -1,443 +1,443 @@
-(* Title: JinjaDCI/BV/Effect.thy
- Author: Gerwin Klein, Susannah Mansky
- Copyright 2000 Technische Universitaet Muenchen, 2019-20 UIUC
-
- Based on the Jinja theory BV/Effect.thy by Gerwin Klein
-*)
-
-section \<open>Effect of Instructions on the State Type\<close>
-
-theory Effect
-imports JVM_SemiType "../JVM/JVMExceptions"
-begin
-
-\<comment> \<open>FIXME\<close>
-locale prog =
- fixes P :: "'a prog"
-
-locale jvm_method = prog +
- fixes mxs :: nat
- fixes mxl\<^sub>0 :: nat
- fixes Ts :: "ty list"
- fixes T\<^sub>r :: ty
- fixes "is" :: "instr list"
- fixes xt :: ex_table
-
- fixes mxl :: nat
- defines mxl_def: "mxl \<equiv> 1+size Ts+mxl\<^sub>0"
-
-text \<open> Program counter of successor instructions: \<close>
-primrec succs :: "instr \<Rightarrow> ty\<^sub>i \<Rightarrow> pc \<Rightarrow> pc list" where
- "succs (Load idx) \<tau> pc = [pc+1]"
-| "succs (Store idx) \<tau> pc = [pc+1]"
-| "succs (Push v) \<tau> pc = [pc+1]"
-| "succs (Getfield F C) \<tau> pc = [pc+1]"
-| "succs (Getstatic C F D) \<tau> pc = [pc+1]"
-| "succs (Putfield F C) \<tau> pc = [pc+1]"
-| "succs (Putstatic C F D) \<tau> pc = [pc+1]"
-| "succs (New C) \<tau> pc = [pc+1]"
-| "succs (Checkcast C) \<tau> pc = [pc+1]"
-| "succs Pop \<tau> pc = [pc+1]"
-| "succs IAdd \<tau> pc = [pc+1]"
-| "succs CmpEq \<tau> pc = [pc+1]"
-| succs_IfFalse:
- "succs (IfFalse b) \<tau> pc = [pc+1, nat (int pc + b)]"
-| succs_Goto:
- "succs (Goto b) \<tau> pc = [nat (int pc + b)]"
-| succs_Return:
- "succs Return \<tau> pc = []"
-| succs_Invoke:
- "succs (Invoke M n) \<tau> pc = (if (fst \<tau>)!n = NT then [] else [pc+1])"
-| succs_Invokestatic:
- "succs (Invokestatic C M n) \<tau> pc = [pc+1]"
-| succs_Throw:
- "succs Throw \<tau> pc = []"
-
-text "Effect of instruction on the state type:"
-
-fun the_class:: "ty \<Rightarrow> cname" where
- "the_class (Class C) = C"
-
-fun eff\<^sub>i :: "instr \<times> 'm prog \<times> ty\<^sub>i \<Rightarrow> ty\<^sub>i" where
- eff\<^sub>i_Load:
- "eff\<^sub>i (Load n, P, (ST, LT)) = (ok_val (LT ! n) # ST, LT)"
-| eff\<^sub>i_Store:
- "eff\<^sub>i (Store n, P, (T#ST, LT)) = (ST, LT[n:= OK T])"
-| eff\<^sub>i_Push:
- "eff\<^sub>i (Push v, P, (ST, LT)) = (the (typeof v) # ST, LT)"
-| eff\<^sub>i_Getfield:
- "eff\<^sub>i (Getfield F C, P, (T#ST, LT)) = (snd (snd (field P C F)) # ST, LT)"
-| eff\<^sub>i_Getstatic:
- "eff\<^sub>i (Getstatic C F D, P, (ST, LT)) = (snd (snd (field P C F)) # ST, LT)"
-| eff\<^sub>i_Putfield:
- "eff\<^sub>i (Putfield F C, P, (T\<^sub>1#T\<^sub>2#ST, LT)) = (ST,LT)"
-| eff\<^sub>i_Putstatic:
- "eff\<^sub>i (Putstatic C F D, P, (T#ST, LT)) = (ST,LT)"
-| eff\<^sub>i_New:
- "eff\<^sub>i (New C, P, (ST,LT)) = (Class C # ST, LT)"
-| eff\<^sub>i_Checkcast:
- "eff\<^sub>i (Checkcast C, P, (T#ST,LT)) = (Class C # ST,LT)"
-| eff\<^sub>i_Pop:
- "eff\<^sub>i (Pop, P, (T#ST,LT)) = (ST,LT)"
-| eff\<^sub>i_IAdd:
- "eff\<^sub>i (IAdd, P,(T\<^sub>1#T\<^sub>2#ST,LT)) = (Integer#ST,LT)"
-| eff\<^sub>i_CmpEq:
- "eff\<^sub>i (CmpEq, P, (T\<^sub>1#T\<^sub>2#ST,LT)) = (Boolean#ST,LT)"
-| eff\<^sub>i_IfFalse:
- "eff\<^sub>i (IfFalse b, P, (T\<^sub>1#ST,LT)) = (ST,LT)"
-| eff\<^sub>i_Invoke:
- "eff\<^sub>i (Invoke M n, P, (ST,LT)) =
- (let C = the_class (ST!n); (D,b,Ts,T\<^sub>r,m) = method P C M
- in (T\<^sub>r # drop (n+1) ST, LT))"
-| eff\<^sub>i_Invokestatic:
- "eff\<^sub>i (Invokestatic C M n, P, (ST,LT)) =
- (let (D,b,Ts,T\<^sub>r,m) = method P C M
- in (T\<^sub>r # drop n ST, LT))"
-| eff\<^sub>i_Goto:
- "eff\<^sub>i (Goto n, P, s) = s"
-
-fun is_relevant_class :: "instr \<Rightarrow> 'm prog \<Rightarrow> cname \<Rightarrow> bool" where
- rel_Getfield:
- "is_relevant_class (Getfield F D)
- = (\<lambda>P C. P \<turnstile> NullPointer \<preceq>\<^sup>* C \<or> P \<turnstile> NoSuchFieldError \<preceq>\<^sup>* C
- \<or> P \<turnstile> IncompatibleClassChangeError \<preceq>\<^sup>* C)"
-| rel_Getstatic:
- "is_relevant_class (Getstatic C F D)
- = (\<lambda>P C. True)"
-| rel_Putfield:
- "is_relevant_class (Putfield F D)
- = (\<lambda>P C. P \<turnstile> NullPointer \<preceq>\<^sup>* C \<or> P \<turnstile> NoSuchFieldError \<preceq>\<^sup>* C
- \<or> P \<turnstile> IncompatibleClassChangeError \<preceq>\<^sup>* C)"
-| rel_Putstatic:
- "is_relevant_class (Putstatic C F D)
- = (\<lambda>P C. True)"
-| rel_Checkcast:
- "is_relevant_class (Checkcast D) = (\<lambda>P C. P \<turnstile> ClassCast \<preceq>\<^sup>* C)"
-| rel_New:
- "is_relevant_class (New D) = (\<lambda>P C. True)"
-| rel_Throw:
- "is_relevant_class Throw = (\<lambda>P C. True)"
-| rel_Invoke:
- "is_relevant_class (Invoke M n) = (\<lambda>P C. True)"
-| rel_Invokestatic:
- "is_relevant_class (Invokestatic C M n) = (\<lambda>P C. True)"
-| rel_default:
- "is_relevant_class i = (\<lambda>P C. False)"
-
-definition is_relevant_entry :: "'m prog \<Rightarrow> instr \<Rightarrow> pc \<Rightarrow> ex_entry \<Rightarrow> bool" where
- "is_relevant_entry P i pc e \<longleftrightarrow> (let (f,t,C,h,d) = e in is_relevant_class i P C \<and> pc \<in> {f..<t})"
-
-definition relevant_entries :: "'m prog \<Rightarrow> instr \<Rightarrow> pc \<Rightarrow> ex_table \<Rightarrow> ex_table" where
- "relevant_entries P i pc = filter (is_relevant_entry P i pc)"
-
-definition xcpt_eff :: "instr \<Rightarrow> 'm prog \<Rightarrow> pc \<Rightarrow> ty\<^sub>i
- \<Rightarrow> ex_table \<Rightarrow> (pc \<times> ty\<^sub>i') list" where
- "xcpt_eff i P pc \<tau> et = (let (ST,LT) = \<tau> in
- map (\<lambda>(f,t,C,h,d). (h, Some (Class C#drop (size ST - d) ST, LT))) (relevant_entries P i pc et))"
-
-definition norm_eff :: "instr \<Rightarrow> 'm prog \<Rightarrow> nat \<Rightarrow> ty\<^sub>i \<Rightarrow> (pc \<times> ty\<^sub>i') list" where
- "norm_eff i P pc \<tau> = map (\<lambda>pc'. (pc',Some (eff\<^sub>i (i,P,\<tau>)))) (succs i \<tau> pc)"
-
-definition eff :: "instr \<Rightarrow> 'm prog \<Rightarrow> pc \<Rightarrow> ex_table \<Rightarrow> ty\<^sub>i' \<Rightarrow> (pc \<times> ty\<^sub>i') list" where
- "eff i P pc et t = (case t of
- None \<Rightarrow> []
- | Some \<tau> \<Rightarrow> (norm_eff i P pc \<tau>) @ (xcpt_eff i P pc \<tau> et))"
-
-
-lemma eff_None:
- "eff i P pc xt None = []"
-by (simp add: eff_def)
-
-lemma eff_Some:
- "eff i P pc xt (Some \<tau>) = norm_eff i P pc \<tau> @ xcpt_eff i P pc \<tau> xt"
-by (simp add: eff_def)
-
-(* FIXME: getfield, \<exists>T D. P \<turnstile> C sees F:T in D \<and> .. *)
-
-text "Conditions under which eff is applicable:"
-
-fun app\<^sub>i :: "instr \<times> 'm prog \<times> pc \<times> nat \<times> ty \<times> ty\<^sub>i \<Rightarrow> bool" where
- app\<^sub>i_Load:
- "app\<^sub>i (Load n, P, pc, mxs, T\<^sub>r, (ST,LT)) =
- (n < length LT \<and> LT ! n \<noteq> Err \<and> length ST < mxs)"
-| app\<^sub>i_Store:
- "app\<^sub>i (Store n, P, pc, mxs, T\<^sub>r, (T#ST, LT)) =
- (n < length LT)"
-| app\<^sub>i_Push:
- "app\<^sub>i (Push v, P, pc, mxs, T\<^sub>r, (ST,LT)) =
- (length ST < mxs \<and> typeof v \<noteq> None)"
-| app\<^sub>i_Getfield:
- "app\<^sub>i (Getfield F C, P, pc, mxs, T\<^sub>r, (T#ST, LT)) =
- (\<exists>T\<^sub>f. P \<turnstile> C sees F,NonStatic:T\<^sub>f in C \<and> P \<turnstile> T \<le> Class C)"
-| app\<^sub>i_Getstatic:
- "app\<^sub>i (Getstatic C F D, P, pc, mxs, T\<^sub>r, (ST, LT)) =
- (length ST < mxs \<and> (\<exists>T\<^sub>f. P \<turnstile> C sees F,Static:T\<^sub>f in D))"
-| app\<^sub>i_Putfield:
- "app\<^sub>i (Putfield F C, P, pc, mxs, T\<^sub>r, (T\<^sub>1#T\<^sub>2#ST, LT)) =
- (\<exists>T\<^sub>f. P \<turnstile> C sees F,NonStatic:T\<^sub>f in C \<and> P \<turnstile> T\<^sub>2 \<le> (Class C) \<and> P \<turnstile> T\<^sub>1 \<le> T\<^sub>f)"
-| app\<^sub>i_Putstatic:
- "app\<^sub>i (Putstatic C F D, P, pc, mxs, T\<^sub>r, (T#ST, LT)) =
- (\<exists>T\<^sub>f. P \<turnstile> C sees F,Static:T\<^sub>f in D \<and> P \<turnstile> T \<le> T\<^sub>f)"
-| app\<^sub>i_New:
- "app\<^sub>i (New C, P, pc, mxs, T\<^sub>r, (ST,LT)) =
- (is_class P C \<and> length ST < mxs)"
-| app\<^sub>i_Checkcast:
- "app\<^sub>i (Checkcast C, P, pc, mxs, T\<^sub>r, (T#ST,LT)) =
- (is_class P C \<and> is_refT T)"
-| app\<^sub>i_Pop:
- "app\<^sub>i (Pop, P, pc, mxs, T\<^sub>r, (T#ST,LT)) =
- True"
-| app\<^sub>i_IAdd:
- "app\<^sub>i (IAdd, P, pc, mxs, T\<^sub>r, (T\<^sub>1#T\<^sub>2#ST,LT)) = (T\<^sub>1 = T\<^sub>2 \<and> T\<^sub>1 = Integer)"
-| app\<^sub>i_CmpEq:
- "app\<^sub>i (CmpEq, P, pc, mxs, T\<^sub>r, (T\<^sub>1#T\<^sub>2#ST,LT)) =
- (T\<^sub>1 = T\<^sub>2 \<or> is_refT T\<^sub>1 \<and> is_refT T\<^sub>2)"
-| app\<^sub>i_IfFalse:
- "app\<^sub>i (IfFalse b, P, pc, mxs, T\<^sub>r, (Boolean#ST,LT)) =
- (0 \<le> int pc + b)"
-| app\<^sub>i_Goto:
- "app\<^sub>i (Goto b, P, pc, mxs, T\<^sub>r, s) =
- (0 \<le> int pc + b)"
-| app\<^sub>i_Return:
- "app\<^sub>i (Return, P, pc, mxs, T\<^sub>r, (T#ST,LT)) =
- (P \<turnstile> T \<le> T\<^sub>r)"
-| app\<^sub>i_Throw:
- "app\<^sub>i (Throw, P, pc, mxs, T\<^sub>r, (T#ST,LT)) =
- is_refT T"
-| app\<^sub>i_Invoke:
- "app\<^sub>i (Invoke M n, P, pc, mxs, T\<^sub>r, (ST,LT)) =
- (n < length ST \<and>
- (ST!n \<noteq> NT \<longrightarrow>
- (\<exists>C D Ts T m. ST!n = Class C \<and> P \<turnstile> C sees M,NonStatic:Ts \<rightarrow> T = m in D \<and>
- P \<turnstile> rev (take n ST) [\<le>] Ts)))"
-| app\<^sub>i_Invokestatic:
- "app\<^sub>i (Invokestatic C M n, P, pc, mxs, T\<^sub>r, (ST,LT)) =
- (length ST - n < mxs \<and> n \<le> length ST \<and> M \<noteq> clinit \<and>
- (\<exists>D Ts T m. P \<turnstile> C sees M,Static:Ts \<rightarrow> T = m in D \<and>
- P \<turnstile> rev (take n ST) [\<le>] Ts))"
-
-| app\<^sub>i_default:
- "app\<^sub>i (i,P, pc,mxs,T\<^sub>r,s) = False"
-
-
-definition xcpt_app :: "instr \<Rightarrow> 'm prog \<Rightarrow> pc \<Rightarrow> nat \<Rightarrow> ex_table \<Rightarrow> ty\<^sub>i \<Rightarrow> bool" where
- "xcpt_app i P pc mxs xt \<tau> \<longleftrightarrow> (\<forall>(f,t,C,h,d) \<in> set (relevant_entries P i pc xt). is_class P C \<and> d \<le> size (fst \<tau>) \<and> d < mxs)"
-
-definition app :: "instr \<Rightarrow> 'm prog \<Rightarrow> nat \<Rightarrow> ty \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> ex_table \<Rightarrow> ty\<^sub>i' \<Rightarrow> bool" where
- "app i P mxs T\<^sub>r pc mpc xt t = (case t of None \<Rightarrow> True | Some \<tau> \<Rightarrow>
- app\<^sub>i (i,P,pc,mxs,T\<^sub>r,\<tau>) \<and> xcpt_app i P pc mxs xt \<tau> \<and>
- (\<forall>(pc',\<tau>') \<in> set (eff i P pc xt t). pc' < mpc))"
-
-
-lemma app_Some:
- "app i P mxs T\<^sub>r pc mpc xt (Some \<tau>) =
- (app\<^sub>i (i,P,pc,mxs,T\<^sub>r,\<tau>) \<and> xcpt_app i P pc mxs xt \<tau> \<and>
- (\<forall>(pc',s') \<in> set (eff i P pc xt (Some \<tau>)). pc' < mpc))"
-by (simp add: app_def)
-
-locale eff = jvm_method +
- fixes eff\<^sub>i and app\<^sub>i and eff and app
- fixes norm_eff and xcpt_app and xcpt_eff
-
- fixes mpc
- defines "mpc \<equiv> size is"
-
- defines "eff\<^sub>i i \<tau> \<equiv> Effect.eff\<^sub>i (i,P,\<tau>)"
- notes eff\<^sub>i_simps [simp] = Effect.eff\<^sub>i.simps [where P = P, folded eff\<^sub>i_def]
-
- defines "app\<^sub>i i pc \<tau> \<equiv> Effect.app\<^sub>i (i, P, pc, mxs, T\<^sub>r, \<tau>)"
- notes app\<^sub>i_simps [simp] = Effect.app\<^sub>i.simps [where P=P and mxs=mxs and T\<^sub>r=T\<^sub>r, folded app\<^sub>i_def]
-
- defines "xcpt_eff i pc \<tau> \<equiv> Effect.xcpt_eff i P pc \<tau> xt"
- notes xcpt_eff = Effect.xcpt_eff_def [of _ P _ _ xt, folded xcpt_eff_def]
-
- defines "norm_eff i pc \<tau> \<equiv> Effect.norm_eff i P pc \<tau>"
- notes norm_eff = Effect.norm_eff_def [of _ P, folded norm_eff_def eff\<^sub>i_def]
-
- defines "eff i pc \<equiv> Effect.eff i P pc xt"
- notes eff = Effect.eff_def [of _ P _ xt, folded eff_def norm_eff_def xcpt_eff_def]
-
- defines "xcpt_app i pc \<tau> \<equiv> Effect.xcpt_app i P pc mxs xt \<tau>"
- notes xcpt_app = Effect.xcpt_app_def [of _ P _ mxs xt, folded xcpt_app_def]
-
- defines "app i pc \<equiv> Effect.app i P mxs T\<^sub>r pc mpc xt"
- notes app = Effect.app_def [of _ P mxs T\<^sub>r _ mpc xt, folded app_def xcpt_app_def app\<^sub>i_def eff_def]
-
-
-lemma length_cases2:
- assumes "\<And>LT. P ([],LT)"
- assumes "\<And>l ST LT. P (l#ST,LT)"
- shows "P s"
- by (cases s, cases "fst s") (auto intro!: assms)
-
-
-lemma length_cases3:
- assumes "\<And>LT. P ([],LT)"
- assumes "\<And>l LT. P ([l],LT)"
- assumes "\<And>l ST LT. P (l#ST,LT)"
- shows "P s"
-(*<*)
-proof -
- obtain xs LT where s: "s = (xs,LT)" by (cases s)
- show ?thesis
- proof (cases xs)
- case Nil with assms s show ?thesis by simp
- next
- fix l xs' assume "xs = l#xs'"
- with assms s show ?thesis by simp
- qed
-qed
-(*>*)
-
-lemma length_cases4:
- assumes "\<And>LT. P ([],LT)"
- assumes "\<And>l LT. P ([l],LT)"
- assumes "\<And>l l' LT. P ([l,l'],LT)"
- assumes "\<And>l l' ST LT. P (l#l'#ST,LT)"
- shows "P s"
-(*<*)
-proof -
- obtain xs LT where s: "s = (xs,LT)" by (cases s)
- show ?thesis
- proof (cases xs)
- case Nil with assms s show ?thesis by simp
- next
- fix l xs' assume xs: "xs = l#xs'"
- thus ?thesis
- proof (cases xs')
- case Nil with assms s xs show ?thesis by simp
- next
- fix l' ST assume "xs' = l'#ST"
- with assms s xs show ?thesis by simp
- qed
- qed
-qed
-(*>*)
-
-text \<open>
-\medskip
-simp rules for @{term app}
-\<close>
-lemma appNone[simp]: "app i P mxs T\<^sub>r pc mpc et None = True"
- by (simp add: app_def)
-
-
-lemma appLoad[simp]:
-"app\<^sub>i (Load idx, P, T\<^sub>r, mxs, pc, s) = (\<exists>ST LT. s = (ST,LT) \<and> idx < length LT \<and> LT!idx \<noteq> Err \<and> length ST < mxs)"
- by (cases s, simp)
-
-lemma appStore[simp]:
-"app\<^sub>i (Store idx,P,pc,mxs,T\<^sub>r,s) = (\<exists>ts ST LT. s = (ts#ST,LT) \<and> idx < length LT)"
- by (rule length_cases2, auto)
-
-lemma appPush[simp]:
-"app\<^sub>i (Push v,P,pc,mxs,T\<^sub>r,s) =
- (\<exists>ST LT. s = (ST,LT) \<and> length ST < mxs \<and> typeof v \<noteq> None)"
- by (cases s, simp)
-
-lemma appGetField[simp]:
-"app\<^sub>i (Getfield F C,P,pc,mxs,T\<^sub>r,s) =
- (\<exists> oT vT ST LT. s = (oT#ST, LT) \<and>
- P \<turnstile> C sees F,NonStatic:vT in C \<and> P \<turnstile> oT \<le> (Class C))"
- by (rule length_cases2 [of _ s]) auto
-
-lemma appGetStatic[simp]:
-"app\<^sub>i (Getstatic C F D,P,pc,mxs,T\<^sub>r,s) =
- (\<exists> vT ST LT. s = (ST, LT) \<and> length ST < mxs \<and> P \<turnstile> C sees F,Static:vT in D)"
- by (rule length_cases2 [of _ s]) auto
-
-lemma appPutField[simp]:
-"app\<^sub>i (Putfield F C,P,pc,mxs,T\<^sub>r,s) =
- (\<exists> vT vT' oT ST LT. s = (vT#oT#ST, LT) \<and>
- P \<turnstile> C sees F,NonStatic:vT' in C \<and> P \<turnstile> oT \<le> (Class C) \<and> P \<turnstile> vT \<le> vT')"
- by (rule length_cases4 [of _ s], auto)
-
-lemma appPutstatic[simp]:
-"app\<^sub>i (Putstatic C F D,P,pc,mxs,T\<^sub>r,s) =
- (\<exists> vT vT' ST LT. s = (vT#ST, LT) \<and>
- P \<turnstile> C sees F,Static:vT' in D \<and> P \<turnstile> vT \<le> vT')"
- by (rule length_cases4 [of _ s], auto)
-
-lemma appNew[simp]:
- "app\<^sub>i (New C,P,pc,mxs,T\<^sub>r,s) =
- (\<exists>ST LT. s=(ST,LT) \<and> is_class P C \<and> length ST < mxs)"
- by (cases s, simp)
-
-lemma appCheckcast[simp]:
- "app\<^sub>i (Checkcast C,P,pc,mxs,T\<^sub>r,s) =
- (\<exists>T ST LT. s = (T#ST,LT) \<and> is_class P C \<and> is_refT T)"
- by (cases s, cases "fst s", simp add: app_def) (cases "hd (fst s)", auto)
-
-lemma app\<^sub>iPop[simp]:
-"app\<^sub>i (Pop,P,pc,mxs,T\<^sub>r,s) = (\<exists>ts ST LT. s = (ts#ST,LT))"
- by (rule length_cases2, auto)
-
-lemma appIAdd[simp]:
-"app\<^sub>i (IAdd,P,pc,mxs,T\<^sub>r,s) = (\<exists>ST LT. s = (Integer#Integer#ST,LT))"
-(*<*)
-proof -
- obtain ST LT where [simp]: "s = (ST,LT)" by (cases s)
- have "ST = [] \<or> (\<exists>T. ST = [T]) \<or> (\<exists>T\<^sub>1 T\<^sub>2 ST'. ST = T\<^sub>1#T\<^sub>2#ST')"
- by (cases ST, auto, case_tac list, auto)
- moreover
- { assume "ST = []" hence ?thesis by simp }
- moreover
- { fix T assume "ST = [T]" hence ?thesis by (cases T, auto) }
- moreover
- { fix T\<^sub>1 T\<^sub>2 ST' assume "ST = T\<^sub>1#T\<^sub>2#ST'"
- hence ?thesis by (cases T\<^sub>1, auto)
- }
- ultimately show ?thesis by blast
-qed
-(*>*)
-
-
-lemma appIfFalse [simp]:
-"app\<^sub>i (IfFalse b,P,pc,mxs,T\<^sub>r,s) =
- (\<exists>ST LT. s = (Boolean#ST,LT) \<and> 0 \<le> int pc + b)"
-(*<*)
- (is "?P s")
-proof(rule length_cases2)
- fix LT show "?P ([],LT)" by simp
-next
- fix l ST LT show "?P (l#ST,LT)"
- by (case_tac l) auto
-qed
-(*>*)
-
-lemma appCmpEq[simp]:
-"app\<^sub>i (CmpEq,P,pc,mxs,T\<^sub>r,s) =
- (\<exists>T\<^sub>1 T\<^sub>2 ST LT. s = (T\<^sub>1#T\<^sub>2#ST,LT) \<and> (\<not>is_refT T\<^sub>1 \<and> T\<^sub>2 = T\<^sub>1 \<or> is_refT T\<^sub>1 \<and> is_refT T\<^sub>2))"
- by (rule length_cases4, auto)
-
-lemma appReturn[simp]:
-"app\<^sub>i (Return,P,pc,mxs,T\<^sub>r,s) = (\<exists>T ST LT. s = (T#ST,LT) \<and> P \<turnstile> T \<le> T\<^sub>r)"
- by (rule length_cases2, auto)
-
-lemma appThrow[simp]:
- "app\<^sub>i (Throw,P,pc,mxs,T\<^sub>r,s) = (\<exists>T ST LT. s=(T#ST,LT) \<and> is_refT T)"
- by (rule length_cases2, auto)
-
-lemma effNone:
- "(pc', s') \<in> set (eff i P pc et None) \<Longrightarrow> s' = None"
- by (auto simp add: eff_def xcpt_eff_def norm_eff_def)
-
-
-text \<open> some helpers to make the specification directly executable: \<close>
-lemma relevant_entries_append [simp]:
- "relevant_entries P i pc (xt @ xt') = relevant_entries P i pc xt @ relevant_entries P i pc xt'"
- by (unfold relevant_entries_def) simp
-
-lemma xcpt_app_append [iff]:
- "xcpt_app i P pc mxs (xt@xt') \<tau> = (xcpt_app i P pc mxs xt \<tau> \<and> xcpt_app i P pc mxs xt' \<tau>)"
- by (unfold xcpt_app_def) fastforce
-
-lemma xcpt_eff_append [simp]:
- "xcpt_eff i P pc \<tau> (xt@xt') = xcpt_eff i P pc \<tau> xt @ xcpt_eff i P pc \<tau> xt'"
- by (unfold xcpt_eff_def, cases \<tau>) simp
-
-lemma app_append [simp]:
- "app i P pc T mxs mpc (xt@xt') \<tau> = (app i P pc T mxs mpc xt \<tau> \<and> app i P pc T mxs mpc xt' \<tau>)"
- by (unfold app_def eff_def) auto
-
-end
+(* Title: JinjaDCI/BV/Effect.thy
+ Author: Gerwin Klein, Susannah Mansky
+ Copyright 2000 Technische Universitaet Muenchen, 2019-20 UIUC
+
+ Based on the Jinja theory BV/Effect.thy by Gerwin Klein
+*)
+
+section \<open>Effect of Instructions on the State Type\<close>
+
+theory Effect
+imports JVM_SemiType "../JVM/JVMExceptions"
+begin
+
+\<comment> \<open>FIXME\<close>
+locale prog =
+ fixes P :: "'a prog"
+
+locale jvm_method = prog +
+ fixes mxs :: nat
+ fixes mxl\<^sub>0 :: nat
+ fixes Ts :: "ty list"
+ fixes T\<^sub>r :: ty
+ fixes "is" :: "instr list"
+ fixes xt :: ex_table
+
+ fixes mxl :: nat
+ defines mxl_def: "mxl \<equiv> 1+size Ts+mxl\<^sub>0"
+
+text \<open> Program counter of successor instructions: \<close>
+primrec succs :: "instr \<Rightarrow> ty\<^sub>i \<Rightarrow> pc \<Rightarrow> pc list" where
+ "succs (Load idx) \<tau> pc = [pc+1]"
+| "succs (Store idx) \<tau> pc = [pc+1]"
+| "succs (Push v) \<tau> pc = [pc+1]"
+| "succs (Getfield F C) \<tau> pc = [pc+1]"
+| "succs (Getstatic C F D) \<tau> pc = [pc+1]"
+| "succs (Putfield F C) \<tau> pc = [pc+1]"
+| "succs (Putstatic C F D) \<tau> pc = [pc+1]"
+| "succs (New C) \<tau> pc = [pc+1]"
+| "succs (Checkcast C) \<tau> pc = [pc+1]"
+| "succs Pop \<tau> pc = [pc+1]"
+| "succs IAdd \<tau> pc = [pc+1]"
+| "succs CmpEq \<tau> pc = [pc+1]"
+| succs_IfFalse:
+ "succs (IfFalse b) \<tau> pc = [pc+1, nat (int pc + b)]"
+| succs_Goto:
+ "succs (Goto b) \<tau> pc = [nat (int pc + b)]"
+| succs_Return:
+ "succs Return \<tau> pc = []"
+| succs_Invoke:
+ "succs (Invoke M n) \<tau> pc = (if (fst \<tau>)!n = NT then [] else [pc+1])"
+| succs_Invokestatic:
+ "succs (Invokestatic C M n) \<tau> pc = [pc+1]"
+| succs_Throw:
+ "succs Throw \<tau> pc = []"
+
+text "Effect of instruction on the state type:"
+
+fun the_class:: "ty \<Rightarrow> cname" where
+ "the_class (Class C) = C"
+
+fun eff\<^sub>i :: "instr \<times> 'm prog \<times> ty\<^sub>i \<Rightarrow> ty\<^sub>i" where
+ eff\<^sub>i_Load:
+ "eff\<^sub>i (Load n, P, (ST, LT)) = (ok_val (LT ! n) # ST, LT)"
+| eff\<^sub>i_Store:
+ "eff\<^sub>i (Store n, P, (T#ST, LT)) = (ST, LT[n:= OK T])"
+| eff\<^sub>i_Push:
+ "eff\<^sub>i (Push v, P, (ST, LT)) = (the (typeof v) # ST, LT)"
+| eff\<^sub>i_Getfield:
+ "eff\<^sub>i (Getfield F C, P, (T#ST, LT)) = (snd (snd (field P C F)) # ST, LT)"
+| eff\<^sub>i_Getstatic:
+ "eff\<^sub>i (Getstatic C F D, P, (ST, LT)) = (snd (snd (field P C F)) # ST, LT)"
+| eff\<^sub>i_Putfield:
+ "eff\<^sub>i (Putfield F C, P, (T\<^sub>1#T\<^sub>2#ST, LT)) = (ST,LT)"
+| eff\<^sub>i_Putstatic:
+ "eff\<^sub>i (Putstatic C F D, P, (T#ST, LT)) = (ST,LT)"
+| eff\<^sub>i_New:
+ "eff\<^sub>i (New C, P, (ST,LT)) = (Class C # ST, LT)"
+| eff\<^sub>i_Checkcast:
+ "eff\<^sub>i (Checkcast C, P, (T#ST,LT)) = (Class C # ST,LT)"
+| eff\<^sub>i_Pop:
+ "eff\<^sub>i (Pop, P, (T#ST,LT)) = (ST,LT)"
+| eff\<^sub>i_IAdd:
+ "eff\<^sub>i (IAdd, P,(T\<^sub>1#T\<^sub>2#ST,LT)) = (Integer#ST,LT)"
+| eff\<^sub>i_CmpEq:
+ "eff\<^sub>i (CmpEq, P, (T\<^sub>1#T\<^sub>2#ST,LT)) = (Boolean#ST,LT)"
+| eff\<^sub>i_IfFalse:
+ "eff\<^sub>i (IfFalse b, P, (T\<^sub>1#ST,LT)) = (ST,LT)"
+| eff\<^sub>i_Invoke:
+ "eff\<^sub>i (Invoke M n, P, (ST,LT)) =
+ (let C = the_class (ST!n); (D,b,Ts,T\<^sub>r,m) = method P C M
+ in (T\<^sub>r # drop (n+1) ST, LT))"
+| eff\<^sub>i_Invokestatic:
+ "eff\<^sub>i (Invokestatic C M n, P, (ST,LT)) =
+ (let (D,b,Ts,T\<^sub>r,m) = method P C M
+ in (T\<^sub>r # drop n ST, LT))"
+| eff\<^sub>i_Goto:
+ "eff\<^sub>i (Goto n, P, s) = s"
+
+fun is_relevant_class :: "instr \<Rightarrow> 'm prog \<Rightarrow> cname \<Rightarrow> bool" where
+ rel_Getfield:
+ "is_relevant_class (Getfield F D)
+ = (\<lambda>P C. P \<turnstile> NullPointer \<preceq>\<^sup>* C \<or> P \<turnstile> NoSuchFieldError \<preceq>\<^sup>* C
+ \<or> P \<turnstile> IncompatibleClassChangeError \<preceq>\<^sup>* C)"
+| rel_Getstatic:
+ "is_relevant_class (Getstatic C F D)
+ = (\<lambda>P C. True)"
+| rel_Putfield:
+ "is_relevant_class (Putfield F D)
+ = (\<lambda>P C. P \<turnstile> NullPointer \<preceq>\<^sup>* C \<or> P \<turnstile> NoSuchFieldError \<preceq>\<^sup>* C
+ \<or> P \<turnstile> IncompatibleClassChangeError \<preceq>\<^sup>* C)"
+| rel_Putstatic:
+ "is_relevant_class (Putstatic C F D)
+ = (\<lambda>P C. True)"
+| rel_Checkcast:
+ "is_relevant_class (Checkcast D) = (\<lambda>P C. P \<turnstile> ClassCast \<preceq>\<^sup>* C)"
+| rel_New:
+ "is_relevant_class (New D) = (\<lambda>P C. True)"
+| rel_Throw:
+ "is_relevant_class Throw = (\<lambda>P C. True)"
+| rel_Invoke:
+ "is_relevant_class (Invoke M n) = (\<lambda>P C. True)"
+| rel_Invokestatic:
+ "is_relevant_class (Invokestatic C M n) = (\<lambda>P C. True)"
+| rel_default:
+ "is_relevant_class i = (\<lambda>P C. False)"
+
+definition is_relevant_entry :: "'m prog \<Rightarrow> instr \<Rightarrow> pc \<Rightarrow> ex_entry \<Rightarrow> bool" where
+ "is_relevant_entry P i pc e \<longleftrightarrow> (let (f,t,C,h,d) = e in is_relevant_class i P C \<and> pc \<in> {f..<t})"
+
+definition relevant_entries :: "'m prog \<Rightarrow> instr \<Rightarrow> pc \<Rightarrow> ex_table \<Rightarrow> ex_table" where
+ "relevant_entries P i pc = filter (is_relevant_entry P i pc)"
+
+definition xcpt_eff :: "instr \<Rightarrow> 'm prog \<Rightarrow> pc \<Rightarrow> ty\<^sub>i
+ \<Rightarrow> ex_table \<Rightarrow> (pc \<times> ty\<^sub>i') list" where
+ "xcpt_eff i P pc \<tau> et = (let (ST,LT) = \<tau> in
+ map (\<lambda>(f,t,C,h,d). (h, Some (Class C#drop (size ST - d) ST, LT))) (relevant_entries P i pc et))"
+
+definition norm_eff :: "instr \<Rightarrow> 'm prog \<Rightarrow> nat \<Rightarrow> ty\<^sub>i \<Rightarrow> (pc \<times> ty\<^sub>i') list" where
+ "norm_eff i P pc \<tau> = map (\<lambda>pc'. (pc',Some (eff\<^sub>i (i,P,\<tau>)))) (succs i \<tau> pc)"
+
+definition eff :: "instr \<Rightarrow> 'm prog \<Rightarrow> pc \<Rightarrow> ex_table \<Rightarrow> ty\<^sub>i' \<Rightarrow> (pc \<times> ty\<^sub>i') list" where
+ "eff i P pc et t = (case t of
+ None \<Rightarrow> []
+ | Some \<tau> \<Rightarrow> (norm_eff i P pc \<tau>) @ (xcpt_eff i P pc \<tau> et))"
+
+
+lemma eff_None:
+ "eff i P pc xt None = []"
+by (simp add: eff_def)
+
+lemma eff_Some:
+ "eff i P pc xt (Some \<tau>) = norm_eff i P pc \<tau> @ xcpt_eff i P pc \<tau> xt"
+by (simp add: eff_def)
+
+(* FIXME: getfield, \<exists>T D. P \<turnstile> C sees F:T in D \<and> .. *)
+
+text "Conditions under which eff is applicable:"
+
+fun app\<^sub>i :: "instr \<times> 'm prog \<times> pc \<times> nat \<times> ty \<times> ty\<^sub>i \<Rightarrow> bool" where
+ app\<^sub>i_Load:
+ "app\<^sub>i (Load n, P, pc, mxs, T\<^sub>r, (ST,LT)) =
+ (n < length LT \<and> LT ! n \<noteq> Err \<and> length ST < mxs)"
+| app\<^sub>i_Store:
+ "app\<^sub>i (Store n, P, pc, mxs, T\<^sub>r, (T#ST, LT)) =
+ (n < length LT)"
+| app\<^sub>i_Push:
+ "app\<^sub>i (Push v, P, pc, mxs, T\<^sub>r, (ST,LT)) =
+ (length ST < mxs \<and> typeof v \<noteq> None)"
+| app\<^sub>i_Getfield:
+ "app\<^sub>i (Getfield F C, P, pc, mxs, T\<^sub>r, (T#ST, LT)) =
+ (\<exists>T\<^sub>f. P \<turnstile> C sees F,NonStatic:T\<^sub>f in C \<and> P \<turnstile> T \<le> Class C)"
+| app\<^sub>i_Getstatic:
+ "app\<^sub>i (Getstatic C F D, P, pc, mxs, T\<^sub>r, (ST, LT)) =
+ (length ST < mxs \<and> (\<exists>T\<^sub>f. P \<turnstile> C sees F,Static:T\<^sub>f in D))"
+| app\<^sub>i_Putfield:
+ "app\<^sub>i (Putfield F C, P, pc, mxs, T\<^sub>r, (T\<^sub>1#T\<^sub>2#ST, LT)) =
+ (\<exists>T\<^sub>f. P \<turnstile> C sees F,NonStatic:T\<^sub>f in C \<and> P \<turnstile> T\<^sub>2 \<le> (Class C) \<and> P \<turnstile> T\<^sub>1 \<le> T\<^sub>f)"
+| app\<^sub>i_Putstatic:
+ "app\<^sub>i (Putstatic C F D, P, pc, mxs, T\<^sub>r, (T#ST, LT)) =
+ (\<exists>T\<^sub>f. P \<turnstile> C sees F,Static:T\<^sub>f in D \<and> P \<turnstile> T \<le> T\<^sub>f)"
+| app\<^sub>i_New:
+ "app\<^sub>i (New C, P, pc, mxs, T\<^sub>r, (ST,LT)) =
+ (is_class P C \<and> length ST < mxs)"
+| app\<^sub>i_Checkcast:
+ "app\<^sub>i (Checkcast C, P, pc, mxs, T\<^sub>r, (T#ST,LT)) =
+ (is_class P C \<and> is_refT T)"
+| app\<^sub>i_Pop:
+ "app\<^sub>i (Pop, P, pc, mxs, T\<^sub>r, (T#ST,LT)) =
+ True"
+| app\<^sub>i_IAdd:
+ "app\<^sub>i (IAdd, P, pc, mxs, T\<^sub>r, (T\<^sub>1#T\<^sub>2#ST,LT)) = (T\<^sub>1 = T\<^sub>2 \<and> T\<^sub>1 = Integer)"
+| app\<^sub>i_CmpEq:
+ "app\<^sub>i (CmpEq, P, pc, mxs, T\<^sub>r, (T\<^sub>1#T\<^sub>2#ST,LT)) =
+ (T\<^sub>1 = T\<^sub>2 \<or> is_refT T\<^sub>1 \<and> is_refT T\<^sub>2)"
+| app\<^sub>i_IfFalse:
+ "app\<^sub>i (IfFalse b, P, pc, mxs, T\<^sub>r, (Boolean#ST,LT)) =
+ (0 \<le> int pc + b)"
+| app\<^sub>i_Goto:
+ "app\<^sub>i (Goto b, P, pc, mxs, T\<^sub>r, s) =
+ (0 \<le> int pc + b)"
+| app\<^sub>i_Return:
+ "app\<^sub>i (Return, P, pc, mxs, T\<^sub>r, (T#ST,LT)) =
+ (P \<turnstile> T \<le> T\<^sub>r)"
+| app\<^sub>i_Throw:
+ "app\<^sub>i (Throw, P, pc, mxs, T\<^sub>r, (T#ST,LT)) =
+ is_refT T"
+| app\<^sub>i_Invoke:
+ "app\<^sub>i (Invoke M n, P, pc, mxs, T\<^sub>r, (ST,LT)) =
+ (n < length ST \<and>
+ (ST!n \<noteq> NT \<longrightarrow>
+ (\<exists>C D Ts T m. ST!n = Class C \<and> P \<turnstile> C sees M,NonStatic:Ts \<rightarrow> T = m in D \<and>
+ P \<turnstile> rev (take n ST) [\<le>] Ts)))"
+| app\<^sub>i_Invokestatic:
+ "app\<^sub>i (Invokestatic C M n, P, pc, mxs, T\<^sub>r, (ST,LT)) =
+ (length ST - n < mxs \<and> n \<le> length ST \<and> M \<noteq> clinit \<and>
+ (\<exists>D Ts T m. P \<turnstile> C sees M,Static:Ts \<rightarrow> T = m in D \<and>
+ P \<turnstile> rev (take n ST) [\<le>] Ts))"
+
+| app\<^sub>i_default:
+ "app\<^sub>i (i,P, pc,mxs,T\<^sub>r,s) = False"
+
+
+definition xcpt_app :: "instr \<Rightarrow> 'm prog \<Rightarrow> pc \<Rightarrow> nat \<Rightarrow> ex_table \<Rightarrow> ty\<^sub>i \<Rightarrow> bool" where
+ "xcpt_app i P pc mxs xt \<tau> \<longleftrightarrow> (\<forall>(f,t,C,h,d) \<in> set (relevant_entries P i pc xt). is_class P C \<and> d \<le> size (fst \<tau>) \<and> d < mxs)"
+
+definition app :: "instr \<Rightarrow> 'm prog \<Rightarrow> nat \<Rightarrow> ty \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> ex_table \<Rightarrow> ty\<^sub>i' \<Rightarrow> bool" where
+ "app i P mxs T\<^sub>r pc mpc xt t = (case t of None \<Rightarrow> True | Some \<tau> \<Rightarrow>
+ app\<^sub>i (i,P,pc,mxs,T\<^sub>r,\<tau>) \<and> xcpt_app i P pc mxs xt \<tau> \<and>
+ (\<forall>(pc',\<tau>') \<in> set (eff i P pc xt t). pc' < mpc))"
+
+
+lemma app_Some:
+ "app i P mxs T\<^sub>r pc mpc xt (Some \<tau>) =
+ (app\<^sub>i (i,P,pc,mxs,T\<^sub>r,\<tau>) \<and> xcpt_app i P pc mxs xt \<tau> \<and>
+ (\<forall>(pc',s') \<in> set (eff i P pc xt (Some \<tau>)). pc' < mpc))"
+by (simp add: app_def)
+
+locale eff = jvm_method +
+ fixes eff\<^sub>i and app\<^sub>i and eff and app
+ fixes norm_eff and xcpt_app and xcpt_eff
+
+ fixes mpc
+ defines "mpc \<equiv> size is"
+
+ defines "eff\<^sub>i i \<tau> \<equiv> Effect.eff\<^sub>i (i,P,\<tau>)"
+ notes eff\<^sub>i_simps [simp] = Effect.eff\<^sub>i.simps [where P = P, folded eff\<^sub>i_def]
+
+ defines "app\<^sub>i i pc \<tau> \<equiv> Effect.app\<^sub>i (i, P, pc, mxs, T\<^sub>r, \<tau>)"
+ notes app\<^sub>i_simps [simp] = Effect.app\<^sub>i.simps [where P=P and mxs=mxs and T\<^sub>r=T\<^sub>r, folded app\<^sub>i_def]
+
+ defines "xcpt_eff i pc \<tau> \<equiv> Effect.xcpt_eff i P pc \<tau> xt"
+ notes xcpt_eff = Effect.xcpt_eff_def [of _ P _ _ xt, folded xcpt_eff_def]
+
+ defines "norm_eff i pc \<tau> \<equiv> Effect.norm_eff i P pc \<tau>"
+ notes norm_eff = Effect.norm_eff_def [of _ P, folded norm_eff_def eff\<^sub>i_def]
+
+ defines "eff i pc \<equiv> Effect.eff i P pc xt"
+ notes eff = Effect.eff_def [of _ P _ xt, folded eff_def norm_eff_def xcpt_eff_def]
+
+ defines "xcpt_app i pc \<tau> \<equiv> Effect.xcpt_app i P pc mxs xt \<tau>"
+ notes xcpt_app = Effect.xcpt_app_def [of _ P _ mxs xt, folded xcpt_app_def]
+
+ defines "app i pc \<equiv> Effect.app i P mxs T\<^sub>r pc mpc xt"
+ notes app = Effect.app_def [of _ P mxs T\<^sub>r _ mpc xt, folded app_def xcpt_app_def app\<^sub>i_def eff_def]
+
+
+lemma length_cases2:
+ assumes "\<And>LT. P ([],LT)"
+ assumes "\<And>l ST LT. P (l#ST,LT)"
+ shows "P s"
+ by (cases s, cases "fst s") (auto intro!: assms)
+
+
+lemma length_cases3:
+ assumes "\<And>LT. P ([],LT)"
+ assumes "\<And>l LT. P ([l],LT)"
+ assumes "\<And>l ST LT. P (l#ST,LT)"
+ shows "P s"
+(*<*)
+proof -
+ obtain xs LT where s: "s = (xs,LT)" by (cases s)
+ show ?thesis
+ proof (cases xs)
+ case Nil with assms s show ?thesis by simp
+ next
+ fix l xs' assume "xs = l#xs'"
+ with assms s show ?thesis by simp
+ qed
+qed
+(*>*)
+
+lemma length_cases4:
+ assumes "\<And>LT. P ([],LT)"
+ assumes "\<And>l LT. P ([l],LT)"
+ assumes "\<And>l l' LT. P ([l,l'],LT)"
+ assumes "\<And>l l' ST LT. P (l#l'#ST,LT)"
+ shows "P s"
+(*<*)
+proof -
+ obtain xs LT where s: "s = (xs,LT)" by (cases s)
+ show ?thesis
+ proof (cases xs)
+ case Nil with assms s show ?thesis by simp
+ next
+ fix l xs' assume xs: "xs = l#xs'"
+ thus ?thesis
+ proof (cases xs')
+ case Nil with assms s xs show ?thesis by simp
+ next
+ fix l' ST assume "xs' = l'#ST"
+ with assms s xs show ?thesis by simp
+ qed
+ qed
+qed
+(*>*)
+
+text \<open>
+\medskip
+simp rules for @{term app}
+\<close>
+lemma appNone[simp]: "app i P mxs T\<^sub>r pc mpc et None = True"
+ by (simp add: app_def)
+
+
+lemma appLoad[simp]:
+"app\<^sub>i (Load idx, P, T\<^sub>r, mxs, pc, s) = (\<exists>ST LT. s = (ST,LT) \<and> idx < length LT \<and> LT!idx \<noteq> Err \<and> length ST < mxs)"
+ by (cases s, simp)
+
+lemma appStore[simp]:
+"app\<^sub>i (Store idx,P,pc,mxs,T\<^sub>r,s) = (\<exists>ts ST LT. s = (ts#ST,LT) \<and> idx < length LT)"
+ by (rule length_cases2, auto)
+
+lemma appPush[simp]:
+"app\<^sub>i (Push v,P,pc,mxs,T\<^sub>r,s) =
+ (\<exists>ST LT. s = (ST,LT) \<and> length ST < mxs \<and> typeof v \<noteq> None)"
+ by (cases s, simp)
+
+lemma appGetField[simp]:
+"app\<^sub>i (Getfield F C,P,pc,mxs,T\<^sub>r,s) =
+ (\<exists> oT vT ST LT. s = (oT#ST, LT) \<and>
+ P \<turnstile> C sees F,NonStatic:vT in C \<and> P \<turnstile> oT \<le> (Class C))"
+ by (rule length_cases2 [of _ s]) auto
+
+lemma appGetStatic[simp]:
+"app\<^sub>i (Getstatic C F D,P,pc,mxs,T\<^sub>r,s) =
+ (\<exists> vT ST LT. s = (ST, LT) \<and> length ST < mxs \<and> P \<turnstile> C sees F,Static:vT in D)"
+ by (rule length_cases2 [of _ s]) auto
+
+lemma appPutField[simp]:
+"app\<^sub>i (Putfield F C,P,pc,mxs,T\<^sub>r,s) =
+ (\<exists> vT vT' oT ST LT. s = (vT#oT#ST, LT) \<and>
+ P \<turnstile> C sees F,NonStatic:vT' in C \<and> P \<turnstile> oT \<le> (Class C) \<and> P \<turnstile> vT \<le> vT')"
+ by (rule length_cases4 [of _ s], auto)
+
+lemma appPutstatic[simp]:
+"app\<^sub>i (Putstatic C F D,P,pc,mxs,T\<^sub>r,s) =
+ (\<exists> vT vT' ST LT. s = (vT#ST, LT) \<and>
+ P \<turnstile> C sees F,Static:vT' in D \<and> P \<turnstile> vT \<le> vT')"
+ by (rule length_cases4 [of _ s], auto)
+
+lemma appNew[simp]:
+ "app\<^sub>i (New C,P,pc,mxs,T\<^sub>r,s) =
+ (\<exists>ST LT. s=(ST,LT) \<and> is_class P C \<and> length ST < mxs)"
+ by (cases s, simp)
+
+lemma appCheckcast[simp]:
+ "app\<^sub>i (Checkcast C,P,pc,mxs,T\<^sub>r,s) =
+ (\<exists>T ST LT. s = (T#ST,LT) \<and> is_class P C \<and> is_refT T)"
+ by (cases s, cases "fst s", simp add: app_def) (cases "hd (fst s)", auto)
+
+lemma app\<^sub>iPop[simp]:
+"app\<^sub>i (Pop,P,pc,mxs,T\<^sub>r,s) = (\<exists>ts ST LT. s = (ts#ST,LT))"
+ by (rule length_cases2, auto)
+
+lemma appIAdd[simp]:
+"app\<^sub>i (IAdd,P,pc,mxs,T\<^sub>r,s) = (\<exists>ST LT. s = (Integer#Integer#ST,LT))"
+(*<*)
+proof -
+ obtain ST LT where [simp]: "s = (ST,LT)" by (cases s)
+ have "ST = [] \<or> (\<exists>T. ST = [T]) \<or> (\<exists>T\<^sub>1 T\<^sub>2 ST'. ST = T\<^sub>1#T\<^sub>2#ST')"
+ by (cases ST, auto, case_tac list, auto)
+ moreover
+ { assume "ST = []" hence ?thesis by simp }
+ moreover
+ { fix T assume "ST = [T]" hence ?thesis by (cases T, auto) }
+ moreover
+ { fix T\<^sub>1 T\<^sub>2 ST' assume "ST = T\<^sub>1#T\<^sub>2#ST'"
+ hence ?thesis by (cases T\<^sub>1, auto)
+ }
+ ultimately show ?thesis by blast
+qed
+(*>*)
+
+
+lemma appIfFalse [simp]:
+"app\<^sub>i (IfFalse b,P,pc,mxs,T\<^sub>r,s) =
+ (\<exists>ST LT. s = (Boolean#ST,LT) \<and> 0 \<le> int pc + b)"
+(*<*)
+ (is "?P s")
+proof(rule length_cases2)
+ fix LT show "?P ([],LT)" by simp
+next
+ fix l ST LT show "?P (l#ST,LT)"
+ by (case_tac l) auto
+qed
+(*>*)
+
+lemma appCmpEq[simp]:
+"app\<^sub>i (CmpEq,P,pc,mxs,T\<^sub>r,s) =
+ (\<exists>T\<^sub>1 T\<^sub>2 ST LT. s = (T\<^sub>1#T\<^sub>2#ST,LT) \<and> (\<not>is_refT T\<^sub>1 \<and> T\<^sub>2 = T\<^sub>1 \<or> is_refT T\<^sub>1 \<and> is_refT T\<^sub>2))"
+ by (rule length_cases4, auto)
+
+lemma appReturn[simp]:
+"app\<^sub>i (Return,P,pc,mxs,T\<^sub>r,s) = (\<exists>T ST LT. s = (T#ST,LT) \<and> P \<turnstile> T \<le> T\<^sub>r)"
+ by (rule length_cases2, auto)
+
+lemma appThrow[simp]:
+ "app\<^sub>i (Throw,P,pc,mxs,T\<^sub>r,s) = (\<exists>T ST LT. s=(T#ST,LT) \<and> is_refT T)"
+ by (rule length_cases2, auto)
+
+lemma effNone:
+ "(pc', s') \<in> set (eff i P pc et None) \<Longrightarrow> s' = None"
+ by (auto simp add: eff_def xcpt_eff_def norm_eff_def)
+
+
+text \<open> some helpers to make the specification directly executable: \<close>
+lemma relevant_entries_append [simp]:
+ "relevant_entries P i pc (xt @ xt') = relevant_entries P i pc xt @ relevant_entries P i pc xt'"
+ by (unfold relevant_entries_def) simp
+
+lemma xcpt_app_append [iff]:
+ "xcpt_app i P pc mxs (xt@xt') \<tau> = (xcpt_app i P pc mxs xt \<tau> \<and> xcpt_app i P pc mxs xt' \<tau>)"
+ by (unfold xcpt_app_def) fastforce
+
+lemma xcpt_eff_append [simp]:
+ "xcpt_eff i P pc \<tau> (xt@xt') = xcpt_eff i P pc \<tau> xt @ xcpt_eff i P pc \<tau> xt'"
+ by (unfold xcpt_eff_def, cases \<tau>) simp
+
+lemma app_append [simp]:
+ "app i P pc T mxs mpc (xt@xt') \<tau> = (app i P pc T mxs mpc xt \<tau> \<and> app i P pc T mxs mpc xt' \<tau>)"
+ by (unfold app_def eff_def) auto
+
+end
diff --git a/thys/JinjaDCI/BV/EffectMono.thy b/thys/JinjaDCI/BV/EffectMono.thy
--- a/thys/JinjaDCI/BV/EffectMono.thy
+++ b/thys/JinjaDCI/BV/EffectMono.thy
@@ -1,235 +1,235 @@
-(* Title: JinjaDCI/BV/EffMono.thy
-
- Author: Gerwin Klein, Susannah Mansky
- Copyright 2000 Technische Universitaet Muenchen, 2019-20 UIUC
-
- Based on the Jinja theory BV/EffectMono.thy by Gerwin Klein
-*)
-
-section \<open> Monotonicity of eff and app \<close>
-
-theory EffectMono imports Effect begin
-
-declare not_Err_eq [iff]
-
-lemma app\<^sub>i_mono:
- assumes wf: "wf_prog p P"
- assumes less: "P \<turnstile> \<tau> \<le>\<^sub>i \<tau>'"
- shows "app\<^sub>i (i,P,mxs,mpc,rT,\<tau>') \<Longrightarrow> app\<^sub>i (i,P,mxs,mpc,rT,\<tau>)"
-(*<*)
-proof -
- assume app: "app\<^sub>i (i,P,mxs,mpc,rT,\<tau>')"
-
- obtain ST LT ST' LT' where
- [simp]: "\<tau> = (ST,LT)" and
- [simp]: "\<tau>' = (ST',LT')"
- by (cases \<tau>, cases \<tau>')
-
- from less have [simp]: "size ST = size ST'" and [simp]: "size LT = size LT'"
- by (auto dest: list_all2_lengthD)
-
- note [iff] = list_all2_Cons2 widen_Class
- note [simp] = fun_of_def
-
- from app less show "app\<^sub>i (i,P,mxs,mpc,rT,\<tau>)"
- proof (cases i)
- case Load
- with app less show ?thesis by (auto dest!: list_all2_nthD)
- next
- case (Invoke M n)
- with app have n: "n < size ST'" by simp
-
- { assume "ST!n = NT" hence ?thesis using n app Invoke by simp }
- moreover {
- assume "ST'!n = NT"
- moreover with n less have "ST!n = NT"
- by (auto dest: list_all2_nthD)
- ultimately have ?thesis using n app Invoke by simp
- }
- moreover {
- assume ST: "ST!n \<noteq> NT" and ST': "ST'!n \<noteq> NT"
-
- from ST' app Invoke obtain D Ts T m C' where
- D: "ST' ! n = Class D" and
- Ts: "P \<turnstile> rev (take n ST') [\<le>] Ts" and
- D_M: "P \<turnstile> D sees M,NonStatic: Ts\<rightarrow>T = m in C'"
- by auto
-
- from n D less have "P \<turnstile> ST!n \<le> ST'!n"
- by (fastforce dest: list_all2_nthD2)
- with D ST obtain D' where
- D': "ST!n = Class D'" and DsubC: "P \<turnstile> D' \<preceq>\<^sup>* D" by auto
-
- from wf D_M DsubC obtain Ts' T' m' C'' where
- D'_M: "P \<turnstile> D' sees M,NonStatic: Ts'\<rightarrow>T' = m' in C''" and
- Ts': "P \<turnstile> Ts [\<le>] Ts'"
- by (blast dest: sees_method_mono)
-
- from less have "P \<turnstile> rev (take n ST) [\<le>] rev (take n ST')" by simp
- also note Ts also note Ts'
- finally have "P \<turnstile> rev (take n ST) [\<le>] Ts'" .
- with D'_M D' app less Invoke have ?thesis by fastforce
- }
- ultimately show ?thesis by blast
- next
- case (Invokestatic D M n)
- moreover {
- from app Invokestatic obtain Ts T m C' where
- Ts: "P \<turnstile> rev (take n ST') [\<le>] Ts" and
- D_M: "P \<turnstile> D sees M,Static: Ts\<rightarrow>T = m in C'"
- by auto
- from wf D_M obtain Ts' T' m' C'' where
- D'_M: "P \<turnstile> D sees M,Static: Ts'\<rightarrow>T' = m' in C''" and
- Ts': "P \<turnstile> Ts [\<le>] Ts'"
- by (blast dest: sees_method_mono)
- from less have "P \<turnstile> rev (take n ST) [\<le>] rev (take n ST')" by simp
- also note Ts also note Ts'
- finally have "P \<turnstile> rev (take n ST) [\<le>] Ts'" .
- with D'_M app less Invokestatic have ?thesis by fastforce
- }
- ultimately show ?thesis by blast
- next
- case Getfield
- with app less show ?thesis by (fastforce intro: rtrancl_trans)
- next
- case (Putfield F C)
- with app less show ?thesis by (fastforce intro: widen_trans rtrancl_trans)
- next
- case (Putstatic C F D)
- with app less show ?thesis by (fastforce intro: widen_trans rtrancl_trans)
- next
- case Return
- with app less show ?thesis by (fastforce intro: widen_trans)
- qed (auto elim!: refTE not_refTE)
-qed
-(*>*)
-
-lemma succs_mono:
- assumes wf: "wf_prog p P" and app\<^sub>i: "app\<^sub>i (i,P,mxs,mpc,rT,\<tau>')"
- shows "P \<turnstile> \<tau> \<le>\<^sub>i \<tau>' \<Longrightarrow> set (succs i \<tau> pc) \<subseteq> set (succs i \<tau>' pc)"
-(*<*)
-proof (cases i)
- case (Invoke M n)
- obtain ST LT ST' LT' where
- [simp]: "\<tau> = (ST,LT)" and [simp]: "\<tau>' = (ST',LT')" by (cases \<tau>, cases \<tau>')
- assume "P \<turnstile> \<tau> \<le>\<^sub>i \<tau>'"
- moreover
- with app\<^sub>i Invoke have "n < size ST" by (auto dest: list_all2_lengthD)
- ultimately
- have "P \<turnstile> ST!n \<le> ST'!n" by (auto simp add: fun_of_def dest: list_all2_nthD)
- with Invoke show ?thesis by auto
-qed auto
-(*>*)
-
-
-lemma app_mono:
- assumes wf: "wf_prog p P"
- assumes less': "P \<turnstile> \<tau> \<le>' \<tau>'"
- shows "app i P m rT pc mpc xt \<tau>' \<Longrightarrow> app i P m rT pc mpc xt \<tau>"
-(*<*)
-proof (cases \<tau>)
- case None thus ?thesis by simp
-next
- case (Some \<tau>\<^sub>1)
- moreover
- with less' obtain \<tau>\<^sub>2 where \<tau>\<^sub>2: "\<tau>' = Some \<tau>\<^sub>2" by (cases \<tau>') auto
- ultimately have less: "P \<turnstile> \<tau>\<^sub>1 \<le>\<^sub>i \<tau>\<^sub>2" using less' by simp
-
- assume "app i P m rT pc mpc xt \<tau>'"
- with Some \<tau>\<^sub>2 obtain
- app\<^sub>i: "app\<^sub>i (i, P, pc, m, rT, \<tau>\<^sub>2)" and
- xcpt: "xcpt_app i P pc m xt \<tau>\<^sub>2" and
- succs: "\<forall>(pc',s')\<in>set (eff i P pc xt (Some \<tau>\<^sub>2)). pc' < mpc"
- by (auto simp add: app_def)
-
- from wf less app\<^sub>i have "app\<^sub>i (i, P, pc, m, rT, \<tau>\<^sub>1)" by (rule app\<^sub>i_mono)
- moreover
- from less have "size (fst \<tau>\<^sub>1) = size (fst \<tau>\<^sub>2)"
- by (cases \<tau>\<^sub>1, cases \<tau>\<^sub>2) (auto dest: list_all2_lengthD)
- with xcpt have "xcpt_app i P pc m xt \<tau>\<^sub>1" by (simp add: xcpt_app_def)
- moreover
- from wf app\<^sub>i less have "\<forall>pc. set (succs i \<tau>\<^sub>1 pc) \<subseteq> set (succs i \<tau>\<^sub>2 pc)"
- by (blast dest: succs_mono)
- with succs
- have "\<forall>(pc',s')\<in>set (eff i P pc xt (Some \<tau>\<^sub>1)). pc' < mpc"
- by (cases \<tau>\<^sub>1, cases \<tau>\<^sub>2)
- (auto simp add: eff_def norm_eff_def xcpt_eff_def dest: bspec)
- ultimately
- show ?thesis using Some by (simp add: app_def)
-qed
-(*>*)
-
-
-lemma eff\<^sub>i_mono:
- assumes wf: "wf_prog p P"
- assumes less: "P \<turnstile> \<tau> \<le>\<^sub>i \<tau>'"
- assumes app\<^sub>i: "app i P m rT pc mpc xt (Some \<tau>')"
- assumes succs: "succs i \<tau> pc \<noteq> []" "succs i \<tau>' pc \<noteq> []"
- shows "P \<turnstile> eff\<^sub>i (i,P,\<tau>) \<le>\<^sub>i eff\<^sub>i (i,P,\<tau>')"
-(*<*)
-proof -
- obtain ST LT ST' LT' where
- [simp]: "\<tau> = (ST,LT)" and
- [simp]: "\<tau>' = (ST',LT')"
- by (cases \<tau>, cases \<tau>')
-
- note [simp] = eff_def app_def fun_of_def
-
- from less have "P \<turnstile> (Some \<tau>) \<le>' (Some \<tau>')" by simp
- from wf this app\<^sub>i
- have app: "app i P m rT pc mpc xt (Some \<tau>)" by (rule app_mono)
-
- from less app app\<^sub>i show ?thesis
- proof (cases i)
- case Throw with succs have False by simp
- thus ?thesis ..
- next
- case Return with succs have False by simp
- thus ?thesis ..
- next
- case (Load i)
- from Load app obtain y where
- y: "i < size LT" "LT!i = OK y" by clarsimp
- from Load app\<^sub>i obtain y' where
- y': "i < size LT'" "LT'!i = OK y'" by clarsimp
-
- from less have "P \<turnstile> LT [\<le>\<^sub>\<top>] LT'" by simp
- with y y' have "P \<turnstile> y \<le> y'" by (auto dest: list_all2_nthD)
- with Load less y y' app app\<^sub>i
- show ?thesis by auto
- next
- case Store with less app app\<^sub>i
- show ?thesis by (auto simp add: list_all2_update_cong)
- next
- case (Invoke M n)
- with app\<^sub>i have n: "n < size ST'" by simp
- from less have [simp]: "size ST = size ST'"
- by (auto dest: list_all2_lengthD)
-
- from Invoke succs have ST: "ST!n \<noteq> NT" and ST': "ST'!n \<noteq> NT"
- by (auto split: if_split_asm)
-
- from ST' app\<^sub>i Invoke obtain D Ts T m C' where
- D: "ST' ! n = Class D" and
- D_M: "P \<turnstile> D sees M,NonStatic: Ts\<rightarrow>T = m in C'"
- by auto
-
- from n D less have "P \<turnstile> ST!n \<le> ST'!n"
- by (fastforce dest: list_all2_nthD2)
- with D ST obtain D' where
- D': "ST ! n = Class D'" and DsubC: "P \<turnstile> D' \<preceq>\<^sup>* D"
- by (auto simp: widen_Class)
-
- from wf D_M DsubC obtain Ts' T' m' C'' where
- D'_M: "P \<turnstile> D' sees M,NonStatic: Ts'\<rightarrow>T' = m' in C''" and
- Ts': "P \<turnstile> T' \<le> T"
- by (blast dest: sees_method_mono)
-
- with Invoke n D D' D_M less
- show ?thesis by (auto intro: list_all2_dropI)
- qed auto
-qed
-(*>*)
-
-end
-
+(* Title: JinjaDCI/BV/EffMono.thy
+
+ Author: Gerwin Klein, Susannah Mansky
+ Copyright 2000 Technische Universitaet Muenchen, 2019-20 UIUC
+
+ Based on the Jinja theory BV/EffectMono.thy by Gerwin Klein
+*)
+
+section \<open> Monotonicity of eff and app \<close>
+
+theory EffectMono imports Effect begin
+
+declare not_Err_eq [iff]
+
+lemma app\<^sub>i_mono:
+ assumes wf: "wf_prog p P"
+ assumes less: "P \<turnstile> \<tau> \<le>\<^sub>i \<tau>'"
+ shows "app\<^sub>i (i,P,mxs,mpc,rT,\<tau>') \<Longrightarrow> app\<^sub>i (i,P,mxs,mpc,rT,\<tau>)"
+(*<*)
+proof -
+ assume app: "app\<^sub>i (i,P,mxs,mpc,rT,\<tau>')"
+
+ obtain ST LT ST' LT' where
+ [simp]: "\<tau> = (ST,LT)" and
+ [simp]: "\<tau>' = (ST',LT')"
+ by (cases \<tau>, cases \<tau>')
+
+ from less have [simp]: "size ST = size ST'" and [simp]: "size LT = size LT'"
+ by (auto dest: list_all2_lengthD)
+
+ note [iff] = list_all2_Cons2 widen_Class
+ note [simp] = fun_of_def
+
+ from app less show "app\<^sub>i (i,P,mxs,mpc,rT,\<tau>)"
+ proof (cases i)
+ case Load
+ with app less show ?thesis by (auto dest!: list_all2_nthD)
+ next
+ case (Invoke M n)
+ with app have n: "n < size ST'" by simp
+
+ { assume "ST!n = NT" hence ?thesis using n app Invoke by simp }
+ moreover {
+ assume "ST'!n = NT"
+ moreover with n less have "ST!n = NT"
+ by (auto dest: list_all2_nthD)
+ ultimately have ?thesis using n app Invoke by simp
+ }
+ moreover {
+ assume ST: "ST!n \<noteq> NT" and ST': "ST'!n \<noteq> NT"
+
+ from ST' app Invoke obtain D Ts T m C' where
+ D: "ST' ! n = Class D" and
+ Ts: "P \<turnstile> rev (take n ST') [\<le>] Ts" and
+ D_M: "P \<turnstile> D sees M,NonStatic: Ts\<rightarrow>T = m in C'"
+ by auto
+
+ from n D less have "P \<turnstile> ST!n \<le> ST'!n"
+ by (fastforce dest: list_all2_nthD2)
+ with D ST obtain D' where
+ D': "ST!n = Class D'" and DsubC: "P \<turnstile> D' \<preceq>\<^sup>* D" by auto
+
+ from wf D_M DsubC obtain Ts' T' m' C'' where
+ D'_M: "P \<turnstile> D' sees M,NonStatic: Ts'\<rightarrow>T' = m' in C''" and
+ Ts': "P \<turnstile> Ts [\<le>] Ts'"
+ by (blast dest: sees_method_mono)
+
+ from less have "P \<turnstile> rev (take n ST) [\<le>] rev (take n ST')" by simp
+ also note Ts also note Ts'
+ finally have "P \<turnstile> rev (take n ST) [\<le>] Ts'" .
+ with D'_M D' app less Invoke have ?thesis by fastforce
+ }
+ ultimately show ?thesis by blast
+ next
+ case (Invokestatic D M n)
+ moreover {
+ from app Invokestatic obtain Ts T m C' where
+ Ts: "P \<turnstile> rev (take n ST') [\<le>] Ts" and
+ D_M: "P \<turnstile> D sees M,Static: Ts\<rightarrow>T = m in C'"
+ by auto
+ from wf D_M obtain Ts' T' m' C'' where
+ D'_M: "P \<turnstile> D sees M,Static: Ts'\<rightarrow>T' = m' in C''" and
+ Ts': "P \<turnstile> Ts [\<le>] Ts'"
+ by (blast dest: sees_method_mono)
+ from less have "P \<turnstile> rev (take n ST) [\<le>] rev (take n ST')" by simp
+ also note Ts also note Ts'
+ finally have "P \<turnstile> rev (take n ST) [\<le>] Ts'" .
+ with D'_M app less Invokestatic have ?thesis by fastforce
+ }
+ ultimately show ?thesis by blast
+ next
+ case Getfield
+ with app less show ?thesis by (fastforce intro: rtrancl_trans)
+ next
+ case (Putfield F C)
+ with app less show ?thesis by (fastforce intro: widen_trans rtrancl_trans)
+ next
+ case (Putstatic C F D)
+ with app less show ?thesis by (fastforce intro: widen_trans rtrancl_trans)
+ next
+ case Return
+ with app less show ?thesis by (fastforce intro: widen_trans)
+ qed (auto elim!: refTE not_refTE)
+qed
+(*>*)
+
+lemma succs_mono:
+ assumes wf: "wf_prog p P" and app\<^sub>i: "app\<^sub>i (i,P,mxs,mpc,rT,\<tau>')"
+ shows "P \<turnstile> \<tau> \<le>\<^sub>i \<tau>' \<Longrightarrow> set (succs i \<tau> pc) \<subseteq> set (succs i \<tau>' pc)"
+(*<*)
+proof (cases i)
+ case (Invoke M n)
+ obtain ST LT ST' LT' where
+ [simp]: "\<tau> = (ST,LT)" and [simp]: "\<tau>' = (ST',LT')" by (cases \<tau>, cases \<tau>')
+ assume "P \<turnstile> \<tau> \<le>\<^sub>i \<tau>'"
+ moreover
+ with app\<^sub>i Invoke have "n < size ST" by (auto dest: list_all2_lengthD)
+ ultimately
+ have "P \<turnstile> ST!n \<le> ST'!n" by (auto simp add: fun_of_def dest: list_all2_nthD)
+ with Invoke show ?thesis by auto
+qed auto
+(*>*)
+
+
+lemma app_mono:
+ assumes wf: "wf_prog p P"
+ assumes less': "P \<turnstile> \<tau> \<le>' \<tau>'"
+ shows "app i P m rT pc mpc xt \<tau>' \<Longrightarrow> app i P m rT pc mpc xt \<tau>"
+(*<*)
+proof (cases \<tau>)
+ case None thus ?thesis by simp
+next
+ case (Some \<tau>\<^sub>1)
+ moreover
+ with less' obtain \<tau>\<^sub>2 where \<tau>\<^sub>2: "\<tau>' = Some \<tau>\<^sub>2" by (cases \<tau>') auto
+ ultimately have less: "P \<turnstile> \<tau>\<^sub>1 \<le>\<^sub>i \<tau>\<^sub>2" using less' by simp
+
+ assume "app i P m rT pc mpc xt \<tau>'"
+ with Some \<tau>\<^sub>2 obtain
+ app\<^sub>i: "app\<^sub>i (i, P, pc, m, rT, \<tau>\<^sub>2)" and
+ xcpt: "xcpt_app i P pc m xt \<tau>\<^sub>2" and
+ succs: "\<forall>(pc',s')\<in>set (eff i P pc xt (Some \<tau>\<^sub>2)). pc' < mpc"
+ by (auto simp add: app_def)
+
+ from wf less app\<^sub>i have "app\<^sub>i (i, P, pc, m, rT, \<tau>\<^sub>1)" by (rule app\<^sub>i_mono)
+ moreover
+ from less have "size (fst \<tau>\<^sub>1) = size (fst \<tau>\<^sub>2)"
+ by (cases \<tau>\<^sub>1, cases \<tau>\<^sub>2) (auto dest: list_all2_lengthD)
+ with xcpt have "xcpt_app i P pc m xt \<tau>\<^sub>1" by (simp add: xcpt_app_def)
+ moreover
+ from wf app\<^sub>i less have "\<forall>pc. set (succs i \<tau>\<^sub>1 pc) \<subseteq> set (succs i \<tau>\<^sub>2 pc)"
+ by (blast dest: succs_mono)
+ with succs
+ have "\<forall>(pc',s')\<in>set (eff i P pc xt (Some \<tau>\<^sub>1)). pc' < mpc"
+ by (cases \<tau>\<^sub>1, cases \<tau>\<^sub>2)
+ (auto simp add: eff_def norm_eff_def xcpt_eff_def dest: bspec)
+ ultimately
+ show ?thesis using Some by (simp add: app_def)
+qed
+(*>*)
+
+
+lemma eff\<^sub>i_mono:
+ assumes wf: "wf_prog p P"
+ assumes less: "P \<turnstile> \<tau> \<le>\<^sub>i \<tau>'"
+ assumes app\<^sub>i: "app i P m rT pc mpc xt (Some \<tau>')"
+ assumes succs: "succs i \<tau> pc \<noteq> []" "succs i \<tau>' pc \<noteq> []"
+ shows "P \<turnstile> eff\<^sub>i (i,P,\<tau>) \<le>\<^sub>i eff\<^sub>i (i,P,\<tau>')"
+(*<*)
+proof -
+ obtain ST LT ST' LT' where
+ [simp]: "\<tau> = (ST,LT)" and
+ [simp]: "\<tau>' = (ST',LT')"
+ by (cases \<tau>, cases \<tau>')
+
+ note [simp] = eff_def app_def fun_of_def
+
+ from less have "P \<turnstile> (Some \<tau>) \<le>' (Some \<tau>')" by simp
+ from wf this app\<^sub>i
+ have app: "app i P m rT pc mpc xt (Some \<tau>)" by (rule app_mono)
+
+ from less app app\<^sub>i show ?thesis
+ proof (cases i)
+ case Throw with succs have False by simp
+ thus ?thesis ..
+ next
+ case Return with succs have False by simp
+ thus ?thesis ..
+ next
+ case (Load i)
+ from Load app obtain y where
+ y: "i < size LT" "LT!i = OK y" by clarsimp
+ from Load app\<^sub>i obtain y' where
+ y': "i < size LT'" "LT'!i = OK y'" by clarsimp
+
+ from less have "P \<turnstile> LT [\<le>\<^sub>\<top>] LT'" by simp
+ with y y' have "P \<turnstile> y \<le> y'" by (auto dest: list_all2_nthD)
+ with Load less y y' app app\<^sub>i
+ show ?thesis by auto
+ next
+ case Store with less app app\<^sub>i
+ show ?thesis by (auto simp add: list_all2_update_cong)
+ next
+ case (Invoke M n)
+ with app\<^sub>i have n: "n < size ST'" by simp
+ from less have [simp]: "size ST = size ST'"
+ by (auto dest: list_all2_lengthD)
+
+ from Invoke succs have ST: "ST!n \<noteq> NT" and ST': "ST'!n \<noteq> NT"
+ by (auto split: if_split_asm)
+
+ from ST' app\<^sub>i Invoke obtain D Ts T m C' where
+ D: "ST' ! n = Class D" and
+ D_M: "P \<turnstile> D sees M,NonStatic: Ts\<rightarrow>T = m in C'"
+ by auto
+
+ from n D less have "P \<turnstile> ST!n \<le> ST'!n"
+ by (fastforce dest: list_all2_nthD2)
+ with D ST obtain D' where
+ D': "ST ! n = Class D'" and DsubC: "P \<turnstile> D' \<preceq>\<^sup>* D"
+ by (auto simp: widen_Class)
+
+ from wf D_M DsubC obtain Ts' T' m' C'' where
+ D'_M: "P \<turnstile> D' sees M,NonStatic: Ts'\<rightarrow>T' = m' in C''" and
+ Ts': "P \<turnstile> T' \<le> T"
+ by (blast dest: sees_method_mono)
+
+ with Invoke n D D' D_M less
+ show ?thesis by (auto intro: list_all2_dropI)
+ qed auto
+qed
+(*>*)
+
+end
+
diff --git a/thys/JinjaDCI/BV/JVM_SemiType.thy b/thys/JinjaDCI/BV/JVM_SemiType.thy
--- a/thys/JinjaDCI/BV/JVM_SemiType.thy
+++ b/thys/JinjaDCI/BV/JVM_SemiType.thy
@@ -1,388 +1,388 @@
-(* Title: HOL/MicroJava/BV/JVM.thy
-
- Author: Gerwin Klein
- Copyright 2000 TUM
-
-*)
-
-section \<open> The JVM Type System as Semilattice \<close>
-
-theory JVM_SemiType imports SemiType begin
-
-type_synonym ty\<^sub>l = "ty err list"
-type_synonym ty\<^sub>s = "ty list"
-type_synonym ty\<^sub>i = "ty\<^sub>s \<times> ty\<^sub>l"
-type_synonym ty\<^sub>i' = "ty\<^sub>i option"
-type_synonym ty\<^sub>m = "ty\<^sub>i' list"
-type_synonym ty\<^sub>P = "mname \<Rightarrow> cname \<Rightarrow> ty\<^sub>m"
-
-
-definition stk_esl :: "'c prog \<Rightarrow> nat \<Rightarrow> ty\<^sub>s esl"
-where
- "stk_esl P mxs \<equiv> upto_esl mxs (SemiType.esl P)"
-
-definition loc_sl :: "'c prog \<Rightarrow> nat \<Rightarrow> ty\<^sub>l sl"
-where
- "loc_sl P mxl \<equiv> Listn.sl mxl (Err.sl (SemiType.esl P))"
-
-definition sl :: "'c prog \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> ty\<^sub>i' err sl"
-where
- "sl P mxs mxl \<equiv>
- Err.sl(Opt.esl(Product.esl (stk_esl P mxs) (Err.esl(loc_sl P mxl))))"
-
-
-definition states :: "'c prog \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> ty\<^sub>i' err set"
-where "states P mxs mxl \<equiv> fst(sl P mxs mxl)"
-
-definition le :: "'c prog \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> ty\<^sub>i' err ord"
-where
- "le P mxs mxl \<equiv> fst(snd(sl P mxs mxl))"
-
-definition sup :: "'c prog \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> ty\<^sub>i' err binop"
-where
- "sup P mxs mxl \<equiv> snd(snd(sl P mxs mxl))"
-
-
-definition sup_ty_opt :: "['c prog,ty err,ty err] \<Rightarrow> bool"
- ("_ \<turnstile> _ \<le>\<^sub>\<top> _" [71,71,71] 70)
-where
- "sup_ty_opt P \<equiv> Err.le (subtype P)"
-
-definition sup_state :: "['c prog,ty\<^sub>i,ty\<^sub>i] \<Rightarrow> bool"
- ("_ \<turnstile> _ \<le>\<^sub>i _" [71,71,71] 70)
-where
- "sup_state P \<equiv> Product.le (Listn.le (subtype P)) (Listn.le (sup_ty_opt P))"
-
-definition sup_state_opt :: "['c prog,ty\<^sub>i',ty\<^sub>i'] \<Rightarrow> bool"
- ("_ \<turnstile> _ \<le>'' _" [71,71,71] 70)
-where
- "sup_state_opt P \<equiv> Opt.le (sup_state P)"
-
-abbreviation
- sup_loc :: "['c prog,ty\<^sub>l,ty\<^sub>l] \<Rightarrow> bool" ("_ \<turnstile> _ [\<le>\<^sub>\<top>] _" [71,71,71] 70)
- where "P \<turnstile> LT [\<le>\<^sub>\<top>] LT' \<equiv> list_all2 (sup_ty_opt P) LT LT'"
-
-notation (ASCII)
- sup_ty_opt ("_ |- _ <=T _" [71,71,71] 70) and
- sup_state ("_ |- _ <=i _" [71,71,71] 70) and
- sup_state_opt ("_ |- _ <=' _" [71,71,71] 70) and
- sup_loc ("_ |- _ [<=T] _" [71,71,71] 70)
-
-
-subsection "Unfolding"
-
-lemma JVM_states_unfold:
- "states P mxs mxl \<equiv> err(opt((Union {list n (types P) |n. n <= mxs}) \<times>
- list mxl (err(types P))))"
-(*<*)
- apply (unfold states_def sl_def Opt.esl_def Err.sl_def
- stk_esl_def loc_sl_def Product.esl_def
- Listn.sl_def upto_esl_def SemiType.esl_def Err.esl_def)
- apply simp
- done
-(*>*)
-
-lemma JVM_le_unfold:
- "le P m n \<equiv>
- Err.le(Opt.le(Product.le(Listn.le(subtype P))(Listn.le(Err.le(subtype P)))))"
-(*<*)
- apply (unfold le_def sl_def Opt.esl_def Err.sl_def
- stk_esl_def loc_sl_def Product.esl_def
- Listn.sl_def upto_esl_def SemiType.esl_def Err.esl_def)
- apply simp
- done
-(*>*)
-
-lemma sl_def2:
- "JVM_SemiType.sl P mxs mxl \<equiv>
- (states P mxs mxl, JVM_SemiType.le P mxs mxl, JVM_SemiType.sup P mxs mxl)"
-(*<*) by (unfold JVM_SemiType.sup_def states_def JVM_SemiType.le_def) simp (*>*)
-
-
-lemma JVM_le_conv:
- "le P m n (OK t1) (OK t2) = P \<turnstile> t1 \<le>' t2"
-(*<*) by (simp add: JVM_le_unfold Err.le_def lesub_def sup_state_opt_def
- sup_state_def sup_ty_opt_def) (*>*)
-
-lemma JVM_le_Err_conv:
- "le P m n = Err.le (sup_state_opt P)"
-(*<*) by (unfold sup_state_opt_def sup_state_def
- sup_ty_opt_def JVM_le_unfold) simp (*>*)
-
-lemma err_le_unfold [iff]:
- "Err.le r (OK a) (OK b) = r a b"
-(*<*) by (simp add: Err.le_def lesub_def) (*>*)
-
-
-subsection \<open> Semilattice \<close>
-lemma order_sup_state_opt' [intro, simp]:
- "wf_prog wf_mb P \<Longrightarrow>
- order (sup_state_opt P) (opt ((\<Union> {list n (types P) |n. n \<le> mxs} ) \<times> list (Suc (length Ts + mxl\<^sub>0)) (err (types P))))"
-(*<*)
- apply (unfold sup_state_opt_def sup_state_def sup_ty_opt_def )
- apply (blast intro:order_le_prodI) \<comment>\<open> use Listn.thy.order_listI2 \<close>
- done
-(*<*)
-lemma order_sup_state_opt'' [intro, simp]:
- "wf_prog wf_mb P \<Longrightarrow>
- order (sup_state_opt P) (opt ((\<Union> {list n (types P) |n. n \<le> mxs} ) \<times> list ((length Ts + mxl\<^sub>0)) (err (types P))))"
-(*<*)
- apply (unfold sup_state_opt_def sup_state_def sup_ty_opt_def )
- apply (blast intro:order_le_prodI) \<comment>\<open> use Listn.thy.order_listI2 \<close>
- done
-(*<*)
-(*
-lemma order_sup_state_opt [intro, simp]:
- "wf_prog wf_mb P \<Longrightarrow> order (sup_state_opt P)"
-(*<*) by (unfold sup_state_opt_def sup_state_def sup_ty_opt_def) blast (*>*)
-*)
-
-lemma semilat_JVM [intro?]:
- "wf_prog wf_mb P \<Longrightarrow> semilat (JVM_SemiType.sl P mxs mxl)"
-(*<*)
- apply (unfold JVM_SemiType.sl_def stk_esl_def loc_sl_def)
- apply (blast intro: err_semilat_Product_esl err_semilat_upto_esl
- Listn_sl err_semilat_JType_esl)
- done
-(*>*)
-
-subsection \<open> Widening with @{text "\<top>"} \<close>
-
-lemma subtype_refl[iff]: "subtype P t t" (*<*) by (simp add: fun_of_def) (*>*)
-
-lemma sup_ty_opt_refl [iff]: "P \<turnstile> T \<le>\<^sub>\<top> T"
-(*<*)
- apply (unfold sup_ty_opt_def)
- apply (fold lesub_def)
- apply (rule le_err_refl)
- apply (simp add: lesub_def)
- done
-(*>*)
-
-lemma Err_any_conv [iff]: "P \<turnstile> Err \<le>\<^sub>\<top> T = (T = Err)"
-(*<*) by (unfold sup_ty_opt_def) (rule Err_le_conv [simplified lesub_def]) (*>*)
-
-lemma any_Err [iff]: "P \<turnstile> T \<le>\<^sub>\<top> Err"
-(*<*) by (unfold sup_ty_opt_def) (rule le_Err [simplified lesub_def]) (*>*)
-
-lemma OK_OK_conv [iff]:
- "P \<turnstile> OK T \<le>\<^sub>\<top> OK T' = P \<turnstile> T \<le> T'"
-(*<*) by (simp add: sup_ty_opt_def fun_of_def) (*>*)
-
-lemma any_OK_conv [iff]:
- "P \<turnstile> X \<le>\<^sub>\<top> OK T' = (\<exists>T. X = OK T \<and> P \<turnstile> T \<le> T')"
-(*<*)
- apply (unfold sup_ty_opt_def)
- apply (rule le_OK_conv [simplified lesub_def])
- done
-(*>*)
-
-lemma OK_any_conv:
- "P \<turnstile> OK T \<le>\<^sub>\<top> X = (X = Err \<or> (\<exists>T'. X = OK T' \<and> P \<turnstile> T \<le> T'))"
-(*<*)
- apply (unfold sup_ty_opt_def)
- apply (rule OK_le_conv [simplified lesub_def])
- done
-(*>*)
-
-lemma sup_ty_opt_trans [intro?, trans]:
- "\<lbrakk>P \<turnstile> a \<le>\<^sub>\<top> b; P \<turnstile> b \<le>\<^sub>\<top> c\<rbrakk> \<Longrightarrow> P \<turnstile> a \<le>\<^sub>\<top> c"
-(*<*) by (auto intro: widen_trans
- simp add: sup_ty_opt_def Err.le_def lesub_def fun_of_def
- split: err.splits) (*>*)
-
-
-subsection "Stack and Registers"
-
-lemma stk_convert:
- "P \<turnstile> ST [\<le>] ST' = Listn.le (subtype P) ST ST'"
-(*<*) by (simp add: Listn.le_def lesub_def) (*>*)
-
-lemma sup_loc_refl [iff]: "P \<turnstile> LT [\<le>\<^sub>\<top>] LT"
-(*<*) by (rule list_all2_refl) simp (*>*)
-
-lemmas sup_loc_Cons1 [iff] = list_all2_Cons1 [of "sup_ty_opt P"] for P
-
-lemma sup_loc_def:
- "P \<turnstile> LT [\<le>\<^sub>\<top>] LT' \<equiv> Listn.le (sup_ty_opt P) LT LT'"
-(*<*) by (simp add: Listn.le_def lesub_def) (*>*)
-
-lemma sup_loc_widens_conv [iff]:
- "P \<turnstile> map OK Ts [\<le>\<^sub>\<top>] map OK Ts' = P \<turnstile> Ts [\<le>] Ts'"
-(*<*)
- by (simp add: list_all2_map1 list_all2_map2)
-(*>*)
-
-
-lemma sup_loc_trans [intro?, trans]:
- "\<lbrakk>P \<turnstile> a [\<le>\<^sub>\<top>] b; P \<turnstile> b [\<le>\<^sub>\<top>] c\<rbrakk> \<Longrightarrow> P \<turnstile> a [\<le>\<^sub>\<top>] c"
-(*<*) by (rule list_all2_trans, rule sup_ty_opt_trans) (*>*)
-
-
-subsection "State Type"
-
-lemma sup_state_conv [iff]:
- "P \<turnstile> (ST,LT) \<le>\<^sub>i (ST',LT') = (P \<turnstile> ST [\<le>] ST' \<and> P \<turnstile> LT [\<le>\<^sub>\<top>] LT')"
-(*<*) by (auto simp add: sup_state_def stk_convert lesub_def Product.le_def sup_loc_def) (*>*)
-
-lemma sup_state_conv2:
- "P \<turnstile> s1 \<le>\<^sub>i s2 = (P \<turnstile> fst s1 [\<le>] fst s2 \<and> P \<turnstile> snd s1 [\<le>\<^sub>\<top>] snd s2)"
-(*<*) by (cases s1, cases s2) simp (*>*)
-
-lemma sup_state_refl [iff]: "P \<turnstile> s \<le>\<^sub>i s"
-(*<*) by (auto simp add: sup_state_conv2) (*>*)
-
-lemma sup_state_trans [intro?, trans]:
- "\<lbrakk>P \<turnstile> a \<le>\<^sub>i b; P \<turnstile> b \<le>\<^sub>i c\<rbrakk> \<Longrightarrow> P \<turnstile> a \<le>\<^sub>i c"
-(*<*) by (auto intro: sup_loc_trans widens_trans simp add: sup_state_conv2) (*>*)
-
-
-lemma sup_state_opt_None_any [iff]:
- "P \<turnstile> None \<le>' s"
-(*<*) by (simp add: sup_state_opt_def Opt.le_def) (*>*)
-
-lemma sup_state_opt_any_None [iff]:
- "P \<turnstile> s \<le>' None = (s = None)"
-(*<*) by (simp add: sup_state_opt_def Opt.le_def) (*>*)
-
-lemma sup_state_opt_Some_Some [iff]:
- "P \<turnstile> Some a \<le>' Some b = P \<turnstile> a \<le>\<^sub>i b"
-(*<*) by (simp add: sup_state_opt_def Opt.le_def lesub_def) (*>*)
-
-lemma sup_state_opt_any_Some:
- "P \<turnstile> (Some s) \<le>' X = (\<exists>s'. X = Some s' \<and> P \<turnstile> s \<le>\<^sub>i s')"
-(*<*) by (simp add: sup_state_opt_def Opt.le_def lesub_def) (*>*)
-
-lemma sup_state_opt_refl [iff]: "P \<turnstile> s \<le>' s"
-(*<*) by (simp add: sup_state_opt_def Opt.le_def lesub_def) (*>*)
-
-lemma sup_state_opt_trans [intro?, trans]:
- "\<lbrakk>P \<turnstile> a \<le>' b; P \<turnstile> b \<le>' c\<rbrakk> \<Longrightarrow> P \<turnstile> a \<le>' c"
-(*<*)
- apply (unfold sup_state_opt_def Opt.le_def lesub_def)
- apply (simp del: split_paired_All)
- apply (rule sup_state_trans, assumption+)
- done
-(*>*)
-
-
-lemma sup_state_opt_err : "(Err.le (sup_state_opt P)) s s"
- apply (unfold JVM_le_unfold Product.le_def Opt.le_def Err.le_def lesssub_def lesub_def Listn.le_def)
- apply (auto split: err.splits)
- done
-
-lemma Cons_less_Conss1 [simp]:
- "x#xs [\<sqsubset>\<^bsub>subtype P\<^esub>] y#ys = (x \<sqsubset>\<^bsub>subtype P\<^esub> y \<and> xs [\<sqsubseteq>\<^bsub>subtype P\<^esub>] ys \<or> x = y \<and> xs [\<sqsubset>\<^bsub>subtype P\<^esub>] ys)"
- apply (unfold lesssub_def )
- apply auto
- apply (simp add:lesssub_def lesub_def) \<comment>\<open>widen_refl, subtype_refl \<close>
- done
-
-lemma Cons_less_Conss2 [simp]:
- "x#xs [\<sqsubset>\<^bsub>Err.le (subtype P)\<^esub>] y#ys = (x \<sqsubset>\<^bsub>Err.le (subtype P)\<^esub> y \<and> xs [\<sqsubseteq>\<^bsub>Err.le (subtype P)\<^esub>] ys \<or> x = y \<and> xs [\<sqsubset>\<^bsub>Err.le (subtype P)\<^esub>] ys)"
- apply (unfold lesssub_def )
- apply auto
- apply (simp add:lesssub_def lesub_def Err.le_def split: err.splits)
- done
-
-lemma acc_le_listI1 [intro!]:
- " acc (subtype P) \<Longrightarrow> acc (Listn.le (subtype P))"
- (*<*)
-apply (unfold acc_def)
-apply (subgoal_tac
- "wf(UN n. {(ys,xs). size xs = n \<and> size ys = n \<and> xs <_(Listn.le (subtype P)) ys})")
- apply (erule wf_subset)
-
- apply (blast intro: lesssub_lengthD)
-apply (rule wf_UN)
- prefer 2
- apply (rename_tac m n)
- apply (case_tac "m=n")
- apply simp
- apply (fast intro!: equals0I dest: not_sym)
-apply (rename_tac n)
-apply (induct_tac n)
- apply (simp add: lesssub_def cong: conj_cong)
-apply (rename_tac k)
-apply (simp add: wf_eq_minimal)
-apply (simp (no_asm) add: length_Suc_conv cong: conj_cong)
-apply clarify
-apply (rename_tac M m)
-apply (case_tac "\<exists>x xs. size xs = k \<and> x#xs \<in> M")
- prefer 2
- apply (erule thin_rl)
- apply (erule thin_rl)
- apply blast
-apply (erule_tac x = "{a. \<exists>xs. size xs = k \<and> a#xs:M}" in allE)
-apply (erule impE)
- apply blast
-apply (thin_tac "\<exists>x xs. P x xs" for P)
-apply clarify
-apply (rename_tac maxA xs)
-apply (erule_tac x = "{ys. size ys = size xs \<and> maxA#ys \<in> M}" in allE)
-apply (erule impE)
- apply blast
-apply clarify
-apply (thin_tac "m \<in> M")
-apply (thin_tac "maxA#xs \<in> M")
-apply (rule bexI)
- prefer 2
- apply assumption
-apply clarify
-apply simp
-apply blast
- done
-
-lemma acc_le_listI2 [intro!]:
- " acc (Err.le (subtype P)) \<Longrightarrow> acc (Listn.le (Err.le (subtype P)))"
- (*<*)
-apply (unfold acc_def)
-apply (subgoal_tac
- "wf(UN n. {(ys,xs). size xs = n \<and> size ys = n \<and> xs <_(Listn.le (Err.le (subtype P))) ys})")
- apply (erule wf_subset)
-
- apply (blast intro: lesssub_lengthD)
-apply (rule wf_UN)
- prefer 2
- apply (rename_tac m n)
- apply (case_tac "m=n")
- apply simp
- apply (fast intro!: equals0I dest: not_sym)
-apply (rename_tac n)
-apply (induct_tac n)
- apply (simp add: lesssub_def cong: conj_cong)
-apply (rename_tac k)
-apply (simp add: wf_eq_minimal)
-apply (simp (no_asm) add: length_Suc_conv cong: conj_cong)
-apply clarify
-apply (rename_tac M m)
-apply (case_tac "\<exists>x xs. size xs = k \<and> x#xs \<in> M")
- prefer 2
- apply (erule thin_rl)
- apply (erule thin_rl)
- apply blast
-apply (erule_tac x = "{a. \<exists>xs. size xs = k \<and> a#xs:M}" in allE)
-apply (erule impE)
- apply blast
-apply (thin_tac "\<exists>x xs. P x xs" for P)
-apply clarify
-apply (rename_tac maxA xs)
-apply (erule_tac x = "{ys. size ys = size xs \<and> maxA#ys \<in> M}" in allE)
-apply (erule impE)
- apply blast
-apply clarify
-apply (thin_tac "m \<in> M")
-apply (thin_tac "maxA#xs \<in> M")
-apply (rule bexI)
- prefer 2
- apply assumption
-apply clarify
-apply simp
-apply blast
- done
-
-lemma acc_JVM [intro]:
- "wf_prog wf_mb P \<Longrightarrow> acc (JVM_SemiType.le P mxs mxl)"
-(*<*) by (unfold JVM_le_unfold) blast (*>*) \<comment>\<open> use acc_listI1, acc_listI2 \<close>
-
-end
+(* Title: HOL/MicroJava/BV/JVM.thy
+
+ Author: Gerwin Klein
+ Copyright 2000 TUM
+
+*)
+
+section \<open> The JVM Type System as Semilattice \<close>
+
+theory JVM_SemiType imports SemiType begin
+
+type_synonym ty\<^sub>l = "ty err list"
+type_synonym ty\<^sub>s = "ty list"
+type_synonym ty\<^sub>i = "ty\<^sub>s \<times> ty\<^sub>l"
+type_synonym ty\<^sub>i' = "ty\<^sub>i option"
+type_synonym ty\<^sub>m = "ty\<^sub>i' list"
+type_synonym ty\<^sub>P = "mname \<Rightarrow> cname \<Rightarrow> ty\<^sub>m"
+
+
+definition stk_esl :: "'c prog \<Rightarrow> nat \<Rightarrow> ty\<^sub>s esl"
+where
+ "stk_esl P mxs \<equiv> upto_esl mxs (SemiType.esl P)"
+
+definition loc_sl :: "'c prog \<Rightarrow> nat \<Rightarrow> ty\<^sub>l sl"
+where
+ "loc_sl P mxl \<equiv> Listn.sl mxl (Err.sl (SemiType.esl P))"
+
+definition sl :: "'c prog \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> ty\<^sub>i' err sl"
+where
+ "sl P mxs mxl \<equiv>
+ Err.sl(Opt.esl(Product.esl (stk_esl P mxs) (Err.esl(loc_sl P mxl))))"
+
+
+definition states :: "'c prog \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> ty\<^sub>i' err set"
+where "states P mxs mxl \<equiv> fst(sl P mxs mxl)"
+
+definition le :: "'c prog \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> ty\<^sub>i' err ord"
+where
+ "le P mxs mxl \<equiv> fst(snd(sl P mxs mxl))"
+
+definition sup :: "'c prog \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> ty\<^sub>i' err binop"
+where
+ "sup P mxs mxl \<equiv> snd(snd(sl P mxs mxl))"
+
+
+definition sup_ty_opt :: "['c prog,ty err,ty err] \<Rightarrow> bool"
+ ("_ \<turnstile> _ \<le>\<^sub>\<top> _" [71,71,71] 70)
+where
+ "sup_ty_opt P \<equiv> Err.le (subtype P)"
+
+definition sup_state :: "['c prog,ty\<^sub>i,ty\<^sub>i] \<Rightarrow> bool"
+ ("_ \<turnstile> _ \<le>\<^sub>i _" [71,71,71] 70)
+where
+ "sup_state P \<equiv> Product.le (Listn.le (subtype P)) (Listn.le (sup_ty_opt P))"
+
+definition sup_state_opt :: "['c prog,ty\<^sub>i',ty\<^sub>i'] \<Rightarrow> bool"
+ ("_ \<turnstile> _ \<le>'' _" [71,71,71] 70)
+where
+ "sup_state_opt P \<equiv> Opt.le (sup_state P)"
+
+abbreviation
+ sup_loc :: "['c prog,ty\<^sub>l,ty\<^sub>l] \<Rightarrow> bool" ("_ \<turnstile> _ [\<le>\<^sub>\<top>] _" [71,71,71] 70)
+ where "P \<turnstile> LT [\<le>\<^sub>\<top>] LT' \<equiv> list_all2 (sup_ty_opt P) LT LT'"
+
+notation (ASCII)
+ sup_ty_opt ("_ |- _ <=T _" [71,71,71] 70) and
+ sup_state ("_ |- _ <=i _" [71,71,71] 70) and
+ sup_state_opt ("_ |- _ <=' _" [71,71,71] 70) and
+ sup_loc ("_ |- _ [<=T] _" [71,71,71] 70)
+
+
+subsection "Unfolding"
+
+lemma JVM_states_unfold:
+ "states P mxs mxl \<equiv> err(opt((Union {list n (types P) |n. n <= mxs}) \<times>
+ list mxl (err(types P))))"
+(*<*)
+ apply (unfold states_def sl_def Opt.esl_def Err.sl_def
+ stk_esl_def loc_sl_def Product.esl_def
+ Listn.sl_def upto_esl_def SemiType.esl_def Err.esl_def)
+ apply simp
+ done
+(*>*)
+
+lemma JVM_le_unfold:
+ "le P m n \<equiv>
+ Err.le(Opt.le(Product.le(Listn.le(subtype P))(Listn.le(Err.le(subtype P)))))"
+(*<*)
+ apply (unfold le_def sl_def Opt.esl_def Err.sl_def
+ stk_esl_def loc_sl_def Product.esl_def
+ Listn.sl_def upto_esl_def SemiType.esl_def Err.esl_def)
+ apply simp
+ done
+(*>*)
+
+lemma sl_def2:
+ "JVM_SemiType.sl P mxs mxl \<equiv>
+ (states P mxs mxl, JVM_SemiType.le P mxs mxl, JVM_SemiType.sup P mxs mxl)"
+(*<*) by (unfold JVM_SemiType.sup_def states_def JVM_SemiType.le_def) simp (*>*)
+
+
+lemma JVM_le_conv:
+ "le P m n (OK t1) (OK t2) = P \<turnstile> t1 \<le>' t2"
+(*<*) by (simp add: JVM_le_unfold Err.le_def lesub_def sup_state_opt_def
+ sup_state_def sup_ty_opt_def) (*>*)
+
+lemma JVM_le_Err_conv:
+ "le P m n = Err.le (sup_state_opt P)"
+(*<*) by (unfold sup_state_opt_def sup_state_def
+ sup_ty_opt_def JVM_le_unfold) simp (*>*)
+
+lemma err_le_unfold [iff]:
+ "Err.le r (OK a) (OK b) = r a b"
+(*<*) by (simp add: Err.le_def lesub_def) (*>*)
+
+
+subsection \<open> Semilattice \<close>
+lemma order_sup_state_opt' [intro, simp]:
+ "wf_prog wf_mb P \<Longrightarrow>
+ order (sup_state_opt P) (opt ((\<Union> {list n (types P) |n. n \<le> mxs} ) \<times> list (Suc (length Ts + mxl\<^sub>0)) (err (types P))))"
+(*<*)
+ apply (unfold sup_state_opt_def sup_state_def sup_ty_opt_def )
+ apply (blast intro:order_le_prodI) \<comment>\<open> use Listn.thy.order_listI2 \<close>
+ done
+(*<*)
+lemma order_sup_state_opt'' [intro, simp]:
+ "wf_prog wf_mb P \<Longrightarrow>
+ order (sup_state_opt P) (opt ((\<Union> {list n (types P) |n. n \<le> mxs} ) \<times> list ((length Ts + mxl\<^sub>0)) (err (types P))))"
+(*<*)
+ apply (unfold sup_state_opt_def sup_state_def sup_ty_opt_def )
+ apply (blast intro:order_le_prodI) \<comment>\<open> use Listn.thy.order_listI2 \<close>
+ done
+(*<*)
+(*
+lemma order_sup_state_opt [intro, simp]:
+ "wf_prog wf_mb P \<Longrightarrow> order (sup_state_opt P)"
+(*<*) by (unfold sup_state_opt_def sup_state_def sup_ty_opt_def) blast (*>*)
+*)
+
+lemma semilat_JVM [intro?]:
+ "wf_prog wf_mb P \<Longrightarrow> semilat (JVM_SemiType.sl P mxs mxl)"
+(*<*)
+ apply (unfold JVM_SemiType.sl_def stk_esl_def loc_sl_def)
+ apply (blast intro: err_semilat_Product_esl err_semilat_upto_esl
+ Listn_sl err_semilat_JType_esl)
+ done
+(*>*)
+
+subsection \<open> Widening with @{text "\<top>"} \<close>
+
+lemma subtype_refl[iff]: "subtype P t t" (*<*) by (simp add: fun_of_def) (*>*)
+
+lemma sup_ty_opt_refl [iff]: "P \<turnstile> T \<le>\<^sub>\<top> T"
+(*<*)
+ apply (unfold sup_ty_opt_def)
+ apply (fold lesub_def)
+ apply (rule le_err_refl)
+ apply (simp add: lesub_def)
+ done
+(*>*)
+
+lemma Err_any_conv [iff]: "P \<turnstile> Err \<le>\<^sub>\<top> T = (T = Err)"
+(*<*) by (unfold sup_ty_opt_def) (rule Err_le_conv [simplified lesub_def]) (*>*)
+
+lemma any_Err [iff]: "P \<turnstile> T \<le>\<^sub>\<top> Err"
+(*<*) by (unfold sup_ty_opt_def) (rule le_Err [simplified lesub_def]) (*>*)
+
+lemma OK_OK_conv [iff]:
+ "P \<turnstile> OK T \<le>\<^sub>\<top> OK T' = P \<turnstile> T \<le> T'"
+(*<*) by (simp add: sup_ty_opt_def fun_of_def) (*>*)
+
+lemma any_OK_conv [iff]:
+ "P \<turnstile> X \<le>\<^sub>\<top> OK T' = (\<exists>T. X = OK T \<and> P \<turnstile> T \<le> T')"
+(*<*)
+ apply (unfold sup_ty_opt_def)
+ apply (rule le_OK_conv [simplified lesub_def])
+ done
+(*>*)
+
+lemma OK_any_conv:
+ "P \<turnstile> OK T \<le>\<^sub>\<top> X = (X = Err \<or> (\<exists>T'. X = OK T' \<and> P \<turnstile> T \<le> T'))"
+(*<*)
+ apply (unfold sup_ty_opt_def)
+ apply (rule OK_le_conv [simplified lesub_def])
+ done
+(*>*)
+
+lemma sup_ty_opt_trans [intro?, trans]:
+ "\<lbrakk>P \<turnstile> a \<le>\<^sub>\<top> b; P \<turnstile> b \<le>\<^sub>\<top> c\<rbrakk> \<Longrightarrow> P \<turnstile> a \<le>\<^sub>\<top> c"
+(*<*) by (auto intro: widen_trans
+ simp add: sup_ty_opt_def Err.le_def lesub_def fun_of_def
+ split: err.splits) (*>*)
+
+
+subsection "Stack and Registers"
+
+lemma stk_convert:
+ "P \<turnstile> ST [\<le>] ST' = Listn.le (subtype P) ST ST'"
+(*<*) by (simp add: Listn.le_def lesub_def) (*>*)
+
+lemma sup_loc_refl [iff]: "P \<turnstile> LT [\<le>\<^sub>\<top>] LT"
+(*<*) by (rule list_all2_refl) simp (*>*)
+
+lemmas sup_loc_Cons1 [iff] = list_all2_Cons1 [of "sup_ty_opt P"] for P
+
+lemma sup_loc_def:
+ "P \<turnstile> LT [\<le>\<^sub>\<top>] LT' \<equiv> Listn.le (sup_ty_opt P) LT LT'"
+(*<*) by (simp add: Listn.le_def lesub_def) (*>*)
+
+lemma sup_loc_widens_conv [iff]:
+ "P \<turnstile> map OK Ts [\<le>\<^sub>\<top>] map OK Ts' = P \<turnstile> Ts [\<le>] Ts'"
+(*<*)
+ by (simp add: list_all2_map1 list_all2_map2)
+(*>*)
+
+
+lemma sup_loc_trans [intro?, trans]:
+ "\<lbrakk>P \<turnstile> a [\<le>\<^sub>\<top>] b; P \<turnstile> b [\<le>\<^sub>\<top>] c\<rbrakk> \<Longrightarrow> P \<turnstile> a [\<le>\<^sub>\<top>] c"
+(*<*) by (rule list_all2_trans, rule sup_ty_opt_trans) (*>*)
+
+
+subsection "State Type"
+
+lemma sup_state_conv [iff]:
+ "P \<turnstile> (ST,LT) \<le>\<^sub>i (ST',LT') = (P \<turnstile> ST [\<le>] ST' \<and> P \<turnstile> LT [\<le>\<^sub>\<top>] LT')"
+(*<*) by (auto simp add: sup_state_def stk_convert lesub_def Product.le_def sup_loc_def) (*>*)
+
+lemma sup_state_conv2:
+ "P \<turnstile> s1 \<le>\<^sub>i s2 = (P \<turnstile> fst s1 [\<le>] fst s2 \<and> P \<turnstile> snd s1 [\<le>\<^sub>\<top>] snd s2)"
+(*<*) by (cases s1, cases s2) simp (*>*)
+
+lemma sup_state_refl [iff]: "P \<turnstile> s \<le>\<^sub>i s"
+(*<*) by (auto simp add: sup_state_conv2) (*>*)
+
+lemma sup_state_trans [intro?, trans]:
+ "\<lbrakk>P \<turnstile> a \<le>\<^sub>i b; P \<turnstile> b \<le>\<^sub>i c\<rbrakk> \<Longrightarrow> P \<turnstile> a \<le>\<^sub>i c"
+(*<*) by (auto intro: sup_loc_trans widens_trans simp add: sup_state_conv2) (*>*)
+
+
+lemma sup_state_opt_None_any [iff]:
+ "P \<turnstile> None \<le>' s"
+(*<*) by (simp add: sup_state_opt_def Opt.le_def) (*>*)
+
+lemma sup_state_opt_any_None [iff]:
+ "P \<turnstile> s \<le>' None = (s = None)"
+(*<*) by (simp add: sup_state_opt_def Opt.le_def) (*>*)
+
+lemma sup_state_opt_Some_Some [iff]:
+ "P \<turnstile> Some a \<le>' Some b = P \<turnstile> a \<le>\<^sub>i b"
+(*<*) by (simp add: sup_state_opt_def Opt.le_def lesub_def) (*>*)
+
+lemma sup_state_opt_any_Some:
+ "P \<turnstile> (Some s) \<le>' X = (\<exists>s'. X = Some s' \<and> P \<turnstile> s \<le>\<^sub>i s')"
+(*<*) by (simp add: sup_state_opt_def Opt.le_def lesub_def) (*>*)
+
+lemma sup_state_opt_refl [iff]: "P \<turnstile> s \<le>' s"
+(*<*) by (simp add: sup_state_opt_def Opt.le_def lesub_def) (*>*)
+
+lemma sup_state_opt_trans [intro?, trans]:
+ "\<lbrakk>P \<turnstile> a \<le>' b; P \<turnstile> b \<le>' c\<rbrakk> \<Longrightarrow> P \<turnstile> a \<le>' c"
+(*<*)
+ apply (unfold sup_state_opt_def Opt.le_def lesub_def)
+ apply (simp del: split_paired_All)
+ apply (rule sup_state_trans, assumption+)
+ done
+(*>*)
+
+
+lemma sup_state_opt_err : "(Err.le (sup_state_opt P)) s s"
+ apply (unfold JVM_le_unfold Product.le_def Opt.le_def Err.le_def lesssub_def lesub_def Listn.le_def)
+ apply (auto split: err.splits)
+ done
+
+lemma Cons_less_Conss1 [simp]:
+ "x#xs [\<sqsubset>\<^bsub>subtype P\<^esub>] y#ys = (x \<sqsubset>\<^bsub>subtype P\<^esub> y \<and> xs [\<sqsubseteq>\<^bsub>subtype P\<^esub>] ys \<or> x = y \<and> xs [\<sqsubset>\<^bsub>subtype P\<^esub>] ys)"
+ apply (unfold lesssub_def )
+ apply auto
+ apply (simp add:lesssub_def lesub_def) \<comment>\<open>widen_refl, subtype_refl \<close>
+ done
+
+lemma Cons_less_Conss2 [simp]:
+ "x#xs [\<sqsubset>\<^bsub>Err.le (subtype P)\<^esub>] y#ys = (x \<sqsubset>\<^bsub>Err.le (subtype P)\<^esub> y \<and> xs [\<sqsubseteq>\<^bsub>Err.le (subtype P)\<^esub>] ys \<or> x = y \<and> xs [\<sqsubset>\<^bsub>Err.le (subtype P)\<^esub>] ys)"
+ apply (unfold lesssub_def )
+ apply auto
+ apply (simp add:lesssub_def lesub_def Err.le_def split: err.splits)
+ done
+
+lemma acc_le_listI1 [intro!]:
+ " acc (subtype P) \<Longrightarrow> acc (Listn.le (subtype P))"
+ (*<*)
+apply (unfold acc_def)
+apply (subgoal_tac
+ "wf(UN n. {(ys,xs). size xs = n \<and> size ys = n \<and> xs <_(Listn.le (subtype P)) ys})")
+ apply (erule wf_subset)
+
+ apply (blast intro: lesssub_lengthD)
+apply (rule wf_UN)
+ prefer 2
+ apply (rename_tac m n)
+ apply (case_tac "m=n")
+ apply simp
+ apply (fast intro!: equals0I dest: not_sym)
+apply (rename_tac n)
+apply (induct_tac n)
+ apply (simp add: lesssub_def cong: conj_cong)
+apply (rename_tac k)
+apply (simp add: wf_eq_minimal)
+apply (simp (no_asm) add: length_Suc_conv cong: conj_cong)
+apply clarify
+apply (rename_tac M m)
+apply (case_tac "\<exists>x xs. size xs = k \<and> x#xs \<in> M")
+ prefer 2
+ apply (erule thin_rl)
+ apply (erule thin_rl)
+ apply blast
+apply (erule_tac x = "{a. \<exists>xs. size xs = k \<and> a#xs:M}" in allE)
+apply (erule impE)
+ apply blast
+apply (thin_tac "\<exists>x xs. P x xs" for P)
+apply clarify
+apply (rename_tac maxA xs)
+apply (erule_tac x = "{ys. size ys = size xs \<and> maxA#ys \<in> M}" in allE)
+apply (erule impE)
+ apply blast
+apply clarify
+apply (thin_tac "m \<in> M")
+apply (thin_tac "maxA#xs \<in> M")
+apply (rule bexI)
+ prefer 2
+ apply assumption
+apply clarify
+apply simp
+apply blast
+ done
+
+lemma acc_le_listI2 [intro!]:
+ " acc (Err.le (subtype P)) \<Longrightarrow> acc (Listn.le (Err.le (subtype P)))"
+ (*<*)
+apply (unfold acc_def)
+apply (subgoal_tac
+ "wf(UN n. {(ys,xs). size xs = n \<and> size ys = n \<and> xs <_(Listn.le (Err.le (subtype P))) ys})")
+ apply (erule wf_subset)
+
+ apply (blast intro: lesssub_lengthD)
+apply (rule wf_UN)
+ prefer 2
+ apply (rename_tac m n)
+ apply (case_tac "m=n")
+ apply simp
+ apply (fast intro!: equals0I dest: not_sym)
+apply (rename_tac n)
+apply (induct_tac n)
+ apply (simp add: lesssub_def cong: conj_cong)
+apply (rename_tac k)
+apply (simp add: wf_eq_minimal)
+apply (simp (no_asm) add: length_Suc_conv cong: conj_cong)
+apply clarify
+apply (rename_tac M m)
+apply (case_tac "\<exists>x xs. size xs = k \<and> x#xs \<in> M")
+ prefer 2
+ apply (erule thin_rl)
+ apply (erule thin_rl)
+ apply blast
+apply (erule_tac x = "{a. \<exists>xs. size xs = k \<and> a#xs:M}" in allE)
+apply (erule impE)
+ apply blast
+apply (thin_tac "\<exists>x xs. P x xs" for P)
+apply clarify
+apply (rename_tac maxA xs)
+apply (erule_tac x = "{ys. size ys = size xs \<and> maxA#ys \<in> M}" in allE)
+apply (erule impE)
+ apply blast
+apply clarify
+apply (thin_tac "m \<in> M")
+apply (thin_tac "maxA#xs \<in> M")
+apply (rule bexI)
+ prefer 2
+ apply assumption
+apply clarify
+apply simp
+apply blast
+ done
+
+lemma acc_JVM [intro]:
+ "wf_prog wf_mb P \<Longrightarrow> acc (JVM_SemiType.le P mxs mxl)"
+(*<*) by (unfold JVM_le_unfold) blast (*>*) \<comment>\<open> use acc_listI1, acc_listI2 \<close>
+
+end
diff --git a/thys/JinjaDCI/BV/LBVJVM.thy b/thys/JinjaDCI/BV/LBVJVM.thy
--- a/thys/JinjaDCI/BV/LBVJVM.thy
+++ b/thys/JinjaDCI/BV/LBVJVM.thy
@@ -1,250 +1,250 @@
-(* Title: JinjaDCI/BV/LBVJVM.thy
-
- Author: Tobias Nipkow, Gerwin Klein, Susannah Mansky
- Copyright 2000 TUM, 2020 UIUC
-
- Based on the Jinja theory BV/LBVJVM.thy by Tobias Nipkow and Gerwin Klein
-*)
-
-section \<open> LBV for the JVM \label{sec:JVM} \<close>
-
-theory LBVJVM
-imports Jinja.Abstract_BV TF_JVM
-begin
-
-type_synonym prog_cert = "cname \<Rightarrow> mname \<Rightarrow> ty\<^sub>i' err list"
-
-definition check_cert :: "jvm_prog \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> ty\<^sub>i' err list \<Rightarrow> bool"
-where
- "check_cert P mxs mxl n cert \<equiv> check_types P mxs mxl cert \<and> size cert = n+1 \<and>
- (\<forall>i<n. cert!i \<noteq> Err) \<and> cert!n = OK None"
-
-definition lbvjvm :: "jvm_prog \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> ty \<Rightarrow> ex_table \<Rightarrow>
- ty\<^sub>i' err list \<Rightarrow> instr list \<Rightarrow> ty\<^sub>i' err \<Rightarrow> ty\<^sub>i' err"
-where
- "lbvjvm P mxs maxr T\<^sub>r et cert bs \<equiv>
- wtl_inst_list bs cert (JVM_SemiType.sup P mxs maxr) (JVM_SemiType.le P mxs maxr) Err (OK None) (exec P mxs T\<^sub>r et bs) 0"
-
-definition wt_lbv :: "jvm_prog \<Rightarrow> cname \<Rightarrow> staticb \<Rightarrow> ty list \<Rightarrow> ty \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow>
- ex_table \<Rightarrow> ty\<^sub>i' err list \<Rightarrow> instr list \<Rightarrow> bool"
-where
- "wt_lbv P C b Ts T\<^sub>r mxs mxl\<^sub>0 et cert ins \<equiv> (b = Static \<or> b = NonStatic) \<and>
- check_cert P mxs ((case b of Static \<Rightarrow> 0 | NonStatic \<Rightarrow> 1)+size Ts+mxl\<^sub>0) (size ins) cert \<and>
- 0 < size ins \<and>
- (let start = Some ([],(case b of Static \<Rightarrow> [] | NonStatic \<Rightarrow> [OK (Class C)])
- @((map OK Ts))@(replicate mxl\<^sub>0 Err));
- result = lbvjvm P mxs ((case b of Static \<Rightarrow> 0 | NonStatic \<Rightarrow> 1)+size Ts+mxl\<^sub>0) T\<^sub>r et cert ins (OK start)
- in result \<noteq> Err)"
-
-definition wt_jvm_prog_lbv :: "jvm_prog \<Rightarrow> prog_cert \<Rightarrow> bool"
-where
- "wt_jvm_prog_lbv P cert \<equiv>
- wf_prog (\<lambda>P C (mn,b,Ts,T\<^sub>r,(mxs,mxl\<^sub>0,ins,et)). wt_lbv P C b Ts T\<^sub>r mxs mxl\<^sub>0 et (cert C mn) ins) P"
-
-definition mk_cert :: "jvm_prog \<Rightarrow> nat \<Rightarrow> ty \<Rightarrow> ex_table \<Rightarrow> instr list
- \<Rightarrow> ty\<^sub>m \<Rightarrow> ty\<^sub>i' err list"
-where
- "mk_cert P mxs T\<^sub>r et bs phi \<equiv> make_cert (exec P mxs T\<^sub>r et bs) (map OK phi) (OK None)"
-
-definition prg_cert :: "jvm_prog \<Rightarrow> ty\<^sub>P \<Rightarrow> prog_cert"
-where
- "prg_cert P phi C mn \<equiv> let (C,b,Ts,T\<^sub>r,(mxs,mxl\<^sub>0,ins,et)) = method P C mn
- in mk_cert P mxs T\<^sub>r et ins (phi C mn)"
-
-lemma check_certD [intro?]:
- "check_cert P mxs mxl n cert \<Longrightarrow> cert_ok cert n Err (OK None) (states P mxs mxl)"
- by (unfold cert_ok_def check_cert_def check_types_def) auto
-
-
-lemma (in start_context) wt_lbv_wt_step:
- assumes lbv: "wt_lbv P C b Ts T\<^sub>r mxs mxl\<^sub>0 xt cert is"
- shows "\<exists>\<tau>s \<in> list (size is) A. wt_step r Err step \<tau>s \<and> OK first \<sqsubseteq>\<^sub>r \<tau>s!0"
-(*<*)
-proof -
- from wf have "semilat (JVM_SemiType.sl P mxs mxl)" ..
- hence "semilat (A, r, f)" by (simp add: sl_def2)
- moreover have "top r Err" by (simp add: JVM_le_Err_conv)
- moreover have "Err \<in> A" by (simp add: JVM_states_unfold)
- moreover have "bottom r (OK None)"
- by (simp add: JVM_le_Err_conv bottom_def lesub_def Err.le_def split: err.split)
- moreover have "OK None \<in> A" by (simp add: JVM_states_unfold)
- moreover note bounded_step
- moreover from lbv have "cert_ok cert (size is) Err (OK None) A"
- by (unfold wt_lbv_def) (auto dest: check_certD)
- moreover note exec_pres_type
- moreover
- from lbv
- have "wtl_inst_list is cert f r Err (OK None) step 0 (OK first) \<noteq> Err"
- by (cases b; simp add: wt_lbv_def lbvjvm_def step_def_exec [symmetric])
- moreover note first_in_A
- moreover from lbv have "0 < size is" by (simp add: wt_lbv_def)
- ultimately show ?thesis by (rule lbvs.wtl_sound_strong [OF lbvs.intro, OF lbv.intro lbvs_axioms.intro, OF Semilat.intro lbv_axioms.intro])
-qed
-(*>*)
-
-
-lemma (in start_context) wt_lbv_wt_method:
- assumes lbv: "wt_lbv P C b Ts T\<^sub>r mxs mxl\<^sub>0 xt cert is"
- shows "\<exists>\<tau>s. wt_method P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt \<tau>s"
-(*<*)
-proof -
- from lbv have l: "is \<noteq> []" and
- stab: "b = Static \<or> b = NonStatic" by (auto simp add: wt_lbv_def)
- moreover
- from wf lbv C Ts obtain \<tau>s where
- list: "\<tau>s \<in> list (size is) A" and
- step: "wt_step r Err step \<tau>s" and
- start: "OK first \<sqsubseteq>\<^sub>r \<tau>s!0"
- by (blast dest: wt_lbv_wt_step)
- from list have [simp]: "size \<tau>s = size is" by simp
- have "size (map ok_val \<tau>s) = size is" by simp
- moreover from l have 0: "0 < size \<tau>s" by simp
- with step obtain \<tau>s0 where "\<tau>s!0 = OK \<tau>s0"
- by (unfold wt_step_def) blast
- with start 0 have "wt_start P C b Ts mxl\<^sub>0 (map ok_val \<tau>s)"
- by (cases b; simp add: wt_start_def JVM_le_Err_conv lesub_def Err.le_def)
- moreover {
- from list have "check_types P mxs mxl \<tau>s" by (simp add: check_types_def)
- also from step have "\<forall>x \<in> set \<tau>s. x \<noteq> Err"
- by (auto simp add: all_set_conv_all_nth wt_step_def)
- hence [symmetric]: "map OK (map ok_val \<tau>s) = \<tau>s"
- by (auto intro!: map_idI)
- finally have "check_types P mxs mxl (map OK (map ok_val \<tau>s))" .
- }
- moreover {
- note bounded_step
- moreover from list have "set \<tau>s \<subseteq> A" by simp
- moreover from step have "wt_err_step (sup_state_opt P) step \<tau>s"
- by (simp add: wt_err_step_def JVM_le_Err_conv)
- ultimately have "wt_app_eff (sup_state_opt P) app eff (map ok_val \<tau>s)"
- by (auto intro: wt_err_imp_wt_app_eff simp add: exec_def states_def)
- }
- ultimately have "wt_method P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt (map ok_val \<tau>s)"
- by (simp add: wt_method_def2 check_types_def del: map_map)
- thus ?thesis ..
-qed
-(*>*)
-
-
-lemma (in start_context) wt_method_wt_lbv:
- assumes wt: "wt_method P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt \<tau>s"
- defines [simp]: "cert \<equiv> mk_cert P mxs T\<^sub>r xt is \<tau>s"
-
- shows "wt_lbv P C b Ts T\<^sub>r mxs mxl\<^sub>0 xt cert is"
-(*<*)
-proof -
- let ?\<tau>s = "map OK \<tau>s"
- let ?cert = "make_cert step ?\<tau>s (OK None)"
-
- from wt obtain
- 0: "0 < size is" and
- size: "size is = size ?\<tau>s" and
- ck_types: "check_types P mxs mxl ?\<tau>s" and
- wt_start: "wt_start P C b Ts mxl\<^sub>0 \<tau>s" and
- app_eff: "wt_app_eff (sup_state_opt P) app eff \<tau>s"
- by (force simp add: wt_method_def2 check_types_def)
-
- from wf have "semilat (JVM_SemiType.sl P mxs mxl)" ..
- hence "semilat (A, r, f)" by (simp add: sl_def2)
- moreover have "top r Err" by (simp add: JVM_le_Err_conv)
- moreover have "Err \<in> A" by (simp add: JVM_states_unfold)
- moreover have "bottom r (OK None)"
- by (simp add: JVM_le_Err_conv bottom_def lesub_def Err.le_def split: err.split)
- moreover have "OK None \<in> A" by (simp add: JVM_states_unfold)
- moreover from wf have "mono r step (size is) A" by (rule step_mono)
- hence "mono r step (size ?\<tau>s) A" by (simp add: size)
- moreover from exec_pres_type
- have "pres_type step (size ?\<tau>s) A" by (simp add: size)
- moreover
- from ck_types have \<tau>s_in_A: "set ?\<tau>s \<subseteq> A" by (simp add: check_types_def)
- hence "\<forall>pc. pc < size ?\<tau>s \<longrightarrow> ?\<tau>s!pc \<in> A \<and> ?\<tau>s!pc \<noteq> Err" by auto
- moreover from bounded_step
- have "bounded step (size ?\<tau>s)" by (simp add: size)
- moreover have "OK None \<noteq> Err" by simp
- moreover from bounded_step size \<tau>s_in_A app_eff
- have "wt_err_step (sup_state_opt P) step ?\<tau>s"
- by (auto intro: wt_app_eff_imp_wt_err simp add: exec_def states_def)
- hence "wt_step r Err step ?\<tau>s"
- by (simp add: wt_err_step_def JVM_le_Err_conv)
- moreover
- from 0 size have "0 < size \<tau>s" by auto
- hence "?\<tau>s!0 = OK (\<tau>s!0)" by simp
- with wt_start have "OK first \<sqsubseteq>\<^sub>r ?\<tau>s!0"
- by (cases b; clarsimp simp add: wt_start_def lesub_def Err.le_def JVM_le_Err_conv)
- moreover note first_in_A
- moreover have "OK first \<noteq> Err" by simp
- moreover note size
- ultimately
- have "wtl_inst_list is ?cert f r Err (OK None) step 0 (OK first) \<noteq> Err"
- by (rule lbvc.wtl_complete [OF lbvc.intro, OF lbv.intro lbvc_axioms.intro, OF Semilat.intro lbv_axioms.intro])
- moreover from 0 size have "\<tau>s \<noteq> []" by auto
- moreover from ck_types have "check_types P mxs mxl ?cert"
- by (fastforce simp: make_cert_def check_types_def JVM_states_unfold
- dest!: nth_mem)
- moreover note 0 size
- ultimately show ?thesis using staticb
- by (simp add: wt_lbv_def lbvjvm_def mk_cert_def step_def_exec [symmetric]
- check_cert_def make_cert_def nth_append)
-qed
-(*>*)
-
-
-
-theorem jvm_lbv_correct:
- "wt_jvm_prog_lbv P Cert \<Longrightarrow> wf_jvm_prog P"
-(*<*)
-proof -
- let ?\<Phi> = "\<lambda>C mn. let (C,b,Ts,T\<^sub>r,(mxs,mxl\<^sub>0,is,xt)) = method P C mn in
- SOME \<tau>s. wt_method P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt \<tau>s"
-
- let ?A = "\<lambda>P C (mn,b,Ts,T\<^sub>r,(mxs,mxl\<^sub>0,ins,et)). wt_lbv P C b Ts T\<^sub>r mxs mxl\<^sub>0 et (Cert C mn) ins"
- let ?B = "\<lambda>P C (M,b,Ts,T\<^sub>r,(mxs,mxl\<^sub>0,is,xt)).
- wt_method P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt (?\<Phi> C M)"
-
- assume wt: "wt_jvm_prog_lbv P Cert"
- then have "wf_prog ?A P" by(simp add: wt_jvm_prog_lbv_def)
- moreover {
- fix wf_md C M b Ts Ca T m bd
-
- assume ass1: "wf_prog wf_md P" and sees: "P \<turnstile> Ca sees M, b : Ts\<rightarrow>T = m in Ca" and
- ass2: "set Ts \<subseteq> types P" and ass3: "bd = (M, b, Ts, T, m)" and
- ass4: "?A P Ca bd"
- from ass3 ass4 have stab: "b = Static \<or> b = NonStatic" by (simp add:wt_lbv_def)
- from ass1 sees ass2 ass3 ass4 have "?B P Ca bd" using sees_method_is_class[OF sees]
- by (auto dest!: start_context.wt_lbv_wt_method [OF start_context_intro_auxi[OF stab]]
- intro: someI)
- }
- ultimately have "wf_prog ?B P" by(rule wf_prog_lift)
- hence "wf_jvm_prog\<^bsub>?\<Phi>\<^esub> P" by (simp add: wf_jvm_prog_phi_def)
- thus ?thesis by (unfold wf_jvm_prog_def) blast
-qed
-(*>*)
-
-theorem jvm_lbv_complete:
- assumes wt: "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P"
- shows "wt_jvm_prog_lbv P (prg_cert P \<Phi>)"
-(*<*)
-proof -
- let ?cert = "prg_cert P \<Phi>"
- let ?A = "\<lambda>P C (M,b,Ts,T\<^sub>r,(mxs,mxl\<^sub>0,is,xt)).
- wt_method P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt (\<Phi> C M)"
- let ?B = "\<lambda>P C (mn,b,Ts,T\<^sub>r,(mxs,mxl\<^sub>0,ins,et)).
- wt_lbv P C b Ts T\<^sub>r mxs mxl\<^sub>0 et (?cert C mn) ins"
-
- from wt have "wf_prog ?A P" by(clarsimp simp: wf_jvm_prog_def wf_jvm_prog_phi_def)
- moreover {
- fix wf_md C M b Ts Ca T m bd
- assume ass1: "wf_prog wf_md P" and sees: "P \<turnstile> Ca sees M, b : Ts\<rightarrow>T = m in Ca" and
- ass2: "set Ts \<subseteq> types P" and ass3: "bd = (M, b, Ts, T, m)" and
- ass4: "?A P Ca bd"
- from ass3 ass4 have stab: "b = Static \<or> b = NonStatic" by (simp add:wt_method_def)
- from ass1 sees ass2 ass3 ass4 have "?B P Ca bd" using sees_method_is_class[OF sees]
- by (auto simp add: prg_cert_def
- intro!: start_context.wt_method_wt_lbv start_context_intro_auxi[OF stab])
- }
- ultimately have "wf_prog ?B P" by(rule wf_prog_lift)
- thus "wt_jvm_prog_lbv P (prg_cert P \<Phi>)" by (simp add: wt_jvm_prog_lbv_def)
-qed
-(*>*)
-
-end
+(* Title: JinjaDCI/BV/LBVJVM.thy
+
+ Author: Tobias Nipkow, Gerwin Klein, Susannah Mansky
+ Copyright 2000 TUM, 2020 UIUC
+
+ Based on the Jinja theory BV/LBVJVM.thy by Tobias Nipkow and Gerwin Klein
+*)
+
+section \<open> LBV for the JVM \label{sec:JVM} \<close>
+
+theory LBVJVM
+imports Jinja.Abstract_BV TF_JVM
+begin
+
+type_synonym prog_cert = "cname \<Rightarrow> mname \<Rightarrow> ty\<^sub>i' err list"
+
+definition check_cert :: "jvm_prog \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> ty\<^sub>i' err list \<Rightarrow> bool"
+where
+ "check_cert P mxs mxl n cert \<equiv> check_types P mxs mxl cert \<and> size cert = n+1 \<and>
+ (\<forall>i<n. cert!i \<noteq> Err) \<and> cert!n = OK None"
+
+definition lbvjvm :: "jvm_prog \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> ty \<Rightarrow> ex_table \<Rightarrow>
+ ty\<^sub>i' err list \<Rightarrow> instr list \<Rightarrow> ty\<^sub>i' err \<Rightarrow> ty\<^sub>i' err"
+where
+ "lbvjvm P mxs maxr T\<^sub>r et cert bs \<equiv>
+ wtl_inst_list bs cert (JVM_SemiType.sup P mxs maxr) (JVM_SemiType.le P mxs maxr) Err (OK None) (exec P mxs T\<^sub>r et bs) 0"
+
+definition wt_lbv :: "jvm_prog \<Rightarrow> cname \<Rightarrow> staticb \<Rightarrow> ty list \<Rightarrow> ty \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow>
+ ex_table \<Rightarrow> ty\<^sub>i' err list \<Rightarrow> instr list \<Rightarrow> bool"
+where
+ "wt_lbv P C b Ts T\<^sub>r mxs mxl\<^sub>0 et cert ins \<equiv> (b = Static \<or> b = NonStatic) \<and>
+ check_cert P mxs ((case b of Static \<Rightarrow> 0 | NonStatic \<Rightarrow> 1)+size Ts+mxl\<^sub>0) (size ins) cert \<and>
+ 0 < size ins \<and>
+ (let start = Some ([],(case b of Static \<Rightarrow> [] | NonStatic \<Rightarrow> [OK (Class C)])
+ @((map OK Ts))@(replicate mxl\<^sub>0 Err));
+ result = lbvjvm P mxs ((case b of Static \<Rightarrow> 0 | NonStatic \<Rightarrow> 1)+size Ts+mxl\<^sub>0) T\<^sub>r et cert ins (OK start)
+ in result \<noteq> Err)"
+
+definition wt_jvm_prog_lbv :: "jvm_prog \<Rightarrow> prog_cert \<Rightarrow> bool"
+where
+ "wt_jvm_prog_lbv P cert \<equiv>
+ wf_prog (\<lambda>P C (mn,b,Ts,T\<^sub>r,(mxs,mxl\<^sub>0,ins,et)). wt_lbv P C b Ts T\<^sub>r mxs mxl\<^sub>0 et (cert C mn) ins) P"
+
+definition mk_cert :: "jvm_prog \<Rightarrow> nat \<Rightarrow> ty \<Rightarrow> ex_table \<Rightarrow> instr list
+ \<Rightarrow> ty\<^sub>m \<Rightarrow> ty\<^sub>i' err list"
+where
+ "mk_cert P mxs T\<^sub>r et bs phi \<equiv> make_cert (exec P mxs T\<^sub>r et bs) (map OK phi) (OK None)"
+
+definition prg_cert :: "jvm_prog \<Rightarrow> ty\<^sub>P \<Rightarrow> prog_cert"
+where
+ "prg_cert P phi C mn \<equiv> let (C,b,Ts,T\<^sub>r,(mxs,mxl\<^sub>0,ins,et)) = method P C mn
+ in mk_cert P mxs T\<^sub>r et ins (phi C mn)"
+
+lemma check_certD [intro?]:
+ "check_cert P mxs mxl n cert \<Longrightarrow> cert_ok cert n Err (OK None) (states P mxs mxl)"
+ by (unfold cert_ok_def check_cert_def check_types_def) auto
+
+
+lemma (in start_context) wt_lbv_wt_step:
+ assumes lbv: "wt_lbv P C b Ts T\<^sub>r mxs mxl\<^sub>0 xt cert is"
+ shows "\<exists>\<tau>s \<in> list (size is) A. wt_step r Err step \<tau>s \<and> OK first \<sqsubseteq>\<^sub>r \<tau>s!0"
+(*<*)
+proof -
+ from wf have "semilat (JVM_SemiType.sl P mxs mxl)" ..
+ hence "semilat (A, r, f)" by (simp add: sl_def2)
+ moreover have "top r Err" by (simp add: JVM_le_Err_conv)
+ moreover have "Err \<in> A" by (simp add: JVM_states_unfold)
+ moreover have "bottom r (OK None)"
+ by (simp add: JVM_le_Err_conv bottom_def lesub_def Err.le_def split: err.split)
+ moreover have "OK None \<in> A" by (simp add: JVM_states_unfold)
+ moreover note bounded_step
+ moreover from lbv have "cert_ok cert (size is) Err (OK None) A"
+ by (unfold wt_lbv_def) (auto dest: check_certD)
+ moreover note exec_pres_type
+ moreover
+ from lbv
+ have "wtl_inst_list is cert f r Err (OK None) step 0 (OK first) \<noteq> Err"
+ by (cases b; simp add: wt_lbv_def lbvjvm_def step_def_exec [symmetric])
+ moreover note first_in_A
+ moreover from lbv have "0 < size is" by (simp add: wt_lbv_def)
+ ultimately show ?thesis by (rule lbvs.wtl_sound_strong [OF lbvs.intro, OF lbv.intro lbvs_axioms.intro, OF Semilat.intro lbv_axioms.intro])
+qed
+(*>*)
+
+
+lemma (in start_context) wt_lbv_wt_method:
+ assumes lbv: "wt_lbv P C b Ts T\<^sub>r mxs mxl\<^sub>0 xt cert is"
+ shows "\<exists>\<tau>s. wt_method P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt \<tau>s"
+(*<*)
+proof -
+ from lbv have l: "is \<noteq> []" and
+ stab: "b = Static \<or> b = NonStatic" by (auto simp add: wt_lbv_def)
+ moreover
+ from wf lbv C Ts obtain \<tau>s where
+ list: "\<tau>s \<in> list (size is) A" and
+ step: "wt_step r Err step \<tau>s" and
+ start: "OK first \<sqsubseteq>\<^sub>r \<tau>s!0"
+ by (blast dest: wt_lbv_wt_step)
+ from list have [simp]: "size \<tau>s = size is" by simp
+ have "size (map ok_val \<tau>s) = size is" by simp
+ moreover from l have 0: "0 < size \<tau>s" by simp
+ with step obtain \<tau>s0 where "\<tau>s!0 = OK \<tau>s0"
+ by (unfold wt_step_def) blast
+ with start 0 have "wt_start P C b Ts mxl\<^sub>0 (map ok_val \<tau>s)"
+ by (cases b; simp add: wt_start_def JVM_le_Err_conv lesub_def Err.le_def)
+ moreover {
+ from list have "check_types P mxs mxl \<tau>s" by (simp add: check_types_def)
+ also from step have "\<forall>x \<in> set \<tau>s. x \<noteq> Err"
+ by (auto simp add: all_set_conv_all_nth wt_step_def)
+ hence [symmetric]: "map OK (map ok_val \<tau>s) = \<tau>s"
+ by (auto intro!: map_idI)
+ finally have "check_types P mxs mxl (map OK (map ok_val \<tau>s))" .
+ }
+ moreover {
+ note bounded_step
+ moreover from list have "set \<tau>s \<subseteq> A" by simp
+ moreover from step have "wt_err_step (sup_state_opt P) step \<tau>s"
+ by (simp add: wt_err_step_def JVM_le_Err_conv)
+ ultimately have "wt_app_eff (sup_state_opt P) app eff (map ok_val \<tau>s)"
+ by (auto intro: wt_err_imp_wt_app_eff simp add: exec_def states_def)
+ }
+ ultimately have "wt_method P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt (map ok_val \<tau>s)"
+ by (simp add: wt_method_def2 check_types_def del: map_map)
+ thus ?thesis ..
+qed
+(*>*)
+
+
+lemma (in start_context) wt_method_wt_lbv:
+ assumes wt: "wt_method P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt \<tau>s"
+ defines [simp]: "cert \<equiv> mk_cert P mxs T\<^sub>r xt is \<tau>s"
+
+ shows "wt_lbv P C b Ts T\<^sub>r mxs mxl\<^sub>0 xt cert is"
+(*<*)
+proof -
+ let ?\<tau>s = "map OK \<tau>s"
+ let ?cert = "make_cert step ?\<tau>s (OK None)"
+
+ from wt obtain
+ 0: "0 < size is" and
+ size: "size is = size ?\<tau>s" and
+ ck_types: "check_types P mxs mxl ?\<tau>s" and
+ wt_start: "wt_start P C b Ts mxl\<^sub>0 \<tau>s" and
+ app_eff: "wt_app_eff (sup_state_opt P) app eff \<tau>s"
+ by (force simp add: wt_method_def2 check_types_def)
+
+ from wf have "semilat (JVM_SemiType.sl P mxs mxl)" ..
+ hence "semilat (A, r, f)" by (simp add: sl_def2)
+ moreover have "top r Err" by (simp add: JVM_le_Err_conv)
+ moreover have "Err \<in> A" by (simp add: JVM_states_unfold)
+ moreover have "bottom r (OK None)"
+ by (simp add: JVM_le_Err_conv bottom_def lesub_def Err.le_def split: err.split)
+ moreover have "OK None \<in> A" by (simp add: JVM_states_unfold)
+ moreover from wf have "mono r step (size is) A" by (rule step_mono)
+ hence "mono r step (size ?\<tau>s) A" by (simp add: size)
+ moreover from exec_pres_type
+ have "pres_type step (size ?\<tau>s) A" by (simp add: size)
+ moreover
+ from ck_types have \<tau>s_in_A: "set ?\<tau>s \<subseteq> A" by (simp add: check_types_def)
+ hence "\<forall>pc. pc < size ?\<tau>s \<longrightarrow> ?\<tau>s!pc \<in> A \<and> ?\<tau>s!pc \<noteq> Err" by auto
+ moreover from bounded_step
+ have "bounded step (size ?\<tau>s)" by (simp add: size)
+ moreover have "OK None \<noteq> Err" by simp
+ moreover from bounded_step size \<tau>s_in_A app_eff
+ have "wt_err_step (sup_state_opt P) step ?\<tau>s"
+ by (auto intro: wt_app_eff_imp_wt_err simp add: exec_def states_def)
+ hence "wt_step r Err step ?\<tau>s"
+ by (simp add: wt_err_step_def JVM_le_Err_conv)
+ moreover
+ from 0 size have "0 < size \<tau>s" by auto
+ hence "?\<tau>s!0 = OK (\<tau>s!0)" by simp
+ with wt_start have "OK first \<sqsubseteq>\<^sub>r ?\<tau>s!0"
+ by (cases b; clarsimp simp add: wt_start_def lesub_def Err.le_def JVM_le_Err_conv)
+ moreover note first_in_A
+ moreover have "OK first \<noteq> Err" by simp
+ moreover note size
+ ultimately
+ have "wtl_inst_list is ?cert f r Err (OK None) step 0 (OK first) \<noteq> Err"
+ by (rule lbvc.wtl_complete [OF lbvc.intro, OF lbv.intro lbvc_axioms.intro, OF Semilat.intro lbv_axioms.intro])
+ moreover from 0 size have "\<tau>s \<noteq> []" by auto
+ moreover from ck_types have "check_types P mxs mxl ?cert"
+ by (fastforce simp: make_cert_def check_types_def JVM_states_unfold
+ dest!: nth_mem)
+ moreover note 0 size
+ ultimately show ?thesis using staticb
+ by (simp add: wt_lbv_def lbvjvm_def mk_cert_def step_def_exec [symmetric]
+ check_cert_def make_cert_def nth_append)
+qed
+(*>*)
+
+
+
+theorem jvm_lbv_correct:
+ "wt_jvm_prog_lbv P Cert \<Longrightarrow> wf_jvm_prog P"
+(*<*)
+proof -
+ let ?\<Phi> = "\<lambda>C mn. let (C,b,Ts,T\<^sub>r,(mxs,mxl\<^sub>0,is,xt)) = method P C mn in
+ SOME \<tau>s. wt_method P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt \<tau>s"
+
+ let ?A = "\<lambda>P C (mn,b,Ts,T\<^sub>r,(mxs,mxl\<^sub>0,ins,et)). wt_lbv P C b Ts T\<^sub>r mxs mxl\<^sub>0 et (Cert C mn) ins"
+ let ?B = "\<lambda>P C (M,b,Ts,T\<^sub>r,(mxs,mxl\<^sub>0,is,xt)).
+ wt_method P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt (?\<Phi> C M)"
+
+ assume wt: "wt_jvm_prog_lbv P Cert"
+ then have "wf_prog ?A P" by(simp add: wt_jvm_prog_lbv_def)
+ moreover {
+ fix wf_md C M b Ts Ca T m bd
+
+ assume ass1: "wf_prog wf_md P" and sees: "P \<turnstile> Ca sees M, b : Ts\<rightarrow>T = m in Ca" and
+ ass2: "set Ts \<subseteq> types P" and ass3: "bd = (M, b, Ts, T, m)" and
+ ass4: "?A P Ca bd"
+ from ass3 ass4 have stab: "b = Static \<or> b = NonStatic" by (simp add:wt_lbv_def)
+ from ass1 sees ass2 ass3 ass4 have "?B P Ca bd" using sees_method_is_class[OF sees]
+ by (auto dest!: start_context.wt_lbv_wt_method [OF start_context_intro_auxi[OF stab]]
+ intro: someI)
+ }
+ ultimately have "wf_prog ?B P" by(rule wf_prog_lift)
+ hence "wf_jvm_prog\<^bsub>?\<Phi>\<^esub> P" by (simp add: wf_jvm_prog_phi_def)
+ thus ?thesis by (unfold wf_jvm_prog_def) blast
+qed
+(*>*)
+
+theorem jvm_lbv_complete:
+ assumes wt: "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P"
+ shows "wt_jvm_prog_lbv P (prg_cert P \<Phi>)"
+(*<*)
+proof -
+ let ?cert = "prg_cert P \<Phi>"
+ let ?A = "\<lambda>P C (M,b,Ts,T\<^sub>r,(mxs,mxl\<^sub>0,is,xt)).
+ wt_method P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt (\<Phi> C M)"
+ let ?B = "\<lambda>P C (mn,b,Ts,T\<^sub>r,(mxs,mxl\<^sub>0,ins,et)).
+ wt_lbv P C b Ts T\<^sub>r mxs mxl\<^sub>0 et (?cert C mn) ins"
+
+ from wt have "wf_prog ?A P" by(clarsimp simp: wf_jvm_prog_def wf_jvm_prog_phi_def)
+ moreover {
+ fix wf_md C M b Ts Ca T m bd
+ assume ass1: "wf_prog wf_md P" and sees: "P \<turnstile> Ca sees M, b : Ts\<rightarrow>T = m in Ca" and
+ ass2: "set Ts \<subseteq> types P" and ass3: "bd = (M, b, Ts, T, m)" and
+ ass4: "?A P Ca bd"
+ from ass3 ass4 have stab: "b = Static \<or> b = NonStatic" by (simp add:wt_method_def)
+ from ass1 sees ass2 ass3 ass4 have "?B P Ca bd" using sees_method_is_class[OF sees]
+ by (auto simp add: prg_cert_def
+ intro!: start_context.wt_method_wt_lbv start_context_intro_auxi[OF stab])
+ }
+ ultimately have "wf_prog ?B P" by(rule wf_prog_lift)
+ thus "wt_jvm_prog_lbv P (prg_cert P \<Phi>)" by (simp add: wt_jvm_prog_lbv_def)
+qed
+(*>*)
+
+end
diff --git a/thys/JinjaDCI/BV/SemiType.thy b/thys/JinjaDCI/BV/SemiType.thy
--- a/thys/JinjaDCI/BV/SemiType.thy
+++ b/thys/JinjaDCI/BV/SemiType.thy
@@ -1,280 +1,280 @@
-(* Title: Jinja/BV/SemiType.thy
-
- Author: Tobias Nipkow, Gerwin Klein
- Copyright 2000 TUM
-*)
-
-section \<open> The Jinja Type System as a Semilattice \<close>
-
-theory SemiType
-imports "../Common/WellForm" Jinja.Semilattices
-begin
-
-definition super :: "'a prog \<Rightarrow> cname \<Rightarrow> cname"
-where "super P C \<equiv> fst (the (class P C))"
-
-lemma superI:
- "(C,D) \<in> subcls1 P \<Longrightarrow> super P C = D"
- by (unfold super_def) (auto dest: subcls1D)
-
-
-primrec the_Class :: "ty \<Rightarrow> cname"
-where
- "the_Class (Class C) = C"
-
-definition sup :: "'c prog \<Rightarrow> ty \<Rightarrow> ty \<Rightarrow> ty err"
-where
- "sup P T\<^sub>1 T\<^sub>2 \<equiv>
- if is_refT T\<^sub>1 \<and> is_refT T\<^sub>2 then
- OK (if T\<^sub>1 = NT then T\<^sub>2 else
- if T\<^sub>2 = NT then T\<^sub>1 else
- (Class (exec_lub (subcls1 P) (super P) (the_Class T\<^sub>1) (the_Class T\<^sub>2))))
- else
- (if T\<^sub>1 = T\<^sub>2 then OK T\<^sub>1 else Err)"
-
-lemma sup_def':
- "sup P = (\<lambda>T\<^sub>1 T\<^sub>2.
- if is_refT T\<^sub>1 \<and> is_refT T\<^sub>2 then
- OK (if T\<^sub>1 = NT then T\<^sub>2 else
- if T\<^sub>2 = NT then T\<^sub>1 else
- (Class (exec_lub (subcls1 P) (super P) (the_Class T\<^sub>1) (the_Class T\<^sub>2))))
- else
- (if T\<^sub>1 = T\<^sub>2 then OK T\<^sub>1 else Err))"
- by (simp add: sup_def fun_eq_iff)
-
-abbreviation
- subtype :: "'c prog \<Rightarrow> ty \<Rightarrow> ty \<Rightarrow> bool"
- where "subtype P \<equiv> widen P"
-
-definition esl :: "'c prog \<Rightarrow> ty esl"
-where
- "esl P \<equiv> (types P, subtype P, sup P)"
-
-
-(* FIXME: move to wellform *)
-lemma is_class_is_subcls:
- "wf_prog m P \<Longrightarrow> is_class P C = P \<turnstile> C \<preceq>\<^sup>* Object"
-(*<*)by (fastforce simp:is_class_def
- elim: subcls_C_Object converse_rtranclE subcls1I
- dest: subcls1D)
-(*>*)
-
-
-(* FIXME: move to wellform *)
-lemma subcls_antisym:
- "\<lbrakk>wf_prog m P; P \<turnstile> C \<preceq>\<^sup>* D; P \<turnstile> D \<preceq>\<^sup>* C\<rbrakk> \<Longrightarrow> C = D"
- (*<*) by (auto dest: acyclic_subcls1 acyclic_impl_antisym_rtrancl antisymD) (*>*)
-
-(* FIXME: move to wellform *)
-lemma widen_antisym:
- "\<lbrakk> wf_prog m P; P \<turnstile> T \<le> U; P \<turnstile> U \<le> T \<rbrakk> \<Longrightarrow> T = U"
-(*<*)
-apply (cases T)
- apply (cases U)
- apply auto
-apply (cases U)
- apply (auto elim!: subcls_antisym)
-done
-(*>*)
-
-lemma order_widen [intro,simp]:
- "wf_prog m P \<Longrightarrow> order (subtype P) (types P)"
-(*<*)
- apply (unfold Semilat.order_def lesub_def)
- apply (auto intro: widen_trans widen_antisym)
- done
-(*>*)
-
-(* FIXME: move to TypeRel *)
-lemma NT_widen:
- "P \<turnstile> NT \<le> T = (T = NT \<or> (\<exists>C. T = Class C))"
-(*<*) by (cases T) auto (*>*)
-
-(* FIXME: move to TypeRel *)
-lemma Class_widen2: "P \<turnstile> Class C \<le> T = (\<exists>D. T = Class D \<and> P \<turnstile> C \<preceq>\<^sup>* D)"
-(*<*) by (cases T) auto (*>*)
-
-lemma wf_converse_subcls1_impl_acc_subtype:
- "wf ((subcls1 P)^-1) \<Longrightarrow> acc (subtype P)"
-(*<*)
-apply (unfold Semilat.acc_def lesssub_def)
-apply (drule_tac p = "(subcls1 P)^-1 - Id" in wf_subset)
- apply blast
-apply (drule wf_trancl)
-apply (simp add: wf_eq_minimal)
-apply clarify
-apply (unfold lesub_def)
-apply (rename_tac M T)
-apply (case_tac "\<exists>C. Class C \<in> M")
- prefer 2
- apply (case_tac T)
- apply fastforce
- apply fastforce
- apply fastforce
- apply simp
- apply (rule_tac x = NT in bexI)
- apply (rule allI)
- apply (rule impI, erule conjE)
- apply (clarsimp simp add: NT_widen)
- apply simp
- apply clarsimp
-apply (erule_tac x = "{C. Class C : M}" in allE)
-apply auto
-apply (rename_tac D)
-apply (rule_tac x = "Class D" in bexI)
- prefer 2
- apply assumption
-apply clarify
-apply (clarsimp simp: Class_widen2)
-apply (insert rtrancl_r_diff_Id [symmetric, of "subcls1 P"])
-apply simp
-apply (erule rtranclE)
- apply blast
-apply (drule rtrancl_converseI)
-apply (subgoal_tac "((subcls1 P)-Id)^-1 = ((subcls1 P)^-1 - Id)")
- prefer 2
- apply blast
-apply simp
-apply (blast intro: rtrancl_into_trancl2)
-done
-(*>*)
-
-lemma wf_subtype_acc [intro, simp]:
- "wf_prog wf_mb P \<Longrightarrow> acc (subtype P)"
-(*<*) by (rule wf_converse_subcls1_impl_acc_subtype, rule wf_subcls1) (*>*)
-
-lemma exec_lub_refl [simp]: "exec_lub r f T T = T"
-(*<*) by (simp add: exec_lub_def while_unfold) (*>*)
-
-lemma closed_err_types:
- "wf_prog wf_mb P \<Longrightarrow> closed (err (types P)) (lift2 (sup P))"
-(*<*)
- apply (unfold closed_def plussub_def lift2_def sup_def')
- apply (frule acyclic_subcls1)
- apply (frule single_valued_subcls1)
- apply (auto simp: is_type_def is_refT_def is_class_is_subcls split: err.split ty.splits)
- apply (blast dest!: is_lub_exec_lub is_lubD is_ubD intro!: is_ubI superI)
- done
-(*>*)
-
-
-lemma sup_subtype_greater:
- "\<lbrakk> wf_prog wf_mb P; is_type P t1; is_type P t2; sup P t1 t2 = OK s \<rbrakk>
- \<Longrightarrow> subtype P t1 s \<and> subtype P t2 s"
-(*<*)
-proof -
- assume wf_prog: "wf_prog wf_mb P"
-
- { fix c1 c2
- assume is_class: "is_class P c1" "is_class P c2"
- with wf_prog
- obtain
- "P \<turnstile> c1 \<preceq>\<^sup>* Object"
- "P \<turnstile> c2 \<preceq>\<^sup>* Object"
- by (blast intro: subcls_C_Object)
- with single_valued_subcls1[OF wf_prog]
- obtain u where
- "is_lub ((subcls1 P)^* ) c1 c2 u"
- by (blast dest: single_valued_has_lubs)
- moreover
- note acyclic_subcls1[OF wf_prog]
- moreover
- have "\<forall>x y. (x, y) \<in> subcls1 P \<longrightarrow> super P x = y"
- by (blast intro: superI)
- ultimately
- have "P \<turnstile> c1 \<preceq>\<^sup>* exec_lub (subcls1 P) (super P) c1 c2 \<and>
- P \<turnstile> c2 \<preceq>\<^sup>* exec_lub (subcls1 P) (super P) c1 c2"
- by (simp add: exec_lub_conv) (blast dest: is_lubD is_ubD)
- } note this [simp]
-
- assume "is_type P t1" "is_type P t2" "sup P t1 t2 = OK s"
- thus ?thesis
- apply (unfold sup_def)
- apply (cases s)
- apply (auto simp add: is_refT_def split: if_split_asm)
- done
-qed
-(*>*)
-
-lemma sup_subtype_smallest:
- "\<lbrakk> wf_prog wf_mb P; is_type P a; is_type P b; is_type P c;
- subtype P a c; subtype P b c; sup P a b = OK d \<rbrakk>
- \<Longrightarrow> subtype P d c"
-(*<*)
-proof -
- assume wf_prog: "wf_prog wf_mb P"
-
- { fix c1 c2 D
- assume is_class: "is_class P c1" "is_class P c2"
- assume le: "P \<turnstile> c1 \<preceq>\<^sup>* D" "P \<turnstile> c2 \<preceq>\<^sup>* D"
- from wf_prog is_class
- obtain
- "P \<turnstile> c1 \<preceq>\<^sup>* Object"
- "P \<turnstile> c2 \<preceq>\<^sup>* Object"
- by (blast intro: subcls_C_Object)
- with single_valued_subcls1[OF wf_prog]
- obtain u where
- lub: "is_lub ((subcls1 P)^* ) c1 c2 u"
- by (blast dest: single_valued_has_lubs)
- with acyclic_subcls1[OF wf_prog]
- have "exec_lub (subcls1 P) (super P) c1 c2 = u"
- by (blast intro: superI exec_lub_conv)
- moreover
- from lub le
- have "P \<turnstile> u \<preceq>\<^sup>* D"
- by (simp add: is_lub_def is_ub_def)
- ultimately
- have "P \<turnstile> exec_lub (subcls1 P) (super P) c1 c2 \<preceq>\<^sup>* D"
- by blast
- } note this [intro]
-
- have [dest!]:
- "\<And>C T. P \<turnstile> Class C \<le> T \<Longrightarrow> \<exists>D. T=Class D \<and> P \<turnstile> C \<preceq>\<^sup>* D"
- by (frule Class_widen, auto)
-
- assume "is_type P a" "is_type P b" "is_type P c"
- "subtype P a c" "subtype P b c" "sup P a b = OK d"
- thus ?thesis
- by (auto simp add: sup_def is_refT_def
- split: if_split_asm)
-qed
-(*>*)
-
-lemma sup_exists:
- "\<lbrakk> subtype P a c; subtype P b c \<rbrakk> \<Longrightarrow> \<exists>T. sup P a b = OK T"
-(*<*)
-apply (unfold sup_def)
-apply (cases b)
-apply auto
-apply (cases a)
-apply auto
-apply (cases a)
-apply auto
-done
-(*>*)
-
-lemma err_semilat_JType_esl:
- "wf_prog wf_mb P \<Longrightarrow> err_semilat (esl P)"
-(*<*)
-proof -
- assume wf_prog: "wf_prog wf_mb P"
- hence "order (subtype P) (types P)"..
- moreover from wf_prog
- have "closed (err (types P)) (lift2 (sup P))"
- by (rule closed_err_types)
- moreover
- from wf_prog have
- "(\<forall>x\<in>err (types P). \<forall>y\<in>err (types P). x \<sqsubseteq>\<^bsub>Err.le (subtype P)\<^esub> x \<squnion>\<^bsub>lift2 (sup P)\<^esub> y) \<and>
- (\<forall>x\<in>err (types P). \<forall>y\<in>err (types P). y \<sqsubseteq>\<^bsub>Err.le (subtype P)\<^esub> x \<squnion>\<^bsub>lift2 (sup P)\<^esub> y)"
- by (auto simp add: lesub_def plussub_def Err.le_def lift2_def sup_subtype_greater split: err.split)
- moreover from wf_prog have
- "\<forall>x\<in>err (types P). \<forall>y\<in>err (types P). \<forall>z\<in>err (types P).
- x \<sqsubseteq>\<^bsub>Err.le (subtype P)\<^esub> z \<and> y \<sqsubseteq>\<^bsub>Err.le (subtype P)\<^esub> z \<longrightarrow> x \<squnion>\<^bsub>lift2 (sup P)\<^esub> y \<sqsubseteq>\<^bsub>Err.le (subtype P)\<^esub> z"
- by (unfold lift2_def plussub_def lesub_def Err.le_def)
- (auto intro: sup_subtype_smallest dest:sup_exists split: err.split)
- ultimately show ?thesis by (simp add: esl_def semilat_def sl_def Err.sl_def)
-qed
-(*>*)
-
-
-end
+(* Title: Jinja/BV/SemiType.thy
+
+ Author: Tobias Nipkow, Gerwin Klein
+ Copyright 2000 TUM
+*)
+
+section \<open> The Jinja Type System as a Semilattice \<close>
+
+theory SemiType
+imports "../Common/WellForm" Jinja.Semilattices
+begin
+
+definition super :: "'a prog \<Rightarrow> cname \<Rightarrow> cname"
+where "super P C \<equiv> fst (the (class P C))"
+
+lemma superI:
+ "(C,D) \<in> subcls1 P \<Longrightarrow> super P C = D"
+ by (unfold super_def) (auto dest: subcls1D)
+
+
+primrec the_Class :: "ty \<Rightarrow> cname"
+where
+ "the_Class (Class C) = C"
+
+definition sup :: "'c prog \<Rightarrow> ty \<Rightarrow> ty \<Rightarrow> ty err"
+where
+ "sup P T\<^sub>1 T\<^sub>2 \<equiv>
+ if is_refT T\<^sub>1 \<and> is_refT T\<^sub>2 then
+ OK (if T\<^sub>1 = NT then T\<^sub>2 else
+ if T\<^sub>2 = NT then T\<^sub>1 else
+ (Class (exec_lub (subcls1 P) (super P) (the_Class T\<^sub>1) (the_Class T\<^sub>2))))
+ else
+ (if T\<^sub>1 = T\<^sub>2 then OK T\<^sub>1 else Err)"
+
+lemma sup_def':
+ "sup P = (\<lambda>T\<^sub>1 T\<^sub>2.
+ if is_refT T\<^sub>1 \<and> is_refT T\<^sub>2 then
+ OK (if T\<^sub>1 = NT then T\<^sub>2 else
+ if T\<^sub>2 = NT then T\<^sub>1 else
+ (Class (exec_lub (subcls1 P) (super P) (the_Class T\<^sub>1) (the_Class T\<^sub>2))))
+ else
+ (if T\<^sub>1 = T\<^sub>2 then OK T\<^sub>1 else Err))"
+ by (simp add: sup_def fun_eq_iff)
+
+abbreviation
+ subtype :: "'c prog \<Rightarrow> ty \<Rightarrow> ty \<Rightarrow> bool"
+ where "subtype P \<equiv> widen P"
+
+definition esl :: "'c prog \<Rightarrow> ty esl"
+where
+ "esl P \<equiv> (types P, subtype P, sup P)"
+
+
+(* FIXME: move to wellform *)
+lemma is_class_is_subcls:
+ "wf_prog m P \<Longrightarrow> is_class P C = P \<turnstile> C \<preceq>\<^sup>* Object"
+(*<*)by (fastforce simp:is_class_def
+ elim: subcls_C_Object converse_rtranclE subcls1I
+ dest: subcls1D)
+(*>*)
+
+
+(* FIXME: move to wellform *)
+lemma subcls_antisym:
+ "\<lbrakk>wf_prog m P; P \<turnstile> C \<preceq>\<^sup>* D; P \<turnstile> D \<preceq>\<^sup>* C\<rbrakk> \<Longrightarrow> C = D"
+ (*<*) by (auto dest: acyclic_subcls1 acyclic_impl_antisym_rtrancl antisymD) (*>*)
+
+(* FIXME: move to wellform *)
+lemma widen_antisym:
+ "\<lbrakk> wf_prog m P; P \<turnstile> T \<le> U; P \<turnstile> U \<le> T \<rbrakk> \<Longrightarrow> T = U"
+(*<*)
+apply (cases T)
+ apply (cases U)
+ apply auto
+apply (cases U)
+ apply (auto elim!: subcls_antisym)
+done
+(*>*)
+
+lemma order_widen [intro,simp]:
+ "wf_prog m P \<Longrightarrow> order (subtype P) (types P)"
+(*<*)
+ apply (unfold Semilat.order_def lesub_def)
+ apply (auto intro: widen_trans widen_antisym)
+ done
+(*>*)
+
+(* FIXME: move to TypeRel *)
+lemma NT_widen:
+ "P \<turnstile> NT \<le> T = (T = NT \<or> (\<exists>C. T = Class C))"
+(*<*) by (cases T) auto (*>*)
+
+(* FIXME: move to TypeRel *)
+lemma Class_widen2: "P \<turnstile> Class C \<le> T = (\<exists>D. T = Class D \<and> P \<turnstile> C \<preceq>\<^sup>* D)"
+(*<*) by (cases T) auto (*>*)
+
+lemma wf_converse_subcls1_impl_acc_subtype:
+ "wf ((subcls1 P)^-1) \<Longrightarrow> acc (subtype P)"
+(*<*)
+apply (unfold Semilat.acc_def lesssub_def)
+apply (drule_tac p = "(subcls1 P)^-1 - Id" in wf_subset)
+ apply blast
+apply (drule wf_trancl)
+apply (simp add: wf_eq_minimal)
+apply clarify
+apply (unfold lesub_def)
+apply (rename_tac M T)
+apply (case_tac "\<exists>C. Class C \<in> M")
+ prefer 2
+ apply (case_tac T)
+ apply fastforce
+ apply fastforce
+ apply fastforce
+ apply simp
+ apply (rule_tac x = NT in bexI)
+ apply (rule allI)
+ apply (rule impI, erule conjE)
+ apply (clarsimp simp add: NT_widen)
+ apply simp
+ apply clarsimp
+apply (erule_tac x = "{C. Class C : M}" in allE)
+apply auto
+apply (rename_tac D)
+apply (rule_tac x = "Class D" in bexI)
+ prefer 2
+ apply assumption
+apply clarify
+apply (clarsimp simp: Class_widen2)
+apply (insert rtrancl_r_diff_Id [symmetric, of "subcls1 P"])
+apply simp
+apply (erule rtranclE)
+ apply blast
+apply (drule rtrancl_converseI)
+apply (subgoal_tac "((subcls1 P)-Id)^-1 = ((subcls1 P)^-1 - Id)")
+ prefer 2
+ apply blast
+apply simp
+apply (blast intro: rtrancl_into_trancl2)
+done
+(*>*)
+
+lemma wf_subtype_acc [intro, simp]:
+ "wf_prog wf_mb P \<Longrightarrow> acc (subtype P)"
+(*<*) by (rule wf_converse_subcls1_impl_acc_subtype, rule wf_subcls1) (*>*)
+
+lemma exec_lub_refl [simp]: "exec_lub r f T T = T"
+(*<*) by (simp add: exec_lub_def while_unfold) (*>*)
+
+lemma closed_err_types:
+ "wf_prog wf_mb P \<Longrightarrow> closed (err (types P)) (lift2 (sup P))"
+(*<*)
+ apply (unfold closed_def plussub_def lift2_def sup_def')
+ apply (frule acyclic_subcls1)
+ apply (frule single_valued_subcls1)
+ apply (auto simp: is_type_def is_refT_def is_class_is_subcls split: err.split ty.splits)
+ apply (blast dest!: is_lub_exec_lub is_lubD is_ubD intro!: is_ubI superI)
+ done
+(*>*)
+
+
+lemma sup_subtype_greater:
+ "\<lbrakk> wf_prog wf_mb P; is_type P t1; is_type P t2; sup P t1 t2 = OK s \<rbrakk>
+ \<Longrightarrow> subtype P t1 s \<and> subtype P t2 s"
+(*<*)
+proof -
+ assume wf_prog: "wf_prog wf_mb P"
+
+ { fix c1 c2
+ assume is_class: "is_class P c1" "is_class P c2"
+ with wf_prog
+ obtain
+ "P \<turnstile> c1 \<preceq>\<^sup>* Object"
+ "P \<turnstile> c2 \<preceq>\<^sup>* Object"
+ by (blast intro: subcls_C_Object)
+ with single_valued_subcls1[OF wf_prog]
+ obtain u where
+ "is_lub ((subcls1 P)^* ) c1 c2 u"
+ by (blast dest: single_valued_has_lubs)
+ moreover
+ note acyclic_subcls1[OF wf_prog]
+ moreover
+ have "\<forall>x y. (x, y) \<in> subcls1 P \<longrightarrow> super P x = y"
+ by (blast intro: superI)
+ ultimately
+ have "P \<turnstile> c1 \<preceq>\<^sup>* exec_lub (subcls1 P) (super P) c1 c2 \<and>
+ P \<turnstile> c2 \<preceq>\<^sup>* exec_lub (subcls1 P) (super P) c1 c2"
+ by (simp add: exec_lub_conv) (blast dest: is_lubD is_ubD)
+ } note this [simp]
+
+ assume "is_type P t1" "is_type P t2" "sup P t1 t2 = OK s"
+ thus ?thesis
+ apply (unfold sup_def)
+ apply (cases s)
+ apply (auto simp add: is_refT_def split: if_split_asm)
+ done
+qed
+(*>*)
+
+lemma sup_subtype_smallest:
+ "\<lbrakk> wf_prog wf_mb P; is_type P a; is_type P b; is_type P c;
+ subtype P a c; subtype P b c; sup P a b = OK d \<rbrakk>
+ \<Longrightarrow> subtype P d c"
+(*<*)
+proof -
+ assume wf_prog: "wf_prog wf_mb P"
+
+ { fix c1 c2 D
+ assume is_class: "is_class P c1" "is_class P c2"
+ assume le: "P \<turnstile> c1 \<preceq>\<^sup>* D" "P \<turnstile> c2 \<preceq>\<^sup>* D"
+ from wf_prog is_class
+ obtain
+ "P \<turnstile> c1 \<preceq>\<^sup>* Object"
+ "P \<turnstile> c2 \<preceq>\<^sup>* Object"
+ by (blast intro: subcls_C_Object)
+ with single_valued_subcls1[OF wf_prog]
+ obtain u where
+ lub: "is_lub ((subcls1 P)^* ) c1 c2 u"
+ by (blast dest: single_valued_has_lubs)
+ with acyclic_subcls1[OF wf_prog]
+ have "exec_lub (subcls1 P) (super P) c1 c2 = u"
+ by (blast intro: superI exec_lub_conv)
+ moreover
+ from lub le
+ have "P \<turnstile> u \<preceq>\<^sup>* D"
+ by (simp add: is_lub_def is_ub_def)
+ ultimately
+ have "P \<turnstile> exec_lub (subcls1 P) (super P) c1 c2 \<preceq>\<^sup>* D"
+ by blast
+ } note this [intro]
+
+ have [dest!]:
+ "\<And>C T. P \<turnstile> Class C \<le> T \<Longrightarrow> \<exists>D. T=Class D \<and> P \<turnstile> C \<preceq>\<^sup>* D"
+ by (frule Class_widen, auto)
+
+ assume "is_type P a" "is_type P b" "is_type P c"
+ "subtype P a c" "subtype P b c" "sup P a b = OK d"
+ thus ?thesis
+ by (auto simp add: sup_def is_refT_def
+ split: if_split_asm)
+qed
+(*>*)
+
+lemma sup_exists:
+ "\<lbrakk> subtype P a c; subtype P b c \<rbrakk> \<Longrightarrow> \<exists>T. sup P a b = OK T"
+(*<*)
+apply (unfold sup_def)
+apply (cases b)
+apply auto
+apply (cases a)
+apply auto
+apply (cases a)
+apply auto
+done
+(*>*)
+
+lemma err_semilat_JType_esl:
+ "wf_prog wf_mb P \<Longrightarrow> err_semilat (esl P)"
+(*<*)
+proof -
+ assume wf_prog: "wf_prog wf_mb P"
+ hence "order (subtype P) (types P)"..
+ moreover from wf_prog
+ have "closed (err (types P)) (lift2 (sup P))"
+ by (rule closed_err_types)
+ moreover
+ from wf_prog have
+ "(\<forall>x\<in>err (types P). \<forall>y\<in>err (types P). x \<sqsubseteq>\<^bsub>Err.le (subtype P)\<^esub> x \<squnion>\<^bsub>lift2 (sup P)\<^esub> y) \<and>
+ (\<forall>x\<in>err (types P). \<forall>y\<in>err (types P). y \<sqsubseteq>\<^bsub>Err.le (subtype P)\<^esub> x \<squnion>\<^bsub>lift2 (sup P)\<^esub> y)"
+ by (auto simp add: lesub_def plussub_def Err.le_def lift2_def sup_subtype_greater split: err.split)
+ moreover from wf_prog have
+ "\<forall>x\<in>err (types P). \<forall>y\<in>err (types P). \<forall>z\<in>err (types P).
+ x \<sqsubseteq>\<^bsub>Err.le (subtype P)\<^esub> z \<and> y \<sqsubseteq>\<^bsub>Err.le (subtype P)\<^esub> z \<longrightarrow> x \<squnion>\<^bsub>lift2 (sup P)\<^esub> y \<sqsubseteq>\<^bsub>Err.le (subtype P)\<^esub> z"
+ by (unfold lift2_def plussub_def lesub_def Err.le_def)
+ (auto intro: sup_subtype_smallest dest:sup_exists split: err.split)
+ ultimately show ?thesis by (simp add: esl_def semilat_def sl_def Err.sl_def)
+qed
+(*>*)
+
+
+end
diff --git a/thys/JinjaDCI/BV/TF_JVM.thy b/thys/JinjaDCI/BV/TF_JVM.thy
--- a/thys/JinjaDCI/BV/TF_JVM.thy
+++ b/thys/JinjaDCI/BV/TF_JVM.thy
@@ -1,355 +1,355 @@
-(* Title: JinjaDCI/BV/TF_JVM.thy
-
- Author: Tobias Nipkow, Gerwin Klein, Susannah Mansky
- Copyright 2000 TUM, 2019-20 UIUC
-
- Based on the Jinja theory BV/TF_JVM.thy by Tobias Nipkow and Gerwin Klein
-*)
-
-section \<open> The Typing Framework for the JVM \label{sec:JVM} \<close>
-
-theory TF_JVM
-imports Jinja.Typing_Framework_err EffectMono BVSpec
-begin
-
-definition exec :: "jvm_prog \<Rightarrow> nat \<Rightarrow> ty \<Rightarrow> ex_table \<Rightarrow> instr list \<Rightarrow> ty\<^sub>i' err step_type"
-where
- "exec G maxs rT et bs \<equiv>
- err_step (size bs) (\<lambda>pc. app (bs!pc) G maxs rT pc (size bs) et)
- (\<lambda>pc. eff (bs!pc) G pc et)"
-
-locale JVM_sl =
- fixes P :: jvm_prog and mxs and mxl\<^sub>0 and n
- fixes b and Ts :: "ty list" and "is" and xt and T\<^sub>r
-
- fixes mxl and A and r and f and app and eff and step
- defines [simp]: "mxl \<equiv> (case b of Static \<Rightarrow> 0 | NonStatic \<Rightarrow> 1)+size Ts+mxl\<^sub>0"
- defines [simp]: "A \<equiv> states P mxs mxl"
- defines [simp]: "r \<equiv> JVM_SemiType.le P mxs mxl"
- defines [simp]: "f \<equiv> JVM_SemiType.sup P mxs mxl"
-
- defines [simp]: "app \<equiv> \<lambda>pc. Effect.app (is!pc) P mxs T\<^sub>r pc (size is) xt"
- defines [simp]: "eff \<equiv> \<lambda>pc. Effect.eff (is!pc) P pc xt"
- defines [simp]: "step \<equiv> err_step (size is) app eff"
-
- defines [simp]: "n \<equiv> size is"
- assumes staticb: "b = Static \<or> b = NonStatic"
-
-locale start_context = JVM_sl +
- fixes p and C
-
- assumes wf: "wf_prog p P"
- assumes C: "is_class P C"
- assumes Ts: "set Ts \<subseteq> types P"
-
-
- fixes first :: ty\<^sub>i' and start
- defines [simp]:
- "first \<equiv> Some ([],(case b of Static \<Rightarrow> [] | NonStatic \<Rightarrow> [OK (Class C)]) @ map OK Ts @ replicate mxl\<^sub>0 Err)"
- defines [simp]:
- "start \<equiv> (OK first) # replicate (size is - 1) (OK None)"
-thm start_context.intro
-
-lemma start_context_intro_auxi:
- fixes P b Ts p C
- assumes "b = Static \<or> b = NonStatic "
- and "wf_prog p P"
- and "is_class P C"
- and "set Ts \<subseteq> types P"
- shows " start_context P b Ts p C"
- using start_context.intro[OF JVM_sl.intro] start_context_axioms_def assms by auto
-
-subsection \<open> Connecting JVM and Framework \<close>
-
-lemma (in start_context) semi: "semilat (A, r, f)"
- apply (insert semilat_JVM[OF wf])
- apply (unfold A_def r_def f_def JVM_SemiType.le_def JVM_SemiType.sup_def states_def)
- apply auto
- done
-
-
-lemma (in JVM_sl) step_def_exec: "step \<equiv> exec P mxs T\<^sub>r xt is"
- by (simp add: exec_def)
-
-lemma special_ex_swap_lemma [iff]:
- "(? X. (? n. X = A n & P n) & Q X) = (? n. Q(A n) & P n)"
- by blast
-
-lemma ex_in_list [iff]:
- "(\<exists>n. ST \<in> list n A \<and> n \<le> mxs) = (set ST \<subseteq> A \<and> size ST \<le> mxs)"
- by (unfold list_def) auto
-
-lemma singleton_list:
- "(\<exists>n. [Class C] \<in> list n (types P) \<and> n \<le> mxs) = (is_class P C \<and> 0 < mxs)"
- by auto
-
-lemma set_drop_subset:
- "set xs \<subseteq> A \<Longrightarrow> set (drop n xs) \<subseteq> A"
- by (auto dest: in_set_dropD)
-
-lemma Suc_minus_minus_le:
- "n < mxs \<Longrightarrow> Suc (n - (n - b)) \<le> mxs"
- by arith
-
-lemma in_listE:
- "\<lbrakk> xs \<in> list n A; \<lbrakk>size xs = n; set xs \<subseteq> A\<rbrakk> \<Longrightarrow> P \<rbrakk> \<Longrightarrow> P"
- by (unfold list_def) blast
-
-declare is_relevant_entry_def [simp]
-declare set_drop_subset [simp]
-
-theorem (in start_context) exec_pres_type:
- "pres_type step (size is) A"
-(*<*)
-proof -
- let ?n = "size is" and ?app = app and ?step = eff
- let ?mxl = "(case b of Static \<Rightarrow> 0 | NonStatic \<Rightarrow> 1) + length Ts + mxl\<^sub>0"
- let ?A = "opt((Union {list n (types P) |n. n <= mxs}) \<times>
- list ?mxl (err(types P)))"
- have "pres_type (err_step ?n ?app ?step) ?n (err ?A)"
- proof(rule pres_type_lift)
- have "\<And>s pc pc' s'. s\<in>?A \<Longrightarrow> pc < ?n \<Longrightarrow> ?app pc s
- \<Longrightarrow> (pc', s')\<in>set (?step pc s) \<Longrightarrow> s' \<in> ?A"
- proof -
- fix s pc pc' s'
- assume asms: "s\<in>?A" "pc < ?n" "?app pc s" "(pc', s')\<in>set (?step pc s)"
- show "s' \<in> ?A"
- proof(cases s)
- case None
- then show ?thesis using asms by (fastforce dest: effNone)
- next
- case (Some ab)
- then show ?thesis using asms proof(cases "is!pc")
- case Load
- then show ?thesis using asms
- by (fastforce simp: Effect.app_def xcpt_app_def Effect.eff_def
- xcpt_eff_def norm_eff_def
- dest: listE_nth_in)
- next
- case Push
- then show ?thesis using asms Some
- by (fastforce simp: Effect.app_def xcpt_app_def Effect.eff_def
- xcpt_eff_def norm_eff_def typeof_lit_is_type)
- next
- case Getfield
- then show ?thesis using asms wf
- by (fastforce simp: Effect.app_def xcpt_app_def Effect.eff_def
- xcpt_eff_def norm_eff_def
- dest: sees_field_is_type)
- next
- case Getstatic
- then show ?thesis using asms wf
- by (fastforce simp: Effect.app_def xcpt_app_def Effect.eff_def
- xcpt_eff_def norm_eff_def
- dest: sees_field_is_type)
- next
- case (Invoke M n)
- obtain a b where [simp]: "s = \<lfloor>(a,b)\<rfloor>" using Some asms(1) by blast
- show ?thesis
- proof(cases "a!n = NT")
- case True
- then show ?thesis using Invoke asms wf
- by (fastforce simp: Effect.app_def xcpt_app_def Effect.eff_def
- xcpt_eff_def norm_eff_def)
- next
- case False
- have "(pc', s') \<in> set (norm_eff (Invoke M n) P pc (a, b)) \<or>
- (pc', s') \<in> set (xcpt_eff (Invoke M n) P pc (a, b) xt)"
- using Invoke asms(4) by (simp add: Effect.eff_def)
- then show ?thesis proof(rule disjE)
- assume "(pc', s') \<in> set (xcpt_eff (Invoke M n) P pc (a, b) xt)"
- then show ?thesis using Invoke asms(1-3)
- by (fastforce simp: Effect.app_def xcpt_app_def xcpt_eff_def)
- next
- assume norm: "(pc', s') \<in> set (norm_eff (Invoke M n) P pc (a, b))"
- also have "Suc (length a - Suc n) \<le> mxs" using Invoke asms(1,3)
- by (simp add: Effect.app_def xcpt_app_def) arith
- ultimately show ?thesis using False Invoke asms(1-3) wf
- by (auto simp: Effect.app_def xcpt_app_def norm_eff_def wf_mdecl_def
- dest!: sees_wf_mdecl)
- qed
- qed
- next
- case (Invokestatic C M n)
- obtain a b where [simp]: "s = \<lfloor>(a,b)\<rfloor>" using Some asms(1) by blast
- have "(pc', s') \<in> set (norm_eff (Invokestatic C M n) P pc (a, b)) \<or>
- (pc', s') \<in> set (xcpt_eff (Invokestatic C M n) P pc (a, b) xt)"
- using Invokestatic asms(4) by (simp add: Effect.eff_def)
- then show ?thesis proof(rule disjE)
- assume "(pc', s') \<in> set (xcpt_eff (Invokestatic C M n) P pc (a, b) xt)"
- then show ?thesis using Invokestatic asms(1-3)
- by (fastforce simp: Effect.app_def xcpt_app_def xcpt_eff_def)
- next
- assume norm: "(pc', s') \<in> set (norm_eff (Invokestatic C M n) P pc (a, b))"
- also have "Suc (length a - Suc n) \<le> mxs" using Invokestatic asms(1,3)
- by (simp add: Effect.app_def xcpt_app_def) arith
- ultimately show ?thesis using Invokestatic asms(1-3) wf
- by (auto simp: Effect.app_def xcpt_app_def norm_eff_def wf_mdecl_def
- dest!: sees_wf_mdecl)
- qed
- qed (fastforce simp: Effect.app_def xcpt_app_def Effect.eff_def
- xcpt_eff_def norm_eff_def)+
- qed
- qed
- then show "\<forall>s\<in>?A. \<forall>p. p < ?n \<longrightarrow> ?app p s \<longrightarrow> (\<forall>(q, s')\<in>set (?step p s). s' \<in> ?A)"
- by clarsimp
- qed
- then show ?thesis by (simp add: JVM_states_unfold)
-qed
-(*>*)
-
-declare is_relevant_entry_def [simp del]
-declare set_drop_subset [simp del]
-
-lemma lesubstep_type_simple:
- "xs [\<sqsubseteq>\<^bsub>Product.le (=) r\<^esub>] ys \<Longrightarrow> set xs {\<sqsubseteq>\<^bsub>r\<^esub>} set ys"
-(*<*)
-proof -
- assume assm: "xs [\<sqsubseteq>\<^bsub>Product.le (=) r\<^esub>] ys"
- have "\<And>a b i y. (a, b) = xs ! i \<Longrightarrow> i < length xs
- \<Longrightarrow> \<exists>\<tau>'. (\<exists>i. (a, \<tau>') = ys ! i \<and> i < length xs) \<and> b \<sqsubseteq>\<^bsub>r\<^esub> \<tau>'"
- proof -
- fix a b i assume ith: "(a, b) = xs ! i" and len: "i < length xs"
- obtain \<tau> where "ys ! i = (a, \<tau>) \<and> r b \<tau>"
- using le_listD[OF assm len] ith
- by (clarsimp simp: lesub_def Product.le_def)
- then have "(a, \<tau>) = ys ! i \<and> b \<sqsubseteq>\<^bsub>r\<^esub> \<tau>"
- by (clarsimp simp: lesub_def)
- with len show "\<exists>\<tau>'. (\<exists>i. (a, \<tau>') = ys ! i \<and> i < length xs) \<and> b \<sqsubseteq>\<^bsub>r\<^esub> \<tau>'"
- by fastforce
- qed
- then show "set xs {\<sqsubseteq>\<^bsub>r\<^esub>} set ys" using assm
- by (clarsimp simp: lesubstep_type_def set_conv_nth)
-qed
-(*>*)
-
-declare is_relevant_entry_def [simp del]
-
-
-lemma conjI2: "\<lbrakk> A; A \<Longrightarrow> B \<rbrakk> \<Longrightarrow> A \<and> B" by blast
-
-lemma (in JVM_sl) eff_mono:
-assumes wf: "wf_prog p P" and "pc < length is" and
- lesub: "s \<sqsubseteq>\<^bsub>sup_state_opt P\<^esub> t" and app: "app pc t"
-shows "set (eff pc s) {\<sqsubseteq>\<^bsub>sup_state_opt P\<^esub>} set (eff pc t)"
-(*<*)
-proof(cases t)
- case None then show ?thesis using lesub
- by (simp add: Effect.eff_def lesub_def)
-next
- case tSome: (Some a)
- show ?thesis proof(cases s)
- case None then show ?thesis using lesub
- by (simp add: Effect.eff_def lesub_def)
- next
- case (Some b)
- let ?norm = "\<lambda>x. norm_eff (is ! pc) P pc x"
- let ?xcpt = "\<lambda>x. xcpt_eff (is ! pc) P pc x xt"
- let ?r = "Product.le (=) (sup_state_opt P)"
- let ?\<tau>' = "\<lfloor>eff\<^sub>i (is ! pc, P, a)\<rfloor>"
- {
- fix x assume xb: "x \<in> set (succs (is ! pc) b pc)"
- then have appi: "app\<^sub>i (is ! pc, P, pc, mxs, T\<^sub>r, a)" and
- bia: "P \<turnstile> b \<le>\<^sub>i a" and appa: "app pc \<lfloor>a\<rfloor>"
- using lesub app tSome Some by (auto simp add: lesub_def Effect.app_def)
- have xa: "x \<in> set (succs (is ! pc) a pc)"
- using xb succs_mono[OF wf appi bia] by auto
- then have "(x, ?\<tau>') \<in> (\<lambda>pc'. (pc', ?\<tau>')) ` set (succs (is ! pc) a pc)"
- by (rule imageI)
- moreover have "P \<turnstile> \<lfloor>eff\<^sub>i (is ! pc, P, b)\<rfloor> \<le>' ?\<tau>'"
- using xb xa eff\<^sub>i_mono[OF wf bia] appa by fastforce
- ultimately have "\<exists>\<tau>'. (x, \<tau>') \<in> (\<lambda>pc'. (pc', \<lfloor>eff\<^sub>i (is ! pc, P, a)\<rfloor>)) ` set (succs (is ! pc) a pc) \<and>
- P \<turnstile> \<lfloor>eff\<^sub>i (is ! pc, P, b)\<rfloor> \<le>' \<tau>'" by blast
- }
- then have norm: "set (?norm b) {\<sqsubseteq>\<^bsub>sup_state_opt P\<^esub>} set (?norm a)"
- using tSome Some by (clarsimp simp: norm_eff_def lesubstep_type_def lesub_def)
- obtain a1 b1 a2 b2 where a: "a = (a1, b1)" and b: "b = (a2, b2)"
- using tSome Some by fastforce
- then have a12: "size a2 = size a1" using lesub tSome Some
- by (clarsimp simp: lesub_def list_all2_lengthD)
- have "length (?xcpt b) = length (?xcpt a)"
- by (simp add: xcpt_eff_def split_beta)
- moreover have "\<And>n. n < length (?xcpt b) \<Longrightarrow> (?xcpt b) ! n \<sqsubseteq>\<^bsub>?r\<^esub> (?xcpt a) ! n"
- using lesub tSome Some a b a12
- by (simp add: xcpt_eff_def split_beta fun_of_def) (simp add: lesub_def)
- ultimately have "?xcpt b [\<sqsubseteq>\<^bsub>?r\<^esub>] ?xcpt a"
- by(rule le_listI)
- then have "set (?xcpt b) {\<sqsubseteq>\<^bsub>sup_state_opt P\<^esub>} set (?xcpt a)"
- by (rule lesubstep_type_simple)
- moreover note norm
- ultimately have
- "set (?norm b) \<union> set (?xcpt b) {\<sqsubseteq>\<^bsub>sup_state_opt P\<^esub>} set (?norm a) \<union> set (?xcpt a)"
- by(intro lesubstep_union)
- then show ?thesis using tSome Some by(simp add: Effect.eff_def)
- qed
-qed
-(*>*)
-
-lemma (in JVM_sl) bounded_step: "bounded step (size is)"
-(*<*)
- by (auto simp: bounded_def err_step_def Effect.app_def Effect.eff_def
- error_def map_snd_def
- split: err.splits option.splits)
-(*>*)
-
-theorem (in JVM_sl) step_mono:
- "wf_prog wf_mb P \<Longrightarrow> mono r step (size is) A"
-(*<*)
- apply (simp add: JVM_le_Err_conv)
- apply (insert bounded_step)
- apply (unfold JVM_states_unfold)
- apply (rule mono_lift)
- apply (subgoal_tac "b = Static \<or> b = NonStatic")
- apply (fastforce split:if_splits)\<comment>\<open> order_sup_state_opt' order_sup_state_opt'' \<close>
- apply (simp only:staticb)
- apply (unfold app_mono_def lesub_def)
- apply clarsimp
- apply (erule (2) app_mono)
- apply simp
- apply clarify
- apply (drule eff_mono)
- apply (auto simp add: lesub_def)
- done
-(*
-proof -
- assume wf: "wf_prog wf_mb P"
- let ?r = "sup_state_opt P" and ?n = "length is" and ?app = app and ?step = eff
- let ?A = "opt (\<Union> {list n (types P) |n. n \<le> mxs} \<times>
- list ((case b of Static \<Rightarrow> 0 | NonStatic \<Rightarrow> 1) + length Ts + mxl\<^sub>0)
- (err (types P)))"
- have "order ?r ?A" using wf by simp
- moreover have "app_mono ?r ?app ?n ?A" using app_mono[OF wf]
- by (clarsimp simp: app_mono_def lesub_def)
- moreover have "bounded (err_step ?n ?app ?step) ?n" using bounded_step
- by simp
- moreover have "\<forall>s p t. s \<in> ?A \<and> p < ?n \<and> s \<sqsubseteq>\<^bsub>?r\<^esub> t \<longrightarrow>
- ?app p t \<longrightarrow> set (?step p s) {\<sqsubseteq>\<^bsub>?r\<^esub>} set (?step p t)"
- using eff_mono[OF wf] by simp
- ultimately have "mono (Err.le ?r) (err_step ?n ?app ?step) ?n (err ?A)"
- by(rule mono_lift)
- then show "mono r step (size is) A" using bounded_step
- by (simp add: JVM_le_Err_conv JVM_states_unfold)
-qed
-*)
-(*>*)
-
-
-lemma (in start_context) first_in_A [iff]: "OK first \<in> A"
- using Ts C by (cases b; force intro!: list_appendI simp add: JVM_states_unfold)
-
-
-lemma (in JVM_sl) wt_method_def2:
- "wt_method P C' b Ts T\<^sub>r mxs mxl\<^sub>0 is xt \<tau>s =
- (is \<noteq> [] \<and>
- (b = Static \<or> b = NonStatic) \<and>
- size \<tau>s = size is \<and>
- OK ` set \<tau>s \<subseteq> states P mxs mxl \<and>
- wt_start P C' b Ts mxl\<^sub>0 \<tau>s \<and>
- wt_app_eff (sup_state_opt P) app eff \<tau>s)"
-(*<*)using staticb
- by (unfold wt_method_def wt_app_eff_def wt_instr_def lesub_def
- check_types_def) auto
-(*>*)
-
-
-end
+(* Title: JinjaDCI/BV/TF_JVM.thy
+
+ Author: Tobias Nipkow, Gerwin Klein, Susannah Mansky
+ Copyright 2000 TUM, 2019-20 UIUC
+
+ Based on the Jinja theory BV/TF_JVM.thy by Tobias Nipkow and Gerwin Klein
+*)
+
+section \<open> The Typing Framework for the JVM \label{sec:JVM} \<close>
+
+theory TF_JVM
+imports Jinja.Typing_Framework_err EffectMono BVSpec
+begin
+
+definition exec :: "jvm_prog \<Rightarrow> nat \<Rightarrow> ty \<Rightarrow> ex_table \<Rightarrow> instr list \<Rightarrow> ty\<^sub>i' err step_type"
+where
+ "exec G maxs rT et bs \<equiv>
+ err_step (size bs) (\<lambda>pc. app (bs!pc) G maxs rT pc (size bs) et)
+ (\<lambda>pc. eff (bs!pc) G pc et)"
+
+locale JVM_sl =
+ fixes P :: jvm_prog and mxs and mxl\<^sub>0 and n
+ fixes b and Ts :: "ty list" and "is" and xt and T\<^sub>r
+
+ fixes mxl and A and r and f and app and eff and step
+ defines [simp]: "mxl \<equiv> (case b of Static \<Rightarrow> 0 | NonStatic \<Rightarrow> 1)+size Ts+mxl\<^sub>0"
+ defines [simp]: "A \<equiv> states P mxs mxl"
+ defines [simp]: "r \<equiv> JVM_SemiType.le P mxs mxl"
+ defines [simp]: "f \<equiv> JVM_SemiType.sup P mxs mxl"
+
+ defines [simp]: "app \<equiv> \<lambda>pc. Effect.app (is!pc) P mxs T\<^sub>r pc (size is) xt"
+ defines [simp]: "eff \<equiv> \<lambda>pc. Effect.eff (is!pc) P pc xt"
+ defines [simp]: "step \<equiv> err_step (size is) app eff"
+
+ defines [simp]: "n \<equiv> size is"
+ assumes staticb: "b = Static \<or> b = NonStatic"
+
+locale start_context = JVM_sl +
+ fixes p and C
+
+ assumes wf: "wf_prog p P"
+ assumes C: "is_class P C"
+ assumes Ts: "set Ts \<subseteq> types P"
+
+
+ fixes first :: ty\<^sub>i' and start
+ defines [simp]:
+ "first \<equiv> Some ([],(case b of Static \<Rightarrow> [] | NonStatic \<Rightarrow> [OK (Class C)]) @ map OK Ts @ replicate mxl\<^sub>0 Err)"
+ defines [simp]:
+ "start \<equiv> (OK first) # replicate (size is - 1) (OK None)"
+thm start_context.intro
+
+lemma start_context_intro_auxi:
+ fixes P b Ts p C
+ assumes "b = Static \<or> b = NonStatic "
+ and "wf_prog p P"
+ and "is_class P C"
+ and "set Ts \<subseteq> types P"
+ shows " start_context P b Ts p C"
+ using start_context.intro[OF JVM_sl.intro] start_context_axioms_def assms by auto
+
+subsection \<open> Connecting JVM and Framework \<close>
+
+lemma (in start_context) semi: "semilat (A, r, f)"
+ apply (insert semilat_JVM[OF wf])
+ apply (unfold A_def r_def f_def JVM_SemiType.le_def JVM_SemiType.sup_def states_def)
+ apply auto
+ done
+
+
+lemma (in JVM_sl) step_def_exec: "step \<equiv> exec P mxs T\<^sub>r xt is"
+ by (simp add: exec_def)
+
+lemma special_ex_swap_lemma [iff]:
+ "(? X. (? n. X = A n & P n) & Q X) = (? n. Q(A n) & P n)"
+ by blast
+
+lemma ex_in_list [iff]:
+ "(\<exists>n. ST \<in> list n A \<and> n \<le> mxs) = (set ST \<subseteq> A \<and> size ST \<le> mxs)"
+ by (unfold list_def) auto
+
+lemma singleton_list:
+ "(\<exists>n. [Class C] \<in> list n (types P) \<and> n \<le> mxs) = (is_class P C \<and> 0 < mxs)"
+ by auto
+
+lemma set_drop_subset:
+ "set xs \<subseteq> A \<Longrightarrow> set (drop n xs) \<subseteq> A"
+ by (auto dest: in_set_dropD)
+
+lemma Suc_minus_minus_le:
+ "n < mxs \<Longrightarrow> Suc (n - (n - b)) \<le> mxs"
+ by arith
+
+lemma in_listE:
+ "\<lbrakk> xs \<in> list n A; \<lbrakk>size xs = n; set xs \<subseteq> A\<rbrakk> \<Longrightarrow> P \<rbrakk> \<Longrightarrow> P"
+ by (unfold list_def) blast
+
+declare is_relevant_entry_def [simp]
+declare set_drop_subset [simp]
+
+theorem (in start_context) exec_pres_type:
+ "pres_type step (size is) A"
+(*<*)
+proof -
+ let ?n = "size is" and ?app = app and ?step = eff
+ let ?mxl = "(case b of Static \<Rightarrow> 0 | NonStatic \<Rightarrow> 1) + length Ts + mxl\<^sub>0"
+ let ?A = "opt((Union {list n (types P) |n. n <= mxs}) \<times>
+ list ?mxl (err(types P)))"
+ have "pres_type (err_step ?n ?app ?step) ?n (err ?A)"
+ proof(rule pres_type_lift)
+ have "\<And>s pc pc' s'. s\<in>?A \<Longrightarrow> pc < ?n \<Longrightarrow> ?app pc s
+ \<Longrightarrow> (pc', s')\<in>set (?step pc s) \<Longrightarrow> s' \<in> ?A"
+ proof -
+ fix s pc pc' s'
+ assume asms: "s\<in>?A" "pc < ?n" "?app pc s" "(pc', s')\<in>set (?step pc s)"
+ show "s' \<in> ?A"
+ proof(cases s)
+ case None
+ then show ?thesis using asms by (fastforce dest: effNone)
+ next
+ case (Some ab)
+ then show ?thesis using asms proof(cases "is!pc")
+ case Load
+ then show ?thesis using asms
+ by (fastforce simp: Effect.app_def xcpt_app_def Effect.eff_def
+ xcpt_eff_def norm_eff_def
+ dest: listE_nth_in)
+ next
+ case Push
+ then show ?thesis using asms Some
+ by (fastforce simp: Effect.app_def xcpt_app_def Effect.eff_def
+ xcpt_eff_def norm_eff_def typeof_lit_is_type)
+ next
+ case Getfield
+ then show ?thesis using asms wf
+ by (fastforce simp: Effect.app_def xcpt_app_def Effect.eff_def
+ xcpt_eff_def norm_eff_def
+ dest: sees_field_is_type)
+ next
+ case Getstatic
+ then show ?thesis using asms wf
+ by (fastforce simp: Effect.app_def xcpt_app_def Effect.eff_def
+ xcpt_eff_def norm_eff_def
+ dest: sees_field_is_type)
+ next
+ case (Invoke M n)
+ obtain a b where [simp]: "s = \<lfloor>(a,b)\<rfloor>" using Some asms(1) by blast
+ show ?thesis
+ proof(cases "a!n = NT")
+ case True
+ then show ?thesis using Invoke asms wf
+ by (fastforce simp: Effect.app_def xcpt_app_def Effect.eff_def
+ xcpt_eff_def norm_eff_def)
+ next
+ case False
+ have "(pc', s') \<in> set (norm_eff (Invoke M n) P pc (a, b)) \<or>
+ (pc', s') \<in> set (xcpt_eff (Invoke M n) P pc (a, b) xt)"
+ using Invoke asms(4) by (simp add: Effect.eff_def)
+ then show ?thesis proof(rule disjE)
+ assume "(pc', s') \<in> set (xcpt_eff (Invoke M n) P pc (a, b) xt)"
+ then show ?thesis using Invoke asms(1-3)
+ by (fastforce simp: Effect.app_def xcpt_app_def xcpt_eff_def)
+ next
+ assume norm: "(pc', s') \<in> set (norm_eff (Invoke M n) P pc (a, b))"
+ also have "Suc (length a - Suc n) \<le> mxs" using Invoke asms(1,3)
+ by (simp add: Effect.app_def xcpt_app_def) arith
+ ultimately show ?thesis using False Invoke asms(1-3) wf
+ by (auto simp: Effect.app_def xcpt_app_def norm_eff_def wf_mdecl_def
+ dest!: sees_wf_mdecl)
+ qed
+ qed
+ next
+ case (Invokestatic C M n)
+ obtain a b where [simp]: "s = \<lfloor>(a,b)\<rfloor>" using Some asms(1) by blast
+ have "(pc', s') \<in> set (norm_eff (Invokestatic C M n) P pc (a, b)) \<or>
+ (pc', s') \<in> set (xcpt_eff (Invokestatic C M n) P pc (a, b) xt)"
+ using Invokestatic asms(4) by (simp add: Effect.eff_def)
+ then show ?thesis proof(rule disjE)
+ assume "(pc', s') \<in> set (xcpt_eff (Invokestatic C M n) P pc (a, b) xt)"
+ then show ?thesis using Invokestatic asms(1-3)
+ by (fastforce simp: Effect.app_def xcpt_app_def xcpt_eff_def)
+ next
+ assume norm: "(pc', s') \<in> set (norm_eff (Invokestatic C M n) P pc (a, b))"
+ also have "Suc (length a - Suc n) \<le> mxs" using Invokestatic asms(1,3)
+ by (simp add: Effect.app_def xcpt_app_def) arith
+ ultimately show ?thesis using Invokestatic asms(1-3) wf
+ by (auto simp: Effect.app_def xcpt_app_def norm_eff_def wf_mdecl_def
+ dest!: sees_wf_mdecl)
+ qed
+ qed (fastforce simp: Effect.app_def xcpt_app_def Effect.eff_def
+ xcpt_eff_def norm_eff_def)+
+ qed
+ qed
+ then show "\<forall>s\<in>?A. \<forall>p. p < ?n \<longrightarrow> ?app p s \<longrightarrow> (\<forall>(q, s')\<in>set (?step p s). s' \<in> ?A)"
+ by clarsimp
+ qed
+ then show ?thesis by (simp add: JVM_states_unfold)
+qed
+(*>*)
+
+declare is_relevant_entry_def [simp del]
+declare set_drop_subset [simp del]
+
+lemma lesubstep_type_simple:
+ "xs [\<sqsubseteq>\<^bsub>Product.le (=) r\<^esub>] ys \<Longrightarrow> set xs {\<sqsubseteq>\<^bsub>r\<^esub>} set ys"
+(*<*)
+proof -
+ assume assm: "xs [\<sqsubseteq>\<^bsub>Product.le (=) r\<^esub>] ys"
+ have "\<And>a b i y. (a, b) = xs ! i \<Longrightarrow> i < length xs
+ \<Longrightarrow> \<exists>\<tau>'. (\<exists>i. (a, \<tau>') = ys ! i \<and> i < length xs) \<and> b \<sqsubseteq>\<^bsub>r\<^esub> \<tau>'"
+ proof -
+ fix a b i assume ith: "(a, b) = xs ! i" and len: "i < length xs"
+ obtain \<tau> where "ys ! i = (a, \<tau>) \<and> r b \<tau>"
+ using le_listD[OF assm len] ith
+ by (clarsimp simp: lesub_def Product.le_def)
+ then have "(a, \<tau>) = ys ! i \<and> b \<sqsubseteq>\<^bsub>r\<^esub> \<tau>"
+ by (clarsimp simp: lesub_def)
+ with len show "\<exists>\<tau>'. (\<exists>i. (a, \<tau>') = ys ! i \<and> i < length xs) \<and> b \<sqsubseteq>\<^bsub>r\<^esub> \<tau>'"
+ by fastforce
+ qed
+ then show "set xs {\<sqsubseteq>\<^bsub>r\<^esub>} set ys" using assm
+ by (clarsimp simp: lesubstep_type_def set_conv_nth)
+qed
+(*>*)
+
+declare is_relevant_entry_def [simp del]
+
+
+lemma conjI2: "\<lbrakk> A; A \<Longrightarrow> B \<rbrakk> \<Longrightarrow> A \<and> B" by blast
+
+lemma (in JVM_sl) eff_mono:
+assumes wf: "wf_prog p P" and "pc < length is" and
+ lesub: "s \<sqsubseteq>\<^bsub>sup_state_opt P\<^esub> t" and app: "app pc t"
+shows "set (eff pc s) {\<sqsubseteq>\<^bsub>sup_state_opt P\<^esub>} set (eff pc t)"
+(*<*)
+proof(cases t)
+ case None then show ?thesis using lesub
+ by (simp add: Effect.eff_def lesub_def)
+next
+ case tSome: (Some a)
+ show ?thesis proof(cases s)
+ case None then show ?thesis using lesub
+ by (simp add: Effect.eff_def lesub_def)
+ next
+ case (Some b)
+ let ?norm = "\<lambda>x. norm_eff (is ! pc) P pc x"
+ let ?xcpt = "\<lambda>x. xcpt_eff (is ! pc) P pc x xt"
+ let ?r = "Product.le (=) (sup_state_opt P)"
+ let ?\<tau>' = "\<lfloor>eff\<^sub>i (is ! pc, P, a)\<rfloor>"
+ {
+ fix x assume xb: "x \<in> set (succs (is ! pc) b pc)"
+ then have appi: "app\<^sub>i (is ! pc, P, pc, mxs, T\<^sub>r, a)" and
+ bia: "P \<turnstile> b \<le>\<^sub>i a" and appa: "app pc \<lfloor>a\<rfloor>"
+ using lesub app tSome Some by (auto simp add: lesub_def Effect.app_def)
+ have xa: "x \<in> set (succs (is ! pc) a pc)"
+ using xb succs_mono[OF wf appi bia] by auto
+ then have "(x, ?\<tau>') \<in> (\<lambda>pc'. (pc', ?\<tau>')) ` set (succs (is ! pc) a pc)"
+ by (rule imageI)
+ moreover have "P \<turnstile> \<lfloor>eff\<^sub>i (is ! pc, P, b)\<rfloor> \<le>' ?\<tau>'"
+ using xb xa eff\<^sub>i_mono[OF wf bia] appa by fastforce
+ ultimately have "\<exists>\<tau>'. (x, \<tau>') \<in> (\<lambda>pc'. (pc', \<lfloor>eff\<^sub>i (is ! pc, P, a)\<rfloor>)) ` set (succs (is ! pc) a pc) \<and>
+ P \<turnstile> \<lfloor>eff\<^sub>i (is ! pc, P, b)\<rfloor> \<le>' \<tau>'" by blast
+ }
+ then have norm: "set (?norm b) {\<sqsubseteq>\<^bsub>sup_state_opt P\<^esub>} set (?norm a)"
+ using tSome Some by (clarsimp simp: norm_eff_def lesubstep_type_def lesub_def)
+ obtain a1 b1 a2 b2 where a: "a = (a1, b1)" and b: "b = (a2, b2)"
+ using tSome Some by fastforce
+ then have a12: "size a2 = size a1" using lesub tSome Some
+ by (clarsimp simp: lesub_def list_all2_lengthD)
+ have "length (?xcpt b) = length (?xcpt a)"
+ by (simp add: xcpt_eff_def split_beta)
+ moreover have "\<And>n. n < length (?xcpt b) \<Longrightarrow> (?xcpt b) ! n \<sqsubseteq>\<^bsub>?r\<^esub> (?xcpt a) ! n"
+ using lesub tSome Some a b a12
+ by (simp add: xcpt_eff_def split_beta fun_of_def) (simp add: lesub_def)
+ ultimately have "?xcpt b [\<sqsubseteq>\<^bsub>?r\<^esub>] ?xcpt a"
+ by(rule le_listI)
+ then have "set (?xcpt b) {\<sqsubseteq>\<^bsub>sup_state_opt P\<^esub>} set (?xcpt a)"
+ by (rule lesubstep_type_simple)
+ moreover note norm
+ ultimately have
+ "set (?norm b) \<union> set (?xcpt b) {\<sqsubseteq>\<^bsub>sup_state_opt P\<^esub>} set (?norm a) \<union> set (?xcpt a)"
+ by(intro lesubstep_union)
+ then show ?thesis using tSome Some by(simp add: Effect.eff_def)
+ qed
+qed
+(*>*)
+
+lemma (in JVM_sl) bounded_step: "bounded step (size is)"
+(*<*)
+ by (auto simp: bounded_def err_step_def Effect.app_def Effect.eff_def
+ error_def map_snd_def
+ split: err.splits option.splits)
+(*>*)
+
+theorem (in JVM_sl) step_mono:
+ "wf_prog wf_mb P \<Longrightarrow> mono r step (size is) A"
+(*<*)
+ apply (simp add: JVM_le_Err_conv)
+ apply (insert bounded_step)
+ apply (unfold JVM_states_unfold)
+ apply (rule mono_lift)
+ apply (subgoal_tac "b = Static \<or> b = NonStatic")
+ apply (fastforce split:if_splits)\<comment>\<open> order_sup_state_opt' order_sup_state_opt'' \<close>
+ apply (simp only:staticb)
+ apply (unfold app_mono_def lesub_def)
+ apply clarsimp
+ apply (erule (2) app_mono)
+ apply simp
+ apply clarify
+ apply (drule eff_mono)
+ apply (auto simp add: lesub_def)
+ done
+(*
+proof -
+ assume wf: "wf_prog wf_mb P"
+ let ?r = "sup_state_opt P" and ?n = "length is" and ?app = app and ?step = eff
+ let ?A = "opt (\<Union> {list n (types P) |n. n \<le> mxs} \<times>
+ list ((case b of Static \<Rightarrow> 0 | NonStatic \<Rightarrow> 1) + length Ts + mxl\<^sub>0)
+ (err (types P)))"
+ have "order ?r ?A" using wf by simp
+ moreover have "app_mono ?r ?app ?n ?A" using app_mono[OF wf]
+ by (clarsimp simp: app_mono_def lesub_def)
+ moreover have "bounded (err_step ?n ?app ?step) ?n" using bounded_step
+ by simp
+ moreover have "\<forall>s p t. s \<in> ?A \<and> p < ?n \<and> s \<sqsubseteq>\<^bsub>?r\<^esub> t \<longrightarrow>
+ ?app p t \<longrightarrow> set (?step p s) {\<sqsubseteq>\<^bsub>?r\<^esub>} set (?step p t)"
+ using eff_mono[OF wf] by simp
+ ultimately have "mono (Err.le ?r) (err_step ?n ?app ?step) ?n (err ?A)"
+ by(rule mono_lift)
+ then show "mono r step (size is) A" using bounded_step
+ by (simp add: JVM_le_Err_conv JVM_states_unfold)
+qed
+*)
+(*>*)
+
+
+lemma (in start_context) first_in_A [iff]: "OK first \<in> A"
+ using Ts C by (cases b; force intro!: list_appendI simp add: JVM_states_unfold)
+
+
+lemma (in JVM_sl) wt_method_def2:
+ "wt_method P C' b Ts T\<^sub>r mxs mxl\<^sub>0 is xt \<tau>s =
+ (is \<noteq> [] \<and>
+ (b = Static \<or> b = NonStatic) \<and>
+ size \<tau>s = size is \<and>
+ OK ` set \<tau>s \<subseteq> states P mxs mxl \<and>
+ wt_start P C' b Ts mxl\<^sub>0 \<tau>s \<and>
+ wt_app_eff (sup_state_opt P) app eff \<tau>s)"
+(*<*)using staticb
+ by (unfold wt_method_def wt_app_eff_def wt_instr_def lesub_def
+ check_types_def) auto
+(*>*)
+
+
+end
diff --git a/thys/JinjaDCI/Common/Auxiliary.thy b/thys/JinjaDCI/Common/Auxiliary.thy
--- a/thys/JinjaDCI/Common/Auxiliary.thy
+++ b/thys/JinjaDCI/Common/Auxiliary.thy
@@ -1,175 +1,175 @@
-(* Title: JinjaDCI/Common/Auxiliary.thy
-
- Author: David von Oheimb, Tobias Nipkow, Susannah Mansky
- Copyright 1999 TU Muenchen, 2019-20 UIUC
-
- Based on the Jinja theory Common/Auxiliary.thy by David von Oheimb and Tobias Nipkow
-*)
-
-chapter \<open> Jinja Source Language \label{cha:j} \<close>
-
-section \<open> Auxiliary Definitions \<close>
-
-theory Auxiliary imports Main begin
-(* FIXME move and possibly turn into a general simproc *)
-lemma nat_add_max_le[simp]:
- "((n::nat) + max i j \<le> m) = (n + i \<le> m \<and> n + j \<le> m)"
- (*<*)by arith(*>*)
-
-lemma Suc_add_max_le[simp]:
- "(Suc(n + max i j) \<le> m) = (Suc(n + i) \<le> m \<and> Suc(n + j) \<le> m)"
-(*<*)by arith(*>*)
-
-
-notation Some ("(\<lfloor>_\<rfloor>)")
-
-(*<*)
-declare
- option.splits[split]
- Let_def[simp]
- subset_insertI2 [simp]
- Cons_eq_map_conv [iff]
-(*>*)
-
-
-subsection \<open>@{text distinct_fst}\<close>
-
-definition distinct_fst :: "('a \<times> 'b) list \<Rightarrow> bool"
-where
- "distinct_fst \<equiv> distinct \<circ> map fst"
-
-lemma distinct_fst_Nil [simp]:
- "distinct_fst []"
- (*<*)
-by (unfold distinct_fst_def) (simp (no_asm))
-(*>*)
-
-lemma distinct_fst_Cons [simp]:
- "distinct_fst ((k,x)#kxs) = (distinct_fst kxs \<and> (\<forall>y. (k,y) \<notin> set kxs))"
-(*<*)
-by (unfold distinct_fst_def) (auto simp:image_def)
-(*>*)
-
-lemma distinct_fst_appendD:
- "distinct_fst(kxs @ kxs') \<Longrightarrow> distinct_fst kxs \<and> distinct_fst kxs'"
-(*<*)by(induct kxs, auto)(*>*)
-
-lemma map_of_SomeI:
- "\<lbrakk> distinct_fst kxs; (k,x) \<in> set kxs \<rbrakk> \<Longrightarrow> map_of kxs k = Some x"
-(*<*)by (induct kxs) (auto simp:fun_upd_apply)(*>*)
-
-
-subsection \<open> Using @{term list_all2} for relations \<close>
-
-definition fun_of :: "('a \<times> 'b) set \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> bool"
-where
- "fun_of S \<equiv> \<lambda>x y. (x,y) \<in> S"
-
-text \<open> Convenience lemmas \<close>
-(*<*)
-declare fun_of_def [simp]
-(*>*)
-lemma rel_list_all2_Cons [iff]:
- "list_all2 (fun_of S) (x#xs) (y#ys) =
- ((x,y) \<in> S \<and> list_all2 (fun_of S) xs ys)"
- (*<*)by simp(*>*)
-
-lemma rel_list_all2_Cons1:
- "list_all2 (fun_of S) (x#xs) ys =
- (\<exists>z zs. ys = z#zs \<and> (x,z) \<in> S \<and> list_all2 (fun_of S) xs zs)"
- (*<*)by (cases ys) auto(*>*)
-
-lemma rel_list_all2_Cons2:
- "list_all2 (fun_of S) xs (y#ys) =
- (\<exists>z zs. xs = z#zs \<and> (z,y) \<in> S \<and> list_all2 (fun_of S) zs ys)"
- (*<*)by (cases xs) auto(*>*)
-
-lemma rel_list_all2_refl:
- "(\<And>x. (x,x) \<in> S) \<Longrightarrow> list_all2 (fun_of S) xs xs"
- (*<*)by (simp add: list_all2_refl)(*>*)
-
-lemma rel_list_all2_antisym:
- "\<lbrakk> (\<And>x y. \<lbrakk>(x,y) \<in> S; (y,x) \<in> T\<rbrakk> \<Longrightarrow> x = y);
- list_all2 (fun_of S) xs ys; list_all2 (fun_of T) ys xs \<rbrakk> \<Longrightarrow> xs = ys"
- (*<*)by (rule list_all2_antisym) auto(*>*)
-
-lemma rel_list_all2_trans:
- "\<lbrakk> \<And>a b c. \<lbrakk>(a,b) \<in> R; (b,c) \<in> S\<rbrakk> \<Longrightarrow> (a,c) \<in> T;
- list_all2 (fun_of R) as bs; list_all2 (fun_of S) bs cs\<rbrakk>
- \<Longrightarrow> list_all2 (fun_of T) as cs"
- (*<*)by (rule list_all2_trans) auto(*>*)
-
-lemma rel_list_all2_update_cong:
- "\<lbrakk> i<size xs; list_all2 (fun_of S) xs ys; (x,y) \<in> S \<rbrakk>
- \<Longrightarrow> list_all2 (fun_of S) (xs[i:=x]) (ys[i:=y])"
- (*<*)by (simp add: list_all2_update_cong)(*>*)
-
-lemma rel_list_all2_nthD:
- "\<lbrakk> list_all2 (fun_of S) xs ys; p < size xs \<rbrakk> \<Longrightarrow> (xs!p,ys!p) \<in> S"
- (*<*)by (drule list_all2_nthD) auto(*>*)
-
-lemma rel_list_all2I:
- "\<lbrakk> length a = length b; \<And>n. n < length a \<Longrightarrow> (a!n,b!n) \<in> S \<rbrakk> \<Longrightarrow> list_all2 (fun_of S) a b"
- (*<*)by (erule list_all2_all_nthI) simp(*>*)
-
-(*<*)declare fun_of_def [simp del](*>*)
-
-subsection \<open> Auxiliary properties of @{text "map_of"} function \<close>
-
-lemma map_of_set_pcs_notin: "C \<notin> (\<lambda>t. snd (fst t)) ` set FDTs \<Longrightarrow> map_of FDTs (F, C) = None"
-(*<*)by (metis image_eqI image_image map_of_eq_None_iff snd_conv)(*>*)
-
-lemma map_of_insertmap_SomeD':
- "map_of fs F = Some y \<Longrightarrow> map_of (map (\<lambda>(F, y). (F, D, y)) fs) F = Some(D,y)"
-(*<*)by (induct fs) (auto simp:fun_upd_apply split: if_split_asm)(*>*)
-
-lemma map_of_reinsert_neq_None:
- "Ca \<noteq> D \<Longrightarrow> map_of (map (\<lambda>(F, y). ((F, Ca), y)) fs) (F, D) = None"
-(*<*)by (induct fs) (auto simp:fun_upd_apply split: if_split_asm)(*>*)
-
-lemma map_of_remap_insertmap:
- "map_of (map ((\<lambda>((F, D), b, T). (F, D, b, T)) \<circ> (\<lambda>(F, y). ((F, D), y))) fs)
- = map_of (map (\<lambda>(F, y). (F, D, y)) fs)"
-(*<*)by (induct fs) (auto simp:fun_upd_apply split: if_split_asm)(*>*)
-
-
-lemma map_of_reinsert_SomeD:
- "map_of (map (\<lambda>(F, y). ((F, D), y)) fs) (F, D) = Some T \<Longrightarrow> map_of fs F = Some T"
-(*<*)by (induct fs) (auto simp:fun_upd_apply split: if_split_asm)(*>*)
-
-lemma map_of_filtered_SomeD:
-"map_of fs (F,D) = Some (a, T) \<Longrightarrow> Q ((F,D),a,T) \<Longrightarrow>
- map_of (map (\<lambda>((F,D), b, T). ((F,D), P T)) (filter Q fs))
- (F,D) = Some (P T)"
-(*<*)by (induct fs) (auto simp:fun_upd_apply split: if_split_asm)(*>*)
-
-
-lemma map_of_remove_filtered_SomeD:
-"map_of fs (F,C) = Some (a, T) \<Longrightarrow> Q ((F,C),a,T) \<Longrightarrow>
- map_of (map (\<lambda>((F,D), b, T). (F, P T)) [((F, D), b, T)\<leftarrow>fs . Q ((F, D), b, T) \<and> D = C])
- F = Some (P T)"
-(*<*)by (induct fs) (auto simp:fun_upd_apply split: if_split_asm)(*>*)
-
-
-lemma map_of_Some_None_split:
-assumes "t = map (\<lambda>(F, y). ((F, C), y)) fs @ t'" "map_of t' (F, C) = None" "map_of t (F, C) = Some y"
-shows "map_of (map (\<lambda>((F, D), b, T). (F, D, b, T)) t) F = Some (C, y)"
-(*<*)
-proof -
- have "map_of (map (\<lambda>(F, y). ((F, C), y)) fs) (F, C) = Some y" using assms by auto
- then have "\<forall>p. map_of fs F = Some p \<or> Some y \<noteq> Some p"
- by (metis map_of_reinsert_SomeD)
- then have "\<forall>f b p pa. ((f ++ map_of (map (\<lambda>(a, p). (a, b::'b, p)) fs)) F = Some p \<or> Some (b, pa) \<noteq> Some p)
- \<or> Some y \<noteq> Some pa"
- by (metis (no_types) map_add_find_right map_of_insertmap_SomeD')
- then have "(map_of (map (\<lambda>((a, b), c, d). (a, b, c, d)) t')
- ++ map_of (map (\<lambda>(a, p). (a, C, p)) fs)) F = Some (C, y)"
- by blast
- then have "(map_of (map (\<lambda>((a, b), c, d). (a, b, c, d)) t')
- ++ map_of (map ((\<lambda>((a, b), c, d). (a, b, c, d)) \<circ> (\<lambda>(a, y). ((a, C), y))) fs)) F = Some (C, y)"
- by (simp add: map_of_remap_insertmap)
- then show ?thesis using assms by auto
-qed
-(*>*)
-
-end
+(* Title: JinjaDCI/Common/Auxiliary.thy
+
+ Author: David von Oheimb, Tobias Nipkow, Susannah Mansky
+ Copyright 1999 TU Muenchen, 2019-20 UIUC
+
+ Based on the Jinja theory Common/Auxiliary.thy by David von Oheimb and Tobias Nipkow
+*)
+
+chapter \<open> Jinja Source Language \label{cha:j} \<close>
+
+section \<open> Auxiliary Definitions \<close>
+
+theory Auxiliary imports Main begin
+(* FIXME move and possibly turn into a general simproc *)
+lemma nat_add_max_le[simp]:
+ "((n::nat) + max i j \<le> m) = (n + i \<le> m \<and> n + j \<le> m)"
+ (*<*)by arith(*>*)
+
+lemma Suc_add_max_le[simp]:
+ "(Suc(n + max i j) \<le> m) = (Suc(n + i) \<le> m \<and> Suc(n + j) \<le> m)"
+(*<*)by arith(*>*)
+
+
+notation Some ("(\<lfloor>_\<rfloor>)")
+
+(*<*)
+declare
+ option.splits[split]
+ Let_def[simp]
+ subset_insertI2 [simp]
+ Cons_eq_map_conv [iff]
+(*>*)
+
+
+subsection \<open>@{text distinct_fst}\<close>
+
+definition distinct_fst :: "('a \<times> 'b) list \<Rightarrow> bool"
+where
+ "distinct_fst \<equiv> distinct \<circ> map fst"
+
+lemma distinct_fst_Nil [simp]:
+ "distinct_fst []"
+ (*<*)
+by (unfold distinct_fst_def) (simp (no_asm))
+(*>*)
+
+lemma distinct_fst_Cons [simp]:
+ "distinct_fst ((k,x)#kxs) = (distinct_fst kxs \<and> (\<forall>y. (k,y) \<notin> set kxs))"
+(*<*)
+by (unfold distinct_fst_def) (auto simp:image_def)
+(*>*)
+
+lemma distinct_fst_appendD:
+ "distinct_fst(kxs @ kxs') \<Longrightarrow> distinct_fst kxs \<and> distinct_fst kxs'"
+(*<*)by(induct kxs, auto)(*>*)
+
+lemma map_of_SomeI:
+ "\<lbrakk> distinct_fst kxs; (k,x) \<in> set kxs \<rbrakk> \<Longrightarrow> map_of kxs k = Some x"
+(*<*)by (induct kxs) (auto simp:fun_upd_apply)(*>*)
+
+
+subsection \<open> Using @{term list_all2} for relations \<close>
+
+definition fun_of :: "('a \<times> 'b) set \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> bool"
+where
+ "fun_of S \<equiv> \<lambda>x y. (x,y) \<in> S"
+
+text \<open> Convenience lemmas \<close>
+(*<*)
+declare fun_of_def [simp]
+(*>*)
+lemma rel_list_all2_Cons [iff]:
+ "list_all2 (fun_of S) (x#xs) (y#ys) =
+ ((x,y) \<in> S \<and> list_all2 (fun_of S) xs ys)"
+ (*<*)by simp(*>*)
+
+lemma rel_list_all2_Cons1:
+ "list_all2 (fun_of S) (x#xs) ys =
+ (\<exists>z zs. ys = z#zs \<and> (x,z) \<in> S \<and> list_all2 (fun_of S) xs zs)"
+ (*<*)by (cases ys) auto(*>*)
+
+lemma rel_list_all2_Cons2:
+ "list_all2 (fun_of S) xs (y#ys) =
+ (\<exists>z zs. xs = z#zs \<and> (z,y) \<in> S \<and> list_all2 (fun_of S) zs ys)"
+ (*<*)by (cases xs) auto(*>*)
+
+lemma rel_list_all2_refl:
+ "(\<And>x. (x,x) \<in> S) \<Longrightarrow> list_all2 (fun_of S) xs xs"
+ (*<*)by (simp add: list_all2_refl)(*>*)
+
+lemma rel_list_all2_antisym:
+ "\<lbrakk> (\<And>x y. \<lbrakk>(x,y) \<in> S; (y,x) \<in> T\<rbrakk> \<Longrightarrow> x = y);
+ list_all2 (fun_of S) xs ys; list_all2 (fun_of T) ys xs \<rbrakk> \<Longrightarrow> xs = ys"
+ (*<*)by (rule list_all2_antisym) auto(*>*)
+
+lemma rel_list_all2_trans:
+ "\<lbrakk> \<And>a b c. \<lbrakk>(a,b) \<in> R; (b,c) \<in> S\<rbrakk> \<Longrightarrow> (a,c) \<in> T;
+ list_all2 (fun_of R) as bs; list_all2 (fun_of S) bs cs\<rbrakk>
+ \<Longrightarrow> list_all2 (fun_of T) as cs"
+ (*<*)by (rule list_all2_trans) auto(*>*)
+
+lemma rel_list_all2_update_cong:
+ "\<lbrakk> i<size xs; list_all2 (fun_of S) xs ys; (x,y) \<in> S \<rbrakk>
+ \<Longrightarrow> list_all2 (fun_of S) (xs[i:=x]) (ys[i:=y])"
+ (*<*)by (simp add: list_all2_update_cong)(*>*)
+
+lemma rel_list_all2_nthD:
+ "\<lbrakk> list_all2 (fun_of S) xs ys; p < size xs \<rbrakk> \<Longrightarrow> (xs!p,ys!p) \<in> S"
+ (*<*)by (drule list_all2_nthD) auto(*>*)
+
+lemma rel_list_all2I:
+ "\<lbrakk> length a = length b; \<And>n. n < length a \<Longrightarrow> (a!n,b!n) \<in> S \<rbrakk> \<Longrightarrow> list_all2 (fun_of S) a b"
+ (*<*)by (erule list_all2_all_nthI) simp(*>*)
+
+(*<*)declare fun_of_def [simp del](*>*)
+
+subsection \<open> Auxiliary properties of @{text "map_of"} function \<close>
+
+lemma map_of_set_pcs_notin: "C \<notin> (\<lambda>t. snd (fst t)) ` set FDTs \<Longrightarrow> map_of FDTs (F, C) = None"
+(*<*)by (metis image_eqI image_image map_of_eq_None_iff snd_conv)(*>*)
+
+lemma map_of_insertmap_SomeD':
+ "map_of fs F = Some y \<Longrightarrow> map_of (map (\<lambda>(F, y). (F, D, y)) fs) F = Some(D,y)"
+(*<*)by (induct fs) (auto simp:fun_upd_apply split: if_split_asm)(*>*)
+
+lemma map_of_reinsert_neq_None:
+ "Ca \<noteq> D \<Longrightarrow> map_of (map (\<lambda>(F, y). ((F, Ca), y)) fs) (F, D) = None"
+(*<*)by (induct fs) (auto simp:fun_upd_apply split: if_split_asm)(*>*)
+
+lemma map_of_remap_insertmap:
+ "map_of (map ((\<lambda>((F, D), b, T). (F, D, b, T)) \<circ> (\<lambda>(F, y). ((F, D), y))) fs)
+ = map_of (map (\<lambda>(F, y). (F, D, y)) fs)"
+(*<*)by (induct fs) (auto simp:fun_upd_apply split: if_split_asm)(*>*)
+
+
+lemma map_of_reinsert_SomeD:
+ "map_of (map (\<lambda>(F, y). ((F, D), y)) fs) (F, D) = Some T \<Longrightarrow> map_of fs F = Some T"
+(*<*)by (induct fs) (auto simp:fun_upd_apply split: if_split_asm)(*>*)
+
+lemma map_of_filtered_SomeD:
+"map_of fs (F,D) = Some (a, T) \<Longrightarrow> Q ((F,D),a,T) \<Longrightarrow>
+ map_of (map (\<lambda>((F,D), b, T). ((F,D), P T)) (filter Q fs))
+ (F,D) = Some (P T)"
+(*<*)by (induct fs) (auto simp:fun_upd_apply split: if_split_asm)(*>*)
+
+
+lemma map_of_remove_filtered_SomeD:
+"map_of fs (F,C) = Some (a, T) \<Longrightarrow> Q ((F,C),a,T) \<Longrightarrow>
+ map_of (map (\<lambda>((F,D), b, T). (F, P T)) [((F, D), b, T)\<leftarrow>fs . Q ((F, D), b, T) \<and> D = C])
+ F = Some (P T)"
+(*<*)by (induct fs) (auto simp:fun_upd_apply split: if_split_asm)(*>*)
+
+
+lemma map_of_Some_None_split:
+assumes "t = map (\<lambda>(F, y). ((F, C), y)) fs @ t'" "map_of t' (F, C) = None" "map_of t (F, C) = Some y"
+shows "map_of (map (\<lambda>((F, D), b, T). (F, D, b, T)) t) F = Some (C, y)"
+(*<*)
+proof -
+ have "map_of (map (\<lambda>(F, y). ((F, C), y)) fs) (F, C) = Some y" using assms by auto
+ then have "\<forall>p. map_of fs F = Some p \<or> Some y \<noteq> Some p"
+ by (metis map_of_reinsert_SomeD)
+ then have "\<forall>f b p pa. ((f ++ map_of (map (\<lambda>(a, p). (a, b::'b, p)) fs)) F = Some p \<or> Some (b, pa) \<noteq> Some p)
+ \<or> Some y \<noteq> Some pa"
+ by (metis (no_types) map_add_find_right map_of_insertmap_SomeD')
+ then have "(map_of (map (\<lambda>((a, b), c, d). (a, b, c, d)) t')
+ ++ map_of (map (\<lambda>(a, p). (a, C, p)) fs)) F = Some (C, y)"
+ by blast
+ then have "(map_of (map (\<lambda>((a, b), c, d). (a, b, c, d)) t')
+ ++ map_of (map ((\<lambda>((a, b), c, d). (a, b, c, d)) \<circ> (\<lambda>(a, y). ((a, C), y))) fs)) F = Some (C, y)"
+ by (simp add: map_of_remap_insertmap)
+ then show ?thesis using assms by auto
+qed
+(*>*)
+
+end
diff --git a/thys/JinjaDCI/Common/Conform.thy b/thys/JinjaDCI/Common/Conform.thy
--- a/thys/JinjaDCI/Common/Conform.thy
+++ b/thys/JinjaDCI/Common/Conform.thy
@@ -1,218 +1,218 @@
-(* Title: JinjaDCI/Common/Conform.thy
-
- Author: David von Oheimb, Tobias Nipkow, Susannah Mansky
- Copyright 1999 Technische Universitaet Muenchen, 2019-20 UIUC
-
- Based on the Jinja theory Common/Conform.thy by David von Oheimb and Tobias Nipkow
-*)
-
-section \<open> Conformance Relations for Type Soundness Proofs \<close>
-
-theory Conform
-imports Exceptions
-begin
-
-definition conf :: "'m prog \<Rightarrow> heap \<Rightarrow> val \<Rightarrow> ty \<Rightarrow> bool" ("_,_ \<turnstile> _ :\<le> _" [51,51,51,51] 50)
-where
- "P,h \<turnstile> v :\<le> T \<equiv>
- \<exists>T'. typeof\<^bsub>h\<^esub> v = Some T' \<and> P \<turnstile> T' \<le> T"
-
-definition oconf :: "'m prog \<Rightarrow> heap \<Rightarrow> obj \<Rightarrow> bool" ("_,_ \<turnstile> _ \<surd>" [51,51,51] 50)
-where
- "P,h \<turnstile> obj \<surd> \<equiv>
- let (C,fs) = obj in \<forall>F D T. P \<turnstile> C has F,NonStatic:T in D \<longrightarrow>
- (\<exists>v. fs(F,D) = Some v \<and> P,h \<turnstile> v :\<le> T)"
-
-definition soconf :: "'m prog \<Rightarrow> heap \<Rightarrow> cname \<Rightarrow> sfields \<Rightarrow> bool" ("_,_,_ \<turnstile>\<^sub>s _ \<surd>" [51,51,51,51] 50)
-where
- "P,h,C \<turnstile>\<^sub>s sfs \<surd> \<equiv>
- \<forall>F T. P \<turnstile> C has F,Static:T in C \<longrightarrow>
- (\<exists>v. sfs F = Some v \<and> P,h \<turnstile> v :\<le> T)"
-
-definition hconf :: "'m prog \<Rightarrow> heap \<Rightarrow> bool" ("_ \<turnstile> _ \<surd>" [51,51] 50)
-where
- "P \<turnstile> h \<surd> \<equiv>
- (\<forall>a obj. h a = Some obj \<longrightarrow> P,h \<turnstile> obj \<surd>) \<and> preallocated h"
-
-definition shconf :: "'m prog \<Rightarrow> heap \<Rightarrow> sheap \<Rightarrow> bool" ("_,_ \<turnstile>\<^sub>s _ \<surd>" [51,51,51] 50)
-where
- "P,h \<turnstile>\<^sub>s sh \<surd> \<equiv>
- (\<forall>C sfs i. sh C = Some(sfs,i) \<longrightarrow> P,h,C \<turnstile>\<^sub>s sfs \<surd>)"
-
-definition lconf :: "'m prog \<Rightarrow> heap \<Rightarrow> (vname \<rightharpoonup> val) \<Rightarrow> (vname \<rightharpoonup> ty) \<Rightarrow> bool" ("_,_ \<turnstile> _ '(:\<le>') _" [51,51,51,51] 50)
-where
- "P,h \<turnstile> l (:\<le>) E \<equiv>
- \<forall>V v. l V = Some v \<longrightarrow> (\<exists>T. E V = Some T \<and> P,h \<turnstile> v :\<le> T)"
-
-abbreviation
- confs :: "'m prog \<Rightarrow> heap \<Rightarrow> val list \<Rightarrow> ty list \<Rightarrow> bool"
- ("_,_ \<turnstile> _ [:\<le>] _" [51,51,51,51] 50) where
- "P,h \<turnstile> vs [:\<le>] Ts \<equiv> list_all2 (conf P h) vs Ts"
-
-
-subsection\<open> Value conformance @{text":\<le>"} \<close>
-
-lemma conf_Null [simp]: "P,h \<turnstile> Null :\<le> T = P \<turnstile> NT \<le> T"
-(*<*)by (simp (no_asm) add: conf_def)(*>*)
-
-lemma typeof_conf[simp]: "typeof\<^bsub>h\<^esub> v = Some T \<Longrightarrow> P,h \<turnstile> v :\<le> T"
-(*<*)by (induct v) (auto simp: conf_def)(*>*)
-
-lemma typeof_lit_conf[simp]: "typeof v = Some T \<Longrightarrow> P,h \<turnstile> v :\<le> T"
-(*<*)by (rule typeof_conf[OF typeof_lit_typeof])(*>*)
-
-lemma defval_conf[simp]: "P,h \<turnstile> default_val T :\<le> T"
-(*<*)by (cases T) (auto simp: conf_def)(*>*)
-
-lemma conf_upd_obj: "h a = Some(C,fs) \<Longrightarrow> (P,h(a\<mapsto>(C,fs')) \<turnstile> x :\<le> T) = (P,h \<turnstile> x :\<le> T)"
-(*<*)by (rule val.induct) (auto simp:conf_def fun_upd_apply)(*>*)
-
-lemma conf_widen: "P,h \<turnstile> v :\<le> T \<Longrightarrow> P \<turnstile> T \<le> T' \<Longrightarrow> P,h \<turnstile> v :\<le> T'"
-(*<*)by (induct v) (auto intro: widen_trans simp: conf_def)(*>*)
-
-lemma conf_hext: "h \<unlhd> h' \<Longrightarrow> P,h \<turnstile> v :\<le> T \<Longrightarrow> P,h' \<turnstile> v :\<le> T"
-(*<*)by (induct v) (auto dest: hext_objD simp: conf_def)(*>*)
-
-lemma conf_ClassD: "P,h \<turnstile> v :\<le> Class C \<Longrightarrow>
- v = Null \<or> (\<exists>a obj T. v = Addr a \<and> h a = Some obj \<and> obj_ty obj = T \<and> P \<turnstile> T \<le> Class C)"
-(*<*)by(induct v) (auto simp: conf_def)(*>*)
-
-lemma conf_NT [iff]: "P,h \<turnstile> v :\<le> NT = (v = Null)"
-(*<*)by (auto simp add: conf_def)(*>*)
-
-lemma non_npD: "\<lbrakk> v \<noteq> Null; P,h \<turnstile> v :\<le> Class C \<rbrakk>
- \<Longrightarrow> \<exists>a C' fs. v = Addr a \<and> h a = Some(C',fs) \<and> P \<turnstile> C' \<preceq>\<^sup>* C"
-(*<*)by (auto dest: conf_ClassD)(*>*)
-
-
-subsection\<open> Value list conformance @{text"[:\<le>]"} \<close>
-
-lemma confs_widens [trans]: "\<lbrakk>P,h \<turnstile> vs [:\<le>] Ts; P \<turnstile> Ts [\<le>] Ts'\<rbrakk> \<Longrightarrow> P,h \<turnstile> vs [:\<le>] Ts'"
-(*<*)by(auto intro: list_all2_trans conf_widen)(*>*)
-
-lemma confs_rev: "P,h \<turnstile> rev s [:\<le>] t = (P,h \<turnstile> s [:\<le>] rev t)"
-(*<*)by (simp add: list_all2_rev1)(*>*)
-
-lemma confs_conv_map:
- "\<And>Ts'. P,h \<turnstile> vs [:\<le>] Ts' = (\<exists>Ts. map typeof\<^bsub>h\<^esub> vs = map Some Ts \<and> P \<turnstile> Ts [\<le>] Ts')"
-(*<*)
-proof(induct vs)
- case (Cons a vs)
- then show ?case by(case_tac Ts') (auto simp add:conf_def)
-qed simp
-(*>*)
-
-lemma confs_hext: "P,h \<turnstile> vs [:\<le>] Ts \<Longrightarrow> h \<unlhd> h' \<Longrightarrow> P,h' \<turnstile> vs [:\<le>] Ts"
-(*<*)by (erule list_all2_mono, erule conf_hext, assumption)(*>*)
-
-lemma confs_Cons2: "P,h \<turnstile> xs [:\<le>] y#ys = (\<exists>z zs. xs = z#zs \<and> P,h \<turnstile> z :\<le> y \<and> P,h \<turnstile> zs [:\<le>] ys)"
-(*<*)by (rule list_all2_Cons2)(*>*)
-
-
-subsection "Object conformance"
-
-lemma oconf_hext: "P,h \<turnstile> obj \<surd> \<Longrightarrow> h \<unlhd> h' \<Longrightarrow> P,h' \<turnstile> obj \<surd>"
-(*<*)by (fastforce elim:conf_hext simp: oconf_def)(*>*)
-
-lemma oconf_blank:
- "P \<turnstile> C has_fields FDTs \<Longrightarrow> P,h \<turnstile> blank P C \<surd>"
-(*<*)
-by(fastforce dest: has_fields_fun
- simp: has_field_def oconf_def blank_def init_fields_def
- map_of_filtered_SomeD)
-(*>*)
-
-lemma oconf_fupd [intro?]:
- "\<lbrakk> P \<turnstile> C has F,NonStatic:T in D; P,h \<turnstile> v :\<le> T; P,h \<turnstile> (C,fs) \<surd> \<rbrakk>
- \<Longrightarrow> P,h \<turnstile> (C, fs((F,D)\<mapsto>v)) \<surd>"
-(*<*)by (auto dest: has_fields_fun simp add: oconf_def has_field_def fun_upd_apply)(*>*)
-
-(*<*)
-lemmas oconf_new = oconf_hext [OF _ hext_new]
-lemmas oconf_upd_obj = oconf_hext [OF _ hext_upd_obj]
-(*>*)
-
-subsection "Static object conformance"
-
-lemma soconf_hext: "P,h,C \<turnstile>\<^sub>s obj \<surd> \<Longrightarrow> h \<unlhd> h' \<Longrightarrow> P,h',C \<turnstile>\<^sub>s obj \<surd>"
-(*<*)by (fastforce elim:conf_hext simp:soconf_def)(*>*)
-
-lemma soconf_sblank:
- "P \<turnstile> C has_fields FDTs \<Longrightarrow> P,h,C \<turnstile>\<^sub>s sblank P C \<surd>"
-(*<*)
-proof -
- let ?sfs = "sblank P C"
- assume FDTs: "P \<turnstile> C has_fields FDTs"
- then have "\<And>F T. P \<turnstile> C has F,Static:T in C
- \<Longrightarrow> \<exists>v. ?sfs F = Some v \<and> P,h \<turnstile> v :\<le> T"
- proof -
- fix F T assume has: "P \<turnstile> C has F,Static:T in C"
- with has_fields_fun[OF FDTs] have map: "map_of FDTs (F, C) = \<lfloor>(Static, T)\<rfloor>"
- by(clarsimp simp: has_field_def)
- then have "map_of (map (\<lambda>((F, D), b, T). (F, default_val T))
- (filter (\<lambda>((F, D), b, T). (case ((F, D), b, T) of (x, xa)
- \<Rightarrow> (case x of (F, D) \<Rightarrow> \<lambda>(b, T). b = Static) xa) \<and> D = C) FDTs)) F
- = \<lfloor>default_val T\<rfloor>"
- by(rule map_of_remove_filtered_SomeD[where P = "default_val"
- and Q = "\<lambda>((F, D), b, T). b = Static"]) simp
- with FDTs show "\<exists>v. ?sfs F = Some v \<and> P,h \<turnstile> v :\<le> T"
- by(clarsimp simp: sblank_def init_sfields_def)
- qed
- then show ?thesis by(simp add: soconf_def)
-qed
-(*>*)
-
-lemma soconf_fupd [intro?]:
- "\<lbrakk> P \<turnstile> C has F,Static:T in C; P,h \<turnstile> v :\<le> T; P,h,C \<turnstile>\<^sub>s sfs \<surd> \<rbrakk>
- \<Longrightarrow> P,h,C \<turnstile>\<^sub>s sfs(F\<mapsto>v) \<surd>"
-(*<*)by (auto dest: has_fields_fun simp add: fun_upd_apply soconf_def has_field_def)(*>*)
-
-(*<*)
-lemmas soconf_new = soconf_hext [OF _ hext_new]
-lemmas soconf_upd_obj = soconf_hext [OF _ hext_upd_obj]
-(*>*)
-
-subsection "Heap conformance"
-
-lemma hconfD: "\<lbrakk> P \<turnstile> h \<surd>; h a = Some obj \<rbrakk> \<Longrightarrow> P,h \<turnstile> obj \<surd>"
-(*<*)by (unfold hconf_def) fast(*>*)
-
-lemma hconf_new: "\<lbrakk> P \<turnstile> h \<surd>; h a = None; P,h \<turnstile> obj \<surd> \<rbrakk> \<Longrightarrow> P \<turnstile> h(a\<mapsto>obj) \<surd>"
-(*<*)by (unfold hconf_def) (auto intro: oconf_new preallocated_new)(*>*)
-
-lemma hconf_upd_obj: "\<lbrakk> P \<turnstile> h\<surd>; h a = Some(C,fs); P,h \<turnstile> (C,fs')\<surd> \<rbrakk> \<Longrightarrow> P \<turnstile> h(a\<mapsto>(C,fs'))\<surd>"
-(*<*)by (unfold hconf_def) (auto intro: oconf_upd_obj preallocated_upd_obj)(*>*)
-
-
-subsection "Class statics conformance"
-
-lemma shconfD: "\<lbrakk> P,h \<turnstile>\<^sub>s sh \<surd>; sh C = Some(sfs,i) \<rbrakk> \<Longrightarrow> P,h,C \<turnstile>\<^sub>s sfs \<surd>"
-(*<*)by (unfold shconf_def) fast(*>*)
-
-lemma shconf_upd_obj: "\<lbrakk> P,h \<turnstile>\<^sub>s sh \<surd>; P,h,C \<turnstile>\<^sub>s sfs' \<surd> \<rbrakk>
- \<Longrightarrow> P,h \<turnstile>\<^sub>s sh(C\<mapsto>(sfs',i'))\<surd>"
-(*<*)by (unfold shconf_def) (auto intro: soconf_upd_obj preallocated_upd_obj)(*>*)
-
-lemma shconf_hnew: "\<lbrakk> P,h \<turnstile>\<^sub>s sh \<surd>; h a = None \<rbrakk> \<Longrightarrow> P,h(a\<mapsto>obj) \<turnstile>\<^sub>s sh \<surd>"
-(*<*)by (unfold shconf_def) (auto intro: soconf_new preallocated_new)(*>*)
-
-lemma shconf_hupd_obj: "\<lbrakk> P,h \<turnstile>\<^sub>s sh \<surd>; h a = Some(C,fs) \<rbrakk> \<Longrightarrow> P,h(a\<mapsto>(C,fs')) \<turnstile>\<^sub>s sh \<surd>"
-(*<*)by (unfold shconf_def) (auto intro: soconf_upd_obj preallocated_upd_obj)(*>*)
-
-subsection "Local variable conformance"
-
-lemma lconf_hext: "\<lbrakk> P,h \<turnstile> l (:\<le>) E; h \<unlhd> h' \<rbrakk> \<Longrightarrow> P,h' \<turnstile> l (:\<le>) E"
-(*<*)by (unfold lconf_def) (fast elim: conf_hext)(*>*)
-
-lemma lconf_upd:
- "\<lbrakk> P,h \<turnstile> l (:\<le>) E; P,h \<turnstile> v :\<le> T; E V = Some T \<rbrakk> \<Longrightarrow> P,h \<turnstile> l(V\<mapsto>v) (:\<le>) E"
-(*<*)by (unfold lconf_def) auto(*>*)
-
-lemma lconf_empty[iff]: "P,h \<turnstile> Map.empty (:\<le>) E"
-(*<*)by(simp add:lconf_def)(*>*)
-
-lemma lconf_upd2: "\<lbrakk>P,h \<turnstile> l (:\<le>) E; P,h \<turnstile> v :\<le> T\<rbrakk> \<Longrightarrow> P,h \<turnstile> l(V\<mapsto>v) (:\<le>) E(V\<mapsto>T)"
-(*<*)by(simp add:lconf_def)(*>*)
-
-
-end
+(* Title: JinjaDCI/Common/Conform.thy
+
+ Author: David von Oheimb, Tobias Nipkow, Susannah Mansky
+ Copyright 1999 Technische Universitaet Muenchen, 2019-20 UIUC
+
+ Based on the Jinja theory Common/Conform.thy by David von Oheimb and Tobias Nipkow
+*)
+
+section \<open> Conformance Relations for Type Soundness Proofs \<close>
+
+theory Conform
+imports Exceptions
+begin
+
+definition conf :: "'m prog \<Rightarrow> heap \<Rightarrow> val \<Rightarrow> ty \<Rightarrow> bool" ("_,_ \<turnstile> _ :\<le> _" [51,51,51,51] 50)
+where
+ "P,h \<turnstile> v :\<le> T \<equiv>
+ \<exists>T'. typeof\<^bsub>h\<^esub> v = Some T' \<and> P \<turnstile> T' \<le> T"
+
+definition oconf :: "'m prog \<Rightarrow> heap \<Rightarrow> obj \<Rightarrow> bool" ("_,_ \<turnstile> _ \<surd>" [51,51,51] 50)
+where
+ "P,h \<turnstile> obj \<surd> \<equiv>
+ let (C,fs) = obj in \<forall>F D T. P \<turnstile> C has F,NonStatic:T in D \<longrightarrow>
+ (\<exists>v. fs(F,D) = Some v \<and> P,h \<turnstile> v :\<le> T)"
+
+definition soconf :: "'m prog \<Rightarrow> heap \<Rightarrow> cname \<Rightarrow> sfields \<Rightarrow> bool" ("_,_,_ \<turnstile>\<^sub>s _ \<surd>" [51,51,51,51] 50)
+where
+ "P,h,C \<turnstile>\<^sub>s sfs \<surd> \<equiv>
+ \<forall>F T. P \<turnstile> C has F,Static:T in C \<longrightarrow>
+ (\<exists>v. sfs F = Some v \<and> P,h \<turnstile> v :\<le> T)"
+
+definition hconf :: "'m prog \<Rightarrow> heap \<Rightarrow> bool" ("_ \<turnstile> _ \<surd>" [51,51] 50)
+where
+ "P \<turnstile> h \<surd> \<equiv>
+ (\<forall>a obj. h a = Some obj \<longrightarrow> P,h \<turnstile> obj \<surd>) \<and> preallocated h"
+
+definition shconf :: "'m prog \<Rightarrow> heap \<Rightarrow> sheap \<Rightarrow> bool" ("_,_ \<turnstile>\<^sub>s _ \<surd>" [51,51,51] 50)
+where
+ "P,h \<turnstile>\<^sub>s sh \<surd> \<equiv>
+ (\<forall>C sfs i. sh C = Some(sfs,i) \<longrightarrow> P,h,C \<turnstile>\<^sub>s sfs \<surd>)"
+
+definition lconf :: "'m prog \<Rightarrow> heap \<Rightarrow> (vname \<rightharpoonup> val) \<Rightarrow> (vname \<rightharpoonup> ty) \<Rightarrow> bool" ("_,_ \<turnstile> _ '(:\<le>') _" [51,51,51,51] 50)
+where
+ "P,h \<turnstile> l (:\<le>) E \<equiv>
+ \<forall>V v. l V = Some v \<longrightarrow> (\<exists>T. E V = Some T \<and> P,h \<turnstile> v :\<le> T)"
+
+abbreviation
+ confs :: "'m prog \<Rightarrow> heap \<Rightarrow> val list \<Rightarrow> ty list \<Rightarrow> bool"
+ ("_,_ \<turnstile> _ [:\<le>] _" [51,51,51,51] 50) where
+ "P,h \<turnstile> vs [:\<le>] Ts \<equiv> list_all2 (conf P h) vs Ts"
+
+
+subsection\<open> Value conformance @{text":\<le>"} \<close>
+
+lemma conf_Null [simp]: "P,h \<turnstile> Null :\<le> T = P \<turnstile> NT \<le> T"
+(*<*)by (simp (no_asm) add: conf_def)(*>*)
+
+lemma typeof_conf[simp]: "typeof\<^bsub>h\<^esub> v = Some T \<Longrightarrow> P,h \<turnstile> v :\<le> T"
+(*<*)by (induct v) (auto simp: conf_def)(*>*)
+
+lemma typeof_lit_conf[simp]: "typeof v = Some T \<Longrightarrow> P,h \<turnstile> v :\<le> T"
+(*<*)by (rule typeof_conf[OF typeof_lit_typeof])(*>*)
+
+lemma defval_conf[simp]: "P,h \<turnstile> default_val T :\<le> T"
+(*<*)by (cases T) (auto simp: conf_def)(*>*)
+
+lemma conf_upd_obj: "h a = Some(C,fs) \<Longrightarrow> (P,h(a\<mapsto>(C,fs')) \<turnstile> x :\<le> T) = (P,h \<turnstile> x :\<le> T)"
+(*<*)by (rule val.induct) (auto simp:conf_def fun_upd_apply)(*>*)
+
+lemma conf_widen: "P,h \<turnstile> v :\<le> T \<Longrightarrow> P \<turnstile> T \<le> T' \<Longrightarrow> P,h \<turnstile> v :\<le> T'"
+(*<*)by (induct v) (auto intro: widen_trans simp: conf_def)(*>*)
+
+lemma conf_hext: "h \<unlhd> h' \<Longrightarrow> P,h \<turnstile> v :\<le> T \<Longrightarrow> P,h' \<turnstile> v :\<le> T"
+(*<*)by (induct v) (auto dest: hext_objD simp: conf_def)(*>*)
+
+lemma conf_ClassD: "P,h \<turnstile> v :\<le> Class C \<Longrightarrow>
+ v = Null \<or> (\<exists>a obj T. v = Addr a \<and> h a = Some obj \<and> obj_ty obj = T \<and> P \<turnstile> T \<le> Class C)"
+(*<*)by(induct v) (auto simp: conf_def)(*>*)
+
+lemma conf_NT [iff]: "P,h \<turnstile> v :\<le> NT = (v = Null)"
+(*<*)by (auto simp add: conf_def)(*>*)
+
+lemma non_npD: "\<lbrakk> v \<noteq> Null; P,h \<turnstile> v :\<le> Class C \<rbrakk>
+ \<Longrightarrow> \<exists>a C' fs. v = Addr a \<and> h a = Some(C',fs) \<and> P \<turnstile> C' \<preceq>\<^sup>* C"
+(*<*)by (auto dest: conf_ClassD)(*>*)
+
+
+subsection\<open> Value list conformance @{text"[:\<le>]"} \<close>
+
+lemma confs_widens [trans]: "\<lbrakk>P,h \<turnstile> vs [:\<le>] Ts; P \<turnstile> Ts [\<le>] Ts'\<rbrakk> \<Longrightarrow> P,h \<turnstile> vs [:\<le>] Ts'"
+(*<*)by(auto intro: list_all2_trans conf_widen)(*>*)
+
+lemma confs_rev: "P,h \<turnstile> rev s [:\<le>] t = (P,h \<turnstile> s [:\<le>] rev t)"
+(*<*)by (simp add: list_all2_rev1)(*>*)
+
+lemma confs_conv_map:
+ "\<And>Ts'. P,h \<turnstile> vs [:\<le>] Ts' = (\<exists>Ts. map typeof\<^bsub>h\<^esub> vs = map Some Ts \<and> P \<turnstile> Ts [\<le>] Ts')"
+(*<*)
+proof(induct vs)
+ case (Cons a vs)
+ then show ?case by(case_tac Ts') (auto simp add:conf_def)
+qed simp
+(*>*)
+
+lemma confs_hext: "P,h \<turnstile> vs [:\<le>] Ts \<Longrightarrow> h \<unlhd> h' \<Longrightarrow> P,h' \<turnstile> vs [:\<le>] Ts"
+(*<*)by (erule list_all2_mono, erule conf_hext, assumption)(*>*)
+
+lemma confs_Cons2: "P,h \<turnstile> xs [:\<le>] y#ys = (\<exists>z zs. xs = z#zs \<and> P,h \<turnstile> z :\<le> y \<and> P,h \<turnstile> zs [:\<le>] ys)"
+(*<*)by (rule list_all2_Cons2)(*>*)
+
+
+subsection "Object conformance"
+
+lemma oconf_hext: "P,h \<turnstile> obj \<surd> \<Longrightarrow> h \<unlhd> h' \<Longrightarrow> P,h' \<turnstile> obj \<surd>"
+(*<*)by (fastforce elim:conf_hext simp: oconf_def)(*>*)
+
+lemma oconf_blank:
+ "P \<turnstile> C has_fields FDTs \<Longrightarrow> P,h \<turnstile> blank P C \<surd>"
+(*<*)
+by(fastforce dest: has_fields_fun
+ simp: has_field_def oconf_def blank_def init_fields_def
+ map_of_filtered_SomeD)
+(*>*)
+
+lemma oconf_fupd [intro?]:
+ "\<lbrakk> P \<turnstile> C has F,NonStatic:T in D; P,h \<turnstile> v :\<le> T; P,h \<turnstile> (C,fs) \<surd> \<rbrakk>
+ \<Longrightarrow> P,h \<turnstile> (C, fs((F,D)\<mapsto>v)) \<surd>"
+(*<*)by (auto dest: has_fields_fun simp add: oconf_def has_field_def fun_upd_apply)(*>*)
+
+(*<*)
+lemmas oconf_new = oconf_hext [OF _ hext_new]
+lemmas oconf_upd_obj = oconf_hext [OF _ hext_upd_obj]
+(*>*)
+
+subsection "Static object conformance"
+
+lemma soconf_hext: "P,h,C \<turnstile>\<^sub>s obj \<surd> \<Longrightarrow> h \<unlhd> h' \<Longrightarrow> P,h',C \<turnstile>\<^sub>s obj \<surd>"
+(*<*)by (fastforce elim:conf_hext simp:soconf_def)(*>*)
+
+lemma soconf_sblank:
+ "P \<turnstile> C has_fields FDTs \<Longrightarrow> P,h,C \<turnstile>\<^sub>s sblank P C \<surd>"
+(*<*)
+proof -
+ let ?sfs = "sblank P C"
+ assume FDTs: "P \<turnstile> C has_fields FDTs"
+ then have "\<And>F T. P \<turnstile> C has F,Static:T in C
+ \<Longrightarrow> \<exists>v. ?sfs F = Some v \<and> P,h \<turnstile> v :\<le> T"
+ proof -
+ fix F T assume has: "P \<turnstile> C has F,Static:T in C"
+ with has_fields_fun[OF FDTs] have map: "map_of FDTs (F, C) = \<lfloor>(Static, T)\<rfloor>"
+ by(clarsimp simp: has_field_def)
+ then have "map_of (map (\<lambda>((F, D), b, T). (F, default_val T))
+ (filter (\<lambda>((F, D), b, T). (case ((F, D), b, T) of (x, xa)
+ \<Rightarrow> (case x of (F, D) \<Rightarrow> \<lambda>(b, T). b = Static) xa) \<and> D = C) FDTs)) F
+ = \<lfloor>default_val T\<rfloor>"
+ by(rule map_of_remove_filtered_SomeD[where P = "default_val"
+ and Q = "\<lambda>((F, D), b, T). b = Static"]) simp
+ with FDTs show "\<exists>v. ?sfs F = Some v \<and> P,h \<turnstile> v :\<le> T"
+ by(clarsimp simp: sblank_def init_sfields_def)
+ qed
+ then show ?thesis by(simp add: soconf_def)
+qed
+(*>*)
+
+lemma soconf_fupd [intro?]:
+ "\<lbrakk> P \<turnstile> C has F,Static:T in C; P,h \<turnstile> v :\<le> T; P,h,C \<turnstile>\<^sub>s sfs \<surd> \<rbrakk>
+ \<Longrightarrow> P,h,C \<turnstile>\<^sub>s sfs(F\<mapsto>v) \<surd>"
+(*<*)by (auto dest: has_fields_fun simp add: fun_upd_apply soconf_def has_field_def)(*>*)
+
+(*<*)
+lemmas soconf_new = soconf_hext [OF _ hext_new]
+lemmas soconf_upd_obj = soconf_hext [OF _ hext_upd_obj]
+(*>*)
+
+subsection "Heap conformance"
+
+lemma hconfD: "\<lbrakk> P \<turnstile> h \<surd>; h a = Some obj \<rbrakk> \<Longrightarrow> P,h \<turnstile> obj \<surd>"
+(*<*)by (unfold hconf_def) fast(*>*)
+
+lemma hconf_new: "\<lbrakk> P \<turnstile> h \<surd>; h a = None; P,h \<turnstile> obj \<surd> \<rbrakk> \<Longrightarrow> P \<turnstile> h(a\<mapsto>obj) \<surd>"
+(*<*)by (unfold hconf_def) (auto intro: oconf_new preallocated_new)(*>*)
+
+lemma hconf_upd_obj: "\<lbrakk> P \<turnstile> h\<surd>; h a = Some(C,fs); P,h \<turnstile> (C,fs')\<surd> \<rbrakk> \<Longrightarrow> P \<turnstile> h(a\<mapsto>(C,fs'))\<surd>"
+(*<*)by (unfold hconf_def) (auto intro: oconf_upd_obj preallocated_upd_obj)(*>*)
+
+
+subsection "Class statics conformance"
+
+lemma shconfD: "\<lbrakk> P,h \<turnstile>\<^sub>s sh \<surd>; sh C = Some(sfs,i) \<rbrakk> \<Longrightarrow> P,h,C \<turnstile>\<^sub>s sfs \<surd>"
+(*<*)by (unfold shconf_def) fast(*>*)
+
+lemma shconf_upd_obj: "\<lbrakk> P,h \<turnstile>\<^sub>s sh \<surd>; P,h,C \<turnstile>\<^sub>s sfs' \<surd> \<rbrakk>
+ \<Longrightarrow> P,h \<turnstile>\<^sub>s sh(C\<mapsto>(sfs',i'))\<surd>"
+(*<*)by (unfold shconf_def) (auto intro: soconf_upd_obj preallocated_upd_obj)(*>*)
+
+lemma shconf_hnew: "\<lbrakk> P,h \<turnstile>\<^sub>s sh \<surd>; h a = None \<rbrakk> \<Longrightarrow> P,h(a\<mapsto>obj) \<turnstile>\<^sub>s sh \<surd>"
+(*<*)by (unfold shconf_def) (auto intro: soconf_new preallocated_new)(*>*)
+
+lemma shconf_hupd_obj: "\<lbrakk> P,h \<turnstile>\<^sub>s sh \<surd>; h a = Some(C,fs) \<rbrakk> \<Longrightarrow> P,h(a\<mapsto>(C,fs')) \<turnstile>\<^sub>s sh \<surd>"
+(*<*)by (unfold shconf_def) (auto intro: soconf_upd_obj preallocated_upd_obj)(*>*)
+
+subsection "Local variable conformance"
+
+lemma lconf_hext: "\<lbrakk> P,h \<turnstile> l (:\<le>) E; h \<unlhd> h' \<rbrakk> \<Longrightarrow> P,h' \<turnstile> l (:\<le>) E"
+(*<*)by (unfold lconf_def) (fast elim: conf_hext)(*>*)
+
+lemma lconf_upd:
+ "\<lbrakk> P,h \<turnstile> l (:\<le>) E; P,h \<turnstile> v :\<le> T; E V = Some T \<rbrakk> \<Longrightarrow> P,h \<turnstile> l(V\<mapsto>v) (:\<le>) E"
+(*<*)by (unfold lconf_def) auto(*>*)
+
+lemma lconf_empty[iff]: "P,h \<turnstile> Map.empty (:\<le>) E"
+(*<*)by(simp add:lconf_def)(*>*)
+
+lemma lconf_upd2: "\<lbrakk>P,h \<turnstile> l (:\<le>) E; P,h \<turnstile> v :\<le> T\<rbrakk> \<Longrightarrow> P,h \<turnstile> l(V\<mapsto>v) (:\<le>) E(V\<mapsto>T)"
+(*<*)by(simp add:lconf_def)(*>*)
+
+
+end
diff --git a/thys/JinjaDCI/Common/Decl.thy b/thys/JinjaDCI/Common/Decl.thy
--- a/thys/JinjaDCI/Common/Decl.thy
+++ b/thys/JinjaDCI/Common/Decl.thy
@@ -1,85 +1,85 @@
-(* Title: JinjaDCI/Common/Decl.thy
-
- Author: David von Oheimb, Susannah Mansky
- Copyright 1999 Technische Universitaet Muenchen, 2019-20 UIUC
-
- Based on the Jinja theory Common/Decl.thy by David von Oheimb
-*)
-
-section \<open> Class Declarations and Programs \<close>
-
-theory Decl imports Type begin
-
-type_synonym
- fdecl = "vname \<times> staticb \<times> ty" \<comment> \<open>field declaration\<close>
-type_synonym
- 'm mdecl = "mname \<times> staticb \<times> ty list \<times> ty \<times> 'm" \<comment> \<open>method = name, static flag, arg.\ types, return type, body\<close>
-type_synonym
- 'm "class" = "cname \<times> fdecl list \<times> 'm mdecl list" \<comment> \<open>class = superclass, fields, methods\<close>
-type_synonym
- 'm cdecl = "cname \<times> 'm class" \<comment> \<open>class declaration\<close>
-type_synonym
- 'm prog = "'m cdecl list" \<comment> \<open>program\<close>
-
-(* replaced all fname, mname, cname in below with `char list' so that
- pretty printing works -SM *)
-(*<*)
-translations
- (type) "fdecl" <= (type) "char list \<times> staticb \<times> ty"
- (type) "'c mdecl" <= (type) "char list \<times> staticb \<times> ty list \<times> ty \<times> 'c"
- (type) "'c class" <= (type) "char list \<times> fdecl list \<times> ('c mdecl) list"
- (type) "'c cdecl" <= (type) "char list \<times> ('c class)"
- (type) "'c prog" <= (type) "('c cdecl) list"
-(*>*)
-
-definition "class" :: "'m prog \<Rightarrow> cname \<rightharpoonup> 'm class"
-where
- "class \<equiv> map_of"
-
-(* Not difficult to prove, but useful for directing particular sequences of equality -SM *)
-lemma class_cons: "\<lbrakk> C \<noteq> fst x \<rbrakk> \<Longrightarrow> class (x # P) C = class P C"
- by (simp add: class_def)
-
-definition is_class :: "'m prog \<Rightarrow> cname \<Rightarrow> bool"
-where
- "is_class P C \<equiv> class P C \<noteq> None"
-
-lemma finite_is_class: "finite {C. is_class P C}"
-(*<*)
-proof -
- have "{C. is_class P C} = dom (map_of P)"
- by (simp add: is_class_def class_def dom_def)
- thus ?thesis by (simp add: finite_dom_map_of)
-qed
-(*>*)
-
-definition is_type :: "'m prog \<Rightarrow> ty \<Rightarrow> bool"
-where
- "is_type P T \<equiv>
- (case T of Void \<Rightarrow> True | Boolean \<Rightarrow> True | Integer \<Rightarrow> True | NT \<Rightarrow> True
- | Class C \<Rightarrow> is_class P C)"
-
-lemma is_type_simps [simp]:
- "is_type P Void \<and> is_type P Boolean \<and> is_type P Integer \<and>
- is_type P NT \<and> is_type P (Class C) = is_class P C"
-(*<*)by(simp add:is_type_def)(*>*)
-
-
-abbreviation
- "types P == Collect (is_type P)"
-
-lemma class_exists_equiv:
- "(\<exists>x. fst x = cn \<and> x \<in> set P) = (class P cn \<noteq> None)"
-proof(rule iffI)
- assume "\<exists>x. fst x = cn \<and> x \<in> set P" then show "class P cn \<noteq> None"
- by (metis class_def image_eqI map_of_eq_None_iff)
-next
- assume "class P cn \<noteq> None" then show "\<exists>x. fst x = cn \<and> x \<in> set P"
- by (metis class_def fst_conv map_of_SomeD option.exhaust)
-qed
-
-lemma class_exists_equiv2:
- "(\<exists>x. fst x = cn \<and> x \<in> set (P1 @ P2)) = (class P1 cn \<noteq> None \<or> class P2 cn \<noteq> None)"
-by (simp only: class_exists_equiv [where P = "P1@P2"], simp add: class_def)
-
-end
+(* Title: JinjaDCI/Common/Decl.thy
+
+ Author: David von Oheimb, Susannah Mansky
+ Copyright 1999 Technische Universitaet Muenchen, 2019-20 UIUC
+
+ Based on the Jinja theory Common/Decl.thy by David von Oheimb
+*)
+
+section \<open> Class Declarations and Programs \<close>
+
+theory Decl imports Type begin
+
+type_synonym
+ fdecl = "vname \<times> staticb \<times> ty" \<comment> \<open>field declaration\<close>
+type_synonym
+ 'm mdecl = "mname \<times> staticb \<times> ty list \<times> ty \<times> 'm" \<comment> \<open>method = name, static flag, arg.\ types, return type, body\<close>
+type_synonym
+ 'm "class" = "cname \<times> fdecl list \<times> 'm mdecl list" \<comment> \<open>class = superclass, fields, methods\<close>
+type_synonym
+ 'm cdecl = "cname \<times> 'm class" \<comment> \<open>class declaration\<close>
+type_synonym
+ 'm prog = "'m cdecl list" \<comment> \<open>program\<close>
+
+(* replaced all fname, mname, cname in below with `char list' so that
+ pretty printing works -SM *)
+(*<*)
+translations
+ (type) "fdecl" <= (type) "char list \<times> staticb \<times> ty"
+ (type) "'c mdecl" <= (type) "char list \<times> staticb \<times> ty list \<times> ty \<times> 'c"
+ (type) "'c class" <= (type) "char list \<times> fdecl list \<times> ('c mdecl) list"
+ (type) "'c cdecl" <= (type) "char list \<times> ('c class)"
+ (type) "'c prog" <= (type) "('c cdecl) list"
+(*>*)
+
+definition "class" :: "'m prog \<Rightarrow> cname \<rightharpoonup> 'm class"
+where
+ "class \<equiv> map_of"
+
+(* Not difficult to prove, but useful for directing particular sequences of equality -SM *)
+lemma class_cons: "\<lbrakk> C \<noteq> fst x \<rbrakk> \<Longrightarrow> class (x # P) C = class P C"
+ by (simp add: class_def)
+
+definition is_class :: "'m prog \<Rightarrow> cname \<Rightarrow> bool"
+where
+ "is_class P C \<equiv> class P C \<noteq> None"
+
+lemma finite_is_class: "finite {C. is_class P C}"
+(*<*)
+proof -
+ have "{C. is_class P C} = dom (map_of P)"
+ by (simp add: is_class_def class_def dom_def)
+ thus ?thesis by (simp add: finite_dom_map_of)
+qed
+(*>*)
+
+definition is_type :: "'m prog \<Rightarrow> ty \<Rightarrow> bool"
+where
+ "is_type P T \<equiv>
+ (case T of Void \<Rightarrow> True | Boolean \<Rightarrow> True | Integer \<Rightarrow> True | NT \<Rightarrow> True
+ | Class C \<Rightarrow> is_class P C)"
+
+lemma is_type_simps [simp]:
+ "is_type P Void \<and> is_type P Boolean \<and> is_type P Integer \<and>
+ is_type P NT \<and> is_type P (Class C) = is_class P C"
+(*<*)by(simp add:is_type_def)(*>*)
+
+
+abbreviation
+ "types P == Collect (is_type P)"
+
+lemma class_exists_equiv:
+ "(\<exists>x. fst x = cn \<and> x \<in> set P) = (class P cn \<noteq> None)"
+proof(rule iffI)
+ assume "\<exists>x. fst x = cn \<and> x \<in> set P" then show "class P cn \<noteq> None"
+ by (metis class_def image_eqI map_of_eq_None_iff)
+next
+ assume "class P cn \<noteq> None" then show "\<exists>x. fst x = cn \<and> x \<in> set P"
+ by (metis class_def fst_conv map_of_SomeD option.exhaust)
+qed
+
+lemma class_exists_equiv2:
+ "(\<exists>x. fst x = cn \<and> x \<in> set (P1 @ P2)) = (class P1 cn \<noteq> None \<or> class P2 cn \<noteq> None)"
+by (simp only: class_exists_equiv [where P = "P1@P2"], simp add: class_def)
+
+end
diff --git a/thys/JinjaDCI/Common/Exceptions.thy b/thys/JinjaDCI/Common/Exceptions.thy
--- a/thys/JinjaDCI/Common/Exceptions.thy
+++ b/thys/JinjaDCI/Common/Exceptions.thy
@@ -1,188 +1,188 @@
-(* Title: JinjaDCI/Common/Exceptions.thy
-
- Author: Gerwin Klein, Martin Strecker, Susannah Mansky
- Copyright 2002 Technische Universitaet Muenchen, 2019-20 UIUC
-
- Based on the Jinja theory Common/Exceptions.thy by Gerwin Klein and Martin Strecker
-*)
-
-section \<open> Exceptions \<close>
-
-theory Exceptions imports Objects begin
-
-definition ErrorCl :: "string" where "ErrorCl = ''Error''"
-definition ThrowCl :: "string" where "ThrowCl = ''Throwable''"
-
-definition NullPointer :: cname
-where
- "NullPointer \<equiv> ''NullPointer''"
-
-definition ClassCast :: cname
-where
- "ClassCast \<equiv> ''ClassCast''"
-
-definition OutOfMemory :: cname
-where
- "OutOfMemory \<equiv> ''OutOfMemory''"
-
-definition NoClassDefFoundError :: cname
-where
- "NoClassDefFoundError \<equiv> ''NoClassDefFoundError''"
-
-definition IncompatibleClassChangeError :: cname
-where
- "IncompatibleClassChangeError \<equiv> ''IncompatibleClassChangeError''"
-
-definition NoSuchFieldError :: cname
-where
- "NoSuchFieldError \<equiv> ''NoSuchFieldError''"
-
-definition NoSuchMethodError :: cname
-where
- "NoSuchMethodError \<equiv> ''NoSuchMethodError''"
-
-definition sys_xcpts :: "cname set"
-where
- "sys_xcpts \<equiv> {NullPointer, ClassCast, OutOfMemory, NoClassDefFoundError,
- IncompatibleClassChangeError,
- NoSuchFieldError, NoSuchMethodError}"
-
-definition addr_of_sys_xcpt :: "cname \<Rightarrow> addr"
-where
- "addr_of_sys_xcpt s \<equiv> if s = NullPointer then 0 else
- if s = ClassCast then 1 else
- if s = OutOfMemory then 2 else
- if s = NoClassDefFoundError then 3 else
- if s = IncompatibleClassChangeError then 4 else
- if s = NoSuchFieldError then 5 else
- if s = NoSuchMethodError then 6 else undefined"
-
-
-lemmas sys_xcpts_defs = NullPointer_def ClassCast_def OutOfMemory_def NoClassDefFoundError_def
- IncompatibleClassChangeError_def NoSuchFieldError_def NoSuchMethodError_def
-
-lemma Start_nsys_xcpts: "Start \<notin> sys_xcpts"
- by(simp add: Start_def sys_xcpts_def sys_xcpts_defs)
-
-lemma Start_nsys_xcpts1 [simp]: "Start \<noteq> NullPointer" "Start \<noteq> ClassCast"
- "Start \<noteq> OutOfMemory" "Start \<noteq> NoClassDefFoundError"
- "Start \<noteq> IncompatibleClassChangeError" "Start \<noteq> NoSuchFieldError"
- "Start \<noteq> NoSuchMethodError"
-using Start_nsys_xcpts by(auto simp: sys_xcpts_def)
-
-lemma Start_nsys_xcpts2 [simp]: "NullPointer \<noteq> Start" "ClassCast \<noteq> Start"
- "OutOfMemory \<noteq> Start" "NoClassDefFoundError \<noteq> Start"
- "IncompatibleClassChangeError \<noteq> Start" "NoSuchFieldError \<noteq> Start"
- "NoSuchMethodError \<noteq> Start"
-using Start_nsys_xcpts by(auto simp: sys_xcpts_def dest: sym)
-
-definition start_heap :: "'c prog \<Rightarrow> heap"
-where
- "start_heap G \<equiv> Map.empty (addr_of_sys_xcpt NullPointer \<mapsto> blank G NullPointer)
- (addr_of_sys_xcpt ClassCast \<mapsto> blank G ClassCast)
- (addr_of_sys_xcpt OutOfMemory \<mapsto> blank G OutOfMemory)
- (addr_of_sys_xcpt NoClassDefFoundError \<mapsto> blank G NoClassDefFoundError)
- (addr_of_sys_xcpt IncompatibleClassChangeError \<mapsto> blank G IncompatibleClassChangeError)
- (addr_of_sys_xcpt NoSuchFieldError \<mapsto> blank G NoSuchFieldError)
- (addr_of_sys_xcpt NoSuchMethodError \<mapsto> blank G NoSuchMethodError)"
-
-definition preallocated :: "heap \<Rightarrow> bool"
-where
- "preallocated h \<equiv> \<forall>C \<in> sys_xcpts. \<exists>fs. h(addr_of_sys_xcpt C) = Some (C,fs)"
-
-subsection "System exceptions"
-
-lemma sys_xcpts_incl [simp]: "NullPointer \<in> sys_xcpts \<and> OutOfMemory \<in> sys_xcpts
- \<and> ClassCast \<in> sys_xcpts \<and> NoClassDefFoundError \<in> sys_xcpts
- \<and> IncompatibleClassChangeError \<in> sys_xcpts \<and> NoSuchFieldError \<in> sys_xcpts
- \<and> NoSuchMethodError \<in> sys_xcpts"
-(*<*)by(simp add: sys_xcpts_def)(*>*)
-
-lemma sys_xcpts_cases [consumes 1, cases set]:
- "\<lbrakk> C \<in> sys_xcpts; P NullPointer; P OutOfMemory; P ClassCast; P NoClassDefFoundError;
- P IncompatibleClassChangeError; P NoSuchFieldError;
- P NoSuchMethodError \<rbrakk> \<Longrightarrow> P C"
-(*<*)by (auto simp: sys_xcpts_def)(*>*)
-
-subsection "Starting heap"
-
-lemma start_heap_sys_xcpts:
-assumes "C \<in> sys_xcpts"
-shows "start_heap P (addr_of_sys_xcpt C) = Some(blank P C)"
-by(rule sys_xcpts_cases[OF assms])
- (auto simp add: start_heap_def sys_xcpts_def addr_of_sys_xcpt_def sys_xcpts_defs)
-
-lemma start_heap_classes:
- "start_heap P a = Some(C,fs) \<Longrightarrow> C \<in> sys_xcpts"
- by(simp add: start_heap_def blank_def split: if_split_asm)
-
-lemma start_heap_nStart: "start_heap P a = Some obj \<Longrightarrow> fst(obj) \<noteq> Start"
- by(cases obj, auto dest!: start_heap_classes simp: Start_nsys_xcpts)
-
-subsection "@{term preallocated}"
-
-lemma preallocated_dom [simp]:
- "\<lbrakk> preallocated h; C \<in> sys_xcpts \<rbrakk> \<Longrightarrow> addr_of_sys_xcpt C \<in> dom h"
-(*<*)by (fastforce simp:preallocated_def dom_def)(*>*)
-
-
-lemma preallocatedD:
- "\<lbrakk> preallocated h; C \<in> sys_xcpts \<rbrakk> \<Longrightarrow> \<exists>fs. h(addr_of_sys_xcpt C) = Some (C, fs)"
-(*<*)by(auto simp: preallocated_def sys_xcpts_def)(*>*)
-
-
-lemma preallocatedE [elim?]:
- "\<lbrakk> preallocated h; C \<in> sys_xcpts; \<And>fs. h(addr_of_sys_xcpt C) = Some(C,fs) \<Longrightarrow> P h C\<rbrakk>
- \<Longrightarrow> P h C"
-(*<*)by (fast dest: preallocatedD)(*>*)
-
-
-lemma cname_of_xcp [simp]:
- "\<lbrakk> preallocated h; C \<in> sys_xcpts \<rbrakk> \<Longrightarrow> cname_of h (addr_of_sys_xcpt C) = C"
-(*<*)by (auto elim: preallocatedE)(*>*)
-
-
-lemma typeof_ClassCast [simp]:
- "preallocated h \<Longrightarrow> typeof\<^bsub>h\<^esub> (Addr(addr_of_sys_xcpt ClassCast)) = Some(Class ClassCast)"
-(*<*)by (auto elim: preallocatedE)(*>*)
-
-
-lemma typeof_OutOfMemory [simp]:
- "preallocated h \<Longrightarrow> typeof\<^bsub>h\<^esub> (Addr(addr_of_sys_xcpt OutOfMemory)) = Some(Class OutOfMemory)"
-(*<*)by (auto elim: preallocatedE)(*>*)
-
-
-lemma typeof_NullPointer [simp]:
- "preallocated h \<Longrightarrow> typeof\<^bsub>h\<^esub> (Addr(addr_of_sys_xcpt NullPointer)) = Some(Class NullPointer)"
-(*<*)by (auto elim: preallocatedE)(*>*)
-
-lemma typeof_NoClassDefFoundError [simp]:
- "preallocated h \<Longrightarrow> typeof\<^bsub>h\<^esub> (Addr(addr_of_sys_xcpt NoClassDefFoundError)) = Some(Class NoClassDefFoundError)"
-(*<*)by (auto elim: preallocatedE)(*>*)
-
-lemma typeof_IncompatibleClassChangeError [simp]:
- "preallocated h \<Longrightarrow> typeof\<^bsub>h\<^esub> (Addr(addr_of_sys_xcpt IncompatibleClassChangeError)) = Some(Class IncompatibleClassChangeError)"
-(*<*)by (auto elim: preallocatedE)(*>*)
-
-lemma typeof_NoSuchFieldError [simp]:
- "preallocated h \<Longrightarrow> typeof\<^bsub>h\<^esub> (Addr(addr_of_sys_xcpt NoSuchFieldError)) = Some(Class NoSuchFieldError)"
-(*<*)by (auto elim: preallocatedE)(*>*)
-
-lemma typeof_NoSuchMethodError [simp]:
- "preallocated h \<Longrightarrow> typeof\<^bsub>h\<^esub> (Addr(addr_of_sys_xcpt NoSuchMethodError)) = Some(Class NoSuchMethodError)"
-(*<*)by (auto elim: preallocatedE)(*>*)
-
-lemma preallocated_hext:
- "\<lbrakk> preallocated h; h \<unlhd> h' \<rbrakk> \<Longrightarrow> preallocated h'"
-(*<*)by (simp add: preallocated_def hext_def)(*>*)
-
-(*<*)
-lemmas preallocated_upd_obj = preallocated_hext [OF _ hext_upd_obj]
-lemmas preallocated_new = preallocated_hext [OF _ hext_new]
-(*>*)
-
-lemma preallocated_start:
- "preallocated (start_heap P)"
- by(auto simp: start_heap_sys_xcpts blank_def preallocated_def)
-
-end
+(* Title: JinjaDCI/Common/Exceptions.thy
+
+ Author: Gerwin Klein, Martin Strecker, Susannah Mansky
+ Copyright 2002 Technische Universitaet Muenchen, 2019-20 UIUC
+
+ Based on the Jinja theory Common/Exceptions.thy by Gerwin Klein and Martin Strecker
+*)
+
+section \<open> Exceptions \<close>
+
+theory Exceptions imports Objects begin
+
+definition ErrorCl :: "string" where "ErrorCl = ''Error''"
+definition ThrowCl :: "string" where "ThrowCl = ''Throwable''"
+
+definition NullPointer :: cname
+where
+ "NullPointer \<equiv> ''NullPointer''"
+
+definition ClassCast :: cname
+where
+ "ClassCast \<equiv> ''ClassCast''"
+
+definition OutOfMemory :: cname
+where
+ "OutOfMemory \<equiv> ''OutOfMemory''"
+
+definition NoClassDefFoundError :: cname
+where
+ "NoClassDefFoundError \<equiv> ''NoClassDefFoundError''"
+
+definition IncompatibleClassChangeError :: cname
+where
+ "IncompatibleClassChangeError \<equiv> ''IncompatibleClassChangeError''"
+
+definition NoSuchFieldError :: cname
+where
+ "NoSuchFieldError \<equiv> ''NoSuchFieldError''"
+
+definition NoSuchMethodError :: cname
+where
+ "NoSuchMethodError \<equiv> ''NoSuchMethodError''"
+
+definition sys_xcpts :: "cname set"
+where
+ "sys_xcpts \<equiv> {NullPointer, ClassCast, OutOfMemory, NoClassDefFoundError,
+ IncompatibleClassChangeError,
+ NoSuchFieldError, NoSuchMethodError}"
+
+definition addr_of_sys_xcpt :: "cname \<Rightarrow> addr"
+where
+ "addr_of_sys_xcpt s \<equiv> if s = NullPointer then 0 else
+ if s = ClassCast then 1 else
+ if s = OutOfMemory then 2 else
+ if s = NoClassDefFoundError then 3 else
+ if s = IncompatibleClassChangeError then 4 else
+ if s = NoSuchFieldError then 5 else
+ if s = NoSuchMethodError then 6 else undefined"
+
+
+lemmas sys_xcpts_defs = NullPointer_def ClassCast_def OutOfMemory_def NoClassDefFoundError_def
+ IncompatibleClassChangeError_def NoSuchFieldError_def NoSuchMethodError_def
+
+lemma Start_nsys_xcpts: "Start \<notin> sys_xcpts"
+ by(simp add: Start_def sys_xcpts_def sys_xcpts_defs)
+
+lemma Start_nsys_xcpts1 [simp]: "Start \<noteq> NullPointer" "Start \<noteq> ClassCast"
+ "Start \<noteq> OutOfMemory" "Start \<noteq> NoClassDefFoundError"
+ "Start \<noteq> IncompatibleClassChangeError" "Start \<noteq> NoSuchFieldError"
+ "Start \<noteq> NoSuchMethodError"
+using Start_nsys_xcpts by(auto simp: sys_xcpts_def)
+
+lemma Start_nsys_xcpts2 [simp]: "NullPointer \<noteq> Start" "ClassCast \<noteq> Start"
+ "OutOfMemory \<noteq> Start" "NoClassDefFoundError \<noteq> Start"
+ "IncompatibleClassChangeError \<noteq> Start" "NoSuchFieldError \<noteq> Start"
+ "NoSuchMethodError \<noteq> Start"
+using Start_nsys_xcpts by(auto simp: sys_xcpts_def dest: sym)
+
+definition start_heap :: "'c prog \<Rightarrow> heap"
+where
+ "start_heap G \<equiv> Map.empty (addr_of_sys_xcpt NullPointer \<mapsto> blank G NullPointer)
+ (addr_of_sys_xcpt ClassCast \<mapsto> blank G ClassCast)
+ (addr_of_sys_xcpt OutOfMemory \<mapsto> blank G OutOfMemory)
+ (addr_of_sys_xcpt NoClassDefFoundError \<mapsto> blank G NoClassDefFoundError)
+ (addr_of_sys_xcpt IncompatibleClassChangeError \<mapsto> blank G IncompatibleClassChangeError)
+ (addr_of_sys_xcpt NoSuchFieldError \<mapsto> blank G NoSuchFieldError)
+ (addr_of_sys_xcpt NoSuchMethodError \<mapsto> blank G NoSuchMethodError)"
+
+definition preallocated :: "heap \<Rightarrow> bool"
+where
+ "preallocated h \<equiv> \<forall>C \<in> sys_xcpts. \<exists>fs. h(addr_of_sys_xcpt C) = Some (C,fs)"
+
+subsection "System exceptions"
+
+lemma sys_xcpts_incl [simp]: "NullPointer \<in> sys_xcpts \<and> OutOfMemory \<in> sys_xcpts
+ \<and> ClassCast \<in> sys_xcpts \<and> NoClassDefFoundError \<in> sys_xcpts
+ \<and> IncompatibleClassChangeError \<in> sys_xcpts \<and> NoSuchFieldError \<in> sys_xcpts
+ \<and> NoSuchMethodError \<in> sys_xcpts"
+(*<*)by(simp add: sys_xcpts_def)(*>*)
+
+lemma sys_xcpts_cases [consumes 1, cases set]:
+ "\<lbrakk> C \<in> sys_xcpts; P NullPointer; P OutOfMemory; P ClassCast; P NoClassDefFoundError;
+ P IncompatibleClassChangeError; P NoSuchFieldError;
+ P NoSuchMethodError \<rbrakk> \<Longrightarrow> P C"
+(*<*)by (auto simp: sys_xcpts_def)(*>*)
+
+subsection "Starting heap"
+
+lemma start_heap_sys_xcpts:
+assumes "C \<in> sys_xcpts"
+shows "start_heap P (addr_of_sys_xcpt C) = Some(blank P C)"
+by(rule sys_xcpts_cases[OF assms])
+ (auto simp add: start_heap_def sys_xcpts_def addr_of_sys_xcpt_def sys_xcpts_defs)
+
+lemma start_heap_classes:
+ "start_heap P a = Some(C,fs) \<Longrightarrow> C \<in> sys_xcpts"
+ by(simp add: start_heap_def blank_def split: if_split_asm)
+
+lemma start_heap_nStart: "start_heap P a = Some obj \<Longrightarrow> fst(obj) \<noteq> Start"
+ by(cases obj, auto dest!: start_heap_classes simp: Start_nsys_xcpts)
+
+subsection "@{term preallocated}"
+
+lemma preallocated_dom [simp]:
+ "\<lbrakk> preallocated h; C \<in> sys_xcpts \<rbrakk> \<Longrightarrow> addr_of_sys_xcpt C \<in> dom h"
+(*<*)by (fastforce simp:preallocated_def dom_def)(*>*)
+
+
+lemma preallocatedD:
+ "\<lbrakk> preallocated h; C \<in> sys_xcpts \<rbrakk> \<Longrightarrow> \<exists>fs. h(addr_of_sys_xcpt C) = Some (C, fs)"
+(*<*)by(auto simp: preallocated_def sys_xcpts_def)(*>*)
+
+
+lemma preallocatedE [elim?]:
+ "\<lbrakk> preallocated h; C \<in> sys_xcpts; \<And>fs. h(addr_of_sys_xcpt C) = Some(C,fs) \<Longrightarrow> P h C\<rbrakk>
+ \<Longrightarrow> P h C"
+(*<*)by (fast dest: preallocatedD)(*>*)
+
+
+lemma cname_of_xcp [simp]:
+ "\<lbrakk> preallocated h; C \<in> sys_xcpts \<rbrakk> \<Longrightarrow> cname_of h (addr_of_sys_xcpt C) = C"
+(*<*)by (auto elim: preallocatedE)(*>*)
+
+
+lemma typeof_ClassCast [simp]:
+ "preallocated h \<Longrightarrow> typeof\<^bsub>h\<^esub> (Addr(addr_of_sys_xcpt ClassCast)) = Some(Class ClassCast)"
+(*<*)by (auto elim: preallocatedE)(*>*)
+
+
+lemma typeof_OutOfMemory [simp]:
+ "preallocated h \<Longrightarrow> typeof\<^bsub>h\<^esub> (Addr(addr_of_sys_xcpt OutOfMemory)) = Some(Class OutOfMemory)"
+(*<*)by (auto elim: preallocatedE)(*>*)
+
+
+lemma typeof_NullPointer [simp]:
+ "preallocated h \<Longrightarrow> typeof\<^bsub>h\<^esub> (Addr(addr_of_sys_xcpt NullPointer)) = Some(Class NullPointer)"
+(*<*)by (auto elim: preallocatedE)(*>*)
+
+lemma typeof_NoClassDefFoundError [simp]:
+ "preallocated h \<Longrightarrow> typeof\<^bsub>h\<^esub> (Addr(addr_of_sys_xcpt NoClassDefFoundError)) = Some(Class NoClassDefFoundError)"
+(*<*)by (auto elim: preallocatedE)(*>*)
+
+lemma typeof_IncompatibleClassChangeError [simp]:
+ "preallocated h \<Longrightarrow> typeof\<^bsub>h\<^esub> (Addr(addr_of_sys_xcpt IncompatibleClassChangeError)) = Some(Class IncompatibleClassChangeError)"
+(*<*)by (auto elim: preallocatedE)(*>*)
+
+lemma typeof_NoSuchFieldError [simp]:
+ "preallocated h \<Longrightarrow> typeof\<^bsub>h\<^esub> (Addr(addr_of_sys_xcpt NoSuchFieldError)) = Some(Class NoSuchFieldError)"
+(*<*)by (auto elim: preallocatedE)(*>*)
+
+lemma typeof_NoSuchMethodError [simp]:
+ "preallocated h \<Longrightarrow> typeof\<^bsub>h\<^esub> (Addr(addr_of_sys_xcpt NoSuchMethodError)) = Some(Class NoSuchMethodError)"
+(*<*)by (auto elim: preallocatedE)(*>*)
+
+lemma preallocated_hext:
+ "\<lbrakk> preallocated h; h \<unlhd> h' \<rbrakk> \<Longrightarrow> preallocated h'"
+(*<*)by (simp add: preallocated_def hext_def)(*>*)
+
+(*<*)
+lemmas preallocated_upd_obj = preallocated_hext [OF _ hext_upd_obj]
+lemmas preallocated_new = preallocated_hext [OF _ hext_new]
+(*>*)
+
+lemma preallocated_start:
+ "preallocated (start_heap P)"
+ by(auto simp: start_heap_sys_xcpts blank_def preallocated_def)
+
+end
diff --git a/thys/JinjaDCI/Common/Objects.thy b/thys/JinjaDCI/Common/Objects.thy
--- a/thys/JinjaDCI/Common/Objects.thy
+++ b/thys/JinjaDCI/Common/Objects.thy
@@ -1,211 +1,211 @@
-(* Title: JinjaDCI/Common/Objects.thy
-
- Author: David von Oheimb, Susannah Mansky
- Copyright 1999 Technische Universitaet Muenchen, 2019-20 UIUC
-
- Based on the Jinja theory Common/Objects.thy by David von Oheimb
-*)
-
-section \<open> Objects and the Heap \<close>
-
-theory Objects imports TypeRel Value begin
-
-subsection\<open> Objects \<close>
-
-type_synonym
- fields = "vname \<times> cname \<rightharpoonup> val" \<comment> \<open>field name, defining class, value\<close>
-type_synonym
- obj = "cname \<times> fields" \<comment> \<open>class instance with class name and fields\<close>
-type_synonym
- sfields = "vname \<rightharpoonup> val" \<comment> \<open>field name to value\<close>
-
-definition obj_ty :: "obj \<Rightarrow> ty"
-where
- "obj_ty obj \<equiv> Class (fst obj)"
-
- \<comment> \<open> initializes a given list of fields \<close>
-definition init_fields :: "((vname \<times> cname) \<times> staticb \<times> ty) list \<Rightarrow> fields"
-where
- "init_fields FDTs \<equiv> (map_of \<circ> map (\<lambda>((F,D),b,T). ((F,D),default_val T))) FDTs"
-
-definition init_sfields :: "((vname \<times> cname) \<times> staticb \<times> ty) list \<Rightarrow> sfields"
-where
- "init_sfields FDTs \<equiv> (map_of \<circ> map (\<lambda>((F,D),b,T). (F,default_val T))) FDTs"
-
- \<comment> \<open>a new, blank object with default values for instance fields:\<close>
-definition blank :: "'m prog \<Rightarrow> cname \<Rightarrow> obj"
-where
- "blank P C \<equiv> (C,init_fields (ifields P C))"
-
- \<comment> \<open>a new, blank object with default values for static fields:\<close>
-definition sblank :: "'m prog \<Rightarrow> cname \<Rightarrow> sfields"
-where
- "sblank P C \<equiv> init_sfields (isfields P C)"
-
-lemma [simp]: "obj_ty (C,fs) = Class C"
-(*<*)by (simp add: obj_ty_def)(*>*)
-
-(* replaced all vname, cname in below with `char list' and \<rightharpoonup> with returned option
- so that pretty printing works -SM *)
-translations
- (type) "fields" <= (type) "char list \<times> char list \<Rightarrow> val option"
- (type) "obj" <= (type) "char list \<times> fields"
- (type) "sfields" <= (type) "char list \<Rightarrow> val option"
-
-subsection\<open> Heap \<close>
-
-type_synonym heap = "addr \<rightharpoonup> obj"
-
-(* replaced addr with nat and \<rightharpoonup> with returned option so that pretty printing works -SM *)
-translations
- (type) "heap" <= (type) "nat \<Rightarrow> obj option"
-
-abbreviation
- cname_of :: "heap \<Rightarrow> addr \<Rightarrow> cname" where
- "cname_of hp a == fst (the (hp a))"
-
-definition new_Addr :: "heap \<Rightarrow> addr option"
-where
- "new_Addr h \<equiv> if \<exists>a. h a = None then Some(LEAST a. h a = None) else None"
-
-definition cast_ok :: "'m prog \<Rightarrow> cname \<Rightarrow> heap \<Rightarrow> val \<Rightarrow> bool"
-where
- "cast_ok P C h v \<equiv> v = Null \<or> P \<turnstile> cname_of h (the_Addr v) \<preceq>\<^sup>* C"
-
-definition hext :: "heap \<Rightarrow> heap \<Rightarrow> bool" ("_ \<unlhd> _" [51,51] 50)
-where
- "h \<unlhd> h' \<equiv> \<forall>a C fs. h a = Some(C,fs) \<longrightarrow> (\<exists>fs'. h' a = Some(C,fs'))"
-
-primrec typeof_h :: "heap \<Rightarrow> val \<Rightarrow> ty option" ("typeof\<^bsub>_\<^esub>")
-where
- "typeof\<^bsub>h\<^esub> Unit = Some Void"
-| "typeof\<^bsub>h\<^esub> Null = Some NT"
-| "typeof\<^bsub>h\<^esub> (Bool b) = Some Boolean"
-| "typeof\<^bsub>h\<^esub> (Intg i) = Some Integer"
-| "typeof\<^bsub>h\<^esub> (Addr a) = (case h a of None \<Rightarrow> None | Some(C,fs) \<Rightarrow> Some(Class C))"
-
-lemma new_Addr_SomeD:
- "new_Addr h = Some a \<Longrightarrow> h a = None"
- (*<*)by(fastforce simp: new_Addr_def split:if_splits intro:LeastI)(*>*)
-
-lemma [simp]: "(typeof\<^bsub>h\<^esub> v = Some Boolean) = (\<exists>b. v = Bool b)"
- (*<*)by(induct v) auto(*>*)
-
-lemma [simp]: "(typeof\<^bsub>h\<^esub> v = Some Integer) = (\<exists>i. v = Intg i)"
-(*<*)by(cases v) auto(*>*)
-
-lemma [simp]: "(typeof\<^bsub>h\<^esub> v = Some NT) = (v = Null)"
- (*<*)by(cases v) auto(*>*)
-
-lemma [simp]: "(typeof\<^bsub>h\<^esub> v = Some(Class C)) = (\<exists>a fs. v = Addr a \<and> h a = Some(C,fs))"
- (*<*)by(cases v) auto(*>*)
-
-lemma [simp]: "h a = Some(C,fs) \<Longrightarrow> typeof\<^bsub>(h(a\<mapsto>(C,fs')))\<^esub> v = typeof\<^bsub>h\<^esub> v"
- (*<*)by(induct v) (auto simp:fun_upd_apply)(*>*)
-
-text\<open> For literal values the first parameter of @{term typeof} can be
-set to @{term empty} because they do not contain addresses: \<close>
-
-abbreviation
- typeof :: "val \<Rightarrow> ty option" where
- "typeof v == typeof_h Map.empty v"
-
-lemma typeof_lit_typeof:
- "typeof v = Some T \<Longrightarrow> typeof\<^bsub>h\<^esub> v = Some T"
- (*<*)by(cases v) auto(*>*)
-
-lemma typeof_lit_is_type:
- "typeof v = Some T \<Longrightarrow> is_type P T"
- (*<*)by (induct v) (auto simp:is_type_def)(*>*)
-
-
-subsection \<open> Heap extension @{text"\<unlhd>"} \<close>
-
-lemma hextI: "\<forall>a C fs. h a = Some(C,fs) \<longrightarrow> (\<exists>fs'. h' a = Some(C,fs')) \<Longrightarrow> h \<unlhd> h'"
-(*<*)by(auto simp: hext_def)(*>*)
-
-lemma hext_objD: "\<lbrakk> h \<unlhd> h'; h a = Some(C,fs) \<rbrakk> \<Longrightarrow> \<exists>fs'. h' a = Some(C,fs')"
-(*<*)by(auto simp: hext_def)(*>*)
-
-lemma hext_refl [iff]: "h \<unlhd> h"
-(*<*)by (rule hextI) fast(*>*)
-
-lemma hext_new [simp]: "h a = None \<Longrightarrow> h \<unlhd> h(a\<mapsto>x)"
-(*<*)by (rule hextI) (auto simp:fun_upd_apply)(*>*)
-
-lemma hext_trans: "\<lbrakk> h \<unlhd> h'; h' \<unlhd> h'' \<rbrakk> \<Longrightarrow> h \<unlhd> h''"
-(*<*)by (rule hextI) (fast dest: hext_objD)(*>*)
-
-lemma hext_upd_obj: "h a = Some (C,fs) \<Longrightarrow> h \<unlhd> h(a\<mapsto>(C,fs'))"
-(*<*)by (rule hextI) (auto simp:fun_upd_apply)(*>*)
-
-lemma hext_typeof_mono: "\<lbrakk> h \<unlhd> h'; typeof\<^bsub>h\<^esub> v = Some T \<rbrakk> \<Longrightarrow> typeof\<^bsub>h'\<^esub> v = Some T"
-(*<*)
-proof(cases v)
- case Addr assume "h \<unlhd> h'" and "typeof\<^bsub>h\<^esub> v = \<lfloor>T\<rfloor>"
- then show ?thesis using Addr by(fastforce simp:hext_def)
-qed simp_all
-(*>*)
-
-subsection\<open> Static field information function \<close>
-
-datatype init_state = Done | Processing | Prepared | Error
- \<comment> \<open>@{term Done} = initialized\<close>
- \<comment> \<open>@{term Processing} = currently being initialized\<close>
- \<comment> \<open>@{term Prepared} = uninitialized and not currently being initialized\<close>
- \<comment> \<open>@{term Error} = previous initialization attempt resulted in erroneous state\<close>
-
-inductive iprog :: "init_state \<Rightarrow> init_state \<Rightarrow> bool" ("_ \<le>\<^sub>i _" [51,51] 50)
-where
- [simp]: "Prepared \<le>\<^sub>i i"
-| [simp]: "Processing \<le>\<^sub>i Done"
-| [simp]: "Processing \<le>\<^sub>i Error"
-| [simp]: "i \<le>\<^sub>i i"
-
-lemma iprog_Done[simp]: "(Done \<le>\<^sub>i i) = (i = Done)"
- by(simp only: iprog.simps, simp)
-
-lemma iprog_Error[simp]: "(Error \<le>\<^sub>i i) = (i = Error)"
- by(simp only: iprog.simps, simp)
-
-lemma iprog_Processing[simp]: "(Processing \<le>\<^sub>i i) = (i = Done \<or> i = Error \<or> i = Processing)"
- by(simp only: iprog.simps, simp)
-
-lemma iprog_trans: "\<lbrakk> i \<le>\<^sub>i i'; i' \<le>\<^sub>i i'' \<rbrakk> \<Longrightarrow> i \<le>\<^sub>i i''"
-(*<*)by(case_tac i; case_tac i') simp_all(*>*)
-
-subsection\<open> Static Heap \<close>
-
-text \<open>The static heap (sheap) is used for storing information about static
- field values and initialization status for classes.\<close>
-
-type_synonym
- sheap = "cname \<rightharpoonup> sfields \<times> init_state"
-
-translations
- (type) "sheap" <= (type) "char list \<Rightarrow> (sfields \<times> init_state) option"
-
-definition shext :: "sheap \<Rightarrow> sheap \<Rightarrow> bool" ("_ \<unlhd>\<^sub>s _" [51,51] 50)
-where
- "sh \<unlhd>\<^sub>s sh' \<equiv> \<forall>C sfs i. sh C = Some(sfs,i) \<longrightarrow> (\<exists>sfs' i'. sh' C = Some(sfs',i') \<and> i \<le>\<^sub>i i')"
-
-
-lemma shextI: "\<forall>C sfs i. sh C = Some(sfs,i) \<longrightarrow> (\<exists>sfs' i'. sh' C = Some(sfs',i') \<and> i \<le>\<^sub>i i') \<Longrightarrow> sh \<unlhd>\<^sub>s sh'"
-(*<*)by(auto simp: shext_def)(*>*)
-
-lemma shext_objD: "\<lbrakk> sh \<unlhd>\<^sub>s sh'; sh C = Some(sfs,i) \<rbrakk> \<Longrightarrow> \<exists>sfs' i'. sh' C = Some(sfs', i') \<and> i \<le>\<^sub>i i'"
-(*<*)by(auto simp: shext_def)(*>*)
-
-lemma shext_refl [iff]: "sh \<unlhd>\<^sub>s sh"
-(*<*)by (rule shextI) auto(*>*)
-
-lemma shext_new [simp]: "sh C = None \<Longrightarrow> sh \<unlhd>\<^sub>s sh(C\<mapsto>x)"
-(*<*)by (rule shextI) (auto simp:fun_upd_apply)(*>*)
-
-lemma shext_trans: "\<lbrakk> sh \<unlhd>\<^sub>s sh'; sh' \<unlhd>\<^sub>s sh'' \<rbrakk> \<Longrightarrow> sh \<unlhd>\<^sub>s sh''"
-(*<*)by (rule shextI) (fast dest: iprog_trans shext_objD)(*>*)
-
-lemma shext_upd_obj: "\<lbrakk> sh C = Some (sfs,i); i \<le>\<^sub>i i' \<rbrakk> \<Longrightarrow> sh \<unlhd>\<^sub>s sh(C\<mapsto>(sfs',i'))"
-(*<*)by (rule shextI) (auto simp:fun_upd_apply)(*>*)
-
-end
+(* Title: JinjaDCI/Common/Objects.thy
+
+ Author: David von Oheimb, Susannah Mansky
+ Copyright 1999 Technische Universitaet Muenchen, 2019-20 UIUC
+
+ Based on the Jinja theory Common/Objects.thy by David von Oheimb
+*)
+
+section \<open> Objects and the Heap \<close>
+
+theory Objects imports TypeRel Value begin
+
+subsection\<open> Objects \<close>
+
+type_synonym
+ fields = "vname \<times> cname \<rightharpoonup> val" \<comment> \<open>field name, defining class, value\<close>
+type_synonym
+ obj = "cname \<times> fields" \<comment> \<open>class instance with class name and fields\<close>
+type_synonym
+ sfields = "vname \<rightharpoonup> val" \<comment> \<open>field name to value\<close>
+
+definition obj_ty :: "obj \<Rightarrow> ty"
+where
+ "obj_ty obj \<equiv> Class (fst obj)"
+
+ \<comment> \<open> initializes a given list of fields \<close>
+definition init_fields :: "((vname \<times> cname) \<times> staticb \<times> ty) list \<Rightarrow> fields"
+where
+ "init_fields FDTs \<equiv> (map_of \<circ> map (\<lambda>((F,D),b,T). ((F,D),default_val T))) FDTs"
+
+definition init_sfields :: "((vname \<times> cname) \<times> staticb \<times> ty) list \<Rightarrow> sfields"
+where
+ "init_sfields FDTs \<equiv> (map_of \<circ> map (\<lambda>((F,D),b,T). (F,default_val T))) FDTs"
+
+ \<comment> \<open>a new, blank object with default values for instance fields:\<close>
+definition blank :: "'m prog \<Rightarrow> cname \<Rightarrow> obj"
+where
+ "blank P C \<equiv> (C,init_fields (ifields P C))"
+
+ \<comment> \<open>a new, blank object with default values for static fields:\<close>
+definition sblank :: "'m prog \<Rightarrow> cname \<Rightarrow> sfields"
+where
+ "sblank P C \<equiv> init_sfields (isfields P C)"
+
+lemma [simp]: "obj_ty (C,fs) = Class C"
+(*<*)by (simp add: obj_ty_def)(*>*)
+
+(* replaced all vname, cname in below with `char list' and \<rightharpoonup> with returned option
+ so that pretty printing works -SM *)
+translations
+ (type) "fields" <= (type) "char list \<times> char list \<Rightarrow> val option"
+ (type) "obj" <= (type) "char list \<times> fields"
+ (type) "sfields" <= (type) "char list \<Rightarrow> val option"
+
+subsection\<open> Heap \<close>
+
+type_synonym heap = "addr \<rightharpoonup> obj"
+
+(* replaced addr with nat and \<rightharpoonup> with returned option so that pretty printing works -SM *)
+translations
+ (type) "heap" <= (type) "nat \<Rightarrow> obj option"
+
+abbreviation
+ cname_of :: "heap \<Rightarrow> addr \<Rightarrow> cname" where
+ "cname_of hp a == fst (the (hp a))"
+
+definition new_Addr :: "heap \<Rightarrow> addr option"
+where
+ "new_Addr h \<equiv> if \<exists>a. h a = None then Some(LEAST a. h a = None) else None"
+
+definition cast_ok :: "'m prog \<Rightarrow> cname \<Rightarrow> heap \<Rightarrow> val \<Rightarrow> bool"
+where
+ "cast_ok P C h v \<equiv> v = Null \<or> P \<turnstile> cname_of h (the_Addr v) \<preceq>\<^sup>* C"
+
+definition hext :: "heap \<Rightarrow> heap \<Rightarrow> bool" ("_ \<unlhd> _" [51,51] 50)
+where
+ "h \<unlhd> h' \<equiv> \<forall>a C fs. h a = Some(C,fs) \<longrightarrow> (\<exists>fs'. h' a = Some(C,fs'))"
+
+primrec typeof_h :: "heap \<Rightarrow> val \<Rightarrow> ty option" ("typeof\<^bsub>_\<^esub>")
+where
+ "typeof\<^bsub>h\<^esub> Unit = Some Void"
+| "typeof\<^bsub>h\<^esub> Null = Some NT"
+| "typeof\<^bsub>h\<^esub> (Bool b) = Some Boolean"
+| "typeof\<^bsub>h\<^esub> (Intg i) = Some Integer"
+| "typeof\<^bsub>h\<^esub> (Addr a) = (case h a of None \<Rightarrow> None | Some(C,fs) \<Rightarrow> Some(Class C))"
+
+lemma new_Addr_SomeD:
+ "new_Addr h = Some a \<Longrightarrow> h a = None"
+ (*<*)by(fastforce simp: new_Addr_def split:if_splits intro:LeastI)(*>*)
+
+lemma [simp]: "(typeof\<^bsub>h\<^esub> v = Some Boolean) = (\<exists>b. v = Bool b)"
+ (*<*)by(induct v) auto(*>*)
+
+lemma [simp]: "(typeof\<^bsub>h\<^esub> v = Some Integer) = (\<exists>i. v = Intg i)"
+(*<*)by(cases v) auto(*>*)
+
+lemma [simp]: "(typeof\<^bsub>h\<^esub> v = Some NT) = (v = Null)"
+ (*<*)by(cases v) auto(*>*)
+
+lemma [simp]: "(typeof\<^bsub>h\<^esub> v = Some(Class C)) = (\<exists>a fs. v = Addr a \<and> h a = Some(C,fs))"
+ (*<*)by(cases v) auto(*>*)
+
+lemma [simp]: "h a = Some(C,fs) \<Longrightarrow> typeof\<^bsub>(h(a\<mapsto>(C,fs')))\<^esub> v = typeof\<^bsub>h\<^esub> v"
+ (*<*)by(induct v) (auto simp:fun_upd_apply)(*>*)
+
+text\<open> For literal values the first parameter of @{term typeof} can be
+set to @{term empty} because they do not contain addresses: \<close>
+
+abbreviation
+ typeof :: "val \<Rightarrow> ty option" where
+ "typeof v == typeof_h Map.empty v"
+
+lemma typeof_lit_typeof:
+ "typeof v = Some T \<Longrightarrow> typeof\<^bsub>h\<^esub> v = Some T"
+ (*<*)by(cases v) auto(*>*)
+
+lemma typeof_lit_is_type:
+ "typeof v = Some T \<Longrightarrow> is_type P T"
+ (*<*)by (induct v) (auto simp:is_type_def)(*>*)
+
+
+subsection \<open> Heap extension @{text"\<unlhd>"} \<close>
+
+lemma hextI: "\<forall>a C fs. h a = Some(C,fs) \<longrightarrow> (\<exists>fs'. h' a = Some(C,fs')) \<Longrightarrow> h \<unlhd> h'"
+(*<*)by(auto simp: hext_def)(*>*)
+
+lemma hext_objD: "\<lbrakk> h \<unlhd> h'; h a = Some(C,fs) \<rbrakk> \<Longrightarrow> \<exists>fs'. h' a = Some(C,fs')"
+(*<*)by(auto simp: hext_def)(*>*)
+
+lemma hext_refl [iff]: "h \<unlhd> h"
+(*<*)by (rule hextI) fast(*>*)
+
+lemma hext_new [simp]: "h a = None \<Longrightarrow> h \<unlhd> h(a\<mapsto>x)"
+(*<*)by (rule hextI) (auto simp:fun_upd_apply)(*>*)
+
+lemma hext_trans: "\<lbrakk> h \<unlhd> h'; h' \<unlhd> h'' \<rbrakk> \<Longrightarrow> h \<unlhd> h''"
+(*<*)by (rule hextI) (fast dest: hext_objD)(*>*)
+
+lemma hext_upd_obj: "h a = Some (C,fs) \<Longrightarrow> h \<unlhd> h(a\<mapsto>(C,fs'))"
+(*<*)by (rule hextI) (auto simp:fun_upd_apply)(*>*)
+
+lemma hext_typeof_mono: "\<lbrakk> h \<unlhd> h'; typeof\<^bsub>h\<^esub> v = Some T \<rbrakk> \<Longrightarrow> typeof\<^bsub>h'\<^esub> v = Some T"
+(*<*)
+proof(cases v)
+ case Addr assume "h \<unlhd> h'" and "typeof\<^bsub>h\<^esub> v = \<lfloor>T\<rfloor>"
+ then show ?thesis using Addr by(fastforce simp:hext_def)
+qed simp_all
+(*>*)
+
+subsection\<open> Static field information function \<close>
+
+datatype init_state = Done | Processing | Prepared | Error
+ \<comment> \<open>@{term Done} = initialized\<close>
+ \<comment> \<open>@{term Processing} = currently being initialized\<close>
+ \<comment> \<open>@{term Prepared} = uninitialized and not currently being initialized\<close>
+ \<comment> \<open>@{term Error} = previous initialization attempt resulted in erroneous state\<close>
+
+inductive iprog :: "init_state \<Rightarrow> init_state \<Rightarrow> bool" ("_ \<le>\<^sub>i _" [51,51] 50)
+where
+ [simp]: "Prepared \<le>\<^sub>i i"
+| [simp]: "Processing \<le>\<^sub>i Done"
+| [simp]: "Processing \<le>\<^sub>i Error"
+| [simp]: "i \<le>\<^sub>i i"
+
+lemma iprog_Done[simp]: "(Done \<le>\<^sub>i i) = (i = Done)"
+ by(simp only: iprog.simps, simp)
+
+lemma iprog_Error[simp]: "(Error \<le>\<^sub>i i) = (i = Error)"
+ by(simp only: iprog.simps, simp)
+
+lemma iprog_Processing[simp]: "(Processing \<le>\<^sub>i i) = (i = Done \<or> i = Error \<or> i = Processing)"
+ by(simp only: iprog.simps, simp)
+
+lemma iprog_trans: "\<lbrakk> i \<le>\<^sub>i i'; i' \<le>\<^sub>i i'' \<rbrakk> \<Longrightarrow> i \<le>\<^sub>i i''"
+(*<*)by(case_tac i; case_tac i') simp_all(*>*)
+
+subsection\<open> Static Heap \<close>
+
+text \<open>The static heap (sheap) is used for storing information about static
+ field values and initialization status for classes.\<close>
+
+type_synonym
+ sheap = "cname \<rightharpoonup> sfields \<times> init_state"
+
+translations
+ (type) "sheap" <= (type) "char list \<Rightarrow> (sfields \<times> init_state) option"
+
+definition shext :: "sheap \<Rightarrow> sheap \<Rightarrow> bool" ("_ \<unlhd>\<^sub>s _" [51,51] 50)
+where
+ "sh \<unlhd>\<^sub>s sh' \<equiv> \<forall>C sfs i. sh C = Some(sfs,i) \<longrightarrow> (\<exists>sfs' i'. sh' C = Some(sfs',i') \<and> i \<le>\<^sub>i i')"
+
+
+lemma shextI: "\<forall>C sfs i. sh C = Some(sfs,i) \<longrightarrow> (\<exists>sfs' i'. sh' C = Some(sfs',i') \<and> i \<le>\<^sub>i i') \<Longrightarrow> sh \<unlhd>\<^sub>s sh'"
+(*<*)by(auto simp: shext_def)(*>*)
+
+lemma shext_objD: "\<lbrakk> sh \<unlhd>\<^sub>s sh'; sh C = Some(sfs,i) \<rbrakk> \<Longrightarrow> \<exists>sfs' i'. sh' C = Some(sfs', i') \<and> i \<le>\<^sub>i i'"
+(*<*)by(auto simp: shext_def)(*>*)
+
+lemma shext_refl [iff]: "sh \<unlhd>\<^sub>s sh"
+(*<*)by (rule shextI) auto(*>*)
+
+lemma shext_new [simp]: "sh C = None \<Longrightarrow> sh \<unlhd>\<^sub>s sh(C\<mapsto>x)"
+(*<*)by (rule shextI) (auto simp:fun_upd_apply)(*>*)
+
+lemma shext_trans: "\<lbrakk> sh \<unlhd>\<^sub>s sh'; sh' \<unlhd>\<^sub>s sh'' \<rbrakk> \<Longrightarrow> sh \<unlhd>\<^sub>s sh''"
+(*<*)by (rule shextI) (fast dest: iprog_trans shext_objD)(*>*)
+
+lemma shext_upd_obj: "\<lbrakk> sh C = Some (sfs,i); i \<le>\<^sub>i i' \<rbrakk> \<Longrightarrow> sh \<unlhd>\<^sub>s sh(C\<mapsto>(sfs',i'))"
+(*<*)by (rule shextI) (auto simp:fun_upd_apply)(*>*)
+
+end
diff --git a/thys/JinjaDCI/Common/SystemClasses.thy b/thys/JinjaDCI/Common/SystemClasses.thy
--- a/thys/JinjaDCI/Common/SystemClasses.thy
+++ b/thys/JinjaDCI/Common/SystemClasses.thy
@@ -1,57 +1,57 @@
-(* Title: JinjaDCI/Common/SystemClasses.thy
-
- Author: Gerwin Klein, Susannah Mansky
- Copyright 2002 Technische Universitaet Muenchen, 2019-20 UIUC
-
- Based on the Jinja theory Common/SystemClasses.thy by Gerwin Klein
-*)
-
-section \<open> System Classes \<close>
-
-theory SystemClasses
-imports Decl Exceptions
-begin
-
-text \<open>
- This theory provides definitions for the @{text Object} class,
- and the system exceptions.
-\<close>
-
-definition ObjectC :: "'m cdecl"
-where
- "ObjectC \<equiv> (Object, (undefined,[],[]))"
-
-definition NullPointerC :: "'m cdecl"
-where
- "NullPointerC \<equiv> (NullPointer, (Object,[],[]))"
-
-definition ClassCastC :: "'m cdecl"
-where
- "ClassCastC \<equiv> (ClassCast, (Object,[],[]))"
-
-definition OutOfMemoryC :: "'m cdecl"
-where
- "OutOfMemoryC \<equiv> (OutOfMemory, (Object,[],[]))"
-
-definition NoClassDefFoundC :: "'m cdecl"
-where
- "NoClassDefFoundC \<equiv> (NoClassDefFoundError, (Object,[],[]))"
-
-definition IncompatibleClassChangeC :: "'m cdecl"
-where
- "IncompatibleClassChangeC \<equiv> (IncompatibleClassChangeError, (Object,[],[]))"
-
-definition NoSuchFieldC :: "'m cdecl"
-where
- "NoSuchFieldC \<equiv> (NoSuchFieldError, (Object,[],[]))"
-
-definition NoSuchMethodC :: "'m cdecl"
-where
- "NoSuchMethodC \<equiv> (NoSuchMethodError, (Object,[],[]))"
-
-definition SystemClasses :: "'m cdecl list"
-where
- "SystemClasses \<equiv> [ObjectC, NullPointerC, ClassCastC, OutOfMemoryC, NoClassDefFoundC,
- IncompatibleClassChangeC, NoSuchFieldC, NoSuchMethodC]"
-
-end
+(* Title: JinjaDCI/Common/SystemClasses.thy
+
+ Author: Gerwin Klein, Susannah Mansky
+ Copyright 2002 Technische Universitaet Muenchen, 2019-20 UIUC
+
+ Based on the Jinja theory Common/SystemClasses.thy by Gerwin Klein
+*)
+
+section \<open> System Classes \<close>
+
+theory SystemClasses
+imports Decl Exceptions
+begin
+
+text \<open>
+ This theory provides definitions for the @{text Object} class,
+ and the system exceptions.
+\<close>
+
+definition ObjectC :: "'m cdecl"
+where
+ "ObjectC \<equiv> (Object, (undefined,[],[]))"
+
+definition NullPointerC :: "'m cdecl"
+where
+ "NullPointerC \<equiv> (NullPointer, (Object,[],[]))"
+
+definition ClassCastC :: "'m cdecl"
+where
+ "ClassCastC \<equiv> (ClassCast, (Object,[],[]))"
+
+definition OutOfMemoryC :: "'m cdecl"
+where
+ "OutOfMemoryC \<equiv> (OutOfMemory, (Object,[],[]))"
+
+definition NoClassDefFoundC :: "'m cdecl"
+where
+ "NoClassDefFoundC \<equiv> (NoClassDefFoundError, (Object,[],[]))"
+
+definition IncompatibleClassChangeC :: "'m cdecl"
+where
+ "IncompatibleClassChangeC \<equiv> (IncompatibleClassChangeError, (Object,[],[]))"
+
+definition NoSuchFieldC :: "'m cdecl"
+where
+ "NoSuchFieldC \<equiv> (NoSuchFieldError, (Object,[],[]))"
+
+definition NoSuchMethodC :: "'m cdecl"
+where
+ "NoSuchMethodC \<equiv> (NoSuchMethodError, (Object,[],[]))"
+
+definition SystemClasses :: "'m cdecl list"
+where
+ "SystemClasses \<equiv> [ObjectC, NullPointerC, ClassCastC, OutOfMemoryC, NoClassDefFoundC,
+ IncompatibleClassChangeC, NoSuchFieldC, NoSuchMethodC]"
+
+end
diff --git a/thys/JinjaDCI/Common/Type.thy b/thys/JinjaDCI/Common/Type.thy
--- a/thys/JinjaDCI/Common/Type.thy
+++ b/thys/JinjaDCI/Common/Type.thy
@@ -1,65 +1,65 @@
-(* Title: JinjaDCI/Common/Type.thy
-
- Author: David von Oheimb, Tobias Nipkow, Susannah Mansky
- Copyright 1999 Technische Universitaet Muenchen, 2019-20 UIUC
-
- Based on the Jinja theory Common/Type.thy by David von Oheimb and Tobias Nipkow
-*)
-
-section \<open> Jinja types \<close>
-
-theory Type imports Auxiliary begin
-
-type_synonym cname = string \<comment> \<open>class names\<close>
-type_synonym mname = string \<comment> \<open>method name\<close>
-type_synonym vname = string \<comment> \<open>names for local/field variables\<close>
-
-definition Object :: cname
-where
- "Object \<equiv> ''Object''"
-
-definition this :: vname
-where
- "this \<equiv> ''this''"
-
-definition clinit :: "string" where "clinit = ''<clinit>''"
-definition init :: "string" where "init = ''<init>''"
-
-definition start_m :: "string" where "start_m = ''<start>''"
-definition Start :: "string" where "Start = ''<Start>''"
-
-lemma start_m_neq_clinit [simp]: "start_m \<noteq> clinit" by(simp add: start_m_def clinit_def)
-lemma Object_neq_Start [simp]: "Object \<noteq> Start" by(simp add: Object_def Start_def)
-lemma Start_neq_Object [simp]: "Start \<noteq> Object" by(simp add: Object_def Start_def)
-
-\<comment> \<open>field/method static flag\<close>
-
-datatype staticb = Static | NonStatic
-
-\<comment> \<open>types\<close>
-datatype ty
- = Void \<comment> \<open>type of statements\<close>
- | Boolean
- | Integer
- | NT \<comment> \<open>null type\<close>
- | Class cname \<comment> \<open>class type\<close>
-
-definition is_refT :: "ty \<Rightarrow> bool"
-where
- "is_refT T \<equiv> T = NT \<or> (\<exists>C. T = Class C)"
-
-lemma [iff]: "is_refT NT"
-(*<*)by(simp add:is_refT_def)(*>*)
-
-lemma [iff]: "is_refT(Class C)"
-(*<*)by(simp add:is_refT_def)(*>*)
-
-lemma refTE:
- "\<lbrakk>is_refT T; T = NT \<Longrightarrow> P; \<And>C. T = Class C \<Longrightarrow> P \<rbrakk> \<Longrightarrow> P"
-(*<*)by (auto simp add: is_refT_def)(*>*)
-
-lemma not_refTE:
- "\<lbrakk> \<not>is_refT T; T = Void \<or> T = Boolean \<or> T = Integer \<Longrightarrow> P \<rbrakk> \<Longrightarrow> P"
-(*<*)by (cases T, auto simp add: is_refT_def)(*>*)
-
-end
+(* Title: JinjaDCI/Common/Type.thy
+
+ Author: David von Oheimb, Tobias Nipkow, Susannah Mansky
+ Copyright 1999 Technische Universitaet Muenchen, 2019-20 UIUC
+
+ Based on the Jinja theory Common/Type.thy by David von Oheimb and Tobias Nipkow
+*)
+
+section \<open> Jinja types \<close>
+
+theory Type imports Auxiliary begin
+
+type_synonym cname = string \<comment> \<open>class names\<close>
+type_synonym mname = string \<comment> \<open>method name\<close>
+type_synonym vname = string \<comment> \<open>names for local/field variables\<close>
+
+definition Object :: cname
+where
+ "Object \<equiv> ''Object''"
+
+definition this :: vname
+where
+ "this \<equiv> ''this''"
+
+definition clinit :: "string" where "clinit = ''<clinit>''"
+definition init :: "string" where "init = ''<init>''"
+
+definition start_m :: "string" where "start_m = ''<start>''"
+definition Start :: "string" where "Start = ''<Start>''"
+
+lemma start_m_neq_clinit [simp]: "start_m \<noteq> clinit" by(simp add: start_m_def clinit_def)
+lemma Object_neq_Start [simp]: "Object \<noteq> Start" by(simp add: Object_def Start_def)
+lemma Start_neq_Object [simp]: "Start \<noteq> Object" by(simp add: Object_def Start_def)
+
+\<comment> \<open>field/method static flag\<close>
+
+datatype staticb = Static | NonStatic
+
+\<comment> \<open>types\<close>
+datatype ty
+ = Void \<comment> \<open>type of statements\<close>
+ | Boolean
+ | Integer
+ | NT \<comment> \<open>null type\<close>
+ | Class cname \<comment> \<open>class type\<close>
+
+definition is_refT :: "ty \<Rightarrow> bool"
+where
+ "is_refT T \<equiv> T = NT \<or> (\<exists>C. T = Class C)"
+
+lemma [iff]: "is_refT NT"
+(*<*)by(simp add:is_refT_def)(*>*)
+
+lemma [iff]: "is_refT(Class C)"
+(*<*)by(simp add:is_refT_def)(*>*)
+
+lemma refTE:
+ "\<lbrakk>is_refT T; T = NT \<Longrightarrow> P; \<And>C. T = Class C \<Longrightarrow> P \<rbrakk> \<Longrightarrow> P"
+(*<*)by (auto simp add: is_refT_def)(*>*)
+
+lemma not_refTE:
+ "\<lbrakk> \<not>is_refT T; T = Void \<or> T = Boolean \<or> T = Integer \<Longrightarrow> P \<rbrakk> \<Longrightarrow> P"
+(*<*)by (cases T, auto simp add: is_refT_def)(*>*)
+
+end
diff --git a/thys/JinjaDCI/Common/TypeRel.thy b/thys/JinjaDCI/Common/TypeRel.thy
--- a/thys/JinjaDCI/Common/TypeRel.thy
+++ b/thys/JinjaDCI/Common/TypeRel.thy
@@ -1,769 +1,769 @@
-(* Title: JinjaDCI/Common/TypeRel.thy
- Author: Tobias Nipkow, Susannah Mansky
- Copyright 2003 Technische Universitaet Muenchen, 2019-20 UIUC
-
- Based on the Jinja theory Common/TypeRel.thy by Tobias Nipkow
-*)
-
-section \<open> Relations between Jinja Types \<close>
-
-theory TypeRel imports
- "HOL-Library.Transitive_Closure_Table"
- Decl
-begin
-
-subsection \<open> The subclass relations \<close>
-
-inductive_set
- subcls1 :: "'m prog \<Rightarrow> (cname \<times> cname) set"
- and subcls1' :: "'m prog \<Rightarrow> [cname, cname] \<Rightarrow> bool" ("_ \<turnstile> _ \<prec>\<^sup>1 _" [71,71,71] 70)
- for P :: "'m prog"
-where
- "P \<turnstile> C \<prec>\<^sup>1 D \<equiv> (C,D) \<in> subcls1 P"
-| subcls1I: "\<lbrakk>class P C = Some (D,rest); C \<noteq> Object\<rbrakk> \<Longrightarrow> P \<turnstile> C \<prec>\<^sup>1 D"
-
-abbreviation
- subcls :: "'m prog \<Rightarrow> [cname, cname] \<Rightarrow> bool" ("_ \<turnstile> _ \<preceq>\<^sup>* _" [71,71,71] 70)
- where "P \<turnstile> C \<preceq>\<^sup>* D \<equiv> (C,D) \<in> (subcls1 P)\<^sup>*"
-
-lemma subcls1D: "P \<turnstile> C \<prec>\<^sup>1 D \<Longrightarrow> C \<noteq> Object \<and> (\<exists>fs ms. class P C = Some (D,fs,ms))"
-(*<*)by(erule subcls1.induct)(fastforce simp add:is_class_def)(*>*)
-
-lemma [iff]: "\<not> P \<turnstile> Object \<prec>\<^sup>1 C"
-(*<*)by(fastforce dest:subcls1D)(*>*)
-
-lemma [iff]: "(P \<turnstile> Object \<preceq>\<^sup>* C) = (C = Object)"
-(*<*)
-proof(rule iffI)
- assume "P \<turnstile> Object \<preceq>\<^sup>* C" then show "C = Object"
- by(auto elim: converse_rtranclE)
-qed simp
-(*>*)
-
-lemma subcls1_def2:
- "subcls1 P =
- (SIGMA C:{C. is_class P C}. {D. C\<noteq>Object \<and> fst (the (class P C))=D})"
-(*<*)
- by (fastforce simp:is_class_def dest: subcls1D elim: subcls1I)
-(*>*)
-
-lemma finite_subcls1: "finite (subcls1 P)"
-(*<*)
-proof -
- let ?SIG = "SIGMA C:{C. is_class P C}. {D. fst (the (class P C)) = D \<and> C \<noteq> Object}"
- have "subcls1 P = ?SIG" by(simp add: subcls1_def2)
- also have "finite ?SIG"
- proof(rule finite_SigmaI [OF finite_is_class])
- fix C assume C_in: "C \<in> {C. is_class P C}"
- then show "finite {D. fst (the (class P C)) = D \<and> C \<noteq> Object}"
- by(rule_tac finite_subset[where B = "{fst (the (class P C))}"]) auto
- qed
- ultimately show ?thesis by simp
-qed
-(*>*)
-
-primrec supercls_lst :: "'m prog \<Rightarrow> cname list \<Rightarrow> bool" where
-"supercls_lst P (C#Cs) = ((\<forall>C' \<in> set Cs. P \<turnstile> C' \<preceq>\<^sup>* C) \<and> supercls_lst P Cs)" |
-"supercls_lst P [] = True"
-
-lemma supercls_lst_app:
- "\<lbrakk> supercls_lst P (C#Cs); P \<turnstile> C \<preceq>\<^sup>* C' \<rbrakk> \<Longrightarrow> supercls_lst P (C'#C#Cs)"
- by auto
-
-subsection\<open> The subtype relations \<close>
-
-inductive
- widen :: "'m prog \<Rightarrow> ty \<Rightarrow> ty \<Rightarrow> bool" ("_ \<turnstile> _ \<le> _" [71,71,71] 70)
- for P :: "'m prog"
-where
- widen_refl[iff]: "P \<turnstile> T \<le> T"
-| widen_subcls: "P \<turnstile> C \<preceq>\<^sup>* D \<Longrightarrow> P \<turnstile> Class C \<le> Class D"
-| widen_null[iff]: "P \<turnstile> NT \<le> Class C"
-
-abbreviation
- widens :: "'m prog \<Rightarrow> ty list \<Rightarrow> ty list \<Rightarrow> bool"
- ("_ \<turnstile> _ [\<le>] _" [71,71,71] 70) where
- "widens P Ts Ts' \<equiv> list_all2 (widen P) Ts Ts'"
-
-lemma [iff]: "(P \<turnstile> T \<le> Void) = (T = Void)"
-(*<*)by (auto elim: widen.cases)(*>*)
-
-lemma [iff]: "(P \<turnstile> T \<le> Boolean) = (T = Boolean)"
-(*<*)by (auto elim: widen.cases)(*>*)
-
-lemma [iff]: "(P \<turnstile> T \<le> Integer) = (T = Integer)"
-(*<*)by (auto elim: widen.cases)(*>*)
-
-lemma [iff]: "(P \<turnstile> Void \<le> T) = (T = Void)"
-(*<*)by (auto elim: widen.cases)(*>*)
-
-lemma [iff]: "(P \<turnstile> Boolean \<le> T) = (T = Boolean)"
-(*<*)by (auto elim: widen.cases)(*>*)
-
-lemma [iff]: "(P \<turnstile> Integer \<le> T) = (T = Integer)"
-(*<*)by (auto elim: widen.cases)(*>*)
-
-
-lemma Class_widen: "P \<turnstile> Class C \<le> T \<Longrightarrow> \<exists>D. T = Class D"
-(*<*)
-by (ind_cases "P \<turnstile> Class C \<le> T") auto
-(*>*)
-
-lemma [iff]: "(P \<turnstile> T \<le> NT) = (T = NT)"
-(*<*)
-by(cases T) (auto dest:Class_widen)
-(*>*)
-
-lemma Class_widen_Class [iff]: "(P \<turnstile> Class C \<le> Class D) = (P \<turnstile> C \<preceq>\<^sup>* D)"
-(*<*)
-proof(rule iffI)
- show "P \<turnstile> Class C \<le> Class D \<Longrightarrow> P \<turnstile> C \<preceq>\<^sup>* D"
- proof(ind_cases "P \<turnstile> Class C \<le> Class D") qed(auto)
-qed(auto elim: widen_subcls)
-(*>*)
-
-lemma widen_Class: "(P \<turnstile> T \<le> Class C) = (T = NT \<or> (\<exists>D. T = Class D \<and> P \<turnstile> D \<preceq>\<^sup>* C))"
-(*<*)by(induct T, auto)(*>*)
-
-
-lemma widen_trans[trans]: "\<lbrakk>P \<turnstile> S \<le> U; P \<turnstile> U \<le> T\<rbrakk> \<Longrightarrow> P \<turnstile> S \<le> T"
-(*<*)
-proof -
- assume "P\<turnstile>S \<le> U" thus "\<And>T. P \<turnstile> U \<le> T \<Longrightarrow> P \<turnstile> S \<le> T"
- proof induct
- case (widen_refl T T') thus "P \<turnstile> T \<le> T'" .
- next
- case (widen_subcls C D T)
- then obtain E where "T = Class E" by (blast dest: Class_widen)
- with widen_subcls show "P \<turnstile> Class C \<le> T" by (auto elim: rtrancl_trans)
- next
- case (widen_null C RT)
- then obtain D where "RT = Class D" by (blast dest: Class_widen)
- thus "P \<turnstile> NT \<le> RT" by auto
- qed
-qed
-(*>*)
-
-lemma widens_trans [trans]: "\<lbrakk>P \<turnstile> Ss [\<le>] Ts; P \<turnstile> Ts [\<le>] Us\<rbrakk> \<Longrightarrow> P \<turnstile> Ss [\<le>] Us"
-(*<*)by (rule list_all2_trans, rule widen_trans)(*>*)
-
-
-(*<*)
-lemmas widens_refl [iff] = list_all2_refl [of "widen P", OF widen_refl] for P
-lemmas widens_Cons [iff] = list_all2_Cons1 [of "widen P"] for P
-(*>*)
-
-
-subsection\<open> Method lookup \<close>
-
-inductive
- Methods :: "['m prog, cname, mname \<rightharpoonup> (staticb \<times> ty list \<times> ty \<times> 'm) \<times> cname] \<Rightarrow> bool"
- ("_ \<turnstile> _ sees'_methods _" [51,51,51] 50)
- for P :: "'m prog"
-where
- sees_methods_Object:
- "\<lbrakk> class P Object = Some(D,fs,ms); Mm = map_option (\<lambda>m. (m,Object)) \<circ> map_of ms \<rbrakk>
- \<Longrightarrow> P \<turnstile> Object sees_methods Mm"
-| sees_methods_rec:
- "\<lbrakk> class P C = Some(D,fs,ms); C \<noteq> Object; P \<turnstile> D sees_methods Mm;
- Mm' = Mm ++ (map_option (\<lambda>m. (m,C)) \<circ> map_of ms) \<rbrakk>
- \<Longrightarrow> P \<turnstile> C sees_methods Mm'"
-
-lemma sees_methods_fun:
-assumes 1: "P \<turnstile> C sees_methods Mm"
-shows "\<And>Mm'. P \<turnstile> C sees_methods Mm' \<Longrightarrow> Mm' = Mm"
- (*<*)
-using 1
-proof induct
- case (sees_methods_rec C D fs ms Dres Cres Cres')
- have "class": "class P C = Some (D, fs, ms)"
- and notObj: "C \<noteq> Object" and Dmethods: "P \<turnstile> D sees_methods Dres"
- and IH: "\<And>Dres'. P \<turnstile> D sees_methods Dres' \<Longrightarrow> Dres' = Dres"
- and Cres: "Cres = Dres ++ (map_option (\<lambda>m. (m,C)) \<circ> map_of ms)"
- and Cmethods': "P \<turnstile> C sees_methods Cres'" by fact+
- from Cmethods' notObj "class" obtain Dres'
- where Dmethods': "P \<turnstile> D sees_methods Dres'"
- and Cres': "Cres' = Dres' ++ (map_option (\<lambda>m. (m,C)) \<circ> map_of ms)"
- by(auto elim: Methods.cases)
- from Cres Cres' IH[OF Dmethods'] show "Cres' = Cres" by simp
-next
- case sees_methods_Object thus ?case by(auto elim: Methods.cases)
-qed
-(*>*)
-
-lemma visible_methods_exist:
- "P \<turnstile> C sees_methods Mm \<Longrightarrow> Mm M = Some(m,D) \<Longrightarrow>
- (\<exists>D' fs ms. class P D = Some(D',fs,ms) \<and> map_of ms M = Some m)"
- (*<*)by(induct rule:Methods.induct) auto(*>*)
-
-lemma sees_methods_decl_above:
-assumes Csees: "P \<turnstile> C sees_methods Mm"
-shows "Mm M = Some(m,D) \<Longrightarrow> P \<turnstile> C \<preceq>\<^sup>* D"
- (*<*)
-using Csees
-proof induct
-next
- case sees_methods_Object thus ?case by auto
-next
- case sees_methods_rec thus ?case
- by(fastforce simp:map_option_case split:option.splits
- elim:converse_rtrancl_into_rtrancl[OF subcls1I])
-qed
-(*>*)
-
-lemma sees_methods_idemp:
-assumes Cmethods: "P \<turnstile> C sees_methods Mm"
-shows "\<And>m D. Mm M = Some(m,D) \<Longrightarrow>
- \<exists>Mm'. (P \<turnstile> D sees_methods Mm') \<and> Mm' M = Some(m,D)"
-(*<*)
-using Cmethods
-proof induct
- case sees_methods_Object thus ?case
- by(fastforce dest: Methods.sees_methods_Object)
-next
- case sees_methods_rec thus ?case
- by(fastforce split:option.splits dest: Methods.sees_methods_rec)
-qed
-(*>*)
-
-(*FIXME something wrong with induct: need to attache [consumes 1]
-directly to proof of thm, declare does not work. *)
-
-lemma sees_methods_decl_mono:
-assumes sub: "P \<turnstile> C' \<preceq>\<^sup>* C"
-shows "P \<turnstile> C sees_methods Mm \<Longrightarrow>
- \<exists>Mm' Mm\<^sub>2. P \<turnstile> C' sees_methods Mm' \<and> Mm' = Mm ++ Mm\<^sub>2 \<and>
- (\<forall>M m D. Mm\<^sub>2 M = Some(m,D) \<longrightarrow> P \<turnstile> D \<preceq>\<^sup>* C)"
-(*<*)
- (is "_ \<Longrightarrow> \<exists>Mm' Mm2. ?Q C' C Mm' Mm2")
-using sub
-proof (induct rule:converse_rtrancl_induct)
- assume "P \<turnstile> C sees_methods Mm"
- hence "?Q C C Mm Map.empty" by simp
- thus "\<exists>Mm' Mm2. ?Q C C Mm' Mm2" by blast
-next
- fix C'' C'
- assume sub1: "P \<turnstile> C'' \<prec>\<^sup>1 C'" and sub: "P \<turnstile> C' \<preceq>\<^sup>* C"
- and IH: "P \<turnstile> C sees_methods Mm \<Longrightarrow>
- \<exists>Mm' Mm2. P \<turnstile> C' sees_methods Mm' \<and>
- Mm' = Mm ++ Mm2 \<and> (\<forall>M m D. Mm2 M = Some(m,D) \<longrightarrow> P \<turnstile> D \<preceq>\<^sup>* C)"
- and Csees: "P \<turnstile> C sees_methods Mm"
- from IH[OF Csees] obtain Mm' Mm2 where C'sees: "P \<turnstile> C' sees_methods Mm'"
- and Mm': "Mm' = Mm ++ Mm2"
- and subC:"\<forall>M m D. Mm2 M = Some(m,D) \<longrightarrow> P \<turnstile> D \<preceq>\<^sup>* C" by blast
- obtain fs ms where "class": "class P C'' = Some(C',fs,ms)" "C'' \<noteq> Object"
- using subcls1D[OF sub1] by blast
- let ?Mm3 = "map_option (\<lambda>m. (m,C'')) \<circ> map_of ms"
- have "P \<turnstile> C'' sees_methods (Mm ++ Mm2) ++ ?Mm3"
- using sees_methods_rec[OF "class" C'sees refl] Mm' by simp
- hence "?Q C'' C ((Mm ++ Mm2) ++ ?Mm3) (Mm2++?Mm3)"
- using converse_rtrancl_into_rtrancl[OF sub1 sub]
- by simp (simp add:map_add_def subC split:option.split)
- thus "\<exists>Mm' Mm2. ?Q C'' C Mm' Mm2" by blast
-qed
-(*>*)
-
-lemma sees_methods_is_class_Object:
- "P \<turnstile> D sees_methods Mm \<Longrightarrow> is_class P Object"
- by(induct rule: Methods.induct; simp add: is_class_def)
-
-lemma sees_methods_sub_Obj: "P \<turnstile> C sees_methods Mm \<Longrightarrow> P \<turnstile> C \<preceq>\<^sup>* Object"
-proof(induct rule: Methods.induct)
- case (sees_methods_rec C D fs ms Mm Mm') show ?case
- using subcls1I[OF sees_methods_rec.hyps(1,2)] sees_methods_rec.hyps(4)
- by(rule converse_rtrancl_into_rtrancl)
-qed(simp)
-
-
-definition Method :: "'m prog \<Rightarrow> cname \<Rightarrow> mname \<Rightarrow> staticb \<Rightarrow> ty list \<Rightarrow> ty \<Rightarrow> 'm \<Rightarrow> cname \<Rightarrow> bool"
- ("_ \<turnstile> _ sees _, _ : _\<rightarrow>_ = _ in _" [51,51,51,51,51,51,51,51] 50)
-where
- "P \<turnstile> C sees M, b: Ts\<rightarrow>T = m in D \<equiv>
- \<exists>Mm. P \<turnstile> C sees_methods Mm \<and> Mm M = Some((b,Ts,T,m),D)"
-
-definition has_method :: "'m prog \<Rightarrow> cname \<Rightarrow> mname \<Rightarrow> staticb \<Rightarrow> bool"
- ("_ \<turnstile> _ has _, _" [51,0,0,51] 50)
-where
- "P \<turnstile> C has M, b \<equiv> \<exists>Ts T m D. P \<turnstile> C sees M,b:Ts\<rightarrow>T = m in D"
-
-lemma sees_method_fun:
- "\<lbrakk>P \<turnstile> C sees M,b:TS\<rightarrow>T = m in D; P \<turnstile> C sees M,b':TS'\<rightarrow>T' = m' in D' \<rbrakk>
- \<Longrightarrow> b = b' \<and> TS' = TS \<and> T' = T \<and> m' = m \<and> D' = D"
- (*<*)by(fastforce dest: sees_methods_fun simp:Method_def)(*>*)
-
-lemma sees_method_decl_above:
- "P \<turnstile> C sees M,b:Ts\<rightarrow>T = m in D \<Longrightarrow> P \<turnstile> C \<preceq>\<^sup>* D"
- (*<*)by(clarsimp simp:Method_def sees_methods_decl_above)(*>*)
-
-lemma visible_method_exists:
- "P \<turnstile> C sees M,b:Ts\<rightarrow>T = m in D \<Longrightarrow>
- \<exists>D' fs ms. class P D = Some(D',fs,ms) \<and> map_of ms M = Some(b,Ts,T,m)"
-(*<*)by(fastforce simp:Method_def dest!: visible_methods_exist)(*>*)
-
-
-lemma sees_method_idemp:
- "P \<turnstile> C sees M,b:Ts\<rightarrow>T=m in D \<Longrightarrow> P \<turnstile> D sees M,b:Ts\<rightarrow>T=m in D"
- (*<*)by(fastforce simp: Method_def intro:sees_methods_idemp)(*>*)
-
-lemma sees_method_decl_mono:
-assumes sub: "P \<turnstile> C' \<preceq>\<^sup>* C" and
- C_sees: "P \<turnstile> C sees M,b:Ts\<rightarrow>T=m in D" and
- C'_sees: "P \<turnstile> C' sees M,b':Ts'\<rightarrow>T'=m' in D'"
-shows "P \<turnstile> D' \<preceq>\<^sup>* D"
- (*<*)
-proof -
- obtain Ms where Ms: "P \<turnstile> C sees_methods Ms"
- using C_sees by(auto simp: Method_def)
- obtain Ms' Ms2 where Ms': "P \<turnstile> C' sees_methods Ms'" and
- Ms'_def: "Ms' = Ms ++ Ms2" and
- Ms2_imp: "(\<forall>M m D. Ms2 M = \<lfloor>(m, D)\<rfloor> \<longrightarrow> P \<turnstile> D \<preceq>\<^sup>* C)"
- using sees_methods_decl_mono[OF sub Ms] by clarsimp
- have "(Ms ++ Ms2) M = \<lfloor>((b', Ts', T', m'), D')\<rfloor>"
- using C'_sees sees_methods_fun[OF Ms'] Ms'_def by(clarsimp simp: Method_def)
- then have "Ms2 M = \<lfloor>((b', Ts', T', m'), D')\<rfloor> \<or>
- Ms2 M = None \<and> b = b' \<and> Ts = Ts' \<and> T = T' \<and> m = m' \<and> D = D'"
- using C_sees sees_methods_fun[OF Ms] by(clarsimp simp: Method_def)
- also have "Ms2 M = \<lfloor>((b', Ts', T', m'), D')\<rfloor> \<Longrightarrow> P \<turnstile> D' \<preceq>\<^sup>* C"
- using Ms2_imp by simp
- ultimately show ?thesis using sub sees_method_decl_above[OF C_sees] by auto
-qed
-(*>*)
-
-lemma sees_methods_is_class: "P \<turnstile> C sees_methods Mm \<Longrightarrow> is_class P C"
-(*<*)by (auto simp add: is_class_def elim: Methods.induct)(*>*)
-
-lemma sees_method_is_class:
- "\<lbrakk> P \<turnstile> C sees M,b:Ts\<rightarrow>T=m in D \<rbrakk> \<Longrightarrow> is_class P C"
-(*<*)by (auto simp add: is_class_def Method_def dest: sees_methods_is_class)(*>*)
-
-lemma sees_method_is_class':
- "\<lbrakk> P \<turnstile> C sees M,b:Ts\<rightarrow>T=m in D \<rbrakk> \<Longrightarrow> is_class P D"
-(*<*)by(drule sees_method_idemp, rule sees_method_is_class, assumption)(*>*)
-
-lemma sees_method_sub_Obj: "P \<turnstile> C sees M,b: Ts\<rightarrow>T = m in D \<Longrightarrow> P \<turnstile> C \<preceq>\<^sup>* Object"
- by(auto simp: Method_def sees_methods_sub_Obj)
-
-subsection\<open> Field lookup \<close>
-
-inductive
- Fields :: "['m prog, cname, ((vname \<times> cname) \<times> staticb \<times> ty) list] \<Rightarrow> bool"
- ("_ \<turnstile> _ has'_fields _" [51,51,51] 50)
- for P :: "'m prog"
-where
- has_fields_rec:
- "\<lbrakk> class P C = Some(D,fs,ms); C \<noteq> Object; P \<turnstile> D has_fields FDTs;
- FDTs' = map (\<lambda>(F,b,T). ((F,C),b,T)) fs @ FDTs \<rbrakk>
- \<Longrightarrow> P \<turnstile> C has_fields FDTs'"
-| has_fields_Object:
- "\<lbrakk> class P Object = Some(D,fs,ms); FDTs = map (\<lambda>(F,b,T). ((F,Object),b,T)) fs \<rbrakk>
- \<Longrightarrow> P \<turnstile> Object has_fields FDTs"
-
-lemma has_fields_is_class:
- "P \<turnstile> C has_fields FDTs \<Longrightarrow> is_class P C"
-(*<*)by (auto simp add: is_class_def elim: Fields.induct)(*>*)
-
-lemma has_fields_fun:
-assumes 1: "P \<turnstile> C has_fields FDTs"
-shows "\<And>FDTs'. P \<turnstile> C has_fields FDTs' \<Longrightarrow> FDTs' = FDTs"
- (*<*)
-using 1
-proof induct
- case (has_fields_rec C D fs ms Dres Cres Cres')
- have "class": "class P C = Some (D, fs, ms)"
- and notObj: "C \<noteq> Object" and DFields: "P \<turnstile> D has_fields Dres"
- and IH: "\<And>Dres'. P \<turnstile> D has_fields Dres' \<Longrightarrow> Dres' = Dres"
- and Cres: "Cres = map (\<lambda>(F,b,T). ((F,C),b,T)) fs @ Dres"
- and CFields': "P \<turnstile> C has_fields Cres'" by fact+
- from CFields' notObj "class" obtain Dres'
- where DFields': "P \<turnstile> D has_fields Dres'"
- and Cres': "Cres' = map (\<lambda>(F,b,T). ((F,C),b,T)) fs @ Dres'"
- by(auto elim: Fields.cases)
- from Cres Cres' IH[OF DFields'] show "Cres' = Cres" by simp
-next
- case has_fields_Object thus ?case by(auto elim: Fields.cases)
-qed
-(*>*)
-
-lemma all_fields_in_has_fields:
-assumes sub: "P \<turnstile> C has_fields FDTs"
-shows "\<lbrakk> P \<turnstile> C \<preceq>\<^sup>* D; class P D = Some(D',fs,ms); (F,b,T) \<in> set fs \<rbrakk>
- \<Longrightarrow> ((F,D),b,T) \<in> set FDTs"
-(*<*)
-using sub proof(induct)
- case (has_fields_rec C D' fs ms FDTs FDTs')
- then have C_D: "P \<turnstile> C \<preceq>\<^sup>* D" by simp
- then show ?case proof(rule converse_rtranclE)
- assume "C = D"
- then show ?case using has_fields_rec by force
- next
- fix y assume sub1: "P \<turnstile> C \<prec>\<^sup>1 y" and sub2: "P \<turnstile> y \<preceq>\<^sup>* D"
- then show ?case using has_fields_rec subcls1D[OF sub1] by simp
- qed
-next
- case (has_fields_Object D fs ms FDTs)
- then show ?case by force
-qed
-(*>*)
-
-lemma has_fields_decl_above:
-assumes fields: "P \<turnstile> C has_fields FDTs"
-shows "((F,D),b,T) \<in> set FDTs \<Longrightarrow> P \<turnstile> C \<preceq>\<^sup>* D"
-(*<*)
-using fields proof(induct)
- case (has_fields_rec C D' fs ms FDTs FDTs')
- then have "((F, D), b, T) \<in> (\<lambda>x. case x of (F, x) \<Rightarrow> ((F, C), x)) ` set fs \<or>
- ((F, D), b, T) \<in> set FDTs" by clarsimp
- then show ?case proof(rule disjE)
- assume "((F, D), b, T) \<in> (\<lambda>x. case x of (F, x) \<Rightarrow> ((F, C), x)) ` set fs"
- then show ?case using has_fields_rec by clarsimp
- next
- assume "((F, D), b, T) \<in> set FDTs"
- then show ?case using has_fields_rec
- by(blast dest:subcls1I converse_rtrancl_into_rtrancl)
- qed
-next
- case (has_fields_Object D fs ms FDTs)
- then show ?case by fastforce
-qed
-(*>*)
-
-
-lemma subcls_notin_has_fields:
-assumes fields: "P \<turnstile> C has_fields FDTs"
-shows "((F,D),b,T) \<in> set FDTs \<Longrightarrow> (D,C) \<notin> (subcls1 P)\<^sup>+"
-(*<*)
-using fields proof(induct)
- case (has_fields_rec C D' fs ms FDTs FDTs')
- then have "((F, D), b, T) \<in> (\<lambda>x. case x of (F, x) \<Rightarrow> ((F, C), x)) ` set fs
- \<or> ((F, D), b, T) \<in> set FDTs" by clarsimp
- then show ?case proof(rule disjE)
- assume "((F, D), b, T) \<in> (\<lambda>x. case x of (F, x) \<Rightarrow> ((F, C), x)) ` set fs"
- then have CD[simp]: "C = D" and fs: "(F, b, T) \<in> set fs" by clarsimp+
- then have "(D, D) \<in> (subcls1 P)\<^sup>+ \<Longrightarrow> False" proof -
- assume DD: "(D, D) \<in> (subcls1 P)\<^sup>+"
- obtain z where z1: "P \<turnstile> D \<prec>\<^sup>1 z" and z_s: "P \<turnstile> z \<preceq>\<^sup>* D"
- using tranclD[OF DD] by clarsimp
- have [simp]: "z = D'" using subcls1D[OF z1] has_fields_rec.hyps(1) by clarsimp
- then have "((F, D), b, T) \<in> set FDTs"
- using z_s all_fields_in_has_fields[OF has_fields_rec.hyps(3) _ has_fields_rec.hyps(1) fs]
- by simp
- then have "(D, z) \<notin> (subcls1 P)\<^sup>+" using has_fields_rec.hyps(4) by simp
- then show False using z1 by auto
- qed
- then show ?case by clarsimp
- next
- assume "((F, D), b, T) \<in> set FDTs"
- then show ?case using has_fields_rec by(blast dest:subcls1I trancl_into_trancl)
- qed
-next
- case (has_fields_Object D fs ms FDTs)
- then show ?case by(fastforce dest: tranclD)
-qed
-(*>*)
-
-lemma subcls_notin_has_fields2:
-assumes fields: "P \<turnstile> C has_fields FDTs"
-shows "\<lbrakk> C \<noteq> Object; P \<turnstile> C \<prec>\<^sup>1 D \<rbrakk> \<Longrightarrow> (D,C) \<notin> (subcls1 P)\<^sup>*"
-using fields proof(induct arbitrary: D)
- case has_fields_rec
- have "\<forall>C C' P. (C, C') \<notin> subcls1 P \<or> C \<noteq> Object \<and> (\<exists>fs ms. class P C = \<lfloor>(C', fs, ms)\<rfloor>)"
- using subcls1D by blast
- then have "(D, D) \<notin> (subcls1 P)\<^sup>+"
- by (metis (no_types) Pair_inject has_fields_rec.hyps(1) has_fields_rec.hyps(4)
- has_fields_rec.prems(2) option.inject tranclD)
- then show ?case
- by (meson has_fields_rec.prems(2) rtrancl_into_trancl1)
-qed(fastforce dest: tranclD)
-
-lemma has_fields_mono_lem:
-assumes sub: "P \<turnstile> D \<preceq>\<^sup>* C"
-shows "P \<turnstile> C has_fields FDTs
- \<Longrightarrow> \<exists>pre. P \<turnstile> D has_fields pre@FDTs \<and> dom(map_of pre) \<inter> dom(map_of FDTs) = {}"
-(*<*)
-using sub proof(induct rule:converse_rtrancl_induct)
- case base
- then show ?case by(rule_tac x = "[]" in exI) simp
-next
- case (step D' D)
- then obtain pre where D_flds: "P \<turnstile> D has_fields pre @ FDTs" and
- dom: "dom (map_of pre) \<inter> dom (map_of FDTs) = {}" by clarsimp
- have "(D',C) \<in> (subcls1 P)^+" by (rule rtrancl_into_trancl2[OF step.hyps(1,2)])
- obtain fs ms where D'_cls: "class P D' = \<lfloor>(D, fs, ms)\<rfloor>" "D' \<noteq> Object"
- using subcls1D[OF step.hyps(1)] by clarsimp+
- have "P \<turnstile> D' has_fields map (\<lambda>(F, T). ((F, D'), T)) fs @ pre @ FDTs"
- using has_fields_rec[OF D'_cls D_flds] by simp
- also have "dom (map_of (map (\<lambda>(F, T). ((F, D'), T)) fs @ pre))
- \<inter> dom (map_of FDTs) = {}"
- using dom subcls_notin_has_fields[OF D_flds, where D=D'] step.hyps(1)
- by(auto simp:dom_map_of_conv_image_fst) fast
- ultimately show ?case
- by(rule_tac x = "map (\<lambda>(F,b,T). ((F,D'),b,T)) fs @ pre" in exI) simp
-qed
-(*>*)
-
-lemma has_fields_declaring_classes:
-shows "P \<turnstile> C has_fields FDTs
- \<Longrightarrow> \<exists>pre FDTs'. FDTs = pre@FDTs'
- \<and> (C \<noteq> Object \<longrightarrow> (\<exists>D fs ms. class P C = \<lfloor>(D,fs,ms)\<rfloor> \<and> P \<turnstile> D has_fields FDTs'))
- \<and> set(map (\<lambda>t. snd(fst t)) pre) \<subseteq> {C}
- \<and> set(map (\<lambda>t. snd(fst t)) FDTs') \<subseteq> {C'. C' \<noteq> C \<and> P \<turnstile> C \<preceq>\<^sup>* C'}"
-proof(induct rule:Fields.induct)
- case (has_fields_rec C D fs ms FDTs FDTs')
- have sup1: "P \<turnstile> C \<prec>\<^sup>1 D" using has_fields_rec.hyps(1,2) by (simp add: subcls1.subcls1I)
- have "P \<turnstile> C has_fields FDTs'"
- using Fields.has_fields_rec[OF has_fields_rec.hyps(1-3)] has_fields_rec by auto
- then have nsup: "(D, C) \<notin> (subcls1 P)\<^sup>*" using subcls_notin_has_fields2 sup1 by auto
- show ?case using has_fields_rec sup1 nsup
- by(rule_tac x = "map (\<lambda>(F, y). ((F, C), y)) fs" in exI, clarsimp) auto
-next
- case has_fields_Object then show ?case by fastforce
-qed
-
-lemma has_fields_mono_lem2:
-assumes hf: "P \<turnstile> C has_fields FDTs"
- and cls: "class P C = Some(D,fs,ms)" and map_of: "map_of FDTs (F,C) = \<lfloor>(b,T)\<rfloor>"
-shows "\<exists>FDTs'. FDTs = (map (\<lambda>(F,b,T). ((F,C),b,T)) fs) @ FDTs' \<and> map_of FDTs' (F,C) = None"
-using assms
-proof(cases "C = Object")
- case False
- let ?pre = "map (\<lambda>(F,b,T). ((F,C),b,T)) fs"
- have sub: "P \<turnstile> C \<preceq>\<^sup>* D" using cls False by (simp add: r_into_rtrancl subcls1.subcls1I)
- obtain FDTs' where fdts': "P \<turnstile> D has_fields FDTs'" "FDTs = ?pre @ FDTs'"
- using False assms(1,2) Fields.simps[of P C FDTs] by clarsimp
- then have int: "dom (map_of ?pre) \<inter> dom (map_of FDTs') = {}"
- using has_fields_mono_lem[OF sub, of FDTs'] has_fields_fun[OF hf] by fastforce
- have "C \<notin> (\<lambda>t. snd (fst t)) ` set FDTs'"
- using has_fields_declaring_classes[OF hf] cls False
- has_fields_fun[OF fdts'(1)] fdts'(2)
- by clarify auto
- then have "map_of FDTs' (F,C) = None" by(rule map_of_set_pcs_notin)
- then show ?thesis using fdts' int by simp
-qed(auto dest: has_fields_Object has_fields_fun)
-
-
-lemma has_fields_is_class_Object:
- "P \<turnstile> D has_fields FDTs \<Longrightarrow> is_class P Object"
- by(induct rule: Fields.induct; simp add: is_class_def)
-
-lemma Object_fields:
- "\<lbrakk> P \<turnstile> Object has_fields FDTs; C \<noteq> Object \<rbrakk> \<Longrightarrow> map_of FDTs (F,C) = None"
- by(drule Fields.cases, auto simp: map_of_reinsert_neq_None)
-
-
-definition has_field :: "'m prog \<Rightarrow> cname \<Rightarrow> vname \<Rightarrow> staticb \<Rightarrow> ty \<Rightarrow> cname \<Rightarrow> bool"
- ("_ \<turnstile> _ has _,_:_ in _" [51,51,51,51,51,51] 50)
-where
- "P \<turnstile> C has F,b:T in D \<equiv>
- \<exists>FDTs. P \<turnstile> C has_fields FDTs \<and> map_of FDTs (F,D) = Some (b,T)"
-
-
-lemma has_field_mono:
-assumes has: " P \<turnstile> C has F,b:T in D" and sub: "P \<turnstile> C' \<preceq>\<^sup>* C"
-shows "P \<turnstile> C' has F,b:T in D"
-(*<*)
-proof -
- obtain FDTs where FDTs:"P \<turnstile> C has_fields FDTs" and "map_of FDTs (F, D) = \<lfloor>(b, T)\<rfloor>"
- using has by(clarsimp simp: has_field_def)
- also obtain pre where "P \<turnstile> C' has_fields pre @ FDTs"
- and "dom (map_of pre) \<inter> dom (map_of FDTs) = {}"
- using has_fields_mono_lem[OF sub FDTs] by clarify
- ultimately show ?thesis by(fastforce simp: has_field_def map_add_def split:option.splits)
-qed
-(*>*)
-
-lemma has_field_fun:
- "\<lbrakk>P \<turnstile> C has F,b:T in D; P \<turnstile> C has F,b':T' in D\<rbrakk> \<Longrightarrow> b = b' \<and> T' = T"
-(*<*)by(fastforce simp:has_field_def dest:has_fields_fun)(*>*)
-
-
-lemma has_field_idemp:
-assumes has: "P \<turnstile> C has F,b:T in D"
-shows "P \<turnstile> D has F,b:T in D"
-(*<*)
-proof -
- obtain FDTs where C_flds: "P \<turnstile> C has_fields FDTs"
- and FDTs: "map_of FDTs (F, D) = \<lfloor>(b, T)\<rfloor>" (is "?FDTs")
- using has by(clarsimp simp: has_field_def)
- have map: "\<And>C' fs. map_of (map (\<lambda>(F, y). ((F, C'), y)) fs) (F, D) = \<lfloor>(b, T)\<rfloor> \<Longrightarrow> D = C'"
- by(frule map_of_SomeD) clarsimp
- have "?FDTs \<longrightarrow> P \<turnstile> D has F,b:T in D"
- using C_flds proof induct
- case NObj: (has_fields_rec C' D' fs ms FDTs FDTs')
- then show ?case using map by (fastforce intro: has_fields_rec simp: has_field_def)
- next
- case Obj: (has_fields_Object D fs ms FDTs)
- then show ?case using map by(fastforce intro: has_fields_Object simp: has_field_def)
- qed
- then show ?thesis using FDTs by(rule_tac mp)
-qed
-(*>*)
-
-lemma visible_fields_exist:
-assumes fields: "P \<turnstile> C has_fields FDTs" and
- FDTs: "map_of FDTs (F,D) = Some (b, T)"
-shows "\<exists>D' fs ms. class P D = Some(D',fs,ms) \<and> map_of fs F = Some(b,T)"
-proof -
- have "map_of FDTs (F,D) = Some (b, T) \<longrightarrow>
- (\<exists>D' fs ms. class P D = Some(D',fs,ms) \<and> map_of fs F = Some(b,T))"
- using fields proof induct
- case (has_fields_rec C' D' fs ms FDTs')
- with assms map_of_reinsert_SomeD map_of_reinsert_neq_None[where D=D and F=F and fs=fs]
- show ?case proof(cases "C' = D") qed auto
- next
- case (has_fields_Object D' fs ms FDTs)
- with assms map_of_reinsert_SomeD map_of_reinsert_neq_None[where D=D and F=F and fs=fs]
- show ?case proof(cases "Object = D") qed auto
- qed
- then show ?thesis using FDTs by simp
-qed
-
-lemma map_of_remap_SomeD:
- "map_of (map (\<lambda>((k,k'),x). (k,(k',x))) t) k = Some (k',x) \<Longrightarrow> map_of t (k, k') = Some x"
-(*<*)by (induct t) (auto simp:fun_upd_apply split: if_split_asm)(*>*)
-
-lemma map_of_remap_SomeD2:
- "map_of (map (\<lambda>((k,k'),x,x'). (k,(k',x,x'))) t) k = Some (k',x,x') \<Longrightarrow> map_of t (k, k') = Some (x, x')"
-(*<*)by (induct t) (auto simp:fun_upd_apply split: if_split_asm)(*>*)
-
-lemma has_field_decl_above:
- "P \<turnstile> C has F,b:T in D \<Longrightarrow> P \<turnstile> C \<preceq>\<^sup>* D"
-(*<*)
-by(auto simp: has_field_def
- intro: has_fields_decl_above map_of_SomeD map_of_remap_SomeD2)
-(*>*)
-
-definition sees_field :: "'m prog \<Rightarrow> cname \<Rightarrow> vname \<Rightarrow> staticb \<Rightarrow> ty \<Rightarrow> cname \<Rightarrow> bool"
- ("_ \<turnstile> _ sees _,_:_ in _" [51,51,51,51,51,51] 50)
-where
- "P \<turnstile> C sees F,b:T in D \<equiv>
- \<exists>FDTs. P \<turnstile> C has_fields FDTs \<and>
- map_of (map (\<lambda>((F,D),b,T). (F,(D,b,T))) FDTs) F = Some(D,b,T)"
-
-lemma has_visible_field:
- "P \<turnstile> C sees F,b:T in D \<Longrightarrow> P \<turnstile> C has F,b:T in D"
-(*<*)by(auto simp add:has_field_def sees_field_def map_of_remap_SomeD2)(*>*)
-
-lemma sees_field_fun:
- "\<lbrakk>P \<turnstile> C sees F,b:T in D; P \<turnstile> C sees F,b':T' in D'\<rbrakk> \<Longrightarrow> b = b' \<and> T' = T \<and> D' = D"
-(*<*)by(fastforce simp:sees_field_def dest:has_fields_fun)(*>*)
-
-lemma sees_field_decl_above:
- "P \<turnstile> C sees F,b:T in D \<Longrightarrow> P \<turnstile> C \<preceq>\<^sup>* D"
-(*<*)
-by(auto simp:sees_field_def
- intro: has_fields_decl_above map_of_SomeD map_of_remap_SomeD2)
-(*>*)
-
-
-lemma sees_field_idemp:
-assumes sees: "P \<turnstile> C sees F,b:T in D"
-shows "P \<turnstile> D sees F,b:T in D"
-(*<*)
-proof -
- obtain FDTs where C_flds: "P \<turnstile> C has_fields FDTs"
- and FDTs: "map_of (map (\<lambda>((F, D), b, T). (F, D, b, T)) FDTs) F = \<lfloor>(D, b, T)\<rfloor>"
- (is "?FDTs")
- using sees by(clarsimp simp: sees_field_def)
- have map: "\<And>C' fs. map_of (map ((\<lambda>((F, D), a). (F, D, a)) \<circ> (\<lambda>(F, y). ((F, C'), y))) fs) F
- = \<lfloor>(D, b, T)\<rfloor>
- \<Longrightarrow> D = C' \<and> (F, b, T) \<in> set fs"
- by(frule map_of_SomeD) clarsimp
-\<comment>\<open> ?FDTs \<longrightarrow> P \<turnstile> D sees F,b:T in D \<close>
- have "?FDTs \<longrightarrow> (\<exists>FDTs. P \<turnstile> D has_fields FDTs
- \<and> map_of (map (\<lambda>((F, D), a). (F, D, a)) FDTs) F = \<lfloor>(D, b, T)\<rfloor>)"
- using C_flds proof induct
- case NObj: (has_fields_rec C' D' fs ms FDTs FDTs')
- then show ?case using map by (fastforce intro: has_fields_rec)
- next
- case Obj: (has_fields_Object D fs ms FDTs)
- then show ?case using map by(fastforce intro: has_fields_Object)
- qed
- then show ?thesis using FDTs
- by (smt map_eq_conv old.prod.case prod_cases3 sees_field_def split_cong)
-qed
-(*>*)
-
-lemma has_field_sees_aux:
-assumes hf: "P \<turnstile> C has_fields FDTs" and map: "map_of FDTs (F, C) = \<lfloor>(b, T)\<rfloor>"
-shows "map_of (map (\<lambda>((F, D), b, T). (F, D, b, T)) FDTs) F = \<lfloor>(C, b, T)\<rfloor>"
-proof -
- obtain D fs ms where fs: "class P C = Some(D,fs,ms)"
- using visible_fields_exist[OF assms] by clarsimp
- then obtain FDTs' where
- "FDTs = map (\<lambda>(F, b, T). ((F, C), b, T)) fs @ FDTs' \<and> map_of FDTs' (F, C) = None"
- using has_fields_mono_lem2[OF hf fs map] by clarsimp
- then show ?thesis using map_of_Some_None_split[OF _ _ map] by auto
-qed
-
-lemma has_field_sees: "P \<turnstile> C has F,b:T in C \<Longrightarrow> P \<turnstile> C sees F,b:T in C"
- by(auto simp:has_field_def sees_field_def has_field_sees_aux)
-
-lemma has_field_is_class:
- "P \<turnstile> C has F,b:T in D \<Longrightarrow> is_class P C"
-(*<*)by (auto simp add: is_class_def has_field_def elim: Fields.induct)(*>*)
-
-lemma has_field_is_class':
- "P \<turnstile> C has F,b:T in D \<Longrightarrow> is_class P D"
-(*<*)by(drule has_field_idemp, rule has_field_is_class, assumption)(*>*)
-
-subsection "Functional lookup"
-
-definition "method" :: "'m prog \<Rightarrow> cname \<Rightarrow> mname \<Rightarrow> cname \<times> staticb \<times> ty list \<times> ty \<times> 'm"
-where
- "method P C M \<equiv> THE (D,b,Ts,T,m). P \<turnstile> C sees M,b:Ts \<rightarrow> T = m in D"
-
-definition field :: "'m prog \<Rightarrow> cname \<Rightarrow> vname \<Rightarrow> cname \<times> staticb \<times> ty"
-where
- "field P C F \<equiv> THE (D,b,T). P \<turnstile> C sees F,b:T in D"
-
-definition fields :: "'m prog \<Rightarrow> cname \<Rightarrow> ((vname \<times> cname) \<times> staticb \<times> ty) list"
-where
- "fields P C \<equiv> THE FDTs. P \<turnstile> C has_fields FDTs"
-
-lemma fields_def2 [simp]: "P \<turnstile> C has_fields FDTs \<Longrightarrow> fields P C = FDTs"
-(*<*)by (unfold fields_def) (auto dest: has_fields_fun)(*>*)
-
-lemma field_def2 [simp]: "P \<turnstile> C sees F,b:T in D \<Longrightarrow> field P C F = (D,b,T)"
-(*<*)by (unfold field_def) (auto dest: sees_field_fun)(*>*)
-
-lemma method_def2 [simp]: "P \<turnstile> C sees M,b: Ts\<rightarrow>T = m in D \<Longrightarrow> method P C M = (D,b,Ts,T,m)"
-(*<*)by (unfold method_def) (auto dest: sees_method_fun)(*>*)
-
-
-text \<open> The following are the fields for initializing an object (non-static fields)
- and a class (just that class's static fields), respectively. \<close>
-
-definition ifields :: "'m prog \<Rightarrow> cname \<Rightarrow> ((vname \<times> cname) \<times> staticb \<times> ty) list"
-where
- "ifields P C \<equiv> filter (\<lambda>((F,D),b,T). b = NonStatic) (fields P C)"
-
-definition isfields :: "'m prog \<Rightarrow> cname \<Rightarrow> ((vname \<times> cname) \<times> staticb \<times> ty) list"
-where
- "isfields P C \<equiv> filter (\<lambda>((F,D),b,T). b = Static \<and> D = C) (fields P C)"
-
-lemma ifields_def2[simp]: "\<lbrakk> P \<turnstile> C has_fields FDTs \<rbrakk> \<Longrightarrow> ifields P C = filter (\<lambda>((F,D),b,T). b = NonStatic) FDTs"
- by (simp add: ifields_def)
-
-lemma isfields_def2[simp]: "\<lbrakk> P \<turnstile> C has_fields FDTs \<rbrakk> \<Longrightarrow> isfields P C = filter (\<lambda>((F,D),b,T). b = Static \<and> D = C) FDTs"
- by (simp add: isfields_def)
-
-lemma ifields_def3: "\<lbrakk> P \<turnstile> C sees F,b:T in D; b = NonStatic \<rbrakk> \<Longrightarrow> (((F,D),b,T) \<in> set (ifields P C))"
-(*<*) by (unfold ifields_def) (auto simp: sees_field_def map_of_SomeD map_of_remap_SomeD2) (*>*)
-
-lemma isfields_def3: "\<lbrakk> P \<turnstile> C sees F,b:T in D; b = Static; D = C \<rbrakk> \<Longrightarrow> (((F,D),b,T) \<in> set (isfields P C))"
-(*<*) by (unfold isfields_def) (auto simp: sees_field_def map_of_SomeD map_of_remap_SomeD2) (*>*)
-
-
-definition seeing_class :: "'m prog \<Rightarrow> cname \<Rightarrow> mname \<Rightarrow> cname option" where
-"seeing_class P C M =
- (if \<exists>Ts T m D. P \<turnstile> C sees M,Static:Ts\<rightarrow>T = m in D
- then Some (fst(method P C M))
- else None)"
-
-lemma seeing_class_def2[simp]:
- "P \<turnstile> C sees M,Static:Ts\<rightarrow>T = m in D \<Longrightarrow> seeing_class P C M = Some D"
- by(fastforce simp: seeing_class_def)
-
-(*<*)
-end
-(*>*)
+(* Title: JinjaDCI/Common/TypeRel.thy
+ Author: Tobias Nipkow, Susannah Mansky
+ Copyright 2003 Technische Universitaet Muenchen, 2019-20 UIUC
+
+ Based on the Jinja theory Common/TypeRel.thy by Tobias Nipkow
+*)
+
+section \<open> Relations between Jinja Types \<close>
+
+theory TypeRel imports
+ "HOL-Library.Transitive_Closure_Table"
+ Decl
+begin
+
+subsection \<open> The subclass relations \<close>
+
+inductive_set
+ subcls1 :: "'m prog \<Rightarrow> (cname \<times> cname) set"
+ and subcls1' :: "'m prog \<Rightarrow> [cname, cname] \<Rightarrow> bool" ("_ \<turnstile> _ \<prec>\<^sup>1 _" [71,71,71] 70)
+ for P :: "'m prog"
+where
+ "P \<turnstile> C \<prec>\<^sup>1 D \<equiv> (C,D) \<in> subcls1 P"
+| subcls1I: "\<lbrakk>class P C = Some (D,rest); C \<noteq> Object\<rbrakk> \<Longrightarrow> P \<turnstile> C \<prec>\<^sup>1 D"
+
+abbreviation
+ subcls :: "'m prog \<Rightarrow> [cname, cname] \<Rightarrow> bool" ("_ \<turnstile> _ \<preceq>\<^sup>* _" [71,71,71] 70)
+ where "P \<turnstile> C \<preceq>\<^sup>* D \<equiv> (C,D) \<in> (subcls1 P)\<^sup>*"
+
+lemma subcls1D: "P \<turnstile> C \<prec>\<^sup>1 D \<Longrightarrow> C \<noteq> Object \<and> (\<exists>fs ms. class P C = Some (D,fs,ms))"
+(*<*)by(erule subcls1.induct)(fastforce simp add:is_class_def)(*>*)
+
+lemma [iff]: "\<not> P \<turnstile> Object \<prec>\<^sup>1 C"
+(*<*)by(fastforce dest:subcls1D)(*>*)
+
+lemma [iff]: "(P \<turnstile> Object \<preceq>\<^sup>* C) = (C = Object)"
+(*<*)
+proof(rule iffI)
+ assume "P \<turnstile> Object \<preceq>\<^sup>* C" then show "C = Object"
+ by(auto elim: converse_rtranclE)
+qed simp
+(*>*)
+
+lemma subcls1_def2:
+ "subcls1 P =
+ (SIGMA C:{C. is_class P C}. {D. C\<noteq>Object \<and> fst (the (class P C))=D})"
+(*<*)
+ by (fastforce simp:is_class_def dest: subcls1D elim: subcls1I)
+(*>*)
+
+lemma finite_subcls1: "finite (subcls1 P)"
+(*<*)
+proof -
+ let ?SIG = "SIGMA C:{C. is_class P C}. {D. fst (the (class P C)) = D \<and> C \<noteq> Object}"
+ have "subcls1 P = ?SIG" by(simp add: subcls1_def2)
+ also have "finite ?SIG"
+ proof(rule finite_SigmaI [OF finite_is_class])
+ fix C assume C_in: "C \<in> {C. is_class P C}"
+ then show "finite {D. fst (the (class P C)) = D \<and> C \<noteq> Object}"
+ by(rule_tac finite_subset[where B = "{fst (the (class P C))}"]) auto
+ qed
+ ultimately show ?thesis by simp
+qed
+(*>*)
+
+primrec supercls_lst :: "'m prog \<Rightarrow> cname list \<Rightarrow> bool" where
+"supercls_lst P (C#Cs) = ((\<forall>C' \<in> set Cs. P \<turnstile> C' \<preceq>\<^sup>* C) \<and> supercls_lst P Cs)" |
+"supercls_lst P [] = True"
+
+lemma supercls_lst_app:
+ "\<lbrakk> supercls_lst P (C#Cs); P \<turnstile> C \<preceq>\<^sup>* C' \<rbrakk> \<Longrightarrow> supercls_lst P (C'#C#Cs)"
+ by auto
+
+subsection\<open> The subtype relations \<close>
+
+inductive
+ widen :: "'m prog \<Rightarrow> ty \<Rightarrow> ty \<Rightarrow> bool" ("_ \<turnstile> _ \<le> _" [71,71,71] 70)
+ for P :: "'m prog"
+where
+ widen_refl[iff]: "P \<turnstile> T \<le> T"
+| widen_subcls: "P \<turnstile> C \<preceq>\<^sup>* D \<Longrightarrow> P \<turnstile> Class C \<le> Class D"
+| widen_null[iff]: "P \<turnstile> NT \<le> Class C"
+
+abbreviation
+ widens :: "'m prog \<Rightarrow> ty list \<Rightarrow> ty list \<Rightarrow> bool"
+ ("_ \<turnstile> _ [\<le>] _" [71,71,71] 70) where
+ "widens P Ts Ts' \<equiv> list_all2 (widen P) Ts Ts'"
+
+lemma [iff]: "(P \<turnstile> T \<le> Void) = (T = Void)"
+(*<*)by (auto elim: widen.cases)(*>*)
+
+lemma [iff]: "(P \<turnstile> T \<le> Boolean) = (T = Boolean)"
+(*<*)by (auto elim: widen.cases)(*>*)
+
+lemma [iff]: "(P \<turnstile> T \<le> Integer) = (T = Integer)"
+(*<*)by (auto elim: widen.cases)(*>*)
+
+lemma [iff]: "(P \<turnstile> Void \<le> T) = (T = Void)"
+(*<*)by (auto elim: widen.cases)(*>*)
+
+lemma [iff]: "(P \<turnstile> Boolean \<le> T) = (T = Boolean)"
+(*<*)by (auto elim: widen.cases)(*>*)
+
+lemma [iff]: "(P \<turnstile> Integer \<le> T) = (T = Integer)"
+(*<*)by (auto elim: widen.cases)(*>*)
+
+
+lemma Class_widen: "P \<turnstile> Class C \<le> T \<Longrightarrow> \<exists>D. T = Class D"
+(*<*)
+by (ind_cases "P \<turnstile> Class C \<le> T") auto
+(*>*)
+
+lemma [iff]: "(P \<turnstile> T \<le> NT) = (T = NT)"
+(*<*)
+by(cases T) (auto dest:Class_widen)
+(*>*)
+
+lemma Class_widen_Class [iff]: "(P \<turnstile> Class C \<le> Class D) = (P \<turnstile> C \<preceq>\<^sup>* D)"
+(*<*)
+proof(rule iffI)
+ show "P \<turnstile> Class C \<le> Class D \<Longrightarrow> P \<turnstile> C \<preceq>\<^sup>* D"
+ proof(ind_cases "P \<turnstile> Class C \<le> Class D") qed(auto)
+qed(auto elim: widen_subcls)
+(*>*)
+
+lemma widen_Class: "(P \<turnstile> T \<le> Class C) = (T = NT \<or> (\<exists>D. T = Class D \<and> P \<turnstile> D \<preceq>\<^sup>* C))"
+(*<*)by(induct T, auto)(*>*)
+
+
+lemma widen_trans[trans]: "\<lbrakk>P \<turnstile> S \<le> U; P \<turnstile> U \<le> T\<rbrakk> \<Longrightarrow> P \<turnstile> S \<le> T"
+(*<*)
+proof -
+ assume "P\<turnstile>S \<le> U" thus "\<And>T. P \<turnstile> U \<le> T \<Longrightarrow> P \<turnstile> S \<le> T"
+ proof induct
+ case (widen_refl T T') thus "P \<turnstile> T \<le> T'" .
+ next
+ case (widen_subcls C D T)
+ then obtain E where "T = Class E" by (blast dest: Class_widen)
+ with widen_subcls show "P \<turnstile> Class C \<le> T" by (auto elim: rtrancl_trans)
+ next
+ case (widen_null C RT)
+ then obtain D where "RT = Class D" by (blast dest: Class_widen)
+ thus "P \<turnstile> NT \<le> RT" by auto
+ qed
+qed
+(*>*)
+
+lemma widens_trans [trans]: "\<lbrakk>P \<turnstile> Ss [\<le>] Ts; P \<turnstile> Ts [\<le>] Us\<rbrakk> \<Longrightarrow> P \<turnstile> Ss [\<le>] Us"
+(*<*)by (rule list_all2_trans, rule widen_trans)(*>*)
+
+
+(*<*)
+lemmas widens_refl [iff] = list_all2_refl [of "widen P", OF widen_refl] for P
+lemmas widens_Cons [iff] = list_all2_Cons1 [of "widen P"] for P
+(*>*)
+
+
+subsection\<open> Method lookup \<close>
+
+inductive
+ Methods :: "['m prog, cname, mname \<rightharpoonup> (staticb \<times> ty list \<times> ty \<times> 'm) \<times> cname] \<Rightarrow> bool"
+ ("_ \<turnstile> _ sees'_methods _" [51,51,51] 50)
+ for P :: "'m prog"
+where
+ sees_methods_Object:
+ "\<lbrakk> class P Object = Some(D,fs,ms); Mm = map_option (\<lambda>m. (m,Object)) \<circ> map_of ms \<rbrakk>
+ \<Longrightarrow> P \<turnstile> Object sees_methods Mm"
+| sees_methods_rec:
+ "\<lbrakk> class P C = Some(D,fs,ms); C \<noteq> Object; P \<turnstile> D sees_methods Mm;
+ Mm' = Mm ++ (map_option (\<lambda>m. (m,C)) \<circ> map_of ms) \<rbrakk>
+ \<Longrightarrow> P \<turnstile> C sees_methods Mm'"
+
+lemma sees_methods_fun:
+assumes 1: "P \<turnstile> C sees_methods Mm"
+shows "\<And>Mm'. P \<turnstile> C sees_methods Mm' \<Longrightarrow> Mm' = Mm"
+ (*<*)
+using 1
+proof induct
+ case (sees_methods_rec C D fs ms Dres Cres Cres')
+ have "class": "class P C = Some (D, fs, ms)"
+ and notObj: "C \<noteq> Object" and Dmethods: "P \<turnstile> D sees_methods Dres"
+ and IH: "\<And>Dres'. P \<turnstile> D sees_methods Dres' \<Longrightarrow> Dres' = Dres"
+ and Cres: "Cres = Dres ++ (map_option (\<lambda>m. (m,C)) \<circ> map_of ms)"
+ and Cmethods': "P \<turnstile> C sees_methods Cres'" by fact+
+ from Cmethods' notObj "class" obtain Dres'
+ where Dmethods': "P \<turnstile> D sees_methods Dres'"
+ and Cres': "Cres' = Dres' ++ (map_option (\<lambda>m. (m,C)) \<circ> map_of ms)"
+ by(auto elim: Methods.cases)
+ from Cres Cres' IH[OF Dmethods'] show "Cres' = Cres" by simp
+next
+ case sees_methods_Object thus ?case by(auto elim: Methods.cases)
+qed
+(*>*)
+
+lemma visible_methods_exist:
+ "P \<turnstile> C sees_methods Mm \<Longrightarrow> Mm M = Some(m,D) \<Longrightarrow>
+ (\<exists>D' fs ms. class P D = Some(D',fs,ms) \<and> map_of ms M = Some m)"
+ (*<*)by(induct rule:Methods.induct) auto(*>*)
+
+lemma sees_methods_decl_above:
+assumes Csees: "P \<turnstile> C sees_methods Mm"
+shows "Mm M = Some(m,D) \<Longrightarrow> P \<turnstile> C \<preceq>\<^sup>* D"
+ (*<*)
+using Csees
+proof induct
+next
+ case sees_methods_Object thus ?case by auto
+next
+ case sees_methods_rec thus ?case
+ by(fastforce simp:map_option_case split:option.splits
+ elim:converse_rtrancl_into_rtrancl[OF subcls1I])
+qed
+(*>*)
+
+lemma sees_methods_idemp:
+assumes Cmethods: "P \<turnstile> C sees_methods Mm"
+shows "\<And>m D. Mm M = Some(m,D) \<Longrightarrow>
+ \<exists>Mm'. (P \<turnstile> D sees_methods Mm') \<and> Mm' M = Some(m,D)"
+(*<*)
+using Cmethods
+proof induct
+ case sees_methods_Object thus ?case
+ by(fastforce dest: Methods.sees_methods_Object)
+next
+ case sees_methods_rec thus ?case
+ by(fastforce split:option.splits dest: Methods.sees_methods_rec)
+qed
+(*>*)
+
+(*FIXME something wrong with induct: need to attache [consumes 1]
+directly to proof of thm, declare does not work. *)
+
+lemma sees_methods_decl_mono:
+assumes sub: "P \<turnstile> C' \<preceq>\<^sup>* C"
+shows "P \<turnstile> C sees_methods Mm \<Longrightarrow>
+ \<exists>Mm' Mm\<^sub>2. P \<turnstile> C' sees_methods Mm' \<and> Mm' = Mm ++ Mm\<^sub>2 \<and>
+ (\<forall>M m D. Mm\<^sub>2 M = Some(m,D) \<longrightarrow> P \<turnstile> D \<preceq>\<^sup>* C)"
+(*<*)
+ (is "_ \<Longrightarrow> \<exists>Mm' Mm2. ?Q C' C Mm' Mm2")
+using sub
+proof (induct rule:converse_rtrancl_induct)
+ assume "P \<turnstile> C sees_methods Mm"
+ hence "?Q C C Mm Map.empty" by simp
+ thus "\<exists>Mm' Mm2. ?Q C C Mm' Mm2" by blast
+next
+ fix C'' C'
+ assume sub1: "P \<turnstile> C'' \<prec>\<^sup>1 C'" and sub: "P \<turnstile> C' \<preceq>\<^sup>* C"
+ and IH: "P \<turnstile> C sees_methods Mm \<Longrightarrow>
+ \<exists>Mm' Mm2. P \<turnstile> C' sees_methods Mm' \<and>
+ Mm' = Mm ++ Mm2 \<and> (\<forall>M m D. Mm2 M = Some(m,D) \<longrightarrow> P \<turnstile> D \<preceq>\<^sup>* C)"
+ and Csees: "P \<turnstile> C sees_methods Mm"
+ from IH[OF Csees] obtain Mm' Mm2 where C'sees: "P \<turnstile> C' sees_methods Mm'"
+ and Mm': "Mm' = Mm ++ Mm2"
+ and subC:"\<forall>M m D. Mm2 M = Some(m,D) \<longrightarrow> P \<turnstile> D \<preceq>\<^sup>* C" by blast
+ obtain fs ms where "class": "class P C'' = Some(C',fs,ms)" "C'' \<noteq> Object"
+ using subcls1D[OF sub1] by blast
+ let ?Mm3 = "map_option (\<lambda>m. (m,C'')) \<circ> map_of ms"
+ have "P \<turnstile> C'' sees_methods (Mm ++ Mm2) ++ ?Mm3"
+ using sees_methods_rec[OF "class" C'sees refl] Mm' by simp
+ hence "?Q C'' C ((Mm ++ Mm2) ++ ?Mm3) (Mm2++?Mm3)"
+ using converse_rtrancl_into_rtrancl[OF sub1 sub]
+ by simp (simp add:map_add_def subC split:option.split)
+ thus "\<exists>Mm' Mm2. ?Q C'' C Mm' Mm2" by blast
+qed
+(*>*)
+
+lemma sees_methods_is_class_Object:
+ "P \<turnstile> D sees_methods Mm \<Longrightarrow> is_class P Object"
+ by(induct rule: Methods.induct; simp add: is_class_def)
+
+lemma sees_methods_sub_Obj: "P \<turnstile> C sees_methods Mm \<Longrightarrow> P \<turnstile> C \<preceq>\<^sup>* Object"
+proof(induct rule: Methods.induct)
+ case (sees_methods_rec C D fs ms Mm Mm') show ?case
+ using subcls1I[OF sees_methods_rec.hyps(1,2)] sees_methods_rec.hyps(4)
+ by(rule converse_rtrancl_into_rtrancl)
+qed(simp)
+
+
+definition Method :: "'m prog \<Rightarrow> cname \<Rightarrow> mname \<Rightarrow> staticb \<Rightarrow> ty list \<Rightarrow> ty \<Rightarrow> 'm \<Rightarrow> cname \<Rightarrow> bool"
+ ("_ \<turnstile> _ sees _, _ : _\<rightarrow>_ = _ in _" [51,51,51,51,51,51,51,51] 50)
+where
+ "P \<turnstile> C sees M, b: Ts\<rightarrow>T = m in D \<equiv>
+ \<exists>Mm. P \<turnstile> C sees_methods Mm \<and> Mm M = Some((b,Ts,T,m),D)"
+
+definition has_method :: "'m prog \<Rightarrow> cname \<Rightarrow> mname \<Rightarrow> staticb \<Rightarrow> bool"
+ ("_ \<turnstile> _ has _, _" [51,0,0,51] 50)
+where
+ "P \<turnstile> C has M, b \<equiv> \<exists>Ts T m D. P \<turnstile> C sees M,b:Ts\<rightarrow>T = m in D"
+
+lemma sees_method_fun:
+ "\<lbrakk>P \<turnstile> C sees M,b:TS\<rightarrow>T = m in D; P \<turnstile> C sees M,b':TS'\<rightarrow>T' = m' in D' \<rbrakk>
+ \<Longrightarrow> b = b' \<and> TS' = TS \<and> T' = T \<and> m' = m \<and> D' = D"
+ (*<*)by(fastforce dest: sees_methods_fun simp:Method_def)(*>*)
+
+lemma sees_method_decl_above:
+ "P \<turnstile> C sees M,b:Ts\<rightarrow>T = m in D \<Longrightarrow> P \<turnstile> C \<preceq>\<^sup>* D"
+ (*<*)by(clarsimp simp:Method_def sees_methods_decl_above)(*>*)
+
+lemma visible_method_exists:
+ "P \<turnstile> C sees M,b:Ts\<rightarrow>T = m in D \<Longrightarrow>
+ \<exists>D' fs ms. class P D = Some(D',fs,ms) \<and> map_of ms M = Some(b,Ts,T,m)"
+(*<*)by(fastforce simp:Method_def dest!: visible_methods_exist)(*>*)
+
+
+lemma sees_method_idemp:
+ "P \<turnstile> C sees M,b:Ts\<rightarrow>T=m in D \<Longrightarrow> P \<turnstile> D sees M,b:Ts\<rightarrow>T=m in D"
+ (*<*)by(fastforce simp: Method_def intro:sees_methods_idemp)(*>*)
+
+lemma sees_method_decl_mono:
+assumes sub: "P \<turnstile> C' \<preceq>\<^sup>* C" and
+ C_sees: "P \<turnstile> C sees M,b:Ts\<rightarrow>T=m in D" and
+ C'_sees: "P \<turnstile> C' sees M,b':Ts'\<rightarrow>T'=m' in D'"
+shows "P \<turnstile> D' \<preceq>\<^sup>* D"
+ (*<*)
+proof -
+ obtain Ms where Ms: "P \<turnstile> C sees_methods Ms"
+ using C_sees by(auto simp: Method_def)
+ obtain Ms' Ms2 where Ms': "P \<turnstile> C' sees_methods Ms'" and
+ Ms'_def: "Ms' = Ms ++ Ms2" and
+ Ms2_imp: "(\<forall>M m D. Ms2 M = \<lfloor>(m, D)\<rfloor> \<longrightarrow> P \<turnstile> D \<preceq>\<^sup>* C)"
+ using sees_methods_decl_mono[OF sub Ms] by clarsimp
+ have "(Ms ++ Ms2) M = \<lfloor>((b', Ts', T', m'), D')\<rfloor>"
+ using C'_sees sees_methods_fun[OF Ms'] Ms'_def by(clarsimp simp: Method_def)
+ then have "Ms2 M = \<lfloor>((b', Ts', T', m'), D')\<rfloor> \<or>
+ Ms2 M = None \<and> b = b' \<and> Ts = Ts' \<and> T = T' \<and> m = m' \<and> D = D'"
+ using C_sees sees_methods_fun[OF Ms] by(clarsimp simp: Method_def)
+ also have "Ms2 M = \<lfloor>((b', Ts', T', m'), D')\<rfloor> \<Longrightarrow> P \<turnstile> D' \<preceq>\<^sup>* C"
+ using Ms2_imp by simp
+ ultimately show ?thesis using sub sees_method_decl_above[OF C_sees] by auto
+qed
+(*>*)
+
+lemma sees_methods_is_class: "P \<turnstile> C sees_methods Mm \<Longrightarrow> is_class P C"
+(*<*)by (auto simp add: is_class_def elim: Methods.induct)(*>*)
+
+lemma sees_method_is_class:
+ "\<lbrakk> P \<turnstile> C sees M,b:Ts\<rightarrow>T=m in D \<rbrakk> \<Longrightarrow> is_class P C"
+(*<*)by (auto simp add: is_class_def Method_def dest: sees_methods_is_class)(*>*)
+
+lemma sees_method_is_class':
+ "\<lbrakk> P \<turnstile> C sees M,b:Ts\<rightarrow>T=m in D \<rbrakk> \<Longrightarrow> is_class P D"
+(*<*)by(drule sees_method_idemp, rule sees_method_is_class, assumption)(*>*)
+
+lemma sees_method_sub_Obj: "P \<turnstile> C sees M,b: Ts\<rightarrow>T = m in D \<Longrightarrow> P \<turnstile> C \<preceq>\<^sup>* Object"
+ by(auto simp: Method_def sees_methods_sub_Obj)
+
+subsection\<open> Field lookup \<close>
+
+inductive
+ Fields :: "['m prog, cname, ((vname \<times> cname) \<times> staticb \<times> ty) list] \<Rightarrow> bool"
+ ("_ \<turnstile> _ has'_fields _" [51,51,51] 50)
+ for P :: "'m prog"
+where
+ has_fields_rec:
+ "\<lbrakk> class P C = Some(D,fs,ms); C \<noteq> Object; P \<turnstile> D has_fields FDTs;
+ FDTs' = map (\<lambda>(F,b,T). ((F,C),b,T)) fs @ FDTs \<rbrakk>
+ \<Longrightarrow> P \<turnstile> C has_fields FDTs'"
+| has_fields_Object:
+ "\<lbrakk> class P Object = Some(D,fs,ms); FDTs = map (\<lambda>(F,b,T). ((F,Object),b,T)) fs \<rbrakk>
+ \<Longrightarrow> P \<turnstile> Object has_fields FDTs"
+
+lemma has_fields_is_class:
+ "P \<turnstile> C has_fields FDTs \<Longrightarrow> is_class P C"
+(*<*)by (auto simp add: is_class_def elim: Fields.induct)(*>*)
+
+lemma has_fields_fun:
+assumes 1: "P \<turnstile> C has_fields FDTs"
+shows "\<And>FDTs'. P \<turnstile> C has_fields FDTs' \<Longrightarrow> FDTs' = FDTs"
+ (*<*)
+using 1
+proof induct
+ case (has_fields_rec C D fs ms Dres Cres Cres')
+ have "class": "class P C = Some (D, fs, ms)"
+ and notObj: "C \<noteq> Object" and DFields: "P \<turnstile> D has_fields Dres"
+ and IH: "\<And>Dres'. P \<turnstile> D has_fields Dres' \<Longrightarrow> Dres' = Dres"
+ and Cres: "Cres = map (\<lambda>(F,b,T). ((F,C),b,T)) fs @ Dres"
+ and CFields': "P \<turnstile> C has_fields Cres'" by fact+
+ from CFields' notObj "class" obtain Dres'
+ where DFields': "P \<turnstile> D has_fields Dres'"
+ and Cres': "Cres' = map (\<lambda>(F,b,T). ((F,C),b,T)) fs @ Dres'"
+ by(auto elim: Fields.cases)
+ from Cres Cres' IH[OF DFields'] show "Cres' = Cres" by simp
+next
+ case has_fields_Object thus ?case by(auto elim: Fields.cases)
+qed
+(*>*)
+
+lemma all_fields_in_has_fields:
+assumes sub: "P \<turnstile> C has_fields FDTs"
+shows "\<lbrakk> P \<turnstile> C \<preceq>\<^sup>* D; class P D = Some(D',fs,ms); (F,b,T) \<in> set fs \<rbrakk>
+ \<Longrightarrow> ((F,D),b,T) \<in> set FDTs"
+(*<*)
+using sub proof(induct)
+ case (has_fields_rec C D' fs ms FDTs FDTs')
+ then have C_D: "P \<turnstile> C \<preceq>\<^sup>* D" by simp
+ then show ?case proof(rule converse_rtranclE)
+ assume "C = D"
+ then show ?case using has_fields_rec by force
+ next
+ fix y assume sub1: "P \<turnstile> C \<prec>\<^sup>1 y" and sub2: "P \<turnstile> y \<preceq>\<^sup>* D"
+ then show ?case using has_fields_rec subcls1D[OF sub1] by simp
+ qed
+next
+ case (has_fields_Object D fs ms FDTs)
+ then show ?case by force
+qed
+(*>*)
+
+lemma has_fields_decl_above:
+assumes fields: "P \<turnstile> C has_fields FDTs"
+shows "((F,D),b,T) \<in> set FDTs \<Longrightarrow> P \<turnstile> C \<preceq>\<^sup>* D"
+(*<*)
+using fields proof(induct)
+ case (has_fields_rec C D' fs ms FDTs FDTs')
+ then have "((F, D), b, T) \<in> (\<lambda>x. case x of (F, x) \<Rightarrow> ((F, C), x)) ` set fs \<or>
+ ((F, D), b, T) \<in> set FDTs" by clarsimp
+ then show ?case proof(rule disjE)
+ assume "((F, D), b, T) \<in> (\<lambda>x. case x of (F, x) \<Rightarrow> ((F, C), x)) ` set fs"
+ then show ?case using has_fields_rec by clarsimp
+ next
+ assume "((F, D), b, T) \<in> set FDTs"
+ then show ?case using has_fields_rec
+ by(blast dest:subcls1I converse_rtrancl_into_rtrancl)
+ qed
+next
+ case (has_fields_Object D fs ms FDTs)
+ then show ?case by fastforce
+qed
+(*>*)
+
+
+lemma subcls_notin_has_fields:
+assumes fields: "P \<turnstile> C has_fields FDTs"
+shows "((F,D),b,T) \<in> set FDTs \<Longrightarrow> (D,C) \<notin> (subcls1 P)\<^sup>+"
+(*<*)
+using fields proof(induct)
+ case (has_fields_rec C D' fs ms FDTs FDTs')
+ then have "((F, D), b, T) \<in> (\<lambda>x. case x of (F, x) \<Rightarrow> ((F, C), x)) ` set fs
+ \<or> ((F, D), b, T) \<in> set FDTs" by clarsimp
+ then show ?case proof(rule disjE)
+ assume "((F, D), b, T) \<in> (\<lambda>x. case x of (F, x) \<Rightarrow> ((F, C), x)) ` set fs"
+ then have CD[simp]: "C = D" and fs: "(F, b, T) \<in> set fs" by clarsimp+
+ then have "(D, D) \<in> (subcls1 P)\<^sup>+ \<Longrightarrow> False" proof -
+ assume DD: "(D, D) \<in> (subcls1 P)\<^sup>+"
+ obtain z where z1: "P \<turnstile> D \<prec>\<^sup>1 z" and z_s: "P \<turnstile> z \<preceq>\<^sup>* D"
+ using tranclD[OF DD] by clarsimp
+ have [simp]: "z = D'" using subcls1D[OF z1] has_fields_rec.hyps(1) by clarsimp
+ then have "((F, D), b, T) \<in> set FDTs"
+ using z_s all_fields_in_has_fields[OF has_fields_rec.hyps(3) _ has_fields_rec.hyps(1) fs]
+ by simp
+ then have "(D, z) \<notin> (subcls1 P)\<^sup>+" using has_fields_rec.hyps(4) by simp
+ then show False using z1 by auto
+ qed
+ then show ?case by clarsimp
+ next
+ assume "((F, D), b, T) \<in> set FDTs"
+ then show ?case using has_fields_rec by(blast dest:subcls1I trancl_into_trancl)
+ qed
+next
+ case (has_fields_Object D fs ms FDTs)
+ then show ?case by(fastforce dest: tranclD)
+qed
+(*>*)
+
+lemma subcls_notin_has_fields2:
+assumes fields: "P \<turnstile> C has_fields FDTs"
+shows "\<lbrakk> C \<noteq> Object; P \<turnstile> C \<prec>\<^sup>1 D \<rbrakk> \<Longrightarrow> (D,C) \<notin> (subcls1 P)\<^sup>*"
+using fields proof(induct arbitrary: D)
+ case has_fields_rec
+ have "\<forall>C C' P. (C, C') \<notin> subcls1 P \<or> C \<noteq> Object \<and> (\<exists>fs ms. class P C = \<lfloor>(C', fs, ms)\<rfloor>)"
+ using subcls1D by blast
+ then have "(D, D) \<notin> (subcls1 P)\<^sup>+"
+ by (metis (no_types) Pair_inject has_fields_rec.hyps(1) has_fields_rec.hyps(4)
+ has_fields_rec.prems(2) option.inject tranclD)
+ then show ?case
+ by (meson has_fields_rec.prems(2) rtrancl_into_trancl1)
+qed(fastforce dest: tranclD)
+
+lemma has_fields_mono_lem:
+assumes sub: "P \<turnstile> D \<preceq>\<^sup>* C"
+shows "P \<turnstile> C has_fields FDTs
+ \<Longrightarrow> \<exists>pre. P \<turnstile> D has_fields pre@FDTs \<and> dom(map_of pre) \<inter> dom(map_of FDTs) = {}"
+(*<*)
+using sub proof(induct rule:converse_rtrancl_induct)
+ case base
+ then show ?case by(rule_tac x = "[]" in exI) simp
+next
+ case (step D' D)
+ then obtain pre where D_flds: "P \<turnstile> D has_fields pre @ FDTs" and
+ dom: "dom (map_of pre) \<inter> dom (map_of FDTs) = {}" by clarsimp
+ have "(D',C) \<in> (subcls1 P)^+" by (rule rtrancl_into_trancl2[OF step.hyps(1,2)])
+ obtain fs ms where D'_cls: "class P D' = \<lfloor>(D, fs, ms)\<rfloor>" "D' \<noteq> Object"
+ using subcls1D[OF step.hyps(1)] by clarsimp+
+ have "P \<turnstile> D' has_fields map (\<lambda>(F, T). ((F, D'), T)) fs @ pre @ FDTs"
+ using has_fields_rec[OF D'_cls D_flds] by simp
+ also have "dom (map_of (map (\<lambda>(F, T). ((F, D'), T)) fs @ pre))
+ \<inter> dom (map_of FDTs) = {}"
+ using dom subcls_notin_has_fields[OF D_flds, where D=D'] step.hyps(1)
+ by(auto simp:dom_map_of_conv_image_fst) fast
+ ultimately show ?case
+ by(rule_tac x = "map (\<lambda>(F,b,T). ((F,D'),b,T)) fs @ pre" in exI) simp
+qed
+(*>*)
+
+lemma has_fields_declaring_classes:
+shows "P \<turnstile> C has_fields FDTs
+ \<Longrightarrow> \<exists>pre FDTs'. FDTs = pre@FDTs'
+ \<and> (C \<noteq> Object \<longrightarrow> (\<exists>D fs ms. class P C = \<lfloor>(D,fs,ms)\<rfloor> \<and> P \<turnstile> D has_fields FDTs'))
+ \<and> set(map (\<lambda>t. snd(fst t)) pre) \<subseteq> {C}
+ \<and> set(map (\<lambda>t. snd(fst t)) FDTs') \<subseteq> {C'. C' \<noteq> C \<and> P \<turnstile> C \<preceq>\<^sup>* C'}"
+proof(induct rule:Fields.induct)
+ case (has_fields_rec C D fs ms FDTs FDTs')
+ have sup1: "P \<turnstile> C \<prec>\<^sup>1 D" using has_fields_rec.hyps(1,2) by (simp add: subcls1.subcls1I)
+ have "P \<turnstile> C has_fields FDTs'"
+ using Fields.has_fields_rec[OF has_fields_rec.hyps(1-3)] has_fields_rec by auto
+ then have nsup: "(D, C) \<notin> (subcls1 P)\<^sup>*" using subcls_notin_has_fields2 sup1 by auto
+ show ?case using has_fields_rec sup1 nsup
+ by(rule_tac x = "map (\<lambda>(F, y). ((F, C), y)) fs" in exI, clarsimp) auto
+next
+ case has_fields_Object then show ?case by fastforce
+qed
+
+lemma has_fields_mono_lem2:
+assumes hf: "P \<turnstile> C has_fields FDTs"
+ and cls: "class P C = Some(D,fs,ms)" and map_of: "map_of FDTs (F,C) = \<lfloor>(b,T)\<rfloor>"
+shows "\<exists>FDTs'. FDTs = (map (\<lambda>(F,b,T). ((F,C),b,T)) fs) @ FDTs' \<and> map_of FDTs' (F,C) = None"
+using assms
+proof(cases "C = Object")
+ case False
+ let ?pre = "map (\<lambda>(F,b,T). ((F,C),b,T)) fs"
+ have sub: "P \<turnstile> C \<preceq>\<^sup>* D" using cls False by (simp add: r_into_rtrancl subcls1.subcls1I)
+ obtain FDTs' where fdts': "P \<turnstile> D has_fields FDTs'" "FDTs = ?pre @ FDTs'"
+ using False assms(1,2) Fields.simps[of P C FDTs] by clarsimp
+ then have int: "dom (map_of ?pre) \<inter> dom (map_of FDTs') = {}"
+ using has_fields_mono_lem[OF sub, of FDTs'] has_fields_fun[OF hf] by fastforce
+ have "C \<notin> (\<lambda>t. snd (fst t)) ` set FDTs'"
+ using has_fields_declaring_classes[OF hf] cls False
+ has_fields_fun[OF fdts'(1)] fdts'(2)
+ by clarify auto
+ then have "map_of FDTs' (F,C) = None" by(rule map_of_set_pcs_notin)
+ then show ?thesis using fdts' int by simp
+qed(auto dest: has_fields_Object has_fields_fun)
+
+
+lemma has_fields_is_class_Object:
+ "P \<turnstile> D has_fields FDTs \<Longrightarrow> is_class P Object"
+ by(induct rule: Fields.induct; simp add: is_class_def)
+
+lemma Object_fields:
+ "\<lbrakk> P \<turnstile> Object has_fields FDTs; C \<noteq> Object \<rbrakk> \<Longrightarrow> map_of FDTs (F,C) = None"
+ by(drule Fields.cases, auto simp: map_of_reinsert_neq_None)
+
+
+definition has_field :: "'m prog \<Rightarrow> cname \<Rightarrow> vname \<Rightarrow> staticb \<Rightarrow> ty \<Rightarrow> cname \<Rightarrow> bool"
+ ("_ \<turnstile> _ has _,_:_ in _" [51,51,51,51,51,51] 50)
+where
+ "P \<turnstile> C has F,b:T in D \<equiv>
+ \<exists>FDTs. P \<turnstile> C has_fields FDTs \<and> map_of FDTs (F,D) = Some (b,T)"
+
+
+lemma has_field_mono:
+assumes has: " P \<turnstile> C has F,b:T in D" and sub: "P \<turnstile> C' \<preceq>\<^sup>* C"
+shows "P \<turnstile> C' has F,b:T in D"
+(*<*)
+proof -
+ obtain FDTs where FDTs:"P \<turnstile> C has_fields FDTs" and "map_of FDTs (F, D) = \<lfloor>(b, T)\<rfloor>"
+ using has by(clarsimp simp: has_field_def)
+ also obtain pre where "P \<turnstile> C' has_fields pre @ FDTs"
+ and "dom (map_of pre) \<inter> dom (map_of FDTs) = {}"
+ using has_fields_mono_lem[OF sub FDTs] by clarify
+ ultimately show ?thesis by(fastforce simp: has_field_def map_add_def split:option.splits)
+qed
+(*>*)
+
+lemma has_field_fun:
+ "\<lbrakk>P \<turnstile> C has F,b:T in D; P \<turnstile> C has F,b':T' in D\<rbrakk> \<Longrightarrow> b = b' \<and> T' = T"
+(*<*)by(fastforce simp:has_field_def dest:has_fields_fun)(*>*)
+
+
+lemma has_field_idemp:
+assumes has: "P \<turnstile> C has F,b:T in D"
+shows "P \<turnstile> D has F,b:T in D"
+(*<*)
+proof -
+ obtain FDTs where C_flds: "P \<turnstile> C has_fields FDTs"
+ and FDTs: "map_of FDTs (F, D) = \<lfloor>(b, T)\<rfloor>" (is "?FDTs")
+ using has by(clarsimp simp: has_field_def)
+ have map: "\<And>C' fs. map_of (map (\<lambda>(F, y). ((F, C'), y)) fs) (F, D) = \<lfloor>(b, T)\<rfloor> \<Longrightarrow> D = C'"
+ by(frule map_of_SomeD) clarsimp
+ have "?FDTs \<longrightarrow> P \<turnstile> D has F,b:T in D"
+ using C_flds proof induct
+ case NObj: (has_fields_rec C' D' fs ms FDTs FDTs')
+ then show ?case using map by (fastforce intro: has_fields_rec simp: has_field_def)
+ next
+ case Obj: (has_fields_Object D fs ms FDTs)
+ then show ?case using map by(fastforce intro: has_fields_Object simp: has_field_def)
+ qed
+ then show ?thesis using FDTs by(rule_tac mp)
+qed
+(*>*)
+
+lemma visible_fields_exist:
+assumes fields: "P \<turnstile> C has_fields FDTs" and
+ FDTs: "map_of FDTs (F,D) = Some (b, T)"
+shows "\<exists>D' fs ms. class P D = Some(D',fs,ms) \<and> map_of fs F = Some(b,T)"
+proof -
+ have "map_of FDTs (F,D) = Some (b, T) \<longrightarrow>
+ (\<exists>D' fs ms. class P D = Some(D',fs,ms) \<and> map_of fs F = Some(b,T))"
+ using fields proof induct
+ case (has_fields_rec C' D' fs ms FDTs')
+ with assms map_of_reinsert_SomeD map_of_reinsert_neq_None[where D=D and F=F and fs=fs]
+ show ?case proof(cases "C' = D") qed auto
+ next
+ case (has_fields_Object D' fs ms FDTs)
+ with assms map_of_reinsert_SomeD map_of_reinsert_neq_None[where D=D and F=F and fs=fs]
+ show ?case proof(cases "Object = D") qed auto
+ qed
+ then show ?thesis using FDTs by simp
+qed
+
+lemma map_of_remap_SomeD:
+ "map_of (map (\<lambda>((k,k'),x). (k,(k',x))) t) k = Some (k',x) \<Longrightarrow> map_of t (k, k') = Some x"
+(*<*)by (induct t) (auto simp:fun_upd_apply split: if_split_asm)(*>*)
+
+lemma map_of_remap_SomeD2:
+ "map_of (map (\<lambda>((k,k'),x,x'). (k,(k',x,x'))) t) k = Some (k',x,x') \<Longrightarrow> map_of t (k, k') = Some (x, x')"
+(*<*)by (induct t) (auto simp:fun_upd_apply split: if_split_asm)(*>*)
+
+lemma has_field_decl_above:
+ "P \<turnstile> C has F,b:T in D \<Longrightarrow> P \<turnstile> C \<preceq>\<^sup>* D"
+(*<*)
+by(auto simp: has_field_def
+ intro: has_fields_decl_above map_of_SomeD map_of_remap_SomeD2)
+(*>*)
+
+definition sees_field :: "'m prog \<Rightarrow> cname \<Rightarrow> vname \<Rightarrow> staticb \<Rightarrow> ty \<Rightarrow> cname \<Rightarrow> bool"
+ ("_ \<turnstile> _ sees _,_:_ in _" [51,51,51,51,51,51] 50)
+where
+ "P \<turnstile> C sees F,b:T in D \<equiv>
+ \<exists>FDTs. P \<turnstile> C has_fields FDTs \<and>
+ map_of (map (\<lambda>((F,D),b,T). (F,(D,b,T))) FDTs) F = Some(D,b,T)"
+
+lemma has_visible_field:
+ "P \<turnstile> C sees F,b:T in D \<Longrightarrow> P \<turnstile> C has F,b:T in D"
+(*<*)by(auto simp add:has_field_def sees_field_def map_of_remap_SomeD2)(*>*)
+
+lemma sees_field_fun:
+ "\<lbrakk>P \<turnstile> C sees F,b:T in D; P \<turnstile> C sees F,b':T' in D'\<rbrakk> \<Longrightarrow> b = b' \<and> T' = T \<and> D' = D"
+(*<*)by(fastforce simp:sees_field_def dest:has_fields_fun)(*>*)
+
+lemma sees_field_decl_above:
+ "P \<turnstile> C sees F,b:T in D \<Longrightarrow> P \<turnstile> C \<preceq>\<^sup>* D"
+(*<*)
+by(auto simp:sees_field_def
+ intro: has_fields_decl_above map_of_SomeD map_of_remap_SomeD2)
+(*>*)
+
+
+lemma sees_field_idemp:
+assumes sees: "P \<turnstile> C sees F,b:T in D"
+shows "P \<turnstile> D sees F,b:T in D"
+(*<*)
+proof -
+ obtain FDTs where C_flds: "P \<turnstile> C has_fields FDTs"
+ and FDTs: "map_of (map (\<lambda>((F, D), b, T). (F, D, b, T)) FDTs) F = \<lfloor>(D, b, T)\<rfloor>"
+ (is "?FDTs")
+ using sees by(clarsimp simp: sees_field_def)
+ have map: "\<And>C' fs. map_of (map ((\<lambda>((F, D), a). (F, D, a)) \<circ> (\<lambda>(F, y). ((F, C'), y))) fs) F
+ = \<lfloor>(D, b, T)\<rfloor>
+ \<Longrightarrow> D = C' \<and> (F, b, T) \<in> set fs"
+ by(frule map_of_SomeD) clarsimp
+\<comment>\<open> ?FDTs \<longrightarrow> P \<turnstile> D sees F,b:T in D \<close>
+ have "?FDTs \<longrightarrow> (\<exists>FDTs. P \<turnstile> D has_fields FDTs
+ \<and> map_of (map (\<lambda>((F, D), a). (F, D, a)) FDTs) F = \<lfloor>(D, b, T)\<rfloor>)"
+ using C_flds proof induct
+ case NObj: (has_fields_rec C' D' fs ms FDTs FDTs')
+ then show ?case using map by (fastforce intro: has_fields_rec)
+ next
+ case Obj: (has_fields_Object D fs ms FDTs)
+ then show ?case using map by(fastforce intro: has_fields_Object)
+ qed
+ then show ?thesis using FDTs
+ by (smt map_eq_conv old.prod.case prod_cases3 sees_field_def split_cong)
+qed
+(*>*)
+
+lemma has_field_sees_aux:
+assumes hf: "P \<turnstile> C has_fields FDTs" and map: "map_of FDTs (F, C) = \<lfloor>(b, T)\<rfloor>"
+shows "map_of (map (\<lambda>((F, D), b, T). (F, D, b, T)) FDTs) F = \<lfloor>(C, b, T)\<rfloor>"
+proof -
+ obtain D fs ms where fs: "class P C = Some(D,fs,ms)"
+ using visible_fields_exist[OF assms] by clarsimp
+ then obtain FDTs' where
+ "FDTs = map (\<lambda>(F, b, T). ((F, C), b, T)) fs @ FDTs' \<and> map_of FDTs' (F, C) = None"
+ using has_fields_mono_lem2[OF hf fs map] by clarsimp
+ then show ?thesis using map_of_Some_None_split[OF _ _ map] by auto
+qed
+
+lemma has_field_sees: "P \<turnstile> C has F,b:T in C \<Longrightarrow> P \<turnstile> C sees F,b:T in C"
+ by(auto simp:has_field_def sees_field_def has_field_sees_aux)
+
+lemma has_field_is_class:
+ "P \<turnstile> C has F,b:T in D \<Longrightarrow> is_class P C"
+(*<*)by (auto simp add: is_class_def has_field_def elim: Fields.induct)(*>*)
+
+lemma has_field_is_class':
+ "P \<turnstile> C has F,b:T in D \<Longrightarrow> is_class P D"
+(*<*)by(drule has_field_idemp, rule has_field_is_class, assumption)(*>*)
+
+subsection "Functional lookup"
+
+definition "method" :: "'m prog \<Rightarrow> cname \<Rightarrow> mname \<Rightarrow> cname \<times> staticb \<times> ty list \<times> ty \<times> 'm"
+where
+ "method P C M \<equiv> THE (D,b,Ts,T,m). P \<turnstile> C sees M,b:Ts \<rightarrow> T = m in D"
+
+definition field :: "'m prog \<Rightarrow> cname \<Rightarrow> vname \<Rightarrow> cname \<times> staticb \<times> ty"
+where
+ "field P C F \<equiv> THE (D,b,T). P \<turnstile> C sees F,b:T in D"
+
+definition fields :: "'m prog \<Rightarrow> cname \<Rightarrow> ((vname \<times> cname) \<times> staticb \<times> ty) list"
+where
+ "fields P C \<equiv> THE FDTs. P \<turnstile> C has_fields FDTs"
+
+lemma fields_def2 [simp]: "P \<turnstile> C has_fields FDTs \<Longrightarrow> fields P C = FDTs"
+(*<*)by (unfold fields_def) (auto dest: has_fields_fun)(*>*)
+
+lemma field_def2 [simp]: "P \<turnstile> C sees F,b:T in D \<Longrightarrow> field P C F = (D,b,T)"
+(*<*)by (unfold field_def) (auto dest: sees_field_fun)(*>*)
+
+lemma method_def2 [simp]: "P \<turnstile> C sees M,b: Ts\<rightarrow>T = m in D \<Longrightarrow> method P C M = (D,b,Ts,T,m)"
+(*<*)by (unfold method_def) (auto dest: sees_method_fun)(*>*)
+
+
+text \<open> The following are the fields for initializing an object (non-static fields)
+ and a class (just that class's static fields), respectively. \<close>
+
+definition ifields :: "'m prog \<Rightarrow> cname \<Rightarrow> ((vname \<times> cname) \<times> staticb \<times> ty) list"
+where
+ "ifields P C \<equiv> filter (\<lambda>((F,D),b,T). b = NonStatic) (fields P C)"
+
+definition isfields :: "'m prog \<Rightarrow> cname \<Rightarrow> ((vname \<times> cname) \<times> staticb \<times> ty) list"
+where
+ "isfields P C \<equiv> filter (\<lambda>((F,D),b,T). b = Static \<and> D = C) (fields P C)"
+
+lemma ifields_def2[simp]: "\<lbrakk> P \<turnstile> C has_fields FDTs \<rbrakk> \<Longrightarrow> ifields P C = filter (\<lambda>((F,D),b,T). b = NonStatic) FDTs"
+ by (simp add: ifields_def)
+
+lemma isfields_def2[simp]: "\<lbrakk> P \<turnstile> C has_fields FDTs \<rbrakk> \<Longrightarrow> isfields P C = filter (\<lambda>((F,D),b,T). b = Static \<and> D = C) FDTs"
+ by (simp add: isfields_def)
+
+lemma ifields_def3: "\<lbrakk> P \<turnstile> C sees F,b:T in D; b = NonStatic \<rbrakk> \<Longrightarrow> (((F,D),b,T) \<in> set (ifields P C))"
+(*<*) by (unfold ifields_def) (auto simp: sees_field_def map_of_SomeD map_of_remap_SomeD2) (*>*)
+
+lemma isfields_def3: "\<lbrakk> P \<turnstile> C sees F,b:T in D; b = Static; D = C \<rbrakk> \<Longrightarrow> (((F,D),b,T) \<in> set (isfields P C))"
+(*<*) by (unfold isfields_def) (auto simp: sees_field_def map_of_SomeD map_of_remap_SomeD2) (*>*)
+
+
+definition seeing_class :: "'m prog \<Rightarrow> cname \<Rightarrow> mname \<Rightarrow> cname option" where
+"seeing_class P C M =
+ (if \<exists>Ts T m D. P \<turnstile> C sees M,Static:Ts\<rightarrow>T = m in D
+ then Some (fst(method P C M))
+ else None)"
+
+lemma seeing_class_def2[simp]:
+ "P \<turnstile> C sees M,Static:Ts\<rightarrow>T = m in D \<Longrightarrow> seeing_class P C M = Some D"
+ by(fastforce simp: seeing_class_def)
+
+(*<*)
+end
+(*>*)
diff --git a/thys/JinjaDCI/Common/Value.thy b/thys/JinjaDCI/Common/Value.thy
--- a/thys/JinjaDCI/Common/Value.thy
+++ b/thys/JinjaDCI/Common/Value.thy
@@ -1,32 +1,32 @@
-(* Title: Jinja/Common/Value.thy
- Author: David von Oheimb, Tobias Nipkow
- Copyright 1999 Technische Universitaet Muenchen
-*)
-
-section \<open> Jinja Values \<close>
-
-theory Value imports TypeRel begin
-
-type_synonym addr = nat
-
-datatype val
- = Unit \<comment> \<open>dummy result value of void expressions\<close>
- | Null \<comment> \<open>null reference\<close>
- | Bool bool \<comment> \<open>Boolean value\<close>
- | Intg int \<comment> \<open>integer value\<close>
- | Addr addr \<comment> \<open>addresses of objects in the heap\<close>
-
-primrec the_Intg :: "val \<Rightarrow> int" where
- "the_Intg (Intg i) = i"
-
-primrec the_Addr :: "val \<Rightarrow> addr" where
- "the_Addr (Addr a) = a"
-
-primrec default_val :: "ty \<Rightarrow> val" \<comment> \<open>default value for all types\<close> where
- "default_val Void = Unit"
-| "default_val Boolean = Bool False"
-| "default_val Integer = Intg 0"
-| "default_val NT = Null"
-| "default_val (Class C) = Null"
-
-end
+(* Title: Jinja/Common/Value.thy
+ Author: David von Oheimb, Tobias Nipkow
+ Copyright 1999 Technische Universitaet Muenchen
+*)
+
+section \<open> Jinja Values \<close>
+
+theory Value imports TypeRel begin
+
+type_synonym addr = nat
+
+datatype val
+ = Unit \<comment> \<open>dummy result value of void expressions\<close>
+ | Null \<comment> \<open>null reference\<close>
+ | Bool bool \<comment> \<open>Boolean value\<close>
+ | Intg int \<comment> \<open>integer value\<close>
+ | Addr addr \<comment> \<open>addresses of objects in the heap\<close>
+
+primrec the_Intg :: "val \<Rightarrow> int" where
+ "the_Intg (Intg i) = i"
+
+primrec the_Addr :: "val \<Rightarrow> addr" where
+ "the_Addr (Addr a) = a"
+
+primrec default_val :: "ty \<Rightarrow> val" \<comment> \<open>default value for all types\<close> where
+ "default_val Void = Unit"
+| "default_val Boolean = Bool False"
+| "default_val Integer = Intg 0"
+| "default_val NT = Null"
+| "default_val (Class C) = Null"
+
+end
diff --git a/thys/JinjaDCI/Common/WellForm.thy b/thys/JinjaDCI/Common/WellForm.thy
--- a/thys/JinjaDCI/Common/WellForm.thy
+++ b/thys/JinjaDCI/Common/WellForm.thy
@@ -1,503 +1,503 @@
-(* Title: JinjaDCI/Common/WellForm.thy
-
- Author: Tobias Nipkow, Susannah Mansky
- Copyright 2003 Technische Universitaet Muenchen, 2019-20 UIUC
-
- Based on the Jinja theory J/WellForm.thy by Tobias Nipkow
-*)
-
-section \<open> Generic Well-formedness of programs \<close>
-
-theory WellForm imports TypeRel SystemClasses begin
-
-text \<open>\noindent This theory defines global well-formedness conditions
-for programs but does not look inside method bodies. Hence it works
-for both Jinja and JVM programs. Well-typing of expressions is defined
-elsewhere (in theory @{text WellType}).
-
-Because Jinja does not have method overloading, its policy for method
-overriding is the classical one: \emph{covariant in the result type
-but contravariant in the argument types.} This means the result type
-of the overriding method becomes more specific, the argument types
-become more general.
-\<close>
-
-type_synonym 'm wf_mdecl_test = "'m prog \<Rightarrow> cname \<Rightarrow> 'm mdecl \<Rightarrow> bool"
-
-definition wf_fdecl :: "'m prog \<Rightarrow> fdecl \<Rightarrow> bool"
-where
- "wf_fdecl P \<equiv> \<lambda>(F,b,T). is_type P T"
-
-definition wf_mdecl :: "'m wf_mdecl_test \<Rightarrow> 'm wf_mdecl_test"
-where
- "wf_mdecl wf_md P C \<equiv> \<lambda>(M,b,Ts,T,m).
- (\<forall>T\<in>set Ts. is_type P T) \<and> is_type P T \<and> wf_md P C (M,b,Ts,T,m)"
-
-definition wf_clinit :: "'m mdecl list \<Rightarrow> bool" where
-"wf_clinit ms = (\<exists>m. (clinit,Static,[],Void,m)\<in>set ms)"
-
-definition wf_cdecl :: "'m wf_mdecl_test \<Rightarrow> 'm prog \<Rightarrow> 'm cdecl \<Rightarrow> bool"
-where
- "wf_cdecl wf_md P \<equiv> \<lambda>(C,(D,fs,ms)).
- (\<forall>f\<in>set fs. wf_fdecl P f) \<and> distinct_fst fs \<and>
- (\<forall>m\<in>set ms. wf_mdecl wf_md P C m) \<and> distinct_fst ms \<and>
- (C \<noteq> Object \<longrightarrow>
- is_class P D \<and> \<not> P \<turnstile> D \<preceq>\<^sup>* C \<and>
- (\<forall>(M,b,Ts,T,m)\<in>set ms.
- \<forall>D' b' Ts' T' m'. P \<turnstile> D sees M,b':Ts' \<rightarrow> T' = m' in D' \<longrightarrow>
- b = b' \<and> P \<turnstile> Ts' [\<le>] Ts \<and> P \<turnstile> T \<le> T')) \<and>
- wf_clinit ms"
-
-definition wf_syscls :: "'m prog \<Rightarrow> bool"
-where
- "wf_syscls P \<equiv> {Object} \<union> sys_xcpts \<subseteq> set(map fst P)"
-
-definition wf_prog :: "'m wf_mdecl_test \<Rightarrow> 'm prog \<Rightarrow> bool"
-where
- "wf_prog wf_md P \<equiv> wf_syscls P \<and> (\<forall>c \<in> set P. wf_cdecl wf_md P c) \<and> distinct_fst P"
-
-
-subsection\<open> Well-formedness lemmas \<close>
-
-lemma class_wf:
- "\<lbrakk>class P C = Some c; wf_prog wf_md P\<rbrakk> \<Longrightarrow> wf_cdecl wf_md P (C,c)"
-(*<*)by (unfold wf_prog_def class_def) (fast dest: map_of_SomeD)(*>*)
-
-
-lemma class_Object [simp]:
- "wf_prog wf_md P \<Longrightarrow> \<exists>C fs ms. class P Object = Some (C,fs,ms)"
-(*<*)by (unfold wf_prog_def wf_syscls_def class_def)
- (auto simp: map_of_SomeI)
-(*>*)
-
-
-lemma is_class_Object [simp]:
- "wf_prog wf_md P \<Longrightarrow> is_class P Object"
-(*<*)by (simp add: is_class_def)(*>*)
-
-lemma is_class_supclass:
-assumes wf: "wf_prog wf_md P" and sub: "P \<turnstile> C \<preceq>\<^sup>* D"
-shows "is_class P C \<Longrightarrow> is_class P D"
-(*<*)
-using sub proof(induct)
- case step then show ?case
- by(auto simp:wf_cdecl_def is_class_def dest!:class_wf[OF _ wf] subcls1D)
-qed simp
-(*>*)
-
-lemma is_class_xcpt:
- "\<lbrakk> C \<in> sys_xcpts; wf_prog wf_md P \<rbrakk> \<Longrightarrow> is_class P C"
-(*<*)
-by (fastforce intro!: map_of_SomeI
- simp add: wf_prog_def wf_syscls_def is_class_def class_def)
-(*>*)
-
-
-lemma subcls1_wfD:
-assumes sub1: "P \<turnstile> C \<prec>\<^sup>1 D" and wf: "wf_prog wf_md P"
-shows "D \<noteq> C \<and> (D,C) \<notin> (subcls1 P)\<^sup>+"
-(*<*)
-proof -
- obtain fs ms where "C \<noteq> Object" and cls: "class P C = \<lfloor>(D, fs, ms)\<rfloor>"
- using subcls1D[OF sub1] by clarify
- then show ?thesis using wf class_wf[OF cls wf] r_into_trancl[OF sub1]
- by(force simp add: wf_cdecl_def reflcl_trancl [THEN sym]
- simp del: reflcl_trancl)
-qed
-(*>*)
-
-
-lemma wf_cdecl_supD:
- "\<lbrakk>wf_cdecl wf_md P (C,D,r); C \<noteq> Object\<rbrakk> \<Longrightarrow> is_class P D"
-(*<*)by (auto simp: wf_cdecl_def)(*>*)
-
-
-lemma subcls_asym:
- "\<lbrakk> wf_prog wf_md P; (C,D) \<in> (subcls1 P)\<^sup>+ \<rbrakk> \<Longrightarrow> (D,C) \<notin> (subcls1 P)\<^sup>+"
-(*<*)by(erule tranclE; fast dest!: subcls1_wfD intro: trancl_trans)(*>*)
-
-
-lemma subcls_irrefl:
- "\<lbrakk> wf_prog wf_md P; (C,D) \<in> (subcls1 P)\<^sup>+ \<rbrakk> \<Longrightarrow> C \<noteq> D"
-(*<*)by (erule trancl_trans_induct) (auto dest: subcls1_wfD subcls_asym)(*>*)
-
-
-lemma acyclic_subcls1:
- "wf_prog wf_md P \<Longrightarrow> acyclic (subcls1 P)"
-(*<*)by (unfold acyclic_def) (fast dest: subcls_irrefl)(*>*)
-
-
-lemma wf_subcls1:
- "wf_prog wf_md P \<Longrightarrow> wf ((subcls1 P)\<inverse>)"
-(*<*)
-proof -
- assume wf: "wf_prog wf_md P"
- have "finite (subcls1 P)" by(rule finite_subcls1)
- then have fin': "finite ((subcls1 P)\<inverse>)" by(subst finite_converse)
-
- from wf have "acyclic (subcls1 P)" by(rule acyclic_subcls1)
- then have acyc': "acyclic ((subcls1 P)\<inverse>)" by (subst acyclic_converse)
-
- from fin' acyc' show ?thesis by (rule finite_acyclic_wf)
-qed
-(*>*)
-
-
-lemma single_valued_subcls1:
- "wf_prog wf_md G \<Longrightarrow> single_valued (subcls1 G)"
-(*<*)
-by(auto simp:wf_prog_def distinct_fst_def single_valued_def dest!:subcls1D)
-(*>*)
-
-
-lemma subcls_induct:
- "\<lbrakk> wf_prog wf_md P; \<And>C. \<forall>D. (C,D) \<in> (subcls1 P)\<^sup>+ \<longrightarrow> Q D \<Longrightarrow> Q C \<rbrakk> \<Longrightarrow> Q C"
-(*<*)
- (is "?A \<Longrightarrow> PROP ?P \<Longrightarrow> _")
-proof -
- assume p: "PROP ?P"
- assume ?A then have wf: "wf_prog wf_md P" by assumption
- have wf':"wf (((subcls1 P)\<^sup>+)\<inverse>)" using wf_trancl[OF wf_subcls1[OF wf]]
- by(simp only: trancl_converse)
- show ?thesis using wf_induct[where a = C and P = Q, OF wf' p] by simp
-qed
-(*>*)
-
-
-lemma subcls1_induct_aux:
-assumes "is_class P C" and wf: "wf_prog wf_md P" and QObj: "Q Object"
-shows
- "\<lbrakk> \<And>C D fs ms.
- \<lbrakk> C \<noteq> Object; is_class P C; class P C = Some (D,fs,ms) \<and>
- wf_cdecl wf_md P (C,D,fs,ms) \<and> P \<turnstile> C \<prec>\<^sup>1 D \<and> is_class P D \<and> Q D\<rbrakk> \<Longrightarrow> Q C \<rbrakk>
- \<Longrightarrow> Q C"
-(*<*)
- (is "PROP ?P \<Longrightarrow> _")
-proof -
- assume p: "PROP ?P"
- have "class P C \<noteq> None \<longrightarrow> Q C"
- proof(induct rule: subcls_induct[OF wf])
- case (1 C)
- have "class P C \<noteq> None \<Longrightarrow> Q C"
- proof(cases "C = Object")
- case True
- then show ?thesis using QObj by fast
- next
- case False
- assume nNone: "class P C \<noteq> None"
- then have is_cls: "is_class P C" by(simp add: is_class_def)
- obtain D fs ms where cls: "class P C = \<lfloor>(D, fs, ms)\<rfloor>" using nNone by safe
- also have wfC: "wf_cdecl wf_md P (C, D, fs, ms)" by(rule class_wf[OF cls wf])
- moreover have D: "is_class P D" by(rule wf_cdecl_supD[OF wfC False])
- moreover have "P \<turnstile> C \<prec>\<^sup>1 D" by(rule subcls1I[OF cls False])
- moreover have "class P D \<noteq> None" using D by(simp add: is_class_def)
- ultimately show ?thesis using 1 by (auto intro: p[OF False is_cls])
- qed
- then show "class P C \<noteq> None \<longrightarrow> Q C" by simp
- qed
- thus ?thesis using assms by(unfold is_class_def) simp
-qed
-(*>*)
-
-(* FIXME can't we prove this one directly?? *)
-lemma subcls1_induct [consumes 2, case_names Object Subcls]:
- "\<lbrakk> wf_prog wf_md P; is_class P C; Q Object;
- \<And>C D. \<lbrakk>C \<noteq> Object; P \<turnstile> C \<prec>\<^sup>1 D; is_class P D; Q D\<rbrakk> \<Longrightarrow> Q C \<rbrakk>
- \<Longrightarrow> Q C"
-(*<*)by (erule (2) subcls1_induct_aux) blast(*>*)
-
-
-lemma subcls_C_Object:
-assumes "class": "is_class P C" and wf: "wf_prog wf_md P"
-shows "P \<turnstile> C \<preceq>\<^sup>* Object"
-(*<*)
-using wf "class"
-proof(induct rule: subcls1_induct)
- case Subcls
- then show ?case by(simp add: converse_rtrancl_into_rtrancl)
-qed fast
-(*>*)
-
-
-lemma is_type_pTs:
-assumes "wf_prog wf_md P" and "(C,S,fs,ms) \<in> set P" and "(M,b,Ts,T,m) \<in> set ms"
-shows "set Ts \<subseteq> types P"
-(*<*)
-proof
- from assms have "wf_mdecl wf_md P C (M,b,Ts,T,m)"
- by (unfold wf_prog_def wf_cdecl_def) auto
- hence "\<forall>t \<in> set Ts. is_type P t" by (unfold wf_mdecl_def) auto
- moreover fix t assume "t \<in> set Ts"
- ultimately have "is_type P t" by blast
- thus "t \<in> types P" ..
-qed
-(*>*)
-
-lemma wf_supercls_distinct_app:
-assumes wf:"wf_prog wf_md P"
- and nObj: "C \<noteq> Object" and cls: "class P C = \<lfloor>(D, fs, ms)\<rfloor>"
- and super: "supercls_lst P (C#Cs)" and dist: "distinct (C#Cs)"
-shows "distinct (D#C#Cs)"
-proof -
- have "\<not> P \<turnstile> D \<preceq>\<^sup>* C" using subcls1_wfD[OF subcls1I[OF cls nObj] wf]
- by (simp add: rtrancl_eq_or_trancl)
- then show ?thesis using assms by auto
-qed
-
-
-subsection\<open> Well-formedness and method lookup \<close>
-
-lemma sees_wf_mdecl:
-assumes wf: "wf_prog wf_md P" and sees: "P \<turnstile> C sees M,b:Ts\<rightarrow>T = m in D"
-shows "wf_mdecl wf_md P D (M,b,Ts,T,m)"
-(*<*)
-using wf visible_method_exists[OF sees] proof(cases b)
-qed (fastforce simp:wf_cdecl_def dest!:class_wf dest:map_of_SomeD)+
-(*>*)
-
-lemma sees_method_mono [rule_format (no_asm)]:
-assumes sub: "P \<turnstile> C' \<preceq>\<^sup>* C" and wf: "wf_prog wf_md P"
-shows "\<forall>D b Ts T m. P \<turnstile> C sees M,b:Ts\<rightarrow>T = m in D \<longrightarrow>
- (\<exists>D' Ts' T' m'. P \<turnstile> C' sees M,b:Ts'\<rightarrow>T' = m' in D' \<and> P \<turnstile> Ts [\<le>] Ts' \<and> P \<turnstile> T' \<le> T)"
-(*<*)
- (is "\<forall>D b Ts T m. ?P C D b Ts T m \<longrightarrow> ?Q C' D b Ts T m")
-proof(rule disjE[OF rtranclD[OF sub]])
- assume "C' = C"
- then show ?thesis using assms by fastforce
-next
- assume "C' \<noteq> C \<and> (C', C) \<in> (subcls1 P)\<^sup>+"
- then have neq: "C' \<noteq> C" and subcls1: "(C', C) \<in> (subcls1 P)\<^sup>+" by simp+
- show ?thesis proof(induct rule: trancl_trans_induct[OF subcls1])
- case (2 x y z)
- then have zy: "\<And>D b Ts T m. ?P z D b Ts T m \<Longrightarrow> ?Q y D b Ts T m" by blast
- have "\<And>D b Ts T m. ?P z D b Ts T m \<Longrightarrow> ?Q x D b Ts T m"
- proof -
- fix D b Ts T m assume P: "?P z D b Ts T m"
- then show "?Q x D b Ts T m" using zy[OF P] 2(2)
- by(fast elim: widen_trans widens_trans)
- qed
- then show ?case by blast
- next
- case (1 x y)
- have "\<And>D b Ts T m. ?P y D b Ts T m \<Longrightarrow> ?Q x D b Ts T m"
- proof -
- fix D b Ts T m assume P: "?P y D b Ts T m"
- then obtain Mm where sees: "P \<turnstile> y sees_methods Mm" and
- M: "Mm M = \<lfloor>((b, Ts, T, m), D)\<rfloor>"
- by(clarsimp simp:Method_def)
- obtain fs ms where nObj: "x \<noteq> Object" and
- cls: "class P x = \<lfloor>(y, fs, ms)\<rfloor>"
- using subcls1D[OF 1] by clarsimp
- have x_meth: "P \<turnstile> x sees_methods Mm ++ (map_option (\<lambda>m. (m, x)) \<circ> map_of ms)"
- using sees_methods_rec[OF cls nObj sees] by simp
- show "?Q x D b Ts T m" proof(cases "map_of ms M")
- case None
- then have "\<exists>m'. P \<turnstile> x sees M, b : Ts\<rightarrow>T = m' in D" using M x_meth
- by(fastforce simp add:Method_def map_add_def split:option.split)
- then show ?thesis by auto
- next
- case (Some a)
- then obtain b' Ts' T' m' where a: "a = (b',Ts',T',m')" by(cases a)
- then have "(\<exists>m' Mm. P \<turnstile> y sees_methods Mm \<and> Mm M = \<lfloor>((b, Ts, T, m'), D)\<rfloor>)
- \<longrightarrow> b' = b \<and> P \<turnstile> Ts [\<le>] Ts' \<and> P \<turnstile> T' \<le> T"
- using nObj class_wf[OF cls wf] map_of_SomeD[OF Some]
- by(clarsimp simp: wf_cdecl_def Method_def) fast
- then show ?thesis using Some a sees M x_meth
- by(fastforce simp:Method_def map_add_def split:option.split)
- qed
- qed
- then show ?case by simp
- qed
-qed
-(*>*)
-
-
-lemma sees_method_mono2:
- "\<lbrakk> P \<turnstile> C' \<preceq>\<^sup>* C; wf_prog wf_md P;
- P \<turnstile> C sees M,b:Ts\<rightarrow>T = m in D; P \<turnstile> C' sees M,b':Ts'\<rightarrow>T' = m' in D' \<rbrakk>
- \<Longrightarrow> b = b' \<and> P \<turnstile> Ts [\<le>] Ts' \<and> P \<turnstile> T' \<le> T"
-(*<*)by(blast dest:sees_method_mono sees_method_fun)(*>*)
-
-lemma mdecls_visible:
-assumes wf: "wf_prog wf_md P" and "class": "is_class P C"
-shows "\<And>D fs ms. class P C = Some(D,fs,ms)
- \<Longrightarrow> \<exists>Mm. P \<turnstile> C sees_methods Mm \<and> (\<forall>(M,b,Ts,T,m) \<in> set ms. Mm M = Some((b,Ts,T,m),C))"
-(*<*)
-using wf "class"
-proof (induct rule:subcls1_induct)
- case Object
- with wf have dfst:"distinct_fst ms"
- by (unfold class_def wf_prog_def wf_cdecl_def) (fastforce dest:map_of_SomeD)
- with dfst have "distinct_fst ms"
- by(blast dest: distinct_fst_appendD)
- with Object show ?case by(fastforce intro!: sees_methods_Object map_of_SomeI)
-next
- case Subcls
- with wf have dfst:"distinct_fst ms"
- by (unfold class_def wf_prog_def wf_cdecl_def) (fastforce dest:map_of_SomeD)
- with dfst have "distinct_fst ms"
- by(blast dest: distinct_fst_appendD)
- with Subcls show ?case
- by(fastforce elim:sees_methods_rec dest:subcls1D map_of_SomeI
- simp:is_class_def)
-qed
-(*>*)
-
-lemma mdecl_visible:
-assumes wf: "wf_prog wf_md P" and C: "(C,S,fs,ms) \<in> set P" and m: "(M,b,Ts,T,m) \<in> set ms"
-shows "P \<turnstile> C sees M,b:Ts\<rightarrow>T = m in C"
-(*<*)
-proof -
- from wf C have "class": "class P C = Some (S,fs,ms)"
- by (auto simp add: wf_prog_def class_def is_class_def intro: map_of_SomeI)
- from "class" have "is_class P C" by(auto simp:is_class_def)
- with assms "class" show ?thesis
- by(bestsimp simp:Method_def dest:mdecls_visible)
-qed
-(*>*)
-
-
-lemma Call_lemma:
-assumes sees: "P \<turnstile> C sees M,b:Ts\<rightarrow>T = m in D" and sub: "P \<turnstile> C' \<preceq>\<^sup>* C" and wf: "wf_prog wf_md P"
-shows "\<exists>D' Ts' T' m'.
- P \<turnstile> C' sees M,b:Ts'\<rightarrow>T' = m' in D' \<and> P \<turnstile> Ts [\<le>] Ts' \<and> P \<turnstile> T' \<le> T \<and> P \<turnstile> C' \<preceq>\<^sup>* D'
- \<and> is_type P T' \<and> (\<forall>T\<in>set Ts'. is_type P T) \<and> wf_md P D' (M,b,Ts',T',m')"
-(*<*)
-using assms sees_method_mono[OF sub wf sees]
-by(fastforce intro:sees_method_decl_above dest:sees_wf_mdecl
- simp: wf_mdecl_def)
-(*>*)
-
-
-lemma wf_prog_lift:
- assumes wf: "wf_prog (\<lambda>P C bd. A P C bd) P"
- and rule:
- "\<And>wf_md C M b Ts C T m bd.
- wf_prog wf_md P \<Longrightarrow>
- P \<turnstile> C sees M,b:Ts\<rightarrow>T = m in C \<Longrightarrow>
- set Ts \<subseteq> types P \<Longrightarrow>
- bd = (M,b,Ts,T,m) \<Longrightarrow>
- A P C bd \<Longrightarrow>
- B P C bd"
- shows "wf_prog (\<lambda>P C bd. B P C bd) P"
-(*<*)
-proof -
- have "\<And>c. c\<in>set P \<Longrightarrow> wf_cdecl A P c \<Longrightarrow> wf_cdecl B P c"
- proof -
- fix c assume "c\<in>set P" and "wf_cdecl A P c"
- then show "wf_cdecl B P c"
- using rule[OF wf mdecl_visible[OF wf] is_type_pTs[OF wf]]
- by (auto simp: wf_cdecl_def wf_mdecl_def)
- qed
- then show ?thesis using wf by (clarsimp simp: wf_prog_def)
-qed
-(*>*)
-
-lemma wf_sees_clinit:
-assumes wf:"wf_prog wf_md P" and ex: "class P C = Some a"
-shows "\<exists>m. P \<turnstile> C sees clinit,Static:[] \<rightarrow> Void = m in C"
-proof -
- from ex obtain D fs ms where "a = (D,fs,ms)" by(cases a)
- then have sP: "(C, D, fs, ms) \<in> set P" using ex map_of_SomeD[of P C a] by(simp add: class_def)
- then have "wf_clinit ms" using assms by(unfold wf_prog_def wf_cdecl_def, auto)
- then obtain m where sm: "(clinit, Static, [], Void, m) \<in> set ms" by (meson wf_clinit_def)
- then have "P \<turnstile> C sees clinit,Static:[] \<rightarrow> Void = m in C"
- using mdecl_visible[OF wf sP sm] by simp
- then show ?thesis by(rule exI)
-qed
-(*>*)
-
-lemma wf_sees_clinit1:
-assumes wf:"wf_prog wf_md P" and ex: "class P C = Some a"
-and "P \<turnstile> C sees clinit,b:Ts \<rightarrow> T = m in D"
-shows "b = Static \<and> Ts = [] \<and> T = Void \<and> D = C"
-proof -
- obtain m' where sees: "P \<turnstile> C sees clinit,Static:[] \<rightarrow> Void = m' in C"
- using wf_sees_clinit[OF wf ex] by clarify
- then show ?thesis using sees wf by (meson assms(3) sees_method_fun)
-qed
-
-lemma wf_NonStatic_nclinit:
-assumes wf: "wf_prog wf_md P" and meth: "P \<turnstile> C sees M,NonStatic:Ts\<rightarrow>T=(mxs,mxl,ins,xt) in D"
-shows "M \<noteq> clinit"
-proof -
- from sees_method_is_class[OF meth] obtain a where cls: "class P C = Some a"
- by(clarsimp simp: is_class_def)
- with wf wf_sees_clinit[OF wf cls]
- obtain m where "P \<turnstile> C sees clinit,Static:[]\<rightarrow>Void=m in C" by clarsimp
- with meth show ?thesis by(auto dest: sees_method_fun)
-qed
-
-subsection\<open> Well-formedness and field lookup \<close>
-
-lemma wf_Fields_Ex:
-assumes wf: "wf_prog wf_md P" and "is_class P C"
-shows "\<exists>FDTs. P \<turnstile> C has_fields FDTs"
-(*<*)
-using assms proof(induct rule:subcls1_induct)
- case Object
- then show ?case using class_Object[OF wf]
- by(blast intro:has_fields_Object)
-next
- case Subcls
- then show ?case by(blast intro:has_fields_rec dest:subcls1D)
-qed
-(*>*)
-
-
-lemma has_fields_types:
- "\<lbrakk> P \<turnstile> C has_fields FDTs; (FD,b,T) \<in> set FDTs; wf_prog wf_md P \<rbrakk> \<Longrightarrow> is_type P T"
-(*<*)
-proof(induct rule:Fields.induct)
-qed(fastforce dest!: class_wf simp: wf_cdecl_def wf_fdecl_def)+
-(*>*)
-
-lemma sees_field_is_type:
- "\<lbrakk> P \<turnstile> C sees F,b:T in D; wf_prog wf_md P \<rbrakk> \<Longrightarrow> is_type P T"
-(*<*)
- by (meson has_field_def has_fields_types has_visible_field map_of_SomeD)
-(*>*)
-
-
-lemma wf_syscls:
- "set SystemClasses \<subseteq> set P \<Longrightarrow> wf_syscls P"
-(*<*)
-by (force simp: image_def SystemClasses_def wf_syscls_def sys_xcpts_def
- ObjectC_def NullPointerC_def ClassCastC_def OutOfMemoryC_def
- NoClassDefFoundC_def
- IncompatibleClassChangeC_def NoSuchFieldC_def NoSuchMethodC_def)
-(*>*)
-
-
-subsection\<open> Well-formedness and subclassing \<close>
-
-lemma wf_subcls_nCls:
-assumes wf: "wf_prog wf_md P" and ns: "\<not> is_class P C"
-shows "\<lbrakk> P \<turnstile> D \<preceq>\<^sup>* D'; D \<noteq> C \<rbrakk> \<Longrightarrow> D' \<noteq> C"
-proof(induct rule: rtrancl.induct)
- case (rtrancl_into_rtrancl a b c)
- with ns show ?case by(clarsimp dest!: subcls1D wf_cdecl_supD[OF class_wf[OF _ wf]])
-qed(simp)
-
-lemma wf_subcls_nCls':
-assumes wf: "wf_prog wf_md P" and ns: "\<not>is_class P C\<^sub>0"
-shows "\<And>cd D'. cd \<in> set P \<Longrightarrow> \<not>P \<turnstile> fst cd \<preceq>\<^sup>* C\<^sub>0"
-proof -
- fix cd D' assume cd: "cd \<in> set P"
- then have cls: "is_class P (fst cd)" using class_exists_equiv is_class_def by blast
- with wf_subcls_nCls[OF wf ns] ns show "\<not>P \<turnstile> fst cd \<preceq>\<^sup>* C\<^sub>0" by(cases "fst cd = D'", auto)
-qed
-
-lemma wf_nclass_nsub:
- "\<lbrakk> wf_prog wf_md P; is_class P C; \<not>is_class P C' \<rbrakk> \<Longrightarrow> \<not>P \<turnstile> C \<preceq>\<^sup>* C'"
- by(rule notI, auto dest: wf_subcls_nCls[where C=C' and D=C])
-
-lemma wf_sys_xcpt_nsub_Start:
-assumes wf: "wf_prog wf_md P" and ns: "\<not>is_class P Start" and sx: "C \<in> sys_xcpts"
-shows "\<not>P \<turnstile> C \<preceq>\<^sup>* Start"
-proof -
- have Cns: "C \<noteq> Start" using Start_nsys_xcpts sx by clarsimp
- show ?thesis using wf_subcls_nCls[OF wf ns _ Cns] by auto
-qed
-
-end
+(* Title: JinjaDCI/Common/WellForm.thy
+
+ Author: Tobias Nipkow, Susannah Mansky
+ Copyright 2003 Technische Universitaet Muenchen, 2019-20 UIUC
+
+ Based on the Jinja theory J/WellForm.thy by Tobias Nipkow
+*)
+
+section \<open> Generic Well-formedness of programs \<close>
+
+theory WellForm imports TypeRel SystemClasses begin
+
+text \<open>\noindent This theory defines global well-formedness conditions
+for programs but does not look inside method bodies. Hence it works
+for both Jinja and JVM programs. Well-typing of expressions is defined
+elsewhere (in theory @{text WellType}).
+
+Because Jinja does not have method overloading, its policy for method
+overriding is the classical one: \emph{covariant in the result type
+but contravariant in the argument types.} This means the result type
+of the overriding method becomes more specific, the argument types
+become more general.
+\<close>
+
+type_synonym 'm wf_mdecl_test = "'m prog \<Rightarrow> cname \<Rightarrow> 'm mdecl \<Rightarrow> bool"
+
+definition wf_fdecl :: "'m prog \<Rightarrow> fdecl \<Rightarrow> bool"
+where
+ "wf_fdecl P \<equiv> \<lambda>(F,b,T). is_type P T"
+
+definition wf_mdecl :: "'m wf_mdecl_test \<Rightarrow> 'm wf_mdecl_test"
+where
+ "wf_mdecl wf_md P C \<equiv> \<lambda>(M,b,Ts,T,m).
+ (\<forall>T\<in>set Ts. is_type P T) \<and> is_type P T \<and> wf_md P C (M,b,Ts,T,m)"
+
+definition wf_clinit :: "'m mdecl list \<Rightarrow> bool" where
+"wf_clinit ms = (\<exists>m. (clinit,Static,[],Void,m)\<in>set ms)"
+
+definition wf_cdecl :: "'m wf_mdecl_test \<Rightarrow> 'm prog \<Rightarrow> 'm cdecl \<Rightarrow> bool"
+where
+ "wf_cdecl wf_md P \<equiv> \<lambda>(C,(D,fs,ms)).
+ (\<forall>f\<in>set fs. wf_fdecl P f) \<and> distinct_fst fs \<and>
+ (\<forall>m\<in>set ms. wf_mdecl wf_md P C m) \<and> distinct_fst ms \<and>
+ (C \<noteq> Object \<longrightarrow>
+ is_class P D \<and> \<not> P \<turnstile> D \<preceq>\<^sup>* C \<and>
+ (\<forall>(M,b,Ts,T,m)\<in>set ms.
+ \<forall>D' b' Ts' T' m'. P \<turnstile> D sees M,b':Ts' \<rightarrow> T' = m' in D' \<longrightarrow>
+ b = b' \<and> P \<turnstile> Ts' [\<le>] Ts \<and> P \<turnstile> T \<le> T')) \<and>
+ wf_clinit ms"
+
+definition wf_syscls :: "'m prog \<Rightarrow> bool"
+where
+ "wf_syscls P \<equiv> {Object} \<union> sys_xcpts \<subseteq> set(map fst P)"
+
+definition wf_prog :: "'m wf_mdecl_test \<Rightarrow> 'm prog \<Rightarrow> bool"
+where
+ "wf_prog wf_md P \<equiv> wf_syscls P \<and> (\<forall>c \<in> set P. wf_cdecl wf_md P c) \<and> distinct_fst P"
+
+
+subsection\<open> Well-formedness lemmas \<close>
+
+lemma class_wf:
+ "\<lbrakk>class P C = Some c; wf_prog wf_md P\<rbrakk> \<Longrightarrow> wf_cdecl wf_md P (C,c)"
+(*<*)by (unfold wf_prog_def class_def) (fast dest: map_of_SomeD)(*>*)
+
+
+lemma class_Object [simp]:
+ "wf_prog wf_md P \<Longrightarrow> \<exists>C fs ms. class P Object = Some (C,fs,ms)"
+(*<*)by (unfold wf_prog_def wf_syscls_def class_def)
+ (auto simp: map_of_SomeI)
+(*>*)
+
+
+lemma is_class_Object [simp]:
+ "wf_prog wf_md P \<Longrightarrow> is_class P Object"
+(*<*)by (simp add: is_class_def)(*>*)
+
+lemma is_class_supclass:
+assumes wf: "wf_prog wf_md P" and sub: "P \<turnstile> C \<preceq>\<^sup>* D"
+shows "is_class P C \<Longrightarrow> is_class P D"
+(*<*)
+using sub proof(induct)
+ case step then show ?case
+ by(auto simp:wf_cdecl_def is_class_def dest!:class_wf[OF _ wf] subcls1D)
+qed simp
+(*>*)
+
+lemma is_class_xcpt:
+ "\<lbrakk> C \<in> sys_xcpts; wf_prog wf_md P \<rbrakk> \<Longrightarrow> is_class P C"
+(*<*)
+by (fastforce intro!: map_of_SomeI
+ simp add: wf_prog_def wf_syscls_def is_class_def class_def)
+(*>*)
+
+
+lemma subcls1_wfD:
+assumes sub1: "P \<turnstile> C \<prec>\<^sup>1 D" and wf: "wf_prog wf_md P"
+shows "D \<noteq> C \<and> (D,C) \<notin> (subcls1 P)\<^sup>+"
+(*<*)
+proof -
+ obtain fs ms where "C \<noteq> Object" and cls: "class P C = \<lfloor>(D, fs, ms)\<rfloor>"
+ using subcls1D[OF sub1] by clarify
+ then show ?thesis using wf class_wf[OF cls wf] r_into_trancl[OF sub1]
+ by(force simp add: wf_cdecl_def reflcl_trancl [THEN sym]
+ simp del: reflcl_trancl)
+qed
+(*>*)
+
+
+lemma wf_cdecl_supD:
+ "\<lbrakk>wf_cdecl wf_md P (C,D,r); C \<noteq> Object\<rbrakk> \<Longrightarrow> is_class P D"
+(*<*)by (auto simp: wf_cdecl_def)(*>*)
+
+
+lemma subcls_asym:
+ "\<lbrakk> wf_prog wf_md P; (C,D) \<in> (subcls1 P)\<^sup>+ \<rbrakk> \<Longrightarrow> (D,C) \<notin> (subcls1 P)\<^sup>+"
+(*<*)by(erule tranclE; fast dest!: subcls1_wfD intro: trancl_trans)(*>*)
+
+
+lemma subcls_irrefl:
+ "\<lbrakk> wf_prog wf_md P; (C,D) \<in> (subcls1 P)\<^sup>+ \<rbrakk> \<Longrightarrow> C \<noteq> D"
+(*<*)by (erule trancl_trans_induct) (auto dest: subcls1_wfD subcls_asym)(*>*)
+
+
+lemma acyclic_subcls1:
+ "wf_prog wf_md P \<Longrightarrow> acyclic (subcls1 P)"
+(*<*)by (unfold acyclic_def) (fast dest: subcls_irrefl)(*>*)
+
+
+lemma wf_subcls1:
+ "wf_prog wf_md P \<Longrightarrow> wf ((subcls1 P)\<inverse>)"
+(*<*)
+proof -
+ assume wf: "wf_prog wf_md P"
+ have "finite (subcls1 P)" by(rule finite_subcls1)
+ then have fin': "finite ((subcls1 P)\<inverse>)" by(subst finite_converse)
+
+ from wf have "acyclic (subcls1 P)" by(rule acyclic_subcls1)
+ then have acyc': "acyclic ((subcls1 P)\<inverse>)" by (subst acyclic_converse)
+
+ from fin' acyc' show ?thesis by (rule finite_acyclic_wf)
+qed
+(*>*)
+
+
+lemma single_valued_subcls1:
+ "wf_prog wf_md G \<Longrightarrow> single_valued (subcls1 G)"
+(*<*)
+by(auto simp:wf_prog_def distinct_fst_def single_valued_def dest!:subcls1D)
+(*>*)
+
+
+lemma subcls_induct:
+ "\<lbrakk> wf_prog wf_md P; \<And>C. \<forall>D. (C,D) \<in> (subcls1 P)\<^sup>+ \<longrightarrow> Q D \<Longrightarrow> Q C \<rbrakk> \<Longrightarrow> Q C"
+(*<*)
+ (is "?A \<Longrightarrow> PROP ?P \<Longrightarrow> _")
+proof -
+ assume p: "PROP ?P"
+ assume ?A then have wf: "wf_prog wf_md P" by assumption
+ have wf':"wf (((subcls1 P)\<^sup>+)\<inverse>)" using wf_trancl[OF wf_subcls1[OF wf]]
+ by(simp only: trancl_converse)
+ show ?thesis using wf_induct[where a = C and P = Q, OF wf' p] by simp
+qed
+(*>*)
+
+
+lemma subcls1_induct_aux:
+assumes "is_class P C" and wf: "wf_prog wf_md P" and QObj: "Q Object"
+shows
+ "\<lbrakk> \<And>C D fs ms.
+ \<lbrakk> C \<noteq> Object; is_class P C; class P C = Some (D,fs,ms) \<and>
+ wf_cdecl wf_md P (C,D,fs,ms) \<and> P \<turnstile> C \<prec>\<^sup>1 D \<and> is_class P D \<and> Q D\<rbrakk> \<Longrightarrow> Q C \<rbrakk>
+ \<Longrightarrow> Q C"
+(*<*)
+ (is "PROP ?P \<Longrightarrow> _")
+proof -
+ assume p: "PROP ?P"
+ have "class P C \<noteq> None \<longrightarrow> Q C"
+ proof(induct rule: subcls_induct[OF wf])
+ case (1 C)
+ have "class P C \<noteq> None \<Longrightarrow> Q C"
+ proof(cases "C = Object")
+ case True
+ then show ?thesis using QObj by fast
+ next
+ case False
+ assume nNone: "class P C \<noteq> None"
+ then have is_cls: "is_class P C" by(simp add: is_class_def)
+ obtain D fs ms where cls: "class P C = \<lfloor>(D, fs, ms)\<rfloor>" using nNone by safe
+ also have wfC: "wf_cdecl wf_md P (C, D, fs, ms)" by(rule class_wf[OF cls wf])
+ moreover have D: "is_class P D" by(rule wf_cdecl_supD[OF wfC False])
+ moreover have "P \<turnstile> C \<prec>\<^sup>1 D" by(rule subcls1I[OF cls False])
+ moreover have "class P D \<noteq> None" using D by(simp add: is_class_def)
+ ultimately show ?thesis using 1 by (auto intro: p[OF False is_cls])
+ qed
+ then show "class P C \<noteq> None \<longrightarrow> Q C" by simp
+ qed
+ thus ?thesis using assms by(unfold is_class_def) simp
+qed
+(*>*)
+
+(* FIXME can't we prove this one directly?? *)
+lemma subcls1_induct [consumes 2, case_names Object Subcls]:
+ "\<lbrakk> wf_prog wf_md P; is_class P C; Q Object;
+ \<And>C D. \<lbrakk>C \<noteq> Object; P \<turnstile> C \<prec>\<^sup>1 D; is_class P D; Q D\<rbrakk> \<Longrightarrow> Q C \<rbrakk>
+ \<Longrightarrow> Q C"
+(*<*)by (erule (2) subcls1_induct_aux) blast(*>*)
+
+
+lemma subcls_C_Object:
+assumes "class": "is_class P C" and wf: "wf_prog wf_md P"
+shows "P \<turnstile> C \<preceq>\<^sup>* Object"
+(*<*)
+using wf "class"
+proof(induct rule: subcls1_induct)
+ case Subcls
+ then show ?case by(simp add: converse_rtrancl_into_rtrancl)
+qed fast
+(*>*)
+
+
+lemma is_type_pTs:
+assumes "wf_prog wf_md P" and "(C,S,fs,ms) \<in> set P" and "(M,b,Ts,T,m) \<in> set ms"
+shows "set Ts \<subseteq> types P"
+(*<*)
+proof
+ from assms have "wf_mdecl wf_md P C (M,b,Ts,T,m)"
+ by (unfold wf_prog_def wf_cdecl_def) auto
+ hence "\<forall>t \<in> set Ts. is_type P t" by (unfold wf_mdecl_def) auto
+ moreover fix t assume "t \<in> set Ts"
+ ultimately have "is_type P t" by blast
+ thus "t \<in> types P" ..
+qed
+(*>*)
+
+lemma wf_supercls_distinct_app:
+assumes wf:"wf_prog wf_md P"
+ and nObj: "C \<noteq> Object" and cls: "class P C = \<lfloor>(D, fs, ms)\<rfloor>"
+ and super: "supercls_lst P (C#Cs)" and dist: "distinct (C#Cs)"
+shows "distinct (D#C#Cs)"
+proof -
+ have "\<not> P \<turnstile> D \<preceq>\<^sup>* C" using subcls1_wfD[OF subcls1I[OF cls nObj] wf]
+ by (simp add: rtrancl_eq_or_trancl)
+ then show ?thesis using assms by auto
+qed
+
+
+subsection\<open> Well-formedness and method lookup \<close>
+
+lemma sees_wf_mdecl:
+assumes wf: "wf_prog wf_md P" and sees: "P \<turnstile> C sees M,b:Ts\<rightarrow>T = m in D"
+shows "wf_mdecl wf_md P D (M,b,Ts,T,m)"
+(*<*)
+using wf visible_method_exists[OF sees] proof(cases b)
+qed (fastforce simp:wf_cdecl_def dest!:class_wf dest:map_of_SomeD)+
+(*>*)
+
+lemma sees_method_mono [rule_format (no_asm)]:
+assumes sub: "P \<turnstile> C' \<preceq>\<^sup>* C" and wf: "wf_prog wf_md P"
+shows "\<forall>D b Ts T m. P \<turnstile> C sees M,b:Ts\<rightarrow>T = m in D \<longrightarrow>
+ (\<exists>D' Ts' T' m'. P \<turnstile> C' sees M,b:Ts'\<rightarrow>T' = m' in D' \<and> P \<turnstile> Ts [\<le>] Ts' \<and> P \<turnstile> T' \<le> T)"
+(*<*)
+ (is "\<forall>D b Ts T m. ?P C D b Ts T m \<longrightarrow> ?Q C' D b Ts T m")
+proof(rule disjE[OF rtranclD[OF sub]])
+ assume "C' = C"
+ then show ?thesis using assms by fastforce
+next
+ assume "C' \<noteq> C \<and> (C', C) \<in> (subcls1 P)\<^sup>+"
+ then have neq: "C' \<noteq> C" and subcls1: "(C', C) \<in> (subcls1 P)\<^sup>+" by simp+
+ show ?thesis proof(induct rule: trancl_trans_induct[OF subcls1])
+ case (2 x y z)
+ then have zy: "\<And>D b Ts T m. ?P z D b Ts T m \<Longrightarrow> ?Q y D b Ts T m" by blast
+ have "\<And>D b Ts T m. ?P z D b Ts T m \<Longrightarrow> ?Q x D b Ts T m"
+ proof -
+ fix D b Ts T m assume P: "?P z D b Ts T m"
+ then show "?Q x D b Ts T m" using zy[OF P] 2(2)
+ by(fast elim: widen_trans widens_trans)
+ qed
+ then show ?case by blast
+ next
+ case (1 x y)
+ have "\<And>D b Ts T m. ?P y D b Ts T m \<Longrightarrow> ?Q x D b Ts T m"
+ proof -
+ fix D b Ts T m assume P: "?P y D b Ts T m"
+ then obtain Mm where sees: "P \<turnstile> y sees_methods Mm" and
+ M: "Mm M = \<lfloor>((b, Ts, T, m), D)\<rfloor>"
+ by(clarsimp simp:Method_def)
+ obtain fs ms where nObj: "x \<noteq> Object" and
+ cls: "class P x = \<lfloor>(y, fs, ms)\<rfloor>"
+ using subcls1D[OF 1] by clarsimp
+ have x_meth: "P \<turnstile> x sees_methods Mm ++ (map_option (\<lambda>m. (m, x)) \<circ> map_of ms)"
+ using sees_methods_rec[OF cls nObj sees] by simp
+ show "?Q x D b Ts T m" proof(cases "map_of ms M")
+ case None
+ then have "\<exists>m'. P \<turnstile> x sees M, b : Ts\<rightarrow>T = m' in D" using M x_meth
+ by(fastforce simp add:Method_def map_add_def split:option.split)
+ then show ?thesis by auto
+ next
+ case (Some a)
+ then obtain b' Ts' T' m' where a: "a = (b',Ts',T',m')" by(cases a)
+ then have "(\<exists>m' Mm. P \<turnstile> y sees_methods Mm \<and> Mm M = \<lfloor>((b, Ts, T, m'), D)\<rfloor>)
+ \<longrightarrow> b' = b \<and> P \<turnstile> Ts [\<le>] Ts' \<and> P \<turnstile> T' \<le> T"
+ using nObj class_wf[OF cls wf] map_of_SomeD[OF Some]
+ by(clarsimp simp: wf_cdecl_def Method_def) fast
+ then show ?thesis using Some a sees M x_meth
+ by(fastforce simp:Method_def map_add_def split:option.split)
+ qed
+ qed
+ then show ?case by simp
+ qed
+qed
+(*>*)
+
+
+lemma sees_method_mono2:
+ "\<lbrakk> P \<turnstile> C' \<preceq>\<^sup>* C; wf_prog wf_md P;
+ P \<turnstile> C sees M,b:Ts\<rightarrow>T = m in D; P \<turnstile> C' sees M,b':Ts'\<rightarrow>T' = m' in D' \<rbrakk>
+ \<Longrightarrow> b = b' \<and> P \<turnstile> Ts [\<le>] Ts' \<and> P \<turnstile> T' \<le> T"
+(*<*)by(blast dest:sees_method_mono sees_method_fun)(*>*)
+
+lemma mdecls_visible:
+assumes wf: "wf_prog wf_md P" and "class": "is_class P C"
+shows "\<And>D fs ms. class P C = Some(D,fs,ms)
+ \<Longrightarrow> \<exists>Mm. P \<turnstile> C sees_methods Mm \<and> (\<forall>(M,b,Ts,T,m) \<in> set ms. Mm M = Some((b,Ts,T,m),C))"
+(*<*)
+using wf "class"
+proof (induct rule:subcls1_induct)
+ case Object
+ with wf have dfst:"distinct_fst ms"
+ by (unfold class_def wf_prog_def wf_cdecl_def) (fastforce dest:map_of_SomeD)
+ with dfst have "distinct_fst ms"
+ by(blast dest: distinct_fst_appendD)
+ with Object show ?case by(fastforce intro!: sees_methods_Object map_of_SomeI)
+next
+ case Subcls
+ with wf have dfst:"distinct_fst ms"
+ by (unfold class_def wf_prog_def wf_cdecl_def) (fastforce dest:map_of_SomeD)
+ with dfst have "distinct_fst ms"
+ by(blast dest: distinct_fst_appendD)
+ with Subcls show ?case
+ by(fastforce elim:sees_methods_rec dest:subcls1D map_of_SomeI
+ simp:is_class_def)
+qed
+(*>*)
+
+lemma mdecl_visible:
+assumes wf: "wf_prog wf_md P" and C: "(C,S,fs,ms) \<in> set P" and m: "(M,b,Ts,T,m) \<in> set ms"
+shows "P \<turnstile> C sees M,b:Ts\<rightarrow>T = m in C"
+(*<*)
+proof -
+ from wf C have "class": "class P C = Some (S,fs,ms)"
+ by (auto simp add: wf_prog_def class_def is_class_def intro: map_of_SomeI)
+ from "class" have "is_class P C" by(auto simp:is_class_def)
+ with assms "class" show ?thesis
+ by(bestsimp simp:Method_def dest:mdecls_visible)
+qed
+(*>*)
+
+
+lemma Call_lemma:
+assumes sees: "P \<turnstile> C sees M,b:Ts\<rightarrow>T = m in D" and sub: "P \<turnstile> C' \<preceq>\<^sup>* C" and wf: "wf_prog wf_md P"
+shows "\<exists>D' Ts' T' m'.
+ P \<turnstile> C' sees M,b:Ts'\<rightarrow>T' = m' in D' \<and> P \<turnstile> Ts [\<le>] Ts' \<and> P \<turnstile> T' \<le> T \<and> P \<turnstile> C' \<preceq>\<^sup>* D'
+ \<and> is_type P T' \<and> (\<forall>T\<in>set Ts'. is_type P T) \<and> wf_md P D' (M,b,Ts',T',m')"
+(*<*)
+using assms sees_method_mono[OF sub wf sees]
+by(fastforce intro:sees_method_decl_above dest:sees_wf_mdecl
+ simp: wf_mdecl_def)
+(*>*)
+
+
+lemma wf_prog_lift:
+ assumes wf: "wf_prog (\<lambda>P C bd. A P C bd) P"
+ and rule:
+ "\<And>wf_md C M b Ts C T m bd.
+ wf_prog wf_md P \<Longrightarrow>
+ P \<turnstile> C sees M,b:Ts\<rightarrow>T = m in C \<Longrightarrow>
+ set Ts \<subseteq> types P \<Longrightarrow>
+ bd = (M,b,Ts,T,m) \<Longrightarrow>
+ A P C bd \<Longrightarrow>
+ B P C bd"
+ shows "wf_prog (\<lambda>P C bd. B P C bd) P"
+(*<*)
+proof -
+ have "\<And>c. c\<in>set P \<Longrightarrow> wf_cdecl A P c \<Longrightarrow> wf_cdecl B P c"
+ proof -
+ fix c assume "c\<in>set P" and "wf_cdecl A P c"
+ then show "wf_cdecl B P c"
+ using rule[OF wf mdecl_visible[OF wf] is_type_pTs[OF wf]]
+ by (auto simp: wf_cdecl_def wf_mdecl_def)
+ qed
+ then show ?thesis using wf by (clarsimp simp: wf_prog_def)
+qed
+(*>*)
+
+lemma wf_sees_clinit:
+assumes wf:"wf_prog wf_md P" and ex: "class P C = Some a"
+shows "\<exists>m. P \<turnstile> C sees clinit,Static:[] \<rightarrow> Void = m in C"
+proof -
+ from ex obtain D fs ms where "a = (D,fs,ms)" by(cases a)
+ then have sP: "(C, D, fs, ms) \<in> set P" using ex map_of_SomeD[of P C a] by(simp add: class_def)
+ then have "wf_clinit ms" using assms by(unfold wf_prog_def wf_cdecl_def, auto)
+ then obtain m where sm: "(clinit, Static, [], Void, m) \<in> set ms" by (meson wf_clinit_def)
+ then have "P \<turnstile> C sees clinit,Static:[] \<rightarrow> Void = m in C"
+ using mdecl_visible[OF wf sP sm] by simp
+ then show ?thesis by(rule exI)
+qed
+(*>*)
+
+lemma wf_sees_clinit1:
+assumes wf:"wf_prog wf_md P" and ex: "class P C = Some a"
+and "P \<turnstile> C sees clinit,b:Ts \<rightarrow> T = m in D"
+shows "b = Static \<and> Ts = [] \<and> T = Void \<and> D = C"
+proof -
+ obtain m' where sees: "P \<turnstile> C sees clinit,Static:[] \<rightarrow> Void = m' in C"
+ using wf_sees_clinit[OF wf ex] by clarify
+ then show ?thesis using sees wf by (meson assms(3) sees_method_fun)
+qed
+
+lemma wf_NonStatic_nclinit:
+assumes wf: "wf_prog wf_md P" and meth: "P \<turnstile> C sees M,NonStatic:Ts\<rightarrow>T=(mxs,mxl,ins,xt) in D"
+shows "M \<noteq> clinit"
+proof -
+ from sees_method_is_class[OF meth] obtain a where cls: "class P C = Some a"
+ by(clarsimp simp: is_class_def)
+ with wf wf_sees_clinit[OF wf cls]
+ obtain m where "P \<turnstile> C sees clinit,Static:[]\<rightarrow>Void=m in C" by clarsimp
+ with meth show ?thesis by(auto dest: sees_method_fun)
+qed
+
+subsection\<open> Well-formedness and field lookup \<close>
+
+lemma wf_Fields_Ex:
+assumes wf: "wf_prog wf_md P" and "is_class P C"
+shows "\<exists>FDTs. P \<turnstile> C has_fields FDTs"
+(*<*)
+using assms proof(induct rule:subcls1_induct)
+ case Object
+ then show ?case using class_Object[OF wf]
+ by(blast intro:has_fields_Object)
+next
+ case Subcls
+ then show ?case by(blast intro:has_fields_rec dest:subcls1D)
+qed
+(*>*)
+
+
+lemma has_fields_types:
+ "\<lbrakk> P \<turnstile> C has_fields FDTs; (FD,b,T) \<in> set FDTs; wf_prog wf_md P \<rbrakk> \<Longrightarrow> is_type P T"
+(*<*)
+proof(induct rule:Fields.induct)
+qed(fastforce dest!: class_wf simp: wf_cdecl_def wf_fdecl_def)+
+(*>*)
+
+lemma sees_field_is_type:
+ "\<lbrakk> P \<turnstile> C sees F,b:T in D; wf_prog wf_md P \<rbrakk> \<Longrightarrow> is_type P T"
+(*<*)
+ by (meson has_field_def has_fields_types has_visible_field map_of_SomeD)
+(*>*)
+
+
+lemma wf_syscls:
+ "set SystemClasses \<subseteq> set P \<Longrightarrow> wf_syscls P"
+(*<*)
+by (force simp: image_def SystemClasses_def wf_syscls_def sys_xcpts_def
+ ObjectC_def NullPointerC_def ClassCastC_def OutOfMemoryC_def
+ NoClassDefFoundC_def
+ IncompatibleClassChangeC_def NoSuchFieldC_def NoSuchMethodC_def)
+(*>*)
+
+
+subsection\<open> Well-formedness and subclassing \<close>
+
+lemma wf_subcls_nCls:
+assumes wf: "wf_prog wf_md P" and ns: "\<not> is_class P C"
+shows "\<lbrakk> P \<turnstile> D \<preceq>\<^sup>* D'; D \<noteq> C \<rbrakk> \<Longrightarrow> D' \<noteq> C"
+proof(induct rule: rtrancl.induct)
+ case (rtrancl_into_rtrancl a b c)
+ with ns show ?case by(clarsimp dest!: subcls1D wf_cdecl_supD[OF class_wf[OF _ wf]])
+qed(simp)
+
+lemma wf_subcls_nCls':
+assumes wf: "wf_prog wf_md P" and ns: "\<not>is_class P C\<^sub>0"
+shows "\<And>cd D'. cd \<in> set P \<Longrightarrow> \<not>P \<turnstile> fst cd \<preceq>\<^sup>* C\<^sub>0"
+proof -
+ fix cd D' assume cd: "cd \<in> set P"
+ then have cls: "is_class P (fst cd)" using class_exists_equiv is_class_def by blast
+ with wf_subcls_nCls[OF wf ns] ns show "\<not>P \<turnstile> fst cd \<preceq>\<^sup>* C\<^sub>0" by(cases "fst cd = D'", auto)
+qed
+
+lemma wf_nclass_nsub:
+ "\<lbrakk> wf_prog wf_md P; is_class P C; \<not>is_class P C' \<rbrakk> \<Longrightarrow> \<not>P \<turnstile> C \<preceq>\<^sup>* C'"
+ by(rule notI, auto dest: wf_subcls_nCls[where C=C' and D=C])
+
+lemma wf_sys_xcpt_nsub_Start:
+assumes wf: "wf_prog wf_md P" and ns: "\<not>is_class P Start" and sx: "C \<in> sys_xcpts"
+shows "\<not>P \<turnstile> C \<preceq>\<^sup>* Start"
+proof -
+ have Cns: "C \<noteq> Start" using Start_nsys_xcpts sx by clarsimp
+ show ?thesis using wf_subcls_nCls[OF wf ns _ Cns] by auto
+qed
+
+end
diff --git a/thys/JinjaDCI/Compiler/Compiler.thy b/thys/JinjaDCI/Compiler/Compiler.thy
--- a/thys/JinjaDCI/Compiler/Compiler.thy
+++ b/thys/JinjaDCI/Compiler/Compiler.thy
@@ -1,72 +1,72 @@
-(* Title: JinjaDCI/Compiler/Compiler.thy
-
- Author: Tobias Nipkow, Susannah Mansky
- Copyright TUM 2003, UIUC 2019-20
-
- Based on the Jinja theory Compiler/Compiler.thy by Tobias Nipkow
-*)
-
-section \<open> Combining Stages 1 and 2 \<close>
-
-theory Compiler
-imports Correctness1 Correctness2
-begin
-
-definition J2JVM :: "J_prog \<Rightarrow> jvm_prog"
-where
- "J2JVM \<equiv> compP\<^sub>2 \<circ> compP\<^sub>1"
-
-theorem comp_correct_NonStatic:
-assumes wf: "wf_J_prog P"
-and "method": "P \<turnstile> C sees M,NonStatic:Ts\<rightarrow>T = (pns,body) in C"
-and eval: "P \<turnstile> \<langle>body,(h,[this#pns [\<mapsto>] vs],sh)\<rangle> \<Rightarrow> \<langle>e',(h',l',sh')\<rangle>"
-and sizes: "size vs = size pns + 1" "size rest = max_vars body"
-shows "J2JVM P \<turnstile> (None,h,[([],vs@rest,C,M,0,No_ics)],sh) -jvm\<rightarrow> (exception e',h',[],sh')"
-(*<*)
-proof -
- let ?P\<^sub>1 = "compP\<^sub>1 P"
- have nclinit: "M \<noteq> clinit" using wf_sees_clinit1[OF wf] visible_method_exists[OF "method"]
- sees_method_idemp[OF "method"] by fastforce
- have wf\<^sub>1: "wf_J\<^sub>1_prog ?P\<^sub>1" by(rule compP\<^sub>1_pres_wf[OF wf])
- have fv: "fv body \<subseteq> set (this#pns)"
- using wf_prog_wwf_prog[OF wf] "method" by(auto dest!:sees_wf_mdecl simp:wf_mdecl_def)
- have init: "[this#pns [\<mapsto>] vs] \<subseteq>\<^sub>m [this#pns [\<mapsto>] vs@rest]"
- using sizes by simp
- have "?P\<^sub>1 \<turnstile> C sees M,NonStatic: Ts\<rightarrow>T = (compE\<^sub>1 (this#pns) body) in C"
- using sees_method_compP[OF "method", of "\<lambda>b (pns,e). compE\<^sub>1 (case b of NonStatic \<Rightarrow> this#pns | Static \<Rightarrow> pns) e"]
- by(simp)
- moreover obtain ls' where
- "?P\<^sub>1 \<turnstile>\<^sub>1 \<langle>compE\<^sub>1 (this#pns) body, (h, vs@rest, sh)\<rangle> \<Rightarrow> \<langle>fin\<^sub>1 e', (h',ls', sh')\<rangle>"
- using eval\<^sub>1_eval[OF wf_prog_wwf_prog[OF wf] eval fv init] sizes by auto
- ultimately show ?thesis using comp\<^sub>2_correct[OF wf\<^sub>1] eval_final[OF eval] nclinit
- by(fastforce simp add:J2JVM_def final_def)
-qed
-(*>*)
-
-theorem comp_correct_Static:
-assumes wf: "wf_J_prog P"
-and "method": "P \<turnstile> C sees M,Static:Ts\<rightarrow>T = (pns,body) in C"
-and eval: "P \<turnstile> \<langle>body,(h,[pns [\<mapsto>] vs],sh)\<rangle> \<Rightarrow> \<langle>e',(h',l',sh')\<rangle>"
-and sizes: "size vs = size pns" "size rest = max_vars body"
-and nclinit: "M \<noteq> clinit"
-shows "J2JVM P \<turnstile> (None,h,[([],vs@rest,C,M,0,No_ics)],sh) -jvm\<rightarrow> (exception e',h',[],sh')"
-(*<*)
-proof -
- let ?P\<^sub>1 = "compP\<^sub>1 P"
- have wf\<^sub>1: "wf_J\<^sub>1_prog ?P\<^sub>1" by(rule compP\<^sub>1_pres_wf[OF wf])
- have fv: "fv body \<subseteq> set pns"
- using wf_prog_wwf_prog[OF wf] "method" by(auto dest!:sees_wf_mdecl simp:wf_mdecl_def)
- have init: "[pns [\<mapsto>] vs] \<subseteq>\<^sub>m [pns [\<mapsto>] vs@rest]"
- using sizes by simp
- have "?P\<^sub>1 \<turnstile> C sees M,Static: Ts\<rightarrow>T = (compE\<^sub>1 pns body) in C"
- using sees_method_compP[OF "method", of "\<lambda>b (pns,e). compE\<^sub>1 (case b of NonStatic \<Rightarrow> this#pns | Static \<Rightarrow> pns) e"]
- by(simp)
- moreover obtain ls' where
- "?P\<^sub>1 \<turnstile>\<^sub>1 \<langle>compE\<^sub>1 pns body, (h, vs@rest, sh)\<rangle> \<Rightarrow> \<langle>fin\<^sub>1 e', (h',ls', sh')\<rangle>"
- using eval\<^sub>1_eval[OF wf_prog_wwf_prog[OF wf] eval fv init] sizes by auto
- ultimately show ?thesis using comp\<^sub>2_correct[OF wf\<^sub>1] eval_final[OF eval] nclinit
- by(fastforce simp add:J2JVM_def final_def)
-qed
-(*>*)
-
-end
+(* Title: JinjaDCI/Compiler/Compiler.thy
+
+ Author: Tobias Nipkow, Susannah Mansky
+ Copyright TUM 2003, UIUC 2019-20
+
+ Based on the Jinja theory Compiler/Compiler.thy by Tobias Nipkow
+*)
+
+section \<open> Combining Stages 1 and 2 \<close>
+
+theory Compiler
+imports Correctness1 Correctness2
+begin
+
+definition J2JVM :: "J_prog \<Rightarrow> jvm_prog"
+where
+ "J2JVM \<equiv> compP\<^sub>2 \<circ> compP\<^sub>1"
+
+theorem comp_correct_NonStatic:
+assumes wf: "wf_J_prog P"
+and "method": "P \<turnstile> C sees M,NonStatic:Ts\<rightarrow>T = (pns,body) in C"
+and eval: "P \<turnstile> \<langle>body,(h,[this#pns [\<mapsto>] vs],sh)\<rangle> \<Rightarrow> \<langle>e',(h',l',sh')\<rangle>"
+and sizes: "size vs = size pns + 1" "size rest = max_vars body"
+shows "J2JVM P \<turnstile> (None,h,[([],vs@rest,C,M,0,No_ics)],sh) -jvm\<rightarrow> (exception e',h',[],sh')"
+(*<*)
+proof -
+ let ?P\<^sub>1 = "compP\<^sub>1 P"
+ have nclinit: "M \<noteq> clinit" using wf_sees_clinit1[OF wf] visible_method_exists[OF "method"]
+ sees_method_idemp[OF "method"] by fastforce
+ have wf\<^sub>1: "wf_J\<^sub>1_prog ?P\<^sub>1" by(rule compP\<^sub>1_pres_wf[OF wf])
+ have fv: "fv body \<subseteq> set (this#pns)"
+ using wf_prog_wwf_prog[OF wf] "method" by(auto dest!:sees_wf_mdecl simp:wf_mdecl_def)
+ have init: "[this#pns [\<mapsto>] vs] \<subseteq>\<^sub>m [this#pns [\<mapsto>] vs@rest]"
+ using sizes by simp
+ have "?P\<^sub>1 \<turnstile> C sees M,NonStatic: Ts\<rightarrow>T = (compE\<^sub>1 (this#pns) body) in C"
+ using sees_method_compP[OF "method", of "\<lambda>b (pns,e). compE\<^sub>1 (case b of NonStatic \<Rightarrow> this#pns | Static \<Rightarrow> pns) e"]
+ by(simp)
+ moreover obtain ls' where
+ "?P\<^sub>1 \<turnstile>\<^sub>1 \<langle>compE\<^sub>1 (this#pns) body, (h, vs@rest, sh)\<rangle> \<Rightarrow> \<langle>fin\<^sub>1 e', (h',ls', sh')\<rangle>"
+ using eval\<^sub>1_eval[OF wf_prog_wwf_prog[OF wf] eval fv init] sizes by auto
+ ultimately show ?thesis using comp\<^sub>2_correct[OF wf\<^sub>1] eval_final[OF eval] nclinit
+ by(fastforce simp add:J2JVM_def final_def)
+qed
+(*>*)
+
+theorem comp_correct_Static:
+assumes wf: "wf_J_prog P"
+and "method": "P \<turnstile> C sees M,Static:Ts\<rightarrow>T = (pns,body) in C"
+and eval: "P \<turnstile> \<langle>body,(h,[pns [\<mapsto>] vs],sh)\<rangle> \<Rightarrow> \<langle>e',(h',l',sh')\<rangle>"
+and sizes: "size vs = size pns" "size rest = max_vars body"
+and nclinit: "M \<noteq> clinit"
+shows "J2JVM P \<turnstile> (None,h,[([],vs@rest,C,M,0,No_ics)],sh) -jvm\<rightarrow> (exception e',h',[],sh')"
+(*<*)
+proof -
+ let ?P\<^sub>1 = "compP\<^sub>1 P"
+ have wf\<^sub>1: "wf_J\<^sub>1_prog ?P\<^sub>1" by(rule compP\<^sub>1_pres_wf[OF wf])
+ have fv: "fv body \<subseteq> set pns"
+ using wf_prog_wwf_prog[OF wf] "method" by(auto dest!:sees_wf_mdecl simp:wf_mdecl_def)
+ have init: "[pns [\<mapsto>] vs] \<subseteq>\<^sub>m [pns [\<mapsto>] vs@rest]"
+ using sizes by simp
+ have "?P\<^sub>1 \<turnstile> C sees M,Static: Ts\<rightarrow>T = (compE\<^sub>1 pns body) in C"
+ using sees_method_compP[OF "method", of "\<lambda>b (pns,e). compE\<^sub>1 (case b of NonStatic \<Rightarrow> this#pns | Static \<Rightarrow> pns) e"]
+ by(simp)
+ moreover obtain ls' where
+ "?P\<^sub>1 \<turnstile>\<^sub>1 \<langle>compE\<^sub>1 pns body, (h, vs@rest, sh)\<rangle> \<Rightarrow> \<langle>fin\<^sub>1 e', (h',ls', sh')\<rangle>"
+ using eval\<^sub>1_eval[OF wf_prog_wwf_prog[OF wf] eval fv init] sizes by auto
+ ultimately show ?thesis using comp\<^sub>2_correct[OF wf\<^sub>1] eval_final[OF eval] nclinit
+ by(fastforce simp add:J2JVM_def final_def)
+qed
+(*>*)
+
+end
diff --git a/thys/JinjaDCI/Compiler/Compiler1.thy b/thys/JinjaDCI/Compiler/Compiler1.thy
--- a/thys/JinjaDCI/Compiler/Compiler1.thy
+++ b/thys/JinjaDCI/Compiler/Compiler1.thy
@@ -1,72 +1,72 @@
-(* Title: JinjaDCI/Compiler/Compiler1.thy
- Author: Tobias Nipkow, Susannah Mansky
- Copyright TUM 2003, UIUC 2019-20
-
- Based on the Jinja theory Compiler/Compiler1.thy by Tobias Nipkow
-*)
-
-section \<open> Compilation Stage 1 \<close>
-
-theory Compiler1 imports PCompiler J1 Hidden begin
-
-text\<open> Replacing variable names by indices. \<close>
-
-primrec compE\<^sub>1 :: "vname list \<Rightarrow> expr \<Rightarrow> expr\<^sub>1"
- and compEs\<^sub>1 :: "vname list \<Rightarrow> expr list \<Rightarrow> expr\<^sub>1 list" where
- "compE\<^sub>1 Vs (new C) = new C"
-| "compE\<^sub>1 Vs (Cast C e) = Cast C (compE\<^sub>1 Vs e)"
-| "compE\<^sub>1 Vs (Val v) = Val v"
-| "compE\<^sub>1 Vs (e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2) = (compE\<^sub>1 Vs e\<^sub>1) \<guillemotleft>bop\<guillemotright> (compE\<^sub>1 Vs e\<^sub>2)"
-| "compE\<^sub>1 Vs (Var V) = Var(last_index Vs V)"
-| "compE\<^sub>1 Vs (V:=e) = (last_index Vs V):= (compE\<^sub>1 Vs e)"
-| "compE\<^sub>1 Vs (e\<bullet>F{D}) = (compE\<^sub>1 Vs e)\<bullet>F{D}"
-| "compE\<^sub>1 Vs (C\<bullet>\<^sub>sF{D}) = C\<bullet>\<^sub>sF{D}"
-| "compE\<^sub>1 Vs (e\<^sub>1\<bullet>F{D}:=e\<^sub>2) = (compE\<^sub>1 Vs e\<^sub>1)\<bullet>F{D} := (compE\<^sub>1 Vs e\<^sub>2)"
-| "compE\<^sub>1 Vs (C\<bullet>\<^sub>sF{D}:=e\<^sub>2) = C\<bullet>\<^sub>sF{D} := (compE\<^sub>1 Vs e\<^sub>2)"
-| "compE\<^sub>1 Vs (e\<bullet>M(es)) = (compE\<^sub>1 Vs e)\<bullet>M(compEs\<^sub>1 Vs es)"
-| "compE\<^sub>1 Vs (C\<bullet>\<^sub>sM(es)) = C\<bullet>\<^sub>sM(compEs\<^sub>1 Vs es)"
-| "compE\<^sub>1 Vs {V:T; e} = {(size Vs):T; compE\<^sub>1 (Vs@[V]) e}"
-| "compE\<^sub>1 Vs (e\<^sub>1;;e\<^sub>2) = (compE\<^sub>1 Vs e\<^sub>1);;(compE\<^sub>1 Vs e\<^sub>2)"
-| "compE\<^sub>1 Vs (if (e) e\<^sub>1 else e\<^sub>2) = if (compE\<^sub>1 Vs e) (compE\<^sub>1 Vs e\<^sub>1) else (compE\<^sub>1 Vs e\<^sub>2)"
-| "compE\<^sub>1 Vs (while (e) c) = while (compE\<^sub>1 Vs e) (compE\<^sub>1 Vs c)"
-| "compE\<^sub>1 Vs (throw e) = throw (compE\<^sub>1 Vs e)"
-| "compE\<^sub>1 Vs (try e\<^sub>1 catch(C V) e\<^sub>2) =
- try(compE\<^sub>1 Vs e\<^sub>1) catch(C (size Vs)) (compE\<^sub>1 (Vs@[V]) e\<^sub>2)"
-| "compE\<^sub>1 Vs (INIT C (Cs,b) \<leftarrow> e) = INIT C (Cs,b) \<leftarrow> (compE\<^sub>1 Vs e)"
-| "compE\<^sub>1 Vs (RI(C,e);Cs \<leftarrow> e') = RI(C,(compE\<^sub>1 Vs e));Cs \<leftarrow> (compE\<^sub>1 Vs e')"
-
-| "compEs\<^sub>1 Vs [] = []"
-| "compEs\<^sub>1 Vs (e#es) = compE\<^sub>1 Vs e # compEs\<^sub>1 Vs es"
-
-lemma [simp]: "compEs\<^sub>1 Vs es = map (compE\<^sub>1 Vs) es"
-(*<*)by(induct es type:list) simp_all(*>*)
-
-lemma [simp]: "\<And>Vs. sub_RI (compE\<^sub>1 Vs e) = sub_RI e"
- and [simp]: "\<And>Vs. sub_RIs (compEs\<^sub>1 Vs es) = sub_RIs es"
-proof(induct rule: sub_RI_sub_RIs_induct) qed(auto)
-
-primrec fin\<^sub>1:: "expr \<Rightarrow> expr\<^sub>1" where
- "fin\<^sub>1(Val v) = Val v"
-| "fin\<^sub>1(throw e) = throw(fin\<^sub>1 e)"
-
-lemma comp_final: "final e \<Longrightarrow> compE\<^sub>1 Vs e = fin\<^sub>1 e"
-(*<*)by(erule finalE, simp_all)(*>*)
-
-
-lemma [simp]:
- "\<And>Vs. max_vars (compE\<^sub>1 Vs e) = max_vars e"
-and "\<And>Vs. max_varss (compEs\<^sub>1 Vs es) = max_varss es"
-(*<*)by (induct e and es rule: max_vars.induct max_varss.induct) simp_all(*>*)
-
-
-text\<open> Compiling programs: \<close>
-
-definition compP\<^sub>1 :: "J_prog \<Rightarrow> J\<^sub>1_prog"
-where
- "compP\<^sub>1 \<equiv> compP (\<lambda>b (pns,body). compE\<^sub>1 (case b of NonStatic \<Rightarrow> this#pns | Static \<Rightarrow> pns) body)"
-
-(*<*)
-declare compP\<^sub>1_def[simp]
-(*>*)
-
-end
+(* Title: JinjaDCI/Compiler/Compiler1.thy
+ Author: Tobias Nipkow, Susannah Mansky
+ Copyright TUM 2003, UIUC 2019-20
+
+ Based on the Jinja theory Compiler/Compiler1.thy by Tobias Nipkow
+*)
+
+section \<open> Compilation Stage 1 \<close>
+
+theory Compiler1 imports PCompiler J1 Hidden begin
+
+text\<open> Replacing variable names by indices. \<close>
+
+primrec compE\<^sub>1 :: "vname list \<Rightarrow> expr \<Rightarrow> expr\<^sub>1"
+ and compEs\<^sub>1 :: "vname list \<Rightarrow> expr list \<Rightarrow> expr\<^sub>1 list" where
+ "compE\<^sub>1 Vs (new C) = new C"
+| "compE\<^sub>1 Vs (Cast C e) = Cast C (compE\<^sub>1 Vs e)"
+| "compE\<^sub>1 Vs (Val v) = Val v"
+| "compE\<^sub>1 Vs (e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2) = (compE\<^sub>1 Vs e\<^sub>1) \<guillemotleft>bop\<guillemotright> (compE\<^sub>1 Vs e\<^sub>2)"
+| "compE\<^sub>1 Vs (Var V) = Var(last_index Vs V)"
+| "compE\<^sub>1 Vs (V:=e) = (last_index Vs V):= (compE\<^sub>1 Vs e)"
+| "compE\<^sub>1 Vs (e\<bullet>F{D}) = (compE\<^sub>1 Vs e)\<bullet>F{D}"
+| "compE\<^sub>1 Vs (C\<bullet>\<^sub>sF{D}) = C\<bullet>\<^sub>sF{D}"
+| "compE\<^sub>1 Vs (e\<^sub>1\<bullet>F{D}:=e\<^sub>2) = (compE\<^sub>1 Vs e\<^sub>1)\<bullet>F{D} := (compE\<^sub>1 Vs e\<^sub>2)"
+| "compE\<^sub>1 Vs (C\<bullet>\<^sub>sF{D}:=e\<^sub>2) = C\<bullet>\<^sub>sF{D} := (compE\<^sub>1 Vs e\<^sub>2)"
+| "compE\<^sub>1 Vs (e\<bullet>M(es)) = (compE\<^sub>1 Vs e)\<bullet>M(compEs\<^sub>1 Vs es)"
+| "compE\<^sub>1 Vs (C\<bullet>\<^sub>sM(es)) = C\<bullet>\<^sub>sM(compEs\<^sub>1 Vs es)"
+| "compE\<^sub>1 Vs {V:T; e} = {(size Vs):T; compE\<^sub>1 (Vs@[V]) e}"
+| "compE\<^sub>1 Vs (e\<^sub>1;;e\<^sub>2) = (compE\<^sub>1 Vs e\<^sub>1);;(compE\<^sub>1 Vs e\<^sub>2)"
+| "compE\<^sub>1 Vs (if (e) e\<^sub>1 else e\<^sub>2) = if (compE\<^sub>1 Vs e) (compE\<^sub>1 Vs e\<^sub>1) else (compE\<^sub>1 Vs e\<^sub>2)"
+| "compE\<^sub>1 Vs (while (e) c) = while (compE\<^sub>1 Vs e) (compE\<^sub>1 Vs c)"
+| "compE\<^sub>1 Vs (throw e) = throw (compE\<^sub>1 Vs e)"
+| "compE\<^sub>1 Vs (try e\<^sub>1 catch(C V) e\<^sub>2) =
+ try(compE\<^sub>1 Vs e\<^sub>1) catch(C (size Vs)) (compE\<^sub>1 (Vs@[V]) e\<^sub>2)"
+| "compE\<^sub>1 Vs (INIT C (Cs,b) \<leftarrow> e) = INIT C (Cs,b) \<leftarrow> (compE\<^sub>1 Vs e)"
+| "compE\<^sub>1 Vs (RI(C,e);Cs \<leftarrow> e') = RI(C,(compE\<^sub>1 Vs e));Cs \<leftarrow> (compE\<^sub>1 Vs e')"
+
+| "compEs\<^sub>1 Vs [] = []"
+| "compEs\<^sub>1 Vs (e#es) = compE\<^sub>1 Vs e # compEs\<^sub>1 Vs es"
+
+lemma [simp]: "compEs\<^sub>1 Vs es = map (compE\<^sub>1 Vs) es"
+(*<*)by(induct es type:list) simp_all(*>*)
+
+lemma [simp]: "\<And>Vs. sub_RI (compE\<^sub>1 Vs e) = sub_RI e"
+ and [simp]: "\<And>Vs. sub_RIs (compEs\<^sub>1 Vs es) = sub_RIs es"
+proof(induct rule: sub_RI_sub_RIs_induct) qed(auto)
+
+primrec fin\<^sub>1:: "expr \<Rightarrow> expr\<^sub>1" where
+ "fin\<^sub>1(Val v) = Val v"
+| "fin\<^sub>1(throw e) = throw(fin\<^sub>1 e)"
+
+lemma comp_final: "final e \<Longrightarrow> compE\<^sub>1 Vs e = fin\<^sub>1 e"
+(*<*)by(erule finalE, simp_all)(*>*)
+
+
+lemma [simp]:
+ "\<And>Vs. max_vars (compE\<^sub>1 Vs e) = max_vars e"
+and "\<And>Vs. max_varss (compEs\<^sub>1 Vs es) = max_varss es"
+(*<*)by (induct e and es rule: max_vars.induct max_varss.induct) simp_all(*>*)
+
+
+text\<open> Compiling programs: \<close>
+
+definition compP\<^sub>1 :: "J_prog \<Rightarrow> J\<^sub>1_prog"
+where
+ "compP\<^sub>1 \<equiv> compP (\<lambda>b (pns,body). compE\<^sub>1 (case b of NonStatic \<Rightarrow> this#pns | Static \<Rightarrow> pns) body)"
+
+(*<*)
+declare compP\<^sub>1_def[simp]
+(*>*)
+
+end
diff --git a/thys/JinjaDCI/Compiler/Compiler2.thy b/thys/JinjaDCI/Compiler/Compiler2.thy
--- a/thys/JinjaDCI/Compiler/Compiler2.thy
+++ b/thys/JinjaDCI/Compiler/Compiler2.thy
@@ -1,156 +1,156 @@
-(* Title: JinjaDCI/Compiler/Compiler2.thy
- Author: Tobias Nipkow, Susannah Mansky
- Copyright TUM 2003, UIUC 2019-20
-
- Based on the Jinja theory Compiler/Compiler2.thy by Tobias Nipkow
-*)
-
-section \<open> Compilation Stage 2 \<close>
-
-theory Compiler2
-imports PCompiler J1 "../JVM/JVMExec"
-begin
-
-lemma bop_expr_length_aux [simp]:
- "length (case bop of Eq \<Rightarrow> [CmpEq] | Add \<Rightarrow> [IAdd]) = Suc 0"
- by(cases bop, simp+)
-
-primrec compE\<^sub>2 :: "expr\<^sub>1 \<Rightarrow> instr list"
- and compEs\<^sub>2 :: "expr\<^sub>1 list \<Rightarrow> instr list" where
- "compE\<^sub>2 (new C) = [New C]"
-| "compE\<^sub>2 (Cast C e) = compE\<^sub>2 e @ [Checkcast C]"
-| "compE\<^sub>2 (Val v) = [Push v]"
-| "compE\<^sub>2 (e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2) = compE\<^sub>2 e\<^sub>1 @ compE\<^sub>2 e\<^sub>2 @
- (case bop of Eq \<Rightarrow> [CmpEq]
- | Add \<Rightarrow> [IAdd])"
-| "compE\<^sub>2 (Var i) = [Load i]"
-| "compE\<^sub>2 (i:=e) = compE\<^sub>2 e @ [Store i, Push Unit]"
-| "compE\<^sub>2 (e\<bullet>F{D}) = compE\<^sub>2 e @ [Getfield F D]"
-| "compE\<^sub>2 (C\<bullet>\<^sub>sF{D}) = [Getstatic C F D]"
-| "compE\<^sub>2 (e\<^sub>1\<bullet>F{D} := e\<^sub>2) =
- compE\<^sub>2 e\<^sub>1 @ compE\<^sub>2 e\<^sub>2 @ [Putfield F D, Push Unit]"
-| "compE\<^sub>2 (C\<bullet>\<^sub>sF{D} := e\<^sub>2) =
- compE\<^sub>2 e\<^sub>2 @ [Putstatic C F D, Push Unit]"
-| "compE\<^sub>2 (e\<bullet>M(es)) = compE\<^sub>2 e @ compEs\<^sub>2 es @ [Invoke M (size es)]"
-| "compE\<^sub>2 (C\<bullet>\<^sub>sM(es)) = compEs\<^sub>2 es @ [Invokestatic C M (size es)]"
-| "compE\<^sub>2 ({i:T; e}) = compE\<^sub>2 e"
-| "compE\<^sub>2 (e\<^sub>1;;e\<^sub>2) = compE\<^sub>2 e\<^sub>1 @ [Pop] @ compE\<^sub>2 e\<^sub>2"
-| "compE\<^sub>2 (if (e) e\<^sub>1 else e\<^sub>2) =
- (let cnd = compE\<^sub>2 e;
- thn = compE\<^sub>2 e\<^sub>1;
- els = compE\<^sub>2 e\<^sub>2;
- test = IfFalse (int(size thn + 2));
- thnex = Goto (int(size els + 1))
- in cnd @ [test] @ thn @ [thnex] @ els)"
-| "compE\<^sub>2 (while (e) c) =
- (let cnd = compE\<^sub>2 e;
- bdy = compE\<^sub>2 c;
- test = IfFalse (int(size bdy + 3));
- loop = Goto (-int(size bdy + size cnd + 2))
- in cnd @ [test] @ bdy @ [Pop] @ [loop] @ [Push Unit])"
-| "compE\<^sub>2 (throw e) = compE\<^sub>2 e @ [instr.Throw]"
-| "compE\<^sub>2 (try e\<^sub>1 catch(C i) e\<^sub>2) =
- (let catch = compE\<^sub>2 e\<^sub>2
- in compE\<^sub>2 e\<^sub>1 @ [Goto (int(size catch)+2), Store i] @ catch)"
-| "compE\<^sub>2 (INIT C (Cs,b) \<leftarrow> e) = []"
-| "compE\<^sub>2 (RI(C,e);Cs \<leftarrow> e') = []"
-
-| "compEs\<^sub>2 [] = []"
-| "compEs\<^sub>2 (e#es) = compE\<^sub>2 e @ compEs\<^sub>2 es"
-
-text\<open> Compilation of exception table. Is given start address of code
-to compute absolute addresses necessary in exception table. \<close>
-
-primrec compxE\<^sub>2 :: "expr\<^sub>1 \<Rightarrow> pc \<Rightarrow> nat \<Rightarrow> ex_table"
- and compxEs\<^sub>2 :: "expr\<^sub>1 list \<Rightarrow> pc \<Rightarrow> nat \<Rightarrow> ex_table" where
- "compxE\<^sub>2 (new C) pc d = []"
-| "compxE\<^sub>2 (Cast C e) pc d = compxE\<^sub>2 e pc d"
-| "compxE\<^sub>2 (Val v) pc d = []"
-| "compxE\<^sub>2 (e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2) pc d =
- compxE\<^sub>2 e\<^sub>1 pc d @ compxE\<^sub>2 e\<^sub>2 (pc + size(compE\<^sub>2 e\<^sub>1)) (d+1)"
-| "compxE\<^sub>2 (Var i) pc d = []"
-| "compxE\<^sub>2 (i:=e) pc d = compxE\<^sub>2 e pc d"
-| "compxE\<^sub>2 (e\<bullet>F{D}) pc d = compxE\<^sub>2 e pc d"
-| "compxE\<^sub>2 (C\<bullet>\<^sub>sF{D}) pc d = []"
-| "compxE\<^sub>2 (e\<^sub>1\<bullet>F{D} := e\<^sub>2) pc d =
- compxE\<^sub>2 e\<^sub>1 pc d @ compxE\<^sub>2 e\<^sub>2 (pc + size(compE\<^sub>2 e\<^sub>1)) (d+1)"
-| "compxE\<^sub>2 (C\<bullet>\<^sub>sF{D} := e\<^sub>2) pc d = compxE\<^sub>2 e\<^sub>2 pc d"
-| "compxE\<^sub>2 (e\<bullet>M(es)) pc d =
- compxE\<^sub>2 e pc d @ compxEs\<^sub>2 es (pc + size(compE\<^sub>2 e)) (d+1)"
-| "compxE\<^sub>2 (C\<bullet>\<^sub>sM(es)) pc d = compxEs\<^sub>2 es pc d"
-| "compxE\<^sub>2 ({i:T; e}) pc d = compxE\<^sub>2 e pc d"
-| "compxE\<^sub>2 (e\<^sub>1;;e\<^sub>2) pc d =
- compxE\<^sub>2 e\<^sub>1 pc d @ compxE\<^sub>2 e\<^sub>2 (pc+size(compE\<^sub>2 e\<^sub>1)+1) d"
-| "compxE\<^sub>2 (if (e) e\<^sub>1 else e\<^sub>2) pc d =
- (let pc\<^sub>1 = pc + size(compE\<^sub>2 e) + 1;
- pc\<^sub>2 = pc\<^sub>1 + size(compE\<^sub>2 e\<^sub>1) + 1
- in compxE\<^sub>2 e pc d @ compxE\<^sub>2 e\<^sub>1 pc\<^sub>1 d @ compxE\<^sub>2 e\<^sub>2 pc\<^sub>2 d)"
-| "compxE\<^sub>2 (while (b) e) pc d =
- compxE\<^sub>2 b pc d @ compxE\<^sub>2 e (pc+size(compE\<^sub>2 b)+1) d"
-| "compxE\<^sub>2 (throw e) pc d = compxE\<^sub>2 e pc d"
-| "compxE\<^sub>2 (try e\<^sub>1 catch(C i) e\<^sub>2) pc d =
- (let pc\<^sub>1 = pc + size(compE\<^sub>2 e\<^sub>1)
- in compxE\<^sub>2 e\<^sub>1 pc d @ compxE\<^sub>2 e\<^sub>2 (pc\<^sub>1+2) d @ [(pc,pc\<^sub>1,C,pc\<^sub>1+1,d)])"
-| "compxE\<^sub>2 (INIT C (Cs, b) \<leftarrow> e) pc d = []"
-| "compxE\<^sub>2 (RI(C, e);Cs \<leftarrow> e') pc d = []"
-
-| "compxEs\<^sub>2 [] pc d = []"
-| "compxEs\<^sub>2 (e#es) pc d = compxE\<^sub>2 e pc d @ compxEs\<^sub>2 es (pc+size(compE\<^sub>2 e)) (d+1)"
-
-primrec max_stack :: "expr\<^sub>1 \<Rightarrow> nat"
- and max_stacks :: "expr\<^sub>1 list \<Rightarrow> nat" where
- "max_stack (new C) = 1"
-| "max_stack (Cast C e) = max_stack e"
-| "max_stack (Val v) = 1"
-| "max_stack (e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2) = max (max_stack e\<^sub>1) (max_stack e\<^sub>2) + 1"
-| "max_stack (Var i) = 1"
-| "max_stack (i:=e) = max_stack e"
-| "max_stack (e\<bullet>F{D}) = max_stack e"
-| "max_stack (C\<bullet>\<^sub>sF{D}) = 1"
-| "max_stack (e\<^sub>1\<bullet>F{D} := e\<^sub>2) = max (max_stack e\<^sub>1) (max_stack e\<^sub>2) + 1"
-| "max_stack (C\<bullet>\<^sub>sF{D} := e\<^sub>2) = max_stack e\<^sub>2"
-| "max_stack (e\<bullet>M(es)) = max (max_stack e) (max_stacks es) + 1"
-| "max_stack (C\<bullet>\<^sub>sM(es)) = max_stacks es + 1"
-| "max_stack ({i:T; e}) = max_stack e"
-| "max_stack (e\<^sub>1;;e\<^sub>2) = max (max_stack e\<^sub>1) (max_stack e\<^sub>2)"
-| "max_stack (if (e) e\<^sub>1 else e\<^sub>2) =
- max (max_stack e) (max (max_stack e\<^sub>1) (max_stack e\<^sub>2))"
-| "max_stack (while (e) c) = max (max_stack e) (max_stack c)"
-| "max_stack (throw e) = max_stack e"
-| "max_stack (try e\<^sub>1 catch(C i) e\<^sub>2) = max (max_stack e\<^sub>1) (max_stack e\<^sub>2)"
-
-| "max_stacks [] = 0"
-| "max_stacks (e#es) = max (max_stack e) (1 + max_stacks es)"
-
-lemma max_stack1': "\<not>sub_RI e \<Longrightarrow> 1 \<le> max_stack e"
-(*<*)by(induct e) (simp_all add:max_def)(*>*)
-
-lemma compE\<^sub>2_not_Nil': "\<not>sub_RI e \<Longrightarrow> compE\<^sub>2 e \<noteq> []"
-(*<*)by(induct e) auto(*>*)
-
-lemma compE\<^sub>2_nRet: "\<And>i. i \<in> set (compE\<^sub>2 e\<^sub>1) \<Longrightarrow> i \<noteq> Return"
- and "\<And>i. i \<in> set (compEs\<^sub>2 es\<^sub>1) \<Longrightarrow> i \<noteq> Return"
- by(induct rule: compE\<^sub>2.induct compEs\<^sub>2.induct, auto simp: nth_append split: bop.splits)
-
-
-definition compMb\<^sub>2 :: "staticb \<Rightarrow> expr\<^sub>1 \<Rightarrow> jvm_method"
-where
- "compMb\<^sub>2 \<equiv> \<lambda>b body.
- let ins = compE\<^sub>2 body @ [Return];
- xt = compxE\<^sub>2 body 0 0
- in (max_stack body, max_vars body, ins, xt)"
-
-definition compP\<^sub>2 :: "J\<^sub>1_prog \<Rightarrow> jvm_prog"
-where
- "compP\<^sub>2 \<equiv> compP compMb\<^sub>2"
-
-(*<*)
-declare compP\<^sub>2_def [simp]
-(*>*)
-
-lemma compMb\<^sub>2 [simp]:
- "compMb\<^sub>2 b e = (max_stack e, max_vars e,
- compE\<^sub>2 e @ [Return], compxE\<^sub>2 e 0 0)"
-(*<*)by (simp add: compMb\<^sub>2_def)(*>*)
-
-end
+(* Title: JinjaDCI/Compiler/Compiler2.thy
+ Author: Tobias Nipkow, Susannah Mansky
+ Copyright TUM 2003, UIUC 2019-20
+
+ Based on the Jinja theory Compiler/Compiler2.thy by Tobias Nipkow
+*)
+
+section \<open> Compilation Stage 2 \<close>
+
+theory Compiler2
+imports PCompiler J1 "../JVM/JVMExec"
+begin
+
+lemma bop_expr_length_aux [simp]:
+ "length (case bop of Eq \<Rightarrow> [CmpEq] | Add \<Rightarrow> [IAdd]) = Suc 0"
+ by(cases bop, simp+)
+
+primrec compE\<^sub>2 :: "expr\<^sub>1 \<Rightarrow> instr list"
+ and compEs\<^sub>2 :: "expr\<^sub>1 list \<Rightarrow> instr list" where
+ "compE\<^sub>2 (new C) = [New C]"
+| "compE\<^sub>2 (Cast C e) = compE\<^sub>2 e @ [Checkcast C]"
+| "compE\<^sub>2 (Val v) = [Push v]"
+| "compE\<^sub>2 (e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2) = compE\<^sub>2 e\<^sub>1 @ compE\<^sub>2 e\<^sub>2 @
+ (case bop of Eq \<Rightarrow> [CmpEq]
+ | Add \<Rightarrow> [IAdd])"
+| "compE\<^sub>2 (Var i) = [Load i]"
+| "compE\<^sub>2 (i:=e) = compE\<^sub>2 e @ [Store i, Push Unit]"
+| "compE\<^sub>2 (e\<bullet>F{D}) = compE\<^sub>2 e @ [Getfield F D]"
+| "compE\<^sub>2 (C\<bullet>\<^sub>sF{D}) = [Getstatic C F D]"
+| "compE\<^sub>2 (e\<^sub>1\<bullet>F{D} := e\<^sub>2) =
+ compE\<^sub>2 e\<^sub>1 @ compE\<^sub>2 e\<^sub>2 @ [Putfield F D, Push Unit]"
+| "compE\<^sub>2 (C\<bullet>\<^sub>sF{D} := e\<^sub>2) =
+ compE\<^sub>2 e\<^sub>2 @ [Putstatic C F D, Push Unit]"
+| "compE\<^sub>2 (e\<bullet>M(es)) = compE\<^sub>2 e @ compEs\<^sub>2 es @ [Invoke M (size es)]"
+| "compE\<^sub>2 (C\<bullet>\<^sub>sM(es)) = compEs\<^sub>2 es @ [Invokestatic C M (size es)]"
+| "compE\<^sub>2 ({i:T; e}) = compE\<^sub>2 e"
+| "compE\<^sub>2 (e\<^sub>1;;e\<^sub>2) = compE\<^sub>2 e\<^sub>1 @ [Pop] @ compE\<^sub>2 e\<^sub>2"
+| "compE\<^sub>2 (if (e) e\<^sub>1 else e\<^sub>2) =
+ (let cnd = compE\<^sub>2 e;
+ thn = compE\<^sub>2 e\<^sub>1;
+ els = compE\<^sub>2 e\<^sub>2;
+ test = IfFalse (int(size thn + 2));
+ thnex = Goto (int(size els + 1))
+ in cnd @ [test] @ thn @ [thnex] @ els)"
+| "compE\<^sub>2 (while (e) c) =
+ (let cnd = compE\<^sub>2 e;
+ bdy = compE\<^sub>2 c;
+ test = IfFalse (int(size bdy + 3));
+ loop = Goto (-int(size bdy + size cnd + 2))
+ in cnd @ [test] @ bdy @ [Pop] @ [loop] @ [Push Unit])"
+| "compE\<^sub>2 (throw e) = compE\<^sub>2 e @ [instr.Throw]"
+| "compE\<^sub>2 (try e\<^sub>1 catch(C i) e\<^sub>2) =
+ (let catch = compE\<^sub>2 e\<^sub>2
+ in compE\<^sub>2 e\<^sub>1 @ [Goto (int(size catch)+2), Store i] @ catch)"
+| "compE\<^sub>2 (INIT C (Cs,b) \<leftarrow> e) = []"
+| "compE\<^sub>2 (RI(C,e);Cs \<leftarrow> e') = []"
+
+| "compEs\<^sub>2 [] = []"
+| "compEs\<^sub>2 (e#es) = compE\<^sub>2 e @ compEs\<^sub>2 es"
+
+text\<open> Compilation of exception table. Is given start address of code
+to compute absolute addresses necessary in exception table. \<close>
+
+primrec compxE\<^sub>2 :: "expr\<^sub>1 \<Rightarrow> pc \<Rightarrow> nat \<Rightarrow> ex_table"
+ and compxEs\<^sub>2 :: "expr\<^sub>1 list \<Rightarrow> pc \<Rightarrow> nat \<Rightarrow> ex_table" where
+ "compxE\<^sub>2 (new C) pc d = []"
+| "compxE\<^sub>2 (Cast C e) pc d = compxE\<^sub>2 e pc d"
+| "compxE\<^sub>2 (Val v) pc d = []"
+| "compxE\<^sub>2 (e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2) pc d =
+ compxE\<^sub>2 e\<^sub>1 pc d @ compxE\<^sub>2 e\<^sub>2 (pc + size(compE\<^sub>2 e\<^sub>1)) (d+1)"
+| "compxE\<^sub>2 (Var i) pc d = []"
+| "compxE\<^sub>2 (i:=e) pc d = compxE\<^sub>2 e pc d"
+| "compxE\<^sub>2 (e\<bullet>F{D}) pc d = compxE\<^sub>2 e pc d"
+| "compxE\<^sub>2 (C\<bullet>\<^sub>sF{D}) pc d = []"
+| "compxE\<^sub>2 (e\<^sub>1\<bullet>F{D} := e\<^sub>2) pc d =
+ compxE\<^sub>2 e\<^sub>1 pc d @ compxE\<^sub>2 e\<^sub>2 (pc + size(compE\<^sub>2 e\<^sub>1)) (d+1)"
+| "compxE\<^sub>2 (C\<bullet>\<^sub>sF{D} := e\<^sub>2) pc d = compxE\<^sub>2 e\<^sub>2 pc d"
+| "compxE\<^sub>2 (e\<bullet>M(es)) pc d =
+ compxE\<^sub>2 e pc d @ compxEs\<^sub>2 es (pc + size(compE\<^sub>2 e)) (d+1)"
+| "compxE\<^sub>2 (C\<bullet>\<^sub>sM(es)) pc d = compxEs\<^sub>2 es pc d"
+| "compxE\<^sub>2 ({i:T; e}) pc d = compxE\<^sub>2 e pc d"
+| "compxE\<^sub>2 (e\<^sub>1;;e\<^sub>2) pc d =
+ compxE\<^sub>2 e\<^sub>1 pc d @ compxE\<^sub>2 e\<^sub>2 (pc+size(compE\<^sub>2 e\<^sub>1)+1) d"
+| "compxE\<^sub>2 (if (e) e\<^sub>1 else e\<^sub>2) pc d =
+ (let pc\<^sub>1 = pc + size(compE\<^sub>2 e) + 1;
+ pc\<^sub>2 = pc\<^sub>1 + size(compE\<^sub>2 e\<^sub>1) + 1
+ in compxE\<^sub>2 e pc d @ compxE\<^sub>2 e\<^sub>1 pc\<^sub>1 d @ compxE\<^sub>2 e\<^sub>2 pc\<^sub>2 d)"
+| "compxE\<^sub>2 (while (b) e) pc d =
+ compxE\<^sub>2 b pc d @ compxE\<^sub>2 e (pc+size(compE\<^sub>2 b)+1) d"
+| "compxE\<^sub>2 (throw e) pc d = compxE\<^sub>2 e pc d"
+| "compxE\<^sub>2 (try e\<^sub>1 catch(C i) e\<^sub>2) pc d =
+ (let pc\<^sub>1 = pc + size(compE\<^sub>2 e\<^sub>1)
+ in compxE\<^sub>2 e\<^sub>1 pc d @ compxE\<^sub>2 e\<^sub>2 (pc\<^sub>1+2) d @ [(pc,pc\<^sub>1,C,pc\<^sub>1+1,d)])"
+| "compxE\<^sub>2 (INIT C (Cs, b) \<leftarrow> e) pc d = []"
+| "compxE\<^sub>2 (RI(C, e);Cs \<leftarrow> e') pc d = []"
+
+| "compxEs\<^sub>2 [] pc d = []"
+| "compxEs\<^sub>2 (e#es) pc d = compxE\<^sub>2 e pc d @ compxEs\<^sub>2 es (pc+size(compE\<^sub>2 e)) (d+1)"
+
+primrec max_stack :: "expr\<^sub>1 \<Rightarrow> nat"
+ and max_stacks :: "expr\<^sub>1 list \<Rightarrow> nat" where
+ "max_stack (new C) = 1"
+| "max_stack (Cast C e) = max_stack e"
+| "max_stack (Val v) = 1"
+| "max_stack (e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2) = max (max_stack e\<^sub>1) (max_stack e\<^sub>2) + 1"
+| "max_stack (Var i) = 1"
+| "max_stack (i:=e) = max_stack e"
+| "max_stack (e\<bullet>F{D}) = max_stack e"
+| "max_stack (C\<bullet>\<^sub>sF{D}) = 1"
+| "max_stack (e\<^sub>1\<bullet>F{D} := e\<^sub>2) = max (max_stack e\<^sub>1) (max_stack e\<^sub>2) + 1"
+| "max_stack (C\<bullet>\<^sub>sF{D} := e\<^sub>2) = max_stack e\<^sub>2"
+| "max_stack (e\<bullet>M(es)) = max (max_stack e) (max_stacks es) + 1"
+| "max_stack (C\<bullet>\<^sub>sM(es)) = max_stacks es + 1"
+| "max_stack ({i:T; e}) = max_stack e"
+| "max_stack (e\<^sub>1;;e\<^sub>2) = max (max_stack e\<^sub>1) (max_stack e\<^sub>2)"
+| "max_stack (if (e) e\<^sub>1 else e\<^sub>2) =
+ max (max_stack e) (max (max_stack e\<^sub>1) (max_stack e\<^sub>2))"
+| "max_stack (while (e) c) = max (max_stack e) (max_stack c)"
+| "max_stack (throw e) = max_stack e"
+| "max_stack (try e\<^sub>1 catch(C i) e\<^sub>2) = max (max_stack e\<^sub>1) (max_stack e\<^sub>2)"
+
+| "max_stacks [] = 0"
+| "max_stacks (e#es) = max (max_stack e) (1 + max_stacks es)"
+
+lemma max_stack1': "\<not>sub_RI e \<Longrightarrow> 1 \<le> max_stack e"
+(*<*)by(induct e) (simp_all add:max_def)(*>*)
+
+lemma compE\<^sub>2_not_Nil': "\<not>sub_RI e \<Longrightarrow> compE\<^sub>2 e \<noteq> []"
+(*<*)by(induct e) auto(*>*)
+
+lemma compE\<^sub>2_nRet: "\<And>i. i \<in> set (compE\<^sub>2 e\<^sub>1) \<Longrightarrow> i \<noteq> Return"
+ and "\<And>i. i \<in> set (compEs\<^sub>2 es\<^sub>1) \<Longrightarrow> i \<noteq> Return"
+ by(induct rule: compE\<^sub>2.induct compEs\<^sub>2.induct, auto simp: nth_append split: bop.splits)
+
+
+definition compMb\<^sub>2 :: "staticb \<Rightarrow> expr\<^sub>1 \<Rightarrow> jvm_method"
+where
+ "compMb\<^sub>2 \<equiv> \<lambda>b body.
+ let ins = compE\<^sub>2 body @ [Return];
+ xt = compxE\<^sub>2 body 0 0
+ in (max_stack body, max_vars body, ins, xt)"
+
+definition compP\<^sub>2 :: "J\<^sub>1_prog \<Rightarrow> jvm_prog"
+where
+ "compP\<^sub>2 \<equiv> compP compMb\<^sub>2"
+
+(*<*)
+declare compP\<^sub>2_def [simp]
+(*>*)
+
+lemma compMb\<^sub>2 [simp]:
+ "compMb\<^sub>2 b e = (max_stack e, max_vars e,
+ compE\<^sub>2 e @ [Return], compxE\<^sub>2 e 0 0)"
+(*<*)by (simp add: compMb\<^sub>2_def)(*>*)
+
+end
diff --git a/thys/JinjaDCI/Compiler/Correctness1.thy b/thys/JinjaDCI/Compiler/Correctness1.thy
--- a/thys/JinjaDCI/Compiler/Correctness1.thy
+++ b/thys/JinjaDCI/Compiler/Correctness1.thy
@@ -1,1064 +1,1064 @@
-(* Title: JinjaDCI/Compiler/Correctness1.thy
- Author: Tobias Nipkow, Susannah Mansky
- Copyright TUM 2003, UIUC 2019-20
-
- Based on the Jinja theory Compiler/Correctness1.thy by Tobias Nipkow
-*)
-
-section \<open> Correctness of Stage 1 \<close>
-
-theory Correctness1
-imports J1WellForm Compiler1
-begin
-
-subsection\<open>Correctness of program compilation \<close>
-
-primrec unmod :: "expr\<^sub>1 \<Rightarrow> nat \<Rightarrow> bool"
- and unmods :: "expr\<^sub>1 list \<Rightarrow> nat \<Rightarrow> bool" where
-"unmod (new C) i = True" |
-"unmod (Cast C e) i = unmod e i" |
-"unmod (Val v) i = True" |
-"unmod (e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2) i = (unmod e\<^sub>1 i \<and> unmod e\<^sub>2 i)" |
-"unmod (Var i) j = True" |
-"unmod (i:=e) j = (i \<noteq> j \<and> unmod e j)" |
-"unmod (e\<bullet>F{D}) i = unmod e i" |
-"unmod (C\<bullet>\<^sub>sF{D}) i = True" |
-"unmod (e\<^sub>1\<bullet>F{D}:=e\<^sub>2) i = (unmod e\<^sub>1 i \<and> unmod e\<^sub>2 i)" |
-"unmod (C\<bullet>\<^sub>sF{D}:=e\<^sub>2) i = unmod e\<^sub>2 i" |
-"unmod (e\<bullet>M(es)) i = (unmod e i \<and> unmods es i)" |
-"unmod (C\<bullet>\<^sub>sM(es)) i = unmods es i" |
-"unmod {j:T; e} i = unmod e i" |
-"unmod (e\<^sub>1;;e\<^sub>2) i = (unmod e\<^sub>1 i \<and> unmod e\<^sub>2 i)" |
-"unmod (if (e) e\<^sub>1 else e\<^sub>2) i = (unmod e i \<and> unmod e\<^sub>1 i \<and> unmod e\<^sub>2 i)" |
-"unmod (while (e) c) i = (unmod e i \<and> unmod c i)" |
-"unmod (throw e) i = unmod e i" |
-"unmod (try e\<^sub>1 catch(C i) e\<^sub>2) j = (unmod e\<^sub>1 j \<and> (if i=j then False else unmod e\<^sub>2 j))" |
-"unmod (INIT C (Cs,b) \<leftarrow> e) i = unmod e i" |
-"unmod (RI(C,e);Cs \<leftarrow> e') i = (unmod e i \<and> unmod e' i)" |
-
-"unmods ([]) i = True" |
-"unmods (e#es) i = (unmod e i \<and> unmods es i)"
-
-lemma hidden_unmod: "\<And>Vs. hidden Vs i \<Longrightarrow> unmod (compE\<^sub>1 Vs e) i" and
- "\<And>Vs. hidden Vs i \<Longrightarrow> unmods (compEs\<^sub>1 Vs es) i"
-(*<*)
-proof(induct e and es rule: compE\<^sub>1.induct compEs\<^sub>1.induct)
- case TryCatch
- then show ?case by(simp add:hidden_inacc)(auto simp add:hidden_def)
-qed (simp_all add:hidden_inacc)
-(*>*)
-
-
-lemma eval\<^sub>1_preserves_unmod:
- "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e,(h,ls,sh)\<rangle> \<Rightarrow> \<langle>e',(h',ls',sh')\<rangle>; unmod e i; i < size ls \<rbrakk>
- \<Longrightarrow> ls ! i = ls' ! i"
-and "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>es,(h,ls,sh)\<rangle> [\<Rightarrow>] \<langle>es',(h',ls',sh')\<rangle>; unmods es i; i < size ls \<rbrakk>
- \<Longrightarrow> ls ! i = ls' ! i"
-(*<*)
-proof(induct rule:eval\<^sub>1_evals\<^sub>1_inducts)
- case (RInitInitFail\<^sub>1 e h l sh a h' l' sh' C sfs i sh'' D Cs e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1)
- have "final (throw a)" using eval\<^sub>1_final[OF RInitInitFail\<^sub>1.hyps(1)] by simp
- then show ?case using RInitInitFail\<^sub>1 by(auto simp: eval\<^sub>1_preserves_len)
-qed(auto dest!:eval\<^sub>1_preserves_len evals\<^sub>1_preserves_len split:if_split_asm)
-(*>*)
-
-
-lemma LAss_lem:
- "\<lbrakk>x \<in> set xs; size xs \<le> size ys \<rbrakk>
- \<Longrightarrow> m\<^sub>1 \<subseteq>\<^sub>m m\<^sub>2(xs[\<mapsto>]ys) \<Longrightarrow> m\<^sub>1(x\<mapsto>y) \<subseteq>\<^sub>m m\<^sub>2(xs[\<mapsto>]ys[last_index xs x := y])"
-(*<*)
-by(simp add:map_le_def fun_upds_apply eq_sym_conv)
-(*>*)
-lemma Block_lem:
-fixes l :: "'a \<rightharpoonup> 'b"
-assumes 0: "l \<subseteq>\<^sub>m [Vs [\<mapsto>] ls]"
- and 1: "l' \<subseteq>\<^sub>m [Vs [\<mapsto>] ls', V\<mapsto>v]"
- and hidden: "V \<in> set Vs \<Longrightarrow> ls ! last_index Vs V = ls' ! last_index Vs V"
- and size: "size ls = size ls'" "size Vs < size ls'"
-shows "l'(V := l V) \<subseteq>\<^sub>m [Vs [\<mapsto>] ls']"
-(*<*)
-proof -
- have "l'(V := l V) \<subseteq>\<^sub>m [Vs [\<mapsto>] ls', V\<mapsto>v](V := l V)"
- using 1 by(rule map_le_upd)
- also have "\<dots> = [Vs [\<mapsto>] ls'](V := l V)" by simp
- also have "\<dots> \<subseteq>\<^sub>m [Vs [\<mapsto>] ls']"
- proof (cases "l V")
- case None thus ?thesis by simp
- next
- case (Some w)
- hence "[Vs [\<mapsto>] ls] V = Some w"
- using 0 by(force simp add: map_le_def split:if_splits)
- hence VinVs: "V \<in> set Vs" and w: "w = ls ! last_index Vs V"
- using size by(auto simp add:fun_upds_apply split:if_splits)
- hence "w = ls' ! last_index Vs V" using hidden[OF VinVs] by simp
- hence "[Vs [\<mapsto>] ls'](V := l V) = [Vs [\<mapsto>] ls']" using Some size VinVs
- by(simp add: map_upds_upd_conv_last_index)
- thus ?thesis by simp
- qed
- finally show ?thesis .
-qed
-(*>*)
-
-(*<*)
-declare fun_upd_apply[simp del]
-(*>*)
-
-
-text\<open>\noindent The main theorem: \<close>
-
-theorem assumes wf: "wwf_J_prog P"
-shows eval\<^sub>1_eval: "P \<turnstile> \<langle>e,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',(h',l',sh')\<rangle>
- \<Longrightarrow> (\<And>Vs ls. \<lbrakk> fv e \<subseteq> set Vs; l \<subseteq>\<^sub>m [Vs[\<mapsto>]ls]; size Vs + max_vars e \<le> size ls \<rbrakk>
- \<Longrightarrow> \<exists>ls'. compP\<^sub>1 P \<turnstile>\<^sub>1 \<langle>compE\<^sub>1 Vs e,(h,ls,sh)\<rangle> \<Rightarrow> \<langle>fin\<^sub>1 e',(h',ls',sh')\<rangle> \<and> l' \<subseteq>\<^sub>m [Vs[\<mapsto>]ls'])"
-(*<*)
- (is "_ \<Longrightarrow> (\<And>Vs ls. PROP ?P e h l sh e' h' l' sh' Vs ls)"
- is "_ \<Longrightarrow> (\<And>Vs ls. \<lbrakk> _; _; _ \<rbrakk> \<Longrightarrow> \<exists>ls'. ?Post e h l sh e' h' l' sh' Vs ls ls')")
-(*>*)
-
-and evals\<^sub>1_evals: "P \<turnstile> \<langle>es,(h,l,sh)\<rangle> [\<Rightarrow>] \<langle>es',(h',l',sh')\<rangle>
- \<Longrightarrow> (\<And>Vs ls. \<lbrakk> fvs es \<subseteq> set Vs; l \<subseteq>\<^sub>m [Vs[\<mapsto>]ls]; size Vs + max_varss es \<le> size ls \<rbrakk>
- \<Longrightarrow> \<exists>ls'. compP\<^sub>1 P \<turnstile>\<^sub>1 \<langle>compEs\<^sub>1 Vs es,(h,ls,sh)\<rangle> [\<Rightarrow>] \<langle>compEs\<^sub>1 Vs es',(h',ls',sh')\<rangle> \<and>
- l' \<subseteq>\<^sub>m [Vs[\<mapsto>]ls'])"
-(*<*)
- (is "_ \<Longrightarrow> (\<And>Vs ls. PROP ?Ps es h l sh es' h' l' sh' Vs ls)"
- is "_ \<Longrightarrow> (\<And>Vs ls. \<lbrakk> _; _; _\<rbrakk> \<Longrightarrow> \<exists>ls'. ?Posts es h l sh es' h' l' sh' Vs ls ls')")
-proof (induct rule:eval_evals_inducts)
- case Nil thus ?case by(fastforce intro!:Nil\<^sub>1)
-next
- case (Cons e h l sh v h' l' sh' es es' h\<^sub>2 l\<^sub>2 sh\<^sub>2)
- have "PROP ?P e h l sh (Val v) h' l' sh' Vs ls" by fact
- with Cons.prems
- obtain ls' where 1: "?Post e h l sh (Val v) h' l' sh' Vs ls ls'"
- "size ls = size ls'" by(auto intro!:eval\<^sub>1_preserves_len)
- have "PROP ?Ps es h' l' sh' es' h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls'" by fact
- with 1 Cons.prems
- obtain ls\<^sub>2 where 2: "?Posts es h' l' sh' es' h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls' ls\<^sub>2" by(auto)
- from 1 2 Cons show ?case by(auto intro!:Cons\<^sub>1)
-next
- case ConsThrow thus ?case
- by(fastforce intro!:ConsThrow\<^sub>1 dest: eval_final)
-next
- case (Block e h l V sh e' h' l' sh' T)
- let ?Vs = "Vs @ [V]"
- have IH:
- "\<lbrakk>fv e \<subseteq> set ?Vs; l(V := None) \<subseteq>\<^sub>m [?Vs [\<mapsto>] ls];
- size ?Vs + max_vars e \<le> size ls\<rbrakk>
- \<Longrightarrow> \<exists>ls'. compP\<^sub>1 P \<turnstile>\<^sub>1 \<langle>compE\<^sub>1 ?Vs e,(h,ls,sh)\<rangle> \<Rightarrow> \<langle>fin\<^sub>1 e',(h', ls',sh')\<rangle> \<and>
- l' \<subseteq>\<^sub>m [?Vs [\<mapsto>] ls']" and
- fv: "fv {V:T; e} \<subseteq> set Vs" and rel: "l \<subseteq>\<^sub>m [Vs [\<mapsto>] ls]" and
- len: "length Vs + max_vars {V:T; e} \<le> length ls" by fact+
- have len': "length Vs < length ls" using len by auto
- have "fv e \<subseteq> set ?Vs" using fv by auto
- moreover have "l(V := None) \<subseteq>\<^sub>m [?Vs [\<mapsto>] ls]" using rel len' by simp
- moreover have "size ?Vs + max_vars e \<le> size ls" using len by simp
- ultimately obtain ls' where
- 1: "compP\<^sub>1 P \<turnstile>\<^sub>1 \<langle>compE\<^sub>1 ?Vs e,(h,ls,sh)\<rangle> \<Rightarrow> \<langle>fin\<^sub>1 e',(h',ls',sh')\<rangle>"
- and rel': "l' \<subseteq>\<^sub>m [?Vs [\<mapsto>] ls']" using IH by blast
- have [simp]: "length ls = length ls'" by(rule eval\<^sub>1_preserves_len[OF 1])
- show "\<exists>ls'. compP\<^sub>1 P \<turnstile>\<^sub>1 \<langle>compE\<^sub>1 Vs {V:T; e},(h,ls,sh)\<rangle> \<Rightarrow> \<langle>fin\<^sub>1 e',(h',ls',sh')\<rangle>
- \<and> l'(V := l V) \<subseteq>\<^sub>m [Vs [\<mapsto>] ls']" (is "\<exists>ls'. ?R ls'")
- proof
- show "?R ls'"
- proof
- show "compP\<^sub>1 P \<turnstile>\<^sub>1 \<langle>compE\<^sub>1 Vs {V:T; e},(h,ls,sh)\<rangle> \<Rightarrow> \<langle>fin\<^sub>1 e',(h',ls',sh')\<rangle>"
- using 1 by(simp add:Block\<^sub>1)
- next
- show "l'(V := l V) \<subseteq>\<^sub>m [Vs [\<mapsto>] ls']"
- proof -
- have "l' \<subseteq>\<^sub>m [Vs [\<mapsto>] ls', V \<mapsto> ls' ! length Vs]"
- using len' rel' by simp
- moreover
- { assume VinVs: "V \<in> set Vs"
- hence "hidden (Vs @ [V]) (last_index Vs V)"
- by(rule hidden_last_index)
- hence "unmod (compE\<^sub>1 (Vs @ [V]) e) (last_index Vs V)"
- by(rule hidden_unmod)
- moreover have "last_index Vs V < length ls"
- using len' VinVs by simp
- ultimately have "ls ! last_index Vs V = ls' ! last_index Vs V"
- by(rule eval\<^sub>1_preserves_unmod[OF 1])
- }
- ultimately show ?thesis using Block_lem[OF rel] len' by auto
- qed
- qed
- qed
-next
- case (TryThrow e' h l sh a h' l' sh' D fs C V e\<^sub>2)
- have "PROP ?P e' h l sh (Throw a) h' l' sh' Vs ls" by fact
- with TryThrow.prems
- obtain ls' where 1: "?Post e' h l sh (Throw a) h' l' sh' Vs ls ls'" by(auto)
- show ?case using 1 TryThrow.hyps by(auto intro!:eval\<^sub>1_evals\<^sub>1.TryThrow\<^sub>1)
-next
- case (TryCatch e\<^sub>1 h l sh a h\<^sub>1 l\<^sub>1 sh\<^sub>1 D fs C e\<^sub>2 V e' h\<^sub>2 l\<^sub>2 sh\<^sub>2)
- let ?e = "try e\<^sub>1 catch(C V) e\<^sub>2"
- have IH\<^sub>1: "\<lbrakk>fv e\<^sub>1 \<subseteq> set Vs; l \<subseteq>\<^sub>m [Vs [\<mapsto>] ls];
- size Vs + max_vars e\<^sub>1 \<le> length ls\<rbrakk>
- \<Longrightarrow> \<exists>ls\<^sub>1. compP\<^sub>1 P \<turnstile>\<^sub>1 \<langle>compE\<^sub>1 Vs e\<^sub>1,(h,ls,sh)\<rangle> \<Rightarrow>
- \<langle>fin\<^sub>1 (Throw a),(h\<^sub>1,ls\<^sub>1,sh\<^sub>1)\<rangle> \<and>
- l\<^sub>1 \<subseteq>\<^sub>m [Vs [\<mapsto>] ls\<^sub>1]" and
- fv: "fv ?e \<subseteq> set Vs" and
- rel: "l \<subseteq>\<^sub>m [Vs [\<mapsto>] ls]" and
- len: "length Vs + max_vars ?e \<le> length ls" by fact+
- have "fv e\<^sub>1 \<subseteq> set Vs" using fv by auto
- moreover have "length Vs + max_vars e\<^sub>1 \<le> length ls" using len by(auto)
- ultimately obtain ls\<^sub>1 where
- 1: "compP\<^sub>1 P \<turnstile>\<^sub>1 \<langle>compE\<^sub>1 Vs e\<^sub>1,(h,ls,sh)\<rangle> \<Rightarrow> \<langle>Throw a,(h\<^sub>1,ls\<^sub>1,sh\<^sub>1)\<rangle>"
- and rel\<^sub>1: "l\<^sub>1 \<subseteq>\<^sub>m [Vs [\<mapsto>] ls\<^sub>1]" using IH\<^sub>1 rel by fastforce
- from 1 have [simp]: "size ls = size ls\<^sub>1" by(rule eval\<^sub>1_preserves_len)
- let ?Vs = "Vs @ [V]" let ?ls = "(ls\<^sub>1[size Vs:=Addr a])"
- have IH\<^sub>2: "\<lbrakk>fv e\<^sub>2 \<subseteq> set ?Vs; l\<^sub>1(V \<mapsto> Addr a) \<subseteq>\<^sub>m [?Vs [\<mapsto>] ?ls];
- length ?Vs + max_vars e\<^sub>2 \<le> length ?ls\<rbrakk> \<Longrightarrow> \<exists>ls\<^sub>2.
- compP\<^sub>1 P \<turnstile>\<^sub>1 \<langle>compE\<^sub>1 ?Vs e\<^sub>2,(h\<^sub>1,?ls,sh\<^sub>1)\<rangle> \<Rightarrow> \<langle>fin\<^sub>1 e',(h\<^sub>2, ls\<^sub>2, sh\<^sub>2)\<rangle> \<and>
- l\<^sub>2 \<subseteq>\<^sub>m [?Vs [\<mapsto>] ls\<^sub>2]" by fact
- have len\<^sub>1: "size Vs < size ls\<^sub>1" using len by(auto)
- have "fv e\<^sub>2 \<subseteq> set ?Vs" using fv by auto
- moreover have "l\<^sub>1(V \<mapsto> Addr a) \<subseteq>\<^sub>m [?Vs [\<mapsto>] ?ls]" using rel\<^sub>1 len\<^sub>1 by simp
- moreover have "length ?Vs + max_vars e\<^sub>2 \<le> length ?ls" using len by(auto)
- ultimately obtain ls\<^sub>2 where
- 2: "compP\<^sub>1 P \<turnstile>\<^sub>1 \<langle>compE\<^sub>1 ?Vs e\<^sub>2,(h\<^sub>1,?ls,sh\<^sub>1)\<rangle> \<Rightarrow> \<langle>fin\<^sub>1 e',(h\<^sub>2, ls\<^sub>2, sh\<^sub>2)\<rangle>"
- and rel\<^sub>2: "l\<^sub>2 \<subseteq>\<^sub>m [?Vs [\<mapsto>] ls\<^sub>2]" using IH\<^sub>2 by blast
- from 2 have [simp]: "size ls\<^sub>1 = size ls\<^sub>2"
- by(fastforce dest: eval\<^sub>1_preserves_len)
- show "\<exists>ls\<^sub>2. compP\<^sub>1 P \<turnstile>\<^sub>1 \<langle>compE\<^sub>1 Vs ?e,(h,ls,sh)\<rangle> \<Rightarrow> \<langle>fin\<^sub>1 e',(h\<^sub>2,ls\<^sub>2,sh\<^sub>2)\<rangle> \<and>
- l\<^sub>2(V := l\<^sub>1 V) \<subseteq>\<^sub>m [Vs [\<mapsto>] ls\<^sub>2]" (is "\<exists>ls\<^sub>2. ?R ls\<^sub>2")
- proof
- show "?R ls\<^sub>2"
- proof
- have hp: "h\<^sub>1 a = Some (D, fs)" by fact
- have "P \<turnstile> D \<preceq>\<^sup>* C" by fact hence caught: "compP\<^sub>1 P \<turnstile> D \<preceq>\<^sup>* C" by simp
- from TryCatch\<^sub>1[OF 1 _ caught len\<^sub>1 2, OF hp]
- show "compP\<^sub>1 P \<turnstile>\<^sub>1 \<langle>compE\<^sub>1 Vs ?e,(h,ls,sh)\<rangle> \<Rightarrow> \<langle>fin\<^sub>1 e',(h\<^sub>2,ls\<^sub>2,sh\<^sub>2)\<rangle>" by simp
- next
- show "l\<^sub>2(V := l\<^sub>1 V) \<subseteq>\<^sub>m [Vs [\<mapsto>] ls\<^sub>2]"
- proof -
- have "l\<^sub>2 \<subseteq>\<^sub>m [Vs [\<mapsto>] ls\<^sub>2, V \<mapsto> ls\<^sub>2 ! length Vs]"
- using len\<^sub>1 rel\<^sub>2 by simp
- moreover
- { assume VinVs: "V \<in> set Vs"
- hence "hidden (Vs @ [V]) (last_index Vs V)" by(rule hidden_last_index)
- hence "unmod (compE\<^sub>1 (Vs @ [V]) e\<^sub>2) (last_index Vs V)"
- by(rule hidden_unmod)
- moreover have "last_index Vs V < length ?ls"
- using len\<^sub>1 VinVs by simp
- ultimately have "?ls ! last_index Vs V = ls\<^sub>2 ! last_index Vs V"
- by(rule eval\<^sub>1_preserves_unmod[OF 2])
- moreover have "last_index Vs V < size Vs" using VinVs by simp
- ultimately have "ls\<^sub>1 ! last_index Vs V = ls\<^sub>2 ! last_index Vs V"
- using len\<^sub>1 by(simp del:size_last_index_conv)
- }
- ultimately show ?thesis using Block_lem[OF rel\<^sub>1] len\<^sub>1 by simp
- qed
- qed
- qed
-next
- case Try thus ?case by(fastforce intro!:Try\<^sub>1)
-next
- case Throw thus ?case by(fastforce intro!:Throw\<^sub>1)
-next
- case ThrowNull thus ?case by(fastforce intro!:ThrowNull\<^sub>1)
-next
- case ThrowThrow thus ?case by(fastforce intro!:ThrowThrow\<^sub>1)
-next
- case (CondT e h l sh h\<^sub>1 l\<^sub>1 sh\<^sub>1 e\<^sub>1 e' h\<^sub>2 l\<^sub>2 sh\<^sub>2 e\<^sub>2)
- have "PROP ?P e h l sh true h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact
- with CondT.prems
- obtain ls\<^sub>1 where 1: "?Post e h l sh true h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1"
- "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len)
- have "PROP ?P e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 e' h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact
- with 1 CondT.prems
- obtain ls\<^sub>2 where 2: "?Post e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 e' h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto)
- from 1 2 show ?case by(auto intro!:CondT\<^sub>1)
-next
- case (CondF e h l sh h\<^sub>1 l\<^sub>1 sh\<^sub>1 e\<^sub>2 e' h\<^sub>2 l\<^sub>2 sh\<^sub>2 e\<^sub>1 Vs ls)
- have "PROP ?P e h l sh false h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact
- with CondF.prems
- obtain ls\<^sub>1 where 1: "?Post e h l sh false h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1"
- "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len)
- have "PROP ?P e\<^sub>2 h\<^sub>1 l\<^sub>1 sh\<^sub>1 e' h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact
- with 1 CondF.prems
- obtain ls\<^sub>2 where 2: "?Post e\<^sub>2 h\<^sub>1 l\<^sub>1 sh\<^sub>1 e' h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto)
- from 1 2 show ?case by(auto intro!:CondF\<^sub>1)
-next
- case CondThrow thus ?case by(fastforce intro!:CondThrow\<^sub>1)
-next
- case (Seq e h l sh v h\<^sub>1 l\<^sub>1 sh\<^sub>1 e\<^sub>1 e' h\<^sub>2 l\<^sub>2 sh\<^sub>2)
- have "PROP ?P e h l sh (Val v) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact
- with Seq.prems
- obtain ls\<^sub>1 where 1: "?Post e h l sh (Val v) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1"
- "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len)
- have "PROP ?P e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 e' h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact
- with 1 Seq.prems
- obtain ls\<^sub>2 where 2: "?Post e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 e' h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto)
- from 1 2 Seq show ?case by(auto intro!:Seq\<^sub>1)
-next
- case SeqThrow thus ?case by(fastforce intro!:SeqThrow\<^sub>1)
-next
- case WhileF thus ?case by(fastforce intro!:eval\<^sub>1_evals\<^sub>1.intros)
-next
- case (WhileT e h l sh h\<^sub>1 l\<^sub>1 sh\<^sub>1 c v h\<^sub>2 l\<^sub>2 sh\<^sub>2 e' h\<^sub>3 l\<^sub>3 sh\<^sub>3)
- have "PROP ?P e h l sh true h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact
- with WhileT.prems
- obtain ls\<^sub>1 where 1: "?Post e h l sh true h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1"
- "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len)
- have "PROP ?P c h\<^sub>1 l\<^sub>1 sh\<^sub>1 (Val v) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact
- with 1 WhileT.prems
- obtain ls\<^sub>2 where 2: "?Post c h\<^sub>1 l\<^sub>1 sh\<^sub>1 (Val v) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2"
- "size ls\<^sub>1 = size ls\<^sub>2" by(auto intro!:eval\<^sub>1_preserves_len)
- have "PROP ?P (While (e) c) h\<^sub>2 l\<^sub>2 sh\<^sub>2 e' h\<^sub>3 l\<^sub>3 sh\<^sub>3 Vs ls\<^sub>2" by fact
- with 1 2 WhileT.prems
- obtain ls\<^sub>3 where 3: "?Post (While (e) c) h\<^sub>2 l\<^sub>2 sh\<^sub>2 e' h\<^sub>3 l\<^sub>3 sh\<^sub>3 Vs ls\<^sub>2 ls\<^sub>3" by(auto)
- from 1 2 3 show ?case by(auto intro!:WhileT\<^sub>1)
-next
- case (WhileBodyThrow e h l sh h\<^sub>1 l\<^sub>1 sh\<^sub>1 c e' h\<^sub>2 l\<^sub>2 sh\<^sub>2)
- have "PROP ?P e h l sh true h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact
- with WhileBodyThrow.prems
- obtain ls\<^sub>1 where 1: "?Post e h l sh true h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1"
- "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len)
- have "PROP ?P c h\<^sub>1 l\<^sub>1 sh\<^sub>1 (throw e') h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact
- with 1 WhileBodyThrow.prems
- obtain ls\<^sub>2 where 2: "?Post c h\<^sub>1 l\<^sub>1 sh\<^sub>1 (throw e') h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by auto
- from 1 2 show ?case by(auto intro!:WhileBodyThrow\<^sub>1)
-next
- case WhileCondThrow thus ?case by(fastforce intro!:WhileCondThrow\<^sub>1)
-next
- case New thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros)
-next
- case NewFail thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros)
-next
- case NewInit then show ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros)
-next
- case NewInitOOM then show ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros)
-next
- case NewInitThrow then show ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros)
-next
- case Cast thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros)
-next
- case CastNull thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros)
-next
- case CastThrow thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros)
-next
- case (CastFail e h l sh a h\<^sub>1 l\<^sub>1 sh\<^sub>1 D fs C)
- have "PROP ?P e h l sh (addr a) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact
- with CastFail.prems
- obtain ls\<^sub>1 where 1: "?Post e h l sh (addr a) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1" by auto
- show ?case using 1 CastFail.hyps
- by(auto intro!:CastFail\<^sub>1[where D=D])
-next
- case Val thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros)
-next
- case (BinOp e h l sh v\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 e\<^sub>1 v\<^sub>2 h\<^sub>2 l\<^sub>2 sh\<^sub>2 bop v)
- have "PROP ?P e h l sh (Val v\<^sub>1) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact
- with BinOp.prems
- obtain ls\<^sub>1 where 1: "?Post e h l sh (Val v\<^sub>1) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1"
- "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len)
- have "PROP ?P e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 (Val v\<^sub>2) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact
- with 1 BinOp.prems
- obtain ls\<^sub>2 where 2: "?Post e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 (Val v\<^sub>2) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto)
- from 1 2 BinOp show ?case by(auto intro!:BinOp\<^sub>1)
-next
- case (BinOpThrow2 e\<^sub>0 h l sh v\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 e\<^sub>1 e h\<^sub>2 l\<^sub>2 sh\<^sub>2 bop)
- have "PROP ?P e\<^sub>0 h l sh (Val v\<^sub>1) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact
- with BinOpThrow2.prems
- obtain ls\<^sub>1 where 1: "?Post e\<^sub>0 h l sh (Val v\<^sub>1) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1"
- "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len)
- have "PROP ?P e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 (throw e) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact
- with 1 BinOpThrow2.prems
- obtain ls\<^sub>2 where 2: "?Post e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 (throw e) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto)
- from 1 2 BinOpThrow2 show ?case by(auto intro!:BinOpThrow\<^sub>2\<^sub>1)
-next
- case BinOpThrow1 thus ?case by(fastforce intro!:eval\<^sub>1_evals\<^sub>1.intros)
-next
- case Var thus ?case
- by(force intro!:Var\<^sub>1 simp add: map_le_def fun_upds_apply)
-next
- case LAss thus ?case
- by(fastforce simp add: LAss_lem intro:eval\<^sub>1_evals\<^sub>1.intros
- dest:eval\<^sub>1_preserves_len)
-next
- case LAssThrow thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros)
-next
- case FAcc thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros)
-next
- case FAccNull thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros)
-next
- case FAccThrow thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros)
-next
- case (FAccNone e h l sh a h' l' sh' C fs F D)
- have "PROP ?P e h l sh (addr a) h' l' sh' Vs ls" by fact
- with FAccNone.prems
- obtain ls\<^sub>2 where 2: "?Post e h l sh (addr a) h' l' sh' Vs ls ls\<^sub>2" by(auto)
- from 2 FAccNone show ?case by(rule_tac x = ls\<^sub>2 in exI, auto elim!: FAccNone\<^sub>1)
-next
- case (FAccStatic e h l sh a h' l' sh' C fs F t D)
- have "PROP ?P e h l sh (addr a) h' l' sh' Vs ls" by fact
- with FAccStatic.prems
- obtain ls\<^sub>2 where 2: "?Post e h l sh (addr a) h' l' sh' Vs ls ls\<^sub>2" by(auto)
- from 2 FAccStatic show ?case by(rule_tac x = ls\<^sub>2 in exI, auto elim!: FAccStatic\<^sub>1)
-next
- case SFAcc then show ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros)
-next
- case (SFAccInit C F t D sh h l v' h' l' sh' sfs i v)
- have "PROP ?P (INIT D ([D],False) \<leftarrow> unit) h l sh (Val v') h' l' sh' Vs ls" by fact
- with SFAccInit.prems
- obtain ls\<^sub>2 where 1: "?Post (INIT D ([D],False) \<leftarrow> unit) h l sh (Val v') h' l' sh' Vs ls ls\<^sub>2" by(auto)
- from 1 SFAccInit show ?case by(rule_tac x = ls\<^sub>2 in exI, auto intro: SFAccInit\<^sub>1)
-next
- case (SFAccInitThrow C F t D sh h l a h' l' sh')
- have "PROP ?P (INIT D ([D],False) \<leftarrow> unit) h l sh (throw a) h' l' sh' Vs ls" by fact
- with SFAccInitThrow.prems
- obtain ls\<^sub>2 where 1: "?Post (INIT D ([D],False) \<leftarrow> unit) h l sh (throw a) h' l' sh' Vs ls ls\<^sub>2" by(auto)
- from 1 SFAccInitThrow show ?case by(rule_tac x = ls\<^sub>2 in exI, auto intro: SFAccInitThrow\<^sub>1)
-next
- case SFAccNone then show ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros)
-next
- case SFAccNonStatic then show ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros)
-next
- case (FAss e\<^sub>1 h l sh a h\<^sub>1 l\<^sub>1 sh\<^sub>1 e\<^sub>2 v h\<^sub>2 l\<^sub>2 sh\<^sub>2 C fs fs' F D h\<^sub>2')
- have "PROP ?P e\<^sub>1 h l sh (addr a) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact
- with FAss.prems
- obtain ls\<^sub>1 where 1: "?Post e\<^sub>1 h l sh (addr a) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1"
- "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len)
- have "PROP ?P e\<^sub>2 h\<^sub>1 l\<^sub>1 sh\<^sub>1 (Val v) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact
- with 1 FAss.prems
- obtain ls\<^sub>2 where 2: "?Post e\<^sub>2 h\<^sub>1 l\<^sub>1 sh\<^sub>1 (Val v) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto)
- from 1 2 FAss show ?case by(auto intro!:FAss\<^sub>1)
-next
- case (FAssNull e\<^sub>1 h l sh h\<^sub>1 l\<^sub>1 sh\<^sub>1 e\<^sub>2 v h\<^sub>2 l\<^sub>2 sh\<^sub>2 F D)
- have "PROP ?P e\<^sub>1 h l sh null h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact
- with FAssNull.prems
- obtain ls\<^sub>1 where 1: "?Post e\<^sub>1 h l sh null h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1"
- "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len)
- have "PROP ?P e\<^sub>2 h\<^sub>1 l\<^sub>1 sh\<^sub>1 (Val v) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact
- with 1 FAssNull.prems
- obtain ls\<^sub>2 where 2: "?Post e\<^sub>2 h\<^sub>1 l\<^sub>1 sh\<^sub>1 (Val v) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto)
- from 1 2 FAssNull show ?case by(auto intro!:FAssNull\<^sub>1)
-next
- case FAssThrow1 thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros)
-next
- case (FAssThrow2 e\<^sub>1 h l sh v h\<^sub>1 l\<^sub>1 sh\<^sub>1 e\<^sub>2 e h\<^sub>2 l\<^sub>2 sh\<^sub>2 F D)
- have "PROP ?P e\<^sub>1 h l sh (Val v) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact
- with FAssThrow2.prems
- obtain ls\<^sub>1 where 1: "?Post e\<^sub>1 h l sh (Val v) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1"
- "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len)
- have "PROP ?P e\<^sub>2 h\<^sub>1 l\<^sub>1 sh\<^sub>1 (throw e) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact
- with 1 FAssThrow2.prems
- obtain ls\<^sub>2 where 2: "?Post e\<^sub>2 h\<^sub>1 l\<^sub>1 sh\<^sub>1 (throw e) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto)
- from 1 2 FAssThrow2 show ?case by(auto intro!:FAssThrow\<^sub>2\<^sub>1)
-next
- case (FAssNone e\<^sub>1 h l sh a h' l' sh' e\<^sub>2 v h\<^sub>2 l\<^sub>2 sh\<^sub>2 C fs F D)
- have "PROP ?P e\<^sub>1 h l sh (addr a) h' l' sh' Vs ls" by fact
- with FAssNone.prems
- obtain ls\<^sub>1 where 1: "?Post e\<^sub>1 h l sh (addr a) h' l' sh' Vs ls ls\<^sub>1"
- "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len)
- have "PROP ?P e\<^sub>2 h' l' sh' (Val v) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact
- with 1 FAssNone.prems
- obtain ls\<^sub>2 where 2: "?Post e\<^sub>2 h' l' sh' (Val v) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto)
- from 1 2 FAssNone show ?case by(auto intro!:FAssNone\<^sub>1)
-next
- case (FAssStatic e\<^sub>1 h l sh a h' l' sh' e\<^sub>2 v h\<^sub>2 l\<^sub>2 sh\<^sub>2 C fs F t D)
- have "PROP ?P e\<^sub>1 h l sh (addr a) h' l' sh' Vs ls" by fact
- with FAssStatic.prems
- obtain ls\<^sub>1 where 1: "?Post e\<^sub>1 h l sh (addr a) h' l' sh' Vs ls ls\<^sub>1"
- "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len)
- have "PROP ?P e\<^sub>2 h' l' sh' (Val v) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact
- with 1 FAssStatic.prems
- obtain ls\<^sub>2 where 2: "?Post e\<^sub>2 h' l' sh' (Val v) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto)
- from 1 2 FAssStatic show ?case by(auto intro!:FAssStatic\<^sub>1)
-next
- case SFAss then show ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros)
-next
- case (SFAssInit e\<^sub>2 h l sh v h\<^sub>1 l\<^sub>1 sh\<^sub>1 C F t D v' h' l' sh' sfs i sfs' sh'')
- have "PROP ?P e\<^sub>2 h l sh (Val v) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact
- with SFAssInit.prems
- obtain ls\<^sub>1 where 1: "?Post e\<^sub>2 h l sh (Val v) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1" "length ls = length ls\<^sub>1"
- by(auto intro!:eval\<^sub>1_preserves_len)
- then have Init_size: "length Vs \<le> length ls\<^sub>1" using SFAssInit.prems(3) by linarith
- have "PROP ?P (INIT D ([D],False) \<leftarrow> unit) h\<^sub>1 l\<^sub>1 sh\<^sub>1 (Val v') h' l' sh' Vs ls\<^sub>1" by fact
- with 1 Init_size SFAssInit.prems
- obtain ls\<^sub>2 where 2: "?Post (INIT D ([D],False) \<leftarrow> unit) h\<^sub>1 l\<^sub>1 sh\<^sub>1 (Val v') h' l' sh' Vs ls\<^sub>1 ls\<^sub>2"
- by auto
- from 1 2 SFAssInit show ?case
- by(auto simp add: comp_def
- intro!: SFAssInit\<^sub>1 dest!:evals_final)
-next
- case (SFAssInitThrow e\<^sub>2 h l sh v h\<^sub>1 l\<^sub>1 sh\<^sub>1 C F t D a h\<^sub>2 l\<^sub>2 sh\<^sub>2)
- have "PROP ?P e\<^sub>2 h l sh (Val v) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact
- with SFAssInitThrow.prems
- obtain ls\<^sub>1 where 1: "?Post e\<^sub>2 h l sh (Val v) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1" "length ls = length ls\<^sub>1"
- by(auto intro!:eval\<^sub>1_preserves_len)
- then have Init_size: "length Vs \<le> length ls\<^sub>1" using SFAssInitThrow.prems(3) by linarith
- have "PROP ?P (INIT D ([D],False) \<leftarrow> unit) h\<^sub>1 l\<^sub>1 sh\<^sub>1 (throw a) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact
- with 1 Init_size SFAssInitThrow.prems
- obtain ls\<^sub>2 where 2: "?Post (INIT D ([D],False) \<leftarrow> unit) h\<^sub>1 l\<^sub>1 sh\<^sub>1 (throw a) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2"
- by auto
- from 1 2 SFAssInitThrow show ?case
- by(auto simp add: comp_def
- intro!: SFAssInitThrow\<^sub>1 dest!:evals_final)
-next
- case SFAssThrow then show ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros)
-next
- case (SFAssNone e\<^sub>2 h l sh v h\<^sub>2 l\<^sub>2 sh\<^sub>2 C F D)
- have "PROP ?P e\<^sub>2 h l sh (Val v) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls" by fact
- with SFAssNone.prems
- obtain ls\<^sub>2 where 2: "?Post e\<^sub>2 h l sh (Val v) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls ls\<^sub>2" by(auto)
- from 2 SFAssNone show ?case by(rule_tac x = ls\<^sub>2 in exI, auto elim!: SFAssNone\<^sub>1)
-next
- case SFAssNonStatic then show ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros)
-next
- case (CallNull e h l sh h\<^sub>1 l\<^sub>1 sh\<^sub>1 es vs h\<^sub>2 l\<^sub>2 sh\<^sub>2 M)
- have "PROP ?P e h l sh null h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact
- with CallNull.prems
- obtain ls\<^sub>1 where 1: "?Post e h l sh null h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1"
- "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len)
- have "PROP ?Ps es h\<^sub>1 l\<^sub>1 sh\<^sub>1 (map Val vs) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact
- with 1 CallNull.prems
- obtain ls\<^sub>2 where 2: "?Posts es h\<^sub>1 l\<^sub>1 sh\<^sub>1 (map Val vs) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto)
- from 1 2 CallNull show ?case
- by (auto simp add: comp_def elim!: CallNull\<^sub>1)
-next
- case CallObjThrow thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros)
-next
- case (CallParamsThrow e h l sh v h\<^sub>1 l\<^sub>1 sh\<^sub>1 es vs ex es' h\<^sub>2 l\<^sub>2 sh\<^sub>2 M)
- have "PROP ?P e h l sh (Val v) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact
- with CallParamsThrow.prems
- obtain ls\<^sub>1 where 1: "?Post e h l sh (Val v) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1"
- "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len)
- have "PROP ?Ps es h\<^sub>1 l\<^sub>1 sh\<^sub>1 (map Val vs @ throw ex # es') h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact
- with 1 CallParamsThrow.prems
- obtain ls\<^sub>2 where 2: "?Posts es h\<^sub>1 l\<^sub>1 sh\<^sub>1 (map Val vs @ throw ex # es') h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto)
- from 1 2 CallParamsThrow show ?case
- by (auto simp add: comp_def
- elim!: CallParamsThrow\<^sub>1 dest!:evals_final)
-next
- case (CallNone e h l sh a h\<^sub>1 l\<^sub>1 sh\<^sub>1 ps vs h\<^sub>2 l\<^sub>2 sh\<^sub>2 C fs M)
- have "PROP ?P e h l sh (addr a) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact
- with CallNone.prems
- obtain ls\<^sub>1 where 1: "?Post e h l sh (addr a) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1"
- "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len)
- have "PROP ?Ps ps h\<^sub>1 l\<^sub>1 sh\<^sub>1 (map Val vs) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact
- with 1 CallNone.prems
- obtain ls\<^sub>2 where 2: "?Posts ps h\<^sub>1 l\<^sub>1 sh\<^sub>1 (map Val vs) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto)
- from 1 2 CallNone show ?case
- by (auto simp add: comp_def
- elim!: CallNone\<^sub>1 dest!:evals_final sees_method_compPD)
-next
- case (CallStatic e h l sh a h\<^sub>1 l\<^sub>1 sh\<^sub>1 ps vs h\<^sub>2 l\<^sub>2 sh\<^sub>2 C fs M Ts T pns body D)
- have "PROP ?P e h l sh (addr a) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact
- with CallStatic.prems
- obtain ls\<^sub>1 where 1: "?Post e h l sh (addr a) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1"
- "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len)
- let ?Vs = pns
- have mdecl: "P \<turnstile> C sees M,Static: Ts\<rightarrow>T = (pns, body) in D" by fact
- have mdecl\<^sub>1: "compP\<^sub>1 P \<turnstile> C sees M,Static: Ts\<rightarrow>T = (compE\<^sub>1 ?Vs body) in D"
- using sees_method_compP[OF mdecl, of "\<lambda>b (pns,e). compE\<^sub>1 (case b of NonStatic \<Rightarrow> this#pns | Static \<Rightarrow> pns) e"]
- by(simp)
- have "PROP ?Ps ps h\<^sub>1 l\<^sub>1 sh\<^sub>1 (map Val vs) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact
- with 1 CallStatic.prems
- obtain ls\<^sub>2 where 2: "?Posts ps h\<^sub>1 l\<^sub>1 sh\<^sub>1 (map Val vs) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto)
- from 1 2 mdecl\<^sub>1 CallStatic show ?case
- by (auto simp add: comp_def
- elim!: CallStatic\<^sub>1 dest!:evals_final)
-next
- case (Call e h l sh a h\<^sub>1 l\<^sub>1 sh\<^sub>1 es vs h\<^sub>2 l\<^sub>2 sh\<^sub>2 C fs M Ts T pns body D l\<^sub>2' b' h\<^sub>3 l\<^sub>3 sh\<^sub>3)
- have "PROP ?P e h l sh (addr a) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact
- with Call.prems
- obtain ls\<^sub>1 where 1: "?Post e h l sh (addr a) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1"
- "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len)
- have "PROP ?Ps es h\<^sub>1 l\<^sub>1 sh\<^sub>1 (map Val vs) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact
- with 1 Call.prems
- obtain ls\<^sub>2 where 2: "?Posts es h\<^sub>1 l\<^sub>1 sh\<^sub>1 (map Val vs) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2"
- "size ls\<^sub>1 = size ls\<^sub>2" by(auto intro!:evals\<^sub>1_preserves_len)
- let ?Vs = "this#pns"
- let ?ls = "Addr a # vs @ replicate (max_vars body) undefined"
- have mdecl: "P \<turnstile> C sees M,NonStatic: Ts\<rightarrow>T = (pns, body) in D" by fact
- have fv_body: "fv body \<subseteq> set ?Vs" and wf_size: "size Ts = size pns"
- using wf mdecl by(auto dest!:sees_wf_mdecl simp:wf_mdecl_def)
- have mdecl\<^sub>1: "compP\<^sub>1 P \<turnstile> C sees M,NonStatic: Ts\<rightarrow>T = (compE\<^sub>1 ?Vs body) in D"
- using sees_method_compP[OF mdecl, of "\<lambda>b (pns,e). compE\<^sub>1 (case b of NonStatic \<Rightarrow> this#pns | Static \<Rightarrow> pns) e"]
- by(simp)
- have [simp]: "l\<^sub>2' = [this \<mapsto> Addr a, pns [\<mapsto>] vs]" by fact
- have Call_size: "size vs = size pns" by fact
- have "PROP ?P body h\<^sub>2 l\<^sub>2' sh\<^sub>2 b' h\<^sub>3 l\<^sub>3 sh\<^sub>3 ?Vs ?ls" by fact
- with 1 2 fv_body Call_size Call.prems
- obtain ls\<^sub>3 where 3: "?Post body h\<^sub>2 l\<^sub>2' sh\<^sub>2 b' h\<^sub>3 l\<^sub>3 sh\<^sub>3 ?Vs ?ls ls\<^sub>3" by(auto)
- have hp: "h\<^sub>2 a = Some (C, fs)" by fact
- from 1 2 3 hp mdecl\<^sub>1 wf_size Call_size show ?case
- by(fastforce simp add: comp_def
- intro!: Call\<^sub>1 dest!:evals_final)
-next
- case (SCallParamsThrow es h l sh vs ex es' h\<^sub>2 l\<^sub>2 sh\<^sub>2 C M)
- have "PROP ?Ps es h l sh (map Val vs @ throw ex # es') h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls" by fact
- with SCallParamsThrow.prems
- obtain ls\<^sub>2 where 2: "?Posts es h l sh (map Val vs @ throw ex # es') h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls ls\<^sub>2" by(auto)
- from 2 SCallParamsThrow show ?case
- by (fastforce simp add: comp_def
- elim!: SCallParamsThrow\<^sub>1 dest!:evals_final)
-next
- case (SCall ps h l sh vs h\<^sub>2 l\<^sub>2 sh\<^sub>2 C M Ts T pns body D sfs l\<^sub>2' e' h\<^sub>3 l\<^sub>3 sh\<^sub>3)
- have "PROP ?Ps ps h l sh (map Val vs) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls" by fact
- with SCall.prems
- obtain ls\<^sub>2 where 2: "?Posts ps h l sh (map Val vs) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls ls\<^sub>2"
- "size ls = size ls\<^sub>2" by(auto intro!:evals\<^sub>1_preserves_len)
- let ?Vs = "pns"
- let ?ls = "vs @ replicate (max_vars body) undefined"
- have mdecl: "P \<turnstile> C sees M,Static: Ts\<rightarrow>T = (pns, body) in D" by fact
- have fv_body: "fv body \<subseteq> set ?Vs" and wf_size: "size Ts = size pns"
- using wf mdecl by(auto dest!:sees_wf_mdecl simp:wf_mdecl_def)
- have mdecl\<^sub>1: "compP\<^sub>1 P \<turnstile> C sees M,Static: Ts\<rightarrow>T = (compE\<^sub>1 ?Vs body) in D"
- using sees_method_compP[OF mdecl, of "\<lambda>b (pns,e). compE\<^sub>1 (case b of NonStatic \<Rightarrow> this#pns | Static \<Rightarrow> pns) e"]
- by(simp)
- have [simp]: "l\<^sub>2' = [pns [\<mapsto>] vs]" by fact
- have SCall_size: "size vs = size pns" by fact
- have "PROP ?P body h\<^sub>2 l\<^sub>2' sh\<^sub>2 e' h\<^sub>3 l\<^sub>3 sh\<^sub>3 ?Vs ?ls" by fact
- with 2 fv_body SCall_size SCall.prems
- obtain ls\<^sub>3 where 3: "?Post body h\<^sub>2 l\<^sub>2' sh\<^sub>2 e' h\<^sub>3 l\<^sub>3 sh\<^sub>3 ?Vs ?ls ls\<^sub>3" by(auto)
- have shp: "sh\<^sub>2 D = \<lfloor>(sfs, Done)\<rfloor> \<or> M = clinit \<and> sh\<^sub>2 D = \<lfloor>(sfs, Processing)\<rfloor>" by fact
- from 2 3 shp mdecl\<^sub>1 wf_size SCall_size show ?case
- by(fastforce simp add: comp_def
- intro!: SCall\<^sub>1 dest!:evals_final)
-next
- case (SCallNone ps h l sh vs h' l' sh' C M)
- have "PROP ?Ps ps h l sh (map Val vs) h' l' sh' Vs ls" by fact
- with SCallNone.prems
- obtain ls\<^sub>2 where 2: "?Posts ps h l sh (map Val vs) h' l' sh' Vs ls ls\<^sub>2" by(auto)
- from 2 SCallNone show ?case
- by(rule_tac x = ls\<^sub>2 in exI,
- auto simp add: comp_def elim!: SCallNone\<^sub>1 dest!:evals_final sees_method_compPD)
-next
- case (SCallNonStatic ps h l sh vs h' l' sh' C M Ts T pns body D)
- let ?Vs = "this#pns"
- have mdecl: "P \<turnstile> C sees M,NonStatic: Ts\<rightarrow>T = (pns, body) in D" by fact
- have mdecl\<^sub>1: "compP\<^sub>1 P \<turnstile> C sees M,NonStatic: Ts\<rightarrow>T = (compE\<^sub>1 ?Vs body) in D"
- using sees_method_compP[OF mdecl, of "\<lambda>b (pns,e). compE\<^sub>1 (case b of NonStatic \<Rightarrow> this#pns | Static \<Rightarrow> pns) e"]
- by(simp)
- have "PROP ?Ps ps h l sh (map Val vs) h' l' sh' Vs ls" by fact
- with SCallNonStatic.prems
- obtain ls\<^sub>2 where 2: "?Posts ps h l sh (map Val vs) h' l' sh' Vs ls ls\<^sub>2" by(auto)
- from 2 mdecl\<^sub>1 SCallNonStatic show ?case
- by (auto simp add: comp_def
- elim!: SCallNonStatic\<^sub>1 dest!:evals_final)
-next
- case (SCallInitThrow ps h l sh vs h\<^sub>1 l\<^sub>1 sh\<^sub>1 C M Ts T pns body D a h\<^sub>2 l\<^sub>2 sh\<^sub>2)
- have "PROP ?Ps ps h l sh (map Val vs) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact
- with SCallInitThrow.prems
- obtain ls\<^sub>1 where 1: "?Posts ps h l sh (map Val vs) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1" "length ls = length ls\<^sub>1"
- by(auto intro!:evals\<^sub>1_preserves_len)
- then have Init_size: "length Vs \<le> length ls\<^sub>1" using SCallInitThrow.prems(3) by linarith
- have "PROP ?P (INIT D ([D],False) \<leftarrow> unit) h\<^sub>1 l\<^sub>1 sh\<^sub>1 (throw a) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact
- with 1 Init_size SCallInitThrow.prems
- obtain ls\<^sub>2 where 2: "?Post (INIT D ([D],False) \<leftarrow> unit) h\<^sub>1 l\<^sub>1 sh\<^sub>1 (throw a) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2"
- by auto
- let ?Vs = "pns"
- have mdecl: "P \<turnstile> C sees M,Static: Ts\<rightarrow>T = (pns, body) in D" by fact
- have mdecl\<^sub>1: "compP\<^sub>1 P \<turnstile> C sees M,Static: Ts\<rightarrow>T = (compE\<^sub>1 ?Vs body) in D"
- using sees_method_compP[OF mdecl, of "\<lambda>b (pns,e). compE\<^sub>1 (case b of NonStatic \<Rightarrow> this#pns | Static \<Rightarrow> pns) e"]
- by(simp)
- from 1 2 mdecl\<^sub>1 SCallInitThrow show ?case
- by(auto simp add: comp_def
- intro!: SCallInitThrow\<^sub>1 dest!:evals_final)
-next
- case (SCallInit ps h l sh vs h\<^sub>1 l\<^sub>1 sh\<^sub>1 C M Ts T pns body D v' h\<^sub>2 l\<^sub>2 sh\<^sub>2 l\<^sub>2' e' h\<^sub>3 l\<^sub>3 sh\<^sub>3)
- have "PROP ?Ps ps h l sh (map Val vs) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact
- with SCallInit.prems
- obtain ls\<^sub>1 where 1: "?Posts ps h l sh (map Val vs) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1" "length ls = length ls\<^sub>1"
- by(auto intro!:evals\<^sub>1_preserves_len)
- then have Init_size: "length Vs \<le> length ls\<^sub>1" using SCallInit.prems(3) by linarith
- have "PROP ?P (INIT D ([D],False) \<leftarrow> unit) h\<^sub>1 l\<^sub>1 sh\<^sub>1 (Val v') h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact
- with 1 Init_size SCallInit.prems
- obtain ls\<^sub>2 where 2: "?Post (INIT D ([D],False) \<leftarrow> unit) h\<^sub>1 l\<^sub>1 sh\<^sub>1 (Val v') h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2"
- by auto
- let ?Vs = "pns"
- let ?ls = "vs @ replicate (max_vars body) undefined"
- have mdecl: "P \<turnstile> C sees M,Static: Ts\<rightarrow>T = (pns, body) in D" by fact
- have fv_body: "fv body \<subseteq> set ?Vs" and wf_size: "size Ts = size pns"
- using wf mdecl by(auto dest!:sees_wf_mdecl simp:wf_mdecl_def)
- have mdecl\<^sub>1: "compP\<^sub>1 P \<turnstile> C sees M,Static: Ts\<rightarrow>T = (compE\<^sub>1 ?Vs body) in D"
- using sees_method_compP[OF mdecl, of "\<lambda>b (pns,e). compE\<^sub>1 (case b of NonStatic \<Rightarrow> this#pns | Static \<Rightarrow> pns) e"]
- by(simp)
- have [simp]: "l\<^sub>2' = [pns [\<mapsto>] vs]" by fact
- have SCall_size: "size vs = size pns" by fact
- have nclinit: "M \<noteq> clinit" by fact
- have "PROP ?P body h\<^sub>2 l\<^sub>2' sh\<^sub>2 e' h\<^sub>3 l\<^sub>3 sh\<^sub>3 ?Vs ?ls" by fact
- with 2 fv_body SCall_size SCallInit.prems
- obtain ls\<^sub>3 where 3: "?Post body h\<^sub>2 l\<^sub>2' sh\<^sub>2 e' h\<^sub>3 l\<^sub>3 sh\<^sub>3 ?Vs ?ls ls\<^sub>3" by(auto)
- have shp: "\<nexists>sfs. sh\<^sub>1 D = \<lfloor>(sfs, Done)\<rfloor>" by fact
- from 1 2 3 shp mdecl\<^sub>1 wf_size SCall_size nclinit show ?case
- by(auto simp add: comp_def
- intro!: SCallInit\<^sub>1 dest!:evals_final)
-next
-\<comment> \<open> init cases \<close>
- case InitFinal then show ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros)
-next
- case (InitNone sh C C' Cs e h l e' h' l' sh')
- let ?sh1 = "sh(C \<mapsto> (sblank P C, Prepared))"
- have "PROP ?P (INIT C' (C # Cs,False) \<leftarrow> e) h l ?sh1 e' h' l' sh' Vs ls" by fact
- with InitNone.prems
- obtain ls\<^sub>2 where 2: "?Post (INIT C' (C # Cs,False) \<leftarrow> e) h l ?sh1 e' h' l' sh' Vs ls ls\<^sub>2" by(auto)
- from 2 InitNone show ?case by (auto elim!: InitNone\<^sub>1)
-next
- case InitDone then show ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros)
-next
- case InitProcessing then show ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros)
-next
- case InitError then show ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros)
-next
- case InitObject then show ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros)
-next
- case (InitNonObject sh C sfs D fs ms sh' C' Cs e h l e' h1 l1 sh1)
- let ?f = "(\<lambda>b (pns,body). compE\<^sub>1 (case b of NonStatic \<Rightarrow> this#pns | Static \<Rightarrow> pns) body)"
- have cls: "class (compP ?f P) C = \<lfloor>(D,fs,map (compM ?f) ms)\<rfloor>"
- by(rule class_compP[OF InitNonObject.hyps(3)])
- have "PROP ?P (INIT C' (D # C # Cs,False) \<leftarrow> e) h l sh' e' h1 l1 sh1 Vs ls" by fact
- with InitNonObject.prems
- obtain ls\<^sub>2 where 2: "?Post (INIT C' (D # C # Cs,False) \<leftarrow> e) h l sh' e' h1 l1 sh1 Vs ls ls\<^sub>2" by(auto)
- from 2 cls InitNonObject show ?case by (auto elim!: InitNonObject\<^sub>1)
-next
- case InitRInit then show ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros)
-next
- case (RInit e h l sh v h' l' sh' C sfs i sh'' C' Cs e' e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1)
- have "PROP ?P e h l sh (Val v) h' l' sh' Vs ls" by fact
- with RInit.prems
- obtain ls\<^sub>1 where 1: "?Post e h l sh (Val v) h' l' sh' Vs ls ls\<^sub>1"
- "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len)
- have "PROP ?P (INIT C' (Cs,True) \<leftarrow> e') h' l' sh'' e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls\<^sub>1" by fact
- with 1 RInit.prems
- obtain ls\<^sub>2 where 2: "?Post (INIT C' (Cs,True) \<leftarrow> e') h' l' sh'' e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls\<^sub>1 ls\<^sub>2" by(auto)
- from 1 2 RInit show ?case by (auto elim!: RInit\<^sub>1)
-next
- case (RInitInitFail e h l sh a h' l' sh' C sfs i sh'' D Cs e' e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1)
- have "PROP ?P e h l sh (throw a) h' l' sh' Vs ls" by fact
- with RInitInitFail.prems
- obtain ls\<^sub>1 where 1: "?Post e h l sh (throw a) h' l' sh' Vs ls ls\<^sub>1"
- "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len)
- have fv: "fv (RI (D,throw a) ; Cs \<leftarrow> e') \<subseteq> set Vs"
- using RInitInitFail.hyps(1) eval_final RInitInitFail.prems(1) subset_eq by fastforce
- have l': "l' \<subseteq>\<^sub>m [Vs [\<mapsto>] ls\<^sub>1]" by (simp add: "1"(1))
- have "PROP ?P (RI (D,throw a) ; Cs \<leftarrow> e') h' l' sh'' e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls\<^sub>1" by fact
- with 1 eval_final[OF RInitInitFail.hyps(1)] RInitInitFail.prems
- obtain ls\<^sub>2 where 2: "?Post (RI (D,throw a) ; Cs \<leftarrow> e') h' l' sh'' e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls\<^sub>1 ls\<^sub>2"
- by fastforce
- from 1 2 RInitInitFail show ?case
- by(fastforce simp add: comp_def
- intro!: RInitInitFail\<^sub>1 dest!:eval_final)
-next
- case RInitFailFinal then show ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros)
-qed
-(*>*)
-
-
-subsection\<open>Preservation of well-formedness\<close>
-
-text\<open> The compiler preserves well-formedness. Is less trivial than it
-may appear. We start with two simple properties: preservation of
-well-typedness \<close>
-
-lemma compE\<^sub>1_pres_wt: "\<And>Vs Ts U.
- \<lbrakk> P,[Vs[\<mapsto>]Ts] \<turnstile> e :: U; size Ts = size Vs \<rbrakk>
- \<Longrightarrow> compP f P,Ts \<turnstile>\<^sub>1 compE\<^sub>1 Vs e :: U"
-and "\<And>Vs Ts Us.
- \<lbrakk> P,[Vs[\<mapsto>]Ts] \<turnstile> es [::] Us; size Ts = size Vs \<rbrakk>
- \<Longrightarrow> compP f P,Ts \<turnstile>\<^sub>1 compEs\<^sub>1 Vs es [::] Us"
-(*<*)
-proof(induct e and es rule: compE\<^sub>1.induct compEs\<^sub>1.induct)
- case Var then show ?case
- by (fastforce simp:map_upds_apply_eq_Some split:if_split_asm)
-next
- case LAss then show ?case
- by (fastforce simp:map_upds_apply_eq_Some split:if_split_asm)
-next
- case Call then show ?case
- by (fastforce dest!: sees_method_compP[where f = f])
-next
- case SCall then show ?case
- by (fastforce dest!: sees_method_compP[where f = f])
-next
- case Block then show ?case by (fastforce simp:nth_append)
-next
- case TryCatch then show ?case by (fastforce simp:nth_append)
-qed fastforce+
-(*>*)
-
-text\<open>\noindent and the correct block numbering: \<close>
-
-lemma \<B>: "\<And>Vs n. size Vs = n \<Longrightarrow> \<B> (compE\<^sub>1 Vs e) n"
-and \<B>s: "\<And>Vs n. size Vs = n \<Longrightarrow> \<B>s (compEs\<^sub>1 Vs es) n"
-(*<*)by (induct e and es rule: \<B>.induct \<B>s.induct)
- (force | simp,metis length_append_singleton)+(*>*)
-
-text\<open> The main complication is preservation of definite assignment
-@{term"\<D>"}. \<close>
-
-lemma image_last_index: "A \<subseteq> set(xs@[x]) \<Longrightarrow> last_index (xs @ [x]) ` A =
- (if x \<in> A then insert (size xs) (last_index xs ` (A-{x})) else last_index xs ` A)"
-(*<*)
-by(auto simp:image_def)
-(*>*)
-
-
-lemma A_compE\<^sub>1_None[simp]:
- "\<And>Vs. \<A> e = None \<Longrightarrow> \<A> (compE\<^sub>1 Vs e) = None"
-and "\<And>Vs. \<A>s es = None \<Longrightarrow> \<A>s (compEs\<^sub>1 Vs es) = None"
-(*<*)by(induct e and es rule: compE\<^sub>1.induct compEs\<^sub>1.induct)(auto simp:hyperset_defs)(*>*)
-
-
-lemma A_compE\<^sub>1:
- "\<And>A Vs. \<lbrakk> \<A> e = \<lfloor>A\<rfloor>; fv e \<subseteq> set Vs \<rbrakk> \<Longrightarrow> \<A> (compE\<^sub>1 Vs e) = \<lfloor>last_index Vs ` A\<rfloor>"
-and "\<And>A Vs. \<lbrakk> \<A>s es = \<lfloor>A\<rfloor>; fvs es \<subseteq> set Vs \<rbrakk> \<Longrightarrow> \<A>s (compEs\<^sub>1 Vs es) = \<lfloor>last_index Vs ` A\<rfloor>"
-(*<*)
-proof(induct e and es rule: fv.induct fvs.induct)
- case (Block V' T e)
- hence "fv e \<subseteq> set (Vs@[V'])" by fastforce
- moreover obtain B where "\<A> e = \<lfloor>B\<rfloor>"
- using Block.prems by(simp add: hyperset_defs)
- moreover from calculation have "B \<subseteq> set (Vs@[V'])" by(auto dest!:A_fv)
- ultimately show ?case using Block
- by(auto simp add: hyperset_defs image_last_index last_index_size_conv)
-next
- case (TryCatch e\<^sub>1 C V' e\<^sub>2)
- hence fve\<^sub>2: "fv e\<^sub>2 \<subseteq> set (Vs@[V'])" by auto
- show ?case
- proof (cases "\<A> e\<^sub>1")
- assume A\<^sub>1: "\<A> e\<^sub>1 = None"
- then obtain A\<^sub>2 where A\<^sub>2: "\<A> e\<^sub>2 = \<lfloor>A\<^sub>2\<rfloor>" using TryCatch
- by(simp add:hyperset_defs)
- hence "A\<^sub>2 \<subseteq> set (Vs@[V'])" using TryCatch.prems A_fv[OF A\<^sub>2] by simp blast
- thus ?thesis using TryCatch fve\<^sub>2 A\<^sub>1 A\<^sub>2
- by(auto simp add:hyperset_defs image_last_index last_index_size_conv)
- next
- fix A\<^sub>1 assume A\<^sub>1: "\<A> e\<^sub>1 = \<lfloor>A\<^sub>1\<rfloor>"
- show ?thesis
- proof (cases "\<A> e\<^sub>2")
- assume A\<^sub>2: "\<A> e\<^sub>2 = None"
- then show ?case using TryCatch A\<^sub>1 by(simp add:hyperset_defs)
- next
- fix A\<^sub>2 assume A\<^sub>2: "\<A> e\<^sub>2 = \<lfloor>A\<^sub>2\<rfloor>"
- have "A\<^sub>1 \<subseteq> set Vs" using TryCatch.prems A_fv[OF A\<^sub>1] by simp blast
- moreover
- have "A\<^sub>2 \<subseteq> set (Vs@[V'])" using TryCatch.prems A_fv[OF A\<^sub>2] by simp blast
- ultimately show ?thesis using TryCatch A\<^sub>1 A\<^sub>2
- by (auto simp add: Diff_subset_conv last_index_size_conv subsetD hyperset_defs
- dest!: sym [of _ A])
- qed
- qed
-next
- case (Cond e e\<^sub>1 e\<^sub>2)
- { assume "\<A> e = None \<or> \<A> e\<^sub>1 = None \<or> \<A> e\<^sub>2 = None"
- hence ?case using Cond by(auto simp add:hyperset_defs image_Un)
- }
- moreover
- { fix A A\<^sub>1 A\<^sub>2
- assume "\<A> e = \<lfloor>A\<rfloor>" and A\<^sub>1: "\<A> e\<^sub>1 = \<lfloor>A\<^sub>1\<rfloor>" and A\<^sub>2: "\<A> e\<^sub>2 = \<lfloor>A\<^sub>2\<rfloor>"
- moreover
- have "A\<^sub>1 \<subseteq> set Vs" using Cond.prems A_fv[OF A\<^sub>1] by simp blast
- moreover
- have "A\<^sub>2 \<subseteq> set Vs" using Cond.prems A_fv[OF A\<^sub>2] by simp blast
- ultimately have ?case using Cond
- by(auto simp add:hyperset_defs image_Un
- inj_on_image_Int[OF inj_on_last_index])
- }
- ultimately show ?case by fastforce
-qed (auto simp add:hyperset_defs)
-(*>*)
-
-
-lemma D_None[iff]: "\<D> (e::'a exp) None" and [iff]: "\<D>s (es::'a exp list) None"
-(*<*)by(induct e and es rule: \<D>.induct \<D>s.induct)(simp_all)(*>*)
-
-
-lemma D_last_index_compE\<^sub>1:
- "\<And>A Vs. \<lbrakk> A \<subseteq> set Vs; fv e \<subseteq> set Vs \<rbrakk> \<Longrightarrow>
- \<D> e \<lfloor>A\<rfloor> \<Longrightarrow> \<D> (compE\<^sub>1 Vs e) \<lfloor>last_index Vs ` A\<rfloor>"
-and "\<And>A Vs. \<lbrakk> A \<subseteq> set Vs; fvs es \<subseteq> set Vs \<rbrakk> \<Longrightarrow>
- \<D>s es \<lfloor>A\<rfloor> \<Longrightarrow> \<D>s (compEs\<^sub>1 Vs es) \<lfloor>last_index Vs ` A\<rfloor>"
-(*<*)
-proof(induct e and es rule: \<D>.induct \<D>s.induct)
- case (BinOp e\<^sub>1 bop e\<^sub>2)
- hence IH\<^sub>1: "\<D> (compE\<^sub>1 Vs e\<^sub>1) \<lfloor>last_index Vs ` A\<rfloor>" by simp
- show ?case
- proof (cases "\<A> e\<^sub>1")
- case None thus ?thesis using BinOp by simp
- next
- case (Some A\<^sub>1)
- have indexA\<^sub>1: "\<A> (compE\<^sub>1 Vs e\<^sub>1) = \<lfloor>last_index Vs ` A\<^sub>1\<rfloor>"
- using A_compE\<^sub>1[OF Some] BinOp.prems by auto
- have "A \<union> A\<^sub>1 \<subseteq> set Vs" using BinOp.prems A_fv[OF Some] by auto
- hence "\<D> (compE\<^sub>1 Vs e\<^sub>2) \<lfloor>last_index Vs ` (A \<union> A\<^sub>1)\<rfloor>" using BinOp Some by auto
- hence "\<D> (compE\<^sub>1 Vs e\<^sub>2) \<lfloor>last_index Vs ` A \<union> last_index Vs ` A\<^sub>1\<rfloor>"
- by(simp add: image_Un)
- thus ?thesis using IH\<^sub>1 indexA\<^sub>1 by auto
- qed
-next
- case (FAss e\<^sub>1 F D e\<^sub>2)
- hence IH\<^sub>1: "\<D> (compE\<^sub>1 Vs e\<^sub>1) \<lfloor>last_index Vs ` A\<rfloor>" by simp
- show ?case
- proof (cases "\<A> e\<^sub>1")
- case None thus ?thesis using FAss by simp
- next
- case (Some A\<^sub>1)
- have indexA\<^sub>1: "\<A> (compE\<^sub>1 Vs e\<^sub>1) = \<lfloor>last_index Vs ` A\<^sub>1\<rfloor>"
- using A_compE\<^sub>1[OF Some] FAss.prems by auto
- have "A \<union> A\<^sub>1 \<subseteq> set Vs" using FAss.prems A_fv[OF Some] by auto
- hence "\<D> (compE\<^sub>1 Vs e\<^sub>2) \<lfloor>last_index Vs ` (A \<union> A\<^sub>1)\<rfloor>" using FAss Some by auto
- hence "\<D> (compE\<^sub>1 Vs e\<^sub>2) \<lfloor>last_index Vs ` A \<union> last_index Vs ` A\<^sub>1\<rfloor>"
- by(simp add: image_Un)
- thus ?thesis using IH\<^sub>1 indexA\<^sub>1 by auto
- qed
-next
- case (Call e\<^sub>1 M es)
- hence IH\<^sub>1: "\<D> (compE\<^sub>1 Vs e\<^sub>1) \<lfloor>last_index Vs ` A\<rfloor>" by simp
- show ?case
- proof (cases "\<A> e\<^sub>1")
- case None thus ?thesis using Call by simp
- next
- case (Some A\<^sub>1)
- have indexA\<^sub>1: "\<A> (compE\<^sub>1 Vs e\<^sub>1) = \<lfloor>last_index Vs ` A\<^sub>1\<rfloor>"
- using A_compE\<^sub>1[OF Some] Call.prems by auto
- have "A \<union> A\<^sub>1 \<subseteq> set Vs" using Call.prems A_fv[OF Some] by auto
- hence "\<D>s (compEs\<^sub>1 Vs es) \<lfloor>last_index Vs ` (A \<union> A\<^sub>1)\<rfloor>" using Call Some by auto
- hence "\<D>s (compEs\<^sub>1 Vs es) \<lfloor>last_index Vs ` A \<union> last_index Vs ` A\<^sub>1\<rfloor>"
- by(simp add: image_Un)
- thus ?thesis using IH\<^sub>1 indexA\<^sub>1 by auto
- qed
-next
- case (TryCatch e\<^sub>1 C V e\<^sub>2)
- have "\<lbrakk> A\<union>{V} \<subseteq> set(Vs@[V]); fv e\<^sub>2 \<subseteq> set(Vs@[V]); \<D> e\<^sub>2 \<lfloor>A\<union>{V}\<rfloor>\<rbrakk> \<Longrightarrow>
- \<D> (compE\<^sub>1 (Vs@[V]) e\<^sub>2) \<lfloor>last_index (Vs@[V]) ` (A\<union>{V})\<rfloor>" by fact
- hence "\<D> (compE\<^sub>1 (Vs@[V]) e\<^sub>2) \<lfloor>last_index (Vs@[V]) ` (A\<union>{V})\<rfloor>"
- using TryCatch.prems by(simp add:Diff_subset_conv)
- moreover have "last_index (Vs@[V]) ` A \<subseteq> last_index Vs ` A \<union> {size Vs}"
- using TryCatch.prems by(auto simp add: image_last_index split:if_split_asm)
- ultimately show ?case using TryCatch
- by(auto simp:hyperset_defs elim!:D_mono')
-next
- case (Seq e\<^sub>1 e\<^sub>2)
- hence IH\<^sub>1: "\<D> (compE\<^sub>1 Vs e\<^sub>1) \<lfloor>last_index Vs ` A\<rfloor>" by simp
- show ?case
- proof (cases "\<A> e\<^sub>1")
- case None thus ?thesis using Seq by simp
- next
- case (Some A\<^sub>1)
- have indexA\<^sub>1: "\<A> (compE\<^sub>1 Vs e\<^sub>1) = \<lfloor>last_index Vs ` A\<^sub>1\<rfloor>"
- using A_compE\<^sub>1[OF Some] Seq.prems by auto
- have "A \<union> A\<^sub>1 \<subseteq> set Vs" using Seq.prems A_fv[OF Some] by auto
- hence "\<D> (compE\<^sub>1 Vs e\<^sub>2) \<lfloor>last_index Vs ` (A \<union> A\<^sub>1)\<rfloor>" using Seq Some by auto
- hence "\<D> (compE\<^sub>1 Vs e\<^sub>2) \<lfloor>last_index Vs ` A \<union> last_index Vs ` A\<^sub>1\<rfloor>"
- by(simp add: image_Un)
- thus ?thesis using IH\<^sub>1 indexA\<^sub>1 by auto
- qed
-next
- case (Cond e e\<^sub>1 e\<^sub>2)
- hence IH\<^sub>1: "\<D> (compE\<^sub>1 Vs e) \<lfloor>last_index Vs ` A\<rfloor>" by simp
- show ?case
- proof (cases "\<A> e")
- case None thus ?thesis using Cond by simp
- next
- case (Some B)
- have indexB: "\<A> (compE\<^sub>1 Vs e) = \<lfloor>last_index Vs ` B\<rfloor>"
- using A_compE\<^sub>1[OF Some] Cond.prems by auto
- have "A \<union> B \<subseteq> set Vs" using Cond.prems A_fv[OF Some] by auto
- hence "\<D> (compE\<^sub>1 Vs e\<^sub>1) \<lfloor>last_index Vs ` (A \<union> B)\<rfloor>"
- and "\<D> (compE\<^sub>1 Vs e\<^sub>2) \<lfloor>last_index Vs ` (A \<union> B)\<rfloor>"
- using Cond Some by auto
- hence "\<D> (compE\<^sub>1 Vs e\<^sub>1) \<lfloor>last_index Vs ` A \<union> last_index Vs ` B\<rfloor>"
- and "\<D> (compE\<^sub>1 Vs e\<^sub>2) \<lfloor>last_index Vs ` A \<union> last_index Vs ` B\<rfloor>"
- by(simp add: image_Un)+
- thus ?thesis using IH\<^sub>1 indexB by auto
- qed
-next
- case (While e\<^sub>1 e\<^sub>2)
- hence IH\<^sub>1: "\<D> (compE\<^sub>1 Vs e\<^sub>1) \<lfloor>last_index Vs ` A\<rfloor>" by simp
- show ?case
- proof (cases "\<A> e\<^sub>1")
- case None thus ?thesis using While by simp
- next
- case (Some A\<^sub>1)
- have indexA\<^sub>1: "\<A> (compE\<^sub>1 Vs e\<^sub>1) = \<lfloor>last_index Vs ` A\<^sub>1\<rfloor>"
- using A_compE\<^sub>1[OF Some] While.prems by auto
- have "A \<union> A\<^sub>1 \<subseteq> set Vs" using While.prems A_fv[OF Some] by auto
- hence "\<D> (compE\<^sub>1 Vs e\<^sub>2) \<lfloor>last_index Vs ` (A \<union> A\<^sub>1)\<rfloor>" using While Some by auto
- hence "\<D> (compE\<^sub>1 Vs e\<^sub>2) \<lfloor>last_index Vs ` A \<union> last_index Vs ` A\<^sub>1\<rfloor>"
- by(simp add: image_Un)
- thus ?thesis using IH\<^sub>1 indexA\<^sub>1 by auto
- qed
-next
- case (Block V T e)
- have "\<lbrakk> A-{V} \<subseteq> set(Vs@[V]); fv e \<subseteq> set(Vs@[V]); \<D> e \<lfloor>A-{V}\<rfloor> \<rbrakk> \<Longrightarrow>
- \<D> (compE\<^sub>1 (Vs@[V]) e) \<lfloor>last_index (Vs@[V]) ` (A-{V})\<rfloor>" by fact
- hence "\<D> (compE\<^sub>1 (Vs@[V]) e) \<lfloor>last_index (Vs@[V]) ` (A-{V})\<rfloor>"
- using Block.prems by(simp add:Diff_subset_conv)
- moreover have "size Vs \<notin> last_index Vs ` A"
- using Block.prems by(auto simp add:image_def size_last_index_conv)
- ultimately show ?case using Block
- by(auto simp add: image_last_index Diff_subset_conv hyperset_defs elim!: D_mono')
-next
- case (Cons_exp e\<^sub>1 es)
- hence IH\<^sub>1: "\<D> (compE\<^sub>1 Vs e\<^sub>1) \<lfloor>last_index Vs ` A\<rfloor>" by simp
- show ?case
- proof (cases "\<A> e\<^sub>1")
- case None thus ?thesis using Cons_exp by simp
- next
- case (Some A\<^sub>1)
- have indexA\<^sub>1: "\<A> (compE\<^sub>1 Vs e\<^sub>1) = \<lfloor>last_index Vs ` A\<^sub>1\<rfloor>"
- using A_compE\<^sub>1[OF Some] Cons_exp.prems by auto
- have "A \<union> A\<^sub>1 \<subseteq> set Vs" using Cons_exp.prems A_fv[OF Some] by auto
- hence "\<D>s (compEs\<^sub>1 Vs es) \<lfloor>last_index Vs ` (A \<union> A\<^sub>1)\<rfloor>" using Cons_exp Some by auto
- hence "\<D>s (compEs\<^sub>1 Vs es) \<lfloor>last_index Vs ` A \<union> last_index Vs ` A\<^sub>1\<rfloor>"
- by(simp add: image_Un)
- thus ?thesis using IH\<^sub>1 indexA\<^sub>1 by auto
- qed
-qed (simp_all add:hyperset_defs)
-(*>*)
-
-
-lemma last_index_image_set: "distinct xs \<Longrightarrow> last_index xs ` set xs = {..<size xs}"
-(*<*)by(induct xs rule:rev_induct) (auto simp add: image_last_index)(*>*)
-
-
-lemma D_compE\<^sub>1:
- "\<lbrakk> \<D> e \<lfloor>set Vs\<rfloor>; fv e \<subseteq> set Vs; distinct Vs \<rbrakk> \<Longrightarrow> \<D> (compE\<^sub>1 Vs e) \<lfloor>{..<length Vs}\<rfloor>"
-(*<*)by(fastforce dest!: D_last_index_compE\<^sub>1[OF subset_refl] simp add:last_index_image_set)(*>*)
-
-
-lemma D_compE\<^sub>1':
-assumes "\<D> e \<lfloor>set(V#Vs)\<rfloor>" and "fv e \<subseteq> set(V#Vs)" and "distinct(V#Vs)"
-shows "\<D> (compE\<^sub>1 (V#Vs) e) \<lfloor>{..length Vs}\<rfloor>"
-(*<*)
-proof -
- have "{..size Vs} = {..<size(V#Vs)}" by auto
- thus ?thesis using assms by (simp only:)(rule D_compE\<^sub>1)
-qed
-(*>*)
-
-lemma compP\<^sub>1_pres_wf: "wf_J_prog P \<Longrightarrow> wf_J\<^sub>1_prog (compP\<^sub>1 P)"
-(*<*)
-proof -
- assume wf: "wf_J_prog P"
- let ?f = "(\<lambda>b (pns, body).
- compE\<^sub>1 (case b of Static \<Rightarrow> pns | NonStatic \<Rightarrow> this # pns) body)"
- let ?wf\<^sub>2 = "wf_J\<^sub>1_mdecl"
-
- { fix C M b Ts T m
- assume cM: "P \<turnstile> C sees M, b : Ts\<rightarrow>T = m in C"
- and wfm: "wf_mdecl wf_J_mdecl P C (M, b, Ts, T, m)"
- obtain pns body where [simp]: "m = (pns, body)" by(cases m) simp
- let ?E = "\<lambda>b. case b of Static \<Rightarrow> [pns [\<mapsto>] Ts] | NonStatic \<Rightarrow> [pns [\<mapsto>] Ts, this \<mapsto> Class C]"
- obtain T' where WT: "P,?E b \<turnstile> body :: T'" and subT: "P \<turnstile> T' \<le> T"
- using wfm by(cases b) (auto simp: wf_mdecl_def)
- have fv: "fv body \<subseteq> dom (?E b)" by(rule WT_fv[OF WT])
- have "wf_mdecl ?wf\<^sub>2 (compP ?f P) C (M, b, Ts, T, ?f b m)"
- proof(cases b)
- case Static then show ?thesis using cM wfm fv
- by(auto simp:wf_mdecl_def wf_J\<^sub>1_mdecl_def
- intro!: compE\<^sub>1_pres_wt D_compE\<^sub>1 \<B>)
- next
- case NonStatic then show ?thesis using cM wfm fv
- by(clarsimp simp add:wf_mdecl_def wf_J\<^sub>1_mdecl_def)
- (fastforce intro!: compE\<^sub>1_pres_wt D_compE\<^sub>1' \<B>)
- qed
- }
- then show ?thesis by simp (rule wf_prog_compPI[OF _ wf])
-qed
-(*>*)
-
-
-end
+(* Title: JinjaDCI/Compiler/Correctness1.thy
+ Author: Tobias Nipkow, Susannah Mansky
+ Copyright TUM 2003, UIUC 2019-20
+
+ Based on the Jinja theory Compiler/Correctness1.thy by Tobias Nipkow
+*)
+
+section \<open> Correctness of Stage 1 \<close>
+
+theory Correctness1
+imports J1WellForm Compiler1
+begin
+
+subsection\<open>Correctness of program compilation \<close>
+
+primrec unmod :: "expr\<^sub>1 \<Rightarrow> nat \<Rightarrow> bool"
+ and unmods :: "expr\<^sub>1 list \<Rightarrow> nat \<Rightarrow> bool" where
+"unmod (new C) i = True" |
+"unmod (Cast C e) i = unmod e i" |
+"unmod (Val v) i = True" |
+"unmod (e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2) i = (unmod e\<^sub>1 i \<and> unmod e\<^sub>2 i)" |
+"unmod (Var i) j = True" |
+"unmod (i:=e) j = (i \<noteq> j \<and> unmod e j)" |
+"unmod (e\<bullet>F{D}) i = unmod e i" |
+"unmod (C\<bullet>\<^sub>sF{D}) i = True" |
+"unmod (e\<^sub>1\<bullet>F{D}:=e\<^sub>2) i = (unmod e\<^sub>1 i \<and> unmod e\<^sub>2 i)" |
+"unmod (C\<bullet>\<^sub>sF{D}:=e\<^sub>2) i = unmod e\<^sub>2 i" |
+"unmod (e\<bullet>M(es)) i = (unmod e i \<and> unmods es i)" |
+"unmod (C\<bullet>\<^sub>sM(es)) i = unmods es i" |
+"unmod {j:T; e} i = unmod e i" |
+"unmod (e\<^sub>1;;e\<^sub>2) i = (unmod e\<^sub>1 i \<and> unmod e\<^sub>2 i)" |
+"unmod (if (e) e\<^sub>1 else e\<^sub>2) i = (unmod e i \<and> unmod e\<^sub>1 i \<and> unmod e\<^sub>2 i)" |
+"unmod (while (e) c) i = (unmod e i \<and> unmod c i)" |
+"unmod (throw e) i = unmod e i" |
+"unmod (try e\<^sub>1 catch(C i) e\<^sub>2) j = (unmod e\<^sub>1 j \<and> (if i=j then False else unmod e\<^sub>2 j))" |
+"unmod (INIT C (Cs,b) \<leftarrow> e) i = unmod e i" |
+"unmod (RI(C,e);Cs \<leftarrow> e') i = (unmod e i \<and> unmod e' i)" |
+
+"unmods ([]) i = True" |
+"unmods (e#es) i = (unmod e i \<and> unmods es i)"
+
+lemma hidden_unmod: "\<And>Vs. hidden Vs i \<Longrightarrow> unmod (compE\<^sub>1 Vs e) i" and
+ "\<And>Vs. hidden Vs i \<Longrightarrow> unmods (compEs\<^sub>1 Vs es) i"
+(*<*)
+proof(induct e and es rule: compE\<^sub>1.induct compEs\<^sub>1.induct)
+ case TryCatch
+ then show ?case by(simp add:hidden_inacc)(auto simp add:hidden_def)
+qed (simp_all add:hidden_inacc)
+(*>*)
+
+
+lemma eval\<^sub>1_preserves_unmod:
+ "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e,(h,ls,sh)\<rangle> \<Rightarrow> \<langle>e',(h',ls',sh')\<rangle>; unmod e i; i < size ls \<rbrakk>
+ \<Longrightarrow> ls ! i = ls' ! i"
+and "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>es,(h,ls,sh)\<rangle> [\<Rightarrow>] \<langle>es',(h',ls',sh')\<rangle>; unmods es i; i < size ls \<rbrakk>
+ \<Longrightarrow> ls ! i = ls' ! i"
+(*<*)
+proof(induct rule:eval\<^sub>1_evals\<^sub>1_inducts)
+ case (RInitInitFail\<^sub>1 e h l sh a h' l' sh' C sfs i sh'' D Cs e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1)
+ have "final (throw a)" using eval\<^sub>1_final[OF RInitInitFail\<^sub>1.hyps(1)] by simp
+ then show ?case using RInitInitFail\<^sub>1 by(auto simp: eval\<^sub>1_preserves_len)
+qed(auto dest!:eval\<^sub>1_preserves_len evals\<^sub>1_preserves_len split:if_split_asm)
+(*>*)
+
+
+lemma LAss_lem:
+ "\<lbrakk>x \<in> set xs; size xs \<le> size ys \<rbrakk>
+ \<Longrightarrow> m\<^sub>1 \<subseteq>\<^sub>m m\<^sub>2(xs[\<mapsto>]ys) \<Longrightarrow> m\<^sub>1(x\<mapsto>y) \<subseteq>\<^sub>m m\<^sub>2(xs[\<mapsto>]ys[last_index xs x := y])"
+(*<*)
+by(simp add:map_le_def fun_upds_apply eq_sym_conv)
+(*>*)
+lemma Block_lem:
+fixes l :: "'a \<rightharpoonup> 'b"
+assumes 0: "l \<subseteq>\<^sub>m [Vs [\<mapsto>] ls]"
+ and 1: "l' \<subseteq>\<^sub>m [Vs [\<mapsto>] ls', V\<mapsto>v]"
+ and hidden: "V \<in> set Vs \<Longrightarrow> ls ! last_index Vs V = ls' ! last_index Vs V"
+ and size: "size ls = size ls'" "size Vs < size ls'"
+shows "l'(V := l V) \<subseteq>\<^sub>m [Vs [\<mapsto>] ls']"
+(*<*)
+proof -
+ have "l'(V := l V) \<subseteq>\<^sub>m [Vs [\<mapsto>] ls', V\<mapsto>v](V := l V)"
+ using 1 by(rule map_le_upd)
+ also have "\<dots> = [Vs [\<mapsto>] ls'](V := l V)" by simp
+ also have "\<dots> \<subseteq>\<^sub>m [Vs [\<mapsto>] ls']"
+ proof (cases "l V")
+ case None thus ?thesis by simp
+ next
+ case (Some w)
+ hence "[Vs [\<mapsto>] ls] V = Some w"
+ using 0 by(force simp add: map_le_def split:if_splits)
+ hence VinVs: "V \<in> set Vs" and w: "w = ls ! last_index Vs V"
+ using size by(auto simp add:fun_upds_apply split:if_splits)
+ hence "w = ls' ! last_index Vs V" using hidden[OF VinVs] by simp
+ hence "[Vs [\<mapsto>] ls'](V := l V) = [Vs [\<mapsto>] ls']" using Some size VinVs
+ by(simp add: map_upds_upd_conv_last_index)
+ thus ?thesis by simp
+ qed
+ finally show ?thesis .
+qed
+(*>*)
+
+(*<*)
+declare fun_upd_apply[simp del]
+(*>*)
+
+
+text\<open>\noindent The main theorem: \<close>
+
+theorem assumes wf: "wwf_J_prog P"
+shows eval\<^sub>1_eval: "P \<turnstile> \<langle>e,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',(h',l',sh')\<rangle>
+ \<Longrightarrow> (\<And>Vs ls. \<lbrakk> fv e \<subseteq> set Vs; l \<subseteq>\<^sub>m [Vs[\<mapsto>]ls]; size Vs + max_vars e \<le> size ls \<rbrakk>
+ \<Longrightarrow> \<exists>ls'. compP\<^sub>1 P \<turnstile>\<^sub>1 \<langle>compE\<^sub>1 Vs e,(h,ls,sh)\<rangle> \<Rightarrow> \<langle>fin\<^sub>1 e',(h',ls',sh')\<rangle> \<and> l' \<subseteq>\<^sub>m [Vs[\<mapsto>]ls'])"
+(*<*)
+ (is "_ \<Longrightarrow> (\<And>Vs ls. PROP ?P e h l sh e' h' l' sh' Vs ls)"
+ is "_ \<Longrightarrow> (\<And>Vs ls. \<lbrakk> _; _; _ \<rbrakk> \<Longrightarrow> \<exists>ls'. ?Post e h l sh e' h' l' sh' Vs ls ls')")
+(*>*)
+
+and evals\<^sub>1_evals: "P \<turnstile> \<langle>es,(h,l,sh)\<rangle> [\<Rightarrow>] \<langle>es',(h',l',sh')\<rangle>
+ \<Longrightarrow> (\<And>Vs ls. \<lbrakk> fvs es \<subseteq> set Vs; l \<subseteq>\<^sub>m [Vs[\<mapsto>]ls]; size Vs + max_varss es \<le> size ls \<rbrakk>
+ \<Longrightarrow> \<exists>ls'. compP\<^sub>1 P \<turnstile>\<^sub>1 \<langle>compEs\<^sub>1 Vs es,(h,ls,sh)\<rangle> [\<Rightarrow>] \<langle>compEs\<^sub>1 Vs es',(h',ls',sh')\<rangle> \<and>
+ l' \<subseteq>\<^sub>m [Vs[\<mapsto>]ls'])"
+(*<*)
+ (is "_ \<Longrightarrow> (\<And>Vs ls. PROP ?Ps es h l sh es' h' l' sh' Vs ls)"
+ is "_ \<Longrightarrow> (\<And>Vs ls. \<lbrakk> _; _; _\<rbrakk> \<Longrightarrow> \<exists>ls'. ?Posts es h l sh es' h' l' sh' Vs ls ls')")
+proof (induct rule:eval_evals_inducts)
+ case Nil thus ?case by(fastforce intro!:Nil\<^sub>1)
+next
+ case (Cons e h l sh v h' l' sh' es es' h\<^sub>2 l\<^sub>2 sh\<^sub>2)
+ have "PROP ?P e h l sh (Val v) h' l' sh' Vs ls" by fact
+ with Cons.prems
+ obtain ls' where 1: "?Post e h l sh (Val v) h' l' sh' Vs ls ls'"
+ "size ls = size ls'" by(auto intro!:eval\<^sub>1_preserves_len)
+ have "PROP ?Ps es h' l' sh' es' h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls'" by fact
+ with 1 Cons.prems
+ obtain ls\<^sub>2 where 2: "?Posts es h' l' sh' es' h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls' ls\<^sub>2" by(auto)
+ from 1 2 Cons show ?case by(auto intro!:Cons\<^sub>1)
+next
+ case ConsThrow thus ?case
+ by(fastforce intro!:ConsThrow\<^sub>1 dest: eval_final)
+next
+ case (Block e h l V sh e' h' l' sh' T)
+ let ?Vs = "Vs @ [V]"
+ have IH:
+ "\<lbrakk>fv e \<subseteq> set ?Vs; l(V := None) \<subseteq>\<^sub>m [?Vs [\<mapsto>] ls];
+ size ?Vs + max_vars e \<le> size ls\<rbrakk>
+ \<Longrightarrow> \<exists>ls'. compP\<^sub>1 P \<turnstile>\<^sub>1 \<langle>compE\<^sub>1 ?Vs e,(h,ls,sh)\<rangle> \<Rightarrow> \<langle>fin\<^sub>1 e',(h', ls',sh')\<rangle> \<and>
+ l' \<subseteq>\<^sub>m [?Vs [\<mapsto>] ls']" and
+ fv: "fv {V:T; e} \<subseteq> set Vs" and rel: "l \<subseteq>\<^sub>m [Vs [\<mapsto>] ls]" and
+ len: "length Vs + max_vars {V:T; e} \<le> length ls" by fact+
+ have len': "length Vs < length ls" using len by auto
+ have "fv e \<subseteq> set ?Vs" using fv by auto
+ moreover have "l(V := None) \<subseteq>\<^sub>m [?Vs [\<mapsto>] ls]" using rel len' by simp
+ moreover have "size ?Vs + max_vars e \<le> size ls" using len by simp
+ ultimately obtain ls' where
+ 1: "compP\<^sub>1 P \<turnstile>\<^sub>1 \<langle>compE\<^sub>1 ?Vs e,(h,ls,sh)\<rangle> \<Rightarrow> \<langle>fin\<^sub>1 e',(h',ls',sh')\<rangle>"
+ and rel': "l' \<subseteq>\<^sub>m [?Vs [\<mapsto>] ls']" using IH by blast
+ have [simp]: "length ls = length ls'" by(rule eval\<^sub>1_preserves_len[OF 1])
+ show "\<exists>ls'. compP\<^sub>1 P \<turnstile>\<^sub>1 \<langle>compE\<^sub>1 Vs {V:T; e},(h,ls,sh)\<rangle> \<Rightarrow> \<langle>fin\<^sub>1 e',(h',ls',sh')\<rangle>
+ \<and> l'(V := l V) \<subseteq>\<^sub>m [Vs [\<mapsto>] ls']" (is "\<exists>ls'. ?R ls'")
+ proof
+ show "?R ls'"
+ proof
+ show "compP\<^sub>1 P \<turnstile>\<^sub>1 \<langle>compE\<^sub>1 Vs {V:T; e},(h,ls,sh)\<rangle> \<Rightarrow> \<langle>fin\<^sub>1 e',(h',ls',sh')\<rangle>"
+ using 1 by(simp add:Block\<^sub>1)
+ next
+ show "l'(V := l V) \<subseteq>\<^sub>m [Vs [\<mapsto>] ls']"
+ proof -
+ have "l' \<subseteq>\<^sub>m [Vs [\<mapsto>] ls', V \<mapsto> ls' ! length Vs]"
+ using len' rel' by simp
+ moreover
+ { assume VinVs: "V \<in> set Vs"
+ hence "hidden (Vs @ [V]) (last_index Vs V)"
+ by(rule hidden_last_index)
+ hence "unmod (compE\<^sub>1 (Vs @ [V]) e) (last_index Vs V)"
+ by(rule hidden_unmod)
+ moreover have "last_index Vs V < length ls"
+ using len' VinVs by simp
+ ultimately have "ls ! last_index Vs V = ls' ! last_index Vs V"
+ by(rule eval\<^sub>1_preserves_unmod[OF 1])
+ }
+ ultimately show ?thesis using Block_lem[OF rel] len' by auto
+ qed
+ qed
+ qed
+next
+ case (TryThrow e' h l sh a h' l' sh' D fs C V e\<^sub>2)
+ have "PROP ?P e' h l sh (Throw a) h' l' sh' Vs ls" by fact
+ with TryThrow.prems
+ obtain ls' where 1: "?Post e' h l sh (Throw a) h' l' sh' Vs ls ls'" by(auto)
+ show ?case using 1 TryThrow.hyps by(auto intro!:eval\<^sub>1_evals\<^sub>1.TryThrow\<^sub>1)
+next
+ case (TryCatch e\<^sub>1 h l sh a h\<^sub>1 l\<^sub>1 sh\<^sub>1 D fs C e\<^sub>2 V e' h\<^sub>2 l\<^sub>2 sh\<^sub>2)
+ let ?e = "try e\<^sub>1 catch(C V) e\<^sub>2"
+ have IH\<^sub>1: "\<lbrakk>fv e\<^sub>1 \<subseteq> set Vs; l \<subseteq>\<^sub>m [Vs [\<mapsto>] ls];
+ size Vs + max_vars e\<^sub>1 \<le> length ls\<rbrakk>
+ \<Longrightarrow> \<exists>ls\<^sub>1. compP\<^sub>1 P \<turnstile>\<^sub>1 \<langle>compE\<^sub>1 Vs e\<^sub>1,(h,ls,sh)\<rangle> \<Rightarrow>
+ \<langle>fin\<^sub>1 (Throw a),(h\<^sub>1,ls\<^sub>1,sh\<^sub>1)\<rangle> \<and>
+ l\<^sub>1 \<subseteq>\<^sub>m [Vs [\<mapsto>] ls\<^sub>1]" and
+ fv: "fv ?e \<subseteq> set Vs" and
+ rel: "l \<subseteq>\<^sub>m [Vs [\<mapsto>] ls]" and
+ len: "length Vs + max_vars ?e \<le> length ls" by fact+
+ have "fv e\<^sub>1 \<subseteq> set Vs" using fv by auto
+ moreover have "length Vs + max_vars e\<^sub>1 \<le> length ls" using len by(auto)
+ ultimately obtain ls\<^sub>1 where
+ 1: "compP\<^sub>1 P \<turnstile>\<^sub>1 \<langle>compE\<^sub>1 Vs e\<^sub>1,(h,ls,sh)\<rangle> \<Rightarrow> \<langle>Throw a,(h\<^sub>1,ls\<^sub>1,sh\<^sub>1)\<rangle>"
+ and rel\<^sub>1: "l\<^sub>1 \<subseteq>\<^sub>m [Vs [\<mapsto>] ls\<^sub>1]" using IH\<^sub>1 rel by fastforce
+ from 1 have [simp]: "size ls = size ls\<^sub>1" by(rule eval\<^sub>1_preserves_len)
+ let ?Vs = "Vs @ [V]" let ?ls = "(ls\<^sub>1[size Vs:=Addr a])"
+ have IH\<^sub>2: "\<lbrakk>fv e\<^sub>2 \<subseteq> set ?Vs; l\<^sub>1(V \<mapsto> Addr a) \<subseteq>\<^sub>m [?Vs [\<mapsto>] ?ls];
+ length ?Vs + max_vars e\<^sub>2 \<le> length ?ls\<rbrakk> \<Longrightarrow> \<exists>ls\<^sub>2.
+ compP\<^sub>1 P \<turnstile>\<^sub>1 \<langle>compE\<^sub>1 ?Vs e\<^sub>2,(h\<^sub>1,?ls,sh\<^sub>1)\<rangle> \<Rightarrow> \<langle>fin\<^sub>1 e',(h\<^sub>2, ls\<^sub>2, sh\<^sub>2)\<rangle> \<and>
+ l\<^sub>2 \<subseteq>\<^sub>m [?Vs [\<mapsto>] ls\<^sub>2]" by fact
+ have len\<^sub>1: "size Vs < size ls\<^sub>1" using len by(auto)
+ have "fv e\<^sub>2 \<subseteq> set ?Vs" using fv by auto
+ moreover have "l\<^sub>1(V \<mapsto> Addr a) \<subseteq>\<^sub>m [?Vs [\<mapsto>] ?ls]" using rel\<^sub>1 len\<^sub>1 by simp
+ moreover have "length ?Vs + max_vars e\<^sub>2 \<le> length ?ls" using len by(auto)
+ ultimately obtain ls\<^sub>2 where
+ 2: "compP\<^sub>1 P \<turnstile>\<^sub>1 \<langle>compE\<^sub>1 ?Vs e\<^sub>2,(h\<^sub>1,?ls,sh\<^sub>1)\<rangle> \<Rightarrow> \<langle>fin\<^sub>1 e',(h\<^sub>2, ls\<^sub>2, sh\<^sub>2)\<rangle>"
+ and rel\<^sub>2: "l\<^sub>2 \<subseteq>\<^sub>m [?Vs [\<mapsto>] ls\<^sub>2]" using IH\<^sub>2 by blast
+ from 2 have [simp]: "size ls\<^sub>1 = size ls\<^sub>2"
+ by(fastforce dest: eval\<^sub>1_preserves_len)
+ show "\<exists>ls\<^sub>2. compP\<^sub>1 P \<turnstile>\<^sub>1 \<langle>compE\<^sub>1 Vs ?e,(h,ls,sh)\<rangle> \<Rightarrow> \<langle>fin\<^sub>1 e',(h\<^sub>2,ls\<^sub>2,sh\<^sub>2)\<rangle> \<and>
+ l\<^sub>2(V := l\<^sub>1 V) \<subseteq>\<^sub>m [Vs [\<mapsto>] ls\<^sub>2]" (is "\<exists>ls\<^sub>2. ?R ls\<^sub>2")
+ proof
+ show "?R ls\<^sub>2"
+ proof
+ have hp: "h\<^sub>1 a = Some (D, fs)" by fact
+ have "P \<turnstile> D \<preceq>\<^sup>* C" by fact hence caught: "compP\<^sub>1 P \<turnstile> D \<preceq>\<^sup>* C" by simp
+ from TryCatch\<^sub>1[OF 1 _ caught len\<^sub>1 2, OF hp]
+ show "compP\<^sub>1 P \<turnstile>\<^sub>1 \<langle>compE\<^sub>1 Vs ?e,(h,ls,sh)\<rangle> \<Rightarrow> \<langle>fin\<^sub>1 e',(h\<^sub>2,ls\<^sub>2,sh\<^sub>2)\<rangle>" by simp
+ next
+ show "l\<^sub>2(V := l\<^sub>1 V) \<subseteq>\<^sub>m [Vs [\<mapsto>] ls\<^sub>2]"
+ proof -
+ have "l\<^sub>2 \<subseteq>\<^sub>m [Vs [\<mapsto>] ls\<^sub>2, V \<mapsto> ls\<^sub>2 ! length Vs]"
+ using len\<^sub>1 rel\<^sub>2 by simp
+ moreover
+ { assume VinVs: "V \<in> set Vs"
+ hence "hidden (Vs @ [V]) (last_index Vs V)" by(rule hidden_last_index)
+ hence "unmod (compE\<^sub>1 (Vs @ [V]) e\<^sub>2) (last_index Vs V)"
+ by(rule hidden_unmod)
+ moreover have "last_index Vs V < length ?ls"
+ using len\<^sub>1 VinVs by simp
+ ultimately have "?ls ! last_index Vs V = ls\<^sub>2 ! last_index Vs V"
+ by(rule eval\<^sub>1_preserves_unmod[OF 2])
+ moreover have "last_index Vs V < size Vs" using VinVs by simp
+ ultimately have "ls\<^sub>1 ! last_index Vs V = ls\<^sub>2 ! last_index Vs V"
+ using len\<^sub>1 by(simp del:size_last_index_conv)
+ }
+ ultimately show ?thesis using Block_lem[OF rel\<^sub>1] len\<^sub>1 by simp
+ qed
+ qed
+ qed
+next
+ case Try thus ?case by(fastforce intro!:Try\<^sub>1)
+next
+ case Throw thus ?case by(fastforce intro!:Throw\<^sub>1)
+next
+ case ThrowNull thus ?case by(fastforce intro!:ThrowNull\<^sub>1)
+next
+ case ThrowThrow thus ?case by(fastforce intro!:ThrowThrow\<^sub>1)
+next
+ case (CondT e h l sh h\<^sub>1 l\<^sub>1 sh\<^sub>1 e\<^sub>1 e' h\<^sub>2 l\<^sub>2 sh\<^sub>2 e\<^sub>2)
+ have "PROP ?P e h l sh true h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact
+ with CondT.prems
+ obtain ls\<^sub>1 where 1: "?Post e h l sh true h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1"
+ "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len)
+ have "PROP ?P e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 e' h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact
+ with 1 CondT.prems
+ obtain ls\<^sub>2 where 2: "?Post e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 e' h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto)
+ from 1 2 show ?case by(auto intro!:CondT\<^sub>1)
+next
+ case (CondF e h l sh h\<^sub>1 l\<^sub>1 sh\<^sub>1 e\<^sub>2 e' h\<^sub>2 l\<^sub>2 sh\<^sub>2 e\<^sub>1 Vs ls)
+ have "PROP ?P e h l sh false h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact
+ with CondF.prems
+ obtain ls\<^sub>1 where 1: "?Post e h l sh false h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1"
+ "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len)
+ have "PROP ?P e\<^sub>2 h\<^sub>1 l\<^sub>1 sh\<^sub>1 e' h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact
+ with 1 CondF.prems
+ obtain ls\<^sub>2 where 2: "?Post e\<^sub>2 h\<^sub>1 l\<^sub>1 sh\<^sub>1 e' h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto)
+ from 1 2 show ?case by(auto intro!:CondF\<^sub>1)
+next
+ case CondThrow thus ?case by(fastforce intro!:CondThrow\<^sub>1)
+next
+ case (Seq e h l sh v h\<^sub>1 l\<^sub>1 sh\<^sub>1 e\<^sub>1 e' h\<^sub>2 l\<^sub>2 sh\<^sub>2)
+ have "PROP ?P e h l sh (Val v) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact
+ with Seq.prems
+ obtain ls\<^sub>1 where 1: "?Post e h l sh (Val v) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1"
+ "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len)
+ have "PROP ?P e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 e' h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact
+ with 1 Seq.prems
+ obtain ls\<^sub>2 where 2: "?Post e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 e' h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto)
+ from 1 2 Seq show ?case by(auto intro!:Seq\<^sub>1)
+next
+ case SeqThrow thus ?case by(fastforce intro!:SeqThrow\<^sub>1)
+next
+ case WhileF thus ?case by(fastforce intro!:eval\<^sub>1_evals\<^sub>1.intros)
+next
+ case (WhileT e h l sh h\<^sub>1 l\<^sub>1 sh\<^sub>1 c v h\<^sub>2 l\<^sub>2 sh\<^sub>2 e' h\<^sub>3 l\<^sub>3 sh\<^sub>3)
+ have "PROP ?P e h l sh true h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact
+ with WhileT.prems
+ obtain ls\<^sub>1 where 1: "?Post e h l sh true h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1"
+ "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len)
+ have "PROP ?P c h\<^sub>1 l\<^sub>1 sh\<^sub>1 (Val v) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact
+ with 1 WhileT.prems
+ obtain ls\<^sub>2 where 2: "?Post c h\<^sub>1 l\<^sub>1 sh\<^sub>1 (Val v) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2"
+ "size ls\<^sub>1 = size ls\<^sub>2" by(auto intro!:eval\<^sub>1_preserves_len)
+ have "PROP ?P (While (e) c) h\<^sub>2 l\<^sub>2 sh\<^sub>2 e' h\<^sub>3 l\<^sub>3 sh\<^sub>3 Vs ls\<^sub>2" by fact
+ with 1 2 WhileT.prems
+ obtain ls\<^sub>3 where 3: "?Post (While (e) c) h\<^sub>2 l\<^sub>2 sh\<^sub>2 e' h\<^sub>3 l\<^sub>3 sh\<^sub>3 Vs ls\<^sub>2 ls\<^sub>3" by(auto)
+ from 1 2 3 show ?case by(auto intro!:WhileT\<^sub>1)
+next
+ case (WhileBodyThrow e h l sh h\<^sub>1 l\<^sub>1 sh\<^sub>1 c e' h\<^sub>2 l\<^sub>2 sh\<^sub>2)
+ have "PROP ?P e h l sh true h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact
+ with WhileBodyThrow.prems
+ obtain ls\<^sub>1 where 1: "?Post e h l sh true h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1"
+ "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len)
+ have "PROP ?P c h\<^sub>1 l\<^sub>1 sh\<^sub>1 (throw e') h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact
+ with 1 WhileBodyThrow.prems
+ obtain ls\<^sub>2 where 2: "?Post c h\<^sub>1 l\<^sub>1 sh\<^sub>1 (throw e') h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by auto
+ from 1 2 show ?case by(auto intro!:WhileBodyThrow\<^sub>1)
+next
+ case WhileCondThrow thus ?case by(fastforce intro!:WhileCondThrow\<^sub>1)
+next
+ case New thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros)
+next
+ case NewFail thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros)
+next
+ case NewInit then show ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros)
+next
+ case NewInitOOM then show ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros)
+next
+ case NewInitThrow then show ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros)
+next
+ case Cast thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros)
+next
+ case CastNull thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros)
+next
+ case CastThrow thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros)
+next
+ case (CastFail e h l sh a h\<^sub>1 l\<^sub>1 sh\<^sub>1 D fs C)
+ have "PROP ?P e h l sh (addr a) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact
+ with CastFail.prems
+ obtain ls\<^sub>1 where 1: "?Post e h l sh (addr a) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1" by auto
+ show ?case using 1 CastFail.hyps
+ by(auto intro!:CastFail\<^sub>1[where D=D])
+next
+ case Val thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros)
+next
+ case (BinOp e h l sh v\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 e\<^sub>1 v\<^sub>2 h\<^sub>2 l\<^sub>2 sh\<^sub>2 bop v)
+ have "PROP ?P e h l sh (Val v\<^sub>1) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact
+ with BinOp.prems
+ obtain ls\<^sub>1 where 1: "?Post e h l sh (Val v\<^sub>1) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1"
+ "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len)
+ have "PROP ?P e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 (Val v\<^sub>2) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact
+ with 1 BinOp.prems
+ obtain ls\<^sub>2 where 2: "?Post e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 (Val v\<^sub>2) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto)
+ from 1 2 BinOp show ?case by(auto intro!:BinOp\<^sub>1)
+next
+ case (BinOpThrow2 e\<^sub>0 h l sh v\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 e\<^sub>1 e h\<^sub>2 l\<^sub>2 sh\<^sub>2 bop)
+ have "PROP ?P e\<^sub>0 h l sh (Val v\<^sub>1) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact
+ with BinOpThrow2.prems
+ obtain ls\<^sub>1 where 1: "?Post e\<^sub>0 h l sh (Val v\<^sub>1) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1"
+ "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len)
+ have "PROP ?P e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 (throw e) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact
+ with 1 BinOpThrow2.prems
+ obtain ls\<^sub>2 where 2: "?Post e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 (throw e) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto)
+ from 1 2 BinOpThrow2 show ?case by(auto intro!:BinOpThrow\<^sub>2\<^sub>1)
+next
+ case BinOpThrow1 thus ?case by(fastforce intro!:eval\<^sub>1_evals\<^sub>1.intros)
+next
+ case Var thus ?case
+ by(force intro!:Var\<^sub>1 simp add: map_le_def fun_upds_apply)
+next
+ case LAss thus ?case
+ by(fastforce simp add: LAss_lem intro:eval\<^sub>1_evals\<^sub>1.intros
+ dest:eval\<^sub>1_preserves_len)
+next
+ case LAssThrow thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros)
+next
+ case FAcc thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros)
+next
+ case FAccNull thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros)
+next
+ case FAccThrow thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros)
+next
+ case (FAccNone e h l sh a h' l' sh' C fs F D)
+ have "PROP ?P e h l sh (addr a) h' l' sh' Vs ls" by fact
+ with FAccNone.prems
+ obtain ls\<^sub>2 where 2: "?Post e h l sh (addr a) h' l' sh' Vs ls ls\<^sub>2" by(auto)
+ from 2 FAccNone show ?case by(rule_tac x = ls\<^sub>2 in exI, auto elim!: FAccNone\<^sub>1)
+next
+ case (FAccStatic e h l sh a h' l' sh' C fs F t D)
+ have "PROP ?P e h l sh (addr a) h' l' sh' Vs ls" by fact
+ with FAccStatic.prems
+ obtain ls\<^sub>2 where 2: "?Post e h l sh (addr a) h' l' sh' Vs ls ls\<^sub>2" by(auto)
+ from 2 FAccStatic show ?case by(rule_tac x = ls\<^sub>2 in exI, auto elim!: FAccStatic\<^sub>1)
+next
+ case SFAcc then show ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros)
+next
+ case (SFAccInit C F t D sh h l v' h' l' sh' sfs i v)
+ have "PROP ?P (INIT D ([D],False) \<leftarrow> unit) h l sh (Val v') h' l' sh' Vs ls" by fact
+ with SFAccInit.prems
+ obtain ls\<^sub>2 where 1: "?Post (INIT D ([D],False) \<leftarrow> unit) h l sh (Val v') h' l' sh' Vs ls ls\<^sub>2" by(auto)
+ from 1 SFAccInit show ?case by(rule_tac x = ls\<^sub>2 in exI, auto intro: SFAccInit\<^sub>1)
+next
+ case (SFAccInitThrow C F t D sh h l a h' l' sh')
+ have "PROP ?P (INIT D ([D],False) \<leftarrow> unit) h l sh (throw a) h' l' sh' Vs ls" by fact
+ with SFAccInitThrow.prems
+ obtain ls\<^sub>2 where 1: "?Post (INIT D ([D],False) \<leftarrow> unit) h l sh (throw a) h' l' sh' Vs ls ls\<^sub>2" by(auto)
+ from 1 SFAccInitThrow show ?case by(rule_tac x = ls\<^sub>2 in exI, auto intro: SFAccInitThrow\<^sub>1)
+next
+ case SFAccNone then show ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros)
+next
+ case SFAccNonStatic then show ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros)
+next
+ case (FAss e\<^sub>1 h l sh a h\<^sub>1 l\<^sub>1 sh\<^sub>1 e\<^sub>2 v h\<^sub>2 l\<^sub>2 sh\<^sub>2 C fs fs' F D h\<^sub>2')
+ have "PROP ?P e\<^sub>1 h l sh (addr a) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact
+ with FAss.prems
+ obtain ls\<^sub>1 where 1: "?Post e\<^sub>1 h l sh (addr a) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1"
+ "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len)
+ have "PROP ?P e\<^sub>2 h\<^sub>1 l\<^sub>1 sh\<^sub>1 (Val v) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact
+ with 1 FAss.prems
+ obtain ls\<^sub>2 where 2: "?Post e\<^sub>2 h\<^sub>1 l\<^sub>1 sh\<^sub>1 (Val v) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto)
+ from 1 2 FAss show ?case by(auto intro!:FAss\<^sub>1)
+next
+ case (FAssNull e\<^sub>1 h l sh h\<^sub>1 l\<^sub>1 sh\<^sub>1 e\<^sub>2 v h\<^sub>2 l\<^sub>2 sh\<^sub>2 F D)
+ have "PROP ?P e\<^sub>1 h l sh null h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact
+ with FAssNull.prems
+ obtain ls\<^sub>1 where 1: "?Post e\<^sub>1 h l sh null h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1"
+ "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len)
+ have "PROP ?P e\<^sub>2 h\<^sub>1 l\<^sub>1 sh\<^sub>1 (Val v) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact
+ with 1 FAssNull.prems
+ obtain ls\<^sub>2 where 2: "?Post e\<^sub>2 h\<^sub>1 l\<^sub>1 sh\<^sub>1 (Val v) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto)
+ from 1 2 FAssNull show ?case by(auto intro!:FAssNull\<^sub>1)
+next
+ case FAssThrow1 thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros)
+next
+ case (FAssThrow2 e\<^sub>1 h l sh v h\<^sub>1 l\<^sub>1 sh\<^sub>1 e\<^sub>2 e h\<^sub>2 l\<^sub>2 sh\<^sub>2 F D)
+ have "PROP ?P e\<^sub>1 h l sh (Val v) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact
+ with FAssThrow2.prems
+ obtain ls\<^sub>1 where 1: "?Post e\<^sub>1 h l sh (Val v) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1"
+ "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len)
+ have "PROP ?P e\<^sub>2 h\<^sub>1 l\<^sub>1 sh\<^sub>1 (throw e) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact
+ with 1 FAssThrow2.prems
+ obtain ls\<^sub>2 where 2: "?Post e\<^sub>2 h\<^sub>1 l\<^sub>1 sh\<^sub>1 (throw e) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto)
+ from 1 2 FAssThrow2 show ?case by(auto intro!:FAssThrow\<^sub>2\<^sub>1)
+next
+ case (FAssNone e\<^sub>1 h l sh a h' l' sh' e\<^sub>2 v h\<^sub>2 l\<^sub>2 sh\<^sub>2 C fs F D)
+ have "PROP ?P e\<^sub>1 h l sh (addr a) h' l' sh' Vs ls" by fact
+ with FAssNone.prems
+ obtain ls\<^sub>1 where 1: "?Post e\<^sub>1 h l sh (addr a) h' l' sh' Vs ls ls\<^sub>1"
+ "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len)
+ have "PROP ?P e\<^sub>2 h' l' sh' (Val v) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact
+ with 1 FAssNone.prems
+ obtain ls\<^sub>2 where 2: "?Post e\<^sub>2 h' l' sh' (Val v) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto)
+ from 1 2 FAssNone show ?case by(auto intro!:FAssNone\<^sub>1)
+next
+ case (FAssStatic e\<^sub>1 h l sh a h' l' sh' e\<^sub>2 v h\<^sub>2 l\<^sub>2 sh\<^sub>2 C fs F t D)
+ have "PROP ?P e\<^sub>1 h l sh (addr a) h' l' sh' Vs ls" by fact
+ with FAssStatic.prems
+ obtain ls\<^sub>1 where 1: "?Post e\<^sub>1 h l sh (addr a) h' l' sh' Vs ls ls\<^sub>1"
+ "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len)
+ have "PROP ?P e\<^sub>2 h' l' sh' (Val v) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact
+ with 1 FAssStatic.prems
+ obtain ls\<^sub>2 where 2: "?Post e\<^sub>2 h' l' sh' (Val v) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto)
+ from 1 2 FAssStatic show ?case by(auto intro!:FAssStatic\<^sub>1)
+next
+ case SFAss then show ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros)
+next
+ case (SFAssInit e\<^sub>2 h l sh v h\<^sub>1 l\<^sub>1 sh\<^sub>1 C F t D v' h' l' sh' sfs i sfs' sh'')
+ have "PROP ?P e\<^sub>2 h l sh (Val v) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact
+ with SFAssInit.prems
+ obtain ls\<^sub>1 where 1: "?Post e\<^sub>2 h l sh (Val v) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1" "length ls = length ls\<^sub>1"
+ by(auto intro!:eval\<^sub>1_preserves_len)
+ then have Init_size: "length Vs \<le> length ls\<^sub>1" using SFAssInit.prems(3) by linarith
+ have "PROP ?P (INIT D ([D],False) \<leftarrow> unit) h\<^sub>1 l\<^sub>1 sh\<^sub>1 (Val v') h' l' sh' Vs ls\<^sub>1" by fact
+ with 1 Init_size SFAssInit.prems
+ obtain ls\<^sub>2 where 2: "?Post (INIT D ([D],False) \<leftarrow> unit) h\<^sub>1 l\<^sub>1 sh\<^sub>1 (Val v') h' l' sh' Vs ls\<^sub>1 ls\<^sub>2"
+ by auto
+ from 1 2 SFAssInit show ?case
+ by(auto simp add: comp_def
+ intro!: SFAssInit\<^sub>1 dest!:evals_final)
+next
+ case (SFAssInitThrow e\<^sub>2 h l sh v h\<^sub>1 l\<^sub>1 sh\<^sub>1 C F t D a h\<^sub>2 l\<^sub>2 sh\<^sub>2)
+ have "PROP ?P e\<^sub>2 h l sh (Val v) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact
+ with SFAssInitThrow.prems
+ obtain ls\<^sub>1 where 1: "?Post e\<^sub>2 h l sh (Val v) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1" "length ls = length ls\<^sub>1"
+ by(auto intro!:eval\<^sub>1_preserves_len)
+ then have Init_size: "length Vs \<le> length ls\<^sub>1" using SFAssInitThrow.prems(3) by linarith
+ have "PROP ?P (INIT D ([D],False) \<leftarrow> unit) h\<^sub>1 l\<^sub>1 sh\<^sub>1 (throw a) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact
+ with 1 Init_size SFAssInitThrow.prems
+ obtain ls\<^sub>2 where 2: "?Post (INIT D ([D],False) \<leftarrow> unit) h\<^sub>1 l\<^sub>1 sh\<^sub>1 (throw a) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2"
+ by auto
+ from 1 2 SFAssInitThrow show ?case
+ by(auto simp add: comp_def
+ intro!: SFAssInitThrow\<^sub>1 dest!:evals_final)
+next
+ case SFAssThrow then show ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros)
+next
+ case (SFAssNone e\<^sub>2 h l sh v h\<^sub>2 l\<^sub>2 sh\<^sub>2 C F D)
+ have "PROP ?P e\<^sub>2 h l sh (Val v) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls" by fact
+ with SFAssNone.prems
+ obtain ls\<^sub>2 where 2: "?Post e\<^sub>2 h l sh (Val v) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls ls\<^sub>2" by(auto)
+ from 2 SFAssNone show ?case by(rule_tac x = ls\<^sub>2 in exI, auto elim!: SFAssNone\<^sub>1)
+next
+ case SFAssNonStatic then show ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros)
+next
+ case (CallNull e h l sh h\<^sub>1 l\<^sub>1 sh\<^sub>1 es vs h\<^sub>2 l\<^sub>2 sh\<^sub>2 M)
+ have "PROP ?P e h l sh null h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact
+ with CallNull.prems
+ obtain ls\<^sub>1 where 1: "?Post e h l sh null h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1"
+ "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len)
+ have "PROP ?Ps es h\<^sub>1 l\<^sub>1 sh\<^sub>1 (map Val vs) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact
+ with 1 CallNull.prems
+ obtain ls\<^sub>2 where 2: "?Posts es h\<^sub>1 l\<^sub>1 sh\<^sub>1 (map Val vs) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto)
+ from 1 2 CallNull show ?case
+ by (auto simp add: comp_def elim!: CallNull\<^sub>1)
+next
+ case CallObjThrow thus ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros)
+next
+ case (CallParamsThrow e h l sh v h\<^sub>1 l\<^sub>1 sh\<^sub>1 es vs ex es' h\<^sub>2 l\<^sub>2 sh\<^sub>2 M)
+ have "PROP ?P e h l sh (Val v) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact
+ with CallParamsThrow.prems
+ obtain ls\<^sub>1 where 1: "?Post e h l sh (Val v) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1"
+ "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len)
+ have "PROP ?Ps es h\<^sub>1 l\<^sub>1 sh\<^sub>1 (map Val vs @ throw ex # es') h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact
+ with 1 CallParamsThrow.prems
+ obtain ls\<^sub>2 where 2: "?Posts es h\<^sub>1 l\<^sub>1 sh\<^sub>1 (map Val vs @ throw ex # es') h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto)
+ from 1 2 CallParamsThrow show ?case
+ by (auto simp add: comp_def
+ elim!: CallParamsThrow\<^sub>1 dest!:evals_final)
+next
+ case (CallNone e h l sh a h\<^sub>1 l\<^sub>1 sh\<^sub>1 ps vs h\<^sub>2 l\<^sub>2 sh\<^sub>2 C fs M)
+ have "PROP ?P e h l sh (addr a) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact
+ with CallNone.prems
+ obtain ls\<^sub>1 where 1: "?Post e h l sh (addr a) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1"
+ "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len)
+ have "PROP ?Ps ps h\<^sub>1 l\<^sub>1 sh\<^sub>1 (map Val vs) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact
+ with 1 CallNone.prems
+ obtain ls\<^sub>2 where 2: "?Posts ps h\<^sub>1 l\<^sub>1 sh\<^sub>1 (map Val vs) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto)
+ from 1 2 CallNone show ?case
+ by (auto simp add: comp_def
+ elim!: CallNone\<^sub>1 dest!:evals_final sees_method_compPD)
+next
+ case (CallStatic e h l sh a h\<^sub>1 l\<^sub>1 sh\<^sub>1 ps vs h\<^sub>2 l\<^sub>2 sh\<^sub>2 C fs M Ts T pns body D)
+ have "PROP ?P e h l sh (addr a) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact
+ with CallStatic.prems
+ obtain ls\<^sub>1 where 1: "?Post e h l sh (addr a) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1"
+ "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len)
+ let ?Vs = pns
+ have mdecl: "P \<turnstile> C sees M,Static: Ts\<rightarrow>T = (pns, body) in D" by fact
+ have mdecl\<^sub>1: "compP\<^sub>1 P \<turnstile> C sees M,Static: Ts\<rightarrow>T = (compE\<^sub>1 ?Vs body) in D"
+ using sees_method_compP[OF mdecl, of "\<lambda>b (pns,e). compE\<^sub>1 (case b of NonStatic \<Rightarrow> this#pns | Static \<Rightarrow> pns) e"]
+ by(simp)
+ have "PROP ?Ps ps h\<^sub>1 l\<^sub>1 sh\<^sub>1 (map Val vs) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact
+ with 1 CallStatic.prems
+ obtain ls\<^sub>2 where 2: "?Posts ps h\<^sub>1 l\<^sub>1 sh\<^sub>1 (map Val vs) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2" by(auto)
+ from 1 2 mdecl\<^sub>1 CallStatic show ?case
+ by (auto simp add: comp_def
+ elim!: CallStatic\<^sub>1 dest!:evals_final)
+next
+ case (Call e h l sh a h\<^sub>1 l\<^sub>1 sh\<^sub>1 es vs h\<^sub>2 l\<^sub>2 sh\<^sub>2 C fs M Ts T pns body D l\<^sub>2' b' h\<^sub>3 l\<^sub>3 sh\<^sub>3)
+ have "PROP ?P e h l sh (addr a) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact
+ with Call.prems
+ obtain ls\<^sub>1 where 1: "?Post e h l sh (addr a) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1"
+ "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len)
+ have "PROP ?Ps es h\<^sub>1 l\<^sub>1 sh\<^sub>1 (map Val vs) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact
+ with 1 Call.prems
+ obtain ls\<^sub>2 where 2: "?Posts es h\<^sub>1 l\<^sub>1 sh\<^sub>1 (map Val vs) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2"
+ "size ls\<^sub>1 = size ls\<^sub>2" by(auto intro!:evals\<^sub>1_preserves_len)
+ let ?Vs = "this#pns"
+ let ?ls = "Addr a # vs @ replicate (max_vars body) undefined"
+ have mdecl: "P \<turnstile> C sees M,NonStatic: Ts\<rightarrow>T = (pns, body) in D" by fact
+ have fv_body: "fv body \<subseteq> set ?Vs" and wf_size: "size Ts = size pns"
+ using wf mdecl by(auto dest!:sees_wf_mdecl simp:wf_mdecl_def)
+ have mdecl\<^sub>1: "compP\<^sub>1 P \<turnstile> C sees M,NonStatic: Ts\<rightarrow>T = (compE\<^sub>1 ?Vs body) in D"
+ using sees_method_compP[OF mdecl, of "\<lambda>b (pns,e). compE\<^sub>1 (case b of NonStatic \<Rightarrow> this#pns | Static \<Rightarrow> pns) e"]
+ by(simp)
+ have [simp]: "l\<^sub>2' = [this \<mapsto> Addr a, pns [\<mapsto>] vs]" by fact
+ have Call_size: "size vs = size pns" by fact
+ have "PROP ?P body h\<^sub>2 l\<^sub>2' sh\<^sub>2 b' h\<^sub>3 l\<^sub>3 sh\<^sub>3 ?Vs ?ls" by fact
+ with 1 2 fv_body Call_size Call.prems
+ obtain ls\<^sub>3 where 3: "?Post body h\<^sub>2 l\<^sub>2' sh\<^sub>2 b' h\<^sub>3 l\<^sub>3 sh\<^sub>3 ?Vs ?ls ls\<^sub>3" by(auto)
+ have hp: "h\<^sub>2 a = Some (C, fs)" by fact
+ from 1 2 3 hp mdecl\<^sub>1 wf_size Call_size show ?case
+ by(fastforce simp add: comp_def
+ intro!: Call\<^sub>1 dest!:evals_final)
+next
+ case (SCallParamsThrow es h l sh vs ex es' h\<^sub>2 l\<^sub>2 sh\<^sub>2 C M)
+ have "PROP ?Ps es h l sh (map Val vs @ throw ex # es') h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls" by fact
+ with SCallParamsThrow.prems
+ obtain ls\<^sub>2 where 2: "?Posts es h l sh (map Val vs @ throw ex # es') h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls ls\<^sub>2" by(auto)
+ from 2 SCallParamsThrow show ?case
+ by (fastforce simp add: comp_def
+ elim!: SCallParamsThrow\<^sub>1 dest!:evals_final)
+next
+ case (SCall ps h l sh vs h\<^sub>2 l\<^sub>2 sh\<^sub>2 C M Ts T pns body D sfs l\<^sub>2' e' h\<^sub>3 l\<^sub>3 sh\<^sub>3)
+ have "PROP ?Ps ps h l sh (map Val vs) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls" by fact
+ with SCall.prems
+ obtain ls\<^sub>2 where 2: "?Posts ps h l sh (map Val vs) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls ls\<^sub>2"
+ "size ls = size ls\<^sub>2" by(auto intro!:evals\<^sub>1_preserves_len)
+ let ?Vs = "pns"
+ let ?ls = "vs @ replicate (max_vars body) undefined"
+ have mdecl: "P \<turnstile> C sees M,Static: Ts\<rightarrow>T = (pns, body) in D" by fact
+ have fv_body: "fv body \<subseteq> set ?Vs" and wf_size: "size Ts = size pns"
+ using wf mdecl by(auto dest!:sees_wf_mdecl simp:wf_mdecl_def)
+ have mdecl\<^sub>1: "compP\<^sub>1 P \<turnstile> C sees M,Static: Ts\<rightarrow>T = (compE\<^sub>1 ?Vs body) in D"
+ using sees_method_compP[OF mdecl, of "\<lambda>b (pns,e). compE\<^sub>1 (case b of NonStatic \<Rightarrow> this#pns | Static \<Rightarrow> pns) e"]
+ by(simp)
+ have [simp]: "l\<^sub>2' = [pns [\<mapsto>] vs]" by fact
+ have SCall_size: "size vs = size pns" by fact
+ have "PROP ?P body h\<^sub>2 l\<^sub>2' sh\<^sub>2 e' h\<^sub>3 l\<^sub>3 sh\<^sub>3 ?Vs ?ls" by fact
+ with 2 fv_body SCall_size SCall.prems
+ obtain ls\<^sub>3 where 3: "?Post body h\<^sub>2 l\<^sub>2' sh\<^sub>2 e' h\<^sub>3 l\<^sub>3 sh\<^sub>3 ?Vs ?ls ls\<^sub>3" by(auto)
+ have shp: "sh\<^sub>2 D = \<lfloor>(sfs, Done)\<rfloor> \<or> M = clinit \<and> sh\<^sub>2 D = \<lfloor>(sfs, Processing)\<rfloor>" by fact
+ from 2 3 shp mdecl\<^sub>1 wf_size SCall_size show ?case
+ by(fastforce simp add: comp_def
+ intro!: SCall\<^sub>1 dest!:evals_final)
+next
+ case (SCallNone ps h l sh vs h' l' sh' C M)
+ have "PROP ?Ps ps h l sh (map Val vs) h' l' sh' Vs ls" by fact
+ with SCallNone.prems
+ obtain ls\<^sub>2 where 2: "?Posts ps h l sh (map Val vs) h' l' sh' Vs ls ls\<^sub>2" by(auto)
+ from 2 SCallNone show ?case
+ by(rule_tac x = ls\<^sub>2 in exI,
+ auto simp add: comp_def elim!: SCallNone\<^sub>1 dest!:evals_final sees_method_compPD)
+next
+ case (SCallNonStatic ps h l sh vs h' l' sh' C M Ts T pns body D)
+ let ?Vs = "this#pns"
+ have mdecl: "P \<turnstile> C sees M,NonStatic: Ts\<rightarrow>T = (pns, body) in D" by fact
+ have mdecl\<^sub>1: "compP\<^sub>1 P \<turnstile> C sees M,NonStatic: Ts\<rightarrow>T = (compE\<^sub>1 ?Vs body) in D"
+ using sees_method_compP[OF mdecl, of "\<lambda>b (pns,e). compE\<^sub>1 (case b of NonStatic \<Rightarrow> this#pns | Static \<Rightarrow> pns) e"]
+ by(simp)
+ have "PROP ?Ps ps h l sh (map Val vs) h' l' sh' Vs ls" by fact
+ with SCallNonStatic.prems
+ obtain ls\<^sub>2 where 2: "?Posts ps h l sh (map Val vs) h' l' sh' Vs ls ls\<^sub>2" by(auto)
+ from 2 mdecl\<^sub>1 SCallNonStatic show ?case
+ by (auto simp add: comp_def
+ elim!: SCallNonStatic\<^sub>1 dest!:evals_final)
+next
+ case (SCallInitThrow ps h l sh vs h\<^sub>1 l\<^sub>1 sh\<^sub>1 C M Ts T pns body D a h\<^sub>2 l\<^sub>2 sh\<^sub>2)
+ have "PROP ?Ps ps h l sh (map Val vs) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact
+ with SCallInitThrow.prems
+ obtain ls\<^sub>1 where 1: "?Posts ps h l sh (map Val vs) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1" "length ls = length ls\<^sub>1"
+ by(auto intro!:evals\<^sub>1_preserves_len)
+ then have Init_size: "length Vs \<le> length ls\<^sub>1" using SCallInitThrow.prems(3) by linarith
+ have "PROP ?P (INIT D ([D],False) \<leftarrow> unit) h\<^sub>1 l\<^sub>1 sh\<^sub>1 (throw a) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact
+ with 1 Init_size SCallInitThrow.prems
+ obtain ls\<^sub>2 where 2: "?Post (INIT D ([D],False) \<leftarrow> unit) h\<^sub>1 l\<^sub>1 sh\<^sub>1 (throw a) h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2"
+ by auto
+ let ?Vs = "pns"
+ have mdecl: "P \<turnstile> C sees M,Static: Ts\<rightarrow>T = (pns, body) in D" by fact
+ have mdecl\<^sub>1: "compP\<^sub>1 P \<turnstile> C sees M,Static: Ts\<rightarrow>T = (compE\<^sub>1 ?Vs body) in D"
+ using sees_method_compP[OF mdecl, of "\<lambda>b (pns,e). compE\<^sub>1 (case b of NonStatic \<Rightarrow> this#pns | Static \<Rightarrow> pns) e"]
+ by(simp)
+ from 1 2 mdecl\<^sub>1 SCallInitThrow show ?case
+ by(auto simp add: comp_def
+ intro!: SCallInitThrow\<^sub>1 dest!:evals_final)
+next
+ case (SCallInit ps h l sh vs h\<^sub>1 l\<^sub>1 sh\<^sub>1 C M Ts T pns body D v' h\<^sub>2 l\<^sub>2 sh\<^sub>2 l\<^sub>2' e' h\<^sub>3 l\<^sub>3 sh\<^sub>3)
+ have "PROP ?Ps ps h l sh (map Val vs) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls" by fact
+ with SCallInit.prems
+ obtain ls\<^sub>1 where 1: "?Posts ps h l sh (map Val vs) h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls ls\<^sub>1" "length ls = length ls\<^sub>1"
+ by(auto intro!:evals\<^sub>1_preserves_len)
+ then have Init_size: "length Vs \<le> length ls\<^sub>1" using SCallInit.prems(3) by linarith
+ have "PROP ?P (INIT D ([D],False) \<leftarrow> unit) h\<^sub>1 l\<^sub>1 sh\<^sub>1 (Val v') h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1" by fact
+ with 1 Init_size SCallInit.prems
+ obtain ls\<^sub>2 where 2: "?Post (INIT D ([D],False) \<leftarrow> unit) h\<^sub>1 l\<^sub>1 sh\<^sub>1 (Val v') h\<^sub>2 l\<^sub>2 sh\<^sub>2 Vs ls\<^sub>1 ls\<^sub>2"
+ by auto
+ let ?Vs = "pns"
+ let ?ls = "vs @ replicate (max_vars body) undefined"
+ have mdecl: "P \<turnstile> C sees M,Static: Ts\<rightarrow>T = (pns, body) in D" by fact
+ have fv_body: "fv body \<subseteq> set ?Vs" and wf_size: "size Ts = size pns"
+ using wf mdecl by(auto dest!:sees_wf_mdecl simp:wf_mdecl_def)
+ have mdecl\<^sub>1: "compP\<^sub>1 P \<turnstile> C sees M,Static: Ts\<rightarrow>T = (compE\<^sub>1 ?Vs body) in D"
+ using sees_method_compP[OF mdecl, of "\<lambda>b (pns,e). compE\<^sub>1 (case b of NonStatic \<Rightarrow> this#pns | Static \<Rightarrow> pns) e"]
+ by(simp)
+ have [simp]: "l\<^sub>2' = [pns [\<mapsto>] vs]" by fact
+ have SCall_size: "size vs = size pns" by fact
+ have nclinit: "M \<noteq> clinit" by fact
+ have "PROP ?P body h\<^sub>2 l\<^sub>2' sh\<^sub>2 e' h\<^sub>3 l\<^sub>3 sh\<^sub>3 ?Vs ?ls" by fact
+ with 2 fv_body SCall_size SCallInit.prems
+ obtain ls\<^sub>3 where 3: "?Post body h\<^sub>2 l\<^sub>2' sh\<^sub>2 e' h\<^sub>3 l\<^sub>3 sh\<^sub>3 ?Vs ?ls ls\<^sub>3" by(auto)
+ have shp: "\<nexists>sfs. sh\<^sub>1 D = \<lfloor>(sfs, Done)\<rfloor>" by fact
+ from 1 2 3 shp mdecl\<^sub>1 wf_size SCall_size nclinit show ?case
+ by(auto simp add: comp_def
+ intro!: SCallInit\<^sub>1 dest!:evals_final)
+next
+\<comment> \<open> init cases \<close>
+ case InitFinal then show ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros)
+next
+ case (InitNone sh C C' Cs e h l e' h' l' sh')
+ let ?sh1 = "sh(C \<mapsto> (sblank P C, Prepared))"
+ have "PROP ?P (INIT C' (C # Cs,False) \<leftarrow> e) h l ?sh1 e' h' l' sh' Vs ls" by fact
+ with InitNone.prems
+ obtain ls\<^sub>2 where 2: "?Post (INIT C' (C # Cs,False) \<leftarrow> e) h l ?sh1 e' h' l' sh' Vs ls ls\<^sub>2" by(auto)
+ from 2 InitNone show ?case by (auto elim!: InitNone\<^sub>1)
+next
+ case InitDone then show ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros)
+next
+ case InitProcessing then show ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros)
+next
+ case InitError then show ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros)
+next
+ case InitObject then show ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros)
+next
+ case (InitNonObject sh C sfs D fs ms sh' C' Cs e h l e' h1 l1 sh1)
+ let ?f = "(\<lambda>b (pns,body). compE\<^sub>1 (case b of NonStatic \<Rightarrow> this#pns | Static \<Rightarrow> pns) body)"
+ have cls: "class (compP ?f P) C = \<lfloor>(D,fs,map (compM ?f) ms)\<rfloor>"
+ by(rule class_compP[OF InitNonObject.hyps(3)])
+ have "PROP ?P (INIT C' (D # C # Cs,False) \<leftarrow> e) h l sh' e' h1 l1 sh1 Vs ls" by fact
+ with InitNonObject.prems
+ obtain ls\<^sub>2 where 2: "?Post (INIT C' (D # C # Cs,False) \<leftarrow> e) h l sh' e' h1 l1 sh1 Vs ls ls\<^sub>2" by(auto)
+ from 2 cls InitNonObject show ?case by (auto elim!: InitNonObject\<^sub>1)
+next
+ case InitRInit then show ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros)
+next
+ case (RInit e h l sh v h' l' sh' C sfs i sh'' C' Cs e' e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1)
+ have "PROP ?P e h l sh (Val v) h' l' sh' Vs ls" by fact
+ with RInit.prems
+ obtain ls\<^sub>1 where 1: "?Post e h l sh (Val v) h' l' sh' Vs ls ls\<^sub>1"
+ "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len)
+ have "PROP ?P (INIT C' (Cs,True) \<leftarrow> e') h' l' sh'' e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls\<^sub>1" by fact
+ with 1 RInit.prems
+ obtain ls\<^sub>2 where 2: "?Post (INIT C' (Cs,True) \<leftarrow> e') h' l' sh'' e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls\<^sub>1 ls\<^sub>2" by(auto)
+ from 1 2 RInit show ?case by (auto elim!: RInit\<^sub>1)
+next
+ case (RInitInitFail e h l sh a h' l' sh' C sfs i sh'' D Cs e' e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1)
+ have "PROP ?P e h l sh (throw a) h' l' sh' Vs ls" by fact
+ with RInitInitFail.prems
+ obtain ls\<^sub>1 where 1: "?Post e h l sh (throw a) h' l' sh' Vs ls ls\<^sub>1"
+ "size ls = size ls\<^sub>1" by(auto intro!:eval\<^sub>1_preserves_len)
+ have fv: "fv (RI (D,throw a) ; Cs \<leftarrow> e') \<subseteq> set Vs"
+ using RInitInitFail.hyps(1) eval_final RInitInitFail.prems(1) subset_eq by fastforce
+ have l': "l' \<subseteq>\<^sub>m [Vs [\<mapsto>] ls\<^sub>1]" by (simp add: "1"(1))
+ have "PROP ?P (RI (D,throw a) ; Cs \<leftarrow> e') h' l' sh'' e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls\<^sub>1" by fact
+ with 1 eval_final[OF RInitInitFail.hyps(1)] RInitInitFail.prems
+ obtain ls\<^sub>2 where 2: "?Post (RI (D,throw a) ; Cs \<leftarrow> e') h' l' sh'' e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 Vs ls\<^sub>1 ls\<^sub>2"
+ by fastforce
+ from 1 2 RInitInitFail show ?case
+ by(fastforce simp add: comp_def
+ intro!: RInitInitFail\<^sub>1 dest!:eval_final)
+next
+ case RInitFailFinal then show ?case by(fastforce intro:eval\<^sub>1_evals\<^sub>1.intros)
+qed
+(*>*)
+
+
+subsection\<open>Preservation of well-formedness\<close>
+
+text\<open> The compiler preserves well-formedness. Is less trivial than it
+may appear. We start with two simple properties: preservation of
+well-typedness \<close>
+
+lemma compE\<^sub>1_pres_wt: "\<And>Vs Ts U.
+ \<lbrakk> P,[Vs[\<mapsto>]Ts] \<turnstile> e :: U; size Ts = size Vs \<rbrakk>
+ \<Longrightarrow> compP f P,Ts \<turnstile>\<^sub>1 compE\<^sub>1 Vs e :: U"
+and "\<And>Vs Ts Us.
+ \<lbrakk> P,[Vs[\<mapsto>]Ts] \<turnstile> es [::] Us; size Ts = size Vs \<rbrakk>
+ \<Longrightarrow> compP f P,Ts \<turnstile>\<^sub>1 compEs\<^sub>1 Vs es [::] Us"
+(*<*)
+proof(induct e and es rule: compE\<^sub>1.induct compEs\<^sub>1.induct)
+ case Var then show ?case
+ by (fastforce simp:map_upds_apply_eq_Some split:if_split_asm)
+next
+ case LAss then show ?case
+ by (fastforce simp:map_upds_apply_eq_Some split:if_split_asm)
+next
+ case Call then show ?case
+ by (fastforce dest!: sees_method_compP[where f = f])
+next
+ case SCall then show ?case
+ by (fastforce dest!: sees_method_compP[where f = f])
+next
+ case Block then show ?case by (fastforce simp:nth_append)
+next
+ case TryCatch then show ?case by (fastforce simp:nth_append)
+qed fastforce+
+(*>*)
+
+text\<open>\noindent and the correct block numbering: \<close>
+
+lemma \<B>: "\<And>Vs n. size Vs = n \<Longrightarrow> \<B> (compE\<^sub>1 Vs e) n"
+and \<B>s: "\<And>Vs n. size Vs = n \<Longrightarrow> \<B>s (compEs\<^sub>1 Vs es) n"
+(*<*)by (induct e and es rule: \<B>.induct \<B>s.induct)
+ (force | simp,metis length_append_singleton)+(*>*)
+
+text\<open> The main complication is preservation of definite assignment
+@{term"\<D>"}. \<close>
+
+lemma image_last_index: "A \<subseteq> set(xs@[x]) \<Longrightarrow> last_index (xs @ [x]) ` A =
+ (if x \<in> A then insert (size xs) (last_index xs ` (A-{x})) else last_index xs ` A)"
+(*<*)
+by(auto simp:image_def)
+(*>*)
+
+
+lemma A_compE\<^sub>1_None[simp]:
+ "\<And>Vs. \<A> e = None \<Longrightarrow> \<A> (compE\<^sub>1 Vs e) = None"
+and "\<And>Vs. \<A>s es = None \<Longrightarrow> \<A>s (compEs\<^sub>1 Vs es) = None"
+(*<*)by(induct e and es rule: compE\<^sub>1.induct compEs\<^sub>1.induct)(auto simp:hyperset_defs)(*>*)
+
+
+lemma A_compE\<^sub>1:
+ "\<And>A Vs. \<lbrakk> \<A> e = \<lfloor>A\<rfloor>; fv e \<subseteq> set Vs \<rbrakk> \<Longrightarrow> \<A> (compE\<^sub>1 Vs e) = \<lfloor>last_index Vs ` A\<rfloor>"
+and "\<And>A Vs. \<lbrakk> \<A>s es = \<lfloor>A\<rfloor>; fvs es \<subseteq> set Vs \<rbrakk> \<Longrightarrow> \<A>s (compEs\<^sub>1 Vs es) = \<lfloor>last_index Vs ` A\<rfloor>"
+(*<*)
+proof(induct e and es rule: fv.induct fvs.induct)
+ case (Block V' T e)
+ hence "fv e \<subseteq> set (Vs@[V'])" by fastforce
+ moreover obtain B where "\<A> e = \<lfloor>B\<rfloor>"
+ using Block.prems by(simp add: hyperset_defs)
+ moreover from calculation have "B \<subseteq> set (Vs@[V'])" by(auto dest!:A_fv)
+ ultimately show ?case using Block
+ by(auto simp add: hyperset_defs image_last_index last_index_size_conv)
+next
+ case (TryCatch e\<^sub>1 C V' e\<^sub>2)
+ hence fve\<^sub>2: "fv e\<^sub>2 \<subseteq> set (Vs@[V'])" by auto
+ show ?case
+ proof (cases "\<A> e\<^sub>1")
+ assume A\<^sub>1: "\<A> e\<^sub>1 = None"
+ then obtain A\<^sub>2 where A\<^sub>2: "\<A> e\<^sub>2 = \<lfloor>A\<^sub>2\<rfloor>" using TryCatch
+ by(simp add:hyperset_defs)
+ hence "A\<^sub>2 \<subseteq> set (Vs@[V'])" using TryCatch.prems A_fv[OF A\<^sub>2] by simp blast
+ thus ?thesis using TryCatch fve\<^sub>2 A\<^sub>1 A\<^sub>2
+ by(auto simp add:hyperset_defs image_last_index last_index_size_conv)
+ next
+ fix A\<^sub>1 assume A\<^sub>1: "\<A> e\<^sub>1 = \<lfloor>A\<^sub>1\<rfloor>"
+ show ?thesis
+ proof (cases "\<A> e\<^sub>2")
+ assume A\<^sub>2: "\<A> e\<^sub>2 = None"
+ then show ?case using TryCatch A\<^sub>1 by(simp add:hyperset_defs)
+ next
+ fix A\<^sub>2 assume A\<^sub>2: "\<A> e\<^sub>2 = \<lfloor>A\<^sub>2\<rfloor>"
+ have "A\<^sub>1 \<subseteq> set Vs" using TryCatch.prems A_fv[OF A\<^sub>1] by simp blast
+ moreover
+ have "A\<^sub>2 \<subseteq> set (Vs@[V'])" using TryCatch.prems A_fv[OF A\<^sub>2] by simp blast
+ ultimately show ?thesis using TryCatch A\<^sub>1 A\<^sub>2
+ by (auto simp add: Diff_subset_conv last_index_size_conv subsetD hyperset_defs
+ dest!: sym [of _ A])
+ qed
+ qed
+next
+ case (Cond e e\<^sub>1 e\<^sub>2)
+ { assume "\<A> e = None \<or> \<A> e\<^sub>1 = None \<or> \<A> e\<^sub>2 = None"
+ hence ?case using Cond by(auto simp add:hyperset_defs image_Un)
+ }
+ moreover
+ { fix A A\<^sub>1 A\<^sub>2
+ assume "\<A> e = \<lfloor>A\<rfloor>" and A\<^sub>1: "\<A> e\<^sub>1 = \<lfloor>A\<^sub>1\<rfloor>" and A\<^sub>2: "\<A> e\<^sub>2 = \<lfloor>A\<^sub>2\<rfloor>"
+ moreover
+ have "A\<^sub>1 \<subseteq> set Vs" using Cond.prems A_fv[OF A\<^sub>1] by simp blast
+ moreover
+ have "A\<^sub>2 \<subseteq> set Vs" using Cond.prems A_fv[OF A\<^sub>2] by simp blast
+ ultimately have ?case using Cond
+ by(auto simp add:hyperset_defs image_Un
+ inj_on_image_Int[OF inj_on_last_index])
+ }
+ ultimately show ?case by fastforce
+qed (auto simp add:hyperset_defs)
+(*>*)
+
+
+lemma D_None[iff]: "\<D> (e::'a exp) None" and [iff]: "\<D>s (es::'a exp list) None"
+(*<*)by(induct e and es rule: \<D>.induct \<D>s.induct)(simp_all)(*>*)
+
+
+lemma D_last_index_compE\<^sub>1:
+ "\<And>A Vs. \<lbrakk> A \<subseteq> set Vs; fv e \<subseteq> set Vs \<rbrakk> \<Longrightarrow>
+ \<D> e \<lfloor>A\<rfloor> \<Longrightarrow> \<D> (compE\<^sub>1 Vs e) \<lfloor>last_index Vs ` A\<rfloor>"
+and "\<And>A Vs. \<lbrakk> A \<subseteq> set Vs; fvs es \<subseteq> set Vs \<rbrakk> \<Longrightarrow>
+ \<D>s es \<lfloor>A\<rfloor> \<Longrightarrow> \<D>s (compEs\<^sub>1 Vs es) \<lfloor>last_index Vs ` A\<rfloor>"
+(*<*)
+proof(induct e and es rule: \<D>.induct \<D>s.induct)
+ case (BinOp e\<^sub>1 bop e\<^sub>2)
+ hence IH\<^sub>1: "\<D> (compE\<^sub>1 Vs e\<^sub>1) \<lfloor>last_index Vs ` A\<rfloor>" by simp
+ show ?case
+ proof (cases "\<A> e\<^sub>1")
+ case None thus ?thesis using BinOp by simp
+ next
+ case (Some A\<^sub>1)
+ have indexA\<^sub>1: "\<A> (compE\<^sub>1 Vs e\<^sub>1) = \<lfloor>last_index Vs ` A\<^sub>1\<rfloor>"
+ using A_compE\<^sub>1[OF Some] BinOp.prems by auto
+ have "A \<union> A\<^sub>1 \<subseteq> set Vs" using BinOp.prems A_fv[OF Some] by auto
+ hence "\<D> (compE\<^sub>1 Vs e\<^sub>2) \<lfloor>last_index Vs ` (A \<union> A\<^sub>1)\<rfloor>" using BinOp Some by auto
+ hence "\<D> (compE\<^sub>1 Vs e\<^sub>2) \<lfloor>last_index Vs ` A \<union> last_index Vs ` A\<^sub>1\<rfloor>"
+ by(simp add: image_Un)
+ thus ?thesis using IH\<^sub>1 indexA\<^sub>1 by auto
+ qed
+next
+ case (FAss e\<^sub>1 F D e\<^sub>2)
+ hence IH\<^sub>1: "\<D> (compE\<^sub>1 Vs e\<^sub>1) \<lfloor>last_index Vs ` A\<rfloor>" by simp
+ show ?case
+ proof (cases "\<A> e\<^sub>1")
+ case None thus ?thesis using FAss by simp
+ next
+ case (Some A\<^sub>1)
+ have indexA\<^sub>1: "\<A> (compE\<^sub>1 Vs e\<^sub>1) = \<lfloor>last_index Vs ` A\<^sub>1\<rfloor>"
+ using A_compE\<^sub>1[OF Some] FAss.prems by auto
+ have "A \<union> A\<^sub>1 \<subseteq> set Vs" using FAss.prems A_fv[OF Some] by auto
+ hence "\<D> (compE\<^sub>1 Vs e\<^sub>2) \<lfloor>last_index Vs ` (A \<union> A\<^sub>1)\<rfloor>" using FAss Some by auto
+ hence "\<D> (compE\<^sub>1 Vs e\<^sub>2) \<lfloor>last_index Vs ` A \<union> last_index Vs ` A\<^sub>1\<rfloor>"
+ by(simp add: image_Un)
+ thus ?thesis using IH\<^sub>1 indexA\<^sub>1 by auto
+ qed
+next
+ case (Call e\<^sub>1 M es)
+ hence IH\<^sub>1: "\<D> (compE\<^sub>1 Vs e\<^sub>1) \<lfloor>last_index Vs ` A\<rfloor>" by simp
+ show ?case
+ proof (cases "\<A> e\<^sub>1")
+ case None thus ?thesis using Call by simp
+ next
+ case (Some A\<^sub>1)
+ have indexA\<^sub>1: "\<A> (compE\<^sub>1 Vs e\<^sub>1) = \<lfloor>last_index Vs ` A\<^sub>1\<rfloor>"
+ using A_compE\<^sub>1[OF Some] Call.prems by auto
+ have "A \<union> A\<^sub>1 \<subseteq> set Vs" using Call.prems A_fv[OF Some] by auto
+ hence "\<D>s (compEs\<^sub>1 Vs es) \<lfloor>last_index Vs ` (A \<union> A\<^sub>1)\<rfloor>" using Call Some by auto
+ hence "\<D>s (compEs\<^sub>1 Vs es) \<lfloor>last_index Vs ` A \<union> last_index Vs ` A\<^sub>1\<rfloor>"
+ by(simp add: image_Un)
+ thus ?thesis using IH\<^sub>1 indexA\<^sub>1 by auto
+ qed
+next
+ case (TryCatch e\<^sub>1 C V e\<^sub>2)
+ have "\<lbrakk> A\<union>{V} \<subseteq> set(Vs@[V]); fv e\<^sub>2 \<subseteq> set(Vs@[V]); \<D> e\<^sub>2 \<lfloor>A\<union>{V}\<rfloor>\<rbrakk> \<Longrightarrow>
+ \<D> (compE\<^sub>1 (Vs@[V]) e\<^sub>2) \<lfloor>last_index (Vs@[V]) ` (A\<union>{V})\<rfloor>" by fact
+ hence "\<D> (compE\<^sub>1 (Vs@[V]) e\<^sub>2) \<lfloor>last_index (Vs@[V]) ` (A\<union>{V})\<rfloor>"
+ using TryCatch.prems by(simp add:Diff_subset_conv)
+ moreover have "last_index (Vs@[V]) ` A \<subseteq> last_index Vs ` A \<union> {size Vs}"
+ using TryCatch.prems by(auto simp add: image_last_index split:if_split_asm)
+ ultimately show ?case using TryCatch
+ by(auto simp:hyperset_defs elim!:D_mono')
+next
+ case (Seq e\<^sub>1 e\<^sub>2)
+ hence IH\<^sub>1: "\<D> (compE\<^sub>1 Vs e\<^sub>1) \<lfloor>last_index Vs ` A\<rfloor>" by simp
+ show ?case
+ proof (cases "\<A> e\<^sub>1")
+ case None thus ?thesis using Seq by simp
+ next
+ case (Some A\<^sub>1)
+ have indexA\<^sub>1: "\<A> (compE\<^sub>1 Vs e\<^sub>1) = \<lfloor>last_index Vs ` A\<^sub>1\<rfloor>"
+ using A_compE\<^sub>1[OF Some] Seq.prems by auto
+ have "A \<union> A\<^sub>1 \<subseteq> set Vs" using Seq.prems A_fv[OF Some] by auto
+ hence "\<D> (compE\<^sub>1 Vs e\<^sub>2) \<lfloor>last_index Vs ` (A \<union> A\<^sub>1)\<rfloor>" using Seq Some by auto
+ hence "\<D> (compE\<^sub>1 Vs e\<^sub>2) \<lfloor>last_index Vs ` A \<union> last_index Vs ` A\<^sub>1\<rfloor>"
+ by(simp add: image_Un)
+ thus ?thesis using IH\<^sub>1 indexA\<^sub>1 by auto
+ qed
+next
+ case (Cond e e\<^sub>1 e\<^sub>2)
+ hence IH\<^sub>1: "\<D> (compE\<^sub>1 Vs e) \<lfloor>last_index Vs ` A\<rfloor>" by simp
+ show ?case
+ proof (cases "\<A> e")
+ case None thus ?thesis using Cond by simp
+ next
+ case (Some B)
+ have indexB: "\<A> (compE\<^sub>1 Vs e) = \<lfloor>last_index Vs ` B\<rfloor>"
+ using A_compE\<^sub>1[OF Some] Cond.prems by auto
+ have "A \<union> B \<subseteq> set Vs" using Cond.prems A_fv[OF Some] by auto
+ hence "\<D> (compE\<^sub>1 Vs e\<^sub>1) \<lfloor>last_index Vs ` (A \<union> B)\<rfloor>"
+ and "\<D> (compE\<^sub>1 Vs e\<^sub>2) \<lfloor>last_index Vs ` (A \<union> B)\<rfloor>"
+ using Cond Some by auto
+ hence "\<D> (compE\<^sub>1 Vs e\<^sub>1) \<lfloor>last_index Vs ` A \<union> last_index Vs ` B\<rfloor>"
+ and "\<D> (compE\<^sub>1 Vs e\<^sub>2) \<lfloor>last_index Vs ` A \<union> last_index Vs ` B\<rfloor>"
+ by(simp add: image_Un)+
+ thus ?thesis using IH\<^sub>1 indexB by auto
+ qed
+next
+ case (While e\<^sub>1 e\<^sub>2)
+ hence IH\<^sub>1: "\<D> (compE\<^sub>1 Vs e\<^sub>1) \<lfloor>last_index Vs ` A\<rfloor>" by simp
+ show ?case
+ proof (cases "\<A> e\<^sub>1")
+ case None thus ?thesis using While by simp
+ next
+ case (Some A\<^sub>1)
+ have indexA\<^sub>1: "\<A> (compE\<^sub>1 Vs e\<^sub>1) = \<lfloor>last_index Vs ` A\<^sub>1\<rfloor>"
+ using A_compE\<^sub>1[OF Some] While.prems by auto
+ have "A \<union> A\<^sub>1 \<subseteq> set Vs" using While.prems A_fv[OF Some] by auto
+ hence "\<D> (compE\<^sub>1 Vs e\<^sub>2) \<lfloor>last_index Vs ` (A \<union> A\<^sub>1)\<rfloor>" using While Some by auto
+ hence "\<D> (compE\<^sub>1 Vs e\<^sub>2) \<lfloor>last_index Vs ` A \<union> last_index Vs ` A\<^sub>1\<rfloor>"
+ by(simp add: image_Un)
+ thus ?thesis using IH\<^sub>1 indexA\<^sub>1 by auto
+ qed
+next
+ case (Block V T e)
+ have "\<lbrakk> A-{V} \<subseteq> set(Vs@[V]); fv e \<subseteq> set(Vs@[V]); \<D> e \<lfloor>A-{V}\<rfloor> \<rbrakk> \<Longrightarrow>
+ \<D> (compE\<^sub>1 (Vs@[V]) e) \<lfloor>last_index (Vs@[V]) ` (A-{V})\<rfloor>" by fact
+ hence "\<D> (compE\<^sub>1 (Vs@[V]) e) \<lfloor>last_index (Vs@[V]) ` (A-{V})\<rfloor>"
+ using Block.prems by(simp add:Diff_subset_conv)
+ moreover have "size Vs \<notin> last_index Vs ` A"
+ using Block.prems by(auto simp add:image_def size_last_index_conv)
+ ultimately show ?case using Block
+ by(auto simp add: image_last_index Diff_subset_conv hyperset_defs elim!: D_mono')
+next
+ case (Cons_exp e\<^sub>1 es)
+ hence IH\<^sub>1: "\<D> (compE\<^sub>1 Vs e\<^sub>1) \<lfloor>last_index Vs ` A\<rfloor>" by simp
+ show ?case
+ proof (cases "\<A> e\<^sub>1")
+ case None thus ?thesis using Cons_exp by simp
+ next
+ case (Some A\<^sub>1)
+ have indexA\<^sub>1: "\<A> (compE\<^sub>1 Vs e\<^sub>1) = \<lfloor>last_index Vs ` A\<^sub>1\<rfloor>"
+ using A_compE\<^sub>1[OF Some] Cons_exp.prems by auto
+ have "A \<union> A\<^sub>1 \<subseteq> set Vs" using Cons_exp.prems A_fv[OF Some] by auto
+ hence "\<D>s (compEs\<^sub>1 Vs es) \<lfloor>last_index Vs ` (A \<union> A\<^sub>1)\<rfloor>" using Cons_exp Some by auto
+ hence "\<D>s (compEs\<^sub>1 Vs es) \<lfloor>last_index Vs ` A \<union> last_index Vs ` A\<^sub>1\<rfloor>"
+ by(simp add: image_Un)
+ thus ?thesis using IH\<^sub>1 indexA\<^sub>1 by auto
+ qed
+qed (simp_all add:hyperset_defs)
+(*>*)
+
+
+lemma last_index_image_set: "distinct xs \<Longrightarrow> last_index xs ` set xs = {..<size xs}"
+(*<*)by(induct xs rule:rev_induct) (auto simp add: image_last_index)(*>*)
+
+
+lemma D_compE\<^sub>1:
+ "\<lbrakk> \<D> e \<lfloor>set Vs\<rfloor>; fv e \<subseteq> set Vs; distinct Vs \<rbrakk> \<Longrightarrow> \<D> (compE\<^sub>1 Vs e) \<lfloor>{..<length Vs}\<rfloor>"
+(*<*)by(fastforce dest!: D_last_index_compE\<^sub>1[OF subset_refl] simp add:last_index_image_set)(*>*)
+
+
+lemma D_compE\<^sub>1':
+assumes "\<D> e \<lfloor>set(V#Vs)\<rfloor>" and "fv e \<subseteq> set(V#Vs)" and "distinct(V#Vs)"
+shows "\<D> (compE\<^sub>1 (V#Vs) e) \<lfloor>{..length Vs}\<rfloor>"
+(*<*)
+proof -
+ have "{..size Vs} = {..<size(V#Vs)}" by auto
+ thus ?thesis using assms by (simp only:)(rule D_compE\<^sub>1)
+qed
+(*>*)
+
+lemma compP\<^sub>1_pres_wf: "wf_J_prog P \<Longrightarrow> wf_J\<^sub>1_prog (compP\<^sub>1 P)"
+(*<*)
+proof -
+ assume wf: "wf_J_prog P"
+ let ?f = "(\<lambda>b (pns, body).
+ compE\<^sub>1 (case b of Static \<Rightarrow> pns | NonStatic \<Rightarrow> this # pns) body)"
+ let ?wf\<^sub>2 = "wf_J\<^sub>1_mdecl"
+
+ { fix C M b Ts T m
+ assume cM: "P \<turnstile> C sees M, b : Ts\<rightarrow>T = m in C"
+ and wfm: "wf_mdecl wf_J_mdecl P C (M, b, Ts, T, m)"
+ obtain pns body where [simp]: "m = (pns, body)" by(cases m) simp
+ let ?E = "\<lambda>b. case b of Static \<Rightarrow> [pns [\<mapsto>] Ts] | NonStatic \<Rightarrow> [pns [\<mapsto>] Ts, this \<mapsto> Class C]"
+ obtain T' where WT: "P,?E b \<turnstile> body :: T'" and subT: "P \<turnstile> T' \<le> T"
+ using wfm by(cases b) (auto simp: wf_mdecl_def)
+ have fv: "fv body \<subseteq> dom (?E b)" by(rule WT_fv[OF WT])
+ have "wf_mdecl ?wf\<^sub>2 (compP ?f P) C (M, b, Ts, T, ?f b m)"
+ proof(cases b)
+ case Static then show ?thesis using cM wfm fv
+ by(auto simp:wf_mdecl_def wf_J\<^sub>1_mdecl_def
+ intro!: compE\<^sub>1_pres_wt D_compE\<^sub>1 \<B>)
+ next
+ case NonStatic then show ?thesis using cM wfm fv
+ by(clarsimp simp add:wf_mdecl_def wf_J\<^sub>1_mdecl_def)
+ (fastforce intro!: compE\<^sub>1_pres_wt D_compE\<^sub>1' \<B>)
+ qed
+ }
+ then show ?thesis by simp (rule wf_prog_compPI[OF _ wf])
+qed
+(*>*)
+
+
+end
diff --git a/thys/JinjaDCI/Compiler/Correctness2.thy b/thys/JinjaDCI/Compiler/Correctness2.thy
--- a/thys/JinjaDCI/Compiler/Correctness2.thy
+++ b/thys/JinjaDCI/Compiler/Correctness2.thy
@@ -1,3314 +1,3314 @@
-(* Title: JinjaDCI/Compiler/Correctness2.thy
- Author: Tobias Nipkow, Susannah Mansky
- Copyright TUM 2003, UIUC 2019-20
-
- Based on the Jinja theory Compiler/Correctness2.thy by Tobias Nipkow
-*)
-
-section \<open> Correctness of Stage 2 \<close>
-
-theory Correctness2
-imports "HOL-Library.Sublist" Compiler2 J1WellForm "../J/EConform"
-begin
-
-(*<*)hide_const (open) Throw(*>*)
-
-subsection\<open> Instruction sequences \<close>
-
-text\<open> How to select individual instructions and subsequences of
-instructions from a program given the class, method and program
-counter. \<close>
-
-definition before :: "jvm_prog \<Rightarrow> cname \<Rightarrow> mname \<Rightarrow> nat \<Rightarrow> instr list \<Rightarrow> bool"
- ("(_,_,_,_/ \<rhd> _)" [51,0,0,0,51] 50) where
- "P,C,M,pc \<rhd> is \<longleftrightarrow> prefix is (drop pc (instrs_of P C M))"
-
-definition at :: "jvm_prog \<Rightarrow> cname \<Rightarrow> mname \<Rightarrow> nat \<Rightarrow> instr \<Rightarrow> bool"
- ("(_,_,_,_/ \<triangleright> _)" [51,0,0,0,51] 50) where
- "P,C,M,pc \<triangleright> i \<longleftrightarrow> (\<exists>is. drop pc (instrs_of P C M) = i#is)"
-
-lemma [simp]: "P,C,M,pc \<rhd> []"
-(*<*)by(simp add:before_def)(*>*)
-
-
-lemma [simp]: "P,C,M,pc \<rhd> (i#is) = (P,C,M,pc \<triangleright> i \<and> P,C,M,pc + 1 \<rhd> is)"
-(*<*)by(fastforce simp add:before_def at_def prefix_def drop_Suc drop_tl)(*>*)
-
-(*<*)
-declare drop_drop[simp del]
-(*>*)
-
-
-lemma [simp]: "P,C,M,pc \<rhd> (is\<^sub>1 @ is\<^sub>2) = (P,C,M,pc \<rhd> is\<^sub>1 \<and> P,C,M,pc + size is\<^sub>1 \<rhd> is\<^sub>2)"
-(*<*)
-by(subst add.commute)
- (fastforce simp add:before_def prefix_def drop_drop[symmetric])
-(*>*)
-
-(*<*)
-declare drop_drop[simp]
-(*>*)
-
-
-lemma [simp]: "P,C,M,pc \<triangleright> i \<Longrightarrow> instrs_of P C M ! pc = i"
-(*<*)by(clarsimp simp add:at_def strict_prefix_def nth_via_drop)(*>*)
-
-lemma beforeM:
- "P \<turnstile> C sees M,b: Ts\<rightarrow>T = body in D \<Longrightarrow>
- compP\<^sub>2 P,D,M,0 \<rhd> compE\<^sub>2 body @ [Return]"
-(*<*)by(drule sees_method_idemp) (simp add:before_def compMb\<^sub>2_def)(*>*)
-
-text\<open> This lemma executes a single instruction by rewriting: \<close>
-
-lemma [simp]:
- "P,C,M,pc \<triangleright> instr \<Longrightarrow>
- (P \<turnstile> (None, h, (vs,ls,C,M,pc,ics) # frs, sh) -jvm\<rightarrow> \<sigma>') =
- ((None, h, (vs,ls,C,M,pc,ics) # frs, sh) = \<sigma>' \<or>
- (\<exists>\<sigma>. exec(P,(None, h, (vs,ls,C,M,pc,ics) # frs, sh)) = Some \<sigma> \<and> P \<turnstile> \<sigma> -jvm\<rightarrow> \<sigma>'))"
-(*<*)
-by(simp only: exec_all_def)
- (blast intro: converse_rtranclE converse_rtrancl_into_rtrancl)
-(*>*)
-
-
-subsection\<open> Exception tables \<close>
-
-definition pcs :: "ex_table \<Rightarrow> nat set"
-where
- "pcs xt \<equiv> \<Union>(f,t,C,h,d) \<in> set xt. {f ..< t}"
-
-lemma pcs_subset:
-shows "(\<And>pc d. pcs(compxE\<^sub>2 e pc d) \<subseteq> {pc..<pc+size(compE\<^sub>2 e)})"
-and "(\<And>pc d. pcs(compxEs\<^sub>2 es pc d) \<subseteq> {pc..<pc+size(compEs\<^sub>2 es)})"
-(*<*)
-proof(induct e and es rule: compxE\<^sub>2.induct compxEs\<^sub>2.induct)
- case Cast then show ?case by (fastforce simp:pcs_def)
-next
- case BinOp then show ?case by (fastforce simp:pcs_def split:bop.splits)
-next
- case LAss then show ?case by (fastforce simp: pcs_def)
-next
- case FAcc then show ?case by (fastforce simp: pcs_def)
-next
- case FAss then show ?case by (fastforce simp: pcs_def)
-next
- case SFAss then show ?case by (fastforce simp: pcs_def)
-next
- case Call then show ?case by (fastforce simp: pcs_def)
-next
- case SCall then show ?case by (fastforce simp: pcs_def)
-next
- case Seq then show ?case by (fastforce simp: pcs_def)
-next
- case Cond then show ?case by (fastforce simp: pcs_def)
-next
- case While then show ?case by (fastforce simp: pcs_def)
-next
- case throw then show ?case by (fastforce simp: pcs_def)
-next
- case TryCatch then show ?case by (fastforce simp: pcs_def)
-next
- case Cons_exp then show ?case by (fastforce simp: pcs_def)
-qed (simp_all add:pcs_def)
-(*>*)
-
-
-lemma [simp]: "pcs [] = {}"
-(*<*)by(simp add:pcs_def)(*>*)
-
-
-lemma [simp]: "pcs (x#xt) = {fst x ..< fst(snd x)} \<union> pcs xt"
-(*<*)by(auto simp add: pcs_def)(*>*)
-
-
-lemma [simp]: "pcs(xt\<^sub>1 @ xt\<^sub>2) = pcs xt\<^sub>1 \<union> pcs xt\<^sub>2"
-(*<*)by(simp add:pcs_def)(*>*)
-
-
-lemma [simp]: "pc < pc\<^sub>0 \<or> pc\<^sub>0+size(compE\<^sub>2 e) \<le> pc \<Longrightarrow> pc \<notin> pcs(compxE\<^sub>2 e pc\<^sub>0 d)"
-(*<*)using pcs_subset by fastforce(*>*)
-
-
-lemma [simp]: "pc < pc\<^sub>0 \<or> pc\<^sub>0+size(compEs\<^sub>2 es) \<le> pc \<Longrightarrow> pc \<notin> pcs(compxEs\<^sub>2 es pc\<^sub>0 d)"
-(*<*)using pcs_subset by fastforce(*>*)
-
-
-lemma [simp]: "pc\<^sub>1 + size(compE\<^sub>2 e\<^sub>1) \<le> pc\<^sub>2 \<Longrightarrow> pcs(compxE\<^sub>2 e\<^sub>1 pc\<^sub>1 d\<^sub>1) \<inter> pcs(compxE\<^sub>2 e\<^sub>2 pc\<^sub>2 d\<^sub>2) = {}"
-(*<*)using pcs_subset by fastforce(*>*)
-
-
-lemma [simp]: "pc\<^sub>1 + size(compE\<^sub>2 e) \<le> pc\<^sub>2 \<Longrightarrow> pcs(compxE\<^sub>2 e pc\<^sub>1 d\<^sub>1) \<inter> pcs(compxEs\<^sub>2 es pc\<^sub>2 d\<^sub>2) = {}"
-(*<*)using pcs_subset by fastforce(*>*)
-
-
-lemma [simp]:
- "pc \<notin> pcs xt\<^sub>0 \<Longrightarrow> match_ex_table P C pc (xt\<^sub>0 @ xt\<^sub>1) = match_ex_table P C pc xt\<^sub>1"
-(*<*)by (induct xt\<^sub>0) (auto simp: matches_ex_entry_def)(*>*)
-
-
-lemma [simp]: "\<lbrakk> x \<in> set xt; pc \<notin> pcs xt \<rbrakk> \<Longrightarrow> \<not> matches_ex_entry P D pc x"
-(*<*)by(auto simp:matches_ex_entry_def pcs_def)(*>*)
-
-
-lemma [simp]:
-assumes xe: "xe \<in> set(compxE\<^sub>2 e pc d)" and outside: "pc' < pc \<or> pc+size(compE\<^sub>2 e) \<le> pc'"
-shows "\<not> matches_ex_entry P C pc' xe"
-(*<*)
-proof
- assume "matches_ex_entry P C pc' xe"
- with xe have "pc' \<in> pcs(compxE\<^sub>2 e pc d)"
- by(force simp add:matches_ex_entry_def pcs_def)
- with outside show False by simp
-qed
-(*>*)
-
-
-lemma [simp]:
-assumes xe: "xe \<in> set(compxEs\<^sub>2 es pc d)" and outside: "pc' < pc \<or> pc+size(compEs\<^sub>2 es) \<le> pc'"
-shows "\<not> matches_ex_entry P C pc' xe"
-(*<*)
-proof
- assume "matches_ex_entry P C pc' xe"
- with xe have "pc' \<in> pcs(compxEs\<^sub>2 es pc d)"
- by(force simp add:matches_ex_entry_def pcs_def)
- with outside show False by simp
-qed
-(*>*)
-
-
-lemma match_ex_table_app[simp]:
- "\<forall>xte \<in> set xt\<^sub>1. \<not> matches_ex_entry P D pc xte \<Longrightarrow>
- match_ex_table P D pc (xt\<^sub>1 @ xt) = match_ex_table P D pc xt"
-(*<*)by(induct xt\<^sub>1) simp_all(*>*)
-
-
-lemma [simp]:
- "\<forall>x \<in> set xtab. \<not> matches_ex_entry P C pc x \<Longrightarrow>
- match_ex_table P C pc xtab = None"
-(*<*)using match_ex_table_app[where ?xt = "[]"] by fastforce(*>*)
-
-
-lemma match_ex_entry:
- "matches_ex_entry P C pc (start, end, catch_type, handler) =
- (start \<le> pc \<and> pc < end \<and> P \<turnstile> C \<preceq>\<^sup>* catch_type)"
-(*<*)by(simp add:matches_ex_entry_def)(*>*)
-
-
-definition caught :: "jvm_prog \<Rightarrow> pc \<Rightarrow> heap \<Rightarrow> addr \<Rightarrow> ex_table \<Rightarrow> bool" where
- "caught P pc h a xt \<longleftrightarrow>
- (\<exists>entry \<in> set xt. matches_ex_entry P (cname_of h a) pc entry)"
-
-definition beforex :: "jvm_prog \<Rightarrow> cname \<Rightarrow> mname \<Rightarrow> ex_table \<Rightarrow> nat set \<Rightarrow> nat \<Rightarrow> bool"
- ("(2_,/_,/_ \<rhd>/ _ /'/ _,/_)" [51,0,0,0,0,51] 50) where
- "P,C,M \<rhd> xt / I,d \<longleftrightarrow>
- (\<exists>xt\<^sub>0 xt\<^sub>1. ex_table_of P C M = xt\<^sub>0 @ xt @ xt\<^sub>1 \<and> pcs xt\<^sub>0 \<inter> I = {} \<and> pcs xt \<subseteq> I \<and>
- (\<forall>pc \<in> I. \<forall>C pc' d'. match_ex_table P C pc xt\<^sub>1 = \<lfloor>(pc',d')\<rfloor> \<longrightarrow> d' \<le> d))"
-
-definition dummyx :: "jvm_prog \<Rightarrow> cname \<Rightarrow> mname \<Rightarrow> ex_table \<Rightarrow> nat set \<Rightarrow> nat \<Rightarrow> bool" ("(2_,_,_ \<triangleright>/ _ '/_,_)" [51,0,0,0,0,51] 50) where
- "P,C,M \<triangleright> xt/I,d \<longleftrightarrow> P,C,M \<rhd> xt/I,d"
-
-abbreviation
-"beforex\<^sub>0 P C M d I xt xt\<^sub>0 xt\<^sub>1
- \<equiv> ex_table_of P C M = xt\<^sub>0 @ xt @ xt\<^sub>1 \<and> pcs xt\<^sub>0 \<inter> I = {}
- \<and> pcs xt \<subseteq> I \<and> (\<forall>pc \<in> I. \<forall>C pc' d'. match_ex_table P C pc xt\<^sub>1 = \<lfloor>(pc',d')\<rfloor> \<longrightarrow> d' \<le> d)"
-
-lemma beforex_beforex\<^sub>0_eq:
- "P,C,M \<rhd> xt / I,d \<equiv> \<exists>xt\<^sub>0 xt\<^sub>1. beforex\<^sub>0 P C M d I xt xt\<^sub>0 xt\<^sub>1"
-using beforex_def by auto
-
-lemma beforexD1: "P,C,M \<rhd> xt / I,d \<Longrightarrow> pcs xt \<subseteq> I"
-(*<*)by(auto simp add:beforex_def)(*>*)
-
-
-lemma beforex_mono: "\<lbrakk> P,C,M \<rhd> xt/I,d'; d' \<le> d \<rbrakk> \<Longrightarrow> P,C,M \<rhd> xt/I,d"
-(*<*)by(fastforce simp:beforex_def)(*>*)
-
-
-lemma [simp]: "P,C,M \<rhd> xt/I,d \<Longrightarrow> P,C,M \<rhd> xt/I,Suc d"
-(*<*)by(fastforce intro:beforex_mono)(*>*)
-
-
-lemma beforex_append[simp]:
- "pcs xt\<^sub>1 \<inter> pcs xt\<^sub>2 = {} \<Longrightarrow>
- P,C,M \<rhd> xt\<^sub>1 @ xt\<^sub>2/I,d =
- (P,C,M \<rhd> xt\<^sub>1/I-pcs xt\<^sub>2,d \<and> P,C,M \<rhd> xt\<^sub>2/I-pcs xt\<^sub>1,d \<and> P,C,M \<triangleright> xt\<^sub>1@xt\<^sub>2/I,d)"
-(*<*)(is "?Q \<Longrightarrow> ?P = (?P1 \<and> ?P2 \<and> ?P3)" is "?Q \<Longrightarrow> ?P = ?P123")
-proof -
- assume pcs: ?Q
- show ?thesis proof(rule iffI)
- assume "?P123" then show ?P by(simp add:dummyx_def)
- next
- assume hyp: ?P
- let ?xt = "xt\<^sub>1 @ xt\<^sub>2"
- let ?beforex = "beforex\<^sub>0 P C M d"
- obtain xt\<^sub>0 xt\<^sub>1' where beforex: "?beforex I ?xt xt\<^sub>0 xt\<^sub>1'"
- using hyp by(clarsimp simp: beforex_def)
- have "\<exists>xt\<^sub>0 xt\<^sub>1'. ?beforex (I - pcs xt\<^sub>2) xt\<^sub>1 xt\<^sub>0 xt\<^sub>1'" \<comment> \<open>?P1\<close>
- using pcs beforex by(rule_tac x=xt\<^sub>0 in exI) auto
- moreover have "\<exists>xt\<^sub>0 xt\<^sub>1'. ?beforex (I - pcs xt\<^sub>1) xt\<^sub>2 xt\<^sub>0 xt\<^sub>1'" \<comment> \<open>?P2\<close>
- using pcs beforex by(rule_tac x="xt\<^sub>0@xt\<^sub>1" in exI) auto
- moreover have ?P3 using hyp by(simp add: dummyx_def)
- ultimately show ?P123 by (simp add: beforex_def)
- qed
-qed
-(*>*)
-
-
-lemma beforex_appendD1:
-assumes bx: "P,C,M \<rhd> xt\<^sub>1 @ xt\<^sub>2 @ [(f,t,D,h,d)] / I,d"
- and pcs: "pcs xt\<^sub>1 \<subseteq> J" and JI: "J \<subseteq> I" and Jpcs: "J \<inter> pcs xt\<^sub>2 = {}"
-shows "P,C,M \<rhd> xt\<^sub>1 / J,d"
-(*<*)
-proof -
- let ?beforex = "beforex\<^sub>0 P C M d"
- obtain xt\<^sub>0 xt\<^sub>1' where bx': "?beforex I (xt\<^sub>1 @ xt\<^sub>2 @ [(f,t,D,h,d)]) xt\<^sub>0 xt\<^sub>1'"
- using bx by(clarsimp simp:beforex_def)
- let ?xt0 = xt\<^sub>0 and ?xt1 = "xt\<^sub>2 @ (f, t, D, h, d) # xt\<^sub>1'"
- have "pcs xt\<^sub>0 \<inter> J = {}" using bx' JI by blast
- moreover {
- fix pc C pc' d' assume pcJ: "pc\<in>J"
- then have "pc \<notin> pcs xt\<^sub>2" using bx' Jpcs by blast
- then have "match_ex_table P C pc (xt\<^sub>2 @ (f, t, D, h, d) # xt\<^sub>1')
- = \<lfloor>(pc', d')\<rfloor> \<longrightarrow> d' \<le> d"
- using bx' JI pcJ by (auto split:if_split_asm)
- }
- ultimately have "?beforex J xt\<^sub>1 ?xt0 ?xt1" using bx' pcs by simp
- then show ?thesis using beforex_def by blast
-qed
-(*>*)
-
-
-lemma beforex_appendD2:
-assumes bx: "P,C,M \<rhd> xt\<^sub>1 @ xt\<^sub>2 @ [(f,t,D,h,d)] / I,d"
- and pcs: "pcs xt\<^sub>2 \<subseteq> J" and JI: "J \<subseteq> I" and Jpcs: "J \<inter> pcs xt\<^sub>1 = {}"
-shows "P,C,M \<rhd> xt\<^sub>2 / J,d"
-(*<*)
-proof -
- let ?beforex = "beforex\<^sub>0 P C M d"
- obtain xt\<^sub>0 xt\<^sub>1' where bx': "?beforex I (xt\<^sub>1 @ xt\<^sub>2 @ [(f,t,D,h,d)]) xt\<^sub>0 xt\<^sub>1'"
- using bx by(clarsimp simp:beforex_def)
- then have "\<exists>xt\<^sub>1''. beforex\<^sub>0 P C M d J xt\<^sub>2 (xt\<^sub>0 @ xt\<^sub>1) xt\<^sub>1''"
- using assms by fastforce
- then show ?thesis using beforex_def by blast
-qed
-(*>*)
-
-
-lemma beforexM:
- "P \<turnstile> C sees M,b: Ts\<rightarrow>T = body in D \<Longrightarrow> compP\<^sub>2 P,D,M \<rhd> compxE\<^sub>2 body 0 0/{..<size(compE\<^sub>2 body)},0"
-(*<*)
-proof -
- assume cM: "P \<turnstile> C sees M,b: Ts\<rightarrow>T = body in D"
- let ?xt0 = "[]"
- have "\<exists>xt1. beforex\<^sub>0 (compP\<^sub>2 P) D M 0 ({..<size(compE\<^sub>2 body)}) (compxE\<^sub>2 body 0 0) ?xt0 xt1"
- using sees_method_compP[where f = compMb\<^sub>2, OF sees_method_idemp[OF cM]]
- pcs_subset by(fastforce simp add: compP\<^sub>2_def compMb\<^sub>2_def)
- then show ?thesis using beforex_def by blast
-qed
-(*>*)
-
-
-lemma match_ex_table_SomeD2:
-assumes met: "match_ex_table P D pc (ex_table_of P C M) = \<lfloor>(pc',d')\<rfloor>"
- and bx: "P,C,M \<rhd> xt/I,d"
- and nmet: "\<forall>x \<in> set xt. \<not> matches_ex_entry P D pc x" and pcI: "pc \<in> I"
-shows "d' \<le> d"
-(*<*)
-proof -
- obtain xt\<^sub>0 xt\<^sub>1 where bx': "beforex\<^sub>0 P C M d I xt xt\<^sub>0 xt\<^sub>1"
- using bx by(clarsimp simp:beforex_def)
- then have "pc \<notin> pcs xt\<^sub>0" using pcI by blast
- then show ?thesis using bx' met nmet pcI by simp
-qed
-(*>*)
-
-
-lemma match_ex_table_SomeD1:
- "\<lbrakk> match_ex_table P D pc (ex_table_of P C M) = \<lfloor>(pc',d')\<rfloor>;
- P,C,M \<rhd> xt / I,d; pc \<in> I; pc \<notin> pcs xt \<rbrakk> \<Longrightarrow> d' \<le> d"
-(*<*)by(auto elim: match_ex_table_SomeD2)(*>*)
-
-
-subsection\<open> The correctness proof \<close>
-
-(*<*)
-declare nat_add_distrib[simp] caught_def[simp]
-declare fun_upd_apply[simp del]
-(*>*)
-
-definition
- handle :: "jvm_prog \<Rightarrow> cname \<Rightarrow> mname \<Rightarrow> addr \<Rightarrow> heap \<Rightarrow> val list \<Rightarrow> val list \<Rightarrow> nat \<Rightarrow> init_call_status \<Rightarrow> frame list \<Rightarrow> sheap
- \<Rightarrow> jvm_state" where
- "handle P C M a h vs ls pc ics frs sh = find_handler P a h ((vs,ls,C,M,pc,ics) # frs) sh"
-
-lemma aux_isin[simp]: "\<lbrakk> B \<subseteq> A; a \<in> B \<rbrakk> \<Longrightarrow> a \<in> A"
-(*<*)by blast(*>*)
-
-lemma handle_frs_tl_neq:
- "ics_of f \<noteq> No_ics
- \<Longrightarrow> (xp, h, f#frs, sh) \<noteq> handle P C M xa h' vs l pc ics frs sh'"
- by(simp add: handle_def find_handler_frs_tl_neq del: find_handler.simps)
-
-subsubsection "Correctness proof inductive hypothesis"
-
-\<comment> \<open> frame definitions for use by correctness proof inductive hypothesis \<close>
-fun calling_to_called :: "frame \<Rightarrow> frame" where
-"calling_to_called (stk,loc,D,M,pc,ics) = (stk,loc,D,M,pc,case ics of Calling C Cs \<Rightarrow> Called (C#Cs))"
-
-fun calling_to_scalled :: "frame \<Rightarrow> frame" where
-"calling_to_scalled (stk,loc,D,M,pc,ics) = (stk,loc,D,M,pc,case ics of Calling C Cs \<Rightarrow> Called Cs)"
-
-fun calling_to_calling :: "frame \<Rightarrow> cname \<Rightarrow> frame" where
-"calling_to_calling (stk,loc,D,M,pc,ics) C' = (stk,loc,D,M,pc,case ics of Calling C Cs \<Rightarrow> Calling C' (C#Cs))"
-
-fun calling_to_throwing :: "frame \<Rightarrow> addr \<Rightarrow> frame" where
-"calling_to_throwing (stk,loc,D,M,pc,ics) a = (stk,loc,D,M,pc,case ics of Calling C Cs \<Rightarrow> Throwing (C#Cs) a)"
-
-fun calling_to_sthrowing :: "frame \<Rightarrow> addr \<Rightarrow> frame" where
-"calling_to_sthrowing (stk,loc,D,M,pc,ics) a = (stk,loc,D,M,pc,case ics of Calling C Cs \<Rightarrow> Throwing Cs a)"
-
-
-\<comment> \<open> pieces of the correctness proof's inductive hypothesis, which depend on the
- expression being compiled) \<close>
-
-fun Jcc_cond :: "J\<^sub>1_prog \<Rightarrow> ty list \<Rightarrow> cname \<Rightarrow> mname \<Rightarrow> val list \<Rightarrow> pc \<Rightarrow> init_call_status
- \<Rightarrow> nat set \<Rightarrow> heap \<Rightarrow> sheap \<Rightarrow> expr\<^sub>1 \<Rightarrow> bool" where
-"Jcc_cond P E C M vs pc ics I h sh (INIT C\<^sub>0 (Cs,b) \<leftarrow> e')
- = ((\<exists>T. P,E,h,sh \<turnstile>\<^sub>1 INIT C\<^sub>0 (Cs,b) \<leftarrow> e' : T) \<and> unit = e' \<and> ics = No_ics)" |
-"Jcc_cond P E C M vs pc ics I h sh (RI(C',e\<^sub>0);Cs \<leftarrow> e')
- = (((e\<^sub>0 = C'\<bullet>\<^sub>sclinit([]) \<and> (\<exists>T. P,E,h,sh \<turnstile>\<^sub>1 RI(C',e\<^sub>0);Cs \<leftarrow> e':T))
- \<or> ((\<exists>a. e\<^sub>0 = Throw a) \<and> (\<forall>C \<in> set(C'#Cs). is_class P C)))
- \<and> unit = e' \<and> ics = No_ics)" |
-"Jcc_cond P E C M vs pc ics I h sh (C'\<bullet>\<^sub>sM'(es))
- = (let e = (C'\<bullet>\<^sub>sM'(es))
- in if M' = clinit \<and> es = [] then (\<exists>T. P,E,h,sh \<turnstile>\<^sub>1 e:T) \<and> (\<exists>Cs. ics = Called Cs)
- else (compP\<^sub>2 P,C,M,pc \<rhd> compE\<^sub>2 e \<and> compP\<^sub>2 P,C,M \<rhd> compxE\<^sub>2 e pc (size vs)/I,size vs
- \<and> {pc..<pc+size(compE\<^sub>2 e)} \<subseteq> I \<and> \<not>sub_RI e \<and> ics = No_ics)
- )" |
-"Jcc_cond P E C M vs pc ics I h sh e
- = (compP\<^sub>2 P,C,M,pc \<rhd> compE\<^sub>2 e \<and> compP\<^sub>2 P,C,M \<rhd> compxE\<^sub>2 e pc (size vs)/I,size vs
- \<and> {pc..<pc+size(compE\<^sub>2 e)} \<subseteq> I \<and> \<not>sub_RI e \<and> ics = No_ics)"
-
-
-fun Jcc_frames :: "jvm_prog \<Rightarrow> cname \<Rightarrow> mname \<Rightarrow> val list \<Rightarrow> val list \<Rightarrow> pc \<Rightarrow> init_call_status
- \<Rightarrow> frame list \<Rightarrow> expr\<^sub>1 \<Rightarrow> frame list" where
-"Jcc_frames P C M vs ls pc ics frs (INIT C\<^sub>0 (C'#Cs,b) \<leftarrow> e')
- = (case b of False \<Rightarrow> (vs,ls,C,M,pc,Calling C' Cs) # frs
- | True \<Rightarrow> (vs,ls,C,M,pc,Called (C'#Cs)) # frs
- )" |
-"Jcc_frames P C M vs ls pc ics frs (INIT C\<^sub>0 (Nil,b) \<leftarrow> e')
- = (vs,ls,C,M,pc,Called [])#frs" |
-"Jcc_frames P C M vs ls pc ics frs (RI(C',e\<^sub>0);Cs \<leftarrow> e')
- = (case e\<^sub>0 of Throw a \<Rightarrow> (vs,ls,C,M,pc,Throwing (C'#Cs) a) # frs
- | _ \<Rightarrow> (vs,ls,C,M,pc,Called (C'#Cs)) # frs )" |
-"Jcc_frames P C M vs ls pc ics frs (C'\<bullet>\<^sub>sM'(es))
- = (if M' = clinit \<and> es = []
- then create_init_frame P C'#(vs,ls,C,M,pc,ics)#frs
- else (vs,ls,C,M,pc,ics)#frs
- )" |
-"Jcc_frames P C M vs ls pc ics frs e
- = (vs,ls,C,M,pc,ics)#frs"
-
-fun Jcc_rhs :: "J\<^sub>1_prog \<Rightarrow> ty list \<Rightarrow> cname \<Rightarrow> mname \<Rightarrow> val list \<Rightarrow> val list \<Rightarrow> pc \<Rightarrow> init_call_status
- \<Rightarrow> frame list \<Rightarrow> heap \<Rightarrow> val list \<Rightarrow> sheap \<Rightarrow> val \<Rightarrow> expr\<^sub>1 \<Rightarrow> jvm_state" where
-"Jcc_rhs P E C M vs ls pc ics frs h' ls' sh' v (INIT C\<^sub>0 (Cs,b) \<leftarrow> e')
- = (None,h',(vs,ls,C,M,pc,Called [])#frs,sh')" |
-"Jcc_rhs P E C M vs ls pc ics frs h' ls' sh' v (RI(C',e\<^sub>0);Cs \<leftarrow> e')
- = (None,h',(vs,ls,C,M,pc,Called [])#frs,sh')" |
-"Jcc_rhs P E C M vs ls pc ics frs h' ls' sh' v (C'\<bullet>\<^sub>sM'(es))
- = (let e = (C'\<bullet>\<^sub>sM'(es))
- in if M' = clinit \<and> es = []
- then (None,h',(vs,ls,C,M,pc,ics)#frs,sh'(C'\<mapsto>(fst(the(sh' C')),Done)))
- else (None,h',(v#vs,ls',C,M,pc+size(compE\<^sub>2 e),ics)#frs,sh')
- )" |
-"Jcc_rhs P E C M vs ls pc ics frs h' ls' sh' v e
- = (None,h',(v#vs,ls',C,M,pc+size(compE\<^sub>2 e),ics)#frs,sh')"
-
-fun Jcc_err :: "jvm_prog \<Rightarrow> cname \<Rightarrow> mname \<Rightarrow> heap \<Rightarrow> val list \<Rightarrow> val list \<Rightarrow> pc \<Rightarrow> init_call_status
- \<Rightarrow> frame list \<Rightarrow> sheap \<Rightarrow> nat set \<Rightarrow> heap \<Rightarrow> val list \<Rightarrow> sheap \<Rightarrow> addr \<Rightarrow> expr\<^sub>1
- \<Rightarrow> bool" where
-"Jcc_err P C M h vs ls pc ics frs sh I h' ls' sh' xa (INIT C\<^sub>0 (Cs,b) \<leftarrow> e')
- = (\<exists>vs'. P \<turnstile> (None,h,Jcc_frames P C M vs ls pc ics frs (INIT C\<^sub>0 (Cs,b) \<leftarrow> e'),sh)
- -jvm\<rightarrow> handle P C M xa h' (vs'@vs) ls pc ics frs sh')" |
-"Jcc_err P C M h vs ls pc ics frs sh I h' ls' sh' xa (RI(C',e\<^sub>0);Cs \<leftarrow> e')
- = (\<exists>vs'. P \<turnstile> (None,h,Jcc_frames P C M vs ls pc ics frs (RI(C',e\<^sub>0);Cs \<leftarrow> e'),sh)
- -jvm\<rightarrow> handle P C M xa h' (vs'@vs) ls pc ics frs sh')" |
-"Jcc_err P C M h vs ls pc ics frs sh I h' ls' sh' xa (C'\<bullet>\<^sub>sM'(es))
- = (let e = (C'\<bullet>\<^sub>sM'(es))
- in if M' = clinit \<and> es = []
- then case ics of
- Called Cs \<Rightarrow> P \<turnstile> (None,h,Jcc_frames P C M vs ls pc ics frs e,sh)
- -jvm\<rightarrow> (None,h',(vs,ls,C,M,pc,Throwing Cs xa)#frs,(sh'(C' \<mapsto> (fst(the(sh' C')),Error))))
- else (\<exists>pc\<^sub>1. pc \<le> pc\<^sub>1 \<and> pc\<^sub>1 < pc + size(compE\<^sub>2 e) \<and>
- \<not> caught P pc\<^sub>1 h' xa (compxE\<^sub>2 e pc (size vs)) \<and>
- (\<exists>vs'. P \<turnstile> (None,h,Jcc_frames P C M vs ls pc ics frs e,sh)
- -jvm\<rightarrow> handle P C M xa h' (vs'@vs) ls' pc\<^sub>1 ics frs sh'))
- )" |
-"Jcc_err P C M h vs ls pc ics frs sh I h' ls' sh' xa e
- = (\<exists>pc\<^sub>1. pc \<le> pc\<^sub>1 \<and> pc\<^sub>1 < pc + size(compE\<^sub>2 e) \<and>
- \<not> caught P pc\<^sub>1 h' xa (compxE\<^sub>2 e pc (size vs)) \<and>
- (\<exists>vs'. P \<turnstile> (None,h,Jcc_frames P C M vs ls pc ics frs e,sh)
- -jvm\<rightarrow> handle P C M xa h' (vs'@vs) ls' pc\<^sub>1 ics frs sh'))"
-
-fun Jcc_pieces :: "J\<^sub>1_prog \<Rightarrow> ty list \<Rightarrow> cname \<Rightarrow> mname \<Rightarrow> heap \<Rightarrow> val list \<Rightarrow> val list \<Rightarrow> pc \<Rightarrow> init_call_status
- \<Rightarrow> frame list \<Rightarrow> sheap \<Rightarrow> nat set \<Rightarrow> heap \<Rightarrow> val list \<Rightarrow> sheap \<Rightarrow> val \<Rightarrow> addr \<Rightarrow> expr\<^sub>1
- \<Rightarrow> bool \<times> frame list \<times> jvm_state \<times> bool" where
-"Jcc_pieces P E C M h vs ls pc ics frs sh I h' ls' sh' v xa e
- = (Jcc_cond P E C M vs pc ics I h sh e, Jcc_frames (compP\<^sub>2 P) C M vs ls pc ics frs e,
- Jcc_rhs P E C M vs ls pc ics frs h' ls' sh' v e,
- Jcc_err (compP\<^sub>2 P) C M h vs ls pc ics frs sh I h' ls' sh' xa e)"
-
-\<comment> \<open> @{text Jcc_pieces} lemmas \<close>
-
-lemma nsub_RI_Jcc_pieces:
-assumes [simp]: "P \<equiv> compP\<^sub>2 P\<^sub>1"
- and nsub: "\<not>sub_RI e"
-shows "Jcc_pieces P\<^sub>1 E C M h vs ls pc ics frs sh I h' ls' sh' v xa e
- = (let cond = P,C,M,pc \<rhd> compE\<^sub>2 e \<and> P,C,M \<rhd> compxE\<^sub>2 e pc (size vs)/I,size vs
- \<and> {pc..<pc+size(compE\<^sub>2 e)} \<subseteq> I \<and> ics = No_ics;
- frs' = (vs,ls,C,M,pc,ics)#frs;
- rhs = (None,h',(v#vs,ls',C,M,pc+size(compE\<^sub>2 e),ics)#frs,sh');
- err = (\<exists>pc\<^sub>1. pc \<le> pc\<^sub>1 \<and> pc\<^sub>1 < pc + size(compE\<^sub>2 e) \<and>
- \<not> caught P pc\<^sub>1 h' xa (compxE\<^sub>2 e pc (size vs)) \<and>
- (\<exists>vs'. P \<turnstile> (None,h,frs',sh) -jvm\<rightarrow> handle P C M xa h' (vs'@vs) ls' pc\<^sub>1 ics frs sh'))
- in (cond, frs',rhs, err)
- )"
-proof -
- have NC: "\<forall>C'. e \<noteq> C'\<bullet>\<^sub>sclinit([])" using assms(2) proof(cases e) qed(simp_all)
- then show ?thesis using assms
- proof(cases e)
- case (SCall C M es)
- then have "M \<noteq> clinit" using nsub by simp
- then show ?thesis using SCall nsub proof(cases es) qed(simp_all)
- qed(simp_all)
-qed
-
-lemma Jcc_pieces_Cast:
-assumes [simp]: "P \<equiv> compP\<^sub>2 P\<^sub>1"
- and "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (Cast C' e)
- = (True, frs\<^sub>0, (xp',h',(v#vs',ls',C\<^sub>0,M',pc',ics')#frs',sh'), err)"
-shows "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v' xa e
- = (True, frs\<^sub>0, (xp',h',(v'#vs',ls',C\<^sub>0,M',pc' - 1,ics')#frs',sh'),
- (\<exists>pc\<^sub>1. pc \<le> pc\<^sub>1 \<and> pc\<^sub>1 < pc + size(compE\<^sub>2 e) \<and>
- \<not> caught P pc\<^sub>1 h\<^sub>1 xa (compxE\<^sub>2 e pc (size vs)) \<and>
- (\<exists>vs'. P \<turnstile> (None,h\<^sub>0,frs\<^sub>0,sh\<^sub>0) -jvm\<rightarrow> handle P C M xa h\<^sub>1 (vs'@vs) ls\<^sub>1 pc\<^sub>1 ics frs sh\<^sub>1)))"
-proof -
- have pc: "{pc..<pc + length (compE\<^sub>2 e)} \<subseteq> I" using assms by clarsimp
- show ?thesis using assms nsub_RI_Jcc_pieces[where e=e] pc by clarsimp
-qed
-
-lemma Jcc_pieces_BinOp1:
-assumes
- "Jcc_pieces P E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v xa (e \<guillemotleft>bop\<guillemotright> e')
- = (True, frs\<^sub>0, (xp',h',(v#vs',ls',C\<^sub>0,M',pc',ics')#frs',sh'), err)"
-shows "\<exists>err. Jcc_pieces P E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0
- (I - pcs (compxE\<^sub>2 e' (pc + length (compE\<^sub>2 e)) (Suc (length vs')))) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v' xa e
- = (True, frs\<^sub>0, (xp',h\<^sub>1,(v'#vs',ls\<^sub>1,C\<^sub>0,M',pc' - size (compE\<^sub>2 e') - 1,ics')#frs',sh\<^sub>1), err)"
-proof -
- have bef: "compP compMb\<^sub>2 P,C\<^sub>0,M' \<rhd> compxE\<^sub>2 e pc (length vs)
- / I - pcs (compxE\<^sub>2 e' (pc + length (compE\<^sub>2 e)) (Suc (length vs'))),length vs"
- using assms by clarsimp
- have vs: "vs = vs'" using assms by simp
- show ?thesis using assms nsub_RI_Jcc_pieces[where e=e] bef vs by clarsimp
-qed
-
-lemma Jcc_pieces_BinOp2:
-assumes [simp]: "P \<equiv> compP\<^sub>2 P\<^sub>1"
- and "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v xa (e \<guillemotleft>bop\<guillemotright> e')
- = (True, frs\<^sub>0, (xp',h',(v#vs',ls',C\<^sub>0,M',pc',ics')#frs',sh'), err)"
-shows "\<exists>err. Jcc_pieces P\<^sub>1 E C M h\<^sub>1 (v\<^sub>1#vs) ls\<^sub>1 (pc + size (compE\<^sub>2 e)) ics frs sh\<^sub>1
- (I - pcs (compxE\<^sub>2 e pc (length vs'))) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v' xa e'
- = (True, (v\<^sub>1#vs,ls\<^sub>1,C,M,pc + size (compE\<^sub>2 e),ics)#frs,
- (xp',h',(v'#v\<^sub>1#vs',ls',C\<^sub>0,M',pc' - 1,ics')#frs',sh'),
- (\<exists>pc\<^sub>1. pc + size (compE\<^sub>2 e) \<le> pc\<^sub>1 \<and> pc\<^sub>1 < pc + size (compE\<^sub>2 e) + length (compE\<^sub>2 e') \<and>
- \<not> caught P pc\<^sub>1 h\<^sub>2 xa (compxE\<^sub>2 e' (pc + size (compE\<^sub>2 e)) (Suc (length vs))) \<and>
- (\<exists>vs'. P \<turnstile> (None,h\<^sub>1,(v\<^sub>1#vs,ls\<^sub>1,C,M,pc + size (compE\<^sub>2 e),ics)#frs,sh\<^sub>1)
- -jvm\<rightarrow> handle P C M xa h\<^sub>2 (vs'@v\<^sub>1#vs) ls\<^sub>2 pc\<^sub>1 ics frs sh\<^sub>2)))"
-proof -
- have bef: "compP compMb\<^sub>2 P\<^sub>1,C\<^sub>0,M' \<rhd> compxE\<^sub>2 e pc (length vs)
- / I - pcs (compxE\<^sub>2 e' (pc + length (compE\<^sub>2 e)) (Suc (length vs'))),length vs"
- using assms by clarsimp
- have vs: "vs = vs'" using assms by simp
- show ?thesis using assms nsub_RI_Jcc_pieces[where e=e'] bef vs by clarsimp
-qed
-
-lemma Jcc_pieces_FAcc:
-assumes
- "Jcc_pieces P E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (e\<bullet>F{D})
- = (True, frs\<^sub>0, (xp',h',(v#vs',ls',C\<^sub>0,M',pc',ics')#frs',sh'), err)"
-shows "\<exists>err. Jcc_pieces P E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v' xa e
- = (True, frs\<^sub>0, (xp',h',(v'#vs',ls',C\<^sub>0,M',pc' - 1,ics')#frs',sh'), err)"
-proof -
- have pc: "{pc..<pc + length (compE\<^sub>2 e)} \<subseteq> I" using assms by clarsimp
- then show ?thesis using assms nsub_RI_Jcc_pieces[where e=e] by clarsimp
-qed
-
-lemma Jcc_pieces_LAss:
-assumes [simp]: "P \<equiv> compP\<^sub>2 P\<^sub>1"
- and "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (i:=e)
- = (True, frs\<^sub>0, (xp',h',(v#vs',ls',C\<^sub>0,M',pc',ics')#frs',sh'), err)"
-shows "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v' xa e
- = (True, frs\<^sub>0, (xp',h',(v'#vs',ls',C\<^sub>0,M',pc' - 2,ics')#frs',sh'),
- (\<exists>pc\<^sub>1. pc \<le> pc\<^sub>1 \<and> pc\<^sub>1 < pc + size(compE\<^sub>2 e) \<and>
- \<not> caught P pc\<^sub>1 h\<^sub>1 xa (compxE\<^sub>2 e pc (size vs)) \<and>
- (\<exists>vs'. P \<turnstile> (None,h\<^sub>0,frs\<^sub>0,sh\<^sub>0) -jvm\<rightarrow> handle P C M xa h\<^sub>1 (vs'@vs) ls\<^sub>1 pc\<^sub>1 ics frs sh\<^sub>1)))"
-proof -
- have pc: "{pc..<pc + length (compE\<^sub>2 e)} \<subseteq> I" using assms by clarsimp
- show ?thesis using assms nsub_RI_Jcc_pieces[where e=e] pc by clarsimp
-qed
-
-lemma Jcc_pieces_FAss1:
-assumes
- "Jcc_pieces P E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v xa (e\<bullet>F{D}:=e')
- = (True, frs\<^sub>0, (xp',h',(v#vs',ls',C\<^sub>0,M',pc',ics')#frs',sh'), err)"
-shows "\<exists>err. Jcc_pieces P E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0
- (I - pcs (compxE\<^sub>2 e' (pc + length (compE\<^sub>2 e)) (Suc (length vs')))) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v' xa e
- = (True, frs\<^sub>0, (xp',h\<^sub>1,(v'#vs',ls\<^sub>1,C\<^sub>0,M',pc' - size (compE\<^sub>2 e') - 2,ics')#frs',sh\<^sub>1), err)"
-proof -
- show ?thesis using assms nsub_RI_Jcc_pieces[where e=e] by clarsimp
-qed
-
-lemma Jcc_pieces_FAss2:
-assumes
- "Jcc_pieces P E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v xa (e\<bullet>F{D}:=e')
- = (True, frs\<^sub>0, (xp',h',(v#vs',ls',C\<^sub>0,M',pc',ics')#frs',sh'), err)"
-shows "Jcc_pieces P E C M h\<^sub>1 (v\<^sub>1#vs) ls\<^sub>1 (pc + size (compE\<^sub>2 e)) ics frs sh\<^sub>1
- (I - pcs (compxE\<^sub>2 e pc (length vs'))) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v' xa e'
- = (True, (v\<^sub>1#vs,ls\<^sub>1,C,M,pc + size (compE\<^sub>2 e),ics)#frs,
- (xp',h',(v'#v\<^sub>1#vs',ls',C\<^sub>0,M',pc' - 2,ics')#frs',sh'),
- (\<exists>pc\<^sub>1. (pc + size (compE\<^sub>2 e)) \<le> pc\<^sub>1 \<and> pc\<^sub>1 < pc + size (compE\<^sub>2 e) + size(compE\<^sub>2 e') \<and>
- \<not> caught (compP\<^sub>2 P) pc\<^sub>1 h\<^sub>2 xa (compxE\<^sub>2 e' (pc + size (compE\<^sub>2 e)) (size (v\<^sub>1#vs))) \<and>
- (\<exists>vs'. (compP\<^sub>2 P) \<turnstile> (None,h\<^sub>1,(v\<^sub>1#vs,ls\<^sub>1,C,M,pc + size (compE\<^sub>2 e),ics)#frs,sh\<^sub>1)
- -jvm\<rightarrow> handle (compP\<^sub>2 P) C M xa h\<^sub>2 (vs'@v\<^sub>1#vs) ls\<^sub>2 pc\<^sub>1 ics frs sh\<^sub>2)))"
-proof -
- show ?thesis using assms nsub_RI_Jcc_pieces[where e=e'] by clarsimp
-qed
-
-lemma Jcc_pieces_SFAss:
-assumes
- "Jcc_pieces P E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h' ls' sh' v xa (C'\<bullet>\<^sub>sF{D}:=e)
- = (True, frs\<^sub>0, (xp',h',(v#vs',ls',C\<^sub>0,M',pc',ics')#frs',sh'), err)"
-shows "\<exists>err. Jcc_pieces P E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v' xa e
- = (True, frs\<^sub>0, (xp',h\<^sub>1,(v'#vs',ls\<^sub>1,C\<^sub>0,M',pc' - 2,ics')#frs',sh\<^sub>1), err)"
-proof -
- have pc: "{pc..<pc + length (compE\<^sub>2 e)} \<subseteq> I" using assms by clarsimp
- show ?thesis using assms nsub_RI_Jcc_pieces[where e=e] pc by clarsimp
-qed
-
-lemma Jcc_pieces_Call1:
-assumes
- "Jcc_pieces P E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>3 ls\<^sub>3 sh\<^sub>3 v xa (e\<bullet>M\<^sub>0(es))
- = (True, frs\<^sub>0, (xp',h',(v#vs',ls',C',M',pc',ics')#frs',sh'), err)"
-shows "\<exists>err. Jcc_pieces P E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0
- (I - pcs (compxEs\<^sub>2 es (pc + length (compE\<^sub>2 e)) (Suc (length vs')))) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v' xa e
- = (True, frs\<^sub>0,
- (xp',h\<^sub>1,(v'#vs',ls\<^sub>1,C',M',pc' - size (compEs\<^sub>2 es) - 1,ics')#frs',sh\<^sub>1), err)"
-proof -
- show ?thesis using assms nsub_RI_Jcc_pieces[where e=e] by clarsimp
-qed
-
-lemma Jcc_pieces_clinit:
-assumes [simp]: "P \<equiv> compP\<^sub>2 P\<^sub>1"
- and cond: "Jcc_cond P\<^sub>1 E C M vs pc ics I h sh (C1\<bullet>\<^sub>sclinit([]))"
-shows "Jcc_pieces P\<^sub>1 E C M h vs ls pc ics frs sh I h' ls' sh' v xa (C1\<bullet>\<^sub>sclinit([]))
- = (True, create_init_frame P C1 # (vs,ls,C,M,pc,ics)#frs,
- (None, h', (vs,ls,C,M,pc,ics)#frs, sh'(C1\<mapsto>(fst(the(sh' C1)),Done))),
- P \<turnstile> (None,h,create_init_frame P C1 # (vs,ls,C,M,pc,ics)#frs,sh) -jvm\<rightarrow>
- (case ics of Called Cs \<Rightarrow> (None,h',(vs,ls,C,M,pc,Throwing Cs xa)#frs,(sh'(C1 \<mapsto> (fst(the(sh' C1)),Error))))))"
-using assms by(auto split: init_call_status.splits list.splits bool.splits)
-
-lemma Jcc_pieces_SCall_clinit_body:
-assumes [simp]: "P \<equiv> compP\<^sub>2 P\<^sub>1" and wf: "wf_J\<^sub>1_prog P\<^sub>1"
- and "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>3 ls\<^sub>2 sh\<^sub>3 v xa (C1\<bullet>\<^sub>sclinit([]))
- = (True, frs', rhs', err')"
- and method: "P\<^sub>1 \<turnstile> C1 sees clinit,Static: []\<rightarrow>Void = body in D"
-shows "Jcc_pieces P\<^sub>1 [] D clinit h\<^sub>2 [] (replicate (max_vars body) undefined) 0
- No_ics (tl frs') sh\<^sub>2 {..<length (compE\<^sub>2 body)} h\<^sub>3 ls\<^sub>3 sh\<^sub>3 v xa body
- = (True, frs',
- (None,h\<^sub>3,([v],ls\<^sub>3,D,clinit,size(compE\<^sub>2 body), No_ics)#tl frs',sh\<^sub>3),
- \<exists>pc\<^sub>1. 0 \<le> pc\<^sub>1 \<and> pc\<^sub>1 < size(compE\<^sub>2 body) \<and>
- \<not> caught P pc\<^sub>1 h\<^sub>3 xa (compxE\<^sub>2 body 0 0) \<and>
- (\<exists>vs'. P \<turnstile> (None,h\<^sub>2,frs',sh\<^sub>2) -jvm\<rightarrow> handle P D clinit xa h\<^sub>3 vs' ls\<^sub>3 pc\<^sub>1
- No_ics (tl frs') sh\<^sub>3))"
-proof -
- have M_in_D: "P\<^sub>1 \<turnstile> D sees clinit,Static: []\<rightarrow>Void = body in D"
- using method by(rule sees_method_idemp)
- hence M_code: "compP\<^sub>2 P\<^sub>1,D,clinit,0 \<rhd> compE\<^sub>2 body @ [Return]"
- and M_xtab: "compP\<^sub>2 P\<^sub>1,D,clinit \<rhd> compxE\<^sub>2 body 0 0/{..<size(compE\<^sub>2 body)},0"
- by(rule beforeM, rule beforexM)
- have nsub: "\<not>sub_RI body" by(rule sees_wf\<^sub>1_nsub_RI[OF wf method])
- then show ?thesis using assms nsub_RI_Jcc_pieces M_code M_xtab by clarsimp
-qed
-
-lemma Jcc_pieces_Cons:
-assumes [simp]: "P \<equiv> compP\<^sub>2 P\<^sub>1"
- and "P,C,M,pc \<rhd> compEs\<^sub>2 (e#es)" and "P,C,M \<rhd> compxEs\<^sub>2 (e#es) pc (size vs)/I,size vs"
- and "{pc..<pc+size(compEs\<^sub>2 (e#es))} \<subseteq> I"
- and "ics = No_ics"
- and "\<not>sub_RIs (e#es)"
-shows "Jcc_pieces P\<^sub>1 E C M h vs ls pc ics frs sh
- (I - pcs (compxEs\<^sub>2 es (pc + length (compE\<^sub>2 e)) (Suc (length vs)))) h' ls' sh' v xa e
- = (True, (vs, ls, C, M, pc, ics) # frs,
- (None, h', (v#vs, ls', C, M, pc + length (compE\<^sub>2 e), ics) # frs, sh'),
- \<exists>pc\<^sub>1\<ge>pc. pc\<^sub>1 < pc + length (compE\<^sub>2 e) \<and> \<not> caught P pc\<^sub>1 h' xa (compxE\<^sub>2 e pc (length vs))
- \<and> (\<exists>vs'. P \<turnstile> (None, h, (vs, ls, C, M, pc, ics) # frs, sh)
- -jvm\<rightarrow> handle P C M xa h' (vs'@vs) ls' pc\<^sub>1 ics frs sh'))"
-proof -
- show ?thesis using assms nsub_RI_Jcc_pieces[where e=e] by auto
-qed
-
-lemma Jcc_pieces_InitNone:
-assumes [simp]: "P \<equiv> compP\<^sub>2 P\<^sub>1"
- and "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' sh' v xa (INIT C' (C\<^sub>0 # Cs,False) \<leftarrow> e)
- = (True, frs', (None, h', (vs, l, C, M, pc, Called []) # frs, sh'), err)"
-shows
- "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs (sh(C\<^sub>0 \<mapsto> (sblank P C\<^sub>0, Prepared)))
- I h' l' sh' v xa (INIT C' (C\<^sub>0 # Cs,False) \<leftarrow> e)
- = (True, frs', (None, h', (vs, l, C, M, pc, Called []) # frs, sh'),
- \<exists>vs'. P \<turnstile> (None,h,frs',(sh(C\<^sub>0 \<mapsto> (sblank P\<^sub>1 C\<^sub>0, Prepared))))
- -jvm\<rightarrow> handle P C M xa h' (vs'@vs) l pc ics frs sh')"
-proof -
- have "Jcc_cond P\<^sub>1 E C M vs pc ics I h sh (INIT C' (C\<^sub>0 # Cs,False) \<leftarrow> e)" using assms by simp
- then obtain T where "P\<^sub>1,E,h,sh \<turnstile>\<^sub>1 INIT C' (C\<^sub>0 # Cs,False) \<leftarrow> unit : T" by fastforce
- then have "P\<^sub>1,E,h,sh(C\<^sub>0 \<mapsto> (sblank P\<^sub>1 C\<^sub>0, Prepared)) \<turnstile>\<^sub>1 INIT C' (C\<^sub>0 # Cs,False) \<leftarrow> unit : T"
- by(auto simp: fun_upd_apply)
- then have "Ex (WTrt2\<^sub>1 P\<^sub>1 E h (sh(C\<^sub>0 \<mapsto> (sblank P\<^sub>1 C\<^sub>0, Prepared))) (INIT C' (C\<^sub>0 # Cs,False) \<leftarrow> unit))"
- by(simp only: exI)
- then show ?thesis using assms by clarsimp
-qed
-
-lemma Jcc_pieces_InitDP:
-assumes [simp]: "P \<equiv> compP\<^sub>2 P\<^sub>1"
- and "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' sh' v xa (INIT C' (C\<^sub>0 # Cs,False) \<leftarrow> e)
- = (True, frs', (None, h', (vs, l, C, M, pc, Called []) # frs, sh'), err)"
-shows
- "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' sh' v xa (INIT C' (Cs,True) \<leftarrow> e)
- = (True, (calling_to_scalled (hd frs'))#(tl frs'),
- (None, h', (vs, l, C, M, pc, Called []) # frs, sh'),
- \<exists>vs'. P \<turnstile> (None,h,calling_to_scalled (hd frs')#(tl frs'),sh)
- -jvm\<rightarrow> handle P C M xa h' (vs'@vs) l pc ics frs sh')"
-proof -
- have "Jcc_cond P\<^sub>1 E C M vs pc ics I h sh (INIT C' (C\<^sub>0 # Cs,False) \<leftarrow> e)" using assms by simp
- then obtain T where "P\<^sub>1,E,h,sh \<turnstile>\<^sub>1 INIT C' (C\<^sub>0 # Cs,False) \<leftarrow> unit : T" by fastforce
- then have "P\<^sub>1,E,h,sh \<turnstile>\<^sub>1 INIT C' (Cs,True) \<leftarrow> unit : T"
- by (auto; metis list.sel(2) list.set_sel(2))
- then have wtrt: "Ex (WTrt2\<^sub>1 P\<^sub>1 E h sh (INIT C' (Cs,True) \<leftarrow> unit))" by(simp only: exI)
- show ?thesis using assms wtrt
- proof(cases Cs)
- case (Cons C1 Cs1)
- then show ?thesis using assms wtrt
- by(case_tac "method P C1 clinit") clarsimp
- qed(clarsimp)
-qed
-
-lemma Jcc_pieces_InitError:
-assumes [simp]: "P \<equiv> compP\<^sub>2 P\<^sub>1"
- and "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' sh' v xa (INIT C' (C\<^sub>0 # Cs,False) \<leftarrow> e)
- = (True, frs', (None, h', (vs, l, C, M, pc, Called []) # frs, sh'), err)"
- and err: "sh C\<^sub>0 = Some(sfs,Error)"
-shows
- "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' sh' v xa (RI (C\<^sub>0, THROW NoClassDefFoundError);Cs \<leftarrow> e)
- = (True, (calling_to_throwing (hd frs') (addr_of_sys_xcpt NoClassDefFoundError))#(tl frs'),
- (None, h', (vs, l, C, M, pc, Called []) # frs, sh'),
- \<exists>vs'. P \<turnstile> (None,h, (calling_to_throwing (hd frs') (addr_of_sys_xcpt NoClassDefFoundError))#(tl frs'),sh)
- -jvm\<rightarrow> handle P C M xa h' (vs'@vs) l pc ics frs sh')"
-proof -
- show ?thesis using assms
- proof(cases Cs)
- case (Cons C1 Cs1)
- then show ?thesis using assms
- by(case_tac "method P C1 clinit", case_tac "method P C\<^sub>0 clinit") clarsimp
- qed(clarsimp)
-qed
-
-lemma Jcc_pieces_InitObj:
-assumes [simp]: "P \<equiv> compP\<^sub>2 P\<^sub>1"
- and "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' (sh(C\<^sub>0 \<mapsto> (sfs,Processing))) v xa (INIT C' (C\<^sub>0 # Cs,False) \<leftarrow> e)
- = (True, frs', (None, h', (vs, l, C, M, pc, Called []) # frs, sh'), err)"
-shows
- "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs (sh(C\<^sub>0 \<mapsto> (sfs,Processing))) I h' l' sh'' v xa (INIT C' (C\<^sub>0 # Cs,True) \<leftarrow> e)
- = (True, calling_to_called (hd frs')#(tl frs'),
- (None, h', (vs, l, C, M, pc, Called []) # frs, sh''),
- \<exists>vs'. P \<turnstile> (None,h,calling_to_called (hd frs')#(tl frs'),sh')
- -jvm\<rightarrow> handle P C M xa h' (vs'@vs) l pc ics frs sh'')"
-proof -
- have "Jcc_cond P\<^sub>1 E C M vs pc ics I h sh (INIT C' (C\<^sub>0 # Cs,False) \<leftarrow> e)" using assms by simp
- then obtain T where "P\<^sub>1,E,h,sh \<turnstile>\<^sub>1 INIT C' (C\<^sub>0 # Cs,False) \<leftarrow> unit : T" by fastforce
- then have "P\<^sub>1,E,h,sh(C\<^sub>0 \<mapsto> (sfs,Processing)) \<turnstile>\<^sub>1 INIT C' (C\<^sub>0 # Cs,True) \<leftarrow> unit : T"
- using assms by clarsimp (auto simp: fun_upd_apply)
- then have wtrt: "Ex (WTrt2\<^sub>1 P\<^sub>1 E h (sh(C\<^sub>0 \<mapsto> (sfs,Processing))) (INIT C' (C\<^sub>0 # Cs,True) \<leftarrow> unit))"
- by(simp only: exI)
- show ?thesis using assms wtrt by clarsimp
-qed
-
-lemma Jcc_pieces_InitNonObj:
-assumes [simp]: "P \<equiv> compP\<^sub>2 P\<^sub>1"
- and "is_class P\<^sub>1 D" and "D \<notin> set (C\<^sub>0#Cs)" and "\<forall>C \<in> set (C\<^sub>0#Cs). P\<^sub>1 \<turnstile> C \<preceq>\<^sup>* D"
- and pcs: "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' (sh(C\<^sub>0 \<mapsto> (sfs,Processing))) v xa (INIT C' (C\<^sub>0 # Cs,False) \<leftarrow> e)
- = (True, frs', (None, h', (vs, l, C, M, pc, Called []) # frs, sh'), err)"
-shows
- "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs (sh(C\<^sub>0 \<mapsto> (sfs,Processing))) I h' l' sh'' v xa (INIT C' (D # C\<^sub>0 # Cs,False) \<leftarrow> e)
- = (True, calling_to_calling (hd frs') D#(tl frs'),
- (None, h', (vs, l, C, M, pc, Called []) # frs, sh''),
- \<exists>vs'. P \<turnstile> (None,h,calling_to_calling (hd frs') D#(tl frs'),sh')
- -jvm\<rightarrow> handle P C M xa h' (vs'@vs) l pc ics frs sh'')"
-proof -
- have "Jcc_cond P\<^sub>1 E C M vs pc ics I h sh (INIT C' (C\<^sub>0 # Cs,False) \<leftarrow> e)" using assms by simp
- then obtain T where "P\<^sub>1,E,h,sh \<turnstile>\<^sub>1 INIT C' (C\<^sub>0 # Cs,False) \<leftarrow> unit : T" by fastforce
- then have "P\<^sub>1,E,h,sh(C\<^sub>0 \<mapsto> (sfs,Processing)) \<turnstile>\<^sub>1 INIT C' (D # C\<^sub>0 # Cs,False) \<leftarrow> unit : T"
- using assms by clarsimp (auto simp: fun_upd_apply)
- then have wtrt: "Ex (WTrt2\<^sub>1 P\<^sub>1 E h (sh(C\<^sub>0 \<mapsto> (sfs,Processing))) (INIT C' (D # C\<^sub>0 # Cs,False) \<leftarrow> unit))"
- by(simp only: exI)
- show ?thesis using assms wtrt by clarsimp
-qed
-
-lemma Jcc_pieces_InitRInit:
-assumes [simp]: "P \<equiv> compP\<^sub>2 P\<^sub>1" and wf: "wf_J\<^sub>1_prog P\<^sub>1"
- and "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' sh' v xa (INIT C' (C\<^sub>0 # Cs,True) \<leftarrow> e)
- = (True, frs', (None, h', (vs, l, C, M, pc, Called []) # frs, sh'), err)"
-shows
- "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' sh' v xa (RI (C\<^sub>0,C\<^sub>0\<bullet>\<^sub>sclinit([])) ; Cs \<leftarrow> e)
- = (True, frs',
- (None, h', (vs, l, C, M, pc, Called []) # frs, sh'),
- \<exists>vs'. P \<turnstile> (None,h,frs',sh)
- -jvm\<rightarrow> handle P C M xa h' (vs'@vs) l pc ics frs sh')"
-proof -
- have cond: "Jcc_cond P\<^sub>1 E C M vs pc ics I h sh (INIT C' (C\<^sub>0 # Cs,True) \<leftarrow> e)" using assms by simp
- then have clinit: "\<exists>T. P\<^sub>1,E,h,sh \<turnstile>\<^sub>1 C\<^sub>0\<bullet>\<^sub>sclinit([]) : T" using wf
- by clarsimp (auto simp: is_class_def intro: wf\<^sub>1_types_clinit)
- then obtain T where cT: "P\<^sub>1,E,h,sh \<turnstile>\<^sub>1 C\<^sub>0\<bullet>\<^sub>sclinit([]) : T" by blast
- obtain T where "P\<^sub>1,E,h,sh \<turnstile>\<^sub>1 INIT C' (C\<^sub>0 # Cs,True) \<leftarrow> unit : T" using cond by fastforce
- then have "P\<^sub>1,E,h,sh \<turnstile>\<^sub>1 RI (C\<^sub>0,C\<^sub>0\<bullet>\<^sub>sclinit([])) ; Cs \<leftarrow> unit : T"
- using assms by (auto intro: cT)
- then have wtrt: "Ex (WTrt2\<^sub>1 P\<^sub>1 E h sh (RI (C\<^sub>0,C\<^sub>0\<bullet>\<^sub>sclinit([])) ; Cs \<leftarrow> unit))"
- by(simp only: exI)
- then show ?thesis using assms by simp
-qed
-
-lemma Jcc_pieces_RInit_clinit:
-assumes [simp]: "P \<equiv> compP\<^sub>2 P\<^sub>1" and wf: "wf_J\<^sub>1_prog P\<^sub>1"
- and "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h\<^sub>1 l\<^sub>1 sh\<^sub>1 v xa (RI (C\<^sub>0,C\<^sub>0\<bullet>\<^sub>sclinit([]));Cs \<leftarrow> e)
- = (True, frs',
- (None, h\<^sub>1, (vs, l, C, M, pc, Called []) # frs, sh\<^sub>1), err)"
-shows
- "Jcc_pieces P\<^sub>1 E C M h vs l pc (Called Cs) (tl frs') sh I h' l' sh' v xa (C\<^sub>0\<bullet>\<^sub>sclinit([]))
- = (True, create_init_frame P C\<^sub>0#(vs,l,C,M,pc,Called Cs)#tl frs',
- (None, h', (vs,l,C,M,pc,Called Cs)#tl frs', sh'(C\<^sub>0\<mapsto>(fst(the(sh' C\<^sub>0)),Done))),
- P \<turnstile> (None,h,create_init_frame P C\<^sub>0#(vs,l,C,M,pc,Called Cs)#tl frs',sh)
- -jvm\<rightarrow> (None,h',(vs, l, C, M, pc, Throwing Cs xa) # tl frs',sh'(C\<^sub>0 \<mapsto> (fst(the(sh' C\<^sub>0)),Error))))"
-proof -
- have cond: "Jcc_cond P\<^sub>1 E C M vs pc ics I h sh (RI (C\<^sub>0,C\<^sub>0\<bullet>\<^sub>sclinit([]));Cs \<leftarrow> e)" using assms by simp
- then have wtrt: "\<exists>T. P\<^sub>1,E,h,sh \<turnstile>\<^sub>1 C\<^sub>0\<bullet>\<^sub>sclinit([]) : T" using wf
- by clarsimp (auto simp: is_class_def intro: wf\<^sub>1_types_clinit)
- then show ?thesis using assms by clarsimp
-qed
-
-lemma Jcc_pieces_RInit_Init:
-assumes [simp]: "P \<equiv> compP\<^sub>2 P\<^sub>1" and wf: "wf_J\<^sub>1_prog P\<^sub>1"
- and proc: "\<forall>C' \<in> set Cs. \<exists>sfs. sh'' C' = \<lfloor>(sfs,Processing)\<rfloor>"
- and "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h\<^sub>1 l\<^sub>1 sh\<^sub>1 v xa (RI (C\<^sub>0,C\<^sub>0\<bullet>\<^sub>sclinit([]));Cs \<leftarrow> e)
- = (True, frs',
- (None, h\<^sub>1, (vs, l, C, M, pc, Called []) # frs, sh\<^sub>1), err)"
-shows
- "Jcc_pieces P\<^sub>1 E C M h' vs l pc ics frs sh'' I h\<^sub>1 l\<^sub>1 sh\<^sub>1 v xa (INIT (last (C\<^sub>0#Cs)) (Cs,True) \<leftarrow> e)
- = (True, (vs, l, C, M, pc, Called Cs) # frs,
- (None, h\<^sub>1, (vs, l, C, M, pc, Called []) # frs, sh\<^sub>1),
- \<exists>vs'. P \<turnstile> (None,h',(vs, l, C, M, pc, Called Cs) # frs,sh'')
- -jvm\<rightarrow> handle P C M xa h\<^sub>1 (vs'@vs) l pc ics frs sh\<^sub>1)"
-proof -
- have "Jcc_cond P\<^sub>1 E C M vs pc ics I h sh (RI (C\<^sub>0,C\<^sub>0\<bullet>\<^sub>sclinit([]));Cs \<leftarrow> e)" using assms by simp
- then have "Ex (WTrt2\<^sub>1 P\<^sub>1 E h sh (RI (C\<^sub>0,C\<^sub>0\<bullet>\<^sub>sclinit([])) ; Cs \<leftarrow> unit))" by simp
- then obtain T where riwt: "P\<^sub>1,E,h,sh \<turnstile>\<^sub>1 RI (C\<^sub>0,C\<^sub>0\<bullet>\<^sub>sclinit([]));Cs \<leftarrow> unit : T" by meson
- then have "P\<^sub>1,E,h',sh'' \<turnstile>\<^sub>1 INIT (last (C\<^sub>0#Cs)) (Cs,True) \<leftarrow> unit : T" using proc
- proof(cases Cs) qed(auto)
- then have wtrt: "Ex (WTrt2\<^sub>1 P\<^sub>1 E h' sh'' (INIT (last (C\<^sub>0#Cs)) (Cs,True) \<leftarrow> unit))" by(simp only: exI)
- show ?thesis using assms wtrt
- proof(cases Cs)
- case (Cons C1 Cs1)
- then show ?thesis using assms wtrt
- by(case_tac "method P C1 clinit") clarsimp
- qed(clarsimp)
-qed
-
-lemma Jcc_pieces_RInit_RInit:
-assumes [simp]: "P \<equiv> compP\<^sub>2 P\<^sub>1"
- and "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h\<^sub>1 l\<^sub>1 sh\<^sub>1 v xa (RI (C\<^sub>0,e);D#Cs \<leftarrow> e')
- = (True, frs', rhs, err)"
- and hd: "hd frs' = (vs1,l1,C1,M1,pc1,ics1)"
-shows
- "Jcc_pieces P\<^sub>1 E C M h' vs l pc ics frs sh'' I h\<^sub>1 l\<^sub>1 sh\<^sub>1 v xa (RI (D,Throw xa) ; Cs \<leftarrow> e')
- = (True, (vs1, l1, C1, M1, pc1, Throwing (D#Cs) xa) # tl frs',
- (None, h\<^sub>1, (vs, l, C, M, pc, Called []) # frs, sh\<^sub>1),
- \<exists>vs'. P \<turnstile> (None,h',(vs1, l1, C1, M1, pc1, Throwing (D#Cs) xa) # tl frs',sh'')
- -jvm\<rightarrow> handle P C M xa h\<^sub>1 (vs'@vs) l pc ics frs sh\<^sub>1)"
-using assms by(case_tac "method P D clinit", cases "e = C\<^sub>0\<bullet>\<^sub>sclinit([])") clarsimp+
-
-
-subsubsection "JVM stepping lemmas"
-
-lemma jvm_Invoke:
-assumes [simp]: "P \<equiv> compP\<^sub>2 P\<^sub>1"
- and "P,C,M,pc \<triangleright> Invoke M' (length Ts)"
- and ha: "h\<^sub>2 a = \<lfloor>(Ca, fs)\<rfloor>" and method: "P\<^sub>1 \<turnstile> Ca sees M', NonStatic : Ts\<rightarrow>T = body in D"
- and len: "length pvs = length Ts" and "ls\<^sub>2' = Addr a # pvs @ replicate (max_vars body) undefined"
-shows "P \<turnstile> (None, h\<^sub>2, (rev pvs @ Addr a # vs, ls\<^sub>2, C, M, pc, No_ics) # frs, sh\<^sub>2) -jvm\<rightarrow>
- (None, h\<^sub>2, ([], ls\<^sub>2', D, M', 0, No_ics) # (rev pvs @ Addr a # vs, ls\<^sub>2, C, M, pc, No_ics) # frs, sh\<^sub>2)"
-proof -
- have cname: "cname_of h\<^sub>2 (the_Addr ((rev pvs @ Addr a # vs) ! length Ts)) = Ca"
- using ha method len by(auto simp: nth_append)
- have r: "(rev pvs @ Addr a # vs) ! (length Ts) = Addr a" using len by(auto simp: nth_append)
- have exm: "\<exists>Ts T m D b. P \<turnstile> Ca sees M',b:Ts \<rightarrow> T = m in D"
- using sees_method_compP[OF method] by fastforce
- show ?thesis using assms cname r exm by simp
-qed
-
-lemma jvm_Invokestatic:
-assumes [simp]: "P \<equiv> compP\<^sub>2 P\<^sub>1"
- and "P,C,M,pc \<triangleright> Invokestatic C' M' (length Ts)"
- and sh: "sh\<^sub>2 D = Some(sfs,Done)"
- and method: "P\<^sub>1 \<turnstile> C' sees M', Static : Ts\<rightarrow>T = body in D"
- and len: "length pvs = length Ts" and "ls\<^sub>2' = pvs @ replicate (max_vars body) undefined"
-shows "P \<turnstile> (None, h\<^sub>2, (rev pvs @ vs, ls\<^sub>2, C, M, pc, No_ics) # frs, sh\<^sub>2) -jvm\<rightarrow>
- (None, h\<^sub>2, ([], ls\<^sub>2', D, M', 0, No_ics) # (rev pvs @ vs, ls\<^sub>2, C, M, pc, No_ics) # frs, sh\<^sub>2)"
-proof -
- have exm: "\<exists>Ts T m D b. P \<turnstile> C' sees M',b:Ts \<rightarrow> T = m in D"
- using sees_method_compP[OF method] by fastforce
- show ?thesis using assms exm by simp
-qed
-
-lemma jvm_Invokestatic_Called:
-assumes [simp]: "P \<equiv> compP\<^sub>2 P\<^sub>1"
- and "P,C,M,pc \<triangleright> Invokestatic C' M' (length Ts)"
- and sh: "sh\<^sub>2 D = Some(sfs,i)"
- and method: "P\<^sub>1 \<turnstile> C' sees M', Static : Ts\<rightarrow>T = body in D"
- and len: "length pvs = length Ts" and "ls\<^sub>2' = pvs @ replicate (max_vars body) undefined"
-shows "P \<turnstile> (None, h\<^sub>2, (rev pvs @ vs, ls\<^sub>2, C, M, pc, Called []) # frs, sh\<^sub>2) -jvm\<rightarrow>
- (None, h\<^sub>2, ([], ls\<^sub>2', D, M', 0, No_ics) # (rev pvs @ vs, ls\<^sub>2, C, M, pc, No_ics) # frs, sh\<^sub>2)"
-proof -
- have exm: "\<exists>Ts T m D b. P \<turnstile> C' sees M',b:Ts \<rightarrow> T = m in D"
- using sees_method_compP[OF method] by fastforce
- show ?thesis using assms exm by simp
-qed
-
-lemma jvm_Return_Init:
-"P,D,clinit,0 \<rhd> compE\<^sub>2 body @ [Return]
- \<Longrightarrow> P \<turnstile> (None, h, (vs, ls, D, clinit, size(compE\<^sub>2 body), No_ics) # frs, sh)
- -jvm\<rightarrow> (None, h, frs, sh(D\<mapsto>(fst(the(sh D)),Done)))"
-(is "?P \<Longrightarrow> P \<turnstile> ?s1 -jvm\<rightarrow> ?s2")
-proof -
- assume ?P
- then have "exec (P, ?s1) = \<lfloor>?s2\<rfloor>" by(cases frs) auto
- then have "(?s1, ?s2) \<in> (exec_1 P)\<^sup>*"
- by(rule exec_1I[THEN r_into_rtrancl])
- then show ?thesis by(simp add: exec_all_def1)
-qed
-
-lemma jvm_InitNone:
- "\<lbrakk> ics_of f = Calling C Cs;
- sh C = None \<rbrakk>
- \<Longrightarrow> P \<turnstile> (None,h,f#frs,sh) -jvm\<rightarrow> (None,h,f#frs,sh(C \<mapsto> (sblank P C, Prepared)))"
-(is "\<lbrakk> ?P; ?Q \<rbrakk> \<Longrightarrow> P \<turnstile> ?s1 -jvm\<rightarrow> ?s2")
-proof -
- assume assms: ?P ?Q
- then obtain stk1 loc1 C1 M1 pc1 ics1 where "f = (stk1,loc1,C1,M1,pc1,ics1)"
- by(cases f) simp
- then have "exec (P, ?s1) = \<lfloor>?s2\<rfloor>" using assms
- by(case_tac ics1) simp_all
- then have "(?s1, ?s2) \<in> (exec_1 P)\<^sup>*"
- by(rule exec_1I[THEN r_into_rtrancl])
- then show ?thesis by(simp add: exec_all_def1)
-qed
-
-lemma jvm_InitDP:
- "\<lbrakk> ics_of f = Calling C Cs;
- sh C = \<lfloor>(sfs,i)\<rfloor>; i = Done \<or> i = Processing \<rbrakk>
- \<Longrightarrow> P \<turnstile> (None,h,f#frs,sh) -jvm\<rightarrow> (None,h,(calling_to_scalled f)#frs,sh)"
-(is "\<lbrakk> ?P; ?Q; ?R \<rbrakk> \<Longrightarrow> P \<turnstile> ?s1 -jvm\<rightarrow> ?s2")
-proof -
- assume assms: ?P ?Q ?R
- then obtain stk1 loc1 C1 M1 pc1 ics1 where "f = (stk1,loc1,C1,M1,pc1,ics1)"
- by(cases f) simp
- then have "exec (P, ?s1) = \<lfloor>?s2\<rfloor>" using assms
- by(case_tac i) simp_all
- then have "(?s1, ?s2) \<in> (exec_1 P)\<^sup>*"
- by(rule exec_1I[THEN r_into_rtrancl])
- then show ?thesis by(simp add: exec_all_def1)
-qed
-
-lemma jvm_InitError:
- "sh C = \<lfloor>(sfs,Error)\<rfloor>
- \<Longrightarrow> P \<turnstile> (None,h,(vs,ls,C\<^sub>0,M,pc,Calling C Cs)#frs,sh)
- -jvm\<rightarrow> (None,h,(vs,ls,C\<^sub>0,M,pc,Throwing Cs (addr_of_sys_xcpt NoClassDefFoundError))#frs,sh)"
- by(clarsimp simp: exec_all_def1 intro!: r_into_rtrancl exec_1I)
-
-lemma exec_ErrorThrowing:
- "sh C = \<lfloor>(sfs,Error)\<rfloor>
- \<Longrightarrow> exec (P, (None,h,calling_to_throwing (stk,loc,D,M,pc,Calling C Cs) a#frs,sh))
- = Some (None,h,calling_to_sthrowing (stk,loc,D,M,pc,Calling C Cs) a #frs,sh)"
- by(clarsimp simp: exec_all_def1 fun_upd_idem_iff intro!: r_into_rtrancl exec_1I)
-
-lemma jvm_InitObj:
- "\<lbrakk> sh C = Some(sfs,Prepared);
- C = Object;
- sh' = sh(C \<mapsto> (sfs,Processing)) \<rbrakk>
-\<Longrightarrow> P \<turnstile> (None, h, (vs,ls,C\<^sub>0,M,pc,Calling C Cs)#frs, sh) -jvm\<rightarrow>
- (None, h, (vs,ls,C\<^sub>0,M,pc,Called (C#Cs))#frs,sh')"
-(is "\<lbrakk> ?P; ?Q; ?R \<rbrakk> \<Longrightarrow> P \<turnstile> ?s1 -jvm\<rightarrow> ?s2")
-proof -
- assume assms: ?P ?Q ?R
- then have "exec (P, ?s1) = \<lfloor>?s2\<rfloor>"
- by(case_tac "method P C clinit") simp
- then have "(?s1, ?s2) \<in> (exec_1 P)\<^sup>*"
- by(rule exec_1I[THEN r_into_rtrancl])
- then show ?thesis by(simp add: exec_all_def1)
-qed
-
-lemma jvm_InitNonObj:
- "\<lbrakk> sh C = Some(sfs,Prepared);
- C \<noteq> Object;
- class P C = Some (D,r);
- sh' = sh(C \<mapsto> (sfs,Processing)) \<rbrakk>
-\<Longrightarrow> P \<turnstile> (None, h, (vs,ls,C\<^sub>0,M,pc,Calling C Cs)#frs, sh) -jvm\<rightarrow>
- (None, h, (vs,ls,C\<^sub>0,M,pc,Calling D (C#Cs))#frs, sh')"
-(is "\<lbrakk> ?P; ?Q; ?R; ?S \<rbrakk> \<Longrightarrow> P \<turnstile> ?s1 -jvm\<rightarrow> ?s2")
-proof -
- assume assms: ?P ?Q ?R ?S
- then have "exec (P, ?s1) = \<lfloor>?s2\<rfloor>"
- by(case_tac "method P C clinit") simp
- then have "(?s1, ?s2) \<in> (exec_1 P)\<^sup>*"
- by(rule exec_1I[THEN r_into_rtrancl])
- then show ?thesis by(simp add: exec_all_def1)
-qed
-
-lemma jvm_RInit_throw:
- "P \<turnstile> (None,h,(vs,l,C,M,pc,Throwing [] xa) # frs,sh)
- -jvm\<rightarrow> handle P C M xa h vs l pc No_ics frs sh"
-(is "P \<turnstile> ?s1 -jvm\<rightarrow> ?s2")
-proof -
- have "exec (P, ?s1) = \<lfloor>?s2\<rfloor>"
- by(simp add: handle_def split: bool.splits)
- then have "(?s1, ?s2) \<in> (exec_1 P)\<^sup>*"
- by(rule exec_1I[THEN r_into_rtrancl])
- then show ?thesis by(simp add: exec_all_def1)
-qed
-
-lemma jvm_RInit_throw':
- "P \<turnstile> (None,h,(vs,l,C,M,pc,Throwing [C'] xa) # frs,sh)
- -jvm\<rightarrow> handle P C M xa h vs l pc No_ics frs (sh(C':=Some(fst(the(sh C')), Error)))"
-(is "P \<turnstile> ?s1 -jvm\<rightarrow> ?s2")
-proof -
- let ?sy = "(None,h,(vs,l,C,M,pc,Throwing [] xa) # frs,sh(C':=Some(fst(the(sh C')), Error)))"
- have "exec (P, ?s1) = \<lfloor>?sy\<rfloor>" by simp
- then have "(?s1, ?sy) \<in> (exec_1 P)\<^sup>*"
- by(rule exec_1I[THEN r_into_rtrancl])
- also have "(?sy, ?s2) \<in> (exec_1 P)\<^sup>*"
- using jvm_RInit_throw by(simp add: exec_all_def1)
- ultimately show ?thesis by(simp add: exec_all_def1)
-qed
-
-lemma jvm_Called:
- "P \<turnstile> (None, h, (vs, l, C, M, pc, Called (C\<^sub>0 # Cs)) # frs, sh) -jvm\<rightarrow>
- (None, h, create_init_frame P C\<^sub>0 # (vs, l, C, M, pc, Called Cs) # frs, sh)"
- by(simp add: exec_all_def1 r_into_rtrancl exec_1I)
-
-lemma jvm_Throwing:
- "P \<turnstile> (None, h, (vs, l, C, M, pc, Throwing (C\<^sub>0#Cs) xa') # frs, sh) -jvm\<rightarrow>
- (None, h, (vs, l, C, M, pc, Throwing Cs xa') # frs, sh(C\<^sub>0 \<mapsto> (fst (the (sh C\<^sub>0)), Error)))"
- by(simp add: exec_all_def1 r_into_rtrancl exec_1I)
-
-subsubsection "Other lemmas for correctness proof"
-
-lemma assumes wf:"wf_prog wf_md P"
- and ex: "class P C = Some a"
-shows create_init_frame_wf_eq: "create_init_frame (compP\<^sub>2 P) C = (stk,loc,D,M,pc,ics) \<Longrightarrow> D=C"
-using wf_sees_clinit[OF wf ex] by(cases "method P C clinit", auto)
-
-lemma beforex_try:
-assumes pcI: "{pc..<pc+size(compE\<^sub>2(try e\<^sub>1 catch(Ci i) e\<^sub>2))} \<subseteq> I"
- and bx: "P,C,M \<rhd> compxE\<^sub>2 (try e\<^sub>1 catch(Ci i) e\<^sub>2) pc (size vs) / I,size vs"
-shows "P,C,M \<rhd> compxE\<^sub>2 e\<^sub>1 pc (size vs) / {pc..<pc + length (compE\<^sub>2 e\<^sub>1)},size vs"
-proof -
- obtain xt\<^sub>0 xt\<^sub>1 where
- "beforex\<^sub>0 P C M (size vs) I (compxE\<^sub>2 (try e\<^sub>1 catch(Ci i) e\<^sub>2) pc (size vs)) xt\<^sub>0 xt\<^sub>1"
- using bx by(clarsimp simp:beforex_def)
- then have "\<exists>xt1. beforex\<^sub>0 P C M (size vs) {pc..<pc + length (compE\<^sub>2 e\<^sub>1)}
- (compxE\<^sub>2 e\<^sub>1 pc (size vs)) xt\<^sub>0 xt1"
- using pcI pcs_subset(1) atLeastLessThan_iff by simp blast
- then show ?thesis using beforex_def by blast
-qed
-
-\<comment> \<open> Evaluation of initialization expressions \<close>
-
-(* --needs J1 and EConform; version for eval found in Equivalence *)
-lemma
-shows eval\<^sub>1_init_return: "P \<turnstile>\<^sub>1 \<langle>e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>
- \<Longrightarrow> iconf (shp\<^sub>1 s) e
- \<Longrightarrow> (\<exists>Cs b. e = INIT C' (Cs,b) \<leftarrow> unit) \<or> (\<exists>C e\<^sub>0 Cs e\<^sub>i. e = RI(C,e\<^sub>0);Cs@[C'] \<leftarrow> unit)
- \<or> (\<exists>e\<^sub>0. e = RI(C',e\<^sub>0);Nil \<leftarrow> unit)
- \<Longrightarrow> (val_of e' = Some v \<longrightarrow> (\<exists>sfs i. shp\<^sub>1 s' C' = \<lfloor>(sfs,i)\<rfloor> \<and> (i = Done \<or> i = Processing)))
- \<and> (throw_of e' = Some a \<longrightarrow> (\<exists>sfs i. shp\<^sub>1 s' C' = \<lfloor>(sfs,Error)\<rfloor>))"
-and "P \<turnstile>\<^sub>1 \<langle>es,s\<rangle> [\<Rightarrow>] \<langle>es',s'\<rangle> \<Longrightarrow> True"
-proof(induct rule: eval\<^sub>1_evals\<^sub>1.inducts)
- case (InitFinal\<^sub>1 e s e' s' C b) then show ?case
- by(auto simp: initPD_def dest: eval\<^sub>1_final_same)
-next
- case (InitDone\<^sub>1 sh C sfs C' Cs e h l e' s')
- then have "final e'" using eval\<^sub>1_final by simp
- then show ?case
- proof(rule finalE)
- fix v assume e': "e' = Val v" then show ?thesis using InitDone\<^sub>1 initPD_def
- proof(cases Cs) qed(auto)
- next
- fix a assume e': "e' = throw a" then show ?thesis using InitDone\<^sub>1 initPD_def
- proof(cases Cs) qed(auto)
- qed
-next
- case (InitProcessing\<^sub>1 sh C sfs C' Cs e h l e' s')
- then have "final e'" using eval\<^sub>1_final by simp
- then show ?case
- proof(rule finalE)
- fix v assume e': "e' = Val v" then show ?thesis using InitProcessing\<^sub>1 initPD_def
- proof(cases Cs) qed(auto)
- next
- fix a assume e': "e' = throw a" then show ?thesis using InitProcessing\<^sub>1 initPD_def
- proof(cases Cs) qed(auto)
- qed
-next
- case (InitError\<^sub>1 sh C sfs Cs e h l e' s' C') show ?case
- proof(cases Cs)
- case Nil then show ?thesis using InitError\<^sub>1 by simp
- next
- case (Cons C2 list)
- then have "final e'" using InitError\<^sub>1 eval\<^sub>1_final by simp
- then show ?thesis
- proof(rule finalE)
- fix v assume e': "e' = Val v" show ?thesis
- using InitError\<^sub>1.hyps(2) e' rinit\<^sub>1_throwE by blast
- next
- fix a assume e': "e' = throw a"
- then show ?thesis using Cons InitError\<^sub>1 cons_to_append[of list] by clarsimp
- qed
- qed
-next
- case (InitRInit\<^sub>1 C Cs h l sh e' s' C') show ?case
- proof(cases Cs)
- case Nil then show ?thesis using InitRInit\<^sub>1 by simp
- next
- case (Cons C' list) then show ?thesis
- using InitRInit\<^sub>1 Cons cons_to_append[of list] by clarsimp
- qed
-next
- case (RInit\<^sub>1 e s v h' l' sh' C sfs i sh'' C' Cs e' e\<^sub>1 s\<^sub>1)
- then have final: "final e\<^sub>1" using eval\<^sub>1_final by simp
- then show ?case
- proof(cases Cs)
- case Nil show ?thesis using final
- proof(rule finalE)
- fix v assume e': "e\<^sub>1 = Val v" show ?thesis
- using RInit\<^sub>1 Nil by(clarsimp, meson fun_upd_same initPD_def)
- next
- fix a assume e': "e\<^sub>1 = throw a" show ?thesis
- using RInit\<^sub>1 Nil by(clarsimp, meson fun_upd_same initPD_def)
- qed
- next
- case (Cons a list) show ?thesis using final
- proof(rule finalE)
- fix v assume e': "e\<^sub>1 = Val v" then show ?thesis
- using RInit\<^sub>1 Cons by(clarsimp, metis last.simps last_appendR list.distinct(1))
- next
- fix a assume e': "e\<^sub>1 = throw a" then show ?thesis
- using RInit\<^sub>1 Cons by(clarsimp, metis last.simps last_appendR list.distinct(1))
- qed
- qed
-next
- case (RInitInitFail\<^sub>1 e s a h' l' sh' C sfs i sh'' D Cs e' e\<^sub>1 s\<^sub>1)
- then have final: "final e\<^sub>1" using eval\<^sub>1_final by simp
- then show ?case
- proof(rule finalE)
- fix v assume e': "e\<^sub>1 = Val v" then show ?thesis
- using RInitInitFail\<^sub>1 by(clarsimp, meson exp.distinct(101) rinit\<^sub>1_throwE)
- next
- fix a' assume e': "e\<^sub>1 = Throw a'"
- then have "iconf (sh'(C \<mapsto> (sfs, Error))) a"
- using RInitInitFail\<^sub>1.hyps(1) eval\<^sub>1_final by fastforce
- then show ?thesis using RInitInitFail\<^sub>1 e'
- by(clarsimp, meson Cons_eq_append_conv list.inject)
- qed
-qed(auto simp: fun_upd_same)
-
-lemma init\<^sub>1_Val_PD: "P \<turnstile>\<^sub>1 \<langle>INIT C' (Cs,b) \<leftarrow> unit,s\<rangle> \<Rightarrow> \<langle>Val v,s'\<rangle>
- \<Longrightarrow> iconf (shp\<^sub>1 s) (INIT C' (Cs,b) \<leftarrow> unit)
- \<Longrightarrow> \<exists>sfs i. shp\<^sub>1 s' C' = \<lfloor>(sfs,i)\<rfloor> \<and> (i = Done \<or> i = Processing)"
- by(drule_tac v = v in eval\<^sub>1_init_return, simp+)
-
-lemma init\<^sub>1_throw_PD: "P \<turnstile>\<^sub>1 \<langle>INIT C' (Cs,b) \<leftarrow> unit,s\<rangle> \<Rightarrow> \<langle>throw a,s'\<rangle>
- \<Longrightarrow> iconf (shp\<^sub>1 s) (INIT C' (Cs,b) \<leftarrow> unit)
- \<Longrightarrow> \<exists>sfs i. shp\<^sub>1 s' C' = \<lfloor>(sfs,Error)\<rfloor>"
- by(drule_tac a = a in eval\<^sub>1_init_return, simp+)
-
-lemma rinit\<^sub>1_Val_PD:
-assumes eval: "P \<turnstile>\<^sub>1 \<langle>RI(C,e\<^sub>0);Cs \<leftarrow> unit,s\<rangle> \<Rightarrow> \<langle>Val v,s'\<rangle>"
- and iconf: "iconf (shp\<^sub>1 s) (RI(C,e\<^sub>0);Cs \<leftarrow> unit)" and last: "last(C#Cs) = C'"
-shows "\<exists>sfs i. shp\<^sub>1 s' C' = \<lfloor>(sfs,i)\<rfloor> \<and> (i = Done \<or> i = Processing)"
-proof(cases Cs)
- case Nil
- then show ?thesis using eval\<^sub>1_init_return[OF eval iconf] last by simp
-next
- case (Cons a list)
- then have nNil: "Cs \<noteq> []" by simp
- then have "\<exists>Cs'. Cs = Cs' @ [C']" using last append_butlast_last_id[OF nNil]
- by(rule_tac x="butlast Cs" in exI) simp
- then show ?thesis using eval\<^sub>1_init_return[OF eval iconf] by simp
-qed
-
-lemma rinit\<^sub>1_throw_PD:
-assumes eval: "P \<turnstile>\<^sub>1 \<langle>RI(C,e\<^sub>0);Cs \<leftarrow> unit,s\<rangle> \<Rightarrow> \<langle>throw a,s'\<rangle>"
- and iconf: "iconf (shp\<^sub>1 s) (RI(C,e\<^sub>0);Cs \<leftarrow> unit)" and last: "last(C#Cs) = C'"
-shows "\<exists>sfs. shp\<^sub>1 s' C' = \<lfloor>(sfs,Error)\<rfloor>"
-proof(cases Cs)
- case Nil
- then show ?thesis using eval\<^sub>1_init_return[OF eval iconf] last by simp
-next
- case (Cons a list)
- then have nNil: "Cs \<noteq> []" by simp
- then have "\<exists>Cs'. Cs = Cs' @ [C']" using last append_butlast_last_id[OF nNil]
- by(rule_tac x="butlast Cs" in exI) simp
- then show ?thesis using eval\<^sub>1_init_return[OF eval iconf] by simp
-qed
-
-subsubsection "The proof"
-
-lemma fixes P\<^sub>1 defines [simp]: "P \<equiv> compP\<^sub>2 P\<^sub>1"
-assumes wf: "wf_J\<^sub>1_prog P\<^sub>1"
-shows Jcc: "P\<^sub>1 \<turnstile>\<^sub>1 \<langle>e,(h\<^sub>0,ls\<^sub>0,sh\<^sub>0)\<rangle> \<Rightarrow> \<langle>ef,(h\<^sub>1,ls\<^sub>1,sh\<^sub>1)\<rangle> \<Longrightarrow>
- (\<And>E C M pc ics v xa vs frs I.
- \<lbrakk> Jcc_cond P\<^sub>1 E C M vs pc ics I h\<^sub>0 sh\<^sub>0 e \<rbrakk> \<Longrightarrow>
- (ef = Val v \<longrightarrow>
- P \<turnstile> (None,h\<^sub>0,Jcc_frames P C M vs ls\<^sub>0 pc ics frs e,sh\<^sub>0)
- -jvm\<rightarrow> Jcc_rhs P\<^sub>1 E C M vs ls\<^sub>0 pc ics frs h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v e)
- \<and>
- (ef = Throw xa \<longrightarrow> Jcc_err P C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 xa e)
- )"
-(*<*)
- (is "_ \<Longrightarrow> (\<And>E C M pc ics v xa vs frs I.
- PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 ef h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics v xa vs frs I)")
-(*>*)
-and "P\<^sub>1 \<turnstile>\<^sub>1 \<langle>es,(h\<^sub>0,ls\<^sub>0,sh\<^sub>0)\<rangle> [\<Rightarrow>] \<langle>fs,(h\<^sub>1,ls\<^sub>1,sh\<^sub>1)\<rangle> \<Longrightarrow>
- (\<And>C M pc ics ws xa es' vs frs I.
- \<lbrakk> P,C,M,pc \<rhd> compEs\<^sub>2 es; P,C,M \<rhd> compxEs\<^sub>2 es pc (size vs)/I,size vs;
- {pc..<pc+size(compEs\<^sub>2 es)} \<subseteq> I; ics = No_ics;
- \<not>sub_RIs es \<rbrakk> \<Longrightarrow>
- (fs = map Val ws \<longrightarrow>
- P \<turnstile> (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0) -jvm\<rightarrow>
- (None,h\<^sub>1,(rev ws @ vs,ls\<^sub>1,C,M,pc+size(compEs\<^sub>2 es),ics)#frs,sh\<^sub>1))
- \<and>
- (fs = map Val ws @ Throw xa # es' \<longrightarrow>
- (\<exists>pc\<^sub>1. pc \<le> pc\<^sub>1 \<and> pc\<^sub>1 < pc + size(compEs\<^sub>2 es) \<and>
- \<not> caught P pc\<^sub>1 h\<^sub>1 xa (compxEs\<^sub>2 es pc (size vs)) \<and>
- (\<exists>vs'. P \<turnstile> (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0)
- -jvm\<rightarrow> handle P C M xa h\<^sub>1 (vs'@vs) ls\<^sub>1 pc\<^sub>1 ics frs sh\<^sub>1))))"
-(*<*)
- (is "_ \<Longrightarrow> (\<And>C M pc ics ws xa es' vs frs I.
- PROP ?Ps es h\<^sub>0 ls\<^sub>0 sh\<^sub>0 fs h\<^sub>1 ls\<^sub>1 sh\<^sub>1 C M pc ics ws xa es' vs frs I)")
-proof (induct rule:eval\<^sub>1_evals\<^sub>1_inducts)
- case New\<^sub>1 thus ?case by auto
-next
- case (NewFail\<^sub>1 sh C' sfs h ls)
- let ?xa = "addr_of_sys_xcpt OutOfMemory"
- have "P \<turnstile> (None,h,(vs,ls,C,M,pc,ics)#frs,sh) -jvm\<rightarrow> handle P C M ?xa h vs ls pc ics frs sh"
- using NewFail\<^sub>1 by(clarsimp simp: handle_def)
- then show ?case by(auto intro!: exI[where x="[]"])
-next
- case (NewInit\<^sub>1 sh C' h ls v' h' ls' sh' a FDTs h'')
- then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs ls pc ics frs sh I h' ls' sh' v xa (new C')
- = (True, frs', (None,h',(v#vs,ls',C,M,pc+size(compE\<^sub>2 (new C')),ics)#frs,sh'), err)"
- using NewInit\<^sub>1.prems(1) by clarsimp
- have "Ex (WTrt2\<^sub>1 P\<^sub>1 E h sh (INIT C' ([C'],False) \<leftarrow> unit))"
- using has_fields_is_class[OF NewInit\<^sub>1.hyps(5)] by auto
- then obtain err' where pcs':
- "Jcc_pieces P\<^sub>1 E C M h vs ls pc ics frs sh I h' ls' sh' v' xa (INIT C' ([C'],False) \<leftarrow> unit)
- = (True, (vs,ls,C,M,pc,Calling C' []) # frs, (None,h',(vs,ls,C,M,pc,Called [])#frs,sh'), err')"
- using NewInit\<^sub>1.prems(1) by auto
- have IH: "PROP ?P (INIT C' ([C'],False) \<leftarrow> unit) h ls sh (Val v')
- h' ls' sh' E C M pc ics v' xa vs frs I" by fact
- have ls: "ls = ls'" by(rule init\<^sub>1_same_loc[OF NewInit\<^sub>1.hyps(2)])
- obtain sfs i where sh': "sh' C' = Some(sfs,i)"
- using init\<^sub>1_Val_PD[OF NewInit\<^sub>1.hyps(2)] by clarsimp
- have "P \<turnstile> (None,h,(vs,ls,C,M,pc,ics)#frs,sh) -jvm\<rightarrow> (None,h,(vs,ls,C,M,pc,Calling C' [])#frs,sh)"
- proof(cases "sh C'")
- case None then show ?thesis using NewInit\<^sub>1.prems by(cases ics) auto
- next
- case (Some a)
- then obtain sfs i where "a = (sfs,i)" by(cases a)
- then show ?thesis using NewInit\<^sub>1.hyps(1) NewInit\<^sub>1.prems Some
- by(cases ics; case_tac i) auto
- qed
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> (None, h', (vs, ls, C, M, pc, Called []) # frs, sh')"
- using IH pcs' by auto
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> (None, h'', (Addr a#vs, ls, C, M, Suc pc, ics) # frs, sh')"
- using NewInit\<^sub>1.hyps(1,2,4-6) NewInit\<^sub>1.prems sh' by(cases ics) auto
- finally show ?case using pcs ls by clarsimp
-next
- case (NewInitOOM\<^sub>1 sh C' h ls v' h' ls' sh')
- let ?xa = "addr_of_sys_xcpt OutOfMemory"
- obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs ls pc ics frs sh I h' ls' sh' v xa (new C')
- = (True, frs', (None,h',(v#vs,ls',C,M,pc+size(compE\<^sub>2 (new C')),ics)#frs,sh'), err)"
- using NewInitOOM\<^sub>1.prems(1) by clarsimp
- have "Ex (WTrt2\<^sub>1 P\<^sub>1 E h sh (INIT C' ([C'],False) \<leftarrow> unit))" using NewInitOOM\<^sub>1.hyps(5) by auto
- then obtain err' where pcs':
- "Jcc_pieces P\<^sub>1 E C M h vs ls pc ics frs sh I h' ls' sh' v' xa (INIT C' ([C'],False) \<leftarrow> unit)
- = (True, (vs,ls,C,M,pc,Calling C' []) # frs, (None,h',(vs,ls,C,M,pc,Called [])#frs,sh'), err')"
- using NewInitOOM\<^sub>1.prems(1) by auto
- have IH: "PROP ?P (INIT C' ([C'],False) \<leftarrow> unit) h ls sh (Val v')
- h' ls' sh' E C M pc ics v' xa vs frs I" by fact
- have ls: "ls = ls'" by(rule init\<^sub>1_same_loc[OF NewInitOOM\<^sub>1.hyps(2)])
- have "iconf (shp\<^sub>1 (h, ls, sh)) (INIT C' ([C'],False) \<leftarrow> unit)" by simp
- then obtain sfs i where sh': "sh' C' = Some(sfs,i)"
- using init\<^sub>1_Val_PD[OF NewInitOOM\<^sub>1.hyps(2)] by clarsimp
- have "P \<turnstile> (None,h,(vs,ls,C,M,pc,ics)#frs,sh) -jvm\<rightarrow> (None,h,(vs,ls,C,M,pc,Calling C' [])#frs,sh)"
- proof(cases "sh C'")
- case None then show ?thesis using NewInitOOM\<^sub>1.prems by(cases ics) auto
- next
- case (Some a)
- then obtain sfs i where "a = (sfs,i)" by(cases a)
- then show ?thesis using NewInitOOM\<^sub>1.hyps(1) NewInitOOM\<^sub>1.prems Some
- by(cases ics; case_tac i) auto
- qed
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> (None, h', (vs, ls, C, M, pc, Called []) # frs, sh')"
- using IH pcs' by auto
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> handle P C M ?xa h' vs ls pc ics frs sh'"
- using NewInitOOM\<^sub>1.hyps(1,2,4,5) NewInitOOM\<^sub>1.prems sh' by(auto simp: handle_def)
- finally show ?case using pcs ls by(simp, metis (no_types) append_Nil le_refl lessI)
-next
- case (NewInitThrow\<^sub>1 sh C' h ls a h' ls' sh')
- obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs ls pc ics frs sh I h' ls' sh' v xa (new C')
- = (True, frs', (None,h',(v#vs,ls',C,M,pc+size(compE\<^sub>2 (new C')),ics)#frs,sh'), err)"
- using NewInitThrow\<^sub>1.prems(1) by clarsimp
- obtain a' where throw: "throw a = Throw a'" using eval\<^sub>1_final[OF NewInitThrow\<^sub>1.hyps(2)] by clarsimp
- have "Ex (WTrt2\<^sub>1 P\<^sub>1 E h sh (INIT C' ([C'],False) \<leftarrow> unit))" using NewInitThrow\<^sub>1.hyps(4) by auto
- then obtain vs' where pcs':
- "Jcc_pieces P\<^sub>1 E C M h vs ls pc ics frs sh I h' ls' sh' v a' (INIT C' ([C'],False) \<leftarrow> unit)
- = (True, (vs,ls,C,M,pc,Calling C' []) # frs, (None,h',(vs,ls,C,M,pc,Called [])#frs,sh'),
- P \<turnstile> (None,h,(vs,ls,C,M,pc,Calling C' []) # frs,sh)
- -jvm\<rightarrow> handle P C M a' h' (vs'@vs) ls pc ics frs sh')"
- using NewInitThrow\<^sub>1.prems(1) by simp blast
- have IH: "PROP ?P (INIT C' ([C'],False) \<leftarrow> unit) h ls sh (throw a)
- h' ls' sh' E C M pc ics v a' vs frs I" by fact
- have ls: "ls = ls'" by(rule init\<^sub>1_same_loc[OF NewInitThrow\<^sub>1.hyps(2)])
- then have "P \<turnstile> (None,h,(vs,ls,C,M,pc,ics)#frs,sh) -jvm\<rightarrow> (None,h,(vs,ls,C,M,pc,Calling C' []) # frs,sh)"
- proof(cases "sh C'")
- case None then show ?thesis using NewInitThrow\<^sub>1.prems by(cases ics) auto
- next
- case (Some a)
- then obtain sfs i where "a = (sfs,i)" by(cases a)
- then show ?thesis using NewInitThrow\<^sub>1.hyps(1) NewInitThrow\<^sub>1.prems Some
- by(cases ics; case_tac i) auto
- qed
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> handle P C M a' h' (vs'@vs) ls pc ics frs sh'" using IH pcs' throw by auto
- finally show ?case using throw ls by auto
-next
- case (Cast\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 a h\<^sub>1 ls\<^sub>1 sh\<^sub>1 D fs C')
- let ?pc = "pc + length(compE\<^sub>2 e)"
- obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (Cast C' e)
- = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (Cast C' e)),ics)#frs,sh\<^sub>1), err)"
- using Cast\<^sub>1.prems(1) by auto
- have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (addr a) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Addr a) xa vs frs I" by fact
- then have "P \<turnstile> (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0) -jvm\<rightarrow>
- (None,h\<^sub>1,(Addr a#vs,ls\<^sub>1,C,M,?pc,ics)#frs,sh\<^sub>1)"
- using Jcc_pieces_Cast[OF assms(1) pcs, of "Addr a"] Cast\<^sub>1.prems pcs by auto
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> (None,h\<^sub>1,(Addr a#vs,ls\<^sub>1,C,M,?pc+1,ics)#frs,sh\<^sub>1)"
- using Cast\<^sub>1 by (auto simp add:cast_ok_def)
- finally show ?case by auto
-next
- case (CastNull\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 C')
- let ?pc = "pc + length(compE\<^sub>2 e)"
- obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (Cast C' e)
- = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (Cast C' e)),ics)#frs,sh\<^sub>1), err)"
- using CastNull\<^sub>1.prems(1) by clarsimp
- have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 null h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics Null xa vs frs I" by fact
- then have "P \<turnstile> (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0) -jvm\<rightarrow>
- (None,h\<^sub>1,(Null#vs,ls\<^sub>1,C,M,?pc,ics)#frs,sh\<^sub>1)"
- using Jcc_pieces_Cast[OF assms(1) pcs, of Null] CastNull\<^sub>1.prems pcs by auto
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> (None,h\<^sub>1,(Null#vs,ls\<^sub>1,C,M,?pc+1,ics)#frs,sh\<^sub>1)"
- using CastNull\<^sub>1 by (auto simp add:cast_ok_def)
- finally show ?case by auto
-next
- case (CastFail\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 a h\<^sub>1 ls\<^sub>1 sh\<^sub>1 D fs C')
- let ?pc = "pc + length(compE\<^sub>2 e)"
- let ?xa = "addr_of_sys_xcpt ClassCast"
- obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (Cast C' e)
- = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (Cast C' e)),ics)#frs,sh\<^sub>1), err)"
- using CastFail\<^sub>1.prems(1) by clarsimp
- have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (addr a) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Addr a) xa vs frs I" by fact
- then have "P \<turnstile> (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0) -jvm\<rightarrow>
- (None,h\<^sub>1,(Addr a#vs,ls\<^sub>1,C,M,?pc,ics)#frs,sh\<^sub>1)"
- using Jcc_pieces_Cast[OF assms(1) pcs, of "Addr a"] CastFail\<^sub>1.prems pcs by auto
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> handle P C M ?xa h\<^sub>1 (Addr a#vs) ls\<^sub>1 ?pc ics frs sh\<^sub>1"
- using CastFail\<^sub>1 by (auto simp add:handle_def cast_ok_def)
- finally have exec: "P \<turnstile> (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0) -jvm\<rightarrow> \<dots>".
- show ?case (is "?N \<and> (?eq \<longrightarrow> ?err)")
- proof
- show ?N by simp
- next
- { assume ?eq
- then have ?err using exec by (auto intro!: exI[where x="?pc"] exI[where x="[Addr a]"])
- }
- thus "?eq \<longrightarrow> ?err" by simp
- qed
-next
- case (CastThrow\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 e' h\<^sub>1 ls\<^sub>1 sh\<^sub>1 C')
- obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (Cast C' e)
- = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (Cast C' e)),ics)#frs,sh\<^sub>1), err)"
- using CastThrow\<^sub>1.prems(1) by clarsimp
- have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (throw e') h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics v xa vs frs I" by fact
- show ?case using IH Jcc_pieces_Cast[OF assms(1) pcs, of v] CastThrow\<^sub>1.prems pcs less_SucI
- by(simp, blast)
-next
- case Val\<^sub>1 thus ?case by auto
-next
- case Var\<^sub>1 thus ?case by auto
-next
- case (BinOp\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 v\<^sub>1 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 e\<^sub>2 v\<^sub>2 h\<^sub>2 ls\<^sub>2 sh\<^sub>2 bop w)
- obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v xa (e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2)
- = (True, frs', (None,h\<^sub>2,(v#vs,ls\<^sub>2,C,M,pc+size(compE\<^sub>2 (e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2)),ics)#frs,sh\<^sub>2), err)"
- using BinOp\<^sub>1.prems(1) by clarsimp
- let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e\<^sub>1)"
- let ?pc\<^sub>2 = "?pc\<^sub>1 + length(compE\<^sub>2 e\<^sub>2)"
- have IH\<^sub>1: "PROP ?P e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (Val v\<^sub>1) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics v\<^sub>1 xa vs frs
- (I - pcs (compxE\<^sub>2 e\<^sub>2 (pc + length (compE\<^sub>2 e\<^sub>1)) (Suc (length vs))))" by fact
- have IH\<^sub>2: "PROP ?P e\<^sub>2 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (Val v\<^sub>2) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 E C M ?pc\<^sub>1 ics v\<^sub>2 xa (v\<^sub>1#vs) frs
- (I - pcs(compxE\<^sub>2 e\<^sub>1 pc (size vs)))" by fact
- have "P \<turnstile> (None,h\<^sub>0,frs',sh\<^sub>0) -jvm\<rightarrow> (None,h\<^sub>1,(v\<^sub>1#vs,ls\<^sub>1,C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)"
- using IH\<^sub>1 Jcc_pieces_BinOp1[OF pcs, of h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v\<^sub>1] by simp
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> (None,h\<^sub>2,(v\<^sub>2#v\<^sub>1#vs,ls\<^sub>2,C,M,?pc\<^sub>2,ics)#frs,sh\<^sub>2)"
- using IH\<^sub>2 Jcc_pieces_BinOp2[OF assms(1) pcs, of h\<^sub>1 v\<^sub>1 ls\<^sub>1 sh\<^sub>1 v\<^sub>2] by (simp add: add.assoc)
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> (None,h\<^sub>2,(w#vs,ls\<^sub>2,C,M,?pc\<^sub>2+1,ics)#frs,sh\<^sub>2)"
- using BinOp\<^sub>1 by(cases bop) auto
- finally show ?case using pcs by (auto split: bop.splits simp:add.assoc)
-next
- case (BinOpThrow\<^sub>1\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 e h\<^sub>1 ls\<^sub>1 sh\<^sub>1 bop e\<^sub>2)
- obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2)
- = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2)),ics)#frs,sh\<^sub>1), err)"
- using BinOpThrow\<^sub>1\<^sub>1.prems(1) by clarsimp
- have IH\<^sub>1: "PROP ?P e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (throw e) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics v xa vs frs
- (I - pcs (compxE\<^sub>2 e\<^sub>2 (pc + length (compE\<^sub>2 e\<^sub>1)) (Suc (length vs))))" by fact
- show ?case using IH\<^sub>1 Jcc_pieces_BinOp1[OF pcs, of h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v] BinOpThrow\<^sub>1\<^sub>1.prems nsub_RI_Jcc_pieces
- by auto
-next
- case (BinOpThrow\<^sub>2\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 v\<^sub>1 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 e\<^sub>2 e h\<^sub>2 ls\<^sub>2 sh\<^sub>2 bop)
- obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v xa (e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2)
- = (True, frs', (None,h\<^sub>2,(v#vs,ls\<^sub>2,C,M,pc+size(compE\<^sub>2 (e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2)),ics)#frs,sh\<^sub>2), err)"
- using BinOpThrow\<^sub>2\<^sub>1.prems(1) by clarsimp
- let ?pc = "pc + length(compE\<^sub>2 e\<^sub>1)"
- have IH\<^sub>1: "PROP ?P e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (Val v\<^sub>1) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics v\<^sub>1 xa vs frs
- (I - pcs (compxE\<^sub>2 e\<^sub>2 (pc + length (compE\<^sub>2 e\<^sub>1)) (Suc (length vs))))" by fact
- have IH\<^sub>2: "PROP ?P e\<^sub>2 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (throw e) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 E C M ?pc ics v xa (v\<^sub>1#vs) frs
- (I - pcs(compxE\<^sub>2 e\<^sub>1 pc (size vs)))" by fact
- let ?\<sigma>\<^sub>1 = "(None,h\<^sub>1,(v\<^sub>1#vs,ls\<^sub>1,C,M,?pc,ics)#frs,sh\<^sub>1)"
- have 1: "P \<turnstile> (None,h\<^sub>0,frs',sh\<^sub>0) -jvm\<rightarrow> ?\<sigma>\<^sub>1"
- using IH\<^sub>1 Jcc_pieces_BinOp1[OF pcs, of h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v\<^sub>1] by simp
- have "(throw e = Val v \<longrightarrow> P \<turnstile> (None, h\<^sub>0, Jcc_frames P C M vs ls\<^sub>0 pc ics frs (e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2), sh\<^sub>0) -jvm\<rightarrow>
- Jcc_rhs P\<^sub>1 E C M vs ls\<^sub>0 pc ics frs h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v (e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2))
- \<and> (throw e = Throw xa \<longrightarrow> (\<exists>pc\<^sub>1. pc \<le> pc\<^sub>1 \<and> pc\<^sub>1 < pc + size(compE\<^sub>2 (e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2)) \<and>
- \<not> caught P pc\<^sub>1 h\<^sub>2 xa (compxE\<^sub>2 (e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2) pc (size vs)) \<and>
- (\<exists>vs'. P \<turnstile> (None,h\<^sub>0,frs',sh\<^sub>0) -jvm\<rightarrow> handle P C M xa h\<^sub>2 (vs'@vs) ls\<^sub>2 pc\<^sub>1 ics frs sh\<^sub>2)))"
- (is "?N \<and> (?eq \<longrightarrow> (\<exists>pc\<^sub>2. ?H pc\<^sub>2))")
- proof
- show ?N by simp
- next
- { assume ?eq
- then obtain pc\<^sub>2 vs' where
- pc\<^sub>2: "?pc \<le> pc\<^sub>2 \<and> pc\<^sub>2 < ?pc + size(compE\<^sub>2 e\<^sub>2) \<and>
- \<not> caught P pc\<^sub>2 h\<^sub>2 xa (compxE\<^sub>2 e\<^sub>2 ?pc (size vs + 1))" and
- 2: "P \<turnstile> ?\<sigma>\<^sub>1 -jvm\<rightarrow> handle P C M xa h\<^sub>2 (vs'@v\<^sub>1#vs) ls\<^sub>2 pc\<^sub>2 ics frs sh\<^sub>2"
- using IH\<^sub>2 Jcc_pieces_BinOp2[OF assms(1) pcs, of h\<^sub>1 v\<^sub>1 ls\<^sub>1 sh\<^sub>1 v] BinOpThrow\<^sub>2\<^sub>1.prems by clarsimp
- then have "?H pc\<^sub>2" using jvm_trans[OF 1 2] by(auto intro!: exI[where x="vs'@[v\<^sub>1]"])
- hence "\<exists>pc\<^sub>2. ?H pc\<^sub>2" by iprover
- }
- thus "?eq \<longrightarrow> (\<exists>pc\<^sub>2. ?H pc\<^sub>2)" by iprover
- qed
- then show ?case using pcs by simp blast
-next
- case (FAcc\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 a h\<^sub>1 ls\<^sub>1 sh\<^sub>1 C' fs F T D w)
- then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (e\<bullet>F{D})
- = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (e\<bullet>F{D})),ics)#frs,sh\<^sub>1), err)"
- using FAcc\<^sub>1.prems(1) by clarsimp
- have "P\<^sub>1 \<turnstile> D sees F,NonStatic:T in D" by(rule has_field_sees[OF has_field_idemp[OF FAcc\<^sub>1.hyps(4)]])
- then have field: "field P D F = (D,NonStatic,T)" by simp
- have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (addr a) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Addr a) xa vs frs I" by fact
- let ?pc = "pc + length(compE\<^sub>2 e)"
- have "P \<turnstile> (None,h\<^sub>0,frs',sh\<^sub>0) -jvm\<rightarrow> (None,h\<^sub>1,(Addr a#vs,ls\<^sub>1,C,M,?pc,ics)#frs,sh\<^sub>1)"
- using IH Jcc_pieces_FAcc[OF pcs, of "Addr a"] pcs by simp
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> (None,h\<^sub>1,(w#vs,ls\<^sub>1,C,M,?pc+1,ics)#frs,sh\<^sub>1)"
- using FAcc\<^sub>1 field by auto
- finally have "P \<turnstile> (None, h\<^sub>0, frs', sh\<^sub>0) -jvm\<rightarrow> (None,h\<^sub>1,(w#vs,ls\<^sub>1,C,M,?pc+1,ics)#frs,sh\<^sub>1)"
- by auto
- then show ?case using pcs by auto
-next
- case (FAccNull\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 F D)
- then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (e\<bullet>F{D})
- = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (e\<bullet>F{D})),ics)#frs,sh\<^sub>1), err)"
- using FAccNull\<^sub>1.prems(1) by clarsimp
- have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 null h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics Null xa vs frs I" by fact
- let ?pc = "pc + length(compE\<^sub>2 e)"
- let ?xa = "addr_of_sys_xcpt NullPointer"
- have "P \<turnstile> (None,h\<^sub>0,frs',sh\<^sub>0) -jvm\<rightarrow> (None,h\<^sub>1,(Null#vs,ls\<^sub>1,C,M,?pc,ics)#frs,sh\<^sub>1)"
- using IH Jcc_pieces_FAcc[OF pcs, of Null] by simp
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> handle P C M ?xa h\<^sub>1 (Null#vs) ls\<^sub>1 ?pc ics frs sh\<^sub>1"
- using FAccNull\<^sub>1.prems
- by(fastforce simp:split_beta handle_def simp del: split_paired_Ex)
- finally show ?case using pcs by (auto intro!: exI[where x = ?pc] exI[where x="[Null]"])
-next
- case (FAccThrow\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 e' h\<^sub>1 ls\<^sub>1 sh\<^sub>1 F D)
- then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (e\<bullet>F{D})
- = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (e\<bullet>F{D})),ics)#frs,sh\<^sub>1), err)"
- using FAccThrow\<^sub>1.prems(1) by clarsimp
- have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (throw e') h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics v xa vs frs I" by fact
- show ?case using IH Jcc_pieces_FAcc[OF pcs, of v] FAccThrow\<^sub>1.prems nsub_RI_Jcc_pieces
- less_Suc_eq by auto
-next
- case (FAccNone\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 a h\<^sub>1 ls\<^sub>1 sh\<^sub>1 C fs F D)
- then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (e\<bullet>F{D})
- = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (e\<bullet>F{D})),ics)#frs,sh\<^sub>1), err)"
- using FAccNone\<^sub>1.prems(1) by clarsimp
- have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (addr a) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Addr a) xa vs frs I" by fact
- let ?pc = "pc + length(compE\<^sub>2 e)"
- let ?xa = "addr_of_sys_xcpt NoSuchFieldError"
- have "P \<turnstile> (None,h\<^sub>0,frs',sh\<^sub>0) -jvm\<rightarrow> (None,h\<^sub>1,(Addr a#vs,ls\<^sub>1,C,M,?pc,ics)#frs,sh\<^sub>1)"
- using IH Jcc_pieces_FAcc[OF pcs, of "Addr a"] by simp
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> handle P C M ?xa h\<^sub>1 (Addr a#vs) ls\<^sub>1 ?pc ics frs sh\<^sub>1"
- using FAccNone\<^sub>1
- by(cases ics; clarsimp simp:split_beta handle_def simp del: split_paired_Ex)
- finally show ?case using pcs by (auto intro!: exI[where x = ?pc] exI[where x="[Addr a]"])
-next
- case (FAccStatic\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 a h\<^sub>1 ls\<^sub>1 sh\<^sub>1 C' fs F T D)
- then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (e\<bullet>F{D})
- = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (e\<bullet>F{D})),ics)#frs,sh\<^sub>1), err)"
- using FAccStatic\<^sub>1.prems(1) by clarsimp
- have "P\<^sub>1 \<turnstile> D sees F,Static:T in D" by(rule has_field_sees[OF has_field_idemp[OF FAccStatic\<^sub>1.hyps(4)]])
- then have field: "field P D F = (D,Static,T)" by simp
- have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (addr a) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Addr a) xa vs frs I" by fact
- let ?pc = "pc + length(compE\<^sub>2 e)"
- let ?xa = "addr_of_sys_xcpt IncompatibleClassChangeError"
- have "P \<turnstile> (None,h\<^sub>0,frs',sh\<^sub>0) -jvm\<rightarrow> (None,h\<^sub>1,(Addr a#vs,ls\<^sub>1,C,M,?pc,ics)#frs,sh\<^sub>1)"
- using IH Jcc_pieces_FAcc[OF pcs, of "Addr a"] by simp
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> handle P C M ?xa h\<^sub>1 (Addr a#vs) ls\<^sub>1 ?pc ics frs sh\<^sub>1"
- using FAccStatic\<^sub>1 field by(fastforce simp:split_beta handle_def simp del: split_paired_Ex)
- finally show ?case using pcs by (auto intro!: exI[where x = ?pc] exI[where x="[Addr a]"])
-next
- case (SFAcc\<^sub>1 C' F t D sh sfs v' h ls)
- have has: "P\<^sub>1 \<turnstile> D has F,Static:t in D" by(rule has_field_idemp[OF SFAcc\<^sub>1.hyps(1)])
- have "P\<^sub>1 \<turnstile> D sees F,Static:t in D" by(rule has_field_sees[OF has])
- then have field: "field P D F = (D,Static,t)" by simp
- then have "P \<turnstile> (None,h,Jcc_frames P C M vs ls pc ics frs (C'\<bullet>\<^sub>sF{D}),sh) -jvm\<rightarrow>
- (None,h,(v'#vs,ls,C,M,Suc pc,ics)#frs,sh)"
- using SFAcc\<^sub>1 has by(cases ics) auto
- then show ?case by clarsimp
-next
- case (SFAccInit\<^sub>1 C' F t D sh h ls v' h' ls' sh' sfs i v'')
- then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs ls pc ics frs sh I h' ls' sh' v xa (C'\<bullet>\<^sub>sF{D})
- = (True, frs', (None,h',(v#vs,ls',C,M,pc+size(compE\<^sub>2 (C'\<bullet>\<^sub>sF{D})),ics)#frs,sh'), err)"
- using SFAccInit\<^sub>1.prems(1) by clarsimp
- have "Ex (WTrt2\<^sub>1 P\<^sub>1 E h sh (INIT D ([D],False) \<leftarrow> unit))"
- using has_field_is_class'[OF SFAccInit\<^sub>1.hyps(1)] by auto
- then obtain err' where pcs':
- "Jcc_pieces P\<^sub>1 E C M h vs ls pc ics frs sh I h' ls' sh' v' xa (INIT D ([D],False) \<leftarrow> unit)
- = (True, (vs,ls,C,M,pc,Calling D []) # frs, (None,h',(vs,ls,C,M,pc,Called [])#frs,sh'), err')"
- using SFAccInit\<^sub>1.prems(1) by auto
- have IH: "PROP ?P (INIT D ([D],False) \<leftarrow> unit) h ls sh (Val v')
- h' ls' sh' E C M pc ics v' xa vs frs I" by fact
- have ls: "ls = ls'" by(rule init\<^sub>1_same_loc[OF SFAccInit\<^sub>1.hyps(3)])
- have has: "P\<^sub>1 \<turnstile> D has F,Static:t in D" by(rule has_field_idemp[OF SFAccInit\<^sub>1.hyps(1)])
- have "P\<^sub>1 \<turnstile> D sees F,Static:t in D" by(rule has_field_sees[OF has])
- then have field: "field P D F = (D,Static,t)" by simp
- have "P \<turnstile> (None,h,(vs,ls,C,M,pc,ics)#frs,sh) -jvm\<rightarrow> (None,h,(vs,ls,C,M,pc,Calling D [])#frs,sh)"
- proof(cases "sh D")
- case None then show ?thesis using SFAccInit\<^sub>1.hyps(1,2,5,6) SFAccInit\<^sub>1.prems field
- by(cases ics) auto
- next
- case (Some a)
- then obtain sfs i where "a = (sfs,i)" by(cases a)
- then show ?thesis using SFAccInit\<^sub>1.hyps(1,2,5,6) SFAccInit\<^sub>1.prems field Some
- by(cases ics; case_tac i) auto
- qed
- also have "P \<turnstile> ... -jvm\<rightarrow> (None, h', (vs, ls, C, M, pc, Called []) # frs, sh')"
- using IH pcs' by auto
- also have "P \<turnstile> ... -jvm\<rightarrow> (None, h', (v''#vs, ls, C, M, Suc pc, ics) # frs, sh')"
- using SFAccInit\<^sub>1.hyps(1,2,5,6) SFAccInit\<^sub>1.prems has field by(cases ics) auto
- finally show ?case using pcs ls by clarsimp
-next
- case (SFAccInitThrow\<^sub>1 C' F t D sh h ls a h' ls' sh')
- obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs ls pc ics frs sh I h' ls' sh' v xa (C'\<bullet>\<^sub>sF{D})
- = (True, frs', (None,h',(v#vs,ls',C,M,pc+size(compE\<^sub>2 (C'\<bullet>\<^sub>sF{D})),ics)#frs,sh'), err)"
- using SFAccInitThrow\<^sub>1.prems(1) by clarsimp
- obtain a' where throw: "throw a = Throw a'" using eval\<^sub>1_final[OF SFAccInitThrow\<^sub>1.hyps(3)] by clarsimp
- have "Ex (WTrt2\<^sub>1 P\<^sub>1 E h sh (INIT D ([D],False) \<leftarrow> unit))"
- using has_field_is_class'[OF SFAccInitThrow\<^sub>1.hyps(1)] by auto
- then obtain vs' where pcs':
- "Jcc_pieces P\<^sub>1 E C M h vs ls pc ics frs sh I h' ls' sh' v a' (INIT D ([D],False) \<leftarrow> unit)
- = (True, (vs,ls,C,M,pc,Calling D []) # frs, (None,h',(vs,ls,C,M,pc,Called [])#frs,sh'),
- P \<turnstile> (None,h,(vs,ls,C,M,pc,Calling D []) # frs,sh)
- -jvm\<rightarrow> handle P C M a' h' (vs'@vs) ls pc ics frs sh')"
- using SFAccInitThrow\<^sub>1.prems(1) by simp blast
- have IH: "PROP ?P (INIT D ([D],False) \<leftarrow> unit) h ls sh (throw a)
- h' ls' sh' E C M pc ics v a' vs frs I" by fact
- have ls: "ls = ls'" by(rule init\<^sub>1_same_loc[OF SFAccInitThrow\<^sub>1.hyps(3)])
- have has: "P\<^sub>1 \<turnstile> D has F,Static:t in D" by(rule has_field_idemp[OF SFAccInitThrow\<^sub>1.hyps(1)])
- have "P\<^sub>1 \<turnstile> D sees F,Static:t in D" by(rule has_field_sees[OF has])
- then have field: "field P D F = (D,Static,t)" by simp
- then have "P \<turnstile> (None,h,(vs,ls,C,M,pc,ics)#frs,sh) -jvm\<rightarrow> (None,h,(vs,ls,C,M,pc,Calling D []) # frs,sh)"
- proof(cases "sh D")
- case None then show ?thesis using SFAccInitThrow\<^sub>1.hyps(1,2) SFAccInitThrow\<^sub>1.prems field
- by(cases ics) auto
- next
- case (Some a)
- then obtain sfs i where "a = (sfs,i)" by(cases a)
- then show ?thesis using SFAccInitThrow\<^sub>1.hyps(1,2) SFAccInitThrow\<^sub>1.prems field Some
- by(cases ics; case_tac i) auto
- qed
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> handle P C M a' h' (vs'@vs) ls pc ics frs sh'"
- using IH pcs' throw by auto
- finally show ?case using throw ls by auto
-next
- case (SFAccNone\<^sub>1 C' F D h\<^sub>1 ls\<^sub>1 sh\<^sub>1)
- then obtain frs' err where pcs:
- "Jcc_pieces P\<^sub>1 E C M h\<^sub>1 vs ls\<^sub>1 pc ics frs sh\<^sub>1 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (C'\<bullet>\<^sub>sF{D})
- = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (C'\<bullet>\<^sub>sF{D})),ics)#frs,sh\<^sub>1), err)"
- by clarsimp
- let ?xa = "addr_of_sys_xcpt NoSuchFieldError"
- have "P \<turnstile> (None,h\<^sub>1,frs',sh\<^sub>1) -jvm\<rightarrow> handle P C M ?xa h\<^sub>1 vs ls\<^sub>1 pc ics frs sh\<^sub>1"
- using SFAccNone\<^sub>1 pcs
- by(cases ics; clarsimp simp:split_beta handle_def simp del: split_paired_Ex)
- then show ?case using pcs by(auto intro!: exI[where x = pc] exI[where x="[]"])
-next
- case (SFAccNonStatic\<^sub>1 C' F t D h\<^sub>1 ls\<^sub>1 sh\<^sub>1)
- let ?frs' = "(vs, ls\<^sub>1, C, M, pc, ics) # frs"
- let ?xa = "addr_of_sys_xcpt IncompatibleClassChangeError"
- have "P\<^sub>1 \<turnstile> D sees F,NonStatic:t in D"
- by(rule has_field_sees[OF has_field_idemp[OF SFAccNonStatic\<^sub>1.hyps(1)]])
- then have field: "field P D F = (D,NonStatic,t)" by simp
- have "P \<turnstile> (None,h\<^sub>1,?frs',sh\<^sub>1) -jvm\<rightarrow> handle P C M ?xa h\<^sub>1 vs ls\<^sub>1 pc ics frs sh\<^sub>1"
- using SFAccNonStatic\<^sub>1
- proof(cases ics)
- case No_ics
- then show ?thesis using SFAccNonStatic\<^sub>1 field
- by (auto simp:split_beta handle_def simp del: split_paired_Ex)
- qed(simp_all)
- then show ?case by (auto intro!: exI[where x = pc] exI[where x="[]"])
-next
- case (LAss\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 w h\<^sub>1 ls\<^sub>1 sh\<^sub>1 i ls\<^sub>2)
- let ?pc = "pc + length(compE\<^sub>2 e)"
- obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (i:=e)
- = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (i:=e)),ics)#frs,sh\<^sub>1), err)"
- using LAss\<^sub>1.prems(1) by auto
- have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (Val w) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics w xa vs frs I" by fact
- then have "P \<turnstile> (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0) -jvm\<rightarrow>
- (None,h\<^sub>1,(w#vs,ls\<^sub>1,C,M,?pc,ics)#frs,sh\<^sub>1)"
- using Jcc_pieces_LAss[OF assms(1) pcs, of w] LAss\<^sub>1.prems pcs by auto
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> (None,h\<^sub>1,(Unit#vs,ls\<^sub>2,C,M,?pc+2,ics)#frs,sh\<^sub>1)"
- using LAss\<^sub>1 by (auto simp add:cast_ok_def)
- finally show ?case by auto
-next
- case (LAssThrow\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 w h\<^sub>1 ls\<^sub>1 sh\<^sub>1 i)
- obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (i:=e)
- = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (i:=e)),ics)#frs,sh\<^sub>1), err)"
- using LAssThrow\<^sub>1.prems(1) by clarsimp
- have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (throw w) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics v xa vs frs I" by fact
- show ?case using IH Jcc_pieces_LAss[OF assms(1) pcs, of v] LAssThrow\<^sub>1.prems pcs less_SucI
- by(simp, blast)
-next
- case (FAss\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 a h\<^sub>1 ls\<^sub>1 sh\<^sub>1 e\<^sub>2 w h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C' fs F T D fs' h\<^sub>2')
- obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v xa (e\<^sub>1\<bullet>F{D} := e\<^sub>2)
- = (True, frs', (None,h\<^sub>2,(v#vs,ls\<^sub>2,C,M,pc+size(compE\<^sub>2 (e\<^sub>1\<bullet>F{D} := e\<^sub>2)),ics)#frs,sh\<^sub>2), err)"
- using FAss\<^sub>1.prems(1) by clarsimp
- have "P\<^sub>1 \<turnstile> D sees F,NonStatic:T in D" by(rule has_field_sees[OF has_field_idemp[OF FAss\<^sub>1.hyps(6)]])
- then have field: "field P D F = (D,NonStatic,T)" by simp
- let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e\<^sub>1)"
- let ?pc\<^sub>2 = "?pc\<^sub>1 + length(compE\<^sub>2 e\<^sub>2)"
- have IH\<^sub>1: "PROP ?P e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (addr a) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Addr a) xa vs frs
- (I - pcs (compxE\<^sub>2 e\<^sub>2 (pc + length (compE\<^sub>2 e\<^sub>1)) (Suc (length vs))))" by fact
- have IH\<^sub>2: "PROP ?P e\<^sub>2 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (Val w) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 E C M ?pc\<^sub>1 ics w xa (Addr a#vs) frs
- (I - pcs(compxE\<^sub>2 e\<^sub>1 pc (size vs)))" by fact
- have "P \<turnstile> (None,h\<^sub>0,frs',sh\<^sub>0) -jvm\<rightarrow> (None,h\<^sub>1,(Addr a#vs,ls\<^sub>1,C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)"
- using IH\<^sub>1 Jcc_pieces_FAss1[OF pcs, of h\<^sub>1 ls\<^sub>1 sh\<^sub>1 "Addr a"] by simp
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> (None,h\<^sub>2,(w#Addr a#vs,ls\<^sub>2,C,M,?pc\<^sub>2,ics)#frs,sh\<^sub>2)"
- using IH\<^sub>2 Jcc_pieces_FAss2[OF pcs, of h\<^sub>1 "Addr a" ls\<^sub>1 sh\<^sub>1 w] by (simp add: add.assoc)
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> (None,h\<^sub>2',(Unit#vs,ls\<^sub>2,C,M,?pc\<^sub>2+2,ics)#frs,sh\<^sub>2)"
- using FAss\<^sub>1 field by auto
- finally show ?case using pcs by (auto simp:add.assoc)
-next
- case (FAssNull\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 e\<^sub>2 w h\<^sub>2 ls\<^sub>2 sh\<^sub>2 F D)
- obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v xa (e\<^sub>1\<bullet>F{D} := e\<^sub>2)
- = (True, frs', (None,h\<^sub>2,(v#vs,ls\<^sub>2,C,M,pc+size(compE\<^sub>2 (e\<^sub>1\<bullet>F{D} := e\<^sub>2)),ics)#frs,sh\<^sub>2), err)"
- using FAssNull\<^sub>1.prems(1) by clarsimp
- let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e\<^sub>1)"
- let ?pc\<^sub>2 = "?pc\<^sub>1 + length(compE\<^sub>2 e\<^sub>2)"
- let ?xa = "addr_of_sys_xcpt NullPointer"
- have IH\<^sub>1: "PROP ?P e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 null h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics Null xa vs frs
- (I - pcs (compxE\<^sub>2 e\<^sub>2 (pc + length (compE\<^sub>2 e\<^sub>1)) (Suc (length vs))))" by fact
- have IH\<^sub>2: "PROP ?P e\<^sub>2 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (Val w) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 E C M ?pc\<^sub>1 ics w xa (Null#vs) frs
- (I - pcs(compxE\<^sub>2 e\<^sub>1 pc (size vs)))" by fact
- have "P \<turnstile> (None,h\<^sub>0,frs',sh\<^sub>0) -jvm\<rightarrow> (None,h\<^sub>1,(Null#vs,ls\<^sub>1,C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)"
- using IH\<^sub>1 Jcc_pieces_FAss1[OF pcs, of h\<^sub>1 ls\<^sub>1 sh\<^sub>1 Null] by simp
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> (None,h\<^sub>2,(w#Null#vs,ls\<^sub>2,C,M,?pc\<^sub>2,ics)#frs,sh\<^sub>2)"
- using IH\<^sub>2 Jcc_pieces_FAss2[OF pcs, of h\<^sub>1 Null ls\<^sub>1 sh\<^sub>1 w] by (simp add: add.assoc)
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> handle P C M ?xa h\<^sub>2 (w#Null#vs) ls\<^sub>2 ?pc\<^sub>2 ics frs sh\<^sub>2"
- using FAssNull\<^sub>1 by(fastforce simp:split_beta handle_def simp del: split_paired_Ex)
- finally show ?case using pcs by (auto intro!: exI[where x = ?pc\<^sub>2] exI[where x="w#[Null]"])
-next
- case (FAssThrow\<^sub>2\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 w h\<^sub>1 ls\<^sub>1 sh\<^sub>1 e\<^sub>2 e' h\<^sub>2 ls\<^sub>2 sh\<^sub>2 F D)
- let ?frs' = "(vs, ls\<^sub>0, C, M, pc, ics) # frs"
- obtain err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v xa (e\<^sub>1\<bullet>F{D} := e\<^sub>2)
- = (True, ?frs', (None,h\<^sub>2,(v#vs,ls\<^sub>2,C,M,pc+size(compE\<^sub>2 (e\<^sub>1\<bullet>F{D} := e\<^sub>2)),ics)#frs,sh\<^sub>2), err)"
- using FAssThrow\<^sub>2\<^sub>1.prems(1) by clarsimp
- let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e\<^sub>1)"
- let ?\<sigma>\<^sub>1 = "(None,h\<^sub>1,(w#vs,ls\<^sub>1,C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)"
- have IH\<^sub>1: "PROP ?P e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (Val w) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics w xa vs frs
- (I - pcs (compxE\<^sub>2 e\<^sub>2 (pc + length (compE\<^sub>2 e\<^sub>1)) (Suc (length vs))))" by fact
- have IH\<^sub>2: "PROP ?P e\<^sub>2 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (throw e') h\<^sub>2 ls\<^sub>2 sh\<^sub>2 E C M ?pc\<^sub>1 ics v xa (w#vs) frs
- (I - pcs(compxE\<^sub>2 e\<^sub>1 pc (size vs)))" by fact
- have 1: "P \<turnstile> (None,h\<^sub>0,?frs',sh\<^sub>0) -jvm\<rightarrow> ?\<sigma>\<^sub>1"
- using IH\<^sub>1 Jcc_pieces_FAss1[OF pcs, of h\<^sub>1 ls\<^sub>1 sh\<^sub>1 w] by simp
- show ?case (is "?N \<and> (?eq \<longrightarrow> ?err)")
- proof
- show ?N by simp
- next
- { assume ?eq
- moreover
- have "PROP ?P e\<^sub>2 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (throw e') h\<^sub>2 ls\<^sub>2 sh\<^sub>2 E C M ?pc\<^sub>1 ics v xa (w#vs) frs
- (I - pcs (compxE\<^sub>2 e\<^sub>1 pc (length vs)))" by fact
- ultimately obtain pc\<^sub>2 vs' where
- pc\<^sub>2: "?pc\<^sub>1 \<le> pc\<^sub>2 \<and> pc\<^sub>2 < ?pc\<^sub>1 + size(compE\<^sub>2 e\<^sub>2) \<and>
- \<not> caught P pc\<^sub>2 h\<^sub>2 xa (compxE\<^sub>2 e\<^sub>2 ?pc\<^sub>1 (size vs + 1))" and
- 2: "P \<turnstile> ?\<sigma>\<^sub>1 -jvm\<rightarrow> handle P C M xa h\<^sub>2 (vs'@w#vs) ls\<^sub>2 pc\<^sub>2 ics frs sh\<^sub>2"
- using FAssThrow\<^sub>2\<^sub>1.prems Jcc_pieces_FAss2[OF pcs, of h\<^sub>1 w ls\<^sub>1 sh\<^sub>1] by auto
- have ?err using Jcc_pieces_FAss2[OF pcs, of h\<^sub>1 w ls\<^sub>1 sh\<^sub>1] pc\<^sub>2 jvm_trans[OF 1 2]
- by(auto intro!: exI[where x=pc\<^sub>2] exI[where x="vs'@[w]"])
- }
- thus "?eq \<longrightarrow> ?err" by simp
- qed
-next
- case (FAssThrow\<^sub>1\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 e' h\<^sub>1 ls\<^sub>1 sh\<^sub>1 F D e\<^sub>2)
- obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (e\<^sub>1\<bullet>F{D} := e\<^sub>2)
- = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (e\<^sub>1\<bullet>F{D} := e\<^sub>2)),ics)#frs,sh\<^sub>1), err)"
- using FAssThrow\<^sub>1\<^sub>1.prems(1) by clarsimp
- have IH\<^sub>1: "PROP ?P e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (throw e') h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics v xa vs frs
- (I - pcs (compxE\<^sub>2 e\<^sub>2 (pc + length (compE\<^sub>2 e\<^sub>1)) (Suc (length vs))))" by fact
- show ?case using IH\<^sub>1 Jcc_pieces_FAss1[OF pcs, of h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v] FAssThrow\<^sub>1\<^sub>1.prems nsub_RI_Jcc_pieces
- by auto
-next
- case (FAssNone\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 a h\<^sub>1 ls\<^sub>1 sh\<^sub>1 e\<^sub>2 w h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C' fs F D)
- obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v xa (e\<^sub>1\<bullet>F{D} := e\<^sub>2)
- = (True, frs', (None,h\<^sub>2,(v#vs,ls\<^sub>2,C,M,pc+size(compE\<^sub>2 (e\<^sub>1\<bullet>F{D} := e\<^sub>2)),ics)#frs,sh\<^sub>2), err)"
- using FAssNone\<^sub>1.prems(1) by clarsimp
- let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e\<^sub>1)"
- let ?pc\<^sub>2 = "?pc\<^sub>1 + length(compE\<^sub>2 e\<^sub>2)"
- let ?xa = "addr_of_sys_xcpt NoSuchFieldError"
- have IH\<^sub>1: "PROP ?P e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (addr a) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Addr a) xa vs frs
- (I - pcs (compxE\<^sub>2 e\<^sub>2 (pc + length (compE\<^sub>2 e\<^sub>1)) (Suc (length vs))))" by fact
- have IH\<^sub>2: "PROP ?P e\<^sub>2 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (Val w) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 E C M ?pc\<^sub>1 ics w xa (Addr a#vs) frs
- (I - pcs(compxE\<^sub>2 e\<^sub>1 pc (size vs)))" by fact
- have "P \<turnstile> (None,h\<^sub>0,frs',sh\<^sub>0) -jvm\<rightarrow> (None,h\<^sub>1,(Addr a#vs,ls\<^sub>1,C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)"
- using IH\<^sub>1 Jcc_pieces_FAss1[OF pcs, of h\<^sub>1 ls\<^sub>1 sh\<^sub>1 "Addr a"] by simp
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> (None,h\<^sub>2,(w#Addr a#vs,ls\<^sub>2,C,M,?pc\<^sub>2,ics)#frs,sh\<^sub>2)"
- using IH\<^sub>2 Jcc_pieces_FAss2[OF pcs, of h\<^sub>1 "Addr a" ls\<^sub>1 sh\<^sub>1 w] by (simp add: add.assoc)
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> handle P C M ?xa h\<^sub>2 (w#Addr a#vs) ls\<^sub>2 ?pc\<^sub>2 ics frs sh\<^sub>2"
- using FAssNone\<^sub>1 by(fastforce simp:split_beta handle_def simp del: split_paired_Ex)
- finally show ?case using pcs by (auto intro!: exI[where x = ?pc\<^sub>2] exI[where x="w#[Addr a]"])
-next
- case (FAssStatic\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 a h\<^sub>1 ls\<^sub>1 sh\<^sub>1 e\<^sub>2 w h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C' fs F T D)
- obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v xa (e\<^sub>1\<bullet>F{D} := e\<^sub>2)
- = (True, frs', (None,h\<^sub>2,(v#vs,ls\<^sub>2,C,M,pc+size(compE\<^sub>2 (e\<^sub>1\<bullet>F{D} := e\<^sub>2)),ics)#frs,sh\<^sub>2), err)"
- using FAssStatic\<^sub>1.prems(1) by clarsimp
- have "P\<^sub>1 \<turnstile> D sees F,Static:T in D" by(rule has_field_sees[OF has_field_idemp[OF FAssStatic\<^sub>1.hyps(6)]])
- then have field: "field P D F = (D,Static,T)" by simp
- let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e\<^sub>1)"
- let ?pc\<^sub>2 = "?pc\<^sub>1 + length(compE\<^sub>2 e\<^sub>2)"
- let ?xa = "addr_of_sys_xcpt IncompatibleClassChangeError"
- have IH\<^sub>1: "PROP ?P e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (addr a) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Addr a) xa vs frs
- (I - pcs (compxE\<^sub>2 e\<^sub>2 (pc + length (compE\<^sub>2 e\<^sub>1)) (Suc (length vs))))" by fact
- have IH\<^sub>2: "PROP ?P e\<^sub>2 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (Val w) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 E C M ?pc\<^sub>1 ics w xa (Addr a#vs) frs
- (I - pcs(compxE\<^sub>2 e\<^sub>1 pc (size vs)))" by fact
- have "P \<turnstile> (None,h\<^sub>0,frs',sh\<^sub>0) -jvm\<rightarrow> (None,h\<^sub>1,(Addr a#vs,ls\<^sub>1,C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)"
- using IH\<^sub>1 Jcc_pieces_FAss1[OF pcs, of h\<^sub>1 ls\<^sub>1 sh\<^sub>1 "Addr a"] by simp
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> (None,h\<^sub>2,(w#Addr a#vs,ls\<^sub>2,C,M,?pc\<^sub>2,ics)#frs,sh\<^sub>2)"
- using IH\<^sub>2 Jcc_pieces_FAss2[OF pcs, of h\<^sub>1 "Addr a" ls\<^sub>1 sh\<^sub>1 w] by (simp add: add.assoc)
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> handle P C M ?xa h\<^sub>2 (w#Addr a#vs) ls\<^sub>2 ?pc\<^sub>2 ics frs sh\<^sub>2"
- using FAssStatic\<^sub>1 field by(fastforce simp:split_beta handle_def simp del: split_paired_Ex)
- finally show ?case using pcs by (auto intro!: exI[where x = ?pc\<^sub>2] exI[where x="w#[Addr a]"])
-next
- case (SFAss\<^sub>1 e\<^sub>2 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 w h\<^sub>1 ls\<^sub>1 sh\<^sub>1 C' F T D sfs sfs' sh\<^sub>1')
- then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (C'\<bullet>\<^sub>sF{D} := e\<^sub>2)
- = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (C'\<bullet>\<^sub>sF{D} := e\<^sub>2)),ics)#frs,sh\<^sub>1), err)"
- using SFAss\<^sub>1.prems(1) by clarsimp
- have "P\<^sub>1 \<turnstile> D sees F,Static:T in D" by(rule has_field_sees[OF has_field_idemp[OF SFAss\<^sub>1.hyps(3)]])
- then have field: "field P D F = (D,Static,T)" by simp
- have IH: "PROP ?P e\<^sub>2 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (Val w) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics w xa vs frs I" by fact
- let ?pc = "pc + length(compE\<^sub>2 e\<^sub>2)"
- have "P \<turnstile> (None,h\<^sub>0,frs',sh\<^sub>0) -jvm\<rightarrow> (None,h\<^sub>1,(w#vs,ls\<^sub>1,C,M,?pc,ics)#frs,sh\<^sub>1)"
- using IH Jcc_pieces_SFAss[OF pcs, where v'=w] pcs by simp
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> (None,h\<^sub>1,(vs,ls\<^sub>1,C,M,?pc+1,ics)#frs,sh\<^sub>1')"
- using SFAss\<^sub>1.hyps(3-6) SFAss\<^sub>1.prems(1) field by auto
- also have "P \<turnstile> ... -jvm\<rightarrow> (None,h\<^sub>1,(Unit#vs,ls\<^sub>1,C,M,?pc+2,ics)#frs,sh\<^sub>1')"
- using SFAss\<^sub>1 by auto
- finally show ?case using pcs by auto
-next
- case (SFAssInit\<^sub>1 e\<^sub>2 h ls sh w h\<^sub>1 ls\<^sub>1 sh\<^sub>1 C' F t D v' h' ls' sh' sfs i sfs' sh'')
- let ?pc = "pc + length(compE\<^sub>2 e\<^sub>2)"
- obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs ls pc ics frs sh I h' ls' sh'' v xa (C'\<bullet>\<^sub>sF{D}:=e\<^sub>2)
- = (True, frs', (None,h',(v#vs,ls',C,M,pc+size(compE\<^sub>2 (C'\<bullet>\<^sub>sF{D}:=e\<^sub>2)),ics)#frs,sh''), err)"
- using SFAssInit\<^sub>1.prems(1) by clarsimp
- have "Ex (WTrt2\<^sub>1 P\<^sub>1 E h\<^sub>1 sh\<^sub>1 (INIT D ([D],False) \<leftarrow> unit))"
- using has_field_is_class'[OF SFAssInit\<^sub>1.hyps(3)] by auto
- then obtain err' where pcs':
- "Jcc_pieces P\<^sub>1 E C M h\<^sub>1 (w#vs) ls\<^sub>1 ?pc ics frs sh\<^sub>1 I h' ls' sh' v' xa (INIT D ([D],False) \<leftarrow> unit)
- = (True, (w#vs,ls\<^sub>1,C,M,?pc,Calling D []) # frs,
- (None,h',(w#vs,ls\<^sub>1,C,M,?pc,Called [])#frs,sh'), err')"
- using SFAssInit\<^sub>1.prems(1) by simp
- have ls: "ls\<^sub>1 = ls'" by(rule init\<^sub>1_same_loc[OF SFAssInit\<^sub>1.hyps(5)])
- have has: "P\<^sub>1 \<turnstile> D has F,Static:t in D" by(rule has_field_idemp[OF SFAssInit\<^sub>1.hyps(3)])
- have "P\<^sub>1 \<turnstile> D sees F,Static:t in D" by(rule has_field_sees[OF has])
- then have field: "field P D F = (D,Static,t)" by simp
- have IH: "PROP ?P e\<^sub>2 h ls sh (Val w) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics w xa vs frs I" by fact
- have IHI: "PROP ?P (INIT D ([D],False) \<leftarrow> unit) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (Val v')
- h' ls' sh' E C M ?pc ics v' xa (w#vs) frs I" by fact
- have "P \<turnstile> (None,h,frs',sh) -jvm\<rightarrow> (None,h\<^sub>1,(w#vs,ls\<^sub>1,C,M,?pc,ics)#frs,sh\<^sub>1)"
- using IH Jcc_pieces_SFAss[OF pcs, where v'=w] by simp
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> (None,h\<^sub>1,(w#vs,ls\<^sub>1,C,M,?pc,Calling D [])#frs,sh\<^sub>1)"
- proof(cases "sh\<^sub>1 D")
- case None then show ?thesis using None SFAssInit\<^sub>1.hyps(1,3-5,7-9) SFAssInit\<^sub>1.prems field
- by(cases ics, auto)
- next
- case (Some a)
- then obtain sfs i where "a = (sfs,i)" by(cases a)
- then show ?thesis using SFAssInit\<^sub>1.hyps(1,3-5,7-9) SFAssInit\<^sub>1.prems field Some
- by(cases ics; case_tac i) auto
- qed
- also have "P \<turnstile> ... -jvm\<rightarrow> (None, h', (w#vs, ls\<^sub>1, C, M, ?pc, Called []) # frs, sh')"
- using IHI pcs' by clarsimp
- also have "P \<turnstile> ... -jvm\<rightarrow> (None, h', (vs, ls\<^sub>1, C, M, ?pc + 1, ics) # frs, sh'')"
- using SFAssInit\<^sub>1.hyps(1,3-5,7-9) SFAssInit\<^sub>1.prems has field by(cases ics) auto
- also have "P \<turnstile> ... -jvm\<rightarrow> (None, h', (Unit#vs, ls\<^sub>1, C, M, ?pc + 2, ics) # frs, sh'')"
- using SFAssInit\<^sub>1.hyps(1,3-5,7-9) SFAssInit\<^sub>1.prems has field by(cases ics) auto
- finally show ?case using pcs ls by simp blast
-next
- case (SFAssInitThrow\<^sub>1 e\<^sub>2 h ls sh w h\<^sub>1 ls\<^sub>1 sh\<^sub>1 C' F t D a h' ls' sh')
- let ?pc = "pc + length(compE\<^sub>2 e\<^sub>2)"
- obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs ls pc ics frs sh I h' ls' sh' v xa (C'\<bullet>\<^sub>sF{D}:=e\<^sub>2)
- = (True, frs', (None,h',(v#vs,ls',C,M,pc+size(compE\<^sub>2 (C'\<bullet>\<^sub>sF{D}:=e\<^sub>2)),ics)#frs,sh'), err)"
- using SFAssInitThrow\<^sub>1.prems(1) by clarsimp
- obtain a' where throw: "throw a = Throw a'" using eval\<^sub>1_final[OF SFAssInitThrow\<^sub>1.hyps(5)] by clarsimp
- have "Ex (WTrt2\<^sub>1 P\<^sub>1 E h\<^sub>1 sh\<^sub>1 (INIT D ([D],False) \<leftarrow> unit))"
- using has_field_is_class'[OF SFAssInitThrow\<^sub>1.hyps(3)] by auto
- then obtain vs' where pcs':
- "Jcc_pieces P\<^sub>1 E C M h\<^sub>1 (w#vs) ls\<^sub>1 ?pc ics frs sh\<^sub>1 I h' ls' sh' v a' (INIT D ([D],False) \<leftarrow> unit)
- = (True, (w#vs,ls\<^sub>1,C,M,?pc,Calling D []) # frs, (None,h',(w#vs,ls\<^sub>1,C,M,?pc,Called [])#frs,sh'),
- P \<turnstile> (None,h\<^sub>1,(w#vs,ls\<^sub>1,C,M,?pc,Calling D []) # frs,sh\<^sub>1)
- -jvm\<rightarrow> handle P C M a' h' (vs'@w#vs) ls\<^sub>1 ?pc ics frs sh')"
- using SFAssInitThrow\<^sub>1.prems(1) by simp blast
- have ls: "ls\<^sub>1 = ls'" by(rule init\<^sub>1_same_loc[OF SFAssInitThrow\<^sub>1.hyps(5)])
- have has: "P\<^sub>1 \<turnstile> D has F,Static:t in D" by(rule has_field_idemp[OF SFAssInitThrow\<^sub>1.hyps(3)])
- have "P\<^sub>1 \<turnstile> D sees F,Static:t in D" by(rule has_field_sees[OF has])
- then have field: "field P D F = (D,Static,t)" by simp
- have IH: "PROP ?P e\<^sub>2 h ls sh (Val w) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics w xa vs frs I" by fact
- have IHI: "PROP ?P (INIT D ([D],False) \<leftarrow> unit) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (throw a)
- h' ls' sh' E C M ?pc ics v a' (w#vs) frs I" by fact
- have "P \<turnstile> (None,h,(vs, ls, C, M, pc, ics) # frs,sh) -jvm\<rightarrow> (None,h\<^sub>1,(w#vs,ls\<^sub>1,C,M,?pc,ics)#frs,sh\<^sub>1)"
- using IH Jcc_pieces_SFAss[OF pcs, where v'=w] pcs by simp blast
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> (None,h\<^sub>1,(w#vs,ls\<^sub>1,C,M,?pc,Calling D [])#frs,sh\<^sub>1)"
- proof(cases "sh\<^sub>1 D")
- case None then show ?thesis using SFAssInitThrow\<^sub>1.hyps(1,3,4,5) SFAssInitThrow\<^sub>1.prems field
- by(cases ics) auto
- next
- case (Some a)
- then obtain sfs i where "a = (sfs,i)" by(cases a)
- then show ?thesis using SFAssInitThrow\<^sub>1.hyps(1,3,4,5) SFAssInitThrow\<^sub>1.prems field Some
- by(cases ics; case_tac i) auto
- qed
- also have "P \<turnstile> ... -jvm\<rightarrow> handle P C M a' h' (vs'@w#vs) ls\<^sub>1 ?pc ics frs sh'"
- using IHI pcs' throw by auto
- finally show ?case using throw ls by(auto intro!: exI[where x = ?pc] exI[where x="vs'@[w]"])
-next
- case (SFAssThrow\<^sub>1 e\<^sub>2 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 e' h\<^sub>1 ls\<^sub>1 sh\<^sub>1 C' F D)
- then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (C'\<bullet>\<^sub>sF{D} := e\<^sub>2)
- = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (C'\<bullet>\<^sub>sF{D} := e\<^sub>2)),ics)#frs,sh\<^sub>1), err)"
- using SFAssThrow\<^sub>1.prems(1) by clarsimp
- have IH: "PROP ?P e\<^sub>2 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (throw e') h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics v xa vs frs I" by fact
- show ?case using IH Jcc_pieces_SFAss[OF pcs, where v'=v] SFAssThrow\<^sub>1.prems nsub_RI_Jcc_pieces
- less_Suc_eq by auto
-next
- case (SFAssNone\<^sub>1 e\<^sub>2 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 w h\<^sub>1 ls\<^sub>1 sh\<^sub>1 C' F D)
- then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (C'\<bullet>\<^sub>sF{D} := e\<^sub>2)
- = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (C'\<bullet>\<^sub>sF{D} := e\<^sub>2)),ics)#frs,sh\<^sub>1), err)"
- using SFAssNone\<^sub>1.prems(1) by clarsimp
- have IH: "PROP ?P e\<^sub>2 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (Val w) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics w xa vs frs I" by fact
- let ?pc = "pc + length(compE\<^sub>2 e\<^sub>2)"
- let ?xa = "addr_of_sys_xcpt NoSuchFieldError"
- have "P \<turnstile> (None,h\<^sub>0,frs',sh\<^sub>0) -jvm\<rightarrow> (None,h\<^sub>1,(w#vs,ls\<^sub>1,C,M,?pc,ics)#frs,sh\<^sub>1)"
- using IH Jcc_pieces_SFAss[OF pcs, where v'=w] pcs by simp
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> handle P C M ?xa h\<^sub>1 (w#vs) ls\<^sub>1 ?pc ics frs sh\<^sub>1"
- using SFAssNone\<^sub>1 by(cases ics; clarsimp simp add: handle_def)
- finally show ?case using pcs by (auto intro!: exI[where x = ?pc] exI[where x="[w]"])
-next
- case (SFAssNonStatic\<^sub>1 e\<^sub>2 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 w h\<^sub>1 ls\<^sub>1 sh\<^sub>1 C' F T D)
- then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (C'\<bullet>\<^sub>sF{D} := e\<^sub>2)
- = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (C'\<bullet>\<^sub>sF{D} := e\<^sub>2)),ics)#frs,sh\<^sub>1), err)"
- using SFAssNonStatic\<^sub>1.prems(1) by clarsimp
- have IH: "PROP ?P e\<^sub>2 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (Val w) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics w xa vs frs I" by fact
- let ?pc = "pc + length(compE\<^sub>2 e\<^sub>2)"
- let ?xa = "addr_of_sys_xcpt IncompatibleClassChangeError"
- have "P\<^sub>1 \<turnstile> D sees F,NonStatic:T in D"
- by(rule has_field_sees[OF has_field_idemp[OF SFAssNonStatic\<^sub>1.hyps(3)]])
- then have field: "field P D F = (D,NonStatic,T)" by simp
- have "P \<turnstile> (None,h\<^sub>0,frs',sh\<^sub>0) -jvm\<rightarrow> (None,h\<^sub>1,(w#vs,ls\<^sub>1,C,M,?pc,ics)#frs,sh\<^sub>1)"
- using IH Jcc_pieces_SFAss[OF pcs, where v'=w] pcs by simp
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> handle P C M ?xa h\<^sub>1 (w#vs) ls\<^sub>1 ?pc ics frs sh\<^sub>1"
- using SFAssNonStatic\<^sub>1
- proof(cases ics)
- case No_ics
- then show ?thesis using SFAssNonStatic\<^sub>1 field
- by (auto simp:split_beta handle_def simp del: split_paired_Ex)
- qed(simp_all)
- finally show ?case using pcs by (auto intro!: exI[where x = ?pc] exI[where x="[w]"])
-next
- case (Call\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 a h\<^sub>1 ls\<^sub>1 sh\<^sub>1 es pvs h\<^sub>2 ls\<^sub>2 sh\<^sub>2 Ca fs M' Ts T body D ls\<^sub>2' f h\<^sub>3 ls\<^sub>3 sh\<^sub>3)
- let ?frs\<^sub>0 = "(vs, ls\<^sub>0, C,M,pc,ics)#frs"
- let ?\<sigma>\<^sub>0 = "(None,h\<^sub>0,?frs\<^sub>0,sh\<^sub>0)"
- let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e)"
- let ?\<sigma>\<^sub>1 = "(None,h\<^sub>1,(Addr a#vs, ls\<^sub>1, C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)"
- let ?pc\<^sub>2 = "?pc\<^sub>1 + length(compEs\<^sub>2 es)"
- let ?frs\<^sub>2 = "(rev pvs @ Addr a # vs, ls\<^sub>2, C,M,?pc\<^sub>2,ics)#frs"
- let ?\<sigma>\<^sub>2 = "(None,h\<^sub>2,?frs\<^sub>2,sh\<^sub>2)"
- let ?frs\<^sub>2' = "([], ls\<^sub>2', D,M',0,No_ics) # ?frs\<^sub>2"
- let ?\<sigma>\<^sub>2' = "(None, h\<^sub>2, ?frs\<^sub>2', sh\<^sub>2)"
- have nclinit: "M' \<noteq> clinit" using wf_sees_clinit1[OF wf] visible_method_exists[OF Call\<^sub>1.hyps(6)]
- sees_method_idemp[OF Call\<^sub>1.hyps(6)] by fastforce
- have "P\<^sub>1 \<turnstile>\<^sub>1 \<langle>es,(h\<^sub>1, ls\<^sub>1, sh\<^sub>1)\<rangle> [\<Rightarrow>] \<langle>map Val pvs,(h\<^sub>2, ls\<^sub>2, sh\<^sub>2)\<rangle>" by fact
- hence [simp]: "length es = length pvs" by(auto dest:evals\<^sub>1_preserves_elen)
- have invoke: "P,C,M,?pc\<^sub>2 \<triangleright> Invoke M' (length Ts)"
- using Call\<^sub>1.hyps(7) Call\<^sub>1.prems(1) by clarsimp
- have nsub: "\<not> sub_RI body" by(rule sees_wf\<^sub>1_nsub_RI[OF wf Call\<^sub>1.hyps(6)])
- obtain err where pcs:
- "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>3 ls\<^sub>2 sh\<^sub>3 v xa (e\<bullet>M'(es)) =
- (True, ?frs\<^sub>0, (None, h\<^sub>3, (v#vs, ls\<^sub>2, C,M,?pc\<^sub>2+1,ics)#frs,sh\<^sub>3), err)"
- using Call\<^sub>1.prems(1) by clarsimp
- have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (addr a) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Addr a) xa vs frs
- (I - pcs (compxEs\<^sub>2 es (pc + length (compE\<^sub>2 e)) (Suc (length vs))))" by fact
- have IH_es: "PROP ?Ps es h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (map Val pvs) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C M ?pc\<^sub>1 ics pvs xa
- (map Val pvs) (Addr a#vs) frs (I - pcs(compxE\<^sub>2 e pc (size vs)))" by fact
- have "P \<turnstile> ?\<sigma>\<^sub>0 -jvm\<rightarrow> ?\<sigma>\<^sub>1" using Jcc_pieces_Call1[OF pcs] IH by clarsimp
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> ?\<sigma>\<^sub>2" using IH_es Call\<^sub>1.prems by fastforce
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> ?\<sigma>\<^sub>2'"
- using jvm_Invoke[OF assms(1) invoke _ Call\<^sub>1.hyps(6-8)] Call\<^sub>1.hyps(5) Call\<^sub>1.prems(1) by simp
- finally have 1: "P \<turnstile> ?\<sigma>\<^sub>0 -jvm\<rightarrow> ?\<sigma>\<^sub>2'".
- have "P\<^sub>1 \<turnstile> Ca sees M',NonStatic: Ts\<rightarrow>T = body in D" by fact
- then have M'_in_D: "P\<^sub>1 \<turnstile> D sees M',NonStatic: Ts\<rightarrow>T = body in D"
- by(rule sees_method_idemp)
- have M'_code: "compP\<^sub>2 P\<^sub>1,D,M',0 \<rhd> compE\<^sub>2 body @ [Return]" using beforeM M'_in_D by simp
- have M'_xtab: "compP\<^sub>2 P\<^sub>1,D,M' \<rhd> compxE\<^sub>2 body 0 0/{..<size(compE\<^sub>2 body)},0"
- using M'_in_D by(rule beforexM)
- have IH_body: "PROP ?P body h\<^sub>2 ls\<^sub>2' sh\<^sub>2 f h\<^sub>3 ls\<^sub>3 sh\<^sub>3 (Class D # Ts) D M' 0 No_ics v xa [] ?frs\<^sub>2
- ({..<size(compE\<^sub>2 body)})" by fact
- have cond: "Jcc_cond P\<^sub>1 (Class D # Ts) D M' [] 0 No_ics {..<length (compE\<^sub>2 body)} h\<^sub>2 sh\<^sub>2 body"
- using nsub_RI_Jcc_pieces[OF assms(1) nsub] M'_code M'_xtab by clarsimp
- show ?case (is "?Norm \<and> ?Err")
- proof
- show ?Norm (is "?val \<longrightarrow> ?trans")
- proof
- assume val: ?val
- note 1
- also have "P \<turnstile> ?\<sigma>\<^sub>2' -jvm\<rightarrow> (None,h\<^sub>3,([v],ls\<^sub>3,D,M',size(compE\<^sub>2 body),No_ics)#?frs\<^sub>2,sh\<^sub>3)"
- using val IH_body Call\<^sub>1.prems M'_code cond nsub_RI_Jcc_pieces nsub by auto
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> (None, h\<^sub>3, (v#vs, ls\<^sub>2, C,M,?pc\<^sub>2+1,ics)#frs,sh\<^sub>3)"
- using Call\<^sub>1.hyps(7) M'_code M'_in_D nclinit by(cases T, auto)
- finally show ?trans by(simp add:add.assoc)
- qed
- next
- show ?Err (is "?throw \<longrightarrow> ?err")
- proof
- assume throw: ?throw
- with IH_body obtain pc\<^sub>2 vs' where
- pc\<^sub>2: "0 \<le> pc\<^sub>2 \<and> pc\<^sub>2 < size(compE\<^sub>2 body) \<and>
- \<not> caught P pc\<^sub>2 h\<^sub>3 xa (compxE\<^sub>2 body 0 0)" and
- 2: "P \<turnstile> ?\<sigma>\<^sub>2' -jvm\<rightarrow> handle P D M' xa h\<^sub>3 vs' ls\<^sub>3 pc\<^sub>2 No_ics ?frs\<^sub>2 sh\<^sub>3"
- using Call\<^sub>1.prems M'_code M'_xtab cond nsub_RI_Jcc_pieces nsub
- by (auto simp del:split_paired_Ex)
- have "handle P D M' xa h\<^sub>3 vs' ls\<^sub>3 pc\<^sub>2 No_ics ?frs\<^sub>2 sh\<^sub>3 =
- handle P C M xa h\<^sub>3 (rev pvs @ Addr a # vs) ls\<^sub>2 ?pc\<^sub>2 ics frs sh\<^sub>3"
- using pc\<^sub>2 M'_in_D nclinit by(auto simp add:handle_def)
- then show "?err" using pc\<^sub>2 jvm_trans[OF 1 2]
- by(auto intro!:exI[where x="?pc\<^sub>2"] exI[where x="rev pvs@[Addr a]"])
- qed
- qed
-next
- case (CallParamsThrow\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 w h\<^sub>1 ls\<^sub>1 sh\<^sub>1 es es' h\<^sub>2 ls\<^sub>2 sh\<^sub>2 pvs ex es'' M')
- let ?frs\<^sub>0 = "(vs, ls\<^sub>0, C,M,pc,ics)#frs"
- let ?\<sigma>\<^sub>0 = "(None,h\<^sub>0,(vs, ls\<^sub>0, C,M,pc,ics)#frs,sh\<^sub>0)"
- let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e)"
- let ?\<sigma>\<^sub>1 = "(None,h\<^sub>1,(w # vs, ls\<^sub>1, C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)"
- let ?pc\<^sub>2 = "?pc\<^sub>1 + length(compEs\<^sub>2 es)"
- obtain err where pcs:
- "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v xa (e\<bullet>M'(es)) =
- (True, ?frs\<^sub>0, (None, h\<^sub>2, (v#vs, ls\<^sub>2, C,M,?pc\<^sub>2+1,ics)#frs,sh\<^sub>2), err)"
- using CallParamsThrow\<^sub>1.prems(1) by clarsimp
- have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (Val w) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics w xa vs frs
- (I - pcs (compxEs\<^sub>2 es (pc + length (compE\<^sub>2 e)) (Suc (length vs))))" by fact
- have 1: "P \<turnstile> ?\<sigma>\<^sub>0 -jvm\<rightarrow> ?\<sigma>\<^sub>1" using Jcc_pieces_Call1[OF pcs] IH by clarsimp
- have Isubs: "{?pc\<^sub>1..<?pc\<^sub>2} \<subseteq> I - pcs (compxE\<^sub>2 e pc (length vs))"
- using CallParamsThrow\<^sub>1.prems by clarsimp
- show ?case (is "?N \<and> (?eq \<longrightarrow> ?err)")
- proof
- show ?N by simp
- next
- { assume ?eq
- moreover
- have "PROP ?Ps es h\<^sub>1 ls\<^sub>1 sh\<^sub>1 es' h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C M ?pc\<^sub>1 ics pvs xa es'' (w#vs) frs
- (I - pcs (compxE\<^sub>2 e pc (length vs)))" by fact
- ultimately obtain vs' where "\<exists>pc\<^sub>2.
- (?pc\<^sub>1 \<le> pc\<^sub>2 \<and> pc\<^sub>2 < ?pc\<^sub>1 + size(compEs\<^sub>2 es) \<and>
- \<not> caught P pc\<^sub>2 h\<^sub>2 xa (compxEs\<^sub>2 es ?pc\<^sub>1 (size vs + 1))) \<and>
- P \<turnstile> ?\<sigma>\<^sub>1 -jvm\<rightarrow> handle P C M xa h\<^sub>2 (vs'@w#vs) ls\<^sub>2 pc\<^sub>2 ics frs sh\<^sub>2"
- (is "\<exists>pc\<^sub>2. ?PC pc\<^sub>2 \<and> ?Exec pc\<^sub>2")
- using CallParamsThrow\<^sub>1 Isubs by auto
- then obtain pc\<^sub>2 where pc\<^sub>2: "?PC pc\<^sub>2" and 2: "?Exec pc\<^sub>2" by iprover
- then have "?err" using pc\<^sub>2 jvm_trans[OF 1 2]
- by(auto intro!: exI[where x="pc\<^sub>2"] exI[where x="vs'@[w]"])
- }
- thus "?eq \<longrightarrow> ?err" by simp
- qed
-next
- case (CallNull\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 es pvs h\<^sub>2 ls\<^sub>2 sh\<^sub>2 M')
- have "P\<^sub>1 \<turnstile>\<^sub>1 \<langle>es,(h\<^sub>1, ls\<^sub>1, sh\<^sub>1)\<rangle> [\<Rightarrow>] \<langle>map Val pvs,(h\<^sub>2, ls\<^sub>2, sh\<^sub>2)\<rangle>" by fact
- hence [simp]: "length es = length pvs" by(auto dest:evals\<^sub>1_preserves_elen)
- let ?frs\<^sub>0 = "(vs, ls\<^sub>0, C,M,pc,ics)#frs"
- let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e)"
- let ?pc\<^sub>2 = "?pc\<^sub>1 + length(compEs\<^sub>2 es)"
- let ?xa = "addr_of_sys_xcpt NullPointer"
- obtain err where pcs:
- "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v xa (e\<bullet>M'(es)) =
- (True, ?frs\<^sub>0, (None, h\<^sub>2, (v#vs, ls\<^sub>2, C,M,?pc\<^sub>2+1,ics)#frs,sh\<^sub>2), err)"
- using CallNull\<^sub>1.prems(1) by clarsimp
- have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 null h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics Null xa vs frs
- (I - pcs (compxEs\<^sub>2 es (pc + length (compE\<^sub>2 e)) (Suc (length vs))))" by fact
- have IH_es: "PROP ?Ps es h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (map Val pvs) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C M ?pc\<^sub>1 ics pvs xa
- (map Val pvs) (Null#vs) frs (I - pcs(compxE\<^sub>2 e pc (size vs)))" by fact
- have Isubs: "{pc + length (compE\<^sub>2 e)..<pc + length (compE\<^sub>2 e) + length (compEs\<^sub>2 es)}
- \<subseteq> I - pcs (compxE\<^sub>2 e pc (length vs))" using CallNull\<^sub>1.prems by clarsimp
- have "P \<turnstile> (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0) -jvm\<rightarrow>
- (None,h\<^sub>1,(Null#vs,ls\<^sub>1,C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)"
- using Jcc_pieces_Call1[OF pcs] IH by clarsimp
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> (None,h\<^sub>2,(rev pvs@Null#vs,ls\<^sub>2,C,M,?pc\<^sub>2,ics)#frs,sh\<^sub>2)"
- using CallNull\<^sub>1 IH_es Isubs by auto
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> handle P C M ?xa h\<^sub>2 (rev pvs@Null#vs) ls\<^sub>2 ?pc\<^sub>2 ics frs sh\<^sub>2"
- using CallNull\<^sub>1.prems
- by(auto simp:split_beta handle_def nth_append simp del: split_paired_Ex)
- finally show ?case by (auto intro!: exI[where x = ?pc\<^sub>2] exI[where x="rev pvs@[Null]"])
-next
- case (CallObjThrow\<^sub>1 e h ls sh e' h' ls' sh' M' es)
- obtain err where pcs:
- "Jcc_pieces P\<^sub>1 E C M h vs ls pc ics frs sh I h' ls' sh' v xa (e\<bullet>M'(es)) =
- (True, (vs, ls, C,M,pc,ics)#frs,
- (None, h', (v#vs, ls', C,M,pc+size(compE\<^sub>2 (e\<bullet>M'(es))),ics)#frs,sh'), err)"
- using CallObjThrow\<^sub>1.prems(1) by clarsimp
- obtain a' where throw: "throw e' = Throw a'"
- using eval\<^sub>1_final[OF CallObjThrow\<^sub>1.hyps(1)] by clarsimp
- have IH: "PROP ?P e h ls sh (throw e') h' ls' sh' E C M pc ics v a' vs frs
- (I - pcs (compxEs\<^sub>2 es (pc + length (compE\<^sub>2 e)) (Suc (length vs))))" by fact
- show ?case using IH Jcc_pieces_Call1[OF pcs] throw CallObjThrow\<^sub>1.prems nsub_RI_Jcc_pieces
- by auto
-next
- case (CallNone\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 a h\<^sub>1 ls\<^sub>1 sh\<^sub>1 es pvs h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C' fs M')
- let ?frs\<^sub>0 = "(vs, ls\<^sub>0, C,M,pc,ics)#frs"
- let ?\<sigma>\<^sub>0 = "(None,h\<^sub>0,?frs\<^sub>0,sh\<^sub>0)"
- let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e)"
- let ?\<sigma>\<^sub>1 = "(None,h\<^sub>1,(Addr a#vs, ls\<^sub>1, C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)"
- let ?pc\<^sub>2 = "?pc\<^sub>1 + length(compEs\<^sub>2 es)"
- let ?frs\<^sub>2 = "(rev pvs @ Addr a # vs, ls\<^sub>2, C,M,?pc\<^sub>2,ics)#frs"
- let ?\<sigma>\<^sub>2 = "(None,h\<^sub>2,?frs\<^sub>2,sh\<^sub>2)"
- let ?xa = "addr_of_sys_xcpt NoSuchMethodError"
- have "P\<^sub>1 \<turnstile>\<^sub>1 \<langle>es,(h\<^sub>1, ls\<^sub>1, sh\<^sub>1)\<rangle> [\<Rightarrow>] \<langle>map Val pvs,(h\<^sub>2, ls\<^sub>2, sh\<^sub>2)\<rangle>" by fact
- hence [simp]: "length es = length pvs" by(auto dest:evals\<^sub>1_preserves_elen)
- have aux: "(rev pvs @ Addr a # vs) ! length pvs = Addr a"
- by (metis length_rev nth_append_length)
- have nmeth: "\<not>(\<exists>b Ts T body D. P \<turnstile> C' sees M', b : Ts\<rightarrow>T = body in D)"
- using sees_method_compPD CallNone\<^sub>1.hyps(6) by fastforce
- obtain err where pcs:
- "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v xa (e\<bullet>M'(es)) =
- (True, ?frs\<^sub>0, (None, h\<^sub>2, (v#vs, ls\<^sub>2, C,M,?pc\<^sub>2+1,ics)#frs,sh\<^sub>2), err)"
- using CallNone\<^sub>1.prems(1) by clarsimp
- have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (addr a) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Addr a) xa vs frs
- (I - pcs (compxEs\<^sub>2 es (pc + length (compE\<^sub>2 e)) (Suc (length vs))))" by fact
- have IH_es: "PROP ?Ps es h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (map Val pvs) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C M ?pc\<^sub>1 ics pvs xa
- (map Val pvs) (Addr a#vs) frs (I - pcs(compxE\<^sub>2 e pc (size vs)))" by fact
- have "P \<turnstile> ?\<sigma>\<^sub>0 -jvm\<rightarrow> ?\<sigma>\<^sub>1" using Jcc_pieces_Call1[OF pcs] IH by clarsimp
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> ?\<sigma>\<^sub>2" using IH_es CallNone\<^sub>1.prems by fastforce
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> handle P C M ?xa h\<^sub>2 (rev pvs@Addr a#vs) ls\<^sub>2 ?pc\<^sub>2 ics frs sh\<^sub>2"
- using CallNone\<^sub>1.hyps(5) CallNone\<^sub>1.prems aux nmeth
- by(cases "method P C' M'", cases "find_handler P ?xa h\<^sub>2 frs sh\<^sub>2", auto simp: handle_def)
- finally show ?case using pcs by (auto intro!: exI[where x = ?pc\<^sub>2] exI[where x="rev pvs@[Addr a]"])
-next
- case (CallStatic\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 a h\<^sub>1 ls\<^sub>1 sh\<^sub>1 es pvs h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C' fs M' Ts T body D)
- let ?frs\<^sub>0 = "(vs, ls\<^sub>0, C,M,pc,ics)#frs"
- let ?\<sigma>\<^sub>0 = "(None,h\<^sub>0,?frs\<^sub>0,sh\<^sub>0)"
- let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e)"
- let ?\<sigma>\<^sub>1 = "(None,h\<^sub>1,(Addr a#vs, ls\<^sub>1, C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)"
- let ?pc\<^sub>2 = "?pc\<^sub>1 + length(compEs\<^sub>2 es)"
- let ?frs\<^sub>2 = "(rev pvs @ Addr a # vs, ls\<^sub>2, C,M,?pc\<^sub>2,ics)#frs"
- let ?\<sigma>\<^sub>2 = "(None,h\<^sub>2,?frs\<^sub>2,sh\<^sub>2)"
- let ?xa = "addr_of_sys_xcpt IncompatibleClassChangeError"
- have "P\<^sub>1 \<turnstile>\<^sub>1 \<langle>es,(h\<^sub>1, ls\<^sub>1, sh\<^sub>1)\<rangle> [\<Rightarrow>] \<langle>map Val pvs,(h\<^sub>2, ls\<^sub>2, sh\<^sub>2)\<rangle>" by fact
- hence [simp]: "length es = length pvs" by(auto dest:evals\<^sub>1_preserves_elen)
- have aux: "(rev pvs @ Addr a # vs) ! length pvs = Addr a"
- by (metis length_rev nth_append_length)
- obtain body' where method: "P \<turnstile> C' sees M', Static : Ts\<rightarrow>T = body' in D"
- by (metis CallStatic\<^sub>1.hyps(6) P_def compP\<^sub>2_def sees_method_compP)
- obtain err where pcs:
- "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v xa (e\<bullet>M'(es)) =
- (True, ?frs\<^sub>0, (None, h\<^sub>2, (v#vs, ls\<^sub>2, C,M,?pc\<^sub>2+1,ics)#frs,sh\<^sub>2), err)"
- using CallStatic\<^sub>1.prems(1) by clarsimp
- have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (addr a) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Addr a) xa vs frs
- (I - pcs (compxEs\<^sub>2 es (pc + length (compE\<^sub>2 e)) (Suc (length vs))))" by fact
- have IH_es: "PROP ?Ps es h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (map Val pvs) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C M ?pc\<^sub>1 ics pvs xa
- (map Val pvs) (Addr a#vs) frs (I - pcs(compxE\<^sub>2 e pc (size vs)))" by fact
- have "P \<turnstile> ?\<sigma>\<^sub>0 -jvm\<rightarrow> ?\<sigma>\<^sub>1" using Jcc_pieces_Call1[OF pcs] IH by clarsimp
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> ?\<sigma>\<^sub>2" using IH_es CallStatic\<^sub>1.prems by fastforce
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> handle P C M ?xa h\<^sub>2 (rev pvs@Addr a#vs) ls\<^sub>2 ?pc\<^sub>2 ics frs sh\<^sub>2"
- using CallStatic\<^sub>1.hyps(5) CallStatic\<^sub>1.prems aux method
- by(cases "method P C' M'", cases "find_handler P ?xa h\<^sub>2 frs sh\<^sub>2")
- (auto simp: handle_def; meson frames_of.cases)
- finally show ?case using pcs by (auto intro!: exI[where x = ?pc\<^sub>2] exI[where x="rev pvs@[Addr a]"])
-next
- case (SCallParamsThrow\<^sub>1 es h\<^sub>1 ls\<^sub>1 sh\<^sub>1 es' h\<^sub>2 ls\<^sub>2 sh\<^sub>2 pvs ex es'' C' M')
- show ?case
- proof(cases "M' = clinit \<and> es = []")
- case clinit: True then show ?thesis
- using SCallParamsThrow\<^sub>1.hyps(1,3) evals\<^sub>1_cases(1) by fastforce
- next
- case nclinit: False
- let ?\<sigma>\<^sub>1 = "(None,h\<^sub>1,(vs, ls\<^sub>1, C,M,pc,ics)#frs,sh\<^sub>1)"
- let ?pc\<^sub>2 = "pc + length(compEs\<^sub>2 es)"
- have Isubs: "{pc..<pc + length (compEs\<^sub>2 es)} \<subseteq> I" using SCallParamsThrow\<^sub>1.prems nclinit by clarsimp
- show ?thesis (is "?N \<and> (?eq \<longrightarrow> ?err)")
- proof
- show ?N by simp
- next
- { assume ?eq
- moreover
- have "PROP ?Ps es h\<^sub>1 ls\<^sub>1 sh\<^sub>1 es' h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C M pc ics pvs xa es'' vs frs I" by fact
- ultimately have "\<exists>pc\<^sub>2.
- (pc \<le> pc\<^sub>2 \<and> pc\<^sub>2 < pc + size(compEs\<^sub>2 es) \<and>
- \<not> caught P pc\<^sub>2 h\<^sub>2 xa (compxEs\<^sub>2 es pc (size vs))) \<and>
- (\<exists>vs'. P \<turnstile> ?\<sigma>\<^sub>1 -jvm\<rightarrow> handle P C M xa h\<^sub>2 (vs'@vs) ls\<^sub>2 pc\<^sub>2 ics frs sh\<^sub>2)"
- (is "\<exists>pc\<^sub>2. ?PC pc\<^sub>2 \<and> ?Exec pc\<^sub>2")
- using SCallParamsThrow\<^sub>1 Isubs nclinit by auto
- then obtain pc\<^sub>2 where pc\<^sub>2: "?PC pc\<^sub>2" and 2: "?Exec pc\<^sub>2" by iprover
- then have "?err" using pc\<^sub>2 2 by(auto intro: exI[where x="pc\<^sub>2"])
- }
- thus "?eq \<longrightarrow> ?err" by iprover
- qed
- qed
-next
- case (SCallNone\<^sub>1 es h\<^sub>1 ls\<^sub>1 sh\<^sub>1 pvs h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C' M')
- show ?case
- proof(cases "M' = clinit \<and> es = []")
- case clinit: True then show ?thesis using SCallNone\<^sub>1.hyps(3) SCallNone\<^sub>1.prems by auto
- next
- case nclinit: False
- let ?\<sigma>\<^sub>1 = "(None,h\<^sub>1,(vs, ls\<^sub>1, C,M,pc,ics)#frs,sh\<^sub>1)"
- let ?pc\<^sub>2 = "pc + length(compEs\<^sub>2 es)"
- let ?frs\<^sub>2 = "(rev pvs @ vs, ls\<^sub>2, C,M,?pc\<^sub>2,ics)#frs"
- let ?\<sigma>\<^sub>2 = "(None,h\<^sub>2,?frs\<^sub>2,sh\<^sub>2)"
- let ?xa = "addr_of_sys_xcpt NoSuchMethodError"
- have "P\<^sub>1 \<turnstile>\<^sub>1 \<langle>es,(h\<^sub>1, ls\<^sub>1, sh\<^sub>1)\<rangle> [\<Rightarrow>] \<langle>map Val pvs,(h\<^sub>2, ls\<^sub>2, sh\<^sub>2)\<rangle>" by fact
- hence [simp]: "length es = length pvs" by(auto dest:evals\<^sub>1_preserves_elen)
- have nmeth: "\<not>(\<exists>b Ts T body D. P \<turnstile> C' sees M', b : Ts\<rightarrow>T = body in D)"
- using sees_method_compPD SCallNone\<^sub>1.hyps(3) by fastforce
- have IH_es: "PROP ?Ps es h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (map Val pvs) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C M pc ics pvs xa
- (map Val pvs) vs frs I" by fact
- have "P \<turnstile> ?\<sigma>\<^sub>1 -jvm\<rightarrow> ?\<sigma>\<^sub>2" using IH_es SCallNone\<^sub>1.prems nclinit by auto fastforce+
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> handle P C M ?xa h\<^sub>2 (rev pvs@vs) ls\<^sub>2 ?pc\<^sub>2 ics frs sh\<^sub>2"
- using SCallNone\<^sub>1.prems nmeth nclinit
- by(cases "method P C' M'", cases "find_handler P ?xa h\<^sub>2 frs sh\<^sub>2", auto simp: handle_def)
- finally show ?thesis using nclinit by (auto intro: exI[where x = ?pc\<^sub>2])
- qed
-next
- case (SCallNonStatic\<^sub>1 es h\<^sub>1 ls\<^sub>1 sh\<^sub>1 pvs h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C' M' Ts T body D)
- show ?case
- proof(cases "M' = clinit \<and> es = []")
- case clinit: True then show ?thesis
- using SCallNonStatic\<^sub>1.hyps(3) SCallNonStatic\<^sub>1.prems sees_method_fun by fastforce
- next
- case nclinit: False
- let ?\<sigma>\<^sub>1 = "(None,h\<^sub>1,(vs, ls\<^sub>1, C,M,pc,ics)#frs,sh\<^sub>1)"
- let ?pc\<^sub>2 = "pc + length(compEs\<^sub>2 es)"
- let ?frs\<^sub>2 = "(rev pvs @ vs, ls\<^sub>2, C,M,?pc\<^sub>2,ics)#frs"
- let ?\<sigma>\<^sub>2 = "(None,h\<^sub>2,?frs\<^sub>2,sh\<^sub>2)"
- let ?xa = "addr_of_sys_xcpt IncompatibleClassChangeError"
- have "P\<^sub>1 \<turnstile>\<^sub>1 \<langle>es,(h\<^sub>1, ls\<^sub>1, sh\<^sub>1)\<rangle> [\<Rightarrow>] \<langle>map Val pvs,(h\<^sub>2, ls\<^sub>2, sh\<^sub>2)\<rangle>" by fact
- hence [simp]: "length es = length pvs" by(auto dest:evals\<^sub>1_preserves_elen)
- obtain body' where method: "P \<turnstile> C' sees M', NonStatic : Ts\<rightarrow>T = body' in D"
- by (metis SCallNonStatic\<^sub>1.hyps(3) P_def compP\<^sub>2_def sees_method_compP)
- have IH_es: "PROP ?Ps es h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (map Val pvs) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C M pc ics pvs xa
- (map Val pvs) vs frs I" by fact
- have "P \<turnstile> ?\<sigma>\<^sub>1 -jvm\<rightarrow> ?\<sigma>\<^sub>2" using IH_es SCallNonStatic\<^sub>1.prems nclinit by auto fastforce+
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> handle P C M ?xa h\<^sub>2 (rev pvs@vs) ls\<^sub>2 ?pc\<^sub>2 ics frs sh\<^sub>2"
- using SCallNonStatic\<^sub>1.prems method nclinit
- by(cases "method P C' M'", cases "find_handler P ?xa h\<^sub>2 frs sh\<^sub>2")
- (auto simp: handle_def; meson frames_of.cases)
- finally show ?thesis using nclinit by (auto intro: exI[where x = ?pc\<^sub>2])
- qed
-next
- case (SCallInitThrow\<^sub>1 es h\<^sub>0 ls\<^sub>0 sh\<^sub>0 pvs h\<^sub>1 ls\<^sub>1 sh\<^sub>1 C' M' Ts T body D a h\<^sub>2 ls\<^sub>2 sh\<^sub>2)
- show ?case
- proof(cases "M' = clinit \<and> es = []")
- case clinit: True then show ?thesis using SCallInitThrow\<^sub>1 by simp
- next
- case nclinit: False
- let ?\<sigma>\<^sub>0 = "(None,h\<^sub>0,(vs, ls\<^sub>0, C,M,pc,ics)#frs,sh\<^sub>0)"
- let ?pc\<^sub>1 = "pc + length(compEs\<^sub>2 es)"
- let ?frs\<^sub>1 = "(rev pvs @ vs, ls\<^sub>1, C,M,?pc\<^sub>1,ics)#frs"
- let ?\<sigma>\<^sub>1 = "(None,h\<^sub>1,?frs\<^sub>1,sh\<^sub>1)"
- let ?frs\<^sub>1' = "(rev pvs@vs,ls\<^sub>1,C,M,?pc\<^sub>1,Calling D [])#frs"
- let ?\<sigma>\<^sub>1' = "(None,h\<^sub>1,?frs\<^sub>1',sh\<^sub>1)"
- let ?frs\<^sub>2 = "(rev pvs@vs,ls\<^sub>1,C,M,?pc\<^sub>1,Called [])#frs"
- let ?\<sigma>\<^sub>2 = "(None,h\<^sub>2,?frs\<^sub>2,sh\<^sub>2)"
- have ls: "ls\<^sub>1 = ls\<^sub>2" by(rule init\<^sub>1_same_loc[OF SCallInitThrow\<^sub>1.hyps(6)])
- have method: "\<exists>m'. P \<turnstile> C' sees M',Static:Ts\<rightarrow>T = m' in D" using SCallInitThrow\<^sub>1.hyps(3)
- by (metis P_def compP\<^sub>2_def sees_method_compP)
- obtain a' where throw: "throw a = Throw a'" using eval\<^sub>1_final[OF SCallInitThrow\<^sub>1.hyps(6)] by clarsimp
- have "Ex (WTrt2\<^sub>1 P\<^sub>1 E h\<^sub>1 sh\<^sub>1 (INIT D ([D],False) \<leftarrow> unit))"
- using sees_method_is_class'[OF SCallInitThrow\<^sub>1.hyps(3)] by auto
- then obtain err' where pcs':
- "Jcc_pieces P\<^sub>1 E C M h\<^sub>1 (rev pvs@vs) ls\<^sub>1 ?pc\<^sub>1 ics frs sh\<^sub>1 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v xa (INIT D ([D],False) \<leftarrow> unit)
- = (True, ?frs\<^sub>1', (None,h\<^sub>2,?frs\<^sub>2,sh\<^sub>2), err')"
- using SCallInitThrow\<^sub>1.prems(1) nclinit by auto
- have IHI: "PROP ?P (INIT D ([D],False) \<leftarrow> unit) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (throw a)
- h\<^sub>2 ls\<^sub>2 sh\<^sub>2 E C M ?pc\<^sub>1 ics v a' (rev pvs@vs) frs I" by fact
- have IH_es: "PROP ?Ps es h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (map Val pvs) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 C M pc ics pvs xa
- (map Val pvs) vs frs I" by fact
- have "P \<turnstile> ?\<sigma>\<^sub>0 -jvm\<rightarrow> ?\<sigma>\<^sub>1" using IH_es SCallInitThrow\<^sub>1.prems nclinit by auto fastforce+
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> ?\<sigma>\<^sub>1'"
- proof(cases "sh\<^sub>1 D")
- case None then show ?thesis using SCallInitThrow\<^sub>1.hyps(1,3-6) SCallInitThrow\<^sub>1.prems method
- by(cases ics) auto
- next
- case (Some a)
- then obtain sfs i where "a = (sfs,i)" by(cases a)
- then show ?thesis using SCallInitThrow\<^sub>1.hyps(1,3-6) SCallInitThrow\<^sub>1.prems method Some
- by(cases ics; case_tac i, auto)
- qed
- also obtain vs' where "P \<turnstile> \<dots> -jvm\<rightarrow> handle P C M a' h\<^sub>2 (vs'@rev pvs@vs) ls\<^sub>1 ?pc\<^sub>1 ics frs sh\<^sub>2"
- using IHI pcs' throw by auto
- finally show ?thesis using nclinit throw ls
- by(auto intro!: exI[where x="?pc\<^sub>1"] exI[where x="vs'@rev pvs"])
- qed
-next
- case (SCallInit\<^sub>1 es h\<^sub>0 ls\<^sub>0 sh\<^sub>0 pvs h\<^sub>1 ls\<^sub>1 sh\<^sub>1 C' M' Ts T body D v' h\<^sub>2 ls\<^sub>2 sh\<^sub>2 ls\<^sub>2' e' h\<^sub>3 ls\<^sub>3 sh\<^sub>3)
- show ?case
- proof(cases "M' = clinit \<and> es = []")
- case clinit: True then show ?thesis using SCallInit\<^sub>1 by simp
- next
- case nclinit: False
- let ?\<sigma>\<^sub>0 = "(None,h\<^sub>0,(vs, ls\<^sub>0, C,M,pc,ics)#frs,sh\<^sub>0)"
- let ?pc\<^sub>1 = "pc + length(compEs\<^sub>2 es)"
- let ?frs\<^sub>1 = "(rev pvs @ vs, ls\<^sub>1, C,M,?pc\<^sub>1,ics)#frs"
- let ?\<sigma>\<^sub>1 = "(None,h\<^sub>1,?frs\<^sub>1,sh\<^sub>1)"
- let ?frs\<^sub>1' = "(rev pvs@vs,ls\<^sub>1,C,M,?pc\<^sub>1,Calling D [])#frs"
- let ?\<sigma>\<^sub>1' = "(None,h\<^sub>1,?frs\<^sub>1',sh\<^sub>1)"
- let ?frs\<^sub>2 = "(rev pvs@vs,ls\<^sub>1,C,M,?pc\<^sub>1,Called [])#frs"
- let ?\<sigma>\<^sub>2 = "(None,h\<^sub>2,?frs\<^sub>2,sh\<^sub>2)"
- let ?frs\<^sub>2' = "([], ls\<^sub>2', D,M',0,No_ics) # ?frs\<^sub>1"
- let ?\<sigma>\<^sub>2' = "(None, h\<^sub>2, ?frs\<^sub>2', sh\<^sub>2)"
- have nclinit': "M' \<noteq> clinit" by fact
- have ics: "ics = No_ics" using SCallInit\<^sub>1.hyps(5) SCallInit\<^sub>1.prems by simp
- have "P\<^sub>1 \<turnstile>\<^sub>1 \<langle>es,(h\<^sub>0, ls\<^sub>0, sh\<^sub>0)\<rangle> [\<Rightarrow>] \<langle>map Val pvs,(h\<^sub>1, ls\<^sub>1, sh\<^sub>1)\<rangle>" by fact
- hence [simp]: "length es = length pvs" by(auto dest:evals\<^sub>1_preserves_elen)
- have invoke: "P,C,M,?pc\<^sub>1 \<triangleright> Invokestatic C' M' (length Ts)"
- using SCallInit\<^sub>1.hyps(8) SCallInit\<^sub>1.prems nclinit by(auto simp: add.assoc)
- have nsub: "\<not> sub_RI body" by(rule sees_wf\<^sub>1_nsub_RI[OF wf SCallInit\<^sub>1.hyps(3)])
- have ls: "ls\<^sub>1 = ls\<^sub>2" by(rule init\<^sub>1_same_loc[OF SCallInit\<^sub>1.hyps(6)])
- obtain sfs i where sh\<^sub>2: "sh\<^sub>2 D = Some(sfs,i)"
- using init\<^sub>1_Val_PD[OF SCallInit\<^sub>1.hyps(6)] by clarsimp
- have method: "\<exists>m'. P \<turnstile> C' sees M',Static:Ts\<rightarrow>T = m' in D" using SCallInit\<^sub>1.hyps(3)
- by (metis P_def compP\<^sub>2_def sees_method_compP)
- have "Ex (WTrt2\<^sub>1 P\<^sub>1 E h\<^sub>1 sh\<^sub>1 (INIT D ([D],False) \<leftarrow> unit))"
- using sees_method_is_class'[OF SCallInit\<^sub>1.hyps(3)] by auto
- then obtain err' where pcs':
- "Jcc_pieces P\<^sub>1 E C M h\<^sub>1 (rev pvs@vs) ls\<^sub>1 ?pc\<^sub>1 ics frs sh\<^sub>1 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v' xa (INIT D ([D],False) \<leftarrow> unit)
- = (True, ?frs\<^sub>1', (None,h\<^sub>2,?frs\<^sub>2,sh\<^sub>2), err')"
- using SCallInit\<^sub>1.prems(1) nclinit by auto
- have IHI: "PROP ?P (INIT D ([D],False) \<leftarrow> unit) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (Val v')
- h\<^sub>2 ls\<^sub>2 sh\<^sub>2 E C M ?pc\<^sub>1 ics v' xa (rev pvs@vs) frs I" by fact
- have IH_es: "PROP ?Ps es h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (map Val pvs) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 C M pc ics pvs xa
- (map Val pvs) vs frs I" by fact
- have "P \<turnstile> ?\<sigma>\<^sub>0 -jvm\<rightarrow> ?\<sigma>\<^sub>1" using IH_es SCallInit\<^sub>1.prems nclinit by auto fastforce+
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> ?\<sigma>\<^sub>1'"
- proof(cases "sh\<^sub>1 D")
- case None then show ?thesis using SCallInit\<^sub>1.hyps(1,3-6,8-10) SCallInit\<^sub>1.prems method
- by(cases ics) auto
- next
- case (Some a)
- then obtain sfs i where "a = (sfs,i)" by(cases a)
- then show ?thesis using SCallInit\<^sub>1.hyps(1,3-6,8-10) SCallInit\<^sub>1.prems method Some
- by(cases ics; case_tac i, auto)
- qed
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> ?\<sigma>\<^sub>2" using IHI pcs' by auto
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> ?\<sigma>\<^sub>2'"
- using jvm_Invokestatic_Called[OF assms(1) invoke _ SCallInit\<^sub>1.hyps(3,8,9)] sh\<^sub>2 ics by auto
- finally have 1: "P \<turnstile> ?\<sigma>\<^sub>0 -jvm\<rightarrow> ?\<sigma>\<^sub>2'".
- have "P\<^sub>1 \<turnstile> C' sees M',Static: Ts\<rightarrow>T = body in D" by fact
- then have M'_in_D: "P\<^sub>1 \<turnstile> D sees M',Static: Ts\<rightarrow>T = body in D"
- by(rule sees_method_idemp)
- have M'_code: "compP\<^sub>2 P\<^sub>1,D,M',0 \<rhd> compE\<^sub>2 body @ [Return]" using beforeM M'_in_D by simp
- have M'_xtab: "compP\<^sub>2 P\<^sub>1,D,M' \<rhd> compxE\<^sub>2 body 0 0/{..<size(compE\<^sub>2 body)},0"
- using M'_in_D by(rule beforexM)
- have IH_body: "PROP ?P body h\<^sub>2 ls\<^sub>2' sh\<^sub>2 e' h\<^sub>3 ls\<^sub>3 sh\<^sub>3 (Class D # Ts) D M' 0 No_ics v xa [] ?frs\<^sub>1
- ({..<size(compE\<^sub>2 body)})" by fact
- have cond: "Jcc_cond P\<^sub>1 (Class D # Ts) D M' [] 0 No_ics {..<length (compE\<^sub>2 body)} h\<^sub>2 sh\<^sub>2 body"
- using nsub_RI_Jcc_pieces[OF assms(1) nsub] M'_code M'_xtab by clarsimp
- show ?thesis (is "?Norm \<and> ?Err")
- proof
- show ?Norm (is "?val \<longrightarrow> ?trans")
- proof
- assume val: ?val
- note 1
- also have "P \<turnstile> ?\<sigma>\<^sub>2' -jvm\<rightarrow> (None,h\<^sub>3,([v],ls\<^sub>3,D,M',size(compE\<^sub>2 body),No_ics)#?frs\<^sub>1,sh\<^sub>3)"
- using val IH_body SCallInit\<^sub>1.prems M'_code cond nsub_RI_Jcc_pieces nsub by auto
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> (None, h\<^sub>3, (v#vs, ls\<^sub>2, C,M,?pc\<^sub>1+1,ics)#frs,sh\<^sub>3)"
- using SCallInit\<^sub>1.hyps(8) M'_code M'_in_D ls nclinit' by(cases T, auto)
- finally show ?trans using nclinit by(auto simp:add.assoc)
- qed
- next
- show ?Err (is "?throw \<longrightarrow> ?err")
- proof
- assume throw: ?throw
- with IH_body obtain pc\<^sub>2 vs' where
- pc\<^sub>2: "0 \<le> pc\<^sub>2 \<and> pc\<^sub>2 < size(compE\<^sub>2 body) \<and>
- \<not> caught P pc\<^sub>2 h\<^sub>3 xa (compxE\<^sub>2 body 0 0)" and
- 2: "P \<turnstile> ?\<sigma>\<^sub>2' -jvm\<rightarrow> handle P D M' xa h\<^sub>3 vs' ls\<^sub>3 pc\<^sub>2 No_ics ?frs\<^sub>1 sh\<^sub>3"
- using SCallInit\<^sub>1.prems M'_code M'_xtab cond nsub_RI_Jcc_pieces nsub
- by (auto simp del:split_paired_Ex)
- have "handle P D M' xa h\<^sub>3 vs' ls\<^sub>3 pc\<^sub>2 No_ics ?frs\<^sub>1 sh\<^sub>3 =
- handle P C M xa h\<^sub>3 (rev pvs @ vs) ls\<^sub>2 ?pc\<^sub>1 ics frs sh\<^sub>3"
- using pc\<^sub>2 M'_in_D ls nclinit' by(auto simp add:handle_def)
- then show "?err" using pc\<^sub>2 jvm_trans[OF 1 2] nclinit
- by(auto intro!:exI[where x="?pc\<^sub>1"] exI[where x="rev pvs"])
- qed
- qed
- qed
-next
- case (SCall\<^sub>1 es h\<^sub>1 ls\<^sub>1 sh\<^sub>1 pvs h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C' M' Ts T body D sfs ls\<^sub>2' e' h\<^sub>3 ls\<^sub>3 sh\<^sub>3)
- show ?case
- proof(cases "M' = clinit \<and> es = []")
- case clinit: True
- then have s1: "pvs = []" "h\<^sub>1 = h\<^sub>2" "ls\<^sub>1 = ls\<^sub>2" "sh\<^sub>1 = sh\<^sub>2"
- using SCall\<^sub>1.hyps(1) evals\<^sub>1_cases(1) by blast+
- then have ls\<^sub>2': "ls\<^sub>2' = replicate (max_vars body) undefined" using SCall\<^sub>1.hyps(6) clinit by simp
- let ?frs = "create_init_frame P C' # (vs, ls\<^sub>1, C,M,pc,ics)#frs"
- let ?\<sigma>\<^sub>1 = "(None,h\<^sub>1,?frs,sh\<^sub>1)"
- have method: "P\<^sub>1 \<turnstile> C' sees clinit,Static: []\<rightarrow>Void = body in C'"
- using SCall\<^sub>1.hyps(3) clinit s1(1) wf_sees_clinit[OF wf]
- by (metis is_class_def option.collapse sees_method_fun sees_method_is_class)
- then have M_code: "compP\<^sub>2 P\<^sub>1,C',clinit,0 \<rhd> compE\<^sub>2 body @ [Return]" by(rule beforeM)
- have pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>1 vs ls\<^sub>1 pc ics frs sh\<^sub>1 I h\<^sub>3 ls\<^sub>2 sh\<^sub>3 v xa (C'\<bullet>\<^sub>sclinit([]))
- = (True, ?frs, (None, h\<^sub>3, tl ?frs, sh\<^sub>3(C'\<mapsto>(fst(the(sh\<^sub>3 C')),Done))),
- P \<turnstile> (None, h\<^sub>1, ?frs, sh\<^sub>1) -jvm\<rightarrow>
- (case ics of
- Called Cs \<Rightarrow> (None, h\<^sub>3, (vs, ls\<^sub>1, C, M, pc, Throwing Cs xa) # frs, sh\<^sub>3(C' \<mapsto> (fst (the (sh\<^sub>3 C')), Error)))))"
- using Jcc_pieces_clinit[OF assms(1),of E C M vs pc ics I h\<^sub>1 sh\<^sub>1 C' ls\<^sub>1 frs h\<^sub>3 ls\<^sub>2 sh\<^sub>3 v xa]
- SCall\<^sub>1.prems(1) clinit s1(1) by clarsimp
- have IH_body: "PROP ?P body h\<^sub>2 ls\<^sub>2' sh\<^sub>2 e' h\<^sub>3 ls\<^sub>3 sh\<^sub>3 [] C' clinit 0 No_ics v xa [] (tl ?frs)
- ({..<size(compE\<^sub>2 body)})" by fact
- show ?thesis (is "?Norm \<and> ?Err")
- proof
- show ?Norm (is "?val \<longrightarrow> ?trans")
- proof
- assume val: ?val
- then have "P \<turnstile> ?\<sigma>\<^sub>1
- -jvm\<rightarrow> (None, h\<^sub>3, ([v], ls\<^sub>3, C', clinit, size(compE\<^sub>2 body), No_ics) # tl ?frs,sh\<^sub>3)"
- using IH_body Jcc_pieces_SCall_clinit_body[OF assms(1) wf pcs method] s1 ls\<^sub>2' by clarsimp
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> (None, h\<^sub>3, tl ?frs, sh\<^sub>3(C'\<mapsto>(fst(the(sh\<^sub>3 C')),Done)))"
- using jvm_Return_Init[OF M_code] by simp
- finally show ?trans using pcs s1 clinit by simp
- qed
- next
- show ?Err (is "?throw \<longrightarrow> ?err")
- proof
- assume throw: ?throw
- with IH_body obtain pc\<^sub>2 vs2 where
- pc\<^sub>2: "0 \<le> pc\<^sub>2 \<and> pc\<^sub>2 < size(compE\<^sub>2 body) \<and>
- \<not> caught P pc\<^sub>2 h\<^sub>3 xa (compxE\<^sub>2 body 0 0)" and
- 2: "P \<turnstile> ?\<sigma>\<^sub>1 -jvm\<rightarrow> handle P C' clinit xa h\<^sub>3 vs2 ls\<^sub>3 pc\<^sub>2 No_ics (tl ?frs) sh\<^sub>3"
- using SCall\<^sub>1.prems Jcc_pieces_SCall_clinit_body[OF assms(1) wf pcs method] s1 ls\<^sub>2' by clarsimp
- show ?err using SCall\<^sub>1.prems(1) clinit
- proof(cases ics)
- case (Called Cs)
- note 2
- also have "handle P C' clinit xa h\<^sub>3 vs2 ls\<^sub>3 pc\<^sub>2 No_ics (tl ?frs) sh\<^sub>3
- = (None, h\<^sub>3, (vs, ls\<^sub>1, C, M, pc, Throwing (C'#Cs) xa) # frs, sh\<^sub>3)"
- using Called pc\<^sub>2 method by(simp add: handle_def)
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> (None, h\<^sub>3, (vs, ls\<^sub>1, C, M, pc, Throwing Cs xa) # frs,
- sh\<^sub>3(C' \<mapsto> (fst (the (sh\<^sub>3 C')), Error)))" using Called jvm_Throwing by simp
- finally show ?thesis using pcs clinit Called by(clarsimp intro!: exI[where x="[]"])
- qed(auto)
- qed
- qed
- next
- case nclinit: False
- let ?\<sigma>\<^sub>1 = "(None,h\<^sub>1,(vs, ls\<^sub>1, C,M,pc,ics)#frs,sh\<^sub>1)"
- let ?pc\<^sub>2 = "pc + length(compEs\<^sub>2 es)"
- let ?frs\<^sub>2 = "(rev pvs @ vs, ls\<^sub>2, C,M,?pc\<^sub>2,ics)#frs"
- let ?\<sigma>\<^sub>2 = "(None,h\<^sub>2,?frs\<^sub>2,sh\<^sub>2)"
- let ?frs\<^sub>2' = "([], ls\<^sub>2', D,M',0,No_ics) # ?frs\<^sub>2"
- let ?\<sigma>\<^sub>2' = "(None, h\<^sub>2, ?frs\<^sub>2', sh\<^sub>2)"
- have nclinit': "M' \<noteq> clinit"
- using wf_sees_clinit1[OF wf] visible_method_exists[OF SCall\<^sub>1.hyps(3)]
- sees_method_idemp[OF SCall\<^sub>1.hyps(3)] nclinit SCall\<^sub>1.hyps(5)
- evals\<^sub>1_preserves_elen[OF SCall\<^sub>1.hyps(1)] by fastforce
- have "P\<^sub>1 \<turnstile>\<^sub>1 \<langle>es,(h\<^sub>1, ls\<^sub>1, sh\<^sub>1)\<rangle> [\<Rightarrow>] \<langle>map Val pvs,(h\<^sub>2, ls\<^sub>2, sh\<^sub>2)\<rangle>" by fact
- hence [simp]: "length es = length pvs" by(auto dest:evals\<^sub>1_preserves_elen)
- have invoke: "P,C,M,?pc\<^sub>2 \<triangleright> Invokestatic C' M' (length Ts)"
- using SCall\<^sub>1.hyps(5) SCall\<^sub>1.prems nclinit by(auto simp: add.assoc)
- have nsub: "\<not> sub_RI body" by(rule sees_wf\<^sub>1_nsub_RI[OF wf SCall\<^sub>1.hyps(3)])
- have IH_es: "PROP ?Ps es h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (map Val pvs) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C M pc ics pvs xa
- (map Val pvs) vs frs I" by fact
- have "P \<turnstile> ?\<sigma>\<^sub>1 -jvm\<rightarrow> ?\<sigma>\<^sub>2" using IH_es SCall\<^sub>1.prems nclinit by auto fastforce+
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> ?\<sigma>\<^sub>2'" using jvm_Invokestatic[OF assms(1) invoke _ SCall\<^sub>1.hyps(3,5,6)]
- SCall\<^sub>1.hyps(4) SCall\<^sub>1.prems nclinit by auto
- finally have 1: "P \<turnstile> ?\<sigma>\<^sub>1 -jvm\<rightarrow> ?\<sigma>\<^sub>2'".
- have "P\<^sub>1 \<turnstile> C' sees M',Static: Ts\<rightarrow>T = body in D" by fact
- then have M'_in_D: "P\<^sub>1 \<turnstile> D sees M',Static: Ts\<rightarrow>T = body in D"
- by(rule sees_method_idemp)
- have M'_code: "compP\<^sub>2 P\<^sub>1,D,M',0 \<rhd> compE\<^sub>2 body @ [Return]" using beforeM M'_in_D by simp
- have M'_xtab: "compP\<^sub>2 P\<^sub>1,D,M' \<rhd> compxE\<^sub>2 body 0 0/{..<size(compE\<^sub>2 body)},0"
- using M'_in_D by(rule beforexM)
- have IH_body: "PROP ?P body h\<^sub>2 ls\<^sub>2' sh\<^sub>2 e' h\<^sub>3 ls\<^sub>3 sh\<^sub>3 (Class D # Ts) D M' 0 No_ics v xa [] ?frs\<^sub>2
- ({..<size(compE\<^sub>2 body)})" by fact
- have cond: "Jcc_cond P\<^sub>1 (Class D # Ts) D M' [] 0 No_ics {..<length (compE\<^sub>2 body)} h\<^sub>2 sh\<^sub>2 body"
- using nsub_RI_Jcc_pieces[OF assms(1) nsub] M'_code M'_xtab by clarsimp
- show ?thesis (is "?Norm \<and> ?Err")
- proof
- show ?Norm (is "?val \<longrightarrow> ?trans")
- proof
- assume val: ?val
- note 1
- also have "P \<turnstile> ?\<sigma>\<^sub>2' -jvm\<rightarrow> (None,h\<^sub>3,([v],ls\<^sub>3,D,M',size(compE\<^sub>2 body),No_ics)#?frs\<^sub>2,sh\<^sub>3)"
- using val IH_body SCall\<^sub>1.prems M'_code cond nsub_RI_Jcc_pieces nsub by auto
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> (None, h\<^sub>3, (v#vs, ls\<^sub>2, C,M,?pc\<^sub>2+1,ics)#frs,sh\<^sub>3)"
- using SCall\<^sub>1.hyps(5) M'_code M'_in_D nclinit' by(cases T, auto)
- finally show ?trans using nclinit by(auto simp:add.assoc)
- qed
- next
- show ?Err (is "?throw \<longrightarrow> ?err")
- proof
- assume throw: ?throw
- with IH_body obtain pc\<^sub>2 vs' where
- pc\<^sub>2: "0 \<le> pc\<^sub>2 \<and> pc\<^sub>2 < size(compE\<^sub>2 body) \<and>
- \<not> caught P pc\<^sub>2 h\<^sub>3 xa (compxE\<^sub>2 body 0 0)" and
- 2: "P \<turnstile> ?\<sigma>\<^sub>2' -jvm\<rightarrow> handle P D M' xa h\<^sub>3 vs' ls\<^sub>3 pc\<^sub>2 No_ics ?frs\<^sub>2 sh\<^sub>3"
- using SCall\<^sub>1.prems M'_code M'_xtab cond nsub_RI_Jcc_pieces nsub
- by (auto simp del:split_paired_Ex)
- have "handle P D M' xa h\<^sub>3 vs' ls\<^sub>3 pc\<^sub>2 No_ics ?frs\<^sub>2 sh\<^sub>3 =
- handle P C M xa h\<^sub>3 (rev pvs @ vs) ls\<^sub>2 ?pc\<^sub>2 ics frs sh\<^sub>3"
- using pc\<^sub>2 M'_in_D nclinit' by(auto simp add:handle_def)
- then show "?err" using pc\<^sub>2 jvm_trans[OF 1 2] nclinit by(auto intro:exI[where x="?pc\<^sub>2"])
- qed
- qed
- qed
-next
- case Block\<^sub>1 then show ?case using nsub_RI_Jcc_pieces by auto
-next
- case (Seq\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 w h\<^sub>1 ls\<^sub>1 sh\<^sub>1 e\<^sub>2 e\<^sub>2' h\<^sub>2 ls\<^sub>2 sh\<^sub>2)
- let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e\<^sub>1)"
- let ?\<sigma>\<^sub>0 = "(None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0)"
- let ?\<sigma>\<^sub>1 = "(None,h\<^sub>1,(vs,ls\<^sub>1,C,M,?pc\<^sub>1+1,ics)#frs,sh\<^sub>1)"
- let ?I = "I - pcs (compxE\<^sub>2 e\<^sub>2 (Suc ?pc\<^sub>1) (length vs))"
- have Isub: "{pc..<pc + length (compE\<^sub>2 e\<^sub>1)} \<subseteq> ?I" using Seq\<^sub>1.prems by clarsimp
- have IH: "PROP ?P e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (Val w) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics w xa vs frs ?I" by fact
- have "P \<turnstile> ?\<sigma>\<^sub>0 -jvm\<rightarrow> (None,h\<^sub>1,(w#vs,ls\<^sub>1,C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)"
- using Seq\<^sub>1.prems nsub_RI_Jcc_pieces IH Isub by auto
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> ?\<sigma>\<^sub>1" using Seq\<^sub>1 by auto
- finally have eval\<^sub>1: "P \<turnstile> ?\<sigma>\<^sub>0 -jvm\<rightarrow> ?\<sigma>\<^sub>1".
- let ?pc\<^sub>2 = "?pc\<^sub>1 + 1 + length(compE\<^sub>2 e\<^sub>2)"
- let ?I' = "I - pcs(compxE\<^sub>2 e\<^sub>1 pc (size vs))"
- have IH\<^sub>2: "PROP ?P e\<^sub>2 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 e\<^sub>2' h\<^sub>2 ls\<^sub>2 sh\<^sub>2 E C M (?pc\<^sub>1+1) ics v xa vs frs
- ?I'" by fact
- have Isub2: "{Suc (pc + length (compE\<^sub>2 e\<^sub>1))..<Suc (pc + length (compE\<^sub>2 e\<^sub>1) + length (compE\<^sub>2 e\<^sub>2))}
- \<subseteq> ?I'" using Seq\<^sub>1.prems by clarsimp
- show ?case (is "?Norm \<and> ?Err")
- proof
- show ?Norm (is "?val \<longrightarrow> ?trans")
- proof
- assume val: ?val
- note eval\<^sub>1
- also have "P \<turnstile> ?\<sigma>\<^sub>1 -jvm\<rightarrow> (None,h\<^sub>2,(v#vs,ls\<^sub>2,C,M,?pc\<^sub>2,ics)#frs,sh\<^sub>2)"
- using val Seq\<^sub>1.prems nsub_RI_Jcc_pieces IH\<^sub>2 Isub2 by auto
- finally show ?trans by(simp add:add.assoc)
- qed
- next
- show ?Err (is "?throw \<longrightarrow> ?err")
- proof
- assume throw: ?throw
- then obtain pc\<^sub>2 vs' where
- pc\<^sub>2: "?pc\<^sub>1+1 \<le> pc\<^sub>2 \<and> pc\<^sub>2 < ?pc\<^sub>2 \<and>
- \<not> caught P pc\<^sub>2 h\<^sub>2 xa (compxE\<^sub>2 e\<^sub>2 (?pc\<^sub>1+1) (size vs))" and
- eval\<^sub>2: "P \<turnstile> ?\<sigma>\<^sub>1 -jvm\<rightarrow> handle P C M xa h\<^sub>2 (vs'@vs) ls\<^sub>2 pc\<^sub>2 ics frs sh\<^sub>2"
- using IH\<^sub>2 Seq\<^sub>1.prems nsub_RI_Jcc_pieces Isub2 by auto
- show "?err" using pc\<^sub>2 jvm_trans[OF eval\<^sub>1 eval\<^sub>2] by(auto intro: exI[where x=pc\<^sub>2])
- qed
- qed
-next
- case (SeqThrow\<^sub>1 e\<^sub>0 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 e h\<^sub>1 ls\<^sub>1 sh\<^sub>1 e\<^sub>1)
- let ?I = "I - pcs (compxE\<^sub>2 e\<^sub>1 (Suc (pc + length (compE\<^sub>2 e\<^sub>0))) (length vs))"
- obtain a' where throw: "throw e = Throw a'" using eval\<^sub>1_final[OF SeqThrow\<^sub>1.hyps(1)] by clarsimp
- have Isub: "{pc..<pc + length (compE\<^sub>2 e\<^sub>0)} \<subseteq> ?I" using SeqThrow\<^sub>1.prems by clarsimp
- have "PROP ?P e\<^sub>0 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (throw e) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics v a' vs frs ?I" by fact
- then show ?case using SeqThrow\<^sub>1.prems throw nsub_RI_Jcc_pieces Isub by auto
-next
- case (CondT\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 e\<^sub>1 e' h\<^sub>2 ls\<^sub>2 sh\<^sub>2 e\<^sub>2)
- let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e)"
- let ?\<sigma>\<^sub>0 = "(None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0)"
- let ?\<sigma>\<^sub>1 = "(None,h\<^sub>1,(vs,ls\<^sub>1,C,M,?pc\<^sub>1+1,ics)#frs,sh\<^sub>1)"
- let ?d = "size vs"
- let ?xt\<^sub>1 = "compxE\<^sub>2 e\<^sub>1 (pc+size(compE\<^sub>2 e)+1) ?d"
- let ?xt\<^sub>2 = "compxE\<^sub>2 e\<^sub>2 (pc+size(compE\<^sub>2 e)+size(compE\<^sub>2 e\<^sub>1)+2) ?d"
- let ?I = "I - (pcs ?xt\<^sub>1 \<union> pcs ?xt\<^sub>2)"
- have Isub: "{pc..<pc + length (compE\<^sub>2 e)} \<subseteq> ?I" using CondT\<^sub>1.prems by clarsimp
- have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 true h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Bool True) xa vs frs ?I" by fact
- have "P \<turnstile> ?\<sigma>\<^sub>0 -jvm\<rightarrow> (None,h\<^sub>1,(Bool(True)#vs,ls\<^sub>1,C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)"
- using CondT\<^sub>1.prems nsub_RI_Jcc_pieces IH Isub by(auto simp: Int_Un_distrib)
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> ?\<sigma>\<^sub>1" using CondT\<^sub>1 by auto
- finally have eval\<^sub>1: "P \<turnstile> ?\<sigma>\<^sub>0 -jvm\<rightarrow> ?\<sigma>\<^sub>1".
- let ?pc\<^sub>1' = "?pc\<^sub>1 + 1 + length(compE\<^sub>2 e\<^sub>1)"
- let ?pc\<^sub>2' = "?pc\<^sub>1' + 1 + length(compE\<^sub>2 e\<^sub>2)"
- let ?I' = "I - pcs(compxE\<^sub>2 e pc ?d) - pcs(compxE\<^sub>2 e\<^sub>2 (?pc\<^sub>1'+1) ?d)"
- have IH2: "PROP ?P e\<^sub>1 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 e' h\<^sub>2 ls\<^sub>2 sh\<^sub>2 E C M (?pc\<^sub>1+1) ics v xa vs frs ?I'" by fact
- show ?case (is "?Norm \<and> ?Err")
- proof
- show ?Norm (is "?val \<longrightarrow> ?trans")
- proof
- assume val: ?val
- note eval\<^sub>1
- also have "P \<turnstile> ?\<sigma>\<^sub>1 -jvm\<rightarrow> (None,h\<^sub>2,(v#vs,ls\<^sub>2,C,M,?pc\<^sub>1',ics)#frs,sh\<^sub>2)"
- using val CondT\<^sub>1.prems nsub_RI_Jcc_pieces IH2 by(fastforce simp:Int_Un_distrib)
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> (None,h\<^sub>2,(v#vs,ls\<^sub>2,C,M,?pc\<^sub>2',ics)#frs,sh\<^sub>2)"
- using CondT\<^sub>1 nsub_RI_Jcc_pieces by(auto simp:add.assoc)
- finally show ?trans by(simp add:add.assoc)
- qed
- next
- show ?Err (is "?throw \<longrightarrow> ?err")
- proof
- assume throw: ?throw
- moreover
- note IH2
- ultimately obtain pc\<^sub>2 vs' where
- pc\<^sub>2: "?pc\<^sub>1+1 \<le> pc\<^sub>2 \<and> pc\<^sub>2 < ?pc\<^sub>1' \<and>
- \<not> caught P pc\<^sub>2 h\<^sub>2 xa (compxE\<^sub>2 e\<^sub>1 (?pc\<^sub>1+1) (size vs))" and
- eval\<^sub>2: "P \<turnstile> ?\<sigma>\<^sub>1 -jvm\<rightarrow> handle P C M xa h\<^sub>2 (vs'@vs) ls\<^sub>2 pc\<^sub>2 ics frs sh\<^sub>2"
- using CondT\<^sub>1.prems nsub_RI_Jcc_pieces by (fastforce simp:Int_Un_distrib)
- show "?err" using pc\<^sub>2 jvm_trans[OF eval\<^sub>1 eval\<^sub>2] by(auto intro: exI[where x=pc\<^sub>2])
- qed
- qed
-next
- case (CondF\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 e\<^sub>2 e' h\<^sub>2 ls\<^sub>2 sh\<^sub>2 e\<^sub>1)
- let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e)"
- let ?pc\<^sub>2 = "?pc\<^sub>1 + 1 + length(compE\<^sub>2 e\<^sub>1)+ 1"
- let ?pc\<^sub>2' = "?pc\<^sub>2 + length(compE\<^sub>2 e\<^sub>2)"
- let ?\<sigma>\<^sub>0 = "(None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0)"
- let ?\<sigma>\<^sub>1 = "(None,h\<^sub>1,(vs,ls\<^sub>1,C,M,?pc\<^sub>2,ics)#frs,sh\<^sub>1)"
- let ?d = "size vs"
- let ?xt\<^sub>1 = "compxE\<^sub>2 e\<^sub>1 (pc+size(compE\<^sub>2 e)+1) ?d"
- let ?xt\<^sub>2 = "compxE\<^sub>2 e\<^sub>2 (pc+size(compE\<^sub>2 e)+size(compE\<^sub>2 e\<^sub>1)+2) ?d"
- let ?I = "I - (pcs ?xt\<^sub>1 \<union> pcs ?xt\<^sub>2)"
- let ?I' = "I - pcs(compxE\<^sub>2 e pc ?d) - pcs(compxE\<^sub>2 e\<^sub>1 (?pc\<^sub>1+1) ?d)"
- have pcs: "pcs(compxE\<^sub>2 e pc ?d) \<inter> pcs(?xt\<^sub>1 @ ?xt\<^sub>2) = {}"
- using CondF\<^sub>1.prems by (simp add:Int_Un_distrib)
- have Isub: "{pc..<pc + length (compE\<^sub>2 e)} \<subseteq> ?I" using CondF\<^sub>1.prems by clarsimp
- have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 false h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Bool False) xa vs frs ?I" by fact
- have IH2: "PROP ?P e\<^sub>2 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 e' h\<^sub>2 ls\<^sub>2 sh\<^sub>2 E C M ?pc\<^sub>2 ics v xa vs frs ?I'" by fact
- have "P \<turnstile> ?\<sigma>\<^sub>0 -jvm\<rightarrow> (None,h\<^sub>1,(Bool(False)#vs,ls\<^sub>1,C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)"
- using CondF\<^sub>1.prems nsub_RI_Jcc_pieces IH Isub pcs by auto
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> ?\<sigma>\<^sub>1" using CondF\<^sub>1 by auto
- finally have eval\<^sub>1: "P \<turnstile> ?\<sigma>\<^sub>0 -jvm\<rightarrow> ?\<sigma>\<^sub>1".
- show ?case (is "?Norm \<and> ?Err")
- proof
- show ?Norm (is "?val \<longrightarrow> ?trans")
- proof
- assume val: ?val
- note eval\<^sub>1
- also have "P \<turnstile> ?\<sigma>\<^sub>1 -jvm\<rightarrow> (None,h\<^sub>2,(v#vs,ls\<^sub>2,C,M,?pc\<^sub>2',ics)#frs,sh\<^sub>2)"
- using val CondF\<^sub>1.prems nsub_RI_Jcc_pieces IH2 by(fastforce simp:Int_Un_distrib)
- finally show ?trans by(simp add:add.assoc)
- qed
- next
- show ?Err (is "?throw \<longrightarrow> ?err")
- proof
- let ?I' = "I - pcs(compxE\<^sub>2 e pc ?d) - pcs(compxE\<^sub>2 e\<^sub>1 (?pc\<^sub>1+1) ?d)"
- assume throw: ?throw
- then obtain pc\<^sub>2 vs' where
- pc\<^sub>2: "?pc\<^sub>2 \<le> pc\<^sub>2 \<and> pc\<^sub>2 < ?pc\<^sub>2' \<and>
- \<not> caught P pc\<^sub>2 h\<^sub>2 xa (compxE\<^sub>2 e\<^sub>2 ?pc\<^sub>2 ?d)" and
- eval\<^sub>2: "P \<turnstile> ?\<sigma>\<^sub>1 -jvm\<rightarrow> handle P C M xa h\<^sub>2 (vs'@vs) ls\<^sub>2 pc\<^sub>2 ics frs sh\<^sub>2"
- using CondF\<^sub>1.prems nsub_RI_Jcc_pieces IH2 by(fastforce simp:Int_Un_distrib)
- show "?err" using pc\<^sub>2 jvm_trans[OF eval\<^sub>1 eval\<^sub>2] by(auto intro: exI[where x=pc\<^sub>2])
- qed
- qed
-next
- case (CondThrow\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 f h\<^sub>1 ls\<^sub>1 sh\<^sub>1 e\<^sub>1 e\<^sub>2)
- let ?d = "size vs"
- let ?xt\<^sub>1 = "compxE\<^sub>2 e\<^sub>1 (pc+size(compE\<^sub>2 e)+1) ?d"
- let ?xt\<^sub>2 = "compxE\<^sub>2 e\<^sub>2 (pc+size(compE\<^sub>2 e)+size(compE\<^sub>2 e\<^sub>1)+2) ?d"
- let ?I = "I - (pcs ?xt\<^sub>1 \<union> pcs ?xt\<^sub>2)"
- have Isub: "{pc..<pc + length (compE\<^sub>2 e)} \<subseteq> ?I" using CondThrow\<^sub>1.prems by clarsimp
- have "pcs(compxE\<^sub>2 e pc ?d) \<inter> pcs(?xt\<^sub>1 @ ?xt\<^sub>2) = {}"
- using CondThrow\<^sub>1.prems by (simp add:Int_Un_distrib)
- moreover have "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (throw f) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics v xa vs frs ?I" by fact
- ultimately show ?case using CondThrow\<^sub>1.prems nsub_RI_Jcc_pieces Isub by auto
-next
- case (WhileF\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 c)
- let ?pc = "pc + length(compE\<^sub>2 e)"
- let ?pc' = "?pc + length(compE\<^sub>2 c) + 3"
- have Isub: "{pc..<pc + length (compE\<^sub>2 e)} \<subseteq> I - pcs (compxE\<^sub>2 c (Suc (pc + length (compE\<^sub>2 e))) (length vs))"
- using WhileF\<^sub>1.prems by clarsimp
- have Isub2: "{Suc (pc + length (compE\<^sub>2 e))..<Suc (pc + length (compE\<^sub>2 e) + length (compE\<^sub>2 c))}
- \<subseteq> I - pcs (compxE\<^sub>2 e pc (length vs))" using WhileF\<^sub>1.prems by clarsimp
- have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 false h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Bool False) xa vs frs
- (I - pcs (compxE\<^sub>2 c (Suc (pc + length (compE\<^sub>2 e))) (length vs)))" by fact
- have "P \<turnstile> (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0) -jvm\<rightarrow>
- (None,h\<^sub>1,(Bool False#vs,ls\<^sub>1,C,M,?pc,ics)#frs,sh\<^sub>1)"
- using WhileF\<^sub>1.prems nsub_RI_Jcc_pieces IH Isub by auto
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> (None,h\<^sub>1,(vs,ls\<^sub>1,C,M,?pc',ics)#frs,sh\<^sub>1)"
- using WhileF\<^sub>1 by (auto simp:add.assoc)
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> (None,h\<^sub>1,(Unit#vs,ls\<^sub>1,C,M,?pc'+1,ics)#frs,sh\<^sub>1)"
- using WhileF\<^sub>1.prems by (auto simp:eval_nat_numeral)
- finally show ?case by (simp add:add.assoc eval_nat_numeral)
-next
- case (WhileT\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 c v\<^sub>1 h\<^sub>2 ls\<^sub>2 sh\<^sub>2 e\<^sub>3 h\<^sub>3 ls\<^sub>3 sh\<^sub>3)
- let ?pc = "pc + length(compE\<^sub>2 e)"
- let ?pc' = "?pc + length(compE\<^sub>2 c) + 1"
- let ?\<sigma>\<^sub>0 = "(None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0)"
- let ?\<sigma>\<^sub>2 = "(None,h\<^sub>2,(vs,ls\<^sub>2,C,M,pc,ics)#frs,sh\<^sub>2)"
- have Isub: "{pc..<pc + length (compE\<^sub>2 e)} \<subseteq> I - pcs (compxE\<^sub>2 c (Suc (pc + length (compE\<^sub>2 e))) (length vs))"
- using WhileT\<^sub>1.prems by clarsimp
- have Isub2: "{Suc (pc + length (compE\<^sub>2 e))..<Suc (pc + length (compE\<^sub>2 e) + length (compE\<^sub>2 c))}
- \<subseteq> I - pcs (compxE\<^sub>2 e pc (length vs))" using WhileT\<^sub>1.prems by clarsimp
- have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 true h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Bool True) xa vs frs
- (I - pcs (compxE\<^sub>2 c (Suc (pc + length (compE\<^sub>2 e))) (length vs)))" by fact
- have IH2: "PROP ?P c h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (Val v\<^sub>1) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 E C M (Suc ?pc) ics v\<^sub>1 xa vs frs
- (I - pcs (compxE\<^sub>2 e pc (length vs)))" by fact
- have "P \<turnstile> ?\<sigma>\<^sub>0 -jvm\<rightarrow> (None,h\<^sub>1,(Bool True#vs,ls\<^sub>1,C,M,?pc,ics)#frs,sh\<^sub>1)"
- using WhileT\<^sub>1.prems nsub_RI_Jcc_pieces IH Isub by auto
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> (None,h\<^sub>1,(vs,ls\<^sub>1,C,M,?pc+1,ics)#frs,sh\<^sub>1)"
- using WhileT\<^sub>1.prems by auto
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> (None,h\<^sub>2,(v\<^sub>1#vs,ls\<^sub>2,C,M,?pc',ics)#frs,sh\<^sub>2)"
- using WhileT\<^sub>1.prems nsub_RI_Jcc_pieces IH2 Isub2 by auto
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> ?\<sigma>\<^sub>2" using WhileT\<^sub>1.prems by auto
- finally have 1: "P \<turnstile> ?\<sigma>\<^sub>0 -jvm\<rightarrow> ?\<sigma>\<^sub>2".
- show ?case (is "?Norm \<and> ?Err")
- proof
- show ?Norm (is "?val \<longrightarrow> ?trans")
- proof
- assume val: ?val
- note 1
- also have "P \<turnstile> ?\<sigma>\<^sub>2 -jvm\<rightarrow> (None,h\<^sub>3,(v#vs,ls\<^sub>3,C,M,?pc'+3,ics)#frs,sh\<^sub>3)"
- using val WhileT\<^sub>1 by (auto simp add:add.assoc eval_nat_numeral)
- finally show ?trans by(simp add:add.assoc eval_nat_numeral)
- qed
- next
- show ?Err (is "?throw \<longrightarrow> ?err")
- proof
- assume throw: ?throw
- moreover
- have "PROP ?P (while (e) c) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 e\<^sub>3 h\<^sub>3 ls\<^sub>3 sh\<^sub>3 E C M pc ics v xa vs frs I" by fact
- ultimately obtain pc\<^sub>2 vs' where
- pc\<^sub>2: "pc \<le> pc\<^sub>2 \<and> pc\<^sub>2 < ?pc'+3 \<and>
- \<not> caught P pc\<^sub>2 h\<^sub>3 xa (compxE\<^sub>2 (while (e) c) pc (size vs))" and
- 2: "P \<turnstile> ?\<sigma>\<^sub>2 -jvm\<rightarrow> handle P C M xa h\<^sub>3 (vs'@vs) ls\<^sub>3 pc\<^sub>2 ics frs sh\<^sub>3"
- using WhileT\<^sub>1.prems by (auto simp:add.assoc eval_nat_numeral)
- show "?err" using pc\<^sub>2 jvm_trans[OF 1 2] by(auto intro: exI[where x=pc\<^sub>2])
- qed
- qed
-next
- case (WhileCondThrow\<^sub>1 e h ls sh e' h' ls' sh' c)
- let ?I = "I - pcs (compxE\<^sub>2 c (Suc (pc + length (compE\<^sub>2 e))) (length vs))"
- obtain a' where throw: "throw e' = Throw a'" using eval\<^sub>1_final[OF WhileCondThrow\<^sub>1.hyps(1)] by clarsimp
- have Isub: "{pc..<pc + length (compE\<^sub>2 e)} \<subseteq> ?I" using WhileCondThrow\<^sub>1.prems by clarsimp
- have "PROP ?P e h ls sh (throw e') h' ls' sh' E C M pc ics v a' vs frs ?I" by fact
- then show ?case using WhileCondThrow\<^sub>1.prems throw nsub_RI_Jcc_pieces Isub by auto
-next
- case (WhileBodyThrow\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 c e' h\<^sub>2 ls\<^sub>2 sh\<^sub>2)
- let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e)"
- let ?\<sigma>\<^sub>0 = "(None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0)"
- let ?\<sigma>\<^sub>1 = "(None,h\<^sub>1,(vs,ls\<^sub>1,C,M,?pc\<^sub>1+1,ics)#frs,sh\<^sub>1)"
- let ?I = "I - pcs (compxE\<^sub>2 c (Suc (pc + length (compE\<^sub>2 e))) (length vs))"
- have Isub: "{pc..<pc + length (compE\<^sub>2 e)} \<subseteq> ?I"
- using WhileBodyThrow\<^sub>1.prems by clarsimp
- have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 true h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Bool True) xa vs frs ?I" by fact
- then have "P \<turnstile> ?\<sigma>\<^sub>0 -jvm\<rightarrow> (None,h\<^sub>1,(Bool(True)#vs,ls\<^sub>1,C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)"
- using WhileBodyThrow\<^sub>1.prems nsub_RI_Jcc_pieces Isub by auto
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> ?\<sigma>\<^sub>1" using WhileBodyThrow\<^sub>1 by auto
- finally have eval\<^sub>1: "P \<turnstile> ?\<sigma>\<^sub>0 -jvm\<rightarrow> ?\<sigma>\<^sub>1".
- let ?pc\<^sub>1' = "?pc\<^sub>1 + 1 + length(compE\<^sub>2 c)"
- show ?case (is "?Norm \<and> ?Err")
- proof
- show ?Norm by simp
- next
- show ?Err (is "?throw \<longrightarrow> ?err")
- proof
- assume throw: ?throw
- moreover
- have "PROP ?P c h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (throw e') h\<^sub>2 ls\<^sub>2 sh\<^sub>2 E C M (?pc\<^sub>1+1) ics v xa vs frs
- (I - pcs (compxE\<^sub>2 e pc (size vs)))" by fact
- ultimately obtain pc\<^sub>2 vs' where
- pc\<^sub>2: "?pc\<^sub>1+1 \<le> pc\<^sub>2 \<and> pc\<^sub>2 < ?pc\<^sub>1' \<and>
- \<not> caught P pc\<^sub>2 h\<^sub>2 xa (compxE\<^sub>2 c (?pc\<^sub>1+1) (size vs))" and
- eval\<^sub>2: "P \<turnstile> ?\<sigma>\<^sub>1 -jvm\<rightarrow> handle P C M xa h\<^sub>2 (vs'@vs) ls\<^sub>2 pc\<^sub>2 ics frs sh\<^sub>2"
- using WhileBodyThrow\<^sub>1.prems nsub_RI_Jcc_pieces by (fastforce simp:Int_Un_distrib)
- show "?err" using pc\<^sub>2 jvm_trans[OF eval\<^sub>1 eval\<^sub>2] by(auto intro: exI[where x=pc\<^sub>2])
- qed
- qed
-next
- case (Throw\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 a h\<^sub>1 ls\<^sub>1 sh\<^sub>1)
- let ?pc = "pc + size(compE\<^sub>2 e)"
- have Isub: "{pc..<pc + length (compE\<^sub>2 e)} \<subseteq> I" using Throw\<^sub>1.prems by clarsimp
- show ?case (is "?Norm \<and> ?Err")
- proof
- show ?Norm by simp
- next
- show ?Err (is "?throw \<longrightarrow> ?err")
- proof
- assume throw:?throw
- have "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (addr a) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Addr a) a vs frs I" by fact
- then have "P \<turnstile> (None, h\<^sub>0, (vs, ls\<^sub>0, C, M, pc, ics) # frs, sh\<^sub>0) -jvm\<rightarrow>
- (None, h\<^sub>1, (Addr xa#vs, ls\<^sub>1, C, M, ?pc, ics) # frs, sh\<^sub>1)"
- using Throw\<^sub>1 nsub_RI_Jcc_pieces Isub throw by auto
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> handle P C M xa h\<^sub>1 (Addr xa#vs) ls\<^sub>1 ?pc ics frs sh\<^sub>1"
- using Throw\<^sub>1.prems by(auto simp add:handle_def)
- finally show "?err" by(auto intro!: exI[where x="?pc"] exI[where x="[Addr xa]"])
- qed
- qed
-next
- case (ThrowNull\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 h\<^sub>1 ls\<^sub>1 sh\<^sub>1)
- let ?pc = "pc + size(compE\<^sub>2 e)"
- let ?xa = "addr_of_sys_xcpt NullPointer"
- have Isub: "{pc..<pc + length (compE\<^sub>2 e)} \<subseteq> I" using ThrowNull\<^sub>1.prems by clarsimp
- show ?case (is "?Norm \<and> ?Err")
- proof
- show ?Norm by simp
- next
- show ?Err (is "?throw \<longrightarrow> ?err")
- proof
- assume throw: ?throw
- have "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 null h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics Null xa vs frs I" by fact
- then have "P \<turnstile> (None, h\<^sub>0, (vs, ls\<^sub>0, C, M, pc, ics) # frs, sh\<^sub>0) -jvm\<rightarrow>
- (None, h\<^sub>1, (Null#vs, ls\<^sub>1, C, M, ?pc, ics) # frs, sh\<^sub>1)"
- using ThrowNull\<^sub>1.prems nsub_RI_Jcc_pieces Isub by auto
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> handle P C M ?xa h\<^sub>1 (Null#vs) ls\<^sub>1 ?pc ics frs sh\<^sub>1"
- using ThrowNull\<^sub>1.prems by(auto simp add:handle_def)
- finally show "?err" using throw by(auto intro!: exI[where x="?pc"] exI[where x="[Null]"])
- qed
- qed
-next
- case (ThrowThrow\<^sub>1 e h ls sh e' h' ls' sh')
- obtain a' where throw: "throw e' = Throw a'" using eval\<^sub>1_final[OF ThrowThrow\<^sub>1.hyps(1)] by clarsimp
- have Isub: "{pc..<pc + length (compE\<^sub>2 e)} \<subseteq> I" using ThrowThrow\<^sub>1.prems by clarsimp
- have "PROP ?P e h ls sh (throw e') h' ls' sh' E C M pc ics v a' vs frs I" by fact
- then show ?case using ThrowThrow\<^sub>1.prems throw nsub_RI_Jcc_pieces Isub by auto
-next
- case (Try\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 v\<^sub>1 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 Ci i e\<^sub>2)
- let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e\<^sub>1)"
- let ?pc\<^sub>1' = "?pc\<^sub>1 + 2 + length(compE\<^sub>2 e\<^sub>2)"
- have "{pc..<pc+size(compE\<^sub>2 (try e\<^sub>1 catch(Ci i) e\<^sub>2))} \<subseteq> I" using Try\<^sub>1.prems by simp
- also have "P,C,M \<rhd> compxE\<^sub>2 (try e\<^sub>1 catch(Ci i) e\<^sub>2) pc (size vs) / I,size vs"
- using Try\<^sub>1.prems by simp
- ultimately have "P,C,M \<rhd> compxE\<^sub>2 e\<^sub>1 pc (size vs) / {pc..<pc + length (compE\<^sub>2 e\<^sub>1)},size vs"
- by(rule beforex_try)
- hence "P \<turnstile> (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0) -jvm\<rightarrow>
- (None,h\<^sub>1,(v\<^sub>1#vs,ls\<^sub>1,C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)"
- using Try\<^sub>1 nsub_RI_Jcc_pieces by auto blast
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> (None,h\<^sub>1,(v\<^sub>1#vs,ls\<^sub>1,C,M,?pc\<^sub>1',ics)#frs,sh\<^sub>1)"
- using Try\<^sub>1.prems by auto
- finally show ?case by (auto simp:add.assoc)
-next
- case (TryCatch\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 a h\<^sub>1 ls\<^sub>1 sh\<^sub>1 D fs Ci i e\<^sub>2 e\<^sub>2' h\<^sub>2 ls\<^sub>2 sh\<^sub>2)
- let ?e = "try e\<^sub>1 catch(Ci i) e\<^sub>2"
- let ?xt = "compxE\<^sub>2 ?e pc (size vs)"
- let ?\<sigma>\<^sub>0 = "(None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0)"
- let ?ls\<^sub>1 = "ls\<^sub>1[i := Addr a]"
- let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e\<^sub>1)"
- let ?pc\<^sub>1' = "?pc\<^sub>1 + 2"
- let ?\<sigma>\<^sub>1 = "(None,h\<^sub>1,(vs,?ls\<^sub>1,C,M, ?pc\<^sub>1',ics) # frs,sh\<^sub>1)"
- have I: "{pc..<pc + length (compE\<^sub>2 (try e\<^sub>1 catch(Ci i) e\<^sub>2))} \<subseteq> I"
- and beforex: "P,C,M \<rhd> ?xt/I,size vs" using TryCatch\<^sub>1.prems by simp+
- have "P \<turnstile> ?\<sigma>\<^sub>0 -jvm\<rightarrow> (None,h\<^sub>1,((Addr a)#vs,ls\<^sub>1,C,M, ?pc\<^sub>1+1,ics) # frs,sh\<^sub>1)"
- proof -
- have ics: "ics = No_ics" using TryCatch\<^sub>1.prems by auto
- have "PROP ?P e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (Throw a) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics v a vs frs {pc..<pc + length (compE\<^sub>2 e\<^sub>1)}"
- by fact
- moreover have "P,C,M \<rhd> compxE\<^sub>2 e\<^sub>1 pc (size vs)/{pc..<?pc\<^sub>1},size vs"
- using beforex I pcs_subset by(force elim!: beforex_appendD1)
- ultimately have
- "\<exists>pc\<^sub>1. pc \<le> pc\<^sub>1 \<and> pc\<^sub>1 < ?pc\<^sub>1 \<and>
- \<not> caught P pc\<^sub>1 h\<^sub>1 a (compxE\<^sub>2 e\<^sub>1 pc (size vs)) \<and>
- (\<exists>vs'. P \<turnstile> ?\<sigma>\<^sub>0 -jvm\<rightarrow> handle P C M a h\<^sub>1 (vs'@vs) ls\<^sub>1 pc\<^sub>1 ics frs sh\<^sub>1)"
- using TryCatch\<^sub>1.prems nsub_RI_Jcc_pieces by auto
- then obtain pc\<^sub>1 vs' where
- pc\<^sub>1_in_e\<^sub>1: "pc \<le> pc\<^sub>1" "pc\<^sub>1 < ?pc\<^sub>1" and
- pc\<^sub>1_not_caught: "\<not> caught P pc\<^sub>1 h\<^sub>1 a (compxE\<^sub>2 e\<^sub>1 pc (size vs))" and
- 0: "P \<turnstile> ?\<sigma>\<^sub>0 -jvm\<rightarrow> handle P C M a h\<^sub>1 (vs'@vs) ls\<^sub>1 pc\<^sub>1 ics frs sh\<^sub>1" by iprover
- from beforex obtain xt\<^sub>0 xt\<^sub>1
- where ex_tab: "ex_table_of P C M = xt\<^sub>0 @ ?xt @ xt\<^sub>1"
- and disj: "pcs xt\<^sub>0 \<inter> I = {}" by(auto simp:beforex_def)
- have hp: "h\<^sub>1 a = Some (D, fs)" "P\<^sub>1 \<turnstile> D \<preceq>\<^sup>* Ci" by fact+
- have "pc\<^sub>1 \<notin> pcs xt\<^sub>0" using pc\<^sub>1_in_e\<^sub>1 I disj by auto
- with pc\<^sub>1_in_e\<^sub>1 pc\<^sub>1_not_caught hp
- show ?thesis using ex_tab 0 ics by(simp add:handle_def matches_ex_entry_def)
- qed
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> ?\<sigma>\<^sub>1" using TryCatch\<^sub>1 by auto
- finally have 1: "P \<turnstile> ?\<sigma>\<^sub>0 -jvm\<rightarrow> ?\<sigma>\<^sub>1" .
- let ?pc\<^sub>2 = "?pc\<^sub>1' + length(compE\<^sub>2 e\<^sub>2)"
- let ?I\<^sub>2 = "{?pc\<^sub>1' ..< ?pc\<^sub>2}"
- have "P,C,M \<rhd> compxE\<^sub>2 ?e pc (size vs) / I,size vs" by fact
- hence beforex\<^sub>2: "P,C,M \<rhd> compxE\<^sub>2 e\<^sub>2 ?pc\<^sub>1' (size vs) / ?I\<^sub>2, size vs"
- using I pcs_subset[of _ ?pc\<^sub>1'] by(auto elim!:beforex_appendD2)
- have IH\<^sub>2: "PROP ?P e\<^sub>2 h\<^sub>1 ?ls\<^sub>1 sh\<^sub>1 e\<^sub>2' h\<^sub>2 ls\<^sub>2 sh\<^sub>2 E C M ?pc\<^sub>1' ics v xa vs frs ?I\<^sub>2" by fact
- show ?case (is "?Norm \<and> ?Err")
- proof
- show ?Norm (is "?val \<longrightarrow> ?trans")
- proof
- assume val: ?val
- note 1 also have "P \<turnstile> ?\<sigma>\<^sub>1 -jvm\<rightarrow> (None,h\<^sub>2,(v#vs,ls\<^sub>2,C,M,?pc\<^sub>2,ics)#frs,sh\<^sub>2)"
- using val beforex\<^sub>2 IH\<^sub>2 TryCatch\<^sub>1.prems nsub_RI_Jcc_pieces by auto
- finally show ?trans by(simp add:add.assoc)
- qed
- next
- show ?Err (is "?throw \<longrightarrow> ?err")
- proof
- assume throw: ?throw
- then obtain pc\<^sub>2 vs' where
- pc\<^sub>2: "?pc\<^sub>1+2 \<le> pc\<^sub>2 \<and> pc\<^sub>2 < ?pc\<^sub>2 \<and>
- \<not> caught P pc\<^sub>2 h\<^sub>2 xa (compxE\<^sub>2 e\<^sub>2 ?pc\<^sub>1' (size vs))" and
- 2: "P \<turnstile> ?\<sigma>\<^sub>1 -jvm\<rightarrow> handle P C M xa h\<^sub>2 (vs'@vs) ls\<^sub>2 pc\<^sub>2 ics frs sh\<^sub>2"
- using IH\<^sub>2 beforex\<^sub>2 TryCatch\<^sub>1.prems nsub_RI_Jcc_pieces by auto
- show ?err using pc\<^sub>2 jvm_trans[OF 1 2]
- by (simp add:match_ex_entry) (auto intro: exI[where x=pc\<^sub>2])
- qed
- qed
-next
- case (TryThrow\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 a h\<^sub>1 ls\<^sub>1 sh\<^sub>1 D fs Ci i e\<^sub>2)
- let ?\<sigma>\<^sub>0 = "(None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0)"
- let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e\<^sub>1)"
- let ?e = "try e\<^sub>1 catch(Ci i) e\<^sub>2"
- let ?xt = "compxE\<^sub>2 ?e pc (size vs)"
- have I: "{pc..<pc + length (compE\<^sub>2 (try e\<^sub>1 catch(Ci i) e\<^sub>2))} \<subseteq> I"
- and beforex: "P,C,M \<rhd> ?xt/I,size vs" using TryThrow\<^sub>1.prems by simp+
- have "PROP ?P e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (Throw a) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics v a vs frs
- {pc..<pc + length (compE\<^sub>2 e\<^sub>1)}" by fact
- moreover have "P,C,M \<rhd> compxE\<^sub>2 e\<^sub>1 pc (size vs)/{pc..<?pc\<^sub>1},size vs"
- using beforex I pcs_subset by(force elim!: beforex_appendD1)
- ultimately have
- "\<exists>pc\<^sub>1. pc \<le> pc\<^sub>1 \<and> pc\<^sub>1 < ?pc\<^sub>1 \<and>
- \<not> caught P pc\<^sub>1 h\<^sub>1 a (compxE\<^sub>2 e\<^sub>1 pc (size vs)) \<and>
- (\<exists>vs'. P \<turnstile> ?\<sigma>\<^sub>0 -jvm\<rightarrow> handle P C M a h\<^sub>1 (vs'@vs) ls\<^sub>1 pc\<^sub>1 ics frs sh\<^sub>1)"
- using TryThrow\<^sub>1.prems nsub_RI_Jcc_pieces by auto
- then obtain pc\<^sub>1 vs' where
- pc\<^sub>1_in_e\<^sub>1: "pc \<le> pc\<^sub>1" "pc\<^sub>1 < ?pc\<^sub>1" and
- pc\<^sub>1_not_caught: "\<not> caught P pc\<^sub>1 h\<^sub>1 a (compxE\<^sub>2 e\<^sub>1 pc (size vs))" and
- 0: "P \<turnstile> ?\<sigma>\<^sub>0 -jvm\<rightarrow> handle P C M a h\<^sub>1 (vs'@vs) ls\<^sub>1 pc\<^sub>1 ics frs sh\<^sub>1" by iprover
- show ?case (is "?N \<and> (?eq \<longrightarrow> ?err)")
- proof
- show ?N by simp
- next
- { assume ?eq
- with TryThrow\<^sub>1 pc\<^sub>1_in_e\<^sub>1 pc\<^sub>1_not_caught 0
- have "?err" by (simp add:match_ex_entry) auto
- }
- thus "?eq \<longrightarrow> ?err" by iprover
- qed
-next
- case Nil\<^sub>1 thus ?case by simp
-next
- case (Cons\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 v h\<^sub>1 ls\<^sub>1 sh\<^sub>1 es fs h\<^sub>2 ls\<^sub>2 sh\<^sub>2)
- let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e)"
- let ?\<sigma>\<^sub>0 = "(None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0)"
- let ?\<sigma>\<^sub>1 = "(None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)"
- have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (Val v) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 [] C M pc ics v xa vs frs
- (I - pcs (compxEs\<^sub>2 es ?pc\<^sub>1 (Suc (length vs))))" by fact
- then have 1: "P \<turnstile> ?\<sigma>\<^sub>0 -jvm\<rightarrow> ?\<sigma>\<^sub>1" using Jcc_pieces_Cons[OF _ Cons\<^sub>1.prems(1-5)] by auto
- let ?pc\<^sub>2 = "?pc\<^sub>1 + length(compEs\<^sub>2 es)"
- have IHs: "PROP ?Ps es h\<^sub>1 ls\<^sub>1 sh\<^sub>1 fs h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C M ?pc\<^sub>1 ics (tl ws) xa es' (v#vs) frs
- (I - pcs (compxE\<^sub>2 e pc (length vs)))" by fact
- show ?case (is "?Norm \<and> ?Err")
- proof
- show ?Norm (is "?val \<longrightarrow> ?trans")
- proof
- assume val: ?val
- note 1
- also have "P \<turnstile> ?\<sigma>\<^sub>1 -jvm\<rightarrow> (None,h\<^sub>2,(rev(ws) @ vs,ls\<^sub>2,C,M,?pc\<^sub>2,ics)#frs,sh\<^sub>2)"
- using val IHs Cons\<^sub>1.prems by fastforce
- finally show ?trans by(simp add:add.assoc)
- qed
- next
- show ?Err (is "?throw \<longrightarrow> (\<exists>pc\<^sub>2. ?H pc\<^sub>2)")
- proof
- assume throw: ?throw
- then obtain pc\<^sub>2 vs' where
- pc\<^sub>2: "?pc\<^sub>1 \<le> pc\<^sub>2 \<and> pc\<^sub>2 < ?pc\<^sub>2 \<and>
- \<not> caught P pc\<^sub>2 h\<^sub>2 xa (compxEs\<^sub>2 es ?pc\<^sub>1 (size vs + 1))" and
- 2: "P \<turnstile> ?\<sigma>\<^sub>1 -jvm\<rightarrow> handle P C M xa h\<^sub>2 (vs'@v#vs) ls\<^sub>2 pc\<^sub>2 ics frs sh\<^sub>2"
- using IHs Cons\<^sub>1.prems by(fastforce simp:Cons_eq_append_conv neq_Nil_conv)
- have "?H pc\<^sub>2" using Cons\<^sub>1.prems pc\<^sub>2 jvm_trans[OF 1 2] by(auto intro!: exI[where x="vs'@[v]"])
- thus "\<exists>pc\<^sub>2. ?H pc\<^sub>2" by iprover
- qed
- qed
-next
- case (ConsThrow\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 a h\<^sub>1 ls\<^sub>1 sh\<^sub>1 es)
- then show ?case using Jcc_pieces_Cons[OF _ ConsThrow\<^sub>1.prems(1-5)]
- by (fastforce simp:Cons_eq_append_conv)
-next
- case InitFinal\<^sub>1 then show ?case using eval\<^sub>1_final_same[OF InitFinal\<^sub>1.hyps(1)] by clarsimp
-next
- case (InitNone\<^sub>1 sh C\<^sub>0 C' Cs e h l e' h' l' sh')
- then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' sh' v xa
- (INIT C' (C\<^sub>0 # Cs,False) \<leftarrow> e)
- = (True, frs', (None, h', (vs, l, C, M, pc, Called []) # frs, sh'), err)"
- using InitNone\<^sub>1.prems(1) by clarsimp
- let ?sh = "(sh(C\<^sub>0 \<mapsto> (sblank P\<^sub>1 C\<^sub>0, Prepared)))"
- obtain ics: "ics_of(hd frs') = Calling C\<^sub>0 Cs"
- and frs\<^sub>1: "frs' \<noteq> Nil" using pcs by clarsimp
- then have 1: "P \<turnstile> (None,h,frs',sh) -jvm\<rightarrow> (None,h,frs',?sh)"
- using InitNone\<^sub>1 jvm_InitNone[where P = P] by(cases frs', simp+)
- show ?case (is "(?e1 \<longrightarrow> ?jvm1) \<and> (?e2 \<longrightarrow> ?err)")
- proof(rule conjI)
- { assume val: ?e1
- note 1
- also have "P \<turnstile> (None,h,frs',?sh) -jvm\<rightarrow> (None,h',(vs,l,C,M,pc,Called [])#frs,sh')"
- using InitNone\<^sub>1.hyps(3)[of E] Jcc_pieces_InitNone[OF assms(1) pcs] InitNone\<^sub>1.prems val
- by clarsimp
- finally have ?jvm1 using pcs by simp
- }
- thus "?e1 \<longrightarrow> ?jvm1" by simp
- next
- { assume throw: ?e2
- note 1
- also obtain vs' where "P \<turnstile> (None,h,frs',?sh)
- -jvm\<rightarrow> handle P C M xa h' (vs'@vs) l pc ics frs sh'"
- using InitNone\<^sub>1.hyps(3)[of E] Jcc_pieces_InitNone[OF assms(1) pcs] throw
- by clarsimp presburger
- finally have ?err using pcs by auto
- }
- thus "?e2 \<longrightarrow> ?err" by simp
- qed
-next
- case (InitDone\<^sub>1 sh C\<^sub>0 sfs C' Cs e h l e' h' l' sh')
- then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' sh' v xa
- (INIT C' (C\<^sub>0 # Cs,False) \<leftarrow> e)
- = (True, frs', (None, h', (vs, l, C, M, pc, Called []) # frs, sh'), err)"
- using InitDone\<^sub>1.prems(1) by clarsimp
- let ?frs' = "(calling_to_scalled (hd frs'))#(tl frs')"
- have IH: "PROP ?P (INIT C' (Cs,True) \<leftarrow> e) h l sh e' h' l' sh' E C M pc ics v xa vs frs I"
- by fact
- obtain ics: "ics_of(hd frs') = Calling C\<^sub>0 Cs"
- and frs\<^sub>1: "frs' \<noteq> Nil" using pcs by clarsimp
- then have 1: "P \<turnstile> (None,h,frs',sh) -jvm\<rightarrow> (None,h,?frs',sh)"
- using InitDone\<^sub>1 jvm_InitDP[where P = P] by(cases frs', simp+)
- show ?case (is "(?e1 \<longrightarrow> ?jvm1) \<and> (?e2 \<longrightarrow> ?err)")
- proof(rule conjI)
- { assume val: ?e1
- note 1
- also have "P \<turnstile> (None,h,?frs',sh) -jvm\<rightarrow> (None,h',(vs,l,C,M,pc,Called [])#frs,sh')"
- using IH Jcc_pieces_InitDP[OF assms(1) pcs] InitDone\<^sub>1.prems val by clarsimp
- finally have ?jvm1 using pcs by simp
- }
- thus "?e1 \<longrightarrow> ?jvm1" by simp
- next
- { assume throw: ?e2
- note 1
- also obtain vs' where "P \<turnstile> (None,h,?frs',sh)
- -jvm\<rightarrow> handle P C M xa h' (vs'@vs) l pc ics frs sh'"
- using IH Jcc_pieces_InitDP[OF assms(1) pcs] InitDone\<^sub>1.prems throw by clarsimp
- finally have ?err using pcs by auto
- }
- thus "?e2 \<longrightarrow> ?err" by simp
- qed
-next
- case (InitProcessing\<^sub>1 sh C\<^sub>0 sfs C' Cs e h l e' h' l' sh')
- then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' sh' v xa
- (INIT C' (C\<^sub>0 # Cs,False) \<leftarrow> e)
- = (True, frs', (None, h', (vs, l, C, M, pc, Called []) # frs, sh'), err)"
- using InitProcessing\<^sub>1.prems(1) by clarsimp
- let ?frs' = "(calling_to_scalled (hd frs'))#(tl frs')"
- have IH: "PROP ?P (INIT C' (Cs,True) \<leftarrow> e) h l sh e' h' l' sh' E C M pc ics v xa vs frs I"
- by fact
- obtain ics: "ics_of(hd frs') = Calling C\<^sub>0 Cs"
- and frs\<^sub>1: "frs' \<noteq> Nil" using pcs by clarsimp
- then have 1: "P \<turnstile> (None,h,frs',sh) -jvm\<rightarrow> (None,h,?frs',sh)"
- using InitProcessing\<^sub>1 jvm_InitDP[where P = P] by(cases frs', simp+)
- show ?case (is "(?e1 \<longrightarrow> ?jvm1) \<and> (?e2 \<longrightarrow> ?err)")
- proof(rule conjI)
- { assume val: ?e1
- note 1
- also have "P \<turnstile> (None,h,?frs',sh) -jvm\<rightarrow> (None,h',(vs,l,C,M,pc,Called [])#frs,sh')"
- using IH Jcc_pieces_InitDP[OF assms(1) pcs] InitProcessing\<^sub>1.prems val by clarsimp
- finally have ?jvm1 using pcs by simp
- }
- thus "?e1 \<longrightarrow> ?jvm1" by simp
- next
- { assume throw: ?e2
- note 1
- also obtain vs' where "P \<turnstile> (None,h,?frs',sh)
- -jvm\<rightarrow> handle P C M xa h' (vs'@vs) l pc ics frs sh'"
- using IH Jcc_pieces_InitDP[OF assms(1) pcs] InitProcessing\<^sub>1.prems throw by clarsimp
- finally have ?err using pcs by auto
- }
- thus "?e2 \<longrightarrow> ?err" by simp
- qed
-next
- case (InitError\<^sub>1 sh C\<^sub>0 sfs Cs e h l e' h' l' sh' C')
- then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' sh' v xa
- (INIT C' (C\<^sub>0 # Cs,False) \<leftarrow> e)
- = (True, frs', (None, h', (vs, l, C, M, pc, Called []) # frs, sh'), err)"
- using InitError\<^sub>1.prems(1) by clarsimp
- let ?e\<^sub>0 = "THROW NoClassDefFoundError"
- let ?frs' = "(calling_to_sthrowing (hd frs') (addr_of_sys_xcpt NoClassDefFoundError))#(tl frs')"
- have IH: "PROP ?P (RI (C\<^sub>0,?e\<^sub>0) ; Cs \<leftarrow> e) h l sh e' h' l' sh' E C M pc ics v xa vs frs I" by fact
- obtain ics: "ics_of(hd frs') = Calling C\<^sub>0 Cs"
- and frs\<^sub>1: "frs' \<noteq> Nil"
- and tl: "tl frs' = frs" using pcs by clarsimp
- then have 1: "P \<turnstile> (None,h,frs',sh) -jvm\<rightarrow> (None,h,?frs',sh)"
- proof(cases frs')
- case (Cons a list)
- obtain vs' l' C' M' pc' ics' where a: "a = (vs',l',C',M',pc',ics')" by(cases a)
- then have "ics' = Calling C\<^sub>0 Cs" using Cons ics by simp
- then show ?thesis
- using Cons a IH InitError\<^sub>1.prems jvm_InitError[where P = P] InitError\<^sub>1.hyps(1) by simp
- qed(simp)
- show ?case (is "(?e1 \<longrightarrow> ?jvm1) \<and> (?e2 \<longrightarrow> ?err)")
- proof(rule conjI)
- { assume val: ?e1
- then have False using val rinit\<^sub>1_throw[OF InitError\<^sub>1.hyps(2)] by blast
- then have ?jvm1 using pcs by simp
- }
- thus "?e1 \<longrightarrow> ?jvm1" by simp
- next
- { assume throw: ?e2
- let ?frs = "(calling_to_throwing (hd frs') (addr_of_sys_xcpt NoClassDefFoundError))#(tl frs')"
- have exec: "exec (P, (None,h,?frs,sh)) = Some (None,h,?frs',sh)"
- using exec_ErrorThrowing[where sh=sh, OF InitError\<^sub>1.hyps(1)] ics by(cases "hd frs'", simp)
- obtain vs' where 2: "P \<turnstile> (None,h,?frs,sh) -jvm\<rightarrow> handle P C M xa h' (vs'@vs) l pc ics frs sh'"
- using IH Jcc_pieces_InitError[OF assms(1) pcs InitError\<^sub>1.hyps(1)] throw by clarsimp
- have neq: "(None, h, ?frs, sh) \<noteq> handle P C M xa h' (vs' @ vs) l pc ics frs sh'"
- using tl ics by(cases "hd frs'", simp add: handle_frs_tl_neq)
-
- note 1
- also have "P \<turnstile> (None,h,?frs',sh) -jvm\<rightarrow> handle P C M xa h' (vs'@vs) l pc ics frs sh'"
- using exec_1_exec_all_conf[OF exec 2] neq by simp
- finally have ?err using pcs by auto
- }
- thus "?e2 \<longrightarrow> ?err" by simp
- qed
-next
- case (InitObject\<^sub>1 sh C\<^sub>0 sfs sh' C' Cs e h l e' h' l' sh'')
- then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l'
- (sh(C\<^sub>0 \<mapsto> (sfs, Processing))) v xa (INIT C' (C\<^sub>0 # Cs,False) \<leftarrow> e)
- = (True, frs', (None, h', (vs, l, C, M, pc, Called []) # frs, sh'), err)"
- using InitObject\<^sub>1.prems(1) by clarsimp
- let ?frs' = "(calling_to_called (hd frs'))#(tl frs')"
- have IH: "PROP ?P (INIT C' (C\<^sub>0#Cs,True) \<leftarrow> e) h l sh' e' h' l' sh'' E C M pc ics v xa vs frs I"
- by fact
- obtain ics: "ics_of(hd frs') = Calling C\<^sub>0 Cs"
- and frs\<^sub>1: "frs' \<noteq> Nil" using pcs by clarsimp
- then have 1: "P \<turnstile> (None,h,frs',sh) -jvm\<rightarrow> (None,h,?frs',sh')"
- proof(cases frs')
- case (Cons a list)
- obtain vs' l' C' M' pc' ics' where a: "a = (vs',l',C',M',pc',ics')" by(cases a)
- then have "ics' = Calling C\<^sub>0 Cs" using Cons ics by simp
- then show ?thesis
- using Cons Nil a IH InitObject\<^sub>1 jvm_InitObj[where P = P] by simp
- qed(simp)
- show ?case (is "(?e1 \<longrightarrow> ?jvm1) \<and> (?e2 \<longrightarrow> ?err)")
- proof(rule conjI)
- { assume val: ?e1
- note 1
- also have "P \<turnstile> (None,h,?frs',sh') -jvm\<rightarrow> (None,h',(vs,l,C,M,pc,Called [])#frs,sh'')"
- using IH Jcc_pieces_InitObj[OF assms(1) pcs] InitObject\<^sub>1 val by simp
- finally have ?jvm1 using pcs by simp
- }
- thus "?e1 \<longrightarrow> ?jvm1" by simp
- next
- { assume throw: ?e2
- note 1
- also obtain vs' where "P \<turnstile> (None,h,?frs',sh')
- -jvm\<rightarrow> handle P C M xa h' (vs'@vs) l pc ics frs sh''"
- using IH Jcc_pieces_InitObj[OF assms(1) pcs] InitObject\<^sub>1 throw by clarsimp
- finally have ?err using pcs by auto
- }
- thus "?e2 \<longrightarrow> ?err" by simp
- qed
-next
- case (InitNonObject\<^sub>1 sh C\<^sub>0 sfs D a b sh' C' Cs e h l e' h' l' sh'')
- then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l'
- (sh(C\<^sub>0 \<mapsto> (sfs,Processing))) v xa (INIT C' (C\<^sub>0 # Cs,False) \<leftarrow> e)
- = (True, frs', (None, h', (vs, l, C, M, pc, Called []) # frs, sh'), err)"
- using InitNonObject\<^sub>1.prems(1) by clarsimp
- let ?frs' = "(calling_to_calling (hd frs') D)#(tl frs')"
- have cls1: "is_class P\<^sub>1 D" using InitNonObject\<^sub>1.hyps(2,3) class_wf wf wf_cdecl_supD by blast
- have cls_aux: "distinct (C\<^sub>0#Cs) \<and> supercls_lst P\<^sub>1 (C\<^sub>0#Cs)" using InitNonObject\<^sub>1.prems(1) by auto
- then have cls2: "D \<notin> set (C\<^sub>0 # Cs)"
- proof -
- have "distinct (D # C\<^sub>0 # Cs)"
- using InitNonObject\<^sub>1.hyps(2,3) cls_aux wf wf_supercls_distinct_app by blast
- then show "D \<notin> set (C\<^sub>0 # Cs)"
- by (metis distinct.simps(2))
- qed
- have cls3: "\<forall>C\<in>set (C\<^sub>0 # Cs). P\<^sub>1 \<turnstile> C \<preceq>\<^sup>* D" using InitNonObject\<^sub>1.hyps(2,3) cls_aux
- by (metis r_into_rtrancl rtrancl_into_rtrancl set_ConsD subcls1.subcls1I supercls_lst.simps(1))
- have IH: "PROP ?P (INIT C' (D # C\<^sub>0 # Cs,False) \<leftarrow> e) h l sh' e' h' l' sh'' E C M pc ics v xa vs frs I"
- by fact
- obtain r where cls: "class P C\<^sub>0 = \<lfloor>(D, r)\<rfloor>" using InitNonObject\<^sub>1.hyps(3)
- by (metis assms class_compP compP\<^sub>2_def)
- obtain ics: "ics_of(hd frs') = Calling C\<^sub>0 Cs"
- and frs\<^sub>1: "frs' \<noteq> Nil" using pcs by clarsimp
- then have 1: "P \<turnstile> (None,h,frs',sh) -jvm\<rightarrow> (None,h,?frs',sh')"
- proof(cases frs')
- case (Cons a list)
- obtain vs' l' C' M' pc' ics' where a: "a = (vs',l',C',M',pc',ics')" by(cases a)
- then have "ics' = Calling C\<^sub>0 Cs" using Cons ics by simp
- then show ?thesis
- using Cons a IH InitNonObject\<^sub>1 jvm_InitNonObj[OF _ _ cls] by simp
- qed(simp)
- show ?case (is "(?e1 \<longrightarrow> ?jvm1) \<and> (?e2 \<longrightarrow> ?err)")
- proof(rule conjI)
- { assume val: ?e1
- note 1
- also have "P \<turnstile> (None,h,?frs',sh') -jvm\<rightarrow> (None,h',(vs,l,C,M,pc,Called [])#frs,sh'')"
- using IH Jcc_pieces_InitNonObj[OF assms(1) cls1 cls2 cls3 pcs] InitNonObject\<^sub>1 val by simp
- finally have ?jvm1 using pcs by simp
- }
- thus "?e1 \<longrightarrow> ?jvm1" by simp
- next
- { assume throw: ?e2
- note 1
- also obtain vs' where "P \<turnstile> (None,h,?frs',sh')
- -jvm\<rightarrow> handle P C M xa h' (vs'@vs) l pc ics frs sh''"
- using IH Jcc_pieces_InitNonObj[OF assms(1) cls1 cls2 cls3 pcs] InitNonObject\<^sub>1 throw by clarsimp
- finally have ?err using pcs by auto
- }
- thus "?e2 \<longrightarrow> ?err" by simp
- qed
-next
- case (InitRInit\<^sub>1 C\<^sub>0 Cs e h l sh e' h' l' sh' C')
- then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' sh' v xa
- (INIT C' (C\<^sub>0 # Cs,True) \<leftarrow> e)
- = (True, frs', (None, h', (vs, l, C, M, pc, Called []) # frs, sh'), err)"
- using InitRInit\<^sub>1.prems(1) by clarsimp
- have IH: "PROP ?P (RI (C\<^sub>0,C\<^sub>0\<bullet>\<^sub>sclinit([])) ; Cs \<leftarrow> e) h l sh e' h' l' sh' E C M pc ics v xa vs frs I"
- by fact
- show ?case (is "(?e1 \<longrightarrow> ?jvm1) \<and> (?e2 \<longrightarrow> ?err)")
- proof(rule conjI)
- { assume val: ?e1
- have "P \<turnstile> (None,h,frs',sh) -jvm\<rightarrow> (None,h',(vs,l,C,M,pc,Called [])#frs,sh')"
- using IH Jcc_pieces_InitRInit[OF assms(1,2) pcs] InitRInit\<^sub>1.prems val by simp
- then have ?jvm1 using pcs by simp
- }
- thus "?e1 \<longrightarrow> ?jvm1" by simp
- next
- { assume throw: ?e2
- obtain vs' where "P \<turnstile> (None,h,frs',sh)
- -jvm\<rightarrow> handle P C M xa h' (vs'@vs) l pc ics frs sh'"
- using IH Jcc_pieces_InitRInit[OF assms(1,2) pcs] InitRInit\<^sub>1 throw by clarsimp
- then have ?err using pcs by auto
- }
- thus "?e2 \<longrightarrow> ?err" by simp
- qed
-next
- case (RInit\<^sub>1 e h l sh v1 h' l' sh' C\<^sub>0 sfs i sh'' C' Cs e' e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1)
- let ?frs = "(vs,l,C,M,pc,Called (C\<^sub>0#Cs)) # frs"
- let ?frs' = "(vs,l,C,M,pc,Called Cs) # frs"
- have clinit: "e = C\<^sub>0\<bullet>\<^sub>sclinit([])" using RInit\<^sub>1
- by (metis Jcc_cond.simps(2) eval\<^sub>1_final_same exp.distinct(101) final_def)
- then obtain err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h\<^sub>1 l\<^sub>1 sh\<^sub>1 v xa
- (RI (C\<^sub>0,C\<^sub>0\<bullet>\<^sub>sclinit([])) ; Cs \<leftarrow> e')
- = (True, ?frs, (None, h\<^sub>1, (vs, l, C, M, pc, Called []) # frs, sh\<^sub>1), err)"
- using RInit\<^sub>1.prems(1) by simp
- have shC: "\<forall>C'\<in>set Cs. \<exists>sfs. sh C' = \<lfloor>(sfs, Processing)\<rfloor>" using RInit\<^sub>1.prems(1) clinit by clarsimp
- then have shC'': "\<forall>C'\<in>set Cs. \<exists>sfs. sh'' C' = \<lfloor>(sfs, Processing)\<rfloor>"
- using clinit\<^sub>1_proc_pres[OF wf] RInit\<^sub>1.hyps(1) clinit RInit\<^sub>1.hyps(4) RInit\<^sub>1.prems(1)
- by (auto simp: fun_upd_apply)
- have loc: "l = l'" using clinit\<^sub>1_loc_pres RInit\<^sub>1.hyps(1) clinit by simp
- have IH: "PROP ?P e h l sh (Val v1) h' l' sh' E C M pc (Called Cs) v1 xa vs (tl ?frs') I" by fact
- then have IH':
- "PROP ?P (C\<^sub>0\<bullet>\<^sub>sclinit([])) h l sh (Val v1) h' l' sh' E C M pc (Called Cs) v1 xa vs (tl ?frs') I"
- using clinit by simp
- have IH2: "PROP ?P (INIT C' (Cs,True) \<leftarrow> e') h' l' sh'' e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 E C M
- pc ics v xa vs frs I" by fact
- have "P \<turnstile> (None,h,?frs,sh) -jvm\<rightarrow> (None,h,create_init_frame P C\<^sub>0 # ?frs',sh)" by(rule jvm_Called)
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> (None,h',?frs',sh'')"
- using IH' Jcc_pieces_RInit_clinit[OF assms(1-2) pcs,of h' l' sh'] RInit\<^sub>1.hyps(3,4) by simp
- finally have jvm1: "P \<turnstile> (None,h,?frs,sh) -jvm\<rightarrow> (None,h',?frs',sh'')" .
- show ?case (is "(?e1 \<longrightarrow> ?jvm1) \<and> (?e2 \<longrightarrow> ?err)")
- proof(rule conjI)
- { assume val: ?e1
- note jvm1
- also have "P \<turnstile> (None,h',?frs',sh'') -jvm\<rightarrow> (None,h\<^sub>1,(vs,l,C,M,pc,Called [])#frs,sh\<^sub>1)"
- using IH2 Jcc_pieces_RInit_Init[OF assms(1-2) shC'' pcs,of h'] RInit\<^sub>1.hyps(5) loc val by auto
- finally have ?jvm1 using pcs clinit by simp
- }
- thus "?e1 \<longrightarrow> ?jvm1" by simp
- next
- { assume throw: ?e2
- note jvm1
- also obtain vs' where "P \<turnstile> (None,h',?frs',sh'')
- -jvm\<rightarrow> handle P C M xa h\<^sub>1 (vs'@vs) l pc ics frs sh\<^sub>1"
- using IH2 Jcc_pieces_RInit_Init[OF assms(1-2) shC'' pcs,of h'] RInit\<^sub>1.hyps(5) loc throw by auto
- finally have ?err using pcs clinit by auto
- }
- thus "?e2 \<longrightarrow> ?err" by simp
- qed
-next
- case (RInitInitFail\<^sub>1 e h l sh a h' l' sh' C\<^sub>0 sfs i sh'' D Cs e' e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1)
- let ?frs = "(vs,l,C,M,pc,Called (C\<^sub>0#D#Cs)) # frs"
- let ?frs' = "(vs,l,C,M,pc,Called (D#Cs)) # frs"
- let "?frsT" = "\<lambda>xa1. (vs,l,C,M,pc,Throwing (C\<^sub>0#D#Cs) xa1) # frs"
- let "?frsT'" = "\<lambda>xa1. (vs,l,C,M,pc,Throwing (D#Cs) xa1) # frs"
- obtain xa' where xa': "throw a = Throw xa'"
- by (metis RInitInitFail\<^sub>1.hyps(1) eval\<^sub>1_final exp.distinct(101) final_def)
- have e\<^sub>1: "e\<^sub>1 = Throw xa'" using xa' rinit\<^sub>1_throw RInitInitFail\<^sub>1.hyps(5) by simp
- show ?case
- proof(cases "e = C\<^sub>0\<bullet>\<^sub>sclinit([])")
- case clinit: True
- then obtain err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h\<^sub>1 l\<^sub>1 sh\<^sub>1 v xa'
- (RI (C\<^sub>0,C\<^sub>0\<bullet>\<^sub>sclinit([])) ; D # Cs \<leftarrow> e')
- = (True, ?frs, (None, h\<^sub>1, (vs, l, C, M, pc, Called []) # frs, sh\<^sub>1), err)"
- using RInitInitFail\<^sub>1.prems(1) by simp
- have loc: "l = l'" using clinit\<^sub>1_loc_pres RInitInitFail\<^sub>1.hyps(1) clinit by simp
- have IH: "PROP ?P e h l sh (throw a) h' l' sh' E C M pc (Called (D#Cs)) v xa' vs frs I"
- by fact
- then have IH':
- "PROP ?P (C\<^sub>0\<bullet>\<^sub>sclinit([])) h l sh (Throw xa') h' l' sh' E C M pc (Called (D#Cs)) v xa' vs
- frs I" using clinit xa' by simp
- have IH2: "PROP ?P (RI (D,throw a) ; Cs \<leftarrow> e') h' l' sh'' e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 E C M
- pc ics v xa' vs frs I" by fact
- have "P \<turnstile> (None,h,?frs,sh) -jvm\<rightarrow> (None,h,create_init_frame P C\<^sub>0 # ?frs',sh)" by(rule jvm_Called)
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> (None,h',(vs, l, C, M, pc, Throwing (D#Cs) xa') # frs,sh'')"
- using IH' Jcc_pieces_RInit_clinit[OF assms(1-2) pcs,of h' l' sh'] RInitInitFail\<^sub>1.hyps(3,4)
- by simp
- also obtain vs'' where "P \<turnstile> \<dots> -jvm\<rightarrow> handle P C M xa' h\<^sub>1 (vs''@vs) l pc ics frs sh\<^sub>1"
- using IH2 pcs Jcc_pieces_RInit_RInit[OF assms(1) pcs] RInitInitFail\<^sub>1.hyps(3,4)
- xa' loc e\<^sub>1 xa' by clarsimp
- finally show ?thesis using pcs e\<^sub>1 clinit by auto
- next
- case throw: False
- then have eT: "e = Throw xa'" "h = h'" "l = l'" "sh = sh'" using xa' RInitInitFail\<^sub>1.prems(1)
- eval\<^sub>1_final_same[OF RInitInitFail\<^sub>1.hyps(1)] by clarsimp+
- obtain a' where "class P\<^sub>1 C\<^sub>0 = \<lfloor>a'\<rfloor>" using RInitInitFail\<^sub>1.prems by(auto simp: is_class_def)
- then obtain stk' loc' M' pc' ics' where "create_init_frame P C\<^sub>0 = (stk',loc',C\<^sub>0,M',pc',ics')"
- using create_init_frame_wf_eq[OF wf] by(cases "create_init_frame P C\<^sub>0", simp)
- then obtain rhs err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' sh'' v xa'
- (RI (C\<^sub>0,e) ; D#Cs \<leftarrow> e') = (True, ?frsT xa', rhs, err)"
- using RInitInitFail\<^sub>1.prems(1) eT by clarsimp
- have IH2: "PROP ?P (RI (D,throw a) ; Cs \<leftarrow> e') h' l' sh'' e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 E C M
- pc ics v xa' vs frs I" by fact
- have "P \<turnstile> (None,h,?frsT xa',sh') -jvm\<rightarrow> (None,h,?frsT' xa',sh'(C\<^sub>0 \<mapsto> (fst (the (sh' C\<^sub>0)), Error)))"
- by(rule jvm_Throwing)
- also obtain vs' where "P \<turnstile> ... -jvm\<rightarrow> handle P C M xa' h\<^sub>1 (vs'@vs) l pc ics frs sh\<^sub>1"
- using IH2 Jcc_pieces_RInit_RInit[OF assms(1) pcs] RInitInitFail\<^sub>1.hyps(3,4)
- eT e\<^sub>1 xa' by clarsimp
- finally show ?thesis using pcs e\<^sub>1 throw eT by auto
- qed
-next
- case (RInitFailFinal\<^sub>1 e h l sh a h' l' sh' C\<^sub>0 sfs i sh'' e'')
- let ?frs = "(vs,l,C,M,pc,Called [C\<^sub>0]) # frs"
- let ?frs' = "(vs,l,C,M,pc,Called []) # frs"
- let "?frsT" = "\<lambda>xa1. (vs,l,C,M,pc,Throwing [C\<^sub>0] xa1) # frs"
- let "?frsT'" = "\<lambda>xa1. (vs,l,C,M,pc,Throwing [] xa1) # frs"
- obtain xa' where xa': "throw a = Throw xa'"
- by (metis RInitFailFinal\<^sub>1.hyps(1) eval\<^sub>1_final exp.distinct(101) final_def)
- show ?case
- proof(cases "e = C\<^sub>0\<bullet>\<^sub>sclinit([])")
- case clinit: True
- then obtain err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' sh'' v xa'
- (RI (C\<^sub>0,C\<^sub>0\<bullet>\<^sub>sclinit([])) ; [] \<leftarrow> unit) = (True, ?frs, (None, h', ?frs', sh''), err)"
- using RInitFailFinal\<^sub>1.prems(1) by clarsimp
- have IH: "PROP ?P e h l sh (throw a) h' l' sh' E C M pc (Called []) v xa' vs frs I" by fact
- then have IH':
- "PROP ?P (C\<^sub>0\<bullet>\<^sub>sclinit([])) h l sh (throw a) h' l' sh' E C M pc (Called []) v xa' vs frs I"
- using clinit by simp
- have "P \<turnstile> (None,h,?frs,sh) -jvm\<rightarrow> (None,h,create_init_frame P C\<^sub>0 # ?frs',sh)"
- by(rule jvm_Called)
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> (None,h',?frsT' xa',sh'')"
- using IH' Jcc_pieces_RInit_clinit[OF assms(1-2) pcs,of h' l' sh'] xa'
- RInitFailFinal\<^sub>1.hyps(3,4) by simp
- also have
- "P \<turnstile> \<dots> -jvm\<rightarrow> handle (compP compMb\<^sub>2 P\<^sub>1) C M xa' h' vs l pc No_ics frs sh''"
- using RInitFailFinal\<^sub>1.hyps(3,4) jvm_RInit_throw[where h=h' and sh=sh''] by simp
- finally show ?thesis using xa' pcs clinit by(clarsimp intro!: exI[where x="[]"])
- next
- case throw: False
- then have eT: "e = Throw xa'" "h = h'" "sh = sh'" using xa' RInitFailFinal\<^sub>1.prems(1)
- eval\<^sub>1_final_same[OF RInitFailFinal\<^sub>1.hyps(1)] by clarsimp+
- obtain a where "class P\<^sub>1 C\<^sub>0 = \<lfloor>a\<rfloor>" using RInitFailFinal\<^sub>1.prems by(auto simp: is_class_def)
- then obtain stk' loc' M' pc' ics' where "create_init_frame P C\<^sub>0 = (stk',loc',C\<^sub>0,M',pc',ics')"
- using create_init_frame_wf_eq[OF wf] by(cases "create_init_frame P C\<^sub>0", simp)
- then obtain rhs err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' sh'' v xa'
- (RI (C\<^sub>0,e) ; [] \<leftarrow> unit) = (True, ?frsT xa', rhs, err)"
- using RInitFailFinal\<^sub>1.prems(1) eT by clarsimp
- have "P \<turnstile> (None,h,?frsT xa',sh) -jvm\<rightarrow> (None,h,?frsT' xa',sh(C\<^sub>0 \<mapsto> (fst (the (sh C\<^sub>0)), Error)))"
- by(rule jvm_Throwing)
- also have "P \<turnstile> \<dots> -jvm\<rightarrow> handle P C M xa' h' vs l pc No_ics frs sh''"
- using RInitFailFinal\<^sub>1.hyps(3,4) jvm_RInit_throw[where h=h and sh=sh''] eT by simp
- finally show ?thesis using pcs xa' by(clarsimp intro!: exI[where x="[]"])
- qed
-qed
-(*>*)
-
-(*FIXME move! *)
-lemma atLeast0AtMost[simp]: "{0::nat..n} = {..n}"
-by auto
-
-lemma atLeast0LessThan[simp]: "{0::nat..<n} = {..<n}"
-by auto
-
-fun exception :: "'a exp \<Rightarrow> addr option" where
- "exception (Throw a) = Some a"
-| "exception e = None"
-
-lemma comp\<^sub>2_correct:
-assumes wf: "wf_J\<^sub>1_prog P\<^sub>1"
- and "method": "P\<^sub>1 \<turnstile> C sees M,b:Ts\<rightarrow>T = body in C"
- and eval: "P\<^sub>1 \<turnstile>\<^sub>1 \<langle>body,(h,ls,sh)\<rangle> \<Rightarrow> \<langle>e',(h',ls',sh')\<rangle>"
- and nclinit: "M \<noteq> clinit"
-shows "compP\<^sub>2 P\<^sub>1 \<turnstile> (None,h,[([],ls,C,M,0,No_ics)],sh) -jvm\<rightarrow> (exception e',h',[],sh')"
-(*<*)
- (is "_ \<turnstile> ?\<sigma>\<^sub>0 -jvm\<rightarrow> ?\<sigma>\<^sub>1")
-proof -
- let ?P = "compP\<^sub>2 P\<^sub>1"
- let ?E = "case b of Static \<Rightarrow> Ts | NonStatic \<Rightarrow> Class C#Ts"
- have nsub: "\<not>sub_RI body" using sees_wf\<^sub>1_nsub_RI[OF wf method] by simp
- have code: "?P,C,M,0 \<rhd> compE\<^sub>2 body" using beforeM[OF "method"] by auto
- have xtab: "?P,C,M \<rhd> compxE\<^sub>2 body 0 (size[])/{..<size(compE\<^sub>2 body)},size[]"
- using beforexM[OF "method"] by auto
- have cond: "Jcc_cond P\<^sub>1 ?E C M [] 0 No_ics {..<size(compE\<^sub>2 body)} h sh body"
- using nsub_RI_Jcc_pieces nsub code xtab by auto
- \<comment> \<open>Distinguish if e' is a value or an exception\<close>
- { fix v assume [simp]: "e' = Val v"
- have "?P \<turnstile> ?\<sigma>\<^sub>0 -jvm\<rightarrow> (None,h',[([v],ls',C,M,size(compE\<^sub>2 body),No_ics)],sh')"
- using Jcc[OF wf eval cond] nsub_RI_Jcc_pieces[OF _ nsub] by auto
- also have "?P \<turnstile> \<dots> -jvm\<rightarrow> ?\<sigma>\<^sub>1" using beforeM[OF "method"] nclinit by auto
- finally have ?thesis .
- }
- moreover
- { fix a assume [simp]: "e' = Throw a"
- obtain pc vs' where pc: "0 \<le> pc \<and> pc < size(compE\<^sub>2 body) \<and>
- \<not> caught ?P pc h' a (compxE\<^sub>2 body 0 0)"
- and 1: "?P \<turnstile> ?\<sigma>\<^sub>0 -jvm\<rightarrow> handle ?P C M a h' vs' ls' pc No_ics [] sh'"
- using Jcc[OF wf eval cond] nsub_RI_Jcc_pieces[OF _ nsub] by auto meson
- from pc have "handle ?P C M a h' vs' ls' pc No_ics [] sh' = ?\<sigma>\<^sub>1" using xtab "method" nclinit
- by(auto simp:handle_def compMb\<^sub>2_def)
- with 1 have ?thesis by simp
- }
- ultimately show ?thesis using eval\<^sub>1_final[OF eval] by(auto simp:final_def)
-qed
-(*>*)
-
-end
+(* Title: JinjaDCI/Compiler/Correctness2.thy
+ Author: Tobias Nipkow, Susannah Mansky
+ Copyright TUM 2003, UIUC 2019-20
+
+ Based on the Jinja theory Compiler/Correctness2.thy by Tobias Nipkow
+*)
+
+section \<open> Correctness of Stage 2 \<close>
+
+theory Correctness2
+imports "HOL-Library.Sublist" Compiler2 J1WellForm "../J/EConform"
+begin
+
+(*<*)hide_const (open) Throw(*>*)
+
+subsection\<open> Instruction sequences \<close>
+
+text\<open> How to select individual instructions and subsequences of
+instructions from a program given the class, method and program
+counter. \<close>
+
+definition before :: "jvm_prog \<Rightarrow> cname \<Rightarrow> mname \<Rightarrow> nat \<Rightarrow> instr list \<Rightarrow> bool"
+ ("(_,_,_,_/ \<rhd> _)" [51,0,0,0,51] 50) where
+ "P,C,M,pc \<rhd> is \<longleftrightarrow> prefix is (drop pc (instrs_of P C M))"
+
+definition at :: "jvm_prog \<Rightarrow> cname \<Rightarrow> mname \<Rightarrow> nat \<Rightarrow> instr \<Rightarrow> bool"
+ ("(_,_,_,_/ \<triangleright> _)" [51,0,0,0,51] 50) where
+ "P,C,M,pc \<triangleright> i \<longleftrightarrow> (\<exists>is. drop pc (instrs_of P C M) = i#is)"
+
+lemma [simp]: "P,C,M,pc \<rhd> []"
+(*<*)by(simp add:before_def)(*>*)
+
+
+lemma [simp]: "P,C,M,pc \<rhd> (i#is) = (P,C,M,pc \<triangleright> i \<and> P,C,M,pc + 1 \<rhd> is)"
+(*<*)by(fastforce simp add:before_def at_def prefix_def drop_Suc drop_tl)(*>*)
+
+(*<*)
+declare drop_drop[simp del]
+(*>*)
+
+
+lemma [simp]: "P,C,M,pc \<rhd> (is\<^sub>1 @ is\<^sub>2) = (P,C,M,pc \<rhd> is\<^sub>1 \<and> P,C,M,pc + size is\<^sub>1 \<rhd> is\<^sub>2)"
+(*<*)
+by(subst add.commute)
+ (fastforce simp add:before_def prefix_def drop_drop[symmetric])
+(*>*)
+
+(*<*)
+declare drop_drop[simp]
+(*>*)
+
+
+lemma [simp]: "P,C,M,pc \<triangleright> i \<Longrightarrow> instrs_of P C M ! pc = i"
+(*<*)by(clarsimp simp add:at_def strict_prefix_def nth_via_drop)(*>*)
+
+lemma beforeM:
+ "P \<turnstile> C sees M,b: Ts\<rightarrow>T = body in D \<Longrightarrow>
+ compP\<^sub>2 P,D,M,0 \<rhd> compE\<^sub>2 body @ [Return]"
+(*<*)by(drule sees_method_idemp) (simp add:before_def compMb\<^sub>2_def)(*>*)
+
+text\<open> This lemma executes a single instruction by rewriting: \<close>
+
+lemma [simp]:
+ "P,C,M,pc \<triangleright> instr \<Longrightarrow>
+ (P \<turnstile> (None, h, (vs,ls,C,M,pc,ics) # frs, sh) -jvm\<rightarrow> \<sigma>') =
+ ((None, h, (vs,ls,C,M,pc,ics) # frs, sh) = \<sigma>' \<or>
+ (\<exists>\<sigma>. exec(P,(None, h, (vs,ls,C,M,pc,ics) # frs, sh)) = Some \<sigma> \<and> P \<turnstile> \<sigma> -jvm\<rightarrow> \<sigma>'))"
+(*<*)
+by(simp only: exec_all_def)
+ (blast intro: converse_rtranclE converse_rtrancl_into_rtrancl)
+(*>*)
+
+
+subsection\<open> Exception tables \<close>
+
+definition pcs :: "ex_table \<Rightarrow> nat set"
+where
+ "pcs xt \<equiv> \<Union>(f,t,C,h,d) \<in> set xt. {f ..< t}"
+
+lemma pcs_subset:
+shows "(\<And>pc d. pcs(compxE\<^sub>2 e pc d) \<subseteq> {pc..<pc+size(compE\<^sub>2 e)})"
+and "(\<And>pc d. pcs(compxEs\<^sub>2 es pc d) \<subseteq> {pc..<pc+size(compEs\<^sub>2 es)})"
+(*<*)
+proof(induct e and es rule: compxE\<^sub>2.induct compxEs\<^sub>2.induct)
+ case Cast then show ?case by (fastforce simp:pcs_def)
+next
+ case BinOp then show ?case by (fastforce simp:pcs_def split:bop.splits)
+next
+ case LAss then show ?case by (fastforce simp: pcs_def)
+next
+ case FAcc then show ?case by (fastforce simp: pcs_def)
+next
+ case FAss then show ?case by (fastforce simp: pcs_def)
+next
+ case SFAss then show ?case by (fastforce simp: pcs_def)
+next
+ case Call then show ?case by (fastforce simp: pcs_def)
+next
+ case SCall then show ?case by (fastforce simp: pcs_def)
+next
+ case Seq then show ?case by (fastforce simp: pcs_def)
+next
+ case Cond then show ?case by (fastforce simp: pcs_def)
+next
+ case While then show ?case by (fastforce simp: pcs_def)
+next
+ case throw then show ?case by (fastforce simp: pcs_def)
+next
+ case TryCatch then show ?case by (fastforce simp: pcs_def)
+next
+ case Cons_exp then show ?case by (fastforce simp: pcs_def)
+qed (simp_all add:pcs_def)
+(*>*)
+
+
+lemma [simp]: "pcs [] = {}"
+(*<*)by(simp add:pcs_def)(*>*)
+
+
+lemma [simp]: "pcs (x#xt) = {fst x ..< fst(snd x)} \<union> pcs xt"
+(*<*)by(auto simp add: pcs_def)(*>*)
+
+
+lemma [simp]: "pcs(xt\<^sub>1 @ xt\<^sub>2) = pcs xt\<^sub>1 \<union> pcs xt\<^sub>2"
+(*<*)by(simp add:pcs_def)(*>*)
+
+
+lemma [simp]: "pc < pc\<^sub>0 \<or> pc\<^sub>0+size(compE\<^sub>2 e) \<le> pc \<Longrightarrow> pc \<notin> pcs(compxE\<^sub>2 e pc\<^sub>0 d)"
+(*<*)using pcs_subset by fastforce(*>*)
+
+
+lemma [simp]: "pc < pc\<^sub>0 \<or> pc\<^sub>0+size(compEs\<^sub>2 es) \<le> pc \<Longrightarrow> pc \<notin> pcs(compxEs\<^sub>2 es pc\<^sub>0 d)"
+(*<*)using pcs_subset by fastforce(*>*)
+
+
+lemma [simp]: "pc\<^sub>1 + size(compE\<^sub>2 e\<^sub>1) \<le> pc\<^sub>2 \<Longrightarrow> pcs(compxE\<^sub>2 e\<^sub>1 pc\<^sub>1 d\<^sub>1) \<inter> pcs(compxE\<^sub>2 e\<^sub>2 pc\<^sub>2 d\<^sub>2) = {}"
+(*<*)using pcs_subset by fastforce(*>*)
+
+
+lemma [simp]: "pc\<^sub>1 + size(compE\<^sub>2 e) \<le> pc\<^sub>2 \<Longrightarrow> pcs(compxE\<^sub>2 e pc\<^sub>1 d\<^sub>1) \<inter> pcs(compxEs\<^sub>2 es pc\<^sub>2 d\<^sub>2) = {}"
+(*<*)using pcs_subset by fastforce(*>*)
+
+
+lemma [simp]:
+ "pc \<notin> pcs xt\<^sub>0 \<Longrightarrow> match_ex_table P C pc (xt\<^sub>0 @ xt\<^sub>1) = match_ex_table P C pc xt\<^sub>1"
+(*<*)by (induct xt\<^sub>0) (auto simp: matches_ex_entry_def)(*>*)
+
+
+lemma [simp]: "\<lbrakk> x \<in> set xt; pc \<notin> pcs xt \<rbrakk> \<Longrightarrow> \<not> matches_ex_entry P D pc x"
+(*<*)by(auto simp:matches_ex_entry_def pcs_def)(*>*)
+
+
+lemma [simp]:
+assumes xe: "xe \<in> set(compxE\<^sub>2 e pc d)" and outside: "pc' < pc \<or> pc+size(compE\<^sub>2 e) \<le> pc'"
+shows "\<not> matches_ex_entry P C pc' xe"
+(*<*)
+proof
+ assume "matches_ex_entry P C pc' xe"
+ with xe have "pc' \<in> pcs(compxE\<^sub>2 e pc d)"
+ by(force simp add:matches_ex_entry_def pcs_def)
+ with outside show False by simp
+qed
+(*>*)
+
+
+lemma [simp]:
+assumes xe: "xe \<in> set(compxEs\<^sub>2 es pc d)" and outside: "pc' < pc \<or> pc+size(compEs\<^sub>2 es) \<le> pc'"
+shows "\<not> matches_ex_entry P C pc' xe"
+(*<*)
+proof
+ assume "matches_ex_entry P C pc' xe"
+ with xe have "pc' \<in> pcs(compxEs\<^sub>2 es pc d)"
+ by(force simp add:matches_ex_entry_def pcs_def)
+ with outside show False by simp
+qed
+(*>*)
+
+
+lemma match_ex_table_app[simp]:
+ "\<forall>xte \<in> set xt\<^sub>1. \<not> matches_ex_entry P D pc xte \<Longrightarrow>
+ match_ex_table P D pc (xt\<^sub>1 @ xt) = match_ex_table P D pc xt"
+(*<*)by(induct xt\<^sub>1) simp_all(*>*)
+
+
+lemma [simp]:
+ "\<forall>x \<in> set xtab. \<not> matches_ex_entry P C pc x \<Longrightarrow>
+ match_ex_table P C pc xtab = None"
+(*<*)using match_ex_table_app[where ?xt = "[]"] by fastforce(*>*)
+
+
+lemma match_ex_entry:
+ "matches_ex_entry P C pc (start, end, catch_type, handler) =
+ (start \<le> pc \<and> pc < end \<and> P \<turnstile> C \<preceq>\<^sup>* catch_type)"
+(*<*)by(simp add:matches_ex_entry_def)(*>*)
+
+
+definition caught :: "jvm_prog \<Rightarrow> pc \<Rightarrow> heap \<Rightarrow> addr \<Rightarrow> ex_table \<Rightarrow> bool" where
+ "caught P pc h a xt \<longleftrightarrow>
+ (\<exists>entry \<in> set xt. matches_ex_entry P (cname_of h a) pc entry)"
+
+definition beforex :: "jvm_prog \<Rightarrow> cname \<Rightarrow> mname \<Rightarrow> ex_table \<Rightarrow> nat set \<Rightarrow> nat \<Rightarrow> bool"
+ ("(2_,/_,/_ \<rhd>/ _ /'/ _,/_)" [51,0,0,0,0,51] 50) where
+ "P,C,M \<rhd> xt / I,d \<longleftrightarrow>
+ (\<exists>xt\<^sub>0 xt\<^sub>1. ex_table_of P C M = xt\<^sub>0 @ xt @ xt\<^sub>1 \<and> pcs xt\<^sub>0 \<inter> I = {} \<and> pcs xt \<subseteq> I \<and>
+ (\<forall>pc \<in> I. \<forall>C pc' d'. match_ex_table P C pc xt\<^sub>1 = \<lfloor>(pc',d')\<rfloor> \<longrightarrow> d' \<le> d))"
+
+definition dummyx :: "jvm_prog \<Rightarrow> cname \<Rightarrow> mname \<Rightarrow> ex_table \<Rightarrow> nat set \<Rightarrow> nat \<Rightarrow> bool" ("(2_,_,_ \<triangleright>/ _ '/_,_)" [51,0,0,0,0,51] 50) where
+ "P,C,M \<triangleright> xt/I,d \<longleftrightarrow> P,C,M \<rhd> xt/I,d"
+
+abbreviation
+"beforex\<^sub>0 P C M d I xt xt\<^sub>0 xt\<^sub>1
+ \<equiv> ex_table_of P C M = xt\<^sub>0 @ xt @ xt\<^sub>1 \<and> pcs xt\<^sub>0 \<inter> I = {}
+ \<and> pcs xt \<subseteq> I \<and> (\<forall>pc \<in> I. \<forall>C pc' d'. match_ex_table P C pc xt\<^sub>1 = \<lfloor>(pc',d')\<rfloor> \<longrightarrow> d' \<le> d)"
+
+lemma beforex_beforex\<^sub>0_eq:
+ "P,C,M \<rhd> xt / I,d \<equiv> \<exists>xt\<^sub>0 xt\<^sub>1. beforex\<^sub>0 P C M d I xt xt\<^sub>0 xt\<^sub>1"
+using beforex_def by auto
+
+lemma beforexD1: "P,C,M \<rhd> xt / I,d \<Longrightarrow> pcs xt \<subseteq> I"
+(*<*)by(auto simp add:beforex_def)(*>*)
+
+
+lemma beforex_mono: "\<lbrakk> P,C,M \<rhd> xt/I,d'; d' \<le> d \<rbrakk> \<Longrightarrow> P,C,M \<rhd> xt/I,d"
+(*<*)by(fastforce simp:beforex_def)(*>*)
+
+
+lemma [simp]: "P,C,M \<rhd> xt/I,d \<Longrightarrow> P,C,M \<rhd> xt/I,Suc d"
+(*<*)by(fastforce intro:beforex_mono)(*>*)
+
+
+lemma beforex_append[simp]:
+ "pcs xt\<^sub>1 \<inter> pcs xt\<^sub>2 = {} \<Longrightarrow>
+ P,C,M \<rhd> xt\<^sub>1 @ xt\<^sub>2/I,d =
+ (P,C,M \<rhd> xt\<^sub>1/I-pcs xt\<^sub>2,d \<and> P,C,M \<rhd> xt\<^sub>2/I-pcs xt\<^sub>1,d \<and> P,C,M \<triangleright> xt\<^sub>1@xt\<^sub>2/I,d)"
+(*<*)(is "?Q \<Longrightarrow> ?P = (?P1 \<and> ?P2 \<and> ?P3)" is "?Q \<Longrightarrow> ?P = ?P123")
+proof -
+ assume pcs: ?Q
+ show ?thesis proof(rule iffI)
+ assume "?P123" then show ?P by(simp add:dummyx_def)
+ next
+ assume hyp: ?P
+ let ?xt = "xt\<^sub>1 @ xt\<^sub>2"
+ let ?beforex = "beforex\<^sub>0 P C M d"
+ obtain xt\<^sub>0 xt\<^sub>1' where beforex: "?beforex I ?xt xt\<^sub>0 xt\<^sub>1'"
+ using hyp by(clarsimp simp: beforex_def)
+ have "\<exists>xt\<^sub>0 xt\<^sub>1'. ?beforex (I - pcs xt\<^sub>2) xt\<^sub>1 xt\<^sub>0 xt\<^sub>1'" \<comment> \<open>?P1\<close>
+ using pcs beforex by(rule_tac x=xt\<^sub>0 in exI) auto
+ moreover have "\<exists>xt\<^sub>0 xt\<^sub>1'. ?beforex (I - pcs xt\<^sub>1) xt\<^sub>2 xt\<^sub>0 xt\<^sub>1'" \<comment> \<open>?P2\<close>
+ using pcs beforex by(rule_tac x="xt\<^sub>0@xt\<^sub>1" in exI) auto
+ moreover have ?P3 using hyp by(simp add: dummyx_def)
+ ultimately show ?P123 by (simp add: beforex_def)
+ qed
+qed
+(*>*)
+
+
+lemma beforex_appendD1:
+assumes bx: "P,C,M \<rhd> xt\<^sub>1 @ xt\<^sub>2 @ [(f,t,D,h,d)] / I,d"
+ and pcs: "pcs xt\<^sub>1 \<subseteq> J" and JI: "J \<subseteq> I" and Jpcs: "J \<inter> pcs xt\<^sub>2 = {}"
+shows "P,C,M \<rhd> xt\<^sub>1 / J,d"
+(*<*)
+proof -
+ let ?beforex = "beforex\<^sub>0 P C M d"
+ obtain xt\<^sub>0 xt\<^sub>1' where bx': "?beforex I (xt\<^sub>1 @ xt\<^sub>2 @ [(f,t,D,h,d)]) xt\<^sub>0 xt\<^sub>1'"
+ using bx by(clarsimp simp:beforex_def)
+ let ?xt0 = xt\<^sub>0 and ?xt1 = "xt\<^sub>2 @ (f, t, D, h, d) # xt\<^sub>1'"
+ have "pcs xt\<^sub>0 \<inter> J = {}" using bx' JI by blast
+ moreover {
+ fix pc C pc' d' assume pcJ: "pc\<in>J"
+ then have "pc \<notin> pcs xt\<^sub>2" using bx' Jpcs by blast
+ then have "match_ex_table P C pc (xt\<^sub>2 @ (f, t, D, h, d) # xt\<^sub>1')
+ = \<lfloor>(pc', d')\<rfloor> \<longrightarrow> d' \<le> d"
+ using bx' JI pcJ by (auto split:if_split_asm)
+ }
+ ultimately have "?beforex J xt\<^sub>1 ?xt0 ?xt1" using bx' pcs by simp
+ then show ?thesis using beforex_def by blast
+qed
+(*>*)
+
+
+lemma beforex_appendD2:
+assumes bx: "P,C,M \<rhd> xt\<^sub>1 @ xt\<^sub>2 @ [(f,t,D,h,d)] / I,d"
+ and pcs: "pcs xt\<^sub>2 \<subseteq> J" and JI: "J \<subseteq> I" and Jpcs: "J \<inter> pcs xt\<^sub>1 = {}"
+shows "P,C,M \<rhd> xt\<^sub>2 / J,d"
+(*<*)
+proof -
+ let ?beforex = "beforex\<^sub>0 P C M d"
+ obtain xt\<^sub>0 xt\<^sub>1' where bx': "?beforex I (xt\<^sub>1 @ xt\<^sub>2 @ [(f,t,D,h,d)]) xt\<^sub>0 xt\<^sub>1'"
+ using bx by(clarsimp simp:beforex_def)
+ then have "\<exists>xt\<^sub>1''. beforex\<^sub>0 P C M d J xt\<^sub>2 (xt\<^sub>0 @ xt\<^sub>1) xt\<^sub>1''"
+ using assms by fastforce
+ then show ?thesis using beforex_def by blast
+qed
+(*>*)
+
+
+lemma beforexM:
+ "P \<turnstile> C sees M,b: Ts\<rightarrow>T = body in D \<Longrightarrow> compP\<^sub>2 P,D,M \<rhd> compxE\<^sub>2 body 0 0/{..<size(compE\<^sub>2 body)},0"
+(*<*)
+proof -
+ assume cM: "P \<turnstile> C sees M,b: Ts\<rightarrow>T = body in D"
+ let ?xt0 = "[]"
+ have "\<exists>xt1. beforex\<^sub>0 (compP\<^sub>2 P) D M 0 ({..<size(compE\<^sub>2 body)}) (compxE\<^sub>2 body 0 0) ?xt0 xt1"
+ using sees_method_compP[where f = compMb\<^sub>2, OF sees_method_idemp[OF cM]]
+ pcs_subset by(fastforce simp add: compP\<^sub>2_def compMb\<^sub>2_def)
+ then show ?thesis using beforex_def by blast
+qed
+(*>*)
+
+
+lemma match_ex_table_SomeD2:
+assumes met: "match_ex_table P D pc (ex_table_of P C M) = \<lfloor>(pc',d')\<rfloor>"
+ and bx: "P,C,M \<rhd> xt/I,d"
+ and nmet: "\<forall>x \<in> set xt. \<not> matches_ex_entry P D pc x" and pcI: "pc \<in> I"
+shows "d' \<le> d"
+(*<*)
+proof -
+ obtain xt\<^sub>0 xt\<^sub>1 where bx': "beforex\<^sub>0 P C M d I xt xt\<^sub>0 xt\<^sub>1"
+ using bx by(clarsimp simp:beforex_def)
+ then have "pc \<notin> pcs xt\<^sub>0" using pcI by blast
+ then show ?thesis using bx' met nmet pcI by simp
+qed
+(*>*)
+
+
+lemma match_ex_table_SomeD1:
+ "\<lbrakk> match_ex_table P D pc (ex_table_of P C M) = \<lfloor>(pc',d')\<rfloor>;
+ P,C,M \<rhd> xt / I,d; pc \<in> I; pc \<notin> pcs xt \<rbrakk> \<Longrightarrow> d' \<le> d"
+(*<*)by(auto elim: match_ex_table_SomeD2)(*>*)
+
+
+subsection\<open> The correctness proof \<close>
+
+(*<*)
+declare nat_add_distrib[simp] caught_def[simp]
+declare fun_upd_apply[simp del]
+(*>*)
+
+definition
+ handle :: "jvm_prog \<Rightarrow> cname \<Rightarrow> mname \<Rightarrow> addr \<Rightarrow> heap \<Rightarrow> val list \<Rightarrow> val list \<Rightarrow> nat \<Rightarrow> init_call_status \<Rightarrow> frame list \<Rightarrow> sheap
+ \<Rightarrow> jvm_state" where
+ "handle P C M a h vs ls pc ics frs sh = find_handler P a h ((vs,ls,C,M,pc,ics) # frs) sh"
+
+lemma aux_isin[simp]: "\<lbrakk> B \<subseteq> A; a \<in> B \<rbrakk> \<Longrightarrow> a \<in> A"
+(*<*)by blast(*>*)
+
+lemma handle_frs_tl_neq:
+ "ics_of f \<noteq> No_ics
+ \<Longrightarrow> (xp, h, f#frs, sh) \<noteq> handle P C M xa h' vs l pc ics frs sh'"
+ by(simp add: handle_def find_handler_frs_tl_neq del: find_handler.simps)
+
+subsubsection "Correctness proof inductive hypothesis"
+
+\<comment> \<open> frame definitions for use by correctness proof inductive hypothesis \<close>
+fun calling_to_called :: "frame \<Rightarrow> frame" where
+"calling_to_called (stk,loc,D,M,pc,ics) = (stk,loc,D,M,pc,case ics of Calling C Cs \<Rightarrow> Called (C#Cs))"
+
+fun calling_to_scalled :: "frame \<Rightarrow> frame" where
+"calling_to_scalled (stk,loc,D,M,pc,ics) = (stk,loc,D,M,pc,case ics of Calling C Cs \<Rightarrow> Called Cs)"
+
+fun calling_to_calling :: "frame \<Rightarrow> cname \<Rightarrow> frame" where
+"calling_to_calling (stk,loc,D,M,pc,ics) C' = (stk,loc,D,M,pc,case ics of Calling C Cs \<Rightarrow> Calling C' (C#Cs))"
+
+fun calling_to_throwing :: "frame \<Rightarrow> addr \<Rightarrow> frame" where
+"calling_to_throwing (stk,loc,D,M,pc,ics) a = (stk,loc,D,M,pc,case ics of Calling C Cs \<Rightarrow> Throwing (C#Cs) a)"
+
+fun calling_to_sthrowing :: "frame \<Rightarrow> addr \<Rightarrow> frame" where
+"calling_to_sthrowing (stk,loc,D,M,pc,ics) a = (stk,loc,D,M,pc,case ics of Calling C Cs \<Rightarrow> Throwing Cs a)"
+
+
+\<comment> \<open> pieces of the correctness proof's inductive hypothesis, which depend on the
+ expression being compiled) \<close>
+
+fun Jcc_cond :: "J\<^sub>1_prog \<Rightarrow> ty list \<Rightarrow> cname \<Rightarrow> mname \<Rightarrow> val list \<Rightarrow> pc \<Rightarrow> init_call_status
+ \<Rightarrow> nat set \<Rightarrow> heap \<Rightarrow> sheap \<Rightarrow> expr\<^sub>1 \<Rightarrow> bool" where
+"Jcc_cond P E C M vs pc ics I h sh (INIT C\<^sub>0 (Cs,b) \<leftarrow> e')
+ = ((\<exists>T. P,E,h,sh \<turnstile>\<^sub>1 INIT C\<^sub>0 (Cs,b) \<leftarrow> e' : T) \<and> unit = e' \<and> ics = No_ics)" |
+"Jcc_cond P E C M vs pc ics I h sh (RI(C',e\<^sub>0);Cs \<leftarrow> e')
+ = (((e\<^sub>0 = C'\<bullet>\<^sub>sclinit([]) \<and> (\<exists>T. P,E,h,sh \<turnstile>\<^sub>1 RI(C',e\<^sub>0);Cs \<leftarrow> e':T))
+ \<or> ((\<exists>a. e\<^sub>0 = Throw a) \<and> (\<forall>C \<in> set(C'#Cs). is_class P C)))
+ \<and> unit = e' \<and> ics = No_ics)" |
+"Jcc_cond P E C M vs pc ics I h sh (C'\<bullet>\<^sub>sM'(es))
+ = (let e = (C'\<bullet>\<^sub>sM'(es))
+ in if M' = clinit \<and> es = [] then (\<exists>T. P,E,h,sh \<turnstile>\<^sub>1 e:T) \<and> (\<exists>Cs. ics = Called Cs)
+ else (compP\<^sub>2 P,C,M,pc \<rhd> compE\<^sub>2 e \<and> compP\<^sub>2 P,C,M \<rhd> compxE\<^sub>2 e pc (size vs)/I,size vs
+ \<and> {pc..<pc+size(compE\<^sub>2 e)} \<subseteq> I \<and> \<not>sub_RI e \<and> ics = No_ics)
+ )" |
+"Jcc_cond P E C M vs pc ics I h sh e
+ = (compP\<^sub>2 P,C,M,pc \<rhd> compE\<^sub>2 e \<and> compP\<^sub>2 P,C,M \<rhd> compxE\<^sub>2 e pc (size vs)/I,size vs
+ \<and> {pc..<pc+size(compE\<^sub>2 e)} \<subseteq> I \<and> \<not>sub_RI e \<and> ics = No_ics)"
+
+
+fun Jcc_frames :: "jvm_prog \<Rightarrow> cname \<Rightarrow> mname \<Rightarrow> val list \<Rightarrow> val list \<Rightarrow> pc \<Rightarrow> init_call_status
+ \<Rightarrow> frame list \<Rightarrow> expr\<^sub>1 \<Rightarrow> frame list" where
+"Jcc_frames P C M vs ls pc ics frs (INIT C\<^sub>0 (C'#Cs,b) \<leftarrow> e')
+ = (case b of False \<Rightarrow> (vs,ls,C,M,pc,Calling C' Cs) # frs
+ | True \<Rightarrow> (vs,ls,C,M,pc,Called (C'#Cs)) # frs
+ )" |
+"Jcc_frames P C M vs ls pc ics frs (INIT C\<^sub>0 (Nil,b) \<leftarrow> e')
+ = (vs,ls,C,M,pc,Called [])#frs" |
+"Jcc_frames P C M vs ls pc ics frs (RI(C',e\<^sub>0);Cs \<leftarrow> e')
+ = (case e\<^sub>0 of Throw a \<Rightarrow> (vs,ls,C,M,pc,Throwing (C'#Cs) a) # frs
+ | _ \<Rightarrow> (vs,ls,C,M,pc,Called (C'#Cs)) # frs )" |
+"Jcc_frames P C M vs ls pc ics frs (C'\<bullet>\<^sub>sM'(es))
+ = (if M' = clinit \<and> es = []
+ then create_init_frame P C'#(vs,ls,C,M,pc,ics)#frs
+ else (vs,ls,C,M,pc,ics)#frs
+ )" |
+"Jcc_frames P C M vs ls pc ics frs e
+ = (vs,ls,C,M,pc,ics)#frs"
+
+fun Jcc_rhs :: "J\<^sub>1_prog \<Rightarrow> ty list \<Rightarrow> cname \<Rightarrow> mname \<Rightarrow> val list \<Rightarrow> val list \<Rightarrow> pc \<Rightarrow> init_call_status
+ \<Rightarrow> frame list \<Rightarrow> heap \<Rightarrow> val list \<Rightarrow> sheap \<Rightarrow> val \<Rightarrow> expr\<^sub>1 \<Rightarrow> jvm_state" where
+"Jcc_rhs P E C M vs ls pc ics frs h' ls' sh' v (INIT C\<^sub>0 (Cs,b) \<leftarrow> e')
+ = (None,h',(vs,ls,C,M,pc,Called [])#frs,sh')" |
+"Jcc_rhs P E C M vs ls pc ics frs h' ls' sh' v (RI(C',e\<^sub>0);Cs \<leftarrow> e')
+ = (None,h',(vs,ls,C,M,pc,Called [])#frs,sh')" |
+"Jcc_rhs P E C M vs ls pc ics frs h' ls' sh' v (C'\<bullet>\<^sub>sM'(es))
+ = (let e = (C'\<bullet>\<^sub>sM'(es))
+ in if M' = clinit \<and> es = []
+ then (None,h',(vs,ls,C,M,pc,ics)#frs,sh'(C'\<mapsto>(fst(the(sh' C')),Done)))
+ else (None,h',(v#vs,ls',C,M,pc+size(compE\<^sub>2 e),ics)#frs,sh')
+ )" |
+"Jcc_rhs P E C M vs ls pc ics frs h' ls' sh' v e
+ = (None,h',(v#vs,ls',C,M,pc+size(compE\<^sub>2 e),ics)#frs,sh')"
+
+fun Jcc_err :: "jvm_prog \<Rightarrow> cname \<Rightarrow> mname \<Rightarrow> heap \<Rightarrow> val list \<Rightarrow> val list \<Rightarrow> pc \<Rightarrow> init_call_status
+ \<Rightarrow> frame list \<Rightarrow> sheap \<Rightarrow> nat set \<Rightarrow> heap \<Rightarrow> val list \<Rightarrow> sheap \<Rightarrow> addr \<Rightarrow> expr\<^sub>1
+ \<Rightarrow> bool" where
+"Jcc_err P C M h vs ls pc ics frs sh I h' ls' sh' xa (INIT C\<^sub>0 (Cs,b) \<leftarrow> e')
+ = (\<exists>vs'. P \<turnstile> (None,h,Jcc_frames P C M vs ls pc ics frs (INIT C\<^sub>0 (Cs,b) \<leftarrow> e'),sh)
+ -jvm\<rightarrow> handle P C M xa h' (vs'@vs) ls pc ics frs sh')" |
+"Jcc_err P C M h vs ls pc ics frs sh I h' ls' sh' xa (RI(C',e\<^sub>0);Cs \<leftarrow> e')
+ = (\<exists>vs'. P \<turnstile> (None,h,Jcc_frames P C M vs ls pc ics frs (RI(C',e\<^sub>0);Cs \<leftarrow> e'),sh)
+ -jvm\<rightarrow> handle P C M xa h' (vs'@vs) ls pc ics frs sh')" |
+"Jcc_err P C M h vs ls pc ics frs sh I h' ls' sh' xa (C'\<bullet>\<^sub>sM'(es))
+ = (let e = (C'\<bullet>\<^sub>sM'(es))
+ in if M' = clinit \<and> es = []
+ then case ics of
+ Called Cs \<Rightarrow> P \<turnstile> (None,h,Jcc_frames P C M vs ls pc ics frs e,sh)
+ -jvm\<rightarrow> (None,h',(vs,ls,C,M,pc,Throwing Cs xa)#frs,(sh'(C' \<mapsto> (fst(the(sh' C')),Error))))
+ else (\<exists>pc\<^sub>1. pc \<le> pc\<^sub>1 \<and> pc\<^sub>1 < pc + size(compE\<^sub>2 e) \<and>
+ \<not> caught P pc\<^sub>1 h' xa (compxE\<^sub>2 e pc (size vs)) \<and>
+ (\<exists>vs'. P \<turnstile> (None,h,Jcc_frames P C M vs ls pc ics frs e,sh)
+ -jvm\<rightarrow> handle P C M xa h' (vs'@vs) ls' pc\<^sub>1 ics frs sh'))
+ )" |
+"Jcc_err P C M h vs ls pc ics frs sh I h' ls' sh' xa e
+ = (\<exists>pc\<^sub>1. pc \<le> pc\<^sub>1 \<and> pc\<^sub>1 < pc + size(compE\<^sub>2 e) \<and>
+ \<not> caught P pc\<^sub>1 h' xa (compxE\<^sub>2 e pc (size vs)) \<and>
+ (\<exists>vs'. P \<turnstile> (None,h,Jcc_frames P C M vs ls pc ics frs e,sh)
+ -jvm\<rightarrow> handle P C M xa h' (vs'@vs) ls' pc\<^sub>1 ics frs sh'))"
+
+fun Jcc_pieces :: "J\<^sub>1_prog \<Rightarrow> ty list \<Rightarrow> cname \<Rightarrow> mname \<Rightarrow> heap \<Rightarrow> val list \<Rightarrow> val list \<Rightarrow> pc \<Rightarrow> init_call_status
+ \<Rightarrow> frame list \<Rightarrow> sheap \<Rightarrow> nat set \<Rightarrow> heap \<Rightarrow> val list \<Rightarrow> sheap \<Rightarrow> val \<Rightarrow> addr \<Rightarrow> expr\<^sub>1
+ \<Rightarrow> bool \<times> frame list \<times> jvm_state \<times> bool" where
+"Jcc_pieces P E C M h vs ls pc ics frs sh I h' ls' sh' v xa e
+ = (Jcc_cond P E C M vs pc ics I h sh e, Jcc_frames (compP\<^sub>2 P) C M vs ls pc ics frs e,
+ Jcc_rhs P E C M vs ls pc ics frs h' ls' sh' v e,
+ Jcc_err (compP\<^sub>2 P) C M h vs ls pc ics frs sh I h' ls' sh' xa e)"
+
+\<comment> \<open> @{text Jcc_pieces} lemmas \<close>
+
+lemma nsub_RI_Jcc_pieces:
+assumes [simp]: "P \<equiv> compP\<^sub>2 P\<^sub>1"
+ and nsub: "\<not>sub_RI e"
+shows "Jcc_pieces P\<^sub>1 E C M h vs ls pc ics frs sh I h' ls' sh' v xa e
+ = (let cond = P,C,M,pc \<rhd> compE\<^sub>2 e \<and> P,C,M \<rhd> compxE\<^sub>2 e pc (size vs)/I,size vs
+ \<and> {pc..<pc+size(compE\<^sub>2 e)} \<subseteq> I \<and> ics = No_ics;
+ frs' = (vs,ls,C,M,pc,ics)#frs;
+ rhs = (None,h',(v#vs,ls',C,M,pc+size(compE\<^sub>2 e),ics)#frs,sh');
+ err = (\<exists>pc\<^sub>1. pc \<le> pc\<^sub>1 \<and> pc\<^sub>1 < pc + size(compE\<^sub>2 e) \<and>
+ \<not> caught P pc\<^sub>1 h' xa (compxE\<^sub>2 e pc (size vs)) \<and>
+ (\<exists>vs'. P \<turnstile> (None,h,frs',sh) -jvm\<rightarrow> handle P C M xa h' (vs'@vs) ls' pc\<^sub>1 ics frs sh'))
+ in (cond, frs',rhs, err)
+ )"
+proof -
+ have NC: "\<forall>C'. e \<noteq> C'\<bullet>\<^sub>sclinit([])" using assms(2) proof(cases e) qed(simp_all)
+ then show ?thesis using assms
+ proof(cases e)
+ case (SCall C M es)
+ then have "M \<noteq> clinit" using nsub by simp
+ then show ?thesis using SCall nsub proof(cases es) qed(simp_all)
+ qed(simp_all)
+qed
+
+lemma Jcc_pieces_Cast:
+assumes [simp]: "P \<equiv> compP\<^sub>2 P\<^sub>1"
+ and "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (Cast C' e)
+ = (True, frs\<^sub>0, (xp',h',(v#vs',ls',C\<^sub>0,M',pc',ics')#frs',sh'), err)"
+shows "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v' xa e
+ = (True, frs\<^sub>0, (xp',h',(v'#vs',ls',C\<^sub>0,M',pc' - 1,ics')#frs',sh'),
+ (\<exists>pc\<^sub>1. pc \<le> pc\<^sub>1 \<and> pc\<^sub>1 < pc + size(compE\<^sub>2 e) \<and>
+ \<not> caught P pc\<^sub>1 h\<^sub>1 xa (compxE\<^sub>2 e pc (size vs)) \<and>
+ (\<exists>vs'. P \<turnstile> (None,h\<^sub>0,frs\<^sub>0,sh\<^sub>0) -jvm\<rightarrow> handle P C M xa h\<^sub>1 (vs'@vs) ls\<^sub>1 pc\<^sub>1 ics frs sh\<^sub>1)))"
+proof -
+ have pc: "{pc..<pc + length (compE\<^sub>2 e)} \<subseteq> I" using assms by clarsimp
+ show ?thesis using assms nsub_RI_Jcc_pieces[where e=e] pc by clarsimp
+qed
+
+lemma Jcc_pieces_BinOp1:
+assumes
+ "Jcc_pieces P E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v xa (e \<guillemotleft>bop\<guillemotright> e')
+ = (True, frs\<^sub>0, (xp',h',(v#vs',ls',C\<^sub>0,M',pc',ics')#frs',sh'), err)"
+shows "\<exists>err. Jcc_pieces P E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0
+ (I - pcs (compxE\<^sub>2 e' (pc + length (compE\<^sub>2 e)) (Suc (length vs')))) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v' xa e
+ = (True, frs\<^sub>0, (xp',h\<^sub>1,(v'#vs',ls\<^sub>1,C\<^sub>0,M',pc' - size (compE\<^sub>2 e') - 1,ics')#frs',sh\<^sub>1), err)"
+proof -
+ have bef: "compP compMb\<^sub>2 P,C\<^sub>0,M' \<rhd> compxE\<^sub>2 e pc (length vs)
+ / I - pcs (compxE\<^sub>2 e' (pc + length (compE\<^sub>2 e)) (Suc (length vs'))),length vs"
+ using assms by clarsimp
+ have vs: "vs = vs'" using assms by simp
+ show ?thesis using assms nsub_RI_Jcc_pieces[where e=e] bef vs by clarsimp
+qed
+
+lemma Jcc_pieces_BinOp2:
+assumes [simp]: "P \<equiv> compP\<^sub>2 P\<^sub>1"
+ and "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v xa (e \<guillemotleft>bop\<guillemotright> e')
+ = (True, frs\<^sub>0, (xp',h',(v#vs',ls',C\<^sub>0,M',pc',ics')#frs',sh'), err)"
+shows "\<exists>err. Jcc_pieces P\<^sub>1 E C M h\<^sub>1 (v\<^sub>1#vs) ls\<^sub>1 (pc + size (compE\<^sub>2 e)) ics frs sh\<^sub>1
+ (I - pcs (compxE\<^sub>2 e pc (length vs'))) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v' xa e'
+ = (True, (v\<^sub>1#vs,ls\<^sub>1,C,M,pc + size (compE\<^sub>2 e),ics)#frs,
+ (xp',h',(v'#v\<^sub>1#vs',ls',C\<^sub>0,M',pc' - 1,ics')#frs',sh'),
+ (\<exists>pc\<^sub>1. pc + size (compE\<^sub>2 e) \<le> pc\<^sub>1 \<and> pc\<^sub>1 < pc + size (compE\<^sub>2 e) + length (compE\<^sub>2 e') \<and>
+ \<not> caught P pc\<^sub>1 h\<^sub>2 xa (compxE\<^sub>2 e' (pc + size (compE\<^sub>2 e)) (Suc (length vs))) \<and>
+ (\<exists>vs'. P \<turnstile> (None,h\<^sub>1,(v\<^sub>1#vs,ls\<^sub>1,C,M,pc + size (compE\<^sub>2 e),ics)#frs,sh\<^sub>1)
+ -jvm\<rightarrow> handle P C M xa h\<^sub>2 (vs'@v\<^sub>1#vs) ls\<^sub>2 pc\<^sub>1 ics frs sh\<^sub>2)))"
+proof -
+ have bef: "compP compMb\<^sub>2 P\<^sub>1,C\<^sub>0,M' \<rhd> compxE\<^sub>2 e pc (length vs)
+ / I - pcs (compxE\<^sub>2 e' (pc + length (compE\<^sub>2 e)) (Suc (length vs'))),length vs"
+ using assms by clarsimp
+ have vs: "vs = vs'" using assms by simp
+ show ?thesis using assms nsub_RI_Jcc_pieces[where e=e'] bef vs by clarsimp
+qed
+
+lemma Jcc_pieces_FAcc:
+assumes
+ "Jcc_pieces P E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (e\<bullet>F{D})
+ = (True, frs\<^sub>0, (xp',h',(v#vs',ls',C\<^sub>0,M',pc',ics')#frs',sh'), err)"
+shows "\<exists>err. Jcc_pieces P E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v' xa e
+ = (True, frs\<^sub>0, (xp',h',(v'#vs',ls',C\<^sub>0,M',pc' - 1,ics')#frs',sh'), err)"
+proof -
+ have pc: "{pc..<pc + length (compE\<^sub>2 e)} \<subseteq> I" using assms by clarsimp
+ then show ?thesis using assms nsub_RI_Jcc_pieces[where e=e] by clarsimp
+qed
+
+lemma Jcc_pieces_LAss:
+assumes [simp]: "P \<equiv> compP\<^sub>2 P\<^sub>1"
+ and "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (i:=e)
+ = (True, frs\<^sub>0, (xp',h',(v#vs',ls',C\<^sub>0,M',pc',ics')#frs',sh'), err)"
+shows "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v' xa e
+ = (True, frs\<^sub>0, (xp',h',(v'#vs',ls',C\<^sub>0,M',pc' - 2,ics')#frs',sh'),
+ (\<exists>pc\<^sub>1. pc \<le> pc\<^sub>1 \<and> pc\<^sub>1 < pc + size(compE\<^sub>2 e) \<and>
+ \<not> caught P pc\<^sub>1 h\<^sub>1 xa (compxE\<^sub>2 e pc (size vs)) \<and>
+ (\<exists>vs'. P \<turnstile> (None,h\<^sub>0,frs\<^sub>0,sh\<^sub>0) -jvm\<rightarrow> handle P C M xa h\<^sub>1 (vs'@vs) ls\<^sub>1 pc\<^sub>1 ics frs sh\<^sub>1)))"
+proof -
+ have pc: "{pc..<pc + length (compE\<^sub>2 e)} \<subseteq> I" using assms by clarsimp
+ show ?thesis using assms nsub_RI_Jcc_pieces[where e=e] pc by clarsimp
+qed
+
+lemma Jcc_pieces_FAss1:
+assumes
+ "Jcc_pieces P E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v xa (e\<bullet>F{D}:=e')
+ = (True, frs\<^sub>0, (xp',h',(v#vs',ls',C\<^sub>0,M',pc',ics')#frs',sh'), err)"
+shows "\<exists>err. Jcc_pieces P E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0
+ (I - pcs (compxE\<^sub>2 e' (pc + length (compE\<^sub>2 e)) (Suc (length vs')))) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v' xa e
+ = (True, frs\<^sub>0, (xp',h\<^sub>1,(v'#vs',ls\<^sub>1,C\<^sub>0,M',pc' - size (compE\<^sub>2 e') - 2,ics')#frs',sh\<^sub>1), err)"
+proof -
+ show ?thesis using assms nsub_RI_Jcc_pieces[where e=e] by clarsimp
+qed
+
+lemma Jcc_pieces_FAss2:
+assumes
+ "Jcc_pieces P E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v xa (e\<bullet>F{D}:=e')
+ = (True, frs\<^sub>0, (xp',h',(v#vs',ls',C\<^sub>0,M',pc',ics')#frs',sh'), err)"
+shows "Jcc_pieces P E C M h\<^sub>1 (v\<^sub>1#vs) ls\<^sub>1 (pc + size (compE\<^sub>2 e)) ics frs sh\<^sub>1
+ (I - pcs (compxE\<^sub>2 e pc (length vs'))) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v' xa e'
+ = (True, (v\<^sub>1#vs,ls\<^sub>1,C,M,pc + size (compE\<^sub>2 e),ics)#frs,
+ (xp',h',(v'#v\<^sub>1#vs',ls',C\<^sub>0,M',pc' - 2,ics')#frs',sh'),
+ (\<exists>pc\<^sub>1. (pc + size (compE\<^sub>2 e)) \<le> pc\<^sub>1 \<and> pc\<^sub>1 < pc + size (compE\<^sub>2 e) + size(compE\<^sub>2 e') \<and>
+ \<not> caught (compP\<^sub>2 P) pc\<^sub>1 h\<^sub>2 xa (compxE\<^sub>2 e' (pc + size (compE\<^sub>2 e)) (size (v\<^sub>1#vs))) \<and>
+ (\<exists>vs'. (compP\<^sub>2 P) \<turnstile> (None,h\<^sub>1,(v\<^sub>1#vs,ls\<^sub>1,C,M,pc + size (compE\<^sub>2 e),ics)#frs,sh\<^sub>1)
+ -jvm\<rightarrow> handle (compP\<^sub>2 P) C M xa h\<^sub>2 (vs'@v\<^sub>1#vs) ls\<^sub>2 pc\<^sub>1 ics frs sh\<^sub>2)))"
+proof -
+ show ?thesis using assms nsub_RI_Jcc_pieces[where e=e'] by clarsimp
+qed
+
+lemma Jcc_pieces_SFAss:
+assumes
+ "Jcc_pieces P E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h' ls' sh' v xa (C'\<bullet>\<^sub>sF{D}:=e)
+ = (True, frs\<^sub>0, (xp',h',(v#vs',ls',C\<^sub>0,M',pc',ics')#frs',sh'), err)"
+shows "\<exists>err. Jcc_pieces P E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v' xa e
+ = (True, frs\<^sub>0, (xp',h\<^sub>1,(v'#vs',ls\<^sub>1,C\<^sub>0,M',pc' - 2,ics')#frs',sh\<^sub>1), err)"
+proof -
+ have pc: "{pc..<pc + length (compE\<^sub>2 e)} \<subseteq> I" using assms by clarsimp
+ show ?thesis using assms nsub_RI_Jcc_pieces[where e=e] pc by clarsimp
+qed
+
+lemma Jcc_pieces_Call1:
+assumes
+ "Jcc_pieces P E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>3 ls\<^sub>3 sh\<^sub>3 v xa (e\<bullet>M\<^sub>0(es))
+ = (True, frs\<^sub>0, (xp',h',(v#vs',ls',C',M',pc',ics')#frs',sh'), err)"
+shows "\<exists>err. Jcc_pieces P E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0
+ (I - pcs (compxEs\<^sub>2 es (pc + length (compE\<^sub>2 e)) (Suc (length vs')))) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v' xa e
+ = (True, frs\<^sub>0,
+ (xp',h\<^sub>1,(v'#vs',ls\<^sub>1,C',M',pc' - size (compEs\<^sub>2 es) - 1,ics')#frs',sh\<^sub>1), err)"
+proof -
+ show ?thesis using assms nsub_RI_Jcc_pieces[where e=e] by clarsimp
+qed
+
+lemma Jcc_pieces_clinit:
+assumes [simp]: "P \<equiv> compP\<^sub>2 P\<^sub>1"
+ and cond: "Jcc_cond P\<^sub>1 E C M vs pc ics I h sh (C1\<bullet>\<^sub>sclinit([]))"
+shows "Jcc_pieces P\<^sub>1 E C M h vs ls pc ics frs sh I h' ls' sh' v xa (C1\<bullet>\<^sub>sclinit([]))
+ = (True, create_init_frame P C1 # (vs,ls,C,M,pc,ics)#frs,
+ (None, h', (vs,ls,C,M,pc,ics)#frs, sh'(C1\<mapsto>(fst(the(sh' C1)),Done))),
+ P \<turnstile> (None,h,create_init_frame P C1 # (vs,ls,C,M,pc,ics)#frs,sh) -jvm\<rightarrow>
+ (case ics of Called Cs \<Rightarrow> (None,h',(vs,ls,C,M,pc,Throwing Cs xa)#frs,(sh'(C1 \<mapsto> (fst(the(sh' C1)),Error))))))"
+using assms by(auto split: init_call_status.splits list.splits bool.splits)
+
+lemma Jcc_pieces_SCall_clinit_body:
+assumes [simp]: "P \<equiv> compP\<^sub>2 P\<^sub>1" and wf: "wf_J\<^sub>1_prog P\<^sub>1"
+ and "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>3 ls\<^sub>2 sh\<^sub>3 v xa (C1\<bullet>\<^sub>sclinit([]))
+ = (True, frs', rhs', err')"
+ and method: "P\<^sub>1 \<turnstile> C1 sees clinit,Static: []\<rightarrow>Void = body in D"
+shows "Jcc_pieces P\<^sub>1 [] D clinit h\<^sub>2 [] (replicate (max_vars body) undefined) 0
+ No_ics (tl frs') sh\<^sub>2 {..<length (compE\<^sub>2 body)} h\<^sub>3 ls\<^sub>3 sh\<^sub>3 v xa body
+ = (True, frs',
+ (None,h\<^sub>3,([v],ls\<^sub>3,D,clinit,size(compE\<^sub>2 body), No_ics)#tl frs',sh\<^sub>3),
+ \<exists>pc\<^sub>1. 0 \<le> pc\<^sub>1 \<and> pc\<^sub>1 < size(compE\<^sub>2 body) \<and>
+ \<not> caught P pc\<^sub>1 h\<^sub>3 xa (compxE\<^sub>2 body 0 0) \<and>
+ (\<exists>vs'. P \<turnstile> (None,h\<^sub>2,frs',sh\<^sub>2) -jvm\<rightarrow> handle P D clinit xa h\<^sub>3 vs' ls\<^sub>3 pc\<^sub>1
+ No_ics (tl frs') sh\<^sub>3))"
+proof -
+ have M_in_D: "P\<^sub>1 \<turnstile> D sees clinit,Static: []\<rightarrow>Void = body in D"
+ using method by(rule sees_method_idemp)
+ hence M_code: "compP\<^sub>2 P\<^sub>1,D,clinit,0 \<rhd> compE\<^sub>2 body @ [Return]"
+ and M_xtab: "compP\<^sub>2 P\<^sub>1,D,clinit \<rhd> compxE\<^sub>2 body 0 0/{..<size(compE\<^sub>2 body)},0"
+ by(rule beforeM, rule beforexM)
+ have nsub: "\<not>sub_RI body" by(rule sees_wf\<^sub>1_nsub_RI[OF wf method])
+ then show ?thesis using assms nsub_RI_Jcc_pieces M_code M_xtab by clarsimp
+qed
+
+lemma Jcc_pieces_Cons:
+assumes [simp]: "P \<equiv> compP\<^sub>2 P\<^sub>1"
+ and "P,C,M,pc \<rhd> compEs\<^sub>2 (e#es)" and "P,C,M \<rhd> compxEs\<^sub>2 (e#es) pc (size vs)/I,size vs"
+ and "{pc..<pc+size(compEs\<^sub>2 (e#es))} \<subseteq> I"
+ and "ics = No_ics"
+ and "\<not>sub_RIs (e#es)"
+shows "Jcc_pieces P\<^sub>1 E C M h vs ls pc ics frs sh
+ (I - pcs (compxEs\<^sub>2 es (pc + length (compE\<^sub>2 e)) (Suc (length vs)))) h' ls' sh' v xa e
+ = (True, (vs, ls, C, M, pc, ics) # frs,
+ (None, h', (v#vs, ls', C, M, pc + length (compE\<^sub>2 e), ics) # frs, sh'),
+ \<exists>pc\<^sub>1\<ge>pc. pc\<^sub>1 < pc + length (compE\<^sub>2 e) \<and> \<not> caught P pc\<^sub>1 h' xa (compxE\<^sub>2 e pc (length vs))
+ \<and> (\<exists>vs'. P \<turnstile> (None, h, (vs, ls, C, M, pc, ics) # frs, sh)
+ -jvm\<rightarrow> handle P C M xa h' (vs'@vs) ls' pc\<^sub>1 ics frs sh'))"
+proof -
+ show ?thesis using assms nsub_RI_Jcc_pieces[where e=e] by auto
+qed
+
+lemma Jcc_pieces_InitNone:
+assumes [simp]: "P \<equiv> compP\<^sub>2 P\<^sub>1"
+ and "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' sh' v xa (INIT C' (C\<^sub>0 # Cs,False) \<leftarrow> e)
+ = (True, frs', (None, h', (vs, l, C, M, pc, Called []) # frs, sh'), err)"
+shows
+ "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs (sh(C\<^sub>0 \<mapsto> (sblank P C\<^sub>0, Prepared)))
+ I h' l' sh' v xa (INIT C' (C\<^sub>0 # Cs,False) \<leftarrow> e)
+ = (True, frs', (None, h', (vs, l, C, M, pc, Called []) # frs, sh'),
+ \<exists>vs'. P \<turnstile> (None,h,frs',(sh(C\<^sub>0 \<mapsto> (sblank P\<^sub>1 C\<^sub>0, Prepared))))
+ -jvm\<rightarrow> handle P C M xa h' (vs'@vs) l pc ics frs sh')"
+proof -
+ have "Jcc_cond P\<^sub>1 E C M vs pc ics I h sh (INIT C' (C\<^sub>0 # Cs,False) \<leftarrow> e)" using assms by simp
+ then obtain T where "P\<^sub>1,E,h,sh \<turnstile>\<^sub>1 INIT C' (C\<^sub>0 # Cs,False) \<leftarrow> unit : T" by fastforce
+ then have "P\<^sub>1,E,h,sh(C\<^sub>0 \<mapsto> (sblank P\<^sub>1 C\<^sub>0, Prepared)) \<turnstile>\<^sub>1 INIT C' (C\<^sub>0 # Cs,False) \<leftarrow> unit : T"
+ by(auto simp: fun_upd_apply)
+ then have "Ex (WTrt2\<^sub>1 P\<^sub>1 E h (sh(C\<^sub>0 \<mapsto> (sblank P\<^sub>1 C\<^sub>0, Prepared))) (INIT C' (C\<^sub>0 # Cs,False) \<leftarrow> unit))"
+ by(simp only: exI)
+ then show ?thesis using assms by clarsimp
+qed
+
+lemma Jcc_pieces_InitDP:
+assumes [simp]: "P \<equiv> compP\<^sub>2 P\<^sub>1"
+ and "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' sh' v xa (INIT C' (C\<^sub>0 # Cs,False) \<leftarrow> e)
+ = (True, frs', (None, h', (vs, l, C, M, pc, Called []) # frs, sh'), err)"
+shows
+ "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' sh' v xa (INIT C' (Cs,True) \<leftarrow> e)
+ = (True, (calling_to_scalled (hd frs'))#(tl frs'),
+ (None, h', (vs, l, C, M, pc, Called []) # frs, sh'),
+ \<exists>vs'. P \<turnstile> (None,h,calling_to_scalled (hd frs')#(tl frs'),sh)
+ -jvm\<rightarrow> handle P C M xa h' (vs'@vs) l pc ics frs sh')"
+proof -
+ have "Jcc_cond P\<^sub>1 E C M vs pc ics I h sh (INIT C' (C\<^sub>0 # Cs,False) \<leftarrow> e)" using assms by simp
+ then obtain T where "P\<^sub>1,E,h,sh \<turnstile>\<^sub>1 INIT C' (C\<^sub>0 # Cs,False) \<leftarrow> unit : T" by fastforce
+ then have "P\<^sub>1,E,h,sh \<turnstile>\<^sub>1 INIT C' (Cs,True) \<leftarrow> unit : T"
+ by (auto; metis list.sel(2) list.set_sel(2))
+ then have wtrt: "Ex (WTrt2\<^sub>1 P\<^sub>1 E h sh (INIT C' (Cs,True) \<leftarrow> unit))" by(simp only: exI)
+ show ?thesis using assms wtrt
+ proof(cases Cs)
+ case (Cons C1 Cs1)
+ then show ?thesis using assms wtrt
+ by(case_tac "method P C1 clinit") clarsimp
+ qed(clarsimp)
+qed
+
+lemma Jcc_pieces_InitError:
+assumes [simp]: "P \<equiv> compP\<^sub>2 P\<^sub>1"
+ and "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' sh' v xa (INIT C' (C\<^sub>0 # Cs,False) \<leftarrow> e)
+ = (True, frs', (None, h', (vs, l, C, M, pc, Called []) # frs, sh'), err)"
+ and err: "sh C\<^sub>0 = Some(sfs,Error)"
+shows
+ "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' sh' v xa (RI (C\<^sub>0, THROW NoClassDefFoundError);Cs \<leftarrow> e)
+ = (True, (calling_to_throwing (hd frs') (addr_of_sys_xcpt NoClassDefFoundError))#(tl frs'),
+ (None, h', (vs, l, C, M, pc, Called []) # frs, sh'),
+ \<exists>vs'. P \<turnstile> (None,h, (calling_to_throwing (hd frs') (addr_of_sys_xcpt NoClassDefFoundError))#(tl frs'),sh)
+ -jvm\<rightarrow> handle P C M xa h' (vs'@vs) l pc ics frs sh')"
+proof -
+ show ?thesis using assms
+ proof(cases Cs)
+ case (Cons C1 Cs1)
+ then show ?thesis using assms
+ by(case_tac "method P C1 clinit", case_tac "method P C\<^sub>0 clinit") clarsimp
+ qed(clarsimp)
+qed
+
+lemma Jcc_pieces_InitObj:
+assumes [simp]: "P \<equiv> compP\<^sub>2 P\<^sub>1"
+ and "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' (sh(C\<^sub>0 \<mapsto> (sfs,Processing))) v xa (INIT C' (C\<^sub>0 # Cs,False) \<leftarrow> e)
+ = (True, frs', (None, h', (vs, l, C, M, pc, Called []) # frs, sh'), err)"
+shows
+ "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs (sh(C\<^sub>0 \<mapsto> (sfs,Processing))) I h' l' sh'' v xa (INIT C' (C\<^sub>0 # Cs,True) \<leftarrow> e)
+ = (True, calling_to_called (hd frs')#(tl frs'),
+ (None, h', (vs, l, C, M, pc, Called []) # frs, sh''),
+ \<exists>vs'. P \<turnstile> (None,h,calling_to_called (hd frs')#(tl frs'),sh')
+ -jvm\<rightarrow> handle P C M xa h' (vs'@vs) l pc ics frs sh'')"
+proof -
+ have "Jcc_cond P\<^sub>1 E C M vs pc ics I h sh (INIT C' (C\<^sub>0 # Cs,False) \<leftarrow> e)" using assms by simp
+ then obtain T where "P\<^sub>1,E,h,sh \<turnstile>\<^sub>1 INIT C' (C\<^sub>0 # Cs,False) \<leftarrow> unit : T" by fastforce
+ then have "P\<^sub>1,E,h,sh(C\<^sub>0 \<mapsto> (sfs,Processing)) \<turnstile>\<^sub>1 INIT C' (C\<^sub>0 # Cs,True) \<leftarrow> unit : T"
+ using assms by clarsimp (auto simp: fun_upd_apply)
+ then have wtrt: "Ex (WTrt2\<^sub>1 P\<^sub>1 E h (sh(C\<^sub>0 \<mapsto> (sfs,Processing))) (INIT C' (C\<^sub>0 # Cs,True) \<leftarrow> unit))"
+ by(simp only: exI)
+ show ?thesis using assms wtrt by clarsimp
+qed
+
+lemma Jcc_pieces_InitNonObj:
+assumes [simp]: "P \<equiv> compP\<^sub>2 P\<^sub>1"
+ and "is_class P\<^sub>1 D" and "D \<notin> set (C\<^sub>0#Cs)" and "\<forall>C \<in> set (C\<^sub>0#Cs). P\<^sub>1 \<turnstile> C \<preceq>\<^sup>* D"
+ and pcs: "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' (sh(C\<^sub>0 \<mapsto> (sfs,Processing))) v xa (INIT C' (C\<^sub>0 # Cs,False) \<leftarrow> e)
+ = (True, frs', (None, h', (vs, l, C, M, pc, Called []) # frs, sh'), err)"
+shows
+ "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs (sh(C\<^sub>0 \<mapsto> (sfs,Processing))) I h' l' sh'' v xa (INIT C' (D # C\<^sub>0 # Cs,False) \<leftarrow> e)
+ = (True, calling_to_calling (hd frs') D#(tl frs'),
+ (None, h', (vs, l, C, M, pc, Called []) # frs, sh''),
+ \<exists>vs'. P \<turnstile> (None,h,calling_to_calling (hd frs') D#(tl frs'),sh')
+ -jvm\<rightarrow> handle P C M xa h' (vs'@vs) l pc ics frs sh'')"
+proof -
+ have "Jcc_cond P\<^sub>1 E C M vs pc ics I h sh (INIT C' (C\<^sub>0 # Cs,False) \<leftarrow> e)" using assms by simp
+ then obtain T where "P\<^sub>1,E,h,sh \<turnstile>\<^sub>1 INIT C' (C\<^sub>0 # Cs,False) \<leftarrow> unit : T" by fastforce
+ then have "P\<^sub>1,E,h,sh(C\<^sub>0 \<mapsto> (sfs,Processing)) \<turnstile>\<^sub>1 INIT C' (D # C\<^sub>0 # Cs,False) \<leftarrow> unit : T"
+ using assms by clarsimp (auto simp: fun_upd_apply)
+ then have wtrt: "Ex (WTrt2\<^sub>1 P\<^sub>1 E h (sh(C\<^sub>0 \<mapsto> (sfs,Processing))) (INIT C' (D # C\<^sub>0 # Cs,False) \<leftarrow> unit))"
+ by(simp only: exI)
+ show ?thesis using assms wtrt by clarsimp
+qed
+
+lemma Jcc_pieces_InitRInit:
+assumes [simp]: "P \<equiv> compP\<^sub>2 P\<^sub>1" and wf: "wf_J\<^sub>1_prog P\<^sub>1"
+ and "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' sh' v xa (INIT C' (C\<^sub>0 # Cs,True) \<leftarrow> e)
+ = (True, frs', (None, h', (vs, l, C, M, pc, Called []) # frs, sh'), err)"
+shows
+ "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' sh' v xa (RI (C\<^sub>0,C\<^sub>0\<bullet>\<^sub>sclinit([])) ; Cs \<leftarrow> e)
+ = (True, frs',
+ (None, h', (vs, l, C, M, pc, Called []) # frs, sh'),
+ \<exists>vs'. P \<turnstile> (None,h,frs',sh)
+ -jvm\<rightarrow> handle P C M xa h' (vs'@vs) l pc ics frs sh')"
+proof -
+ have cond: "Jcc_cond P\<^sub>1 E C M vs pc ics I h sh (INIT C' (C\<^sub>0 # Cs,True) \<leftarrow> e)" using assms by simp
+ then have clinit: "\<exists>T. P\<^sub>1,E,h,sh \<turnstile>\<^sub>1 C\<^sub>0\<bullet>\<^sub>sclinit([]) : T" using wf
+ by clarsimp (auto simp: is_class_def intro: wf\<^sub>1_types_clinit)
+ then obtain T where cT: "P\<^sub>1,E,h,sh \<turnstile>\<^sub>1 C\<^sub>0\<bullet>\<^sub>sclinit([]) : T" by blast
+ obtain T where "P\<^sub>1,E,h,sh \<turnstile>\<^sub>1 INIT C' (C\<^sub>0 # Cs,True) \<leftarrow> unit : T" using cond by fastforce
+ then have "P\<^sub>1,E,h,sh \<turnstile>\<^sub>1 RI (C\<^sub>0,C\<^sub>0\<bullet>\<^sub>sclinit([])) ; Cs \<leftarrow> unit : T"
+ using assms by (auto intro: cT)
+ then have wtrt: "Ex (WTrt2\<^sub>1 P\<^sub>1 E h sh (RI (C\<^sub>0,C\<^sub>0\<bullet>\<^sub>sclinit([])) ; Cs \<leftarrow> unit))"
+ by(simp only: exI)
+ then show ?thesis using assms by simp
+qed
+
+lemma Jcc_pieces_RInit_clinit:
+assumes [simp]: "P \<equiv> compP\<^sub>2 P\<^sub>1" and wf: "wf_J\<^sub>1_prog P\<^sub>1"
+ and "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h\<^sub>1 l\<^sub>1 sh\<^sub>1 v xa (RI (C\<^sub>0,C\<^sub>0\<bullet>\<^sub>sclinit([]));Cs \<leftarrow> e)
+ = (True, frs',
+ (None, h\<^sub>1, (vs, l, C, M, pc, Called []) # frs, sh\<^sub>1), err)"
+shows
+ "Jcc_pieces P\<^sub>1 E C M h vs l pc (Called Cs) (tl frs') sh I h' l' sh' v xa (C\<^sub>0\<bullet>\<^sub>sclinit([]))
+ = (True, create_init_frame P C\<^sub>0#(vs,l,C,M,pc,Called Cs)#tl frs',
+ (None, h', (vs,l,C,M,pc,Called Cs)#tl frs', sh'(C\<^sub>0\<mapsto>(fst(the(sh' C\<^sub>0)),Done))),
+ P \<turnstile> (None,h,create_init_frame P C\<^sub>0#(vs,l,C,M,pc,Called Cs)#tl frs',sh)
+ -jvm\<rightarrow> (None,h',(vs, l, C, M, pc, Throwing Cs xa) # tl frs',sh'(C\<^sub>0 \<mapsto> (fst(the(sh' C\<^sub>0)),Error))))"
+proof -
+ have cond: "Jcc_cond P\<^sub>1 E C M vs pc ics I h sh (RI (C\<^sub>0,C\<^sub>0\<bullet>\<^sub>sclinit([]));Cs \<leftarrow> e)" using assms by simp
+ then have wtrt: "\<exists>T. P\<^sub>1,E,h,sh \<turnstile>\<^sub>1 C\<^sub>0\<bullet>\<^sub>sclinit([]) : T" using wf
+ by clarsimp (auto simp: is_class_def intro: wf\<^sub>1_types_clinit)
+ then show ?thesis using assms by clarsimp
+qed
+
+lemma Jcc_pieces_RInit_Init:
+assumes [simp]: "P \<equiv> compP\<^sub>2 P\<^sub>1" and wf: "wf_J\<^sub>1_prog P\<^sub>1"
+ and proc: "\<forall>C' \<in> set Cs. \<exists>sfs. sh'' C' = \<lfloor>(sfs,Processing)\<rfloor>"
+ and "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h\<^sub>1 l\<^sub>1 sh\<^sub>1 v xa (RI (C\<^sub>0,C\<^sub>0\<bullet>\<^sub>sclinit([]));Cs \<leftarrow> e)
+ = (True, frs',
+ (None, h\<^sub>1, (vs, l, C, M, pc, Called []) # frs, sh\<^sub>1), err)"
+shows
+ "Jcc_pieces P\<^sub>1 E C M h' vs l pc ics frs sh'' I h\<^sub>1 l\<^sub>1 sh\<^sub>1 v xa (INIT (last (C\<^sub>0#Cs)) (Cs,True) \<leftarrow> e)
+ = (True, (vs, l, C, M, pc, Called Cs) # frs,
+ (None, h\<^sub>1, (vs, l, C, M, pc, Called []) # frs, sh\<^sub>1),
+ \<exists>vs'. P \<turnstile> (None,h',(vs, l, C, M, pc, Called Cs) # frs,sh'')
+ -jvm\<rightarrow> handle P C M xa h\<^sub>1 (vs'@vs) l pc ics frs sh\<^sub>1)"
+proof -
+ have "Jcc_cond P\<^sub>1 E C M vs pc ics I h sh (RI (C\<^sub>0,C\<^sub>0\<bullet>\<^sub>sclinit([]));Cs \<leftarrow> e)" using assms by simp
+ then have "Ex (WTrt2\<^sub>1 P\<^sub>1 E h sh (RI (C\<^sub>0,C\<^sub>0\<bullet>\<^sub>sclinit([])) ; Cs \<leftarrow> unit))" by simp
+ then obtain T where riwt: "P\<^sub>1,E,h,sh \<turnstile>\<^sub>1 RI (C\<^sub>0,C\<^sub>0\<bullet>\<^sub>sclinit([]));Cs \<leftarrow> unit : T" by meson
+ then have "P\<^sub>1,E,h',sh'' \<turnstile>\<^sub>1 INIT (last (C\<^sub>0#Cs)) (Cs,True) \<leftarrow> unit : T" using proc
+ proof(cases Cs) qed(auto)
+ then have wtrt: "Ex (WTrt2\<^sub>1 P\<^sub>1 E h' sh'' (INIT (last (C\<^sub>0#Cs)) (Cs,True) \<leftarrow> unit))" by(simp only: exI)
+ show ?thesis using assms wtrt
+ proof(cases Cs)
+ case (Cons C1 Cs1)
+ then show ?thesis using assms wtrt
+ by(case_tac "method P C1 clinit") clarsimp
+ qed(clarsimp)
+qed
+
+lemma Jcc_pieces_RInit_RInit:
+assumes [simp]: "P \<equiv> compP\<^sub>2 P\<^sub>1"
+ and "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h\<^sub>1 l\<^sub>1 sh\<^sub>1 v xa (RI (C\<^sub>0,e);D#Cs \<leftarrow> e')
+ = (True, frs', rhs, err)"
+ and hd: "hd frs' = (vs1,l1,C1,M1,pc1,ics1)"
+shows
+ "Jcc_pieces P\<^sub>1 E C M h' vs l pc ics frs sh'' I h\<^sub>1 l\<^sub>1 sh\<^sub>1 v xa (RI (D,Throw xa) ; Cs \<leftarrow> e')
+ = (True, (vs1, l1, C1, M1, pc1, Throwing (D#Cs) xa) # tl frs',
+ (None, h\<^sub>1, (vs, l, C, M, pc, Called []) # frs, sh\<^sub>1),
+ \<exists>vs'. P \<turnstile> (None,h',(vs1, l1, C1, M1, pc1, Throwing (D#Cs) xa) # tl frs',sh'')
+ -jvm\<rightarrow> handle P C M xa h\<^sub>1 (vs'@vs) l pc ics frs sh\<^sub>1)"
+using assms by(case_tac "method P D clinit", cases "e = C\<^sub>0\<bullet>\<^sub>sclinit([])") clarsimp+
+
+
+subsubsection "JVM stepping lemmas"
+
+lemma jvm_Invoke:
+assumes [simp]: "P \<equiv> compP\<^sub>2 P\<^sub>1"
+ and "P,C,M,pc \<triangleright> Invoke M' (length Ts)"
+ and ha: "h\<^sub>2 a = \<lfloor>(Ca, fs)\<rfloor>" and method: "P\<^sub>1 \<turnstile> Ca sees M', NonStatic : Ts\<rightarrow>T = body in D"
+ and len: "length pvs = length Ts" and "ls\<^sub>2' = Addr a # pvs @ replicate (max_vars body) undefined"
+shows "P \<turnstile> (None, h\<^sub>2, (rev pvs @ Addr a # vs, ls\<^sub>2, C, M, pc, No_ics) # frs, sh\<^sub>2) -jvm\<rightarrow>
+ (None, h\<^sub>2, ([], ls\<^sub>2', D, M', 0, No_ics) # (rev pvs @ Addr a # vs, ls\<^sub>2, C, M, pc, No_ics) # frs, sh\<^sub>2)"
+proof -
+ have cname: "cname_of h\<^sub>2 (the_Addr ((rev pvs @ Addr a # vs) ! length Ts)) = Ca"
+ using ha method len by(auto simp: nth_append)
+ have r: "(rev pvs @ Addr a # vs) ! (length Ts) = Addr a" using len by(auto simp: nth_append)
+ have exm: "\<exists>Ts T m D b. P \<turnstile> Ca sees M',b:Ts \<rightarrow> T = m in D"
+ using sees_method_compP[OF method] by fastforce
+ show ?thesis using assms cname r exm by simp
+qed
+
+lemma jvm_Invokestatic:
+assumes [simp]: "P \<equiv> compP\<^sub>2 P\<^sub>1"
+ and "P,C,M,pc \<triangleright> Invokestatic C' M' (length Ts)"
+ and sh: "sh\<^sub>2 D = Some(sfs,Done)"
+ and method: "P\<^sub>1 \<turnstile> C' sees M', Static : Ts\<rightarrow>T = body in D"
+ and len: "length pvs = length Ts" and "ls\<^sub>2' = pvs @ replicate (max_vars body) undefined"
+shows "P \<turnstile> (None, h\<^sub>2, (rev pvs @ vs, ls\<^sub>2, C, M, pc, No_ics) # frs, sh\<^sub>2) -jvm\<rightarrow>
+ (None, h\<^sub>2, ([], ls\<^sub>2', D, M', 0, No_ics) # (rev pvs @ vs, ls\<^sub>2, C, M, pc, No_ics) # frs, sh\<^sub>2)"
+proof -
+ have exm: "\<exists>Ts T m D b. P \<turnstile> C' sees M',b:Ts \<rightarrow> T = m in D"
+ using sees_method_compP[OF method] by fastforce
+ show ?thesis using assms exm by simp
+qed
+
+lemma jvm_Invokestatic_Called:
+assumes [simp]: "P \<equiv> compP\<^sub>2 P\<^sub>1"
+ and "P,C,M,pc \<triangleright> Invokestatic C' M' (length Ts)"
+ and sh: "sh\<^sub>2 D = Some(sfs,i)"
+ and method: "P\<^sub>1 \<turnstile> C' sees M', Static : Ts\<rightarrow>T = body in D"
+ and len: "length pvs = length Ts" and "ls\<^sub>2' = pvs @ replicate (max_vars body) undefined"
+shows "P \<turnstile> (None, h\<^sub>2, (rev pvs @ vs, ls\<^sub>2, C, M, pc, Called []) # frs, sh\<^sub>2) -jvm\<rightarrow>
+ (None, h\<^sub>2, ([], ls\<^sub>2', D, M', 0, No_ics) # (rev pvs @ vs, ls\<^sub>2, C, M, pc, No_ics) # frs, sh\<^sub>2)"
+proof -
+ have exm: "\<exists>Ts T m D b. P \<turnstile> C' sees M',b:Ts \<rightarrow> T = m in D"
+ using sees_method_compP[OF method] by fastforce
+ show ?thesis using assms exm by simp
+qed
+
+lemma jvm_Return_Init:
+"P,D,clinit,0 \<rhd> compE\<^sub>2 body @ [Return]
+ \<Longrightarrow> P \<turnstile> (None, h, (vs, ls, D, clinit, size(compE\<^sub>2 body), No_ics) # frs, sh)
+ -jvm\<rightarrow> (None, h, frs, sh(D\<mapsto>(fst(the(sh D)),Done)))"
+(is "?P \<Longrightarrow> P \<turnstile> ?s1 -jvm\<rightarrow> ?s2")
+proof -
+ assume ?P
+ then have "exec (P, ?s1) = \<lfloor>?s2\<rfloor>" by(cases frs) auto
+ then have "(?s1, ?s2) \<in> (exec_1 P)\<^sup>*"
+ by(rule exec_1I[THEN r_into_rtrancl])
+ then show ?thesis by(simp add: exec_all_def1)
+qed
+
+lemma jvm_InitNone:
+ "\<lbrakk> ics_of f = Calling C Cs;
+ sh C = None \<rbrakk>
+ \<Longrightarrow> P \<turnstile> (None,h,f#frs,sh) -jvm\<rightarrow> (None,h,f#frs,sh(C \<mapsto> (sblank P C, Prepared)))"
+(is "\<lbrakk> ?P; ?Q \<rbrakk> \<Longrightarrow> P \<turnstile> ?s1 -jvm\<rightarrow> ?s2")
+proof -
+ assume assms: ?P ?Q
+ then obtain stk1 loc1 C1 M1 pc1 ics1 where "f = (stk1,loc1,C1,M1,pc1,ics1)"
+ by(cases f) simp
+ then have "exec (P, ?s1) = \<lfloor>?s2\<rfloor>" using assms
+ by(case_tac ics1) simp_all
+ then have "(?s1, ?s2) \<in> (exec_1 P)\<^sup>*"
+ by(rule exec_1I[THEN r_into_rtrancl])
+ then show ?thesis by(simp add: exec_all_def1)
+qed
+
+lemma jvm_InitDP:
+ "\<lbrakk> ics_of f = Calling C Cs;
+ sh C = \<lfloor>(sfs,i)\<rfloor>; i = Done \<or> i = Processing \<rbrakk>
+ \<Longrightarrow> P \<turnstile> (None,h,f#frs,sh) -jvm\<rightarrow> (None,h,(calling_to_scalled f)#frs,sh)"
+(is "\<lbrakk> ?P; ?Q; ?R \<rbrakk> \<Longrightarrow> P \<turnstile> ?s1 -jvm\<rightarrow> ?s2")
+proof -
+ assume assms: ?P ?Q ?R
+ then obtain stk1 loc1 C1 M1 pc1 ics1 where "f = (stk1,loc1,C1,M1,pc1,ics1)"
+ by(cases f) simp
+ then have "exec (P, ?s1) = \<lfloor>?s2\<rfloor>" using assms
+ by(case_tac i) simp_all
+ then have "(?s1, ?s2) \<in> (exec_1 P)\<^sup>*"
+ by(rule exec_1I[THEN r_into_rtrancl])
+ then show ?thesis by(simp add: exec_all_def1)
+qed
+
+lemma jvm_InitError:
+ "sh C = \<lfloor>(sfs,Error)\<rfloor>
+ \<Longrightarrow> P \<turnstile> (None,h,(vs,ls,C\<^sub>0,M,pc,Calling C Cs)#frs,sh)
+ -jvm\<rightarrow> (None,h,(vs,ls,C\<^sub>0,M,pc,Throwing Cs (addr_of_sys_xcpt NoClassDefFoundError))#frs,sh)"
+ by(clarsimp simp: exec_all_def1 intro!: r_into_rtrancl exec_1I)
+
+lemma exec_ErrorThrowing:
+ "sh C = \<lfloor>(sfs,Error)\<rfloor>
+ \<Longrightarrow> exec (P, (None,h,calling_to_throwing (stk,loc,D,M,pc,Calling C Cs) a#frs,sh))
+ = Some (None,h,calling_to_sthrowing (stk,loc,D,M,pc,Calling C Cs) a #frs,sh)"
+ by(clarsimp simp: exec_all_def1 fun_upd_idem_iff intro!: r_into_rtrancl exec_1I)
+
+lemma jvm_InitObj:
+ "\<lbrakk> sh C = Some(sfs,Prepared);
+ C = Object;
+ sh' = sh(C \<mapsto> (sfs,Processing)) \<rbrakk>
+\<Longrightarrow> P \<turnstile> (None, h, (vs,ls,C\<^sub>0,M,pc,Calling C Cs)#frs, sh) -jvm\<rightarrow>
+ (None, h, (vs,ls,C\<^sub>0,M,pc,Called (C#Cs))#frs,sh')"
+(is "\<lbrakk> ?P; ?Q; ?R \<rbrakk> \<Longrightarrow> P \<turnstile> ?s1 -jvm\<rightarrow> ?s2")
+proof -
+ assume assms: ?P ?Q ?R
+ then have "exec (P, ?s1) = \<lfloor>?s2\<rfloor>"
+ by(case_tac "method P C clinit") simp
+ then have "(?s1, ?s2) \<in> (exec_1 P)\<^sup>*"
+ by(rule exec_1I[THEN r_into_rtrancl])
+ then show ?thesis by(simp add: exec_all_def1)
+qed
+
+lemma jvm_InitNonObj:
+ "\<lbrakk> sh C = Some(sfs,Prepared);
+ C \<noteq> Object;
+ class P C = Some (D,r);
+ sh' = sh(C \<mapsto> (sfs,Processing)) \<rbrakk>
+\<Longrightarrow> P \<turnstile> (None, h, (vs,ls,C\<^sub>0,M,pc,Calling C Cs)#frs, sh) -jvm\<rightarrow>
+ (None, h, (vs,ls,C\<^sub>0,M,pc,Calling D (C#Cs))#frs, sh')"
+(is "\<lbrakk> ?P; ?Q; ?R; ?S \<rbrakk> \<Longrightarrow> P \<turnstile> ?s1 -jvm\<rightarrow> ?s2")
+proof -
+ assume assms: ?P ?Q ?R ?S
+ then have "exec (P, ?s1) = \<lfloor>?s2\<rfloor>"
+ by(case_tac "method P C clinit") simp
+ then have "(?s1, ?s2) \<in> (exec_1 P)\<^sup>*"
+ by(rule exec_1I[THEN r_into_rtrancl])
+ then show ?thesis by(simp add: exec_all_def1)
+qed
+
+lemma jvm_RInit_throw:
+ "P \<turnstile> (None,h,(vs,l,C,M,pc,Throwing [] xa) # frs,sh)
+ -jvm\<rightarrow> handle P C M xa h vs l pc No_ics frs sh"
+(is "P \<turnstile> ?s1 -jvm\<rightarrow> ?s2")
+proof -
+ have "exec (P, ?s1) = \<lfloor>?s2\<rfloor>"
+ by(simp add: handle_def split: bool.splits)
+ then have "(?s1, ?s2) \<in> (exec_1 P)\<^sup>*"
+ by(rule exec_1I[THEN r_into_rtrancl])
+ then show ?thesis by(simp add: exec_all_def1)
+qed
+
+lemma jvm_RInit_throw':
+ "P \<turnstile> (None,h,(vs,l,C,M,pc,Throwing [C'] xa) # frs,sh)
+ -jvm\<rightarrow> handle P C M xa h vs l pc No_ics frs (sh(C':=Some(fst(the(sh C')), Error)))"
+(is "P \<turnstile> ?s1 -jvm\<rightarrow> ?s2")
+proof -
+ let ?sy = "(None,h,(vs,l,C,M,pc,Throwing [] xa) # frs,sh(C':=Some(fst(the(sh C')), Error)))"
+ have "exec (P, ?s1) = \<lfloor>?sy\<rfloor>" by simp
+ then have "(?s1, ?sy) \<in> (exec_1 P)\<^sup>*"
+ by(rule exec_1I[THEN r_into_rtrancl])
+ also have "(?sy, ?s2) \<in> (exec_1 P)\<^sup>*"
+ using jvm_RInit_throw by(simp add: exec_all_def1)
+ ultimately show ?thesis by(simp add: exec_all_def1)
+qed
+
+lemma jvm_Called:
+ "P \<turnstile> (None, h, (vs, l, C, M, pc, Called (C\<^sub>0 # Cs)) # frs, sh) -jvm\<rightarrow>
+ (None, h, create_init_frame P C\<^sub>0 # (vs, l, C, M, pc, Called Cs) # frs, sh)"
+ by(simp add: exec_all_def1 r_into_rtrancl exec_1I)
+
+lemma jvm_Throwing:
+ "P \<turnstile> (None, h, (vs, l, C, M, pc, Throwing (C\<^sub>0#Cs) xa') # frs, sh) -jvm\<rightarrow>
+ (None, h, (vs, l, C, M, pc, Throwing Cs xa') # frs, sh(C\<^sub>0 \<mapsto> (fst (the (sh C\<^sub>0)), Error)))"
+ by(simp add: exec_all_def1 r_into_rtrancl exec_1I)
+
+subsubsection "Other lemmas for correctness proof"
+
+lemma assumes wf:"wf_prog wf_md P"
+ and ex: "class P C = Some a"
+shows create_init_frame_wf_eq: "create_init_frame (compP\<^sub>2 P) C = (stk,loc,D,M,pc,ics) \<Longrightarrow> D=C"
+using wf_sees_clinit[OF wf ex] by(cases "method P C clinit", auto)
+
+lemma beforex_try:
+assumes pcI: "{pc..<pc+size(compE\<^sub>2(try e\<^sub>1 catch(Ci i) e\<^sub>2))} \<subseteq> I"
+ and bx: "P,C,M \<rhd> compxE\<^sub>2 (try e\<^sub>1 catch(Ci i) e\<^sub>2) pc (size vs) / I,size vs"
+shows "P,C,M \<rhd> compxE\<^sub>2 e\<^sub>1 pc (size vs) / {pc..<pc + length (compE\<^sub>2 e\<^sub>1)},size vs"
+proof -
+ obtain xt\<^sub>0 xt\<^sub>1 where
+ "beforex\<^sub>0 P C M (size vs) I (compxE\<^sub>2 (try e\<^sub>1 catch(Ci i) e\<^sub>2) pc (size vs)) xt\<^sub>0 xt\<^sub>1"
+ using bx by(clarsimp simp:beforex_def)
+ then have "\<exists>xt1. beforex\<^sub>0 P C M (size vs) {pc..<pc + length (compE\<^sub>2 e\<^sub>1)}
+ (compxE\<^sub>2 e\<^sub>1 pc (size vs)) xt\<^sub>0 xt1"
+ using pcI pcs_subset(1) atLeastLessThan_iff by simp blast
+ then show ?thesis using beforex_def by blast
+qed
+
+\<comment> \<open> Evaluation of initialization expressions \<close>
+
+(* --needs J1 and EConform; version for eval found in Equivalence *)
+lemma
+shows eval\<^sub>1_init_return: "P \<turnstile>\<^sub>1 \<langle>e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>
+ \<Longrightarrow> iconf (shp\<^sub>1 s) e
+ \<Longrightarrow> (\<exists>Cs b. e = INIT C' (Cs,b) \<leftarrow> unit) \<or> (\<exists>C e\<^sub>0 Cs e\<^sub>i. e = RI(C,e\<^sub>0);Cs@[C'] \<leftarrow> unit)
+ \<or> (\<exists>e\<^sub>0. e = RI(C',e\<^sub>0);Nil \<leftarrow> unit)
+ \<Longrightarrow> (val_of e' = Some v \<longrightarrow> (\<exists>sfs i. shp\<^sub>1 s' C' = \<lfloor>(sfs,i)\<rfloor> \<and> (i = Done \<or> i = Processing)))
+ \<and> (throw_of e' = Some a \<longrightarrow> (\<exists>sfs i. shp\<^sub>1 s' C' = \<lfloor>(sfs,Error)\<rfloor>))"
+and "P \<turnstile>\<^sub>1 \<langle>es,s\<rangle> [\<Rightarrow>] \<langle>es',s'\<rangle> \<Longrightarrow> True"
+proof(induct rule: eval\<^sub>1_evals\<^sub>1.inducts)
+ case (InitFinal\<^sub>1 e s e' s' C b) then show ?case
+ by(auto simp: initPD_def dest: eval\<^sub>1_final_same)
+next
+ case (InitDone\<^sub>1 sh C sfs C' Cs e h l e' s')
+ then have "final e'" using eval\<^sub>1_final by simp
+ then show ?case
+ proof(rule finalE)
+ fix v assume e': "e' = Val v" then show ?thesis using InitDone\<^sub>1 initPD_def
+ proof(cases Cs) qed(auto)
+ next
+ fix a assume e': "e' = throw a" then show ?thesis using InitDone\<^sub>1 initPD_def
+ proof(cases Cs) qed(auto)
+ qed
+next
+ case (InitProcessing\<^sub>1 sh C sfs C' Cs e h l e' s')
+ then have "final e'" using eval\<^sub>1_final by simp
+ then show ?case
+ proof(rule finalE)
+ fix v assume e': "e' = Val v" then show ?thesis using InitProcessing\<^sub>1 initPD_def
+ proof(cases Cs) qed(auto)
+ next
+ fix a assume e': "e' = throw a" then show ?thesis using InitProcessing\<^sub>1 initPD_def
+ proof(cases Cs) qed(auto)
+ qed
+next
+ case (InitError\<^sub>1 sh C sfs Cs e h l e' s' C') show ?case
+ proof(cases Cs)
+ case Nil then show ?thesis using InitError\<^sub>1 by simp
+ next
+ case (Cons C2 list)
+ then have "final e'" using InitError\<^sub>1 eval\<^sub>1_final by simp
+ then show ?thesis
+ proof(rule finalE)
+ fix v assume e': "e' = Val v" show ?thesis
+ using InitError\<^sub>1.hyps(2) e' rinit\<^sub>1_throwE by blast
+ next
+ fix a assume e': "e' = throw a"
+ then show ?thesis using Cons InitError\<^sub>1 cons_to_append[of list] by clarsimp
+ qed
+ qed
+next
+ case (InitRInit\<^sub>1 C Cs h l sh e' s' C') show ?case
+ proof(cases Cs)
+ case Nil then show ?thesis using InitRInit\<^sub>1 by simp
+ next
+ case (Cons C' list) then show ?thesis
+ using InitRInit\<^sub>1 Cons cons_to_append[of list] by clarsimp
+ qed
+next
+ case (RInit\<^sub>1 e s v h' l' sh' C sfs i sh'' C' Cs e' e\<^sub>1 s\<^sub>1)
+ then have final: "final e\<^sub>1" using eval\<^sub>1_final by simp
+ then show ?case
+ proof(cases Cs)
+ case Nil show ?thesis using final
+ proof(rule finalE)
+ fix v assume e': "e\<^sub>1 = Val v" show ?thesis
+ using RInit\<^sub>1 Nil by(clarsimp, meson fun_upd_same initPD_def)
+ next
+ fix a assume e': "e\<^sub>1 = throw a" show ?thesis
+ using RInit\<^sub>1 Nil by(clarsimp, meson fun_upd_same initPD_def)
+ qed
+ next
+ case (Cons a list) show ?thesis using final
+ proof(rule finalE)
+ fix v assume e': "e\<^sub>1 = Val v" then show ?thesis
+ using RInit\<^sub>1 Cons by(clarsimp, metis last.simps last_appendR list.distinct(1))
+ next
+ fix a assume e': "e\<^sub>1 = throw a" then show ?thesis
+ using RInit\<^sub>1 Cons by(clarsimp, metis last.simps last_appendR list.distinct(1))
+ qed
+ qed
+next
+ case (RInitInitFail\<^sub>1 e s a h' l' sh' C sfs i sh'' D Cs e' e\<^sub>1 s\<^sub>1)
+ then have final: "final e\<^sub>1" using eval\<^sub>1_final by simp
+ then show ?case
+ proof(rule finalE)
+ fix v assume e': "e\<^sub>1 = Val v" then show ?thesis
+ using RInitInitFail\<^sub>1 by(clarsimp, meson exp.distinct(101) rinit\<^sub>1_throwE)
+ next
+ fix a' assume e': "e\<^sub>1 = Throw a'"
+ then have "iconf (sh'(C \<mapsto> (sfs, Error))) a"
+ using RInitInitFail\<^sub>1.hyps(1) eval\<^sub>1_final by fastforce
+ then show ?thesis using RInitInitFail\<^sub>1 e'
+ by(clarsimp, meson Cons_eq_append_conv list.inject)
+ qed
+qed(auto simp: fun_upd_same)
+
+lemma init\<^sub>1_Val_PD: "P \<turnstile>\<^sub>1 \<langle>INIT C' (Cs,b) \<leftarrow> unit,s\<rangle> \<Rightarrow> \<langle>Val v,s'\<rangle>
+ \<Longrightarrow> iconf (shp\<^sub>1 s) (INIT C' (Cs,b) \<leftarrow> unit)
+ \<Longrightarrow> \<exists>sfs i. shp\<^sub>1 s' C' = \<lfloor>(sfs,i)\<rfloor> \<and> (i = Done \<or> i = Processing)"
+ by(drule_tac v = v in eval\<^sub>1_init_return, simp+)
+
+lemma init\<^sub>1_throw_PD: "P \<turnstile>\<^sub>1 \<langle>INIT C' (Cs,b) \<leftarrow> unit,s\<rangle> \<Rightarrow> \<langle>throw a,s'\<rangle>
+ \<Longrightarrow> iconf (shp\<^sub>1 s) (INIT C' (Cs,b) \<leftarrow> unit)
+ \<Longrightarrow> \<exists>sfs i. shp\<^sub>1 s' C' = \<lfloor>(sfs,Error)\<rfloor>"
+ by(drule_tac a = a in eval\<^sub>1_init_return, simp+)
+
+lemma rinit\<^sub>1_Val_PD:
+assumes eval: "P \<turnstile>\<^sub>1 \<langle>RI(C,e\<^sub>0);Cs \<leftarrow> unit,s\<rangle> \<Rightarrow> \<langle>Val v,s'\<rangle>"
+ and iconf: "iconf (shp\<^sub>1 s) (RI(C,e\<^sub>0);Cs \<leftarrow> unit)" and last: "last(C#Cs) = C'"
+shows "\<exists>sfs i. shp\<^sub>1 s' C' = \<lfloor>(sfs,i)\<rfloor> \<and> (i = Done \<or> i = Processing)"
+proof(cases Cs)
+ case Nil
+ then show ?thesis using eval\<^sub>1_init_return[OF eval iconf] last by simp
+next
+ case (Cons a list)
+ then have nNil: "Cs \<noteq> []" by simp
+ then have "\<exists>Cs'. Cs = Cs' @ [C']" using last append_butlast_last_id[OF nNil]
+ by(rule_tac x="butlast Cs" in exI) simp
+ then show ?thesis using eval\<^sub>1_init_return[OF eval iconf] by simp
+qed
+
+lemma rinit\<^sub>1_throw_PD:
+assumes eval: "P \<turnstile>\<^sub>1 \<langle>RI(C,e\<^sub>0);Cs \<leftarrow> unit,s\<rangle> \<Rightarrow> \<langle>throw a,s'\<rangle>"
+ and iconf: "iconf (shp\<^sub>1 s) (RI(C,e\<^sub>0);Cs \<leftarrow> unit)" and last: "last(C#Cs) = C'"
+shows "\<exists>sfs. shp\<^sub>1 s' C' = \<lfloor>(sfs,Error)\<rfloor>"
+proof(cases Cs)
+ case Nil
+ then show ?thesis using eval\<^sub>1_init_return[OF eval iconf] last by simp
+next
+ case (Cons a list)
+ then have nNil: "Cs \<noteq> []" by simp
+ then have "\<exists>Cs'. Cs = Cs' @ [C']" using last append_butlast_last_id[OF nNil]
+ by(rule_tac x="butlast Cs" in exI) simp
+ then show ?thesis using eval\<^sub>1_init_return[OF eval iconf] by simp
+qed
+
+subsubsection "The proof"
+
+lemma fixes P\<^sub>1 defines [simp]: "P \<equiv> compP\<^sub>2 P\<^sub>1"
+assumes wf: "wf_J\<^sub>1_prog P\<^sub>1"
+shows Jcc: "P\<^sub>1 \<turnstile>\<^sub>1 \<langle>e,(h\<^sub>0,ls\<^sub>0,sh\<^sub>0)\<rangle> \<Rightarrow> \<langle>ef,(h\<^sub>1,ls\<^sub>1,sh\<^sub>1)\<rangle> \<Longrightarrow>
+ (\<And>E C M pc ics v xa vs frs I.
+ \<lbrakk> Jcc_cond P\<^sub>1 E C M vs pc ics I h\<^sub>0 sh\<^sub>0 e \<rbrakk> \<Longrightarrow>
+ (ef = Val v \<longrightarrow>
+ P \<turnstile> (None,h\<^sub>0,Jcc_frames P C M vs ls\<^sub>0 pc ics frs e,sh\<^sub>0)
+ -jvm\<rightarrow> Jcc_rhs P\<^sub>1 E C M vs ls\<^sub>0 pc ics frs h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v e)
+ \<and>
+ (ef = Throw xa \<longrightarrow> Jcc_err P C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 xa e)
+ )"
+(*<*)
+ (is "_ \<Longrightarrow> (\<And>E C M pc ics v xa vs frs I.
+ PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 ef h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics v xa vs frs I)")
+(*>*)
+and "P\<^sub>1 \<turnstile>\<^sub>1 \<langle>es,(h\<^sub>0,ls\<^sub>0,sh\<^sub>0)\<rangle> [\<Rightarrow>] \<langle>fs,(h\<^sub>1,ls\<^sub>1,sh\<^sub>1)\<rangle> \<Longrightarrow>
+ (\<And>C M pc ics ws xa es' vs frs I.
+ \<lbrakk> P,C,M,pc \<rhd> compEs\<^sub>2 es; P,C,M \<rhd> compxEs\<^sub>2 es pc (size vs)/I,size vs;
+ {pc..<pc+size(compEs\<^sub>2 es)} \<subseteq> I; ics = No_ics;
+ \<not>sub_RIs es \<rbrakk> \<Longrightarrow>
+ (fs = map Val ws \<longrightarrow>
+ P \<turnstile> (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0) -jvm\<rightarrow>
+ (None,h\<^sub>1,(rev ws @ vs,ls\<^sub>1,C,M,pc+size(compEs\<^sub>2 es),ics)#frs,sh\<^sub>1))
+ \<and>
+ (fs = map Val ws @ Throw xa # es' \<longrightarrow>
+ (\<exists>pc\<^sub>1. pc \<le> pc\<^sub>1 \<and> pc\<^sub>1 < pc + size(compEs\<^sub>2 es) \<and>
+ \<not> caught P pc\<^sub>1 h\<^sub>1 xa (compxEs\<^sub>2 es pc (size vs)) \<and>
+ (\<exists>vs'. P \<turnstile> (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0)
+ -jvm\<rightarrow> handle P C M xa h\<^sub>1 (vs'@vs) ls\<^sub>1 pc\<^sub>1 ics frs sh\<^sub>1))))"
+(*<*)
+ (is "_ \<Longrightarrow> (\<And>C M pc ics ws xa es' vs frs I.
+ PROP ?Ps es h\<^sub>0 ls\<^sub>0 sh\<^sub>0 fs h\<^sub>1 ls\<^sub>1 sh\<^sub>1 C M pc ics ws xa es' vs frs I)")
+proof (induct rule:eval\<^sub>1_evals\<^sub>1_inducts)
+ case New\<^sub>1 thus ?case by auto
+next
+ case (NewFail\<^sub>1 sh C' sfs h ls)
+ let ?xa = "addr_of_sys_xcpt OutOfMemory"
+ have "P \<turnstile> (None,h,(vs,ls,C,M,pc,ics)#frs,sh) -jvm\<rightarrow> handle P C M ?xa h vs ls pc ics frs sh"
+ using NewFail\<^sub>1 by(clarsimp simp: handle_def)
+ then show ?case by(auto intro!: exI[where x="[]"])
+next
+ case (NewInit\<^sub>1 sh C' h ls v' h' ls' sh' a FDTs h'')
+ then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs ls pc ics frs sh I h' ls' sh' v xa (new C')
+ = (True, frs', (None,h',(v#vs,ls',C,M,pc+size(compE\<^sub>2 (new C')),ics)#frs,sh'), err)"
+ using NewInit\<^sub>1.prems(1) by clarsimp
+ have "Ex (WTrt2\<^sub>1 P\<^sub>1 E h sh (INIT C' ([C'],False) \<leftarrow> unit))"
+ using has_fields_is_class[OF NewInit\<^sub>1.hyps(5)] by auto
+ then obtain err' where pcs':
+ "Jcc_pieces P\<^sub>1 E C M h vs ls pc ics frs sh I h' ls' sh' v' xa (INIT C' ([C'],False) \<leftarrow> unit)
+ = (True, (vs,ls,C,M,pc,Calling C' []) # frs, (None,h',(vs,ls,C,M,pc,Called [])#frs,sh'), err')"
+ using NewInit\<^sub>1.prems(1) by auto
+ have IH: "PROP ?P (INIT C' ([C'],False) \<leftarrow> unit) h ls sh (Val v')
+ h' ls' sh' E C M pc ics v' xa vs frs I" by fact
+ have ls: "ls = ls'" by(rule init\<^sub>1_same_loc[OF NewInit\<^sub>1.hyps(2)])
+ obtain sfs i where sh': "sh' C' = Some(sfs,i)"
+ using init\<^sub>1_Val_PD[OF NewInit\<^sub>1.hyps(2)] by clarsimp
+ have "P \<turnstile> (None,h,(vs,ls,C,M,pc,ics)#frs,sh) -jvm\<rightarrow> (None,h,(vs,ls,C,M,pc,Calling C' [])#frs,sh)"
+ proof(cases "sh C'")
+ case None then show ?thesis using NewInit\<^sub>1.prems by(cases ics) auto
+ next
+ case (Some a)
+ then obtain sfs i where "a = (sfs,i)" by(cases a)
+ then show ?thesis using NewInit\<^sub>1.hyps(1) NewInit\<^sub>1.prems Some
+ by(cases ics; case_tac i) auto
+ qed
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> (None, h', (vs, ls, C, M, pc, Called []) # frs, sh')"
+ using IH pcs' by auto
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> (None, h'', (Addr a#vs, ls, C, M, Suc pc, ics) # frs, sh')"
+ using NewInit\<^sub>1.hyps(1,2,4-6) NewInit\<^sub>1.prems sh' by(cases ics) auto
+ finally show ?case using pcs ls by clarsimp
+next
+ case (NewInitOOM\<^sub>1 sh C' h ls v' h' ls' sh')
+ let ?xa = "addr_of_sys_xcpt OutOfMemory"
+ obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs ls pc ics frs sh I h' ls' sh' v xa (new C')
+ = (True, frs', (None,h',(v#vs,ls',C,M,pc+size(compE\<^sub>2 (new C')),ics)#frs,sh'), err)"
+ using NewInitOOM\<^sub>1.prems(1) by clarsimp
+ have "Ex (WTrt2\<^sub>1 P\<^sub>1 E h sh (INIT C' ([C'],False) \<leftarrow> unit))" using NewInitOOM\<^sub>1.hyps(5) by auto
+ then obtain err' where pcs':
+ "Jcc_pieces P\<^sub>1 E C M h vs ls pc ics frs sh I h' ls' sh' v' xa (INIT C' ([C'],False) \<leftarrow> unit)
+ = (True, (vs,ls,C,M,pc,Calling C' []) # frs, (None,h',(vs,ls,C,M,pc,Called [])#frs,sh'), err')"
+ using NewInitOOM\<^sub>1.prems(1) by auto
+ have IH: "PROP ?P (INIT C' ([C'],False) \<leftarrow> unit) h ls sh (Val v')
+ h' ls' sh' E C M pc ics v' xa vs frs I" by fact
+ have ls: "ls = ls'" by(rule init\<^sub>1_same_loc[OF NewInitOOM\<^sub>1.hyps(2)])
+ have "iconf (shp\<^sub>1 (h, ls, sh)) (INIT C' ([C'],False) \<leftarrow> unit)" by simp
+ then obtain sfs i where sh': "sh' C' = Some(sfs,i)"
+ using init\<^sub>1_Val_PD[OF NewInitOOM\<^sub>1.hyps(2)] by clarsimp
+ have "P \<turnstile> (None,h,(vs,ls,C,M,pc,ics)#frs,sh) -jvm\<rightarrow> (None,h,(vs,ls,C,M,pc,Calling C' [])#frs,sh)"
+ proof(cases "sh C'")
+ case None then show ?thesis using NewInitOOM\<^sub>1.prems by(cases ics) auto
+ next
+ case (Some a)
+ then obtain sfs i where "a = (sfs,i)" by(cases a)
+ then show ?thesis using NewInitOOM\<^sub>1.hyps(1) NewInitOOM\<^sub>1.prems Some
+ by(cases ics; case_tac i) auto
+ qed
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> (None, h', (vs, ls, C, M, pc, Called []) # frs, sh')"
+ using IH pcs' by auto
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> handle P C M ?xa h' vs ls pc ics frs sh'"
+ using NewInitOOM\<^sub>1.hyps(1,2,4,5) NewInitOOM\<^sub>1.prems sh' by(auto simp: handle_def)
+ finally show ?case using pcs ls by(simp, metis (no_types) append_Nil le_refl lessI)
+next
+ case (NewInitThrow\<^sub>1 sh C' h ls a h' ls' sh')
+ obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs ls pc ics frs sh I h' ls' sh' v xa (new C')
+ = (True, frs', (None,h',(v#vs,ls',C,M,pc+size(compE\<^sub>2 (new C')),ics)#frs,sh'), err)"
+ using NewInitThrow\<^sub>1.prems(1) by clarsimp
+ obtain a' where throw: "throw a = Throw a'" using eval\<^sub>1_final[OF NewInitThrow\<^sub>1.hyps(2)] by clarsimp
+ have "Ex (WTrt2\<^sub>1 P\<^sub>1 E h sh (INIT C' ([C'],False) \<leftarrow> unit))" using NewInitThrow\<^sub>1.hyps(4) by auto
+ then obtain vs' where pcs':
+ "Jcc_pieces P\<^sub>1 E C M h vs ls pc ics frs sh I h' ls' sh' v a' (INIT C' ([C'],False) \<leftarrow> unit)
+ = (True, (vs,ls,C,M,pc,Calling C' []) # frs, (None,h',(vs,ls,C,M,pc,Called [])#frs,sh'),
+ P \<turnstile> (None,h,(vs,ls,C,M,pc,Calling C' []) # frs,sh)
+ -jvm\<rightarrow> handle P C M a' h' (vs'@vs) ls pc ics frs sh')"
+ using NewInitThrow\<^sub>1.prems(1) by simp blast
+ have IH: "PROP ?P (INIT C' ([C'],False) \<leftarrow> unit) h ls sh (throw a)
+ h' ls' sh' E C M pc ics v a' vs frs I" by fact
+ have ls: "ls = ls'" by(rule init\<^sub>1_same_loc[OF NewInitThrow\<^sub>1.hyps(2)])
+ then have "P \<turnstile> (None,h,(vs,ls,C,M,pc,ics)#frs,sh) -jvm\<rightarrow> (None,h,(vs,ls,C,M,pc,Calling C' []) # frs,sh)"
+ proof(cases "sh C'")
+ case None then show ?thesis using NewInitThrow\<^sub>1.prems by(cases ics) auto
+ next
+ case (Some a)
+ then obtain sfs i where "a = (sfs,i)" by(cases a)
+ then show ?thesis using NewInitThrow\<^sub>1.hyps(1) NewInitThrow\<^sub>1.prems Some
+ by(cases ics; case_tac i) auto
+ qed
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> handle P C M a' h' (vs'@vs) ls pc ics frs sh'" using IH pcs' throw by auto
+ finally show ?case using throw ls by auto
+next
+ case (Cast\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 a h\<^sub>1 ls\<^sub>1 sh\<^sub>1 D fs C')
+ let ?pc = "pc + length(compE\<^sub>2 e)"
+ obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (Cast C' e)
+ = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (Cast C' e)),ics)#frs,sh\<^sub>1), err)"
+ using Cast\<^sub>1.prems(1) by auto
+ have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (addr a) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Addr a) xa vs frs I" by fact
+ then have "P \<turnstile> (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0) -jvm\<rightarrow>
+ (None,h\<^sub>1,(Addr a#vs,ls\<^sub>1,C,M,?pc,ics)#frs,sh\<^sub>1)"
+ using Jcc_pieces_Cast[OF assms(1) pcs, of "Addr a"] Cast\<^sub>1.prems pcs by auto
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> (None,h\<^sub>1,(Addr a#vs,ls\<^sub>1,C,M,?pc+1,ics)#frs,sh\<^sub>1)"
+ using Cast\<^sub>1 by (auto simp add:cast_ok_def)
+ finally show ?case by auto
+next
+ case (CastNull\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 C')
+ let ?pc = "pc + length(compE\<^sub>2 e)"
+ obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (Cast C' e)
+ = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (Cast C' e)),ics)#frs,sh\<^sub>1), err)"
+ using CastNull\<^sub>1.prems(1) by clarsimp
+ have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 null h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics Null xa vs frs I" by fact
+ then have "P \<turnstile> (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0) -jvm\<rightarrow>
+ (None,h\<^sub>1,(Null#vs,ls\<^sub>1,C,M,?pc,ics)#frs,sh\<^sub>1)"
+ using Jcc_pieces_Cast[OF assms(1) pcs, of Null] CastNull\<^sub>1.prems pcs by auto
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> (None,h\<^sub>1,(Null#vs,ls\<^sub>1,C,M,?pc+1,ics)#frs,sh\<^sub>1)"
+ using CastNull\<^sub>1 by (auto simp add:cast_ok_def)
+ finally show ?case by auto
+next
+ case (CastFail\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 a h\<^sub>1 ls\<^sub>1 sh\<^sub>1 D fs C')
+ let ?pc = "pc + length(compE\<^sub>2 e)"
+ let ?xa = "addr_of_sys_xcpt ClassCast"
+ obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (Cast C' e)
+ = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (Cast C' e)),ics)#frs,sh\<^sub>1), err)"
+ using CastFail\<^sub>1.prems(1) by clarsimp
+ have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (addr a) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Addr a) xa vs frs I" by fact
+ then have "P \<turnstile> (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0) -jvm\<rightarrow>
+ (None,h\<^sub>1,(Addr a#vs,ls\<^sub>1,C,M,?pc,ics)#frs,sh\<^sub>1)"
+ using Jcc_pieces_Cast[OF assms(1) pcs, of "Addr a"] CastFail\<^sub>1.prems pcs by auto
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> handle P C M ?xa h\<^sub>1 (Addr a#vs) ls\<^sub>1 ?pc ics frs sh\<^sub>1"
+ using CastFail\<^sub>1 by (auto simp add:handle_def cast_ok_def)
+ finally have exec: "P \<turnstile> (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0) -jvm\<rightarrow> \<dots>".
+ show ?case (is "?N \<and> (?eq \<longrightarrow> ?err)")
+ proof
+ show ?N by simp
+ next
+ { assume ?eq
+ then have ?err using exec by (auto intro!: exI[where x="?pc"] exI[where x="[Addr a]"])
+ }
+ thus "?eq \<longrightarrow> ?err" by simp
+ qed
+next
+ case (CastThrow\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 e' h\<^sub>1 ls\<^sub>1 sh\<^sub>1 C')
+ obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (Cast C' e)
+ = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (Cast C' e)),ics)#frs,sh\<^sub>1), err)"
+ using CastThrow\<^sub>1.prems(1) by clarsimp
+ have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (throw e') h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics v xa vs frs I" by fact
+ show ?case using IH Jcc_pieces_Cast[OF assms(1) pcs, of v] CastThrow\<^sub>1.prems pcs less_SucI
+ by(simp, blast)
+next
+ case Val\<^sub>1 thus ?case by auto
+next
+ case Var\<^sub>1 thus ?case by auto
+next
+ case (BinOp\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 v\<^sub>1 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 e\<^sub>2 v\<^sub>2 h\<^sub>2 ls\<^sub>2 sh\<^sub>2 bop w)
+ obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v xa (e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2)
+ = (True, frs', (None,h\<^sub>2,(v#vs,ls\<^sub>2,C,M,pc+size(compE\<^sub>2 (e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2)),ics)#frs,sh\<^sub>2), err)"
+ using BinOp\<^sub>1.prems(1) by clarsimp
+ let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e\<^sub>1)"
+ let ?pc\<^sub>2 = "?pc\<^sub>1 + length(compE\<^sub>2 e\<^sub>2)"
+ have IH\<^sub>1: "PROP ?P e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (Val v\<^sub>1) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics v\<^sub>1 xa vs frs
+ (I - pcs (compxE\<^sub>2 e\<^sub>2 (pc + length (compE\<^sub>2 e\<^sub>1)) (Suc (length vs))))" by fact
+ have IH\<^sub>2: "PROP ?P e\<^sub>2 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (Val v\<^sub>2) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 E C M ?pc\<^sub>1 ics v\<^sub>2 xa (v\<^sub>1#vs) frs
+ (I - pcs(compxE\<^sub>2 e\<^sub>1 pc (size vs)))" by fact
+ have "P \<turnstile> (None,h\<^sub>0,frs',sh\<^sub>0) -jvm\<rightarrow> (None,h\<^sub>1,(v\<^sub>1#vs,ls\<^sub>1,C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)"
+ using IH\<^sub>1 Jcc_pieces_BinOp1[OF pcs, of h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v\<^sub>1] by simp
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> (None,h\<^sub>2,(v\<^sub>2#v\<^sub>1#vs,ls\<^sub>2,C,M,?pc\<^sub>2,ics)#frs,sh\<^sub>2)"
+ using IH\<^sub>2 Jcc_pieces_BinOp2[OF assms(1) pcs, of h\<^sub>1 v\<^sub>1 ls\<^sub>1 sh\<^sub>1 v\<^sub>2] by (simp add: add.assoc)
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> (None,h\<^sub>2,(w#vs,ls\<^sub>2,C,M,?pc\<^sub>2+1,ics)#frs,sh\<^sub>2)"
+ using BinOp\<^sub>1 by(cases bop) auto
+ finally show ?case using pcs by (auto split: bop.splits simp:add.assoc)
+next
+ case (BinOpThrow\<^sub>1\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 e h\<^sub>1 ls\<^sub>1 sh\<^sub>1 bop e\<^sub>2)
+ obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2)
+ = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2)),ics)#frs,sh\<^sub>1), err)"
+ using BinOpThrow\<^sub>1\<^sub>1.prems(1) by clarsimp
+ have IH\<^sub>1: "PROP ?P e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (throw e) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics v xa vs frs
+ (I - pcs (compxE\<^sub>2 e\<^sub>2 (pc + length (compE\<^sub>2 e\<^sub>1)) (Suc (length vs))))" by fact
+ show ?case using IH\<^sub>1 Jcc_pieces_BinOp1[OF pcs, of h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v] BinOpThrow\<^sub>1\<^sub>1.prems nsub_RI_Jcc_pieces
+ by auto
+next
+ case (BinOpThrow\<^sub>2\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 v\<^sub>1 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 e\<^sub>2 e h\<^sub>2 ls\<^sub>2 sh\<^sub>2 bop)
+ obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v xa (e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2)
+ = (True, frs', (None,h\<^sub>2,(v#vs,ls\<^sub>2,C,M,pc+size(compE\<^sub>2 (e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2)),ics)#frs,sh\<^sub>2), err)"
+ using BinOpThrow\<^sub>2\<^sub>1.prems(1) by clarsimp
+ let ?pc = "pc + length(compE\<^sub>2 e\<^sub>1)"
+ have IH\<^sub>1: "PROP ?P e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (Val v\<^sub>1) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics v\<^sub>1 xa vs frs
+ (I - pcs (compxE\<^sub>2 e\<^sub>2 (pc + length (compE\<^sub>2 e\<^sub>1)) (Suc (length vs))))" by fact
+ have IH\<^sub>2: "PROP ?P e\<^sub>2 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (throw e) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 E C M ?pc ics v xa (v\<^sub>1#vs) frs
+ (I - pcs(compxE\<^sub>2 e\<^sub>1 pc (size vs)))" by fact
+ let ?\<sigma>\<^sub>1 = "(None,h\<^sub>1,(v\<^sub>1#vs,ls\<^sub>1,C,M,?pc,ics)#frs,sh\<^sub>1)"
+ have 1: "P \<turnstile> (None,h\<^sub>0,frs',sh\<^sub>0) -jvm\<rightarrow> ?\<sigma>\<^sub>1"
+ using IH\<^sub>1 Jcc_pieces_BinOp1[OF pcs, of h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v\<^sub>1] by simp
+ have "(throw e = Val v \<longrightarrow> P \<turnstile> (None, h\<^sub>0, Jcc_frames P C M vs ls\<^sub>0 pc ics frs (e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2), sh\<^sub>0) -jvm\<rightarrow>
+ Jcc_rhs P\<^sub>1 E C M vs ls\<^sub>0 pc ics frs h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v (e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2))
+ \<and> (throw e = Throw xa \<longrightarrow> (\<exists>pc\<^sub>1. pc \<le> pc\<^sub>1 \<and> pc\<^sub>1 < pc + size(compE\<^sub>2 (e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2)) \<and>
+ \<not> caught P pc\<^sub>1 h\<^sub>2 xa (compxE\<^sub>2 (e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2) pc (size vs)) \<and>
+ (\<exists>vs'. P \<turnstile> (None,h\<^sub>0,frs',sh\<^sub>0) -jvm\<rightarrow> handle P C M xa h\<^sub>2 (vs'@vs) ls\<^sub>2 pc\<^sub>1 ics frs sh\<^sub>2)))"
+ (is "?N \<and> (?eq \<longrightarrow> (\<exists>pc\<^sub>2. ?H pc\<^sub>2))")
+ proof
+ show ?N by simp
+ next
+ { assume ?eq
+ then obtain pc\<^sub>2 vs' where
+ pc\<^sub>2: "?pc \<le> pc\<^sub>2 \<and> pc\<^sub>2 < ?pc + size(compE\<^sub>2 e\<^sub>2) \<and>
+ \<not> caught P pc\<^sub>2 h\<^sub>2 xa (compxE\<^sub>2 e\<^sub>2 ?pc (size vs + 1))" and
+ 2: "P \<turnstile> ?\<sigma>\<^sub>1 -jvm\<rightarrow> handle P C M xa h\<^sub>2 (vs'@v\<^sub>1#vs) ls\<^sub>2 pc\<^sub>2 ics frs sh\<^sub>2"
+ using IH\<^sub>2 Jcc_pieces_BinOp2[OF assms(1) pcs, of h\<^sub>1 v\<^sub>1 ls\<^sub>1 sh\<^sub>1 v] BinOpThrow\<^sub>2\<^sub>1.prems by clarsimp
+ then have "?H pc\<^sub>2" using jvm_trans[OF 1 2] by(auto intro!: exI[where x="vs'@[v\<^sub>1]"])
+ hence "\<exists>pc\<^sub>2. ?H pc\<^sub>2" by iprover
+ }
+ thus "?eq \<longrightarrow> (\<exists>pc\<^sub>2. ?H pc\<^sub>2)" by iprover
+ qed
+ then show ?case using pcs by simp blast
+next
+ case (FAcc\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 a h\<^sub>1 ls\<^sub>1 sh\<^sub>1 C' fs F T D w)
+ then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (e\<bullet>F{D})
+ = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (e\<bullet>F{D})),ics)#frs,sh\<^sub>1), err)"
+ using FAcc\<^sub>1.prems(1) by clarsimp
+ have "P\<^sub>1 \<turnstile> D sees F,NonStatic:T in D" by(rule has_field_sees[OF has_field_idemp[OF FAcc\<^sub>1.hyps(4)]])
+ then have field: "field P D F = (D,NonStatic,T)" by simp
+ have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (addr a) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Addr a) xa vs frs I" by fact
+ let ?pc = "pc + length(compE\<^sub>2 e)"
+ have "P \<turnstile> (None,h\<^sub>0,frs',sh\<^sub>0) -jvm\<rightarrow> (None,h\<^sub>1,(Addr a#vs,ls\<^sub>1,C,M,?pc,ics)#frs,sh\<^sub>1)"
+ using IH Jcc_pieces_FAcc[OF pcs, of "Addr a"] pcs by simp
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> (None,h\<^sub>1,(w#vs,ls\<^sub>1,C,M,?pc+1,ics)#frs,sh\<^sub>1)"
+ using FAcc\<^sub>1 field by auto
+ finally have "P \<turnstile> (None, h\<^sub>0, frs', sh\<^sub>0) -jvm\<rightarrow> (None,h\<^sub>1,(w#vs,ls\<^sub>1,C,M,?pc+1,ics)#frs,sh\<^sub>1)"
+ by auto
+ then show ?case using pcs by auto
+next
+ case (FAccNull\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 F D)
+ then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (e\<bullet>F{D})
+ = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (e\<bullet>F{D})),ics)#frs,sh\<^sub>1), err)"
+ using FAccNull\<^sub>1.prems(1) by clarsimp
+ have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 null h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics Null xa vs frs I" by fact
+ let ?pc = "pc + length(compE\<^sub>2 e)"
+ let ?xa = "addr_of_sys_xcpt NullPointer"
+ have "P \<turnstile> (None,h\<^sub>0,frs',sh\<^sub>0) -jvm\<rightarrow> (None,h\<^sub>1,(Null#vs,ls\<^sub>1,C,M,?pc,ics)#frs,sh\<^sub>1)"
+ using IH Jcc_pieces_FAcc[OF pcs, of Null] by simp
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> handle P C M ?xa h\<^sub>1 (Null#vs) ls\<^sub>1 ?pc ics frs sh\<^sub>1"
+ using FAccNull\<^sub>1.prems
+ by(fastforce simp:split_beta handle_def simp del: split_paired_Ex)
+ finally show ?case using pcs by (auto intro!: exI[where x = ?pc] exI[where x="[Null]"])
+next
+ case (FAccThrow\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 e' h\<^sub>1 ls\<^sub>1 sh\<^sub>1 F D)
+ then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (e\<bullet>F{D})
+ = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (e\<bullet>F{D})),ics)#frs,sh\<^sub>1), err)"
+ using FAccThrow\<^sub>1.prems(1) by clarsimp
+ have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (throw e') h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics v xa vs frs I" by fact
+ show ?case using IH Jcc_pieces_FAcc[OF pcs, of v] FAccThrow\<^sub>1.prems nsub_RI_Jcc_pieces
+ less_Suc_eq by auto
+next
+ case (FAccNone\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 a h\<^sub>1 ls\<^sub>1 sh\<^sub>1 C fs F D)
+ then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (e\<bullet>F{D})
+ = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (e\<bullet>F{D})),ics)#frs,sh\<^sub>1), err)"
+ using FAccNone\<^sub>1.prems(1) by clarsimp
+ have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (addr a) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Addr a) xa vs frs I" by fact
+ let ?pc = "pc + length(compE\<^sub>2 e)"
+ let ?xa = "addr_of_sys_xcpt NoSuchFieldError"
+ have "P \<turnstile> (None,h\<^sub>0,frs',sh\<^sub>0) -jvm\<rightarrow> (None,h\<^sub>1,(Addr a#vs,ls\<^sub>1,C,M,?pc,ics)#frs,sh\<^sub>1)"
+ using IH Jcc_pieces_FAcc[OF pcs, of "Addr a"] by simp
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> handle P C M ?xa h\<^sub>1 (Addr a#vs) ls\<^sub>1 ?pc ics frs sh\<^sub>1"
+ using FAccNone\<^sub>1
+ by(cases ics; clarsimp simp:split_beta handle_def simp del: split_paired_Ex)
+ finally show ?case using pcs by (auto intro!: exI[where x = ?pc] exI[where x="[Addr a]"])
+next
+ case (FAccStatic\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 a h\<^sub>1 ls\<^sub>1 sh\<^sub>1 C' fs F T D)
+ then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (e\<bullet>F{D})
+ = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (e\<bullet>F{D})),ics)#frs,sh\<^sub>1), err)"
+ using FAccStatic\<^sub>1.prems(1) by clarsimp
+ have "P\<^sub>1 \<turnstile> D sees F,Static:T in D" by(rule has_field_sees[OF has_field_idemp[OF FAccStatic\<^sub>1.hyps(4)]])
+ then have field: "field P D F = (D,Static,T)" by simp
+ have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (addr a) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Addr a) xa vs frs I" by fact
+ let ?pc = "pc + length(compE\<^sub>2 e)"
+ let ?xa = "addr_of_sys_xcpt IncompatibleClassChangeError"
+ have "P \<turnstile> (None,h\<^sub>0,frs',sh\<^sub>0) -jvm\<rightarrow> (None,h\<^sub>1,(Addr a#vs,ls\<^sub>1,C,M,?pc,ics)#frs,sh\<^sub>1)"
+ using IH Jcc_pieces_FAcc[OF pcs, of "Addr a"] by simp
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> handle P C M ?xa h\<^sub>1 (Addr a#vs) ls\<^sub>1 ?pc ics frs sh\<^sub>1"
+ using FAccStatic\<^sub>1 field by(fastforce simp:split_beta handle_def simp del: split_paired_Ex)
+ finally show ?case using pcs by (auto intro!: exI[where x = ?pc] exI[where x="[Addr a]"])
+next
+ case (SFAcc\<^sub>1 C' F t D sh sfs v' h ls)
+ have has: "P\<^sub>1 \<turnstile> D has F,Static:t in D" by(rule has_field_idemp[OF SFAcc\<^sub>1.hyps(1)])
+ have "P\<^sub>1 \<turnstile> D sees F,Static:t in D" by(rule has_field_sees[OF has])
+ then have field: "field P D F = (D,Static,t)" by simp
+ then have "P \<turnstile> (None,h,Jcc_frames P C M vs ls pc ics frs (C'\<bullet>\<^sub>sF{D}),sh) -jvm\<rightarrow>
+ (None,h,(v'#vs,ls,C,M,Suc pc,ics)#frs,sh)"
+ using SFAcc\<^sub>1 has by(cases ics) auto
+ then show ?case by clarsimp
+next
+ case (SFAccInit\<^sub>1 C' F t D sh h ls v' h' ls' sh' sfs i v'')
+ then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs ls pc ics frs sh I h' ls' sh' v xa (C'\<bullet>\<^sub>sF{D})
+ = (True, frs', (None,h',(v#vs,ls',C,M,pc+size(compE\<^sub>2 (C'\<bullet>\<^sub>sF{D})),ics)#frs,sh'), err)"
+ using SFAccInit\<^sub>1.prems(1) by clarsimp
+ have "Ex (WTrt2\<^sub>1 P\<^sub>1 E h sh (INIT D ([D],False) \<leftarrow> unit))"
+ using has_field_is_class'[OF SFAccInit\<^sub>1.hyps(1)] by auto
+ then obtain err' where pcs':
+ "Jcc_pieces P\<^sub>1 E C M h vs ls pc ics frs sh I h' ls' sh' v' xa (INIT D ([D],False) \<leftarrow> unit)
+ = (True, (vs,ls,C,M,pc,Calling D []) # frs, (None,h',(vs,ls,C,M,pc,Called [])#frs,sh'), err')"
+ using SFAccInit\<^sub>1.prems(1) by auto
+ have IH: "PROP ?P (INIT D ([D],False) \<leftarrow> unit) h ls sh (Val v')
+ h' ls' sh' E C M pc ics v' xa vs frs I" by fact
+ have ls: "ls = ls'" by(rule init\<^sub>1_same_loc[OF SFAccInit\<^sub>1.hyps(3)])
+ have has: "P\<^sub>1 \<turnstile> D has F,Static:t in D" by(rule has_field_idemp[OF SFAccInit\<^sub>1.hyps(1)])
+ have "P\<^sub>1 \<turnstile> D sees F,Static:t in D" by(rule has_field_sees[OF has])
+ then have field: "field P D F = (D,Static,t)" by simp
+ have "P \<turnstile> (None,h,(vs,ls,C,M,pc,ics)#frs,sh) -jvm\<rightarrow> (None,h,(vs,ls,C,M,pc,Calling D [])#frs,sh)"
+ proof(cases "sh D")
+ case None then show ?thesis using SFAccInit\<^sub>1.hyps(1,2,5,6) SFAccInit\<^sub>1.prems field
+ by(cases ics) auto
+ next
+ case (Some a)
+ then obtain sfs i where "a = (sfs,i)" by(cases a)
+ then show ?thesis using SFAccInit\<^sub>1.hyps(1,2,5,6) SFAccInit\<^sub>1.prems field Some
+ by(cases ics; case_tac i) auto
+ qed
+ also have "P \<turnstile> ... -jvm\<rightarrow> (None, h', (vs, ls, C, M, pc, Called []) # frs, sh')"
+ using IH pcs' by auto
+ also have "P \<turnstile> ... -jvm\<rightarrow> (None, h', (v''#vs, ls, C, M, Suc pc, ics) # frs, sh')"
+ using SFAccInit\<^sub>1.hyps(1,2,5,6) SFAccInit\<^sub>1.prems has field by(cases ics) auto
+ finally show ?case using pcs ls by clarsimp
+next
+ case (SFAccInitThrow\<^sub>1 C' F t D sh h ls a h' ls' sh')
+ obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs ls pc ics frs sh I h' ls' sh' v xa (C'\<bullet>\<^sub>sF{D})
+ = (True, frs', (None,h',(v#vs,ls',C,M,pc+size(compE\<^sub>2 (C'\<bullet>\<^sub>sF{D})),ics)#frs,sh'), err)"
+ using SFAccInitThrow\<^sub>1.prems(1) by clarsimp
+ obtain a' where throw: "throw a = Throw a'" using eval\<^sub>1_final[OF SFAccInitThrow\<^sub>1.hyps(3)] by clarsimp
+ have "Ex (WTrt2\<^sub>1 P\<^sub>1 E h sh (INIT D ([D],False) \<leftarrow> unit))"
+ using has_field_is_class'[OF SFAccInitThrow\<^sub>1.hyps(1)] by auto
+ then obtain vs' where pcs':
+ "Jcc_pieces P\<^sub>1 E C M h vs ls pc ics frs sh I h' ls' sh' v a' (INIT D ([D],False) \<leftarrow> unit)
+ = (True, (vs,ls,C,M,pc,Calling D []) # frs, (None,h',(vs,ls,C,M,pc,Called [])#frs,sh'),
+ P \<turnstile> (None,h,(vs,ls,C,M,pc,Calling D []) # frs,sh)
+ -jvm\<rightarrow> handle P C M a' h' (vs'@vs) ls pc ics frs sh')"
+ using SFAccInitThrow\<^sub>1.prems(1) by simp blast
+ have IH: "PROP ?P (INIT D ([D],False) \<leftarrow> unit) h ls sh (throw a)
+ h' ls' sh' E C M pc ics v a' vs frs I" by fact
+ have ls: "ls = ls'" by(rule init\<^sub>1_same_loc[OF SFAccInitThrow\<^sub>1.hyps(3)])
+ have has: "P\<^sub>1 \<turnstile> D has F,Static:t in D" by(rule has_field_idemp[OF SFAccInitThrow\<^sub>1.hyps(1)])
+ have "P\<^sub>1 \<turnstile> D sees F,Static:t in D" by(rule has_field_sees[OF has])
+ then have field: "field P D F = (D,Static,t)" by simp
+ then have "P \<turnstile> (None,h,(vs,ls,C,M,pc,ics)#frs,sh) -jvm\<rightarrow> (None,h,(vs,ls,C,M,pc,Calling D []) # frs,sh)"
+ proof(cases "sh D")
+ case None then show ?thesis using SFAccInitThrow\<^sub>1.hyps(1,2) SFAccInitThrow\<^sub>1.prems field
+ by(cases ics) auto
+ next
+ case (Some a)
+ then obtain sfs i where "a = (sfs,i)" by(cases a)
+ then show ?thesis using SFAccInitThrow\<^sub>1.hyps(1,2) SFAccInitThrow\<^sub>1.prems field Some
+ by(cases ics; case_tac i) auto
+ qed
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> handle P C M a' h' (vs'@vs) ls pc ics frs sh'"
+ using IH pcs' throw by auto
+ finally show ?case using throw ls by auto
+next
+ case (SFAccNone\<^sub>1 C' F D h\<^sub>1 ls\<^sub>1 sh\<^sub>1)
+ then obtain frs' err where pcs:
+ "Jcc_pieces P\<^sub>1 E C M h\<^sub>1 vs ls\<^sub>1 pc ics frs sh\<^sub>1 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (C'\<bullet>\<^sub>sF{D})
+ = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (C'\<bullet>\<^sub>sF{D})),ics)#frs,sh\<^sub>1), err)"
+ by clarsimp
+ let ?xa = "addr_of_sys_xcpt NoSuchFieldError"
+ have "P \<turnstile> (None,h\<^sub>1,frs',sh\<^sub>1) -jvm\<rightarrow> handle P C M ?xa h\<^sub>1 vs ls\<^sub>1 pc ics frs sh\<^sub>1"
+ using SFAccNone\<^sub>1 pcs
+ by(cases ics; clarsimp simp:split_beta handle_def simp del: split_paired_Ex)
+ then show ?case using pcs by(auto intro!: exI[where x = pc] exI[where x="[]"])
+next
+ case (SFAccNonStatic\<^sub>1 C' F t D h\<^sub>1 ls\<^sub>1 sh\<^sub>1)
+ let ?frs' = "(vs, ls\<^sub>1, C, M, pc, ics) # frs"
+ let ?xa = "addr_of_sys_xcpt IncompatibleClassChangeError"
+ have "P\<^sub>1 \<turnstile> D sees F,NonStatic:t in D"
+ by(rule has_field_sees[OF has_field_idemp[OF SFAccNonStatic\<^sub>1.hyps(1)]])
+ then have field: "field P D F = (D,NonStatic,t)" by simp
+ have "P \<turnstile> (None,h\<^sub>1,?frs',sh\<^sub>1) -jvm\<rightarrow> handle P C M ?xa h\<^sub>1 vs ls\<^sub>1 pc ics frs sh\<^sub>1"
+ using SFAccNonStatic\<^sub>1
+ proof(cases ics)
+ case No_ics
+ then show ?thesis using SFAccNonStatic\<^sub>1 field
+ by (auto simp:split_beta handle_def simp del: split_paired_Ex)
+ qed(simp_all)
+ then show ?case by (auto intro!: exI[where x = pc] exI[where x="[]"])
+next
+ case (LAss\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 w h\<^sub>1 ls\<^sub>1 sh\<^sub>1 i ls\<^sub>2)
+ let ?pc = "pc + length(compE\<^sub>2 e)"
+ obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (i:=e)
+ = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (i:=e)),ics)#frs,sh\<^sub>1), err)"
+ using LAss\<^sub>1.prems(1) by auto
+ have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (Val w) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics w xa vs frs I" by fact
+ then have "P \<turnstile> (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0) -jvm\<rightarrow>
+ (None,h\<^sub>1,(w#vs,ls\<^sub>1,C,M,?pc,ics)#frs,sh\<^sub>1)"
+ using Jcc_pieces_LAss[OF assms(1) pcs, of w] LAss\<^sub>1.prems pcs by auto
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> (None,h\<^sub>1,(Unit#vs,ls\<^sub>2,C,M,?pc+2,ics)#frs,sh\<^sub>1)"
+ using LAss\<^sub>1 by (auto simp add:cast_ok_def)
+ finally show ?case by auto
+next
+ case (LAssThrow\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 w h\<^sub>1 ls\<^sub>1 sh\<^sub>1 i)
+ obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (i:=e)
+ = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (i:=e)),ics)#frs,sh\<^sub>1), err)"
+ using LAssThrow\<^sub>1.prems(1) by clarsimp
+ have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (throw w) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics v xa vs frs I" by fact
+ show ?case using IH Jcc_pieces_LAss[OF assms(1) pcs, of v] LAssThrow\<^sub>1.prems pcs less_SucI
+ by(simp, blast)
+next
+ case (FAss\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 a h\<^sub>1 ls\<^sub>1 sh\<^sub>1 e\<^sub>2 w h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C' fs F T D fs' h\<^sub>2')
+ obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v xa (e\<^sub>1\<bullet>F{D} := e\<^sub>2)
+ = (True, frs', (None,h\<^sub>2,(v#vs,ls\<^sub>2,C,M,pc+size(compE\<^sub>2 (e\<^sub>1\<bullet>F{D} := e\<^sub>2)),ics)#frs,sh\<^sub>2), err)"
+ using FAss\<^sub>1.prems(1) by clarsimp
+ have "P\<^sub>1 \<turnstile> D sees F,NonStatic:T in D" by(rule has_field_sees[OF has_field_idemp[OF FAss\<^sub>1.hyps(6)]])
+ then have field: "field P D F = (D,NonStatic,T)" by simp
+ let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e\<^sub>1)"
+ let ?pc\<^sub>2 = "?pc\<^sub>1 + length(compE\<^sub>2 e\<^sub>2)"
+ have IH\<^sub>1: "PROP ?P e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (addr a) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Addr a) xa vs frs
+ (I - pcs (compxE\<^sub>2 e\<^sub>2 (pc + length (compE\<^sub>2 e\<^sub>1)) (Suc (length vs))))" by fact
+ have IH\<^sub>2: "PROP ?P e\<^sub>2 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (Val w) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 E C M ?pc\<^sub>1 ics w xa (Addr a#vs) frs
+ (I - pcs(compxE\<^sub>2 e\<^sub>1 pc (size vs)))" by fact
+ have "P \<turnstile> (None,h\<^sub>0,frs',sh\<^sub>0) -jvm\<rightarrow> (None,h\<^sub>1,(Addr a#vs,ls\<^sub>1,C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)"
+ using IH\<^sub>1 Jcc_pieces_FAss1[OF pcs, of h\<^sub>1 ls\<^sub>1 sh\<^sub>1 "Addr a"] by simp
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> (None,h\<^sub>2,(w#Addr a#vs,ls\<^sub>2,C,M,?pc\<^sub>2,ics)#frs,sh\<^sub>2)"
+ using IH\<^sub>2 Jcc_pieces_FAss2[OF pcs, of h\<^sub>1 "Addr a" ls\<^sub>1 sh\<^sub>1 w] by (simp add: add.assoc)
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> (None,h\<^sub>2',(Unit#vs,ls\<^sub>2,C,M,?pc\<^sub>2+2,ics)#frs,sh\<^sub>2)"
+ using FAss\<^sub>1 field by auto
+ finally show ?case using pcs by (auto simp:add.assoc)
+next
+ case (FAssNull\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 e\<^sub>2 w h\<^sub>2 ls\<^sub>2 sh\<^sub>2 F D)
+ obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v xa (e\<^sub>1\<bullet>F{D} := e\<^sub>2)
+ = (True, frs', (None,h\<^sub>2,(v#vs,ls\<^sub>2,C,M,pc+size(compE\<^sub>2 (e\<^sub>1\<bullet>F{D} := e\<^sub>2)),ics)#frs,sh\<^sub>2), err)"
+ using FAssNull\<^sub>1.prems(1) by clarsimp
+ let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e\<^sub>1)"
+ let ?pc\<^sub>2 = "?pc\<^sub>1 + length(compE\<^sub>2 e\<^sub>2)"
+ let ?xa = "addr_of_sys_xcpt NullPointer"
+ have IH\<^sub>1: "PROP ?P e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 null h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics Null xa vs frs
+ (I - pcs (compxE\<^sub>2 e\<^sub>2 (pc + length (compE\<^sub>2 e\<^sub>1)) (Suc (length vs))))" by fact
+ have IH\<^sub>2: "PROP ?P e\<^sub>2 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (Val w) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 E C M ?pc\<^sub>1 ics w xa (Null#vs) frs
+ (I - pcs(compxE\<^sub>2 e\<^sub>1 pc (size vs)))" by fact
+ have "P \<turnstile> (None,h\<^sub>0,frs',sh\<^sub>0) -jvm\<rightarrow> (None,h\<^sub>1,(Null#vs,ls\<^sub>1,C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)"
+ using IH\<^sub>1 Jcc_pieces_FAss1[OF pcs, of h\<^sub>1 ls\<^sub>1 sh\<^sub>1 Null] by simp
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> (None,h\<^sub>2,(w#Null#vs,ls\<^sub>2,C,M,?pc\<^sub>2,ics)#frs,sh\<^sub>2)"
+ using IH\<^sub>2 Jcc_pieces_FAss2[OF pcs, of h\<^sub>1 Null ls\<^sub>1 sh\<^sub>1 w] by (simp add: add.assoc)
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> handle P C M ?xa h\<^sub>2 (w#Null#vs) ls\<^sub>2 ?pc\<^sub>2 ics frs sh\<^sub>2"
+ using FAssNull\<^sub>1 by(fastforce simp:split_beta handle_def simp del: split_paired_Ex)
+ finally show ?case using pcs by (auto intro!: exI[where x = ?pc\<^sub>2] exI[where x="w#[Null]"])
+next
+ case (FAssThrow\<^sub>2\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 w h\<^sub>1 ls\<^sub>1 sh\<^sub>1 e\<^sub>2 e' h\<^sub>2 ls\<^sub>2 sh\<^sub>2 F D)
+ let ?frs' = "(vs, ls\<^sub>0, C, M, pc, ics) # frs"
+ obtain err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v xa (e\<^sub>1\<bullet>F{D} := e\<^sub>2)
+ = (True, ?frs', (None,h\<^sub>2,(v#vs,ls\<^sub>2,C,M,pc+size(compE\<^sub>2 (e\<^sub>1\<bullet>F{D} := e\<^sub>2)),ics)#frs,sh\<^sub>2), err)"
+ using FAssThrow\<^sub>2\<^sub>1.prems(1) by clarsimp
+ let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e\<^sub>1)"
+ let ?\<sigma>\<^sub>1 = "(None,h\<^sub>1,(w#vs,ls\<^sub>1,C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)"
+ have IH\<^sub>1: "PROP ?P e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (Val w) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics w xa vs frs
+ (I - pcs (compxE\<^sub>2 e\<^sub>2 (pc + length (compE\<^sub>2 e\<^sub>1)) (Suc (length vs))))" by fact
+ have IH\<^sub>2: "PROP ?P e\<^sub>2 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (throw e') h\<^sub>2 ls\<^sub>2 sh\<^sub>2 E C M ?pc\<^sub>1 ics v xa (w#vs) frs
+ (I - pcs(compxE\<^sub>2 e\<^sub>1 pc (size vs)))" by fact
+ have 1: "P \<turnstile> (None,h\<^sub>0,?frs',sh\<^sub>0) -jvm\<rightarrow> ?\<sigma>\<^sub>1"
+ using IH\<^sub>1 Jcc_pieces_FAss1[OF pcs, of h\<^sub>1 ls\<^sub>1 sh\<^sub>1 w] by simp
+ show ?case (is "?N \<and> (?eq \<longrightarrow> ?err)")
+ proof
+ show ?N by simp
+ next
+ { assume ?eq
+ moreover
+ have "PROP ?P e\<^sub>2 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (throw e') h\<^sub>2 ls\<^sub>2 sh\<^sub>2 E C M ?pc\<^sub>1 ics v xa (w#vs) frs
+ (I - pcs (compxE\<^sub>2 e\<^sub>1 pc (length vs)))" by fact
+ ultimately obtain pc\<^sub>2 vs' where
+ pc\<^sub>2: "?pc\<^sub>1 \<le> pc\<^sub>2 \<and> pc\<^sub>2 < ?pc\<^sub>1 + size(compE\<^sub>2 e\<^sub>2) \<and>
+ \<not> caught P pc\<^sub>2 h\<^sub>2 xa (compxE\<^sub>2 e\<^sub>2 ?pc\<^sub>1 (size vs + 1))" and
+ 2: "P \<turnstile> ?\<sigma>\<^sub>1 -jvm\<rightarrow> handle P C M xa h\<^sub>2 (vs'@w#vs) ls\<^sub>2 pc\<^sub>2 ics frs sh\<^sub>2"
+ using FAssThrow\<^sub>2\<^sub>1.prems Jcc_pieces_FAss2[OF pcs, of h\<^sub>1 w ls\<^sub>1 sh\<^sub>1] by auto
+ have ?err using Jcc_pieces_FAss2[OF pcs, of h\<^sub>1 w ls\<^sub>1 sh\<^sub>1] pc\<^sub>2 jvm_trans[OF 1 2]
+ by(auto intro!: exI[where x=pc\<^sub>2] exI[where x="vs'@[w]"])
+ }
+ thus "?eq \<longrightarrow> ?err" by simp
+ qed
+next
+ case (FAssThrow\<^sub>1\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 e' h\<^sub>1 ls\<^sub>1 sh\<^sub>1 F D e\<^sub>2)
+ obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (e\<^sub>1\<bullet>F{D} := e\<^sub>2)
+ = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (e\<^sub>1\<bullet>F{D} := e\<^sub>2)),ics)#frs,sh\<^sub>1), err)"
+ using FAssThrow\<^sub>1\<^sub>1.prems(1) by clarsimp
+ have IH\<^sub>1: "PROP ?P e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (throw e') h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics v xa vs frs
+ (I - pcs (compxE\<^sub>2 e\<^sub>2 (pc + length (compE\<^sub>2 e\<^sub>1)) (Suc (length vs))))" by fact
+ show ?case using IH\<^sub>1 Jcc_pieces_FAss1[OF pcs, of h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v] FAssThrow\<^sub>1\<^sub>1.prems nsub_RI_Jcc_pieces
+ by auto
+next
+ case (FAssNone\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 a h\<^sub>1 ls\<^sub>1 sh\<^sub>1 e\<^sub>2 w h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C' fs F D)
+ obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v xa (e\<^sub>1\<bullet>F{D} := e\<^sub>2)
+ = (True, frs', (None,h\<^sub>2,(v#vs,ls\<^sub>2,C,M,pc+size(compE\<^sub>2 (e\<^sub>1\<bullet>F{D} := e\<^sub>2)),ics)#frs,sh\<^sub>2), err)"
+ using FAssNone\<^sub>1.prems(1) by clarsimp
+ let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e\<^sub>1)"
+ let ?pc\<^sub>2 = "?pc\<^sub>1 + length(compE\<^sub>2 e\<^sub>2)"
+ let ?xa = "addr_of_sys_xcpt NoSuchFieldError"
+ have IH\<^sub>1: "PROP ?P e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (addr a) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Addr a) xa vs frs
+ (I - pcs (compxE\<^sub>2 e\<^sub>2 (pc + length (compE\<^sub>2 e\<^sub>1)) (Suc (length vs))))" by fact
+ have IH\<^sub>2: "PROP ?P e\<^sub>2 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (Val w) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 E C M ?pc\<^sub>1 ics w xa (Addr a#vs) frs
+ (I - pcs(compxE\<^sub>2 e\<^sub>1 pc (size vs)))" by fact
+ have "P \<turnstile> (None,h\<^sub>0,frs',sh\<^sub>0) -jvm\<rightarrow> (None,h\<^sub>1,(Addr a#vs,ls\<^sub>1,C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)"
+ using IH\<^sub>1 Jcc_pieces_FAss1[OF pcs, of h\<^sub>1 ls\<^sub>1 sh\<^sub>1 "Addr a"] by simp
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> (None,h\<^sub>2,(w#Addr a#vs,ls\<^sub>2,C,M,?pc\<^sub>2,ics)#frs,sh\<^sub>2)"
+ using IH\<^sub>2 Jcc_pieces_FAss2[OF pcs, of h\<^sub>1 "Addr a" ls\<^sub>1 sh\<^sub>1 w] by (simp add: add.assoc)
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> handle P C M ?xa h\<^sub>2 (w#Addr a#vs) ls\<^sub>2 ?pc\<^sub>2 ics frs sh\<^sub>2"
+ using FAssNone\<^sub>1 by(fastforce simp:split_beta handle_def simp del: split_paired_Ex)
+ finally show ?case using pcs by (auto intro!: exI[where x = ?pc\<^sub>2] exI[where x="w#[Addr a]"])
+next
+ case (FAssStatic\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 a h\<^sub>1 ls\<^sub>1 sh\<^sub>1 e\<^sub>2 w h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C' fs F T D)
+ obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v xa (e\<^sub>1\<bullet>F{D} := e\<^sub>2)
+ = (True, frs', (None,h\<^sub>2,(v#vs,ls\<^sub>2,C,M,pc+size(compE\<^sub>2 (e\<^sub>1\<bullet>F{D} := e\<^sub>2)),ics)#frs,sh\<^sub>2), err)"
+ using FAssStatic\<^sub>1.prems(1) by clarsimp
+ have "P\<^sub>1 \<turnstile> D sees F,Static:T in D" by(rule has_field_sees[OF has_field_idemp[OF FAssStatic\<^sub>1.hyps(6)]])
+ then have field: "field P D F = (D,Static,T)" by simp
+ let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e\<^sub>1)"
+ let ?pc\<^sub>2 = "?pc\<^sub>1 + length(compE\<^sub>2 e\<^sub>2)"
+ let ?xa = "addr_of_sys_xcpt IncompatibleClassChangeError"
+ have IH\<^sub>1: "PROP ?P e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (addr a) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Addr a) xa vs frs
+ (I - pcs (compxE\<^sub>2 e\<^sub>2 (pc + length (compE\<^sub>2 e\<^sub>1)) (Suc (length vs))))" by fact
+ have IH\<^sub>2: "PROP ?P e\<^sub>2 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (Val w) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 E C M ?pc\<^sub>1 ics w xa (Addr a#vs) frs
+ (I - pcs(compxE\<^sub>2 e\<^sub>1 pc (size vs)))" by fact
+ have "P \<turnstile> (None,h\<^sub>0,frs',sh\<^sub>0) -jvm\<rightarrow> (None,h\<^sub>1,(Addr a#vs,ls\<^sub>1,C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)"
+ using IH\<^sub>1 Jcc_pieces_FAss1[OF pcs, of h\<^sub>1 ls\<^sub>1 sh\<^sub>1 "Addr a"] by simp
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> (None,h\<^sub>2,(w#Addr a#vs,ls\<^sub>2,C,M,?pc\<^sub>2,ics)#frs,sh\<^sub>2)"
+ using IH\<^sub>2 Jcc_pieces_FAss2[OF pcs, of h\<^sub>1 "Addr a" ls\<^sub>1 sh\<^sub>1 w] by (simp add: add.assoc)
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> handle P C M ?xa h\<^sub>2 (w#Addr a#vs) ls\<^sub>2 ?pc\<^sub>2 ics frs sh\<^sub>2"
+ using FAssStatic\<^sub>1 field by(fastforce simp:split_beta handle_def simp del: split_paired_Ex)
+ finally show ?case using pcs by (auto intro!: exI[where x = ?pc\<^sub>2] exI[where x="w#[Addr a]"])
+next
+ case (SFAss\<^sub>1 e\<^sub>2 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 w h\<^sub>1 ls\<^sub>1 sh\<^sub>1 C' F T D sfs sfs' sh\<^sub>1')
+ then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (C'\<bullet>\<^sub>sF{D} := e\<^sub>2)
+ = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (C'\<bullet>\<^sub>sF{D} := e\<^sub>2)),ics)#frs,sh\<^sub>1), err)"
+ using SFAss\<^sub>1.prems(1) by clarsimp
+ have "P\<^sub>1 \<turnstile> D sees F,Static:T in D" by(rule has_field_sees[OF has_field_idemp[OF SFAss\<^sub>1.hyps(3)]])
+ then have field: "field P D F = (D,Static,T)" by simp
+ have IH: "PROP ?P e\<^sub>2 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (Val w) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics w xa vs frs I" by fact
+ let ?pc = "pc + length(compE\<^sub>2 e\<^sub>2)"
+ have "P \<turnstile> (None,h\<^sub>0,frs',sh\<^sub>0) -jvm\<rightarrow> (None,h\<^sub>1,(w#vs,ls\<^sub>1,C,M,?pc,ics)#frs,sh\<^sub>1)"
+ using IH Jcc_pieces_SFAss[OF pcs, where v'=w] pcs by simp
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> (None,h\<^sub>1,(vs,ls\<^sub>1,C,M,?pc+1,ics)#frs,sh\<^sub>1')"
+ using SFAss\<^sub>1.hyps(3-6) SFAss\<^sub>1.prems(1) field by auto
+ also have "P \<turnstile> ... -jvm\<rightarrow> (None,h\<^sub>1,(Unit#vs,ls\<^sub>1,C,M,?pc+2,ics)#frs,sh\<^sub>1')"
+ using SFAss\<^sub>1 by auto
+ finally show ?case using pcs by auto
+next
+ case (SFAssInit\<^sub>1 e\<^sub>2 h ls sh w h\<^sub>1 ls\<^sub>1 sh\<^sub>1 C' F t D v' h' ls' sh' sfs i sfs' sh'')
+ let ?pc = "pc + length(compE\<^sub>2 e\<^sub>2)"
+ obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs ls pc ics frs sh I h' ls' sh'' v xa (C'\<bullet>\<^sub>sF{D}:=e\<^sub>2)
+ = (True, frs', (None,h',(v#vs,ls',C,M,pc+size(compE\<^sub>2 (C'\<bullet>\<^sub>sF{D}:=e\<^sub>2)),ics)#frs,sh''), err)"
+ using SFAssInit\<^sub>1.prems(1) by clarsimp
+ have "Ex (WTrt2\<^sub>1 P\<^sub>1 E h\<^sub>1 sh\<^sub>1 (INIT D ([D],False) \<leftarrow> unit))"
+ using has_field_is_class'[OF SFAssInit\<^sub>1.hyps(3)] by auto
+ then obtain err' where pcs':
+ "Jcc_pieces P\<^sub>1 E C M h\<^sub>1 (w#vs) ls\<^sub>1 ?pc ics frs sh\<^sub>1 I h' ls' sh' v' xa (INIT D ([D],False) \<leftarrow> unit)
+ = (True, (w#vs,ls\<^sub>1,C,M,?pc,Calling D []) # frs,
+ (None,h',(w#vs,ls\<^sub>1,C,M,?pc,Called [])#frs,sh'), err')"
+ using SFAssInit\<^sub>1.prems(1) by simp
+ have ls: "ls\<^sub>1 = ls'" by(rule init\<^sub>1_same_loc[OF SFAssInit\<^sub>1.hyps(5)])
+ have has: "P\<^sub>1 \<turnstile> D has F,Static:t in D" by(rule has_field_idemp[OF SFAssInit\<^sub>1.hyps(3)])
+ have "P\<^sub>1 \<turnstile> D sees F,Static:t in D" by(rule has_field_sees[OF has])
+ then have field: "field P D F = (D,Static,t)" by simp
+ have IH: "PROP ?P e\<^sub>2 h ls sh (Val w) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics w xa vs frs I" by fact
+ have IHI: "PROP ?P (INIT D ([D],False) \<leftarrow> unit) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (Val v')
+ h' ls' sh' E C M ?pc ics v' xa (w#vs) frs I" by fact
+ have "P \<turnstile> (None,h,frs',sh) -jvm\<rightarrow> (None,h\<^sub>1,(w#vs,ls\<^sub>1,C,M,?pc,ics)#frs,sh\<^sub>1)"
+ using IH Jcc_pieces_SFAss[OF pcs, where v'=w] by simp
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> (None,h\<^sub>1,(w#vs,ls\<^sub>1,C,M,?pc,Calling D [])#frs,sh\<^sub>1)"
+ proof(cases "sh\<^sub>1 D")
+ case None then show ?thesis using None SFAssInit\<^sub>1.hyps(1,3-5,7-9) SFAssInit\<^sub>1.prems field
+ by(cases ics, auto)
+ next
+ case (Some a)
+ then obtain sfs i where "a = (sfs,i)" by(cases a)
+ then show ?thesis using SFAssInit\<^sub>1.hyps(1,3-5,7-9) SFAssInit\<^sub>1.prems field Some
+ by(cases ics; case_tac i) auto
+ qed
+ also have "P \<turnstile> ... -jvm\<rightarrow> (None, h', (w#vs, ls\<^sub>1, C, M, ?pc, Called []) # frs, sh')"
+ using IHI pcs' by clarsimp
+ also have "P \<turnstile> ... -jvm\<rightarrow> (None, h', (vs, ls\<^sub>1, C, M, ?pc + 1, ics) # frs, sh'')"
+ using SFAssInit\<^sub>1.hyps(1,3-5,7-9) SFAssInit\<^sub>1.prems has field by(cases ics) auto
+ also have "P \<turnstile> ... -jvm\<rightarrow> (None, h', (Unit#vs, ls\<^sub>1, C, M, ?pc + 2, ics) # frs, sh'')"
+ using SFAssInit\<^sub>1.hyps(1,3-5,7-9) SFAssInit\<^sub>1.prems has field by(cases ics) auto
+ finally show ?case using pcs ls by simp blast
+next
+ case (SFAssInitThrow\<^sub>1 e\<^sub>2 h ls sh w h\<^sub>1 ls\<^sub>1 sh\<^sub>1 C' F t D a h' ls' sh')
+ let ?pc = "pc + length(compE\<^sub>2 e\<^sub>2)"
+ obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs ls pc ics frs sh I h' ls' sh' v xa (C'\<bullet>\<^sub>sF{D}:=e\<^sub>2)
+ = (True, frs', (None,h',(v#vs,ls',C,M,pc+size(compE\<^sub>2 (C'\<bullet>\<^sub>sF{D}:=e\<^sub>2)),ics)#frs,sh'), err)"
+ using SFAssInitThrow\<^sub>1.prems(1) by clarsimp
+ obtain a' where throw: "throw a = Throw a'" using eval\<^sub>1_final[OF SFAssInitThrow\<^sub>1.hyps(5)] by clarsimp
+ have "Ex (WTrt2\<^sub>1 P\<^sub>1 E h\<^sub>1 sh\<^sub>1 (INIT D ([D],False) \<leftarrow> unit))"
+ using has_field_is_class'[OF SFAssInitThrow\<^sub>1.hyps(3)] by auto
+ then obtain vs' where pcs':
+ "Jcc_pieces P\<^sub>1 E C M h\<^sub>1 (w#vs) ls\<^sub>1 ?pc ics frs sh\<^sub>1 I h' ls' sh' v a' (INIT D ([D],False) \<leftarrow> unit)
+ = (True, (w#vs,ls\<^sub>1,C,M,?pc,Calling D []) # frs, (None,h',(w#vs,ls\<^sub>1,C,M,?pc,Called [])#frs,sh'),
+ P \<turnstile> (None,h\<^sub>1,(w#vs,ls\<^sub>1,C,M,?pc,Calling D []) # frs,sh\<^sub>1)
+ -jvm\<rightarrow> handle P C M a' h' (vs'@w#vs) ls\<^sub>1 ?pc ics frs sh')"
+ using SFAssInitThrow\<^sub>1.prems(1) by simp blast
+ have ls: "ls\<^sub>1 = ls'" by(rule init\<^sub>1_same_loc[OF SFAssInitThrow\<^sub>1.hyps(5)])
+ have has: "P\<^sub>1 \<turnstile> D has F,Static:t in D" by(rule has_field_idemp[OF SFAssInitThrow\<^sub>1.hyps(3)])
+ have "P\<^sub>1 \<turnstile> D sees F,Static:t in D" by(rule has_field_sees[OF has])
+ then have field: "field P D F = (D,Static,t)" by simp
+ have IH: "PROP ?P e\<^sub>2 h ls sh (Val w) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics w xa vs frs I" by fact
+ have IHI: "PROP ?P (INIT D ([D],False) \<leftarrow> unit) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (throw a)
+ h' ls' sh' E C M ?pc ics v a' (w#vs) frs I" by fact
+ have "P \<turnstile> (None,h,(vs, ls, C, M, pc, ics) # frs,sh) -jvm\<rightarrow> (None,h\<^sub>1,(w#vs,ls\<^sub>1,C,M,?pc,ics)#frs,sh\<^sub>1)"
+ using IH Jcc_pieces_SFAss[OF pcs, where v'=w] pcs by simp blast
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> (None,h\<^sub>1,(w#vs,ls\<^sub>1,C,M,?pc,Calling D [])#frs,sh\<^sub>1)"
+ proof(cases "sh\<^sub>1 D")
+ case None then show ?thesis using SFAssInitThrow\<^sub>1.hyps(1,3,4,5) SFAssInitThrow\<^sub>1.prems field
+ by(cases ics) auto
+ next
+ case (Some a)
+ then obtain sfs i where "a = (sfs,i)" by(cases a)
+ then show ?thesis using SFAssInitThrow\<^sub>1.hyps(1,3,4,5) SFAssInitThrow\<^sub>1.prems field Some
+ by(cases ics; case_tac i) auto
+ qed
+ also have "P \<turnstile> ... -jvm\<rightarrow> handle P C M a' h' (vs'@w#vs) ls\<^sub>1 ?pc ics frs sh'"
+ using IHI pcs' throw by auto
+ finally show ?case using throw ls by(auto intro!: exI[where x = ?pc] exI[where x="vs'@[w]"])
+next
+ case (SFAssThrow\<^sub>1 e\<^sub>2 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 e' h\<^sub>1 ls\<^sub>1 sh\<^sub>1 C' F D)
+ then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (C'\<bullet>\<^sub>sF{D} := e\<^sub>2)
+ = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (C'\<bullet>\<^sub>sF{D} := e\<^sub>2)),ics)#frs,sh\<^sub>1), err)"
+ using SFAssThrow\<^sub>1.prems(1) by clarsimp
+ have IH: "PROP ?P e\<^sub>2 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (throw e') h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics v xa vs frs I" by fact
+ show ?case using IH Jcc_pieces_SFAss[OF pcs, where v'=v] SFAssThrow\<^sub>1.prems nsub_RI_Jcc_pieces
+ less_Suc_eq by auto
+next
+ case (SFAssNone\<^sub>1 e\<^sub>2 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 w h\<^sub>1 ls\<^sub>1 sh\<^sub>1 C' F D)
+ then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (C'\<bullet>\<^sub>sF{D} := e\<^sub>2)
+ = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (C'\<bullet>\<^sub>sF{D} := e\<^sub>2)),ics)#frs,sh\<^sub>1), err)"
+ using SFAssNone\<^sub>1.prems(1) by clarsimp
+ have IH: "PROP ?P e\<^sub>2 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (Val w) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics w xa vs frs I" by fact
+ let ?pc = "pc + length(compE\<^sub>2 e\<^sub>2)"
+ let ?xa = "addr_of_sys_xcpt NoSuchFieldError"
+ have "P \<turnstile> (None,h\<^sub>0,frs',sh\<^sub>0) -jvm\<rightarrow> (None,h\<^sub>1,(w#vs,ls\<^sub>1,C,M,?pc,ics)#frs,sh\<^sub>1)"
+ using IH Jcc_pieces_SFAss[OF pcs, where v'=w] pcs by simp
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> handle P C M ?xa h\<^sub>1 (w#vs) ls\<^sub>1 ?pc ics frs sh\<^sub>1"
+ using SFAssNone\<^sub>1 by(cases ics; clarsimp simp add: handle_def)
+ finally show ?case using pcs by (auto intro!: exI[where x = ?pc] exI[where x="[w]"])
+next
+ case (SFAssNonStatic\<^sub>1 e\<^sub>2 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 w h\<^sub>1 ls\<^sub>1 sh\<^sub>1 C' F T D)
+ then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>1 ls\<^sub>1 sh\<^sub>1 v xa (C'\<bullet>\<^sub>sF{D} := e\<^sub>2)
+ = (True, frs', (None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,pc+size(compE\<^sub>2 (C'\<bullet>\<^sub>sF{D} := e\<^sub>2)),ics)#frs,sh\<^sub>1), err)"
+ using SFAssNonStatic\<^sub>1.prems(1) by clarsimp
+ have IH: "PROP ?P e\<^sub>2 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (Val w) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics w xa vs frs I" by fact
+ let ?pc = "pc + length(compE\<^sub>2 e\<^sub>2)"
+ let ?xa = "addr_of_sys_xcpt IncompatibleClassChangeError"
+ have "P\<^sub>1 \<turnstile> D sees F,NonStatic:T in D"
+ by(rule has_field_sees[OF has_field_idemp[OF SFAssNonStatic\<^sub>1.hyps(3)]])
+ then have field: "field P D F = (D,NonStatic,T)" by simp
+ have "P \<turnstile> (None,h\<^sub>0,frs',sh\<^sub>0) -jvm\<rightarrow> (None,h\<^sub>1,(w#vs,ls\<^sub>1,C,M,?pc,ics)#frs,sh\<^sub>1)"
+ using IH Jcc_pieces_SFAss[OF pcs, where v'=w] pcs by simp
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> handle P C M ?xa h\<^sub>1 (w#vs) ls\<^sub>1 ?pc ics frs sh\<^sub>1"
+ using SFAssNonStatic\<^sub>1
+ proof(cases ics)
+ case No_ics
+ then show ?thesis using SFAssNonStatic\<^sub>1 field
+ by (auto simp:split_beta handle_def simp del: split_paired_Ex)
+ qed(simp_all)
+ finally show ?case using pcs by (auto intro!: exI[where x = ?pc] exI[where x="[w]"])
+next
+ case (Call\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 a h\<^sub>1 ls\<^sub>1 sh\<^sub>1 es pvs h\<^sub>2 ls\<^sub>2 sh\<^sub>2 Ca fs M' Ts T body D ls\<^sub>2' f h\<^sub>3 ls\<^sub>3 sh\<^sub>3)
+ let ?frs\<^sub>0 = "(vs, ls\<^sub>0, C,M,pc,ics)#frs"
+ let ?\<sigma>\<^sub>0 = "(None,h\<^sub>0,?frs\<^sub>0,sh\<^sub>0)"
+ let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e)"
+ let ?\<sigma>\<^sub>1 = "(None,h\<^sub>1,(Addr a#vs, ls\<^sub>1, C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)"
+ let ?pc\<^sub>2 = "?pc\<^sub>1 + length(compEs\<^sub>2 es)"
+ let ?frs\<^sub>2 = "(rev pvs @ Addr a # vs, ls\<^sub>2, C,M,?pc\<^sub>2,ics)#frs"
+ let ?\<sigma>\<^sub>2 = "(None,h\<^sub>2,?frs\<^sub>2,sh\<^sub>2)"
+ let ?frs\<^sub>2' = "([], ls\<^sub>2', D,M',0,No_ics) # ?frs\<^sub>2"
+ let ?\<sigma>\<^sub>2' = "(None, h\<^sub>2, ?frs\<^sub>2', sh\<^sub>2)"
+ have nclinit: "M' \<noteq> clinit" using wf_sees_clinit1[OF wf] visible_method_exists[OF Call\<^sub>1.hyps(6)]
+ sees_method_idemp[OF Call\<^sub>1.hyps(6)] by fastforce
+ have "P\<^sub>1 \<turnstile>\<^sub>1 \<langle>es,(h\<^sub>1, ls\<^sub>1, sh\<^sub>1)\<rangle> [\<Rightarrow>] \<langle>map Val pvs,(h\<^sub>2, ls\<^sub>2, sh\<^sub>2)\<rangle>" by fact
+ hence [simp]: "length es = length pvs" by(auto dest:evals\<^sub>1_preserves_elen)
+ have invoke: "P,C,M,?pc\<^sub>2 \<triangleright> Invoke M' (length Ts)"
+ using Call\<^sub>1.hyps(7) Call\<^sub>1.prems(1) by clarsimp
+ have nsub: "\<not> sub_RI body" by(rule sees_wf\<^sub>1_nsub_RI[OF wf Call\<^sub>1.hyps(6)])
+ obtain err where pcs:
+ "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>3 ls\<^sub>2 sh\<^sub>3 v xa (e\<bullet>M'(es)) =
+ (True, ?frs\<^sub>0, (None, h\<^sub>3, (v#vs, ls\<^sub>2, C,M,?pc\<^sub>2+1,ics)#frs,sh\<^sub>3), err)"
+ using Call\<^sub>1.prems(1) by clarsimp
+ have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (addr a) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Addr a) xa vs frs
+ (I - pcs (compxEs\<^sub>2 es (pc + length (compE\<^sub>2 e)) (Suc (length vs))))" by fact
+ have IH_es: "PROP ?Ps es h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (map Val pvs) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C M ?pc\<^sub>1 ics pvs xa
+ (map Val pvs) (Addr a#vs) frs (I - pcs(compxE\<^sub>2 e pc (size vs)))" by fact
+ have "P \<turnstile> ?\<sigma>\<^sub>0 -jvm\<rightarrow> ?\<sigma>\<^sub>1" using Jcc_pieces_Call1[OF pcs] IH by clarsimp
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> ?\<sigma>\<^sub>2" using IH_es Call\<^sub>1.prems by fastforce
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> ?\<sigma>\<^sub>2'"
+ using jvm_Invoke[OF assms(1) invoke _ Call\<^sub>1.hyps(6-8)] Call\<^sub>1.hyps(5) Call\<^sub>1.prems(1) by simp
+ finally have 1: "P \<turnstile> ?\<sigma>\<^sub>0 -jvm\<rightarrow> ?\<sigma>\<^sub>2'".
+ have "P\<^sub>1 \<turnstile> Ca sees M',NonStatic: Ts\<rightarrow>T = body in D" by fact
+ then have M'_in_D: "P\<^sub>1 \<turnstile> D sees M',NonStatic: Ts\<rightarrow>T = body in D"
+ by(rule sees_method_idemp)
+ have M'_code: "compP\<^sub>2 P\<^sub>1,D,M',0 \<rhd> compE\<^sub>2 body @ [Return]" using beforeM M'_in_D by simp
+ have M'_xtab: "compP\<^sub>2 P\<^sub>1,D,M' \<rhd> compxE\<^sub>2 body 0 0/{..<size(compE\<^sub>2 body)},0"
+ using M'_in_D by(rule beforexM)
+ have IH_body: "PROP ?P body h\<^sub>2 ls\<^sub>2' sh\<^sub>2 f h\<^sub>3 ls\<^sub>3 sh\<^sub>3 (Class D # Ts) D M' 0 No_ics v xa [] ?frs\<^sub>2
+ ({..<size(compE\<^sub>2 body)})" by fact
+ have cond: "Jcc_cond P\<^sub>1 (Class D # Ts) D M' [] 0 No_ics {..<length (compE\<^sub>2 body)} h\<^sub>2 sh\<^sub>2 body"
+ using nsub_RI_Jcc_pieces[OF assms(1) nsub] M'_code M'_xtab by clarsimp
+ show ?case (is "?Norm \<and> ?Err")
+ proof
+ show ?Norm (is "?val \<longrightarrow> ?trans")
+ proof
+ assume val: ?val
+ note 1
+ also have "P \<turnstile> ?\<sigma>\<^sub>2' -jvm\<rightarrow> (None,h\<^sub>3,([v],ls\<^sub>3,D,M',size(compE\<^sub>2 body),No_ics)#?frs\<^sub>2,sh\<^sub>3)"
+ using val IH_body Call\<^sub>1.prems M'_code cond nsub_RI_Jcc_pieces nsub by auto
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> (None, h\<^sub>3, (v#vs, ls\<^sub>2, C,M,?pc\<^sub>2+1,ics)#frs,sh\<^sub>3)"
+ using Call\<^sub>1.hyps(7) M'_code M'_in_D nclinit by(cases T, auto)
+ finally show ?trans by(simp add:add.assoc)
+ qed
+ next
+ show ?Err (is "?throw \<longrightarrow> ?err")
+ proof
+ assume throw: ?throw
+ with IH_body obtain pc\<^sub>2 vs' where
+ pc\<^sub>2: "0 \<le> pc\<^sub>2 \<and> pc\<^sub>2 < size(compE\<^sub>2 body) \<and>
+ \<not> caught P pc\<^sub>2 h\<^sub>3 xa (compxE\<^sub>2 body 0 0)" and
+ 2: "P \<turnstile> ?\<sigma>\<^sub>2' -jvm\<rightarrow> handle P D M' xa h\<^sub>3 vs' ls\<^sub>3 pc\<^sub>2 No_ics ?frs\<^sub>2 sh\<^sub>3"
+ using Call\<^sub>1.prems M'_code M'_xtab cond nsub_RI_Jcc_pieces nsub
+ by (auto simp del:split_paired_Ex)
+ have "handle P D M' xa h\<^sub>3 vs' ls\<^sub>3 pc\<^sub>2 No_ics ?frs\<^sub>2 sh\<^sub>3 =
+ handle P C M xa h\<^sub>3 (rev pvs @ Addr a # vs) ls\<^sub>2 ?pc\<^sub>2 ics frs sh\<^sub>3"
+ using pc\<^sub>2 M'_in_D nclinit by(auto simp add:handle_def)
+ then show "?err" using pc\<^sub>2 jvm_trans[OF 1 2]
+ by(auto intro!:exI[where x="?pc\<^sub>2"] exI[where x="rev pvs@[Addr a]"])
+ qed
+ qed
+next
+ case (CallParamsThrow\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 w h\<^sub>1 ls\<^sub>1 sh\<^sub>1 es es' h\<^sub>2 ls\<^sub>2 sh\<^sub>2 pvs ex es'' M')
+ let ?frs\<^sub>0 = "(vs, ls\<^sub>0, C,M,pc,ics)#frs"
+ let ?\<sigma>\<^sub>0 = "(None,h\<^sub>0,(vs, ls\<^sub>0, C,M,pc,ics)#frs,sh\<^sub>0)"
+ let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e)"
+ let ?\<sigma>\<^sub>1 = "(None,h\<^sub>1,(w # vs, ls\<^sub>1, C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)"
+ let ?pc\<^sub>2 = "?pc\<^sub>1 + length(compEs\<^sub>2 es)"
+ obtain err where pcs:
+ "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v xa (e\<bullet>M'(es)) =
+ (True, ?frs\<^sub>0, (None, h\<^sub>2, (v#vs, ls\<^sub>2, C,M,?pc\<^sub>2+1,ics)#frs,sh\<^sub>2), err)"
+ using CallParamsThrow\<^sub>1.prems(1) by clarsimp
+ have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (Val w) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics w xa vs frs
+ (I - pcs (compxEs\<^sub>2 es (pc + length (compE\<^sub>2 e)) (Suc (length vs))))" by fact
+ have 1: "P \<turnstile> ?\<sigma>\<^sub>0 -jvm\<rightarrow> ?\<sigma>\<^sub>1" using Jcc_pieces_Call1[OF pcs] IH by clarsimp
+ have Isubs: "{?pc\<^sub>1..<?pc\<^sub>2} \<subseteq> I - pcs (compxE\<^sub>2 e pc (length vs))"
+ using CallParamsThrow\<^sub>1.prems by clarsimp
+ show ?case (is "?N \<and> (?eq \<longrightarrow> ?err)")
+ proof
+ show ?N by simp
+ next
+ { assume ?eq
+ moreover
+ have "PROP ?Ps es h\<^sub>1 ls\<^sub>1 sh\<^sub>1 es' h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C M ?pc\<^sub>1 ics pvs xa es'' (w#vs) frs
+ (I - pcs (compxE\<^sub>2 e pc (length vs)))" by fact
+ ultimately obtain vs' where "\<exists>pc\<^sub>2.
+ (?pc\<^sub>1 \<le> pc\<^sub>2 \<and> pc\<^sub>2 < ?pc\<^sub>1 + size(compEs\<^sub>2 es) \<and>
+ \<not> caught P pc\<^sub>2 h\<^sub>2 xa (compxEs\<^sub>2 es ?pc\<^sub>1 (size vs + 1))) \<and>
+ P \<turnstile> ?\<sigma>\<^sub>1 -jvm\<rightarrow> handle P C M xa h\<^sub>2 (vs'@w#vs) ls\<^sub>2 pc\<^sub>2 ics frs sh\<^sub>2"
+ (is "\<exists>pc\<^sub>2. ?PC pc\<^sub>2 \<and> ?Exec pc\<^sub>2")
+ using CallParamsThrow\<^sub>1 Isubs by auto
+ then obtain pc\<^sub>2 where pc\<^sub>2: "?PC pc\<^sub>2" and 2: "?Exec pc\<^sub>2" by iprover
+ then have "?err" using pc\<^sub>2 jvm_trans[OF 1 2]
+ by(auto intro!: exI[where x="pc\<^sub>2"] exI[where x="vs'@[w]"])
+ }
+ thus "?eq \<longrightarrow> ?err" by simp
+ qed
+next
+ case (CallNull\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 es pvs h\<^sub>2 ls\<^sub>2 sh\<^sub>2 M')
+ have "P\<^sub>1 \<turnstile>\<^sub>1 \<langle>es,(h\<^sub>1, ls\<^sub>1, sh\<^sub>1)\<rangle> [\<Rightarrow>] \<langle>map Val pvs,(h\<^sub>2, ls\<^sub>2, sh\<^sub>2)\<rangle>" by fact
+ hence [simp]: "length es = length pvs" by(auto dest:evals\<^sub>1_preserves_elen)
+ let ?frs\<^sub>0 = "(vs, ls\<^sub>0, C,M,pc,ics)#frs"
+ let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e)"
+ let ?pc\<^sub>2 = "?pc\<^sub>1 + length(compEs\<^sub>2 es)"
+ let ?xa = "addr_of_sys_xcpt NullPointer"
+ obtain err where pcs:
+ "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v xa (e\<bullet>M'(es)) =
+ (True, ?frs\<^sub>0, (None, h\<^sub>2, (v#vs, ls\<^sub>2, C,M,?pc\<^sub>2+1,ics)#frs,sh\<^sub>2), err)"
+ using CallNull\<^sub>1.prems(1) by clarsimp
+ have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 null h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics Null xa vs frs
+ (I - pcs (compxEs\<^sub>2 es (pc + length (compE\<^sub>2 e)) (Suc (length vs))))" by fact
+ have IH_es: "PROP ?Ps es h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (map Val pvs) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C M ?pc\<^sub>1 ics pvs xa
+ (map Val pvs) (Null#vs) frs (I - pcs(compxE\<^sub>2 e pc (size vs)))" by fact
+ have Isubs: "{pc + length (compE\<^sub>2 e)..<pc + length (compE\<^sub>2 e) + length (compEs\<^sub>2 es)}
+ \<subseteq> I - pcs (compxE\<^sub>2 e pc (length vs))" using CallNull\<^sub>1.prems by clarsimp
+ have "P \<turnstile> (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0) -jvm\<rightarrow>
+ (None,h\<^sub>1,(Null#vs,ls\<^sub>1,C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)"
+ using Jcc_pieces_Call1[OF pcs] IH by clarsimp
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> (None,h\<^sub>2,(rev pvs@Null#vs,ls\<^sub>2,C,M,?pc\<^sub>2,ics)#frs,sh\<^sub>2)"
+ using CallNull\<^sub>1 IH_es Isubs by auto
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> handle P C M ?xa h\<^sub>2 (rev pvs@Null#vs) ls\<^sub>2 ?pc\<^sub>2 ics frs sh\<^sub>2"
+ using CallNull\<^sub>1.prems
+ by(auto simp:split_beta handle_def nth_append simp del: split_paired_Ex)
+ finally show ?case by (auto intro!: exI[where x = ?pc\<^sub>2] exI[where x="rev pvs@[Null]"])
+next
+ case (CallObjThrow\<^sub>1 e h ls sh e' h' ls' sh' M' es)
+ obtain err where pcs:
+ "Jcc_pieces P\<^sub>1 E C M h vs ls pc ics frs sh I h' ls' sh' v xa (e\<bullet>M'(es)) =
+ (True, (vs, ls, C,M,pc,ics)#frs,
+ (None, h', (v#vs, ls', C,M,pc+size(compE\<^sub>2 (e\<bullet>M'(es))),ics)#frs,sh'), err)"
+ using CallObjThrow\<^sub>1.prems(1) by clarsimp
+ obtain a' where throw: "throw e' = Throw a'"
+ using eval\<^sub>1_final[OF CallObjThrow\<^sub>1.hyps(1)] by clarsimp
+ have IH: "PROP ?P e h ls sh (throw e') h' ls' sh' E C M pc ics v a' vs frs
+ (I - pcs (compxEs\<^sub>2 es (pc + length (compE\<^sub>2 e)) (Suc (length vs))))" by fact
+ show ?case using IH Jcc_pieces_Call1[OF pcs] throw CallObjThrow\<^sub>1.prems nsub_RI_Jcc_pieces
+ by auto
+next
+ case (CallNone\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 a h\<^sub>1 ls\<^sub>1 sh\<^sub>1 es pvs h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C' fs M')
+ let ?frs\<^sub>0 = "(vs, ls\<^sub>0, C,M,pc,ics)#frs"
+ let ?\<sigma>\<^sub>0 = "(None,h\<^sub>0,?frs\<^sub>0,sh\<^sub>0)"
+ let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e)"
+ let ?\<sigma>\<^sub>1 = "(None,h\<^sub>1,(Addr a#vs, ls\<^sub>1, C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)"
+ let ?pc\<^sub>2 = "?pc\<^sub>1 + length(compEs\<^sub>2 es)"
+ let ?frs\<^sub>2 = "(rev pvs @ Addr a # vs, ls\<^sub>2, C,M,?pc\<^sub>2,ics)#frs"
+ let ?\<sigma>\<^sub>2 = "(None,h\<^sub>2,?frs\<^sub>2,sh\<^sub>2)"
+ let ?xa = "addr_of_sys_xcpt NoSuchMethodError"
+ have "P\<^sub>1 \<turnstile>\<^sub>1 \<langle>es,(h\<^sub>1, ls\<^sub>1, sh\<^sub>1)\<rangle> [\<Rightarrow>] \<langle>map Val pvs,(h\<^sub>2, ls\<^sub>2, sh\<^sub>2)\<rangle>" by fact
+ hence [simp]: "length es = length pvs" by(auto dest:evals\<^sub>1_preserves_elen)
+ have aux: "(rev pvs @ Addr a # vs) ! length pvs = Addr a"
+ by (metis length_rev nth_append_length)
+ have nmeth: "\<not>(\<exists>b Ts T body D. P \<turnstile> C' sees M', b : Ts\<rightarrow>T = body in D)"
+ using sees_method_compPD CallNone\<^sub>1.hyps(6) by fastforce
+ obtain err where pcs:
+ "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v xa (e\<bullet>M'(es)) =
+ (True, ?frs\<^sub>0, (None, h\<^sub>2, (v#vs, ls\<^sub>2, C,M,?pc\<^sub>2+1,ics)#frs,sh\<^sub>2), err)"
+ using CallNone\<^sub>1.prems(1) by clarsimp
+ have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (addr a) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Addr a) xa vs frs
+ (I - pcs (compxEs\<^sub>2 es (pc + length (compE\<^sub>2 e)) (Suc (length vs))))" by fact
+ have IH_es: "PROP ?Ps es h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (map Val pvs) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C M ?pc\<^sub>1 ics pvs xa
+ (map Val pvs) (Addr a#vs) frs (I - pcs(compxE\<^sub>2 e pc (size vs)))" by fact
+ have "P \<turnstile> ?\<sigma>\<^sub>0 -jvm\<rightarrow> ?\<sigma>\<^sub>1" using Jcc_pieces_Call1[OF pcs] IH by clarsimp
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> ?\<sigma>\<^sub>2" using IH_es CallNone\<^sub>1.prems by fastforce
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> handle P C M ?xa h\<^sub>2 (rev pvs@Addr a#vs) ls\<^sub>2 ?pc\<^sub>2 ics frs sh\<^sub>2"
+ using CallNone\<^sub>1.hyps(5) CallNone\<^sub>1.prems aux nmeth
+ by(cases "method P C' M'", cases "find_handler P ?xa h\<^sub>2 frs sh\<^sub>2", auto simp: handle_def)
+ finally show ?case using pcs by (auto intro!: exI[where x = ?pc\<^sub>2] exI[where x="rev pvs@[Addr a]"])
+next
+ case (CallStatic\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 a h\<^sub>1 ls\<^sub>1 sh\<^sub>1 es pvs h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C' fs M' Ts T body D)
+ let ?frs\<^sub>0 = "(vs, ls\<^sub>0, C,M,pc,ics)#frs"
+ let ?\<sigma>\<^sub>0 = "(None,h\<^sub>0,?frs\<^sub>0,sh\<^sub>0)"
+ let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e)"
+ let ?\<sigma>\<^sub>1 = "(None,h\<^sub>1,(Addr a#vs, ls\<^sub>1, C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)"
+ let ?pc\<^sub>2 = "?pc\<^sub>1 + length(compEs\<^sub>2 es)"
+ let ?frs\<^sub>2 = "(rev pvs @ Addr a # vs, ls\<^sub>2, C,M,?pc\<^sub>2,ics)#frs"
+ let ?\<sigma>\<^sub>2 = "(None,h\<^sub>2,?frs\<^sub>2,sh\<^sub>2)"
+ let ?xa = "addr_of_sys_xcpt IncompatibleClassChangeError"
+ have "P\<^sub>1 \<turnstile>\<^sub>1 \<langle>es,(h\<^sub>1, ls\<^sub>1, sh\<^sub>1)\<rangle> [\<Rightarrow>] \<langle>map Val pvs,(h\<^sub>2, ls\<^sub>2, sh\<^sub>2)\<rangle>" by fact
+ hence [simp]: "length es = length pvs" by(auto dest:evals\<^sub>1_preserves_elen)
+ have aux: "(rev pvs @ Addr a # vs) ! length pvs = Addr a"
+ by (metis length_rev nth_append_length)
+ obtain body' where method: "P \<turnstile> C' sees M', Static : Ts\<rightarrow>T = body' in D"
+ by (metis CallStatic\<^sub>1.hyps(6) P_def compP\<^sub>2_def sees_method_compP)
+ obtain err where pcs:
+ "Jcc_pieces P\<^sub>1 E C M h\<^sub>0 vs ls\<^sub>0 pc ics frs sh\<^sub>0 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v xa (e\<bullet>M'(es)) =
+ (True, ?frs\<^sub>0, (None, h\<^sub>2, (v#vs, ls\<^sub>2, C,M,?pc\<^sub>2+1,ics)#frs,sh\<^sub>2), err)"
+ using CallStatic\<^sub>1.prems(1) by clarsimp
+ have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (addr a) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Addr a) xa vs frs
+ (I - pcs (compxEs\<^sub>2 es (pc + length (compE\<^sub>2 e)) (Suc (length vs))))" by fact
+ have IH_es: "PROP ?Ps es h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (map Val pvs) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C M ?pc\<^sub>1 ics pvs xa
+ (map Val pvs) (Addr a#vs) frs (I - pcs(compxE\<^sub>2 e pc (size vs)))" by fact
+ have "P \<turnstile> ?\<sigma>\<^sub>0 -jvm\<rightarrow> ?\<sigma>\<^sub>1" using Jcc_pieces_Call1[OF pcs] IH by clarsimp
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> ?\<sigma>\<^sub>2" using IH_es CallStatic\<^sub>1.prems by fastforce
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> handle P C M ?xa h\<^sub>2 (rev pvs@Addr a#vs) ls\<^sub>2 ?pc\<^sub>2 ics frs sh\<^sub>2"
+ using CallStatic\<^sub>1.hyps(5) CallStatic\<^sub>1.prems aux method
+ by(cases "method P C' M'", cases "find_handler P ?xa h\<^sub>2 frs sh\<^sub>2")
+ (auto simp: handle_def; meson frames_of.cases)
+ finally show ?case using pcs by (auto intro!: exI[where x = ?pc\<^sub>2] exI[where x="rev pvs@[Addr a]"])
+next
+ case (SCallParamsThrow\<^sub>1 es h\<^sub>1 ls\<^sub>1 sh\<^sub>1 es' h\<^sub>2 ls\<^sub>2 sh\<^sub>2 pvs ex es'' C' M')
+ show ?case
+ proof(cases "M' = clinit \<and> es = []")
+ case clinit: True then show ?thesis
+ using SCallParamsThrow\<^sub>1.hyps(1,3) evals\<^sub>1_cases(1) by fastforce
+ next
+ case nclinit: False
+ let ?\<sigma>\<^sub>1 = "(None,h\<^sub>1,(vs, ls\<^sub>1, C,M,pc,ics)#frs,sh\<^sub>1)"
+ let ?pc\<^sub>2 = "pc + length(compEs\<^sub>2 es)"
+ have Isubs: "{pc..<pc + length (compEs\<^sub>2 es)} \<subseteq> I" using SCallParamsThrow\<^sub>1.prems nclinit by clarsimp
+ show ?thesis (is "?N \<and> (?eq \<longrightarrow> ?err)")
+ proof
+ show ?N by simp
+ next
+ { assume ?eq
+ moreover
+ have "PROP ?Ps es h\<^sub>1 ls\<^sub>1 sh\<^sub>1 es' h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C M pc ics pvs xa es'' vs frs I" by fact
+ ultimately have "\<exists>pc\<^sub>2.
+ (pc \<le> pc\<^sub>2 \<and> pc\<^sub>2 < pc + size(compEs\<^sub>2 es) \<and>
+ \<not> caught P pc\<^sub>2 h\<^sub>2 xa (compxEs\<^sub>2 es pc (size vs))) \<and>
+ (\<exists>vs'. P \<turnstile> ?\<sigma>\<^sub>1 -jvm\<rightarrow> handle P C M xa h\<^sub>2 (vs'@vs) ls\<^sub>2 pc\<^sub>2 ics frs sh\<^sub>2)"
+ (is "\<exists>pc\<^sub>2. ?PC pc\<^sub>2 \<and> ?Exec pc\<^sub>2")
+ using SCallParamsThrow\<^sub>1 Isubs nclinit by auto
+ then obtain pc\<^sub>2 where pc\<^sub>2: "?PC pc\<^sub>2" and 2: "?Exec pc\<^sub>2" by iprover
+ then have "?err" using pc\<^sub>2 2 by(auto intro: exI[where x="pc\<^sub>2"])
+ }
+ thus "?eq \<longrightarrow> ?err" by iprover
+ qed
+ qed
+next
+ case (SCallNone\<^sub>1 es h\<^sub>1 ls\<^sub>1 sh\<^sub>1 pvs h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C' M')
+ show ?case
+ proof(cases "M' = clinit \<and> es = []")
+ case clinit: True then show ?thesis using SCallNone\<^sub>1.hyps(3) SCallNone\<^sub>1.prems by auto
+ next
+ case nclinit: False
+ let ?\<sigma>\<^sub>1 = "(None,h\<^sub>1,(vs, ls\<^sub>1, C,M,pc,ics)#frs,sh\<^sub>1)"
+ let ?pc\<^sub>2 = "pc + length(compEs\<^sub>2 es)"
+ let ?frs\<^sub>2 = "(rev pvs @ vs, ls\<^sub>2, C,M,?pc\<^sub>2,ics)#frs"
+ let ?\<sigma>\<^sub>2 = "(None,h\<^sub>2,?frs\<^sub>2,sh\<^sub>2)"
+ let ?xa = "addr_of_sys_xcpt NoSuchMethodError"
+ have "P\<^sub>1 \<turnstile>\<^sub>1 \<langle>es,(h\<^sub>1, ls\<^sub>1, sh\<^sub>1)\<rangle> [\<Rightarrow>] \<langle>map Val pvs,(h\<^sub>2, ls\<^sub>2, sh\<^sub>2)\<rangle>" by fact
+ hence [simp]: "length es = length pvs" by(auto dest:evals\<^sub>1_preserves_elen)
+ have nmeth: "\<not>(\<exists>b Ts T body D. P \<turnstile> C' sees M', b : Ts\<rightarrow>T = body in D)"
+ using sees_method_compPD SCallNone\<^sub>1.hyps(3) by fastforce
+ have IH_es: "PROP ?Ps es h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (map Val pvs) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C M pc ics pvs xa
+ (map Val pvs) vs frs I" by fact
+ have "P \<turnstile> ?\<sigma>\<^sub>1 -jvm\<rightarrow> ?\<sigma>\<^sub>2" using IH_es SCallNone\<^sub>1.prems nclinit by auto fastforce+
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> handle P C M ?xa h\<^sub>2 (rev pvs@vs) ls\<^sub>2 ?pc\<^sub>2 ics frs sh\<^sub>2"
+ using SCallNone\<^sub>1.prems nmeth nclinit
+ by(cases "method P C' M'", cases "find_handler P ?xa h\<^sub>2 frs sh\<^sub>2", auto simp: handle_def)
+ finally show ?thesis using nclinit by (auto intro: exI[where x = ?pc\<^sub>2])
+ qed
+next
+ case (SCallNonStatic\<^sub>1 es h\<^sub>1 ls\<^sub>1 sh\<^sub>1 pvs h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C' M' Ts T body D)
+ show ?case
+ proof(cases "M' = clinit \<and> es = []")
+ case clinit: True then show ?thesis
+ using SCallNonStatic\<^sub>1.hyps(3) SCallNonStatic\<^sub>1.prems sees_method_fun by fastforce
+ next
+ case nclinit: False
+ let ?\<sigma>\<^sub>1 = "(None,h\<^sub>1,(vs, ls\<^sub>1, C,M,pc,ics)#frs,sh\<^sub>1)"
+ let ?pc\<^sub>2 = "pc + length(compEs\<^sub>2 es)"
+ let ?frs\<^sub>2 = "(rev pvs @ vs, ls\<^sub>2, C,M,?pc\<^sub>2,ics)#frs"
+ let ?\<sigma>\<^sub>2 = "(None,h\<^sub>2,?frs\<^sub>2,sh\<^sub>2)"
+ let ?xa = "addr_of_sys_xcpt IncompatibleClassChangeError"
+ have "P\<^sub>1 \<turnstile>\<^sub>1 \<langle>es,(h\<^sub>1, ls\<^sub>1, sh\<^sub>1)\<rangle> [\<Rightarrow>] \<langle>map Val pvs,(h\<^sub>2, ls\<^sub>2, sh\<^sub>2)\<rangle>" by fact
+ hence [simp]: "length es = length pvs" by(auto dest:evals\<^sub>1_preserves_elen)
+ obtain body' where method: "P \<turnstile> C' sees M', NonStatic : Ts\<rightarrow>T = body' in D"
+ by (metis SCallNonStatic\<^sub>1.hyps(3) P_def compP\<^sub>2_def sees_method_compP)
+ have IH_es: "PROP ?Ps es h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (map Val pvs) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C M pc ics pvs xa
+ (map Val pvs) vs frs I" by fact
+ have "P \<turnstile> ?\<sigma>\<^sub>1 -jvm\<rightarrow> ?\<sigma>\<^sub>2" using IH_es SCallNonStatic\<^sub>1.prems nclinit by auto fastforce+
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> handle P C M ?xa h\<^sub>2 (rev pvs@vs) ls\<^sub>2 ?pc\<^sub>2 ics frs sh\<^sub>2"
+ using SCallNonStatic\<^sub>1.prems method nclinit
+ by(cases "method P C' M'", cases "find_handler P ?xa h\<^sub>2 frs sh\<^sub>2")
+ (auto simp: handle_def; meson frames_of.cases)
+ finally show ?thesis using nclinit by (auto intro: exI[where x = ?pc\<^sub>2])
+ qed
+next
+ case (SCallInitThrow\<^sub>1 es h\<^sub>0 ls\<^sub>0 sh\<^sub>0 pvs h\<^sub>1 ls\<^sub>1 sh\<^sub>1 C' M' Ts T body D a h\<^sub>2 ls\<^sub>2 sh\<^sub>2)
+ show ?case
+ proof(cases "M' = clinit \<and> es = []")
+ case clinit: True then show ?thesis using SCallInitThrow\<^sub>1 by simp
+ next
+ case nclinit: False
+ let ?\<sigma>\<^sub>0 = "(None,h\<^sub>0,(vs, ls\<^sub>0, C,M,pc,ics)#frs,sh\<^sub>0)"
+ let ?pc\<^sub>1 = "pc + length(compEs\<^sub>2 es)"
+ let ?frs\<^sub>1 = "(rev pvs @ vs, ls\<^sub>1, C,M,?pc\<^sub>1,ics)#frs"
+ let ?\<sigma>\<^sub>1 = "(None,h\<^sub>1,?frs\<^sub>1,sh\<^sub>1)"
+ let ?frs\<^sub>1' = "(rev pvs@vs,ls\<^sub>1,C,M,?pc\<^sub>1,Calling D [])#frs"
+ let ?\<sigma>\<^sub>1' = "(None,h\<^sub>1,?frs\<^sub>1',sh\<^sub>1)"
+ let ?frs\<^sub>2 = "(rev pvs@vs,ls\<^sub>1,C,M,?pc\<^sub>1,Called [])#frs"
+ let ?\<sigma>\<^sub>2 = "(None,h\<^sub>2,?frs\<^sub>2,sh\<^sub>2)"
+ have ls: "ls\<^sub>1 = ls\<^sub>2" by(rule init\<^sub>1_same_loc[OF SCallInitThrow\<^sub>1.hyps(6)])
+ have method: "\<exists>m'. P \<turnstile> C' sees M',Static:Ts\<rightarrow>T = m' in D" using SCallInitThrow\<^sub>1.hyps(3)
+ by (metis P_def compP\<^sub>2_def sees_method_compP)
+ obtain a' where throw: "throw a = Throw a'" using eval\<^sub>1_final[OF SCallInitThrow\<^sub>1.hyps(6)] by clarsimp
+ have "Ex (WTrt2\<^sub>1 P\<^sub>1 E h\<^sub>1 sh\<^sub>1 (INIT D ([D],False) \<leftarrow> unit))"
+ using sees_method_is_class'[OF SCallInitThrow\<^sub>1.hyps(3)] by auto
+ then obtain err' where pcs':
+ "Jcc_pieces P\<^sub>1 E C M h\<^sub>1 (rev pvs@vs) ls\<^sub>1 ?pc\<^sub>1 ics frs sh\<^sub>1 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v xa (INIT D ([D],False) \<leftarrow> unit)
+ = (True, ?frs\<^sub>1', (None,h\<^sub>2,?frs\<^sub>2,sh\<^sub>2), err')"
+ using SCallInitThrow\<^sub>1.prems(1) nclinit by auto
+ have IHI: "PROP ?P (INIT D ([D],False) \<leftarrow> unit) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (throw a)
+ h\<^sub>2 ls\<^sub>2 sh\<^sub>2 E C M ?pc\<^sub>1 ics v a' (rev pvs@vs) frs I" by fact
+ have IH_es: "PROP ?Ps es h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (map Val pvs) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 C M pc ics pvs xa
+ (map Val pvs) vs frs I" by fact
+ have "P \<turnstile> ?\<sigma>\<^sub>0 -jvm\<rightarrow> ?\<sigma>\<^sub>1" using IH_es SCallInitThrow\<^sub>1.prems nclinit by auto fastforce+
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> ?\<sigma>\<^sub>1'"
+ proof(cases "sh\<^sub>1 D")
+ case None then show ?thesis using SCallInitThrow\<^sub>1.hyps(1,3-6) SCallInitThrow\<^sub>1.prems method
+ by(cases ics) auto
+ next
+ case (Some a)
+ then obtain sfs i where "a = (sfs,i)" by(cases a)
+ then show ?thesis using SCallInitThrow\<^sub>1.hyps(1,3-6) SCallInitThrow\<^sub>1.prems method Some
+ by(cases ics; case_tac i, auto)
+ qed
+ also obtain vs' where "P \<turnstile> \<dots> -jvm\<rightarrow> handle P C M a' h\<^sub>2 (vs'@rev pvs@vs) ls\<^sub>1 ?pc\<^sub>1 ics frs sh\<^sub>2"
+ using IHI pcs' throw by auto
+ finally show ?thesis using nclinit throw ls
+ by(auto intro!: exI[where x="?pc\<^sub>1"] exI[where x="vs'@rev pvs"])
+ qed
+next
+ case (SCallInit\<^sub>1 es h\<^sub>0 ls\<^sub>0 sh\<^sub>0 pvs h\<^sub>1 ls\<^sub>1 sh\<^sub>1 C' M' Ts T body D v' h\<^sub>2 ls\<^sub>2 sh\<^sub>2 ls\<^sub>2' e' h\<^sub>3 ls\<^sub>3 sh\<^sub>3)
+ show ?case
+ proof(cases "M' = clinit \<and> es = []")
+ case clinit: True then show ?thesis using SCallInit\<^sub>1 by simp
+ next
+ case nclinit: False
+ let ?\<sigma>\<^sub>0 = "(None,h\<^sub>0,(vs, ls\<^sub>0, C,M,pc,ics)#frs,sh\<^sub>0)"
+ let ?pc\<^sub>1 = "pc + length(compEs\<^sub>2 es)"
+ let ?frs\<^sub>1 = "(rev pvs @ vs, ls\<^sub>1, C,M,?pc\<^sub>1,ics)#frs"
+ let ?\<sigma>\<^sub>1 = "(None,h\<^sub>1,?frs\<^sub>1,sh\<^sub>1)"
+ let ?frs\<^sub>1' = "(rev pvs@vs,ls\<^sub>1,C,M,?pc\<^sub>1,Calling D [])#frs"
+ let ?\<sigma>\<^sub>1' = "(None,h\<^sub>1,?frs\<^sub>1',sh\<^sub>1)"
+ let ?frs\<^sub>2 = "(rev pvs@vs,ls\<^sub>1,C,M,?pc\<^sub>1,Called [])#frs"
+ let ?\<sigma>\<^sub>2 = "(None,h\<^sub>2,?frs\<^sub>2,sh\<^sub>2)"
+ let ?frs\<^sub>2' = "([], ls\<^sub>2', D,M',0,No_ics) # ?frs\<^sub>1"
+ let ?\<sigma>\<^sub>2' = "(None, h\<^sub>2, ?frs\<^sub>2', sh\<^sub>2)"
+ have nclinit': "M' \<noteq> clinit" by fact
+ have ics: "ics = No_ics" using SCallInit\<^sub>1.hyps(5) SCallInit\<^sub>1.prems by simp
+ have "P\<^sub>1 \<turnstile>\<^sub>1 \<langle>es,(h\<^sub>0, ls\<^sub>0, sh\<^sub>0)\<rangle> [\<Rightarrow>] \<langle>map Val pvs,(h\<^sub>1, ls\<^sub>1, sh\<^sub>1)\<rangle>" by fact
+ hence [simp]: "length es = length pvs" by(auto dest:evals\<^sub>1_preserves_elen)
+ have invoke: "P,C,M,?pc\<^sub>1 \<triangleright> Invokestatic C' M' (length Ts)"
+ using SCallInit\<^sub>1.hyps(8) SCallInit\<^sub>1.prems nclinit by(auto simp: add.assoc)
+ have nsub: "\<not> sub_RI body" by(rule sees_wf\<^sub>1_nsub_RI[OF wf SCallInit\<^sub>1.hyps(3)])
+ have ls: "ls\<^sub>1 = ls\<^sub>2" by(rule init\<^sub>1_same_loc[OF SCallInit\<^sub>1.hyps(6)])
+ obtain sfs i where sh\<^sub>2: "sh\<^sub>2 D = Some(sfs,i)"
+ using init\<^sub>1_Val_PD[OF SCallInit\<^sub>1.hyps(6)] by clarsimp
+ have method: "\<exists>m'. P \<turnstile> C' sees M',Static:Ts\<rightarrow>T = m' in D" using SCallInit\<^sub>1.hyps(3)
+ by (metis P_def compP\<^sub>2_def sees_method_compP)
+ have "Ex (WTrt2\<^sub>1 P\<^sub>1 E h\<^sub>1 sh\<^sub>1 (INIT D ([D],False) \<leftarrow> unit))"
+ using sees_method_is_class'[OF SCallInit\<^sub>1.hyps(3)] by auto
+ then obtain err' where pcs':
+ "Jcc_pieces P\<^sub>1 E C M h\<^sub>1 (rev pvs@vs) ls\<^sub>1 ?pc\<^sub>1 ics frs sh\<^sub>1 I h\<^sub>2 ls\<^sub>2 sh\<^sub>2 v' xa (INIT D ([D],False) \<leftarrow> unit)
+ = (True, ?frs\<^sub>1', (None,h\<^sub>2,?frs\<^sub>2,sh\<^sub>2), err')"
+ using SCallInit\<^sub>1.prems(1) nclinit by auto
+ have IHI: "PROP ?P (INIT D ([D],False) \<leftarrow> unit) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (Val v')
+ h\<^sub>2 ls\<^sub>2 sh\<^sub>2 E C M ?pc\<^sub>1 ics v' xa (rev pvs@vs) frs I" by fact
+ have IH_es: "PROP ?Ps es h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (map Val pvs) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 C M pc ics pvs xa
+ (map Val pvs) vs frs I" by fact
+ have "P \<turnstile> ?\<sigma>\<^sub>0 -jvm\<rightarrow> ?\<sigma>\<^sub>1" using IH_es SCallInit\<^sub>1.prems nclinit by auto fastforce+
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> ?\<sigma>\<^sub>1'"
+ proof(cases "sh\<^sub>1 D")
+ case None then show ?thesis using SCallInit\<^sub>1.hyps(1,3-6,8-10) SCallInit\<^sub>1.prems method
+ by(cases ics) auto
+ next
+ case (Some a)
+ then obtain sfs i where "a = (sfs,i)" by(cases a)
+ then show ?thesis using SCallInit\<^sub>1.hyps(1,3-6,8-10) SCallInit\<^sub>1.prems method Some
+ by(cases ics; case_tac i, auto)
+ qed
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> ?\<sigma>\<^sub>2" using IHI pcs' by auto
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> ?\<sigma>\<^sub>2'"
+ using jvm_Invokestatic_Called[OF assms(1) invoke _ SCallInit\<^sub>1.hyps(3,8,9)] sh\<^sub>2 ics by auto
+ finally have 1: "P \<turnstile> ?\<sigma>\<^sub>0 -jvm\<rightarrow> ?\<sigma>\<^sub>2'".
+ have "P\<^sub>1 \<turnstile> C' sees M',Static: Ts\<rightarrow>T = body in D" by fact
+ then have M'_in_D: "P\<^sub>1 \<turnstile> D sees M',Static: Ts\<rightarrow>T = body in D"
+ by(rule sees_method_idemp)
+ have M'_code: "compP\<^sub>2 P\<^sub>1,D,M',0 \<rhd> compE\<^sub>2 body @ [Return]" using beforeM M'_in_D by simp
+ have M'_xtab: "compP\<^sub>2 P\<^sub>1,D,M' \<rhd> compxE\<^sub>2 body 0 0/{..<size(compE\<^sub>2 body)},0"
+ using M'_in_D by(rule beforexM)
+ have IH_body: "PROP ?P body h\<^sub>2 ls\<^sub>2' sh\<^sub>2 e' h\<^sub>3 ls\<^sub>3 sh\<^sub>3 (Class D # Ts) D M' 0 No_ics v xa [] ?frs\<^sub>1
+ ({..<size(compE\<^sub>2 body)})" by fact
+ have cond: "Jcc_cond P\<^sub>1 (Class D # Ts) D M' [] 0 No_ics {..<length (compE\<^sub>2 body)} h\<^sub>2 sh\<^sub>2 body"
+ using nsub_RI_Jcc_pieces[OF assms(1) nsub] M'_code M'_xtab by clarsimp
+ show ?thesis (is "?Norm \<and> ?Err")
+ proof
+ show ?Norm (is "?val \<longrightarrow> ?trans")
+ proof
+ assume val: ?val
+ note 1
+ also have "P \<turnstile> ?\<sigma>\<^sub>2' -jvm\<rightarrow> (None,h\<^sub>3,([v],ls\<^sub>3,D,M',size(compE\<^sub>2 body),No_ics)#?frs\<^sub>1,sh\<^sub>3)"
+ using val IH_body SCallInit\<^sub>1.prems M'_code cond nsub_RI_Jcc_pieces nsub by auto
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> (None, h\<^sub>3, (v#vs, ls\<^sub>2, C,M,?pc\<^sub>1+1,ics)#frs,sh\<^sub>3)"
+ using SCallInit\<^sub>1.hyps(8) M'_code M'_in_D ls nclinit' by(cases T, auto)
+ finally show ?trans using nclinit by(auto simp:add.assoc)
+ qed
+ next
+ show ?Err (is "?throw \<longrightarrow> ?err")
+ proof
+ assume throw: ?throw
+ with IH_body obtain pc\<^sub>2 vs' where
+ pc\<^sub>2: "0 \<le> pc\<^sub>2 \<and> pc\<^sub>2 < size(compE\<^sub>2 body) \<and>
+ \<not> caught P pc\<^sub>2 h\<^sub>3 xa (compxE\<^sub>2 body 0 0)" and
+ 2: "P \<turnstile> ?\<sigma>\<^sub>2' -jvm\<rightarrow> handle P D M' xa h\<^sub>3 vs' ls\<^sub>3 pc\<^sub>2 No_ics ?frs\<^sub>1 sh\<^sub>3"
+ using SCallInit\<^sub>1.prems M'_code M'_xtab cond nsub_RI_Jcc_pieces nsub
+ by (auto simp del:split_paired_Ex)
+ have "handle P D M' xa h\<^sub>3 vs' ls\<^sub>3 pc\<^sub>2 No_ics ?frs\<^sub>1 sh\<^sub>3 =
+ handle P C M xa h\<^sub>3 (rev pvs @ vs) ls\<^sub>2 ?pc\<^sub>1 ics frs sh\<^sub>3"
+ using pc\<^sub>2 M'_in_D ls nclinit' by(auto simp add:handle_def)
+ then show "?err" using pc\<^sub>2 jvm_trans[OF 1 2] nclinit
+ by(auto intro!:exI[where x="?pc\<^sub>1"] exI[where x="rev pvs"])
+ qed
+ qed
+ qed
+next
+ case (SCall\<^sub>1 es h\<^sub>1 ls\<^sub>1 sh\<^sub>1 pvs h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C' M' Ts T body D sfs ls\<^sub>2' e' h\<^sub>3 ls\<^sub>3 sh\<^sub>3)
+ show ?case
+ proof(cases "M' = clinit \<and> es = []")
+ case clinit: True
+ then have s1: "pvs = []" "h\<^sub>1 = h\<^sub>2" "ls\<^sub>1 = ls\<^sub>2" "sh\<^sub>1 = sh\<^sub>2"
+ using SCall\<^sub>1.hyps(1) evals\<^sub>1_cases(1) by blast+
+ then have ls\<^sub>2': "ls\<^sub>2' = replicate (max_vars body) undefined" using SCall\<^sub>1.hyps(6) clinit by simp
+ let ?frs = "create_init_frame P C' # (vs, ls\<^sub>1, C,M,pc,ics)#frs"
+ let ?\<sigma>\<^sub>1 = "(None,h\<^sub>1,?frs,sh\<^sub>1)"
+ have method: "P\<^sub>1 \<turnstile> C' sees clinit,Static: []\<rightarrow>Void = body in C'"
+ using SCall\<^sub>1.hyps(3) clinit s1(1) wf_sees_clinit[OF wf]
+ by (metis is_class_def option.collapse sees_method_fun sees_method_is_class)
+ then have M_code: "compP\<^sub>2 P\<^sub>1,C',clinit,0 \<rhd> compE\<^sub>2 body @ [Return]" by(rule beforeM)
+ have pcs: "Jcc_pieces P\<^sub>1 E C M h\<^sub>1 vs ls\<^sub>1 pc ics frs sh\<^sub>1 I h\<^sub>3 ls\<^sub>2 sh\<^sub>3 v xa (C'\<bullet>\<^sub>sclinit([]))
+ = (True, ?frs, (None, h\<^sub>3, tl ?frs, sh\<^sub>3(C'\<mapsto>(fst(the(sh\<^sub>3 C')),Done))),
+ P \<turnstile> (None, h\<^sub>1, ?frs, sh\<^sub>1) -jvm\<rightarrow>
+ (case ics of
+ Called Cs \<Rightarrow> (None, h\<^sub>3, (vs, ls\<^sub>1, C, M, pc, Throwing Cs xa) # frs, sh\<^sub>3(C' \<mapsto> (fst (the (sh\<^sub>3 C')), Error)))))"
+ using Jcc_pieces_clinit[OF assms(1),of E C M vs pc ics I h\<^sub>1 sh\<^sub>1 C' ls\<^sub>1 frs h\<^sub>3 ls\<^sub>2 sh\<^sub>3 v xa]
+ SCall\<^sub>1.prems(1) clinit s1(1) by clarsimp
+ have IH_body: "PROP ?P body h\<^sub>2 ls\<^sub>2' sh\<^sub>2 e' h\<^sub>3 ls\<^sub>3 sh\<^sub>3 [] C' clinit 0 No_ics v xa [] (tl ?frs)
+ ({..<size(compE\<^sub>2 body)})" by fact
+ show ?thesis (is "?Norm \<and> ?Err")
+ proof
+ show ?Norm (is "?val \<longrightarrow> ?trans")
+ proof
+ assume val: ?val
+ then have "P \<turnstile> ?\<sigma>\<^sub>1
+ -jvm\<rightarrow> (None, h\<^sub>3, ([v], ls\<^sub>3, C', clinit, size(compE\<^sub>2 body), No_ics) # tl ?frs,sh\<^sub>3)"
+ using IH_body Jcc_pieces_SCall_clinit_body[OF assms(1) wf pcs method] s1 ls\<^sub>2' by clarsimp
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> (None, h\<^sub>3, tl ?frs, sh\<^sub>3(C'\<mapsto>(fst(the(sh\<^sub>3 C')),Done)))"
+ using jvm_Return_Init[OF M_code] by simp
+ finally show ?trans using pcs s1 clinit by simp
+ qed
+ next
+ show ?Err (is "?throw \<longrightarrow> ?err")
+ proof
+ assume throw: ?throw
+ with IH_body obtain pc\<^sub>2 vs2 where
+ pc\<^sub>2: "0 \<le> pc\<^sub>2 \<and> pc\<^sub>2 < size(compE\<^sub>2 body) \<and>
+ \<not> caught P pc\<^sub>2 h\<^sub>3 xa (compxE\<^sub>2 body 0 0)" and
+ 2: "P \<turnstile> ?\<sigma>\<^sub>1 -jvm\<rightarrow> handle P C' clinit xa h\<^sub>3 vs2 ls\<^sub>3 pc\<^sub>2 No_ics (tl ?frs) sh\<^sub>3"
+ using SCall\<^sub>1.prems Jcc_pieces_SCall_clinit_body[OF assms(1) wf pcs method] s1 ls\<^sub>2' by clarsimp
+ show ?err using SCall\<^sub>1.prems(1) clinit
+ proof(cases ics)
+ case (Called Cs)
+ note 2
+ also have "handle P C' clinit xa h\<^sub>3 vs2 ls\<^sub>3 pc\<^sub>2 No_ics (tl ?frs) sh\<^sub>3
+ = (None, h\<^sub>3, (vs, ls\<^sub>1, C, M, pc, Throwing (C'#Cs) xa) # frs, sh\<^sub>3)"
+ using Called pc\<^sub>2 method by(simp add: handle_def)
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> (None, h\<^sub>3, (vs, ls\<^sub>1, C, M, pc, Throwing Cs xa) # frs,
+ sh\<^sub>3(C' \<mapsto> (fst (the (sh\<^sub>3 C')), Error)))" using Called jvm_Throwing by simp
+ finally show ?thesis using pcs clinit Called by(clarsimp intro!: exI[where x="[]"])
+ qed(auto)
+ qed
+ qed
+ next
+ case nclinit: False
+ let ?\<sigma>\<^sub>1 = "(None,h\<^sub>1,(vs, ls\<^sub>1, C,M,pc,ics)#frs,sh\<^sub>1)"
+ let ?pc\<^sub>2 = "pc + length(compEs\<^sub>2 es)"
+ let ?frs\<^sub>2 = "(rev pvs @ vs, ls\<^sub>2, C,M,?pc\<^sub>2,ics)#frs"
+ let ?\<sigma>\<^sub>2 = "(None,h\<^sub>2,?frs\<^sub>2,sh\<^sub>2)"
+ let ?frs\<^sub>2' = "([], ls\<^sub>2', D,M',0,No_ics) # ?frs\<^sub>2"
+ let ?\<sigma>\<^sub>2' = "(None, h\<^sub>2, ?frs\<^sub>2', sh\<^sub>2)"
+ have nclinit': "M' \<noteq> clinit"
+ using wf_sees_clinit1[OF wf] visible_method_exists[OF SCall\<^sub>1.hyps(3)]
+ sees_method_idemp[OF SCall\<^sub>1.hyps(3)] nclinit SCall\<^sub>1.hyps(5)
+ evals\<^sub>1_preserves_elen[OF SCall\<^sub>1.hyps(1)] by fastforce
+ have "P\<^sub>1 \<turnstile>\<^sub>1 \<langle>es,(h\<^sub>1, ls\<^sub>1, sh\<^sub>1)\<rangle> [\<Rightarrow>] \<langle>map Val pvs,(h\<^sub>2, ls\<^sub>2, sh\<^sub>2)\<rangle>" by fact
+ hence [simp]: "length es = length pvs" by(auto dest:evals\<^sub>1_preserves_elen)
+ have invoke: "P,C,M,?pc\<^sub>2 \<triangleright> Invokestatic C' M' (length Ts)"
+ using SCall\<^sub>1.hyps(5) SCall\<^sub>1.prems nclinit by(auto simp: add.assoc)
+ have nsub: "\<not> sub_RI body" by(rule sees_wf\<^sub>1_nsub_RI[OF wf SCall\<^sub>1.hyps(3)])
+ have IH_es: "PROP ?Ps es h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (map Val pvs) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C M pc ics pvs xa
+ (map Val pvs) vs frs I" by fact
+ have "P \<turnstile> ?\<sigma>\<^sub>1 -jvm\<rightarrow> ?\<sigma>\<^sub>2" using IH_es SCall\<^sub>1.prems nclinit by auto fastforce+
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> ?\<sigma>\<^sub>2'" using jvm_Invokestatic[OF assms(1) invoke _ SCall\<^sub>1.hyps(3,5,6)]
+ SCall\<^sub>1.hyps(4) SCall\<^sub>1.prems nclinit by auto
+ finally have 1: "P \<turnstile> ?\<sigma>\<^sub>1 -jvm\<rightarrow> ?\<sigma>\<^sub>2'".
+ have "P\<^sub>1 \<turnstile> C' sees M',Static: Ts\<rightarrow>T = body in D" by fact
+ then have M'_in_D: "P\<^sub>1 \<turnstile> D sees M',Static: Ts\<rightarrow>T = body in D"
+ by(rule sees_method_idemp)
+ have M'_code: "compP\<^sub>2 P\<^sub>1,D,M',0 \<rhd> compE\<^sub>2 body @ [Return]" using beforeM M'_in_D by simp
+ have M'_xtab: "compP\<^sub>2 P\<^sub>1,D,M' \<rhd> compxE\<^sub>2 body 0 0/{..<size(compE\<^sub>2 body)},0"
+ using M'_in_D by(rule beforexM)
+ have IH_body: "PROP ?P body h\<^sub>2 ls\<^sub>2' sh\<^sub>2 e' h\<^sub>3 ls\<^sub>3 sh\<^sub>3 (Class D # Ts) D M' 0 No_ics v xa [] ?frs\<^sub>2
+ ({..<size(compE\<^sub>2 body)})" by fact
+ have cond: "Jcc_cond P\<^sub>1 (Class D # Ts) D M' [] 0 No_ics {..<length (compE\<^sub>2 body)} h\<^sub>2 sh\<^sub>2 body"
+ using nsub_RI_Jcc_pieces[OF assms(1) nsub] M'_code M'_xtab by clarsimp
+ show ?thesis (is "?Norm \<and> ?Err")
+ proof
+ show ?Norm (is "?val \<longrightarrow> ?trans")
+ proof
+ assume val: ?val
+ note 1
+ also have "P \<turnstile> ?\<sigma>\<^sub>2' -jvm\<rightarrow> (None,h\<^sub>3,([v],ls\<^sub>3,D,M',size(compE\<^sub>2 body),No_ics)#?frs\<^sub>2,sh\<^sub>3)"
+ using val IH_body SCall\<^sub>1.prems M'_code cond nsub_RI_Jcc_pieces nsub by auto
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> (None, h\<^sub>3, (v#vs, ls\<^sub>2, C,M,?pc\<^sub>2+1,ics)#frs,sh\<^sub>3)"
+ using SCall\<^sub>1.hyps(5) M'_code M'_in_D nclinit' by(cases T, auto)
+ finally show ?trans using nclinit by(auto simp:add.assoc)
+ qed
+ next
+ show ?Err (is "?throw \<longrightarrow> ?err")
+ proof
+ assume throw: ?throw
+ with IH_body obtain pc\<^sub>2 vs' where
+ pc\<^sub>2: "0 \<le> pc\<^sub>2 \<and> pc\<^sub>2 < size(compE\<^sub>2 body) \<and>
+ \<not> caught P pc\<^sub>2 h\<^sub>3 xa (compxE\<^sub>2 body 0 0)" and
+ 2: "P \<turnstile> ?\<sigma>\<^sub>2' -jvm\<rightarrow> handle P D M' xa h\<^sub>3 vs' ls\<^sub>3 pc\<^sub>2 No_ics ?frs\<^sub>2 sh\<^sub>3"
+ using SCall\<^sub>1.prems M'_code M'_xtab cond nsub_RI_Jcc_pieces nsub
+ by (auto simp del:split_paired_Ex)
+ have "handle P D M' xa h\<^sub>3 vs' ls\<^sub>3 pc\<^sub>2 No_ics ?frs\<^sub>2 sh\<^sub>3 =
+ handle P C M xa h\<^sub>3 (rev pvs @ vs) ls\<^sub>2 ?pc\<^sub>2 ics frs sh\<^sub>3"
+ using pc\<^sub>2 M'_in_D nclinit' by(auto simp add:handle_def)
+ then show "?err" using pc\<^sub>2 jvm_trans[OF 1 2] nclinit by(auto intro:exI[where x="?pc\<^sub>2"])
+ qed
+ qed
+ qed
+next
+ case Block\<^sub>1 then show ?case using nsub_RI_Jcc_pieces by auto
+next
+ case (Seq\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 w h\<^sub>1 ls\<^sub>1 sh\<^sub>1 e\<^sub>2 e\<^sub>2' h\<^sub>2 ls\<^sub>2 sh\<^sub>2)
+ let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e\<^sub>1)"
+ let ?\<sigma>\<^sub>0 = "(None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0)"
+ let ?\<sigma>\<^sub>1 = "(None,h\<^sub>1,(vs,ls\<^sub>1,C,M,?pc\<^sub>1+1,ics)#frs,sh\<^sub>1)"
+ let ?I = "I - pcs (compxE\<^sub>2 e\<^sub>2 (Suc ?pc\<^sub>1) (length vs))"
+ have Isub: "{pc..<pc + length (compE\<^sub>2 e\<^sub>1)} \<subseteq> ?I" using Seq\<^sub>1.prems by clarsimp
+ have IH: "PROP ?P e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (Val w) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics w xa vs frs ?I" by fact
+ have "P \<turnstile> ?\<sigma>\<^sub>0 -jvm\<rightarrow> (None,h\<^sub>1,(w#vs,ls\<^sub>1,C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)"
+ using Seq\<^sub>1.prems nsub_RI_Jcc_pieces IH Isub by auto
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> ?\<sigma>\<^sub>1" using Seq\<^sub>1 by auto
+ finally have eval\<^sub>1: "P \<turnstile> ?\<sigma>\<^sub>0 -jvm\<rightarrow> ?\<sigma>\<^sub>1".
+ let ?pc\<^sub>2 = "?pc\<^sub>1 + 1 + length(compE\<^sub>2 e\<^sub>2)"
+ let ?I' = "I - pcs(compxE\<^sub>2 e\<^sub>1 pc (size vs))"
+ have IH\<^sub>2: "PROP ?P e\<^sub>2 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 e\<^sub>2' h\<^sub>2 ls\<^sub>2 sh\<^sub>2 E C M (?pc\<^sub>1+1) ics v xa vs frs
+ ?I'" by fact
+ have Isub2: "{Suc (pc + length (compE\<^sub>2 e\<^sub>1))..<Suc (pc + length (compE\<^sub>2 e\<^sub>1) + length (compE\<^sub>2 e\<^sub>2))}
+ \<subseteq> ?I'" using Seq\<^sub>1.prems by clarsimp
+ show ?case (is "?Norm \<and> ?Err")
+ proof
+ show ?Norm (is "?val \<longrightarrow> ?trans")
+ proof
+ assume val: ?val
+ note eval\<^sub>1
+ also have "P \<turnstile> ?\<sigma>\<^sub>1 -jvm\<rightarrow> (None,h\<^sub>2,(v#vs,ls\<^sub>2,C,M,?pc\<^sub>2,ics)#frs,sh\<^sub>2)"
+ using val Seq\<^sub>1.prems nsub_RI_Jcc_pieces IH\<^sub>2 Isub2 by auto
+ finally show ?trans by(simp add:add.assoc)
+ qed
+ next
+ show ?Err (is "?throw \<longrightarrow> ?err")
+ proof
+ assume throw: ?throw
+ then obtain pc\<^sub>2 vs' where
+ pc\<^sub>2: "?pc\<^sub>1+1 \<le> pc\<^sub>2 \<and> pc\<^sub>2 < ?pc\<^sub>2 \<and>
+ \<not> caught P pc\<^sub>2 h\<^sub>2 xa (compxE\<^sub>2 e\<^sub>2 (?pc\<^sub>1+1) (size vs))" and
+ eval\<^sub>2: "P \<turnstile> ?\<sigma>\<^sub>1 -jvm\<rightarrow> handle P C M xa h\<^sub>2 (vs'@vs) ls\<^sub>2 pc\<^sub>2 ics frs sh\<^sub>2"
+ using IH\<^sub>2 Seq\<^sub>1.prems nsub_RI_Jcc_pieces Isub2 by auto
+ show "?err" using pc\<^sub>2 jvm_trans[OF eval\<^sub>1 eval\<^sub>2] by(auto intro: exI[where x=pc\<^sub>2])
+ qed
+ qed
+next
+ case (SeqThrow\<^sub>1 e\<^sub>0 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 e h\<^sub>1 ls\<^sub>1 sh\<^sub>1 e\<^sub>1)
+ let ?I = "I - pcs (compxE\<^sub>2 e\<^sub>1 (Suc (pc + length (compE\<^sub>2 e\<^sub>0))) (length vs))"
+ obtain a' where throw: "throw e = Throw a'" using eval\<^sub>1_final[OF SeqThrow\<^sub>1.hyps(1)] by clarsimp
+ have Isub: "{pc..<pc + length (compE\<^sub>2 e\<^sub>0)} \<subseteq> ?I" using SeqThrow\<^sub>1.prems by clarsimp
+ have "PROP ?P e\<^sub>0 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (throw e) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics v a' vs frs ?I" by fact
+ then show ?case using SeqThrow\<^sub>1.prems throw nsub_RI_Jcc_pieces Isub by auto
+next
+ case (CondT\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 e\<^sub>1 e' h\<^sub>2 ls\<^sub>2 sh\<^sub>2 e\<^sub>2)
+ let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e)"
+ let ?\<sigma>\<^sub>0 = "(None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0)"
+ let ?\<sigma>\<^sub>1 = "(None,h\<^sub>1,(vs,ls\<^sub>1,C,M,?pc\<^sub>1+1,ics)#frs,sh\<^sub>1)"
+ let ?d = "size vs"
+ let ?xt\<^sub>1 = "compxE\<^sub>2 e\<^sub>1 (pc+size(compE\<^sub>2 e)+1) ?d"
+ let ?xt\<^sub>2 = "compxE\<^sub>2 e\<^sub>2 (pc+size(compE\<^sub>2 e)+size(compE\<^sub>2 e\<^sub>1)+2) ?d"
+ let ?I = "I - (pcs ?xt\<^sub>1 \<union> pcs ?xt\<^sub>2)"
+ have Isub: "{pc..<pc + length (compE\<^sub>2 e)} \<subseteq> ?I" using CondT\<^sub>1.prems by clarsimp
+ have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 true h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Bool True) xa vs frs ?I" by fact
+ have "P \<turnstile> ?\<sigma>\<^sub>0 -jvm\<rightarrow> (None,h\<^sub>1,(Bool(True)#vs,ls\<^sub>1,C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)"
+ using CondT\<^sub>1.prems nsub_RI_Jcc_pieces IH Isub by(auto simp: Int_Un_distrib)
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> ?\<sigma>\<^sub>1" using CondT\<^sub>1 by auto
+ finally have eval\<^sub>1: "P \<turnstile> ?\<sigma>\<^sub>0 -jvm\<rightarrow> ?\<sigma>\<^sub>1".
+ let ?pc\<^sub>1' = "?pc\<^sub>1 + 1 + length(compE\<^sub>2 e\<^sub>1)"
+ let ?pc\<^sub>2' = "?pc\<^sub>1' + 1 + length(compE\<^sub>2 e\<^sub>2)"
+ let ?I' = "I - pcs(compxE\<^sub>2 e pc ?d) - pcs(compxE\<^sub>2 e\<^sub>2 (?pc\<^sub>1'+1) ?d)"
+ have IH2: "PROP ?P e\<^sub>1 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 e' h\<^sub>2 ls\<^sub>2 sh\<^sub>2 E C M (?pc\<^sub>1+1) ics v xa vs frs ?I'" by fact
+ show ?case (is "?Norm \<and> ?Err")
+ proof
+ show ?Norm (is "?val \<longrightarrow> ?trans")
+ proof
+ assume val: ?val
+ note eval\<^sub>1
+ also have "P \<turnstile> ?\<sigma>\<^sub>1 -jvm\<rightarrow> (None,h\<^sub>2,(v#vs,ls\<^sub>2,C,M,?pc\<^sub>1',ics)#frs,sh\<^sub>2)"
+ using val CondT\<^sub>1.prems nsub_RI_Jcc_pieces IH2 by(fastforce simp:Int_Un_distrib)
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> (None,h\<^sub>2,(v#vs,ls\<^sub>2,C,M,?pc\<^sub>2',ics)#frs,sh\<^sub>2)"
+ using CondT\<^sub>1 nsub_RI_Jcc_pieces by(auto simp:add.assoc)
+ finally show ?trans by(simp add:add.assoc)
+ qed
+ next
+ show ?Err (is "?throw \<longrightarrow> ?err")
+ proof
+ assume throw: ?throw
+ moreover
+ note IH2
+ ultimately obtain pc\<^sub>2 vs' where
+ pc\<^sub>2: "?pc\<^sub>1+1 \<le> pc\<^sub>2 \<and> pc\<^sub>2 < ?pc\<^sub>1' \<and>
+ \<not> caught P pc\<^sub>2 h\<^sub>2 xa (compxE\<^sub>2 e\<^sub>1 (?pc\<^sub>1+1) (size vs))" and
+ eval\<^sub>2: "P \<turnstile> ?\<sigma>\<^sub>1 -jvm\<rightarrow> handle P C M xa h\<^sub>2 (vs'@vs) ls\<^sub>2 pc\<^sub>2 ics frs sh\<^sub>2"
+ using CondT\<^sub>1.prems nsub_RI_Jcc_pieces by (fastforce simp:Int_Un_distrib)
+ show "?err" using pc\<^sub>2 jvm_trans[OF eval\<^sub>1 eval\<^sub>2] by(auto intro: exI[where x=pc\<^sub>2])
+ qed
+ qed
+next
+ case (CondF\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 e\<^sub>2 e' h\<^sub>2 ls\<^sub>2 sh\<^sub>2 e\<^sub>1)
+ let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e)"
+ let ?pc\<^sub>2 = "?pc\<^sub>1 + 1 + length(compE\<^sub>2 e\<^sub>1)+ 1"
+ let ?pc\<^sub>2' = "?pc\<^sub>2 + length(compE\<^sub>2 e\<^sub>2)"
+ let ?\<sigma>\<^sub>0 = "(None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0)"
+ let ?\<sigma>\<^sub>1 = "(None,h\<^sub>1,(vs,ls\<^sub>1,C,M,?pc\<^sub>2,ics)#frs,sh\<^sub>1)"
+ let ?d = "size vs"
+ let ?xt\<^sub>1 = "compxE\<^sub>2 e\<^sub>1 (pc+size(compE\<^sub>2 e)+1) ?d"
+ let ?xt\<^sub>2 = "compxE\<^sub>2 e\<^sub>2 (pc+size(compE\<^sub>2 e)+size(compE\<^sub>2 e\<^sub>1)+2) ?d"
+ let ?I = "I - (pcs ?xt\<^sub>1 \<union> pcs ?xt\<^sub>2)"
+ let ?I' = "I - pcs(compxE\<^sub>2 e pc ?d) - pcs(compxE\<^sub>2 e\<^sub>1 (?pc\<^sub>1+1) ?d)"
+ have pcs: "pcs(compxE\<^sub>2 e pc ?d) \<inter> pcs(?xt\<^sub>1 @ ?xt\<^sub>2) = {}"
+ using CondF\<^sub>1.prems by (simp add:Int_Un_distrib)
+ have Isub: "{pc..<pc + length (compE\<^sub>2 e)} \<subseteq> ?I" using CondF\<^sub>1.prems by clarsimp
+ have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 false h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Bool False) xa vs frs ?I" by fact
+ have IH2: "PROP ?P e\<^sub>2 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 e' h\<^sub>2 ls\<^sub>2 sh\<^sub>2 E C M ?pc\<^sub>2 ics v xa vs frs ?I'" by fact
+ have "P \<turnstile> ?\<sigma>\<^sub>0 -jvm\<rightarrow> (None,h\<^sub>1,(Bool(False)#vs,ls\<^sub>1,C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)"
+ using CondF\<^sub>1.prems nsub_RI_Jcc_pieces IH Isub pcs by auto
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> ?\<sigma>\<^sub>1" using CondF\<^sub>1 by auto
+ finally have eval\<^sub>1: "P \<turnstile> ?\<sigma>\<^sub>0 -jvm\<rightarrow> ?\<sigma>\<^sub>1".
+ show ?case (is "?Norm \<and> ?Err")
+ proof
+ show ?Norm (is "?val \<longrightarrow> ?trans")
+ proof
+ assume val: ?val
+ note eval\<^sub>1
+ also have "P \<turnstile> ?\<sigma>\<^sub>1 -jvm\<rightarrow> (None,h\<^sub>2,(v#vs,ls\<^sub>2,C,M,?pc\<^sub>2',ics)#frs,sh\<^sub>2)"
+ using val CondF\<^sub>1.prems nsub_RI_Jcc_pieces IH2 by(fastforce simp:Int_Un_distrib)
+ finally show ?trans by(simp add:add.assoc)
+ qed
+ next
+ show ?Err (is "?throw \<longrightarrow> ?err")
+ proof
+ let ?I' = "I - pcs(compxE\<^sub>2 e pc ?d) - pcs(compxE\<^sub>2 e\<^sub>1 (?pc\<^sub>1+1) ?d)"
+ assume throw: ?throw
+ then obtain pc\<^sub>2 vs' where
+ pc\<^sub>2: "?pc\<^sub>2 \<le> pc\<^sub>2 \<and> pc\<^sub>2 < ?pc\<^sub>2' \<and>
+ \<not> caught P pc\<^sub>2 h\<^sub>2 xa (compxE\<^sub>2 e\<^sub>2 ?pc\<^sub>2 ?d)" and
+ eval\<^sub>2: "P \<turnstile> ?\<sigma>\<^sub>1 -jvm\<rightarrow> handle P C M xa h\<^sub>2 (vs'@vs) ls\<^sub>2 pc\<^sub>2 ics frs sh\<^sub>2"
+ using CondF\<^sub>1.prems nsub_RI_Jcc_pieces IH2 by(fastforce simp:Int_Un_distrib)
+ show "?err" using pc\<^sub>2 jvm_trans[OF eval\<^sub>1 eval\<^sub>2] by(auto intro: exI[where x=pc\<^sub>2])
+ qed
+ qed
+next
+ case (CondThrow\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 f h\<^sub>1 ls\<^sub>1 sh\<^sub>1 e\<^sub>1 e\<^sub>2)
+ let ?d = "size vs"
+ let ?xt\<^sub>1 = "compxE\<^sub>2 e\<^sub>1 (pc+size(compE\<^sub>2 e)+1) ?d"
+ let ?xt\<^sub>2 = "compxE\<^sub>2 e\<^sub>2 (pc+size(compE\<^sub>2 e)+size(compE\<^sub>2 e\<^sub>1)+2) ?d"
+ let ?I = "I - (pcs ?xt\<^sub>1 \<union> pcs ?xt\<^sub>2)"
+ have Isub: "{pc..<pc + length (compE\<^sub>2 e)} \<subseteq> ?I" using CondThrow\<^sub>1.prems by clarsimp
+ have "pcs(compxE\<^sub>2 e pc ?d) \<inter> pcs(?xt\<^sub>1 @ ?xt\<^sub>2) = {}"
+ using CondThrow\<^sub>1.prems by (simp add:Int_Un_distrib)
+ moreover have "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (throw f) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics v xa vs frs ?I" by fact
+ ultimately show ?case using CondThrow\<^sub>1.prems nsub_RI_Jcc_pieces Isub by auto
+next
+ case (WhileF\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 c)
+ let ?pc = "pc + length(compE\<^sub>2 e)"
+ let ?pc' = "?pc + length(compE\<^sub>2 c) + 3"
+ have Isub: "{pc..<pc + length (compE\<^sub>2 e)} \<subseteq> I - pcs (compxE\<^sub>2 c (Suc (pc + length (compE\<^sub>2 e))) (length vs))"
+ using WhileF\<^sub>1.prems by clarsimp
+ have Isub2: "{Suc (pc + length (compE\<^sub>2 e))..<Suc (pc + length (compE\<^sub>2 e) + length (compE\<^sub>2 c))}
+ \<subseteq> I - pcs (compxE\<^sub>2 e pc (length vs))" using WhileF\<^sub>1.prems by clarsimp
+ have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 false h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Bool False) xa vs frs
+ (I - pcs (compxE\<^sub>2 c (Suc (pc + length (compE\<^sub>2 e))) (length vs)))" by fact
+ have "P \<turnstile> (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0) -jvm\<rightarrow>
+ (None,h\<^sub>1,(Bool False#vs,ls\<^sub>1,C,M,?pc,ics)#frs,sh\<^sub>1)"
+ using WhileF\<^sub>1.prems nsub_RI_Jcc_pieces IH Isub by auto
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> (None,h\<^sub>1,(vs,ls\<^sub>1,C,M,?pc',ics)#frs,sh\<^sub>1)"
+ using WhileF\<^sub>1 by (auto simp:add.assoc)
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> (None,h\<^sub>1,(Unit#vs,ls\<^sub>1,C,M,?pc'+1,ics)#frs,sh\<^sub>1)"
+ using WhileF\<^sub>1.prems by (auto simp:eval_nat_numeral)
+ finally show ?case by (simp add:add.assoc eval_nat_numeral)
+next
+ case (WhileT\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 c v\<^sub>1 h\<^sub>2 ls\<^sub>2 sh\<^sub>2 e\<^sub>3 h\<^sub>3 ls\<^sub>3 sh\<^sub>3)
+ let ?pc = "pc + length(compE\<^sub>2 e)"
+ let ?pc' = "?pc + length(compE\<^sub>2 c) + 1"
+ let ?\<sigma>\<^sub>0 = "(None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0)"
+ let ?\<sigma>\<^sub>2 = "(None,h\<^sub>2,(vs,ls\<^sub>2,C,M,pc,ics)#frs,sh\<^sub>2)"
+ have Isub: "{pc..<pc + length (compE\<^sub>2 e)} \<subseteq> I - pcs (compxE\<^sub>2 c (Suc (pc + length (compE\<^sub>2 e))) (length vs))"
+ using WhileT\<^sub>1.prems by clarsimp
+ have Isub2: "{Suc (pc + length (compE\<^sub>2 e))..<Suc (pc + length (compE\<^sub>2 e) + length (compE\<^sub>2 c))}
+ \<subseteq> I - pcs (compxE\<^sub>2 e pc (length vs))" using WhileT\<^sub>1.prems by clarsimp
+ have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 true h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Bool True) xa vs frs
+ (I - pcs (compxE\<^sub>2 c (Suc (pc + length (compE\<^sub>2 e))) (length vs)))" by fact
+ have IH2: "PROP ?P c h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (Val v\<^sub>1) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 E C M (Suc ?pc) ics v\<^sub>1 xa vs frs
+ (I - pcs (compxE\<^sub>2 e pc (length vs)))" by fact
+ have "P \<turnstile> ?\<sigma>\<^sub>0 -jvm\<rightarrow> (None,h\<^sub>1,(Bool True#vs,ls\<^sub>1,C,M,?pc,ics)#frs,sh\<^sub>1)"
+ using WhileT\<^sub>1.prems nsub_RI_Jcc_pieces IH Isub by auto
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> (None,h\<^sub>1,(vs,ls\<^sub>1,C,M,?pc+1,ics)#frs,sh\<^sub>1)"
+ using WhileT\<^sub>1.prems by auto
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> (None,h\<^sub>2,(v\<^sub>1#vs,ls\<^sub>2,C,M,?pc',ics)#frs,sh\<^sub>2)"
+ using WhileT\<^sub>1.prems nsub_RI_Jcc_pieces IH2 Isub2 by auto
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> ?\<sigma>\<^sub>2" using WhileT\<^sub>1.prems by auto
+ finally have 1: "P \<turnstile> ?\<sigma>\<^sub>0 -jvm\<rightarrow> ?\<sigma>\<^sub>2".
+ show ?case (is "?Norm \<and> ?Err")
+ proof
+ show ?Norm (is "?val \<longrightarrow> ?trans")
+ proof
+ assume val: ?val
+ note 1
+ also have "P \<turnstile> ?\<sigma>\<^sub>2 -jvm\<rightarrow> (None,h\<^sub>3,(v#vs,ls\<^sub>3,C,M,?pc'+3,ics)#frs,sh\<^sub>3)"
+ using val WhileT\<^sub>1 by (auto simp add:add.assoc eval_nat_numeral)
+ finally show ?trans by(simp add:add.assoc eval_nat_numeral)
+ qed
+ next
+ show ?Err (is "?throw \<longrightarrow> ?err")
+ proof
+ assume throw: ?throw
+ moreover
+ have "PROP ?P (while (e) c) h\<^sub>2 ls\<^sub>2 sh\<^sub>2 e\<^sub>3 h\<^sub>3 ls\<^sub>3 sh\<^sub>3 E C M pc ics v xa vs frs I" by fact
+ ultimately obtain pc\<^sub>2 vs' where
+ pc\<^sub>2: "pc \<le> pc\<^sub>2 \<and> pc\<^sub>2 < ?pc'+3 \<and>
+ \<not> caught P pc\<^sub>2 h\<^sub>3 xa (compxE\<^sub>2 (while (e) c) pc (size vs))" and
+ 2: "P \<turnstile> ?\<sigma>\<^sub>2 -jvm\<rightarrow> handle P C M xa h\<^sub>3 (vs'@vs) ls\<^sub>3 pc\<^sub>2 ics frs sh\<^sub>3"
+ using WhileT\<^sub>1.prems by (auto simp:add.assoc eval_nat_numeral)
+ show "?err" using pc\<^sub>2 jvm_trans[OF 1 2] by(auto intro: exI[where x=pc\<^sub>2])
+ qed
+ qed
+next
+ case (WhileCondThrow\<^sub>1 e h ls sh e' h' ls' sh' c)
+ let ?I = "I - pcs (compxE\<^sub>2 c (Suc (pc + length (compE\<^sub>2 e))) (length vs))"
+ obtain a' where throw: "throw e' = Throw a'" using eval\<^sub>1_final[OF WhileCondThrow\<^sub>1.hyps(1)] by clarsimp
+ have Isub: "{pc..<pc + length (compE\<^sub>2 e)} \<subseteq> ?I" using WhileCondThrow\<^sub>1.prems by clarsimp
+ have "PROP ?P e h ls sh (throw e') h' ls' sh' E C M pc ics v a' vs frs ?I" by fact
+ then show ?case using WhileCondThrow\<^sub>1.prems throw nsub_RI_Jcc_pieces Isub by auto
+next
+ case (WhileBodyThrow\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 c e' h\<^sub>2 ls\<^sub>2 sh\<^sub>2)
+ let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e)"
+ let ?\<sigma>\<^sub>0 = "(None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0)"
+ let ?\<sigma>\<^sub>1 = "(None,h\<^sub>1,(vs,ls\<^sub>1,C,M,?pc\<^sub>1+1,ics)#frs,sh\<^sub>1)"
+ let ?I = "I - pcs (compxE\<^sub>2 c (Suc (pc + length (compE\<^sub>2 e))) (length vs))"
+ have Isub: "{pc..<pc + length (compE\<^sub>2 e)} \<subseteq> ?I"
+ using WhileBodyThrow\<^sub>1.prems by clarsimp
+ have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 true h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Bool True) xa vs frs ?I" by fact
+ then have "P \<turnstile> ?\<sigma>\<^sub>0 -jvm\<rightarrow> (None,h\<^sub>1,(Bool(True)#vs,ls\<^sub>1,C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)"
+ using WhileBodyThrow\<^sub>1.prems nsub_RI_Jcc_pieces Isub by auto
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> ?\<sigma>\<^sub>1" using WhileBodyThrow\<^sub>1 by auto
+ finally have eval\<^sub>1: "P \<turnstile> ?\<sigma>\<^sub>0 -jvm\<rightarrow> ?\<sigma>\<^sub>1".
+ let ?pc\<^sub>1' = "?pc\<^sub>1 + 1 + length(compE\<^sub>2 c)"
+ show ?case (is "?Norm \<and> ?Err")
+ proof
+ show ?Norm by simp
+ next
+ show ?Err (is "?throw \<longrightarrow> ?err")
+ proof
+ assume throw: ?throw
+ moreover
+ have "PROP ?P c h\<^sub>1 ls\<^sub>1 sh\<^sub>1 (throw e') h\<^sub>2 ls\<^sub>2 sh\<^sub>2 E C M (?pc\<^sub>1+1) ics v xa vs frs
+ (I - pcs (compxE\<^sub>2 e pc (size vs)))" by fact
+ ultimately obtain pc\<^sub>2 vs' where
+ pc\<^sub>2: "?pc\<^sub>1+1 \<le> pc\<^sub>2 \<and> pc\<^sub>2 < ?pc\<^sub>1' \<and>
+ \<not> caught P pc\<^sub>2 h\<^sub>2 xa (compxE\<^sub>2 c (?pc\<^sub>1+1) (size vs))" and
+ eval\<^sub>2: "P \<turnstile> ?\<sigma>\<^sub>1 -jvm\<rightarrow> handle P C M xa h\<^sub>2 (vs'@vs) ls\<^sub>2 pc\<^sub>2 ics frs sh\<^sub>2"
+ using WhileBodyThrow\<^sub>1.prems nsub_RI_Jcc_pieces by (fastforce simp:Int_Un_distrib)
+ show "?err" using pc\<^sub>2 jvm_trans[OF eval\<^sub>1 eval\<^sub>2] by(auto intro: exI[where x=pc\<^sub>2])
+ qed
+ qed
+next
+ case (Throw\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 a h\<^sub>1 ls\<^sub>1 sh\<^sub>1)
+ let ?pc = "pc + size(compE\<^sub>2 e)"
+ have Isub: "{pc..<pc + length (compE\<^sub>2 e)} \<subseteq> I" using Throw\<^sub>1.prems by clarsimp
+ show ?case (is "?Norm \<and> ?Err")
+ proof
+ show ?Norm by simp
+ next
+ show ?Err (is "?throw \<longrightarrow> ?err")
+ proof
+ assume throw:?throw
+ have "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (addr a) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics (Addr a) a vs frs I" by fact
+ then have "P \<turnstile> (None, h\<^sub>0, (vs, ls\<^sub>0, C, M, pc, ics) # frs, sh\<^sub>0) -jvm\<rightarrow>
+ (None, h\<^sub>1, (Addr xa#vs, ls\<^sub>1, C, M, ?pc, ics) # frs, sh\<^sub>1)"
+ using Throw\<^sub>1 nsub_RI_Jcc_pieces Isub throw by auto
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> handle P C M xa h\<^sub>1 (Addr xa#vs) ls\<^sub>1 ?pc ics frs sh\<^sub>1"
+ using Throw\<^sub>1.prems by(auto simp add:handle_def)
+ finally show "?err" by(auto intro!: exI[where x="?pc"] exI[where x="[Addr xa]"])
+ qed
+ qed
+next
+ case (ThrowNull\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 h\<^sub>1 ls\<^sub>1 sh\<^sub>1)
+ let ?pc = "pc + size(compE\<^sub>2 e)"
+ let ?xa = "addr_of_sys_xcpt NullPointer"
+ have Isub: "{pc..<pc + length (compE\<^sub>2 e)} \<subseteq> I" using ThrowNull\<^sub>1.prems by clarsimp
+ show ?case (is "?Norm \<and> ?Err")
+ proof
+ show ?Norm by simp
+ next
+ show ?Err (is "?throw \<longrightarrow> ?err")
+ proof
+ assume throw: ?throw
+ have "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 null h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics Null xa vs frs I" by fact
+ then have "P \<turnstile> (None, h\<^sub>0, (vs, ls\<^sub>0, C, M, pc, ics) # frs, sh\<^sub>0) -jvm\<rightarrow>
+ (None, h\<^sub>1, (Null#vs, ls\<^sub>1, C, M, ?pc, ics) # frs, sh\<^sub>1)"
+ using ThrowNull\<^sub>1.prems nsub_RI_Jcc_pieces Isub by auto
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> handle P C M ?xa h\<^sub>1 (Null#vs) ls\<^sub>1 ?pc ics frs sh\<^sub>1"
+ using ThrowNull\<^sub>1.prems by(auto simp add:handle_def)
+ finally show "?err" using throw by(auto intro!: exI[where x="?pc"] exI[where x="[Null]"])
+ qed
+ qed
+next
+ case (ThrowThrow\<^sub>1 e h ls sh e' h' ls' sh')
+ obtain a' where throw: "throw e' = Throw a'" using eval\<^sub>1_final[OF ThrowThrow\<^sub>1.hyps(1)] by clarsimp
+ have Isub: "{pc..<pc + length (compE\<^sub>2 e)} \<subseteq> I" using ThrowThrow\<^sub>1.prems by clarsimp
+ have "PROP ?P e h ls sh (throw e') h' ls' sh' E C M pc ics v a' vs frs I" by fact
+ then show ?case using ThrowThrow\<^sub>1.prems throw nsub_RI_Jcc_pieces Isub by auto
+next
+ case (Try\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 v\<^sub>1 h\<^sub>1 ls\<^sub>1 sh\<^sub>1 Ci i e\<^sub>2)
+ let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e\<^sub>1)"
+ let ?pc\<^sub>1' = "?pc\<^sub>1 + 2 + length(compE\<^sub>2 e\<^sub>2)"
+ have "{pc..<pc+size(compE\<^sub>2 (try e\<^sub>1 catch(Ci i) e\<^sub>2))} \<subseteq> I" using Try\<^sub>1.prems by simp
+ also have "P,C,M \<rhd> compxE\<^sub>2 (try e\<^sub>1 catch(Ci i) e\<^sub>2) pc (size vs) / I,size vs"
+ using Try\<^sub>1.prems by simp
+ ultimately have "P,C,M \<rhd> compxE\<^sub>2 e\<^sub>1 pc (size vs) / {pc..<pc + length (compE\<^sub>2 e\<^sub>1)},size vs"
+ by(rule beforex_try)
+ hence "P \<turnstile> (None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0) -jvm\<rightarrow>
+ (None,h\<^sub>1,(v\<^sub>1#vs,ls\<^sub>1,C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)"
+ using Try\<^sub>1 nsub_RI_Jcc_pieces by auto blast
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> (None,h\<^sub>1,(v\<^sub>1#vs,ls\<^sub>1,C,M,?pc\<^sub>1',ics)#frs,sh\<^sub>1)"
+ using Try\<^sub>1.prems by auto
+ finally show ?case by (auto simp:add.assoc)
+next
+ case (TryCatch\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 a h\<^sub>1 ls\<^sub>1 sh\<^sub>1 D fs Ci i e\<^sub>2 e\<^sub>2' h\<^sub>2 ls\<^sub>2 sh\<^sub>2)
+ let ?e = "try e\<^sub>1 catch(Ci i) e\<^sub>2"
+ let ?xt = "compxE\<^sub>2 ?e pc (size vs)"
+ let ?\<sigma>\<^sub>0 = "(None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0)"
+ let ?ls\<^sub>1 = "ls\<^sub>1[i := Addr a]"
+ let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e\<^sub>1)"
+ let ?pc\<^sub>1' = "?pc\<^sub>1 + 2"
+ let ?\<sigma>\<^sub>1 = "(None,h\<^sub>1,(vs,?ls\<^sub>1,C,M, ?pc\<^sub>1',ics) # frs,sh\<^sub>1)"
+ have I: "{pc..<pc + length (compE\<^sub>2 (try e\<^sub>1 catch(Ci i) e\<^sub>2))} \<subseteq> I"
+ and beforex: "P,C,M \<rhd> ?xt/I,size vs" using TryCatch\<^sub>1.prems by simp+
+ have "P \<turnstile> ?\<sigma>\<^sub>0 -jvm\<rightarrow> (None,h\<^sub>1,((Addr a)#vs,ls\<^sub>1,C,M, ?pc\<^sub>1+1,ics) # frs,sh\<^sub>1)"
+ proof -
+ have ics: "ics = No_ics" using TryCatch\<^sub>1.prems by auto
+ have "PROP ?P e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (Throw a) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics v a vs frs {pc..<pc + length (compE\<^sub>2 e\<^sub>1)}"
+ by fact
+ moreover have "P,C,M \<rhd> compxE\<^sub>2 e\<^sub>1 pc (size vs)/{pc..<?pc\<^sub>1},size vs"
+ using beforex I pcs_subset by(force elim!: beforex_appendD1)
+ ultimately have
+ "\<exists>pc\<^sub>1. pc \<le> pc\<^sub>1 \<and> pc\<^sub>1 < ?pc\<^sub>1 \<and>
+ \<not> caught P pc\<^sub>1 h\<^sub>1 a (compxE\<^sub>2 e\<^sub>1 pc (size vs)) \<and>
+ (\<exists>vs'. P \<turnstile> ?\<sigma>\<^sub>0 -jvm\<rightarrow> handle P C M a h\<^sub>1 (vs'@vs) ls\<^sub>1 pc\<^sub>1 ics frs sh\<^sub>1)"
+ using TryCatch\<^sub>1.prems nsub_RI_Jcc_pieces by auto
+ then obtain pc\<^sub>1 vs' where
+ pc\<^sub>1_in_e\<^sub>1: "pc \<le> pc\<^sub>1" "pc\<^sub>1 < ?pc\<^sub>1" and
+ pc\<^sub>1_not_caught: "\<not> caught P pc\<^sub>1 h\<^sub>1 a (compxE\<^sub>2 e\<^sub>1 pc (size vs))" and
+ 0: "P \<turnstile> ?\<sigma>\<^sub>0 -jvm\<rightarrow> handle P C M a h\<^sub>1 (vs'@vs) ls\<^sub>1 pc\<^sub>1 ics frs sh\<^sub>1" by iprover
+ from beforex obtain xt\<^sub>0 xt\<^sub>1
+ where ex_tab: "ex_table_of P C M = xt\<^sub>0 @ ?xt @ xt\<^sub>1"
+ and disj: "pcs xt\<^sub>0 \<inter> I = {}" by(auto simp:beforex_def)
+ have hp: "h\<^sub>1 a = Some (D, fs)" "P\<^sub>1 \<turnstile> D \<preceq>\<^sup>* Ci" by fact+
+ have "pc\<^sub>1 \<notin> pcs xt\<^sub>0" using pc\<^sub>1_in_e\<^sub>1 I disj by auto
+ with pc\<^sub>1_in_e\<^sub>1 pc\<^sub>1_not_caught hp
+ show ?thesis using ex_tab 0 ics by(simp add:handle_def matches_ex_entry_def)
+ qed
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> ?\<sigma>\<^sub>1" using TryCatch\<^sub>1 by auto
+ finally have 1: "P \<turnstile> ?\<sigma>\<^sub>0 -jvm\<rightarrow> ?\<sigma>\<^sub>1" .
+ let ?pc\<^sub>2 = "?pc\<^sub>1' + length(compE\<^sub>2 e\<^sub>2)"
+ let ?I\<^sub>2 = "{?pc\<^sub>1' ..< ?pc\<^sub>2}"
+ have "P,C,M \<rhd> compxE\<^sub>2 ?e pc (size vs) / I,size vs" by fact
+ hence beforex\<^sub>2: "P,C,M \<rhd> compxE\<^sub>2 e\<^sub>2 ?pc\<^sub>1' (size vs) / ?I\<^sub>2, size vs"
+ using I pcs_subset[of _ ?pc\<^sub>1'] by(auto elim!:beforex_appendD2)
+ have IH\<^sub>2: "PROP ?P e\<^sub>2 h\<^sub>1 ?ls\<^sub>1 sh\<^sub>1 e\<^sub>2' h\<^sub>2 ls\<^sub>2 sh\<^sub>2 E C M ?pc\<^sub>1' ics v xa vs frs ?I\<^sub>2" by fact
+ show ?case (is "?Norm \<and> ?Err")
+ proof
+ show ?Norm (is "?val \<longrightarrow> ?trans")
+ proof
+ assume val: ?val
+ note 1 also have "P \<turnstile> ?\<sigma>\<^sub>1 -jvm\<rightarrow> (None,h\<^sub>2,(v#vs,ls\<^sub>2,C,M,?pc\<^sub>2,ics)#frs,sh\<^sub>2)"
+ using val beforex\<^sub>2 IH\<^sub>2 TryCatch\<^sub>1.prems nsub_RI_Jcc_pieces by auto
+ finally show ?trans by(simp add:add.assoc)
+ qed
+ next
+ show ?Err (is "?throw \<longrightarrow> ?err")
+ proof
+ assume throw: ?throw
+ then obtain pc\<^sub>2 vs' where
+ pc\<^sub>2: "?pc\<^sub>1+2 \<le> pc\<^sub>2 \<and> pc\<^sub>2 < ?pc\<^sub>2 \<and>
+ \<not> caught P pc\<^sub>2 h\<^sub>2 xa (compxE\<^sub>2 e\<^sub>2 ?pc\<^sub>1' (size vs))" and
+ 2: "P \<turnstile> ?\<sigma>\<^sub>1 -jvm\<rightarrow> handle P C M xa h\<^sub>2 (vs'@vs) ls\<^sub>2 pc\<^sub>2 ics frs sh\<^sub>2"
+ using IH\<^sub>2 beforex\<^sub>2 TryCatch\<^sub>1.prems nsub_RI_Jcc_pieces by auto
+ show ?err using pc\<^sub>2 jvm_trans[OF 1 2]
+ by (simp add:match_ex_entry) (auto intro: exI[where x=pc\<^sub>2])
+ qed
+ qed
+next
+ case (TryThrow\<^sub>1 e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 a h\<^sub>1 ls\<^sub>1 sh\<^sub>1 D fs Ci i e\<^sub>2)
+ let ?\<sigma>\<^sub>0 = "(None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0)"
+ let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e\<^sub>1)"
+ let ?e = "try e\<^sub>1 catch(Ci i) e\<^sub>2"
+ let ?xt = "compxE\<^sub>2 ?e pc (size vs)"
+ have I: "{pc..<pc + length (compE\<^sub>2 (try e\<^sub>1 catch(Ci i) e\<^sub>2))} \<subseteq> I"
+ and beforex: "P,C,M \<rhd> ?xt/I,size vs" using TryThrow\<^sub>1.prems by simp+
+ have "PROP ?P e\<^sub>1 h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (Throw a) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 E C M pc ics v a vs frs
+ {pc..<pc + length (compE\<^sub>2 e\<^sub>1)}" by fact
+ moreover have "P,C,M \<rhd> compxE\<^sub>2 e\<^sub>1 pc (size vs)/{pc..<?pc\<^sub>1},size vs"
+ using beforex I pcs_subset by(force elim!: beforex_appendD1)
+ ultimately have
+ "\<exists>pc\<^sub>1. pc \<le> pc\<^sub>1 \<and> pc\<^sub>1 < ?pc\<^sub>1 \<and>
+ \<not> caught P pc\<^sub>1 h\<^sub>1 a (compxE\<^sub>2 e\<^sub>1 pc (size vs)) \<and>
+ (\<exists>vs'. P \<turnstile> ?\<sigma>\<^sub>0 -jvm\<rightarrow> handle P C M a h\<^sub>1 (vs'@vs) ls\<^sub>1 pc\<^sub>1 ics frs sh\<^sub>1)"
+ using TryThrow\<^sub>1.prems nsub_RI_Jcc_pieces by auto
+ then obtain pc\<^sub>1 vs' where
+ pc\<^sub>1_in_e\<^sub>1: "pc \<le> pc\<^sub>1" "pc\<^sub>1 < ?pc\<^sub>1" and
+ pc\<^sub>1_not_caught: "\<not> caught P pc\<^sub>1 h\<^sub>1 a (compxE\<^sub>2 e\<^sub>1 pc (size vs))" and
+ 0: "P \<turnstile> ?\<sigma>\<^sub>0 -jvm\<rightarrow> handle P C M a h\<^sub>1 (vs'@vs) ls\<^sub>1 pc\<^sub>1 ics frs sh\<^sub>1" by iprover
+ show ?case (is "?N \<and> (?eq \<longrightarrow> ?err)")
+ proof
+ show ?N by simp
+ next
+ { assume ?eq
+ with TryThrow\<^sub>1 pc\<^sub>1_in_e\<^sub>1 pc\<^sub>1_not_caught 0
+ have "?err" by (simp add:match_ex_entry) auto
+ }
+ thus "?eq \<longrightarrow> ?err" by iprover
+ qed
+next
+ case Nil\<^sub>1 thus ?case by simp
+next
+ case (Cons\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 v h\<^sub>1 ls\<^sub>1 sh\<^sub>1 es fs h\<^sub>2 ls\<^sub>2 sh\<^sub>2)
+ let ?pc\<^sub>1 = "pc + length(compE\<^sub>2 e)"
+ let ?\<sigma>\<^sub>0 = "(None,h\<^sub>0,(vs,ls\<^sub>0,C,M,pc,ics)#frs,sh\<^sub>0)"
+ let ?\<sigma>\<^sub>1 = "(None,h\<^sub>1,(v#vs,ls\<^sub>1,C,M,?pc\<^sub>1,ics)#frs,sh\<^sub>1)"
+ have IH: "PROP ?P e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 (Val v) h\<^sub>1 ls\<^sub>1 sh\<^sub>1 [] C M pc ics v xa vs frs
+ (I - pcs (compxEs\<^sub>2 es ?pc\<^sub>1 (Suc (length vs))))" by fact
+ then have 1: "P \<turnstile> ?\<sigma>\<^sub>0 -jvm\<rightarrow> ?\<sigma>\<^sub>1" using Jcc_pieces_Cons[OF _ Cons\<^sub>1.prems(1-5)] by auto
+ let ?pc\<^sub>2 = "?pc\<^sub>1 + length(compEs\<^sub>2 es)"
+ have IHs: "PROP ?Ps es h\<^sub>1 ls\<^sub>1 sh\<^sub>1 fs h\<^sub>2 ls\<^sub>2 sh\<^sub>2 C M ?pc\<^sub>1 ics (tl ws) xa es' (v#vs) frs
+ (I - pcs (compxE\<^sub>2 e pc (length vs)))" by fact
+ show ?case (is "?Norm \<and> ?Err")
+ proof
+ show ?Norm (is "?val \<longrightarrow> ?trans")
+ proof
+ assume val: ?val
+ note 1
+ also have "P \<turnstile> ?\<sigma>\<^sub>1 -jvm\<rightarrow> (None,h\<^sub>2,(rev(ws) @ vs,ls\<^sub>2,C,M,?pc\<^sub>2,ics)#frs,sh\<^sub>2)"
+ using val IHs Cons\<^sub>1.prems by fastforce
+ finally show ?trans by(simp add:add.assoc)
+ qed
+ next
+ show ?Err (is "?throw \<longrightarrow> (\<exists>pc\<^sub>2. ?H pc\<^sub>2)")
+ proof
+ assume throw: ?throw
+ then obtain pc\<^sub>2 vs' where
+ pc\<^sub>2: "?pc\<^sub>1 \<le> pc\<^sub>2 \<and> pc\<^sub>2 < ?pc\<^sub>2 \<and>
+ \<not> caught P pc\<^sub>2 h\<^sub>2 xa (compxEs\<^sub>2 es ?pc\<^sub>1 (size vs + 1))" and
+ 2: "P \<turnstile> ?\<sigma>\<^sub>1 -jvm\<rightarrow> handle P C M xa h\<^sub>2 (vs'@v#vs) ls\<^sub>2 pc\<^sub>2 ics frs sh\<^sub>2"
+ using IHs Cons\<^sub>1.prems by(fastforce simp:Cons_eq_append_conv neq_Nil_conv)
+ have "?H pc\<^sub>2" using Cons\<^sub>1.prems pc\<^sub>2 jvm_trans[OF 1 2] by(auto intro!: exI[where x="vs'@[v]"])
+ thus "\<exists>pc\<^sub>2. ?H pc\<^sub>2" by iprover
+ qed
+ qed
+next
+ case (ConsThrow\<^sub>1 e h\<^sub>0 ls\<^sub>0 sh\<^sub>0 a h\<^sub>1 ls\<^sub>1 sh\<^sub>1 es)
+ then show ?case using Jcc_pieces_Cons[OF _ ConsThrow\<^sub>1.prems(1-5)]
+ by (fastforce simp:Cons_eq_append_conv)
+next
+ case InitFinal\<^sub>1 then show ?case using eval\<^sub>1_final_same[OF InitFinal\<^sub>1.hyps(1)] by clarsimp
+next
+ case (InitNone\<^sub>1 sh C\<^sub>0 C' Cs e h l e' h' l' sh')
+ then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' sh' v xa
+ (INIT C' (C\<^sub>0 # Cs,False) \<leftarrow> e)
+ = (True, frs', (None, h', (vs, l, C, M, pc, Called []) # frs, sh'), err)"
+ using InitNone\<^sub>1.prems(1) by clarsimp
+ let ?sh = "(sh(C\<^sub>0 \<mapsto> (sblank P\<^sub>1 C\<^sub>0, Prepared)))"
+ obtain ics: "ics_of(hd frs') = Calling C\<^sub>0 Cs"
+ and frs\<^sub>1: "frs' \<noteq> Nil" using pcs by clarsimp
+ then have 1: "P \<turnstile> (None,h,frs',sh) -jvm\<rightarrow> (None,h,frs',?sh)"
+ using InitNone\<^sub>1 jvm_InitNone[where P = P] by(cases frs', simp+)
+ show ?case (is "(?e1 \<longrightarrow> ?jvm1) \<and> (?e2 \<longrightarrow> ?err)")
+ proof(rule conjI)
+ { assume val: ?e1
+ note 1
+ also have "P \<turnstile> (None,h,frs',?sh) -jvm\<rightarrow> (None,h',(vs,l,C,M,pc,Called [])#frs,sh')"
+ using InitNone\<^sub>1.hyps(3)[of E] Jcc_pieces_InitNone[OF assms(1) pcs] InitNone\<^sub>1.prems val
+ by clarsimp
+ finally have ?jvm1 using pcs by simp
+ }
+ thus "?e1 \<longrightarrow> ?jvm1" by simp
+ next
+ { assume throw: ?e2
+ note 1
+ also obtain vs' where "P \<turnstile> (None,h,frs',?sh)
+ -jvm\<rightarrow> handle P C M xa h' (vs'@vs) l pc ics frs sh'"
+ using InitNone\<^sub>1.hyps(3)[of E] Jcc_pieces_InitNone[OF assms(1) pcs] throw
+ by clarsimp presburger
+ finally have ?err using pcs by auto
+ }
+ thus "?e2 \<longrightarrow> ?err" by simp
+ qed
+next
+ case (InitDone\<^sub>1 sh C\<^sub>0 sfs C' Cs e h l e' h' l' sh')
+ then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' sh' v xa
+ (INIT C' (C\<^sub>0 # Cs,False) \<leftarrow> e)
+ = (True, frs', (None, h', (vs, l, C, M, pc, Called []) # frs, sh'), err)"
+ using InitDone\<^sub>1.prems(1) by clarsimp
+ let ?frs' = "(calling_to_scalled (hd frs'))#(tl frs')"
+ have IH: "PROP ?P (INIT C' (Cs,True) \<leftarrow> e) h l sh e' h' l' sh' E C M pc ics v xa vs frs I"
+ by fact
+ obtain ics: "ics_of(hd frs') = Calling C\<^sub>0 Cs"
+ and frs\<^sub>1: "frs' \<noteq> Nil" using pcs by clarsimp
+ then have 1: "P \<turnstile> (None,h,frs',sh) -jvm\<rightarrow> (None,h,?frs',sh)"
+ using InitDone\<^sub>1 jvm_InitDP[where P = P] by(cases frs', simp+)
+ show ?case (is "(?e1 \<longrightarrow> ?jvm1) \<and> (?e2 \<longrightarrow> ?err)")
+ proof(rule conjI)
+ { assume val: ?e1
+ note 1
+ also have "P \<turnstile> (None,h,?frs',sh) -jvm\<rightarrow> (None,h',(vs,l,C,M,pc,Called [])#frs,sh')"
+ using IH Jcc_pieces_InitDP[OF assms(1) pcs] InitDone\<^sub>1.prems val by clarsimp
+ finally have ?jvm1 using pcs by simp
+ }
+ thus "?e1 \<longrightarrow> ?jvm1" by simp
+ next
+ { assume throw: ?e2
+ note 1
+ also obtain vs' where "P \<turnstile> (None,h,?frs',sh)
+ -jvm\<rightarrow> handle P C M xa h' (vs'@vs) l pc ics frs sh'"
+ using IH Jcc_pieces_InitDP[OF assms(1) pcs] InitDone\<^sub>1.prems throw by clarsimp
+ finally have ?err using pcs by auto
+ }
+ thus "?e2 \<longrightarrow> ?err" by simp
+ qed
+next
+ case (InitProcessing\<^sub>1 sh C\<^sub>0 sfs C' Cs e h l e' h' l' sh')
+ then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' sh' v xa
+ (INIT C' (C\<^sub>0 # Cs,False) \<leftarrow> e)
+ = (True, frs', (None, h', (vs, l, C, M, pc, Called []) # frs, sh'), err)"
+ using InitProcessing\<^sub>1.prems(1) by clarsimp
+ let ?frs' = "(calling_to_scalled (hd frs'))#(tl frs')"
+ have IH: "PROP ?P (INIT C' (Cs,True) \<leftarrow> e) h l sh e' h' l' sh' E C M pc ics v xa vs frs I"
+ by fact
+ obtain ics: "ics_of(hd frs') = Calling C\<^sub>0 Cs"
+ and frs\<^sub>1: "frs' \<noteq> Nil" using pcs by clarsimp
+ then have 1: "P \<turnstile> (None,h,frs',sh) -jvm\<rightarrow> (None,h,?frs',sh)"
+ using InitProcessing\<^sub>1 jvm_InitDP[where P = P] by(cases frs', simp+)
+ show ?case (is "(?e1 \<longrightarrow> ?jvm1) \<and> (?e2 \<longrightarrow> ?err)")
+ proof(rule conjI)
+ { assume val: ?e1
+ note 1
+ also have "P \<turnstile> (None,h,?frs',sh) -jvm\<rightarrow> (None,h',(vs,l,C,M,pc,Called [])#frs,sh')"
+ using IH Jcc_pieces_InitDP[OF assms(1) pcs] InitProcessing\<^sub>1.prems val by clarsimp
+ finally have ?jvm1 using pcs by simp
+ }
+ thus "?e1 \<longrightarrow> ?jvm1" by simp
+ next
+ { assume throw: ?e2
+ note 1
+ also obtain vs' where "P \<turnstile> (None,h,?frs',sh)
+ -jvm\<rightarrow> handle P C M xa h' (vs'@vs) l pc ics frs sh'"
+ using IH Jcc_pieces_InitDP[OF assms(1) pcs] InitProcessing\<^sub>1.prems throw by clarsimp
+ finally have ?err using pcs by auto
+ }
+ thus "?e2 \<longrightarrow> ?err" by simp
+ qed
+next
+ case (InitError\<^sub>1 sh C\<^sub>0 sfs Cs e h l e' h' l' sh' C')
+ then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' sh' v xa
+ (INIT C' (C\<^sub>0 # Cs,False) \<leftarrow> e)
+ = (True, frs', (None, h', (vs, l, C, M, pc, Called []) # frs, sh'), err)"
+ using InitError\<^sub>1.prems(1) by clarsimp
+ let ?e\<^sub>0 = "THROW NoClassDefFoundError"
+ let ?frs' = "(calling_to_sthrowing (hd frs') (addr_of_sys_xcpt NoClassDefFoundError))#(tl frs')"
+ have IH: "PROP ?P (RI (C\<^sub>0,?e\<^sub>0) ; Cs \<leftarrow> e) h l sh e' h' l' sh' E C M pc ics v xa vs frs I" by fact
+ obtain ics: "ics_of(hd frs') = Calling C\<^sub>0 Cs"
+ and frs\<^sub>1: "frs' \<noteq> Nil"
+ and tl: "tl frs' = frs" using pcs by clarsimp
+ then have 1: "P \<turnstile> (None,h,frs',sh) -jvm\<rightarrow> (None,h,?frs',sh)"
+ proof(cases frs')
+ case (Cons a list)
+ obtain vs' l' C' M' pc' ics' where a: "a = (vs',l',C',M',pc',ics')" by(cases a)
+ then have "ics' = Calling C\<^sub>0 Cs" using Cons ics by simp
+ then show ?thesis
+ using Cons a IH InitError\<^sub>1.prems jvm_InitError[where P = P] InitError\<^sub>1.hyps(1) by simp
+ qed(simp)
+ show ?case (is "(?e1 \<longrightarrow> ?jvm1) \<and> (?e2 \<longrightarrow> ?err)")
+ proof(rule conjI)
+ { assume val: ?e1
+ then have False using val rinit\<^sub>1_throw[OF InitError\<^sub>1.hyps(2)] by blast
+ then have ?jvm1 using pcs by simp
+ }
+ thus "?e1 \<longrightarrow> ?jvm1" by simp
+ next
+ { assume throw: ?e2
+ let ?frs = "(calling_to_throwing (hd frs') (addr_of_sys_xcpt NoClassDefFoundError))#(tl frs')"
+ have exec: "exec (P, (None,h,?frs,sh)) = Some (None,h,?frs',sh)"
+ using exec_ErrorThrowing[where sh=sh, OF InitError\<^sub>1.hyps(1)] ics by(cases "hd frs'", simp)
+ obtain vs' where 2: "P \<turnstile> (None,h,?frs,sh) -jvm\<rightarrow> handle P C M xa h' (vs'@vs) l pc ics frs sh'"
+ using IH Jcc_pieces_InitError[OF assms(1) pcs InitError\<^sub>1.hyps(1)] throw by clarsimp
+ have neq: "(None, h, ?frs, sh) \<noteq> handle P C M xa h' (vs' @ vs) l pc ics frs sh'"
+ using tl ics by(cases "hd frs'", simp add: handle_frs_tl_neq)
+
+ note 1
+ also have "P \<turnstile> (None,h,?frs',sh) -jvm\<rightarrow> handle P C M xa h' (vs'@vs) l pc ics frs sh'"
+ using exec_1_exec_all_conf[OF exec 2] neq by simp
+ finally have ?err using pcs by auto
+ }
+ thus "?e2 \<longrightarrow> ?err" by simp
+ qed
+next
+ case (InitObject\<^sub>1 sh C\<^sub>0 sfs sh' C' Cs e h l e' h' l' sh'')
+ then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l'
+ (sh(C\<^sub>0 \<mapsto> (sfs, Processing))) v xa (INIT C' (C\<^sub>0 # Cs,False) \<leftarrow> e)
+ = (True, frs', (None, h', (vs, l, C, M, pc, Called []) # frs, sh'), err)"
+ using InitObject\<^sub>1.prems(1) by clarsimp
+ let ?frs' = "(calling_to_called (hd frs'))#(tl frs')"
+ have IH: "PROP ?P (INIT C' (C\<^sub>0#Cs,True) \<leftarrow> e) h l sh' e' h' l' sh'' E C M pc ics v xa vs frs I"
+ by fact
+ obtain ics: "ics_of(hd frs') = Calling C\<^sub>0 Cs"
+ and frs\<^sub>1: "frs' \<noteq> Nil" using pcs by clarsimp
+ then have 1: "P \<turnstile> (None,h,frs',sh) -jvm\<rightarrow> (None,h,?frs',sh')"
+ proof(cases frs')
+ case (Cons a list)
+ obtain vs' l' C' M' pc' ics' where a: "a = (vs',l',C',M',pc',ics')" by(cases a)
+ then have "ics' = Calling C\<^sub>0 Cs" using Cons ics by simp
+ then show ?thesis
+ using Cons Nil a IH InitObject\<^sub>1 jvm_InitObj[where P = P] by simp
+ qed(simp)
+ show ?case (is "(?e1 \<longrightarrow> ?jvm1) \<and> (?e2 \<longrightarrow> ?err)")
+ proof(rule conjI)
+ { assume val: ?e1
+ note 1
+ also have "P \<turnstile> (None,h,?frs',sh') -jvm\<rightarrow> (None,h',(vs,l,C,M,pc,Called [])#frs,sh'')"
+ using IH Jcc_pieces_InitObj[OF assms(1) pcs] InitObject\<^sub>1 val by simp
+ finally have ?jvm1 using pcs by simp
+ }
+ thus "?e1 \<longrightarrow> ?jvm1" by simp
+ next
+ { assume throw: ?e2
+ note 1
+ also obtain vs' where "P \<turnstile> (None,h,?frs',sh')
+ -jvm\<rightarrow> handle P C M xa h' (vs'@vs) l pc ics frs sh''"
+ using IH Jcc_pieces_InitObj[OF assms(1) pcs] InitObject\<^sub>1 throw by clarsimp
+ finally have ?err using pcs by auto
+ }
+ thus "?e2 \<longrightarrow> ?err" by simp
+ qed
+next
+ case (InitNonObject\<^sub>1 sh C\<^sub>0 sfs D a b sh' C' Cs e h l e' h' l' sh'')
+ then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l'
+ (sh(C\<^sub>0 \<mapsto> (sfs,Processing))) v xa (INIT C' (C\<^sub>0 # Cs,False) \<leftarrow> e)
+ = (True, frs', (None, h', (vs, l, C, M, pc, Called []) # frs, sh'), err)"
+ using InitNonObject\<^sub>1.prems(1) by clarsimp
+ let ?frs' = "(calling_to_calling (hd frs') D)#(tl frs')"
+ have cls1: "is_class P\<^sub>1 D" using InitNonObject\<^sub>1.hyps(2,3) class_wf wf wf_cdecl_supD by blast
+ have cls_aux: "distinct (C\<^sub>0#Cs) \<and> supercls_lst P\<^sub>1 (C\<^sub>0#Cs)" using InitNonObject\<^sub>1.prems(1) by auto
+ then have cls2: "D \<notin> set (C\<^sub>0 # Cs)"
+ proof -
+ have "distinct (D # C\<^sub>0 # Cs)"
+ using InitNonObject\<^sub>1.hyps(2,3) cls_aux wf wf_supercls_distinct_app by blast
+ then show "D \<notin> set (C\<^sub>0 # Cs)"
+ by (metis distinct.simps(2))
+ qed
+ have cls3: "\<forall>C\<in>set (C\<^sub>0 # Cs). P\<^sub>1 \<turnstile> C \<preceq>\<^sup>* D" using InitNonObject\<^sub>1.hyps(2,3) cls_aux
+ by (metis r_into_rtrancl rtrancl_into_rtrancl set_ConsD subcls1.subcls1I supercls_lst.simps(1))
+ have IH: "PROP ?P (INIT C' (D # C\<^sub>0 # Cs,False) \<leftarrow> e) h l sh' e' h' l' sh'' E C M pc ics v xa vs frs I"
+ by fact
+ obtain r where cls: "class P C\<^sub>0 = \<lfloor>(D, r)\<rfloor>" using InitNonObject\<^sub>1.hyps(3)
+ by (metis assms class_compP compP\<^sub>2_def)
+ obtain ics: "ics_of(hd frs') = Calling C\<^sub>0 Cs"
+ and frs\<^sub>1: "frs' \<noteq> Nil" using pcs by clarsimp
+ then have 1: "P \<turnstile> (None,h,frs',sh) -jvm\<rightarrow> (None,h,?frs',sh')"
+ proof(cases frs')
+ case (Cons a list)
+ obtain vs' l' C' M' pc' ics' where a: "a = (vs',l',C',M',pc',ics')" by(cases a)
+ then have "ics' = Calling C\<^sub>0 Cs" using Cons ics by simp
+ then show ?thesis
+ using Cons a IH InitNonObject\<^sub>1 jvm_InitNonObj[OF _ _ cls] by simp
+ qed(simp)
+ show ?case (is "(?e1 \<longrightarrow> ?jvm1) \<and> (?e2 \<longrightarrow> ?err)")
+ proof(rule conjI)
+ { assume val: ?e1
+ note 1
+ also have "P \<turnstile> (None,h,?frs',sh') -jvm\<rightarrow> (None,h',(vs,l,C,M,pc,Called [])#frs,sh'')"
+ using IH Jcc_pieces_InitNonObj[OF assms(1) cls1 cls2 cls3 pcs] InitNonObject\<^sub>1 val by simp
+ finally have ?jvm1 using pcs by simp
+ }
+ thus "?e1 \<longrightarrow> ?jvm1" by simp
+ next
+ { assume throw: ?e2
+ note 1
+ also obtain vs' where "P \<turnstile> (None,h,?frs',sh')
+ -jvm\<rightarrow> handle P C M xa h' (vs'@vs) l pc ics frs sh''"
+ using IH Jcc_pieces_InitNonObj[OF assms(1) cls1 cls2 cls3 pcs] InitNonObject\<^sub>1 throw by clarsimp
+ finally have ?err using pcs by auto
+ }
+ thus "?e2 \<longrightarrow> ?err" by simp
+ qed
+next
+ case (InitRInit\<^sub>1 C\<^sub>0 Cs e h l sh e' h' l' sh' C')
+ then obtain frs' err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' sh' v xa
+ (INIT C' (C\<^sub>0 # Cs,True) \<leftarrow> e)
+ = (True, frs', (None, h', (vs, l, C, M, pc, Called []) # frs, sh'), err)"
+ using InitRInit\<^sub>1.prems(1) by clarsimp
+ have IH: "PROP ?P (RI (C\<^sub>0,C\<^sub>0\<bullet>\<^sub>sclinit([])) ; Cs \<leftarrow> e) h l sh e' h' l' sh' E C M pc ics v xa vs frs I"
+ by fact
+ show ?case (is "(?e1 \<longrightarrow> ?jvm1) \<and> (?e2 \<longrightarrow> ?err)")
+ proof(rule conjI)
+ { assume val: ?e1
+ have "P \<turnstile> (None,h,frs',sh) -jvm\<rightarrow> (None,h',(vs,l,C,M,pc,Called [])#frs,sh')"
+ using IH Jcc_pieces_InitRInit[OF assms(1,2) pcs] InitRInit\<^sub>1.prems val by simp
+ then have ?jvm1 using pcs by simp
+ }
+ thus "?e1 \<longrightarrow> ?jvm1" by simp
+ next
+ { assume throw: ?e2
+ obtain vs' where "P \<turnstile> (None,h,frs',sh)
+ -jvm\<rightarrow> handle P C M xa h' (vs'@vs) l pc ics frs sh'"
+ using IH Jcc_pieces_InitRInit[OF assms(1,2) pcs] InitRInit\<^sub>1 throw by clarsimp
+ then have ?err using pcs by auto
+ }
+ thus "?e2 \<longrightarrow> ?err" by simp
+ qed
+next
+ case (RInit\<^sub>1 e h l sh v1 h' l' sh' C\<^sub>0 sfs i sh'' C' Cs e' e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1)
+ let ?frs = "(vs,l,C,M,pc,Called (C\<^sub>0#Cs)) # frs"
+ let ?frs' = "(vs,l,C,M,pc,Called Cs) # frs"
+ have clinit: "e = C\<^sub>0\<bullet>\<^sub>sclinit([])" using RInit\<^sub>1
+ by (metis Jcc_cond.simps(2) eval\<^sub>1_final_same exp.distinct(101) final_def)
+ then obtain err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h\<^sub>1 l\<^sub>1 sh\<^sub>1 v xa
+ (RI (C\<^sub>0,C\<^sub>0\<bullet>\<^sub>sclinit([])) ; Cs \<leftarrow> e')
+ = (True, ?frs, (None, h\<^sub>1, (vs, l, C, M, pc, Called []) # frs, sh\<^sub>1), err)"
+ using RInit\<^sub>1.prems(1) by simp
+ have shC: "\<forall>C'\<in>set Cs. \<exists>sfs. sh C' = \<lfloor>(sfs, Processing)\<rfloor>" using RInit\<^sub>1.prems(1) clinit by clarsimp
+ then have shC'': "\<forall>C'\<in>set Cs. \<exists>sfs. sh'' C' = \<lfloor>(sfs, Processing)\<rfloor>"
+ using clinit\<^sub>1_proc_pres[OF wf] RInit\<^sub>1.hyps(1) clinit RInit\<^sub>1.hyps(4) RInit\<^sub>1.prems(1)
+ by (auto simp: fun_upd_apply)
+ have loc: "l = l'" using clinit\<^sub>1_loc_pres RInit\<^sub>1.hyps(1) clinit by simp
+ have IH: "PROP ?P e h l sh (Val v1) h' l' sh' E C M pc (Called Cs) v1 xa vs (tl ?frs') I" by fact
+ then have IH':
+ "PROP ?P (C\<^sub>0\<bullet>\<^sub>sclinit([])) h l sh (Val v1) h' l' sh' E C M pc (Called Cs) v1 xa vs (tl ?frs') I"
+ using clinit by simp
+ have IH2: "PROP ?P (INIT C' (Cs,True) \<leftarrow> e') h' l' sh'' e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 E C M
+ pc ics v xa vs frs I" by fact
+ have "P \<turnstile> (None,h,?frs,sh) -jvm\<rightarrow> (None,h,create_init_frame P C\<^sub>0 # ?frs',sh)" by(rule jvm_Called)
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> (None,h',?frs',sh'')"
+ using IH' Jcc_pieces_RInit_clinit[OF assms(1-2) pcs,of h' l' sh'] RInit\<^sub>1.hyps(3,4) by simp
+ finally have jvm1: "P \<turnstile> (None,h,?frs,sh) -jvm\<rightarrow> (None,h',?frs',sh'')" .
+ show ?case (is "(?e1 \<longrightarrow> ?jvm1) \<and> (?e2 \<longrightarrow> ?err)")
+ proof(rule conjI)
+ { assume val: ?e1
+ note jvm1
+ also have "P \<turnstile> (None,h',?frs',sh'') -jvm\<rightarrow> (None,h\<^sub>1,(vs,l,C,M,pc,Called [])#frs,sh\<^sub>1)"
+ using IH2 Jcc_pieces_RInit_Init[OF assms(1-2) shC'' pcs,of h'] RInit\<^sub>1.hyps(5) loc val by auto
+ finally have ?jvm1 using pcs clinit by simp
+ }
+ thus "?e1 \<longrightarrow> ?jvm1" by simp
+ next
+ { assume throw: ?e2
+ note jvm1
+ also obtain vs' where "P \<turnstile> (None,h',?frs',sh'')
+ -jvm\<rightarrow> handle P C M xa h\<^sub>1 (vs'@vs) l pc ics frs sh\<^sub>1"
+ using IH2 Jcc_pieces_RInit_Init[OF assms(1-2) shC'' pcs,of h'] RInit\<^sub>1.hyps(5) loc throw by auto
+ finally have ?err using pcs clinit by auto
+ }
+ thus "?e2 \<longrightarrow> ?err" by simp
+ qed
+next
+ case (RInitInitFail\<^sub>1 e h l sh a h' l' sh' C\<^sub>0 sfs i sh'' D Cs e' e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1)
+ let ?frs = "(vs,l,C,M,pc,Called (C\<^sub>0#D#Cs)) # frs"
+ let ?frs' = "(vs,l,C,M,pc,Called (D#Cs)) # frs"
+ let "?frsT" = "\<lambda>xa1. (vs,l,C,M,pc,Throwing (C\<^sub>0#D#Cs) xa1) # frs"
+ let "?frsT'" = "\<lambda>xa1. (vs,l,C,M,pc,Throwing (D#Cs) xa1) # frs"
+ obtain xa' where xa': "throw a = Throw xa'"
+ by (metis RInitInitFail\<^sub>1.hyps(1) eval\<^sub>1_final exp.distinct(101) final_def)
+ have e\<^sub>1: "e\<^sub>1 = Throw xa'" using xa' rinit\<^sub>1_throw RInitInitFail\<^sub>1.hyps(5) by simp
+ show ?case
+ proof(cases "e = C\<^sub>0\<bullet>\<^sub>sclinit([])")
+ case clinit: True
+ then obtain err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h\<^sub>1 l\<^sub>1 sh\<^sub>1 v xa'
+ (RI (C\<^sub>0,C\<^sub>0\<bullet>\<^sub>sclinit([])) ; D # Cs \<leftarrow> e')
+ = (True, ?frs, (None, h\<^sub>1, (vs, l, C, M, pc, Called []) # frs, sh\<^sub>1), err)"
+ using RInitInitFail\<^sub>1.prems(1) by simp
+ have loc: "l = l'" using clinit\<^sub>1_loc_pres RInitInitFail\<^sub>1.hyps(1) clinit by simp
+ have IH: "PROP ?P e h l sh (throw a) h' l' sh' E C M pc (Called (D#Cs)) v xa' vs frs I"
+ by fact
+ then have IH':
+ "PROP ?P (C\<^sub>0\<bullet>\<^sub>sclinit([])) h l sh (Throw xa') h' l' sh' E C M pc (Called (D#Cs)) v xa' vs
+ frs I" using clinit xa' by simp
+ have IH2: "PROP ?P (RI (D,throw a) ; Cs \<leftarrow> e') h' l' sh'' e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 E C M
+ pc ics v xa' vs frs I" by fact
+ have "P \<turnstile> (None,h,?frs,sh) -jvm\<rightarrow> (None,h,create_init_frame P C\<^sub>0 # ?frs',sh)" by(rule jvm_Called)
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> (None,h',(vs, l, C, M, pc, Throwing (D#Cs) xa') # frs,sh'')"
+ using IH' Jcc_pieces_RInit_clinit[OF assms(1-2) pcs,of h' l' sh'] RInitInitFail\<^sub>1.hyps(3,4)
+ by simp
+ also obtain vs'' where "P \<turnstile> \<dots> -jvm\<rightarrow> handle P C M xa' h\<^sub>1 (vs''@vs) l pc ics frs sh\<^sub>1"
+ using IH2 pcs Jcc_pieces_RInit_RInit[OF assms(1) pcs] RInitInitFail\<^sub>1.hyps(3,4)
+ xa' loc e\<^sub>1 xa' by clarsimp
+ finally show ?thesis using pcs e\<^sub>1 clinit by auto
+ next
+ case throw: False
+ then have eT: "e = Throw xa'" "h = h'" "l = l'" "sh = sh'" using xa' RInitInitFail\<^sub>1.prems(1)
+ eval\<^sub>1_final_same[OF RInitInitFail\<^sub>1.hyps(1)] by clarsimp+
+ obtain a' where "class P\<^sub>1 C\<^sub>0 = \<lfloor>a'\<rfloor>" using RInitInitFail\<^sub>1.prems by(auto simp: is_class_def)
+ then obtain stk' loc' M' pc' ics' where "create_init_frame P C\<^sub>0 = (stk',loc',C\<^sub>0,M',pc',ics')"
+ using create_init_frame_wf_eq[OF wf] by(cases "create_init_frame P C\<^sub>0", simp)
+ then obtain rhs err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' sh'' v xa'
+ (RI (C\<^sub>0,e) ; D#Cs \<leftarrow> e') = (True, ?frsT xa', rhs, err)"
+ using RInitInitFail\<^sub>1.prems(1) eT by clarsimp
+ have IH2: "PROP ?P (RI (D,throw a) ; Cs \<leftarrow> e') h' l' sh'' e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 E C M
+ pc ics v xa' vs frs I" by fact
+ have "P \<turnstile> (None,h,?frsT xa',sh') -jvm\<rightarrow> (None,h,?frsT' xa',sh'(C\<^sub>0 \<mapsto> (fst (the (sh' C\<^sub>0)), Error)))"
+ by(rule jvm_Throwing)
+ also obtain vs' where "P \<turnstile> ... -jvm\<rightarrow> handle P C M xa' h\<^sub>1 (vs'@vs) l pc ics frs sh\<^sub>1"
+ using IH2 Jcc_pieces_RInit_RInit[OF assms(1) pcs] RInitInitFail\<^sub>1.hyps(3,4)
+ eT e\<^sub>1 xa' by clarsimp
+ finally show ?thesis using pcs e\<^sub>1 throw eT by auto
+ qed
+next
+ case (RInitFailFinal\<^sub>1 e h l sh a h' l' sh' C\<^sub>0 sfs i sh'' e'')
+ let ?frs = "(vs,l,C,M,pc,Called [C\<^sub>0]) # frs"
+ let ?frs' = "(vs,l,C,M,pc,Called []) # frs"
+ let "?frsT" = "\<lambda>xa1. (vs,l,C,M,pc,Throwing [C\<^sub>0] xa1) # frs"
+ let "?frsT'" = "\<lambda>xa1. (vs,l,C,M,pc,Throwing [] xa1) # frs"
+ obtain xa' where xa': "throw a = Throw xa'"
+ by (metis RInitFailFinal\<^sub>1.hyps(1) eval\<^sub>1_final exp.distinct(101) final_def)
+ show ?case
+ proof(cases "e = C\<^sub>0\<bullet>\<^sub>sclinit([])")
+ case clinit: True
+ then obtain err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' sh'' v xa'
+ (RI (C\<^sub>0,C\<^sub>0\<bullet>\<^sub>sclinit([])) ; [] \<leftarrow> unit) = (True, ?frs, (None, h', ?frs', sh''), err)"
+ using RInitFailFinal\<^sub>1.prems(1) by clarsimp
+ have IH: "PROP ?P e h l sh (throw a) h' l' sh' E C M pc (Called []) v xa' vs frs I" by fact
+ then have IH':
+ "PROP ?P (C\<^sub>0\<bullet>\<^sub>sclinit([])) h l sh (throw a) h' l' sh' E C M pc (Called []) v xa' vs frs I"
+ using clinit by simp
+ have "P \<turnstile> (None,h,?frs,sh) -jvm\<rightarrow> (None,h,create_init_frame P C\<^sub>0 # ?frs',sh)"
+ by(rule jvm_Called)
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> (None,h',?frsT' xa',sh'')"
+ using IH' Jcc_pieces_RInit_clinit[OF assms(1-2) pcs,of h' l' sh'] xa'
+ RInitFailFinal\<^sub>1.hyps(3,4) by simp
+ also have
+ "P \<turnstile> \<dots> -jvm\<rightarrow> handle (compP compMb\<^sub>2 P\<^sub>1) C M xa' h' vs l pc No_ics frs sh''"
+ using RInitFailFinal\<^sub>1.hyps(3,4) jvm_RInit_throw[where h=h' and sh=sh''] by simp
+ finally show ?thesis using xa' pcs clinit by(clarsimp intro!: exI[where x="[]"])
+ next
+ case throw: False
+ then have eT: "e = Throw xa'" "h = h'" "sh = sh'" using xa' RInitFailFinal\<^sub>1.prems(1)
+ eval\<^sub>1_final_same[OF RInitFailFinal\<^sub>1.hyps(1)] by clarsimp+
+ obtain a where "class P\<^sub>1 C\<^sub>0 = \<lfloor>a\<rfloor>" using RInitFailFinal\<^sub>1.prems by(auto simp: is_class_def)
+ then obtain stk' loc' M' pc' ics' where "create_init_frame P C\<^sub>0 = (stk',loc',C\<^sub>0,M',pc',ics')"
+ using create_init_frame_wf_eq[OF wf] by(cases "create_init_frame P C\<^sub>0", simp)
+ then obtain rhs err where pcs: "Jcc_pieces P\<^sub>1 E C M h vs l pc ics frs sh I h' l' sh'' v xa'
+ (RI (C\<^sub>0,e) ; [] \<leftarrow> unit) = (True, ?frsT xa', rhs, err)"
+ using RInitFailFinal\<^sub>1.prems(1) eT by clarsimp
+ have "P \<turnstile> (None,h,?frsT xa',sh) -jvm\<rightarrow> (None,h,?frsT' xa',sh(C\<^sub>0 \<mapsto> (fst (the (sh C\<^sub>0)), Error)))"
+ by(rule jvm_Throwing)
+ also have "P \<turnstile> \<dots> -jvm\<rightarrow> handle P C M xa' h' vs l pc No_ics frs sh''"
+ using RInitFailFinal\<^sub>1.hyps(3,4) jvm_RInit_throw[where h=h and sh=sh''] eT by simp
+ finally show ?thesis using pcs xa' by(clarsimp intro!: exI[where x="[]"])
+ qed
+qed
+(*>*)
+
+(*FIXME move! *)
+lemma atLeast0AtMost[simp]: "{0::nat..n} = {..n}"
+by auto
+
+lemma atLeast0LessThan[simp]: "{0::nat..<n} = {..<n}"
+by auto
+
+fun exception :: "'a exp \<Rightarrow> addr option" where
+ "exception (Throw a) = Some a"
+| "exception e = None"
+
+lemma comp\<^sub>2_correct:
+assumes wf: "wf_J\<^sub>1_prog P\<^sub>1"
+ and "method": "P\<^sub>1 \<turnstile> C sees M,b:Ts\<rightarrow>T = body in C"
+ and eval: "P\<^sub>1 \<turnstile>\<^sub>1 \<langle>body,(h,ls,sh)\<rangle> \<Rightarrow> \<langle>e',(h',ls',sh')\<rangle>"
+ and nclinit: "M \<noteq> clinit"
+shows "compP\<^sub>2 P\<^sub>1 \<turnstile> (None,h,[([],ls,C,M,0,No_ics)],sh) -jvm\<rightarrow> (exception e',h',[],sh')"
+(*<*)
+ (is "_ \<turnstile> ?\<sigma>\<^sub>0 -jvm\<rightarrow> ?\<sigma>\<^sub>1")
+proof -
+ let ?P = "compP\<^sub>2 P\<^sub>1"
+ let ?E = "case b of Static \<Rightarrow> Ts | NonStatic \<Rightarrow> Class C#Ts"
+ have nsub: "\<not>sub_RI body" using sees_wf\<^sub>1_nsub_RI[OF wf method] by simp
+ have code: "?P,C,M,0 \<rhd> compE\<^sub>2 body" using beforeM[OF "method"] by auto
+ have xtab: "?P,C,M \<rhd> compxE\<^sub>2 body 0 (size[])/{..<size(compE\<^sub>2 body)},size[]"
+ using beforexM[OF "method"] by auto
+ have cond: "Jcc_cond P\<^sub>1 ?E C M [] 0 No_ics {..<size(compE\<^sub>2 body)} h sh body"
+ using nsub_RI_Jcc_pieces nsub code xtab by auto
+ \<comment> \<open>Distinguish if e' is a value or an exception\<close>
+ { fix v assume [simp]: "e' = Val v"
+ have "?P \<turnstile> ?\<sigma>\<^sub>0 -jvm\<rightarrow> (None,h',[([v],ls',C,M,size(compE\<^sub>2 body),No_ics)],sh')"
+ using Jcc[OF wf eval cond] nsub_RI_Jcc_pieces[OF _ nsub] by auto
+ also have "?P \<turnstile> \<dots> -jvm\<rightarrow> ?\<sigma>\<^sub>1" using beforeM[OF "method"] nclinit by auto
+ finally have ?thesis .
+ }
+ moreover
+ { fix a assume [simp]: "e' = Throw a"
+ obtain pc vs' where pc: "0 \<le> pc \<and> pc < size(compE\<^sub>2 body) \<and>
+ \<not> caught ?P pc h' a (compxE\<^sub>2 body 0 0)"
+ and 1: "?P \<turnstile> ?\<sigma>\<^sub>0 -jvm\<rightarrow> handle ?P C M a h' vs' ls' pc No_ics [] sh'"
+ using Jcc[OF wf eval cond] nsub_RI_Jcc_pieces[OF _ nsub] by auto meson
+ from pc have "handle ?P C M a h' vs' ls' pc No_ics [] sh' = ?\<sigma>\<^sub>1" using xtab "method" nclinit
+ by(auto simp:handle_def compMb\<^sub>2_def)
+ with 1 have ?thesis by simp
+ }
+ ultimately show ?thesis using eval\<^sub>1_final[OF eval] by(auto simp:final_def)
+qed
+(*>*)
+
+end
diff --git a/thys/JinjaDCI/Compiler/Hidden.thy b/thys/JinjaDCI/Compiler/Hidden.thy
--- a/thys/JinjaDCI/Compiler/Hidden.thy
+++ b/thys/JinjaDCI/Compiler/Hidden.thy
@@ -1,52 +1,52 @@
-theory Hidden
-imports "List-Index.List_Index"
-begin
-
-definition hidden :: "'a list \<Rightarrow> nat \<Rightarrow> bool" where
-"hidden xs i \<equiv> i < size xs \<and> xs!i \<in> set(drop (i+1) xs)"
-
-
-lemma hidden_last_index: "x \<in> set xs \<Longrightarrow> hidden (xs @ [x]) (last_index xs x)"
-by(auto simp add: hidden_def nth_append rev_nth[symmetric]
- dest: last_index_less[OF _ le_refl])
-
-lemma hidden_inacc: "hidden xs i \<Longrightarrow> last_index xs x \<noteq> i"
-by(auto simp add: hidden_def last_index_drop last_index_less_size_conv)
-
-
-lemma [simp]: "hidden xs i \<Longrightarrow> hidden (xs@[x]) i"
-by(auto simp add:hidden_def nth_append)
-
-
-lemma fun_upds_apply:
- "(m(xs[\<mapsto>]ys)) x =
- (let xs' = take (size ys) xs
- in if x \<in> set xs' then Some(ys ! last_index xs' x) else m x)"
-proof(induct xs arbitrary: m ys)
- case Nil then show ?case by(simp add: Let_def)
-next
- case Cons show ?case
- proof(cases ys)
- case Nil
- then show ?thesis by(simp add:Let_def)
- next
- case Cons': Cons
- then show ?thesis using Cons by(simp add: Let_def last_index_Cons)
- qed
-qed
-
-
-
-lemma map_upds_apply_eq_Some:
- "((m(xs[\<mapsto>]ys)) x = Some y) =
- (let xs' = take (size ys) xs
- in if x \<in> set xs' then ys ! last_index xs' x = y else m x = Some y)"
-by(simp add:fun_upds_apply Let_def)
-
-
-lemma map_upds_upd_conv_last_index:
- "\<lbrakk>x \<in> set xs; size xs \<le> size ys \<rbrakk>
- \<Longrightarrow> m(xs[\<mapsto>]ys)(x\<mapsto>y) = m(xs[\<mapsto>]ys[last_index xs x := y])"
-by(rule ext) (simp add:fun_upds_apply eq_sym_conv Let_def)
-
-end
+theory Hidden
+imports "List-Index.List_Index"
+begin
+
+definition hidden :: "'a list \<Rightarrow> nat \<Rightarrow> bool" where
+"hidden xs i \<equiv> i < size xs \<and> xs!i \<in> set(drop (i+1) xs)"
+
+
+lemma hidden_last_index: "x \<in> set xs \<Longrightarrow> hidden (xs @ [x]) (last_index xs x)"
+by(auto simp add: hidden_def nth_append rev_nth[symmetric]
+ dest: last_index_less[OF _ le_refl])
+
+lemma hidden_inacc: "hidden xs i \<Longrightarrow> last_index xs x \<noteq> i"
+by(auto simp add: hidden_def last_index_drop last_index_less_size_conv)
+
+
+lemma [simp]: "hidden xs i \<Longrightarrow> hidden (xs@[x]) i"
+by(auto simp add:hidden_def nth_append)
+
+
+lemma fun_upds_apply:
+ "(m(xs[\<mapsto>]ys)) x =
+ (let xs' = take (size ys) xs
+ in if x \<in> set xs' then Some(ys ! last_index xs' x) else m x)"
+proof(induct xs arbitrary: m ys)
+ case Nil then show ?case by(simp add: Let_def)
+next
+ case Cons show ?case
+ proof(cases ys)
+ case Nil
+ then show ?thesis by(simp add:Let_def)
+ next
+ case Cons': Cons
+ then show ?thesis using Cons by(simp add: Let_def last_index_Cons)
+ qed
+qed
+
+
+
+lemma map_upds_apply_eq_Some:
+ "((m(xs[\<mapsto>]ys)) x = Some y) =
+ (let xs' = take (size ys) xs
+ in if x \<in> set xs' then ys ! last_index xs' x = y else m x = Some y)"
+by(simp add:fun_upds_apply Let_def)
+
+
+lemma map_upds_upd_conv_last_index:
+ "\<lbrakk>x \<in> set xs; size xs \<le> size ys \<rbrakk>
+ \<Longrightarrow> m(xs[\<mapsto>]ys)(x\<mapsto>y) = m(xs[\<mapsto>]ys[last_index xs x := y])"
+by(rule ext) (simp add:fun_upds_apply eq_sym_conv Let_def)
+
+end
diff --git a/thys/JinjaDCI/Compiler/J1.thy b/thys/JinjaDCI/Compiler/J1.thy
--- a/thys/JinjaDCI/Compiler/J1.thy
+++ b/thys/JinjaDCI/Compiler/J1.thy
@@ -1,548 +1,548 @@
-(* Title: JinjaDCI/Compiler/J1.thy
- Author: Tobias Nipkow, Susannah Mansky
- Copyright 2003 Technische Universitaet Muenchen, 2019-20 UIUC
-
- Based on the Jinja theory Compiler/J1.thy by Tobias Nipkow
-*)
-
-chapter \<open> Compilation \label{cha:comp} \<close>
-
-section \<open> An Intermediate Language \<close>
-
-theory J1 imports "../J/BigStep" begin
-
-type_synonym expr\<^sub>1 = "nat exp"
-type_synonym J\<^sub>1_prog = "expr\<^sub>1 prog"
-type_synonym state\<^sub>1 = "heap \<times> (val list) \<times> sheap"
-
-definition hp\<^sub>1 :: "state\<^sub>1 \<Rightarrow> heap"
-where
- "hp\<^sub>1 \<equiv> fst"
-definition lcl\<^sub>1 :: "state\<^sub>1 \<Rightarrow> val list"
-where
- "lcl\<^sub>1 \<equiv> fst \<circ> snd"
-definition shp\<^sub>1 :: "state\<^sub>1 \<Rightarrow> sheap"
-where
- "shp\<^sub>1 \<equiv> snd \<circ> snd"
-
-(*<*)
-declare hp\<^sub>1_def[simp] lcl\<^sub>1_def[simp] shp\<^sub>1_def[simp]
-(*>*)
-
-primrec
- max_vars :: "'a exp \<Rightarrow> nat"
- and max_varss :: "'a exp list \<Rightarrow> nat"
-where
- "max_vars(new C) = 0"
-| "max_vars(Cast C e) = max_vars e"
-| "max_vars(Val v) = 0"
-| "max_vars(e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2) = max (max_vars e\<^sub>1) (max_vars e\<^sub>2)"
-| "max_vars(Var V) = 0"
-| "max_vars(V:=e) = max_vars e"
-| "max_vars(e\<bullet>F{D}) = max_vars e"
-| "max_vars(C\<bullet>\<^sub>sF{D}) = 0"
-| "max_vars(FAss e\<^sub>1 F D e\<^sub>2) = max (max_vars e\<^sub>1) (max_vars e\<^sub>2)"
-| "max_vars(SFAss C F D e\<^sub>2) = max_vars e\<^sub>2"
-| "max_vars(e\<bullet>M(es)) = max (max_vars e) (max_varss es)"
-| "max_vars(C\<bullet>\<^sub>sM(es)) = max_varss es"
-| "max_vars({V:T; e}) = max_vars e + 1"
-| "max_vars(e\<^sub>1;;e\<^sub>2) = max (max_vars e\<^sub>1) (max_vars e\<^sub>2)"
-| "max_vars(if (e) e\<^sub>1 else e\<^sub>2) =
- max (max_vars e) (max (max_vars e\<^sub>1) (max_vars e\<^sub>2))"
-| "max_vars(while (b) e) = max (max_vars b) (max_vars e)"
-| "max_vars(throw e) = max_vars e"
-| "max_vars(try e\<^sub>1 catch(C V) e\<^sub>2) = max (max_vars e\<^sub>1) (max_vars e\<^sub>2 + 1)"
-| "max_vars(INIT C (Cs,b) \<leftarrow> e) = max_vars e"
-| "max_vars(RI(C,e);Cs \<leftarrow> e') = max (max_vars e) (max_vars e')"
-
-| "max_varss [] = 0"
-| "max_varss (e#es) = max (max_vars e) (max_varss es)"
-
-inductive
- eval\<^sub>1 :: "J\<^sub>1_prog \<Rightarrow> expr\<^sub>1 \<Rightarrow> state\<^sub>1 \<Rightarrow> expr\<^sub>1 \<Rightarrow> state\<^sub>1 \<Rightarrow> bool"
- ("_ \<turnstile>\<^sub>1 ((1\<langle>_,/_\<rangle>) \<Rightarrow>/ (1\<langle>_,/_\<rangle>))" [51,0,0,0,0] 81)
- and evals\<^sub>1 :: "J\<^sub>1_prog \<Rightarrow> expr\<^sub>1 list \<Rightarrow> state\<^sub>1 \<Rightarrow> expr\<^sub>1 list \<Rightarrow> state\<^sub>1 \<Rightarrow> bool"
- ("_ \<turnstile>\<^sub>1 ((1\<langle>_,/_\<rangle>) [\<Rightarrow>]/ (1\<langle>_,/_\<rangle>))" [51,0,0,0,0] 81)
- for P :: J\<^sub>1_prog
-where
-
- New\<^sub>1:
- "\<lbrakk> sh C = Some (sfs, Done); new_Addr h = Some a;
- P \<turnstile> C has_fields FDTs; h' = h(a\<mapsto>blank P C) \<rbrakk>
- \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>new C,(h,l,sh)\<rangle> \<Rightarrow> \<langle>addr a,(h',l,sh)\<rangle>"
-| NewFail\<^sub>1:
- "\<lbrakk> sh C = Some (sfs, Done); new_Addr h = None \<rbrakk> \<Longrightarrow>
- P \<turnstile>\<^sub>1 \<langle>new C, (h,l,sh)\<rangle> \<Rightarrow> \<langle>THROW OutOfMemory,(h,l,sh)\<rangle>"
-| NewInit\<^sub>1:
- "\<lbrakk> \<nexists>sfs. sh C = Some (sfs, Done); P \<turnstile>\<^sub>1 \<langle>INIT C ([C],False) \<leftarrow> unit,(h,l,sh)\<rangle> \<Rightarrow> \<langle>Val v',(h',l',sh')\<rangle>;
- new_Addr h' = Some a; P \<turnstile> C has_fields FDTs; h'' = h'(a\<mapsto>blank P C) \<rbrakk>
- \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>new C,(h,l,sh)\<rangle> \<Rightarrow> \<langle>addr a,(h'',l',sh')\<rangle>"
-| NewInitOOM\<^sub>1:
- "\<lbrakk> \<nexists>sfs. sh C = Some (sfs, Done); P \<turnstile>\<^sub>1 \<langle>INIT C ([C],False) \<leftarrow> unit,(h,l,sh)\<rangle> \<Rightarrow> \<langle>Val v',(h',l',sh')\<rangle>;
- new_Addr h' = None; is_class P C \<rbrakk>
- \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>new C,(h,l,sh)\<rangle> \<Rightarrow> \<langle>THROW OutOfMemory,(h',l',sh')\<rangle>"
-| NewInitThrow\<^sub>1:
- "\<lbrakk> \<nexists>sfs. sh C = Some (sfs, Done); P \<turnstile>\<^sub>1 \<langle>INIT C ([C],False) \<leftarrow> unit,(h,l,sh)\<rangle> \<Rightarrow> \<langle>throw a,s'\<rangle>;
- is_class P C \<rbrakk>
- \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>new C,(h,l,sh)\<rangle> \<Rightarrow> \<langle>throw a,s'\<rangle>"
-
-| Cast\<^sub>1:
- "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>addr a,(h,l,sh)\<rangle>; h a = Some(D,fs); P \<turnstile> D \<preceq>\<^sup>* C \<rbrakk>
- \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>Cast C e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>addr a,(h,l,sh)\<rangle>"
-| CastNull\<^sub>1:
- "P \<turnstile>\<^sub>1 \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>null,s\<^sub>1\<rangle> \<Longrightarrow>
- P \<turnstile>\<^sub>1 \<langle>Cast C e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>null,s\<^sub>1\<rangle>"
-| CastFail\<^sub>1:
- "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>addr a,(h,l,sh)\<rangle>; h a = Some(D,fs); \<not> P \<turnstile> D \<preceq>\<^sup>* C \<rbrakk>
- \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>Cast C e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>THROW ClassCast,(h,l,sh)\<rangle>"
-| CastThrow\<^sub>1:
- "P \<turnstile>\<^sub>1 \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle> \<Longrightarrow>
- P \<turnstile>\<^sub>1 \<langle>Cast C e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle>"
-
-| Val\<^sub>1:
- "P \<turnstile>\<^sub>1 \<langle>Val v,s\<rangle> \<Rightarrow> \<langle>Val v,s\<rangle>"
-
-| BinOp\<^sub>1:
- "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Val v\<^sub>1,s\<^sub>1\<rangle>; P \<turnstile>\<^sub>1 \<langle>e\<^sub>2,s\<^sub>1\<rangle> \<Rightarrow> \<langle>Val v\<^sub>2,s\<^sub>2\<rangle>; binop(bop,v\<^sub>1,v\<^sub>2) = Some v \<rbrakk>
- \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Val v,s\<^sub>2\<rangle>"
-| BinOpThrow\<^sub>1\<^sub>1:
- "P \<turnstile>\<^sub>1 \<langle>e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e,s\<^sub>1\<rangle> \<Longrightarrow>
- P \<turnstile>\<^sub>1 \<langle>e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2, s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e,s\<^sub>1\<rangle>"
-| BinOpThrow\<^sub>2\<^sub>1:
- "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Val v\<^sub>1,s\<^sub>1\<rangle>; P \<turnstile>\<^sub>1 \<langle>e\<^sub>2,s\<^sub>1\<rangle> \<Rightarrow> \<langle>throw e,s\<^sub>2\<rangle> \<rbrakk>
- \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e,s\<^sub>2\<rangle>"
-
-| Var\<^sub>1:
- "\<lbrakk> ls!i = v; i < size ls \<rbrakk> \<Longrightarrow>
- P \<turnstile>\<^sub>1 \<langle>Var i,(h,ls,sh)\<rangle> \<Rightarrow> \<langle>Val v,(h,ls,sh)\<rangle>"
-
-| LAss\<^sub>1:
- "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Val v,(h,ls,sh)\<rangle>; i < size ls; ls' = ls[i := v] \<rbrakk>
- \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>i:= e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>unit,(h,ls',sh)\<rangle>"
-| LAssThrow\<^sub>1:
- "P \<turnstile>\<^sub>1 \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle> \<Longrightarrow>
- P \<turnstile>\<^sub>1 \<langle>i:= e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle>"
-
-| FAcc\<^sub>1:
- "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>addr a,(h,ls,sh)\<rangle>; h a = Some(C,fs);
- P \<turnstile> C has F,NonStatic:t in D;
- fs(F,D) = Some v \<rbrakk>
- \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>e\<bullet>F{D},s\<^sub>0\<rangle> \<Rightarrow> \<langle>Val v,(h,ls,sh)\<rangle>"
-| FAccNull\<^sub>1:
- "P \<turnstile>\<^sub>1 \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>null,s\<^sub>1\<rangle> \<Longrightarrow>
- P \<turnstile>\<^sub>1 \<langle>e\<bullet>F{D},s\<^sub>0\<rangle> \<Rightarrow> \<langle>THROW NullPointer,s\<^sub>1\<rangle>"
-| FAccThrow\<^sub>1:
- "P \<turnstile>\<^sub>1 \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle> \<Longrightarrow>
- P \<turnstile>\<^sub>1 \<langle>e\<bullet>F{D},s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle>"
-| FAccNone\<^sub>1:
- "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>addr a,(h,ls,sh)\<rangle>; h a = Some(C,fs);
- \<not>(\<exists>b t. P \<turnstile> C has F,b:t in D) \<rbrakk>
- \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>e\<bullet>F{D},s\<^sub>0\<rangle> \<Rightarrow> \<langle>THROW NoSuchFieldError,(h,ls,sh)\<rangle>"
-| FAccStatic\<^sub>1:
- "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>addr a,(h,ls,sh)\<rangle>; h a = Some(C,fs);
- P \<turnstile> C has F,Static:t in D \<rbrakk>
- \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>e\<bullet>F{D},s\<^sub>0\<rangle> \<Rightarrow> \<langle>THROW IncompatibleClassChangeError,(h,ls,sh)\<rangle>"
-
-| SFAcc\<^sub>1:
- "\<lbrakk> P \<turnstile> C has F,Static:t in D;
- sh D = Some (sfs,Done);
- sfs F = Some v \<rbrakk>
- \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>C\<bullet>\<^sub>sF{D},(h,ls,sh)\<rangle> \<Rightarrow> \<langle>Val v,(h,ls,sh)\<rangle>"
-| SFAccInit\<^sub>1:
- "\<lbrakk> P \<turnstile> C has F,Static:t in D;
- \<nexists>sfs. sh D = Some (sfs,Done); P \<turnstile>\<^sub>1 \<langle>INIT D ([D],False) \<leftarrow> unit,(h,ls,sh)\<rangle> \<Rightarrow> \<langle>Val v',(h',ls',sh')\<rangle>;
- sh' D = Some (sfs,i);
- sfs F = Some v \<rbrakk>
- \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>C\<bullet>\<^sub>sF{D},(h,ls,sh)\<rangle> \<Rightarrow> \<langle>Val v,(h',ls',sh')\<rangle>"
-| SFAccInitThrow\<^sub>1:
- "\<lbrakk> P \<turnstile> C has F,Static:t in D;
- \<nexists>sfs. sh D = Some (sfs,Done); P \<turnstile>\<^sub>1 \<langle>INIT D ([D],False) \<leftarrow> unit,(h,ls,sh)\<rangle> \<Rightarrow> \<langle>throw a,s'\<rangle> \<rbrakk>
- \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>C\<bullet>\<^sub>sF{D},(h,ls,sh)\<rangle> \<Rightarrow> \<langle>throw a,s'\<rangle>"
-| SFAccNone\<^sub>1:
- "\<lbrakk> \<not>(\<exists>b t. P \<turnstile> C has F,b:t in D) \<rbrakk>
- \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>C\<bullet>\<^sub>sF{D},s\<rangle> \<Rightarrow> \<langle>THROW NoSuchFieldError,s\<rangle>"
-| SFAccNonStatic\<^sub>1:
- "\<lbrakk> P \<turnstile> C has F,NonStatic:t in D \<rbrakk>
- \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>C\<bullet>\<^sub>sF{D},s\<rangle> \<Rightarrow> \<langle>THROW IncompatibleClassChangeError,s\<rangle>"
-
-
-| FAss\<^sub>1:
- "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow> \<langle>addr a,s\<^sub>1\<rangle>; P \<turnstile>\<^sub>1 \<langle>e\<^sub>2,s\<^sub>1\<rangle> \<Rightarrow> \<langle>Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\<rangle>;
- h\<^sub>2 a = Some(C,fs); P \<turnstile> C has F,NonStatic:t in D;
- fs' = fs((F,D)\<mapsto>v); h\<^sub>2' = h\<^sub>2(a\<mapsto>(C,fs')) \<rbrakk>
- \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>e\<^sub>1\<bullet>F{D}:=e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>unit,(h\<^sub>2',l\<^sub>2,sh\<^sub>2)\<rangle>"
-| FAssNull\<^sub>1:
- "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow> \<langle>null,s\<^sub>1\<rangle>; P \<turnstile>\<^sub>1 \<langle>e\<^sub>2,s\<^sub>1\<rangle> \<Rightarrow> \<langle>Val v,s\<^sub>2\<rangle> \<rbrakk> \<Longrightarrow>
- P \<turnstile>\<^sub>1 \<langle>e\<^sub>1\<bullet>F{D}:=e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>THROW NullPointer,s\<^sub>2\<rangle>"
-| FAssThrow\<^sub>1\<^sub>1:
- "P \<turnstile>\<^sub>1 \<langle>e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle> \<Longrightarrow>
- P \<turnstile>\<^sub>1 \<langle>e\<^sub>1\<bullet>F{D}:=e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle>"
-| FAssThrow\<^sub>2\<^sub>1:
- "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Val v,s\<^sub>1\<rangle>; P \<turnstile>\<^sub>1 \<langle>e\<^sub>2,s\<^sub>1\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>2\<rangle> \<rbrakk>
- \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>e\<^sub>1\<bullet>F{D}:=e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>2\<rangle>"
-| FAssNone\<^sub>1:
- "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow> \<langle>addr a,s\<^sub>1\<rangle>; P \<turnstile>\<^sub>1 \<langle>e\<^sub>2,s\<^sub>1\<rangle> \<Rightarrow> \<langle>Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\<rangle>;
- h\<^sub>2 a = Some(C,fs); \<not>(\<exists>b t. P \<turnstile> C has F,b:t in D) \<rbrakk>
- \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>e\<^sub>1\<bullet>F{D}:=e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>THROW NoSuchFieldError,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\<rangle>"
-| FAssStatic\<^sub>1:
- "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow> \<langle>addr a,s\<^sub>1\<rangle>; P \<turnstile>\<^sub>1 \<langle>e\<^sub>2,s\<^sub>1\<rangle> \<Rightarrow> \<langle>Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\<rangle>;
- h\<^sub>2 a = Some(C,fs); P \<turnstile> C has F,Static:t in D \<rbrakk>
- \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>e\<^sub>1\<bullet>F{D}:=e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>THROW IncompatibleClassChangeError,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\<rangle>"
-
-| SFAss\<^sub>1:
- "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Val v,(h\<^sub>1,l\<^sub>1,sh\<^sub>1)\<rangle>;
- P \<turnstile> C has F,Static:t in D;
- sh\<^sub>1 D = Some(sfs,Done); sfs' = sfs(F\<mapsto>v); sh\<^sub>1' = sh\<^sub>1(D\<mapsto>(sfs',Done)) \<rbrakk>
- \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>C\<bullet>\<^sub>sF{D}:=e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>unit,(h\<^sub>1,l\<^sub>1,sh\<^sub>1')\<rangle>"
-| SFAssInit\<^sub>1:
- "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Val v,(h\<^sub>1,l\<^sub>1,sh\<^sub>1)\<rangle>;
- P \<turnstile> C has F,Static:t in D;
- \<nexists>sfs. sh\<^sub>1 D = Some(sfs,Done); P \<turnstile>\<^sub>1 \<langle>INIT D ([D],False) \<leftarrow> unit,(h\<^sub>1,l\<^sub>1,sh\<^sub>1)\<rangle> \<Rightarrow> \<langle>Val v',(h',l',sh')\<rangle>;
- sh' D = Some(sfs,i);
- sfs' = sfs(F\<mapsto>v); sh'' = sh'(D\<mapsto>(sfs',i)) \<rbrakk>
- \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>C\<bullet>\<^sub>sF{D}:=e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>unit,(h',l',sh'')\<rangle>"
-| SFAssInitThrow\<^sub>1:
- "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Val v,(h\<^sub>1,l\<^sub>1,sh\<^sub>1)\<rangle>;
- P \<turnstile> C has F,Static:t in D;
- \<nexists>sfs. sh\<^sub>1 D = Some(sfs,Done); P \<turnstile>\<^sub>1 \<langle>INIT D ([D],False) \<leftarrow> unit,(h\<^sub>1,l\<^sub>1,sh\<^sub>1)\<rangle> \<Rightarrow> \<langle>throw a,s'\<rangle> \<rbrakk>
- \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>C\<bullet>\<^sub>sF{D}:=e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw a,s'\<rangle>"
-| SFAssThrow\<^sub>1:
- "P \<turnstile>\<^sub>1 \<langle>e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>2\<rangle>
- \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>C\<bullet>\<^sub>sF{D}:=e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>2\<rangle>"
-| SFAssNone\<^sub>1:
- "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\<rangle>;
- \<not>(\<exists>b t. P \<turnstile> C has F,b:t in D) \<rbrakk>
- \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>C\<bullet>\<^sub>sF{D}:=e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>THROW NoSuchFieldError,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\<rangle>"
-| SFAssNonStatic\<^sub>1:
- "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\<rangle>;
- P \<turnstile> C has F,NonStatic:t in D \<rbrakk>
- \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>C\<bullet>\<^sub>sF{D}:=e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>THROW IncompatibleClassChangeError,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\<rangle>"
-
-| CallObjThrow\<^sub>1:
- "P \<turnstile>\<^sub>1 \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle> \<Longrightarrow>
- P \<turnstile>\<^sub>1 \<langle>e\<bullet>M(es),s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle>"
-| CallNull\<^sub>1:
- "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>null,s\<^sub>1\<rangle>; P \<turnstile>\<^sub>1 \<langle>es,s\<^sub>1\<rangle> [\<Rightarrow>] \<langle>map Val vs,s\<^sub>2\<rangle> \<rbrakk>
- \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>e\<bullet>M(es),s\<^sub>0\<rangle> \<Rightarrow> \<langle>THROW NullPointer,s\<^sub>2\<rangle>"
-| Call\<^sub>1:
- "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>addr a,s\<^sub>1\<rangle>; P \<turnstile>\<^sub>1 \<langle>es,s\<^sub>1\<rangle> [\<Rightarrow>] \<langle>map Val vs,(h\<^sub>2,ls\<^sub>2,sh\<^sub>2)\<rangle>;
- h\<^sub>2 a = Some(C,fs); P \<turnstile> C sees M,NonStatic:Ts\<rightarrow>T = body in D;
- size vs = size Ts; ls\<^sub>2' = (Addr a) # vs @ replicate (max_vars body) undefined;
- P \<turnstile>\<^sub>1 \<langle>body,(h\<^sub>2,ls\<^sub>2',sh\<^sub>2)\<rangle> \<Rightarrow> \<langle>e',(h\<^sub>3,ls\<^sub>3,sh\<^sub>3)\<rangle> \<rbrakk>
- \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>e\<bullet>M(es),s\<^sub>0\<rangle> \<Rightarrow> \<langle>e',(h\<^sub>3,ls\<^sub>2,sh\<^sub>3)\<rangle>"
-| CallParamsThrow\<^sub>1:
- "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Val v,s\<^sub>1\<rangle>; P \<turnstile>\<^sub>1 \<langle>es,s\<^sub>1\<rangle> [\<Rightarrow>] \<langle>es',s\<^sub>2\<rangle>;
- es' = map Val vs @ throw ex # es\<^sub>2 \<rbrakk>
- \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>e\<bullet>M(es),s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw ex,s\<^sub>2\<rangle>"
-| CallNone\<^sub>1:
- "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>addr a,s\<^sub>1\<rangle>; P \<turnstile>\<^sub>1 \<langle>ps,s\<^sub>1\<rangle> [\<Rightarrow>] \<langle>map Val vs,(h\<^sub>2,ls\<^sub>2,sh\<^sub>2)\<rangle>;
- h\<^sub>2 a = Some(C,fs); \<not>(\<exists>b Ts T body D. P \<turnstile> C sees M,b:Ts\<rightarrow>T = body in D) \<rbrakk>
- \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>e\<bullet>M(ps),s\<^sub>0\<rangle> \<Rightarrow> \<langle>THROW NoSuchMethodError,(h\<^sub>2,ls\<^sub>2,sh\<^sub>2)\<rangle>"
-| CallStatic\<^sub>1:
- "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>addr a,s\<^sub>1\<rangle>; P \<turnstile>\<^sub>1 \<langle>ps,s\<^sub>1\<rangle> [\<Rightarrow>] \<langle>map Val vs,(h\<^sub>2,ls\<^sub>2,sh\<^sub>2)\<rangle>;
- h\<^sub>2 a = Some(C,fs); P \<turnstile> C sees M,Static:Ts\<rightarrow>T = body in D \<rbrakk>
- \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>e\<bullet>M(ps),s\<^sub>0\<rangle> \<Rightarrow> \<langle>THROW IncompatibleClassChangeError,(h\<^sub>2,ls\<^sub>2,sh\<^sub>2)\<rangle>"
-
-| SCallParamsThrow\<^sub>1:
- "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>es,s\<^sub>0\<rangle> [\<Rightarrow>] \<langle>es',s\<^sub>2\<rangle>; es' = map Val vs @ throw ex # es\<^sub>2 \<rbrakk>
- \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>C\<bullet>\<^sub>sM(es),s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw ex,s\<^sub>2\<rangle>"
-| SCallNone\<^sub>1:
- "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>ps,s\<^sub>0\<rangle> [\<Rightarrow>] \<langle>map Val vs,s\<^sub>2\<rangle>;
- \<not>(\<exists>b Ts T body D. P \<turnstile> C sees M,b:Ts\<rightarrow>T = body in D) \<rbrakk>
- \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>C\<bullet>\<^sub>sM(ps),s\<^sub>0\<rangle> \<Rightarrow> \<langle>THROW NoSuchMethodError,s\<^sub>2\<rangle>"
-| SCallNonStatic\<^sub>1:
- "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>ps,s\<^sub>0\<rangle> [\<Rightarrow>] \<langle>map Val vs,s\<^sub>2\<rangle>;
- P \<turnstile> C sees M,NonStatic:Ts\<rightarrow>T = body in D \<rbrakk>
- \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>C\<bullet>\<^sub>sM(ps),s\<^sub>0\<rangle> \<Rightarrow> \<langle>THROW IncompatibleClassChangeError,s\<^sub>2\<rangle>"
-| SCallInitThrow\<^sub>1:
- "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>ps,s\<^sub>0\<rangle> [\<Rightarrow>] \<langle>map Val vs,(h\<^sub>1,ls\<^sub>1,sh\<^sub>1)\<rangle>;
- P \<turnstile> C sees M,Static:Ts\<rightarrow>T = body in D;
- \<nexists>sfs. sh\<^sub>1 D = Some(sfs,Done); M \<noteq> clinit;
- P \<turnstile>\<^sub>1 \<langle>INIT D ([D],False) \<leftarrow> unit,(h\<^sub>1,ls\<^sub>1,sh\<^sub>1)\<rangle> \<Rightarrow> \<langle>throw a,s'\<rangle> \<rbrakk>
- \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>C\<bullet>\<^sub>sM(ps),s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw a,s'\<rangle>"
-| SCallInit\<^sub>1:
- "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>ps,s\<^sub>0\<rangle> [\<Rightarrow>] \<langle>map Val vs,(h\<^sub>1,ls\<^sub>1,sh\<^sub>1)\<rangle>;
- P \<turnstile> C sees M,Static:Ts\<rightarrow>T = body in D;
- \<nexists>sfs. sh\<^sub>1 D = Some(sfs,Done); M \<noteq> clinit;
- P \<turnstile>\<^sub>1 \<langle>INIT D ([D],False) \<leftarrow> unit,(h\<^sub>1,ls\<^sub>1,sh\<^sub>1)\<rangle> \<Rightarrow> \<langle>Val v',(h\<^sub>2,ls\<^sub>2,sh\<^sub>2)\<rangle>;
- size vs = size Ts; ls\<^sub>2' = vs @ replicate (max_vars body) undefined;
- P \<turnstile>\<^sub>1 \<langle>body,(h\<^sub>2,ls\<^sub>2',sh\<^sub>2)\<rangle> \<Rightarrow> \<langle>e',(h\<^sub>3,ls\<^sub>3,sh\<^sub>3)\<rangle> \<rbrakk>
- \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>C\<bullet>\<^sub>sM(ps),s\<^sub>0\<rangle> \<Rightarrow> \<langle>e',(h\<^sub>3,ls\<^sub>2,sh\<^sub>3)\<rangle>"
-| SCall\<^sub>1:
- "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>ps,s\<^sub>0\<rangle> [\<Rightarrow>] \<langle>map Val vs,(h\<^sub>2,ls\<^sub>2,sh\<^sub>2)\<rangle>;
- P \<turnstile> C sees M,Static:Ts\<rightarrow>T = body in D;
- sh\<^sub>2 D = Some(sfs,Done) \<or> (M = clinit \<and> sh\<^sub>2 D = \<lfloor>(sfs, Processing)\<rfloor>);
- size vs = size Ts; ls\<^sub>2' = vs @ replicate (max_vars body) undefined;
- P \<turnstile>\<^sub>1 \<langle>body,(h\<^sub>2,ls\<^sub>2',sh\<^sub>2)\<rangle> \<Rightarrow> \<langle>e',(h\<^sub>3,ls\<^sub>3,sh\<^sub>3)\<rangle> \<rbrakk>
- \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>C\<bullet>\<^sub>sM(ps),s\<^sub>0\<rangle> \<Rightarrow> \<langle>e',(h\<^sub>3,ls\<^sub>2,sh\<^sub>3)\<rangle>"
-
-| Block\<^sub>1:
- "P \<turnstile>\<^sub>1 \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>e',s\<^sub>1\<rangle> \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>Block i T e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>e',s\<^sub>1\<rangle>"
-
-| Seq\<^sub>1:
- "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e\<^sub>0,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Val v,s\<^sub>1\<rangle>; P \<turnstile>\<^sub>1 \<langle>e\<^sub>1,s\<^sub>1\<rangle> \<Rightarrow> \<langle>e\<^sub>2,s\<^sub>2\<rangle> \<rbrakk>
- \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>e\<^sub>0;;e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow> \<langle>e\<^sub>2,s\<^sub>2\<rangle>"
-| SeqThrow\<^sub>1:
- "P \<turnstile>\<^sub>1 \<langle>e\<^sub>0,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e,s\<^sub>1\<rangle> \<Longrightarrow>
- P \<turnstile>\<^sub>1 \<langle>e\<^sub>0;;e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e,s\<^sub>1\<rangle>"
-
-| CondT\<^sub>1:
- "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>true,s\<^sub>1\<rangle>; P \<turnstile>\<^sub>1 \<langle>e\<^sub>1,s\<^sub>1\<rangle> \<Rightarrow> \<langle>e',s\<^sub>2\<rangle> \<rbrakk>
- \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>if (e) e\<^sub>1 else e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>e',s\<^sub>2\<rangle>"
-| CondF\<^sub>1:
- "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>false,s\<^sub>1\<rangle>; P \<turnstile>\<^sub>1 \<langle>e\<^sub>2,s\<^sub>1\<rangle> \<Rightarrow> \<langle>e',s\<^sub>2\<rangle> \<rbrakk>
- \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>if (e) e\<^sub>1 else e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>e',s\<^sub>2\<rangle>"
-| CondThrow\<^sub>1:
- "P \<turnstile>\<^sub>1 \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle> \<Longrightarrow>
- P \<turnstile>\<^sub>1 \<langle>if (e) e\<^sub>1 else e\<^sub>2, s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle>"
-
-| WhileF\<^sub>1:
- "P \<turnstile>\<^sub>1 \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>false,s\<^sub>1\<rangle> \<Longrightarrow>
- P \<turnstile>\<^sub>1 \<langle>while (e) c,s\<^sub>0\<rangle> \<Rightarrow> \<langle>unit,s\<^sub>1\<rangle>"
-| WhileT\<^sub>1:
- "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>true,s\<^sub>1\<rangle>; P \<turnstile>\<^sub>1 \<langle>c,s\<^sub>1\<rangle> \<Rightarrow> \<langle>Val v\<^sub>1,s\<^sub>2\<rangle>;
- P \<turnstile>\<^sub>1 \<langle>while (e) c,s\<^sub>2\<rangle> \<Rightarrow> \<langle>e\<^sub>3,s\<^sub>3\<rangle> \<rbrakk>
- \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>while (e) c,s\<^sub>0\<rangle> \<Rightarrow> \<langle>e\<^sub>3,s\<^sub>3\<rangle>"
-| WhileCondThrow\<^sub>1:
- "P \<turnstile>\<^sub>1 \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle> \<Longrightarrow>
- P \<turnstile>\<^sub>1 \<langle>while (e) c,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle>"
-| WhileBodyThrow\<^sub>1:
- "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>true,s\<^sub>1\<rangle>; P \<turnstile>\<^sub>1 \<langle>c,s\<^sub>1\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>2\<rangle>\<rbrakk>
- \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>while (e) c,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>2\<rangle>"
-
-| Throw\<^sub>1:
- "P \<turnstile>\<^sub>1 \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>addr a,s\<^sub>1\<rangle> \<Longrightarrow>
- P \<turnstile>\<^sub>1 \<langle>throw e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Throw a,s\<^sub>1\<rangle>"
-| ThrowNull\<^sub>1:
- "P \<turnstile>\<^sub>1 \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>null,s\<^sub>1\<rangle> \<Longrightarrow>
- P \<turnstile>\<^sub>1 \<langle>throw e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>THROW NullPointer,s\<^sub>1\<rangle>"
-| ThrowThrow\<^sub>1:
- "P \<turnstile>\<^sub>1 \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle> \<Longrightarrow>
- P \<turnstile>\<^sub>1 \<langle>throw e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle>"
-
-| Try\<^sub>1:
- "P \<turnstile>\<^sub>1 \<langle>e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Val v\<^sub>1,s\<^sub>1\<rangle> \<Longrightarrow>
- P \<turnstile>\<^sub>1 \<langle>try e\<^sub>1 catch(C i) e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Val v\<^sub>1,s\<^sub>1\<rangle>"
-| TryCatch\<^sub>1:
- "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Throw a,(h\<^sub>1,ls\<^sub>1,sh\<^sub>1)\<rangle>;
- h\<^sub>1 a = Some(D,fs); P \<turnstile> D \<preceq>\<^sup>* C; i < length ls\<^sub>1;
- P \<turnstile>\<^sub>1 \<langle>e\<^sub>2,(h\<^sub>1,ls\<^sub>1[i:=Addr a],sh\<^sub>1)\<rangle> \<Rightarrow> \<langle>e\<^sub>2',(h\<^sub>2,ls\<^sub>2,sh\<^sub>2)\<rangle> \<rbrakk>
- \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>try e\<^sub>1 catch(C i) e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>e\<^sub>2',(h\<^sub>2,ls\<^sub>2,sh\<^sub>2)\<rangle>"
-| TryThrow\<^sub>1:
- "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Throw a,(h\<^sub>1,ls\<^sub>1,sh\<^sub>1)\<rangle>; h\<^sub>1 a = Some(D,fs); \<not> P \<turnstile> D \<preceq>\<^sup>* C \<rbrakk>
- \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>try e\<^sub>1 catch(C i) e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Throw a,(h\<^sub>1,ls\<^sub>1,sh\<^sub>1)\<rangle>"
-
-| Nil\<^sub>1:
- "P \<turnstile>\<^sub>1 \<langle>[],s\<rangle> [\<Rightarrow>] \<langle>[],s\<rangle>"
-
-| Cons\<^sub>1:
- "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Val v,s\<^sub>1\<rangle>; P \<turnstile>\<^sub>1 \<langle>es,s\<^sub>1\<rangle> [\<Rightarrow>] \<langle>es',s\<^sub>2\<rangle> \<rbrakk>
- \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>e#es,s\<^sub>0\<rangle> [\<Rightarrow>] \<langle>Val v # es',s\<^sub>2\<rangle>"
-| ConsThrow\<^sub>1:
- "P \<turnstile>\<^sub>1 \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle> \<Longrightarrow>
- P \<turnstile>\<^sub>1 \<langle>e#es,s\<^sub>0\<rangle> [\<Rightarrow>] \<langle>throw e' # es, s\<^sub>1\<rangle>"
-
-\<comment> \<open> init rules \<close>
-
-| InitFinal\<^sub>1:
- "P \<turnstile>\<^sub>1 \<langle>e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle> \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>INIT C (Nil,b) \<leftarrow> e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
-| InitNone\<^sub>1:
- "\<lbrakk> sh C = None; P \<turnstile>\<^sub>1 \<langle>INIT C' (C#Cs,False) \<leftarrow> e,(h,l,sh(C \<mapsto> (sblank P C, Prepared)))\<rangle> \<Rightarrow> \<langle>e',s'\<rangle> \<rbrakk>
- \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>INIT C' (C#Cs,False) \<leftarrow> e,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
-| InitDone\<^sub>1:
- "\<lbrakk> sh C = Some(sfs,Done); P \<turnstile>\<^sub>1 \<langle>INIT C' (Cs,True) \<leftarrow> e,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle> \<rbrakk>
- \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>INIT C' (C#Cs,False) \<leftarrow> e,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
-| InitProcessing\<^sub>1:
- "\<lbrakk> sh C = Some(sfs,Processing); P \<turnstile>\<^sub>1 \<langle>INIT C' (Cs,True) \<leftarrow> e,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle> \<rbrakk>
- \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>INIT C' (C#Cs,False) \<leftarrow> e,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
-| InitError\<^sub>1:
- "\<lbrakk> sh C = Some(sfs,Error);
- P \<turnstile>\<^sub>1 \<langle>RI (C, THROW NoClassDefFoundError);Cs \<leftarrow> e,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle> \<rbrakk>
- \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>INIT C' (C#Cs,False) \<leftarrow> e,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
-| InitObject\<^sub>1:
- "\<lbrakk> sh C = Some(sfs,Prepared);
- C = Object;
- sh' = sh(C \<mapsto> (sfs,Processing));
- P \<turnstile>\<^sub>1 \<langle>INIT C' (C#Cs,True) \<leftarrow> e,(h,l,sh')\<rangle> \<Rightarrow> \<langle>e',s'\<rangle> \<rbrakk>
- \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>INIT C' (C#Cs,False) \<leftarrow> e,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
-| InitNonObject\<^sub>1:
- "\<lbrakk> sh C = Some(sfs,Prepared);
- C \<noteq> Object;
- class P C = Some (D,r);
- sh' = sh(C \<mapsto> (sfs,Processing));
- P \<turnstile>\<^sub>1 \<langle>INIT C' (D#C#Cs,False) \<leftarrow> e,(h,l,sh')\<rangle> \<Rightarrow> \<langle>e',s'\<rangle> \<rbrakk>
- \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>INIT C' (C#Cs,False) \<leftarrow> e,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
-| InitRInit\<^sub>1:
- "P \<turnstile>\<^sub>1 \<langle>RI (C,C\<bullet>\<^sub>sclinit([]));Cs \<leftarrow> e,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>
- \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>INIT C' (C#Cs,True) \<leftarrow> e,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
-
-| RInit\<^sub>1:
- "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e,s\<rangle> \<Rightarrow> \<langle>Val v, (h',l',sh')\<rangle>;
- sh' C = Some(sfs, i); sh'' = sh'(C \<mapsto> (sfs, Done));
- C' = last(C#Cs);
- P \<turnstile>\<^sub>1 \<langle>INIT C' (Cs,True) \<leftarrow> e', (h',l',sh'')\<rangle> \<Rightarrow> \<langle>e\<^sub>1,s\<^sub>1\<rangle> \<rbrakk>
- \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>RI (C,e);Cs \<leftarrow> e',s\<rangle> \<Rightarrow> \<langle>e\<^sub>1,s\<^sub>1\<rangle>"
-| RInitInitFail\<^sub>1:
- "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e,s\<rangle> \<Rightarrow> \<langle>throw a, (h',l',sh')\<rangle>;
- sh' C = Some(sfs, i); sh'' = sh'(C \<mapsto> (sfs, Error));
- P \<turnstile>\<^sub>1 \<langle>RI (D,throw a);Cs \<leftarrow> e', (h',l',sh'')\<rangle> \<Rightarrow> \<langle>e\<^sub>1,s\<^sub>1\<rangle> \<rbrakk>
- \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>RI (C,e);D#Cs \<leftarrow> e',s\<rangle> \<Rightarrow> \<langle>e\<^sub>1,s\<^sub>1\<rangle>"
-| RInitFailFinal\<^sub>1:
- "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e,s\<rangle> \<Rightarrow> \<langle>throw a, (h',l',sh')\<rangle>;
- sh' C = Some(sfs, i); sh'' = sh'(C \<mapsto> (sfs, Error)) \<rbrakk>
- \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>RI (C,e);Nil \<leftarrow> e',s\<rangle> \<Rightarrow> \<langle>throw a, (h',l',sh'')\<rangle>"
-
-
-(*<*)
-lemmas eval\<^sub>1_evals\<^sub>1_induct = eval\<^sub>1_evals\<^sub>1.induct [split_format (complete)]
- and eval\<^sub>1_evals\<^sub>1_inducts = eval\<^sub>1_evals\<^sub>1.inducts [split_format (complete)]
-(*>*)
-
-
-inductive_cases eval\<^sub>1_cases [cases set]:
- "P \<turnstile>\<^sub>1 \<langle>new C,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
- "P \<turnstile>\<^sub>1 \<langle>Cast C e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
- "P \<turnstile>\<^sub>1 \<langle>Val v,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
- "P \<turnstile>\<^sub>1 \<langle>e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
- "P \<turnstile>\<^sub>1 \<langle>Var v,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
- "P \<turnstile>\<^sub>1 \<langle>V:=e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
- "P \<turnstile>\<^sub>1 \<langle>e\<bullet>F{D},s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
- "P \<turnstile>\<^sub>1 \<langle>C\<bullet>\<^sub>sF{D},s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
- "P \<turnstile>\<^sub>1 \<langle>e\<^sub>1\<bullet>F{D}:=e\<^sub>2,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
- "P \<turnstile>\<^sub>1 \<langle>C\<bullet>\<^sub>sF{D}:=e\<^sub>2,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
- "P \<turnstile>\<^sub>1 \<langle>e\<bullet>M(es),s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
- "P \<turnstile>\<^sub>1 \<langle>C\<bullet>\<^sub>sM(es),s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
- "P \<turnstile>\<^sub>1 \<langle>{V:T;e\<^sub>1},s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
- "P \<turnstile>\<^sub>1 \<langle>e\<^sub>1;;e\<^sub>2,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
- "P \<turnstile>\<^sub>1 \<langle>if (e) e\<^sub>1 else e\<^sub>2,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
- "P \<turnstile>\<^sub>1 \<langle>while (b) c,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
- "P \<turnstile>\<^sub>1 \<langle>throw e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
- "P \<turnstile>\<^sub>1 \<langle>try e\<^sub>1 catch(C V) e\<^sub>2,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
- "P \<turnstile>\<^sub>1 \<langle>INIT C (Cs,b) \<leftarrow> e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
- "P \<turnstile>\<^sub>1 \<langle>RI (C,e);Cs \<leftarrow> e\<^sub>0,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
-
-inductive_cases evals\<^sub>1_cases [cases set]:
- "P \<turnstile>\<^sub>1 \<langle>[],s\<rangle> [\<Rightarrow>] \<langle>e',s'\<rangle>"
- "P \<turnstile>\<^sub>1 \<langle>e#es,s\<rangle> [\<Rightarrow>] \<langle>e',s'\<rangle>"
-(*>*)
-
-
-lemma eval\<^sub>1_final: "P \<turnstile>\<^sub>1 \<langle>e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle> \<Longrightarrow> final e'"
- and evals\<^sub>1_final: "P \<turnstile>\<^sub>1 \<langle>es,s\<rangle> [\<Rightarrow>] \<langle>es',s'\<rangle> \<Longrightarrow> finals es'"
-(*<*)by(induct rule:eval\<^sub>1_evals\<^sub>1.inducts, simp_all)(*>*)
-
-
-lemma eval\<^sub>1_final_same:
-assumes eval: "P \<turnstile>\<^sub>1 \<langle>e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>" shows "final e \<Longrightarrow> e = e' \<and> s = s'"
-(*<*)
-proof(erule finalE)
- fix v assume Val: "e = Val v"
- then show ?thesis using eval eval\<^sub>1_cases(3) by blast
-next
- fix a assume "e = Throw a"
- then show ?thesis using eval
- by (metis eval\<^sub>1_cases(3,17) exp.distinct(101) exp.inject(3) val.distinct(13))
-qed
-(*>*)
-
-subsection "Property preservation"
-
-lemma eval\<^sub>1_preserves_len:
- "P \<turnstile>\<^sub>1 \<langle>e\<^sub>0,(h\<^sub>0,ls\<^sub>0,sh\<^sub>0)\<rangle> \<Rightarrow> \<langle>e\<^sub>1,(h\<^sub>1,ls\<^sub>1,sh\<^sub>1)\<rangle> \<Longrightarrow> length ls\<^sub>0 = length ls\<^sub>1"
-and evals\<^sub>1_preserves_len:
- "P \<turnstile>\<^sub>1 \<langle>es\<^sub>0,(h\<^sub>0,ls\<^sub>0,sh\<^sub>0)\<rangle> [\<Rightarrow>] \<langle>es\<^sub>1,(h\<^sub>1,ls\<^sub>1,sh\<^sub>1)\<rangle> \<Longrightarrow> length ls\<^sub>0 = length ls\<^sub>1"
-(*<*)by (induct rule:eval\<^sub>1_evals\<^sub>1_inducts, simp_all)(*>*)
-
-
-lemma evals\<^sub>1_preserves_elen:
- "\<And>es' s s'. P \<turnstile>\<^sub>1 \<langle>es,s\<rangle> [\<Rightarrow>] \<langle>es',s'\<rangle> \<Longrightarrow> length es = length es'"
-(*<*)by(induct es type:list) (auto elim:evals\<^sub>1.cases)(*>*)
-
-
-lemma clinit\<^sub>1_loc_pres:
- "P \<turnstile>\<^sub>1 \<langle>C\<^sub>0\<bullet>\<^sub>sclinit([]),(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',(h',l',sh')\<rangle> \<Longrightarrow> l = l'"
- by(auto elim!: eval\<^sub>1_cases(12) evals\<^sub>1_cases(1))
-
-lemma
-shows init\<^sub>1_ri\<^sub>1_same_loc: "P \<turnstile>\<^sub>1 \<langle>e,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',(h',l',sh')\<rangle>
- \<Longrightarrow> (\<And>C Cs b M a. e = INIT C (Cs,b) \<leftarrow> unit \<or> e = C\<bullet>\<^sub>sM([]) \<or> e = RI (C,Throw a) ; Cs \<leftarrow> unit
- \<or> e = RI (C,C\<bullet>\<^sub>sclinit([])) ; Cs \<leftarrow> unit
- \<Longrightarrow> l = l')"
- and "P \<turnstile>\<^sub>1 \<langle>es,(h,l,sh)\<rangle> [\<Rightarrow>] \<langle>es',(h',l',sh')\<rangle> \<Longrightarrow> True"
-proof(induct rule: eval\<^sub>1_evals\<^sub>1_inducts)
- case (RInitInitFail\<^sub>1 e h l sh a')
- then show ?case using eval\<^sub>1_final[of _ _ _ "throw a'"]
- by(fastforce dest: eval\<^sub>1_final_same[of _ "Throw a"])
-next
- case RInitFailFinal\<^sub>1 then show ?case by(auto dest: eval\<^sub>1_final_same)
-qed(auto dest: evals\<^sub>1_cases eval\<^sub>1_cases(17) eval\<^sub>1_final_same)
-
-lemma init\<^sub>1_same_loc: "P \<turnstile>\<^sub>1 \<langle>INIT C (Cs,b) \<leftarrow> unit,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',(h',l',sh')\<rangle> \<Longrightarrow> l = l'"
- by(simp add: init\<^sub>1_ri\<^sub>1_same_loc)
-
-
-theorem eval\<^sub>1_hext: "P \<turnstile>\<^sub>1 \<langle>e,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',(h',l',sh')\<rangle> \<Longrightarrow> h \<unlhd> h'"
-and evals\<^sub>1_hext: "P \<turnstile>\<^sub>1 \<langle>es,(h,l,sh)\<rangle> [\<Rightarrow>] \<langle>es',(h',l',sh')\<rangle> \<Longrightarrow> h \<unlhd> h'"
-(*<*)
-proof (induct rule: eval\<^sub>1_evals\<^sub>1_inducts)
- case New\<^sub>1 thus ?case
- by(fastforce intro!: hext_new intro:LeastI simp:new_Addr_def
- split:if_split_asm simp del:fun_upd_apply)
-next
- case NewInit\<^sub>1 thus ?case
- by (meson hext_new hext_trans new_Addr_SomeD)
-next
- case FAss\<^sub>1 thus ?case
- by(auto simp:sym[THEN hext_upd_obj] simp del:fun_upd_apply
- elim!: hext_trans)
-qed (auto elim!: hext_trans)
-(*>*)
-
-subsection "Initialization"
-
-lemma rinit\<^sub>1_throw:
- "P\<^sub>1 \<turnstile>\<^sub>1 \<langle>RI (D,Throw xa) ; Cs \<leftarrow> e,(h, l, sh)\<rangle> \<Rightarrow> \<langle>e',(h', l', sh')\<rangle>
- \<Longrightarrow> e' = Throw xa"
-proof(induct Cs arbitrary: D h l sh h' l' sh')
- case Nil then show ?case
- proof(rule eval\<^sub>1_cases(20)) qed(auto elim: eval\<^sub>1_cases)
-next
- case (Cons C Cs) show ?case using Cons.prems
- proof(induct rule: eval\<^sub>1_cases(20))
- case RInit\<^sub>1: 1
- then show ?case using Cons.hyps by(auto elim: eval\<^sub>1_cases)
- next
- case RInitInitFail\<^sub>1: 2
- then show ?case using Cons.hyps eval\<^sub>1_final_same final_def by blast
- next
- case RInitFailFinal\<^sub>1: 3 then show ?case by simp
- qed
-qed
-
-lemma rinit\<^sub>1_throwE:
-"P \<turnstile>\<^sub>1 \<langle>RI (C,throw e) ; Cs \<leftarrow> e\<^sub>0,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>
- \<Longrightarrow> \<exists>a s\<^sub>t. e' = throw a \<and> P \<turnstile>\<^sub>1 \<langle>throw e,s\<rangle> \<Rightarrow> \<langle>throw a,s\<^sub>t\<rangle>"
-proof(induct Cs arbitrary: C e s)
- case Nil
- then show ?case
- proof(rule eval\<^sub>1_cases(20)) \<comment> \<open> RI \<close>
- fix v h' l' sh' assume "P \<turnstile>\<^sub>1 \<langle>throw e,s\<rangle> \<Rightarrow> \<langle>Val v,(h', l', sh')\<rangle>"
- then show ?case using eval\<^sub>1_cases(17) by blast
- qed(auto)
-next
- case (Cons C' Cs')
- show ?case using Cons.prems(1)
- proof(rule eval\<^sub>1_cases(20)) \<comment> \<open> RI \<close>
- fix v h' l' sh' assume "P \<turnstile>\<^sub>1 \<langle>throw e,s\<rangle> \<Rightarrow> \<langle>Val v,(h', l', sh')\<rangle>"
- then show ?case using eval\<^sub>1_cases(17) by blast
- next
- fix a h' l' sh' sfs i D Cs''
- assume e''step: "P \<turnstile>\<^sub>1 \<langle>throw e,s\<rangle> \<Rightarrow> \<langle>throw a,(h', l', sh')\<rangle>"
- and shC: "sh' C = \<lfloor>(sfs, i)\<rfloor>"
- and riD: "P \<turnstile>\<^sub>1 \<langle>RI (D,throw a) ; Cs'' \<leftarrow> e\<^sub>0,(h', l', sh'(C \<mapsto> (sfs, Error)))\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
- and "C' # Cs' = D # Cs''"
- then show ?thesis using Cons.hyps eval\<^sub>1_final eval\<^sub>1_final_same by blast
- qed(simp)
-qed
-
-end
+(* Title: JinjaDCI/Compiler/J1.thy
+ Author: Tobias Nipkow, Susannah Mansky
+ Copyright 2003 Technische Universitaet Muenchen, 2019-20 UIUC
+
+ Based on the Jinja theory Compiler/J1.thy by Tobias Nipkow
+*)
+
+chapter \<open> Compilation \label{cha:comp} \<close>
+
+section \<open> An Intermediate Language \<close>
+
+theory J1 imports "../J/BigStep" begin
+
+type_synonym expr\<^sub>1 = "nat exp"
+type_synonym J\<^sub>1_prog = "expr\<^sub>1 prog"
+type_synonym state\<^sub>1 = "heap \<times> (val list) \<times> sheap"
+
+definition hp\<^sub>1 :: "state\<^sub>1 \<Rightarrow> heap"
+where
+ "hp\<^sub>1 \<equiv> fst"
+definition lcl\<^sub>1 :: "state\<^sub>1 \<Rightarrow> val list"
+where
+ "lcl\<^sub>1 \<equiv> fst \<circ> snd"
+definition shp\<^sub>1 :: "state\<^sub>1 \<Rightarrow> sheap"
+where
+ "shp\<^sub>1 \<equiv> snd \<circ> snd"
+
+(*<*)
+declare hp\<^sub>1_def[simp] lcl\<^sub>1_def[simp] shp\<^sub>1_def[simp]
+(*>*)
+
+primrec
+ max_vars :: "'a exp \<Rightarrow> nat"
+ and max_varss :: "'a exp list \<Rightarrow> nat"
+where
+ "max_vars(new C) = 0"
+| "max_vars(Cast C e) = max_vars e"
+| "max_vars(Val v) = 0"
+| "max_vars(e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2) = max (max_vars e\<^sub>1) (max_vars e\<^sub>2)"
+| "max_vars(Var V) = 0"
+| "max_vars(V:=e) = max_vars e"
+| "max_vars(e\<bullet>F{D}) = max_vars e"
+| "max_vars(C\<bullet>\<^sub>sF{D}) = 0"
+| "max_vars(FAss e\<^sub>1 F D e\<^sub>2) = max (max_vars e\<^sub>1) (max_vars e\<^sub>2)"
+| "max_vars(SFAss C F D e\<^sub>2) = max_vars e\<^sub>2"
+| "max_vars(e\<bullet>M(es)) = max (max_vars e) (max_varss es)"
+| "max_vars(C\<bullet>\<^sub>sM(es)) = max_varss es"
+| "max_vars({V:T; e}) = max_vars e + 1"
+| "max_vars(e\<^sub>1;;e\<^sub>2) = max (max_vars e\<^sub>1) (max_vars e\<^sub>2)"
+| "max_vars(if (e) e\<^sub>1 else e\<^sub>2) =
+ max (max_vars e) (max (max_vars e\<^sub>1) (max_vars e\<^sub>2))"
+| "max_vars(while (b) e) = max (max_vars b) (max_vars e)"
+| "max_vars(throw e) = max_vars e"
+| "max_vars(try e\<^sub>1 catch(C V) e\<^sub>2) = max (max_vars e\<^sub>1) (max_vars e\<^sub>2 + 1)"
+| "max_vars(INIT C (Cs,b) \<leftarrow> e) = max_vars e"
+| "max_vars(RI(C,e);Cs \<leftarrow> e') = max (max_vars e) (max_vars e')"
+
+| "max_varss [] = 0"
+| "max_varss (e#es) = max (max_vars e) (max_varss es)"
+
+inductive
+ eval\<^sub>1 :: "J\<^sub>1_prog \<Rightarrow> expr\<^sub>1 \<Rightarrow> state\<^sub>1 \<Rightarrow> expr\<^sub>1 \<Rightarrow> state\<^sub>1 \<Rightarrow> bool"
+ ("_ \<turnstile>\<^sub>1 ((1\<langle>_,/_\<rangle>) \<Rightarrow>/ (1\<langle>_,/_\<rangle>))" [51,0,0,0,0] 81)
+ and evals\<^sub>1 :: "J\<^sub>1_prog \<Rightarrow> expr\<^sub>1 list \<Rightarrow> state\<^sub>1 \<Rightarrow> expr\<^sub>1 list \<Rightarrow> state\<^sub>1 \<Rightarrow> bool"
+ ("_ \<turnstile>\<^sub>1 ((1\<langle>_,/_\<rangle>) [\<Rightarrow>]/ (1\<langle>_,/_\<rangle>))" [51,0,0,0,0] 81)
+ for P :: J\<^sub>1_prog
+where
+
+ New\<^sub>1:
+ "\<lbrakk> sh C = Some (sfs, Done); new_Addr h = Some a;
+ P \<turnstile> C has_fields FDTs; h' = h(a\<mapsto>blank P C) \<rbrakk>
+ \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>new C,(h,l,sh)\<rangle> \<Rightarrow> \<langle>addr a,(h',l,sh)\<rangle>"
+| NewFail\<^sub>1:
+ "\<lbrakk> sh C = Some (sfs, Done); new_Addr h = None \<rbrakk> \<Longrightarrow>
+ P \<turnstile>\<^sub>1 \<langle>new C, (h,l,sh)\<rangle> \<Rightarrow> \<langle>THROW OutOfMemory,(h,l,sh)\<rangle>"
+| NewInit\<^sub>1:
+ "\<lbrakk> \<nexists>sfs. sh C = Some (sfs, Done); P \<turnstile>\<^sub>1 \<langle>INIT C ([C],False) \<leftarrow> unit,(h,l,sh)\<rangle> \<Rightarrow> \<langle>Val v',(h',l',sh')\<rangle>;
+ new_Addr h' = Some a; P \<turnstile> C has_fields FDTs; h'' = h'(a\<mapsto>blank P C) \<rbrakk>
+ \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>new C,(h,l,sh)\<rangle> \<Rightarrow> \<langle>addr a,(h'',l',sh')\<rangle>"
+| NewInitOOM\<^sub>1:
+ "\<lbrakk> \<nexists>sfs. sh C = Some (sfs, Done); P \<turnstile>\<^sub>1 \<langle>INIT C ([C],False) \<leftarrow> unit,(h,l,sh)\<rangle> \<Rightarrow> \<langle>Val v',(h',l',sh')\<rangle>;
+ new_Addr h' = None; is_class P C \<rbrakk>
+ \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>new C,(h,l,sh)\<rangle> \<Rightarrow> \<langle>THROW OutOfMemory,(h',l',sh')\<rangle>"
+| NewInitThrow\<^sub>1:
+ "\<lbrakk> \<nexists>sfs. sh C = Some (sfs, Done); P \<turnstile>\<^sub>1 \<langle>INIT C ([C],False) \<leftarrow> unit,(h,l,sh)\<rangle> \<Rightarrow> \<langle>throw a,s'\<rangle>;
+ is_class P C \<rbrakk>
+ \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>new C,(h,l,sh)\<rangle> \<Rightarrow> \<langle>throw a,s'\<rangle>"
+
+| Cast\<^sub>1:
+ "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>addr a,(h,l,sh)\<rangle>; h a = Some(D,fs); P \<turnstile> D \<preceq>\<^sup>* C \<rbrakk>
+ \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>Cast C e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>addr a,(h,l,sh)\<rangle>"
+| CastNull\<^sub>1:
+ "P \<turnstile>\<^sub>1 \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>null,s\<^sub>1\<rangle> \<Longrightarrow>
+ P \<turnstile>\<^sub>1 \<langle>Cast C e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>null,s\<^sub>1\<rangle>"
+| CastFail\<^sub>1:
+ "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>addr a,(h,l,sh)\<rangle>; h a = Some(D,fs); \<not> P \<turnstile> D \<preceq>\<^sup>* C \<rbrakk>
+ \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>Cast C e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>THROW ClassCast,(h,l,sh)\<rangle>"
+| CastThrow\<^sub>1:
+ "P \<turnstile>\<^sub>1 \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle> \<Longrightarrow>
+ P \<turnstile>\<^sub>1 \<langle>Cast C e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle>"
+
+| Val\<^sub>1:
+ "P \<turnstile>\<^sub>1 \<langle>Val v,s\<rangle> \<Rightarrow> \<langle>Val v,s\<rangle>"
+
+| BinOp\<^sub>1:
+ "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Val v\<^sub>1,s\<^sub>1\<rangle>; P \<turnstile>\<^sub>1 \<langle>e\<^sub>2,s\<^sub>1\<rangle> \<Rightarrow> \<langle>Val v\<^sub>2,s\<^sub>2\<rangle>; binop(bop,v\<^sub>1,v\<^sub>2) = Some v \<rbrakk>
+ \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Val v,s\<^sub>2\<rangle>"
+| BinOpThrow\<^sub>1\<^sub>1:
+ "P \<turnstile>\<^sub>1 \<langle>e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e,s\<^sub>1\<rangle> \<Longrightarrow>
+ P \<turnstile>\<^sub>1 \<langle>e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2, s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e,s\<^sub>1\<rangle>"
+| BinOpThrow\<^sub>2\<^sub>1:
+ "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Val v\<^sub>1,s\<^sub>1\<rangle>; P \<turnstile>\<^sub>1 \<langle>e\<^sub>2,s\<^sub>1\<rangle> \<Rightarrow> \<langle>throw e,s\<^sub>2\<rangle> \<rbrakk>
+ \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e,s\<^sub>2\<rangle>"
+
+| Var\<^sub>1:
+ "\<lbrakk> ls!i = v; i < size ls \<rbrakk> \<Longrightarrow>
+ P \<turnstile>\<^sub>1 \<langle>Var i,(h,ls,sh)\<rangle> \<Rightarrow> \<langle>Val v,(h,ls,sh)\<rangle>"
+
+| LAss\<^sub>1:
+ "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Val v,(h,ls,sh)\<rangle>; i < size ls; ls' = ls[i := v] \<rbrakk>
+ \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>i:= e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>unit,(h,ls',sh)\<rangle>"
+| LAssThrow\<^sub>1:
+ "P \<turnstile>\<^sub>1 \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle> \<Longrightarrow>
+ P \<turnstile>\<^sub>1 \<langle>i:= e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle>"
+
+| FAcc\<^sub>1:
+ "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>addr a,(h,ls,sh)\<rangle>; h a = Some(C,fs);
+ P \<turnstile> C has F,NonStatic:t in D;
+ fs(F,D) = Some v \<rbrakk>
+ \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>e\<bullet>F{D},s\<^sub>0\<rangle> \<Rightarrow> \<langle>Val v,(h,ls,sh)\<rangle>"
+| FAccNull\<^sub>1:
+ "P \<turnstile>\<^sub>1 \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>null,s\<^sub>1\<rangle> \<Longrightarrow>
+ P \<turnstile>\<^sub>1 \<langle>e\<bullet>F{D},s\<^sub>0\<rangle> \<Rightarrow> \<langle>THROW NullPointer,s\<^sub>1\<rangle>"
+| FAccThrow\<^sub>1:
+ "P \<turnstile>\<^sub>1 \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle> \<Longrightarrow>
+ P \<turnstile>\<^sub>1 \<langle>e\<bullet>F{D},s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle>"
+| FAccNone\<^sub>1:
+ "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>addr a,(h,ls,sh)\<rangle>; h a = Some(C,fs);
+ \<not>(\<exists>b t. P \<turnstile> C has F,b:t in D) \<rbrakk>
+ \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>e\<bullet>F{D},s\<^sub>0\<rangle> \<Rightarrow> \<langle>THROW NoSuchFieldError,(h,ls,sh)\<rangle>"
+| FAccStatic\<^sub>1:
+ "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>addr a,(h,ls,sh)\<rangle>; h a = Some(C,fs);
+ P \<turnstile> C has F,Static:t in D \<rbrakk>
+ \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>e\<bullet>F{D},s\<^sub>0\<rangle> \<Rightarrow> \<langle>THROW IncompatibleClassChangeError,(h,ls,sh)\<rangle>"
+
+| SFAcc\<^sub>1:
+ "\<lbrakk> P \<turnstile> C has F,Static:t in D;
+ sh D = Some (sfs,Done);
+ sfs F = Some v \<rbrakk>
+ \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>C\<bullet>\<^sub>sF{D},(h,ls,sh)\<rangle> \<Rightarrow> \<langle>Val v,(h,ls,sh)\<rangle>"
+| SFAccInit\<^sub>1:
+ "\<lbrakk> P \<turnstile> C has F,Static:t in D;
+ \<nexists>sfs. sh D = Some (sfs,Done); P \<turnstile>\<^sub>1 \<langle>INIT D ([D],False) \<leftarrow> unit,(h,ls,sh)\<rangle> \<Rightarrow> \<langle>Val v',(h',ls',sh')\<rangle>;
+ sh' D = Some (sfs,i);
+ sfs F = Some v \<rbrakk>
+ \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>C\<bullet>\<^sub>sF{D},(h,ls,sh)\<rangle> \<Rightarrow> \<langle>Val v,(h',ls',sh')\<rangle>"
+| SFAccInitThrow\<^sub>1:
+ "\<lbrakk> P \<turnstile> C has F,Static:t in D;
+ \<nexists>sfs. sh D = Some (sfs,Done); P \<turnstile>\<^sub>1 \<langle>INIT D ([D],False) \<leftarrow> unit,(h,ls,sh)\<rangle> \<Rightarrow> \<langle>throw a,s'\<rangle> \<rbrakk>
+ \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>C\<bullet>\<^sub>sF{D},(h,ls,sh)\<rangle> \<Rightarrow> \<langle>throw a,s'\<rangle>"
+| SFAccNone\<^sub>1:
+ "\<lbrakk> \<not>(\<exists>b t. P \<turnstile> C has F,b:t in D) \<rbrakk>
+ \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>C\<bullet>\<^sub>sF{D},s\<rangle> \<Rightarrow> \<langle>THROW NoSuchFieldError,s\<rangle>"
+| SFAccNonStatic\<^sub>1:
+ "\<lbrakk> P \<turnstile> C has F,NonStatic:t in D \<rbrakk>
+ \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>C\<bullet>\<^sub>sF{D},s\<rangle> \<Rightarrow> \<langle>THROW IncompatibleClassChangeError,s\<rangle>"
+
+
+| FAss\<^sub>1:
+ "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow> \<langle>addr a,s\<^sub>1\<rangle>; P \<turnstile>\<^sub>1 \<langle>e\<^sub>2,s\<^sub>1\<rangle> \<Rightarrow> \<langle>Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\<rangle>;
+ h\<^sub>2 a = Some(C,fs); P \<turnstile> C has F,NonStatic:t in D;
+ fs' = fs((F,D)\<mapsto>v); h\<^sub>2' = h\<^sub>2(a\<mapsto>(C,fs')) \<rbrakk>
+ \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>e\<^sub>1\<bullet>F{D}:=e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>unit,(h\<^sub>2',l\<^sub>2,sh\<^sub>2)\<rangle>"
+| FAssNull\<^sub>1:
+ "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow> \<langle>null,s\<^sub>1\<rangle>; P \<turnstile>\<^sub>1 \<langle>e\<^sub>2,s\<^sub>1\<rangle> \<Rightarrow> \<langle>Val v,s\<^sub>2\<rangle> \<rbrakk> \<Longrightarrow>
+ P \<turnstile>\<^sub>1 \<langle>e\<^sub>1\<bullet>F{D}:=e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>THROW NullPointer,s\<^sub>2\<rangle>"
+| FAssThrow\<^sub>1\<^sub>1:
+ "P \<turnstile>\<^sub>1 \<langle>e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle> \<Longrightarrow>
+ P \<turnstile>\<^sub>1 \<langle>e\<^sub>1\<bullet>F{D}:=e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle>"
+| FAssThrow\<^sub>2\<^sub>1:
+ "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Val v,s\<^sub>1\<rangle>; P \<turnstile>\<^sub>1 \<langle>e\<^sub>2,s\<^sub>1\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>2\<rangle> \<rbrakk>
+ \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>e\<^sub>1\<bullet>F{D}:=e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>2\<rangle>"
+| FAssNone\<^sub>1:
+ "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow> \<langle>addr a,s\<^sub>1\<rangle>; P \<turnstile>\<^sub>1 \<langle>e\<^sub>2,s\<^sub>1\<rangle> \<Rightarrow> \<langle>Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\<rangle>;
+ h\<^sub>2 a = Some(C,fs); \<not>(\<exists>b t. P \<turnstile> C has F,b:t in D) \<rbrakk>
+ \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>e\<^sub>1\<bullet>F{D}:=e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>THROW NoSuchFieldError,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\<rangle>"
+| FAssStatic\<^sub>1:
+ "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow> \<langle>addr a,s\<^sub>1\<rangle>; P \<turnstile>\<^sub>1 \<langle>e\<^sub>2,s\<^sub>1\<rangle> \<Rightarrow> \<langle>Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\<rangle>;
+ h\<^sub>2 a = Some(C,fs); P \<turnstile> C has F,Static:t in D \<rbrakk>
+ \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>e\<^sub>1\<bullet>F{D}:=e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>THROW IncompatibleClassChangeError,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\<rangle>"
+
+| SFAss\<^sub>1:
+ "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Val v,(h\<^sub>1,l\<^sub>1,sh\<^sub>1)\<rangle>;
+ P \<turnstile> C has F,Static:t in D;
+ sh\<^sub>1 D = Some(sfs,Done); sfs' = sfs(F\<mapsto>v); sh\<^sub>1' = sh\<^sub>1(D\<mapsto>(sfs',Done)) \<rbrakk>
+ \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>C\<bullet>\<^sub>sF{D}:=e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>unit,(h\<^sub>1,l\<^sub>1,sh\<^sub>1')\<rangle>"
+| SFAssInit\<^sub>1:
+ "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Val v,(h\<^sub>1,l\<^sub>1,sh\<^sub>1)\<rangle>;
+ P \<turnstile> C has F,Static:t in D;
+ \<nexists>sfs. sh\<^sub>1 D = Some(sfs,Done); P \<turnstile>\<^sub>1 \<langle>INIT D ([D],False) \<leftarrow> unit,(h\<^sub>1,l\<^sub>1,sh\<^sub>1)\<rangle> \<Rightarrow> \<langle>Val v',(h',l',sh')\<rangle>;
+ sh' D = Some(sfs,i);
+ sfs' = sfs(F\<mapsto>v); sh'' = sh'(D\<mapsto>(sfs',i)) \<rbrakk>
+ \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>C\<bullet>\<^sub>sF{D}:=e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>unit,(h',l',sh'')\<rangle>"
+| SFAssInitThrow\<^sub>1:
+ "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Val v,(h\<^sub>1,l\<^sub>1,sh\<^sub>1)\<rangle>;
+ P \<turnstile> C has F,Static:t in D;
+ \<nexists>sfs. sh\<^sub>1 D = Some(sfs,Done); P \<turnstile>\<^sub>1 \<langle>INIT D ([D],False) \<leftarrow> unit,(h\<^sub>1,l\<^sub>1,sh\<^sub>1)\<rangle> \<Rightarrow> \<langle>throw a,s'\<rangle> \<rbrakk>
+ \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>C\<bullet>\<^sub>sF{D}:=e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw a,s'\<rangle>"
+| SFAssThrow\<^sub>1:
+ "P \<turnstile>\<^sub>1 \<langle>e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>2\<rangle>
+ \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>C\<bullet>\<^sub>sF{D}:=e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>2\<rangle>"
+| SFAssNone\<^sub>1:
+ "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\<rangle>;
+ \<not>(\<exists>b t. P \<turnstile> C has F,b:t in D) \<rbrakk>
+ \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>C\<bullet>\<^sub>sF{D}:=e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>THROW NoSuchFieldError,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\<rangle>"
+| SFAssNonStatic\<^sub>1:
+ "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\<rangle>;
+ P \<turnstile> C has F,NonStatic:t in D \<rbrakk>
+ \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>C\<bullet>\<^sub>sF{D}:=e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>THROW IncompatibleClassChangeError,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\<rangle>"
+
+| CallObjThrow\<^sub>1:
+ "P \<turnstile>\<^sub>1 \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle> \<Longrightarrow>
+ P \<turnstile>\<^sub>1 \<langle>e\<bullet>M(es),s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle>"
+| CallNull\<^sub>1:
+ "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>null,s\<^sub>1\<rangle>; P \<turnstile>\<^sub>1 \<langle>es,s\<^sub>1\<rangle> [\<Rightarrow>] \<langle>map Val vs,s\<^sub>2\<rangle> \<rbrakk>
+ \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>e\<bullet>M(es),s\<^sub>0\<rangle> \<Rightarrow> \<langle>THROW NullPointer,s\<^sub>2\<rangle>"
+| Call\<^sub>1:
+ "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>addr a,s\<^sub>1\<rangle>; P \<turnstile>\<^sub>1 \<langle>es,s\<^sub>1\<rangle> [\<Rightarrow>] \<langle>map Val vs,(h\<^sub>2,ls\<^sub>2,sh\<^sub>2)\<rangle>;
+ h\<^sub>2 a = Some(C,fs); P \<turnstile> C sees M,NonStatic:Ts\<rightarrow>T = body in D;
+ size vs = size Ts; ls\<^sub>2' = (Addr a) # vs @ replicate (max_vars body) undefined;
+ P \<turnstile>\<^sub>1 \<langle>body,(h\<^sub>2,ls\<^sub>2',sh\<^sub>2)\<rangle> \<Rightarrow> \<langle>e',(h\<^sub>3,ls\<^sub>3,sh\<^sub>3)\<rangle> \<rbrakk>
+ \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>e\<bullet>M(es),s\<^sub>0\<rangle> \<Rightarrow> \<langle>e',(h\<^sub>3,ls\<^sub>2,sh\<^sub>3)\<rangle>"
+| CallParamsThrow\<^sub>1:
+ "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Val v,s\<^sub>1\<rangle>; P \<turnstile>\<^sub>1 \<langle>es,s\<^sub>1\<rangle> [\<Rightarrow>] \<langle>es',s\<^sub>2\<rangle>;
+ es' = map Val vs @ throw ex # es\<^sub>2 \<rbrakk>
+ \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>e\<bullet>M(es),s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw ex,s\<^sub>2\<rangle>"
+| CallNone\<^sub>1:
+ "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>addr a,s\<^sub>1\<rangle>; P \<turnstile>\<^sub>1 \<langle>ps,s\<^sub>1\<rangle> [\<Rightarrow>] \<langle>map Val vs,(h\<^sub>2,ls\<^sub>2,sh\<^sub>2)\<rangle>;
+ h\<^sub>2 a = Some(C,fs); \<not>(\<exists>b Ts T body D. P \<turnstile> C sees M,b:Ts\<rightarrow>T = body in D) \<rbrakk>
+ \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>e\<bullet>M(ps),s\<^sub>0\<rangle> \<Rightarrow> \<langle>THROW NoSuchMethodError,(h\<^sub>2,ls\<^sub>2,sh\<^sub>2)\<rangle>"
+| CallStatic\<^sub>1:
+ "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>addr a,s\<^sub>1\<rangle>; P \<turnstile>\<^sub>1 \<langle>ps,s\<^sub>1\<rangle> [\<Rightarrow>] \<langle>map Val vs,(h\<^sub>2,ls\<^sub>2,sh\<^sub>2)\<rangle>;
+ h\<^sub>2 a = Some(C,fs); P \<turnstile> C sees M,Static:Ts\<rightarrow>T = body in D \<rbrakk>
+ \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>e\<bullet>M(ps),s\<^sub>0\<rangle> \<Rightarrow> \<langle>THROW IncompatibleClassChangeError,(h\<^sub>2,ls\<^sub>2,sh\<^sub>2)\<rangle>"
+
+| SCallParamsThrow\<^sub>1:
+ "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>es,s\<^sub>0\<rangle> [\<Rightarrow>] \<langle>es',s\<^sub>2\<rangle>; es' = map Val vs @ throw ex # es\<^sub>2 \<rbrakk>
+ \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>C\<bullet>\<^sub>sM(es),s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw ex,s\<^sub>2\<rangle>"
+| SCallNone\<^sub>1:
+ "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>ps,s\<^sub>0\<rangle> [\<Rightarrow>] \<langle>map Val vs,s\<^sub>2\<rangle>;
+ \<not>(\<exists>b Ts T body D. P \<turnstile> C sees M,b:Ts\<rightarrow>T = body in D) \<rbrakk>
+ \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>C\<bullet>\<^sub>sM(ps),s\<^sub>0\<rangle> \<Rightarrow> \<langle>THROW NoSuchMethodError,s\<^sub>2\<rangle>"
+| SCallNonStatic\<^sub>1:
+ "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>ps,s\<^sub>0\<rangle> [\<Rightarrow>] \<langle>map Val vs,s\<^sub>2\<rangle>;
+ P \<turnstile> C sees M,NonStatic:Ts\<rightarrow>T = body in D \<rbrakk>
+ \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>C\<bullet>\<^sub>sM(ps),s\<^sub>0\<rangle> \<Rightarrow> \<langle>THROW IncompatibleClassChangeError,s\<^sub>2\<rangle>"
+| SCallInitThrow\<^sub>1:
+ "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>ps,s\<^sub>0\<rangle> [\<Rightarrow>] \<langle>map Val vs,(h\<^sub>1,ls\<^sub>1,sh\<^sub>1)\<rangle>;
+ P \<turnstile> C sees M,Static:Ts\<rightarrow>T = body in D;
+ \<nexists>sfs. sh\<^sub>1 D = Some(sfs,Done); M \<noteq> clinit;
+ P \<turnstile>\<^sub>1 \<langle>INIT D ([D],False) \<leftarrow> unit,(h\<^sub>1,ls\<^sub>1,sh\<^sub>1)\<rangle> \<Rightarrow> \<langle>throw a,s'\<rangle> \<rbrakk>
+ \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>C\<bullet>\<^sub>sM(ps),s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw a,s'\<rangle>"
+| SCallInit\<^sub>1:
+ "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>ps,s\<^sub>0\<rangle> [\<Rightarrow>] \<langle>map Val vs,(h\<^sub>1,ls\<^sub>1,sh\<^sub>1)\<rangle>;
+ P \<turnstile> C sees M,Static:Ts\<rightarrow>T = body in D;
+ \<nexists>sfs. sh\<^sub>1 D = Some(sfs,Done); M \<noteq> clinit;
+ P \<turnstile>\<^sub>1 \<langle>INIT D ([D],False) \<leftarrow> unit,(h\<^sub>1,ls\<^sub>1,sh\<^sub>1)\<rangle> \<Rightarrow> \<langle>Val v',(h\<^sub>2,ls\<^sub>2,sh\<^sub>2)\<rangle>;
+ size vs = size Ts; ls\<^sub>2' = vs @ replicate (max_vars body) undefined;
+ P \<turnstile>\<^sub>1 \<langle>body,(h\<^sub>2,ls\<^sub>2',sh\<^sub>2)\<rangle> \<Rightarrow> \<langle>e',(h\<^sub>3,ls\<^sub>3,sh\<^sub>3)\<rangle> \<rbrakk>
+ \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>C\<bullet>\<^sub>sM(ps),s\<^sub>0\<rangle> \<Rightarrow> \<langle>e',(h\<^sub>3,ls\<^sub>2,sh\<^sub>3)\<rangle>"
+| SCall\<^sub>1:
+ "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>ps,s\<^sub>0\<rangle> [\<Rightarrow>] \<langle>map Val vs,(h\<^sub>2,ls\<^sub>2,sh\<^sub>2)\<rangle>;
+ P \<turnstile> C sees M,Static:Ts\<rightarrow>T = body in D;
+ sh\<^sub>2 D = Some(sfs,Done) \<or> (M = clinit \<and> sh\<^sub>2 D = \<lfloor>(sfs, Processing)\<rfloor>);
+ size vs = size Ts; ls\<^sub>2' = vs @ replicate (max_vars body) undefined;
+ P \<turnstile>\<^sub>1 \<langle>body,(h\<^sub>2,ls\<^sub>2',sh\<^sub>2)\<rangle> \<Rightarrow> \<langle>e',(h\<^sub>3,ls\<^sub>3,sh\<^sub>3)\<rangle> \<rbrakk>
+ \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>C\<bullet>\<^sub>sM(ps),s\<^sub>0\<rangle> \<Rightarrow> \<langle>e',(h\<^sub>3,ls\<^sub>2,sh\<^sub>3)\<rangle>"
+
+| Block\<^sub>1:
+ "P \<turnstile>\<^sub>1 \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>e',s\<^sub>1\<rangle> \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>Block i T e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>e',s\<^sub>1\<rangle>"
+
+| Seq\<^sub>1:
+ "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e\<^sub>0,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Val v,s\<^sub>1\<rangle>; P \<turnstile>\<^sub>1 \<langle>e\<^sub>1,s\<^sub>1\<rangle> \<Rightarrow> \<langle>e\<^sub>2,s\<^sub>2\<rangle> \<rbrakk>
+ \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>e\<^sub>0;;e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow> \<langle>e\<^sub>2,s\<^sub>2\<rangle>"
+| SeqThrow\<^sub>1:
+ "P \<turnstile>\<^sub>1 \<langle>e\<^sub>0,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e,s\<^sub>1\<rangle> \<Longrightarrow>
+ P \<turnstile>\<^sub>1 \<langle>e\<^sub>0;;e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e,s\<^sub>1\<rangle>"
+
+| CondT\<^sub>1:
+ "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>true,s\<^sub>1\<rangle>; P \<turnstile>\<^sub>1 \<langle>e\<^sub>1,s\<^sub>1\<rangle> \<Rightarrow> \<langle>e',s\<^sub>2\<rangle> \<rbrakk>
+ \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>if (e) e\<^sub>1 else e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>e',s\<^sub>2\<rangle>"
+| CondF\<^sub>1:
+ "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>false,s\<^sub>1\<rangle>; P \<turnstile>\<^sub>1 \<langle>e\<^sub>2,s\<^sub>1\<rangle> \<Rightarrow> \<langle>e',s\<^sub>2\<rangle> \<rbrakk>
+ \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>if (e) e\<^sub>1 else e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>e',s\<^sub>2\<rangle>"
+| CondThrow\<^sub>1:
+ "P \<turnstile>\<^sub>1 \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle> \<Longrightarrow>
+ P \<turnstile>\<^sub>1 \<langle>if (e) e\<^sub>1 else e\<^sub>2, s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle>"
+
+| WhileF\<^sub>1:
+ "P \<turnstile>\<^sub>1 \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>false,s\<^sub>1\<rangle> \<Longrightarrow>
+ P \<turnstile>\<^sub>1 \<langle>while (e) c,s\<^sub>0\<rangle> \<Rightarrow> \<langle>unit,s\<^sub>1\<rangle>"
+| WhileT\<^sub>1:
+ "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>true,s\<^sub>1\<rangle>; P \<turnstile>\<^sub>1 \<langle>c,s\<^sub>1\<rangle> \<Rightarrow> \<langle>Val v\<^sub>1,s\<^sub>2\<rangle>;
+ P \<turnstile>\<^sub>1 \<langle>while (e) c,s\<^sub>2\<rangle> \<Rightarrow> \<langle>e\<^sub>3,s\<^sub>3\<rangle> \<rbrakk>
+ \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>while (e) c,s\<^sub>0\<rangle> \<Rightarrow> \<langle>e\<^sub>3,s\<^sub>3\<rangle>"
+| WhileCondThrow\<^sub>1:
+ "P \<turnstile>\<^sub>1 \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle> \<Longrightarrow>
+ P \<turnstile>\<^sub>1 \<langle>while (e) c,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle>"
+| WhileBodyThrow\<^sub>1:
+ "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>true,s\<^sub>1\<rangle>; P \<turnstile>\<^sub>1 \<langle>c,s\<^sub>1\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>2\<rangle>\<rbrakk>
+ \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>while (e) c,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>2\<rangle>"
+
+| Throw\<^sub>1:
+ "P \<turnstile>\<^sub>1 \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>addr a,s\<^sub>1\<rangle> \<Longrightarrow>
+ P \<turnstile>\<^sub>1 \<langle>throw e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Throw a,s\<^sub>1\<rangle>"
+| ThrowNull\<^sub>1:
+ "P \<turnstile>\<^sub>1 \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>null,s\<^sub>1\<rangle> \<Longrightarrow>
+ P \<turnstile>\<^sub>1 \<langle>throw e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>THROW NullPointer,s\<^sub>1\<rangle>"
+| ThrowThrow\<^sub>1:
+ "P \<turnstile>\<^sub>1 \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle> \<Longrightarrow>
+ P \<turnstile>\<^sub>1 \<langle>throw e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle>"
+
+| Try\<^sub>1:
+ "P \<turnstile>\<^sub>1 \<langle>e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Val v\<^sub>1,s\<^sub>1\<rangle> \<Longrightarrow>
+ P \<turnstile>\<^sub>1 \<langle>try e\<^sub>1 catch(C i) e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Val v\<^sub>1,s\<^sub>1\<rangle>"
+| TryCatch\<^sub>1:
+ "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Throw a,(h\<^sub>1,ls\<^sub>1,sh\<^sub>1)\<rangle>;
+ h\<^sub>1 a = Some(D,fs); P \<turnstile> D \<preceq>\<^sup>* C; i < length ls\<^sub>1;
+ P \<turnstile>\<^sub>1 \<langle>e\<^sub>2,(h\<^sub>1,ls\<^sub>1[i:=Addr a],sh\<^sub>1)\<rangle> \<Rightarrow> \<langle>e\<^sub>2',(h\<^sub>2,ls\<^sub>2,sh\<^sub>2)\<rangle> \<rbrakk>
+ \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>try e\<^sub>1 catch(C i) e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>e\<^sub>2',(h\<^sub>2,ls\<^sub>2,sh\<^sub>2)\<rangle>"
+| TryThrow\<^sub>1:
+ "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Throw a,(h\<^sub>1,ls\<^sub>1,sh\<^sub>1)\<rangle>; h\<^sub>1 a = Some(D,fs); \<not> P \<turnstile> D \<preceq>\<^sup>* C \<rbrakk>
+ \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>try e\<^sub>1 catch(C i) e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Throw a,(h\<^sub>1,ls\<^sub>1,sh\<^sub>1)\<rangle>"
+
+| Nil\<^sub>1:
+ "P \<turnstile>\<^sub>1 \<langle>[],s\<rangle> [\<Rightarrow>] \<langle>[],s\<rangle>"
+
+| Cons\<^sub>1:
+ "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Val v,s\<^sub>1\<rangle>; P \<turnstile>\<^sub>1 \<langle>es,s\<^sub>1\<rangle> [\<Rightarrow>] \<langle>es',s\<^sub>2\<rangle> \<rbrakk>
+ \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>e#es,s\<^sub>0\<rangle> [\<Rightarrow>] \<langle>Val v # es',s\<^sub>2\<rangle>"
+| ConsThrow\<^sub>1:
+ "P \<turnstile>\<^sub>1 \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle> \<Longrightarrow>
+ P \<turnstile>\<^sub>1 \<langle>e#es,s\<^sub>0\<rangle> [\<Rightarrow>] \<langle>throw e' # es, s\<^sub>1\<rangle>"
+
+\<comment> \<open> init rules \<close>
+
+| InitFinal\<^sub>1:
+ "P \<turnstile>\<^sub>1 \<langle>e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle> \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>INIT C (Nil,b) \<leftarrow> e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+| InitNone\<^sub>1:
+ "\<lbrakk> sh C = None; P \<turnstile>\<^sub>1 \<langle>INIT C' (C#Cs,False) \<leftarrow> e,(h,l,sh(C \<mapsto> (sblank P C, Prepared)))\<rangle> \<Rightarrow> \<langle>e',s'\<rangle> \<rbrakk>
+ \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>INIT C' (C#Cs,False) \<leftarrow> e,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+| InitDone\<^sub>1:
+ "\<lbrakk> sh C = Some(sfs,Done); P \<turnstile>\<^sub>1 \<langle>INIT C' (Cs,True) \<leftarrow> e,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle> \<rbrakk>
+ \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>INIT C' (C#Cs,False) \<leftarrow> e,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+| InitProcessing\<^sub>1:
+ "\<lbrakk> sh C = Some(sfs,Processing); P \<turnstile>\<^sub>1 \<langle>INIT C' (Cs,True) \<leftarrow> e,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle> \<rbrakk>
+ \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>INIT C' (C#Cs,False) \<leftarrow> e,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+| InitError\<^sub>1:
+ "\<lbrakk> sh C = Some(sfs,Error);
+ P \<turnstile>\<^sub>1 \<langle>RI (C, THROW NoClassDefFoundError);Cs \<leftarrow> e,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle> \<rbrakk>
+ \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>INIT C' (C#Cs,False) \<leftarrow> e,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+| InitObject\<^sub>1:
+ "\<lbrakk> sh C = Some(sfs,Prepared);
+ C = Object;
+ sh' = sh(C \<mapsto> (sfs,Processing));
+ P \<turnstile>\<^sub>1 \<langle>INIT C' (C#Cs,True) \<leftarrow> e,(h,l,sh')\<rangle> \<Rightarrow> \<langle>e',s'\<rangle> \<rbrakk>
+ \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>INIT C' (C#Cs,False) \<leftarrow> e,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+| InitNonObject\<^sub>1:
+ "\<lbrakk> sh C = Some(sfs,Prepared);
+ C \<noteq> Object;
+ class P C = Some (D,r);
+ sh' = sh(C \<mapsto> (sfs,Processing));
+ P \<turnstile>\<^sub>1 \<langle>INIT C' (D#C#Cs,False) \<leftarrow> e,(h,l,sh')\<rangle> \<Rightarrow> \<langle>e',s'\<rangle> \<rbrakk>
+ \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>INIT C' (C#Cs,False) \<leftarrow> e,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+| InitRInit\<^sub>1:
+ "P \<turnstile>\<^sub>1 \<langle>RI (C,C\<bullet>\<^sub>sclinit([]));Cs \<leftarrow> e,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>
+ \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>INIT C' (C#Cs,True) \<leftarrow> e,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+
+| RInit\<^sub>1:
+ "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e,s\<rangle> \<Rightarrow> \<langle>Val v, (h',l',sh')\<rangle>;
+ sh' C = Some(sfs, i); sh'' = sh'(C \<mapsto> (sfs, Done));
+ C' = last(C#Cs);
+ P \<turnstile>\<^sub>1 \<langle>INIT C' (Cs,True) \<leftarrow> e', (h',l',sh'')\<rangle> \<Rightarrow> \<langle>e\<^sub>1,s\<^sub>1\<rangle> \<rbrakk>
+ \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>RI (C,e);Cs \<leftarrow> e',s\<rangle> \<Rightarrow> \<langle>e\<^sub>1,s\<^sub>1\<rangle>"
+| RInitInitFail\<^sub>1:
+ "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e,s\<rangle> \<Rightarrow> \<langle>throw a, (h',l',sh')\<rangle>;
+ sh' C = Some(sfs, i); sh'' = sh'(C \<mapsto> (sfs, Error));
+ P \<turnstile>\<^sub>1 \<langle>RI (D,throw a);Cs \<leftarrow> e', (h',l',sh'')\<rangle> \<Rightarrow> \<langle>e\<^sub>1,s\<^sub>1\<rangle> \<rbrakk>
+ \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>RI (C,e);D#Cs \<leftarrow> e',s\<rangle> \<Rightarrow> \<langle>e\<^sub>1,s\<^sub>1\<rangle>"
+| RInitFailFinal\<^sub>1:
+ "\<lbrakk> P \<turnstile>\<^sub>1 \<langle>e,s\<rangle> \<Rightarrow> \<langle>throw a, (h',l',sh')\<rangle>;
+ sh' C = Some(sfs, i); sh'' = sh'(C \<mapsto> (sfs, Error)) \<rbrakk>
+ \<Longrightarrow> P \<turnstile>\<^sub>1 \<langle>RI (C,e);Nil \<leftarrow> e',s\<rangle> \<Rightarrow> \<langle>throw a, (h',l',sh'')\<rangle>"
+
+
+(*<*)
+lemmas eval\<^sub>1_evals\<^sub>1_induct = eval\<^sub>1_evals\<^sub>1.induct [split_format (complete)]
+ and eval\<^sub>1_evals\<^sub>1_inducts = eval\<^sub>1_evals\<^sub>1.inducts [split_format (complete)]
+(*>*)
+
+
+inductive_cases eval\<^sub>1_cases [cases set]:
+ "P \<turnstile>\<^sub>1 \<langle>new C,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+ "P \<turnstile>\<^sub>1 \<langle>Cast C e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+ "P \<turnstile>\<^sub>1 \<langle>Val v,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+ "P \<turnstile>\<^sub>1 \<langle>e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+ "P \<turnstile>\<^sub>1 \<langle>Var v,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+ "P \<turnstile>\<^sub>1 \<langle>V:=e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+ "P \<turnstile>\<^sub>1 \<langle>e\<bullet>F{D},s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+ "P \<turnstile>\<^sub>1 \<langle>C\<bullet>\<^sub>sF{D},s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+ "P \<turnstile>\<^sub>1 \<langle>e\<^sub>1\<bullet>F{D}:=e\<^sub>2,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+ "P \<turnstile>\<^sub>1 \<langle>C\<bullet>\<^sub>sF{D}:=e\<^sub>2,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+ "P \<turnstile>\<^sub>1 \<langle>e\<bullet>M(es),s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+ "P \<turnstile>\<^sub>1 \<langle>C\<bullet>\<^sub>sM(es),s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+ "P \<turnstile>\<^sub>1 \<langle>{V:T;e\<^sub>1},s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+ "P \<turnstile>\<^sub>1 \<langle>e\<^sub>1;;e\<^sub>2,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+ "P \<turnstile>\<^sub>1 \<langle>if (e) e\<^sub>1 else e\<^sub>2,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+ "P \<turnstile>\<^sub>1 \<langle>while (b) c,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+ "P \<turnstile>\<^sub>1 \<langle>throw e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+ "P \<turnstile>\<^sub>1 \<langle>try e\<^sub>1 catch(C V) e\<^sub>2,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+ "P \<turnstile>\<^sub>1 \<langle>INIT C (Cs,b) \<leftarrow> e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+ "P \<turnstile>\<^sub>1 \<langle>RI (C,e);Cs \<leftarrow> e\<^sub>0,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+
+inductive_cases evals\<^sub>1_cases [cases set]:
+ "P \<turnstile>\<^sub>1 \<langle>[],s\<rangle> [\<Rightarrow>] \<langle>e',s'\<rangle>"
+ "P \<turnstile>\<^sub>1 \<langle>e#es,s\<rangle> [\<Rightarrow>] \<langle>e',s'\<rangle>"
+(*>*)
+
+
+lemma eval\<^sub>1_final: "P \<turnstile>\<^sub>1 \<langle>e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle> \<Longrightarrow> final e'"
+ and evals\<^sub>1_final: "P \<turnstile>\<^sub>1 \<langle>es,s\<rangle> [\<Rightarrow>] \<langle>es',s'\<rangle> \<Longrightarrow> finals es'"
+(*<*)by(induct rule:eval\<^sub>1_evals\<^sub>1.inducts, simp_all)(*>*)
+
+
+lemma eval\<^sub>1_final_same:
+assumes eval: "P \<turnstile>\<^sub>1 \<langle>e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>" shows "final e \<Longrightarrow> e = e' \<and> s = s'"
+(*<*)
+proof(erule finalE)
+ fix v assume Val: "e = Val v"
+ then show ?thesis using eval eval\<^sub>1_cases(3) by blast
+next
+ fix a assume "e = Throw a"
+ then show ?thesis using eval
+ by (metis eval\<^sub>1_cases(3,17) exp.distinct(101) exp.inject(3) val.distinct(13))
+qed
+(*>*)
+
+subsection "Property preservation"
+
+lemma eval\<^sub>1_preserves_len:
+ "P \<turnstile>\<^sub>1 \<langle>e\<^sub>0,(h\<^sub>0,ls\<^sub>0,sh\<^sub>0)\<rangle> \<Rightarrow> \<langle>e\<^sub>1,(h\<^sub>1,ls\<^sub>1,sh\<^sub>1)\<rangle> \<Longrightarrow> length ls\<^sub>0 = length ls\<^sub>1"
+and evals\<^sub>1_preserves_len:
+ "P \<turnstile>\<^sub>1 \<langle>es\<^sub>0,(h\<^sub>0,ls\<^sub>0,sh\<^sub>0)\<rangle> [\<Rightarrow>] \<langle>es\<^sub>1,(h\<^sub>1,ls\<^sub>1,sh\<^sub>1)\<rangle> \<Longrightarrow> length ls\<^sub>0 = length ls\<^sub>1"
+(*<*)by (induct rule:eval\<^sub>1_evals\<^sub>1_inducts, simp_all)(*>*)
+
+
+lemma evals\<^sub>1_preserves_elen:
+ "\<And>es' s s'. P \<turnstile>\<^sub>1 \<langle>es,s\<rangle> [\<Rightarrow>] \<langle>es',s'\<rangle> \<Longrightarrow> length es = length es'"
+(*<*)by(induct es type:list) (auto elim:evals\<^sub>1.cases)(*>*)
+
+
+lemma clinit\<^sub>1_loc_pres:
+ "P \<turnstile>\<^sub>1 \<langle>C\<^sub>0\<bullet>\<^sub>sclinit([]),(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',(h',l',sh')\<rangle> \<Longrightarrow> l = l'"
+ by(auto elim!: eval\<^sub>1_cases(12) evals\<^sub>1_cases(1))
+
+lemma
+shows init\<^sub>1_ri\<^sub>1_same_loc: "P \<turnstile>\<^sub>1 \<langle>e,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',(h',l',sh')\<rangle>
+ \<Longrightarrow> (\<And>C Cs b M a. e = INIT C (Cs,b) \<leftarrow> unit \<or> e = C\<bullet>\<^sub>sM([]) \<or> e = RI (C,Throw a) ; Cs \<leftarrow> unit
+ \<or> e = RI (C,C\<bullet>\<^sub>sclinit([])) ; Cs \<leftarrow> unit
+ \<Longrightarrow> l = l')"
+ and "P \<turnstile>\<^sub>1 \<langle>es,(h,l,sh)\<rangle> [\<Rightarrow>] \<langle>es',(h',l',sh')\<rangle> \<Longrightarrow> True"
+proof(induct rule: eval\<^sub>1_evals\<^sub>1_inducts)
+ case (RInitInitFail\<^sub>1 e h l sh a')
+ then show ?case using eval\<^sub>1_final[of _ _ _ "throw a'"]
+ by(fastforce dest: eval\<^sub>1_final_same[of _ "Throw a"])
+next
+ case RInitFailFinal\<^sub>1 then show ?case by(auto dest: eval\<^sub>1_final_same)
+qed(auto dest: evals\<^sub>1_cases eval\<^sub>1_cases(17) eval\<^sub>1_final_same)
+
+lemma init\<^sub>1_same_loc: "P \<turnstile>\<^sub>1 \<langle>INIT C (Cs,b) \<leftarrow> unit,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',(h',l',sh')\<rangle> \<Longrightarrow> l = l'"
+ by(simp add: init\<^sub>1_ri\<^sub>1_same_loc)
+
+
+theorem eval\<^sub>1_hext: "P \<turnstile>\<^sub>1 \<langle>e,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',(h',l',sh')\<rangle> \<Longrightarrow> h \<unlhd> h'"
+and evals\<^sub>1_hext: "P \<turnstile>\<^sub>1 \<langle>es,(h,l,sh)\<rangle> [\<Rightarrow>] \<langle>es',(h',l',sh')\<rangle> \<Longrightarrow> h \<unlhd> h'"
+(*<*)
+proof (induct rule: eval\<^sub>1_evals\<^sub>1_inducts)
+ case New\<^sub>1 thus ?case
+ by(fastforce intro!: hext_new intro:LeastI simp:new_Addr_def
+ split:if_split_asm simp del:fun_upd_apply)
+next
+ case NewInit\<^sub>1 thus ?case
+ by (meson hext_new hext_trans new_Addr_SomeD)
+next
+ case FAss\<^sub>1 thus ?case
+ by(auto simp:sym[THEN hext_upd_obj] simp del:fun_upd_apply
+ elim!: hext_trans)
+qed (auto elim!: hext_trans)
+(*>*)
+
+subsection "Initialization"
+
+lemma rinit\<^sub>1_throw:
+ "P\<^sub>1 \<turnstile>\<^sub>1 \<langle>RI (D,Throw xa) ; Cs \<leftarrow> e,(h, l, sh)\<rangle> \<Rightarrow> \<langle>e',(h', l', sh')\<rangle>
+ \<Longrightarrow> e' = Throw xa"
+proof(induct Cs arbitrary: D h l sh h' l' sh')
+ case Nil then show ?case
+ proof(rule eval\<^sub>1_cases(20)) qed(auto elim: eval\<^sub>1_cases)
+next
+ case (Cons C Cs) show ?case using Cons.prems
+ proof(induct rule: eval\<^sub>1_cases(20))
+ case RInit\<^sub>1: 1
+ then show ?case using Cons.hyps by(auto elim: eval\<^sub>1_cases)
+ next
+ case RInitInitFail\<^sub>1: 2
+ then show ?case using Cons.hyps eval\<^sub>1_final_same final_def by blast
+ next
+ case RInitFailFinal\<^sub>1: 3 then show ?case by simp
+ qed
+qed
+
+lemma rinit\<^sub>1_throwE:
+"P \<turnstile>\<^sub>1 \<langle>RI (C,throw e) ; Cs \<leftarrow> e\<^sub>0,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>
+ \<Longrightarrow> \<exists>a s\<^sub>t. e' = throw a \<and> P \<turnstile>\<^sub>1 \<langle>throw e,s\<rangle> \<Rightarrow> \<langle>throw a,s\<^sub>t\<rangle>"
+proof(induct Cs arbitrary: C e s)
+ case Nil
+ then show ?case
+ proof(rule eval\<^sub>1_cases(20)) \<comment> \<open> RI \<close>
+ fix v h' l' sh' assume "P \<turnstile>\<^sub>1 \<langle>throw e,s\<rangle> \<Rightarrow> \<langle>Val v,(h', l', sh')\<rangle>"
+ then show ?case using eval\<^sub>1_cases(17) by blast
+ qed(auto)
+next
+ case (Cons C' Cs')
+ show ?case using Cons.prems(1)
+ proof(rule eval\<^sub>1_cases(20)) \<comment> \<open> RI \<close>
+ fix v h' l' sh' assume "P \<turnstile>\<^sub>1 \<langle>throw e,s\<rangle> \<Rightarrow> \<langle>Val v,(h', l', sh')\<rangle>"
+ then show ?case using eval\<^sub>1_cases(17) by blast
+ next
+ fix a h' l' sh' sfs i D Cs''
+ assume e''step: "P \<turnstile>\<^sub>1 \<langle>throw e,s\<rangle> \<Rightarrow> \<langle>throw a,(h', l', sh')\<rangle>"
+ and shC: "sh' C = \<lfloor>(sfs, i)\<rfloor>"
+ and riD: "P \<turnstile>\<^sub>1 \<langle>RI (D,throw a) ; Cs'' \<leftarrow> e\<^sub>0,(h', l', sh'(C \<mapsto> (sfs, Error)))\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+ and "C' # Cs' = D # Cs''"
+ then show ?thesis using Cons.hyps eval\<^sub>1_final eval\<^sub>1_final_same by blast
+ qed(simp)
+qed
+
+end
diff --git a/thys/JinjaDCI/Compiler/J1WellForm.thy b/thys/JinjaDCI/Compiler/J1WellForm.thy
--- a/thys/JinjaDCI/Compiler/J1WellForm.thy
+++ b/thys/JinjaDCI/Compiler/J1WellForm.thy
@@ -1,519 +1,519 @@
-(* Title: JinjaDCI/Compiler/J1WellForm.thy
-
- Author: Tobias Nipkow, Susannah Mansky
- Copyright 2003 Technische Universitaet Muenchen, 2019-20 UIUC
-
- Based on the Jinja theory Compiler/J1WellForm.thy by Tobias Nipkow
-*)
-
-section \<open> Well-Formedness of Intermediate Language \<close>
-
-theory J1WellForm
-imports "../J/JWellForm" J1
-begin
-
-subsection "Well-Typedness"
-
-type_synonym
- env\<^sub>1 = "ty list" \<comment> \<open>type environment indexed by variable number\<close>
-
-inductive
- WT\<^sub>1 :: "[J\<^sub>1_prog,env\<^sub>1, expr\<^sub>1 , ty ] \<Rightarrow> bool"
- ("(_,_ \<turnstile>\<^sub>1/ _ :: _)" [51,51,51]50)
- and WTs\<^sub>1 :: "[J\<^sub>1_prog,env\<^sub>1, expr\<^sub>1 list, ty list] \<Rightarrow> bool"
- ("(_,_ \<turnstile>\<^sub>1/ _ [::] _)" [51,51,51]50)
- for P :: J\<^sub>1_prog
-where
-
- WTNew\<^sub>1:
- "is_class P C \<Longrightarrow>
- P,E \<turnstile>\<^sub>1 new C :: Class C"
-
-| WTCast\<^sub>1:
- "\<lbrakk> P,E \<turnstile>\<^sub>1 e :: Class D; is_class P C; P \<turnstile> C \<preceq>\<^sup>* D \<or> P \<turnstile> D \<preceq>\<^sup>* C \<rbrakk>
- \<Longrightarrow> P,E \<turnstile>\<^sub>1 Cast C e :: Class C"
-
-| WTVal\<^sub>1:
- "typeof v = Some T \<Longrightarrow>
- P,E \<turnstile>\<^sub>1 Val v :: T"
-
-| WTVar\<^sub>1:
- "\<lbrakk> E!i = T; i < size E \<rbrakk>
- \<Longrightarrow> P,E \<turnstile>\<^sub>1 Var i :: T"
-
-| WTBinOp\<^sub>1:
- "\<lbrakk> P,E \<turnstile>\<^sub>1 e\<^sub>1 :: T\<^sub>1; P,E \<turnstile>\<^sub>1 e\<^sub>2 :: T\<^sub>2;
- case bop of Eq \<Rightarrow> (P \<turnstile> T\<^sub>1 \<le> T\<^sub>2 \<or> P \<turnstile> T\<^sub>2 \<le> T\<^sub>1) \<and> T = Boolean
- | Add \<Rightarrow> T\<^sub>1 = Integer \<and> T\<^sub>2 = Integer \<and> T = Integer \<rbrakk>
- \<Longrightarrow> P,E \<turnstile>\<^sub>1 e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2 :: T"
-
-| WTLAss\<^sub>1:
- "\<lbrakk> E!i = T; i < size E; P,E \<turnstile>\<^sub>1 e :: T'; P \<turnstile> T' \<le> T \<rbrakk>
- \<Longrightarrow> P,E \<turnstile>\<^sub>1 i:=e :: Void"
-
-| WTFAcc\<^sub>1:
- "\<lbrakk> P,E \<turnstile>\<^sub>1 e :: Class C; P \<turnstile> C sees F,NonStatic:T in D \<rbrakk>
- \<Longrightarrow> P,E \<turnstile>\<^sub>1 e\<bullet>F{D} :: T"
-
-| WTSFAcc\<^sub>1:
- "\<lbrakk> P \<turnstile> C sees F,Static:T in D \<rbrakk>
- \<Longrightarrow> P,E \<turnstile>\<^sub>1 C\<bullet>\<^sub>sF{D} :: T"
-
-| WTFAss\<^sub>1:
- "\<lbrakk> P,E \<turnstile>\<^sub>1 e\<^sub>1 :: Class C; P \<turnstile> C sees F,NonStatic:T in D; P,E \<turnstile>\<^sub>1 e\<^sub>2 :: T'; P \<turnstile> T' \<le> T \<rbrakk>
- \<Longrightarrow> P,E \<turnstile>\<^sub>1 e\<^sub>1\<bullet>F{D} := e\<^sub>2 :: Void"
-
-| WTSFAss\<^sub>1:
- "\<lbrakk> P \<turnstile> C sees F,Static:T in D; P,E \<turnstile>\<^sub>1 e\<^sub>2 :: T'; P \<turnstile> T' \<le> T \<rbrakk>
- \<Longrightarrow> P,E \<turnstile>\<^sub>1 C\<bullet>\<^sub>sF{D}:=e\<^sub>2 :: Void"
-
-| WTCall\<^sub>1:
- "\<lbrakk> P,E \<turnstile>\<^sub>1 e :: Class C; P \<turnstile> C sees M,NonStatic:Ts' \<rightarrow> T = m in D;
- P,E \<turnstile>\<^sub>1 es [::] Ts; P \<turnstile> Ts [\<le>] Ts' \<rbrakk>
- \<Longrightarrow> P,E \<turnstile>\<^sub>1 e\<bullet>M(es) :: T"
-
-| WTSCall\<^sub>1:
- "\<lbrakk> P \<turnstile> C sees M,Static:Ts \<rightarrow> T = m in D;
- P,E \<turnstile>\<^sub>1 es [::] Ts'; P \<turnstile> Ts' [\<le>] Ts; M \<noteq> clinit \<rbrakk>
- \<Longrightarrow> P,E \<turnstile>\<^sub>1 C\<bullet>\<^sub>sM(es) :: T"
-
-| WTBlock\<^sub>1:
- "\<lbrakk> is_type P T; P,E@[T] \<turnstile>\<^sub>1 e::T' \<rbrakk>
- \<Longrightarrow> P,E \<turnstile>\<^sub>1 {i:T; e} :: T'"
-
-| WTSeq\<^sub>1:
- "\<lbrakk> P,E \<turnstile>\<^sub>1 e\<^sub>1::T\<^sub>1; P,E \<turnstile>\<^sub>1 e\<^sub>2::T\<^sub>2 \<rbrakk>
- \<Longrightarrow> P,E \<turnstile>\<^sub>1 e\<^sub>1;;e\<^sub>2 :: T\<^sub>2"
-
-| WTCond\<^sub>1:
- "\<lbrakk> P,E \<turnstile>\<^sub>1 e :: Boolean; P,E \<turnstile>\<^sub>1 e\<^sub>1::T\<^sub>1; P,E \<turnstile>\<^sub>1 e\<^sub>2::T\<^sub>2;
- P \<turnstile> T\<^sub>1 \<le> T\<^sub>2 \<or> P \<turnstile> T\<^sub>2 \<le> T\<^sub>1; P \<turnstile> T\<^sub>1 \<le> T\<^sub>2 \<longrightarrow> T = T\<^sub>2; P \<turnstile> T\<^sub>2 \<le> T\<^sub>1 \<longrightarrow> T = T\<^sub>1 \<rbrakk>
- \<Longrightarrow> P,E \<turnstile>\<^sub>1 if (e) e\<^sub>1 else e\<^sub>2 :: T"
-
-| WTWhile\<^sub>1:
- "\<lbrakk> P,E \<turnstile>\<^sub>1 e :: Boolean; P,E \<turnstile>\<^sub>1 c::T \<rbrakk>
- \<Longrightarrow> P,E \<turnstile>\<^sub>1 while (e) c :: Void"
-
-| WTThrow\<^sub>1:
- "P,E \<turnstile>\<^sub>1 e :: Class C \<Longrightarrow>
- P,E \<turnstile>\<^sub>1 throw e :: Void"
-
-| WTTry\<^sub>1:
- "\<lbrakk> P,E \<turnstile>\<^sub>1 e\<^sub>1 :: T; P,E@[Class C] \<turnstile>\<^sub>1 e\<^sub>2 :: T; is_class P C \<rbrakk>
- \<Longrightarrow> P,E \<turnstile>\<^sub>1 try e\<^sub>1 catch(C i) e\<^sub>2 :: T"
-
-| WTNil\<^sub>1:
- "P,E \<turnstile>\<^sub>1 [] [::] []"
-
-| WTCons\<^sub>1:
- "\<lbrakk> P,E \<turnstile>\<^sub>1 e :: T; P,E \<turnstile>\<^sub>1 es [::] Ts \<rbrakk>
- \<Longrightarrow> P,E \<turnstile>\<^sub>1 e#es [::] T#Ts"
-
-(*<*)
-declare WT\<^sub>1_WTs\<^sub>1.intros[intro!]
-declare WTNil\<^sub>1[iff]
-
-lemmas WT\<^sub>1_WTs\<^sub>1_induct = WT\<^sub>1_WTs\<^sub>1.induct [split_format (complete)]
- and WT\<^sub>1_WTs\<^sub>1_inducts = WT\<^sub>1_WTs\<^sub>1.inducts [split_format (complete)]
-
-inductive_cases eee[elim!]:
- "P,E \<turnstile>\<^sub>1 Val v :: T"
- "P,E \<turnstile>\<^sub>1 Var i :: T"
- "P,E \<turnstile>\<^sub>1 Cast D e :: T"
- "P,E \<turnstile>\<^sub>1 i:=e :: T"
- "P,E \<turnstile>\<^sub>1 {i:U; e} :: T"
- "P,E \<turnstile>\<^sub>1 e\<^sub>1;;e\<^sub>2 :: T"
- "P,E \<turnstile>\<^sub>1 if (e) e\<^sub>1 else e\<^sub>2 :: T"
- "P,E \<turnstile>\<^sub>1 while (e) c :: T"
- "P,E \<turnstile>\<^sub>1 throw e :: T"
- "P,E \<turnstile>\<^sub>1 try e\<^sub>1 catch(C i) e\<^sub>2 :: T"
- "P,E \<turnstile>\<^sub>1 e\<bullet>F{D} :: T"
- "P,E \<turnstile>\<^sub>1 C\<bullet>\<^sub>sF{D} :: T"
- "P,E \<turnstile>\<^sub>1 e\<^sub>1\<bullet>F{D}:=e\<^sub>2 :: T"
- "P,E \<turnstile>\<^sub>1 C\<bullet>\<^sub>sF{D}:=e\<^sub>2 :: T"
- "P,E \<turnstile>\<^sub>1 e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2 :: T"
- "P,E \<turnstile>\<^sub>1 new C :: T"
- "P,E \<turnstile>\<^sub>1 e\<bullet>M(es) :: T"
- "P,E \<turnstile>\<^sub>1 C\<bullet>\<^sub>sM(es) :: T"
- "P,E \<turnstile>\<^sub>1 [] [::] Ts"
- "P,E \<turnstile>\<^sub>1 e#es [::] Ts"
-(*>*)
-
-lemma init_nWT\<^sub>1 [simp]:"\<not>P,E \<turnstile>\<^sub>1 INIT C (Cs,b) \<leftarrow> e :: T"
- by(auto elim:WT\<^sub>1.cases)
-lemma rinit_nWT\<^sub>1 [simp]:"\<not>P,E \<turnstile>\<^sub>1 RI(C,e);Cs \<leftarrow> e' :: T"
- by(auto elim:WT\<^sub>1.cases)
-
-lemma WTs\<^sub>1_same_size: "\<And>Ts. P,E \<turnstile>\<^sub>1 es [::] Ts \<Longrightarrow> size es = size Ts"
-(*<*)by (induct es type:list) auto(*>*)
-
-
-lemma WT\<^sub>1_unique:
- "P,E \<turnstile>\<^sub>1 e :: T\<^sub>1 \<Longrightarrow> (\<And>T\<^sub>2. P,E \<turnstile>\<^sub>1 e :: T\<^sub>2 \<Longrightarrow> T\<^sub>1 = T\<^sub>2)" and
- WTs\<^sub>1_unique: "P,E \<turnstile>\<^sub>1 es [::] Ts\<^sub>1 \<Longrightarrow> (\<And>Ts\<^sub>2. P,E \<turnstile>\<^sub>1 es [::] Ts\<^sub>2 \<Longrightarrow> Ts\<^sub>1 = Ts\<^sub>2)"
-(*<*)
-proof(induct rule:WT\<^sub>1_WTs\<^sub>1.inducts)
- case WTVal\<^sub>1 then show ?case by clarsimp
-next
- case (WTBinOp\<^sub>1 E e\<^sub>1 T\<^sub>1 e\<^sub>2 T\<^sub>2 bop T)
- then show ?case by(case_tac bop) force+
-next
- case WTFAcc\<^sub>1 then show ?case
- by (blast dest:sees_field_idemp sees_field_fun)
-next
- case WTSFAcc\<^sub>1 then show ?case by (blast dest:sees_field_fun)
-next
- case WTSFAss\<^sub>1 then show ?case by (blast dest:sees_field_fun)
-next
- case WTCall\<^sub>1 then show ?case
- by (blast dest:sees_method_idemp sees_method_fun)
-next
- case WTSCall\<^sub>1 then show ?case by (blast dest:sees_method_fun)
-qed blast+
-(*>*)
-
-
-lemma assumes wf: "wf_prog p P"
-shows WT\<^sub>1_is_type: "P,E \<turnstile>\<^sub>1 e :: T \<Longrightarrow> set E \<subseteq> types P \<Longrightarrow> is_type P T"
-and "P,E \<turnstile>\<^sub>1 es [::] Ts \<Longrightarrow> True"
-(*<*)
-proof(induct rule:WT\<^sub>1_WTs\<^sub>1.inducts)
- case WTVal\<^sub>1 then show ?case by (simp add:typeof_lit_is_type)
-next
- case WTVar\<^sub>1 then show ?case by (blast intro:nth_mem)
-next
- case WTBinOp\<^sub>1 then show ?case by (simp split:bop.splits)
-next
- case WTFAcc\<^sub>1 then show ?case
- by (simp add:sees_field_is_type[OF _ wf])
-next
- case WTSFAcc\<^sub>1 then show ?case
- by (simp add:sees_field_is_type[OF _ wf])
-next
- case WTCall\<^sub>1 then show ?case
- by (fastforce dest!: sees_wf_mdecl[OF wf] simp:wf_mdecl_def)
-next
- case WTSCall\<^sub>1 then show ?case
- by (fastforce dest!: sees_wf_mdecl[OF wf] simp:wf_mdecl_def)
-next
- case WTCond\<^sub>1 then show ?case by blast
-qed simp+
-(*>*)
-
-lemma WT\<^sub>1_nsub_RI: "P,E \<turnstile>\<^sub>1 e :: T \<Longrightarrow> \<not>sub_RI e"
- and WTs\<^sub>1_nsub_RIs: "P,E \<turnstile>\<^sub>1 es [::] Ts \<Longrightarrow> \<not>sub_RIs es"
-proof(induct rule: WT\<^sub>1_WTs\<^sub>1.inducts) qed(simp_all)
-
-subsection\<open> Runtime Well-Typedness \<close>
-
-inductive
- WTrt\<^sub>1 :: "J\<^sub>1_prog \<Rightarrow> heap \<Rightarrow> sheap \<Rightarrow> env\<^sub>1 \<Rightarrow> expr\<^sub>1 \<Rightarrow> ty \<Rightarrow> bool"
- and WTrts\<^sub>1 :: "J\<^sub>1_prog \<Rightarrow> heap \<Rightarrow> sheap \<Rightarrow> env\<^sub>1 \<Rightarrow> expr\<^sub>1 list \<Rightarrow> ty list \<Rightarrow> bool"
- and WTrt2\<^sub>1 :: "[J\<^sub>1_prog,env\<^sub>1,heap,sheap,expr\<^sub>1,ty] \<Rightarrow> bool"
- ("_,_,_,_ \<turnstile>\<^sub>1 _ : _" [51,51,51,51]50)
- and WTrts2\<^sub>1 :: "[J\<^sub>1_prog,env\<^sub>1,heap,sheap,expr\<^sub>1 list, ty list] \<Rightarrow> bool"
- ("_,_,_,_ \<turnstile>\<^sub>1 _ [:] _" [51,51,51,51]50)
- for P :: J\<^sub>1_prog and h :: heap and sh :: sheap
-where
-
- "P,E,h,sh \<turnstile>\<^sub>1 e : T \<equiv> WTrt\<^sub>1 P h sh E e T"
-| "P,E,h,sh \<turnstile>\<^sub>1 es[:]Ts \<equiv> WTrts\<^sub>1 P h sh E es Ts"
-
-| WTrtNew\<^sub>1:
- "is_class P C \<Longrightarrow>
- P,E,h,sh \<turnstile>\<^sub>1 new C : Class C"
-
-| WTrtCast\<^sub>1:
- "\<lbrakk> P,E,h,sh \<turnstile>\<^sub>1 e : T; is_refT T; is_class P C \<rbrakk>
- \<Longrightarrow> P,E,h,sh \<turnstile>\<^sub>1 Cast C e : Class C"
-
-| WTrtVal\<^sub>1:
- "typeof\<^bsub>h\<^esub> v = Some T \<Longrightarrow>
- P,E,h,sh \<turnstile>\<^sub>1 Val v : T"
-
-| WTrtVar\<^sub>1:
- "\<lbrakk> E!i = T; i < size E \<rbrakk> \<Longrightarrow>
- P,E,h,sh \<turnstile>\<^sub>1 Var i : T"
-
-| WTrtBinOpEq\<^sub>1:
- "\<lbrakk> P,E,h,sh \<turnstile>\<^sub>1 e\<^sub>1 : T\<^sub>1; P,E,h,sh \<turnstile>\<^sub>1 e\<^sub>2 : T\<^sub>2 \<rbrakk>
- \<Longrightarrow> P,E,h,sh \<turnstile>\<^sub>1 e\<^sub>1 \<guillemotleft>Eq\<guillemotright> e\<^sub>2 : Boolean"
-
-| WTrtBinOpAdd\<^sub>1:
- "\<lbrakk> P,E,h,sh \<turnstile>\<^sub>1 e\<^sub>1 : Integer; P,E,h,sh \<turnstile>\<^sub>1 e\<^sub>2 : Integer \<rbrakk>
- \<Longrightarrow> P,E,h,sh \<turnstile>\<^sub>1 e\<^sub>1 \<guillemotleft>Add\<guillemotright> e\<^sub>2 : Integer"
-
-| WTrtLAss\<^sub>1:
- "\<lbrakk> E!i = T; i < size E; P,E,h,sh \<turnstile>\<^sub>1 e : T'; P \<turnstile> T' \<le> T \<rbrakk>
- \<Longrightarrow> P,E,h,sh \<turnstile>\<^sub>1 i:=e : Void"
-
-| WTrtFAcc\<^sub>1:
- "\<lbrakk> P,E,h,sh \<turnstile>\<^sub>1 e : Class C; P \<turnstile> C has F,NonStatic:T in D \<rbrakk> \<Longrightarrow>
- P,E,h,sh \<turnstile>\<^sub>1 e\<bullet>F{D} : T"
-
-| WTrtFAccNT\<^sub>1:
- "P,E,h,sh \<turnstile>\<^sub>1 e : NT \<Longrightarrow>
- P,E,h,sh \<turnstile>\<^sub>1 e\<bullet>F{D} : T"
-
-| WTrtSFAcc\<^sub>1:
- "\<lbrakk> P \<turnstile> C has F,Static:T in D \<rbrakk> \<Longrightarrow>
- P,E,h,sh \<turnstile>\<^sub>1 C\<bullet>\<^sub>sF{D} : T"
-
-| WTrtFAss\<^sub>1:
- "\<lbrakk> P,E,h,sh \<turnstile>\<^sub>1 e\<^sub>1 : Class C; P \<turnstile> C has F,NonStatic:T in D; P,E,h,sh \<turnstile>\<^sub>1 e\<^sub>2 : T\<^sub>2; P \<turnstile> T\<^sub>2 \<le> T \<rbrakk>
- \<Longrightarrow> P,E,h,sh \<turnstile>\<^sub>1 e\<^sub>1\<bullet>F{D}:=e\<^sub>2 : Void"
-
-| WTrtFAssNT\<^sub>1:
- "\<lbrakk> P,E,h,sh \<turnstile>\<^sub>1 e\<^sub>1:NT; P,E,h,sh \<turnstile>\<^sub>1 e\<^sub>2 : T\<^sub>2 \<rbrakk>
- \<Longrightarrow> P,E,h,sh \<turnstile>\<^sub>1 e\<^sub>1\<bullet>F{D}:=e\<^sub>2 : Void"
-
-| WTrtSFAss\<^sub>1:
- "\<lbrakk> P,E,h,sh \<turnstile>\<^sub>1 e\<^sub>2 : T\<^sub>2; P \<turnstile> C has F,Static:T in D; P \<turnstile> T\<^sub>2 \<le> T \<rbrakk>
- \<Longrightarrow> P,E,h,sh \<turnstile>\<^sub>1 C\<bullet>\<^sub>sF{D}:=e\<^sub>2 : Void"
-
-| WTrtCall\<^sub>1:
- "\<lbrakk> P,E,h,sh \<turnstile>\<^sub>1 e : Class C; P \<turnstile> C sees M,NonStatic:Ts \<rightarrow> T = m in D;
- P,E,h,sh \<turnstile>\<^sub>1 es [:] Ts'; P \<turnstile> Ts' [\<le>] Ts \<rbrakk>
- \<Longrightarrow> P,E,h,sh \<turnstile>\<^sub>1 e\<bullet>M(es) : T"
-
-| WTrtCallNT\<^sub>1:
- "\<lbrakk> P,E,h,sh \<turnstile>\<^sub>1 e : NT; P,E,h,sh \<turnstile>\<^sub>1 es [:] Ts \<rbrakk>
- \<Longrightarrow> P,E,h,sh \<turnstile>\<^sub>1 e\<bullet>M(es) : T"
-
-| WTrtSCall\<^sub>1:
- "\<lbrakk> P \<turnstile> C sees M,Static:Ts \<rightarrow> T = m in D;
- P,E,h,sh \<turnstile>\<^sub>1 es [:] Ts'; P \<turnstile> Ts' [\<le>] Ts;
- M = clinit \<longrightarrow> sh D = \<lfloor>(sfs,Processing)\<rfloor> \<and> es = map Val vs \<rbrakk>
- \<Longrightarrow> P,E,h,sh \<turnstile>\<^sub>1 C\<bullet>\<^sub>sM(es) : T"
-
-| WTrtBlock\<^sub>1:
- "P,E@[T],h,sh \<turnstile>\<^sub>1 e : T' \<Longrightarrow>
- P,E,h,sh \<turnstile>\<^sub>1 {i:T; e} : T'"
-
-| WTrtSeq\<^sub>1:
- "\<lbrakk> P,E,h,sh \<turnstile>\<^sub>1 e\<^sub>1:T\<^sub>1; P,E,h,sh \<turnstile>\<^sub>1 e\<^sub>2:T\<^sub>2 \<rbrakk>
- \<Longrightarrow> P,E,h,sh \<turnstile>\<^sub>1 e\<^sub>1;;e\<^sub>2 : T\<^sub>2"
-
-| WTrtCond\<^sub>1:
- "\<lbrakk> P,E,h,sh \<turnstile>\<^sub>1 e : Boolean; P,E,h,sh \<turnstile>\<^sub>1 e\<^sub>1:T\<^sub>1; P,E,h,sh \<turnstile>\<^sub>1 e\<^sub>2:T\<^sub>2;
- P \<turnstile> T\<^sub>1 \<le> T\<^sub>2 \<or> P \<turnstile> T\<^sub>2 \<le> T\<^sub>1; P \<turnstile> T\<^sub>1 \<le> T\<^sub>2 \<longrightarrow> T = T\<^sub>2; P \<turnstile> T\<^sub>2 \<le> T\<^sub>1 \<longrightarrow> T = T\<^sub>1 \<rbrakk>
- \<Longrightarrow> P,E,h,sh \<turnstile>\<^sub>1 if (e) e\<^sub>1 else e\<^sub>2 : T"
-
-| WTrtWhile\<^sub>1:
- "\<lbrakk> P,E,h,sh \<turnstile>\<^sub>1 e : Boolean; P,E,h,sh \<turnstile>\<^sub>1 c:T \<rbrakk>
- \<Longrightarrow> P,E,h,sh \<turnstile>\<^sub>1 while(e) c : Void"
-
-| WTrtThrow\<^sub>1:
- "\<lbrakk> P,E,h,sh \<turnstile>\<^sub>1 e : T\<^sub>r; is_refT T\<^sub>r \<rbrakk> \<Longrightarrow>
- P,E,h,sh \<turnstile>\<^sub>1 throw e : T"
-
-| WTrtTry\<^sub>1:
- "\<lbrakk> P,E,h,sh \<turnstile>\<^sub>1 e\<^sub>1 : T\<^sub>1; P,E@[Class C],h,sh \<turnstile>\<^sub>1 e\<^sub>2 : T\<^sub>2; P \<turnstile> T\<^sub>1 \<le> T\<^sub>2 \<rbrakk>
- \<Longrightarrow> P,E,h,sh \<turnstile>\<^sub>1 try e\<^sub>1 catch(C i) e\<^sub>2 : T\<^sub>2"
-
-| WTrtInit\<^sub>1:
- "\<lbrakk> P,E,h,sh \<turnstile>\<^sub>1 e : T; \<forall>C' \<in> set (C#Cs). is_class P C'; \<not>sub_RI e;
- \<forall>C' \<in> set (tl Cs). \<exists>sfs. sh C' = \<lfloor>(sfs,Processing)\<rfloor>;
- b \<longrightarrow> (\<forall>C' \<in> set Cs. \<exists>sfs. sh C' = \<lfloor>(sfs,Processing)\<rfloor>);
- distinct Cs; supercls_lst P Cs \<rbrakk>
- \<Longrightarrow> P,E,h,sh \<turnstile>\<^sub>1 INIT C (Cs, b) \<leftarrow> e : T"
-
-| WTrtRI\<^sub>1:
- "\<lbrakk> P,E,h,sh \<turnstile>\<^sub>1 e : T; P,E,h,sh \<turnstile>\<^sub>1 e' : T'; \<forall>C' \<in> set (C#Cs). is_class P C'; \<not>sub_RI e';
- \<forall>C' \<in> set (C#Cs). not_init C' e;
- \<forall>C' \<in> set Cs. \<exists>sfs. sh C' = \<lfloor>(sfs,Processing)\<rfloor>;
- \<exists>sfs. sh C = \<lfloor>(sfs, Processing)\<rfloor> \<or> (sh C = \<lfloor>(sfs, Error)\<rfloor> \<and> e = THROW NoClassDefFoundError);
- distinct (C#Cs); supercls_lst P (C#Cs) \<rbrakk>
- \<Longrightarrow> P,E,h,sh \<turnstile>\<^sub>1 RI(C, e);Cs \<leftarrow> e' : T'"
-
-\<comment> \<open>well-typed expression lists\<close>
-
-| WTrtNil\<^sub>1:
- "P,E,h,sh \<turnstile>\<^sub>1 [] [:] []"
-
-| WTrtCons\<^sub>1:
- "\<lbrakk> P,E,h,sh \<turnstile>\<^sub>1 e : T; P,E,h,sh \<turnstile>\<^sub>1 es [:] Ts \<rbrakk>
- \<Longrightarrow> P,E,h,sh \<turnstile>\<^sub>1 e#es [:] T#Ts"
-
-(*<*)
-declare WTrt\<^sub>1_WTrts\<^sub>1.intros[intro!] WTrtNil\<^sub>1[iff]
-declare
- WTrtFAcc\<^sub>1[rule del] WTrtFAccNT\<^sub>1[rule del] WTrtSFAcc\<^sub>1[rule del]
- WTrtFAss\<^sub>1[rule del] WTrtFAssNT\<^sub>1[rule del] WTrtSFAss\<^sub>1[rule del]
- WTrtCall\<^sub>1[rule del] WTrtCallNT\<^sub>1[rule del] WTrtSCall\<^sub>1[rule del]
-
-lemmas WTrt\<^sub>1_induct = WTrt\<^sub>1_WTrts\<^sub>1.induct [split_format (complete)]
- and WTrt\<^sub>1_inducts = WTrt\<^sub>1_WTrts\<^sub>1.inducts [split_format (complete)]
-(*>*)
-
-(*<*)
-inductive_cases WTrt\<^sub>1_elim_cases[elim!]:
- "P,E,h,sh \<turnstile>\<^sub>1 Val v : T"
- "P,E,h,sh \<turnstile>\<^sub>1 Var i : T"
- "P,E,h,sh \<turnstile>\<^sub>1 v :=e : T"
- "P,E,h,sh \<turnstile>\<^sub>1 {i:U; e} : T"
- "P,E,h,sh \<turnstile>\<^sub>1 e\<^sub>1;;e\<^sub>2 : T\<^sub>2"
- "P,E,h,sh \<turnstile>\<^sub>1 if (e) e\<^sub>1 else e\<^sub>2 : T"
- "P,E,h,sh \<turnstile>\<^sub>1 while(e) c : T"
- "P,E,h,sh \<turnstile>\<^sub>1 throw e : T"
- "P,E,h,sh \<turnstile>\<^sub>1 try e\<^sub>1 catch(C V) e\<^sub>2 : T"
- "P,E,h,sh \<turnstile>\<^sub>1 Cast D e : T"
- "P,E,h,sh \<turnstile>\<^sub>1 e\<bullet>F{D} : T"
- "P,E,h,sh \<turnstile>\<^sub>1 C\<bullet>\<^sub>sF{D} : T"
- "P,E,h,sh \<turnstile>\<^sub>1 e\<bullet>F{D} := v : T"
- "P,E,h,sh \<turnstile>\<^sub>1 C\<bullet>\<^sub>sF{D} := v : T"
- "P,E,h,sh \<turnstile>\<^sub>1 e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2 : T"
- "P,E,h,sh \<turnstile>\<^sub>1 new C : T"
- "P,E,h,sh \<turnstile>\<^sub>1 e\<bullet>M{D}(es) : T"
- "P,E,h,sh \<turnstile>\<^sub>1 C\<bullet>\<^sub>sM{D}(es) : T"
- "P,E,h,sh \<turnstile>\<^sub>1 INIT C (Cs,b) \<leftarrow> e : T"
- "P,E,h,sh \<turnstile>\<^sub>1 RI(C,e);Cs \<leftarrow> e' : T"
- "P,E,h,sh \<turnstile>\<^sub>1 [] [:] Ts"
- "P,E,h,sh \<turnstile>\<^sub>1 e#es [:] Ts"
-(*>*)
-
-lemma WT\<^sub>1_implies_WTrt\<^sub>1: "P,E \<turnstile>\<^sub>1 e :: T \<Longrightarrow> P,E,h,sh \<turnstile>\<^sub>1 e : T"
-and WTs\<^sub>1_implies_WTrts\<^sub>1: "P,E \<turnstile>\<^sub>1 es [::] Ts \<Longrightarrow> P,E,h,sh \<turnstile>\<^sub>1 es [:] Ts"
-(*<*)
-proof(induct rule: WT\<^sub>1_WTs\<^sub>1_inducts)
- case WTVal\<^sub>1 then show ?case by (fastforce dest:typeof_lit_typeof)
-next
- case (WTBinOp\<^sub>1 E e\<^sub>1 T\<^sub>1 e\<^sub>2 T\<^sub>2 bop T)
- then show ?case by (case_tac bop) fastforce+
-next
- case WTFAcc\<^sub>1 then show ?case
- by (fastforce simp: WTrtFAcc\<^sub>1 has_visible_field)
-next
- case WTSFAcc\<^sub>1 then show ?case
- by (fastforce simp: WTrtSFAcc\<^sub>1 has_visible_field)
-next
- case WTFAss\<^sub>1 then show ?case by (meson WTrtFAss\<^sub>1 has_visible_field)
-next
- case WTSFAss\<^sub>1 then show ?case by (meson WTrtSFAss\<^sub>1 has_visible_field)
-next
- case WTCall\<^sub>1 then show ?case by (fastforce simp: WTrtCall\<^sub>1)
-next
- case WTSCall\<^sub>1 then show ?case by (fastforce simp: WTrtSCall\<^sub>1)
-qed fastforce+
-(*>*)
-
-subsection\<open> Well-formedness\<close>
-
-\<comment> \<open>Indices in blocks increase by 1\<close>
-
-primrec \<B> :: "expr\<^sub>1 \<Rightarrow> nat \<Rightarrow> bool"
- and \<B>s :: "expr\<^sub>1 list \<Rightarrow> nat \<Rightarrow> bool" where
-"\<B> (new C) i = True" |
-"\<B> (Cast C e) i = \<B> e i" |
-"\<B> (Val v) i = True" |
-"\<B> (e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2) i = (\<B> e\<^sub>1 i \<and> \<B> e\<^sub>2 i)" |
-"\<B> (Var j) i = True" |
-"\<B> (e\<bullet>F{D}) i = \<B> e i" |
-"\<B> (C\<bullet>\<^sub>sF{D}) i = True" |
-"\<B> (j:=e) i = \<B> e i" |
-"\<B> (e\<^sub>1\<bullet>F{D} := e\<^sub>2) i = (\<B> e\<^sub>1 i \<and> \<B> e\<^sub>2 i)" |
-"\<B> (C\<bullet>\<^sub>sF{D} := e\<^sub>2) i = \<B> e\<^sub>2 i" |
-"\<B> (e\<bullet>M(es)) i = (\<B> e i \<and> \<B>s es i)" |
-"\<B> (C\<bullet>\<^sub>sM(es)) i = \<B>s es i" |
-"\<B> ({j:T ; e}) i = (i = j \<and> \<B> e (i+1))" |
-"\<B> (e\<^sub>1;;e\<^sub>2) i = (\<B> e\<^sub>1 i \<and> \<B> e\<^sub>2 i)" |
-"\<B> (if (e) e\<^sub>1 else e\<^sub>2) i = (\<B> e i \<and> \<B> e\<^sub>1 i \<and> \<B> e\<^sub>2 i)" |
-"\<B> (throw e) i = \<B> e i" |
-"\<B> (while (e) c) i = (\<B> e i \<and> \<B> c i)" |
-"\<B> (try e\<^sub>1 catch(C j) e\<^sub>2) i = (\<B> e\<^sub>1 i \<and> i=j \<and> \<B> e\<^sub>2 (i+1))" |
-"\<B> (INIT C (Cs,b) \<leftarrow> e) i = \<B> e i" |
-"\<B> (RI(C,e);Cs \<leftarrow> e') i = (\<B> e i \<and> \<B> e' i)" |
-
-"\<B>s [] i = True" |
-"\<B>s (e#es) i = (\<B> e i \<and> \<B>s es i)"
-
-
-definition wf_J\<^sub>1_mdecl :: "J\<^sub>1_prog \<Rightarrow> cname \<Rightarrow> expr\<^sub>1 mdecl \<Rightarrow> bool"
-where
- "wf_J\<^sub>1_mdecl P C \<equiv> \<lambda>(M,b,Ts,T,body).
- \<not>sub_RI body \<and>
- (case b of
- NonStatic \<Rightarrow>
- (\<exists>T'. P,Class C#Ts \<turnstile>\<^sub>1 body :: T' \<and> P \<turnstile> T' \<le> T) \<and>
- \<D> body \<lfloor>{..size Ts}\<rfloor> \<and> \<B> body (size Ts + 1)
- | Static \<Rightarrow> (\<exists>T'. P,Ts \<turnstile>\<^sub>1 body :: T' \<and> P \<turnstile> T' \<le> T) \<and>
- \<D> body \<lfloor>{..<size Ts}\<rfloor> \<and> \<B> body (size Ts))"
-
-lemma wf_J\<^sub>1_mdecl_NonStatic[simp]:
- "wf_J\<^sub>1_mdecl P C (M,NonStatic,Ts,T,body) \<equiv>
- (\<not>sub_RI body \<and>
- (\<exists>T'. P,Class C#Ts \<turnstile>\<^sub>1 body :: T' \<and> P \<turnstile> T' \<le> T) \<and>
- \<D> body \<lfloor>{..size Ts}\<rfloor> \<and> \<B> body (size Ts + 1))"
-(*<*)by (simp add:wf_J\<^sub>1_mdecl_def)(*>*)
-
-lemma wf_J\<^sub>1_mdecl_Static[simp]:
- "wf_J\<^sub>1_mdecl P C (M,Static,Ts,T,body) \<equiv>
- (\<not>sub_RI body \<and>
- (\<exists>T'. P,Ts \<turnstile>\<^sub>1 body :: T' \<and> P \<turnstile> T' \<le> T) \<and>
- \<D> body \<lfloor>{..<size Ts}\<rfloor> \<and> \<B> body (size Ts))"
-(*<*)by (simp add:wf_J\<^sub>1_mdecl_def)(*>*)
-
-abbreviation "wf_J\<^sub>1_prog == wf_prog wf_J\<^sub>1_mdecl"
-
-lemma sees_wf\<^sub>1_nsub_RI:
-assumes wf: "wf_J\<^sub>1_prog P" and cM: "P \<turnstile> C sees M,b : Ts\<rightarrow>T = body in D"
-shows "\<not>sub_RI body"
-using sees_wf_mdecl[OF wf cM] by(simp add: wf_J\<^sub>1_mdecl_def wf_mdecl_def)
-
-lemma wf\<^sub>1_types_clinit:
-assumes wf:"wf_prog wf_md P" and ex: "class P C = Some a" and proc: "sh C = \<lfloor>(sfs, Processing)\<rfloor>"
-shows "P,E,h,sh \<turnstile>\<^sub>1 C\<bullet>\<^sub>sclinit([]) : Void"
-proof -
- from ex obtain D fs ms where "a = (D,fs,ms)" by(cases a)
- then have sP: "(C, D, fs, ms) \<in> set P" using ex map_of_SomeD[of P C a] by(simp add: class_def)
- then have "wf_clinit ms" using assms by(unfold wf_prog_def wf_cdecl_def, auto)
- then obtain m where sm: "(clinit, Static, [], Void, m) \<in> set ms"
- by(unfold wf_clinit_def) auto
- then have "P \<turnstile> C sees clinit,Static:[] \<rightarrow> Void = m in C"
- using mdecl_visible[OF wf sP sm] by simp
- then show ?thesis using WTrtSCall\<^sub>1 proc by blast
-qed
-
-
-lemma assumes wf: "wf_J\<^sub>1_prog P"
-shows eval\<^sub>1_proc_pres: "P \<turnstile>\<^sub>1 \<langle>e,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',(h',l',sh')\<rangle>
- \<Longrightarrow> not_init C e \<Longrightarrow> \<exists>sfs. sh C = \<lfloor>(sfs, Processing)\<rfloor> \<Longrightarrow> \<exists>sfs'. sh' C = \<lfloor>(sfs', Processing)\<rfloor>"
- and evals\<^sub>1_proc_pres: "P \<turnstile>\<^sub>1 \<langle>es,(h,l,sh)\<rangle> [\<Rightarrow>] \<langle>es',(h',l',sh')\<rangle>
- \<Longrightarrow> not_inits C es \<Longrightarrow> \<exists>sfs. sh C = \<lfloor>(sfs, Processing)\<rfloor> \<Longrightarrow> \<exists>sfs'. sh' C = \<lfloor>(sfs', Processing)\<rfloor>"
-(*<*)
-proof(induct rule:eval\<^sub>1_evals\<^sub>1_inducts)
- case Call\<^sub>1 then show ?case using sees_wf\<^sub>1_nsub_RI[OF wf Call\<^sub>1.hyps(6)] nsub_RI_not_init by auto
-next
- case (SCallInit\<^sub>1 ps h l sh vs h\<^sub>1 l\<^sub>1 sh\<^sub>1 C' M Ts T body D v' h\<^sub>2 l\<^sub>2 sh\<^sub>2 l\<^sub>2' e' h\<^sub>3 l\<^sub>3 sh\<^sub>3)
- then show ?case
- using SCallInit\<^sub>1 sees_wf\<^sub>1_nsub_RI[OF wf SCallInit\<^sub>1.hyps(3)] nsub_RI_not_init[of body] by auto
-next
- case SCall\<^sub>1 then show ?case using sees_wf\<^sub>1_nsub_RI[OF wf SCall\<^sub>1.hyps(3)] nsub_RI_not_init by auto
-next
- case (InitNone\<^sub>1 sh C1 C' Cs h l e' a a b) then show ?case by(cases "C = C1") auto
-next
- case (InitDone\<^sub>1 sh C sfs C' Cs h l e' a a b) then show ?case by(cases Cs, auto)
-next
- case (InitProcessing\<^sub>1 sh C sfs C' Cs h l e' a a b) then show ?case by(cases Cs, auto)
-next
- case (InitError\<^sub>1 sh C1 sfs Cs h l e' a a b C') then show ?case by(cases "C = C1") auto
-next
- case (InitObject\<^sub>1 sh C1 sfs sh' C' Cs h l e' a a b) then show ?case by(cases "C = C1") auto
-next
- case (InitNonObject\<^sub>1 sh C1 sfs D a b sh' C' Cs h l e' a a b)
- then show ?case by(cases "C = C1") auto
-next
- case (RInit\<^sub>1 e a a b v h' l' sh' C sfs i sh'' C' Cs e\<^sub>1 a a b) then show ?case by(cases Cs, auto)
-next
- case (RInitInitFail\<^sub>1 e h l sh a h' l' sh' C1 sfs i sh'' D Cs e\<^sub>1 h1 l1 sh1)
- then show ?case using eval\<^sub>1_final by fastforce
-qed(auto)
-(*>*)
-
-lemma clinit\<^sub>1_proc_pres:
- "\<lbrakk> wf_J\<^sub>1_prog P; P \<turnstile>\<^sub>1 \<langle>C\<^sub>0\<bullet>\<^sub>sclinit([]),(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',(h',l',sh')\<rangle>;
- sh C' = \<lfloor>(sfs,Processing)\<rfloor> \<rbrakk>
- \<Longrightarrow> \<exists>sfs. sh' C' = \<lfloor>(sfs,Processing)\<rfloor>"
- by(auto dest: eval\<^sub>1_proc_pres)
-
-end
+(* Title: JinjaDCI/Compiler/J1WellForm.thy
+
+ Author: Tobias Nipkow, Susannah Mansky
+ Copyright 2003 Technische Universitaet Muenchen, 2019-20 UIUC
+
+ Based on the Jinja theory Compiler/J1WellForm.thy by Tobias Nipkow
+*)
+
+section \<open> Well-Formedness of Intermediate Language \<close>
+
+theory J1WellForm
+imports "../J/JWellForm" J1
+begin
+
+subsection "Well-Typedness"
+
+type_synonym
+ env\<^sub>1 = "ty list" \<comment> \<open>type environment indexed by variable number\<close>
+
+inductive
+ WT\<^sub>1 :: "[J\<^sub>1_prog,env\<^sub>1, expr\<^sub>1 , ty ] \<Rightarrow> bool"
+ ("(_,_ \<turnstile>\<^sub>1/ _ :: _)" [51,51,51]50)
+ and WTs\<^sub>1 :: "[J\<^sub>1_prog,env\<^sub>1, expr\<^sub>1 list, ty list] \<Rightarrow> bool"
+ ("(_,_ \<turnstile>\<^sub>1/ _ [::] _)" [51,51,51]50)
+ for P :: J\<^sub>1_prog
+where
+
+ WTNew\<^sub>1:
+ "is_class P C \<Longrightarrow>
+ P,E \<turnstile>\<^sub>1 new C :: Class C"
+
+| WTCast\<^sub>1:
+ "\<lbrakk> P,E \<turnstile>\<^sub>1 e :: Class D; is_class P C; P \<turnstile> C \<preceq>\<^sup>* D \<or> P \<turnstile> D \<preceq>\<^sup>* C \<rbrakk>
+ \<Longrightarrow> P,E \<turnstile>\<^sub>1 Cast C e :: Class C"
+
+| WTVal\<^sub>1:
+ "typeof v = Some T \<Longrightarrow>
+ P,E \<turnstile>\<^sub>1 Val v :: T"
+
+| WTVar\<^sub>1:
+ "\<lbrakk> E!i = T; i < size E \<rbrakk>
+ \<Longrightarrow> P,E \<turnstile>\<^sub>1 Var i :: T"
+
+| WTBinOp\<^sub>1:
+ "\<lbrakk> P,E \<turnstile>\<^sub>1 e\<^sub>1 :: T\<^sub>1; P,E \<turnstile>\<^sub>1 e\<^sub>2 :: T\<^sub>2;
+ case bop of Eq \<Rightarrow> (P \<turnstile> T\<^sub>1 \<le> T\<^sub>2 \<or> P \<turnstile> T\<^sub>2 \<le> T\<^sub>1) \<and> T = Boolean
+ | Add \<Rightarrow> T\<^sub>1 = Integer \<and> T\<^sub>2 = Integer \<and> T = Integer \<rbrakk>
+ \<Longrightarrow> P,E \<turnstile>\<^sub>1 e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2 :: T"
+
+| WTLAss\<^sub>1:
+ "\<lbrakk> E!i = T; i < size E; P,E \<turnstile>\<^sub>1 e :: T'; P \<turnstile> T' \<le> T \<rbrakk>
+ \<Longrightarrow> P,E \<turnstile>\<^sub>1 i:=e :: Void"
+
+| WTFAcc\<^sub>1:
+ "\<lbrakk> P,E \<turnstile>\<^sub>1 e :: Class C; P \<turnstile> C sees F,NonStatic:T in D \<rbrakk>
+ \<Longrightarrow> P,E \<turnstile>\<^sub>1 e\<bullet>F{D} :: T"
+
+| WTSFAcc\<^sub>1:
+ "\<lbrakk> P \<turnstile> C sees F,Static:T in D \<rbrakk>
+ \<Longrightarrow> P,E \<turnstile>\<^sub>1 C\<bullet>\<^sub>sF{D} :: T"
+
+| WTFAss\<^sub>1:
+ "\<lbrakk> P,E \<turnstile>\<^sub>1 e\<^sub>1 :: Class C; P \<turnstile> C sees F,NonStatic:T in D; P,E \<turnstile>\<^sub>1 e\<^sub>2 :: T'; P \<turnstile> T' \<le> T \<rbrakk>
+ \<Longrightarrow> P,E \<turnstile>\<^sub>1 e\<^sub>1\<bullet>F{D} := e\<^sub>2 :: Void"
+
+| WTSFAss\<^sub>1:
+ "\<lbrakk> P \<turnstile> C sees F,Static:T in D; P,E \<turnstile>\<^sub>1 e\<^sub>2 :: T'; P \<turnstile> T' \<le> T \<rbrakk>
+ \<Longrightarrow> P,E \<turnstile>\<^sub>1 C\<bullet>\<^sub>sF{D}:=e\<^sub>2 :: Void"
+
+| WTCall\<^sub>1:
+ "\<lbrakk> P,E \<turnstile>\<^sub>1 e :: Class C; P \<turnstile> C sees M,NonStatic:Ts' \<rightarrow> T = m in D;
+ P,E \<turnstile>\<^sub>1 es [::] Ts; P \<turnstile> Ts [\<le>] Ts' \<rbrakk>
+ \<Longrightarrow> P,E \<turnstile>\<^sub>1 e\<bullet>M(es) :: T"
+
+| WTSCall\<^sub>1:
+ "\<lbrakk> P \<turnstile> C sees M,Static:Ts \<rightarrow> T = m in D;
+ P,E \<turnstile>\<^sub>1 es [::] Ts'; P \<turnstile> Ts' [\<le>] Ts; M \<noteq> clinit \<rbrakk>
+ \<Longrightarrow> P,E \<turnstile>\<^sub>1 C\<bullet>\<^sub>sM(es) :: T"
+
+| WTBlock\<^sub>1:
+ "\<lbrakk> is_type P T; P,E@[T] \<turnstile>\<^sub>1 e::T' \<rbrakk>
+ \<Longrightarrow> P,E \<turnstile>\<^sub>1 {i:T; e} :: T'"
+
+| WTSeq\<^sub>1:
+ "\<lbrakk> P,E \<turnstile>\<^sub>1 e\<^sub>1::T\<^sub>1; P,E \<turnstile>\<^sub>1 e\<^sub>2::T\<^sub>2 \<rbrakk>
+ \<Longrightarrow> P,E \<turnstile>\<^sub>1 e\<^sub>1;;e\<^sub>2 :: T\<^sub>2"
+
+| WTCond\<^sub>1:
+ "\<lbrakk> P,E \<turnstile>\<^sub>1 e :: Boolean; P,E \<turnstile>\<^sub>1 e\<^sub>1::T\<^sub>1; P,E \<turnstile>\<^sub>1 e\<^sub>2::T\<^sub>2;
+ P \<turnstile> T\<^sub>1 \<le> T\<^sub>2 \<or> P \<turnstile> T\<^sub>2 \<le> T\<^sub>1; P \<turnstile> T\<^sub>1 \<le> T\<^sub>2 \<longrightarrow> T = T\<^sub>2; P \<turnstile> T\<^sub>2 \<le> T\<^sub>1 \<longrightarrow> T = T\<^sub>1 \<rbrakk>
+ \<Longrightarrow> P,E \<turnstile>\<^sub>1 if (e) e\<^sub>1 else e\<^sub>2 :: T"
+
+| WTWhile\<^sub>1:
+ "\<lbrakk> P,E \<turnstile>\<^sub>1 e :: Boolean; P,E \<turnstile>\<^sub>1 c::T \<rbrakk>
+ \<Longrightarrow> P,E \<turnstile>\<^sub>1 while (e) c :: Void"
+
+| WTThrow\<^sub>1:
+ "P,E \<turnstile>\<^sub>1 e :: Class C \<Longrightarrow>
+ P,E \<turnstile>\<^sub>1 throw e :: Void"
+
+| WTTry\<^sub>1:
+ "\<lbrakk> P,E \<turnstile>\<^sub>1 e\<^sub>1 :: T; P,E@[Class C] \<turnstile>\<^sub>1 e\<^sub>2 :: T; is_class P C \<rbrakk>
+ \<Longrightarrow> P,E \<turnstile>\<^sub>1 try e\<^sub>1 catch(C i) e\<^sub>2 :: T"
+
+| WTNil\<^sub>1:
+ "P,E \<turnstile>\<^sub>1 [] [::] []"
+
+| WTCons\<^sub>1:
+ "\<lbrakk> P,E \<turnstile>\<^sub>1 e :: T; P,E \<turnstile>\<^sub>1 es [::] Ts \<rbrakk>
+ \<Longrightarrow> P,E \<turnstile>\<^sub>1 e#es [::] T#Ts"
+
+(*<*)
+declare WT\<^sub>1_WTs\<^sub>1.intros[intro!]
+declare WTNil\<^sub>1[iff]
+
+lemmas WT\<^sub>1_WTs\<^sub>1_induct = WT\<^sub>1_WTs\<^sub>1.induct [split_format (complete)]
+ and WT\<^sub>1_WTs\<^sub>1_inducts = WT\<^sub>1_WTs\<^sub>1.inducts [split_format (complete)]
+
+inductive_cases eee[elim!]:
+ "P,E \<turnstile>\<^sub>1 Val v :: T"
+ "P,E \<turnstile>\<^sub>1 Var i :: T"
+ "P,E \<turnstile>\<^sub>1 Cast D e :: T"
+ "P,E \<turnstile>\<^sub>1 i:=e :: T"
+ "P,E \<turnstile>\<^sub>1 {i:U; e} :: T"
+ "P,E \<turnstile>\<^sub>1 e\<^sub>1;;e\<^sub>2 :: T"
+ "P,E \<turnstile>\<^sub>1 if (e) e\<^sub>1 else e\<^sub>2 :: T"
+ "P,E \<turnstile>\<^sub>1 while (e) c :: T"
+ "P,E \<turnstile>\<^sub>1 throw e :: T"
+ "P,E \<turnstile>\<^sub>1 try e\<^sub>1 catch(C i) e\<^sub>2 :: T"
+ "P,E \<turnstile>\<^sub>1 e\<bullet>F{D} :: T"
+ "P,E \<turnstile>\<^sub>1 C\<bullet>\<^sub>sF{D} :: T"
+ "P,E \<turnstile>\<^sub>1 e\<^sub>1\<bullet>F{D}:=e\<^sub>2 :: T"
+ "P,E \<turnstile>\<^sub>1 C\<bullet>\<^sub>sF{D}:=e\<^sub>2 :: T"
+ "P,E \<turnstile>\<^sub>1 e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2 :: T"
+ "P,E \<turnstile>\<^sub>1 new C :: T"
+ "P,E \<turnstile>\<^sub>1 e\<bullet>M(es) :: T"
+ "P,E \<turnstile>\<^sub>1 C\<bullet>\<^sub>sM(es) :: T"
+ "P,E \<turnstile>\<^sub>1 [] [::] Ts"
+ "P,E \<turnstile>\<^sub>1 e#es [::] Ts"
+(*>*)
+
+lemma init_nWT\<^sub>1 [simp]:"\<not>P,E \<turnstile>\<^sub>1 INIT C (Cs,b) \<leftarrow> e :: T"
+ by(auto elim:WT\<^sub>1.cases)
+lemma rinit_nWT\<^sub>1 [simp]:"\<not>P,E \<turnstile>\<^sub>1 RI(C,e);Cs \<leftarrow> e' :: T"
+ by(auto elim:WT\<^sub>1.cases)
+
+lemma WTs\<^sub>1_same_size: "\<And>Ts. P,E \<turnstile>\<^sub>1 es [::] Ts \<Longrightarrow> size es = size Ts"
+(*<*)by (induct es type:list) auto(*>*)
+
+
+lemma WT\<^sub>1_unique:
+ "P,E \<turnstile>\<^sub>1 e :: T\<^sub>1 \<Longrightarrow> (\<And>T\<^sub>2. P,E \<turnstile>\<^sub>1 e :: T\<^sub>2 \<Longrightarrow> T\<^sub>1 = T\<^sub>2)" and
+ WTs\<^sub>1_unique: "P,E \<turnstile>\<^sub>1 es [::] Ts\<^sub>1 \<Longrightarrow> (\<And>Ts\<^sub>2. P,E \<turnstile>\<^sub>1 es [::] Ts\<^sub>2 \<Longrightarrow> Ts\<^sub>1 = Ts\<^sub>2)"
+(*<*)
+proof(induct rule:WT\<^sub>1_WTs\<^sub>1.inducts)
+ case WTVal\<^sub>1 then show ?case by clarsimp
+next
+ case (WTBinOp\<^sub>1 E e\<^sub>1 T\<^sub>1 e\<^sub>2 T\<^sub>2 bop T)
+ then show ?case by(case_tac bop) force+
+next
+ case WTFAcc\<^sub>1 then show ?case
+ by (blast dest:sees_field_idemp sees_field_fun)
+next
+ case WTSFAcc\<^sub>1 then show ?case by (blast dest:sees_field_fun)
+next
+ case WTSFAss\<^sub>1 then show ?case by (blast dest:sees_field_fun)
+next
+ case WTCall\<^sub>1 then show ?case
+ by (blast dest:sees_method_idemp sees_method_fun)
+next
+ case WTSCall\<^sub>1 then show ?case by (blast dest:sees_method_fun)
+qed blast+
+(*>*)
+
+
+lemma assumes wf: "wf_prog p P"
+shows WT\<^sub>1_is_type: "P,E \<turnstile>\<^sub>1 e :: T \<Longrightarrow> set E \<subseteq> types P \<Longrightarrow> is_type P T"
+and "P,E \<turnstile>\<^sub>1 es [::] Ts \<Longrightarrow> True"
+(*<*)
+proof(induct rule:WT\<^sub>1_WTs\<^sub>1.inducts)
+ case WTVal\<^sub>1 then show ?case by (simp add:typeof_lit_is_type)
+next
+ case WTVar\<^sub>1 then show ?case by (blast intro:nth_mem)
+next
+ case WTBinOp\<^sub>1 then show ?case by (simp split:bop.splits)
+next
+ case WTFAcc\<^sub>1 then show ?case
+ by (simp add:sees_field_is_type[OF _ wf])
+next
+ case WTSFAcc\<^sub>1 then show ?case
+ by (simp add:sees_field_is_type[OF _ wf])
+next
+ case WTCall\<^sub>1 then show ?case
+ by (fastforce dest!: sees_wf_mdecl[OF wf] simp:wf_mdecl_def)
+next
+ case WTSCall\<^sub>1 then show ?case
+ by (fastforce dest!: sees_wf_mdecl[OF wf] simp:wf_mdecl_def)
+next
+ case WTCond\<^sub>1 then show ?case by blast
+qed simp+
+(*>*)
+
+lemma WT\<^sub>1_nsub_RI: "P,E \<turnstile>\<^sub>1 e :: T \<Longrightarrow> \<not>sub_RI e"
+ and WTs\<^sub>1_nsub_RIs: "P,E \<turnstile>\<^sub>1 es [::] Ts \<Longrightarrow> \<not>sub_RIs es"
+proof(induct rule: WT\<^sub>1_WTs\<^sub>1.inducts) qed(simp_all)
+
+subsection\<open> Runtime Well-Typedness \<close>
+
+inductive
+ WTrt\<^sub>1 :: "J\<^sub>1_prog \<Rightarrow> heap \<Rightarrow> sheap \<Rightarrow> env\<^sub>1 \<Rightarrow> expr\<^sub>1 \<Rightarrow> ty \<Rightarrow> bool"
+ and WTrts\<^sub>1 :: "J\<^sub>1_prog \<Rightarrow> heap \<Rightarrow> sheap \<Rightarrow> env\<^sub>1 \<Rightarrow> expr\<^sub>1 list \<Rightarrow> ty list \<Rightarrow> bool"
+ and WTrt2\<^sub>1 :: "[J\<^sub>1_prog,env\<^sub>1,heap,sheap,expr\<^sub>1,ty] \<Rightarrow> bool"
+ ("_,_,_,_ \<turnstile>\<^sub>1 _ : _" [51,51,51,51]50)
+ and WTrts2\<^sub>1 :: "[J\<^sub>1_prog,env\<^sub>1,heap,sheap,expr\<^sub>1 list, ty list] \<Rightarrow> bool"
+ ("_,_,_,_ \<turnstile>\<^sub>1 _ [:] _" [51,51,51,51]50)
+ for P :: J\<^sub>1_prog and h :: heap and sh :: sheap
+where
+
+ "P,E,h,sh \<turnstile>\<^sub>1 e : T \<equiv> WTrt\<^sub>1 P h sh E e T"
+| "P,E,h,sh \<turnstile>\<^sub>1 es[:]Ts \<equiv> WTrts\<^sub>1 P h sh E es Ts"
+
+| WTrtNew\<^sub>1:
+ "is_class P C \<Longrightarrow>
+ P,E,h,sh \<turnstile>\<^sub>1 new C : Class C"
+
+| WTrtCast\<^sub>1:
+ "\<lbrakk> P,E,h,sh \<turnstile>\<^sub>1 e : T; is_refT T; is_class P C \<rbrakk>
+ \<Longrightarrow> P,E,h,sh \<turnstile>\<^sub>1 Cast C e : Class C"
+
+| WTrtVal\<^sub>1:
+ "typeof\<^bsub>h\<^esub> v = Some T \<Longrightarrow>
+ P,E,h,sh \<turnstile>\<^sub>1 Val v : T"
+
+| WTrtVar\<^sub>1:
+ "\<lbrakk> E!i = T; i < size E \<rbrakk> \<Longrightarrow>
+ P,E,h,sh \<turnstile>\<^sub>1 Var i : T"
+
+| WTrtBinOpEq\<^sub>1:
+ "\<lbrakk> P,E,h,sh \<turnstile>\<^sub>1 e\<^sub>1 : T\<^sub>1; P,E,h,sh \<turnstile>\<^sub>1 e\<^sub>2 : T\<^sub>2 \<rbrakk>
+ \<Longrightarrow> P,E,h,sh \<turnstile>\<^sub>1 e\<^sub>1 \<guillemotleft>Eq\<guillemotright> e\<^sub>2 : Boolean"
+
+| WTrtBinOpAdd\<^sub>1:
+ "\<lbrakk> P,E,h,sh \<turnstile>\<^sub>1 e\<^sub>1 : Integer; P,E,h,sh \<turnstile>\<^sub>1 e\<^sub>2 : Integer \<rbrakk>
+ \<Longrightarrow> P,E,h,sh \<turnstile>\<^sub>1 e\<^sub>1 \<guillemotleft>Add\<guillemotright> e\<^sub>2 : Integer"
+
+| WTrtLAss\<^sub>1:
+ "\<lbrakk> E!i = T; i < size E; P,E,h,sh \<turnstile>\<^sub>1 e : T'; P \<turnstile> T' \<le> T \<rbrakk>
+ \<Longrightarrow> P,E,h,sh \<turnstile>\<^sub>1 i:=e : Void"
+
+| WTrtFAcc\<^sub>1:
+ "\<lbrakk> P,E,h,sh \<turnstile>\<^sub>1 e : Class C; P \<turnstile> C has F,NonStatic:T in D \<rbrakk> \<Longrightarrow>
+ P,E,h,sh \<turnstile>\<^sub>1 e\<bullet>F{D} : T"
+
+| WTrtFAccNT\<^sub>1:
+ "P,E,h,sh \<turnstile>\<^sub>1 e : NT \<Longrightarrow>
+ P,E,h,sh \<turnstile>\<^sub>1 e\<bullet>F{D} : T"
+
+| WTrtSFAcc\<^sub>1:
+ "\<lbrakk> P \<turnstile> C has F,Static:T in D \<rbrakk> \<Longrightarrow>
+ P,E,h,sh \<turnstile>\<^sub>1 C\<bullet>\<^sub>sF{D} : T"
+
+| WTrtFAss\<^sub>1:
+ "\<lbrakk> P,E,h,sh \<turnstile>\<^sub>1 e\<^sub>1 : Class C; P \<turnstile> C has F,NonStatic:T in D; P,E,h,sh \<turnstile>\<^sub>1 e\<^sub>2 : T\<^sub>2; P \<turnstile> T\<^sub>2 \<le> T \<rbrakk>
+ \<Longrightarrow> P,E,h,sh \<turnstile>\<^sub>1 e\<^sub>1\<bullet>F{D}:=e\<^sub>2 : Void"
+
+| WTrtFAssNT\<^sub>1:
+ "\<lbrakk> P,E,h,sh \<turnstile>\<^sub>1 e\<^sub>1:NT; P,E,h,sh \<turnstile>\<^sub>1 e\<^sub>2 : T\<^sub>2 \<rbrakk>
+ \<Longrightarrow> P,E,h,sh \<turnstile>\<^sub>1 e\<^sub>1\<bullet>F{D}:=e\<^sub>2 : Void"
+
+| WTrtSFAss\<^sub>1:
+ "\<lbrakk> P,E,h,sh \<turnstile>\<^sub>1 e\<^sub>2 : T\<^sub>2; P \<turnstile> C has F,Static:T in D; P \<turnstile> T\<^sub>2 \<le> T \<rbrakk>
+ \<Longrightarrow> P,E,h,sh \<turnstile>\<^sub>1 C\<bullet>\<^sub>sF{D}:=e\<^sub>2 : Void"
+
+| WTrtCall\<^sub>1:
+ "\<lbrakk> P,E,h,sh \<turnstile>\<^sub>1 e : Class C; P \<turnstile> C sees M,NonStatic:Ts \<rightarrow> T = m in D;
+ P,E,h,sh \<turnstile>\<^sub>1 es [:] Ts'; P \<turnstile> Ts' [\<le>] Ts \<rbrakk>
+ \<Longrightarrow> P,E,h,sh \<turnstile>\<^sub>1 e\<bullet>M(es) : T"
+
+| WTrtCallNT\<^sub>1:
+ "\<lbrakk> P,E,h,sh \<turnstile>\<^sub>1 e : NT; P,E,h,sh \<turnstile>\<^sub>1 es [:] Ts \<rbrakk>
+ \<Longrightarrow> P,E,h,sh \<turnstile>\<^sub>1 e\<bullet>M(es) : T"
+
+| WTrtSCall\<^sub>1:
+ "\<lbrakk> P \<turnstile> C sees M,Static:Ts \<rightarrow> T = m in D;
+ P,E,h,sh \<turnstile>\<^sub>1 es [:] Ts'; P \<turnstile> Ts' [\<le>] Ts;
+ M = clinit \<longrightarrow> sh D = \<lfloor>(sfs,Processing)\<rfloor> \<and> es = map Val vs \<rbrakk>
+ \<Longrightarrow> P,E,h,sh \<turnstile>\<^sub>1 C\<bullet>\<^sub>sM(es) : T"
+
+| WTrtBlock\<^sub>1:
+ "P,E@[T],h,sh \<turnstile>\<^sub>1 e : T' \<Longrightarrow>
+ P,E,h,sh \<turnstile>\<^sub>1 {i:T; e} : T'"
+
+| WTrtSeq\<^sub>1:
+ "\<lbrakk> P,E,h,sh \<turnstile>\<^sub>1 e\<^sub>1:T\<^sub>1; P,E,h,sh \<turnstile>\<^sub>1 e\<^sub>2:T\<^sub>2 \<rbrakk>
+ \<Longrightarrow> P,E,h,sh \<turnstile>\<^sub>1 e\<^sub>1;;e\<^sub>2 : T\<^sub>2"
+
+| WTrtCond\<^sub>1:
+ "\<lbrakk> P,E,h,sh \<turnstile>\<^sub>1 e : Boolean; P,E,h,sh \<turnstile>\<^sub>1 e\<^sub>1:T\<^sub>1; P,E,h,sh \<turnstile>\<^sub>1 e\<^sub>2:T\<^sub>2;
+ P \<turnstile> T\<^sub>1 \<le> T\<^sub>2 \<or> P \<turnstile> T\<^sub>2 \<le> T\<^sub>1; P \<turnstile> T\<^sub>1 \<le> T\<^sub>2 \<longrightarrow> T = T\<^sub>2; P \<turnstile> T\<^sub>2 \<le> T\<^sub>1 \<longrightarrow> T = T\<^sub>1 \<rbrakk>
+ \<Longrightarrow> P,E,h,sh \<turnstile>\<^sub>1 if (e) e\<^sub>1 else e\<^sub>2 : T"
+
+| WTrtWhile\<^sub>1:
+ "\<lbrakk> P,E,h,sh \<turnstile>\<^sub>1 e : Boolean; P,E,h,sh \<turnstile>\<^sub>1 c:T \<rbrakk>
+ \<Longrightarrow> P,E,h,sh \<turnstile>\<^sub>1 while(e) c : Void"
+
+| WTrtThrow\<^sub>1:
+ "\<lbrakk> P,E,h,sh \<turnstile>\<^sub>1 e : T\<^sub>r; is_refT T\<^sub>r \<rbrakk> \<Longrightarrow>
+ P,E,h,sh \<turnstile>\<^sub>1 throw e : T"
+
+| WTrtTry\<^sub>1:
+ "\<lbrakk> P,E,h,sh \<turnstile>\<^sub>1 e\<^sub>1 : T\<^sub>1; P,E@[Class C],h,sh \<turnstile>\<^sub>1 e\<^sub>2 : T\<^sub>2; P \<turnstile> T\<^sub>1 \<le> T\<^sub>2 \<rbrakk>
+ \<Longrightarrow> P,E,h,sh \<turnstile>\<^sub>1 try e\<^sub>1 catch(C i) e\<^sub>2 : T\<^sub>2"
+
+| WTrtInit\<^sub>1:
+ "\<lbrakk> P,E,h,sh \<turnstile>\<^sub>1 e : T; \<forall>C' \<in> set (C#Cs). is_class P C'; \<not>sub_RI e;
+ \<forall>C' \<in> set (tl Cs). \<exists>sfs. sh C' = \<lfloor>(sfs,Processing)\<rfloor>;
+ b \<longrightarrow> (\<forall>C' \<in> set Cs. \<exists>sfs. sh C' = \<lfloor>(sfs,Processing)\<rfloor>);
+ distinct Cs; supercls_lst P Cs \<rbrakk>
+ \<Longrightarrow> P,E,h,sh \<turnstile>\<^sub>1 INIT C (Cs, b) \<leftarrow> e : T"
+
+| WTrtRI\<^sub>1:
+ "\<lbrakk> P,E,h,sh \<turnstile>\<^sub>1 e : T; P,E,h,sh \<turnstile>\<^sub>1 e' : T'; \<forall>C' \<in> set (C#Cs). is_class P C'; \<not>sub_RI e';
+ \<forall>C' \<in> set (C#Cs). not_init C' e;
+ \<forall>C' \<in> set Cs. \<exists>sfs. sh C' = \<lfloor>(sfs,Processing)\<rfloor>;
+ \<exists>sfs. sh C = \<lfloor>(sfs, Processing)\<rfloor> \<or> (sh C = \<lfloor>(sfs, Error)\<rfloor> \<and> e = THROW NoClassDefFoundError);
+ distinct (C#Cs); supercls_lst P (C#Cs) \<rbrakk>
+ \<Longrightarrow> P,E,h,sh \<turnstile>\<^sub>1 RI(C, e);Cs \<leftarrow> e' : T'"
+
+\<comment> \<open>well-typed expression lists\<close>
+
+| WTrtNil\<^sub>1:
+ "P,E,h,sh \<turnstile>\<^sub>1 [] [:] []"
+
+| WTrtCons\<^sub>1:
+ "\<lbrakk> P,E,h,sh \<turnstile>\<^sub>1 e : T; P,E,h,sh \<turnstile>\<^sub>1 es [:] Ts \<rbrakk>
+ \<Longrightarrow> P,E,h,sh \<turnstile>\<^sub>1 e#es [:] T#Ts"
+
+(*<*)
+declare WTrt\<^sub>1_WTrts\<^sub>1.intros[intro!] WTrtNil\<^sub>1[iff]
+declare
+ WTrtFAcc\<^sub>1[rule del] WTrtFAccNT\<^sub>1[rule del] WTrtSFAcc\<^sub>1[rule del]
+ WTrtFAss\<^sub>1[rule del] WTrtFAssNT\<^sub>1[rule del] WTrtSFAss\<^sub>1[rule del]
+ WTrtCall\<^sub>1[rule del] WTrtCallNT\<^sub>1[rule del] WTrtSCall\<^sub>1[rule del]
+
+lemmas WTrt\<^sub>1_induct = WTrt\<^sub>1_WTrts\<^sub>1.induct [split_format (complete)]
+ and WTrt\<^sub>1_inducts = WTrt\<^sub>1_WTrts\<^sub>1.inducts [split_format (complete)]
+(*>*)
+
+(*<*)
+inductive_cases WTrt\<^sub>1_elim_cases[elim!]:
+ "P,E,h,sh \<turnstile>\<^sub>1 Val v : T"
+ "P,E,h,sh \<turnstile>\<^sub>1 Var i : T"
+ "P,E,h,sh \<turnstile>\<^sub>1 v :=e : T"
+ "P,E,h,sh \<turnstile>\<^sub>1 {i:U; e} : T"
+ "P,E,h,sh \<turnstile>\<^sub>1 e\<^sub>1;;e\<^sub>2 : T\<^sub>2"
+ "P,E,h,sh \<turnstile>\<^sub>1 if (e) e\<^sub>1 else e\<^sub>2 : T"
+ "P,E,h,sh \<turnstile>\<^sub>1 while(e) c : T"
+ "P,E,h,sh \<turnstile>\<^sub>1 throw e : T"
+ "P,E,h,sh \<turnstile>\<^sub>1 try e\<^sub>1 catch(C V) e\<^sub>2 : T"
+ "P,E,h,sh \<turnstile>\<^sub>1 Cast D e : T"
+ "P,E,h,sh \<turnstile>\<^sub>1 e\<bullet>F{D} : T"
+ "P,E,h,sh \<turnstile>\<^sub>1 C\<bullet>\<^sub>sF{D} : T"
+ "P,E,h,sh \<turnstile>\<^sub>1 e\<bullet>F{D} := v : T"
+ "P,E,h,sh \<turnstile>\<^sub>1 C\<bullet>\<^sub>sF{D} := v : T"
+ "P,E,h,sh \<turnstile>\<^sub>1 e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2 : T"
+ "P,E,h,sh \<turnstile>\<^sub>1 new C : T"
+ "P,E,h,sh \<turnstile>\<^sub>1 e\<bullet>M{D}(es) : T"
+ "P,E,h,sh \<turnstile>\<^sub>1 C\<bullet>\<^sub>sM{D}(es) : T"
+ "P,E,h,sh \<turnstile>\<^sub>1 INIT C (Cs,b) \<leftarrow> e : T"
+ "P,E,h,sh \<turnstile>\<^sub>1 RI(C,e);Cs \<leftarrow> e' : T"
+ "P,E,h,sh \<turnstile>\<^sub>1 [] [:] Ts"
+ "P,E,h,sh \<turnstile>\<^sub>1 e#es [:] Ts"
+(*>*)
+
+lemma WT\<^sub>1_implies_WTrt\<^sub>1: "P,E \<turnstile>\<^sub>1 e :: T \<Longrightarrow> P,E,h,sh \<turnstile>\<^sub>1 e : T"
+and WTs\<^sub>1_implies_WTrts\<^sub>1: "P,E \<turnstile>\<^sub>1 es [::] Ts \<Longrightarrow> P,E,h,sh \<turnstile>\<^sub>1 es [:] Ts"
+(*<*)
+proof(induct rule: WT\<^sub>1_WTs\<^sub>1_inducts)
+ case WTVal\<^sub>1 then show ?case by (fastforce dest:typeof_lit_typeof)
+next
+ case (WTBinOp\<^sub>1 E e\<^sub>1 T\<^sub>1 e\<^sub>2 T\<^sub>2 bop T)
+ then show ?case by (case_tac bop) fastforce+
+next
+ case WTFAcc\<^sub>1 then show ?case
+ by (fastforce simp: WTrtFAcc\<^sub>1 has_visible_field)
+next
+ case WTSFAcc\<^sub>1 then show ?case
+ by (fastforce simp: WTrtSFAcc\<^sub>1 has_visible_field)
+next
+ case WTFAss\<^sub>1 then show ?case by (meson WTrtFAss\<^sub>1 has_visible_field)
+next
+ case WTSFAss\<^sub>1 then show ?case by (meson WTrtSFAss\<^sub>1 has_visible_field)
+next
+ case WTCall\<^sub>1 then show ?case by (fastforce simp: WTrtCall\<^sub>1)
+next
+ case WTSCall\<^sub>1 then show ?case by (fastforce simp: WTrtSCall\<^sub>1)
+qed fastforce+
+(*>*)
+
+subsection\<open> Well-formedness\<close>
+
+\<comment> \<open>Indices in blocks increase by 1\<close>
+
+primrec \<B> :: "expr\<^sub>1 \<Rightarrow> nat \<Rightarrow> bool"
+ and \<B>s :: "expr\<^sub>1 list \<Rightarrow> nat \<Rightarrow> bool" where
+"\<B> (new C) i = True" |
+"\<B> (Cast C e) i = \<B> e i" |
+"\<B> (Val v) i = True" |
+"\<B> (e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2) i = (\<B> e\<^sub>1 i \<and> \<B> e\<^sub>2 i)" |
+"\<B> (Var j) i = True" |
+"\<B> (e\<bullet>F{D}) i = \<B> e i" |
+"\<B> (C\<bullet>\<^sub>sF{D}) i = True" |
+"\<B> (j:=e) i = \<B> e i" |
+"\<B> (e\<^sub>1\<bullet>F{D} := e\<^sub>2) i = (\<B> e\<^sub>1 i \<and> \<B> e\<^sub>2 i)" |
+"\<B> (C\<bullet>\<^sub>sF{D} := e\<^sub>2) i = \<B> e\<^sub>2 i" |
+"\<B> (e\<bullet>M(es)) i = (\<B> e i \<and> \<B>s es i)" |
+"\<B> (C\<bullet>\<^sub>sM(es)) i = \<B>s es i" |
+"\<B> ({j:T ; e}) i = (i = j \<and> \<B> e (i+1))" |
+"\<B> (e\<^sub>1;;e\<^sub>2) i = (\<B> e\<^sub>1 i \<and> \<B> e\<^sub>2 i)" |
+"\<B> (if (e) e\<^sub>1 else e\<^sub>2) i = (\<B> e i \<and> \<B> e\<^sub>1 i \<and> \<B> e\<^sub>2 i)" |
+"\<B> (throw e) i = \<B> e i" |
+"\<B> (while (e) c) i = (\<B> e i \<and> \<B> c i)" |
+"\<B> (try e\<^sub>1 catch(C j) e\<^sub>2) i = (\<B> e\<^sub>1 i \<and> i=j \<and> \<B> e\<^sub>2 (i+1))" |
+"\<B> (INIT C (Cs,b) \<leftarrow> e) i = \<B> e i" |
+"\<B> (RI(C,e);Cs \<leftarrow> e') i = (\<B> e i \<and> \<B> e' i)" |
+
+"\<B>s [] i = True" |
+"\<B>s (e#es) i = (\<B> e i \<and> \<B>s es i)"
+
+
+definition wf_J\<^sub>1_mdecl :: "J\<^sub>1_prog \<Rightarrow> cname \<Rightarrow> expr\<^sub>1 mdecl \<Rightarrow> bool"
+where
+ "wf_J\<^sub>1_mdecl P C \<equiv> \<lambda>(M,b,Ts,T,body).
+ \<not>sub_RI body \<and>
+ (case b of
+ NonStatic \<Rightarrow>
+ (\<exists>T'. P,Class C#Ts \<turnstile>\<^sub>1 body :: T' \<and> P \<turnstile> T' \<le> T) \<and>
+ \<D> body \<lfloor>{..size Ts}\<rfloor> \<and> \<B> body (size Ts + 1)
+ | Static \<Rightarrow> (\<exists>T'. P,Ts \<turnstile>\<^sub>1 body :: T' \<and> P \<turnstile> T' \<le> T) \<and>
+ \<D> body \<lfloor>{..<size Ts}\<rfloor> \<and> \<B> body (size Ts))"
+
+lemma wf_J\<^sub>1_mdecl_NonStatic[simp]:
+ "wf_J\<^sub>1_mdecl P C (M,NonStatic,Ts,T,body) \<equiv>
+ (\<not>sub_RI body \<and>
+ (\<exists>T'. P,Class C#Ts \<turnstile>\<^sub>1 body :: T' \<and> P \<turnstile> T' \<le> T) \<and>
+ \<D> body \<lfloor>{..size Ts}\<rfloor> \<and> \<B> body (size Ts + 1))"
+(*<*)by (simp add:wf_J\<^sub>1_mdecl_def)(*>*)
+
+lemma wf_J\<^sub>1_mdecl_Static[simp]:
+ "wf_J\<^sub>1_mdecl P C (M,Static,Ts,T,body) \<equiv>
+ (\<not>sub_RI body \<and>
+ (\<exists>T'. P,Ts \<turnstile>\<^sub>1 body :: T' \<and> P \<turnstile> T' \<le> T) \<and>
+ \<D> body \<lfloor>{..<size Ts}\<rfloor> \<and> \<B> body (size Ts))"
+(*<*)by (simp add:wf_J\<^sub>1_mdecl_def)(*>*)
+
+abbreviation "wf_J\<^sub>1_prog == wf_prog wf_J\<^sub>1_mdecl"
+
+lemma sees_wf\<^sub>1_nsub_RI:
+assumes wf: "wf_J\<^sub>1_prog P" and cM: "P \<turnstile> C sees M,b : Ts\<rightarrow>T = body in D"
+shows "\<not>sub_RI body"
+using sees_wf_mdecl[OF wf cM] by(simp add: wf_J\<^sub>1_mdecl_def wf_mdecl_def)
+
+lemma wf\<^sub>1_types_clinit:
+assumes wf:"wf_prog wf_md P" and ex: "class P C = Some a" and proc: "sh C = \<lfloor>(sfs, Processing)\<rfloor>"
+shows "P,E,h,sh \<turnstile>\<^sub>1 C\<bullet>\<^sub>sclinit([]) : Void"
+proof -
+ from ex obtain D fs ms where "a = (D,fs,ms)" by(cases a)
+ then have sP: "(C, D, fs, ms) \<in> set P" using ex map_of_SomeD[of P C a] by(simp add: class_def)
+ then have "wf_clinit ms" using assms by(unfold wf_prog_def wf_cdecl_def, auto)
+ then obtain m where sm: "(clinit, Static, [], Void, m) \<in> set ms"
+ by(unfold wf_clinit_def) auto
+ then have "P \<turnstile> C sees clinit,Static:[] \<rightarrow> Void = m in C"
+ using mdecl_visible[OF wf sP sm] by simp
+ then show ?thesis using WTrtSCall\<^sub>1 proc by blast
+qed
+
+
+lemma assumes wf: "wf_J\<^sub>1_prog P"
+shows eval\<^sub>1_proc_pres: "P \<turnstile>\<^sub>1 \<langle>e,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',(h',l',sh')\<rangle>
+ \<Longrightarrow> not_init C e \<Longrightarrow> \<exists>sfs. sh C = \<lfloor>(sfs, Processing)\<rfloor> \<Longrightarrow> \<exists>sfs'. sh' C = \<lfloor>(sfs', Processing)\<rfloor>"
+ and evals\<^sub>1_proc_pres: "P \<turnstile>\<^sub>1 \<langle>es,(h,l,sh)\<rangle> [\<Rightarrow>] \<langle>es',(h',l',sh')\<rangle>
+ \<Longrightarrow> not_inits C es \<Longrightarrow> \<exists>sfs. sh C = \<lfloor>(sfs, Processing)\<rfloor> \<Longrightarrow> \<exists>sfs'. sh' C = \<lfloor>(sfs', Processing)\<rfloor>"
+(*<*)
+proof(induct rule:eval\<^sub>1_evals\<^sub>1_inducts)
+ case Call\<^sub>1 then show ?case using sees_wf\<^sub>1_nsub_RI[OF wf Call\<^sub>1.hyps(6)] nsub_RI_not_init by auto
+next
+ case (SCallInit\<^sub>1 ps h l sh vs h\<^sub>1 l\<^sub>1 sh\<^sub>1 C' M Ts T body D v' h\<^sub>2 l\<^sub>2 sh\<^sub>2 l\<^sub>2' e' h\<^sub>3 l\<^sub>3 sh\<^sub>3)
+ then show ?case
+ using SCallInit\<^sub>1 sees_wf\<^sub>1_nsub_RI[OF wf SCallInit\<^sub>1.hyps(3)] nsub_RI_not_init[of body] by auto
+next
+ case SCall\<^sub>1 then show ?case using sees_wf\<^sub>1_nsub_RI[OF wf SCall\<^sub>1.hyps(3)] nsub_RI_not_init by auto
+next
+ case (InitNone\<^sub>1 sh C1 C' Cs h l e' a a b) then show ?case by(cases "C = C1") auto
+next
+ case (InitDone\<^sub>1 sh C sfs C' Cs h l e' a a b) then show ?case by(cases Cs, auto)
+next
+ case (InitProcessing\<^sub>1 sh C sfs C' Cs h l e' a a b) then show ?case by(cases Cs, auto)
+next
+ case (InitError\<^sub>1 sh C1 sfs Cs h l e' a a b C') then show ?case by(cases "C = C1") auto
+next
+ case (InitObject\<^sub>1 sh C1 sfs sh' C' Cs h l e' a a b) then show ?case by(cases "C = C1") auto
+next
+ case (InitNonObject\<^sub>1 sh C1 sfs D a b sh' C' Cs h l e' a a b)
+ then show ?case by(cases "C = C1") auto
+next
+ case (RInit\<^sub>1 e a a b v h' l' sh' C sfs i sh'' C' Cs e\<^sub>1 a a b) then show ?case by(cases Cs, auto)
+next
+ case (RInitInitFail\<^sub>1 e h l sh a h' l' sh' C1 sfs i sh'' D Cs e\<^sub>1 h1 l1 sh1)
+ then show ?case using eval\<^sub>1_final by fastforce
+qed(auto)
+(*>*)
+
+lemma clinit\<^sub>1_proc_pres:
+ "\<lbrakk> wf_J\<^sub>1_prog P; P \<turnstile>\<^sub>1 \<langle>C\<^sub>0\<bullet>\<^sub>sclinit([]),(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',(h',l',sh')\<rangle>;
+ sh C' = \<lfloor>(sfs,Processing)\<rfloor> \<rbrakk>
+ \<Longrightarrow> \<exists>sfs. sh' C' = \<lfloor>(sfs,Processing)\<rfloor>"
+ by(auto dest: eval\<^sub>1_proc_pres)
+
+end
diff --git a/thys/JinjaDCI/Compiler/PCompiler.thy b/thys/JinjaDCI/Compiler/PCompiler.thy
--- a/thys/JinjaDCI/Compiler/PCompiler.thy
+++ b/thys/JinjaDCI/Compiler/PCompiler.thy
@@ -1,393 +1,393 @@
-(* Title: JinjaDCI/Compiler/PCompiler.thy
-
- Author: Tobias Nipkow, Susannah Mansky
- Copyright TUM 2003, UIUC 2019-20
-
- Based on the Jinja theory Common/PCompiler.thy by Tobias Nipkow
-*)
-
-section \<open> Program Compilation \<close>
-
-theory PCompiler
-imports "../Common/WellForm"
-begin
-
-definition compM :: "(staticb \<Rightarrow> 'a \<Rightarrow> 'b) \<Rightarrow> 'a mdecl \<Rightarrow> 'b mdecl"
-where
- "compM f \<equiv> \<lambda>(M, b, Ts, T, m). (M, b, Ts, T, f b m)"
-
-definition compC :: "(staticb \<Rightarrow> 'a \<Rightarrow> 'b) \<Rightarrow> 'a cdecl \<Rightarrow> 'b cdecl"
-where
- "compC f \<equiv> \<lambda>(C,D,Fdecls,Mdecls). (C,D,Fdecls, map (compM f) Mdecls)"
-
-definition compP :: "(staticb \<Rightarrow> 'a \<Rightarrow> 'b) \<Rightarrow> 'a prog \<Rightarrow> 'b prog"
-where
- "compP f \<equiv> map (compC f)"
-
-text\<open> Compilation preserves the program structure. Therefore lookup
-functions either commute with compilation (like method lookup) or are
-preserved by it (like the subclass relation). \<close>
-
-lemma map_of_map4:
- "map_of (map (\<lambda>(x,a,b,c).(x,a,b,f c)) ts) =
- map_option (\<lambda>(a,b,c).(a,b,f c)) \<circ> (map_of ts)"
-(*<*)
-proof(induct ts)
- case Nil then show ?case by simp
-qed fastforce
-(*>*)
-
-lemma map_of_map245:
- "map_of (map (\<lambda>(x,a,b,c,d).(x,a,b,c,f a c d)) ts) =
- map_option (\<lambda>(a,b,c,d).(a,b,c,f a c d)) \<circ> (map_of ts)"
-(*<*)
-proof(induct ts)
- case Nil then show ?case by simp
-qed fastforce
-(*>*)
-
-
-lemma class_compP:
- "class P C = Some (D, fs, ms)
- \<Longrightarrow> class (compP f P) C = Some (D, fs, map (compM f) ms)"
-(*<*)by(simp add:class_def compP_def compC_def map_of_map4)(*>*)
-
-
-lemma class_compPD:
- "class (compP f P) C = Some (D, fs, cms)
- \<Longrightarrow> \<exists>ms. class P C = Some(D,fs,ms) \<and> cms = map (compM f) ms"
-(*<*)by(clarsimp simp add:class_def compP_def compC_def map_of_map4)(*>*)
-
-
-lemma [simp]: "is_class (compP f P) C = is_class P C"
-(*<*)by(auto simp:is_class_def dest: class_compP class_compPD)(*>*)
-
-
-lemma [simp]: "class (compP f P) C = map_option (\<lambda>c. snd(compC f (C,c))) (class P C)"
-(*<*)
-by(simp add:compP_def compC_def class_def map_of_map4)
- (simp add:split_def)
-(*>*)
-
-
-lemma sees_methods_compP:
- "P \<turnstile> C sees_methods Mm \<Longrightarrow>
- compP f P \<turnstile> C sees_methods (map_option (\<lambda>((b,Ts,T,m),D). ((b,Ts,T,f b m),D)) \<circ> Mm)"
-(*<*)(is "?P \<Longrightarrow> compP f P \<turnstile> C sees_methods (?map Mm)")
-proof(induct rule: Methods.induct)
- case Object: (sees_methods_Object D fs ms Mm)
- let ?Mm1 = "\<lambda>x. map_option ((\<lambda>m. (m, Object)) \<circ> (\<lambda>(b, Ts, T, m). (b, Ts, T, f b m))) (map_of ms x)"
- let ?Mm2 = "\<lambda>x. map_option (case_prod (\<lambda>(b, Ts, T, m).
- Pair (b, Ts, T, f b m)) \<circ> (\<lambda>m. (m, Object))) (map_of ms x)"
- have Mm_eq: "\<And>x. ?Mm1 x = ?Mm2 x"
- proof -
- fix x show "?Mm1 x = ?Mm2 x"
- proof(cases "map_of ms x")
- case None then show ?thesis by simp
- qed fastforce
- qed
-
- have Mm: "Mm = map_option (\<lambda>m. (m, Object)) \<circ> map_of ms" by fact
- let ?Mm = "map_option (\<lambda>m. (m, Object)) \<circ> map_of (map (compM f) ms)"
- let ?Mm' = "?map Mm"
- have "?Mm' = ?Mm"
- by(rule ext) (simp add:Mm Mm_eq compM_def map_of_map245 option.map_comp)
- then show ?case by(rule sees_methods_Object[OF class_compP[OF Object(1)]])
-next
- case rec: (sees_methods_rec C D fs ms Mm Mm')
- have Mm': "Mm' = Mm ++ (map_option (\<lambda>m. (m, C)) \<circ> map_of ms)" by fact
- let ?Mm' = "?map Mm'"
- let ?Mm'' = "(?map Mm) ++ (map_option (\<lambda>m. (m, C)) \<circ> map_of (map (compM f) ms))"
- have "?Mm' = ?Mm''"
- by(rule ext) (simp add:Mm' map_add_def compM_def map_of_map245)
- moreover have "compP f P \<turnstile> C sees_methods ?Mm''"
- using sees_methods_rec[OF class_compP[OF rec(1)] rec(2,4)] by fast
- ultimately show "compP f P \<turnstile> C sees_methods ?Mm'" by simp
-qed
-(*>*)
-
-
-lemma sees_method_compP:
- "P \<turnstile> C sees M,b: Ts\<rightarrow>T = m in D \<Longrightarrow>
- compP f P \<turnstile> C sees M,b: Ts\<rightarrow>T = (f b m) in D"
-(*<*)by(fastforce elim:sees_methods_compP simp add:Method_def)(*>*)
-
-
-lemma [simp]:
- "P \<turnstile> C sees M,b: Ts\<rightarrow>T = m in D \<Longrightarrow>
- method (compP f P) C M = (D,b,Ts,T,f b m)"
-(*<*)
-proof -
- let ?P = "\<lambda>(D, b, Ts, T, m). compP f P \<turnstile> C sees M, b : Ts\<rightarrow>T = m in D"
- let ?a = "(D, b, Ts, T, f b m)"
- assume cM: "P \<turnstile> C sees M,b: Ts\<rightarrow>T = m in D"
- have compP_cM: "?P ?a" using sees_method_compP[OF cM] by simp
- moreover {
- fix x assume "?P x" then have "x = ?a"
- using compP_cM by(fastforce dest:sees_method_fun)
- }
- ultimately have "(THE x. ?P x) = ?a" by(rule the_equality)
- then show ?thesis by(simp add:method_def)
-qed
-(*>*)
-
-
-lemma sees_methods_compPD:
- "\<lbrakk> cP \<turnstile> C sees_methods Mm'; cP = compP f P \<rbrakk> \<Longrightarrow>
- \<exists>Mm. P \<turnstile> C sees_methods Mm \<and>
- Mm' = (map_option (\<lambda>((b,Ts,T,m),D). ((b,Ts,T,f b m),D)) \<circ> Mm)"
-(*<*)(is "\<lbrakk> ?P; ?Q \<rbrakk> \<Longrightarrow> \<exists>Mm. P \<turnstile> C sees_methods Mm \<and> Mm' = (?map Mm)")
-proof(induct rule: Methods.induct)
- case Object: (sees_methods_Object D fs ms Mm)
- then obtain ms' where P_Obj: "class P Object = \<lfloor>(D, fs, ms')\<rfloor>"
- and ms: "ms = map (compM f) ms'" by(clarsimp simp:compC_def)
-
- let ?Mm1 = "\<lambda>x. map_option ((\<lambda>m. (m, Object)) \<circ> (\<lambda>(b, Ts, T, m). (b, Ts, T, f b m))) (map_of ms' x)"
- let ?Mm2 = "\<lambda>x. map_option (case_prod (\<lambda>(b, Ts, T, m). Pair (b, Ts, T, f b m)) \<circ> (\<lambda>m. (m, Object)))
- (map_of ms' x)"
- have Mm_eq: "\<And>x. ?Mm1 x = ?Mm2 x"
- proof -
- fix x show "?Mm1 x = ?Mm2 x"
- proof(cases "map_of ms' x")
- case None then show ?thesis by simp
- qed fastforce
- qed
-
- let ?Mm = "map_option (\<lambda>m. (m,Object)) \<circ> map_of ms'"
- let ?Mm' = "?map ?Mm"
- have Mm: "Mm = map_option (\<lambda>m. (m, Object)) \<circ> map_of ms" by fact
- have "P \<turnstile> Object sees_methods ?Mm"
- using sees_methods_Object[OF P_Obj] by simp
- moreover have "Mm = ?Mm'"
- by(rule ext) (simp add:Mm_eq Mm ms compM_def map_of_map245 option.map_comp)
- ultimately show ?case by fast
-next
- case rec: (sees_methods_rec C D fs ms Mm Mm')
- then obtain ms' Mm\<^sub>D where P_D: "class P C = \<lfloor>(D, fs, ms')\<rfloor>"
- and ms: "ms = map (compM f) ms'" and C_nObj: "C \<noteq> Object"
- and Mm\<^sub>D: "P \<turnstile> D sees_methods Mm\<^sub>D"
- and Mm: "Mm = (\<lambda>a. map_option (case_prod (\<lambda>(b, Ts, T, m). Pair (b, Ts, T, f b m))) (Mm\<^sub>D a))"
- by(clarsimp simp:compC_def)
-
- let ?Mm = "Mm\<^sub>D ++ (map_option (\<lambda>m. (m, C)) \<circ> map_of ms')"
- let ?Mm1 = "Mm ++ (map_option (\<lambda>m. (m, C)) \<circ> map_of ms)"
- let ?Mm2 = "Mm ++ (map_option (\<lambda>m. (m, C)) \<circ> map_of (map (compM f) ms'))"
- let ?Mm3 = "?map ?Mm"
- have "Mm' = ?Mm1" by fact
- also have "\<dots> = ?Mm2" using ms by simp
- also have "\<dots> = ?Mm3"
- by(rule ext)(simp add:Mm map_add_def compM_def map_of_map245)
- moreover have "P \<turnstile> C sees_methods ?Mm"
- using sees_methods_rec[OF P_D C_nObj Mm\<^sub>D] by simp
- ultimately show ?case by fast
-qed
-(*>*)
-
-
-lemma sees_method_compPD:
- "compP f P \<turnstile> C sees M,b: Ts\<rightarrow>T = fm in D \<Longrightarrow>
- \<exists>m. P \<turnstile> C sees M,b: Ts\<rightarrow>T = m in D \<and> f b m = fm"
-(*<*)
-proof -
- assume "compP f P \<turnstile> C sees M,b: Ts\<rightarrow>T = fm in D"
- then obtain Mm where Mm: "compP f P \<turnstile> C sees_methods Mm"
- and MmM: "Mm M = \<lfloor>((b, Ts, T, fm), D)\<rfloor>"
- by(clarsimp simp:Method_def)
- show ?thesis using sees_methods_compPD[OF Mm refl] MmM
- by(fastforce simp: Method_def)
-qed
-(*>*)
-
-
-lemma [simp]: "subcls1(compP f P) = subcls1 P"
-(*<*)
-by(fastforce simp add: is_class_def compC_def intro:subcls1I order_antisym dest:subcls1D)
-(*>*)
-
-
-lemma compP_widen[simp]: "(compP f P \<turnstile> T \<le> T') = (P \<turnstile> T \<le> T')"
-(*<*)by(cases T')(simp_all add:widen_Class)(*>*)
-
-
-lemma [simp]: "(compP f P \<turnstile> Ts [\<le>] Ts') = (P \<turnstile> Ts [\<le>] Ts')"
-(*<*)
-proof(induct Ts)
- case (Cons a Ts)
- then show ?case by(cases Ts')(auto simp:fun_of_def)
-qed simp
-(*>*)
-
-
-lemma [simp]: "is_type (compP f P) T = is_type P T"
-(*<*)by(cases T) simp_all(*>*)
-
-
-lemma [simp]: "(compP (f::staticb\<Rightarrow>'a\<Rightarrow>'b) P \<turnstile> C has_fields FDTs) = (P \<turnstile> C has_fields FDTs)"
-(*<*)
- (is "?A = ?B")
-proof
- { fix cP::"'b prog" assume "cP \<turnstile> C has_fields FDTs"
- hence "cP = compP f P \<Longrightarrow> P \<turnstile> C has_fields FDTs"
- proof induct
- case has_fields_Object
- thus ?case by(fast intro:Fields.has_fields_Object dest:class_compPD)
- next
- case has_fields_rec
- thus ?case by(fast intro:Fields.has_fields_rec dest:class_compPD)
- qed
- } note lem = this
- assume ?A
- with lem show ?B by blast
-next
- assume ?B
- thus ?A
- proof induct
- case has_fields_Object
- thus ?case by(fast intro:Fields.has_fields_Object class_compP)
- next
- case has_fields_rec
- thus ?case by(fast intro:Fields.has_fields_rec class_compP)
- qed
-qed
-(*>*)
-
-
-lemma fields_compP [simp]: "fields (compP f P) C = fields P C"
-(*<*)by(simp add:fields_def)(*>*)
-
-lemma ifields_compP [simp]: "ifields (compP f P) C = ifields P C"
-(*<*)by(simp add:ifields_def)(*>*)
-
-lemma blank_compP [simp]: "blank (compP f P) C = blank P C"
-(*<*)by(simp add:blank_def)(*>*)
-
-lemma isfields_compP [simp]: "isfields (compP f P) C = isfields P C"
-(*<*)by(simp add:isfields_def)(*>*)
-
-lemma sblank_compP [simp]: "sblank (compP f P) C = sblank P C"
-(*<*)by(simp add:sblank_def)(*>*)
-
-lemma sees_fields_compP [simp]: "(compP f P \<turnstile> C sees F,b:T in D) = (P \<turnstile> C sees F,b:T in D)"
-(*<*)by(simp add:sees_field_def)(*>*)
-
-lemma has_field_compP [simp]: "(compP f P \<turnstile> C has F,b:T in D) = (P \<turnstile> C has F,b:T in D)"
-(*<*)by(simp add:has_field_def)(*>*)
-
-lemma field_compP [simp]: "field (compP f P) F D = field P F D"
-(*<*)by(simp add:field_def)(*>*)
-
-
-subsection\<open>Invariance of @{term wf_prog} under compilation \<close>
-
-lemma [iff]: "distinct_fst (compP f P) = distinct_fst P"
-(*<*)
-by (induct P)
- (auto simp:distinct_fst_def compP_def compC_def image_iff)
-(*>*)
-
-
-lemma [iff]: "distinct_fst (map (compM f) ms) = distinct_fst ms"
-(*<*)
-by (induct ms)
- (auto simp:distinct_fst_def compM_def image_iff)
-(*>*)
-
-
-lemma [iff]: "wf_syscls (compP f P) = wf_syscls P"
-(*<*)by(simp add:wf_syscls_def compP_def compC_def image_def Bex_def)(*>*)
-
-
-lemma [iff]: "wf_fdecl (compP f P) = wf_fdecl P"
-(*<*)by(simp add:wf_fdecl_def)(*>*)
-
-
-lemma wf_clinit_compM [iff]: "wf_clinit (map (compM f) ms) = wf_clinit ms"
-(*<*)
-proof(rule iffI)
- assume "wf_clinit (map (compM f) ms)"
- then obtain m where "(clinit, Static, [], Void, m) \<in> set ms"
- by(clarsimp simp: wf_clinit_def compM_def)
- then show "wf_clinit ms" by(fastforce simp: wf_clinit_def)
-next
- assume "wf_clinit ms"
- then obtain m where "(clinit, Static, [], Void, m) \<in> set ms"
- by(clarsimp simp: wf_clinit_def compM_def)
- then have "\<exists>m. (clinit, Static, [], Void, m)
- \<in> (\<lambda>x. case x of (M, b, Ts, T, m) \<Rightarrow> (M, b, Ts, T, f b m)) ` set ms"
- by(rule_tac x = "f Static m" in exI) (simp add: rev_image_eqI)
- then show "wf_clinit (map (compM f) ms)"
- by(simp add: wf_clinit_def compM_def)
-qed
-(*>*)
-
-lemma set_compP:
- "((C,D,fs,ms') \<in> set(compP f P)) =
- (\<exists>ms. (C,D,fs,ms) \<in> set P \<and> ms' = map (compM f) ms)"
-(*<*)by(fastforce simp add:compP_def compC_def image_iff Bex_def)(*>*)
-
-lemma wf_cdecl_compPI:
- "\<lbrakk> \<And>C M b Ts T m.
- \<lbrakk> wf_mdecl wf\<^sub>1 P C (M,b,Ts,T,m); P \<turnstile> C sees M,b:Ts\<rightarrow>T = m in C \<rbrakk>
- \<Longrightarrow> wf_mdecl wf\<^sub>2 (compP f P) C (M,b,Ts,T, f b m);
- \<forall>x\<in>set P. wf_cdecl wf\<^sub>1 P x; x \<in> set (compP f P); wf_prog p P \<rbrakk>
- \<Longrightarrow> wf_cdecl wf\<^sub>2 (compP f P) x"
-(*<*)
-proof -
- assume
- wfm: "\<And>C M b Ts T m. \<lbrakk> wf_mdecl wf\<^sub>1 P C (M,b,Ts,T,m); P \<turnstile> C sees M,b:Ts\<rightarrow>T = m in C \<rbrakk>
- \<Longrightarrow> wf_mdecl wf\<^sub>2 (compP f P) C (M,b,Ts,T, f b m)"
- and wfc: "\<forall>x\<in>set P. wf_cdecl wf\<^sub>1 P x"
- and compP: "x \<in> set (compP f P)" and wf: "wf_prog p P"
- obtain C D fs ms where x: "x = (C, D, fs, map (compM f) ms)"
- and x_set: "(C, D, fs, ms) \<in> set P"
- using compP by(case_tac x) (clarsimp simp: set_compP)
- have wfc': "wf_cdecl wf\<^sub>1 P (C, D, fs, ms)" using wfc x_set by fast
- let ?P = "compP f P" and ?ms = "compM f ` set ms"
- { fix M b Ts T m
- assume M: "(M,b,Ts,T,m) \<in> set ms"
- then have "wf_mdecl wf\<^sub>1 P C (M, b, Ts, T, m)" using wfc'
- by(simp add:wf_cdecl_def)
- moreover have cM: "P \<turnstile> C sees M, b : Ts\<rightarrow>T = m in C" using M
- by(rule mdecl_visible[OF wf x_set])
- ultimately have "wf_mdecl wf\<^sub>2 (compP f P) C (M, b, Ts, T, f b m)"
- by(rule wfm)
- }
- then have "\<forall>m \<in> ?ms. wf_mdecl wf\<^sub>2 ?P C m"
- by (clarsimp simp:compM_def)
- moreover have "C \<noteq> Object \<longrightarrow>
- (\<forall>(M,b,Ts,T,m)\<in>?ms.
- \<forall>D' b' Ts' T' m'. ?P \<turnstile> D sees M,b':Ts' \<rightarrow> T' = m' in D' \<longrightarrow>
- b = b' \<and> P \<turnstile> Ts' [\<le>] Ts \<and> P \<turnstile> T \<le> T')"
- proof -
- { fix M b Ts T m D' b' Ts' T' m'
- assume "C \<noteq> Object" and "(M,b,Ts,T,m)\<in>?ms"
- and dM: "?P \<turnstile> D sees M,b':Ts' \<rightarrow> T' = m' in D'"
- then have "b = b' \<and> P \<turnstile> Ts' [\<le>] Ts \<and> P \<turnstile> T \<le> T'"
- using wfc' sees_method_compPD[OF dM]
- by(fastforce simp:wf_cdecl_def image_iff compM_def)
- }
- then show ?thesis by fast
- qed
- moreover have "(\<forall>f\<in>set fs. wf_fdecl P f) \<and> distinct_fst fs
- \<and> distinct_fst ms \<and> wf_clinit ms
- \<and> (C \<noteq> Object \<longrightarrow> is_class P D \<and> \<not> P \<turnstile> D \<preceq>\<^sup>* C)" using wfc'
- by(simp add: wf_cdecl_def)
- ultimately show ?thesis using x by(simp add:wf_cdecl_def)
-qed
-(*>*)
-
-
-lemma wf_prog_compPI:
-assumes lift:
- "\<And>C M b Ts T m.
- \<lbrakk> P \<turnstile> C sees M,b:Ts\<rightarrow>T = m in C; wf_mdecl wf\<^sub>1 P C (M,b,Ts,T,m) \<rbrakk>
- \<Longrightarrow> wf_mdecl wf\<^sub>2 (compP f P) C (M,b,Ts,T, f b m)"
-and wf: "wf_prog wf\<^sub>1 P"
-shows "wf_prog wf\<^sub>2 (compP f P)"
-(*<*)
-using wf
-by (simp add:wf_prog_def) (blast intro:wf_cdecl_compPI lift wf)
-(*>*)
-
-
-end
+(* Title: JinjaDCI/Compiler/PCompiler.thy
+
+ Author: Tobias Nipkow, Susannah Mansky
+ Copyright TUM 2003, UIUC 2019-20
+
+ Based on the Jinja theory Common/PCompiler.thy by Tobias Nipkow
+*)
+
+section \<open> Program Compilation \<close>
+
+theory PCompiler
+imports "../Common/WellForm"
+begin
+
+definition compM :: "(staticb \<Rightarrow> 'a \<Rightarrow> 'b) \<Rightarrow> 'a mdecl \<Rightarrow> 'b mdecl"
+where
+ "compM f \<equiv> \<lambda>(M, b, Ts, T, m). (M, b, Ts, T, f b m)"
+
+definition compC :: "(staticb \<Rightarrow> 'a \<Rightarrow> 'b) \<Rightarrow> 'a cdecl \<Rightarrow> 'b cdecl"
+where
+ "compC f \<equiv> \<lambda>(C,D,Fdecls,Mdecls). (C,D,Fdecls, map (compM f) Mdecls)"
+
+definition compP :: "(staticb \<Rightarrow> 'a \<Rightarrow> 'b) \<Rightarrow> 'a prog \<Rightarrow> 'b prog"
+where
+ "compP f \<equiv> map (compC f)"
+
+text\<open> Compilation preserves the program structure. Therefore lookup
+functions either commute with compilation (like method lookup) or are
+preserved by it (like the subclass relation). \<close>
+
+lemma map_of_map4:
+ "map_of (map (\<lambda>(x,a,b,c).(x,a,b,f c)) ts) =
+ map_option (\<lambda>(a,b,c).(a,b,f c)) \<circ> (map_of ts)"
+(*<*)
+proof(induct ts)
+ case Nil then show ?case by simp
+qed fastforce
+(*>*)
+
+lemma map_of_map245:
+ "map_of (map (\<lambda>(x,a,b,c,d).(x,a,b,c,f a c d)) ts) =
+ map_option (\<lambda>(a,b,c,d).(a,b,c,f a c d)) \<circ> (map_of ts)"
+(*<*)
+proof(induct ts)
+ case Nil then show ?case by simp
+qed fastforce
+(*>*)
+
+
+lemma class_compP:
+ "class P C = Some (D, fs, ms)
+ \<Longrightarrow> class (compP f P) C = Some (D, fs, map (compM f) ms)"
+(*<*)by(simp add:class_def compP_def compC_def map_of_map4)(*>*)
+
+
+lemma class_compPD:
+ "class (compP f P) C = Some (D, fs, cms)
+ \<Longrightarrow> \<exists>ms. class P C = Some(D,fs,ms) \<and> cms = map (compM f) ms"
+(*<*)by(clarsimp simp add:class_def compP_def compC_def map_of_map4)(*>*)
+
+
+lemma [simp]: "is_class (compP f P) C = is_class P C"
+(*<*)by(auto simp:is_class_def dest: class_compP class_compPD)(*>*)
+
+
+lemma [simp]: "class (compP f P) C = map_option (\<lambda>c. snd(compC f (C,c))) (class P C)"
+(*<*)
+by(simp add:compP_def compC_def class_def map_of_map4)
+ (simp add:split_def)
+(*>*)
+
+
+lemma sees_methods_compP:
+ "P \<turnstile> C sees_methods Mm \<Longrightarrow>
+ compP f P \<turnstile> C sees_methods (map_option (\<lambda>((b,Ts,T,m),D). ((b,Ts,T,f b m),D)) \<circ> Mm)"
+(*<*)(is "?P \<Longrightarrow> compP f P \<turnstile> C sees_methods (?map Mm)")
+proof(induct rule: Methods.induct)
+ case Object: (sees_methods_Object D fs ms Mm)
+ let ?Mm1 = "\<lambda>x. map_option ((\<lambda>m. (m, Object)) \<circ> (\<lambda>(b, Ts, T, m). (b, Ts, T, f b m))) (map_of ms x)"
+ let ?Mm2 = "\<lambda>x. map_option (case_prod (\<lambda>(b, Ts, T, m).
+ Pair (b, Ts, T, f b m)) \<circ> (\<lambda>m. (m, Object))) (map_of ms x)"
+ have Mm_eq: "\<And>x. ?Mm1 x = ?Mm2 x"
+ proof -
+ fix x show "?Mm1 x = ?Mm2 x"
+ proof(cases "map_of ms x")
+ case None then show ?thesis by simp
+ qed fastforce
+ qed
+
+ have Mm: "Mm = map_option (\<lambda>m. (m, Object)) \<circ> map_of ms" by fact
+ let ?Mm = "map_option (\<lambda>m. (m, Object)) \<circ> map_of (map (compM f) ms)"
+ let ?Mm' = "?map Mm"
+ have "?Mm' = ?Mm"
+ by(rule ext) (simp add:Mm Mm_eq compM_def map_of_map245 option.map_comp)
+ then show ?case by(rule sees_methods_Object[OF class_compP[OF Object(1)]])
+next
+ case rec: (sees_methods_rec C D fs ms Mm Mm')
+ have Mm': "Mm' = Mm ++ (map_option (\<lambda>m. (m, C)) \<circ> map_of ms)" by fact
+ let ?Mm' = "?map Mm'"
+ let ?Mm'' = "(?map Mm) ++ (map_option (\<lambda>m. (m, C)) \<circ> map_of (map (compM f) ms))"
+ have "?Mm' = ?Mm''"
+ by(rule ext) (simp add:Mm' map_add_def compM_def map_of_map245)
+ moreover have "compP f P \<turnstile> C sees_methods ?Mm''"
+ using sees_methods_rec[OF class_compP[OF rec(1)] rec(2,4)] by fast
+ ultimately show "compP f P \<turnstile> C sees_methods ?Mm'" by simp
+qed
+(*>*)
+
+
+lemma sees_method_compP:
+ "P \<turnstile> C sees M,b: Ts\<rightarrow>T = m in D \<Longrightarrow>
+ compP f P \<turnstile> C sees M,b: Ts\<rightarrow>T = (f b m) in D"
+(*<*)by(fastforce elim:sees_methods_compP simp add:Method_def)(*>*)
+
+
+lemma [simp]:
+ "P \<turnstile> C sees M,b: Ts\<rightarrow>T = m in D \<Longrightarrow>
+ method (compP f P) C M = (D,b,Ts,T,f b m)"
+(*<*)
+proof -
+ let ?P = "\<lambda>(D, b, Ts, T, m). compP f P \<turnstile> C sees M, b : Ts\<rightarrow>T = m in D"
+ let ?a = "(D, b, Ts, T, f b m)"
+ assume cM: "P \<turnstile> C sees M,b: Ts\<rightarrow>T = m in D"
+ have compP_cM: "?P ?a" using sees_method_compP[OF cM] by simp
+ moreover {
+ fix x assume "?P x" then have "x = ?a"
+ using compP_cM by(fastforce dest:sees_method_fun)
+ }
+ ultimately have "(THE x. ?P x) = ?a" by(rule the_equality)
+ then show ?thesis by(simp add:method_def)
+qed
+(*>*)
+
+
+lemma sees_methods_compPD:
+ "\<lbrakk> cP \<turnstile> C sees_methods Mm'; cP = compP f P \<rbrakk> \<Longrightarrow>
+ \<exists>Mm. P \<turnstile> C sees_methods Mm \<and>
+ Mm' = (map_option (\<lambda>((b,Ts,T,m),D). ((b,Ts,T,f b m),D)) \<circ> Mm)"
+(*<*)(is "\<lbrakk> ?P; ?Q \<rbrakk> \<Longrightarrow> \<exists>Mm. P \<turnstile> C sees_methods Mm \<and> Mm' = (?map Mm)")
+proof(induct rule: Methods.induct)
+ case Object: (sees_methods_Object D fs ms Mm)
+ then obtain ms' where P_Obj: "class P Object = \<lfloor>(D, fs, ms')\<rfloor>"
+ and ms: "ms = map (compM f) ms'" by(clarsimp simp:compC_def)
+
+ let ?Mm1 = "\<lambda>x. map_option ((\<lambda>m. (m, Object)) \<circ> (\<lambda>(b, Ts, T, m). (b, Ts, T, f b m))) (map_of ms' x)"
+ let ?Mm2 = "\<lambda>x. map_option (case_prod (\<lambda>(b, Ts, T, m). Pair (b, Ts, T, f b m)) \<circ> (\<lambda>m. (m, Object)))
+ (map_of ms' x)"
+ have Mm_eq: "\<And>x. ?Mm1 x = ?Mm2 x"
+ proof -
+ fix x show "?Mm1 x = ?Mm2 x"
+ proof(cases "map_of ms' x")
+ case None then show ?thesis by simp
+ qed fastforce
+ qed
+
+ let ?Mm = "map_option (\<lambda>m. (m,Object)) \<circ> map_of ms'"
+ let ?Mm' = "?map ?Mm"
+ have Mm: "Mm = map_option (\<lambda>m. (m, Object)) \<circ> map_of ms" by fact
+ have "P \<turnstile> Object sees_methods ?Mm"
+ using sees_methods_Object[OF P_Obj] by simp
+ moreover have "Mm = ?Mm'"
+ by(rule ext) (simp add:Mm_eq Mm ms compM_def map_of_map245 option.map_comp)
+ ultimately show ?case by fast
+next
+ case rec: (sees_methods_rec C D fs ms Mm Mm')
+ then obtain ms' Mm\<^sub>D where P_D: "class P C = \<lfloor>(D, fs, ms')\<rfloor>"
+ and ms: "ms = map (compM f) ms'" and C_nObj: "C \<noteq> Object"
+ and Mm\<^sub>D: "P \<turnstile> D sees_methods Mm\<^sub>D"
+ and Mm: "Mm = (\<lambda>a. map_option (case_prod (\<lambda>(b, Ts, T, m). Pair (b, Ts, T, f b m))) (Mm\<^sub>D a))"
+ by(clarsimp simp:compC_def)
+
+ let ?Mm = "Mm\<^sub>D ++ (map_option (\<lambda>m. (m, C)) \<circ> map_of ms')"
+ let ?Mm1 = "Mm ++ (map_option (\<lambda>m. (m, C)) \<circ> map_of ms)"
+ let ?Mm2 = "Mm ++ (map_option (\<lambda>m. (m, C)) \<circ> map_of (map (compM f) ms'))"
+ let ?Mm3 = "?map ?Mm"
+ have "Mm' = ?Mm1" by fact
+ also have "\<dots> = ?Mm2" using ms by simp
+ also have "\<dots> = ?Mm3"
+ by(rule ext)(simp add:Mm map_add_def compM_def map_of_map245)
+ moreover have "P \<turnstile> C sees_methods ?Mm"
+ using sees_methods_rec[OF P_D C_nObj Mm\<^sub>D] by simp
+ ultimately show ?case by fast
+qed
+(*>*)
+
+
+lemma sees_method_compPD:
+ "compP f P \<turnstile> C sees M,b: Ts\<rightarrow>T = fm in D \<Longrightarrow>
+ \<exists>m. P \<turnstile> C sees M,b: Ts\<rightarrow>T = m in D \<and> f b m = fm"
+(*<*)
+proof -
+ assume "compP f P \<turnstile> C sees M,b: Ts\<rightarrow>T = fm in D"
+ then obtain Mm where Mm: "compP f P \<turnstile> C sees_methods Mm"
+ and MmM: "Mm M = \<lfloor>((b, Ts, T, fm), D)\<rfloor>"
+ by(clarsimp simp:Method_def)
+ show ?thesis using sees_methods_compPD[OF Mm refl] MmM
+ by(fastforce simp: Method_def)
+qed
+(*>*)
+
+
+lemma [simp]: "subcls1(compP f P) = subcls1 P"
+(*<*)
+by(fastforce simp add: is_class_def compC_def intro:subcls1I order_antisym dest:subcls1D)
+(*>*)
+
+
+lemma compP_widen[simp]: "(compP f P \<turnstile> T \<le> T') = (P \<turnstile> T \<le> T')"
+(*<*)by(cases T')(simp_all add:widen_Class)(*>*)
+
+
+lemma [simp]: "(compP f P \<turnstile> Ts [\<le>] Ts') = (P \<turnstile> Ts [\<le>] Ts')"
+(*<*)
+proof(induct Ts)
+ case (Cons a Ts)
+ then show ?case by(cases Ts')(auto simp:fun_of_def)
+qed simp
+(*>*)
+
+
+lemma [simp]: "is_type (compP f P) T = is_type P T"
+(*<*)by(cases T) simp_all(*>*)
+
+
+lemma [simp]: "(compP (f::staticb\<Rightarrow>'a\<Rightarrow>'b) P \<turnstile> C has_fields FDTs) = (P \<turnstile> C has_fields FDTs)"
+(*<*)
+ (is "?A = ?B")
+proof
+ { fix cP::"'b prog" assume "cP \<turnstile> C has_fields FDTs"
+ hence "cP = compP f P \<Longrightarrow> P \<turnstile> C has_fields FDTs"
+ proof induct
+ case has_fields_Object
+ thus ?case by(fast intro:Fields.has_fields_Object dest:class_compPD)
+ next
+ case has_fields_rec
+ thus ?case by(fast intro:Fields.has_fields_rec dest:class_compPD)
+ qed
+ } note lem = this
+ assume ?A
+ with lem show ?B by blast
+next
+ assume ?B
+ thus ?A
+ proof induct
+ case has_fields_Object
+ thus ?case by(fast intro:Fields.has_fields_Object class_compP)
+ next
+ case has_fields_rec
+ thus ?case by(fast intro:Fields.has_fields_rec class_compP)
+ qed
+qed
+(*>*)
+
+
+lemma fields_compP [simp]: "fields (compP f P) C = fields P C"
+(*<*)by(simp add:fields_def)(*>*)
+
+lemma ifields_compP [simp]: "ifields (compP f P) C = ifields P C"
+(*<*)by(simp add:ifields_def)(*>*)
+
+lemma blank_compP [simp]: "blank (compP f P) C = blank P C"
+(*<*)by(simp add:blank_def)(*>*)
+
+lemma isfields_compP [simp]: "isfields (compP f P) C = isfields P C"
+(*<*)by(simp add:isfields_def)(*>*)
+
+lemma sblank_compP [simp]: "sblank (compP f P) C = sblank P C"
+(*<*)by(simp add:sblank_def)(*>*)
+
+lemma sees_fields_compP [simp]: "(compP f P \<turnstile> C sees F,b:T in D) = (P \<turnstile> C sees F,b:T in D)"
+(*<*)by(simp add:sees_field_def)(*>*)
+
+lemma has_field_compP [simp]: "(compP f P \<turnstile> C has F,b:T in D) = (P \<turnstile> C has F,b:T in D)"
+(*<*)by(simp add:has_field_def)(*>*)
+
+lemma field_compP [simp]: "field (compP f P) F D = field P F D"
+(*<*)by(simp add:field_def)(*>*)
+
+
+subsection\<open>Invariance of @{term wf_prog} under compilation \<close>
+
+lemma [iff]: "distinct_fst (compP f P) = distinct_fst P"
+(*<*)
+by (induct P)
+ (auto simp:distinct_fst_def compP_def compC_def image_iff)
+(*>*)
+
+
+lemma [iff]: "distinct_fst (map (compM f) ms) = distinct_fst ms"
+(*<*)
+by (induct ms)
+ (auto simp:distinct_fst_def compM_def image_iff)
+(*>*)
+
+
+lemma [iff]: "wf_syscls (compP f P) = wf_syscls P"
+(*<*)by(simp add:wf_syscls_def compP_def compC_def image_def Bex_def)(*>*)
+
+
+lemma [iff]: "wf_fdecl (compP f P) = wf_fdecl P"
+(*<*)by(simp add:wf_fdecl_def)(*>*)
+
+
+lemma wf_clinit_compM [iff]: "wf_clinit (map (compM f) ms) = wf_clinit ms"
+(*<*)
+proof(rule iffI)
+ assume "wf_clinit (map (compM f) ms)"
+ then obtain m where "(clinit, Static, [], Void, m) \<in> set ms"
+ by(clarsimp simp: wf_clinit_def compM_def)
+ then show "wf_clinit ms" by(fastforce simp: wf_clinit_def)
+next
+ assume "wf_clinit ms"
+ then obtain m where "(clinit, Static, [], Void, m) \<in> set ms"
+ by(clarsimp simp: wf_clinit_def compM_def)
+ then have "\<exists>m. (clinit, Static, [], Void, m)
+ \<in> (\<lambda>x. case x of (M, b, Ts, T, m) \<Rightarrow> (M, b, Ts, T, f b m)) ` set ms"
+ by(rule_tac x = "f Static m" in exI) (simp add: rev_image_eqI)
+ then show "wf_clinit (map (compM f) ms)"
+ by(simp add: wf_clinit_def compM_def)
+qed
+(*>*)
+
+lemma set_compP:
+ "((C,D,fs,ms') \<in> set(compP f P)) =
+ (\<exists>ms. (C,D,fs,ms) \<in> set P \<and> ms' = map (compM f) ms)"
+(*<*)by(fastforce simp add:compP_def compC_def image_iff Bex_def)(*>*)
+
+lemma wf_cdecl_compPI:
+ "\<lbrakk> \<And>C M b Ts T m.
+ \<lbrakk> wf_mdecl wf\<^sub>1 P C (M,b,Ts,T,m); P \<turnstile> C sees M,b:Ts\<rightarrow>T = m in C \<rbrakk>
+ \<Longrightarrow> wf_mdecl wf\<^sub>2 (compP f P) C (M,b,Ts,T, f b m);
+ \<forall>x\<in>set P. wf_cdecl wf\<^sub>1 P x; x \<in> set (compP f P); wf_prog p P \<rbrakk>
+ \<Longrightarrow> wf_cdecl wf\<^sub>2 (compP f P) x"
+(*<*)
+proof -
+ assume
+ wfm: "\<And>C M b Ts T m. \<lbrakk> wf_mdecl wf\<^sub>1 P C (M,b,Ts,T,m); P \<turnstile> C sees M,b:Ts\<rightarrow>T = m in C \<rbrakk>
+ \<Longrightarrow> wf_mdecl wf\<^sub>2 (compP f P) C (M,b,Ts,T, f b m)"
+ and wfc: "\<forall>x\<in>set P. wf_cdecl wf\<^sub>1 P x"
+ and compP: "x \<in> set (compP f P)" and wf: "wf_prog p P"
+ obtain C D fs ms where x: "x = (C, D, fs, map (compM f) ms)"
+ and x_set: "(C, D, fs, ms) \<in> set P"
+ using compP by(case_tac x) (clarsimp simp: set_compP)
+ have wfc': "wf_cdecl wf\<^sub>1 P (C, D, fs, ms)" using wfc x_set by fast
+ let ?P = "compP f P" and ?ms = "compM f ` set ms"
+ { fix M b Ts T m
+ assume M: "(M,b,Ts,T,m) \<in> set ms"
+ then have "wf_mdecl wf\<^sub>1 P C (M, b, Ts, T, m)" using wfc'
+ by(simp add:wf_cdecl_def)
+ moreover have cM: "P \<turnstile> C sees M, b : Ts\<rightarrow>T = m in C" using M
+ by(rule mdecl_visible[OF wf x_set])
+ ultimately have "wf_mdecl wf\<^sub>2 (compP f P) C (M, b, Ts, T, f b m)"
+ by(rule wfm)
+ }
+ then have "\<forall>m \<in> ?ms. wf_mdecl wf\<^sub>2 ?P C m"
+ by (clarsimp simp:compM_def)
+ moreover have "C \<noteq> Object \<longrightarrow>
+ (\<forall>(M,b,Ts,T,m)\<in>?ms.
+ \<forall>D' b' Ts' T' m'. ?P \<turnstile> D sees M,b':Ts' \<rightarrow> T' = m' in D' \<longrightarrow>
+ b = b' \<and> P \<turnstile> Ts' [\<le>] Ts \<and> P \<turnstile> T \<le> T')"
+ proof -
+ { fix M b Ts T m D' b' Ts' T' m'
+ assume "C \<noteq> Object" and "(M,b,Ts,T,m)\<in>?ms"
+ and dM: "?P \<turnstile> D sees M,b':Ts' \<rightarrow> T' = m' in D'"
+ then have "b = b' \<and> P \<turnstile> Ts' [\<le>] Ts \<and> P \<turnstile> T \<le> T'"
+ using wfc' sees_method_compPD[OF dM]
+ by(fastforce simp:wf_cdecl_def image_iff compM_def)
+ }
+ then show ?thesis by fast
+ qed
+ moreover have "(\<forall>f\<in>set fs. wf_fdecl P f) \<and> distinct_fst fs
+ \<and> distinct_fst ms \<and> wf_clinit ms
+ \<and> (C \<noteq> Object \<longrightarrow> is_class P D \<and> \<not> P \<turnstile> D \<preceq>\<^sup>* C)" using wfc'
+ by(simp add: wf_cdecl_def)
+ ultimately show ?thesis using x by(simp add:wf_cdecl_def)
+qed
+(*>*)
+
+
+lemma wf_prog_compPI:
+assumes lift:
+ "\<And>C M b Ts T m.
+ \<lbrakk> P \<turnstile> C sees M,b:Ts\<rightarrow>T = m in C; wf_mdecl wf\<^sub>1 P C (M,b,Ts,T,m) \<rbrakk>
+ \<Longrightarrow> wf_mdecl wf\<^sub>2 (compP f P) C (M,b,Ts,T, f b m)"
+and wf: "wf_prog wf\<^sub>1 P"
+shows "wf_prog wf\<^sub>2 (compP f P)"
+(*<*)
+using wf
+by (simp add:wf_prog_def) (blast intro:wf_cdecl_compPI lift wf)
+(*>*)
+
+
+end
diff --git a/thys/JinjaDCI/Compiler/TypeComp.thy b/thys/JinjaDCI/Compiler/TypeComp.thy
--- a/thys/JinjaDCI/Compiler/TypeComp.thy
+++ b/thys/JinjaDCI/Compiler/TypeComp.thy
@@ -1,1686 +1,1686 @@
-(* Title: JinjaDCI/Compiler/TypeComp.thy
-
- Author: Tobias Nipkow, Susannah Mansky
- Copyright TUM 2003, UIUC 2019-20
-
- Based on the Jinja theory Compiler/TypeComp.thy by Tobias Nipkow
-*)
-
-section \<open> Preservation of Well-Typedness \<close>
-
-theory TypeComp
-imports Compiler "../BV/BVSpec"
-begin
-
-(*<*)
-declare nth_append[simp]
-(*>*)
-
-lemma max_stack1: "P,E \<turnstile>\<^sub>1 e :: T \<Longrightarrow> 1 \<le> max_stack e"
-(*<*)using max_stack1'[OF WT\<^sub>1_nsub_RI] by simp(*>*)
-
-locale TC0 =
- fixes P :: "J\<^sub>1_prog" and mxl :: nat
-begin
-
-definition "ty E e = (THE T. P,E \<turnstile>\<^sub>1 e :: T)"
-
-definition "ty\<^sub>l E A' = map (\<lambda>i. if i \<in> A' \<and> i < size E then OK(E!i) else Err) [0..<mxl]"
-
-definition "ty\<^sub>i' ST E A = (case A of None \<Rightarrow> None | \<lfloor>A'\<rfloor> \<Rightarrow> Some(ST, ty\<^sub>l E A'))"
-
-definition "after E A ST e = ty\<^sub>i' (ty E e # ST) E (A \<squnion> \<A> e)"
-
-end
-
-lemma (in TC0) ty_def2 [simp]: "P,E \<turnstile>\<^sub>1 e :: T \<Longrightarrow> ty E e = T"
-(*<*)by(unfold ty_def) (blast intro: the_equality WT\<^sub>1_unique)(*>*)
-
-lemma (in TC0) [simp]: "ty\<^sub>i' ST E None = None"
-(*<*)by (simp add: ty\<^sub>i'_def)(*>*)
-
-lemma (in TC0) ty\<^sub>l_app_diff[simp]:
- "ty\<^sub>l (E@[T]) (A - {size E}) = ty\<^sub>l E A"
-(*<*)by(auto simp add:ty\<^sub>l_def hyperset_defs)(*>*)
-
-
-lemma (in TC0) ty\<^sub>i'_app_diff[simp]:
- "ty\<^sub>i' ST (E @ [T]) (A \<ominus> size E) = ty\<^sub>i' ST E A"
-(*<*)by(auto simp add:ty\<^sub>i'_def hyperset_defs)(*>*)
-
-
-lemma (in TC0) ty\<^sub>l_antimono:
- "A \<subseteq> A' \<Longrightarrow> P \<turnstile> ty\<^sub>l E A' [\<le>\<^sub>\<top>] ty\<^sub>l E A"
-(*<*)by(auto simp:ty\<^sub>l_def list_all2_conv_all_nth)(*>*)
-
-
-lemma (in TC0) ty\<^sub>i'_antimono:
- "A \<subseteq> A' \<Longrightarrow> P \<turnstile> ty\<^sub>i' ST E \<lfloor>A'\<rfloor> \<le>' ty\<^sub>i' ST E \<lfloor>A\<rfloor>"
-(*<*)by(auto simp:ty\<^sub>i'_def ty\<^sub>l_def list_all2_conv_all_nth)(*>*)
-
-
-lemma (in TC0) ty\<^sub>l_env_antimono:
- "P \<turnstile> ty\<^sub>l (E@[T]) A [\<le>\<^sub>\<top>] ty\<^sub>l E A"
-(*<*)by(auto simp:ty\<^sub>l_def list_all2_conv_all_nth)(*>*)
-
-
-lemma (in TC0) ty\<^sub>i'_env_antimono:
- "P \<turnstile> ty\<^sub>i' ST (E@[T]) A \<le>' ty\<^sub>i' ST E A"
-(*<*)by(auto simp:ty\<^sub>i'_def ty\<^sub>l_def list_all2_conv_all_nth)(*>*)
-
-
-lemma (in TC0) ty\<^sub>i'_incr:
- "P \<turnstile> ty\<^sub>i' ST (E @ [T]) \<lfloor>insert (size E) A\<rfloor> \<le>' ty\<^sub>i' ST E \<lfloor>A\<rfloor>"
-(*<*)by(auto simp:ty\<^sub>i'_def ty\<^sub>l_def list_all2_conv_all_nth)(*>*)
-
-
-lemma (in TC0) ty\<^sub>l_incr:
- "P \<turnstile> ty\<^sub>l (E @ [T]) (insert (size E) A) [\<le>\<^sub>\<top>] ty\<^sub>l E A"
-(*<*)by(auto simp: hyperset_defs ty\<^sub>l_def list_all2_conv_all_nth)(*>*)
-
-
-lemma (in TC0) ty\<^sub>l_in_types:
- "set E \<subseteq> types P \<Longrightarrow> ty\<^sub>l E A \<in> list mxl (err (types P))"
-(*<*)by(auto simp add:ty\<^sub>l_def intro!:listI dest!: nth_mem)(*>*)
-
-locale TC1 = TC0
-begin
-
-primrec compT :: "ty list \<Rightarrow> nat hyperset \<Rightarrow> ty list \<Rightarrow> expr\<^sub>1 \<Rightarrow> ty\<^sub>i' list" and
- compTs :: "ty list \<Rightarrow> nat hyperset \<Rightarrow> ty list \<Rightarrow> expr\<^sub>1 list \<Rightarrow> ty\<^sub>i' list" where
-"compT E A ST (new C) = []"
-| "compT E A ST (Cast C e) =
- compT E A ST e @ [after E A ST e]"
-| "compT E A ST (Val v) = []"
-| "compT E A ST (e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2) =
- (let ST\<^sub>1 = ty E e\<^sub>1#ST; A\<^sub>1 = A \<squnion> \<A> e\<^sub>1 in
- compT E A ST e\<^sub>1 @ [after E A ST e\<^sub>1] @
- compT E A\<^sub>1 ST\<^sub>1 e\<^sub>2 @ [after E A\<^sub>1 ST\<^sub>1 e\<^sub>2])"
-| "compT E A ST (Var i) = []"
-| "compT E A ST (i := e) = compT E A ST e @
- [after E A ST e, ty\<^sub>i' ST E (A \<squnion> \<A> e \<squnion> \<lfloor>{i}\<rfloor>)]"
-| "compT E A ST (e\<bullet>F{D}) =
- compT E A ST e @ [after E A ST e]"
-| "compT E A ST (C\<bullet>\<^sub>sF{D}) = []"
-| "compT E A ST (e\<^sub>1\<bullet>F{D} := e\<^sub>2) =
- (let ST\<^sub>1 = ty E e\<^sub>1#ST; A\<^sub>1 = A \<squnion> \<A> e\<^sub>1; A\<^sub>2 = A\<^sub>1 \<squnion> \<A> e\<^sub>2 in
- compT E A ST e\<^sub>1 @ [after E A ST e\<^sub>1] @
- compT E A\<^sub>1 ST\<^sub>1 e\<^sub>2 @ [after E A\<^sub>1 ST\<^sub>1 e\<^sub>2] @
- [ty\<^sub>i' ST E A\<^sub>2])"
-| "compT E A ST (C\<bullet>\<^sub>sF{D} := e\<^sub>2) = compT E A ST e\<^sub>2 @ [after E A ST e\<^sub>2] @ [ty\<^sub>i' ST E (A \<squnion> \<A> e\<^sub>2)]"
-| "compT E A ST {i:T; e} = compT (E@[T]) (A\<ominus>i) ST e"
-| "compT E A ST (e\<^sub>1;;e\<^sub>2) =
- (let A\<^sub>1 = A \<squnion> \<A> e\<^sub>1 in
- compT E A ST e\<^sub>1 @ [after E A ST e\<^sub>1, ty\<^sub>i' ST E A\<^sub>1] @
- compT E A\<^sub>1 ST e\<^sub>2)"
-| "compT E A ST (if (e) e\<^sub>1 else e\<^sub>2) =
- (let A\<^sub>0 = A \<squnion> \<A> e; \<tau> = ty\<^sub>i' ST E A\<^sub>0 in
- compT E A ST e @ [after E A ST e, \<tau>] @
- compT E A\<^sub>0 ST e\<^sub>1 @ [after E A\<^sub>0 ST e\<^sub>1, \<tau>] @
- compT E A\<^sub>0 ST e\<^sub>2)"
-| "compT E A ST (while (e) c) =
- (let A\<^sub>0 = A \<squnion> \<A> e; A\<^sub>1 = A\<^sub>0 \<squnion> \<A> c; \<tau> = ty\<^sub>i' ST E A\<^sub>0 in
- compT E A ST e @ [after E A ST e, \<tau>] @
- compT E A\<^sub>0 ST c @ [after E A\<^sub>0 ST c, ty\<^sub>i' ST E A\<^sub>1, ty\<^sub>i' ST E A\<^sub>0])"
-| "compT E A ST (throw e) = compT E A ST e @ [after E A ST e]"
-| "compT E A ST (e\<bullet>M(es)) =
- compT E A ST e @ [after E A ST e] @
- compTs E (A \<squnion> \<A> e) (ty E e # ST) es"
-| "compT E A ST (C\<bullet>\<^sub>sM(es)) = compTs E A ST es"
-| "compT E A ST (try e\<^sub>1 catch(C i) e\<^sub>2) =
- compT E A ST e\<^sub>1 @ [after E A ST e\<^sub>1] @
- [ty\<^sub>i' (Class C#ST) E A, ty\<^sub>i' ST (E@[Class C]) (A \<squnion> \<lfloor>{i}\<rfloor>)] @
- compT (E@[Class C]) (A \<squnion> \<lfloor>{i}\<rfloor>) ST e\<^sub>2"
-| "compT E A ST (INIT C (Cs,b) \<leftarrow> e) = []"
-| "compT E A ST (RI(C,e');Cs \<leftarrow> e) = []"
-| "compTs E A ST [] = []"
-| "compTs E A ST (e#es) = compT E A ST e @ [after E A ST e] @
- compTs E (A \<squnion> (\<A> e)) (ty E e # ST) es"
-
-definition compT\<^sub>a :: "ty list \<Rightarrow> nat hyperset \<Rightarrow> ty list \<Rightarrow> expr\<^sub>1 \<Rightarrow> ty\<^sub>i' list" where
- "compT\<^sub>a E A ST e = compT E A ST e @ [after E A ST e]"
-
-end
-
-lemma compE\<^sub>2_not_Nil[simp]: "P,E \<turnstile>\<^sub>1 e :: T \<Longrightarrow> compE\<^sub>2 e \<noteq> []"
-(*<*)by(simp add: compE\<^sub>2_not_Nil' WT\<^sub>1_nsub_RI)(*>*)
-
-lemma (in TC1) compT_sizes':
-shows "\<And>E A ST. \<not>sub_RI e \<Longrightarrow> size(compT E A ST e) = size(compE\<^sub>2 e) - 1"
-and "\<And>E A ST. \<not>sub_RIs es \<Longrightarrow> size(compTs E A ST es) = size(compEs\<^sub>2 es)"
-(*<*)
-by(induct e and es rule: compE\<^sub>2.induct compEs\<^sub>2.induct)
- (auto split:bop.splits nat_diff_split simp: compE\<^sub>2_not_Nil')
-(*>*)
-
-lemma (in TC1) compT_sizes[simp]:
-shows "\<And>E A ST. P,E \<turnstile>\<^sub>1 e :: T \<Longrightarrow> size(compT E A ST e) = size(compE\<^sub>2 e) - 1"
-and "\<And>E A ST. P,E \<turnstile>\<^sub>1 es [::] Ts \<Longrightarrow> size(compTs E A ST es) = size(compEs\<^sub>2 es)"
-(*<*)using compT_sizes' WT\<^sub>1_nsub_RI WTs\<^sub>1_nsub_RIs by auto(*>*)
-
-
-lemma (in TC1) [simp]: "\<And>ST E. \<lfloor>\<tau>\<rfloor> \<notin> set (compT E None ST e)"
-and [simp]: "\<And>ST E. \<lfloor>\<tau>\<rfloor> \<notin> set (compTs E None ST es)"
-(*<*)by(induct e and es rule: compT.induct compTs.induct) (simp_all add:after_def)(*>*)
-
-
-lemma (in TC0) pair_eq_ty\<^sub>i'_conv:
- "(\<lfloor>(ST, LT)\<rfloor> = ty\<^sub>i' ST\<^sub>0 E A) =
- (case A of None \<Rightarrow> False | Some A \<Rightarrow> (ST = ST\<^sub>0 \<and> LT = ty\<^sub>l E A))"
-(*<*)by(simp add:ty\<^sub>i'_def)(*>*)
-
-
-lemma (in TC0) pair_conv_ty\<^sub>i':
- "\<lfloor>(ST, ty\<^sub>l E A)\<rfloor> = ty\<^sub>i' ST E \<lfloor>A\<rfloor>"
-(*<*)by(simp add:ty\<^sub>i'_def)(*>*)
-
-(*<*)
-declare (in TC0)
- ty\<^sub>i'_antimono [intro!] after_def[simp]
- pair_conv_ty\<^sub>i'[simp] pair_eq_ty\<^sub>i'_conv[simp]
-(*>*)
-
-
-lemma (in TC1) compT_LT_prefix:
- "\<And>E A ST\<^sub>0. \<lbrakk> \<lfloor>(ST,LT)\<rfloor> \<in> set(compT E A ST\<^sub>0 e); \<B> e (size E) \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<lfloor>(ST,LT)\<rfloor> \<le>' ty\<^sub>i' ST E A"
-and
- "\<And>E A ST\<^sub>0. \<lbrakk> \<lfloor>(ST,LT)\<rfloor> \<in> set(compTs E A ST\<^sub>0 es); \<B>s es (size E) \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<lfloor>(ST,LT)\<rfloor> \<le>' ty\<^sub>i' ST E A"
-(*<*)
-proof(induct e and es rule: compT.induct compTs.induct)
- case FAss thus ?case by(fastforce simp:hyperset_defs elim!:sup_state_opt_trans)
-next
- case BinOp thus ?case
- by(fastforce simp:hyperset_defs elim!:sup_state_opt_trans split:bop.splits)
-next
- case Seq thus ?case by(fastforce simp:hyperset_defs elim!:sup_state_opt_trans)
-next
- case While thus ?case by(fastforce simp:hyperset_defs elim!:sup_state_opt_trans)
-next
- case Cond thus ?case by(fastforce simp:hyperset_defs elim!:sup_state_opt_trans)
-next
- case Block thus ?case
- by(force simp add:hyperset_defs ty\<^sub>i'_def simp del:pair_conv_ty\<^sub>i'
- elim!:sup_state_opt_trans)
-next
- case Call thus ?case by(fastforce simp:hyperset_defs elim!:sup_state_opt_trans)
-next
- case Cons_exp thus ?case
- by(fastforce simp:hyperset_defs elim!:sup_state_opt_trans)
-next
- case TryCatch thus ?case
- by(fastforce simp:hyperset_defs intro!: ty\<^sub>i'_incr
- elim!:sup_state_opt_trans)
-qed (auto simp:hyperset_defs)
-
-declare (in TC0)
- ty\<^sub>i'_antimono [rule del] after_def[simp del]
- pair_conv_ty\<^sub>i'[simp del] pair_eq_ty\<^sub>i'_conv[simp del]
-(*>*)
-
-
-lemma [iff]: "OK None \<in> states P mxs mxl"
-(*<*)by(simp add: JVM_states_unfold)(*>*)
-
-lemma (in TC0) after_in_states:
-assumes wf: "wf_prog p P" and wt: "P,E \<turnstile>\<^sub>1 e :: T"
- and Etypes: "set E \<subseteq> types P" and STtypes: "set ST \<subseteq> types P"
- and stack: "size ST + max_stack e \<le> mxs"
-shows "OK (after E A ST e) \<in> states P mxs mxl"
-(*<*)
-proof -
- have "size ST + 1 \<le> mxs" using max_stack1[where e=e] wt stack
- by fastforce
- then show ?thesis using assms
- by(simp add: after_def ty\<^sub>i'_def JVM_states_unfold ty\<^sub>l_in_types)
- (blast intro!:listI WT\<^sub>1_is_type)
-qed
-(*>*)
-
-
-lemma (in TC0) OK_ty\<^sub>i'_in_statesI[simp]:
- "\<lbrakk> set E \<subseteq> types P; set ST \<subseteq> types P; size ST \<le> mxs \<rbrakk>
- \<Longrightarrow> OK (ty\<^sub>i' ST E A) \<in> states P mxs mxl"
-(*<*)
-by(simp add:ty\<^sub>i'_def JVM_states_unfold ty\<^sub>l_in_types)
- (blast intro!:listI)
-(*>*)
-
-
-lemma is_class_type_aux: "is_class P C \<Longrightarrow> is_type P (Class C)"
-(*<*)by(simp)(*>*)
-
-(*<*)
-declare is_type_simps[simp del] subsetI[rule del]
-(*>*)
-
-theorem (in TC1) compT_states:
-assumes wf: "wf_prog p P"
-shows "\<And>E T A ST.
- \<lbrakk> P,E \<turnstile>\<^sub>1 e :: T; set E \<subseteq> types P; set ST \<subseteq> types P;
- size ST + max_stack e \<le> mxs; size E + max_vars e \<le> mxl \<rbrakk>
- \<Longrightarrow> OK ` set(compT E A ST e) \<subseteq> states P mxs mxl"
-(*<*)(is "\<And>E T A ST. PROP ?P e E T A ST")(*>*)
-
-and "\<And>E Ts A ST.
- \<lbrakk> P,E \<turnstile>\<^sub>1 es[::]Ts; set E \<subseteq> types P; set ST \<subseteq> types P;
- size ST + max_stacks es \<le> mxs; size E + max_varss es \<le> mxl \<rbrakk>
- \<Longrightarrow> OK ` set(compTs E A ST es) \<subseteq> states P mxs mxl"
-(*<*)(is "\<And>E Ts A ST. PROP ?Ps es E Ts A ST")
-proof(induct e and es rule: compT.induct compTs.induct)
- case new thus ?case by(simp)
-next
- case (Cast C e) thus ?case by (auto simp:after_in_states[OF wf])
-next
- case Val thus ?case by(simp)
-next
- case Var thus ?case by(simp)
-next
- case LAss thus ?case by(auto simp:after_in_states[OF wf])
-next
- case FAcc thus ?case by(auto simp:after_in_states[OF wf])
-next
- case SFAcc thus ?case by(auto simp:after_in_states[OF wf])
-next
- case FAss thus ?case
- by(auto simp:image_Un WT\<^sub>1_is_type[OF wf] after_in_states[OF wf])
-next
- case SFAss thus ?case
- by(auto simp:image_Un WT\<^sub>1_is_type[OF wf] after_in_states[OF wf])
-next
- case Seq thus ?case
- by(auto simp:image_Un after_in_states[OF wf])
-next
- case BinOp thus ?case
- by(auto simp:image_Un WT\<^sub>1_is_type[OF wf] after_in_states[OF wf])
-next
- case Cond thus ?case
- by(force simp:image_Un WT\<^sub>1_is_type[OF wf] after_in_states[OF wf])
-next
- case While thus ?case
- by(auto simp:image_Un WT\<^sub>1_is_type[OF wf] after_in_states[OF wf])
-next
- case Block thus ?case by(auto)
-next
- case (TryCatch e\<^sub>1 C i e\<^sub>2)
- moreover have "size ST + 1 \<le> mxs"
- using TryCatch.prems max_stack1[where e=e\<^sub>1] by fastforce
- ultimately show ?case
- by(auto simp:image_Un WT\<^sub>1_is_type[OF wf] after_in_states[OF wf]
- is_class_type_aux)
-next
- case Nil_exp thus ?case by simp
-next
- case Cons_exp thus ?case
- by(auto simp:image_Un WT\<^sub>1_is_type[OF wf] after_in_states[OF wf])
-next
- case throw thus ?case
- by(auto simp: WT\<^sub>1_is_type[OF wf] after_in_states[OF wf])
-next
- case Call thus ?case
- by(auto simp:image_Un WT\<^sub>1_is_type[OF wf] after_in_states[OF wf])
-next
- case SCall thus ?case
- by(auto simp:image_Un WT\<^sub>1_is_type[OF wf] after_in_states[OF wf])
-next
- case INIT thus ?case by simp
-next
- case RI thus ?case by simp
-qed
-
-declare is_type_simps[simp] subsetI[intro!]
-(*>*)
-
-
-definition shift :: "nat \<Rightarrow> ex_table \<Rightarrow> ex_table"
-where
- "shift n xt \<equiv> map (\<lambda>(from,to,C,handler,depth). (from+n,to+n,C,handler+n,depth)) xt"
-
-
-lemma [simp]: "shift 0 xt = xt"
-(*<*)by(induct xt)(auto simp:shift_def)(*>*)
-
-lemma [simp]: "shift n [] = []"
-(*<*)by(simp add:shift_def)(*>*)
-
-lemma [simp]: "shift n (xt\<^sub>1 @ xt\<^sub>2) = shift n xt\<^sub>1 @ shift n xt\<^sub>2"
-(*<*)by(simp add:shift_def)(*>*)
-
-lemma [simp]: "shift m (shift n xt) = shift (m+n) xt"
-(*<*)by(induct xt)(auto simp:shift_def)(*>*)
-
-lemma [simp]: "pcs (shift n xt) = {pc+n|pc. pc \<in> pcs xt}"
-(*<*)
-proof -
- { fix x f t C h d
- assume "(f,t,C,h,d) \<in> set xt" and "f + n \<le> x"
- and "x < t + n"
- then have "\<exists>pc. x = pc + n \<and> (\<exists>x\<in>set xt. pc \<in> (case x of (f, t, C, h, d) \<Rightarrow> {f..<t}))"
- by(rule_tac x = "x-n" in exI) (force split:nat_diff_split)
- }
- then show ?thesis by(auto simp:shift_def pcs_def) fast
-qed
-(*>*)
-
-
-lemma shift_compxE\<^sub>2:
-shows "\<And>pc pc' d. shift pc (compxE\<^sub>2 e pc' d) = compxE\<^sub>2 e (pc' + pc) d"
-and "\<And>pc pc' d. shift pc (compxEs\<^sub>2 es pc' d) = compxEs\<^sub>2 es (pc' + pc) d"
-(*<*)
-by(induct e and es rule: compxE\<^sub>2.induct compxEs\<^sub>2.induct)
- (auto simp:shift_def ac_simps)
-(*>*)
-
-
-lemma compxE\<^sub>2_size_convs[simp]:
-shows "n \<noteq> 0 \<Longrightarrow> compxE\<^sub>2 e n d = shift n (compxE\<^sub>2 e 0 d)"
-and "n \<noteq> 0 \<Longrightarrow> compxEs\<^sub>2 es n d = shift n (compxEs\<^sub>2 es 0 d)"
-(*<*)by(simp_all add:shift_compxE\<^sub>2)(*>*)
-
-locale TC2 = TC1 +
- fixes T\<^sub>r :: ty and mxs :: pc
-begin
-
-definition
- wt_instrs :: "instr list \<Rightarrow> ex_table \<Rightarrow> ty\<^sub>i' list \<Rightarrow> bool"
- ("(\<turnstile> _, _ /[::]/ _)" [0,0,51] 50) where
- "\<turnstile> is,xt [::] \<tau>s \<longleftrightarrow> size is < size \<tau>s \<and> pcs xt \<subseteq> {0..<size is} \<and>
- (\<forall>pc< size is. P,T\<^sub>r,mxs,size \<tau>s,xt \<turnstile> is!pc,pc :: \<tau>s)"
-
-end
-
-notation TC2.wt_instrs ("(_,_,_ \<turnstile>/ _, _ /[::]/ _)" [50,50,50,50,50,51] 50)
-
-(*<*)
-lemmas (in TC2) wt_defs =
- wt_instrs_def wt_instr_def app_def eff_def norm_eff_def
-(*>*)
-
-lemma (in TC2) [simp]: "\<tau>s \<noteq> [] \<Longrightarrow> \<turnstile> [],[] [::] \<tau>s"
-(*<*) by (simp add: wt_defs) (*>*)
-
-lemma [simp]: "eff i P pc et None = []"
-(*<*)by (simp add: Effect.eff_def)(*>*)
-
-(*<*)
-declare split_comp_eq[simp del]
-(*>*)
-
-lemma wt_instr_appR:
- "\<lbrakk> P,T,m,mpc,xt \<turnstile> is!pc,pc :: \<tau>s;
- pc < size is; size is < size \<tau>s; mpc \<le> size \<tau>s; mpc \<le> mpc' \<rbrakk>
- \<Longrightarrow> P,T,m,mpc',xt \<turnstile> is!pc,pc :: \<tau>s@\<tau>s'"
-(*<*)by (fastforce simp:wt_instr_def app_def)(*>*)
-
-
-lemma relevant_entries_shift [simp]:
- "relevant_entries P i (pc+n) (shift n xt) = shift n (relevant_entries P i pc xt)"
-(*<*)
-proof(induct xt)
- case Nil
- then show ?case by (simp add: relevant_entries_def shift_def)
-next
- case (Cons a xt)
- then show ?case
- by (auto simp add: relevant_entries_def shift_def is_relevant_entry_def)
-qed
-(*>*)
-
-
-lemma [simp]:
- "xcpt_eff i P (pc+n) \<tau> (shift n xt) =
- map (\<lambda>(pc,\<tau>). (pc + n, \<tau>)) (xcpt_eff i P pc \<tau> xt)"
-(*<*)
-proof -
- obtain ST LT where "\<tau> = (ST, LT)" by(cases \<tau>) simp
- then show ?thesis by(simp add: xcpt_eff_def) (auto simp add: shift_def)
-qed
-(*>*)
-
-
-lemma [simp]:
- "app\<^sub>i (i, P, pc, m, T, \<tau>) \<Longrightarrow>
- eff i P (pc+n) (shift n xt) (Some \<tau>) =
- map (\<lambda>(pc,\<tau>). (pc+n,\<tau>)) (eff i P pc xt (Some \<tau>))"
-(*<*)by(cases "i") (auto simp add:eff_def norm_eff_def)(*>*)
-
-
-lemma [simp]:
- "xcpt_app i P (pc+n) mxs (shift n xt) \<tau> = xcpt_app i P pc mxs xt \<tau>"
-(*<*)by (simp add: xcpt_app_def) (auto simp add: shift_def)(*>*)
-
-
-lemma wt_instr_appL:
-assumes "P,T,m,mpc,xt \<turnstile> i,pc :: \<tau>s" and "pc < size \<tau>s" and "mpc \<le> size \<tau>s"
-shows "P,T,m,mpc + size \<tau>s',shift (size \<tau>s') xt \<turnstile> i,pc+size \<tau>s' :: \<tau>s'@\<tau>s"
-(*<*)
-proof -
- let ?t = "(\<tau>s'@\<tau>s)!(pc+size \<tau>s')"
- show ?thesis
- proof(cases ?t)
- case (Some \<tau>)
- obtain ST LT where [simp]: "\<tau> = (ST, LT)" by(cases \<tau>) simp
- have "app\<^sub>i (i, P, pc + length \<tau>s', m, T, \<tau>)"
- using Some assms by(cases "i") (auto simp:wt_instr_def app_def)
- moreover {
- fix pc' \<tau>' assume "(pc',\<tau>') \<in> set (eff i P pc xt ?t)"
- then have "P \<turnstile> \<tau>' \<le>' \<tau>s!pc'" and "pc' < mpc"
- using Some assms by(auto simp:wt_instr_def app_def)
- }
- ultimately show ?thesis using Some assms
- by(fastforce simp:wt_instr_def app_def)
- qed (auto simp:wt_instr_def app_def)
-qed
-(*>*)
-
-
-lemma wt_instr_Cons:
-assumes wti: "P,T,m,mpc - 1,[] \<turnstile> i,pc - 1 :: \<tau>s"
- and pcl: "0 < pc" and mpcl: "0 < mpc"
- and pcu: "pc < size \<tau>s + 1" and mpcu: "mpc \<le> size \<tau>s + 1"
-shows "P,T,m,mpc,[] \<turnstile> i,pc :: \<tau>#\<tau>s"
-(*<*)
-proof -
- have "pc - 1 < length \<tau>s" using pcl pcu by arith
- moreover have "mpc - 1 \<le> length \<tau>s" using mpcl mpcu by arith
- ultimately have
- "P,T,m,mpc - 1 + length [\<tau>],shift (length [\<tau>]) [] \<turnstile> i,pc - 1 + length [\<tau>] :: [\<tau>] @ \<tau>s"
- by(rule wt_instr_appL[where \<tau>s' = "[\<tau>]", OF wti])
- then show ?thesis using pcl mpcl by (simp split:nat_diff_split_asm)
-qed
-(*>*)
-
-
-lemma wt_instr_append:
-assumes wti: "P,T,m,mpc - size \<tau>s',[] \<turnstile> i,pc - size \<tau>s' :: \<tau>s"
- and pcl: "size \<tau>s' \<le> pc" and mpcl: "size \<tau>s' \<le> mpc"
- and pcu: "pc < size \<tau>s + size \<tau>s'" and mpcu: "mpc \<le> size \<tau>s + size \<tau>s'"
-shows "P,T,m,mpc,[] \<turnstile> i,pc :: \<tau>s'@\<tau>s"
-(*<*)
-proof -
- have "pc - length \<tau>s' < length \<tau>s" using pcl pcu by arith
- moreover have "mpc - length \<tau>s' \<le> length \<tau>s" using mpcl mpcu by arith
-thm wt_instr_appL[where \<tau>s' = "\<tau>s'", OF wti]
- ultimately have "P,T,m,mpc - length \<tau>s' + length \<tau>s',shift (length \<tau>s') []
- \<turnstile> i,pc - length \<tau>s' + length \<tau>s' :: \<tau>s' @ \<tau>s"
- by(rule wt_instr_appL[where \<tau>s' = "\<tau>s'", OF wti])
- then show ?thesis using pcl mpcl by (simp split:nat_diff_split_asm)
-qed
-(*>*)
-
-
-lemma xcpt_app_pcs:
- "pc \<notin> pcs xt \<Longrightarrow> xcpt_app i P pc mxs xt \<tau>"
-(*<*)
-by (auto simp add: xcpt_app_def relevant_entries_def is_relevant_entry_def pcs_def)
-(*>*)
-
-
-lemma xcpt_eff_pcs:
- "pc \<notin> pcs xt \<Longrightarrow> xcpt_eff i P pc \<tau> xt = []"
-(*<*)
-by (cases \<tau>)
- (auto simp add: is_relevant_entry_def xcpt_eff_def relevant_entries_def pcs_def
- intro!: filter_False)
-(*>*)
-
-
-lemma pcs_shift:
- "pc < n \<Longrightarrow> pc \<notin> pcs (shift n xt)"
-(*<*)by (auto simp add: shift_def pcs_def)(*>*)
-
-
-lemma wt_instr_appRx:
- "\<lbrakk> P,T,m,mpc,xt \<turnstile> is!pc,pc :: \<tau>s; pc < size is; size is < size \<tau>s; mpc \<le> size \<tau>s \<rbrakk>
- \<Longrightarrow> P,T,m,mpc,xt @ shift (size is) xt' \<turnstile> is!pc,pc :: \<tau>s"
-(*<*)by (auto simp:wt_instr_def eff_def app_def xcpt_app_pcs xcpt_eff_pcs)(*>*)
-
-
-lemma wt_instr_appLx:
- "\<lbrakk> P,T,m,mpc,xt \<turnstile> i,pc :: \<tau>s; pc \<notin> pcs xt' \<rbrakk>
- \<Longrightarrow> P,T,m,mpc,xt'@xt \<turnstile> i,pc :: \<tau>s"
-(*<*)by (auto simp:wt_instr_def app_def eff_def xcpt_app_pcs xcpt_eff_pcs)(*>*)
-
-
-lemma (in TC2) wt_instrs_extR:
- "\<turnstile> is,xt [::] \<tau>s \<Longrightarrow> \<turnstile> is,xt [::] \<tau>s @ \<tau>s'"
-(*<*)by(auto simp add:wt_instrs_def wt_instr_appR)(*>*)
-
-
-lemma (in TC2) wt_instrs_ext:
-assumes wt\<^sub>1: "\<turnstile> is\<^sub>1,xt\<^sub>1 [::] \<tau>s\<^sub>1@\<tau>s\<^sub>2" and wt\<^sub>2: "\<turnstile> is\<^sub>2,xt\<^sub>2 [::] \<tau>s\<^sub>2"
- and \<tau>s_size: "size \<tau>s\<^sub>1 = size is\<^sub>1"
- shows "\<turnstile> is\<^sub>1@is\<^sub>2, xt\<^sub>1 @ shift (size is\<^sub>1) xt\<^sub>2 [::] \<tau>s\<^sub>1@\<tau>s\<^sub>2"
-(*<*)
-proof -
- let ?is = "is\<^sub>1@is\<^sub>2" and ?xt = "xt\<^sub>1 @ shift (size is\<^sub>1) xt\<^sub>2"
- and ?\<tau>s = "\<tau>s\<^sub>1@\<tau>s\<^sub>2"
- have "size ?is < size ?\<tau>s" using wt\<^sub>2 \<tau>s_size by(fastforce simp:wt_instrs_def)
- moreover have "pcs ?xt \<subseteq> {0..<size ?is}" using wt\<^sub>1 wt\<^sub>2
- by(fastforce simp:wt_instrs_def)
- moreover {
- fix pc assume pc: "pc<size ?is"
- have "P,T\<^sub>r,mxs,size ?\<tau>s,?xt \<turnstile> ?is!pc,pc :: ?\<tau>s"
- proof(cases "pc < length is\<^sub>1")
- case True then show ?thesis using wt\<^sub>1 pc
- by(fastforce simp: wt_instrs_def wt_instr_appRx)
- next
- case False
- then have "pc - length is\<^sub>1 < length is\<^sub>2" using pc by fastforce
- then have "P,T\<^sub>r,mxs,length \<tau>s\<^sub>2,xt\<^sub>2 \<turnstile> is\<^sub>2 ! (pc - length is\<^sub>1),pc - length is\<^sub>1 :: \<tau>s\<^sub>2"
- using wt\<^sub>2 by(clarsimp simp: wt_instrs_def)
- moreover have "pc - length is\<^sub>1 < length \<tau>s\<^sub>2" using pc wt\<^sub>2
- by(clarsimp simp: wt_instrs_def) arith
- moreover have "length \<tau>s\<^sub>2 \<le> length \<tau>s\<^sub>2" by simp
- moreover have "pc - length is\<^sub>1 + length \<tau>s\<^sub>1 \<notin> pcs xt\<^sub>1" using wt\<^sub>1 \<tau>s_size
- by(fastforce simp: wt_instrs_def)
- ultimately have "P,T\<^sub>r,mxs,length \<tau>s\<^sub>2 + length \<tau>s\<^sub>1,xt\<^sub>1 @ shift (length \<tau>s\<^sub>1) xt\<^sub>2
- \<turnstile> is\<^sub>2 ! (pc - length is\<^sub>1),pc - length is\<^sub>1 + length \<tau>s\<^sub>1 :: \<tau>s\<^sub>1 @ \<tau>s\<^sub>2"
- by(rule wt_instr_appLx[OF wt_instr_appL[where \<tau>s' = "\<tau>s\<^sub>1"]])
- then show ?thesis using False \<tau>s_size by(simp add:add.commute)
- qed
- }
- ultimately show ?thesis by(clarsimp simp:wt_instrs_def)
-qed
-(*>*)
-
-corollary (in TC2) wt_instrs_ext2:
- "\<lbrakk> \<turnstile> is\<^sub>2,xt\<^sub>2 [::] \<tau>s\<^sub>2; \<turnstile> is\<^sub>1,xt\<^sub>1 [::] \<tau>s\<^sub>1@\<tau>s\<^sub>2; size \<tau>s\<^sub>1 = size is\<^sub>1 \<rbrakk>
- \<Longrightarrow> \<turnstile> is\<^sub>1@is\<^sub>2, xt\<^sub>1 @ shift (size is\<^sub>1) xt\<^sub>2 [::] \<tau>s\<^sub>1@\<tau>s\<^sub>2"
-(*<*)by(rule wt_instrs_ext)(*>*)
-
-
-corollary (in TC2) wt_instrs_ext_prefix [trans]:
- "\<lbrakk> \<turnstile> is\<^sub>1,xt\<^sub>1 [::] \<tau>s\<^sub>1@\<tau>s\<^sub>2; \<turnstile> is\<^sub>2,xt\<^sub>2 [::] \<tau>s\<^sub>3;
- size \<tau>s\<^sub>1 = size is\<^sub>1; prefix \<tau>s\<^sub>3 \<tau>s\<^sub>2 \<rbrakk>
- \<Longrightarrow> \<turnstile> is\<^sub>1@is\<^sub>2, xt\<^sub>1 @ shift (size is\<^sub>1) xt\<^sub>2 [::] \<tau>s\<^sub>1@\<tau>s\<^sub>2"
-(*<*)by(bestsimp simp:prefix_def elim: wt_instrs_ext dest:wt_instrs_extR)(*>*)
-
-
-corollary (in TC2) wt_instrs_app:
- assumes is\<^sub>1: "\<turnstile> is\<^sub>1,xt\<^sub>1 [::] \<tau>s\<^sub>1@[\<tau>]"
- assumes is\<^sub>2: "\<turnstile> is\<^sub>2,xt\<^sub>2 [::] \<tau>#\<tau>s\<^sub>2"
- assumes s: "size \<tau>s\<^sub>1 = size is\<^sub>1"
- shows "\<turnstile> is\<^sub>1@is\<^sub>2, xt\<^sub>1@shift (size is\<^sub>1) xt\<^sub>2 [::] \<tau>s\<^sub>1@\<tau>#\<tau>s\<^sub>2"
-(*<*)
-proof -
- from is\<^sub>1 have "\<turnstile> is\<^sub>1,xt\<^sub>1 [::] (\<tau>s\<^sub>1@[\<tau>])@\<tau>s\<^sub>2"
- by (rule wt_instrs_extR)
- hence "\<turnstile> is\<^sub>1,xt\<^sub>1 [::] \<tau>s\<^sub>1@\<tau>#\<tau>s\<^sub>2" by simp
- from this is\<^sub>2 s show ?thesis by (rule wt_instrs_ext)
-qed
-(*>*)
-
-
-corollary (in TC2) wt_instrs_app_last[trans]:
-assumes "\<turnstile> is\<^sub>2,xt\<^sub>2 [::] \<tau>#\<tau>s\<^sub>2" "\<turnstile> is\<^sub>1,xt\<^sub>1 [::] \<tau>s\<^sub>1"
- "last \<tau>s\<^sub>1 = \<tau>" "size \<tau>s\<^sub>1 = size is\<^sub>1+1"
-shows "\<turnstile> is\<^sub>1@is\<^sub>2, xt\<^sub>1@shift (size is\<^sub>1) xt\<^sub>2 [::] \<tau>s\<^sub>1@\<tau>s\<^sub>2"
-(*<*)
-using assms proof(cases \<tau>s\<^sub>1 rule:rev_cases)
- case (snoc ys y)
- then show ?thesis using assms by(simp add:wt_instrs_app)
-qed simp
-(*>*)
-
-
-corollary (in TC2) wt_instrs_append_last[trans]:
-assumes wtis: "\<turnstile> is,xt [::] \<tau>s" and wti: "P,T\<^sub>r,mxs,mpc,[] \<turnstile> i,pc :: \<tau>s"
- and pc: "pc = size is" and mpc: "mpc = size \<tau>s" and is_\<tau>s: "size is + 1 < size \<tau>s"
-shows "\<turnstile> is@[i],xt [::] \<tau>s"
-(*<*)
-proof -
- have pc_xt: "pc \<notin> pcs xt" using wtis pc by (fastforce simp:wt_instrs_def)
- have "pcs xt \<subseteq> {..<Suc (length is)}" using wtis by (fastforce simp:wt_instrs_def)
- moreover {
- fix pc' assume pc': "\<not> pc' < length is" "pc' < Suc (length is)"
- have "P,T\<^sub>r,mxs,length \<tau>s,xt \<turnstile> i,pc' :: \<tau>s"
- using wt_instr_appLx[where xt = "[]",simplified,OF wti pc_xt]
- less_antisym[OF pc'] pc mpc by simp
- }
- ultimately show ?thesis using wtis is_\<tau>s by(clarsimp simp add:wt_instrs_def)
-qed
-(*>*)
-
-
-corollary (in TC2) wt_instrs_app2:
- "\<lbrakk> \<turnstile> is\<^sub>2,xt\<^sub>2 [::] \<tau>'#\<tau>s\<^sub>2; \<turnstile> is\<^sub>1,xt\<^sub>1 [::] \<tau>#\<tau>s\<^sub>1@[\<tau>'];
- xt' = xt\<^sub>1 @ shift (size is\<^sub>1) xt\<^sub>2; size \<tau>s\<^sub>1+1 = size is\<^sub>1 \<rbrakk>
- \<Longrightarrow> \<turnstile> is\<^sub>1@is\<^sub>2,xt' [::] \<tau>#\<tau>s\<^sub>1@\<tau>'#\<tau>s\<^sub>2"
-(*<*)using wt_instrs_app[where ?\<tau>s\<^sub>1.0 = "\<tau> # \<tau>s\<^sub>1"] by simp (*>*)
-
-
-corollary (in TC2) wt_instrs_app2_simp[trans,simp]:
- "\<lbrakk> \<turnstile> is\<^sub>2,xt\<^sub>2 [::] \<tau>'#\<tau>s\<^sub>2; \<turnstile> is\<^sub>1,xt\<^sub>1 [::] \<tau>#\<tau>s\<^sub>1@[\<tau>']; size \<tau>s\<^sub>1+1 = size is\<^sub>1 \<rbrakk>
- \<Longrightarrow> \<turnstile> is\<^sub>1@is\<^sub>2, xt\<^sub>1@shift (size is\<^sub>1) xt\<^sub>2 [::] \<tau>#\<tau>s\<^sub>1@\<tau>'#\<tau>s\<^sub>2"
-(*<*)using wt_instrs_app[where ?\<tau>s\<^sub>1.0 = "\<tau> # \<tau>s\<^sub>1"] by simp(*>*)
-
-
-corollary (in TC2) wt_instrs_Cons[simp]:
- "\<lbrakk> \<tau>s \<noteq> []; \<turnstile> [i],[] [::] [\<tau>,\<tau>']; \<turnstile> is,xt [::] \<tau>'#\<tau>s \<rbrakk>
- \<Longrightarrow> \<turnstile> i#is,shift 1 xt [::] \<tau>#\<tau>'#\<tau>s"
-(*<*)
-using wt_instrs_app2[where ?is\<^sub>1.0 = "[i]" and ?\<tau>s\<^sub>1.0 = "[]" and ?is\<^sub>2.0 = "is"
- and ?xt\<^sub>1.0 = "[]"]
-by simp
-
-
-corollary (in TC2) wt_instrs_Cons2[trans]:
- assumes \<tau>s: "\<turnstile> is,xt [::] \<tau>s"
- assumes i: "P,T\<^sub>r,mxs,mpc,[] \<turnstile> i,0 :: \<tau>#\<tau>s"
- assumes mpc: "mpc = size \<tau>s + 1"
- shows "\<turnstile> i#is,shift 1 xt [::] \<tau>#\<tau>s"
-(*<*)
-proof -
- from \<tau>s have "\<tau>s \<noteq> []" by (auto simp: wt_instrs_def)
- with mpc i have "\<turnstile> [i],[] [::] [\<tau>]@\<tau>s" by (simp add: wt_instrs_def)
- with \<tau>s show ?thesis by (fastforce dest: wt_instrs_ext)
-qed
-(*>*)
-
-
-lemma (in TC2) wt_instrs_last_incr[trans]:
-assumes wtis: "\<turnstile> is,xt [::] \<tau>s@[\<tau>]" and ss: "P \<turnstile> \<tau> \<le>' \<tau>'"
-shows "\<turnstile> is,xt [::] \<tau>s@[\<tau>']"
-(*<*)
-proof -
- let ?\<tau>s = "\<tau>s@[\<tau>]" and ?\<tau>s' = "\<tau>s@[\<tau>']"
- { fix pc assume pc: "pc< size is"
- let ?i = "is!pc"
- have app_pc: "app (is ! pc) P mxs T\<^sub>r pc (length ?\<tau>s) xt (\<tau>s ! pc)"
- using wtis pc by(clarsimp simp add:wt_instrs_def wt_instr_def)
- then have Apc\<tau>': "\<And>pc' \<tau>'. (pc',\<tau>') \<in> set (eff ?i P pc xt (?\<tau>s!pc))
- \<Longrightarrow> pc' < length ?\<tau>s"
- using wtis pc by(fastforce simp add:wt_instrs_def app_def)
- have Aepc\<tau>': "\<And>pc' \<tau>'. (pc',\<tau>') \<in> set (eff ?i P pc xt (?\<tau>s!pc))
- \<Longrightarrow> P \<turnstile> \<tau>' \<le>' ?\<tau>s!pc'"
- using wtis pc by(fastforce simp add:wt_instrs_def wt_instr_def)
- { fix pc1 \<tau>1 assume pc\<tau>1: "(pc1,\<tau>1) \<in> set (eff ?i P pc xt (?\<tau>s'!pc))"
- then have epc\<tau>': "(pc1,\<tau>1) \<in> set (eff ?i P pc xt (?\<tau>s!pc))"
- using wtis pc by(simp add:wt_instrs_def)
- have "P \<turnstile> \<tau>1 \<le>' ?\<tau>s'!pc1"
- proof(cases "pc1 < length \<tau>s")
- case True
- then show ?thesis using wtis pc pc\<tau>1
- by(fastforce simp add:wt_instrs_def wt_instr_def)
- next
- case False
- then have "pc1 < length ?\<tau>s" using Apc\<tau>'[OF epc\<tau>'] by simp
- then have [simp]: "pc1 = size \<tau>s" using False by clarsimp
- have "P \<turnstile> \<tau>1 \<le>' \<tau>" using Aepc\<tau>'[OF epc\<tau>'] by simp
- then have "P \<turnstile> \<tau>1 \<le>' \<tau>'" by(rule sup_state_opt_trans[OF _ ss])
- then show ?thesis by simp
- qed
- }
- then have "P,T\<^sub>r,mxs,size ?\<tau>s',xt \<turnstile> is!pc,pc :: ?\<tau>s'" using wtis pc
- by(clarsimp simp add:wt_instrs_def wt_instr_def)
- }
- then show ?thesis using wtis by(simp add:wt_instrs_def)
-qed
-(*>*)
-
-
-lemma [iff]: "xcpt_app i P pc mxs [] \<tau>"
-(*<*)by (simp add: xcpt_app_def relevant_entries_def)(*>*)
-
-
-lemma [simp]: "xcpt_eff i P pc \<tau> [] = []"
-(*<*)by (simp add: xcpt_eff_def relevant_entries_def)(*>*)
-
-
-lemma (in TC2) wt_New:
- "\<lbrakk> is_class P C; size ST < mxs \<rbrakk> \<Longrightarrow>
- \<turnstile> [New C],[] [::] [ty\<^sub>i' ST E A, ty\<^sub>i' (Class C#ST) E A]"
-(*<*)by(simp add:wt_defs ty\<^sub>i'_def)(*>*)
-
-
-lemma (in TC2) wt_Cast:
- "is_class P C \<Longrightarrow>
- \<turnstile> [Checkcast C],[] [::] [ty\<^sub>i' (Class D # ST) E A, ty\<^sub>i' (Class C # ST) E A]"
-(*<*)by(simp add: ty\<^sub>i'_def wt_defs)(*>*)
-
-
-lemma (in TC2) wt_Push:
- "\<lbrakk> size ST < mxs; typeof v = Some T \<rbrakk>
- \<Longrightarrow> \<turnstile> [Push v],[] [::] [ty\<^sub>i' ST E A, ty\<^sub>i' (T#ST) E A]"
-(*<*)by(simp add: ty\<^sub>i'_def wt_defs)(*>*)
-
-
-lemma (in TC2) wt_Pop:
- "\<turnstile> [Pop],[] [::] (ty\<^sub>i' (T#ST) E A # ty\<^sub>i' ST E A # \<tau>s)"
-(*<*)by(simp add: ty\<^sub>i'_def wt_defs)(*>*)
-
-
-lemma (in TC2) wt_CmpEq:
- "\<lbrakk> P \<turnstile> T\<^sub>1 \<le> T\<^sub>2 \<or> P \<turnstile> T\<^sub>2 \<le> T\<^sub>1\<rbrakk>
- \<Longrightarrow> \<turnstile> [CmpEq],[] [::] [ty\<^sub>i' (T\<^sub>2 # T\<^sub>1 # ST) E A, ty\<^sub>i' (Boolean # ST) E A]"
-(*<*) by(auto simp:ty\<^sub>i'_def wt_defs elim!: refTE not_refTE) (*>*)
-
-
-lemma (in TC2) wt_IAdd:
- "\<turnstile> [IAdd],[] [::] [ty\<^sub>i' (Integer#Integer#ST) E A, ty\<^sub>i' (Integer#ST) E A]"
-(*<*)by(simp add:ty\<^sub>i'_def wt_defs)(*>*)
-
-
-lemma (in TC2) wt_Load:
- "\<lbrakk> size ST < mxs; size E \<le> mxl; i \<in>\<in> A; i < size E \<rbrakk>
- \<Longrightarrow> \<turnstile> [Load i],[] [::] [ty\<^sub>i' ST E A, ty\<^sub>i' (E!i # ST) E A]"
-(*<*)by(auto simp add:ty\<^sub>i'_def wt_defs ty\<^sub>l_def hyperset_defs)(*>*)
-
-
-lemma (in TC2) wt_Store:
- "\<lbrakk> P \<turnstile> T \<le> E!i; i < size E; size E \<le> mxl \<rbrakk> \<Longrightarrow>
- \<turnstile> [Store i],[] [::] [ty\<^sub>i' (T#ST) E A, ty\<^sub>i' ST E (\<lfloor>{i}\<rfloor> \<squnion> A)]"
-(*<*)
-by(auto simp:hyperset_defs nth_list_update ty\<^sub>i'_def wt_defs ty\<^sub>l_def
- intro:list_all2_all_nthI)
-(*>*)
-
-
-lemma (in TC2) wt_Get:
- "\<lbrakk> P \<turnstile> C sees F,NonStatic:T in D \<rbrakk> \<Longrightarrow>
- \<turnstile> [Getfield F D],[] [::] [ty\<^sub>i' (Class C # ST) E A, ty\<^sub>i' (T # ST) E A]"
-(*<*)by(auto simp: ty\<^sub>i'_def wt_defs dest: sees_field_idemp sees_field_decl_above)(*>*)
-
-lemma (in TC2) wt_GetS:
- "\<lbrakk> size ST < mxs; P \<turnstile> C sees F,Static:T in D \<rbrakk> \<Longrightarrow>
- \<turnstile> [Getstatic C F D],[] [::] [ty\<^sub>i' ST E A, ty\<^sub>i' (T # ST) E A]"
-(*<*)by(auto simp: ty\<^sub>i'_def wt_defs dest: sees_field_idemp sees_field_decl_above)(*>*)
-
-lemma (in TC2) wt_Put:
- "\<lbrakk> P \<turnstile> C sees F,NonStatic:T in D; P \<turnstile> T' \<le> T \<rbrakk> \<Longrightarrow>
- \<turnstile> [Putfield F D],[] [::] [ty\<^sub>i' (T' # Class C # ST) E A, ty\<^sub>i' ST E A]"
-(*<*)by(auto intro: sees_field_idemp sees_field_decl_above simp: ty\<^sub>i'_def wt_defs)(*>*)
-
-lemma (in TC2) wt_PutS:
- "\<lbrakk> P \<turnstile> C sees F,Static:T in D; P \<turnstile> T' \<le> T \<rbrakk> \<Longrightarrow>
- \<turnstile> [Putstatic C F D],[] [::] [ty\<^sub>i' (T' # ST) E A, ty\<^sub>i' ST E A]"
-(*<*)by(auto intro: sees_field_idemp sees_field_decl_above simp: ty\<^sub>i'_def wt_defs)(*>*)
-
-
-lemma (in TC2) wt_Throw:
- "\<turnstile> [Throw],[] [::] [ty\<^sub>i' (Class C # ST) E A, \<tau>']"
-(*<*)by(auto simp: ty\<^sub>i'_def wt_defs)(*>*)
-
-
-lemma (in TC2) wt_IfFalse:
- "\<lbrakk> 2 \<le> i; nat i < size \<tau>s + 2; P \<turnstile> ty\<^sub>i' ST E A \<le>' \<tau>s ! nat(i - 2) \<rbrakk>
- \<Longrightarrow> \<turnstile> [IfFalse i],[] [::] ty\<^sub>i' (Boolean # ST) E A # ty\<^sub>i' ST E A # \<tau>s"
-(*<*)
-by(simp add: ty\<^sub>i'_def wt_defs eval_nat_numeral nat_diff_distrib)
-(*>*)
-
-
-lemma wt_Goto:
- "\<lbrakk> 0 \<le> int pc + i; nat (int pc + i) < size \<tau>s; size \<tau>s \<le> mpc;
- P \<turnstile> \<tau>s!pc \<le>' \<tau>s ! nat (int pc + i) \<rbrakk>
- \<Longrightarrow> P,T,mxs,mpc,[] \<turnstile> Goto i,pc :: \<tau>s"
-(*<*)by(clarsimp simp add: TC2.wt_defs)(*>*)
-
-
-lemma (in TC2) wt_Invoke:
- "\<lbrakk> size es = size Ts'; P \<turnstile> C sees M,NonStatic: Ts\<rightarrow>T = m in D; P \<turnstile> Ts' [\<le>] Ts \<rbrakk>
- \<Longrightarrow> \<turnstile> [Invoke M (size es)],[] [::] [ty\<^sub>i' (rev Ts' @ Class C # ST) E A, ty\<^sub>i' (T#ST) E A]"
-(*<*)by(fastforce simp add: ty\<^sub>i'_def wt_defs)(*>*)
-
-lemma (in TC2) wt_Invokestatic:
- "\<lbrakk> size ST < mxs; size es = size Ts'; M \<noteq> clinit;
- P \<turnstile> C sees M,Static: Ts\<rightarrow>T = m in D; P \<turnstile> Ts' [\<le>] Ts \<rbrakk>
- \<Longrightarrow> \<turnstile> [Invokestatic C M (size es)],[] [::] [ty\<^sub>i' (rev Ts' @ ST) E A, ty\<^sub>i' (T#ST) E A]"
-(*<*)by(fastforce simp add: ty\<^sub>i'_def wt_defs)(*>*)
-
-
-corollary (in TC2) wt_instrs_app3[simp]:
- "\<lbrakk> \<turnstile> is\<^sub>2,[] [::] (\<tau>' # \<tau>s\<^sub>2); \<turnstile> is\<^sub>1,xt\<^sub>1 [::] \<tau> # \<tau>s\<^sub>1 @ [\<tau>']; size \<tau>s\<^sub>1+1 = size is\<^sub>1\<rbrakk>
- \<Longrightarrow> \<turnstile> (is\<^sub>1 @ is\<^sub>2),xt\<^sub>1 [::] \<tau> # \<tau>s\<^sub>1 @ \<tau>' # \<tau>s\<^sub>2"
-(*<*)using wt_instrs_app2[where ?xt\<^sub>2.0 = "[]"] by (simp add:shift_def)(*>*)
-
-
-corollary (in TC2) wt_instrs_Cons3[simp]:
- "\<lbrakk> \<tau>s \<noteq> []; \<turnstile> [i],[] [::] [\<tau>,\<tau>']; \<turnstile> is,[] [::] \<tau>'#\<tau>s \<rbrakk>
- \<Longrightarrow> \<turnstile> (i # is),[] [::] \<tau> # \<tau>' # \<tau>s"
-(*<*)
-using wt_instrs_Cons[where ?xt = "[]"]
-by (simp add:shift_def)
-
-(*<*)
-declare nth_append[simp del]
-declare [[simproc del: list_to_set_comprehension]]
-(*>*)
-
-lemma (in TC2) wt_instrs_xapp[trans]:
-assumes wtis: "\<turnstile> is\<^sub>1 @ is\<^sub>2, xt [::] \<tau>s\<^sub>1 @ ty\<^sub>i' (Class C # ST) E A # \<tau>s\<^sub>2"
- and P_\<tau>s\<^sub>1: "\<forall>\<tau> \<in> set \<tau>s\<^sub>1. \<forall>ST' LT'. \<tau> = Some(ST',LT') \<longrightarrow>
- size ST \<le> size ST' \<and> P \<turnstile> Some (drop (size ST' - size ST) ST',LT') \<le>' ty\<^sub>i' ST E A"
- and is_\<tau>s: "size is\<^sub>1 = size \<tau>s\<^sub>1" and PC: "is_class P C" and ST_mxs: "size ST < mxs"
-shows "\<turnstile> is\<^sub>1 @ is\<^sub>2, xt @ [(0,size is\<^sub>1 - 1,C,size is\<^sub>1,size ST)] [::] \<tau>s\<^sub>1 @ ty\<^sub>i' (Class C # ST) E A # \<tau>s\<^sub>2"
-(*<*)(is "\<turnstile> ?is, xt@[?xte] [::] ?\<tau>s" is "\<turnstile> ?is, ?xt' [::] ?\<tau>s")
-proof -
- have P_\<tau>s\<^sub>1': "\<And>\<tau>. \<tau> \<in> set \<tau>s\<^sub>1 \<Longrightarrow> (\<forall>ST' LT'. \<tau> = Some(ST',LT') \<longrightarrow>
- size ST \<le> size ST' \<and> P \<turnstile> Some (drop (size ST' - size ST) ST',LT') \<le>' ty\<^sub>i' ST E A)"
- using P_\<tau>s\<^sub>1 by fast
- have "size ?is < size ?\<tau>s" and "pcs ?xt' \<subseteq> {0..<size ?is}"
- and P_pc: "\<And>pc. pc< size ?is \<Longrightarrow> P,T\<^sub>r,mxs,size ?\<tau>s,xt \<turnstile> ?is!pc,pc :: ?\<tau>s"
- using wtis by(simp_all add:wt_instrs_def)
- moreover {
- fix pc let ?mpc = "size ?\<tau>s" and ?i = "?is!pc" and ?t = "?\<tau>s!pc"
- assume "pc< size ?is"
- then have wti: "P,T\<^sub>r,mxs,?mpc,xt \<turnstile> ?i,pc :: ?\<tau>s" by(rule P_pc)
- then have app: "app ?i P mxs T\<^sub>r pc ?mpc xt ?t"
- and eff_ss: "\<And>pc' \<tau>'. (pc',\<tau>') \<in> set (eff ?i P pc xt (?\<tau>s!pc))
- \<Longrightarrow> P \<turnstile> \<tau>' \<le>' ?\<tau>s!pc'"
- by(fastforce simp add: wt_instr_def)+
- have "app ?i P mxs T\<^sub>r pc ?mpc ?xt' ?t
- \<and> (\<forall>(pc',\<tau>') \<in> set (eff ?i P pc ?xt' ?t). P \<turnstile> \<tau>' \<le>' ?\<tau>s!pc')"
- proof (cases ?t)
- case (Some \<tau>)
- obtain ST' LT' where \<tau>: "\<tau> = (ST', LT')" by(cases \<tau>) simp
- have app\<^sub>i: "app\<^sub>i (?i,P,pc,mxs,T\<^sub>r,\<tau>)" and xcpt_app: "xcpt_app ?i P pc mxs xt \<tau>"
- and eff_pc: "\<And>pc' \<tau>'. (pc',\<tau>') \<in> set (eff ?i P pc xt ?t) \<Longrightarrow> pc' < ?mpc"
- using Some app by(fastforce simp add: app_def)+
- have "xcpt_app ?i P pc mxs ?xt' \<tau>"
- proof(cases "pc < length \<tau>s\<^sub>1 - 1")
- case False then show ?thesis using Some \<tau> is_\<tau>s xcpt_app
- by (clarsimp simp: xcpt_app_def relevant_entries_def
- is_relevant_entry_def)
- next
- case True
- then have True': "pc < length \<tau>s\<^sub>1" by simp
- then have "\<tau>s\<^sub>1 ! pc = ?t" by(fastforce simp: nth_append)
- moreover have \<tau>s\<^sub>1_pc: "\<tau>s\<^sub>1 ! pc \<in> set \<tau>s\<^sub>1" by(rule nth_mem[OF True'])
- ultimately show ?thesis
- using Some \<tau> PC ST_mxs xcpt_app P_\<tau>s\<^sub>1'[OF \<tau>s\<^sub>1_pc]
- by (simp add: xcpt_app_def relevant_entries_def)
- qed
- moreover {
- fix pc' \<tau>' assume efft: "(pc',\<tau>') \<in> set (eff ?i P pc ?xt' ?t)"
- have "pc' < ?mpc \<and> P \<turnstile> \<tau>' \<le>' ?\<tau>s!pc'" (is "?P1 \<and> ?P2")
- proof(cases "(pc',\<tau>') \<in> set (eff ?i P pc xt ?t)")
- case True
- have ?P1 using True by(rule eff_pc)
- moreover have ?P2 using True by(rule eff_ss)
- ultimately show ?thesis by simp
- next
- case False
- then have xte: "(pc',\<tau>') \<in> set (xcpt_eff ?i P pc \<tau> [?xte])"
- using efft Some by(clarsimp simp: eff_def)
- then have ?P1 using Some \<tau> is_\<tau>s
- by(clarsimp simp: xcpt_eff_def relevant_entries_def split: if_split_asm)
- moreover have ?P2
- proof(cases "pc < length \<tau>s\<^sub>1 - 1")
- case False': False
- then show ?thesis using False Some \<tau> xte is_\<tau>s
- by(simp add: xcpt_eff_def relevant_entries_def is_relevant_entry_def)
- next
- case True
- then have True': "pc < length \<tau>s\<^sub>1" by simp
- have \<tau>s\<^sub>1_pc: "\<tau>s\<^sub>1 ! pc \<in> set \<tau>s\<^sub>1" by(rule nth_mem[OF True'])
- have "P \<turnstile> \<lfloor>(Class C # drop (length ST' - length ST) ST', LT')\<rfloor>
- \<le>' ty\<^sub>i' (Class C # ST) E A"
- using True' Some \<tau> P_\<tau>s\<^sub>1'[OF \<tau>s\<^sub>1_pc]
- by (fastforce simp: nth_append ty\<^sub>i'_def)
- then show ?thesis using \<tau> xte is_\<tau>s
- by(simp add: xcpt_eff_def relevant_entries_def split: if_split_asm)
- qed
- ultimately show ?thesis by simp
- qed
- }
- ultimately show ?thesis using Some app\<^sub>i by(fastforce simp add: app_def)
- qed simp
- then have "P,T\<^sub>r,mxs,size ?\<tau>s,?xt' \<turnstile> ?is!pc,pc :: ?\<tau>s"
- by(simp add: wt_instr_def)
- }
- ultimately show ?thesis by(simp add:wt_instrs_def)
-qed
-
-declare [[simproc add: list_to_set_comprehension]]
-declare nth_append[simp]
-(*>*)
-
-lemma drop_Cons_Suc:
- "\<And>xs. drop n xs = y#ys \<Longrightarrow> drop (Suc n) xs = ys"
-proof(induct n)
- case (Suc n) then show ?case by(simp add: drop_Suc)
-qed simp
-
-lemma drop_mess:
-assumes "Suc (length xs\<^sub>0) \<le> length xs"
- and "drop (length xs - Suc (length xs\<^sub>0)) xs = x # xs\<^sub>0"
-shows "drop (length xs - length xs\<^sub>0) xs = xs\<^sub>0"
-using assms proof(cases xs)
- case (Cons a list) then show ?thesis using assms
- proof(cases "length list - length xs\<^sub>0")
- case Suc then show ?thesis using Cons assms
- by (simp add: Suc_diff_le drop_Cons_Suc)
- qed simp
-qed simp
-
-(*<*)
-declare (in TC0)
- after_def[simp] pair_eq_ty\<^sub>i'_conv[simp]
-(*>*)
-
-lemma (in TC1) compT_ST_prefix:
- "\<And>E A ST\<^sub>0. \<lfloor>(ST,LT)\<rfloor> \<in> set(compT E A ST\<^sub>0 e) \<Longrightarrow>
- size ST\<^sub>0 \<le> size ST \<and> drop (size ST - size ST\<^sub>0) ST = ST\<^sub>0"
-and
- "\<And>E A ST\<^sub>0. \<lfloor>(ST,LT)\<rfloor> \<in> set(compTs E A ST\<^sub>0 es) \<Longrightarrow>
- size ST\<^sub>0 \<le> size ST \<and> drop (size ST - size ST\<^sub>0) ST = ST\<^sub>0"
-(*<*)
-proof(induct e and es rule: compT.induct compTs.induct)
- case (FAss e\<^sub>1 F D e\<^sub>2)
- moreover {
- let ?ST\<^sub>0 = "ty E e\<^sub>1 # ST\<^sub>0"
- fix A assume "\<lfloor>(ST, LT)\<rfloor> \<in> set (compT E A ?ST\<^sub>0 e\<^sub>2)"
- with FAss
- have "length ?ST\<^sub>0 \<le> length ST \<and> drop (size ST - size ?ST\<^sub>0) ST = ?ST\<^sub>0" by blast
- hence ?case by (clarsimp simp add: drop_mess)
- }
- ultimately show ?case by auto
-next
- case TryCatch thus ?case by auto
-next
- case Block thus ?case by auto
-next
- case Seq thus ?case by auto
-next
- case While thus ?case by auto
-next
- case Cond thus ?case by auto
-next
- case (Call e M es)
- moreover {
- let ?ST\<^sub>0 = "ty E e # ST\<^sub>0"
- fix A assume "\<lfloor>(ST, LT)\<rfloor> \<in> set (compTs E A ?ST\<^sub>0 es)"
- with Call
- have "length ?ST\<^sub>0 \<le> length ST \<and> drop (size ST - size ?ST\<^sub>0) ST = ?ST\<^sub>0" by blast
- hence ?case by (clarsimp simp add: drop_mess)
- }
- ultimately show ?case by auto
-next
- case (Cons_exp e es)
- moreover {
- let ?ST\<^sub>0 = "ty E e # ST\<^sub>0"
- fix A assume "\<lfloor>(ST, LT)\<rfloor> \<in> set (compTs E A ?ST\<^sub>0 es)"
- with Cons_exp
- have "length ?ST\<^sub>0 \<le> length ST \<and> drop (size ST - size ?ST\<^sub>0) ST = ?ST\<^sub>0" by blast
- hence ?case by (clarsimp simp add: drop_mess)
- }
- ultimately show ?case by auto
-next
- case (BinOp e\<^sub>1 bop e\<^sub>2)
- moreover {
- let ?ST\<^sub>0 = "ty E e\<^sub>1 # ST\<^sub>0"
- fix A assume "\<lfloor>(ST, LT)\<rfloor> \<in> set (compT E A ?ST\<^sub>0 e\<^sub>2)"
- with BinOp
- have "length ?ST\<^sub>0 \<le> length ST \<and> drop (size ST - size ?ST\<^sub>0) ST = ?ST\<^sub>0" by blast
- hence ?case by (clarsimp simp add: drop_mess)
- }
- ultimately show ?case by auto
-next
- case new thus ?case by auto
-next
- case Val thus ?case by auto
-next
- case Cast thus ?case by auto
-next
- case Var thus ?case by auto
-next
- case LAss thus ?case by auto
-next
- case throw thus ?case by auto
-next
- case FAcc thus ?case by auto
-next
- case SFAcc thus ?case by auto
-next
- case SFAss thus ?case by auto
-next
- case (SCall C M es) thus ?case by auto
-next
- case INIT thus ?case by auto
-next
- case RI thus ?case by auto
-next
- case Nil_exp thus ?case by auto
-qed
-
-declare (in TC0)
- after_def[simp del] pair_eq_ty\<^sub>i'_conv[simp del]
-(*>*)
-
-(* FIXME *)
-lemma fun_of_simp [simp]: "fun_of S x y = ((x,y) \<in> S)"
-(*<*) by (simp add: fun_of_def)(*>*)
-
-theorem (in TC2) compT_wt_instrs: "\<And>E T A ST.
- \<lbrakk> P,E \<turnstile>\<^sub>1 e :: T; \<D> e A; \<B> e (size E);
- size ST + max_stack e \<le> mxs; size E + max_vars e \<le> mxl \<rbrakk>
- \<Longrightarrow> \<turnstile> compE\<^sub>2 e, compxE\<^sub>2 e 0 (size ST) [::]
- ty\<^sub>i' ST E A # compT E A ST e @ [after E A ST e]"
-(*<*)(is "\<And>E T A ST. PROP ?P e E T A ST")(*>*)
-
-and "\<And>E Ts A ST.
- \<lbrakk> P,E \<turnstile>\<^sub>1 es[::]Ts; \<D>s es A; \<B>s es (size E);
- size ST + max_stacks es \<le> mxs; size E + max_varss es \<le> mxl \<rbrakk>
- \<Longrightarrow> let \<tau>s = ty\<^sub>i' ST E A # compTs E A ST es in
- \<turnstile> compEs\<^sub>2 es,compxEs\<^sub>2 es 0 (size ST) [::] \<tau>s \<and>
- last \<tau>s = ty\<^sub>i' (rev Ts @ ST) E (A \<squnion> \<A>s es)"
-(*<*)
-(is "\<And>E Ts A ST. PROP ?Ps es E Ts A ST")
-proof(induct e and es rule: compxE\<^sub>2.induct compxEs\<^sub>2.induct)
- case (TryCatch e\<^sub>1 C i e\<^sub>2)
- hence [simp]: "i = size E" by simp
- have wt\<^sub>1: "P,E \<turnstile>\<^sub>1 e\<^sub>1 :: T" and wt\<^sub>2: "P,E@[Class C] \<turnstile>\<^sub>1 e\<^sub>2 :: T"
- and "class": "is_class P C" using TryCatch by auto
- let ?A\<^sub>1 = "A \<squnion> \<A> e\<^sub>1" let ?A\<^sub>i = "A \<squnion> \<lfloor>{i}\<rfloor>" let ?E\<^sub>i = "E @ [Class C]"
- let ?\<tau> = "ty\<^sub>i' ST E A" let ?\<tau>s\<^sub>1 = "compT E A ST e\<^sub>1"
- let ?\<tau>\<^sub>1 = "ty\<^sub>i' (T#ST) E ?A\<^sub>1" let ?\<tau>\<^sub>2 = "ty\<^sub>i' (Class C#ST) E A"
- let ?\<tau>\<^sub>3 = "ty\<^sub>i' ST ?E\<^sub>i ?A\<^sub>i" let ?\<tau>s\<^sub>2 = "compT ?E\<^sub>i ?A\<^sub>i ST e\<^sub>2"
- let ?\<tau>\<^sub>2' = "ty\<^sub>i' (T#ST) ?E\<^sub>i (?A\<^sub>i \<squnion> \<A> e\<^sub>2)"
- let ?\<tau>' = "ty\<^sub>i' (T#ST) E (A \<squnion> \<A> e\<^sub>1 \<sqinter> (\<A> e\<^sub>2 \<ominus> i))"
- let ?go = "Goto (int(size(compE\<^sub>2 e\<^sub>2)) + 2)"
- have "PROP ?P e\<^sub>2 ?E\<^sub>i T ?A\<^sub>i ST" by fact
- hence "\<turnstile> compE\<^sub>2 e\<^sub>2,compxE\<^sub>2 e\<^sub>2 0 (size ST) [::] (?\<tau>\<^sub>3 # ?\<tau>s\<^sub>2) @ [?\<tau>\<^sub>2']"
- using TryCatch.prems by(auto simp:after_def)
- also have "?A\<^sub>i \<squnion> \<A> e\<^sub>2 = (A \<squnion> \<A> e\<^sub>2) \<squnion> \<lfloor>{size E}\<rfloor>"
- by(fastforce simp:hyperset_defs)
- also have "P \<turnstile> ty\<^sub>i' (T#ST) ?E\<^sub>i \<dots> \<le>' ty\<^sub>i' (T#ST) E (A \<squnion> \<A> e\<^sub>2)"
- by(simp add:hyperset_defs ty\<^sub>l_incr ty\<^sub>i'_def)
- also have "P \<turnstile> \<dots> \<le>' ty\<^sub>i' (T#ST) E (A \<squnion> \<A> e\<^sub>1 \<sqinter> (\<A> e\<^sub>2 \<ominus> i))"
- by(auto intro!: ty\<^sub>l_antimono simp:hyperset_defs ty\<^sub>i'_def)
- also have "(?\<tau>\<^sub>3 # ?\<tau>s\<^sub>2) @ [?\<tau>'] = ?\<tau>\<^sub>3 # ?\<tau>s\<^sub>2 @ [?\<tau>']" by simp
- also have "\<turnstile> [Store i],[] [::] ?\<tau>\<^sub>2 # [] @ [?\<tau>\<^sub>3]"
- using TryCatch.prems
- by(auto simp:nth_list_update wt_defs ty\<^sub>i'_def ty\<^sub>l_def
- list_all2_conv_all_nth hyperset_defs)
- also have "[] @ (?\<tau>\<^sub>3 # ?\<tau>s\<^sub>2 @ [?\<tau>']) = (?\<tau>\<^sub>3 # ?\<tau>s\<^sub>2 @ [?\<tau>'])" by simp
- also have "P,T\<^sub>r,mxs,size(compE\<^sub>2 e\<^sub>2)+3,[] \<turnstile> ?go,0 :: ?\<tau>\<^sub>1#?\<tau>\<^sub>2#?\<tau>\<^sub>3#?\<tau>s\<^sub>2 @ [?\<tau>']" using wt\<^sub>2
- by (auto simp: hyperset_defs ty\<^sub>i'_def wt_defs nth_Cons nat_add_distrib
- fun_of_def intro: ty\<^sub>l_antimono list_all2_refl split:nat.split)
- also have "\<turnstile> compE\<^sub>2 e\<^sub>1,compxE\<^sub>2 e\<^sub>1 0 (size ST) [::] ?\<tau> # ?\<tau>s\<^sub>1 @ [?\<tau>\<^sub>1]"
- using TryCatch by(auto simp:after_def)
- also have "?\<tau> # ?\<tau>s\<^sub>1 @ ?\<tau>\<^sub>1 # ?\<tau>\<^sub>2 # ?\<tau>\<^sub>3 # ?\<tau>s\<^sub>2 @ [?\<tau>'] =
- (?\<tau> # ?\<tau>s\<^sub>1 @ [?\<tau>\<^sub>1]) @ ?\<tau>\<^sub>2 # ?\<tau>\<^sub>3 # ?\<tau>s\<^sub>2 @ [?\<tau>']" by simp
- also have "compE\<^sub>2 e\<^sub>1 @ ?go # [Store i] @ compE\<^sub>2 e\<^sub>2 =
- (compE\<^sub>2 e\<^sub>1 @ [?go]) @ (Store i # compE\<^sub>2 e\<^sub>2)" by simp
- also
- let "?Q \<tau>" = "\<forall>ST' LT'. \<tau> = \<lfloor>(ST', LT')\<rfloor> \<longrightarrow>
- size ST \<le> size ST' \<and> P \<turnstile> Some (drop (size ST' - size ST) ST',LT') \<le>' ty\<^sub>i' ST E A"
- {
- have "?Q (ty\<^sub>i' ST E A)" by (clarsimp simp add: ty\<^sub>i'_def)
- moreover have "?Q (ty\<^sub>i' (T # ST) E ?A\<^sub>1)"
- by (fastforce simp add: ty\<^sub>i'_def hyperset_defs intro!: ty\<^sub>l_antimono)
- moreover have "\<And>\<tau>. \<tau> \<in> set (compT E A ST e\<^sub>1) \<Longrightarrow> ?Q \<tau>" using TryCatch.prems
- by clarsimp (frule compT_ST_prefix,
- fastforce dest!: compT_LT_prefix simp add: ty\<^sub>i'_def)
- ultimately
- have "\<forall>\<tau>\<in>set (ty\<^sub>i' ST E A # compT E A ST e\<^sub>1 @ [ty\<^sub>i' (T # ST) E ?A\<^sub>1]). ?Q \<tau>"
- by auto
- }
- also from TryCatch.prems max_stack1[OF wt\<^sub>1] have "size ST + 1 \<le> mxs" by auto
- ultimately show ?case using wt\<^sub>1 wt\<^sub>2 TryCatch.prems "class"
- by (simp add:after_def)
-next
- case new thus ?case by(auto simp add:after_def wt_New)
-next
- case (BinOp e\<^sub>1 bop e\<^sub>2)
- let ?op = "case bop of Eq \<Rightarrow> [CmpEq] | Add \<Rightarrow> [IAdd]"
- have T: "P,E \<turnstile>\<^sub>1 e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2 :: T" by fact
- then obtain T\<^sub>1 T\<^sub>2 where T\<^sub>1: "P,E \<turnstile>\<^sub>1 e\<^sub>1 :: T\<^sub>1" and T\<^sub>2: "P,E \<turnstile>\<^sub>1 e\<^sub>2 :: T\<^sub>2" and
- bopT: "case bop of Eq \<Rightarrow> (P \<turnstile> T\<^sub>1 \<le> T\<^sub>2 \<or> P \<turnstile> T\<^sub>2 \<le> T\<^sub>1) \<and> T = Boolean
- | Add \<Rightarrow> T\<^sub>1 = Integer \<and> T\<^sub>2 = Integer \<and> T = Integer" by auto
- let ?A\<^sub>1 = "A \<squnion> \<A> e\<^sub>1" let ?A\<^sub>2 = "?A\<^sub>1 \<squnion> \<A> e\<^sub>2"
- let ?\<tau> = "ty\<^sub>i' ST E A" let ?\<tau>s\<^sub>1 = "compT E A ST e\<^sub>1"
- let ?\<tau>\<^sub>1 = "ty\<^sub>i' (T\<^sub>1#ST) E ?A\<^sub>1" let ?\<tau>s\<^sub>2 = "compT E ?A\<^sub>1 (T\<^sub>1#ST) e\<^sub>2"
- let ?\<tau>\<^sub>2 = "ty\<^sub>i' (T\<^sub>2#T\<^sub>1#ST) E ?A\<^sub>2" let ?\<tau>' = "ty\<^sub>i' (T#ST) E ?A\<^sub>2"
- from bopT have "\<turnstile> ?op,[] [::] [?\<tau>\<^sub>2,?\<tau>']"
- by (cases bop) (auto simp add: wt_CmpEq wt_IAdd)
- also have "PROP ?P e\<^sub>2 E T\<^sub>2 ?A\<^sub>1 (T\<^sub>1#ST)" by fact
- with BinOp.prems T\<^sub>2
- have "\<turnstile> compE\<^sub>2 e\<^sub>2, compxE\<^sub>2 e\<^sub>2 0 (size (T\<^sub>1#ST)) [::] ?\<tau>\<^sub>1#?\<tau>s\<^sub>2@[?\<tau>\<^sub>2]"
- by (auto simp: after_def)
- also from BinOp T\<^sub>1 have "\<turnstile> compE\<^sub>2 e\<^sub>1, compxE\<^sub>2 e\<^sub>1 0 (size ST) [::] ?\<tau>#?\<tau>s\<^sub>1@[?\<tau>\<^sub>1]"
- by (auto simp: after_def)
- finally show ?case using T T\<^sub>1 T\<^sub>2 by (simp add: after_def hyperUn_assoc)
-next
- case (Cons_exp e es)
- have "P,E \<turnstile>\<^sub>1 e # es [::] Ts" by fact
- then obtain T\<^sub>e Ts' where
- T\<^sub>e: "P,E \<turnstile>\<^sub>1 e :: T\<^sub>e" and Ts': "P,E \<turnstile>\<^sub>1 es [::] Ts'" and
- Ts: "Ts = T\<^sub>e#Ts'" by auto
- let ?A\<^sub>e = "A \<squnion> \<A> e"
- let ?\<tau> = "ty\<^sub>i' ST E A" let ?\<tau>s\<^sub>e = "compT E A ST e"
- let ?\<tau>\<^sub>e = "ty\<^sub>i' (T\<^sub>e#ST) E ?A\<^sub>e" let ?\<tau>s' = "compTs E ?A\<^sub>e (T\<^sub>e#ST) es"
- let ?\<tau>s = "?\<tau> # ?\<tau>s\<^sub>e @ (?\<tau>\<^sub>e # ?\<tau>s')"
- have Ps: "PROP ?Ps es E Ts' ?A\<^sub>e (T\<^sub>e#ST)" by fact
- with Cons_exp.prems T\<^sub>e Ts'
- have "\<turnstile> compEs\<^sub>2 es, compxEs\<^sub>2 es 0 (size (T\<^sub>e#ST)) [::] ?\<tau>\<^sub>e#?\<tau>s'" by (simp add: after_def)
- also from Cons_exp T\<^sub>e have "\<turnstile> compE\<^sub>2 e, compxE\<^sub>2 e 0 (size ST) [::] ?\<tau>#?\<tau>s\<^sub>e@[?\<tau>\<^sub>e]"
- by (auto simp: after_def)
- moreover
- from Ps Cons_exp.prems T\<^sub>e Ts' Ts
- have "last ?\<tau>s = ty\<^sub>i' (rev Ts@ST) E (?A\<^sub>e \<squnion> \<A>s es)" by simp
- ultimately show ?case using T\<^sub>e by (simp add: after_def hyperUn_assoc)
-next
- case (FAss e\<^sub>1 F D e\<^sub>2)
- hence Void: "P,E \<turnstile>\<^sub>1 e\<^sub>1\<bullet>F{D} := e\<^sub>2 :: Void" by auto
- then obtain C T T' where
- C: "P,E \<turnstile>\<^sub>1 e\<^sub>1 :: Class C" and sees: "P \<turnstile> C sees F,NonStatic:T in D" and
- T': "P,E \<turnstile>\<^sub>1 e\<^sub>2 :: T'" and T'_T: "P \<turnstile> T' \<le> T" by auto
- let ?A\<^sub>1 = "A \<squnion> \<A> e\<^sub>1" let ?A\<^sub>2 = "?A\<^sub>1 \<squnion> \<A> e\<^sub>2"
- let ?\<tau> = "ty\<^sub>i' ST E A" let ?\<tau>s\<^sub>1 = "compT E A ST e\<^sub>1"
- let ?\<tau>\<^sub>1 = "ty\<^sub>i' (Class C#ST) E ?A\<^sub>1" let ?\<tau>s\<^sub>2 = "compT E ?A\<^sub>1 (Class C#ST) e\<^sub>2"
- let ?\<tau>\<^sub>2 = "ty\<^sub>i' (T'#Class C#ST) E ?A\<^sub>2" let ?\<tau>\<^sub>3 = "ty\<^sub>i' ST E ?A\<^sub>2"
- let ?\<tau>' = "ty\<^sub>i' (Void#ST) E ?A\<^sub>2"
- from FAss.prems sees T'_T
- have "\<turnstile> [Putfield F D,Push Unit],[] [::] [?\<tau>\<^sub>2,?\<tau>\<^sub>3,?\<tau>']"
- by (fastforce simp add: wt_Push wt_Put)
- also have "PROP ?P e\<^sub>2 E T' ?A\<^sub>1 (Class C#ST)" by fact
- with FAss.prems T'
- have "\<turnstile> compE\<^sub>2 e\<^sub>2, compxE\<^sub>2 e\<^sub>2 0 (size ST+1) [::] ?\<tau>\<^sub>1#?\<tau>s\<^sub>2@[?\<tau>\<^sub>2]"
- by (auto simp add: after_def hyperUn_assoc)
- also from FAss C have "\<turnstile> compE\<^sub>2 e\<^sub>1, compxE\<^sub>2 e\<^sub>1 0 (size ST) [::] ?\<tau>#?\<tau>s\<^sub>1@[?\<tau>\<^sub>1]"
- by (auto simp add: after_def)
- finally show ?case using Void C T' by (simp add: after_def hyperUn_assoc)
-next
- case (SFAss C F D e\<^sub>2)
- hence Void: "P,E \<turnstile>\<^sub>1 C\<bullet>\<^sub>sF{D} := e\<^sub>2 :: Void" by auto
- then obtain T T' where
- sees: "P \<turnstile> C sees F,Static:T in D" and
- T': "P,E \<turnstile>\<^sub>1 e\<^sub>2 :: T'" and T'_T: "P \<turnstile> T' \<le> T" by auto
- let ?A\<^sub>2 = "A \<squnion> \<A> e\<^sub>2"
- let ?\<tau> = "ty\<^sub>i' ST E A" let ?\<tau>s\<^sub>2 = "compT E A ST e\<^sub>2"
- let ?\<tau>\<^sub>2 = "ty\<^sub>i' (T'#ST) E ?A\<^sub>2" let ?\<tau>\<^sub>3 = "ty\<^sub>i' ST E ?A\<^sub>2"
- let ?\<tau>' = "ty\<^sub>i' (Void#ST) E ?A\<^sub>2"
- from SFAss.prems sees T'_T max_stack1[OF T']
- have "\<turnstile> [Putstatic C F D,Push Unit],[] [::] [?\<tau>\<^sub>2,?\<tau>\<^sub>3,?\<tau>']"
- by (fastforce simp add: wt_Push wt_PutS)
- also have "PROP ?P e\<^sub>2 E T' A ST" by fact
- with SFAss.prems T'
- have "\<turnstile> compE\<^sub>2 e\<^sub>2, compxE\<^sub>2 e\<^sub>2 0 (size ST) [::] ?\<tau>#?\<tau>s\<^sub>2@[?\<tau>\<^sub>2]"
- by (auto simp add: after_def hyperUn_assoc)
- finally show ?case using Void T' by (simp add: after_def hyperUn_assoc)
-next
- case Val thus ?case by(auto simp:after_def wt_Push)
-next
- case Cast thus ?case by (auto simp:after_def wt_Cast)
-next
- case (Block i T\<^sub>i e)
- let ?\<tau>s = "ty\<^sub>i' ST E A # compT (E @ [T\<^sub>i]) (A\<ominus>i) ST e"
- have IH: "PROP ?P e (E@[T\<^sub>i]) T (A\<ominus>i) ST" by fact
- hence "\<turnstile> compE\<^sub>2 e, compxE\<^sub>2 e 0 (size ST) [::]
- ?\<tau>s @ [ty\<^sub>i' (T#ST) (E@[T\<^sub>i]) (A\<ominus>(size E) \<squnion> \<A> e)]"
- using Block.prems by (auto simp add: after_def)
- also have "P \<turnstile> ty\<^sub>i' (T # ST) (E@[T\<^sub>i]) (A \<ominus> size E \<squnion> \<A> e) \<le>'
- ty\<^sub>i' (T # ST) (E@[T\<^sub>i]) ((A \<squnion> \<A> e) \<ominus> size E)"
- by(auto simp add:hyperset_defs intro: ty\<^sub>i'_antimono)
- also have "\<dots> = ty\<^sub>i' (T # ST) E (A \<squnion> \<A> e)" by simp
- also have "P \<turnstile> \<dots> \<le>' ty\<^sub>i' (T # ST) E (A \<squnion> (\<A> e \<ominus> i))"
- by(auto simp add:hyperset_defs intro: ty\<^sub>i'_antimono)
- finally show ?case using Block.prems by(simp add: after_def)
-next
- case Var thus ?case by(auto simp:after_def wt_Load)
-next
- case FAcc thus ?case by(auto simp:after_def wt_Get)
-next
- case SFAcc thus ?case by(auto simp: after_def wt_GetS)
-next
- case (LAss i e)
- then obtain T' where wt: "P,E \<turnstile>\<^sub>1 e :: T'" by auto
- show ?case using max_stack1[OF wt] LAss
- by(auto simp: hyper_insert_comm after_def wt_Store wt_Push)
-next
- case Nil_exp thus ?case by auto
-next
- case throw thus ?case by(auto simp add: after_def wt_Throw)
-next
- case (While e c)
- obtain Tc where wte: "P,E \<turnstile>\<^sub>1 e :: Boolean" and wtc: "P,E \<turnstile>\<^sub>1 c :: Tc"
- and [simp]: "T = Void" using While by auto
- have [simp]: "ty E (while (e) c) = Void" using While by simp
- let ?A\<^sub>0 = "A \<squnion> \<A> e" let ?A\<^sub>1 = "?A\<^sub>0 \<squnion> \<A> c"
- let ?\<tau> = "ty\<^sub>i' ST E A" let ?\<tau>s\<^sub>e = "compT E A ST e"
- let ?\<tau>\<^sub>e = "ty\<^sub>i' (Boolean#ST) E ?A\<^sub>0" let ?\<tau>\<^sub>1 = "ty\<^sub>i' ST E ?A\<^sub>0"
- let ?\<tau>s\<^sub>c = "compT E ?A\<^sub>0 ST c" let ?\<tau>\<^sub>c = "ty\<^sub>i' (Tc#ST) E ?A\<^sub>1"
- let ?\<tau>\<^sub>2 = "ty\<^sub>i' ST E ?A\<^sub>1" let ?\<tau>' = "ty\<^sub>i' (Void#ST) E ?A\<^sub>0"
- let ?\<tau>s = "(?\<tau> # ?\<tau>s\<^sub>e @ [?\<tau>\<^sub>e]) @ ?\<tau>\<^sub>1 # ?\<tau>s\<^sub>c @ [?\<tau>\<^sub>c, ?\<tau>\<^sub>2, ?\<tau>\<^sub>1, ?\<tau>']"
- have "\<turnstile> [],[] [::] [] @ ?\<tau>s" by(simp add:wt_instrs_def)
- also
- have "PROP ?P e E Boolean A ST" by fact
- hence "\<turnstile> compE\<^sub>2 e,compxE\<^sub>2 e 0 (size ST) [::] ?\<tau> # ?\<tau>s\<^sub>e @ [?\<tau>\<^sub>e]"
- using While.prems by (auto simp:after_def)
- also
- have "[] @ ?\<tau>s = (?\<tau> # ?\<tau>s\<^sub>e) @ ?\<tau>\<^sub>e # ?\<tau>\<^sub>1 # ?\<tau>s\<^sub>c @ [?\<tau>\<^sub>c,?\<tau>\<^sub>2,?\<tau>\<^sub>1,?\<tau>']" by simp
- also
- let ?n\<^sub>e = "size(compE\<^sub>2 e)" let ?n\<^sub>c = "size(compE\<^sub>2 c)"
- thm compT_sizes(1)
- let ?if = "IfFalse (int ?n\<^sub>c + 3)"
- from wtc wte have "compE\<^sub>2 c \<noteq> []" and "compE\<^sub>2 e \<noteq> []" using WT\<^sub>1_nsub_RI by auto
- then have compe2c: "length (compE\<^sub>2 c) > 0" and compe2e: "length (compE\<^sub>2 e) > 0" by auto
- have Suc_pred'_auxi: "\<And>n. 0 < n \<Longrightarrow> n = Suc (n - Suc 0)" using Suc_pred'[OF compe2c] by auto
- have "\<turnstile> [?if],[] [::] ?\<tau>\<^sub>e # ?\<tau>\<^sub>1 # ?\<tau>s\<^sub>c @ [?\<tau>\<^sub>c, ?\<tau>\<^sub>2, ?\<tau>\<^sub>1, ?\<tau>']"
- proof-{
- show ?thesis
- apply (rule wt_IfFalse)
- apply simp
- apply simp
- apply (subgoal_tac "length (compE\<^sub>2 c) = length (compT E (A \<squnion> \<A> e) ST c) + 1"
- "nat (int (length (compE\<^sub>2 c)) + 3 - 2) > length (compT E (A \<squnion> \<A> e) ST c)")
- apply (insert Suc_pred'_auxi[OF compe2c])
- apply (simp add:compT_sizes(1)[OF wtc] )+
- done
- }
- qed
- also
- have "(?\<tau> # ?\<tau>s\<^sub>e) @ (?\<tau>\<^sub>e # ?\<tau>\<^sub>1 # ?\<tau>s\<^sub>c @ [?\<tau>\<^sub>c, ?\<tau>\<^sub>2, ?\<tau>\<^sub>1, ?\<tau>']) = ?\<tau>s" by simp
- also
- have "PROP ?P c E Tc ?A\<^sub>0 ST" by fact
- hence "\<turnstile> compE\<^sub>2 c,compxE\<^sub>2 c 0 (size ST) [::] ?\<tau>\<^sub>1 # ?\<tau>s\<^sub>c @ [?\<tau>\<^sub>c]"
- using While.prems wtc by (auto simp:after_def)
- also have "?\<tau>s = (?\<tau> # ?\<tau>s\<^sub>e @ [?\<tau>\<^sub>e,?\<tau>\<^sub>1] @ ?\<tau>s\<^sub>c) @ [?\<tau>\<^sub>c,?\<tau>\<^sub>2,?\<tau>\<^sub>1,?\<tau>']" by simp
- also have "\<turnstile> [Pop],[] [::] [?\<tau>\<^sub>c, ?\<tau>\<^sub>2]" by(simp add:wt_Pop)
- also have "(?\<tau> # ?\<tau>s\<^sub>e @ [?\<tau>\<^sub>e,?\<tau>\<^sub>1] @ ?\<tau>s\<^sub>c) @ [?\<tau>\<^sub>c,?\<tau>\<^sub>2,?\<tau>\<^sub>1,?\<tau>'] = ?\<tau>s" by simp
- also let ?go = "Goto (-int(?n\<^sub>c+?n\<^sub>e+2))"
- have "P \<turnstile> ?\<tau>\<^sub>2 \<le>' ?\<tau>" by(fastforce intro: ty\<^sub>i'_antimono simp: hyperset_defs)
- hence "P,T\<^sub>r,mxs,size ?\<tau>s,[] \<turnstile> ?go,?n\<^sub>e+?n\<^sub>c+2 :: ?\<tau>s" using wte wtc
- proof-{
- let ?t1 = "ty\<^sub>i' ST E A" let ?t2 = "ty\<^sub>i' (Boolean # ST) E (A \<squnion> \<A> e)" let ?t3 = "ty\<^sub>i' ST E (A \<squnion> \<A> e)"
- let ?t4 = "[ty\<^sub>i' (Tc # ST) E (A \<squnion> \<A> e \<squnion> \<A> c), ty\<^sub>i' ST E (A \<squnion> \<A> e \<squnion> \<A> c), ty\<^sub>i' ST E (A \<squnion> \<A> e), ty\<^sub>i' (Void # ST) E (A \<squnion> \<A> e)]"
- let ?c = " compT E (A \<squnion> \<A> e) ST c" let ?e = "compT E A ST e"
-
- assume ass: "P \<turnstile> ty\<^sub>i' ST E (A \<squnion> \<A> e \<squnion> \<A> c) \<le>' ?t1"
- show ?thesis
- apply (rule wt_Goto)
- apply simp
- apply simp
- apply simp
- proof-{
- let ?s1 = "((?t1 # ?e @ [?t2]) @ ?t3 # ?c)"
- have len1: "length ?c = length (compE\<^sub>2 c) - 1" using compT_sizes(1)[OF wtc] by simp
- have len2: "length ?e = length (compE\<^sub>2 e) - 1" using compT_sizes(1)[OF wte] by simp
- hence "length (compE\<^sub>2 e) + length (compE\<^sub>2 c) + 2 > length ?s1"
- using len1 compe2c compe2e by auto
- hence "(?s1 @ ?t4) ! (length (compE\<^sub>2 e) + length (compE\<^sub>2 c) + 2) =
- ?t4 ! (length (compE\<^sub>2 e) + length (compE\<^sub>2 c) + 2 - length ?s1)"
- by (auto simp only:nth_append split: if_splits)
- hence "(?s1 @ ?t4) ! (length (compE\<^sub>2 e) + length (compE\<^sub>2 c) + 2) = ty\<^sub>i' ST E (A \<squnion> \<A> e \<squnion> \<A> c)"
- using len1 len2 compe2c compe2e by auto
- hence "P \<turnstile> (?s1 @ ?t4) ! (length (compE\<^sub>2 e) + length (compE\<^sub>2 c) + 2) \<le>' ty\<^sub>i' ST E A"
- using ass by simp
- thus "P \<turnstile> ((?t1 # ?e @ [?t2]) @ ?t3 # ?c @ ?t4) ! (length (compE\<^sub>2 e) + length (compE\<^sub>2 c) + 2) \<le>'
- ((?t1 # ?e @ [?t2]) @ ?t3 # ?c @ ?t4) ! nat (int (length (compE\<^sub>2 e) + length (compE\<^sub>2 c) + 2) +
- - int (length (compE\<^sub>2 c) + length (compE\<^sub>2 e) + 2))"
- by simp
- }qed
- }qed
- also have "?\<tau>s = (?\<tau> # ?\<tau>s\<^sub>e @ [?\<tau>\<^sub>e,?\<tau>\<^sub>1] @ ?\<tau>s\<^sub>c @ [?\<tau>\<^sub>c, ?\<tau>\<^sub>2]) @ [?\<tau>\<^sub>1, ?\<tau>']"
- by simp
- also have "\<turnstile> [Push Unit],[] [::] [?\<tau>\<^sub>1,?\<tau>']"
- using While.prems max_stack1[OF wtc] by(auto simp add:wt_Push)
- finally show ?case using wtc wte
- by (simp add:after_def)
-next
- case (Cond e e\<^sub>1 e\<^sub>2)
- obtain T\<^sub>1 T\<^sub>2 where wte: "P,E \<turnstile>\<^sub>1 e :: Boolean"
- and wt\<^sub>1: "P,E \<turnstile>\<^sub>1 e\<^sub>1 :: T\<^sub>1" and wt\<^sub>2: "P,E \<turnstile>\<^sub>1 e\<^sub>2 :: T\<^sub>2"
- and sub\<^sub>1: "P \<turnstile> T\<^sub>1 \<le> T" and sub\<^sub>2: "P \<turnstile> T\<^sub>2 \<le> T"
- using Cond by auto
- from wte wt\<^sub>1 wt\<^sub>2 have "compE\<^sub>2 e\<^sub>1 \<noteq> []" and "compE\<^sub>2 e\<^sub>2 \<noteq> []" and "compE\<^sub>2 e \<noteq> []" using WT\<^sub>1_nsub_RI by auto
- then have compe: "length (compE\<^sub>2 e) > 0"
- and compe1: "length (compE\<^sub>2 e\<^sub>1) > 0"
- and compe2: "length (compE\<^sub>2 e\<^sub>2) > 0" by auto
-
-
- have [simp]: "ty E (if (e) e\<^sub>1 else e\<^sub>2) = T" using Cond by simp
- let ?A\<^sub>0 = "A \<squnion> \<A> e" let ?A\<^sub>2 = "?A\<^sub>0 \<squnion> \<A> e\<^sub>2" let ?A\<^sub>1 = "?A\<^sub>0 \<squnion> \<A> e\<^sub>1"
- let ?A' = "?A\<^sub>0 \<squnion> \<A> e\<^sub>1 \<sqinter> \<A> e\<^sub>2"
- let ?\<tau>\<^sub>2 = "ty\<^sub>i' ST E ?A\<^sub>0" let ?\<tau>' = "ty\<^sub>i' (T#ST) E ?A'"
- let ?\<tau>s\<^sub>2 = "compT E ?A\<^sub>0 ST e\<^sub>2"
- have "PROP ?P e\<^sub>2 E T\<^sub>2 ?A\<^sub>0 ST" by fact
- hence "\<turnstile> compE\<^sub>2 e\<^sub>2, compxE\<^sub>2 e\<^sub>2 0 (size ST) [::] (?\<tau>\<^sub>2#?\<tau>s\<^sub>2) @ [ty\<^sub>i' (T\<^sub>2#ST) E ?A\<^sub>2]"
- using Cond.prems wt\<^sub>2 by(auto simp add:after_def)
- also have "P \<turnstile> ty\<^sub>i' (T\<^sub>2#ST) E ?A\<^sub>2 \<le>' ?\<tau>'" using sub\<^sub>2
- by(auto simp add: hyperset_defs ty\<^sub>i'_def intro!: ty\<^sub>l_antimono)
- also
- let ?\<tau>\<^sub>3 = "ty\<^sub>i' (T\<^sub>1 # ST) E ?A\<^sub>1"
- let ?g\<^sub>2 = "Goto(int (size (compE\<^sub>2 e\<^sub>2) + 1))"
- from sub\<^sub>1 have "P,T\<^sub>r,mxs,size(compE\<^sub>2 e\<^sub>2)+2,[] \<turnstile> ?g\<^sub>2,0 :: ?\<tau>\<^sub>3#(?\<tau>\<^sub>2#?\<tau>s\<^sub>2)@[?\<tau>']" using wt\<^sub>2
- by(auto simp: hyperset_defs wt_defs nth_Cons ty\<^sub>i'_def
- split:nat.split intro!: ty\<^sub>l_antimono)
- also
- let ?\<tau>s\<^sub>1 = "compT E ?A\<^sub>0 ST e\<^sub>1"
- have "PROP ?P e\<^sub>1 E T\<^sub>1 ?A\<^sub>0 ST" by fact
- hence "\<turnstile> compE\<^sub>2 e\<^sub>1,compxE\<^sub>2 e\<^sub>1 0 (size ST) [::] ?\<tau>\<^sub>2 # ?\<tau>s\<^sub>1 @ [?\<tau>\<^sub>3]"
- using Cond.prems wt\<^sub>1 by(auto simp add:after_def)
- also
- let ?\<tau>s\<^sub>1\<^sub>2 = "?\<tau>\<^sub>2 # ?\<tau>s\<^sub>1 @ ?\<tau>\<^sub>3 # (?\<tau>\<^sub>2 # ?\<tau>s\<^sub>2) @ [?\<tau>']"
- let ?\<tau>\<^sub>1 = "ty\<^sub>i' (Boolean#ST) E ?A\<^sub>0"
- let ?g\<^sub>1 = "IfFalse(int (size (compE\<^sub>2 e\<^sub>1) + 2))"
- let ?code = "compE\<^sub>2 e\<^sub>1 @ ?g\<^sub>2 # compE\<^sub>2 e\<^sub>2"
-
- have len1: "length ?\<tau>s\<^sub>1 = length (compE\<^sub>2 e\<^sub>1) - 1" using compT_sizes(1)[OF wt\<^sub>1] by simp
- have len2: "length ?\<tau>s\<^sub>2 = length (compE\<^sub>2 e\<^sub>2) - 1" using compT_sizes(1)[OF wt\<^sub>2] by simp
- have len_auxi: "length (compE\<^sub>2 e\<^sub>1) - (length (compE\<^sub>2 e\<^sub>1) - Suc 0) = Suc 0" using compe1
- by (simp add:diff_Suc split:nat.split)
- have "\<turnstile> [?g\<^sub>1],[] [::] [?\<tau>\<^sub>1] @ ?\<tau>s\<^sub>1\<^sub>2"
- proof-{
- show ?thesis
- apply clarsimp
- apply (rule wt_IfFalse)
- apply (simp only:nat_add_distrib)
- apply simp
- apply (auto simp only:nth_append split:if_splits)
- apply (simp add:len1)
- apply (simp only: len1 compe1)
- apply (simp add:len1 len2 compe1 compe2 )
- apply (insert len_auxi)
- apply simp
- done
- }qed
- also (wt_instrs_ext2) have "[?\<tau>\<^sub>1] @ ?\<tau>s\<^sub>1\<^sub>2 = ?\<tau>\<^sub>1 # ?\<tau>s\<^sub>1\<^sub>2" by simp also
- let ?\<tau> = "ty\<^sub>i' ST E A"
- have "PROP ?P e E Boolean A ST" by fact
- hence "\<turnstile> compE\<^sub>2 e, compxE\<^sub>2 e 0 (size ST) [::] ?\<tau> # compT E A ST e @ [?\<tau>\<^sub>1]"
- using Cond.prems wte by(auto simp add:after_def)
- finally show ?case using wte wt\<^sub>1 wt\<^sub>2 by(simp add:after_def hyperUn_assoc)
-next
- case (Call e M es)
- obtain C D Ts m Ts' where C: "P,E \<turnstile>\<^sub>1 e :: Class C"
- and "method": "P \<turnstile> C sees M,NonStatic:Ts \<rightarrow> T = m in D"
- and wtes: "P,E \<turnstile>\<^sub>1 es [::] Ts'" and subs: "P \<turnstile> Ts' [\<le>] Ts"
- using Call.prems by auto
- from wtes have same_size: "size es = size Ts'" by(rule WTs\<^sub>1_same_size)
- let ?A\<^sub>0 = "A \<squnion> \<A> e" let ?A\<^sub>1 = "?A\<^sub>0 \<squnion> \<A>s es"
- let ?\<tau> = "ty\<^sub>i' ST E A" let ?\<tau>s\<^sub>e = "compT E A ST e"
- let ?\<tau>\<^sub>e = "ty\<^sub>i' (Class C # ST) E ?A\<^sub>0"
- let ?\<tau>s\<^sub>e\<^sub>s = "compTs E ?A\<^sub>0 (Class C # ST) es"
- let ?\<tau>\<^sub>1 = "ty\<^sub>i' (rev Ts' @ Class C # ST) E ?A\<^sub>1"
- let ?\<tau>' = "ty\<^sub>i' (T # ST) E ?A\<^sub>1"
- have "\<turnstile> [Invoke M (size es)],[] [::] [?\<tau>\<^sub>1,?\<tau>']"
- by(rule wt_Invoke[OF same_size "method" subs])
- also
- have "PROP ?Ps es E Ts' ?A\<^sub>0 (Class C # ST)" by fact
- hence "\<turnstile> compEs\<^sub>2 es,compxEs\<^sub>2 es 0 (size ST+1) [::] ?\<tau>\<^sub>e # ?\<tau>s\<^sub>e\<^sub>s"
- "last (?\<tau>\<^sub>e # ?\<tau>s\<^sub>e\<^sub>s) = ?\<tau>\<^sub>1"
- using Call.prems wtes by(auto simp add:after_def)
- also have "(?\<tau>\<^sub>e # ?\<tau>s\<^sub>e\<^sub>s) @ [?\<tau>'] = ?\<tau>\<^sub>e # ?\<tau>s\<^sub>e\<^sub>s @ [?\<tau>']" by simp
- also have "\<turnstile> compE\<^sub>2 e,compxE\<^sub>2 e 0 (size ST) [::] ?\<tau> # ?\<tau>s\<^sub>e @ [?\<tau>\<^sub>e]"
- using Call C by(auto simp add:after_def)
- finally show ?case using Call.prems C wtes by(simp add:after_def hyperUn_assoc)
-next
- case (SCall C M es)
- obtain D Ts m Ts' where "method": "P \<turnstile> C sees M,Static:Ts \<rightarrow> T = m in D"
- and wtes: "P,E \<turnstile>\<^sub>1 es [::] Ts'" and subs: "P \<turnstile> Ts' [\<le>] Ts"
- using SCall.prems by auto
- from SCall.prems(1) have nclinit: "M \<noteq> clinit" by auto
- from wtes have same_size: "size es = size Ts'" by(rule WTs\<^sub>1_same_size)
- have mxs: "length ST < mxs" using WT\<^sub>1_nsub_RI[OF SCall.prems(1)] SCall.prems(4) by simp
- let ?A\<^sub>1 = "A \<squnion> \<A>s es"
- let ?\<tau> = "ty\<^sub>i' ST E A"
- let ?\<tau>s\<^sub>e\<^sub>s = "compTs E A ST es"
- let ?\<tau>\<^sub>1 = "ty\<^sub>i' (rev Ts' @ ST) E ?A\<^sub>1"
- let ?\<tau>' = "ty\<^sub>i' (T # ST) E ?A\<^sub>1"
- have "\<turnstile> [Invokestatic C M (size es)],[] [::] [?\<tau>\<^sub>1,?\<tau>']"
- by(rule wt_Invokestatic[OF mxs same_size nclinit "method" subs])
- also
- have "PROP ?Ps es E Ts' A ST" by fact
- hence "\<turnstile> compEs\<^sub>2 es,compxEs\<^sub>2 es 0 (size ST) [::] ?\<tau> # ?\<tau>s\<^sub>e\<^sub>s"
- "last (?\<tau> # ?\<tau>s\<^sub>e\<^sub>s) = ?\<tau>\<^sub>1"
- using SCall.prems wtes by(auto simp add:after_def)
- also have "(?\<tau> # ?\<tau>s\<^sub>e\<^sub>s) @ [?\<tau>'] = ?\<tau> # ?\<tau>s\<^sub>e\<^sub>s @ [?\<tau>']" by simp
- finally show ?case using SCall.prems wtes by(simp add:after_def hyperUn_assoc)
-next
- case Seq thus ?case
- by(auto simp:after_def)
- (fastforce simp:wt_Push wt_Pop hyperUn_assoc
- intro:wt_instrs_app2 wt_instrs_Cons)
-next
- case (INIT C Cs b e)
- have "P,E \<turnstile>\<^sub>1 INIT C (Cs,b) \<leftarrow> e :: T" by fact
- thus ?case using WT\<^sub>1_nsub_RI by simp
-next
- case (RI C e' Cs e)
- have "P,E \<turnstile>\<^sub>1 RI (C,e') ; Cs \<leftarrow> e :: T" by fact
- thus ?case using WT\<^sub>1_nsub_RI by simp
-qed
-(*>*)
-
-
-lemma [simp]: "types (compP f P) = types P"
-(*<*)by auto(*>*)
-
-lemma [simp]: "states (compP f P) mxs mxl = states P mxs mxl"
-(*<*)by (simp add: JVM_states_unfold)(*>*)
-
-lemma [simp]: "app\<^sub>i (i, compP f P, pc, mpc, T, \<tau>) = app\<^sub>i (i, P, pc, mpc, T, \<tau>)"
-(*<*)(is "?A = ?B")
-proof -
- obtain ST LT where \<tau>: "\<tau> = (ST, LT)" by(cases \<tau>) simp
- then show ?thesis proof(cases i)
- case Invoke show ?thesis
- proof(rule iffI)
- assume ?A then show ?B using Invoke \<tau>
- by auto (fastforce dest!: sees_method_compPD)
- next
- assume ?B then show ?A using Invoke \<tau>
- by auto (force dest: sees_method_compP)
- qed
- next
- case (Invokestatic x111 x112 x113) show ?thesis
- proof(rule iffI)
- assume ?A then show ?B using Invokestatic \<tau>
- by auto (fastforce dest!: sees_method_compPD)
- next
- assume ?B then show ?A using Invokestatic \<tau>
- by auto (force dest: sees_method_compP)
- qed
- qed auto
-qed
-(*>*)
-
-lemma [simp]: "is_relevant_entry (compP f P) i = is_relevant_entry P i"
-(*<*)
-proof -
- { fix pc e
- have "is_relevant_entry (compP f P) i pc e = is_relevant_entry P i pc e"
- by(cases i) (auto simp: is_relevant_entry_def)
- }
- then show ?thesis by fast
-qed
-(*>*)
-
-lemma [simp]: "relevant_entries (compP f P) i pc xt = relevant_entries P i pc xt"
-(*<*) by (simp add: relevant_entries_def)(*>*)
-
-lemma [simp]: "app i (compP f P) mpc T pc mxl xt \<tau> = app i P mpc T pc mxl xt \<tau>"
-(*<*)
-by (simp add: app_def xcpt_app_def eff_def xcpt_eff_def norm_eff_def)
- (fastforce simp add: image_def)
-(*>*)
-
-lemma [simp]:
-assumes "app i P mpc T pc mxl xt \<tau>"
-shows "eff i (compP f P) pc xt \<tau> = eff i P pc xt \<tau>"
-(*<*)
-using assms
-proof(clarsimp simp: eff_def norm_eff_def xcpt_eff_def app_def, cases i)
-qed auto
-(*>*)
-
-lemma [simp]: "subtype (compP f P) = subtype P"
-(*<*)by (rule ext)+ simp(*>*)
-
-lemma [simp]: "compP f P \<turnstile> \<tau> \<le>' \<tau>' = P \<turnstile> \<tau> \<le>' \<tau>'"
-(*<*) by (simp add: sup_state_opt_def sup_state_def sup_ty_opt_def)(*>*)
-
-lemma [simp]: "compP f P,T,mpc,mxl,xt \<turnstile> i,pc :: \<tau>s = P,T,mpc,mxl,xt \<turnstile> i,pc :: \<tau>s"
-(*<*)by (simp add: wt_instr_def cong: conj_cong)(*>*)
-
-declare TC1.compT_sizes[simp] TC0.ty_def2[simp]
-
-context TC2
-begin
-
-lemma compT_method_NonStatic:
- fixes e and A and C and Ts and mxl\<^sub>0
- defines [simp]: "E \<equiv> Class C # Ts"
- and [simp]: "A \<equiv> \<lfloor>{..size Ts}\<rfloor>"
- and [simp]: "A' \<equiv> A \<squnion> \<A> e"
- and [simp]: "mxl\<^sub>0 \<equiv> max_vars e"
- assumes mxs: "max_stack e = mxs"
- and mxl: "Suc (length Ts + max_vars e) = mxl"
- assumes wf: "wf_prog p P" and wte: "P,E \<turnstile>\<^sub>1 e :: T" and \<D>: "\<D> e A"
- and \<B>: "\<B> e (size E)" and E_P: "set E \<subseteq> types P" and wid: "P \<turnstile> T \<le> T\<^sub>r"
- shows "wt_method (compP\<^sub>2 P) C NonStatic Ts T\<^sub>r mxs mxl\<^sub>0 (compE\<^sub>2 e @ [Return])
- (compxE\<^sub>2 e 0 0) (ty\<^sub>i' [] E A # compT\<^sub>a E A [] e)"
-(*<*)(is "wt_method ?P C ?b Ts T\<^sub>r mxs mxl\<^sub>0 ?is ?xt ?\<tau>s")
-proof -
- let ?n = "length E + mxl\<^sub>0"
- have wt_compE: "P,T\<^sub>r,mxs \<turnstile> compE\<^sub>2 e, compxE\<^sub>2 e 0 (length []) [::]
- TC0.ty\<^sub>i' ?n [] E A # TC1.compT P ?n E A [] e @[TC0.after P ?n E A [] e]"
- using mxs TC2.compT_wt_instrs [OF wte \<D> \<B>, of "[]" mxs ?n T\<^sub>r] by simp
-
- have "OK (ty\<^sub>i' [T] E A') \<in> states P mxs mxl"
- using mxs WT\<^sub>1_is_type[OF wf wte E_P] max_stack1[OF wte] OK_ty\<^sub>i'_in_statesI[OF E_P]
- by simp
- moreover have "OK ` set (compT E A [] e) \<subseteq> states P mxs mxl"
- using mxs mxl wid compT_states(1)[OF wf wte E_P] by simp
- ultimately have "check_types ?P mxs ?n (map OK ?\<tau>s)"
- using mxl wte E_P by (simp add: compT\<^sub>a_def after_def check_types_def)
- moreover have "wt_start ?P C ?b Ts mxl\<^sub>0 ?\<tau>s" using mxl
- by (auto simp: wt_start_def ty\<^sub>i'_def ty\<^sub>l_def list_all2_conv_all_nth nth_Cons
- split: nat.split)
- moreover {
- fix pc assume pc: "pc < size ?is"
- then have "?P,T\<^sub>r,mxs,size ?is,?xt \<turnstile> ?is!pc,pc :: ?\<tau>s"
- proof(cases "pc < length (compE\<^sub>2 e)")
- case True
- then show ?thesis using mxs wte wt_compE
- by (clarsimp simp: compT\<^sub>a_def mxl after_def wt_instrs_def)
- next
- case False
- have "length (compE\<^sub>2 e) = pc" using less_antisym[OF False] pc by simp
- then show ?thesis using mxl wte E_P wid
- by (clarsimp simp: compT\<^sub>a_def after_def wt_defs xcpt_app_pcs xcpt_eff_pcs ty\<^sub>i'_def)
- qed
- }
- moreover have "0 < size ?is" and "size ?\<tau>s = size ?is"
- using wte by (simp_all add: compT\<^sub>a_def)
- ultimately show ?thesis by(simp add: wt_method_def)
-qed
-(*>*)
-
-lemma compT_method_Static:
- fixes e and A and C and Ts and mxl\<^sub>0
- defines [simp]: "E \<equiv> Ts"
- and [simp]: "A \<equiv> \<lfloor>{..<size Ts}\<rfloor>"
- and [simp]: "A' \<equiv> A \<squnion> \<A> e"
- and [simp]: "mxl\<^sub>0 \<equiv> max_vars e"
- assumes mxs: "max_stack e = mxs"
- and mxl: "length Ts + max_vars e = mxl"
- assumes wf: "wf_prog p P" and wte: "P,E \<turnstile>\<^sub>1 e :: T" and \<D>: "\<D> e A"
- and \<B>: "\<B> e (size E)" and E_P: "set E \<subseteq> types P" and wid: "P \<turnstile> T \<le> T\<^sub>r"
- shows "wt_method (compP\<^sub>2 P) C Static Ts T\<^sub>r mxs mxl\<^sub>0 (compE\<^sub>2 e @ [Return])
- (compxE\<^sub>2 e 0 0) (ty\<^sub>i' [] E A # compT\<^sub>a E A [] e)"
-(*<*)(is "wt_method ?P C ?b Ts T\<^sub>r mxs mxl\<^sub>0 ?is ?xt ?\<tau>s")
-proof -
- let ?n = "length E + mxl\<^sub>0"
- have wt_compE: "P,T\<^sub>r,mxs \<turnstile> compE\<^sub>2 e, compxE\<^sub>2 e 0 (length []) [::]
- TC0.ty\<^sub>i' ?n [] E A # TC1.compT P ?n E A [] e @[TC0.after P ?n E A [] e]"
- using mxs TC2.compT_wt_instrs [OF wte \<D> \<B>, of "[]" mxs ?n T\<^sub>r] by simp
-
- have "OK (ty\<^sub>i' [T] E A') \<in> states P mxs mxl"
- using mxs WT\<^sub>1_is_type[OF wf wte E_P] max_stack1[OF wte] OK_ty\<^sub>i'_in_statesI[OF E_P]
- by simp
- moreover have "OK ` set (compT E A [] e) \<subseteq> states P mxs mxl"
- using mxs mxl wid compT_states(1)[OF wf wte E_P] by simp
- ultimately have "check_types ?P mxs ?n (map OK ?\<tau>s)"
- using mxl wte E_P by (simp add: compT\<^sub>a_def after_def check_types_def)
- moreover have "wt_start ?P C ?b Ts mxl\<^sub>0 ?\<tau>s" using mxl
- by (auto simp: wt_start_def ty\<^sub>i'_def ty\<^sub>l_def list_all2_conv_all_nth nth_Cons
- split: nat.split)
- moreover {
- fix pc assume pc: "pc < size ?is"
- then have "?P,T\<^sub>r,mxs,size ?is,?xt \<turnstile> ?is!pc,pc :: ?\<tau>s"
- proof(cases "pc < length (compE\<^sub>2 e)")
- case True
- then show ?thesis using mxs wte wt_compE
- by (clarsimp simp: compT\<^sub>a_def mxl after_def wt_instrs_def)
- next
- case False
- have "length (compE\<^sub>2 e) = pc" using less_antisym[OF False] pc by simp
- then show ?thesis using mxl wte E_P wid
- by (clarsimp simp: compT\<^sub>a_def after_def wt_defs xcpt_app_pcs xcpt_eff_pcs ty\<^sub>i'_def)
- qed
- }
- moreover have "0 < size ?is" and "size ?\<tau>s = size ?is"
- using wte by (simp_all add: compT\<^sub>a_def)
- ultimately show ?thesis by(simp add: wt_method_def)
-qed
-(*>*)
-
-end
-
-definition compTP :: "J\<^sub>1_prog \<Rightarrow> ty\<^sub>P" where
- "compTP P C M = (
- let (D,b,Ts,T,e) = method P C M;
- E = case b of Static \<Rightarrow> Ts | NonStatic \<Rightarrow> Class C # Ts;
- A = case b of Static \<Rightarrow> \<lfloor>{..<size Ts}\<rfloor> | NonStatic \<Rightarrow> \<lfloor>{..size Ts}\<rfloor>;
- mxl = (case b of Static \<Rightarrow> 0 | NonStatic \<Rightarrow> 1) + size Ts + max_vars e
- in (TC0.ty\<^sub>i' mxl [] E A # TC1.compT\<^sub>a P mxl E A [] e))"
-
-theorem wt_compP\<^sub>2:
-assumes wf: "wf_J\<^sub>1_prog P" shows "wf_jvm_prog (compP\<^sub>2 P)"
-(*<*)
-proof -
- let ?\<Phi> = "compTP P" and ?f = compMb\<^sub>2
- let ?wf\<^sub>2 = "\<lambda>P C (M, b, Ts, T\<^sub>r, mxs, mxl\<^sub>0, is, xt).
- wt_method P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt (?\<Phi> C M)"
- and ?P = "compP ?f P"
- { fix C M b Ts T m let ?md = "(M,b,Ts,T,m)::expr\<^sub>1 mdecl"
- assume cM: "P \<turnstile> C sees M, b : Ts\<rightarrow>T = m in C"
- and wfm: "wf_mdecl wf_J\<^sub>1_mdecl P C ?md"
- then have Ts_types: "\<forall>T'\<in>set Ts. is_type P T'"
- and T_type: "is_type P T" and wfm\<^sub>1: "wf_J\<^sub>1_mdecl P C ?md"
- by(simp_all add: wf_mdecl_def)
- have Ts_P: "set Ts \<subseteq> types P" using sees_wf_mdecl[OF wf cM]
- by (clarsimp simp: wf_mdecl_def)
- then have CTs_P: "is_class P C \<and> set Ts \<subseteq> types P"
- using sees_method_is_class[OF cM] by simp
-
- have "?wf\<^sub>2 ?P C (M,b,Ts,T,?f b m)"
- proof(cases b)
- case Static
- then obtain T' where wte: "P,Ts \<turnstile>\<^sub>1 m :: T'" and wid: "P \<turnstile> T' \<le> T"
- and \<D>: "\<D> m \<lfloor>{..<size Ts}\<rfloor>" and \<B>: "\<B> m (size Ts)"
- using wfm\<^sub>1 by(auto simp: wf_mdecl_def)
- show ?thesis using Static cM TC2.compT_method_Static [OF _ _ wf wte \<D> \<B> Ts_P wid]
- by(simp add: compTP_def)
- next
- case NonStatic
- then obtain T' where wte: "P,Class C#Ts \<turnstile>\<^sub>1 m :: T'" and wid: "P \<turnstile> T' \<le> T"
- and \<D>: "\<D> m \<lfloor>{..size Ts}\<rfloor>" and \<B>: "\<B> m (Suc (size Ts))"
- using wfm\<^sub>1 by(auto simp: wf_mdecl_def)
- have Ts_P: "set Ts \<subseteq> types P" using sees_wf_mdecl[OF wf cM]
- by (fastforce simp: wf_mdecl_def intro: sees_method_is_class)
- show ?thesis using NonStatic cM
- TC2.compT_method_NonStatic [simplified, OF _ _ wf wte \<D> \<B> CTs_P wid]
- by(simp add: compTP_def)
- qed
- then have "wf_mdecl ?wf\<^sub>2 ?P C (M, b, Ts, T, ?f b m)"
- using Ts_types T_type by(simp add: wf_mdecl_def)
- }
- then have "wf_prog ?wf\<^sub>2 (compP ?f P)" by(rule wf_prog_compPI[OF _ wf])
- then show ?thesis by (simp add: wf_jvm_prog_def wf_jvm_prog_phi_def) fast
-qed
-(*>*)
-
-theorem wt_J2JVM:
- "wf_J_prog P \<Longrightarrow> wf_jvm_prog (J2JVM P)"
-(*<*)
-by(simp only:o_def J2JVM_def)
- (blast intro:wt_compP\<^sub>2 compP\<^sub>1_pres_wf)
-
-end
+(* Title: JinjaDCI/Compiler/TypeComp.thy
+
+ Author: Tobias Nipkow, Susannah Mansky
+ Copyright TUM 2003, UIUC 2019-20
+
+ Based on the Jinja theory Compiler/TypeComp.thy by Tobias Nipkow
+*)
+
+section \<open> Preservation of Well-Typedness \<close>
+
+theory TypeComp
+imports Compiler "../BV/BVSpec"
+begin
+
+(*<*)
+declare nth_append[simp]
+(*>*)
+
+lemma max_stack1: "P,E \<turnstile>\<^sub>1 e :: T \<Longrightarrow> 1 \<le> max_stack e"
+(*<*)using max_stack1'[OF WT\<^sub>1_nsub_RI] by simp(*>*)
+
+locale TC0 =
+ fixes P :: "J\<^sub>1_prog" and mxl :: nat
+begin
+
+definition "ty E e = (THE T. P,E \<turnstile>\<^sub>1 e :: T)"
+
+definition "ty\<^sub>l E A' = map (\<lambda>i. if i \<in> A' \<and> i < size E then OK(E!i) else Err) [0..<mxl]"
+
+definition "ty\<^sub>i' ST E A = (case A of None \<Rightarrow> None | \<lfloor>A'\<rfloor> \<Rightarrow> Some(ST, ty\<^sub>l E A'))"
+
+definition "after E A ST e = ty\<^sub>i' (ty E e # ST) E (A \<squnion> \<A> e)"
+
+end
+
+lemma (in TC0) ty_def2 [simp]: "P,E \<turnstile>\<^sub>1 e :: T \<Longrightarrow> ty E e = T"
+(*<*)by(unfold ty_def) (blast intro: the_equality WT\<^sub>1_unique)(*>*)
+
+lemma (in TC0) [simp]: "ty\<^sub>i' ST E None = None"
+(*<*)by (simp add: ty\<^sub>i'_def)(*>*)
+
+lemma (in TC0) ty\<^sub>l_app_diff[simp]:
+ "ty\<^sub>l (E@[T]) (A - {size E}) = ty\<^sub>l E A"
+(*<*)by(auto simp add:ty\<^sub>l_def hyperset_defs)(*>*)
+
+
+lemma (in TC0) ty\<^sub>i'_app_diff[simp]:
+ "ty\<^sub>i' ST (E @ [T]) (A \<ominus> size E) = ty\<^sub>i' ST E A"
+(*<*)by(auto simp add:ty\<^sub>i'_def hyperset_defs)(*>*)
+
+
+lemma (in TC0) ty\<^sub>l_antimono:
+ "A \<subseteq> A' \<Longrightarrow> P \<turnstile> ty\<^sub>l E A' [\<le>\<^sub>\<top>] ty\<^sub>l E A"
+(*<*)by(auto simp:ty\<^sub>l_def list_all2_conv_all_nth)(*>*)
+
+
+lemma (in TC0) ty\<^sub>i'_antimono:
+ "A \<subseteq> A' \<Longrightarrow> P \<turnstile> ty\<^sub>i' ST E \<lfloor>A'\<rfloor> \<le>' ty\<^sub>i' ST E \<lfloor>A\<rfloor>"
+(*<*)by(auto simp:ty\<^sub>i'_def ty\<^sub>l_def list_all2_conv_all_nth)(*>*)
+
+
+lemma (in TC0) ty\<^sub>l_env_antimono:
+ "P \<turnstile> ty\<^sub>l (E@[T]) A [\<le>\<^sub>\<top>] ty\<^sub>l E A"
+(*<*)by(auto simp:ty\<^sub>l_def list_all2_conv_all_nth)(*>*)
+
+
+lemma (in TC0) ty\<^sub>i'_env_antimono:
+ "P \<turnstile> ty\<^sub>i' ST (E@[T]) A \<le>' ty\<^sub>i' ST E A"
+(*<*)by(auto simp:ty\<^sub>i'_def ty\<^sub>l_def list_all2_conv_all_nth)(*>*)
+
+
+lemma (in TC0) ty\<^sub>i'_incr:
+ "P \<turnstile> ty\<^sub>i' ST (E @ [T]) \<lfloor>insert (size E) A\<rfloor> \<le>' ty\<^sub>i' ST E \<lfloor>A\<rfloor>"
+(*<*)by(auto simp:ty\<^sub>i'_def ty\<^sub>l_def list_all2_conv_all_nth)(*>*)
+
+
+lemma (in TC0) ty\<^sub>l_incr:
+ "P \<turnstile> ty\<^sub>l (E @ [T]) (insert (size E) A) [\<le>\<^sub>\<top>] ty\<^sub>l E A"
+(*<*)by(auto simp: hyperset_defs ty\<^sub>l_def list_all2_conv_all_nth)(*>*)
+
+
+lemma (in TC0) ty\<^sub>l_in_types:
+ "set E \<subseteq> types P \<Longrightarrow> ty\<^sub>l E A \<in> list mxl (err (types P))"
+(*<*)by(auto simp add:ty\<^sub>l_def intro!:listI dest!: nth_mem)(*>*)
+
+locale TC1 = TC0
+begin
+
+primrec compT :: "ty list \<Rightarrow> nat hyperset \<Rightarrow> ty list \<Rightarrow> expr\<^sub>1 \<Rightarrow> ty\<^sub>i' list" and
+ compTs :: "ty list \<Rightarrow> nat hyperset \<Rightarrow> ty list \<Rightarrow> expr\<^sub>1 list \<Rightarrow> ty\<^sub>i' list" where
+"compT E A ST (new C) = []"
+| "compT E A ST (Cast C e) =
+ compT E A ST e @ [after E A ST e]"
+| "compT E A ST (Val v) = []"
+| "compT E A ST (e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2) =
+ (let ST\<^sub>1 = ty E e\<^sub>1#ST; A\<^sub>1 = A \<squnion> \<A> e\<^sub>1 in
+ compT E A ST e\<^sub>1 @ [after E A ST e\<^sub>1] @
+ compT E A\<^sub>1 ST\<^sub>1 e\<^sub>2 @ [after E A\<^sub>1 ST\<^sub>1 e\<^sub>2])"
+| "compT E A ST (Var i) = []"
+| "compT E A ST (i := e) = compT E A ST e @
+ [after E A ST e, ty\<^sub>i' ST E (A \<squnion> \<A> e \<squnion> \<lfloor>{i}\<rfloor>)]"
+| "compT E A ST (e\<bullet>F{D}) =
+ compT E A ST e @ [after E A ST e]"
+| "compT E A ST (C\<bullet>\<^sub>sF{D}) = []"
+| "compT E A ST (e\<^sub>1\<bullet>F{D} := e\<^sub>2) =
+ (let ST\<^sub>1 = ty E e\<^sub>1#ST; A\<^sub>1 = A \<squnion> \<A> e\<^sub>1; A\<^sub>2 = A\<^sub>1 \<squnion> \<A> e\<^sub>2 in
+ compT E A ST e\<^sub>1 @ [after E A ST e\<^sub>1] @
+ compT E A\<^sub>1 ST\<^sub>1 e\<^sub>2 @ [after E A\<^sub>1 ST\<^sub>1 e\<^sub>2] @
+ [ty\<^sub>i' ST E A\<^sub>2])"
+| "compT E A ST (C\<bullet>\<^sub>sF{D} := e\<^sub>2) = compT E A ST e\<^sub>2 @ [after E A ST e\<^sub>2] @ [ty\<^sub>i' ST E (A \<squnion> \<A> e\<^sub>2)]"
+| "compT E A ST {i:T; e} = compT (E@[T]) (A\<ominus>i) ST e"
+| "compT E A ST (e\<^sub>1;;e\<^sub>2) =
+ (let A\<^sub>1 = A \<squnion> \<A> e\<^sub>1 in
+ compT E A ST e\<^sub>1 @ [after E A ST e\<^sub>1, ty\<^sub>i' ST E A\<^sub>1] @
+ compT E A\<^sub>1 ST e\<^sub>2)"
+| "compT E A ST (if (e) e\<^sub>1 else e\<^sub>2) =
+ (let A\<^sub>0 = A \<squnion> \<A> e; \<tau> = ty\<^sub>i' ST E A\<^sub>0 in
+ compT E A ST e @ [after E A ST e, \<tau>] @
+ compT E A\<^sub>0 ST e\<^sub>1 @ [after E A\<^sub>0 ST e\<^sub>1, \<tau>] @
+ compT E A\<^sub>0 ST e\<^sub>2)"
+| "compT E A ST (while (e) c) =
+ (let A\<^sub>0 = A \<squnion> \<A> e; A\<^sub>1 = A\<^sub>0 \<squnion> \<A> c; \<tau> = ty\<^sub>i' ST E A\<^sub>0 in
+ compT E A ST e @ [after E A ST e, \<tau>] @
+ compT E A\<^sub>0 ST c @ [after E A\<^sub>0 ST c, ty\<^sub>i' ST E A\<^sub>1, ty\<^sub>i' ST E A\<^sub>0])"
+| "compT E A ST (throw e) = compT E A ST e @ [after E A ST e]"
+| "compT E A ST (e\<bullet>M(es)) =
+ compT E A ST e @ [after E A ST e] @
+ compTs E (A \<squnion> \<A> e) (ty E e # ST) es"
+| "compT E A ST (C\<bullet>\<^sub>sM(es)) = compTs E A ST es"
+| "compT E A ST (try e\<^sub>1 catch(C i) e\<^sub>2) =
+ compT E A ST e\<^sub>1 @ [after E A ST e\<^sub>1] @
+ [ty\<^sub>i' (Class C#ST) E A, ty\<^sub>i' ST (E@[Class C]) (A \<squnion> \<lfloor>{i}\<rfloor>)] @
+ compT (E@[Class C]) (A \<squnion> \<lfloor>{i}\<rfloor>) ST e\<^sub>2"
+| "compT E A ST (INIT C (Cs,b) \<leftarrow> e) = []"
+| "compT E A ST (RI(C,e');Cs \<leftarrow> e) = []"
+| "compTs E A ST [] = []"
+| "compTs E A ST (e#es) = compT E A ST e @ [after E A ST e] @
+ compTs E (A \<squnion> (\<A> e)) (ty E e # ST) es"
+
+definition compT\<^sub>a :: "ty list \<Rightarrow> nat hyperset \<Rightarrow> ty list \<Rightarrow> expr\<^sub>1 \<Rightarrow> ty\<^sub>i' list" where
+ "compT\<^sub>a E A ST e = compT E A ST e @ [after E A ST e]"
+
+end
+
+lemma compE\<^sub>2_not_Nil[simp]: "P,E \<turnstile>\<^sub>1 e :: T \<Longrightarrow> compE\<^sub>2 e \<noteq> []"
+(*<*)by(simp add: compE\<^sub>2_not_Nil' WT\<^sub>1_nsub_RI)(*>*)
+
+lemma (in TC1) compT_sizes':
+shows "\<And>E A ST. \<not>sub_RI e \<Longrightarrow> size(compT E A ST e) = size(compE\<^sub>2 e) - 1"
+and "\<And>E A ST. \<not>sub_RIs es \<Longrightarrow> size(compTs E A ST es) = size(compEs\<^sub>2 es)"
+(*<*)
+by(induct e and es rule: compE\<^sub>2.induct compEs\<^sub>2.induct)
+ (auto split:bop.splits nat_diff_split simp: compE\<^sub>2_not_Nil')
+(*>*)
+
+lemma (in TC1) compT_sizes[simp]:
+shows "\<And>E A ST. P,E \<turnstile>\<^sub>1 e :: T \<Longrightarrow> size(compT E A ST e) = size(compE\<^sub>2 e) - 1"
+and "\<And>E A ST. P,E \<turnstile>\<^sub>1 es [::] Ts \<Longrightarrow> size(compTs E A ST es) = size(compEs\<^sub>2 es)"
+(*<*)using compT_sizes' WT\<^sub>1_nsub_RI WTs\<^sub>1_nsub_RIs by auto(*>*)
+
+
+lemma (in TC1) [simp]: "\<And>ST E. \<lfloor>\<tau>\<rfloor> \<notin> set (compT E None ST e)"
+and [simp]: "\<And>ST E. \<lfloor>\<tau>\<rfloor> \<notin> set (compTs E None ST es)"
+(*<*)by(induct e and es rule: compT.induct compTs.induct) (simp_all add:after_def)(*>*)
+
+
+lemma (in TC0) pair_eq_ty\<^sub>i'_conv:
+ "(\<lfloor>(ST, LT)\<rfloor> = ty\<^sub>i' ST\<^sub>0 E A) =
+ (case A of None \<Rightarrow> False | Some A \<Rightarrow> (ST = ST\<^sub>0 \<and> LT = ty\<^sub>l E A))"
+(*<*)by(simp add:ty\<^sub>i'_def)(*>*)
+
+
+lemma (in TC0) pair_conv_ty\<^sub>i':
+ "\<lfloor>(ST, ty\<^sub>l E A)\<rfloor> = ty\<^sub>i' ST E \<lfloor>A\<rfloor>"
+(*<*)by(simp add:ty\<^sub>i'_def)(*>*)
+
+(*<*)
+declare (in TC0)
+ ty\<^sub>i'_antimono [intro!] after_def[simp]
+ pair_conv_ty\<^sub>i'[simp] pair_eq_ty\<^sub>i'_conv[simp]
+(*>*)
+
+
+lemma (in TC1) compT_LT_prefix:
+ "\<And>E A ST\<^sub>0. \<lbrakk> \<lfloor>(ST,LT)\<rfloor> \<in> set(compT E A ST\<^sub>0 e); \<B> e (size E) \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<lfloor>(ST,LT)\<rfloor> \<le>' ty\<^sub>i' ST E A"
+and
+ "\<And>E A ST\<^sub>0. \<lbrakk> \<lfloor>(ST,LT)\<rfloor> \<in> set(compTs E A ST\<^sub>0 es); \<B>s es (size E) \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<lfloor>(ST,LT)\<rfloor> \<le>' ty\<^sub>i' ST E A"
+(*<*)
+proof(induct e and es rule: compT.induct compTs.induct)
+ case FAss thus ?case by(fastforce simp:hyperset_defs elim!:sup_state_opt_trans)
+next
+ case BinOp thus ?case
+ by(fastforce simp:hyperset_defs elim!:sup_state_opt_trans split:bop.splits)
+next
+ case Seq thus ?case by(fastforce simp:hyperset_defs elim!:sup_state_opt_trans)
+next
+ case While thus ?case by(fastforce simp:hyperset_defs elim!:sup_state_opt_trans)
+next
+ case Cond thus ?case by(fastforce simp:hyperset_defs elim!:sup_state_opt_trans)
+next
+ case Block thus ?case
+ by(force simp add:hyperset_defs ty\<^sub>i'_def simp del:pair_conv_ty\<^sub>i'
+ elim!:sup_state_opt_trans)
+next
+ case Call thus ?case by(fastforce simp:hyperset_defs elim!:sup_state_opt_trans)
+next
+ case Cons_exp thus ?case
+ by(fastforce simp:hyperset_defs elim!:sup_state_opt_trans)
+next
+ case TryCatch thus ?case
+ by(fastforce simp:hyperset_defs intro!: ty\<^sub>i'_incr
+ elim!:sup_state_opt_trans)
+qed (auto simp:hyperset_defs)
+
+declare (in TC0)
+ ty\<^sub>i'_antimono [rule del] after_def[simp del]
+ pair_conv_ty\<^sub>i'[simp del] pair_eq_ty\<^sub>i'_conv[simp del]
+(*>*)
+
+
+lemma [iff]: "OK None \<in> states P mxs mxl"
+(*<*)by(simp add: JVM_states_unfold)(*>*)
+
+lemma (in TC0) after_in_states:
+assumes wf: "wf_prog p P" and wt: "P,E \<turnstile>\<^sub>1 e :: T"
+ and Etypes: "set E \<subseteq> types P" and STtypes: "set ST \<subseteq> types P"
+ and stack: "size ST + max_stack e \<le> mxs"
+shows "OK (after E A ST e) \<in> states P mxs mxl"
+(*<*)
+proof -
+ have "size ST + 1 \<le> mxs" using max_stack1[where e=e] wt stack
+ by fastforce
+ then show ?thesis using assms
+ by(simp add: after_def ty\<^sub>i'_def JVM_states_unfold ty\<^sub>l_in_types)
+ (blast intro!:listI WT\<^sub>1_is_type)
+qed
+(*>*)
+
+
+lemma (in TC0) OK_ty\<^sub>i'_in_statesI[simp]:
+ "\<lbrakk> set E \<subseteq> types P; set ST \<subseteq> types P; size ST \<le> mxs \<rbrakk>
+ \<Longrightarrow> OK (ty\<^sub>i' ST E A) \<in> states P mxs mxl"
+(*<*)
+by(simp add:ty\<^sub>i'_def JVM_states_unfold ty\<^sub>l_in_types)
+ (blast intro!:listI)
+(*>*)
+
+
+lemma is_class_type_aux: "is_class P C \<Longrightarrow> is_type P (Class C)"
+(*<*)by(simp)(*>*)
+
+(*<*)
+declare is_type_simps[simp del] subsetI[rule del]
+(*>*)
+
+theorem (in TC1) compT_states:
+assumes wf: "wf_prog p P"
+shows "\<And>E T A ST.
+ \<lbrakk> P,E \<turnstile>\<^sub>1 e :: T; set E \<subseteq> types P; set ST \<subseteq> types P;
+ size ST + max_stack e \<le> mxs; size E + max_vars e \<le> mxl \<rbrakk>
+ \<Longrightarrow> OK ` set(compT E A ST e) \<subseteq> states P mxs mxl"
+(*<*)(is "\<And>E T A ST. PROP ?P e E T A ST")(*>*)
+
+and "\<And>E Ts A ST.
+ \<lbrakk> P,E \<turnstile>\<^sub>1 es[::]Ts; set E \<subseteq> types P; set ST \<subseteq> types P;
+ size ST + max_stacks es \<le> mxs; size E + max_varss es \<le> mxl \<rbrakk>
+ \<Longrightarrow> OK ` set(compTs E A ST es) \<subseteq> states P mxs mxl"
+(*<*)(is "\<And>E Ts A ST. PROP ?Ps es E Ts A ST")
+proof(induct e and es rule: compT.induct compTs.induct)
+ case new thus ?case by(simp)
+next
+ case (Cast C e) thus ?case by (auto simp:after_in_states[OF wf])
+next
+ case Val thus ?case by(simp)
+next
+ case Var thus ?case by(simp)
+next
+ case LAss thus ?case by(auto simp:after_in_states[OF wf])
+next
+ case FAcc thus ?case by(auto simp:after_in_states[OF wf])
+next
+ case SFAcc thus ?case by(auto simp:after_in_states[OF wf])
+next
+ case FAss thus ?case
+ by(auto simp:image_Un WT\<^sub>1_is_type[OF wf] after_in_states[OF wf])
+next
+ case SFAss thus ?case
+ by(auto simp:image_Un WT\<^sub>1_is_type[OF wf] after_in_states[OF wf])
+next
+ case Seq thus ?case
+ by(auto simp:image_Un after_in_states[OF wf])
+next
+ case BinOp thus ?case
+ by(auto simp:image_Un WT\<^sub>1_is_type[OF wf] after_in_states[OF wf])
+next
+ case Cond thus ?case
+ by(force simp:image_Un WT\<^sub>1_is_type[OF wf] after_in_states[OF wf])
+next
+ case While thus ?case
+ by(auto simp:image_Un WT\<^sub>1_is_type[OF wf] after_in_states[OF wf])
+next
+ case Block thus ?case by(auto)
+next
+ case (TryCatch e\<^sub>1 C i e\<^sub>2)
+ moreover have "size ST + 1 \<le> mxs"
+ using TryCatch.prems max_stack1[where e=e\<^sub>1] by fastforce
+ ultimately show ?case
+ by(auto simp:image_Un WT\<^sub>1_is_type[OF wf] after_in_states[OF wf]
+ is_class_type_aux)
+next
+ case Nil_exp thus ?case by simp
+next
+ case Cons_exp thus ?case
+ by(auto simp:image_Un WT\<^sub>1_is_type[OF wf] after_in_states[OF wf])
+next
+ case throw thus ?case
+ by(auto simp: WT\<^sub>1_is_type[OF wf] after_in_states[OF wf])
+next
+ case Call thus ?case
+ by(auto simp:image_Un WT\<^sub>1_is_type[OF wf] after_in_states[OF wf])
+next
+ case SCall thus ?case
+ by(auto simp:image_Un WT\<^sub>1_is_type[OF wf] after_in_states[OF wf])
+next
+ case INIT thus ?case by simp
+next
+ case RI thus ?case by simp
+qed
+
+declare is_type_simps[simp] subsetI[intro!]
+(*>*)
+
+
+definition shift :: "nat \<Rightarrow> ex_table \<Rightarrow> ex_table"
+where
+ "shift n xt \<equiv> map (\<lambda>(from,to,C,handler,depth). (from+n,to+n,C,handler+n,depth)) xt"
+
+
+lemma [simp]: "shift 0 xt = xt"
+(*<*)by(induct xt)(auto simp:shift_def)(*>*)
+
+lemma [simp]: "shift n [] = []"
+(*<*)by(simp add:shift_def)(*>*)
+
+lemma [simp]: "shift n (xt\<^sub>1 @ xt\<^sub>2) = shift n xt\<^sub>1 @ shift n xt\<^sub>2"
+(*<*)by(simp add:shift_def)(*>*)
+
+lemma [simp]: "shift m (shift n xt) = shift (m+n) xt"
+(*<*)by(induct xt)(auto simp:shift_def)(*>*)
+
+lemma [simp]: "pcs (shift n xt) = {pc+n|pc. pc \<in> pcs xt}"
+(*<*)
+proof -
+ { fix x f t C h d
+ assume "(f,t,C,h,d) \<in> set xt" and "f + n \<le> x"
+ and "x < t + n"
+ then have "\<exists>pc. x = pc + n \<and> (\<exists>x\<in>set xt. pc \<in> (case x of (f, t, C, h, d) \<Rightarrow> {f..<t}))"
+ by(rule_tac x = "x-n" in exI) (force split:nat_diff_split)
+ }
+ then show ?thesis by(auto simp:shift_def pcs_def) fast
+qed
+(*>*)
+
+
+lemma shift_compxE\<^sub>2:
+shows "\<And>pc pc' d. shift pc (compxE\<^sub>2 e pc' d) = compxE\<^sub>2 e (pc' + pc) d"
+and "\<And>pc pc' d. shift pc (compxEs\<^sub>2 es pc' d) = compxEs\<^sub>2 es (pc' + pc) d"
+(*<*)
+by(induct e and es rule: compxE\<^sub>2.induct compxEs\<^sub>2.induct)
+ (auto simp:shift_def ac_simps)
+(*>*)
+
+
+lemma compxE\<^sub>2_size_convs[simp]:
+shows "n \<noteq> 0 \<Longrightarrow> compxE\<^sub>2 e n d = shift n (compxE\<^sub>2 e 0 d)"
+and "n \<noteq> 0 \<Longrightarrow> compxEs\<^sub>2 es n d = shift n (compxEs\<^sub>2 es 0 d)"
+(*<*)by(simp_all add:shift_compxE\<^sub>2)(*>*)
+
+locale TC2 = TC1 +
+ fixes T\<^sub>r :: ty and mxs :: pc
+begin
+
+definition
+ wt_instrs :: "instr list \<Rightarrow> ex_table \<Rightarrow> ty\<^sub>i' list \<Rightarrow> bool"
+ ("(\<turnstile> _, _ /[::]/ _)" [0,0,51] 50) where
+ "\<turnstile> is,xt [::] \<tau>s \<longleftrightarrow> size is < size \<tau>s \<and> pcs xt \<subseteq> {0..<size is} \<and>
+ (\<forall>pc< size is. P,T\<^sub>r,mxs,size \<tau>s,xt \<turnstile> is!pc,pc :: \<tau>s)"
+
+end
+
+notation TC2.wt_instrs ("(_,_,_ \<turnstile>/ _, _ /[::]/ _)" [50,50,50,50,50,51] 50)
+
+(*<*)
+lemmas (in TC2) wt_defs =
+ wt_instrs_def wt_instr_def app_def eff_def norm_eff_def
+(*>*)
+
+lemma (in TC2) [simp]: "\<tau>s \<noteq> [] \<Longrightarrow> \<turnstile> [],[] [::] \<tau>s"
+(*<*) by (simp add: wt_defs) (*>*)
+
+lemma [simp]: "eff i P pc et None = []"
+(*<*)by (simp add: Effect.eff_def)(*>*)
+
+(*<*)
+declare split_comp_eq[simp del]
+(*>*)
+
+lemma wt_instr_appR:
+ "\<lbrakk> P,T,m,mpc,xt \<turnstile> is!pc,pc :: \<tau>s;
+ pc < size is; size is < size \<tau>s; mpc \<le> size \<tau>s; mpc \<le> mpc' \<rbrakk>
+ \<Longrightarrow> P,T,m,mpc',xt \<turnstile> is!pc,pc :: \<tau>s@\<tau>s'"
+(*<*)by (fastforce simp:wt_instr_def app_def)(*>*)
+
+
+lemma relevant_entries_shift [simp]:
+ "relevant_entries P i (pc+n) (shift n xt) = shift n (relevant_entries P i pc xt)"
+(*<*)
+proof(induct xt)
+ case Nil
+ then show ?case by (simp add: relevant_entries_def shift_def)
+next
+ case (Cons a xt)
+ then show ?case
+ by (auto simp add: relevant_entries_def shift_def is_relevant_entry_def)
+qed
+(*>*)
+
+
+lemma [simp]:
+ "xcpt_eff i P (pc+n) \<tau> (shift n xt) =
+ map (\<lambda>(pc,\<tau>). (pc + n, \<tau>)) (xcpt_eff i P pc \<tau> xt)"
+(*<*)
+proof -
+ obtain ST LT where "\<tau> = (ST, LT)" by(cases \<tau>) simp
+ then show ?thesis by(simp add: xcpt_eff_def) (auto simp add: shift_def)
+qed
+(*>*)
+
+
+lemma [simp]:
+ "app\<^sub>i (i, P, pc, m, T, \<tau>) \<Longrightarrow>
+ eff i P (pc+n) (shift n xt) (Some \<tau>) =
+ map (\<lambda>(pc,\<tau>). (pc+n,\<tau>)) (eff i P pc xt (Some \<tau>))"
+(*<*)by(cases "i") (auto simp add:eff_def norm_eff_def)(*>*)
+
+
+lemma [simp]:
+ "xcpt_app i P (pc+n) mxs (shift n xt) \<tau> = xcpt_app i P pc mxs xt \<tau>"
+(*<*)by (simp add: xcpt_app_def) (auto simp add: shift_def)(*>*)
+
+
+lemma wt_instr_appL:
+assumes "P,T,m,mpc,xt \<turnstile> i,pc :: \<tau>s" and "pc < size \<tau>s" and "mpc \<le> size \<tau>s"
+shows "P,T,m,mpc + size \<tau>s',shift (size \<tau>s') xt \<turnstile> i,pc+size \<tau>s' :: \<tau>s'@\<tau>s"
+(*<*)
+proof -
+ let ?t = "(\<tau>s'@\<tau>s)!(pc+size \<tau>s')"
+ show ?thesis
+ proof(cases ?t)
+ case (Some \<tau>)
+ obtain ST LT where [simp]: "\<tau> = (ST, LT)" by(cases \<tau>) simp
+ have "app\<^sub>i (i, P, pc + length \<tau>s', m, T, \<tau>)"
+ using Some assms by(cases "i") (auto simp:wt_instr_def app_def)
+ moreover {
+ fix pc' \<tau>' assume "(pc',\<tau>') \<in> set (eff i P pc xt ?t)"
+ then have "P \<turnstile> \<tau>' \<le>' \<tau>s!pc'" and "pc' < mpc"
+ using Some assms by(auto simp:wt_instr_def app_def)
+ }
+ ultimately show ?thesis using Some assms
+ by(fastforce simp:wt_instr_def app_def)
+ qed (auto simp:wt_instr_def app_def)
+qed
+(*>*)
+
+
+lemma wt_instr_Cons:
+assumes wti: "P,T,m,mpc - 1,[] \<turnstile> i,pc - 1 :: \<tau>s"
+ and pcl: "0 < pc" and mpcl: "0 < mpc"
+ and pcu: "pc < size \<tau>s + 1" and mpcu: "mpc \<le> size \<tau>s + 1"
+shows "P,T,m,mpc,[] \<turnstile> i,pc :: \<tau>#\<tau>s"
+(*<*)
+proof -
+ have "pc - 1 < length \<tau>s" using pcl pcu by arith
+ moreover have "mpc - 1 \<le> length \<tau>s" using mpcl mpcu by arith
+ ultimately have
+ "P,T,m,mpc - 1 + length [\<tau>],shift (length [\<tau>]) [] \<turnstile> i,pc - 1 + length [\<tau>] :: [\<tau>] @ \<tau>s"
+ by(rule wt_instr_appL[where \<tau>s' = "[\<tau>]", OF wti])
+ then show ?thesis using pcl mpcl by (simp split:nat_diff_split_asm)
+qed
+(*>*)
+
+
+lemma wt_instr_append:
+assumes wti: "P,T,m,mpc - size \<tau>s',[] \<turnstile> i,pc - size \<tau>s' :: \<tau>s"
+ and pcl: "size \<tau>s' \<le> pc" and mpcl: "size \<tau>s' \<le> mpc"
+ and pcu: "pc < size \<tau>s + size \<tau>s'" and mpcu: "mpc \<le> size \<tau>s + size \<tau>s'"
+shows "P,T,m,mpc,[] \<turnstile> i,pc :: \<tau>s'@\<tau>s"
+(*<*)
+proof -
+ have "pc - length \<tau>s' < length \<tau>s" using pcl pcu by arith
+ moreover have "mpc - length \<tau>s' \<le> length \<tau>s" using mpcl mpcu by arith
+thm wt_instr_appL[where \<tau>s' = "\<tau>s'", OF wti]
+ ultimately have "P,T,m,mpc - length \<tau>s' + length \<tau>s',shift (length \<tau>s') []
+ \<turnstile> i,pc - length \<tau>s' + length \<tau>s' :: \<tau>s' @ \<tau>s"
+ by(rule wt_instr_appL[where \<tau>s' = "\<tau>s'", OF wti])
+ then show ?thesis using pcl mpcl by (simp split:nat_diff_split_asm)
+qed
+(*>*)
+
+
+lemma xcpt_app_pcs:
+ "pc \<notin> pcs xt \<Longrightarrow> xcpt_app i P pc mxs xt \<tau>"
+(*<*)
+by (auto simp add: xcpt_app_def relevant_entries_def is_relevant_entry_def pcs_def)
+(*>*)
+
+
+lemma xcpt_eff_pcs:
+ "pc \<notin> pcs xt \<Longrightarrow> xcpt_eff i P pc \<tau> xt = []"
+(*<*)
+by (cases \<tau>)
+ (auto simp add: is_relevant_entry_def xcpt_eff_def relevant_entries_def pcs_def
+ intro!: filter_False)
+(*>*)
+
+
+lemma pcs_shift:
+ "pc < n \<Longrightarrow> pc \<notin> pcs (shift n xt)"
+(*<*)by (auto simp add: shift_def pcs_def)(*>*)
+
+
+lemma wt_instr_appRx:
+ "\<lbrakk> P,T,m,mpc,xt \<turnstile> is!pc,pc :: \<tau>s; pc < size is; size is < size \<tau>s; mpc \<le> size \<tau>s \<rbrakk>
+ \<Longrightarrow> P,T,m,mpc,xt @ shift (size is) xt' \<turnstile> is!pc,pc :: \<tau>s"
+(*<*)by (auto simp:wt_instr_def eff_def app_def xcpt_app_pcs xcpt_eff_pcs)(*>*)
+
+
+lemma wt_instr_appLx:
+ "\<lbrakk> P,T,m,mpc,xt \<turnstile> i,pc :: \<tau>s; pc \<notin> pcs xt' \<rbrakk>
+ \<Longrightarrow> P,T,m,mpc,xt'@xt \<turnstile> i,pc :: \<tau>s"
+(*<*)by (auto simp:wt_instr_def app_def eff_def xcpt_app_pcs xcpt_eff_pcs)(*>*)
+
+
+lemma (in TC2) wt_instrs_extR:
+ "\<turnstile> is,xt [::] \<tau>s \<Longrightarrow> \<turnstile> is,xt [::] \<tau>s @ \<tau>s'"
+(*<*)by(auto simp add:wt_instrs_def wt_instr_appR)(*>*)
+
+
+lemma (in TC2) wt_instrs_ext:
+assumes wt\<^sub>1: "\<turnstile> is\<^sub>1,xt\<^sub>1 [::] \<tau>s\<^sub>1@\<tau>s\<^sub>2" and wt\<^sub>2: "\<turnstile> is\<^sub>2,xt\<^sub>2 [::] \<tau>s\<^sub>2"
+ and \<tau>s_size: "size \<tau>s\<^sub>1 = size is\<^sub>1"
+ shows "\<turnstile> is\<^sub>1@is\<^sub>2, xt\<^sub>1 @ shift (size is\<^sub>1) xt\<^sub>2 [::] \<tau>s\<^sub>1@\<tau>s\<^sub>2"
+(*<*)
+proof -
+ let ?is = "is\<^sub>1@is\<^sub>2" and ?xt = "xt\<^sub>1 @ shift (size is\<^sub>1) xt\<^sub>2"
+ and ?\<tau>s = "\<tau>s\<^sub>1@\<tau>s\<^sub>2"
+ have "size ?is < size ?\<tau>s" using wt\<^sub>2 \<tau>s_size by(fastforce simp:wt_instrs_def)
+ moreover have "pcs ?xt \<subseteq> {0..<size ?is}" using wt\<^sub>1 wt\<^sub>2
+ by(fastforce simp:wt_instrs_def)
+ moreover {
+ fix pc assume pc: "pc<size ?is"
+ have "P,T\<^sub>r,mxs,size ?\<tau>s,?xt \<turnstile> ?is!pc,pc :: ?\<tau>s"
+ proof(cases "pc < length is\<^sub>1")
+ case True then show ?thesis using wt\<^sub>1 pc
+ by(fastforce simp: wt_instrs_def wt_instr_appRx)
+ next
+ case False
+ then have "pc - length is\<^sub>1 < length is\<^sub>2" using pc by fastforce
+ then have "P,T\<^sub>r,mxs,length \<tau>s\<^sub>2,xt\<^sub>2 \<turnstile> is\<^sub>2 ! (pc - length is\<^sub>1),pc - length is\<^sub>1 :: \<tau>s\<^sub>2"
+ using wt\<^sub>2 by(clarsimp simp: wt_instrs_def)
+ moreover have "pc - length is\<^sub>1 < length \<tau>s\<^sub>2" using pc wt\<^sub>2
+ by(clarsimp simp: wt_instrs_def) arith
+ moreover have "length \<tau>s\<^sub>2 \<le> length \<tau>s\<^sub>2" by simp
+ moreover have "pc - length is\<^sub>1 + length \<tau>s\<^sub>1 \<notin> pcs xt\<^sub>1" using wt\<^sub>1 \<tau>s_size
+ by(fastforce simp: wt_instrs_def)
+ ultimately have "P,T\<^sub>r,mxs,length \<tau>s\<^sub>2 + length \<tau>s\<^sub>1,xt\<^sub>1 @ shift (length \<tau>s\<^sub>1) xt\<^sub>2
+ \<turnstile> is\<^sub>2 ! (pc - length is\<^sub>1),pc - length is\<^sub>1 + length \<tau>s\<^sub>1 :: \<tau>s\<^sub>1 @ \<tau>s\<^sub>2"
+ by(rule wt_instr_appLx[OF wt_instr_appL[where \<tau>s' = "\<tau>s\<^sub>1"]])
+ then show ?thesis using False \<tau>s_size by(simp add:add.commute)
+ qed
+ }
+ ultimately show ?thesis by(clarsimp simp:wt_instrs_def)
+qed
+(*>*)
+
+corollary (in TC2) wt_instrs_ext2:
+ "\<lbrakk> \<turnstile> is\<^sub>2,xt\<^sub>2 [::] \<tau>s\<^sub>2; \<turnstile> is\<^sub>1,xt\<^sub>1 [::] \<tau>s\<^sub>1@\<tau>s\<^sub>2; size \<tau>s\<^sub>1 = size is\<^sub>1 \<rbrakk>
+ \<Longrightarrow> \<turnstile> is\<^sub>1@is\<^sub>2, xt\<^sub>1 @ shift (size is\<^sub>1) xt\<^sub>2 [::] \<tau>s\<^sub>1@\<tau>s\<^sub>2"
+(*<*)by(rule wt_instrs_ext)(*>*)
+
+
+corollary (in TC2) wt_instrs_ext_prefix [trans]:
+ "\<lbrakk> \<turnstile> is\<^sub>1,xt\<^sub>1 [::] \<tau>s\<^sub>1@\<tau>s\<^sub>2; \<turnstile> is\<^sub>2,xt\<^sub>2 [::] \<tau>s\<^sub>3;
+ size \<tau>s\<^sub>1 = size is\<^sub>1; prefix \<tau>s\<^sub>3 \<tau>s\<^sub>2 \<rbrakk>
+ \<Longrightarrow> \<turnstile> is\<^sub>1@is\<^sub>2, xt\<^sub>1 @ shift (size is\<^sub>1) xt\<^sub>2 [::] \<tau>s\<^sub>1@\<tau>s\<^sub>2"
+(*<*)by(bestsimp simp:prefix_def elim: wt_instrs_ext dest:wt_instrs_extR)(*>*)
+
+
+corollary (in TC2) wt_instrs_app:
+ assumes is\<^sub>1: "\<turnstile> is\<^sub>1,xt\<^sub>1 [::] \<tau>s\<^sub>1@[\<tau>]"
+ assumes is\<^sub>2: "\<turnstile> is\<^sub>2,xt\<^sub>2 [::] \<tau>#\<tau>s\<^sub>2"
+ assumes s: "size \<tau>s\<^sub>1 = size is\<^sub>1"
+ shows "\<turnstile> is\<^sub>1@is\<^sub>2, xt\<^sub>1@shift (size is\<^sub>1) xt\<^sub>2 [::] \<tau>s\<^sub>1@\<tau>#\<tau>s\<^sub>2"
+(*<*)
+proof -
+ from is\<^sub>1 have "\<turnstile> is\<^sub>1,xt\<^sub>1 [::] (\<tau>s\<^sub>1@[\<tau>])@\<tau>s\<^sub>2"
+ by (rule wt_instrs_extR)
+ hence "\<turnstile> is\<^sub>1,xt\<^sub>1 [::] \<tau>s\<^sub>1@\<tau>#\<tau>s\<^sub>2" by simp
+ from this is\<^sub>2 s show ?thesis by (rule wt_instrs_ext)
+qed
+(*>*)
+
+
+corollary (in TC2) wt_instrs_app_last[trans]:
+assumes "\<turnstile> is\<^sub>2,xt\<^sub>2 [::] \<tau>#\<tau>s\<^sub>2" "\<turnstile> is\<^sub>1,xt\<^sub>1 [::] \<tau>s\<^sub>1"
+ "last \<tau>s\<^sub>1 = \<tau>" "size \<tau>s\<^sub>1 = size is\<^sub>1+1"
+shows "\<turnstile> is\<^sub>1@is\<^sub>2, xt\<^sub>1@shift (size is\<^sub>1) xt\<^sub>2 [::] \<tau>s\<^sub>1@\<tau>s\<^sub>2"
+(*<*)
+using assms proof(cases \<tau>s\<^sub>1 rule:rev_cases)
+ case (snoc ys y)
+ then show ?thesis using assms by(simp add:wt_instrs_app)
+qed simp
+(*>*)
+
+
+corollary (in TC2) wt_instrs_append_last[trans]:
+assumes wtis: "\<turnstile> is,xt [::] \<tau>s" and wti: "P,T\<^sub>r,mxs,mpc,[] \<turnstile> i,pc :: \<tau>s"
+ and pc: "pc = size is" and mpc: "mpc = size \<tau>s" and is_\<tau>s: "size is + 1 < size \<tau>s"
+shows "\<turnstile> is@[i],xt [::] \<tau>s"
+(*<*)
+proof -
+ have pc_xt: "pc \<notin> pcs xt" using wtis pc by (fastforce simp:wt_instrs_def)
+ have "pcs xt \<subseteq> {..<Suc (length is)}" using wtis by (fastforce simp:wt_instrs_def)
+ moreover {
+ fix pc' assume pc': "\<not> pc' < length is" "pc' < Suc (length is)"
+ have "P,T\<^sub>r,mxs,length \<tau>s,xt \<turnstile> i,pc' :: \<tau>s"
+ using wt_instr_appLx[where xt = "[]",simplified,OF wti pc_xt]
+ less_antisym[OF pc'] pc mpc by simp
+ }
+ ultimately show ?thesis using wtis is_\<tau>s by(clarsimp simp add:wt_instrs_def)
+qed
+(*>*)
+
+
+corollary (in TC2) wt_instrs_app2:
+ "\<lbrakk> \<turnstile> is\<^sub>2,xt\<^sub>2 [::] \<tau>'#\<tau>s\<^sub>2; \<turnstile> is\<^sub>1,xt\<^sub>1 [::] \<tau>#\<tau>s\<^sub>1@[\<tau>'];
+ xt' = xt\<^sub>1 @ shift (size is\<^sub>1) xt\<^sub>2; size \<tau>s\<^sub>1+1 = size is\<^sub>1 \<rbrakk>
+ \<Longrightarrow> \<turnstile> is\<^sub>1@is\<^sub>2,xt' [::] \<tau>#\<tau>s\<^sub>1@\<tau>'#\<tau>s\<^sub>2"
+(*<*)using wt_instrs_app[where ?\<tau>s\<^sub>1.0 = "\<tau> # \<tau>s\<^sub>1"] by simp (*>*)
+
+
+corollary (in TC2) wt_instrs_app2_simp[trans,simp]:
+ "\<lbrakk> \<turnstile> is\<^sub>2,xt\<^sub>2 [::] \<tau>'#\<tau>s\<^sub>2; \<turnstile> is\<^sub>1,xt\<^sub>1 [::] \<tau>#\<tau>s\<^sub>1@[\<tau>']; size \<tau>s\<^sub>1+1 = size is\<^sub>1 \<rbrakk>
+ \<Longrightarrow> \<turnstile> is\<^sub>1@is\<^sub>2, xt\<^sub>1@shift (size is\<^sub>1) xt\<^sub>2 [::] \<tau>#\<tau>s\<^sub>1@\<tau>'#\<tau>s\<^sub>2"
+(*<*)using wt_instrs_app[where ?\<tau>s\<^sub>1.0 = "\<tau> # \<tau>s\<^sub>1"] by simp(*>*)
+
+
+corollary (in TC2) wt_instrs_Cons[simp]:
+ "\<lbrakk> \<tau>s \<noteq> []; \<turnstile> [i],[] [::] [\<tau>,\<tau>']; \<turnstile> is,xt [::] \<tau>'#\<tau>s \<rbrakk>
+ \<Longrightarrow> \<turnstile> i#is,shift 1 xt [::] \<tau>#\<tau>'#\<tau>s"
+(*<*)
+using wt_instrs_app2[where ?is\<^sub>1.0 = "[i]" and ?\<tau>s\<^sub>1.0 = "[]" and ?is\<^sub>2.0 = "is"
+ and ?xt\<^sub>1.0 = "[]"]
+by simp
+
+
+corollary (in TC2) wt_instrs_Cons2[trans]:
+ assumes \<tau>s: "\<turnstile> is,xt [::] \<tau>s"
+ assumes i: "P,T\<^sub>r,mxs,mpc,[] \<turnstile> i,0 :: \<tau>#\<tau>s"
+ assumes mpc: "mpc = size \<tau>s + 1"
+ shows "\<turnstile> i#is,shift 1 xt [::] \<tau>#\<tau>s"
+(*<*)
+proof -
+ from \<tau>s have "\<tau>s \<noteq> []" by (auto simp: wt_instrs_def)
+ with mpc i have "\<turnstile> [i],[] [::] [\<tau>]@\<tau>s" by (simp add: wt_instrs_def)
+ with \<tau>s show ?thesis by (fastforce dest: wt_instrs_ext)
+qed
+(*>*)
+
+
+lemma (in TC2) wt_instrs_last_incr[trans]:
+assumes wtis: "\<turnstile> is,xt [::] \<tau>s@[\<tau>]" and ss: "P \<turnstile> \<tau> \<le>' \<tau>'"
+shows "\<turnstile> is,xt [::] \<tau>s@[\<tau>']"
+(*<*)
+proof -
+ let ?\<tau>s = "\<tau>s@[\<tau>]" and ?\<tau>s' = "\<tau>s@[\<tau>']"
+ { fix pc assume pc: "pc< size is"
+ let ?i = "is!pc"
+ have app_pc: "app (is ! pc) P mxs T\<^sub>r pc (length ?\<tau>s) xt (\<tau>s ! pc)"
+ using wtis pc by(clarsimp simp add:wt_instrs_def wt_instr_def)
+ then have Apc\<tau>': "\<And>pc' \<tau>'. (pc',\<tau>') \<in> set (eff ?i P pc xt (?\<tau>s!pc))
+ \<Longrightarrow> pc' < length ?\<tau>s"
+ using wtis pc by(fastforce simp add:wt_instrs_def app_def)
+ have Aepc\<tau>': "\<And>pc' \<tau>'. (pc',\<tau>') \<in> set (eff ?i P pc xt (?\<tau>s!pc))
+ \<Longrightarrow> P \<turnstile> \<tau>' \<le>' ?\<tau>s!pc'"
+ using wtis pc by(fastforce simp add:wt_instrs_def wt_instr_def)
+ { fix pc1 \<tau>1 assume pc\<tau>1: "(pc1,\<tau>1) \<in> set (eff ?i P pc xt (?\<tau>s'!pc))"
+ then have epc\<tau>': "(pc1,\<tau>1) \<in> set (eff ?i P pc xt (?\<tau>s!pc))"
+ using wtis pc by(simp add:wt_instrs_def)
+ have "P \<turnstile> \<tau>1 \<le>' ?\<tau>s'!pc1"
+ proof(cases "pc1 < length \<tau>s")
+ case True
+ then show ?thesis using wtis pc pc\<tau>1
+ by(fastforce simp add:wt_instrs_def wt_instr_def)
+ next
+ case False
+ then have "pc1 < length ?\<tau>s" using Apc\<tau>'[OF epc\<tau>'] by simp
+ then have [simp]: "pc1 = size \<tau>s" using False by clarsimp
+ have "P \<turnstile> \<tau>1 \<le>' \<tau>" using Aepc\<tau>'[OF epc\<tau>'] by simp
+ then have "P \<turnstile> \<tau>1 \<le>' \<tau>'" by(rule sup_state_opt_trans[OF _ ss])
+ then show ?thesis by simp
+ qed
+ }
+ then have "P,T\<^sub>r,mxs,size ?\<tau>s',xt \<turnstile> is!pc,pc :: ?\<tau>s'" using wtis pc
+ by(clarsimp simp add:wt_instrs_def wt_instr_def)
+ }
+ then show ?thesis using wtis by(simp add:wt_instrs_def)
+qed
+(*>*)
+
+
+lemma [iff]: "xcpt_app i P pc mxs [] \<tau>"
+(*<*)by (simp add: xcpt_app_def relevant_entries_def)(*>*)
+
+
+lemma [simp]: "xcpt_eff i P pc \<tau> [] = []"
+(*<*)by (simp add: xcpt_eff_def relevant_entries_def)(*>*)
+
+
+lemma (in TC2) wt_New:
+ "\<lbrakk> is_class P C; size ST < mxs \<rbrakk> \<Longrightarrow>
+ \<turnstile> [New C],[] [::] [ty\<^sub>i' ST E A, ty\<^sub>i' (Class C#ST) E A]"
+(*<*)by(simp add:wt_defs ty\<^sub>i'_def)(*>*)
+
+
+lemma (in TC2) wt_Cast:
+ "is_class P C \<Longrightarrow>
+ \<turnstile> [Checkcast C],[] [::] [ty\<^sub>i' (Class D # ST) E A, ty\<^sub>i' (Class C # ST) E A]"
+(*<*)by(simp add: ty\<^sub>i'_def wt_defs)(*>*)
+
+
+lemma (in TC2) wt_Push:
+ "\<lbrakk> size ST < mxs; typeof v = Some T \<rbrakk>
+ \<Longrightarrow> \<turnstile> [Push v],[] [::] [ty\<^sub>i' ST E A, ty\<^sub>i' (T#ST) E A]"
+(*<*)by(simp add: ty\<^sub>i'_def wt_defs)(*>*)
+
+
+lemma (in TC2) wt_Pop:
+ "\<turnstile> [Pop],[] [::] (ty\<^sub>i' (T#ST) E A # ty\<^sub>i' ST E A # \<tau>s)"
+(*<*)by(simp add: ty\<^sub>i'_def wt_defs)(*>*)
+
+
+lemma (in TC2) wt_CmpEq:
+ "\<lbrakk> P \<turnstile> T\<^sub>1 \<le> T\<^sub>2 \<or> P \<turnstile> T\<^sub>2 \<le> T\<^sub>1\<rbrakk>
+ \<Longrightarrow> \<turnstile> [CmpEq],[] [::] [ty\<^sub>i' (T\<^sub>2 # T\<^sub>1 # ST) E A, ty\<^sub>i' (Boolean # ST) E A]"
+(*<*) by(auto simp:ty\<^sub>i'_def wt_defs elim!: refTE not_refTE) (*>*)
+
+
+lemma (in TC2) wt_IAdd:
+ "\<turnstile> [IAdd],[] [::] [ty\<^sub>i' (Integer#Integer#ST) E A, ty\<^sub>i' (Integer#ST) E A]"
+(*<*)by(simp add:ty\<^sub>i'_def wt_defs)(*>*)
+
+
+lemma (in TC2) wt_Load:
+ "\<lbrakk> size ST < mxs; size E \<le> mxl; i \<in>\<in> A; i < size E \<rbrakk>
+ \<Longrightarrow> \<turnstile> [Load i],[] [::] [ty\<^sub>i' ST E A, ty\<^sub>i' (E!i # ST) E A]"
+(*<*)by(auto simp add:ty\<^sub>i'_def wt_defs ty\<^sub>l_def hyperset_defs)(*>*)
+
+
+lemma (in TC2) wt_Store:
+ "\<lbrakk> P \<turnstile> T \<le> E!i; i < size E; size E \<le> mxl \<rbrakk> \<Longrightarrow>
+ \<turnstile> [Store i],[] [::] [ty\<^sub>i' (T#ST) E A, ty\<^sub>i' ST E (\<lfloor>{i}\<rfloor> \<squnion> A)]"
+(*<*)
+by(auto simp:hyperset_defs nth_list_update ty\<^sub>i'_def wt_defs ty\<^sub>l_def
+ intro:list_all2_all_nthI)
+(*>*)
+
+
+lemma (in TC2) wt_Get:
+ "\<lbrakk> P \<turnstile> C sees F,NonStatic:T in D \<rbrakk> \<Longrightarrow>
+ \<turnstile> [Getfield F D],[] [::] [ty\<^sub>i' (Class C # ST) E A, ty\<^sub>i' (T # ST) E A]"
+(*<*)by(auto simp: ty\<^sub>i'_def wt_defs dest: sees_field_idemp sees_field_decl_above)(*>*)
+
+lemma (in TC2) wt_GetS:
+ "\<lbrakk> size ST < mxs; P \<turnstile> C sees F,Static:T in D \<rbrakk> \<Longrightarrow>
+ \<turnstile> [Getstatic C F D],[] [::] [ty\<^sub>i' ST E A, ty\<^sub>i' (T # ST) E A]"
+(*<*)by(auto simp: ty\<^sub>i'_def wt_defs dest: sees_field_idemp sees_field_decl_above)(*>*)
+
+lemma (in TC2) wt_Put:
+ "\<lbrakk> P \<turnstile> C sees F,NonStatic:T in D; P \<turnstile> T' \<le> T \<rbrakk> \<Longrightarrow>
+ \<turnstile> [Putfield F D],[] [::] [ty\<^sub>i' (T' # Class C # ST) E A, ty\<^sub>i' ST E A]"
+(*<*)by(auto intro: sees_field_idemp sees_field_decl_above simp: ty\<^sub>i'_def wt_defs)(*>*)
+
+lemma (in TC2) wt_PutS:
+ "\<lbrakk> P \<turnstile> C sees F,Static:T in D; P \<turnstile> T' \<le> T \<rbrakk> \<Longrightarrow>
+ \<turnstile> [Putstatic C F D],[] [::] [ty\<^sub>i' (T' # ST) E A, ty\<^sub>i' ST E A]"
+(*<*)by(auto intro: sees_field_idemp sees_field_decl_above simp: ty\<^sub>i'_def wt_defs)(*>*)
+
+
+lemma (in TC2) wt_Throw:
+ "\<turnstile> [Throw],[] [::] [ty\<^sub>i' (Class C # ST) E A, \<tau>']"
+(*<*)by(auto simp: ty\<^sub>i'_def wt_defs)(*>*)
+
+
+lemma (in TC2) wt_IfFalse:
+ "\<lbrakk> 2 \<le> i; nat i < size \<tau>s + 2; P \<turnstile> ty\<^sub>i' ST E A \<le>' \<tau>s ! nat(i - 2) \<rbrakk>
+ \<Longrightarrow> \<turnstile> [IfFalse i],[] [::] ty\<^sub>i' (Boolean # ST) E A # ty\<^sub>i' ST E A # \<tau>s"
+(*<*)
+by(simp add: ty\<^sub>i'_def wt_defs eval_nat_numeral nat_diff_distrib)
+(*>*)
+
+
+lemma wt_Goto:
+ "\<lbrakk> 0 \<le> int pc + i; nat (int pc + i) < size \<tau>s; size \<tau>s \<le> mpc;
+ P \<turnstile> \<tau>s!pc \<le>' \<tau>s ! nat (int pc + i) \<rbrakk>
+ \<Longrightarrow> P,T,mxs,mpc,[] \<turnstile> Goto i,pc :: \<tau>s"
+(*<*)by(clarsimp simp add: TC2.wt_defs)(*>*)
+
+
+lemma (in TC2) wt_Invoke:
+ "\<lbrakk> size es = size Ts'; P \<turnstile> C sees M,NonStatic: Ts\<rightarrow>T = m in D; P \<turnstile> Ts' [\<le>] Ts \<rbrakk>
+ \<Longrightarrow> \<turnstile> [Invoke M (size es)],[] [::] [ty\<^sub>i' (rev Ts' @ Class C # ST) E A, ty\<^sub>i' (T#ST) E A]"
+(*<*)by(fastforce simp add: ty\<^sub>i'_def wt_defs)(*>*)
+
+lemma (in TC2) wt_Invokestatic:
+ "\<lbrakk> size ST < mxs; size es = size Ts'; M \<noteq> clinit;
+ P \<turnstile> C sees M,Static: Ts\<rightarrow>T = m in D; P \<turnstile> Ts' [\<le>] Ts \<rbrakk>
+ \<Longrightarrow> \<turnstile> [Invokestatic C M (size es)],[] [::] [ty\<^sub>i' (rev Ts' @ ST) E A, ty\<^sub>i' (T#ST) E A]"
+(*<*)by(fastforce simp add: ty\<^sub>i'_def wt_defs)(*>*)
+
+
+corollary (in TC2) wt_instrs_app3[simp]:
+ "\<lbrakk> \<turnstile> is\<^sub>2,[] [::] (\<tau>' # \<tau>s\<^sub>2); \<turnstile> is\<^sub>1,xt\<^sub>1 [::] \<tau> # \<tau>s\<^sub>1 @ [\<tau>']; size \<tau>s\<^sub>1+1 = size is\<^sub>1\<rbrakk>
+ \<Longrightarrow> \<turnstile> (is\<^sub>1 @ is\<^sub>2),xt\<^sub>1 [::] \<tau> # \<tau>s\<^sub>1 @ \<tau>' # \<tau>s\<^sub>2"
+(*<*)using wt_instrs_app2[where ?xt\<^sub>2.0 = "[]"] by (simp add:shift_def)(*>*)
+
+
+corollary (in TC2) wt_instrs_Cons3[simp]:
+ "\<lbrakk> \<tau>s \<noteq> []; \<turnstile> [i],[] [::] [\<tau>,\<tau>']; \<turnstile> is,[] [::] \<tau>'#\<tau>s \<rbrakk>
+ \<Longrightarrow> \<turnstile> (i # is),[] [::] \<tau> # \<tau>' # \<tau>s"
+(*<*)
+using wt_instrs_Cons[where ?xt = "[]"]
+by (simp add:shift_def)
+
+(*<*)
+declare nth_append[simp del]
+declare [[simproc del: list_to_set_comprehension]]
+(*>*)
+
+lemma (in TC2) wt_instrs_xapp[trans]:
+assumes wtis: "\<turnstile> is\<^sub>1 @ is\<^sub>2, xt [::] \<tau>s\<^sub>1 @ ty\<^sub>i' (Class C # ST) E A # \<tau>s\<^sub>2"
+ and P_\<tau>s\<^sub>1: "\<forall>\<tau> \<in> set \<tau>s\<^sub>1. \<forall>ST' LT'. \<tau> = Some(ST',LT') \<longrightarrow>
+ size ST \<le> size ST' \<and> P \<turnstile> Some (drop (size ST' - size ST) ST',LT') \<le>' ty\<^sub>i' ST E A"
+ and is_\<tau>s: "size is\<^sub>1 = size \<tau>s\<^sub>1" and PC: "is_class P C" and ST_mxs: "size ST < mxs"
+shows "\<turnstile> is\<^sub>1 @ is\<^sub>2, xt @ [(0,size is\<^sub>1 - 1,C,size is\<^sub>1,size ST)] [::] \<tau>s\<^sub>1 @ ty\<^sub>i' (Class C # ST) E A # \<tau>s\<^sub>2"
+(*<*)(is "\<turnstile> ?is, xt@[?xte] [::] ?\<tau>s" is "\<turnstile> ?is, ?xt' [::] ?\<tau>s")
+proof -
+ have P_\<tau>s\<^sub>1': "\<And>\<tau>. \<tau> \<in> set \<tau>s\<^sub>1 \<Longrightarrow> (\<forall>ST' LT'. \<tau> = Some(ST',LT') \<longrightarrow>
+ size ST \<le> size ST' \<and> P \<turnstile> Some (drop (size ST' - size ST) ST',LT') \<le>' ty\<^sub>i' ST E A)"
+ using P_\<tau>s\<^sub>1 by fast
+ have "size ?is < size ?\<tau>s" and "pcs ?xt' \<subseteq> {0..<size ?is}"
+ and P_pc: "\<And>pc. pc< size ?is \<Longrightarrow> P,T\<^sub>r,mxs,size ?\<tau>s,xt \<turnstile> ?is!pc,pc :: ?\<tau>s"
+ using wtis by(simp_all add:wt_instrs_def)
+ moreover {
+ fix pc let ?mpc = "size ?\<tau>s" and ?i = "?is!pc" and ?t = "?\<tau>s!pc"
+ assume "pc< size ?is"
+ then have wti: "P,T\<^sub>r,mxs,?mpc,xt \<turnstile> ?i,pc :: ?\<tau>s" by(rule P_pc)
+ then have app: "app ?i P mxs T\<^sub>r pc ?mpc xt ?t"
+ and eff_ss: "\<And>pc' \<tau>'. (pc',\<tau>') \<in> set (eff ?i P pc xt (?\<tau>s!pc))
+ \<Longrightarrow> P \<turnstile> \<tau>' \<le>' ?\<tau>s!pc'"
+ by(fastforce simp add: wt_instr_def)+
+ have "app ?i P mxs T\<^sub>r pc ?mpc ?xt' ?t
+ \<and> (\<forall>(pc',\<tau>') \<in> set (eff ?i P pc ?xt' ?t). P \<turnstile> \<tau>' \<le>' ?\<tau>s!pc')"
+ proof (cases ?t)
+ case (Some \<tau>)
+ obtain ST' LT' where \<tau>: "\<tau> = (ST', LT')" by(cases \<tau>) simp
+ have app\<^sub>i: "app\<^sub>i (?i,P,pc,mxs,T\<^sub>r,\<tau>)" and xcpt_app: "xcpt_app ?i P pc mxs xt \<tau>"
+ and eff_pc: "\<And>pc' \<tau>'. (pc',\<tau>') \<in> set (eff ?i P pc xt ?t) \<Longrightarrow> pc' < ?mpc"
+ using Some app by(fastforce simp add: app_def)+
+ have "xcpt_app ?i P pc mxs ?xt' \<tau>"
+ proof(cases "pc < length \<tau>s\<^sub>1 - 1")
+ case False then show ?thesis using Some \<tau> is_\<tau>s xcpt_app
+ by (clarsimp simp: xcpt_app_def relevant_entries_def
+ is_relevant_entry_def)
+ next
+ case True
+ then have True': "pc < length \<tau>s\<^sub>1" by simp
+ then have "\<tau>s\<^sub>1 ! pc = ?t" by(fastforce simp: nth_append)
+ moreover have \<tau>s\<^sub>1_pc: "\<tau>s\<^sub>1 ! pc \<in> set \<tau>s\<^sub>1" by(rule nth_mem[OF True'])
+ ultimately show ?thesis
+ using Some \<tau> PC ST_mxs xcpt_app P_\<tau>s\<^sub>1'[OF \<tau>s\<^sub>1_pc]
+ by (simp add: xcpt_app_def relevant_entries_def)
+ qed
+ moreover {
+ fix pc' \<tau>' assume efft: "(pc',\<tau>') \<in> set (eff ?i P pc ?xt' ?t)"
+ have "pc' < ?mpc \<and> P \<turnstile> \<tau>' \<le>' ?\<tau>s!pc'" (is "?P1 \<and> ?P2")
+ proof(cases "(pc',\<tau>') \<in> set (eff ?i P pc xt ?t)")
+ case True
+ have ?P1 using True by(rule eff_pc)
+ moreover have ?P2 using True by(rule eff_ss)
+ ultimately show ?thesis by simp
+ next
+ case False
+ then have xte: "(pc',\<tau>') \<in> set (xcpt_eff ?i P pc \<tau> [?xte])"
+ using efft Some by(clarsimp simp: eff_def)
+ then have ?P1 using Some \<tau> is_\<tau>s
+ by(clarsimp simp: xcpt_eff_def relevant_entries_def split: if_split_asm)
+ moreover have ?P2
+ proof(cases "pc < length \<tau>s\<^sub>1 - 1")
+ case False': False
+ then show ?thesis using False Some \<tau> xte is_\<tau>s
+ by(simp add: xcpt_eff_def relevant_entries_def is_relevant_entry_def)
+ next
+ case True
+ then have True': "pc < length \<tau>s\<^sub>1" by simp
+ have \<tau>s\<^sub>1_pc: "\<tau>s\<^sub>1 ! pc \<in> set \<tau>s\<^sub>1" by(rule nth_mem[OF True'])
+ have "P \<turnstile> \<lfloor>(Class C # drop (length ST' - length ST) ST', LT')\<rfloor>
+ \<le>' ty\<^sub>i' (Class C # ST) E A"
+ using True' Some \<tau> P_\<tau>s\<^sub>1'[OF \<tau>s\<^sub>1_pc]
+ by (fastforce simp: nth_append ty\<^sub>i'_def)
+ then show ?thesis using \<tau> xte is_\<tau>s
+ by(simp add: xcpt_eff_def relevant_entries_def split: if_split_asm)
+ qed
+ ultimately show ?thesis by simp
+ qed
+ }
+ ultimately show ?thesis using Some app\<^sub>i by(fastforce simp add: app_def)
+ qed simp
+ then have "P,T\<^sub>r,mxs,size ?\<tau>s,?xt' \<turnstile> ?is!pc,pc :: ?\<tau>s"
+ by(simp add: wt_instr_def)
+ }
+ ultimately show ?thesis by(simp add:wt_instrs_def)
+qed
+
+declare [[simproc add: list_to_set_comprehension]]
+declare nth_append[simp]
+(*>*)
+
+lemma drop_Cons_Suc:
+ "\<And>xs. drop n xs = y#ys \<Longrightarrow> drop (Suc n) xs = ys"
+proof(induct n)
+ case (Suc n) then show ?case by(simp add: drop_Suc)
+qed simp
+
+lemma drop_mess:
+assumes "Suc (length xs\<^sub>0) \<le> length xs"
+ and "drop (length xs - Suc (length xs\<^sub>0)) xs = x # xs\<^sub>0"
+shows "drop (length xs - length xs\<^sub>0) xs = xs\<^sub>0"
+using assms proof(cases xs)
+ case (Cons a list) then show ?thesis using assms
+ proof(cases "length list - length xs\<^sub>0")
+ case Suc then show ?thesis using Cons assms
+ by (simp add: Suc_diff_le drop_Cons_Suc)
+ qed simp
+qed simp
+
+(*<*)
+declare (in TC0)
+ after_def[simp] pair_eq_ty\<^sub>i'_conv[simp]
+(*>*)
+
+lemma (in TC1) compT_ST_prefix:
+ "\<And>E A ST\<^sub>0. \<lfloor>(ST,LT)\<rfloor> \<in> set(compT E A ST\<^sub>0 e) \<Longrightarrow>
+ size ST\<^sub>0 \<le> size ST \<and> drop (size ST - size ST\<^sub>0) ST = ST\<^sub>0"
+and
+ "\<And>E A ST\<^sub>0. \<lfloor>(ST,LT)\<rfloor> \<in> set(compTs E A ST\<^sub>0 es) \<Longrightarrow>
+ size ST\<^sub>0 \<le> size ST \<and> drop (size ST - size ST\<^sub>0) ST = ST\<^sub>0"
+(*<*)
+proof(induct e and es rule: compT.induct compTs.induct)
+ case (FAss e\<^sub>1 F D e\<^sub>2)
+ moreover {
+ let ?ST\<^sub>0 = "ty E e\<^sub>1 # ST\<^sub>0"
+ fix A assume "\<lfloor>(ST, LT)\<rfloor> \<in> set (compT E A ?ST\<^sub>0 e\<^sub>2)"
+ with FAss
+ have "length ?ST\<^sub>0 \<le> length ST \<and> drop (size ST - size ?ST\<^sub>0) ST = ?ST\<^sub>0" by blast
+ hence ?case by (clarsimp simp add: drop_mess)
+ }
+ ultimately show ?case by auto
+next
+ case TryCatch thus ?case by auto
+next
+ case Block thus ?case by auto
+next
+ case Seq thus ?case by auto
+next
+ case While thus ?case by auto
+next
+ case Cond thus ?case by auto
+next
+ case (Call e M es)
+ moreover {
+ let ?ST\<^sub>0 = "ty E e # ST\<^sub>0"
+ fix A assume "\<lfloor>(ST, LT)\<rfloor> \<in> set (compTs E A ?ST\<^sub>0 es)"
+ with Call
+ have "length ?ST\<^sub>0 \<le> length ST \<and> drop (size ST - size ?ST\<^sub>0) ST = ?ST\<^sub>0" by blast
+ hence ?case by (clarsimp simp add: drop_mess)
+ }
+ ultimately show ?case by auto
+next
+ case (Cons_exp e es)
+ moreover {
+ let ?ST\<^sub>0 = "ty E e # ST\<^sub>0"
+ fix A assume "\<lfloor>(ST, LT)\<rfloor> \<in> set (compTs E A ?ST\<^sub>0 es)"
+ with Cons_exp
+ have "length ?ST\<^sub>0 \<le> length ST \<and> drop (size ST - size ?ST\<^sub>0) ST = ?ST\<^sub>0" by blast
+ hence ?case by (clarsimp simp add: drop_mess)
+ }
+ ultimately show ?case by auto
+next
+ case (BinOp e\<^sub>1 bop e\<^sub>2)
+ moreover {
+ let ?ST\<^sub>0 = "ty E e\<^sub>1 # ST\<^sub>0"
+ fix A assume "\<lfloor>(ST, LT)\<rfloor> \<in> set (compT E A ?ST\<^sub>0 e\<^sub>2)"
+ with BinOp
+ have "length ?ST\<^sub>0 \<le> length ST \<and> drop (size ST - size ?ST\<^sub>0) ST = ?ST\<^sub>0" by blast
+ hence ?case by (clarsimp simp add: drop_mess)
+ }
+ ultimately show ?case by auto
+next
+ case new thus ?case by auto
+next
+ case Val thus ?case by auto
+next
+ case Cast thus ?case by auto
+next
+ case Var thus ?case by auto
+next
+ case LAss thus ?case by auto
+next
+ case throw thus ?case by auto
+next
+ case FAcc thus ?case by auto
+next
+ case SFAcc thus ?case by auto
+next
+ case SFAss thus ?case by auto
+next
+ case (SCall C M es) thus ?case by auto
+next
+ case INIT thus ?case by auto
+next
+ case RI thus ?case by auto
+next
+ case Nil_exp thus ?case by auto
+qed
+
+declare (in TC0)
+ after_def[simp del] pair_eq_ty\<^sub>i'_conv[simp del]
+(*>*)
+
+(* FIXME *)
+lemma fun_of_simp [simp]: "fun_of S x y = ((x,y) \<in> S)"
+(*<*) by (simp add: fun_of_def)(*>*)
+
+theorem (in TC2) compT_wt_instrs: "\<And>E T A ST.
+ \<lbrakk> P,E \<turnstile>\<^sub>1 e :: T; \<D> e A; \<B> e (size E);
+ size ST + max_stack e \<le> mxs; size E + max_vars e \<le> mxl \<rbrakk>
+ \<Longrightarrow> \<turnstile> compE\<^sub>2 e, compxE\<^sub>2 e 0 (size ST) [::]
+ ty\<^sub>i' ST E A # compT E A ST e @ [after E A ST e]"
+(*<*)(is "\<And>E T A ST. PROP ?P e E T A ST")(*>*)
+
+and "\<And>E Ts A ST.
+ \<lbrakk> P,E \<turnstile>\<^sub>1 es[::]Ts; \<D>s es A; \<B>s es (size E);
+ size ST + max_stacks es \<le> mxs; size E + max_varss es \<le> mxl \<rbrakk>
+ \<Longrightarrow> let \<tau>s = ty\<^sub>i' ST E A # compTs E A ST es in
+ \<turnstile> compEs\<^sub>2 es,compxEs\<^sub>2 es 0 (size ST) [::] \<tau>s \<and>
+ last \<tau>s = ty\<^sub>i' (rev Ts @ ST) E (A \<squnion> \<A>s es)"
+(*<*)
+(is "\<And>E Ts A ST. PROP ?Ps es E Ts A ST")
+proof(induct e and es rule: compxE\<^sub>2.induct compxEs\<^sub>2.induct)
+ case (TryCatch e\<^sub>1 C i e\<^sub>2)
+ hence [simp]: "i = size E" by simp
+ have wt\<^sub>1: "P,E \<turnstile>\<^sub>1 e\<^sub>1 :: T" and wt\<^sub>2: "P,E@[Class C] \<turnstile>\<^sub>1 e\<^sub>2 :: T"
+ and "class": "is_class P C" using TryCatch by auto
+ let ?A\<^sub>1 = "A \<squnion> \<A> e\<^sub>1" let ?A\<^sub>i = "A \<squnion> \<lfloor>{i}\<rfloor>" let ?E\<^sub>i = "E @ [Class C]"
+ let ?\<tau> = "ty\<^sub>i' ST E A" let ?\<tau>s\<^sub>1 = "compT E A ST e\<^sub>1"
+ let ?\<tau>\<^sub>1 = "ty\<^sub>i' (T#ST) E ?A\<^sub>1" let ?\<tau>\<^sub>2 = "ty\<^sub>i' (Class C#ST) E A"
+ let ?\<tau>\<^sub>3 = "ty\<^sub>i' ST ?E\<^sub>i ?A\<^sub>i" let ?\<tau>s\<^sub>2 = "compT ?E\<^sub>i ?A\<^sub>i ST e\<^sub>2"
+ let ?\<tau>\<^sub>2' = "ty\<^sub>i' (T#ST) ?E\<^sub>i (?A\<^sub>i \<squnion> \<A> e\<^sub>2)"
+ let ?\<tau>' = "ty\<^sub>i' (T#ST) E (A \<squnion> \<A> e\<^sub>1 \<sqinter> (\<A> e\<^sub>2 \<ominus> i))"
+ let ?go = "Goto (int(size(compE\<^sub>2 e\<^sub>2)) + 2)"
+ have "PROP ?P e\<^sub>2 ?E\<^sub>i T ?A\<^sub>i ST" by fact
+ hence "\<turnstile> compE\<^sub>2 e\<^sub>2,compxE\<^sub>2 e\<^sub>2 0 (size ST) [::] (?\<tau>\<^sub>3 # ?\<tau>s\<^sub>2) @ [?\<tau>\<^sub>2']"
+ using TryCatch.prems by(auto simp:after_def)
+ also have "?A\<^sub>i \<squnion> \<A> e\<^sub>2 = (A \<squnion> \<A> e\<^sub>2) \<squnion> \<lfloor>{size E}\<rfloor>"
+ by(fastforce simp:hyperset_defs)
+ also have "P \<turnstile> ty\<^sub>i' (T#ST) ?E\<^sub>i \<dots> \<le>' ty\<^sub>i' (T#ST) E (A \<squnion> \<A> e\<^sub>2)"
+ by(simp add:hyperset_defs ty\<^sub>l_incr ty\<^sub>i'_def)
+ also have "P \<turnstile> \<dots> \<le>' ty\<^sub>i' (T#ST) E (A \<squnion> \<A> e\<^sub>1 \<sqinter> (\<A> e\<^sub>2 \<ominus> i))"
+ by(auto intro!: ty\<^sub>l_antimono simp:hyperset_defs ty\<^sub>i'_def)
+ also have "(?\<tau>\<^sub>3 # ?\<tau>s\<^sub>2) @ [?\<tau>'] = ?\<tau>\<^sub>3 # ?\<tau>s\<^sub>2 @ [?\<tau>']" by simp
+ also have "\<turnstile> [Store i],[] [::] ?\<tau>\<^sub>2 # [] @ [?\<tau>\<^sub>3]"
+ using TryCatch.prems
+ by(auto simp:nth_list_update wt_defs ty\<^sub>i'_def ty\<^sub>l_def
+ list_all2_conv_all_nth hyperset_defs)
+ also have "[] @ (?\<tau>\<^sub>3 # ?\<tau>s\<^sub>2 @ [?\<tau>']) = (?\<tau>\<^sub>3 # ?\<tau>s\<^sub>2 @ [?\<tau>'])" by simp
+ also have "P,T\<^sub>r,mxs,size(compE\<^sub>2 e\<^sub>2)+3,[] \<turnstile> ?go,0 :: ?\<tau>\<^sub>1#?\<tau>\<^sub>2#?\<tau>\<^sub>3#?\<tau>s\<^sub>2 @ [?\<tau>']" using wt\<^sub>2
+ by (auto simp: hyperset_defs ty\<^sub>i'_def wt_defs nth_Cons nat_add_distrib
+ fun_of_def intro: ty\<^sub>l_antimono list_all2_refl split:nat.split)
+ also have "\<turnstile> compE\<^sub>2 e\<^sub>1,compxE\<^sub>2 e\<^sub>1 0 (size ST) [::] ?\<tau> # ?\<tau>s\<^sub>1 @ [?\<tau>\<^sub>1]"
+ using TryCatch by(auto simp:after_def)
+ also have "?\<tau> # ?\<tau>s\<^sub>1 @ ?\<tau>\<^sub>1 # ?\<tau>\<^sub>2 # ?\<tau>\<^sub>3 # ?\<tau>s\<^sub>2 @ [?\<tau>'] =
+ (?\<tau> # ?\<tau>s\<^sub>1 @ [?\<tau>\<^sub>1]) @ ?\<tau>\<^sub>2 # ?\<tau>\<^sub>3 # ?\<tau>s\<^sub>2 @ [?\<tau>']" by simp
+ also have "compE\<^sub>2 e\<^sub>1 @ ?go # [Store i] @ compE\<^sub>2 e\<^sub>2 =
+ (compE\<^sub>2 e\<^sub>1 @ [?go]) @ (Store i # compE\<^sub>2 e\<^sub>2)" by simp
+ also
+ let "?Q \<tau>" = "\<forall>ST' LT'. \<tau> = \<lfloor>(ST', LT')\<rfloor> \<longrightarrow>
+ size ST \<le> size ST' \<and> P \<turnstile> Some (drop (size ST' - size ST) ST',LT') \<le>' ty\<^sub>i' ST E A"
+ {
+ have "?Q (ty\<^sub>i' ST E A)" by (clarsimp simp add: ty\<^sub>i'_def)
+ moreover have "?Q (ty\<^sub>i' (T # ST) E ?A\<^sub>1)"
+ by (fastforce simp add: ty\<^sub>i'_def hyperset_defs intro!: ty\<^sub>l_antimono)
+ moreover have "\<And>\<tau>. \<tau> \<in> set (compT E A ST e\<^sub>1) \<Longrightarrow> ?Q \<tau>" using TryCatch.prems
+ by clarsimp (frule compT_ST_prefix,
+ fastforce dest!: compT_LT_prefix simp add: ty\<^sub>i'_def)
+ ultimately
+ have "\<forall>\<tau>\<in>set (ty\<^sub>i' ST E A # compT E A ST e\<^sub>1 @ [ty\<^sub>i' (T # ST) E ?A\<^sub>1]). ?Q \<tau>"
+ by auto
+ }
+ also from TryCatch.prems max_stack1[OF wt\<^sub>1] have "size ST + 1 \<le> mxs" by auto
+ ultimately show ?case using wt\<^sub>1 wt\<^sub>2 TryCatch.prems "class"
+ by (simp add:after_def)
+next
+ case new thus ?case by(auto simp add:after_def wt_New)
+next
+ case (BinOp e\<^sub>1 bop e\<^sub>2)
+ let ?op = "case bop of Eq \<Rightarrow> [CmpEq] | Add \<Rightarrow> [IAdd]"
+ have T: "P,E \<turnstile>\<^sub>1 e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2 :: T" by fact
+ then obtain T\<^sub>1 T\<^sub>2 where T\<^sub>1: "P,E \<turnstile>\<^sub>1 e\<^sub>1 :: T\<^sub>1" and T\<^sub>2: "P,E \<turnstile>\<^sub>1 e\<^sub>2 :: T\<^sub>2" and
+ bopT: "case bop of Eq \<Rightarrow> (P \<turnstile> T\<^sub>1 \<le> T\<^sub>2 \<or> P \<turnstile> T\<^sub>2 \<le> T\<^sub>1) \<and> T = Boolean
+ | Add \<Rightarrow> T\<^sub>1 = Integer \<and> T\<^sub>2 = Integer \<and> T = Integer" by auto
+ let ?A\<^sub>1 = "A \<squnion> \<A> e\<^sub>1" let ?A\<^sub>2 = "?A\<^sub>1 \<squnion> \<A> e\<^sub>2"
+ let ?\<tau> = "ty\<^sub>i' ST E A" let ?\<tau>s\<^sub>1 = "compT E A ST e\<^sub>1"
+ let ?\<tau>\<^sub>1 = "ty\<^sub>i' (T\<^sub>1#ST) E ?A\<^sub>1" let ?\<tau>s\<^sub>2 = "compT E ?A\<^sub>1 (T\<^sub>1#ST) e\<^sub>2"
+ let ?\<tau>\<^sub>2 = "ty\<^sub>i' (T\<^sub>2#T\<^sub>1#ST) E ?A\<^sub>2" let ?\<tau>' = "ty\<^sub>i' (T#ST) E ?A\<^sub>2"
+ from bopT have "\<turnstile> ?op,[] [::] [?\<tau>\<^sub>2,?\<tau>']"
+ by (cases bop) (auto simp add: wt_CmpEq wt_IAdd)
+ also have "PROP ?P e\<^sub>2 E T\<^sub>2 ?A\<^sub>1 (T\<^sub>1#ST)" by fact
+ with BinOp.prems T\<^sub>2
+ have "\<turnstile> compE\<^sub>2 e\<^sub>2, compxE\<^sub>2 e\<^sub>2 0 (size (T\<^sub>1#ST)) [::] ?\<tau>\<^sub>1#?\<tau>s\<^sub>2@[?\<tau>\<^sub>2]"
+ by (auto simp: after_def)
+ also from BinOp T\<^sub>1 have "\<turnstile> compE\<^sub>2 e\<^sub>1, compxE\<^sub>2 e\<^sub>1 0 (size ST) [::] ?\<tau>#?\<tau>s\<^sub>1@[?\<tau>\<^sub>1]"
+ by (auto simp: after_def)
+ finally show ?case using T T\<^sub>1 T\<^sub>2 by (simp add: after_def hyperUn_assoc)
+next
+ case (Cons_exp e es)
+ have "P,E \<turnstile>\<^sub>1 e # es [::] Ts" by fact
+ then obtain T\<^sub>e Ts' where
+ T\<^sub>e: "P,E \<turnstile>\<^sub>1 e :: T\<^sub>e" and Ts': "P,E \<turnstile>\<^sub>1 es [::] Ts'" and
+ Ts: "Ts = T\<^sub>e#Ts'" by auto
+ let ?A\<^sub>e = "A \<squnion> \<A> e"
+ let ?\<tau> = "ty\<^sub>i' ST E A" let ?\<tau>s\<^sub>e = "compT E A ST e"
+ let ?\<tau>\<^sub>e = "ty\<^sub>i' (T\<^sub>e#ST) E ?A\<^sub>e" let ?\<tau>s' = "compTs E ?A\<^sub>e (T\<^sub>e#ST) es"
+ let ?\<tau>s = "?\<tau> # ?\<tau>s\<^sub>e @ (?\<tau>\<^sub>e # ?\<tau>s')"
+ have Ps: "PROP ?Ps es E Ts' ?A\<^sub>e (T\<^sub>e#ST)" by fact
+ with Cons_exp.prems T\<^sub>e Ts'
+ have "\<turnstile> compEs\<^sub>2 es, compxEs\<^sub>2 es 0 (size (T\<^sub>e#ST)) [::] ?\<tau>\<^sub>e#?\<tau>s'" by (simp add: after_def)
+ also from Cons_exp T\<^sub>e have "\<turnstile> compE\<^sub>2 e, compxE\<^sub>2 e 0 (size ST) [::] ?\<tau>#?\<tau>s\<^sub>e@[?\<tau>\<^sub>e]"
+ by (auto simp: after_def)
+ moreover
+ from Ps Cons_exp.prems T\<^sub>e Ts' Ts
+ have "last ?\<tau>s = ty\<^sub>i' (rev Ts@ST) E (?A\<^sub>e \<squnion> \<A>s es)" by simp
+ ultimately show ?case using T\<^sub>e by (simp add: after_def hyperUn_assoc)
+next
+ case (FAss e\<^sub>1 F D e\<^sub>2)
+ hence Void: "P,E \<turnstile>\<^sub>1 e\<^sub>1\<bullet>F{D} := e\<^sub>2 :: Void" by auto
+ then obtain C T T' where
+ C: "P,E \<turnstile>\<^sub>1 e\<^sub>1 :: Class C" and sees: "P \<turnstile> C sees F,NonStatic:T in D" and
+ T': "P,E \<turnstile>\<^sub>1 e\<^sub>2 :: T'" and T'_T: "P \<turnstile> T' \<le> T" by auto
+ let ?A\<^sub>1 = "A \<squnion> \<A> e\<^sub>1" let ?A\<^sub>2 = "?A\<^sub>1 \<squnion> \<A> e\<^sub>2"
+ let ?\<tau> = "ty\<^sub>i' ST E A" let ?\<tau>s\<^sub>1 = "compT E A ST e\<^sub>1"
+ let ?\<tau>\<^sub>1 = "ty\<^sub>i' (Class C#ST) E ?A\<^sub>1" let ?\<tau>s\<^sub>2 = "compT E ?A\<^sub>1 (Class C#ST) e\<^sub>2"
+ let ?\<tau>\<^sub>2 = "ty\<^sub>i' (T'#Class C#ST) E ?A\<^sub>2" let ?\<tau>\<^sub>3 = "ty\<^sub>i' ST E ?A\<^sub>2"
+ let ?\<tau>' = "ty\<^sub>i' (Void#ST) E ?A\<^sub>2"
+ from FAss.prems sees T'_T
+ have "\<turnstile> [Putfield F D,Push Unit],[] [::] [?\<tau>\<^sub>2,?\<tau>\<^sub>3,?\<tau>']"
+ by (fastforce simp add: wt_Push wt_Put)
+ also have "PROP ?P e\<^sub>2 E T' ?A\<^sub>1 (Class C#ST)" by fact
+ with FAss.prems T'
+ have "\<turnstile> compE\<^sub>2 e\<^sub>2, compxE\<^sub>2 e\<^sub>2 0 (size ST+1) [::] ?\<tau>\<^sub>1#?\<tau>s\<^sub>2@[?\<tau>\<^sub>2]"
+ by (auto simp add: after_def hyperUn_assoc)
+ also from FAss C have "\<turnstile> compE\<^sub>2 e\<^sub>1, compxE\<^sub>2 e\<^sub>1 0 (size ST) [::] ?\<tau>#?\<tau>s\<^sub>1@[?\<tau>\<^sub>1]"
+ by (auto simp add: after_def)
+ finally show ?case using Void C T' by (simp add: after_def hyperUn_assoc)
+next
+ case (SFAss C F D e\<^sub>2)
+ hence Void: "P,E \<turnstile>\<^sub>1 C\<bullet>\<^sub>sF{D} := e\<^sub>2 :: Void" by auto
+ then obtain T T' where
+ sees: "P \<turnstile> C sees F,Static:T in D" and
+ T': "P,E \<turnstile>\<^sub>1 e\<^sub>2 :: T'" and T'_T: "P \<turnstile> T' \<le> T" by auto
+ let ?A\<^sub>2 = "A \<squnion> \<A> e\<^sub>2"
+ let ?\<tau> = "ty\<^sub>i' ST E A" let ?\<tau>s\<^sub>2 = "compT E A ST e\<^sub>2"
+ let ?\<tau>\<^sub>2 = "ty\<^sub>i' (T'#ST) E ?A\<^sub>2" let ?\<tau>\<^sub>3 = "ty\<^sub>i' ST E ?A\<^sub>2"
+ let ?\<tau>' = "ty\<^sub>i' (Void#ST) E ?A\<^sub>2"
+ from SFAss.prems sees T'_T max_stack1[OF T']
+ have "\<turnstile> [Putstatic C F D,Push Unit],[] [::] [?\<tau>\<^sub>2,?\<tau>\<^sub>3,?\<tau>']"
+ by (fastforce simp add: wt_Push wt_PutS)
+ also have "PROP ?P e\<^sub>2 E T' A ST" by fact
+ with SFAss.prems T'
+ have "\<turnstile> compE\<^sub>2 e\<^sub>2, compxE\<^sub>2 e\<^sub>2 0 (size ST) [::] ?\<tau>#?\<tau>s\<^sub>2@[?\<tau>\<^sub>2]"
+ by (auto simp add: after_def hyperUn_assoc)
+ finally show ?case using Void T' by (simp add: after_def hyperUn_assoc)
+next
+ case Val thus ?case by(auto simp:after_def wt_Push)
+next
+ case Cast thus ?case by (auto simp:after_def wt_Cast)
+next
+ case (Block i T\<^sub>i e)
+ let ?\<tau>s = "ty\<^sub>i' ST E A # compT (E @ [T\<^sub>i]) (A\<ominus>i) ST e"
+ have IH: "PROP ?P e (E@[T\<^sub>i]) T (A\<ominus>i) ST" by fact
+ hence "\<turnstile> compE\<^sub>2 e, compxE\<^sub>2 e 0 (size ST) [::]
+ ?\<tau>s @ [ty\<^sub>i' (T#ST) (E@[T\<^sub>i]) (A\<ominus>(size E) \<squnion> \<A> e)]"
+ using Block.prems by (auto simp add: after_def)
+ also have "P \<turnstile> ty\<^sub>i' (T # ST) (E@[T\<^sub>i]) (A \<ominus> size E \<squnion> \<A> e) \<le>'
+ ty\<^sub>i' (T # ST) (E@[T\<^sub>i]) ((A \<squnion> \<A> e) \<ominus> size E)"
+ by(auto simp add:hyperset_defs intro: ty\<^sub>i'_antimono)
+ also have "\<dots> = ty\<^sub>i' (T # ST) E (A \<squnion> \<A> e)" by simp
+ also have "P \<turnstile> \<dots> \<le>' ty\<^sub>i' (T # ST) E (A \<squnion> (\<A> e \<ominus> i))"
+ by(auto simp add:hyperset_defs intro: ty\<^sub>i'_antimono)
+ finally show ?case using Block.prems by(simp add: after_def)
+next
+ case Var thus ?case by(auto simp:after_def wt_Load)
+next
+ case FAcc thus ?case by(auto simp:after_def wt_Get)
+next
+ case SFAcc thus ?case by(auto simp: after_def wt_GetS)
+next
+ case (LAss i e)
+ then obtain T' where wt: "P,E \<turnstile>\<^sub>1 e :: T'" by auto
+ show ?case using max_stack1[OF wt] LAss
+ by(auto simp: hyper_insert_comm after_def wt_Store wt_Push)
+next
+ case Nil_exp thus ?case by auto
+next
+ case throw thus ?case by(auto simp add: after_def wt_Throw)
+next
+ case (While e c)
+ obtain Tc where wte: "P,E \<turnstile>\<^sub>1 e :: Boolean" and wtc: "P,E \<turnstile>\<^sub>1 c :: Tc"
+ and [simp]: "T = Void" using While by auto
+ have [simp]: "ty E (while (e) c) = Void" using While by simp
+ let ?A\<^sub>0 = "A \<squnion> \<A> e" let ?A\<^sub>1 = "?A\<^sub>0 \<squnion> \<A> c"
+ let ?\<tau> = "ty\<^sub>i' ST E A" let ?\<tau>s\<^sub>e = "compT E A ST e"
+ let ?\<tau>\<^sub>e = "ty\<^sub>i' (Boolean#ST) E ?A\<^sub>0" let ?\<tau>\<^sub>1 = "ty\<^sub>i' ST E ?A\<^sub>0"
+ let ?\<tau>s\<^sub>c = "compT E ?A\<^sub>0 ST c" let ?\<tau>\<^sub>c = "ty\<^sub>i' (Tc#ST) E ?A\<^sub>1"
+ let ?\<tau>\<^sub>2 = "ty\<^sub>i' ST E ?A\<^sub>1" let ?\<tau>' = "ty\<^sub>i' (Void#ST) E ?A\<^sub>0"
+ let ?\<tau>s = "(?\<tau> # ?\<tau>s\<^sub>e @ [?\<tau>\<^sub>e]) @ ?\<tau>\<^sub>1 # ?\<tau>s\<^sub>c @ [?\<tau>\<^sub>c, ?\<tau>\<^sub>2, ?\<tau>\<^sub>1, ?\<tau>']"
+ have "\<turnstile> [],[] [::] [] @ ?\<tau>s" by(simp add:wt_instrs_def)
+ also
+ have "PROP ?P e E Boolean A ST" by fact
+ hence "\<turnstile> compE\<^sub>2 e,compxE\<^sub>2 e 0 (size ST) [::] ?\<tau> # ?\<tau>s\<^sub>e @ [?\<tau>\<^sub>e]"
+ using While.prems by (auto simp:after_def)
+ also
+ have "[] @ ?\<tau>s = (?\<tau> # ?\<tau>s\<^sub>e) @ ?\<tau>\<^sub>e # ?\<tau>\<^sub>1 # ?\<tau>s\<^sub>c @ [?\<tau>\<^sub>c,?\<tau>\<^sub>2,?\<tau>\<^sub>1,?\<tau>']" by simp
+ also
+ let ?n\<^sub>e = "size(compE\<^sub>2 e)" let ?n\<^sub>c = "size(compE\<^sub>2 c)"
+ thm compT_sizes(1)
+ let ?if = "IfFalse (int ?n\<^sub>c + 3)"
+ from wtc wte have "compE\<^sub>2 c \<noteq> []" and "compE\<^sub>2 e \<noteq> []" using WT\<^sub>1_nsub_RI by auto
+ then have compe2c: "length (compE\<^sub>2 c) > 0" and compe2e: "length (compE\<^sub>2 e) > 0" by auto
+ have Suc_pred'_auxi: "\<And>n. 0 < n \<Longrightarrow> n = Suc (n - Suc 0)" using Suc_pred'[OF compe2c] by auto
+ have "\<turnstile> [?if],[] [::] ?\<tau>\<^sub>e # ?\<tau>\<^sub>1 # ?\<tau>s\<^sub>c @ [?\<tau>\<^sub>c, ?\<tau>\<^sub>2, ?\<tau>\<^sub>1, ?\<tau>']"
+ proof-{
+ show ?thesis
+ apply (rule wt_IfFalse)
+ apply simp
+ apply simp
+ apply (subgoal_tac "length (compE\<^sub>2 c) = length (compT E (A \<squnion> \<A> e) ST c) + 1"
+ "nat (int (length (compE\<^sub>2 c)) + 3 - 2) > length (compT E (A \<squnion> \<A> e) ST c)")
+ apply (insert Suc_pred'_auxi[OF compe2c])
+ apply (simp add:compT_sizes(1)[OF wtc] )+
+ done
+ }
+ qed
+ also
+ have "(?\<tau> # ?\<tau>s\<^sub>e) @ (?\<tau>\<^sub>e # ?\<tau>\<^sub>1 # ?\<tau>s\<^sub>c @ [?\<tau>\<^sub>c, ?\<tau>\<^sub>2, ?\<tau>\<^sub>1, ?\<tau>']) = ?\<tau>s" by simp
+ also
+ have "PROP ?P c E Tc ?A\<^sub>0 ST" by fact
+ hence "\<turnstile> compE\<^sub>2 c,compxE\<^sub>2 c 0 (size ST) [::] ?\<tau>\<^sub>1 # ?\<tau>s\<^sub>c @ [?\<tau>\<^sub>c]"
+ using While.prems wtc by (auto simp:after_def)
+ also have "?\<tau>s = (?\<tau> # ?\<tau>s\<^sub>e @ [?\<tau>\<^sub>e,?\<tau>\<^sub>1] @ ?\<tau>s\<^sub>c) @ [?\<tau>\<^sub>c,?\<tau>\<^sub>2,?\<tau>\<^sub>1,?\<tau>']" by simp
+ also have "\<turnstile> [Pop],[] [::] [?\<tau>\<^sub>c, ?\<tau>\<^sub>2]" by(simp add:wt_Pop)
+ also have "(?\<tau> # ?\<tau>s\<^sub>e @ [?\<tau>\<^sub>e,?\<tau>\<^sub>1] @ ?\<tau>s\<^sub>c) @ [?\<tau>\<^sub>c,?\<tau>\<^sub>2,?\<tau>\<^sub>1,?\<tau>'] = ?\<tau>s" by simp
+ also let ?go = "Goto (-int(?n\<^sub>c+?n\<^sub>e+2))"
+ have "P \<turnstile> ?\<tau>\<^sub>2 \<le>' ?\<tau>" by(fastforce intro: ty\<^sub>i'_antimono simp: hyperset_defs)
+ hence "P,T\<^sub>r,mxs,size ?\<tau>s,[] \<turnstile> ?go,?n\<^sub>e+?n\<^sub>c+2 :: ?\<tau>s" using wte wtc
+ proof-{
+ let ?t1 = "ty\<^sub>i' ST E A" let ?t2 = "ty\<^sub>i' (Boolean # ST) E (A \<squnion> \<A> e)" let ?t3 = "ty\<^sub>i' ST E (A \<squnion> \<A> e)"
+ let ?t4 = "[ty\<^sub>i' (Tc # ST) E (A \<squnion> \<A> e \<squnion> \<A> c), ty\<^sub>i' ST E (A \<squnion> \<A> e \<squnion> \<A> c), ty\<^sub>i' ST E (A \<squnion> \<A> e), ty\<^sub>i' (Void # ST) E (A \<squnion> \<A> e)]"
+ let ?c = " compT E (A \<squnion> \<A> e) ST c" let ?e = "compT E A ST e"
+
+ assume ass: "P \<turnstile> ty\<^sub>i' ST E (A \<squnion> \<A> e \<squnion> \<A> c) \<le>' ?t1"
+ show ?thesis
+ apply (rule wt_Goto)
+ apply simp
+ apply simp
+ apply simp
+ proof-{
+ let ?s1 = "((?t1 # ?e @ [?t2]) @ ?t3 # ?c)"
+ have len1: "length ?c = length (compE\<^sub>2 c) - 1" using compT_sizes(1)[OF wtc] by simp
+ have len2: "length ?e = length (compE\<^sub>2 e) - 1" using compT_sizes(1)[OF wte] by simp
+ hence "length (compE\<^sub>2 e) + length (compE\<^sub>2 c) + 2 > length ?s1"
+ using len1 compe2c compe2e by auto
+ hence "(?s1 @ ?t4) ! (length (compE\<^sub>2 e) + length (compE\<^sub>2 c) + 2) =
+ ?t4 ! (length (compE\<^sub>2 e) + length (compE\<^sub>2 c) + 2 - length ?s1)"
+ by (auto simp only:nth_append split: if_splits)
+ hence "(?s1 @ ?t4) ! (length (compE\<^sub>2 e) + length (compE\<^sub>2 c) + 2) = ty\<^sub>i' ST E (A \<squnion> \<A> e \<squnion> \<A> c)"
+ using len1 len2 compe2c compe2e by auto
+ hence "P \<turnstile> (?s1 @ ?t4) ! (length (compE\<^sub>2 e) + length (compE\<^sub>2 c) + 2) \<le>' ty\<^sub>i' ST E A"
+ using ass by simp
+ thus "P \<turnstile> ((?t1 # ?e @ [?t2]) @ ?t3 # ?c @ ?t4) ! (length (compE\<^sub>2 e) + length (compE\<^sub>2 c) + 2) \<le>'
+ ((?t1 # ?e @ [?t2]) @ ?t3 # ?c @ ?t4) ! nat (int (length (compE\<^sub>2 e) + length (compE\<^sub>2 c) + 2) +
+ - int (length (compE\<^sub>2 c) + length (compE\<^sub>2 e) + 2))"
+ by simp
+ }qed
+ }qed
+ also have "?\<tau>s = (?\<tau> # ?\<tau>s\<^sub>e @ [?\<tau>\<^sub>e,?\<tau>\<^sub>1] @ ?\<tau>s\<^sub>c @ [?\<tau>\<^sub>c, ?\<tau>\<^sub>2]) @ [?\<tau>\<^sub>1, ?\<tau>']"
+ by simp
+ also have "\<turnstile> [Push Unit],[] [::] [?\<tau>\<^sub>1,?\<tau>']"
+ using While.prems max_stack1[OF wtc] by(auto simp add:wt_Push)
+ finally show ?case using wtc wte
+ by (simp add:after_def)
+next
+ case (Cond e e\<^sub>1 e\<^sub>2)
+ obtain T\<^sub>1 T\<^sub>2 where wte: "P,E \<turnstile>\<^sub>1 e :: Boolean"
+ and wt\<^sub>1: "P,E \<turnstile>\<^sub>1 e\<^sub>1 :: T\<^sub>1" and wt\<^sub>2: "P,E \<turnstile>\<^sub>1 e\<^sub>2 :: T\<^sub>2"
+ and sub\<^sub>1: "P \<turnstile> T\<^sub>1 \<le> T" and sub\<^sub>2: "P \<turnstile> T\<^sub>2 \<le> T"
+ using Cond by auto
+ from wte wt\<^sub>1 wt\<^sub>2 have "compE\<^sub>2 e\<^sub>1 \<noteq> []" and "compE\<^sub>2 e\<^sub>2 \<noteq> []" and "compE\<^sub>2 e \<noteq> []" using WT\<^sub>1_nsub_RI by auto
+ then have compe: "length (compE\<^sub>2 e) > 0"
+ and compe1: "length (compE\<^sub>2 e\<^sub>1) > 0"
+ and compe2: "length (compE\<^sub>2 e\<^sub>2) > 0" by auto
+
+
+ have [simp]: "ty E (if (e) e\<^sub>1 else e\<^sub>2) = T" using Cond by simp
+ let ?A\<^sub>0 = "A \<squnion> \<A> e" let ?A\<^sub>2 = "?A\<^sub>0 \<squnion> \<A> e\<^sub>2" let ?A\<^sub>1 = "?A\<^sub>0 \<squnion> \<A> e\<^sub>1"
+ let ?A' = "?A\<^sub>0 \<squnion> \<A> e\<^sub>1 \<sqinter> \<A> e\<^sub>2"
+ let ?\<tau>\<^sub>2 = "ty\<^sub>i' ST E ?A\<^sub>0" let ?\<tau>' = "ty\<^sub>i' (T#ST) E ?A'"
+ let ?\<tau>s\<^sub>2 = "compT E ?A\<^sub>0 ST e\<^sub>2"
+ have "PROP ?P e\<^sub>2 E T\<^sub>2 ?A\<^sub>0 ST" by fact
+ hence "\<turnstile> compE\<^sub>2 e\<^sub>2, compxE\<^sub>2 e\<^sub>2 0 (size ST) [::] (?\<tau>\<^sub>2#?\<tau>s\<^sub>2) @ [ty\<^sub>i' (T\<^sub>2#ST) E ?A\<^sub>2]"
+ using Cond.prems wt\<^sub>2 by(auto simp add:after_def)
+ also have "P \<turnstile> ty\<^sub>i' (T\<^sub>2#ST) E ?A\<^sub>2 \<le>' ?\<tau>'" using sub\<^sub>2
+ by(auto simp add: hyperset_defs ty\<^sub>i'_def intro!: ty\<^sub>l_antimono)
+ also
+ let ?\<tau>\<^sub>3 = "ty\<^sub>i' (T\<^sub>1 # ST) E ?A\<^sub>1"
+ let ?g\<^sub>2 = "Goto(int (size (compE\<^sub>2 e\<^sub>2) + 1))"
+ from sub\<^sub>1 have "P,T\<^sub>r,mxs,size(compE\<^sub>2 e\<^sub>2)+2,[] \<turnstile> ?g\<^sub>2,0 :: ?\<tau>\<^sub>3#(?\<tau>\<^sub>2#?\<tau>s\<^sub>2)@[?\<tau>']" using wt\<^sub>2
+ by(auto simp: hyperset_defs wt_defs nth_Cons ty\<^sub>i'_def
+ split:nat.split intro!: ty\<^sub>l_antimono)
+ also
+ let ?\<tau>s\<^sub>1 = "compT E ?A\<^sub>0 ST e\<^sub>1"
+ have "PROP ?P e\<^sub>1 E T\<^sub>1 ?A\<^sub>0 ST" by fact
+ hence "\<turnstile> compE\<^sub>2 e\<^sub>1,compxE\<^sub>2 e\<^sub>1 0 (size ST) [::] ?\<tau>\<^sub>2 # ?\<tau>s\<^sub>1 @ [?\<tau>\<^sub>3]"
+ using Cond.prems wt\<^sub>1 by(auto simp add:after_def)
+ also
+ let ?\<tau>s\<^sub>1\<^sub>2 = "?\<tau>\<^sub>2 # ?\<tau>s\<^sub>1 @ ?\<tau>\<^sub>3 # (?\<tau>\<^sub>2 # ?\<tau>s\<^sub>2) @ [?\<tau>']"
+ let ?\<tau>\<^sub>1 = "ty\<^sub>i' (Boolean#ST) E ?A\<^sub>0"
+ let ?g\<^sub>1 = "IfFalse(int (size (compE\<^sub>2 e\<^sub>1) + 2))"
+ let ?code = "compE\<^sub>2 e\<^sub>1 @ ?g\<^sub>2 # compE\<^sub>2 e\<^sub>2"
+
+ have len1: "length ?\<tau>s\<^sub>1 = length (compE\<^sub>2 e\<^sub>1) - 1" using compT_sizes(1)[OF wt\<^sub>1] by simp
+ have len2: "length ?\<tau>s\<^sub>2 = length (compE\<^sub>2 e\<^sub>2) - 1" using compT_sizes(1)[OF wt\<^sub>2] by simp
+ have len_auxi: "length (compE\<^sub>2 e\<^sub>1) - (length (compE\<^sub>2 e\<^sub>1) - Suc 0) = Suc 0" using compe1
+ by (simp add:diff_Suc split:nat.split)
+ have "\<turnstile> [?g\<^sub>1],[] [::] [?\<tau>\<^sub>1] @ ?\<tau>s\<^sub>1\<^sub>2"
+ proof-{
+ show ?thesis
+ apply clarsimp
+ apply (rule wt_IfFalse)
+ apply (simp only:nat_add_distrib)
+ apply simp
+ apply (auto simp only:nth_append split:if_splits)
+ apply (simp add:len1)
+ apply (simp only: len1 compe1)
+ apply (simp add:len1 len2 compe1 compe2 )
+ apply (insert len_auxi)
+ apply simp
+ done
+ }qed
+ also (wt_instrs_ext2) have "[?\<tau>\<^sub>1] @ ?\<tau>s\<^sub>1\<^sub>2 = ?\<tau>\<^sub>1 # ?\<tau>s\<^sub>1\<^sub>2" by simp also
+ let ?\<tau> = "ty\<^sub>i' ST E A"
+ have "PROP ?P e E Boolean A ST" by fact
+ hence "\<turnstile> compE\<^sub>2 e, compxE\<^sub>2 e 0 (size ST) [::] ?\<tau> # compT E A ST e @ [?\<tau>\<^sub>1]"
+ using Cond.prems wte by(auto simp add:after_def)
+ finally show ?case using wte wt\<^sub>1 wt\<^sub>2 by(simp add:after_def hyperUn_assoc)
+next
+ case (Call e M es)
+ obtain C D Ts m Ts' where C: "P,E \<turnstile>\<^sub>1 e :: Class C"
+ and "method": "P \<turnstile> C sees M,NonStatic:Ts \<rightarrow> T = m in D"
+ and wtes: "P,E \<turnstile>\<^sub>1 es [::] Ts'" and subs: "P \<turnstile> Ts' [\<le>] Ts"
+ using Call.prems by auto
+ from wtes have same_size: "size es = size Ts'" by(rule WTs\<^sub>1_same_size)
+ let ?A\<^sub>0 = "A \<squnion> \<A> e" let ?A\<^sub>1 = "?A\<^sub>0 \<squnion> \<A>s es"
+ let ?\<tau> = "ty\<^sub>i' ST E A" let ?\<tau>s\<^sub>e = "compT E A ST e"
+ let ?\<tau>\<^sub>e = "ty\<^sub>i' (Class C # ST) E ?A\<^sub>0"
+ let ?\<tau>s\<^sub>e\<^sub>s = "compTs E ?A\<^sub>0 (Class C # ST) es"
+ let ?\<tau>\<^sub>1 = "ty\<^sub>i' (rev Ts' @ Class C # ST) E ?A\<^sub>1"
+ let ?\<tau>' = "ty\<^sub>i' (T # ST) E ?A\<^sub>1"
+ have "\<turnstile> [Invoke M (size es)],[] [::] [?\<tau>\<^sub>1,?\<tau>']"
+ by(rule wt_Invoke[OF same_size "method" subs])
+ also
+ have "PROP ?Ps es E Ts' ?A\<^sub>0 (Class C # ST)" by fact
+ hence "\<turnstile> compEs\<^sub>2 es,compxEs\<^sub>2 es 0 (size ST+1) [::] ?\<tau>\<^sub>e # ?\<tau>s\<^sub>e\<^sub>s"
+ "last (?\<tau>\<^sub>e # ?\<tau>s\<^sub>e\<^sub>s) = ?\<tau>\<^sub>1"
+ using Call.prems wtes by(auto simp add:after_def)
+ also have "(?\<tau>\<^sub>e # ?\<tau>s\<^sub>e\<^sub>s) @ [?\<tau>'] = ?\<tau>\<^sub>e # ?\<tau>s\<^sub>e\<^sub>s @ [?\<tau>']" by simp
+ also have "\<turnstile> compE\<^sub>2 e,compxE\<^sub>2 e 0 (size ST) [::] ?\<tau> # ?\<tau>s\<^sub>e @ [?\<tau>\<^sub>e]"
+ using Call C by(auto simp add:after_def)
+ finally show ?case using Call.prems C wtes by(simp add:after_def hyperUn_assoc)
+next
+ case (SCall C M es)
+ obtain D Ts m Ts' where "method": "P \<turnstile> C sees M,Static:Ts \<rightarrow> T = m in D"
+ and wtes: "P,E \<turnstile>\<^sub>1 es [::] Ts'" and subs: "P \<turnstile> Ts' [\<le>] Ts"
+ using SCall.prems by auto
+ from SCall.prems(1) have nclinit: "M \<noteq> clinit" by auto
+ from wtes have same_size: "size es = size Ts'" by(rule WTs\<^sub>1_same_size)
+ have mxs: "length ST < mxs" using WT\<^sub>1_nsub_RI[OF SCall.prems(1)] SCall.prems(4) by simp
+ let ?A\<^sub>1 = "A \<squnion> \<A>s es"
+ let ?\<tau> = "ty\<^sub>i' ST E A"
+ let ?\<tau>s\<^sub>e\<^sub>s = "compTs E A ST es"
+ let ?\<tau>\<^sub>1 = "ty\<^sub>i' (rev Ts' @ ST) E ?A\<^sub>1"
+ let ?\<tau>' = "ty\<^sub>i' (T # ST) E ?A\<^sub>1"
+ have "\<turnstile> [Invokestatic C M (size es)],[] [::] [?\<tau>\<^sub>1,?\<tau>']"
+ by(rule wt_Invokestatic[OF mxs same_size nclinit "method" subs])
+ also
+ have "PROP ?Ps es E Ts' A ST" by fact
+ hence "\<turnstile> compEs\<^sub>2 es,compxEs\<^sub>2 es 0 (size ST) [::] ?\<tau> # ?\<tau>s\<^sub>e\<^sub>s"
+ "last (?\<tau> # ?\<tau>s\<^sub>e\<^sub>s) = ?\<tau>\<^sub>1"
+ using SCall.prems wtes by(auto simp add:after_def)
+ also have "(?\<tau> # ?\<tau>s\<^sub>e\<^sub>s) @ [?\<tau>'] = ?\<tau> # ?\<tau>s\<^sub>e\<^sub>s @ [?\<tau>']" by simp
+ finally show ?case using SCall.prems wtes by(simp add:after_def hyperUn_assoc)
+next
+ case Seq thus ?case
+ by(auto simp:after_def)
+ (fastforce simp:wt_Push wt_Pop hyperUn_assoc
+ intro:wt_instrs_app2 wt_instrs_Cons)
+next
+ case (INIT C Cs b e)
+ have "P,E \<turnstile>\<^sub>1 INIT C (Cs,b) \<leftarrow> e :: T" by fact
+ thus ?case using WT\<^sub>1_nsub_RI by simp
+next
+ case (RI C e' Cs e)
+ have "P,E \<turnstile>\<^sub>1 RI (C,e') ; Cs \<leftarrow> e :: T" by fact
+ thus ?case using WT\<^sub>1_nsub_RI by simp
+qed
+(*>*)
+
+
+lemma [simp]: "types (compP f P) = types P"
+(*<*)by auto(*>*)
+
+lemma [simp]: "states (compP f P) mxs mxl = states P mxs mxl"
+(*<*)by (simp add: JVM_states_unfold)(*>*)
+
+lemma [simp]: "app\<^sub>i (i, compP f P, pc, mpc, T, \<tau>) = app\<^sub>i (i, P, pc, mpc, T, \<tau>)"
+(*<*)(is "?A = ?B")
+proof -
+ obtain ST LT where \<tau>: "\<tau> = (ST, LT)" by(cases \<tau>) simp
+ then show ?thesis proof(cases i)
+ case Invoke show ?thesis
+ proof(rule iffI)
+ assume ?A then show ?B using Invoke \<tau>
+ by auto (fastforce dest!: sees_method_compPD)
+ next
+ assume ?B then show ?A using Invoke \<tau>
+ by auto (force dest: sees_method_compP)
+ qed
+ next
+ case (Invokestatic x111 x112 x113) show ?thesis
+ proof(rule iffI)
+ assume ?A then show ?B using Invokestatic \<tau>
+ by auto (fastforce dest!: sees_method_compPD)
+ next
+ assume ?B then show ?A using Invokestatic \<tau>
+ by auto (force dest: sees_method_compP)
+ qed
+ qed auto
+qed
+(*>*)
+
+lemma [simp]: "is_relevant_entry (compP f P) i = is_relevant_entry P i"
+(*<*)
+proof -
+ { fix pc e
+ have "is_relevant_entry (compP f P) i pc e = is_relevant_entry P i pc e"
+ by(cases i) (auto simp: is_relevant_entry_def)
+ }
+ then show ?thesis by fast
+qed
+(*>*)
+
+lemma [simp]: "relevant_entries (compP f P) i pc xt = relevant_entries P i pc xt"
+(*<*) by (simp add: relevant_entries_def)(*>*)
+
+lemma [simp]: "app i (compP f P) mpc T pc mxl xt \<tau> = app i P mpc T pc mxl xt \<tau>"
+(*<*)
+by (simp add: app_def xcpt_app_def eff_def xcpt_eff_def norm_eff_def)
+ (fastforce simp add: image_def)
+(*>*)
+
+lemma [simp]:
+assumes "app i P mpc T pc mxl xt \<tau>"
+shows "eff i (compP f P) pc xt \<tau> = eff i P pc xt \<tau>"
+(*<*)
+using assms
+proof(clarsimp simp: eff_def norm_eff_def xcpt_eff_def app_def, cases i)
+qed auto
+(*>*)
+
+lemma [simp]: "subtype (compP f P) = subtype P"
+(*<*)by (rule ext)+ simp(*>*)
+
+lemma [simp]: "compP f P \<turnstile> \<tau> \<le>' \<tau>' = P \<turnstile> \<tau> \<le>' \<tau>'"
+(*<*) by (simp add: sup_state_opt_def sup_state_def sup_ty_opt_def)(*>*)
+
+lemma [simp]: "compP f P,T,mpc,mxl,xt \<turnstile> i,pc :: \<tau>s = P,T,mpc,mxl,xt \<turnstile> i,pc :: \<tau>s"
+(*<*)by (simp add: wt_instr_def cong: conj_cong)(*>*)
+
+declare TC1.compT_sizes[simp] TC0.ty_def2[simp]
+
+context TC2
+begin
+
+lemma compT_method_NonStatic:
+ fixes e and A and C and Ts and mxl\<^sub>0
+ defines [simp]: "E \<equiv> Class C # Ts"
+ and [simp]: "A \<equiv> \<lfloor>{..size Ts}\<rfloor>"
+ and [simp]: "A' \<equiv> A \<squnion> \<A> e"
+ and [simp]: "mxl\<^sub>0 \<equiv> max_vars e"
+ assumes mxs: "max_stack e = mxs"
+ and mxl: "Suc (length Ts + max_vars e) = mxl"
+ assumes wf: "wf_prog p P" and wte: "P,E \<turnstile>\<^sub>1 e :: T" and \<D>: "\<D> e A"
+ and \<B>: "\<B> e (size E)" and E_P: "set E \<subseteq> types P" and wid: "P \<turnstile> T \<le> T\<^sub>r"
+ shows "wt_method (compP\<^sub>2 P) C NonStatic Ts T\<^sub>r mxs mxl\<^sub>0 (compE\<^sub>2 e @ [Return])
+ (compxE\<^sub>2 e 0 0) (ty\<^sub>i' [] E A # compT\<^sub>a E A [] e)"
+(*<*)(is "wt_method ?P C ?b Ts T\<^sub>r mxs mxl\<^sub>0 ?is ?xt ?\<tau>s")
+proof -
+ let ?n = "length E + mxl\<^sub>0"
+ have wt_compE: "P,T\<^sub>r,mxs \<turnstile> compE\<^sub>2 e, compxE\<^sub>2 e 0 (length []) [::]
+ TC0.ty\<^sub>i' ?n [] E A # TC1.compT P ?n E A [] e @[TC0.after P ?n E A [] e]"
+ using mxs TC2.compT_wt_instrs [OF wte \<D> \<B>, of "[]" mxs ?n T\<^sub>r] by simp
+
+ have "OK (ty\<^sub>i' [T] E A') \<in> states P mxs mxl"
+ using mxs WT\<^sub>1_is_type[OF wf wte E_P] max_stack1[OF wte] OK_ty\<^sub>i'_in_statesI[OF E_P]
+ by simp
+ moreover have "OK ` set (compT E A [] e) \<subseteq> states P mxs mxl"
+ using mxs mxl wid compT_states(1)[OF wf wte E_P] by simp
+ ultimately have "check_types ?P mxs ?n (map OK ?\<tau>s)"
+ using mxl wte E_P by (simp add: compT\<^sub>a_def after_def check_types_def)
+ moreover have "wt_start ?P C ?b Ts mxl\<^sub>0 ?\<tau>s" using mxl
+ by (auto simp: wt_start_def ty\<^sub>i'_def ty\<^sub>l_def list_all2_conv_all_nth nth_Cons
+ split: nat.split)
+ moreover {
+ fix pc assume pc: "pc < size ?is"
+ then have "?P,T\<^sub>r,mxs,size ?is,?xt \<turnstile> ?is!pc,pc :: ?\<tau>s"
+ proof(cases "pc < length (compE\<^sub>2 e)")
+ case True
+ then show ?thesis using mxs wte wt_compE
+ by (clarsimp simp: compT\<^sub>a_def mxl after_def wt_instrs_def)
+ next
+ case False
+ have "length (compE\<^sub>2 e) = pc" using less_antisym[OF False] pc by simp
+ then show ?thesis using mxl wte E_P wid
+ by (clarsimp simp: compT\<^sub>a_def after_def wt_defs xcpt_app_pcs xcpt_eff_pcs ty\<^sub>i'_def)
+ qed
+ }
+ moreover have "0 < size ?is" and "size ?\<tau>s = size ?is"
+ using wte by (simp_all add: compT\<^sub>a_def)
+ ultimately show ?thesis by(simp add: wt_method_def)
+qed
+(*>*)
+
+lemma compT_method_Static:
+ fixes e and A and C and Ts and mxl\<^sub>0
+ defines [simp]: "E \<equiv> Ts"
+ and [simp]: "A \<equiv> \<lfloor>{..<size Ts}\<rfloor>"
+ and [simp]: "A' \<equiv> A \<squnion> \<A> e"
+ and [simp]: "mxl\<^sub>0 \<equiv> max_vars e"
+ assumes mxs: "max_stack e = mxs"
+ and mxl: "length Ts + max_vars e = mxl"
+ assumes wf: "wf_prog p P" and wte: "P,E \<turnstile>\<^sub>1 e :: T" and \<D>: "\<D> e A"
+ and \<B>: "\<B> e (size E)" and E_P: "set E \<subseteq> types P" and wid: "P \<turnstile> T \<le> T\<^sub>r"
+ shows "wt_method (compP\<^sub>2 P) C Static Ts T\<^sub>r mxs mxl\<^sub>0 (compE\<^sub>2 e @ [Return])
+ (compxE\<^sub>2 e 0 0) (ty\<^sub>i' [] E A # compT\<^sub>a E A [] e)"
+(*<*)(is "wt_method ?P C ?b Ts T\<^sub>r mxs mxl\<^sub>0 ?is ?xt ?\<tau>s")
+proof -
+ let ?n = "length E + mxl\<^sub>0"
+ have wt_compE: "P,T\<^sub>r,mxs \<turnstile> compE\<^sub>2 e, compxE\<^sub>2 e 0 (length []) [::]
+ TC0.ty\<^sub>i' ?n [] E A # TC1.compT P ?n E A [] e @[TC0.after P ?n E A [] e]"
+ using mxs TC2.compT_wt_instrs [OF wte \<D> \<B>, of "[]" mxs ?n T\<^sub>r] by simp
+
+ have "OK (ty\<^sub>i' [T] E A') \<in> states P mxs mxl"
+ using mxs WT\<^sub>1_is_type[OF wf wte E_P] max_stack1[OF wte] OK_ty\<^sub>i'_in_statesI[OF E_P]
+ by simp
+ moreover have "OK ` set (compT E A [] e) \<subseteq> states P mxs mxl"
+ using mxs mxl wid compT_states(1)[OF wf wte E_P] by simp
+ ultimately have "check_types ?P mxs ?n (map OK ?\<tau>s)"
+ using mxl wte E_P by (simp add: compT\<^sub>a_def after_def check_types_def)
+ moreover have "wt_start ?P C ?b Ts mxl\<^sub>0 ?\<tau>s" using mxl
+ by (auto simp: wt_start_def ty\<^sub>i'_def ty\<^sub>l_def list_all2_conv_all_nth nth_Cons
+ split: nat.split)
+ moreover {
+ fix pc assume pc: "pc < size ?is"
+ then have "?P,T\<^sub>r,mxs,size ?is,?xt \<turnstile> ?is!pc,pc :: ?\<tau>s"
+ proof(cases "pc < length (compE\<^sub>2 e)")
+ case True
+ then show ?thesis using mxs wte wt_compE
+ by (clarsimp simp: compT\<^sub>a_def mxl after_def wt_instrs_def)
+ next
+ case False
+ have "length (compE\<^sub>2 e) = pc" using less_antisym[OF False] pc by simp
+ then show ?thesis using mxl wte E_P wid
+ by (clarsimp simp: compT\<^sub>a_def after_def wt_defs xcpt_app_pcs xcpt_eff_pcs ty\<^sub>i'_def)
+ qed
+ }
+ moreover have "0 < size ?is" and "size ?\<tau>s = size ?is"
+ using wte by (simp_all add: compT\<^sub>a_def)
+ ultimately show ?thesis by(simp add: wt_method_def)
+qed
+(*>*)
+
+end
+
+definition compTP :: "J\<^sub>1_prog \<Rightarrow> ty\<^sub>P" where
+ "compTP P C M = (
+ let (D,b,Ts,T,e) = method P C M;
+ E = case b of Static \<Rightarrow> Ts | NonStatic \<Rightarrow> Class C # Ts;
+ A = case b of Static \<Rightarrow> \<lfloor>{..<size Ts}\<rfloor> | NonStatic \<Rightarrow> \<lfloor>{..size Ts}\<rfloor>;
+ mxl = (case b of Static \<Rightarrow> 0 | NonStatic \<Rightarrow> 1) + size Ts + max_vars e
+ in (TC0.ty\<^sub>i' mxl [] E A # TC1.compT\<^sub>a P mxl E A [] e))"
+
+theorem wt_compP\<^sub>2:
+assumes wf: "wf_J\<^sub>1_prog P" shows "wf_jvm_prog (compP\<^sub>2 P)"
+(*<*)
+proof -
+ let ?\<Phi> = "compTP P" and ?f = compMb\<^sub>2
+ let ?wf\<^sub>2 = "\<lambda>P C (M, b, Ts, T\<^sub>r, mxs, mxl\<^sub>0, is, xt).
+ wt_method P C b Ts T\<^sub>r mxs mxl\<^sub>0 is xt (?\<Phi> C M)"
+ and ?P = "compP ?f P"
+ { fix C M b Ts T m let ?md = "(M,b,Ts,T,m)::expr\<^sub>1 mdecl"
+ assume cM: "P \<turnstile> C sees M, b : Ts\<rightarrow>T = m in C"
+ and wfm: "wf_mdecl wf_J\<^sub>1_mdecl P C ?md"
+ then have Ts_types: "\<forall>T'\<in>set Ts. is_type P T'"
+ and T_type: "is_type P T" and wfm\<^sub>1: "wf_J\<^sub>1_mdecl P C ?md"
+ by(simp_all add: wf_mdecl_def)
+ have Ts_P: "set Ts \<subseteq> types P" using sees_wf_mdecl[OF wf cM]
+ by (clarsimp simp: wf_mdecl_def)
+ then have CTs_P: "is_class P C \<and> set Ts \<subseteq> types P"
+ using sees_method_is_class[OF cM] by simp
+
+ have "?wf\<^sub>2 ?P C (M,b,Ts,T,?f b m)"
+ proof(cases b)
+ case Static
+ then obtain T' where wte: "P,Ts \<turnstile>\<^sub>1 m :: T'" and wid: "P \<turnstile> T' \<le> T"
+ and \<D>: "\<D> m \<lfloor>{..<size Ts}\<rfloor>" and \<B>: "\<B> m (size Ts)"
+ using wfm\<^sub>1 by(auto simp: wf_mdecl_def)
+ show ?thesis using Static cM TC2.compT_method_Static [OF _ _ wf wte \<D> \<B> Ts_P wid]
+ by(simp add: compTP_def)
+ next
+ case NonStatic
+ then obtain T' where wte: "P,Class C#Ts \<turnstile>\<^sub>1 m :: T'" and wid: "P \<turnstile> T' \<le> T"
+ and \<D>: "\<D> m \<lfloor>{..size Ts}\<rfloor>" and \<B>: "\<B> m (Suc (size Ts))"
+ using wfm\<^sub>1 by(auto simp: wf_mdecl_def)
+ have Ts_P: "set Ts \<subseteq> types P" using sees_wf_mdecl[OF wf cM]
+ by (fastforce simp: wf_mdecl_def intro: sees_method_is_class)
+ show ?thesis using NonStatic cM
+ TC2.compT_method_NonStatic [simplified, OF _ _ wf wte \<D> \<B> CTs_P wid]
+ by(simp add: compTP_def)
+ qed
+ then have "wf_mdecl ?wf\<^sub>2 ?P C (M, b, Ts, T, ?f b m)"
+ using Ts_types T_type by(simp add: wf_mdecl_def)
+ }
+ then have "wf_prog ?wf\<^sub>2 (compP ?f P)" by(rule wf_prog_compPI[OF _ wf])
+ then show ?thesis by (simp add: wf_jvm_prog_def wf_jvm_prog_phi_def) fast
+qed
+(*>*)
+
+theorem wt_J2JVM:
+ "wf_J_prog P \<Longrightarrow> wf_jvm_prog (J2JVM P)"
+(*<*)
+by(simp only:o_def J2JVM_def)
+ (blast intro:wt_compP\<^sub>2 compP\<^sub>1_pres_wf)
+
+end
diff --git a/thys/JinjaDCI/J/Annotate.thy b/thys/JinjaDCI/J/Annotate.thy
--- a/thys/JinjaDCI/J/Annotate.thy
+++ b/thys/JinjaDCI/J/Annotate.thy
@@ -1,78 +1,78 @@
-(* Title: JinjaDCI/J/Annotate.thy
-
- Author: Tobias Nipkow, Susannah Mansky
- Copyright 2003 Technische Universitaet Muenchen, 2019-20 UIUC
-
- Based on the Jinja theory J/Annotate.thy by Tobias Nipkow
-*)
-
-section \<open> Program annotation \<close>
-
-theory Annotate imports WellType begin
-
-(*<*)
-abbreviation (output)
- unanFAcc :: "expr \<Rightarrow> vname \<Rightarrow> expr" ("(_\<bullet>_)" [10,10] 90) where
- "unanFAcc e F == FAcc e F []"
-
-abbreviation (output)
- unanFAss :: "expr \<Rightarrow> vname \<Rightarrow> expr \<Rightarrow> expr" ("(_\<bullet>_ := _)" [10,0,90] 90) where
- "unanFAss e F e' == FAss e F [] e'"
-(*>*)
-
-inductive
- Anno :: "[J_prog,env, expr , expr] \<Rightarrow> bool"
- ("_,_ \<turnstile> _ \<leadsto> _" [51,0,0,51]50)
- and Annos :: "[J_prog,env, expr list, expr list] \<Rightarrow> bool"
- ("_,_ \<turnstile> _ [\<leadsto>] _" [51,0,0,51]50)
- for P :: J_prog
-where
-
- AnnoNew: "P,E \<turnstile> new C \<leadsto> new C"
-| AnnoCast: "P,E \<turnstile> e \<leadsto> e' \<Longrightarrow> P,E \<turnstile> Cast C e \<leadsto> Cast C e'"
-| AnnoVal: "P,E \<turnstile> Val v \<leadsto> Val v"
-| AnnoVarVar: "E V = \<lfloor>T\<rfloor> \<Longrightarrow> P,E \<turnstile> Var V \<leadsto> Var V"
-| AnnoVarField: "\<lbrakk> E V = None; E this = \<lfloor>Class C\<rfloor>; P \<turnstile> C sees V,NonStatic:T in D \<rbrakk>
- \<Longrightarrow> P,E \<turnstile> Var V \<leadsto> Var this\<bullet>V{D}"
-| AnnoBinOp:
- "\<lbrakk> P,E \<turnstile> e1 \<leadsto> e1'; P,E \<turnstile> e2 \<leadsto> e2' \<rbrakk>
- \<Longrightarrow> P,E \<turnstile> e1 \<guillemotleft>bop\<guillemotright> e2 \<leadsto> e1' \<guillemotleft>bop\<guillemotright> e2'"
-| AnnoLAssVar:
- "\<lbrakk> E V = \<lfloor>T\<rfloor>; P,E \<turnstile> e \<leadsto> e' \<rbrakk> \<Longrightarrow> P,E \<turnstile> V:=e \<leadsto> V:=e'"
-| AnnoLAssField:
- "\<lbrakk> E V = None; E this = \<lfloor>Class C\<rfloor>; P \<turnstile> C sees V,NonStatic:T in D; P,E \<turnstile> e \<leadsto> e' \<rbrakk>
- \<Longrightarrow> P,E \<turnstile> V:=e \<leadsto> Var this\<bullet>V{D} := e'"
-| AnnoFAcc:
- "\<lbrakk> P,E \<turnstile> e \<leadsto> e'; P,E \<turnstile> e' :: Class C; P \<turnstile> C sees F,NonStatic:T in D \<rbrakk>
- \<Longrightarrow> P,E \<turnstile> e\<bullet>F{[]} \<leadsto> e'\<bullet>F{D}"
-| AnnoSFAcc:
- "\<lbrakk> P \<turnstile> C sees F,Static:T in D \<rbrakk>
- \<Longrightarrow> P,E \<turnstile> C\<bullet>\<^sub>sF{[]} \<leadsto> C\<bullet>\<^sub>sF{D}"
-| AnnoFAss: "\<lbrakk> P,E \<turnstile> e1 \<leadsto> e1'; P,E \<turnstile> e2 \<leadsto> e2';
- P,E \<turnstile> e1' :: Class C; P \<turnstile> C sees F,NonStatic:T in D \<rbrakk>
- \<Longrightarrow> P,E \<turnstile> e1\<bullet>F{[]} := e2 \<leadsto> e1'\<bullet>F{D} := e2'"
-| AnnoSFAss: "\<lbrakk> P,E \<turnstile> e2 \<leadsto> e2'; P \<turnstile> C sees F,Static:T in D \<rbrakk>
- \<Longrightarrow> P,E \<turnstile> C\<bullet>\<^sub>sF{[]} := e2 \<leadsto> C\<bullet>\<^sub>sF{D} := e2'"
-| AnnoCall:
- "\<lbrakk> P,E \<turnstile> e \<leadsto> e'; P,E \<turnstile> es [\<leadsto>] es' \<rbrakk>
- \<Longrightarrow> P,E \<turnstile> Call e M es \<leadsto> Call e' M es'"
-| AnnoSCall:
- "\<lbrakk> P,E \<turnstile> es [\<leadsto>] es' \<rbrakk>
- \<Longrightarrow> P,E \<turnstile> SCall C M es \<leadsto> SCall C M es'"
-| AnnoBlock:
- "P,E(V \<mapsto> T) \<turnstile> e \<leadsto> e' \<Longrightarrow> P,E \<turnstile> {V:T; e} \<leadsto> {V:T; e'}"
-| AnnoComp: "\<lbrakk> P,E \<turnstile> e1 \<leadsto> e1'; P,E \<turnstile> e2 \<leadsto> e2' \<rbrakk>
- \<Longrightarrow> P,E \<turnstile> e1;;e2 \<leadsto> e1';;e2'"
-| AnnoCond: "\<lbrakk> P,E \<turnstile> e \<leadsto> e'; P,E \<turnstile> e1 \<leadsto> e1'; P,E \<turnstile> e2 \<leadsto> e2' \<rbrakk>
- \<Longrightarrow> P,E \<turnstile> if (e) e1 else e2 \<leadsto> if (e') e1' else e2'"
-| AnnoLoop: "\<lbrakk> P,E \<turnstile> e \<leadsto> e'; P,E \<turnstile> c \<leadsto> c' \<rbrakk>
- \<Longrightarrow> P,E \<turnstile> while (e) c \<leadsto> while (e') c'"
-| AnnoThrow: "P,E \<turnstile> e \<leadsto> e' \<Longrightarrow> P,E \<turnstile> throw e \<leadsto> throw e'"
-| AnnoTry: "\<lbrakk> P,E \<turnstile> e1 \<leadsto> e1'; P,E(V \<mapsto> Class C) \<turnstile> e2 \<leadsto> e2' \<rbrakk>
- \<Longrightarrow> P,E \<turnstile> try e1 catch(C V) e2 \<leadsto> try e1' catch(C V) e2'"
-
-| AnnoNil: "P,E \<turnstile> [] [\<leadsto>] []"
-| AnnoCons: "\<lbrakk> P,E \<turnstile> e \<leadsto> e'; P,E \<turnstile> es [\<leadsto>] es' \<rbrakk>
- \<Longrightarrow> P,E \<turnstile> e#es [\<leadsto>] e'#es'"
-
-end
+(* Title: JinjaDCI/J/Annotate.thy
+
+ Author: Tobias Nipkow, Susannah Mansky
+ Copyright 2003 Technische Universitaet Muenchen, 2019-20 UIUC
+
+ Based on the Jinja theory J/Annotate.thy by Tobias Nipkow
+*)
+
+section \<open> Program annotation \<close>
+
+theory Annotate imports WellType begin
+
+(*<*)
+abbreviation (output)
+ unanFAcc :: "expr \<Rightarrow> vname \<Rightarrow> expr" ("(_\<bullet>_)" [10,10] 90) where
+ "unanFAcc e F == FAcc e F []"
+
+abbreviation (output)
+ unanFAss :: "expr \<Rightarrow> vname \<Rightarrow> expr \<Rightarrow> expr" ("(_\<bullet>_ := _)" [10,0,90] 90) where
+ "unanFAss e F e' == FAss e F [] e'"
+(*>*)
+
+inductive
+ Anno :: "[J_prog,env, expr , expr] \<Rightarrow> bool"
+ ("_,_ \<turnstile> _ \<leadsto> _" [51,0,0,51]50)
+ and Annos :: "[J_prog,env, expr list, expr list] \<Rightarrow> bool"
+ ("_,_ \<turnstile> _ [\<leadsto>] _" [51,0,0,51]50)
+ for P :: J_prog
+where
+
+ AnnoNew: "P,E \<turnstile> new C \<leadsto> new C"
+| AnnoCast: "P,E \<turnstile> e \<leadsto> e' \<Longrightarrow> P,E \<turnstile> Cast C e \<leadsto> Cast C e'"
+| AnnoVal: "P,E \<turnstile> Val v \<leadsto> Val v"
+| AnnoVarVar: "E V = \<lfloor>T\<rfloor> \<Longrightarrow> P,E \<turnstile> Var V \<leadsto> Var V"
+| AnnoVarField: "\<lbrakk> E V = None; E this = \<lfloor>Class C\<rfloor>; P \<turnstile> C sees V,NonStatic:T in D \<rbrakk>
+ \<Longrightarrow> P,E \<turnstile> Var V \<leadsto> Var this\<bullet>V{D}"
+| AnnoBinOp:
+ "\<lbrakk> P,E \<turnstile> e1 \<leadsto> e1'; P,E \<turnstile> e2 \<leadsto> e2' \<rbrakk>
+ \<Longrightarrow> P,E \<turnstile> e1 \<guillemotleft>bop\<guillemotright> e2 \<leadsto> e1' \<guillemotleft>bop\<guillemotright> e2'"
+| AnnoLAssVar:
+ "\<lbrakk> E V = \<lfloor>T\<rfloor>; P,E \<turnstile> e \<leadsto> e' \<rbrakk> \<Longrightarrow> P,E \<turnstile> V:=e \<leadsto> V:=e'"
+| AnnoLAssField:
+ "\<lbrakk> E V = None; E this = \<lfloor>Class C\<rfloor>; P \<turnstile> C sees V,NonStatic:T in D; P,E \<turnstile> e \<leadsto> e' \<rbrakk>
+ \<Longrightarrow> P,E \<turnstile> V:=e \<leadsto> Var this\<bullet>V{D} := e'"
+| AnnoFAcc:
+ "\<lbrakk> P,E \<turnstile> e \<leadsto> e'; P,E \<turnstile> e' :: Class C; P \<turnstile> C sees F,NonStatic:T in D \<rbrakk>
+ \<Longrightarrow> P,E \<turnstile> e\<bullet>F{[]} \<leadsto> e'\<bullet>F{D}"
+| AnnoSFAcc:
+ "\<lbrakk> P \<turnstile> C sees F,Static:T in D \<rbrakk>
+ \<Longrightarrow> P,E \<turnstile> C\<bullet>\<^sub>sF{[]} \<leadsto> C\<bullet>\<^sub>sF{D}"
+| AnnoFAss: "\<lbrakk> P,E \<turnstile> e1 \<leadsto> e1'; P,E \<turnstile> e2 \<leadsto> e2';
+ P,E \<turnstile> e1' :: Class C; P \<turnstile> C sees F,NonStatic:T in D \<rbrakk>
+ \<Longrightarrow> P,E \<turnstile> e1\<bullet>F{[]} := e2 \<leadsto> e1'\<bullet>F{D} := e2'"
+| AnnoSFAss: "\<lbrakk> P,E \<turnstile> e2 \<leadsto> e2'; P \<turnstile> C sees F,Static:T in D \<rbrakk>
+ \<Longrightarrow> P,E \<turnstile> C\<bullet>\<^sub>sF{[]} := e2 \<leadsto> C\<bullet>\<^sub>sF{D} := e2'"
+| AnnoCall:
+ "\<lbrakk> P,E \<turnstile> e \<leadsto> e'; P,E \<turnstile> es [\<leadsto>] es' \<rbrakk>
+ \<Longrightarrow> P,E \<turnstile> Call e M es \<leadsto> Call e' M es'"
+| AnnoSCall:
+ "\<lbrakk> P,E \<turnstile> es [\<leadsto>] es' \<rbrakk>
+ \<Longrightarrow> P,E \<turnstile> SCall C M es \<leadsto> SCall C M es'"
+| AnnoBlock:
+ "P,E(V \<mapsto> T) \<turnstile> e \<leadsto> e' \<Longrightarrow> P,E \<turnstile> {V:T; e} \<leadsto> {V:T; e'}"
+| AnnoComp: "\<lbrakk> P,E \<turnstile> e1 \<leadsto> e1'; P,E \<turnstile> e2 \<leadsto> e2' \<rbrakk>
+ \<Longrightarrow> P,E \<turnstile> e1;;e2 \<leadsto> e1';;e2'"
+| AnnoCond: "\<lbrakk> P,E \<turnstile> e \<leadsto> e'; P,E \<turnstile> e1 \<leadsto> e1'; P,E \<turnstile> e2 \<leadsto> e2' \<rbrakk>
+ \<Longrightarrow> P,E \<turnstile> if (e) e1 else e2 \<leadsto> if (e') e1' else e2'"
+| AnnoLoop: "\<lbrakk> P,E \<turnstile> e \<leadsto> e'; P,E \<turnstile> c \<leadsto> c' \<rbrakk>
+ \<Longrightarrow> P,E \<turnstile> while (e) c \<leadsto> while (e') c'"
+| AnnoThrow: "P,E \<turnstile> e \<leadsto> e' \<Longrightarrow> P,E \<turnstile> throw e \<leadsto> throw e'"
+| AnnoTry: "\<lbrakk> P,E \<turnstile> e1 \<leadsto> e1'; P,E(V \<mapsto> Class C) \<turnstile> e2 \<leadsto> e2' \<rbrakk>
+ \<Longrightarrow> P,E \<turnstile> try e1 catch(C V) e2 \<leadsto> try e1' catch(C V) e2'"
+
+| AnnoNil: "P,E \<turnstile> [] [\<leadsto>] []"
+| AnnoCons: "\<lbrakk> P,E \<turnstile> e \<leadsto> e'; P,E \<turnstile> es [\<leadsto>] es' \<rbrakk>
+ \<Longrightarrow> P,E \<turnstile> e#es [\<leadsto>] e'#es'"
+
+end
diff --git a/thys/JinjaDCI/J/BigStep.thy b/thys/JinjaDCI/J/BigStep.thy
--- a/thys/JinjaDCI/J/BigStep.thy
+++ b/thys/JinjaDCI/J/BigStep.thy
@@ -1,699 +1,699 @@
-(* Title: JinjaDCI/J/BigStep.thy
-
- Author: Tobias Nipkow, Susannah Mansky
- Copyright 2003 Technische Universitaet Muenchen, 2019-20 UIUC
-
- Based on the Jinja theory J/BigStep.thy by Tobias Nipkow
-*)
-
-section \<open> Big Step Semantics \<close>
-
-theory BigStep imports Expr State WWellForm begin
-
-inductive
- eval :: "J_prog \<Rightarrow> expr \<Rightarrow> state \<Rightarrow> expr \<Rightarrow> state \<Rightarrow> bool"
- ("_ \<turnstile> ((1\<langle>_,/_\<rangle>) \<Rightarrow>/ (1\<langle>_,/_\<rangle>))" [51,0,0,0,0] 81)
- and evals :: "J_prog \<Rightarrow> expr list \<Rightarrow> state \<Rightarrow> expr list \<Rightarrow> state \<Rightarrow> bool"
- ("_ \<turnstile> ((1\<langle>_,/_\<rangle>) [\<Rightarrow>]/ (1\<langle>_,/_\<rangle>))" [51,0,0,0,0] 81)
- for P :: J_prog
-where
-
- New:
- "\<lbrakk> sh C = Some (sfs, Done); new_Addr h = Some a;
- P \<turnstile> C has_fields FDTs; h' = h(a\<mapsto>blank P C) \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>new C,(h,l,sh)\<rangle> \<Rightarrow> \<langle>addr a,(h',l,sh)\<rangle>"
-
-| NewFail:
- "\<lbrakk> sh C = Some (sfs, Done); new_Addr h = None; is_class P C \<rbrakk> \<Longrightarrow>
- P \<turnstile> \<langle>new C, (h,l,sh)\<rangle> \<Rightarrow> \<langle>THROW OutOfMemory,(h,l,sh)\<rangle>"
-
-| NewInit:
- "\<lbrakk> \<nexists>sfs. sh C = Some (sfs, Done); P \<turnstile> \<langle>INIT C ([C],False) \<leftarrow> unit,(h,l,sh)\<rangle> \<Rightarrow> \<langle>Val v',(h',l',sh')\<rangle>;
- new_Addr h' = Some a; P \<turnstile> C has_fields FDTs; h'' = h'(a\<mapsto>blank P C) \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>new C,(h,l,sh)\<rangle> \<Rightarrow> \<langle>addr a,(h'',l',sh')\<rangle>"
-
-| NewInitOOM:
- "\<lbrakk> \<nexists>sfs. sh C = Some (sfs, Done); P \<turnstile> \<langle>INIT C ([C],False) \<leftarrow> unit,(h,l,sh)\<rangle> \<Rightarrow> \<langle>Val v',(h',l',sh')\<rangle>;
- new_Addr h' = None; is_class P C \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>new C,(h,l,sh)\<rangle> \<Rightarrow> \<langle>THROW OutOfMemory,(h',l',sh')\<rangle>"
-
-| NewInitThrow:
- "\<lbrakk> \<nexists>sfs. sh C = Some (sfs, Done); P \<turnstile> \<langle>INIT C ([C],False) \<leftarrow> unit,(h,l,sh)\<rangle> \<Rightarrow> \<langle>throw a,s'\<rangle>;
- is_class P C \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>new C,(h,l,sh)\<rangle> \<Rightarrow> \<langle>throw a,s'\<rangle>"
-
-| Cast:
- "\<lbrakk> P \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>addr a,(h,l,sh)\<rangle>; h a = Some(D,fs); P \<turnstile> D \<preceq>\<^sup>* C \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>Cast C e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>addr a,(h,l,sh)\<rangle>"
-
-| CastNull:
- "P \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>null,s\<^sub>1\<rangle> \<Longrightarrow>
- P \<turnstile> \<langle>Cast C e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>null,s\<^sub>1\<rangle>"
-
-| CastFail:
- "\<lbrakk> P \<turnstile> \<langle>e,s\<^sub>0\<rangle>\<Rightarrow> \<langle>addr a,(h,l,sh)\<rangle>; h a = Some(D,fs); \<not> P \<turnstile> D \<preceq>\<^sup>* C \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>Cast C e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>THROW ClassCast,(h,l,sh)\<rangle>"
-
-| CastThrow:
- "P \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle> \<Longrightarrow>
- P \<turnstile> \<langle>Cast C e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle>"
-
-| Val:
- "P \<turnstile> \<langle>Val v,s\<rangle> \<Rightarrow> \<langle>Val v,s\<rangle>"
-
-| BinOp:
- "\<lbrakk> P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Val v\<^sub>1,s\<^sub>1\<rangle>; P \<turnstile> \<langle>e\<^sub>2,s\<^sub>1\<rangle> \<Rightarrow> \<langle>Val v\<^sub>2,s\<^sub>2\<rangle>; binop(bop,v\<^sub>1,v\<^sub>2) = Some v \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Val v,s\<^sub>2\<rangle>"
-
-| BinOpThrow1:
- "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e,s\<^sub>1\<rangle> \<Longrightarrow>
- P \<turnstile> \<langle>e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2, s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e,s\<^sub>1\<rangle>"
-
-| BinOpThrow2:
- "\<lbrakk> P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Val v\<^sub>1,s\<^sub>1\<rangle>; P \<turnstile> \<langle>e\<^sub>2,s\<^sub>1\<rangle> \<Rightarrow> \<langle>throw e,s\<^sub>2\<rangle> \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e,s\<^sub>2\<rangle>"
-
-| Var:
- "l V = Some v \<Longrightarrow>
- P \<turnstile> \<langle>Var V,(h,l,sh)\<rangle> \<Rightarrow> \<langle>Val v,(h,l,sh)\<rangle>"
-
-| LAss:
- "\<lbrakk> P \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Val v,(h,l,sh)\<rangle>; l' = l(V\<mapsto>v) \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>V:=e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>unit,(h,l',sh)\<rangle>"
-
-| LAssThrow:
- "P \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle> \<Longrightarrow>
- P \<turnstile> \<langle>V:=e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle>"
-
-| FAcc:
- "\<lbrakk> P \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>addr a,(h,l,sh)\<rangle>; h a = Some(C,fs);
- P \<turnstile> C has F,NonStatic:t in D;
- fs(F,D) = Some v \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>e\<bullet>F{D},s\<^sub>0\<rangle> \<Rightarrow> \<langle>Val v,(h,l,sh)\<rangle>"
-
-| FAccNull:
- "P \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>null,s\<^sub>1\<rangle> \<Longrightarrow>
- P \<turnstile> \<langle>e\<bullet>F{D},s\<^sub>0\<rangle> \<Rightarrow> \<langle>THROW NullPointer,s\<^sub>1\<rangle>"
-
-| FAccThrow:
- "P \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle> \<Longrightarrow>
- P \<turnstile> \<langle>e\<bullet>F{D},s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle>"
-
-| FAccNone:
- "\<lbrakk> P \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>addr a,(h,l,sh)\<rangle>; h a = Some(C,fs);
- \<not>(\<exists>b t. P \<turnstile> C has F,b:t in D) \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>e\<bullet>F{D},s\<^sub>0\<rangle> \<Rightarrow> \<langle>THROW NoSuchFieldError,(h,l,sh)\<rangle>"
-
-| FAccStatic:
- "\<lbrakk> P \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>addr a,(h,l,sh)\<rangle>; h a = Some(C,fs);
- P \<turnstile> C has F,Static:t in D \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>e\<bullet>F{D},s\<^sub>0\<rangle> \<Rightarrow> \<langle>THROW IncompatibleClassChangeError,(h,l,sh)\<rangle>"
-
-| SFAcc:
- "\<lbrakk> P \<turnstile> C has F,Static:t in D;
- sh D = Some (sfs,Done);
- sfs F = Some v \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D},(h,l,sh)\<rangle> \<Rightarrow> \<langle>Val v,(h,l,sh)\<rangle>"
-
-| SFAccInit:
- "\<lbrakk> P \<turnstile> C has F,Static:t in D;
- \<nexists>sfs. sh D = Some (sfs,Done); P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(h,l,sh)\<rangle> \<Rightarrow> \<langle>Val v',(h',l',sh')\<rangle>;
- sh' D = Some (sfs,i);
- sfs F = Some v \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D},(h,l,sh)\<rangle> \<Rightarrow> \<langle>Val v,(h',l',sh')\<rangle>"
-
-| SFAccInitThrow:
- "\<lbrakk> P \<turnstile> C has F,Static:t in D;
- \<nexists>sfs. sh D = Some (sfs,Done); P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(h,l,sh)\<rangle> \<Rightarrow> \<langle>throw a,s'\<rangle> \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D},(h,l,sh)\<rangle> \<Rightarrow> \<langle>throw a,s'\<rangle>"
-
-| SFAccNone:
- "\<lbrakk> \<not>(\<exists>b t. P \<turnstile> C has F,b:t in D) \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D},s\<rangle> \<Rightarrow> \<langle>THROW NoSuchFieldError,s\<rangle>"
-
-| SFAccNonStatic:
- "\<lbrakk> P \<turnstile> C has F,NonStatic:t in D \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D},s\<rangle> \<Rightarrow> \<langle>THROW IncompatibleClassChangeError,s\<rangle>"
-
-| FAss:
- "\<lbrakk> P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow> \<langle>addr a,s\<^sub>1\<rangle>; P \<turnstile> \<langle>e\<^sub>2,s\<^sub>1\<rangle> \<Rightarrow> \<langle>Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\<rangle>;
- h\<^sub>2 a = Some(C,fs); P \<turnstile> C has F,NonStatic:t in D;
- fs' = fs((F,D)\<mapsto>v); h\<^sub>2' = h\<^sub>2(a\<mapsto>(C,fs')) \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>e\<^sub>1\<bullet>F{D}:=e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>unit,(h\<^sub>2',l\<^sub>2,sh\<^sub>2)\<rangle>"
-
-| FAssNull:
- "\<lbrakk> P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow> \<langle>null,s\<^sub>1\<rangle>; P \<turnstile> \<langle>e\<^sub>2,s\<^sub>1\<rangle> \<Rightarrow> \<langle>Val v,s\<^sub>2\<rangle> \<rbrakk> \<Longrightarrow>
- P \<turnstile> \<langle>e\<^sub>1\<bullet>F{D}:=e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>THROW NullPointer,s\<^sub>2\<rangle>"
-
-| FAssThrow1:
- "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle> \<Longrightarrow>
- P \<turnstile> \<langle>e\<^sub>1\<bullet>F{D}:=e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle>"
-
-| FAssThrow2:
- "\<lbrakk> P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Val v,s\<^sub>1\<rangle>; P \<turnstile> \<langle>e\<^sub>2,s\<^sub>1\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>2\<rangle> \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>e\<^sub>1\<bullet>F{D}:=e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>2\<rangle>"
-
-| FAssNone:
- "\<lbrakk> P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow> \<langle>addr a,s\<^sub>1\<rangle>; P \<turnstile> \<langle>e\<^sub>2,s\<^sub>1\<rangle> \<Rightarrow> \<langle>Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\<rangle>;
- h\<^sub>2 a = Some(C,fs); \<not>(\<exists>b t. P \<turnstile> C has F,b:t in D) \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>e\<^sub>1\<bullet>F{D}:=e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>THROW NoSuchFieldError,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\<rangle>"
-
-| FAssStatic:
- "\<lbrakk> P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow> \<langle>addr a,s\<^sub>1\<rangle>; P \<turnstile> \<langle>e\<^sub>2,s\<^sub>1\<rangle> \<Rightarrow> \<langle>Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\<rangle>;
- h\<^sub>2 a = Some(C,fs); P \<turnstile> C has F,Static:t in D \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>e\<^sub>1\<bullet>F{D}:=e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>THROW IncompatibleClassChangeError,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\<rangle>"
-
-| SFAss:
- "\<lbrakk> P \<turnstile> \<langle>e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Val v,(h\<^sub>1,l\<^sub>1,sh\<^sub>1)\<rangle>;
- P \<turnstile> C has F,Static:t in D;
- sh\<^sub>1 D = Some(sfs,Done); sfs' = sfs(F\<mapsto>v); sh\<^sub>1' = sh\<^sub>1(D\<mapsto>(sfs',Done)) \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D}:=e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>unit,(h\<^sub>1,l\<^sub>1,sh\<^sub>1')\<rangle>"
-
-| SFAssInit:
- "\<lbrakk> P \<turnstile> \<langle>e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Val v,(h\<^sub>1,l\<^sub>1,sh\<^sub>1)\<rangle>;
- P \<turnstile> C has F,Static:t in D;
- \<nexists>sfs. sh\<^sub>1 D = Some(sfs,Done); P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(h\<^sub>1,l\<^sub>1,sh\<^sub>1)\<rangle> \<Rightarrow> \<langle>Val v',(h',l',sh')\<rangle>;
- sh' D = Some(sfs,i);
- sfs' = sfs(F\<mapsto>v); sh'' = sh'(D\<mapsto>(sfs',i)) \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D}:=e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>unit,(h',l',sh'')\<rangle>"
-
-| SFAssInitThrow:
- "\<lbrakk> P \<turnstile> \<langle>e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Val v,(h\<^sub>1,l\<^sub>1,sh\<^sub>1)\<rangle>;
- P \<turnstile> C has F,Static:t in D;
- \<nexists>sfs. sh\<^sub>1 D = Some(sfs,Done); P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(h\<^sub>1,l\<^sub>1,sh\<^sub>1)\<rangle> \<Rightarrow> \<langle>throw a,s'\<rangle> \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D}:=e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw a,s'\<rangle>"
-
-| SFAssThrow:
- "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>2\<rangle>
- \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D}:=e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>2\<rangle>"
-
-| SFAssNone:
- "\<lbrakk> P \<turnstile> \<langle>e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\<rangle>;
- \<not>(\<exists>b t. P \<turnstile> C has F,b:t in D) \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D}:=e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>THROW NoSuchFieldError,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\<rangle>"
-
-| SFAssNonStatic:
- "\<lbrakk> P \<turnstile> \<langle>e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\<rangle>;
- P \<turnstile> C has F,NonStatic:t in D \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D}:=e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>THROW IncompatibleClassChangeError,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\<rangle>"
-
-| CallObjThrow:
- "P \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle> \<Longrightarrow>
- P \<turnstile> \<langle>e\<bullet>M(ps),s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle>"
-
-| CallParamsThrow:
- "\<lbrakk> P \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Val v,s\<^sub>1\<rangle>; P \<turnstile> \<langle>es,s\<^sub>1\<rangle> [\<Rightarrow>] \<langle>map Val vs @ throw ex # es',s\<^sub>2\<rangle> \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>e\<bullet>M(es),s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw ex,s\<^sub>2\<rangle>"
-
-| CallNull:
- "\<lbrakk> P \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>null,s\<^sub>1\<rangle>; P \<turnstile> \<langle>ps,s\<^sub>1\<rangle> [\<Rightarrow>] \<langle>map Val vs,s\<^sub>2\<rangle> \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>e\<bullet>M(ps),s\<^sub>0\<rangle> \<Rightarrow> \<langle>THROW NullPointer,s\<^sub>2\<rangle>"
-
-| CallNone:
- "\<lbrakk> P \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>addr a,s\<^sub>1\<rangle>; P \<turnstile> \<langle>ps,s\<^sub>1\<rangle> [\<Rightarrow>] \<langle>map Val vs,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\<rangle>;
- h\<^sub>2 a = Some(C,fs); \<not>(\<exists>b Ts T m D. P \<turnstile> C sees M,b:Ts\<rightarrow>T = m in D) \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>e\<bullet>M(ps),s\<^sub>0\<rangle> \<Rightarrow> \<langle>THROW NoSuchMethodError,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\<rangle>"
-
-| CallStatic:
- "\<lbrakk> P \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>addr a,s\<^sub>1\<rangle>; P \<turnstile> \<langle>ps,s\<^sub>1\<rangle> [\<Rightarrow>] \<langle>map Val vs,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\<rangle>;
- h\<^sub>2 a = Some(C,fs); P \<turnstile> C sees M,Static:Ts\<rightarrow>T = m in D \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>e\<bullet>M(ps),s\<^sub>0\<rangle> \<Rightarrow> \<langle>THROW IncompatibleClassChangeError,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\<rangle>"
-
-| Call:
- "\<lbrakk> P \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>addr a,s\<^sub>1\<rangle>; P \<turnstile> \<langle>ps,s\<^sub>1\<rangle> [\<Rightarrow>] \<langle>map Val vs,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\<rangle>;
- h\<^sub>2 a = Some(C,fs); P \<turnstile> C sees M,NonStatic:Ts\<rightarrow>T = (pns,body) in D;
- length vs = length pns; l\<^sub>2' = [this\<mapsto>Addr a, pns[\<mapsto>]vs];
- P \<turnstile> \<langle>body,(h\<^sub>2,l\<^sub>2',sh\<^sub>2)\<rangle> \<Rightarrow> \<langle>e',(h\<^sub>3,l\<^sub>3,sh\<^sub>3)\<rangle> \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>e\<bullet>M(ps),s\<^sub>0\<rangle> \<Rightarrow> \<langle>e',(h\<^sub>3,l\<^sub>2,sh\<^sub>3)\<rangle>"
-
-| SCallParamsThrow:
- "\<lbrakk> P \<turnstile> \<langle>es,s\<^sub>0\<rangle> [\<Rightarrow>] \<langle>map Val vs @ throw ex # es',s\<^sub>2\<rangle> \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sM(es),s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw ex,s\<^sub>2\<rangle>"
-
-| SCallNone:
- "\<lbrakk> P \<turnstile> \<langle>ps,s\<^sub>0\<rangle> [\<Rightarrow>] \<langle>map Val vs,s\<^sub>2\<rangle>;
- \<not>(\<exists>b Ts T m D. P \<turnstile> C sees M,b:Ts\<rightarrow>T = m in D) \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sM(ps),s\<^sub>0\<rangle> \<Rightarrow> \<langle>THROW NoSuchMethodError,s\<^sub>2\<rangle>"
-
-| SCallNonStatic:
- "\<lbrakk> P \<turnstile> \<langle>ps,s\<^sub>0\<rangle> [\<Rightarrow>] \<langle>map Val vs,s\<^sub>2\<rangle>;
- P \<turnstile> C sees M,NonStatic:Ts\<rightarrow>T = m in D \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sM(ps),s\<^sub>0\<rangle> \<Rightarrow> \<langle>THROW IncompatibleClassChangeError,s\<^sub>2\<rangle>"
-
-| SCallInitThrow:
- "\<lbrakk> P \<turnstile> \<langle>ps,s\<^sub>0\<rangle> [\<Rightarrow>] \<langle>map Val vs,(h\<^sub>1,l\<^sub>1,sh\<^sub>1)\<rangle>;
- P \<turnstile> C sees M,Static:Ts\<rightarrow>T = (pns,body) in D;
- \<nexists>sfs. sh\<^sub>1 D = Some(sfs,Done); M \<noteq> clinit;
- P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(h\<^sub>1,l\<^sub>1,sh\<^sub>1)\<rangle> \<Rightarrow> \<langle>throw a,s'\<rangle> \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sM(ps),s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw a,s'\<rangle>"
-
-| SCallInit:
- "\<lbrakk> P \<turnstile> \<langle>ps,s\<^sub>0\<rangle> [\<Rightarrow>] \<langle>map Val vs,(h\<^sub>1,l\<^sub>1,sh\<^sub>1)\<rangle>;
- P \<turnstile> C sees M,Static:Ts\<rightarrow>T = (pns,body) in D;
- \<nexists>sfs. sh\<^sub>1 D = Some(sfs,Done); M \<noteq> clinit;
- P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(h\<^sub>1,l\<^sub>1,sh\<^sub>1)\<rangle> \<Rightarrow> \<langle>Val v',(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\<rangle>;
- length vs = length pns; l\<^sub>2' = [pns[\<mapsto>]vs];
- P \<turnstile> \<langle>body,(h\<^sub>2,l\<^sub>2',sh\<^sub>2)\<rangle> \<Rightarrow> \<langle>e',(h\<^sub>3,l\<^sub>3,sh\<^sub>3)\<rangle> \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sM(ps),s\<^sub>0\<rangle> \<Rightarrow> \<langle>e',(h\<^sub>3,l\<^sub>2,sh\<^sub>3)\<rangle>"
-
-| SCall:
- "\<lbrakk> P \<turnstile> \<langle>ps,s\<^sub>0\<rangle> [\<Rightarrow>] \<langle>map Val vs,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\<rangle>;
- P \<turnstile> C sees M,Static:Ts\<rightarrow>T = (pns,body) in D;
- sh\<^sub>2 D = Some(sfs,Done) \<or> (M = clinit \<and> sh\<^sub>2 D = Some(sfs,Processing));
- length vs = length pns; l\<^sub>2' = [pns[\<mapsto>]vs];
- P \<turnstile> \<langle>body,(h\<^sub>2,l\<^sub>2',sh\<^sub>2)\<rangle> \<Rightarrow> \<langle>e',(h\<^sub>3,l\<^sub>3,sh\<^sub>3)\<rangle> \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sM(ps),s\<^sub>0\<rangle> \<Rightarrow> \<langle>e',(h\<^sub>3,l\<^sub>2,sh\<^sub>3)\<rangle>"
-
-| Block:
- "P \<turnstile> \<langle>e\<^sub>0,(h\<^sub>0,l\<^sub>0(V:=None),sh\<^sub>0)\<rangle> \<Rightarrow> \<langle>e\<^sub>1,(h\<^sub>1,l\<^sub>1,sh\<^sub>1)\<rangle> \<Longrightarrow>
- P \<turnstile> \<langle>{V:T; e\<^sub>0},(h\<^sub>0,l\<^sub>0,sh\<^sub>0)\<rangle> \<Rightarrow> \<langle>e\<^sub>1,(h\<^sub>1,l\<^sub>1(V:=l\<^sub>0 V),sh\<^sub>1)\<rangle>"
-
-| Seq:
- "\<lbrakk> P \<turnstile> \<langle>e\<^sub>0,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Val v,s\<^sub>1\<rangle>; P \<turnstile> \<langle>e\<^sub>1,s\<^sub>1\<rangle> \<Rightarrow> \<langle>e\<^sub>2,s\<^sub>2\<rangle> \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>e\<^sub>0;;e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow> \<langle>e\<^sub>2,s\<^sub>2\<rangle>"
-
-| SeqThrow:
- "P \<turnstile> \<langle>e\<^sub>0,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e,s\<^sub>1\<rangle> \<Longrightarrow>
- P \<turnstile> \<langle>e\<^sub>0;;e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e,s\<^sub>1\<rangle>"
-
-| CondT:
- "\<lbrakk> P \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>true,s\<^sub>1\<rangle>; P \<turnstile> \<langle>e\<^sub>1,s\<^sub>1\<rangle> \<Rightarrow> \<langle>e',s\<^sub>2\<rangle> \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>if (e) e\<^sub>1 else e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>e',s\<^sub>2\<rangle>"
-
-| CondF:
- "\<lbrakk> P \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>false,s\<^sub>1\<rangle>; P \<turnstile> \<langle>e\<^sub>2,s\<^sub>1\<rangle> \<Rightarrow> \<langle>e',s\<^sub>2\<rangle> \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>if (e) e\<^sub>1 else e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>e',s\<^sub>2\<rangle>"
-
-| CondThrow:
- "P \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle> \<Longrightarrow>
- P \<turnstile> \<langle>if (e) e\<^sub>1 else e\<^sub>2, s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle>"
-
-| WhileF:
- "P \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>false,s\<^sub>1\<rangle> \<Longrightarrow>
- P \<turnstile> \<langle>while (e) c,s\<^sub>0\<rangle> \<Rightarrow> \<langle>unit,s\<^sub>1\<rangle>"
-
-| WhileT:
- "\<lbrakk> P \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>true,s\<^sub>1\<rangle>; P \<turnstile> \<langle>c,s\<^sub>1\<rangle> \<Rightarrow> \<langle>Val v\<^sub>1,s\<^sub>2\<rangle>; P \<turnstile> \<langle>while (e) c,s\<^sub>2\<rangle> \<Rightarrow> \<langle>e\<^sub>3,s\<^sub>3\<rangle> \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>while (e) c,s\<^sub>0\<rangle> \<Rightarrow> \<langle>e\<^sub>3,s\<^sub>3\<rangle>"
-
-| WhileCondThrow:
- "P \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle> \<Longrightarrow>
- P \<turnstile> \<langle>while (e) c,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle>"
-
-| WhileBodyThrow:
- "\<lbrakk> P \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>true,s\<^sub>1\<rangle>; P \<turnstile> \<langle>c,s\<^sub>1\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>2\<rangle>\<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>while (e) c,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>2\<rangle>"
-
-| Throw:
- "P \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>addr a,s\<^sub>1\<rangle> \<Longrightarrow>
- P \<turnstile> \<langle>throw e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Throw a,s\<^sub>1\<rangle>"
-
-| ThrowNull:
- "P \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>null,s\<^sub>1\<rangle> \<Longrightarrow>
- P \<turnstile> \<langle>throw e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>THROW NullPointer,s\<^sub>1\<rangle>"
-
-| ThrowThrow:
- "P \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle> \<Longrightarrow>
- P \<turnstile> \<langle>throw e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle>"
-
-| Try:
- "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Val v\<^sub>1,s\<^sub>1\<rangle> \<Longrightarrow>
- P \<turnstile> \<langle>try e\<^sub>1 catch(C V) e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Val v\<^sub>1,s\<^sub>1\<rangle>"
-
-| TryCatch:
- "\<lbrakk> P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Throw a,(h\<^sub>1,l\<^sub>1,sh\<^sub>1)\<rangle>; h\<^sub>1 a = Some(D,fs); P \<turnstile> D \<preceq>\<^sup>* C;
- P \<turnstile> \<langle>e\<^sub>2,(h\<^sub>1,l\<^sub>1(V\<mapsto>Addr a),sh\<^sub>1)\<rangle> \<Rightarrow> \<langle>e\<^sub>2',(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\<rangle> \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>try e\<^sub>1 catch(C V) e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>e\<^sub>2',(h\<^sub>2,l\<^sub>2(V:=l\<^sub>1 V),sh\<^sub>2)\<rangle>"
-
-| TryThrow:
- "\<lbrakk> P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Throw a,(h\<^sub>1,l\<^sub>1,sh\<^sub>1)\<rangle>; h\<^sub>1 a = Some(D,fs); \<not> P \<turnstile> D \<preceq>\<^sup>* C \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>try e\<^sub>1 catch(C V) e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Throw a,(h\<^sub>1,l\<^sub>1,sh\<^sub>1)\<rangle>"
-
-| Nil:
- "P \<turnstile> \<langle>[],s\<rangle> [\<Rightarrow>] \<langle>[],s\<rangle>"
-
-| Cons:
- "\<lbrakk> P \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Val v,s\<^sub>1\<rangle>; P \<turnstile> \<langle>es,s\<^sub>1\<rangle> [\<Rightarrow>] \<langle>es',s\<^sub>2\<rangle> \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>e#es,s\<^sub>0\<rangle> [\<Rightarrow>] \<langle>Val v # es',s\<^sub>2\<rangle>"
-
-| ConsThrow:
- "P \<turnstile> \<langle>e, s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e', s\<^sub>1\<rangle> \<Longrightarrow>
- P \<turnstile> \<langle>e#es, s\<^sub>0\<rangle> [\<Rightarrow>] \<langle>throw e' # es, s\<^sub>1\<rangle>"
-
-\<comment> \<open> init rules \<close>
-
-| InitFinal:
- "P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>INIT C (Nil,b) \<leftarrow> e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
-
-| InitNone:
- "\<lbrakk> sh C = None; P \<turnstile> \<langle>INIT C' (C#Cs,False) \<leftarrow> e,(h,l,sh(C \<mapsto> (sblank P C, Prepared)))\<rangle> \<Rightarrow> \<langle>e',s'\<rangle> \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>INIT C' (C#Cs,False) \<leftarrow> e,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
-
-| InitDone:
- "\<lbrakk> sh C = Some(sfs,Done); P \<turnstile> \<langle>INIT C' (Cs,True) \<leftarrow> e,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle> \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>INIT C' (C#Cs,False) \<leftarrow> e,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
-
-| InitProcessing:
- "\<lbrakk> sh C = Some(sfs,Processing); P \<turnstile> \<langle>INIT C' (Cs,True) \<leftarrow> e,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle> \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>INIT C' (C#Cs,False) \<leftarrow> e,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
-
-\<comment> \<open> note that @{text RI} will mark all classes in the list Cs with the Error flag \<close>
-| InitError:
- "\<lbrakk> sh C = Some(sfs,Error);
- P \<turnstile> \<langle>RI (C, THROW NoClassDefFoundError);Cs \<leftarrow> e,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle> \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>INIT C' (C#Cs,False) \<leftarrow> e,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
-
-| InitObject:
- "\<lbrakk> sh C = Some(sfs,Prepared);
- C = Object;
- sh' = sh(C \<mapsto> (sfs,Processing));
- P \<turnstile> \<langle>INIT C' (C#Cs,True) \<leftarrow> e,(h,l,sh')\<rangle> \<Rightarrow> \<langle>e',s'\<rangle> \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>INIT C' (C#Cs,False) \<leftarrow> e,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
-
-| InitNonObject:
- "\<lbrakk> sh C = Some(sfs,Prepared);
- C \<noteq> Object;
- class P C = Some (D,r);
- sh' = sh(C \<mapsto> (sfs,Processing));
- P \<turnstile> \<langle>INIT C' (D#C#Cs,False) \<leftarrow> e,(h,l,sh')\<rangle> \<Rightarrow> \<langle>e',s'\<rangle> \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>INIT C' (C#Cs,False) \<leftarrow> e,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
-
-| InitRInit:
- "P \<turnstile> \<langle>RI (C,C\<bullet>\<^sub>sclinit([]));Cs \<leftarrow> e,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>
- \<Longrightarrow> P \<turnstile> \<langle>INIT C' (C#Cs,True) \<leftarrow> e,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
-
-| RInit:
- "\<lbrakk> P \<turnstile> \<langle>e',s\<rangle> \<Rightarrow> \<langle>Val v, (h',l',sh')\<rangle>;
- sh' C = Some(sfs, i); sh'' = sh'(C \<mapsto> (sfs, Done));
- C' = last(C#Cs);
- P \<turnstile> \<langle>INIT C' (Cs,True) \<leftarrow> e, (h',l',sh'')\<rangle> \<Rightarrow> \<langle>e\<^sub>1,s\<^sub>1\<rangle> \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>RI (C,e');Cs \<leftarrow> e,s\<rangle> \<Rightarrow> \<langle>e\<^sub>1,s\<^sub>1\<rangle>"
-
-| RInitInitFail:
- "\<lbrakk> P \<turnstile> \<langle>e',s\<rangle> \<Rightarrow> \<langle>throw a, (h',l',sh')\<rangle>;
- sh' C = Some(sfs, i); sh'' = sh'(C \<mapsto> (sfs, Error));
- P \<turnstile> \<langle>RI (D,throw a);Cs \<leftarrow> e, (h',l',sh'')\<rangle> \<Rightarrow> \<langle>e\<^sub>1,s\<^sub>1\<rangle> \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>RI (C,e');D#Cs \<leftarrow> e,s\<rangle> \<Rightarrow> \<langle>e\<^sub>1,s\<^sub>1\<rangle>"
-
-| RInitFailFinal:
- "\<lbrakk> P \<turnstile> \<langle>e',s\<rangle> \<Rightarrow> \<langle>throw a, (h',l',sh')\<rangle>;
- sh' C = Some(sfs, i); sh'' = sh'(C \<mapsto> (sfs, Error)) \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>RI (C,e');Nil \<leftarrow> e,s\<rangle> \<Rightarrow> \<langle>throw a, (h',l',sh'')\<rangle>"
-
-(*<*)
-lemmas eval_evals_induct = eval_evals.induct [split_format (complete)]
- and eval_evals_inducts = eval_evals.inducts [split_format (complete)]
-
-inductive_cases eval_cases [cases set]:
- "P \<turnstile> \<langle>new C,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
- "P \<turnstile> \<langle>Cast C e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
- "P \<turnstile> \<langle>Val v,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
- "P \<turnstile> \<langle>e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
- "P \<turnstile> \<langle>Var v,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
- "P \<turnstile> \<langle>V:=e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
- "P \<turnstile> \<langle>e\<bullet>F{D},s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
- "P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D},s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
- "P \<turnstile> \<langle>e\<^sub>1\<bullet>F{D}:=e\<^sub>2,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
- "P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D}:=e\<^sub>2,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
- "P \<turnstile> \<langle>e\<bullet>M(es),s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
- "P \<turnstile> \<langle>C\<bullet>\<^sub>sM(es),s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
- "P \<turnstile> \<langle>{V:T;e\<^sub>1},s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
- "P \<turnstile> \<langle>e\<^sub>1;;e\<^sub>2,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
- "P \<turnstile> \<langle>if (e) e\<^sub>1 else e\<^sub>2,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
- "P \<turnstile> \<langle>while (b) c,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
- "P \<turnstile> \<langle>throw e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
- "P \<turnstile> \<langle>try e\<^sub>1 catch(C V) e\<^sub>2,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
- "P \<turnstile> \<langle>INIT C (Cs,b) \<leftarrow> e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
- "P \<turnstile> \<langle>RI (C,e);Cs \<leftarrow> e\<^sub>0,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
-
-inductive_cases evals_cases [cases set]:
- "P \<turnstile> \<langle>[],s\<rangle> [\<Rightarrow>] \<langle>e',s'\<rangle>"
- "P \<turnstile> \<langle>e#es,s\<rangle> [\<Rightarrow>] \<langle>e',s'\<rangle>"
-(*>*)
-
-subsection "Final expressions"
-
-lemma eval_final: "P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle> \<Longrightarrow> final e'"
- and evals_final: "P \<turnstile> \<langle>es,s\<rangle> [\<Rightarrow>] \<langle>es',s'\<rangle> \<Longrightarrow> finals es'"
-(*<*)by(induct rule:eval_evals.inducts, simp_all)(*>*)
-
-text\<open> Only used later, in the small to big translation, but is already a
-good sanity check: \<close>
-
-lemma eval_finalId: "final e \<Longrightarrow> P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>e,s\<rangle>"
-(*<*)by (erule finalE) (iprover intro: eval_evals.intros)+(*>*)
-
-lemma eval_final_same: "\<lbrakk> P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>; final e \<rbrakk> \<Longrightarrow> e = e' \<and> s = s'"
-(*<*)by(auto elim!: finalE eval_cases)(*>*)
-
-lemma eval_finalsId:
-assumes finals: "finals es" shows "P \<turnstile> \<langle>es,s\<rangle> [\<Rightarrow>] \<langle>es,s\<rangle>"
-(*<*)
- using finals
-proof (induct es type: list)
- case Nil show ?case by (rule eval_evals.intros)
-next
- case (Cons e es)
- have hyp: "finals es \<Longrightarrow> P \<turnstile> \<langle>es,s\<rangle> [\<Rightarrow>] \<langle>es,s\<rangle>"
- and finals: "finals (e # es)" by fact+
- show "P \<turnstile> \<langle>e # es,s\<rangle> [\<Rightarrow>] \<langle>e # es,s\<rangle>"
- proof cases
- assume "final e"
- thus ?thesis
- proof (cases rule: finalE)
- fix v assume e: "e = Val v"
- have "P \<turnstile> \<langle>Val v,s\<rangle> \<Rightarrow> \<langle>Val v,s\<rangle>" by (simp add: eval_finalId)
- moreover from finals e have "P \<turnstile> \<langle>es,s\<rangle> [\<Rightarrow>] \<langle>es,s\<rangle>" by(fast intro:hyp)
- ultimately have "P \<turnstile> \<langle>Val v#es,s\<rangle> [\<Rightarrow>] \<langle>Val v#es,s\<rangle>"
- by (rule eval_evals.intros)
- with e show ?thesis by simp
- next
- fix a assume e: "e = Throw a"
- have "P \<turnstile> \<langle>Throw a,s\<rangle> \<Rightarrow> \<langle>Throw a,s\<rangle>" by (simp add: eval_finalId)
- hence "P \<turnstile> \<langle>Throw a#es,s\<rangle> [\<Rightarrow>] \<langle>Throw a#es,s\<rangle>" by (rule eval_evals.intros)
- with e show ?thesis by simp
- qed
- next
- assume "\<not> final e"
- with not_finals_ConsI finals have False by blast
- thus ?thesis ..
- qed
-qed
-(*>*)
-
-lemma evals_finals_same:
-assumes finals: "finals es"
-shows "P \<turnstile> \<langle>es,s\<rangle> [\<Rightarrow>] \<langle>es',s'\<rangle> \<Longrightarrow> es = es' \<and> s = s'"
- using finals
-proof (induct es arbitrary: es' type: list)
- case Nil then show ?case using evals_cases(1) by blast
-next
- case (Cons e es)
- have IH: "\<And>es'. P \<turnstile> \<langle>es,s\<rangle> [\<Rightarrow>] \<langle>es',s'\<rangle> \<Longrightarrow> finals es \<Longrightarrow> es = es' \<and> s = s'"
- and step: "P \<turnstile> \<langle>e # es,s\<rangle> [\<Rightarrow>] \<langle>es',s'\<rangle>" and finals: "finals (e # es)" by fact+
- then obtain e' es'' where es': "es' = e'#es''" by (meson Cons.prems(1) evals_cases(2))
- have fe: "final e" using finals not_finals_ConsI by auto
- show ?case
- proof(rule evals_cases(2)[OF step])
- fix v s\<^sub>1 es1 assume es1: "es' = Val v # es1"
- and estep: "P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>Val v,s\<^sub>1\<rangle>" and esstep: "P \<turnstile> \<langle>es,s\<^sub>1\<rangle> [\<Rightarrow>] \<langle>es1,s'\<rangle>"
- then have "e = Val v" using eval_final_same fe by auto
- then have "finals es" using es' not_finals_ConsI2 finals by blast
- then show ?thesis using es' IH estep esstep es1 eval_final_same fe by blast
- next
- fix e' assume es1: "es' = throw e' # es" and estep: "P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>throw e',s'\<rangle>"
- then have "e = throw e'" using eval_final_same fe by auto
- then show ?thesis using es' estep es1 eval_final_same fe by blast
- qed
-qed
-(*>*)
-
-subsection "Property preservation"
-
-lemma evals_length: "P \<turnstile> \<langle>es,s\<rangle> [\<Rightarrow>] \<langle>es',s'\<rangle> \<Longrightarrow> length es = length es'"
- by(induct es arbitrary:es' s s', auto elim: evals_cases)
-
-corollary evals_empty: "P \<turnstile> \<langle>es,s\<rangle> [\<Rightarrow>] \<langle>es',s'\<rangle> \<Longrightarrow> (es = []) = (es' = [])"
- by(drule evals_length, fastforce)
-
-(****)
-
-theorem eval_hext: "P \<turnstile> \<langle>e,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',(h',l',sh')\<rangle> \<Longrightarrow> h \<unlhd> h'"
-and evals_hext: "P \<turnstile> \<langle>es,(h,l,sh)\<rangle> [\<Rightarrow>] \<langle>es',(h',l',sh')\<rangle> \<Longrightarrow> h \<unlhd> h'"
-(*<*)
-proof (induct rule: eval_evals_inducts)
- case New thus ?case
- by(fastforce intro!: hext_new intro:LeastI simp:new_Addr_def
- split:if_split_asm simp del:fun_upd_apply)
-next
- case NewInit thus ?case
- by (meson hext_new hext_trans new_Addr_SomeD)
-next
- case FAss thus ?case
- by(auto simp:sym[THEN hext_upd_obj] simp del:fun_upd_apply
- elim!: hext_trans)
-qed (auto elim!: hext_trans)
-(*>*)
-
-
-lemma eval_lcl_incr: "P \<turnstile> \<langle>e,(h\<^sub>0,l\<^sub>0,sh\<^sub>0)\<rangle> \<Rightarrow> \<langle>e',(h\<^sub>1,l\<^sub>1,sh\<^sub>1)\<rangle> \<Longrightarrow> dom l\<^sub>0 \<subseteq> dom l\<^sub>1"
- and evals_lcl_incr: "P \<turnstile> \<langle>es,(h\<^sub>0,l\<^sub>0,sh\<^sub>0)\<rangle> [\<Rightarrow>] \<langle>es',(h\<^sub>1,l\<^sub>1,sh\<^sub>1)\<rangle> \<Longrightarrow> dom l\<^sub>0 \<subseteq> dom l\<^sub>1"
-(*<*)
-proof (induct rule: eval_evals_inducts)
- case BinOp show ?case by(rule subset_trans)(rule BinOp.hyps)+
-next
- case Call thus ?case
- by(simp del: fun_upd_apply)
-next
- case Seq show ?case by(rule subset_trans)(rule Seq.hyps)+
-next
- case CondT show ?case by(rule subset_trans)(rule CondT.hyps)+
-next
- case CondF show ?case by(rule subset_trans)(rule CondF.hyps)+
-next
- case WhileT thus ?case by(blast)
-next
- case TryCatch thus ?case by(clarsimp simp:dom_def split:if_split_asm) blast
-next
- case Cons show ?case by(rule subset_trans)(rule Cons.hyps)+
-next
- case Block thus ?case by(auto simp del:fun_upd_apply)
-qed auto
-(*>*)
-
-lemma
-shows init_ri_same_loc: "P \<turnstile> \<langle>e,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',(h',l',sh')\<rangle>
- \<Longrightarrow> (\<And>C Cs b M a. e = INIT C (Cs,b) \<leftarrow> unit \<or> e = C\<bullet>\<^sub>sM([]) \<or> e = RI (C,Throw a) ; Cs \<leftarrow> unit
- \<or> e = RI (C,C\<bullet>\<^sub>sclinit([])) ; Cs \<leftarrow> unit
- \<Longrightarrow> l = l')"
- and "P \<turnstile> \<langle>es,(h,l,sh)\<rangle> [\<Rightarrow>] \<langle>es',(h',l',sh')\<rangle> \<Longrightarrow> True"
-proof(induct rule: eval_evals_inducts)
- case (RInitInitFail e h l sh a')
- then show ?case using eval_final[of _ _ _ "throw a'"]
- by(fastforce dest: eval_final_same[of _ "Throw a"])
-next
- case RInitFailFinal then show ?case by(auto dest: eval_final_same)
-qed(auto dest: evals_cases eval_cases(17) eval_final_same)
-
-lemma init_same_loc: "P \<turnstile> \<langle>INIT C (Cs,b) \<leftarrow> unit,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',(h',l',sh')\<rangle> \<Longrightarrow> l = l'"
- by(simp add: init_ri_same_loc)
-
-(****)
-
-lemma assumes wf: "wwf_J_prog P"
-shows eval_proc_pres': "P \<turnstile> \<langle>e,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',(h',l',sh')\<rangle>
- \<Longrightarrow> not_init C e \<Longrightarrow> \<exists>sfs. sh C = \<lfloor>(sfs, Processing)\<rfloor> \<Longrightarrow> \<exists>sfs'. sh' C = \<lfloor>(sfs', Processing)\<rfloor>"
- and evals_proc_pres': "P \<turnstile> \<langle>es,(h,l,sh)\<rangle> [\<Rightarrow>] \<langle>es',(h',l',sh')\<rangle>
- \<Longrightarrow> not_inits C es \<Longrightarrow> \<exists>sfs. sh C = \<lfloor>(sfs, Processing)\<rfloor> \<Longrightarrow> \<exists>sfs'. sh' C = \<lfloor>(sfs', Processing)\<rfloor>"
-(*<*)
-proof(induct rule:eval_evals_inducts)
- case Call then show ?case using sees_wwf_nsub_RI[OF wf Call.hyps(6)] nsub_RI_not_init by auto
-next
- case (SCallInit ps h l sh vs h\<^sub>1 l\<^sub>1 sh\<^sub>1 C' M Ts T pns body D v' h\<^sub>2 l\<^sub>2 sh\<^sub>2 l\<^sub>2' e' h\<^sub>3 l\<^sub>3 sh\<^sub>3)
- then show ?case
- using SCallInit sees_wwf_nsub_RI[OF wf SCallInit.hyps(3)] nsub_RI_not_init[of body] by auto
-next
- case SCall then show ?case using sees_wwf_nsub_RI[OF wf SCall.hyps(3)] nsub_RI_not_init by auto
-next
- case (InitNone sh C1 C' Cs h l e' a a b) then show ?case by(cases "C = C1") auto
-next
- case (InitDone sh C sfs C' Cs h l e' a a b) then show ?case by(cases Cs, auto)
-next
- case (InitProcessing sh C sfs C' Cs h l e' a a b) then show ?case by(cases Cs, auto)
-next
- case (InitError sh C1 sfs Cs h l e' a a b C') then show ?case by(cases "C = C1") auto
-next
- case (InitObject sh C1 sfs sh' C' Cs h l e' a a b) then show ?case by(cases "C = C1") auto
-next
- case (InitNonObject sh C1 sfs D a b sh' C' Cs h l e' a a b)
- then show ?case by(cases "C = C1") auto
-next
- case (RInit e a a b v h' l' sh' C sfs i sh'' C' Cs e\<^sub>1 a a b) then show ?case by(cases Cs, auto)
-next
- case (RInitInitFail e h l sh a h' l' sh' C1 sfs i sh'' D Cs e\<^sub>1 h1 l1 sh1)
- then show ?case using eval_final by fastforce
-qed(auto)
-
-(************************************************)
-
-subsection\<open>Init Elimination rules\<close>
-
-lemma init_NilE:
-assumes init: "P \<turnstile> \<langle>INIT C (Nil,b) \<leftarrow> unit,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
-shows "e' = unit \<and> s' = s"
-proof(rule eval_cases(19)[OF init]) \<comment> \<open> Init \<close> qed(auto dest: eval_final_same)
-
-lemma init_ProcessingE:
-assumes shC: "sh C = \<lfloor>(sfs, Processing)\<rfloor>"
- and init: "P \<turnstile> \<langle>INIT C ([C],False) \<leftarrow> unit,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
-shows "e' = unit \<and> s' = (h,l,sh)"
-proof(rule eval_cases(19)[OF init]) \<comment> \<open> Init \<close>
- fix sha Ca sfs Cs ha la
- assume "(h, l, sh) = (ha, la, sha)" and "sha Ca = \<lfloor>(sfs, Processing)\<rfloor>"
- and "P \<turnstile> \<langle>INIT C (Cs,True) \<leftarrow> unit,(ha, la, sha)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>" and "[C] = Ca # Cs"
- then show ?thesis using init_NilE by simp
-next
- fix sha sfs Cs ha la
- assume "(h, l, sh) = (ha, la, sha)" and "sha Object = \<lfloor>(sfs, Prepared)\<rfloor>"
- and "[C] = Object # Cs"
- then show ?thesis using shC by clarsimp
-qed(auto simp: assms)
-
-
-lemma rinit_throwE:
-"P \<turnstile> \<langle>RI (C,throw e) ; Cs \<leftarrow> e\<^sub>0,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>
- \<Longrightarrow> \<exists>a s\<^sub>t. e' = throw a \<and> P \<turnstile> \<langle>throw e,s\<rangle> \<Rightarrow> \<langle>throw a,s\<^sub>t\<rangle>"
-proof(induct Cs arbitrary: C e s)
- case Nil
- then show ?case
- proof(rule eval_cases(20)) \<comment> \<open> RI \<close>
- fix v h' l' sh' assume "P \<turnstile> \<langle>throw e,s\<rangle> \<Rightarrow> \<langle>Val v,(h', l', sh')\<rangle>"
- then show ?case using eval_cases(17) by blast
- qed(auto)
-next
- case (Cons C' Cs')
- show ?case using Cons.prems(1)
- proof(rule eval_cases(20)) \<comment> \<open> RI \<close>
- fix v h' l' sh' assume "P \<turnstile> \<langle>throw e,s\<rangle> \<Rightarrow> \<langle>Val v,(h', l', sh')\<rangle>"
- then show ?case using eval_cases(17) by blast
- next
- fix a h' l' sh' sfs i D Cs''
- assume e''step: "P \<turnstile> \<langle>throw e,s\<rangle> \<Rightarrow> \<langle>throw a,(h', l', sh')\<rangle>"
- and shC: "sh' C = \<lfloor>(sfs, i)\<rfloor>"
- and riD: "P \<turnstile> \<langle>RI (D,throw a) ; Cs'' \<leftarrow> e\<^sub>0,(h', l', sh'(C \<mapsto> (sfs, Error)))\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
- and "C' # Cs' = D # Cs''"
- then show ?thesis using Cons.hyps eval_final eval_final_same by blast
- qed(simp)
-qed
-
-lemma rinit_ValE:
-assumes ri: "P \<turnstile> \<langle>RI (C,e) ; Cs \<leftarrow> unit,s\<rangle> \<Rightarrow> \<langle>Val v',s'\<rangle>"
-shows "\<exists>v'' s''. P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>Val v'',s''\<rangle>"
-proof(rule eval_cases(20)[OF ri]) \<comment> \<open> RI \<close>
- fix a h' l' sh' sfs i D Cs'
- assume "P \<turnstile> \<langle>RI (D,throw a) ; Cs' \<leftarrow> unit,(h', l', sh'(C \<mapsto> (sfs, Error)))\<rangle> \<Rightarrow> \<langle>Val v',s'\<rangle>"
- then show ?thesis using rinit_throwE by blast
-qed(auto)
-
-subsection "Some specific evaluations"
-
-lemma lass_val_of_eval:
- "\<lbrakk> lass_val_of e = \<lfloor>a\<rfloor>; P \<turnstile> \<langle>e,(h, l, sh)\<rangle> \<Rightarrow> \<langle>e',(h', l', sh')\<rangle> \<rbrakk>
- \<Longrightarrow> e' = unit \<and> h' = h \<and> l' = l(fst a\<mapsto>snd a) \<and> sh' = sh"
- by(drule lass_val_of_spec, auto elim: eval.cases)
-
-lemma eval_throw_nonVal:
-assumes eval: "P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>throw a,s'\<rangle>"
-shows "val_of e = None"
-proof(cases "val_of e")
- case (Some v) show ?thesis using eval by(auto simp: val_of_spec[OF Some] intro: eval.cases)
-qed(simp)
-
-lemma eval_throw_nonLAss:
-assumes eval: "P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>throw a,s'\<rangle>"
-shows "lass_val_of e = None"
-proof(cases "lass_val_of e")
- case (Some v) show ?thesis using eval by(auto simp: lass_val_of_spec[OF Some] intro: eval.cases)
-qed(simp)
-
-end
+(* Title: JinjaDCI/J/BigStep.thy
+
+ Author: Tobias Nipkow, Susannah Mansky
+ Copyright 2003 Technische Universitaet Muenchen, 2019-20 UIUC
+
+ Based on the Jinja theory J/BigStep.thy by Tobias Nipkow
+*)
+
+section \<open> Big Step Semantics \<close>
+
+theory BigStep imports Expr State WWellForm begin
+
+inductive
+ eval :: "J_prog \<Rightarrow> expr \<Rightarrow> state \<Rightarrow> expr \<Rightarrow> state \<Rightarrow> bool"
+ ("_ \<turnstile> ((1\<langle>_,/_\<rangle>) \<Rightarrow>/ (1\<langle>_,/_\<rangle>))" [51,0,0,0,0] 81)
+ and evals :: "J_prog \<Rightarrow> expr list \<Rightarrow> state \<Rightarrow> expr list \<Rightarrow> state \<Rightarrow> bool"
+ ("_ \<turnstile> ((1\<langle>_,/_\<rangle>) [\<Rightarrow>]/ (1\<langle>_,/_\<rangle>))" [51,0,0,0,0] 81)
+ for P :: J_prog
+where
+
+ New:
+ "\<lbrakk> sh C = Some (sfs, Done); new_Addr h = Some a;
+ P \<turnstile> C has_fields FDTs; h' = h(a\<mapsto>blank P C) \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>new C,(h,l,sh)\<rangle> \<Rightarrow> \<langle>addr a,(h',l,sh)\<rangle>"
+
+| NewFail:
+ "\<lbrakk> sh C = Some (sfs, Done); new_Addr h = None; is_class P C \<rbrakk> \<Longrightarrow>
+ P \<turnstile> \<langle>new C, (h,l,sh)\<rangle> \<Rightarrow> \<langle>THROW OutOfMemory,(h,l,sh)\<rangle>"
+
+| NewInit:
+ "\<lbrakk> \<nexists>sfs. sh C = Some (sfs, Done); P \<turnstile> \<langle>INIT C ([C],False) \<leftarrow> unit,(h,l,sh)\<rangle> \<Rightarrow> \<langle>Val v',(h',l',sh')\<rangle>;
+ new_Addr h' = Some a; P \<turnstile> C has_fields FDTs; h'' = h'(a\<mapsto>blank P C) \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>new C,(h,l,sh)\<rangle> \<Rightarrow> \<langle>addr a,(h'',l',sh')\<rangle>"
+
+| NewInitOOM:
+ "\<lbrakk> \<nexists>sfs. sh C = Some (sfs, Done); P \<turnstile> \<langle>INIT C ([C],False) \<leftarrow> unit,(h,l,sh)\<rangle> \<Rightarrow> \<langle>Val v',(h',l',sh')\<rangle>;
+ new_Addr h' = None; is_class P C \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>new C,(h,l,sh)\<rangle> \<Rightarrow> \<langle>THROW OutOfMemory,(h',l',sh')\<rangle>"
+
+| NewInitThrow:
+ "\<lbrakk> \<nexists>sfs. sh C = Some (sfs, Done); P \<turnstile> \<langle>INIT C ([C],False) \<leftarrow> unit,(h,l,sh)\<rangle> \<Rightarrow> \<langle>throw a,s'\<rangle>;
+ is_class P C \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>new C,(h,l,sh)\<rangle> \<Rightarrow> \<langle>throw a,s'\<rangle>"
+
+| Cast:
+ "\<lbrakk> P \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>addr a,(h,l,sh)\<rangle>; h a = Some(D,fs); P \<turnstile> D \<preceq>\<^sup>* C \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>Cast C e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>addr a,(h,l,sh)\<rangle>"
+
+| CastNull:
+ "P \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>null,s\<^sub>1\<rangle> \<Longrightarrow>
+ P \<turnstile> \<langle>Cast C e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>null,s\<^sub>1\<rangle>"
+
+| CastFail:
+ "\<lbrakk> P \<turnstile> \<langle>e,s\<^sub>0\<rangle>\<Rightarrow> \<langle>addr a,(h,l,sh)\<rangle>; h a = Some(D,fs); \<not> P \<turnstile> D \<preceq>\<^sup>* C \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>Cast C e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>THROW ClassCast,(h,l,sh)\<rangle>"
+
+| CastThrow:
+ "P \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle> \<Longrightarrow>
+ P \<turnstile> \<langle>Cast C e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle>"
+
+| Val:
+ "P \<turnstile> \<langle>Val v,s\<rangle> \<Rightarrow> \<langle>Val v,s\<rangle>"
+
+| BinOp:
+ "\<lbrakk> P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Val v\<^sub>1,s\<^sub>1\<rangle>; P \<turnstile> \<langle>e\<^sub>2,s\<^sub>1\<rangle> \<Rightarrow> \<langle>Val v\<^sub>2,s\<^sub>2\<rangle>; binop(bop,v\<^sub>1,v\<^sub>2) = Some v \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Val v,s\<^sub>2\<rangle>"
+
+| BinOpThrow1:
+ "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e,s\<^sub>1\<rangle> \<Longrightarrow>
+ P \<turnstile> \<langle>e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2, s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e,s\<^sub>1\<rangle>"
+
+| BinOpThrow2:
+ "\<lbrakk> P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Val v\<^sub>1,s\<^sub>1\<rangle>; P \<turnstile> \<langle>e\<^sub>2,s\<^sub>1\<rangle> \<Rightarrow> \<langle>throw e,s\<^sub>2\<rangle> \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e,s\<^sub>2\<rangle>"
+
+| Var:
+ "l V = Some v \<Longrightarrow>
+ P \<turnstile> \<langle>Var V,(h,l,sh)\<rangle> \<Rightarrow> \<langle>Val v,(h,l,sh)\<rangle>"
+
+| LAss:
+ "\<lbrakk> P \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Val v,(h,l,sh)\<rangle>; l' = l(V\<mapsto>v) \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>V:=e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>unit,(h,l',sh)\<rangle>"
+
+| LAssThrow:
+ "P \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle> \<Longrightarrow>
+ P \<turnstile> \<langle>V:=e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle>"
+
+| FAcc:
+ "\<lbrakk> P \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>addr a,(h,l,sh)\<rangle>; h a = Some(C,fs);
+ P \<turnstile> C has F,NonStatic:t in D;
+ fs(F,D) = Some v \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>e\<bullet>F{D},s\<^sub>0\<rangle> \<Rightarrow> \<langle>Val v,(h,l,sh)\<rangle>"
+
+| FAccNull:
+ "P \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>null,s\<^sub>1\<rangle> \<Longrightarrow>
+ P \<turnstile> \<langle>e\<bullet>F{D},s\<^sub>0\<rangle> \<Rightarrow> \<langle>THROW NullPointer,s\<^sub>1\<rangle>"
+
+| FAccThrow:
+ "P \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle> \<Longrightarrow>
+ P \<turnstile> \<langle>e\<bullet>F{D},s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle>"
+
+| FAccNone:
+ "\<lbrakk> P \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>addr a,(h,l,sh)\<rangle>; h a = Some(C,fs);
+ \<not>(\<exists>b t. P \<turnstile> C has F,b:t in D) \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>e\<bullet>F{D},s\<^sub>0\<rangle> \<Rightarrow> \<langle>THROW NoSuchFieldError,(h,l,sh)\<rangle>"
+
+| FAccStatic:
+ "\<lbrakk> P \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>addr a,(h,l,sh)\<rangle>; h a = Some(C,fs);
+ P \<turnstile> C has F,Static:t in D \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>e\<bullet>F{D},s\<^sub>0\<rangle> \<Rightarrow> \<langle>THROW IncompatibleClassChangeError,(h,l,sh)\<rangle>"
+
+| SFAcc:
+ "\<lbrakk> P \<turnstile> C has F,Static:t in D;
+ sh D = Some (sfs,Done);
+ sfs F = Some v \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D},(h,l,sh)\<rangle> \<Rightarrow> \<langle>Val v,(h,l,sh)\<rangle>"
+
+| SFAccInit:
+ "\<lbrakk> P \<turnstile> C has F,Static:t in D;
+ \<nexists>sfs. sh D = Some (sfs,Done); P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(h,l,sh)\<rangle> \<Rightarrow> \<langle>Val v',(h',l',sh')\<rangle>;
+ sh' D = Some (sfs,i);
+ sfs F = Some v \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D},(h,l,sh)\<rangle> \<Rightarrow> \<langle>Val v,(h',l',sh')\<rangle>"
+
+| SFAccInitThrow:
+ "\<lbrakk> P \<turnstile> C has F,Static:t in D;
+ \<nexists>sfs. sh D = Some (sfs,Done); P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(h,l,sh)\<rangle> \<Rightarrow> \<langle>throw a,s'\<rangle> \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D},(h,l,sh)\<rangle> \<Rightarrow> \<langle>throw a,s'\<rangle>"
+
+| SFAccNone:
+ "\<lbrakk> \<not>(\<exists>b t. P \<turnstile> C has F,b:t in D) \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D},s\<rangle> \<Rightarrow> \<langle>THROW NoSuchFieldError,s\<rangle>"
+
+| SFAccNonStatic:
+ "\<lbrakk> P \<turnstile> C has F,NonStatic:t in D \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D},s\<rangle> \<Rightarrow> \<langle>THROW IncompatibleClassChangeError,s\<rangle>"
+
+| FAss:
+ "\<lbrakk> P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow> \<langle>addr a,s\<^sub>1\<rangle>; P \<turnstile> \<langle>e\<^sub>2,s\<^sub>1\<rangle> \<Rightarrow> \<langle>Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\<rangle>;
+ h\<^sub>2 a = Some(C,fs); P \<turnstile> C has F,NonStatic:t in D;
+ fs' = fs((F,D)\<mapsto>v); h\<^sub>2' = h\<^sub>2(a\<mapsto>(C,fs')) \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>e\<^sub>1\<bullet>F{D}:=e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>unit,(h\<^sub>2',l\<^sub>2,sh\<^sub>2)\<rangle>"
+
+| FAssNull:
+ "\<lbrakk> P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow> \<langle>null,s\<^sub>1\<rangle>; P \<turnstile> \<langle>e\<^sub>2,s\<^sub>1\<rangle> \<Rightarrow> \<langle>Val v,s\<^sub>2\<rangle> \<rbrakk> \<Longrightarrow>
+ P \<turnstile> \<langle>e\<^sub>1\<bullet>F{D}:=e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>THROW NullPointer,s\<^sub>2\<rangle>"
+
+| FAssThrow1:
+ "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle> \<Longrightarrow>
+ P \<turnstile> \<langle>e\<^sub>1\<bullet>F{D}:=e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle>"
+
+| FAssThrow2:
+ "\<lbrakk> P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Val v,s\<^sub>1\<rangle>; P \<turnstile> \<langle>e\<^sub>2,s\<^sub>1\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>2\<rangle> \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>e\<^sub>1\<bullet>F{D}:=e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>2\<rangle>"
+
+| FAssNone:
+ "\<lbrakk> P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow> \<langle>addr a,s\<^sub>1\<rangle>; P \<turnstile> \<langle>e\<^sub>2,s\<^sub>1\<rangle> \<Rightarrow> \<langle>Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\<rangle>;
+ h\<^sub>2 a = Some(C,fs); \<not>(\<exists>b t. P \<turnstile> C has F,b:t in D) \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>e\<^sub>1\<bullet>F{D}:=e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>THROW NoSuchFieldError,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\<rangle>"
+
+| FAssStatic:
+ "\<lbrakk> P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow> \<langle>addr a,s\<^sub>1\<rangle>; P \<turnstile> \<langle>e\<^sub>2,s\<^sub>1\<rangle> \<Rightarrow> \<langle>Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\<rangle>;
+ h\<^sub>2 a = Some(C,fs); P \<turnstile> C has F,Static:t in D \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>e\<^sub>1\<bullet>F{D}:=e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>THROW IncompatibleClassChangeError,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\<rangle>"
+
+| SFAss:
+ "\<lbrakk> P \<turnstile> \<langle>e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Val v,(h\<^sub>1,l\<^sub>1,sh\<^sub>1)\<rangle>;
+ P \<turnstile> C has F,Static:t in D;
+ sh\<^sub>1 D = Some(sfs,Done); sfs' = sfs(F\<mapsto>v); sh\<^sub>1' = sh\<^sub>1(D\<mapsto>(sfs',Done)) \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D}:=e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>unit,(h\<^sub>1,l\<^sub>1,sh\<^sub>1')\<rangle>"
+
+| SFAssInit:
+ "\<lbrakk> P \<turnstile> \<langle>e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Val v,(h\<^sub>1,l\<^sub>1,sh\<^sub>1)\<rangle>;
+ P \<turnstile> C has F,Static:t in D;
+ \<nexists>sfs. sh\<^sub>1 D = Some(sfs,Done); P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(h\<^sub>1,l\<^sub>1,sh\<^sub>1)\<rangle> \<Rightarrow> \<langle>Val v',(h',l',sh')\<rangle>;
+ sh' D = Some(sfs,i);
+ sfs' = sfs(F\<mapsto>v); sh'' = sh'(D\<mapsto>(sfs',i)) \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D}:=e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>unit,(h',l',sh'')\<rangle>"
+
+| SFAssInitThrow:
+ "\<lbrakk> P \<turnstile> \<langle>e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Val v,(h\<^sub>1,l\<^sub>1,sh\<^sub>1)\<rangle>;
+ P \<turnstile> C has F,Static:t in D;
+ \<nexists>sfs. sh\<^sub>1 D = Some(sfs,Done); P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(h\<^sub>1,l\<^sub>1,sh\<^sub>1)\<rangle> \<Rightarrow> \<langle>throw a,s'\<rangle> \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D}:=e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw a,s'\<rangle>"
+
+| SFAssThrow:
+ "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>2\<rangle>
+ \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D}:=e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>2\<rangle>"
+
+| SFAssNone:
+ "\<lbrakk> P \<turnstile> \<langle>e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\<rangle>;
+ \<not>(\<exists>b t. P \<turnstile> C has F,b:t in D) \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D}:=e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>THROW NoSuchFieldError,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\<rangle>"
+
+| SFAssNonStatic:
+ "\<lbrakk> P \<turnstile> \<langle>e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\<rangle>;
+ P \<turnstile> C has F,NonStatic:t in D \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D}:=e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>THROW IncompatibleClassChangeError,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\<rangle>"
+
+| CallObjThrow:
+ "P \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle> \<Longrightarrow>
+ P \<turnstile> \<langle>e\<bullet>M(ps),s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle>"
+
+| CallParamsThrow:
+ "\<lbrakk> P \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Val v,s\<^sub>1\<rangle>; P \<turnstile> \<langle>es,s\<^sub>1\<rangle> [\<Rightarrow>] \<langle>map Val vs @ throw ex # es',s\<^sub>2\<rangle> \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>e\<bullet>M(es),s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw ex,s\<^sub>2\<rangle>"
+
+| CallNull:
+ "\<lbrakk> P \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>null,s\<^sub>1\<rangle>; P \<turnstile> \<langle>ps,s\<^sub>1\<rangle> [\<Rightarrow>] \<langle>map Val vs,s\<^sub>2\<rangle> \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>e\<bullet>M(ps),s\<^sub>0\<rangle> \<Rightarrow> \<langle>THROW NullPointer,s\<^sub>2\<rangle>"
+
+| CallNone:
+ "\<lbrakk> P \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>addr a,s\<^sub>1\<rangle>; P \<turnstile> \<langle>ps,s\<^sub>1\<rangle> [\<Rightarrow>] \<langle>map Val vs,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\<rangle>;
+ h\<^sub>2 a = Some(C,fs); \<not>(\<exists>b Ts T m D. P \<turnstile> C sees M,b:Ts\<rightarrow>T = m in D) \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>e\<bullet>M(ps),s\<^sub>0\<rangle> \<Rightarrow> \<langle>THROW NoSuchMethodError,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\<rangle>"
+
+| CallStatic:
+ "\<lbrakk> P \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>addr a,s\<^sub>1\<rangle>; P \<turnstile> \<langle>ps,s\<^sub>1\<rangle> [\<Rightarrow>] \<langle>map Val vs,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\<rangle>;
+ h\<^sub>2 a = Some(C,fs); P \<turnstile> C sees M,Static:Ts\<rightarrow>T = m in D \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>e\<bullet>M(ps),s\<^sub>0\<rangle> \<Rightarrow> \<langle>THROW IncompatibleClassChangeError,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\<rangle>"
+
+| Call:
+ "\<lbrakk> P \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>addr a,s\<^sub>1\<rangle>; P \<turnstile> \<langle>ps,s\<^sub>1\<rangle> [\<Rightarrow>] \<langle>map Val vs,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\<rangle>;
+ h\<^sub>2 a = Some(C,fs); P \<turnstile> C sees M,NonStatic:Ts\<rightarrow>T = (pns,body) in D;
+ length vs = length pns; l\<^sub>2' = [this\<mapsto>Addr a, pns[\<mapsto>]vs];
+ P \<turnstile> \<langle>body,(h\<^sub>2,l\<^sub>2',sh\<^sub>2)\<rangle> \<Rightarrow> \<langle>e',(h\<^sub>3,l\<^sub>3,sh\<^sub>3)\<rangle> \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>e\<bullet>M(ps),s\<^sub>0\<rangle> \<Rightarrow> \<langle>e',(h\<^sub>3,l\<^sub>2,sh\<^sub>3)\<rangle>"
+
+| SCallParamsThrow:
+ "\<lbrakk> P \<turnstile> \<langle>es,s\<^sub>0\<rangle> [\<Rightarrow>] \<langle>map Val vs @ throw ex # es',s\<^sub>2\<rangle> \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sM(es),s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw ex,s\<^sub>2\<rangle>"
+
+| SCallNone:
+ "\<lbrakk> P \<turnstile> \<langle>ps,s\<^sub>0\<rangle> [\<Rightarrow>] \<langle>map Val vs,s\<^sub>2\<rangle>;
+ \<not>(\<exists>b Ts T m D. P \<turnstile> C sees M,b:Ts\<rightarrow>T = m in D) \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sM(ps),s\<^sub>0\<rangle> \<Rightarrow> \<langle>THROW NoSuchMethodError,s\<^sub>2\<rangle>"
+
+| SCallNonStatic:
+ "\<lbrakk> P \<turnstile> \<langle>ps,s\<^sub>0\<rangle> [\<Rightarrow>] \<langle>map Val vs,s\<^sub>2\<rangle>;
+ P \<turnstile> C sees M,NonStatic:Ts\<rightarrow>T = m in D \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sM(ps),s\<^sub>0\<rangle> \<Rightarrow> \<langle>THROW IncompatibleClassChangeError,s\<^sub>2\<rangle>"
+
+| SCallInitThrow:
+ "\<lbrakk> P \<turnstile> \<langle>ps,s\<^sub>0\<rangle> [\<Rightarrow>] \<langle>map Val vs,(h\<^sub>1,l\<^sub>1,sh\<^sub>1)\<rangle>;
+ P \<turnstile> C sees M,Static:Ts\<rightarrow>T = (pns,body) in D;
+ \<nexists>sfs. sh\<^sub>1 D = Some(sfs,Done); M \<noteq> clinit;
+ P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(h\<^sub>1,l\<^sub>1,sh\<^sub>1)\<rangle> \<Rightarrow> \<langle>throw a,s'\<rangle> \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sM(ps),s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw a,s'\<rangle>"
+
+| SCallInit:
+ "\<lbrakk> P \<turnstile> \<langle>ps,s\<^sub>0\<rangle> [\<Rightarrow>] \<langle>map Val vs,(h\<^sub>1,l\<^sub>1,sh\<^sub>1)\<rangle>;
+ P \<turnstile> C sees M,Static:Ts\<rightarrow>T = (pns,body) in D;
+ \<nexists>sfs. sh\<^sub>1 D = Some(sfs,Done); M \<noteq> clinit;
+ P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(h\<^sub>1,l\<^sub>1,sh\<^sub>1)\<rangle> \<Rightarrow> \<langle>Val v',(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\<rangle>;
+ length vs = length pns; l\<^sub>2' = [pns[\<mapsto>]vs];
+ P \<turnstile> \<langle>body,(h\<^sub>2,l\<^sub>2',sh\<^sub>2)\<rangle> \<Rightarrow> \<langle>e',(h\<^sub>3,l\<^sub>3,sh\<^sub>3)\<rangle> \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sM(ps),s\<^sub>0\<rangle> \<Rightarrow> \<langle>e',(h\<^sub>3,l\<^sub>2,sh\<^sub>3)\<rangle>"
+
+| SCall:
+ "\<lbrakk> P \<turnstile> \<langle>ps,s\<^sub>0\<rangle> [\<Rightarrow>] \<langle>map Val vs,(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\<rangle>;
+ P \<turnstile> C sees M,Static:Ts\<rightarrow>T = (pns,body) in D;
+ sh\<^sub>2 D = Some(sfs,Done) \<or> (M = clinit \<and> sh\<^sub>2 D = Some(sfs,Processing));
+ length vs = length pns; l\<^sub>2' = [pns[\<mapsto>]vs];
+ P \<turnstile> \<langle>body,(h\<^sub>2,l\<^sub>2',sh\<^sub>2)\<rangle> \<Rightarrow> \<langle>e',(h\<^sub>3,l\<^sub>3,sh\<^sub>3)\<rangle> \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sM(ps),s\<^sub>0\<rangle> \<Rightarrow> \<langle>e',(h\<^sub>3,l\<^sub>2,sh\<^sub>3)\<rangle>"
+
+| Block:
+ "P \<turnstile> \<langle>e\<^sub>0,(h\<^sub>0,l\<^sub>0(V:=None),sh\<^sub>0)\<rangle> \<Rightarrow> \<langle>e\<^sub>1,(h\<^sub>1,l\<^sub>1,sh\<^sub>1)\<rangle> \<Longrightarrow>
+ P \<turnstile> \<langle>{V:T; e\<^sub>0},(h\<^sub>0,l\<^sub>0,sh\<^sub>0)\<rangle> \<Rightarrow> \<langle>e\<^sub>1,(h\<^sub>1,l\<^sub>1(V:=l\<^sub>0 V),sh\<^sub>1)\<rangle>"
+
+| Seq:
+ "\<lbrakk> P \<turnstile> \<langle>e\<^sub>0,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Val v,s\<^sub>1\<rangle>; P \<turnstile> \<langle>e\<^sub>1,s\<^sub>1\<rangle> \<Rightarrow> \<langle>e\<^sub>2,s\<^sub>2\<rangle> \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>e\<^sub>0;;e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow> \<langle>e\<^sub>2,s\<^sub>2\<rangle>"
+
+| SeqThrow:
+ "P \<turnstile> \<langle>e\<^sub>0,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e,s\<^sub>1\<rangle> \<Longrightarrow>
+ P \<turnstile> \<langle>e\<^sub>0;;e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e,s\<^sub>1\<rangle>"
+
+| CondT:
+ "\<lbrakk> P \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>true,s\<^sub>1\<rangle>; P \<turnstile> \<langle>e\<^sub>1,s\<^sub>1\<rangle> \<Rightarrow> \<langle>e',s\<^sub>2\<rangle> \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>if (e) e\<^sub>1 else e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>e',s\<^sub>2\<rangle>"
+
+| CondF:
+ "\<lbrakk> P \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>false,s\<^sub>1\<rangle>; P \<turnstile> \<langle>e\<^sub>2,s\<^sub>1\<rangle> \<Rightarrow> \<langle>e',s\<^sub>2\<rangle> \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>if (e) e\<^sub>1 else e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>e',s\<^sub>2\<rangle>"
+
+| CondThrow:
+ "P \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle> \<Longrightarrow>
+ P \<turnstile> \<langle>if (e) e\<^sub>1 else e\<^sub>2, s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle>"
+
+| WhileF:
+ "P \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>false,s\<^sub>1\<rangle> \<Longrightarrow>
+ P \<turnstile> \<langle>while (e) c,s\<^sub>0\<rangle> \<Rightarrow> \<langle>unit,s\<^sub>1\<rangle>"
+
+| WhileT:
+ "\<lbrakk> P \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>true,s\<^sub>1\<rangle>; P \<turnstile> \<langle>c,s\<^sub>1\<rangle> \<Rightarrow> \<langle>Val v\<^sub>1,s\<^sub>2\<rangle>; P \<turnstile> \<langle>while (e) c,s\<^sub>2\<rangle> \<Rightarrow> \<langle>e\<^sub>3,s\<^sub>3\<rangle> \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>while (e) c,s\<^sub>0\<rangle> \<Rightarrow> \<langle>e\<^sub>3,s\<^sub>3\<rangle>"
+
+| WhileCondThrow:
+ "P \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle> \<Longrightarrow>
+ P \<turnstile> \<langle>while (e) c,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle>"
+
+| WhileBodyThrow:
+ "\<lbrakk> P \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>true,s\<^sub>1\<rangle>; P \<turnstile> \<langle>c,s\<^sub>1\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>2\<rangle>\<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>while (e) c,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>2\<rangle>"
+
+| Throw:
+ "P \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>addr a,s\<^sub>1\<rangle> \<Longrightarrow>
+ P \<turnstile> \<langle>throw e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Throw a,s\<^sub>1\<rangle>"
+
+| ThrowNull:
+ "P \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>null,s\<^sub>1\<rangle> \<Longrightarrow>
+ P \<turnstile> \<langle>throw e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>THROW NullPointer,s\<^sub>1\<rangle>"
+
+| ThrowThrow:
+ "P \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle> \<Longrightarrow>
+ P \<turnstile> \<langle>throw e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e',s\<^sub>1\<rangle>"
+
+| Try:
+ "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Val v\<^sub>1,s\<^sub>1\<rangle> \<Longrightarrow>
+ P \<turnstile> \<langle>try e\<^sub>1 catch(C V) e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Val v\<^sub>1,s\<^sub>1\<rangle>"
+
+| TryCatch:
+ "\<lbrakk> P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Throw a,(h\<^sub>1,l\<^sub>1,sh\<^sub>1)\<rangle>; h\<^sub>1 a = Some(D,fs); P \<turnstile> D \<preceq>\<^sup>* C;
+ P \<turnstile> \<langle>e\<^sub>2,(h\<^sub>1,l\<^sub>1(V\<mapsto>Addr a),sh\<^sub>1)\<rangle> \<Rightarrow> \<langle>e\<^sub>2',(h\<^sub>2,l\<^sub>2,sh\<^sub>2)\<rangle> \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>try e\<^sub>1 catch(C V) e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>e\<^sub>2',(h\<^sub>2,l\<^sub>2(V:=l\<^sub>1 V),sh\<^sub>2)\<rangle>"
+
+| TryThrow:
+ "\<lbrakk> P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Throw a,(h\<^sub>1,l\<^sub>1,sh\<^sub>1)\<rangle>; h\<^sub>1 a = Some(D,fs); \<not> P \<turnstile> D \<preceq>\<^sup>* C \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>try e\<^sub>1 catch(C V) e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Throw a,(h\<^sub>1,l\<^sub>1,sh\<^sub>1)\<rangle>"
+
+| Nil:
+ "P \<turnstile> \<langle>[],s\<rangle> [\<Rightarrow>] \<langle>[],s\<rangle>"
+
+| Cons:
+ "\<lbrakk> P \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Val v,s\<^sub>1\<rangle>; P \<turnstile> \<langle>es,s\<^sub>1\<rangle> [\<Rightarrow>] \<langle>es',s\<^sub>2\<rangle> \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>e#es,s\<^sub>0\<rangle> [\<Rightarrow>] \<langle>Val v # es',s\<^sub>2\<rangle>"
+
+| ConsThrow:
+ "P \<turnstile> \<langle>e, s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw e', s\<^sub>1\<rangle> \<Longrightarrow>
+ P \<turnstile> \<langle>e#es, s\<^sub>0\<rangle> [\<Rightarrow>] \<langle>throw e' # es, s\<^sub>1\<rangle>"
+
+\<comment> \<open> init rules \<close>
+
+| InitFinal:
+ "P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>INIT C (Nil,b) \<leftarrow> e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+
+| InitNone:
+ "\<lbrakk> sh C = None; P \<turnstile> \<langle>INIT C' (C#Cs,False) \<leftarrow> e,(h,l,sh(C \<mapsto> (sblank P C, Prepared)))\<rangle> \<Rightarrow> \<langle>e',s'\<rangle> \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>INIT C' (C#Cs,False) \<leftarrow> e,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+
+| InitDone:
+ "\<lbrakk> sh C = Some(sfs,Done); P \<turnstile> \<langle>INIT C' (Cs,True) \<leftarrow> e,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle> \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>INIT C' (C#Cs,False) \<leftarrow> e,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+
+| InitProcessing:
+ "\<lbrakk> sh C = Some(sfs,Processing); P \<turnstile> \<langle>INIT C' (Cs,True) \<leftarrow> e,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle> \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>INIT C' (C#Cs,False) \<leftarrow> e,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+
+\<comment> \<open> note that @{text RI} will mark all classes in the list Cs with the Error flag \<close>
+| InitError:
+ "\<lbrakk> sh C = Some(sfs,Error);
+ P \<turnstile> \<langle>RI (C, THROW NoClassDefFoundError);Cs \<leftarrow> e,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle> \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>INIT C' (C#Cs,False) \<leftarrow> e,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+
+| InitObject:
+ "\<lbrakk> sh C = Some(sfs,Prepared);
+ C = Object;
+ sh' = sh(C \<mapsto> (sfs,Processing));
+ P \<turnstile> \<langle>INIT C' (C#Cs,True) \<leftarrow> e,(h,l,sh')\<rangle> \<Rightarrow> \<langle>e',s'\<rangle> \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>INIT C' (C#Cs,False) \<leftarrow> e,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+
+| InitNonObject:
+ "\<lbrakk> sh C = Some(sfs,Prepared);
+ C \<noteq> Object;
+ class P C = Some (D,r);
+ sh' = sh(C \<mapsto> (sfs,Processing));
+ P \<turnstile> \<langle>INIT C' (D#C#Cs,False) \<leftarrow> e,(h,l,sh')\<rangle> \<Rightarrow> \<langle>e',s'\<rangle> \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>INIT C' (C#Cs,False) \<leftarrow> e,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+
+| InitRInit:
+ "P \<turnstile> \<langle>RI (C,C\<bullet>\<^sub>sclinit([]));Cs \<leftarrow> e,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>
+ \<Longrightarrow> P \<turnstile> \<langle>INIT C' (C#Cs,True) \<leftarrow> e,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+
+| RInit:
+ "\<lbrakk> P \<turnstile> \<langle>e',s\<rangle> \<Rightarrow> \<langle>Val v, (h',l',sh')\<rangle>;
+ sh' C = Some(sfs, i); sh'' = sh'(C \<mapsto> (sfs, Done));
+ C' = last(C#Cs);
+ P \<turnstile> \<langle>INIT C' (Cs,True) \<leftarrow> e, (h',l',sh'')\<rangle> \<Rightarrow> \<langle>e\<^sub>1,s\<^sub>1\<rangle> \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>RI (C,e');Cs \<leftarrow> e,s\<rangle> \<Rightarrow> \<langle>e\<^sub>1,s\<^sub>1\<rangle>"
+
+| RInitInitFail:
+ "\<lbrakk> P \<turnstile> \<langle>e',s\<rangle> \<Rightarrow> \<langle>throw a, (h',l',sh')\<rangle>;
+ sh' C = Some(sfs, i); sh'' = sh'(C \<mapsto> (sfs, Error));
+ P \<turnstile> \<langle>RI (D,throw a);Cs \<leftarrow> e, (h',l',sh'')\<rangle> \<Rightarrow> \<langle>e\<^sub>1,s\<^sub>1\<rangle> \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>RI (C,e');D#Cs \<leftarrow> e,s\<rangle> \<Rightarrow> \<langle>e\<^sub>1,s\<^sub>1\<rangle>"
+
+| RInitFailFinal:
+ "\<lbrakk> P \<turnstile> \<langle>e',s\<rangle> \<Rightarrow> \<langle>throw a, (h',l',sh')\<rangle>;
+ sh' C = Some(sfs, i); sh'' = sh'(C \<mapsto> (sfs, Error)) \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>RI (C,e');Nil \<leftarrow> e,s\<rangle> \<Rightarrow> \<langle>throw a, (h',l',sh'')\<rangle>"
+
+(*<*)
+lemmas eval_evals_induct = eval_evals.induct [split_format (complete)]
+ and eval_evals_inducts = eval_evals.inducts [split_format (complete)]
+
+inductive_cases eval_cases [cases set]:
+ "P \<turnstile> \<langle>new C,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+ "P \<turnstile> \<langle>Cast C e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+ "P \<turnstile> \<langle>Val v,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+ "P \<turnstile> \<langle>e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+ "P \<turnstile> \<langle>Var v,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+ "P \<turnstile> \<langle>V:=e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+ "P \<turnstile> \<langle>e\<bullet>F{D},s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+ "P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D},s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+ "P \<turnstile> \<langle>e\<^sub>1\<bullet>F{D}:=e\<^sub>2,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+ "P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D}:=e\<^sub>2,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+ "P \<turnstile> \<langle>e\<bullet>M(es),s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+ "P \<turnstile> \<langle>C\<bullet>\<^sub>sM(es),s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+ "P \<turnstile> \<langle>{V:T;e\<^sub>1},s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+ "P \<turnstile> \<langle>e\<^sub>1;;e\<^sub>2,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+ "P \<turnstile> \<langle>if (e) e\<^sub>1 else e\<^sub>2,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+ "P \<turnstile> \<langle>while (b) c,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+ "P \<turnstile> \<langle>throw e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+ "P \<turnstile> \<langle>try e\<^sub>1 catch(C V) e\<^sub>2,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+ "P \<turnstile> \<langle>INIT C (Cs,b) \<leftarrow> e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+ "P \<turnstile> \<langle>RI (C,e);Cs \<leftarrow> e\<^sub>0,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+
+inductive_cases evals_cases [cases set]:
+ "P \<turnstile> \<langle>[],s\<rangle> [\<Rightarrow>] \<langle>e',s'\<rangle>"
+ "P \<turnstile> \<langle>e#es,s\<rangle> [\<Rightarrow>] \<langle>e',s'\<rangle>"
+(*>*)
+
+subsection "Final expressions"
+
+lemma eval_final: "P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle> \<Longrightarrow> final e'"
+ and evals_final: "P \<turnstile> \<langle>es,s\<rangle> [\<Rightarrow>] \<langle>es',s'\<rangle> \<Longrightarrow> finals es'"
+(*<*)by(induct rule:eval_evals.inducts, simp_all)(*>*)
+
+text\<open> Only used later, in the small to big translation, but is already a
+good sanity check: \<close>
+
+lemma eval_finalId: "final e \<Longrightarrow> P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>e,s\<rangle>"
+(*<*)by (erule finalE) (iprover intro: eval_evals.intros)+(*>*)
+
+lemma eval_final_same: "\<lbrakk> P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>; final e \<rbrakk> \<Longrightarrow> e = e' \<and> s = s'"
+(*<*)by(auto elim!: finalE eval_cases)(*>*)
+
+lemma eval_finalsId:
+assumes finals: "finals es" shows "P \<turnstile> \<langle>es,s\<rangle> [\<Rightarrow>] \<langle>es,s\<rangle>"
+(*<*)
+ using finals
+proof (induct es type: list)
+ case Nil show ?case by (rule eval_evals.intros)
+next
+ case (Cons e es)
+ have hyp: "finals es \<Longrightarrow> P \<turnstile> \<langle>es,s\<rangle> [\<Rightarrow>] \<langle>es,s\<rangle>"
+ and finals: "finals (e # es)" by fact+
+ show "P \<turnstile> \<langle>e # es,s\<rangle> [\<Rightarrow>] \<langle>e # es,s\<rangle>"
+ proof cases
+ assume "final e"
+ thus ?thesis
+ proof (cases rule: finalE)
+ fix v assume e: "e = Val v"
+ have "P \<turnstile> \<langle>Val v,s\<rangle> \<Rightarrow> \<langle>Val v,s\<rangle>" by (simp add: eval_finalId)
+ moreover from finals e have "P \<turnstile> \<langle>es,s\<rangle> [\<Rightarrow>] \<langle>es,s\<rangle>" by(fast intro:hyp)
+ ultimately have "P \<turnstile> \<langle>Val v#es,s\<rangle> [\<Rightarrow>] \<langle>Val v#es,s\<rangle>"
+ by (rule eval_evals.intros)
+ with e show ?thesis by simp
+ next
+ fix a assume e: "e = Throw a"
+ have "P \<turnstile> \<langle>Throw a,s\<rangle> \<Rightarrow> \<langle>Throw a,s\<rangle>" by (simp add: eval_finalId)
+ hence "P \<turnstile> \<langle>Throw a#es,s\<rangle> [\<Rightarrow>] \<langle>Throw a#es,s\<rangle>" by (rule eval_evals.intros)
+ with e show ?thesis by simp
+ qed
+ next
+ assume "\<not> final e"
+ with not_finals_ConsI finals have False by blast
+ thus ?thesis ..
+ qed
+qed
+(*>*)
+
+lemma evals_finals_same:
+assumes finals: "finals es"
+shows "P \<turnstile> \<langle>es,s\<rangle> [\<Rightarrow>] \<langle>es',s'\<rangle> \<Longrightarrow> es = es' \<and> s = s'"
+ using finals
+proof (induct es arbitrary: es' type: list)
+ case Nil then show ?case using evals_cases(1) by blast
+next
+ case (Cons e es)
+ have IH: "\<And>es'. P \<turnstile> \<langle>es,s\<rangle> [\<Rightarrow>] \<langle>es',s'\<rangle> \<Longrightarrow> finals es \<Longrightarrow> es = es' \<and> s = s'"
+ and step: "P \<turnstile> \<langle>e # es,s\<rangle> [\<Rightarrow>] \<langle>es',s'\<rangle>" and finals: "finals (e # es)" by fact+
+ then obtain e' es'' where es': "es' = e'#es''" by (meson Cons.prems(1) evals_cases(2))
+ have fe: "final e" using finals not_finals_ConsI by auto
+ show ?case
+ proof(rule evals_cases(2)[OF step])
+ fix v s\<^sub>1 es1 assume es1: "es' = Val v # es1"
+ and estep: "P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>Val v,s\<^sub>1\<rangle>" and esstep: "P \<turnstile> \<langle>es,s\<^sub>1\<rangle> [\<Rightarrow>] \<langle>es1,s'\<rangle>"
+ then have "e = Val v" using eval_final_same fe by auto
+ then have "finals es" using es' not_finals_ConsI2 finals by blast
+ then show ?thesis using es' IH estep esstep es1 eval_final_same fe by blast
+ next
+ fix e' assume es1: "es' = throw e' # es" and estep: "P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>throw e',s'\<rangle>"
+ then have "e = throw e'" using eval_final_same fe by auto
+ then show ?thesis using es' estep es1 eval_final_same fe by blast
+ qed
+qed
+(*>*)
+
+subsection "Property preservation"
+
+lemma evals_length: "P \<turnstile> \<langle>es,s\<rangle> [\<Rightarrow>] \<langle>es',s'\<rangle> \<Longrightarrow> length es = length es'"
+ by(induct es arbitrary:es' s s', auto elim: evals_cases)
+
+corollary evals_empty: "P \<turnstile> \<langle>es,s\<rangle> [\<Rightarrow>] \<langle>es',s'\<rangle> \<Longrightarrow> (es = []) = (es' = [])"
+ by(drule evals_length, fastforce)
+
+(****)
+
+theorem eval_hext: "P \<turnstile> \<langle>e,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',(h',l',sh')\<rangle> \<Longrightarrow> h \<unlhd> h'"
+and evals_hext: "P \<turnstile> \<langle>es,(h,l,sh)\<rangle> [\<Rightarrow>] \<langle>es',(h',l',sh')\<rangle> \<Longrightarrow> h \<unlhd> h'"
+(*<*)
+proof (induct rule: eval_evals_inducts)
+ case New thus ?case
+ by(fastforce intro!: hext_new intro:LeastI simp:new_Addr_def
+ split:if_split_asm simp del:fun_upd_apply)
+next
+ case NewInit thus ?case
+ by (meson hext_new hext_trans new_Addr_SomeD)
+next
+ case FAss thus ?case
+ by(auto simp:sym[THEN hext_upd_obj] simp del:fun_upd_apply
+ elim!: hext_trans)
+qed (auto elim!: hext_trans)
+(*>*)
+
+
+lemma eval_lcl_incr: "P \<turnstile> \<langle>e,(h\<^sub>0,l\<^sub>0,sh\<^sub>0)\<rangle> \<Rightarrow> \<langle>e',(h\<^sub>1,l\<^sub>1,sh\<^sub>1)\<rangle> \<Longrightarrow> dom l\<^sub>0 \<subseteq> dom l\<^sub>1"
+ and evals_lcl_incr: "P \<turnstile> \<langle>es,(h\<^sub>0,l\<^sub>0,sh\<^sub>0)\<rangle> [\<Rightarrow>] \<langle>es',(h\<^sub>1,l\<^sub>1,sh\<^sub>1)\<rangle> \<Longrightarrow> dom l\<^sub>0 \<subseteq> dom l\<^sub>1"
+(*<*)
+proof (induct rule: eval_evals_inducts)
+ case BinOp show ?case by(rule subset_trans)(rule BinOp.hyps)+
+next
+ case Call thus ?case
+ by(simp del: fun_upd_apply)
+next
+ case Seq show ?case by(rule subset_trans)(rule Seq.hyps)+
+next
+ case CondT show ?case by(rule subset_trans)(rule CondT.hyps)+
+next
+ case CondF show ?case by(rule subset_trans)(rule CondF.hyps)+
+next
+ case WhileT thus ?case by(blast)
+next
+ case TryCatch thus ?case by(clarsimp simp:dom_def split:if_split_asm) blast
+next
+ case Cons show ?case by(rule subset_trans)(rule Cons.hyps)+
+next
+ case Block thus ?case by(auto simp del:fun_upd_apply)
+qed auto
+(*>*)
+
+lemma
+shows init_ri_same_loc: "P \<turnstile> \<langle>e,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',(h',l',sh')\<rangle>
+ \<Longrightarrow> (\<And>C Cs b M a. e = INIT C (Cs,b) \<leftarrow> unit \<or> e = C\<bullet>\<^sub>sM([]) \<or> e = RI (C,Throw a) ; Cs \<leftarrow> unit
+ \<or> e = RI (C,C\<bullet>\<^sub>sclinit([])) ; Cs \<leftarrow> unit
+ \<Longrightarrow> l = l')"
+ and "P \<turnstile> \<langle>es,(h,l,sh)\<rangle> [\<Rightarrow>] \<langle>es',(h',l',sh')\<rangle> \<Longrightarrow> True"
+proof(induct rule: eval_evals_inducts)
+ case (RInitInitFail e h l sh a')
+ then show ?case using eval_final[of _ _ _ "throw a'"]
+ by(fastforce dest: eval_final_same[of _ "Throw a"])
+next
+ case RInitFailFinal then show ?case by(auto dest: eval_final_same)
+qed(auto dest: evals_cases eval_cases(17) eval_final_same)
+
+lemma init_same_loc: "P \<turnstile> \<langle>INIT C (Cs,b) \<leftarrow> unit,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',(h',l',sh')\<rangle> \<Longrightarrow> l = l'"
+ by(simp add: init_ri_same_loc)
+
+(****)
+
+lemma assumes wf: "wwf_J_prog P"
+shows eval_proc_pres': "P \<turnstile> \<langle>e,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',(h',l',sh')\<rangle>
+ \<Longrightarrow> not_init C e \<Longrightarrow> \<exists>sfs. sh C = \<lfloor>(sfs, Processing)\<rfloor> \<Longrightarrow> \<exists>sfs'. sh' C = \<lfloor>(sfs', Processing)\<rfloor>"
+ and evals_proc_pres': "P \<turnstile> \<langle>es,(h,l,sh)\<rangle> [\<Rightarrow>] \<langle>es',(h',l',sh')\<rangle>
+ \<Longrightarrow> not_inits C es \<Longrightarrow> \<exists>sfs. sh C = \<lfloor>(sfs, Processing)\<rfloor> \<Longrightarrow> \<exists>sfs'. sh' C = \<lfloor>(sfs', Processing)\<rfloor>"
+(*<*)
+proof(induct rule:eval_evals_inducts)
+ case Call then show ?case using sees_wwf_nsub_RI[OF wf Call.hyps(6)] nsub_RI_not_init by auto
+next
+ case (SCallInit ps h l sh vs h\<^sub>1 l\<^sub>1 sh\<^sub>1 C' M Ts T pns body D v' h\<^sub>2 l\<^sub>2 sh\<^sub>2 l\<^sub>2' e' h\<^sub>3 l\<^sub>3 sh\<^sub>3)
+ then show ?case
+ using SCallInit sees_wwf_nsub_RI[OF wf SCallInit.hyps(3)] nsub_RI_not_init[of body] by auto
+next
+ case SCall then show ?case using sees_wwf_nsub_RI[OF wf SCall.hyps(3)] nsub_RI_not_init by auto
+next
+ case (InitNone sh C1 C' Cs h l e' a a b) then show ?case by(cases "C = C1") auto
+next
+ case (InitDone sh C sfs C' Cs h l e' a a b) then show ?case by(cases Cs, auto)
+next
+ case (InitProcessing sh C sfs C' Cs h l e' a a b) then show ?case by(cases Cs, auto)
+next
+ case (InitError sh C1 sfs Cs h l e' a a b C') then show ?case by(cases "C = C1") auto
+next
+ case (InitObject sh C1 sfs sh' C' Cs h l e' a a b) then show ?case by(cases "C = C1") auto
+next
+ case (InitNonObject sh C1 sfs D a b sh' C' Cs h l e' a a b)
+ then show ?case by(cases "C = C1") auto
+next
+ case (RInit e a a b v h' l' sh' C sfs i sh'' C' Cs e\<^sub>1 a a b) then show ?case by(cases Cs, auto)
+next
+ case (RInitInitFail e h l sh a h' l' sh' C1 sfs i sh'' D Cs e\<^sub>1 h1 l1 sh1)
+ then show ?case using eval_final by fastforce
+qed(auto)
+
+(************************************************)
+
+subsection\<open>Init Elimination rules\<close>
+
+lemma init_NilE:
+assumes init: "P \<turnstile> \<langle>INIT C (Nil,b) \<leftarrow> unit,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+shows "e' = unit \<and> s' = s"
+proof(rule eval_cases(19)[OF init]) \<comment> \<open> Init \<close> qed(auto dest: eval_final_same)
+
+lemma init_ProcessingE:
+assumes shC: "sh C = \<lfloor>(sfs, Processing)\<rfloor>"
+ and init: "P \<turnstile> \<langle>INIT C ([C],False) \<leftarrow> unit,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+shows "e' = unit \<and> s' = (h,l,sh)"
+proof(rule eval_cases(19)[OF init]) \<comment> \<open> Init \<close>
+ fix sha Ca sfs Cs ha la
+ assume "(h, l, sh) = (ha, la, sha)" and "sha Ca = \<lfloor>(sfs, Processing)\<rfloor>"
+ and "P \<turnstile> \<langle>INIT C (Cs,True) \<leftarrow> unit,(ha, la, sha)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>" and "[C] = Ca # Cs"
+ then show ?thesis using init_NilE by simp
+next
+ fix sha sfs Cs ha la
+ assume "(h, l, sh) = (ha, la, sha)" and "sha Object = \<lfloor>(sfs, Prepared)\<rfloor>"
+ and "[C] = Object # Cs"
+ then show ?thesis using shC by clarsimp
+qed(auto simp: assms)
+
+
+lemma rinit_throwE:
+"P \<turnstile> \<langle>RI (C,throw e) ; Cs \<leftarrow> e\<^sub>0,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>
+ \<Longrightarrow> \<exists>a s\<^sub>t. e' = throw a \<and> P \<turnstile> \<langle>throw e,s\<rangle> \<Rightarrow> \<langle>throw a,s\<^sub>t\<rangle>"
+proof(induct Cs arbitrary: C e s)
+ case Nil
+ then show ?case
+ proof(rule eval_cases(20)) \<comment> \<open> RI \<close>
+ fix v h' l' sh' assume "P \<turnstile> \<langle>throw e,s\<rangle> \<Rightarrow> \<langle>Val v,(h', l', sh')\<rangle>"
+ then show ?case using eval_cases(17) by blast
+ qed(auto)
+next
+ case (Cons C' Cs')
+ show ?case using Cons.prems(1)
+ proof(rule eval_cases(20)) \<comment> \<open> RI \<close>
+ fix v h' l' sh' assume "P \<turnstile> \<langle>throw e,s\<rangle> \<Rightarrow> \<langle>Val v,(h', l', sh')\<rangle>"
+ then show ?case using eval_cases(17) by blast
+ next
+ fix a h' l' sh' sfs i D Cs''
+ assume e''step: "P \<turnstile> \<langle>throw e,s\<rangle> \<Rightarrow> \<langle>throw a,(h', l', sh')\<rangle>"
+ and shC: "sh' C = \<lfloor>(sfs, i)\<rfloor>"
+ and riD: "P \<turnstile> \<langle>RI (D,throw a) ; Cs'' \<leftarrow> e\<^sub>0,(h', l', sh'(C \<mapsto> (sfs, Error)))\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+ and "C' # Cs' = D # Cs''"
+ then show ?thesis using Cons.hyps eval_final eval_final_same by blast
+ qed(simp)
+qed
+
+lemma rinit_ValE:
+assumes ri: "P \<turnstile> \<langle>RI (C,e) ; Cs \<leftarrow> unit,s\<rangle> \<Rightarrow> \<langle>Val v',s'\<rangle>"
+shows "\<exists>v'' s''. P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>Val v'',s''\<rangle>"
+proof(rule eval_cases(20)[OF ri]) \<comment> \<open> RI \<close>
+ fix a h' l' sh' sfs i D Cs'
+ assume "P \<turnstile> \<langle>RI (D,throw a) ; Cs' \<leftarrow> unit,(h', l', sh'(C \<mapsto> (sfs, Error)))\<rangle> \<Rightarrow> \<langle>Val v',s'\<rangle>"
+ then show ?thesis using rinit_throwE by blast
+qed(auto)
+
+subsection "Some specific evaluations"
+
+lemma lass_val_of_eval:
+ "\<lbrakk> lass_val_of e = \<lfloor>a\<rfloor>; P \<turnstile> \<langle>e,(h, l, sh)\<rangle> \<Rightarrow> \<langle>e',(h', l', sh')\<rangle> \<rbrakk>
+ \<Longrightarrow> e' = unit \<and> h' = h \<and> l' = l(fst a\<mapsto>snd a) \<and> sh' = sh"
+ by(drule lass_val_of_spec, auto elim: eval.cases)
+
+lemma eval_throw_nonVal:
+assumes eval: "P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>throw a,s'\<rangle>"
+shows "val_of e = None"
+proof(cases "val_of e")
+ case (Some v) show ?thesis using eval by(auto simp: val_of_spec[OF Some] intro: eval.cases)
+qed(simp)
+
+lemma eval_throw_nonLAss:
+assumes eval: "P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>throw a,s'\<rangle>"
+shows "lass_val_of e = None"
+proof(cases "lass_val_of e")
+ case (Some v) show ?thesis using eval by(auto simp: lass_val_of_spec[OF Some] intro: eval.cases)
+qed(simp)
+
+end
diff --git a/thys/JinjaDCI/J/DefAss.thy b/thys/JinjaDCI/J/DefAss.thy
--- a/thys/JinjaDCI/J/DefAss.thy
+++ b/thys/JinjaDCI/J/DefAss.thy
@@ -1,192 +1,192 @@
-(* Title: JinjaDCI/J/DefAss.thy
- Author: Tobias Nipkow, Susannah Mansky
- Copyright 2003 Technische Universitaet Muenchen, 2019-20 UIUC
-
- Based on the Jinja theory J/DefAss.thy by Tobias Nipkow
-*)
-
-section \<open> Definite assignment \<close>
-
-theory DefAss imports BigStep begin
-
-subsection "Hypersets"
-
-type_synonym 'a hyperset = "'a set option"
-
-definition hyperUn :: "'a hyperset \<Rightarrow> 'a hyperset \<Rightarrow> 'a hyperset" (infixl "\<squnion>" 65)
-where
- "A \<squnion> B \<equiv> case A of None \<Rightarrow> None
- | \<lfloor>A\<rfloor> \<Rightarrow> (case B of None \<Rightarrow> None | \<lfloor>B\<rfloor> \<Rightarrow> \<lfloor>A \<union> B\<rfloor>)"
-
-definition hyperInt :: "'a hyperset \<Rightarrow> 'a hyperset \<Rightarrow> 'a hyperset" (infixl "\<sqinter>" 70)
-where
- "A \<sqinter> B \<equiv> case A of None \<Rightarrow> B
- | \<lfloor>A\<rfloor> \<Rightarrow> (case B of None \<Rightarrow> \<lfloor>A\<rfloor> | \<lfloor>B\<rfloor> \<Rightarrow> \<lfloor>A \<inter> B\<rfloor>)"
-
-definition hyperDiff1 :: "'a hyperset \<Rightarrow> 'a \<Rightarrow> 'a hyperset" (infixl "\<ominus>" 65)
-where
- "A \<ominus> a \<equiv> case A of None \<Rightarrow> None | \<lfloor>A\<rfloor> \<Rightarrow> \<lfloor>A - {a}\<rfloor>"
-
-definition hyper_isin :: "'a \<Rightarrow> 'a hyperset \<Rightarrow> bool" (infix "\<in>\<in>" 50)
-where
- "a \<in>\<in> A \<equiv> case A of None \<Rightarrow> True | \<lfloor>A\<rfloor> \<Rightarrow> a \<in> A"
-
-definition hyper_subset :: "'a hyperset \<Rightarrow> 'a hyperset \<Rightarrow> bool" (infix "\<sqsubseteq>" 50)
-where
- "A \<sqsubseteq> B \<equiv> case B of None \<Rightarrow> True
- | \<lfloor>B\<rfloor> \<Rightarrow> (case A of None \<Rightarrow> False | \<lfloor>A\<rfloor> \<Rightarrow> A \<subseteq> B)"
-
-lemmas hyperset_defs =
- hyperUn_def hyperInt_def hyperDiff1_def hyper_isin_def hyper_subset_def
-
-lemma [simp]: "\<lfloor>{}\<rfloor> \<squnion> A = A \<and> A \<squnion> \<lfloor>{}\<rfloor> = A"
-(*<*)by(simp add:hyperset_defs)(*>*)
-
-lemma [simp]: "\<lfloor>A\<rfloor> \<squnion> \<lfloor>B\<rfloor> = \<lfloor>A \<union> B\<rfloor> \<and> \<lfloor>A\<rfloor> \<ominus> a = \<lfloor>A - {a}\<rfloor>"
-(*<*)by(simp add:hyperset_defs)(*>*)
-
-lemma [simp]: "None \<squnion> A = None \<and> A \<squnion> None = None"
-(*<*)by(simp add:hyperset_defs)(*>*)
-
-lemma [simp]: "a \<in>\<in> None \<and> None \<ominus> a = None"
-(*<*)by(simp add:hyperset_defs)(*>*)
-
-lemma hyper_isin_union: "x \<in>\<in> \<lfloor>A\<rfloor> \<Longrightarrow> x \<in>\<in> \<lfloor>A\<rfloor> \<squnion> B"
- by(case_tac B, auto simp: hyper_isin_def)
-
-lemma hyperUn_assoc: "(A \<squnion> B) \<squnion> C = A \<squnion> (B \<squnion> C)"
-(*<*)by(simp add:hyperset_defs Un_assoc)(*>*)
-
-lemma hyper_insert_comm: "A \<squnion> \<lfloor>{a}\<rfloor> = \<lfloor>{a}\<rfloor> \<squnion> A \<and> A \<squnion> (\<lfloor>{a}\<rfloor> \<squnion> B) = \<lfloor>{a}\<rfloor> \<squnion> (A \<squnion> B)"
-(*<*)by(simp add:hyperset_defs)(*>*)
-
-lemma hyper_comm: "A \<squnion> B = B \<squnion> A \<and> A \<squnion> B \<squnion> C = B \<squnion> A \<squnion> C"
-(*<*)
-proof(cases A)
- case SomeA: Some then show ?thesis
- proof(cases B)
- case SomeB: Some with SomeA show ?thesis
- proof(cases C)
- case SomeC: Some with SomeA SomeB show ?thesis
- by(simp add: Un_left_commute Un_commute)
- qed (simp add: Un_commute)
- qed simp
-qed simp
-(*>*)
-
-subsection "Definite assignment"
-
-primrec
- \<A> :: "'a exp \<Rightarrow> 'a hyperset"
- and \<A>s :: "'a exp list \<Rightarrow> 'a hyperset"
-where
- "\<A> (new C) = \<lfloor>{}\<rfloor>"
-| "\<A> (Cast C e) = \<A> e"
-| "\<A> (Val v) = \<lfloor>{}\<rfloor>"
-| "\<A> (e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2) = \<A> e\<^sub>1 \<squnion> \<A> e\<^sub>2"
-| "\<A> (Var V) = \<lfloor>{}\<rfloor>"
-| "\<A> (LAss V e) = \<lfloor>{V}\<rfloor> \<squnion> \<A> e"
-| "\<A> (e\<bullet>F{D}) = \<A> e"
-| "\<A> (C\<bullet>\<^sub>sF{D}) = \<lfloor>{}\<rfloor>"
-| "\<A> (e\<^sub>1\<bullet>F{D}:=e\<^sub>2) = \<A> e\<^sub>1 \<squnion> \<A> e\<^sub>2"
-| "\<A> (C\<bullet>\<^sub>sF{D}:=e\<^sub>2) = \<A> e\<^sub>2"
-| "\<A> (e\<bullet>M(es)) = \<A> e \<squnion> \<A>s es"
-| "\<A> (C\<bullet>\<^sub>sM(es)) = \<A>s es"
-| "\<A> ({V:T; e}) = \<A> e \<ominus> V"
-| "\<A> (e\<^sub>1;;e\<^sub>2) = \<A> e\<^sub>1 \<squnion> \<A> e\<^sub>2"
-| "\<A> (if (e) e\<^sub>1 else e\<^sub>2) = \<A> e \<squnion> (\<A> e\<^sub>1 \<sqinter> \<A> e\<^sub>2)"
-| "\<A> (while (b) e) = \<A> b"
-| "\<A> (throw e) = None"
-| "\<A> (try e\<^sub>1 catch(C V) e\<^sub>2) = \<A> e\<^sub>1 \<sqinter> (\<A> e\<^sub>2 \<ominus> V)"
-| "\<A> (INIT C (Cs,b) \<leftarrow> e) = \<lfloor>{}\<rfloor>"
-| "\<A> (RI (C,e);Cs \<leftarrow> e') = \<A> e"
-
-| "\<A>s ([]) = \<lfloor>{}\<rfloor>"
-| "\<A>s (e#es) = \<A> e \<squnion> \<A>s es"
-
-primrec
- \<D> :: "'a exp \<Rightarrow> 'a hyperset \<Rightarrow> bool"
- and \<D>s :: "'a exp list \<Rightarrow> 'a hyperset \<Rightarrow> bool"
-where
- "\<D> (new C) A = True"
-| "\<D> (Cast C e) A = \<D> e A"
-| "\<D> (Val v) A = True"
-| "\<D> (e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2) A = (\<D> e\<^sub>1 A \<and> \<D> e\<^sub>2 (A \<squnion> \<A> e\<^sub>1))"
-| "\<D> (Var V) A = (V \<in>\<in> A)"
-| "\<D> (LAss V e) A = \<D> e A"
-| "\<D> (e\<bullet>F{D}) A = \<D> e A"
-| "\<D> (C\<bullet>\<^sub>sF{D}) A = True"
-| "\<D> (e\<^sub>1\<bullet>F{D}:=e\<^sub>2) A = (\<D> e\<^sub>1 A \<and> \<D> e\<^sub>2 (A \<squnion> \<A> e\<^sub>1))"
-| "\<D> (C\<bullet>\<^sub>sF{D}:=e\<^sub>2) A = \<D> e\<^sub>2 A"
-| "\<D> (e\<bullet>M(es)) A = (\<D> e A \<and> \<D>s es (A \<squnion> \<A> e))"
-| "\<D> (C\<bullet>\<^sub>sM(es)) A = \<D>s es A"
-| "\<D> ({V:T; e}) A = \<D> e (A \<ominus> V)"
-| "\<D> (e\<^sub>1;;e\<^sub>2) A = (\<D> e\<^sub>1 A \<and> \<D> e\<^sub>2 (A \<squnion> \<A> e\<^sub>1))"
-| "\<D> (if (e) e\<^sub>1 else e\<^sub>2) A =
- (\<D> e A \<and> \<D> e\<^sub>1 (A \<squnion> \<A> e) \<and> \<D> e\<^sub>2 (A \<squnion> \<A> e))"
-| "\<D> (while (e) c) A = (\<D> e A \<and> \<D> c (A \<squnion> \<A> e))"
-| "\<D> (throw e) A = \<D> e A"
-| "\<D> (try e\<^sub>1 catch(C V) e\<^sub>2) A = (\<D> e\<^sub>1 A \<and> \<D> e\<^sub>2 (A \<squnion> \<lfloor>{V}\<rfloor>))"
-| "\<D> (INIT C (Cs,b) \<leftarrow> e) A = \<D> e A"
-| "\<D> (RI (C,e);Cs \<leftarrow> e') A = (\<D> e A \<and> \<D> e' A)"
-
-| "\<D>s ([]) A = True"
-| "\<D>s (e#es) A = (\<D> e A \<and> \<D>s es (A \<squnion> \<A> e))"
-
-lemma As_map_Val[simp]: "\<A>s (map Val vs) = \<lfloor>{}\<rfloor>"
-(*<*)by (induct vs) simp_all(*>*)
-
-lemma D_append[iff]: "\<And>A. \<D>s (es @ es') A = (\<D>s es A \<and> \<D>s es' (A \<squnion> \<A>s es))"
-(*<*)by (induct es type:list) (auto simp:hyperUn_assoc)(*>*)
-
-
-lemma A_fv: "\<And>A. \<A> e = \<lfloor>A\<rfloor> \<Longrightarrow> A \<subseteq> fv e"
-and "\<And>A. \<A>s es = \<lfloor>A\<rfloor> \<Longrightarrow> A \<subseteq> fvs es"
-(*<*)
-by (induct e and es rule: \<A>.induct \<A>s.induct)
- (fastforce simp add:hyperset_defs)+
-(*>*)
-
-
-lemma sqUn_lem: "A \<sqsubseteq> A' \<Longrightarrow> A \<squnion> B \<sqsubseteq> A' \<squnion> B"
-(*<*)by(simp add:hyperset_defs) blast(*>*)
-
-lemma diff_lem: "A \<sqsubseteq> A' \<Longrightarrow> A \<ominus> b \<sqsubseteq> A' \<ominus> b"
-(*<*)by(simp add:hyperset_defs) blast(*>*)
-
-(* This order of the premises avoids looping of the simplifier *)
-lemma D_mono: "\<And>A A'. A \<sqsubseteq> A' \<Longrightarrow> \<D> e A \<Longrightarrow> \<D> (e::'a exp) A'"
-and Ds_mono: "\<And>A A'. A \<sqsubseteq> A' \<Longrightarrow> \<D>s es A \<Longrightarrow> \<D>s (es::'a exp list) A'"
-(*<*)
-proof(induct e and es rule: \<D>.induct \<D>s.induct)
- case BinOp then show ?case by simp (iprover dest:sqUn_lem)
-next
- case Var then show ?case by (fastforce simp add:hyperset_defs)
-next
- case FAss then show ?case by simp (iprover dest:sqUn_lem)
-next
- case Call then show ?case by simp (iprover dest:sqUn_lem)
-next
- case Block then show ?case by simp (iprover dest:diff_lem)
-next
- case Seq then show ?case by simp (iprover dest:sqUn_lem)
-next
- case Cond then show ?case by simp (iprover dest:sqUn_lem)
-next
- case While then show ?case by simp (iprover dest:sqUn_lem)
-next
- case TryCatch then show ?case by simp (iprover dest:sqUn_lem)
-next
- case Cons_exp then show ?case by simp (iprover dest:sqUn_lem)
-qed simp_all
-(*>*)
-
-(* And this is the order of premises preferred during application: *)
-lemma D_mono': "\<D> e A \<Longrightarrow> A \<sqsubseteq> A' \<Longrightarrow> \<D> e A'"
-and Ds_mono': "\<D>s es A \<Longrightarrow> A \<sqsubseteq> A' \<Longrightarrow> \<D>s es A'"
-(*<*)by(blast intro:D_mono, blast intro:Ds_mono)(*>*)
-
-
-lemma Ds_Vals: "\<D>s (map Val vs) A" by(induct vs, auto)
-
-end
+(* Title: JinjaDCI/J/DefAss.thy
+ Author: Tobias Nipkow, Susannah Mansky
+ Copyright 2003 Technische Universitaet Muenchen, 2019-20 UIUC
+
+ Based on the Jinja theory J/DefAss.thy by Tobias Nipkow
+*)
+
+section \<open> Definite assignment \<close>
+
+theory DefAss imports BigStep begin
+
+subsection "Hypersets"
+
+type_synonym 'a hyperset = "'a set option"
+
+definition hyperUn :: "'a hyperset \<Rightarrow> 'a hyperset \<Rightarrow> 'a hyperset" (infixl "\<squnion>" 65)
+where
+ "A \<squnion> B \<equiv> case A of None \<Rightarrow> None
+ | \<lfloor>A\<rfloor> \<Rightarrow> (case B of None \<Rightarrow> None | \<lfloor>B\<rfloor> \<Rightarrow> \<lfloor>A \<union> B\<rfloor>)"
+
+definition hyperInt :: "'a hyperset \<Rightarrow> 'a hyperset \<Rightarrow> 'a hyperset" (infixl "\<sqinter>" 70)
+where
+ "A \<sqinter> B \<equiv> case A of None \<Rightarrow> B
+ | \<lfloor>A\<rfloor> \<Rightarrow> (case B of None \<Rightarrow> \<lfloor>A\<rfloor> | \<lfloor>B\<rfloor> \<Rightarrow> \<lfloor>A \<inter> B\<rfloor>)"
+
+definition hyperDiff1 :: "'a hyperset \<Rightarrow> 'a \<Rightarrow> 'a hyperset" (infixl "\<ominus>" 65)
+where
+ "A \<ominus> a \<equiv> case A of None \<Rightarrow> None | \<lfloor>A\<rfloor> \<Rightarrow> \<lfloor>A - {a}\<rfloor>"
+
+definition hyper_isin :: "'a \<Rightarrow> 'a hyperset \<Rightarrow> bool" (infix "\<in>\<in>" 50)
+where
+ "a \<in>\<in> A \<equiv> case A of None \<Rightarrow> True | \<lfloor>A\<rfloor> \<Rightarrow> a \<in> A"
+
+definition hyper_subset :: "'a hyperset \<Rightarrow> 'a hyperset \<Rightarrow> bool" (infix "\<sqsubseteq>" 50)
+where
+ "A \<sqsubseteq> B \<equiv> case B of None \<Rightarrow> True
+ | \<lfloor>B\<rfloor> \<Rightarrow> (case A of None \<Rightarrow> False | \<lfloor>A\<rfloor> \<Rightarrow> A \<subseteq> B)"
+
+lemmas hyperset_defs =
+ hyperUn_def hyperInt_def hyperDiff1_def hyper_isin_def hyper_subset_def
+
+lemma [simp]: "\<lfloor>{}\<rfloor> \<squnion> A = A \<and> A \<squnion> \<lfloor>{}\<rfloor> = A"
+(*<*)by(simp add:hyperset_defs)(*>*)
+
+lemma [simp]: "\<lfloor>A\<rfloor> \<squnion> \<lfloor>B\<rfloor> = \<lfloor>A \<union> B\<rfloor> \<and> \<lfloor>A\<rfloor> \<ominus> a = \<lfloor>A - {a}\<rfloor>"
+(*<*)by(simp add:hyperset_defs)(*>*)
+
+lemma [simp]: "None \<squnion> A = None \<and> A \<squnion> None = None"
+(*<*)by(simp add:hyperset_defs)(*>*)
+
+lemma [simp]: "a \<in>\<in> None \<and> None \<ominus> a = None"
+(*<*)by(simp add:hyperset_defs)(*>*)
+
+lemma hyper_isin_union: "x \<in>\<in> \<lfloor>A\<rfloor> \<Longrightarrow> x \<in>\<in> \<lfloor>A\<rfloor> \<squnion> B"
+ by(case_tac B, auto simp: hyper_isin_def)
+
+lemma hyperUn_assoc: "(A \<squnion> B) \<squnion> C = A \<squnion> (B \<squnion> C)"
+(*<*)by(simp add:hyperset_defs Un_assoc)(*>*)
+
+lemma hyper_insert_comm: "A \<squnion> \<lfloor>{a}\<rfloor> = \<lfloor>{a}\<rfloor> \<squnion> A \<and> A \<squnion> (\<lfloor>{a}\<rfloor> \<squnion> B) = \<lfloor>{a}\<rfloor> \<squnion> (A \<squnion> B)"
+(*<*)by(simp add:hyperset_defs)(*>*)
+
+lemma hyper_comm: "A \<squnion> B = B \<squnion> A \<and> A \<squnion> B \<squnion> C = B \<squnion> A \<squnion> C"
+(*<*)
+proof(cases A)
+ case SomeA: Some then show ?thesis
+ proof(cases B)
+ case SomeB: Some with SomeA show ?thesis
+ proof(cases C)
+ case SomeC: Some with SomeA SomeB show ?thesis
+ by(simp add: Un_left_commute Un_commute)
+ qed (simp add: Un_commute)
+ qed simp
+qed simp
+(*>*)
+
+subsection "Definite assignment"
+
+primrec
+ \<A> :: "'a exp \<Rightarrow> 'a hyperset"
+ and \<A>s :: "'a exp list \<Rightarrow> 'a hyperset"
+where
+ "\<A> (new C) = \<lfloor>{}\<rfloor>"
+| "\<A> (Cast C e) = \<A> e"
+| "\<A> (Val v) = \<lfloor>{}\<rfloor>"
+| "\<A> (e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2) = \<A> e\<^sub>1 \<squnion> \<A> e\<^sub>2"
+| "\<A> (Var V) = \<lfloor>{}\<rfloor>"
+| "\<A> (LAss V e) = \<lfloor>{V}\<rfloor> \<squnion> \<A> e"
+| "\<A> (e\<bullet>F{D}) = \<A> e"
+| "\<A> (C\<bullet>\<^sub>sF{D}) = \<lfloor>{}\<rfloor>"
+| "\<A> (e\<^sub>1\<bullet>F{D}:=e\<^sub>2) = \<A> e\<^sub>1 \<squnion> \<A> e\<^sub>2"
+| "\<A> (C\<bullet>\<^sub>sF{D}:=e\<^sub>2) = \<A> e\<^sub>2"
+| "\<A> (e\<bullet>M(es)) = \<A> e \<squnion> \<A>s es"
+| "\<A> (C\<bullet>\<^sub>sM(es)) = \<A>s es"
+| "\<A> ({V:T; e}) = \<A> e \<ominus> V"
+| "\<A> (e\<^sub>1;;e\<^sub>2) = \<A> e\<^sub>1 \<squnion> \<A> e\<^sub>2"
+| "\<A> (if (e) e\<^sub>1 else e\<^sub>2) = \<A> e \<squnion> (\<A> e\<^sub>1 \<sqinter> \<A> e\<^sub>2)"
+| "\<A> (while (b) e) = \<A> b"
+| "\<A> (throw e) = None"
+| "\<A> (try e\<^sub>1 catch(C V) e\<^sub>2) = \<A> e\<^sub>1 \<sqinter> (\<A> e\<^sub>2 \<ominus> V)"
+| "\<A> (INIT C (Cs,b) \<leftarrow> e) = \<lfloor>{}\<rfloor>"
+| "\<A> (RI (C,e);Cs \<leftarrow> e') = \<A> e"
+
+| "\<A>s ([]) = \<lfloor>{}\<rfloor>"
+| "\<A>s (e#es) = \<A> e \<squnion> \<A>s es"
+
+primrec
+ \<D> :: "'a exp \<Rightarrow> 'a hyperset \<Rightarrow> bool"
+ and \<D>s :: "'a exp list \<Rightarrow> 'a hyperset \<Rightarrow> bool"
+where
+ "\<D> (new C) A = True"
+| "\<D> (Cast C e) A = \<D> e A"
+| "\<D> (Val v) A = True"
+| "\<D> (e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2) A = (\<D> e\<^sub>1 A \<and> \<D> e\<^sub>2 (A \<squnion> \<A> e\<^sub>1))"
+| "\<D> (Var V) A = (V \<in>\<in> A)"
+| "\<D> (LAss V e) A = \<D> e A"
+| "\<D> (e\<bullet>F{D}) A = \<D> e A"
+| "\<D> (C\<bullet>\<^sub>sF{D}) A = True"
+| "\<D> (e\<^sub>1\<bullet>F{D}:=e\<^sub>2) A = (\<D> e\<^sub>1 A \<and> \<D> e\<^sub>2 (A \<squnion> \<A> e\<^sub>1))"
+| "\<D> (C\<bullet>\<^sub>sF{D}:=e\<^sub>2) A = \<D> e\<^sub>2 A"
+| "\<D> (e\<bullet>M(es)) A = (\<D> e A \<and> \<D>s es (A \<squnion> \<A> e))"
+| "\<D> (C\<bullet>\<^sub>sM(es)) A = \<D>s es A"
+| "\<D> ({V:T; e}) A = \<D> e (A \<ominus> V)"
+| "\<D> (e\<^sub>1;;e\<^sub>2) A = (\<D> e\<^sub>1 A \<and> \<D> e\<^sub>2 (A \<squnion> \<A> e\<^sub>1))"
+| "\<D> (if (e) e\<^sub>1 else e\<^sub>2) A =
+ (\<D> e A \<and> \<D> e\<^sub>1 (A \<squnion> \<A> e) \<and> \<D> e\<^sub>2 (A \<squnion> \<A> e))"
+| "\<D> (while (e) c) A = (\<D> e A \<and> \<D> c (A \<squnion> \<A> e))"
+| "\<D> (throw e) A = \<D> e A"
+| "\<D> (try e\<^sub>1 catch(C V) e\<^sub>2) A = (\<D> e\<^sub>1 A \<and> \<D> e\<^sub>2 (A \<squnion> \<lfloor>{V}\<rfloor>))"
+| "\<D> (INIT C (Cs,b) \<leftarrow> e) A = \<D> e A"
+| "\<D> (RI (C,e);Cs \<leftarrow> e') A = (\<D> e A \<and> \<D> e' A)"
+
+| "\<D>s ([]) A = True"
+| "\<D>s (e#es) A = (\<D> e A \<and> \<D>s es (A \<squnion> \<A> e))"
+
+lemma As_map_Val[simp]: "\<A>s (map Val vs) = \<lfloor>{}\<rfloor>"
+(*<*)by (induct vs) simp_all(*>*)
+
+lemma D_append[iff]: "\<And>A. \<D>s (es @ es') A = (\<D>s es A \<and> \<D>s es' (A \<squnion> \<A>s es))"
+(*<*)by (induct es type:list) (auto simp:hyperUn_assoc)(*>*)
+
+
+lemma A_fv: "\<And>A. \<A> e = \<lfloor>A\<rfloor> \<Longrightarrow> A \<subseteq> fv e"
+and "\<And>A. \<A>s es = \<lfloor>A\<rfloor> \<Longrightarrow> A \<subseteq> fvs es"
+(*<*)
+by (induct e and es rule: \<A>.induct \<A>s.induct)
+ (fastforce simp add:hyperset_defs)+
+(*>*)
+
+
+lemma sqUn_lem: "A \<sqsubseteq> A' \<Longrightarrow> A \<squnion> B \<sqsubseteq> A' \<squnion> B"
+(*<*)by(simp add:hyperset_defs) blast(*>*)
+
+lemma diff_lem: "A \<sqsubseteq> A' \<Longrightarrow> A \<ominus> b \<sqsubseteq> A' \<ominus> b"
+(*<*)by(simp add:hyperset_defs) blast(*>*)
+
+(* This order of the premises avoids looping of the simplifier *)
+lemma D_mono: "\<And>A A'. A \<sqsubseteq> A' \<Longrightarrow> \<D> e A \<Longrightarrow> \<D> (e::'a exp) A'"
+and Ds_mono: "\<And>A A'. A \<sqsubseteq> A' \<Longrightarrow> \<D>s es A \<Longrightarrow> \<D>s (es::'a exp list) A'"
+(*<*)
+proof(induct e and es rule: \<D>.induct \<D>s.induct)
+ case BinOp then show ?case by simp (iprover dest:sqUn_lem)
+next
+ case Var then show ?case by (fastforce simp add:hyperset_defs)
+next
+ case FAss then show ?case by simp (iprover dest:sqUn_lem)
+next
+ case Call then show ?case by simp (iprover dest:sqUn_lem)
+next
+ case Block then show ?case by simp (iprover dest:diff_lem)
+next
+ case Seq then show ?case by simp (iprover dest:sqUn_lem)
+next
+ case Cond then show ?case by simp (iprover dest:sqUn_lem)
+next
+ case While then show ?case by simp (iprover dest:sqUn_lem)
+next
+ case TryCatch then show ?case by simp (iprover dest:sqUn_lem)
+next
+ case Cons_exp then show ?case by simp (iprover dest:sqUn_lem)
+qed simp_all
+(*>*)
+
+(* And this is the order of premises preferred during application: *)
+lemma D_mono': "\<D> e A \<Longrightarrow> A \<sqsubseteq> A' \<Longrightarrow> \<D> e A'"
+and Ds_mono': "\<D>s es A \<Longrightarrow> A \<sqsubseteq> A' \<Longrightarrow> \<D>s es A'"
+(*<*)by(blast intro:D_mono, blast intro:Ds_mono)(*>*)
+
+
+lemma Ds_Vals: "\<D>s (map Val vs) A" by(induct vs, auto)
+
+end
diff --git a/thys/JinjaDCI/J/EConform.thy b/thys/JinjaDCI/J/EConform.thy
--- a/thys/JinjaDCI/J/EConform.thy
+++ b/thys/JinjaDCI/J/EConform.thy
@@ -1,337 +1,337 @@
-(* Title: JinjaDCI/J/EConform.thy
- Author: Susannah Mansky
- 2019-20 UIUC
-*)
-
-section \<open> Expression conformance properties \<close>
-
-theory EConform
-imports SmallStep BigStep
-begin
-
-lemma cons_to_append: "list \<noteq> [] \<longrightarrow> (\<exists>ls. a # list = ls @ [last list])"
- by (metis append_butlast_last_id last_ConsR list.simps(3))
-
-subsection "Initialization conformance"
-
-\<comment> \<open> returns class that can be initialized (if any) by top-level expression \<close>
-fun init_class :: "'m prog \<Rightarrow> 'a exp \<Rightarrow> cname option" where
-"init_class P (new C) = Some C" |
-"init_class P (C\<bullet>\<^sub>sF{D}) = Some D" |
-"init_class P (C\<bullet>\<^sub>sF{D}:=e\<^sub>2) = Some D" |
-"init_class P (C\<bullet>\<^sub>sM(es)) = seeing_class P C M" |
-"init_class _ _ = None"
-
-lemma icheck_init_class: "icheck P C e \<Longrightarrow> init_class P e = \<lfloor>C\<rfloor>"
-proof(induct e)
- case (SFAss x1 x2 x3 e')
- then show ?case by(case_tac e') auto
-qed auto
-
-\<comment> \<open> exp to take next small step (in particular, subexp that may contain initialization) \<close>
-fun ss_exp :: "'a exp \<Rightarrow> 'a exp" and ss_exps :: "'a exp list \<Rightarrow> 'a exp option" where
- "ss_exp (new C) = new C"
-| "ss_exp (Cast C e) = (case val_of e of Some v \<Rightarrow> Cast C e | _ \<Rightarrow> ss_exp e)"
-| "ss_exp (Val v) = Val v"
-| "ss_exp (e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2) = (case val_of e\<^sub>1 of Some v \<Rightarrow> (case val_of e\<^sub>2 of Some v \<Rightarrow> e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2 | _ \<Rightarrow> ss_exp e\<^sub>2)
- | _ \<Rightarrow> ss_exp e\<^sub>1)"
-| "ss_exp (Var V) = Var V"
-| "ss_exp (LAss V e) = (case val_of e of Some v \<Rightarrow> LAss V e | _ \<Rightarrow> ss_exp e)"
-| "ss_exp (e\<bullet>F{D}) = (case val_of e of Some v \<Rightarrow> e\<bullet>F{D} | _ \<Rightarrow> ss_exp e)"
-| "ss_exp (C\<bullet>\<^sub>sF{D}) = C\<bullet>\<^sub>sF{D}"
-| "ss_exp (e\<^sub>1\<bullet>F{D}:=e\<^sub>2) = (case val_of e\<^sub>1 of Some v \<Rightarrow> (case val_of e\<^sub>2 of Some v \<Rightarrow> e\<^sub>1\<bullet>F{D}:=e\<^sub>2 | _ \<Rightarrow> ss_exp e\<^sub>2)
- | _ \<Rightarrow> ss_exp e\<^sub>1)"
-| "ss_exp (C\<bullet>\<^sub>sF{D}:=e\<^sub>2) = (case val_of e\<^sub>2 of Some v \<Rightarrow> C\<bullet>\<^sub>sF{D}:=e\<^sub>2 | _ \<Rightarrow> ss_exp e\<^sub>2)"
-| "ss_exp (e\<bullet>M(es)) = (case val_of e of Some v \<Rightarrow> (case map_vals_of es of Some t \<Rightarrow> e\<bullet>M(es) | _ \<Rightarrow> the(ss_exps es))
- | _ \<Rightarrow> ss_exp e)"
-| "ss_exp (C\<bullet>\<^sub>sM(es)) = (case map_vals_of es of Some t \<Rightarrow> C\<bullet>\<^sub>sM(es) | _ \<Rightarrow> the(ss_exps es))"
-| "ss_exp ({V:T; e}) = ss_exp e"
-| "ss_exp (e\<^sub>1;;e\<^sub>2) = (case val_of e\<^sub>1 of Some v \<Rightarrow> ss_exp e\<^sub>2
- | None \<Rightarrow> (case lass_val_of e\<^sub>1 of Some p \<Rightarrow> ss_exp e\<^sub>2
- | None \<Rightarrow> ss_exp e\<^sub>1))"
-| "ss_exp (if (b) e\<^sub>1 else e\<^sub>2) = (case bool_of b of Some True \<Rightarrow> if (b) e\<^sub>1 else e\<^sub>2
- | Some False \<Rightarrow> if (b) e\<^sub>1 else e\<^sub>2
- | _ \<Rightarrow> ss_exp b)"
-| "ss_exp (while (b) e) = while (b) e"
-| "ss_exp (throw e) = (case val_of e of Some v \<Rightarrow> throw e | _ \<Rightarrow> ss_exp e)"
-| "ss_exp (try e\<^sub>1 catch(C V) e\<^sub>2) = (case val_of e\<^sub>1 of Some v \<Rightarrow> try e\<^sub>1 catch(C V) e\<^sub>2
- | _ \<Rightarrow> ss_exp e\<^sub>1)"
-| "ss_exp (INIT C (Cs,b) \<leftarrow> e) = INIT C (Cs,b) \<leftarrow> e"
-| "ss_exp (RI (C,e);Cs \<leftarrow> e') = (case val_of e of Some v \<Rightarrow> RI (C,e);Cs \<leftarrow> e | _ \<Rightarrow> ss_exp e)"
-| "ss_exps([]) = None"
-| "ss_exps(e#es) = (case val_of e of Some v \<Rightarrow> ss_exps es | _ \<Rightarrow> Some (ss_exp e))"
-
-(*<*)
-lemmas ss_exp_ss_exps_induct = ss_exp_ss_exps.induct
- [ case_names New Cast Val BinOp Var LAss FAcc SFAcc FAss SFAss Call SCall
- Block Seq Cond While Throw Try Init RI Nil Cons ]
-(*>*)
-
-lemma icheck_ss_exp:
-assumes "icheck P C e" shows "ss_exp e = e"
-using assms
-proof(cases e)
- case (SFAss C F D e) then show ?thesis using assms
- proof(cases e)qed(auto)
-qed(auto)
-
-lemma ss_exps_Vals_None[simp]:
- "ss_exps (map Val vs) = None"
- by(induct vs) (auto)
-
-lemma ss_exps_Vals_NoneI:
- "ss_exps es = None \<Longrightarrow> \<exists>vs. es = map Val vs"
-using val_of_spec by(induct es) (auto)
-
-lemma ss_exps_throw_nVal:
- "\<lbrakk> val_of e = None; ss_exps (map Val vs @ throw e # es') = \<lfloor>e'\<rfloor> \<rbrakk>
- \<Longrightarrow> e' = ss_exp e"
- by(induct vs) (auto)
-
-lemma ss_exps_throw_Val:
- "\<lbrakk> val_of e = \<lfloor>a\<rfloor>; ss_exps (map Val vs @ throw e # es') = \<lfloor>e'\<rfloor> \<rbrakk>
- \<Longrightarrow> e' = throw e"
- by(induct vs) (auto)
-
-
-abbreviation curr_init :: "'m prog \<Rightarrow> 'a exp \<Rightarrow> cname option" where
-"curr_init P e \<equiv> init_class P (ss_exp e)"
-abbreviation curr_inits :: "'m prog \<Rightarrow> 'a exp list \<Rightarrow> cname option" where
-"curr_inits P es \<equiv> case ss_exps es of Some e \<Rightarrow> init_class P e | _ \<Rightarrow> None"
-
-lemma icheck_curr_init': "\<And>e'. ss_exp e = e' \<Longrightarrow> icheck P C e' \<Longrightarrow> curr_init P e = \<lfloor>C\<rfloor>"
- and icheck_curr_inits': "\<And>e. ss_exps es = \<lfloor>e\<rfloor> \<Longrightarrow> icheck P C e \<Longrightarrow> curr_inits P es = \<lfloor>C\<rfloor>"
-proof(induct rule: ss_exp_ss_exps_induct)
-qed(simp_all add: icheck_init_class)
-
-lemma icheck_curr_init: "icheck P C e' \<Longrightarrow> ss_exp e = e' \<Longrightarrow> curr_init P e = \<lfloor>C\<rfloor>"
- by(rule icheck_curr_init')
-
-lemma icheck_curr_inits: "icheck P C e \<Longrightarrow> ss_exps es = \<lfloor>e\<rfloor> \<Longrightarrow> curr_inits P es = \<lfloor>C\<rfloor>"
- by(rule icheck_curr_inits')
-
-definition initPD :: "sheap \<Rightarrow> cname \<Rightarrow> bool" where
-"initPD sh C \<equiv> \<exists>sfs i. sh C = Some (sfs, i) \<and> (i = Done \<or> i = Processing)"
-
-\<comment> \<open> checks that @{text INIT} and @{text RI} conform and are only in the main computation \<close>
-fun iconf :: "sheap \<Rightarrow> 'a exp \<Rightarrow> bool" and iconfs :: " sheap \<Rightarrow> 'a exp list \<Rightarrow> bool" where
- "iconf sh (new C) = True"
-| "iconf sh (Cast C e) = iconf sh e"
-| "iconf sh (Val v) = True"
-| "iconf sh (e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2) = (case val_of e\<^sub>1 of Some v \<Rightarrow> iconf sh e\<^sub>2 | _ \<Rightarrow> iconf sh e\<^sub>1 \<and> \<not>sub_RI e\<^sub>2)"
-| "iconf sh (Var V) = True"
-| "iconf sh (LAss V e) = iconf sh e"
-| "iconf sh (e\<bullet>F{D}) = iconf sh e"
-| "iconf sh (C\<bullet>\<^sub>sF{D}) = True"
-| "iconf sh (e\<^sub>1\<bullet>F{D}:=e\<^sub>2) = (case val_of e\<^sub>1 of Some v \<Rightarrow> iconf sh e\<^sub>2 | _ \<Rightarrow> iconf sh e\<^sub>1 \<and> \<not>sub_RI e\<^sub>2)"
-| "iconf sh (C\<bullet>\<^sub>sF{D}:=e\<^sub>2) = iconf sh e\<^sub>2"
-| "iconf sh (e\<bullet>M(es)) = (case val_of e of Some v \<Rightarrow> iconfs sh es | _ \<Rightarrow> iconf sh e \<and> \<not>sub_RIs es)"
-| "iconf sh (C\<bullet>\<^sub>sM(es)) = iconfs sh es"
-| "iconf sh ({V:T; e}) = iconf sh e"
-| "iconf sh (e\<^sub>1;;e\<^sub>2) = (case val_of e\<^sub>1 of Some v \<Rightarrow> iconf sh e\<^sub>2
- | None \<Rightarrow> (case lass_val_of e\<^sub>1 of Some p \<Rightarrow> iconf sh e\<^sub>2
- | None \<Rightarrow> iconf sh e\<^sub>1 \<and> \<not>sub_RI e\<^sub>2))"
-| "iconf sh (if (b) e\<^sub>1 else e\<^sub>2) = (iconf sh b \<and> \<not>sub_RI e\<^sub>1 \<and> \<not>sub_RI e\<^sub>2)"
-| "iconf sh (while (b) e) = (\<not>sub_RI b \<and> \<not>sub_RI e)"
-| "iconf sh (throw e) = iconf sh e"
-| "iconf sh (try e\<^sub>1 catch(C V) e\<^sub>2) = (iconf sh e\<^sub>1 \<and> \<not>sub_RI e\<^sub>2)"
-| "iconf sh (INIT C (Cs,b) \<leftarrow> e) = ((case Cs of Nil \<Rightarrow> initPD sh C | C'#Cs' \<Rightarrow> last Cs = C) \<and> \<not>sub_RI e)"
-| "iconf sh (RI (C,e);Cs \<leftarrow> e') = (iconf sh e \<and> \<not>sub_RI e')"
-| "iconfs sh ([]) = True"
-| "iconfs sh (e#es) = (case val_of e of Some v \<Rightarrow> iconfs sh es | _ \<Rightarrow> iconf sh e \<and> \<not>sub_RIs es)"
-
-lemma iconfs_map_throw: "iconfs sh (map Val vs @ throw e # es') \<Longrightarrow> iconf sh e"
- by(induct vs,auto)
-
-lemma nsub_RI_iconf_aux:
- "(\<not>sub_RI (e::'a exp) \<longrightarrow> (\<forall>e'. e' \<in> subexp e \<longrightarrow> \<not>sub_RI e' \<longrightarrow> iconf sh e') \<longrightarrow> iconf sh e)
- \<and> (\<not>sub_RIs (es::'a exp list) \<longrightarrow> (\<forall>e'. e' \<in> subexps es \<longrightarrow> \<not>sub_RI e' \<longrightarrow> iconf sh e') \<longrightarrow> iconfs sh es)"
-proof(induct rule: sub_RI_sub_RIs.induct) qed(auto)
-
-lemma nsub_RI_iconf_aux':
- "(\<And>e'. subexp_of e' e \<Longrightarrow> \<not>sub_RI e' \<longrightarrow> iconf sh e') \<Longrightarrow> (\<not>sub_RI e \<Longrightarrow> iconf sh e)"
- by(simp add: nsub_RI_iconf_aux)
-
-lemma nsub_RI_iconf: "\<not>sub_RI e \<Longrightarrow> iconf sh e"
- and nsub_RIs_iconfs: "\<not>sub_RIs es \<Longrightarrow> iconfs sh es"
-proof -
- let ?R = "\<lambda>e. \<not>sub_RI e \<longrightarrow> iconf sh e"
- let ?Rs = "\<lambda>es. \<not>sub_RIs es \<longrightarrow> iconfs sh es"
- have "(\<forall>e'. subexp_of e' e \<longrightarrow> ?R e') \<and> ?R e"
- by(rule subexp_induct[where ?Rs = ?Rs]; clarsimp simp: nsub_RI_iconf_aux)
- moreover have "(\<forall>e'. e' \<in> subexps es \<longrightarrow> ?R e') \<and> ?Rs es"
- by(rule subexps_induct; clarsimp simp: nsub_RI_iconf_aux)
- ultimately show "\<not>sub_RI e \<Longrightarrow> iconf sh e"
- and "\<not>sub_RIs es \<Longrightarrow> iconfs sh es" by simp+
-qed
-
-lemma lass_val_of_iconf: "lass_val_of e = \<lfloor>a\<rfloor> \<Longrightarrow> iconf sh e"
- by(drule lass_val_of_nsub_RI, erule nsub_RI_iconf)
-
-lemma icheck_iconf:
-assumes "icheck P C e" shows "iconf sh e"
-using assms
-proof(cases e)
- case (SFAss C F D e) then show ?thesis using assms
- proof(cases e)qed(auto)
-next
- case (SCall C M es) then show ?thesis using assms
- by (auto simp: nsub_RIs_iconfs)
-next
-qed(auto)
-
-
-subsection "Indicator boolean conformance"
-
-\<comment> \<open> checks that the given expression, indicator boolean pair is allowed in small-step
- (i.e., if @{term b} is True, then @{term e} is an initialization-calling expression to
- a class that is marked either @{term Processing} or @{term Done}) \<close>
-definition bconf :: "'m prog \<Rightarrow> sheap \<Rightarrow> 'a exp \<Rightarrow> bool \<Rightarrow> bool" ("_,_ \<turnstile>\<^sub>b '(_,_') \<surd>" [51,51,0,0] 50)
-where
- "P,sh \<turnstile>\<^sub>b (e,b) \<surd> \<equiv> b \<longrightarrow> (\<exists>C. icheck P C (ss_exp e) \<and> initPD sh C)"
-
-definition bconfs :: "'m prog \<Rightarrow> sheap \<Rightarrow> 'a exp list \<Rightarrow> bool \<Rightarrow> bool" ("_,_ \<turnstile>\<^sub>b '(_,_') \<surd>" [51,51,0,0] 50)
-where
- "P,sh \<turnstile>\<^sub>b (es,b) \<surd> \<equiv> b \<longrightarrow> (\<exists>C. (icheck P C (the(ss_exps es))
- \<and> (curr_inits P es = Some C) \<and> initPD sh C))"
-
-
-\<comment> \<open> bconf helper lemmas \<close>
-
-lemma bconf_nonVal[simp]:
- "P,sh \<turnstile>\<^sub>b (e,True) \<surd> \<Longrightarrow> val_of e = None"
- by(cases e) (auto simp: bconf_def)
-
-lemma bconfs_nonVals[simp]:
- "P,sh \<turnstile>\<^sub>b (es,True) \<surd> \<Longrightarrow> map_vals_of es = None"
- by(induct es) (auto simp: bconfs_def)
-
-lemma bconf_Cast[iff]:
- "P,sh \<turnstile>\<^sub>b (Cast C e,b) \<surd> \<longleftrightarrow> P,sh \<turnstile>\<^sub>b (e,b) \<surd>"
- by(cases b) (auto simp: bconf_def dest: val_of_spec)
-
-lemma bconf_BinOp[iff]:
- "P,sh \<turnstile>\<^sub>b (e1 \<guillemotleft>bop\<guillemotright> e2,b) \<surd>
- \<longleftrightarrow> (case val_of e1 of Some v \<Rightarrow> P,sh \<turnstile>\<^sub>b (e2,b) \<surd> | _ \<Rightarrow> P,sh \<turnstile>\<^sub>b (e1,b) \<surd>)"
- by(cases b) (auto simp: bconf_def dest: val_of_spec)
-
-lemma bconf_LAss[iff]:
- "P,sh \<turnstile>\<^sub>b (LAss V e,b) \<surd> \<longleftrightarrow> P,sh \<turnstile>\<^sub>b (e,b) \<surd>"
- by(cases b) (auto simp: bconf_def dest: val_of_spec)
-
-lemma bconf_FAcc[iff]:
- "P,sh \<turnstile>\<^sub>b (e\<bullet>F{D},b) \<surd> \<longleftrightarrow> P,sh \<turnstile>\<^sub>b (e,b) \<surd>"
- by(cases b) (auto simp: bconf_def dest: val_of_spec)
-
-lemma bconf_FAss[iff]:
- "P,sh \<turnstile>\<^sub>b (FAss e1 F D e2,b) \<surd>
- \<longleftrightarrow> (case val_of e1 of Some v \<Rightarrow> P,sh \<turnstile>\<^sub>b (e2,b) \<surd> | _ \<Rightarrow> P,sh \<turnstile>\<^sub>b (e1,b) \<surd>)"
- by(cases b) (auto simp: bconf_def dest: val_of_spec)
-
-lemma bconf_SFAss[iff]:
-"val_of e2 = None \<Longrightarrow> P,sh \<turnstile>\<^sub>b (SFAss C F D e2,b) \<surd> \<longleftrightarrow> P,sh \<turnstile>\<^sub>b (e2,b) \<surd>"
- by(cases b) (auto simp: bconf_def)
-
-lemma bconfs_Vals[iff]:
- "P,sh \<turnstile>\<^sub>b (map Val vs, b) \<surd> \<longleftrightarrow> \<not> b"
- by(unfold bconfs_def) simp
-
-lemma bconf_Call[iff]:
- "P,sh \<turnstile>\<^sub>b (e\<bullet>M(es),b) \<surd>
- \<longleftrightarrow> (case val_of e of Some v \<Rightarrow> P,sh \<turnstile>\<^sub>b (es,b) \<surd> | _ \<Rightarrow> P,sh \<turnstile>\<^sub>b (e,b) \<surd>)"
-proof(cases b)
- case True
- then show ?thesis
- proof(cases "ss_exps es")
- case None
- then obtain vs where "es = map Val vs" using ss_exps_Vals_NoneI by auto
- then have mv: "map_vals_of es = \<lfloor>vs\<rfloor>" by simp
- then show ?thesis by(auto simp: bconf_def) (simp add: bconfs_def)
- next
- case (Some a)
- then show ?thesis by(auto simp: bconf_def) (auto simp: bconfs_def icheck_init_class)
- qed
-qed(simp add: bconf_def bconfs_def)
-
-lemma bconf_SCall[iff]:
-assumes mvn: "map_vals_of es = None"
-shows "P,sh \<turnstile>\<^sub>b (C\<bullet>\<^sub>sM(es),b) \<surd> \<longleftrightarrow> P,sh \<turnstile>\<^sub>b (es,b) \<surd>"
-proof(cases b)
- case True
- then show ?thesis
- proof(cases "ss_exps es")
- case None
- then have "\<exists>vs. es = map Val vs" using ss_exps_Vals_NoneI by auto
- then show ?thesis using mvn finals_def by clarsimp
- next
- case (Some a)
- then show ?thesis by(auto simp: bconf_def) (auto simp: bconfs_def icheck_init_class)
- qed
-qed(simp add: bconf_def bconfs_def)
-
-lemma bconf_Cons[iff]:
- "P,sh \<turnstile>\<^sub>b (e#es,b) \<surd>
- \<longleftrightarrow> (case val_of e of Some v \<Rightarrow> P,sh \<turnstile>\<^sub>b (es,b) \<surd> | _ \<Rightarrow> P,sh \<turnstile>\<^sub>b (e,b) \<surd>)"
-proof(cases b)
- case True
- then show ?thesis
- proof(cases "ss_exps es")
- case None
- then have "\<exists>vs. es = map Val vs" using ss_exps_Vals_NoneI by auto
- then show ?thesis using None by(auto simp: bconf_def bconfs_def icheck_init_class)
- next
- case (Some a)
- then show ?thesis by(auto simp: bconf_def bconfs_def icheck_init_class)
- qed
-qed(simp add: bconf_def bconfs_def)
-
-lemma bconf_InitBlock[iff]:
- "P,sh \<turnstile>\<^sub>b ({V:T; V:=Val v;; e\<^sub>2},b) \<surd> \<longleftrightarrow> P,sh \<turnstile>\<^sub>b (e\<^sub>2,b) \<surd>"
- by(cases b) (auto simp: bconf_def assigned_def)
-
-lemma bconf_Block[iff]:
- "P,sh \<turnstile>\<^sub>b ({V:T; e},b) \<surd> \<longleftrightarrow> P,sh \<turnstile>\<^sub>b (e,b) \<surd>"
- by(cases b) (auto simp: bconf_def)
-
-lemma bconf_Seq[iff]:
- "P,sh \<turnstile>\<^sub>b (e1;;e2,b) \<surd>
- \<longleftrightarrow> (case val_of e1 of Some v \<Rightarrow> P,sh \<turnstile>\<^sub>b (e2,b) \<surd>
- | _ \<Rightarrow> (case lass_val_of e1 of Some p \<Rightarrow> P,sh \<turnstile>\<^sub>b (e2,b) \<surd>
- | None \<Rightarrow> P,sh \<turnstile>\<^sub>b (e1,b) \<surd>))"
- by(cases b) (auto simp: bconf_def dest: val_of_spec lass_val_of_spec)
-
-lemma bconf_Cond[iff]:
- "P,sh \<turnstile>\<^sub>b (if (b) e\<^sub>1 else e\<^sub>2,b') \<surd> \<longleftrightarrow> P,sh \<turnstile>\<^sub>b (b,b') \<surd>"
-proof(cases "bool_of b")
- case None
- then show ?thesis by(auto simp: bconf_def)
-next
- case (Some a)
- then show ?thesis by(case_tac a) (auto simp: bconf_def dest: bool_of_specT bool_of_specF)
-qed
-
-lemma bconf_While[iff]:
- "P,sh \<turnstile>\<^sub>b (while (b) e,b') \<surd> \<longleftrightarrow> \<not>b'"
- by(cases b) (auto simp: bconf_def)
-
-lemma bconf_Throw[iff]:
- "P,sh \<turnstile>\<^sub>b (throw e,b) \<surd> \<longleftrightarrow> P,sh \<turnstile>\<^sub>b (e,b) \<surd>"
- by(cases b) (auto simp: bconf_def dest: val_of_spec)
-
-lemma bconf_Try[iff]:
- "P,sh \<turnstile>\<^sub>b (try e\<^sub>1 catch(C V) e\<^sub>2,b) \<surd> \<longleftrightarrow> P,sh \<turnstile>\<^sub>b (e\<^sub>1,b) \<surd>"
- by(cases b) (auto simp: bconf_def dest: val_of_spec)
-
-lemma bconf_INIT[iff]:
- "P,sh \<turnstile>\<^sub>b (INIT C (Cs,b') \<leftarrow> e,b) \<surd> \<longleftrightarrow> \<not>b"
- by(cases b) (auto simp: bconf_def)
-
-lemma bconf_RI[iff]:
- "P,sh \<turnstile>\<^sub>b (RI(C,e);Cs \<leftarrow> e',b) \<surd> \<longleftrightarrow> P,sh \<turnstile>\<^sub>b (e,b) \<surd>"
- by(cases b) (auto simp: bconf_def dest: val_of_spec)
-
-lemma bconfs_map_throw[iff]:
- "P,sh \<turnstile>\<^sub>b (map Val vs @ throw e # es',b) \<surd> \<longleftrightarrow> P,sh \<turnstile>\<^sub>b (e,b) \<surd>"
- by(induct vs) auto
-
-end
+(* Title: JinjaDCI/J/EConform.thy
+ Author: Susannah Mansky
+ 2019-20 UIUC
+*)
+
+section \<open> Expression conformance properties \<close>
+
+theory EConform
+imports SmallStep BigStep
+begin
+
+lemma cons_to_append: "list \<noteq> [] \<longrightarrow> (\<exists>ls. a # list = ls @ [last list])"
+ by (metis append_butlast_last_id last_ConsR list.simps(3))
+
+subsection "Initialization conformance"
+
+\<comment> \<open> returns class that can be initialized (if any) by top-level expression \<close>
+fun init_class :: "'m prog \<Rightarrow> 'a exp \<Rightarrow> cname option" where
+"init_class P (new C) = Some C" |
+"init_class P (C\<bullet>\<^sub>sF{D}) = Some D" |
+"init_class P (C\<bullet>\<^sub>sF{D}:=e\<^sub>2) = Some D" |
+"init_class P (C\<bullet>\<^sub>sM(es)) = seeing_class P C M" |
+"init_class _ _ = None"
+
+lemma icheck_init_class: "icheck P C e \<Longrightarrow> init_class P e = \<lfloor>C\<rfloor>"
+proof(induct e)
+ case (SFAss x1 x2 x3 e')
+ then show ?case by(case_tac e') auto
+qed auto
+
+\<comment> \<open> exp to take next small step (in particular, subexp that may contain initialization) \<close>
+fun ss_exp :: "'a exp \<Rightarrow> 'a exp" and ss_exps :: "'a exp list \<Rightarrow> 'a exp option" where
+ "ss_exp (new C) = new C"
+| "ss_exp (Cast C e) = (case val_of e of Some v \<Rightarrow> Cast C e | _ \<Rightarrow> ss_exp e)"
+| "ss_exp (Val v) = Val v"
+| "ss_exp (e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2) = (case val_of e\<^sub>1 of Some v \<Rightarrow> (case val_of e\<^sub>2 of Some v \<Rightarrow> e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2 | _ \<Rightarrow> ss_exp e\<^sub>2)
+ | _ \<Rightarrow> ss_exp e\<^sub>1)"
+| "ss_exp (Var V) = Var V"
+| "ss_exp (LAss V e) = (case val_of e of Some v \<Rightarrow> LAss V e | _ \<Rightarrow> ss_exp e)"
+| "ss_exp (e\<bullet>F{D}) = (case val_of e of Some v \<Rightarrow> e\<bullet>F{D} | _ \<Rightarrow> ss_exp e)"
+| "ss_exp (C\<bullet>\<^sub>sF{D}) = C\<bullet>\<^sub>sF{D}"
+| "ss_exp (e\<^sub>1\<bullet>F{D}:=e\<^sub>2) = (case val_of e\<^sub>1 of Some v \<Rightarrow> (case val_of e\<^sub>2 of Some v \<Rightarrow> e\<^sub>1\<bullet>F{D}:=e\<^sub>2 | _ \<Rightarrow> ss_exp e\<^sub>2)
+ | _ \<Rightarrow> ss_exp e\<^sub>1)"
+| "ss_exp (C\<bullet>\<^sub>sF{D}:=e\<^sub>2) = (case val_of e\<^sub>2 of Some v \<Rightarrow> C\<bullet>\<^sub>sF{D}:=e\<^sub>2 | _ \<Rightarrow> ss_exp e\<^sub>2)"
+| "ss_exp (e\<bullet>M(es)) = (case val_of e of Some v \<Rightarrow> (case map_vals_of es of Some t \<Rightarrow> e\<bullet>M(es) | _ \<Rightarrow> the(ss_exps es))
+ | _ \<Rightarrow> ss_exp e)"
+| "ss_exp (C\<bullet>\<^sub>sM(es)) = (case map_vals_of es of Some t \<Rightarrow> C\<bullet>\<^sub>sM(es) | _ \<Rightarrow> the(ss_exps es))"
+| "ss_exp ({V:T; e}) = ss_exp e"
+| "ss_exp (e\<^sub>1;;e\<^sub>2) = (case val_of e\<^sub>1 of Some v \<Rightarrow> ss_exp e\<^sub>2
+ | None \<Rightarrow> (case lass_val_of e\<^sub>1 of Some p \<Rightarrow> ss_exp e\<^sub>2
+ | None \<Rightarrow> ss_exp e\<^sub>1))"
+| "ss_exp (if (b) e\<^sub>1 else e\<^sub>2) = (case bool_of b of Some True \<Rightarrow> if (b) e\<^sub>1 else e\<^sub>2
+ | Some False \<Rightarrow> if (b) e\<^sub>1 else e\<^sub>2
+ | _ \<Rightarrow> ss_exp b)"
+| "ss_exp (while (b) e) = while (b) e"
+| "ss_exp (throw e) = (case val_of e of Some v \<Rightarrow> throw e | _ \<Rightarrow> ss_exp e)"
+| "ss_exp (try e\<^sub>1 catch(C V) e\<^sub>2) = (case val_of e\<^sub>1 of Some v \<Rightarrow> try e\<^sub>1 catch(C V) e\<^sub>2
+ | _ \<Rightarrow> ss_exp e\<^sub>1)"
+| "ss_exp (INIT C (Cs,b) \<leftarrow> e) = INIT C (Cs,b) \<leftarrow> e"
+| "ss_exp (RI (C,e);Cs \<leftarrow> e') = (case val_of e of Some v \<Rightarrow> RI (C,e);Cs \<leftarrow> e | _ \<Rightarrow> ss_exp e)"
+| "ss_exps([]) = None"
+| "ss_exps(e#es) = (case val_of e of Some v \<Rightarrow> ss_exps es | _ \<Rightarrow> Some (ss_exp e))"
+
+(*<*)
+lemmas ss_exp_ss_exps_induct = ss_exp_ss_exps.induct
+ [ case_names New Cast Val BinOp Var LAss FAcc SFAcc FAss SFAss Call SCall
+ Block Seq Cond While Throw Try Init RI Nil Cons ]
+(*>*)
+
+lemma icheck_ss_exp:
+assumes "icheck P C e" shows "ss_exp e = e"
+using assms
+proof(cases e)
+ case (SFAss C F D e) then show ?thesis using assms
+ proof(cases e)qed(auto)
+qed(auto)
+
+lemma ss_exps_Vals_None[simp]:
+ "ss_exps (map Val vs) = None"
+ by(induct vs) (auto)
+
+lemma ss_exps_Vals_NoneI:
+ "ss_exps es = None \<Longrightarrow> \<exists>vs. es = map Val vs"
+using val_of_spec by(induct es) (auto)
+
+lemma ss_exps_throw_nVal:
+ "\<lbrakk> val_of e = None; ss_exps (map Val vs @ throw e # es') = \<lfloor>e'\<rfloor> \<rbrakk>
+ \<Longrightarrow> e' = ss_exp e"
+ by(induct vs) (auto)
+
+lemma ss_exps_throw_Val:
+ "\<lbrakk> val_of e = \<lfloor>a\<rfloor>; ss_exps (map Val vs @ throw e # es') = \<lfloor>e'\<rfloor> \<rbrakk>
+ \<Longrightarrow> e' = throw e"
+ by(induct vs) (auto)
+
+
+abbreviation curr_init :: "'m prog \<Rightarrow> 'a exp \<Rightarrow> cname option" where
+"curr_init P e \<equiv> init_class P (ss_exp e)"
+abbreviation curr_inits :: "'m prog \<Rightarrow> 'a exp list \<Rightarrow> cname option" where
+"curr_inits P es \<equiv> case ss_exps es of Some e \<Rightarrow> init_class P e | _ \<Rightarrow> None"
+
+lemma icheck_curr_init': "\<And>e'. ss_exp e = e' \<Longrightarrow> icheck P C e' \<Longrightarrow> curr_init P e = \<lfloor>C\<rfloor>"
+ and icheck_curr_inits': "\<And>e. ss_exps es = \<lfloor>e\<rfloor> \<Longrightarrow> icheck P C e \<Longrightarrow> curr_inits P es = \<lfloor>C\<rfloor>"
+proof(induct rule: ss_exp_ss_exps_induct)
+qed(simp_all add: icheck_init_class)
+
+lemma icheck_curr_init: "icheck P C e' \<Longrightarrow> ss_exp e = e' \<Longrightarrow> curr_init P e = \<lfloor>C\<rfloor>"
+ by(rule icheck_curr_init')
+
+lemma icheck_curr_inits: "icheck P C e \<Longrightarrow> ss_exps es = \<lfloor>e\<rfloor> \<Longrightarrow> curr_inits P es = \<lfloor>C\<rfloor>"
+ by(rule icheck_curr_inits')
+
+definition initPD :: "sheap \<Rightarrow> cname \<Rightarrow> bool" where
+"initPD sh C \<equiv> \<exists>sfs i. sh C = Some (sfs, i) \<and> (i = Done \<or> i = Processing)"
+
+\<comment> \<open> checks that @{text INIT} and @{text RI} conform and are only in the main computation \<close>
+fun iconf :: "sheap \<Rightarrow> 'a exp \<Rightarrow> bool" and iconfs :: " sheap \<Rightarrow> 'a exp list \<Rightarrow> bool" where
+ "iconf sh (new C) = True"
+| "iconf sh (Cast C e) = iconf sh e"
+| "iconf sh (Val v) = True"
+| "iconf sh (e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2) = (case val_of e\<^sub>1 of Some v \<Rightarrow> iconf sh e\<^sub>2 | _ \<Rightarrow> iconf sh e\<^sub>1 \<and> \<not>sub_RI e\<^sub>2)"
+| "iconf sh (Var V) = True"
+| "iconf sh (LAss V e) = iconf sh e"
+| "iconf sh (e\<bullet>F{D}) = iconf sh e"
+| "iconf sh (C\<bullet>\<^sub>sF{D}) = True"
+| "iconf sh (e\<^sub>1\<bullet>F{D}:=e\<^sub>2) = (case val_of e\<^sub>1 of Some v \<Rightarrow> iconf sh e\<^sub>2 | _ \<Rightarrow> iconf sh e\<^sub>1 \<and> \<not>sub_RI e\<^sub>2)"
+| "iconf sh (C\<bullet>\<^sub>sF{D}:=e\<^sub>2) = iconf sh e\<^sub>2"
+| "iconf sh (e\<bullet>M(es)) = (case val_of e of Some v \<Rightarrow> iconfs sh es | _ \<Rightarrow> iconf sh e \<and> \<not>sub_RIs es)"
+| "iconf sh (C\<bullet>\<^sub>sM(es)) = iconfs sh es"
+| "iconf sh ({V:T; e}) = iconf sh e"
+| "iconf sh (e\<^sub>1;;e\<^sub>2) = (case val_of e\<^sub>1 of Some v \<Rightarrow> iconf sh e\<^sub>2
+ | None \<Rightarrow> (case lass_val_of e\<^sub>1 of Some p \<Rightarrow> iconf sh e\<^sub>2
+ | None \<Rightarrow> iconf sh e\<^sub>1 \<and> \<not>sub_RI e\<^sub>2))"
+| "iconf sh (if (b) e\<^sub>1 else e\<^sub>2) = (iconf sh b \<and> \<not>sub_RI e\<^sub>1 \<and> \<not>sub_RI e\<^sub>2)"
+| "iconf sh (while (b) e) = (\<not>sub_RI b \<and> \<not>sub_RI e)"
+| "iconf sh (throw e) = iconf sh e"
+| "iconf sh (try e\<^sub>1 catch(C V) e\<^sub>2) = (iconf sh e\<^sub>1 \<and> \<not>sub_RI e\<^sub>2)"
+| "iconf sh (INIT C (Cs,b) \<leftarrow> e) = ((case Cs of Nil \<Rightarrow> initPD sh C | C'#Cs' \<Rightarrow> last Cs = C) \<and> \<not>sub_RI e)"
+| "iconf sh (RI (C,e);Cs \<leftarrow> e') = (iconf sh e \<and> \<not>sub_RI e')"
+| "iconfs sh ([]) = True"
+| "iconfs sh (e#es) = (case val_of e of Some v \<Rightarrow> iconfs sh es | _ \<Rightarrow> iconf sh e \<and> \<not>sub_RIs es)"
+
+lemma iconfs_map_throw: "iconfs sh (map Val vs @ throw e # es') \<Longrightarrow> iconf sh e"
+ by(induct vs,auto)
+
+lemma nsub_RI_iconf_aux:
+ "(\<not>sub_RI (e::'a exp) \<longrightarrow> (\<forall>e'. e' \<in> subexp e \<longrightarrow> \<not>sub_RI e' \<longrightarrow> iconf sh e') \<longrightarrow> iconf sh e)
+ \<and> (\<not>sub_RIs (es::'a exp list) \<longrightarrow> (\<forall>e'. e' \<in> subexps es \<longrightarrow> \<not>sub_RI e' \<longrightarrow> iconf sh e') \<longrightarrow> iconfs sh es)"
+proof(induct rule: sub_RI_sub_RIs.induct) qed(auto)
+
+lemma nsub_RI_iconf_aux':
+ "(\<And>e'. subexp_of e' e \<Longrightarrow> \<not>sub_RI e' \<longrightarrow> iconf sh e') \<Longrightarrow> (\<not>sub_RI e \<Longrightarrow> iconf sh e)"
+ by(simp add: nsub_RI_iconf_aux)
+
+lemma nsub_RI_iconf: "\<not>sub_RI e \<Longrightarrow> iconf sh e"
+ and nsub_RIs_iconfs: "\<not>sub_RIs es \<Longrightarrow> iconfs sh es"
+proof -
+ let ?R = "\<lambda>e. \<not>sub_RI e \<longrightarrow> iconf sh e"
+ let ?Rs = "\<lambda>es. \<not>sub_RIs es \<longrightarrow> iconfs sh es"
+ have "(\<forall>e'. subexp_of e' e \<longrightarrow> ?R e') \<and> ?R e"
+ by(rule subexp_induct[where ?Rs = ?Rs]; clarsimp simp: nsub_RI_iconf_aux)
+ moreover have "(\<forall>e'. e' \<in> subexps es \<longrightarrow> ?R e') \<and> ?Rs es"
+ by(rule subexps_induct; clarsimp simp: nsub_RI_iconf_aux)
+ ultimately show "\<not>sub_RI e \<Longrightarrow> iconf sh e"
+ and "\<not>sub_RIs es \<Longrightarrow> iconfs sh es" by simp+
+qed
+
+lemma lass_val_of_iconf: "lass_val_of e = \<lfloor>a\<rfloor> \<Longrightarrow> iconf sh e"
+ by(drule lass_val_of_nsub_RI, erule nsub_RI_iconf)
+
+lemma icheck_iconf:
+assumes "icheck P C e" shows "iconf sh e"
+using assms
+proof(cases e)
+ case (SFAss C F D e) then show ?thesis using assms
+ proof(cases e)qed(auto)
+next
+ case (SCall C M es) then show ?thesis using assms
+ by (auto simp: nsub_RIs_iconfs)
+next
+qed(auto)
+
+
+subsection "Indicator boolean conformance"
+
+\<comment> \<open> checks that the given expression, indicator boolean pair is allowed in small-step
+ (i.e., if @{term b} is True, then @{term e} is an initialization-calling expression to
+ a class that is marked either @{term Processing} or @{term Done}) \<close>
+definition bconf :: "'m prog \<Rightarrow> sheap \<Rightarrow> 'a exp \<Rightarrow> bool \<Rightarrow> bool" ("_,_ \<turnstile>\<^sub>b '(_,_') \<surd>" [51,51,0,0] 50)
+where
+ "P,sh \<turnstile>\<^sub>b (e,b) \<surd> \<equiv> b \<longrightarrow> (\<exists>C. icheck P C (ss_exp e) \<and> initPD sh C)"
+
+definition bconfs :: "'m prog \<Rightarrow> sheap \<Rightarrow> 'a exp list \<Rightarrow> bool \<Rightarrow> bool" ("_,_ \<turnstile>\<^sub>b '(_,_') \<surd>" [51,51,0,0] 50)
+where
+ "P,sh \<turnstile>\<^sub>b (es,b) \<surd> \<equiv> b \<longrightarrow> (\<exists>C. (icheck P C (the(ss_exps es))
+ \<and> (curr_inits P es = Some C) \<and> initPD sh C))"
+
+
+\<comment> \<open> bconf helper lemmas \<close>
+
+lemma bconf_nonVal[simp]:
+ "P,sh \<turnstile>\<^sub>b (e,True) \<surd> \<Longrightarrow> val_of e = None"
+ by(cases e) (auto simp: bconf_def)
+
+lemma bconfs_nonVals[simp]:
+ "P,sh \<turnstile>\<^sub>b (es,True) \<surd> \<Longrightarrow> map_vals_of es = None"
+ by(induct es) (auto simp: bconfs_def)
+
+lemma bconf_Cast[iff]:
+ "P,sh \<turnstile>\<^sub>b (Cast C e,b) \<surd> \<longleftrightarrow> P,sh \<turnstile>\<^sub>b (e,b) \<surd>"
+ by(cases b) (auto simp: bconf_def dest: val_of_spec)
+
+lemma bconf_BinOp[iff]:
+ "P,sh \<turnstile>\<^sub>b (e1 \<guillemotleft>bop\<guillemotright> e2,b) \<surd>
+ \<longleftrightarrow> (case val_of e1 of Some v \<Rightarrow> P,sh \<turnstile>\<^sub>b (e2,b) \<surd> | _ \<Rightarrow> P,sh \<turnstile>\<^sub>b (e1,b) \<surd>)"
+ by(cases b) (auto simp: bconf_def dest: val_of_spec)
+
+lemma bconf_LAss[iff]:
+ "P,sh \<turnstile>\<^sub>b (LAss V e,b) \<surd> \<longleftrightarrow> P,sh \<turnstile>\<^sub>b (e,b) \<surd>"
+ by(cases b) (auto simp: bconf_def dest: val_of_spec)
+
+lemma bconf_FAcc[iff]:
+ "P,sh \<turnstile>\<^sub>b (e\<bullet>F{D},b) \<surd> \<longleftrightarrow> P,sh \<turnstile>\<^sub>b (e,b) \<surd>"
+ by(cases b) (auto simp: bconf_def dest: val_of_spec)
+
+lemma bconf_FAss[iff]:
+ "P,sh \<turnstile>\<^sub>b (FAss e1 F D e2,b) \<surd>
+ \<longleftrightarrow> (case val_of e1 of Some v \<Rightarrow> P,sh \<turnstile>\<^sub>b (e2,b) \<surd> | _ \<Rightarrow> P,sh \<turnstile>\<^sub>b (e1,b) \<surd>)"
+ by(cases b) (auto simp: bconf_def dest: val_of_spec)
+
+lemma bconf_SFAss[iff]:
+"val_of e2 = None \<Longrightarrow> P,sh \<turnstile>\<^sub>b (SFAss C F D e2,b) \<surd> \<longleftrightarrow> P,sh \<turnstile>\<^sub>b (e2,b) \<surd>"
+ by(cases b) (auto simp: bconf_def)
+
+lemma bconfs_Vals[iff]:
+ "P,sh \<turnstile>\<^sub>b (map Val vs, b) \<surd> \<longleftrightarrow> \<not> b"
+ by(unfold bconfs_def) simp
+
+lemma bconf_Call[iff]:
+ "P,sh \<turnstile>\<^sub>b (e\<bullet>M(es),b) \<surd>
+ \<longleftrightarrow> (case val_of e of Some v \<Rightarrow> P,sh \<turnstile>\<^sub>b (es,b) \<surd> | _ \<Rightarrow> P,sh \<turnstile>\<^sub>b (e,b) \<surd>)"
+proof(cases b)
+ case True
+ then show ?thesis
+ proof(cases "ss_exps es")
+ case None
+ then obtain vs where "es = map Val vs" using ss_exps_Vals_NoneI by auto
+ then have mv: "map_vals_of es = \<lfloor>vs\<rfloor>" by simp
+ then show ?thesis by(auto simp: bconf_def) (simp add: bconfs_def)
+ next
+ case (Some a)
+ then show ?thesis by(auto simp: bconf_def) (auto simp: bconfs_def icheck_init_class)
+ qed
+qed(simp add: bconf_def bconfs_def)
+
+lemma bconf_SCall[iff]:
+assumes mvn: "map_vals_of es = None"
+shows "P,sh \<turnstile>\<^sub>b (C\<bullet>\<^sub>sM(es),b) \<surd> \<longleftrightarrow> P,sh \<turnstile>\<^sub>b (es,b) \<surd>"
+proof(cases b)
+ case True
+ then show ?thesis
+ proof(cases "ss_exps es")
+ case None
+ then have "\<exists>vs. es = map Val vs" using ss_exps_Vals_NoneI by auto
+ then show ?thesis using mvn finals_def by clarsimp
+ next
+ case (Some a)
+ then show ?thesis by(auto simp: bconf_def) (auto simp: bconfs_def icheck_init_class)
+ qed
+qed(simp add: bconf_def bconfs_def)
+
+lemma bconf_Cons[iff]:
+ "P,sh \<turnstile>\<^sub>b (e#es,b) \<surd>
+ \<longleftrightarrow> (case val_of e of Some v \<Rightarrow> P,sh \<turnstile>\<^sub>b (es,b) \<surd> | _ \<Rightarrow> P,sh \<turnstile>\<^sub>b (e,b) \<surd>)"
+proof(cases b)
+ case True
+ then show ?thesis
+ proof(cases "ss_exps es")
+ case None
+ then have "\<exists>vs. es = map Val vs" using ss_exps_Vals_NoneI by auto
+ then show ?thesis using None by(auto simp: bconf_def bconfs_def icheck_init_class)
+ next
+ case (Some a)
+ then show ?thesis by(auto simp: bconf_def bconfs_def icheck_init_class)
+ qed
+qed(simp add: bconf_def bconfs_def)
+
+lemma bconf_InitBlock[iff]:
+ "P,sh \<turnstile>\<^sub>b ({V:T; V:=Val v;; e\<^sub>2},b) \<surd> \<longleftrightarrow> P,sh \<turnstile>\<^sub>b (e\<^sub>2,b) \<surd>"
+ by(cases b) (auto simp: bconf_def assigned_def)
+
+lemma bconf_Block[iff]:
+ "P,sh \<turnstile>\<^sub>b ({V:T; e},b) \<surd> \<longleftrightarrow> P,sh \<turnstile>\<^sub>b (e,b) \<surd>"
+ by(cases b) (auto simp: bconf_def)
+
+lemma bconf_Seq[iff]:
+ "P,sh \<turnstile>\<^sub>b (e1;;e2,b) \<surd>
+ \<longleftrightarrow> (case val_of e1 of Some v \<Rightarrow> P,sh \<turnstile>\<^sub>b (e2,b) \<surd>
+ | _ \<Rightarrow> (case lass_val_of e1 of Some p \<Rightarrow> P,sh \<turnstile>\<^sub>b (e2,b) \<surd>
+ | None \<Rightarrow> P,sh \<turnstile>\<^sub>b (e1,b) \<surd>))"
+ by(cases b) (auto simp: bconf_def dest: val_of_spec lass_val_of_spec)
+
+lemma bconf_Cond[iff]:
+ "P,sh \<turnstile>\<^sub>b (if (b) e\<^sub>1 else e\<^sub>2,b') \<surd> \<longleftrightarrow> P,sh \<turnstile>\<^sub>b (b,b') \<surd>"
+proof(cases "bool_of b")
+ case None
+ then show ?thesis by(auto simp: bconf_def)
+next
+ case (Some a)
+ then show ?thesis by(case_tac a) (auto simp: bconf_def dest: bool_of_specT bool_of_specF)
+qed
+
+lemma bconf_While[iff]:
+ "P,sh \<turnstile>\<^sub>b (while (b) e,b') \<surd> \<longleftrightarrow> \<not>b'"
+ by(cases b) (auto simp: bconf_def)
+
+lemma bconf_Throw[iff]:
+ "P,sh \<turnstile>\<^sub>b (throw e,b) \<surd> \<longleftrightarrow> P,sh \<turnstile>\<^sub>b (e,b) \<surd>"
+ by(cases b) (auto simp: bconf_def dest: val_of_spec)
+
+lemma bconf_Try[iff]:
+ "P,sh \<turnstile>\<^sub>b (try e\<^sub>1 catch(C V) e\<^sub>2,b) \<surd> \<longleftrightarrow> P,sh \<turnstile>\<^sub>b (e\<^sub>1,b) \<surd>"
+ by(cases b) (auto simp: bconf_def dest: val_of_spec)
+
+lemma bconf_INIT[iff]:
+ "P,sh \<turnstile>\<^sub>b (INIT C (Cs,b') \<leftarrow> e,b) \<surd> \<longleftrightarrow> \<not>b"
+ by(cases b) (auto simp: bconf_def)
+
+lemma bconf_RI[iff]:
+ "P,sh \<turnstile>\<^sub>b (RI(C,e);Cs \<leftarrow> e',b) \<surd> \<longleftrightarrow> P,sh \<turnstile>\<^sub>b (e,b) \<surd>"
+ by(cases b) (auto simp: bconf_def dest: val_of_spec)
+
+lemma bconfs_map_throw[iff]:
+ "P,sh \<turnstile>\<^sub>b (map Val vs @ throw e # es',b) \<surd> \<longleftrightarrow> P,sh \<turnstile>\<^sub>b (e,b) \<surd>"
+ by(induct vs) auto
+
+end
diff --git a/thys/JinjaDCI/J/Equivalence.thy b/thys/JinjaDCI/J/Equivalence.thy
--- a/thys/JinjaDCI/J/Equivalence.thy
+++ b/thys/JinjaDCI/J/Equivalence.thy
@@ -1,4191 +1,4191 @@
-(* Title: JinjaDCI/J/Equivalence.thy
- Author: Tobias Nipkow, Susannah Mansky
- Copyright 2003 Technische Universitaet Muenchen, 2019-20 UIUC
-
- Based on the Jinja theory J/Equivalence.thy by Tobias Nipkow
-*)
-
-section \<open> Equivalence of Big Step and Small Step Semantics \<close>
-
-theory Equivalence imports TypeSafe WWellForm begin
-
-subsection\<open>Small steps simulate big step\<close>
-
-subsubsection "Init"
-
-text "The reduction of initialization expressions doesn't touch or use
- their on-hold expressions (the subexpression to the right of @{text \<leftarrow>})
- until the initialization operation completes. This function is used to prove
- this and related properties. It is then used for general reduction of
- initialization expressions separately from their on-hold expressions by
- using the on-hold expression @{term unit}, then putting the real on-hold
- expression back in at the end."
-
-fun init_switch :: "'a exp \<Rightarrow> 'a exp \<Rightarrow> 'a exp" where
-"init_switch (INIT C (Cs,b) \<leftarrow> e\<^sub>i) e = (INIT C (Cs,b) \<leftarrow> e)" |
-"init_switch (RI(C,e');Cs \<leftarrow> e\<^sub>i) e = (RI(C,e');Cs \<leftarrow> e)" |
-"init_switch e' e = e'"
-
-fun INIT_class :: "'a exp \<Rightarrow> cname option" where
-"INIT_class (INIT C (Cs,b) \<leftarrow> e) = (if C = last (C#Cs) then Some C else None)" |
-"INIT_class (RI(C,e\<^sub>0);Cs \<leftarrow> e) = Some (last (C#Cs))" |
-"INIT_class _ = None"
-
-lemma init_red_init:
-"\<lbrakk> init_exp_of e\<^sub>0 = \<lfloor>e\<rfloor>; P \<turnstile> \<langle>e\<^sub>0,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow> \<langle>e\<^sub>1,s\<^sub>1,b\<^sub>1\<rangle> \<rbrakk>
- \<Longrightarrow> (init_exp_of e\<^sub>1 = \<lfloor>e\<rfloor> \<and> (INIT_class e\<^sub>0 = \<lfloor>C\<rfloor> \<longrightarrow> INIT_class e\<^sub>1 = \<lfloor>C\<rfloor>))
- \<or> (e\<^sub>1 = e \<and> b\<^sub>1 = icheck P (the(INIT_class e\<^sub>0)) e) \<or> (\<exists>a. e\<^sub>1 = throw a)"
- by(erule red.cases, auto)
-
-lemma init_exp_switch[simp]:
-"init_exp_of e\<^sub>0 = \<lfloor>e\<rfloor> \<Longrightarrow> init_exp_of (init_switch e\<^sub>0 e') = \<lfloor>e'\<rfloor>"
- by(cases e\<^sub>0, simp_all)
-
-lemma init_red_sync:
-"\<lbrakk> P \<turnstile> \<langle>e\<^sub>0,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow> \<langle>e\<^sub>1,s\<^sub>1,b\<^sub>1\<rangle>; init_exp_of e\<^sub>0 = \<lfloor>e\<rfloor>; e\<^sub>1 \<noteq> e \<rbrakk>
- \<Longrightarrow> (\<And>e'. P \<turnstile> \<langle>init_switch e\<^sub>0 e',s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow> \<langle>init_switch e\<^sub>1 e',s\<^sub>1,b\<^sub>1\<rangle>)"
-proof(induct rule: red.cases) qed(simp_all add: red_reds.intros)
-
-lemma init_red_sync_end:
-"\<lbrakk> P \<turnstile> \<langle>e\<^sub>0,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow> \<langle>e\<^sub>1,s\<^sub>1,b\<^sub>1\<rangle>; init_exp_of e\<^sub>0 = \<lfloor>e\<rfloor>; e\<^sub>1 = e; throw_of e = None \<rbrakk>
- \<Longrightarrow> (\<And>e'. \<not>sub_RI e' \<Longrightarrow> P \<turnstile> \<langle>init_switch e\<^sub>0 e',s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow> \<langle>e',s\<^sub>1,icheck P (the(INIT_class e\<^sub>0)) e'\<rangle>)"
-proof(induct rule: red.cases) qed(simp_all add: red_reds.intros)
-
-lemma init_reds_sync_unit':
- "\<lbrakk> P \<turnstile> \<langle>e\<^sub>0,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>Val v',s\<^sub>1,b\<^sub>1\<rangle>; init_exp_of e\<^sub>0 = \<lfloor>unit\<rfloor>; INIT_class e\<^sub>0 = \<lfloor>C\<rfloor> \<rbrakk>
- \<Longrightarrow> (\<And>e'. \<not>sub_RI e' \<Longrightarrow> P \<turnstile> \<langle>init_switch e\<^sub>0 e',s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>e',s\<^sub>1,icheck P (the(INIT_class e\<^sub>0)) e'\<rangle>)"
-proof(induct rule:converse_rtrancl_induct3)
-case refl then show ?case by simp
-next
- case (step e0 s0 b0 e1 s1 b1)
- have "(init_exp_of e1 = \<lfloor>unit\<rfloor> \<and> (INIT_class e0 = \<lfloor>C\<rfloor> \<longrightarrow> INIT_class e1 = \<lfloor>C\<rfloor>))
- \<or> (e1 = unit \<and> b1 = icheck P (the(INIT_class e0)) unit) \<or> (\<exists>a. e1 = throw a)"
- using init_red_init[OF step.prems(1) step.hyps(1)] by simp
- then show ?case
- proof(rule disjE)
- assume assm: "init_exp_of e1 = \<lfloor>unit\<rfloor> \<and> (INIT_class e0 = \<lfloor>C\<rfloor> \<longrightarrow> INIT_class e1 = \<lfloor>C\<rfloor>)"
- then have red: "P \<turnstile> \<langle>init_switch e0 e',s0,b0\<rangle> \<rightarrow> \<langle>init_switch e1 e',s1,b1\<rangle>"
- using init_red_sync[OF step.hyps(1) step.prems(1)] by simp
- have reds: "P \<turnstile> \<langle>init_switch e1 e',s1,b1\<rangle> \<rightarrow>* \<langle>e',s\<^sub>1,icheck P (the (INIT_class e0)) e'\<rangle>"
- using step.hyps(3)[OF _ _ step.prems(3)] assm step.prems(2) by simp
- show ?thesis by(rule converse_rtrancl_into_rtrancl[OF red reds])
- next
- assume "(e1 = unit \<and> b1 = icheck P (the(INIT_class e0)) unit) \<or> (\<exists>a. e1 = throw a)"
- then show ?thesis
- proof(rule disjE)
- assume assm: "e1 = unit \<and> b1 = icheck P (the(INIT_class e0)) unit" then have e1: "e1 = unit" by simp
- have sb: "s1 = s\<^sub>1" "b1 = b\<^sub>1" using reds_final_same[OF step.hyps(2)] assm
- by(simp_all add: final_def)
- then have step': "P \<turnstile> \<langle>init_switch e0 e',s0,b0\<rangle> \<rightarrow> \<langle>e',s\<^sub>1,icheck P (the (INIT_class e0)) e'\<rangle>"
- using init_red_sync_end[OF step.hyps(1) step.prems(1) e1 _ step.prems(3)] by auto
- then have "P \<turnstile> \<langle>init_switch e0 e',s0,b0\<rangle> \<rightarrow>* \<langle>e',s\<^sub>1,icheck P (the (INIT_class e0)) e'\<rangle>"
- using r_into_rtrancl by auto
- then show ?thesis using step assm sb by simp
- next
- assume "\<exists>a. e1 = throw a" then obtain a where "e1 = throw a" by clarsimp
- then have tof: "throw_of e1 = \<lfloor>a\<rfloor>" by simp
- then show ?thesis using reds_throw[OF step.hyps(2) tof] by simp
- qed
- qed
-qed
-
-lemma init_reds_sync_unit_throw':
- "\<lbrakk> P \<turnstile> \<langle>e\<^sub>0,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>throw a,s\<^sub>1,b\<^sub>1\<rangle>; init_exp_of e\<^sub>0 = \<lfloor>unit\<rfloor> \<rbrakk>
- \<Longrightarrow> (\<And>e'. P \<turnstile> \<langle>init_switch e\<^sub>0 e',s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>throw a,s\<^sub>1,b\<^sub>1\<rangle>)"
-proof(induct rule:converse_rtrancl_induct3)
-case refl then show ?case by simp
-next
- case (step e0 s0 b0 e1 s1 b1)
- have "init_exp_of e1 = \<lfloor>unit\<rfloor> \<and> (\<forall>C. INIT_class e0 = \<lfloor>C\<rfloor> \<longrightarrow> INIT_class e1 = \<lfloor>C\<rfloor>) \<or>
- e1 = unit \<and> b1 = icheck P (the (INIT_class e0)) unit \<or> (\<exists>a. e1 = throw a)"
- using init_red_init[OF step.prems(1) step.hyps(1)] by auto
- then show ?case
- proof(rule disjE)
- assume assm: "init_exp_of e1 = \<lfloor>unit\<rfloor> \<and> (\<forall>C. INIT_class e0 = \<lfloor>C\<rfloor> \<longrightarrow> INIT_class e1 = \<lfloor>C\<rfloor>)"
- then have "P \<turnstile> \<langle>init_switch e0 e',s0,b0\<rangle> \<rightarrow> \<langle>init_switch e1 e',s1,b1\<rangle>"
- using step init_red_sync[OF step.hyps(1) step.prems] by simp
- then show ?thesis using step assm by (meson converse_rtrancl_into_rtrancl)
- next
- assume "e1 = unit \<and> b1 = icheck P (the (INIT_class e0)) unit \<or> (\<exists>a. e1 = throw a)"
- then show ?thesis
- proof(rule disjE)
- assume "e1 = unit \<and> b1 = icheck P (the (INIT_class e0)) unit"
- then show ?thesis using step using final_def reds_final_same by blast
- next
- assume assm: "\<exists>a. e1 = throw a"
- then have "P \<turnstile> \<langle>init_switch e0 e',s0,b0\<rangle> \<rightarrow> \<langle>e1,s1,b1\<rangle>"
- using init_red_sync[OF step.hyps(1) step.prems] by clarsimp
- then show ?thesis using step by simp
- qed
- qed
-qed
-
-lemma init_reds_sync_unit:
-assumes "P \<turnstile> \<langle>e\<^sub>0,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>Val v',s\<^sub>1,b\<^sub>1\<rangle>" and "init_exp_of e\<^sub>0 = \<lfloor>unit\<rfloor>" and "INIT_class e\<^sub>0 = \<lfloor>C\<rfloor>"
- and "\<not>sub_RI e'"
-shows "P \<turnstile> \<langle>init_switch e\<^sub>0 e',s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>e',s\<^sub>1,icheck P (the(INIT_class e\<^sub>0)) e'\<rangle>"
-using init_reds_sync_unit'[OF assms] by clarsimp
-
-lemma init_reds_sync_unit_throw:
-assumes "P \<turnstile> \<langle>e\<^sub>0,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>throw a,s\<^sub>1,b\<^sub>1\<rangle>" and "init_exp_of e\<^sub>0 = \<lfloor>unit\<rfloor>"
-shows "P \<turnstile> \<langle>init_switch e\<^sub>0 e',s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>throw a,s\<^sub>1,b\<^sub>1\<rangle>"
-using init_reds_sync_unit_throw'[OF assms] by clarsimp
-
-\<comment> \<open> init reds lemmas \<close>
-
-lemma InitSeqReds:
-assumes "P \<turnstile> \<langle>INIT C ([C],b) \<leftarrow> unit,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>Val v',s\<^sub>1,b\<^sub>1\<rangle>"
- and "P \<turnstile> \<langle>e,s\<^sub>1,icheck P C e\<rangle> \<rightarrow>* \<langle>e\<^sub>2,s\<^sub>2,b\<^sub>2\<rangle>" and "\<not>sub_RI e"
-shows "P \<turnstile> \<langle>INIT C ([C],b) \<leftarrow> e,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>e\<^sub>2,s\<^sub>2,b\<^sub>2\<rangle>"
-using assms
-proof -
- have "P \<turnstile> \<langle>init_switch (INIT C ([C],b) \<leftarrow> unit) e,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>e,s\<^sub>1,icheck P C e\<rangle>"
- using init_reds_sync_unit[OF assms(1) _ _ assms(3)] by simp
- then show ?thesis using assms(2) by simp
-qed
-
-lemma InitSeqThrowReds:
-assumes "P \<turnstile> \<langle>INIT C ([C],b) \<leftarrow> unit,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>throw a,s\<^sub>1,b\<^sub>1\<rangle>"
-shows "P \<turnstile> \<langle>INIT C ([C],b) \<leftarrow> e,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>throw a,s\<^sub>1,b\<^sub>1\<rangle>"
-using assms
-proof -
- have "P \<turnstile> \<langle>init_switch (INIT C ([C],b) \<leftarrow> unit) e,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>throw a,s\<^sub>1,b\<^sub>1\<rangle>"
- using init_reds_sync_unit_throw[OF assms(1)] by simp
- then show ?thesis by simp
-qed
-
-lemma InitNoneReds:
- "\<lbrakk> sh C = None;
- P \<turnstile> \<langle>INIT C' (C # Cs,False) \<leftarrow> e,(h, l, sh(C \<mapsto> (sblank P C, Prepared))),b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle> \<rbrakk>
-\<Longrightarrow> P \<turnstile> \<langle>INIT C' (C#Cs,False) \<leftarrow> e,(h,l,sh),b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle>"
-(*<*)by(rule InitNoneRed[THEN converse_rtrancl_into_rtrancl])(*>*)
-
-lemma InitDoneReds:
- "\<lbrakk> sh C = Some(sfs,Done); P \<turnstile> \<langle>INIT C' (Cs,True) \<leftarrow> e,(h,l,sh),b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle> \<rbrakk>
-\<Longrightarrow> P \<turnstile> \<langle>INIT C' (C#Cs,False) \<leftarrow> e,(h,l,sh),b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle>"
-(*<*)by(rule RedInitDone[THEN converse_rtrancl_into_rtrancl])(*>*)
-
-lemma InitProcessingReds:
- "\<lbrakk> sh C = Some(sfs,Processing); P \<turnstile> \<langle>INIT C' (Cs,True) \<leftarrow> e,(h,l,sh),b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle> \<rbrakk>
-\<Longrightarrow> P \<turnstile> \<langle>INIT C' (C#Cs,False) \<leftarrow> e,(h,l,sh),b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle>"
-(*<*)by(rule RedInitProcessing[THEN converse_rtrancl_into_rtrancl])(*>*)
-
-lemma InitErrorReds:
- "\<lbrakk> sh C = Some(sfs,Error); P \<turnstile> \<langle>RI (C,THROW NoClassDefFoundError);Cs \<leftarrow> e,(h,l,sh),b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle> \<rbrakk>
-\<Longrightarrow> P \<turnstile> \<langle>INIT C' (C#Cs,False) \<leftarrow> e,(h,l,sh),b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle>"
-(*<*)by(rule RedInitError[THEN converse_rtrancl_into_rtrancl])(*>*)
-
-lemma InitObjectReds:
- "\<lbrakk> sh C = Some(sfs,Prepared); C = Object; sh' = sh(C \<mapsto> (sfs,Processing));
- P \<turnstile> \<langle>INIT C' (C#Cs,True) \<leftarrow> e,(h,l,sh'),b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle> \<rbrakk>
-\<Longrightarrow> P \<turnstile> \<langle>INIT C' (C#Cs,False) \<leftarrow> e,(h,l,sh),b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle>"
-(*<*)by(rule InitObjectRed[THEN converse_rtrancl_into_rtrancl])(*>*)
-
-lemma InitNonObjectReds:
- "\<lbrakk> sh C = Some(sfs,Prepared); C \<noteq> Object; class P C = Some (D,r);
- sh' = sh(C \<mapsto> (sfs,Processing));
- P \<turnstile> \<langle>INIT C' (D#C#Cs,False) \<leftarrow> e,(h,l,sh'),b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle> \<rbrakk>
-\<Longrightarrow> P \<turnstile> \<langle>INIT C' (C#Cs,False) \<leftarrow> e,(h,l,sh),b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle>"
-(*<*)by(rule InitNonObjectSuperRed[THEN converse_rtrancl_into_rtrancl])(*>*)
-
-lemma RedsInitRInit:
- "P \<turnstile> \<langle>RI (C,C\<bullet>\<^sub>sclinit([]));Cs \<leftarrow> e,(h,l,sh),b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle>
-\<Longrightarrow> P \<turnstile> \<langle>INIT C' (C#Cs,True) \<leftarrow> e,(h,l,sh),b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle>"
-(*<*)by(rule RedInitRInit[THEN converse_rtrancl_into_rtrancl])(*>*)
-
-lemmas rtrancl_induct3 =
- rtrancl_induct[of "(ax,ay,az)" "(bx,by,bz)", split_format (complete), consumes 1, case_names refl step]
-
-lemma RInitReds:
- "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle>
-\<Longrightarrow> P \<turnstile> \<langle>RI (C,e);Cs \<leftarrow> e\<^sub>0, s, b\<rangle> \<rightarrow>* \<langle>RI (C,e');Cs \<leftarrow> e\<^sub>0, s', b'\<rangle>"
-(*<*)
-proof(induct rule: rtrancl_induct3)
- case refl show ?case by blast
-next
- case step show ?case
- by(rule rtrancl_into_rtrancl[OF step(3) RInitRed[OF step(2)]])
-qed
-(*>*)
-
-lemma RedsRInit:
- "\<lbrakk> P \<turnstile> \<langle>e\<^sub>0,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>Val v,(h\<^sub>1,l\<^sub>1,sh\<^sub>1),b\<^sub>1\<rangle>;
- sh\<^sub>1 C = Some (sfs, i); sh\<^sub>2 = sh\<^sub>1(C \<mapsto> (sfs,Done)); C' = last(C#Cs);
- P \<turnstile> \<langle>INIT C' (Cs,True) \<leftarrow> e,(h\<^sub>1,l\<^sub>1,sh\<^sub>2),b\<^sub>1\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle> \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>RI (C, e\<^sub>0);Cs \<leftarrow> e,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle>"
-(*<*)
-by(rule rtrancl_trans[OF RInitReds
- RedRInit[THEN converse_rtrancl_into_rtrancl]])
-(*>*)
-
-
-lemma RInitInitThrowReds:
- "\<lbrakk> P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>Throw a,(h',l',sh'),b'\<rangle>;
- sh' C = Some (sfs, i); sh'' = sh'(C \<mapsto> (sfs,Error));
- P \<turnstile> \<langle>RI (D,Throw a);Cs \<leftarrow> e\<^sub>0, (h',l',sh''),b'\<rangle> \<rightarrow>* \<langle>e\<^sub>1,s\<^sub>1,b\<^sub>1\<rangle> \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>RI (C, e);D#Cs \<leftarrow> e\<^sub>0,s,b\<rangle> \<rightarrow>* \<langle>e\<^sub>1,s\<^sub>1,b\<^sub>1\<rangle>"
-(*<*)
-by(rule rtrancl_trans[OF RInitReds
- RInitInitThrow[THEN converse_rtrancl_into_rtrancl]])
-(*>*)
-
-lemma RInitThrowReds:
- "\<lbrakk> P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>Throw a, (h',l',sh'),b'\<rangle>;
- sh' C = Some(sfs, i); sh'' = sh'(C \<mapsto> (sfs, Error)) \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>RI (C,e);Nil \<leftarrow> e\<^sub>0,s,b\<rangle> \<rightarrow>* \<langle>Throw a, (h',l',sh''),b'\<rangle>"
-(*<*)by(rule RInitReds[THEN rtrancl_into_rtrancl, OF _ RInitThrow])(*>*)
-
-subsubsection "New"
-
-lemma NewInitDoneReds:
- "\<lbrakk> sh C = Some (sfs, Done); new_Addr h = Some a;
- P \<turnstile> C has_fields FDTs; h' = h(a\<mapsto>blank P C) \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>new C,(h,l,sh),False\<rangle> \<rightarrow>* \<langle>addr a,(h',l,sh),False\<rangle>"
-(*<*)
-by(rule NewInitDoneRed[THEN converse_rtrancl_into_rtrancl,
- OF _ RedNew[THEN r_into_rtrancl]])
-(*>*)
-
-lemma NewInitDoneReds2:
- "\<lbrakk> sh C = Some (sfs, Done); new_Addr h = None; is_class P C \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>new C,(h,l,sh),False\<rangle> \<rightarrow>* \<langle>THROW OutOfMemory, (h,l,sh), False\<rangle>"
-(*<*)
-by(rule NewInitDoneRed[THEN converse_rtrancl_into_rtrancl,
- OF _ RedNewFail[THEN r_into_rtrancl]])
-(*>*)
-
-lemma NewInitReds:
-assumes nDone: "\<nexists>sfs. shp s C = Some (sfs, Done)"
- and INIT_steps: "P \<turnstile> \<langle>INIT C ([C],False) \<leftarrow> unit,s,False\<rangle> \<rightarrow>* \<langle>Val v',(h',l',sh'),b'\<rangle>"
- and Addr: "new_Addr h' = Some a" and has: "P \<turnstile> C has_fields FDTs"
- and h'': "h'' = h'(a\<mapsto>blank P C)" and cls_C: "is_class P C"
-shows "P \<turnstile> \<langle>new C,s,False\<rangle> \<rightarrow>* \<langle>addr a,(h'',l',sh'),False\<rangle>"
-(*<*)(is "(?a, ?c) \<in> (red P)\<^sup>*")
-proof -
- let ?b = "(INIT C ([C],False) \<leftarrow> new C,s,False)"
- let ?b' = "(new C,(h', l', sh'),icheck P C (new C::expr))"
- have b'c: "(?b', ?c) \<in> (red P)\<^sup>*"
- using RedNew[OF Addr has h'', THEN r_into_rtrancl] by simp
- obtain h l sh where [simp]: "s = (h, l, sh)" by(cases s) simp
- have "(?a, ?b) \<in> (red P)\<^sup>*"
- using NewInitRed[OF _ cls_C] nDone by fastforce
- also have "(?b, ?c) \<in> (red P)\<^sup>*"
- by(rule InitSeqReds[OF INIT_steps b'c]) simp
- ultimately show ?thesis by simp
-qed
-(*>*)
-
-lemma NewInitOOMReds:
-assumes nDone: "\<nexists>sfs. shp s C = Some (sfs, Done)"
- and INIT_steps: "P \<turnstile> \<langle>INIT C ([C],False) \<leftarrow> unit,s,False\<rangle> \<rightarrow>* \<langle>Val v',(h',l',sh'),b'\<rangle>"
- and Addr: "new_Addr h' = None" and cls_C: "is_class P C"
-shows "P \<turnstile> \<langle>new C,s,False\<rangle> \<rightarrow>* \<langle>THROW OutOfMemory,(h',l',sh'),False\<rangle>"
-(*<*)(is "(?a, ?c) \<in> (red P)\<^sup>*")
-proof -
- let ?b = "(INIT C ([C],False) \<leftarrow> new C,s,False)"
- let ?b' = "(new C,(h', l', sh'),icheck P C (new C::expr))"
- have b'c: "(?b', ?c) \<in> (red P)\<^sup>*"
- using RedNewFail[OF Addr cls_C, THEN r_into_rtrancl] by simp
- obtain h l sh where [simp]: "s = (h, l, sh)" by(cases s) simp
- have "(?a, ?b) \<in> (red P)\<^sup>*"
- using NewInitRed[OF _ cls_C] nDone by fastforce
- also have "(?b, ?c) \<in> (red P)\<^sup>*"
- by(rule InitSeqReds[OF INIT_steps b'c]) simp
- ultimately show ?thesis by simp
-qed
-(*>*)
-
-lemma NewInitThrowReds:
-assumes nDone: "\<nexists>sfs. shp s C = Some (sfs, Done)"
- and cls_C: "is_class P C"
- and INIT_steps: "P \<turnstile> \<langle>INIT C ([C],False) \<leftarrow> unit,s,False\<rangle> \<rightarrow>* \<langle>throw a,s',b'\<rangle>"
-shows "P \<turnstile> \<langle>new C,s,False\<rangle> \<rightarrow>* \<langle>throw a,s',b'\<rangle>"
-(*<*)(is "(?a, ?c) \<in> (red P)\<^sup>*")
-proof -
- let ?b = "(INIT C ([C],False) \<leftarrow> new C,s,False)"
- obtain h l sh where [simp]: "s = (h, l, sh)" by(cases s) simp
- have "(?a, ?b) \<in> (red P)\<^sup>*"
- using NewInitRed[OF _ cls_C] nDone by fastforce
- also have "(?b, ?c) \<in> (red P)\<^sup>*"
- using InitSeqThrowReds[OF INIT_steps] by simp
- ultimately show ?thesis by simp
-qed
-(*>*)
-
-subsubsection "Cast"
-
-lemma CastReds:
- "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>Cast C e,s,b\<rangle> \<rightarrow>* \<langle>Cast C e',s',b'\<rangle>"
-(*<*)
-proof(induct rule: rtrancl_induct3)
- case refl show ?case by blast
-next
- case step show ?case
- by(rule rtrancl_into_rtrancl[OF step(3) CastRed[OF step(2)]])
-qed
-(*>*)
-
-lemma CastRedsNull:
- "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>null,s',b'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>Cast C e,s,b\<rangle> \<rightarrow>* \<langle>null,s',b'\<rangle>"
-(*<*)by(rule CastReds[THEN rtrancl_into_rtrancl, OF _ RedCastNull])(*>*)
-
-lemma CastRedsAddr:
- "\<lbrakk> P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>addr a,s',b'\<rangle>; hp s' a = Some(D,fs); P \<turnstile> D \<preceq>\<^sup>* C \<rbrakk> \<Longrightarrow>
- P \<turnstile> \<langle>Cast C e,s,b\<rangle> \<rightarrow>* \<langle>addr a,s',b'\<rangle>"
-(*<*)by(cases s', simp) (rule CastReds[THEN rtrancl_into_rtrancl, OF _ RedCast])(*>*)
-
-lemma CastRedsFail:
- "\<lbrakk> P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>addr a,s',b'\<rangle>; hp s' a = Some(D,fs); \<not> P \<turnstile> D \<preceq>\<^sup>* C \<rbrakk> \<Longrightarrow>
- P \<turnstile> \<langle>Cast C e,s,b\<rangle> \<rightarrow>* \<langle>THROW ClassCast,s',b'\<rangle>"
-(*<*)by(cases s', simp) (rule CastReds[THEN rtrancl_into_rtrancl, OF _ RedCastFail])(*>*)
-
-lemma CastRedsThrow:
- "\<lbrakk> P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>throw a,s',b'\<rangle> \<rbrakk> \<Longrightarrow> P \<turnstile> \<langle>Cast C e,s,b\<rangle> \<rightarrow>* \<langle>throw a,s',b'\<rangle>"
-(*<*)by(rule CastReds[THEN rtrancl_into_rtrancl, OF _ red_reds.CastThrow])(*>*)
-
-subsubsection "LAss"
-
-lemma LAssReds:
- "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle> \<Longrightarrow> P \<turnstile> \<langle> V:=e,s,b\<rangle> \<rightarrow>* \<langle> V:=e',s',b'\<rangle>"
-(*<*)
-proof(induct rule: rtrancl_induct3)
- case refl show ?case by blast
-next
- case step show ?case
- by(rule rtrancl_into_rtrancl[OF step(3) LAssRed[OF step(2)]])
-qed
-(*>*)
-
-lemma LAssRedsVal:
- "\<lbrakk> P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>Val v,(h',l',sh'),b'\<rangle> \<rbrakk> \<Longrightarrow> P \<turnstile> \<langle> V:=e,s,b\<rangle> \<rightarrow>* \<langle>unit,(h',l'(V\<mapsto>v),sh'),b'\<rangle>"
-(*<*)by(rule LAssReds[THEN rtrancl_into_rtrancl, OF _ RedLAss])(*>*)
-
-lemma LAssRedsThrow:
- "\<lbrakk> P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>throw a,s',b'\<rangle> \<rbrakk> \<Longrightarrow> P \<turnstile> \<langle> V:=e,s,b\<rangle> \<rightarrow>* \<langle>throw a,s',b'\<rangle>"
-(*<*)by(rule LAssReds[THEN rtrancl_into_rtrancl, OF _ red_reds.LAssThrow])(*>*)
-
-subsubsection "BinOp"
-
-lemma BinOp1Reds:
- "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle> \<Longrightarrow> P \<turnstile> \<langle> e \<guillemotleft>bop\<guillemotright> e\<^sub>2, s,b\<rangle> \<rightarrow>* \<langle>e' \<guillemotleft>bop\<guillemotright> e\<^sub>2, s',b'\<rangle>"
-(*<*)
-proof(induct rule: rtrancl_induct3)
- case refl show ?case by blast
-next
- case step show ?case
- by(rule rtrancl_into_rtrancl[OF step(3) BinOpRed1[OF step(2)]])
-qed
-(*>*)
-
-lemma BinOp2Reds:
- "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>(Val v) \<guillemotleft>bop\<guillemotright> e, s,b\<rangle> \<rightarrow>* \<langle>(Val v) \<guillemotleft>bop\<guillemotright> e', s',b'\<rangle>"
-(*<*)
-proof(induct rule: rtrancl_induct3)
- case refl show ?case by blast
-next
- case step show ?case
- by(rule rtrancl_into_rtrancl[OF step(3) BinOpRed2[OF step(2)]])
-qed
-(*>*)
-
-lemma BinOpRedsVal:
-assumes e\<^sub>1_steps: "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>Val v\<^sub>1,s\<^sub>1,b\<^sub>1\<rangle>"
- and e\<^sub>2_steps: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>1,b\<^sub>1\<rangle> \<rightarrow>* \<langle>Val v\<^sub>2,s\<^sub>2,b\<^sub>2\<rangle>"
- and op: "binop(bop,v\<^sub>1,v\<^sub>2) = Some v"
-shows "P \<turnstile> \<langle>e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2, s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>Val v,s\<^sub>2,b\<^sub>2\<rangle>"
-(*<*)(is "(?x, ?z) \<in> (red P)\<^sup>*")
-proof -
- let ?y = "(Val v\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2, s\<^sub>1,b\<^sub>1)"
- let ?y' = "(Val v\<^sub>1 \<guillemotleft>bop\<guillemotright> Val v\<^sub>2, s\<^sub>2,b\<^sub>2)"
- have "(?x, ?y) \<in> (red P)\<^sup>*" by(rule BinOp1Reds[OF e\<^sub>1_steps])
- also have "(?y, ?y') \<in> (red P)\<^sup>*" by(rule BinOp2Reds[OF e\<^sub>2_steps])
- also have "(?y', ?z) \<in> (red P)" by(rule RedBinOp[OF op])
- ultimately show ?thesis by simp
-qed
-(*>*)
-
-lemma BinOpRedsThrow1:
- "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>throw e',s',b'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>e \<guillemotleft>bop\<guillemotright> e\<^sub>2, s,b\<rangle> \<rightarrow>* \<langle>throw e', s',b'\<rangle>"
-(*<*)by(rule BinOp1Reds[THEN rtrancl_into_rtrancl, OF _ red_reds.BinOpThrow1])(*>*)
-
-lemma BinOpRedsThrow2:
-assumes e\<^sub>1_steps: "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>Val v\<^sub>1,s\<^sub>1,b\<^sub>1\<rangle>"
- and e\<^sub>2_steps: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>1,b\<^sub>1\<rangle> \<rightarrow>* \<langle>throw e,s\<^sub>2,b\<^sub>2\<rangle>"
-shows "P \<turnstile> \<langle>e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2, s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>throw e,s\<^sub>2,b\<^sub>2\<rangle>"
-(*<*)(is "(?x, ?z) \<in> (red P)\<^sup>*")
-proof -
- let ?y = "(Val v\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2, s\<^sub>1,b\<^sub>1)"
- let ?y' = "(Val v\<^sub>1 \<guillemotleft>bop\<guillemotright> throw e, s\<^sub>2,b\<^sub>2)"
- have "(?x, ?y) \<in> (red P)\<^sup>*" by(rule BinOp1Reds[OF e\<^sub>1_steps])
- also have "(?y, ?y') \<in> (red P)\<^sup>*" by(rule BinOp2Reds[OF e\<^sub>2_steps])
- also have "(?y', ?z) \<in> (red P)" by(rule red_reds.BinOpThrow2)
- ultimately show ?thesis by simp
-qed
-(*>*)
-
-subsubsection "FAcc"
-
-lemma FAccReds:
- "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>e\<bullet>F{D}, s,b\<rangle> \<rightarrow>* \<langle>e'\<bullet>F{D}, s',b'\<rangle>"
-(*<*)
-proof(induct rule: rtrancl_induct3)
- case refl show ?case by blast
-next
- case step show ?case
- by(rule rtrancl_into_rtrancl[OF step(3) FAccRed[OF step(2)]])
-qed
-(*>*)
-
-lemma FAccRedsVal:
- "\<lbrakk> P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>addr a,s',b'\<rangle>; hp s' a = Some(C,fs); fs(F,D) = Some v;
- P \<turnstile> C has F,NonStatic:t in D \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>e\<bullet>F{D},s,b\<rangle> \<rightarrow>* \<langle>Val v,s',b'\<rangle>"
-(*<*)by(cases s',simp) (rule FAccReds[THEN rtrancl_into_rtrancl, OF _ RedFAcc])(*>*)
-
-lemma FAccRedsNull:
- "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>null,s',b'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>e\<bullet>F{D},s,b\<rangle> \<rightarrow>* \<langle>THROW NullPointer,s',b'\<rangle>"
-(*<*)by(rule FAccReds[THEN rtrancl_into_rtrancl, OF _ RedFAccNull])(*>*)
-
-lemma FAccRedsNone:
- "\<lbrakk> P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>addr a,s',b'\<rangle>;
- hp s' a = Some(C,fs);
- \<not>(\<exists>b t. P \<turnstile> C has F,b:t in D) \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>e\<bullet>F{D},s,b\<rangle> \<rightarrow>* \<langle>THROW NoSuchFieldError,s',b'\<rangle>"
-(*<*)by(cases s',simp) (auto intro: FAccReds[THEN rtrancl_into_rtrancl, OF _ RedFAccNone])(*>*)
-
-lemma FAccRedsStatic:
- "\<lbrakk> P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>addr a,s',b'\<rangle>;
- hp s' a = Some(C,fs);
- P \<turnstile> C has F,Static:t in D \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>e\<bullet>F{D},s,b\<rangle> \<rightarrow>* \<langle>THROW IncompatibleClassChangeError,s',b'\<rangle>"
-(*<*)by(cases s',simp) (rule FAccReds[THEN rtrancl_into_rtrancl, OF _ RedFAccStatic])(*>*)
-
-lemma FAccRedsThrow:
- "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>throw a,s',b'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>e\<bullet>F{D},s,b\<rangle> \<rightarrow>* \<langle>throw a,s',b'\<rangle>"
-(*<*)by(rule FAccReds[THEN rtrancl_into_rtrancl, OF _ red_reds.FAccThrow])(*>*)
-
-subsubsection "SFAcc"
-
-lemma SFAccReds:
- "\<lbrakk> P \<turnstile> C has F,Static:t in D;
- shp s D = Some(sfs,Done); sfs F = Some v \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D},s,True\<rangle> \<rightarrow>* \<langle>Val v,s,False\<rangle>"
-(*<*)by(cases s,simp) (rule RedSFAcc[THEN r_into_rtrancl])(*>*)
-
-lemma SFAccRedsNone:
- "\<not>(\<exists>b t. P \<turnstile> C has F,b:t in D)
- \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D},s,b\<rangle> \<rightarrow>* \<langle>THROW NoSuchFieldError,s,False\<rangle>"
-(*<*)by(cases s,simp) (auto intro: RedSFAccNone[THEN r_into_rtrancl])(*>*)
-
-lemma SFAccRedsNonStatic:
- "P \<turnstile> C has F,NonStatic:t in D
- \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D},s,b\<rangle> \<rightarrow>* \<langle>THROW IncompatibleClassChangeError,s,False\<rangle>"
-(*<*)by(cases s,simp) (rule RedSFAccNonStatic[THEN r_into_rtrancl])(*>*)
-
-lemma SFAccInitDoneReds:
-assumes cF: "P \<turnstile> C has F,Static:t in D"
- and shp: "shp s D = Some (sfs,Done)" and sfs: "sfs F = Some v"
-shows "P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D}, s, b\<rangle> \<rightarrow>* \<langle>Val v, s, False\<rangle>"
-(*<*)(is "(?a, ?c) \<in> (red P)\<^sup>*")
-proof -
- obtain h l sh where [simp]: "s = (h, l, sh)" by(cases s) simp
- show ?thesis proof(cases b)
- case True
- then show ?thesis using assms
- by simp (rule RedSFAcc[THEN r_into_rtrancl])
- next
- case False
- let ?b = "(C\<bullet>\<^sub>sF{D},s,True)"
- have "(?a, ?b) \<in> (red P)\<^sup>*"
- using False SFAccInitDoneRed[where sh="shp s", OF cF shp] by fastforce
- also have "(?b, ?c) \<in> (red P)\<^sup>*" by(rule SFAccReds[OF assms])
- ultimately show ?thesis by simp
- qed
-qed
-(*>*)
-
-lemma SFAccInitReds:
-assumes cF: "P \<turnstile> C has F,Static:t in D"
- and nDone: "\<nexists>sfs. shp s D = Some (sfs,Done)"
- and INIT_steps: "P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,s,False\<rangle> \<rightarrow>* \<langle>Val v',s',b'\<rangle>"
- and shp': "shp s' D = Some (sfs,i)" and sfs: "sfs F = Some v"
-shows "P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D},s,False\<rangle> \<rightarrow>* \<langle>Val v,s',False\<rangle>"
-(*<*)(is "(?a, ?c) \<in> (red P)\<^sup>*")
-proof -
- let ?b = "(INIT D ([D],False) \<leftarrow> C\<bullet>\<^sub>sF{D},s,False)"
- let ?b' = "(C\<bullet>\<^sub>sF{D},s',icheck P D (C\<bullet>\<^sub>sF{D}::expr))"
- obtain h l sh where [simp]: "s = (h, l, sh)" by(cases s) simp
- obtain h' l' sh' where [simp]: "s' = (h', l', sh')" by(cases s') simp
- have icheck: "icheck P D (C\<bullet>\<^sub>sF{D}) = True" using cF by fastforce
- then have b'c: "(?b', ?c) \<in> (red P)\<^sup>*"
- using RedSFAcc[THEN r_into_rtrancl, OF cF] shp' sfs by simp
- have "(?a, ?b) \<in> (red P)" using SFAccInitRed[OF cF] nDone by simp
- also have "(?b, ?c) \<in> (red P)\<^sup>*"
- by(rule InitSeqReds[OF INIT_steps b'c]) simp
- ultimately show ?thesis by simp
-qed
-(*>*)
-
-lemma SFAccInitThrowReds:
- "\<lbrakk> P \<turnstile> C has F,Static:t in D;
- \<nexists>sfs. shp s D = Some (sfs,Done);
- P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,s,False\<rangle> \<rightarrow>* \<langle>throw a,s',b'\<rangle> \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D},s,False\<rangle> \<rightarrow>* \<langle>throw a,s',b'\<rangle>"
-(*<*)
-by(cases s, simp)
- (auto elim: converse_rtrancl_into_rtrancl[OF SFAccInitRed InitSeqThrowReds])
-(*>*)
-
-subsubsection "FAss"
-
-lemma FAssReds1:
- "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>e\<bullet>F{D}:=e\<^sub>2, s,b\<rangle> \<rightarrow>* \<langle>e'\<bullet>F{D}:=e\<^sub>2, s',b'\<rangle>"
-(*<*)
-proof(induct rule: rtrancl_induct3)
- case refl show ?case by blast
-next
- case step show ?case
- by(rule rtrancl_into_rtrancl[OF step(3) FAssRed1[OF step(2)]])
-qed
-(*>*)
-
-lemma FAssReds2:
- "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>Val v\<bullet>F{D}:=e, s,b\<rangle> \<rightarrow>* \<langle>Val v\<bullet>F{D}:=e', s',b'\<rangle>"
-(*<*)
-proof(induct rule: rtrancl_induct3)
- case refl show ?case by blast
-next
- case step show ?case
- by(rule rtrancl_into_rtrancl[OF step(3) FAssRed2[OF step(2)]])
-qed
-(*>*)
-
-lemma FAssRedsVal:
-assumes e\<^sub>1_steps: "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>addr a,s\<^sub>1,b\<^sub>1\<rangle>"
- and e\<^sub>2_steps: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>1,b\<^sub>1\<rangle> \<rightarrow>* \<langle>Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\<rangle>"
- and cF: "P \<turnstile> C has F,NonStatic:t in D" and hC: "Some(C,fs) = h\<^sub>2 a"
-shows "P \<turnstile> \<langle>e\<^sub>1\<bullet>F{D}:=e\<^sub>2, s\<^sub>0, b\<^sub>0\<rangle> \<rightarrow>* \<langle>unit, (h\<^sub>2(a\<mapsto>(C,fs((F,D)\<mapsto>v))),l\<^sub>2,sh\<^sub>2),b\<^sub>2\<rangle>"
-(*<*)(is "(?x, ?z) \<in> (red P)\<^sup>*")
-proof -
- let ?y = "(addr a\<bullet>F{D}:=e\<^sub>2, s\<^sub>1, b\<^sub>1)"
- let ?y' = "(addr a\<bullet>F{D}:=Val v::expr,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2)"
- have "(?x, ?y) \<in> (red P)\<^sup>*" by(rule FAssReds1[OF e\<^sub>1_steps])
- also have "(?y, ?y') \<in> (red P)\<^sup>*" by(rule FAssReds2[OF e\<^sub>2_steps])
- also have "(?y', ?z) \<in> (red P)" using RedFAss[OF cF] hC by simp
- ultimately show ?thesis by simp
-qed
-(*>*)
-
-lemma FAssRedsNull:
-assumes e\<^sub>1_steps: "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>null,s\<^sub>1,b\<^sub>1\<rangle>"
- and e\<^sub>2_steps: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>1,b\<^sub>1\<rangle> \<rightarrow>* \<langle>Val v,s\<^sub>2,b\<^sub>2\<rangle>"
-shows "P \<turnstile> \<langle>e\<^sub>1\<bullet>F{D}:=e\<^sub>2, s\<^sub>0, b\<^sub>0\<rangle> \<rightarrow>* \<langle>THROW NullPointer, s\<^sub>2, b\<^sub>2\<rangle>"
-(*<*)(is "(?x, ?z) \<in> (red P)\<^sup>*")
-proof -
- let ?y = "(null\<bullet>F{D}:=e\<^sub>2, s\<^sub>1, b\<^sub>1)"
- let ?y' = "(null\<bullet>F{D}:=Val v::expr,s\<^sub>2,b\<^sub>2)"
- have "(?x, ?y) \<in> (red P)\<^sup>*" by(rule FAssReds1[OF e\<^sub>1_steps])
- also have "(?y, ?y') \<in> (red P)\<^sup>*" by(rule FAssReds2[OF e\<^sub>2_steps])
- also have "(?y', ?z) \<in> (red P)" by(rule RedFAssNull)
- ultimately show ?thesis by simp
-qed
-(*>*)
-
-lemma FAssRedsThrow1:
- "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>throw e',s',b'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>e\<bullet>F{D}:=e\<^sub>2, s,b\<rangle> \<rightarrow>* \<langle>throw e', s',b'\<rangle>"
-(*<*)by(rule FAssReds1[THEN rtrancl_into_rtrancl, OF _ red_reds.FAssThrow1])(*>*)
-
-lemma FAssRedsThrow2:
-assumes e\<^sub>1_steps: "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>Val v,s\<^sub>1,b\<^sub>1\<rangle>"
- and e\<^sub>2_steps: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>1,b\<^sub>1\<rangle> \<rightarrow>* \<langle>throw e,s\<^sub>2,b\<^sub>2\<rangle>"
-shows "P \<turnstile> \<langle>e\<^sub>1\<bullet>F{D}:=e\<^sub>2,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>throw e,s\<^sub>2,b\<^sub>2\<rangle>"
-(*<*)(is "(?x, ?z) \<in> (red P)\<^sup>*")
-proof -
- let ?y = "(Val v\<bullet>F{D}:=e\<^sub>2,s\<^sub>1,b\<^sub>1)"
- let ?y' = "(Val v\<bullet>F{D}:=throw e,s\<^sub>2,b\<^sub>2)"
- have "(?x, ?y) \<in> (red P)\<^sup>*" by(rule FAssReds1[OF e\<^sub>1_steps])
- also have "(?y, ?y') \<in> (red P)\<^sup>*" by(rule FAssReds2[OF e\<^sub>2_steps])
- also have "(?y', ?z) \<in> (red P)" by(rule red_reds.FAssThrow2)
- ultimately show ?thesis by simp
-qed
-(*>*)
-
-lemma FAssRedsNone:
-assumes e\<^sub>1_steps: "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>addr a,s\<^sub>1,b\<^sub>1\<rangle>"
- and e\<^sub>2_steps: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>1,b\<^sub>1\<rangle> \<rightarrow>* \<langle>Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\<rangle>"
- and hC: "h\<^sub>2 a = Some(C,fs)" and ncF: "\<not>(\<exists>b t. P \<turnstile> C has F,b:t in D)"
-shows "P \<turnstile> \<langle>e\<^sub>1\<bullet>F{D}:=e\<^sub>2, s\<^sub>0, b\<^sub>0\<rangle> \<rightarrow>* \<langle>THROW NoSuchFieldError, (h\<^sub>2,l\<^sub>2,sh\<^sub>2), b\<^sub>2\<rangle>"
-(*<*)(is "(?x, ?z) \<in> (red P)\<^sup>*")
-proof -
- let ?y = "(addr a\<bullet>F{D}:=e\<^sub>2, s\<^sub>1, b\<^sub>1)"
- let ?y' = "(addr a\<bullet>F{D}:=Val v::expr,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2)"
- have "(?x, ?y) \<in> (red P)\<^sup>*" by(rule FAssReds1[OF e\<^sub>1_steps])
- also have "(?y, ?y') \<in> (red P)\<^sup>*" by(rule FAssReds2[OF e\<^sub>2_steps])
- also have "(?y', ?z) \<in> (red P)"
- using RedFAssNone[OF _ ncF] hC by simp
- ultimately show ?thesis by simp
-qed
-(*>*)
-
-lemma FAssRedsStatic:
-assumes e\<^sub>1_steps: "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>addr a,s\<^sub>1,b\<^sub>1\<rangle>"
- and e\<^sub>2_steps: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>1,b\<^sub>1\<rangle> \<rightarrow>* \<langle>Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\<rangle>"
- and hC: "h\<^sub>2 a = Some(C,fs)" and cF_Static: "P \<turnstile> C has F,Static:t in D"
-shows "P \<turnstile> \<langle>e\<^sub>1\<bullet>F{D}:=e\<^sub>2, s\<^sub>0, b\<^sub>0\<rangle> \<rightarrow>* \<langle>THROW IncompatibleClassChangeError, (h\<^sub>2,l\<^sub>2,sh\<^sub>2), b\<^sub>2\<rangle>"
-(*<*)(is "(?x, ?z) \<in> (red P)\<^sup>*")
-proof -
- let ?y = "(addr a\<bullet>F{D}:=e\<^sub>2, s\<^sub>1, b\<^sub>1)"
- let ?y' = "(addr a\<bullet>F{D}:=Val v::expr,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2)"
- have "(?x, ?y) \<in> (red P)\<^sup>*" by(rule FAssReds1[OF e\<^sub>1_steps])
- also have "(?y, ?y') \<in> (red P)\<^sup>*" by(rule FAssReds2[OF e\<^sub>2_steps])
- also have "(?y', ?z) \<in> (red P)"
- using RedFAssStatic[OF _ cF_Static] hC by simp
- ultimately show ?thesis by simp
-qed
-(*>*)
-
-subsubsection "SFAss"
-
-lemma SFAssReds:
- "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D}:=e,s,b\<rangle> \<rightarrow>* \<langle>C\<bullet>\<^sub>sF{D}:=e',s',b'\<rangle>"
-(*<*)
-proof(induct rule: rtrancl_induct3)
- case refl show ?case by blast
-next
- case step show ?case
- by(rule rtrancl_into_rtrancl[OF step(3) SFAssRed[OF step(2)]])
-qed
-(*>*)
-
-lemma SFAssRedsVal:
-assumes e\<^sub>2_steps: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\<rangle>"
- and cF: "P \<turnstile> C has F,Static:t in D"
- and shD: "sh\<^sub>2 D = \<lfloor>(sfs,Done)\<rfloor>"
-shows "P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D}:=e\<^sub>2, s\<^sub>0, b\<^sub>0\<rangle> \<rightarrow>* \<langle>unit, (h\<^sub>2,l\<^sub>2,sh\<^sub>2(D\<mapsto>(sfs(F\<mapsto>v), Done))),False\<rangle>"
-(*<*)(is "(?a, ?c) \<in> (red P)\<^sup>*")
-proof -
- let ?b = "(C\<bullet>\<^sub>sF{D}:=Val v::expr, (h\<^sub>2,l\<^sub>2,sh\<^sub>2), b\<^sub>2)"
- have "(?a, ?b) \<in> (red P)\<^sup>*" by(rule SFAssReds[OF e\<^sub>2_steps])
- also have "(?b, ?c) \<in> (red P)\<^sup>*" proof(cases b\<^sub>2)
- case True
- then show ?thesis
- using RedSFAss[THEN r_into_rtrancl, OF cF] shD by simp
- next
- case False
- let ?b' = "(C\<bullet>\<^sub>sF{D}:=Val v::expr, (h\<^sub>2,l\<^sub>2,sh\<^sub>2), True)"
- have "(?b, ?b') \<in> (red P)\<^sup>*"
- using False SFAssInitDoneRed[THEN converse_rtrancl_into_rtrancl, OF cF] shD
- by simp
- also have "(?b', ?c) \<in> (red P)\<^sup>*"
- using RedSFAss[THEN r_into_rtrancl, OF cF] shD by simp
- ultimately show ?thesis by simp
- qed
- ultimately show ?thesis by simp
-qed
-(*>*)
-
-lemma SFAssRedsThrow:
- "\<lbrakk> P \<turnstile> \<langle>e\<^sub>2,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>throw e,s\<^sub>2,b\<^sub>2\<rangle> \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D}:=e\<^sub>2,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>throw e,s\<^sub>2,b\<^sub>2\<rangle>"
-(*<*)by(rule SFAssReds[THEN rtrancl_into_rtrancl, OF _ red_reds.SFAssThrow])(*>*)
-
-lemma SFAssRedsNone:
- "\<lbrakk> P \<turnstile> \<langle>e\<^sub>2,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\<rangle>;
- \<not>(\<exists>b t. P \<turnstile> C has F,b:t in D) \<rbrakk> \<Longrightarrow>
- P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D}:=e\<^sub>2,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>THROW NoSuchFieldError, (h\<^sub>2,l\<^sub>2,sh\<^sub>2),False\<rangle>"
-(*<*)by(rule SFAssReds[THEN rtrancl_into_rtrancl, OF _ RedSFAssNone])(*>*)
-
-lemma SFAssRedsNonStatic:
- "\<lbrakk> P \<turnstile> \<langle>e\<^sub>2,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\<rangle>;
- P \<turnstile> C has F,NonStatic:t in D \<rbrakk> \<Longrightarrow>
- P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D}:=e\<^sub>2,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>THROW IncompatibleClassChangeError,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),False\<rangle>"
-(*<*)by(rule SFAssReds[THEN rtrancl_into_rtrancl, OF _ RedSFAssNonStatic])(*>*)
-
-lemma SFAssInitReds:
-assumes e\<^sub>2_steps: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),False\<rangle>"
- and cF: "P \<turnstile> C has F,Static:t in D"
- and nDone: "\<nexists>sfs. sh\<^sub>2 D = Some (sfs, Done)"
- and INIT_steps: "P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),False\<rangle> \<rightarrow>* \<langle>Val v',(h',l',sh'),b'\<rangle>"
- and sh'D: "sh' D = Some(sfs,i)"
- and sfs': "sfs' = sfs(F\<mapsto>v)" and sh'': "sh'' = sh'(D\<mapsto>(sfs',i))"
-shows "P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D}:=e\<^sub>2,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>unit,(h',l',sh''),False\<rangle>"
-(*<*)(is "(?x, ?z) \<in> (red P)\<^sup>*")
-proof -
- let ?y = "(C\<bullet>\<^sub>sF{D} := Val v::expr,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False)"
- let ?y' = "(INIT D ([D],False) \<leftarrow> C\<bullet>\<^sub>sF{D} := Val v::expr,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False)"
- let ?y'' = "(C\<bullet>\<^sub>sF{D} := Val v::expr,(h', l', sh'),icheck P D (C\<bullet>\<^sub>sF{D} := Val v::expr))"
- have icheck: "icheck P D (C\<bullet>\<^sub>sF{D} := Val v::expr)" using cF by fastforce
- then have y''z: "(?y'',?z) \<in> (red P)"
- using RedSFAss[OF cF _ sfs' sh''] sh'D by simp
- have "(?x, ?y) \<in> (red P)\<^sup>*" by(rule SFAssReds[OF e\<^sub>2_steps])
- also have "(?y, ?y') \<in> (red P)\<^sup>*"
- using SFAssInitRed[THEN converse_rtrancl_into_rtrancl, OF cF] nDone
- by simp
- also have "(?y', ?z) \<in> (red P)\<^sup>*"
- using InitSeqReds[OF INIT_steps y''z[THEN r_into_rtrancl]] by simp
- ultimately show ?thesis by simp
-qed
-(*>*)
-
-lemma SFAssInitThrowReds:
-assumes e\<^sub>2_steps: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),False\<rangle>"
- and cF: "P \<turnstile> C has F,Static:t in D"
- and nDone: "\<nexists>sfs. sh\<^sub>2 D = Some (sfs, Done)"
- and INIT_steps: "P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),False\<rangle> \<rightarrow>* \<langle>throw a,s',b'\<rangle>"
-shows "P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D}:=e\<^sub>2,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>throw a,s',b'\<rangle>"
-(*<*)(is "(?x, ?z) \<in> (red P)\<^sup>*")
-proof -
- let ?y = "(C\<bullet>\<^sub>sF{D} := Val v::expr,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False)"
- let ?y' = "(INIT D ([D],False) \<leftarrow> C\<bullet>\<^sub>sF{D} := Val v::expr,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False)"
- have "(?x, ?y) \<in> (red P)\<^sup>*" by(rule SFAssReds[OF e\<^sub>2_steps])
- also have "(?y, ?y') \<in> (red P)\<^sup>*"
- using SFAssInitRed[THEN converse_rtrancl_into_rtrancl, OF cF] nDone
- by simp
- also have "(?y', ?z) \<in> (red P)\<^sup>*"
- using InitSeqThrowReds[OF INIT_steps] by simp
- ultimately show ?thesis by simp
-qed
-(*>*)
-
-subsubsection";;"
-
-lemma SeqReds:
- "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>e;;e\<^sub>2, s,b\<rangle> \<rightarrow>* \<langle>e';;e\<^sub>2, s',b'\<rangle>"
-(*<*)
-proof(induct rule: rtrancl_induct3)
- case refl show ?case by blast
-next
- case step show ?case
- by(rule rtrancl_into_rtrancl[OF step(3) SeqRed[OF step(2)]])
-qed
-(*>*)
-
-lemma SeqRedsThrow:
- "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>throw e',s',b'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>e;;e\<^sub>2, s,b\<rangle> \<rightarrow>* \<langle>throw e', s',b'\<rangle>"
-(*<*)by(rule SeqReds[THEN rtrancl_into_rtrancl, OF _ red_reds.SeqThrow])(*>*)
-
-lemma SeqReds2:
-assumes e\<^sub>1_steps: "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>Val v\<^sub>1,s\<^sub>1,b\<^sub>1\<rangle>"
- and e\<^sub>2_steps: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>1,b\<^sub>1\<rangle> \<rightarrow>* \<langle>e\<^sub>2',s\<^sub>2,b\<^sub>2\<rangle>"
-shows "P \<turnstile> \<langle>e\<^sub>1;;e\<^sub>2, s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>e\<^sub>2',s\<^sub>2,b\<^sub>2\<rangle>"
-(*<*)(is "(?x, ?z) \<in> (red P)\<^sup>*")
-proof -
- let ?y = "(Val v\<^sub>1;; e\<^sub>2,s\<^sub>1,b\<^sub>1)"
- have "(?x, ?y) \<in> (red P)\<^sup>*" by(rule SeqReds[OF e\<^sub>1_steps])
- also have "(?y, ?z) \<in> (red P)\<^sup>*"
- by(rule RedSeq[THEN converse_rtrancl_into_rtrancl, OF e\<^sub>2_steps])
- ultimately show ?thesis by simp
-qed
-(*>*)
-
-
-subsubsection"If"
-
-lemma CondReds:
- "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>if (e) e\<^sub>1 else e\<^sub>2,s,b\<rangle> \<rightarrow>* \<langle>if (e') e\<^sub>1 else e\<^sub>2,s',b'\<rangle>"
-(*<*)
-proof(induct rule: rtrancl_induct3)
- case refl show ?case by blast
-next
- case step show ?case
- by(rule rtrancl_into_rtrancl[OF step(3) CondRed[OF step(2)]])
-qed
-(*>*)
-
-lemma CondRedsThrow:
- "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>throw a,s',b'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>if (e) e\<^sub>1 else e\<^sub>2, s,b\<rangle> \<rightarrow>* \<langle>throw a,s',b'\<rangle>"
-(*<*)by(rule CondReds[THEN rtrancl_into_rtrancl, OF _ red_reds.CondThrow])(*>*)
-
-lemma CondReds2T:
-assumes e_steps: "P \<turnstile> \<langle>e,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>true,s\<^sub>1,b\<^sub>1\<rangle>"
- and e\<^sub>1_steps: "P \<turnstile> \<langle>e\<^sub>1, s\<^sub>1,b\<^sub>1\<rangle> \<rightarrow>* \<langle>e',s\<^sub>2,b\<^sub>2\<rangle>"
-shows "P \<turnstile> \<langle>if (e) e\<^sub>1 else e\<^sub>2, s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>e',s\<^sub>2,b\<^sub>2\<rangle>"
-(*<*)(is "(?x, ?z) \<in> (red P)\<^sup>*")
-proof -
- let ?y = "(if (true) e\<^sub>1 else e\<^sub>2,s\<^sub>1,b\<^sub>1)"
- have "(?x, ?y) \<in> (red P)\<^sup>*" by(rule CondReds[OF e_steps])
- also have "(?y, ?z) \<in> (red P)\<^sup>*"
- by(rule RedCondT[THEN converse_rtrancl_into_rtrancl, OF e\<^sub>1_steps])
- ultimately show ?thesis by simp
-qed
-(*>*)
-
-lemma CondReds2F:
-assumes e_steps: "P \<turnstile> \<langle>e,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>false,s\<^sub>1,b\<^sub>1\<rangle>"
- and e\<^sub>2_steps: "P \<turnstile> \<langle>e\<^sub>2, s\<^sub>1,b\<^sub>1\<rangle> \<rightarrow>* \<langle>e',s\<^sub>2,b\<^sub>2\<rangle>"
-shows "P \<turnstile> \<langle>if (e) e\<^sub>1 else e\<^sub>2, s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>e',s\<^sub>2,b\<^sub>2\<rangle>"
-(*<*)(is "(?x, ?z) \<in> (red P)\<^sup>*")
-proof -
- let ?y = "(if (false) e\<^sub>1 else e\<^sub>2,s\<^sub>1,b\<^sub>1)"
- have "(?x, ?y) \<in> (red P)\<^sup>*" by(rule CondReds[OF e_steps])
- also have "(?y, ?z) \<in> (red P)\<^sup>*"
- by(rule RedCondF[THEN converse_rtrancl_into_rtrancl, OF e\<^sub>2_steps])
- ultimately show ?thesis by simp
-qed
-(*>*)
-
-
-subsubsection "While"
-
-lemma WhileFReds:
-assumes b_steps: "P \<turnstile> \<langle>b,s,b\<^sub>0\<rangle> \<rightarrow>* \<langle>false,s',b'\<rangle>"
-shows "P \<turnstile> \<langle>while (b) c,s,b\<^sub>0\<rangle> \<rightarrow>* \<langle>unit,s',b'\<rangle>"
-(*<*)
-by(rule RedWhile[THEN converse_rtrancl_into_rtrancl,
- OF CondReds[THEN rtrancl_into_rtrancl,
- OF b_steps RedCondF]])
-(*>*)
-
-lemma WhileRedsThrow:
-assumes b_steps: "P \<turnstile> \<langle>b,s,b\<^sub>0\<rangle> \<rightarrow>* \<langle>throw e,s',b'\<rangle>"
-shows "P \<turnstile> \<langle>while (b) c,s,b\<^sub>0\<rangle> \<rightarrow>* \<langle>throw e,s',b'\<rangle>"
-(*<*)
-by(rule RedWhile[THEN converse_rtrancl_into_rtrancl,
- OF CondReds[THEN rtrancl_into_rtrancl,
- OF b_steps red_reds.CondThrow]])
-(*>*)
-
-lemma WhileTReds:
-assumes b_steps: "P \<turnstile> \<langle>b,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>true,s\<^sub>1,b\<^sub>1\<rangle>"
- and c_steps: "P \<turnstile> \<langle>c,s\<^sub>1,b\<^sub>1\<rangle> \<rightarrow>* \<langle>Val v\<^sub>1,s\<^sub>2,b\<^sub>2\<rangle>"
- and while_steps: "P \<turnstile> \<langle>while (b) c,s\<^sub>2,b\<^sub>2\<rangle> \<rightarrow>* \<langle>e,s\<^sub>3,b\<^sub>3\<rangle>"
-shows "P \<turnstile> \<langle>while (b) c,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>e,s\<^sub>3,b\<^sub>3\<rangle>"
-(*<*)(is "(?a, ?c) \<in> (red P)\<^sup>*")
-proof -
- let ?b = "(if (b) (c;; while (b) c) else unit,s\<^sub>0,b\<^sub>0)"
- let ?y = "(if (true) (c;; while (b) c) else unit,s\<^sub>1,b\<^sub>1)"
- let ?b' = "(c;; while (b) c,s\<^sub>1,b\<^sub>1)"
- let ?y' = "(Val v\<^sub>1;; while (b) c,s\<^sub>2,b\<^sub>2)"
- have "(?a, ?b) \<in> (red P)\<^sup>*"
- using RedWhile[THEN converse_rtrancl_into_rtrancl] by simp
- also have "(?b, ?y) \<in> (red P)\<^sup>*" by(rule CondReds[OF b_steps])
- also have "(?y, ?b') \<in> (red P)\<^sup>*"
- using RedCondT[THEN converse_rtrancl_into_rtrancl] by simp
- also have "(?b', ?y') \<in> (red P)\<^sup>*" by(rule SeqReds[OF c_steps])
- also have "(?y', ?c) \<in> (red P)\<^sup>*"
- by(rule RedSeq[THEN converse_rtrancl_into_rtrancl, OF while_steps])
- ultimately show ?thesis by simp
-qed
-(*>*)
-
-lemma WhileTRedsThrow:
-assumes b_steps: "P \<turnstile> \<langle>b,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>true,s\<^sub>1,b\<^sub>1\<rangle>"
- and c_steps: "P \<turnstile> \<langle>c,s\<^sub>1,b\<^sub>1\<rangle> \<rightarrow>* \<langle>throw e,s\<^sub>2,b\<^sub>2\<rangle>"
-shows "P \<turnstile> \<langle>while (b) c,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>throw e,s\<^sub>2,b\<^sub>2\<rangle>"
-(*<*)(is "(?a, ?c) \<in> (red P)\<^sup>*")
-proof -
- let ?b = "(if (b) (c;; while (b) c) else unit,s\<^sub>0,b\<^sub>0)"
- let ?y = "(if (true) (c;; while (b) c) else unit,s\<^sub>1,b\<^sub>1)"
- let ?b' = "(c;; while (b) c,s\<^sub>1,b\<^sub>1)"
- let ?y' = "(throw e;; while (b) c,s\<^sub>2,b\<^sub>2)"
- have "(?a, ?b) \<in> (red P)\<^sup>*"
- using RedWhile[THEN converse_rtrancl_into_rtrancl] by simp
- also have "(?b, ?y) \<in> (red P)\<^sup>*" by(rule CondReds[OF b_steps])
- also have "(?y, ?b') \<in> (red P)\<^sup>*"
- using RedCondT[THEN converse_rtrancl_into_rtrancl] by simp
- also have "(?b', ?y') \<in> (red P)\<^sup>*" by(rule SeqReds[OF c_steps])
- also have "(?y', ?c) \<in> (red P)" by(rule red_reds.SeqThrow)
- ultimately show ?thesis by simp
-qed
-(*>*)
-
-subsubsection"Throw"
-
-lemma ThrowReds:
- "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>throw e,s,b\<rangle> \<rightarrow>* \<langle>throw e',s',b'\<rangle>"
-(*<*)
-proof(induct rule: rtrancl_induct3)
- case refl show ?case by blast
-next
- case step show ?case
- by(rule rtrancl_into_rtrancl[OF step(3) ThrowRed[OF step(2)]])
-qed
-(*>*)
-
-lemma ThrowRedsNull:
- "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>null,s',b'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>throw e,s,b\<rangle> \<rightarrow>* \<langle>THROW NullPointer,s',b'\<rangle>"
-(*<*)by(rule ThrowReds[THEN rtrancl_into_rtrancl, OF _ RedThrowNull])(*>*)
-
-lemma ThrowRedsThrow:
- "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>throw a,s',b'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>throw e,s,b\<rangle> \<rightarrow>* \<langle>throw a,s',b'\<rangle>"
-(*<*)by(rule ThrowReds[THEN rtrancl_into_rtrancl, OF _ red_reds.ThrowThrow])(*>*)
-
-subsubsection "InitBlock"
-
-lemma InitBlockReds_aux:
- "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle> \<Longrightarrow>
- \<forall>h l sh h' l' sh' v. s = (h,l(V\<mapsto>v),sh) \<longrightarrow> s' = (h',l',sh') \<longrightarrow>
- P \<turnstile> \<langle>{V:T := Val v; e},(h,l,sh),b\<rangle> \<rightarrow>* \<langle>{V:T := Val(the(l' V)); e'},(h',l'(V:=(l V)),sh'),b'\<rangle>"
-(*<*)
-proof(induct rule: converse_rtrancl_induct3)
- case refl then show ?case
- by(fastforce simp: fun_upd_same simp del:fun_upd_apply)
-next
- case (step e0 s0 b0 e1 s1 b1)
- obtain h1 l1 sh1 where s1[simp]: "s1 = (h1, l1, sh1)" by(cases s1) simp
- { fix h0 l0 sh0 h2 l2 sh2 v0
- assume [simp]: "s0 = (h0, l0(V \<mapsto> v0), sh0)" and s'[simp]: "s' = (h2, l2, sh2)"
- then have "V \<in> dom l1" using step(1) by(auto dest!: red_lcl_incr)
- then obtain v1 where l1V[simp]: "l1 V = \<lfloor>v1\<rfloor>" by blast
-
- let ?a = "({V:T; V:=Val v0;; e0},(h0, l0, sh0),b0)"
- let ?b = "({V:T; V:=Val v1;; e1},(h1, l1(V := l0 V), sh1),b1)"
- let ?c = "({V:T; V:=Val (the (l2 V));; e'},(h2, l2(V := l0 V), sh2),b')"
- let ?l = "l1(V := l0 V)" and ?v = v1
-
- have e0_steps: "P \<turnstile> \<langle>e0,(h0, l0(V \<mapsto> v0), sh0),b0\<rangle> \<rightarrow> \<langle>e1,(h1, l1, sh1),b1\<rangle>"
- using step(1) by simp
-
- have lv: "\<And>l v. l1 = l(V \<mapsto> v) \<longrightarrow>
- P \<turnstile> \<langle>{V:T; V:=Val v;; e1},(h1, l, sh1),b1\<rangle> \<rightarrow>*
- \<langle>{V:T; V:=Val (the (l2 V));; e'},(h2, l2(V := l V), sh2),b'\<rangle>"
- using step(3) s' s1 by blast
- moreover have "l1 = ?l(V \<mapsto> ?v)" by(rule ext) (simp add:fun_upd_def)
- ultimately have "(?b, ?c) \<in> (red P)\<^sup>*" using lv[of ?l ?v] by simp
- then have "(?a, ?c) \<in> (red P)\<^sup>*"
- by(rule InitBlockRed[THEN converse_rtrancl_into_rtrancl, OF e0_steps l1V])
- }
- then show ?case by blast
-qed
-(*>*)
-
-lemma InitBlockReds:
- "P \<turnstile> \<langle>e, (h,l(V\<mapsto>v),sh),b\<rangle> \<rightarrow>* \<langle>e', (h',l',sh'),b'\<rangle> \<Longrightarrow>
- P \<turnstile> \<langle>{V:T := Val v; e}, (h,l,sh),b\<rangle> \<rightarrow>* \<langle>{V:T := Val(the(l' V)); e'}, (h',l'(V:=(l V)),sh'),b'\<rangle>"
-(*<*)by(blast dest:InitBlockReds_aux)(*>*)
-
-lemma InitBlockRedsFinal:
- "\<lbrakk> P \<turnstile> \<langle>e,(h,l(V\<mapsto>v),sh),b\<rangle> \<rightarrow>* \<langle>e',(h',l',sh'),b'\<rangle>; final e' \<rbrakk> \<Longrightarrow>
- P \<turnstile> \<langle>{V:T := Val v; e},(h,l,sh),b\<rangle> \<rightarrow>* \<langle>e',(h', l'(V := l V),sh'),b'\<rangle>"
-(*<*)
-by(fast elim!:InitBlockReds[THEN rtrancl_into_rtrancl] finalE
- intro:RedInitBlock InitBlockThrow)
-(*>*)
-
-
-subsubsection "Block"
-
-lemmas converse_rtranclE3 = converse_rtranclE [of "(xa,xb,xc)" "(za,zb,zc)", split_rule]
-
-lemma BlockRedsFinal:
-assumes reds: "P \<turnstile> \<langle>e\<^sub>0,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>e\<^sub>2,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\<rangle>" and fin: "final e\<^sub>2"
-shows "\<And>h\<^sub>0 l\<^sub>0 sh\<^sub>0. s\<^sub>0 = (h\<^sub>0,l\<^sub>0(V:=None),sh\<^sub>0) \<Longrightarrow> P \<turnstile> \<langle>{V:T; e\<^sub>0},(h\<^sub>0,l\<^sub>0,sh\<^sub>0),b\<^sub>0\<rangle> \<rightarrow>* \<langle>e\<^sub>2,(h\<^sub>2,l\<^sub>2(V:=l\<^sub>0 V),sh\<^sub>2),b\<^sub>2\<rangle>"
-(*<*)
-using reds
-proof (induct rule:converse_rtrancl_induct3)
- case refl thus ?case
- by(fastforce intro:finalE[OF fin] RedBlock BlockThrow
- simp del:fun_upd_apply)
-next
- case (step e\<^sub>0 s\<^sub>0 b\<^sub>0 e\<^sub>1 s\<^sub>1 b\<^sub>1)
- have red: "P \<turnstile> \<langle>e\<^sub>0,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow> \<langle>e\<^sub>1,s\<^sub>1,b\<^sub>1\<rangle>"
- and reds: "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>1,b\<^sub>1\<rangle> \<rightarrow>* \<langle>e\<^sub>2,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\<rangle>"
- and IH: "\<And>h l sh. s\<^sub>1 = (h,l(V := None),sh)
- \<Longrightarrow> P \<turnstile> \<langle>{V:T; e\<^sub>1},(h,l,sh),b\<^sub>1\<rangle> \<rightarrow>* \<langle>e\<^sub>2,(h\<^sub>2, l\<^sub>2(V := l V),sh\<^sub>2),b\<^sub>2\<rangle>"
- and s\<^sub>0: "s\<^sub>0 = (h\<^sub>0, l\<^sub>0(V := None),sh\<^sub>0)" by fact+
- obtain h\<^sub>1 l\<^sub>1 sh\<^sub>1 where s\<^sub>1: "s\<^sub>1 = (h\<^sub>1,l\<^sub>1,sh\<^sub>1)"
- using prod_cases3 by blast
- show ?case
- proof cases
- assume "assigned V e\<^sub>0"
- then obtain v e where e\<^sub>0: "e\<^sub>0 = V := Val v;; e"
- by (unfold assigned_def)blast
- from red e\<^sub>0 s\<^sub>0 have e\<^sub>1: "e\<^sub>1 = unit;;e" and s\<^sub>1: "s\<^sub>1 = (h\<^sub>0, l\<^sub>0(V \<mapsto> v),sh\<^sub>0)" and b\<^sub>1: "b\<^sub>1 = b\<^sub>0"
- by auto
- from e\<^sub>1 fin have "e\<^sub>1 \<noteq> e\<^sub>2" by (auto simp:final_def)
- then obtain e' s' b' where red1: "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>1,b\<^sub>1\<rangle> \<rightarrow> \<langle>e',s',b'\<rangle>"
- and reds': "P \<turnstile> \<langle>e',s',b'\<rangle> \<rightarrow>* \<langle>e\<^sub>2,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\<rangle>"
- using converse_rtranclE3[OF reds] by blast
- from red1 e\<^sub>1 b\<^sub>1 have es': "e' = e" "s' = s\<^sub>1" "b' = b\<^sub>0" by auto
- show ?case using e\<^sub>0 s\<^sub>1 es' reds'
- by(auto intro!: InitBlockRedsFinal[OF _ fin] simp del:fun_upd_apply)
- next
- assume unass: "\<not> assigned V e\<^sub>0"
- show ?thesis
- proof (cases "l\<^sub>1 V")
- assume None: "l\<^sub>1 V = None"
- hence "P \<turnstile> \<langle>{V:T; e\<^sub>0},(h\<^sub>0,l\<^sub>0,sh\<^sub>0),b\<^sub>0\<rangle> \<rightarrow> \<langle>{V:T; e\<^sub>1},(h\<^sub>1, l\<^sub>1(V := l\<^sub>0 V),sh\<^sub>1),b\<^sub>1\<rangle>"
- using s\<^sub>0 s\<^sub>1 red by(simp add: BlockRedNone[OF _ _ unass])
- moreover
- have "P \<turnstile> \<langle>{V:T; e\<^sub>1},(h\<^sub>1, l\<^sub>1(V := l\<^sub>0 V),sh\<^sub>1),b\<^sub>1\<rangle> \<rightarrow>* \<langle>e\<^sub>2,(h\<^sub>2, l\<^sub>2(V := l\<^sub>0 V),sh\<^sub>2),b\<^sub>2\<rangle>"
- using IH[of _ "l\<^sub>1(V := l\<^sub>0 V)"] s\<^sub>1 None by(simp add:fun_upd_idem)
- ultimately show ?case by(rule converse_rtrancl_into_rtrancl)
- next
- fix v assume Some: "l\<^sub>1 V = Some v"
- hence "P \<turnstile> \<langle>{V:T;e\<^sub>0},(h\<^sub>0,l\<^sub>0,sh\<^sub>0),b\<^sub>0\<rangle> \<rightarrow> \<langle>{V:T := Val v; e\<^sub>1},(h\<^sub>1,l\<^sub>1(V := l\<^sub>0 V),sh\<^sub>1),b\<^sub>1\<rangle>"
- using s\<^sub>0 s\<^sub>1 red by(simp add: BlockRedSome[OF _ _ unass])
- moreover
- have "P \<turnstile> \<langle>{V:T := Val v; e\<^sub>1},(h\<^sub>1,l\<^sub>1(V:= l\<^sub>0 V),sh\<^sub>1),b\<^sub>1\<rangle> \<rightarrow>*
- \<langle>e\<^sub>2,(h\<^sub>2,l\<^sub>2(V:=l\<^sub>0 V),sh\<^sub>2),b\<^sub>2\<rangle>"
- using InitBlockRedsFinal[OF _ fin,of _ _ "l\<^sub>1(V:=l\<^sub>0 V)" V]
- Some reds s\<^sub>1 by(simp add:fun_upd_idem)
- ultimately show ?case by(rule converse_rtrancl_into_rtrancl)
- qed
- qed
-qed
-(*>*)
-
-
-subsubsection "try-catch"
-
-lemma TryReds:
- "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>try e catch(C V) e\<^sub>2,s,b\<rangle> \<rightarrow>* \<langle>try e' catch(C V) e\<^sub>2,s',b'\<rangle>"
-(*<*)
-proof(induct rule: rtrancl_induct3)
- case refl show ?case by blast
-next
- case step show ?case
- by(rule rtrancl_into_rtrancl[OF step(3) TryRed[OF step(2)]])
-qed
-(*>*)
-
-lemma TryRedsVal:
- "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>Val v,s',b'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>try e catch(C V) e\<^sub>2,s,b\<rangle> \<rightarrow>* \<langle>Val v,s',b'\<rangle>"
-(*<*)by(rule TryReds[THEN rtrancl_into_rtrancl, OF _ RedTry])(*>*)
-
-lemma TryCatchRedsFinal:
-assumes e\<^sub>1_steps: "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>Throw a,(h\<^sub>1,l\<^sub>1,sh\<^sub>1),b\<^sub>1\<rangle>"
- and h\<^sub>1a: "h\<^sub>1 a = Some(D,fs)" and sub: "P \<turnstile> D \<preceq>\<^sup>* C"
- and e\<^sub>2_steps: "P \<turnstile> \<langle>e\<^sub>2, (h\<^sub>1, l\<^sub>1(V \<mapsto> Addr a),sh\<^sub>1),b\<^sub>1\<rangle> \<rightarrow>* \<langle>e\<^sub>2', (h\<^sub>2,l\<^sub>2,sh\<^sub>2), b\<^sub>2\<rangle>"
- and final: "final e\<^sub>2'"
-shows "P \<turnstile> \<langle>try e\<^sub>1 catch(C V) e\<^sub>2, s\<^sub>0, b\<^sub>0\<rangle> \<rightarrow>* \<langle>e\<^sub>2', (h\<^sub>2, (l\<^sub>2::locals)(V := l\<^sub>1 V),sh\<^sub>2),b\<^sub>2\<rangle>"
-(*<*)(is "(?x, ?z) \<in> (red P)\<^sup>*")
-proof -
- let ?y = "(try Throw a catch(C V) e\<^sub>2,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),b\<^sub>1)"
- let ?b = "({V:Class C; V:=addr a;; e\<^sub>2},(h\<^sub>1, l\<^sub>1, sh\<^sub>1),b\<^sub>1)"
- have bz: "(?b, ?z) \<in> (red P)\<^sup>*"
- by(rule InitBlockRedsFinal[OF e\<^sub>2_steps final])
- have hp: "hp (h\<^sub>1, l\<^sub>1, sh\<^sub>1) a = \<lfloor>(D, fs)\<rfloor>" using h\<^sub>1a by simp
- have "(?x, ?y) \<in> (red P)\<^sup>*" by(rule TryReds[OF e\<^sub>1_steps])
- also have "(?y, ?z) \<in> (red P)\<^sup>*"
- by(rule RedTryCatch[THEN converse_rtrancl_into_rtrancl, OF hp sub bz])
- ultimately show ?thesis by simp
-qed
-(*>*)
-
-lemma TryRedsFail:
- "\<lbrakk> P \<turnstile> \<langle>e\<^sub>1,s,b\<rangle> \<rightarrow>* \<langle>Throw a,(h,l,sh),b'\<rangle>; h a = Some(D,fs); \<not> P \<turnstile> D \<preceq>\<^sup>* C \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>try e\<^sub>1 catch(C V) e\<^sub>2,s,b\<rangle> \<rightarrow>* \<langle>Throw a,(h,l,sh),b'\<rangle>"
-(*<*)by(fastforce intro!: TryReds[THEN rtrancl_into_rtrancl, OF _ RedTryFail])(*>*)
-
-subsubsection "List"
-
-lemma ListReds1:
- "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>e#es,s,b\<rangle> [\<rightarrow>]* \<langle>e' # es,s',b'\<rangle>"
-(*<*)
-proof(induct rule: rtrancl_induct3)
- case refl show ?case by blast
-next
- case step show ?case
- by(rule rtrancl_into_rtrancl[OF step(3) ListRed1[OF step(2)]])
-qed
-(*>*)
-
-lemma ListReds2:
- "P \<turnstile> \<langle>es,s,b\<rangle> [\<rightarrow>]* \<langle>es',s',b'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>Val v # es,s,b\<rangle> [\<rightarrow>]* \<langle>Val v # es',s',b'\<rangle>"
-(*<*)
-proof(induct rule: rtrancl_induct3)
- case refl show ?case by blast
-next
- case step show ?case
- by(rule rtrancl_into_rtrancl[OF step(3) ListRed2[OF step(2)]])
-qed
-(*>*)
-
-lemma ListRedsVal:
- "\<lbrakk> P \<turnstile> \<langle>e,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>Val v,s\<^sub>1,b\<^sub>1\<rangle>; P \<turnstile> \<langle>es,s\<^sub>1,b\<^sub>1\<rangle> [\<rightarrow>]* \<langle>es',s\<^sub>2,b\<^sub>2\<rangle> \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>e#es,s\<^sub>0,b\<^sub>0\<rangle> [\<rightarrow>]* \<langle>Val v # es',s\<^sub>2,b\<^sub>2\<rangle>"
-(*<*)by(rule rtrancl_trans[OF ListReds1 ListReds2])(*>*)
-
-subsubsection"Call"
-
-text\<open> First a few lemmas on what happens to free variables during redction. \<close>
-
-lemma assumes wf: "wwf_J_prog P"
-shows Red_fv: "P \<turnstile> \<langle>e,(h,l,sh),b\<rangle> \<rightarrow> \<langle>e',(h',l',sh'),b'\<rangle> \<Longrightarrow> fv e' \<subseteq> fv e"
- and "P \<turnstile> \<langle>es,(h,l,sh),b\<rangle> [\<rightarrow>] \<langle>es',(h',l',sh'),b'\<rangle> \<Longrightarrow> fvs es' \<subseteq> fvs es"
-(*<*)
-proof (induct rule:red_reds_inducts)
- case (RedCall h a C fs M Ts T pns body D vs l sh b)
- hence "fv body \<subseteq> {this} \<union> set pns"
- using assms by(fastforce dest!:sees_wf_mdecl simp:wf_mdecl_def)
- with RedCall.hyps show ?case by fastforce
-next
- case (RedSCall C M Ts T pns body D vs)
- hence "fv body \<subseteq> set pns"
- using assms by(fastforce dest!:sees_wf_mdecl simp:wf_mdecl_def)
- with RedSCall.hyps show ?case by fastforce
-qed auto
-(*>*)
-
-
-lemma Red_dom_lcl:
- "P \<turnstile> \<langle>e,(h,l,sh),b\<rangle> \<rightarrow> \<langle>e',(h',l',sh'),b'\<rangle> \<Longrightarrow> dom l' \<subseteq> dom l \<union> fv e" and
- "P \<turnstile> \<langle>es,(h,l,sh),b\<rangle> [\<rightarrow>] \<langle>es',(h',l',sh'),b'\<rangle> \<Longrightarrow> dom l' \<subseteq> dom l \<union> fvs es"
-(*<*)
-proof (induct rule:red_reds_inducts)
- case RedLAss thus ?case by(force split:if_splits)
-next
- case CallParams thus ?case by(force split:if_splits)
-next
- case BlockRedNone thus ?case by clarsimp (fastforce split:if_splits)
-next
- case BlockRedSome thus ?case by clarsimp (fastforce split:if_splits)
-next
- case InitBlockRed thus ?case by clarsimp (fastforce split:if_splits)
-qed auto
-(*>*)
-
-lemma Reds_dom_lcl:
-assumes wf: "wwf_J_prog P"
-shows "P \<turnstile> \<langle>e,(h,l,sh),b\<rangle> \<rightarrow>* \<langle>e',(h',l',sh'),b'\<rangle> \<Longrightarrow> dom l' \<subseteq> dom l \<union> fv e"
-(*<*)
-proof(induct rule: converse_rtrancl_induct_red)
- case 1 then show ?case by blast
-next
- case 2 then show ?case using wf by(blast dest: Red_fv Red_dom_lcl)
-qed
-(*>*)
-
-text\<open> Now a few lemmas on the behaviour of blocks during reduction. \<close>
-
-lemma override_on_upd_lemma:
- "(override_on f (g(a\<mapsto>b)) A)(a := g a) = override_on f g (insert a A)"
-(*<*)by(rule ext) (simp add:override_on_def)
-
-declare fun_upd_apply[simp del] map_upds_twist[simp del]
-(*>*)
-
-
-lemma blocksReds:
- "\<And>l. \<lbrakk> length Vs = length Ts; length vs = length Ts; distinct Vs;
- P \<turnstile> \<langle>e, (h,l(Vs [\<mapsto>] vs),sh),b\<rangle> \<rightarrow>* \<langle>e', (h',l',sh'),b'\<rangle> \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>blocks(Vs,Ts,vs,e), (h,l,sh),b\<rangle> \<rightarrow>* \<langle>blocks(Vs,Ts,map (the \<circ> l') Vs,e'), (h',override_on l' l (set Vs),sh'),b'\<rangle>"
-(*<*)
-proof(induct Vs Ts vs e rule:blocks_induct)
- case (1 V Vs T Ts v vs e) show ?case
- using InitBlockReds[OF "1.hyps"[of "l(V\<mapsto>v)"]] "1.prems"
- by(auto simp:override_on_upd_lemma)
-qed auto
-(*>*)
-
-
-lemma blocksFinal:
- "\<And>l. \<lbrakk> length Vs = length Ts; length vs = length Ts; final e \<rbrakk> \<Longrightarrow>
- P \<turnstile> \<langle>blocks(Vs,Ts,vs,e), (h,l,sh),b\<rangle> \<rightarrow>* \<langle>e, (h,l,sh),b\<rangle>"
-(*<*)
-proof(induct Vs Ts vs e rule:blocks_induct)
- case 1
- show ?case using "1.prems" InitBlockReds[OF "1.hyps"]
- by(fastforce elim!:finalE elim: rtrancl_into_rtrancl[OF _ RedInitBlock]
- rtrancl_into_rtrancl[OF _ InitBlockThrow])
-qed auto
-(*>*)
-
-
-lemma blocksRedsFinal:
-assumes wf: "length Vs = length Ts" "length vs = length Ts" "distinct Vs"
- and reds: "P \<turnstile> \<langle>e, (h,l(Vs [\<mapsto>] vs),sh),b\<rangle> \<rightarrow>* \<langle>e', (h',l',sh'),b'\<rangle>"
- and fin: "final e'" and l'': "l'' = override_on l' l (set Vs)"
-shows "P \<turnstile> \<langle>blocks(Vs,Ts,vs,e), (h,l,sh),b\<rangle> \<rightarrow>* \<langle>e', (h',l'',sh'),b'\<rangle>"
-(*<*)
-proof -
- let ?bv = "blocks(Vs,Ts,map (the o l') Vs,e')"
- have "P \<turnstile> \<langle>blocks(Vs,Ts,vs,e), (h,l,sh),b\<rangle> \<rightarrow>* \<langle>?bv, (h',l'',sh'),b'\<rangle>"
- using l'' by simp (rule blocksReds[OF wf reds])
- also have "P \<turnstile> \<langle>?bv, (h',l'',sh'),b'\<rangle> \<rightarrow>* \<langle>e', (h',l'',sh'),b'\<rangle>"
- using wf by(fastforce intro:blocksFinal fin)
- finally show ?thesis .
-qed
-(*>*)
-
-text\<open> An now the actual method call reduction lemmas. \<close>
-
-lemma CallRedsObj:
- "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>e\<bullet>M(es),s,b\<rangle> \<rightarrow>* \<langle>e'\<bullet>M(es),s',b'\<rangle>"
-(*<*)
-proof(induct rule: rtrancl_induct3)
- case refl show ?case by blast
-next
- case step show ?case
- by(rule rtrancl_into_rtrancl[OF step(3) CallObj[OF step(2)]])
-qed
-(*>*)
-
-
-lemma CallRedsParams:
- "P \<turnstile> \<langle>es,s,b\<rangle> [\<rightarrow>]* \<langle>es',s',b'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>(Val v)\<bullet>M(es),s,b\<rangle> \<rightarrow>* \<langle>(Val v)\<bullet>M(es'),s',b'\<rangle>"
-(*<*)
-proof(induct rule: rtrancl_induct3)
- case refl show ?case by blast
-next
- case step show ?case
- by(rule rtrancl_into_rtrancl[OF step(3) CallParams[OF step(2)]])
-qed
-(*>*)
-
-
-lemma CallRedsFinal:
-assumes wwf: "wwf_J_prog P"
-and "P \<turnstile> \<langle>e,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>addr a,s\<^sub>1,b\<^sub>1\<rangle>"
- "P \<turnstile> \<langle>es,s\<^sub>1,b\<^sub>1\<rangle> [\<rightarrow>]* \<langle>map Val vs,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\<rangle>"
- "h\<^sub>2 a = Some(C,fs)" "P \<turnstile> C sees M,NonStatic:Ts\<rightarrow>T = (pns,body) in D"
- "size vs = size pns"
-and l\<^sub>2': "l\<^sub>2' = [this \<mapsto> Addr a, pns[\<mapsto>]vs]"
-and body: "P \<turnstile> \<langle>body,(h\<^sub>2,l\<^sub>2',sh\<^sub>2),b\<^sub>2\<rangle> \<rightarrow>* \<langle>ef,(h\<^sub>3,l\<^sub>3,sh\<^sub>3),b\<^sub>3\<rangle>"
-and "final ef"
-shows "P \<turnstile> \<langle>e\<bullet>M(es), s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>ef,(h\<^sub>3,l\<^sub>2,sh\<^sub>3),b\<^sub>3\<rangle>"
-(*<*)
-proof -
- have wf: "size Ts = size pns \<and> distinct pns \<and> this \<notin> set pns"
- and wt: "fv body \<subseteq> {this} \<union> set pns"
- using assms by(fastforce dest!:sees_wf_mdecl simp:wf_mdecl_def)+
- from body[THEN Red_lcl_add, of l\<^sub>2]
- have body': "P \<turnstile> \<langle>body,(h\<^sub>2,l\<^sub>2(this\<mapsto> Addr a, pns[\<mapsto>]vs),sh\<^sub>2),b\<^sub>2\<rangle> \<rightarrow>* \<langle>ef,(h\<^sub>3,l\<^sub>2++l\<^sub>3,sh\<^sub>3),b\<^sub>3\<rangle>"
- by (simp add:l\<^sub>2')
- have "dom l\<^sub>3 \<subseteq> {this} \<union> set pns"
- using Reds_dom_lcl[OF wwf body] wt l\<^sub>2' set_take_subset by force
- hence eql\<^sub>2: "override_on (l\<^sub>2++l\<^sub>3) l\<^sub>2 ({this} \<union> set pns) = l\<^sub>2"
- by(fastforce simp add:map_add_def override_on_def fun_eq_iff)
- have "P \<turnstile> \<langle>e\<bullet>M(es),s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>(addr a)\<bullet>M(es),s\<^sub>1,b\<^sub>1\<rangle>" by(rule CallRedsObj)(rule assms(2))
- also have "P \<turnstile> \<langle>(addr a)\<bullet>M(es),s\<^sub>1,b\<^sub>1\<rangle> \<rightarrow>*
- \<langle>(addr a)\<bullet>M(map Val vs),(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\<rangle>"
- by(rule CallRedsParams)(rule assms(3))
- also have "P \<turnstile> \<langle>(addr a)\<bullet>M(map Val vs), (h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\<rangle> \<rightarrow>
- \<langle>blocks(this#pns, Class D#Ts, Addr a#vs, body), (h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\<rangle>"
- by(rule RedCall)(auto simp: assms wf, rule assms(5))
- also (rtrancl_into_rtrancl) have "P \<turnstile> \<langle>blocks(this#pns, Class D#Ts, Addr a#vs, body), (h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\<rangle>
- \<rightarrow>* \<langle>ef,(h\<^sub>3,override_on (l\<^sub>2++l\<^sub>3) l\<^sub>2 ({this} \<union> set pns),sh\<^sub>3),b\<^sub>3\<rangle>"
- by(rule blocksRedsFinal, insert assms wf body', simp_all)
- finally show ?thesis using eql\<^sub>2 by simp
-qed
-(*>*)
-
-
-lemma CallRedsThrowParams:
-assumes e_steps: "P \<turnstile> \<langle>e,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>Val v,s\<^sub>1,b\<^sub>1\<rangle>"
- and es_steps: "P \<turnstile> \<langle>es,s\<^sub>1,b\<^sub>1\<rangle> [\<rightarrow>]* \<langle>map Val vs\<^sub>1 @ throw a # es\<^sub>2,s\<^sub>2,b\<^sub>2\<rangle>"
-shows "P \<turnstile> \<langle>e\<bullet>M(es),s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>throw a,s\<^sub>2,b\<^sub>2\<rangle>"
-(*<*)(is "(?x, ?z) \<in> (red P)\<^sup>*")
-proof -
- let ?y = "(Val v\<bullet>M(es),s\<^sub>1,b\<^sub>1)"
- let ?y' = "(Val v\<bullet>M(map Val vs\<^sub>1 @ throw a # es\<^sub>2),s\<^sub>2,b\<^sub>2)"
- have "(?x, ?y) \<in> (red P)\<^sup>*" by(rule CallRedsObj[OF e_steps])
- also have "(?y, ?y') \<in> (red P)\<^sup>*" by(rule CallRedsParams[OF es_steps])
- also have "(?y', ?z) \<in> (red P)\<^sup>*" using CallThrowParams by fast
- ultimately show ?thesis by simp
-qed
-(*>*)
-
-
-lemma CallRedsThrowObj:
- "P \<turnstile> \<langle>e,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>throw a,s\<^sub>1,b\<^sub>1\<rangle> \<Longrightarrow> P \<turnstile> \<langle>e\<bullet>M(es),s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>throw a,s\<^sub>1,b\<^sub>1\<rangle>"
-(*<*)by(rule CallRedsObj[THEN rtrancl_into_rtrancl, OF _ CallThrowObj])(*>*)
-
-
-lemma CallRedsNull:
-assumes e_steps: "P \<turnstile> \<langle>e,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>null,s\<^sub>1,b\<^sub>1\<rangle>"
- and es_steps: "P \<turnstile> \<langle>es,s\<^sub>1,b\<^sub>1\<rangle> [\<rightarrow>]* \<langle>map Val vs,s\<^sub>2,b\<^sub>2\<rangle>"
-shows "P \<turnstile> \<langle>e\<bullet>M(es),s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>THROW NullPointer,s\<^sub>2,b\<^sub>2\<rangle>"
-(*<*)(is "(?x, ?z) \<in> (red P)\<^sup>*")
-proof -
- let ?y = "(null\<bullet>M(es),s\<^sub>1,b\<^sub>1)"
- let ?y' = "(null\<bullet>M(map Val vs),s\<^sub>2,b\<^sub>2)"
- have "(?x, ?y) \<in> (red P)\<^sup>*" by(rule CallRedsObj[OF e_steps])
- also have "(?y, ?y') \<in> (red P)\<^sup>*" by(rule CallRedsParams[OF es_steps])
- also have "(?y', ?z) \<in> (red P)" by(rule RedCallNull)
- ultimately show ?thesis by simp
-qed
-(*>*)
-
-lemma CallRedsNone:
-assumes e_steps: "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>addr a,s\<^sub>1,b\<^sub>1\<rangle>"
- and es_steps: "P \<turnstile> \<langle>es,s\<^sub>1,b\<^sub>1\<rangle> [\<rightarrow>]* \<langle>map Val vs,s\<^sub>2,b\<^sub>2\<rangle>"
- and hp\<^sub>2a: "hp s\<^sub>2 a = Some(C,fs)"
- and ncM: "\<not>(\<exists>b Ts T m D. P \<turnstile> C sees M,b:Ts\<rightarrow>T = m in D)"
-shows "P \<turnstile> \<langle>e\<bullet>M(es),s,b\<rangle> \<rightarrow>* \<langle>THROW NoSuchMethodError,s\<^sub>2,b\<^sub>2\<rangle>"
-(*<*)(is "(?x, ?z) \<in> (red P)\<^sup>*")
-proof -
- let ?y = "(addr a\<bullet>M(es),s\<^sub>1,b\<^sub>1)"
- let ?y' = "(addr a\<bullet>M(map Val vs),s\<^sub>2,b\<^sub>2)"
- have "(?x, ?y) \<in> (red P)\<^sup>*" by(rule CallRedsObj[OF e_steps])
- also have "(?y, ?y') \<in> (red P)\<^sup>*" by(rule CallRedsParams[OF es_steps])
- also have "(?y', ?z) \<in> (red P)" using RedCallNone[OF _ ncM] hp\<^sub>2a
- by(cases s\<^sub>2) simp
- ultimately show ?thesis by simp
-qed
-(*>*)
-
-lemma CallRedsStatic:
-assumes e_steps: "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>addr a,s\<^sub>1,b\<^sub>1\<rangle>"
- and es_steps: "P \<turnstile> \<langle>es,s\<^sub>1,b\<^sub>1\<rangle> [\<rightarrow>]* \<langle>map Val vs,s\<^sub>2,b\<^sub>2\<rangle>"
- and hp\<^sub>2a: "hp s\<^sub>2 a = Some(C,fs)"
- and cM_Static: "P \<turnstile> C sees M,Static:Ts\<rightarrow>T = m in D"
-shows "P \<turnstile> \<langle>e\<bullet>M(es),s,b\<rangle> \<rightarrow>* \<langle>THROW IncompatibleClassChangeError,s\<^sub>2,b\<^sub>2\<rangle>"
-(*<*)(is "(?x, ?z) \<in> (red P)\<^sup>*")
-proof -
- let ?y = "(addr a\<bullet>M(es),s\<^sub>1,b\<^sub>1)"
- let ?y' = "(addr a\<bullet>M(map Val vs),s\<^sub>2,b\<^sub>2)"
- have "(?x, ?y) \<in> (red P)\<^sup>*" by(rule CallRedsObj[OF e_steps])
- also have "(?y, ?y') \<in> (red P)\<^sup>*" by(rule CallRedsParams[OF es_steps])
- also have "(?y', ?z) \<in> (red P)" using RedCallStatic[OF _ cM_Static] hp\<^sub>2a
- by(cases s\<^sub>2) simp
- ultimately show ?thesis by simp
-qed
-(*>*)
-
-subsection\<open>SCall\<close>
-
-lemma SCallRedsParams:
- "P \<turnstile> \<langle>es,s,b\<rangle> [\<rightarrow>]* \<langle>es',s',b'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sM(es),s,b\<rangle> \<rightarrow>* \<langle>C\<bullet>\<^sub>sM(es'),s',b'\<rangle>"
-(*<*)
-proof(induct rule: rtrancl_induct3)
- case refl show ?case by blast
-next
- case step show ?case
- by(rule rtrancl_into_rtrancl[OF step(3) SCallParams[OF step(2)]])
-qed
-(*>*)
-
-lemma SCallRedsFinal:
-assumes wwf: "wwf_J_prog P"
-and "P \<turnstile> \<langle>es,s\<^sub>0,b\<^sub>0\<rangle> [\<rightarrow>]* \<langle>map Val vs,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\<rangle>"
- "P \<turnstile> C sees M,Static:Ts\<rightarrow>T = (pns,body) in D"
- "sh\<^sub>2 D = Some(sfs,Done) \<or> (M = clinit \<and> sh\<^sub>2 D = \<lfloor>(sfs, Processing)\<rfloor>)"
- "size vs = size pns"
-and l\<^sub>2': "l\<^sub>2' = [pns[\<mapsto>]vs]"
-and body: "P \<turnstile> \<langle>body,(h\<^sub>2,l\<^sub>2',sh\<^sub>2),False\<rangle> \<rightarrow>* \<langle>ef,(h\<^sub>3,l\<^sub>3,sh\<^sub>3),b\<^sub>3\<rangle>"
-and "final ef"
-shows "P \<turnstile> \<langle>C\<bullet>\<^sub>sM(es), s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>ef,(h\<^sub>3,l\<^sub>2,sh\<^sub>3),b\<^sub>3\<rangle>"
-(*<*)
-proof -
- have wf: "size Ts = size pns \<and> distinct pns"
- and wt: "fv body \<subseteq> set pns"
- using assms by(fastforce dest!:sees_wf_mdecl simp:wf_mdecl_def)+
- from body[THEN Red_lcl_add, of l\<^sub>2]
- have body': "P \<turnstile> \<langle>body,(h\<^sub>2,l\<^sub>2(pns[\<mapsto>]vs),sh\<^sub>2),False\<rangle> \<rightarrow>* \<langle>ef,(h\<^sub>3,l\<^sub>2++l\<^sub>3,sh\<^sub>3),b\<^sub>3\<rangle>"
- by (simp add:l\<^sub>2')
- have "dom l\<^sub>3 \<subseteq> set pns"
- using Reds_dom_lcl[OF wwf body] wt l\<^sub>2' set_take_subset by force
- hence eql\<^sub>2: "override_on (l\<^sub>2++l\<^sub>3) l\<^sub>2 (set pns) = l\<^sub>2"
- by(fastforce simp add:map_add_def override_on_def fun_eq_iff)
- have b2T: "P \<turnstile> \<langle>C\<bullet>\<^sub>sM(map Val vs), (h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\<rangle> \<rightarrow>* \<langle>C\<bullet>\<^sub>sM(map Val vs), (h\<^sub>2,l\<^sub>2,sh\<^sub>2),True\<rangle>"
- proof(cases b\<^sub>2)
- case True then show ?thesis by simp
- next
- case False then show ?thesis using assms(3,4) by(auto elim: SCallInitDoneRed)
- qed
- have "P \<turnstile> \<langle>C\<bullet>\<^sub>sM(es),s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>C\<bullet>\<^sub>sM(map Val vs),(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\<rangle>"
- by(rule SCallRedsParams)(rule assms(2))
- also have "P \<turnstile> \<langle>C\<bullet>\<^sub>sM(map Val vs), (h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\<rangle> \<rightarrow>* \<langle>blocks(pns, Ts, vs, body), (h\<^sub>2,l\<^sub>2,sh\<^sub>2),False\<rangle>"
- by(auto intro!: rtrancl_into_rtrancl[OF b2T] RedSCall assms(3) simp: assms wf)
- also (rtrancl_trans) have "P \<turnstile> \<langle>blocks(pns, Ts, vs, body), (h\<^sub>2,l\<^sub>2,sh\<^sub>2),False\<rangle>
- \<rightarrow>* \<langle>ef,(h\<^sub>3,override_on (l\<^sub>2++l\<^sub>3) l\<^sub>2 (set pns),sh\<^sub>3),b\<^sub>3\<rangle>"
- by(rule blocksRedsFinal, insert assms wf body', simp_all)
- finally show ?thesis using eql\<^sub>2 by simp
-qed
-(*>*)
-
-lemma SCallRedsThrowParams:
- "\<lbrakk> P \<turnstile> \<langle>es,s\<^sub>0,b\<^sub>0\<rangle> [\<rightarrow>]* \<langle>map Val vs\<^sub>1 @ throw a # es\<^sub>2,s\<^sub>2,b\<^sub>2\<rangle> \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sM(es),s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>throw a,s\<^sub>2,b\<^sub>2\<rangle>"
-(*<*)
-by(erule SCallRedsParams[THEN rtrancl_into_rtrancl, OF _ SCallThrowParams])
- simp
-(*>*)
-
-lemma SCallRedsNone:
- "\<lbrakk> P \<turnstile> \<langle>es,s,b\<rangle> [\<rightarrow>]* \<langle>map Val vs,s\<^sub>2,False\<rangle>;
- \<not>(\<exists>b Ts T m D. P \<turnstile> C sees M,b:Ts\<rightarrow>T = m in D) \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sM(es),s,b\<rangle> \<rightarrow>* \<langle>THROW NoSuchMethodError,s\<^sub>2,False\<rangle>"
-(*<*)
-by(erule SCallRedsParams[THEN rtrancl_into_rtrancl, OF _ RedSCallNone])
- simp
-(*>*)
-
-lemma SCallRedsNonStatic:
- "\<lbrakk> P \<turnstile> \<langle>es,s,b\<rangle> [\<rightarrow>]* \<langle>map Val vs,s\<^sub>2,False\<rangle>;
- P \<turnstile> C sees M,NonStatic:Ts\<rightarrow>T = m in D \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sM(es),s,b\<rangle> \<rightarrow>* \<langle>THROW IncompatibleClassChangeError,s\<^sub>2,False\<rangle>"
-(*<*)
-by(erule SCallRedsParams[THEN rtrancl_into_rtrancl, OF _ RedSCallNonStatic])
- simp
-(*>*)
-
-lemma SCallInitThrowReds:
-assumes wwf: "wwf_J_prog P"
-and "P \<turnstile> \<langle>es,s\<^sub>0,b\<^sub>0\<rangle> [\<rightarrow>]* \<langle>map Val vs,(h\<^sub>1,l\<^sub>1,sh\<^sub>1),False\<rangle>"
- "P \<turnstile> C sees M,Static:Ts\<rightarrow>T = (pns,body) in D"
- "\<nexists>sfs. sh\<^sub>1 D = Some(sfs,Done)"
- "M \<noteq> clinit"
-and "P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(h\<^sub>1,l\<^sub>1,sh\<^sub>1),False\<rangle> \<rightarrow>* \<langle>throw a,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\<rangle>"
-shows "P \<turnstile> \<langle>C\<bullet>\<^sub>sM(es), s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>throw a,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\<rangle>"
-(*<*)(is "(?x, ?z) \<in> (red P)\<^sup>*")
-proof -
- let ?y = "(C\<bullet>\<^sub>sM(map Val vs),(h\<^sub>1,l\<^sub>1,sh\<^sub>1),False)"
- let ?y' = "(INIT D ([D],False) \<leftarrow> C\<bullet>\<^sub>sM(map Val vs),(h\<^sub>1,l\<^sub>1,sh\<^sub>1),False)"
- have "(?x, ?y) \<in> (red P)\<^sup>*" by(rule SCallRedsParams[OF assms(2)])
- also have "(?y, ?y') \<in> (red P)\<^sup>*"
- using SCallInitRed[OF assms(3)] assms(4-5) by auto
- also have "(?y', ?z) \<in> (red P)\<^sup>*" by(rule InitSeqThrowReds[OF assms(6)])
- finally show ?thesis by simp
-qed
-(*>*)
-
-lemma SCallInitReds:
-assumes wwf: "wwf_J_prog P"
-and "P \<turnstile> \<langle>es,s\<^sub>0,b\<^sub>0\<rangle> [\<rightarrow>]* \<langle>map Val vs,(h\<^sub>1,l\<^sub>1,sh\<^sub>1),False\<rangle>"
- "P \<turnstile> C sees M,Static:Ts\<rightarrow>T = (pns,body) in D"
- "\<nexists>sfs. sh\<^sub>1 D = Some(sfs,Done)"
- "M \<noteq> clinit"
-and "P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(h\<^sub>1,l\<^sub>1,sh\<^sub>1),False\<rangle> \<rightarrow>* \<langle>Val v',(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\<rangle>"
-and "size vs = size pns"
-and l\<^sub>2': "l\<^sub>2' = [pns[\<mapsto>]vs]"
-and body: "P \<turnstile> \<langle>body,(h\<^sub>2,l\<^sub>2',sh\<^sub>2),False\<rangle> \<rightarrow>* \<langle>ef,(h\<^sub>3,l\<^sub>3,sh\<^sub>3),b\<^sub>3\<rangle>"
-and "final ef"
-shows "P \<turnstile> \<langle>C\<bullet>\<^sub>sM(es),s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>ef,(h\<^sub>3,l\<^sub>2,sh\<^sub>3),b\<^sub>3\<rangle>"
-(*<*)
-proof -
- have wf: "size Ts = size pns \<and> distinct pns"
- and wt: "fv body \<subseteq> set pns"
- using assms by(fastforce dest!:sees_wf_mdecl simp:wf_mdecl_def)+
- from body[THEN Red_lcl_add, of l\<^sub>2]
- have body': "P \<turnstile> \<langle>body,(h\<^sub>2,l\<^sub>2(pns[\<mapsto>]vs),sh\<^sub>2),False\<rangle> \<rightarrow>* \<langle>ef,(h\<^sub>3,l\<^sub>2++l\<^sub>3,sh\<^sub>3),b\<^sub>3\<rangle>"
- by (simp add:l\<^sub>2')
- have "dom l\<^sub>3 \<subseteq> set pns"
- using Reds_dom_lcl[OF wwf body] wt l\<^sub>2' set_take_subset by force
- hence eql\<^sub>2: "override_on (l\<^sub>2++l\<^sub>3) l\<^sub>2 (set pns) = l\<^sub>2"
- by(fastforce simp add:map_add_def override_on_def fun_eq_iff)
- have "icheck P D (C\<bullet>\<^sub>sM(map Val vs)::'a exp)" using assms(3) by auto
- then have "P \<turnstile> \<langle>C\<bullet>\<^sub>sM(map Val vs),(h\<^sub>2, l\<^sub>2, sh\<^sub>2),icheck P D (C\<bullet>\<^sub>sM(map Val vs))\<rangle>
- \<rightarrow> \<langle>blocks(pns, Ts, vs, body), (h\<^sub>2, l\<^sub>2, sh\<^sub>2), False\<rangle>"
- by (metis (full_types) assms(3) assms(7) local.wf red_reds.RedSCall)
- also have "P \<turnstile> \<langle>blocks(pns, Ts, vs, body), (h\<^sub>2, l\<^sub>2, sh\<^sub>2), False\<rangle>
- \<rightarrow>* \<langle>ef,(h\<^sub>3,override_on (l\<^sub>2++l\<^sub>3) l\<^sub>2 (set pns),sh\<^sub>3),b\<^sub>3\<rangle>"
- by(rule blocksRedsFinal, insert assms wf body', simp_all)
- finally have trueReds: "P \<turnstile> \<langle>C\<bullet>\<^sub>sM(map Val vs),(h\<^sub>2, l\<^sub>2, sh\<^sub>2),icheck P D (C\<bullet>\<^sub>sM(map Val vs))\<rangle>
- \<rightarrow>* \<langle>ef,(h\<^sub>3,override_on (l\<^sub>2++l\<^sub>3) l\<^sub>2 (set pns),sh\<^sub>3),b\<^sub>3\<rangle>" by simp
- have "P \<turnstile> \<langle>C\<bullet>\<^sub>sM(es),s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>C\<bullet>\<^sub>sM(map Val vs),(h\<^sub>1,l\<^sub>1,sh\<^sub>1),False\<rangle>"
- by(rule SCallRedsParams)(rule assms(2))
- also have "P \<turnstile> \<langle>C\<bullet>\<^sub>sM(map Val vs),(h\<^sub>1,l\<^sub>1,sh\<^sub>1),False\<rangle> \<rightarrow> \<langle>INIT D ([D],False) \<leftarrow> C\<bullet>\<^sub>sM(map Val vs),(h\<^sub>1,l\<^sub>1,sh\<^sub>1),False\<rangle>"
- using SCallInitRed[OF assms(3)] assms(4-5) by auto
- also (rtrancl_into_rtrancl) have "P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> C\<bullet>\<^sub>sM(map Val vs),(h\<^sub>1,l\<^sub>1,sh\<^sub>1),False\<rangle>
- \<rightarrow>* \<langle>ef,(h\<^sub>3,override_on (l\<^sub>2++l\<^sub>3) l\<^sub>2 (set pns),sh\<^sub>3),b\<^sub>3\<rangle>"
- using InitSeqReds[OF assms(6) trueReds] assms(5) by simp
- finally show ?thesis using eql\<^sub>2 by simp
-qed
-(*>*)
-
-lemma SCallInitProcessingReds:
-assumes wwf: "wwf_J_prog P"
-and "P \<turnstile> \<langle>es,s\<^sub>0,b\<^sub>0\<rangle> [\<rightarrow>]* \<langle>map Val vs,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\<rangle>"
- "P \<turnstile> C sees M,Static:Ts\<rightarrow>T = (pns,body) in D"
- "sh\<^sub>2 D = Some(sfs,Processing)"
-and "size vs = size pns"
-and l\<^sub>2': "l\<^sub>2' = [pns[\<mapsto>]vs]"
-and body: "P \<turnstile> \<langle>body,(h\<^sub>2,l\<^sub>2',sh\<^sub>2),False\<rangle> \<rightarrow>* \<langle>ef,(h\<^sub>3,l\<^sub>3,sh\<^sub>3),b\<^sub>3\<rangle>"
-and "final ef"
-shows "P \<turnstile> \<langle>C\<bullet>\<^sub>sM(es),s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>ef,(h\<^sub>3,l\<^sub>2,sh\<^sub>3),b\<^sub>3\<rangle>"
-(*<*)
-proof -
- have wf: "size Ts = size pns \<and> distinct pns"
- and wt: "fv body \<subseteq> set pns"
- using assms by(fastforce dest!:sees_wf_mdecl simp:wf_mdecl_def)+
- from body[THEN Red_lcl_add, of l\<^sub>2]
- have body': "P \<turnstile> \<langle>body,(h\<^sub>2,l\<^sub>2(pns[\<mapsto>]vs),sh\<^sub>2),False\<rangle> \<rightarrow>* \<langle>ef,(h\<^sub>3,l\<^sub>2++l\<^sub>3,sh\<^sub>3),b\<^sub>3\<rangle>"
- by (simp add:l\<^sub>2')
- have "dom l\<^sub>3 \<subseteq> set pns"
- using Reds_dom_lcl[OF wwf body] wt l\<^sub>2' set_take_subset by force
- hence eql\<^sub>2: "override_on (l\<^sub>2++l\<^sub>3) l\<^sub>2 (set pns) = l\<^sub>2"
- by(fastforce simp add:map_add_def override_on_def fun_eq_iff)
- have b2T: "P \<turnstile> \<langle>C\<bullet>\<^sub>sM(map Val vs), (h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\<rangle> \<rightarrow>* \<langle>C\<bullet>\<^sub>sM(map Val vs), (h\<^sub>2,l\<^sub>2,sh\<^sub>2),True\<rangle>"
- proof(cases b\<^sub>2)
- case True then show ?thesis by simp
- next
- case False
- show ?thesis
- proof(cases "M = clinit")
- case True then show ?thesis using False assms(3) red_reds.SCallInitDoneRed assms(4)
- by (simp add: r_into_rtrancl)
- next
- case nclinit: False
- have icheck: "icheck P D (C\<bullet>\<^sub>sM(map Val vs))" using assms(3) by auto
- have "P \<turnstile> \<langle>C\<bullet>\<^sub>sM(map Val vs),(h\<^sub>2, l\<^sub>2, sh\<^sub>2),b\<^sub>2\<rangle>
- \<rightarrow> \<langle>INIT D ([D],False) \<leftarrow> C\<bullet>\<^sub>sM(map Val vs),(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\<rangle>"
- using SCallInitRed[OF assms(3)] assms(4) False nclinit by simp
- also have "P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> C\<bullet>\<^sub>sM(map Val vs),(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\<rangle>
- \<rightarrow> \<langle>INIT D (Nil,True) \<leftarrow> C\<bullet>\<^sub>sM(map Val vs),(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\<rangle>"
- using RedInitProcessing assms(4) by simp
- also have "P \<turnstile> \<langle>INIT D (Nil,True) \<leftarrow> C\<bullet>\<^sub>sM(map Val vs),(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\<rangle>
- \<rightarrow> \<langle>C\<bullet>\<^sub>sM(map Val vs),(h\<^sub>2, l\<^sub>2, sh\<^sub>2),True\<rangle>"
- using RedInit[of "C\<bullet>\<^sub>sM(map Val vs)" D _ _ _ P] icheck nclinit
- by (metis (full_types) nsub_RI_Vals sub_RI.simps(12))
- finally show ?thesis by simp
- qed
- qed
- have "P \<turnstile> \<langle>C\<bullet>\<^sub>sM(es),s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>C\<bullet>\<^sub>sM(map Val vs),(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\<rangle>"
- by(rule SCallRedsParams)(rule assms(2))
- also have "P \<turnstile> \<langle>C\<bullet>\<^sub>sM(map Val vs), (h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\<rangle> \<rightarrow>* \<langle>blocks(pns, Ts, vs, body), (h\<^sub>2,l\<^sub>2,sh\<^sub>2),False\<rangle>"
- by(auto intro!: rtrancl_into_rtrancl[OF b2T] RedSCall assms(3) simp: assms wf)
- also (rtrancl_trans) have "P \<turnstile> \<langle>blocks(pns, Ts, vs, body), (h\<^sub>2,l\<^sub>2,sh\<^sub>2),False\<rangle>
- \<rightarrow>* \<langle>ef,(h\<^sub>3,override_on (l\<^sub>2++l\<^sub>3) l\<^sub>2 (set pns),sh\<^sub>3),b\<^sub>3\<rangle>"
- by(rule blocksRedsFinal, insert assms wf body', simp_all)
- finally show ?thesis using eql\<^sub>2 by simp
-qed
-(*>*)
-
-(***********************************)
-
-subsubsection "The main Theorem"
-
-lemma assumes wwf: "wwf_J_prog P"
-shows big_by_small: "P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>
- \<Longrightarrow> (\<And>b. iconf (shp s) e \<Longrightarrow> P,shp s \<turnstile>\<^sub>b (e,b) \<surd> \<Longrightarrow> P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e',s',False\<rangle>)"
-and bigs_by_smalls: "P \<turnstile> \<langle>es,s\<rangle> [\<Rightarrow>] \<langle>es',s'\<rangle>
- \<Longrightarrow> (\<And>b. iconfs (shp s) es \<Longrightarrow> P,shp s \<turnstile>\<^sub>b (es,b) \<surd> \<Longrightarrow> P \<turnstile> \<langle>es,s,b\<rangle> [\<rightarrow>]* \<langle>es',s',False\<rangle>)"
-(*<*)
-proof (induct rule: eval_evals.inducts)
- case New show ?case
- proof(cases b)
- case True then show ?thesis using RedNew[OF New.hyps(2-4)] by auto
- next
- case False then show ?thesis using New.hyps(1) NewInitDoneReds[OF _ New.hyps(2-4)] by auto
- qed
-next
- case NewFail show ?case
- proof(cases b)
- case True then show ?thesis using RedNewFail[OF NewFail.hyps(2)] NewFail.hyps(3) by fastforce
- next
- case False
- then show ?thesis using NewInitDoneReds2[OF _ NewFail.hyps(2)] NewFail by fastforce
- qed
-next
- case (NewInit sh C h l v' h' l' sh' a FDTs h'') show ?case
- proof(cases b)
- case True
- then obtain sfs where shC: "sh C = \<lfloor>(sfs, Processing)\<rfloor>"
- using NewInit.hyps(1) NewInit.prems by(clarsimp simp: bconf_def initPD_def)
- then have s': "(h',l',sh') = (h,l,sh)" using NewInit.hyps(2) init_ProcessingE by clarsimp
- then show ?thesis using RedNew[OF NewInit.hyps(4-6)] True by auto
- next
- case False
- then have init: "P \<turnstile> \<langle>INIT C ([C],False) \<leftarrow> unit,(h, l, sh),False\<rangle> \<rightarrow>* \<langle>Val v',(h', l', sh'),False\<rangle>"
- using NewInit.hyps(3) by(auto simp: bconf_def)
- then show ?thesis using NewInit NewInitReds[OF _ init NewInit.hyps(4-6)] False
- has_fields_is_class[OF NewInit.hyps(5)] by auto
- qed
-next
- case (NewInitOOM sh C h l v' h' l' sh') show ?case
- proof(cases b)
- case True
- then obtain sfs where shC: "sh C = \<lfloor>(sfs, Processing)\<rfloor>"
- using NewInitOOM.hyps(1) NewInitOOM.prems by(clarsimp simp: bconf_def initPD_def)
- then have s': "(h',l',sh') = (h,l,sh)" using NewInitOOM.hyps(2) init_ProcessingE by clarsimp
- then show ?thesis using RedNewFail[OF NewInitOOM.hyps(4)] True r_into_rtrancl NewInitOOM.hyps(5)
- by auto
- next
- case False
- then have init: "P \<turnstile> \<langle>INIT C ([C],False) \<leftarrow> unit,(h, l, sh),False\<rangle> \<rightarrow>* \<langle>Val v',(h', l', sh'),False\<rangle>"
- using NewInitOOM.hyps(3) by(auto simp: bconf_def)
- then show ?thesis using NewInitOOM.hyps(1) NewInitOOMReds[OF _ init NewInitOOM.hyps(4)] False
- NewInitOOM.hyps(5) by auto
- qed
-next
- case (NewInitThrow sh C h l a s') show ?case
- proof(cases b)
- case True
- then obtain sfs where shC: "sh C = \<lfloor>(sfs, Processing)\<rfloor>"
- using NewInitThrow.hyps(1) NewInitThrow.prems by(clarsimp simp: bconf_def initPD_def)
- then show ?thesis using NewInitThrow.hyps(2) init_ProcessingE by blast
- next
- case False
- then have init: "P \<turnstile> \<langle>INIT C ([C],False) \<leftarrow> unit,(h, l, sh),b\<rangle> \<rightarrow>* \<langle>throw a,s',False\<rangle>"
- using NewInitThrow.hyps(3) by(auto simp: bconf_def)
- then show ?thesis using NewInitThrow NewInitThrowReds[of "(h,l,sh)" C P a s'] False by auto
- qed
-next
- case Cast then show ?case by(fastforce intro:CastRedsAddr)
-next
- case CastNull then show ?case by(fastforce intro: CastRedsNull)
-next
- case CastFail thus ?case by(fastforce intro!:CastRedsFail)
-next
- case CastThrow thus ?case by(fastforce dest!:eval_final intro:CastRedsThrow)
-next
- case Val then show ?case using exI[of _ b] by(simp add: bconf_def)
-next
- case (BinOp e\<^sub>1 s\<^sub>0 v\<^sub>1 s\<^sub>1 e\<^sub>2 v\<^sub>2 s\<^sub>2 bop v)
- show ?case
- proof(cases "val_of e\<^sub>1")
- case None
- then have iconf: "iconf (shp s\<^sub>0) e\<^sub>1" using None BinOp.prems by auto
- have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e\<^sub>1,b) \<surd>" using None BinOp.prems by auto
- then have b1: "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Val v\<^sub>1,s\<^sub>1,False\<rangle>" using iconf BinOp.hyps(2) by auto
- have binop: "P \<turnstile> \<langle>e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Val v\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2,s\<^sub>1,False\<rangle>" by(rule BinOp1Reds[OF b1])
- then have iconf2': "iconf (shp s\<^sub>1) e\<^sub>2" using Red_preserves_iconf[OF wwf binop] None BinOp by auto
- have "P,shp s\<^sub>1 \<turnstile>\<^sub>b (e\<^sub>2,False) \<surd>" by(simp add: bconf_def)
- then have b2: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>1,False\<rangle> \<rightarrow>* \<langle>Val v\<^sub>2,s\<^sub>2,False\<rangle>" using BinOp.hyps(4)[OF iconf2'] by auto
- then show ?thesis using BinOpRedsVal[OF b1 b2 BinOp.hyps(5)] by fast
- next
- case (Some a)
- then obtain b1 where b1: "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Val v\<^sub>1,s\<^sub>1,b1\<rangle>"
- by (metis (no_types, lifting) BinOp.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec)
- have binop: "P \<turnstile> \<langle>e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Val v\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2,s\<^sub>1,b1\<rangle>" by(rule BinOp1Reds[OF b1])
- then have iconf2': "iconf (shp s\<^sub>1) e\<^sub>2" using Red_preserves_iconf[OF wwf binop] BinOp by auto
- have bconf2: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e\<^sub>2,b) \<surd>" using BinOp.prems Some by simp
- then have "P,shp s\<^sub>1 \<turnstile>\<^sub>b (e\<^sub>2,b1) \<surd>" using Red_preserves_bconf[OF wwf binop BinOp.prems] by simp
- then have b2: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>1,b1\<rangle> \<rightarrow>* \<langle>Val v\<^sub>2,s\<^sub>2,False\<rangle>" using BinOp.hyps(4)[OF iconf2'] by auto
- then show ?thesis using BinOpRedsVal[OF b1 b2 BinOp.hyps(5)] by fast
- qed
-next
- case (BinOpThrow1 e\<^sub>1 s\<^sub>0 e s\<^sub>1 bop e\<^sub>2) show ?case
- proof(cases "val_of e\<^sub>1")
- case None
- then have "iconf (shp s\<^sub>0) e\<^sub>1" and "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e\<^sub>1,b) \<surd>" using BinOpThrow1.prems by auto
- then have b1: "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>throw e,s\<^sub>1,False\<rangle>" using BinOpThrow1.hyps(2) by auto
- then have "P \<turnstile> \<langle>e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>throw e,s\<^sub>1,False\<rangle>"
- using BinOpThrow1 None by(auto dest!:eval_final simp: BinOpRedsThrow1[OF b1])
- then show ?thesis by fast
- next
- case (Some a)
- then show ?thesis using eval_final_same[OF BinOpThrow1.hyps(1)] val_of_spec[OF Some] by auto
- qed
-next
- case (BinOpThrow2 e\<^sub>1 s\<^sub>0 v\<^sub>1 s\<^sub>1 e\<^sub>2 e s\<^sub>2 bop)
- show ?case
- proof(cases "val_of e\<^sub>1")
- case None
- then have iconf: "iconf (shp s\<^sub>0) e\<^sub>1" using None BinOpThrow2.prems by auto
- have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e\<^sub>1,b) \<surd>" using None BinOpThrow2.prems by auto
- then have b1: "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Val v\<^sub>1,s\<^sub>1,False\<rangle>" using iconf BinOpThrow2.hyps(2) by auto
- have binop: "P \<turnstile> \<langle>e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Val v\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2,s\<^sub>1,False\<rangle>" by(rule BinOp1Reds[OF b1])
- then have iconf2': "iconf (shp s\<^sub>1) e\<^sub>2" using Red_preserves_iconf[OF wwf binop] None BinOpThrow2 by auto
- have "P,shp s\<^sub>1 \<turnstile>\<^sub>b (e\<^sub>2,False) \<surd>" by(simp add: bconf_def)
- then have b2: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>1,False\<rangle> \<rightarrow>* \<langle>throw e,s\<^sub>2,False\<rangle>" using BinOpThrow2.hyps(4)[OF iconf2'] by auto
- then show ?thesis using BinOpRedsThrow2[OF b1 b2] by fast
- next
- case (Some a)
- then obtain b1 where b1: "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Val v\<^sub>1,s\<^sub>1,b1\<rangle>"
- by (metis (no_types, lifting) BinOpThrow2.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec)
- have binop: "P \<turnstile> \<langle>e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Val v\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2,s\<^sub>1,b1\<rangle>" by(rule BinOp1Reds[OF b1])
- then have iconf2': "iconf (shp s\<^sub>1) e\<^sub>2" using Red_preserves_iconf[OF wwf binop] BinOpThrow2 by auto
- have bconf2: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e\<^sub>2,b) \<surd>" using BinOpThrow2.prems Some by simp
- then have "P,shp s\<^sub>1 \<turnstile>\<^sub>b (e\<^sub>2,b1) \<surd>" using Red_preserves_bconf[OF wwf binop BinOpThrow2.prems] by simp
- then have b2: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>1,b1\<rangle> \<rightarrow>* \<langle>throw e,s\<^sub>2,False\<rangle>" using BinOpThrow2.hyps(4)[OF iconf2'] by auto
- then show ?thesis using BinOpRedsThrow2[OF b1 b2] by fast
- qed
-next
- case Var thus ?case by(auto dest:RedVar simp: bconf_def)
-next
- case LAss thus ?case by(auto dest: LAssRedsVal)
-next
- case LAssThrow thus ?case by(auto dest!:eval_final dest: LAssRedsThrow)
-next
- case FAcc thus ?case by(fastforce intro:FAccRedsVal)
-next
- case FAccNull thus ?case by(auto dest:FAccRedsNull)
-next
- case FAccThrow thus ?case by(auto dest!:eval_final dest:FAccRedsThrow)
-next
- case FAccNone then show ?case by(fastforce intro: FAccRedsNone)
-next
- case FAccStatic then show ?case by(fastforce intro: FAccRedsStatic)
-next
- case SFAcc show ?case
- proof(cases b)
- case True then show ?thesis using RedSFAcc SFAcc.hyps by auto
- next
- case False then show ?thesis using SFAcc.hyps SFAccInitDoneReds[OF SFAcc.hyps(1)] by auto
- qed
-next
- case (SFAccInit C F t D sh h l v' h' l' sh' sfs i v) show ?case
- proof(cases b)
- case True
- then obtain sfs where shC: "sh D = \<lfloor>(sfs, Processing)\<rfloor>"
- using SFAccInit.hyps(2) SFAccInit.prems by(clarsimp simp: bconf_def initPD_def)
- then have s': "(h',l',sh') = (h,l,sh)" using SFAccInit.hyps(3) init_ProcessingE by clarsimp
- then show ?thesis using RedSFAcc SFAccInit.hyps True by auto
- next
- case False
- then have init: "P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(h, l, sh),False\<rangle> \<rightarrow>* \<langle>Val v',(h', l', sh'),False\<rangle>"
- using SFAccInit.hyps(4) by(auto simp: bconf_def)
- then show ?thesis using SFAccInit SFAccInitReds[OF _ _ init] False by auto
- qed
-next
- case (SFAccInitThrow C F t D sh h l a s') show ?case
- proof(cases b)
- case True
- then obtain sfs where shC: "sh D = \<lfloor>(sfs, Processing)\<rfloor>"
- using SFAccInitThrow.hyps(2) SFAccInitThrow.prems(2) by(clarsimp simp: bconf_def initPD_def)
- then show ?thesis using SFAccInitThrow.hyps(3) init_ProcessingE by blast
- next
- case False
- then have init: "P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(h, l, sh),b\<rangle> \<rightarrow>* \<langle>throw a,s',False\<rangle>"
- using SFAccInitThrow.hyps(4) by(auto simp: bconf_def)
- then show ?thesis using SFAccInitThrow SFAccInitThrowReds False by auto
- qed
-next
- case SFAccNone then show ?case by(fastforce intro: SFAccRedsNone)
-next
- case SFAccNonStatic then show ?case by(fastforce intro: SFAccRedsNonStatic)
-next
- case (FAss e\<^sub>1 s\<^sub>0 a s\<^sub>1 e\<^sub>2 v h\<^sub>2 l\<^sub>2 sh\<^sub>2 C fs F t D fs' h\<^sub>2')
- show ?case
- proof(cases "val_of e\<^sub>1")
- case None
- then have iconf: "iconf (shp s\<^sub>0) e\<^sub>1" using None FAss.prems by auto
- have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e\<^sub>1,b) \<surd>" using None FAss.prems by auto
- then have b1: "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>addr a,s\<^sub>1,False\<rangle>" using iconf FAss.hyps(2) by auto
- have fass: "P \<turnstile> \<langle>e\<^sub>1\<bullet>F{D} := e\<^sub>2,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>addr a\<bullet>F{D} := e\<^sub>2,s\<^sub>1,False\<rangle>" by(rule FAssReds1[OF b1])
- then have iconf2': "iconf (shp s\<^sub>1) e\<^sub>2" using Red_preserves_iconf[OF wwf fass] None FAss by auto
- have "P,shp s\<^sub>1 \<turnstile>\<^sub>b (e\<^sub>2,False) \<surd>" by(simp add: bconf_def)
- then have b2: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>1,False\<rangle> \<rightarrow>* \<langle>Val v,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\<rangle>" using FAss.hyps(4)[OF iconf2'] by auto
- then show ?thesis using FAssRedsVal[OF b1 b2 FAss.hyps(6) FAss.hyps(5)[THEN sym]] FAss.hyps(7,8) by fast
- next
- case (Some a')
- then obtain b1 where b1: "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>addr a,s\<^sub>1,b1\<rangle>"
- by (metis (no_types, lifting) FAss.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec)
- have fass: "P \<turnstile> \<langle>e\<^sub>1\<bullet>F{D} := e\<^sub>2,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>addr a\<bullet>F{D} := e\<^sub>2,s\<^sub>1,b1\<rangle>" by(rule FAssReds1[OF b1])
- then have iconf2': "iconf (shp s\<^sub>1) e\<^sub>2" using Red_preserves_iconf[OF wwf fass] FAss by auto
- have bconf2: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e\<^sub>2,b) \<surd>" using FAss.prems Some by simp
- then have "P,shp s\<^sub>1 \<turnstile>\<^sub>b (e\<^sub>2,b1) \<surd>" using Red_preserves_bconf[OF wwf fass FAss.prems] by simp
- then have b2: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>1,b1\<rangle> \<rightarrow>* \<langle>Val v,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\<rangle>" using FAss.hyps(4)[OF iconf2'] by auto
- then show ?thesis using FAssRedsVal[OF b1 b2] FAss.hyps(5)[THEN sym] FAss.hyps(6-8) by fast
- qed
-next
- case (FAssNull e\<^sub>1 s\<^sub>0 s\<^sub>1 e\<^sub>2 v s\<^sub>2 F D)
- show ?case
- proof(cases "val_of e\<^sub>1")
- case None
- then have iconf: "iconf (shp s\<^sub>0) e\<^sub>1" using FAssNull.prems(1) by simp
- have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e\<^sub>1,b) \<surd>" using FAssNull.prems(2) None by simp
- then have b1: "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>null,s\<^sub>1,False\<rangle>" using FAssNull.hyps(2)[OF iconf] by auto
- have fass: "P \<turnstile> \<langle>e\<^sub>1\<bullet>F{D} := e\<^sub>2,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>null\<bullet>F{D} := e\<^sub>2,s\<^sub>1,False\<rangle>" by(rule FAssReds1[OF b1])
- then have iconf2': "iconf (shp s\<^sub>1) e\<^sub>2" using Red_preserves_iconf[OF wwf fass] None FAssNull by auto
- have "P,shp s\<^sub>1 \<turnstile>\<^sub>b (e\<^sub>2,False) \<surd>" by(simp add: bconf_def)
- then have b2: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>1,False\<rangle> \<rightarrow>* \<langle>Val v,s\<^sub>2,False\<rangle>" using FAssNull.hyps(4)[OF iconf2'] by auto
- then show ?thesis using FAssRedsNull[OF b1 b2] by fast
- next
- case (Some a')
- then obtain b1 where b1: "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>null,s\<^sub>1,b1\<rangle>"
- by (metis (no_types, lifting) FAssNull.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec)
- have fass: "P \<turnstile> \<langle>e\<^sub>1\<bullet>F{D} := e\<^sub>2,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>null\<bullet>F{D} := e\<^sub>2,s\<^sub>1,b1\<rangle>" by(rule FAssReds1[OF b1])
- then have iconf2': "iconf (shp s\<^sub>1) e\<^sub>2" using Red_preserves_iconf[OF wwf fass] FAssNull by auto
- have bconf2: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e\<^sub>2,b) \<surd>" using FAssNull.prems Some by simp
- then have "P,shp s\<^sub>1 \<turnstile>\<^sub>b (e\<^sub>2,b1) \<surd>" using Red_preserves_bconf[OF wwf fass FAssNull.prems] by simp
- then have b2: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>1,b1\<rangle> \<rightarrow>* \<langle>Val v,s\<^sub>2,False\<rangle>" using FAssNull.hyps(4)[OF iconf2'] by auto
- then show ?thesis using FAssRedsNull[OF b1 b2] by fast
- qed
-next
- case (FAssThrow1 e\<^sub>1 s\<^sub>0 e' s\<^sub>1 F D e\<^sub>2) show ?case
- proof(cases "val_of e\<^sub>1")
- case None
- then have "iconf (shp s\<^sub>0) e\<^sub>1" and "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e\<^sub>1,b) \<surd>" using FAssThrow1.prems by auto
- then have b1: "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>throw e',s\<^sub>1,False\<rangle>" using FAssThrow1.hyps(2) by auto
- then have "P \<turnstile> \<langle>e\<^sub>1\<bullet>F{D} := e\<^sub>2,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>throw e',s\<^sub>1,False\<rangle>"
- using FAssThrow1 None by(auto dest!:eval_final simp: FAssRedsThrow1[OF b1])
- then show ?thesis by fast
- next
- case (Some a)
- then show ?thesis using eval_final_same[OF FAssThrow1.hyps(1)] val_of_spec[OF Some] by auto
- qed
-next
- case (FAssThrow2 e\<^sub>1 s\<^sub>0 v s\<^sub>1 e\<^sub>2 e' s\<^sub>2 F D)
- show ?case
- proof(cases "val_of e\<^sub>1")
- case None
- then have iconf: "iconf (shp s\<^sub>0) e\<^sub>1" using None FAssThrow2.prems by auto
- have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e\<^sub>1,b) \<surd>" using None FAssThrow2.prems by auto
- then have b1: "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Val v,s\<^sub>1,False\<rangle>" using iconf FAssThrow2.hyps(2) by auto
- have fass: "P \<turnstile> \<langle>e\<^sub>1\<bullet>F{D} := e\<^sub>2,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Val v\<bullet>F{D} := e\<^sub>2,s\<^sub>1,False\<rangle>" by(rule FAssReds1[OF b1])
- then have iconf2': "iconf (shp s\<^sub>1) e\<^sub>2" using Red_preserves_iconf[OF wwf fass] None FAssThrow2 by auto
- have "P,shp s\<^sub>1 \<turnstile>\<^sub>b (e\<^sub>2,False) \<surd>" by(simp add: bconf_def)
- then have b2: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>1,False\<rangle> \<rightarrow>* \<langle>throw e',s\<^sub>2,False\<rangle>" using FAssThrow2.hyps(4)[OF iconf2'] by auto
- then show ?thesis using FAssRedsThrow2[OF b1 b2] by fast
- next
- case (Some a')
- then obtain b1 where b1: "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Val v,s\<^sub>1,b1\<rangle>"
- by (metis (no_types, lifting) FAssThrow2.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec)
- have fass: "P \<turnstile> \<langle>e\<^sub>1\<bullet>F{D} := e\<^sub>2,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Val v\<bullet>F{D} := e\<^sub>2,s\<^sub>1,b1\<rangle>" by(rule FAssReds1[OF b1])
- then have iconf2': "iconf (shp s\<^sub>1) e\<^sub>2" using Red_preserves_iconf[OF wwf fass] FAssThrow2 by auto
- have bconf2: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e\<^sub>2,b) \<surd>" using FAssThrow2.prems Some by simp
- then have "P,shp s\<^sub>1 \<turnstile>\<^sub>b (e\<^sub>2,b1) \<surd>" using Red_preserves_bconf[OF wwf fass FAssThrow2.prems] by simp
- then have b2: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>1,b1\<rangle> \<rightarrow>* \<langle>throw e',s\<^sub>2,False\<rangle>" using FAssThrow2.hyps(4)[OF iconf2'] by auto
- then show ?thesis using FAssRedsThrow2[OF b1 b2] by fast
- qed
-next
- case (FAssNone e\<^sub>1 s\<^sub>0 a s\<^sub>1 e\<^sub>2 v h\<^sub>2 l\<^sub>2 sh\<^sub>2 C fs F D)
- show ?case
- proof(cases "val_of e\<^sub>1")
- case None
- then have iconf: "iconf (shp s\<^sub>0) e\<^sub>1" using None FAssNone.prems by auto
- have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e\<^sub>1,b) \<surd>" using None FAssNone.prems by auto
- then have b1: "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>addr a,s\<^sub>1,False\<rangle>" using iconf FAssNone.hyps(2) by auto
- have fass: "P \<turnstile> \<langle>e\<^sub>1\<bullet>F{D} := e\<^sub>2,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>addr a\<bullet>F{D} := e\<^sub>2,s\<^sub>1,False\<rangle>" by(rule FAssReds1[OF b1])
- then have iconf2': "iconf (shp s\<^sub>1) e\<^sub>2" using Red_preserves_iconf[OF wwf fass] None FAssNone by auto
- have "P,shp s\<^sub>1 \<turnstile>\<^sub>b (e\<^sub>2,False) \<surd>" by(simp add: bconf_def)
- then have b2: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>1,False\<rangle> \<rightarrow>* \<langle>Val v,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\<rangle>" using FAssNone.hyps(4)[OF iconf2'] by auto
- then show ?thesis using FAssRedsNone[OF b1 b2 FAssNone.hyps(5,6)] by fast
- next
- case (Some a')
- then obtain b1 where b1: "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>addr a,s\<^sub>1,b1\<rangle>"
- by (metis (no_types, lifting) FAssNone.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec)
- have fass: "P \<turnstile> \<langle>e\<^sub>1\<bullet>F{D} := e\<^sub>2,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>addr a\<bullet>F{D} := e\<^sub>2,s\<^sub>1,b1\<rangle>" by(rule FAssReds1[OF b1])
- then have iconf2': "iconf (shp s\<^sub>1) e\<^sub>2" using Red_preserves_iconf[OF wwf fass] FAssNone by auto
- have bconf2: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e\<^sub>2,b) \<surd>" using FAssNone.prems Some by simp
- then have "P,shp s\<^sub>1 \<turnstile>\<^sub>b (e\<^sub>2,b1) \<surd>" using Red_preserves_bconf[OF wwf fass FAssNone.prems] by simp
- then have b2: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>1,b1\<rangle> \<rightarrow>* \<langle>Val v,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\<rangle>" using FAssNone.hyps(4)[OF iconf2'] by auto
- then show ?thesis using FAssRedsNone[OF b1 b2 FAssNone.hyps(5,6)] by fast
- qed
-next
- case (FAssStatic e\<^sub>1 s\<^sub>0 a s\<^sub>1 e\<^sub>2 v h\<^sub>2 l\<^sub>2 sh\<^sub>2 C fs F t D)
- show ?case
- proof(cases "val_of e\<^sub>1")
- case None
- then have iconf: "iconf (shp s\<^sub>0) e\<^sub>1" using None FAssStatic.prems by auto
- have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e\<^sub>1,b) \<surd>" using None FAssStatic.prems by auto
- then have b1: "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>addr a,s\<^sub>1,False\<rangle>" using iconf FAssStatic.hyps(2) by auto
- have fass: "P \<turnstile> \<langle>e\<^sub>1\<bullet>F{D} := e\<^sub>2,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>addr a\<bullet>F{D} := e\<^sub>2,s\<^sub>1,False\<rangle>" by(rule FAssReds1[OF b1])
- then have iconf2': "iconf (shp s\<^sub>1) e\<^sub>2" using Red_preserves_iconf[OF wwf fass] None FAssStatic by auto
- have "P,shp s\<^sub>1 \<turnstile>\<^sub>b (e\<^sub>2,False) \<surd>" by(simp add: bconf_def)
- then have b2: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>1,False\<rangle> \<rightarrow>* \<langle>Val v,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\<rangle>" using FAssStatic.hyps(4)[OF iconf2'] by auto
- then show ?thesis using FAssRedsStatic[OF b1 b2 FAssStatic.hyps(5,6)] by fast
- next
- case (Some a')
- then obtain b1 where b1: "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>addr a,s\<^sub>1,b1\<rangle>"
- by (metis (no_types, lifting) FAssStatic.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec)
- have fass: "P \<turnstile> \<langle>e\<^sub>1\<bullet>F{D} := e\<^sub>2,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>addr a\<bullet>F{D} := e\<^sub>2,s\<^sub>1,b1\<rangle>" by(rule FAssReds1[OF b1])
- then have iconf2': "iconf (shp s\<^sub>1) e\<^sub>2" using Red_preserves_iconf[OF wwf fass] FAssStatic by auto
- have bconf2: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e\<^sub>2,b) \<surd>" using FAssStatic.prems Some by simp
- then have "P,shp s\<^sub>1 \<turnstile>\<^sub>b (e\<^sub>2,b1) \<surd>" using Red_preserves_bconf[OF wwf fass FAssStatic.prems] by simp
- then have b2: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>1,b1\<rangle> \<rightarrow>* \<langle>Val v,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\<rangle>" using FAssStatic.hyps(4)[OF iconf2'] by auto
- then show ?thesis using FAssRedsStatic[OF b1 b2 FAssStatic.hyps(5,6)] by fast
- qed
-next
- case (SFAss e\<^sub>2 s\<^sub>0 v h\<^sub>1 l\<^sub>1 sh\<^sub>1 C F t D sfs sfs' sh\<^sub>1')
- show ?case
- proof(cases "val_of e\<^sub>2")
- case None
- then have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e\<^sub>2,b) \<surd>" using SFAss.prems(2) by simp
- then have b1: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Val v,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\<rangle>" using SFAss by auto
- thus ?thesis using SFAssRedsVal[OF b1 SFAss.hyps(3,4)] SFAss.hyps(5,6) by fast
- next
- case (Some a)
- then obtain b1 where b1: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Val v,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),b1\<rangle>"
- by (metis (no_types, lifting) SFAss.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec)
- thus ?thesis using SFAssRedsVal[OF b1 SFAss.hyps(3,4)] SFAss.hyps(5,6) by fast
- qed
-next
- case (SFAssInit e\<^sub>2 s\<^sub>0 v h\<^sub>1 l\<^sub>1 sh\<^sub>1 C F t D v' h' l' sh' sfs i sfs' sh'')
- then have iconf: "iconf (shp s\<^sub>0) e\<^sub>2" by simp
- show ?case
- proof(cases "val_of e\<^sub>2")
- case None
- then have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e\<^sub>2,b) \<surd>" using SFAssInit.prems(2) by simp
- then have reds: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Val v,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\<rangle>"
- using SFAssInit.hyps(2)[OF iconf bconf] by auto
- then have init: "P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\<rangle> \<rightarrow>* \<langle>Val v',(h', l', sh'),False\<rangle>"
- using SFAssInit.hyps(6) by(auto simp: bconf_def)
- then show ?thesis using SFAssInit SFAssInitReds[OF reds SFAssInit.hyps(3) _ init] by auto
- next
- case (Some v2) show ?thesis
- proof(cases b)
- case False
- then have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e\<^sub>2,b) \<surd>" by(simp add: bconf_def)
- then have reds: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Val v,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\<rangle>"
- using SFAssInit.hyps(2)[OF iconf bconf] by auto
- then have init: "P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\<rangle> \<rightarrow>* \<langle>Val v',(h', l', sh'),False\<rangle>"
- using SFAssInit.hyps(6) by(auto simp: bconf_def)
- then show ?thesis using SFAssInit SFAssInitReds[OF reds SFAssInit.hyps(3) _ init] by auto
- next
- case True
- have e\<^sub>2: "e\<^sub>2 = Val v2" using val_of_spec[OF Some] by simp
- then have vs: "v2 = v \<and> s\<^sub>0 = (h\<^sub>1, l\<^sub>1, sh\<^sub>1)" using eval_final_same[OF SFAssInit.hyps(1)] by simp
- then obtain sfs where shC: "sh\<^sub>1 D = \<lfloor>(sfs, Processing)\<rfloor>"
- using SFAssInit.hyps(3,4) SFAssInit.prems(2) Some True
- by(cases e\<^sub>2, auto simp: bconf_def initPD_def dest: sees_method_fun)
- then have s': "(h',l',sh') = (h\<^sub>1, l\<^sub>1, sh\<^sub>1)" using SFAssInit.hyps(5) init_ProcessingE by clarsimp
- then show ?thesis using SFAssInit.hyps(3,7-9) True e\<^sub>2 red_reds.RedSFAss vs by auto
- qed
- qed
-next
- case (SFAssInitThrow e\<^sub>2 s\<^sub>0 v h\<^sub>1 l\<^sub>1 sh\<^sub>1 C F t D a s')
- then have iconf: "iconf (shp s\<^sub>0) e\<^sub>2" by simp
- show ?case
- proof(cases "val_of e\<^sub>2")
- case None
- then have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e\<^sub>2,b) \<surd>" using SFAssInitThrow.prems(2) by simp
- then have reds: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Val v,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\<rangle>"
- using SFAssInitThrow.hyps(2)[OF iconf bconf] by auto
- then have init: "P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\<rangle> \<rightarrow>* \<langle>throw a,s',False\<rangle>"
- using SFAssInitThrow.hyps(6) by(auto simp: bconf_def)
- then show ?thesis using SFAssInitThrow SFAssInitThrowReds[OF reds _ _ init] by auto
- next
- case (Some v2) show ?thesis
- proof(cases b)
- case False
- then have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e\<^sub>2,b) \<surd>" by(simp add: bconf_def)
- then have reds: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Val v,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\<rangle>"
- using SFAssInitThrow.hyps(2)[OF iconf bconf] by auto
- then have init: "P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\<rangle> \<rightarrow>* \<langle>throw a,s',False\<rangle>"
- using SFAssInitThrow.hyps(6) by(auto simp: bconf_def)
- then show ?thesis using SFAssInitThrow SFAssInitThrowReds[OF reds _ _ init] by auto
- next
- case True
- obtain v2 where e\<^sub>2: "e\<^sub>2 = Val v2" using val_of_spec[OF Some] by simp
- then have vs: "v2 = v \<and> s\<^sub>0 = (h\<^sub>1, l\<^sub>1, sh\<^sub>1)"
- using eval_final_same[OF SFAssInitThrow.hyps(1)] by simp
- then obtain sfs where shC: "sh\<^sub>1 D = \<lfloor>(sfs, Processing)\<rfloor>"
- using SFAssInitThrow.hyps(4) SFAssInitThrow.prems(2) Some True
- by(cases e\<^sub>2, auto simp: bconf_def initPD_def dest: sees_method_fun)
- then show ?thesis using SFAssInitThrow.hyps(5) init_ProcessingE by blast
- qed
- qed
-next
- case (SFAssThrow e\<^sub>2 s\<^sub>0 e' s\<^sub>2 C F D)
- show ?case
- proof(cases "val_of e\<^sub>2")
- case None
- then have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e\<^sub>2,b) \<surd>" using SFAssThrow.prems(2) None by simp
- then have b1: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>throw e',s\<^sub>2,False\<rangle>" using SFAssThrow by auto
- thus ?thesis using SFAssRedsThrow[OF b1] by fast
- next
- case (Some a)
- then show ?thesis using eval_final_same[OF SFAssThrow.hyps(1)] val_of_spec[OF Some] by auto
- qed
-next
- case (SFAssNone e\<^sub>2 s\<^sub>0 v h\<^sub>2 l\<^sub>2 sh\<^sub>2 C F D)
- show ?case
- proof(cases "val_of e\<^sub>2")
- case None
- then have iconf: "iconf (shp s\<^sub>0) e\<^sub>2" using SFAssNone by simp
- then have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e\<^sub>2,b) \<surd>" using SFAssNone.prems(2) None by simp
- then have b1: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Val v,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\<rangle>" using SFAssNone.hyps(2) iconf by auto
- thus ?thesis using SFAssRedsNone[OF b1 SFAssNone.hyps(3)] by fast
- next
- case (Some a)
- then obtain b1 where b1: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Val v,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),b1\<rangle>"
- by (metis (no_types, lifting) SFAssNone.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec)
- thus ?thesis using SFAssRedsNone[OF b1 SFAssNone.hyps(3)] by fast
- qed
-next
- case (SFAssNonStatic e\<^sub>2 s\<^sub>0 v h\<^sub>2 l\<^sub>2 sh\<^sub>2 C F t D) show ?case
- proof(cases "val_of e\<^sub>2")
- case None
- then have iconf: "iconf (shp s\<^sub>0) e\<^sub>2" using SFAssNonStatic by simp
- then have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e\<^sub>2,b) \<surd>" using SFAssNonStatic.prems(2) None by simp
- then have b1: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Val v,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\<rangle>" using SFAssNonStatic.hyps(2) iconf by auto
- thus ?thesis using SFAssRedsNonStatic[OF b1 SFAssNonStatic.hyps(3)] by fast
- next
- case (Some a)
- then obtain b' where b1: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Val v,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),b'\<rangle>"
- by (metis (no_types, lifting) SFAssNonStatic.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec)
- thus ?thesis using SFAssRedsNonStatic[OF b1 SFAssNonStatic.hyps(3)] by fast
- qed
-next
- case (CallObjThrow e s\<^sub>0 e' s\<^sub>1 M ps) show ?case
- proof(cases "val_of e")
- case None
- then have "iconf (shp s\<^sub>0) e" and "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e,b) \<surd>" using CallObjThrow.prems by auto
- then have b1: "P \<turnstile> \<langle>e,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>throw e',s\<^sub>1,False\<rangle>" using CallObjThrow.hyps(2) by auto
- then have "P \<turnstile> \<langle>e\<bullet>M(ps),s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>throw e',s\<^sub>1,False\<rangle>"
- using CallObjThrow None by(auto dest!:eval_final simp: CallRedsThrowObj[OF b1])
- then show ?thesis by fast
- next
- case (Some a)
- then show ?thesis using eval_final_same[OF CallObjThrow.hyps(1)] val_of_spec[OF Some] by auto
- qed
-next
- case (CallNull e s\<^sub>0 s\<^sub>1 ps vs s\<^sub>2 M) show ?case
- proof(cases "val_of e")
- case None
- then have iconf: "iconf (shp s\<^sub>0) e" using CallNull.prems(1) by simp
- have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e,b) \<surd>" using CallNull.prems(2) None by simp
- then have b1: "P \<turnstile> \<langle>e,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>null,s\<^sub>1,False\<rangle>" using CallNull.hyps(2)[OF iconf] by auto
- have call: "P \<turnstile> \<langle>e\<bullet>M(ps),s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>null\<bullet>M(ps),s\<^sub>1,False\<rangle>" by(rule CallRedsObj[OF b1])
- then have iconf2': "iconfs (shp s\<^sub>1) ps" using Red_preserves_iconf[OF wwf call] None CallNull by auto
- have "P,shp s\<^sub>1 \<turnstile>\<^sub>b (ps,False) \<surd>" by(simp add: bconfs_def)
- then have b2: "P \<turnstile> \<langle>ps,s\<^sub>1,False\<rangle> [\<rightarrow>]* \<langle>map Val vs,s\<^sub>2,False\<rangle>" using CallNull.hyps(4)[OF iconf2'] by auto
- then show ?thesis using CallRedsNull[OF b1 b2] by fast
- next
- case (Some a')
- then obtain b1 where b1: "P \<turnstile> \<langle>e,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>null,s\<^sub>1,b1\<rangle>"
- by (metis (no_types, lifting) CallNull.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec)
- have fass: "P \<turnstile> \<langle>e\<bullet>M(ps),s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>null\<bullet>M(ps),s\<^sub>1,b1\<rangle>" by(rule CallRedsObj[OF b1])
- then have iconf2': "iconfs (shp s\<^sub>1) ps" using Red_preserves_iconf[OF wwf fass] CallNull by auto
- have bconf2: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (ps,b) \<surd>" using CallNull.prems Some by simp
- then have "P,shp s\<^sub>1 \<turnstile>\<^sub>b (ps,b1) \<surd>" using Red_preserves_bconf[OF wwf fass CallNull.prems] by simp
- then have b2: "P \<turnstile> \<langle>ps,s\<^sub>1,b1\<rangle> [\<rightarrow>]* \<langle>map Val vs,s\<^sub>2,False\<rangle>" using CallNull.hyps(4)[OF iconf2'] by auto
- then show ?thesis using CallRedsNull[OF b1 b2] by fast
- qed
-next
- case (CallParamsThrow e s\<^sub>0 v s\<^sub>1 es vs ex es' s\<^sub>2 M) show ?case
- proof(cases "val_of e")
- case None
- then have iconf: "iconf (shp s\<^sub>0) e" using CallParamsThrow.prems(1) by simp
- have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e,b) \<surd>" using CallParamsThrow.prems(2) None by simp
- then have b1: "P \<turnstile> \<langle>e,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Val v,s\<^sub>1,False\<rangle>" using CallParamsThrow.hyps(2)[OF iconf] by auto
- have call: "P \<turnstile> \<langle>e\<bullet>M(es),s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Val v\<bullet>M(es),s\<^sub>1,False\<rangle>" by(rule CallRedsObj[OF b1])
- then have iconf2': "iconfs (shp s\<^sub>1) es" using Red_preserves_iconf[OF wwf call] None CallParamsThrow by auto
- have "P,shp s\<^sub>1 \<turnstile>\<^sub>b (es,False) \<surd>" by(simp add: bconfs_def)
- then have b2: "P \<turnstile> \<langle>es,s\<^sub>1,False\<rangle> [\<rightarrow>]* \<langle>map Val vs @ throw ex # es',s\<^sub>2,False\<rangle>"
- using CallParamsThrow.hyps(4)[OF iconf2'] by auto
- then show ?thesis using CallRedsThrowParams[OF b1 b2] by fast
- next
- case (Some a')
- then obtain b1 where b1: "P \<turnstile> \<langle>e,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Val v,s\<^sub>1,b1\<rangle>"
- by (metis (no_types, lifting) CallParamsThrow.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec)
- have fass: "P \<turnstile> \<langle>e\<bullet>M(es),s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Val v\<bullet>M(es),s\<^sub>1,b1\<rangle>" by(rule CallRedsObj[OF b1])
- then have iconf2': "iconfs (shp s\<^sub>1) es" using Red_preserves_iconf[OF wwf fass] CallParamsThrow by auto
- have bconf2: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (es,b) \<surd>" using CallParamsThrow.prems Some by simp
- then have "P,shp s\<^sub>1 \<turnstile>\<^sub>b (es,b1) \<surd>" using Red_preserves_bconf[OF wwf fass CallParamsThrow.prems] by simp
- then have b2: "P \<turnstile> \<langle>es,s\<^sub>1,b1\<rangle> [\<rightarrow>]* \<langle>map Val vs @ throw ex # es',s\<^sub>2,False\<rangle>"
- using CallParamsThrow.hyps(4)[OF iconf2'] by auto
- then show ?thesis using CallRedsThrowParams[OF b1 b2] by fast
- qed
-next
- case (CallNone e s\<^sub>0 a s\<^sub>1 ps vs h\<^sub>2 l\<^sub>2 sh\<^sub>2 C fs M) show ?case
- proof(cases "val_of e")
- case None
- then have iconf: "iconf (shp s\<^sub>0) e" using CallNone.prems(1) by simp
- have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e,b) \<surd>" using CallNone.prems(2) None by simp
- then have b1: "P \<turnstile> \<langle>e,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>addr a,s\<^sub>1,False\<rangle>" using CallNone.hyps(2)[OF iconf] by auto
- have call: "P \<turnstile> \<langle>e\<bullet>M(ps),s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>addr a\<bullet>M(ps),s\<^sub>1,False\<rangle>" by(rule CallRedsObj[OF b1])
- then have iconf2': "iconfs (shp s\<^sub>1) ps" using Red_preserves_iconf[OF wwf call] None CallNone by auto
- have "P,shp s\<^sub>1 \<turnstile>\<^sub>b (ps,False) \<surd>" by(simp add: bconfs_def)
- then have b2: "P \<turnstile> \<langle>ps,s\<^sub>1,False\<rangle> [\<rightarrow>]* \<langle>map Val vs,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\<rangle>"
- using CallNone.hyps(4)[OF iconf2'] by auto
- then show ?thesis using CallRedsNone[OF b1 b2 _ CallNone.hyps(6)] CallNone.hyps(5) by fastforce
- next
- case (Some a')
- then obtain b1 where b1: "P \<turnstile> \<langle>e,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>addr a,s\<^sub>1,b1\<rangle>"
- by (metis (no_types, lifting) CallNone.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec)
- have fass: "P \<turnstile> \<langle>e\<bullet>M(ps),s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>addr a\<bullet>M(ps),s\<^sub>1,b1\<rangle>" by(rule CallRedsObj[OF b1])
- then have iconf2': "iconfs (shp s\<^sub>1) ps" using Red_preserves_iconf[OF wwf fass] CallNone by auto
- have bconf2: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (ps,b) \<surd>" using CallNone.prems Some by simp
- then have "P,shp s\<^sub>1 \<turnstile>\<^sub>b (ps,b1) \<surd>" using Red_preserves_bconf[OF wwf fass CallNone.prems] by simp
- then have b2: "P \<turnstile> \<langle>ps,s\<^sub>1,b1\<rangle> [\<rightarrow>]* \<langle>map Val vs,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\<rangle>"
- using CallNone.hyps(4)[OF iconf2'] by auto
- then show ?thesis using CallRedsNone[OF b1 b2 _ CallNone.hyps(6)] CallNone.hyps(5) by fastforce
- qed
-next
- case (CallStatic e s\<^sub>0 a s\<^sub>1 ps vs h\<^sub>2 l\<^sub>2 sh\<^sub>2 C fs M Ts T m D) show ?case
- proof(cases "val_of e")
- case None
- then have iconf: "iconf (shp s\<^sub>0) e" using CallStatic.prems(1) by simp
- have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e,b) \<surd>" using CallStatic.prems(2) None by simp
- then have b1: "P \<turnstile> \<langle>e,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>addr a,s\<^sub>1,False\<rangle>" using CallStatic.hyps(2)[OF iconf] by auto
- have call: "P \<turnstile> \<langle>e\<bullet>M(ps),s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>addr a\<bullet>M(ps),s\<^sub>1,False\<rangle>" by(rule CallRedsObj[OF b1])
- then have iconf2': "iconfs (shp s\<^sub>1) ps" using Red_preserves_iconf[OF wwf call] None CallStatic by auto
- have "P,shp s\<^sub>1 \<turnstile>\<^sub>b (ps,False) \<surd>" by(simp add: bconfs_def)
- then have b2: "P \<turnstile> \<langle>ps,s\<^sub>1,False\<rangle> [\<rightarrow>]* \<langle>map Val vs,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\<rangle>"
- using CallStatic.hyps(4)[OF iconf2'] by auto
- then show ?thesis using CallRedsStatic[OF b1 b2 _ CallStatic.hyps(6)] CallStatic.hyps(5) by fastforce
- next
- case (Some a')
- then obtain b1 where b1: "P \<turnstile> \<langle>e,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>addr a,s\<^sub>1,b1\<rangle>"
- by (metis (no_types, lifting) CallStatic.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec)
- have call: "P \<turnstile> \<langle>e\<bullet>M(ps),s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>addr a\<bullet>M(ps),s\<^sub>1,b1\<rangle>" by(rule CallRedsObj[OF b1])
- then have iconf2': "iconfs (shp s\<^sub>1) ps" using Red_preserves_iconf[OF wwf call] CallStatic by auto
- have bconf2: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (ps,b) \<surd>" using CallStatic.prems Some by simp
- then have "P,shp s\<^sub>1 \<turnstile>\<^sub>b (ps,b1) \<surd>" using Red_preserves_bconf[OF wwf call CallStatic.prems] by simp
- then have b2: "P \<turnstile> \<langle>ps,s\<^sub>1,b1\<rangle> [\<rightarrow>]* \<langle>map Val vs,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\<rangle>"
- using CallStatic.hyps(4)[OF iconf2'] by auto
- then show ?thesis using CallRedsStatic[OF b1 b2 _ CallStatic.hyps(6)] CallStatic.hyps(5) by fastforce
- qed
-next
- case (Call e s\<^sub>0 a s\<^sub>1 ps vs h\<^sub>2 l\<^sub>2 sh\<^sub>2 C fs M Ts T pns body D l\<^sub>2' e' h\<^sub>3 l\<^sub>3 sh\<^sub>3) show ?case
- proof(cases "val_of e")
- case None
- then have iconf: "iconf (shp s\<^sub>0) e" using Call.prems(1) by simp
- have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e,b) \<surd>" using Call.prems(2) None by simp
- then have b1: "P \<turnstile> \<langle>e,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>addr a,s\<^sub>1,False\<rangle>" using Call.hyps(2)[OF iconf] by auto
- have call: "P \<turnstile> \<langle>e\<bullet>M(ps),s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>addr a\<bullet>M(ps),s\<^sub>1,False\<rangle>" by(rule CallRedsObj[OF b1])
- then have iconf2: "iconfs (shp s\<^sub>1) ps" using Red_preserves_iconf[OF wwf call] None Call by auto
- have "P,shp s\<^sub>1 \<turnstile>\<^sub>b (ps,False) \<surd>" by(simp add: bconfs_def)
- then have b2: "P \<turnstile> \<langle>ps,s\<^sub>1,False\<rangle> [\<rightarrow>]* \<langle>map Val vs,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\<rangle>"
- using Call.hyps(4)[OF iconf2] by simp
- have iconf3: "iconf (shp (h\<^sub>2, l\<^sub>2', sh\<^sub>2)) body"
- by(rule nsub_RI_iconf[OF sees_wwf_nsub_RI[OF wwf Call.hyps(6)]])
- have "P,shp (h\<^sub>2, l\<^sub>2', sh\<^sub>2) \<turnstile>\<^sub>b (body,False) \<surd>" by(simp add: bconf_def)
- then have b3: "P \<turnstile> \<langle>body,(h\<^sub>2, l\<^sub>2', sh\<^sub>2),False\<rangle> \<rightarrow>* \<langle>e',(h\<^sub>3, l\<^sub>3, sh\<^sub>3),False\<rangle>"
- using Call.hyps(10)[OF iconf3] by simp
- show ?thesis by(rule CallRedsFinal[OF wwf b1 b2 Call.hyps(5-8) b3 eval_final[OF Call.hyps(9)]])
- next
- case (Some a')
- then obtain b1 where b1: "P \<turnstile> \<langle>e,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>addr a,s\<^sub>1,b1\<rangle>"
- by (metis (no_types, lifting) Call.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec)
- have call: "P \<turnstile> \<langle>e\<bullet>M(ps),s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>addr a\<bullet>M(ps),s\<^sub>1,b1\<rangle>" by(rule CallRedsObj[OF b1])
- then have iconf2': "iconfs (shp s\<^sub>1) ps" using Red_preserves_iconf[OF wwf call] Call by auto
- have bconf2: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (ps,b) \<surd>" using Call.prems Some by simp
- then have "P,shp s\<^sub>1 \<turnstile>\<^sub>b (ps,b1) \<surd>" using Red_preserves_bconf[OF wwf call Call.prems] by simp
- then have b2: "P \<turnstile> \<langle>ps,s\<^sub>1,b1\<rangle> [\<rightarrow>]* \<langle>map Val vs,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\<rangle>"
- using Call.hyps(4)[OF iconf2'] by auto
- have iconf3: "iconf (shp (h\<^sub>2, l\<^sub>2', sh\<^sub>2)) body"
- by(rule nsub_RI_iconf[OF sees_wwf_nsub_RI[OF wwf Call.hyps(6)]])
- have "P,shp (h\<^sub>2, l\<^sub>2', sh\<^sub>2) \<turnstile>\<^sub>b (body,False) \<surd>" by(simp add: bconf_def)
- then have b3: "P \<turnstile> \<langle>body,(h\<^sub>2, l\<^sub>2', sh\<^sub>2),False\<rangle> \<rightarrow>* \<langle>e',(h\<^sub>3, l\<^sub>3, sh\<^sub>3),False\<rangle>"
- using Call.hyps(10)[OF iconf3] by simp
- show ?thesis by(rule CallRedsFinal[OF wwf b1 b2 Call.hyps(5-8) b3 eval_final[OF Call.hyps(9)]])
- qed
-next
- case (SCallParamsThrow es s\<^sub>0 vs ex es' s\<^sub>2 C M) show ?case
- proof(cases "map_vals_of es")
- case None
- then have iconf: "iconfs (shp s\<^sub>0) es" using SCallParamsThrow.prems(1) by simp
- have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (es,b) \<surd>" using SCallParamsThrow.prems(2) None by simp
- then have b1: "P \<turnstile> \<langle>es,s\<^sub>0,b\<rangle> [\<rightarrow>]* \<langle>map Val vs @ throw ex # es',s\<^sub>2,False\<rangle>"
- using SCallParamsThrow.hyps(2)[OF iconf] by simp
- show ?thesis using SCallRedsThrowParams[OF b1] by simp
- next
- case (Some vs')
- then have "es = map Val vs'" by(rule map_vals_of_spec)
- then show ?thesis using evals_finals_same[OF _ SCallParamsThrow.hyps(1)] map_Val_nthrow_neq
- by auto
- qed
-next
- case (SCallNone ps s\<^sub>0 vs s\<^sub>2 C M) show ?case
- proof(cases "map_vals_of ps")
- case None
- then have iconf: "iconfs (shp s\<^sub>0) ps" using SCallNone.prems(1) by simp
- have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (ps,b) \<surd>" using SCallNone.prems(2) None by simp
- then have b1: "P \<turnstile> \<langle>ps,s\<^sub>0,b\<rangle> [\<rightarrow>]* \<langle>map Val vs,s\<^sub>2,False\<rangle>" using SCallNone.hyps(2)[OF iconf] by auto
- then show ?thesis using SCallRedsNone[OF b1 SCallNone.hyps(3)] SCallNone.hyps(1) by simp
- next
- case (Some vs')
- then have ps: "ps = map Val vs'" by(rule map_vals_of_spec)
- then have s\<^sub>0: "s\<^sub>0 = s\<^sub>2" using SCallNone.hyps(1) evals_finals_same by blast
- then show ?thesis using RedSCallNone[OF SCallNone.hyps(3)] ps by(cases s\<^sub>2, auto)
- qed
-next
- case (SCallNonStatic ps s\<^sub>0 vs s\<^sub>2 C M Ts T m D) show ?case
- proof(cases "map_vals_of ps")
- case None
- then have iconf: "iconfs (shp s\<^sub>0) ps" using SCallNonStatic.prems(1) by simp
- have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (ps,b) \<surd>" using SCallNonStatic.prems(2) None by simp
- then have b1: "P \<turnstile> \<langle>ps,s\<^sub>0,b\<rangle> [\<rightarrow>]* \<langle>map Val vs,s\<^sub>2,False\<rangle>" using SCallNonStatic.hyps(2)[OF iconf] by auto
- then show ?thesis using SCallRedsNonStatic[OF b1 SCallNonStatic.hyps(3)] SCallNonStatic.hyps(1) by simp
- next
- case (Some vs')
- then have ps: "ps = map Val vs'" by(rule map_vals_of_spec)
- then have s\<^sub>0: "s\<^sub>0 = s\<^sub>2" using SCallNonStatic.hyps(1) evals_finals_same by blast
- then show ?thesis using RedSCallNonStatic[OF SCallNonStatic.hyps(3)] ps by(cases s\<^sub>2, auto)
- qed
-next
- case (SCallInitThrow ps s\<^sub>0 vs h\<^sub>1 l\<^sub>1 sh\<^sub>1 C M Ts T pns body D a s') show ?case
- proof(cases "map_vals_of ps")
- case None
- then have iconf: "iconfs (shp s\<^sub>0) ps" using SCallInitThrow.prems(1) by simp
- have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (ps,b) \<surd>" using SCallInitThrow.prems(2) None by simp
- then have b1: "P \<turnstile> \<langle>ps,s\<^sub>0,b\<rangle> [\<rightarrow>]* \<langle>map Val vs,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\<rangle>"
- using SCallInitThrow.hyps(2)[OF iconf] by auto
- have bconf2: "P,shp (h\<^sub>1, l\<^sub>1, sh\<^sub>1) \<turnstile>\<^sub>b (INIT D ([D],False) \<leftarrow> unit,False) \<surd>" by(simp add: bconf_def)
- then have b2: "P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\<rangle> \<rightarrow>* \<langle>throw a,s',False\<rangle>"
- using SCallInitThrow.hyps(7) by auto
- then show ?thesis using SCallInitThrowReds[OF wwf b1 SCallInitThrow.hyps(3-5)]
- by(cases s', auto)
- next
- case (Some vs')
- have ps: "ps = map Val vs'" by(rule map_vals_of_spec[OF Some])
- then have vs: "vs = vs' \<and> s\<^sub>0 = (h\<^sub>1, l\<^sub>1, sh\<^sub>1)"
- using evals_finals_same[OF _ SCallInitThrow.hyps(1)] map_Val_eq by auto
- show ?thesis
- proof(cases b)
- case True
- obtain sfs where shC: "sh\<^sub>1 D = \<lfloor>(sfs, Processing)\<rfloor>"
- using SCallInitThrow.hyps(3,4) SCallInitThrow.prems(2) True Some vs
- by(auto simp: bconf_def initPD_def dest: sees_method_fun)
- then show ?thesis using init_ProcessingE[OF _ SCallInitThrow.hyps(6)] by blast
- next
- case False
- then have b1: "P \<turnstile> \<langle>ps,s\<^sub>0,b\<rangle> [\<rightarrow>]* \<langle>map Val vs,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\<rangle>" using ps vs by simp
- have bconf2: "P,shp (h\<^sub>1, l\<^sub>1, sh\<^sub>1) \<turnstile>\<^sub>b (INIT D ([D],False) \<leftarrow> unit,False) \<surd>" by(simp add: bconf_def)
- then have b2: "P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\<rangle> \<rightarrow>* \<langle>throw a,s',False\<rangle>"
- using SCallInitThrow.hyps(7) by auto
- then show ?thesis using SCallInitThrowReds[OF wwf b1 SCallInitThrow.hyps(3-5)] by(cases s', auto)
- qed
- qed
-next
- case (SCallInit ps s\<^sub>0 vs h\<^sub>1 l\<^sub>1 sh\<^sub>1 C M Ts T pns body D v' h\<^sub>2 l\<^sub>2 sh\<^sub>2 l\<^sub>2' e' h\<^sub>3 l\<^sub>3 sh\<^sub>3) show ?case
- proof(cases "map_vals_of ps")
- case None
- then have iconf: "iconfs (shp s\<^sub>0) ps" using SCallInit.prems(1) by simp
- have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (ps,b) \<surd>" using SCallInit.prems(2) None by simp
- then have b1: "P \<turnstile> \<langle>ps,s\<^sub>0,b\<rangle> [\<rightarrow>]* \<langle>map Val vs,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\<rangle>"
- using SCallInit.hyps(2)[OF iconf] by auto
- have bconf2: "P,shp (h\<^sub>1, l\<^sub>1, sh\<^sub>1) \<turnstile>\<^sub>b (INIT D ([D],False) \<leftarrow> unit,False) \<surd>" by(simp add: bconf_def)
- then have b2: "P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\<rangle> \<rightarrow>* \<langle>Val v',(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\<rangle>"
- using SCallInit.hyps(7) by auto
- have iconf3: "iconf (shp (h\<^sub>2, l\<^sub>2', sh\<^sub>2)) body"
- by(rule nsub_RI_iconf[OF sees_wwf_nsub_RI[OF wwf SCallInit.hyps(3)]])
- have "P,shp (h\<^sub>2, l\<^sub>2', sh\<^sub>2) \<turnstile>\<^sub>b (body,False) \<surd>" by(simp add: bconf_def)
- then have b3: "P \<turnstile> \<langle>body,(h\<^sub>2, l\<^sub>2', sh\<^sub>2),False\<rangle> \<rightarrow>* \<langle>e',(h\<^sub>3, l\<^sub>3, sh\<^sub>3),False\<rangle>"
- using SCallInit.hyps(11)[OF iconf3] by simp
- show ?thesis by(rule SCallInitReds[OF wwf b1 SCallInit.hyps(3-5) b2 SCallInit.hyps(8-9)
- b3 eval_final[OF SCallInit.hyps(10)]])
- next
- case (Some vs')
- then have ps: "ps = map Val vs'" by(rule map_vals_of_spec)
- then have vs: "vs = vs' \<and> s\<^sub>0 = (h\<^sub>1, l\<^sub>1, sh\<^sub>1)"
- using evals_finals_same[OF _ SCallInit.hyps(1)] map_Val_eq by auto
- show ?thesis
- proof(cases b)
- case True
- then have b1: "P \<turnstile> \<langle>ps,s\<^sub>0,b\<rangle> [\<rightarrow>]* \<langle>map Val vs,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),b\<rangle>" using ps vs by simp
- obtain sfs where shC: "sh\<^sub>1 D = \<lfloor>(sfs, Processing)\<rfloor>"
- using SCallInit.hyps(3,4) SCallInit.prems(2) True Some vs
- by(auto simp: bconf_def initPD_def dest: sees_method_fun)
- then have s': "(h\<^sub>1, l\<^sub>1, sh\<^sub>1) = (h\<^sub>2, l\<^sub>2, sh\<^sub>2)" using init_ProcessingE[OF _ SCallInit.hyps(6)] by simp
- have iconf3: "iconf (shp (h\<^sub>2, l\<^sub>2', sh\<^sub>2)) body"
- by(rule nsub_RI_iconf[OF sees_wwf_nsub_RI[OF wwf SCallInit.hyps(3)]])
- have "P,shp (h\<^sub>2, l\<^sub>2', sh\<^sub>2) \<turnstile>\<^sub>b (body,False) \<surd>" by(simp add: bconf_def)
- then have b3: "P \<turnstile> \<langle>body,(h\<^sub>2, l\<^sub>2', sh\<^sub>2),False\<rangle> \<rightarrow>* \<langle>e',(h\<^sub>3, l\<^sub>3, sh\<^sub>3),False\<rangle>"
- using SCallInit.hyps(11)[OF iconf3] by simp
- then have b3': "P \<turnstile> \<langle>body,(h\<^sub>1, l\<^sub>2', sh\<^sub>1),False\<rangle> \<rightarrow>* \<langle>e',(h\<^sub>3, l\<^sub>3, sh\<^sub>3),False\<rangle>"
- using s' by simp
- then show ?thesis using SCallInitProcessingReds[OF wwf b1 SCallInit.hyps(3) shC
- SCallInit.hyps(8-9) b3' eval_final[OF SCallInit.hyps(10)]] s' by simp
- next
- case False
- then have b1: "P \<turnstile> \<langle>ps,s\<^sub>0,b\<rangle> [\<rightarrow>]* \<langle>map Val vs,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\<rangle>" using ps vs by simp
- have bconf2: "P,shp (h\<^sub>1, l\<^sub>1, sh\<^sub>1) \<turnstile>\<^sub>b (INIT D ([D],False) \<leftarrow> unit,False) \<surd>" by(simp add: bconf_def)
- then have b2: "P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\<rangle> \<rightarrow>* \<langle>Val v',(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\<rangle>"
- using SCallInit.hyps(7) by auto
- have iconf3: "iconf (shp (h\<^sub>2, l\<^sub>2', sh\<^sub>2)) body"
- by(rule nsub_RI_iconf[OF sees_wwf_nsub_RI[OF wwf SCallInit.hyps(3)]])
- have "P,shp (h\<^sub>2, l\<^sub>2', sh\<^sub>2) \<turnstile>\<^sub>b (body,False) \<surd>" by(simp add: bconf_def)
- then have b3: "P \<turnstile> \<langle>body,(h\<^sub>2, l\<^sub>2', sh\<^sub>2),False\<rangle> \<rightarrow>* \<langle>e',(h\<^sub>3, l\<^sub>3, sh\<^sub>3),False\<rangle>"
- using SCallInit.hyps(11)[OF iconf3] by simp
- show ?thesis by(rule SCallInitReds[OF wwf b1 SCallInit.hyps(3-5) b2 SCallInit.hyps(8-9)
- b3 eval_final[OF SCallInit.hyps(10)]])
- qed
- qed
-next
- case (SCall ps s\<^sub>0 vs h\<^sub>2 l\<^sub>2 sh\<^sub>2 C M Ts T pns body D sfs l\<^sub>2' e' h\<^sub>3 l\<^sub>3 sh\<^sub>3) show ?case
- proof(cases "map_vals_of ps")
- case None
- then have iconf: "iconfs (shp s\<^sub>0) ps" using SCall.prems(1) by simp
- have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (ps,b) \<surd>" using SCall.prems(2) None by simp
- then have b1: "P \<turnstile> \<langle>ps,s\<^sub>0,b\<rangle> [\<rightarrow>]* \<langle>map Val vs,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\<rangle>"
- using SCall.hyps(2)[OF iconf] by auto
- have iconf3: "iconf (shp (h\<^sub>2, l\<^sub>2', sh\<^sub>2)) body"
- by(rule nsub_RI_iconf[OF sees_wwf_nsub_RI[OF wwf SCall.hyps(3)]])
- have "P,shp (h\<^sub>2, l\<^sub>2', sh\<^sub>2) \<turnstile>\<^sub>b (body,False) \<surd>" by(simp add: bconf_def)
- then have b2: "P \<turnstile> \<langle>body,(h\<^sub>2, l\<^sub>2', sh\<^sub>2),False\<rangle> \<rightarrow>* \<langle>e',(h\<^sub>3, l\<^sub>3, sh\<^sub>3),False\<rangle>"
- using SCall.hyps(8)[OF iconf3] by simp
- show ?thesis by(rule SCallRedsFinal[OF wwf b1 SCall.hyps(3-6) b2 eval_final[OF SCall.hyps(7)]])
- next
- case (Some vs')
- then have ps: "ps = map Val vs'" by(rule map_vals_of_spec)
- then have vs: "vs = vs' \<and> s\<^sub>0 = (h\<^sub>2, l\<^sub>2, sh\<^sub>2)"
- using evals_finals_same[OF _ SCall.hyps(1)] map_Val_eq by auto
- then have b1: "P \<turnstile> \<langle>ps,s\<^sub>0,b\<rangle> [\<rightarrow>]* \<langle>map Val vs,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),b\<rangle>" using ps by simp
- have iconf3: "iconf (shp (h\<^sub>2, l\<^sub>2', sh\<^sub>2)) body"
- by(rule nsub_RI_iconf[OF sees_wwf_nsub_RI[OF wwf SCall.hyps(3)]])
- have "P,shp (h\<^sub>2, l\<^sub>2', sh\<^sub>2) \<turnstile>\<^sub>b (body,False) \<surd>" by(simp add: bconf_def)
- then have b2: "P \<turnstile> \<langle>body,(h\<^sub>2, l\<^sub>2', sh\<^sub>2),False\<rangle> \<rightarrow>* \<langle>e',(h\<^sub>3, l\<^sub>3, sh\<^sub>3),False\<rangle>"
- using SCall.hyps(8)[OF iconf3] by simp
- show ?thesis by(rule SCallRedsFinal[OF wwf b1 SCall.hyps(3-6) b2 eval_final[OF SCall.hyps(7)]])
- qed
-next
- case (Block e\<^sub>0 h\<^sub>0 l\<^sub>0 V sh\<^sub>0 e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 T)
- have iconf: "iconf (shp (h\<^sub>0, l\<^sub>0(V := None), sh\<^sub>0)) e\<^sub>0"
- using Block.prems(1) by (auto simp: assigned_def)
- have bconf: "P,shp (h\<^sub>0, l\<^sub>0(V := None), sh\<^sub>0) \<turnstile>\<^sub>b (e\<^sub>0,b) \<surd>" using Block.prems(2)
- by(auto simp: bconf_def)
- then have b': "P \<turnstile> \<langle>e\<^sub>0,(h\<^sub>0, l\<^sub>0(V := None), sh\<^sub>0),b\<rangle> \<rightarrow>* \<langle>e\<^sub>1,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\<rangle>"
- using Block.hyps(2)[OF iconf] by auto
- have fin: "final e\<^sub>1" using Block by(auto dest: eval_final)
- thus ?case using BlockRedsFinal[OF b' fin] by simp
-next
- case (Seq e\<^sub>0 s\<^sub>0 v s\<^sub>1 e\<^sub>1 e\<^sub>2 s\<^sub>2)
- then have iconf: "iconf (shp s\<^sub>0) e\<^sub>0" using Seq.prems(1)
- by(auto dest: val_of_spec lass_val_of_spec)
- have b1: "\<exists>b1. P \<turnstile> \<langle>e\<^sub>0,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Val v,s\<^sub>1,b1\<rangle>"
- proof(cases "val_of e\<^sub>0")
- case None show ?thesis
- proof(cases "lass_val_of e\<^sub>0")
- case lNone:None
- then have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e\<^sub>0,b) \<surd>" using Seq.prems(2) None by simp
- then have "P \<turnstile> \<langle>e\<^sub>0,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Val v,s\<^sub>1,False\<rangle>" using iconf Seq.hyps(2) by auto
- then show ?thesis by fast
- next
- case (Some p)
- obtain V' v' where p: "p = (V',v')" and e\<^sub>0: "e\<^sub>0 = V':=Val v'"
- using lass_val_of_spec[OF Some] by(cases p, auto)
- obtain h l sh h' l' sh' where s\<^sub>0: "s\<^sub>0 = (h,l,sh)" and s\<^sub>1: "s\<^sub>1 = (h',l',sh')" by(cases s\<^sub>0, cases s\<^sub>1)
- then have eval: "P \<turnstile> \<langle>e\<^sub>0,(h,l,sh)\<rangle> \<Rightarrow> \<langle>Val v,(h',l',sh')\<rangle>" using Seq.hyps(1) by simp
- then have s\<^sub>1': "Val v = unit \<and> h' = h \<and> l' = l(V' \<mapsto> v') \<and> sh' = sh"
- using lass_val_of_eval[OF Some eval] p e\<^sub>0 by simp
- then have "P \<turnstile> \<langle>e\<^sub>0,s\<^sub>0,b\<rangle> \<rightarrow> \<langle>Val v,s\<^sub>1,b\<rangle>" using e\<^sub>0 s\<^sub>0 s\<^sub>1 by(auto intro: RedLAss)
- then show ?thesis by auto
- qed
- next
- case (Some a)
- then have "e\<^sub>0 = Val v" and "s\<^sub>0 = s\<^sub>1" using Seq.hyps(1) eval_cases(3) val_of_spec by blast+
- then show ?thesis using Seq by auto
- qed
- then obtain b1 where b1': "P \<turnstile> \<langle>e\<^sub>0,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Val v,s\<^sub>1,b1\<rangle>" by clarsimp
- have seq: "P \<turnstile> \<langle>e\<^sub>0;;e\<^sub>1,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Val v;;e\<^sub>1,s\<^sub>1,b1\<rangle>" by(rule SeqReds[OF b1'])
- then have iconf2: "iconf (shp s\<^sub>1) e\<^sub>1" using Red_preserves_iconf[OF wwf seq] Seq nsub_RI_iconf
- by auto
- have "P,shp s\<^sub>1 \<turnstile>\<^sub>b (Val v;; e\<^sub>1,b1) \<surd>" by(rule Red_preserves_bconf[OF wwf seq Seq.prems])
- then have bconf2: "P,shp s\<^sub>1 \<turnstile>\<^sub>b (e\<^sub>1,b1) \<surd>" by simp
- have b2: "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>1,b1\<rangle> \<rightarrow>* \<langle>e\<^sub>2,s\<^sub>2,False\<rangle>" by(rule Seq.hyps(4)[OF iconf2 bconf2])
- then show ?case using SeqReds2[OF b1' b2] by fast
-next
- case (SeqThrow e\<^sub>0 s\<^sub>0 a s\<^sub>1 e\<^sub>1 b)
- have notVal: "val_of e\<^sub>0 = None" "lass_val_of e\<^sub>0 = None"
- using SeqThrow.hyps(1) eval_throw_nonVal eval_throw_nonLAss by auto
- thus ?case using SeqThrow notVal by(auto dest!:eval_final dest: SeqRedsThrow)
-next
- case (CondT e s\<^sub>0 s\<^sub>1 e\<^sub>1 e' s\<^sub>2 e\<^sub>2)
- then have iconf: "iconf (shp s\<^sub>0) e" using CondT.prems(1) by auto
- have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e,b) \<surd>" using CondT.prems(2) by auto
- then have b1: "P \<turnstile> \<langle>e,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>true,s\<^sub>1,False\<rangle>" using iconf CondT.hyps(2) by auto
- have cond: "P \<turnstile> \<langle>if (e) e\<^sub>1 else e\<^sub>2,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>if (true) e\<^sub>1 else e\<^sub>2,s\<^sub>1,False\<rangle>" by(rule CondReds[OF b1])
- then have iconf2': "iconf (shp s\<^sub>1) e\<^sub>1" using Red_preserves_iconf[OF wwf cond] CondT nsub_RI_iconf
- by auto
- have "P,shp s\<^sub>1 \<turnstile>\<^sub>b (e\<^sub>1,False) \<surd>" by(simp add: bconf_def)
- then have b2: "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>1,False\<rangle> \<rightarrow>* \<langle>e',s\<^sub>2,False\<rangle>" using CondT.hyps(4)[OF iconf2'] by auto
- then show ?case using CondReds2T[OF b1 b2] by fast
-next
- case (CondF e s\<^sub>0 s\<^sub>1 e\<^sub>2 e' s\<^sub>2 e\<^sub>1)
- then have iconf: "iconf (shp s\<^sub>0) e" using CondF.prems(1) by auto
- have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e,b) \<surd>" using CondF.prems(2) by auto
- then have b1: "P \<turnstile> \<langle>e,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>false,s\<^sub>1,False\<rangle>" using iconf CondF.hyps(2) by auto
- have cond: "P \<turnstile> \<langle>if (e) e\<^sub>1 else e\<^sub>2,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>if (false) e\<^sub>1 else e\<^sub>2,s\<^sub>1,False\<rangle>" by(rule CondReds[OF b1])
- then have iconf2': "iconf (shp s\<^sub>1) e\<^sub>2" using Red_preserves_iconf[OF wwf cond] CondF nsub_RI_iconf
- by auto
- have "P,shp s\<^sub>1 \<turnstile>\<^sub>b (e\<^sub>2,False) \<surd>" by(simp add: bconf_def)
- then have b2: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>1,False\<rangle> \<rightarrow>* \<langle>e',s\<^sub>2,False\<rangle>" using CondF.hyps(4)[OF iconf2'] by auto
- then show ?case using CondReds2F[OF b1 b2] by fast
-next
- case CondThrow thus ?case by(auto dest!:eval_final dest:CondRedsThrow)
-next
- case (WhileF e s\<^sub>0 s\<^sub>1 c)
- then have iconf: "iconf (shp s\<^sub>0) e" using nsub_RI_iconf by auto
- then have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e,b) \<surd>" using WhileF.prems(2) by(simp add: bconf_def)
- then have b': "P \<turnstile> \<langle>e,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>false,s\<^sub>1,False\<rangle>" using WhileF.hyps(2) iconf by auto
- thus ?case using WhileFReds[OF b'] by fast
-next
- case (WhileT e s\<^sub>0 s\<^sub>1 c v\<^sub>1 s\<^sub>2 e\<^sub>3 s\<^sub>3)
- then have iconf: "iconf (shp s\<^sub>0) e" using nsub_RI_iconf by auto
- then have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e,b) \<surd>" using WhileT.prems(2) by(simp add: bconf_def)
- then have b1: "P \<turnstile> \<langle>e,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>true,s\<^sub>1,False\<rangle>" using WhileT.hyps(2) iconf by auto
- have iconf2: "iconf (shp s\<^sub>1) c" using WhileT.prems(1) nsub_RI_iconf by auto
- have bconf2: "P,shp s\<^sub>1 \<turnstile>\<^sub>b (c,False) \<surd>" by(simp add: bconf_def)
- then have b2: "P \<turnstile> \<langle>c,s\<^sub>1,False\<rangle> \<rightarrow>* \<langle>Val v\<^sub>1,s\<^sub>2,False\<rangle>" using WhileT.hyps(4) iconf2 by auto
- have iconf3: "iconf (shp s\<^sub>2) (while (e) c)" using WhileT.prems(1) by auto
- have "P,shp s\<^sub>2 \<turnstile>\<^sub>b (while (e) c,False) \<surd>" by(simp add: bconf_def)
- then have b3: "P \<turnstile> \<langle>while (e) c,s\<^sub>2,False\<rangle> \<rightarrow>* \<langle>e\<^sub>3,s\<^sub>3,False\<rangle>" using WhileT.hyps(6) iconf3 by auto
- show ?case using WhileTReds[OF b1 b2 b3] by fast
-next
- case WhileCondThrow thus ?case
- by (metis (no_types, lifting) WhileRedsThrow iconf.simps(16) bconf_While bconf_def nsub_RI_iconf)
-next
- case (WhileBodyThrow e s\<^sub>0 s\<^sub>1 c e' s\<^sub>2)
- then have iconf: "iconf (shp s\<^sub>0) e" using nsub_RI_iconf by auto
- then have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e,b) \<surd>" using WhileBodyThrow.prems(2) by(simp add: bconf_def)
- then have b1: "P \<turnstile> \<langle>e,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>true,s\<^sub>1,False\<rangle>" using WhileBodyThrow.hyps(2) iconf by auto
- have iconf2: "iconf (shp s\<^sub>1) c" using WhileBodyThrow.prems(1) nsub_RI_iconf by auto
- have bconf2: "P,shp s\<^sub>1 \<turnstile>\<^sub>b (c,False) \<surd>" by(simp add: bconf_def)
- then have b2: "P \<turnstile> \<langle>c,s\<^sub>1,False\<rangle> \<rightarrow>* \<langle>throw e',s\<^sub>2,False\<rangle>" using WhileBodyThrow.hyps(4) iconf2 by auto
- show ?case using WhileTRedsThrow[OF b1 b2] by fast
-next
- case Throw thus ?case by (meson ThrowReds iconf.simps(17) bconf_Throw)
-next
- case ThrowNull thus ?case by (meson ThrowRedsNull iconf.simps(17) bconf_Throw)
-next
- case ThrowThrow thus ?case by (meson ThrowRedsThrow iconf.simps(17) bconf_Throw)
-next
- case Try thus ?case by (meson TryRedsVal iconf.simps(18) bconf_Try)
-next
- case (TryCatch e\<^sub>1 s\<^sub>0 a h\<^sub>1 l\<^sub>1 sh\<^sub>1 D fs C e\<^sub>2 V e\<^sub>2' h\<^sub>2 l\<^sub>2 sh\<^sub>2)
- then have b1: "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Throw a,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\<rangle>" by auto
- have Try: "P \<turnstile> \<langle>try e\<^sub>1 catch(C V) e\<^sub>2,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>try (Throw a) catch(C V) e\<^sub>2,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\<rangle>"
- by(rule TryReds[OF b1])
- have iconf: "iconf sh\<^sub>1 e\<^sub>2" using Red_preserves_iconf[OF wwf Try] TryCatch nsub_RI_iconf
- by auto
- have bconf: "P,shp (h\<^sub>1, l\<^sub>1(V \<mapsto> Addr a), sh\<^sub>1) \<turnstile>\<^sub>b (e\<^sub>2,False) \<surd>" by(simp add: bconf_def)
- then have b2: "P \<turnstile> \<langle>e\<^sub>2,(h\<^sub>1, l\<^sub>1(V \<mapsto> Addr a), sh\<^sub>1),False\<rangle> \<rightarrow>* \<langle>e\<^sub>2',(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\<rangle>"
- using TryCatch.hyps(6) iconf by auto
- thus ?case using TryCatchRedsFinal[OF b1] TryCatch.hyps(3-5) by (meson eval_final)
-next
- case TryThrow thus ?case by (meson TryRedsFail iconf.simps(18) bconf_Try)
-next
- case Nil thus ?case by(auto simp: bconfs_def)
-next
- case (Cons e s\<^sub>0 v s\<^sub>1 es es' s\<^sub>2) show ?case
- proof(cases "val_of e")
- case None
- then have iconf: "iconf (shp s\<^sub>0) e" using Cons.prems(1) by simp
- have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e,b) \<surd>" using Cons.prems(2) None by simp
- then have b1: "P \<turnstile> \<langle>e,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Val v,s\<^sub>1,False\<rangle>" using Cons.hyps(2) iconf by auto
- have cons: "P \<turnstile> \<langle>e # es,s\<^sub>0,b\<rangle> [\<rightarrow>]* \<langle>Val v # es,s\<^sub>1,False\<rangle>" by(rule ListReds1[OF b1])
- then have iconf2': "iconfs (shp s\<^sub>1) es" using Reds_preserves_iconf[OF wwf cons] None Cons by auto
- have "P,shp s\<^sub>1 \<turnstile>\<^sub>b (es,False) \<surd>" by(simp add: bconfs_def)
- then have b2: "P \<turnstile> \<langle>es,s\<^sub>1,False\<rangle> [\<rightarrow>]* \<langle>es',s\<^sub>2,False\<rangle>" using Cons.hyps(4)[OF iconf2'] by auto
- show ?thesis using ListRedsVal[OF b1 b2] by auto
- next
- case (Some a)
- then obtain b1 where b1: "P \<turnstile> \<langle>e,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Val v,s\<^sub>1,b1\<rangle>"
- by (metis (no_types, lifting) Cons.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec)
- have cons: "P \<turnstile> \<langle>e # es,s\<^sub>0,b\<rangle> [\<rightarrow>]* \<langle>Val v # es,s\<^sub>1,b1\<rangle>" by(rule ListReds1[OF b1])
- then have iconf2': "iconfs (shp s\<^sub>1) es" using Reds_preserves_iconf[OF wwf cons] Cons by auto
- have bconf2: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (es,b) \<surd>" using Cons.prems Some by simp
- then have "P,shp s\<^sub>1 \<turnstile>\<^sub>b (es,b1) \<surd>" using Reds_preserves_bconf[OF wwf cons Cons.prems] by simp
- then have b2: "P \<turnstile> \<langle>es,s\<^sub>1,b1\<rangle> [\<rightarrow>]* \<langle>es',s\<^sub>2,False\<rangle>" using Cons.hyps(4)[OF iconf2'] by auto
- show ?thesis using ListRedsVal[OF b1 b2] by auto
- qed
-next
- case (ConsThrow e s\<^sub>0 e' s\<^sub>1 es) show ?case
- proof(cases "val_of e")
- case None
- then have iconf: "iconf (shp s\<^sub>0) e" using ConsThrow.prems(1) by simp
- have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e,b) \<surd>" using ConsThrow.prems(2) None by simp
- then have b1: "P \<turnstile> \<langle>e,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>throw e',s\<^sub>1,False\<rangle>" using ConsThrow.hyps(2) iconf by auto
- have cons: "P \<turnstile> \<langle>e # es,s\<^sub>0,b\<rangle> [\<rightarrow>]* \<langle>throw e' # es,s\<^sub>1,False\<rangle>" by(rule ListReds1[OF b1])
- then show ?thesis by fast
- next
- case (Some a)
- then show ?thesis using eval_final_same[OF ConsThrow.hyps(1)] val_of_spec[OF Some] by auto
- qed
-next
- case (InitFinal e s e' s' C b')
- then have "\<not>sub_RI e" by simp
- then show ?case using InitFinal RedInit[of e C b' s b P]
- by (meson converse_rtrancl_into_rtrancl nsub_RI_iconf red_preserves_bconf RedInit)
-next
- case (InitNone sh C C' Cs e h l e' s')
- then have init: "P \<turnstile> \<langle>INIT C' (C # Cs,False) \<leftarrow> e,(h, l, sh(C \<mapsto> (sblank P C, Prepared))),b\<rangle> \<rightarrow>* \<langle>e',s',False\<rangle>"
- by(simp add: bconf_def)
- show ?case by(rule InitNoneReds[OF InitNone.hyps(1) init])
-next
- case (InitDone sh C sfs C' Cs e h l e' s')
- then have "iconf (shp (h, l, sh)) (INIT C' (Cs,True) \<leftarrow> e)" using InitDone.hyps(1)
- proof(cases Cs)
- case Nil
- then have "C = C'" "\<not>sub_RI e" using InitDone.prems(1) by simp+
- then show ?thesis using Nil InitDone.hyps(1) by(simp add: initPD_def)
- qed(auto)
- then have init: "P \<turnstile> \<langle>INIT C' (Cs,True) \<leftarrow> e,(h, l, sh),b\<rangle> \<rightarrow>* \<langle>e',s',False\<rangle>"
- using InitDone by(simp add: bconf_def)
- show ?case by(rule InitDoneReds[OF InitDone.hyps(1) init])
-next
- case (InitProcessing sh C sfs C' Cs e h l e' s')
- then have "iconf (shp (h, l, sh)) (INIT C' (Cs,True) \<leftarrow> e)" using InitProcessing.hyps(1)
- proof(cases Cs)
- case Nil
- then have "C = C'" "\<not>sub_RI e" using InitProcessing.prems(1) by simp+
- then show ?thesis using Nil InitProcessing.hyps(1) by(simp add: initPD_def)
- qed(auto)
- then have init: "P \<turnstile> \<langle>INIT C' (Cs,True) \<leftarrow> e,(h, l, sh),b\<rangle> \<rightarrow>* \<langle>e',s',False\<rangle>"
- using InitProcessing by(simp add: bconf_def)
- show ?case by(rule InitProcessingReds[OF InitProcessing.hyps(1) init])
-next
- case InitError thus ?case by(fastforce intro: InitErrorReds simp: bconf_def)
-next
- case InitObject thus ?case by(fastforce intro: InitObjectReds simp: bconf_def)
-next
- case InitNonObject thus ?case by(fastforce intro: InitNonObjectReds simp: bconf_def)
-next
- case InitRInit thus ?case by(fastforce intro: RedsInitRInit simp: bconf_def)
-next
- case (RInit e s v h' l' sh' C sfs i sh'' C' Cs e' e\<^sub>1 s\<^sub>1)
- then have iconf2: "iconf (shp (h', l', sh'')) (INIT C' (Cs,True) \<leftarrow> e')"
- by(auto simp: initPD_def fun_upd_same list_nonempty_induct)
- show ?case
- proof(cases "val_of e")
- case None
- then have iconf: "iconf (shp s) e" using RInit.prems(1) by simp
- have bconf: "P,shp s \<turnstile>\<^sub>b (e,b) \<surd>" using RInit.prems(2) None by simp
- then have b1: "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>Val v,(h',l',sh'),False\<rangle>" using RInit.hyps(2)[OF iconf] by auto
- have "P,shp (h', l', sh'') \<turnstile>\<^sub>b (INIT C' (Cs,True) \<leftarrow> e',False) \<surd>" by(simp add: bconf_def)
- then have b2: "P \<turnstile> \<langle>INIT C' (Cs,True) \<leftarrow> e',(h',l',sh''),False\<rangle> \<rightarrow>* \<langle>e\<^sub>1,s\<^sub>1,False\<rangle>"
- using RInit.hyps(7)[OF iconf2] by auto
- then show ?thesis using RedsRInit[OF b1 RInit.hyps(3-5) b2] by fast
- next
- case (Some a')
- then obtain b1 where b1: "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>Val v,(h',l',sh'),b1\<rangle>"
- by (metis (no_types, lifting) RInit.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec)
- have fin: "final e" by(simp add: val_of_spec[OF Some])
- have "\<not>b" using RInit.prems(2) Some by(simp add: bconf_def)
- then have nb1: "\<not>b1" using reds_final_same[OF b1 fin] by simp
- have "P,shp (h', l', sh'') \<turnstile>\<^sub>b (INIT C' (Cs,True) \<leftarrow> e',b1) \<surd>" using nb1
- by(simp add: bconf_def)
- then have b2: "P \<turnstile> \<langle>INIT C' (Cs,True) \<leftarrow> e',(h', l', sh''),b1\<rangle> \<rightarrow>* \<langle>e\<^sub>1,s\<^sub>1,False\<rangle>"
- using RInit.hyps(7)[OF iconf2] by auto
- then show ?thesis using RedsRInit[OF b1 RInit.hyps(3-5) b2] by fast
- qed
-next
- case (RInitInitFail e s a h' l' sh' C sfs i sh'' D Cs e' e\<^sub>1 s\<^sub>1)
- have fin: "final (throw a)" using eval_final[OF RInitInitFail.hyps(1)] by simp
- then obtain a' where a': "throw a = Throw a'" by auto
- have iconf2: "iconf (shp (h', l', sh'')) (RI (D,Throw a') ; Cs \<leftarrow> e')"
- using RInitInitFail.prems(1) by auto
- show ?case
- proof(cases "val_of e")
- case None
- then have iconf: "iconf (shp s) e" using RInitInitFail.prems(1) by simp
- have bconf: "P,shp s \<turnstile>\<^sub>b (e,b) \<surd>" using RInitInitFail.prems(2) None by simp
- then have b1: "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>Throw a',(h',l',sh'),False\<rangle>"
- using RInitInitFail.hyps(2)[OF iconf] a' by auto
- have "P,shp (h', l', sh'') \<turnstile>\<^sub>b (RI (D,Throw a') ; Cs \<leftarrow> e',False) \<surd>" by(simp add: bconf_def)
- then have b2: "P \<turnstile> \<langle>RI (D,Throw a') ; Cs \<leftarrow> e',(h',l',sh''),False\<rangle> \<rightarrow>* \<langle>e\<^sub>1,s\<^sub>1,False\<rangle>"
- using RInitInitFail.hyps(6) iconf2 a' by auto
- show ?thesis using RInitInitThrowReds[OF b1 RInitInitFail.hyps(3-4) b2] by fast
- next
- case (Some a1)
- then obtain b1 where b1: "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>Throw a',(h',l',sh'),b1\<rangle>" using a'
- by (metis (no_types, lifting) RInitInitFail.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec)
- have fin: "final e" by(simp add: val_of_spec[OF Some])
- have "\<not>b" using RInitInitFail.prems(2) Some by(simp add: bconf_def)
- then have nb1: "\<not>b1" using reds_final_same[OF b1 fin] by simp
- have "P,shp (h', l', sh'') \<turnstile>\<^sub>b (RI (D,Throw a') ; Cs \<leftarrow> e',b1) \<surd>" using nb1
- by(simp add: bconf_def)
- then have b2: "P \<turnstile> \<langle>RI (D,Throw a') ; Cs \<leftarrow> e',(h', l', sh''),b1\<rangle> \<rightarrow>* \<langle>e\<^sub>1,s\<^sub>1,False\<rangle>"
- using RInitInitFail.hyps(6) iconf2 a' by auto
- show ?thesis using RInitInitThrowReds[OF b1 RInitInitFail.hyps(3-4) b2] by fast
- qed
-next
- case (RInitFailFinal e s a h' l' sh' C sfs i sh'' e')
- have fin: "final (throw a)" using eval_final[OF RInitFailFinal.hyps(1)] by simp
- then obtain a' where a': "throw a = Throw a'" by auto
- show ?case
- proof(cases "val_of e")
- case None
- then have iconf: "iconf (shp s) e" using RInitFailFinal.prems(1) by simp
- have bconf: "P,shp s \<turnstile>\<^sub>b (e,b) \<surd>" using RInitFailFinal.prems(2) None by simp
- then have b1: "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>Throw a',(h',l',sh'),False\<rangle>"
- using RInitFailFinal.hyps(2)[OF iconf] a' by auto
- show ?thesis using RInitThrowReds[OF b1 RInitFailFinal.hyps(3-4)] a' by fast
- next
- case (Some a1)
- then show ?thesis using eval_final_same[OF RInitFailFinal.hyps(1)] val_of_spec[OF Some] by auto
- qed
-qed
-(*>*)
-
-
-subsection\<open>Big steps simulates small step\<close>
-
-text\<open> This direction was carried out by Norbert Schirmer and Daniel
-Wasserrab (and modified to include statics and DCI by Susannah Mansky). \<close>
-
-text \<open> The big step equivalent of @{text RedWhile}: \<close>
-
-lemma unfold_while:
- "P \<turnstile> \<langle>while(b) c,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle> = P \<turnstile> \<langle>if(b) (c;;while(b) c) else (unit),s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
-(*<*)
-proof
- assume "P \<turnstile> \<langle>while (b) c,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
- thus "P \<turnstile> \<langle>if (b) (c;; while (b) c) else unit,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
- by cases (fastforce intro: eval_evals.intros)+
-next
- assume "P \<turnstile> \<langle>if (b) (c;; while (b) c) else unit,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
- thus "P \<turnstile> \<langle>while (b) c,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
- proof (cases)
- fix a
- assume e': "e' = throw a"
- assume "P \<turnstile> \<langle>b,s\<rangle> \<Rightarrow> \<langle>throw a,s'\<rangle>"
- hence "P \<turnstile> \<langle>while(b) c,s\<rangle> \<Rightarrow> \<langle>throw a,s'\<rangle>" by (rule WhileCondThrow)
- with e' show ?thesis by simp
- next
- fix s\<^sub>1
- assume eval_false: "P \<turnstile> \<langle>b,s\<rangle> \<Rightarrow> \<langle>false,s\<^sub>1\<rangle>"
- and eval_unit: "P \<turnstile> \<langle>unit,s\<^sub>1\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
- with eval_unit have "s' = s\<^sub>1" "e' = unit" by (auto elim: eval_cases)
- moreover from eval_false have "P \<turnstile> \<langle>while (b) c,s\<rangle> \<Rightarrow> \<langle>unit,s\<^sub>1\<rangle>"
- by - (rule WhileF, simp)
- ultimately show ?thesis by simp
- next
- fix s\<^sub>1
- assume eval_true: "P \<turnstile> \<langle>b,s\<rangle> \<Rightarrow> \<langle>true,s\<^sub>1\<rangle>"
- and eval_rest: "P \<turnstile> \<langle>c;; while (b) c,s\<^sub>1\<rangle>\<Rightarrow>\<langle>e',s'\<rangle>"
- from eval_rest show ?thesis
- proof (cases)
- fix s\<^sub>2 v\<^sub>1
- assume "P \<turnstile> \<langle>c,s\<^sub>1\<rangle> \<Rightarrow> \<langle>Val v\<^sub>1,s\<^sub>2\<rangle>" "P \<turnstile> \<langle>while (b) c,s\<^sub>2\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
- with eval_true show "P \<turnstile> \<langle>while(b) c,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>" by (rule WhileT)
- next
- fix a
- assume "P \<turnstile> \<langle>c,s\<^sub>1\<rangle> \<Rightarrow> \<langle>throw a,s'\<rangle>" "e' = throw a"
- with eval_true show "P \<turnstile> \<langle>while(b) c,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
- by (iprover intro: WhileBodyThrow)
- qed
- qed
-qed
-(*>*)
-
-
-lemma blocksEval:
- "\<And>Ts vs l l'. \<lbrakk>size ps = size Ts; size ps = size vs; P \<turnstile> \<langle>blocks(ps,Ts,vs,e),(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',(h',l',sh')\<rangle> \<rbrakk>
- \<Longrightarrow> \<exists> l''. P \<turnstile> \<langle>e,(h,l(ps[\<mapsto>]vs),sh)\<rangle> \<Rightarrow> \<langle>e',(h',l'',sh')\<rangle>"
-(*<*)
-proof (induct ps)
- case Nil then show ?case by fastforce
-next
- case (Cons p ps')
- have length_eqs: "length (p # ps') = length Ts"
- "length (p # ps') = length vs" by fact+
- then obtain T Ts' where Ts: "Ts = T#Ts'" by (cases "Ts") simp
- obtain v vs' where vs: "vs = v#vs'" using length_eqs by (cases "vs") simp
- have "P \<turnstile> \<langle>blocks (p # ps', Ts, vs, e),(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',(h', l',sh')\<rangle>" by fact
- with Ts vs
- have "P \<turnstile> \<langle>{p:T := Val v; blocks (ps', Ts', vs', e)},(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',(h', l',sh')\<rangle>"
- by simp
- then obtain l''' where
- eval_ps': "P \<turnstile> \<langle>blocks (ps', Ts', vs', e),(h, l(p\<mapsto>v), sh)\<rangle> \<Rightarrow> \<langle>e',(h', l''', sh')\<rangle>"
- and l''': "l'=l'''(p:=l p)"
- by (auto elim!: eval_cases)
- then obtain l'' where
- hyp: "P \<turnstile> \<langle>e,(h, l(p\<mapsto>v)(ps'[\<mapsto>]vs'), sh)\<rangle> \<Rightarrow> \<langle>e',(h', l'', sh')\<rangle>"
- using length_eqs Ts vs Cons.hyps [OF _ _ eval_ps'] by auto
- from hyp
- show "\<exists>l''. P \<turnstile> \<langle>e,(h, l(p # ps'[\<mapsto>]vs), sh)\<rangle> \<Rightarrow> \<langle>e',(h', l'', sh')\<rangle>"
- using Ts vs by auto
-qed
-(*>*)
-
-lemma
-assumes wf: "wwf_J_prog P"
-shows eval_restrict_lcl:
- "P \<turnstile> \<langle>e,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',(h',l',sh')\<rangle> \<Longrightarrow> (\<And>W. fv e \<subseteq> W \<Longrightarrow> P \<turnstile> \<langle>e,(h,l|`W,sh)\<rangle> \<Rightarrow> \<langle>e',(h',l'|`W,sh')\<rangle>)"
-and "P \<turnstile> \<langle>es,(h,l,sh)\<rangle> [\<Rightarrow>] \<langle>es',(h',l',sh')\<rangle> \<Longrightarrow> (\<And>W. fvs es \<subseteq> W \<Longrightarrow> P \<turnstile> \<langle>es,(h,l|`W,sh)\<rangle> [\<Rightarrow>] \<langle>es',(h',l'|`W,sh')\<rangle>)"
-(*<*)
-proof(induct rule:eval_evals_inducts)
- case (Block e\<^sub>0 h\<^sub>0 l\<^sub>0 V sh\<^sub>0 e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 T)
- have IH: "\<And>W. fv e\<^sub>0 \<subseteq> W \<Longrightarrow> P \<turnstile> \<langle>e\<^sub>0,(h\<^sub>0,l\<^sub>0(V:=None)|`W,sh\<^sub>0)\<rangle> \<Rightarrow> \<langle>e\<^sub>1,(h\<^sub>1,l\<^sub>1|`W,sh\<^sub>1)\<rangle>" by fact
- have "fv({V:T; e\<^sub>0}) \<subseteq> W" by fact+
- hence "fv e\<^sub>0 - {V} \<subseteq> W" by simp_all
- hence "fv e\<^sub>0 \<subseteq> insert V W" by fast
- from IH[OF this]
- have "P \<turnstile> \<langle>e\<^sub>0,(h\<^sub>0, (l\<^sub>0|`W)(V := None), sh\<^sub>0)\<rangle> \<Rightarrow> \<langle>e\<^sub>1,(h\<^sub>1, l\<^sub>1|`insert V W, sh\<^sub>1)\<rangle>"
- by fastforce
- from eval_evals.Block[OF this] show ?case by fastforce
-next
- case Seq thus ?case by simp (blast intro:eval_evals.Seq)
-next
- case New thus ?case by(simp add:eval_evals.intros)
-next
- case NewFail thus ?case by(simp add:eval_evals.intros)
-next
- case (NewInit sh C h l v' h' l' sh' a h'')
- have "fv(INIT C ([C],False) \<leftarrow> unit) \<subseteq> W" by simp
- then have "P \<turnstile> \<langle>INIT C ([C],False) \<leftarrow> unit,(h, l |` W, sh)\<rangle> \<Rightarrow> \<langle>Val v',(h', l' |` W, sh')\<rangle>"
- by (simp add: NewInit.hyps(3))
- thus ?case using NewInit.hyps(1,4-6) eval_evals.NewInit by blast
-next
- case (NewInitOOM sh C h l v' h' l' sh')
- have "fv(INIT C ([C],False) \<leftarrow> unit) \<subseteq> W" by simp
- then have "P \<turnstile> \<langle>INIT C ([C],False) \<leftarrow> unit,(h, l |` W, sh)\<rangle> \<Rightarrow> \<langle>Val v',(h', l' |` W, sh')\<rangle>"
- by (simp add: NewInitOOM.hyps(3))
- thus ?case
- using NewInitOOM.hyps(1,4,5) eval_evals.NewInitOOM by auto
-next
- case NewInitThrow thus ?case by(simp add:eval_evals.intros)
-next
- case Cast thus ?case by simp (blast intro:eval_evals.Cast)
-next
- case CastNull thus ?case by simp (blast intro:eval_evals.CastNull)
-next
- case CastFail thus ?case by simp (blast intro:eval_evals.CastFail)
-next
- case CastThrow thus ?case by(simp add:eval_evals.intros)
-next
- case Val thus ?case by(simp add:eval_evals.intros)
-next
- case BinOp thus ?case by simp (blast intro:eval_evals.BinOp)
-next
- case BinOpThrow1 thus ?case by simp (blast intro:eval_evals.BinOpThrow1)
-next
- case BinOpThrow2 thus ?case by simp (blast intro:eval_evals.BinOpThrow2)
-next
- case Var thus ?case by(simp add:eval_evals.intros)
-next
- case (LAss e h\<^sub>0 l\<^sub>0 sh\<^sub>0 v h l sh l' V)
- have IH: "\<And>W. fv e \<subseteq> W \<Longrightarrow> P \<turnstile> \<langle>e,(h\<^sub>0,l\<^sub>0|`W,sh\<^sub>0)\<rangle> \<Rightarrow> \<langle>Val v,(h,l|`W,sh)\<rangle>"
- and [simp]: "l' = l(V \<mapsto> v)" by fact+
- have "fv (V:=e) \<subseteq> W" by fact
- hence fv: "fv e \<subseteq> W" and VinW: "V \<in> W" by auto
- from eval_evals.LAss[OF IH[OF fv] refl, of V] VinW
- show ?case by simp
-next
- case LAssThrow thus ?case by(fastforce intro: eval_evals.LAssThrow)
-next
- case FAcc thus ?case by simp (blast intro: eval_evals.FAcc)
-next
- case FAccNull thus ?case by(fastforce intro: eval_evals.FAccNull)
-next
- case FAccThrow thus ?case by(fastforce intro: eval_evals.FAccThrow)
-next
- case FAccNone thus ?case by(metis eval_evals.FAccNone fv.simps(7))
-next
- case FAccStatic thus ?case by(metis eval_evals.FAccStatic fv.simps(7))
-next
- case SFAcc thus ?case by simp (blast intro: eval_evals.SFAcc)
-next
- case SFAccInit thus ?case by simp (blast intro: eval_evals.SFAccInit)
-next
- case SFAccInitThrow thus ?case by simp (blast intro: eval_evals.SFAccInitThrow)
-next
- case SFAccNone thus ?case by simp (blast intro: eval_evals.SFAccNone)
-next
- case SFAccNonStatic thus ?case by simp (blast intro: eval_evals.SFAccNonStatic)
-next
- case FAss thus ?case by simp (blast intro: eval_evals.FAss)
-next
- case FAssNull thus ?case by simp (blast intro: eval_evals.FAssNull)
-next
- case FAssThrow1 thus ?case by simp (blast intro: eval_evals.FAssThrow1)
-next
- case FAssThrow2 thus ?case by simp (blast intro: eval_evals.FAssThrow2)
-next
- case FAssNone thus ?case by simp (blast intro: eval_evals.FAssNone)
-next
- case FAssStatic thus ?case by simp (blast intro: eval_evals.FAssStatic)
-next
- case SFAss thus ?case by simp (blast intro: eval_evals.SFAss)
-next
- case SFAssInit thus ?case by simp (blast intro: eval_evals.SFAssInit)
-next
- case SFAssInitThrow thus ?case by simp (blast intro: eval_evals.SFAssInitThrow)
-next
- case SFAssThrow thus ?case by simp (blast intro: eval_evals.SFAssThrow)
-next
- case SFAssNone thus ?case by simp (blast intro: eval_evals.SFAssNone)
-next
- case SFAssNonStatic thus ?case by simp (blast intro: eval_evals.SFAssNonStatic)
-next
- case CallObjThrow thus ?case by simp (blast intro: eval_evals.intros)
-next
- case CallNull thus ?case by simp (blast intro: eval_evals.CallNull)
-next
- case (CallNone e h l sh a h' l' sh' ps vs h\<^sub>2 l\<^sub>2 sh\<^sub>2 C fs M)
- have f1: "P \<turnstile> \<langle>e,(h, l |` W, sh)\<rangle> \<Rightarrow> \<langle>addr a,(h', l' |` W, sh')\<rangle>"
- by (metis (no_types) fv.simps(11) le_sup_iff local.CallNone(2) local.CallNone(7))
- have "P \<turnstile> \<langle>ps,(h', l' |` W, sh')\<rangle> [\<Rightarrow>] \<langle>map Val vs, (h\<^sub>2, l\<^sub>2 |` W, sh\<^sub>2)\<rangle>"
- using local.CallNone(4) local.CallNone(7) by auto
- then show ?case
- using f1 eval_evals.CallNone local.CallNone(5) local.CallNone(6) by auto
-next
- case CallStatic thus ?case
- by (metis (no_types, lifting) eval_evals.CallStatic fv.simps(11) le_sup_iff)
-next
- case CallParamsThrow thus ?case
- by simp (blast intro: eval_evals.CallParamsThrow)
-next
- case (Call e h\<^sub>0 l\<^sub>0 sh\<^sub>0 a h\<^sub>1 l\<^sub>1 sh\<^sub>1 ps vs h\<^sub>2 l\<^sub>2 sh\<^sub>2 C fs M Ts T pns body
- D l\<^sub>2' e' h\<^sub>3 l\<^sub>3 sh\<^sub>3)
- have IHe: "\<And>W. fv e \<subseteq> W \<Longrightarrow> P \<turnstile> \<langle>e,(h\<^sub>0,l\<^sub>0|`W,sh\<^sub>0)\<rangle> \<Rightarrow> \<langle>addr a,(h\<^sub>1,l\<^sub>1|`W,sh\<^sub>1)\<rangle>"
- and IHps: "\<And>W. fvs ps \<subseteq> W \<Longrightarrow> P \<turnstile> \<langle>ps,(h\<^sub>1,l\<^sub>1|`W,sh\<^sub>1)\<rangle> [\<Rightarrow>] \<langle>map Val vs,(h\<^sub>2,l\<^sub>2|`W,sh\<^sub>2)\<rangle>"
- and IHbd: "\<And>W. fv body \<subseteq> W \<Longrightarrow> P \<turnstile> \<langle>body,(h\<^sub>2,l\<^sub>2'|`W,sh\<^sub>2)\<rangle> \<Rightarrow> \<langle>e',(h\<^sub>3,l\<^sub>3|`W,sh\<^sub>3)\<rangle>"
- and h\<^sub>2a: "h\<^sub>2 a = Some (C, fs)"
- and "method": "P \<turnstile> C sees M,NonStatic: Ts\<rightarrow>T = (pns, body) in D"
- and same_len: "size vs = size pns"
- and l\<^sub>2': "l\<^sub>2' = [this \<mapsto> Addr a, pns [\<mapsto>] vs]" by fact+
- have "fv (e\<bullet>M(ps)) \<subseteq> W" by fact
- hence fve: "fv e \<subseteq> W" and fvps: "fvs(ps) \<subseteq> W" by auto
- have wfmethod: "size Ts = size pns \<and> this \<notin> set pns" and
- fvbd: "fv body \<subseteq> {this} \<union> set pns"
- using "method" wf by(fastforce dest!:sees_wf_mdecl simp:wf_mdecl_def)+
- show ?case
- using IHbd[OF fvbd] l\<^sub>2' same_len wfmethod h\<^sub>2a
- eval_evals.Call[OF IHe[OF fve] IHps[OF fvps] _ "method" same_len l\<^sub>2']
- by (simp add:subset_insertI)
-next
- case (SCallNone ps h l sh vs h' l' sh' C M)
- have "P \<turnstile> \<langle>ps,(h, l |` W, sh)\<rangle> [\<Rightarrow>] \<langle>map Val vs,(h', l' |` W, sh')\<rangle>"
- using SCallNone.hyps(2) SCallNone.prems by auto
- then show ?case using SCallNone.hyps(3) eval_evals.SCallNone by auto
-next
- case SCallNonStatic thus ?case by (metis eval_evals.SCallNonStatic fv.simps(12))
-next
- case SCallParamsThrow thus ?case
- by simp (blast intro: eval_evals.SCallParamsThrow)
-next
- case SCallInitThrow thus ?case by simp (blast intro: eval_evals.SCallInitThrow)
-next
- case SCallInit thus ?case by simp (blast intro: eval_evals.SCallInit)
-next
- case (SCall ps h\<^sub>0 l\<^sub>0 sh\<^sub>0 vs h\<^sub>2 l\<^sub>2 sh\<^sub>2 C M Ts T pns body D sfs l\<^sub>2' e' h\<^sub>3 l\<^sub>3 sh\<^sub>3)
- have IHps: "\<And>W. fvs ps \<subseteq> W \<Longrightarrow> P \<turnstile> \<langle>ps,(h\<^sub>0,l\<^sub>0|`W,sh\<^sub>0)\<rangle> [\<Rightarrow>] \<langle>map Val vs,(h\<^sub>2,l\<^sub>2|`W,sh\<^sub>2)\<rangle>"
- and IHbd: "\<And>W. fv body \<subseteq> W \<Longrightarrow> P \<turnstile> \<langle>body,(h\<^sub>2,l\<^sub>2'|`W,sh\<^sub>2)\<rangle> \<Rightarrow> \<langle>e',(h\<^sub>3,l\<^sub>3|`W,sh\<^sub>3)\<rangle>"
- and sh\<^sub>2D: "sh\<^sub>2 D = Some (sfs, Done) \<or> M = clinit \<and> sh\<^sub>2 D = \<lfloor>(sfs, Processing)\<rfloor>"
- and "method": "P \<turnstile> C sees M,Static: Ts\<rightarrow>T = (pns, body) in D"
- and same_len: "size vs = size pns"
- and l\<^sub>2': "l\<^sub>2' = [pns [\<mapsto>] vs]" by fact+
- have "fv (C\<bullet>\<^sub>sM(ps)) \<subseteq> W" by fact
- hence fvps: "fvs(ps) \<subseteq> W" by auto
- have wfmethod: "size Ts = size pns" and fvbd: "fv body \<subseteq> set pns"
- using "method" wf by(fastforce dest!:sees_wf_mdecl simp:wf_mdecl_def)+
- show ?case
- using IHbd[OF fvbd] l\<^sub>2' same_len wfmethod sh\<^sub>2D
- eval_evals.SCall[OF IHps[OF fvps] "method" _ same_len l\<^sub>2']
- by (simp add:subset_insertI)
-next
- case SeqThrow thus ?case by simp (blast intro: eval_evals.SeqThrow)
-next
- case CondT thus ?case by simp (blast intro: eval_evals.CondT)
-next
- case CondF thus ?case by simp (blast intro: eval_evals.CondF)
-next
- case CondThrow thus ?case by simp (blast intro: eval_evals.CondThrow)
-next
- case WhileF thus ?case by simp (blast intro: eval_evals.WhileF)
-next
- case WhileT thus ?case by simp (blast intro: eval_evals.WhileT)
-next
- case WhileCondThrow thus ?case by simp (blast intro: eval_evals.WhileCondThrow)
-next
- case WhileBodyThrow thus ?case by simp (blast intro: eval_evals.WhileBodyThrow)
-next
- case Throw thus ?case by simp (blast intro: eval_evals.Throw)
-next
- case ThrowNull thus ?case by simp (blast intro: eval_evals.ThrowNull)
-next
- case ThrowThrow thus ?case by simp (blast intro: eval_evals.ThrowThrow)
-next
- case Try thus ?case by simp (blast intro: eval_evals.Try)
-next
- case (TryCatch e\<^sub>1 h\<^sub>0 l\<^sub>0 sh\<^sub>0 a h\<^sub>1 l\<^sub>1 sh\<^sub>1 D fs C e\<^sub>2 V e\<^sub>2' h\<^sub>2 l\<^sub>2 sh\<^sub>2)
- have IH\<^sub>1: "\<And>W. fv e\<^sub>1 \<subseteq> W \<Longrightarrow> P \<turnstile> \<langle>e\<^sub>1,(h\<^sub>0,l\<^sub>0|`W,sh\<^sub>0)\<rangle> \<Rightarrow> \<langle>Throw a,(h\<^sub>1,l\<^sub>1|`W,sh\<^sub>1)\<rangle>"
- and IH\<^sub>2: "\<And>W. fv e\<^sub>2 \<subseteq> W \<Longrightarrow> P \<turnstile> \<langle>e\<^sub>2,(h\<^sub>1,l\<^sub>1(V\<mapsto>Addr a)|`W,sh\<^sub>1)\<rangle> \<Rightarrow> \<langle>e\<^sub>2',(h\<^sub>2,l\<^sub>2|`W,sh\<^sub>2)\<rangle>"
- and lookup: "h\<^sub>1 a = Some(D, fs)" and subtype: "P \<turnstile> D \<preceq>\<^sup>* C" by fact+
- have "fv (try e\<^sub>1 catch(C V) e\<^sub>2) \<subseteq> W" by fact
- hence fv\<^sub>1: "fv e\<^sub>1 \<subseteq> W" and fv\<^sub>2: "fv e\<^sub>2 \<subseteq> insert V W" by auto
- have IH\<^sub>2': "P \<turnstile> \<langle>e\<^sub>2,(h\<^sub>1,(l\<^sub>1|`W)(V \<mapsto> Addr a),sh\<^sub>1)\<rangle> \<Rightarrow> \<langle>e\<^sub>2',(h\<^sub>2,l\<^sub>2|`insert V W,sh\<^sub>2)\<rangle>"
- using IH\<^sub>2[OF fv\<^sub>2] fun_upd_restrict[of l\<^sub>1 W] (*FIXME just l|W instead of l|(W-V) in simp rule??*) by simp
- with eval_evals.TryCatch[OF IH\<^sub>1[OF fv\<^sub>1] _ subtype IH\<^sub>2'] lookup
- show ?case by fastforce
-next
- case TryThrow thus ?case by simp (blast intro: eval_evals.TryThrow)
-next
- case Nil thus ?case by (simp add: eval_evals.Nil)
-next
- case Cons thus ?case by simp (blast intro: eval_evals.Cons)
-next
- case ConsThrow thus ?case by simp (blast intro: eval_evals.ConsThrow)
-next
- case InitFinal thus ?case by (simp add: eval_evals.InitFinal)
-next
- case InitNone thus ?case by(blast intro: eval_evals.InitNone)
-next
- case InitDone thus ?case
- by (simp add: InitDone.hyps(2) InitDone.prems eval_evals.InitDone)
-next
- case InitProcessing thus ?case by (simp add: eval_evals.InitProcessing)
-next
- case InitError thus ?case using eval_evals.InitError by auto
-next
- case InitObject thus ?case by(simp add: eval_evals.InitObject)
-next
- case InitNonObject thus ?case by(simp add: eval_evals.InitNonObject)
-next
- case InitRInit thus ?case by(simp add: eval_evals.InitRInit)
-next
- case (RInit e h l sh v h' l' sh' C sfs i sh'' C' Cs e' e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1)
- have f1: "fv e \<subseteq> W \<and> fv (INIT C' (Cs,True) \<leftarrow> e') \<subseteq> W"
- using RInit.prems by auto
- then have f2: "P \<turnstile> \<langle>e,(h, l |` W, sh)\<rangle> \<Rightarrow> \<langle>Val v,(h', l' |` W, sh')\<rangle>"
- using RInit.hyps(2) by blast
- have "P \<turnstile> \<langle>INIT C' (Cs,True) \<leftarrow> e', (h', l' |` W, sh'')\<rangle> \<Rightarrow> \<langle>e\<^sub>1,(h\<^sub>1, l\<^sub>1 |` W, sh\<^sub>1)\<rangle>"
- using f1 by (meson RInit.hyps(7))
- then show ?case
- using f2 RInit.hyps(3) RInit.hyps(4) RInit.hyps(5) eval_evals.RInit by blast
-next
- case (RInitInitFail e h l sh a h' l' sh' C sfs i sh'' D Cs e' e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1)
- have f1: "fv e \<subseteq> W" "fv e' \<subseteq> W"
- using RInitInitFail.prems by auto
- then have f2: "P \<turnstile> \<langle>e,(h, l |` W, sh)\<rangle> \<Rightarrow> \<langle>throw a,(h', l' |` W, sh')\<rangle>"
- using RInitInitFail.hyps(2) by blast
- then have f2': "fv (throw a) \<subseteq> W"
- using eval_final[OF f2] by auto
- then have f1': "fv (RI (D,throw a);Cs \<leftarrow> e') \<subseteq> W"
- using f1 by auto
- have "P \<turnstile> \<langle>RI (D,throw a);Cs \<leftarrow> e', (h', l' |` W, sh'')\<rangle> \<Rightarrow> \<langle>e\<^sub>1,(h\<^sub>1, l\<^sub>1 |` W, sh\<^sub>1)\<rangle>"
- using f1' by (meson RInitInitFail.hyps(6))
- then show ?case
- using f2 by (simp add: RInitInitFail.hyps(3,4) eval_evals.RInitInitFail)
-next
- case (RInitFailFinal e h l sh a h' l' sh' sh'' C)
- have f1: "fv e \<subseteq> W"
- using RInitFailFinal.prems by auto
- then have f2: "P \<turnstile> \<langle>e,(h, l |` W, sh)\<rangle> \<Rightarrow> \<langle>throw a,(h', l' |` W, sh')\<rangle>"
- using RInitFailFinal.hyps(2) by blast
- then have f2': "fv (throw a) \<subseteq> W"
- using eval_final[OF f2] by auto
- then show ?case using f2 RInitFailFinal.hyps(3,4) eval_evals.RInitFailFinal by blast
-qed
-(*>*)
-
-
-lemma eval_notfree_unchanged:
- "P \<turnstile> \<langle>e,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',(h',l',sh')\<rangle> \<Longrightarrow> (\<And>V. V \<notin> fv e \<Longrightarrow> l' V = l V)"
-and "P \<turnstile> \<langle>es,(h,l,sh)\<rangle> [\<Rightarrow>] \<langle>es',(h',l',sh')\<rangle> \<Longrightarrow> (\<And>V. V \<notin> fvs es \<Longrightarrow> l' V = l V)"
-(*<*)
-proof(induct rule:eval_evals_inducts)
- case LAss thus ?case by(simp add:fun_upd_apply)
-next
- case Block thus ?case
- by (simp only:fun_upd_apply split:if_splits) fastforce
-next
- case TryCatch thus ?case
- by (simp only:fun_upd_apply split:if_splits) fastforce
-next
- case (RInitInitFail e h l sh a h' l' sh' C sfs i sh'' D Cs e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1)
- have "fv (throw a) = {}"
- using RInitInitFail.hyps(1) eval_final final_fv by blast
- then have "fv e \<union> fv (RI (D,throw a) ; Cs \<leftarrow> unit) \<subseteq> fv (RI (C,e) ; D#Cs \<leftarrow> unit)"
- by auto
- then show ?case using RInitInitFail.hyps(2,6) RInitInitFail.prems by fastforce
-qed simp_all
-(*>*)
-
-
-lemma eval_closed_lcl_unchanged:
- "\<lbrakk> P \<turnstile> \<langle>e,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',(h',l',sh')\<rangle>; fv e = {} \<rbrakk> \<Longrightarrow> l' = l"
-(*<*)by(fastforce dest:eval_notfree_unchanged simp add:fun_eq_iff [where 'b="val option"])(*>*)
-
-
-lemma list_eval_Throw:
-assumes eval_e: "P \<turnstile> \<langle>throw x,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
-shows "P \<turnstile> \<langle>map Val vs @ throw x # es',s\<rangle> [\<Rightarrow>] \<langle>map Val vs @ e' # es',s'\<rangle>"
-(*<*)
-proof -
- from eval_e
- obtain a where e': "e' = Throw a"
- by (cases) (auto dest!: eval_final)
- {
- fix es
- have "\<And>vs. es = map Val vs @ throw x # es'
- \<Longrightarrow> P \<turnstile> \<langle>es,s\<rangle>[\<Rightarrow>]\<langle>map Val vs @ e' # es',s'\<rangle>"
- proof (induct es type: list)
- case Nil thus ?case by simp
- next
- case (Cons e es vs)
- have e_es: "e # es = map Val vs @ throw x # es'" by fact
- show "P \<turnstile> \<langle>e # es,s\<rangle> [\<Rightarrow>] \<langle>map Val vs @ e' # es',s'\<rangle>"
- proof (cases vs)
- case Nil
- with e_es obtain "e=throw x" "es=es'" by simp
- moreover from eval_e e'
- have "P \<turnstile> \<langle>throw x # es,s\<rangle> [\<Rightarrow>] \<langle>Throw a # es,s'\<rangle>"
- by (iprover intro: ConsThrow)
- ultimately show ?thesis using Nil e' by simp
- next
- case (Cons v vs')
- have vs: "vs = v # vs'" by fact
- with e_es obtain
- e: "e=Val v" and es:"es= map Val vs' @ throw x # es'"
- by simp
- from e
- have "P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>Val v,s\<rangle>"
- by (iprover intro: eval_evals.Val)
- moreover from es
- have "P \<turnstile> \<langle>es,s\<rangle> [\<Rightarrow>] \<langle>map Val vs' @ e' # es',s'\<rangle>"
- by (rule Cons.hyps)
- ultimately show
- "P \<turnstile> \<langle>e#es,s\<rangle> [\<Rightarrow>] \<langle>map Val vs @ e' # es',s'\<rangle>"
- using vs by (auto intro: eval_evals.Cons)
- qed
- qed
- }
- thus ?thesis
- by simp
-qed
-(*>*)
-
-\<comment> \<open> separate evaluation of first subexp of a sequence \<close>
-lemma seq_ext:
-assumes IH: "\<And>e' s'. P \<turnstile> \<langle>e'',s''\<rangle> \<Rightarrow> \<langle>e',s'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
- and seq: "P \<turnstile> \<langle>e'' ;; e\<^sub>0,s''\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
-shows "P \<turnstile> \<langle>e ;; e\<^sub>0,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
-proof(rule eval_cases(14)[OF seq]) \<comment> \<open> Seq \<close>
- fix v' s\<^sub>1 assume e'': "P \<turnstile> \<langle>e'',s''\<rangle> \<Rightarrow> \<langle>Val v',s\<^sub>1\<rangle>" and estep: "P \<turnstile> \<langle>e\<^sub>0,s\<^sub>1\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
- have "P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>Val v',s\<^sub>1\<rangle>" using e'' IH by simp
- then show ?thesis using estep Seq by simp
-next
- fix e\<^sub>t assume e'': "P \<turnstile> \<langle>e'',s''\<rangle> \<Rightarrow> \<langle>throw e\<^sub>t,s'\<rangle>" and e': "e' = throw e\<^sub>t"
- have "P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>throw e\<^sub>t,s'\<rangle>" using e'' IH by simp
- then show ?thesis using eval_evals.SeqThrow e' by simp
-qed
-
-\<comment> \<open> separate evaluation of @{text RI} subexp, val case \<close>
-lemma rinit_Val_ext:
-assumes ri: "P \<turnstile> \<langle>RI (C,e'') ; Cs \<leftarrow> e\<^sub>0,s''\<rangle> \<Rightarrow> \<langle>Val v',s\<^sub>1\<rangle>"
- and IH: "\<And>e' s'. P \<turnstile> \<langle>e'',s''\<rangle> \<Rightarrow> \<langle>e',s'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
-shows "P \<turnstile> \<langle>RI (C,e) ; Cs \<leftarrow> e\<^sub>0,s\<rangle> \<Rightarrow> \<langle>Val v',s\<^sub>1\<rangle>"
-proof(rule eval_cases(20)[OF ri]) \<comment> \<open> RI \<close>
- fix v'' h' l' sh' sfs i
- assume e''step: "P \<turnstile> \<langle>e'',s''\<rangle> \<Rightarrow> \<langle>Val v'',(h', l', sh')\<rangle>"
- and shC: "sh' C = \<lfloor>(sfs, i)\<rfloor>"
- and init: "P \<turnstile> \<langle>INIT (if Cs = [] then C else last Cs) (Cs,True) \<leftarrow> e\<^sub>0,(h', l', sh'(C \<mapsto> (sfs, Done)))\<rangle> \<Rightarrow>
- \<langle>Val v',s\<^sub>1\<rangle>"
- have "P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>Val v'',(h', l', sh')\<rangle>" using IH[OF e''step] by simp
- then show ?thesis using RInit init shC by auto
-next
- fix a h' l' sh' sfs i D Cs'
- assume e''step: "P \<turnstile> \<langle>e'',s''\<rangle> \<Rightarrow> \<langle>throw a,(h', l', sh')\<rangle>"
- and riD: "P \<turnstile> \<langle>RI (D,throw a) ; Cs' \<leftarrow> e\<^sub>0,(h', l', sh'(C \<mapsto> (sfs, Error)))\<rangle> \<Rightarrow> \<langle>Val v',s\<^sub>1\<rangle>"
- have "P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>throw a,(h', l', sh')\<rangle>" using IH[OF e''step] by simp
- then show ?thesis using riD rinit_throwE by blast
-qed(simp)
-
-\<comment> \<open> separate evaluation of @{text RI} subexp, throw case \<close>
-lemma rinit_throw_ext:
-assumes ri: "P \<turnstile> \<langle>RI (C,e'') ; Cs \<leftarrow> e\<^sub>0,s''\<rangle> \<Rightarrow> \<langle>throw e\<^sub>t,s'\<rangle>"
- and IH: "\<And>e' s'. P \<turnstile> \<langle>e'',s''\<rangle> \<Rightarrow> \<langle>e',s'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
-shows "P \<turnstile> \<langle>RI (C,e) ; Cs \<leftarrow> e\<^sub>0,s\<rangle> \<Rightarrow> \<langle>throw e\<^sub>t,s'\<rangle>"
-proof(rule eval_cases(20)[OF ri]) \<comment> \<open> RI \<close>
- fix v h' l' sh' sfs i
- assume e''step: "P \<turnstile> \<langle>e'',s''\<rangle> \<Rightarrow> \<langle>Val v,(h', l', sh')\<rangle>"
- and shC: "sh' C = \<lfloor>(sfs, i)\<rfloor>"
- and init: "P \<turnstile> \<langle>INIT (if Cs = [] then C else last Cs) (Cs,True) \<leftarrow> e\<^sub>0,(h', l', sh'(C \<mapsto> (sfs, Done)))\<rangle> \<Rightarrow>
- \<langle>throw e\<^sub>t,s'\<rangle>"
- have "P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>Val v,(h', l', sh')\<rangle>" using IH[OF e''step] by simp
- then show ?thesis using RInit init shC by auto
-next
- fix a h' l' sh' sfs i D Cs'
- assume e''step: "P \<turnstile> \<langle>e'',s''\<rangle> \<Rightarrow> \<langle>throw a,(h', l', sh')\<rangle>"
- and shC: "sh' C = \<lfloor>(sfs, i)\<rfloor>"
- and riD: "P \<turnstile> \<langle>RI (D,throw a) ; Cs' \<leftarrow> e\<^sub>0,(h', l', sh'(C \<mapsto> (sfs, Error)))\<rangle> \<Rightarrow> \<langle>throw e\<^sub>t,s'\<rangle>"
- and cons: "Cs = D # Cs'"
- have estep': "P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>throw a,(h', l', sh')\<rangle>" using IH[OF e''step] by simp
- then show ?thesis using RInitInitFail cons riD shC by simp
-next
- fix a h' l' sh' sfs i
- assume "throw e\<^sub>t = throw a"
- and "s' = (h', l', sh'(C \<mapsto> (sfs, Error)))"
- and "P \<turnstile> \<langle>e'',s''\<rangle> \<Rightarrow> \<langle>throw a,(h', l', sh')\<rangle>"
- and "sh' C = \<lfloor>(sfs, i)\<rfloor>"
- and "Cs = []"
- then show ?thesis using RInitFailFinal IH by auto
-qed
-
-\<comment> \<open> separate evaluation of @{text RI} subexp \<close>
-lemma rinit_ext:
-assumes IH: "\<And>e' s'. P \<turnstile> \<langle>e'',s''\<rangle> \<Rightarrow> \<langle>e',s'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
-shows "\<And>e' s'. P \<turnstile> \<langle>RI (C,e'') ; Cs \<leftarrow> e\<^sub>0,s''\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>
- \<Longrightarrow> P \<turnstile> \<langle>RI (C,e) ; Cs \<leftarrow> e\<^sub>0,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
-proof -
- fix e' s' assume ri'': "P \<turnstile> \<langle>RI (C,e'') ; Cs \<leftarrow> e\<^sub>0,s''\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
- then have "final e'" using eval_final by simp
- then show "P \<turnstile> \<langle>RI (C,e) ; Cs \<leftarrow> e\<^sub>0,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
- proof(rule finalE)
- fix v assume "e' = Val v" then show ?thesis using rinit_Val_ext[OF _ IH] ri'' by simp
- next
- fix a assume "e' = throw a" then show ?thesis using rinit_throw_ext[OF _ IH] ri'' by simp
- qed
-qed
-
-\<comment> \<open> @{text INIT} and @{text RI} return either @{text Val} with @{text Done} or
- @{text Processing} flag or @{text Throw} with @{text Error} flag \<close>
-lemma
-shows eval_init_return: "P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>
- \<Longrightarrow> iconf (shp s) e
- \<Longrightarrow> (\<exists>Cs b. e = INIT C' (Cs,b) \<leftarrow> unit) \<or> (\<exists>C e\<^sub>0 Cs e\<^sub>i. e = RI(C,e\<^sub>0);Cs@[C'] \<leftarrow> unit)
- \<or> (\<exists>e\<^sub>0. e = RI(C',e\<^sub>0);Nil \<leftarrow> unit)
- \<Longrightarrow> (val_of e' = Some v \<longrightarrow> (\<exists>sfs i. shp s' C' = \<lfloor>(sfs,i)\<rfloor> \<and> (i = Done \<or> i = Processing)))
- \<and> (throw_of e' = Some a \<longrightarrow> (\<exists>sfs i. shp s' C' = \<lfloor>(sfs,Error)\<rfloor>))"
-and "P \<turnstile> \<langle>es,s\<rangle> [\<Rightarrow>] \<langle>es',s'\<rangle> \<Longrightarrow> True"
-proof(induct rule: eval_evals.inducts)
- case (InitFinal e s e' s' C b) then show ?case
- by(fastforce simp: initPD_def dest: eval_final_same)
-next
- case (InitDone sh C sfs C' Cs e h l e' s')
- then have "final e'" using eval_final by simp
- then show ?case
- proof(rule finalE)
- fix v assume e': "e' = Val v" then show ?thesis using InitDone initPD_def
- proof(cases Cs) qed(auto)
- next
- fix a assume e': "e' = throw a" then show ?thesis using InitDone initPD_def
- proof(cases Cs) qed(auto)
- qed
-next
- case (InitProcessing sh C sfs C' Cs e h l e' s')
- then have "final e'" using eval_final by simp
- then show ?case
- proof(rule finalE)
- fix v assume e': "e' = Val v" then show ?thesis using InitProcessing initPD_def
- proof(cases Cs) qed(auto)
- next
- fix a assume e': "e' = throw a" then show ?thesis using InitProcessing initPD_def
- proof(cases Cs) qed(auto)
- qed
-next
- case (InitError sh C sfs Cs e h l e' s' C') show ?case
- proof(cases Cs)
- case Nil then show ?thesis using InitError by simp
- next
- case (Cons C2 list)
- then have "final e'" using InitError eval_final by simp
- then show ?thesis
- proof(rule finalE)
- fix v assume e': "e' = Val v" then show ?thesis
- using InitError
- proof -
- obtain ccss :: "char list list" and bb :: bool where
- "INIT C' (C # Cs,False) \<leftarrow> e = INIT C' (ccss,bb) \<leftarrow> unit"
- using InitError.prems(2) by blast
- then show ?thesis using InitError.hyps(2) e' by(auto dest!: rinit_throwE)
- qed
- next
- fix a assume e': "e' = throw a"
- then show ?thesis using Cons InitError cons_to_append[of list] by clarsimp
- qed
- qed
-next
- case (InitRInit C Cs h l sh e' s' C') show ?case
- proof(cases Cs)
- case Nil then show ?thesis using InitRInit by simp
- next
- case (Cons C' list) then show ?thesis
- using InitRInit Cons cons_to_append[of list] by clarsimp
- qed
-next
- case (RInit e s v h' l' sh' C sfs i sh'' C' Cs e' e\<^sub>1 s\<^sub>1)
- then have final: "final e\<^sub>1" using eval_final by simp
- then show ?case
- proof(cases Cs)
- case Nil show ?thesis using final
- proof(rule finalE)
- fix v assume e': "e\<^sub>1 = Val v" show ?thesis
- using RInit Nil by(auto simp: fun_upd_same initPD_def)
- next
- fix a assume e': "e\<^sub>1 = throw a" show ?thesis
- using RInit Nil by(auto simp: fun_upd_same initPD_def)
- qed
- next
- case (Cons a list) show ?thesis
- proof(rule finalE[OF final])
- fix v assume e': "e\<^sub>1 = Val v" then show ?thesis
- using RInit Cons by clarsimp (metis last.simps last_appendR list.distinct(1))
- next
- fix a assume e': "e\<^sub>1 = throw a" then show ?thesis
- using RInit Cons by clarsimp (metis last.simps last_appendR list.distinct(1))
- qed
- qed
-next
- case (RInitInitFail e s a h' l' sh' C sfs i sh'' D Cs e' e\<^sub>1 s\<^sub>1)
- then have final: "final e\<^sub>1" using eval_final by simp
- then show ?case
- proof(rule finalE)
- fix v assume e': "e\<^sub>1 = Val v" then show ?thesis
- using RInitInitFail by clarsimp (meson exp.distinct(101) rinit_throwE)
- next
- fix a' assume e': "e\<^sub>1 = Throw a'"
- then have "iconf (sh'(C \<mapsto> (sfs, Error))) a"
- using RInitInitFail.hyps(1) eval_final by fastforce
- then show ?thesis using RInitInitFail e'
- by clarsimp (meson Cons_eq_append_conv list.inject)
- qed
-qed(auto simp: fun_upd_same)
-
-lemma init_Val_PD: "P \<turnstile> \<langle>INIT C' (Cs,b) \<leftarrow> unit,s\<rangle> \<Rightarrow> \<langle>Val v,s'\<rangle>
- \<Longrightarrow> iconf (shp s) (INIT C' (Cs,b) \<leftarrow> unit)
- \<Longrightarrow> \<exists>sfs i. shp s' C' = \<lfloor>(sfs,i)\<rfloor> \<and> (i = Done \<or> i = Processing)"
- by(drule_tac v = v in eval_init_return, simp+)
-
-lemma init_throw_PD: "P \<turnstile> \<langle>INIT C' (Cs,b) \<leftarrow> unit,s\<rangle> \<Rightarrow> \<langle>throw a,s'\<rangle>
- \<Longrightarrow> iconf (shp s) (INIT C' (Cs,b) \<leftarrow> unit)
- \<Longrightarrow> \<exists>sfs i. shp s' C' = \<lfloor>(sfs,Error)\<rfloor>"
- by(drule_tac a = a in eval_init_return, simp+)
-
-lemma rinit_Val_PD: "P \<turnstile> \<langle>RI(C,e\<^sub>0);Cs \<leftarrow> unit,s\<rangle> \<Rightarrow> \<langle>Val v,s'\<rangle>
- \<Longrightarrow> iconf (shp s) (RI(C,e\<^sub>0);Cs \<leftarrow> unit) \<Longrightarrow> last(C#Cs) = C'
- \<Longrightarrow> \<exists>sfs i. shp s' C' = \<lfloor>(sfs,i)\<rfloor> \<and> (i = Done \<or> i = Processing)"
-by(auto dest!: eval_init_return [where C'=C']
- append_butlast_last_id[THEN sym])
-
-lemma rinit_throw_PD: "P \<turnstile> \<langle>RI(C,e\<^sub>0);Cs \<leftarrow> unit,s\<rangle> \<Rightarrow> \<langle>throw a,s'\<rangle>
- \<Longrightarrow> iconf (shp s) (RI(C,e\<^sub>0);Cs \<leftarrow> unit) \<Longrightarrow> last(C#Cs) = C'
- \<Longrightarrow> \<exists>sfs i. shp s' C' = \<lfloor>(sfs,Error)\<rfloor>"
-by(auto dest!: eval_init_return [where C'=C']
- append_butlast_last_id[THEN sym])
-
-\<comment> \<open> combining the above to get evaluation of @{text INIT} in a sequence \<close>
-
-(* Hiermit kann man die ganze pair-Splitterei in den automatischen Taktiken
-abschalten. Wieder anschalten siehe nach dem Beweis. *)
-(*<*)
-declare split_paired_All [simp del] split_paired_Ex [simp del]
-(*>*)
-
-lemma eval_init_seq': "P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>
- \<Longrightarrow> (\<exists>C Cs b e\<^sub>i. e = INIT C (Cs,b) \<leftarrow> e\<^sub>i) \<or> (\<exists>C e\<^sub>0 Cs e\<^sub>i. e = RI(C,e\<^sub>0);Cs \<leftarrow> e\<^sub>i)
- \<Longrightarrow> (\<exists>C Cs b e\<^sub>i. e = INIT C (Cs,b) \<leftarrow> e\<^sub>i \<and> P \<turnstile> \<langle>(INIT C (Cs,b) \<leftarrow> unit);; e\<^sub>i,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>)
- \<or> (\<exists>C e\<^sub>0 Cs e\<^sub>i. e = RI(C,e\<^sub>0);Cs \<leftarrow> e\<^sub>i \<and> P \<turnstile> \<langle>(RI(C,e\<^sub>0);Cs \<leftarrow> unit);; e\<^sub>i,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>)"
-and "P \<turnstile> \<langle>es,s\<rangle> [\<Rightarrow>] \<langle>es',s'\<rangle> \<Longrightarrow> True"
-proof(induct rule: eval_evals.inducts)
- case InitFinal then show ?case by(auto simp: Seq[OF eval_evals.InitFinal[OF Val[where v=Unit]]])
-next
- case (InitNone sh) then show ?case
- using seq_ext[OF eval_evals.InitNone[where sh=sh and e=unit, OF InitNone.hyps(1)]] by fastforce
-next
- case (InitDone sh) then show ?case
- using seq_ext[OF eval_evals.InitDone[where sh=sh and e=unit, OF InitDone.hyps(1)]] by fastforce
-next
- case (InitProcessing sh) then show ?case
- using seq_ext[OF eval_evals.InitProcessing[where sh=sh and e=unit, OF InitProcessing.hyps(1)]]
- by fastforce
-next
- case (InitError sh) then show ?case
- using seq_ext[OF eval_evals.InitError[where sh=sh and e=unit, OF InitError.hyps(1)]] by fastforce
-next
- case (InitObject sh) then show ?case
- using seq_ext[OF eval_evals.InitObject[where sh=sh and e=unit, OF InitObject.hyps(1)]]
- by fastforce
-next
- case (InitNonObject sh) then show ?case
- using seq_ext[OF eval_evals.InitNonObject[where sh=sh and e=unit, OF InitNonObject.hyps(1)]]
- by fastforce
-next
- case (InitRInit C Cs e h l sh e' s' C') then show ?case
- using seq_ext[OF eval_evals.InitRInit[where e=unit]] by fastforce
-next
- case RInit then show ?case
- using seq_ext[OF eval_evals.RInit[where e=unit, OF RInit.hyps(1)]] by fastforce
-next
- case RInitInitFail then show ?case
- using seq_ext[OF eval_evals.RInitInitFail[where e=unit, OF RInitInitFail.hyps(1)]] by fastforce
-next
- case RInitFailFinal
- then show ?case using eval_evals.RInitFailFinal eval_evals.SeqThrow by auto
-qed(auto)
-
-lemma eval_init_seq: "P \<turnstile> \<langle>INIT C (Cs,b) \<leftarrow> e,(h, l, sh)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>
- \<Longrightarrow> P \<turnstile> \<langle>(INIT C (Cs,b) \<leftarrow> unit);; e,(h, l, sh)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
- by(auto dest: eval_init_seq')
-
-
-text \<open> The key lemma: \<close>
-lemma
-assumes wf: "wwf_J_prog P"
-shows extend_1_eval: "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow> \<langle>e'',s'',b''\<rangle> \<Longrightarrow> P,shp s \<turnstile>\<^sub>b (e,b) \<surd>
- \<Longrightarrow> (\<And>s' e'. P \<turnstile> \<langle>e'',s''\<rangle> \<Rightarrow> \<langle>e',s'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>)"
-and extend_1_evals: "P \<turnstile> \<langle>es,s,b\<rangle> [\<rightarrow>] \<langle>es'',s'',b''\<rangle> \<Longrightarrow> P,shp s \<turnstile>\<^sub>b (es,b) \<surd>
- \<Longrightarrow> (\<And>s' es'. P \<turnstile> \<langle>es'',s''\<rangle> [\<Rightarrow>] \<langle>es',s'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>es,s\<rangle> [\<Rightarrow>] \<langle>es',s'\<rangle>)"
-proof (induct rule: red_reds.inducts)
- case (RedNew h a C FDTs h' l sh)
- then have e':"e' = addr a" and s':"s' = (h(a \<mapsto> blank P C), l, sh)"
- using eval_cases(3) by fastforce+
- obtain sfs i where shC: "sh C = \<lfloor>(sfs, i)\<rfloor>" and "i = Done \<or> i = Processing"
- using RedNew by (clarsimp simp: bconf_def initPD_def)
- then show ?case
- proof(cases i)
- case Done then show ?thesis using RedNew shC e' s' New by simp
- next
- case Processing
- then have shC': "\<nexists>sfs. sh C = Some(sfs,Done)" and shP: "sh C = Some(sfs,Processing)"
- using shC by simp+
- then have init: "P \<turnstile> \<langle>INIT C ([C],False) \<leftarrow> unit,(h,l,sh)\<rangle> \<Rightarrow> \<langle>unit,(h,l,sh)\<rangle>"
- by(simp add: InitFinal InitProcessing Val)
- have "P \<turnstile> \<langle>new C,(h, l, sh)\<rangle> \<Rightarrow> \<langle>addr a,(h(a \<mapsto> blank P C),l,sh)\<rangle>"
- using RedNew shC' by(auto intro: NewInit[OF _ init])
- then show ?thesis using e' s' by simp
- qed(auto)
-next
- case (RedNewFail h C l sh)
- then have e':"e' = THROW OutOfMemory" and s':"s' = (h, l, sh)"
- using eval_final_same final_def by fastforce+
- obtain sfs i where shC: "sh C = \<lfloor>(sfs, i)\<rfloor>" and "i = Done \<or> i = Processing"
- using RedNewFail by (clarsimp simp: bconf_def initPD_def)
- then show ?case
- proof(cases i)
- case Done then show ?thesis using RedNewFail shC e' s' NewFail by simp
- next
- case Processing
- then have shC': "\<nexists>sfs. sh C = Some(sfs,Done)" and shP: "sh C = Some(sfs,Processing)"
- using shC by simp+
- then have init: "P \<turnstile> \<langle>INIT C ([C],False) \<leftarrow> unit,(h,l,sh)\<rangle> \<Rightarrow> \<langle>unit,(h,l,sh)\<rangle>"
- by(simp add: InitFinal InitProcessing Val)
- have "P \<turnstile> \<langle>new C,(h, l, sh)\<rangle> \<Rightarrow> \<langle>THROW OutOfMemory,(h,l,sh)\<rangle>"
- using RedNewFail shC' by(auto intro: NewInitOOM[OF _ init])
- then show ?thesis using e' s' by simp
- qed(auto)
-next
- case (NewInitRed sh C h l)
- then have seq: "P \<turnstile> \<langle>(INIT C ([C],False) \<leftarrow> unit);; new C,(h, l, sh)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
- using eval_init_seq by simp
- then show ?case
- proof(rule eval_cases(14)) \<comment> \<open> Seq \<close>
- fix v s\<^sub>1 assume init: "P \<turnstile> \<langle>INIT C ([C],False) \<leftarrow> unit,(h, l, sh)\<rangle> \<Rightarrow> \<langle>Val v,s\<^sub>1\<rangle>"
- and new: "P \<turnstile> \<langle>new C,s\<^sub>1\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
- obtain h\<^sub>1 l\<^sub>1 sh\<^sub>1 where s\<^sub>1: "s\<^sub>1 = (h\<^sub>1,l\<^sub>1,sh\<^sub>1)" by(cases s\<^sub>1)
- then obtain sfs i where shC: "sh\<^sub>1 C = \<lfloor>(sfs, i)\<rfloor>" and iDP: "i = Done \<or> i = Processing"
- using init_Val_PD[OF init] by auto
- show ?thesis
- proof(rule eval_cases(1)[OF new]) \<comment> \<open> New \<close>
- fix sha ha a FDTs la
- assume s\<^sub>1a: "s\<^sub>1 = (ha, la, sha)" and e': "e' = addr a"
- and s': "s' = (ha(a \<mapsto> blank P C), la, sha)"
- and addr: "new_Addr ha = \<lfloor>a\<rfloor>" and fields: "P \<turnstile> C has_fields FDTs"
- then show ?thesis using NewInit[OF _ _ addr fields] NewInitRed.hyps init by simp
- next
- fix sha ha la
- assume "s\<^sub>1 = (ha, la, sha)" and "e' = THROW OutOfMemory"
- and "s' = (ha, la, sha)" and "new_Addr ha = None"
- then show ?thesis using NewInitOOM NewInitRed.hyps init by simp
- next
- fix sha ha la v' h' l' sh' a FDTs
- assume s\<^sub>1a: "s\<^sub>1 = (ha, la, sha)" and e': "e' = addr a"
- and s': "s' = (h'(a \<mapsto> blank P C), l', sh')"
- and shaC: "\<forall>sfs. sha C \<noteq> \<lfloor>(sfs, Done)\<rfloor>"
- and init': "P \<turnstile> \<langle>INIT C ([C],False) \<leftarrow> unit,(ha, la, sha)\<rangle> \<Rightarrow> \<langle>Val v',(h', l', sh')\<rangle>"
- and addr: "new_Addr h' = \<lfloor>a\<rfloor>" and fields: "P \<turnstile> C has_fields FDTs"
- then have i: "i = Processing" using iDP shC s\<^sub>1 by simp
- then have "(h', l', sh') = (ha, la, sha)" using init' init_ProcessingE s\<^sub>1 s\<^sub>1a shC by blast
- then show ?thesis using NewInit NewInitRed.hyps s\<^sub>1a addr fields init shaC e' s' by auto
- next
- fix sha ha la v' h' l' sh'
- assume s\<^sub>1a: "s\<^sub>1 = (ha, la, sha)" and e': "e' = THROW OutOfMemory"
- and s': "s' = (h', l', sh')" and "\<forall>sfs. sha C \<noteq> \<lfloor>(sfs, Done)\<rfloor>"
- and init': "P \<turnstile> \<langle>INIT C ([C],False) \<leftarrow> unit,(ha, la, sha)\<rangle> \<Rightarrow> \<langle>Val v',(h', l', sh')\<rangle>"
- and addr: "new_Addr h' = None"
- then have i: "i = Processing" using iDP shC s\<^sub>1 by simp
- then have "(h', l', sh') = (ha, la, sha)" using init' init_ProcessingE s\<^sub>1 s\<^sub>1a shC by blast
- then show ?thesis
- using NewInitOOM NewInitRed.hyps e' addr s' s\<^sub>1a init by auto
- next
- fix sha ha la a
- assume s\<^sub>1a: "s\<^sub>1 = (ha, la, sha)"
- and "\<forall>sfs. sha C \<noteq> \<lfloor>(sfs, Done)\<rfloor>"
- and init': "P \<turnstile> \<langle>INIT C ([C],False) \<leftarrow> unit,(ha, la, sha)\<rangle> \<Rightarrow> \<langle>throw a,s'\<rangle>"
- then have i: "i = Processing" using iDP shC s\<^sub>1 by simp
- then show ?thesis using init' init_ProcessingE s\<^sub>1 s\<^sub>1a shC by blast
- qed
- next
- fix e assume e': "e' = throw e"
- and init: "P \<turnstile> \<langle>INIT C ([C],False) \<leftarrow> unit,(h, l, sh)\<rangle> \<Rightarrow> \<langle>throw e,s'\<rangle>"
- obtain h' l' sh' where s': "s' = (h',l',sh')" by(cases s')
- then obtain sfs i where shC: "sh' C = \<lfloor>(sfs, i)\<rfloor>" and iDP: "i = Error"
- using init_throw_PD[OF init] by auto
- then show ?thesis by (simp add: NewInitRed.hyps NewInitThrow e' init)
- qed
-next
- case CastRed then show ?case
- by(fastforce elim!: eval_cases intro: eval_evals.intros intro!: CastFail)
-next
- case RedCastNull
- then show ?case
- by simp (iprover elim: eval_cases intro: eval_evals.intros)
-next
- case RedCast
- then show ?case
- by (auto elim: eval_cases intro: eval_evals.intros)
-next
- case RedCastFail
- then show ?case
- by (auto elim!: eval_cases intro: eval_evals.intros)
-next
- case BinOpRed1 then show ?case
- by(fastforce elim!: eval_cases intro: eval_evals.intros simp: val_no_step)
-next
- case BinOpRed2
- thus ?case
- by (fastforce elim!: eval_cases intro: eval_evals.intros eval_finalId)
-next
- case RedBinOp
- thus ?case
- by simp (iprover elim: eval_cases intro: eval_evals.intros)
-next
- case RedVar
- thus ?case
- by (fastforce elim: eval_cases intro: eval_evals.intros)
-next
- case LAssRed thus ?case
- by(fastforce elim: eval_cases intro: eval_evals.intros)
-next
- case RedLAss
- thus ?case
- by (fastforce elim: eval_cases intro: eval_evals.intros)
-next
- case FAccRed thus ?case
- by(fastforce elim: eval_cases intro: eval_evals.intros)
-next
- case RedFAcc then show ?case
- by (fastforce elim: eval_cases intro: eval_evals.intros)
-next
- case RedFAccNull then show ?case
- by (fastforce elim!: eval_cases intro: eval_evals.intros)
-next
- case RedFAccNone thus ?case
- by(fastforce elim: eval_cases intro: eval_evals.intros)
-next
- case RedFAccStatic thus ?case
- by(fastforce elim: eval_cases intro: eval_evals.intros)
-next
- case (RedSFAcc C F t D sh sfs i v h l)
- then have e':"e' = Val v" and s':"s' = (h, l, sh)"
- using eval_cases(3) by fastforce+
- have "i = Done \<or> i = Processing" using RedSFAcc by (clarsimp simp: bconf_def initPD_def)
- then show ?case
- proof(cases i)
- case Done then show ?thesis using RedSFAcc e' s' SFAcc by simp
- next
- case Processing
- then have shC': "\<nexists>sfs. sh D = Some(sfs,Done)" and shP: "sh D = Some(sfs,Processing)"
- using RedSFAcc by simp+
- then have init: "P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(h,l,sh)\<rangle> \<Rightarrow> \<langle>unit,(h,l,sh)\<rangle>"
- by(simp add: InitFinal InitProcessing Val)
- have "P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D},(h, l, sh)\<rangle> \<Rightarrow> \<langle>Val v,(h,l,sh)\<rangle>"
- by(rule SFAccInit[OF RedSFAcc.hyps(1) shC' init shP RedSFAcc.hyps(3)])
- then show ?thesis using e' s' by simp
- qed(auto)
-next
- case (SFAccInitRed C F t D sh h l)
- then have seq: "P \<turnstile> \<langle>(INIT D ([D],False) \<leftarrow> unit);; C\<bullet>\<^sub>sF{D},(h, l, sh)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
- using eval_init_seq by simp
- then show ?case
- proof(rule eval_cases(14)) \<comment> \<open> Seq \<close>
- fix v s\<^sub>1 assume init: "P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(h, l, sh)\<rangle> \<Rightarrow> \<langle>Val v,s\<^sub>1\<rangle>"
- and acc: "P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D},s\<^sub>1\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
- obtain h\<^sub>1 l\<^sub>1 sh\<^sub>1 where s\<^sub>1: "s\<^sub>1 = (h\<^sub>1,l\<^sub>1,sh\<^sub>1)" by(cases s\<^sub>1)
- then obtain sfs i where shD: "sh\<^sub>1 D = \<lfloor>(sfs, i)\<rfloor>" and iDP: "i = Done \<or> i = Processing"
- using init_Val_PD[OF init] by auto
- show ?thesis
- proof(rule eval_cases(8)[OF acc]) \<comment> \<open> SFAcc \<close>
- fix t sha sfs v ha la
- assume "s\<^sub>1 = (ha, la, sha)" and "e' = Val v"
- and "s' = (ha, la, sha)" and "P \<turnstile> C has F,Static:t in D"
- and "sha D = \<lfloor>(sfs, Done)\<rfloor>" and "sfs F = \<lfloor>v\<rfloor>"
- then show ?thesis using SFAccInit SFAccInitRed.hyps(2) init by auto
- next
- fix t sha ha la v' h' l' sh' sfs i' v
- assume s\<^sub>1a: "s\<^sub>1 = (ha, la, sha)" and e': "e' = Val v"
- and s': "s' = (h', l', sh')" and field: "P \<turnstile> C has F,Static:t in D"
- and "\<forall>sfs. sha D \<noteq> \<lfloor>(sfs, Done)\<rfloor>"
- and init': "P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(ha, la, sha)\<rangle> \<Rightarrow> \<langle>Val v',(h', l', sh')\<rangle>"
- and shD': "sh' D = \<lfloor>(sfs, i')\<rfloor>" and sfsF: "sfs F = \<lfloor>v\<rfloor>"
- then have i: "i = Processing" using iDP shD s\<^sub>1 by simp
- then have "(h', l', sh') = (ha, la, sha)" using init' init_ProcessingE s\<^sub>1 s\<^sub>1a shD by blast
- then show ?thesis using SFAccInit SFAccInitRed.hyps(2) e' s' field init s\<^sub>1a sfsF shD' by auto
- next
- fix t sha ha la a
- assume s\<^sub>1a: "s\<^sub>1 = (ha, la, sha)" and "e' = throw a"
- and "P \<turnstile> C has F,Static:t in D" and "\<forall>sfs. sha D \<noteq> \<lfloor>(sfs, Done)\<rfloor>"
- and init': "P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(ha, la, sha)\<rangle> \<Rightarrow> \<langle>throw a,s'\<rangle>"
- then have i: "i = Processing" using iDP shD s\<^sub>1 by simp
- then show ?thesis using init' init_ProcessingE s\<^sub>1 s\<^sub>1a shD by blast
- next
- assume "\<forall>b t. \<not> P \<turnstile> C has F,b:t in D"
- then show ?thesis using SFAccInitRed.hyps(1) by blast
- next
- fix t assume field: "P \<turnstile> C has F,NonStatic:t in D"
- then show ?thesis using has_field_fun[OF SFAccInitRed.hyps(1) field] by simp
- qed
- next
- fix e assume e': "e' = throw e"
- and init: "P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(h, l, sh)\<rangle> \<Rightarrow> \<langle>throw e,s'\<rangle>"
- obtain h' l' sh' where s': "s' = (h',l',sh')" by(cases s')
- then obtain sfs i where shC: "sh' D = \<lfloor>(sfs, i)\<rfloor>" and iDP: "i = Error"
- using init_throw_PD[OF init] by auto
- then show ?thesis
- using SFAccInitRed.hyps(1) SFAccInitRed.hyps(2) SFAccInitThrow e' init by auto
- qed
-next
- case RedSFAccNone thus ?case
- by(fastforce elim: eval_cases intro: eval_evals.intros)
-next
- case RedSFAccNonStatic thus ?case
- by(fastforce elim: eval_cases intro: eval_evals.intros)
-next
- case (FAssRed1 e s b e1 s1 b1 F D e\<^sub>2)
- obtain h' l' sh' where "s'=(h',l',sh')" by(cases s')
- with FAssRed1 show ?case
- by(fastforce elim!: eval_cases(9)[where e\<^sub>1=e1] intro: eval_evals.intros simp: val_no_step
- intro!: FAss FAssNull FAssNone FAssStatic FAssThrow2)
-next
- case FAssRed2
- obtain h' l' sh' where "s'=(h',l',sh')" by(cases s')
- with FAssRed2 show ?case
- by(auto elim!: eval_cases intro: eval_evals.intros
- intro!: FAss FAssNull FAssNone FAssStatic FAssThrow2 Val)
-next
- case RedFAss
- thus ?case
- by (fastforce elim!: eval_cases intro: eval_evals.intros)
-next
- case RedFAssNull
- thus ?case
- by (fastforce elim!: eval_cases intro: eval_evals.intros)
-next
- case RedFAssNone
- then show ?case
- by(auto elim!: eval_cases intro: eval_evals.intros eval_finalId)
-next
- case RedFAssStatic
- then show ?case
- by(auto elim!: eval_cases intro: eval_evals.intros eval_finalId)
-next
- case (SFAssRed e s b e'' s'' b'' C F D)
- obtain h l sh where [simp]: "s = (h,l,sh)" by(cases s)
- obtain h' l' sh' where [simp]: "s'=(h',l',sh')" by(cases s')
- have "val_of e = None" using val_no_step SFAssRed.hyps(1) by(meson option.exhaust)
- then have bconf: "P,sh \<turnstile>\<^sub>b (e,b) \<surd>" using SFAssRed by simp
- show ?case using SFAssRed.prems(2) SFAssRed bconf
- proof cases
- case 2 with SFAssRed bconf show ?thesis by(auto intro!: SFAssInit)
- next
- case 3 with SFAssRed bconf show ?thesis by(auto intro!: SFAssInitThrow)
- qed(auto intro: eval_evals.intros intro!: SFAss SFAssInit SFAssNone SFAssNonStatic)
-next
- case (RedSFAss C F t D sh sfs i sfs' v sh' h l)
- let ?sfs' = "sfs(F \<mapsto> v)"
- have e':"e' = unit" and s':"s' = (h, l, sh(D \<mapsto> (?sfs', i)))"
- using RedSFAss eval_cases(3) by fastforce+
- have "i = Done \<or> i = Processing" using RedSFAss by(clarsimp simp: bconf_def initPD_def)
- then show ?case
- proof(cases i)
- case Done then show ?thesis using RedSFAss e' s' SFAss Val by auto
- next
- case Processing
- then have shC': "\<nexists>sfs. sh D = Some(sfs,Done)" and shP: "sh D = Some(sfs,Processing)"
- using RedSFAss by simp+
- then have init: "P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(h,l,sh)\<rangle> \<Rightarrow> \<langle>unit,(h,l,sh)\<rangle>"
- by(simp add: InitFinal InitProcessing Val)
- have "P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D} := Val v,(h, l, sh)\<rangle> \<Rightarrow> \<langle>unit,(h,l,sh(D \<mapsto> (?sfs', i)))\<rangle>"
- using Processing by(auto intro: SFAssInit[OF Val RedSFAss.hyps(1) shC' init shP])
- then show ?thesis using e' s' by simp
- qed(auto)
-next
- case (SFAssInitRed C F t D sh v h l)
- then have seq: "P \<turnstile> \<langle>(INIT D ([D],False) \<leftarrow> unit);; C\<bullet>\<^sub>sF{D} := Val v,(h, l, sh)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
- using eval_init_seq by simp
- then show ?case
- proof(rule eval_cases(14)) \<comment> \<open> Seq \<close>
- fix v' s\<^sub>1 assume init: "P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(h, l, sh)\<rangle> \<Rightarrow> \<langle>Val v',s\<^sub>1\<rangle>"
- and acc: "P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D} := Val v,s\<^sub>1\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
- obtain h\<^sub>1 l\<^sub>1 sh\<^sub>1 where s\<^sub>1: "s\<^sub>1 = (h\<^sub>1,l\<^sub>1,sh\<^sub>1)" by(cases s\<^sub>1)
- then obtain sfs i where shD: "sh\<^sub>1 D = \<lfloor>(sfs, i)\<rfloor>" and iDP: "i = Done \<or> i = Processing"
- using init_Val_PD[OF init] by auto
- show ?thesis
- proof(rule eval_cases(10)[OF acc]) \<comment> \<open> SFAss \<close>
- fix va h\<^sub>1 l\<^sub>1 sh\<^sub>1 t sfs
- assume e': "e' = unit" and s': "s' = (h\<^sub>1, l\<^sub>1, sh\<^sub>1(D \<mapsto> (sfs(F \<mapsto> va), Done)))"
- and val: "P \<turnstile> \<langle>Val v,s\<^sub>1\<rangle> \<Rightarrow> \<langle>Val va,(h\<^sub>1, l\<^sub>1, sh\<^sub>1)\<rangle>"
- and field: "P \<turnstile> C has F,Static:t in D" and shD': "sh\<^sub>1 D = \<lfloor>(sfs, Done)\<rfloor>"
- have "v = va" and "s\<^sub>1 = (h\<^sub>1, l\<^sub>1, sh\<^sub>1)" using eval_final_same[OF val] by auto
- then show ?thesis using SFAssInit field SFAssInitRed.hyps(2) shD' e' s' init val
- by (metis eval_final eval_finalId)
- next
- fix va h\<^sub>1 l\<^sub>1 sh\<^sub>1 t v' h' l' sh' sfs i'
- assume e': "e' = unit" and s': "s' = (h', l', sh'(D \<mapsto> (sfs(F \<mapsto> va), i')))"
- and val: "P \<turnstile> \<langle>Val v,s\<^sub>1\<rangle> \<Rightarrow> \<langle>Val va,(h\<^sub>1, l\<^sub>1, sh\<^sub>1)\<rangle>"
- and field: "P \<turnstile> C has F,Static:t in D" and nDone: "\<forall>sfs. sh\<^sub>1 D \<noteq> \<lfloor>(sfs, Done)\<rfloor>"
- and init': "P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(h\<^sub>1, l\<^sub>1, sh\<^sub>1)\<rangle> \<Rightarrow> \<langle>Val v',(h', l', sh')\<rangle>"
- and shD': "sh' D = \<lfloor>(sfs, i')\<rfloor>"
- have v: "v = va" and s\<^sub>1a: "s\<^sub>1 = (h\<^sub>1, l\<^sub>1, sh\<^sub>1)" using eval_final_same[OF val] by auto
- then have i: "i = Processing" using iDP shD s\<^sub>1 nDone by simp
- then have "(h\<^sub>1, l\<^sub>1, sh\<^sub>1) = (h', l', sh')" using init' init_ProcessingE s\<^sub>1 s\<^sub>1a shD by blast
- then show ?thesis using SFAssInit SFAssInitRed.hyps(2) e' s' field init v s\<^sub>1a shD' val
- by (metis eval_final eval_finalId)
- next
- fix va h\<^sub>1 l\<^sub>1 sh\<^sub>1 t a
- assume "e' = throw a" and val: "P \<turnstile> \<langle>Val v,s\<^sub>1\<rangle> \<Rightarrow> \<langle>Val va,(h\<^sub>1, l\<^sub>1, sh\<^sub>1)\<rangle>"
- and "P \<turnstile> C has F,Static:t in D" and nDone: "\<forall>sfs. sh\<^sub>1 D \<noteq> \<lfloor>(sfs, Done)\<rfloor>"
- and init': "P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(h\<^sub>1, l\<^sub>1, sh\<^sub>1)\<rangle> \<Rightarrow> \<langle>throw a,s'\<rangle>"
- have v: "v = va" and s\<^sub>1a: "s\<^sub>1 = (h\<^sub>1, l\<^sub>1, sh\<^sub>1)" using eval_final_same[OF val] by auto
- then have i: "i = Processing" using iDP shD s\<^sub>1 nDone by simp
- then have "(h\<^sub>1, l\<^sub>1, sh\<^sub>1) = s'" using init' init_ProcessingE s\<^sub>1 s\<^sub>1a shD by blast
- then show ?thesis using init' init_ProcessingE s\<^sub>1 s\<^sub>1a shD i by blast
- next
- fix e'' assume val:"P \<turnstile> \<langle>Val v,s\<^sub>1\<rangle> \<Rightarrow> \<langle>throw e'',s'\<rangle>"
- then show ?thesis using eval_final_same[OF val] by simp
- next
- assume "\<forall>b t. \<not> P \<turnstile> C has F,b:t in D"
- then show ?thesis using SFAssInitRed.hyps(1) by blast
- next
- fix t assume field: "P \<turnstile> C has F,NonStatic:t in D"
- then show ?thesis using has_field_fun[OF SFAssInitRed.hyps(1) field] by simp
- qed
- next
- fix e assume e': "e' = throw e"
- and init: "P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(h, l, sh)\<rangle> \<Rightarrow> \<langle>throw e,s'\<rangle>"
- obtain h' l' sh' where s': "s' = (h',l',sh')" by(cases s')
- then obtain sfs i where shC: "sh' D = \<lfloor>(sfs, i)\<rfloor>" and iDP: "i = Error"
- using init_throw_PD[OF init] by auto
- then show ?thesis using SFAssInitRed.hyps(1) SFAssInitRed.hyps(2) SFAssInitThrow Val
- by (metis e' init)
- qed
-next
- case (RedSFAssNone C F D v s b) then show ?case
- by(cases s) (auto elim!: eval_cases intro: eval_evals.intros eval_finalId)
-next
- case (RedSFAssNonStatic C F t D v s b) then show ?case
- by(cases s) (auto elim!: eval_cases intro: eval_evals.intros eval_finalId)
-next
- case CallObj
- note val_no_step[simp]
- from CallObj.prems(2) CallObj show ?case
- proof cases
- case 2 with CallObj show ?thesis by(fastforce intro!: CallParamsThrow)
- next
- case 3 with CallObj show ?thesis by(fastforce intro!: CallNull)
- next
- case 4 with CallObj show ?thesis by(fastforce intro!: CallNone)
- next
- case 5 with CallObj show ?thesis by(fastforce intro!: CallStatic)
- qed(fastforce intro!: CallObjThrow Call)+
-next
- case (CallParams es s b es'' s'' b'' v M s')
- then obtain h' l' sh' where "s' = (h',l',sh')" by(cases s')
- with CallParams show ?case
- by(auto elim!: eval_cases intro!: CallNone eval_finalId CallStatic Val)
- (auto intro!: CallParamsThrow CallNull Call Val)
-next
- case (RedCall h a C fs M Ts T pns body D vs l sh b)
- have "P \<turnstile> \<langle>addr a,(h,l,sh)\<rangle> \<Rightarrow> \<langle>addr a,(h,l,sh)\<rangle>" by (rule eval_evals.intros)
- moreover
- have finals: "finals(map Val vs)" by simp
- with finals have "P \<turnstile> \<langle>map Val vs,(h,l,sh)\<rangle> [\<Rightarrow>] \<langle>map Val vs,(h,l,sh)\<rangle>"
- by (iprover intro: eval_finalsId)
- moreover have "h a = Some (C, fs)" using RedCall by simp
- moreover have "method": "P \<turnstile> C sees M,NonStatic: Ts\<rightarrow>T = (pns, body) in D" by fact
- moreover have same_len\<^sub>1: "length Ts = length pns"
- and this_distinct: "this \<notin> set pns" and fv: "fv (body) \<subseteq> {this} \<union> set pns"
- using "method" wf by (fastforce dest!:sees_wf_mdecl simp:wf_mdecl_def)+
- have same_len: "length vs = length pns" by fact
- moreover
- obtain l\<^sub>2' where l\<^sub>2': "l\<^sub>2' = [this\<mapsto>Addr a,pns[\<mapsto>]vs]" by simp
- moreover
- obtain h\<^sub>3 l\<^sub>3 sh\<^sub>3 where s': "s' = (h\<^sub>3,l\<^sub>3,sh\<^sub>3)" by (cases s')
- have eval_blocks:
- "P \<turnstile> \<langle>(blocks (this # pns, Class D # Ts, Addr a # vs, body)),(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>" by fact
- hence id: "l\<^sub>3 = l" using fv s' same_len\<^sub>1 same_len
- by(fastforce elim: eval_closed_lcl_unchanged)
- from eval_blocks obtain l\<^sub>3' where "P \<turnstile> \<langle>body,(h,l\<^sub>2',sh)\<rangle> \<Rightarrow> \<langle>e',(h\<^sub>3,l\<^sub>3',sh\<^sub>3)\<rangle>"
- proof -
- from same_len\<^sub>1 have "length(this#pns) = length(Class D#Ts)" by simp
- moreover from same_len\<^sub>1 same_len
- have same_len\<^sub>2: "length (this#pns) = length (Addr a#vs)" by simp
- moreover from eval_blocks
- have "P \<turnstile> \<langle>blocks (this#pns,Class D#Ts,Addr a#vs,body),(h,l,sh)\<rangle>
- \<Rightarrow>\<langle>e',(h\<^sub>3,l\<^sub>3,sh\<^sub>3)\<rangle>" using s' same_len\<^sub>1 same_len\<^sub>2 by simp
- ultimately obtain l''
- where "P \<turnstile> \<langle>body,(h,l(this # pns[\<mapsto>]Addr a # vs),sh)\<rangle>\<Rightarrow>\<langle>e',(h\<^sub>3, l'',sh\<^sub>3)\<rangle>"
- by (blast dest:blocksEval)
- from eval_restrict_lcl[OF wf this fv] this_distinct same_len\<^sub>1 same_len
- have "P \<turnstile> \<langle>body,(h,[this # pns[\<mapsto>]Addr a # vs],sh)\<rangle> \<Rightarrow>
- \<langle>e',(h\<^sub>3, l''|`(set(this#pns)),sh\<^sub>3)\<rangle>" using wf method
- by(simp add:subset_insert_iff insert_Diff_if)
- thus ?thesis by(fastforce intro!:that simp add: l\<^sub>2')
- qed
- ultimately
- have "P \<turnstile> \<langle>(addr a)\<bullet>M(map Val vs),(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',(h\<^sub>3,l,sh\<^sub>3)\<rangle>" by (rule Call)
- with s' id show ?case by simp
-next
- case RedCallNull
- thus ?case
- by (fastforce elim: eval_cases intro: eval_evals.intros eval_finalsId)
-next
- case (RedCallNone h a C fs M vs l sh b)
- then have tes: "THROW NoSuchMethodError = e' \<and> (h,l,sh) = s'"
- using eval_final_same by simp
- have "P \<turnstile> \<langle>addr a,(h,l,sh)\<rangle> \<Rightarrow> \<langle>addr a,(h,l,sh)\<rangle>" and "P \<turnstile> \<langle>map Val vs,(h,l,sh)\<rangle> [\<Rightarrow>] \<langle>map Val vs,(h,l,sh)\<rangle>"
- using eval_finalId eval_finalsId by auto
- then show ?case using RedCallNone CallNone tes by auto
-next
- case (RedCallStatic h a C fs M Ts T m D vs l sh b)
- then have tes: "THROW IncompatibleClassChangeError = e' \<and> (h,l,sh) = s'"
- using eval_final_same by simp
- have "P \<turnstile> \<langle>addr a,(h,l,sh)\<rangle> \<Rightarrow> \<langle>addr a,(h,l,sh)\<rangle>" and "P \<turnstile> \<langle>map Val vs,(h,l,sh)\<rangle> [\<Rightarrow>] \<langle>map Val vs,(h,l,sh)\<rangle>"
- using eval_finalId eval_finalsId by auto
- then show ?case using RedCallStatic CallStatic tes by fastforce
-next
- case (SCallParams es s b es'' s'' b' C M s')
- obtain h' l' sh' where s'[simp]: "s' = (h',l',sh')" by(cases s')
- obtain h l sh where s[simp]: "s = (h,l,sh)" by(cases s)
- have es: "map_vals_of es = None" using vals_no_step SCallParams.hyps(1) by (meson not_Some_eq)
- have bconf: "P,sh \<turnstile>\<^sub>b (es,b) \<surd>" using s SCallParams.prems(1) by (simp add: bconf_SCall[OF es])
- from SCallParams.prems(2) SCallParams bconf show ?case
- proof cases
- case 2 with SCallParams bconf show ?thesis by(auto intro!: SCallNone)
- next
- case 3 with SCallParams bconf show ?thesis by(auto intro!: SCallNonStatic)
- next
- case 4 with SCallParams bconf show ?thesis by(auto intro!: SCallInitThrow)
- next
- case 5 with SCallParams bconf show ?thesis by(auto intro!: SCallInit)
- qed(auto intro!: SCallParamsThrow SCall)
-next
- case (RedSCall C M Ts T pns body D vs s)
- then obtain h l sh where s:"s = (h,l,sh)" by(cases s)
- then obtain sfs i where shC: "sh D = \<lfloor>(sfs, i)\<rfloor>" and "i = Done \<or> i = Processing"
- using RedSCall by(auto simp: bconf_def initPD_def dest: sees_method_fun)
- have finals: "finals(map Val vs)" by simp
- with finals have mVs: "P \<turnstile> \<langle>map Val vs,(h,l,sh)\<rangle> [\<Rightarrow>] \<langle>map Val vs,(h,l,sh)\<rangle>"
- by (iprover intro: eval_finalsId)
- obtain sfs i where shC: "sh D = \<lfloor>(sfs, i)\<rfloor>"
- using RedSCall s by(auto simp: bconf_def initPD_def dest: sees_method_fun)
- then have iDP: "i = Done \<or> i = Processing" using RedSCall s
- by (auto simp: bconf_def initPD_def dest: sees_method_fun[OF RedSCall.hyps(1)])
- have "method": "P \<turnstile> C sees M,Static: Ts\<rightarrow>T = (pns, body) in D" by fact
- have same_len\<^sub>1: "length Ts = length pns" and fv: "fv (body) \<subseteq> set pns"
- using "method" wf by (fastforce dest!:sees_wf_mdecl simp:wf_mdecl_def)+
- have same_len: "length vs = length pns" by fact
- obtain l\<^sub>2' where l\<^sub>2': "l\<^sub>2' = [pns[\<mapsto>]vs]" by simp
- obtain h\<^sub>3 l\<^sub>3 sh\<^sub>3 where s': "s' = (h\<^sub>3,l\<^sub>3,sh\<^sub>3)" by (cases s')
- have eval_blocks:
- "P \<turnstile> \<langle>(blocks (pns, Ts, vs, body)),(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>" using RedSCall.prems(2) s by simp
- hence id: "l\<^sub>3 = l" using fv s' same_len\<^sub>1 same_len
- by(fastforce elim: eval_closed_lcl_unchanged)
- from eval_blocks obtain l\<^sub>3' where body: "P \<turnstile> \<langle>body,(h,l\<^sub>2',sh)\<rangle> \<Rightarrow> \<langle>e',(h\<^sub>3,l\<^sub>3',sh\<^sub>3)\<rangle>"
- proof -
- from eval_blocks
- have "P \<turnstile> \<langle>blocks (pns,Ts,vs,body),(h,l,sh)\<rangle>
- \<Rightarrow>\<langle>e',(h\<^sub>3,l\<^sub>3,sh\<^sub>3)\<rangle>" using s' same_len same_len\<^sub>1 by simp
- then obtain l''
- where "P \<turnstile> \<langle>body,(h,l(pns[\<mapsto>]vs),sh)\<rangle> \<Rightarrow> \<langle>e',(h\<^sub>3, l'',sh\<^sub>3)\<rangle>"
- by (blast dest:blocksEval[OF same_len\<^sub>1[THEN sym] same_len[THEN sym]])
- from eval_restrict_lcl[OF wf this fv] same_len\<^sub>1 same_len
- have "P \<turnstile> \<langle>body,(h,[pns[\<mapsto>]vs],sh)\<rangle> \<Rightarrow> \<langle>e',(h\<^sub>3, l''|`(set(pns)),sh\<^sub>3)\<rangle>" using wf method
- by(simp add:subset_insert_iff insert_Diff_if)
- thus ?thesis by(fastforce intro!:that simp add: l\<^sub>2')
- qed
- show ?case using iDP
- proof(cases i)
- case Done
- then have shC': "sh D = \<lfloor>(sfs, Done)\<rfloor> \<or> M = clinit \<and> sh D = \<lfloor>(sfs, Processing)\<rfloor>"
- using shC by simp
- have "P \<turnstile> \<langle>C\<bullet>\<^sub>sM(map Val vs),(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',(h\<^sub>3,l,sh\<^sub>3)\<rangle>"
- by (rule SCall[OF mVs method shC' same_len l\<^sub>2' body])
- with s s' id show ?thesis by simp
- next
- case Processing
- then have shC': "\<nexists>sfs. sh D = Some(sfs,Done)" and shP: "sh D = Some(sfs,Processing)"
- using shC by simp+
- then have init: "P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(h,l,sh)\<rangle> \<Rightarrow> \<langle>unit,(h,l,sh)\<rangle>"
- by(simp add: InitFinal InitProcessing Val)
- have "P \<turnstile> \<langle>C\<bullet>\<^sub>sM(map Val vs),(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',(h\<^sub>3,l,sh\<^sub>3)\<rangle>"
- proof(cases "M = clinit")
- case False show ?thesis by(rule SCallInit[OF mVs method shC' False init same_len l\<^sub>2' body])
- next
- case True
- then have shC': "sh D = \<lfloor>(sfs, Done)\<rfloor> \<or> M = clinit \<and> sh D = \<lfloor>(sfs, Processing)\<rfloor>"
- using shC Processing by simp
- have "P \<turnstile> \<langle>C\<bullet>\<^sub>sM(map Val vs),(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',(h\<^sub>3,l,sh\<^sub>3)\<rangle>"
- by (rule SCall[OF mVs method shC' same_len l\<^sub>2' body])
- with s s' id show ?thesis by simp
- qed
- with s s' id show ?thesis by simp
- qed(auto)
-next
- case (SCallInitRed C M Ts T pns body D sh vs h l)
- then have seq: "P \<turnstile> \<langle>(INIT D ([D],False) \<leftarrow> unit);; C\<bullet>\<^sub>sM(map Val vs),(h, l, sh)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
- using eval_init_seq by simp
- then show ?case
- proof(rule eval_cases(14)) \<comment> \<open> Seq \<close>
- fix v' s\<^sub>1 assume init: "P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(h, l, sh)\<rangle> \<Rightarrow> \<langle>Val v',s\<^sub>1\<rangle>"
- and call: "P \<turnstile> \<langle>C\<bullet>\<^sub>sM(map Val vs),s\<^sub>1\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
- obtain h\<^sub>1 l\<^sub>1 sh\<^sub>1 where s\<^sub>1: "s\<^sub>1 = (h\<^sub>1,l\<^sub>1,sh\<^sub>1)" by(cases s\<^sub>1)
- then obtain sfs i where shD: "sh\<^sub>1 D = \<lfloor>(sfs, i)\<rfloor>" and iDP: "i = Done \<or> i = Processing"
- using init_Val_PD[OF init] by auto
- show ?thesis
- proof(rule eval_cases(12)[OF call]) \<comment> \<open> SCall \<close>
- fix vsa ex es' assume "P \<turnstile> \<langle>map Val vs,s\<^sub>1\<rangle> [\<Rightarrow>] \<langle>map Val vsa @ throw ex # es',s'\<rangle>"
- then show ?thesis using evals_finals_same by (meson finals_def map_Val_nthrow_neq)
- next
- assume "\<forall>b Ts T a ba x. \<not> P \<turnstile> C sees M, b : Ts\<rightarrow>T = (a, ba) in x"
- then show ?thesis using SCallInitRed.hyps(1) by auto
- next
- fix Ts T m D assume "P \<turnstile> C sees M, NonStatic : Ts\<rightarrow>T = m in D"
- then show ?thesis using sees_method_fun[OF SCallInitRed.hyps(1)] by blast
- next
- fix vsa h1 l1 sh1 Ts T pns body D' a
- assume "e' = throw a" and vals: "P \<turnstile> \<langle>map Val vs,s\<^sub>1\<rangle> [\<Rightarrow>] \<langle>map Val vsa,(h1, l1, sh1)\<rangle>"
- and method: "P \<turnstile> C sees M, Static : Ts\<rightarrow>T = (pns, body) in D'"
- and nDone: "\<forall>sfs. sh1 D' \<noteq> \<lfloor>(sfs, Done)\<rfloor>"
- and init': "P \<turnstile> \<langle>INIT D' ([D'],False) \<leftarrow> unit,(h1, l1, sh1)\<rangle> \<Rightarrow> \<langle>throw a,s'\<rangle>"
- have vs: "vs = vsa" and s\<^sub>1a: "s\<^sub>1 = (h1, l1, sh1)"
- using evals_finals_same[OF _ vals] map_Val_eq by auto
- have D: "D = D'" using sees_method_fun[OF SCallInitRed.hyps(1) method] by simp
- then have i: "i = Processing" using iDP shD s\<^sub>1 s\<^sub>1a nDone by simp
- then show ?thesis using D init' init_ProcessingE s\<^sub>1 s\<^sub>1a shD by blast
- next
- fix vsa h1 l1 sh1 Ts T pns' body' D' v' h\<^sub>2 l\<^sub>2 sh\<^sub>2 h\<^sub>3 l\<^sub>3 sh\<^sub>3
- assume s': "s' = (h\<^sub>3, l\<^sub>2, sh\<^sub>3)"
- and vals: "P \<turnstile> \<langle>map Val vs,s\<^sub>1\<rangle> [\<Rightarrow>] \<langle>map Val vsa,(h1, l1, sh1)\<rangle>"
- and method: "P \<turnstile> C sees M, Static : Ts\<rightarrow>T = (pns', body') in D'"
- and nDone: "\<forall>sfs. sh1 D' \<noteq> \<lfloor>(sfs, Done)\<rfloor>"
- and init': "P \<turnstile> \<langle>INIT D' ([D'],False) \<leftarrow> unit,(h1, l1, sh1)\<rangle> \<Rightarrow> \<langle>Val v',(h\<^sub>2, l\<^sub>2, sh\<^sub>2)\<rangle>"
- and len: "length vsa = length pns'"
- and bstep: "P \<turnstile> \<langle>body',(h\<^sub>2, [pns' [\<mapsto>] vsa], sh\<^sub>2)\<rangle> \<Rightarrow> \<langle>e',(h\<^sub>3, l\<^sub>3, sh\<^sub>3)\<rangle>"
- have vs: "vs = vsa" and s\<^sub>1a: "s\<^sub>1 = (h1, l1, sh1)"
- using evals_finals_same[OF _ vals] map_Val_eq by auto
- have D: "D = D'" and pns: "pns = pns'" and body: "body = body'"
- using sees_method_fun[OF SCallInitRed.hyps(1) method] by auto
- then have i: "i = Processing" using iDP shD s\<^sub>1 s\<^sub>1a nDone by simp
- then have s2: "(h\<^sub>2, l\<^sub>2, sh\<^sub>2) = s\<^sub>1" using D init' init_ProcessingE s\<^sub>1 s\<^sub>1a shD by blast
- then show ?thesis
- using eval_finalId SCallInit[OF eval_finalsId[of "map Val vs" P "(h,l,sh)"]
- SCallInitRed.hyps(1)] init init' len bstep nDone D pns body s' s\<^sub>1 s\<^sub>1a shD vals vs
- SCallInitRed.hyps(2-3) s2 by auto
- next
- fix vsa h\<^sub>2 l\<^sub>2 sh\<^sub>2 Ts T pns' body' D' sfs h\<^sub>3 l\<^sub>3 sh\<^sub>3
- assume s': "s' = (h\<^sub>3, l\<^sub>2, sh\<^sub>3)" and vals: "P \<turnstile> \<langle>map Val vs,s\<^sub>1\<rangle> [\<Rightarrow>] \<langle>map Val vsa,(h\<^sub>2, l\<^sub>2, sh\<^sub>2)\<rangle>"
- and method: "P \<turnstile> C sees M, Static : Ts\<rightarrow>T = (pns', body') in D'"
- and "sh\<^sub>2 D' = \<lfloor>(sfs, Done)\<rfloor> \<or> M = clinit \<and> sh\<^sub>2 D' = \<lfloor>(sfs, Processing)\<rfloor>"
- and len: "length vsa = length pns'"
- and bstep: "P \<turnstile> \<langle>body',(h\<^sub>2, [pns' [\<mapsto>] vsa], sh\<^sub>2)\<rangle> \<Rightarrow> \<langle>e',(h\<^sub>3, l\<^sub>3, sh\<^sub>3)\<rangle>"
- have vs: "vs = vsa" and s\<^sub>1a: "s\<^sub>1 = (h\<^sub>2, l\<^sub>2, sh\<^sub>2)"
- using evals_finals_same[OF _ vals] map_Val_eq by auto
- have D: "D = D'" and pns: "pns = pns'" and body: "body = body'"
- using sees_method_fun[OF SCallInitRed.hyps(1) method] by auto
- then show ?thesis using SCallInit[OF eval_finalsId[of "map Val vs" P "(h,l,sh)"]
- SCallInitRed.hyps(1)] bstep SCallInitRed.hyps(2-3) len s' s\<^sub>1a vals vs init by auto
- qed
- next
- fix e assume e': "e' = throw e"
- and init: "P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(h, l, sh)\<rangle> \<Rightarrow> \<langle>throw e,s'\<rangle>"
- obtain h' l' sh' where s': "s' = (h',l',sh')" by(cases s')
- then obtain sfs i where shC: "sh' D = \<lfloor>(sfs, i)\<rfloor>" and iDP: "i = Error"
- using init_throw_PD[OF init] by auto
- then show ?thesis using SCallInitRed.hyps(2-3) init e'
- SCallInitThrow[OF eval_finalsId[of "map Val vs" _] SCallInitRed.hyps(1)]
- by auto
- qed
-next
- case (RedSCallNone C M vs s b)
- then have tes: "THROW NoSuchMethodError = e' \<and> s = s'"
- using eval_final_same by simp
- have "P \<turnstile> \<langle>map Val vs,s\<rangle> [\<Rightarrow>] \<langle>map Val vs,s\<rangle>" using eval_finalsId by simp
- then show ?case using RedSCallNone eval_evals.SCallNone tes by auto
-next
- case (RedSCallNonStatic C M Ts T m D vs s b)
- then have tes: "THROW IncompatibleClassChangeError = e' \<and> s = s'"
- using eval_final_same by simp
- have "P \<turnstile> \<langle>map Val vs,s\<rangle> [\<Rightarrow>] \<langle>map Val vs,s\<rangle>" using eval_finalsId by simp
- then show ?case using RedSCallNonStatic eval_evals.SCallNonStatic tes by auto
-next
- case InitBlockRed
- thus ?case
- by (fastforce elim!: eval_cases intro: eval_evals.intros eval_finalId
- simp: assigned_def map_upd_triv fun_upd_same)
-next
- case (RedInitBlock V T v u s b)
- then have "P \<turnstile> \<langle>Val u,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>" by simp
- then obtain s': "s'=s" and e': "e'=Val u" by cases simp
- obtain h l sh where s: "s=(h,l,sh)" by (cases s)
- have "P \<turnstile> \<langle>{V:T :=Val v; Val u},(h,l,sh)\<rangle> \<Rightarrow> \<langle>Val u,(h, (l(V\<mapsto>v))(V:=l V), sh)\<rangle>"
- by (fastforce intro!: eval_evals.intros)
- then have "P \<turnstile> \<langle>{V:T := Val v; Val u},s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>" using s s' e' by simp
- then show ?case by simp
-next
- case BlockRedNone
- thus ?case
- by (fastforce elim!: eval_cases intro: eval_evals.intros
- simp add: fun_upd_same fun_upd_idem)
-next
- case BlockRedSome
- thus ?case
- by (fastforce elim!: eval_cases intro: eval_evals.intros
- simp add: fun_upd_same fun_upd_idem)
-next
- case (RedBlock V T v s b)
- then have "P \<turnstile> \<langle>Val v,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>" by simp
- then obtain s': "s'=s" and e': "e'=Val v"
- by cases simp
- obtain h l sh where s: "s=(h,l,sh)" by (cases s)
- have "P \<turnstile> \<langle>Val v,(h,l(V:=None),sh)\<rangle> \<Rightarrow> \<langle>Val v,(h,l(V:=None),sh)\<rangle>"
- by (rule eval_evals.intros)
- hence "P \<turnstile> \<langle>{V:T;Val v},(h,l,sh)\<rangle> \<Rightarrow> \<langle>Val v,(h,(l(V:=None))(V:=l V),sh)\<rangle>"
- by (rule eval_evals.Block)
- then have "P \<turnstile> \<langle>{V:T; Val v},s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>" using s s' e' by simp
- then show ?case by simp
-next
- case (SeqRed e s b e1 s1 b1 e\<^sub>2) show ?case
- proof(cases "val_of e")
- case None show ?thesis
- proof(cases "lass_val_of e")
- case lNone:None
- then have bconf: "P,shp s \<turnstile>\<^sub>b (e,b) \<surd>" using SeqRed.prems(1) None by simp
- then show ?thesis using SeqRed using seq_ext by fastforce
- next
- case (Some p)
- obtain V' v' where p: "p = (V',v')" and e: "e = V':=Val v'"
- using lass_val_of_spec[OF Some] by(cases p, auto)
- obtain h l sh h' l' sh' where s: "s = (h,l,sh)" and s1: "s1 = (h',l',sh')" by(cases s, cases s1)
- then have red: "P \<turnstile> \<langle>e,(h,l,sh),b\<rangle> \<rightarrow> \<langle>e1,(h',l',sh'),b1\<rangle>" using SeqRed.hyps(1) by simp
- then have s\<^sub>1': "e1 = unit \<and> h' = h \<and> l' = l(V' \<mapsto> v') \<and> sh' = sh"
- using lass_val_of_red[OF Some red] p e by simp
- then have eval: "P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>e1,s1\<rangle>" using e s s1 LAss Val by auto
- then show ?thesis
- by (metis SeqRed.prems(2) eval_final eval_final_same seq_ext)
- qed
- next
- case (Some a) then show ?thesis using SeqRed.hyps(1) val_no_step by blast
- qed
-next
- case RedSeq
- thus ?case
- by (fastforce elim: eval_cases intro: eval_evals.intros)
-next
- case CondRed
- thus ?case
- by (fastforce elim: eval_cases intro: eval_evals.intros simp: val_no_step)
-next
- case RedCondT
- thus ?case
- by (fastforce elim: eval_cases intro: eval_evals.intros)
-next
- case RedCondF
- thus ?case
- by (fastforce elim: eval_cases intro: eval_evals.intros)
-next
- case RedWhile
- thus ?case
- by (auto simp add: unfold_while intro:eval_evals.intros elim:eval_cases)
-next
- case ThrowRed then show ?case by(fastforce elim: eval_cases simp: eval_evals.intros)
-next
- case RedThrowNull
- thus ?case
- by (fastforce elim: eval_cases intro: eval_evals.intros)
-next
- case TryRed thus ?case
- by(fastforce elim: eval_cases intro: eval_evals.intros)
-next
- case RedTryCatch
- thus ?case
- by (fastforce elim: eval_cases intro: eval_evals.intros)
-next
- case (RedTryFail s a D fs C V e\<^sub>2 b)
- thus ?case
- by (cases s)(auto elim!: eval_cases intro: eval_evals.intros)
-next
- case ListRed1
- thus ?case
- by (fastforce elim: evals_cases intro: eval_evals.intros simp: val_no_step)
-next
- case ListRed2
- thus ?case
- by (fastforce elim!: evals_cases eval_cases
- intro: eval_evals.intros eval_finalId)
-next
- case (RedInit e1 C b s1 b') then show ?case using InitFinal by simp
-next
- case (InitNoneRed sh C C' Cs e h l b)
- show ?case using InitNone InitNoneRed.hyps InitNoneRed.prems(2) by auto
-next
- case (RedInitDone sh C sfs C' Cs e h l b)
- show ?case using InitDone RedInitDone.hyps RedInitDone.prems(2) by auto
-next
- case (RedInitProcessing sh C sfs C' Cs e h l b)
- show ?case using InitProcessing RedInitProcessing.hyps RedInitProcessing.prems(2) by auto
-next
- case (RedInitError sh C sfs C' Cs e h l b)
- show ?case using InitError RedInitError.hyps RedInitError.prems(2) by auto
-next
- case (InitObjectRed sh C sfs sh' C' Cs e h l b) show ?case using InitObject InitObjectRed by auto
-next
- case (InitNonObjectSuperRed sh C sfs D r sh' C' Cs e h l b)
- show ?case using InitNonObject InitNonObjectSuperRed by auto
-next
- case (RedInitRInit C' C Cs e h l sh b)
- show ?case using InitRInit RedInitRInit by auto
-next
- case (RInitRed e s b e'' s'' b'' C Cs e\<^sub>0)
- then have IH: "\<And>e' s'. P \<turnstile> \<langle>e'',s''\<rangle> \<Rightarrow> \<langle>e',s'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>" by simp
- show ?case using RInitRed rinit_ext[OF IH] by simp
-next
- case (RedRInit sh C sfs i sh' C' Cs v e h l b s' e')
- then have init: "P \<turnstile> \<langle>(INIT C' (Cs,True) \<leftarrow> e), (h, l, sh(C \<mapsto> (sfs, Done)))\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
- using RedRInit by simp
- then show ?case using RInit RedRInit.hyps(1) RedRInit.hyps(3) Val by fastforce
-next
- case BinOpThrow2
- thus ?case
- by (fastforce elim: eval_cases intro: eval_evals.intros)
-next
- case FAssThrow2
- thus ?case
- by (fastforce elim: eval_cases intro: eval_evals.intros)
-next
- case SFAssThrow
- then show ?case
- by (fastforce elim: eval_cases intro: eval_evals.intros)
-next
- case (CallThrowParams es vs e es' v M s b)
- have val: "P \<turnstile> \<langle>Val v,s\<rangle> \<Rightarrow> \<langle>Val v,s\<rangle>" by (rule eval_evals.intros)
- have eval_e: "P \<turnstile> \<langle>throw (e),s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>" using CallThrowParams by simp
- then obtain xa where e': "e' = Throw xa" by (cases) (auto dest!: eval_final)
- with list_eval_Throw [OF eval_e]
- have vals: "P \<turnstile> \<langle>es,s\<rangle> [\<Rightarrow>] \<langle>map Val vs @ Throw xa # es',s'\<rangle>"
- using CallThrowParams.hyps(1) eval_e list_eval_Throw by blast
- then have "P \<turnstile> \<langle>Val v\<bullet>M(es),s\<rangle> \<Rightarrow> \<langle>Throw xa,s'\<rangle>"
- using eval_evals.CallParamsThrow[OF val vals] by simp
- thus ?case using e' by simp
-next
- case (SCallThrowParams es vs e es' C M s b)
- have eval_e: "P \<turnstile> \<langle>throw (e),s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>" using SCallThrowParams by simp
- then obtain xa where e': "e' = Throw xa" by (cases) (auto dest!: eval_final)
- then have "P \<turnstile> \<langle>es,s\<rangle> [\<Rightarrow>] \<langle>map Val vs @ Throw xa # es',s'\<rangle>"
- using SCallThrowParams.hyps(1) eval_e list_eval_Throw by blast
- then have "P \<turnstile> \<langle>C\<bullet>\<^sub>sM(es),s\<rangle> \<Rightarrow> \<langle>Throw xa,s'\<rangle>"
- by (rule eval_evals.SCallParamsThrow)
- thus ?case using e' by simp
-next
- case (BlockThrow V T a s b)
- then have "P \<turnstile> \<langle>Throw a, s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>" by simp
- then obtain s': "s' = s" and e': "e' = Throw a"
- by cases (auto elim!:eval_cases)
- obtain h l sh where s: "s=(h,l,sh)" by (cases s)
- have "P \<turnstile> \<langle>Throw a, (h,l(V:=None),sh)\<rangle> \<Rightarrow> \<langle>Throw a, (h,l(V:=None),sh)\<rangle>"
- by (simp add:eval_evals.intros eval_finalId)
- hence "P\<turnstile>\<langle>{V:T;Throw a},(h,l,sh)\<rangle> \<Rightarrow> \<langle>Throw a, (h,(l(V:=None))(V:=l V),sh)\<rangle>"
- by (rule eval_evals.Block)
- then have "P \<turnstile> \<langle>{V:T; Throw a},s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>" using s s' e' by simp
- then show ?case by simp
-next
- case (InitBlockThrow V T v a s b)
- then have "P \<turnstile> \<langle>Throw a,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>" by simp
- then obtain s': "s' = s" and e': "e' = Throw a"
- by cases (auto elim!:eval_cases)
- obtain h l sh where s: "s = (h,l,sh)" by (cases s)
- have "P \<turnstile> \<langle>{V:T :=Val v; Throw a},(h,l,sh)\<rangle> \<Rightarrow> \<langle>Throw a, (h, (l(V\<mapsto>v))(V:=l V),sh)\<rangle>"
- by(fastforce intro:eval_evals.intros)
- then have "P \<turnstile> \<langle>{V:T := Val v; Throw a},s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>" using s s' e' by simp
- then show ?case by simp
-next
- case (RInitInitThrow sh C sfs i sh' a D Cs e h l b)
- have IH: "\<And>e' s'. P \<turnstile> \<langle>RI (D,Throw a) ; Cs \<leftarrow> e,(h, l, sh(C \<mapsto> (sfs, Error)))\<rangle> \<Rightarrow> \<langle>e',s'\<rangle> \<Longrightarrow>
- P \<turnstile> \<langle>RI (C,Throw a) ; D # Cs \<leftarrow> e,(h, l, sh)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
- using RInitInitFail[OF eval_finalId] RInitInitThrow by simp
- then show ?case using RInitInitThrow.hyps(2) RInitInitThrow.prems(2) by auto
-next
- case (RInitThrow sh C sfs i sh' a e h l b)
- then have e': "e' = Throw a" and s': "s' = (h,l,sh')"
- using eval_final_same final_def by fastforce+
- show ?case using RInitFailFinal RInitThrow.hyps(1) RInitThrow.hyps(2) e' eval_finalId s' by auto
-qed(auto elim: eval_cases simp: eval_evals.intros)
-(*>*)
-
-(*<*)
-(* ... und wieder anschalten: *)
-declare split_paired_All [simp] split_paired_Ex [simp]
-(*>*)
-
-text \<open> Its extension to @{text"\<rightarrow>*"}: \<close>
-
-lemma extend_eval:
-assumes wf: "wwf_J_prog P"
-shows "\<lbrakk> P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e'',s'',b''\<rangle>; P \<turnstile> \<langle>e'',s''\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>;
- iconf (shp s) e; P,shp s \<turnstile>\<^sub>b (e::expr,b) \<surd> \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
-(*<*)
-proof (induct rule: converse_rtrancl_induct3)
- case refl then show ?case by simp
-next
- case (step e1 s1 b1 e2 s2 b2)
- then have ic: "iconf (shp s2) e2" using Red_preserves_iconf local.wf by blast
- then have ec: "P,shp s2 \<turnstile>\<^sub>b (e2,b2) \<surd>"
- using Red_preserves_bconf local.wf step.hyps(1) step.prems(2) step.prems(3) by blast
- show ?case using step ic ec extend_1_eval[OF wf step.hyps(1)] by simp
-qed
-(*>*)
-
-
-lemma extend_evals:
-assumes wf: "wwf_J_prog P"
-shows "\<lbrakk> P \<turnstile> \<langle>es,s,b\<rangle> [\<rightarrow>]* \<langle>es'',s'',b''\<rangle>; P \<turnstile> \<langle>es'',s''\<rangle> [\<Rightarrow>] \<langle>es',s'\<rangle>;
- iconfs (shp s) es; P,shp s \<turnstile>\<^sub>b (es::expr list,b) \<surd> \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>es,s\<rangle> [\<Rightarrow>] \<langle>es',s'\<rangle>"
-(*<*)
-proof (induct rule: converse_rtrancl_induct3)
- case refl then show ?case by simp
-next
- case (step es1 s1 b1 es2 s2 b2)
- then have ic: "iconfs (shp s2) es2" using Reds_preserves_iconf local.wf by blast
- then have ec: "P,shp s2 \<turnstile>\<^sub>b (es2,b2) \<surd>"
- using Reds_preserves_bconf local.wf step.hyps(1) step.prems(2) step.prems(3) by blast
- show ?case using step ic ec extend_1_evals[OF wf step.hyps(1)] by simp
-qed
-(*>*)
-
-text \<open> Finally, small step semantics can be simulated by big step semantics:
-\<close>
-
-theorem
-assumes wf: "wwf_J_prog P"
-shows small_by_big:
- "\<lbrakk>P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle>; iconf (shp s) e; P,shp s \<turnstile>\<^sub>b (e,b) \<surd>; final e'\<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
-and "\<lbrakk>P \<turnstile> \<langle>es,s,b\<rangle> [\<rightarrow>]* \<langle>es',s',b'\<rangle>; iconfs (shp s) es; P,shp s \<turnstile>\<^sub>b (es,b) \<surd>; finals es'\<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>es,s\<rangle> [\<Rightarrow>] \<langle>es',s'\<rangle>"
-(*<*)
-proof -
- note wf
- moreover assume "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle>"
- moreover assume "final e'"
- then have "P \<turnstile> \<langle>e',s'\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
- by (simp add: eval_finalId)
- moreover assume "iconf (shp s) e"
- moreover assume "P,shp s \<turnstile>\<^sub>b (e,b) \<surd>"
- ultimately show "P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
- by (rule extend_eval)
-next
- assume fins: "finals es'"
- note wf
- moreover assume "P \<turnstile> \<langle>es,s,b\<rangle> [\<rightarrow>]* \<langle>es',s',b'\<rangle>"
- moreover have "P \<turnstile> \<langle>es',s'\<rangle> [\<Rightarrow>] \<langle>es',s'\<rangle>" using fins
- by (rule eval_finalsId)
- moreover assume "iconfs (shp s) es"
- moreover assume "P,shp s \<turnstile>\<^sub>b (es,b) \<surd>"
- ultimately show "P \<turnstile> \<langle>es,s\<rangle> [\<Rightarrow>] \<langle>es',s'\<rangle>"
- by (rule extend_evals)
-qed
-(*>*)
-
-subsection "Equivalence"
-
-text\<open> And now, the crowning achievement: \<close>
-
-corollary big_iff_small:
-"\<lbrakk> wwf_J_prog P; iconf (shp s) e; P,shp s \<turnstile>\<^sub>b (e::expr,b) \<surd> \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle> = (P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e',s',False\<rangle> \<and> final e')"
-(*<*)by(blast dest: big_by_small eval_final small_by_big)(*>*)
-
-corollary big_iff_small_WT:
- "wwf_J_prog P \<Longrightarrow> P,E \<turnstile> e::T \<Longrightarrow> P,shp s \<turnstile>\<^sub>b (e,b) \<surd> \<Longrightarrow>
- P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle> = (P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e',s',False\<rangle> \<and> final e')"
-(*<*)by(blast dest: big_iff_small WT_nsub_RI nsub_RI_iconf)(*>*)
-
-
-subsection \<open> Lifting type safety to @{text"\<Rightarrow>"} \<close>
-
-text\<open> \dots and now to the big step semantics, just for fun. \<close>
-
-lemma eval_preserves_sconf:
-fixes s::state and s'::state
-assumes "wf_J_prog P" and "P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>" and "iconf (shp s) e"
- and "P,E \<turnstile> e::T" and "P,E \<turnstile> s\<surd>"
-shows "P,E \<turnstile> s'\<surd>"
-(*<*)
-proof -
- have "P,shp s \<turnstile>\<^sub>b (e,False) \<surd>" by(simp add: bconf_def)
- with assms show ?thesis
- by(blast intro:Red_preserves_sconf Red_preserves_iconf Red_preserves_bconf big_by_small
- WT_implies_WTrt wf_prog_wwf_prog)
-qed
-(*>*)
-
-
-lemma eval_preserves_type:
-fixes s::state
-assumes wf: "wf_J_prog P"
- and "P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>" and "P,E \<turnstile> s\<surd>" and "iconf (shp s) e" and "P,E \<turnstile> e::T"
-shows "\<exists>T'. P \<turnstile> T' \<le> T \<and> P,E,hp s',shp s' \<turnstile> e':T'"
-(*<*)
-proof -
- have "P,shp s \<turnstile>\<^sub>b (e,False) \<surd>" by(simp add: bconf_def)
- with assms show ?thesis by(blast dest:big_by_small[OF wf_prog_wwf_prog[OF wf]]
- WT_implies_WTrt Red_preserves_type[OF wf])
-qed
-(*>*)
-
-
-end
+(* Title: JinjaDCI/J/Equivalence.thy
+ Author: Tobias Nipkow, Susannah Mansky
+ Copyright 2003 Technische Universitaet Muenchen, 2019-20 UIUC
+
+ Based on the Jinja theory J/Equivalence.thy by Tobias Nipkow
+*)
+
+section \<open> Equivalence of Big Step and Small Step Semantics \<close>
+
+theory Equivalence imports TypeSafe WWellForm begin
+
+subsection\<open>Small steps simulate big step\<close>
+
+subsubsection "Init"
+
+text "The reduction of initialization expressions doesn't touch or use
+ their on-hold expressions (the subexpression to the right of @{text \<leftarrow>})
+ until the initialization operation completes. This function is used to prove
+ this and related properties. It is then used for general reduction of
+ initialization expressions separately from their on-hold expressions by
+ using the on-hold expression @{term unit}, then putting the real on-hold
+ expression back in at the end."
+
+fun init_switch :: "'a exp \<Rightarrow> 'a exp \<Rightarrow> 'a exp" where
+"init_switch (INIT C (Cs,b) \<leftarrow> e\<^sub>i) e = (INIT C (Cs,b) \<leftarrow> e)" |
+"init_switch (RI(C,e');Cs \<leftarrow> e\<^sub>i) e = (RI(C,e');Cs \<leftarrow> e)" |
+"init_switch e' e = e'"
+
+fun INIT_class :: "'a exp \<Rightarrow> cname option" where
+"INIT_class (INIT C (Cs,b) \<leftarrow> e) = (if C = last (C#Cs) then Some C else None)" |
+"INIT_class (RI(C,e\<^sub>0);Cs \<leftarrow> e) = Some (last (C#Cs))" |
+"INIT_class _ = None"
+
+lemma init_red_init:
+"\<lbrakk> init_exp_of e\<^sub>0 = \<lfloor>e\<rfloor>; P \<turnstile> \<langle>e\<^sub>0,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow> \<langle>e\<^sub>1,s\<^sub>1,b\<^sub>1\<rangle> \<rbrakk>
+ \<Longrightarrow> (init_exp_of e\<^sub>1 = \<lfloor>e\<rfloor> \<and> (INIT_class e\<^sub>0 = \<lfloor>C\<rfloor> \<longrightarrow> INIT_class e\<^sub>1 = \<lfloor>C\<rfloor>))
+ \<or> (e\<^sub>1 = e \<and> b\<^sub>1 = icheck P (the(INIT_class e\<^sub>0)) e) \<or> (\<exists>a. e\<^sub>1 = throw a)"
+ by(erule red.cases, auto)
+
+lemma init_exp_switch[simp]:
+"init_exp_of e\<^sub>0 = \<lfloor>e\<rfloor> \<Longrightarrow> init_exp_of (init_switch e\<^sub>0 e') = \<lfloor>e'\<rfloor>"
+ by(cases e\<^sub>0, simp_all)
+
+lemma init_red_sync:
+"\<lbrakk> P \<turnstile> \<langle>e\<^sub>0,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow> \<langle>e\<^sub>1,s\<^sub>1,b\<^sub>1\<rangle>; init_exp_of e\<^sub>0 = \<lfloor>e\<rfloor>; e\<^sub>1 \<noteq> e \<rbrakk>
+ \<Longrightarrow> (\<And>e'. P \<turnstile> \<langle>init_switch e\<^sub>0 e',s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow> \<langle>init_switch e\<^sub>1 e',s\<^sub>1,b\<^sub>1\<rangle>)"
+proof(induct rule: red.cases) qed(simp_all add: red_reds.intros)
+
+lemma init_red_sync_end:
+"\<lbrakk> P \<turnstile> \<langle>e\<^sub>0,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow> \<langle>e\<^sub>1,s\<^sub>1,b\<^sub>1\<rangle>; init_exp_of e\<^sub>0 = \<lfloor>e\<rfloor>; e\<^sub>1 = e; throw_of e = None \<rbrakk>
+ \<Longrightarrow> (\<And>e'. \<not>sub_RI e' \<Longrightarrow> P \<turnstile> \<langle>init_switch e\<^sub>0 e',s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow> \<langle>e',s\<^sub>1,icheck P (the(INIT_class e\<^sub>0)) e'\<rangle>)"
+proof(induct rule: red.cases) qed(simp_all add: red_reds.intros)
+
+lemma init_reds_sync_unit':
+ "\<lbrakk> P \<turnstile> \<langle>e\<^sub>0,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>Val v',s\<^sub>1,b\<^sub>1\<rangle>; init_exp_of e\<^sub>0 = \<lfloor>unit\<rfloor>; INIT_class e\<^sub>0 = \<lfloor>C\<rfloor> \<rbrakk>
+ \<Longrightarrow> (\<And>e'. \<not>sub_RI e' \<Longrightarrow> P \<turnstile> \<langle>init_switch e\<^sub>0 e',s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>e',s\<^sub>1,icheck P (the(INIT_class e\<^sub>0)) e'\<rangle>)"
+proof(induct rule:converse_rtrancl_induct3)
+case refl then show ?case by simp
+next
+ case (step e0 s0 b0 e1 s1 b1)
+ have "(init_exp_of e1 = \<lfloor>unit\<rfloor> \<and> (INIT_class e0 = \<lfloor>C\<rfloor> \<longrightarrow> INIT_class e1 = \<lfloor>C\<rfloor>))
+ \<or> (e1 = unit \<and> b1 = icheck P (the(INIT_class e0)) unit) \<or> (\<exists>a. e1 = throw a)"
+ using init_red_init[OF step.prems(1) step.hyps(1)] by simp
+ then show ?case
+ proof(rule disjE)
+ assume assm: "init_exp_of e1 = \<lfloor>unit\<rfloor> \<and> (INIT_class e0 = \<lfloor>C\<rfloor> \<longrightarrow> INIT_class e1 = \<lfloor>C\<rfloor>)"
+ then have red: "P \<turnstile> \<langle>init_switch e0 e',s0,b0\<rangle> \<rightarrow> \<langle>init_switch e1 e',s1,b1\<rangle>"
+ using init_red_sync[OF step.hyps(1) step.prems(1)] by simp
+ have reds: "P \<turnstile> \<langle>init_switch e1 e',s1,b1\<rangle> \<rightarrow>* \<langle>e',s\<^sub>1,icheck P (the (INIT_class e0)) e'\<rangle>"
+ using step.hyps(3)[OF _ _ step.prems(3)] assm step.prems(2) by simp
+ show ?thesis by(rule converse_rtrancl_into_rtrancl[OF red reds])
+ next
+ assume "(e1 = unit \<and> b1 = icheck P (the(INIT_class e0)) unit) \<or> (\<exists>a. e1 = throw a)"
+ then show ?thesis
+ proof(rule disjE)
+ assume assm: "e1 = unit \<and> b1 = icheck P (the(INIT_class e0)) unit" then have e1: "e1 = unit" by simp
+ have sb: "s1 = s\<^sub>1" "b1 = b\<^sub>1" using reds_final_same[OF step.hyps(2)] assm
+ by(simp_all add: final_def)
+ then have step': "P \<turnstile> \<langle>init_switch e0 e',s0,b0\<rangle> \<rightarrow> \<langle>e',s\<^sub>1,icheck P (the (INIT_class e0)) e'\<rangle>"
+ using init_red_sync_end[OF step.hyps(1) step.prems(1) e1 _ step.prems(3)] by auto
+ then have "P \<turnstile> \<langle>init_switch e0 e',s0,b0\<rangle> \<rightarrow>* \<langle>e',s\<^sub>1,icheck P (the (INIT_class e0)) e'\<rangle>"
+ using r_into_rtrancl by auto
+ then show ?thesis using step assm sb by simp
+ next
+ assume "\<exists>a. e1 = throw a" then obtain a where "e1 = throw a" by clarsimp
+ then have tof: "throw_of e1 = \<lfloor>a\<rfloor>" by simp
+ then show ?thesis using reds_throw[OF step.hyps(2) tof] by simp
+ qed
+ qed
+qed
+
+lemma init_reds_sync_unit_throw':
+ "\<lbrakk> P \<turnstile> \<langle>e\<^sub>0,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>throw a,s\<^sub>1,b\<^sub>1\<rangle>; init_exp_of e\<^sub>0 = \<lfloor>unit\<rfloor> \<rbrakk>
+ \<Longrightarrow> (\<And>e'. P \<turnstile> \<langle>init_switch e\<^sub>0 e',s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>throw a,s\<^sub>1,b\<^sub>1\<rangle>)"
+proof(induct rule:converse_rtrancl_induct3)
+case refl then show ?case by simp
+next
+ case (step e0 s0 b0 e1 s1 b1)
+ have "init_exp_of e1 = \<lfloor>unit\<rfloor> \<and> (\<forall>C. INIT_class e0 = \<lfloor>C\<rfloor> \<longrightarrow> INIT_class e1 = \<lfloor>C\<rfloor>) \<or>
+ e1 = unit \<and> b1 = icheck P (the (INIT_class e0)) unit \<or> (\<exists>a. e1 = throw a)"
+ using init_red_init[OF step.prems(1) step.hyps(1)] by auto
+ then show ?case
+ proof(rule disjE)
+ assume assm: "init_exp_of e1 = \<lfloor>unit\<rfloor> \<and> (\<forall>C. INIT_class e0 = \<lfloor>C\<rfloor> \<longrightarrow> INIT_class e1 = \<lfloor>C\<rfloor>)"
+ then have "P \<turnstile> \<langle>init_switch e0 e',s0,b0\<rangle> \<rightarrow> \<langle>init_switch e1 e',s1,b1\<rangle>"
+ using step init_red_sync[OF step.hyps(1) step.prems] by simp
+ then show ?thesis using step assm by (meson converse_rtrancl_into_rtrancl)
+ next
+ assume "e1 = unit \<and> b1 = icheck P (the (INIT_class e0)) unit \<or> (\<exists>a. e1 = throw a)"
+ then show ?thesis
+ proof(rule disjE)
+ assume "e1 = unit \<and> b1 = icheck P (the (INIT_class e0)) unit"
+ then show ?thesis using step using final_def reds_final_same by blast
+ next
+ assume assm: "\<exists>a. e1 = throw a"
+ then have "P \<turnstile> \<langle>init_switch e0 e',s0,b0\<rangle> \<rightarrow> \<langle>e1,s1,b1\<rangle>"
+ using init_red_sync[OF step.hyps(1) step.prems] by clarsimp
+ then show ?thesis using step by simp
+ qed
+ qed
+qed
+
+lemma init_reds_sync_unit:
+assumes "P \<turnstile> \<langle>e\<^sub>0,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>Val v',s\<^sub>1,b\<^sub>1\<rangle>" and "init_exp_of e\<^sub>0 = \<lfloor>unit\<rfloor>" and "INIT_class e\<^sub>0 = \<lfloor>C\<rfloor>"
+ and "\<not>sub_RI e'"
+shows "P \<turnstile> \<langle>init_switch e\<^sub>0 e',s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>e',s\<^sub>1,icheck P (the(INIT_class e\<^sub>0)) e'\<rangle>"
+using init_reds_sync_unit'[OF assms] by clarsimp
+
+lemma init_reds_sync_unit_throw:
+assumes "P \<turnstile> \<langle>e\<^sub>0,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>throw a,s\<^sub>1,b\<^sub>1\<rangle>" and "init_exp_of e\<^sub>0 = \<lfloor>unit\<rfloor>"
+shows "P \<turnstile> \<langle>init_switch e\<^sub>0 e',s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>throw a,s\<^sub>1,b\<^sub>1\<rangle>"
+using init_reds_sync_unit_throw'[OF assms] by clarsimp
+
+\<comment> \<open> init reds lemmas \<close>
+
+lemma InitSeqReds:
+assumes "P \<turnstile> \<langle>INIT C ([C],b) \<leftarrow> unit,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>Val v',s\<^sub>1,b\<^sub>1\<rangle>"
+ and "P \<turnstile> \<langle>e,s\<^sub>1,icheck P C e\<rangle> \<rightarrow>* \<langle>e\<^sub>2,s\<^sub>2,b\<^sub>2\<rangle>" and "\<not>sub_RI e"
+shows "P \<turnstile> \<langle>INIT C ([C],b) \<leftarrow> e,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>e\<^sub>2,s\<^sub>2,b\<^sub>2\<rangle>"
+using assms
+proof -
+ have "P \<turnstile> \<langle>init_switch (INIT C ([C],b) \<leftarrow> unit) e,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>e,s\<^sub>1,icheck P C e\<rangle>"
+ using init_reds_sync_unit[OF assms(1) _ _ assms(3)] by simp
+ then show ?thesis using assms(2) by simp
+qed
+
+lemma InitSeqThrowReds:
+assumes "P \<turnstile> \<langle>INIT C ([C],b) \<leftarrow> unit,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>throw a,s\<^sub>1,b\<^sub>1\<rangle>"
+shows "P \<turnstile> \<langle>INIT C ([C],b) \<leftarrow> e,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>throw a,s\<^sub>1,b\<^sub>1\<rangle>"
+using assms
+proof -
+ have "P \<turnstile> \<langle>init_switch (INIT C ([C],b) \<leftarrow> unit) e,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>throw a,s\<^sub>1,b\<^sub>1\<rangle>"
+ using init_reds_sync_unit_throw[OF assms(1)] by simp
+ then show ?thesis by simp
+qed
+
+lemma InitNoneReds:
+ "\<lbrakk> sh C = None;
+ P \<turnstile> \<langle>INIT C' (C # Cs,False) \<leftarrow> e,(h, l, sh(C \<mapsto> (sblank P C, Prepared))),b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle> \<rbrakk>
+\<Longrightarrow> P \<turnstile> \<langle>INIT C' (C#Cs,False) \<leftarrow> e,(h,l,sh),b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle>"
+(*<*)by(rule InitNoneRed[THEN converse_rtrancl_into_rtrancl])(*>*)
+
+lemma InitDoneReds:
+ "\<lbrakk> sh C = Some(sfs,Done); P \<turnstile> \<langle>INIT C' (Cs,True) \<leftarrow> e,(h,l,sh),b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle> \<rbrakk>
+\<Longrightarrow> P \<turnstile> \<langle>INIT C' (C#Cs,False) \<leftarrow> e,(h,l,sh),b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle>"
+(*<*)by(rule RedInitDone[THEN converse_rtrancl_into_rtrancl])(*>*)
+
+lemma InitProcessingReds:
+ "\<lbrakk> sh C = Some(sfs,Processing); P \<turnstile> \<langle>INIT C' (Cs,True) \<leftarrow> e,(h,l,sh),b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle> \<rbrakk>
+\<Longrightarrow> P \<turnstile> \<langle>INIT C' (C#Cs,False) \<leftarrow> e,(h,l,sh),b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle>"
+(*<*)by(rule RedInitProcessing[THEN converse_rtrancl_into_rtrancl])(*>*)
+
+lemma InitErrorReds:
+ "\<lbrakk> sh C = Some(sfs,Error); P \<turnstile> \<langle>RI (C,THROW NoClassDefFoundError);Cs \<leftarrow> e,(h,l,sh),b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle> \<rbrakk>
+\<Longrightarrow> P \<turnstile> \<langle>INIT C' (C#Cs,False) \<leftarrow> e,(h,l,sh),b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle>"
+(*<*)by(rule RedInitError[THEN converse_rtrancl_into_rtrancl])(*>*)
+
+lemma InitObjectReds:
+ "\<lbrakk> sh C = Some(sfs,Prepared); C = Object; sh' = sh(C \<mapsto> (sfs,Processing));
+ P \<turnstile> \<langle>INIT C' (C#Cs,True) \<leftarrow> e,(h,l,sh'),b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle> \<rbrakk>
+\<Longrightarrow> P \<turnstile> \<langle>INIT C' (C#Cs,False) \<leftarrow> e,(h,l,sh),b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle>"
+(*<*)by(rule InitObjectRed[THEN converse_rtrancl_into_rtrancl])(*>*)
+
+lemma InitNonObjectReds:
+ "\<lbrakk> sh C = Some(sfs,Prepared); C \<noteq> Object; class P C = Some (D,r);
+ sh' = sh(C \<mapsto> (sfs,Processing));
+ P \<turnstile> \<langle>INIT C' (D#C#Cs,False) \<leftarrow> e,(h,l,sh'),b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle> \<rbrakk>
+\<Longrightarrow> P \<turnstile> \<langle>INIT C' (C#Cs,False) \<leftarrow> e,(h,l,sh),b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle>"
+(*<*)by(rule InitNonObjectSuperRed[THEN converse_rtrancl_into_rtrancl])(*>*)
+
+lemma RedsInitRInit:
+ "P \<turnstile> \<langle>RI (C,C\<bullet>\<^sub>sclinit([]));Cs \<leftarrow> e,(h,l,sh),b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle>
+\<Longrightarrow> P \<turnstile> \<langle>INIT C' (C#Cs,True) \<leftarrow> e,(h,l,sh),b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle>"
+(*<*)by(rule RedInitRInit[THEN converse_rtrancl_into_rtrancl])(*>*)
+
+lemmas rtrancl_induct3 =
+ rtrancl_induct[of "(ax,ay,az)" "(bx,by,bz)", split_format (complete), consumes 1, case_names refl step]
+
+lemma RInitReds:
+ "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle>
+\<Longrightarrow> P \<turnstile> \<langle>RI (C,e);Cs \<leftarrow> e\<^sub>0, s, b\<rangle> \<rightarrow>* \<langle>RI (C,e');Cs \<leftarrow> e\<^sub>0, s', b'\<rangle>"
+(*<*)
+proof(induct rule: rtrancl_induct3)
+ case refl show ?case by blast
+next
+ case step show ?case
+ by(rule rtrancl_into_rtrancl[OF step(3) RInitRed[OF step(2)]])
+qed
+(*>*)
+
+lemma RedsRInit:
+ "\<lbrakk> P \<turnstile> \<langle>e\<^sub>0,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>Val v,(h\<^sub>1,l\<^sub>1,sh\<^sub>1),b\<^sub>1\<rangle>;
+ sh\<^sub>1 C = Some (sfs, i); sh\<^sub>2 = sh\<^sub>1(C \<mapsto> (sfs,Done)); C' = last(C#Cs);
+ P \<turnstile> \<langle>INIT C' (Cs,True) \<leftarrow> e,(h\<^sub>1,l\<^sub>1,sh\<^sub>2),b\<^sub>1\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle> \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>RI (C, e\<^sub>0);Cs \<leftarrow> e,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle>"
+(*<*)
+by(rule rtrancl_trans[OF RInitReds
+ RedRInit[THEN converse_rtrancl_into_rtrancl]])
+(*>*)
+
+
+lemma RInitInitThrowReds:
+ "\<lbrakk> P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>Throw a,(h',l',sh'),b'\<rangle>;
+ sh' C = Some (sfs, i); sh'' = sh'(C \<mapsto> (sfs,Error));
+ P \<turnstile> \<langle>RI (D,Throw a);Cs \<leftarrow> e\<^sub>0, (h',l',sh''),b'\<rangle> \<rightarrow>* \<langle>e\<^sub>1,s\<^sub>1,b\<^sub>1\<rangle> \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>RI (C, e);D#Cs \<leftarrow> e\<^sub>0,s,b\<rangle> \<rightarrow>* \<langle>e\<^sub>1,s\<^sub>1,b\<^sub>1\<rangle>"
+(*<*)
+by(rule rtrancl_trans[OF RInitReds
+ RInitInitThrow[THEN converse_rtrancl_into_rtrancl]])
+(*>*)
+
+lemma RInitThrowReds:
+ "\<lbrakk> P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>Throw a, (h',l',sh'),b'\<rangle>;
+ sh' C = Some(sfs, i); sh'' = sh'(C \<mapsto> (sfs, Error)) \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>RI (C,e);Nil \<leftarrow> e\<^sub>0,s,b\<rangle> \<rightarrow>* \<langle>Throw a, (h',l',sh''),b'\<rangle>"
+(*<*)by(rule RInitReds[THEN rtrancl_into_rtrancl, OF _ RInitThrow])(*>*)
+
+subsubsection "New"
+
+lemma NewInitDoneReds:
+ "\<lbrakk> sh C = Some (sfs, Done); new_Addr h = Some a;
+ P \<turnstile> C has_fields FDTs; h' = h(a\<mapsto>blank P C) \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>new C,(h,l,sh),False\<rangle> \<rightarrow>* \<langle>addr a,(h',l,sh),False\<rangle>"
+(*<*)
+by(rule NewInitDoneRed[THEN converse_rtrancl_into_rtrancl,
+ OF _ RedNew[THEN r_into_rtrancl]])
+(*>*)
+
+lemma NewInitDoneReds2:
+ "\<lbrakk> sh C = Some (sfs, Done); new_Addr h = None; is_class P C \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>new C,(h,l,sh),False\<rangle> \<rightarrow>* \<langle>THROW OutOfMemory, (h,l,sh), False\<rangle>"
+(*<*)
+by(rule NewInitDoneRed[THEN converse_rtrancl_into_rtrancl,
+ OF _ RedNewFail[THEN r_into_rtrancl]])
+(*>*)
+
+lemma NewInitReds:
+assumes nDone: "\<nexists>sfs. shp s C = Some (sfs, Done)"
+ and INIT_steps: "P \<turnstile> \<langle>INIT C ([C],False) \<leftarrow> unit,s,False\<rangle> \<rightarrow>* \<langle>Val v',(h',l',sh'),b'\<rangle>"
+ and Addr: "new_Addr h' = Some a" and has: "P \<turnstile> C has_fields FDTs"
+ and h'': "h'' = h'(a\<mapsto>blank P C)" and cls_C: "is_class P C"
+shows "P \<turnstile> \<langle>new C,s,False\<rangle> \<rightarrow>* \<langle>addr a,(h'',l',sh'),False\<rangle>"
+(*<*)(is "(?a, ?c) \<in> (red P)\<^sup>*")
+proof -
+ let ?b = "(INIT C ([C],False) \<leftarrow> new C,s,False)"
+ let ?b' = "(new C,(h', l', sh'),icheck P C (new C::expr))"
+ have b'c: "(?b', ?c) \<in> (red P)\<^sup>*"
+ using RedNew[OF Addr has h'', THEN r_into_rtrancl] by simp
+ obtain h l sh where [simp]: "s = (h, l, sh)" by(cases s) simp
+ have "(?a, ?b) \<in> (red P)\<^sup>*"
+ using NewInitRed[OF _ cls_C] nDone by fastforce
+ also have "(?b, ?c) \<in> (red P)\<^sup>*"
+ by(rule InitSeqReds[OF INIT_steps b'c]) simp
+ ultimately show ?thesis by simp
+qed
+(*>*)
+
+lemma NewInitOOMReds:
+assumes nDone: "\<nexists>sfs. shp s C = Some (sfs, Done)"
+ and INIT_steps: "P \<turnstile> \<langle>INIT C ([C],False) \<leftarrow> unit,s,False\<rangle> \<rightarrow>* \<langle>Val v',(h',l',sh'),b'\<rangle>"
+ and Addr: "new_Addr h' = None" and cls_C: "is_class P C"
+shows "P \<turnstile> \<langle>new C,s,False\<rangle> \<rightarrow>* \<langle>THROW OutOfMemory,(h',l',sh'),False\<rangle>"
+(*<*)(is "(?a, ?c) \<in> (red P)\<^sup>*")
+proof -
+ let ?b = "(INIT C ([C],False) \<leftarrow> new C,s,False)"
+ let ?b' = "(new C,(h', l', sh'),icheck P C (new C::expr))"
+ have b'c: "(?b', ?c) \<in> (red P)\<^sup>*"
+ using RedNewFail[OF Addr cls_C, THEN r_into_rtrancl] by simp
+ obtain h l sh where [simp]: "s = (h, l, sh)" by(cases s) simp
+ have "(?a, ?b) \<in> (red P)\<^sup>*"
+ using NewInitRed[OF _ cls_C] nDone by fastforce
+ also have "(?b, ?c) \<in> (red P)\<^sup>*"
+ by(rule InitSeqReds[OF INIT_steps b'c]) simp
+ ultimately show ?thesis by simp
+qed
+(*>*)
+
+lemma NewInitThrowReds:
+assumes nDone: "\<nexists>sfs. shp s C = Some (sfs, Done)"
+ and cls_C: "is_class P C"
+ and INIT_steps: "P \<turnstile> \<langle>INIT C ([C],False) \<leftarrow> unit,s,False\<rangle> \<rightarrow>* \<langle>throw a,s',b'\<rangle>"
+shows "P \<turnstile> \<langle>new C,s,False\<rangle> \<rightarrow>* \<langle>throw a,s',b'\<rangle>"
+(*<*)(is "(?a, ?c) \<in> (red P)\<^sup>*")
+proof -
+ let ?b = "(INIT C ([C],False) \<leftarrow> new C,s,False)"
+ obtain h l sh where [simp]: "s = (h, l, sh)" by(cases s) simp
+ have "(?a, ?b) \<in> (red P)\<^sup>*"
+ using NewInitRed[OF _ cls_C] nDone by fastforce
+ also have "(?b, ?c) \<in> (red P)\<^sup>*"
+ using InitSeqThrowReds[OF INIT_steps] by simp
+ ultimately show ?thesis by simp
+qed
+(*>*)
+
+subsubsection "Cast"
+
+lemma CastReds:
+ "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>Cast C e,s,b\<rangle> \<rightarrow>* \<langle>Cast C e',s',b'\<rangle>"
+(*<*)
+proof(induct rule: rtrancl_induct3)
+ case refl show ?case by blast
+next
+ case step show ?case
+ by(rule rtrancl_into_rtrancl[OF step(3) CastRed[OF step(2)]])
+qed
+(*>*)
+
+lemma CastRedsNull:
+ "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>null,s',b'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>Cast C e,s,b\<rangle> \<rightarrow>* \<langle>null,s',b'\<rangle>"
+(*<*)by(rule CastReds[THEN rtrancl_into_rtrancl, OF _ RedCastNull])(*>*)
+
+lemma CastRedsAddr:
+ "\<lbrakk> P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>addr a,s',b'\<rangle>; hp s' a = Some(D,fs); P \<turnstile> D \<preceq>\<^sup>* C \<rbrakk> \<Longrightarrow>
+ P \<turnstile> \<langle>Cast C e,s,b\<rangle> \<rightarrow>* \<langle>addr a,s',b'\<rangle>"
+(*<*)by(cases s', simp) (rule CastReds[THEN rtrancl_into_rtrancl, OF _ RedCast])(*>*)
+
+lemma CastRedsFail:
+ "\<lbrakk> P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>addr a,s',b'\<rangle>; hp s' a = Some(D,fs); \<not> P \<turnstile> D \<preceq>\<^sup>* C \<rbrakk> \<Longrightarrow>
+ P \<turnstile> \<langle>Cast C e,s,b\<rangle> \<rightarrow>* \<langle>THROW ClassCast,s',b'\<rangle>"
+(*<*)by(cases s', simp) (rule CastReds[THEN rtrancl_into_rtrancl, OF _ RedCastFail])(*>*)
+
+lemma CastRedsThrow:
+ "\<lbrakk> P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>throw a,s',b'\<rangle> \<rbrakk> \<Longrightarrow> P \<turnstile> \<langle>Cast C e,s,b\<rangle> \<rightarrow>* \<langle>throw a,s',b'\<rangle>"
+(*<*)by(rule CastReds[THEN rtrancl_into_rtrancl, OF _ red_reds.CastThrow])(*>*)
+
+subsubsection "LAss"
+
+lemma LAssReds:
+ "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle> \<Longrightarrow> P \<turnstile> \<langle> V:=e,s,b\<rangle> \<rightarrow>* \<langle> V:=e',s',b'\<rangle>"
+(*<*)
+proof(induct rule: rtrancl_induct3)
+ case refl show ?case by blast
+next
+ case step show ?case
+ by(rule rtrancl_into_rtrancl[OF step(3) LAssRed[OF step(2)]])
+qed
+(*>*)
+
+lemma LAssRedsVal:
+ "\<lbrakk> P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>Val v,(h',l',sh'),b'\<rangle> \<rbrakk> \<Longrightarrow> P \<turnstile> \<langle> V:=e,s,b\<rangle> \<rightarrow>* \<langle>unit,(h',l'(V\<mapsto>v),sh'),b'\<rangle>"
+(*<*)by(rule LAssReds[THEN rtrancl_into_rtrancl, OF _ RedLAss])(*>*)
+
+lemma LAssRedsThrow:
+ "\<lbrakk> P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>throw a,s',b'\<rangle> \<rbrakk> \<Longrightarrow> P \<turnstile> \<langle> V:=e,s,b\<rangle> \<rightarrow>* \<langle>throw a,s',b'\<rangle>"
+(*<*)by(rule LAssReds[THEN rtrancl_into_rtrancl, OF _ red_reds.LAssThrow])(*>*)
+
+subsubsection "BinOp"
+
+lemma BinOp1Reds:
+ "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle> \<Longrightarrow> P \<turnstile> \<langle> e \<guillemotleft>bop\<guillemotright> e\<^sub>2, s,b\<rangle> \<rightarrow>* \<langle>e' \<guillemotleft>bop\<guillemotright> e\<^sub>2, s',b'\<rangle>"
+(*<*)
+proof(induct rule: rtrancl_induct3)
+ case refl show ?case by blast
+next
+ case step show ?case
+ by(rule rtrancl_into_rtrancl[OF step(3) BinOpRed1[OF step(2)]])
+qed
+(*>*)
+
+lemma BinOp2Reds:
+ "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>(Val v) \<guillemotleft>bop\<guillemotright> e, s,b\<rangle> \<rightarrow>* \<langle>(Val v) \<guillemotleft>bop\<guillemotright> e', s',b'\<rangle>"
+(*<*)
+proof(induct rule: rtrancl_induct3)
+ case refl show ?case by blast
+next
+ case step show ?case
+ by(rule rtrancl_into_rtrancl[OF step(3) BinOpRed2[OF step(2)]])
+qed
+(*>*)
+
+lemma BinOpRedsVal:
+assumes e\<^sub>1_steps: "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>Val v\<^sub>1,s\<^sub>1,b\<^sub>1\<rangle>"
+ and e\<^sub>2_steps: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>1,b\<^sub>1\<rangle> \<rightarrow>* \<langle>Val v\<^sub>2,s\<^sub>2,b\<^sub>2\<rangle>"
+ and op: "binop(bop,v\<^sub>1,v\<^sub>2) = Some v"
+shows "P \<turnstile> \<langle>e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2, s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>Val v,s\<^sub>2,b\<^sub>2\<rangle>"
+(*<*)(is "(?x, ?z) \<in> (red P)\<^sup>*")
+proof -
+ let ?y = "(Val v\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2, s\<^sub>1,b\<^sub>1)"
+ let ?y' = "(Val v\<^sub>1 \<guillemotleft>bop\<guillemotright> Val v\<^sub>2, s\<^sub>2,b\<^sub>2)"
+ have "(?x, ?y) \<in> (red P)\<^sup>*" by(rule BinOp1Reds[OF e\<^sub>1_steps])
+ also have "(?y, ?y') \<in> (red P)\<^sup>*" by(rule BinOp2Reds[OF e\<^sub>2_steps])
+ also have "(?y', ?z) \<in> (red P)" by(rule RedBinOp[OF op])
+ ultimately show ?thesis by simp
+qed
+(*>*)
+
+lemma BinOpRedsThrow1:
+ "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>throw e',s',b'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>e \<guillemotleft>bop\<guillemotright> e\<^sub>2, s,b\<rangle> \<rightarrow>* \<langle>throw e', s',b'\<rangle>"
+(*<*)by(rule BinOp1Reds[THEN rtrancl_into_rtrancl, OF _ red_reds.BinOpThrow1])(*>*)
+
+lemma BinOpRedsThrow2:
+assumes e\<^sub>1_steps: "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>Val v\<^sub>1,s\<^sub>1,b\<^sub>1\<rangle>"
+ and e\<^sub>2_steps: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>1,b\<^sub>1\<rangle> \<rightarrow>* \<langle>throw e,s\<^sub>2,b\<^sub>2\<rangle>"
+shows "P \<turnstile> \<langle>e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2, s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>throw e,s\<^sub>2,b\<^sub>2\<rangle>"
+(*<*)(is "(?x, ?z) \<in> (red P)\<^sup>*")
+proof -
+ let ?y = "(Val v\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2, s\<^sub>1,b\<^sub>1)"
+ let ?y' = "(Val v\<^sub>1 \<guillemotleft>bop\<guillemotright> throw e, s\<^sub>2,b\<^sub>2)"
+ have "(?x, ?y) \<in> (red P)\<^sup>*" by(rule BinOp1Reds[OF e\<^sub>1_steps])
+ also have "(?y, ?y') \<in> (red P)\<^sup>*" by(rule BinOp2Reds[OF e\<^sub>2_steps])
+ also have "(?y', ?z) \<in> (red P)" by(rule red_reds.BinOpThrow2)
+ ultimately show ?thesis by simp
+qed
+(*>*)
+
+subsubsection "FAcc"
+
+lemma FAccReds:
+ "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>e\<bullet>F{D}, s,b\<rangle> \<rightarrow>* \<langle>e'\<bullet>F{D}, s',b'\<rangle>"
+(*<*)
+proof(induct rule: rtrancl_induct3)
+ case refl show ?case by blast
+next
+ case step show ?case
+ by(rule rtrancl_into_rtrancl[OF step(3) FAccRed[OF step(2)]])
+qed
+(*>*)
+
+lemma FAccRedsVal:
+ "\<lbrakk> P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>addr a,s',b'\<rangle>; hp s' a = Some(C,fs); fs(F,D) = Some v;
+ P \<turnstile> C has F,NonStatic:t in D \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>e\<bullet>F{D},s,b\<rangle> \<rightarrow>* \<langle>Val v,s',b'\<rangle>"
+(*<*)by(cases s',simp) (rule FAccReds[THEN rtrancl_into_rtrancl, OF _ RedFAcc])(*>*)
+
+lemma FAccRedsNull:
+ "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>null,s',b'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>e\<bullet>F{D},s,b\<rangle> \<rightarrow>* \<langle>THROW NullPointer,s',b'\<rangle>"
+(*<*)by(rule FAccReds[THEN rtrancl_into_rtrancl, OF _ RedFAccNull])(*>*)
+
+lemma FAccRedsNone:
+ "\<lbrakk> P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>addr a,s',b'\<rangle>;
+ hp s' a = Some(C,fs);
+ \<not>(\<exists>b t. P \<turnstile> C has F,b:t in D) \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>e\<bullet>F{D},s,b\<rangle> \<rightarrow>* \<langle>THROW NoSuchFieldError,s',b'\<rangle>"
+(*<*)by(cases s',simp) (auto intro: FAccReds[THEN rtrancl_into_rtrancl, OF _ RedFAccNone])(*>*)
+
+lemma FAccRedsStatic:
+ "\<lbrakk> P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>addr a,s',b'\<rangle>;
+ hp s' a = Some(C,fs);
+ P \<turnstile> C has F,Static:t in D \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>e\<bullet>F{D},s,b\<rangle> \<rightarrow>* \<langle>THROW IncompatibleClassChangeError,s',b'\<rangle>"
+(*<*)by(cases s',simp) (rule FAccReds[THEN rtrancl_into_rtrancl, OF _ RedFAccStatic])(*>*)
+
+lemma FAccRedsThrow:
+ "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>throw a,s',b'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>e\<bullet>F{D},s,b\<rangle> \<rightarrow>* \<langle>throw a,s',b'\<rangle>"
+(*<*)by(rule FAccReds[THEN rtrancl_into_rtrancl, OF _ red_reds.FAccThrow])(*>*)
+
+subsubsection "SFAcc"
+
+lemma SFAccReds:
+ "\<lbrakk> P \<turnstile> C has F,Static:t in D;
+ shp s D = Some(sfs,Done); sfs F = Some v \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D},s,True\<rangle> \<rightarrow>* \<langle>Val v,s,False\<rangle>"
+(*<*)by(cases s,simp) (rule RedSFAcc[THEN r_into_rtrancl])(*>*)
+
+lemma SFAccRedsNone:
+ "\<not>(\<exists>b t. P \<turnstile> C has F,b:t in D)
+ \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D},s,b\<rangle> \<rightarrow>* \<langle>THROW NoSuchFieldError,s,False\<rangle>"
+(*<*)by(cases s,simp) (auto intro: RedSFAccNone[THEN r_into_rtrancl])(*>*)
+
+lemma SFAccRedsNonStatic:
+ "P \<turnstile> C has F,NonStatic:t in D
+ \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D},s,b\<rangle> \<rightarrow>* \<langle>THROW IncompatibleClassChangeError,s,False\<rangle>"
+(*<*)by(cases s,simp) (rule RedSFAccNonStatic[THEN r_into_rtrancl])(*>*)
+
+lemma SFAccInitDoneReds:
+assumes cF: "P \<turnstile> C has F,Static:t in D"
+ and shp: "shp s D = Some (sfs,Done)" and sfs: "sfs F = Some v"
+shows "P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D}, s, b\<rangle> \<rightarrow>* \<langle>Val v, s, False\<rangle>"
+(*<*)(is "(?a, ?c) \<in> (red P)\<^sup>*")
+proof -
+ obtain h l sh where [simp]: "s = (h, l, sh)" by(cases s) simp
+ show ?thesis proof(cases b)
+ case True
+ then show ?thesis using assms
+ by simp (rule RedSFAcc[THEN r_into_rtrancl])
+ next
+ case False
+ let ?b = "(C\<bullet>\<^sub>sF{D},s,True)"
+ have "(?a, ?b) \<in> (red P)\<^sup>*"
+ using False SFAccInitDoneRed[where sh="shp s", OF cF shp] by fastforce
+ also have "(?b, ?c) \<in> (red P)\<^sup>*" by(rule SFAccReds[OF assms])
+ ultimately show ?thesis by simp
+ qed
+qed
+(*>*)
+
+lemma SFAccInitReds:
+assumes cF: "P \<turnstile> C has F,Static:t in D"
+ and nDone: "\<nexists>sfs. shp s D = Some (sfs,Done)"
+ and INIT_steps: "P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,s,False\<rangle> \<rightarrow>* \<langle>Val v',s',b'\<rangle>"
+ and shp': "shp s' D = Some (sfs,i)" and sfs: "sfs F = Some v"
+shows "P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D},s,False\<rangle> \<rightarrow>* \<langle>Val v,s',False\<rangle>"
+(*<*)(is "(?a, ?c) \<in> (red P)\<^sup>*")
+proof -
+ let ?b = "(INIT D ([D],False) \<leftarrow> C\<bullet>\<^sub>sF{D},s,False)"
+ let ?b' = "(C\<bullet>\<^sub>sF{D},s',icheck P D (C\<bullet>\<^sub>sF{D}::expr))"
+ obtain h l sh where [simp]: "s = (h, l, sh)" by(cases s) simp
+ obtain h' l' sh' where [simp]: "s' = (h', l', sh')" by(cases s') simp
+ have icheck: "icheck P D (C\<bullet>\<^sub>sF{D}) = True" using cF by fastforce
+ then have b'c: "(?b', ?c) \<in> (red P)\<^sup>*"
+ using RedSFAcc[THEN r_into_rtrancl, OF cF] shp' sfs by simp
+ have "(?a, ?b) \<in> (red P)" using SFAccInitRed[OF cF] nDone by simp
+ also have "(?b, ?c) \<in> (red P)\<^sup>*"
+ by(rule InitSeqReds[OF INIT_steps b'c]) simp
+ ultimately show ?thesis by simp
+qed
+(*>*)
+
+lemma SFAccInitThrowReds:
+ "\<lbrakk> P \<turnstile> C has F,Static:t in D;
+ \<nexists>sfs. shp s D = Some (sfs,Done);
+ P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,s,False\<rangle> \<rightarrow>* \<langle>throw a,s',b'\<rangle> \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D},s,False\<rangle> \<rightarrow>* \<langle>throw a,s',b'\<rangle>"
+(*<*)
+by(cases s, simp)
+ (auto elim: converse_rtrancl_into_rtrancl[OF SFAccInitRed InitSeqThrowReds])
+(*>*)
+
+subsubsection "FAss"
+
+lemma FAssReds1:
+ "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>e\<bullet>F{D}:=e\<^sub>2, s,b\<rangle> \<rightarrow>* \<langle>e'\<bullet>F{D}:=e\<^sub>2, s',b'\<rangle>"
+(*<*)
+proof(induct rule: rtrancl_induct3)
+ case refl show ?case by blast
+next
+ case step show ?case
+ by(rule rtrancl_into_rtrancl[OF step(3) FAssRed1[OF step(2)]])
+qed
+(*>*)
+
+lemma FAssReds2:
+ "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>Val v\<bullet>F{D}:=e, s,b\<rangle> \<rightarrow>* \<langle>Val v\<bullet>F{D}:=e', s',b'\<rangle>"
+(*<*)
+proof(induct rule: rtrancl_induct3)
+ case refl show ?case by blast
+next
+ case step show ?case
+ by(rule rtrancl_into_rtrancl[OF step(3) FAssRed2[OF step(2)]])
+qed
+(*>*)
+
+lemma FAssRedsVal:
+assumes e\<^sub>1_steps: "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>addr a,s\<^sub>1,b\<^sub>1\<rangle>"
+ and e\<^sub>2_steps: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>1,b\<^sub>1\<rangle> \<rightarrow>* \<langle>Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\<rangle>"
+ and cF: "P \<turnstile> C has F,NonStatic:t in D" and hC: "Some(C,fs) = h\<^sub>2 a"
+shows "P \<turnstile> \<langle>e\<^sub>1\<bullet>F{D}:=e\<^sub>2, s\<^sub>0, b\<^sub>0\<rangle> \<rightarrow>* \<langle>unit, (h\<^sub>2(a\<mapsto>(C,fs((F,D)\<mapsto>v))),l\<^sub>2,sh\<^sub>2),b\<^sub>2\<rangle>"
+(*<*)(is "(?x, ?z) \<in> (red P)\<^sup>*")
+proof -
+ let ?y = "(addr a\<bullet>F{D}:=e\<^sub>2, s\<^sub>1, b\<^sub>1)"
+ let ?y' = "(addr a\<bullet>F{D}:=Val v::expr,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2)"
+ have "(?x, ?y) \<in> (red P)\<^sup>*" by(rule FAssReds1[OF e\<^sub>1_steps])
+ also have "(?y, ?y') \<in> (red P)\<^sup>*" by(rule FAssReds2[OF e\<^sub>2_steps])
+ also have "(?y', ?z) \<in> (red P)" using RedFAss[OF cF] hC by simp
+ ultimately show ?thesis by simp
+qed
+(*>*)
+
+lemma FAssRedsNull:
+assumes e\<^sub>1_steps: "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>null,s\<^sub>1,b\<^sub>1\<rangle>"
+ and e\<^sub>2_steps: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>1,b\<^sub>1\<rangle> \<rightarrow>* \<langle>Val v,s\<^sub>2,b\<^sub>2\<rangle>"
+shows "P \<turnstile> \<langle>e\<^sub>1\<bullet>F{D}:=e\<^sub>2, s\<^sub>0, b\<^sub>0\<rangle> \<rightarrow>* \<langle>THROW NullPointer, s\<^sub>2, b\<^sub>2\<rangle>"
+(*<*)(is "(?x, ?z) \<in> (red P)\<^sup>*")
+proof -
+ let ?y = "(null\<bullet>F{D}:=e\<^sub>2, s\<^sub>1, b\<^sub>1)"
+ let ?y' = "(null\<bullet>F{D}:=Val v::expr,s\<^sub>2,b\<^sub>2)"
+ have "(?x, ?y) \<in> (red P)\<^sup>*" by(rule FAssReds1[OF e\<^sub>1_steps])
+ also have "(?y, ?y') \<in> (red P)\<^sup>*" by(rule FAssReds2[OF e\<^sub>2_steps])
+ also have "(?y', ?z) \<in> (red P)" by(rule RedFAssNull)
+ ultimately show ?thesis by simp
+qed
+(*>*)
+
+lemma FAssRedsThrow1:
+ "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>throw e',s',b'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>e\<bullet>F{D}:=e\<^sub>2, s,b\<rangle> \<rightarrow>* \<langle>throw e', s',b'\<rangle>"
+(*<*)by(rule FAssReds1[THEN rtrancl_into_rtrancl, OF _ red_reds.FAssThrow1])(*>*)
+
+lemma FAssRedsThrow2:
+assumes e\<^sub>1_steps: "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>Val v,s\<^sub>1,b\<^sub>1\<rangle>"
+ and e\<^sub>2_steps: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>1,b\<^sub>1\<rangle> \<rightarrow>* \<langle>throw e,s\<^sub>2,b\<^sub>2\<rangle>"
+shows "P \<turnstile> \<langle>e\<^sub>1\<bullet>F{D}:=e\<^sub>2,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>throw e,s\<^sub>2,b\<^sub>2\<rangle>"
+(*<*)(is "(?x, ?z) \<in> (red P)\<^sup>*")
+proof -
+ let ?y = "(Val v\<bullet>F{D}:=e\<^sub>2,s\<^sub>1,b\<^sub>1)"
+ let ?y' = "(Val v\<bullet>F{D}:=throw e,s\<^sub>2,b\<^sub>2)"
+ have "(?x, ?y) \<in> (red P)\<^sup>*" by(rule FAssReds1[OF e\<^sub>1_steps])
+ also have "(?y, ?y') \<in> (red P)\<^sup>*" by(rule FAssReds2[OF e\<^sub>2_steps])
+ also have "(?y', ?z) \<in> (red P)" by(rule red_reds.FAssThrow2)
+ ultimately show ?thesis by simp
+qed
+(*>*)
+
+lemma FAssRedsNone:
+assumes e\<^sub>1_steps: "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>addr a,s\<^sub>1,b\<^sub>1\<rangle>"
+ and e\<^sub>2_steps: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>1,b\<^sub>1\<rangle> \<rightarrow>* \<langle>Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\<rangle>"
+ and hC: "h\<^sub>2 a = Some(C,fs)" and ncF: "\<not>(\<exists>b t. P \<turnstile> C has F,b:t in D)"
+shows "P \<turnstile> \<langle>e\<^sub>1\<bullet>F{D}:=e\<^sub>2, s\<^sub>0, b\<^sub>0\<rangle> \<rightarrow>* \<langle>THROW NoSuchFieldError, (h\<^sub>2,l\<^sub>2,sh\<^sub>2), b\<^sub>2\<rangle>"
+(*<*)(is "(?x, ?z) \<in> (red P)\<^sup>*")
+proof -
+ let ?y = "(addr a\<bullet>F{D}:=e\<^sub>2, s\<^sub>1, b\<^sub>1)"
+ let ?y' = "(addr a\<bullet>F{D}:=Val v::expr,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2)"
+ have "(?x, ?y) \<in> (red P)\<^sup>*" by(rule FAssReds1[OF e\<^sub>1_steps])
+ also have "(?y, ?y') \<in> (red P)\<^sup>*" by(rule FAssReds2[OF e\<^sub>2_steps])
+ also have "(?y', ?z) \<in> (red P)"
+ using RedFAssNone[OF _ ncF] hC by simp
+ ultimately show ?thesis by simp
+qed
+(*>*)
+
+lemma FAssRedsStatic:
+assumes e\<^sub>1_steps: "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>addr a,s\<^sub>1,b\<^sub>1\<rangle>"
+ and e\<^sub>2_steps: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>1,b\<^sub>1\<rangle> \<rightarrow>* \<langle>Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\<rangle>"
+ and hC: "h\<^sub>2 a = Some(C,fs)" and cF_Static: "P \<turnstile> C has F,Static:t in D"
+shows "P \<turnstile> \<langle>e\<^sub>1\<bullet>F{D}:=e\<^sub>2, s\<^sub>0, b\<^sub>0\<rangle> \<rightarrow>* \<langle>THROW IncompatibleClassChangeError, (h\<^sub>2,l\<^sub>2,sh\<^sub>2), b\<^sub>2\<rangle>"
+(*<*)(is "(?x, ?z) \<in> (red P)\<^sup>*")
+proof -
+ let ?y = "(addr a\<bullet>F{D}:=e\<^sub>2, s\<^sub>1, b\<^sub>1)"
+ let ?y' = "(addr a\<bullet>F{D}:=Val v::expr,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2)"
+ have "(?x, ?y) \<in> (red P)\<^sup>*" by(rule FAssReds1[OF e\<^sub>1_steps])
+ also have "(?y, ?y') \<in> (red P)\<^sup>*" by(rule FAssReds2[OF e\<^sub>2_steps])
+ also have "(?y', ?z) \<in> (red P)"
+ using RedFAssStatic[OF _ cF_Static] hC by simp
+ ultimately show ?thesis by simp
+qed
+(*>*)
+
+subsubsection "SFAss"
+
+lemma SFAssReds:
+ "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D}:=e,s,b\<rangle> \<rightarrow>* \<langle>C\<bullet>\<^sub>sF{D}:=e',s',b'\<rangle>"
+(*<*)
+proof(induct rule: rtrancl_induct3)
+ case refl show ?case by blast
+next
+ case step show ?case
+ by(rule rtrancl_into_rtrancl[OF step(3) SFAssRed[OF step(2)]])
+qed
+(*>*)
+
+lemma SFAssRedsVal:
+assumes e\<^sub>2_steps: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\<rangle>"
+ and cF: "P \<turnstile> C has F,Static:t in D"
+ and shD: "sh\<^sub>2 D = \<lfloor>(sfs,Done)\<rfloor>"
+shows "P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D}:=e\<^sub>2, s\<^sub>0, b\<^sub>0\<rangle> \<rightarrow>* \<langle>unit, (h\<^sub>2,l\<^sub>2,sh\<^sub>2(D\<mapsto>(sfs(F\<mapsto>v), Done))),False\<rangle>"
+(*<*)(is "(?a, ?c) \<in> (red P)\<^sup>*")
+proof -
+ let ?b = "(C\<bullet>\<^sub>sF{D}:=Val v::expr, (h\<^sub>2,l\<^sub>2,sh\<^sub>2), b\<^sub>2)"
+ have "(?a, ?b) \<in> (red P)\<^sup>*" by(rule SFAssReds[OF e\<^sub>2_steps])
+ also have "(?b, ?c) \<in> (red P)\<^sup>*" proof(cases b\<^sub>2)
+ case True
+ then show ?thesis
+ using RedSFAss[THEN r_into_rtrancl, OF cF] shD by simp
+ next
+ case False
+ let ?b' = "(C\<bullet>\<^sub>sF{D}:=Val v::expr, (h\<^sub>2,l\<^sub>2,sh\<^sub>2), True)"
+ have "(?b, ?b') \<in> (red P)\<^sup>*"
+ using False SFAssInitDoneRed[THEN converse_rtrancl_into_rtrancl, OF cF] shD
+ by simp
+ also have "(?b', ?c) \<in> (red P)\<^sup>*"
+ using RedSFAss[THEN r_into_rtrancl, OF cF] shD by simp
+ ultimately show ?thesis by simp
+ qed
+ ultimately show ?thesis by simp
+qed
+(*>*)
+
+lemma SFAssRedsThrow:
+ "\<lbrakk> P \<turnstile> \<langle>e\<^sub>2,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>throw e,s\<^sub>2,b\<^sub>2\<rangle> \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D}:=e\<^sub>2,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>throw e,s\<^sub>2,b\<^sub>2\<rangle>"
+(*<*)by(rule SFAssReds[THEN rtrancl_into_rtrancl, OF _ red_reds.SFAssThrow])(*>*)
+
+lemma SFAssRedsNone:
+ "\<lbrakk> P \<turnstile> \<langle>e\<^sub>2,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\<rangle>;
+ \<not>(\<exists>b t. P \<turnstile> C has F,b:t in D) \<rbrakk> \<Longrightarrow>
+ P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D}:=e\<^sub>2,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>THROW NoSuchFieldError, (h\<^sub>2,l\<^sub>2,sh\<^sub>2),False\<rangle>"
+(*<*)by(rule SFAssReds[THEN rtrancl_into_rtrancl, OF _ RedSFAssNone])(*>*)
+
+lemma SFAssRedsNonStatic:
+ "\<lbrakk> P \<turnstile> \<langle>e\<^sub>2,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\<rangle>;
+ P \<turnstile> C has F,NonStatic:t in D \<rbrakk> \<Longrightarrow>
+ P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D}:=e\<^sub>2,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>THROW IncompatibleClassChangeError,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),False\<rangle>"
+(*<*)by(rule SFAssReds[THEN rtrancl_into_rtrancl, OF _ RedSFAssNonStatic])(*>*)
+
+lemma SFAssInitReds:
+assumes e\<^sub>2_steps: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),False\<rangle>"
+ and cF: "P \<turnstile> C has F,Static:t in D"
+ and nDone: "\<nexists>sfs. sh\<^sub>2 D = Some (sfs, Done)"
+ and INIT_steps: "P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),False\<rangle> \<rightarrow>* \<langle>Val v',(h',l',sh'),b'\<rangle>"
+ and sh'D: "sh' D = Some(sfs,i)"
+ and sfs': "sfs' = sfs(F\<mapsto>v)" and sh'': "sh'' = sh'(D\<mapsto>(sfs',i))"
+shows "P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D}:=e\<^sub>2,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>unit,(h',l',sh''),False\<rangle>"
+(*<*)(is "(?x, ?z) \<in> (red P)\<^sup>*")
+proof -
+ let ?y = "(C\<bullet>\<^sub>sF{D} := Val v::expr,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False)"
+ let ?y' = "(INIT D ([D],False) \<leftarrow> C\<bullet>\<^sub>sF{D} := Val v::expr,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False)"
+ let ?y'' = "(C\<bullet>\<^sub>sF{D} := Val v::expr,(h', l', sh'),icheck P D (C\<bullet>\<^sub>sF{D} := Val v::expr))"
+ have icheck: "icheck P D (C\<bullet>\<^sub>sF{D} := Val v::expr)" using cF by fastforce
+ then have y''z: "(?y'',?z) \<in> (red P)"
+ using RedSFAss[OF cF _ sfs' sh''] sh'D by simp
+ have "(?x, ?y) \<in> (red P)\<^sup>*" by(rule SFAssReds[OF e\<^sub>2_steps])
+ also have "(?y, ?y') \<in> (red P)\<^sup>*"
+ using SFAssInitRed[THEN converse_rtrancl_into_rtrancl, OF cF] nDone
+ by simp
+ also have "(?y', ?z) \<in> (red P)\<^sup>*"
+ using InitSeqReds[OF INIT_steps y''z[THEN r_into_rtrancl]] by simp
+ ultimately show ?thesis by simp
+qed
+(*>*)
+
+lemma SFAssInitThrowReds:
+assumes e\<^sub>2_steps: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>Val v,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),False\<rangle>"
+ and cF: "P \<turnstile> C has F,Static:t in D"
+ and nDone: "\<nexists>sfs. sh\<^sub>2 D = Some (sfs, Done)"
+ and INIT_steps: "P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),False\<rangle> \<rightarrow>* \<langle>throw a,s',b'\<rangle>"
+shows "P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D}:=e\<^sub>2,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>throw a,s',b'\<rangle>"
+(*<*)(is "(?x, ?z) \<in> (red P)\<^sup>*")
+proof -
+ let ?y = "(C\<bullet>\<^sub>sF{D} := Val v::expr,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False)"
+ let ?y' = "(INIT D ([D],False) \<leftarrow> C\<bullet>\<^sub>sF{D} := Val v::expr,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False)"
+ have "(?x, ?y) \<in> (red P)\<^sup>*" by(rule SFAssReds[OF e\<^sub>2_steps])
+ also have "(?y, ?y') \<in> (red P)\<^sup>*"
+ using SFAssInitRed[THEN converse_rtrancl_into_rtrancl, OF cF] nDone
+ by simp
+ also have "(?y', ?z) \<in> (red P)\<^sup>*"
+ using InitSeqThrowReds[OF INIT_steps] by simp
+ ultimately show ?thesis by simp
+qed
+(*>*)
+
+subsubsection";;"
+
+lemma SeqReds:
+ "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>e;;e\<^sub>2, s,b\<rangle> \<rightarrow>* \<langle>e';;e\<^sub>2, s',b'\<rangle>"
+(*<*)
+proof(induct rule: rtrancl_induct3)
+ case refl show ?case by blast
+next
+ case step show ?case
+ by(rule rtrancl_into_rtrancl[OF step(3) SeqRed[OF step(2)]])
+qed
+(*>*)
+
+lemma SeqRedsThrow:
+ "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>throw e',s',b'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>e;;e\<^sub>2, s,b\<rangle> \<rightarrow>* \<langle>throw e', s',b'\<rangle>"
+(*<*)by(rule SeqReds[THEN rtrancl_into_rtrancl, OF _ red_reds.SeqThrow])(*>*)
+
+lemma SeqReds2:
+assumes e\<^sub>1_steps: "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>Val v\<^sub>1,s\<^sub>1,b\<^sub>1\<rangle>"
+ and e\<^sub>2_steps: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>1,b\<^sub>1\<rangle> \<rightarrow>* \<langle>e\<^sub>2',s\<^sub>2,b\<^sub>2\<rangle>"
+shows "P \<turnstile> \<langle>e\<^sub>1;;e\<^sub>2, s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>e\<^sub>2',s\<^sub>2,b\<^sub>2\<rangle>"
+(*<*)(is "(?x, ?z) \<in> (red P)\<^sup>*")
+proof -
+ let ?y = "(Val v\<^sub>1;; e\<^sub>2,s\<^sub>1,b\<^sub>1)"
+ have "(?x, ?y) \<in> (red P)\<^sup>*" by(rule SeqReds[OF e\<^sub>1_steps])
+ also have "(?y, ?z) \<in> (red P)\<^sup>*"
+ by(rule RedSeq[THEN converse_rtrancl_into_rtrancl, OF e\<^sub>2_steps])
+ ultimately show ?thesis by simp
+qed
+(*>*)
+
+
+subsubsection"If"
+
+lemma CondReds:
+ "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>if (e) e\<^sub>1 else e\<^sub>2,s,b\<rangle> \<rightarrow>* \<langle>if (e') e\<^sub>1 else e\<^sub>2,s',b'\<rangle>"
+(*<*)
+proof(induct rule: rtrancl_induct3)
+ case refl show ?case by blast
+next
+ case step show ?case
+ by(rule rtrancl_into_rtrancl[OF step(3) CondRed[OF step(2)]])
+qed
+(*>*)
+
+lemma CondRedsThrow:
+ "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>throw a,s',b'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>if (e) e\<^sub>1 else e\<^sub>2, s,b\<rangle> \<rightarrow>* \<langle>throw a,s',b'\<rangle>"
+(*<*)by(rule CondReds[THEN rtrancl_into_rtrancl, OF _ red_reds.CondThrow])(*>*)
+
+lemma CondReds2T:
+assumes e_steps: "P \<turnstile> \<langle>e,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>true,s\<^sub>1,b\<^sub>1\<rangle>"
+ and e\<^sub>1_steps: "P \<turnstile> \<langle>e\<^sub>1, s\<^sub>1,b\<^sub>1\<rangle> \<rightarrow>* \<langle>e',s\<^sub>2,b\<^sub>2\<rangle>"
+shows "P \<turnstile> \<langle>if (e) e\<^sub>1 else e\<^sub>2, s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>e',s\<^sub>2,b\<^sub>2\<rangle>"
+(*<*)(is "(?x, ?z) \<in> (red P)\<^sup>*")
+proof -
+ let ?y = "(if (true) e\<^sub>1 else e\<^sub>2,s\<^sub>1,b\<^sub>1)"
+ have "(?x, ?y) \<in> (red P)\<^sup>*" by(rule CondReds[OF e_steps])
+ also have "(?y, ?z) \<in> (red P)\<^sup>*"
+ by(rule RedCondT[THEN converse_rtrancl_into_rtrancl, OF e\<^sub>1_steps])
+ ultimately show ?thesis by simp
+qed
+(*>*)
+
+lemma CondReds2F:
+assumes e_steps: "P \<turnstile> \<langle>e,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>false,s\<^sub>1,b\<^sub>1\<rangle>"
+ and e\<^sub>2_steps: "P \<turnstile> \<langle>e\<^sub>2, s\<^sub>1,b\<^sub>1\<rangle> \<rightarrow>* \<langle>e',s\<^sub>2,b\<^sub>2\<rangle>"
+shows "P \<turnstile> \<langle>if (e) e\<^sub>1 else e\<^sub>2, s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>e',s\<^sub>2,b\<^sub>2\<rangle>"
+(*<*)(is "(?x, ?z) \<in> (red P)\<^sup>*")
+proof -
+ let ?y = "(if (false) e\<^sub>1 else e\<^sub>2,s\<^sub>1,b\<^sub>1)"
+ have "(?x, ?y) \<in> (red P)\<^sup>*" by(rule CondReds[OF e_steps])
+ also have "(?y, ?z) \<in> (red P)\<^sup>*"
+ by(rule RedCondF[THEN converse_rtrancl_into_rtrancl, OF e\<^sub>2_steps])
+ ultimately show ?thesis by simp
+qed
+(*>*)
+
+
+subsubsection "While"
+
+lemma WhileFReds:
+assumes b_steps: "P \<turnstile> \<langle>b,s,b\<^sub>0\<rangle> \<rightarrow>* \<langle>false,s',b'\<rangle>"
+shows "P \<turnstile> \<langle>while (b) c,s,b\<^sub>0\<rangle> \<rightarrow>* \<langle>unit,s',b'\<rangle>"
+(*<*)
+by(rule RedWhile[THEN converse_rtrancl_into_rtrancl,
+ OF CondReds[THEN rtrancl_into_rtrancl,
+ OF b_steps RedCondF]])
+(*>*)
+
+lemma WhileRedsThrow:
+assumes b_steps: "P \<turnstile> \<langle>b,s,b\<^sub>0\<rangle> \<rightarrow>* \<langle>throw e,s',b'\<rangle>"
+shows "P \<turnstile> \<langle>while (b) c,s,b\<^sub>0\<rangle> \<rightarrow>* \<langle>throw e,s',b'\<rangle>"
+(*<*)
+by(rule RedWhile[THEN converse_rtrancl_into_rtrancl,
+ OF CondReds[THEN rtrancl_into_rtrancl,
+ OF b_steps red_reds.CondThrow]])
+(*>*)
+
+lemma WhileTReds:
+assumes b_steps: "P \<turnstile> \<langle>b,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>true,s\<^sub>1,b\<^sub>1\<rangle>"
+ and c_steps: "P \<turnstile> \<langle>c,s\<^sub>1,b\<^sub>1\<rangle> \<rightarrow>* \<langle>Val v\<^sub>1,s\<^sub>2,b\<^sub>2\<rangle>"
+ and while_steps: "P \<turnstile> \<langle>while (b) c,s\<^sub>2,b\<^sub>2\<rangle> \<rightarrow>* \<langle>e,s\<^sub>3,b\<^sub>3\<rangle>"
+shows "P \<turnstile> \<langle>while (b) c,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>e,s\<^sub>3,b\<^sub>3\<rangle>"
+(*<*)(is "(?a, ?c) \<in> (red P)\<^sup>*")
+proof -
+ let ?b = "(if (b) (c;; while (b) c) else unit,s\<^sub>0,b\<^sub>0)"
+ let ?y = "(if (true) (c;; while (b) c) else unit,s\<^sub>1,b\<^sub>1)"
+ let ?b' = "(c;; while (b) c,s\<^sub>1,b\<^sub>1)"
+ let ?y' = "(Val v\<^sub>1;; while (b) c,s\<^sub>2,b\<^sub>2)"
+ have "(?a, ?b) \<in> (red P)\<^sup>*"
+ using RedWhile[THEN converse_rtrancl_into_rtrancl] by simp
+ also have "(?b, ?y) \<in> (red P)\<^sup>*" by(rule CondReds[OF b_steps])
+ also have "(?y, ?b') \<in> (red P)\<^sup>*"
+ using RedCondT[THEN converse_rtrancl_into_rtrancl] by simp
+ also have "(?b', ?y') \<in> (red P)\<^sup>*" by(rule SeqReds[OF c_steps])
+ also have "(?y', ?c) \<in> (red P)\<^sup>*"
+ by(rule RedSeq[THEN converse_rtrancl_into_rtrancl, OF while_steps])
+ ultimately show ?thesis by simp
+qed
+(*>*)
+
+lemma WhileTRedsThrow:
+assumes b_steps: "P \<turnstile> \<langle>b,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>true,s\<^sub>1,b\<^sub>1\<rangle>"
+ and c_steps: "P \<turnstile> \<langle>c,s\<^sub>1,b\<^sub>1\<rangle> \<rightarrow>* \<langle>throw e,s\<^sub>2,b\<^sub>2\<rangle>"
+shows "P \<turnstile> \<langle>while (b) c,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>throw e,s\<^sub>2,b\<^sub>2\<rangle>"
+(*<*)(is "(?a, ?c) \<in> (red P)\<^sup>*")
+proof -
+ let ?b = "(if (b) (c;; while (b) c) else unit,s\<^sub>0,b\<^sub>0)"
+ let ?y = "(if (true) (c;; while (b) c) else unit,s\<^sub>1,b\<^sub>1)"
+ let ?b' = "(c;; while (b) c,s\<^sub>1,b\<^sub>1)"
+ let ?y' = "(throw e;; while (b) c,s\<^sub>2,b\<^sub>2)"
+ have "(?a, ?b) \<in> (red P)\<^sup>*"
+ using RedWhile[THEN converse_rtrancl_into_rtrancl] by simp
+ also have "(?b, ?y) \<in> (red P)\<^sup>*" by(rule CondReds[OF b_steps])
+ also have "(?y, ?b') \<in> (red P)\<^sup>*"
+ using RedCondT[THEN converse_rtrancl_into_rtrancl] by simp
+ also have "(?b', ?y') \<in> (red P)\<^sup>*" by(rule SeqReds[OF c_steps])
+ also have "(?y', ?c) \<in> (red P)" by(rule red_reds.SeqThrow)
+ ultimately show ?thesis by simp
+qed
+(*>*)
+
+subsubsection"Throw"
+
+lemma ThrowReds:
+ "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>throw e,s,b\<rangle> \<rightarrow>* \<langle>throw e',s',b'\<rangle>"
+(*<*)
+proof(induct rule: rtrancl_induct3)
+ case refl show ?case by blast
+next
+ case step show ?case
+ by(rule rtrancl_into_rtrancl[OF step(3) ThrowRed[OF step(2)]])
+qed
+(*>*)
+
+lemma ThrowRedsNull:
+ "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>null,s',b'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>throw e,s,b\<rangle> \<rightarrow>* \<langle>THROW NullPointer,s',b'\<rangle>"
+(*<*)by(rule ThrowReds[THEN rtrancl_into_rtrancl, OF _ RedThrowNull])(*>*)
+
+lemma ThrowRedsThrow:
+ "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>throw a,s',b'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>throw e,s,b\<rangle> \<rightarrow>* \<langle>throw a,s',b'\<rangle>"
+(*<*)by(rule ThrowReds[THEN rtrancl_into_rtrancl, OF _ red_reds.ThrowThrow])(*>*)
+
+subsubsection "InitBlock"
+
+lemma InitBlockReds_aux:
+ "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle> \<Longrightarrow>
+ \<forall>h l sh h' l' sh' v. s = (h,l(V\<mapsto>v),sh) \<longrightarrow> s' = (h',l',sh') \<longrightarrow>
+ P \<turnstile> \<langle>{V:T := Val v; e},(h,l,sh),b\<rangle> \<rightarrow>* \<langle>{V:T := Val(the(l' V)); e'},(h',l'(V:=(l V)),sh'),b'\<rangle>"
+(*<*)
+proof(induct rule: converse_rtrancl_induct3)
+ case refl then show ?case
+ by(fastforce simp: fun_upd_same simp del:fun_upd_apply)
+next
+ case (step e0 s0 b0 e1 s1 b1)
+ obtain h1 l1 sh1 where s1[simp]: "s1 = (h1, l1, sh1)" by(cases s1) simp
+ { fix h0 l0 sh0 h2 l2 sh2 v0
+ assume [simp]: "s0 = (h0, l0(V \<mapsto> v0), sh0)" and s'[simp]: "s' = (h2, l2, sh2)"
+ then have "V \<in> dom l1" using step(1) by(auto dest!: red_lcl_incr)
+ then obtain v1 where l1V[simp]: "l1 V = \<lfloor>v1\<rfloor>" by blast
+
+ let ?a = "({V:T; V:=Val v0;; e0},(h0, l0, sh0),b0)"
+ let ?b = "({V:T; V:=Val v1;; e1},(h1, l1(V := l0 V), sh1),b1)"
+ let ?c = "({V:T; V:=Val (the (l2 V));; e'},(h2, l2(V := l0 V), sh2),b')"
+ let ?l = "l1(V := l0 V)" and ?v = v1
+
+ have e0_steps: "P \<turnstile> \<langle>e0,(h0, l0(V \<mapsto> v0), sh0),b0\<rangle> \<rightarrow> \<langle>e1,(h1, l1, sh1),b1\<rangle>"
+ using step(1) by simp
+
+ have lv: "\<And>l v. l1 = l(V \<mapsto> v) \<longrightarrow>
+ P \<turnstile> \<langle>{V:T; V:=Val v;; e1},(h1, l, sh1),b1\<rangle> \<rightarrow>*
+ \<langle>{V:T; V:=Val (the (l2 V));; e'},(h2, l2(V := l V), sh2),b'\<rangle>"
+ using step(3) s' s1 by blast
+ moreover have "l1 = ?l(V \<mapsto> ?v)" by(rule ext) (simp add:fun_upd_def)
+ ultimately have "(?b, ?c) \<in> (red P)\<^sup>*" using lv[of ?l ?v] by simp
+ then have "(?a, ?c) \<in> (red P)\<^sup>*"
+ by(rule InitBlockRed[THEN converse_rtrancl_into_rtrancl, OF e0_steps l1V])
+ }
+ then show ?case by blast
+qed
+(*>*)
+
+lemma InitBlockReds:
+ "P \<turnstile> \<langle>e, (h,l(V\<mapsto>v),sh),b\<rangle> \<rightarrow>* \<langle>e', (h',l',sh'),b'\<rangle> \<Longrightarrow>
+ P \<turnstile> \<langle>{V:T := Val v; e}, (h,l,sh),b\<rangle> \<rightarrow>* \<langle>{V:T := Val(the(l' V)); e'}, (h',l'(V:=(l V)),sh'),b'\<rangle>"
+(*<*)by(blast dest:InitBlockReds_aux)(*>*)
+
+lemma InitBlockRedsFinal:
+ "\<lbrakk> P \<turnstile> \<langle>e,(h,l(V\<mapsto>v),sh),b\<rangle> \<rightarrow>* \<langle>e',(h',l',sh'),b'\<rangle>; final e' \<rbrakk> \<Longrightarrow>
+ P \<turnstile> \<langle>{V:T := Val v; e},(h,l,sh),b\<rangle> \<rightarrow>* \<langle>e',(h', l'(V := l V),sh'),b'\<rangle>"
+(*<*)
+by(fast elim!:InitBlockReds[THEN rtrancl_into_rtrancl] finalE
+ intro:RedInitBlock InitBlockThrow)
+(*>*)
+
+
+subsubsection "Block"
+
+lemmas converse_rtranclE3 = converse_rtranclE [of "(xa,xb,xc)" "(za,zb,zc)", split_rule]
+
+lemma BlockRedsFinal:
+assumes reds: "P \<turnstile> \<langle>e\<^sub>0,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>e\<^sub>2,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\<rangle>" and fin: "final e\<^sub>2"
+shows "\<And>h\<^sub>0 l\<^sub>0 sh\<^sub>0. s\<^sub>0 = (h\<^sub>0,l\<^sub>0(V:=None),sh\<^sub>0) \<Longrightarrow> P \<turnstile> \<langle>{V:T; e\<^sub>0},(h\<^sub>0,l\<^sub>0,sh\<^sub>0),b\<^sub>0\<rangle> \<rightarrow>* \<langle>e\<^sub>2,(h\<^sub>2,l\<^sub>2(V:=l\<^sub>0 V),sh\<^sub>2),b\<^sub>2\<rangle>"
+(*<*)
+using reds
+proof (induct rule:converse_rtrancl_induct3)
+ case refl thus ?case
+ by(fastforce intro:finalE[OF fin] RedBlock BlockThrow
+ simp del:fun_upd_apply)
+next
+ case (step e\<^sub>0 s\<^sub>0 b\<^sub>0 e\<^sub>1 s\<^sub>1 b\<^sub>1)
+ have red: "P \<turnstile> \<langle>e\<^sub>0,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow> \<langle>e\<^sub>1,s\<^sub>1,b\<^sub>1\<rangle>"
+ and reds: "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>1,b\<^sub>1\<rangle> \<rightarrow>* \<langle>e\<^sub>2,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\<rangle>"
+ and IH: "\<And>h l sh. s\<^sub>1 = (h,l(V := None),sh)
+ \<Longrightarrow> P \<turnstile> \<langle>{V:T; e\<^sub>1},(h,l,sh),b\<^sub>1\<rangle> \<rightarrow>* \<langle>e\<^sub>2,(h\<^sub>2, l\<^sub>2(V := l V),sh\<^sub>2),b\<^sub>2\<rangle>"
+ and s\<^sub>0: "s\<^sub>0 = (h\<^sub>0, l\<^sub>0(V := None),sh\<^sub>0)" by fact+
+ obtain h\<^sub>1 l\<^sub>1 sh\<^sub>1 where s\<^sub>1: "s\<^sub>1 = (h\<^sub>1,l\<^sub>1,sh\<^sub>1)"
+ using prod_cases3 by blast
+ show ?case
+ proof cases
+ assume "assigned V e\<^sub>0"
+ then obtain v e where e\<^sub>0: "e\<^sub>0 = V := Val v;; e"
+ by (unfold assigned_def)blast
+ from red e\<^sub>0 s\<^sub>0 have e\<^sub>1: "e\<^sub>1 = unit;;e" and s\<^sub>1: "s\<^sub>1 = (h\<^sub>0, l\<^sub>0(V \<mapsto> v),sh\<^sub>0)" and b\<^sub>1: "b\<^sub>1 = b\<^sub>0"
+ by auto
+ from e\<^sub>1 fin have "e\<^sub>1 \<noteq> e\<^sub>2" by (auto simp:final_def)
+ then obtain e' s' b' where red1: "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>1,b\<^sub>1\<rangle> \<rightarrow> \<langle>e',s',b'\<rangle>"
+ and reds': "P \<turnstile> \<langle>e',s',b'\<rangle> \<rightarrow>* \<langle>e\<^sub>2,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\<rangle>"
+ using converse_rtranclE3[OF reds] by blast
+ from red1 e\<^sub>1 b\<^sub>1 have es': "e' = e" "s' = s\<^sub>1" "b' = b\<^sub>0" by auto
+ show ?case using e\<^sub>0 s\<^sub>1 es' reds'
+ by(auto intro!: InitBlockRedsFinal[OF _ fin] simp del:fun_upd_apply)
+ next
+ assume unass: "\<not> assigned V e\<^sub>0"
+ show ?thesis
+ proof (cases "l\<^sub>1 V")
+ assume None: "l\<^sub>1 V = None"
+ hence "P \<turnstile> \<langle>{V:T; e\<^sub>0},(h\<^sub>0,l\<^sub>0,sh\<^sub>0),b\<^sub>0\<rangle> \<rightarrow> \<langle>{V:T; e\<^sub>1},(h\<^sub>1, l\<^sub>1(V := l\<^sub>0 V),sh\<^sub>1),b\<^sub>1\<rangle>"
+ using s\<^sub>0 s\<^sub>1 red by(simp add: BlockRedNone[OF _ _ unass])
+ moreover
+ have "P \<turnstile> \<langle>{V:T; e\<^sub>1},(h\<^sub>1, l\<^sub>1(V := l\<^sub>0 V),sh\<^sub>1),b\<^sub>1\<rangle> \<rightarrow>* \<langle>e\<^sub>2,(h\<^sub>2, l\<^sub>2(V := l\<^sub>0 V),sh\<^sub>2),b\<^sub>2\<rangle>"
+ using IH[of _ "l\<^sub>1(V := l\<^sub>0 V)"] s\<^sub>1 None by(simp add:fun_upd_idem)
+ ultimately show ?case by(rule converse_rtrancl_into_rtrancl)
+ next
+ fix v assume Some: "l\<^sub>1 V = Some v"
+ hence "P \<turnstile> \<langle>{V:T;e\<^sub>0},(h\<^sub>0,l\<^sub>0,sh\<^sub>0),b\<^sub>0\<rangle> \<rightarrow> \<langle>{V:T := Val v; e\<^sub>1},(h\<^sub>1,l\<^sub>1(V := l\<^sub>0 V),sh\<^sub>1),b\<^sub>1\<rangle>"
+ using s\<^sub>0 s\<^sub>1 red by(simp add: BlockRedSome[OF _ _ unass])
+ moreover
+ have "P \<turnstile> \<langle>{V:T := Val v; e\<^sub>1},(h\<^sub>1,l\<^sub>1(V:= l\<^sub>0 V),sh\<^sub>1),b\<^sub>1\<rangle> \<rightarrow>*
+ \<langle>e\<^sub>2,(h\<^sub>2,l\<^sub>2(V:=l\<^sub>0 V),sh\<^sub>2),b\<^sub>2\<rangle>"
+ using InitBlockRedsFinal[OF _ fin,of _ _ "l\<^sub>1(V:=l\<^sub>0 V)" V]
+ Some reds s\<^sub>1 by(simp add:fun_upd_idem)
+ ultimately show ?case by(rule converse_rtrancl_into_rtrancl)
+ qed
+ qed
+qed
+(*>*)
+
+
+subsubsection "try-catch"
+
+lemma TryReds:
+ "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>try e catch(C V) e\<^sub>2,s,b\<rangle> \<rightarrow>* \<langle>try e' catch(C V) e\<^sub>2,s',b'\<rangle>"
+(*<*)
+proof(induct rule: rtrancl_induct3)
+ case refl show ?case by blast
+next
+ case step show ?case
+ by(rule rtrancl_into_rtrancl[OF step(3) TryRed[OF step(2)]])
+qed
+(*>*)
+
+lemma TryRedsVal:
+ "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>Val v,s',b'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>try e catch(C V) e\<^sub>2,s,b\<rangle> \<rightarrow>* \<langle>Val v,s',b'\<rangle>"
+(*<*)by(rule TryReds[THEN rtrancl_into_rtrancl, OF _ RedTry])(*>*)
+
+lemma TryCatchRedsFinal:
+assumes e\<^sub>1_steps: "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>Throw a,(h\<^sub>1,l\<^sub>1,sh\<^sub>1),b\<^sub>1\<rangle>"
+ and h\<^sub>1a: "h\<^sub>1 a = Some(D,fs)" and sub: "P \<turnstile> D \<preceq>\<^sup>* C"
+ and e\<^sub>2_steps: "P \<turnstile> \<langle>e\<^sub>2, (h\<^sub>1, l\<^sub>1(V \<mapsto> Addr a),sh\<^sub>1),b\<^sub>1\<rangle> \<rightarrow>* \<langle>e\<^sub>2', (h\<^sub>2,l\<^sub>2,sh\<^sub>2), b\<^sub>2\<rangle>"
+ and final: "final e\<^sub>2'"
+shows "P \<turnstile> \<langle>try e\<^sub>1 catch(C V) e\<^sub>2, s\<^sub>0, b\<^sub>0\<rangle> \<rightarrow>* \<langle>e\<^sub>2', (h\<^sub>2, (l\<^sub>2::locals)(V := l\<^sub>1 V),sh\<^sub>2),b\<^sub>2\<rangle>"
+(*<*)(is "(?x, ?z) \<in> (red P)\<^sup>*")
+proof -
+ let ?y = "(try Throw a catch(C V) e\<^sub>2,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),b\<^sub>1)"
+ let ?b = "({V:Class C; V:=addr a;; e\<^sub>2},(h\<^sub>1, l\<^sub>1, sh\<^sub>1),b\<^sub>1)"
+ have bz: "(?b, ?z) \<in> (red P)\<^sup>*"
+ by(rule InitBlockRedsFinal[OF e\<^sub>2_steps final])
+ have hp: "hp (h\<^sub>1, l\<^sub>1, sh\<^sub>1) a = \<lfloor>(D, fs)\<rfloor>" using h\<^sub>1a by simp
+ have "(?x, ?y) \<in> (red P)\<^sup>*" by(rule TryReds[OF e\<^sub>1_steps])
+ also have "(?y, ?z) \<in> (red P)\<^sup>*"
+ by(rule RedTryCatch[THEN converse_rtrancl_into_rtrancl, OF hp sub bz])
+ ultimately show ?thesis by simp
+qed
+(*>*)
+
+lemma TryRedsFail:
+ "\<lbrakk> P \<turnstile> \<langle>e\<^sub>1,s,b\<rangle> \<rightarrow>* \<langle>Throw a,(h,l,sh),b'\<rangle>; h a = Some(D,fs); \<not> P \<turnstile> D \<preceq>\<^sup>* C \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>try e\<^sub>1 catch(C V) e\<^sub>2,s,b\<rangle> \<rightarrow>* \<langle>Throw a,(h,l,sh),b'\<rangle>"
+(*<*)by(fastforce intro!: TryReds[THEN rtrancl_into_rtrancl, OF _ RedTryFail])(*>*)
+
+subsubsection "List"
+
+lemma ListReds1:
+ "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>e#es,s,b\<rangle> [\<rightarrow>]* \<langle>e' # es,s',b'\<rangle>"
+(*<*)
+proof(induct rule: rtrancl_induct3)
+ case refl show ?case by blast
+next
+ case step show ?case
+ by(rule rtrancl_into_rtrancl[OF step(3) ListRed1[OF step(2)]])
+qed
+(*>*)
+
+lemma ListReds2:
+ "P \<turnstile> \<langle>es,s,b\<rangle> [\<rightarrow>]* \<langle>es',s',b'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>Val v # es,s,b\<rangle> [\<rightarrow>]* \<langle>Val v # es',s',b'\<rangle>"
+(*<*)
+proof(induct rule: rtrancl_induct3)
+ case refl show ?case by blast
+next
+ case step show ?case
+ by(rule rtrancl_into_rtrancl[OF step(3) ListRed2[OF step(2)]])
+qed
+(*>*)
+
+lemma ListRedsVal:
+ "\<lbrakk> P \<turnstile> \<langle>e,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>Val v,s\<^sub>1,b\<^sub>1\<rangle>; P \<turnstile> \<langle>es,s\<^sub>1,b\<^sub>1\<rangle> [\<rightarrow>]* \<langle>es',s\<^sub>2,b\<^sub>2\<rangle> \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>e#es,s\<^sub>0,b\<^sub>0\<rangle> [\<rightarrow>]* \<langle>Val v # es',s\<^sub>2,b\<^sub>2\<rangle>"
+(*<*)by(rule rtrancl_trans[OF ListReds1 ListReds2])(*>*)
+
+subsubsection"Call"
+
+text\<open> First a few lemmas on what happens to free variables during redction. \<close>
+
+lemma assumes wf: "wwf_J_prog P"
+shows Red_fv: "P \<turnstile> \<langle>e,(h,l,sh),b\<rangle> \<rightarrow> \<langle>e',(h',l',sh'),b'\<rangle> \<Longrightarrow> fv e' \<subseteq> fv e"
+ and "P \<turnstile> \<langle>es,(h,l,sh),b\<rangle> [\<rightarrow>] \<langle>es',(h',l',sh'),b'\<rangle> \<Longrightarrow> fvs es' \<subseteq> fvs es"
+(*<*)
+proof (induct rule:red_reds_inducts)
+ case (RedCall h a C fs M Ts T pns body D vs l sh b)
+ hence "fv body \<subseteq> {this} \<union> set pns"
+ using assms by(fastforce dest!:sees_wf_mdecl simp:wf_mdecl_def)
+ with RedCall.hyps show ?case by fastforce
+next
+ case (RedSCall C M Ts T pns body D vs)
+ hence "fv body \<subseteq> set pns"
+ using assms by(fastforce dest!:sees_wf_mdecl simp:wf_mdecl_def)
+ with RedSCall.hyps show ?case by fastforce
+qed auto
+(*>*)
+
+
+lemma Red_dom_lcl:
+ "P \<turnstile> \<langle>e,(h,l,sh),b\<rangle> \<rightarrow> \<langle>e',(h',l',sh'),b'\<rangle> \<Longrightarrow> dom l' \<subseteq> dom l \<union> fv e" and
+ "P \<turnstile> \<langle>es,(h,l,sh),b\<rangle> [\<rightarrow>] \<langle>es',(h',l',sh'),b'\<rangle> \<Longrightarrow> dom l' \<subseteq> dom l \<union> fvs es"
+(*<*)
+proof (induct rule:red_reds_inducts)
+ case RedLAss thus ?case by(force split:if_splits)
+next
+ case CallParams thus ?case by(force split:if_splits)
+next
+ case BlockRedNone thus ?case by clarsimp (fastforce split:if_splits)
+next
+ case BlockRedSome thus ?case by clarsimp (fastforce split:if_splits)
+next
+ case InitBlockRed thus ?case by clarsimp (fastforce split:if_splits)
+qed auto
+(*>*)
+
+lemma Reds_dom_lcl:
+assumes wf: "wwf_J_prog P"
+shows "P \<turnstile> \<langle>e,(h,l,sh),b\<rangle> \<rightarrow>* \<langle>e',(h',l',sh'),b'\<rangle> \<Longrightarrow> dom l' \<subseteq> dom l \<union> fv e"
+(*<*)
+proof(induct rule: converse_rtrancl_induct_red)
+ case 1 then show ?case by blast
+next
+ case 2 then show ?case using wf by(blast dest: Red_fv Red_dom_lcl)
+qed
+(*>*)
+
+text\<open> Now a few lemmas on the behaviour of blocks during reduction. \<close>
+
+lemma override_on_upd_lemma:
+ "(override_on f (g(a\<mapsto>b)) A)(a := g a) = override_on f g (insert a A)"
+(*<*)by(rule ext) (simp add:override_on_def)
+
+declare fun_upd_apply[simp del] map_upds_twist[simp del]
+(*>*)
+
+
+lemma blocksReds:
+ "\<And>l. \<lbrakk> length Vs = length Ts; length vs = length Ts; distinct Vs;
+ P \<turnstile> \<langle>e, (h,l(Vs [\<mapsto>] vs),sh),b\<rangle> \<rightarrow>* \<langle>e', (h',l',sh'),b'\<rangle> \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>blocks(Vs,Ts,vs,e), (h,l,sh),b\<rangle> \<rightarrow>* \<langle>blocks(Vs,Ts,map (the \<circ> l') Vs,e'), (h',override_on l' l (set Vs),sh'),b'\<rangle>"
+(*<*)
+proof(induct Vs Ts vs e rule:blocks_induct)
+ case (1 V Vs T Ts v vs e) show ?case
+ using InitBlockReds[OF "1.hyps"[of "l(V\<mapsto>v)"]] "1.prems"
+ by(auto simp:override_on_upd_lemma)
+qed auto
+(*>*)
+
+
+lemma blocksFinal:
+ "\<And>l. \<lbrakk> length Vs = length Ts; length vs = length Ts; final e \<rbrakk> \<Longrightarrow>
+ P \<turnstile> \<langle>blocks(Vs,Ts,vs,e), (h,l,sh),b\<rangle> \<rightarrow>* \<langle>e, (h,l,sh),b\<rangle>"
+(*<*)
+proof(induct Vs Ts vs e rule:blocks_induct)
+ case 1
+ show ?case using "1.prems" InitBlockReds[OF "1.hyps"]
+ by(fastforce elim!:finalE elim: rtrancl_into_rtrancl[OF _ RedInitBlock]
+ rtrancl_into_rtrancl[OF _ InitBlockThrow])
+qed auto
+(*>*)
+
+
+lemma blocksRedsFinal:
+assumes wf: "length Vs = length Ts" "length vs = length Ts" "distinct Vs"
+ and reds: "P \<turnstile> \<langle>e, (h,l(Vs [\<mapsto>] vs),sh),b\<rangle> \<rightarrow>* \<langle>e', (h',l',sh'),b'\<rangle>"
+ and fin: "final e'" and l'': "l'' = override_on l' l (set Vs)"
+shows "P \<turnstile> \<langle>blocks(Vs,Ts,vs,e), (h,l,sh),b\<rangle> \<rightarrow>* \<langle>e', (h',l'',sh'),b'\<rangle>"
+(*<*)
+proof -
+ let ?bv = "blocks(Vs,Ts,map (the o l') Vs,e')"
+ have "P \<turnstile> \<langle>blocks(Vs,Ts,vs,e), (h,l,sh),b\<rangle> \<rightarrow>* \<langle>?bv, (h',l'',sh'),b'\<rangle>"
+ using l'' by simp (rule blocksReds[OF wf reds])
+ also have "P \<turnstile> \<langle>?bv, (h',l'',sh'),b'\<rangle> \<rightarrow>* \<langle>e', (h',l'',sh'),b'\<rangle>"
+ using wf by(fastforce intro:blocksFinal fin)
+ finally show ?thesis .
+qed
+(*>*)
+
+text\<open> An now the actual method call reduction lemmas. \<close>
+
+lemma CallRedsObj:
+ "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>e\<bullet>M(es),s,b\<rangle> \<rightarrow>* \<langle>e'\<bullet>M(es),s',b'\<rangle>"
+(*<*)
+proof(induct rule: rtrancl_induct3)
+ case refl show ?case by blast
+next
+ case step show ?case
+ by(rule rtrancl_into_rtrancl[OF step(3) CallObj[OF step(2)]])
+qed
+(*>*)
+
+
+lemma CallRedsParams:
+ "P \<turnstile> \<langle>es,s,b\<rangle> [\<rightarrow>]* \<langle>es',s',b'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>(Val v)\<bullet>M(es),s,b\<rangle> \<rightarrow>* \<langle>(Val v)\<bullet>M(es'),s',b'\<rangle>"
+(*<*)
+proof(induct rule: rtrancl_induct3)
+ case refl show ?case by blast
+next
+ case step show ?case
+ by(rule rtrancl_into_rtrancl[OF step(3) CallParams[OF step(2)]])
+qed
+(*>*)
+
+
+lemma CallRedsFinal:
+assumes wwf: "wwf_J_prog P"
+and "P \<turnstile> \<langle>e,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>addr a,s\<^sub>1,b\<^sub>1\<rangle>"
+ "P \<turnstile> \<langle>es,s\<^sub>1,b\<^sub>1\<rangle> [\<rightarrow>]* \<langle>map Val vs,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\<rangle>"
+ "h\<^sub>2 a = Some(C,fs)" "P \<turnstile> C sees M,NonStatic:Ts\<rightarrow>T = (pns,body) in D"
+ "size vs = size pns"
+and l\<^sub>2': "l\<^sub>2' = [this \<mapsto> Addr a, pns[\<mapsto>]vs]"
+and body: "P \<turnstile> \<langle>body,(h\<^sub>2,l\<^sub>2',sh\<^sub>2),b\<^sub>2\<rangle> \<rightarrow>* \<langle>ef,(h\<^sub>3,l\<^sub>3,sh\<^sub>3),b\<^sub>3\<rangle>"
+and "final ef"
+shows "P \<turnstile> \<langle>e\<bullet>M(es), s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>ef,(h\<^sub>3,l\<^sub>2,sh\<^sub>3),b\<^sub>3\<rangle>"
+(*<*)
+proof -
+ have wf: "size Ts = size pns \<and> distinct pns \<and> this \<notin> set pns"
+ and wt: "fv body \<subseteq> {this} \<union> set pns"
+ using assms by(fastforce dest!:sees_wf_mdecl simp:wf_mdecl_def)+
+ from body[THEN Red_lcl_add, of l\<^sub>2]
+ have body': "P \<turnstile> \<langle>body,(h\<^sub>2,l\<^sub>2(this\<mapsto> Addr a, pns[\<mapsto>]vs),sh\<^sub>2),b\<^sub>2\<rangle> \<rightarrow>* \<langle>ef,(h\<^sub>3,l\<^sub>2++l\<^sub>3,sh\<^sub>3),b\<^sub>3\<rangle>"
+ by (simp add:l\<^sub>2')
+ have "dom l\<^sub>3 \<subseteq> {this} \<union> set pns"
+ using Reds_dom_lcl[OF wwf body] wt l\<^sub>2' set_take_subset by force
+ hence eql\<^sub>2: "override_on (l\<^sub>2++l\<^sub>3) l\<^sub>2 ({this} \<union> set pns) = l\<^sub>2"
+ by(fastforce simp add:map_add_def override_on_def fun_eq_iff)
+ have "P \<turnstile> \<langle>e\<bullet>M(es),s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>(addr a)\<bullet>M(es),s\<^sub>1,b\<^sub>1\<rangle>" by(rule CallRedsObj)(rule assms(2))
+ also have "P \<turnstile> \<langle>(addr a)\<bullet>M(es),s\<^sub>1,b\<^sub>1\<rangle> \<rightarrow>*
+ \<langle>(addr a)\<bullet>M(map Val vs),(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\<rangle>"
+ by(rule CallRedsParams)(rule assms(3))
+ also have "P \<turnstile> \<langle>(addr a)\<bullet>M(map Val vs), (h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\<rangle> \<rightarrow>
+ \<langle>blocks(this#pns, Class D#Ts, Addr a#vs, body), (h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\<rangle>"
+ by(rule RedCall)(auto simp: assms wf, rule assms(5))
+ also (rtrancl_into_rtrancl) have "P \<turnstile> \<langle>blocks(this#pns, Class D#Ts, Addr a#vs, body), (h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\<rangle>
+ \<rightarrow>* \<langle>ef,(h\<^sub>3,override_on (l\<^sub>2++l\<^sub>3) l\<^sub>2 ({this} \<union> set pns),sh\<^sub>3),b\<^sub>3\<rangle>"
+ by(rule blocksRedsFinal, insert assms wf body', simp_all)
+ finally show ?thesis using eql\<^sub>2 by simp
+qed
+(*>*)
+
+
+lemma CallRedsThrowParams:
+assumes e_steps: "P \<turnstile> \<langle>e,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>Val v,s\<^sub>1,b\<^sub>1\<rangle>"
+ and es_steps: "P \<turnstile> \<langle>es,s\<^sub>1,b\<^sub>1\<rangle> [\<rightarrow>]* \<langle>map Val vs\<^sub>1 @ throw a # es\<^sub>2,s\<^sub>2,b\<^sub>2\<rangle>"
+shows "P \<turnstile> \<langle>e\<bullet>M(es),s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>throw a,s\<^sub>2,b\<^sub>2\<rangle>"
+(*<*)(is "(?x, ?z) \<in> (red P)\<^sup>*")
+proof -
+ let ?y = "(Val v\<bullet>M(es),s\<^sub>1,b\<^sub>1)"
+ let ?y' = "(Val v\<bullet>M(map Val vs\<^sub>1 @ throw a # es\<^sub>2),s\<^sub>2,b\<^sub>2)"
+ have "(?x, ?y) \<in> (red P)\<^sup>*" by(rule CallRedsObj[OF e_steps])
+ also have "(?y, ?y') \<in> (red P)\<^sup>*" by(rule CallRedsParams[OF es_steps])
+ also have "(?y', ?z) \<in> (red P)\<^sup>*" using CallThrowParams by fast
+ ultimately show ?thesis by simp
+qed
+(*>*)
+
+
+lemma CallRedsThrowObj:
+ "P \<turnstile> \<langle>e,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>throw a,s\<^sub>1,b\<^sub>1\<rangle> \<Longrightarrow> P \<turnstile> \<langle>e\<bullet>M(es),s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>throw a,s\<^sub>1,b\<^sub>1\<rangle>"
+(*<*)by(rule CallRedsObj[THEN rtrancl_into_rtrancl, OF _ CallThrowObj])(*>*)
+
+
+lemma CallRedsNull:
+assumes e_steps: "P \<turnstile> \<langle>e,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>null,s\<^sub>1,b\<^sub>1\<rangle>"
+ and es_steps: "P \<turnstile> \<langle>es,s\<^sub>1,b\<^sub>1\<rangle> [\<rightarrow>]* \<langle>map Val vs,s\<^sub>2,b\<^sub>2\<rangle>"
+shows "P \<turnstile> \<langle>e\<bullet>M(es),s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>THROW NullPointer,s\<^sub>2,b\<^sub>2\<rangle>"
+(*<*)(is "(?x, ?z) \<in> (red P)\<^sup>*")
+proof -
+ let ?y = "(null\<bullet>M(es),s\<^sub>1,b\<^sub>1)"
+ let ?y' = "(null\<bullet>M(map Val vs),s\<^sub>2,b\<^sub>2)"
+ have "(?x, ?y) \<in> (red P)\<^sup>*" by(rule CallRedsObj[OF e_steps])
+ also have "(?y, ?y') \<in> (red P)\<^sup>*" by(rule CallRedsParams[OF es_steps])
+ also have "(?y', ?z) \<in> (red P)" by(rule RedCallNull)
+ ultimately show ?thesis by simp
+qed
+(*>*)
+
+lemma CallRedsNone:
+assumes e_steps: "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>addr a,s\<^sub>1,b\<^sub>1\<rangle>"
+ and es_steps: "P \<turnstile> \<langle>es,s\<^sub>1,b\<^sub>1\<rangle> [\<rightarrow>]* \<langle>map Val vs,s\<^sub>2,b\<^sub>2\<rangle>"
+ and hp\<^sub>2a: "hp s\<^sub>2 a = Some(C,fs)"
+ and ncM: "\<not>(\<exists>b Ts T m D. P \<turnstile> C sees M,b:Ts\<rightarrow>T = m in D)"
+shows "P \<turnstile> \<langle>e\<bullet>M(es),s,b\<rangle> \<rightarrow>* \<langle>THROW NoSuchMethodError,s\<^sub>2,b\<^sub>2\<rangle>"
+(*<*)(is "(?x, ?z) \<in> (red P)\<^sup>*")
+proof -
+ let ?y = "(addr a\<bullet>M(es),s\<^sub>1,b\<^sub>1)"
+ let ?y' = "(addr a\<bullet>M(map Val vs),s\<^sub>2,b\<^sub>2)"
+ have "(?x, ?y) \<in> (red P)\<^sup>*" by(rule CallRedsObj[OF e_steps])
+ also have "(?y, ?y') \<in> (red P)\<^sup>*" by(rule CallRedsParams[OF es_steps])
+ also have "(?y', ?z) \<in> (red P)" using RedCallNone[OF _ ncM] hp\<^sub>2a
+ by(cases s\<^sub>2) simp
+ ultimately show ?thesis by simp
+qed
+(*>*)
+
+lemma CallRedsStatic:
+assumes e_steps: "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>addr a,s\<^sub>1,b\<^sub>1\<rangle>"
+ and es_steps: "P \<turnstile> \<langle>es,s\<^sub>1,b\<^sub>1\<rangle> [\<rightarrow>]* \<langle>map Val vs,s\<^sub>2,b\<^sub>2\<rangle>"
+ and hp\<^sub>2a: "hp s\<^sub>2 a = Some(C,fs)"
+ and cM_Static: "P \<turnstile> C sees M,Static:Ts\<rightarrow>T = m in D"
+shows "P \<turnstile> \<langle>e\<bullet>M(es),s,b\<rangle> \<rightarrow>* \<langle>THROW IncompatibleClassChangeError,s\<^sub>2,b\<^sub>2\<rangle>"
+(*<*)(is "(?x, ?z) \<in> (red P)\<^sup>*")
+proof -
+ let ?y = "(addr a\<bullet>M(es),s\<^sub>1,b\<^sub>1)"
+ let ?y' = "(addr a\<bullet>M(map Val vs),s\<^sub>2,b\<^sub>2)"
+ have "(?x, ?y) \<in> (red P)\<^sup>*" by(rule CallRedsObj[OF e_steps])
+ also have "(?y, ?y') \<in> (red P)\<^sup>*" by(rule CallRedsParams[OF es_steps])
+ also have "(?y', ?z) \<in> (red P)" using RedCallStatic[OF _ cM_Static] hp\<^sub>2a
+ by(cases s\<^sub>2) simp
+ ultimately show ?thesis by simp
+qed
+(*>*)
+
+subsection\<open>SCall\<close>
+
+lemma SCallRedsParams:
+ "P \<turnstile> \<langle>es,s,b\<rangle> [\<rightarrow>]* \<langle>es',s',b'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sM(es),s,b\<rangle> \<rightarrow>* \<langle>C\<bullet>\<^sub>sM(es'),s',b'\<rangle>"
+(*<*)
+proof(induct rule: rtrancl_induct3)
+ case refl show ?case by blast
+next
+ case step show ?case
+ by(rule rtrancl_into_rtrancl[OF step(3) SCallParams[OF step(2)]])
+qed
+(*>*)
+
+lemma SCallRedsFinal:
+assumes wwf: "wwf_J_prog P"
+and "P \<turnstile> \<langle>es,s\<^sub>0,b\<^sub>0\<rangle> [\<rightarrow>]* \<langle>map Val vs,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\<rangle>"
+ "P \<turnstile> C sees M,Static:Ts\<rightarrow>T = (pns,body) in D"
+ "sh\<^sub>2 D = Some(sfs,Done) \<or> (M = clinit \<and> sh\<^sub>2 D = \<lfloor>(sfs, Processing)\<rfloor>)"
+ "size vs = size pns"
+and l\<^sub>2': "l\<^sub>2' = [pns[\<mapsto>]vs]"
+and body: "P \<turnstile> \<langle>body,(h\<^sub>2,l\<^sub>2',sh\<^sub>2),False\<rangle> \<rightarrow>* \<langle>ef,(h\<^sub>3,l\<^sub>3,sh\<^sub>3),b\<^sub>3\<rangle>"
+and "final ef"
+shows "P \<turnstile> \<langle>C\<bullet>\<^sub>sM(es), s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>ef,(h\<^sub>3,l\<^sub>2,sh\<^sub>3),b\<^sub>3\<rangle>"
+(*<*)
+proof -
+ have wf: "size Ts = size pns \<and> distinct pns"
+ and wt: "fv body \<subseteq> set pns"
+ using assms by(fastforce dest!:sees_wf_mdecl simp:wf_mdecl_def)+
+ from body[THEN Red_lcl_add, of l\<^sub>2]
+ have body': "P \<turnstile> \<langle>body,(h\<^sub>2,l\<^sub>2(pns[\<mapsto>]vs),sh\<^sub>2),False\<rangle> \<rightarrow>* \<langle>ef,(h\<^sub>3,l\<^sub>2++l\<^sub>3,sh\<^sub>3),b\<^sub>3\<rangle>"
+ by (simp add:l\<^sub>2')
+ have "dom l\<^sub>3 \<subseteq> set pns"
+ using Reds_dom_lcl[OF wwf body] wt l\<^sub>2' set_take_subset by force
+ hence eql\<^sub>2: "override_on (l\<^sub>2++l\<^sub>3) l\<^sub>2 (set pns) = l\<^sub>2"
+ by(fastforce simp add:map_add_def override_on_def fun_eq_iff)
+ have b2T: "P \<turnstile> \<langle>C\<bullet>\<^sub>sM(map Val vs), (h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\<rangle> \<rightarrow>* \<langle>C\<bullet>\<^sub>sM(map Val vs), (h\<^sub>2,l\<^sub>2,sh\<^sub>2),True\<rangle>"
+ proof(cases b\<^sub>2)
+ case True then show ?thesis by simp
+ next
+ case False then show ?thesis using assms(3,4) by(auto elim: SCallInitDoneRed)
+ qed
+ have "P \<turnstile> \<langle>C\<bullet>\<^sub>sM(es),s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>C\<bullet>\<^sub>sM(map Val vs),(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\<rangle>"
+ by(rule SCallRedsParams)(rule assms(2))
+ also have "P \<turnstile> \<langle>C\<bullet>\<^sub>sM(map Val vs), (h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\<rangle> \<rightarrow>* \<langle>blocks(pns, Ts, vs, body), (h\<^sub>2,l\<^sub>2,sh\<^sub>2),False\<rangle>"
+ by(auto intro!: rtrancl_into_rtrancl[OF b2T] RedSCall assms(3) simp: assms wf)
+ also (rtrancl_trans) have "P \<turnstile> \<langle>blocks(pns, Ts, vs, body), (h\<^sub>2,l\<^sub>2,sh\<^sub>2),False\<rangle>
+ \<rightarrow>* \<langle>ef,(h\<^sub>3,override_on (l\<^sub>2++l\<^sub>3) l\<^sub>2 (set pns),sh\<^sub>3),b\<^sub>3\<rangle>"
+ by(rule blocksRedsFinal, insert assms wf body', simp_all)
+ finally show ?thesis using eql\<^sub>2 by simp
+qed
+(*>*)
+
+lemma SCallRedsThrowParams:
+ "\<lbrakk> P \<turnstile> \<langle>es,s\<^sub>0,b\<^sub>0\<rangle> [\<rightarrow>]* \<langle>map Val vs\<^sub>1 @ throw a # es\<^sub>2,s\<^sub>2,b\<^sub>2\<rangle> \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sM(es),s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>throw a,s\<^sub>2,b\<^sub>2\<rangle>"
+(*<*)
+by(erule SCallRedsParams[THEN rtrancl_into_rtrancl, OF _ SCallThrowParams])
+ simp
+(*>*)
+
+lemma SCallRedsNone:
+ "\<lbrakk> P \<turnstile> \<langle>es,s,b\<rangle> [\<rightarrow>]* \<langle>map Val vs,s\<^sub>2,False\<rangle>;
+ \<not>(\<exists>b Ts T m D. P \<turnstile> C sees M,b:Ts\<rightarrow>T = m in D) \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sM(es),s,b\<rangle> \<rightarrow>* \<langle>THROW NoSuchMethodError,s\<^sub>2,False\<rangle>"
+(*<*)
+by(erule SCallRedsParams[THEN rtrancl_into_rtrancl, OF _ RedSCallNone])
+ simp
+(*>*)
+
+lemma SCallRedsNonStatic:
+ "\<lbrakk> P \<turnstile> \<langle>es,s,b\<rangle> [\<rightarrow>]* \<langle>map Val vs,s\<^sub>2,False\<rangle>;
+ P \<turnstile> C sees M,NonStatic:Ts\<rightarrow>T = m in D \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sM(es),s,b\<rangle> \<rightarrow>* \<langle>THROW IncompatibleClassChangeError,s\<^sub>2,False\<rangle>"
+(*<*)
+by(erule SCallRedsParams[THEN rtrancl_into_rtrancl, OF _ RedSCallNonStatic])
+ simp
+(*>*)
+
+lemma SCallInitThrowReds:
+assumes wwf: "wwf_J_prog P"
+and "P \<turnstile> \<langle>es,s\<^sub>0,b\<^sub>0\<rangle> [\<rightarrow>]* \<langle>map Val vs,(h\<^sub>1,l\<^sub>1,sh\<^sub>1),False\<rangle>"
+ "P \<turnstile> C sees M,Static:Ts\<rightarrow>T = (pns,body) in D"
+ "\<nexists>sfs. sh\<^sub>1 D = Some(sfs,Done)"
+ "M \<noteq> clinit"
+and "P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(h\<^sub>1,l\<^sub>1,sh\<^sub>1),False\<rangle> \<rightarrow>* \<langle>throw a,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\<rangle>"
+shows "P \<turnstile> \<langle>C\<bullet>\<^sub>sM(es), s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>throw a,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\<rangle>"
+(*<*)(is "(?x, ?z) \<in> (red P)\<^sup>*")
+proof -
+ let ?y = "(C\<bullet>\<^sub>sM(map Val vs),(h\<^sub>1,l\<^sub>1,sh\<^sub>1),False)"
+ let ?y' = "(INIT D ([D],False) \<leftarrow> C\<bullet>\<^sub>sM(map Val vs),(h\<^sub>1,l\<^sub>1,sh\<^sub>1),False)"
+ have "(?x, ?y) \<in> (red P)\<^sup>*" by(rule SCallRedsParams[OF assms(2)])
+ also have "(?y, ?y') \<in> (red P)\<^sup>*"
+ using SCallInitRed[OF assms(3)] assms(4-5) by auto
+ also have "(?y', ?z) \<in> (red P)\<^sup>*" by(rule InitSeqThrowReds[OF assms(6)])
+ finally show ?thesis by simp
+qed
+(*>*)
+
+lemma SCallInitReds:
+assumes wwf: "wwf_J_prog P"
+and "P \<turnstile> \<langle>es,s\<^sub>0,b\<^sub>0\<rangle> [\<rightarrow>]* \<langle>map Val vs,(h\<^sub>1,l\<^sub>1,sh\<^sub>1),False\<rangle>"
+ "P \<turnstile> C sees M,Static:Ts\<rightarrow>T = (pns,body) in D"
+ "\<nexists>sfs. sh\<^sub>1 D = Some(sfs,Done)"
+ "M \<noteq> clinit"
+and "P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(h\<^sub>1,l\<^sub>1,sh\<^sub>1),False\<rangle> \<rightarrow>* \<langle>Val v',(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\<rangle>"
+and "size vs = size pns"
+and l\<^sub>2': "l\<^sub>2' = [pns[\<mapsto>]vs]"
+and body: "P \<turnstile> \<langle>body,(h\<^sub>2,l\<^sub>2',sh\<^sub>2),False\<rangle> \<rightarrow>* \<langle>ef,(h\<^sub>3,l\<^sub>3,sh\<^sub>3),b\<^sub>3\<rangle>"
+and "final ef"
+shows "P \<turnstile> \<langle>C\<bullet>\<^sub>sM(es),s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>ef,(h\<^sub>3,l\<^sub>2,sh\<^sub>3),b\<^sub>3\<rangle>"
+(*<*)
+proof -
+ have wf: "size Ts = size pns \<and> distinct pns"
+ and wt: "fv body \<subseteq> set pns"
+ using assms by(fastforce dest!:sees_wf_mdecl simp:wf_mdecl_def)+
+ from body[THEN Red_lcl_add, of l\<^sub>2]
+ have body': "P \<turnstile> \<langle>body,(h\<^sub>2,l\<^sub>2(pns[\<mapsto>]vs),sh\<^sub>2),False\<rangle> \<rightarrow>* \<langle>ef,(h\<^sub>3,l\<^sub>2++l\<^sub>3,sh\<^sub>3),b\<^sub>3\<rangle>"
+ by (simp add:l\<^sub>2')
+ have "dom l\<^sub>3 \<subseteq> set pns"
+ using Reds_dom_lcl[OF wwf body] wt l\<^sub>2' set_take_subset by force
+ hence eql\<^sub>2: "override_on (l\<^sub>2++l\<^sub>3) l\<^sub>2 (set pns) = l\<^sub>2"
+ by(fastforce simp add:map_add_def override_on_def fun_eq_iff)
+ have "icheck P D (C\<bullet>\<^sub>sM(map Val vs)::'a exp)" using assms(3) by auto
+ then have "P \<turnstile> \<langle>C\<bullet>\<^sub>sM(map Val vs),(h\<^sub>2, l\<^sub>2, sh\<^sub>2),icheck P D (C\<bullet>\<^sub>sM(map Val vs))\<rangle>
+ \<rightarrow> \<langle>blocks(pns, Ts, vs, body), (h\<^sub>2, l\<^sub>2, sh\<^sub>2), False\<rangle>"
+ by (metis (full_types) assms(3) assms(7) local.wf red_reds.RedSCall)
+ also have "P \<turnstile> \<langle>blocks(pns, Ts, vs, body), (h\<^sub>2, l\<^sub>2, sh\<^sub>2), False\<rangle>
+ \<rightarrow>* \<langle>ef,(h\<^sub>3,override_on (l\<^sub>2++l\<^sub>3) l\<^sub>2 (set pns),sh\<^sub>3),b\<^sub>3\<rangle>"
+ by(rule blocksRedsFinal, insert assms wf body', simp_all)
+ finally have trueReds: "P \<turnstile> \<langle>C\<bullet>\<^sub>sM(map Val vs),(h\<^sub>2, l\<^sub>2, sh\<^sub>2),icheck P D (C\<bullet>\<^sub>sM(map Val vs))\<rangle>
+ \<rightarrow>* \<langle>ef,(h\<^sub>3,override_on (l\<^sub>2++l\<^sub>3) l\<^sub>2 (set pns),sh\<^sub>3),b\<^sub>3\<rangle>" by simp
+ have "P \<turnstile> \<langle>C\<bullet>\<^sub>sM(es),s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>C\<bullet>\<^sub>sM(map Val vs),(h\<^sub>1,l\<^sub>1,sh\<^sub>1),False\<rangle>"
+ by(rule SCallRedsParams)(rule assms(2))
+ also have "P \<turnstile> \<langle>C\<bullet>\<^sub>sM(map Val vs),(h\<^sub>1,l\<^sub>1,sh\<^sub>1),False\<rangle> \<rightarrow> \<langle>INIT D ([D],False) \<leftarrow> C\<bullet>\<^sub>sM(map Val vs),(h\<^sub>1,l\<^sub>1,sh\<^sub>1),False\<rangle>"
+ using SCallInitRed[OF assms(3)] assms(4-5) by auto
+ also (rtrancl_into_rtrancl) have "P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> C\<bullet>\<^sub>sM(map Val vs),(h\<^sub>1,l\<^sub>1,sh\<^sub>1),False\<rangle>
+ \<rightarrow>* \<langle>ef,(h\<^sub>3,override_on (l\<^sub>2++l\<^sub>3) l\<^sub>2 (set pns),sh\<^sub>3),b\<^sub>3\<rangle>"
+ using InitSeqReds[OF assms(6) trueReds] assms(5) by simp
+ finally show ?thesis using eql\<^sub>2 by simp
+qed
+(*>*)
+
+lemma SCallInitProcessingReds:
+assumes wwf: "wwf_J_prog P"
+and "P \<turnstile> \<langle>es,s\<^sub>0,b\<^sub>0\<rangle> [\<rightarrow>]* \<langle>map Val vs,(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\<rangle>"
+ "P \<turnstile> C sees M,Static:Ts\<rightarrow>T = (pns,body) in D"
+ "sh\<^sub>2 D = Some(sfs,Processing)"
+and "size vs = size pns"
+and l\<^sub>2': "l\<^sub>2' = [pns[\<mapsto>]vs]"
+and body: "P \<turnstile> \<langle>body,(h\<^sub>2,l\<^sub>2',sh\<^sub>2),False\<rangle> \<rightarrow>* \<langle>ef,(h\<^sub>3,l\<^sub>3,sh\<^sub>3),b\<^sub>3\<rangle>"
+and "final ef"
+shows "P \<turnstile> \<langle>C\<bullet>\<^sub>sM(es),s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>ef,(h\<^sub>3,l\<^sub>2,sh\<^sub>3),b\<^sub>3\<rangle>"
+(*<*)
+proof -
+ have wf: "size Ts = size pns \<and> distinct pns"
+ and wt: "fv body \<subseteq> set pns"
+ using assms by(fastforce dest!:sees_wf_mdecl simp:wf_mdecl_def)+
+ from body[THEN Red_lcl_add, of l\<^sub>2]
+ have body': "P \<turnstile> \<langle>body,(h\<^sub>2,l\<^sub>2(pns[\<mapsto>]vs),sh\<^sub>2),False\<rangle> \<rightarrow>* \<langle>ef,(h\<^sub>3,l\<^sub>2++l\<^sub>3,sh\<^sub>3),b\<^sub>3\<rangle>"
+ by (simp add:l\<^sub>2')
+ have "dom l\<^sub>3 \<subseteq> set pns"
+ using Reds_dom_lcl[OF wwf body] wt l\<^sub>2' set_take_subset by force
+ hence eql\<^sub>2: "override_on (l\<^sub>2++l\<^sub>3) l\<^sub>2 (set pns) = l\<^sub>2"
+ by(fastforce simp add:map_add_def override_on_def fun_eq_iff)
+ have b2T: "P \<turnstile> \<langle>C\<bullet>\<^sub>sM(map Val vs), (h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\<rangle> \<rightarrow>* \<langle>C\<bullet>\<^sub>sM(map Val vs), (h\<^sub>2,l\<^sub>2,sh\<^sub>2),True\<rangle>"
+ proof(cases b\<^sub>2)
+ case True then show ?thesis by simp
+ next
+ case False
+ show ?thesis
+ proof(cases "M = clinit")
+ case True then show ?thesis using False assms(3) red_reds.SCallInitDoneRed assms(4)
+ by (simp add: r_into_rtrancl)
+ next
+ case nclinit: False
+ have icheck: "icheck P D (C\<bullet>\<^sub>sM(map Val vs))" using assms(3) by auto
+ have "P \<turnstile> \<langle>C\<bullet>\<^sub>sM(map Val vs),(h\<^sub>2, l\<^sub>2, sh\<^sub>2),b\<^sub>2\<rangle>
+ \<rightarrow> \<langle>INIT D ([D],False) \<leftarrow> C\<bullet>\<^sub>sM(map Val vs),(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\<rangle>"
+ using SCallInitRed[OF assms(3)] assms(4) False nclinit by simp
+ also have "P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> C\<bullet>\<^sub>sM(map Val vs),(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\<rangle>
+ \<rightarrow> \<langle>INIT D (Nil,True) \<leftarrow> C\<bullet>\<^sub>sM(map Val vs),(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\<rangle>"
+ using RedInitProcessing assms(4) by simp
+ also have "P \<turnstile> \<langle>INIT D (Nil,True) \<leftarrow> C\<bullet>\<^sub>sM(map Val vs),(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\<rangle>
+ \<rightarrow> \<langle>C\<bullet>\<^sub>sM(map Val vs),(h\<^sub>2, l\<^sub>2, sh\<^sub>2),True\<rangle>"
+ using RedInit[of "C\<bullet>\<^sub>sM(map Val vs)" D _ _ _ P] icheck nclinit
+ by (metis (full_types) nsub_RI_Vals sub_RI.simps(12))
+ finally show ?thesis by simp
+ qed
+ qed
+ have "P \<turnstile> \<langle>C\<bullet>\<^sub>sM(es),s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow>* \<langle>C\<bullet>\<^sub>sM(map Val vs),(h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\<rangle>"
+ by(rule SCallRedsParams)(rule assms(2))
+ also have "P \<turnstile> \<langle>C\<bullet>\<^sub>sM(map Val vs), (h\<^sub>2,l\<^sub>2,sh\<^sub>2),b\<^sub>2\<rangle> \<rightarrow>* \<langle>blocks(pns, Ts, vs, body), (h\<^sub>2,l\<^sub>2,sh\<^sub>2),False\<rangle>"
+ by(auto intro!: rtrancl_into_rtrancl[OF b2T] RedSCall assms(3) simp: assms wf)
+ also (rtrancl_trans) have "P \<turnstile> \<langle>blocks(pns, Ts, vs, body), (h\<^sub>2,l\<^sub>2,sh\<^sub>2),False\<rangle>
+ \<rightarrow>* \<langle>ef,(h\<^sub>3,override_on (l\<^sub>2++l\<^sub>3) l\<^sub>2 (set pns),sh\<^sub>3),b\<^sub>3\<rangle>"
+ by(rule blocksRedsFinal, insert assms wf body', simp_all)
+ finally show ?thesis using eql\<^sub>2 by simp
+qed
+(*>*)
+
+(***********************************)
+
+subsubsection "The main Theorem"
+
+lemma assumes wwf: "wwf_J_prog P"
+shows big_by_small: "P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>
+ \<Longrightarrow> (\<And>b. iconf (shp s) e \<Longrightarrow> P,shp s \<turnstile>\<^sub>b (e,b) \<surd> \<Longrightarrow> P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e',s',False\<rangle>)"
+and bigs_by_smalls: "P \<turnstile> \<langle>es,s\<rangle> [\<Rightarrow>] \<langle>es',s'\<rangle>
+ \<Longrightarrow> (\<And>b. iconfs (shp s) es \<Longrightarrow> P,shp s \<turnstile>\<^sub>b (es,b) \<surd> \<Longrightarrow> P \<turnstile> \<langle>es,s,b\<rangle> [\<rightarrow>]* \<langle>es',s',False\<rangle>)"
+(*<*)
+proof (induct rule: eval_evals.inducts)
+ case New show ?case
+ proof(cases b)
+ case True then show ?thesis using RedNew[OF New.hyps(2-4)] by auto
+ next
+ case False then show ?thesis using New.hyps(1) NewInitDoneReds[OF _ New.hyps(2-4)] by auto
+ qed
+next
+ case NewFail show ?case
+ proof(cases b)
+ case True then show ?thesis using RedNewFail[OF NewFail.hyps(2)] NewFail.hyps(3) by fastforce
+ next
+ case False
+ then show ?thesis using NewInitDoneReds2[OF _ NewFail.hyps(2)] NewFail by fastforce
+ qed
+next
+ case (NewInit sh C h l v' h' l' sh' a FDTs h'') show ?case
+ proof(cases b)
+ case True
+ then obtain sfs where shC: "sh C = \<lfloor>(sfs, Processing)\<rfloor>"
+ using NewInit.hyps(1) NewInit.prems by(clarsimp simp: bconf_def initPD_def)
+ then have s': "(h',l',sh') = (h,l,sh)" using NewInit.hyps(2) init_ProcessingE by clarsimp
+ then show ?thesis using RedNew[OF NewInit.hyps(4-6)] True by auto
+ next
+ case False
+ then have init: "P \<turnstile> \<langle>INIT C ([C],False) \<leftarrow> unit,(h, l, sh),False\<rangle> \<rightarrow>* \<langle>Val v',(h', l', sh'),False\<rangle>"
+ using NewInit.hyps(3) by(auto simp: bconf_def)
+ then show ?thesis using NewInit NewInitReds[OF _ init NewInit.hyps(4-6)] False
+ has_fields_is_class[OF NewInit.hyps(5)] by auto
+ qed
+next
+ case (NewInitOOM sh C h l v' h' l' sh') show ?case
+ proof(cases b)
+ case True
+ then obtain sfs where shC: "sh C = \<lfloor>(sfs, Processing)\<rfloor>"
+ using NewInitOOM.hyps(1) NewInitOOM.prems by(clarsimp simp: bconf_def initPD_def)
+ then have s': "(h',l',sh') = (h,l,sh)" using NewInitOOM.hyps(2) init_ProcessingE by clarsimp
+ then show ?thesis using RedNewFail[OF NewInitOOM.hyps(4)] True r_into_rtrancl NewInitOOM.hyps(5)
+ by auto
+ next
+ case False
+ then have init: "P \<turnstile> \<langle>INIT C ([C],False) \<leftarrow> unit,(h, l, sh),False\<rangle> \<rightarrow>* \<langle>Val v',(h', l', sh'),False\<rangle>"
+ using NewInitOOM.hyps(3) by(auto simp: bconf_def)
+ then show ?thesis using NewInitOOM.hyps(1) NewInitOOMReds[OF _ init NewInitOOM.hyps(4)] False
+ NewInitOOM.hyps(5) by auto
+ qed
+next
+ case (NewInitThrow sh C h l a s') show ?case
+ proof(cases b)
+ case True
+ then obtain sfs where shC: "sh C = \<lfloor>(sfs, Processing)\<rfloor>"
+ using NewInitThrow.hyps(1) NewInitThrow.prems by(clarsimp simp: bconf_def initPD_def)
+ then show ?thesis using NewInitThrow.hyps(2) init_ProcessingE by blast
+ next
+ case False
+ then have init: "P \<turnstile> \<langle>INIT C ([C],False) \<leftarrow> unit,(h, l, sh),b\<rangle> \<rightarrow>* \<langle>throw a,s',False\<rangle>"
+ using NewInitThrow.hyps(3) by(auto simp: bconf_def)
+ then show ?thesis using NewInitThrow NewInitThrowReds[of "(h,l,sh)" C P a s'] False by auto
+ qed
+next
+ case Cast then show ?case by(fastforce intro:CastRedsAddr)
+next
+ case CastNull then show ?case by(fastforce intro: CastRedsNull)
+next
+ case CastFail thus ?case by(fastforce intro!:CastRedsFail)
+next
+ case CastThrow thus ?case by(fastforce dest!:eval_final intro:CastRedsThrow)
+next
+ case Val then show ?case using exI[of _ b] by(simp add: bconf_def)
+next
+ case (BinOp e\<^sub>1 s\<^sub>0 v\<^sub>1 s\<^sub>1 e\<^sub>2 v\<^sub>2 s\<^sub>2 bop v)
+ show ?case
+ proof(cases "val_of e\<^sub>1")
+ case None
+ then have iconf: "iconf (shp s\<^sub>0) e\<^sub>1" using None BinOp.prems by auto
+ have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e\<^sub>1,b) \<surd>" using None BinOp.prems by auto
+ then have b1: "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Val v\<^sub>1,s\<^sub>1,False\<rangle>" using iconf BinOp.hyps(2) by auto
+ have binop: "P \<turnstile> \<langle>e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Val v\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2,s\<^sub>1,False\<rangle>" by(rule BinOp1Reds[OF b1])
+ then have iconf2': "iconf (shp s\<^sub>1) e\<^sub>2" using Red_preserves_iconf[OF wwf binop] None BinOp by auto
+ have "P,shp s\<^sub>1 \<turnstile>\<^sub>b (e\<^sub>2,False) \<surd>" by(simp add: bconf_def)
+ then have b2: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>1,False\<rangle> \<rightarrow>* \<langle>Val v\<^sub>2,s\<^sub>2,False\<rangle>" using BinOp.hyps(4)[OF iconf2'] by auto
+ then show ?thesis using BinOpRedsVal[OF b1 b2 BinOp.hyps(5)] by fast
+ next
+ case (Some a)
+ then obtain b1 where b1: "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Val v\<^sub>1,s\<^sub>1,b1\<rangle>"
+ by (metis (no_types, lifting) BinOp.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec)
+ have binop: "P \<turnstile> \<langle>e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Val v\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2,s\<^sub>1,b1\<rangle>" by(rule BinOp1Reds[OF b1])
+ then have iconf2': "iconf (shp s\<^sub>1) e\<^sub>2" using Red_preserves_iconf[OF wwf binop] BinOp by auto
+ have bconf2: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e\<^sub>2,b) \<surd>" using BinOp.prems Some by simp
+ then have "P,shp s\<^sub>1 \<turnstile>\<^sub>b (e\<^sub>2,b1) \<surd>" using Red_preserves_bconf[OF wwf binop BinOp.prems] by simp
+ then have b2: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>1,b1\<rangle> \<rightarrow>* \<langle>Val v\<^sub>2,s\<^sub>2,False\<rangle>" using BinOp.hyps(4)[OF iconf2'] by auto
+ then show ?thesis using BinOpRedsVal[OF b1 b2 BinOp.hyps(5)] by fast
+ qed
+next
+ case (BinOpThrow1 e\<^sub>1 s\<^sub>0 e s\<^sub>1 bop e\<^sub>2) show ?case
+ proof(cases "val_of e\<^sub>1")
+ case None
+ then have "iconf (shp s\<^sub>0) e\<^sub>1" and "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e\<^sub>1,b) \<surd>" using BinOpThrow1.prems by auto
+ then have b1: "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>throw e,s\<^sub>1,False\<rangle>" using BinOpThrow1.hyps(2) by auto
+ then have "P \<turnstile> \<langle>e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>throw e,s\<^sub>1,False\<rangle>"
+ using BinOpThrow1 None by(auto dest!:eval_final simp: BinOpRedsThrow1[OF b1])
+ then show ?thesis by fast
+ next
+ case (Some a)
+ then show ?thesis using eval_final_same[OF BinOpThrow1.hyps(1)] val_of_spec[OF Some] by auto
+ qed
+next
+ case (BinOpThrow2 e\<^sub>1 s\<^sub>0 v\<^sub>1 s\<^sub>1 e\<^sub>2 e s\<^sub>2 bop)
+ show ?case
+ proof(cases "val_of e\<^sub>1")
+ case None
+ then have iconf: "iconf (shp s\<^sub>0) e\<^sub>1" using None BinOpThrow2.prems by auto
+ have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e\<^sub>1,b) \<surd>" using None BinOpThrow2.prems by auto
+ then have b1: "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Val v\<^sub>1,s\<^sub>1,False\<rangle>" using iconf BinOpThrow2.hyps(2) by auto
+ have binop: "P \<turnstile> \<langle>e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Val v\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2,s\<^sub>1,False\<rangle>" by(rule BinOp1Reds[OF b1])
+ then have iconf2': "iconf (shp s\<^sub>1) e\<^sub>2" using Red_preserves_iconf[OF wwf binop] None BinOpThrow2 by auto
+ have "P,shp s\<^sub>1 \<turnstile>\<^sub>b (e\<^sub>2,False) \<surd>" by(simp add: bconf_def)
+ then have b2: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>1,False\<rangle> \<rightarrow>* \<langle>throw e,s\<^sub>2,False\<rangle>" using BinOpThrow2.hyps(4)[OF iconf2'] by auto
+ then show ?thesis using BinOpRedsThrow2[OF b1 b2] by fast
+ next
+ case (Some a)
+ then obtain b1 where b1: "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Val v\<^sub>1,s\<^sub>1,b1\<rangle>"
+ by (metis (no_types, lifting) BinOpThrow2.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec)
+ have binop: "P \<turnstile> \<langle>e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Val v\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2,s\<^sub>1,b1\<rangle>" by(rule BinOp1Reds[OF b1])
+ then have iconf2': "iconf (shp s\<^sub>1) e\<^sub>2" using Red_preserves_iconf[OF wwf binop] BinOpThrow2 by auto
+ have bconf2: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e\<^sub>2,b) \<surd>" using BinOpThrow2.prems Some by simp
+ then have "P,shp s\<^sub>1 \<turnstile>\<^sub>b (e\<^sub>2,b1) \<surd>" using Red_preserves_bconf[OF wwf binop BinOpThrow2.prems] by simp
+ then have b2: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>1,b1\<rangle> \<rightarrow>* \<langle>throw e,s\<^sub>2,False\<rangle>" using BinOpThrow2.hyps(4)[OF iconf2'] by auto
+ then show ?thesis using BinOpRedsThrow2[OF b1 b2] by fast
+ qed
+next
+ case Var thus ?case by(auto dest:RedVar simp: bconf_def)
+next
+ case LAss thus ?case by(auto dest: LAssRedsVal)
+next
+ case LAssThrow thus ?case by(auto dest!:eval_final dest: LAssRedsThrow)
+next
+ case FAcc thus ?case by(fastforce intro:FAccRedsVal)
+next
+ case FAccNull thus ?case by(auto dest:FAccRedsNull)
+next
+ case FAccThrow thus ?case by(auto dest!:eval_final dest:FAccRedsThrow)
+next
+ case FAccNone then show ?case by(fastforce intro: FAccRedsNone)
+next
+ case FAccStatic then show ?case by(fastforce intro: FAccRedsStatic)
+next
+ case SFAcc show ?case
+ proof(cases b)
+ case True then show ?thesis using RedSFAcc SFAcc.hyps by auto
+ next
+ case False then show ?thesis using SFAcc.hyps SFAccInitDoneReds[OF SFAcc.hyps(1)] by auto
+ qed
+next
+ case (SFAccInit C F t D sh h l v' h' l' sh' sfs i v) show ?case
+ proof(cases b)
+ case True
+ then obtain sfs where shC: "sh D = \<lfloor>(sfs, Processing)\<rfloor>"
+ using SFAccInit.hyps(2) SFAccInit.prems by(clarsimp simp: bconf_def initPD_def)
+ then have s': "(h',l',sh') = (h,l,sh)" using SFAccInit.hyps(3) init_ProcessingE by clarsimp
+ then show ?thesis using RedSFAcc SFAccInit.hyps True by auto
+ next
+ case False
+ then have init: "P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(h, l, sh),False\<rangle> \<rightarrow>* \<langle>Val v',(h', l', sh'),False\<rangle>"
+ using SFAccInit.hyps(4) by(auto simp: bconf_def)
+ then show ?thesis using SFAccInit SFAccInitReds[OF _ _ init] False by auto
+ qed
+next
+ case (SFAccInitThrow C F t D sh h l a s') show ?case
+ proof(cases b)
+ case True
+ then obtain sfs where shC: "sh D = \<lfloor>(sfs, Processing)\<rfloor>"
+ using SFAccInitThrow.hyps(2) SFAccInitThrow.prems(2) by(clarsimp simp: bconf_def initPD_def)
+ then show ?thesis using SFAccInitThrow.hyps(3) init_ProcessingE by blast
+ next
+ case False
+ then have init: "P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(h, l, sh),b\<rangle> \<rightarrow>* \<langle>throw a,s',False\<rangle>"
+ using SFAccInitThrow.hyps(4) by(auto simp: bconf_def)
+ then show ?thesis using SFAccInitThrow SFAccInitThrowReds False by auto
+ qed
+next
+ case SFAccNone then show ?case by(fastforce intro: SFAccRedsNone)
+next
+ case SFAccNonStatic then show ?case by(fastforce intro: SFAccRedsNonStatic)
+next
+ case (FAss e\<^sub>1 s\<^sub>0 a s\<^sub>1 e\<^sub>2 v h\<^sub>2 l\<^sub>2 sh\<^sub>2 C fs F t D fs' h\<^sub>2')
+ show ?case
+ proof(cases "val_of e\<^sub>1")
+ case None
+ then have iconf: "iconf (shp s\<^sub>0) e\<^sub>1" using None FAss.prems by auto
+ have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e\<^sub>1,b) \<surd>" using None FAss.prems by auto
+ then have b1: "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>addr a,s\<^sub>1,False\<rangle>" using iconf FAss.hyps(2) by auto
+ have fass: "P \<turnstile> \<langle>e\<^sub>1\<bullet>F{D} := e\<^sub>2,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>addr a\<bullet>F{D} := e\<^sub>2,s\<^sub>1,False\<rangle>" by(rule FAssReds1[OF b1])
+ then have iconf2': "iconf (shp s\<^sub>1) e\<^sub>2" using Red_preserves_iconf[OF wwf fass] None FAss by auto
+ have "P,shp s\<^sub>1 \<turnstile>\<^sub>b (e\<^sub>2,False) \<surd>" by(simp add: bconf_def)
+ then have b2: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>1,False\<rangle> \<rightarrow>* \<langle>Val v,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\<rangle>" using FAss.hyps(4)[OF iconf2'] by auto
+ then show ?thesis using FAssRedsVal[OF b1 b2 FAss.hyps(6) FAss.hyps(5)[THEN sym]] FAss.hyps(7,8) by fast
+ next
+ case (Some a')
+ then obtain b1 where b1: "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>addr a,s\<^sub>1,b1\<rangle>"
+ by (metis (no_types, lifting) FAss.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec)
+ have fass: "P \<turnstile> \<langle>e\<^sub>1\<bullet>F{D} := e\<^sub>2,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>addr a\<bullet>F{D} := e\<^sub>2,s\<^sub>1,b1\<rangle>" by(rule FAssReds1[OF b1])
+ then have iconf2': "iconf (shp s\<^sub>1) e\<^sub>2" using Red_preserves_iconf[OF wwf fass] FAss by auto
+ have bconf2: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e\<^sub>2,b) \<surd>" using FAss.prems Some by simp
+ then have "P,shp s\<^sub>1 \<turnstile>\<^sub>b (e\<^sub>2,b1) \<surd>" using Red_preserves_bconf[OF wwf fass FAss.prems] by simp
+ then have b2: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>1,b1\<rangle> \<rightarrow>* \<langle>Val v,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\<rangle>" using FAss.hyps(4)[OF iconf2'] by auto
+ then show ?thesis using FAssRedsVal[OF b1 b2] FAss.hyps(5)[THEN sym] FAss.hyps(6-8) by fast
+ qed
+next
+ case (FAssNull e\<^sub>1 s\<^sub>0 s\<^sub>1 e\<^sub>2 v s\<^sub>2 F D)
+ show ?case
+ proof(cases "val_of e\<^sub>1")
+ case None
+ then have iconf: "iconf (shp s\<^sub>0) e\<^sub>1" using FAssNull.prems(1) by simp
+ have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e\<^sub>1,b) \<surd>" using FAssNull.prems(2) None by simp
+ then have b1: "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>null,s\<^sub>1,False\<rangle>" using FAssNull.hyps(2)[OF iconf] by auto
+ have fass: "P \<turnstile> \<langle>e\<^sub>1\<bullet>F{D} := e\<^sub>2,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>null\<bullet>F{D} := e\<^sub>2,s\<^sub>1,False\<rangle>" by(rule FAssReds1[OF b1])
+ then have iconf2': "iconf (shp s\<^sub>1) e\<^sub>2" using Red_preserves_iconf[OF wwf fass] None FAssNull by auto
+ have "P,shp s\<^sub>1 \<turnstile>\<^sub>b (e\<^sub>2,False) \<surd>" by(simp add: bconf_def)
+ then have b2: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>1,False\<rangle> \<rightarrow>* \<langle>Val v,s\<^sub>2,False\<rangle>" using FAssNull.hyps(4)[OF iconf2'] by auto
+ then show ?thesis using FAssRedsNull[OF b1 b2] by fast
+ next
+ case (Some a')
+ then obtain b1 where b1: "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>null,s\<^sub>1,b1\<rangle>"
+ by (metis (no_types, lifting) FAssNull.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec)
+ have fass: "P \<turnstile> \<langle>e\<^sub>1\<bullet>F{D} := e\<^sub>2,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>null\<bullet>F{D} := e\<^sub>2,s\<^sub>1,b1\<rangle>" by(rule FAssReds1[OF b1])
+ then have iconf2': "iconf (shp s\<^sub>1) e\<^sub>2" using Red_preserves_iconf[OF wwf fass] FAssNull by auto
+ have bconf2: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e\<^sub>2,b) \<surd>" using FAssNull.prems Some by simp
+ then have "P,shp s\<^sub>1 \<turnstile>\<^sub>b (e\<^sub>2,b1) \<surd>" using Red_preserves_bconf[OF wwf fass FAssNull.prems] by simp
+ then have b2: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>1,b1\<rangle> \<rightarrow>* \<langle>Val v,s\<^sub>2,False\<rangle>" using FAssNull.hyps(4)[OF iconf2'] by auto
+ then show ?thesis using FAssRedsNull[OF b1 b2] by fast
+ qed
+next
+ case (FAssThrow1 e\<^sub>1 s\<^sub>0 e' s\<^sub>1 F D e\<^sub>2) show ?case
+ proof(cases "val_of e\<^sub>1")
+ case None
+ then have "iconf (shp s\<^sub>0) e\<^sub>1" and "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e\<^sub>1,b) \<surd>" using FAssThrow1.prems by auto
+ then have b1: "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>throw e',s\<^sub>1,False\<rangle>" using FAssThrow1.hyps(2) by auto
+ then have "P \<turnstile> \<langle>e\<^sub>1\<bullet>F{D} := e\<^sub>2,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>throw e',s\<^sub>1,False\<rangle>"
+ using FAssThrow1 None by(auto dest!:eval_final simp: FAssRedsThrow1[OF b1])
+ then show ?thesis by fast
+ next
+ case (Some a)
+ then show ?thesis using eval_final_same[OF FAssThrow1.hyps(1)] val_of_spec[OF Some] by auto
+ qed
+next
+ case (FAssThrow2 e\<^sub>1 s\<^sub>0 v s\<^sub>1 e\<^sub>2 e' s\<^sub>2 F D)
+ show ?case
+ proof(cases "val_of e\<^sub>1")
+ case None
+ then have iconf: "iconf (shp s\<^sub>0) e\<^sub>1" using None FAssThrow2.prems by auto
+ have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e\<^sub>1,b) \<surd>" using None FAssThrow2.prems by auto
+ then have b1: "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Val v,s\<^sub>1,False\<rangle>" using iconf FAssThrow2.hyps(2) by auto
+ have fass: "P \<turnstile> \<langle>e\<^sub>1\<bullet>F{D} := e\<^sub>2,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Val v\<bullet>F{D} := e\<^sub>2,s\<^sub>1,False\<rangle>" by(rule FAssReds1[OF b1])
+ then have iconf2': "iconf (shp s\<^sub>1) e\<^sub>2" using Red_preserves_iconf[OF wwf fass] None FAssThrow2 by auto
+ have "P,shp s\<^sub>1 \<turnstile>\<^sub>b (e\<^sub>2,False) \<surd>" by(simp add: bconf_def)
+ then have b2: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>1,False\<rangle> \<rightarrow>* \<langle>throw e',s\<^sub>2,False\<rangle>" using FAssThrow2.hyps(4)[OF iconf2'] by auto
+ then show ?thesis using FAssRedsThrow2[OF b1 b2] by fast
+ next
+ case (Some a')
+ then obtain b1 where b1: "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Val v,s\<^sub>1,b1\<rangle>"
+ by (metis (no_types, lifting) FAssThrow2.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec)
+ have fass: "P \<turnstile> \<langle>e\<^sub>1\<bullet>F{D} := e\<^sub>2,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Val v\<bullet>F{D} := e\<^sub>2,s\<^sub>1,b1\<rangle>" by(rule FAssReds1[OF b1])
+ then have iconf2': "iconf (shp s\<^sub>1) e\<^sub>2" using Red_preserves_iconf[OF wwf fass] FAssThrow2 by auto
+ have bconf2: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e\<^sub>2,b) \<surd>" using FAssThrow2.prems Some by simp
+ then have "P,shp s\<^sub>1 \<turnstile>\<^sub>b (e\<^sub>2,b1) \<surd>" using Red_preserves_bconf[OF wwf fass FAssThrow2.prems] by simp
+ then have b2: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>1,b1\<rangle> \<rightarrow>* \<langle>throw e',s\<^sub>2,False\<rangle>" using FAssThrow2.hyps(4)[OF iconf2'] by auto
+ then show ?thesis using FAssRedsThrow2[OF b1 b2] by fast
+ qed
+next
+ case (FAssNone e\<^sub>1 s\<^sub>0 a s\<^sub>1 e\<^sub>2 v h\<^sub>2 l\<^sub>2 sh\<^sub>2 C fs F D)
+ show ?case
+ proof(cases "val_of e\<^sub>1")
+ case None
+ then have iconf: "iconf (shp s\<^sub>0) e\<^sub>1" using None FAssNone.prems by auto
+ have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e\<^sub>1,b) \<surd>" using None FAssNone.prems by auto
+ then have b1: "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>addr a,s\<^sub>1,False\<rangle>" using iconf FAssNone.hyps(2) by auto
+ have fass: "P \<turnstile> \<langle>e\<^sub>1\<bullet>F{D} := e\<^sub>2,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>addr a\<bullet>F{D} := e\<^sub>2,s\<^sub>1,False\<rangle>" by(rule FAssReds1[OF b1])
+ then have iconf2': "iconf (shp s\<^sub>1) e\<^sub>2" using Red_preserves_iconf[OF wwf fass] None FAssNone by auto
+ have "P,shp s\<^sub>1 \<turnstile>\<^sub>b (e\<^sub>2,False) \<surd>" by(simp add: bconf_def)
+ then have b2: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>1,False\<rangle> \<rightarrow>* \<langle>Val v,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\<rangle>" using FAssNone.hyps(4)[OF iconf2'] by auto
+ then show ?thesis using FAssRedsNone[OF b1 b2 FAssNone.hyps(5,6)] by fast
+ next
+ case (Some a')
+ then obtain b1 where b1: "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>addr a,s\<^sub>1,b1\<rangle>"
+ by (metis (no_types, lifting) FAssNone.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec)
+ have fass: "P \<turnstile> \<langle>e\<^sub>1\<bullet>F{D} := e\<^sub>2,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>addr a\<bullet>F{D} := e\<^sub>2,s\<^sub>1,b1\<rangle>" by(rule FAssReds1[OF b1])
+ then have iconf2': "iconf (shp s\<^sub>1) e\<^sub>2" using Red_preserves_iconf[OF wwf fass] FAssNone by auto
+ have bconf2: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e\<^sub>2,b) \<surd>" using FAssNone.prems Some by simp
+ then have "P,shp s\<^sub>1 \<turnstile>\<^sub>b (e\<^sub>2,b1) \<surd>" using Red_preserves_bconf[OF wwf fass FAssNone.prems] by simp
+ then have b2: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>1,b1\<rangle> \<rightarrow>* \<langle>Val v,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\<rangle>" using FAssNone.hyps(4)[OF iconf2'] by auto
+ then show ?thesis using FAssRedsNone[OF b1 b2 FAssNone.hyps(5,6)] by fast
+ qed
+next
+ case (FAssStatic e\<^sub>1 s\<^sub>0 a s\<^sub>1 e\<^sub>2 v h\<^sub>2 l\<^sub>2 sh\<^sub>2 C fs F t D)
+ show ?case
+ proof(cases "val_of e\<^sub>1")
+ case None
+ then have iconf: "iconf (shp s\<^sub>0) e\<^sub>1" using None FAssStatic.prems by auto
+ have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e\<^sub>1,b) \<surd>" using None FAssStatic.prems by auto
+ then have b1: "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>addr a,s\<^sub>1,False\<rangle>" using iconf FAssStatic.hyps(2) by auto
+ have fass: "P \<turnstile> \<langle>e\<^sub>1\<bullet>F{D} := e\<^sub>2,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>addr a\<bullet>F{D} := e\<^sub>2,s\<^sub>1,False\<rangle>" by(rule FAssReds1[OF b1])
+ then have iconf2': "iconf (shp s\<^sub>1) e\<^sub>2" using Red_preserves_iconf[OF wwf fass] None FAssStatic by auto
+ have "P,shp s\<^sub>1 \<turnstile>\<^sub>b (e\<^sub>2,False) \<surd>" by(simp add: bconf_def)
+ then have b2: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>1,False\<rangle> \<rightarrow>* \<langle>Val v,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\<rangle>" using FAssStatic.hyps(4)[OF iconf2'] by auto
+ then show ?thesis using FAssRedsStatic[OF b1 b2 FAssStatic.hyps(5,6)] by fast
+ next
+ case (Some a')
+ then obtain b1 where b1: "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>addr a,s\<^sub>1,b1\<rangle>"
+ by (metis (no_types, lifting) FAssStatic.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec)
+ have fass: "P \<turnstile> \<langle>e\<^sub>1\<bullet>F{D} := e\<^sub>2,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>addr a\<bullet>F{D} := e\<^sub>2,s\<^sub>1,b1\<rangle>" by(rule FAssReds1[OF b1])
+ then have iconf2': "iconf (shp s\<^sub>1) e\<^sub>2" using Red_preserves_iconf[OF wwf fass] FAssStatic by auto
+ have bconf2: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e\<^sub>2,b) \<surd>" using FAssStatic.prems Some by simp
+ then have "P,shp s\<^sub>1 \<turnstile>\<^sub>b (e\<^sub>2,b1) \<surd>" using Red_preserves_bconf[OF wwf fass FAssStatic.prems] by simp
+ then have b2: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>1,b1\<rangle> \<rightarrow>* \<langle>Val v,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\<rangle>" using FAssStatic.hyps(4)[OF iconf2'] by auto
+ then show ?thesis using FAssRedsStatic[OF b1 b2 FAssStatic.hyps(5,6)] by fast
+ qed
+next
+ case (SFAss e\<^sub>2 s\<^sub>0 v h\<^sub>1 l\<^sub>1 sh\<^sub>1 C F t D sfs sfs' sh\<^sub>1')
+ show ?case
+ proof(cases "val_of e\<^sub>2")
+ case None
+ then have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e\<^sub>2,b) \<surd>" using SFAss.prems(2) by simp
+ then have b1: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Val v,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\<rangle>" using SFAss by auto
+ thus ?thesis using SFAssRedsVal[OF b1 SFAss.hyps(3,4)] SFAss.hyps(5,6) by fast
+ next
+ case (Some a)
+ then obtain b1 where b1: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Val v,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),b1\<rangle>"
+ by (metis (no_types, lifting) SFAss.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec)
+ thus ?thesis using SFAssRedsVal[OF b1 SFAss.hyps(3,4)] SFAss.hyps(5,6) by fast
+ qed
+next
+ case (SFAssInit e\<^sub>2 s\<^sub>0 v h\<^sub>1 l\<^sub>1 sh\<^sub>1 C F t D v' h' l' sh' sfs i sfs' sh'')
+ then have iconf: "iconf (shp s\<^sub>0) e\<^sub>2" by simp
+ show ?case
+ proof(cases "val_of e\<^sub>2")
+ case None
+ then have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e\<^sub>2,b) \<surd>" using SFAssInit.prems(2) by simp
+ then have reds: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Val v,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\<rangle>"
+ using SFAssInit.hyps(2)[OF iconf bconf] by auto
+ then have init: "P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\<rangle> \<rightarrow>* \<langle>Val v',(h', l', sh'),False\<rangle>"
+ using SFAssInit.hyps(6) by(auto simp: bconf_def)
+ then show ?thesis using SFAssInit SFAssInitReds[OF reds SFAssInit.hyps(3) _ init] by auto
+ next
+ case (Some v2) show ?thesis
+ proof(cases b)
+ case False
+ then have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e\<^sub>2,b) \<surd>" by(simp add: bconf_def)
+ then have reds: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Val v,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\<rangle>"
+ using SFAssInit.hyps(2)[OF iconf bconf] by auto
+ then have init: "P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\<rangle> \<rightarrow>* \<langle>Val v',(h', l', sh'),False\<rangle>"
+ using SFAssInit.hyps(6) by(auto simp: bconf_def)
+ then show ?thesis using SFAssInit SFAssInitReds[OF reds SFAssInit.hyps(3) _ init] by auto
+ next
+ case True
+ have e\<^sub>2: "e\<^sub>2 = Val v2" using val_of_spec[OF Some] by simp
+ then have vs: "v2 = v \<and> s\<^sub>0 = (h\<^sub>1, l\<^sub>1, sh\<^sub>1)" using eval_final_same[OF SFAssInit.hyps(1)] by simp
+ then obtain sfs where shC: "sh\<^sub>1 D = \<lfloor>(sfs, Processing)\<rfloor>"
+ using SFAssInit.hyps(3,4) SFAssInit.prems(2) Some True
+ by(cases e\<^sub>2, auto simp: bconf_def initPD_def dest: sees_method_fun)
+ then have s': "(h',l',sh') = (h\<^sub>1, l\<^sub>1, sh\<^sub>1)" using SFAssInit.hyps(5) init_ProcessingE by clarsimp
+ then show ?thesis using SFAssInit.hyps(3,7-9) True e\<^sub>2 red_reds.RedSFAss vs by auto
+ qed
+ qed
+next
+ case (SFAssInitThrow e\<^sub>2 s\<^sub>0 v h\<^sub>1 l\<^sub>1 sh\<^sub>1 C F t D a s')
+ then have iconf: "iconf (shp s\<^sub>0) e\<^sub>2" by simp
+ show ?case
+ proof(cases "val_of e\<^sub>2")
+ case None
+ then have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e\<^sub>2,b) \<surd>" using SFAssInitThrow.prems(2) by simp
+ then have reds: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Val v,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\<rangle>"
+ using SFAssInitThrow.hyps(2)[OF iconf bconf] by auto
+ then have init: "P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\<rangle> \<rightarrow>* \<langle>throw a,s',False\<rangle>"
+ using SFAssInitThrow.hyps(6) by(auto simp: bconf_def)
+ then show ?thesis using SFAssInitThrow SFAssInitThrowReds[OF reds _ _ init] by auto
+ next
+ case (Some v2) show ?thesis
+ proof(cases b)
+ case False
+ then have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e\<^sub>2,b) \<surd>" by(simp add: bconf_def)
+ then have reds: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Val v,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\<rangle>"
+ using SFAssInitThrow.hyps(2)[OF iconf bconf] by auto
+ then have init: "P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\<rangle> \<rightarrow>* \<langle>throw a,s',False\<rangle>"
+ using SFAssInitThrow.hyps(6) by(auto simp: bconf_def)
+ then show ?thesis using SFAssInitThrow SFAssInitThrowReds[OF reds _ _ init] by auto
+ next
+ case True
+ obtain v2 where e\<^sub>2: "e\<^sub>2 = Val v2" using val_of_spec[OF Some] by simp
+ then have vs: "v2 = v \<and> s\<^sub>0 = (h\<^sub>1, l\<^sub>1, sh\<^sub>1)"
+ using eval_final_same[OF SFAssInitThrow.hyps(1)] by simp
+ then obtain sfs where shC: "sh\<^sub>1 D = \<lfloor>(sfs, Processing)\<rfloor>"
+ using SFAssInitThrow.hyps(4) SFAssInitThrow.prems(2) Some True
+ by(cases e\<^sub>2, auto simp: bconf_def initPD_def dest: sees_method_fun)
+ then show ?thesis using SFAssInitThrow.hyps(5) init_ProcessingE by blast
+ qed
+ qed
+next
+ case (SFAssThrow e\<^sub>2 s\<^sub>0 e' s\<^sub>2 C F D)
+ show ?case
+ proof(cases "val_of e\<^sub>2")
+ case None
+ then have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e\<^sub>2,b) \<surd>" using SFAssThrow.prems(2) None by simp
+ then have b1: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>throw e',s\<^sub>2,False\<rangle>" using SFAssThrow by auto
+ thus ?thesis using SFAssRedsThrow[OF b1] by fast
+ next
+ case (Some a)
+ then show ?thesis using eval_final_same[OF SFAssThrow.hyps(1)] val_of_spec[OF Some] by auto
+ qed
+next
+ case (SFAssNone e\<^sub>2 s\<^sub>0 v h\<^sub>2 l\<^sub>2 sh\<^sub>2 C F D)
+ show ?case
+ proof(cases "val_of e\<^sub>2")
+ case None
+ then have iconf: "iconf (shp s\<^sub>0) e\<^sub>2" using SFAssNone by simp
+ then have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e\<^sub>2,b) \<surd>" using SFAssNone.prems(2) None by simp
+ then have b1: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Val v,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\<rangle>" using SFAssNone.hyps(2) iconf by auto
+ thus ?thesis using SFAssRedsNone[OF b1 SFAssNone.hyps(3)] by fast
+ next
+ case (Some a)
+ then obtain b1 where b1: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Val v,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),b1\<rangle>"
+ by (metis (no_types, lifting) SFAssNone.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec)
+ thus ?thesis using SFAssRedsNone[OF b1 SFAssNone.hyps(3)] by fast
+ qed
+next
+ case (SFAssNonStatic e\<^sub>2 s\<^sub>0 v h\<^sub>2 l\<^sub>2 sh\<^sub>2 C F t D) show ?case
+ proof(cases "val_of e\<^sub>2")
+ case None
+ then have iconf: "iconf (shp s\<^sub>0) e\<^sub>2" using SFAssNonStatic by simp
+ then have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e\<^sub>2,b) \<surd>" using SFAssNonStatic.prems(2) None by simp
+ then have b1: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Val v,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\<rangle>" using SFAssNonStatic.hyps(2) iconf by auto
+ thus ?thesis using SFAssRedsNonStatic[OF b1 SFAssNonStatic.hyps(3)] by fast
+ next
+ case (Some a)
+ then obtain b' where b1: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Val v,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),b'\<rangle>"
+ by (metis (no_types, lifting) SFAssNonStatic.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec)
+ thus ?thesis using SFAssRedsNonStatic[OF b1 SFAssNonStatic.hyps(3)] by fast
+ qed
+next
+ case (CallObjThrow e s\<^sub>0 e' s\<^sub>1 M ps) show ?case
+ proof(cases "val_of e")
+ case None
+ then have "iconf (shp s\<^sub>0) e" and "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e,b) \<surd>" using CallObjThrow.prems by auto
+ then have b1: "P \<turnstile> \<langle>e,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>throw e',s\<^sub>1,False\<rangle>" using CallObjThrow.hyps(2) by auto
+ then have "P \<turnstile> \<langle>e\<bullet>M(ps),s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>throw e',s\<^sub>1,False\<rangle>"
+ using CallObjThrow None by(auto dest!:eval_final simp: CallRedsThrowObj[OF b1])
+ then show ?thesis by fast
+ next
+ case (Some a)
+ then show ?thesis using eval_final_same[OF CallObjThrow.hyps(1)] val_of_spec[OF Some] by auto
+ qed
+next
+ case (CallNull e s\<^sub>0 s\<^sub>1 ps vs s\<^sub>2 M) show ?case
+ proof(cases "val_of e")
+ case None
+ then have iconf: "iconf (shp s\<^sub>0) e" using CallNull.prems(1) by simp
+ have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e,b) \<surd>" using CallNull.prems(2) None by simp
+ then have b1: "P \<turnstile> \<langle>e,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>null,s\<^sub>1,False\<rangle>" using CallNull.hyps(2)[OF iconf] by auto
+ have call: "P \<turnstile> \<langle>e\<bullet>M(ps),s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>null\<bullet>M(ps),s\<^sub>1,False\<rangle>" by(rule CallRedsObj[OF b1])
+ then have iconf2': "iconfs (shp s\<^sub>1) ps" using Red_preserves_iconf[OF wwf call] None CallNull by auto
+ have "P,shp s\<^sub>1 \<turnstile>\<^sub>b (ps,False) \<surd>" by(simp add: bconfs_def)
+ then have b2: "P \<turnstile> \<langle>ps,s\<^sub>1,False\<rangle> [\<rightarrow>]* \<langle>map Val vs,s\<^sub>2,False\<rangle>" using CallNull.hyps(4)[OF iconf2'] by auto
+ then show ?thesis using CallRedsNull[OF b1 b2] by fast
+ next
+ case (Some a')
+ then obtain b1 where b1: "P \<turnstile> \<langle>e,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>null,s\<^sub>1,b1\<rangle>"
+ by (metis (no_types, lifting) CallNull.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec)
+ have fass: "P \<turnstile> \<langle>e\<bullet>M(ps),s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>null\<bullet>M(ps),s\<^sub>1,b1\<rangle>" by(rule CallRedsObj[OF b1])
+ then have iconf2': "iconfs (shp s\<^sub>1) ps" using Red_preserves_iconf[OF wwf fass] CallNull by auto
+ have bconf2: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (ps,b) \<surd>" using CallNull.prems Some by simp
+ then have "P,shp s\<^sub>1 \<turnstile>\<^sub>b (ps,b1) \<surd>" using Red_preserves_bconf[OF wwf fass CallNull.prems] by simp
+ then have b2: "P \<turnstile> \<langle>ps,s\<^sub>1,b1\<rangle> [\<rightarrow>]* \<langle>map Val vs,s\<^sub>2,False\<rangle>" using CallNull.hyps(4)[OF iconf2'] by auto
+ then show ?thesis using CallRedsNull[OF b1 b2] by fast
+ qed
+next
+ case (CallParamsThrow e s\<^sub>0 v s\<^sub>1 es vs ex es' s\<^sub>2 M) show ?case
+ proof(cases "val_of e")
+ case None
+ then have iconf: "iconf (shp s\<^sub>0) e" using CallParamsThrow.prems(1) by simp
+ have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e,b) \<surd>" using CallParamsThrow.prems(2) None by simp
+ then have b1: "P \<turnstile> \<langle>e,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Val v,s\<^sub>1,False\<rangle>" using CallParamsThrow.hyps(2)[OF iconf] by auto
+ have call: "P \<turnstile> \<langle>e\<bullet>M(es),s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Val v\<bullet>M(es),s\<^sub>1,False\<rangle>" by(rule CallRedsObj[OF b1])
+ then have iconf2': "iconfs (shp s\<^sub>1) es" using Red_preserves_iconf[OF wwf call] None CallParamsThrow by auto
+ have "P,shp s\<^sub>1 \<turnstile>\<^sub>b (es,False) \<surd>" by(simp add: bconfs_def)
+ then have b2: "P \<turnstile> \<langle>es,s\<^sub>1,False\<rangle> [\<rightarrow>]* \<langle>map Val vs @ throw ex # es',s\<^sub>2,False\<rangle>"
+ using CallParamsThrow.hyps(4)[OF iconf2'] by auto
+ then show ?thesis using CallRedsThrowParams[OF b1 b2] by fast
+ next
+ case (Some a')
+ then obtain b1 where b1: "P \<turnstile> \<langle>e,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Val v,s\<^sub>1,b1\<rangle>"
+ by (metis (no_types, lifting) CallParamsThrow.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec)
+ have fass: "P \<turnstile> \<langle>e\<bullet>M(es),s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Val v\<bullet>M(es),s\<^sub>1,b1\<rangle>" by(rule CallRedsObj[OF b1])
+ then have iconf2': "iconfs (shp s\<^sub>1) es" using Red_preserves_iconf[OF wwf fass] CallParamsThrow by auto
+ have bconf2: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (es,b) \<surd>" using CallParamsThrow.prems Some by simp
+ then have "P,shp s\<^sub>1 \<turnstile>\<^sub>b (es,b1) \<surd>" using Red_preserves_bconf[OF wwf fass CallParamsThrow.prems] by simp
+ then have b2: "P \<turnstile> \<langle>es,s\<^sub>1,b1\<rangle> [\<rightarrow>]* \<langle>map Val vs @ throw ex # es',s\<^sub>2,False\<rangle>"
+ using CallParamsThrow.hyps(4)[OF iconf2'] by auto
+ then show ?thesis using CallRedsThrowParams[OF b1 b2] by fast
+ qed
+next
+ case (CallNone e s\<^sub>0 a s\<^sub>1 ps vs h\<^sub>2 l\<^sub>2 sh\<^sub>2 C fs M) show ?case
+ proof(cases "val_of e")
+ case None
+ then have iconf: "iconf (shp s\<^sub>0) e" using CallNone.prems(1) by simp
+ have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e,b) \<surd>" using CallNone.prems(2) None by simp
+ then have b1: "P \<turnstile> \<langle>e,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>addr a,s\<^sub>1,False\<rangle>" using CallNone.hyps(2)[OF iconf] by auto
+ have call: "P \<turnstile> \<langle>e\<bullet>M(ps),s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>addr a\<bullet>M(ps),s\<^sub>1,False\<rangle>" by(rule CallRedsObj[OF b1])
+ then have iconf2': "iconfs (shp s\<^sub>1) ps" using Red_preserves_iconf[OF wwf call] None CallNone by auto
+ have "P,shp s\<^sub>1 \<turnstile>\<^sub>b (ps,False) \<surd>" by(simp add: bconfs_def)
+ then have b2: "P \<turnstile> \<langle>ps,s\<^sub>1,False\<rangle> [\<rightarrow>]* \<langle>map Val vs,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\<rangle>"
+ using CallNone.hyps(4)[OF iconf2'] by auto
+ then show ?thesis using CallRedsNone[OF b1 b2 _ CallNone.hyps(6)] CallNone.hyps(5) by fastforce
+ next
+ case (Some a')
+ then obtain b1 where b1: "P \<turnstile> \<langle>e,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>addr a,s\<^sub>1,b1\<rangle>"
+ by (metis (no_types, lifting) CallNone.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec)
+ have fass: "P \<turnstile> \<langle>e\<bullet>M(ps),s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>addr a\<bullet>M(ps),s\<^sub>1,b1\<rangle>" by(rule CallRedsObj[OF b1])
+ then have iconf2': "iconfs (shp s\<^sub>1) ps" using Red_preserves_iconf[OF wwf fass] CallNone by auto
+ have bconf2: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (ps,b) \<surd>" using CallNone.prems Some by simp
+ then have "P,shp s\<^sub>1 \<turnstile>\<^sub>b (ps,b1) \<surd>" using Red_preserves_bconf[OF wwf fass CallNone.prems] by simp
+ then have b2: "P \<turnstile> \<langle>ps,s\<^sub>1,b1\<rangle> [\<rightarrow>]* \<langle>map Val vs,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\<rangle>"
+ using CallNone.hyps(4)[OF iconf2'] by auto
+ then show ?thesis using CallRedsNone[OF b1 b2 _ CallNone.hyps(6)] CallNone.hyps(5) by fastforce
+ qed
+next
+ case (CallStatic e s\<^sub>0 a s\<^sub>1 ps vs h\<^sub>2 l\<^sub>2 sh\<^sub>2 C fs M Ts T m D) show ?case
+ proof(cases "val_of e")
+ case None
+ then have iconf: "iconf (shp s\<^sub>0) e" using CallStatic.prems(1) by simp
+ have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e,b) \<surd>" using CallStatic.prems(2) None by simp
+ then have b1: "P \<turnstile> \<langle>e,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>addr a,s\<^sub>1,False\<rangle>" using CallStatic.hyps(2)[OF iconf] by auto
+ have call: "P \<turnstile> \<langle>e\<bullet>M(ps),s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>addr a\<bullet>M(ps),s\<^sub>1,False\<rangle>" by(rule CallRedsObj[OF b1])
+ then have iconf2': "iconfs (shp s\<^sub>1) ps" using Red_preserves_iconf[OF wwf call] None CallStatic by auto
+ have "P,shp s\<^sub>1 \<turnstile>\<^sub>b (ps,False) \<surd>" by(simp add: bconfs_def)
+ then have b2: "P \<turnstile> \<langle>ps,s\<^sub>1,False\<rangle> [\<rightarrow>]* \<langle>map Val vs,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\<rangle>"
+ using CallStatic.hyps(4)[OF iconf2'] by auto
+ then show ?thesis using CallRedsStatic[OF b1 b2 _ CallStatic.hyps(6)] CallStatic.hyps(5) by fastforce
+ next
+ case (Some a')
+ then obtain b1 where b1: "P \<turnstile> \<langle>e,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>addr a,s\<^sub>1,b1\<rangle>"
+ by (metis (no_types, lifting) CallStatic.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec)
+ have call: "P \<turnstile> \<langle>e\<bullet>M(ps),s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>addr a\<bullet>M(ps),s\<^sub>1,b1\<rangle>" by(rule CallRedsObj[OF b1])
+ then have iconf2': "iconfs (shp s\<^sub>1) ps" using Red_preserves_iconf[OF wwf call] CallStatic by auto
+ have bconf2: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (ps,b) \<surd>" using CallStatic.prems Some by simp
+ then have "P,shp s\<^sub>1 \<turnstile>\<^sub>b (ps,b1) \<surd>" using Red_preserves_bconf[OF wwf call CallStatic.prems] by simp
+ then have b2: "P \<turnstile> \<langle>ps,s\<^sub>1,b1\<rangle> [\<rightarrow>]* \<langle>map Val vs,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\<rangle>"
+ using CallStatic.hyps(4)[OF iconf2'] by auto
+ then show ?thesis using CallRedsStatic[OF b1 b2 _ CallStatic.hyps(6)] CallStatic.hyps(5) by fastforce
+ qed
+next
+ case (Call e s\<^sub>0 a s\<^sub>1 ps vs h\<^sub>2 l\<^sub>2 sh\<^sub>2 C fs M Ts T pns body D l\<^sub>2' e' h\<^sub>3 l\<^sub>3 sh\<^sub>3) show ?case
+ proof(cases "val_of e")
+ case None
+ then have iconf: "iconf (shp s\<^sub>0) e" using Call.prems(1) by simp
+ have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e,b) \<surd>" using Call.prems(2) None by simp
+ then have b1: "P \<turnstile> \<langle>e,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>addr a,s\<^sub>1,False\<rangle>" using Call.hyps(2)[OF iconf] by auto
+ have call: "P \<turnstile> \<langle>e\<bullet>M(ps),s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>addr a\<bullet>M(ps),s\<^sub>1,False\<rangle>" by(rule CallRedsObj[OF b1])
+ then have iconf2: "iconfs (shp s\<^sub>1) ps" using Red_preserves_iconf[OF wwf call] None Call by auto
+ have "P,shp s\<^sub>1 \<turnstile>\<^sub>b (ps,False) \<surd>" by(simp add: bconfs_def)
+ then have b2: "P \<turnstile> \<langle>ps,s\<^sub>1,False\<rangle> [\<rightarrow>]* \<langle>map Val vs,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\<rangle>"
+ using Call.hyps(4)[OF iconf2] by simp
+ have iconf3: "iconf (shp (h\<^sub>2, l\<^sub>2', sh\<^sub>2)) body"
+ by(rule nsub_RI_iconf[OF sees_wwf_nsub_RI[OF wwf Call.hyps(6)]])
+ have "P,shp (h\<^sub>2, l\<^sub>2', sh\<^sub>2) \<turnstile>\<^sub>b (body,False) \<surd>" by(simp add: bconf_def)
+ then have b3: "P \<turnstile> \<langle>body,(h\<^sub>2, l\<^sub>2', sh\<^sub>2),False\<rangle> \<rightarrow>* \<langle>e',(h\<^sub>3, l\<^sub>3, sh\<^sub>3),False\<rangle>"
+ using Call.hyps(10)[OF iconf3] by simp
+ show ?thesis by(rule CallRedsFinal[OF wwf b1 b2 Call.hyps(5-8) b3 eval_final[OF Call.hyps(9)]])
+ next
+ case (Some a')
+ then obtain b1 where b1: "P \<turnstile> \<langle>e,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>addr a,s\<^sub>1,b1\<rangle>"
+ by (metis (no_types, lifting) Call.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec)
+ have call: "P \<turnstile> \<langle>e\<bullet>M(ps),s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>addr a\<bullet>M(ps),s\<^sub>1,b1\<rangle>" by(rule CallRedsObj[OF b1])
+ then have iconf2': "iconfs (shp s\<^sub>1) ps" using Red_preserves_iconf[OF wwf call] Call by auto
+ have bconf2: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (ps,b) \<surd>" using Call.prems Some by simp
+ then have "P,shp s\<^sub>1 \<turnstile>\<^sub>b (ps,b1) \<surd>" using Red_preserves_bconf[OF wwf call Call.prems] by simp
+ then have b2: "P \<turnstile> \<langle>ps,s\<^sub>1,b1\<rangle> [\<rightarrow>]* \<langle>map Val vs,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\<rangle>"
+ using Call.hyps(4)[OF iconf2'] by auto
+ have iconf3: "iconf (shp (h\<^sub>2, l\<^sub>2', sh\<^sub>2)) body"
+ by(rule nsub_RI_iconf[OF sees_wwf_nsub_RI[OF wwf Call.hyps(6)]])
+ have "P,shp (h\<^sub>2, l\<^sub>2', sh\<^sub>2) \<turnstile>\<^sub>b (body,False) \<surd>" by(simp add: bconf_def)
+ then have b3: "P \<turnstile> \<langle>body,(h\<^sub>2, l\<^sub>2', sh\<^sub>2),False\<rangle> \<rightarrow>* \<langle>e',(h\<^sub>3, l\<^sub>3, sh\<^sub>3),False\<rangle>"
+ using Call.hyps(10)[OF iconf3] by simp
+ show ?thesis by(rule CallRedsFinal[OF wwf b1 b2 Call.hyps(5-8) b3 eval_final[OF Call.hyps(9)]])
+ qed
+next
+ case (SCallParamsThrow es s\<^sub>0 vs ex es' s\<^sub>2 C M) show ?case
+ proof(cases "map_vals_of es")
+ case None
+ then have iconf: "iconfs (shp s\<^sub>0) es" using SCallParamsThrow.prems(1) by simp
+ have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (es,b) \<surd>" using SCallParamsThrow.prems(2) None by simp
+ then have b1: "P \<turnstile> \<langle>es,s\<^sub>0,b\<rangle> [\<rightarrow>]* \<langle>map Val vs @ throw ex # es',s\<^sub>2,False\<rangle>"
+ using SCallParamsThrow.hyps(2)[OF iconf] by simp
+ show ?thesis using SCallRedsThrowParams[OF b1] by simp
+ next
+ case (Some vs')
+ then have "es = map Val vs'" by(rule map_vals_of_spec)
+ then show ?thesis using evals_finals_same[OF _ SCallParamsThrow.hyps(1)] map_Val_nthrow_neq
+ by auto
+ qed
+next
+ case (SCallNone ps s\<^sub>0 vs s\<^sub>2 C M) show ?case
+ proof(cases "map_vals_of ps")
+ case None
+ then have iconf: "iconfs (shp s\<^sub>0) ps" using SCallNone.prems(1) by simp
+ have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (ps,b) \<surd>" using SCallNone.prems(2) None by simp
+ then have b1: "P \<turnstile> \<langle>ps,s\<^sub>0,b\<rangle> [\<rightarrow>]* \<langle>map Val vs,s\<^sub>2,False\<rangle>" using SCallNone.hyps(2)[OF iconf] by auto
+ then show ?thesis using SCallRedsNone[OF b1 SCallNone.hyps(3)] SCallNone.hyps(1) by simp
+ next
+ case (Some vs')
+ then have ps: "ps = map Val vs'" by(rule map_vals_of_spec)
+ then have s\<^sub>0: "s\<^sub>0 = s\<^sub>2" using SCallNone.hyps(1) evals_finals_same by blast
+ then show ?thesis using RedSCallNone[OF SCallNone.hyps(3)] ps by(cases s\<^sub>2, auto)
+ qed
+next
+ case (SCallNonStatic ps s\<^sub>0 vs s\<^sub>2 C M Ts T m D) show ?case
+ proof(cases "map_vals_of ps")
+ case None
+ then have iconf: "iconfs (shp s\<^sub>0) ps" using SCallNonStatic.prems(1) by simp
+ have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (ps,b) \<surd>" using SCallNonStatic.prems(2) None by simp
+ then have b1: "P \<turnstile> \<langle>ps,s\<^sub>0,b\<rangle> [\<rightarrow>]* \<langle>map Val vs,s\<^sub>2,False\<rangle>" using SCallNonStatic.hyps(2)[OF iconf] by auto
+ then show ?thesis using SCallRedsNonStatic[OF b1 SCallNonStatic.hyps(3)] SCallNonStatic.hyps(1) by simp
+ next
+ case (Some vs')
+ then have ps: "ps = map Val vs'" by(rule map_vals_of_spec)
+ then have s\<^sub>0: "s\<^sub>0 = s\<^sub>2" using SCallNonStatic.hyps(1) evals_finals_same by blast
+ then show ?thesis using RedSCallNonStatic[OF SCallNonStatic.hyps(3)] ps by(cases s\<^sub>2, auto)
+ qed
+next
+ case (SCallInitThrow ps s\<^sub>0 vs h\<^sub>1 l\<^sub>1 sh\<^sub>1 C M Ts T pns body D a s') show ?case
+ proof(cases "map_vals_of ps")
+ case None
+ then have iconf: "iconfs (shp s\<^sub>0) ps" using SCallInitThrow.prems(1) by simp
+ have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (ps,b) \<surd>" using SCallInitThrow.prems(2) None by simp
+ then have b1: "P \<turnstile> \<langle>ps,s\<^sub>0,b\<rangle> [\<rightarrow>]* \<langle>map Val vs,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\<rangle>"
+ using SCallInitThrow.hyps(2)[OF iconf] by auto
+ have bconf2: "P,shp (h\<^sub>1, l\<^sub>1, sh\<^sub>1) \<turnstile>\<^sub>b (INIT D ([D],False) \<leftarrow> unit,False) \<surd>" by(simp add: bconf_def)
+ then have b2: "P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\<rangle> \<rightarrow>* \<langle>throw a,s',False\<rangle>"
+ using SCallInitThrow.hyps(7) by auto
+ then show ?thesis using SCallInitThrowReds[OF wwf b1 SCallInitThrow.hyps(3-5)]
+ by(cases s', auto)
+ next
+ case (Some vs')
+ have ps: "ps = map Val vs'" by(rule map_vals_of_spec[OF Some])
+ then have vs: "vs = vs' \<and> s\<^sub>0 = (h\<^sub>1, l\<^sub>1, sh\<^sub>1)"
+ using evals_finals_same[OF _ SCallInitThrow.hyps(1)] map_Val_eq by auto
+ show ?thesis
+ proof(cases b)
+ case True
+ obtain sfs where shC: "sh\<^sub>1 D = \<lfloor>(sfs, Processing)\<rfloor>"
+ using SCallInitThrow.hyps(3,4) SCallInitThrow.prems(2) True Some vs
+ by(auto simp: bconf_def initPD_def dest: sees_method_fun)
+ then show ?thesis using init_ProcessingE[OF _ SCallInitThrow.hyps(6)] by blast
+ next
+ case False
+ then have b1: "P \<turnstile> \<langle>ps,s\<^sub>0,b\<rangle> [\<rightarrow>]* \<langle>map Val vs,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\<rangle>" using ps vs by simp
+ have bconf2: "P,shp (h\<^sub>1, l\<^sub>1, sh\<^sub>1) \<turnstile>\<^sub>b (INIT D ([D],False) \<leftarrow> unit,False) \<surd>" by(simp add: bconf_def)
+ then have b2: "P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\<rangle> \<rightarrow>* \<langle>throw a,s',False\<rangle>"
+ using SCallInitThrow.hyps(7) by auto
+ then show ?thesis using SCallInitThrowReds[OF wwf b1 SCallInitThrow.hyps(3-5)] by(cases s', auto)
+ qed
+ qed
+next
+ case (SCallInit ps s\<^sub>0 vs h\<^sub>1 l\<^sub>1 sh\<^sub>1 C M Ts T pns body D v' h\<^sub>2 l\<^sub>2 sh\<^sub>2 l\<^sub>2' e' h\<^sub>3 l\<^sub>3 sh\<^sub>3) show ?case
+ proof(cases "map_vals_of ps")
+ case None
+ then have iconf: "iconfs (shp s\<^sub>0) ps" using SCallInit.prems(1) by simp
+ have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (ps,b) \<surd>" using SCallInit.prems(2) None by simp
+ then have b1: "P \<turnstile> \<langle>ps,s\<^sub>0,b\<rangle> [\<rightarrow>]* \<langle>map Val vs,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\<rangle>"
+ using SCallInit.hyps(2)[OF iconf] by auto
+ have bconf2: "P,shp (h\<^sub>1, l\<^sub>1, sh\<^sub>1) \<turnstile>\<^sub>b (INIT D ([D],False) \<leftarrow> unit,False) \<surd>" by(simp add: bconf_def)
+ then have b2: "P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\<rangle> \<rightarrow>* \<langle>Val v',(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\<rangle>"
+ using SCallInit.hyps(7) by auto
+ have iconf3: "iconf (shp (h\<^sub>2, l\<^sub>2', sh\<^sub>2)) body"
+ by(rule nsub_RI_iconf[OF sees_wwf_nsub_RI[OF wwf SCallInit.hyps(3)]])
+ have "P,shp (h\<^sub>2, l\<^sub>2', sh\<^sub>2) \<turnstile>\<^sub>b (body,False) \<surd>" by(simp add: bconf_def)
+ then have b3: "P \<turnstile> \<langle>body,(h\<^sub>2, l\<^sub>2', sh\<^sub>2),False\<rangle> \<rightarrow>* \<langle>e',(h\<^sub>3, l\<^sub>3, sh\<^sub>3),False\<rangle>"
+ using SCallInit.hyps(11)[OF iconf3] by simp
+ show ?thesis by(rule SCallInitReds[OF wwf b1 SCallInit.hyps(3-5) b2 SCallInit.hyps(8-9)
+ b3 eval_final[OF SCallInit.hyps(10)]])
+ next
+ case (Some vs')
+ then have ps: "ps = map Val vs'" by(rule map_vals_of_spec)
+ then have vs: "vs = vs' \<and> s\<^sub>0 = (h\<^sub>1, l\<^sub>1, sh\<^sub>1)"
+ using evals_finals_same[OF _ SCallInit.hyps(1)] map_Val_eq by auto
+ show ?thesis
+ proof(cases b)
+ case True
+ then have b1: "P \<turnstile> \<langle>ps,s\<^sub>0,b\<rangle> [\<rightarrow>]* \<langle>map Val vs,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),b\<rangle>" using ps vs by simp
+ obtain sfs where shC: "sh\<^sub>1 D = \<lfloor>(sfs, Processing)\<rfloor>"
+ using SCallInit.hyps(3,4) SCallInit.prems(2) True Some vs
+ by(auto simp: bconf_def initPD_def dest: sees_method_fun)
+ then have s': "(h\<^sub>1, l\<^sub>1, sh\<^sub>1) = (h\<^sub>2, l\<^sub>2, sh\<^sub>2)" using init_ProcessingE[OF _ SCallInit.hyps(6)] by simp
+ have iconf3: "iconf (shp (h\<^sub>2, l\<^sub>2', sh\<^sub>2)) body"
+ by(rule nsub_RI_iconf[OF sees_wwf_nsub_RI[OF wwf SCallInit.hyps(3)]])
+ have "P,shp (h\<^sub>2, l\<^sub>2', sh\<^sub>2) \<turnstile>\<^sub>b (body,False) \<surd>" by(simp add: bconf_def)
+ then have b3: "P \<turnstile> \<langle>body,(h\<^sub>2, l\<^sub>2', sh\<^sub>2),False\<rangle> \<rightarrow>* \<langle>e',(h\<^sub>3, l\<^sub>3, sh\<^sub>3),False\<rangle>"
+ using SCallInit.hyps(11)[OF iconf3] by simp
+ then have b3': "P \<turnstile> \<langle>body,(h\<^sub>1, l\<^sub>2', sh\<^sub>1),False\<rangle> \<rightarrow>* \<langle>e',(h\<^sub>3, l\<^sub>3, sh\<^sub>3),False\<rangle>"
+ using s' by simp
+ then show ?thesis using SCallInitProcessingReds[OF wwf b1 SCallInit.hyps(3) shC
+ SCallInit.hyps(8-9) b3' eval_final[OF SCallInit.hyps(10)]] s' by simp
+ next
+ case False
+ then have b1: "P \<turnstile> \<langle>ps,s\<^sub>0,b\<rangle> [\<rightarrow>]* \<langle>map Val vs,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\<rangle>" using ps vs by simp
+ have bconf2: "P,shp (h\<^sub>1, l\<^sub>1, sh\<^sub>1) \<turnstile>\<^sub>b (INIT D ([D],False) \<leftarrow> unit,False) \<surd>" by(simp add: bconf_def)
+ then have b2: "P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\<rangle> \<rightarrow>* \<langle>Val v',(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\<rangle>"
+ using SCallInit.hyps(7) by auto
+ have iconf3: "iconf (shp (h\<^sub>2, l\<^sub>2', sh\<^sub>2)) body"
+ by(rule nsub_RI_iconf[OF sees_wwf_nsub_RI[OF wwf SCallInit.hyps(3)]])
+ have "P,shp (h\<^sub>2, l\<^sub>2', sh\<^sub>2) \<turnstile>\<^sub>b (body,False) \<surd>" by(simp add: bconf_def)
+ then have b3: "P \<turnstile> \<langle>body,(h\<^sub>2, l\<^sub>2', sh\<^sub>2),False\<rangle> \<rightarrow>* \<langle>e',(h\<^sub>3, l\<^sub>3, sh\<^sub>3),False\<rangle>"
+ using SCallInit.hyps(11)[OF iconf3] by simp
+ show ?thesis by(rule SCallInitReds[OF wwf b1 SCallInit.hyps(3-5) b2 SCallInit.hyps(8-9)
+ b3 eval_final[OF SCallInit.hyps(10)]])
+ qed
+ qed
+next
+ case (SCall ps s\<^sub>0 vs h\<^sub>2 l\<^sub>2 sh\<^sub>2 C M Ts T pns body D sfs l\<^sub>2' e' h\<^sub>3 l\<^sub>3 sh\<^sub>3) show ?case
+ proof(cases "map_vals_of ps")
+ case None
+ then have iconf: "iconfs (shp s\<^sub>0) ps" using SCall.prems(1) by simp
+ have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (ps,b) \<surd>" using SCall.prems(2) None by simp
+ then have b1: "P \<turnstile> \<langle>ps,s\<^sub>0,b\<rangle> [\<rightarrow>]* \<langle>map Val vs,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\<rangle>"
+ using SCall.hyps(2)[OF iconf] by auto
+ have iconf3: "iconf (shp (h\<^sub>2, l\<^sub>2', sh\<^sub>2)) body"
+ by(rule nsub_RI_iconf[OF sees_wwf_nsub_RI[OF wwf SCall.hyps(3)]])
+ have "P,shp (h\<^sub>2, l\<^sub>2', sh\<^sub>2) \<turnstile>\<^sub>b (body,False) \<surd>" by(simp add: bconf_def)
+ then have b2: "P \<turnstile> \<langle>body,(h\<^sub>2, l\<^sub>2', sh\<^sub>2),False\<rangle> \<rightarrow>* \<langle>e',(h\<^sub>3, l\<^sub>3, sh\<^sub>3),False\<rangle>"
+ using SCall.hyps(8)[OF iconf3] by simp
+ show ?thesis by(rule SCallRedsFinal[OF wwf b1 SCall.hyps(3-6) b2 eval_final[OF SCall.hyps(7)]])
+ next
+ case (Some vs')
+ then have ps: "ps = map Val vs'" by(rule map_vals_of_spec)
+ then have vs: "vs = vs' \<and> s\<^sub>0 = (h\<^sub>2, l\<^sub>2, sh\<^sub>2)"
+ using evals_finals_same[OF _ SCall.hyps(1)] map_Val_eq by auto
+ then have b1: "P \<turnstile> \<langle>ps,s\<^sub>0,b\<rangle> [\<rightarrow>]* \<langle>map Val vs,(h\<^sub>2, l\<^sub>2, sh\<^sub>2),b\<rangle>" using ps by simp
+ have iconf3: "iconf (shp (h\<^sub>2, l\<^sub>2', sh\<^sub>2)) body"
+ by(rule nsub_RI_iconf[OF sees_wwf_nsub_RI[OF wwf SCall.hyps(3)]])
+ have "P,shp (h\<^sub>2, l\<^sub>2', sh\<^sub>2) \<turnstile>\<^sub>b (body,False) \<surd>" by(simp add: bconf_def)
+ then have b2: "P \<turnstile> \<langle>body,(h\<^sub>2, l\<^sub>2', sh\<^sub>2),False\<rangle> \<rightarrow>* \<langle>e',(h\<^sub>3, l\<^sub>3, sh\<^sub>3),False\<rangle>"
+ using SCall.hyps(8)[OF iconf3] by simp
+ show ?thesis by(rule SCallRedsFinal[OF wwf b1 SCall.hyps(3-6) b2 eval_final[OF SCall.hyps(7)]])
+ qed
+next
+ case (Block e\<^sub>0 h\<^sub>0 l\<^sub>0 V sh\<^sub>0 e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 T)
+ have iconf: "iconf (shp (h\<^sub>0, l\<^sub>0(V := None), sh\<^sub>0)) e\<^sub>0"
+ using Block.prems(1) by (auto simp: assigned_def)
+ have bconf: "P,shp (h\<^sub>0, l\<^sub>0(V := None), sh\<^sub>0) \<turnstile>\<^sub>b (e\<^sub>0,b) \<surd>" using Block.prems(2)
+ by(auto simp: bconf_def)
+ then have b': "P \<turnstile> \<langle>e\<^sub>0,(h\<^sub>0, l\<^sub>0(V := None), sh\<^sub>0),b\<rangle> \<rightarrow>* \<langle>e\<^sub>1,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\<rangle>"
+ using Block.hyps(2)[OF iconf] by auto
+ have fin: "final e\<^sub>1" using Block by(auto dest: eval_final)
+ thus ?case using BlockRedsFinal[OF b' fin] by simp
+next
+ case (Seq e\<^sub>0 s\<^sub>0 v s\<^sub>1 e\<^sub>1 e\<^sub>2 s\<^sub>2)
+ then have iconf: "iconf (shp s\<^sub>0) e\<^sub>0" using Seq.prems(1)
+ by(auto dest: val_of_spec lass_val_of_spec)
+ have b1: "\<exists>b1. P \<turnstile> \<langle>e\<^sub>0,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Val v,s\<^sub>1,b1\<rangle>"
+ proof(cases "val_of e\<^sub>0")
+ case None show ?thesis
+ proof(cases "lass_val_of e\<^sub>0")
+ case lNone:None
+ then have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e\<^sub>0,b) \<surd>" using Seq.prems(2) None by simp
+ then have "P \<turnstile> \<langle>e\<^sub>0,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Val v,s\<^sub>1,False\<rangle>" using iconf Seq.hyps(2) by auto
+ then show ?thesis by fast
+ next
+ case (Some p)
+ obtain V' v' where p: "p = (V',v')" and e\<^sub>0: "e\<^sub>0 = V':=Val v'"
+ using lass_val_of_spec[OF Some] by(cases p, auto)
+ obtain h l sh h' l' sh' where s\<^sub>0: "s\<^sub>0 = (h,l,sh)" and s\<^sub>1: "s\<^sub>1 = (h',l',sh')" by(cases s\<^sub>0, cases s\<^sub>1)
+ then have eval: "P \<turnstile> \<langle>e\<^sub>0,(h,l,sh)\<rangle> \<Rightarrow> \<langle>Val v,(h',l',sh')\<rangle>" using Seq.hyps(1) by simp
+ then have s\<^sub>1': "Val v = unit \<and> h' = h \<and> l' = l(V' \<mapsto> v') \<and> sh' = sh"
+ using lass_val_of_eval[OF Some eval] p e\<^sub>0 by simp
+ then have "P \<turnstile> \<langle>e\<^sub>0,s\<^sub>0,b\<rangle> \<rightarrow> \<langle>Val v,s\<^sub>1,b\<rangle>" using e\<^sub>0 s\<^sub>0 s\<^sub>1 by(auto intro: RedLAss)
+ then show ?thesis by auto
+ qed
+ next
+ case (Some a)
+ then have "e\<^sub>0 = Val v" and "s\<^sub>0 = s\<^sub>1" using Seq.hyps(1) eval_cases(3) val_of_spec by blast+
+ then show ?thesis using Seq by auto
+ qed
+ then obtain b1 where b1': "P \<turnstile> \<langle>e\<^sub>0,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Val v,s\<^sub>1,b1\<rangle>" by clarsimp
+ have seq: "P \<turnstile> \<langle>e\<^sub>0;;e\<^sub>1,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Val v;;e\<^sub>1,s\<^sub>1,b1\<rangle>" by(rule SeqReds[OF b1'])
+ then have iconf2: "iconf (shp s\<^sub>1) e\<^sub>1" using Red_preserves_iconf[OF wwf seq] Seq nsub_RI_iconf
+ by auto
+ have "P,shp s\<^sub>1 \<turnstile>\<^sub>b (Val v;; e\<^sub>1,b1) \<surd>" by(rule Red_preserves_bconf[OF wwf seq Seq.prems])
+ then have bconf2: "P,shp s\<^sub>1 \<turnstile>\<^sub>b (e\<^sub>1,b1) \<surd>" by simp
+ have b2: "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>1,b1\<rangle> \<rightarrow>* \<langle>e\<^sub>2,s\<^sub>2,False\<rangle>" by(rule Seq.hyps(4)[OF iconf2 bconf2])
+ then show ?case using SeqReds2[OF b1' b2] by fast
+next
+ case (SeqThrow e\<^sub>0 s\<^sub>0 a s\<^sub>1 e\<^sub>1 b)
+ have notVal: "val_of e\<^sub>0 = None" "lass_val_of e\<^sub>0 = None"
+ using SeqThrow.hyps(1) eval_throw_nonVal eval_throw_nonLAss by auto
+ thus ?case using SeqThrow notVal by(auto dest!:eval_final dest: SeqRedsThrow)
+next
+ case (CondT e s\<^sub>0 s\<^sub>1 e\<^sub>1 e' s\<^sub>2 e\<^sub>2)
+ then have iconf: "iconf (shp s\<^sub>0) e" using CondT.prems(1) by auto
+ have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e,b) \<surd>" using CondT.prems(2) by auto
+ then have b1: "P \<turnstile> \<langle>e,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>true,s\<^sub>1,False\<rangle>" using iconf CondT.hyps(2) by auto
+ have cond: "P \<turnstile> \<langle>if (e) e\<^sub>1 else e\<^sub>2,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>if (true) e\<^sub>1 else e\<^sub>2,s\<^sub>1,False\<rangle>" by(rule CondReds[OF b1])
+ then have iconf2': "iconf (shp s\<^sub>1) e\<^sub>1" using Red_preserves_iconf[OF wwf cond] CondT nsub_RI_iconf
+ by auto
+ have "P,shp s\<^sub>1 \<turnstile>\<^sub>b (e\<^sub>1,False) \<surd>" by(simp add: bconf_def)
+ then have b2: "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>1,False\<rangle> \<rightarrow>* \<langle>e',s\<^sub>2,False\<rangle>" using CondT.hyps(4)[OF iconf2'] by auto
+ then show ?case using CondReds2T[OF b1 b2] by fast
+next
+ case (CondF e s\<^sub>0 s\<^sub>1 e\<^sub>2 e' s\<^sub>2 e\<^sub>1)
+ then have iconf: "iconf (shp s\<^sub>0) e" using CondF.prems(1) by auto
+ have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e,b) \<surd>" using CondF.prems(2) by auto
+ then have b1: "P \<turnstile> \<langle>e,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>false,s\<^sub>1,False\<rangle>" using iconf CondF.hyps(2) by auto
+ have cond: "P \<turnstile> \<langle>if (e) e\<^sub>1 else e\<^sub>2,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>if (false) e\<^sub>1 else e\<^sub>2,s\<^sub>1,False\<rangle>" by(rule CondReds[OF b1])
+ then have iconf2': "iconf (shp s\<^sub>1) e\<^sub>2" using Red_preserves_iconf[OF wwf cond] CondF nsub_RI_iconf
+ by auto
+ have "P,shp s\<^sub>1 \<turnstile>\<^sub>b (e\<^sub>2,False) \<surd>" by(simp add: bconf_def)
+ then have b2: "P \<turnstile> \<langle>e\<^sub>2,s\<^sub>1,False\<rangle> \<rightarrow>* \<langle>e',s\<^sub>2,False\<rangle>" using CondF.hyps(4)[OF iconf2'] by auto
+ then show ?case using CondReds2F[OF b1 b2] by fast
+next
+ case CondThrow thus ?case by(auto dest!:eval_final dest:CondRedsThrow)
+next
+ case (WhileF e s\<^sub>0 s\<^sub>1 c)
+ then have iconf: "iconf (shp s\<^sub>0) e" using nsub_RI_iconf by auto
+ then have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e,b) \<surd>" using WhileF.prems(2) by(simp add: bconf_def)
+ then have b': "P \<turnstile> \<langle>e,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>false,s\<^sub>1,False\<rangle>" using WhileF.hyps(2) iconf by auto
+ thus ?case using WhileFReds[OF b'] by fast
+next
+ case (WhileT e s\<^sub>0 s\<^sub>1 c v\<^sub>1 s\<^sub>2 e\<^sub>3 s\<^sub>3)
+ then have iconf: "iconf (shp s\<^sub>0) e" using nsub_RI_iconf by auto
+ then have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e,b) \<surd>" using WhileT.prems(2) by(simp add: bconf_def)
+ then have b1: "P \<turnstile> \<langle>e,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>true,s\<^sub>1,False\<rangle>" using WhileT.hyps(2) iconf by auto
+ have iconf2: "iconf (shp s\<^sub>1) c" using WhileT.prems(1) nsub_RI_iconf by auto
+ have bconf2: "P,shp s\<^sub>1 \<turnstile>\<^sub>b (c,False) \<surd>" by(simp add: bconf_def)
+ then have b2: "P \<turnstile> \<langle>c,s\<^sub>1,False\<rangle> \<rightarrow>* \<langle>Val v\<^sub>1,s\<^sub>2,False\<rangle>" using WhileT.hyps(4) iconf2 by auto
+ have iconf3: "iconf (shp s\<^sub>2) (while (e) c)" using WhileT.prems(1) by auto
+ have "P,shp s\<^sub>2 \<turnstile>\<^sub>b (while (e) c,False) \<surd>" by(simp add: bconf_def)
+ then have b3: "P \<turnstile> \<langle>while (e) c,s\<^sub>2,False\<rangle> \<rightarrow>* \<langle>e\<^sub>3,s\<^sub>3,False\<rangle>" using WhileT.hyps(6) iconf3 by auto
+ show ?case using WhileTReds[OF b1 b2 b3] by fast
+next
+ case WhileCondThrow thus ?case
+ by (metis (no_types, lifting) WhileRedsThrow iconf.simps(16) bconf_While bconf_def nsub_RI_iconf)
+next
+ case (WhileBodyThrow e s\<^sub>0 s\<^sub>1 c e' s\<^sub>2)
+ then have iconf: "iconf (shp s\<^sub>0) e" using nsub_RI_iconf by auto
+ then have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e,b) \<surd>" using WhileBodyThrow.prems(2) by(simp add: bconf_def)
+ then have b1: "P \<turnstile> \<langle>e,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>true,s\<^sub>1,False\<rangle>" using WhileBodyThrow.hyps(2) iconf by auto
+ have iconf2: "iconf (shp s\<^sub>1) c" using WhileBodyThrow.prems(1) nsub_RI_iconf by auto
+ have bconf2: "P,shp s\<^sub>1 \<turnstile>\<^sub>b (c,False) \<surd>" by(simp add: bconf_def)
+ then have b2: "P \<turnstile> \<langle>c,s\<^sub>1,False\<rangle> \<rightarrow>* \<langle>throw e',s\<^sub>2,False\<rangle>" using WhileBodyThrow.hyps(4) iconf2 by auto
+ show ?case using WhileTRedsThrow[OF b1 b2] by fast
+next
+ case Throw thus ?case by (meson ThrowReds iconf.simps(17) bconf_Throw)
+next
+ case ThrowNull thus ?case by (meson ThrowRedsNull iconf.simps(17) bconf_Throw)
+next
+ case ThrowThrow thus ?case by (meson ThrowRedsThrow iconf.simps(17) bconf_Throw)
+next
+ case Try thus ?case by (meson TryRedsVal iconf.simps(18) bconf_Try)
+next
+ case (TryCatch e\<^sub>1 s\<^sub>0 a h\<^sub>1 l\<^sub>1 sh\<^sub>1 D fs C e\<^sub>2 V e\<^sub>2' h\<^sub>2 l\<^sub>2 sh\<^sub>2)
+ then have b1: "P \<turnstile> \<langle>e\<^sub>1,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Throw a,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\<rangle>" by auto
+ have Try: "P \<turnstile> \<langle>try e\<^sub>1 catch(C V) e\<^sub>2,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>try (Throw a) catch(C V) e\<^sub>2,(h\<^sub>1, l\<^sub>1, sh\<^sub>1),False\<rangle>"
+ by(rule TryReds[OF b1])
+ have iconf: "iconf sh\<^sub>1 e\<^sub>2" using Red_preserves_iconf[OF wwf Try] TryCatch nsub_RI_iconf
+ by auto
+ have bconf: "P,shp (h\<^sub>1, l\<^sub>1(V \<mapsto> Addr a), sh\<^sub>1) \<turnstile>\<^sub>b (e\<^sub>2,False) \<surd>" by(simp add: bconf_def)
+ then have b2: "P \<turnstile> \<langle>e\<^sub>2,(h\<^sub>1, l\<^sub>1(V \<mapsto> Addr a), sh\<^sub>1),False\<rangle> \<rightarrow>* \<langle>e\<^sub>2',(h\<^sub>2, l\<^sub>2, sh\<^sub>2),False\<rangle>"
+ using TryCatch.hyps(6) iconf by auto
+ thus ?case using TryCatchRedsFinal[OF b1] TryCatch.hyps(3-5) by (meson eval_final)
+next
+ case TryThrow thus ?case by (meson TryRedsFail iconf.simps(18) bconf_Try)
+next
+ case Nil thus ?case by(auto simp: bconfs_def)
+next
+ case (Cons e s\<^sub>0 v s\<^sub>1 es es' s\<^sub>2) show ?case
+ proof(cases "val_of e")
+ case None
+ then have iconf: "iconf (shp s\<^sub>0) e" using Cons.prems(1) by simp
+ have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e,b) \<surd>" using Cons.prems(2) None by simp
+ then have b1: "P \<turnstile> \<langle>e,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Val v,s\<^sub>1,False\<rangle>" using Cons.hyps(2) iconf by auto
+ have cons: "P \<turnstile> \<langle>e # es,s\<^sub>0,b\<rangle> [\<rightarrow>]* \<langle>Val v # es,s\<^sub>1,False\<rangle>" by(rule ListReds1[OF b1])
+ then have iconf2': "iconfs (shp s\<^sub>1) es" using Reds_preserves_iconf[OF wwf cons] None Cons by auto
+ have "P,shp s\<^sub>1 \<turnstile>\<^sub>b (es,False) \<surd>" by(simp add: bconfs_def)
+ then have b2: "P \<turnstile> \<langle>es,s\<^sub>1,False\<rangle> [\<rightarrow>]* \<langle>es',s\<^sub>2,False\<rangle>" using Cons.hyps(4)[OF iconf2'] by auto
+ show ?thesis using ListRedsVal[OF b1 b2] by auto
+ next
+ case (Some a)
+ then obtain b1 where b1: "P \<turnstile> \<langle>e,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>Val v,s\<^sub>1,b1\<rangle>"
+ by (metis (no_types, lifting) Cons.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec)
+ have cons: "P \<turnstile> \<langle>e # es,s\<^sub>0,b\<rangle> [\<rightarrow>]* \<langle>Val v # es,s\<^sub>1,b1\<rangle>" by(rule ListReds1[OF b1])
+ then have iconf2': "iconfs (shp s\<^sub>1) es" using Reds_preserves_iconf[OF wwf cons] Cons by auto
+ have bconf2: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (es,b) \<surd>" using Cons.prems Some by simp
+ then have "P,shp s\<^sub>1 \<turnstile>\<^sub>b (es,b1) \<surd>" using Reds_preserves_bconf[OF wwf cons Cons.prems] by simp
+ then have b2: "P \<turnstile> \<langle>es,s\<^sub>1,b1\<rangle> [\<rightarrow>]* \<langle>es',s\<^sub>2,False\<rangle>" using Cons.hyps(4)[OF iconf2'] by auto
+ show ?thesis using ListRedsVal[OF b1 b2] by auto
+ qed
+next
+ case (ConsThrow e s\<^sub>0 e' s\<^sub>1 es) show ?case
+ proof(cases "val_of e")
+ case None
+ then have iconf: "iconf (shp s\<^sub>0) e" using ConsThrow.prems(1) by simp
+ have bconf: "P,shp s\<^sub>0 \<turnstile>\<^sub>b (e,b) \<surd>" using ConsThrow.prems(2) None by simp
+ then have b1: "P \<turnstile> \<langle>e,s\<^sub>0,b\<rangle> \<rightarrow>* \<langle>throw e',s\<^sub>1,False\<rangle>" using ConsThrow.hyps(2) iconf by auto
+ have cons: "P \<turnstile> \<langle>e # es,s\<^sub>0,b\<rangle> [\<rightarrow>]* \<langle>throw e' # es,s\<^sub>1,False\<rangle>" by(rule ListReds1[OF b1])
+ then show ?thesis by fast
+ next
+ case (Some a)
+ then show ?thesis using eval_final_same[OF ConsThrow.hyps(1)] val_of_spec[OF Some] by auto
+ qed
+next
+ case (InitFinal e s e' s' C b')
+ then have "\<not>sub_RI e" by simp
+ then show ?case using InitFinal RedInit[of e C b' s b P]
+ by (meson converse_rtrancl_into_rtrancl nsub_RI_iconf red_preserves_bconf RedInit)
+next
+ case (InitNone sh C C' Cs e h l e' s')
+ then have init: "P \<turnstile> \<langle>INIT C' (C # Cs,False) \<leftarrow> e,(h, l, sh(C \<mapsto> (sblank P C, Prepared))),b\<rangle> \<rightarrow>* \<langle>e',s',False\<rangle>"
+ by(simp add: bconf_def)
+ show ?case by(rule InitNoneReds[OF InitNone.hyps(1) init])
+next
+ case (InitDone sh C sfs C' Cs e h l e' s')
+ then have "iconf (shp (h, l, sh)) (INIT C' (Cs,True) \<leftarrow> e)" using InitDone.hyps(1)
+ proof(cases Cs)
+ case Nil
+ then have "C = C'" "\<not>sub_RI e" using InitDone.prems(1) by simp+
+ then show ?thesis using Nil InitDone.hyps(1) by(simp add: initPD_def)
+ qed(auto)
+ then have init: "P \<turnstile> \<langle>INIT C' (Cs,True) \<leftarrow> e,(h, l, sh),b\<rangle> \<rightarrow>* \<langle>e',s',False\<rangle>"
+ using InitDone by(simp add: bconf_def)
+ show ?case by(rule InitDoneReds[OF InitDone.hyps(1) init])
+next
+ case (InitProcessing sh C sfs C' Cs e h l e' s')
+ then have "iconf (shp (h, l, sh)) (INIT C' (Cs,True) \<leftarrow> e)" using InitProcessing.hyps(1)
+ proof(cases Cs)
+ case Nil
+ then have "C = C'" "\<not>sub_RI e" using InitProcessing.prems(1) by simp+
+ then show ?thesis using Nil InitProcessing.hyps(1) by(simp add: initPD_def)
+ qed(auto)
+ then have init: "P \<turnstile> \<langle>INIT C' (Cs,True) \<leftarrow> e,(h, l, sh),b\<rangle> \<rightarrow>* \<langle>e',s',False\<rangle>"
+ using InitProcessing by(simp add: bconf_def)
+ show ?case by(rule InitProcessingReds[OF InitProcessing.hyps(1) init])
+next
+ case InitError thus ?case by(fastforce intro: InitErrorReds simp: bconf_def)
+next
+ case InitObject thus ?case by(fastforce intro: InitObjectReds simp: bconf_def)
+next
+ case InitNonObject thus ?case by(fastforce intro: InitNonObjectReds simp: bconf_def)
+next
+ case InitRInit thus ?case by(fastforce intro: RedsInitRInit simp: bconf_def)
+next
+ case (RInit e s v h' l' sh' C sfs i sh'' C' Cs e' e\<^sub>1 s\<^sub>1)
+ then have iconf2: "iconf (shp (h', l', sh'')) (INIT C' (Cs,True) \<leftarrow> e')"
+ by(auto simp: initPD_def fun_upd_same list_nonempty_induct)
+ show ?case
+ proof(cases "val_of e")
+ case None
+ then have iconf: "iconf (shp s) e" using RInit.prems(1) by simp
+ have bconf: "P,shp s \<turnstile>\<^sub>b (e,b) \<surd>" using RInit.prems(2) None by simp
+ then have b1: "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>Val v,(h',l',sh'),False\<rangle>" using RInit.hyps(2)[OF iconf] by auto
+ have "P,shp (h', l', sh'') \<turnstile>\<^sub>b (INIT C' (Cs,True) \<leftarrow> e',False) \<surd>" by(simp add: bconf_def)
+ then have b2: "P \<turnstile> \<langle>INIT C' (Cs,True) \<leftarrow> e',(h',l',sh''),False\<rangle> \<rightarrow>* \<langle>e\<^sub>1,s\<^sub>1,False\<rangle>"
+ using RInit.hyps(7)[OF iconf2] by auto
+ then show ?thesis using RedsRInit[OF b1 RInit.hyps(3-5) b2] by fast
+ next
+ case (Some a')
+ then obtain b1 where b1: "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>Val v,(h',l',sh'),b1\<rangle>"
+ by (metis (no_types, lifting) RInit.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec)
+ have fin: "final e" by(simp add: val_of_spec[OF Some])
+ have "\<not>b" using RInit.prems(2) Some by(simp add: bconf_def)
+ then have nb1: "\<not>b1" using reds_final_same[OF b1 fin] by simp
+ have "P,shp (h', l', sh'') \<turnstile>\<^sub>b (INIT C' (Cs,True) \<leftarrow> e',b1) \<surd>" using nb1
+ by(simp add: bconf_def)
+ then have b2: "P \<turnstile> \<langle>INIT C' (Cs,True) \<leftarrow> e',(h', l', sh''),b1\<rangle> \<rightarrow>* \<langle>e\<^sub>1,s\<^sub>1,False\<rangle>"
+ using RInit.hyps(7)[OF iconf2] by auto
+ then show ?thesis using RedsRInit[OF b1 RInit.hyps(3-5) b2] by fast
+ qed
+next
+ case (RInitInitFail e s a h' l' sh' C sfs i sh'' D Cs e' e\<^sub>1 s\<^sub>1)
+ have fin: "final (throw a)" using eval_final[OF RInitInitFail.hyps(1)] by simp
+ then obtain a' where a': "throw a = Throw a'" by auto
+ have iconf2: "iconf (shp (h', l', sh'')) (RI (D,Throw a') ; Cs \<leftarrow> e')"
+ using RInitInitFail.prems(1) by auto
+ show ?case
+ proof(cases "val_of e")
+ case None
+ then have iconf: "iconf (shp s) e" using RInitInitFail.prems(1) by simp
+ have bconf: "P,shp s \<turnstile>\<^sub>b (e,b) \<surd>" using RInitInitFail.prems(2) None by simp
+ then have b1: "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>Throw a',(h',l',sh'),False\<rangle>"
+ using RInitInitFail.hyps(2)[OF iconf] a' by auto
+ have "P,shp (h', l', sh'') \<turnstile>\<^sub>b (RI (D,Throw a') ; Cs \<leftarrow> e',False) \<surd>" by(simp add: bconf_def)
+ then have b2: "P \<turnstile> \<langle>RI (D,Throw a') ; Cs \<leftarrow> e',(h',l',sh''),False\<rangle> \<rightarrow>* \<langle>e\<^sub>1,s\<^sub>1,False\<rangle>"
+ using RInitInitFail.hyps(6) iconf2 a' by auto
+ show ?thesis using RInitInitThrowReds[OF b1 RInitInitFail.hyps(3-4) b2] by fast
+ next
+ case (Some a1)
+ then obtain b1 where b1: "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>Throw a',(h',l',sh'),b1\<rangle>" using a'
+ by (metis (no_types, lifting) RInitInitFail.hyps(1) eval_cases(3) rtrancl.rtrancl_refl val_of_spec)
+ have fin: "final e" by(simp add: val_of_spec[OF Some])
+ have "\<not>b" using RInitInitFail.prems(2) Some by(simp add: bconf_def)
+ then have nb1: "\<not>b1" using reds_final_same[OF b1 fin] by simp
+ have "P,shp (h', l', sh'') \<turnstile>\<^sub>b (RI (D,Throw a') ; Cs \<leftarrow> e',b1) \<surd>" using nb1
+ by(simp add: bconf_def)
+ then have b2: "P \<turnstile> \<langle>RI (D,Throw a') ; Cs \<leftarrow> e',(h', l', sh''),b1\<rangle> \<rightarrow>* \<langle>e\<^sub>1,s\<^sub>1,False\<rangle>"
+ using RInitInitFail.hyps(6) iconf2 a' by auto
+ show ?thesis using RInitInitThrowReds[OF b1 RInitInitFail.hyps(3-4) b2] by fast
+ qed
+next
+ case (RInitFailFinal e s a h' l' sh' C sfs i sh'' e')
+ have fin: "final (throw a)" using eval_final[OF RInitFailFinal.hyps(1)] by simp
+ then obtain a' where a': "throw a = Throw a'" by auto
+ show ?case
+ proof(cases "val_of e")
+ case None
+ then have iconf: "iconf (shp s) e" using RInitFailFinal.prems(1) by simp
+ have bconf: "P,shp s \<turnstile>\<^sub>b (e,b) \<surd>" using RInitFailFinal.prems(2) None by simp
+ then have b1: "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>Throw a',(h',l',sh'),False\<rangle>"
+ using RInitFailFinal.hyps(2)[OF iconf] a' by auto
+ show ?thesis using RInitThrowReds[OF b1 RInitFailFinal.hyps(3-4)] a' by fast
+ next
+ case (Some a1)
+ then show ?thesis using eval_final_same[OF RInitFailFinal.hyps(1)] val_of_spec[OF Some] by auto
+ qed
+qed
+(*>*)
+
+
+subsection\<open>Big steps simulates small step\<close>
+
+text\<open> This direction was carried out by Norbert Schirmer and Daniel
+Wasserrab (and modified to include statics and DCI by Susannah Mansky). \<close>
+
+text \<open> The big step equivalent of @{text RedWhile}: \<close>
+
+lemma unfold_while:
+ "P \<turnstile> \<langle>while(b) c,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle> = P \<turnstile> \<langle>if(b) (c;;while(b) c) else (unit),s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+(*<*)
+proof
+ assume "P \<turnstile> \<langle>while (b) c,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+ thus "P \<turnstile> \<langle>if (b) (c;; while (b) c) else unit,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+ by cases (fastforce intro: eval_evals.intros)+
+next
+ assume "P \<turnstile> \<langle>if (b) (c;; while (b) c) else unit,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+ thus "P \<turnstile> \<langle>while (b) c,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+ proof (cases)
+ fix a
+ assume e': "e' = throw a"
+ assume "P \<turnstile> \<langle>b,s\<rangle> \<Rightarrow> \<langle>throw a,s'\<rangle>"
+ hence "P \<turnstile> \<langle>while(b) c,s\<rangle> \<Rightarrow> \<langle>throw a,s'\<rangle>" by (rule WhileCondThrow)
+ with e' show ?thesis by simp
+ next
+ fix s\<^sub>1
+ assume eval_false: "P \<turnstile> \<langle>b,s\<rangle> \<Rightarrow> \<langle>false,s\<^sub>1\<rangle>"
+ and eval_unit: "P \<turnstile> \<langle>unit,s\<^sub>1\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+ with eval_unit have "s' = s\<^sub>1" "e' = unit" by (auto elim: eval_cases)
+ moreover from eval_false have "P \<turnstile> \<langle>while (b) c,s\<rangle> \<Rightarrow> \<langle>unit,s\<^sub>1\<rangle>"
+ by - (rule WhileF, simp)
+ ultimately show ?thesis by simp
+ next
+ fix s\<^sub>1
+ assume eval_true: "P \<turnstile> \<langle>b,s\<rangle> \<Rightarrow> \<langle>true,s\<^sub>1\<rangle>"
+ and eval_rest: "P \<turnstile> \<langle>c;; while (b) c,s\<^sub>1\<rangle>\<Rightarrow>\<langle>e',s'\<rangle>"
+ from eval_rest show ?thesis
+ proof (cases)
+ fix s\<^sub>2 v\<^sub>1
+ assume "P \<turnstile> \<langle>c,s\<^sub>1\<rangle> \<Rightarrow> \<langle>Val v\<^sub>1,s\<^sub>2\<rangle>" "P \<turnstile> \<langle>while (b) c,s\<^sub>2\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+ with eval_true show "P \<turnstile> \<langle>while(b) c,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>" by (rule WhileT)
+ next
+ fix a
+ assume "P \<turnstile> \<langle>c,s\<^sub>1\<rangle> \<Rightarrow> \<langle>throw a,s'\<rangle>" "e' = throw a"
+ with eval_true show "P \<turnstile> \<langle>while(b) c,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+ by (iprover intro: WhileBodyThrow)
+ qed
+ qed
+qed
+(*>*)
+
+
+lemma blocksEval:
+ "\<And>Ts vs l l'. \<lbrakk>size ps = size Ts; size ps = size vs; P \<turnstile> \<langle>blocks(ps,Ts,vs,e),(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',(h',l',sh')\<rangle> \<rbrakk>
+ \<Longrightarrow> \<exists> l''. P \<turnstile> \<langle>e,(h,l(ps[\<mapsto>]vs),sh)\<rangle> \<Rightarrow> \<langle>e',(h',l'',sh')\<rangle>"
+(*<*)
+proof (induct ps)
+ case Nil then show ?case by fastforce
+next
+ case (Cons p ps')
+ have length_eqs: "length (p # ps') = length Ts"
+ "length (p # ps') = length vs" by fact+
+ then obtain T Ts' where Ts: "Ts = T#Ts'" by (cases "Ts") simp
+ obtain v vs' where vs: "vs = v#vs'" using length_eqs by (cases "vs") simp
+ have "P \<turnstile> \<langle>blocks (p # ps', Ts, vs, e),(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',(h', l',sh')\<rangle>" by fact
+ with Ts vs
+ have "P \<turnstile> \<langle>{p:T := Val v; blocks (ps', Ts', vs', e)},(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',(h', l',sh')\<rangle>"
+ by simp
+ then obtain l''' where
+ eval_ps': "P \<turnstile> \<langle>blocks (ps', Ts', vs', e),(h, l(p\<mapsto>v), sh)\<rangle> \<Rightarrow> \<langle>e',(h', l''', sh')\<rangle>"
+ and l''': "l'=l'''(p:=l p)"
+ by (auto elim!: eval_cases)
+ then obtain l'' where
+ hyp: "P \<turnstile> \<langle>e,(h, l(p\<mapsto>v)(ps'[\<mapsto>]vs'), sh)\<rangle> \<Rightarrow> \<langle>e',(h', l'', sh')\<rangle>"
+ using length_eqs Ts vs Cons.hyps [OF _ _ eval_ps'] by auto
+ from hyp
+ show "\<exists>l''. P \<turnstile> \<langle>e,(h, l(p # ps'[\<mapsto>]vs), sh)\<rangle> \<Rightarrow> \<langle>e',(h', l'', sh')\<rangle>"
+ using Ts vs by auto
+qed
+(*>*)
+
+lemma
+assumes wf: "wwf_J_prog P"
+shows eval_restrict_lcl:
+ "P \<turnstile> \<langle>e,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',(h',l',sh')\<rangle> \<Longrightarrow> (\<And>W. fv e \<subseteq> W \<Longrightarrow> P \<turnstile> \<langle>e,(h,l|`W,sh)\<rangle> \<Rightarrow> \<langle>e',(h',l'|`W,sh')\<rangle>)"
+and "P \<turnstile> \<langle>es,(h,l,sh)\<rangle> [\<Rightarrow>] \<langle>es',(h',l',sh')\<rangle> \<Longrightarrow> (\<And>W. fvs es \<subseteq> W \<Longrightarrow> P \<turnstile> \<langle>es,(h,l|`W,sh)\<rangle> [\<Rightarrow>] \<langle>es',(h',l'|`W,sh')\<rangle>)"
+(*<*)
+proof(induct rule:eval_evals_inducts)
+ case (Block e\<^sub>0 h\<^sub>0 l\<^sub>0 V sh\<^sub>0 e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 T)
+ have IH: "\<And>W. fv e\<^sub>0 \<subseteq> W \<Longrightarrow> P \<turnstile> \<langle>e\<^sub>0,(h\<^sub>0,l\<^sub>0(V:=None)|`W,sh\<^sub>0)\<rangle> \<Rightarrow> \<langle>e\<^sub>1,(h\<^sub>1,l\<^sub>1|`W,sh\<^sub>1)\<rangle>" by fact
+ have "fv({V:T; e\<^sub>0}) \<subseteq> W" by fact+
+ hence "fv e\<^sub>0 - {V} \<subseteq> W" by simp_all
+ hence "fv e\<^sub>0 \<subseteq> insert V W" by fast
+ from IH[OF this]
+ have "P \<turnstile> \<langle>e\<^sub>0,(h\<^sub>0, (l\<^sub>0|`W)(V := None), sh\<^sub>0)\<rangle> \<Rightarrow> \<langle>e\<^sub>1,(h\<^sub>1, l\<^sub>1|`insert V W, sh\<^sub>1)\<rangle>"
+ by fastforce
+ from eval_evals.Block[OF this] show ?case by fastforce
+next
+ case Seq thus ?case by simp (blast intro:eval_evals.Seq)
+next
+ case New thus ?case by(simp add:eval_evals.intros)
+next
+ case NewFail thus ?case by(simp add:eval_evals.intros)
+next
+ case (NewInit sh C h l v' h' l' sh' a h'')
+ have "fv(INIT C ([C],False) \<leftarrow> unit) \<subseteq> W" by simp
+ then have "P \<turnstile> \<langle>INIT C ([C],False) \<leftarrow> unit,(h, l |` W, sh)\<rangle> \<Rightarrow> \<langle>Val v',(h', l' |` W, sh')\<rangle>"
+ by (simp add: NewInit.hyps(3))
+ thus ?case using NewInit.hyps(1,4-6) eval_evals.NewInit by blast
+next
+ case (NewInitOOM sh C h l v' h' l' sh')
+ have "fv(INIT C ([C],False) \<leftarrow> unit) \<subseteq> W" by simp
+ then have "P \<turnstile> \<langle>INIT C ([C],False) \<leftarrow> unit,(h, l |` W, sh)\<rangle> \<Rightarrow> \<langle>Val v',(h', l' |` W, sh')\<rangle>"
+ by (simp add: NewInitOOM.hyps(3))
+ thus ?case
+ using NewInitOOM.hyps(1,4,5) eval_evals.NewInitOOM by auto
+next
+ case NewInitThrow thus ?case by(simp add:eval_evals.intros)
+next
+ case Cast thus ?case by simp (blast intro:eval_evals.Cast)
+next
+ case CastNull thus ?case by simp (blast intro:eval_evals.CastNull)
+next
+ case CastFail thus ?case by simp (blast intro:eval_evals.CastFail)
+next
+ case CastThrow thus ?case by(simp add:eval_evals.intros)
+next
+ case Val thus ?case by(simp add:eval_evals.intros)
+next
+ case BinOp thus ?case by simp (blast intro:eval_evals.BinOp)
+next
+ case BinOpThrow1 thus ?case by simp (blast intro:eval_evals.BinOpThrow1)
+next
+ case BinOpThrow2 thus ?case by simp (blast intro:eval_evals.BinOpThrow2)
+next
+ case Var thus ?case by(simp add:eval_evals.intros)
+next
+ case (LAss e h\<^sub>0 l\<^sub>0 sh\<^sub>0 v h l sh l' V)
+ have IH: "\<And>W. fv e \<subseteq> W \<Longrightarrow> P \<turnstile> \<langle>e,(h\<^sub>0,l\<^sub>0|`W,sh\<^sub>0)\<rangle> \<Rightarrow> \<langle>Val v,(h,l|`W,sh)\<rangle>"
+ and [simp]: "l' = l(V \<mapsto> v)" by fact+
+ have "fv (V:=e) \<subseteq> W" by fact
+ hence fv: "fv e \<subseteq> W" and VinW: "V \<in> W" by auto
+ from eval_evals.LAss[OF IH[OF fv] refl, of V] VinW
+ show ?case by simp
+next
+ case LAssThrow thus ?case by(fastforce intro: eval_evals.LAssThrow)
+next
+ case FAcc thus ?case by simp (blast intro: eval_evals.FAcc)
+next
+ case FAccNull thus ?case by(fastforce intro: eval_evals.FAccNull)
+next
+ case FAccThrow thus ?case by(fastforce intro: eval_evals.FAccThrow)
+next
+ case FAccNone thus ?case by(metis eval_evals.FAccNone fv.simps(7))
+next
+ case FAccStatic thus ?case by(metis eval_evals.FAccStatic fv.simps(7))
+next
+ case SFAcc thus ?case by simp (blast intro: eval_evals.SFAcc)
+next
+ case SFAccInit thus ?case by simp (blast intro: eval_evals.SFAccInit)
+next
+ case SFAccInitThrow thus ?case by simp (blast intro: eval_evals.SFAccInitThrow)
+next
+ case SFAccNone thus ?case by simp (blast intro: eval_evals.SFAccNone)
+next
+ case SFAccNonStatic thus ?case by simp (blast intro: eval_evals.SFAccNonStatic)
+next
+ case FAss thus ?case by simp (blast intro: eval_evals.FAss)
+next
+ case FAssNull thus ?case by simp (blast intro: eval_evals.FAssNull)
+next
+ case FAssThrow1 thus ?case by simp (blast intro: eval_evals.FAssThrow1)
+next
+ case FAssThrow2 thus ?case by simp (blast intro: eval_evals.FAssThrow2)
+next
+ case FAssNone thus ?case by simp (blast intro: eval_evals.FAssNone)
+next
+ case FAssStatic thus ?case by simp (blast intro: eval_evals.FAssStatic)
+next
+ case SFAss thus ?case by simp (blast intro: eval_evals.SFAss)
+next
+ case SFAssInit thus ?case by simp (blast intro: eval_evals.SFAssInit)
+next
+ case SFAssInitThrow thus ?case by simp (blast intro: eval_evals.SFAssInitThrow)
+next
+ case SFAssThrow thus ?case by simp (blast intro: eval_evals.SFAssThrow)
+next
+ case SFAssNone thus ?case by simp (blast intro: eval_evals.SFAssNone)
+next
+ case SFAssNonStatic thus ?case by simp (blast intro: eval_evals.SFAssNonStatic)
+next
+ case CallObjThrow thus ?case by simp (blast intro: eval_evals.intros)
+next
+ case CallNull thus ?case by simp (blast intro: eval_evals.CallNull)
+next
+ case (CallNone e h l sh a h' l' sh' ps vs h\<^sub>2 l\<^sub>2 sh\<^sub>2 C fs M)
+ have f1: "P \<turnstile> \<langle>e,(h, l |` W, sh)\<rangle> \<Rightarrow> \<langle>addr a,(h', l' |` W, sh')\<rangle>"
+ by (metis (no_types) fv.simps(11) le_sup_iff local.CallNone(2) local.CallNone(7))
+ have "P \<turnstile> \<langle>ps,(h', l' |` W, sh')\<rangle> [\<Rightarrow>] \<langle>map Val vs, (h\<^sub>2, l\<^sub>2 |` W, sh\<^sub>2)\<rangle>"
+ using local.CallNone(4) local.CallNone(7) by auto
+ then show ?case
+ using f1 eval_evals.CallNone local.CallNone(5) local.CallNone(6) by auto
+next
+ case CallStatic thus ?case
+ by (metis (no_types, lifting) eval_evals.CallStatic fv.simps(11) le_sup_iff)
+next
+ case CallParamsThrow thus ?case
+ by simp (blast intro: eval_evals.CallParamsThrow)
+next
+ case (Call e h\<^sub>0 l\<^sub>0 sh\<^sub>0 a h\<^sub>1 l\<^sub>1 sh\<^sub>1 ps vs h\<^sub>2 l\<^sub>2 sh\<^sub>2 C fs M Ts T pns body
+ D l\<^sub>2' e' h\<^sub>3 l\<^sub>3 sh\<^sub>3)
+ have IHe: "\<And>W. fv e \<subseteq> W \<Longrightarrow> P \<turnstile> \<langle>e,(h\<^sub>0,l\<^sub>0|`W,sh\<^sub>0)\<rangle> \<Rightarrow> \<langle>addr a,(h\<^sub>1,l\<^sub>1|`W,sh\<^sub>1)\<rangle>"
+ and IHps: "\<And>W. fvs ps \<subseteq> W \<Longrightarrow> P \<turnstile> \<langle>ps,(h\<^sub>1,l\<^sub>1|`W,sh\<^sub>1)\<rangle> [\<Rightarrow>] \<langle>map Val vs,(h\<^sub>2,l\<^sub>2|`W,sh\<^sub>2)\<rangle>"
+ and IHbd: "\<And>W. fv body \<subseteq> W \<Longrightarrow> P \<turnstile> \<langle>body,(h\<^sub>2,l\<^sub>2'|`W,sh\<^sub>2)\<rangle> \<Rightarrow> \<langle>e',(h\<^sub>3,l\<^sub>3|`W,sh\<^sub>3)\<rangle>"
+ and h\<^sub>2a: "h\<^sub>2 a = Some (C, fs)"
+ and "method": "P \<turnstile> C sees M,NonStatic: Ts\<rightarrow>T = (pns, body) in D"
+ and same_len: "size vs = size pns"
+ and l\<^sub>2': "l\<^sub>2' = [this \<mapsto> Addr a, pns [\<mapsto>] vs]" by fact+
+ have "fv (e\<bullet>M(ps)) \<subseteq> W" by fact
+ hence fve: "fv e \<subseteq> W" and fvps: "fvs(ps) \<subseteq> W" by auto
+ have wfmethod: "size Ts = size pns \<and> this \<notin> set pns" and
+ fvbd: "fv body \<subseteq> {this} \<union> set pns"
+ using "method" wf by(fastforce dest!:sees_wf_mdecl simp:wf_mdecl_def)+
+ show ?case
+ using IHbd[OF fvbd] l\<^sub>2' same_len wfmethod h\<^sub>2a
+ eval_evals.Call[OF IHe[OF fve] IHps[OF fvps] _ "method" same_len l\<^sub>2']
+ by (simp add:subset_insertI)
+next
+ case (SCallNone ps h l sh vs h' l' sh' C M)
+ have "P \<turnstile> \<langle>ps,(h, l |` W, sh)\<rangle> [\<Rightarrow>] \<langle>map Val vs,(h', l' |` W, sh')\<rangle>"
+ using SCallNone.hyps(2) SCallNone.prems by auto
+ then show ?case using SCallNone.hyps(3) eval_evals.SCallNone by auto
+next
+ case SCallNonStatic thus ?case by (metis eval_evals.SCallNonStatic fv.simps(12))
+next
+ case SCallParamsThrow thus ?case
+ by simp (blast intro: eval_evals.SCallParamsThrow)
+next
+ case SCallInitThrow thus ?case by simp (blast intro: eval_evals.SCallInitThrow)
+next
+ case SCallInit thus ?case by simp (blast intro: eval_evals.SCallInit)
+next
+ case (SCall ps h\<^sub>0 l\<^sub>0 sh\<^sub>0 vs h\<^sub>2 l\<^sub>2 sh\<^sub>2 C M Ts T pns body D sfs l\<^sub>2' e' h\<^sub>3 l\<^sub>3 sh\<^sub>3)
+ have IHps: "\<And>W. fvs ps \<subseteq> W \<Longrightarrow> P \<turnstile> \<langle>ps,(h\<^sub>0,l\<^sub>0|`W,sh\<^sub>0)\<rangle> [\<Rightarrow>] \<langle>map Val vs,(h\<^sub>2,l\<^sub>2|`W,sh\<^sub>2)\<rangle>"
+ and IHbd: "\<And>W. fv body \<subseteq> W \<Longrightarrow> P \<turnstile> \<langle>body,(h\<^sub>2,l\<^sub>2'|`W,sh\<^sub>2)\<rangle> \<Rightarrow> \<langle>e',(h\<^sub>3,l\<^sub>3|`W,sh\<^sub>3)\<rangle>"
+ and sh\<^sub>2D: "sh\<^sub>2 D = Some (sfs, Done) \<or> M = clinit \<and> sh\<^sub>2 D = \<lfloor>(sfs, Processing)\<rfloor>"
+ and "method": "P \<turnstile> C sees M,Static: Ts\<rightarrow>T = (pns, body) in D"
+ and same_len: "size vs = size pns"
+ and l\<^sub>2': "l\<^sub>2' = [pns [\<mapsto>] vs]" by fact+
+ have "fv (C\<bullet>\<^sub>sM(ps)) \<subseteq> W" by fact
+ hence fvps: "fvs(ps) \<subseteq> W" by auto
+ have wfmethod: "size Ts = size pns" and fvbd: "fv body \<subseteq> set pns"
+ using "method" wf by(fastforce dest!:sees_wf_mdecl simp:wf_mdecl_def)+
+ show ?case
+ using IHbd[OF fvbd] l\<^sub>2' same_len wfmethod sh\<^sub>2D
+ eval_evals.SCall[OF IHps[OF fvps] "method" _ same_len l\<^sub>2']
+ by (simp add:subset_insertI)
+next
+ case SeqThrow thus ?case by simp (blast intro: eval_evals.SeqThrow)
+next
+ case CondT thus ?case by simp (blast intro: eval_evals.CondT)
+next
+ case CondF thus ?case by simp (blast intro: eval_evals.CondF)
+next
+ case CondThrow thus ?case by simp (blast intro: eval_evals.CondThrow)
+next
+ case WhileF thus ?case by simp (blast intro: eval_evals.WhileF)
+next
+ case WhileT thus ?case by simp (blast intro: eval_evals.WhileT)
+next
+ case WhileCondThrow thus ?case by simp (blast intro: eval_evals.WhileCondThrow)
+next
+ case WhileBodyThrow thus ?case by simp (blast intro: eval_evals.WhileBodyThrow)
+next
+ case Throw thus ?case by simp (blast intro: eval_evals.Throw)
+next
+ case ThrowNull thus ?case by simp (blast intro: eval_evals.ThrowNull)
+next
+ case ThrowThrow thus ?case by simp (blast intro: eval_evals.ThrowThrow)
+next
+ case Try thus ?case by simp (blast intro: eval_evals.Try)
+next
+ case (TryCatch e\<^sub>1 h\<^sub>0 l\<^sub>0 sh\<^sub>0 a h\<^sub>1 l\<^sub>1 sh\<^sub>1 D fs C e\<^sub>2 V e\<^sub>2' h\<^sub>2 l\<^sub>2 sh\<^sub>2)
+ have IH\<^sub>1: "\<And>W. fv e\<^sub>1 \<subseteq> W \<Longrightarrow> P \<turnstile> \<langle>e\<^sub>1,(h\<^sub>0,l\<^sub>0|`W,sh\<^sub>0)\<rangle> \<Rightarrow> \<langle>Throw a,(h\<^sub>1,l\<^sub>1|`W,sh\<^sub>1)\<rangle>"
+ and IH\<^sub>2: "\<And>W. fv e\<^sub>2 \<subseteq> W \<Longrightarrow> P \<turnstile> \<langle>e\<^sub>2,(h\<^sub>1,l\<^sub>1(V\<mapsto>Addr a)|`W,sh\<^sub>1)\<rangle> \<Rightarrow> \<langle>e\<^sub>2',(h\<^sub>2,l\<^sub>2|`W,sh\<^sub>2)\<rangle>"
+ and lookup: "h\<^sub>1 a = Some(D, fs)" and subtype: "P \<turnstile> D \<preceq>\<^sup>* C" by fact+
+ have "fv (try e\<^sub>1 catch(C V) e\<^sub>2) \<subseteq> W" by fact
+ hence fv\<^sub>1: "fv e\<^sub>1 \<subseteq> W" and fv\<^sub>2: "fv e\<^sub>2 \<subseteq> insert V W" by auto
+ have IH\<^sub>2': "P \<turnstile> \<langle>e\<^sub>2,(h\<^sub>1,(l\<^sub>1|`W)(V \<mapsto> Addr a),sh\<^sub>1)\<rangle> \<Rightarrow> \<langle>e\<^sub>2',(h\<^sub>2,l\<^sub>2|`insert V W,sh\<^sub>2)\<rangle>"
+ using IH\<^sub>2[OF fv\<^sub>2] fun_upd_restrict[of l\<^sub>1 W] (*FIXME just l|W instead of l|(W-V) in simp rule??*) by simp
+ with eval_evals.TryCatch[OF IH\<^sub>1[OF fv\<^sub>1] _ subtype IH\<^sub>2'] lookup
+ show ?case by fastforce
+next
+ case TryThrow thus ?case by simp (blast intro: eval_evals.TryThrow)
+next
+ case Nil thus ?case by (simp add: eval_evals.Nil)
+next
+ case Cons thus ?case by simp (blast intro: eval_evals.Cons)
+next
+ case ConsThrow thus ?case by simp (blast intro: eval_evals.ConsThrow)
+next
+ case InitFinal thus ?case by (simp add: eval_evals.InitFinal)
+next
+ case InitNone thus ?case by(blast intro: eval_evals.InitNone)
+next
+ case InitDone thus ?case
+ by (simp add: InitDone.hyps(2) InitDone.prems eval_evals.InitDone)
+next
+ case InitProcessing thus ?case by (simp add: eval_evals.InitProcessing)
+next
+ case InitError thus ?case using eval_evals.InitError by auto
+next
+ case InitObject thus ?case by(simp add: eval_evals.InitObject)
+next
+ case InitNonObject thus ?case by(simp add: eval_evals.InitNonObject)
+next
+ case InitRInit thus ?case by(simp add: eval_evals.InitRInit)
+next
+ case (RInit e h l sh v h' l' sh' C sfs i sh'' C' Cs e' e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1)
+ have f1: "fv e \<subseteq> W \<and> fv (INIT C' (Cs,True) \<leftarrow> e') \<subseteq> W"
+ using RInit.prems by auto
+ then have f2: "P \<turnstile> \<langle>e,(h, l |` W, sh)\<rangle> \<Rightarrow> \<langle>Val v,(h', l' |` W, sh')\<rangle>"
+ using RInit.hyps(2) by blast
+ have "P \<turnstile> \<langle>INIT C' (Cs,True) \<leftarrow> e', (h', l' |` W, sh'')\<rangle> \<Rightarrow> \<langle>e\<^sub>1,(h\<^sub>1, l\<^sub>1 |` W, sh\<^sub>1)\<rangle>"
+ using f1 by (meson RInit.hyps(7))
+ then show ?case
+ using f2 RInit.hyps(3) RInit.hyps(4) RInit.hyps(5) eval_evals.RInit by blast
+next
+ case (RInitInitFail e h l sh a h' l' sh' C sfs i sh'' D Cs e' e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1)
+ have f1: "fv e \<subseteq> W" "fv e' \<subseteq> W"
+ using RInitInitFail.prems by auto
+ then have f2: "P \<turnstile> \<langle>e,(h, l |` W, sh)\<rangle> \<Rightarrow> \<langle>throw a,(h', l' |` W, sh')\<rangle>"
+ using RInitInitFail.hyps(2) by blast
+ then have f2': "fv (throw a) \<subseteq> W"
+ using eval_final[OF f2] by auto
+ then have f1': "fv (RI (D,throw a);Cs \<leftarrow> e') \<subseteq> W"
+ using f1 by auto
+ have "P \<turnstile> \<langle>RI (D,throw a);Cs \<leftarrow> e', (h', l' |` W, sh'')\<rangle> \<Rightarrow> \<langle>e\<^sub>1,(h\<^sub>1, l\<^sub>1 |` W, sh\<^sub>1)\<rangle>"
+ using f1' by (meson RInitInitFail.hyps(6))
+ then show ?case
+ using f2 by (simp add: RInitInitFail.hyps(3,4) eval_evals.RInitInitFail)
+next
+ case (RInitFailFinal e h l sh a h' l' sh' sh'' C)
+ have f1: "fv e \<subseteq> W"
+ using RInitFailFinal.prems by auto
+ then have f2: "P \<turnstile> \<langle>e,(h, l |` W, sh)\<rangle> \<Rightarrow> \<langle>throw a,(h', l' |` W, sh')\<rangle>"
+ using RInitFailFinal.hyps(2) by blast
+ then have f2': "fv (throw a) \<subseteq> W"
+ using eval_final[OF f2] by auto
+ then show ?case using f2 RInitFailFinal.hyps(3,4) eval_evals.RInitFailFinal by blast
+qed
+(*>*)
+
+
+lemma eval_notfree_unchanged:
+ "P \<turnstile> \<langle>e,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',(h',l',sh')\<rangle> \<Longrightarrow> (\<And>V. V \<notin> fv e \<Longrightarrow> l' V = l V)"
+and "P \<turnstile> \<langle>es,(h,l,sh)\<rangle> [\<Rightarrow>] \<langle>es',(h',l',sh')\<rangle> \<Longrightarrow> (\<And>V. V \<notin> fvs es \<Longrightarrow> l' V = l V)"
+(*<*)
+proof(induct rule:eval_evals_inducts)
+ case LAss thus ?case by(simp add:fun_upd_apply)
+next
+ case Block thus ?case
+ by (simp only:fun_upd_apply split:if_splits) fastforce
+next
+ case TryCatch thus ?case
+ by (simp only:fun_upd_apply split:if_splits) fastforce
+next
+ case (RInitInitFail e h l sh a h' l' sh' C sfs i sh'' D Cs e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1)
+ have "fv (throw a) = {}"
+ using RInitInitFail.hyps(1) eval_final final_fv by blast
+ then have "fv e \<union> fv (RI (D,throw a) ; Cs \<leftarrow> unit) \<subseteq> fv (RI (C,e) ; D#Cs \<leftarrow> unit)"
+ by auto
+ then show ?case using RInitInitFail.hyps(2,6) RInitInitFail.prems by fastforce
+qed simp_all
+(*>*)
+
+
+lemma eval_closed_lcl_unchanged:
+ "\<lbrakk> P \<turnstile> \<langle>e,(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',(h',l',sh')\<rangle>; fv e = {} \<rbrakk> \<Longrightarrow> l' = l"
+(*<*)by(fastforce dest:eval_notfree_unchanged simp add:fun_eq_iff [where 'b="val option"])(*>*)
+
+
+lemma list_eval_Throw:
+assumes eval_e: "P \<turnstile> \<langle>throw x,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+shows "P \<turnstile> \<langle>map Val vs @ throw x # es',s\<rangle> [\<Rightarrow>] \<langle>map Val vs @ e' # es',s'\<rangle>"
+(*<*)
+proof -
+ from eval_e
+ obtain a where e': "e' = Throw a"
+ by (cases) (auto dest!: eval_final)
+ {
+ fix es
+ have "\<And>vs. es = map Val vs @ throw x # es'
+ \<Longrightarrow> P \<turnstile> \<langle>es,s\<rangle>[\<Rightarrow>]\<langle>map Val vs @ e' # es',s'\<rangle>"
+ proof (induct es type: list)
+ case Nil thus ?case by simp
+ next
+ case (Cons e es vs)
+ have e_es: "e # es = map Val vs @ throw x # es'" by fact
+ show "P \<turnstile> \<langle>e # es,s\<rangle> [\<Rightarrow>] \<langle>map Val vs @ e' # es',s'\<rangle>"
+ proof (cases vs)
+ case Nil
+ with e_es obtain "e=throw x" "es=es'" by simp
+ moreover from eval_e e'
+ have "P \<turnstile> \<langle>throw x # es,s\<rangle> [\<Rightarrow>] \<langle>Throw a # es,s'\<rangle>"
+ by (iprover intro: ConsThrow)
+ ultimately show ?thesis using Nil e' by simp
+ next
+ case (Cons v vs')
+ have vs: "vs = v # vs'" by fact
+ with e_es obtain
+ e: "e=Val v" and es:"es= map Val vs' @ throw x # es'"
+ by simp
+ from e
+ have "P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>Val v,s\<rangle>"
+ by (iprover intro: eval_evals.Val)
+ moreover from es
+ have "P \<turnstile> \<langle>es,s\<rangle> [\<Rightarrow>] \<langle>map Val vs' @ e' # es',s'\<rangle>"
+ by (rule Cons.hyps)
+ ultimately show
+ "P \<turnstile> \<langle>e#es,s\<rangle> [\<Rightarrow>] \<langle>map Val vs @ e' # es',s'\<rangle>"
+ using vs by (auto intro: eval_evals.Cons)
+ qed
+ qed
+ }
+ thus ?thesis
+ by simp
+qed
+(*>*)
+
+\<comment> \<open> separate evaluation of first subexp of a sequence \<close>
+lemma seq_ext:
+assumes IH: "\<And>e' s'. P \<turnstile> \<langle>e'',s''\<rangle> \<Rightarrow> \<langle>e',s'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+ and seq: "P \<turnstile> \<langle>e'' ;; e\<^sub>0,s''\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+shows "P \<turnstile> \<langle>e ;; e\<^sub>0,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+proof(rule eval_cases(14)[OF seq]) \<comment> \<open> Seq \<close>
+ fix v' s\<^sub>1 assume e'': "P \<turnstile> \<langle>e'',s''\<rangle> \<Rightarrow> \<langle>Val v',s\<^sub>1\<rangle>" and estep: "P \<turnstile> \<langle>e\<^sub>0,s\<^sub>1\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+ have "P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>Val v',s\<^sub>1\<rangle>" using e'' IH by simp
+ then show ?thesis using estep Seq by simp
+next
+ fix e\<^sub>t assume e'': "P \<turnstile> \<langle>e'',s''\<rangle> \<Rightarrow> \<langle>throw e\<^sub>t,s'\<rangle>" and e': "e' = throw e\<^sub>t"
+ have "P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>throw e\<^sub>t,s'\<rangle>" using e'' IH by simp
+ then show ?thesis using eval_evals.SeqThrow e' by simp
+qed
+
+\<comment> \<open> separate evaluation of @{text RI} subexp, val case \<close>
+lemma rinit_Val_ext:
+assumes ri: "P \<turnstile> \<langle>RI (C,e'') ; Cs \<leftarrow> e\<^sub>0,s''\<rangle> \<Rightarrow> \<langle>Val v',s\<^sub>1\<rangle>"
+ and IH: "\<And>e' s'. P \<turnstile> \<langle>e'',s''\<rangle> \<Rightarrow> \<langle>e',s'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+shows "P \<turnstile> \<langle>RI (C,e) ; Cs \<leftarrow> e\<^sub>0,s\<rangle> \<Rightarrow> \<langle>Val v',s\<^sub>1\<rangle>"
+proof(rule eval_cases(20)[OF ri]) \<comment> \<open> RI \<close>
+ fix v'' h' l' sh' sfs i
+ assume e''step: "P \<turnstile> \<langle>e'',s''\<rangle> \<Rightarrow> \<langle>Val v'',(h', l', sh')\<rangle>"
+ and shC: "sh' C = \<lfloor>(sfs, i)\<rfloor>"
+ and init: "P \<turnstile> \<langle>INIT (if Cs = [] then C else last Cs) (Cs,True) \<leftarrow> e\<^sub>0,(h', l', sh'(C \<mapsto> (sfs, Done)))\<rangle> \<Rightarrow>
+ \<langle>Val v',s\<^sub>1\<rangle>"
+ have "P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>Val v'',(h', l', sh')\<rangle>" using IH[OF e''step] by simp
+ then show ?thesis using RInit init shC by auto
+next
+ fix a h' l' sh' sfs i D Cs'
+ assume e''step: "P \<turnstile> \<langle>e'',s''\<rangle> \<Rightarrow> \<langle>throw a,(h', l', sh')\<rangle>"
+ and riD: "P \<turnstile> \<langle>RI (D,throw a) ; Cs' \<leftarrow> e\<^sub>0,(h', l', sh'(C \<mapsto> (sfs, Error)))\<rangle> \<Rightarrow> \<langle>Val v',s\<^sub>1\<rangle>"
+ have "P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>throw a,(h', l', sh')\<rangle>" using IH[OF e''step] by simp
+ then show ?thesis using riD rinit_throwE by blast
+qed(simp)
+
+\<comment> \<open> separate evaluation of @{text RI} subexp, throw case \<close>
+lemma rinit_throw_ext:
+assumes ri: "P \<turnstile> \<langle>RI (C,e'') ; Cs \<leftarrow> e\<^sub>0,s''\<rangle> \<Rightarrow> \<langle>throw e\<^sub>t,s'\<rangle>"
+ and IH: "\<And>e' s'. P \<turnstile> \<langle>e'',s''\<rangle> \<Rightarrow> \<langle>e',s'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+shows "P \<turnstile> \<langle>RI (C,e) ; Cs \<leftarrow> e\<^sub>0,s\<rangle> \<Rightarrow> \<langle>throw e\<^sub>t,s'\<rangle>"
+proof(rule eval_cases(20)[OF ri]) \<comment> \<open> RI \<close>
+ fix v h' l' sh' sfs i
+ assume e''step: "P \<turnstile> \<langle>e'',s''\<rangle> \<Rightarrow> \<langle>Val v,(h', l', sh')\<rangle>"
+ and shC: "sh' C = \<lfloor>(sfs, i)\<rfloor>"
+ and init: "P \<turnstile> \<langle>INIT (if Cs = [] then C else last Cs) (Cs,True) \<leftarrow> e\<^sub>0,(h', l', sh'(C \<mapsto> (sfs, Done)))\<rangle> \<Rightarrow>
+ \<langle>throw e\<^sub>t,s'\<rangle>"
+ have "P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>Val v,(h', l', sh')\<rangle>" using IH[OF e''step] by simp
+ then show ?thesis using RInit init shC by auto
+next
+ fix a h' l' sh' sfs i D Cs'
+ assume e''step: "P \<turnstile> \<langle>e'',s''\<rangle> \<Rightarrow> \<langle>throw a,(h', l', sh')\<rangle>"
+ and shC: "sh' C = \<lfloor>(sfs, i)\<rfloor>"
+ and riD: "P \<turnstile> \<langle>RI (D,throw a) ; Cs' \<leftarrow> e\<^sub>0,(h', l', sh'(C \<mapsto> (sfs, Error)))\<rangle> \<Rightarrow> \<langle>throw e\<^sub>t,s'\<rangle>"
+ and cons: "Cs = D # Cs'"
+ have estep': "P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>throw a,(h', l', sh')\<rangle>" using IH[OF e''step] by simp
+ then show ?thesis using RInitInitFail cons riD shC by simp
+next
+ fix a h' l' sh' sfs i
+ assume "throw e\<^sub>t = throw a"
+ and "s' = (h', l', sh'(C \<mapsto> (sfs, Error)))"
+ and "P \<turnstile> \<langle>e'',s''\<rangle> \<Rightarrow> \<langle>throw a,(h', l', sh')\<rangle>"
+ and "sh' C = \<lfloor>(sfs, i)\<rfloor>"
+ and "Cs = []"
+ then show ?thesis using RInitFailFinal IH by auto
+qed
+
+\<comment> \<open> separate evaluation of @{text RI} subexp \<close>
+lemma rinit_ext:
+assumes IH: "\<And>e' s'. P \<turnstile> \<langle>e'',s''\<rangle> \<Rightarrow> \<langle>e',s'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+shows "\<And>e' s'. P \<turnstile> \<langle>RI (C,e'') ; Cs \<leftarrow> e\<^sub>0,s''\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>
+ \<Longrightarrow> P \<turnstile> \<langle>RI (C,e) ; Cs \<leftarrow> e\<^sub>0,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+proof -
+ fix e' s' assume ri'': "P \<turnstile> \<langle>RI (C,e'') ; Cs \<leftarrow> e\<^sub>0,s''\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+ then have "final e'" using eval_final by simp
+ then show "P \<turnstile> \<langle>RI (C,e) ; Cs \<leftarrow> e\<^sub>0,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+ proof(rule finalE)
+ fix v assume "e' = Val v" then show ?thesis using rinit_Val_ext[OF _ IH] ri'' by simp
+ next
+ fix a assume "e' = throw a" then show ?thesis using rinit_throw_ext[OF _ IH] ri'' by simp
+ qed
+qed
+
+\<comment> \<open> @{text INIT} and @{text RI} return either @{text Val} with @{text Done} or
+ @{text Processing} flag or @{text Throw} with @{text Error} flag \<close>
+lemma
+shows eval_init_return: "P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>
+ \<Longrightarrow> iconf (shp s) e
+ \<Longrightarrow> (\<exists>Cs b. e = INIT C' (Cs,b) \<leftarrow> unit) \<or> (\<exists>C e\<^sub>0 Cs e\<^sub>i. e = RI(C,e\<^sub>0);Cs@[C'] \<leftarrow> unit)
+ \<or> (\<exists>e\<^sub>0. e = RI(C',e\<^sub>0);Nil \<leftarrow> unit)
+ \<Longrightarrow> (val_of e' = Some v \<longrightarrow> (\<exists>sfs i. shp s' C' = \<lfloor>(sfs,i)\<rfloor> \<and> (i = Done \<or> i = Processing)))
+ \<and> (throw_of e' = Some a \<longrightarrow> (\<exists>sfs i. shp s' C' = \<lfloor>(sfs,Error)\<rfloor>))"
+and "P \<turnstile> \<langle>es,s\<rangle> [\<Rightarrow>] \<langle>es',s'\<rangle> \<Longrightarrow> True"
+proof(induct rule: eval_evals.inducts)
+ case (InitFinal e s e' s' C b) then show ?case
+ by(fastforce simp: initPD_def dest: eval_final_same)
+next
+ case (InitDone sh C sfs C' Cs e h l e' s')
+ then have "final e'" using eval_final by simp
+ then show ?case
+ proof(rule finalE)
+ fix v assume e': "e' = Val v" then show ?thesis using InitDone initPD_def
+ proof(cases Cs) qed(auto)
+ next
+ fix a assume e': "e' = throw a" then show ?thesis using InitDone initPD_def
+ proof(cases Cs) qed(auto)
+ qed
+next
+ case (InitProcessing sh C sfs C' Cs e h l e' s')
+ then have "final e'" using eval_final by simp
+ then show ?case
+ proof(rule finalE)
+ fix v assume e': "e' = Val v" then show ?thesis using InitProcessing initPD_def
+ proof(cases Cs) qed(auto)
+ next
+ fix a assume e': "e' = throw a" then show ?thesis using InitProcessing initPD_def
+ proof(cases Cs) qed(auto)
+ qed
+next
+ case (InitError sh C sfs Cs e h l e' s' C') show ?case
+ proof(cases Cs)
+ case Nil then show ?thesis using InitError by simp
+ next
+ case (Cons C2 list)
+ then have "final e'" using InitError eval_final by simp
+ then show ?thesis
+ proof(rule finalE)
+ fix v assume e': "e' = Val v" then show ?thesis
+ using InitError
+ proof -
+ obtain ccss :: "char list list" and bb :: bool where
+ "INIT C' (C # Cs,False) \<leftarrow> e = INIT C' (ccss,bb) \<leftarrow> unit"
+ using InitError.prems(2) by blast
+ then show ?thesis using InitError.hyps(2) e' by(auto dest!: rinit_throwE)
+ qed
+ next
+ fix a assume e': "e' = throw a"
+ then show ?thesis using Cons InitError cons_to_append[of list] by clarsimp
+ qed
+ qed
+next
+ case (InitRInit C Cs h l sh e' s' C') show ?case
+ proof(cases Cs)
+ case Nil then show ?thesis using InitRInit by simp
+ next
+ case (Cons C' list) then show ?thesis
+ using InitRInit Cons cons_to_append[of list] by clarsimp
+ qed
+next
+ case (RInit e s v h' l' sh' C sfs i sh'' C' Cs e' e\<^sub>1 s\<^sub>1)
+ then have final: "final e\<^sub>1" using eval_final by simp
+ then show ?case
+ proof(cases Cs)
+ case Nil show ?thesis using final
+ proof(rule finalE)
+ fix v assume e': "e\<^sub>1 = Val v" show ?thesis
+ using RInit Nil by(auto simp: fun_upd_same initPD_def)
+ next
+ fix a assume e': "e\<^sub>1 = throw a" show ?thesis
+ using RInit Nil by(auto simp: fun_upd_same initPD_def)
+ qed
+ next
+ case (Cons a list) show ?thesis
+ proof(rule finalE[OF final])
+ fix v assume e': "e\<^sub>1 = Val v" then show ?thesis
+ using RInit Cons by clarsimp (metis last.simps last_appendR list.distinct(1))
+ next
+ fix a assume e': "e\<^sub>1 = throw a" then show ?thesis
+ using RInit Cons by clarsimp (metis last.simps last_appendR list.distinct(1))
+ qed
+ qed
+next
+ case (RInitInitFail e s a h' l' sh' C sfs i sh'' D Cs e' e\<^sub>1 s\<^sub>1)
+ then have final: "final e\<^sub>1" using eval_final by simp
+ then show ?case
+ proof(rule finalE)
+ fix v assume e': "e\<^sub>1 = Val v" then show ?thesis
+ using RInitInitFail by clarsimp (meson exp.distinct(101) rinit_throwE)
+ next
+ fix a' assume e': "e\<^sub>1 = Throw a'"
+ then have "iconf (sh'(C \<mapsto> (sfs, Error))) a"
+ using RInitInitFail.hyps(1) eval_final by fastforce
+ then show ?thesis using RInitInitFail e'
+ by clarsimp (meson Cons_eq_append_conv list.inject)
+ qed
+qed(auto simp: fun_upd_same)
+
+lemma init_Val_PD: "P \<turnstile> \<langle>INIT C' (Cs,b) \<leftarrow> unit,s\<rangle> \<Rightarrow> \<langle>Val v,s'\<rangle>
+ \<Longrightarrow> iconf (shp s) (INIT C' (Cs,b) \<leftarrow> unit)
+ \<Longrightarrow> \<exists>sfs i. shp s' C' = \<lfloor>(sfs,i)\<rfloor> \<and> (i = Done \<or> i = Processing)"
+ by(drule_tac v = v in eval_init_return, simp+)
+
+lemma init_throw_PD: "P \<turnstile> \<langle>INIT C' (Cs,b) \<leftarrow> unit,s\<rangle> \<Rightarrow> \<langle>throw a,s'\<rangle>
+ \<Longrightarrow> iconf (shp s) (INIT C' (Cs,b) \<leftarrow> unit)
+ \<Longrightarrow> \<exists>sfs i. shp s' C' = \<lfloor>(sfs,Error)\<rfloor>"
+ by(drule_tac a = a in eval_init_return, simp+)
+
+lemma rinit_Val_PD: "P \<turnstile> \<langle>RI(C,e\<^sub>0);Cs \<leftarrow> unit,s\<rangle> \<Rightarrow> \<langle>Val v,s'\<rangle>
+ \<Longrightarrow> iconf (shp s) (RI(C,e\<^sub>0);Cs \<leftarrow> unit) \<Longrightarrow> last(C#Cs) = C'
+ \<Longrightarrow> \<exists>sfs i. shp s' C' = \<lfloor>(sfs,i)\<rfloor> \<and> (i = Done \<or> i = Processing)"
+by(auto dest!: eval_init_return [where C'=C']
+ append_butlast_last_id[THEN sym])
+
+lemma rinit_throw_PD: "P \<turnstile> \<langle>RI(C,e\<^sub>0);Cs \<leftarrow> unit,s\<rangle> \<Rightarrow> \<langle>throw a,s'\<rangle>
+ \<Longrightarrow> iconf (shp s) (RI(C,e\<^sub>0);Cs \<leftarrow> unit) \<Longrightarrow> last(C#Cs) = C'
+ \<Longrightarrow> \<exists>sfs i. shp s' C' = \<lfloor>(sfs,Error)\<rfloor>"
+by(auto dest!: eval_init_return [where C'=C']
+ append_butlast_last_id[THEN sym])
+
+\<comment> \<open> combining the above to get evaluation of @{text INIT} in a sequence \<close>
+
+(* Hiermit kann man die ganze pair-Splitterei in den automatischen Taktiken
+abschalten. Wieder anschalten siehe nach dem Beweis. *)
+(*<*)
+declare split_paired_All [simp del] split_paired_Ex [simp del]
+(*>*)
+
+lemma eval_init_seq': "P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>
+ \<Longrightarrow> (\<exists>C Cs b e\<^sub>i. e = INIT C (Cs,b) \<leftarrow> e\<^sub>i) \<or> (\<exists>C e\<^sub>0 Cs e\<^sub>i. e = RI(C,e\<^sub>0);Cs \<leftarrow> e\<^sub>i)
+ \<Longrightarrow> (\<exists>C Cs b e\<^sub>i. e = INIT C (Cs,b) \<leftarrow> e\<^sub>i \<and> P \<turnstile> \<langle>(INIT C (Cs,b) \<leftarrow> unit);; e\<^sub>i,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>)
+ \<or> (\<exists>C e\<^sub>0 Cs e\<^sub>i. e = RI(C,e\<^sub>0);Cs \<leftarrow> e\<^sub>i \<and> P \<turnstile> \<langle>(RI(C,e\<^sub>0);Cs \<leftarrow> unit);; e\<^sub>i,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>)"
+and "P \<turnstile> \<langle>es,s\<rangle> [\<Rightarrow>] \<langle>es',s'\<rangle> \<Longrightarrow> True"
+proof(induct rule: eval_evals.inducts)
+ case InitFinal then show ?case by(auto simp: Seq[OF eval_evals.InitFinal[OF Val[where v=Unit]]])
+next
+ case (InitNone sh) then show ?case
+ using seq_ext[OF eval_evals.InitNone[where sh=sh and e=unit, OF InitNone.hyps(1)]] by fastforce
+next
+ case (InitDone sh) then show ?case
+ using seq_ext[OF eval_evals.InitDone[where sh=sh and e=unit, OF InitDone.hyps(1)]] by fastforce
+next
+ case (InitProcessing sh) then show ?case
+ using seq_ext[OF eval_evals.InitProcessing[where sh=sh and e=unit, OF InitProcessing.hyps(1)]]
+ by fastforce
+next
+ case (InitError sh) then show ?case
+ using seq_ext[OF eval_evals.InitError[where sh=sh and e=unit, OF InitError.hyps(1)]] by fastforce
+next
+ case (InitObject sh) then show ?case
+ using seq_ext[OF eval_evals.InitObject[where sh=sh and e=unit, OF InitObject.hyps(1)]]
+ by fastforce
+next
+ case (InitNonObject sh) then show ?case
+ using seq_ext[OF eval_evals.InitNonObject[where sh=sh and e=unit, OF InitNonObject.hyps(1)]]
+ by fastforce
+next
+ case (InitRInit C Cs e h l sh e' s' C') then show ?case
+ using seq_ext[OF eval_evals.InitRInit[where e=unit]] by fastforce
+next
+ case RInit then show ?case
+ using seq_ext[OF eval_evals.RInit[where e=unit, OF RInit.hyps(1)]] by fastforce
+next
+ case RInitInitFail then show ?case
+ using seq_ext[OF eval_evals.RInitInitFail[where e=unit, OF RInitInitFail.hyps(1)]] by fastforce
+next
+ case RInitFailFinal
+ then show ?case using eval_evals.RInitFailFinal eval_evals.SeqThrow by auto
+qed(auto)
+
+lemma eval_init_seq: "P \<turnstile> \<langle>INIT C (Cs,b) \<leftarrow> e,(h, l, sh)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>
+ \<Longrightarrow> P \<turnstile> \<langle>(INIT C (Cs,b) \<leftarrow> unit);; e,(h, l, sh)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+ by(auto dest: eval_init_seq')
+
+
+text \<open> The key lemma: \<close>
+lemma
+assumes wf: "wwf_J_prog P"
+shows extend_1_eval: "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow> \<langle>e'',s'',b''\<rangle> \<Longrightarrow> P,shp s \<turnstile>\<^sub>b (e,b) \<surd>
+ \<Longrightarrow> (\<And>s' e'. P \<turnstile> \<langle>e'',s''\<rangle> \<Rightarrow> \<langle>e',s'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>)"
+and extend_1_evals: "P \<turnstile> \<langle>es,s,b\<rangle> [\<rightarrow>] \<langle>es'',s'',b''\<rangle> \<Longrightarrow> P,shp s \<turnstile>\<^sub>b (es,b) \<surd>
+ \<Longrightarrow> (\<And>s' es'. P \<turnstile> \<langle>es'',s''\<rangle> [\<Rightarrow>] \<langle>es',s'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>es,s\<rangle> [\<Rightarrow>] \<langle>es',s'\<rangle>)"
+proof (induct rule: red_reds.inducts)
+ case (RedNew h a C FDTs h' l sh)
+ then have e':"e' = addr a" and s':"s' = (h(a \<mapsto> blank P C), l, sh)"
+ using eval_cases(3) by fastforce+
+ obtain sfs i where shC: "sh C = \<lfloor>(sfs, i)\<rfloor>" and "i = Done \<or> i = Processing"
+ using RedNew by (clarsimp simp: bconf_def initPD_def)
+ then show ?case
+ proof(cases i)
+ case Done then show ?thesis using RedNew shC e' s' New by simp
+ next
+ case Processing
+ then have shC': "\<nexists>sfs. sh C = Some(sfs,Done)" and shP: "sh C = Some(sfs,Processing)"
+ using shC by simp+
+ then have init: "P \<turnstile> \<langle>INIT C ([C],False) \<leftarrow> unit,(h,l,sh)\<rangle> \<Rightarrow> \<langle>unit,(h,l,sh)\<rangle>"
+ by(simp add: InitFinal InitProcessing Val)
+ have "P \<turnstile> \<langle>new C,(h, l, sh)\<rangle> \<Rightarrow> \<langle>addr a,(h(a \<mapsto> blank P C),l,sh)\<rangle>"
+ using RedNew shC' by(auto intro: NewInit[OF _ init])
+ then show ?thesis using e' s' by simp
+ qed(auto)
+next
+ case (RedNewFail h C l sh)
+ then have e':"e' = THROW OutOfMemory" and s':"s' = (h, l, sh)"
+ using eval_final_same final_def by fastforce+
+ obtain sfs i where shC: "sh C = \<lfloor>(sfs, i)\<rfloor>" and "i = Done \<or> i = Processing"
+ using RedNewFail by (clarsimp simp: bconf_def initPD_def)
+ then show ?case
+ proof(cases i)
+ case Done then show ?thesis using RedNewFail shC e' s' NewFail by simp
+ next
+ case Processing
+ then have shC': "\<nexists>sfs. sh C = Some(sfs,Done)" and shP: "sh C = Some(sfs,Processing)"
+ using shC by simp+
+ then have init: "P \<turnstile> \<langle>INIT C ([C],False) \<leftarrow> unit,(h,l,sh)\<rangle> \<Rightarrow> \<langle>unit,(h,l,sh)\<rangle>"
+ by(simp add: InitFinal InitProcessing Val)
+ have "P \<turnstile> \<langle>new C,(h, l, sh)\<rangle> \<Rightarrow> \<langle>THROW OutOfMemory,(h,l,sh)\<rangle>"
+ using RedNewFail shC' by(auto intro: NewInitOOM[OF _ init])
+ then show ?thesis using e' s' by simp
+ qed(auto)
+next
+ case (NewInitRed sh C h l)
+ then have seq: "P \<turnstile> \<langle>(INIT C ([C],False) \<leftarrow> unit);; new C,(h, l, sh)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+ using eval_init_seq by simp
+ then show ?case
+ proof(rule eval_cases(14)) \<comment> \<open> Seq \<close>
+ fix v s\<^sub>1 assume init: "P \<turnstile> \<langle>INIT C ([C],False) \<leftarrow> unit,(h, l, sh)\<rangle> \<Rightarrow> \<langle>Val v,s\<^sub>1\<rangle>"
+ and new: "P \<turnstile> \<langle>new C,s\<^sub>1\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+ obtain h\<^sub>1 l\<^sub>1 sh\<^sub>1 where s\<^sub>1: "s\<^sub>1 = (h\<^sub>1,l\<^sub>1,sh\<^sub>1)" by(cases s\<^sub>1)
+ then obtain sfs i where shC: "sh\<^sub>1 C = \<lfloor>(sfs, i)\<rfloor>" and iDP: "i = Done \<or> i = Processing"
+ using init_Val_PD[OF init] by auto
+ show ?thesis
+ proof(rule eval_cases(1)[OF new]) \<comment> \<open> New \<close>
+ fix sha ha a FDTs la
+ assume s\<^sub>1a: "s\<^sub>1 = (ha, la, sha)" and e': "e' = addr a"
+ and s': "s' = (ha(a \<mapsto> blank P C), la, sha)"
+ and addr: "new_Addr ha = \<lfloor>a\<rfloor>" and fields: "P \<turnstile> C has_fields FDTs"
+ then show ?thesis using NewInit[OF _ _ addr fields] NewInitRed.hyps init by simp
+ next
+ fix sha ha la
+ assume "s\<^sub>1 = (ha, la, sha)" and "e' = THROW OutOfMemory"
+ and "s' = (ha, la, sha)" and "new_Addr ha = None"
+ then show ?thesis using NewInitOOM NewInitRed.hyps init by simp
+ next
+ fix sha ha la v' h' l' sh' a FDTs
+ assume s\<^sub>1a: "s\<^sub>1 = (ha, la, sha)" and e': "e' = addr a"
+ and s': "s' = (h'(a \<mapsto> blank P C), l', sh')"
+ and shaC: "\<forall>sfs. sha C \<noteq> \<lfloor>(sfs, Done)\<rfloor>"
+ and init': "P \<turnstile> \<langle>INIT C ([C],False) \<leftarrow> unit,(ha, la, sha)\<rangle> \<Rightarrow> \<langle>Val v',(h', l', sh')\<rangle>"
+ and addr: "new_Addr h' = \<lfloor>a\<rfloor>" and fields: "P \<turnstile> C has_fields FDTs"
+ then have i: "i = Processing" using iDP shC s\<^sub>1 by simp
+ then have "(h', l', sh') = (ha, la, sha)" using init' init_ProcessingE s\<^sub>1 s\<^sub>1a shC by blast
+ then show ?thesis using NewInit NewInitRed.hyps s\<^sub>1a addr fields init shaC e' s' by auto
+ next
+ fix sha ha la v' h' l' sh'
+ assume s\<^sub>1a: "s\<^sub>1 = (ha, la, sha)" and e': "e' = THROW OutOfMemory"
+ and s': "s' = (h', l', sh')" and "\<forall>sfs. sha C \<noteq> \<lfloor>(sfs, Done)\<rfloor>"
+ and init': "P \<turnstile> \<langle>INIT C ([C],False) \<leftarrow> unit,(ha, la, sha)\<rangle> \<Rightarrow> \<langle>Val v',(h', l', sh')\<rangle>"
+ and addr: "new_Addr h' = None"
+ then have i: "i = Processing" using iDP shC s\<^sub>1 by simp
+ then have "(h', l', sh') = (ha, la, sha)" using init' init_ProcessingE s\<^sub>1 s\<^sub>1a shC by blast
+ then show ?thesis
+ using NewInitOOM NewInitRed.hyps e' addr s' s\<^sub>1a init by auto
+ next
+ fix sha ha la a
+ assume s\<^sub>1a: "s\<^sub>1 = (ha, la, sha)"
+ and "\<forall>sfs. sha C \<noteq> \<lfloor>(sfs, Done)\<rfloor>"
+ and init': "P \<turnstile> \<langle>INIT C ([C],False) \<leftarrow> unit,(ha, la, sha)\<rangle> \<Rightarrow> \<langle>throw a,s'\<rangle>"
+ then have i: "i = Processing" using iDP shC s\<^sub>1 by simp
+ then show ?thesis using init' init_ProcessingE s\<^sub>1 s\<^sub>1a shC by blast
+ qed
+ next
+ fix e assume e': "e' = throw e"
+ and init: "P \<turnstile> \<langle>INIT C ([C],False) \<leftarrow> unit,(h, l, sh)\<rangle> \<Rightarrow> \<langle>throw e,s'\<rangle>"
+ obtain h' l' sh' where s': "s' = (h',l',sh')" by(cases s')
+ then obtain sfs i where shC: "sh' C = \<lfloor>(sfs, i)\<rfloor>" and iDP: "i = Error"
+ using init_throw_PD[OF init] by auto
+ then show ?thesis by (simp add: NewInitRed.hyps NewInitThrow e' init)
+ qed
+next
+ case CastRed then show ?case
+ by(fastforce elim!: eval_cases intro: eval_evals.intros intro!: CastFail)
+next
+ case RedCastNull
+ then show ?case
+ by simp (iprover elim: eval_cases intro: eval_evals.intros)
+next
+ case RedCast
+ then show ?case
+ by (auto elim: eval_cases intro: eval_evals.intros)
+next
+ case RedCastFail
+ then show ?case
+ by (auto elim!: eval_cases intro: eval_evals.intros)
+next
+ case BinOpRed1 then show ?case
+ by(fastforce elim!: eval_cases intro: eval_evals.intros simp: val_no_step)
+next
+ case BinOpRed2
+ thus ?case
+ by (fastforce elim!: eval_cases intro: eval_evals.intros eval_finalId)
+next
+ case RedBinOp
+ thus ?case
+ by simp (iprover elim: eval_cases intro: eval_evals.intros)
+next
+ case RedVar
+ thus ?case
+ by (fastforce elim: eval_cases intro: eval_evals.intros)
+next
+ case LAssRed thus ?case
+ by(fastforce elim: eval_cases intro: eval_evals.intros)
+next
+ case RedLAss
+ thus ?case
+ by (fastforce elim: eval_cases intro: eval_evals.intros)
+next
+ case FAccRed thus ?case
+ by(fastforce elim: eval_cases intro: eval_evals.intros)
+next
+ case RedFAcc then show ?case
+ by (fastforce elim: eval_cases intro: eval_evals.intros)
+next
+ case RedFAccNull then show ?case
+ by (fastforce elim!: eval_cases intro: eval_evals.intros)
+next
+ case RedFAccNone thus ?case
+ by(fastforce elim: eval_cases intro: eval_evals.intros)
+next
+ case RedFAccStatic thus ?case
+ by(fastforce elim: eval_cases intro: eval_evals.intros)
+next
+ case (RedSFAcc C F t D sh sfs i v h l)
+ then have e':"e' = Val v" and s':"s' = (h, l, sh)"
+ using eval_cases(3) by fastforce+
+ have "i = Done \<or> i = Processing" using RedSFAcc by (clarsimp simp: bconf_def initPD_def)
+ then show ?case
+ proof(cases i)
+ case Done then show ?thesis using RedSFAcc e' s' SFAcc by simp
+ next
+ case Processing
+ then have shC': "\<nexists>sfs. sh D = Some(sfs,Done)" and shP: "sh D = Some(sfs,Processing)"
+ using RedSFAcc by simp+
+ then have init: "P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(h,l,sh)\<rangle> \<Rightarrow> \<langle>unit,(h,l,sh)\<rangle>"
+ by(simp add: InitFinal InitProcessing Val)
+ have "P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D},(h, l, sh)\<rangle> \<Rightarrow> \<langle>Val v,(h,l,sh)\<rangle>"
+ by(rule SFAccInit[OF RedSFAcc.hyps(1) shC' init shP RedSFAcc.hyps(3)])
+ then show ?thesis using e' s' by simp
+ qed(auto)
+next
+ case (SFAccInitRed C F t D sh h l)
+ then have seq: "P \<turnstile> \<langle>(INIT D ([D],False) \<leftarrow> unit);; C\<bullet>\<^sub>sF{D},(h, l, sh)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+ using eval_init_seq by simp
+ then show ?case
+ proof(rule eval_cases(14)) \<comment> \<open> Seq \<close>
+ fix v s\<^sub>1 assume init: "P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(h, l, sh)\<rangle> \<Rightarrow> \<langle>Val v,s\<^sub>1\<rangle>"
+ and acc: "P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D},s\<^sub>1\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+ obtain h\<^sub>1 l\<^sub>1 sh\<^sub>1 where s\<^sub>1: "s\<^sub>1 = (h\<^sub>1,l\<^sub>1,sh\<^sub>1)" by(cases s\<^sub>1)
+ then obtain sfs i where shD: "sh\<^sub>1 D = \<lfloor>(sfs, i)\<rfloor>" and iDP: "i = Done \<or> i = Processing"
+ using init_Val_PD[OF init] by auto
+ show ?thesis
+ proof(rule eval_cases(8)[OF acc]) \<comment> \<open> SFAcc \<close>
+ fix t sha sfs v ha la
+ assume "s\<^sub>1 = (ha, la, sha)" and "e' = Val v"
+ and "s' = (ha, la, sha)" and "P \<turnstile> C has F,Static:t in D"
+ and "sha D = \<lfloor>(sfs, Done)\<rfloor>" and "sfs F = \<lfloor>v\<rfloor>"
+ then show ?thesis using SFAccInit SFAccInitRed.hyps(2) init by auto
+ next
+ fix t sha ha la v' h' l' sh' sfs i' v
+ assume s\<^sub>1a: "s\<^sub>1 = (ha, la, sha)" and e': "e' = Val v"
+ and s': "s' = (h', l', sh')" and field: "P \<turnstile> C has F,Static:t in D"
+ and "\<forall>sfs. sha D \<noteq> \<lfloor>(sfs, Done)\<rfloor>"
+ and init': "P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(ha, la, sha)\<rangle> \<Rightarrow> \<langle>Val v',(h', l', sh')\<rangle>"
+ and shD': "sh' D = \<lfloor>(sfs, i')\<rfloor>" and sfsF: "sfs F = \<lfloor>v\<rfloor>"
+ then have i: "i = Processing" using iDP shD s\<^sub>1 by simp
+ then have "(h', l', sh') = (ha, la, sha)" using init' init_ProcessingE s\<^sub>1 s\<^sub>1a shD by blast
+ then show ?thesis using SFAccInit SFAccInitRed.hyps(2) e' s' field init s\<^sub>1a sfsF shD' by auto
+ next
+ fix t sha ha la a
+ assume s\<^sub>1a: "s\<^sub>1 = (ha, la, sha)" and "e' = throw a"
+ and "P \<turnstile> C has F,Static:t in D" and "\<forall>sfs. sha D \<noteq> \<lfloor>(sfs, Done)\<rfloor>"
+ and init': "P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(ha, la, sha)\<rangle> \<Rightarrow> \<langle>throw a,s'\<rangle>"
+ then have i: "i = Processing" using iDP shD s\<^sub>1 by simp
+ then show ?thesis using init' init_ProcessingE s\<^sub>1 s\<^sub>1a shD by blast
+ next
+ assume "\<forall>b t. \<not> P \<turnstile> C has F,b:t in D"
+ then show ?thesis using SFAccInitRed.hyps(1) by blast
+ next
+ fix t assume field: "P \<turnstile> C has F,NonStatic:t in D"
+ then show ?thesis using has_field_fun[OF SFAccInitRed.hyps(1) field] by simp
+ qed
+ next
+ fix e assume e': "e' = throw e"
+ and init: "P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(h, l, sh)\<rangle> \<Rightarrow> \<langle>throw e,s'\<rangle>"
+ obtain h' l' sh' where s': "s' = (h',l',sh')" by(cases s')
+ then obtain sfs i where shC: "sh' D = \<lfloor>(sfs, i)\<rfloor>" and iDP: "i = Error"
+ using init_throw_PD[OF init] by auto
+ then show ?thesis
+ using SFAccInitRed.hyps(1) SFAccInitRed.hyps(2) SFAccInitThrow e' init by auto
+ qed
+next
+ case RedSFAccNone thus ?case
+ by(fastforce elim: eval_cases intro: eval_evals.intros)
+next
+ case RedSFAccNonStatic thus ?case
+ by(fastforce elim: eval_cases intro: eval_evals.intros)
+next
+ case (FAssRed1 e s b e1 s1 b1 F D e\<^sub>2)
+ obtain h' l' sh' where "s'=(h',l',sh')" by(cases s')
+ with FAssRed1 show ?case
+ by(fastforce elim!: eval_cases(9)[where e\<^sub>1=e1] intro: eval_evals.intros simp: val_no_step
+ intro!: FAss FAssNull FAssNone FAssStatic FAssThrow2)
+next
+ case FAssRed2
+ obtain h' l' sh' where "s'=(h',l',sh')" by(cases s')
+ with FAssRed2 show ?case
+ by(auto elim!: eval_cases intro: eval_evals.intros
+ intro!: FAss FAssNull FAssNone FAssStatic FAssThrow2 Val)
+next
+ case RedFAss
+ thus ?case
+ by (fastforce elim!: eval_cases intro: eval_evals.intros)
+next
+ case RedFAssNull
+ thus ?case
+ by (fastforce elim!: eval_cases intro: eval_evals.intros)
+next
+ case RedFAssNone
+ then show ?case
+ by(auto elim!: eval_cases intro: eval_evals.intros eval_finalId)
+next
+ case RedFAssStatic
+ then show ?case
+ by(auto elim!: eval_cases intro: eval_evals.intros eval_finalId)
+next
+ case (SFAssRed e s b e'' s'' b'' C F D)
+ obtain h l sh where [simp]: "s = (h,l,sh)" by(cases s)
+ obtain h' l' sh' where [simp]: "s'=(h',l',sh')" by(cases s')
+ have "val_of e = None" using val_no_step SFAssRed.hyps(1) by(meson option.exhaust)
+ then have bconf: "P,sh \<turnstile>\<^sub>b (e,b) \<surd>" using SFAssRed by simp
+ show ?case using SFAssRed.prems(2) SFAssRed bconf
+ proof cases
+ case 2 with SFAssRed bconf show ?thesis by(auto intro!: SFAssInit)
+ next
+ case 3 with SFAssRed bconf show ?thesis by(auto intro!: SFAssInitThrow)
+ qed(auto intro: eval_evals.intros intro!: SFAss SFAssInit SFAssNone SFAssNonStatic)
+next
+ case (RedSFAss C F t D sh sfs i sfs' v sh' h l)
+ let ?sfs' = "sfs(F \<mapsto> v)"
+ have e':"e' = unit" and s':"s' = (h, l, sh(D \<mapsto> (?sfs', i)))"
+ using RedSFAss eval_cases(3) by fastforce+
+ have "i = Done \<or> i = Processing" using RedSFAss by(clarsimp simp: bconf_def initPD_def)
+ then show ?case
+ proof(cases i)
+ case Done then show ?thesis using RedSFAss e' s' SFAss Val by auto
+ next
+ case Processing
+ then have shC': "\<nexists>sfs. sh D = Some(sfs,Done)" and shP: "sh D = Some(sfs,Processing)"
+ using RedSFAss by simp+
+ then have init: "P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(h,l,sh)\<rangle> \<Rightarrow> \<langle>unit,(h,l,sh)\<rangle>"
+ by(simp add: InitFinal InitProcessing Val)
+ have "P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D} := Val v,(h, l, sh)\<rangle> \<Rightarrow> \<langle>unit,(h,l,sh(D \<mapsto> (?sfs', i)))\<rangle>"
+ using Processing by(auto intro: SFAssInit[OF Val RedSFAss.hyps(1) shC' init shP])
+ then show ?thesis using e' s' by simp
+ qed(auto)
+next
+ case (SFAssInitRed C F t D sh v h l)
+ then have seq: "P \<turnstile> \<langle>(INIT D ([D],False) \<leftarrow> unit);; C\<bullet>\<^sub>sF{D} := Val v,(h, l, sh)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+ using eval_init_seq by simp
+ then show ?case
+ proof(rule eval_cases(14)) \<comment> \<open> Seq \<close>
+ fix v' s\<^sub>1 assume init: "P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(h, l, sh)\<rangle> \<Rightarrow> \<langle>Val v',s\<^sub>1\<rangle>"
+ and acc: "P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D} := Val v,s\<^sub>1\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+ obtain h\<^sub>1 l\<^sub>1 sh\<^sub>1 where s\<^sub>1: "s\<^sub>1 = (h\<^sub>1,l\<^sub>1,sh\<^sub>1)" by(cases s\<^sub>1)
+ then obtain sfs i where shD: "sh\<^sub>1 D = \<lfloor>(sfs, i)\<rfloor>" and iDP: "i = Done \<or> i = Processing"
+ using init_Val_PD[OF init] by auto
+ show ?thesis
+ proof(rule eval_cases(10)[OF acc]) \<comment> \<open> SFAss \<close>
+ fix va h\<^sub>1 l\<^sub>1 sh\<^sub>1 t sfs
+ assume e': "e' = unit" and s': "s' = (h\<^sub>1, l\<^sub>1, sh\<^sub>1(D \<mapsto> (sfs(F \<mapsto> va), Done)))"
+ and val: "P \<turnstile> \<langle>Val v,s\<^sub>1\<rangle> \<Rightarrow> \<langle>Val va,(h\<^sub>1, l\<^sub>1, sh\<^sub>1)\<rangle>"
+ and field: "P \<turnstile> C has F,Static:t in D" and shD': "sh\<^sub>1 D = \<lfloor>(sfs, Done)\<rfloor>"
+ have "v = va" and "s\<^sub>1 = (h\<^sub>1, l\<^sub>1, sh\<^sub>1)" using eval_final_same[OF val] by auto
+ then show ?thesis using SFAssInit field SFAssInitRed.hyps(2) shD' e' s' init val
+ by (metis eval_final eval_finalId)
+ next
+ fix va h\<^sub>1 l\<^sub>1 sh\<^sub>1 t v' h' l' sh' sfs i'
+ assume e': "e' = unit" and s': "s' = (h', l', sh'(D \<mapsto> (sfs(F \<mapsto> va), i')))"
+ and val: "P \<turnstile> \<langle>Val v,s\<^sub>1\<rangle> \<Rightarrow> \<langle>Val va,(h\<^sub>1, l\<^sub>1, sh\<^sub>1)\<rangle>"
+ and field: "P \<turnstile> C has F,Static:t in D" and nDone: "\<forall>sfs. sh\<^sub>1 D \<noteq> \<lfloor>(sfs, Done)\<rfloor>"
+ and init': "P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(h\<^sub>1, l\<^sub>1, sh\<^sub>1)\<rangle> \<Rightarrow> \<langle>Val v',(h', l', sh')\<rangle>"
+ and shD': "sh' D = \<lfloor>(sfs, i')\<rfloor>"
+ have v: "v = va" and s\<^sub>1a: "s\<^sub>1 = (h\<^sub>1, l\<^sub>1, sh\<^sub>1)" using eval_final_same[OF val] by auto
+ then have i: "i = Processing" using iDP shD s\<^sub>1 nDone by simp
+ then have "(h\<^sub>1, l\<^sub>1, sh\<^sub>1) = (h', l', sh')" using init' init_ProcessingE s\<^sub>1 s\<^sub>1a shD by blast
+ then show ?thesis using SFAssInit SFAssInitRed.hyps(2) e' s' field init v s\<^sub>1a shD' val
+ by (metis eval_final eval_finalId)
+ next
+ fix va h\<^sub>1 l\<^sub>1 sh\<^sub>1 t a
+ assume "e' = throw a" and val: "P \<turnstile> \<langle>Val v,s\<^sub>1\<rangle> \<Rightarrow> \<langle>Val va,(h\<^sub>1, l\<^sub>1, sh\<^sub>1)\<rangle>"
+ and "P \<turnstile> C has F,Static:t in D" and nDone: "\<forall>sfs. sh\<^sub>1 D \<noteq> \<lfloor>(sfs, Done)\<rfloor>"
+ and init': "P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(h\<^sub>1, l\<^sub>1, sh\<^sub>1)\<rangle> \<Rightarrow> \<langle>throw a,s'\<rangle>"
+ have v: "v = va" and s\<^sub>1a: "s\<^sub>1 = (h\<^sub>1, l\<^sub>1, sh\<^sub>1)" using eval_final_same[OF val] by auto
+ then have i: "i = Processing" using iDP shD s\<^sub>1 nDone by simp
+ then have "(h\<^sub>1, l\<^sub>1, sh\<^sub>1) = s'" using init' init_ProcessingE s\<^sub>1 s\<^sub>1a shD by blast
+ then show ?thesis using init' init_ProcessingE s\<^sub>1 s\<^sub>1a shD i by blast
+ next
+ fix e'' assume val:"P \<turnstile> \<langle>Val v,s\<^sub>1\<rangle> \<Rightarrow> \<langle>throw e'',s'\<rangle>"
+ then show ?thesis using eval_final_same[OF val] by simp
+ next
+ assume "\<forall>b t. \<not> P \<turnstile> C has F,b:t in D"
+ then show ?thesis using SFAssInitRed.hyps(1) by blast
+ next
+ fix t assume field: "P \<turnstile> C has F,NonStatic:t in D"
+ then show ?thesis using has_field_fun[OF SFAssInitRed.hyps(1) field] by simp
+ qed
+ next
+ fix e assume e': "e' = throw e"
+ and init: "P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(h, l, sh)\<rangle> \<Rightarrow> \<langle>throw e,s'\<rangle>"
+ obtain h' l' sh' where s': "s' = (h',l',sh')" by(cases s')
+ then obtain sfs i where shC: "sh' D = \<lfloor>(sfs, i)\<rfloor>" and iDP: "i = Error"
+ using init_throw_PD[OF init] by auto
+ then show ?thesis using SFAssInitRed.hyps(1) SFAssInitRed.hyps(2) SFAssInitThrow Val
+ by (metis e' init)
+ qed
+next
+ case (RedSFAssNone C F D v s b) then show ?case
+ by(cases s) (auto elim!: eval_cases intro: eval_evals.intros eval_finalId)
+next
+ case (RedSFAssNonStatic C F t D v s b) then show ?case
+ by(cases s) (auto elim!: eval_cases intro: eval_evals.intros eval_finalId)
+next
+ case CallObj
+ note val_no_step[simp]
+ from CallObj.prems(2) CallObj show ?case
+ proof cases
+ case 2 with CallObj show ?thesis by(fastforce intro!: CallParamsThrow)
+ next
+ case 3 with CallObj show ?thesis by(fastforce intro!: CallNull)
+ next
+ case 4 with CallObj show ?thesis by(fastforce intro!: CallNone)
+ next
+ case 5 with CallObj show ?thesis by(fastforce intro!: CallStatic)
+ qed(fastforce intro!: CallObjThrow Call)+
+next
+ case (CallParams es s b es'' s'' b'' v M s')
+ then obtain h' l' sh' where "s' = (h',l',sh')" by(cases s')
+ with CallParams show ?case
+ by(auto elim!: eval_cases intro!: CallNone eval_finalId CallStatic Val)
+ (auto intro!: CallParamsThrow CallNull Call Val)
+next
+ case (RedCall h a C fs M Ts T pns body D vs l sh b)
+ have "P \<turnstile> \<langle>addr a,(h,l,sh)\<rangle> \<Rightarrow> \<langle>addr a,(h,l,sh)\<rangle>" by (rule eval_evals.intros)
+ moreover
+ have finals: "finals(map Val vs)" by simp
+ with finals have "P \<turnstile> \<langle>map Val vs,(h,l,sh)\<rangle> [\<Rightarrow>] \<langle>map Val vs,(h,l,sh)\<rangle>"
+ by (iprover intro: eval_finalsId)
+ moreover have "h a = Some (C, fs)" using RedCall by simp
+ moreover have "method": "P \<turnstile> C sees M,NonStatic: Ts\<rightarrow>T = (pns, body) in D" by fact
+ moreover have same_len\<^sub>1: "length Ts = length pns"
+ and this_distinct: "this \<notin> set pns" and fv: "fv (body) \<subseteq> {this} \<union> set pns"
+ using "method" wf by (fastforce dest!:sees_wf_mdecl simp:wf_mdecl_def)+
+ have same_len: "length vs = length pns" by fact
+ moreover
+ obtain l\<^sub>2' where l\<^sub>2': "l\<^sub>2' = [this\<mapsto>Addr a,pns[\<mapsto>]vs]" by simp
+ moreover
+ obtain h\<^sub>3 l\<^sub>3 sh\<^sub>3 where s': "s' = (h\<^sub>3,l\<^sub>3,sh\<^sub>3)" by (cases s')
+ have eval_blocks:
+ "P \<turnstile> \<langle>(blocks (this # pns, Class D # Ts, Addr a # vs, body)),(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>" by fact
+ hence id: "l\<^sub>3 = l" using fv s' same_len\<^sub>1 same_len
+ by(fastforce elim: eval_closed_lcl_unchanged)
+ from eval_blocks obtain l\<^sub>3' where "P \<turnstile> \<langle>body,(h,l\<^sub>2',sh)\<rangle> \<Rightarrow> \<langle>e',(h\<^sub>3,l\<^sub>3',sh\<^sub>3)\<rangle>"
+ proof -
+ from same_len\<^sub>1 have "length(this#pns) = length(Class D#Ts)" by simp
+ moreover from same_len\<^sub>1 same_len
+ have same_len\<^sub>2: "length (this#pns) = length (Addr a#vs)" by simp
+ moreover from eval_blocks
+ have "P \<turnstile> \<langle>blocks (this#pns,Class D#Ts,Addr a#vs,body),(h,l,sh)\<rangle>
+ \<Rightarrow>\<langle>e',(h\<^sub>3,l\<^sub>3,sh\<^sub>3)\<rangle>" using s' same_len\<^sub>1 same_len\<^sub>2 by simp
+ ultimately obtain l''
+ where "P \<turnstile> \<langle>body,(h,l(this # pns[\<mapsto>]Addr a # vs),sh)\<rangle>\<Rightarrow>\<langle>e',(h\<^sub>3, l'',sh\<^sub>3)\<rangle>"
+ by (blast dest:blocksEval)
+ from eval_restrict_lcl[OF wf this fv] this_distinct same_len\<^sub>1 same_len
+ have "P \<turnstile> \<langle>body,(h,[this # pns[\<mapsto>]Addr a # vs],sh)\<rangle> \<Rightarrow>
+ \<langle>e',(h\<^sub>3, l''|`(set(this#pns)),sh\<^sub>3)\<rangle>" using wf method
+ by(simp add:subset_insert_iff insert_Diff_if)
+ thus ?thesis by(fastforce intro!:that simp add: l\<^sub>2')
+ qed
+ ultimately
+ have "P \<turnstile> \<langle>(addr a)\<bullet>M(map Val vs),(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',(h\<^sub>3,l,sh\<^sub>3)\<rangle>" by (rule Call)
+ with s' id show ?case by simp
+next
+ case RedCallNull
+ thus ?case
+ by (fastforce elim: eval_cases intro: eval_evals.intros eval_finalsId)
+next
+ case (RedCallNone h a C fs M vs l sh b)
+ then have tes: "THROW NoSuchMethodError = e' \<and> (h,l,sh) = s'"
+ using eval_final_same by simp
+ have "P \<turnstile> \<langle>addr a,(h,l,sh)\<rangle> \<Rightarrow> \<langle>addr a,(h,l,sh)\<rangle>" and "P \<turnstile> \<langle>map Val vs,(h,l,sh)\<rangle> [\<Rightarrow>] \<langle>map Val vs,(h,l,sh)\<rangle>"
+ using eval_finalId eval_finalsId by auto
+ then show ?case using RedCallNone CallNone tes by auto
+next
+ case (RedCallStatic h a C fs M Ts T m D vs l sh b)
+ then have tes: "THROW IncompatibleClassChangeError = e' \<and> (h,l,sh) = s'"
+ using eval_final_same by simp
+ have "P \<turnstile> \<langle>addr a,(h,l,sh)\<rangle> \<Rightarrow> \<langle>addr a,(h,l,sh)\<rangle>" and "P \<turnstile> \<langle>map Val vs,(h,l,sh)\<rangle> [\<Rightarrow>] \<langle>map Val vs,(h,l,sh)\<rangle>"
+ using eval_finalId eval_finalsId by auto
+ then show ?case using RedCallStatic CallStatic tes by fastforce
+next
+ case (SCallParams es s b es'' s'' b' C M s')
+ obtain h' l' sh' where s'[simp]: "s' = (h',l',sh')" by(cases s')
+ obtain h l sh where s[simp]: "s = (h,l,sh)" by(cases s)
+ have es: "map_vals_of es = None" using vals_no_step SCallParams.hyps(1) by (meson not_Some_eq)
+ have bconf: "P,sh \<turnstile>\<^sub>b (es,b) \<surd>" using s SCallParams.prems(1) by (simp add: bconf_SCall[OF es])
+ from SCallParams.prems(2) SCallParams bconf show ?case
+ proof cases
+ case 2 with SCallParams bconf show ?thesis by(auto intro!: SCallNone)
+ next
+ case 3 with SCallParams bconf show ?thesis by(auto intro!: SCallNonStatic)
+ next
+ case 4 with SCallParams bconf show ?thesis by(auto intro!: SCallInitThrow)
+ next
+ case 5 with SCallParams bconf show ?thesis by(auto intro!: SCallInit)
+ qed(auto intro!: SCallParamsThrow SCall)
+next
+ case (RedSCall C M Ts T pns body D vs s)
+ then obtain h l sh where s:"s = (h,l,sh)" by(cases s)
+ then obtain sfs i where shC: "sh D = \<lfloor>(sfs, i)\<rfloor>" and "i = Done \<or> i = Processing"
+ using RedSCall by(auto simp: bconf_def initPD_def dest: sees_method_fun)
+ have finals: "finals(map Val vs)" by simp
+ with finals have mVs: "P \<turnstile> \<langle>map Val vs,(h,l,sh)\<rangle> [\<Rightarrow>] \<langle>map Val vs,(h,l,sh)\<rangle>"
+ by (iprover intro: eval_finalsId)
+ obtain sfs i where shC: "sh D = \<lfloor>(sfs, i)\<rfloor>"
+ using RedSCall s by(auto simp: bconf_def initPD_def dest: sees_method_fun)
+ then have iDP: "i = Done \<or> i = Processing" using RedSCall s
+ by (auto simp: bconf_def initPD_def dest: sees_method_fun[OF RedSCall.hyps(1)])
+ have "method": "P \<turnstile> C sees M,Static: Ts\<rightarrow>T = (pns, body) in D" by fact
+ have same_len\<^sub>1: "length Ts = length pns" and fv: "fv (body) \<subseteq> set pns"
+ using "method" wf by (fastforce dest!:sees_wf_mdecl simp:wf_mdecl_def)+
+ have same_len: "length vs = length pns" by fact
+ obtain l\<^sub>2' where l\<^sub>2': "l\<^sub>2' = [pns[\<mapsto>]vs]" by simp
+ obtain h\<^sub>3 l\<^sub>3 sh\<^sub>3 where s': "s' = (h\<^sub>3,l\<^sub>3,sh\<^sub>3)" by (cases s')
+ have eval_blocks:
+ "P \<turnstile> \<langle>(blocks (pns, Ts, vs, body)),(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>" using RedSCall.prems(2) s by simp
+ hence id: "l\<^sub>3 = l" using fv s' same_len\<^sub>1 same_len
+ by(fastforce elim: eval_closed_lcl_unchanged)
+ from eval_blocks obtain l\<^sub>3' where body: "P \<turnstile> \<langle>body,(h,l\<^sub>2',sh)\<rangle> \<Rightarrow> \<langle>e',(h\<^sub>3,l\<^sub>3',sh\<^sub>3)\<rangle>"
+ proof -
+ from eval_blocks
+ have "P \<turnstile> \<langle>blocks (pns,Ts,vs,body),(h,l,sh)\<rangle>
+ \<Rightarrow>\<langle>e',(h\<^sub>3,l\<^sub>3,sh\<^sub>3)\<rangle>" using s' same_len same_len\<^sub>1 by simp
+ then obtain l''
+ where "P \<turnstile> \<langle>body,(h,l(pns[\<mapsto>]vs),sh)\<rangle> \<Rightarrow> \<langle>e',(h\<^sub>3, l'',sh\<^sub>3)\<rangle>"
+ by (blast dest:blocksEval[OF same_len\<^sub>1[THEN sym] same_len[THEN sym]])
+ from eval_restrict_lcl[OF wf this fv] same_len\<^sub>1 same_len
+ have "P \<turnstile> \<langle>body,(h,[pns[\<mapsto>]vs],sh)\<rangle> \<Rightarrow> \<langle>e',(h\<^sub>3, l''|`(set(pns)),sh\<^sub>3)\<rangle>" using wf method
+ by(simp add:subset_insert_iff insert_Diff_if)
+ thus ?thesis by(fastforce intro!:that simp add: l\<^sub>2')
+ qed
+ show ?case using iDP
+ proof(cases i)
+ case Done
+ then have shC': "sh D = \<lfloor>(sfs, Done)\<rfloor> \<or> M = clinit \<and> sh D = \<lfloor>(sfs, Processing)\<rfloor>"
+ using shC by simp
+ have "P \<turnstile> \<langle>C\<bullet>\<^sub>sM(map Val vs),(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',(h\<^sub>3,l,sh\<^sub>3)\<rangle>"
+ by (rule SCall[OF mVs method shC' same_len l\<^sub>2' body])
+ with s s' id show ?thesis by simp
+ next
+ case Processing
+ then have shC': "\<nexists>sfs. sh D = Some(sfs,Done)" and shP: "sh D = Some(sfs,Processing)"
+ using shC by simp+
+ then have init: "P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(h,l,sh)\<rangle> \<Rightarrow> \<langle>unit,(h,l,sh)\<rangle>"
+ by(simp add: InitFinal InitProcessing Val)
+ have "P \<turnstile> \<langle>C\<bullet>\<^sub>sM(map Val vs),(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',(h\<^sub>3,l,sh\<^sub>3)\<rangle>"
+ proof(cases "M = clinit")
+ case False show ?thesis by(rule SCallInit[OF mVs method shC' False init same_len l\<^sub>2' body])
+ next
+ case True
+ then have shC': "sh D = \<lfloor>(sfs, Done)\<rfloor> \<or> M = clinit \<and> sh D = \<lfloor>(sfs, Processing)\<rfloor>"
+ using shC Processing by simp
+ have "P \<turnstile> \<langle>C\<bullet>\<^sub>sM(map Val vs),(h,l,sh)\<rangle> \<Rightarrow> \<langle>e',(h\<^sub>3,l,sh\<^sub>3)\<rangle>"
+ by (rule SCall[OF mVs method shC' same_len l\<^sub>2' body])
+ with s s' id show ?thesis by simp
+ qed
+ with s s' id show ?thesis by simp
+ qed(auto)
+next
+ case (SCallInitRed C M Ts T pns body D sh vs h l)
+ then have seq: "P \<turnstile> \<langle>(INIT D ([D],False) \<leftarrow> unit);; C\<bullet>\<^sub>sM(map Val vs),(h, l, sh)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+ using eval_init_seq by simp
+ then show ?case
+ proof(rule eval_cases(14)) \<comment> \<open> Seq \<close>
+ fix v' s\<^sub>1 assume init: "P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(h, l, sh)\<rangle> \<Rightarrow> \<langle>Val v',s\<^sub>1\<rangle>"
+ and call: "P \<turnstile> \<langle>C\<bullet>\<^sub>sM(map Val vs),s\<^sub>1\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+ obtain h\<^sub>1 l\<^sub>1 sh\<^sub>1 where s\<^sub>1: "s\<^sub>1 = (h\<^sub>1,l\<^sub>1,sh\<^sub>1)" by(cases s\<^sub>1)
+ then obtain sfs i where shD: "sh\<^sub>1 D = \<lfloor>(sfs, i)\<rfloor>" and iDP: "i = Done \<or> i = Processing"
+ using init_Val_PD[OF init] by auto
+ show ?thesis
+ proof(rule eval_cases(12)[OF call]) \<comment> \<open> SCall \<close>
+ fix vsa ex es' assume "P \<turnstile> \<langle>map Val vs,s\<^sub>1\<rangle> [\<Rightarrow>] \<langle>map Val vsa @ throw ex # es',s'\<rangle>"
+ then show ?thesis using evals_finals_same by (meson finals_def map_Val_nthrow_neq)
+ next
+ assume "\<forall>b Ts T a ba x. \<not> P \<turnstile> C sees M, b : Ts\<rightarrow>T = (a, ba) in x"
+ then show ?thesis using SCallInitRed.hyps(1) by auto
+ next
+ fix Ts T m D assume "P \<turnstile> C sees M, NonStatic : Ts\<rightarrow>T = m in D"
+ then show ?thesis using sees_method_fun[OF SCallInitRed.hyps(1)] by blast
+ next
+ fix vsa h1 l1 sh1 Ts T pns body D' a
+ assume "e' = throw a" and vals: "P \<turnstile> \<langle>map Val vs,s\<^sub>1\<rangle> [\<Rightarrow>] \<langle>map Val vsa,(h1, l1, sh1)\<rangle>"
+ and method: "P \<turnstile> C sees M, Static : Ts\<rightarrow>T = (pns, body) in D'"
+ and nDone: "\<forall>sfs. sh1 D' \<noteq> \<lfloor>(sfs, Done)\<rfloor>"
+ and init': "P \<turnstile> \<langle>INIT D' ([D'],False) \<leftarrow> unit,(h1, l1, sh1)\<rangle> \<Rightarrow> \<langle>throw a,s'\<rangle>"
+ have vs: "vs = vsa" and s\<^sub>1a: "s\<^sub>1 = (h1, l1, sh1)"
+ using evals_finals_same[OF _ vals] map_Val_eq by auto
+ have D: "D = D'" using sees_method_fun[OF SCallInitRed.hyps(1) method] by simp
+ then have i: "i = Processing" using iDP shD s\<^sub>1 s\<^sub>1a nDone by simp
+ then show ?thesis using D init' init_ProcessingE s\<^sub>1 s\<^sub>1a shD by blast
+ next
+ fix vsa h1 l1 sh1 Ts T pns' body' D' v' h\<^sub>2 l\<^sub>2 sh\<^sub>2 h\<^sub>3 l\<^sub>3 sh\<^sub>3
+ assume s': "s' = (h\<^sub>3, l\<^sub>2, sh\<^sub>3)"
+ and vals: "P \<turnstile> \<langle>map Val vs,s\<^sub>1\<rangle> [\<Rightarrow>] \<langle>map Val vsa,(h1, l1, sh1)\<rangle>"
+ and method: "P \<turnstile> C sees M, Static : Ts\<rightarrow>T = (pns', body') in D'"
+ and nDone: "\<forall>sfs. sh1 D' \<noteq> \<lfloor>(sfs, Done)\<rfloor>"
+ and init': "P \<turnstile> \<langle>INIT D' ([D'],False) \<leftarrow> unit,(h1, l1, sh1)\<rangle> \<Rightarrow> \<langle>Val v',(h\<^sub>2, l\<^sub>2, sh\<^sub>2)\<rangle>"
+ and len: "length vsa = length pns'"
+ and bstep: "P \<turnstile> \<langle>body',(h\<^sub>2, [pns' [\<mapsto>] vsa], sh\<^sub>2)\<rangle> \<Rightarrow> \<langle>e',(h\<^sub>3, l\<^sub>3, sh\<^sub>3)\<rangle>"
+ have vs: "vs = vsa" and s\<^sub>1a: "s\<^sub>1 = (h1, l1, sh1)"
+ using evals_finals_same[OF _ vals] map_Val_eq by auto
+ have D: "D = D'" and pns: "pns = pns'" and body: "body = body'"
+ using sees_method_fun[OF SCallInitRed.hyps(1) method] by auto
+ then have i: "i = Processing" using iDP shD s\<^sub>1 s\<^sub>1a nDone by simp
+ then have s2: "(h\<^sub>2, l\<^sub>2, sh\<^sub>2) = s\<^sub>1" using D init' init_ProcessingE s\<^sub>1 s\<^sub>1a shD by blast
+ then show ?thesis
+ using eval_finalId SCallInit[OF eval_finalsId[of "map Val vs" P "(h,l,sh)"]
+ SCallInitRed.hyps(1)] init init' len bstep nDone D pns body s' s\<^sub>1 s\<^sub>1a shD vals vs
+ SCallInitRed.hyps(2-3) s2 by auto
+ next
+ fix vsa h\<^sub>2 l\<^sub>2 sh\<^sub>2 Ts T pns' body' D' sfs h\<^sub>3 l\<^sub>3 sh\<^sub>3
+ assume s': "s' = (h\<^sub>3, l\<^sub>2, sh\<^sub>3)" and vals: "P \<turnstile> \<langle>map Val vs,s\<^sub>1\<rangle> [\<Rightarrow>] \<langle>map Val vsa,(h\<^sub>2, l\<^sub>2, sh\<^sub>2)\<rangle>"
+ and method: "P \<turnstile> C sees M, Static : Ts\<rightarrow>T = (pns', body') in D'"
+ and "sh\<^sub>2 D' = \<lfloor>(sfs, Done)\<rfloor> \<or> M = clinit \<and> sh\<^sub>2 D' = \<lfloor>(sfs, Processing)\<rfloor>"
+ and len: "length vsa = length pns'"
+ and bstep: "P \<turnstile> \<langle>body',(h\<^sub>2, [pns' [\<mapsto>] vsa], sh\<^sub>2)\<rangle> \<Rightarrow> \<langle>e',(h\<^sub>3, l\<^sub>3, sh\<^sub>3)\<rangle>"
+ have vs: "vs = vsa" and s\<^sub>1a: "s\<^sub>1 = (h\<^sub>2, l\<^sub>2, sh\<^sub>2)"
+ using evals_finals_same[OF _ vals] map_Val_eq by auto
+ have D: "D = D'" and pns: "pns = pns'" and body: "body = body'"
+ using sees_method_fun[OF SCallInitRed.hyps(1) method] by auto
+ then show ?thesis using SCallInit[OF eval_finalsId[of "map Val vs" P "(h,l,sh)"]
+ SCallInitRed.hyps(1)] bstep SCallInitRed.hyps(2-3) len s' s\<^sub>1a vals vs init by auto
+ qed
+ next
+ fix e assume e': "e' = throw e"
+ and init: "P \<turnstile> \<langle>INIT D ([D],False) \<leftarrow> unit,(h, l, sh)\<rangle> \<Rightarrow> \<langle>throw e,s'\<rangle>"
+ obtain h' l' sh' where s': "s' = (h',l',sh')" by(cases s')
+ then obtain sfs i where shC: "sh' D = \<lfloor>(sfs, i)\<rfloor>" and iDP: "i = Error"
+ using init_throw_PD[OF init] by auto
+ then show ?thesis using SCallInitRed.hyps(2-3) init e'
+ SCallInitThrow[OF eval_finalsId[of "map Val vs" _] SCallInitRed.hyps(1)]
+ by auto
+ qed
+next
+ case (RedSCallNone C M vs s b)
+ then have tes: "THROW NoSuchMethodError = e' \<and> s = s'"
+ using eval_final_same by simp
+ have "P \<turnstile> \<langle>map Val vs,s\<rangle> [\<Rightarrow>] \<langle>map Val vs,s\<rangle>" using eval_finalsId by simp
+ then show ?case using RedSCallNone eval_evals.SCallNone tes by auto
+next
+ case (RedSCallNonStatic C M Ts T m D vs s b)
+ then have tes: "THROW IncompatibleClassChangeError = e' \<and> s = s'"
+ using eval_final_same by simp
+ have "P \<turnstile> \<langle>map Val vs,s\<rangle> [\<Rightarrow>] \<langle>map Val vs,s\<rangle>" using eval_finalsId by simp
+ then show ?case using RedSCallNonStatic eval_evals.SCallNonStatic tes by auto
+next
+ case InitBlockRed
+ thus ?case
+ by (fastforce elim!: eval_cases intro: eval_evals.intros eval_finalId
+ simp: assigned_def map_upd_triv fun_upd_same)
+next
+ case (RedInitBlock V T v u s b)
+ then have "P \<turnstile> \<langle>Val u,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>" by simp
+ then obtain s': "s'=s" and e': "e'=Val u" by cases simp
+ obtain h l sh where s: "s=(h,l,sh)" by (cases s)
+ have "P \<turnstile> \<langle>{V:T :=Val v; Val u},(h,l,sh)\<rangle> \<Rightarrow> \<langle>Val u,(h, (l(V\<mapsto>v))(V:=l V), sh)\<rangle>"
+ by (fastforce intro!: eval_evals.intros)
+ then have "P \<turnstile> \<langle>{V:T := Val v; Val u},s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>" using s s' e' by simp
+ then show ?case by simp
+next
+ case BlockRedNone
+ thus ?case
+ by (fastforce elim!: eval_cases intro: eval_evals.intros
+ simp add: fun_upd_same fun_upd_idem)
+next
+ case BlockRedSome
+ thus ?case
+ by (fastforce elim!: eval_cases intro: eval_evals.intros
+ simp add: fun_upd_same fun_upd_idem)
+next
+ case (RedBlock V T v s b)
+ then have "P \<turnstile> \<langle>Val v,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>" by simp
+ then obtain s': "s'=s" and e': "e'=Val v"
+ by cases simp
+ obtain h l sh where s: "s=(h,l,sh)" by (cases s)
+ have "P \<turnstile> \<langle>Val v,(h,l(V:=None),sh)\<rangle> \<Rightarrow> \<langle>Val v,(h,l(V:=None),sh)\<rangle>"
+ by (rule eval_evals.intros)
+ hence "P \<turnstile> \<langle>{V:T;Val v},(h,l,sh)\<rangle> \<Rightarrow> \<langle>Val v,(h,(l(V:=None))(V:=l V),sh)\<rangle>"
+ by (rule eval_evals.Block)
+ then have "P \<turnstile> \<langle>{V:T; Val v},s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>" using s s' e' by simp
+ then show ?case by simp
+next
+ case (SeqRed e s b e1 s1 b1 e\<^sub>2) show ?case
+ proof(cases "val_of e")
+ case None show ?thesis
+ proof(cases "lass_val_of e")
+ case lNone:None
+ then have bconf: "P,shp s \<turnstile>\<^sub>b (e,b) \<surd>" using SeqRed.prems(1) None by simp
+ then show ?thesis using SeqRed using seq_ext by fastforce
+ next
+ case (Some p)
+ obtain V' v' where p: "p = (V',v')" and e: "e = V':=Val v'"
+ using lass_val_of_spec[OF Some] by(cases p, auto)
+ obtain h l sh h' l' sh' where s: "s = (h,l,sh)" and s1: "s1 = (h',l',sh')" by(cases s, cases s1)
+ then have red: "P \<turnstile> \<langle>e,(h,l,sh),b\<rangle> \<rightarrow> \<langle>e1,(h',l',sh'),b1\<rangle>" using SeqRed.hyps(1) by simp
+ then have s\<^sub>1': "e1 = unit \<and> h' = h \<and> l' = l(V' \<mapsto> v') \<and> sh' = sh"
+ using lass_val_of_red[OF Some red] p e by simp
+ then have eval: "P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>e1,s1\<rangle>" using e s s1 LAss Val by auto
+ then show ?thesis
+ by (metis SeqRed.prems(2) eval_final eval_final_same seq_ext)
+ qed
+ next
+ case (Some a) then show ?thesis using SeqRed.hyps(1) val_no_step by blast
+ qed
+next
+ case RedSeq
+ thus ?case
+ by (fastforce elim: eval_cases intro: eval_evals.intros)
+next
+ case CondRed
+ thus ?case
+ by (fastforce elim: eval_cases intro: eval_evals.intros simp: val_no_step)
+next
+ case RedCondT
+ thus ?case
+ by (fastforce elim: eval_cases intro: eval_evals.intros)
+next
+ case RedCondF
+ thus ?case
+ by (fastforce elim: eval_cases intro: eval_evals.intros)
+next
+ case RedWhile
+ thus ?case
+ by (auto simp add: unfold_while intro:eval_evals.intros elim:eval_cases)
+next
+ case ThrowRed then show ?case by(fastforce elim: eval_cases simp: eval_evals.intros)
+next
+ case RedThrowNull
+ thus ?case
+ by (fastforce elim: eval_cases intro: eval_evals.intros)
+next
+ case TryRed thus ?case
+ by(fastforce elim: eval_cases intro: eval_evals.intros)
+next
+ case RedTryCatch
+ thus ?case
+ by (fastforce elim: eval_cases intro: eval_evals.intros)
+next
+ case (RedTryFail s a D fs C V e\<^sub>2 b)
+ thus ?case
+ by (cases s)(auto elim!: eval_cases intro: eval_evals.intros)
+next
+ case ListRed1
+ thus ?case
+ by (fastforce elim: evals_cases intro: eval_evals.intros simp: val_no_step)
+next
+ case ListRed2
+ thus ?case
+ by (fastforce elim!: evals_cases eval_cases
+ intro: eval_evals.intros eval_finalId)
+next
+ case (RedInit e1 C b s1 b') then show ?case using InitFinal by simp
+next
+ case (InitNoneRed sh C C' Cs e h l b)
+ show ?case using InitNone InitNoneRed.hyps InitNoneRed.prems(2) by auto
+next
+ case (RedInitDone sh C sfs C' Cs e h l b)
+ show ?case using InitDone RedInitDone.hyps RedInitDone.prems(2) by auto
+next
+ case (RedInitProcessing sh C sfs C' Cs e h l b)
+ show ?case using InitProcessing RedInitProcessing.hyps RedInitProcessing.prems(2) by auto
+next
+ case (RedInitError sh C sfs C' Cs e h l b)
+ show ?case using InitError RedInitError.hyps RedInitError.prems(2) by auto
+next
+ case (InitObjectRed sh C sfs sh' C' Cs e h l b) show ?case using InitObject InitObjectRed by auto
+next
+ case (InitNonObjectSuperRed sh C sfs D r sh' C' Cs e h l b)
+ show ?case using InitNonObject InitNonObjectSuperRed by auto
+next
+ case (RedInitRInit C' C Cs e h l sh b)
+ show ?case using InitRInit RedInitRInit by auto
+next
+ case (RInitRed e s b e'' s'' b'' C Cs e\<^sub>0)
+ then have IH: "\<And>e' s'. P \<turnstile> \<langle>e'',s''\<rangle> \<Rightarrow> \<langle>e',s'\<rangle> \<Longrightarrow> P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>" by simp
+ show ?case using RInitRed rinit_ext[OF IH] by simp
+next
+ case (RedRInit sh C sfs i sh' C' Cs v e h l b s' e')
+ then have init: "P \<turnstile> \<langle>(INIT C' (Cs,True) \<leftarrow> e), (h, l, sh(C \<mapsto> (sfs, Done)))\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+ using RedRInit by simp
+ then show ?case using RInit RedRInit.hyps(1) RedRInit.hyps(3) Val by fastforce
+next
+ case BinOpThrow2
+ thus ?case
+ by (fastforce elim: eval_cases intro: eval_evals.intros)
+next
+ case FAssThrow2
+ thus ?case
+ by (fastforce elim: eval_cases intro: eval_evals.intros)
+next
+ case SFAssThrow
+ then show ?case
+ by (fastforce elim: eval_cases intro: eval_evals.intros)
+next
+ case (CallThrowParams es vs e es' v M s b)
+ have val: "P \<turnstile> \<langle>Val v,s\<rangle> \<Rightarrow> \<langle>Val v,s\<rangle>" by (rule eval_evals.intros)
+ have eval_e: "P \<turnstile> \<langle>throw (e),s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>" using CallThrowParams by simp
+ then obtain xa where e': "e' = Throw xa" by (cases) (auto dest!: eval_final)
+ with list_eval_Throw [OF eval_e]
+ have vals: "P \<turnstile> \<langle>es,s\<rangle> [\<Rightarrow>] \<langle>map Val vs @ Throw xa # es',s'\<rangle>"
+ using CallThrowParams.hyps(1) eval_e list_eval_Throw by blast
+ then have "P \<turnstile> \<langle>Val v\<bullet>M(es),s\<rangle> \<Rightarrow> \<langle>Throw xa,s'\<rangle>"
+ using eval_evals.CallParamsThrow[OF val vals] by simp
+ thus ?case using e' by simp
+next
+ case (SCallThrowParams es vs e es' C M s b)
+ have eval_e: "P \<turnstile> \<langle>throw (e),s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>" using SCallThrowParams by simp
+ then obtain xa where e': "e' = Throw xa" by (cases) (auto dest!: eval_final)
+ then have "P \<turnstile> \<langle>es,s\<rangle> [\<Rightarrow>] \<langle>map Val vs @ Throw xa # es',s'\<rangle>"
+ using SCallThrowParams.hyps(1) eval_e list_eval_Throw by blast
+ then have "P \<turnstile> \<langle>C\<bullet>\<^sub>sM(es),s\<rangle> \<Rightarrow> \<langle>Throw xa,s'\<rangle>"
+ by (rule eval_evals.SCallParamsThrow)
+ thus ?case using e' by simp
+next
+ case (BlockThrow V T a s b)
+ then have "P \<turnstile> \<langle>Throw a, s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>" by simp
+ then obtain s': "s' = s" and e': "e' = Throw a"
+ by cases (auto elim!:eval_cases)
+ obtain h l sh where s: "s=(h,l,sh)" by (cases s)
+ have "P \<turnstile> \<langle>Throw a, (h,l(V:=None),sh)\<rangle> \<Rightarrow> \<langle>Throw a, (h,l(V:=None),sh)\<rangle>"
+ by (simp add:eval_evals.intros eval_finalId)
+ hence "P\<turnstile>\<langle>{V:T;Throw a},(h,l,sh)\<rangle> \<Rightarrow> \<langle>Throw a, (h,(l(V:=None))(V:=l V),sh)\<rangle>"
+ by (rule eval_evals.Block)
+ then have "P \<turnstile> \<langle>{V:T; Throw a},s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>" using s s' e' by simp
+ then show ?case by simp
+next
+ case (InitBlockThrow V T v a s b)
+ then have "P \<turnstile> \<langle>Throw a,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>" by simp
+ then obtain s': "s' = s" and e': "e' = Throw a"
+ by cases (auto elim!:eval_cases)
+ obtain h l sh where s: "s = (h,l,sh)" by (cases s)
+ have "P \<turnstile> \<langle>{V:T :=Val v; Throw a},(h,l,sh)\<rangle> \<Rightarrow> \<langle>Throw a, (h, (l(V\<mapsto>v))(V:=l V),sh)\<rangle>"
+ by(fastforce intro:eval_evals.intros)
+ then have "P \<turnstile> \<langle>{V:T := Val v; Throw a},s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>" using s s' e' by simp
+ then show ?case by simp
+next
+ case (RInitInitThrow sh C sfs i sh' a D Cs e h l b)
+ have IH: "\<And>e' s'. P \<turnstile> \<langle>RI (D,Throw a) ; Cs \<leftarrow> e,(h, l, sh(C \<mapsto> (sfs, Error)))\<rangle> \<Rightarrow> \<langle>e',s'\<rangle> \<Longrightarrow>
+ P \<turnstile> \<langle>RI (C,Throw a) ; D # Cs \<leftarrow> e,(h, l, sh)\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+ using RInitInitFail[OF eval_finalId] RInitInitThrow by simp
+ then show ?case using RInitInitThrow.hyps(2) RInitInitThrow.prems(2) by auto
+next
+ case (RInitThrow sh C sfs i sh' a e h l b)
+ then have e': "e' = Throw a" and s': "s' = (h,l,sh')"
+ using eval_final_same final_def by fastforce+
+ show ?case using RInitFailFinal RInitThrow.hyps(1) RInitThrow.hyps(2) e' eval_finalId s' by auto
+qed(auto elim: eval_cases simp: eval_evals.intros)
+(*>*)
+
+(*<*)
+(* ... und wieder anschalten: *)
+declare split_paired_All [simp] split_paired_Ex [simp]
+(*>*)
+
+text \<open> Its extension to @{text"\<rightarrow>*"}: \<close>
+
+lemma extend_eval:
+assumes wf: "wwf_J_prog P"
+shows "\<lbrakk> P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e'',s'',b''\<rangle>; P \<turnstile> \<langle>e'',s''\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>;
+ iconf (shp s) e; P,shp s \<turnstile>\<^sub>b (e::expr,b) \<surd> \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+(*<*)
+proof (induct rule: converse_rtrancl_induct3)
+ case refl then show ?case by simp
+next
+ case (step e1 s1 b1 e2 s2 b2)
+ then have ic: "iconf (shp s2) e2" using Red_preserves_iconf local.wf by blast
+ then have ec: "P,shp s2 \<turnstile>\<^sub>b (e2,b2) \<surd>"
+ using Red_preserves_bconf local.wf step.hyps(1) step.prems(2) step.prems(3) by blast
+ show ?case using step ic ec extend_1_eval[OF wf step.hyps(1)] by simp
+qed
+(*>*)
+
+
+lemma extend_evals:
+assumes wf: "wwf_J_prog P"
+shows "\<lbrakk> P \<turnstile> \<langle>es,s,b\<rangle> [\<rightarrow>]* \<langle>es'',s'',b''\<rangle>; P \<turnstile> \<langle>es'',s''\<rangle> [\<Rightarrow>] \<langle>es',s'\<rangle>;
+ iconfs (shp s) es; P,shp s \<turnstile>\<^sub>b (es::expr list,b) \<surd> \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>es,s\<rangle> [\<Rightarrow>] \<langle>es',s'\<rangle>"
+(*<*)
+proof (induct rule: converse_rtrancl_induct3)
+ case refl then show ?case by simp
+next
+ case (step es1 s1 b1 es2 s2 b2)
+ then have ic: "iconfs (shp s2) es2" using Reds_preserves_iconf local.wf by blast
+ then have ec: "P,shp s2 \<turnstile>\<^sub>b (es2,b2) \<surd>"
+ using Reds_preserves_bconf local.wf step.hyps(1) step.prems(2) step.prems(3) by blast
+ show ?case using step ic ec extend_1_evals[OF wf step.hyps(1)] by simp
+qed
+(*>*)
+
+text \<open> Finally, small step semantics can be simulated by big step semantics:
+\<close>
+
+theorem
+assumes wf: "wwf_J_prog P"
+shows small_by_big:
+ "\<lbrakk>P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle>; iconf (shp s) e; P,shp s \<turnstile>\<^sub>b (e,b) \<surd>; final e'\<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+and "\<lbrakk>P \<turnstile> \<langle>es,s,b\<rangle> [\<rightarrow>]* \<langle>es',s',b'\<rangle>; iconfs (shp s) es; P,shp s \<turnstile>\<^sub>b (es,b) \<surd>; finals es'\<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>es,s\<rangle> [\<Rightarrow>] \<langle>es',s'\<rangle>"
+(*<*)
+proof -
+ note wf
+ moreover assume "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle>"
+ moreover assume "final e'"
+ then have "P \<turnstile> \<langle>e',s'\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+ by (simp add: eval_finalId)
+ moreover assume "iconf (shp s) e"
+ moreover assume "P,shp s \<turnstile>\<^sub>b (e,b) \<surd>"
+ ultimately show "P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>"
+ by (rule extend_eval)
+next
+ assume fins: "finals es'"
+ note wf
+ moreover assume "P \<turnstile> \<langle>es,s,b\<rangle> [\<rightarrow>]* \<langle>es',s',b'\<rangle>"
+ moreover have "P \<turnstile> \<langle>es',s'\<rangle> [\<Rightarrow>] \<langle>es',s'\<rangle>" using fins
+ by (rule eval_finalsId)
+ moreover assume "iconfs (shp s) es"
+ moreover assume "P,shp s \<turnstile>\<^sub>b (es,b) \<surd>"
+ ultimately show "P \<turnstile> \<langle>es,s\<rangle> [\<Rightarrow>] \<langle>es',s'\<rangle>"
+ by (rule extend_evals)
+qed
+(*>*)
+
+subsection "Equivalence"
+
+text\<open> And now, the crowning achievement: \<close>
+
+corollary big_iff_small:
+"\<lbrakk> wwf_J_prog P; iconf (shp s) e; P,shp s \<turnstile>\<^sub>b (e::expr,b) \<surd> \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle> = (P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e',s',False\<rangle> \<and> final e')"
+(*<*)by(blast dest: big_by_small eval_final small_by_big)(*>*)
+
+corollary big_iff_small_WT:
+ "wwf_J_prog P \<Longrightarrow> P,E \<turnstile> e::T \<Longrightarrow> P,shp s \<turnstile>\<^sub>b (e,b) \<surd> \<Longrightarrow>
+ P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle> = (P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e',s',False\<rangle> \<and> final e')"
+(*<*)by(blast dest: big_iff_small WT_nsub_RI nsub_RI_iconf)(*>*)
+
+
+subsection \<open> Lifting type safety to @{text"\<Rightarrow>"} \<close>
+
+text\<open> \dots and now to the big step semantics, just for fun. \<close>
+
+lemma eval_preserves_sconf:
+fixes s::state and s'::state
+assumes "wf_J_prog P" and "P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>" and "iconf (shp s) e"
+ and "P,E \<turnstile> e::T" and "P,E \<turnstile> s\<surd>"
+shows "P,E \<turnstile> s'\<surd>"
+(*<*)
+proof -
+ have "P,shp s \<turnstile>\<^sub>b (e,False) \<surd>" by(simp add: bconf_def)
+ with assms show ?thesis
+ by(blast intro:Red_preserves_sconf Red_preserves_iconf Red_preserves_bconf big_by_small
+ WT_implies_WTrt wf_prog_wwf_prog)
+qed
+(*>*)
+
+
+lemma eval_preserves_type:
+fixes s::state
+assumes wf: "wf_J_prog P"
+ and "P \<turnstile> \<langle>e,s\<rangle> \<Rightarrow> \<langle>e',s'\<rangle>" and "P,E \<turnstile> s\<surd>" and "iconf (shp s) e" and "P,E \<turnstile> e::T"
+shows "\<exists>T'. P \<turnstile> T' \<le> T \<and> P,E,hp s',shp s' \<turnstile> e':T'"
+(*<*)
+proof -
+ have "P,shp s \<turnstile>\<^sub>b (e,False) \<surd>" by(simp add: bconf_def)
+ with assms show ?thesis by(blast dest:big_by_small[OF wf_prog_wwf_prog[OF wf]]
+ WT_implies_WTrt Red_preserves_type[OF wf])
+qed
+(*>*)
+
+
+end
diff --git a/thys/JinjaDCI/J/Expr.thy b/thys/JinjaDCI/J/Expr.thy
--- a/thys/JinjaDCI/J/Expr.thy
+++ b/thys/JinjaDCI/J/Expr.thy
@@ -1,579 +1,579 @@
-(* Title: JinjaDCI/J/Expr.thy
- Author: Tobias Nipkow, Susannah Mansky
- Copyright 2003 Technische Universitaet Muenchen, 2019-20 UIUC
-
- Based on the Jinja theory J/Expr.thy by Tobias Nipkow
-*)
-
-section \<open> Expressions \<close>
-
-theory Expr
-imports "../Common/Exceptions"
-begin
-
-datatype bop = Eq | Add \<comment> \<open>names of binary operations\<close>
-
-datatype 'a exp
- = new cname \<comment> \<open>class instance creation\<close>
- | Cast cname "('a exp)" \<comment> \<open>type cast\<close>
- | Val val \<comment> \<open>value\<close>
- | BinOp "('a exp)" bop "('a exp)" ("_ \<guillemotleft>_\<guillemotright> _" [80,0,81] 80) \<comment> \<open>binary operation\<close>
- | Var 'a \<comment> \<open>local variable (incl. parameter)\<close>
- | LAss 'a "('a exp)" ("_:=_" [90,90]90) \<comment> \<open>local assignment\<close>
- | FAcc "('a exp)" vname cname ("_\<bullet>_{_}" [10,90,99]90) \<comment> \<open>field access\<close>
- | SFAcc cname vname cname ("_\<bullet>\<^sub>s_{_}" [10,90,99]90) \<comment> \<open>static field access\<close>
- | FAss "('a exp)" vname cname "('a exp)" ("_\<bullet>_{_} := _" [10,90,99,90]90) \<comment> \<open>field assignment\<close>
- | SFAss cname vname cname "('a exp)" ("_\<bullet>\<^sub>s_{_} := _" [10,90,99,90]90) \<comment> \<open>static field assignment\<close>
- | Call "('a exp)" mname "('a exp list)" ("_\<bullet>_'(_')" [90,99,0] 90) \<comment> \<open>method call\<close>
- | SCall cname mname "('a exp list)" ("_\<bullet>\<^sub>s_'(_')" [90,99,0] 90) \<comment> \<open>static method call\<close>
- | Block 'a ty "('a exp)" ("'{_:_; _}")
- | Seq "('a exp)" "('a exp)" ("_;;/ _" [61,60]60)
- | Cond "('a exp)" "('a exp)" "('a exp)" ("if '(_') _/ else _" [80,79,79]70)
- | While "('a exp)" "('a exp)" ("while '(_') _" [80,79]70)
- | throw "('a exp)"
- | TryCatch "('a exp)" cname 'a "('a exp)" ("try _/ catch'(_ _') _" [0,99,80,79] 70)
- | INIT cname "cname list" bool "('a exp)" ("INIT _ '(_,_') \<leftarrow> _" [60,60,60,60] 60) \<comment> \<open>internal initialization command: class, list of superclasses to initialize, preparation flag; command on hold\<close>
- | RI cname "('a exp)" "cname list" "('a exp)" ("RI '(_,_') ; _ \<leftarrow> _" [60,60,60,60] 60) \<comment> \<open>running of the initialization procedure for class with expression, classes still to initialize command on hold\<close>
-
-type_synonym
- expr = "vname exp" \<comment> \<open>Jinja expression\<close>
-type_synonym
- J_mb = "vname list \<times> expr" \<comment> \<open>Jinja method body: parameter names and expression\<close>
-type_synonym
- J_prog = "J_mb prog" \<comment> \<open>Jinja program\<close>
-
-type_synonym
- init_stack = "expr list \<times> bool" \<comment> \<open>Stack of expressions waiting on initialization in small step; indicator boolean True if current expression has been init checked\<close>
-
-text\<open>The semantics of binary operators: \<close>
-
-fun binop :: "bop \<times> val \<times> val \<Rightarrow> val option" where
- "binop(Eq,v\<^sub>1,v\<^sub>2) = Some(Bool (v\<^sub>1 = v\<^sub>2))"
-| "binop(Add,Intg i\<^sub>1,Intg i\<^sub>2) = Some(Intg(i\<^sub>1+i\<^sub>2))"
-| "binop(bop,v\<^sub>1,v\<^sub>2) = None"
-
-lemma [simp]:
- "(binop(Add,v\<^sub>1,v\<^sub>2) = Some v) = (\<exists>i\<^sub>1 i\<^sub>2. v\<^sub>1 = Intg i\<^sub>1 \<and> v\<^sub>2 = Intg i\<^sub>2 \<and> v = Intg(i\<^sub>1+i\<^sub>2))"
-(*<*)by(cases v\<^sub>1; cases v\<^sub>2) auto(*>*)
-
-
-lemma map_Val_throw_eq:
- "map Val vs @ throw ex # es = map Val vs' @ throw ex' # es' \<Longrightarrow> ex = ex'"
-(*<*)by(induct vs arbitrary: vs'; case_tac vs') auto(*>*)
-
-lemma map_Val_nthrow_neq:
- "map Val vs = map Val vs' @ throw ex' # es' \<Longrightarrow> False"
-(*<*)by(induct vs arbitrary: vs'; case_tac vs') auto(*>*)
-
-lemma map_Val_eq:
- "map Val vs = map Val vs' \<Longrightarrow> vs = vs'"
-(*<*)by(induct vs arbitrary: vs'; case_tac vs') auto(*>*)
-
-
-lemma init_rhs_neq [simp]: "e \<noteq> INIT C (Cs,b) \<leftarrow> e"
-proof -
- have "size e \<noteq> size (INIT C (Cs,b) \<leftarrow> e)" by auto
- then show ?thesis by fastforce
-qed
-
-lemma init_rhs_neq' [simp]: "INIT C (Cs,b) \<leftarrow> e \<noteq> e"
-proof -
- have "size e \<noteq> size (INIT C (Cs,b) \<leftarrow> e)" by auto
- then show ?thesis by fastforce
-qed
-
-lemma ri_rhs_neq [simp]: "e \<noteq> RI(C,e');Cs \<leftarrow> e"
-proof -
- have "size e \<noteq> size (RI(C,e');Cs \<leftarrow> e)" by auto
- then show ?thesis by fastforce
-qed
-
-lemma ri_rhs_neq' [simp]: "RI(C,e');Cs \<leftarrow> e \<noteq> e"
-proof -
- have "size e \<noteq> size (RI(C,e');Cs \<leftarrow> e)" by auto
- then show ?thesis by fastforce
-qed
-
-subsection "Syntactic sugar"
-
-abbreviation (input)
- InitBlock:: "'a \<Rightarrow> ty \<Rightarrow> 'a exp \<Rightarrow> 'a exp \<Rightarrow> 'a exp" ("(1'{_:_ := _;/ _})") where
- "InitBlock V T e1 e2 == {V:T; V := e1;; e2}"
-
-abbreviation unit where "unit == Val Unit"
-abbreviation null where "null == Val Null"
-abbreviation "addr a == Val(Addr a)"
-abbreviation "true == Val(Bool True)"
-abbreviation "false == Val(Bool False)"
-
-abbreviation
- Throw :: "addr \<Rightarrow> 'a exp" where
- "Throw a == throw(Val(Addr a))"
-
-abbreviation
- THROW :: "cname \<Rightarrow> 'a exp" where
- "THROW xc == Throw(addr_of_sys_xcpt xc)"
-
-
-subsection\<open>Free Variables\<close>
-
-primrec fv :: "expr \<Rightarrow> vname set" and fvs :: "expr list \<Rightarrow> vname set" where
- "fv(new C) = {}"
-| "fv(Cast C e) = fv e"
-| "fv(Val v) = {}"
-| "fv(e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2) = fv e\<^sub>1 \<union> fv e\<^sub>2"
-| "fv(Var V) = {V}"
-| "fv(LAss V e) = {V} \<union> fv e"
-| "fv(e\<bullet>F{D}) = fv e"
-| "fv(C\<bullet>\<^sub>sF{D}) = {}"
-| "fv(e\<^sub>1\<bullet>F{D}:=e\<^sub>2) = fv e\<^sub>1 \<union> fv e\<^sub>2"
-| "fv(C\<bullet>\<^sub>sF{D}:=e\<^sub>2) = fv e\<^sub>2"
-| "fv(e\<bullet>M(es)) = fv e \<union> fvs es"
-| "fv(C\<bullet>\<^sub>sM(es)) = fvs es"
-| "fv({V:T; e}) = fv e - {V}"
-| "fv(e\<^sub>1;;e\<^sub>2) = fv e\<^sub>1 \<union> fv e\<^sub>2"
-| "fv(if (b) e\<^sub>1 else e\<^sub>2) = fv b \<union> fv e\<^sub>1 \<union> fv e\<^sub>2"
-| "fv(while (b) e) = fv b \<union> fv e"
-| "fv(throw e) = fv e"
-| "fv(try e\<^sub>1 catch(C V) e\<^sub>2) = fv e\<^sub>1 \<union> (fv e\<^sub>2 - {V})"
-| "fv(INIT C (Cs,b) \<leftarrow> e) = fv e"
-| "fv(RI (C,e);Cs \<leftarrow> e') = fv e \<union> fv e'"
-| "fvs([]) = {}"
-| "fvs(e#es) = fv e \<union> fvs es"
-
-lemma [simp]: "fvs(es\<^sub>1 @ es\<^sub>2) = fvs es\<^sub>1 \<union> fvs es\<^sub>2"
-(*<*)by (induct es\<^sub>1 type:list) auto(*>*)
-
-lemma [simp]: "fvs(map Val vs) = {}"
-(*<*)by (induct vs) auto(*>*)
-
-
-subsection\<open>Accessing expression constructor arguments\<close>
-
-fun val_of :: "'a exp \<Rightarrow> val option" where
-"val_of (Val v) = Some v" |
-"val_of _ = None"
-
-lemma val_of_spec: "val_of e = Some v \<Longrightarrow> e = Val v"
-proof(cases e) qed(auto)
-
-fun lass_val_of :: "'a exp \<Rightarrow> ('a \<times> val) option" where
-"lass_val_of (V:=Val v) = Some (V, v)" |
-"lass_val_of _ = None"
-
-lemma lass_val_of_spec:
-assumes "lass_val_of e = \<lfloor>a\<rfloor>"
-shows "e = (fst a:=Val (snd a))"
-using assms proof(cases e)
- case (LAss V e') then show ?thesis using assms proof(cases e')qed(auto)
-qed(auto)
-
-fun map_vals_of :: "'a exp list \<Rightarrow> val list option" where
-"map_vals_of (e#es) = (case val_of e of Some v \<Rightarrow> (case map_vals_of es of Some vs \<Rightarrow> Some (v#vs)
- | _ \<Rightarrow> None)
- | _ \<Rightarrow> None)" |
-"map_vals_of [] = Some []"
-
-lemma map_vals_of_spec: "map_vals_of es = Some vs \<Longrightarrow> es = map Val vs"
-proof(induct es arbitrary: vs) qed(auto simp: val_of_spec)
-
-lemma map_vals_of_Vals[simp]: "map_vals_of (map Val vs) = \<lfloor>vs\<rfloor>" by(induct vs, auto)
-
-lemma map_vals_of_throw[simp]:
- "map_vals_of (map Val vs @ throw e # es') = None"
- by(induct vs, auto)
-
-
-fun bool_of :: "'a exp \<Rightarrow> bool option" where
-"bool_of true = Some True" |
-"bool_of false = Some False" |
-"bool_of _ = None"
-
-lemma bool_of_specT:
-assumes "bool_of e = Some True" shows "e = true"
-proof -
- have "bool_of e = Some True" by fact
- then show ?thesis
- proof(cases e)
- case (Val x3) with assms show ?thesis
- proof(cases x3)
- case (Bool x) with assms Val show ?thesis
- proof(cases x)qed(auto)
- qed(simp_all)
- qed(auto)
-qed
-
-lemma bool_of_specF:
-assumes "bool_of e = Some False" shows "e = false"
-proof -
- have "bool_of e = Some False" by fact
- then show ?thesis
- proof(cases e)
- case (Val x3) with assms show ?thesis
- proof(cases x3)
- case (Bool x) with assms Val show ?thesis
- proof(cases x)qed(auto)
- qed(simp_all)
- qed(auto)
-qed
-
-
-fun throw_of :: "'a exp \<Rightarrow> 'a exp option" where
-"throw_of (throw e') = Some e'" |
-"throw_of _ = None"
-
-lemma throw_of_spec: "throw_of e = Some e' \<Longrightarrow> e = throw e'"
-proof(cases e) qed(auto)
-
-fun init_exp_of :: "'a exp \<Rightarrow> 'a exp option" where
-"init_exp_of (INIT C (Cs,b) \<leftarrow> e) = Some e" |
-"init_exp_of (RI(C,e');Cs \<leftarrow> e) = Some e" |
-"init_exp_of _ = None"
-
-lemma init_exp_of_neq [simp]: "init_exp_of e = \<lfloor>e'\<rfloor> \<Longrightarrow> e' \<noteq> e" by(cases e, auto)
-lemma init_exp_of_neq'[simp]: "init_exp_of e = \<lfloor>e'\<rfloor> \<Longrightarrow> e \<noteq> e'" by(cases e, auto)
-
-
-subsection\<open>Class initialization\<close>
-
-text \<open> This section defines a few functions that return information
- about an expression's current initialization status. \<close>
-
- \<comment> \<open> True if expression contains @{text INIT}, @{text RI}, or a call to a static method @{term clinit} \<close>
-primrec sub_RI :: "'a exp \<Rightarrow> bool" and sub_RIs :: "'a exp list \<Rightarrow> bool" where
- "sub_RI(new C) = False"
-| "sub_RI(Cast C e) = sub_RI e"
-| "sub_RI(Val v) = False"
-| "sub_RI(e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2) = (sub_RI e\<^sub>1 \<or> sub_RI e\<^sub>2)"
-| "sub_RI(Var V) = False"
-| "sub_RI(LAss V e) = sub_RI e"
-| "sub_RI(e\<bullet>F{D}) = sub_RI e"
-| "sub_RI(C\<bullet>\<^sub>sF{D}) = False"
-| "sub_RI(e\<^sub>1\<bullet>F{D}:=e\<^sub>2) = (sub_RI e\<^sub>1 \<or> sub_RI e\<^sub>2)"
-| "sub_RI(C\<bullet>\<^sub>sF{D}:=e\<^sub>2) = sub_RI e\<^sub>2"
-| "sub_RI(e\<bullet>M(es)) = (sub_RI e \<or> sub_RIs es)"
-| "sub_RI(C\<bullet>\<^sub>sM(es)) = (M = clinit \<or> sub_RIs es)"
-| "sub_RI({V:T; e}) = sub_RI e"
-| "sub_RI(e\<^sub>1;;e\<^sub>2) = (sub_RI e\<^sub>1 \<or> sub_RI e\<^sub>2)"
-| "sub_RI(if (b) e\<^sub>1 else e\<^sub>2) = (sub_RI b \<or> sub_RI e\<^sub>1 \<or> sub_RI e\<^sub>2)"
-| "sub_RI(while (b) e) = (sub_RI b \<or> sub_RI e)"
-| "sub_RI(throw e) = sub_RI e"
-| "sub_RI(try e\<^sub>1 catch(C V) e\<^sub>2) = (sub_RI e\<^sub>1 \<or> sub_RI e\<^sub>2)"
-| "sub_RI(INIT C (Cs,b) \<leftarrow> e) = True"
-| "sub_RI(RI (C,e);Cs \<leftarrow> e') = True"
-| "sub_RIs([]) = False"
-| "sub_RIs(e#es) = (sub_RI e \<or> sub_RIs es)"
-
-
-lemmas sub_RI_sub_RIs_induct = sub_RI.induct sub_RIs.induct
-
-lemma nsub_RIs_def[simp]:
- "\<not>sub_RIs es \<Longrightarrow> \<forall>e \<in> set es. \<not>sub_RI e"
- by(induct es, auto)
-
-lemma sub_RI_base:
- "e = INIT C (Cs, b) \<leftarrow> e' \<or> e = RI(C,e\<^sub>0);Cs \<leftarrow> e' \<Longrightarrow> sub_RI e"
- by(cases e, auto)
-
-lemma nsub_RI_Vals[simp]: "\<not>sub_RIs (map Val vs)"
- by(induct vs, auto)
-
-lemma lass_val_of_nsub_RI: "lass_val_of e = \<lfloor>a\<rfloor> \<Longrightarrow> \<not>sub_RI e"
- by(drule lass_val_of_spec, simp)
-
-
- \<comment> \<open> is not currently initializing class @{text C'} (point past checking flag) \<close>
-primrec not_init :: "cname \<Rightarrow> 'a exp \<Rightarrow> bool" and not_inits :: "cname \<Rightarrow> 'a exp list \<Rightarrow> bool" where
- "not_init C' (new C) = True"
-| "not_init C' (Cast C e) = not_init C' e"
-| "not_init C' (Val v) = True"
-| "not_init C' (e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2) = (not_init C' e\<^sub>1 \<and> not_init C' e\<^sub>2)"
-| "not_init C' (Var V) = True"
-| "not_init C' (LAss V e) = not_init C' e"
-| "not_init C' (e\<bullet>F{D}) = not_init C' e"
-| "not_init C' (C\<bullet>\<^sub>sF{D}) = True"
-| "not_init C' (e\<^sub>1\<bullet>F{D}:=e\<^sub>2) = (not_init C' e\<^sub>1 \<and> not_init C' e\<^sub>2)"
-| "not_init C' (C\<bullet>\<^sub>sF{D}:=e\<^sub>2) = not_init C' e\<^sub>2"
-| "not_init C' (e\<bullet>M(es)) = (not_init C' e \<and> not_inits C' es)"
-| "not_init C' (C\<bullet>\<^sub>sM(es)) = not_inits C' es"
-| "not_init C' ({V:T; e}) = not_init C' e"
-| "not_init C' (e\<^sub>1;;e\<^sub>2) = (not_init C' e\<^sub>1 \<and> not_init C' e\<^sub>2)"
-| "not_init C' (if (b) e\<^sub>1 else e\<^sub>2) = (not_init C' b \<and> not_init C' e\<^sub>1 \<and> not_init C' e\<^sub>2)"
-| "not_init C' (while (b) e) = (not_init C' b \<and> not_init C' e)"
-| "not_init C' (throw e) = not_init C' e"
-| "not_init C' (try e\<^sub>1 catch(C V) e\<^sub>2) = (not_init C' e\<^sub>1 \<and> not_init C' e\<^sub>2)"
-| "not_init C' (INIT C (Cs,b) \<leftarrow> e) = ((b \<longrightarrow> Cs = Nil \<or> C' \<noteq> hd Cs) \<and> C' \<notin> set(tl Cs) \<and> not_init C' e)"
-| "not_init C' (RI (C,e);Cs \<leftarrow> e') = (C' \<notin> set (C#Cs) \<and> not_init C' e \<and> not_init C' e')"
-| "not_inits C' ([]) = True"
-| "not_inits C' (e#es) = (not_init C' e \<and> not_inits C' es)"
-
-lemma not_inits_def'[simp]:
- "not_inits C es \<Longrightarrow> \<forall>e \<in> set es. not_init C e"
- by(induct es, auto)
-
-lemma nsub_RIs_not_inits_aux: "\<forall>e \<in> set es. \<not>sub_RI e \<longrightarrow> not_init C e
- \<Longrightarrow> \<not>sub_RIs es \<Longrightarrow> not_inits C es"
- by(induct es, auto)
-
-lemma nsub_RI_not_init: "\<not>sub_RI e \<Longrightarrow> not_init C e"
-proof(induct e) qed(auto intro: nsub_RIs_not_inits_aux)
-
-lemma nsub_RIs_not_inits: "\<not>sub_RIs es \<Longrightarrow> not_inits C es"
-by(rule nsub_RIs_not_inits_aux) (simp_all add: nsub_RI_not_init)
-
-subsection\<open>Subexpressions\<close>
-
- \<comment> \<open> all strictly smaller subexpressions; does not include self \<close>
- primrec subexp :: "'a exp \<Rightarrow> 'a exp set" and subexps :: "'a exp list \<Rightarrow> 'a exp set" where
- "subexp(new C) = {}"
-| "subexp(Cast C e) = {e} \<union> subexp e"
-| "subexp(Val v) = {}"
-| "subexp(e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2) = {e\<^sub>1, e\<^sub>2} \<union> subexp e\<^sub>1 \<union> subexp e\<^sub>2"
-| "subexp(Var V) = {}"
-| "subexp(LAss V e) = {e} \<union> subexp e"
-| "subexp(e\<bullet>F{D}) = {e} \<union> subexp e"
-| "subexp(C\<bullet>\<^sub>sF{D}) = {}"
-| "subexp(e\<^sub>1\<bullet>F{D}:=e\<^sub>2) = {e\<^sub>1, e\<^sub>2} \<union> subexp e\<^sub>1 \<union> subexp e\<^sub>2"
-| "subexp(C\<bullet>\<^sub>sF{D}:=e\<^sub>2) = {e\<^sub>2} \<union>subexp e\<^sub>2"
-| "subexp(e\<bullet>M(es)) = {e} \<union> set es \<union> subexp e \<union> subexps es"
-| "subexp(C\<bullet>\<^sub>sM(es)) = set es \<union> subexps es"
-| "subexp({V:T; e}) = {e} \<union> subexp e"
-| "subexp(e\<^sub>1;;e\<^sub>2) = {e\<^sub>1, e\<^sub>2} \<union> subexp e\<^sub>1 \<union> subexp e\<^sub>2"
-| "subexp(if (b) e\<^sub>1 else e\<^sub>2) = {b, e\<^sub>1, e\<^sub>2} \<union> subexp b \<union> subexp e\<^sub>1 \<union> subexp e\<^sub>2"
-| "subexp(while (b) e) = {b, e} \<union> subexp b \<union> subexp e"
-| "subexp(throw e) = {e} \<union> subexp e"
-| "subexp(try e\<^sub>1 catch(C V) e\<^sub>2) = {e\<^sub>1, e\<^sub>2} \<union> subexp e\<^sub>1 \<union> subexp e\<^sub>2"
-| "subexp(INIT C (Cs,b) \<leftarrow> e) = {e} \<union> subexp e"
-| "subexp(RI (C,e);Cs \<leftarrow> e') = {e, e'} \<union> subexp e \<union> subexp e'"
-| "subexps([]) = {}"
-| "subexps(e#es) = {e} \<union> subexp e \<union> subexps es"
-
-
-lemmas subexp_subexps_induct = subexp.induct subexps.induct
-
-abbreviation subexp_of :: "'a exp \<Rightarrow> 'a exp \<Rightarrow> bool" where
- "subexp_of e e' \<equiv> e \<in> subexp e'"
-
-lemma subexp_size_le:
- "(e' \<in> subexp e \<longrightarrow> size e' < size e) \<and> (e' \<in> subexps es \<longrightarrow> size e' < size_list size es)"
-proof(induct rule: subexp_subexps.induct)
- case Call:11 then show ?case using not_less_eq size_list_estimation by fastforce
-next
- case SCall:12 then show ?case using not_less_eq size_list_estimation by fastforce
-qed(auto)
-
-lemma subexps_def2: "subexps es = set es \<union> (\<Union>e \<in> set es. subexp e)" by(induct es, auto)
-
- \<comment> \<open> strong induction \<close>
-lemma shows subexp_induct[consumes 1]:
-"(\<And>e. subexp e = {} \<Longrightarrow> R e) \<Longrightarrow> (\<And>e. (\<And>e'. e' \<in> subexp e \<Longrightarrow> R e') \<Longrightarrow> R e)
- \<Longrightarrow> (\<And>es. (\<And>e'. e' \<in> subexps es \<Longrightarrow> R e') \<Longrightarrow> Rs es) \<Longrightarrow> (\<forall>e'. e' \<in> subexp e \<longrightarrow> R e') \<and> R e"
-and subexps_induct[consumes 1]:
- "(\<And>es. subexps es = {} \<Longrightarrow> Rs es) \<Longrightarrow> (\<And>e. (\<And>e'. e' \<in> subexp e \<Longrightarrow> R e') \<Longrightarrow> R e)
- \<Longrightarrow> (\<And>es. (\<And>e'. e' \<in> subexps es \<Longrightarrow> R e') \<Longrightarrow> Rs es) \<Longrightarrow> (\<forall>e'. e' \<in> subexps es \<longrightarrow> R e') \<and> Rs es"
-proof(induct rule: subexp_subexps_induct)
- case (Cast x1 x2)
- then have "(\<forall>e'. subexp_of e' x2 \<longrightarrow> R e') \<and> R x2" by fast
- then have "(\<forall>e'. subexp_of e' (Cast x1 x2) \<longrightarrow> R e')" by auto
- then show ?case using Cast.prems(2) by fast
-next
- case (BinOp x1 x2 x3)
- then have "(\<forall>e'. subexp_of e' x1 \<longrightarrow> R e') \<and> R x1" and "(\<forall>e'. subexp_of e' x3 \<longrightarrow> R e') \<and> R x3"
- by fast+
- then have "(\<forall>e'. subexp_of e' (x1 \<guillemotleft>x2\<guillemotright> x3) \<longrightarrow> R e')" by auto
- then show ?case using BinOp.prems(2) by fast
-next
- case (LAss x1 x2)
- then have "(\<forall>e'. subexp_of e' x2 \<longrightarrow> R e') \<and> R x2" by fast
- then have "(\<forall>e'. subexp_of e' (LAss x1 x2) \<longrightarrow> R e')" by auto
- then show ?case using LAss.prems(2) by fast
-next
- case (FAcc x1 x2 x3)
- then have "(\<forall>e'. subexp_of e' x1 \<longrightarrow> R e') \<and> R x1" by fast
- then have "(\<forall>e'. subexp_of e' (x1\<bullet>x2{x3}) \<longrightarrow> R e')" by auto
- then show ?case using FAcc.prems(2) by fast
-next
- case (FAss x1 x2 x3 x4)
- then have "(\<forall>e'. subexp_of e' x1 \<longrightarrow> R e') \<and> R x1" and "(\<forall>e'. subexp_of e' x4 \<longrightarrow> R e') \<and> R x4"
- by fast+
- then have "(\<forall>e'. subexp_of e' (x1\<bullet>x2{x3} := x4) \<longrightarrow> R e')" by auto
- then show ?case using FAss.prems(2) by fast
-next
- case (SFAss x1 x2 x3 x4)
- then have "(\<forall>e'. subexp_of e' x4 \<longrightarrow> R e') \<and> R x4" by fast
- then have "(\<forall>e'. subexp_of e' (x1\<bullet>\<^sub>sx2{x3} := x4) \<longrightarrow> R e')" by auto
- then show ?case using SFAss.prems(2) by fast
-next
- case (Call x1 x2 x3)
- then have "(\<forall>e'. subexp_of e' x1 \<longrightarrow> R e') \<and> R x1" and "(\<forall>e'. e' \<in> subexps x3 \<longrightarrow> R e') \<and> Rs x3"
- by fast+
- then have "(\<forall>e'. subexp_of e' (x1\<bullet>x2(x3)) \<longrightarrow> R e')" using subexps_def2 by auto
- then show ?case using Call.prems(2) by fast
-next
- case (SCall x1 x2 x3)
- then have "(\<forall>e'. e' \<in> subexps x3 \<longrightarrow> R e') \<and> Rs x3" by fast
- then have "(\<forall>e'. subexp_of e' (x1\<bullet>\<^sub>sx2(x3)) \<longrightarrow> R e')" using subexps_def2 by auto
- then show ?case using SCall.prems(2) by fast
-next
- case (Block x1 x2 x3)
- then have "(\<forall>e'. subexp_of e' x3 \<longrightarrow> R e') \<and> R x3" by fast
- then have "(\<forall>e'. subexp_of e' {x1:x2; x3} \<longrightarrow> R e')" by auto
- then show ?case using Block.prems(2) by fast
-next
- case (Seq x1 x2)
- then have "(\<forall>e'. subexp_of e' x1 \<longrightarrow> R e') \<and> R x1" and "(\<forall>e'. subexp_of e' x2 \<longrightarrow> R e') \<and> R x2"
- by fast+
- then have "(\<forall>e'. subexp_of e' (x1;; x2) \<longrightarrow> R e')" by auto
- then show ?case using Seq.prems(2) by fast
-next
- case (Cond x1 x2 x3)
- then have "(\<forall>e'. subexp_of e' x1 \<longrightarrow> R e') \<and> R x1" and "(\<forall>e'. subexp_of e' x2 \<longrightarrow> R e') \<and> R x2"
- and "(\<forall>e'. subexp_of e' x3 \<longrightarrow> R e') \<and> R x3" by fast+
- then have "(\<forall>e'. subexp_of e' (if (x1) x2 else x3) \<longrightarrow> R e')" by auto
- then show ?case using Cond.prems(2) by fast
-next
- case (While x1 x2)
- then have "(\<forall>e'. subexp_of e' x1 \<longrightarrow> R e') \<and> R x1" and "(\<forall>e'. subexp_of e' x2 \<longrightarrow> R e') \<and> R x2"
- by fast+
- then have "(\<forall>e'. subexp_of e' (while (x1) x2) \<longrightarrow> R e')" by auto
- then show ?case using While.prems(2) by fast
-next
- case (throw x)
- then have "(\<forall>e'. subexp_of e' x \<longrightarrow> R e') \<and> R x" by fast
- then have "(\<forall>e'. subexp_of e' (throw x) \<longrightarrow> R e')" by auto
- then show ?case using throw.prems(2) by fast
-next
- case (TryCatch x1 x2 x3 x4)
- then have "(\<forall>e'. subexp_of e' x1 \<longrightarrow> R e') \<and> R x1" and "(\<forall>e'. subexp_of e' x4 \<longrightarrow> R e') \<and> R x4"
- by fast+
- then have "(\<forall>e'. subexp_of e' (try x1 catch(x2 x3) x4) \<longrightarrow> R e')" by auto
- then show ?case using TryCatch.prems(2) by fast
-next
- case (INIT x1 x2 x3 x4)
- then have "(\<forall>e'. subexp_of e' x4 \<longrightarrow> R e') \<and> R x4" by fast
- then have "(\<forall>e'. subexp_of e' (INIT x1 (x2,x3) \<leftarrow> x4) \<longrightarrow> R e')" by auto
- then show ?case using INIT.prems(2) by fast
-next
- case (RI x1 x2 x3 x4)
- then have "(\<forall>e'. subexp_of e' x2 \<longrightarrow> R e') \<and> R x2" and "(\<forall>e'. subexp_of e' x4 \<longrightarrow> R e') \<and> R x4"
- by fast+
- then have "(\<forall>e'. subexp_of e' (RI (x1,x2) ; x3 \<leftarrow> x4) \<longrightarrow> R e')" by auto
- then show ?case using RI.prems(2) by fast
-next
- case (Cons_exp x1 x2)
- then have "(\<forall>e'. subexp_of e' x1 \<longrightarrow> R e') \<and> R x1" and "(\<forall>e'. e' \<in> subexps x2 \<longrightarrow> R e') \<and> Rs x2"
- by fast+
- then have "(\<forall>e'. e' \<in> subexps (x1 # x2) \<longrightarrow> R e')" using subexps_def2 by auto
- then show ?case using Cons_exp.prems(3) by fast
-qed(auto)
-
-
-subsection"Final expressions"
-(* these definitions and most of the lemmas were in BigStep.thy in the original Jinja *)
-
-definition final :: "'a exp \<Rightarrow> bool"
-where
- "final e \<equiv> (\<exists>v. e = Val v) \<or> (\<exists>a. e = Throw a)"
-
-definition finals:: "'a exp list \<Rightarrow> bool"
-where
- "finals es \<equiv> (\<exists>vs. es = map Val vs) \<or> (\<exists>vs a es'. es = map Val vs @ Throw a # es')"
-
-lemma [simp]: "final(Val v)"
-(*<*)by(simp add:final_def)(*>*)
-
-lemma [simp]: "final(throw e) = (\<exists>a. e = addr a)"
-(*<*)by(simp add:final_def)(*>*)
-
-lemma finalE: "\<lbrakk> final e; \<And>v. e = Val v \<Longrightarrow> R; \<And>a. e = Throw a \<Longrightarrow> R \<rbrakk> \<Longrightarrow> R"
-(*<*)by(auto simp:final_def)(*>*)
-
-lemma final_fv[iff]: "final e \<Longrightarrow> fv e = {}"
- by (auto simp: final_def)
-
-lemma finalsE:
- "\<lbrakk> finals es; \<And>vs. es = map Val vs \<Longrightarrow> R; \<And>vs a es'. es = map Val vs @ Throw a # es' \<Longrightarrow> R \<rbrakk> \<Longrightarrow> R"
-(*<*)by(auto simp:finals_def)(*>*)
-
-lemma [iff]: "finals []"
-(*<*)by(simp add:finals_def)(*>*)
-
-lemma [iff]: "finals (Val v # es) = finals es"
-(*<*)
-proof(rule iffI)
- assume "finals (Val v # es)"
- moreover {
- fix vs a es'
- assume "\<forall>vs a es'. es \<noteq> map Val vs @ Throw a # es'"
- and "Val v # es = map Val vs @ Throw a # es'"
- then have "\<exists>vs. es = map Val vs" by(case_tac vs; simp)
- }
- ultimately show "finals es" by(clarsimp simp add: finals_def)
-next
- assume "finals es"
- moreover {
- fix vs a es'
- assume "es = map Val vs @ Throw a # es'"
- then have "\<exists>vs' a' es''. Val v # map Val vs @ Throw a # es' = map Val vs' @ Throw a' # es''"
- by(rule_tac x = "v#vs" in exI) simp
- }
- ultimately show "finals (Val v # es)" by(clarsimp simp add: finals_def)
-qed
-(*>*)
-
-lemma finals_app_map[iff]: "finals (map Val vs @ es) = finals es"
-(*<*)by(induct_tac vs, auto)(*>*)
-
-lemma [iff]: "finals (map Val vs)"
-(*<*)using finals_app_map[of vs "[]"]by(simp)(*>*)
-
-lemma [iff]: "finals (throw e # es) = (\<exists>a. e = addr a)"
-(*<*)
-proof(rule iffI)
- assume "finals (throw e # es)"
- moreover {
- fix vs a es'
- assume "throw e # es = map Val vs @ Throw a # es'"
- then have "\<exists>a. e = addr a" by(case_tac vs; simp)
- }
- ultimately show "\<exists>a. e = addr a" by(clarsimp simp add: finals_def)
-next
- assume "\<exists>a. e = addr a"
- moreover {
- fix vs a es'
- assume "e = addr a"
- then have "\<exists>vs aa es'. Throw a # es = map Val vs @ Throw aa # es'"
- by(rule_tac x = "[]" in exI) simp
- }
- ultimately show "finals (throw e # es)" by(clarsimp simp add: finals_def)
-qed
-(*>*)
-
-lemma not_finals_ConsI: "\<not> final e \<Longrightarrow> \<not> finals(e#es)"
-(*<*)
-proof -
- assume "\<not> final e"
- moreover {
- fix vs a es'
- assume "\<forall>v. e \<noteq> Val v" and "\<forall>a. e \<noteq> Throw a"
- then have "e # es \<noteq> map Val vs @ Throw a # es'" by(case_tac vs; simp)
- }
- ultimately show ?thesis by(clarsimp simp add:finals_def final_def)
-qed
-(*>*)
-
-lemma not_finals_ConsI2: "e = Val v \<Longrightarrow> \<not> finals es \<Longrightarrow> \<not> finals(e#es)"
-(*<*)
-proof -
- assume [simp]: "e = Val v" and "\<not> finals es"
- moreover {
- fix vs a es'
- assume "\<forall>vs. es \<noteq> map Val vs" and "\<forall>vs a es'. es \<noteq> map Val vs @ Throw a # es'"
- then have "e # es \<noteq> map Val vs @ Throw a # es'" by(case_tac vs; simp)
- }
- ultimately show ?thesis by(clarsimp simp add:finals_def final_def)
-qed
-(*>*)
-
-
-end
+(* Title: JinjaDCI/J/Expr.thy
+ Author: Tobias Nipkow, Susannah Mansky
+ Copyright 2003 Technische Universitaet Muenchen, 2019-20 UIUC
+
+ Based on the Jinja theory J/Expr.thy by Tobias Nipkow
+*)
+
+section \<open> Expressions \<close>
+
+theory Expr
+imports "../Common/Exceptions"
+begin
+
+datatype bop = Eq | Add \<comment> \<open>names of binary operations\<close>
+
+datatype 'a exp
+ = new cname \<comment> \<open>class instance creation\<close>
+ | Cast cname "('a exp)" \<comment> \<open>type cast\<close>
+ | Val val \<comment> \<open>value\<close>
+ | BinOp "('a exp)" bop "('a exp)" ("_ \<guillemotleft>_\<guillemotright> _" [80,0,81] 80) \<comment> \<open>binary operation\<close>
+ | Var 'a \<comment> \<open>local variable (incl. parameter)\<close>
+ | LAss 'a "('a exp)" ("_:=_" [90,90]90) \<comment> \<open>local assignment\<close>
+ | FAcc "('a exp)" vname cname ("_\<bullet>_{_}" [10,90,99]90) \<comment> \<open>field access\<close>
+ | SFAcc cname vname cname ("_\<bullet>\<^sub>s_{_}" [10,90,99]90) \<comment> \<open>static field access\<close>
+ | FAss "('a exp)" vname cname "('a exp)" ("_\<bullet>_{_} := _" [10,90,99,90]90) \<comment> \<open>field assignment\<close>
+ | SFAss cname vname cname "('a exp)" ("_\<bullet>\<^sub>s_{_} := _" [10,90,99,90]90) \<comment> \<open>static field assignment\<close>
+ | Call "('a exp)" mname "('a exp list)" ("_\<bullet>_'(_')" [90,99,0] 90) \<comment> \<open>method call\<close>
+ | SCall cname mname "('a exp list)" ("_\<bullet>\<^sub>s_'(_')" [90,99,0] 90) \<comment> \<open>static method call\<close>
+ | Block 'a ty "('a exp)" ("'{_:_; _}")
+ | Seq "('a exp)" "('a exp)" ("_;;/ _" [61,60]60)
+ | Cond "('a exp)" "('a exp)" "('a exp)" ("if '(_') _/ else _" [80,79,79]70)
+ | While "('a exp)" "('a exp)" ("while '(_') _" [80,79]70)
+ | throw "('a exp)"
+ | TryCatch "('a exp)" cname 'a "('a exp)" ("try _/ catch'(_ _') _" [0,99,80,79] 70)
+ | INIT cname "cname list" bool "('a exp)" ("INIT _ '(_,_') \<leftarrow> _" [60,60,60,60] 60) \<comment> \<open>internal initialization command: class, list of superclasses to initialize, preparation flag; command on hold\<close>
+ | RI cname "('a exp)" "cname list" "('a exp)" ("RI '(_,_') ; _ \<leftarrow> _" [60,60,60,60] 60) \<comment> \<open>running of the initialization procedure for class with expression, classes still to initialize command on hold\<close>
+
+type_synonym
+ expr = "vname exp" \<comment> \<open>Jinja expression\<close>
+type_synonym
+ J_mb = "vname list \<times> expr" \<comment> \<open>Jinja method body: parameter names and expression\<close>
+type_synonym
+ J_prog = "J_mb prog" \<comment> \<open>Jinja program\<close>
+
+type_synonym
+ init_stack = "expr list \<times> bool" \<comment> \<open>Stack of expressions waiting on initialization in small step; indicator boolean True if current expression has been init checked\<close>
+
+text\<open>The semantics of binary operators: \<close>
+
+fun binop :: "bop \<times> val \<times> val \<Rightarrow> val option" where
+ "binop(Eq,v\<^sub>1,v\<^sub>2) = Some(Bool (v\<^sub>1 = v\<^sub>2))"
+| "binop(Add,Intg i\<^sub>1,Intg i\<^sub>2) = Some(Intg(i\<^sub>1+i\<^sub>2))"
+| "binop(bop,v\<^sub>1,v\<^sub>2) = None"
+
+lemma [simp]:
+ "(binop(Add,v\<^sub>1,v\<^sub>2) = Some v) = (\<exists>i\<^sub>1 i\<^sub>2. v\<^sub>1 = Intg i\<^sub>1 \<and> v\<^sub>2 = Intg i\<^sub>2 \<and> v = Intg(i\<^sub>1+i\<^sub>2))"
+(*<*)by(cases v\<^sub>1; cases v\<^sub>2) auto(*>*)
+
+
+lemma map_Val_throw_eq:
+ "map Val vs @ throw ex # es = map Val vs' @ throw ex' # es' \<Longrightarrow> ex = ex'"
+(*<*)by(induct vs arbitrary: vs'; case_tac vs') auto(*>*)
+
+lemma map_Val_nthrow_neq:
+ "map Val vs = map Val vs' @ throw ex' # es' \<Longrightarrow> False"
+(*<*)by(induct vs arbitrary: vs'; case_tac vs') auto(*>*)
+
+lemma map_Val_eq:
+ "map Val vs = map Val vs' \<Longrightarrow> vs = vs'"
+(*<*)by(induct vs arbitrary: vs'; case_tac vs') auto(*>*)
+
+
+lemma init_rhs_neq [simp]: "e \<noteq> INIT C (Cs,b) \<leftarrow> e"
+proof -
+ have "size e \<noteq> size (INIT C (Cs,b) \<leftarrow> e)" by auto
+ then show ?thesis by fastforce
+qed
+
+lemma init_rhs_neq' [simp]: "INIT C (Cs,b) \<leftarrow> e \<noteq> e"
+proof -
+ have "size e \<noteq> size (INIT C (Cs,b) \<leftarrow> e)" by auto
+ then show ?thesis by fastforce
+qed
+
+lemma ri_rhs_neq [simp]: "e \<noteq> RI(C,e');Cs \<leftarrow> e"
+proof -
+ have "size e \<noteq> size (RI(C,e');Cs \<leftarrow> e)" by auto
+ then show ?thesis by fastforce
+qed
+
+lemma ri_rhs_neq' [simp]: "RI(C,e');Cs \<leftarrow> e \<noteq> e"
+proof -
+ have "size e \<noteq> size (RI(C,e');Cs \<leftarrow> e)" by auto
+ then show ?thesis by fastforce
+qed
+
+subsection "Syntactic sugar"
+
+abbreviation (input)
+ InitBlock:: "'a \<Rightarrow> ty \<Rightarrow> 'a exp \<Rightarrow> 'a exp \<Rightarrow> 'a exp" ("(1'{_:_ := _;/ _})") where
+ "InitBlock V T e1 e2 == {V:T; V := e1;; e2}"
+
+abbreviation unit where "unit == Val Unit"
+abbreviation null where "null == Val Null"
+abbreviation "addr a == Val(Addr a)"
+abbreviation "true == Val(Bool True)"
+abbreviation "false == Val(Bool False)"
+
+abbreviation
+ Throw :: "addr \<Rightarrow> 'a exp" where
+ "Throw a == throw(Val(Addr a))"
+
+abbreviation
+ THROW :: "cname \<Rightarrow> 'a exp" where
+ "THROW xc == Throw(addr_of_sys_xcpt xc)"
+
+
+subsection\<open>Free Variables\<close>
+
+primrec fv :: "expr \<Rightarrow> vname set" and fvs :: "expr list \<Rightarrow> vname set" where
+ "fv(new C) = {}"
+| "fv(Cast C e) = fv e"
+| "fv(Val v) = {}"
+| "fv(e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2) = fv e\<^sub>1 \<union> fv e\<^sub>2"
+| "fv(Var V) = {V}"
+| "fv(LAss V e) = {V} \<union> fv e"
+| "fv(e\<bullet>F{D}) = fv e"
+| "fv(C\<bullet>\<^sub>sF{D}) = {}"
+| "fv(e\<^sub>1\<bullet>F{D}:=e\<^sub>2) = fv e\<^sub>1 \<union> fv e\<^sub>2"
+| "fv(C\<bullet>\<^sub>sF{D}:=e\<^sub>2) = fv e\<^sub>2"
+| "fv(e\<bullet>M(es)) = fv e \<union> fvs es"
+| "fv(C\<bullet>\<^sub>sM(es)) = fvs es"
+| "fv({V:T; e}) = fv e - {V}"
+| "fv(e\<^sub>1;;e\<^sub>2) = fv e\<^sub>1 \<union> fv e\<^sub>2"
+| "fv(if (b) e\<^sub>1 else e\<^sub>2) = fv b \<union> fv e\<^sub>1 \<union> fv e\<^sub>2"
+| "fv(while (b) e) = fv b \<union> fv e"
+| "fv(throw e) = fv e"
+| "fv(try e\<^sub>1 catch(C V) e\<^sub>2) = fv e\<^sub>1 \<union> (fv e\<^sub>2 - {V})"
+| "fv(INIT C (Cs,b) \<leftarrow> e) = fv e"
+| "fv(RI (C,e);Cs \<leftarrow> e') = fv e \<union> fv e'"
+| "fvs([]) = {}"
+| "fvs(e#es) = fv e \<union> fvs es"
+
+lemma [simp]: "fvs(es\<^sub>1 @ es\<^sub>2) = fvs es\<^sub>1 \<union> fvs es\<^sub>2"
+(*<*)by (induct es\<^sub>1 type:list) auto(*>*)
+
+lemma [simp]: "fvs(map Val vs) = {}"
+(*<*)by (induct vs) auto(*>*)
+
+
+subsection\<open>Accessing expression constructor arguments\<close>
+
+fun val_of :: "'a exp \<Rightarrow> val option" where
+"val_of (Val v) = Some v" |
+"val_of _ = None"
+
+lemma val_of_spec: "val_of e = Some v \<Longrightarrow> e = Val v"
+proof(cases e) qed(auto)
+
+fun lass_val_of :: "'a exp \<Rightarrow> ('a \<times> val) option" where
+"lass_val_of (V:=Val v) = Some (V, v)" |
+"lass_val_of _ = None"
+
+lemma lass_val_of_spec:
+assumes "lass_val_of e = \<lfloor>a\<rfloor>"
+shows "e = (fst a:=Val (snd a))"
+using assms proof(cases e)
+ case (LAss V e') then show ?thesis using assms proof(cases e')qed(auto)
+qed(auto)
+
+fun map_vals_of :: "'a exp list \<Rightarrow> val list option" where
+"map_vals_of (e#es) = (case val_of e of Some v \<Rightarrow> (case map_vals_of es of Some vs \<Rightarrow> Some (v#vs)
+ | _ \<Rightarrow> None)
+ | _ \<Rightarrow> None)" |
+"map_vals_of [] = Some []"
+
+lemma map_vals_of_spec: "map_vals_of es = Some vs \<Longrightarrow> es = map Val vs"
+proof(induct es arbitrary: vs) qed(auto simp: val_of_spec)
+
+lemma map_vals_of_Vals[simp]: "map_vals_of (map Val vs) = \<lfloor>vs\<rfloor>" by(induct vs, auto)
+
+lemma map_vals_of_throw[simp]:
+ "map_vals_of (map Val vs @ throw e # es') = None"
+ by(induct vs, auto)
+
+
+fun bool_of :: "'a exp \<Rightarrow> bool option" where
+"bool_of true = Some True" |
+"bool_of false = Some False" |
+"bool_of _ = None"
+
+lemma bool_of_specT:
+assumes "bool_of e = Some True" shows "e = true"
+proof -
+ have "bool_of e = Some True" by fact
+ then show ?thesis
+ proof(cases e)
+ case (Val x3) with assms show ?thesis
+ proof(cases x3)
+ case (Bool x) with assms Val show ?thesis
+ proof(cases x)qed(auto)
+ qed(simp_all)
+ qed(auto)
+qed
+
+lemma bool_of_specF:
+assumes "bool_of e = Some False" shows "e = false"
+proof -
+ have "bool_of e = Some False" by fact
+ then show ?thesis
+ proof(cases e)
+ case (Val x3) with assms show ?thesis
+ proof(cases x3)
+ case (Bool x) with assms Val show ?thesis
+ proof(cases x)qed(auto)
+ qed(simp_all)
+ qed(auto)
+qed
+
+
+fun throw_of :: "'a exp \<Rightarrow> 'a exp option" where
+"throw_of (throw e') = Some e'" |
+"throw_of _ = None"
+
+lemma throw_of_spec: "throw_of e = Some e' \<Longrightarrow> e = throw e'"
+proof(cases e) qed(auto)
+
+fun init_exp_of :: "'a exp \<Rightarrow> 'a exp option" where
+"init_exp_of (INIT C (Cs,b) \<leftarrow> e) = Some e" |
+"init_exp_of (RI(C,e');Cs \<leftarrow> e) = Some e" |
+"init_exp_of _ = None"
+
+lemma init_exp_of_neq [simp]: "init_exp_of e = \<lfloor>e'\<rfloor> \<Longrightarrow> e' \<noteq> e" by(cases e, auto)
+lemma init_exp_of_neq'[simp]: "init_exp_of e = \<lfloor>e'\<rfloor> \<Longrightarrow> e \<noteq> e'" by(cases e, auto)
+
+
+subsection\<open>Class initialization\<close>
+
+text \<open> This section defines a few functions that return information
+ about an expression's current initialization status. \<close>
+
+ \<comment> \<open> True if expression contains @{text INIT}, @{text RI}, or a call to a static method @{term clinit} \<close>
+primrec sub_RI :: "'a exp \<Rightarrow> bool" and sub_RIs :: "'a exp list \<Rightarrow> bool" where
+ "sub_RI(new C) = False"
+| "sub_RI(Cast C e) = sub_RI e"
+| "sub_RI(Val v) = False"
+| "sub_RI(e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2) = (sub_RI e\<^sub>1 \<or> sub_RI e\<^sub>2)"
+| "sub_RI(Var V) = False"
+| "sub_RI(LAss V e) = sub_RI e"
+| "sub_RI(e\<bullet>F{D}) = sub_RI e"
+| "sub_RI(C\<bullet>\<^sub>sF{D}) = False"
+| "sub_RI(e\<^sub>1\<bullet>F{D}:=e\<^sub>2) = (sub_RI e\<^sub>1 \<or> sub_RI e\<^sub>2)"
+| "sub_RI(C\<bullet>\<^sub>sF{D}:=e\<^sub>2) = sub_RI e\<^sub>2"
+| "sub_RI(e\<bullet>M(es)) = (sub_RI e \<or> sub_RIs es)"
+| "sub_RI(C\<bullet>\<^sub>sM(es)) = (M = clinit \<or> sub_RIs es)"
+| "sub_RI({V:T; e}) = sub_RI e"
+| "sub_RI(e\<^sub>1;;e\<^sub>2) = (sub_RI e\<^sub>1 \<or> sub_RI e\<^sub>2)"
+| "sub_RI(if (b) e\<^sub>1 else e\<^sub>2) = (sub_RI b \<or> sub_RI e\<^sub>1 \<or> sub_RI e\<^sub>2)"
+| "sub_RI(while (b) e) = (sub_RI b \<or> sub_RI e)"
+| "sub_RI(throw e) = sub_RI e"
+| "sub_RI(try e\<^sub>1 catch(C V) e\<^sub>2) = (sub_RI e\<^sub>1 \<or> sub_RI e\<^sub>2)"
+| "sub_RI(INIT C (Cs,b) \<leftarrow> e) = True"
+| "sub_RI(RI (C,e);Cs \<leftarrow> e') = True"
+| "sub_RIs([]) = False"
+| "sub_RIs(e#es) = (sub_RI e \<or> sub_RIs es)"
+
+
+lemmas sub_RI_sub_RIs_induct = sub_RI.induct sub_RIs.induct
+
+lemma nsub_RIs_def[simp]:
+ "\<not>sub_RIs es \<Longrightarrow> \<forall>e \<in> set es. \<not>sub_RI e"
+ by(induct es, auto)
+
+lemma sub_RI_base:
+ "e = INIT C (Cs, b) \<leftarrow> e' \<or> e = RI(C,e\<^sub>0);Cs \<leftarrow> e' \<Longrightarrow> sub_RI e"
+ by(cases e, auto)
+
+lemma nsub_RI_Vals[simp]: "\<not>sub_RIs (map Val vs)"
+ by(induct vs, auto)
+
+lemma lass_val_of_nsub_RI: "lass_val_of e = \<lfloor>a\<rfloor> \<Longrightarrow> \<not>sub_RI e"
+ by(drule lass_val_of_spec, simp)
+
+
+ \<comment> \<open> is not currently initializing class @{text C'} (point past checking flag) \<close>
+primrec not_init :: "cname \<Rightarrow> 'a exp \<Rightarrow> bool" and not_inits :: "cname \<Rightarrow> 'a exp list \<Rightarrow> bool" where
+ "not_init C' (new C) = True"
+| "not_init C' (Cast C e) = not_init C' e"
+| "not_init C' (Val v) = True"
+| "not_init C' (e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2) = (not_init C' e\<^sub>1 \<and> not_init C' e\<^sub>2)"
+| "not_init C' (Var V) = True"
+| "not_init C' (LAss V e) = not_init C' e"
+| "not_init C' (e\<bullet>F{D}) = not_init C' e"
+| "not_init C' (C\<bullet>\<^sub>sF{D}) = True"
+| "not_init C' (e\<^sub>1\<bullet>F{D}:=e\<^sub>2) = (not_init C' e\<^sub>1 \<and> not_init C' e\<^sub>2)"
+| "not_init C' (C\<bullet>\<^sub>sF{D}:=e\<^sub>2) = not_init C' e\<^sub>2"
+| "not_init C' (e\<bullet>M(es)) = (not_init C' e \<and> not_inits C' es)"
+| "not_init C' (C\<bullet>\<^sub>sM(es)) = not_inits C' es"
+| "not_init C' ({V:T; e}) = not_init C' e"
+| "not_init C' (e\<^sub>1;;e\<^sub>2) = (not_init C' e\<^sub>1 \<and> not_init C' e\<^sub>2)"
+| "not_init C' (if (b) e\<^sub>1 else e\<^sub>2) = (not_init C' b \<and> not_init C' e\<^sub>1 \<and> not_init C' e\<^sub>2)"
+| "not_init C' (while (b) e) = (not_init C' b \<and> not_init C' e)"
+| "not_init C' (throw e) = not_init C' e"
+| "not_init C' (try e\<^sub>1 catch(C V) e\<^sub>2) = (not_init C' e\<^sub>1 \<and> not_init C' e\<^sub>2)"
+| "not_init C' (INIT C (Cs,b) \<leftarrow> e) = ((b \<longrightarrow> Cs = Nil \<or> C' \<noteq> hd Cs) \<and> C' \<notin> set(tl Cs) \<and> not_init C' e)"
+| "not_init C' (RI (C,e);Cs \<leftarrow> e') = (C' \<notin> set (C#Cs) \<and> not_init C' e \<and> not_init C' e')"
+| "not_inits C' ([]) = True"
+| "not_inits C' (e#es) = (not_init C' e \<and> not_inits C' es)"
+
+lemma not_inits_def'[simp]:
+ "not_inits C es \<Longrightarrow> \<forall>e \<in> set es. not_init C e"
+ by(induct es, auto)
+
+lemma nsub_RIs_not_inits_aux: "\<forall>e \<in> set es. \<not>sub_RI e \<longrightarrow> not_init C e
+ \<Longrightarrow> \<not>sub_RIs es \<Longrightarrow> not_inits C es"
+ by(induct es, auto)
+
+lemma nsub_RI_not_init: "\<not>sub_RI e \<Longrightarrow> not_init C e"
+proof(induct e) qed(auto intro: nsub_RIs_not_inits_aux)
+
+lemma nsub_RIs_not_inits: "\<not>sub_RIs es \<Longrightarrow> not_inits C es"
+by(rule nsub_RIs_not_inits_aux) (simp_all add: nsub_RI_not_init)
+
+subsection\<open>Subexpressions\<close>
+
+ \<comment> \<open> all strictly smaller subexpressions; does not include self \<close>
+ primrec subexp :: "'a exp \<Rightarrow> 'a exp set" and subexps :: "'a exp list \<Rightarrow> 'a exp set" where
+ "subexp(new C) = {}"
+| "subexp(Cast C e) = {e} \<union> subexp e"
+| "subexp(Val v) = {}"
+| "subexp(e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2) = {e\<^sub>1, e\<^sub>2} \<union> subexp e\<^sub>1 \<union> subexp e\<^sub>2"
+| "subexp(Var V) = {}"
+| "subexp(LAss V e) = {e} \<union> subexp e"
+| "subexp(e\<bullet>F{D}) = {e} \<union> subexp e"
+| "subexp(C\<bullet>\<^sub>sF{D}) = {}"
+| "subexp(e\<^sub>1\<bullet>F{D}:=e\<^sub>2) = {e\<^sub>1, e\<^sub>2} \<union> subexp e\<^sub>1 \<union> subexp e\<^sub>2"
+| "subexp(C\<bullet>\<^sub>sF{D}:=e\<^sub>2) = {e\<^sub>2} \<union>subexp e\<^sub>2"
+| "subexp(e\<bullet>M(es)) = {e} \<union> set es \<union> subexp e \<union> subexps es"
+| "subexp(C\<bullet>\<^sub>sM(es)) = set es \<union> subexps es"
+| "subexp({V:T; e}) = {e} \<union> subexp e"
+| "subexp(e\<^sub>1;;e\<^sub>2) = {e\<^sub>1, e\<^sub>2} \<union> subexp e\<^sub>1 \<union> subexp e\<^sub>2"
+| "subexp(if (b) e\<^sub>1 else e\<^sub>2) = {b, e\<^sub>1, e\<^sub>2} \<union> subexp b \<union> subexp e\<^sub>1 \<union> subexp e\<^sub>2"
+| "subexp(while (b) e) = {b, e} \<union> subexp b \<union> subexp e"
+| "subexp(throw e) = {e} \<union> subexp e"
+| "subexp(try e\<^sub>1 catch(C V) e\<^sub>2) = {e\<^sub>1, e\<^sub>2} \<union> subexp e\<^sub>1 \<union> subexp e\<^sub>2"
+| "subexp(INIT C (Cs,b) \<leftarrow> e) = {e} \<union> subexp e"
+| "subexp(RI (C,e);Cs \<leftarrow> e') = {e, e'} \<union> subexp e \<union> subexp e'"
+| "subexps([]) = {}"
+| "subexps(e#es) = {e} \<union> subexp e \<union> subexps es"
+
+
+lemmas subexp_subexps_induct = subexp.induct subexps.induct
+
+abbreviation subexp_of :: "'a exp \<Rightarrow> 'a exp \<Rightarrow> bool" where
+ "subexp_of e e' \<equiv> e \<in> subexp e'"
+
+lemma subexp_size_le:
+ "(e' \<in> subexp e \<longrightarrow> size e' < size e) \<and> (e' \<in> subexps es \<longrightarrow> size e' < size_list size es)"
+proof(induct rule: subexp_subexps.induct)
+ case Call:11 then show ?case using not_less_eq size_list_estimation by fastforce
+next
+ case SCall:12 then show ?case using not_less_eq size_list_estimation by fastforce
+qed(auto)
+
+lemma subexps_def2: "subexps es = set es \<union> (\<Union>e \<in> set es. subexp e)" by(induct es, auto)
+
+ \<comment> \<open> strong induction \<close>
+lemma shows subexp_induct[consumes 1]:
+"(\<And>e. subexp e = {} \<Longrightarrow> R e) \<Longrightarrow> (\<And>e. (\<And>e'. e' \<in> subexp e \<Longrightarrow> R e') \<Longrightarrow> R e)
+ \<Longrightarrow> (\<And>es. (\<And>e'. e' \<in> subexps es \<Longrightarrow> R e') \<Longrightarrow> Rs es) \<Longrightarrow> (\<forall>e'. e' \<in> subexp e \<longrightarrow> R e') \<and> R e"
+and subexps_induct[consumes 1]:
+ "(\<And>es. subexps es = {} \<Longrightarrow> Rs es) \<Longrightarrow> (\<And>e. (\<And>e'. e' \<in> subexp e \<Longrightarrow> R e') \<Longrightarrow> R e)
+ \<Longrightarrow> (\<And>es. (\<And>e'. e' \<in> subexps es \<Longrightarrow> R e') \<Longrightarrow> Rs es) \<Longrightarrow> (\<forall>e'. e' \<in> subexps es \<longrightarrow> R e') \<and> Rs es"
+proof(induct rule: subexp_subexps_induct)
+ case (Cast x1 x2)
+ then have "(\<forall>e'. subexp_of e' x2 \<longrightarrow> R e') \<and> R x2" by fast
+ then have "(\<forall>e'. subexp_of e' (Cast x1 x2) \<longrightarrow> R e')" by auto
+ then show ?case using Cast.prems(2) by fast
+next
+ case (BinOp x1 x2 x3)
+ then have "(\<forall>e'. subexp_of e' x1 \<longrightarrow> R e') \<and> R x1" and "(\<forall>e'. subexp_of e' x3 \<longrightarrow> R e') \<and> R x3"
+ by fast+
+ then have "(\<forall>e'. subexp_of e' (x1 \<guillemotleft>x2\<guillemotright> x3) \<longrightarrow> R e')" by auto
+ then show ?case using BinOp.prems(2) by fast
+next
+ case (LAss x1 x2)
+ then have "(\<forall>e'. subexp_of e' x2 \<longrightarrow> R e') \<and> R x2" by fast
+ then have "(\<forall>e'. subexp_of e' (LAss x1 x2) \<longrightarrow> R e')" by auto
+ then show ?case using LAss.prems(2) by fast
+next
+ case (FAcc x1 x2 x3)
+ then have "(\<forall>e'. subexp_of e' x1 \<longrightarrow> R e') \<and> R x1" by fast
+ then have "(\<forall>e'. subexp_of e' (x1\<bullet>x2{x3}) \<longrightarrow> R e')" by auto
+ then show ?case using FAcc.prems(2) by fast
+next
+ case (FAss x1 x2 x3 x4)
+ then have "(\<forall>e'. subexp_of e' x1 \<longrightarrow> R e') \<and> R x1" and "(\<forall>e'. subexp_of e' x4 \<longrightarrow> R e') \<and> R x4"
+ by fast+
+ then have "(\<forall>e'. subexp_of e' (x1\<bullet>x2{x3} := x4) \<longrightarrow> R e')" by auto
+ then show ?case using FAss.prems(2) by fast
+next
+ case (SFAss x1 x2 x3 x4)
+ then have "(\<forall>e'. subexp_of e' x4 \<longrightarrow> R e') \<and> R x4" by fast
+ then have "(\<forall>e'. subexp_of e' (x1\<bullet>\<^sub>sx2{x3} := x4) \<longrightarrow> R e')" by auto
+ then show ?case using SFAss.prems(2) by fast
+next
+ case (Call x1 x2 x3)
+ then have "(\<forall>e'. subexp_of e' x1 \<longrightarrow> R e') \<and> R x1" and "(\<forall>e'. e' \<in> subexps x3 \<longrightarrow> R e') \<and> Rs x3"
+ by fast+
+ then have "(\<forall>e'. subexp_of e' (x1\<bullet>x2(x3)) \<longrightarrow> R e')" using subexps_def2 by auto
+ then show ?case using Call.prems(2) by fast
+next
+ case (SCall x1 x2 x3)
+ then have "(\<forall>e'. e' \<in> subexps x3 \<longrightarrow> R e') \<and> Rs x3" by fast
+ then have "(\<forall>e'. subexp_of e' (x1\<bullet>\<^sub>sx2(x3)) \<longrightarrow> R e')" using subexps_def2 by auto
+ then show ?case using SCall.prems(2) by fast
+next
+ case (Block x1 x2 x3)
+ then have "(\<forall>e'. subexp_of e' x3 \<longrightarrow> R e') \<and> R x3" by fast
+ then have "(\<forall>e'. subexp_of e' {x1:x2; x3} \<longrightarrow> R e')" by auto
+ then show ?case using Block.prems(2) by fast
+next
+ case (Seq x1 x2)
+ then have "(\<forall>e'. subexp_of e' x1 \<longrightarrow> R e') \<and> R x1" and "(\<forall>e'. subexp_of e' x2 \<longrightarrow> R e') \<and> R x2"
+ by fast+
+ then have "(\<forall>e'. subexp_of e' (x1;; x2) \<longrightarrow> R e')" by auto
+ then show ?case using Seq.prems(2) by fast
+next
+ case (Cond x1 x2 x3)
+ then have "(\<forall>e'. subexp_of e' x1 \<longrightarrow> R e') \<and> R x1" and "(\<forall>e'. subexp_of e' x2 \<longrightarrow> R e') \<and> R x2"
+ and "(\<forall>e'. subexp_of e' x3 \<longrightarrow> R e') \<and> R x3" by fast+
+ then have "(\<forall>e'. subexp_of e' (if (x1) x2 else x3) \<longrightarrow> R e')" by auto
+ then show ?case using Cond.prems(2) by fast
+next
+ case (While x1 x2)
+ then have "(\<forall>e'. subexp_of e' x1 \<longrightarrow> R e') \<and> R x1" and "(\<forall>e'. subexp_of e' x2 \<longrightarrow> R e') \<and> R x2"
+ by fast+
+ then have "(\<forall>e'. subexp_of e' (while (x1) x2) \<longrightarrow> R e')" by auto
+ then show ?case using While.prems(2) by fast
+next
+ case (throw x)
+ then have "(\<forall>e'. subexp_of e' x \<longrightarrow> R e') \<and> R x" by fast
+ then have "(\<forall>e'. subexp_of e' (throw x) \<longrightarrow> R e')" by auto
+ then show ?case using throw.prems(2) by fast
+next
+ case (TryCatch x1 x2 x3 x4)
+ then have "(\<forall>e'. subexp_of e' x1 \<longrightarrow> R e') \<and> R x1" and "(\<forall>e'. subexp_of e' x4 \<longrightarrow> R e') \<and> R x4"
+ by fast+
+ then have "(\<forall>e'. subexp_of e' (try x1 catch(x2 x3) x4) \<longrightarrow> R e')" by auto
+ then show ?case using TryCatch.prems(2) by fast
+next
+ case (INIT x1 x2 x3 x4)
+ then have "(\<forall>e'. subexp_of e' x4 \<longrightarrow> R e') \<and> R x4" by fast
+ then have "(\<forall>e'. subexp_of e' (INIT x1 (x2,x3) \<leftarrow> x4) \<longrightarrow> R e')" by auto
+ then show ?case using INIT.prems(2) by fast
+next
+ case (RI x1 x2 x3 x4)
+ then have "(\<forall>e'. subexp_of e' x2 \<longrightarrow> R e') \<and> R x2" and "(\<forall>e'. subexp_of e' x4 \<longrightarrow> R e') \<and> R x4"
+ by fast+
+ then have "(\<forall>e'. subexp_of e' (RI (x1,x2) ; x3 \<leftarrow> x4) \<longrightarrow> R e')" by auto
+ then show ?case using RI.prems(2) by fast
+next
+ case (Cons_exp x1 x2)
+ then have "(\<forall>e'. subexp_of e' x1 \<longrightarrow> R e') \<and> R x1" and "(\<forall>e'. e' \<in> subexps x2 \<longrightarrow> R e') \<and> Rs x2"
+ by fast+
+ then have "(\<forall>e'. e' \<in> subexps (x1 # x2) \<longrightarrow> R e')" using subexps_def2 by auto
+ then show ?case using Cons_exp.prems(3) by fast
+qed(auto)
+
+
+subsection"Final expressions"
+(* these definitions and most of the lemmas were in BigStep.thy in the original Jinja *)
+
+definition final :: "'a exp \<Rightarrow> bool"
+where
+ "final e \<equiv> (\<exists>v. e = Val v) \<or> (\<exists>a. e = Throw a)"
+
+definition finals:: "'a exp list \<Rightarrow> bool"
+where
+ "finals es \<equiv> (\<exists>vs. es = map Val vs) \<or> (\<exists>vs a es'. es = map Val vs @ Throw a # es')"
+
+lemma [simp]: "final(Val v)"
+(*<*)by(simp add:final_def)(*>*)
+
+lemma [simp]: "final(throw e) = (\<exists>a. e = addr a)"
+(*<*)by(simp add:final_def)(*>*)
+
+lemma finalE: "\<lbrakk> final e; \<And>v. e = Val v \<Longrightarrow> R; \<And>a. e = Throw a \<Longrightarrow> R \<rbrakk> \<Longrightarrow> R"
+(*<*)by(auto simp:final_def)(*>*)
+
+lemma final_fv[iff]: "final e \<Longrightarrow> fv e = {}"
+ by (auto simp: final_def)
+
+lemma finalsE:
+ "\<lbrakk> finals es; \<And>vs. es = map Val vs \<Longrightarrow> R; \<And>vs a es'. es = map Val vs @ Throw a # es' \<Longrightarrow> R \<rbrakk> \<Longrightarrow> R"
+(*<*)by(auto simp:finals_def)(*>*)
+
+lemma [iff]: "finals []"
+(*<*)by(simp add:finals_def)(*>*)
+
+lemma [iff]: "finals (Val v # es) = finals es"
+(*<*)
+proof(rule iffI)
+ assume "finals (Val v # es)"
+ moreover {
+ fix vs a es'
+ assume "\<forall>vs a es'. es \<noteq> map Val vs @ Throw a # es'"
+ and "Val v # es = map Val vs @ Throw a # es'"
+ then have "\<exists>vs. es = map Val vs" by(case_tac vs; simp)
+ }
+ ultimately show "finals es" by(clarsimp simp add: finals_def)
+next
+ assume "finals es"
+ moreover {
+ fix vs a es'
+ assume "es = map Val vs @ Throw a # es'"
+ then have "\<exists>vs' a' es''. Val v # map Val vs @ Throw a # es' = map Val vs' @ Throw a' # es''"
+ by(rule_tac x = "v#vs" in exI) simp
+ }
+ ultimately show "finals (Val v # es)" by(clarsimp simp add: finals_def)
+qed
+(*>*)
+
+lemma finals_app_map[iff]: "finals (map Val vs @ es) = finals es"
+(*<*)by(induct_tac vs, auto)(*>*)
+
+lemma [iff]: "finals (map Val vs)"
+(*<*)using finals_app_map[of vs "[]"]by(simp)(*>*)
+
+lemma [iff]: "finals (throw e # es) = (\<exists>a. e = addr a)"
+(*<*)
+proof(rule iffI)
+ assume "finals (throw e # es)"
+ moreover {
+ fix vs a es'
+ assume "throw e # es = map Val vs @ Throw a # es'"
+ then have "\<exists>a. e = addr a" by(case_tac vs; simp)
+ }
+ ultimately show "\<exists>a. e = addr a" by(clarsimp simp add: finals_def)
+next
+ assume "\<exists>a. e = addr a"
+ moreover {
+ fix vs a es'
+ assume "e = addr a"
+ then have "\<exists>vs aa es'. Throw a # es = map Val vs @ Throw aa # es'"
+ by(rule_tac x = "[]" in exI) simp
+ }
+ ultimately show "finals (throw e # es)" by(clarsimp simp add: finals_def)
+qed
+(*>*)
+
+lemma not_finals_ConsI: "\<not> final e \<Longrightarrow> \<not> finals(e#es)"
+(*<*)
+proof -
+ assume "\<not> final e"
+ moreover {
+ fix vs a es'
+ assume "\<forall>v. e \<noteq> Val v" and "\<forall>a. e \<noteq> Throw a"
+ then have "e # es \<noteq> map Val vs @ Throw a # es'" by(case_tac vs; simp)
+ }
+ ultimately show ?thesis by(clarsimp simp add:finals_def final_def)
+qed
+(*>*)
+
+lemma not_finals_ConsI2: "e = Val v \<Longrightarrow> \<not> finals es \<Longrightarrow> \<not> finals(e#es)"
+(*<*)
+proof -
+ assume [simp]: "e = Val v" and "\<not> finals es"
+ moreover {
+ fix vs a es'
+ assume "\<forall>vs. es \<noteq> map Val vs" and "\<forall>vs a es'. es \<noteq> map Val vs @ Throw a # es'"
+ then have "e # es \<noteq> map Val vs @ Throw a # es'" by(case_tac vs; simp)
+ }
+ ultimately show ?thesis by(clarsimp simp add:finals_def final_def)
+qed
+(*>*)
+
+
+end
diff --git a/thys/JinjaDCI/J/JWellForm.thy b/thys/JinjaDCI/J/JWellForm.thy
--- a/thys/JinjaDCI/J/JWellForm.thy
+++ b/thys/JinjaDCI/J/JWellForm.thy
@@ -1,73 +1,73 @@
-(* Title: JinjaDCI/J/JWellForm.thy
-
- Author: Tobias Nipkow, Susannah Mansky
- Copyright 2003 Technische Universitaet Muenchen, 2019-20 UIUC
-
- Based on the Jinja theory J/JWellForm.thy by Tobias Nipkow
-*)
-
-section \<open> Well-formedness Constraints \<close>
-
-theory JWellForm
-imports "../Common/WellForm" WWellForm WellType DefAss
-begin
-
-definition wf_J_mdecl :: "J_prog \<Rightarrow> cname \<Rightarrow> J_mb mdecl \<Rightarrow> bool"
-where
- "wf_J_mdecl P C \<equiv> \<lambda>(M,b,Ts,T,(pns,body)).
- length Ts = length pns \<and>
- distinct pns \<and>
- \<not>sub_RI body \<and>
- (case b of
- NonStatic \<Rightarrow> this \<notin> set pns \<and>
- (\<exists>T'. P,[this\<mapsto>Class C,pns[\<mapsto>]Ts] \<turnstile> body :: T' \<and> P \<turnstile> T' \<le> T) \<and>
- \<D> body \<lfloor>{this} \<union> set pns\<rfloor>
- | Static \<Rightarrow> (\<exists>T'. P,[pns[\<mapsto>]Ts] \<turnstile> body :: T' \<and> P \<turnstile> T' \<le> T) \<and>
- \<D> body \<lfloor>set pns\<rfloor>)"
-
-lemma wf_J_mdecl_NonStatic[simp]:
- "wf_J_mdecl P C (M,NonStatic,Ts,T,pns,body) \<equiv>
- (length Ts = length pns \<and>
- distinct pns \<and>
- \<not>sub_RI body \<and>
- this \<notin> set pns \<and>
- (\<exists>T'. P,[this\<mapsto>Class C,pns[\<mapsto>]Ts] \<turnstile> body :: T' \<and> P \<turnstile> T' \<le> T) \<and>
- \<D> body \<lfloor>{this} \<union> set pns\<rfloor>)"
-(*<*)by(simp add:wf_J_mdecl_def)(*>*)
-
-lemma wf_J_mdecl_Static[simp]:
- "wf_J_mdecl P C (M,Static,Ts,T,pns,body) \<equiv>
- (length Ts = length pns \<and>
- distinct pns \<and>
- \<not>sub_RI body \<and>
- (\<exists>T'. P,[pns[\<mapsto>]Ts] \<turnstile> body :: T' \<and> P \<turnstile> T' \<le> T) \<and>
- \<D> body \<lfloor>set pns\<rfloor>)"
-(*<*)by(simp add:wf_J_mdecl_def)(*>*)
-
-
-abbreviation
- wf_J_prog :: "J_prog \<Rightarrow> bool" where
- "wf_J_prog == wf_prog wf_J_mdecl"
-
-lemma wf_J_prog_wf_J_mdecl:
- "\<lbrakk> wf_J_prog P; (C, D, fds, mths) \<in> set P; jmdcl \<in> set mths \<rbrakk>
- \<Longrightarrow> wf_J_mdecl P C jmdcl"
-(*<*)by(fastforce simp: wf_prog_def wf_cdecl_def wf_mdecl_def)(*>*)
-
-lemma wf_mdecl_wwf_mdecl: "wf_J_mdecl P C Md \<Longrightarrow> wwf_J_mdecl P C Md"
-(*<*)
-proof -
- obtain M b Ts T pns body where "Md = (M, b, Ts, T, pns, body)" by(cases Md) simp
- then show "wf_J_mdecl P C Md \<Longrightarrow> wwf_J_mdecl P C Md"
- by(case_tac b) (fastforce simp:wwf_J_mdecl_def dest!:WT_fv)+
-qed
-(*>*)
-
-lemma wf_prog_wwf_prog: "wf_J_prog P \<Longrightarrow> wwf_J_prog P"
-(*<*)
-by (simp add:wf_prog_def wf_cdecl_def wf_mdecl_def)
- (fast intro:wf_mdecl_wwf_mdecl)
-(*>*)
-
-
-end
+(* Title: JinjaDCI/J/JWellForm.thy
+
+ Author: Tobias Nipkow, Susannah Mansky
+ Copyright 2003 Technische Universitaet Muenchen, 2019-20 UIUC
+
+ Based on the Jinja theory J/JWellForm.thy by Tobias Nipkow
+*)
+
+section \<open> Well-formedness Constraints \<close>
+
+theory JWellForm
+imports "../Common/WellForm" WWellForm WellType DefAss
+begin
+
+definition wf_J_mdecl :: "J_prog \<Rightarrow> cname \<Rightarrow> J_mb mdecl \<Rightarrow> bool"
+where
+ "wf_J_mdecl P C \<equiv> \<lambda>(M,b,Ts,T,(pns,body)).
+ length Ts = length pns \<and>
+ distinct pns \<and>
+ \<not>sub_RI body \<and>
+ (case b of
+ NonStatic \<Rightarrow> this \<notin> set pns \<and>
+ (\<exists>T'. P,[this\<mapsto>Class C,pns[\<mapsto>]Ts] \<turnstile> body :: T' \<and> P \<turnstile> T' \<le> T) \<and>
+ \<D> body \<lfloor>{this} \<union> set pns\<rfloor>
+ | Static \<Rightarrow> (\<exists>T'. P,[pns[\<mapsto>]Ts] \<turnstile> body :: T' \<and> P \<turnstile> T' \<le> T) \<and>
+ \<D> body \<lfloor>set pns\<rfloor>)"
+
+lemma wf_J_mdecl_NonStatic[simp]:
+ "wf_J_mdecl P C (M,NonStatic,Ts,T,pns,body) \<equiv>
+ (length Ts = length pns \<and>
+ distinct pns \<and>
+ \<not>sub_RI body \<and>
+ this \<notin> set pns \<and>
+ (\<exists>T'. P,[this\<mapsto>Class C,pns[\<mapsto>]Ts] \<turnstile> body :: T' \<and> P \<turnstile> T' \<le> T) \<and>
+ \<D> body \<lfloor>{this} \<union> set pns\<rfloor>)"
+(*<*)by(simp add:wf_J_mdecl_def)(*>*)
+
+lemma wf_J_mdecl_Static[simp]:
+ "wf_J_mdecl P C (M,Static,Ts,T,pns,body) \<equiv>
+ (length Ts = length pns \<and>
+ distinct pns \<and>
+ \<not>sub_RI body \<and>
+ (\<exists>T'. P,[pns[\<mapsto>]Ts] \<turnstile> body :: T' \<and> P \<turnstile> T' \<le> T) \<and>
+ \<D> body \<lfloor>set pns\<rfloor>)"
+(*<*)by(simp add:wf_J_mdecl_def)(*>*)
+
+
+abbreviation
+ wf_J_prog :: "J_prog \<Rightarrow> bool" where
+ "wf_J_prog == wf_prog wf_J_mdecl"
+
+lemma wf_J_prog_wf_J_mdecl:
+ "\<lbrakk> wf_J_prog P; (C, D, fds, mths) \<in> set P; jmdcl \<in> set mths \<rbrakk>
+ \<Longrightarrow> wf_J_mdecl P C jmdcl"
+(*<*)by(fastforce simp: wf_prog_def wf_cdecl_def wf_mdecl_def)(*>*)
+
+lemma wf_mdecl_wwf_mdecl: "wf_J_mdecl P C Md \<Longrightarrow> wwf_J_mdecl P C Md"
+(*<*)
+proof -
+ obtain M b Ts T pns body where "Md = (M, b, Ts, T, pns, body)" by(cases Md) simp
+ then show "wf_J_mdecl P C Md \<Longrightarrow> wwf_J_mdecl P C Md"
+ by(case_tac b) (fastforce simp:wwf_J_mdecl_def dest!:WT_fv)+
+qed
+(*>*)
+
+lemma wf_prog_wwf_prog: "wf_J_prog P \<Longrightarrow> wwf_J_prog P"
+(*<*)
+by (simp add:wf_prog_def wf_cdecl_def wf_mdecl_def)
+ (fast intro:wf_mdecl_wwf_mdecl)
+(*>*)
+
+
+end
diff --git a/thys/JinjaDCI/J/Progress.thy b/thys/JinjaDCI/J/Progress.thy
--- a/thys/JinjaDCI/J/Progress.thy
+++ b/thys/JinjaDCI/J/Progress.thy
@@ -1,880 +1,880 @@
-(* Title: JinjaDCI/J/Progress.thy
-
- Author: Tobias Nipkow, Susannah Mansky
- Copyright 2003 Technische Universitaet Muenchen, 2019-20 UIUC
-
- Based on the Jinja theory J/Progress.thy by Tobias Nipkow
-*)
-
-section \<open> Progress of Small Step Semantics \<close>
-
-theory Progress
-imports WellTypeRT DefAss "../Common/Conform" EConform
-begin
-
-lemma final_addrE:
- "\<lbrakk> P,E,h,sh \<turnstile> e : Class C; final e;
- \<And>a. e = addr a \<Longrightarrow> R;
- \<And>a. e = Throw a \<Longrightarrow> R \<rbrakk> \<Longrightarrow> R"
-(*<*)by(auto simp:final_def)(*>*)
-
-
-lemma finalRefE:
- "\<lbrakk> P,E,h,sh \<turnstile> e : T; is_refT T; final e;
- e = null \<Longrightarrow> R;
- \<And>a C. \<lbrakk> e = addr a; T = Class C \<rbrakk> \<Longrightarrow> R;
- \<And>a. e = Throw a \<Longrightarrow> R \<rbrakk> \<Longrightarrow> R"
-(*<*)by(auto simp:final_def is_refT_def)(*>*)
-
-
-text\<open> Derivation of new induction scheme for well typing: \<close>
-
-inductive
- WTrt' :: "[J_prog,heap,sheap,env,expr,ty] \<Rightarrow> bool"
- and WTrts' :: "[J_prog,heap,sheap,env,expr list, ty list] \<Rightarrow> bool"
- and WTrt2' :: "[J_prog,env,heap,sheap,expr,ty] \<Rightarrow> bool"
- ("_,_,_,_ \<turnstile> _ :'' _" [51,51,51,51]50)
- and WTrts2' :: "[J_prog,env,heap,sheap,expr list, ty list] \<Rightarrow> bool"
- ("_,_,_,_ \<turnstile> _ [:''] _" [51,51,51,51]50)
- for P :: J_prog and h :: heap and sh :: sheap
-where
- "P,E,h,sh \<turnstile> e :' T \<equiv> WTrt' P h sh E e T"
-| "P,E,h,sh \<turnstile> es [:'] Ts \<equiv> WTrts' P h sh E es Ts"
-
-| "is_class P C \<Longrightarrow> P,E,h,sh \<turnstile> new C :' Class C"
-| "\<lbrakk> P,E,h,sh \<turnstile> e :' T; is_refT T; is_class P C \<rbrakk>
- \<Longrightarrow> P,E,h,sh \<turnstile> Cast C e :' Class C"
-| "typeof\<^bsub>h\<^esub> v = Some T \<Longrightarrow> P,E,h,sh \<turnstile> Val v :' T"
-| "E v = Some T \<Longrightarrow> P,E,h,sh \<turnstile> Var v :' T"
-| "\<lbrakk> P,E,h,sh \<turnstile> e\<^sub>1 :' T\<^sub>1; P,E,h,sh \<turnstile> e\<^sub>2 :' T\<^sub>2 \<rbrakk>
- \<Longrightarrow> P,E,h,sh \<turnstile> e\<^sub>1 \<guillemotleft>Eq\<guillemotright> e\<^sub>2 :' Boolean"
-| "\<lbrakk> P,E,h,sh \<turnstile> e\<^sub>1 :' Integer; P,E,h,sh \<turnstile> e\<^sub>2 :' Integer \<rbrakk>
- \<Longrightarrow> P,E,h,sh \<turnstile> e\<^sub>1 \<guillemotleft>Add\<guillemotright> e\<^sub>2 :' Integer"
-| "\<lbrakk> P,E,h,sh \<turnstile> Var V :' T; P,E,h,sh \<turnstile> e :' T'; P \<turnstile> T' \<le> T \<rbrakk>
- \<Longrightarrow> P,E,h,sh \<turnstile> V:=e :' Void"
-| "\<lbrakk> P,E,h,sh \<turnstile> e :' Class C; P \<turnstile> C has F,NonStatic:T in D \<rbrakk> \<Longrightarrow> P,E,h,sh \<turnstile> e\<bullet>F{D} :' T"
-| "P,E,h,sh \<turnstile> e :' NT \<Longrightarrow> P,E,h,sh \<turnstile> e\<bullet>F{D} :' T"
-| "\<lbrakk> P \<turnstile> C has F,Static:T in D \<rbrakk> \<Longrightarrow> P,E,h,sh \<turnstile> C\<bullet>\<^sub>sF{D} :' T"
-| "\<lbrakk> P,E,h,sh \<turnstile> e\<^sub>1 :' Class C; P \<turnstile> C has F,NonStatic:T in D;
- P,E,h,sh \<turnstile> e\<^sub>2 :' T\<^sub>2; P \<turnstile> T\<^sub>2 \<le> T \<rbrakk>
- \<Longrightarrow> P,E,h,sh \<turnstile> e\<^sub>1\<bullet>F{D}:=e\<^sub>2 :' Void"
-| "\<lbrakk> P,E,h,sh \<turnstile> e\<^sub>1:'NT; P,E,h,sh \<turnstile> e\<^sub>2 :' T\<^sub>2 \<rbrakk> \<Longrightarrow> P,E,h,sh \<turnstile> e\<^sub>1\<bullet>F{D}:=e\<^sub>2 :' Void"
-| "\<lbrakk> P \<turnstile> C has F,Static:T in D;
- P,E,h,sh \<turnstile> e\<^sub>2 :' T\<^sub>2; P \<turnstile> T\<^sub>2 \<le> T \<rbrakk>
- \<Longrightarrow> P,E,h,sh \<turnstile> C\<bullet>\<^sub>sF{D}:=e\<^sub>2 :' Void"
-| "\<lbrakk> P,E,h,sh \<turnstile> e :' Class C; P \<turnstile> C sees M,NonStatic:Ts \<rightarrow> T = (pns,body) in D;
- P,E,h,sh \<turnstile> es [:'] Ts'; P \<turnstile> Ts' [\<le>] Ts \<rbrakk>
- \<Longrightarrow> P,E,h,sh \<turnstile> e\<bullet>M(es) :' T"
-| "\<lbrakk> P,E,h,sh \<turnstile> e :' NT; P,E,h,sh \<turnstile> es [:'] Ts \<rbrakk> \<Longrightarrow> P,E,h,sh \<turnstile> e\<bullet>M(es) :' T"
-| "\<lbrakk> P \<turnstile> C sees M,Static:Ts \<rightarrow> T = (pns,body) in D;
- P,E,h,sh \<turnstile> es [:'] Ts'; P \<turnstile> Ts' [\<le>] Ts;
- M = clinit \<longrightarrow> sh D = \<lfloor>(sfs,Processing)\<rfloor> \<and> es = map Val vs \<rbrakk>
- \<Longrightarrow> P,E,h,sh \<turnstile> C\<bullet>\<^sub>sM(es) :' T"
-| "P,E,h,sh \<turnstile> [] [:'] []"
-| "\<lbrakk> P,E,h,sh \<turnstile> e :' T; P,E,h,sh \<turnstile> es [:'] Ts \<rbrakk> \<Longrightarrow> P,E,h,sh \<turnstile> e#es [:'] T#Ts"
-| "\<lbrakk> typeof\<^bsub>h\<^esub> v = Some T\<^sub>1; P \<turnstile> T\<^sub>1 \<le> T; P,E(V\<mapsto>T),h,sh \<turnstile> e\<^sub>2 :' T\<^sub>2 \<rbrakk>
- \<Longrightarrow> P,E,h,sh \<turnstile> {V:T := Val v; e\<^sub>2} :' T\<^sub>2"
-| "\<lbrakk> P,E(V\<mapsto>T),h,sh \<turnstile> e :' T'; \<not> assigned V e \<rbrakk> \<Longrightarrow> P,E,h,sh \<turnstile> {V:T; e} :' T'"
-| "\<lbrakk> P,E,h,sh \<turnstile> e\<^sub>1:' T\<^sub>1; P,E,h,sh \<turnstile> e\<^sub>2:'T\<^sub>2 \<rbrakk> \<Longrightarrow> P,E,h,sh \<turnstile> e\<^sub>1;;e\<^sub>2 :' T\<^sub>2"
-| "\<lbrakk> P,E,h,sh \<turnstile> e :' Boolean; P,E,h,sh \<turnstile> e\<^sub>1:' T\<^sub>1; P,E,h,sh \<turnstile> e\<^sub>2:' T\<^sub>2;
- P \<turnstile> T\<^sub>1 \<le> T\<^sub>2 \<or> P \<turnstile> T\<^sub>2 \<le> T\<^sub>1;
- P \<turnstile> T\<^sub>1 \<le> T\<^sub>2 \<longrightarrow> T = T\<^sub>2; P \<turnstile> T\<^sub>2 \<le> T\<^sub>1 \<longrightarrow> T = T\<^sub>1 \<rbrakk>
- \<Longrightarrow> P,E,h,sh \<turnstile> if (e) e\<^sub>1 else e\<^sub>2 :' T"
-| "\<lbrakk> P,E,h,sh \<turnstile> e :' Boolean; P,E,h,sh \<turnstile> c:' T \<rbrakk>
- \<Longrightarrow> P,E,h,sh \<turnstile> while(e) c :' Void"
-| "\<lbrakk> P,E,h,sh \<turnstile> e :' T\<^sub>r; is_refT T\<^sub>r \<rbrakk> \<Longrightarrow> P,E,h,sh \<turnstile> throw e :' T"
-| "\<lbrakk> P,E,h,sh \<turnstile> e\<^sub>1 :' T\<^sub>1; P,E(V \<mapsto> Class C),h,sh \<turnstile> e\<^sub>2 :' T\<^sub>2; P \<turnstile> T\<^sub>1 \<le> T\<^sub>2 \<rbrakk>
- \<Longrightarrow> P,E,h,sh \<turnstile> try e\<^sub>1 catch(C V) e\<^sub>2 :' T\<^sub>2"
-| "\<lbrakk> P,E,h,sh \<turnstile> e :' T; \<forall>C' \<in> set (C#Cs). is_class P C'; \<not>sub_RI e;
- \<forall>C' \<in> set (tl Cs). \<exists>sfs. sh C' = \<lfloor>(sfs,Processing)\<rfloor>;
- b \<longrightarrow> (\<forall>C' \<in> set Cs. \<exists>sfs. sh C' = \<lfloor>(sfs,Processing)\<rfloor>);
- distinct Cs; supercls_lst P Cs \<rbrakk> \<Longrightarrow> P,E,h,sh \<turnstile> INIT C (Cs, b) \<leftarrow> e :' T"
-| "\<lbrakk> P,E,h,sh \<turnstile> e :' T; P,E,h,sh \<turnstile> e' :' T'; \<forall>C' \<in> set (C#Cs). is_class P C'; \<not>sub_RI e';
- \<forall>C' \<in> set (C#Cs). not_init C' e;
- \<forall>C' \<in> set Cs. \<exists>sfs. sh C' = \<lfloor>(sfs,Processing)\<rfloor>;
- \<exists>sfs. sh C = \<lfloor>(sfs, Processing)\<rfloor> \<or> (sh C = \<lfloor>(sfs, Error)\<rfloor> \<and> e = THROW NoClassDefFoundError);
- distinct (C#Cs); supercls_lst P (C#Cs) \<rbrakk>
- \<Longrightarrow> P,E,h,sh \<turnstile> RI(C, e);Cs \<leftarrow> e' :' T'"
-
-(*<*)
-lemmas WTrt'_induct = WTrt'_WTrts'.induct [split_format (complete)]
- and WTrt'_inducts = WTrt'_WTrts'.inducts [split_format (complete)]
-
-inductive_cases WTrt'_elim_cases[elim!]:
- "P,E,h,sh \<turnstile> V :=e :' T"
-(*>*)
-
-lemma [iff]: "P,E,h,sh \<turnstile> e\<^sub>1;;e\<^sub>2 :' T\<^sub>2 = (\<exists>T\<^sub>1. P,E,h,sh \<turnstile> e\<^sub>1:' T\<^sub>1 \<and> P,E,h,sh \<turnstile> e\<^sub>2:' T\<^sub>2)"
-(*<*)by(rule iffI) (auto elim: WTrt'.cases intro!:WTrt'_WTrts'.intros)(*>*)
-
-lemma [iff]: "P,E,h,sh \<turnstile> Val v :' T = (typeof\<^bsub>h\<^esub> v = Some T)"
-(*<*)by(rule iffI) (auto elim: WTrt'.cases intro!:WTrt'_WTrts'.intros)(*>*)
-
-lemma [iff]: "P,E,h,sh \<turnstile> Var v :' T = (E v = Some T)"
-(*<*)by(rule iffI) (auto elim: WTrt'.cases intro!:WTrt'_WTrts'.intros)(*>*)
-
-
-lemma wt_wt': "P,E,h,sh \<turnstile> e : T \<Longrightarrow> P,E,h,sh \<turnstile> e :' T"
-and wts_wts': "P,E,h,sh \<turnstile> es [:] Ts \<Longrightarrow> P,E,h,sh \<turnstile> es [:'] Ts"
-(*<*)
-proof(induct rule:WTrt_inducts)
- case (WTrtBlock E V T e T')
- then show ?case
- proof(cases "assigned V e")
- case True then show ?thesis using WTrtBlock.hyps(2)
- by(clarsimp simp add:fun_upd_same assigned_def WTrt'_WTrts'.intros
- simp del:fun_upd_apply)
- next
- case False then show ?thesis
- by (simp add: WTrtBlock.hyps(2) WTrt'_WTrts'.intros)
- qed
-qed (blast intro:WTrt'_WTrts'.intros)+
-(*>*)
-
-
-lemma wt'_wt: "P,E,h,sh \<turnstile> e :' T \<Longrightarrow> P,E,h,sh \<turnstile> e : T"
-and wts'_wts: "P,E,h,sh \<turnstile> es [:'] Ts \<Longrightarrow> P,E,h,sh \<turnstile> es [:] Ts"
-(*<*)
-proof(induct rule:WTrt'_inducts)
- case Block: (19 v T\<^sub>1 T E V e\<^sub>2 T\<^sub>2)
- let ?E = "E(V \<mapsto> T)"
- have "P,?E,h,sh \<turnstile> Val v : T\<^sub>1" using Block.hyps(1) by simp
- moreover have "P \<turnstile> T\<^sub>1 \<le> T" by(rule Block.hyps(2))
- ultimately have "P,?E,h,sh \<turnstile> V:=Val v : Void" using WTrtLAss by simp
- moreover have "P,?E,h,sh \<turnstile> e\<^sub>2 : T\<^sub>2" by(rule Block.hyps(4))
- ultimately have "P,?E,h,sh \<turnstile> V:=Val v;; e\<^sub>2 : T\<^sub>2" by blast
- then show ?case by simp
-qed (blast intro:WTrt_WTrts.intros)+
-(*>*)
-
-
-corollary wt'_iff_wt: "(P,E,h,sh \<turnstile> e :' T) = (P,E,h,sh \<turnstile> e : T)"
-(*<*)by(blast intro:wt_wt' wt'_wt)(*>*)
-
-
-corollary wts'_iff_wts: "(P,E,h,sh \<turnstile> es [:'] Ts) = (P,E,h,sh \<turnstile> es [:] Ts)"
-(*<*)by(blast intro:wts_wts' wts'_wts)(*>*)
-
-(*<*)
-lemmas WTrt_inducts2 = WTrt'_inducts [unfolded wt'_iff_wt wts'_iff_wts,
- case_names WTrtNew WTrtCast WTrtVal WTrtVar WTrtBinOpEq WTrtBinOpAdd WTrtLAss
- WTrtFAcc WTrtFAccNT WTrtSFAcc WTrtFAss WTrtFAssNT WTrtSFAss WTrtCall WTrtCallNT WTrtSCall
- WTrtNil WTrtCons WTrtInitBlock WTrtBlock WTrtSeq WTrtCond WTrtWhile WTrtThrow WTrtTry
- WTrtInit WTrtRI, consumes 1]
-(*>*)
-
-theorem assumes wf: "wwf_J_prog P" and hconf: "P \<turnstile> h \<surd>" and shconf: "P,h \<turnstile>\<^sub>s sh \<surd>"
-shows progress: "P,E,h,sh \<turnstile> e : T \<Longrightarrow>
- (\<And>l. \<lbrakk> \<D> e \<lfloor>dom l\<rfloor>; P,sh \<turnstile>\<^sub>b (e,b) \<surd>; \<not> final e \<rbrakk> \<Longrightarrow> \<exists>e' s' b'. P \<turnstile> \<langle>e,(h,l,sh),b\<rangle> \<rightarrow> \<langle>e',s',b'\<rangle>)"
-and "P,E,h,sh \<turnstile> es [:] Ts \<Longrightarrow>
- (\<And>l. \<lbrakk> \<D>s es \<lfloor>dom l\<rfloor>; P,sh \<turnstile>\<^sub>b (es,b) \<surd>; \<not> finals es \<rbrakk> \<Longrightarrow> \<exists>es' s' b'. P \<turnstile> \<langle>es,(h,l,sh),b\<rangle> [\<rightarrow>] \<langle>es',s',b'\<rangle>)"
-(*<*)
-proof (induct rule:WTrt_inducts2)
- case (WTrtNew C) show ?case
- proof (cases b)
- case True then show ?thesis
- proof cases
- assume "\<exists>a. h a = None"
- with assms WTrtNew True show ?thesis
- by (fastforce del:exE intro!:RedNew simp add:new_Addr_def
- elim!:wf_Fields_Ex[THEN exE])
- next
- assume "\<not>(\<exists>a. h a = None)"
- with assms WTrtNew True show ?thesis
- by(fastforce intro:RedNewFail simp:new_Addr_def)
- qed
- next
- case False then show ?thesis
- proof cases
- assume "\<exists>sfs. sh C = Some (sfs, Done)"
- with assms WTrtNew False show ?thesis
- by(fastforce intro:NewInitDoneRed simp:new_Addr_def)
- next
- assume "\<nexists>sfs. sh C = Some (sfs, Done)"
- with assms WTrtNew False show ?thesis
- by(fastforce intro:NewInitRed simp:new_Addr_def)
- qed
- qed
-next
- case (WTrtCast E e T C)
- have wte: "P,E,h,sh \<turnstile> e : T" and ref: "is_refT T"
- and IH: "\<And>l. \<lbrakk>\<D> e \<lfloor>dom l\<rfloor>; P,sh \<turnstile>\<^sub>b (e,b) \<surd>; \<not> final e\<rbrakk>
- \<Longrightarrow> \<exists>e' s' b'. P \<turnstile> \<langle>e,(h,l,sh),b\<rangle> \<rightarrow> \<langle>e',s',b'\<rangle>"
- and D: "\<D> (Cast C e) \<lfloor>dom l\<rfloor>"
- and castconf: "P,sh \<turnstile>\<^sub>b (Cast C e,b) \<surd>" by fact+
- from D have De: "\<D> e \<lfloor>dom l\<rfloor>" by auto
- have bconf: "P,sh \<turnstile>\<^sub>b (e,b) \<surd>" using castconf bconf_Cast by fast
- show ?case
- proof cases
- assume "final e"
- with wte ref show ?thesis
- proof (rule finalRefE)
- assume "e = null" thus ?case by(fastforce intro:RedCastNull)
- next
- fix D a assume A: "T = Class D" "e = addr a"
- show ?thesis
- proof cases
- assume "P \<turnstile> D \<preceq>\<^sup>* C"
- thus ?thesis using A wte by(fastforce intro:RedCast)
- next
- assume "\<not> P \<turnstile> D \<preceq>\<^sup>* C"
- thus ?thesis using A wte by(fastforce elim!:RedCastFail)
- qed
- next
- fix a assume "e = Throw a"
- thus ?thesis by(blast intro!:red_reds.CastThrow)
- qed
- next
- assume nf: "\<not> final e"
- from IH[OF De bconf nf] show ?thesis by (blast intro:CastRed)
- qed
-next
- case WTrtVal thus ?case by(simp add:final_def)
-next
- case WTrtVar thus ?case by(fastforce intro:RedVar simp:hyper_isin_def)
-next
- case (WTrtBinOpEq E e1 T1 e2 T2) show ?case
- proof cases
- assume "final e1"
- thus ?thesis
- proof (rule finalE)
- fix v1 assume eV[simp]: "e1 = Val v1"
- show ?thesis
- proof cases
- assume "final e2"
- thus ?thesis
- proof (rule finalE)
- fix v2 assume "e2 = Val v2"
- thus ?thesis using WTrtBinOpEq by(fastforce intro:RedBinOp)
- next
- fix a assume "e2 = Throw a"
- thus ?thesis using eV by(blast intro:red_reds.BinOpThrow2)
- qed
- next
- assume nf: "\<not> final e2"
- then have "P,sh \<turnstile>\<^sub>b (e2,b) \<surd>" using WTrtBinOpEq.prems(2) by(simp add:bconf_BinOp)
- with WTrtBinOpEq nf show ?thesis
- by simp (fast intro!:BinOpRed2)
- qed
- next
- fix a assume "e1 = Throw a"
- thus ?thesis by (fast intro:red_reds.BinOpThrow1)
- qed
- next
- assume nf: "\<not> final e1"
- then have e1: "val_of e1 = None" proof(cases e1)qed(auto)
- then have "P,sh \<turnstile>\<^sub>b (e1,b) \<surd>" using WTrtBinOpEq.prems(2) by(simp add:bconf_BinOp)
- with WTrtBinOpEq nf e1 show ?thesis
- by simp (fast intro:BinOpRed1)
- qed
-next
- case (WTrtBinOpAdd E e1 e2) show ?case
- proof cases
- assume "final e1"
- thus ?thesis
- proof (rule finalE)
- fix v1 assume eV[simp]: "e1 = Val v1"
- show ?thesis
- proof cases
- assume "final e2"
- thus ?thesis
- proof (rule finalE)
- fix v2 assume eV2:"e2 = Val v2"
- then obtain i1 i2 where "v1 = Intg i1 \<and> v2 = Intg i2" using WTrtBinOpAdd by clarsimp
- thus ?thesis using WTrtBinOpAdd eV eV2 by(fastforce intro:RedBinOp)
- next
- fix a assume "e2 = Throw a"
- thus ?thesis using eV by(blast intro:red_reds.BinOpThrow2)
- qed
- next
- assume nf:"\<not> final e2"
- then have "P,sh \<turnstile>\<^sub>b (e2,b) \<surd>" using WTrtBinOpAdd.prems(2) by(simp add:bconf_BinOp)
- with WTrtBinOpAdd nf show ?thesis
- by simp (fast intro!:BinOpRed2)
- qed
- next
- fix a assume "e1 = Throw a"
- thus ?thesis by(fast intro:red_reds.BinOpThrow1)
- qed
- next
- assume nf: "\<not> final e1"
- then have e1: "val_of e1 = None" proof(cases e1)qed(auto)
- then have "P,sh \<turnstile>\<^sub>b (e1,b) \<surd>" using WTrtBinOpAdd.prems(2) by(simp add:bconf_BinOp)
- with WTrtBinOpAdd nf e1 show ?thesis
- by simp (fast intro:BinOpRed1)
- qed
-next
- case (WTrtLAss E V T e T')
- then have bconf: "P,sh \<turnstile>\<^sub>b (e,b) \<surd>" using bconf_LAss by fast
- show ?case
- proof cases
- assume "final e" with WTrtLAss show ?thesis
- by(fastforce simp:final_def intro: red_reds.RedLAss red_reds.LAssThrow)
- next
- assume "\<not> final e" with WTrtLAss bconf show ?thesis
- by simp (fast intro:LAssRed)
- qed
-next
- case (WTrtFAcc E e C F T D)
- then have bconf: "P,sh \<turnstile>\<^sub>b (e,b) \<surd>" using bconf_FAcc by fast
- have wte: "P,E,h,sh \<turnstile> e : Class C"
- and field: "P \<turnstile> C has F,NonStatic:T in D" by fact+
- show ?case
- proof cases
- assume "final e"
- with wte show ?thesis
- proof (rule final_addrE)
- fix a assume e: "e = addr a"
- with wte obtain fs where hp: "h a = Some(C,fs)" by auto
- with hconf have "P,h \<turnstile> (C,fs) \<surd>" using hconf_def by fastforce
- then obtain v where "fs(F,D) = Some v" using field
- by(fastforce dest:has_fields_fun simp:oconf_def has_field_def)
- with hp e show ?thesis by (meson field red_reds.RedFAcc)
- next
- fix a assume "e = Throw a"
- thus ?thesis by(fastforce intro:red_reds.FAccThrow)
- qed
- next
- assume "\<not> final e" with WTrtFAcc bconf show ?thesis
- by(fastforce intro!:FAccRed)
- qed
-next
- case (WTrtFAccNT E e F D T)
- then have bconf: "P,sh \<turnstile>\<^sub>b (e,b) \<surd>" using bconf_FAcc by fast
- show ?case
- proof cases
- assume "final e" \<comment> \<open>@{term e} is @{term null} or @{term throw}\<close>
- with WTrtFAccNT show ?thesis
- by(fastforce simp:final_def intro: red_reds.RedFAccNull red_reds.FAccThrow)
- next
- assume "\<not> final e" \<comment> \<open>@{term e} reduces by IH\<close>
- with WTrtFAccNT bconf show ?thesis by simp (fast intro:FAccRed)
- qed
-next
-case (WTrtSFAcc C F T D E) then show ?case
- proof (cases b)
- case True
- then obtain sfs i where shD: "sh D = \<lfloor>(sfs,i)\<rfloor>"
- using bconf_def[of P sh "C\<bullet>\<^sub>sF{D}" b] WTrtSFAcc.prems(2) initPD_def by auto
- with shconf have "P,h,D \<turnstile>\<^sub>s sfs \<surd>" using shconf_def[of P h sh] by auto
- then obtain v where sfsF: "sfs F = Some v" using WTrtSFAcc.hyps
- by(unfold soconf_def) (auto dest:has_field_idemp)
- then show ?thesis using WTrtSFAcc.hyps shD sfsF True
- by(fastforce elim:RedSFAcc)
- next
- case False
- with assms WTrtSFAcc show ?thesis
- by(metis (full_types) SFAccInitDoneRed SFAccInitRed)
- qed
-next
- case (WTrtFAss E e1 C F T D e2 T2)
- have wte1: "P,E,h,sh \<turnstile> e1 : Class C" and field: "P \<turnstile> C has F,NonStatic:T in D" by fact+
- show ?case
- proof cases
- assume "final e1"
- with wte1 show ?thesis
- proof (rule final_addrE)
- fix a assume e1: "e1 = addr a"
- show ?thesis
- proof cases
- assume "final e2"
- thus ?thesis
- proof (rule finalE)
- fix v assume "e2 = Val v"
- thus ?thesis using e1 wte1 by(fastforce intro: RedFAss[OF field])
- next
- fix a assume "e2 = Throw a"
- thus ?thesis using e1 by(fastforce intro:red_reds.FAssThrow2)
- qed
- next
- assume nf: "\<not> final e2"
- then have "P,sh \<turnstile>\<^sub>b (e2,b) \<surd>" using WTrtFAss.prems(2) e1 by(simp add:bconf_FAss)
- with WTrtFAss e1 nf show ?thesis
- by simp (fast intro!:FAssRed2)
- qed
- next
- fix a assume "e1 = Throw a"
- thus ?thesis by(fastforce intro:red_reds.FAssThrow1)
- qed
- next
- assume nf: "\<not> final e1"
- then have e1: "val_of e1 = None" proof(cases e1)qed(auto)
- then have "P,sh \<turnstile>\<^sub>b (e1,b) \<surd>" using WTrtFAss.prems(2) by(simp add:bconf_FAss)
- with WTrtFAss nf e1 show ?thesis
- by simp (blast intro!:FAssRed1)
- qed
-next
- case (WTrtFAssNT E e\<^sub>1 e\<^sub>2 T\<^sub>2 F D)
- show ?case
- proof cases
- assume e1: "final e\<^sub>1" \<comment> \<open>@{term e\<^sub>1} is @{term null} or @{term throw}\<close>
- show ?thesis
- proof cases
- assume "final e\<^sub>2" \<comment> \<open>@{term e\<^sub>2} is @{term Val} or @{term throw}\<close>
- with WTrtFAssNT e1 show ?thesis
- by(fastforce simp:final_def
- intro: red_reds.RedFAssNull red_reds.FAssThrow1 red_reds.FAssThrow2)
- next
- assume nf: "\<not> final e\<^sub>2" \<comment> \<open>@{term e\<^sub>2} reduces by IH\<close>
- show ?thesis
- proof (rule finalE[OF e1])
- fix v assume ev: "e\<^sub>1 = Val v"
- then have "P,sh \<turnstile>\<^sub>b (e\<^sub>2,b) \<surd>" using WTrtFAssNT.prems(2) nf by(simp add:bconf_FAss)
- with WTrtFAssNT ev nf show ?thesis by auto (meson red_reds.FAssRed2)
- next
- fix a assume et: "e\<^sub>1 = Throw a"
- then have "P,sh \<turnstile>\<^sub>b (e\<^sub>1,b) \<surd>" using WTrtFAssNT.prems(2) nf by(simp add:bconf_FAss)
- with WTrtFAssNT et nf show ?thesis by(fastforce intro: red_reds.FAssThrow1)
- qed
- qed
- next
- assume nf: "\<not> final e\<^sub>1" \<comment> \<open>@{term e\<^sub>1} reduces by IH\<close>
- then have e1: "val_of e\<^sub>1 = None" proof(cases e\<^sub>1)qed(auto)
- then have "P,sh \<turnstile>\<^sub>b (e\<^sub>1,b) \<surd>" using WTrtFAssNT.prems(2) by(simp add:bconf_FAss)
- with WTrtFAssNT nf e1 show ?thesis
- by simp (blast intro!:FAssRed1)
- qed
-next
- case (WTrtSFAss C F T D E e2 T\<^sub>2)
- have field: "P \<turnstile> C has F,Static:T in D" by fact+
- show ?case
- proof cases
- assume "final e2"
- thus ?thesis
- proof (rule finalE)
- fix v assume ev: "e2 = Val v"
- then show ?case
- proof (cases b)
- case True
- then obtain sfs i where shD: "sh D = \<lfloor>(sfs,i)\<rfloor>"
- using bconf_def[of P _ "C\<bullet>\<^sub>sF{D} := e2"] WTrtSFAss.prems(2) initPD_def ev by auto
- with shconf have "P,h,D \<turnstile>\<^sub>s sfs \<surd>" using shconf_def[of P] by auto
- then obtain v where sfsF: "sfs F = Some v" using field
- by(unfold soconf_def) (auto dest:has_field_idemp)
- then show ?thesis using WTrtSFAss.hyps shD sfsF True ev
- by(fastforce elim:RedSFAss)
- next
- case False
- with assms WTrtSFAss ev show ?thesis
- by(metis (full_types) SFAssInitDoneRed SFAssInitRed)
- qed
- next
- fix a assume "e2 = Throw a"
- thus ?thesis by(fastforce intro:red_reds.SFAssThrow)
- qed
- next
- assume nf: "\<not> final e2"
- then have "val_of e2 = None" using final_def val_of_spec by fastforce
- then have "P,sh \<turnstile>\<^sub>b (e2,b) \<surd>" using WTrtSFAss.prems(2) by(simp add:bconf_SFAss)
- with WTrtSFAss nf show ?thesis
- by simp (fast intro!:SFAssRed)
- qed
-next
- case (WTrtCall E e C M Ts T pns body D es Ts')
- have wte: "P,E,h,sh \<turnstile> e : Class C"
- and "method": "P \<turnstile> C sees M,NonStatic:Ts\<rightarrow>T = (pns,body) in D"
- and wtes: "P,E,h,sh \<turnstile> es [:] Ts'"and sub: "P \<turnstile> Ts' [\<le>] Ts"
- and IHes: "\<And>l.
- \<lbrakk>\<D>s es \<lfloor>dom l\<rfloor>; P,sh \<turnstile>\<^sub>b (es,b) \<surd>; \<not> finals es\<rbrakk>
- \<Longrightarrow> \<exists>es' s' b'. P \<turnstile> \<langle>es,(h,l,sh),b\<rangle> [\<rightarrow>] \<langle>es',s',b'\<rangle>"
- and D: "\<D> (e\<bullet>M(es)) \<lfloor>dom l\<rfloor>" by fact+
- show ?case
- proof cases
- assume "final e"
- with wte show ?thesis
- proof (rule final_addrE)
- fix a assume e_addr: "e = addr a"
- show ?thesis
- proof cases
- assume es: "\<exists>vs. es = map Val vs"
- from wte e_addr obtain fs where ha: "h a = Some(C,fs)" by auto
- show ?thesis
- using e_addr ha "method" WTrts_same_length[OF wtes] sub es sees_wf_mdecl[OF wf "method"]
- by(fastforce intro!: RedCall simp:list_all2_iff wf_mdecl_def)
- next
- assume "\<not>(\<exists>vs. es = map Val vs)"
- hence not_all_Val: "\<not>(\<forall>e \<in> set es. \<exists>v. e = Val v)"
- by(simp add:ex_map_conv)
- let ?ves = "takeWhile (\<lambda>e. \<exists>v. e = Val v) es"
- let ?rest = "dropWhile (\<lambda>e. \<exists>v. e = Val v) es"
- let ?ex = "hd ?rest" let ?rst = "tl ?rest"
- from not_all_Val have nonempty: "?rest \<noteq> []" by auto
- hence es: "es = ?ves @ ?ex # ?rst" by simp
- have "\<forall>e \<in> set ?ves. \<exists>v. e = Val v" by(fastforce dest:set_takeWhileD)
- then obtain vs where ves: "?ves = map Val vs"
- using ex_map_conv by blast
- show ?thesis
- proof cases
- assume "final ?ex"
- moreover from nonempty have "\<not>(\<exists>v. ?ex = Val v)"
- by(auto simp:neq_Nil_conv simp del:dropWhile_eq_Nil_conv)
- (simp add:dropWhile_eq_Cons_conv)
- ultimately obtain b where ex_Throw: "?ex = Throw b"
- by(fast elim!:finalE)
- show ?thesis using e_addr es ex_Throw ves
- by(fastforce intro:CallThrowParams)
- next
- assume not_fin: "\<not> final ?ex"
- have "finals es = finals(?ves @ ?ex # ?rst)" using es
- by(rule arg_cong)
- also have "\<dots> = finals(?ex # ?rst)" using ves by simp
- finally have "finals es = finals(?ex # ?rst)" .
- hence fes: "\<not> finals es" using not_finals_ConsI[OF not_fin] by blast
- have "P,sh \<turnstile>\<^sub>b (es,b) \<surd>" using bconf_Call WTrtCall.prems(2)
- by (metis e_addr option.simps(5) val_of.simps(1))
- thus ?thesis using fes e_addr D IHes by(fastforce intro!:CallParams)
- qed
- qed
- next
- fix a assume "e = Throw a"
- with WTrtCall.prems show ?thesis by(fast intro!:CallThrowObj)
- qed
- next
- assume nf: "\<not> final e"
- then have e1: "val_of e = None" proof(cases e)qed(auto)
- then have "P,sh \<turnstile>\<^sub>b (e,b) \<surd>" using WTrtCall.prems(2) by(simp add:bconf_Call)
- with WTrtCall nf e1 show ?thesis by simp (blast intro!:CallObj)
- qed
-next
- case (WTrtCallNT E e es Ts M T) show ?case
- proof cases
- assume "final e"
- moreover
- { fix v assume e: "e = Val v"
- hence "e = null" using WTrtCallNT by simp
- have ?case
- proof cases
- assume "finals es"
- moreover
- { fix vs assume "es = map Val vs"
- with WTrtCallNT e have ?thesis by(fastforce intro: RedCallNull) }
- moreover
- { fix vs a es' assume "es = map Val vs @ Throw a # es'"
- with WTrtCallNT e have ?thesis by(fastforce intro: CallThrowParams) }
- ultimately show ?thesis by(fastforce simp:finals_def)
- next
- assume nf: "\<not> finals es" \<comment> \<open>@{term es} reduces by IH\<close>
- have "P,sh \<turnstile>\<^sub>b (es,b) \<surd>" using WTrtCallNT.prems(2) e by (simp add: bconf_Call)
- with WTrtCallNT e nf show ?thesis by(fastforce intro: CallParams)
- qed
- }
- moreover
- { fix a assume "e = Throw a"
- with WTrtCallNT have ?case by(fastforce intro: CallThrowObj) }
- ultimately show ?thesis by(fastforce simp:final_def)
- next
- assume nf: "\<not> final e" \<comment> \<open>@{term e} reduces by IH\<close>
- then have "val_of e = None" proof(cases e)qed(auto)
- then have "P,sh \<turnstile>\<^sub>b (e,b) \<surd>" using WTrtCallNT.prems(2) by(simp add:bconf_Call)
- with WTrtCallNT nf show ?thesis by (fastforce intro:CallObj)
- qed
-next
- case (WTrtSCall C M Ts T pns body D E es Ts' sfs vs)
- have "method": "P \<turnstile> C sees M,Static:Ts\<rightarrow>T = (pns,body) in D"
- and wtes: "P,E,h,sh \<turnstile> es [:] Ts'"and sub: "P \<turnstile> Ts' [\<le>] Ts"
- and IHes: "\<And>l.
- \<lbrakk>\<D>s es \<lfloor>dom l\<rfloor>; P,sh \<turnstile>\<^sub>b (es,b) \<surd>; \<not> finals es\<rbrakk>
- \<Longrightarrow> \<exists>es' s' b'. P \<turnstile> \<langle>es,(h,l,sh),b\<rangle> [\<rightarrow>] \<langle>es',s',b'\<rangle>"
- and clinit: "M = clinit \<longrightarrow> sh D = \<lfloor>(sfs, Processing)\<rfloor> \<and> es = map Val vs"
- and D: "\<D> (C\<bullet>\<^sub>sM(es)) \<lfloor>dom l\<rfloor>" by fact+
- show ?case
- proof cases
- assume es: "\<exists>vs. es = map Val vs"
- show ?thesis
- proof (cases b)
- case True
- then show ?thesis
- using "method" WTrts_same_length[OF wtes] sub es sees_wf_mdecl[OF wf "method"] True
- by(fastforce intro!: RedSCall simp:list_all2_iff wf_mdecl_def)
- next
- case False
- show ?thesis
- using "method" clinit WTrts_same_length[OF wtes] sub es False
- by (metis (full_types) red_reds.SCallInitDoneRed red_reds.SCallInitRed)
- qed
- next
- assume nmap: "\<not>(\<exists>vs. es = map Val vs)"
- hence not_all_Val: "\<not>(\<forall>e \<in> set es. \<exists>v. e = Val v)"
- by(simp add:ex_map_conv)
- let ?ves = "takeWhile (\<lambda>e. \<exists>v. e = Val v) es"
- let ?rest = "dropWhile (\<lambda>e. \<exists>v. e = Val v) es"
- let ?ex = "hd ?rest" let ?rst = "tl ?rest"
- from not_all_Val have nonempty: "?rest \<noteq> []" by auto
- hence es: "es = ?ves @ ?ex # ?rst" by simp
- have "\<forall>e \<in> set ?ves. \<exists>v. e = Val v" by(fastforce dest:set_takeWhileD)
- then obtain vs where ves: "?ves = map Val vs"
- using ex_map_conv by blast
- show ?thesis
- proof cases
- assume "final ?ex"
- moreover from nonempty have "\<not>(\<exists>v. ?ex = Val v)"
- by(auto simp:neq_Nil_conv simp del:dropWhile_eq_Nil_conv)
- (simp add:dropWhile_eq_Cons_conv)
- ultimately obtain b where ex_Throw: "?ex = Throw b"
- by(fast elim!:finalE)
- show ?thesis using es ex_Throw ves
- by(fastforce intro:SCallThrowParams)
- next
- assume not_fin: "\<not> final ?ex"
- have "finals es = finals(?ves @ ?ex # ?rst)" using es
- by(rule arg_cong)
- also have "\<dots> = finals(?ex # ?rst)" using ves by simp
- finally have "finals es = finals(?ex # ?rst)" .
- hence fes: "\<not> finals es" using not_finals_ConsI[OF not_fin] by blast
- have "P,sh \<turnstile>\<^sub>b (es,b) \<surd>"
- by (meson WTrtSCall.prems(2) nmap bconf_SCall map_vals_of_spec not_None_eq)
- thus ?thesis using fes D IHes by(fastforce intro!:SCallParams)
- qed
- qed
-next
- case WTrtNil thus ?case by simp
-next
- case (WTrtCons E e T es Ts)
- have IHe: "\<And>l. \<lbrakk>\<D> e \<lfloor>dom l\<rfloor>; P,sh \<turnstile>\<^sub>b (e,b) \<surd>; \<not> final e\<rbrakk>
- \<Longrightarrow> \<exists>e' s' b'. P \<turnstile> \<langle>e,(h,l,sh),b\<rangle> \<rightarrow> \<langle>e',s',b'\<rangle>"
- and IHes: "\<And>l. \<lbrakk>\<D>s es \<lfloor>dom l\<rfloor>; P,sh \<turnstile>\<^sub>b (es,b) \<surd>; \<not> finals es\<rbrakk>
- \<Longrightarrow> \<exists>es' s' b'. P \<turnstile> \<langle>es,(h,l,sh),b\<rangle> [\<rightarrow>] \<langle>es',s',b'\<rangle>"
- and D: "\<D>s (e#es) \<lfloor>dom l\<rfloor>" and not_fins: "\<not> finals(e # es)" by fact+
- have De: "\<D> e \<lfloor>dom l\<rfloor>" and Des: "\<D>s es (\<lfloor>dom l\<rfloor> \<squnion> \<A> e)"
- using D by auto
- show ?case
- proof cases
- assume "final e"
- thus ?thesis
- proof (rule finalE)
- fix v assume e: "e = Val v"
- hence Des': "\<D>s es \<lfloor>dom l\<rfloor>" using De Des by auto
- have bconfs: "P,sh \<turnstile>\<^sub>b (es,b) \<surd>" using WTrtCons.prems(2) e by(simp add: bconf_Cons)
- have not_fins_tl: "\<not> finals es" using not_fins e by simp
- show ?thesis using e IHes[OF Des' bconfs not_fins_tl]
- by (blast intro!:ListRed2)
- next
- fix a assume "e = Throw a"
- hence False using not_fins by simp
- thus ?thesis ..
- qed
- next
- assume nf:"\<not> final e"
- then have "val_of e = None" proof(cases e)qed(auto)
- then have bconf: "P,sh \<turnstile>\<^sub>b (e,b) \<surd>" using WTrtCons.prems(2) by(simp add: bconf_Cons)
- with IHe[OF De bconf nf] show ?thesis by(fast intro!:ListRed1)
- qed
-next
- case (WTrtInitBlock v T\<^sub>1 T E V e\<^sub>2 T\<^sub>2)
- have IH2: "\<And>l. \<lbrakk>\<D> e\<^sub>2 \<lfloor>dom l\<rfloor>; P,sh \<turnstile>\<^sub>b (e\<^sub>2,b) \<surd>; \<not> final e\<^sub>2\<rbrakk>
- \<Longrightarrow> \<exists>e' s' b'. P \<turnstile> \<langle>e\<^sub>2,(h,l,sh),b\<rangle> \<rightarrow> \<langle>e',s',b'\<rangle>"
- and D: "\<D> {V:T := Val v; e\<^sub>2} \<lfloor>dom l\<rfloor>" by fact+
- show ?case
- proof cases
- assume "final e\<^sub>2"
- then show ?thesis
- proof (rule finalE)
- fix v\<^sub>2 assume "e\<^sub>2 = Val v\<^sub>2"
- thus ?thesis by(fast intro:RedInitBlock)
- next
- fix a assume "e\<^sub>2 = Throw a"
- thus ?thesis by(fast intro:red_reds.InitBlockThrow)
- qed
- next
- assume not_fin2: "\<not> final e\<^sub>2"
- then have "val_of e\<^sub>2 = None" proof(cases e\<^sub>2)qed(auto)
- from D have D2: "\<D> e\<^sub>2 \<lfloor>dom(l(V\<mapsto>v))\<rfloor>" by (auto simp:hyperset_defs)
- have e2conf: "P,sh \<turnstile>\<^sub>b (e\<^sub>2,b) \<surd>" using WTrtInitBlock.prems(2) by(simp add: bconf_InitBlock)
- from IH2[OF D2 e2conf not_fin2]
- obtain h' l' sh' e' b' where red2: "P \<turnstile> \<langle>e\<^sub>2,(h, l(V\<mapsto>v),sh),b\<rangle> \<rightarrow> \<langle>e',(h', l',sh'),b'\<rangle>"
- by auto
- from red_lcl_incr[OF red2] have "V \<in> dom l'" by auto
- with red2 show ?thesis by(fastforce intro:InitBlockRed)
- qed
-next
- case (WTrtBlock E V T e T')
- have IH: "\<And>l. \<lbrakk>\<D> e \<lfloor>dom l\<rfloor>; P,sh \<turnstile>\<^sub>b (e,b) \<surd>; \<not> final e\<rbrakk>
- \<Longrightarrow> \<exists>e' s' b'. P \<turnstile> \<langle>e,(h,l,sh),b\<rangle> \<rightarrow> \<langle>e',s',b'\<rangle>"
- and unass: "\<not> assigned V e" and D: "\<D> {V:T; e} \<lfloor>dom l\<rfloor>" by fact+
- show ?case
- proof cases
- assume "final e"
- thus ?thesis
- proof (rule finalE)
- fix v assume "e = Val v" thus ?thesis by(fast intro:RedBlock)
- next
- fix a assume "e = Throw a"
- thus ?thesis by(fast intro:red_reds.BlockThrow)
- qed
- next
- assume not_fin: "\<not> final e"
- then have "val_of e = None" proof(cases e)qed(auto)
- from D have De: "\<D> e \<lfloor>dom(l(V:=None))\<rfloor>" by(simp add:hyperset_defs)
- have bconf: "P,sh \<turnstile>\<^sub>b (e,b) \<surd>" using WTrtBlock by(simp add: bconf_Block)
- from IH[OF De bconf not_fin]
- obtain h' l' sh' e' b' where red: "P \<turnstile> \<langle>e,(h,l(V:=None),sh),b\<rangle> \<rightarrow> \<langle>e',(h',l',sh'),b'\<rangle>"
- by auto
- show ?thesis
- proof (cases "l' V")
- assume "l' V = None"
- with red unass show ?thesis by(blast intro: BlockRedNone)
- next
- fix v assume "l' V = Some v"
- with red unass show ?thesis by(blast intro: BlockRedSome)
- qed
- qed
-next
- case (WTrtSeq E e1 T1 e2 T2) show ?case
- proof cases
- assume "final e1"
- thus ?thesis
- by(fast elim:finalE intro:RedSeq red_reds.SeqThrow)
- next
- assume nf: "\<not> final e1"
- then have e1: "val_of e1 = None" proof(cases e1)qed(auto)
- then show ?thesis
- proof(cases "lass_val_of e1")
- case None
- then have "P,sh \<turnstile>\<^sub>b (e1,b) \<surd>" using WTrtSeq.prems(2) e1 by(simp add: bconf_Seq)
- with WTrtSeq nf e1 None show ?thesis by simp (blast intro:SeqRed)
- next
- case (Some p)
- obtain V v where "e1 = V:=Val v" using lass_val_of_spec[OF Some] by simp
- then show ?thesis using SeqRed[OF RedLAss] by blast
- qed
- qed
-next
- case (WTrtCond E e e\<^sub>1 T\<^sub>1 e\<^sub>2 T\<^sub>2 T)
- have wt: "P,E,h,sh \<turnstile> e : Boolean" by fact
- show ?case
- proof cases
- assume "final e"
- thus ?thesis
- proof (rule finalE)
- fix v assume val: "e = Val v"
- then obtain b where v: "v = Bool b" using wt by auto
- show ?thesis
- proof (cases b)
- case True with val v show ?thesis by(fastforce intro:RedCondT simp: prod_cases3)
- next
- case False with val v show ?thesis by(fastforce intro:RedCondF simp: prod_cases3)
- qed
- next
- fix a assume "e = Throw a"
- thus ?thesis by(fast intro:red_reds.CondThrow)
- qed
- next
- assume nf: "\<not> final e"
- then have "bool_of e = None" proof(cases e)qed(auto)
- then have "P,sh \<turnstile>\<^sub>b (e,b) \<surd>" using WTrtCond.prems(2) by(simp add: bconf_Cond)
- with WTrtCond nf show ?thesis by simp (blast intro:CondRed)
- qed
-next
- case WTrtWhile show ?case by(fast intro:RedWhile)
-next
- case (WTrtThrow E e T\<^sub>r T) show ?case
- proof cases
- assume "final e" \<comment> \<open>Then @{term e} must be @{term throw} or @{term null}\<close>
- with WTrtThrow show ?thesis
- by(fastforce simp:final_def is_refT_def
- intro:red_reds.ThrowThrow red_reds.RedThrowNull)
- next
- assume nf: "\<not> final e" \<comment> \<open>Then @{term e} must reduce\<close>
- then have "val_of e = None" proof(cases e)qed(auto)
- then have "P,sh \<turnstile>\<^sub>b (e,b) \<surd>" using WTrtThrow.prems(2) by(simp add: bconf_Throw)
- with WTrtThrow nf show ?thesis by simp (blast intro:ThrowRed)
- qed
-next
- case (WTrtTry E e1 T1 V C e2 T2)
- have wt1: "P,E,h,sh \<turnstile> e1 : T1" by fact
- show ?case
- proof cases
- assume "final e1"
- thus ?thesis
- proof (rule finalE)
- fix v assume "e1 = Val v"
- thus ?thesis by(fast intro:RedTry)
- next
- fix a assume e1_Throw: "e1 = Throw a"
- with wt1 obtain D fs where ha: "h a = Some(D,fs)" by fastforce
- show ?thesis
- proof cases
- assume "P \<turnstile> D \<preceq>\<^sup>* C"
- with e1_Throw ha show ?thesis by(fastforce intro!:RedTryCatch)
- next
- assume "\<not> P \<turnstile> D \<preceq>\<^sup>* C"
- with e1_Throw ha show ?thesis by(fastforce intro!:RedTryFail)
- qed
- qed
- next
- assume nf: "\<not> final e1"
- then have "val_of e1 = None" proof(cases e1)qed(auto)
- then have "P,sh \<turnstile>\<^sub>b (e1,b) \<surd>" using WTrtTry.prems(2) by(simp add: bconf_Try)
- with WTrtTry nf show ?thesis by simp (fast intro:TryRed)
- qed
-next
- case (WTrtInit E e T\<^sub>r C Cs b) show ?case
- proof(cases Cs)
- case Nil then show ?thesis using WTrtInit by(fastforce intro!: RedInit)
- next
- case (Cons C' Cs')
- show ?thesis
- proof(cases b)
- case True then show ?thesis using Cons by(fastforce intro!: RedInitRInit)
- next
- case False
- show ?thesis
- proof(cases "sh C'")
- case None
- then show ?thesis using False Cons by(fastforce intro!: InitNoneRed)
- next
- case (Some sfsi)
- then obtain sfs i where sfsi:"sfsi = (sfs,i)" by(cases sfsi)
- show ?thesis
- proof(cases i)
- case Done
- then show ?thesis using False Some sfsi Cons by(fastforce intro!: RedInitDone)
- next
- case Processing
- then show ?thesis using False Some sfsi Cons by(fastforce intro!: RedInitProcessing)
- next
- case Error
- then show ?thesis using False Some sfsi Cons by(fastforce intro!: RedInitError)
- next
- case Prepared
- show ?thesis
- proof cases
- assume "C' = Object"
- then show ?thesis using False Some sfsi Prepared Cons by(fastforce intro: InitObjectRed)
- next
- assume "C' \<noteq> Object"
- then show ?thesis using False Some sfsi Prepared WTrtInit.hyps(3) Cons
- by(simp only: is_class_def)(fastforce intro!: InitNonObjectSuperRed)
- qed
- qed
- qed
- qed
- qed
-next
- case (WTrtRI E e T\<^sub>r e' T\<^sub>r' C Cs)
- obtain sfs i where shC: "sh C = \<lfloor>(sfs, i)\<rfloor>" using WTrtRI.hyps(9) by blast
- show ?case
- proof cases
- assume fin: "final e" then show ?thesis
- proof (rule finalE)
- fix v assume e: "e = Val v"
- then show ?thesis using shC e by(fast intro:RedRInit)
- next
- fix a assume eThrow: "e = Throw a"
- show ?thesis
- proof(cases Cs)
- case Nil then show ?thesis using eThrow shC by(fastforce intro!: RInitThrow)
- next
- case Cons then show ?thesis using eThrow shC by(fastforce intro!: RInitInitThrow)
- qed
- qed
- next
- assume nf: "\<not> final e"
- then have "val_of e = None" proof(cases e)qed(auto)
- then have "P,sh \<turnstile>\<^sub>b (e,b) \<surd>" using WTrtRI.prems(2) by(simp add: bconf_RI)
- with WTrtRI nf show ?thesis by simp (meson red_reds.RInitRed)
- qed
-qed
-(*>*)
-
-end
+(* Title: JinjaDCI/J/Progress.thy
+
+ Author: Tobias Nipkow, Susannah Mansky
+ Copyright 2003 Technische Universitaet Muenchen, 2019-20 UIUC
+
+ Based on the Jinja theory J/Progress.thy by Tobias Nipkow
+*)
+
+section \<open> Progress of Small Step Semantics \<close>
+
+theory Progress
+imports WellTypeRT DefAss "../Common/Conform" EConform
+begin
+
+lemma final_addrE:
+ "\<lbrakk> P,E,h,sh \<turnstile> e : Class C; final e;
+ \<And>a. e = addr a \<Longrightarrow> R;
+ \<And>a. e = Throw a \<Longrightarrow> R \<rbrakk> \<Longrightarrow> R"
+(*<*)by(auto simp:final_def)(*>*)
+
+
+lemma finalRefE:
+ "\<lbrakk> P,E,h,sh \<turnstile> e : T; is_refT T; final e;
+ e = null \<Longrightarrow> R;
+ \<And>a C. \<lbrakk> e = addr a; T = Class C \<rbrakk> \<Longrightarrow> R;
+ \<And>a. e = Throw a \<Longrightarrow> R \<rbrakk> \<Longrightarrow> R"
+(*<*)by(auto simp:final_def is_refT_def)(*>*)
+
+
+text\<open> Derivation of new induction scheme for well typing: \<close>
+
+inductive
+ WTrt' :: "[J_prog,heap,sheap,env,expr,ty] \<Rightarrow> bool"
+ and WTrts' :: "[J_prog,heap,sheap,env,expr list, ty list] \<Rightarrow> bool"
+ and WTrt2' :: "[J_prog,env,heap,sheap,expr,ty] \<Rightarrow> bool"
+ ("_,_,_,_ \<turnstile> _ :'' _" [51,51,51,51]50)
+ and WTrts2' :: "[J_prog,env,heap,sheap,expr list, ty list] \<Rightarrow> bool"
+ ("_,_,_,_ \<turnstile> _ [:''] _" [51,51,51,51]50)
+ for P :: J_prog and h :: heap and sh :: sheap
+where
+ "P,E,h,sh \<turnstile> e :' T \<equiv> WTrt' P h sh E e T"
+| "P,E,h,sh \<turnstile> es [:'] Ts \<equiv> WTrts' P h sh E es Ts"
+
+| "is_class P C \<Longrightarrow> P,E,h,sh \<turnstile> new C :' Class C"
+| "\<lbrakk> P,E,h,sh \<turnstile> e :' T; is_refT T; is_class P C \<rbrakk>
+ \<Longrightarrow> P,E,h,sh \<turnstile> Cast C e :' Class C"
+| "typeof\<^bsub>h\<^esub> v = Some T \<Longrightarrow> P,E,h,sh \<turnstile> Val v :' T"
+| "E v = Some T \<Longrightarrow> P,E,h,sh \<turnstile> Var v :' T"
+| "\<lbrakk> P,E,h,sh \<turnstile> e\<^sub>1 :' T\<^sub>1; P,E,h,sh \<turnstile> e\<^sub>2 :' T\<^sub>2 \<rbrakk>
+ \<Longrightarrow> P,E,h,sh \<turnstile> e\<^sub>1 \<guillemotleft>Eq\<guillemotright> e\<^sub>2 :' Boolean"
+| "\<lbrakk> P,E,h,sh \<turnstile> e\<^sub>1 :' Integer; P,E,h,sh \<turnstile> e\<^sub>2 :' Integer \<rbrakk>
+ \<Longrightarrow> P,E,h,sh \<turnstile> e\<^sub>1 \<guillemotleft>Add\<guillemotright> e\<^sub>2 :' Integer"
+| "\<lbrakk> P,E,h,sh \<turnstile> Var V :' T; P,E,h,sh \<turnstile> e :' T'; P \<turnstile> T' \<le> T \<rbrakk>
+ \<Longrightarrow> P,E,h,sh \<turnstile> V:=e :' Void"
+| "\<lbrakk> P,E,h,sh \<turnstile> e :' Class C; P \<turnstile> C has F,NonStatic:T in D \<rbrakk> \<Longrightarrow> P,E,h,sh \<turnstile> e\<bullet>F{D} :' T"
+| "P,E,h,sh \<turnstile> e :' NT \<Longrightarrow> P,E,h,sh \<turnstile> e\<bullet>F{D} :' T"
+| "\<lbrakk> P \<turnstile> C has F,Static:T in D \<rbrakk> \<Longrightarrow> P,E,h,sh \<turnstile> C\<bullet>\<^sub>sF{D} :' T"
+| "\<lbrakk> P,E,h,sh \<turnstile> e\<^sub>1 :' Class C; P \<turnstile> C has F,NonStatic:T in D;
+ P,E,h,sh \<turnstile> e\<^sub>2 :' T\<^sub>2; P \<turnstile> T\<^sub>2 \<le> T \<rbrakk>
+ \<Longrightarrow> P,E,h,sh \<turnstile> e\<^sub>1\<bullet>F{D}:=e\<^sub>2 :' Void"
+| "\<lbrakk> P,E,h,sh \<turnstile> e\<^sub>1:'NT; P,E,h,sh \<turnstile> e\<^sub>2 :' T\<^sub>2 \<rbrakk> \<Longrightarrow> P,E,h,sh \<turnstile> e\<^sub>1\<bullet>F{D}:=e\<^sub>2 :' Void"
+| "\<lbrakk> P \<turnstile> C has F,Static:T in D;
+ P,E,h,sh \<turnstile> e\<^sub>2 :' T\<^sub>2; P \<turnstile> T\<^sub>2 \<le> T \<rbrakk>
+ \<Longrightarrow> P,E,h,sh \<turnstile> C\<bullet>\<^sub>sF{D}:=e\<^sub>2 :' Void"
+| "\<lbrakk> P,E,h,sh \<turnstile> e :' Class C; P \<turnstile> C sees M,NonStatic:Ts \<rightarrow> T = (pns,body) in D;
+ P,E,h,sh \<turnstile> es [:'] Ts'; P \<turnstile> Ts' [\<le>] Ts \<rbrakk>
+ \<Longrightarrow> P,E,h,sh \<turnstile> e\<bullet>M(es) :' T"
+| "\<lbrakk> P,E,h,sh \<turnstile> e :' NT; P,E,h,sh \<turnstile> es [:'] Ts \<rbrakk> \<Longrightarrow> P,E,h,sh \<turnstile> e\<bullet>M(es) :' T"
+| "\<lbrakk> P \<turnstile> C sees M,Static:Ts \<rightarrow> T = (pns,body) in D;
+ P,E,h,sh \<turnstile> es [:'] Ts'; P \<turnstile> Ts' [\<le>] Ts;
+ M = clinit \<longrightarrow> sh D = \<lfloor>(sfs,Processing)\<rfloor> \<and> es = map Val vs \<rbrakk>
+ \<Longrightarrow> P,E,h,sh \<turnstile> C\<bullet>\<^sub>sM(es) :' T"
+| "P,E,h,sh \<turnstile> [] [:'] []"
+| "\<lbrakk> P,E,h,sh \<turnstile> e :' T; P,E,h,sh \<turnstile> es [:'] Ts \<rbrakk> \<Longrightarrow> P,E,h,sh \<turnstile> e#es [:'] T#Ts"
+| "\<lbrakk> typeof\<^bsub>h\<^esub> v = Some T\<^sub>1; P \<turnstile> T\<^sub>1 \<le> T; P,E(V\<mapsto>T),h,sh \<turnstile> e\<^sub>2 :' T\<^sub>2 \<rbrakk>
+ \<Longrightarrow> P,E,h,sh \<turnstile> {V:T := Val v; e\<^sub>2} :' T\<^sub>2"
+| "\<lbrakk> P,E(V\<mapsto>T),h,sh \<turnstile> e :' T'; \<not> assigned V e \<rbrakk> \<Longrightarrow> P,E,h,sh \<turnstile> {V:T; e} :' T'"
+| "\<lbrakk> P,E,h,sh \<turnstile> e\<^sub>1:' T\<^sub>1; P,E,h,sh \<turnstile> e\<^sub>2:'T\<^sub>2 \<rbrakk> \<Longrightarrow> P,E,h,sh \<turnstile> e\<^sub>1;;e\<^sub>2 :' T\<^sub>2"
+| "\<lbrakk> P,E,h,sh \<turnstile> e :' Boolean; P,E,h,sh \<turnstile> e\<^sub>1:' T\<^sub>1; P,E,h,sh \<turnstile> e\<^sub>2:' T\<^sub>2;
+ P \<turnstile> T\<^sub>1 \<le> T\<^sub>2 \<or> P \<turnstile> T\<^sub>2 \<le> T\<^sub>1;
+ P \<turnstile> T\<^sub>1 \<le> T\<^sub>2 \<longrightarrow> T = T\<^sub>2; P \<turnstile> T\<^sub>2 \<le> T\<^sub>1 \<longrightarrow> T = T\<^sub>1 \<rbrakk>
+ \<Longrightarrow> P,E,h,sh \<turnstile> if (e) e\<^sub>1 else e\<^sub>2 :' T"
+| "\<lbrakk> P,E,h,sh \<turnstile> e :' Boolean; P,E,h,sh \<turnstile> c:' T \<rbrakk>
+ \<Longrightarrow> P,E,h,sh \<turnstile> while(e) c :' Void"
+| "\<lbrakk> P,E,h,sh \<turnstile> e :' T\<^sub>r; is_refT T\<^sub>r \<rbrakk> \<Longrightarrow> P,E,h,sh \<turnstile> throw e :' T"
+| "\<lbrakk> P,E,h,sh \<turnstile> e\<^sub>1 :' T\<^sub>1; P,E(V \<mapsto> Class C),h,sh \<turnstile> e\<^sub>2 :' T\<^sub>2; P \<turnstile> T\<^sub>1 \<le> T\<^sub>2 \<rbrakk>
+ \<Longrightarrow> P,E,h,sh \<turnstile> try e\<^sub>1 catch(C V) e\<^sub>2 :' T\<^sub>2"
+| "\<lbrakk> P,E,h,sh \<turnstile> e :' T; \<forall>C' \<in> set (C#Cs). is_class P C'; \<not>sub_RI e;
+ \<forall>C' \<in> set (tl Cs). \<exists>sfs. sh C' = \<lfloor>(sfs,Processing)\<rfloor>;
+ b \<longrightarrow> (\<forall>C' \<in> set Cs. \<exists>sfs. sh C' = \<lfloor>(sfs,Processing)\<rfloor>);
+ distinct Cs; supercls_lst P Cs \<rbrakk> \<Longrightarrow> P,E,h,sh \<turnstile> INIT C (Cs, b) \<leftarrow> e :' T"
+| "\<lbrakk> P,E,h,sh \<turnstile> e :' T; P,E,h,sh \<turnstile> e' :' T'; \<forall>C' \<in> set (C#Cs). is_class P C'; \<not>sub_RI e';
+ \<forall>C' \<in> set (C#Cs). not_init C' e;
+ \<forall>C' \<in> set Cs. \<exists>sfs. sh C' = \<lfloor>(sfs,Processing)\<rfloor>;
+ \<exists>sfs. sh C = \<lfloor>(sfs, Processing)\<rfloor> \<or> (sh C = \<lfloor>(sfs, Error)\<rfloor> \<and> e = THROW NoClassDefFoundError);
+ distinct (C#Cs); supercls_lst P (C#Cs) \<rbrakk>
+ \<Longrightarrow> P,E,h,sh \<turnstile> RI(C, e);Cs \<leftarrow> e' :' T'"
+
+(*<*)
+lemmas WTrt'_induct = WTrt'_WTrts'.induct [split_format (complete)]
+ and WTrt'_inducts = WTrt'_WTrts'.inducts [split_format (complete)]
+
+inductive_cases WTrt'_elim_cases[elim!]:
+ "P,E,h,sh \<turnstile> V :=e :' T"
+(*>*)
+
+lemma [iff]: "P,E,h,sh \<turnstile> e\<^sub>1;;e\<^sub>2 :' T\<^sub>2 = (\<exists>T\<^sub>1. P,E,h,sh \<turnstile> e\<^sub>1:' T\<^sub>1 \<and> P,E,h,sh \<turnstile> e\<^sub>2:' T\<^sub>2)"
+(*<*)by(rule iffI) (auto elim: WTrt'.cases intro!:WTrt'_WTrts'.intros)(*>*)
+
+lemma [iff]: "P,E,h,sh \<turnstile> Val v :' T = (typeof\<^bsub>h\<^esub> v = Some T)"
+(*<*)by(rule iffI) (auto elim: WTrt'.cases intro!:WTrt'_WTrts'.intros)(*>*)
+
+lemma [iff]: "P,E,h,sh \<turnstile> Var v :' T = (E v = Some T)"
+(*<*)by(rule iffI) (auto elim: WTrt'.cases intro!:WTrt'_WTrts'.intros)(*>*)
+
+
+lemma wt_wt': "P,E,h,sh \<turnstile> e : T \<Longrightarrow> P,E,h,sh \<turnstile> e :' T"
+and wts_wts': "P,E,h,sh \<turnstile> es [:] Ts \<Longrightarrow> P,E,h,sh \<turnstile> es [:'] Ts"
+(*<*)
+proof(induct rule:WTrt_inducts)
+ case (WTrtBlock E V T e T')
+ then show ?case
+ proof(cases "assigned V e")
+ case True then show ?thesis using WTrtBlock.hyps(2)
+ by(clarsimp simp add:fun_upd_same assigned_def WTrt'_WTrts'.intros
+ simp del:fun_upd_apply)
+ next
+ case False then show ?thesis
+ by (simp add: WTrtBlock.hyps(2) WTrt'_WTrts'.intros)
+ qed
+qed (blast intro:WTrt'_WTrts'.intros)+
+(*>*)
+
+
+lemma wt'_wt: "P,E,h,sh \<turnstile> e :' T \<Longrightarrow> P,E,h,sh \<turnstile> e : T"
+and wts'_wts: "P,E,h,sh \<turnstile> es [:'] Ts \<Longrightarrow> P,E,h,sh \<turnstile> es [:] Ts"
+(*<*)
+proof(induct rule:WTrt'_inducts)
+ case Block: (19 v T\<^sub>1 T E V e\<^sub>2 T\<^sub>2)
+ let ?E = "E(V \<mapsto> T)"
+ have "P,?E,h,sh \<turnstile> Val v : T\<^sub>1" using Block.hyps(1) by simp
+ moreover have "P \<turnstile> T\<^sub>1 \<le> T" by(rule Block.hyps(2))
+ ultimately have "P,?E,h,sh \<turnstile> V:=Val v : Void" using WTrtLAss by simp
+ moreover have "P,?E,h,sh \<turnstile> e\<^sub>2 : T\<^sub>2" by(rule Block.hyps(4))
+ ultimately have "P,?E,h,sh \<turnstile> V:=Val v;; e\<^sub>2 : T\<^sub>2" by blast
+ then show ?case by simp
+qed (blast intro:WTrt_WTrts.intros)+
+(*>*)
+
+
+corollary wt'_iff_wt: "(P,E,h,sh \<turnstile> e :' T) = (P,E,h,sh \<turnstile> e : T)"
+(*<*)by(blast intro:wt_wt' wt'_wt)(*>*)
+
+
+corollary wts'_iff_wts: "(P,E,h,sh \<turnstile> es [:'] Ts) = (P,E,h,sh \<turnstile> es [:] Ts)"
+(*<*)by(blast intro:wts_wts' wts'_wts)(*>*)
+
+(*<*)
+lemmas WTrt_inducts2 = WTrt'_inducts [unfolded wt'_iff_wt wts'_iff_wts,
+ case_names WTrtNew WTrtCast WTrtVal WTrtVar WTrtBinOpEq WTrtBinOpAdd WTrtLAss
+ WTrtFAcc WTrtFAccNT WTrtSFAcc WTrtFAss WTrtFAssNT WTrtSFAss WTrtCall WTrtCallNT WTrtSCall
+ WTrtNil WTrtCons WTrtInitBlock WTrtBlock WTrtSeq WTrtCond WTrtWhile WTrtThrow WTrtTry
+ WTrtInit WTrtRI, consumes 1]
+(*>*)
+
+theorem assumes wf: "wwf_J_prog P" and hconf: "P \<turnstile> h \<surd>" and shconf: "P,h \<turnstile>\<^sub>s sh \<surd>"
+shows progress: "P,E,h,sh \<turnstile> e : T \<Longrightarrow>
+ (\<And>l. \<lbrakk> \<D> e \<lfloor>dom l\<rfloor>; P,sh \<turnstile>\<^sub>b (e,b) \<surd>; \<not> final e \<rbrakk> \<Longrightarrow> \<exists>e' s' b'. P \<turnstile> \<langle>e,(h,l,sh),b\<rangle> \<rightarrow> \<langle>e',s',b'\<rangle>)"
+and "P,E,h,sh \<turnstile> es [:] Ts \<Longrightarrow>
+ (\<And>l. \<lbrakk> \<D>s es \<lfloor>dom l\<rfloor>; P,sh \<turnstile>\<^sub>b (es,b) \<surd>; \<not> finals es \<rbrakk> \<Longrightarrow> \<exists>es' s' b'. P \<turnstile> \<langle>es,(h,l,sh),b\<rangle> [\<rightarrow>] \<langle>es',s',b'\<rangle>)"
+(*<*)
+proof (induct rule:WTrt_inducts2)
+ case (WTrtNew C) show ?case
+ proof (cases b)
+ case True then show ?thesis
+ proof cases
+ assume "\<exists>a. h a = None"
+ with assms WTrtNew True show ?thesis
+ by (fastforce del:exE intro!:RedNew simp add:new_Addr_def
+ elim!:wf_Fields_Ex[THEN exE])
+ next
+ assume "\<not>(\<exists>a. h a = None)"
+ with assms WTrtNew True show ?thesis
+ by(fastforce intro:RedNewFail simp:new_Addr_def)
+ qed
+ next
+ case False then show ?thesis
+ proof cases
+ assume "\<exists>sfs. sh C = Some (sfs, Done)"
+ with assms WTrtNew False show ?thesis
+ by(fastforce intro:NewInitDoneRed simp:new_Addr_def)
+ next
+ assume "\<nexists>sfs. sh C = Some (sfs, Done)"
+ with assms WTrtNew False show ?thesis
+ by(fastforce intro:NewInitRed simp:new_Addr_def)
+ qed
+ qed
+next
+ case (WTrtCast E e T C)
+ have wte: "P,E,h,sh \<turnstile> e : T" and ref: "is_refT T"
+ and IH: "\<And>l. \<lbrakk>\<D> e \<lfloor>dom l\<rfloor>; P,sh \<turnstile>\<^sub>b (e,b) \<surd>; \<not> final e\<rbrakk>
+ \<Longrightarrow> \<exists>e' s' b'. P \<turnstile> \<langle>e,(h,l,sh),b\<rangle> \<rightarrow> \<langle>e',s',b'\<rangle>"
+ and D: "\<D> (Cast C e) \<lfloor>dom l\<rfloor>"
+ and castconf: "P,sh \<turnstile>\<^sub>b (Cast C e,b) \<surd>" by fact+
+ from D have De: "\<D> e \<lfloor>dom l\<rfloor>" by auto
+ have bconf: "P,sh \<turnstile>\<^sub>b (e,b) \<surd>" using castconf bconf_Cast by fast
+ show ?case
+ proof cases
+ assume "final e"
+ with wte ref show ?thesis
+ proof (rule finalRefE)
+ assume "e = null" thus ?case by(fastforce intro:RedCastNull)
+ next
+ fix D a assume A: "T = Class D" "e = addr a"
+ show ?thesis
+ proof cases
+ assume "P \<turnstile> D \<preceq>\<^sup>* C"
+ thus ?thesis using A wte by(fastforce intro:RedCast)
+ next
+ assume "\<not> P \<turnstile> D \<preceq>\<^sup>* C"
+ thus ?thesis using A wte by(fastforce elim!:RedCastFail)
+ qed
+ next
+ fix a assume "e = Throw a"
+ thus ?thesis by(blast intro!:red_reds.CastThrow)
+ qed
+ next
+ assume nf: "\<not> final e"
+ from IH[OF De bconf nf] show ?thesis by (blast intro:CastRed)
+ qed
+next
+ case WTrtVal thus ?case by(simp add:final_def)
+next
+ case WTrtVar thus ?case by(fastforce intro:RedVar simp:hyper_isin_def)
+next
+ case (WTrtBinOpEq E e1 T1 e2 T2) show ?case
+ proof cases
+ assume "final e1"
+ thus ?thesis
+ proof (rule finalE)
+ fix v1 assume eV[simp]: "e1 = Val v1"
+ show ?thesis
+ proof cases
+ assume "final e2"
+ thus ?thesis
+ proof (rule finalE)
+ fix v2 assume "e2 = Val v2"
+ thus ?thesis using WTrtBinOpEq by(fastforce intro:RedBinOp)
+ next
+ fix a assume "e2 = Throw a"
+ thus ?thesis using eV by(blast intro:red_reds.BinOpThrow2)
+ qed
+ next
+ assume nf: "\<not> final e2"
+ then have "P,sh \<turnstile>\<^sub>b (e2,b) \<surd>" using WTrtBinOpEq.prems(2) by(simp add:bconf_BinOp)
+ with WTrtBinOpEq nf show ?thesis
+ by simp (fast intro!:BinOpRed2)
+ qed
+ next
+ fix a assume "e1 = Throw a"
+ thus ?thesis by (fast intro:red_reds.BinOpThrow1)
+ qed
+ next
+ assume nf: "\<not> final e1"
+ then have e1: "val_of e1 = None" proof(cases e1)qed(auto)
+ then have "P,sh \<turnstile>\<^sub>b (e1,b) \<surd>" using WTrtBinOpEq.prems(2) by(simp add:bconf_BinOp)
+ with WTrtBinOpEq nf e1 show ?thesis
+ by simp (fast intro:BinOpRed1)
+ qed
+next
+ case (WTrtBinOpAdd E e1 e2) show ?case
+ proof cases
+ assume "final e1"
+ thus ?thesis
+ proof (rule finalE)
+ fix v1 assume eV[simp]: "e1 = Val v1"
+ show ?thesis
+ proof cases
+ assume "final e2"
+ thus ?thesis
+ proof (rule finalE)
+ fix v2 assume eV2:"e2 = Val v2"
+ then obtain i1 i2 where "v1 = Intg i1 \<and> v2 = Intg i2" using WTrtBinOpAdd by clarsimp
+ thus ?thesis using WTrtBinOpAdd eV eV2 by(fastforce intro:RedBinOp)
+ next
+ fix a assume "e2 = Throw a"
+ thus ?thesis using eV by(blast intro:red_reds.BinOpThrow2)
+ qed
+ next
+ assume nf:"\<not> final e2"
+ then have "P,sh \<turnstile>\<^sub>b (e2,b) \<surd>" using WTrtBinOpAdd.prems(2) by(simp add:bconf_BinOp)
+ with WTrtBinOpAdd nf show ?thesis
+ by simp (fast intro!:BinOpRed2)
+ qed
+ next
+ fix a assume "e1 = Throw a"
+ thus ?thesis by(fast intro:red_reds.BinOpThrow1)
+ qed
+ next
+ assume nf: "\<not> final e1"
+ then have e1: "val_of e1 = None" proof(cases e1)qed(auto)
+ then have "P,sh \<turnstile>\<^sub>b (e1,b) \<surd>" using WTrtBinOpAdd.prems(2) by(simp add:bconf_BinOp)
+ with WTrtBinOpAdd nf e1 show ?thesis
+ by simp (fast intro:BinOpRed1)
+ qed
+next
+ case (WTrtLAss E V T e T')
+ then have bconf: "P,sh \<turnstile>\<^sub>b (e,b) \<surd>" using bconf_LAss by fast
+ show ?case
+ proof cases
+ assume "final e" with WTrtLAss show ?thesis
+ by(fastforce simp:final_def intro: red_reds.RedLAss red_reds.LAssThrow)
+ next
+ assume "\<not> final e" with WTrtLAss bconf show ?thesis
+ by simp (fast intro:LAssRed)
+ qed
+next
+ case (WTrtFAcc E e C F T D)
+ then have bconf: "P,sh \<turnstile>\<^sub>b (e,b) \<surd>" using bconf_FAcc by fast
+ have wte: "P,E,h,sh \<turnstile> e : Class C"
+ and field: "P \<turnstile> C has F,NonStatic:T in D" by fact+
+ show ?case
+ proof cases
+ assume "final e"
+ with wte show ?thesis
+ proof (rule final_addrE)
+ fix a assume e: "e = addr a"
+ with wte obtain fs where hp: "h a = Some(C,fs)" by auto
+ with hconf have "P,h \<turnstile> (C,fs) \<surd>" using hconf_def by fastforce
+ then obtain v where "fs(F,D) = Some v" using field
+ by(fastforce dest:has_fields_fun simp:oconf_def has_field_def)
+ with hp e show ?thesis by (meson field red_reds.RedFAcc)
+ next
+ fix a assume "e = Throw a"
+ thus ?thesis by(fastforce intro:red_reds.FAccThrow)
+ qed
+ next
+ assume "\<not> final e" with WTrtFAcc bconf show ?thesis
+ by(fastforce intro!:FAccRed)
+ qed
+next
+ case (WTrtFAccNT E e F D T)
+ then have bconf: "P,sh \<turnstile>\<^sub>b (e,b) \<surd>" using bconf_FAcc by fast
+ show ?case
+ proof cases
+ assume "final e" \<comment> \<open>@{term e} is @{term null} or @{term throw}\<close>
+ with WTrtFAccNT show ?thesis
+ by(fastforce simp:final_def intro: red_reds.RedFAccNull red_reds.FAccThrow)
+ next
+ assume "\<not> final e" \<comment> \<open>@{term e} reduces by IH\<close>
+ with WTrtFAccNT bconf show ?thesis by simp (fast intro:FAccRed)
+ qed
+next
+case (WTrtSFAcc C F T D E) then show ?case
+ proof (cases b)
+ case True
+ then obtain sfs i where shD: "sh D = \<lfloor>(sfs,i)\<rfloor>"
+ using bconf_def[of P sh "C\<bullet>\<^sub>sF{D}" b] WTrtSFAcc.prems(2) initPD_def by auto
+ with shconf have "P,h,D \<turnstile>\<^sub>s sfs \<surd>" using shconf_def[of P h sh] by auto
+ then obtain v where sfsF: "sfs F = Some v" using WTrtSFAcc.hyps
+ by(unfold soconf_def) (auto dest:has_field_idemp)
+ then show ?thesis using WTrtSFAcc.hyps shD sfsF True
+ by(fastforce elim:RedSFAcc)
+ next
+ case False
+ with assms WTrtSFAcc show ?thesis
+ by(metis (full_types) SFAccInitDoneRed SFAccInitRed)
+ qed
+next
+ case (WTrtFAss E e1 C F T D e2 T2)
+ have wte1: "P,E,h,sh \<turnstile> e1 : Class C" and field: "P \<turnstile> C has F,NonStatic:T in D" by fact+
+ show ?case
+ proof cases
+ assume "final e1"
+ with wte1 show ?thesis
+ proof (rule final_addrE)
+ fix a assume e1: "e1 = addr a"
+ show ?thesis
+ proof cases
+ assume "final e2"
+ thus ?thesis
+ proof (rule finalE)
+ fix v assume "e2 = Val v"
+ thus ?thesis using e1 wte1 by(fastforce intro: RedFAss[OF field])
+ next
+ fix a assume "e2 = Throw a"
+ thus ?thesis using e1 by(fastforce intro:red_reds.FAssThrow2)
+ qed
+ next
+ assume nf: "\<not> final e2"
+ then have "P,sh \<turnstile>\<^sub>b (e2,b) \<surd>" using WTrtFAss.prems(2) e1 by(simp add:bconf_FAss)
+ with WTrtFAss e1 nf show ?thesis
+ by simp (fast intro!:FAssRed2)
+ qed
+ next
+ fix a assume "e1 = Throw a"
+ thus ?thesis by(fastforce intro:red_reds.FAssThrow1)
+ qed
+ next
+ assume nf: "\<not> final e1"
+ then have e1: "val_of e1 = None" proof(cases e1)qed(auto)
+ then have "P,sh \<turnstile>\<^sub>b (e1,b) \<surd>" using WTrtFAss.prems(2) by(simp add:bconf_FAss)
+ with WTrtFAss nf e1 show ?thesis
+ by simp (blast intro!:FAssRed1)
+ qed
+next
+ case (WTrtFAssNT E e\<^sub>1 e\<^sub>2 T\<^sub>2 F D)
+ show ?case
+ proof cases
+ assume e1: "final e\<^sub>1" \<comment> \<open>@{term e\<^sub>1} is @{term null} or @{term throw}\<close>
+ show ?thesis
+ proof cases
+ assume "final e\<^sub>2" \<comment> \<open>@{term e\<^sub>2} is @{term Val} or @{term throw}\<close>
+ with WTrtFAssNT e1 show ?thesis
+ by(fastforce simp:final_def
+ intro: red_reds.RedFAssNull red_reds.FAssThrow1 red_reds.FAssThrow2)
+ next
+ assume nf: "\<not> final e\<^sub>2" \<comment> \<open>@{term e\<^sub>2} reduces by IH\<close>
+ show ?thesis
+ proof (rule finalE[OF e1])
+ fix v assume ev: "e\<^sub>1 = Val v"
+ then have "P,sh \<turnstile>\<^sub>b (e\<^sub>2,b) \<surd>" using WTrtFAssNT.prems(2) nf by(simp add:bconf_FAss)
+ with WTrtFAssNT ev nf show ?thesis by auto (meson red_reds.FAssRed2)
+ next
+ fix a assume et: "e\<^sub>1 = Throw a"
+ then have "P,sh \<turnstile>\<^sub>b (e\<^sub>1,b) \<surd>" using WTrtFAssNT.prems(2) nf by(simp add:bconf_FAss)
+ with WTrtFAssNT et nf show ?thesis by(fastforce intro: red_reds.FAssThrow1)
+ qed
+ qed
+ next
+ assume nf: "\<not> final e\<^sub>1" \<comment> \<open>@{term e\<^sub>1} reduces by IH\<close>
+ then have e1: "val_of e\<^sub>1 = None" proof(cases e\<^sub>1)qed(auto)
+ then have "P,sh \<turnstile>\<^sub>b (e\<^sub>1,b) \<surd>" using WTrtFAssNT.prems(2) by(simp add:bconf_FAss)
+ with WTrtFAssNT nf e1 show ?thesis
+ by simp (blast intro!:FAssRed1)
+ qed
+next
+ case (WTrtSFAss C F T D E e2 T\<^sub>2)
+ have field: "P \<turnstile> C has F,Static:T in D" by fact+
+ show ?case
+ proof cases
+ assume "final e2"
+ thus ?thesis
+ proof (rule finalE)
+ fix v assume ev: "e2 = Val v"
+ then show ?case
+ proof (cases b)
+ case True
+ then obtain sfs i where shD: "sh D = \<lfloor>(sfs,i)\<rfloor>"
+ using bconf_def[of P _ "C\<bullet>\<^sub>sF{D} := e2"] WTrtSFAss.prems(2) initPD_def ev by auto
+ with shconf have "P,h,D \<turnstile>\<^sub>s sfs \<surd>" using shconf_def[of P] by auto
+ then obtain v where sfsF: "sfs F = Some v" using field
+ by(unfold soconf_def) (auto dest:has_field_idemp)
+ then show ?thesis using WTrtSFAss.hyps shD sfsF True ev
+ by(fastforce elim:RedSFAss)
+ next
+ case False
+ with assms WTrtSFAss ev show ?thesis
+ by(metis (full_types) SFAssInitDoneRed SFAssInitRed)
+ qed
+ next
+ fix a assume "e2 = Throw a"
+ thus ?thesis by(fastforce intro:red_reds.SFAssThrow)
+ qed
+ next
+ assume nf: "\<not> final e2"
+ then have "val_of e2 = None" using final_def val_of_spec by fastforce
+ then have "P,sh \<turnstile>\<^sub>b (e2,b) \<surd>" using WTrtSFAss.prems(2) by(simp add:bconf_SFAss)
+ with WTrtSFAss nf show ?thesis
+ by simp (fast intro!:SFAssRed)
+ qed
+next
+ case (WTrtCall E e C M Ts T pns body D es Ts')
+ have wte: "P,E,h,sh \<turnstile> e : Class C"
+ and "method": "P \<turnstile> C sees M,NonStatic:Ts\<rightarrow>T = (pns,body) in D"
+ and wtes: "P,E,h,sh \<turnstile> es [:] Ts'"and sub: "P \<turnstile> Ts' [\<le>] Ts"
+ and IHes: "\<And>l.
+ \<lbrakk>\<D>s es \<lfloor>dom l\<rfloor>; P,sh \<turnstile>\<^sub>b (es,b) \<surd>; \<not> finals es\<rbrakk>
+ \<Longrightarrow> \<exists>es' s' b'. P \<turnstile> \<langle>es,(h,l,sh),b\<rangle> [\<rightarrow>] \<langle>es',s',b'\<rangle>"
+ and D: "\<D> (e\<bullet>M(es)) \<lfloor>dom l\<rfloor>" by fact+
+ show ?case
+ proof cases
+ assume "final e"
+ with wte show ?thesis
+ proof (rule final_addrE)
+ fix a assume e_addr: "e = addr a"
+ show ?thesis
+ proof cases
+ assume es: "\<exists>vs. es = map Val vs"
+ from wte e_addr obtain fs where ha: "h a = Some(C,fs)" by auto
+ show ?thesis
+ using e_addr ha "method" WTrts_same_length[OF wtes] sub es sees_wf_mdecl[OF wf "method"]
+ by(fastforce intro!: RedCall simp:list_all2_iff wf_mdecl_def)
+ next
+ assume "\<not>(\<exists>vs. es = map Val vs)"
+ hence not_all_Val: "\<not>(\<forall>e \<in> set es. \<exists>v. e = Val v)"
+ by(simp add:ex_map_conv)
+ let ?ves = "takeWhile (\<lambda>e. \<exists>v. e = Val v) es"
+ let ?rest = "dropWhile (\<lambda>e. \<exists>v. e = Val v) es"
+ let ?ex = "hd ?rest" let ?rst = "tl ?rest"
+ from not_all_Val have nonempty: "?rest \<noteq> []" by auto
+ hence es: "es = ?ves @ ?ex # ?rst" by simp
+ have "\<forall>e \<in> set ?ves. \<exists>v. e = Val v" by(fastforce dest:set_takeWhileD)
+ then obtain vs where ves: "?ves = map Val vs"
+ using ex_map_conv by blast
+ show ?thesis
+ proof cases
+ assume "final ?ex"
+ moreover from nonempty have "\<not>(\<exists>v. ?ex = Val v)"
+ by(auto simp:neq_Nil_conv simp del:dropWhile_eq_Nil_conv)
+ (simp add:dropWhile_eq_Cons_conv)
+ ultimately obtain b where ex_Throw: "?ex = Throw b"
+ by(fast elim!:finalE)
+ show ?thesis using e_addr es ex_Throw ves
+ by(fastforce intro:CallThrowParams)
+ next
+ assume not_fin: "\<not> final ?ex"
+ have "finals es = finals(?ves @ ?ex # ?rst)" using es
+ by(rule arg_cong)
+ also have "\<dots> = finals(?ex # ?rst)" using ves by simp
+ finally have "finals es = finals(?ex # ?rst)" .
+ hence fes: "\<not> finals es" using not_finals_ConsI[OF not_fin] by blast
+ have "P,sh \<turnstile>\<^sub>b (es,b) \<surd>" using bconf_Call WTrtCall.prems(2)
+ by (metis e_addr option.simps(5) val_of.simps(1))
+ thus ?thesis using fes e_addr D IHes by(fastforce intro!:CallParams)
+ qed
+ qed
+ next
+ fix a assume "e = Throw a"
+ with WTrtCall.prems show ?thesis by(fast intro!:CallThrowObj)
+ qed
+ next
+ assume nf: "\<not> final e"
+ then have e1: "val_of e = None" proof(cases e)qed(auto)
+ then have "P,sh \<turnstile>\<^sub>b (e,b) \<surd>" using WTrtCall.prems(2) by(simp add:bconf_Call)
+ with WTrtCall nf e1 show ?thesis by simp (blast intro!:CallObj)
+ qed
+next
+ case (WTrtCallNT E e es Ts M T) show ?case
+ proof cases
+ assume "final e"
+ moreover
+ { fix v assume e: "e = Val v"
+ hence "e = null" using WTrtCallNT by simp
+ have ?case
+ proof cases
+ assume "finals es"
+ moreover
+ { fix vs assume "es = map Val vs"
+ with WTrtCallNT e have ?thesis by(fastforce intro: RedCallNull) }
+ moreover
+ { fix vs a es' assume "es = map Val vs @ Throw a # es'"
+ with WTrtCallNT e have ?thesis by(fastforce intro: CallThrowParams) }
+ ultimately show ?thesis by(fastforce simp:finals_def)
+ next
+ assume nf: "\<not> finals es" \<comment> \<open>@{term es} reduces by IH\<close>
+ have "P,sh \<turnstile>\<^sub>b (es,b) \<surd>" using WTrtCallNT.prems(2) e by (simp add: bconf_Call)
+ with WTrtCallNT e nf show ?thesis by(fastforce intro: CallParams)
+ qed
+ }
+ moreover
+ { fix a assume "e = Throw a"
+ with WTrtCallNT have ?case by(fastforce intro: CallThrowObj) }
+ ultimately show ?thesis by(fastforce simp:final_def)
+ next
+ assume nf: "\<not> final e" \<comment> \<open>@{term e} reduces by IH\<close>
+ then have "val_of e = None" proof(cases e)qed(auto)
+ then have "P,sh \<turnstile>\<^sub>b (e,b) \<surd>" using WTrtCallNT.prems(2) by(simp add:bconf_Call)
+ with WTrtCallNT nf show ?thesis by (fastforce intro:CallObj)
+ qed
+next
+ case (WTrtSCall C M Ts T pns body D E es Ts' sfs vs)
+ have "method": "P \<turnstile> C sees M,Static:Ts\<rightarrow>T = (pns,body) in D"
+ and wtes: "P,E,h,sh \<turnstile> es [:] Ts'"and sub: "P \<turnstile> Ts' [\<le>] Ts"
+ and IHes: "\<And>l.
+ \<lbrakk>\<D>s es \<lfloor>dom l\<rfloor>; P,sh \<turnstile>\<^sub>b (es,b) \<surd>; \<not> finals es\<rbrakk>
+ \<Longrightarrow> \<exists>es' s' b'. P \<turnstile> \<langle>es,(h,l,sh),b\<rangle> [\<rightarrow>] \<langle>es',s',b'\<rangle>"
+ and clinit: "M = clinit \<longrightarrow> sh D = \<lfloor>(sfs, Processing)\<rfloor> \<and> es = map Val vs"
+ and D: "\<D> (C\<bullet>\<^sub>sM(es)) \<lfloor>dom l\<rfloor>" by fact+
+ show ?case
+ proof cases
+ assume es: "\<exists>vs. es = map Val vs"
+ show ?thesis
+ proof (cases b)
+ case True
+ then show ?thesis
+ using "method" WTrts_same_length[OF wtes] sub es sees_wf_mdecl[OF wf "method"] True
+ by(fastforce intro!: RedSCall simp:list_all2_iff wf_mdecl_def)
+ next
+ case False
+ show ?thesis
+ using "method" clinit WTrts_same_length[OF wtes] sub es False
+ by (metis (full_types) red_reds.SCallInitDoneRed red_reds.SCallInitRed)
+ qed
+ next
+ assume nmap: "\<not>(\<exists>vs. es = map Val vs)"
+ hence not_all_Val: "\<not>(\<forall>e \<in> set es. \<exists>v. e = Val v)"
+ by(simp add:ex_map_conv)
+ let ?ves = "takeWhile (\<lambda>e. \<exists>v. e = Val v) es"
+ let ?rest = "dropWhile (\<lambda>e. \<exists>v. e = Val v) es"
+ let ?ex = "hd ?rest" let ?rst = "tl ?rest"
+ from not_all_Val have nonempty: "?rest \<noteq> []" by auto
+ hence es: "es = ?ves @ ?ex # ?rst" by simp
+ have "\<forall>e \<in> set ?ves. \<exists>v. e = Val v" by(fastforce dest:set_takeWhileD)
+ then obtain vs where ves: "?ves = map Val vs"
+ using ex_map_conv by blast
+ show ?thesis
+ proof cases
+ assume "final ?ex"
+ moreover from nonempty have "\<not>(\<exists>v. ?ex = Val v)"
+ by(auto simp:neq_Nil_conv simp del:dropWhile_eq_Nil_conv)
+ (simp add:dropWhile_eq_Cons_conv)
+ ultimately obtain b where ex_Throw: "?ex = Throw b"
+ by(fast elim!:finalE)
+ show ?thesis using es ex_Throw ves
+ by(fastforce intro:SCallThrowParams)
+ next
+ assume not_fin: "\<not> final ?ex"
+ have "finals es = finals(?ves @ ?ex # ?rst)" using es
+ by(rule arg_cong)
+ also have "\<dots> = finals(?ex # ?rst)" using ves by simp
+ finally have "finals es = finals(?ex # ?rst)" .
+ hence fes: "\<not> finals es" using not_finals_ConsI[OF not_fin] by blast
+ have "P,sh \<turnstile>\<^sub>b (es,b) \<surd>"
+ by (meson WTrtSCall.prems(2) nmap bconf_SCall map_vals_of_spec not_None_eq)
+ thus ?thesis using fes D IHes by(fastforce intro!:SCallParams)
+ qed
+ qed
+next
+ case WTrtNil thus ?case by simp
+next
+ case (WTrtCons E e T es Ts)
+ have IHe: "\<And>l. \<lbrakk>\<D> e \<lfloor>dom l\<rfloor>; P,sh \<turnstile>\<^sub>b (e,b) \<surd>; \<not> final e\<rbrakk>
+ \<Longrightarrow> \<exists>e' s' b'. P \<turnstile> \<langle>e,(h,l,sh),b\<rangle> \<rightarrow> \<langle>e',s',b'\<rangle>"
+ and IHes: "\<And>l. \<lbrakk>\<D>s es \<lfloor>dom l\<rfloor>; P,sh \<turnstile>\<^sub>b (es,b) \<surd>; \<not> finals es\<rbrakk>
+ \<Longrightarrow> \<exists>es' s' b'. P \<turnstile> \<langle>es,(h,l,sh),b\<rangle> [\<rightarrow>] \<langle>es',s',b'\<rangle>"
+ and D: "\<D>s (e#es) \<lfloor>dom l\<rfloor>" and not_fins: "\<not> finals(e # es)" by fact+
+ have De: "\<D> e \<lfloor>dom l\<rfloor>" and Des: "\<D>s es (\<lfloor>dom l\<rfloor> \<squnion> \<A> e)"
+ using D by auto
+ show ?case
+ proof cases
+ assume "final e"
+ thus ?thesis
+ proof (rule finalE)
+ fix v assume e: "e = Val v"
+ hence Des': "\<D>s es \<lfloor>dom l\<rfloor>" using De Des by auto
+ have bconfs: "P,sh \<turnstile>\<^sub>b (es,b) \<surd>" using WTrtCons.prems(2) e by(simp add: bconf_Cons)
+ have not_fins_tl: "\<not> finals es" using not_fins e by simp
+ show ?thesis using e IHes[OF Des' bconfs not_fins_tl]
+ by (blast intro!:ListRed2)
+ next
+ fix a assume "e = Throw a"
+ hence False using not_fins by simp
+ thus ?thesis ..
+ qed
+ next
+ assume nf:"\<not> final e"
+ then have "val_of e = None" proof(cases e)qed(auto)
+ then have bconf: "P,sh \<turnstile>\<^sub>b (e,b) \<surd>" using WTrtCons.prems(2) by(simp add: bconf_Cons)
+ with IHe[OF De bconf nf] show ?thesis by(fast intro!:ListRed1)
+ qed
+next
+ case (WTrtInitBlock v T\<^sub>1 T E V e\<^sub>2 T\<^sub>2)
+ have IH2: "\<And>l. \<lbrakk>\<D> e\<^sub>2 \<lfloor>dom l\<rfloor>; P,sh \<turnstile>\<^sub>b (e\<^sub>2,b) \<surd>; \<not> final e\<^sub>2\<rbrakk>
+ \<Longrightarrow> \<exists>e' s' b'. P \<turnstile> \<langle>e\<^sub>2,(h,l,sh),b\<rangle> \<rightarrow> \<langle>e',s',b'\<rangle>"
+ and D: "\<D> {V:T := Val v; e\<^sub>2} \<lfloor>dom l\<rfloor>" by fact+
+ show ?case
+ proof cases
+ assume "final e\<^sub>2"
+ then show ?thesis
+ proof (rule finalE)
+ fix v\<^sub>2 assume "e\<^sub>2 = Val v\<^sub>2"
+ thus ?thesis by(fast intro:RedInitBlock)
+ next
+ fix a assume "e\<^sub>2 = Throw a"
+ thus ?thesis by(fast intro:red_reds.InitBlockThrow)
+ qed
+ next
+ assume not_fin2: "\<not> final e\<^sub>2"
+ then have "val_of e\<^sub>2 = None" proof(cases e\<^sub>2)qed(auto)
+ from D have D2: "\<D> e\<^sub>2 \<lfloor>dom(l(V\<mapsto>v))\<rfloor>" by (auto simp:hyperset_defs)
+ have e2conf: "P,sh \<turnstile>\<^sub>b (e\<^sub>2,b) \<surd>" using WTrtInitBlock.prems(2) by(simp add: bconf_InitBlock)
+ from IH2[OF D2 e2conf not_fin2]
+ obtain h' l' sh' e' b' where red2: "P \<turnstile> \<langle>e\<^sub>2,(h, l(V\<mapsto>v),sh),b\<rangle> \<rightarrow> \<langle>e',(h', l',sh'),b'\<rangle>"
+ by auto
+ from red_lcl_incr[OF red2] have "V \<in> dom l'" by auto
+ with red2 show ?thesis by(fastforce intro:InitBlockRed)
+ qed
+next
+ case (WTrtBlock E V T e T')
+ have IH: "\<And>l. \<lbrakk>\<D> e \<lfloor>dom l\<rfloor>; P,sh \<turnstile>\<^sub>b (e,b) \<surd>; \<not> final e\<rbrakk>
+ \<Longrightarrow> \<exists>e' s' b'. P \<turnstile> \<langle>e,(h,l,sh),b\<rangle> \<rightarrow> \<langle>e',s',b'\<rangle>"
+ and unass: "\<not> assigned V e" and D: "\<D> {V:T; e} \<lfloor>dom l\<rfloor>" by fact+
+ show ?case
+ proof cases
+ assume "final e"
+ thus ?thesis
+ proof (rule finalE)
+ fix v assume "e = Val v" thus ?thesis by(fast intro:RedBlock)
+ next
+ fix a assume "e = Throw a"
+ thus ?thesis by(fast intro:red_reds.BlockThrow)
+ qed
+ next
+ assume not_fin: "\<not> final e"
+ then have "val_of e = None" proof(cases e)qed(auto)
+ from D have De: "\<D> e \<lfloor>dom(l(V:=None))\<rfloor>" by(simp add:hyperset_defs)
+ have bconf: "P,sh \<turnstile>\<^sub>b (e,b) \<surd>" using WTrtBlock by(simp add: bconf_Block)
+ from IH[OF De bconf not_fin]
+ obtain h' l' sh' e' b' where red: "P \<turnstile> \<langle>e,(h,l(V:=None),sh),b\<rangle> \<rightarrow> \<langle>e',(h',l',sh'),b'\<rangle>"
+ by auto
+ show ?thesis
+ proof (cases "l' V")
+ assume "l' V = None"
+ with red unass show ?thesis by(blast intro: BlockRedNone)
+ next
+ fix v assume "l' V = Some v"
+ with red unass show ?thesis by(blast intro: BlockRedSome)
+ qed
+ qed
+next
+ case (WTrtSeq E e1 T1 e2 T2) show ?case
+ proof cases
+ assume "final e1"
+ thus ?thesis
+ by(fast elim:finalE intro:RedSeq red_reds.SeqThrow)
+ next
+ assume nf: "\<not> final e1"
+ then have e1: "val_of e1 = None" proof(cases e1)qed(auto)
+ then show ?thesis
+ proof(cases "lass_val_of e1")
+ case None
+ then have "P,sh \<turnstile>\<^sub>b (e1,b) \<surd>" using WTrtSeq.prems(2) e1 by(simp add: bconf_Seq)
+ with WTrtSeq nf e1 None show ?thesis by simp (blast intro:SeqRed)
+ next
+ case (Some p)
+ obtain V v where "e1 = V:=Val v" using lass_val_of_spec[OF Some] by simp
+ then show ?thesis using SeqRed[OF RedLAss] by blast
+ qed
+ qed
+next
+ case (WTrtCond E e e\<^sub>1 T\<^sub>1 e\<^sub>2 T\<^sub>2 T)
+ have wt: "P,E,h,sh \<turnstile> e : Boolean" by fact
+ show ?case
+ proof cases
+ assume "final e"
+ thus ?thesis
+ proof (rule finalE)
+ fix v assume val: "e = Val v"
+ then obtain b where v: "v = Bool b" using wt by auto
+ show ?thesis
+ proof (cases b)
+ case True with val v show ?thesis by(fastforce intro:RedCondT simp: prod_cases3)
+ next
+ case False with val v show ?thesis by(fastforce intro:RedCondF simp: prod_cases3)
+ qed
+ next
+ fix a assume "e = Throw a"
+ thus ?thesis by(fast intro:red_reds.CondThrow)
+ qed
+ next
+ assume nf: "\<not> final e"
+ then have "bool_of e = None" proof(cases e)qed(auto)
+ then have "P,sh \<turnstile>\<^sub>b (e,b) \<surd>" using WTrtCond.prems(2) by(simp add: bconf_Cond)
+ with WTrtCond nf show ?thesis by simp (blast intro:CondRed)
+ qed
+next
+ case WTrtWhile show ?case by(fast intro:RedWhile)
+next
+ case (WTrtThrow E e T\<^sub>r T) show ?case
+ proof cases
+ assume "final e" \<comment> \<open>Then @{term e} must be @{term throw} or @{term null}\<close>
+ with WTrtThrow show ?thesis
+ by(fastforce simp:final_def is_refT_def
+ intro:red_reds.ThrowThrow red_reds.RedThrowNull)
+ next
+ assume nf: "\<not> final e" \<comment> \<open>Then @{term e} must reduce\<close>
+ then have "val_of e = None" proof(cases e)qed(auto)
+ then have "P,sh \<turnstile>\<^sub>b (e,b) \<surd>" using WTrtThrow.prems(2) by(simp add: bconf_Throw)
+ with WTrtThrow nf show ?thesis by simp (blast intro:ThrowRed)
+ qed
+next
+ case (WTrtTry E e1 T1 V C e2 T2)
+ have wt1: "P,E,h,sh \<turnstile> e1 : T1" by fact
+ show ?case
+ proof cases
+ assume "final e1"
+ thus ?thesis
+ proof (rule finalE)
+ fix v assume "e1 = Val v"
+ thus ?thesis by(fast intro:RedTry)
+ next
+ fix a assume e1_Throw: "e1 = Throw a"
+ with wt1 obtain D fs where ha: "h a = Some(D,fs)" by fastforce
+ show ?thesis
+ proof cases
+ assume "P \<turnstile> D \<preceq>\<^sup>* C"
+ with e1_Throw ha show ?thesis by(fastforce intro!:RedTryCatch)
+ next
+ assume "\<not> P \<turnstile> D \<preceq>\<^sup>* C"
+ with e1_Throw ha show ?thesis by(fastforce intro!:RedTryFail)
+ qed
+ qed
+ next
+ assume nf: "\<not> final e1"
+ then have "val_of e1 = None" proof(cases e1)qed(auto)
+ then have "P,sh \<turnstile>\<^sub>b (e1,b) \<surd>" using WTrtTry.prems(2) by(simp add: bconf_Try)
+ with WTrtTry nf show ?thesis by simp (fast intro:TryRed)
+ qed
+next
+ case (WTrtInit E e T\<^sub>r C Cs b) show ?case
+ proof(cases Cs)
+ case Nil then show ?thesis using WTrtInit by(fastforce intro!: RedInit)
+ next
+ case (Cons C' Cs')
+ show ?thesis
+ proof(cases b)
+ case True then show ?thesis using Cons by(fastforce intro!: RedInitRInit)
+ next
+ case False
+ show ?thesis
+ proof(cases "sh C'")
+ case None
+ then show ?thesis using False Cons by(fastforce intro!: InitNoneRed)
+ next
+ case (Some sfsi)
+ then obtain sfs i where sfsi:"sfsi = (sfs,i)" by(cases sfsi)
+ show ?thesis
+ proof(cases i)
+ case Done
+ then show ?thesis using False Some sfsi Cons by(fastforce intro!: RedInitDone)
+ next
+ case Processing
+ then show ?thesis using False Some sfsi Cons by(fastforce intro!: RedInitProcessing)
+ next
+ case Error
+ then show ?thesis using False Some sfsi Cons by(fastforce intro!: RedInitError)
+ next
+ case Prepared
+ show ?thesis
+ proof cases
+ assume "C' = Object"
+ then show ?thesis using False Some sfsi Prepared Cons by(fastforce intro: InitObjectRed)
+ next
+ assume "C' \<noteq> Object"
+ then show ?thesis using False Some sfsi Prepared WTrtInit.hyps(3) Cons
+ by(simp only: is_class_def)(fastforce intro!: InitNonObjectSuperRed)
+ qed
+ qed
+ qed
+ qed
+ qed
+next
+ case (WTrtRI E e T\<^sub>r e' T\<^sub>r' C Cs)
+ obtain sfs i where shC: "sh C = \<lfloor>(sfs, i)\<rfloor>" using WTrtRI.hyps(9) by blast
+ show ?case
+ proof cases
+ assume fin: "final e" then show ?thesis
+ proof (rule finalE)
+ fix v assume e: "e = Val v"
+ then show ?thesis using shC e by(fast intro:RedRInit)
+ next
+ fix a assume eThrow: "e = Throw a"
+ show ?thesis
+ proof(cases Cs)
+ case Nil then show ?thesis using eThrow shC by(fastforce intro!: RInitThrow)
+ next
+ case Cons then show ?thesis using eThrow shC by(fastforce intro!: RInitInitThrow)
+ qed
+ qed
+ next
+ assume nf: "\<not> final e"
+ then have "val_of e = None" proof(cases e)qed(auto)
+ then have "P,sh \<turnstile>\<^sub>b (e,b) \<surd>" using WTrtRI.prems(2) by(simp add: bconf_RI)
+ with WTrtRI nf show ?thesis by simp (meson red_reds.RInitRed)
+ qed
+qed
+(*>*)
+
+end
diff --git a/thys/JinjaDCI/J/SmallStep.thy b/thys/JinjaDCI/J/SmallStep.thy
--- a/thys/JinjaDCI/J/SmallStep.thy
+++ b/thys/JinjaDCI/J/SmallStep.thy
@@ -1,646 +1,646 @@
-(* Title: JinjaDCI/J/SmallStep.thy
- Author: Tobias Nipkow, Susannah Mansky
- Copyright 2003 Technische Universitaet Muenchen, 2019-20 UIUC
-
- Based on the Jinja theory J/SmallStep.thy by Tobias Nipkow
-*)
-
-section \<open> Small Step Semantics \<close>
-
-theory SmallStep
-imports Expr State WWellForm
-begin
-
-fun blocks :: "vname list * ty list * val list * expr \<Rightarrow> expr"
-where
- "blocks(V#Vs, T#Ts, v#vs, e) = {V:T := Val v; blocks(Vs,Ts,vs,e)}"
-|"blocks([],[],[],e) = e"
-
-lemmas blocks_induct = blocks.induct[split_format (complete)]
-
-lemma [simp]:
- "\<lbrakk> size vs = size Vs; size Ts = size Vs \<rbrakk> \<Longrightarrow> fv(blocks(Vs,Ts,vs,e)) = fv e - set Vs"
-(*<*)
-by (induct rule:blocks_induct) auto
-(*>*)
-
-
-lemma sub_RI_blocks_body[iff]: "length vs = length pns \<Longrightarrow> length Ts = length pns
- \<Longrightarrow> sub_RI (blocks (pns, Ts, vs, body)) \<longleftrightarrow> sub_RI body"
-proof(induct pns arbitrary: Ts vs)
- case Nil then show ?case by simp
-next
- case Cons then show ?case by(cases vs; cases Ts) auto
-qed
-
-
-definition assigned :: "'a \<Rightarrow> 'a exp \<Rightarrow> bool"
-where
- "assigned V e \<equiv> \<exists>v e'. e = (V := Val v;; e')"
-
-\<comment> \<open> expression is okay to go the right side of @{text INIT} or @{text "RI \<leftarrow>"}
- or to have indicator Boolean be True (in latter case, given that class is
- also verified initialized) \<close>
-fun icheck :: "'m prog \<Rightarrow> cname \<Rightarrow> 'a exp \<Rightarrow> bool" where
-"icheck P C' (new C) = (C' = C)" |
-"icheck P D' (C\<bullet>\<^sub>sF{D}) = ((D' = D) \<and> (\<exists>T. P \<turnstile> C has F,Static:T in D))" |
-"icheck P D' (C\<bullet>\<^sub>sF{D}:=(Val v)) = ((D' = D) \<and> (\<exists>T. P \<turnstile> C has F,Static:T in D))" |
-"icheck P D (C\<bullet>\<^sub>sM(es)) = ((\<exists>vs. es = map Val vs) \<and> (\<exists>Ts T m. P \<turnstile> C sees M,Static:Ts\<rightarrow>T = m in D))" |
-"icheck P _ _ = False"
-
-lemma nicheck_SFAss_nonVal: "val_of e\<^sub>2 = None \<Longrightarrow> \<not>icheck P C' (C\<bullet>\<^sub>sF{D} := (e\<^sub>2::'a exp))"
- by(rule notI, cases e\<^sub>2, auto)
-
-inductive_set
- red :: "J_prog \<Rightarrow> ((expr \<times> state \<times> bool) \<times> (expr \<times> state \<times> bool)) set"
- and reds :: "J_prog \<Rightarrow> ((expr list \<times> state \<times> bool) \<times> (expr list \<times> state \<times> bool)) set"
- and red' :: "J_prog \<Rightarrow> expr \<Rightarrow> state \<Rightarrow> bool \<Rightarrow> expr \<Rightarrow> state \<Rightarrow> bool \<Rightarrow> bool"
- ("_ \<turnstile> ((1\<langle>_,/_,/_\<rangle>) \<rightarrow>/ (1\<langle>_,/_,/_\<rangle>))" [51,0,0,0,0,0,0] 81)
- and reds' :: "J_prog \<Rightarrow> expr list \<Rightarrow> state \<Rightarrow> bool \<Rightarrow> expr list \<Rightarrow> state \<Rightarrow> bool \<Rightarrow> bool"
- ("_ \<turnstile> ((1\<langle>_,/_,/_\<rangle>) [\<rightarrow>]/ (1\<langle>_,/_,/_\<rangle>))" [51,0,0,0,0,0,0] 81)
- for P :: J_prog
-where
-
- "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow> \<langle>e',s',b'\<rangle> \<equiv> ((e,s,b), e',s',b') \<in> red P"
-| "P \<turnstile> \<langle>es,s,b\<rangle> [\<rightarrow>] \<langle>es',s',b'\<rangle> \<equiv> ((es,s,b), es',s',b') \<in> reds P"
-
-| RedNew:
- "\<lbrakk> new_Addr h = Some a; P \<turnstile> C has_fields FDTs; h' = h(a\<mapsto>blank P C) \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>new C, (h,l,sh), True\<rangle> \<rightarrow> \<langle>addr a, (h',l,sh), False\<rangle>"
-
-| RedNewFail:
- "\<lbrakk> new_Addr h = None; is_class P C \<rbrakk> \<Longrightarrow>
- P \<turnstile> \<langle>new C, (h,l,sh), True\<rangle> \<rightarrow> \<langle>THROW OutOfMemory, (h,l,sh), False\<rangle>"
-
-| NewInitDoneRed:
- "sh C = Some (sfs, Done) \<Longrightarrow>
- P \<turnstile> \<langle>new C, (h,l,sh), False\<rangle> \<rightarrow> \<langle>new C, (h,l,sh), True\<rangle>"
-
-| NewInitRed:
- "\<lbrakk> \<nexists>sfs. sh C = Some (sfs, Done); is_class P C \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>new C,(h,l,sh),False\<rangle> \<rightarrow> \<langle>INIT C ([C],False) \<leftarrow> new C,(h,l,sh),False\<rangle>"
-
-| CastRed:
- "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow> \<langle>e',s',b'\<rangle> \<Longrightarrow>
- P \<turnstile> \<langle>Cast C e, s, b\<rangle> \<rightarrow> \<langle>Cast C e', s', b'\<rangle>"
-
-| RedCastNull:
- "P \<turnstile> \<langle>Cast C null, s, b\<rangle> \<rightarrow> \<langle>null,s,b\<rangle>"
-
-| RedCast:
- "\<lbrakk> h a = Some(D,fs); P \<turnstile> D \<preceq>\<^sup>* C \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>Cast C (addr a), (h,l,sh), b\<rangle> \<rightarrow> \<langle>addr a, (h,l,sh), b\<rangle>"
-
-| RedCastFail:
- "\<lbrakk> h a = Some(D,fs); \<not> P \<turnstile> D \<preceq>\<^sup>* C \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>Cast C (addr a), (h,l,sh), b\<rangle> \<rightarrow> \<langle>THROW ClassCast, (h,l,sh), b\<rangle>"
-
-| BinOpRed1:
- "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow> \<langle>e',s',b'\<rangle> \<Longrightarrow>
- P \<turnstile> \<langle>e \<guillemotleft>bop\<guillemotright> e\<^sub>2, s, b\<rangle> \<rightarrow> \<langle>e' \<guillemotleft>bop\<guillemotright> e\<^sub>2, s', b'\<rangle>"
-
-| BinOpRed2:
- "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow> \<langle>e',s',b'\<rangle> \<Longrightarrow>
- P \<turnstile> \<langle>(Val v\<^sub>1) \<guillemotleft>bop\<guillemotright> e, s, b\<rangle> \<rightarrow> \<langle>(Val v\<^sub>1) \<guillemotleft>bop\<guillemotright> e', s', b'\<rangle>"
-
-| RedBinOp:
- "binop(bop,v\<^sub>1,v\<^sub>2) = Some v \<Longrightarrow>
- P \<turnstile> \<langle>(Val v\<^sub>1) \<guillemotleft>bop\<guillemotright> (Val v\<^sub>2), s, b\<rangle> \<rightarrow> \<langle>Val v,s,b\<rangle>"
-
-| RedVar:
- "l V = Some v \<Longrightarrow>
- P \<turnstile> \<langle>Var V,(h,l,sh),b\<rangle> \<rightarrow> \<langle>Val v,(h,l,sh),b\<rangle>"
-
-| LAssRed:
- "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow> \<langle>e',s',b'\<rangle> \<Longrightarrow>
- P \<turnstile> \<langle>V:=e,s,b\<rangle> \<rightarrow> \<langle>V:=e',s',b'\<rangle>"
-
-| RedLAss:
- "P \<turnstile> \<langle>V:=(Val v), (h,l,sh), b\<rangle> \<rightarrow> \<langle>unit, (h,l(V\<mapsto>v),sh), b\<rangle>"
-
-| FAccRed:
- "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow> \<langle>e',s',b'\<rangle> \<Longrightarrow>
- P \<turnstile> \<langle>e\<bullet>F{D}, s, b\<rangle> \<rightarrow> \<langle>e'\<bullet>F{D}, s', b'\<rangle>"
-
-| RedFAcc:
- "\<lbrakk> h a = Some(C,fs); fs(F,D) = Some v;
- P \<turnstile> C has F,NonStatic:t in D \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>(addr a)\<bullet>F{D}, (h,l,sh), b\<rangle> \<rightarrow> \<langle>Val v,(h,l,sh),b\<rangle>"
-
-| RedFAccNull:
- "P \<turnstile> \<langle>null\<bullet>F{D}, s, b\<rangle> \<rightarrow> \<langle>THROW NullPointer, s, b\<rangle>"
-
-| RedFAccNone:
- "\<lbrakk> h a = Some(C,fs); \<not>(\<exists>b t. P \<turnstile> C has F,b:t in D) \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>(addr a)\<bullet>F{D},(h,l,sh),b\<rangle> \<rightarrow> \<langle>THROW NoSuchFieldError,(h,l,sh),b\<rangle>"
-
-| RedFAccStatic:
- "\<lbrakk> h a = Some(C,fs); P \<turnstile> C has F,Static:t in D \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>(addr a)\<bullet>F{D},(h,l,sh),b\<rangle> \<rightarrow> \<langle>THROW IncompatibleClassChangeError,(h,l,sh),b\<rangle>"
-
-| RedSFAcc:
- "\<lbrakk> P \<turnstile> C has F,Static:t in D;
- sh D = Some (sfs,i);
- sfs F = Some v \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D},(h,l,sh),True\<rangle> \<rightarrow> \<langle>Val v,(h,l,sh),False\<rangle>"
-
-| SFAccInitDoneRed:
- "\<lbrakk> P \<turnstile> C has F,Static:t in D;
- sh D = Some (sfs,Done) \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D},(h,l,sh),False\<rangle> \<rightarrow> \<langle>C\<bullet>\<^sub>sF{D},(h,l,sh),True\<rangle>"
-
-| SFAccInitRed:
- "\<lbrakk> P \<turnstile> C has F,Static:t in D;
- \<nexists>sfs. sh D = Some (sfs,Done) \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D},(h,l,sh),False\<rangle> \<rightarrow> \<langle>INIT D ([D],False) \<leftarrow> C\<bullet>\<^sub>sF{D},(h,l,sh),False\<rangle>"
-
-| RedSFAccNone:
- "\<not>(\<exists>b t. P \<turnstile> C has F,b:t in D)
- \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D},(h,l,sh),b\<rangle> \<rightarrow> \<langle>THROW NoSuchFieldError,(h,l,sh),False\<rangle>"
-
-| RedSFAccNonStatic:
- "P \<turnstile> C has F,NonStatic:t in D
- \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D},(h,l,sh),b\<rangle> \<rightarrow> \<langle>THROW IncompatibleClassChangeError,(h,l,sh),False\<rangle>"
-
-| FAssRed1:
- "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow> \<langle>e',s',b'\<rangle> \<Longrightarrow>
- P \<turnstile> \<langle>e\<bullet>F{D}:=e\<^sub>2, s, b\<rangle> \<rightarrow> \<langle>e'\<bullet>F{D}:=e\<^sub>2, s', b'\<rangle>"
-
-| FAssRed2:
- "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow> \<langle>e',s',b'\<rangle> \<Longrightarrow>
- P \<turnstile> \<langle>Val v\<bullet>F{D}:=e, s, b\<rangle> \<rightarrow> \<langle>Val v\<bullet>F{D}:=e', s', b'\<rangle>"
-
-| RedFAss:
- "\<lbrakk> P \<turnstile> C has F,NonStatic:t in D; h a = Some(C,fs) \<rbrakk> \<Longrightarrow>
- P \<turnstile> \<langle>(addr a)\<bullet>F{D}:=(Val v), (h,l,sh), b\<rangle> \<rightarrow> \<langle>unit, (h(a \<mapsto> (C,fs((F,D) \<mapsto> v))),l,sh), b\<rangle>"
-
-| RedFAssNull:
- "P \<turnstile> \<langle>null\<bullet>F{D}:=Val v, s, b\<rangle> \<rightarrow> \<langle>THROW NullPointer, s, b\<rangle>"
-
-| RedFAssNone:
- "\<lbrakk> h a = Some(C,fs); \<not>(\<exists>b t. P \<turnstile> C has F,b:t in D) \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>(addr a)\<bullet>F{D}:=(Val v),(h,l,sh),b\<rangle> \<rightarrow> \<langle>THROW NoSuchFieldError,(h,l,sh),b\<rangle>"
-
-| RedFAssStatic:
- "\<lbrakk> h a = Some(C,fs); P \<turnstile> C has F,Static:t in D \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>(addr a)\<bullet>F{D}:=(Val v),(h,l,sh),b\<rangle> \<rightarrow> \<langle>THROW IncompatibleClassChangeError,(h,l,sh),b\<rangle>"
-
-| SFAssRed:
- "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow> \<langle>e',s',b'\<rangle> \<Longrightarrow>
- P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D}:=e, s, b\<rangle> \<rightarrow> \<langle>C\<bullet>\<^sub>sF{D}:=e', s', b'\<rangle>"
-
-| RedSFAss:
- "\<lbrakk> P \<turnstile> C has F,Static:t in D;
- sh D = Some(sfs,i);
- sfs' = sfs(F\<mapsto>v); sh' = sh(D\<mapsto>(sfs',i)) \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D}:=(Val v),(h,l,sh),True\<rangle> \<rightarrow> \<langle>unit,(h,l,sh'),False\<rangle>"
-
-| SFAssInitDoneRed:
- "\<lbrakk> P \<turnstile> C has F,Static:t in D;
- sh D = Some(sfs,Done) \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D}:=(Val v),(h,l,sh),False\<rangle> \<rightarrow> \<langle>C\<bullet>\<^sub>sF{D}:=(Val v),(h,l,sh),True\<rangle>"
-
-| SFAssInitRed:
- "\<lbrakk> P \<turnstile> C has F,Static:t in D;
- \<nexists>sfs. sh D = Some(sfs,Done) \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D}:=(Val v),(h,l,sh),False\<rangle> \<rightarrow> \<langle>INIT D ([D],False)\<leftarrow> C\<bullet>\<^sub>sF{D}:=(Val v),(h,l,sh),False\<rangle>"
-
-| RedSFAssNone:
- "\<not>(\<exists>b t. P \<turnstile> C has F,b:t in D)
- \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D}:=(Val v),s,b\<rangle> \<rightarrow> \<langle>THROW NoSuchFieldError,s,False\<rangle>"
-
-| RedSFAssNonStatic:
- "P \<turnstile> C has F,NonStatic:t in D
- \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D}:=(Val v),s,b\<rangle> \<rightarrow> \<langle>THROW IncompatibleClassChangeError,s,False\<rangle>"
-
-| CallObj:
- "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow> \<langle>e',s',b'\<rangle> \<Longrightarrow>
- P \<turnstile> \<langle>e\<bullet>M(es),s,b\<rangle> \<rightarrow> \<langle>e'\<bullet>M(es),s',b'\<rangle>"
-
-| CallParams:
- "P \<turnstile> \<langle>es,s,b\<rangle> [\<rightarrow>] \<langle>es',s',b'\<rangle> \<Longrightarrow>
- P \<turnstile> \<langle>(Val v)\<bullet>M(es),s,b\<rangle> \<rightarrow> \<langle>(Val v)\<bullet>M(es'),s',b'\<rangle>"
-
-| RedCall:
- "\<lbrakk> h a = Some(C,fs); P \<turnstile> C sees M,NonStatic:Ts\<rightarrow>T = (pns,body) in D; size vs = size pns; size Ts = size pns \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>(addr a)\<bullet>M(map Val vs), (h,l,sh), b\<rangle> \<rightarrow> \<langle>blocks(this#pns, Class D#Ts, Addr a#vs, body), (h,l,sh), b\<rangle>"
-
-| RedCallNull:
- "P \<turnstile> \<langle>null\<bullet>M(map Val vs),s,b\<rangle> \<rightarrow> \<langle>THROW NullPointer,s,b\<rangle>"
-
-| RedCallNone:
- "\<lbrakk> h a = Some(C,fs); \<not>(\<exists>b Ts T m D. P \<turnstile> C sees M,b:Ts\<rightarrow>T = m in D) \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>(addr a)\<bullet>M(map Val vs),(h,l,sh),b\<rangle> \<rightarrow> \<langle>THROW NoSuchMethodError,(h,l,sh),b\<rangle>"
-
-| RedCallStatic:
- "\<lbrakk> h a = Some(C,fs); P \<turnstile> C sees M,Static:Ts\<rightarrow>T = m in D \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>(addr a)\<bullet>M(map Val vs),(h,l,sh),b\<rangle> \<rightarrow> \<langle>THROW IncompatibleClassChangeError,(h,l,sh),b\<rangle>"
-
-| SCallParams:
- "P \<turnstile> \<langle>es,s,b\<rangle> [\<rightarrow>] \<langle>es',s',b'\<rangle> \<Longrightarrow>
- P \<turnstile> \<langle>C\<bullet>\<^sub>sM(es),s,b\<rangle> \<rightarrow> \<langle>C\<bullet>\<^sub>sM(es'),s',b'\<rangle>"
-
-| RedSCall:
- "\<lbrakk> P \<turnstile> C sees M,Static:Ts\<rightarrow>T = (pns,body) in D;
- length vs = length pns; size Ts = size pns \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sM(map Val vs),s,True\<rangle> \<rightarrow> \<langle>blocks(pns, Ts, vs, body), s, False\<rangle>"
-
-| SCallInitDoneRed:
- "\<lbrakk> P \<turnstile> C sees M,Static:Ts\<rightarrow>T = (pns,body) in D;
- sh D = Some(sfs,Done) \<or> (M = clinit \<and> sh D = Some(sfs,Processing)) \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sM(map Val vs),(h,l,sh), False\<rangle> \<rightarrow> \<langle>C\<bullet>\<^sub>sM(map Val vs),(h,l,sh), True\<rangle>"
-
-| SCallInitRed:
- "\<lbrakk> P \<turnstile> C sees M,Static:Ts\<rightarrow>T = (pns,body) in D;
- \<nexists>sfs. sh D = Some(sfs,Done); M \<noteq> clinit \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sM(map Val vs),(h,l,sh), False\<rangle> \<rightarrow> \<langle>INIT D ([D],False) \<leftarrow> C\<bullet>\<^sub>sM(map Val vs),(h,l,sh),False\<rangle>"
-
-| RedSCallNone:
- "\<lbrakk> \<not>(\<exists>b Ts T m D. P \<turnstile> C sees M,b:Ts\<rightarrow>T = m in D) \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sM(map Val vs),s,b\<rangle> \<rightarrow> \<langle>THROW NoSuchMethodError,s,False\<rangle>"
-
-| RedSCallNonStatic:
- "\<lbrakk> P \<turnstile> C sees M,NonStatic:Ts\<rightarrow>T = m in D \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sM(map Val vs),s,b\<rangle> \<rightarrow> \<langle>THROW IncompatibleClassChangeError,s,False\<rangle>"
-
-| BlockRedNone:
- "\<lbrakk> P \<turnstile> \<langle>e, (h,l(V:=None),sh), b\<rangle> \<rightarrow> \<langle>e', (h',l',sh'), b'\<rangle>; l' V = None; \<not> assigned V e \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>{V:T; e}, (h,l,sh), b\<rangle> \<rightarrow> \<langle>{V:T; e'}, (h',l'(V := l V),sh'), b'\<rangle>"
-
-| BlockRedSome:
- "\<lbrakk> P \<turnstile> \<langle>e, (h,l(V:=None),sh), b\<rangle> \<rightarrow> \<langle>e', (h',l',sh'), b'\<rangle>; l' V = Some v;\<not> assigned V e \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>{V:T; e}, (h,l,sh), b\<rangle> \<rightarrow> \<langle>{V:T := Val v; e'}, (h',l'(V := l V),sh'), b'\<rangle>"
-
-| InitBlockRed:
- "\<lbrakk> P \<turnstile> \<langle>e, (h,l(V\<mapsto>v),sh), b\<rangle> \<rightarrow> \<langle>e', (h',l',sh'), b'\<rangle>; l' V = Some v' \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>{V:T := Val v; e}, (h,l,sh), b\<rangle> \<rightarrow> \<langle>{V:T := Val v'; e'}, (h',l'(V := l V),sh'), b'\<rangle>"
-
-| RedBlock:
- "P \<turnstile> \<langle>{V:T; Val u}, s, b\<rangle> \<rightarrow> \<langle>Val u, s, b\<rangle>"
-
-| RedInitBlock:
- "P \<turnstile> \<langle>{V:T := Val v; Val u}, s, b\<rangle> \<rightarrow> \<langle>Val u, s, b\<rangle>"
-
-| SeqRed:
- "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow> \<langle>e',s',b'\<rangle> \<Longrightarrow>
- P \<turnstile> \<langle>e;;e\<^sub>2, s, b\<rangle> \<rightarrow> \<langle>e';;e\<^sub>2, s', b'\<rangle>"
-
-| RedSeq:
- "P \<turnstile> \<langle>(Val v);;e\<^sub>2, s, b\<rangle> \<rightarrow> \<langle>e\<^sub>2, s, b\<rangle>"
-
-| CondRed:
- "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow> \<langle>e',s',b'\<rangle> \<Longrightarrow>
- P \<turnstile> \<langle>if (e) e\<^sub>1 else e\<^sub>2, s, b\<rangle> \<rightarrow> \<langle>if (e') e\<^sub>1 else e\<^sub>2, s', b'\<rangle>"
-
-| RedCondT:
- "P \<turnstile> \<langle>if (true) e\<^sub>1 else e\<^sub>2, s, b\<rangle> \<rightarrow> \<langle>e\<^sub>1, s, b\<rangle>"
-
-| RedCondF:
- "P \<turnstile> \<langle>if (false) e\<^sub>1 else e\<^sub>2, s, b\<rangle> \<rightarrow> \<langle>e\<^sub>2, s, b\<rangle>"
-
-| RedWhile:
- "P \<turnstile> \<langle>while(b) c, s, b'\<rangle> \<rightarrow> \<langle>if(b) (c;;while(b) c) else unit, s, b'\<rangle>"
-
-| ThrowRed:
- "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow> \<langle>e',s',b'\<rangle> \<Longrightarrow>
- P \<turnstile> \<langle>throw e, s, b\<rangle> \<rightarrow> \<langle>throw e', s', b'\<rangle>"
-
-| RedThrowNull:
- "P \<turnstile> \<langle>throw null, s, b\<rangle> \<rightarrow> \<langle>THROW NullPointer, s, b\<rangle>"
-
-| TryRed:
- "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow> \<langle>e',s',b'\<rangle> \<Longrightarrow>
- P \<turnstile> \<langle>try e catch(C V) e\<^sub>2, s, b\<rangle> \<rightarrow> \<langle>try e' catch(C V) e\<^sub>2, s', b'\<rangle>"
-
-| RedTry:
- "P \<turnstile> \<langle>try (Val v) catch(C V) e\<^sub>2, s, b\<rangle> \<rightarrow> \<langle>Val v, s, b\<rangle>"
-
-| RedTryCatch:
- "\<lbrakk> hp s a = Some(D,fs); P \<turnstile> D \<preceq>\<^sup>* C \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>try (Throw a) catch(C V) e\<^sub>2, s, b\<rangle> \<rightarrow> \<langle>{V:Class C := addr a; e\<^sub>2}, s, b\<rangle>"
-
-| RedTryFail:
- "\<lbrakk> hp s a = Some(D,fs); \<not> P \<turnstile> D \<preceq>\<^sup>* C \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>try (Throw a) catch(C V) e\<^sub>2, s, b\<rangle> \<rightarrow> \<langle>Throw a, s, b\<rangle>"
-
-| ListRed1:
- "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow> \<langle>e',s',b'\<rangle> \<Longrightarrow>
- P \<turnstile> \<langle>e#es,s,b\<rangle> [\<rightarrow>] \<langle>e'#es,s',b'\<rangle>"
-
-| ListRed2:
- "P \<turnstile> \<langle>es,s,b\<rangle> [\<rightarrow>] \<langle>es',s',b'\<rangle> \<Longrightarrow>
- P \<turnstile> \<langle>Val v # es,s,b\<rangle> [\<rightarrow>] \<langle>Val v # es',s',b'\<rangle>"
-
-\<comment> \<open>Initialization procedure\<close>
-
-| RedInit:
- "\<not>sub_RI e \<Longrightarrow> P \<turnstile> \<langle>INIT C (Nil,b) \<leftarrow> e,s,b'\<rangle> \<rightarrow> \<langle>e,s,icheck P C e\<rangle>"
-
-| InitNoneRed:
- "sh C = None
- \<Longrightarrow> P \<turnstile> \<langle>INIT C' (C#Cs,False) \<leftarrow> e,(h,l,sh),b\<rangle> \<rightarrow> \<langle>INIT C' (C#Cs,False) \<leftarrow> e,(h,l,sh(C \<mapsto> (sblank P C, Prepared))),b\<rangle>"
-
-| RedInitDone:
- "sh C = Some(sfs,Done)
- \<Longrightarrow> P \<turnstile> \<langle>INIT C' (C#Cs,False) \<leftarrow> e,(h,l,sh),b\<rangle> \<rightarrow> \<langle>INIT C' (Cs,True) \<leftarrow> e,(h,l,sh),b\<rangle>"
-
-| RedInitProcessing:
- "sh C = Some(sfs,Processing)
- \<Longrightarrow> P \<turnstile> \<langle>INIT C' (C#Cs,False) \<leftarrow> e,(h,l,sh),b\<rangle> \<rightarrow> \<langle>INIT C' (Cs,True) \<leftarrow> e,(h,l,sh),b\<rangle>"
-
-| RedInitError:
- "sh C = Some(sfs,Error)
- \<Longrightarrow> P \<turnstile> \<langle>INIT C' (C#Cs,False) \<leftarrow> e,(h,l,sh),b\<rangle> \<rightarrow> \<langle>RI (C,THROW NoClassDefFoundError);Cs \<leftarrow> e,(h,l,sh),b\<rangle>"
-
-| InitObjectRed:
- "\<lbrakk> sh C = Some(sfs,Prepared);
- C = Object;
- sh' = sh(C \<mapsto> (sfs,Processing)) \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>INIT C' (C#Cs,False) \<leftarrow> e,(h,l,sh),b\<rangle> \<rightarrow> \<langle>INIT C' (C#Cs,True) \<leftarrow> e,(h,l,sh'),b\<rangle>"
-
-| InitNonObjectSuperRed:
- "\<lbrakk> sh C = Some(sfs,Prepared);
- C \<noteq> Object;
- class P C = Some (D,r);
- sh' = sh(C \<mapsto> (sfs,Processing)) \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<langle>INIT C' (C#Cs,False) \<leftarrow> e,(h,l,sh),b\<rangle> \<rightarrow> \<langle>INIT C' (D#C#Cs,False) \<leftarrow> e,(h,l,sh'),b\<rangle>"
-
-| RedInitRInit:
- "P \<turnstile> \<langle>INIT C' (C#Cs,True) \<leftarrow> e,(h,l,sh),b\<rangle> \<rightarrow> \<langle>RI (C,C\<bullet>\<^sub>sclinit([]));Cs \<leftarrow> e,(h,l,sh),b\<rangle>"
-
-| RInitRed:
- "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow> \<langle>e',s',b'\<rangle> \<Longrightarrow>
- P \<turnstile> \<langle>RI (C,e);Cs \<leftarrow> e\<^sub>0, s, b\<rangle> \<rightarrow> \<langle>RI (C,e');Cs \<leftarrow> e\<^sub>0, s', b'\<rangle>"
-
-| RedRInit:
- "\<lbrakk> sh C = Some (sfs, i);
- sh' = sh(C \<mapsto> (sfs,Done));
- C' = last(C#Cs) \<rbrakk> \<Longrightarrow>
- P \<turnstile> \<langle>RI (C, Val v);Cs \<leftarrow> e, (h,l,sh), b\<rangle> \<rightarrow> \<langle>INIT C' (Cs,True) \<leftarrow> e, (h,l,sh'), b\<rangle>"
-
-\<comment> \<open>Exception propagation\<close>
-
-| CastThrow: "P \<turnstile> \<langle>Cast C (throw e), s, b\<rangle> \<rightarrow> \<langle>throw e, s, b\<rangle>"
-| BinOpThrow1: "P \<turnstile> \<langle>(throw e) \<guillemotleft>bop\<guillemotright> e\<^sub>2, s, b\<rangle> \<rightarrow> \<langle>throw e, s, b\<rangle>"
-| BinOpThrow2: "P \<turnstile> \<langle>(Val v\<^sub>1) \<guillemotleft>bop\<guillemotright> (throw e), s, b\<rangle> \<rightarrow> \<langle>throw e, s, b\<rangle>"
-| LAssThrow: "P \<turnstile> \<langle>V:=(throw e), s, b\<rangle> \<rightarrow> \<langle>throw e, s, b\<rangle>"
-| FAccThrow: "P \<turnstile> \<langle>(throw e)\<bullet>F{D}, s, b\<rangle> \<rightarrow> \<langle>throw e, s, b\<rangle>"
-| FAssThrow1: "P \<turnstile> \<langle>(throw e)\<bullet>F{D}:=e\<^sub>2, s, b\<rangle> \<rightarrow> \<langle>throw e, s, b\<rangle>"
-| FAssThrow2: "P \<turnstile> \<langle>Val v\<bullet>F{D}:=(throw e), s, b\<rangle> \<rightarrow> \<langle>throw e, s, b\<rangle>"
-| SFAssThrow: "P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D}:=(throw e), s, b\<rangle> \<rightarrow> \<langle>throw e, s, b\<rangle>"
-| CallThrowObj: "P \<turnstile> \<langle>(throw e)\<bullet>M(es), s, b\<rangle> \<rightarrow> \<langle>throw e, s, b\<rangle>"
-| CallThrowParams: "\<lbrakk> es = map Val vs @ throw e # es' \<rbrakk> \<Longrightarrow> P \<turnstile> \<langle>(Val v)\<bullet>M(es), s, b\<rangle> \<rightarrow> \<langle>throw e, s, b\<rangle>"
-| SCallThrowParams: "\<lbrakk> es = map Val vs @ throw e # es' \<rbrakk> \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sM(es), s, b\<rangle> \<rightarrow> \<langle>throw e, s, b\<rangle>"
-| BlockThrow: "P \<turnstile> \<langle>{V:T; Throw a}, s, b\<rangle> \<rightarrow> \<langle>Throw a, s, b\<rangle>"
-| InitBlockThrow: "P \<turnstile> \<langle>{V:T := Val v; Throw a}, s, b\<rangle> \<rightarrow> \<langle>Throw a, s, b\<rangle>"
-| SeqThrow: "P \<turnstile> \<langle>(throw e);;e\<^sub>2, s, b\<rangle> \<rightarrow> \<langle>throw e, s, b\<rangle>"
-| CondThrow: "P \<turnstile> \<langle>if (throw e) e\<^sub>1 else e\<^sub>2, s, b\<rangle> \<rightarrow> \<langle>throw e, s, b\<rangle>"
-| ThrowThrow: "P \<turnstile> \<langle>throw(throw e), s, b\<rangle> \<rightarrow> \<langle>throw e, s, b\<rangle>"
-| RInitInitThrow: "\<lbrakk> sh C = Some(sfs,i); sh' = sh(C \<mapsto> (sfs,Error)) \<rbrakk> \<Longrightarrow>
- P \<turnstile> \<langle>RI (C,Throw a);D#Cs \<leftarrow> e,(h,l,sh),b\<rangle> \<rightarrow> \<langle>RI (D,Throw a);Cs \<leftarrow> e,(h,l,sh'),b\<rangle>"
-| RInitThrow: "\<lbrakk> sh C = Some(sfs, i); sh' = sh(C \<mapsto> (sfs,Error)) \<rbrakk> \<Longrightarrow>
- P \<turnstile> \<langle>RI (C,Throw a);Nil \<leftarrow> e,(h,l,sh),b\<rangle> \<rightarrow> \<langle>Throw a,(h,l,sh'),b\<rangle>"
-(*<*)
-lemmas red_reds_induct = red_reds.induct [split_format (complete)]
- and red_reds_inducts = red_reds.inducts [split_format (complete)]
-
-inductive_cases [elim!]:
- "P \<turnstile> \<langle>V:=e,s,b\<rangle> \<rightarrow> \<langle>e',s',b'\<rangle>"
- "P \<turnstile> \<langle>e1;;e2,s,b\<rangle> \<rightarrow> \<langle>e',s',b'\<rangle>"
-(*>*)
-
-subsection\<open> The reflexive transitive closure \<close>
-
-abbreviation
- Step :: "J_prog \<Rightarrow> expr \<Rightarrow> state \<Rightarrow> bool \<Rightarrow> expr \<Rightarrow> state \<Rightarrow> bool \<Rightarrow> bool"
- ("_ \<turnstile> ((1\<langle>_,/_,/_\<rangle>) \<rightarrow>*/ (1\<langle>_,/_,/_\<rangle>))" [51,0,0,0,0,0,0] 81)
- where "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle> \<equiv> ((e,s,b), e',s',b') \<in> (red P)\<^sup>*"
-
-abbreviation
- Steps :: "J_prog \<Rightarrow> expr list \<Rightarrow> state \<Rightarrow> bool \<Rightarrow> expr list \<Rightarrow> state \<Rightarrow> bool \<Rightarrow> bool"
- ("_ \<turnstile> ((1\<langle>_,/_,/_\<rangle>) [\<rightarrow>]*/ (1\<langle>_,/_,/_\<rangle>))" [51,0,0,0,0,0,0] 81)
- where "P \<turnstile> \<langle>es,s,b\<rangle> [\<rightarrow>]* \<langle>es',s',b'\<rangle> \<equiv> ((es,s,b), es',s',b') \<in> (reds P)\<^sup>*"
-
-
-lemmas converse_rtrancl_induct3 =
- converse_rtrancl_induct [of "(ax, ay, az)" "(bx, by, bz)", split_format (complete),
- consumes 1, case_names refl step]
-
-lemma converse_rtrancl_induct_red[consumes 1]:
-assumes "P \<turnstile> \<langle>e,(h,l,sh),b\<rangle> \<rightarrow>* \<langle>e',(h',l',sh'),b'\<rangle>"
-and "\<And>e h l sh b. R e h l sh b e h l sh b"
-and "\<And>e\<^sub>0 h\<^sub>0 l\<^sub>0 sh\<^sub>0 b\<^sub>0 e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 b\<^sub>1 e' h' l' sh' b'.
- \<lbrakk> P \<turnstile> \<langle>e\<^sub>0,(h\<^sub>0,l\<^sub>0,sh\<^sub>0),b\<^sub>0\<rangle> \<rightarrow> \<langle>e\<^sub>1,(h\<^sub>1,l\<^sub>1,sh\<^sub>1),b\<^sub>1\<rangle>; R e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 b\<^sub>1 e' h' l' sh' b' \<rbrakk>
- \<Longrightarrow> R e\<^sub>0 h\<^sub>0 l\<^sub>0 sh\<^sub>0 b\<^sub>0 e' h' l' sh' b'"
-shows "R e h l sh b e' h' l' sh' b'"
-(*<*)
-proof -
- { fix s s'
- assume reds: "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle>"
- and base: "\<And>e s b. R e (hp s) (lcl s) (shp s) b e (hp s) (lcl s) (shp s) b"
- and red\<^sub>1: "\<And>e\<^sub>0 s\<^sub>0 b\<^sub>0 e\<^sub>1 s\<^sub>1 b\<^sub>1 e' s' b'.
- \<lbrakk> P \<turnstile> \<langle>e\<^sub>0,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow> \<langle>e\<^sub>1,s\<^sub>1,b\<^sub>1\<rangle>; R e\<^sub>1 (hp s\<^sub>1) (lcl s\<^sub>1) (shp s\<^sub>1) b\<^sub>1 e' (hp s') (lcl s') (shp s') b' \<rbrakk>
- \<Longrightarrow> R e\<^sub>0 (hp s\<^sub>0) (lcl s\<^sub>0) (shp s\<^sub>0) b\<^sub>0 e' (hp s') (lcl s') (shp s') b'"
- from reds have "R e (hp s) (lcl s) (shp s) b e' (hp s') (lcl s') (shp s') b'"
- proof (induct rule:converse_rtrancl_induct3)
- case refl show ?case by(rule base)
- next
- case step
- thus ?case by(blast intro:red\<^sub>1)
- qed
- }
- with assms show ?thesis by fastforce
-qed
-(*>*)
-
-
-subsection\<open>Some easy lemmas\<close>
-
-lemma [iff]: "\<not> P \<turnstile> \<langle>[],s,b\<rangle> [\<rightarrow>] \<langle>es',s',b'\<rangle>"
-(*<*)by(blast elim: reds.cases)(*>*)
-
-lemma [iff]: "\<not> P \<turnstile> \<langle>Val v,s,b\<rangle> \<rightarrow> \<langle>e',s',b'\<rangle>"
-(*<*)by(fastforce elim: red.cases)(*>*)
-
-lemma val_no_step: "val_of e = \<lfloor>v\<rfloor> \<Longrightarrow> \<not> P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow> \<langle>e',s',b'\<rangle>"
-(*<*)by(drule val_of_spec, simp)(*>*)
-
-lemma [iff]: "\<not> P \<turnstile> \<langle>Throw a,s,b\<rangle> \<rightarrow> \<langle>e',s',b'\<rangle>"
-(*<*)by(fastforce elim: red.cases)(*>*)
-
-
-lemma map_Vals_no_step [iff]: "\<not> P \<turnstile> \<langle>map Val vs,s,b\<rangle> [\<rightarrow>] \<langle>es',s',b'\<rangle>"
-(*<*)
-proof(induct vs arbitrary: es')
- case (Cons a vs)
- { assume "P \<turnstile> \<langle>map Val (a # vs),s,b\<rangle> [\<rightarrow>] \<langle>es',s',b'\<rangle>"
- then have False by(rule reds.cases) (insert Cons, auto)
- }
- then show ?case by clarsimp
-qed simp
-(*>*)
-
-lemma vals_no_step: "map_vals_of es = \<lfloor>vs\<rfloor> \<Longrightarrow> \<not> P \<turnstile> \<langle>es,s,b\<rangle> [\<rightarrow>] \<langle>es',s',b'\<rangle>"
-(*<*)by(drule map_vals_of_spec, simp)(*>*)
-
-lemma vals_throw_no_step [iff]: "\<not> P \<turnstile> \<langle>map Val vs @ Throw a # es,s,b\<rangle> [\<rightarrow>] \<langle>es',s',b'\<rangle>"
-(*<*)
-proof(induct vs arbitrary: es')
- case Nil
- { assume "P \<turnstile> \<langle>Throw a # es,s,b\<rangle> [\<rightarrow>] \<langle>es',s',b'\<rangle>"
- then have False by(rule reds.cases) (insert Cons, auto)
- }
- then show ?case by clarsimp
-next
- case (Cons a' vs')
- { assume "P \<turnstile> \<langle>map Val (a'#vs') @ Throw a # es,s,b\<rangle> [\<rightarrow>] \<langle>es',s',b'\<rangle>"
- then have False by(rule reds.cases) (insert Cons, auto)
- }
- then show ?case by clarsimp
-qed
-(*>*)
-
-lemma lass_val_of_red:
- "\<lbrakk> lass_val_of e = \<lfloor>a\<rfloor>; P \<turnstile> \<langle>e,(h, l, sh),b\<rangle> \<rightarrow> \<langle>e',(h', l', sh'),b'\<rangle> \<rbrakk>
- \<Longrightarrow> e' = unit \<and> h' = h \<and> l' = l(fst a\<mapsto>snd a) \<and> sh' = sh \<and> b = b'"
-(*<*)by(drule lass_val_of_spec, auto)(*>*)
-
-
-lemma final_no_step [iff]: "final e \<Longrightarrow> \<not> P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow> \<langle>e',s',b'\<rangle>"
-(*<*)by(erule finalE, simp+)(*>*)
-
-lemma finals_no_step [iff]: "finals es \<Longrightarrow> \<not> P \<turnstile> \<langle>es,s,b\<rangle> [\<rightarrow>] \<langle>es',s',b'\<rangle>"
-(*<*)by(erule finalsE, simp+)(*>*)
-
-lemma reds_final_same:
-"P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle> \<Longrightarrow> final e \<Longrightarrow> e = e' \<and> s = s' \<and> b = b'"
-proof(induct rule:converse_rtrancl_induct3)
- case refl show ?case by simp
-next
- case (step e0 s0 b0 e1 s1 b1) show ?case
- proof(rule finalE[OF step.prems(1)])
- fix v assume "e0 = Val v" then show ?thesis using step by simp
- next
- fix a assume "e0 = Throw a" then show ?thesis using step by simp
- qed
-qed
-
-lemma reds_throw:
-"P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle> \<Longrightarrow> (\<And>e\<^sub>t. throw_of e = \<lfloor>e\<^sub>t\<rfloor> \<Longrightarrow> \<exists>e\<^sub>t'. throw_of e' = \<lfloor>e\<^sub>t'\<rfloor>)"
-proof(induct rule:converse_rtrancl_induct3)
- case refl then show ?case by simp
-next
- case (step e0 s0 b0 e1 s1 b1)
- then show ?case by(auto elim: red.cases)
-qed
-
-lemma red_hext_incr: "P \<turnstile> \<langle>e,(h,l,sh),b\<rangle> \<rightarrow> \<langle>e',(h',l',sh'),b'\<rangle> \<Longrightarrow> h \<unlhd> h'"
- and reds_hext_incr: "P \<turnstile> \<langle>es,(h,l,sh),b\<rangle> [\<rightarrow>] \<langle>es',(h',l',sh'),b'\<rangle> \<Longrightarrow> h \<unlhd> h'"
-(*<*)
-proof(induct rule:red_reds_inducts)
- case RedNew thus ?case
- by(fastforce dest:new_Addr_SomeD simp:hext_def split:if_splits)
-next
- case RedFAss thus ?case by(simp add:hext_def split:if_splits)
-qed simp_all
-(*>*)
-
-
-lemma red_lcl_incr: "P \<turnstile> \<langle>e,(h\<^sub>0,l\<^sub>0,sh\<^sub>0),b\<rangle> \<rightarrow> \<langle>e',(h\<^sub>1,l\<^sub>1,sh\<^sub>1),b'\<rangle> \<Longrightarrow> dom l\<^sub>0 \<subseteq> dom l\<^sub>1"
-and reds_lcl_incr: "P \<turnstile> \<langle>es,(h\<^sub>0,l\<^sub>0,sh\<^sub>0),b\<rangle> [\<rightarrow>] \<langle>es',(h\<^sub>1,l\<^sub>1,sh\<^sub>1),b'\<rangle> \<Longrightarrow> dom l\<^sub>0 \<subseteq> dom l\<^sub>1"
-(*<*)by(induct rule: red_reds_inducts)(auto simp del:fun_upd_apply)(*>*)
-
-lemma red_lcl_add: "P \<turnstile> \<langle>e,(h,l,sh),b\<rangle> \<rightarrow> \<langle>e',(h',l',sh'),b'\<rangle> \<Longrightarrow> (\<And>l\<^sub>0. P \<turnstile> \<langle>e,(h,l\<^sub>0++l,sh),b\<rangle> \<rightarrow> \<langle>e',(h',l\<^sub>0++l',sh'),b'\<rangle>)"
-and reds_lcl_add: "P \<turnstile> \<langle>es,(h,l,sh),b\<rangle> [\<rightarrow>] \<langle>es',(h',l',sh'),b'\<rangle> \<Longrightarrow> (\<And>l\<^sub>0. P \<turnstile> \<langle>es,(h,l\<^sub>0++l,sh),b\<rangle> [\<rightarrow>] \<langle>es',(h',l\<^sub>0++l',sh'),b'\<rangle>)"
-(*<*)
-proof (induct rule:red_reds_inducts)
- case RedCast thus ?case by(fastforce intro:red_reds.intros)
-next
- case RedCastFail thus ?case by(force intro:red_reds.intros)
-next
- case RedFAcc thus ?case by(fastforce intro:red_reds.intros)
-next
- case RedCall thus ?case by(fastforce intro:red_reds.intros)
-next
- case (InitBlockRed e h l V v sh b e' h' l' sh' b' v' T l\<^sub>0)
- have IH: "\<And>l\<^sub>0. P \<turnstile> \<langle>e,(h, l\<^sub>0 ++ l(V \<mapsto> v), sh),b\<rangle> \<rightarrow> \<langle>e',(h', l\<^sub>0 ++ l', sh'),b'\<rangle>"
- and l'V: "l' V = Some v'" by fact+
- from IH have IH': "P \<turnstile> \<langle>e,(h, (l\<^sub>0 ++ l)(V \<mapsto> v), sh),b\<rangle> \<rightarrow> \<langle>e',(h', l\<^sub>0 ++ l', sh'),b'\<rangle>"
- by simp
- have "(l\<^sub>0 ++ l')(V := (l\<^sub>0 ++ l) V) = l\<^sub>0 ++ l'(V := l V)"
- by(rule ext)(simp add:map_add_def)
- with red_reds.InitBlockRed[OF IH'] l'V show ?case by(simp del:fun_upd_apply)
-next
- case (BlockRedNone e h l V sh b e' h' l' sh' b' T l\<^sub>0)
- have IH: "\<And>l\<^sub>0. P \<turnstile> \<langle>e,(h, l\<^sub>0 ++ l(V := None), sh),b\<rangle> \<rightarrow> \<langle>e',(h', l\<^sub>0 ++ l', sh'),b'\<rangle>"
- and l'V: "l' V = None" and unass: "\<not> assigned V e" by fact+
- have "l\<^sub>0(V := None) ++ l(V := None) = (l\<^sub>0 ++ l)(V := None)"
- by(simp add:fun_eq_iff map_add_def)
- hence IH': "P \<turnstile> \<langle>e,(h, (l\<^sub>0++l)(V := None), sh),b\<rangle> \<rightarrow> \<langle>e',(h', l\<^sub>0(V := None) ++ l', sh'),b'\<rangle>"
- using IH[of "l\<^sub>0(V := None)"] by simp
- have "(l\<^sub>0(V := None) ++ l')(V := (l\<^sub>0 ++ l) V) = l\<^sub>0 ++ l'(V := l V)"
- by(simp add:fun_eq_iff map_add_def)
- with red_reds.BlockRedNone[OF IH' _ unass] l'V show ?case
- by(simp add: map_add_def)
-next
- case (BlockRedSome e h l V sh b e' h' l' sh' b' v T l\<^sub>0)
- have IH: "\<And>l\<^sub>0. P \<turnstile> \<langle>e,(h, l\<^sub>0 ++ l(V := None), sh),b\<rangle> \<rightarrow> \<langle>e',(h', l\<^sub>0 ++ l', sh'),b'\<rangle>"
- and l'V: "l' V = Some v" and unass: "\<not> assigned V e" by fact+
- have "l\<^sub>0(V := None) ++ l(V := None) = (l\<^sub>0 ++ l)(V := None)"
- by(simp add:fun_eq_iff map_add_def)
- hence IH': "P \<turnstile> \<langle>e,(h, (l\<^sub>0++l)(V := None), sh),b\<rangle> \<rightarrow> \<langle>e',(h', l\<^sub>0(V := None) ++ l', sh'),b'\<rangle>"
- using IH[of "l\<^sub>0(V := None)"] by simp
- have "(l\<^sub>0(V := None) ++ l')(V := (l\<^sub>0 ++ l) V) = l\<^sub>0 ++ l'(V := l V)"
- by(simp add:fun_eq_iff map_add_def)
- with red_reds.BlockRedSome[OF IH' _ unass] l'V show ?case
- by(simp add:map_add_def)
-next
- case RedTryCatch thus ?case by(fastforce intro:red_reds.intros)
-next
- case RedTryFail thus ?case by(force intro!:red_reds.intros)
-qed (simp_all add:red_reds.intros)
-(*>*)
-
-
-lemma Red_lcl_add:
-assumes "P \<turnstile> \<langle>e,(h,l,sh), b\<rangle> \<rightarrow>* \<langle>e',(h',l',sh'), b'\<rangle>" shows "P \<turnstile> \<langle>e,(h,l\<^sub>0++l,sh),b\<rangle> \<rightarrow>* \<langle>e',(h',l\<^sub>0++l',sh'),b'\<rangle>"
-(*<*)
-using assms
-proof(induct rule:converse_rtrancl_induct_red)
- case 1 thus ?case by simp
-next
- case 2 thus ?case
- by (blast dest: red_lcl_add intro: converse_rtrancl_into_rtrancl)
-qed
-(*>*)
-
-lemma assumes wf: "wwf_J_prog P"
-shows red_proc_pres: "P \<turnstile> \<langle>e,(h,l,sh),b\<rangle> \<rightarrow> \<langle>e',(h',l',sh'),b'\<rangle>
- \<Longrightarrow> not_init C e \<Longrightarrow> sh C = \<lfloor>(sfs, Processing)\<rfloor> \<Longrightarrow> not_init C e' \<and> (\<exists>sfs'. sh' C = \<lfloor>(sfs', Processing)\<rfloor>)"
- and reds_proc_pres: "P \<turnstile> \<langle>es,(h,l,sh),b\<rangle> [\<rightarrow>] \<langle>es',(h',l',sh'),b'\<rangle>
- \<Longrightarrow> not_inits C es \<Longrightarrow> sh C = \<lfloor>(sfs, Processing)\<rfloor> \<Longrightarrow> not_inits C es' \<and> (\<exists>sfs'. sh' C = \<lfloor>(sfs', Processing)\<rfloor>)"
-(*<*)
-proof(induct rule:red_reds_inducts)
- case RedCall then show ?case
- using sees_wwf_nsub_RI[OF wf RedCall.hyps(2)] sub_RI_blocks_body nsub_RI_not_init by auto
-next
- case RedSCall then show ?case
- using sees_wwf_nsub_RI[OF wf RedSCall.hyps(1)] sub_RI_blocks_body nsub_RI_not_init by auto
-next
- case (RedInitDone sh C sfs C' Cs e h l b)
- then show ?case by(cases Cs, auto)
-next
- case (RedInitProcessing sh C sfs C' Cs e h l b)
- then show ?case by(cases Cs, auto)
-next
- case (RedRInit sh C sfs i sh' C' Cs v e h l b)
- then show ?case by(cases Cs, auto)
-next
- case (CallThrowParams es vs e es' v M h l sh b)
- then show ?case by(auto dest: not_inits_def')
-next
- case (SCallThrowParams es vs e es' C M h l sh b)
- then show ?case by(auto dest: not_inits_def')
-qed(auto)
-
-end
+(* Title: JinjaDCI/J/SmallStep.thy
+ Author: Tobias Nipkow, Susannah Mansky
+ Copyright 2003 Technische Universitaet Muenchen, 2019-20 UIUC
+
+ Based on the Jinja theory J/SmallStep.thy by Tobias Nipkow
+*)
+
+section \<open> Small Step Semantics \<close>
+
+theory SmallStep
+imports Expr State WWellForm
+begin
+
+fun blocks :: "vname list * ty list * val list * expr \<Rightarrow> expr"
+where
+ "blocks(V#Vs, T#Ts, v#vs, e) = {V:T := Val v; blocks(Vs,Ts,vs,e)}"
+|"blocks([],[],[],e) = e"
+
+lemmas blocks_induct = blocks.induct[split_format (complete)]
+
+lemma [simp]:
+ "\<lbrakk> size vs = size Vs; size Ts = size Vs \<rbrakk> \<Longrightarrow> fv(blocks(Vs,Ts,vs,e)) = fv e - set Vs"
+(*<*)
+by (induct rule:blocks_induct) auto
+(*>*)
+
+
+lemma sub_RI_blocks_body[iff]: "length vs = length pns \<Longrightarrow> length Ts = length pns
+ \<Longrightarrow> sub_RI (blocks (pns, Ts, vs, body)) \<longleftrightarrow> sub_RI body"
+proof(induct pns arbitrary: Ts vs)
+ case Nil then show ?case by simp
+next
+ case Cons then show ?case by(cases vs; cases Ts) auto
+qed
+
+
+definition assigned :: "'a \<Rightarrow> 'a exp \<Rightarrow> bool"
+where
+ "assigned V e \<equiv> \<exists>v e'. e = (V := Val v;; e')"
+
+\<comment> \<open> expression is okay to go the right side of @{text INIT} or @{text "RI \<leftarrow>"}
+ or to have indicator Boolean be True (in latter case, given that class is
+ also verified initialized) \<close>
+fun icheck :: "'m prog \<Rightarrow> cname \<Rightarrow> 'a exp \<Rightarrow> bool" where
+"icheck P C' (new C) = (C' = C)" |
+"icheck P D' (C\<bullet>\<^sub>sF{D}) = ((D' = D) \<and> (\<exists>T. P \<turnstile> C has F,Static:T in D))" |
+"icheck P D' (C\<bullet>\<^sub>sF{D}:=(Val v)) = ((D' = D) \<and> (\<exists>T. P \<turnstile> C has F,Static:T in D))" |
+"icheck P D (C\<bullet>\<^sub>sM(es)) = ((\<exists>vs. es = map Val vs) \<and> (\<exists>Ts T m. P \<turnstile> C sees M,Static:Ts\<rightarrow>T = m in D))" |
+"icheck P _ _ = False"
+
+lemma nicheck_SFAss_nonVal: "val_of e\<^sub>2 = None \<Longrightarrow> \<not>icheck P C' (C\<bullet>\<^sub>sF{D} := (e\<^sub>2::'a exp))"
+ by(rule notI, cases e\<^sub>2, auto)
+
+inductive_set
+ red :: "J_prog \<Rightarrow> ((expr \<times> state \<times> bool) \<times> (expr \<times> state \<times> bool)) set"
+ and reds :: "J_prog \<Rightarrow> ((expr list \<times> state \<times> bool) \<times> (expr list \<times> state \<times> bool)) set"
+ and red' :: "J_prog \<Rightarrow> expr \<Rightarrow> state \<Rightarrow> bool \<Rightarrow> expr \<Rightarrow> state \<Rightarrow> bool \<Rightarrow> bool"
+ ("_ \<turnstile> ((1\<langle>_,/_,/_\<rangle>) \<rightarrow>/ (1\<langle>_,/_,/_\<rangle>))" [51,0,0,0,0,0,0] 81)
+ and reds' :: "J_prog \<Rightarrow> expr list \<Rightarrow> state \<Rightarrow> bool \<Rightarrow> expr list \<Rightarrow> state \<Rightarrow> bool \<Rightarrow> bool"
+ ("_ \<turnstile> ((1\<langle>_,/_,/_\<rangle>) [\<rightarrow>]/ (1\<langle>_,/_,/_\<rangle>))" [51,0,0,0,0,0,0] 81)
+ for P :: J_prog
+where
+
+ "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow> \<langle>e',s',b'\<rangle> \<equiv> ((e,s,b), e',s',b') \<in> red P"
+| "P \<turnstile> \<langle>es,s,b\<rangle> [\<rightarrow>] \<langle>es',s',b'\<rangle> \<equiv> ((es,s,b), es',s',b') \<in> reds P"
+
+| RedNew:
+ "\<lbrakk> new_Addr h = Some a; P \<turnstile> C has_fields FDTs; h' = h(a\<mapsto>blank P C) \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>new C, (h,l,sh), True\<rangle> \<rightarrow> \<langle>addr a, (h',l,sh), False\<rangle>"
+
+| RedNewFail:
+ "\<lbrakk> new_Addr h = None; is_class P C \<rbrakk> \<Longrightarrow>
+ P \<turnstile> \<langle>new C, (h,l,sh), True\<rangle> \<rightarrow> \<langle>THROW OutOfMemory, (h,l,sh), False\<rangle>"
+
+| NewInitDoneRed:
+ "sh C = Some (sfs, Done) \<Longrightarrow>
+ P \<turnstile> \<langle>new C, (h,l,sh), False\<rangle> \<rightarrow> \<langle>new C, (h,l,sh), True\<rangle>"
+
+| NewInitRed:
+ "\<lbrakk> \<nexists>sfs. sh C = Some (sfs, Done); is_class P C \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>new C,(h,l,sh),False\<rangle> \<rightarrow> \<langle>INIT C ([C],False) \<leftarrow> new C,(h,l,sh),False\<rangle>"
+
+| CastRed:
+ "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow> \<langle>e',s',b'\<rangle> \<Longrightarrow>
+ P \<turnstile> \<langle>Cast C e, s, b\<rangle> \<rightarrow> \<langle>Cast C e', s', b'\<rangle>"
+
+| RedCastNull:
+ "P \<turnstile> \<langle>Cast C null, s, b\<rangle> \<rightarrow> \<langle>null,s,b\<rangle>"
+
+| RedCast:
+ "\<lbrakk> h a = Some(D,fs); P \<turnstile> D \<preceq>\<^sup>* C \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>Cast C (addr a), (h,l,sh), b\<rangle> \<rightarrow> \<langle>addr a, (h,l,sh), b\<rangle>"
+
+| RedCastFail:
+ "\<lbrakk> h a = Some(D,fs); \<not> P \<turnstile> D \<preceq>\<^sup>* C \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>Cast C (addr a), (h,l,sh), b\<rangle> \<rightarrow> \<langle>THROW ClassCast, (h,l,sh), b\<rangle>"
+
+| BinOpRed1:
+ "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow> \<langle>e',s',b'\<rangle> \<Longrightarrow>
+ P \<turnstile> \<langle>e \<guillemotleft>bop\<guillemotright> e\<^sub>2, s, b\<rangle> \<rightarrow> \<langle>e' \<guillemotleft>bop\<guillemotright> e\<^sub>2, s', b'\<rangle>"
+
+| BinOpRed2:
+ "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow> \<langle>e',s',b'\<rangle> \<Longrightarrow>
+ P \<turnstile> \<langle>(Val v\<^sub>1) \<guillemotleft>bop\<guillemotright> e, s, b\<rangle> \<rightarrow> \<langle>(Val v\<^sub>1) \<guillemotleft>bop\<guillemotright> e', s', b'\<rangle>"
+
+| RedBinOp:
+ "binop(bop,v\<^sub>1,v\<^sub>2) = Some v \<Longrightarrow>
+ P \<turnstile> \<langle>(Val v\<^sub>1) \<guillemotleft>bop\<guillemotright> (Val v\<^sub>2), s, b\<rangle> \<rightarrow> \<langle>Val v,s,b\<rangle>"
+
+| RedVar:
+ "l V = Some v \<Longrightarrow>
+ P \<turnstile> \<langle>Var V,(h,l,sh),b\<rangle> \<rightarrow> \<langle>Val v,(h,l,sh),b\<rangle>"
+
+| LAssRed:
+ "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow> \<langle>e',s',b'\<rangle> \<Longrightarrow>
+ P \<turnstile> \<langle>V:=e,s,b\<rangle> \<rightarrow> \<langle>V:=e',s',b'\<rangle>"
+
+| RedLAss:
+ "P \<turnstile> \<langle>V:=(Val v), (h,l,sh), b\<rangle> \<rightarrow> \<langle>unit, (h,l(V\<mapsto>v),sh), b\<rangle>"
+
+| FAccRed:
+ "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow> \<langle>e',s',b'\<rangle> \<Longrightarrow>
+ P \<turnstile> \<langle>e\<bullet>F{D}, s, b\<rangle> \<rightarrow> \<langle>e'\<bullet>F{D}, s', b'\<rangle>"
+
+| RedFAcc:
+ "\<lbrakk> h a = Some(C,fs); fs(F,D) = Some v;
+ P \<turnstile> C has F,NonStatic:t in D \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>(addr a)\<bullet>F{D}, (h,l,sh), b\<rangle> \<rightarrow> \<langle>Val v,(h,l,sh),b\<rangle>"
+
+| RedFAccNull:
+ "P \<turnstile> \<langle>null\<bullet>F{D}, s, b\<rangle> \<rightarrow> \<langle>THROW NullPointer, s, b\<rangle>"
+
+| RedFAccNone:
+ "\<lbrakk> h a = Some(C,fs); \<not>(\<exists>b t. P \<turnstile> C has F,b:t in D) \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>(addr a)\<bullet>F{D},(h,l,sh),b\<rangle> \<rightarrow> \<langle>THROW NoSuchFieldError,(h,l,sh),b\<rangle>"
+
+| RedFAccStatic:
+ "\<lbrakk> h a = Some(C,fs); P \<turnstile> C has F,Static:t in D \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>(addr a)\<bullet>F{D},(h,l,sh),b\<rangle> \<rightarrow> \<langle>THROW IncompatibleClassChangeError,(h,l,sh),b\<rangle>"
+
+| RedSFAcc:
+ "\<lbrakk> P \<turnstile> C has F,Static:t in D;
+ sh D = Some (sfs,i);
+ sfs F = Some v \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D},(h,l,sh),True\<rangle> \<rightarrow> \<langle>Val v,(h,l,sh),False\<rangle>"
+
+| SFAccInitDoneRed:
+ "\<lbrakk> P \<turnstile> C has F,Static:t in D;
+ sh D = Some (sfs,Done) \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D},(h,l,sh),False\<rangle> \<rightarrow> \<langle>C\<bullet>\<^sub>sF{D},(h,l,sh),True\<rangle>"
+
+| SFAccInitRed:
+ "\<lbrakk> P \<turnstile> C has F,Static:t in D;
+ \<nexists>sfs. sh D = Some (sfs,Done) \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D},(h,l,sh),False\<rangle> \<rightarrow> \<langle>INIT D ([D],False) \<leftarrow> C\<bullet>\<^sub>sF{D},(h,l,sh),False\<rangle>"
+
+| RedSFAccNone:
+ "\<not>(\<exists>b t. P \<turnstile> C has F,b:t in D)
+ \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D},(h,l,sh),b\<rangle> \<rightarrow> \<langle>THROW NoSuchFieldError,(h,l,sh),False\<rangle>"
+
+| RedSFAccNonStatic:
+ "P \<turnstile> C has F,NonStatic:t in D
+ \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D},(h,l,sh),b\<rangle> \<rightarrow> \<langle>THROW IncompatibleClassChangeError,(h,l,sh),False\<rangle>"
+
+| FAssRed1:
+ "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow> \<langle>e',s',b'\<rangle> \<Longrightarrow>
+ P \<turnstile> \<langle>e\<bullet>F{D}:=e\<^sub>2, s, b\<rangle> \<rightarrow> \<langle>e'\<bullet>F{D}:=e\<^sub>2, s', b'\<rangle>"
+
+| FAssRed2:
+ "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow> \<langle>e',s',b'\<rangle> \<Longrightarrow>
+ P \<turnstile> \<langle>Val v\<bullet>F{D}:=e, s, b\<rangle> \<rightarrow> \<langle>Val v\<bullet>F{D}:=e', s', b'\<rangle>"
+
+| RedFAss:
+ "\<lbrakk> P \<turnstile> C has F,NonStatic:t in D; h a = Some(C,fs) \<rbrakk> \<Longrightarrow>
+ P \<turnstile> \<langle>(addr a)\<bullet>F{D}:=(Val v), (h,l,sh), b\<rangle> \<rightarrow> \<langle>unit, (h(a \<mapsto> (C,fs((F,D) \<mapsto> v))),l,sh), b\<rangle>"
+
+| RedFAssNull:
+ "P \<turnstile> \<langle>null\<bullet>F{D}:=Val v, s, b\<rangle> \<rightarrow> \<langle>THROW NullPointer, s, b\<rangle>"
+
+| RedFAssNone:
+ "\<lbrakk> h a = Some(C,fs); \<not>(\<exists>b t. P \<turnstile> C has F,b:t in D) \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>(addr a)\<bullet>F{D}:=(Val v),(h,l,sh),b\<rangle> \<rightarrow> \<langle>THROW NoSuchFieldError,(h,l,sh),b\<rangle>"
+
+| RedFAssStatic:
+ "\<lbrakk> h a = Some(C,fs); P \<turnstile> C has F,Static:t in D \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>(addr a)\<bullet>F{D}:=(Val v),(h,l,sh),b\<rangle> \<rightarrow> \<langle>THROW IncompatibleClassChangeError,(h,l,sh),b\<rangle>"
+
+| SFAssRed:
+ "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow> \<langle>e',s',b'\<rangle> \<Longrightarrow>
+ P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D}:=e, s, b\<rangle> \<rightarrow> \<langle>C\<bullet>\<^sub>sF{D}:=e', s', b'\<rangle>"
+
+| RedSFAss:
+ "\<lbrakk> P \<turnstile> C has F,Static:t in D;
+ sh D = Some(sfs,i);
+ sfs' = sfs(F\<mapsto>v); sh' = sh(D\<mapsto>(sfs',i)) \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D}:=(Val v),(h,l,sh),True\<rangle> \<rightarrow> \<langle>unit,(h,l,sh'),False\<rangle>"
+
+| SFAssInitDoneRed:
+ "\<lbrakk> P \<turnstile> C has F,Static:t in D;
+ sh D = Some(sfs,Done) \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D}:=(Val v),(h,l,sh),False\<rangle> \<rightarrow> \<langle>C\<bullet>\<^sub>sF{D}:=(Val v),(h,l,sh),True\<rangle>"
+
+| SFAssInitRed:
+ "\<lbrakk> P \<turnstile> C has F,Static:t in D;
+ \<nexists>sfs. sh D = Some(sfs,Done) \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D}:=(Val v),(h,l,sh),False\<rangle> \<rightarrow> \<langle>INIT D ([D],False)\<leftarrow> C\<bullet>\<^sub>sF{D}:=(Val v),(h,l,sh),False\<rangle>"
+
+| RedSFAssNone:
+ "\<not>(\<exists>b t. P \<turnstile> C has F,b:t in D)
+ \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D}:=(Val v),s,b\<rangle> \<rightarrow> \<langle>THROW NoSuchFieldError,s,False\<rangle>"
+
+| RedSFAssNonStatic:
+ "P \<turnstile> C has F,NonStatic:t in D
+ \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D}:=(Val v),s,b\<rangle> \<rightarrow> \<langle>THROW IncompatibleClassChangeError,s,False\<rangle>"
+
+| CallObj:
+ "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow> \<langle>e',s',b'\<rangle> \<Longrightarrow>
+ P \<turnstile> \<langle>e\<bullet>M(es),s,b\<rangle> \<rightarrow> \<langle>e'\<bullet>M(es),s',b'\<rangle>"
+
+| CallParams:
+ "P \<turnstile> \<langle>es,s,b\<rangle> [\<rightarrow>] \<langle>es',s',b'\<rangle> \<Longrightarrow>
+ P \<turnstile> \<langle>(Val v)\<bullet>M(es),s,b\<rangle> \<rightarrow> \<langle>(Val v)\<bullet>M(es'),s',b'\<rangle>"
+
+| RedCall:
+ "\<lbrakk> h a = Some(C,fs); P \<turnstile> C sees M,NonStatic:Ts\<rightarrow>T = (pns,body) in D; size vs = size pns; size Ts = size pns \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>(addr a)\<bullet>M(map Val vs), (h,l,sh), b\<rangle> \<rightarrow> \<langle>blocks(this#pns, Class D#Ts, Addr a#vs, body), (h,l,sh), b\<rangle>"
+
+| RedCallNull:
+ "P \<turnstile> \<langle>null\<bullet>M(map Val vs),s,b\<rangle> \<rightarrow> \<langle>THROW NullPointer,s,b\<rangle>"
+
+| RedCallNone:
+ "\<lbrakk> h a = Some(C,fs); \<not>(\<exists>b Ts T m D. P \<turnstile> C sees M,b:Ts\<rightarrow>T = m in D) \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>(addr a)\<bullet>M(map Val vs),(h,l,sh),b\<rangle> \<rightarrow> \<langle>THROW NoSuchMethodError,(h,l,sh),b\<rangle>"
+
+| RedCallStatic:
+ "\<lbrakk> h a = Some(C,fs); P \<turnstile> C sees M,Static:Ts\<rightarrow>T = m in D \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>(addr a)\<bullet>M(map Val vs),(h,l,sh),b\<rangle> \<rightarrow> \<langle>THROW IncompatibleClassChangeError,(h,l,sh),b\<rangle>"
+
+| SCallParams:
+ "P \<turnstile> \<langle>es,s,b\<rangle> [\<rightarrow>] \<langle>es',s',b'\<rangle> \<Longrightarrow>
+ P \<turnstile> \<langle>C\<bullet>\<^sub>sM(es),s,b\<rangle> \<rightarrow> \<langle>C\<bullet>\<^sub>sM(es'),s',b'\<rangle>"
+
+| RedSCall:
+ "\<lbrakk> P \<turnstile> C sees M,Static:Ts\<rightarrow>T = (pns,body) in D;
+ length vs = length pns; size Ts = size pns \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sM(map Val vs),s,True\<rangle> \<rightarrow> \<langle>blocks(pns, Ts, vs, body), s, False\<rangle>"
+
+| SCallInitDoneRed:
+ "\<lbrakk> P \<turnstile> C sees M,Static:Ts\<rightarrow>T = (pns,body) in D;
+ sh D = Some(sfs,Done) \<or> (M = clinit \<and> sh D = Some(sfs,Processing)) \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sM(map Val vs),(h,l,sh), False\<rangle> \<rightarrow> \<langle>C\<bullet>\<^sub>sM(map Val vs),(h,l,sh), True\<rangle>"
+
+| SCallInitRed:
+ "\<lbrakk> P \<turnstile> C sees M,Static:Ts\<rightarrow>T = (pns,body) in D;
+ \<nexists>sfs. sh D = Some(sfs,Done); M \<noteq> clinit \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sM(map Val vs),(h,l,sh), False\<rangle> \<rightarrow> \<langle>INIT D ([D],False) \<leftarrow> C\<bullet>\<^sub>sM(map Val vs),(h,l,sh),False\<rangle>"
+
+| RedSCallNone:
+ "\<lbrakk> \<not>(\<exists>b Ts T m D. P \<turnstile> C sees M,b:Ts\<rightarrow>T = m in D) \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sM(map Val vs),s,b\<rangle> \<rightarrow> \<langle>THROW NoSuchMethodError,s,False\<rangle>"
+
+| RedSCallNonStatic:
+ "\<lbrakk> P \<turnstile> C sees M,NonStatic:Ts\<rightarrow>T = m in D \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sM(map Val vs),s,b\<rangle> \<rightarrow> \<langle>THROW IncompatibleClassChangeError,s,False\<rangle>"
+
+| BlockRedNone:
+ "\<lbrakk> P \<turnstile> \<langle>e, (h,l(V:=None),sh), b\<rangle> \<rightarrow> \<langle>e', (h',l',sh'), b'\<rangle>; l' V = None; \<not> assigned V e \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>{V:T; e}, (h,l,sh), b\<rangle> \<rightarrow> \<langle>{V:T; e'}, (h',l'(V := l V),sh'), b'\<rangle>"
+
+| BlockRedSome:
+ "\<lbrakk> P \<turnstile> \<langle>e, (h,l(V:=None),sh), b\<rangle> \<rightarrow> \<langle>e', (h',l',sh'), b'\<rangle>; l' V = Some v;\<not> assigned V e \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>{V:T; e}, (h,l,sh), b\<rangle> \<rightarrow> \<langle>{V:T := Val v; e'}, (h',l'(V := l V),sh'), b'\<rangle>"
+
+| InitBlockRed:
+ "\<lbrakk> P \<turnstile> \<langle>e, (h,l(V\<mapsto>v),sh), b\<rangle> \<rightarrow> \<langle>e', (h',l',sh'), b'\<rangle>; l' V = Some v' \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>{V:T := Val v; e}, (h,l,sh), b\<rangle> \<rightarrow> \<langle>{V:T := Val v'; e'}, (h',l'(V := l V),sh'), b'\<rangle>"
+
+| RedBlock:
+ "P \<turnstile> \<langle>{V:T; Val u}, s, b\<rangle> \<rightarrow> \<langle>Val u, s, b\<rangle>"
+
+| RedInitBlock:
+ "P \<turnstile> \<langle>{V:T := Val v; Val u}, s, b\<rangle> \<rightarrow> \<langle>Val u, s, b\<rangle>"
+
+| SeqRed:
+ "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow> \<langle>e',s',b'\<rangle> \<Longrightarrow>
+ P \<turnstile> \<langle>e;;e\<^sub>2, s, b\<rangle> \<rightarrow> \<langle>e';;e\<^sub>2, s', b'\<rangle>"
+
+| RedSeq:
+ "P \<turnstile> \<langle>(Val v);;e\<^sub>2, s, b\<rangle> \<rightarrow> \<langle>e\<^sub>2, s, b\<rangle>"
+
+| CondRed:
+ "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow> \<langle>e',s',b'\<rangle> \<Longrightarrow>
+ P \<turnstile> \<langle>if (e) e\<^sub>1 else e\<^sub>2, s, b\<rangle> \<rightarrow> \<langle>if (e') e\<^sub>1 else e\<^sub>2, s', b'\<rangle>"
+
+| RedCondT:
+ "P \<turnstile> \<langle>if (true) e\<^sub>1 else e\<^sub>2, s, b\<rangle> \<rightarrow> \<langle>e\<^sub>1, s, b\<rangle>"
+
+| RedCondF:
+ "P \<turnstile> \<langle>if (false) e\<^sub>1 else e\<^sub>2, s, b\<rangle> \<rightarrow> \<langle>e\<^sub>2, s, b\<rangle>"
+
+| RedWhile:
+ "P \<turnstile> \<langle>while(b) c, s, b'\<rangle> \<rightarrow> \<langle>if(b) (c;;while(b) c) else unit, s, b'\<rangle>"
+
+| ThrowRed:
+ "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow> \<langle>e',s',b'\<rangle> \<Longrightarrow>
+ P \<turnstile> \<langle>throw e, s, b\<rangle> \<rightarrow> \<langle>throw e', s', b'\<rangle>"
+
+| RedThrowNull:
+ "P \<turnstile> \<langle>throw null, s, b\<rangle> \<rightarrow> \<langle>THROW NullPointer, s, b\<rangle>"
+
+| TryRed:
+ "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow> \<langle>e',s',b'\<rangle> \<Longrightarrow>
+ P \<turnstile> \<langle>try e catch(C V) e\<^sub>2, s, b\<rangle> \<rightarrow> \<langle>try e' catch(C V) e\<^sub>2, s', b'\<rangle>"
+
+| RedTry:
+ "P \<turnstile> \<langle>try (Val v) catch(C V) e\<^sub>2, s, b\<rangle> \<rightarrow> \<langle>Val v, s, b\<rangle>"
+
+| RedTryCatch:
+ "\<lbrakk> hp s a = Some(D,fs); P \<turnstile> D \<preceq>\<^sup>* C \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>try (Throw a) catch(C V) e\<^sub>2, s, b\<rangle> \<rightarrow> \<langle>{V:Class C := addr a; e\<^sub>2}, s, b\<rangle>"
+
+| RedTryFail:
+ "\<lbrakk> hp s a = Some(D,fs); \<not> P \<turnstile> D \<preceq>\<^sup>* C \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>try (Throw a) catch(C V) e\<^sub>2, s, b\<rangle> \<rightarrow> \<langle>Throw a, s, b\<rangle>"
+
+| ListRed1:
+ "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow> \<langle>e',s',b'\<rangle> \<Longrightarrow>
+ P \<turnstile> \<langle>e#es,s,b\<rangle> [\<rightarrow>] \<langle>e'#es,s',b'\<rangle>"
+
+| ListRed2:
+ "P \<turnstile> \<langle>es,s,b\<rangle> [\<rightarrow>] \<langle>es',s',b'\<rangle> \<Longrightarrow>
+ P \<turnstile> \<langle>Val v # es,s,b\<rangle> [\<rightarrow>] \<langle>Val v # es',s',b'\<rangle>"
+
+\<comment> \<open>Initialization procedure\<close>
+
+| RedInit:
+ "\<not>sub_RI e \<Longrightarrow> P \<turnstile> \<langle>INIT C (Nil,b) \<leftarrow> e,s,b'\<rangle> \<rightarrow> \<langle>e,s,icheck P C e\<rangle>"
+
+| InitNoneRed:
+ "sh C = None
+ \<Longrightarrow> P \<turnstile> \<langle>INIT C' (C#Cs,False) \<leftarrow> e,(h,l,sh),b\<rangle> \<rightarrow> \<langle>INIT C' (C#Cs,False) \<leftarrow> e,(h,l,sh(C \<mapsto> (sblank P C, Prepared))),b\<rangle>"
+
+| RedInitDone:
+ "sh C = Some(sfs,Done)
+ \<Longrightarrow> P \<turnstile> \<langle>INIT C' (C#Cs,False) \<leftarrow> e,(h,l,sh),b\<rangle> \<rightarrow> \<langle>INIT C' (Cs,True) \<leftarrow> e,(h,l,sh),b\<rangle>"
+
+| RedInitProcessing:
+ "sh C = Some(sfs,Processing)
+ \<Longrightarrow> P \<turnstile> \<langle>INIT C' (C#Cs,False) \<leftarrow> e,(h,l,sh),b\<rangle> \<rightarrow> \<langle>INIT C' (Cs,True) \<leftarrow> e,(h,l,sh),b\<rangle>"
+
+| RedInitError:
+ "sh C = Some(sfs,Error)
+ \<Longrightarrow> P \<turnstile> \<langle>INIT C' (C#Cs,False) \<leftarrow> e,(h,l,sh),b\<rangle> \<rightarrow> \<langle>RI (C,THROW NoClassDefFoundError);Cs \<leftarrow> e,(h,l,sh),b\<rangle>"
+
+| InitObjectRed:
+ "\<lbrakk> sh C = Some(sfs,Prepared);
+ C = Object;
+ sh' = sh(C \<mapsto> (sfs,Processing)) \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>INIT C' (C#Cs,False) \<leftarrow> e,(h,l,sh),b\<rangle> \<rightarrow> \<langle>INIT C' (C#Cs,True) \<leftarrow> e,(h,l,sh'),b\<rangle>"
+
+| InitNonObjectSuperRed:
+ "\<lbrakk> sh C = Some(sfs,Prepared);
+ C \<noteq> Object;
+ class P C = Some (D,r);
+ sh' = sh(C \<mapsto> (sfs,Processing)) \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<langle>INIT C' (C#Cs,False) \<leftarrow> e,(h,l,sh),b\<rangle> \<rightarrow> \<langle>INIT C' (D#C#Cs,False) \<leftarrow> e,(h,l,sh'),b\<rangle>"
+
+| RedInitRInit:
+ "P \<turnstile> \<langle>INIT C' (C#Cs,True) \<leftarrow> e,(h,l,sh),b\<rangle> \<rightarrow> \<langle>RI (C,C\<bullet>\<^sub>sclinit([]));Cs \<leftarrow> e,(h,l,sh),b\<rangle>"
+
+| RInitRed:
+ "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow> \<langle>e',s',b'\<rangle> \<Longrightarrow>
+ P \<turnstile> \<langle>RI (C,e);Cs \<leftarrow> e\<^sub>0, s, b\<rangle> \<rightarrow> \<langle>RI (C,e');Cs \<leftarrow> e\<^sub>0, s', b'\<rangle>"
+
+| RedRInit:
+ "\<lbrakk> sh C = Some (sfs, i);
+ sh' = sh(C \<mapsto> (sfs,Done));
+ C' = last(C#Cs) \<rbrakk> \<Longrightarrow>
+ P \<turnstile> \<langle>RI (C, Val v);Cs \<leftarrow> e, (h,l,sh), b\<rangle> \<rightarrow> \<langle>INIT C' (Cs,True) \<leftarrow> e, (h,l,sh'), b\<rangle>"
+
+\<comment> \<open>Exception propagation\<close>
+
+| CastThrow: "P \<turnstile> \<langle>Cast C (throw e), s, b\<rangle> \<rightarrow> \<langle>throw e, s, b\<rangle>"
+| BinOpThrow1: "P \<turnstile> \<langle>(throw e) \<guillemotleft>bop\<guillemotright> e\<^sub>2, s, b\<rangle> \<rightarrow> \<langle>throw e, s, b\<rangle>"
+| BinOpThrow2: "P \<turnstile> \<langle>(Val v\<^sub>1) \<guillemotleft>bop\<guillemotright> (throw e), s, b\<rangle> \<rightarrow> \<langle>throw e, s, b\<rangle>"
+| LAssThrow: "P \<turnstile> \<langle>V:=(throw e), s, b\<rangle> \<rightarrow> \<langle>throw e, s, b\<rangle>"
+| FAccThrow: "P \<turnstile> \<langle>(throw e)\<bullet>F{D}, s, b\<rangle> \<rightarrow> \<langle>throw e, s, b\<rangle>"
+| FAssThrow1: "P \<turnstile> \<langle>(throw e)\<bullet>F{D}:=e\<^sub>2, s, b\<rangle> \<rightarrow> \<langle>throw e, s, b\<rangle>"
+| FAssThrow2: "P \<turnstile> \<langle>Val v\<bullet>F{D}:=(throw e), s, b\<rangle> \<rightarrow> \<langle>throw e, s, b\<rangle>"
+| SFAssThrow: "P \<turnstile> \<langle>C\<bullet>\<^sub>sF{D}:=(throw e), s, b\<rangle> \<rightarrow> \<langle>throw e, s, b\<rangle>"
+| CallThrowObj: "P \<turnstile> \<langle>(throw e)\<bullet>M(es), s, b\<rangle> \<rightarrow> \<langle>throw e, s, b\<rangle>"
+| CallThrowParams: "\<lbrakk> es = map Val vs @ throw e # es' \<rbrakk> \<Longrightarrow> P \<turnstile> \<langle>(Val v)\<bullet>M(es), s, b\<rangle> \<rightarrow> \<langle>throw e, s, b\<rangle>"
+| SCallThrowParams: "\<lbrakk> es = map Val vs @ throw e # es' \<rbrakk> \<Longrightarrow> P \<turnstile> \<langle>C\<bullet>\<^sub>sM(es), s, b\<rangle> \<rightarrow> \<langle>throw e, s, b\<rangle>"
+| BlockThrow: "P \<turnstile> \<langle>{V:T; Throw a}, s, b\<rangle> \<rightarrow> \<langle>Throw a, s, b\<rangle>"
+| InitBlockThrow: "P \<turnstile> \<langle>{V:T := Val v; Throw a}, s, b\<rangle> \<rightarrow> \<langle>Throw a, s, b\<rangle>"
+| SeqThrow: "P \<turnstile> \<langle>(throw e);;e\<^sub>2, s, b\<rangle> \<rightarrow> \<langle>throw e, s, b\<rangle>"
+| CondThrow: "P \<turnstile> \<langle>if (throw e) e\<^sub>1 else e\<^sub>2, s, b\<rangle> \<rightarrow> \<langle>throw e, s, b\<rangle>"
+| ThrowThrow: "P \<turnstile> \<langle>throw(throw e), s, b\<rangle> \<rightarrow> \<langle>throw e, s, b\<rangle>"
+| RInitInitThrow: "\<lbrakk> sh C = Some(sfs,i); sh' = sh(C \<mapsto> (sfs,Error)) \<rbrakk> \<Longrightarrow>
+ P \<turnstile> \<langle>RI (C,Throw a);D#Cs \<leftarrow> e,(h,l,sh),b\<rangle> \<rightarrow> \<langle>RI (D,Throw a);Cs \<leftarrow> e,(h,l,sh'),b\<rangle>"
+| RInitThrow: "\<lbrakk> sh C = Some(sfs, i); sh' = sh(C \<mapsto> (sfs,Error)) \<rbrakk> \<Longrightarrow>
+ P \<turnstile> \<langle>RI (C,Throw a);Nil \<leftarrow> e,(h,l,sh),b\<rangle> \<rightarrow> \<langle>Throw a,(h,l,sh'),b\<rangle>"
+(*<*)
+lemmas red_reds_induct = red_reds.induct [split_format (complete)]
+ and red_reds_inducts = red_reds.inducts [split_format (complete)]
+
+inductive_cases [elim!]:
+ "P \<turnstile> \<langle>V:=e,s,b\<rangle> \<rightarrow> \<langle>e',s',b'\<rangle>"
+ "P \<turnstile> \<langle>e1;;e2,s,b\<rangle> \<rightarrow> \<langle>e',s',b'\<rangle>"
+(*>*)
+
+subsection\<open> The reflexive transitive closure \<close>
+
+abbreviation
+ Step :: "J_prog \<Rightarrow> expr \<Rightarrow> state \<Rightarrow> bool \<Rightarrow> expr \<Rightarrow> state \<Rightarrow> bool \<Rightarrow> bool"
+ ("_ \<turnstile> ((1\<langle>_,/_,/_\<rangle>) \<rightarrow>*/ (1\<langle>_,/_,/_\<rangle>))" [51,0,0,0,0,0,0] 81)
+ where "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle> \<equiv> ((e,s,b), e',s',b') \<in> (red P)\<^sup>*"
+
+abbreviation
+ Steps :: "J_prog \<Rightarrow> expr list \<Rightarrow> state \<Rightarrow> bool \<Rightarrow> expr list \<Rightarrow> state \<Rightarrow> bool \<Rightarrow> bool"
+ ("_ \<turnstile> ((1\<langle>_,/_,/_\<rangle>) [\<rightarrow>]*/ (1\<langle>_,/_,/_\<rangle>))" [51,0,0,0,0,0,0] 81)
+ where "P \<turnstile> \<langle>es,s,b\<rangle> [\<rightarrow>]* \<langle>es',s',b'\<rangle> \<equiv> ((es,s,b), es',s',b') \<in> (reds P)\<^sup>*"
+
+
+lemmas converse_rtrancl_induct3 =
+ converse_rtrancl_induct [of "(ax, ay, az)" "(bx, by, bz)", split_format (complete),
+ consumes 1, case_names refl step]
+
+lemma converse_rtrancl_induct_red[consumes 1]:
+assumes "P \<turnstile> \<langle>e,(h,l,sh),b\<rangle> \<rightarrow>* \<langle>e',(h',l',sh'),b'\<rangle>"
+and "\<And>e h l sh b. R e h l sh b e h l sh b"
+and "\<And>e\<^sub>0 h\<^sub>0 l\<^sub>0 sh\<^sub>0 b\<^sub>0 e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 b\<^sub>1 e' h' l' sh' b'.
+ \<lbrakk> P \<turnstile> \<langle>e\<^sub>0,(h\<^sub>0,l\<^sub>0,sh\<^sub>0),b\<^sub>0\<rangle> \<rightarrow> \<langle>e\<^sub>1,(h\<^sub>1,l\<^sub>1,sh\<^sub>1),b\<^sub>1\<rangle>; R e\<^sub>1 h\<^sub>1 l\<^sub>1 sh\<^sub>1 b\<^sub>1 e' h' l' sh' b' \<rbrakk>
+ \<Longrightarrow> R e\<^sub>0 h\<^sub>0 l\<^sub>0 sh\<^sub>0 b\<^sub>0 e' h' l' sh' b'"
+shows "R e h l sh b e' h' l' sh' b'"
+(*<*)
+proof -
+ { fix s s'
+ assume reds: "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle>"
+ and base: "\<And>e s b. R e (hp s) (lcl s) (shp s) b e (hp s) (lcl s) (shp s) b"
+ and red\<^sub>1: "\<And>e\<^sub>0 s\<^sub>0 b\<^sub>0 e\<^sub>1 s\<^sub>1 b\<^sub>1 e' s' b'.
+ \<lbrakk> P \<turnstile> \<langle>e\<^sub>0,s\<^sub>0,b\<^sub>0\<rangle> \<rightarrow> \<langle>e\<^sub>1,s\<^sub>1,b\<^sub>1\<rangle>; R e\<^sub>1 (hp s\<^sub>1) (lcl s\<^sub>1) (shp s\<^sub>1) b\<^sub>1 e' (hp s') (lcl s') (shp s') b' \<rbrakk>
+ \<Longrightarrow> R e\<^sub>0 (hp s\<^sub>0) (lcl s\<^sub>0) (shp s\<^sub>0) b\<^sub>0 e' (hp s') (lcl s') (shp s') b'"
+ from reds have "R e (hp s) (lcl s) (shp s) b e' (hp s') (lcl s') (shp s') b'"
+ proof (induct rule:converse_rtrancl_induct3)
+ case refl show ?case by(rule base)
+ next
+ case step
+ thus ?case by(blast intro:red\<^sub>1)
+ qed
+ }
+ with assms show ?thesis by fastforce
+qed
+(*>*)
+
+
+subsection\<open>Some easy lemmas\<close>
+
+lemma [iff]: "\<not> P \<turnstile> \<langle>[],s,b\<rangle> [\<rightarrow>] \<langle>es',s',b'\<rangle>"
+(*<*)by(blast elim: reds.cases)(*>*)
+
+lemma [iff]: "\<not> P \<turnstile> \<langle>Val v,s,b\<rangle> \<rightarrow> \<langle>e',s',b'\<rangle>"
+(*<*)by(fastforce elim: red.cases)(*>*)
+
+lemma val_no_step: "val_of e = \<lfloor>v\<rfloor> \<Longrightarrow> \<not> P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow> \<langle>e',s',b'\<rangle>"
+(*<*)by(drule val_of_spec, simp)(*>*)
+
+lemma [iff]: "\<not> P \<turnstile> \<langle>Throw a,s,b\<rangle> \<rightarrow> \<langle>e',s',b'\<rangle>"
+(*<*)by(fastforce elim: red.cases)(*>*)
+
+
+lemma map_Vals_no_step [iff]: "\<not> P \<turnstile> \<langle>map Val vs,s,b\<rangle> [\<rightarrow>] \<langle>es',s',b'\<rangle>"
+(*<*)
+proof(induct vs arbitrary: es')
+ case (Cons a vs)
+ { assume "P \<turnstile> \<langle>map Val (a # vs),s,b\<rangle> [\<rightarrow>] \<langle>es',s',b'\<rangle>"
+ then have False by(rule reds.cases) (insert Cons, auto)
+ }
+ then show ?case by clarsimp
+qed simp
+(*>*)
+
+lemma vals_no_step: "map_vals_of es = \<lfloor>vs\<rfloor> \<Longrightarrow> \<not> P \<turnstile> \<langle>es,s,b\<rangle> [\<rightarrow>] \<langle>es',s',b'\<rangle>"
+(*<*)by(drule map_vals_of_spec, simp)(*>*)
+
+lemma vals_throw_no_step [iff]: "\<not> P \<turnstile> \<langle>map Val vs @ Throw a # es,s,b\<rangle> [\<rightarrow>] \<langle>es',s',b'\<rangle>"
+(*<*)
+proof(induct vs arbitrary: es')
+ case Nil
+ { assume "P \<turnstile> \<langle>Throw a # es,s,b\<rangle> [\<rightarrow>] \<langle>es',s',b'\<rangle>"
+ then have False by(rule reds.cases) (insert Cons, auto)
+ }
+ then show ?case by clarsimp
+next
+ case (Cons a' vs')
+ { assume "P \<turnstile> \<langle>map Val (a'#vs') @ Throw a # es,s,b\<rangle> [\<rightarrow>] \<langle>es',s',b'\<rangle>"
+ then have False by(rule reds.cases) (insert Cons, auto)
+ }
+ then show ?case by clarsimp
+qed
+(*>*)
+
+lemma lass_val_of_red:
+ "\<lbrakk> lass_val_of e = \<lfloor>a\<rfloor>; P \<turnstile> \<langle>e,(h, l, sh),b\<rangle> \<rightarrow> \<langle>e',(h', l', sh'),b'\<rangle> \<rbrakk>
+ \<Longrightarrow> e' = unit \<and> h' = h \<and> l' = l(fst a\<mapsto>snd a) \<and> sh' = sh \<and> b = b'"
+(*<*)by(drule lass_val_of_spec, auto)(*>*)
+
+
+lemma final_no_step [iff]: "final e \<Longrightarrow> \<not> P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow> \<langle>e',s',b'\<rangle>"
+(*<*)by(erule finalE, simp+)(*>*)
+
+lemma finals_no_step [iff]: "finals es \<Longrightarrow> \<not> P \<turnstile> \<langle>es,s,b\<rangle> [\<rightarrow>] \<langle>es',s',b'\<rangle>"
+(*<*)by(erule finalsE, simp+)(*>*)
+
+lemma reds_final_same:
+"P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle> \<Longrightarrow> final e \<Longrightarrow> e = e' \<and> s = s' \<and> b = b'"
+proof(induct rule:converse_rtrancl_induct3)
+ case refl show ?case by simp
+next
+ case (step e0 s0 b0 e1 s1 b1) show ?case
+ proof(rule finalE[OF step.prems(1)])
+ fix v assume "e0 = Val v" then show ?thesis using step by simp
+ next
+ fix a assume "e0 = Throw a" then show ?thesis using step by simp
+ qed
+qed
+
+lemma reds_throw:
+"P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle> \<Longrightarrow> (\<And>e\<^sub>t. throw_of e = \<lfloor>e\<^sub>t\<rfloor> \<Longrightarrow> \<exists>e\<^sub>t'. throw_of e' = \<lfloor>e\<^sub>t'\<rfloor>)"
+proof(induct rule:converse_rtrancl_induct3)
+ case refl then show ?case by simp
+next
+ case (step e0 s0 b0 e1 s1 b1)
+ then show ?case by(auto elim: red.cases)
+qed
+
+lemma red_hext_incr: "P \<turnstile> \<langle>e,(h,l,sh),b\<rangle> \<rightarrow> \<langle>e',(h',l',sh'),b'\<rangle> \<Longrightarrow> h \<unlhd> h'"
+ and reds_hext_incr: "P \<turnstile> \<langle>es,(h,l,sh),b\<rangle> [\<rightarrow>] \<langle>es',(h',l',sh'),b'\<rangle> \<Longrightarrow> h \<unlhd> h'"
+(*<*)
+proof(induct rule:red_reds_inducts)
+ case RedNew thus ?case
+ by(fastforce dest:new_Addr_SomeD simp:hext_def split:if_splits)
+next
+ case RedFAss thus ?case by(simp add:hext_def split:if_splits)
+qed simp_all
+(*>*)
+
+
+lemma red_lcl_incr: "P \<turnstile> \<langle>e,(h\<^sub>0,l\<^sub>0,sh\<^sub>0),b\<rangle> \<rightarrow> \<langle>e',(h\<^sub>1,l\<^sub>1,sh\<^sub>1),b'\<rangle> \<Longrightarrow> dom l\<^sub>0 \<subseteq> dom l\<^sub>1"
+and reds_lcl_incr: "P \<turnstile> \<langle>es,(h\<^sub>0,l\<^sub>0,sh\<^sub>0),b\<rangle> [\<rightarrow>] \<langle>es',(h\<^sub>1,l\<^sub>1,sh\<^sub>1),b'\<rangle> \<Longrightarrow> dom l\<^sub>0 \<subseteq> dom l\<^sub>1"
+(*<*)by(induct rule: red_reds_inducts)(auto simp del:fun_upd_apply)(*>*)
+
+lemma red_lcl_add: "P \<turnstile> \<langle>e,(h,l,sh),b\<rangle> \<rightarrow> \<langle>e',(h',l',sh'),b'\<rangle> \<Longrightarrow> (\<And>l\<^sub>0. P \<turnstile> \<langle>e,(h,l\<^sub>0++l,sh),b\<rangle> \<rightarrow> \<langle>e',(h',l\<^sub>0++l',sh'),b'\<rangle>)"
+and reds_lcl_add: "P \<turnstile> \<langle>es,(h,l,sh),b\<rangle> [\<rightarrow>] \<langle>es',(h',l',sh'),b'\<rangle> \<Longrightarrow> (\<And>l\<^sub>0. P \<turnstile> \<langle>es,(h,l\<^sub>0++l,sh),b\<rangle> [\<rightarrow>] \<langle>es',(h',l\<^sub>0++l',sh'),b'\<rangle>)"
+(*<*)
+proof (induct rule:red_reds_inducts)
+ case RedCast thus ?case by(fastforce intro:red_reds.intros)
+next
+ case RedCastFail thus ?case by(force intro:red_reds.intros)
+next
+ case RedFAcc thus ?case by(fastforce intro:red_reds.intros)
+next
+ case RedCall thus ?case by(fastforce intro:red_reds.intros)
+next
+ case (InitBlockRed e h l V v sh b e' h' l' sh' b' v' T l\<^sub>0)
+ have IH: "\<And>l\<^sub>0. P \<turnstile> \<langle>e,(h, l\<^sub>0 ++ l(V \<mapsto> v), sh),b\<rangle> \<rightarrow> \<langle>e',(h', l\<^sub>0 ++ l', sh'),b'\<rangle>"
+ and l'V: "l' V = Some v'" by fact+
+ from IH have IH': "P \<turnstile> \<langle>e,(h, (l\<^sub>0 ++ l)(V \<mapsto> v), sh),b\<rangle> \<rightarrow> \<langle>e',(h', l\<^sub>0 ++ l', sh'),b'\<rangle>"
+ by simp
+ have "(l\<^sub>0 ++ l')(V := (l\<^sub>0 ++ l) V) = l\<^sub>0 ++ l'(V := l V)"
+ by(rule ext)(simp add:map_add_def)
+ with red_reds.InitBlockRed[OF IH'] l'V show ?case by(simp del:fun_upd_apply)
+next
+ case (BlockRedNone e h l V sh b e' h' l' sh' b' T l\<^sub>0)
+ have IH: "\<And>l\<^sub>0. P \<turnstile> \<langle>e,(h, l\<^sub>0 ++ l(V := None), sh),b\<rangle> \<rightarrow> \<langle>e',(h', l\<^sub>0 ++ l', sh'),b'\<rangle>"
+ and l'V: "l' V = None" and unass: "\<not> assigned V e" by fact+
+ have "l\<^sub>0(V := None) ++ l(V := None) = (l\<^sub>0 ++ l)(V := None)"
+ by(simp add:fun_eq_iff map_add_def)
+ hence IH': "P \<turnstile> \<langle>e,(h, (l\<^sub>0++l)(V := None), sh),b\<rangle> \<rightarrow> \<langle>e',(h', l\<^sub>0(V := None) ++ l', sh'),b'\<rangle>"
+ using IH[of "l\<^sub>0(V := None)"] by simp
+ have "(l\<^sub>0(V := None) ++ l')(V := (l\<^sub>0 ++ l) V) = l\<^sub>0 ++ l'(V := l V)"
+ by(simp add:fun_eq_iff map_add_def)
+ with red_reds.BlockRedNone[OF IH' _ unass] l'V show ?case
+ by(simp add: map_add_def)
+next
+ case (BlockRedSome e h l V sh b e' h' l' sh' b' v T l\<^sub>0)
+ have IH: "\<And>l\<^sub>0. P \<turnstile> \<langle>e,(h, l\<^sub>0 ++ l(V := None), sh),b\<rangle> \<rightarrow> \<langle>e',(h', l\<^sub>0 ++ l', sh'),b'\<rangle>"
+ and l'V: "l' V = Some v" and unass: "\<not> assigned V e" by fact+
+ have "l\<^sub>0(V := None) ++ l(V := None) = (l\<^sub>0 ++ l)(V := None)"
+ by(simp add:fun_eq_iff map_add_def)
+ hence IH': "P \<turnstile> \<langle>e,(h, (l\<^sub>0++l)(V := None), sh),b\<rangle> \<rightarrow> \<langle>e',(h', l\<^sub>0(V := None) ++ l', sh'),b'\<rangle>"
+ using IH[of "l\<^sub>0(V := None)"] by simp
+ have "(l\<^sub>0(V := None) ++ l')(V := (l\<^sub>0 ++ l) V) = l\<^sub>0 ++ l'(V := l V)"
+ by(simp add:fun_eq_iff map_add_def)
+ with red_reds.BlockRedSome[OF IH' _ unass] l'V show ?case
+ by(simp add:map_add_def)
+next
+ case RedTryCatch thus ?case by(fastforce intro:red_reds.intros)
+next
+ case RedTryFail thus ?case by(force intro!:red_reds.intros)
+qed (simp_all add:red_reds.intros)
+(*>*)
+
+
+lemma Red_lcl_add:
+assumes "P \<turnstile> \<langle>e,(h,l,sh), b\<rangle> \<rightarrow>* \<langle>e',(h',l',sh'), b'\<rangle>" shows "P \<turnstile> \<langle>e,(h,l\<^sub>0++l,sh),b\<rangle> \<rightarrow>* \<langle>e',(h',l\<^sub>0++l',sh'),b'\<rangle>"
+(*<*)
+using assms
+proof(induct rule:converse_rtrancl_induct_red)
+ case 1 thus ?case by simp
+next
+ case 2 thus ?case
+ by (blast dest: red_lcl_add intro: converse_rtrancl_into_rtrancl)
+qed
+(*>*)
+
+lemma assumes wf: "wwf_J_prog P"
+shows red_proc_pres: "P \<turnstile> \<langle>e,(h,l,sh),b\<rangle> \<rightarrow> \<langle>e',(h',l',sh'),b'\<rangle>
+ \<Longrightarrow> not_init C e \<Longrightarrow> sh C = \<lfloor>(sfs, Processing)\<rfloor> \<Longrightarrow> not_init C e' \<and> (\<exists>sfs'. sh' C = \<lfloor>(sfs', Processing)\<rfloor>)"
+ and reds_proc_pres: "P \<turnstile> \<langle>es,(h,l,sh),b\<rangle> [\<rightarrow>] \<langle>es',(h',l',sh'),b'\<rangle>
+ \<Longrightarrow> not_inits C es \<Longrightarrow> sh C = \<lfloor>(sfs, Processing)\<rfloor> \<Longrightarrow> not_inits C es' \<and> (\<exists>sfs'. sh' C = \<lfloor>(sfs', Processing)\<rfloor>)"
+(*<*)
+proof(induct rule:red_reds_inducts)
+ case RedCall then show ?case
+ using sees_wwf_nsub_RI[OF wf RedCall.hyps(2)] sub_RI_blocks_body nsub_RI_not_init by auto
+next
+ case RedSCall then show ?case
+ using sees_wwf_nsub_RI[OF wf RedSCall.hyps(1)] sub_RI_blocks_body nsub_RI_not_init by auto
+next
+ case (RedInitDone sh C sfs C' Cs e h l b)
+ then show ?case by(cases Cs, auto)
+next
+ case (RedInitProcessing sh C sfs C' Cs e h l b)
+ then show ?case by(cases Cs, auto)
+next
+ case (RedRInit sh C sfs i sh' C' Cs v e h l b)
+ then show ?case by(cases Cs, auto)
+next
+ case (CallThrowParams es vs e es' v M h l sh b)
+ then show ?case by(auto dest: not_inits_def')
+next
+ case (SCallThrowParams es vs e es' C M h l sh b)
+ then show ?case by(auto dest: not_inits_def')
+qed(auto)
+
+end
diff --git a/thys/JinjaDCI/J/State.thy b/thys/JinjaDCI/J/State.thy
--- a/thys/JinjaDCI/J/State.thy
+++ b/thys/JinjaDCI/J/State.thy
@@ -1,32 +1,32 @@
-(* Title: JinjaDCI/J/State.thy
-
- Author: Tobias Nipkow, Susannah Mansky
- Copyright 2003 Technische Universitaet Muenchen, 2019-20 UIUC
-
- Based on the Jinja theory J/State.thy by Tobias Nipkow
-*)
-
-section \<open> Program State \<close>
-
-theory State imports "../Common/Exceptions" begin
-
-type_synonym
- locals = "vname \<rightharpoonup> val" \<comment> \<open>local vars, incl. params and ``this''\<close>
-type_synonym
- state = "heap \<times> locals \<times> sheap"
-
-definition hp :: "state \<Rightarrow> heap"
-where
- "hp \<equiv> fst"
-definition lcl :: "state \<Rightarrow> locals"
-where
- "lcl \<equiv> fst \<circ> snd"
-definition shp :: "state \<Rightarrow> sheap"
-where
- "shp \<equiv> snd \<circ> snd"
-
-(*<*)
-declare hp_def[simp] lcl_def[simp] shp_def[simp]
-(*>*)
-
-end
+(* Title: JinjaDCI/J/State.thy
+
+ Author: Tobias Nipkow, Susannah Mansky
+ Copyright 2003 Technische Universitaet Muenchen, 2019-20 UIUC
+
+ Based on the Jinja theory J/State.thy by Tobias Nipkow
+*)
+
+section \<open> Program State \<close>
+
+theory State imports "../Common/Exceptions" begin
+
+type_synonym
+ locals = "vname \<rightharpoonup> val" \<comment> \<open>local vars, incl. params and ``this''\<close>
+type_synonym
+ state = "heap \<times> locals \<times> sheap"
+
+definition hp :: "state \<Rightarrow> heap"
+where
+ "hp \<equiv> fst"
+definition lcl :: "state \<Rightarrow> locals"
+where
+ "lcl \<equiv> fst \<circ> snd"
+definition shp :: "state \<Rightarrow> sheap"
+where
+ "shp \<equiv> snd \<circ> snd"
+
+(*<*)
+declare hp_def[simp] lcl_def[simp] shp_def[simp]
+(*>*)
+
+end
diff --git a/thys/JinjaDCI/J/TypeSafe.thy b/thys/JinjaDCI/J/TypeSafe.thy
--- a/thys/JinjaDCI/J/TypeSafe.thy
+++ b/thys/JinjaDCI/J/TypeSafe.thy
@@ -1,1382 +1,1382 @@
-(* Title: JinjaDCI/J/TypeSafe.thy
-
- Author: Tobias Nipkow, Susannah Mansky
- Copyright 2003 Technische Universitaet Muenchen, 2019-20 UIUC
-
- Based on the Jinja theory J/TypeSafe.thy by Tobias Nipkow
-*)
-
-section \<open> Type Safety Proof \<close>
-
-theory TypeSafe
-imports Progress BigStep SmallStep JWellForm
-begin
-
-(* here because it requires well-typing def *)
-lemma red_shext_incr: "P \<turnstile> \<langle>e,(h,l,sh),b\<rangle> \<rightarrow> \<langle>e',(h',l',sh'),b'\<rangle>
- \<Longrightarrow> (\<And>E T. P,E,h,sh \<turnstile> e : T \<Longrightarrow> sh \<unlhd>\<^sub>s sh')"
- and reds_shext_incr: "P \<turnstile> \<langle>es,(h,l,sh),b\<rangle> [\<rightarrow>] \<langle>es',(h',l',sh'),b'\<rangle>
- \<Longrightarrow> (\<And>E Ts. P,E,h,sh \<turnstile> es [:] Ts \<Longrightarrow> sh \<unlhd>\<^sub>s sh')"
-(*<*)
-proof(induct rule:red_reds_inducts) qed(auto simp: shext_def)
-(*>*)
-
-lemma wf_types_clinit:
-assumes wf:"wf_prog wf_md P" and ex: "class P C = Some a" and proc: "sh C = \<lfloor>(sfs, Processing)\<rfloor>"
-shows "P,E,h,sh \<turnstile> C\<bullet>\<^sub>sclinit([]) : Void"
-proof -
- from ex obtain D fs ms where "a = (D,fs,ms)" by(cases a)
- then have sP: "(C, D, fs, ms) \<in> set P" using ex map_of_SomeD[of P C a] by(simp add: class_def)
- then have "wf_clinit ms" using assms by(unfold wf_prog_def wf_cdecl_def, auto)
- then obtain pns body where sm: "(clinit, Static, [], Void, pns, body) \<in> set ms"
- by(unfold wf_clinit_def) auto
- then have "P \<turnstile> C sees clinit,Static:[] \<rightarrow> Void = (pns,body) in C"
- using mdecl_visible[OF wf sP sm] by simp
- then show ?thesis using WTrtSCall proc by simp
-qed
-
-subsection\<open>Basic preservation lemmas\<close>
-
-text\<open> First some easy preservation lemmas. \<close>
-
-theorem red_preserves_hconf:
- "P \<turnstile> \<langle>e,(h,l,sh),b\<rangle> \<rightarrow> \<langle>e',(h',l',sh'),b'\<rangle> \<Longrightarrow> (\<And>T E. \<lbrakk> P,E,h,sh \<turnstile> e : T; P \<turnstile> h \<surd> \<rbrakk> \<Longrightarrow> P \<turnstile> h' \<surd>)"
-and reds_preserves_hconf:
- "P \<turnstile> \<langle>es,(h,l,sh),b\<rangle> [\<rightarrow>] \<langle>es',(h',l',sh'),b'\<rangle> \<Longrightarrow> (\<And>Ts E. \<lbrakk> P,E,h,sh \<turnstile> es [:] Ts; P \<turnstile> h \<surd> \<rbrakk> \<Longrightarrow> P \<turnstile> h' \<surd>)"
-(*<*)
-proof (induct rule:red_reds_inducts)
- case (RedNew h a C FDTs h' l sh es)
- have new: "new_Addr h = Some a" and fields: "P \<turnstile> C has_fields FDTs"
- and h': "h' = h(a\<mapsto>blank P C)"
- and hconf: "P \<turnstile> h \<surd>" by fact+
- from new have None: "h a = None" by(rule new_Addr_SomeD)
- moreover have "P,h \<turnstile> blank P C \<surd>"
- using fields by(rule oconf_blank)
- ultimately show "P \<turnstile> h' \<surd>" using h' by(fast intro: hconf_new[OF hconf])
-next
- case (RedFAss C F t D h a fs v l sh b')
- let ?fs' = "fs((F,D)\<mapsto>v)"
- have hconf: "P \<turnstile> h \<surd>" and ha: "h a = Some(C,fs)"
- and wt: "P,E,h,sh \<turnstile> addr a\<bullet>F{D}:=Val v : T" by fact+
- from wt ha obtain TF Tv where typeofv: "typeof\<^bsub>h\<^esub> v = Some Tv"
- and has: "P \<turnstile> C has F,NonStatic:TF in D"
- and sub: "P \<turnstile> Tv \<le> TF" by auto
- have "P,h \<turnstile> (C, ?fs') \<surd>"
- proof (rule oconf_fupd[OF has])
- show "P,h \<turnstile> (C, fs) \<surd>" using hconf ha by(simp add:hconf_def)
- show "P,h \<turnstile> v :\<le> TF" using sub typeofv by(simp add:conf_def)
- qed
- with hconf ha show "P \<turnstile> h(a\<mapsto>(C, ?fs')) \<surd>" by (rule hconf_upd_obj)
-qed(auto elim: WTrt.cases)
-(*>*)
-
-
-theorem red_preserves_lconf:
- "P \<turnstile> \<langle>e,(h,l,sh),b\<rangle> \<rightarrow> \<langle>e',(h',l',sh'),b'\<rangle> \<Longrightarrow>
- (\<And>T E. \<lbrakk> P,E,h,sh \<turnstile> e:T; P,h \<turnstile> l (:\<le>) E \<rbrakk> \<Longrightarrow> P,h' \<turnstile> l' (:\<le>) E)"
-and reds_preserves_lconf:
- "P \<turnstile> \<langle>es,(h,l,sh),b\<rangle> [\<rightarrow>] \<langle>es',(h',l',sh'),b'\<rangle> \<Longrightarrow>
- (\<And>Ts E. \<lbrakk> P,E,h,sh \<turnstile> es[:]Ts; P,h \<turnstile> l (:\<le>) E \<rbrakk> \<Longrightarrow> P,h' \<turnstile> l' (:\<le>) E)"
-(*<*)
-proof(induct rule:red_reds_inducts)
- case RedNew thus ?case
- by(fast intro:lconf_hext red_hext_incr[OF red_reds.RedNew])
-next
- case RedLAss thus ?case by(fastforce elim: lconf_upd simp:conf_def)
-next
- case RedFAss thus ?case
- by(fast intro:lconf_hext red_hext_incr[OF red_reds.RedFAss])
-next
- case (InitBlockRed e h l V v sh b e' h' l' sh' b' v' T T')
- have red: "P \<turnstile> \<langle>e, (h, l(V\<mapsto>v),sh),b\<rangle> \<rightarrow> \<langle>e',(h', l',sh'),b'\<rangle>"
- and IH: "\<And>T E . \<lbrakk> P,E,h,sh \<turnstile> e:T; P,h \<turnstile> l(V\<mapsto>v) (:\<le>) E \<rbrakk>
- \<Longrightarrow> P,h' \<turnstile> l' (:\<le>) E"
- and l'V: "l' V = Some v'" and lconf: "P,h \<turnstile> l (:\<le>) E"
- and wt: "P,E,h,sh \<turnstile> {V:T := Val v; e} : T'" by fact+
- from lconf_hext[OF lconf red_hext_incr[OF red]]
- have "P,h' \<turnstile> l (:\<le>) E" .
- moreover from IH lconf wt have "P,h' \<turnstile> l' (:\<le>) E(V\<mapsto>T)"
- by(auto simp del: fun_upd_apply simp: fun_upd_same lconf_upd2 conf_def)
- ultimately show "P,h' \<turnstile> l'(V := l V) (:\<le>) E"
- by (fastforce simp:lconf_def split:if_split_asm)
-next
- case (BlockRedNone e h l V sh b e' h' l' sh' b' T T')
- have red: "P \<turnstile> \<langle>e,(h, l(V := None),sh),b\<rangle> \<rightarrow> \<langle>e',(h', l',sh'),b'\<rangle>"
- and IH: "\<And>E T. \<lbrakk> P,E,h,sh \<turnstile> e : T; P,h \<turnstile> l(V:=None) (:\<le>) E \<rbrakk>
- \<Longrightarrow> P,h' \<turnstile> l' (:\<le>) E"
- and lconf: "P,h \<turnstile> l (:\<le>) E" and wt: "P,E,h,sh \<turnstile> {V:T; e} : T'" by fact+
- from lconf_hext[OF lconf red_hext_incr[OF red]]
- have "P,h' \<turnstile> l (:\<le>) E" .
- moreover have "P,h' \<turnstile> l' (:\<le>) E(V\<mapsto>T)"
- by(rule IH, insert lconf wt, auto simp:lconf_def)
- ultimately show "P,h' \<turnstile> l'(V := l V) (:\<le>) E"
- by (fastforce simp:lconf_def split:if_split_asm)
-next
- case (BlockRedSome e h l V sh b e' h' l' sh' b' v T T')
- have red: "P \<turnstile> \<langle>e,(h, l(V := None),sh),b\<rangle> \<rightarrow> \<langle>e',(h', l',sh'),b'\<rangle>"
- and IH: "\<And>E T. \<lbrakk>P,E,h,sh \<turnstile> e : T; P,h \<turnstile> l(V:=None) (:\<le>) E\<rbrakk>
- \<Longrightarrow> P,h' \<turnstile> l' (:\<le>) E"
- and lconf: "P,h \<turnstile> l (:\<le>) E" and wt: "P,E,h,sh \<turnstile> {V:T; e} : T'" by fact+
- from lconf_hext[OF lconf red_hext_incr[OF red]]
- have "P,h' \<turnstile> l (:\<le>) E" .
- moreover have "P,h' \<turnstile> l' (:\<le>) E(V\<mapsto>T)"
- by(rule IH, insert lconf wt, auto simp:lconf_def)
- ultimately show "P,h' \<turnstile> l'(V := l V) (:\<le>) E"
- by (fastforce simp:lconf_def split:if_split_asm)
-qed(auto elim: WTrt.cases)
-(*>*)
-
-
-theorem red_preserves_shconf:
- "P \<turnstile> \<langle>e,(h,l,sh),b\<rangle> \<rightarrow> \<langle>e',(h',l',sh'),b'\<rangle> \<Longrightarrow> (\<And>T E. \<lbrakk> P,E,h,sh \<turnstile> e : T; P,h \<turnstile>\<^sub>s sh \<surd> \<rbrakk> \<Longrightarrow> P,h' \<turnstile>\<^sub>s sh' \<surd>)"
-and reds_preserves_shconf:
- "P \<turnstile> \<langle>es,(h,l,sh),b\<rangle> [\<rightarrow>] \<langle>es',(h',l',sh'),b'\<rangle> \<Longrightarrow> (\<And>Ts E. \<lbrakk> P,E,h,sh \<turnstile> es [:] Ts; P,h \<turnstile>\<^sub>s sh \<surd> \<rbrakk> \<Longrightarrow> P,h' \<turnstile>\<^sub>s sh' \<surd>)"
-(*<*)
-proof (induct rule:red_reds_inducts)
- case (RedNew h a C FDTs h' l sh es)
- have new: "new_Addr h = Some a"
- and h': "h' = h(a\<mapsto>blank P C)"
- and shconf: "P,h \<turnstile>\<^sub>s sh \<surd>" by fact+
- from new have None: "h a = None" by(rule new_Addr_SomeD)
- then show "P,h' \<turnstile>\<^sub>s sh \<surd>" using h' by(fast intro: shconf_hnew[OF shconf])
-next
- case (RedFAss C F t D h a fs v l sh b)
- let ?fs' = "fs((F,D)\<mapsto>v)"
- have shconf: "P,h \<turnstile>\<^sub>s sh \<surd>" and ha: "h a = Some(C,fs)" by fact+
- then show "P,h(a\<mapsto>(C, ?fs')) \<turnstile>\<^sub>s sh \<surd>" by (rule shconf_hupd_obj)
-next
- case (RedSFAss C F t D sh sfs i sfs' v sh' h l)
- let ?sfs' = "sfs(F\<mapsto>v)"
- have shconf: "P,h \<turnstile>\<^sub>s sh \<surd>" and shD: "sh D = \<lfloor>(sfs, i)\<rfloor>"
- and wt: "P,E,h,sh \<turnstile> C\<bullet>\<^sub>sF{D} := Val v : T" by fact+
- from wt obtain TF Tv where typeofv: "typeof\<^bsub>h\<^esub> v = Some Tv"
- and has: "P \<turnstile> C has F,Static:TF in D"
- and sub: "P \<turnstile> Tv \<le> TF" by (auto elim: WTrt.cases)
- have has': "P \<turnstile> D has F,Static:TF in D" using has by(rule has_field_idemp)
- have "P,h,D \<turnstile>\<^sub>s ?sfs' \<surd>"
- proof (rule soconf_fupd[OF has'])
- show "P,h,D \<turnstile>\<^sub>s sfs \<surd>" using shconf shD by(simp add:shconf_def)
- show "P,h \<turnstile> v :\<le> TF" using sub typeofv by(simp add:conf_def)
- qed
- with shconf have "P,h \<turnstile>\<^sub>s sh(D\<mapsto>(?sfs',i)) \<surd>" by (rule shconf_upd_obj)
- then show ?case using RedSFAss.hyps(3) RedSFAss.hyps(4) by blast
-next
- case (InitNoneRed sh C C' Cs e h l)
- let ?sfs' = "sblank P C"
- have "P,h \<turnstile>\<^sub>s sh(C \<mapsto> (?sfs', Prepared)) \<surd>"
- proof(rule shconf_upd_obj)
- show "P,h \<turnstile>\<^sub>s sh \<surd>" using InitNoneRed by simp
- show "P,h,C \<turnstile>\<^sub>s sblank P C \<surd>" by (metis has_field_def soconf_def soconf_sblank)
- qed
- then show ?case by blast
-next
- case (InitObjectRed sh C sfs sh' C' Cs e h l)
- have sh': "sh' = sh(C \<mapsto> (sfs, Processing))" by fact
- have "P,h \<turnstile>\<^sub>s sh(C \<mapsto> (sfs, Processing)) \<surd>"
- proof(rule shconf_upd_obj)
- show "P,h \<turnstile>\<^sub>s sh \<surd>" using InitObjectRed by simp
- moreover have "sh C = \<lfloor>(sfs, Prepared)\<rfloor>" using InitObjectRed by simp
- ultimately show "P,h,C \<turnstile>\<^sub>s sfs \<surd>" using shconfD by blast
- qed
- then show ?case using sh' by blast
-next
- case (InitNonObjectSuperRed sh C sfs D a b sh' C' Cs e h l)
- have sh': "sh' = sh(C \<mapsto> (sfs, Processing))" by fact
- have "P,h \<turnstile>\<^sub>s sh(C \<mapsto> (sfs, Processing)) \<surd>"
- proof(rule shconf_upd_obj)
- show "P,h \<turnstile>\<^sub>s sh \<surd>" using InitNonObjectSuperRed by simp
- moreover have "sh C = \<lfloor>(sfs, Prepared)\<rfloor>" using InitNonObjectSuperRed by simp
- ultimately show "P,h,C \<turnstile>\<^sub>s sfs \<surd>" using shconfD by blast
- qed
- then show ?case using sh' by blast
-next
- case (RedRInit sh C sfs i sh' C' Cs e v h l)
- have sh': "sh' = sh(C \<mapsto> (sfs, Done))" by fact
- have "P,h \<turnstile>\<^sub>s sh(C \<mapsto> (sfs, Done)) \<surd>"
- proof(rule shconf_upd_obj)
- show "P,h \<turnstile>\<^sub>s sh \<surd>" using RedRInit by simp
- moreover have "sh C = \<lfloor>(sfs, i)\<rfloor>" using RedRInit by simp
- ultimately show "P,h,C \<turnstile>\<^sub>s sfs \<surd>" using shconfD by blast
- qed
- then show ?case using sh' by blast
-next
- case (RInitInitThrow sh C sfs i sh' a D Cs e h l)
- have sh': "sh' = sh(C \<mapsto> (sfs, Error))" by fact
- have "P,h \<turnstile>\<^sub>s sh(C \<mapsto> (sfs, Error)) \<surd>"
- proof(rule shconf_upd_obj)
- show "P,h \<turnstile>\<^sub>s sh \<surd>" using RInitInitThrow by simp
- moreover have "sh C = \<lfloor>(sfs, i)\<rfloor>" using RInitInitThrow by simp
- ultimately show "P,h,C \<turnstile>\<^sub>s sfs \<surd>" using shconfD by blast
- qed
- then show ?case using sh' by blast
-next
- case (RInitThrow sh C sfs i sh' a e' h l)
- have sh': "sh' = sh(C \<mapsto> (sfs, Error))" by fact
- have "P,h \<turnstile>\<^sub>s sh(C \<mapsto> (sfs, Error)) \<surd>"
- proof(rule shconf_upd_obj)
- show "P,h \<turnstile>\<^sub>s sh \<surd>" using RInitThrow by simp
- moreover have "sh C = \<lfloor>(sfs, i)\<rfloor>" using RInitThrow by simp
- ultimately show "P,h,C \<turnstile>\<^sub>s sfs \<surd>" using shconfD by blast
- qed
- then show ?case using sh' by blast
-qed(auto elim: WTrt.cases)
-(*>*)
-
-theorem assumes wf: "wwf_J_prog P"
-shows red_preserves_iconf:
- "P \<turnstile> \<langle>e,(h,l,sh),b\<rangle> \<rightarrow> \<langle>e',(h',l',sh'),b'\<rangle> \<Longrightarrow> iconf sh e \<Longrightarrow> iconf sh' e'"
-and reds_preserves_iconf:
- "P \<turnstile> \<langle>es,(h,l,sh),b\<rangle> [\<rightarrow>] \<langle>es',(h',l',sh'),b'\<rangle> \<Longrightarrow> iconfs sh es \<Longrightarrow> iconfs sh' es'"
-(*<*)
-proof (induct rule:red_reds_inducts)
- case (BinOpRed1 e h l sh b e' h' l' sh' b' bop e\<^sub>2)
- then show ?case using BinOpRed1 nsub_RI_iconf[of e\<^sub>2 sh'] val_of_spec
- proof(cases "val_of e") qed(simp,fast)
-next
- case (FAssRed1 e h l sh b e' h' l' sh' b' F D e\<^sub>2)
- then show ?case using FAssRed1 nsub_RI_iconf[of e\<^sub>2 sh'] val_of_spec
- proof(cases "val_of e") qed(simp,fast)
-next
- case (CallObj e h l sh b e' h' l' sh' b' M es)
- then show ?case using CallObj nsub_RIs_iconfs[of es sh'] val_of_spec
- proof(cases "val_of e") qed(simp,fast)
-next
- case RedCall then show ?case using sees_wwf_nsub_RI[OF wf RedCall.hyps(2)]
- by (clarsimp simp: assigned_def nsub_RI_iconf)
-next
- case (RedSCall C M Ts T pns body D vs a a b)
- then have "\<not>sub_RI (blocks (pns, Ts, vs, body))"
- using sees_wwf_nsub_RI[OF wf RedSCall.hyps(1)] by simp
- then show ?case by(rule nsub_RI_iconf)
-next
- case SCallInitRed then show ?case by fastforce
-next
- case (BlockRedNone e h l V sh b e' h' l' sh' b' T)
- then show ?case by auto
-next
- case (SeqRed e h l sh b e' h' l' sh' b' e\<^sub>2)
- then show ?case
- proof(cases "lass_val_of e")
- case None then show ?thesis using SeqRed nsub_RI_iconf by(auto dest: val_of_spec)
- next
- case (Some a)
- have "e' = unit \<and> sh' = sh" by(simp add: lass_val_of_red[OF Some SeqRed(1)])
- then show ?thesis using SeqRed Some by(auto dest: val_of_spec)
- qed
-next
- case (ListRed1 e h l sh b e' h' l' sh' b' es)
- then show ?case using ListRed1 nsub_RIs_iconfs[of es sh'] val_of_spec
- proof(cases "val_of e") qed(simp,fast)
-next
- case RedInit then show ?case by(auto simp: nsub_RI_iconf)
-next
- case (RedInitDone sh C sfs C' Cs e h l b)
- then show ?case proof(cases Cs) qed(auto simp: initPD_def)
-next
- case (RedInitProcessing sh C sfs C' Cs e h l b)
- then show ?case proof(cases Cs) qed(auto simp: initPD_def)
-next
- case (RedRInit sh C sfs i sh' C' Cs v e h l b)
- then show ?case proof(cases Cs) qed(auto simp: initPD_def)
-next
- case CallThrowParams then show ?case by(auto simp: iconfs_map_throw)
-next
- case SCallThrowParams then show ?case by(auto simp: iconfs_map_throw)
-qed(auto simp: nsub_RI_iconf lass_val_of_iconf)
-(*>*)
-
-
-lemma Seq_bconf_preserve_aux:
-assumes "P \<turnstile> \<langle>e,(h, l, sh),b\<rangle> \<rightarrow> \<langle>e',(h', l', sh'),b'\<rangle>" and "P,sh \<turnstile>\<^sub>b (e;; e\<^sub>2,b) \<surd>"
- and "P,sh \<turnstile>\<^sub>b (e::expr,b) \<surd> \<longrightarrow> P,sh' \<turnstile>\<^sub>b (e'::expr,b') \<surd>"
-shows "P,sh' \<turnstile>\<^sub>b (e';;e\<^sub>2,b') \<surd>"
-proof(cases "val_of e")
- case None show ?thesis
- proof(cases "lass_val_of e")
- case lNone: None show ?thesis
- proof(cases "lass_val_of e'")
- case lNone': None
- then show ?thesis using None assms lNone
- by(auto dest: val_of_spec simp: bconf_def option.distinct(1))
- next
- case (Some a)
- then show ?thesis using None assms lNone by(auto dest: lass_val_of_spec simp: bconf_def)
- qed
- next
- case (Some a)
- then show ?thesis using None assms by(auto dest: lass_val_of_spec)
- qed
-next
- case (Some a)
- then show ?thesis using assms by(auto dest: val_of_spec)
-qed
-
-theorem red_preserves_bconf:
- "P \<turnstile> \<langle>e,(h,l,sh),b\<rangle> \<rightarrow> \<langle>e',(h',l',sh'),b'\<rangle> \<Longrightarrow> iconf sh e \<Longrightarrow> P,sh \<turnstile>\<^sub>b (e,b) \<surd> \<Longrightarrow> P,sh' \<turnstile>\<^sub>b (e',b') \<surd>"
-and reds_preserves_bconf:
- "P \<turnstile> \<langle>es,(h,l,sh),b\<rangle> [\<rightarrow>] \<langle>es',(h',l',sh'),b'\<rangle> \<Longrightarrow> iconfs sh es \<Longrightarrow> P,sh \<turnstile>\<^sub>b (es,b) \<surd> \<Longrightarrow> P,sh' \<turnstile>\<^sub>b (es',b') \<surd>"
-(*<*)
-proof (induct rule:red_reds_inducts)
- case (CastRed e a a b b e' a a b b' C) then show ?case
- proof(cases b') qed(simp, simp add: bconf_def)
-next
- case (BinOpRed1 e h l sh b e' h' l' sh' b' bop e\<^sub>2) show ?case
- proof(cases b')
- case True with BinOpRed1 val_of_spec show ?thesis
- proof(cases "val_of e") qed(simp,fast)
- next
- case False then show ?thesis by (simp add: bconf_def)
- qed
-next
-case (BinOpRed2 e a a b b e' a a b b' v\<^sub>1 bop) then show ?case
- proof(cases b') qed(simp, simp add: bconf_def)
-next
- case (LAssRed e a a b b e' a a b b' V) then show ?case
- proof(cases b') qed(simp, simp add: bconf_def)
-next
- case (FAccRed e a a b b e' a a b b' F D) then show ?case
- proof(cases b') qed(simp, simp add: bconf_def)
-next
- case (RedSFAccNonStatic C F t D h l sh b) then show ?case
- using has_field_fun[of P C F NonStatic] by (auto simp: bconf_def)
-next
- case (FAssRed1 e h l sh b e' h' l' sh' b' F D e\<^sub>2) show ?case
- proof(cases b')
- case True with FAssRed1 val_of_spec show ?thesis
- proof(cases "val_of e'")qed((simp,fast)+)
- next
- case False then show ?thesis by(simp add: bconf_def)
- qed
-next
- case (FAssRed2 e a a b b e' a a b b' v F D) then show ?case
- proof(cases b') qed(simp, simp add: bconf_def)
-next
- case (SFAssRed e h l sh b e' h' l' sh' b' C F D) then show ?case
- proof(cases b') qed(fastforce simp: bconf_def val_no_step)+
-next
- case (RedSFAssNonStatic C F t D v a a b b) then show ?case
- using has_field_fun[of P C F NonStatic] by (auto simp: bconf_def)
-next
- case (CallObj e h l sh b e' h' l' sh' b' M es) show ?case
- proof(cases b')
- case True with CallObj val_of_spec show ?thesis
- proof(cases "val_of e'")qed((simp,fast)+)
- next
- case False then show ?thesis by(simp add: bconf_def)
- qed
-next
- case (CallParams es a a b b es' a a b b' v M) then show ?case
- proof(cases b') qed(simp, simp add: bconf_def)
-next
- case (SCallParams es h l sh b es' h' l' sh' b' C M) show ?case
- proof(cases b')
- case b': True show ?thesis
- proof(cases "map_vals_of es'")
- case None
- then show ?thesis using SCallParams b' vals_no_step
- proof(cases "map_vals_of es")qed(clarsimp,fast)
- next
- case f: Some
- then show ?thesis using SCallParams b' vals_no_step
- proof(cases "map_vals_of es")qed(clarsimp,fast)
- qed
- next
- case False then show ?thesis by(simp add: bconf_def)
- qed
-next
- case (SCallInitDoneRed C M Ts T pns body D sh sfs vs h l)
- then show ?case by(auto simp: bconf_def initPD_def) (rule_tac x=D in exI, auto)+
-next
- case (RedSCallNonStatic C M Ts T a b D vs h l sh b) then show ?case
- using sees_method_fun[of P C M NonStatic] by (auto simp: bconf_def)
-next
- case (BlockRedNone e h l V sh b e' h' l' sh' b' T) show ?case
- proof(cases "assigned V e'")
- case True
- then obtain v e2 where "e' = V := Val v;;e2" by(clarsimp simp: assigned_def)
- then show ?thesis using BlockRedNone by(clarsimp simp: bconf_def)
- next
- case False then show ?thesis using BlockRedNone by simp
- qed
-next
- case (BlockRedSome e h l V sh b e' h' l' sh' b' v T) then show ?case
- proof(cases b') qed(simp, simp add: bconf_def)
-next
- case (InitBlockRed e h l V v sh b e' h' l' sh' b' v' T) show ?case
- proof(cases b')
- case True
- then show ?thesis using InitBlockRed by (simp add: assigned_def)
- next
- case False then show ?thesis by(simp add: bconf_def)
- qed
-next
- case (RedBlock V T u)
- then have "\<not>assigned V (Val u)" by(clarsimp simp: assigned_def)
- then show ?case using RedBlock by(simp add: bconf_def)
-next
- case (SeqRed e h l sh b e' h' l' sh' b' e\<^sub>2)
- have "iconf sh e"
- proof(cases "lass_val_of e")
- case (Some a) show ?thesis by(rule lass_val_of_iconf[OF Some])
- next
- case None then show ?thesis using SeqRed by(auto dest: val_of_spec)
- qed
- then show ?case using SeqRed Seq_bconf_preserve_aux by simp
-next
- case (CondRed e a a b b e' a a b b' e\<^sub>1 e\<^sub>2) then show ?case
- proof(cases b') qed(simp, simp add: bconf_def)
-next
- case (ThrowRed e a a b b e' a a b b') then show ?case
- proof(cases b') qed(simp, simp add: bconf_def)
-next
- case (TryRed e a a b b e' a a b b' C V e\<^sub>2) then show ?case
- proof(cases b') qed(simp, simp add: bconf_def)
-next
- case (ListRed1 e h l sh b e' h' l' sh' b' es)
- with val_of_spec show ?case
- proof(cases b') qed((clarsimp,fast),simp add: bconfs_def)
-next
- case (RedInit C b' e X Y b b'')
- then show ?case
- by(auto simp: bconf_def icheck_ss_exp icheck_init_class icheck_curr_init)
-next
- case (RInitRed e a a b b e' a a b b' C Cs e\<^sub>0) then show ?case
- proof(cases b') qed(simp, simp add: bconf_def)
-next
- case (BlockThrow V T a)
- then have "\<not>assigned V (Throw a)" by(simp add: assigned_def)
- then show ?case using BlockThrow by simp
-qed(simp_all, auto simp: bconf_def initPD_def)
-(*>*)
-
-text\<open> Preservation of definite assignment more complex and requires a
-few lemmas first. \<close>
-
-lemma [iff]: "\<And>A. \<lbrakk> length Vs = length Ts; length vs = length Ts\<rbrakk> \<Longrightarrow>
- \<D> (blocks (Vs,Ts,vs,e)) A = \<D> e (A \<squnion> \<lfloor>set Vs\<rfloor>)"
-(*<*)
-by (induct Vs Ts vs e rule:blocks_induct)
- (simp_all add:hyperset_defs)
-(*>*)
-
-
-lemma red_lA_incr: "P \<turnstile> \<langle>e,(h,l,sh),b\<rangle> \<rightarrow> \<langle>e',(h',l',sh'),b'\<rangle>
- \<Longrightarrow> \<lfloor>dom l\<rfloor> \<squnion> \<A> e \<sqsubseteq> \<lfloor>dom l'\<rfloor> \<squnion> \<A> e'"
-and reds_lA_incr: "P \<turnstile> \<langle>es,(h,l,sh),b\<rangle> [\<rightarrow>] \<langle>es',(h',l',sh'),b'\<rangle>
- \<Longrightarrow> \<lfloor>dom l\<rfloor> \<squnion> \<A>s es \<sqsubseteq> \<lfloor>dom l'\<rfloor> \<squnion> \<A>s es'"
-(*<*)
-proof(induct rule:red_reds_inducts)
- case TryRed then show ?case
- by(simp del:fun_upd_apply add:hyperset_defs)
- (blast dest:red_lcl_incr)+
-next
- case BinOpRed1 then show ?case
- by(auto simp del:fun_upd_apply simp:hyperset_defs)
-next
- case BinOpRed2 then show ?case
- by(auto simp del:fun_upd_apply simp:hyperset_defs)
-next
- case LAssRed then show ?case
- by(auto simp del:fun_upd_apply simp:hyperset_defs)
-next
- case FAssRed1 then show ?case
- by(auto simp del:fun_upd_apply simp:hyperset_defs)
-next
- case FAssRed2 then show ?case
- by(auto simp del:fun_upd_apply simp:hyperset_defs)
-next
- case CallObj then show ?case
- by(auto simp del:fun_upd_apply simp:hyperset_defs)
-next
- case CallParams then show ?case
- by(auto simp del:fun_upd_apply simp:hyperset_defs)
-next
- case BlockRedNone then show ?case
- by(auto simp del:fun_upd_apply simp:hyperset_defs)
-next
- case BlockRedSome then show ?case
- by(auto simp del:fun_upd_apply simp:hyperset_defs)
-next
- case InitBlockRed then show ?case
- by(auto simp del:fun_upd_apply simp:hyperset_defs)
-next
- case SeqRed then show ?case
- by(auto simp del:fun_upd_apply simp:hyperset_defs)
-next
- case CondRed then show ?case
- by(auto simp del:fun_upd_apply simp:hyperset_defs)
-next
- case RedCondT then show ?case
- by(auto simp del:fun_upd_apply simp:hyperset_defs)
-next
- case RedCondF then show ?case
- by(auto simp del:fun_upd_apply simp:hyperset_defs)
-next
- case ListRed1 then show ?case
- by(auto simp del:fun_upd_apply simp:hyperset_defs)
-next
- case ListRed2 then show ?case
- by(auto simp del:fun_upd_apply simp:hyperset_defs)
-qed (simp_all del:fun_upd_apply add:hyperset_defs)
-(*>*)
-
-text\<open> Now preservation of definite assignment. \<close>
-
-lemma assumes wf: "wf_J_prog P"
-shows red_preserves_defass:
- "P \<turnstile> \<langle>e,(h,l,sh),b\<rangle> \<rightarrow> \<langle>e',(h',l',sh'),b'\<rangle> \<Longrightarrow> \<D> e \<lfloor>dom l\<rfloor> \<Longrightarrow> \<D> e' \<lfloor>dom l'\<rfloor>"
-and "P \<turnstile> \<langle>es,(h,l,sh),b\<rangle> [\<rightarrow>] \<langle>es',(h',l',sh'),b'\<rangle> \<Longrightarrow> \<D>s es \<lfloor>dom l\<rfloor> \<Longrightarrow> \<D>s es' \<lfloor>dom l'\<rfloor>"
-(*<*)
-proof (induct rule:red_reds_inducts)
- case BinOpRed1 thus ?case by (auto elim!: D_mono[OF red_lA_incr])
-next
- case FAssRed1 thus ?case by (auto elim!: D_mono[OF red_lA_incr])
-next
- case CallObj thus ?case by (auto elim!: Ds_mono[OF red_lA_incr])
-next
- case RedCall thus ?case
- by (auto dest!:sees_wf_mdecl[OF wf] simp:wf_mdecl_def hyperset_defs elim!:D_mono')
-next
- case RedSCall thus ?case
- by (auto dest!:sees_wf_mdecl[OF wf] simp:wf_mdecl_def hyperset_defs elim!:D_mono')
-next
- case SCallInitRed
- then show ?case by(auto simp:hyperset_defs Ds_Vals)
-next
- case InitBlockRed thus ?case
- by(auto simp:hyperset_defs elim!:D_mono' simp del:fun_upd_apply)
-next
- case BlockRedNone thus ?case
- by(auto simp:hyperset_defs elim!:D_mono' simp del:fun_upd_apply)
-next
- case BlockRedSome thus ?case
- by(auto simp:hyperset_defs elim!:D_mono' simp del:fun_upd_apply)
-next
- case SeqRed thus ?case by (auto elim!: D_mono[OF red_lA_incr])
-next
- case CondRed thus ?case by (auto elim!: D_mono[OF red_lA_incr])
-next
- case TryRed thus ?case
- by (fastforce dest:red_lcl_incr intro:D_mono' simp:hyperset_defs)
-next
- case ListRed1 thus ?case by (auto elim!: Ds_mono[OF red_lA_incr])
-next
- case RedWhile thus ?case by(auto simp:hyperset_defs elim!:D_mono')
-next
- case RedInit then show ?case by (auto intro: D_mono' simp: hyperset_defs)
-next
- case (RInitRed e h l sh b e' h' l' sh' b' C Cs e\<^sub>0)
- then show ?case by(auto simp:hyperset_defs dest:red_lcl_incr elim!:D_mono')
-qed(auto simp:hyperset_defs)
-(*>*)
-
-
-text\<open> Combining conformance of heap, static heap, and local variables: \<close>
-
-definition sconf :: "J_prog \<Rightarrow> env \<Rightarrow> state \<Rightarrow> bool" ("_,_ \<turnstile> _ \<surd>" [51,51,51]50)
-where
- "P,E \<turnstile> s \<surd> \<equiv> let (h,l,sh) = s in P \<turnstile> h \<surd> \<and> P,h \<turnstile> l (:\<le>) E \<and> P,h \<turnstile>\<^sub>s sh \<surd>"
-
-lemma red_preserves_sconf:
- "\<lbrakk> P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow> \<langle>e',s',b'\<rangle>; P,E,hp s,shp s \<turnstile> e : T; P,E \<turnstile> s \<surd> \<rbrakk> \<Longrightarrow> P,E \<turnstile> s' \<surd>"
-(*<*)
-by(fastforce intro:red_preserves_hconf red_preserves_lconf red_preserves_shconf
- simp add:sconf_def)
-(*>*)
-
-lemma reds_preserves_sconf:
- "\<lbrakk> P \<turnstile> \<langle>es,s,b\<rangle> [\<rightarrow>] \<langle>es',s',b'\<rangle>; P,E,hp s,shp s \<turnstile> es [:] Ts; P,E \<turnstile> s \<surd> \<rbrakk> \<Longrightarrow> P,E \<turnstile> s' \<surd>"
-(*<*)
-by(fastforce intro:reds_preserves_hconf reds_preserves_lconf reds_preserves_shconf
- simp add:sconf_def)
-(*>*)
-
-
-subsection "Subject reduction"
-
-lemma wt_blocks:
- "\<And>E. \<lbrakk> length Vs = length Ts; length vs = length Ts \<rbrakk> \<Longrightarrow>
- (P,E,h,sh \<turnstile> blocks(Vs,Ts,vs,e) : T) =
- (P,E(Vs[\<mapsto>]Ts),h,sh \<turnstile> e:T \<and> (\<exists>Ts'. map (typeof\<^bsub>h\<^esub>) vs = map Some Ts' \<and> P \<turnstile> Ts' [\<le>] Ts))"
-(*<*)
-proof(induct Vs Ts vs e rule:blocks_induct)
- case (1 V Vs T Ts v vs e)
- then show ?case by(force simp add:rel_list_all2_Cons2)
-qed simp_all
-(*>*)
-
-theorem assumes wf: "wf_J_prog P"
-shows subject_reduction2: "P \<turnstile> \<langle>e,(h,l,sh),b\<rangle> \<rightarrow> \<langle>e',(h',l',sh'),b'\<rangle> \<Longrightarrow>
- (\<And>E T. \<lbrakk> P,E \<turnstile> (h,l,sh) \<surd>; iconf sh e; P,E,h,sh \<turnstile> e:T \<rbrakk>
- \<Longrightarrow> \<exists>T'. P,E,h',sh' \<turnstile> e':T' \<and> P \<turnstile> T' \<le> T)"
-and subjects_reduction2: "P \<turnstile> \<langle>es,(h,l,sh),b\<rangle> [\<rightarrow>] \<langle>es',(h',l',sh'),b'\<rangle> \<Longrightarrow>
- (\<And>E Ts. \<lbrakk> P,E \<turnstile> (h,l,sh) \<surd>; iconfs sh es; P,E,h,sh \<turnstile> es [:] Ts \<rbrakk>
- \<Longrightarrow> \<exists>Ts'. P,E,h',sh' \<turnstile> es' [:] Ts' \<and> P \<turnstile> Ts' [\<le>] Ts)"
-(*<*)
-proof (induct rule:red_reds_inducts)
- case RedNew then show ?case by (auto simp: blank_def)
-next
- case RedNewFail thus ?case
- by (unfold sconf_def hconf_def) (fastforce elim!:typeof_OutOfMemory)
-next
- case CastRed thus ?case
- by(clarsimp simp:is_refT_def)
- (blast intro: widens_trans dest!:widen_Class[THEN iffD1])
-next
- case RedCastFail thus ?case
- by (unfold sconf_def hconf_def) (fastforce elim!:typeof_ClassCast)
-next
- case (BinOpRed1 e\<^sub>1 h l sh b e\<^sub>1' h' l' sh' b' bop e\<^sub>2)
- have red: "P \<turnstile> \<langle>e\<^sub>1,(h,l,sh),b\<rangle> \<rightarrow> \<langle>e\<^sub>1',(h',l',sh'),b'\<rangle>"
- and IH: "\<And>E T. \<lbrakk>P,E \<turnstile> (h,l,sh) \<surd>; iconf sh e\<^sub>1; P,E,h,sh \<turnstile> e\<^sub>1:T\<rbrakk>
- \<Longrightarrow> \<exists>U. P,E,h',sh' \<turnstile> e\<^sub>1' : U \<and> P \<turnstile> U \<le> T"
- and conf: "P,E \<turnstile> (h,l,sh) \<surd>" and iconf: "iconf sh (e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2)"
- and wt: "P,E,h,sh \<turnstile> e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2 : T" by fact+
- have val: "val_of e\<^sub>1 = None" using red iconf val_no_step by auto
- then have iconf1: "iconf sh e\<^sub>1" and nsub_RI2: "\<not>sub_RI e\<^sub>2" using iconf by simp+
- have "P,E,h',sh' \<turnstile> e\<^sub>1' \<guillemotleft>bop\<guillemotright> e\<^sub>2 : T"
- proof (cases bop)
- assume [simp]: "bop = Eq"
- from wt obtain T\<^sub>1 T\<^sub>2 where [simp]: "T = Boolean"
- and wt\<^sub>1: "P,E,h,sh \<turnstile> e\<^sub>1 : T\<^sub>1" and wt\<^sub>2: "P,E,h,sh \<turnstile> e\<^sub>2 : T\<^sub>2" by auto
- show ?thesis
- using WTrt_hext_shext_mono[OF wt\<^sub>2 red_hext_incr[OF red] red_shext_incr[OF red wt\<^sub>1] nsub_RI2]
- IH[OF conf iconf1 wt\<^sub>1] by auto
- next
- assume [simp]: "bop = Add"
- from wt have [simp]: "T = Integer"
- and wt\<^sub>1: "P,E,h,sh \<turnstile> e\<^sub>1 : Integer" and wt\<^sub>2: "P,E,h,sh \<turnstile> e\<^sub>2 : Integer"
- by auto
- show ?thesis
- using WTrt_hext_shext_mono[OF wt\<^sub>2 red_hext_incr[OF red] red_shext_incr[OF red wt\<^sub>1] nsub_RI2]
- IH[OF conf iconf1 wt\<^sub>1] by auto
- qed
- thus ?case by auto
-next
- case (BinOpRed2 e\<^sub>2 h l sh b e\<^sub>2' h' l' sh' b' v\<^sub>1 bop)
- have red: "P \<turnstile> \<langle>e\<^sub>2,(h,l,sh),b\<rangle> \<rightarrow> \<langle>e\<^sub>2',(h',l',sh'),b'\<rangle>"
- and IH: "\<And>E T. \<lbrakk>P,E \<turnstile> (h,l,sh) \<surd>; iconf sh e\<^sub>2; P,E,h,sh \<turnstile> e\<^sub>2:T\<rbrakk>
- \<Longrightarrow> \<exists>U. P,E,h',sh' \<turnstile> e\<^sub>2' : U \<and> P \<turnstile> U \<le> T"
- and conf: "P,E \<turnstile> (h,l,sh) \<surd>" and iconf: "iconf sh (Val v\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2)"
- and wt: "P,E,h,sh \<turnstile> (Val v\<^sub>1) \<guillemotleft>bop\<guillemotright> e\<^sub>2 : T" by fact+
- have iconf2: "iconf sh e\<^sub>2" using iconf by simp
- have "P,E,h',sh' \<turnstile> (Val v\<^sub>1) \<guillemotleft>bop\<guillemotright> e\<^sub>2' : T"
- proof (cases bop)
- assume [simp]: "bop = Eq"
- from wt obtain T\<^sub>1 T\<^sub>2 where [simp]: "T = Boolean"
- and wt\<^sub>1: "P,E,h,sh \<turnstile> Val v\<^sub>1 : T\<^sub>1" and wt\<^sub>2: "P,E,h,sh \<turnstile> e\<^sub>2:T\<^sub>2" by auto
- show ?thesis
- using IH[OF conf iconf2 wt\<^sub>2] WTrt_hext_mono[OF wt\<^sub>1 red_hext_incr[OF red]]
- by auto
- next
- assume [simp]: "bop = Add"
- from wt have [simp]: "T = Integer"
- and wt\<^sub>1: "P,E,h,sh \<turnstile> Val v\<^sub>1 : Integer" and wt\<^sub>2: "P,E,h,sh \<turnstile> e\<^sub>2 : Integer"
- by auto
- show ?thesis
- using IH[OF conf iconf2 wt\<^sub>2] WTrt_hext_mono[OF wt\<^sub>1 red_hext_incr[OF red]]
- by auto
- qed
- thus ?case by auto
-next
- case (RedBinOp bop) thus ?case
- proof (cases bop)
- case Eq thus ?thesis using RedBinOp by auto
- next
- case Add thus ?thesis using RedBinOp by auto
- qed
-next
- case RedVar thus ?case by (fastforce simp:sconf_def lconf_def conf_def)
-next
- case (LAssRed e h l sh b e' h' l' sh' b' V)
- obtain Te where Te: "P,E,h,sh \<turnstile> e : Te \<and> P \<turnstile> Te \<le> the(E V)" using LAssRed.prems(3) by auto
- then have wide: "P \<turnstile> Te \<le> the(E V)" using LAssRed by simp
- then have "\<exists>T'. P,E,h',sh' \<turnstile> e' : T' \<and> P \<turnstile> T' \<le> Te"
- using LAssRed.hyps(2) LAssRed.prems(1,2) Te widen_trans[OF _ wide] by auto
- then obtain T' where wt: "P,E,h',sh' \<turnstile> e' : T' \<and> P \<turnstile> T' \<le> Te" by clarsimp
- have "P,E,h',sh' \<turnstile> V:=e' : Void" using LAssRed wt widen_trans[OF _ wide] by auto
- then show ?case using LAssRed by(rule_tac x = Void in exI) auto
-next
- case (FAccRed e h l sh b e' h' l' sh' b' F D)
- have IH: "\<And>E T. \<lbrakk>P,E \<turnstile> (h,l,sh) \<surd>; iconf sh e; P,E,h,sh \<turnstile> e : T\<rbrakk>
- \<Longrightarrow> \<exists>U. P,E,h',sh' \<turnstile> e' : U \<and> P \<turnstile> U \<le> T"
- and conf: "P,E \<turnstile> (h,l,sh) \<surd>" and iconf: "iconf sh (e\<bullet>F{D})"
- and wt: "P,E,h,sh \<turnstile> e\<bullet>F{D} : T" by fact+
- have iconf': "iconf sh e" using iconf by simp+
- \<comment> \<open>The goal: ?case = @{prop ?case}\<close>
- \<comment> \<open>Now distinguish the two cases how wt can have arisen.\<close>
- { fix C assume wte: "P,E,h,sh \<turnstile> e : Class C"
- and has: "P \<turnstile> C has F,NonStatic:T in D"
- from IH[OF conf iconf' wte]
- obtain U where wte': "P,E,h',sh' \<turnstile> e' : U" and UsubC: "P \<turnstile> U \<le> Class C"
- by auto
- \<comment> \<open>Now distinguish what @{term U} can be.\<close>
- { assume "U = NT" hence ?case using wte'
- by(blast intro:WTrtFAccNT widen_refl) }
- moreover
- { fix C' assume U: "U = Class C'" and C'subC: "P \<turnstile> C' \<preceq>\<^sup>* C"
- from has_field_mono[OF has C'subC] wte' U
- have ?case by(blast intro:WTrtFAcc) }
- ultimately have ?case using UsubC by(simp add: widen_Class) blast }
- moreover
- { assume "P,E,h,sh \<turnstile> e : NT"
- hence "P,E,h',sh' \<turnstile> e' : NT" using IH[OF conf iconf'] by fastforce
- hence ?case by(fastforce intro:WTrtFAccNT widen_refl) }
- ultimately show ?case using wt by blast
-next
- case RedFAcc thus ?case
- by(fastforce simp:sconf_def hconf_def oconf_def conf_def has_field_def
- dest:has_fields_fun)
-next
- case RedFAccNull thus ?case
- by(fastforce intro: widen_refl WTThrow[OF WTVal] elim!: typeof_NullPointer
- simp: sconf_def hconf_def)
-next
- case RedSFAccNone then show ?case
- by(fastforce intro: WTrtThrow[OF WTrtVal] elim!: typeof_NoSuchFieldError
- simp: sconf_def hconf_def)
-next
- case RedFAccStatic then show ?case
- by(fastforce intro: WTrtThrow[OF WTrtVal] elim!: typeof_IncompatibleClassChangeError
- simp: sconf_def hconf_def)
-next
- case (RedSFAcc C F t D sh sfs i v h l es)
- then have "P \<turnstile> C has F,Static:T in D" by fast
- then have dM: "P \<turnstile> D has F,Static:T in D" by(rule has_field_idemp)
- then show ?case using RedSFAcc by(fastforce simp:sconf_def shconf_def soconf_def conf_def)
-next
- case SFAccInitDoneRed then show ?case by (meson widen_refl)
-next
- case (SFAccInitRed C F t D sh h l E T)
- have "is_class P D" using SFAccInitRed.hyps(1) by(rule has_field_is_class')
- then have "P,E,h,sh \<turnstile> INIT D ([D],False) \<leftarrow> C\<bullet>\<^sub>sF{D} : T \<and> P \<turnstile> T \<le> T"
- using SFAccInitRed WTrtInit[OF SFAccInitRed.prems(3)] by clarsimp
- then show ?case by(rule exI)
-next
- case RedSFAccNonStatic then show ?case
- by(fastforce intro: WTrtThrow[OF WTrtVal] elim!: typeof_IncompatibleClassChangeError
- simp: sconf_def hconf_def)
-next
- case (FAssRed1 e h l sh b e' h' l' sh' b' F D e\<^sub>2)
- have red: "P \<turnstile> \<langle>e,(h,l,sh),b\<rangle> \<rightarrow> \<langle>e',(h',l',sh'),b'\<rangle>"
- and IH: "\<And>E T. \<lbrakk>P,E \<turnstile> (h,l,sh) \<surd>; iconf sh e; P,E,h,sh \<turnstile> e : T\<rbrakk>
- \<Longrightarrow> \<exists>U. P,E,h',sh' \<turnstile> e' : U \<and> P \<turnstile> U \<le> T"
- and conf: "P,E \<turnstile> (h,l,sh) \<surd>" and iconf: "iconf sh (e\<bullet>F{D} := e\<^sub>2)"
- and wt: "P,E,h,sh \<turnstile> e\<bullet>F{D}:=e\<^sub>2 : T" by fact+
- have val: "val_of e = None" using red iconf val_no_step by auto
- then have iconf': "iconf sh e" and nsub_RI2: "\<not>sub_RI e\<^sub>2" using iconf by simp+
- from wt have void: "T = Void" by blast
- \<comment> \<open>We distinguish if @{term e} has type @{term NT} or a Class type\<close>
- \<comment> \<open>Remember ?case = @{term ?case}\<close>
- { assume wt':"P,E,h,sh \<turnstile> e : NT"
- hence "P,E,h',sh' \<turnstile> e' : NT" using IH[OF conf iconf'] by fastforce
- moreover obtain T\<^sub>2 where "P,E,h,sh \<turnstile> e\<^sub>2 : T\<^sub>2" using wt by auto
- from this red_hext_incr[OF red] red_shext_incr[OF red wt'] nsub_RI2 have "P,E,h',sh' \<turnstile> e\<^sub>2 : T\<^sub>2"
- by(rule WTrt_hext_shext_mono)
- ultimately have ?case using void by(blast intro!:WTrtFAssNT)
- }
- moreover
- { fix C TF T\<^sub>2 assume wt\<^sub>1: "P,E,h,sh \<turnstile> e : Class C" and wt\<^sub>2: "P,E,h,sh \<turnstile> e\<^sub>2 : T\<^sub>2"
- and has: "P \<turnstile> C has F,NonStatic:TF in D" and sub: "P \<turnstile> T\<^sub>2 \<le> TF"
- obtain U where wt\<^sub>1': "P,E,h',sh' \<turnstile> e' : U" and UsubC: "P \<turnstile> U \<le> Class C"
- using IH[OF conf iconf' wt\<^sub>1] by blast
- have wt\<^sub>2': "P,E,h',sh' \<turnstile> e\<^sub>2 : T\<^sub>2"
- by(rule WTrt_hext_shext_mono[OF wt\<^sub>2 red_hext_incr[OF red] red_shext_incr[OF red wt\<^sub>1] nsub_RI2])
- \<comment> \<open>Is @{term U} the null type or a class type?\<close>
- { assume "U = NT" with wt\<^sub>1' wt\<^sub>2' void have ?case
- by(blast intro!:WTrtFAssNT) }
- moreover
- { fix C' assume UClass: "U = Class C'" and "subclass": "P \<turnstile> C' \<preceq>\<^sup>* C"
- have "P,E,h',sh' \<turnstile> e' : Class C'" using wt\<^sub>1' UClass by auto
- moreover have "P \<turnstile> C' has F,NonStatic:TF in D"
- by(rule has_field_mono[OF has "subclass"])
- ultimately have ?case using wt\<^sub>2' sub void by(blast intro:WTrtFAss) }
- ultimately have ?case using UsubC by(auto simp add:widen_Class) }
- ultimately show ?case using wt by blast
-next
- case (FAssRed2 e\<^sub>2 h l sh b e\<^sub>2' h' l' sh' b' v F D)
- have red: "P \<turnstile> \<langle>e\<^sub>2,(h,l,sh),b\<rangle> \<rightarrow> \<langle>e\<^sub>2',(h',l',sh'),b'\<rangle>"
- and IH: "\<And>E T. \<lbrakk>P,E \<turnstile> (h,l,sh) \<surd>; iconf sh e\<^sub>2; P,E,h,sh \<turnstile> e\<^sub>2 : T\<rbrakk>
- \<Longrightarrow> \<exists>U. P,E,h',sh' \<turnstile> e\<^sub>2' : U \<and> P \<turnstile> U \<le> T"
- and conf: "P,E \<turnstile> (h,l,sh) \<surd>" and iconf: "iconf sh (Val v\<bullet>F{D} := e\<^sub>2)"
- and wt: "P,E,h,sh \<turnstile> Val v\<bullet>F{D}:=e\<^sub>2 : T" by fact+
- have iconf2: "iconf sh e\<^sub>2" using iconf by simp
- from wt have [simp]: "T = Void" by auto
- from wt show ?case
- proof (rule WTrt_elim_cases)
- fix C TF T\<^sub>2
- assume wt\<^sub>1: "P,E,h,sh \<turnstile> Val v : Class C"
- and has: "P \<turnstile> C has F,NonStatic:TF in D"
- and wt\<^sub>2: "P,E,h,sh \<turnstile> e\<^sub>2 : T\<^sub>2" and TsubTF: "P \<turnstile> T\<^sub>2 \<le> TF"
- have wt\<^sub>1': "P,E,h',sh' \<turnstile> Val v : Class C"
- using WTrt_hext_mono[OF wt\<^sub>1 red_hext_incr[OF red]] by auto
- obtain T\<^sub>2' where wt\<^sub>2': "P,E,h',sh' \<turnstile> e\<^sub>2' : T\<^sub>2'" and T'subT: "P \<turnstile> T\<^sub>2' \<le> T\<^sub>2"
- using IH[OF conf iconf2 wt\<^sub>2] by blast
- have "P,E,h',sh' \<turnstile> Val v\<bullet>F{D}:=e\<^sub>2' : Void"
- by(rule WTrtFAss[OF wt\<^sub>1' has wt\<^sub>2' widen_trans[OF T'subT TsubTF]])
- thus ?case by auto
- next
- fix T\<^sub>2 assume null: "P,E,h,sh \<turnstile> Val v : NT" and wt\<^sub>2: "P,E,h,sh \<turnstile> e\<^sub>2 : T\<^sub>2"
- from null have "v = Null" by simp
- moreover
- obtain T\<^sub>2' where "P,E,h',sh' \<turnstile> e\<^sub>2' : T\<^sub>2' \<and> P \<turnstile> T\<^sub>2' \<le> T\<^sub>2"
- using IH[OF conf iconf2 wt\<^sub>2] by blast
- ultimately show ?thesis by(fastforce intro:WTrtFAssNT)
- qed
-next
- case RedFAss thus ?case by(auto simp del:fun_upd_apply)
-next
- case RedFAssNull thus ?case
- by(fastforce intro: WTThrow[OF WTVal] elim!:typeof_NullPointer simp:sconf_def hconf_def)
-next
- case RedFAssStatic then show ?case
- by(fastforce intro: WTrtThrow[OF WTrtVal] elim!: typeof_IncompatibleClassChangeError
- simp: sconf_def hconf_def)
-next
- case (SFAssRed e h l sh b e' h' l' sh' b' C F D E T)
- have IH: "\<And>E T. \<lbrakk>P,E \<turnstile> (h,l,sh) \<surd>; iconf sh e; P,E,h,sh \<turnstile> e : T\<rbrakk>
- \<Longrightarrow> \<exists>U. P,E,h',sh' \<turnstile> e' : U \<and> P \<turnstile> U \<le> T"
- and conf: "P,E \<turnstile> (h,l,sh) \<surd>" and iconf: "iconf sh (C\<bullet>\<^sub>sF{D} := e)"
- and wt: "P,E,h,sh \<turnstile> C\<bullet>\<^sub>sF{D}:=e : T" by fact+
- have iconf': "iconf sh e" using iconf by simp
- from wt have [simp]: "T = Void" by auto
- from wt show ?case
- proof (rule WTrt_elim_cases)
- fix TF T1
- assume has: "P \<turnstile> C has F,Static:TF in D"
- and wt1: "P,E,h,sh \<turnstile> e : T1" and TsubTF: "P \<turnstile> T1 \<le> TF"
- obtain T' where wt1': "P,E,h',sh' \<turnstile> e' : T'" and T'subT: "P \<turnstile> T' \<le> T1"
- using IH[OF conf iconf' wt1] by blast
- have "P,E,h',sh' \<turnstile> C\<bullet>\<^sub>sF{D}:=e' : Void"
- by(rule WTrtSFAss[OF wt1' has widen_trans[OF T'subT TsubTF]])
- thus ?case by auto
- qed
-next
- case SFAssInitDoneRed then show ?case by (meson widen_refl)
-next
- case (SFAssInitRed C F t D sh v h l E T)
- have "is_class P D" using SFAssInitRed.hyps(1) by(rule has_field_is_class')
- then have "P,E,h,sh \<turnstile> INIT D ([D],False) \<leftarrow> C\<bullet>\<^sub>sF{D} := Val v : T \<and> P \<turnstile> T \<le> T"
- using SFAssInitRed WTrtInit[OF SFAssInitRed.prems(3)] by clarsimp
- then show ?case by(rule exI)
-next
- case RedSFAssNone then show ?case
- by(fastforce intro: WTrtThrow[OF WTrtVal] elim!: typeof_NoSuchFieldError
- simp: sconf_def hconf_def)
-next
- case RedSFAssNonStatic then show ?case
- by(fastforce intro: WTrtThrow[OF WTrtVal] elim!: typeof_IncompatibleClassChangeError
- simp: sconf_def hconf_def)
-next
- case (CallObj e h l sh b e' h' l' sh' b' M es)
- have red: "P \<turnstile> \<langle>e,(h,l,sh),b\<rangle> \<rightarrow> \<langle>e',(h',l',sh'),b'\<rangle>"
- and IH: "\<And>E T. \<lbrakk>P,E \<turnstile> (h,l,sh) \<surd>; iconf sh e; P,E,h,sh \<turnstile> e : T\<rbrakk>
- \<Longrightarrow> \<exists>U. P,E,h',sh' \<turnstile> e' : U \<and> P \<turnstile> U \<le> T"
- and conf: "P,E \<turnstile> (h,l,sh) \<surd>" and iconf: "iconf sh (e\<bullet>M(es))"
- and wt: "P,E,h,sh \<turnstile> e\<bullet>M(es) : T" by fact+
- have val: "val_of e = None" using red iconf val_no_step by auto
- then have iconf': "iconf sh e" and nsub_RIs: "\<not>sub_RIs es" using iconf by simp+
- \<comment> \<open>We distinguish if @{term e} has type @{term NT} or a Class type\<close>
- \<comment> \<open>Remember ?case = @{term ?case}\<close>
- { assume wt':"P,E,h,sh \<turnstile> e:NT"
- hence "P,E,h',sh' \<turnstile> e' : NT" using IH[OF conf iconf'] by fastforce
- moreover
- fix Ts assume wtes: "P,E,h,sh \<turnstile> es [:] Ts"
- have "P,E,h',sh' \<turnstile> es [:] Ts"
- by(rule WTrts_hext_shext_mono[OF wtes red_hext_incr[OF red] red_shext_incr[OF red wt'] nsub_RIs])
- ultimately have ?case by(blast intro!:WTrtCallNT) }
- moreover
- { fix C D Ts Us pns body
- assume wte: "P,E,h,sh \<turnstile> e : Class C"
- and "method": "P \<turnstile> C sees M,NonStatic:Ts\<rightarrow>T = (pns,body) in D"
- and wtes: "P,E,h,sh \<turnstile> es [:] Us" and subs: "P \<turnstile> Us [\<le>] Ts"
- obtain U where wte': "P,E,h',sh' \<turnstile> e' : U" and UsubC: "P \<turnstile> U \<le> Class C"
- using IH[OF conf iconf' wte] by blast
- \<comment> \<open>Is @{term U} the null type or a class type?\<close>
- { assume "U = NT"
- moreover have "P,E,h',sh' \<turnstile> es [:] Us"
- by(rule WTrts_hext_shext_mono[OF wtes red_hext_incr[OF red] red_shext_incr[OF red wte] nsub_RIs])
- ultimately have ?case using wte' by(blast intro!:WTrtCallNT) }
- moreover
- { fix C' assume UClass: "U = Class C'" and "subclass": "P \<turnstile> C' \<preceq>\<^sup>* C"
- have "P,E,h',sh' \<turnstile> e' : Class C'" using wte' UClass by auto
- moreover obtain Ts' T' pns' body' D'
- where method': "P \<turnstile> C' sees M,NonStatic:Ts'\<rightarrow>T' = (pns',body') in D'"
- and subs': "P \<turnstile> Ts [\<le>] Ts'" and sub': "P \<turnstile> T' \<le> T"
- using Call_lemma[OF "method" "subclass" wf] by fast
- moreover have "P,E,h',sh' \<turnstile> es [:] Us"
- by(rule WTrts_hext_shext_mono[OF wtes red_hext_incr[OF red] red_shext_incr[OF red wte] nsub_RIs])
- ultimately have ?case
- using subs by(blast intro:WTrtCall rtrancl_trans widens_trans) }
- ultimately have ?case using UsubC by(auto simp add:widen_Class) }
- ultimately show ?case using wt by auto
-next
- case (CallParams es h l sh b es' h' l' sh' b' v M)
- have reds: "P \<turnstile> \<langle>es,(h,l,sh),b\<rangle> [\<rightarrow>] \<langle>es',(h',l',sh'),b'\<rangle>"
- and IH: "\<And>E Ts. \<lbrakk>P,E \<turnstile> (h,l,sh) \<surd>; iconfs sh es; P,E,h,sh \<turnstile> es [:] Ts\<rbrakk>
- \<Longrightarrow> \<exists>Us. P,E,h',sh' \<turnstile> es' [:] Us \<and> P \<turnstile> Us [\<le>] Ts"
- and conf: "P,E \<turnstile> (h,l,sh) \<surd>" and iconf: "iconf sh (Val v\<bullet>M(es))"
- and wt: "P,E,h,sh \<turnstile> Val v\<bullet>M(es) : T" by fact+
- have iconfs: "iconfs sh es" using iconf by simp
- from wt show ?case
- proof (rule WTrt_elim_cases)
- fix C D Ts Us pns body
- assume wte: "P,E,h,sh \<turnstile> Val v : Class C"
- and "P \<turnstile> C sees M,NonStatic:Ts\<rightarrow>T = (pns,body) in D"
- and wtes: "P,E,h,sh \<turnstile> es [:] Us" and "P \<turnstile> Us [\<le>] Ts"
- moreover have "P,E,h',sh' \<turnstile> Val v : Class C"
- using WTrt_hext_mono[OF wte reds_hext_incr[OF reds]] by auto
- moreover
- obtain Us' where "P,E,h',sh' \<turnstile> es' [:] Us' \<and> P \<turnstile> Us' [\<le>] Us"
- using IH[OF conf iconfs wtes] by blast
- ultimately show ?thesis by(blast intro:WTrtCall widens_trans)
- next
- fix Us
- assume null: "P,E,h,sh \<turnstile> Val v : NT" and wtes: "P,E,h,sh \<turnstile> es [:] Us"
- from null have "v = Null" by simp
- moreover
- obtain Us' where "P,E,h',sh' \<turnstile> es' [:] Us' \<and> P \<turnstile> Us' [\<le>] Us"
- using IH[OF conf iconfs wtes] by blast
- ultimately show ?thesis by(fastforce intro:WTrtCallNT)
- qed
-next
- case (RedCall h a C fs M Ts T pns body D vs l sh b E T')
- have hp: "h a = Some(C,fs)"
- and "method": "P \<turnstile> C sees M,NonStatic: Ts\<rightarrow>T = (pns,body) in D"
- and wt: "P,E,h,sh \<turnstile> addr a\<bullet>M(map Val vs) : T'" by fact+
- obtain Ts' where wtes: "P,E,h,sh \<turnstile> map Val vs [:] Ts'"
- and subs: "P \<turnstile> Ts' [\<le>] Ts" and T'isT: "T' = T"
- using wt "method" hp by (auto dest:sees_method_fun)
- from wtes subs have length_vs: "length vs = length Ts"
- by(fastforce simp:list_all2_iff dest!:WTrts_same_length)
- from sees_wf_mdecl[OF wf "method"] obtain T''
- where wtabody: "P,[this#pns [\<mapsto>] Class D#Ts] \<turnstile> body :: T''"
- and T''subT: "P \<turnstile> T'' \<le> T" and length_pns: "length pns = length Ts"
- by(fastforce simp:wf_mdecl_def simp del:map_upds_twist)
- from wtabody have "P,Map.empty(this#pns [\<mapsto>] Class D#Ts),h,sh \<turnstile> body : T''"
- by(rule WT_implies_WTrt)
- hence "P,E(this#pns [\<mapsto>] Class D#Ts),h,sh \<turnstile> body : T''"
- by(rule WTrt_env_mono) simp
- hence "P,E,h,sh \<turnstile> blocks(this#pns, Class D#Ts, Addr a#vs, body) : T''"
- using wtes subs hp sees_method_decl_above[OF "method"] length_vs length_pns
- by(fastforce simp add:wt_blocks rel_list_all2_Cons2)
- with T''subT T'isT show ?case by blast
-next
- case RedCallNull thus ?case
- by(fastforce intro: WTThrow[OF WTVal] elim!:typeof_NullPointer simp: sconf_def hconf_def)
-next
- case RedCallStatic then show ?case
- by(fastforce intro: WTrtThrow[OF WTrtVal] elim!: typeof_IncompatibleClassChangeError
- simp: sconf_def hconf_def)
-next
- case (SCallParams es h l sh b es' h' l' sh' b' C M)
- have IH: "\<And>E Ts. \<lbrakk>P,E \<turnstile> (h,l,sh) \<surd>; iconfs sh es; P,E,h,sh \<turnstile> es [:] Ts\<rbrakk>
- \<Longrightarrow> \<exists>Us. P,E,h',sh' \<turnstile> es' [:] Us \<and> P \<turnstile> Us [\<le>] Ts"
- and conf: "P,E \<turnstile> (h,l,sh) \<surd>" and iconf: "iconf sh (C\<bullet>\<^sub>sM(es))"
- and wt: "P,E,h,sh \<turnstile> C\<bullet>\<^sub>sM(es) : T" by fact+
- have iconfs: "iconfs sh es" using iconf by simp
- from wt show ?case
- proof (rule WTrt_elim_cases)
- fix D Ts Us pns body sfs vs
- assume method: "P \<turnstile> C sees M,Static:Ts\<rightarrow>T = (pns,body) in D"
- and wtes: "P,E,h,sh \<turnstile> es [:] Us" and us: "P \<turnstile> Us [\<le>] Ts"
- and clinit: "M = clinit \<longrightarrow> sh D = \<lfloor>(sfs,Processing)\<rfloor> \<and> es = map Val vs"
- obtain Us' where es': "P,E,h',sh' \<turnstile> es' [:] Us'" and us': "P \<turnstile> Us' [\<le>] Us"
- using IH[OF conf iconfs wtes] by blast
- show ?thesis
- proof(cases "M = clinit")
- case True then show ?thesis using clinit SCallParams.hyps(1) by blast
- next
- case False
- then show ?thesis using es' method us us' by(blast intro:WTrtSCall widens_trans)
- qed
- qed
-next
- case (RedSCall C M Ts T pns body D vs h l sh E T')
- have "method": "P \<turnstile> C sees M,Static: Ts\<rightarrow>T = (pns,body) in D"
- and wt: "P,E,h,sh \<turnstile> C\<bullet>\<^sub>sM(map Val vs) : T'" by fact+
- obtain Ts' where wtes: "P,E,h,sh \<turnstile> map Val vs [:] Ts'"
- and subs: "P \<turnstile> Ts' [\<le>] Ts" and T'isT: "T' = T"
- using wt "method" map_Val_eq by(auto dest:sees_method_fun)+
- from wtes subs have length_vs: "length vs = length Ts"
- by(fastforce simp:list_all2_iff dest!:WTrts_same_length)
- from sees_wf_mdecl[OF wf "method"] obtain T''
- where wtabody: "P,[pns [\<mapsto>] Ts] \<turnstile> body :: T''"
- and T''subT: "P \<turnstile> T'' \<le> T" and length_pns: "length pns = length Ts"
- by(fastforce simp:wf_mdecl_def simp del:map_upds_twist)
- from wtabody have "P,Map.empty(pns [\<mapsto>] Ts),h,sh \<turnstile> body : T''"
- by(rule WT_implies_WTrt)
- hence "P,E(pns [\<mapsto>] Ts),h,sh \<turnstile> body : T''"
- by(rule WTrt_env_mono) simp
- hence "P,E,h,sh \<turnstile> blocks(pns, Ts, vs, body) : T''"
- using wtes subs sees_method_decl_above[OF "method"] length_vs length_pns
- by(fastforce simp add:wt_blocks rel_list_all2_Cons2)
- with T''subT T'isT show ?case by blast
-next
- case SCallInitDoneRed then show ?case by (meson widen_refl)
-next
- case (SCallInitRed C F Ts t pns body D sh v h l E T)
- have "is_class P D" using SCallInitRed.hyps(1) by(rule sees_method_is_class')
- then have "P,E,h,sh \<turnstile> INIT D ([D],False) \<leftarrow> C\<bullet>\<^sub>sF(map Val v) : T \<and> P \<turnstile> T \<le> T"
- using SCallInitRed WTrtInit[OF SCallInitRed.prems(3)] by clarsimp
- then show ?case by(rule exI)
-next
- case RedSCallNone then show ?case
- by(fastforce intro: WTrtThrow[OF WTrtVal] elim!: typeof_NoSuchMethodError
- simp: sconf_def hconf_def)
-next
- case RedSCallNonStatic then show ?case
- by(fastforce intro: WTrtThrow[OF WTrtVal] elim!: typeof_IncompatibleClassChangeError
- simp: sconf_def hconf_def)
-next
- case BlockRedNone thus ?case
- by(auto simp del:fun_upd_apply)(fastforce simp:sconf_def lconf_def)
-next
- case (BlockRedSome e h l V sh b e' h' l' sh' b' v T E Te)
- have red: "P \<turnstile> \<langle>e,(h,l(V:=None),sh),b\<rangle> \<rightarrow> \<langle>e',(h',l',sh'),b'\<rangle>"
- and IH: "\<And>E T. \<lbrakk>P,E \<turnstile> (h,l(V:=None),sh) \<surd>; iconf sh e; P,E,h,sh \<turnstile> e : T\<rbrakk>
- \<Longrightarrow> \<exists>T'. P,E,h',sh' \<turnstile> e' : T' \<and> P \<turnstile> T' \<le> T"
- and Some: "l' V = Some v" and conf: "P,E \<turnstile> (h,l,sh) \<surd>"
- and iconf: "iconf sh {V:T; e}"
- and wt: "P,E,h,sh \<turnstile> {V:T; e} : Te" by fact+
- obtain Te' where IH': "P,E(V\<mapsto>T),h',sh' \<turnstile> e' : Te' \<and> P \<turnstile> Te' \<le> Te"
- using IH conf iconf wt by(fastforce simp:sconf_def lconf_def)
- have "P,h' \<turnstile> l' (:\<le>) E(V\<mapsto>T)" using conf wt
- by(fastforce intro:red_preserves_lconf[OF red] simp:sconf_def lconf_def)
- hence "P,h' \<turnstile> v :\<le> T" using Some by(fastforce simp:lconf_def)
- with IH' show ?case
- by(fastforce simp:sconf_def conf_def fun_upd_same simp del:fun_upd_apply)
-next
- case (InitBlockRed e h l V v sh b e' h' l' sh' b' v' T E T')
- have red: "P \<turnstile> \<langle>e, (h,l(V\<mapsto>v),sh),b\<rangle> \<rightarrow> \<langle>e',(h',l',sh'),b'\<rangle>"
- and IH: "\<And>E T. \<lbrakk>P,E \<turnstile> (h,l(V\<mapsto>v),sh) \<surd>; iconf sh e; P,E,h,sh \<turnstile> e : T\<rbrakk>
- \<Longrightarrow> \<exists>U. P,E,h',sh' \<turnstile> e' : U \<and> P \<turnstile> U \<le> T"
- and v': "l' V = Some v'" and conf: "P,E \<turnstile> (h,l,sh) \<surd>"
- and iconf: "iconf sh {V:T; V:=Val v;; e}"
- and wt: "P,E,h,sh \<turnstile> {V:T := Val v; e} : T'" by fact+
- from wt obtain T\<^sub>1 where wt\<^sub>1: "typeof\<^bsub>h\<^esub> v = Some T\<^sub>1"
- and T1subT: "P \<turnstile> T\<^sub>1 \<le> T" and wt\<^sub>2: "P,E(V\<mapsto>T),h,sh \<turnstile> e : T'" by auto
- have lconf\<^sub>2: "P,h \<turnstile> l(V\<mapsto>v) (:\<le>) E(V\<mapsto>T)" using conf wt\<^sub>1 T1subT
- by(simp add:sconf_def lconf_upd2 conf_def)
- have "\<exists>T\<^sub>1'. typeof\<^bsub>h'\<^esub> v' = Some T\<^sub>1' \<and> P \<turnstile> T\<^sub>1' \<le> T"
- using v' red_preserves_lconf[OF red wt\<^sub>2 lconf\<^sub>2]
- by(fastforce simp:lconf_def conf_def)
- with IH conf iconf lconf\<^sub>2 wt\<^sub>2 show ?case by (fastforce simp add:sconf_def)
-next
- case (SeqRed e h l sh b e' h' l' sh' b' e\<^sub>2)
- then have val: "val_of e = None" by (simp add: val_no_step)
- show ?case
- proof(cases "lass_val_of e")
- case None
- then show ?thesis
- using SeqRed val by(auto elim: WTrt_hext_shext_mono[OF _ red_hext_incr red_shext_incr])
- next
- case (Some a)
- have "sh = sh'" using SeqRed lass_val_of_spec[OF Some] by auto
- then show ?thesis using SeqRed val Some
- by(auto intro: lass_val_of_iconf[OF Some] elim: WTrt_hext_mono[OF _ red_hext_incr])
- qed
-next
- case CondRed thus ?case
- by auto (blast intro:WTrt_hext_shext_mono[OF _ red_hext_incr red_shext_incr])+
-next
- case ThrowRed thus ?case
- by(auto simp:is_refT_def)(blast dest:widen_Class[THEN iffD1])+
-next
- case RedThrowNull thus ?case
- by(fastforce intro: WTThrow[OF WTVal] elim!:typeof_NullPointer simp:sconf_def hconf_def)
-next
- case TryRed thus ?case
- by auto (blast intro:widen_trans WTrt_hext_shext_mono[OF _ red_hext_incr red_shext_incr])
-next
- case RedTryFail thus ?case
- by(fastforce intro: WTrtThrow[OF WTrtVal] simp:sconf_def hconf_def)
-next
- case (ListRed1 e h l sh b e' h' l' sh' b' es)
- then have val: "val_of e = None" by(simp add: val_no_step)
- obtain U Us where Ts: "Ts = U # Us" using ListRed1 by auto
- then have nsub_RI: "\<not> sub_RIs es" and wts: "P,E,h,sh \<turnstile> es [:] Us" and wt: "P,E,h,sh \<turnstile> e : U"
- and IH: "\<And>E T. \<lbrakk>P,E \<turnstile> (h, l, sh) \<surd>; P,E,h,sh \<turnstile> e : T\<rbrakk> \<Longrightarrow> \<exists>T'. P,E,h',sh' \<turnstile> e' : T' \<and> P \<turnstile> T' \<le> T"
- using ListRed1 val by auto
- obtain T' where
- "\<forall>E0 E1. (\<exists>T2. P,E1,h',sh' \<turnstile> e' : T2 \<and> P \<turnstile> T2 \<le> E0) = (P,E1,h',sh' \<turnstile> e' : T' E0 E1 \<and> P \<turnstile> T' E0 E1 \<le> E0)"
- by moura
- then have disj: "\<forall>E t. \<not> P,E \<turnstile> (h, l, sh) \<surd> \<or> \<not> P,E,h,sh \<turnstile> e : t \<or> P,E,h',sh' \<turnstile> e' : T' t E \<and> P \<turnstile> T' t E \<le> t"
- using IH by presburger
- have "P,E,h',sh' \<turnstile> es [:] Us"
- using nsub_RI wts wt by (metis (no_types) ListRed1.hyps(1) WTrts_hext_shext_mono red_hext_incr red_shext_incr)
- then have "\<exists>ts. (\<exists>t tsa. ts = t # tsa \<and> P,E,h',sh' \<turnstile> e' : t \<and> P,E,h',sh' \<turnstile> es [:] tsa) \<and> P \<turnstile> ts [\<le>] (U # Us)"
- using disj wt ListRed1.prems(1) by blast
- then show ?case using Ts by auto
-next
- case ListRed2 thus ?case
- by(fastforce dest: hext_typeof_mono[OF reds_hext_incr])
-next
- case (InitNoneRed sh C C' Cs e h l b)
- then have sh: "sh \<unlhd>\<^sub>s sh(C \<mapsto> (sblank P C, Prepared))" by(simp add: shext_def)
- have wt: "P,E,h,sh(C \<mapsto> (sblank P C, Prepared)) \<turnstile> INIT C' (C # Cs,False) \<leftarrow> e : T"
- using InitNoneRed WTrt_shext_mono[OF _ sh] by fastforce
- then show ?case by(rule_tac x = T in exI) (simp add: fun_upd_def)
-next
- case (RedInitDone sh C sfs C' Cs e h l b)
- then have "P,E,h,sh \<turnstile> INIT C' (Cs,True) \<leftarrow> e : T" by auto (metis Nil_tl list.set_sel(2))
- then show ?case by(rule_tac x = T in exI) simp
-next
- case (RedInitProcessing sh C sfs C' Cs e h l b)
- then have "P,E,h,sh \<turnstile> INIT C' (Cs,True) \<leftarrow> e : T" by auto (metis Nil_tl list.set_sel(2))+
- then show ?case by(rule_tac x = T in exI) simp
-next
- case RedInitError then show ?case
- by(fastforce intro: WTrtThrow[OF WTrtVal] elim!: typeof_NoClassDefFoundError
- simp: sconf_def hconf_def)
-next
- case (InitObjectRed sh C sfs sh' C' Cs e h l b)
- then have sh: "sh \<unlhd>\<^sub>s sh(Object \<mapsto> (sfs, Processing))" by(simp add: shext_def)
- have "P,E,h,sh' \<turnstile> INIT C' (C # Cs,True) \<leftarrow> e : T"
- using InitObjectRed WTrt_shext_mono[OF _ sh] by auto
- then show ?case by(rule_tac x = T in exI) (simp add: fun_upd_def)
-next
- case (InitNonObjectSuperRed sh C sfs D fs ms sh' C' Cs e h l b)
- then have sh: "sh \<unlhd>\<^sub>s sh(C \<mapsto> (sfs, Processing))" by(simp add: shext_def)
- then have cd: "is_class P D" using InitNonObjectSuperRed class_wf wf wf_cdecl_supD by blast
- have sup': "supercls_lst P (C # Cs)" using InitNonObjectSuperRed.prems(3) by auto
- then have sup: "supercls_lst P (D # C # Cs)"
- using supercls_lst_app[of P C Cs D] subcls1I[OF InitNonObjectSuperRed.hyps(3,2)] by auto
- have "distinct (C # Cs)" using InitNonObjectSuperRed.prems(3) by auto
- then have dist: "distinct (D # C # Cs)"
- using wf_supercls_distinct_app[OF wf InitNonObjectSuperRed.hyps(2-3) sup'] by simp
- have "P,E,h,sh' \<turnstile> INIT C' (D # C # Cs,False) \<leftarrow> e : T"
- using InitNonObjectSuperRed WTrt_shext_mono[OF _ sh] cd sup dist by auto
- then show ?case by(rule_tac x = T in exI) simp
-next
- case (RedInitRInit C' C Cs e' h l sh b E T)
- then obtain a sfs where C: "class P C = \<lfloor>a\<rfloor>" and proc: "sh C = \<lfloor>(sfs, Processing)\<rfloor>"
- using WTrtInit by(auto simp: is_class_def)
- then have T': "P,E,h,sh \<turnstile> C\<bullet>\<^sub>sclinit([]) : Void" using wf_types_clinit[OF wf C] by simp
- have "P,E,h,sh \<turnstile> RI (C,C\<bullet>\<^sub>sclinit([])) ; Cs \<leftarrow> e' : T"
- using RedInitRInit by(auto intro: T')
- then show ?case by(rule_tac x = T in exI) simp
-next
- case (RInitRed e h l sh b e' h' l' sh' b' C Cs e\<^sub>0 E T)
- then have "(\<And>E T. P,E \<turnstile> (h, l, sh) \<surd> \<Longrightarrow> P,E,h,sh \<turnstile> e : T \<Longrightarrow> \<exists>T'. P,E,h',sh' \<turnstile> e' : T' \<and> P \<turnstile> T' \<le> T)"
- by auto
- then have "\<exists>T'. P,E,h',sh' \<turnstile> e' : T'" using RInitRed by blast
- then obtain T' where e': "P,E,h',sh' \<turnstile> e' : T'" by auto
- have wt\<^sub>0: "P,E,h',sh' \<turnstile> e\<^sub>0 : T"
- using RInitRed by simp (auto intro: WTrt_hext_shext_mono[OF _ red_hext_incr red_shext_incr])
- have nip: "\<forall>C' \<in> set (C#Cs). not_init C' e' \<and> (\<exists>sfs. sh' C' = \<lfloor>(sfs, Processing)\<rfloor>)"
- using RInitRed red_proc_pres[OF wf_prog_wwf_prog[OF wf]] by auto
- have shC: "\<exists>sfs. sh' C = \<lfloor>(sfs, Processing)\<rfloor> \<or> sh' C = \<lfloor>(sfs, Error)\<rfloor> \<and> e' = THROW NoClassDefFoundError"
- using RInitRed red_proc_pres[OF wf_prog_wwf_prog[OF wf] RInitRed.hyps(1)] by blast
- have "P,E,h',sh' \<turnstile> RI (C,e') ; Cs \<leftarrow> e\<^sub>0 : T" using RInitRed e' wt\<^sub>0 nip shC by auto
- then show ?case by(rule_tac x = T in exI) simp
-next
- case (RedRInit sh C sfs i sh' C' Cs v e h l b)
- then have sh: "sh \<unlhd>\<^sub>s sh(C \<mapsto> (sfs, Done))" by(auto simp: shext_def)
- have wt: "P,E,h,sh(C \<mapsto> (sfs, Done)) \<turnstile> e : T"
- using RedRInit WTrt_shext_mono[OF _ sh] by auto
- have shC: "\<forall>C' \<in> set(tl Cs). \<exists>sfs. sh C' = \<lfloor>(sfs, Processing)\<rfloor>" using RedRInit by(cases Cs, auto)
- have "P,E,h,sh' \<turnstile> INIT C' (Cs,True) \<leftarrow> e : T" using RedRInit wt shC by(cases Cs, auto)
- then show ?case by(rule_tac x = T in exI) simp
-next
- case (SCallThrowParams es vs e es' C M h l sh b)
- then show ?case using map_Val_nthrow_neq[of _ vs e es'] by fastforce
-next
- case (RInitInitThrow sh C sfs i sh' a D Cs e h l b)
- then have sh: "sh \<unlhd>\<^sub>s sh(C \<mapsto> (sfs, Error))" by(auto simp: shext_def)
- have wt: "P,E,h,sh(C \<mapsto> (sfs, Error)) \<turnstile> e : T"
- using RInitInitThrow WTrt_shext_mono[OF _ sh] by clarsimp
- then have "P,E,h,sh' \<turnstile> RI (D,Throw a) ; Cs \<leftarrow> e : T" using RInitInitThrow by auto
- then show ?case by(rule_tac x = T in exI) simp
-qed fastforce+ (* esp all Throw propagation rules except RInitInit are dealt with here *)
-(*>*)
-
-
-corollary subject_reduction:
- "\<lbrakk> wf_J_prog P; P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow> \<langle>e',s',b'\<rangle>; P,E \<turnstile> s \<surd>; iconf (shp s) e; P,E,hp s,shp s \<turnstile> e:T \<rbrakk>
- \<Longrightarrow> \<exists>T'. P,E,hp s',shp s' \<turnstile> e':T' \<and> P \<turnstile> T' \<le> T"
-(*<*)by(cases s, cases s', fastforce dest:subject_reduction2)(*>*)
-
-corollary subjects_reduction:
- "\<lbrakk> wf_J_prog P; P \<turnstile> \<langle>es,s,b\<rangle> [\<rightarrow>] \<langle>es',s',b'\<rangle>; P,E \<turnstile> s \<surd>; iconfs (shp s) es; P,E,hp s,shp s \<turnstile> es[:]Ts \<rbrakk>
- \<Longrightarrow> \<exists>Ts'. P,E,hp s',shp s' \<turnstile> es'[:]Ts' \<and> P \<turnstile> Ts' [\<le>] Ts"
-(*<*)by(cases s, cases s', fastforce dest:subjects_reduction2)(*>*)
-
-
-subsection \<open> Lifting to @{text"\<rightarrow>*"} \<close>
-
-text\<open> Now all these preservation lemmas are first lifted to the transitive
-closure \dots \<close>
-
-lemma Red_preserves_sconf:
-assumes wf: "wf_J_prog P" and Red: "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle>"
-shows "\<And>T. \<lbrakk> P,E,hp s,shp s \<turnstile> e : T; iconf (shp s) e; P,E \<turnstile> s \<surd> \<rbrakk> \<Longrightarrow> P,E \<turnstile> s' \<surd>"
-(*<*)
-using Red
-proof (induct rule:converse_rtrancl_induct3)
- case refl show ?case by fact
-next
- case (step e s b e' s' b')
- obtain h l sh h' l' sh' where s:"s = (h,l,sh)" and s':"s' = (h',l',sh')"
- by(cases s, cases s')
- then have "P \<turnstile> \<langle>e,(h,l,sh),b\<rangle> \<rightarrow> \<langle>e',(h',l',sh'),b'\<rangle>" using step.hyps(1) by simp
- then have iconf': "iconf (shp s') e'" using red_preserves_iconf[OF wf_prog_wwf_prog[OF wf]]
- step.prems(2) s s' by simp
- thus ?case using step
- by(blast intro:red_preserves_sconf dest: subject_reduction[OF wf])
-qed
-(*>*)
-
-lemma Red_preserves_iconf:
-assumes wf: "wwf_J_prog P" and Red: "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle>"
-shows "iconf (shp s) e \<Longrightarrow> iconf (shp s') e'"
-(*<*)
-using Red
-proof (induct rule:converse_rtrancl_induct3)
- case refl show ?case by fact
-next
- case (step e s b e' s' b')
- thus ?case using wf step by(cases s, cases s', simp) (blast intro:red_preserves_iconf)
-qed
-(*>*)
-
-lemma Reds_preserves_iconf:
-assumes wf: "wwf_J_prog P" and Red: "P \<turnstile> \<langle>es,s,b\<rangle> [\<rightarrow>]* \<langle>es',s',b'\<rangle>"
-shows "iconfs (shp s) es \<Longrightarrow> iconfs (shp s') es'"
-(*<*)
-using Red
-proof (induct rule:converse_rtrancl_induct3)
- case refl show ?case by fact
-next
- case (step e s b e' s' b')
- thus ?case using wf step by(cases s, cases s', simp) (blast intro:reds_preserves_iconf)
-qed
-(*>*)
-
-lemma Red_preserves_bconf:
-assumes wf: "wwf_J_prog P" and Red: "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle>"
-shows "iconf (shp s) e \<Longrightarrow> P,(shp s) \<turnstile>\<^sub>b (e,b) \<surd> \<Longrightarrow> P,(shp s') \<turnstile>\<^sub>b (e'::expr,b') \<surd>"
-(*<*)
-using Red
-proof (induct rule:converse_rtrancl_induct3)
- case refl show ?case by fact
-next
- case (step e s1 b e' s2 b')
- then have "iconf (shp s2) e'" using step red_preserves_iconf[OF wf]
- by(cases s1, cases s2) auto
- thus ?case using step by(cases s1, cases s2, simp) (blast intro:red_preserves_bconf)
-qed
-(*>*)
-
-lemma Reds_preserves_bconf:
-assumes wf: "wwf_J_prog P" and Red: "P \<turnstile> \<langle>es,s,b\<rangle> [\<rightarrow>]* \<langle>es',s',b'\<rangle>"
-shows "iconfs (shp s) es \<Longrightarrow> P,(shp s) \<turnstile>\<^sub>b (es,b) \<surd> \<Longrightarrow> P,(shp s') \<turnstile>\<^sub>b (es'::expr list,b') \<surd>"
-(*<*)
-using Red
-proof (induct rule:converse_rtrancl_induct3)
- case refl show ?case by fact
-next
- case (step es s1 b es' s2 b')
- then have "iconfs (shp s2) es'" using step reds_preserves_iconf[OF wf]
- by(cases s1, cases s2) auto
- thus ?case using step by(cases s1, cases s2, simp) (blast intro:reds_preserves_bconf)
-qed
-(*>*)
-
-lemma Red_preserves_defass:
-assumes wf: "wf_J_prog P" and reds: "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle>"
-shows "\<D> e \<lfloor>dom(lcl s)\<rfloor> \<Longrightarrow> \<D> e' \<lfloor>dom(lcl s')\<rfloor>"
-using reds
-proof (induct rule:converse_rtrancl_induct3)
- case refl thus ?case .
-next
- case (step e s b e' s' b') thus ?case
- by(cases s,cases s')(auto dest:red_preserves_defass[OF wf])
-qed
-
-
-lemma Red_preserves_type:
-assumes wf: "wf_J_prog P" and Red: "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle>"
-shows "!!T. \<lbrakk> P,E \<turnstile> s\<surd>; iconf (shp s) e; P,E,hp s,shp s \<turnstile> e:T \<rbrakk>
- \<Longrightarrow> \<exists>T'. P \<turnstile> T' \<le> T \<and> P,E,hp s',shp s' \<turnstile> e':T'"
-(*<*)
-using Red
-proof (induct rule:converse_rtrancl_induct3)
- case refl thus ?case by blast
-next
- case step thus ?case
- by(blast intro:widen_trans red_preserves_sconf Red_preserves_iconf[OF wf_prog_wwf_prog[OF wf]]
- dest:subject_reduction[OF wf])
-qed
-(*>*)
-
-
-subsection "The final polish"
-
-text\<open> The above preservation lemmas are now combined and packed nicely. \<close>
-
-definition wf_config :: "J_prog \<Rightarrow> env \<Rightarrow> state \<Rightarrow> expr \<Rightarrow> ty \<Rightarrow> bool" ("_,_,_ \<turnstile> _ : _ \<surd>" [51,0,0,0,0]50)
-where
- "P,E,s \<turnstile> e:T \<surd> \<equiv> P,E \<turnstile> s \<surd> \<and> iconf (shp s) e \<and> P,E,hp s,shp s \<turnstile> e:T"
-
-theorem Subject_reduction: assumes wf: "wf_J_prog P"
-shows "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow> \<langle>e',s',b'\<rangle> \<Longrightarrow> P,E,s \<turnstile> e : T \<surd>
- \<Longrightarrow> \<exists>T'. P,E,s' \<turnstile> e' : T' \<surd> \<and> P \<turnstile> T' \<le> T"
-(*<*)
-by(cases s, cases s')
- (force simp: wf_config_def
- elim:red_preserves_sconf red_preserves_iconf[OF wf_prog_wwf_prog[OF wf]]
- dest:subject_reduction[OF wf])
-(*>*)
-
-
-theorem Subject_reductions:
-assumes wf: "wf_J_prog P" and reds: "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle>"
-shows "\<And>T. P,E,s \<turnstile> e:T \<surd> \<Longrightarrow> \<exists>T'. P,E,s' \<turnstile> e':T' \<surd> \<and> P \<turnstile> T' \<le> T"
-(*<*)
-using reds
-proof (induct rule:converse_rtrancl_induct3)
- case refl thus ?case by blast
-next
- case step thus ?case
- by(blast dest:Subject_reduction[OF wf] intro:widen_trans)
-qed
-(*>*)
-
-
-corollary Progress: assumes wf: "wf_J_prog P"
-shows "\<lbrakk> P,E,s \<turnstile> e : T \<surd>; \<D> e \<lfloor>dom(lcl s)\<rfloor>; P,shp s \<turnstile>\<^sub>b (e,b) \<surd>; \<not> final e \<rbrakk>
- \<Longrightarrow> \<exists>e' s' b'. P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow> \<langle>e',s',b'\<rangle>"
-(*<*)
-using progress[OF wf_prog_wwf_prog[OF wf]]
-by(cases b) (auto simp:wf_config_def sconf_def)
-(*>*)
-
-corollary TypeSafety:
-fixes s::state and e::expr
-assumes wf: "wf_J_prog P" and sconf: "P,E \<turnstile> s \<surd>" and wt: "P,E \<turnstile> e::T"
- and \<D>: "\<D> e \<lfloor>dom(lcl s)\<rfloor>"
- and iconf: "iconf (shp s) e" and bconf: "P,(shp s) \<turnstile>\<^sub>b (e,b) \<surd>"
- and steps: "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle>"
- and nstep: "\<not>(\<exists>e'' s'' b''. P \<turnstile> \<langle>e',s',b'\<rangle> \<rightarrow> \<langle>e'',s'',b''\<rangle>)"
-shows "(\<exists>v. e' = Val v \<and> P,hp s' \<turnstile> v :\<le> T) \<or>
- (\<exists>a. e' = Throw a \<and> a \<in> dom(hp s'))"
-(*<*)
-proof -
- have wwf: "wwf_J_prog P" by(rule wf_prog_wwf_prog[OF wf])
- have wfc: "P,E,s \<turnstile> e:T \<surd>" using WT_implies_WTrt[OF wt] sconf iconf
- by(simp add:wf_config_def)
- obtain T' where wfc': "P,E,s' \<turnstile> e' : T' \<surd>" and T': "P \<turnstile> T' \<le> T"
- using Subject_reductions[OF wf steps wfc] by clarsimp
- have \<D>': "\<D> e' \<lfloor>dom (lcl s')\<rfloor>"
- by(rule Red_preserves_defass[OF wf steps \<D>])
- have bconf': "P,(shp s') \<turnstile>\<^sub>b (e',b') \<surd>"
- by(rule Red_preserves_bconf[OF wwf steps iconf bconf])
- have fin': "final e'" using Progress[OF wf wfc' \<D>' bconf'] nstep by blast
- then show ?thesis using wfc wfc' T'
- by(fastforce simp:wf_config_def final_def conf_def)
-qed
-(*>*)
-
-
-end
+(* Title: JinjaDCI/J/TypeSafe.thy
+
+ Author: Tobias Nipkow, Susannah Mansky
+ Copyright 2003 Technische Universitaet Muenchen, 2019-20 UIUC
+
+ Based on the Jinja theory J/TypeSafe.thy by Tobias Nipkow
+*)
+
+section \<open> Type Safety Proof \<close>
+
+theory TypeSafe
+imports Progress BigStep SmallStep JWellForm
+begin
+
+(* here because it requires well-typing def *)
+lemma red_shext_incr: "P \<turnstile> \<langle>e,(h,l,sh),b\<rangle> \<rightarrow> \<langle>e',(h',l',sh'),b'\<rangle>
+ \<Longrightarrow> (\<And>E T. P,E,h,sh \<turnstile> e : T \<Longrightarrow> sh \<unlhd>\<^sub>s sh')"
+ and reds_shext_incr: "P \<turnstile> \<langle>es,(h,l,sh),b\<rangle> [\<rightarrow>] \<langle>es',(h',l',sh'),b'\<rangle>
+ \<Longrightarrow> (\<And>E Ts. P,E,h,sh \<turnstile> es [:] Ts \<Longrightarrow> sh \<unlhd>\<^sub>s sh')"
+(*<*)
+proof(induct rule:red_reds_inducts) qed(auto simp: shext_def)
+(*>*)
+
+lemma wf_types_clinit:
+assumes wf:"wf_prog wf_md P" and ex: "class P C = Some a" and proc: "sh C = \<lfloor>(sfs, Processing)\<rfloor>"
+shows "P,E,h,sh \<turnstile> C\<bullet>\<^sub>sclinit([]) : Void"
+proof -
+ from ex obtain D fs ms where "a = (D,fs,ms)" by(cases a)
+ then have sP: "(C, D, fs, ms) \<in> set P" using ex map_of_SomeD[of P C a] by(simp add: class_def)
+ then have "wf_clinit ms" using assms by(unfold wf_prog_def wf_cdecl_def, auto)
+ then obtain pns body where sm: "(clinit, Static, [], Void, pns, body) \<in> set ms"
+ by(unfold wf_clinit_def) auto
+ then have "P \<turnstile> C sees clinit,Static:[] \<rightarrow> Void = (pns,body) in C"
+ using mdecl_visible[OF wf sP sm] by simp
+ then show ?thesis using WTrtSCall proc by simp
+qed
+
+subsection\<open>Basic preservation lemmas\<close>
+
+text\<open> First some easy preservation lemmas. \<close>
+
+theorem red_preserves_hconf:
+ "P \<turnstile> \<langle>e,(h,l,sh),b\<rangle> \<rightarrow> \<langle>e',(h',l',sh'),b'\<rangle> \<Longrightarrow> (\<And>T E. \<lbrakk> P,E,h,sh \<turnstile> e : T; P \<turnstile> h \<surd> \<rbrakk> \<Longrightarrow> P \<turnstile> h' \<surd>)"
+and reds_preserves_hconf:
+ "P \<turnstile> \<langle>es,(h,l,sh),b\<rangle> [\<rightarrow>] \<langle>es',(h',l',sh'),b'\<rangle> \<Longrightarrow> (\<And>Ts E. \<lbrakk> P,E,h,sh \<turnstile> es [:] Ts; P \<turnstile> h \<surd> \<rbrakk> \<Longrightarrow> P \<turnstile> h' \<surd>)"
+(*<*)
+proof (induct rule:red_reds_inducts)
+ case (RedNew h a C FDTs h' l sh es)
+ have new: "new_Addr h = Some a" and fields: "P \<turnstile> C has_fields FDTs"
+ and h': "h' = h(a\<mapsto>blank P C)"
+ and hconf: "P \<turnstile> h \<surd>" by fact+
+ from new have None: "h a = None" by(rule new_Addr_SomeD)
+ moreover have "P,h \<turnstile> blank P C \<surd>"
+ using fields by(rule oconf_blank)
+ ultimately show "P \<turnstile> h' \<surd>" using h' by(fast intro: hconf_new[OF hconf])
+next
+ case (RedFAss C F t D h a fs v l sh b')
+ let ?fs' = "fs((F,D)\<mapsto>v)"
+ have hconf: "P \<turnstile> h \<surd>" and ha: "h a = Some(C,fs)"
+ and wt: "P,E,h,sh \<turnstile> addr a\<bullet>F{D}:=Val v : T" by fact+
+ from wt ha obtain TF Tv where typeofv: "typeof\<^bsub>h\<^esub> v = Some Tv"
+ and has: "P \<turnstile> C has F,NonStatic:TF in D"
+ and sub: "P \<turnstile> Tv \<le> TF" by auto
+ have "P,h \<turnstile> (C, ?fs') \<surd>"
+ proof (rule oconf_fupd[OF has])
+ show "P,h \<turnstile> (C, fs) \<surd>" using hconf ha by(simp add:hconf_def)
+ show "P,h \<turnstile> v :\<le> TF" using sub typeofv by(simp add:conf_def)
+ qed
+ with hconf ha show "P \<turnstile> h(a\<mapsto>(C, ?fs')) \<surd>" by (rule hconf_upd_obj)
+qed(auto elim: WTrt.cases)
+(*>*)
+
+
+theorem red_preserves_lconf:
+ "P \<turnstile> \<langle>e,(h,l,sh),b\<rangle> \<rightarrow> \<langle>e',(h',l',sh'),b'\<rangle> \<Longrightarrow>
+ (\<And>T E. \<lbrakk> P,E,h,sh \<turnstile> e:T; P,h \<turnstile> l (:\<le>) E \<rbrakk> \<Longrightarrow> P,h' \<turnstile> l' (:\<le>) E)"
+and reds_preserves_lconf:
+ "P \<turnstile> \<langle>es,(h,l,sh),b\<rangle> [\<rightarrow>] \<langle>es',(h',l',sh'),b'\<rangle> \<Longrightarrow>
+ (\<And>Ts E. \<lbrakk> P,E,h,sh \<turnstile> es[:]Ts; P,h \<turnstile> l (:\<le>) E \<rbrakk> \<Longrightarrow> P,h' \<turnstile> l' (:\<le>) E)"
+(*<*)
+proof(induct rule:red_reds_inducts)
+ case RedNew thus ?case
+ by(fast intro:lconf_hext red_hext_incr[OF red_reds.RedNew])
+next
+ case RedLAss thus ?case by(fastforce elim: lconf_upd simp:conf_def)
+next
+ case RedFAss thus ?case
+ by(fast intro:lconf_hext red_hext_incr[OF red_reds.RedFAss])
+next
+ case (InitBlockRed e h l V v sh b e' h' l' sh' b' v' T T')
+ have red: "P \<turnstile> \<langle>e, (h, l(V\<mapsto>v),sh),b\<rangle> \<rightarrow> \<langle>e',(h', l',sh'),b'\<rangle>"
+ and IH: "\<And>T E . \<lbrakk> P,E,h,sh \<turnstile> e:T; P,h \<turnstile> l(V\<mapsto>v) (:\<le>) E \<rbrakk>
+ \<Longrightarrow> P,h' \<turnstile> l' (:\<le>) E"
+ and l'V: "l' V = Some v'" and lconf: "P,h \<turnstile> l (:\<le>) E"
+ and wt: "P,E,h,sh \<turnstile> {V:T := Val v; e} : T'" by fact+
+ from lconf_hext[OF lconf red_hext_incr[OF red]]
+ have "P,h' \<turnstile> l (:\<le>) E" .
+ moreover from IH lconf wt have "P,h' \<turnstile> l' (:\<le>) E(V\<mapsto>T)"
+ by(auto simp del: fun_upd_apply simp: fun_upd_same lconf_upd2 conf_def)
+ ultimately show "P,h' \<turnstile> l'(V := l V) (:\<le>) E"
+ by (fastforce simp:lconf_def split:if_split_asm)
+next
+ case (BlockRedNone e h l V sh b e' h' l' sh' b' T T')
+ have red: "P \<turnstile> \<langle>e,(h, l(V := None),sh),b\<rangle> \<rightarrow> \<langle>e',(h', l',sh'),b'\<rangle>"
+ and IH: "\<And>E T. \<lbrakk> P,E,h,sh \<turnstile> e : T; P,h \<turnstile> l(V:=None) (:\<le>) E \<rbrakk>
+ \<Longrightarrow> P,h' \<turnstile> l' (:\<le>) E"
+ and lconf: "P,h \<turnstile> l (:\<le>) E" and wt: "P,E,h,sh \<turnstile> {V:T; e} : T'" by fact+
+ from lconf_hext[OF lconf red_hext_incr[OF red]]
+ have "P,h' \<turnstile> l (:\<le>) E" .
+ moreover have "P,h' \<turnstile> l' (:\<le>) E(V\<mapsto>T)"
+ by(rule IH, insert lconf wt, auto simp:lconf_def)
+ ultimately show "P,h' \<turnstile> l'(V := l V) (:\<le>) E"
+ by (fastforce simp:lconf_def split:if_split_asm)
+next
+ case (BlockRedSome e h l V sh b e' h' l' sh' b' v T T')
+ have red: "P \<turnstile> \<langle>e,(h, l(V := None),sh),b\<rangle> \<rightarrow> \<langle>e',(h', l',sh'),b'\<rangle>"
+ and IH: "\<And>E T. \<lbrakk>P,E,h,sh \<turnstile> e : T; P,h \<turnstile> l(V:=None) (:\<le>) E\<rbrakk>
+ \<Longrightarrow> P,h' \<turnstile> l' (:\<le>) E"
+ and lconf: "P,h \<turnstile> l (:\<le>) E" and wt: "P,E,h,sh \<turnstile> {V:T; e} : T'" by fact+
+ from lconf_hext[OF lconf red_hext_incr[OF red]]
+ have "P,h' \<turnstile> l (:\<le>) E" .
+ moreover have "P,h' \<turnstile> l' (:\<le>) E(V\<mapsto>T)"
+ by(rule IH, insert lconf wt, auto simp:lconf_def)
+ ultimately show "P,h' \<turnstile> l'(V := l V) (:\<le>) E"
+ by (fastforce simp:lconf_def split:if_split_asm)
+qed(auto elim: WTrt.cases)
+(*>*)
+
+
+theorem red_preserves_shconf:
+ "P \<turnstile> \<langle>e,(h,l,sh),b\<rangle> \<rightarrow> \<langle>e',(h',l',sh'),b'\<rangle> \<Longrightarrow> (\<And>T E. \<lbrakk> P,E,h,sh \<turnstile> e : T; P,h \<turnstile>\<^sub>s sh \<surd> \<rbrakk> \<Longrightarrow> P,h' \<turnstile>\<^sub>s sh' \<surd>)"
+and reds_preserves_shconf:
+ "P \<turnstile> \<langle>es,(h,l,sh),b\<rangle> [\<rightarrow>] \<langle>es',(h',l',sh'),b'\<rangle> \<Longrightarrow> (\<And>Ts E. \<lbrakk> P,E,h,sh \<turnstile> es [:] Ts; P,h \<turnstile>\<^sub>s sh \<surd> \<rbrakk> \<Longrightarrow> P,h' \<turnstile>\<^sub>s sh' \<surd>)"
+(*<*)
+proof (induct rule:red_reds_inducts)
+ case (RedNew h a C FDTs h' l sh es)
+ have new: "new_Addr h = Some a"
+ and h': "h' = h(a\<mapsto>blank P C)"
+ and shconf: "P,h \<turnstile>\<^sub>s sh \<surd>" by fact+
+ from new have None: "h a = None" by(rule new_Addr_SomeD)
+ then show "P,h' \<turnstile>\<^sub>s sh \<surd>" using h' by(fast intro: shconf_hnew[OF shconf])
+next
+ case (RedFAss C F t D h a fs v l sh b)
+ let ?fs' = "fs((F,D)\<mapsto>v)"
+ have shconf: "P,h \<turnstile>\<^sub>s sh \<surd>" and ha: "h a = Some(C,fs)" by fact+
+ then show "P,h(a\<mapsto>(C, ?fs')) \<turnstile>\<^sub>s sh \<surd>" by (rule shconf_hupd_obj)
+next
+ case (RedSFAss C F t D sh sfs i sfs' v sh' h l)
+ let ?sfs' = "sfs(F\<mapsto>v)"
+ have shconf: "P,h \<turnstile>\<^sub>s sh \<surd>" and shD: "sh D = \<lfloor>(sfs, i)\<rfloor>"
+ and wt: "P,E,h,sh \<turnstile> C\<bullet>\<^sub>sF{D} := Val v : T" by fact+
+ from wt obtain TF Tv where typeofv: "typeof\<^bsub>h\<^esub> v = Some Tv"
+ and has: "P \<turnstile> C has F,Static:TF in D"
+ and sub: "P \<turnstile> Tv \<le> TF" by (auto elim: WTrt.cases)
+ have has': "P \<turnstile> D has F,Static:TF in D" using has by(rule has_field_idemp)
+ have "P,h,D \<turnstile>\<^sub>s ?sfs' \<surd>"
+ proof (rule soconf_fupd[OF has'])
+ show "P,h,D \<turnstile>\<^sub>s sfs \<surd>" using shconf shD by(simp add:shconf_def)
+ show "P,h \<turnstile> v :\<le> TF" using sub typeofv by(simp add:conf_def)
+ qed
+ with shconf have "P,h \<turnstile>\<^sub>s sh(D\<mapsto>(?sfs',i)) \<surd>" by (rule shconf_upd_obj)
+ then show ?case using RedSFAss.hyps(3) RedSFAss.hyps(4) by blast
+next
+ case (InitNoneRed sh C C' Cs e h l)
+ let ?sfs' = "sblank P C"
+ have "P,h \<turnstile>\<^sub>s sh(C \<mapsto> (?sfs', Prepared)) \<surd>"
+ proof(rule shconf_upd_obj)
+ show "P,h \<turnstile>\<^sub>s sh \<surd>" using InitNoneRed by simp
+ show "P,h,C \<turnstile>\<^sub>s sblank P C \<surd>" by (metis has_field_def soconf_def soconf_sblank)
+ qed
+ then show ?case by blast
+next
+ case (InitObjectRed sh C sfs sh' C' Cs e h l)
+ have sh': "sh' = sh(C \<mapsto> (sfs, Processing))" by fact
+ have "P,h \<turnstile>\<^sub>s sh(C \<mapsto> (sfs, Processing)) \<surd>"
+ proof(rule shconf_upd_obj)
+ show "P,h \<turnstile>\<^sub>s sh \<surd>" using InitObjectRed by simp
+ moreover have "sh C = \<lfloor>(sfs, Prepared)\<rfloor>" using InitObjectRed by simp
+ ultimately show "P,h,C \<turnstile>\<^sub>s sfs \<surd>" using shconfD by blast
+ qed
+ then show ?case using sh' by blast
+next
+ case (InitNonObjectSuperRed sh C sfs D a b sh' C' Cs e h l)
+ have sh': "sh' = sh(C \<mapsto> (sfs, Processing))" by fact
+ have "P,h \<turnstile>\<^sub>s sh(C \<mapsto> (sfs, Processing)) \<surd>"
+ proof(rule shconf_upd_obj)
+ show "P,h \<turnstile>\<^sub>s sh \<surd>" using InitNonObjectSuperRed by simp
+ moreover have "sh C = \<lfloor>(sfs, Prepared)\<rfloor>" using InitNonObjectSuperRed by simp
+ ultimately show "P,h,C \<turnstile>\<^sub>s sfs \<surd>" using shconfD by blast
+ qed
+ then show ?case using sh' by blast
+next
+ case (RedRInit sh C sfs i sh' C' Cs e v h l)
+ have sh': "sh' = sh(C \<mapsto> (sfs, Done))" by fact
+ have "P,h \<turnstile>\<^sub>s sh(C \<mapsto> (sfs, Done)) \<surd>"
+ proof(rule shconf_upd_obj)
+ show "P,h \<turnstile>\<^sub>s sh \<surd>" using RedRInit by simp
+ moreover have "sh C = \<lfloor>(sfs, i)\<rfloor>" using RedRInit by simp
+ ultimately show "P,h,C \<turnstile>\<^sub>s sfs \<surd>" using shconfD by blast
+ qed
+ then show ?case using sh' by blast
+next
+ case (RInitInitThrow sh C sfs i sh' a D Cs e h l)
+ have sh': "sh' = sh(C \<mapsto> (sfs, Error))" by fact
+ have "P,h \<turnstile>\<^sub>s sh(C \<mapsto> (sfs, Error)) \<surd>"
+ proof(rule shconf_upd_obj)
+ show "P,h \<turnstile>\<^sub>s sh \<surd>" using RInitInitThrow by simp
+ moreover have "sh C = \<lfloor>(sfs, i)\<rfloor>" using RInitInitThrow by simp
+ ultimately show "P,h,C \<turnstile>\<^sub>s sfs \<surd>" using shconfD by blast
+ qed
+ then show ?case using sh' by blast
+next
+ case (RInitThrow sh C sfs i sh' a e' h l)
+ have sh': "sh' = sh(C \<mapsto> (sfs, Error))" by fact
+ have "P,h \<turnstile>\<^sub>s sh(C \<mapsto> (sfs, Error)) \<surd>"
+ proof(rule shconf_upd_obj)
+ show "P,h \<turnstile>\<^sub>s sh \<surd>" using RInitThrow by simp
+ moreover have "sh C = \<lfloor>(sfs, i)\<rfloor>" using RInitThrow by simp
+ ultimately show "P,h,C \<turnstile>\<^sub>s sfs \<surd>" using shconfD by blast
+ qed
+ then show ?case using sh' by blast
+qed(auto elim: WTrt.cases)
+(*>*)
+
+theorem assumes wf: "wwf_J_prog P"
+shows red_preserves_iconf:
+ "P \<turnstile> \<langle>e,(h,l,sh),b\<rangle> \<rightarrow> \<langle>e',(h',l',sh'),b'\<rangle> \<Longrightarrow> iconf sh e \<Longrightarrow> iconf sh' e'"
+and reds_preserves_iconf:
+ "P \<turnstile> \<langle>es,(h,l,sh),b\<rangle> [\<rightarrow>] \<langle>es',(h',l',sh'),b'\<rangle> \<Longrightarrow> iconfs sh es \<Longrightarrow> iconfs sh' es'"
+(*<*)
+proof (induct rule:red_reds_inducts)
+ case (BinOpRed1 e h l sh b e' h' l' sh' b' bop e\<^sub>2)
+ then show ?case using BinOpRed1 nsub_RI_iconf[of e\<^sub>2 sh'] val_of_spec
+ proof(cases "val_of e") qed(simp,fast)
+next
+ case (FAssRed1 e h l sh b e' h' l' sh' b' F D e\<^sub>2)
+ then show ?case using FAssRed1 nsub_RI_iconf[of e\<^sub>2 sh'] val_of_spec
+ proof(cases "val_of e") qed(simp,fast)
+next
+ case (CallObj e h l sh b e' h' l' sh' b' M es)
+ then show ?case using CallObj nsub_RIs_iconfs[of es sh'] val_of_spec
+ proof(cases "val_of e") qed(simp,fast)
+next
+ case RedCall then show ?case using sees_wwf_nsub_RI[OF wf RedCall.hyps(2)]
+ by (clarsimp simp: assigned_def nsub_RI_iconf)
+next
+ case (RedSCall C M Ts T pns body D vs a a b)
+ then have "\<not>sub_RI (blocks (pns, Ts, vs, body))"
+ using sees_wwf_nsub_RI[OF wf RedSCall.hyps(1)] by simp
+ then show ?case by(rule nsub_RI_iconf)
+next
+ case SCallInitRed then show ?case by fastforce
+next
+ case (BlockRedNone e h l V sh b e' h' l' sh' b' T)
+ then show ?case by auto
+next
+ case (SeqRed e h l sh b e' h' l' sh' b' e\<^sub>2)
+ then show ?case
+ proof(cases "lass_val_of e")
+ case None then show ?thesis using SeqRed nsub_RI_iconf by(auto dest: val_of_spec)
+ next
+ case (Some a)
+ have "e' = unit \<and> sh' = sh" by(simp add: lass_val_of_red[OF Some SeqRed(1)])
+ then show ?thesis using SeqRed Some by(auto dest: val_of_spec)
+ qed
+next
+ case (ListRed1 e h l sh b e' h' l' sh' b' es)
+ then show ?case using ListRed1 nsub_RIs_iconfs[of es sh'] val_of_spec
+ proof(cases "val_of e") qed(simp,fast)
+next
+ case RedInit then show ?case by(auto simp: nsub_RI_iconf)
+next
+ case (RedInitDone sh C sfs C' Cs e h l b)
+ then show ?case proof(cases Cs) qed(auto simp: initPD_def)
+next
+ case (RedInitProcessing sh C sfs C' Cs e h l b)
+ then show ?case proof(cases Cs) qed(auto simp: initPD_def)
+next
+ case (RedRInit sh C sfs i sh' C' Cs v e h l b)
+ then show ?case proof(cases Cs) qed(auto simp: initPD_def)
+next
+ case CallThrowParams then show ?case by(auto simp: iconfs_map_throw)
+next
+ case SCallThrowParams then show ?case by(auto simp: iconfs_map_throw)
+qed(auto simp: nsub_RI_iconf lass_val_of_iconf)
+(*>*)
+
+
+lemma Seq_bconf_preserve_aux:
+assumes "P \<turnstile> \<langle>e,(h, l, sh),b\<rangle> \<rightarrow> \<langle>e',(h', l', sh'),b'\<rangle>" and "P,sh \<turnstile>\<^sub>b (e;; e\<^sub>2,b) \<surd>"
+ and "P,sh \<turnstile>\<^sub>b (e::expr,b) \<surd> \<longrightarrow> P,sh' \<turnstile>\<^sub>b (e'::expr,b') \<surd>"
+shows "P,sh' \<turnstile>\<^sub>b (e';;e\<^sub>2,b') \<surd>"
+proof(cases "val_of e")
+ case None show ?thesis
+ proof(cases "lass_val_of e")
+ case lNone: None show ?thesis
+ proof(cases "lass_val_of e'")
+ case lNone': None
+ then show ?thesis using None assms lNone
+ by(auto dest: val_of_spec simp: bconf_def option.distinct(1))
+ next
+ case (Some a)
+ then show ?thesis using None assms lNone by(auto dest: lass_val_of_spec simp: bconf_def)
+ qed
+ next
+ case (Some a)
+ then show ?thesis using None assms by(auto dest: lass_val_of_spec)
+ qed
+next
+ case (Some a)
+ then show ?thesis using assms by(auto dest: val_of_spec)
+qed
+
+theorem red_preserves_bconf:
+ "P \<turnstile> \<langle>e,(h,l,sh),b\<rangle> \<rightarrow> \<langle>e',(h',l',sh'),b'\<rangle> \<Longrightarrow> iconf sh e \<Longrightarrow> P,sh \<turnstile>\<^sub>b (e,b) \<surd> \<Longrightarrow> P,sh' \<turnstile>\<^sub>b (e',b') \<surd>"
+and reds_preserves_bconf:
+ "P \<turnstile> \<langle>es,(h,l,sh),b\<rangle> [\<rightarrow>] \<langle>es',(h',l',sh'),b'\<rangle> \<Longrightarrow> iconfs sh es \<Longrightarrow> P,sh \<turnstile>\<^sub>b (es,b) \<surd> \<Longrightarrow> P,sh' \<turnstile>\<^sub>b (es',b') \<surd>"
+(*<*)
+proof (induct rule:red_reds_inducts)
+ case (CastRed e a a b b e' a a b b' C) then show ?case
+ proof(cases b') qed(simp, simp add: bconf_def)
+next
+ case (BinOpRed1 e h l sh b e' h' l' sh' b' bop e\<^sub>2) show ?case
+ proof(cases b')
+ case True with BinOpRed1 val_of_spec show ?thesis
+ proof(cases "val_of e") qed(simp,fast)
+ next
+ case False then show ?thesis by (simp add: bconf_def)
+ qed
+next
+case (BinOpRed2 e a a b b e' a a b b' v\<^sub>1 bop) then show ?case
+ proof(cases b') qed(simp, simp add: bconf_def)
+next
+ case (LAssRed e a a b b e' a a b b' V) then show ?case
+ proof(cases b') qed(simp, simp add: bconf_def)
+next
+ case (FAccRed e a a b b e' a a b b' F D) then show ?case
+ proof(cases b') qed(simp, simp add: bconf_def)
+next
+ case (RedSFAccNonStatic C F t D h l sh b) then show ?case
+ using has_field_fun[of P C F NonStatic] by (auto simp: bconf_def)
+next
+ case (FAssRed1 e h l sh b e' h' l' sh' b' F D e\<^sub>2) show ?case
+ proof(cases b')
+ case True with FAssRed1 val_of_spec show ?thesis
+ proof(cases "val_of e'")qed((simp,fast)+)
+ next
+ case False then show ?thesis by(simp add: bconf_def)
+ qed
+next
+ case (FAssRed2 e a a b b e' a a b b' v F D) then show ?case
+ proof(cases b') qed(simp, simp add: bconf_def)
+next
+ case (SFAssRed e h l sh b e' h' l' sh' b' C F D) then show ?case
+ proof(cases b') qed(fastforce simp: bconf_def val_no_step)+
+next
+ case (RedSFAssNonStatic C F t D v a a b b) then show ?case
+ using has_field_fun[of P C F NonStatic] by (auto simp: bconf_def)
+next
+ case (CallObj e h l sh b e' h' l' sh' b' M es) show ?case
+ proof(cases b')
+ case True with CallObj val_of_spec show ?thesis
+ proof(cases "val_of e'")qed((simp,fast)+)
+ next
+ case False then show ?thesis by(simp add: bconf_def)
+ qed
+next
+ case (CallParams es a a b b es' a a b b' v M) then show ?case
+ proof(cases b') qed(simp, simp add: bconf_def)
+next
+ case (SCallParams es h l sh b es' h' l' sh' b' C M) show ?case
+ proof(cases b')
+ case b': True show ?thesis
+ proof(cases "map_vals_of es'")
+ case None
+ then show ?thesis using SCallParams b' vals_no_step
+ proof(cases "map_vals_of es")qed(clarsimp,fast)
+ next
+ case f: Some
+ then show ?thesis using SCallParams b' vals_no_step
+ proof(cases "map_vals_of es")qed(clarsimp,fast)
+ qed
+ next
+ case False then show ?thesis by(simp add: bconf_def)
+ qed
+next
+ case (SCallInitDoneRed C M Ts T pns body D sh sfs vs h l)
+ then show ?case by(auto simp: bconf_def initPD_def) (rule_tac x=D in exI, auto)+
+next
+ case (RedSCallNonStatic C M Ts T a b D vs h l sh b) then show ?case
+ using sees_method_fun[of P C M NonStatic] by (auto simp: bconf_def)
+next
+ case (BlockRedNone e h l V sh b e' h' l' sh' b' T) show ?case
+ proof(cases "assigned V e'")
+ case True
+ then obtain v e2 where "e' = V := Val v;;e2" by(clarsimp simp: assigned_def)
+ then show ?thesis using BlockRedNone by(clarsimp simp: bconf_def)
+ next
+ case False then show ?thesis using BlockRedNone by simp
+ qed
+next
+ case (BlockRedSome e h l V sh b e' h' l' sh' b' v T) then show ?case
+ proof(cases b') qed(simp, simp add: bconf_def)
+next
+ case (InitBlockRed e h l V v sh b e' h' l' sh' b' v' T) show ?case
+ proof(cases b')
+ case True
+ then show ?thesis using InitBlockRed by (simp add: assigned_def)
+ next
+ case False then show ?thesis by(simp add: bconf_def)
+ qed
+next
+ case (RedBlock V T u)
+ then have "\<not>assigned V (Val u)" by(clarsimp simp: assigned_def)
+ then show ?case using RedBlock by(simp add: bconf_def)
+next
+ case (SeqRed e h l sh b e' h' l' sh' b' e\<^sub>2)
+ have "iconf sh e"
+ proof(cases "lass_val_of e")
+ case (Some a) show ?thesis by(rule lass_val_of_iconf[OF Some])
+ next
+ case None then show ?thesis using SeqRed by(auto dest: val_of_spec)
+ qed
+ then show ?case using SeqRed Seq_bconf_preserve_aux by simp
+next
+ case (CondRed e a a b b e' a a b b' e\<^sub>1 e\<^sub>2) then show ?case
+ proof(cases b') qed(simp, simp add: bconf_def)
+next
+ case (ThrowRed e a a b b e' a a b b') then show ?case
+ proof(cases b') qed(simp, simp add: bconf_def)
+next
+ case (TryRed e a a b b e' a a b b' C V e\<^sub>2) then show ?case
+ proof(cases b') qed(simp, simp add: bconf_def)
+next
+ case (ListRed1 e h l sh b e' h' l' sh' b' es)
+ with val_of_spec show ?case
+ proof(cases b') qed((clarsimp,fast),simp add: bconfs_def)
+next
+ case (RedInit C b' e X Y b b'')
+ then show ?case
+ by(auto simp: bconf_def icheck_ss_exp icheck_init_class icheck_curr_init)
+next
+ case (RInitRed e a a b b e' a a b b' C Cs e\<^sub>0) then show ?case
+ proof(cases b') qed(simp, simp add: bconf_def)
+next
+ case (BlockThrow V T a)
+ then have "\<not>assigned V (Throw a)" by(simp add: assigned_def)
+ then show ?case using BlockThrow by simp
+qed(simp_all, auto simp: bconf_def initPD_def)
+(*>*)
+
+text\<open> Preservation of definite assignment more complex and requires a
+few lemmas first. \<close>
+
+lemma [iff]: "\<And>A. \<lbrakk> length Vs = length Ts; length vs = length Ts\<rbrakk> \<Longrightarrow>
+ \<D> (blocks (Vs,Ts,vs,e)) A = \<D> e (A \<squnion> \<lfloor>set Vs\<rfloor>)"
+(*<*)
+by (induct Vs Ts vs e rule:blocks_induct)
+ (simp_all add:hyperset_defs)
+(*>*)
+
+
+lemma red_lA_incr: "P \<turnstile> \<langle>e,(h,l,sh),b\<rangle> \<rightarrow> \<langle>e',(h',l',sh'),b'\<rangle>
+ \<Longrightarrow> \<lfloor>dom l\<rfloor> \<squnion> \<A> e \<sqsubseteq> \<lfloor>dom l'\<rfloor> \<squnion> \<A> e'"
+and reds_lA_incr: "P \<turnstile> \<langle>es,(h,l,sh),b\<rangle> [\<rightarrow>] \<langle>es',(h',l',sh'),b'\<rangle>
+ \<Longrightarrow> \<lfloor>dom l\<rfloor> \<squnion> \<A>s es \<sqsubseteq> \<lfloor>dom l'\<rfloor> \<squnion> \<A>s es'"
+(*<*)
+proof(induct rule:red_reds_inducts)
+ case TryRed then show ?case
+ by(simp del:fun_upd_apply add:hyperset_defs)
+ (blast dest:red_lcl_incr)+
+next
+ case BinOpRed1 then show ?case
+ by(auto simp del:fun_upd_apply simp:hyperset_defs)
+next
+ case BinOpRed2 then show ?case
+ by(auto simp del:fun_upd_apply simp:hyperset_defs)
+next
+ case LAssRed then show ?case
+ by(auto simp del:fun_upd_apply simp:hyperset_defs)
+next
+ case FAssRed1 then show ?case
+ by(auto simp del:fun_upd_apply simp:hyperset_defs)
+next
+ case FAssRed2 then show ?case
+ by(auto simp del:fun_upd_apply simp:hyperset_defs)
+next
+ case CallObj then show ?case
+ by(auto simp del:fun_upd_apply simp:hyperset_defs)
+next
+ case CallParams then show ?case
+ by(auto simp del:fun_upd_apply simp:hyperset_defs)
+next
+ case BlockRedNone then show ?case
+ by(auto simp del:fun_upd_apply simp:hyperset_defs)
+next
+ case BlockRedSome then show ?case
+ by(auto simp del:fun_upd_apply simp:hyperset_defs)
+next
+ case InitBlockRed then show ?case
+ by(auto simp del:fun_upd_apply simp:hyperset_defs)
+next
+ case SeqRed then show ?case
+ by(auto simp del:fun_upd_apply simp:hyperset_defs)
+next
+ case CondRed then show ?case
+ by(auto simp del:fun_upd_apply simp:hyperset_defs)
+next
+ case RedCondT then show ?case
+ by(auto simp del:fun_upd_apply simp:hyperset_defs)
+next
+ case RedCondF then show ?case
+ by(auto simp del:fun_upd_apply simp:hyperset_defs)
+next
+ case ListRed1 then show ?case
+ by(auto simp del:fun_upd_apply simp:hyperset_defs)
+next
+ case ListRed2 then show ?case
+ by(auto simp del:fun_upd_apply simp:hyperset_defs)
+qed (simp_all del:fun_upd_apply add:hyperset_defs)
+(*>*)
+
+text\<open> Now preservation of definite assignment. \<close>
+
+lemma assumes wf: "wf_J_prog P"
+shows red_preserves_defass:
+ "P \<turnstile> \<langle>e,(h,l,sh),b\<rangle> \<rightarrow> \<langle>e',(h',l',sh'),b'\<rangle> \<Longrightarrow> \<D> e \<lfloor>dom l\<rfloor> \<Longrightarrow> \<D> e' \<lfloor>dom l'\<rfloor>"
+and "P \<turnstile> \<langle>es,(h,l,sh),b\<rangle> [\<rightarrow>] \<langle>es',(h',l',sh'),b'\<rangle> \<Longrightarrow> \<D>s es \<lfloor>dom l\<rfloor> \<Longrightarrow> \<D>s es' \<lfloor>dom l'\<rfloor>"
+(*<*)
+proof (induct rule:red_reds_inducts)
+ case BinOpRed1 thus ?case by (auto elim!: D_mono[OF red_lA_incr])
+next
+ case FAssRed1 thus ?case by (auto elim!: D_mono[OF red_lA_incr])
+next
+ case CallObj thus ?case by (auto elim!: Ds_mono[OF red_lA_incr])
+next
+ case RedCall thus ?case
+ by (auto dest!:sees_wf_mdecl[OF wf] simp:wf_mdecl_def hyperset_defs elim!:D_mono')
+next
+ case RedSCall thus ?case
+ by (auto dest!:sees_wf_mdecl[OF wf] simp:wf_mdecl_def hyperset_defs elim!:D_mono')
+next
+ case SCallInitRed
+ then show ?case by(auto simp:hyperset_defs Ds_Vals)
+next
+ case InitBlockRed thus ?case
+ by(auto simp:hyperset_defs elim!:D_mono' simp del:fun_upd_apply)
+next
+ case BlockRedNone thus ?case
+ by(auto simp:hyperset_defs elim!:D_mono' simp del:fun_upd_apply)
+next
+ case BlockRedSome thus ?case
+ by(auto simp:hyperset_defs elim!:D_mono' simp del:fun_upd_apply)
+next
+ case SeqRed thus ?case by (auto elim!: D_mono[OF red_lA_incr])
+next
+ case CondRed thus ?case by (auto elim!: D_mono[OF red_lA_incr])
+next
+ case TryRed thus ?case
+ by (fastforce dest:red_lcl_incr intro:D_mono' simp:hyperset_defs)
+next
+ case ListRed1 thus ?case by (auto elim!: Ds_mono[OF red_lA_incr])
+next
+ case RedWhile thus ?case by(auto simp:hyperset_defs elim!:D_mono')
+next
+ case RedInit then show ?case by (auto intro: D_mono' simp: hyperset_defs)
+next
+ case (RInitRed e h l sh b e' h' l' sh' b' C Cs e\<^sub>0)
+ then show ?case by(auto simp:hyperset_defs dest:red_lcl_incr elim!:D_mono')
+qed(auto simp:hyperset_defs)
+(*>*)
+
+
+text\<open> Combining conformance of heap, static heap, and local variables: \<close>
+
+definition sconf :: "J_prog \<Rightarrow> env \<Rightarrow> state \<Rightarrow> bool" ("_,_ \<turnstile> _ \<surd>" [51,51,51]50)
+where
+ "P,E \<turnstile> s \<surd> \<equiv> let (h,l,sh) = s in P \<turnstile> h \<surd> \<and> P,h \<turnstile> l (:\<le>) E \<and> P,h \<turnstile>\<^sub>s sh \<surd>"
+
+lemma red_preserves_sconf:
+ "\<lbrakk> P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow> \<langle>e',s',b'\<rangle>; P,E,hp s,shp s \<turnstile> e : T; P,E \<turnstile> s \<surd> \<rbrakk> \<Longrightarrow> P,E \<turnstile> s' \<surd>"
+(*<*)
+by(fastforce intro:red_preserves_hconf red_preserves_lconf red_preserves_shconf
+ simp add:sconf_def)
+(*>*)
+
+lemma reds_preserves_sconf:
+ "\<lbrakk> P \<turnstile> \<langle>es,s,b\<rangle> [\<rightarrow>] \<langle>es',s',b'\<rangle>; P,E,hp s,shp s \<turnstile> es [:] Ts; P,E \<turnstile> s \<surd> \<rbrakk> \<Longrightarrow> P,E \<turnstile> s' \<surd>"
+(*<*)
+by(fastforce intro:reds_preserves_hconf reds_preserves_lconf reds_preserves_shconf
+ simp add:sconf_def)
+(*>*)
+
+
+subsection "Subject reduction"
+
+lemma wt_blocks:
+ "\<And>E. \<lbrakk> length Vs = length Ts; length vs = length Ts \<rbrakk> \<Longrightarrow>
+ (P,E,h,sh \<turnstile> blocks(Vs,Ts,vs,e) : T) =
+ (P,E(Vs[\<mapsto>]Ts),h,sh \<turnstile> e:T \<and> (\<exists>Ts'. map (typeof\<^bsub>h\<^esub>) vs = map Some Ts' \<and> P \<turnstile> Ts' [\<le>] Ts))"
+(*<*)
+proof(induct Vs Ts vs e rule:blocks_induct)
+ case (1 V Vs T Ts v vs e)
+ then show ?case by(force simp add:rel_list_all2_Cons2)
+qed simp_all
+(*>*)
+
+theorem assumes wf: "wf_J_prog P"
+shows subject_reduction2: "P \<turnstile> \<langle>e,(h,l,sh),b\<rangle> \<rightarrow> \<langle>e',(h',l',sh'),b'\<rangle> \<Longrightarrow>
+ (\<And>E T. \<lbrakk> P,E \<turnstile> (h,l,sh) \<surd>; iconf sh e; P,E,h,sh \<turnstile> e:T \<rbrakk>
+ \<Longrightarrow> \<exists>T'. P,E,h',sh' \<turnstile> e':T' \<and> P \<turnstile> T' \<le> T)"
+and subjects_reduction2: "P \<turnstile> \<langle>es,(h,l,sh),b\<rangle> [\<rightarrow>] \<langle>es',(h',l',sh'),b'\<rangle> \<Longrightarrow>
+ (\<And>E Ts. \<lbrakk> P,E \<turnstile> (h,l,sh) \<surd>; iconfs sh es; P,E,h,sh \<turnstile> es [:] Ts \<rbrakk>
+ \<Longrightarrow> \<exists>Ts'. P,E,h',sh' \<turnstile> es' [:] Ts' \<and> P \<turnstile> Ts' [\<le>] Ts)"
+(*<*)
+proof (induct rule:red_reds_inducts)
+ case RedNew then show ?case by (auto simp: blank_def)
+next
+ case RedNewFail thus ?case
+ by (unfold sconf_def hconf_def) (fastforce elim!:typeof_OutOfMemory)
+next
+ case CastRed thus ?case
+ by(clarsimp simp:is_refT_def)
+ (blast intro: widens_trans dest!:widen_Class[THEN iffD1])
+next
+ case RedCastFail thus ?case
+ by (unfold sconf_def hconf_def) (fastforce elim!:typeof_ClassCast)
+next
+ case (BinOpRed1 e\<^sub>1 h l sh b e\<^sub>1' h' l' sh' b' bop e\<^sub>2)
+ have red: "P \<turnstile> \<langle>e\<^sub>1,(h,l,sh),b\<rangle> \<rightarrow> \<langle>e\<^sub>1',(h',l',sh'),b'\<rangle>"
+ and IH: "\<And>E T. \<lbrakk>P,E \<turnstile> (h,l,sh) \<surd>; iconf sh e\<^sub>1; P,E,h,sh \<turnstile> e\<^sub>1:T\<rbrakk>
+ \<Longrightarrow> \<exists>U. P,E,h',sh' \<turnstile> e\<^sub>1' : U \<and> P \<turnstile> U \<le> T"
+ and conf: "P,E \<turnstile> (h,l,sh) \<surd>" and iconf: "iconf sh (e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2)"
+ and wt: "P,E,h,sh \<turnstile> e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2 : T" by fact+
+ have val: "val_of e\<^sub>1 = None" using red iconf val_no_step by auto
+ then have iconf1: "iconf sh e\<^sub>1" and nsub_RI2: "\<not>sub_RI e\<^sub>2" using iconf by simp+
+ have "P,E,h',sh' \<turnstile> e\<^sub>1' \<guillemotleft>bop\<guillemotright> e\<^sub>2 : T"
+ proof (cases bop)
+ assume [simp]: "bop = Eq"
+ from wt obtain T\<^sub>1 T\<^sub>2 where [simp]: "T = Boolean"
+ and wt\<^sub>1: "P,E,h,sh \<turnstile> e\<^sub>1 : T\<^sub>1" and wt\<^sub>2: "P,E,h,sh \<turnstile> e\<^sub>2 : T\<^sub>2" by auto
+ show ?thesis
+ using WTrt_hext_shext_mono[OF wt\<^sub>2 red_hext_incr[OF red] red_shext_incr[OF red wt\<^sub>1] nsub_RI2]
+ IH[OF conf iconf1 wt\<^sub>1] by auto
+ next
+ assume [simp]: "bop = Add"
+ from wt have [simp]: "T = Integer"
+ and wt\<^sub>1: "P,E,h,sh \<turnstile> e\<^sub>1 : Integer" and wt\<^sub>2: "P,E,h,sh \<turnstile> e\<^sub>2 : Integer"
+ by auto
+ show ?thesis
+ using WTrt_hext_shext_mono[OF wt\<^sub>2 red_hext_incr[OF red] red_shext_incr[OF red wt\<^sub>1] nsub_RI2]
+ IH[OF conf iconf1 wt\<^sub>1] by auto
+ qed
+ thus ?case by auto
+next
+ case (BinOpRed2 e\<^sub>2 h l sh b e\<^sub>2' h' l' sh' b' v\<^sub>1 bop)
+ have red: "P \<turnstile> \<langle>e\<^sub>2,(h,l,sh),b\<rangle> \<rightarrow> \<langle>e\<^sub>2',(h',l',sh'),b'\<rangle>"
+ and IH: "\<And>E T. \<lbrakk>P,E \<turnstile> (h,l,sh) \<surd>; iconf sh e\<^sub>2; P,E,h,sh \<turnstile> e\<^sub>2:T\<rbrakk>
+ \<Longrightarrow> \<exists>U. P,E,h',sh' \<turnstile> e\<^sub>2' : U \<and> P \<turnstile> U \<le> T"
+ and conf: "P,E \<turnstile> (h,l,sh) \<surd>" and iconf: "iconf sh (Val v\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2)"
+ and wt: "P,E,h,sh \<turnstile> (Val v\<^sub>1) \<guillemotleft>bop\<guillemotright> e\<^sub>2 : T" by fact+
+ have iconf2: "iconf sh e\<^sub>2" using iconf by simp
+ have "P,E,h',sh' \<turnstile> (Val v\<^sub>1) \<guillemotleft>bop\<guillemotright> e\<^sub>2' : T"
+ proof (cases bop)
+ assume [simp]: "bop = Eq"
+ from wt obtain T\<^sub>1 T\<^sub>2 where [simp]: "T = Boolean"
+ and wt\<^sub>1: "P,E,h,sh \<turnstile> Val v\<^sub>1 : T\<^sub>1" and wt\<^sub>2: "P,E,h,sh \<turnstile> e\<^sub>2:T\<^sub>2" by auto
+ show ?thesis
+ using IH[OF conf iconf2 wt\<^sub>2] WTrt_hext_mono[OF wt\<^sub>1 red_hext_incr[OF red]]
+ by auto
+ next
+ assume [simp]: "bop = Add"
+ from wt have [simp]: "T = Integer"
+ and wt\<^sub>1: "P,E,h,sh \<turnstile> Val v\<^sub>1 : Integer" and wt\<^sub>2: "P,E,h,sh \<turnstile> e\<^sub>2 : Integer"
+ by auto
+ show ?thesis
+ using IH[OF conf iconf2 wt\<^sub>2] WTrt_hext_mono[OF wt\<^sub>1 red_hext_incr[OF red]]
+ by auto
+ qed
+ thus ?case by auto
+next
+ case (RedBinOp bop) thus ?case
+ proof (cases bop)
+ case Eq thus ?thesis using RedBinOp by auto
+ next
+ case Add thus ?thesis using RedBinOp by auto
+ qed
+next
+ case RedVar thus ?case by (fastforce simp:sconf_def lconf_def conf_def)
+next
+ case (LAssRed e h l sh b e' h' l' sh' b' V)
+ obtain Te where Te: "P,E,h,sh \<turnstile> e : Te \<and> P \<turnstile> Te \<le> the(E V)" using LAssRed.prems(3) by auto
+ then have wide: "P \<turnstile> Te \<le> the(E V)" using LAssRed by simp
+ then have "\<exists>T'. P,E,h',sh' \<turnstile> e' : T' \<and> P \<turnstile> T' \<le> Te"
+ using LAssRed.hyps(2) LAssRed.prems(1,2) Te widen_trans[OF _ wide] by auto
+ then obtain T' where wt: "P,E,h',sh' \<turnstile> e' : T' \<and> P \<turnstile> T' \<le> Te" by clarsimp
+ have "P,E,h',sh' \<turnstile> V:=e' : Void" using LAssRed wt widen_trans[OF _ wide] by auto
+ then show ?case using LAssRed by(rule_tac x = Void in exI) auto
+next
+ case (FAccRed e h l sh b e' h' l' sh' b' F D)
+ have IH: "\<And>E T. \<lbrakk>P,E \<turnstile> (h,l,sh) \<surd>; iconf sh e; P,E,h,sh \<turnstile> e : T\<rbrakk>
+ \<Longrightarrow> \<exists>U. P,E,h',sh' \<turnstile> e' : U \<and> P \<turnstile> U \<le> T"
+ and conf: "P,E \<turnstile> (h,l,sh) \<surd>" and iconf: "iconf sh (e\<bullet>F{D})"
+ and wt: "P,E,h,sh \<turnstile> e\<bullet>F{D} : T" by fact+
+ have iconf': "iconf sh e" using iconf by simp+
+ \<comment> \<open>The goal: ?case = @{prop ?case}\<close>
+ \<comment> \<open>Now distinguish the two cases how wt can have arisen.\<close>
+ { fix C assume wte: "P,E,h,sh \<turnstile> e : Class C"
+ and has: "P \<turnstile> C has F,NonStatic:T in D"
+ from IH[OF conf iconf' wte]
+ obtain U where wte': "P,E,h',sh' \<turnstile> e' : U" and UsubC: "P \<turnstile> U \<le> Class C"
+ by auto
+ \<comment> \<open>Now distinguish what @{term U} can be.\<close>
+ { assume "U = NT" hence ?case using wte'
+ by(blast intro:WTrtFAccNT widen_refl) }
+ moreover
+ { fix C' assume U: "U = Class C'" and C'subC: "P \<turnstile> C' \<preceq>\<^sup>* C"
+ from has_field_mono[OF has C'subC] wte' U
+ have ?case by(blast intro:WTrtFAcc) }
+ ultimately have ?case using UsubC by(simp add: widen_Class) blast }
+ moreover
+ { assume "P,E,h,sh \<turnstile> e : NT"
+ hence "P,E,h',sh' \<turnstile> e' : NT" using IH[OF conf iconf'] by fastforce
+ hence ?case by(fastforce intro:WTrtFAccNT widen_refl) }
+ ultimately show ?case using wt by blast
+next
+ case RedFAcc thus ?case
+ by(fastforce simp:sconf_def hconf_def oconf_def conf_def has_field_def
+ dest:has_fields_fun)
+next
+ case RedFAccNull thus ?case
+ by(fastforce intro: widen_refl WTThrow[OF WTVal] elim!: typeof_NullPointer
+ simp: sconf_def hconf_def)
+next
+ case RedSFAccNone then show ?case
+ by(fastforce intro: WTrtThrow[OF WTrtVal] elim!: typeof_NoSuchFieldError
+ simp: sconf_def hconf_def)
+next
+ case RedFAccStatic then show ?case
+ by(fastforce intro: WTrtThrow[OF WTrtVal] elim!: typeof_IncompatibleClassChangeError
+ simp: sconf_def hconf_def)
+next
+ case (RedSFAcc C F t D sh sfs i v h l es)
+ then have "P \<turnstile> C has F,Static:T in D" by fast
+ then have dM: "P \<turnstile> D has F,Static:T in D" by(rule has_field_idemp)
+ then show ?case using RedSFAcc by(fastforce simp:sconf_def shconf_def soconf_def conf_def)
+next
+ case SFAccInitDoneRed then show ?case by (meson widen_refl)
+next
+ case (SFAccInitRed C F t D sh h l E T)
+ have "is_class P D" using SFAccInitRed.hyps(1) by(rule has_field_is_class')
+ then have "P,E,h,sh \<turnstile> INIT D ([D],False) \<leftarrow> C\<bullet>\<^sub>sF{D} : T \<and> P \<turnstile> T \<le> T"
+ using SFAccInitRed WTrtInit[OF SFAccInitRed.prems(3)] by clarsimp
+ then show ?case by(rule exI)
+next
+ case RedSFAccNonStatic then show ?case
+ by(fastforce intro: WTrtThrow[OF WTrtVal] elim!: typeof_IncompatibleClassChangeError
+ simp: sconf_def hconf_def)
+next
+ case (FAssRed1 e h l sh b e' h' l' sh' b' F D e\<^sub>2)
+ have red: "P \<turnstile> \<langle>e,(h,l,sh),b\<rangle> \<rightarrow> \<langle>e',(h',l',sh'),b'\<rangle>"
+ and IH: "\<And>E T. \<lbrakk>P,E \<turnstile> (h,l,sh) \<surd>; iconf sh e; P,E,h,sh \<turnstile> e : T\<rbrakk>
+ \<Longrightarrow> \<exists>U. P,E,h',sh' \<turnstile> e' : U \<and> P \<turnstile> U \<le> T"
+ and conf: "P,E \<turnstile> (h,l,sh) \<surd>" and iconf: "iconf sh (e\<bullet>F{D} := e\<^sub>2)"
+ and wt: "P,E,h,sh \<turnstile> e\<bullet>F{D}:=e\<^sub>2 : T" by fact+
+ have val: "val_of e = None" using red iconf val_no_step by auto
+ then have iconf': "iconf sh e" and nsub_RI2: "\<not>sub_RI e\<^sub>2" using iconf by simp+
+ from wt have void: "T = Void" by blast
+ \<comment> \<open>We distinguish if @{term e} has type @{term NT} or a Class type\<close>
+ \<comment> \<open>Remember ?case = @{term ?case}\<close>
+ { assume wt':"P,E,h,sh \<turnstile> e : NT"
+ hence "P,E,h',sh' \<turnstile> e' : NT" using IH[OF conf iconf'] by fastforce
+ moreover obtain T\<^sub>2 where "P,E,h,sh \<turnstile> e\<^sub>2 : T\<^sub>2" using wt by auto
+ from this red_hext_incr[OF red] red_shext_incr[OF red wt'] nsub_RI2 have "P,E,h',sh' \<turnstile> e\<^sub>2 : T\<^sub>2"
+ by(rule WTrt_hext_shext_mono)
+ ultimately have ?case using void by(blast intro!:WTrtFAssNT)
+ }
+ moreover
+ { fix C TF T\<^sub>2 assume wt\<^sub>1: "P,E,h,sh \<turnstile> e : Class C" and wt\<^sub>2: "P,E,h,sh \<turnstile> e\<^sub>2 : T\<^sub>2"
+ and has: "P \<turnstile> C has F,NonStatic:TF in D" and sub: "P \<turnstile> T\<^sub>2 \<le> TF"
+ obtain U where wt\<^sub>1': "P,E,h',sh' \<turnstile> e' : U" and UsubC: "P \<turnstile> U \<le> Class C"
+ using IH[OF conf iconf' wt\<^sub>1] by blast
+ have wt\<^sub>2': "P,E,h',sh' \<turnstile> e\<^sub>2 : T\<^sub>2"
+ by(rule WTrt_hext_shext_mono[OF wt\<^sub>2 red_hext_incr[OF red] red_shext_incr[OF red wt\<^sub>1] nsub_RI2])
+ \<comment> \<open>Is @{term U} the null type or a class type?\<close>
+ { assume "U = NT" with wt\<^sub>1' wt\<^sub>2' void have ?case
+ by(blast intro!:WTrtFAssNT) }
+ moreover
+ { fix C' assume UClass: "U = Class C'" and "subclass": "P \<turnstile> C' \<preceq>\<^sup>* C"
+ have "P,E,h',sh' \<turnstile> e' : Class C'" using wt\<^sub>1' UClass by auto
+ moreover have "P \<turnstile> C' has F,NonStatic:TF in D"
+ by(rule has_field_mono[OF has "subclass"])
+ ultimately have ?case using wt\<^sub>2' sub void by(blast intro:WTrtFAss) }
+ ultimately have ?case using UsubC by(auto simp add:widen_Class) }
+ ultimately show ?case using wt by blast
+next
+ case (FAssRed2 e\<^sub>2 h l sh b e\<^sub>2' h' l' sh' b' v F D)
+ have red: "P \<turnstile> \<langle>e\<^sub>2,(h,l,sh),b\<rangle> \<rightarrow> \<langle>e\<^sub>2',(h',l',sh'),b'\<rangle>"
+ and IH: "\<And>E T. \<lbrakk>P,E \<turnstile> (h,l,sh) \<surd>; iconf sh e\<^sub>2; P,E,h,sh \<turnstile> e\<^sub>2 : T\<rbrakk>
+ \<Longrightarrow> \<exists>U. P,E,h',sh' \<turnstile> e\<^sub>2' : U \<and> P \<turnstile> U \<le> T"
+ and conf: "P,E \<turnstile> (h,l,sh) \<surd>" and iconf: "iconf sh (Val v\<bullet>F{D} := e\<^sub>2)"
+ and wt: "P,E,h,sh \<turnstile> Val v\<bullet>F{D}:=e\<^sub>2 : T" by fact+
+ have iconf2: "iconf sh e\<^sub>2" using iconf by simp
+ from wt have [simp]: "T = Void" by auto
+ from wt show ?case
+ proof (rule WTrt_elim_cases)
+ fix C TF T\<^sub>2
+ assume wt\<^sub>1: "P,E,h,sh \<turnstile> Val v : Class C"
+ and has: "P \<turnstile> C has F,NonStatic:TF in D"
+ and wt\<^sub>2: "P,E,h,sh \<turnstile> e\<^sub>2 : T\<^sub>2" and TsubTF: "P \<turnstile> T\<^sub>2 \<le> TF"
+ have wt\<^sub>1': "P,E,h',sh' \<turnstile> Val v : Class C"
+ using WTrt_hext_mono[OF wt\<^sub>1 red_hext_incr[OF red]] by auto
+ obtain T\<^sub>2' where wt\<^sub>2': "P,E,h',sh' \<turnstile> e\<^sub>2' : T\<^sub>2'" and T'subT: "P \<turnstile> T\<^sub>2' \<le> T\<^sub>2"
+ using IH[OF conf iconf2 wt\<^sub>2] by blast
+ have "P,E,h',sh' \<turnstile> Val v\<bullet>F{D}:=e\<^sub>2' : Void"
+ by(rule WTrtFAss[OF wt\<^sub>1' has wt\<^sub>2' widen_trans[OF T'subT TsubTF]])
+ thus ?case by auto
+ next
+ fix T\<^sub>2 assume null: "P,E,h,sh \<turnstile> Val v : NT" and wt\<^sub>2: "P,E,h,sh \<turnstile> e\<^sub>2 : T\<^sub>2"
+ from null have "v = Null" by simp
+ moreover
+ obtain T\<^sub>2' where "P,E,h',sh' \<turnstile> e\<^sub>2' : T\<^sub>2' \<and> P \<turnstile> T\<^sub>2' \<le> T\<^sub>2"
+ using IH[OF conf iconf2 wt\<^sub>2] by blast
+ ultimately show ?thesis by(fastforce intro:WTrtFAssNT)
+ qed
+next
+ case RedFAss thus ?case by(auto simp del:fun_upd_apply)
+next
+ case RedFAssNull thus ?case
+ by(fastforce intro: WTThrow[OF WTVal] elim!:typeof_NullPointer simp:sconf_def hconf_def)
+next
+ case RedFAssStatic then show ?case
+ by(fastforce intro: WTrtThrow[OF WTrtVal] elim!: typeof_IncompatibleClassChangeError
+ simp: sconf_def hconf_def)
+next
+ case (SFAssRed e h l sh b e' h' l' sh' b' C F D E T)
+ have IH: "\<And>E T. \<lbrakk>P,E \<turnstile> (h,l,sh) \<surd>; iconf sh e; P,E,h,sh \<turnstile> e : T\<rbrakk>
+ \<Longrightarrow> \<exists>U. P,E,h',sh' \<turnstile> e' : U \<and> P \<turnstile> U \<le> T"
+ and conf: "P,E \<turnstile> (h,l,sh) \<surd>" and iconf: "iconf sh (C\<bullet>\<^sub>sF{D} := e)"
+ and wt: "P,E,h,sh \<turnstile> C\<bullet>\<^sub>sF{D}:=e : T" by fact+
+ have iconf': "iconf sh e" using iconf by simp
+ from wt have [simp]: "T = Void" by auto
+ from wt show ?case
+ proof (rule WTrt_elim_cases)
+ fix TF T1
+ assume has: "P \<turnstile> C has F,Static:TF in D"
+ and wt1: "P,E,h,sh \<turnstile> e : T1" and TsubTF: "P \<turnstile> T1 \<le> TF"
+ obtain T' where wt1': "P,E,h',sh' \<turnstile> e' : T'" and T'subT: "P \<turnstile> T' \<le> T1"
+ using IH[OF conf iconf' wt1] by blast
+ have "P,E,h',sh' \<turnstile> C\<bullet>\<^sub>sF{D}:=e' : Void"
+ by(rule WTrtSFAss[OF wt1' has widen_trans[OF T'subT TsubTF]])
+ thus ?case by auto
+ qed
+next
+ case SFAssInitDoneRed then show ?case by (meson widen_refl)
+next
+ case (SFAssInitRed C F t D sh v h l E T)
+ have "is_class P D" using SFAssInitRed.hyps(1) by(rule has_field_is_class')
+ then have "P,E,h,sh \<turnstile> INIT D ([D],False) \<leftarrow> C\<bullet>\<^sub>sF{D} := Val v : T \<and> P \<turnstile> T \<le> T"
+ using SFAssInitRed WTrtInit[OF SFAssInitRed.prems(3)] by clarsimp
+ then show ?case by(rule exI)
+next
+ case RedSFAssNone then show ?case
+ by(fastforce intro: WTrtThrow[OF WTrtVal] elim!: typeof_NoSuchFieldError
+ simp: sconf_def hconf_def)
+next
+ case RedSFAssNonStatic then show ?case
+ by(fastforce intro: WTrtThrow[OF WTrtVal] elim!: typeof_IncompatibleClassChangeError
+ simp: sconf_def hconf_def)
+next
+ case (CallObj e h l sh b e' h' l' sh' b' M es)
+ have red: "P \<turnstile> \<langle>e,(h,l,sh),b\<rangle> \<rightarrow> \<langle>e',(h',l',sh'),b'\<rangle>"
+ and IH: "\<And>E T. \<lbrakk>P,E \<turnstile> (h,l,sh) \<surd>; iconf sh e; P,E,h,sh \<turnstile> e : T\<rbrakk>
+ \<Longrightarrow> \<exists>U. P,E,h',sh' \<turnstile> e' : U \<and> P \<turnstile> U \<le> T"
+ and conf: "P,E \<turnstile> (h,l,sh) \<surd>" and iconf: "iconf sh (e\<bullet>M(es))"
+ and wt: "P,E,h,sh \<turnstile> e\<bullet>M(es) : T" by fact+
+ have val: "val_of e = None" using red iconf val_no_step by auto
+ then have iconf': "iconf sh e" and nsub_RIs: "\<not>sub_RIs es" using iconf by simp+
+ \<comment> \<open>We distinguish if @{term e} has type @{term NT} or a Class type\<close>
+ \<comment> \<open>Remember ?case = @{term ?case}\<close>
+ { assume wt':"P,E,h,sh \<turnstile> e:NT"
+ hence "P,E,h',sh' \<turnstile> e' : NT" using IH[OF conf iconf'] by fastforce
+ moreover
+ fix Ts assume wtes: "P,E,h,sh \<turnstile> es [:] Ts"
+ have "P,E,h',sh' \<turnstile> es [:] Ts"
+ by(rule WTrts_hext_shext_mono[OF wtes red_hext_incr[OF red] red_shext_incr[OF red wt'] nsub_RIs])
+ ultimately have ?case by(blast intro!:WTrtCallNT) }
+ moreover
+ { fix C D Ts Us pns body
+ assume wte: "P,E,h,sh \<turnstile> e : Class C"
+ and "method": "P \<turnstile> C sees M,NonStatic:Ts\<rightarrow>T = (pns,body) in D"
+ and wtes: "P,E,h,sh \<turnstile> es [:] Us" and subs: "P \<turnstile> Us [\<le>] Ts"
+ obtain U where wte': "P,E,h',sh' \<turnstile> e' : U" and UsubC: "P \<turnstile> U \<le> Class C"
+ using IH[OF conf iconf' wte] by blast
+ \<comment> \<open>Is @{term U} the null type or a class type?\<close>
+ { assume "U = NT"
+ moreover have "P,E,h',sh' \<turnstile> es [:] Us"
+ by(rule WTrts_hext_shext_mono[OF wtes red_hext_incr[OF red] red_shext_incr[OF red wte] nsub_RIs])
+ ultimately have ?case using wte' by(blast intro!:WTrtCallNT) }
+ moreover
+ { fix C' assume UClass: "U = Class C'" and "subclass": "P \<turnstile> C' \<preceq>\<^sup>* C"
+ have "P,E,h',sh' \<turnstile> e' : Class C'" using wte' UClass by auto
+ moreover obtain Ts' T' pns' body' D'
+ where method': "P \<turnstile> C' sees M,NonStatic:Ts'\<rightarrow>T' = (pns',body') in D'"
+ and subs': "P \<turnstile> Ts [\<le>] Ts'" and sub': "P \<turnstile> T' \<le> T"
+ using Call_lemma[OF "method" "subclass" wf] by fast
+ moreover have "P,E,h',sh' \<turnstile> es [:] Us"
+ by(rule WTrts_hext_shext_mono[OF wtes red_hext_incr[OF red] red_shext_incr[OF red wte] nsub_RIs])
+ ultimately have ?case
+ using subs by(blast intro:WTrtCall rtrancl_trans widens_trans) }
+ ultimately have ?case using UsubC by(auto simp add:widen_Class) }
+ ultimately show ?case using wt by auto
+next
+ case (CallParams es h l sh b es' h' l' sh' b' v M)
+ have reds: "P \<turnstile> \<langle>es,(h,l,sh),b\<rangle> [\<rightarrow>] \<langle>es',(h',l',sh'),b'\<rangle>"
+ and IH: "\<And>E Ts. \<lbrakk>P,E \<turnstile> (h,l,sh) \<surd>; iconfs sh es; P,E,h,sh \<turnstile> es [:] Ts\<rbrakk>
+ \<Longrightarrow> \<exists>Us. P,E,h',sh' \<turnstile> es' [:] Us \<and> P \<turnstile> Us [\<le>] Ts"
+ and conf: "P,E \<turnstile> (h,l,sh) \<surd>" and iconf: "iconf sh (Val v\<bullet>M(es))"
+ and wt: "P,E,h,sh \<turnstile> Val v\<bullet>M(es) : T" by fact+
+ have iconfs: "iconfs sh es" using iconf by simp
+ from wt show ?case
+ proof (rule WTrt_elim_cases)
+ fix C D Ts Us pns body
+ assume wte: "P,E,h,sh \<turnstile> Val v : Class C"
+ and "P \<turnstile> C sees M,NonStatic:Ts\<rightarrow>T = (pns,body) in D"
+ and wtes: "P,E,h,sh \<turnstile> es [:] Us" and "P \<turnstile> Us [\<le>] Ts"
+ moreover have "P,E,h',sh' \<turnstile> Val v : Class C"
+ using WTrt_hext_mono[OF wte reds_hext_incr[OF reds]] by auto
+ moreover
+ obtain Us' where "P,E,h',sh' \<turnstile> es' [:] Us' \<and> P \<turnstile> Us' [\<le>] Us"
+ using IH[OF conf iconfs wtes] by blast
+ ultimately show ?thesis by(blast intro:WTrtCall widens_trans)
+ next
+ fix Us
+ assume null: "P,E,h,sh \<turnstile> Val v : NT" and wtes: "P,E,h,sh \<turnstile> es [:] Us"
+ from null have "v = Null" by simp
+ moreover
+ obtain Us' where "P,E,h',sh' \<turnstile> es' [:] Us' \<and> P \<turnstile> Us' [\<le>] Us"
+ using IH[OF conf iconfs wtes] by blast
+ ultimately show ?thesis by(fastforce intro:WTrtCallNT)
+ qed
+next
+ case (RedCall h a C fs M Ts T pns body D vs l sh b E T')
+ have hp: "h a = Some(C,fs)"
+ and "method": "P \<turnstile> C sees M,NonStatic: Ts\<rightarrow>T = (pns,body) in D"
+ and wt: "P,E,h,sh \<turnstile> addr a\<bullet>M(map Val vs) : T'" by fact+
+ obtain Ts' where wtes: "P,E,h,sh \<turnstile> map Val vs [:] Ts'"
+ and subs: "P \<turnstile> Ts' [\<le>] Ts" and T'isT: "T' = T"
+ using wt "method" hp by (auto dest:sees_method_fun)
+ from wtes subs have length_vs: "length vs = length Ts"
+ by(fastforce simp:list_all2_iff dest!:WTrts_same_length)
+ from sees_wf_mdecl[OF wf "method"] obtain T''
+ where wtabody: "P,[this#pns [\<mapsto>] Class D#Ts] \<turnstile> body :: T''"
+ and T''subT: "P \<turnstile> T'' \<le> T" and length_pns: "length pns = length Ts"
+ by(fastforce simp:wf_mdecl_def simp del:map_upds_twist)
+ from wtabody have "P,Map.empty(this#pns [\<mapsto>] Class D#Ts),h,sh \<turnstile> body : T''"
+ by(rule WT_implies_WTrt)
+ hence "P,E(this#pns [\<mapsto>] Class D#Ts),h,sh \<turnstile> body : T''"
+ by(rule WTrt_env_mono) simp
+ hence "P,E,h,sh \<turnstile> blocks(this#pns, Class D#Ts, Addr a#vs, body) : T''"
+ using wtes subs hp sees_method_decl_above[OF "method"] length_vs length_pns
+ by(fastforce simp add:wt_blocks rel_list_all2_Cons2)
+ with T''subT T'isT show ?case by blast
+next
+ case RedCallNull thus ?case
+ by(fastforce intro: WTThrow[OF WTVal] elim!:typeof_NullPointer simp: sconf_def hconf_def)
+next
+ case RedCallStatic then show ?case
+ by(fastforce intro: WTrtThrow[OF WTrtVal] elim!: typeof_IncompatibleClassChangeError
+ simp: sconf_def hconf_def)
+next
+ case (SCallParams es h l sh b es' h' l' sh' b' C M)
+ have IH: "\<And>E Ts. \<lbrakk>P,E \<turnstile> (h,l,sh) \<surd>; iconfs sh es; P,E,h,sh \<turnstile> es [:] Ts\<rbrakk>
+ \<Longrightarrow> \<exists>Us. P,E,h',sh' \<turnstile> es' [:] Us \<and> P \<turnstile> Us [\<le>] Ts"
+ and conf: "P,E \<turnstile> (h,l,sh) \<surd>" and iconf: "iconf sh (C\<bullet>\<^sub>sM(es))"
+ and wt: "P,E,h,sh \<turnstile> C\<bullet>\<^sub>sM(es) : T" by fact+
+ have iconfs: "iconfs sh es" using iconf by simp
+ from wt show ?case
+ proof (rule WTrt_elim_cases)
+ fix D Ts Us pns body sfs vs
+ assume method: "P \<turnstile> C sees M,Static:Ts\<rightarrow>T = (pns,body) in D"
+ and wtes: "P,E,h,sh \<turnstile> es [:] Us" and us: "P \<turnstile> Us [\<le>] Ts"
+ and clinit: "M = clinit \<longrightarrow> sh D = \<lfloor>(sfs,Processing)\<rfloor> \<and> es = map Val vs"
+ obtain Us' where es': "P,E,h',sh' \<turnstile> es' [:] Us'" and us': "P \<turnstile> Us' [\<le>] Us"
+ using IH[OF conf iconfs wtes] by blast
+ show ?thesis
+ proof(cases "M = clinit")
+ case True then show ?thesis using clinit SCallParams.hyps(1) by blast
+ next
+ case False
+ then show ?thesis using es' method us us' by(blast intro:WTrtSCall widens_trans)
+ qed
+ qed
+next
+ case (RedSCall C M Ts T pns body D vs h l sh E T')
+ have "method": "P \<turnstile> C sees M,Static: Ts\<rightarrow>T = (pns,body) in D"
+ and wt: "P,E,h,sh \<turnstile> C\<bullet>\<^sub>sM(map Val vs) : T'" by fact+
+ obtain Ts' where wtes: "P,E,h,sh \<turnstile> map Val vs [:] Ts'"
+ and subs: "P \<turnstile> Ts' [\<le>] Ts" and T'isT: "T' = T"
+ using wt "method" map_Val_eq by(auto dest:sees_method_fun)+
+ from wtes subs have length_vs: "length vs = length Ts"
+ by(fastforce simp:list_all2_iff dest!:WTrts_same_length)
+ from sees_wf_mdecl[OF wf "method"] obtain T''
+ where wtabody: "P,[pns [\<mapsto>] Ts] \<turnstile> body :: T''"
+ and T''subT: "P \<turnstile> T'' \<le> T" and length_pns: "length pns = length Ts"
+ by(fastforce simp:wf_mdecl_def simp del:map_upds_twist)
+ from wtabody have "P,Map.empty(pns [\<mapsto>] Ts),h,sh \<turnstile> body : T''"
+ by(rule WT_implies_WTrt)
+ hence "P,E(pns [\<mapsto>] Ts),h,sh \<turnstile> body : T''"
+ by(rule WTrt_env_mono) simp
+ hence "P,E,h,sh \<turnstile> blocks(pns, Ts, vs, body) : T''"
+ using wtes subs sees_method_decl_above[OF "method"] length_vs length_pns
+ by(fastforce simp add:wt_blocks rel_list_all2_Cons2)
+ with T''subT T'isT show ?case by blast
+next
+ case SCallInitDoneRed then show ?case by (meson widen_refl)
+next
+ case (SCallInitRed C F Ts t pns body D sh v h l E T)
+ have "is_class P D" using SCallInitRed.hyps(1) by(rule sees_method_is_class')
+ then have "P,E,h,sh \<turnstile> INIT D ([D],False) \<leftarrow> C\<bullet>\<^sub>sF(map Val v) : T \<and> P \<turnstile> T \<le> T"
+ using SCallInitRed WTrtInit[OF SCallInitRed.prems(3)] by clarsimp
+ then show ?case by(rule exI)
+next
+ case RedSCallNone then show ?case
+ by(fastforce intro: WTrtThrow[OF WTrtVal] elim!: typeof_NoSuchMethodError
+ simp: sconf_def hconf_def)
+next
+ case RedSCallNonStatic then show ?case
+ by(fastforce intro: WTrtThrow[OF WTrtVal] elim!: typeof_IncompatibleClassChangeError
+ simp: sconf_def hconf_def)
+next
+ case BlockRedNone thus ?case
+ by(auto simp del:fun_upd_apply)(fastforce simp:sconf_def lconf_def)
+next
+ case (BlockRedSome e h l V sh b e' h' l' sh' b' v T E Te)
+ have red: "P \<turnstile> \<langle>e,(h,l(V:=None),sh),b\<rangle> \<rightarrow> \<langle>e',(h',l',sh'),b'\<rangle>"
+ and IH: "\<And>E T. \<lbrakk>P,E \<turnstile> (h,l(V:=None),sh) \<surd>; iconf sh e; P,E,h,sh \<turnstile> e : T\<rbrakk>
+ \<Longrightarrow> \<exists>T'. P,E,h',sh' \<turnstile> e' : T' \<and> P \<turnstile> T' \<le> T"
+ and Some: "l' V = Some v" and conf: "P,E \<turnstile> (h,l,sh) \<surd>"
+ and iconf: "iconf sh {V:T; e}"
+ and wt: "P,E,h,sh \<turnstile> {V:T; e} : Te" by fact+
+ obtain Te' where IH': "P,E(V\<mapsto>T),h',sh' \<turnstile> e' : Te' \<and> P \<turnstile> Te' \<le> Te"
+ using IH conf iconf wt by(fastforce simp:sconf_def lconf_def)
+ have "P,h' \<turnstile> l' (:\<le>) E(V\<mapsto>T)" using conf wt
+ by(fastforce intro:red_preserves_lconf[OF red] simp:sconf_def lconf_def)
+ hence "P,h' \<turnstile> v :\<le> T" using Some by(fastforce simp:lconf_def)
+ with IH' show ?case
+ by(fastforce simp:sconf_def conf_def fun_upd_same simp del:fun_upd_apply)
+next
+ case (InitBlockRed e h l V v sh b e' h' l' sh' b' v' T E T')
+ have red: "P \<turnstile> \<langle>e, (h,l(V\<mapsto>v),sh),b\<rangle> \<rightarrow> \<langle>e',(h',l',sh'),b'\<rangle>"
+ and IH: "\<And>E T. \<lbrakk>P,E \<turnstile> (h,l(V\<mapsto>v),sh) \<surd>; iconf sh e; P,E,h,sh \<turnstile> e : T\<rbrakk>
+ \<Longrightarrow> \<exists>U. P,E,h',sh' \<turnstile> e' : U \<and> P \<turnstile> U \<le> T"
+ and v': "l' V = Some v'" and conf: "P,E \<turnstile> (h,l,sh) \<surd>"
+ and iconf: "iconf sh {V:T; V:=Val v;; e}"
+ and wt: "P,E,h,sh \<turnstile> {V:T := Val v; e} : T'" by fact+
+ from wt obtain T\<^sub>1 where wt\<^sub>1: "typeof\<^bsub>h\<^esub> v = Some T\<^sub>1"
+ and T1subT: "P \<turnstile> T\<^sub>1 \<le> T" and wt\<^sub>2: "P,E(V\<mapsto>T),h,sh \<turnstile> e : T'" by auto
+ have lconf\<^sub>2: "P,h \<turnstile> l(V\<mapsto>v) (:\<le>) E(V\<mapsto>T)" using conf wt\<^sub>1 T1subT
+ by(simp add:sconf_def lconf_upd2 conf_def)
+ have "\<exists>T\<^sub>1'. typeof\<^bsub>h'\<^esub> v' = Some T\<^sub>1' \<and> P \<turnstile> T\<^sub>1' \<le> T"
+ using v' red_preserves_lconf[OF red wt\<^sub>2 lconf\<^sub>2]
+ by(fastforce simp:lconf_def conf_def)
+ with IH conf iconf lconf\<^sub>2 wt\<^sub>2 show ?case by (fastforce simp add:sconf_def)
+next
+ case (SeqRed e h l sh b e' h' l' sh' b' e\<^sub>2)
+ then have val: "val_of e = None" by (simp add: val_no_step)
+ show ?case
+ proof(cases "lass_val_of e")
+ case None
+ then show ?thesis
+ using SeqRed val by(auto elim: WTrt_hext_shext_mono[OF _ red_hext_incr red_shext_incr])
+ next
+ case (Some a)
+ have "sh = sh'" using SeqRed lass_val_of_spec[OF Some] by auto
+ then show ?thesis using SeqRed val Some
+ by(auto intro: lass_val_of_iconf[OF Some] elim: WTrt_hext_mono[OF _ red_hext_incr])
+ qed
+next
+ case CondRed thus ?case
+ by auto (blast intro:WTrt_hext_shext_mono[OF _ red_hext_incr red_shext_incr])+
+next
+ case ThrowRed thus ?case
+ by(auto simp:is_refT_def)(blast dest:widen_Class[THEN iffD1])+
+next
+ case RedThrowNull thus ?case
+ by(fastforce intro: WTThrow[OF WTVal] elim!:typeof_NullPointer simp:sconf_def hconf_def)
+next
+ case TryRed thus ?case
+ by auto (blast intro:widen_trans WTrt_hext_shext_mono[OF _ red_hext_incr red_shext_incr])
+next
+ case RedTryFail thus ?case
+ by(fastforce intro: WTrtThrow[OF WTrtVal] simp:sconf_def hconf_def)
+next
+ case (ListRed1 e h l sh b e' h' l' sh' b' es)
+ then have val: "val_of e = None" by(simp add: val_no_step)
+ obtain U Us where Ts: "Ts = U # Us" using ListRed1 by auto
+ then have nsub_RI: "\<not> sub_RIs es" and wts: "P,E,h,sh \<turnstile> es [:] Us" and wt: "P,E,h,sh \<turnstile> e : U"
+ and IH: "\<And>E T. \<lbrakk>P,E \<turnstile> (h, l, sh) \<surd>; P,E,h,sh \<turnstile> e : T\<rbrakk> \<Longrightarrow> \<exists>T'. P,E,h',sh' \<turnstile> e' : T' \<and> P \<turnstile> T' \<le> T"
+ using ListRed1 val by auto
+ obtain T' where
+ "\<forall>E0 E1. (\<exists>T2. P,E1,h',sh' \<turnstile> e' : T2 \<and> P \<turnstile> T2 \<le> E0) = (P,E1,h',sh' \<turnstile> e' : T' E0 E1 \<and> P \<turnstile> T' E0 E1 \<le> E0)"
+ by moura
+ then have disj: "\<forall>E t. \<not> P,E \<turnstile> (h, l, sh) \<surd> \<or> \<not> P,E,h,sh \<turnstile> e : t \<or> P,E,h',sh' \<turnstile> e' : T' t E \<and> P \<turnstile> T' t E \<le> t"
+ using IH by presburger
+ have "P,E,h',sh' \<turnstile> es [:] Us"
+ using nsub_RI wts wt by (metis (no_types) ListRed1.hyps(1) WTrts_hext_shext_mono red_hext_incr red_shext_incr)
+ then have "\<exists>ts. (\<exists>t tsa. ts = t # tsa \<and> P,E,h',sh' \<turnstile> e' : t \<and> P,E,h',sh' \<turnstile> es [:] tsa) \<and> P \<turnstile> ts [\<le>] (U # Us)"
+ using disj wt ListRed1.prems(1) by blast
+ then show ?case using Ts by auto
+next
+ case ListRed2 thus ?case
+ by(fastforce dest: hext_typeof_mono[OF reds_hext_incr])
+next
+ case (InitNoneRed sh C C' Cs e h l b)
+ then have sh: "sh \<unlhd>\<^sub>s sh(C \<mapsto> (sblank P C, Prepared))" by(simp add: shext_def)
+ have wt: "P,E,h,sh(C \<mapsto> (sblank P C, Prepared)) \<turnstile> INIT C' (C # Cs,False) \<leftarrow> e : T"
+ using InitNoneRed WTrt_shext_mono[OF _ sh] by fastforce
+ then show ?case by(rule_tac x = T in exI) (simp add: fun_upd_def)
+next
+ case (RedInitDone sh C sfs C' Cs e h l b)
+ then have "P,E,h,sh \<turnstile> INIT C' (Cs,True) \<leftarrow> e : T" by auto (metis Nil_tl list.set_sel(2))
+ then show ?case by(rule_tac x = T in exI) simp
+next
+ case (RedInitProcessing sh C sfs C' Cs e h l b)
+ then have "P,E,h,sh \<turnstile> INIT C' (Cs,True) \<leftarrow> e : T" by auto (metis Nil_tl list.set_sel(2))+
+ then show ?case by(rule_tac x = T in exI) simp
+next
+ case RedInitError then show ?case
+ by(fastforce intro: WTrtThrow[OF WTrtVal] elim!: typeof_NoClassDefFoundError
+ simp: sconf_def hconf_def)
+next
+ case (InitObjectRed sh C sfs sh' C' Cs e h l b)
+ then have sh: "sh \<unlhd>\<^sub>s sh(Object \<mapsto> (sfs, Processing))" by(simp add: shext_def)
+ have "P,E,h,sh' \<turnstile> INIT C' (C # Cs,True) \<leftarrow> e : T"
+ using InitObjectRed WTrt_shext_mono[OF _ sh] by auto
+ then show ?case by(rule_tac x = T in exI) (simp add: fun_upd_def)
+next
+ case (InitNonObjectSuperRed sh C sfs D fs ms sh' C' Cs e h l b)
+ then have sh: "sh \<unlhd>\<^sub>s sh(C \<mapsto> (sfs, Processing))" by(simp add: shext_def)
+ then have cd: "is_class P D" using InitNonObjectSuperRed class_wf wf wf_cdecl_supD by blast
+ have sup': "supercls_lst P (C # Cs)" using InitNonObjectSuperRed.prems(3) by auto
+ then have sup: "supercls_lst P (D # C # Cs)"
+ using supercls_lst_app[of P C Cs D] subcls1I[OF InitNonObjectSuperRed.hyps(3,2)] by auto
+ have "distinct (C # Cs)" using InitNonObjectSuperRed.prems(3) by auto
+ then have dist: "distinct (D # C # Cs)"
+ using wf_supercls_distinct_app[OF wf InitNonObjectSuperRed.hyps(2-3) sup'] by simp
+ have "P,E,h,sh' \<turnstile> INIT C' (D # C # Cs,False) \<leftarrow> e : T"
+ using InitNonObjectSuperRed WTrt_shext_mono[OF _ sh] cd sup dist by auto
+ then show ?case by(rule_tac x = T in exI) simp
+next
+ case (RedInitRInit C' C Cs e' h l sh b E T)
+ then obtain a sfs where C: "class P C = \<lfloor>a\<rfloor>" and proc: "sh C = \<lfloor>(sfs, Processing)\<rfloor>"
+ using WTrtInit by(auto simp: is_class_def)
+ then have T': "P,E,h,sh \<turnstile> C\<bullet>\<^sub>sclinit([]) : Void" using wf_types_clinit[OF wf C] by simp
+ have "P,E,h,sh \<turnstile> RI (C,C\<bullet>\<^sub>sclinit([])) ; Cs \<leftarrow> e' : T"
+ using RedInitRInit by(auto intro: T')
+ then show ?case by(rule_tac x = T in exI) simp
+next
+ case (RInitRed e h l sh b e' h' l' sh' b' C Cs e\<^sub>0 E T)
+ then have "(\<And>E T. P,E \<turnstile> (h, l, sh) \<surd> \<Longrightarrow> P,E,h,sh \<turnstile> e : T \<Longrightarrow> \<exists>T'. P,E,h',sh' \<turnstile> e' : T' \<and> P \<turnstile> T' \<le> T)"
+ by auto
+ then have "\<exists>T'. P,E,h',sh' \<turnstile> e' : T'" using RInitRed by blast
+ then obtain T' where e': "P,E,h',sh' \<turnstile> e' : T'" by auto
+ have wt\<^sub>0: "P,E,h',sh' \<turnstile> e\<^sub>0 : T"
+ using RInitRed by simp (auto intro: WTrt_hext_shext_mono[OF _ red_hext_incr red_shext_incr])
+ have nip: "\<forall>C' \<in> set (C#Cs). not_init C' e' \<and> (\<exists>sfs. sh' C' = \<lfloor>(sfs, Processing)\<rfloor>)"
+ using RInitRed red_proc_pres[OF wf_prog_wwf_prog[OF wf]] by auto
+ have shC: "\<exists>sfs. sh' C = \<lfloor>(sfs, Processing)\<rfloor> \<or> sh' C = \<lfloor>(sfs, Error)\<rfloor> \<and> e' = THROW NoClassDefFoundError"
+ using RInitRed red_proc_pres[OF wf_prog_wwf_prog[OF wf] RInitRed.hyps(1)] by blast
+ have "P,E,h',sh' \<turnstile> RI (C,e') ; Cs \<leftarrow> e\<^sub>0 : T" using RInitRed e' wt\<^sub>0 nip shC by auto
+ then show ?case by(rule_tac x = T in exI) simp
+next
+ case (RedRInit sh C sfs i sh' C' Cs v e h l b)
+ then have sh: "sh \<unlhd>\<^sub>s sh(C \<mapsto> (sfs, Done))" by(auto simp: shext_def)
+ have wt: "P,E,h,sh(C \<mapsto> (sfs, Done)) \<turnstile> e : T"
+ using RedRInit WTrt_shext_mono[OF _ sh] by auto
+ have shC: "\<forall>C' \<in> set(tl Cs). \<exists>sfs. sh C' = \<lfloor>(sfs, Processing)\<rfloor>" using RedRInit by(cases Cs, auto)
+ have "P,E,h,sh' \<turnstile> INIT C' (Cs,True) \<leftarrow> e : T" using RedRInit wt shC by(cases Cs, auto)
+ then show ?case by(rule_tac x = T in exI) simp
+next
+ case (SCallThrowParams es vs e es' C M h l sh b)
+ then show ?case using map_Val_nthrow_neq[of _ vs e es'] by fastforce
+next
+ case (RInitInitThrow sh C sfs i sh' a D Cs e h l b)
+ then have sh: "sh \<unlhd>\<^sub>s sh(C \<mapsto> (sfs, Error))" by(auto simp: shext_def)
+ have wt: "P,E,h,sh(C \<mapsto> (sfs, Error)) \<turnstile> e : T"
+ using RInitInitThrow WTrt_shext_mono[OF _ sh] by clarsimp
+ then have "P,E,h,sh' \<turnstile> RI (D,Throw a) ; Cs \<leftarrow> e : T" using RInitInitThrow by auto
+ then show ?case by(rule_tac x = T in exI) simp
+qed fastforce+ (* esp all Throw propagation rules except RInitInit are dealt with here *)
+(*>*)
+
+
+corollary subject_reduction:
+ "\<lbrakk> wf_J_prog P; P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow> \<langle>e',s',b'\<rangle>; P,E \<turnstile> s \<surd>; iconf (shp s) e; P,E,hp s,shp s \<turnstile> e:T \<rbrakk>
+ \<Longrightarrow> \<exists>T'. P,E,hp s',shp s' \<turnstile> e':T' \<and> P \<turnstile> T' \<le> T"
+(*<*)by(cases s, cases s', fastforce dest:subject_reduction2)(*>*)
+
+corollary subjects_reduction:
+ "\<lbrakk> wf_J_prog P; P \<turnstile> \<langle>es,s,b\<rangle> [\<rightarrow>] \<langle>es',s',b'\<rangle>; P,E \<turnstile> s \<surd>; iconfs (shp s) es; P,E,hp s,shp s \<turnstile> es[:]Ts \<rbrakk>
+ \<Longrightarrow> \<exists>Ts'. P,E,hp s',shp s' \<turnstile> es'[:]Ts' \<and> P \<turnstile> Ts' [\<le>] Ts"
+(*<*)by(cases s, cases s', fastforce dest:subjects_reduction2)(*>*)
+
+
+subsection \<open> Lifting to @{text"\<rightarrow>*"} \<close>
+
+text\<open> Now all these preservation lemmas are first lifted to the transitive
+closure \dots \<close>
+
+lemma Red_preserves_sconf:
+assumes wf: "wf_J_prog P" and Red: "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle>"
+shows "\<And>T. \<lbrakk> P,E,hp s,shp s \<turnstile> e : T; iconf (shp s) e; P,E \<turnstile> s \<surd> \<rbrakk> \<Longrightarrow> P,E \<turnstile> s' \<surd>"
+(*<*)
+using Red
+proof (induct rule:converse_rtrancl_induct3)
+ case refl show ?case by fact
+next
+ case (step e s b e' s' b')
+ obtain h l sh h' l' sh' where s:"s = (h,l,sh)" and s':"s' = (h',l',sh')"
+ by(cases s, cases s')
+ then have "P \<turnstile> \<langle>e,(h,l,sh),b\<rangle> \<rightarrow> \<langle>e',(h',l',sh'),b'\<rangle>" using step.hyps(1) by simp
+ then have iconf': "iconf (shp s') e'" using red_preserves_iconf[OF wf_prog_wwf_prog[OF wf]]
+ step.prems(2) s s' by simp
+ thus ?case using step
+ by(blast intro:red_preserves_sconf dest: subject_reduction[OF wf])
+qed
+(*>*)
+
+lemma Red_preserves_iconf:
+assumes wf: "wwf_J_prog P" and Red: "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle>"
+shows "iconf (shp s) e \<Longrightarrow> iconf (shp s') e'"
+(*<*)
+using Red
+proof (induct rule:converse_rtrancl_induct3)
+ case refl show ?case by fact
+next
+ case (step e s b e' s' b')
+ thus ?case using wf step by(cases s, cases s', simp) (blast intro:red_preserves_iconf)
+qed
+(*>*)
+
+lemma Reds_preserves_iconf:
+assumes wf: "wwf_J_prog P" and Red: "P \<turnstile> \<langle>es,s,b\<rangle> [\<rightarrow>]* \<langle>es',s',b'\<rangle>"
+shows "iconfs (shp s) es \<Longrightarrow> iconfs (shp s') es'"
+(*<*)
+using Red
+proof (induct rule:converse_rtrancl_induct3)
+ case refl show ?case by fact
+next
+ case (step e s b e' s' b')
+ thus ?case using wf step by(cases s, cases s', simp) (blast intro:reds_preserves_iconf)
+qed
+(*>*)
+
+lemma Red_preserves_bconf:
+assumes wf: "wwf_J_prog P" and Red: "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle>"
+shows "iconf (shp s) e \<Longrightarrow> P,(shp s) \<turnstile>\<^sub>b (e,b) \<surd> \<Longrightarrow> P,(shp s') \<turnstile>\<^sub>b (e'::expr,b') \<surd>"
+(*<*)
+using Red
+proof (induct rule:converse_rtrancl_induct3)
+ case refl show ?case by fact
+next
+ case (step e s1 b e' s2 b')
+ then have "iconf (shp s2) e'" using step red_preserves_iconf[OF wf]
+ by(cases s1, cases s2) auto
+ thus ?case using step by(cases s1, cases s2, simp) (blast intro:red_preserves_bconf)
+qed
+(*>*)
+
+lemma Reds_preserves_bconf:
+assumes wf: "wwf_J_prog P" and Red: "P \<turnstile> \<langle>es,s,b\<rangle> [\<rightarrow>]* \<langle>es',s',b'\<rangle>"
+shows "iconfs (shp s) es \<Longrightarrow> P,(shp s) \<turnstile>\<^sub>b (es,b) \<surd> \<Longrightarrow> P,(shp s') \<turnstile>\<^sub>b (es'::expr list,b') \<surd>"
+(*<*)
+using Red
+proof (induct rule:converse_rtrancl_induct3)
+ case refl show ?case by fact
+next
+ case (step es s1 b es' s2 b')
+ then have "iconfs (shp s2) es'" using step reds_preserves_iconf[OF wf]
+ by(cases s1, cases s2) auto
+ thus ?case using step by(cases s1, cases s2, simp) (blast intro:reds_preserves_bconf)
+qed
+(*>*)
+
+lemma Red_preserves_defass:
+assumes wf: "wf_J_prog P" and reds: "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle>"
+shows "\<D> e \<lfloor>dom(lcl s)\<rfloor> \<Longrightarrow> \<D> e' \<lfloor>dom(lcl s')\<rfloor>"
+using reds
+proof (induct rule:converse_rtrancl_induct3)
+ case refl thus ?case .
+next
+ case (step e s b e' s' b') thus ?case
+ by(cases s,cases s')(auto dest:red_preserves_defass[OF wf])
+qed
+
+
+lemma Red_preserves_type:
+assumes wf: "wf_J_prog P" and Red: "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle>"
+shows "!!T. \<lbrakk> P,E \<turnstile> s\<surd>; iconf (shp s) e; P,E,hp s,shp s \<turnstile> e:T \<rbrakk>
+ \<Longrightarrow> \<exists>T'. P \<turnstile> T' \<le> T \<and> P,E,hp s',shp s' \<turnstile> e':T'"
+(*<*)
+using Red
+proof (induct rule:converse_rtrancl_induct3)
+ case refl thus ?case by blast
+next
+ case step thus ?case
+ by(blast intro:widen_trans red_preserves_sconf Red_preserves_iconf[OF wf_prog_wwf_prog[OF wf]]
+ dest:subject_reduction[OF wf])
+qed
+(*>*)
+
+
+subsection "The final polish"
+
+text\<open> The above preservation lemmas are now combined and packed nicely. \<close>
+
+definition wf_config :: "J_prog \<Rightarrow> env \<Rightarrow> state \<Rightarrow> expr \<Rightarrow> ty \<Rightarrow> bool" ("_,_,_ \<turnstile> _ : _ \<surd>" [51,0,0,0,0]50)
+where
+ "P,E,s \<turnstile> e:T \<surd> \<equiv> P,E \<turnstile> s \<surd> \<and> iconf (shp s) e \<and> P,E,hp s,shp s \<turnstile> e:T"
+
+theorem Subject_reduction: assumes wf: "wf_J_prog P"
+shows "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow> \<langle>e',s',b'\<rangle> \<Longrightarrow> P,E,s \<turnstile> e : T \<surd>
+ \<Longrightarrow> \<exists>T'. P,E,s' \<turnstile> e' : T' \<surd> \<and> P \<turnstile> T' \<le> T"
+(*<*)
+by(cases s, cases s')
+ (force simp: wf_config_def
+ elim:red_preserves_sconf red_preserves_iconf[OF wf_prog_wwf_prog[OF wf]]
+ dest:subject_reduction[OF wf])
+(*>*)
+
+
+theorem Subject_reductions:
+assumes wf: "wf_J_prog P" and reds: "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle>"
+shows "\<And>T. P,E,s \<turnstile> e:T \<surd> \<Longrightarrow> \<exists>T'. P,E,s' \<turnstile> e':T' \<surd> \<and> P \<turnstile> T' \<le> T"
+(*<*)
+using reds
+proof (induct rule:converse_rtrancl_induct3)
+ case refl thus ?case by blast
+next
+ case step thus ?case
+ by(blast dest:Subject_reduction[OF wf] intro:widen_trans)
+qed
+(*>*)
+
+
+corollary Progress: assumes wf: "wf_J_prog P"
+shows "\<lbrakk> P,E,s \<turnstile> e : T \<surd>; \<D> e \<lfloor>dom(lcl s)\<rfloor>; P,shp s \<turnstile>\<^sub>b (e,b) \<surd>; \<not> final e \<rbrakk>
+ \<Longrightarrow> \<exists>e' s' b'. P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow> \<langle>e',s',b'\<rangle>"
+(*<*)
+using progress[OF wf_prog_wwf_prog[OF wf]]
+by(cases b) (auto simp:wf_config_def sconf_def)
+(*>*)
+
+corollary TypeSafety:
+fixes s::state and e::expr
+assumes wf: "wf_J_prog P" and sconf: "P,E \<turnstile> s \<surd>" and wt: "P,E \<turnstile> e::T"
+ and \<D>: "\<D> e \<lfloor>dom(lcl s)\<rfloor>"
+ and iconf: "iconf (shp s) e" and bconf: "P,(shp s) \<turnstile>\<^sub>b (e,b) \<surd>"
+ and steps: "P \<turnstile> \<langle>e,s,b\<rangle> \<rightarrow>* \<langle>e',s',b'\<rangle>"
+ and nstep: "\<not>(\<exists>e'' s'' b''. P \<turnstile> \<langle>e',s',b'\<rangle> \<rightarrow> \<langle>e'',s'',b''\<rangle>)"
+shows "(\<exists>v. e' = Val v \<and> P,hp s' \<turnstile> v :\<le> T) \<or>
+ (\<exists>a. e' = Throw a \<and> a \<in> dom(hp s'))"
+(*<*)
+proof -
+ have wwf: "wwf_J_prog P" by(rule wf_prog_wwf_prog[OF wf])
+ have wfc: "P,E,s \<turnstile> e:T \<surd>" using WT_implies_WTrt[OF wt] sconf iconf
+ by(simp add:wf_config_def)
+ obtain T' where wfc': "P,E,s' \<turnstile> e' : T' \<surd>" and T': "P \<turnstile> T' \<le> T"
+ using Subject_reductions[OF wf steps wfc] by clarsimp
+ have \<D>': "\<D> e' \<lfloor>dom (lcl s')\<rfloor>"
+ by(rule Red_preserves_defass[OF wf steps \<D>])
+ have bconf': "P,(shp s') \<turnstile>\<^sub>b (e',b') \<surd>"
+ by(rule Red_preserves_bconf[OF wwf steps iconf bconf])
+ have fin': "final e'" using Progress[OF wf wfc' \<D>' bconf'] nstep by blast
+ then show ?thesis using wfc wfc' T'
+ by(fastforce simp:wf_config_def final_def conf_def)
+qed
+(*>*)
+
+
+end
diff --git a/thys/JinjaDCI/J/WWellForm.thy b/thys/JinjaDCI/J/WWellForm.thy
--- a/thys/JinjaDCI/J/WWellForm.thy
+++ b/thys/JinjaDCI/J/WWellForm.thy
@@ -1,40 +1,40 @@
-(* Title: JinjaDCI/J/WWellForm.thy
-
- Author: Tobias Nipkow, Susannah Mansky
- Copyright 2003 Technische Universitaet Muenchen, 2019-20 UIUC
-
- Based on the Jinja theory J/WWellForm.thy by Tobias Nipkow
-*)
-
-section \<open> Weak well-formedness of Jinja programs \<close>
-
-theory WWellForm imports "../Common/WellForm" Expr begin
-
-definition wwf_J_mdecl :: "J_prog \<Rightarrow> cname \<Rightarrow> J_mb mdecl \<Rightarrow> bool"
-where
- "wwf_J_mdecl P C \<equiv> \<lambda>(M,b,Ts,T,(pns,body)).
- length Ts = length pns \<and> distinct pns \<and> \<not>sub_RI body \<and>
- (case b of
- NonStatic \<Rightarrow> this \<notin> set pns \<and> fv body \<subseteq> {this} \<union> set pns
- | Static \<Rightarrow> fv body \<subseteq> set pns)"
-
-lemma wwf_J_mdecl_NonStatic[simp]:
- "wwf_J_mdecl P C (M,NonStatic,Ts,T,pns,body) =
- (length Ts = length pns \<and> distinct pns \<and> \<not>sub_RI body \<and> this \<notin> set pns \<and> fv body \<subseteq> {this} \<union> set pns)"
-(*<*)by(simp add:wwf_J_mdecl_def)(*>*)
-
-lemma wwf_J_mdecl_Static[simp]:
- "wwf_J_mdecl P C (M,Static,Ts,T,pns,body) =
- (length Ts = length pns \<and> distinct pns \<and> \<not>sub_RI body \<and> fv body \<subseteq> set pns)"
-(*<*)by(simp add:wwf_J_mdecl_def)(*>*)
-
-abbreviation
- wwf_J_prog :: "J_prog \<Rightarrow> bool" where
- "wwf_J_prog \<equiv> wf_prog wwf_J_mdecl"
-
-
-lemma sees_wwf_nsub_RI:
- "\<lbrakk> wwf_J_prog P; P \<turnstile> C sees M,b : Ts\<rightarrow>T = (pns, body) in D \<rbrakk> \<Longrightarrow> \<not>sub_RI body"
-(*<*)by(auto dest!: sees_wf_mdecl simp: wwf_J_mdecl_def wf_mdecl_def)(*>*)
-
-end
+(* Title: JinjaDCI/J/WWellForm.thy
+
+ Author: Tobias Nipkow, Susannah Mansky
+ Copyright 2003 Technische Universitaet Muenchen, 2019-20 UIUC
+
+ Based on the Jinja theory J/WWellForm.thy by Tobias Nipkow
+*)
+
+section \<open> Weak well-formedness of Jinja programs \<close>
+
+theory WWellForm imports "../Common/WellForm" Expr begin
+
+definition wwf_J_mdecl :: "J_prog \<Rightarrow> cname \<Rightarrow> J_mb mdecl \<Rightarrow> bool"
+where
+ "wwf_J_mdecl P C \<equiv> \<lambda>(M,b,Ts,T,(pns,body)).
+ length Ts = length pns \<and> distinct pns \<and> \<not>sub_RI body \<and>
+ (case b of
+ NonStatic \<Rightarrow> this \<notin> set pns \<and> fv body \<subseteq> {this} \<union> set pns
+ | Static \<Rightarrow> fv body \<subseteq> set pns)"
+
+lemma wwf_J_mdecl_NonStatic[simp]:
+ "wwf_J_mdecl P C (M,NonStatic,Ts,T,pns,body) =
+ (length Ts = length pns \<and> distinct pns \<and> \<not>sub_RI body \<and> this \<notin> set pns \<and> fv body \<subseteq> {this} \<union> set pns)"
+(*<*)by(simp add:wwf_J_mdecl_def)(*>*)
+
+lemma wwf_J_mdecl_Static[simp]:
+ "wwf_J_mdecl P C (M,Static,Ts,T,pns,body) =
+ (length Ts = length pns \<and> distinct pns \<and> \<not>sub_RI body \<and> fv body \<subseteq> set pns)"
+(*<*)by(simp add:wwf_J_mdecl_def)(*>*)
+
+abbreviation
+ wwf_J_prog :: "J_prog \<Rightarrow> bool" where
+ "wwf_J_prog \<equiv> wf_prog wwf_J_mdecl"
+
+
+lemma sees_wwf_nsub_RI:
+ "\<lbrakk> wwf_J_prog P; P \<turnstile> C sees M,b : Ts\<rightarrow>T = (pns, body) in D \<rbrakk> \<Longrightarrow> \<not>sub_RI body"
+(*<*)by(auto dest!: sees_wf_mdecl simp: wwf_J_mdecl_def wf_mdecl_def)(*>*)
+
+end
diff --git a/thys/JinjaDCI/J/WellType.thy b/thys/JinjaDCI/J/WellType.thy
--- a/thys/JinjaDCI/J/WellType.thy
+++ b/thys/JinjaDCI/J/WellType.thy
@@ -1,219 +1,219 @@
-(* Title: JinjaDCI/J/WellType.thy
-
- Author: Tobias Nipkow, Susannah Mansky
- Copyright 2003 Technische Universitaet Muenchen, 2019-20 UIUC
-
- Based on the Jinja theory J/WellType.thy by Tobias Nipkow
-*)
-
-section \<open> Well-typedness of Jinja expressions \<close>
-
-theory WellType
-imports "../Common/Objects" Expr
-begin
-
-type_synonym
- env = "vname \<rightharpoonup> ty"
-
-inductive
- WT :: "[J_prog,env, expr , ty ] \<Rightarrow> bool"
- ("_,_ \<turnstile> _ :: _" [51,51,51]50)
- and WTs :: "[J_prog,env, expr list, ty list] \<Rightarrow> bool"
- ("_,_ \<turnstile> _ [::] _" [51,51,51]50)
- for P :: J_prog
-where
-
- WTNew:
- "is_class P C \<Longrightarrow>
- P,E \<turnstile> new C :: Class C"
-
-| WTCast:
- "\<lbrakk> P,E \<turnstile> e :: Class D; is_class P C; P \<turnstile> C \<preceq>\<^sup>* D \<or> P \<turnstile> D \<preceq>\<^sup>* C \<rbrakk>
- \<Longrightarrow> P,E \<turnstile> Cast C e :: Class C"
-
-| WTVal:
- "typeof v = Some T \<Longrightarrow>
- P,E \<turnstile> Val v :: T"
-
-| WTVar:
- "E V = Some T \<Longrightarrow>
- P,E \<turnstile> Var V :: T"
-
-| WTBinOpEq:
- "\<lbrakk> P,E \<turnstile> e\<^sub>1 :: T\<^sub>1; P,E \<turnstile> e\<^sub>2 :: T\<^sub>2; P \<turnstile> T\<^sub>1 \<le> T\<^sub>2 \<or> P \<turnstile> T\<^sub>2 \<le> T\<^sub>1 \<rbrakk>
- \<Longrightarrow> P,E \<turnstile> e\<^sub>1 \<guillemotleft>Eq\<guillemotright> e\<^sub>2 :: Boolean"
-
-| WTBinOpAdd:
- "\<lbrakk> P,E \<turnstile> e\<^sub>1 :: Integer; P,E \<turnstile> e\<^sub>2 :: Integer \<rbrakk>
- \<Longrightarrow> P,E \<turnstile> e\<^sub>1 \<guillemotleft>Add\<guillemotright> e\<^sub>2 :: Integer"
-
-| WTLAss:
- "\<lbrakk> E V = Some T; P,E \<turnstile> e :: T'; P \<turnstile> T' \<le> T; V \<noteq> this \<rbrakk>
- \<Longrightarrow> P,E \<turnstile> V:=e :: Void"
-
-| WTFAcc:
- "\<lbrakk> P,E \<turnstile> e :: Class C; P \<turnstile> C sees F,NonStatic:T in D \<rbrakk>
- \<Longrightarrow> P,E \<turnstile> e\<bullet>F{D} :: T"
-
-| WTSFAcc:
- "\<lbrakk> P \<turnstile> C sees F,Static:T in D \<rbrakk>
- \<Longrightarrow> P,E \<turnstile> C\<bullet>\<^sub>sF{D} :: T"
-
-| WTFAss:
- "\<lbrakk> P,E \<turnstile> e\<^sub>1 :: Class C; P \<turnstile> C sees F,NonStatic:T in D; P,E \<turnstile> e\<^sub>2 :: T'; P \<turnstile> T' \<le> T \<rbrakk>
- \<Longrightarrow> P,E \<turnstile> e\<^sub>1\<bullet>F{D}:=e\<^sub>2 :: Void"
-
-| WTSFAss:
- "\<lbrakk> P \<turnstile> C sees F,Static:T in D; P,E \<turnstile> e\<^sub>2 :: T'; P \<turnstile> T' \<le> T \<rbrakk>
- \<Longrightarrow> P,E \<turnstile> C\<bullet>\<^sub>sF{D}:=e\<^sub>2 :: Void"
-
-| WTCall:
- "\<lbrakk> P,E \<turnstile> e :: Class C; P \<turnstile> C sees M,NonStatic:Ts \<rightarrow> T = (pns,body) in D;
- P,E \<turnstile> es [::] Ts'; P \<turnstile> Ts' [\<le>] Ts \<rbrakk>
- \<Longrightarrow> P,E \<turnstile> e\<bullet>M(es) :: T"
-
-| WTSCall:
- "\<lbrakk> P \<turnstile> C sees M,Static:Ts \<rightarrow> T = (pns,body) in D;
- P,E \<turnstile> es [::] Ts'; P \<turnstile> Ts' [\<le>] Ts; M \<noteq> clinit \<rbrakk>
- \<Longrightarrow> P,E \<turnstile> C\<bullet>\<^sub>sM(es) :: T"
-
-| WTBlock:
- "\<lbrakk> is_type P T; P,E(V \<mapsto> T) \<turnstile> e :: T' \<rbrakk>
- \<Longrightarrow> P,E \<turnstile> {V:T; e} :: T'"
-
-| WTSeq:
- "\<lbrakk> P,E \<turnstile> e\<^sub>1::T\<^sub>1; P,E \<turnstile> e\<^sub>2::T\<^sub>2 \<rbrakk>
- \<Longrightarrow> P,E \<turnstile> e\<^sub>1;;e\<^sub>2 :: T\<^sub>2"
-
-| WTCond:
- "\<lbrakk> P,E \<turnstile> e :: Boolean; P,E \<turnstile> e\<^sub>1::T\<^sub>1; P,E \<turnstile> e\<^sub>2::T\<^sub>2;
- P \<turnstile> T\<^sub>1 \<le> T\<^sub>2 \<or> P \<turnstile> T\<^sub>2 \<le> T\<^sub>1; P \<turnstile> T\<^sub>1 \<le> T\<^sub>2 \<longrightarrow> T = T\<^sub>2; P \<turnstile> T\<^sub>2 \<le> T\<^sub>1 \<longrightarrow> T = T\<^sub>1 \<rbrakk>
- \<Longrightarrow> P,E \<turnstile> if (e) e\<^sub>1 else e\<^sub>2 :: T"
-
-| WTWhile:
- "\<lbrakk> P,E \<turnstile> e :: Boolean; P,E \<turnstile> c::T \<rbrakk>
- \<Longrightarrow> P,E \<turnstile> while (e) c :: Void"
-
-| WTThrow:
- "P,E \<turnstile> e :: Class C \<Longrightarrow>
- P,E \<turnstile> throw e :: Void"
-
-| WTTry:
- "\<lbrakk> P,E \<turnstile> e\<^sub>1 :: T; P,E(V \<mapsto> Class C) \<turnstile> e\<^sub>2 :: T; is_class P C \<rbrakk>
- \<Longrightarrow> P,E \<turnstile> try e\<^sub>1 catch(C V) e\<^sub>2 :: T"
-
-\<comment> \<open>well-typed expression lists\<close>
-
-| WTNil:
- "P,E \<turnstile> [] [::] []"
-
-| WTCons:
- "\<lbrakk> P,E \<turnstile> e :: T; P,E \<turnstile> es [::] Ts \<rbrakk>
- \<Longrightarrow> P,E \<turnstile> e#es [::] T#Ts"
-
-(*<*)
-declare WT_WTs.intros[intro!] (* WTNil[iff] *)
-
-lemmas WT_WTs_induct = WT_WTs.induct [split_format (complete)]
- and WT_WTs_inducts = WT_WTs.inducts [split_format (complete)]
-(*>*)
-
-lemma init_nwt [simp]:"\<not>P,E \<turnstile> INIT C (Cs,b) \<leftarrow> e :: T"
- by(auto elim:WT.cases)
-lemma ri_nwt [simp]:"\<not>P,E \<turnstile> RI(C,e);Cs \<leftarrow> e' :: T"
- by(auto elim:WT.cases)
-
-lemma [iff]: "(P,E \<turnstile> [] [::] Ts) = (Ts = [])"
-(*<*)by (rule iffI) (auto elim: WTs.cases)(*>*)
-
-lemma [iff]: "(P,E \<turnstile> e#es [::] T#Ts) = (P,E \<turnstile> e :: T \<and> P,E \<turnstile> es [::] Ts)"
-(*<*)by (rule iffI) (auto elim: WTs.cases)(*>*)
-
-lemma [iff]: "(P,E \<turnstile> (e#es) [::] Ts) =
- (\<exists>U Us. Ts = U#Us \<and> P,E \<turnstile> e :: U \<and> P,E \<turnstile> es [::] Us)"
-(*<*)by (rule iffI) (auto elim: WTs.cases)(*>*)
-
-lemma [iff]: "\<And>Ts. (P,E \<turnstile> es\<^sub>1 @ es\<^sub>2 [::] Ts) =
- (\<exists>Ts\<^sub>1 Ts\<^sub>2. Ts = Ts\<^sub>1 @ Ts\<^sub>2 \<and> P,E \<turnstile> es\<^sub>1 [::] Ts\<^sub>1 \<and> P,E \<turnstile> es\<^sub>2[::]Ts\<^sub>2)"
-(*<*)
-proof(induct es\<^sub>1 type:list)
- case (Cons a list)
- let ?lhs = "(\<exists>U Us. Ts = U # Us \<and> P,E \<turnstile> a :: U \<and>
- (\<exists>Ts\<^sub>1 Ts\<^sub>2. Us = Ts\<^sub>1 @ Ts\<^sub>2 \<and> P,E \<turnstile> list [::] Ts\<^sub>1 \<and> P,E \<turnstile> es\<^sub>2 [::] Ts\<^sub>2))"
- let ?rhs = "(\<exists>Ts\<^sub>1 Ts\<^sub>2. Ts = Ts\<^sub>1 @ Ts\<^sub>2 \<and>
- (\<exists>U Us. Ts\<^sub>1 = U # Us \<and> P,E \<turnstile> a :: U \<and> P,E \<turnstile> list [::] Us) \<and> P,E \<turnstile> es\<^sub>2 [::] Ts\<^sub>2)"
- { assume ?lhs
- then have ?rhs by (auto intro: Cons_eq_appendI)
- }
- moreover {
- assume ?rhs
- then have ?lhs by fastforce
- }
- ultimately have "?lhs = ?rhs" by(rule iffI)
- then show ?case by (clarsimp simp: Cons)
-qed simp
-(*>*)
-
-lemma [iff]: "P,E \<turnstile> Val v :: T = (typeof v = Some T)"
-(*<*)proof(rule iffI) qed (auto elim: WT.cases)(*>*)
-
-lemma [iff]: "P,E \<turnstile> Var V :: T = (E V = Some T)"
-(*<*)proof(rule iffI) qed (auto elim: WT.cases)(*>*)
-
-lemma [iff]: "P,E \<turnstile> e\<^sub>1;;e\<^sub>2 :: T\<^sub>2 = (\<exists>T\<^sub>1. P,E \<turnstile> e\<^sub>1::T\<^sub>1 \<and> P,E \<turnstile> e\<^sub>2::T\<^sub>2)"
-(*<*)proof(rule iffI) qed (auto elim: WT.cases)(*>*)
-
-lemma [iff]: "(P,E \<turnstile> {V:T; e} :: T') = (is_type P T \<and> P,E(V\<mapsto>T) \<turnstile> e :: T')"
-(*<*)proof(rule iffI) qed (auto elim: WT.cases)(*>*)
-
-(*<*)
-inductive_cases WT_elim_cases[elim!]:
- "P,E \<turnstile> V :=e :: T"
- "P,E \<turnstile> if (e) e\<^sub>1 else e\<^sub>2 :: T"
- "P,E \<turnstile> while (e) c :: T"
- "P,E \<turnstile> throw e :: T"
- "P,E \<turnstile> try e\<^sub>1 catch(C V) e\<^sub>2 :: T"
- "P,E \<turnstile> Cast D e :: T"
- "P,E \<turnstile> a\<bullet>F{D} :: T"
- "P,E \<turnstile> C\<bullet>\<^sub>sF{D} :: T"
- "P,E \<turnstile> a\<bullet>F{D} := v :: T"
- "P,E \<turnstile> C\<bullet>\<^sub>sF{D} := v :: T"
- "P,E \<turnstile> e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2 :: T"
- "P,E \<turnstile> new C :: T"
- "P,E \<turnstile> e\<bullet>M(ps) :: T"
- "P,E \<turnstile> C\<bullet>\<^sub>sM(ps) :: T"
-(*>*)
-
-
-lemma wt_env_mono:
- "P,E \<turnstile> e :: T \<Longrightarrow> (\<And>E'. E \<subseteq>\<^sub>m E' \<Longrightarrow> P,E' \<turnstile> e :: T)" and
- "P,E \<turnstile> es [::] Ts \<Longrightarrow> (\<And>E'. E \<subseteq>\<^sub>m E' \<Longrightarrow> P,E' \<turnstile> es [::] Ts)"
-(*<*)
-proof(induct rule: WT_WTs_inducts)
- case WTVar then show ?case by(simp add: map_le_def dom_def)
-next
- case WTLAss then show ?case by(force simp:map_le_def)
-qed fastforce+
-(*>*)
-
-
-lemma WT_fv: "P,E \<turnstile> e :: T \<Longrightarrow> fv e \<subseteq> dom E"
-and "P,E \<turnstile> es [::] Ts \<Longrightarrow> fvs es \<subseteq> dom E"
-(*<*)
-proof(induct rule:WT_WTs.inducts)
- case WTVar then show ?case by fastforce
-next
- case WTLAss then show ?case by fastforce
-next
- case WTBlock then show ?case by fastforce
-next
- case WTTry then show ?case by fastforce
-qed simp_all
-(*>*)
-
-lemma WT_nsub_RI: "P,E \<turnstile> e :: T \<Longrightarrow> \<not>sub_RI e"
- and WTs_nsub_RIs: "P,E \<turnstile> es [::] Ts \<Longrightarrow> \<not>sub_RIs es"
-(*<*)proof(induct rule: WT_WTs.inducts) qed(simp_all)
-
-end
-(*>*)
+(* Title: JinjaDCI/J/WellType.thy
+
+ Author: Tobias Nipkow, Susannah Mansky
+ Copyright 2003 Technische Universitaet Muenchen, 2019-20 UIUC
+
+ Based on the Jinja theory J/WellType.thy by Tobias Nipkow
+*)
+
+section \<open> Well-typedness of Jinja expressions \<close>
+
+theory WellType
+imports "../Common/Objects" Expr
+begin
+
+type_synonym
+ env = "vname \<rightharpoonup> ty"
+
+inductive
+ WT :: "[J_prog,env, expr , ty ] \<Rightarrow> bool"
+ ("_,_ \<turnstile> _ :: _" [51,51,51]50)
+ and WTs :: "[J_prog,env, expr list, ty list] \<Rightarrow> bool"
+ ("_,_ \<turnstile> _ [::] _" [51,51,51]50)
+ for P :: J_prog
+where
+
+ WTNew:
+ "is_class P C \<Longrightarrow>
+ P,E \<turnstile> new C :: Class C"
+
+| WTCast:
+ "\<lbrakk> P,E \<turnstile> e :: Class D; is_class P C; P \<turnstile> C \<preceq>\<^sup>* D \<or> P \<turnstile> D \<preceq>\<^sup>* C \<rbrakk>
+ \<Longrightarrow> P,E \<turnstile> Cast C e :: Class C"
+
+| WTVal:
+ "typeof v = Some T \<Longrightarrow>
+ P,E \<turnstile> Val v :: T"
+
+| WTVar:
+ "E V = Some T \<Longrightarrow>
+ P,E \<turnstile> Var V :: T"
+
+| WTBinOpEq:
+ "\<lbrakk> P,E \<turnstile> e\<^sub>1 :: T\<^sub>1; P,E \<turnstile> e\<^sub>2 :: T\<^sub>2; P \<turnstile> T\<^sub>1 \<le> T\<^sub>2 \<or> P \<turnstile> T\<^sub>2 \<le> T\<^sub>1 \<rbrakk>
+ \<Longrightarrow> P,E \<turnstile> e\<^sub>1 \<guillemotleft>Eq\<guillemotright> e\<^sub>2 :: Boolean"
+
+| WTBinOpAdd:
+ "\<lbrakk> P,E \<turnstile> e\<^sub>1 :: Integer; P,E \<turnstile> e\<^sub>2 :: Integer \<rbrakk>
+ \<Longrightarrow> P,E \<turnstile> e\<^sub>1 \<guillemotleft>Add\<guillemotright> e\<^sub>2 :: Integer"
+
+| WTLAss:
+ "\<lbrakk> E V = Some T; P,E \<turnstile> e :: T'; P \<turnstile> T' \<le> T; V \<noteq> this \<rbrakk>
+ \<Longrightarrow> P,E \<turnstile> V:=e :: Void"
+
+| WTFAcc:
+ "\<lbrakk> P,E \<turnstile> e :: Class C; P \<turnstile> C sees F,NonStatic:T in D \<rbrakk>
+ \<Longrightarrow> P,E \<turnstile> e\<bullet>F{D} :: T"
+
+| WTSFAcc:
+ "\<lbrakk> P \<turnstile> C sees F,Static:T in D \<rbrakk>
+ \<Longrightarrow> P,E \<turnstile> C\<bullet>\<^sub>sF{D} :: T"
+
+| WTFAss:
+ "\<lbrakk> P,E \<turnstile> e\<^sub>1 :: Class C; P \<turnstile> C sees F,NonStatic:T in D; P,E \<turnstile> e\<^sub>2 :: T'; P \<turnstile> T' \<le> T \<rbrakk>
+ \<Longrightarrow> P,E \<turnstile> e\<^sub>1\<bullet>F{D}:=e\<^sub>2 :: Void"
+
+| WTSFAss:
+ "\<lbrakk> P \<turnstile> C sees F,Static:T in D; P,E \<turnstile> e\<^sub>2 :: T'; P \<turnstile> T' \<le> T \<rbrakk>
+ \<Longrightarrow> P,E \<turnstile> C\<bullet>\<^sub>sF{D}:=e\<^sub>2 :: Void"
+
+| WTCall:
+ "\<lbrakk> P,E \<turnstile> e :: Class C; P \<turnstile> C sees M,NonStatic:Ts \<rightarrow> T = (pns,body) in D;
+ P,E \<turnstile> es [::] Ts'; P \<turnstile> Ts' [\<le>] Ts \<rbrakk>
+ \<Longrightarrow> P,E \<turnstile> e\<bullet>M(es) :: T"
+
+| WTSCall:
+ "\<lbrakk> P \<turnstile> C sees M,Static:Ts \<rightarrow> T = (pns,body) in D;
+ P,E \<turnstile> es [::] Ts'; P \<turnstile> Ts' [\<le>] Ts; M \<noteq> clinit \<rbrakk>
+ \<Longrightarrow> P,E \<turnstile> C\<bullet>\<^sub>sM(es) :: T"
+
+| WTBlock:
+ "\<lbrakk> is_type P T; P,E(V \<mapsto> T) \<turnstile> e :: T' \<rbrakk>
+ \<Longrightarrow> P,E \<turnstile> {V:T; e} :: T'"
+
+| WTSeq:
+ "\<lbrakk> P,E \<turnstile> e\<^sub>1::T\<^sub>1; P,E \<turnstile> e\<^sub>2::T\<^sub>2 \<rbrakk>
+ \<Longrightarrow> P,E \<turnstile> e\<^sub>1;;e\<^sub>2 :: T\<^sub>2"
+
+| WTCond:
+ "\<lbrakk> P,E \<turnstile> e :: Boolean; P,E \<turnstile> e\<^sub>1::T\<^sub>1; P,E \<turnstile> e\<^sub>2::T\<^sub>2;
+ P \<turnstile> T\<^sub>1 \<le> T\<^sub>2 \<or> P \<turnstile> T\<^sub>2 \<le> T\<^sub>1; P \<turnstile> T\<^sub>1 \<le> T\<^sub>2 \<longrightarrow> T = T\<^sub>2; P \<turnstile> T\<^sub>2 \<le> T\<^sub>1 \<longrightarrow> T = T\<^sub>1 \<rbrakk>
+ \<Longrightarrow> P,E \<turnstile> if (e) e\<^sub>1 else e\<^sub>2 :: T"
+
+| WTWhile:
+ "\<lbrakk> P,E \<turnstile> e :: Boolean; P,E \<turnstile> c::T \<rbrakk>
+ \<Longrightarrow> P,E \<turnstile> while (e) c :: Void"
+
+| WTThrow:
+ "P,E \<turnstile> e :: Class C \<Longrightarrow>
+ P,E \<turnstile> throw e :: Void"
+
+| WTTry:
+ "\<lbrakk> P,E \<turnstile> e\<^sub>1 :: T; P,E(V \<mapsto> Class C) \<turnstile> e\<^sub>2 :: T; is_class P C \<rbrakk>
+ \<Longrightarrow> P,E \<turnstile> try e\<^sub>1 catch(C V) e\<^sub>2 :: T"
+
+\<comment> \<open>well-typed expression lists\<close>
+
+| WTNil:
+ "P,E \<turnstile> [] [::] []"
+
+| WTCons:
+ "\<lbrakk> P,E \<turnstile> e :: T; P,E \<turnstile> es [::] Ts \<rbrakk>
+ \<Longrightarrow> P,E \<turnstile> e#es [::] T#Ts"
+
+(*<*)
+declare WT_WTs.intros[intro!] (* WTNil[iff] *)
+
+lemmas WT_WTs_induct = WT_WTs.induct [split_format (complete)]
+ and WT_WTs_inducts = WT_WTs.inducts [split_format (complete)]
+(*>*)
+
+lemma init_nwt [simp]:"\<not>P,E \<turnstile> INIT C (Cs,b) \<leftarrow> e :: T"
+ by(auto elim:WT.cases)
+lemma ri_nwt [simp]:"\<not>P,E \<turnstile> RI(C,e);Cs \<leftarrow> e' :: T"
+ by(auto elim:WT.cases)
+
+lemma [iff]: "(P,E \<turnstile> [] [::] Ts) = (Ts = [])"
+(*<*)by (rule iffI) (auto elim: WTs.cases)(*>*)
+
+lemma [iff]: "(P,E \<turnstile> e#es [::] T#Ts) = (P,E \<turnstile> e :: T \<and> P,E \<turnstile> es [::] Ts)"
+(*<*)by (rule iffI) (auto elim: WTs.cases)(*>*)
+
+lemma [iff]: "(P,E \<turnstile> (e#es) [::] Ts) =
+ (\<exists>U Us. Ts = U#Us \<and> P,E \<turnstile> e :: U \<and> P,E \<turnstile> es [::] Us)"
+(*<*)by (rule iffI) (auto elim: WTs.cases)(*>*)
+
+lemma [iff]: "\<And>Ts. (P,E \<turnstile> es\<^sub>1 @ es\<^sub>2 [::] Ts) =
+ (\<exists>Ts\<^sub>1 Ts\<^sub>2. Ts = Ts\<^sub>1 @ Ts\<^sub>2 \<and> P,E \<turnstile> es\<^sub>1 [::] Ts\<^sub>1 \<and> P,E \<turnstile> es\<^sub>2[::]Ts\<^sub>2)"
+(*<*)
+proof(induct es\<^sub>1 type:list)
+ case (Cons a list)
+ let ?lhs = "(\<exists>U Us. Ts = U # Us \<and> P,E \<turnstile> a :: U \<and>
+ (\<exists>Ts\<^sub>1 Ts\<^sub>2. Us = Ts\<^sub>1 @ Ts\<^sub>2 \<and> P,E \<turnstile> list [::] Ts\<^sub>1 \<and> P,E \<turnstile> es\<^sub>2 [::] Ts\<^sub>2))"
+ let ?rhs = "(\<exists>Ts\<^sub>1 Ts\<^sub>2. Ts = Ts\<^sub>1 @ Ts\<^sub>2 \<and>
+ (\<exists>U Us. Ts\<^sub>1 = U # Us \<and> P,E \<turnstile> a :: U \<and> P,E \<turnstile> list [::] Us) \<and> P,E \<turnstile> es\<^sub>2 [::] Ts\<^sub>2)"
+ { assume ?lhs
+ then have ?rhs by (auto intro: Cons_eq_appendI)
+ }
+ moreover {
+ assume ?rhs
+ then have ?lhs by fastforce
+ }
+ ultimately have "?lhs = ?rhs" by(rule iffI)
+ then show ?case by (clarsimp simp: Cons)
+qed simp
+(*>*)
+
+lemma [iff]: "P,E \<turnstile> Val v :: T = (typeof v = Some T)"
+(*<*)proof(rule iffI) qed (auto elim: WT.cases)(*>*)
+
+lemma [iff]: "P,E \<turnstile> Var V :: T = (E V = Some T)"
+(*<*)proof(rule iffI) qed (auto elim: WT.cases)(*>*)
+
+lemma [iff]: "P,E \<turnstile> e\<^sub>1;;e\<^sub>2 :: T\<^sub>2 = (\<exists>T\<^sub>1. P,E \<turnstile> e\<^sub>1::T\<^sub>1 \<and> P,E \<turnstile> e\<^sub>2::T\<^sub>2)"
+(*<*)proof(rule iffI) qed (auto elim: WT.cases)(*>*)
+
+lemma [iff]: "(P,E \<turnstile> {V:T; e} :: T') = (is_type P T \<and> P,E(V\<mapsto>T) \<turnstile> e :: T')"
+(*<*)proof(rule iffI) qed (auto elim: WT.cases)(*>*)
+
+(*<*)
+inductive_cases WT_elim_cases[elim!]:
+ "P,E \<turnstile> V :=e :: T"
+ "P,E \<turnstile> if (e) e\<^sub>1 else e\<^sub>2 :: T"
+ "P,E \<turnstile> while (e) c :: T"
+ "P,E \<turnstile> throw e :: T"
+ "P,E \<turnstile> try e\<^sub>1 catch(C V) e\<^sub>2 :: T"
+ "P,E \<turnstile> Cast D e :: T"
+ "P,E \<turnstile> a\<bullet>F{D} :: T"
+ "P,E \<turnstile> C\<bullet>\<^sub>sF{D} :: T"
+ "P,E \<turnstile> a\<bullet>F{D} := v :: T"
+ "P,E \<turnstile> C\<bullet>\<^sub>sF{D} := v :: T"
+ "P,E \<turnstile> e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2 :: T"
+ "P,E \<turnstile> new C :: T"
+ "P,E \<turnstile> e\<bullet>M(ps) :: T"
+ "P,E \<turnstile> C\<bullet>\<^sub>sM(ps) :: T"
+(*>*)
+
+
+lemma wt_env_mono:
+ "P,E \<turnstile> e :: T \<Longrightarrow> (\<And>E'. E \<subseteq>\<^sub>m E' \<Longrightarrow> P,E' \<turnstile> e :: T)" and
+ "P,E \<turnstile> es [::] Ts \<Longrightarrow> (\<And>E'. E \<subseteq>\<^sub>m E' \<Longrightarrow> P,E' \<turnstile> es [::] Ts)"
+(*<*)
+proof(induct rule: WT_WTs_inducts)
+ case WTVar then show ?case by(simp add: map_le_def dom_def)
+next
+ case WTLAss then show ?case by(force simp:map_le_def)
+qed fastforce+
+(*>*)
+
+
+lemma WT_fv: "P,E \<turnstile> e :: T \<Longrightarrow> fv e \<subseteq> dom E"
+and "P,E \<turnstile> es [::] Ts \<Longrightarrow> fvs es \<subseteq> dom E"
+(*<*)
+proof(induct rule:WT_WTs.inducts)
+ case WTVar then show ?case by fastforce
+next
+ case WTLAss then show ?case by fastforce
+next
+ case WTBlock then show ?case by fastforce
+next
+ case WTTry then show ?case by fastforce
+qed simp_all
+(*>*)
+
+lemma WT_nsub_RI: "P,E \<turnstile> e :: T \<Longrightarrow> \<not>sub_RI e"
+ and WTs_nsub_RIs: "P,E \<turnstile> es [::] Ts \<Longrightarrow> \<not>sub_RIs es"
+(*<*)proof(induct rule: WT_WTs.inducts) qed(simp_all)
+
+end
+(*>*)
diff --git a/thys/JinjaDCI/J/WellTypeRT.thy b/thys/JinjaDCI/J/WellTypeRT.thy
--- a/thys/JinjaDCI/J/WellTypeRT.thy
+++ b/thys/JinjaDCI/J/WellTypeRT.thy
@@ -1,293 +1,293 @@
-(* Title: JinjaDCI/J/WellTypeRT.thy
-
- Author: Tobias Nipkow, Susannah Mansky
- Copyright 2003 Technische Universitaet Muenchen, 2019-20 UIUC
-
- Based on the Jinja theory J/WellTypeRT.thy by Tobias Nipkow
-*)
-
-section \<open> Runtime Well-typedness \<close>
-
-theory WellTypeRT
-imports WellType
-begin
-
-inductive
- WTrt :: "J_prog \<Rightarrow> heap \<Rightarrow> sheap \<Rightarrow> env \<Rightarrow> expr \<Rightarrow> ty \<Rightarrow> bool"
- and WTrts :: "J_prog \<Rightarrow> heap \<Rightarrow> sheap \<Rightarrow> env \<Rightarrow> expr list \<Rightarrow> ty list \<Rightarrow> bool"
- and WTrt2 :: "[J_prog,env,heap,sheap,expr,ty] \<Rightarrow> bool"
- ("_,_,_,_ \<turnstile> _ : _" [51,51,51,51]50)
- and WTrts2 :: "[J_prog,env,heap,sheap,expr list, ty list] \<Rightarrow> bool"
- ("_,_,_,_ \<turnstile> _ [:] _" [51,51,51,51]50)
- for P :: J_prog and h :: heap and sh :: sheap
-where
-
- "P,E,h,sh \<turnstile> e : T \<equiv> WTrt P h sh E e T"
-| "P,E,h,sh \<turnstile> es[:]Ts \<equiv> WTrts P h sh E es Ts"
-
-| WTrtNew:
- "is_class P C \<Longrightarrow>
- P,E,h,sh \<turnstile> new C : Class C"
-
-| WTrtCast:
- "\<lbrakk> P,E,h,sh \<turnstile> e : T; is_refT T; is_class P C \<rbrakk>
- \<Longrightarrow> P,E,h,sh \<turnstile> Cast C e : Class C"
-
-| WTrtVal:
- "typeof\<^bsub>h\<^esub> v = Some T \<Longrightarrow>
- P,E,h,sh \<turnstile> Val v : T"
-
-| WTrtVar:
- "E V = Some T \<Longrightarrow>
- P,E,h,sh \<turnstile> Var V : T"
-
-| WTrtBinOpEq:
- "\<lbrakk> P,E,h,sh \<turnstile> e\<^sub>1 : T\<^sub>1; P,E,h,sh \<turnstile> e\<^sub>2 : T\<^sub>2 \<rbrakk>
- \<Longrightarrow> P,E,h,sh \<turnstile> e\<^sub>1 \<guillemotleft>Eq\<guillemotright> e\<^sub>2 : Boolean"
-
-| WTrtBinOpAdd:
- "\<lbrakk> P,E,h,sh \<turnstile> e\<^sub>1 : Integer; P,E,h,sh \<turnstile> e\<^sub>2 : Integer \<rbrakk>
- \<Longrightarrow> P,E,h,sh \<turnstile> e\<^sub>1 \<guillemotleft>Add\<guillemotright> e\<^sub>2 : Integer"
-
-| WTrtLAss:
- "\<lbrakk> E V = Some T; P,E,h,sh \<turnstile> e : T'; P \<turnstile> T' \<le> T \<rbrakk>
- \<Longrightarrow> P,E,h,sh \<turnstile> V:=e : Void"
-
-| WTrtFAcc:
- "\<lbrakk> P,E,h,sh \<turnstile> e : Class C; P \<turnstile> C has F,NonStatic:T in D \<rbrakk> \<Longrightarrow>
- P,E,h,sh \<turnstile> e\<bullet>F{D} : T"
-
-| WTrtFAccNT:
- "P,E,h,sh \<turnstile> e : NT \<Longrightarrow>
- P,E,h,sh \<turnstile> e\<bullet>F{D} : T"
-
-| WTrtSFAcc:
- "\<lbrakk> P \<turnstile> C has F,Static:T in D \<rbrakk> \<Longrightarrow>
- P,E,h,sh \<turnstile> C\<bullet>\<^sub>sF{D} : T"
-
-| WTrtFAss:
- "\<lbrakk> P,E,h,sh \<turnstile> e\<^sub>1 : Class C; P \<turnstile> C has F,NonStatic:T in D; P,E,h,sh \<turnstile> e\<^sub>2 : T\<^sub>2; P \<turnstile> T\<^sub>2 \<le> T \<rbrakk>
- \<Longrightarrow> P,E,h,sh \<turnstile> e\<^sub>1\<bullet>F{D}:=e\<^sub>2 : Void"
-
-| WTrtFAssNT:
- "\<lbrakk> P,E,h,sh \<turnstile> e\<^sub>1:NT; P,E,h,sh \<turnstile> e\<^sub>2 : T\<^sub>2 \<rbrakk>
- \<Longrightarrow> P,E,h,sh \<turnstile> e\<^sub>1\<bullet>F{D}:=e\<^sub>2 : Void"
-
-| WTrtSFAss:
- "\<lbrakk> P,E,h,sh \<turnstile> e\<^sub>2 : T\<^sub>2; P \<turnstile> C has F,Static:T in D; P \<turnstile> T\<^sub>2 \<le> T \<rbrakk>
- \<Longrightarrow> P,E,h,sh \<turnstile> C\<bullet>\<^sub>sF{D}:=e\<^sub>2 : Void"
-
-| WTrtCall:
- "\<lbrakk> P,E,h,sh \<turnstile> e : Class C; P \<turnstile> C sees M,NonStatic:Ts \<rightarrow> T = (pns,body) in D;
- P,E,h,sh \<turnstile> es [:] Ts'; P \<turnstile> Ts' [\<le>] Ts \<rbrakk>
- \<Longrightarrow> P,E,h,sh \<turnstile> e\<bullet>M(es) : T"
-
-| WTrtCallNT:
- "\<lbrakk> P,E,h,sh \<turnstile> e : NT; P,E,h,sh \<turnstile> es [:] Ts \<rbrakk>
- \<Longrightarrow> P,E,h,sh \<turnstile> e\<bullet>M(es) : T"
-
-| WTrtSCall:
- "\<lbrakk> P \<turnstile> C sees M,Static:Ts \<rightarrow> T = (pns,body) in D;
- P,E,h,sh \<turnstile> es [:] Ts'; P \<turnstile> Ts' [\<le>] Ts;
- M = clinit \<longrightarrow> sh D = \<lfloor>(sfs,Processing)\<rfloor> \<and> es = map Val vs \<rbrakk>
- \<Longrightarrow> P,E,h,sh \<turnstile> C\<bullet>\<^sub>sM(es) : T"
-
-| WTrtBlock:
- "P,E(V\<mapsto>T),h,sh \<turnstile> e : T' \<Longrightarrow>
- P,E,h,sh \<turnstile> {V:T; e} : T'"
-
-| WTrtSeq:
- "\<lbrakk> P,E,h,sh \<turnstile> e\<^sub>1:T\<^sub>1; P,E,h,sh \<turnstile> e\<^sub>2:T\<^sub>2 \<rbrakk>
- \<Longrightarrow> P,E,h,sh \<turnstile> e\<^sub>1;;e\<^sub>2 : T\<^sub>2"
-
-| WTrtCond:
- "\<lbrakk> P,E,h,sh \<turnstile> e : Boolean; P,E,h,sh \<turnstile> e\<^sub>1:T\<^sub>1; P,E,h,sh \<turnstile> e\<^sub>2:T\<^sub>2;
- P \<turnstile> T\<^sub>1 \<le> T\<^sub>2 \<or> P \<turnstile> T\<^sub>2 \<le> T\<^sub>1; P \<turnstile> T\<^sub>1 \<le> T\<^sub>2 \<longrightarrow> T = T\<^sub>2; P \<turnstile> T\<^sub>2 \<le> T\<^sub>1 \<longrightarrow> T = T\<^sub>1 \<rbrakk>
- \<Longrightarrow> P,E,h,sh \<turnstile> if (e) e\<^sub>1 else e\<^sub>2 : T"
-
-| WTrtWhile:
- "\<lbrakk> P,E,h,sh \<turnstile> e : Boolean; P,E,h,sh \<turnstile> c:T \<rbrakk>
- \<Longrightarrow> P,E,h,sh \<turnstile> while(e) c : Void"
-
-| WTrtThrow:
- "\<lbrakk> P,E,h,sh \<turnstile> e : T\<^sub>r; is_refT T\<^sub>r \<rbrakk> \<Longrightarrow>
- P,E,h,sh \<turnstile> throw e : T"
-
-| WTrtTry:
- "\<lbrakk> P,E,h,sh \<turnstile> e\<^sub>1 : T\<^sub>1; P,E(V \<mapsto> Class C),h,sh \<turnstile> e\<^sub>2 : T\<^sub>2; P \<turnstile> T\<^sub>1 \<le> T\<^sub>2 \<rbrakk>
- \<Longrightarrow> P,E,h,sh \<turnstile> try e\<^sub>1 catch(C V) e\<^sub>2 : T\<^sub>2"
-
-| WTrtInit:
- "\<lbrakk> P,E,h,sh \<turnstile> e : T; \<forall>C' \<in> set (C#Cs). is_class P C'; \<not>sub_RI e;
- \<forall>C' \<in> set (tl Cs). \<exists>sfs. sh C' = \<lfloor>(sfs,Processing)\<rfloor>;
- b \<longrightarrow> (\<forall>C' \<in> set Cs. \<exists>sfs. sh C' = \<lfloor>(sfs,Processing)\<rfloor>);
- distinct Cs; supercls_lst P Cs \<rbrakk>
- \<Longrightarrow> P,E,h,sh \<turnstile> INIT C (Cs, b) \<leftarrow> e : T"
-
-| WTrtRI:
- "\<lbrakk> P,E,h,sh \<turnstile> e : T; P,E,h,sh \<turnstile> e' : T'; \<forall>C' \<in> set (C#Cs). is_class P C'; \<not>sub_RI e';
- \<forall>C' \<in> set (C#Cs). not_init C' e;
- \<forall>C' \<in> set Cs. \<exists>sfs. sh C' = \<lfloor>(sfs,Processing)\<rfloor>;
- \<exists>sfs. sh C = \<lfloor>(sfs, Processing)\<rfloor> \<or> (sh C = \<lfloor>(sfs, Error)\<rfloor> \<and> e = THROW NoClassDefFoundError);
- distinct (C#Cs); supercls_lst P (C#Cs) \<rbrakk>
- \<Longrightarrow> P,E,h,sh \<turnstile> RI(C, e);Cs \<leftarrow> e' : T'"
-
-\<comment> \<open>well-typed expression lists\<close>
-
-| WTrtNil:
- "P,E,h,sh \<turnstile> [] [:] []"
-
-| WTrtCons:
- "\<lbrakk> P,E,h,sh \<turnstile> e : T; P,E,h,sh \<turnstile> es [:] Ts \<rbrakk>
- \<Longrightarrow> P,E,h,sh \<turnstile> e#es [:] T#Ts"
-
-(*<*)
-declare WTrt_WTrts.intros[intro!] WTrtNil[iff]
-declare
- WTrtFAcc[rule del] WTrtFAccNT[rule del] WTrtSFAcc[rule del]
- WTrtFAss[rule del] WTrtFAssNT[rule del] WTrtSFAss[rule del]
- WTrtCall[rule del] WTrtCallNT[rule del] WTrtSCall[rule del]
-
-lemmas WTrt_induct = WTrt_WTrts.induct [split_format (complete)]
- and WTrt_inducts = WTrt_WTrts.inducts [split_format (complete)]
-(*>*)
-
-
-subsection\<open>Easy consequences\<close>
-
-lemma [iff]: "(P,E,h,sh \<turnstile> [] [:] Ts) = (Ts = [])"
-(*<*)by(rule iffI) (auto elim: WTrts.cases)(*>*)
-
-lemma [iff]: "(P,E,h,sh \<turnstile> e#es [:] T#Ts) = (P,E,h,sh \<turnstile> e : T \<and> P,E,h,sh \<turnstile> es [:] Ts)"
-(*<*)by(rule iffI) (auto elim: WTrts.cases)(*>*)
-
-lemma [iff]: "(P,E,h,sh \<turnstile> (e#es) [:] Ts) =
- (\<exists>U Us. Ts = U#Us \<and> P,E,h,sh \<turnstile> e : U \<and> P,E,h,sh \<turnstile> es [:] Us)"
-(*<*)by(rule iffI) (auto elim: WTrts.cases)(*>*)
-
-lemma [simp]: "\<forall>Ts. (P,E,h,sh \<turnstile> es\<^sub>1 @ es\<^sub>2 [:] Ts) =
- (\<exists>Ts\<^sub>1 Ts\<^sub>2. Ts = Ts\<^sub>1 @ Ts\<^sub>2 \<and> P,E,h,sh \<turnstile> es\<^sub>1 [:] Ts\<^sub>1 & P,E,h,sh \<turnstile> es\<^sub>2[:]Ts\<^sub>2)"
-(*<*)
-proof(induct es\<^sub>1)
- case (Cons a list)
- let ?lhs = "\<lambda>Ts. (\<exists>U Us. Ts = U # Us \<and> P,E,h,sh \<turnstile> a : U \<and>
- (\<exists>Ts\<^sub>1 Ts\<^sub>2. Us = Ts\<^sub>1 @ Ts\<^sub>2 \<and> P,E,h,sh \<turnstile> list [:] Ts\<^sub>1 \<and> P,E,h,sh \<turnstile> es\<^sub>2 [:] Ts\<^sub>2))"
- let ?rhs = "\<lambda>Ts. (\<exists>Ts\<^sub>1 Ts\<^sub>2. Ts = Ts\<^sub>1 @ Ts\<^sub>2 \<and>
- (\<exists>U Us. Ts\<^sub>1 = U # Us \<and> P,E,h,sh \<turnstile> a : U \<and> P,E,h,sh \<turnstile> list [:] Us) \<and> P,E,h,sh \<turnstile> es\<^sub>2 [:] Ts\<^sub>2)"
- { fix Ts assume "?lhs Ts"
- then have "?rhs Ts" by (auto intro: Cons_eq_appendI)
- }
- moreover {
- fix Ts assume "?rhs Ts"
- then have "?lhs Ts" by fastforce
- }
- ultimately have "\<And>Ts. ?lhs Ts = ?rhs Ts" by(rule iffI)
- then show ?case by (clarsimp simp: Cons)
-qed simp
-(*>*)
-
-lemma [iff]: "P,E,h,sh \<turnstile> Val v : T = (typeof\<^bsub>h\<^esub> v = Some T)"
-(*<*)proof(rule iffI) qed (auto elim: WTrt.cases)(*>*)
-
-lemma [iff]: "P,E,h,sh \<turnstile> Var v : T = (E v = Some T)"
-(*<*)proof(rule iffI) qed (auto elim: WTrt.cases)(*>*)
-
-lemma [iff]: "P,E,h,sh \<turnstile> e\<^sub>1;;e\<^sub>2 : T\<^sub>2 = (\<exists>T\<^sub>1. P,E,h,sh \<turnstile> e\<^sub>1:T\<^sub>1 \<and> P,E,h,sh \<turnstile> e\<^sub>2:T\<^sub>2)"
-(*<*)proof(rule iffI) qed (auto elim: WTrt.cases)(*>*)
-
-lemma [iff]: "P,E,h,sh \<turnstile> {V:T; e} : T' = (P,E(V\<mapsto>T),h,sh \<turnstile> e : T')"
-(*<*)proof(rule iffI) qed (auto elim: WTrt.cases)(*>*)
-
-(*<*)
-inductive_cases WTrt_elim_cases[elim!]:
- "P,E,h,sh \<turnstile> v :=e : T"
- "P,E,h,sh \<turnstile> if (e) e\<^sub>1 else e\<^sub>2 : T"
- "P,E,h,sh \<turnstile> while(e) c : T"
- "P,E,h,sh \<turnstile> throw e : T"
- "P,E,h,sh \<turnstile> try e\<^sub>1 catch(C V) e\<^sub>2 : T"
- "P,E,h,sh \<turnstile> Cast D e : T"
- "P,E,h,sh \<turnstile> e\<bullet>F{D} : T"
- "P,E,h,sh \<turnstile> C\<bullet>\<^sub>sF{D} : T"
- "P,E,h,sh \<turnstile> e\<bullet>F{D} := v : T"
- "P,E,h,sh \<turnstile> C\<bullet>\<^sub>sF{D} := v : T"
- "P,E,h,sh \<turnstile> e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2 : T"
- "P,E,h,sh \<turnstile> new C : T"
- "P,E,h,sh \<turnstile> e\<bullet>M{D}(es) : T"
- "P,E,h,sh \<turnstile> C\<bullet>\<^sub>sM{D}(es) : T"
- "P,E,h,sh \<turnstile> INIT C (Cs,b) \<leftarrow> e : T"
- "P,E,h,sh \<turnstile> RI(C,e);Cs \<leftarrow> e' : T"
-(*>*)
-
-subsection\<open>Some interesting lemmas\<close>
-
-lemma WTrts_Val[simp]:
- "\<And>Ts. (P,E,h,sh \<turnstile> map Val vs [:] Ts) = (map (typeof\<^bsub>h\<^esub>) vs = map Some Ts)"
-(*<*)
-proof(induct vs)
- case (Cons a vs)
- then show ?case by(case_tac Ts; simp)
-qed simp
-(*>*)
-
-
-lemma WTrts_same_length: "\<And>Ts. P,E,h,sh \<turnstile> es [:] Ts \<Longrightarrow> length es = length Ts"
-(*<*)by(induct es type:list)auto(*>*)
-
-
-lemma WTrt_env_mono:
- "P,E,h,sh \<turnstile> e : T \<Longrightarrow> (\<And>E'. E \<subseteq>\<^sub>m E' \<Longrightarrow> P,E',h,sh \<turnstile> e : T)" and
- "P,E,h,sh \<turnstile> es [:] Ts \<Longrightarrow> (\<And>E'. E \<subseteq>\<^sub>m E' \<Longrightarrow> P,E',h,sh \<turnstile> es [:] Ts)"
-(*<*)
-proof(induct rule: WTrt_inducts)
- case WTrtVar then show ?case by(simp add: map_le_def dom_def)
-next
- case WTrtLAss then show ?case by(force simp:map_le_def)
-qed(fastforce intro: WTrt_WTrts.intros)+
-(*>*)
-
-
-lemma WTrt_hext_mono: "P,E,h,sh \<turnstile> e : T \<Longrightarrow> h \<unlhd> h' \<Longrightarrow> P,E,h',sh \<turnstile> e : T"
-and WTrts_hext_mono: "P,E,h,sh \<turnstile> es [:] Ts \<Longrightarrow> h \<unlhd> h' \<Longrightarrow> P,E,h',sh \<turnstile> es [:] Ts"
-(*<*)
-proof(induct rule: WTrt_inducts)
- case WTrtVal then show ?case by(simp add: hext_typeof_mono)
-qed(fastforce intro: WTrt_WTrts.intros)+
-(*>*)
-
-lemma WTrt_shext_mono: "P,E,h,sh \<turnstile> e : T \<Longrightarrow> sh \<unlhd>\<^sub>s sh' \<Longrightarrow> \<not>sub_RI e \<Longrightarrow> P,E,h,sh' \<turnstile> e : T"
-and WTrts_shext_mono: "P,E,h,sh \<turnstile> es [:] Ts \<Longrightarrow> sh \<unlhd>\<^sub>s sh' \<Longrightarrow> \<not>sub_RIs es \<Longrightarrow> P,E,h,sh' \<turnstile> es [:] Ts"
-(*<*)
-by(induct rule: WTrt_inducts)
- (auto simp add: WTrt_WTrts.intros)
-(*>*)
-
-lemma WTrt_hext_shext_mono: "P,E,h,sh \<turnstile> e : T
- \<Longrightarrow> h \<unlhd> h' \<Longrightarrow> sh \<unlhd>\<^sub>s sh' \<Longrightarrow> \<not>sub_RI e \<Longrightarrow> P,E,h',sh' \<turnstile> e : T"
- by(auto intro: WTrt_hext_mono WTrt_shext_mono)
-
-lemma WTrts_hext_shext_mono: "P,E,h,sh \<turnstile> es [:] Ts
- \<Longrightarrow> h \<unlhd> h' \<Longrightarrow> sh \<unlhd>\<^sub>s sh' \<Longrightarrow> \<not>sub_RIs es \<Longrightarrow> P,E,h',sh' \<turnstile> es [:] Ts"
- by(auto intro: WTrts_hext_mono WTrts_shext_mono)
-
-
-lemma WT_implies_WTrt: "P,E \<turnstile> e :: T \<Longrightarrow> P,E,h,sh \<turnstile> e : T"
-and WTs_implies_WTrts: "P,E \<turnstile> es [::] Ts \<Longrightarrow> P,E,h,sh \<turnstile> es [:] Ts"
-(*<*)
-proof(induct rule: WT_WTs_inducts)
- case WTVal then show ?case by(fastforce dest: typeof_lit_typeof)
-next
- case WTFAcc
- then show ?case by(fastforce simp: WTrtFAcc has_visible_field)
-next
- case WTSFAcc
- then show ?case by(fastforce simp: WTrtSFAcc has_visible_field)
-next
- case WTFAss
- then show ?case by(fastforce simp: WTrtFAss dest: has_visible_field)
-next
- case WTSFAss
- then show ?case by(fastforce simp: WTrtSFAss dest: has_visible_field)
-qed(fastforce intro: WTrt_WTrts.intros)+
-(*>*)
-
-end
+(* Title: JinjaDCI/J/WellTypeRT.thy
+
+ Author: Tobias Nipkow, Susannah Mansky
+ Copyright 2003 Technische Universitaet Muenchen, 2019-20 UIUC
+
+ Based on the Jinja theory J/WellTypeRT.thy by Tobias Nipkow
+*)
+
+section \<open> Runtime Well-typedness \<close>
+
+theory WellTypeRT
+imports WellType
+begin
+
+inductive
+ WTrt :: "J_prog \<Rightarrow> heap \<Rightarrow> sheap \<Rightarrow> env \<Rightarrow> expr \<Rightarrow> ty \<Rightarrow> bool"
+ and WTrts :: "J_prog \<Rightarrow> heap \<Rightarrow> sheap \<Rightarrow> env \<Rightarrow> expr list \<Rightarrow> ty list \<Rightarrow> bool"
+ and WTrt2 :: "[J_prog,env,heap,sheap,expr,ty] \<Rightarrow> bool"
+ ("_,_,_,_ \<turnstile> _ : _" [51,51,51,51]50)
+ and WTrts2 :: "[J_prog,env,heap,sheap,expr list, ty list] \<Rightarrow> bool"
+ ("_,_,_,_ \<turnstile> _ [:] _" [51,51,51,51]50)
+ for P :: J_prog and h :: heap and sh :: sheap
+where
+
+ "P,E,h,sh \<turnstile> e : T \<equiv> WTrt P h sh E e T"
+| "P,E,h,sh \<turnstile> es[:]Ts \<equiv> WTrts P h sh E es Ts"
+
+| WTrtNew:
+ "is_class P C \<Longrightarrow>
+ P,E,h,sh \<turnstile> new C : Class C"
+
+| WTrtCast:
+ "\<lbrakk> P,E,h,sh \<turnstile> e : T; is_refT T; is_class P C \<rbrakk>
+ \<Longrightarrow> P,E,h,sh \<turnstile> Cast C e : Class C"
+
+| WTrtVal:
+ "typeof\<^bsub>h\<^esub> v = Some T \<Longrightarrow>
+ P,E,h,sh \<turnstile> Val v : T"
+
+| WTrtVar:
+ "E V = Some T \<Longrightarrow>
+ P,E,h,sh \<turnstile> Var V : T"
+
+| WTrtBinOpEq:
+ "\<lbrakk> P,E,h,sh \<turnstile> e\<^sub>1 : T\<^sub>1; P,E,h,sh \<turnstile> e\<^sub>2 : T\<^sub>2 \<rbrakk>
+ \<Longrightarrow> P,E,h,sh \<turnstile> e\<^sub>1 \<guillemotleft>Eq\<guillemotright> e\<^sub>2 : Boolean"
+
+| WTrtBinOpAdd:
+ "\<lbrakk> P,E,h,sh \<turnstile> e\<^sub>1 : Integer; P,E,h,sh \<turnstile> e\<^sub>2 : Integer \<rbrakk>
+ \<Longrightarrow> P,E,h,sh \<turnstile> e\<^sub>1 \<guillemotleft>Add\<guillemotright> e\<^sub>2 : Integer"
+
+| WTrtLAss:
+ "\<lbrakk> E V = Some T; P,E,h,sh \<turnstile> e : T'; P \<turnstile> T' \<le> T \<rbrakk>
+ \<Longrightarrow> P,E,h,sh \<turnstile> V:=e : Void"
+
+| WTrtFAcc:
+ "\<lbrakk> P,E,h,sh \<turnstile> e : Class C; P \<turnstile> C has F,NonStatic:T in D \<rbrakk> \<Longrightarrow>
+ P,E,h,sh \<turnstile> e\<bullet>F{D} : T"
+
+| WTrtFAccNT:
+ "P,E,h,sh \<turnstile> e : NT \<Longrightarrow>
+ P,E,h,sh \<turnstile> e\<bullet>F{D} : T"
+
+| WTrtSFAcc:
+ "\<lbrakk> P \<turnstile> C has F,Static:T in D \<rbrakk> \<Longrightarrow>
+ P,E,h,sh \<turnstile> C\<bullet>\<^sub>sF{D} : T"
+
+| WTrtFAss:
+ "\<lbrakk> P,E,h,sh \<turnstile> e\<^sub>1 : Class C; P \<turnstile> C has F,NonStatic:T in D; P,E,h,sh \<turnstile> e\<^sub>2 : T\<^sub>2; P \<turnstile> T\<^sub>2 \<le> T \<rbrakk>
+ \<Longrightarrow> P,E,h,sh \<turnstile> e\<^sub>1\<bullet>F{D}:=e\<^sub>2 : Void"
+
+| WTrtFAssNT:
+ "\<lbrakk> P,E,h,sh \<turnstile> e\<^sub>1:NT; P,E,h,sh \<turnstile> e\<^sub>2 : T\<^sub>2 \<rbrakk>
+ \<Longrightarrow> P,E,h,sh \<turnstile> e\<^sub>1\<bullet>F{D}:=e\<^sub>2 : Void"
+
+| WTrtSFAss:
+ "\<lbrakk> P,E,h,sh \<turnstile> e\<^sub>2 : T\<^sub>2; P \<turnstile> C has F,Static:T in D; P \<turnstile> T\<^sub>2 \<le> T \<rbrakk>
+ \<Longrightarrow> P,E,h,sh \<turnstile> C\<bullet>\<^sub>sF{D}:=e\<^sub>2 : Void"
+
+| WTrtCall:
+ "\<lbrakk> P,E,h,sh \<turnstile> e : Class C; P \<turnstile> C sees M,NonStatic:Ts \<rightarrow> T = (pns,body) in D;
+ P,E,h,sh \<turnstile> es [:] Ts'; P \<turnstile> Ts' [\<le>] Ts \<rbrakk>
+ \<Longrightarrow> P,E,h,sh \<turnstile> e\<bullet>M(es) : T"
+
+| WTrtCallNT:
+ "\<lbrakk> P,E,h,sh \<turnstile> e : NT; P,E,h,sh \<turnstile> es [:] Ts \<rbrakk>
+ \<Longrightarrow> P,E,h,sh \<turnstile> e\<bullet>M(es) : T"
+
+| WTrtSCall:
+ "\<lbrakk> P \<turnstile> C sees M,Static:Ts \<rightarrow> T = (pns,body) in D;
+ P,E,h,sh \<turnstile> es [:] Ts'; P \<turnstile> Ts' [\<le>] Ts;
+ M = clinit \<longrightarrow> sh D = \<lfloor>(sfs,Processing)\<rfloor> \<and> es = map Val vs \<rbrakk>
+ \<Longrightarrow> P,E,h,sh \<turnstile> C\<bullet>\<^sub>sM(es) : T"
+
+| WTrtBlock:
+ "P,E(V\<mapsto>T),h,sh \<turnstile> e : T' \<Longrightarrow>
+ P,E,h,sh \<turnstile> {V:T; e} : T'"
+
+| WTrtSeq:
+ "\<lbrakk> P,E,h,sh \<turnstile> e\<^sub>1:T\<^sub>1; P,E,h,sh \<turnstile> e\<^sub>2:T\<^sub>2 \<rbrakk>
+ \<Longrightarrow> P,E,h,sh \<turnstile> e\<^sub>1;;e\<^sub>2 : T\<^sub>2"
+
+| WTrtCond:
+ "\<lbrakk> P,E,h,sh \<turnstile> e : Boolean; P,E,h,sh \<turnstile> e\<^sub>1:T\<^sub>1; P,E,h,sh \<turnstile> e\<^sub>2:T\<^sub>2;
+ P \<turnstile> T\<^sub>1 \<le> T\<^sub>2 \<or> P \<turnstile> T\<^sub>2 \<le> T\<^sub>1; P \<turnstile> T\<^sub>1 \<le> T\<^sub>2 \<longrightarrow> T = T\<^sub>2; P \<turnstile> T\<^sub>2 \<le> T\<^sub>1 \<longrightarrow> T = T\<^sub>1 \<rbrakk>
+ \<Longrightarrow> P,E,h,sh \<turnstile> if (e) e\<^sub>1 else e\<^sub>2 : T"
+
+| WTrtWhile:
+ "\<lbrakk> P,E,h,sh \<turnstile> e : Boolean; P,E,h,sh \<turnstile> c:T \<rbrakk>
+ \<Longrightarrow> P,E,h,sh \<turnstile> while(e) c : Void"
+
+| WTrtThrow:
+ "\<lbrakk> P,E,h,sh \<turnstile> e : T\<^sub>r; is_refT T\<^sub>r \<rbrakk> \<Longrightarrow>
+ P,E,h,sh \<turnstile> throw e : T"
+
+| WTrtTry:
+ "\<lbrakk> P,E,h,sh \<turnstile> e\<^sub>1 : T\<^sub>1; P,E(V \<mapsto> Class C),h,sh \<turnstile> e\<^sub>2 : T\<^sub>2; P \<turnstile> T\<^sub>1 \<le> T\<^sub>2 \<rbrakk>
+ \<Longrightarrow> P,E,h,sh \<turnstile> try e\<^sub>1 catch(C V) e\<^sub>2 : T\<^sub>2"
+
+| WTrtInit:
+ "\<lbrakk> P,E,h,sh \<turnstile> e : T; \<forall>C' \<in> set (C#Cs). is_class P C'; \<not>sub_RI e;
+ \<forall>C' \<in> set (tl Cs). \<exists>sfs. sh C' = \<lfloor>(sfs,Processing)\<rfloor>;
+ b \<longrightarrow> (\<forall>C' \<in> set Cs. \<exists>sfs. sh C' = \<lfloor>(sfs,Processing)\<rfloor>);
+ distinct Cs; supercls_lst P Cs \<rbrakk>
+ \<Longrightarrow> P,E,h,sh \<turnstile> INIT C (Cs, b) \<leftarrow> e : T"
+
+| WTrtRI:
+ "\<lbrakk> P,E,h,sh \<turnstile> e : T; P,E,h,sh \<turnstile> e' : T'; \<forall>C' \<in> set (C#Cs). is_class P C'; \<not>sub_RI e';
+ \<forall>C' \<in> set (C#Cs). not_init C' e;
+ \<forall>C' \<in> set Cs. \<exists>sfs. sh C' = \<lfloor>(sfs,Processing)\<rfloor>;
+ \<exists>sfs. sh C = \<lfloor>(sfs, Processing)\<rfloor> \<or> (sh C = \<lfloor>(sfs, Error)\<rfloor> \<and> e = THROW NoClassDefFoundError);
+ distinct (C#Cs); supercls_lst P (C#Cs) \<rbrakk>
+ \<Longrightarrow> P,E,h,sh \<turnstile> RI(C, e);Cs \<leftarrow> e' : T'"
+
+\<comment> \<open>well-typed expression lists\<close>
+
+| WTrtNil:
+ "P,E,h,sh \<turnstile> [] [:] []"
+
+| WTrtCons:
+ "\<lbrakk> P,E,h,sh \<turnstile> e : T; P,E,h,sh \<turnstile> es [:] Ts \<rbrakk>
+ \<Longrightarrow> P,E,h,sh \<turnstile> e#es [:] T#Ts"
+
+(*<*)
+declare WTrt_WTrts.intros[intro!] WTrtNil[iff]
+declare
+ WTrtFAcc[rule del] WTrtFAccNT[rule del] WTrtSFAcc[rule del]
+ WTrtFAss[rule del] WTrtFAssNT[rule del] WTrtSFAss[rule del]
+ WTrtCall[rule del] WTrtCallNT[rule del] WTrtSCall[rule del]
+
+lemmas WTrt_induct = WTrt_WTrts.induct [split_format (complete)]
+ and WTrt_inducts = WTrt_WTrts.inducts [split_format (complete)]
+(*>*)
+
+
+subsection\<open>Easy consequences\<close>
+
+lemma [iff]: "(P,E,h,sh \<turnstile> [] [:] Ts) = (Ts = [])"
+(*<*)by(rule iffI) (auto elim: WTrts.cases)(*>*)
+
+lemma [iff]: "(P,E,h,sh \<turnstile> e#es [:] T#Ts) = (P,E,h,sh \<turnstile> e : T \<and> P,E,h,sh \<turnstile> es [:] Ts)"
+(*<*)by(rule iffI) (auto elim: WTrts.cases)(*>*)
+
+lemma [iff]: "(P,E,h,sh \<turnstile> (e#es) [:] Ts) =
+ (\<exists>U Us. Ts = U#Us \<and> P,E,h,sh \<turnstile> e : U \<and> P,E,h,sh \<turnstile> es [:] Us)"
+(*<*)by(rule iffI) (auto elim: WTrts.cases)(*>*)
+
+lemma [simp]: "\<forall>Ts. (P,E,h,sh \<turnstile> es\<^sub>1 @ es\<^sub>2 [:] Ts) =
+ (\<exists>Ts\<^sub>1 Ts\<^sub>2. Ts = Ts\<^sub>1 @ Ts\<^sub>2 \<and> P,E,h,sh \<turnstile> es\<^sub>1 [:] Ts\<^sub>1 & P,E,h,sh \<turnstile> es\<^sub>2[:]Ts\<^sub>2)"
+(*<*)
+proof(induct es\<^sub>1)
+ case (Cons a list)
+ let ?lhs = "\<lambda>Ts. (\<exists>U Us. Ts = U # Us \<and> P,E,h,sh \<turnstile> a : U \<and>
+ (\<exists>Ts\<^sub>1 Ts\<^sub>2. Us = Ts\<^sub>1 @ Ts\<^sub>2 \<and> P,E,h,sh \<turnstile> list [:] Ts\<^sub>1 \<and> P,E,h,sh \<turnstile> es\<^sub>2 [:] Ts\<^sub>2))"
+ let ?rhs = "\<lambda>Ts. (\<exists>Ts\<^sub>1 Ts\<^sub>2. Ts = Ts\<^sub>1 @ Ts\<^sub>2 \<and>
+ (\<exists>U Us. Ts\<^sub>1 = U # Us \<and> P,E,h,sh \<turnstile> a : U \<and> P,E,h,sh \<turnstile> list [:] Us) \<and> P,E,h,sh \<turnstile> es\<^sub>2 [:] Ts\<^sub>2)"
+ { fix Ts assume "?lhs Ts"
+ then have "?rhs Ts" by (auto intro: Cons_eq_appendI)
+ }
+ moreover {
+ fix Ts assume "?rhs Ts"
+ then have "?lhs Ts" by fastforce
+ }
+ ultimately have "\<And>Ts. ?lhs Ts = ?rhs Ts" by(rule iffI)
+ then show ?case by (clarsimp simp: Cons)
+qed simp
+(*>*)
+
+lemma [iff]: "P,E,h,sh \<turnstile> Val v : T = (typeof\<^bsub>h\<^esub> v = Some T)"
+(*<*)proof(rule iffI) qed (auto elim: WTrt.cases)(*>*)
+
+lemma [iff]: "P,E,h,sh \<turnstile> Var v : T = (E v = Some T)"
+(*<*)proof(rule iffI) qed (auto elim: WTrt.cases)(*>*)
+
+lemma [iff]: "P,E,h,sh \<turnstile> e\<^sub>1;;e\<^sub>2 : T\<^sub>2 = (\<exists>T\<^sub>1. P,E,h,sh \<turnstile> e\<^sub>1:T\<^sub>1 \<and> P,E,h,sh \<turnstile> e\<^sub>2:T\<^sub>2)"
+(*<*)proof(rule iffI) qed (auto elim: WTrt.cases)(*>*)
+
+lemma [iff]: "P,E,h,sh \<turnstile> {V:T; e} : T' = (P,E(V\<mapsto>T),h,sh \<turnstile> e : T')"
+(*<*)proof(rule iffI) qed (auto elim: WTrt.cases)(*>*)
+
+(*<*)
+inductive_cases WTrt_elim_cases[elim!]:
+ "P,E,h,sh \<turnstile> v :=e : T"
+ "P,E,h,sh \<turnstile> if (e) e\<^sub>1 else e\<^sub>2 : T"
+ "P,E,h,sh \<turnstile> while(e) c : T"
+ "P,E,h,sh \<turnstile> throw e : T"
+ "P,E,h,sh \<turnstile> try e\<^sub>1 catch(C V) e\<^sub>2 : T"
+ "P,E,h,sh \<turnstile> Cast D e : T"
+ "P,E,h,sh \<turnstile> e\<bullet>F{D} : T"
+ "P,E,h,sh \<turnstile> C\<bullet>\<^sub>sF{D} : T"
+ "P,E,h,sh \<turnstile> e\<bullet>F{D} := v : T"
+ "P,E,h,sh \<turnstile> C\<bullet>\<^sub>sF{D} := v : T"
+ "P,E,h,sh \<turnstile> e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2 : T"
+ "P,E,h,sh \<turnstile> new C : T"
+ "P,E,h,sh \<turnstile> e\<bullet>M{D}(es) : T"
+ "P,E,h,sh \<turnstile> C\<bullet>\<^sub>sM{D}(es) : T"
+ "P,E,h,sh \<turnstile> INIT C (Cs,b) \<leftarrow> e : T"
+ "P,E,h,sh \<turnstile> RI(C,e);Cs \<leftarrow> e' : T"
+(*>*)
+
+subsection\<open>Some interesting lemmas\<close>
+
+lemma WTrts_Val[simp]:
+ "\<And>Ts. (P,E,h,sh \<turnstile> map Val vs [:] Ts) = (map (typeof\<^bsub>h\<^esub>) vs = map Some Ts)"
+(*<*)
+proof(induct vs)
+ case (Cons a vs)
+ then show ?case by(case_tac Ts; simp)
+qed simp
+(*>*)
+
+
+lemma WTrts_same_length: "\<And>Ts. P,E,h,sh \<turnstile> es [:] Ts \<Longrightarrow> length es = length Ts"
+(*<*)by(induct es type:list)auto(*>*)
+
+
+lemma WTrt_env_mono:
+ "P,E,h,sh \<turnstile> e : T \<Longrightarrow> (\<And>E'. E \<subseteq>\<^sub>m E' \<Longrightarrow> P,E',h,sh \<turnstile> e : T)" and
+ "P,E,h,sh \<turnstile> es [:] Ts \<Longrightarrow> (\<And>E'. E \<subseteq>\<^sub>m E' \<Longrightarrow> P,E',h,sh \<turnstile> es [:] Ts)"
+(*<*)
+proof(induct rule: WTrt_inducts)
+ case WTrtVar then show ?case by(simp add: map_le_def dom_def)
+next
+ case WTrtLAss then show ?case by(force simp:map_le_def)
+qed(fastforce intro: WTrt_WTrts.intros)+
+(*>*)
+
+
+lemma WTrt_hext_mono: "P,E,h,sh \<turnstile> e : T \<Longrightarrow> h \<unlhd> h' \<Longrightarrow> P,E,h',sh \<turnstile> e : T"
+and WTrts_hext_mono: "P,E,h,sh \<turnstile> es [:] Ts \<Longrightarrow> h \<unlhd> h' \<Longrightarrow> P,E,h',sh \<turnstile> es [:] Ts"
+(*<*)
+proof(induct rule: WTrt_inducts)
+ case WTrtVal then show ?case by(simp add: hext_typeof_mono)
+qed(fastforce intro: WTrt_WTrts.intros)+
+(*>*)
+
+lemma WTrt_shext_mono: "P,E,h,sh \<turnstile> e : T \<Longrightarrow> sh \<unlhd>\<^sub>s sh' \<Longrightarrow> \<not>sub_RI e \<Longrightarrow> P,E,h,sh' \<turnstile> e : T"
+and WTrts_shext_mono: "P,E,h,sh \<turnstile> es [:] Ts \<Longrightarrow> sh \<unlhd>\<^sub>s sh' \<Longrightarrow> \<not>sub_RIs es \<Longrightarrow> P,E,h,sh' \<turnstile> es [:] Ts"
+(*<*)
+by(induct rule: WTrt_inducts)
+ (auto simp add: WTrt_WTrts.intros)
+(*>*)
+
+lemma WTrt_hext_shext_mono: "P,E,h,sh \<turnstile> e : T
+ \<Longrightarrow> h \<unlhd> h' \<Longrightarrow> sh \<unlhd>\<^sub>s sh' \<Longrightarrow> \<not>sub_RI e \<Longrightarrow> P,E,h',sh' \<turnstile> e : T"
+ by(auto intro: WTrt_hext_mono WTrt_shext_mono)
+
+lemma WTrts_hext_shext_mono: "P,E,h,sh \<turnstile> es [:] Ts
+ \<Longrightarrow> h \<unlhd> h' \<Longrightarrow> sh \<unlhd>\<^sub>s sh' \<Longrightarrow> \<not>sub_RIs es \<Longrightarrow> P,E,h',sh' \<turnstile> es [:] Ts"
+ by(auto intro: WTrts_hext_mono WTrts_shext_mono)
+
+
+lemma WT_implies_WTrt: "P,E \<turnstile> e :: T \<Longrightarrow> P,E,h,sh \<turnstile> e : T"
+and WTs_implies_WTrts: "P,E \<turnstile> es [::] Ts \<Longrightarrow> P,E,h,sh \<turnstile> es [:] Ts"
+(*<*)
+proof(induct rule: WT_WTs_inducts)
+ case WTVal then show ?case by(fastforce dest: typeof_lit_typeof)
+next
+ case WTFAcc
+ then show ?case by(fastforce simp: WTrtFAcc has_visible_field)
+next
+ case WTSFAcc
+ then show ?case by(fastforce simp: WTrtSFAcc has_visible_field)
+next
+ case WTFAss
+ then show ?case by(fastforce simp: WTrtFAss dest: has_visible_field)
+next
+ case WTSFAss
+ then show ?case by(fastforce simp: WTrtSFAss dest: has_visible_field)
+qed(fastforce intro: WTrt_WTrts.intros)+
+(*>*)
+
+end
diff --git a/thys/JinjaDCI/JVM/JVMDefensive.thy b/thys/JinjaDCI/JVM/JVMDefensive.thy
--- a/thys/JinjaDCI/JVM/JVMDefensive.thy
+++ b/thys/JinjaDCI/JVM/JVMDefensive.thy
@@ -1,220 +1,220 @@
-(* Title: JinjaDCI/JVM/JVMDefensive.thy
- Author: Gerwin Klein, Susannah Mansky
- Copyright GPL
-
- Based on the Jinja theory JVM/JVMDefensive.thy by Gerwin Klein
-*)
-
-section \<open> A Defensive JVM \<close>
-
-theory JVMDefensive
-imports JVMExec "../Common/Conform"
-begin
-
-text \<open>
- Extend the state space by one element indicating a type error (or
- other abnormal termination) \<close>
-datatype 'a type_error = TypeError | Normal 'a
-
-fun is_Addr :: "val \<Rightarrow> bool" where
- "is_Addr (Addr a) \<longleftrightarrow> True"
-| "is_Addr v \<longleftrightarrow> False"
-
-fun is_Intg :: "val \<Rightarrow> bool" where
- "is_Intg (Intg i) \<longleftrightarrow> True"
-| "is_Intg v \<longleftrightarrow> False"
-
-fun is_Bool :: "val \<Rightarrow> bool" where
- "is_Bool (Bool b) \<longleftrightarrow> True"
-| "is_Bool v \<longleftrightarrow> False"
-
-definition is_Ref :: "val \<Rightarrow> bool" where
- "is_Ref v \<longleftrightarrow> v = Null \<or> is_Addr v"
-
-primrec check_instr :: "[instr, jvm_prog, heap, val list, val list,
- cname, mname, pc, frame list, sheap] \<Rightarrow> bool" where
- check_instr_Load:
- "check_instr (Load n) P h stk loc C M\<^sub>0 pc frs sh =
- (n < length loc)"
-
-| check_instr_Store:
- "check_instr (Store n) P h stk loc C\<^sub>0 M\<^sub>0 pc frs sh =
- (0 < length stk \<and> n < length loc)"
-
-| check_instr_Push:
- "check_instr (Push v) P h stk loc C\<^sub>0 M\<^sub>0 pc frs sh =
- (\<not>is_Addr v)"
-
-| check_instr_New:
- "check_instr (New C) P h stk loc C\<^sub>0 M\<^sub>0 pc frs sh =
- is_class P C"
-
-| check_instr_Getfield:
- "check_instr (Getfield F C) P h stk loc C\<^sub>0 M\<^sub>0 pc frs sh =
- (0 < length stk \<and> (\<exists>C' T. P \<turnstile> C sees F,NonStatic:T in C') \<and>
- (let (C', b, T) = field P C F; ref = hd stk in
- C' = C \<and> is_Ref ref \<and> (ref \<noteq> Null \<longrightarrow>
- h (the_Addr ref) \<noteq> None \<and>
- (let (D,vs) = the (h (the_Addr ref)) in
- P \<turnstile> D \<preceq>\<^sup>* C \<and> vs (F,C) \<noteq> None \<and> P,h \<turnstile> the (vs (F,C)) :\<le> T))))"
-
-| check_instr_Getstatic:
- "check_instr (Getstatic C F D) P h stk loc C\<^sub>0 M\<^sub>0 pc frs sh =
- ((\<exists>T. P \<turnstile> C sees F,Static:T in D) \<and>
- (let (C', b, T) = field P C F in
- C' = D \<and> (sh D \<noteq> None \<longrightarrow>
- (let (sfs,i) = the (sh D) in
- sfs F \<noteq> None \<and> P,h \<turnstile> the (sfs F) :\<le> T))))"
-
-| check_instr_Putfield:
- "check_instr (Putfield F C) P h stk loc C\<^sub>0 M\<^sub>0 pc frs sh =
- (1 < length stk \<and> (\<exists>C' T. P \<turnstile> C sees F,NonStatic:T in C') \<and>
- (let (C', b, T) = field P C F; v = hd stk; ref = hd (tl stk) in
- C' = C \<and> is_Ref ref \<and> (ref \<noteq> Null \<longrightarrow>
- h (the_Addr ref) \<noteq> None \<and>
- (let D = fst (the (h (the_Addr ref))) in
- P \<turnstile> D \<preceq>\<^sup>* C \<and> P,h \<turnstile> v :\<le> T))))"
-
-| check_instr_Putstatic:
- "check_instr (Putstatic C F D) P h stk loc C\<^sub>0 M\<^sub>0 pc frs sh =
- (0 < length stk \<and> (\<exists>T. P \<turnstile> C sees F,Static:T in D) \<and>
- (let (C', b, T) = field P C F; v = hd stk in
- C' = D \<and> P,h \<turnstile> v :\<le> T))"
-
-| check_instr_Checkcast:
- "check_instr (Checkcast C) P h stk loc C\<^sub>0 M\<^sub>0 pc frs sh =
- (0 < length stk \<and> is_class P C \<and> is_Ref (hd stk))"
-
-| check_instr_Invoke:
- "check_instr (Invoke M n) P h stk loc C\<^sub>0 M\<^sub>0 pc frs sh =
- (n < length stk \<and> is_Ref (stk!n) \<and>
- (stk!n \<noteq> Null \<longrightarrow>
- (let a = the_Addr (stk!n);
- C = cname_of h a;
- Ts = fst (snd (snd (method P C M)))
- in h a \<noteq> None \<and> P \<turnstile> C has M,NonStatic \<and>
- P,h \<turnstile> rev (take n stk) [:\<le>] Ts)))"
-
-| check_instr_Invokestatic:
- "check_instr (Invokestatic C M n) P h stk loc C\<^sub>0 M\<^sub>0 pc frs sh =
- (n \<le> length stk \<and>
- (let Ts = fst (snd (snd (method P C M)))
- in P \<turnstile> C has M,Static \<and>
- P,h \<turnstile> rev (take n stk) [:\<le>] Ts))"
-
-| check_instr_Return:
- "check_instr Return P h stk loc C\<^sub>0 M\<^sub>0 pc frs sh =
- (case (M\<^sub>0 = clinit) of False \<Rightarrow> (0 < length stk \<and> ((0 < length frs) \<longrightarrow>
- (\<exists>b. P \<turnstile> C\<^sub>0 has M\<^sub>0,b) \<and>
- (let v = hd stk;
- T = fst (snd (snd (snd (method P C\<^sub>0 M\<^sub>0))))
- in P,h \<turnstile> v :\<le> T)))
- | True \<Rightarrow> P \<turnstile> C\<^sub>0 has M\<^sub>0,Static)"
-
-| check_instr_Pop:
- "check_instr Pop P h stk loc C\<^sub>0 M\<^sub>0 pc frs sh =
- (0 < length stk)"
-
-| check_instr_IAdd:
- "check_instr IAdd P h stk loc C\<^sub>0 M\<^sub>0 pc frs sh =
- (1 < length stk \<and> is_Intg (hd stk) \<and> is_Intg (hd (tl stk)))"
-
-| check_instr_IfFalse:
- "check_instr (IfFalse b) P h stk loc C\<^sub>0 M\<^sub>0 pc frs sh =
- (0 < length stk \<and> is_Bool (hd stk) \<and> 0 \<le> int pc+b)"
-
-| check_instr_CmpEq:
- "check_instr CmpEq P h stk loc C\<^sub>0 M\<^sub>0 pc frs sh =
- (1 < length stk)"
-
-| check_instr_Goto:
- "check_instr (Goto b) P h stk loc C\<^sub>0 M\<^sub>0 pc frs sh =
- (0 \<le> int pc+b)"
-
-| check_instr_Throw:
- "check_instr Throw P h stk loc C\<^sub>0 M\<^sub>0 pc frs sh =
- (0 < length stk \<and> is_Ref (hd stk))"
-
-definition check :: "jvm_prog \<Rightarrow> jvm_state \<Rightarrow> bool" where
- "check P \<sigma> = (let (xcpt, h, frs, sh) = \<sigma> in
- (case frs of [] \<Rightarrow> True | (stk,loc,C,M,pc,ics)#frs' \<Rightarrow>
- \<exists>b. P \<turnstile> C has M, b \<and>
- (let (C',b,Ts,T,mxs,mxl\<^sub>0,ins,xt) = method P C M; i = ins!pc in
- pc < size ins \<and> size stk \<le> mxs \<and>
- check_instr i P h stk loc C M pc frs' sh)))"
-
-
-definition exec_d :: "jvm_prog \<Rightarrow> jvm_state \<Rightarrow> jvm_state option type_error" where
- "exec_d P \<sigma> = (if check P \<sigma> then Normal (exec (P, \<sigma>)) else TypeError)"
-
-
-inductive_set
- exec_1_d :: "jvm_prog \<Rightarrow> (jvm_state type_error \<times> jvm_state type_error) set"
- and exec_1_d' :: "jvm_prog \<Rightarrow> jvm_state type_error \<Rightarrow> jvm_state type_error \<Rightarrow> bool"
- ("_ \<turnstile> _ -jvmd\<rightarrow>\<^sub>1 _" [61,61,61]60)
- for P :: jvm_prog
-where
- "P \<turnstile> \<sigma> -jvmd\<rightarrow>\<^sub>1 \<sigma>' \<equiv> (\<sigma>,\<sigma>') \<in> exec_1_d P"
-| exec_1_d_ErrorI: "exec_d P \<sigma> = TypeError \<Longrightarrow> P \<turnstile> Normal \<sigma> -jvmd\<rightarrow>\<^sub>1 TypeError"
-| exec_1_d_NormalI: "exec_d P \<sigma> = Normal (Some \<sigma>') \<Longrightarrow> P \<turnstile> Normal \<sigma> -jvmd\<rightarrow>\<^sub>1 Normal \<sigma>'"
-
-\<comment> \<open>reflexive transitive closure:\<close>
-definition exec_all_d :: "jvm_prog \<Rightarrow> jvm_state type_error \<Rightarrow> jvm_state type_error \<Rightarrow> bool"
- ("_ \<turnstile> _ -jvmd\<rightarrow> _" [61,61,61]60) where
- exec_all_d_def1: "P \<turnstile> \<sigma> -jvmd\<rightarrow> \<sigma>' \<longleftrightarrow> (\<sigma>,\<sigma>') \<in> (exec_1_d P)\<^sup>*"
-
-notation (ASCII)
- "exec_all_d" ("_ |- _ -jvmd-> _" [61,61,61]60)
-
-lemma exec_1_d_eq:
- "exec_1_d P = {(s,t). \<exists>\<sigma>. s = Normal \<sigma> \<and> t = TypeError \<and> exec_d P \<sigma> = TypeError} \<union>
- {(s,t). \<exists>\<sigma> \<sigma>'. s = Normal \<sigma> \<and> t = Normal \<sigma>' \<and> exec_d P \<sigma> = Normal (Some \<sigma>')}"
-by (auto elim!: exec_1_d.cases intro!: exec_1_d.intros)
-
-
-declare split_paired_All [simp del]
-declare split_paired_Ex [simp del]
-
-lemma if_neq [dest!]:
- "(if P then A else B) \<noteq> B \<Longrightarrow> P"
- by (cases P, auto)
-
-lemma exec_d_no_errorI [intro]:
- "check P \<sigma> \<Longrightarrow> exec_d P \<sigma> \<noteq> TypeError"
- by (unfold exec_d_def) simp
-
-theorem no_type_error_commutes:
- "exec_d P \<sigma> \<noteq> TypeError \<Longrightarrow> exec_d P \<sigma> = Normal (exec (P, \<sigma>))"
- by (unfold exec_d_def, auto)
-
-
-lemma defensive_imp_aggressive:
- "P \<turnstile> (Normal \<sigma>) -jvmd\<rightarrow> (Normal \<sigma>') \<Longrightarrow> P \<turnstile> \<sigma> -jvm\<rightarrow> \<sigma>'"
-(*<*)
-proof -
- have "\<And>x y. P \<turnstile> x -jvmd\<rightarrow> y \<Longrightarrow> \<forall>\<sigma> \<sigma>'. x = Normal \<sigma> \<longrightarrow> y = Normal \<sigma>' \<longrightarrow> P \<turnstile> \<sigma> -jvm\<rightarrow> \<sigma>'"
- proof -
- fix x y assume xy: "P \<turnstile> x -jvmd\<rightarrow> y"
- then have "(x, y) \<in> (exec_1_d P)\<^sup>*" by (unfold exec_all_d_def1)
- then show "\<forall>\<sigma> \<sigma>'. x = Normal \<sigma> \<longrightarrow> y = Normal \<sigma>' \<longrightarrow> P \<turnstile> \<sigma> -jvm\<rightarrow> \<sigma>'"
- proof(induct rule: rtrancl_induct)
- case base
- then show ?case by (simp add: exec_all_def)
- next
- case (step y' z')
- show ?case proof(induct rule: exec_1_d.cases[OF step.hyps(2)])
- case (2 \<sigma> \<sigma>')
- then have "(\<sigma>, \<sigma>') \<in> {(\<sigma>, \<sigma>'). exec (P, \<sigma>) = \<lfloor>\<sigma>'\<rfloor>}\<^sup>*" using step
- by(fastforce simp: exec_d_def split: type_error.splits if_split_asm)
- then show ?case using step 2 by (auto simp: exec_all_def)
- qed simp
- qed
- qed
- moreover
- assume "P \<turnstile> (Normal \<sigma>) -jvmd\<rightarrow> (Normal \<sigma>')"
- ultimately
- show "P \<turnstile> \<sigma> -jvm\<rightarrow> \<sigma>'" by blast
-qed
-(*>*)
-
-end
+(* Title: JinjaDCI/JVM/JVMDefensive.thy
+ Author: Gerwin Klein, Susannah Mansky
+ Copyright GPL
+
+ Based on the Jinja theory JVM/JVMDefensive.thy by Gerwin Klein
+*)
+
+section \<open> A Defensive JVM \<close>
+
+theory JVMDefensive
+imports JVMExec "../Common/Conform"
+begin
+
+text \<open>
+ Extend the state space by one element indicating a type error (or
+ other abnormal termination) \<close>
+datatype 'a type_error = TypeError | Normal 'a
+
+fun is_Addr :: "val \<Rightarrow> bool" where
+ "is_Addr (Addr a) \<longleftrightarrow> True"
+| "is_Addr v \<longleftrightarrow> False"
+
+fun is_Intg :: "val \<Rightarrow> bool" where
+ "is_Intg (Intg i) \<longleftrightarrow> True"
+| "is_Intg v \<longleftrightarrow> False"
+
+fun is_Bool :: "val \<Rightarrow> bool" where
+ "is_Bool (Bool b) \<longleftrightarrow> True"
+| "is_Bool v \<longleftrightarrow> False"
+
+definition is_Ref :: "val \<Rightarrow> bool" where
+ "is_Ref v \<longleftrightarrow> v = Null \<or> is_Addr v"
+
+primrec check_instr :: "[instr, jvm_prog, heap, val list, val list,
+ cname, mname, pc, frame list, sheap] \<Rightarrow> bool" where
+ check_instr_Load:
+ "check_instr (Load n) P h stk loc C M\<^sub>0 pc frs sh =
+ (n < length loc)"
+
+| check_instr_Store:
+ "check_instr (Store n) P h stk loc C\<^sub>0 M\<^sub>0 pc frs sh =
+ (0 < length stk \<and> n < length loc)"
+
+| check_instr_Push:
+ "check_instr (Push v) P h stk loc C\<^sub>0 M\<^sub>0 pc frs sh =
+ (\<not>is_Addr v)"
+
+| check_instr_New:
+ "check_instr (New C) P h stk loc C\<^sub>0 M\<^sub>0 pc frs sh =
+ is_class P C"
+
+| check_instr_Getfield:
+ "check_instr (Getfield F C) P h stk loc C\<^sub>0 M\<^sub>0 pc frs sh =
+ (0 < length stk \<and> (\<exists>C' T. P \<turnstile> C sees F,NonStatic:T in C') \<and>
+ (let (C', b, T) = field P C F; ref = hd stk in
+ C' = C \<and> is_Ref ref \<and> (ref \<noteq> Null \<longrightarrow>
+ h (the_Addr ref) \<noteq> None \<and>
+ (let (D,vs) = the (h (the_Addr ref)) in
+ P \<turnstile> D \<preceq>\<^sup>* C \<and> vs (F,C) \<noteq> None \<and> P,h \<turnstile> the (vs (F,C)) :\<le> T))))"
+
+| check_instr_Getstatic:
+ "check_instr (Getstatic C F D) P h stk loc C\<^sub>0 M\<^sub>0 pc frs sh =
+ ((\<exists>T. P \<turnstile> C sees F,Static:T in D) \<and>
+ (let (C', b, T) = field P C F in
+ C' = D \<and> (sh D \<noteq> None \<longrightarrow>
+ (let (sfs,i) = the (sh D) in
+ sfs F \<noteq> None \<and> P,h \<turnstile> the (sfs F) :\<le> T))))"
+
+| check_instr_Putfield:
+ "check_instr (Putfield F C) P h stk loc C\<^sub>0 M\<^sub>0 pc frs sh =
+ (1 < length stk \<and> (\<exists>C' T. P \<turnstile> C sees F,NonStatic:T in C') \<and>
+ (let (C', b, T) = field P C F; v = hd stk; ref = hd (tl stk) in
+ C' = C \<and> is_Ref ref \<and> (ref \<noteq> Null \<longrightarrow>
+ h (the_Addr ref) \<noteq> None \<and>
+ (let D = fst (the (h (the_Addr ref))) in
+ P \<turnstile> D \<preceq>\<^sup>* C \<and> P,h \<turnstile> v :\<le> T))))"
+
+| check_instr_Putstatic:
+ "check_instr (Putstatic C F D) P h stk loc C\<^sub>0 M\<^sub>0 pc frs sh =
+ (0 < length stk \<and> (\<exists>T. P \<turnstile> C sees F,Static:T in D) \<and>
+ (let (C', b, T) = field P C F; v = hd stk in
+ C' = D \<and> P,h \<turnstile> v :\<le> T))"
+
+| check_instr_Checkcast:
+ "check_instr (Checkcast C) P h stk loc C\<^sub>0 M\<^sub>0 pc frs sh =
+ (0 < length stk \<and> is_class P C \<and> is_Ref (hd stk))"
+
+| check_instr_Invoke:
+ "check_instr (Invoke M n) P h stk loc C\<^sub>0 M\<^sub>0 pc frs sh =
+ (n < length stk \<and> is_Ref (stk!n) \<and>
+ (stk!n \<noteq> Null \<longrightarrow>
+ (let a = the_Addr (stk!n);
+ C = cname_of h a;
+ Ts = fst (snd (snd (method P C M)))
+ in h a \<noteq> None \<and> P \<turnstile> C has M,NonStatic \<and>
+ P,h \<turnstile> rev (take n stk) [:\<le>] Ts)))"
+
+| check_instr_Invokestatic:
+ "check_instr (Invokestatic C M n) P h stk loc C\<^sub>0 M\<^sub>0 pc frs sh =
+ (n \<le> length stk \<and>
+ (let Ts = fst (snd (snd (method P C M)))
+ in P \<turnstile> C has M,Static \<and>
+ P,h \<turnstile> rev (take n stk) [:\<le>] Ts))"
+
+| check_instr_Return:
+ "check_instr Return P h stk loc C\<^sub>0 M\<^sub>0 pc frs sh =
+ (case (M\<^sub>0 = clinit) of False \<Rightarrow> (0 < length stk \<and> ((0 < length frs) \<longrightarrow>
+ (\<exists>b. P \<turnstile> C\<^sub>0 has M\<^sub>0,b) \<and>
+ (let v = hd stk;
+ T = fst (snd (snd (snd (method P C\<^sub>0 M\<^sub>0))))
+ in P,h \<turnstile> v :\<le> T)))
+ | True \<Rightarrow> P \<turnstile> C\<^sub>0 has M\<^sub>0,Static)"
+
+| check_instr_Pop:
+ "check_instr Pop P h stk loc C\<^sub>0 M\<^sub>0 pc frs sh =
+ (0 < length stk)"
+
+| check_instr_IAdd:
+ "check_instr IAdd P h stk loc C\<^sub>0 M\<^sub>0 pc frs sh =
+ (1 < length stk \<and> is_Intg (hd stk) \<and> is_Intg (hd (tl stk)))"
+
+| check_instr_IfFalse:
+ "check_instr (IfFalse b) P h stk loc C\<^sub>0 M\<^sub>0 pc frs sh =
+ (0 < length stk \<and> is_Bool (hd stk) \<and> 0 \<le> int pc+b)"
+
+| check_instr_CmpEq:
+ "check_instr CmpEq P h stk loc C\<^sub>0 M\<^sub>0 pc frs sh =
+ (1 < length stk)"
+
+| check_instr_Goto:
+ "check_instr (Goto b) P h stk loc C\<^sub>0 M\<^sub>0 pc frs sh =
+ (0 \<le> int pc+b)"
+
+| check_instr_Throw:
+ "check_instr Throw P h stk loc C\<^sub>0 M\<^sub>0 pc frs sh =
+ (0 < length stk \<and> is_Ref (hd stk))"
+
+definition check :: "jvm_prog \<Rightarrow> jvm_state \<Rightarrow> bool" where
+ "check P \<sigma> = (let (xcpt, h, frs, sh) = \<sigma> in
+ (case frs of [] \<Rightarrow> True | (stk,loc,C,M,pc,ics)#frs' \<Rightarrow>
+ \<exists>b. P \<turnstile> C has M, b \<and>
+ (let (C',b,Ts,T,mxs,mxl\<^sub>0,ins,xt) = method P C M; i = ins!pc in
+ pc < size ins \<and> size stk \<le> mxs \<and>
+ check_instr i P h stk loc C M pc frs' sh)))"
+
+
+definition exec_d :: "jvm_prog \<Rightarrow> jvm_state \<Rightarrow> jvm_state option type_error" where
+ "exec_d P \<sigma> = (if check P \<sigma> then Normal (exec (P, \<sigma>)) else TypeError)"
+
+
+inductive_set
+ exec_1_d :: "jvm_prog \<Rightarrow> (jvm_state type_error \<times> jvm_state type_error) set"
+ and exec_1_d' :: "jvm_prog \<Rightarrow> jvm_state type_error \<Rightarrow> jvm_state type_error \<Rightarrow> bool"
+ ("_ \<turnstile> _ -jvmd\<rightarrow>\<^sub>1 _" [61,61,61]60)
+ for P :: jvm_prog
+where
+ "P \<turnstile> \<sigma> -jvmd\<rightarrow>\<^sub>1 \<sigma>' \<equiv> (\<sigma>,\<sigma>') \<in> exec_1_d P"
+| exec_1_d_ErrorI: "exec_d P \<sigma> = TypeError \<Longrightarrow> P \<turnstile> Normal \<sigma> -jvmd\<rightarrow>\<^sub>1 TypeError"
+| exec_1_d_NormalI: "exec_d P \<sigma> = Normal (Some \<sigma>') \<Longrightarrow> P \<turnstile> Normal \<sigma> -jvmd\<rightarrow>\<^sub>1 Normal \<sigma>'"
+
+\<comment> \<open>reflexive transitive closure:\<close>
+definition exec_all_d :: "jvm_prog \<Rightarrow> jvm_state type_error \<Rightarrow> jvm_state type_error \<Rightarrow> bool"
+ ("_ \<turnstile> _ -jvmd\<rightarrow> _" [61,61,61]60) where
+ exec_all_d_def1: "P \<turnstile> \<sigma> -jvmd\<rightarrow> \<sigma>' \<longleftrightarrow> (\<sigma>,\<sigma>') \<in> (exec_1_d P)\<^sup>*"
+
+notation (ASCII)
+ "exec_all_d" ("_ |- _ -jvmd-> _" [61,61,61]60)
+
+lemma exec_1_d_eq:
+ "exec_1_d P = {(s,t). \<exists>\<sigma>. s = Normal \<sigma> \<and> t = TypeError \<and> exec_d P \<sigma> = TypeError} \<union>
+ {(s,t). \<exists>\<sigma> \<sigma>'. s = Normal \<sigma> \<and> t = Normal \<sigma>' \<and> exec_d P \<sigma> = Normal (Some \<sigma>')}"
+by (auto elim!: exec_1_d.cases intro!: exec_1_d.intros)
+
+
+declare split_paired_All [simp del]
+declare split_paired_Ex [simp del]
+
+lemma if_neq [dest!]:
+ "(if P then A else B) \<noteq> B \<Longrightarrow> P"
+ by (cases P, auto)
+
+lemma exec_d_no_errorI [intro]:
+ "check P \<sigma> \<Longrightarrow> exec_d P \<sigma> \<noteq> TypeError"
+ by (unfold exec_d_def) simp
+
+theorem no_type_error_commutes:
+ "exec_d P \<sigma> \<noteq> TypeError \<Longrightarrow> exec_d P \<sigma> = Normal (exec (P, \<sigma>))"
+ by (unfold exec_d_def, auto)
+
+
+lemma defensive_imp_aggressive:
+ "P \<turnstile> (Normal \<sigma>) -jvmd\<rightarrow> (Normal \<sigma>') \<Longrightarrow> P \<turnstile> \<sigma> -jvm\<rightarrow> \<sigma>'"
+(*<*)
+proof -
+ have "\<And>x y. P \<turnstile> x -jvmd\<rightarrow> y \<Longrightarrow> \<forall>\<sigma> \<sigma>'. x = Normal \<sigma> \<longrightarrow> y = Normal \<sigma>' \<longrightarrow> P \<turnstile> \<sigma> -jvm\<rightarrow> \<sigma>'"
+ proof -
+ fix x y assume xy: "P \<turnstile> x -jvmd\<rightarrow> y"
+ then have "(x, y) \<in> (exec_1_d P)\<^sup>*" by (unfold exec_all_d_def1)
+ then show "\<forall>\<sigma> \<sigma>'. x = Normal \<sigma> \<longrightarrow> y = Normal \<sigma>' \<longrightarrow> P \<turnstile> \<sigma> -jvm\<rightarrow> \<sigma>'"
+ proof(induct rule: rtrancl_induct)
+ case base
+ then show ?case by (simp add: exec_all_def)
+ next
+ case (step y' z')
+ show ?case proof(induct rule: exec_1_d.cases[OF step.hyps(2)])
+ case (2 \<sigma> \<sigma>')
+ then have "(\<sigma>, \<sigma>') \<in> {(\<sigma>, \<sigma>'). exec (P, \<sigma>) = \<lfloor>\<sigma>'\<rfloor>}\<^sup>*" using step
+ by(fastforce simp: exec_d_def split: type_error.splits if_split_asm)
+ then show ?case using step 2 by (auto simp: exec_all_def)
+ qed simp
+ qed
+ qed
+ moreover
+ assume "P \<turnstile> (Normal \<sigma>) -jvmd\<rightarrow> (Normal \<sigma>')"
+ ultimately
+ show "P \<turnstile> \<sigma> -jvm\<rightarrow> \<sigma>'" by blast
+qed
+(*>*)
+
+end
diff --git a/thys/JinjaDCI/JVM/JVMExceptions.thy b/thys/JinjaDCI/JVM/JVMExceptions.thy
--- a/thys/JinjaDCI/JVM/JVMExceptions.thy
+++ b/thys/JinjaDCI/JVM/JVMExceptions.thy
@@ -1,116 +1,116 @@
-(* Title: JinjaDCI/JVM/JVMExceptions.thy
- Author: Gerwin Klein, Martin Strecker, Susannah Mansky
- Copyright 2001 Technische Universitaet Muenchen, 2019-20 UIUC
-
- Based on the Jinja theory JVM/JVMExceptions.thy by Gerwin Klein and Martin Strecker
-*)
-
-section \<open> Exception handling in the JVM \<close>
-
-theory JVMExceptions imports "../Common/Exceptions" JVMInstructions
-begin
-
-definition matches_ex_entry :: "'m prog \<Rightarrow> cname \<Rightarrow> pc \<Rightarrow> ex_entry \<Rightarrow> bool"
-where
- "matches_ex_entry P C pc xcp \<equiv>
- let (s, e, C', h, d) = xcp in
- s \<le> pc \<and> pc < e \<and> P \<turnstile> C \<preceq>\<^sup>* C'"
-
-
-primrec match_ex_table :: "'m prog \<Rightarrow> cname \<Rightarrow> pc \<Rightarrow> ex_table \<Rightarrow> (pc \<times> nat) option"
-where
- "match_ex_table P C pc [] = None"
-| "match_ex_table P C pc (e#es) = (if matches_ex_entry P C pc e
- then Some (snd(snd(snd e)))
- else match_ex_table P C pc es)"
-
-abbreviation
- ex_table_of :: "jvm_prog \<Rightarrow> cname \<Rightarrow> mname \<Rightarrow> ex_table" where
- "ex_table_of P C M == snd (snd (snd (snd (snd (snd (snd (method P C M)))))))"
-
-
-fun find_handler :: "jvm_prog \<Rightarrow> addr \<Rightarrow> heap \<Rightarrow> frame list \<Rightarrow> sheap \<Rightarrow> jvm_state"
-where
- "find_handler P a h [] sh = (Some a, h, [], sh)"
-| "find_handler P a h (fr#frs) sh =
- (let (stk,loc,C,M,pc,ics) = fr in
- case match_ex_table P (cname_of h a) pc (ex_table_of P C M) of
- None \<Rightarrow>
- (case M = clinit of
- True \<Rightarrow> (case frs of (stk',loc',C',M',pc',ics')#frs'
- \<Rightarrow> (case ics' of Called Cs \<Rightarrow> (None, h, (stk',loc',C',M',pc',Throwing (C#Cs) a)#frs', sh)
- | _ \<Rightarrow> (None, h, (stk',loc',C',M',pc',ics')#frs', sh) \<comment> \<open>this won't happen in wf code\<close>
- )
- | [] \<Rightarrow> (Some a, h, [], sh)
- )
- | _ \<Rightarrow> find_handler P a h frs sh
- )
- | Some pc_d \<Rightarrow> (None, h, (Addr a # drop (size stk - snd pc_d) stk, loc, C, M, fst pc_d, No_ics)#frs, sh))"
-
-lemma find_handler_cases:
- "find_handler P a h frs sh = js
- \<Longrightarrow> (\<exists>frs'. frs' \<noteq> [] \<and> js = (None, h, frs', sh)) \<or> (js = (Some a, h, [], sh))"
-proof(induct P a h frs sh rule: find_handler.induct)
- case 1 then show ?case by clarsimp
-next
- case (2 P a h fr frs sh) then show ?case
- by(cases fr, auto split: bool.splits list.splits init_call_status.splits)
-qed
-
-lemma find_handler_heap[simp]:
-"find_handler P a h frs sh = (xp',h',frs',sh') \<Longrightarrow> h' = h"
- by(auto dest: find_handler_cases)
-
-lemma find_handler_sheap[simp]:
-"find_handler P a h frs sh = (xp',h',frs',sh') \<Longrightarrow> sh' = sh"
- by(auto dest: find_handler_cases)
-
-lemma find_handler_frames[simp]:
-"find_handler P a h frs sh = (xp',h',frs',sh') \<Longrightarrow> length frs' \<le> length frs"
-proof(induct frs)
- case Nil then show ?case by simp
-next
- case (Cons a frs) then show ?case
- by(auto simp: split_beta split: bool.splits list.splits init_call_status.splits)
-qed
-
-lemma find_handler_None:
- "find_handler P a h frs sh = (None, h, frs', sh') \<Longrightarrow> frs' \<noteq> []"
- by (drule find_handler_cases, clarsimp)
-
-lemma find_handler_Some:
- "find_handler P a h frs sh = (Some x, h, frs', sh') \<Longrightarrow> frs' = []"
- by (drule find_handler_cases, clarsimp)
-
-lemma find_handler_Some_same_error_same_heap[simp]:
- "find_handler P a h frs sh = (Some x, h', frs', sh') \<Longrightarrow> x = a \<and> h = h' \<and> sh = sh'"
- by(auto dest: find_handler_cases)
-
-lemma find_handler_prealloc_pres:
-assumes "preallocated h"
-and fh: "find_handler P a h frs sh = (xp',h',frs',sh')"
-shows "preallocated h'"
-using assms find_handler_heap[OF fh] by simp
-
-lemma find_handler_frs_tl_neq:
- "ics_of f \<noteq> No_ics
- \<Longrightarrow> (xp, h, f#frs, sh) \<noteq> find_handler P xa h' (f' # frs) sh'"
-proof(induct frs arbitrary: f f')
- case Nil then show ?case by(auto simp: split_beta split: bool.splits)
-next
- case (Cons a frs)
- obtain xp1 h1 frs1 sh1 where fh: "find_handler P xa h' (a # frs) sh' = (xp1,h1,frs1,sh1)"
- by(cases "find_handler P xa h' (a # frs) sh'")
- then have "length frs1 \<le> length (a#frs)"
- by(rule find_handler_frames[where P=P and a=xa and h=h' and frs="a#frs" and sh=sh'])
- then have neq: "f#a#frs \<noteq> frs1" by(clarsimp dest: impossible_Cons)
- then show ?case
- proof(cases "find_handler P xa h' (f' # a # frs) sh' = find_handler P xa h' (a # frs) sh'")
- case True then show ?thesis using neq fh by simp
- next
- case False then show ?thesis using Cons.prems
- by(fastforce simp: split_beta split: bool.splits init_call_status.splits list.splits)
- qed
-qed
-
-end
+(* Title: JinjaDCI/JVM/JVMExceptions.thy
+ Author: Gerwin Klein, Martin Strecker, Susannah Mansky
+ Copyright 2001 Technische Universitaet Muenchen, 2019-20 UIUC
+
+ Based on the Jinja theory JVM/JVMExceptions.thy by Gerwin Klein and Martin Strecker
+*)
+
+section \<open> Exception handling in the JVM \<close>
+
+theory JVMExceptions imports "../Common/Exceptions" JVMInstructions
+begin
+
+definition matches_ex_entry :: "'m prog \<Rightarrow> cname \<Rightarrow> pc \<Rightarrow> ex_entry \<Rightarrow> bool"
+where
+ "matches_ex_entry P C pc xcp \<equiv>
+ let (s, e, C', h, d) = xcp in
+ s \<le> pc \<and> pc < e \<and> P \<turnstile> C \<preceq>\<^sup>* C'"
+
+
+primrec match_ex_table :: "'m prog \<Rightarrow> cname \<Rightarrow> pc \<Rightarrow> ex_table \<Rightarrow> (pc \<times> nat) option"
+where
+ "match_ex_table P C pc [] = None"
+| "match_ex_table P C pc (e#es) = (if matches_ex_entry P C pc e
+ then Some (snd(snd(snd e)))
+ else match_ex_table P C pc es)"
+
+abbreviation
+ ex_table_of :: "jvm_prog \<Rightarrow> cname \<Rightarrow> mname \<Rightarrow> ex_table" where
+ "ex_table_of P C M == snd (snd (snd (snd (snd (snd (snd (method P C M)))))))"
+
+
+fun find_handler :: "jvm_prog \<Rightarrow> addr \<Rightarrow> heap \<Rightarrow> frame list \<Rightarrow> sheap \<Rightarrow> jvm_state"
+where
+ "find_handler P a h [] sh = (Some a, h, [], sh)"
+| "find_handler P a h (fr#frs) sh =
+ (let (stk,loc,C,M,pc,ics) = fr in
+ case match_ex_table P (cname_of h a) pc (ex_table_of P C M) of
+ None \<Rightarrow>
+ (case M = clinit of
+ True \<Rightarrow> (case frs of (stk',loc',C',M',pc',ics')#frs'
+ \<Rightarrow> (case ics' of Called Cs \<Rightarrow> (None, h, (stk',loc',C',M',pc',Throwing (C#Cs) a)#frs', sh)
+ | _ \<Rightarrow> (None, h, (stk',loc',C',M',pc',ics')#frs', sh) \<comment> \<open>this won't happen in wf code\<close>
+ )
+ | [] \<Rightarrow> (Some a, h, [], sh)
+ )
+ | _ \<Rightarrow> find_handler P a h frs sh
+ )
+ | Some pc_d \<Rightarrow> (None, h, (Addr a # drop (size stk - snd pc_d) stk, loc, C, M, fst pc_d, No_ics)#frs, sh))"
+
+lemma find_handler_cases:
+ "find_handler P a h frs sh = js
+ \<Longrightarrow> (\<exists>frs'. frs' \<noteq> [] \<and> js = (None, h, frs', sh)) \<or> (js = (Some a, h, [], sh))"
+proof(induct P a h frs sh rule: find_handler.induct)
+ case 1 then show ?case by clarsimp
+next
+ case (2 P a h fr frs sh) then show ?case
+ by(cases fr, auto split: bool.splits list.splits init_call_status.splits)
+qed
+
+lemma find_handler_heap[simp]:
+"find_handler P a h frs sh = (xp',h',frs',sh') \<Longrightarrow> h' = h"
+ by(auto dest: find_handler_cases)
+
+lemma find_handler_sheap[simp]:
+"find_handler P a h frs sh = (xp',h',frs',sh') \<Longrightarrow> sh' = sh"
+ by(auto dest: find_handler_cases)
+
+lemma find_handler_frames[simp]:
+"find_handler P a h frs sh = (xp',h',frs',sh') \<Longrightarrow> length frs' \<le> length frs"
+proof(induct frs)
+ case Nil then show ?case by simp
+next
+ case (Cons a frs) then show ?case
+ by(auto simp: split_beta split: bool.splits list.splits init_call_status.splits)
+qed
+
+lemma find_handler_None:
+ "find_handler P a h frs sh = (None, h, frs', sh') \<Longrightarrow> frs' \<noteq> []"
+ by (drule find_handler_cases, clarsimp)
+
+lemma find_handler_Some:
+ "find_handler P a h frs sh = (Some x, h, frs', sh') \<Longrightarrow> frs' = []"
+ by (drule find_handler_cases, clarsimp)
+
+lemma find_handler_Some_same_error_same_heap[simp]:
+ "find_handler P a h frs sh = (Some x, h', frs', sh') \<Longrightarrow> x = a \<and> h = h' \<and> sh = sh'"
+ by(auto dest: find_handler_cases)
+
+lemma find_handler_prealloc_pres:
+assumes "preallocated h"
+and fh: "find_handler P a h frs sh = (xp',h',frs',sh')"
+shows "preallocated h'"
+using assms find_handler_heap[OF fh] by simp
+
+lemma find_handler_frs_tl_neq:
+ "ics_of f \<noteq> No_ics
+ \<Longrightarrow> (xp, h, f#frs, sh) \<noteq> find_handler P xa h' (f' # frs) sh'"
+proof(induct frs arbitrary: f f')
+ case Nil then show ?case by(auto simp: split_beta split: bool.splits)
+next
+ case (Cons a frs)
+ obtain xp1 h1 frs1 sh1 where fh: "find_handler P xa h' (a # frs) sh' = (xp1,h1,frs1,sh1)"
+ by(cases "find_handler P xa h' (a # frs) sh'")
+ then have "length frs1 \<le> length (a#frs)"
+ by(rule find_handler_frames[where P=P and a=xa and h=h' and frs="a#frs" and sh=sh'])
+ then have neq: "f#a#frs \<noteq> frs1" by(clarsimp dest: impossible_Cons)
+ then show ?case
+ proof(cases "find_handler P xa h' (f' # a # frs) sh' = find_handler P xa h' (a # frs) sh'")
+ case True then show ?thesis using neq fh by simp
+ next
+ case False then show ?thesis using Cons.prems
+ by(fastforce simp: split_beta split: bool.splits init_call_status.splits list.splits)
+ qed
+qed
+
+end
diff --git a/thys/JinjaDCI/JVM/JVMExec.thy b/thys/JinjaDCI/JVM/JVMExec.thy
--- a/thys/JinjaDCI/JVM/JVMExec.thy
+++ b/thys/JinjaDCI/JVM/JVMExec.thy
@@ -1,238 +1,238 @@
-
-(* Title: JinjaDCI/JVM/JVMExec.thy
- Author: Cornelia Pusch, Gerwin Klein, Susannah Mansky
- Copyright 1999 Technische Universitaet Muenchen, 2019-20
-
- Based on the Jinja theory JVM/JVMExec.thy by Cornelia Pusch and Gerwin Klein
-*)
-
-section \<open> Program Execution in the JVM in full small step style \<close>
-
-theory JVMExec
-imports JVMExecInstr
-begin
-
-abbreviation
- instrs_of :: "jvm_prog \<Rightarrow> cname \<Rightarrow> mname \<Rightarrow> instr list" where
- "instrs_of P C M == fst(snd(snd(snd(snd(snd(snd(method P C M)))))))"
-
-fun curr_instr :: "jvm_prog \<Rightarrow> frame \<Rightarrow> instr" where
-"curr_instr P (stk,loc,C,M,pc,ics) = instrs_of P C M ! pc"
-
-\<comment> \<open> execution of single step of the initialization procedure \<close>
-fun exec_Calling :: "[cname, cname list, jvm_prog, heap, val list, val list,
- cname, mname, pc, frame list, sheap] \<Rightarrow> jvm_state"
-where
-"exec_Calling C Cs P h stk loc C\<^sub>0 M\<^sub>0 pc frs sh =
- (case sh C of
- None \<Rightarrow> (None, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, Calling C Cs)#frs, sh(C := Some (sblank P C, Prepared)))
- | Some (obj, iflag) \<Rightarrow>
- (case iflag of
- Done \<Rightarrow> (None, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, Called Cs)#frs, sh)
- | Processing \<Rightarrow> (None, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, Called Cs)#frs, sh)
- | Error \<Rightarrow> (None, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc,
- Throwing Cs (addr_of_sys_xcpt NoClassDefFoundError))#frs, sh)
- | Prepared \<Rightarrow>
- let sh' = sh(C:=Some(fst(the(sh C)), Processing));
- D = fst(the(class P C))
- in if C = Object
- then (None, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, Called (C#Cs))#frs, sh')
- else (None, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, Calling D (C#Cs))#frs, sh')
- )
- )"
-
-\<comment> \<open> single step of execution without error handling \<close>
-fun exec_step :: "[jvm_prog, heap, val list, val list,
- cname, mname, pc, init_call_status, frame list, sheap] \<Rightarrow> jvm_state"
-where
-"exec_step P h stk loc C M pc (Calling C' Cs) frs sh
- = exec_Calling C' Cs P h stk loc C M pc frs sh" |
-"exec_step P h stk loc C M pc (Called (C'#Cs)) frs sh
- = (None, h, create_init_frame P C'#(stk, loc, C, M, pc, Called Cs)#frs, sh)" |
-"exec_step P h stk loc C M pc (Throwing (C'#Cs) a) frs sh
- = (None, h, (stk,loc,C,M,pc,Throwing Cs a)#frs, sh(C':=Some(fst(the(sh C')), Error)))" |
-"exec_step P h stk loc C M pc (Throwing [] a) frs sh
- = (\<lfloor>a\<rfloor>, h, (stk,loc,C,M,pc,No_ics)#frs, sh)" |
-"exec_step P h stk loc C M pc ics frs sh
- = exec_instr (instrs_of P C M ! pc) P h stk loc C M pc ics frs sh"
-
-\<comment> \<open> execution including error handling \<close>
-fun exec :: "jvm_prog \<times> jvm_state \<Rightarrow> jvm_state option" \<comment> \<open>single step execution\<close> where
-"exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh) =
- (let (xp', h', frs', sh') = exec_step P h stk loc C M pc ics frs sh
- in case xp' of
- None \<Rightarrow> Some (None,h',frs',sh')
- | Some x \<Rightarrow> Some (find_handler P x h ((stk,loc,C,M,pc,ics)#frs) sh)
- )"
-| "exec _ = None"
-
-\<comment> \<open>relational view\<close>
-inductive_set
- exec_1 :: "jvm_prog \<Rightarrow> (jvm_state \<times> jvm_state) set"
- and exec_1' :: "jvm_prog \<Rightarrow> jvm_state \<Rightarrow> jvm_state \<Rightarrow> bool"
- ("_ \<turnstile>/ _ -jvm\<rightarrow>\<^sub>1/ _" [61,61,61] 60)
- for P :: jvm_prog
-where
- "P \<turnstile> \<sigma> -jvm\<rightarrow>\<^sub>1 \<sigma>' \<equiv> (\<sigma>,\<sigma>') \<in> exec_1 P"
-| exec_1I: "exec (P,\<sigma>) = Some \<sigma>' \<Longrightarrow> P \<turnstile> \<sigma> -jvm\<rightarrow>\<^sub>1 \<sigma>'"
-
-\<comment> \<open>reflexive transitive closure:\<close>
-definition exec_all :: "jvm_prog \<Rightarrow> jvm_state \<Rightarrow> jvm_state \<Rightarrow> bool"
- ("(_ \<turnstile>/ _ -jvm\<rightarrow>/ _)" [61,61,61]60) where
- exec_all_def1: "P \<turnstile> \<sigma> -jvm\<rightarrow> \<sigma>' \<longleftrightarrow> (\<sigma>,\<sigma>') \<in> (exec_1 P)\<^sup>*"
-
-notation (ASCII)
- exec_all ("_ |-/ _ -jvm->/ _" [61,61,61]60)
-
-
-lemma exec_1_eq:
- "exec_1 P = {(\<sigma>,\<sigma>'). exec (P,\<sigma>) = Some \<sigma>'}"
-(*<*)by (auto intro: exec_1I elim: exec_1.cases)(*>*)
-
-lemma exec_1_iff:
- "P \<turnstile> \<sigma> -jvm\<rightarrow>\<^sub>1 \<sigma>' = (exec (P,\<sigma>) = Some \<sigma>')"
-(*<*)by (simp add: exec_1_eq)(*>*)
-
-lemma exec_all_def:
- "P \<turnstile> \<sigma> -jvm\<rightarrow> \<sigma>' = ((\<sigma>,\<sigma>') \<in> {(\<sigma>,\<sigma>'). exec (P,\<sigma>) = Some \<sigma>'}\<^sup>*)"
-(*<*)by (simp add: exec_all_def1 exec_1_eq)(*>*)
-
-lemma jvm_refl[iff]: "P \<turnstile> \<sigma> -jvm\<rightarrow> \<sigma>"
-(*<*)by(simp add: exec_all_def)(*>*)
-
-lemma jvm_trans[trans]:
- "\<lbrakk> P \<turnstile> \<sigma> -jvm\<rightarrow> \<sigma>'; P \<turnstile> \<sigma>' -jvm\<rightarrow> \<sigma>'' \<rbrakk> \<Longrightarrow> P \<turnstile> \<sigma> -jvm\<rightarrow> \<sigma>''"
-(*<*)by(simp add: exec_all_def)(*>*)
-
-lemma jvm_one_step1[trans]:
- "\<lbrakk> P \<turnstile> \<sigma> -jvm\<rightarrow>\<^sub>1 \<sigma>'; P \<turnstile> \<sigma>' -jvm\<rightarrow> \<sigma>'' \<rbrakk> \<Longrightarrow> P \<turnstile> \<sigma> -jvm\<rightarrow> \<sigma>''"
-(*<*) by (simp add: exec_all_def1) (*>*)
-
-lemma jvm_one_step2[trans]:
- "\<lbrakk> P \<turnstile> \<sigma> -jvm\<rightarrow> \<sigma>'; P \<turnstile> \<sigma>' -jvm\<rightarrow>\<^sub>1 \<sigma>'' \<rbrakk> \<Longrightarrow> P \<turnstile> \<sigma> -jvm\<rightarrow> \<sigma>''"
-(*<*) by (simp add: exec_all_def1) (*>*)
-
-lemma exec_all_conf:
- "\<lbrakk> P \<turnstile> \<sigma> -jvm\<rightarrow> \<sigma>'; P \<turnstile> \<sigma> -jvm\<rightarrow> \<sigma>'' \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<sigma>' -jvm\<rightarrow> \<sigma>'' \<or> P \<turnstile> \<sigma>'' -jvm\<rightarrow> \<sigma>'"
-(*<*)by(simp add: exec_all_def single_valued_def single_valued_confluent)(*>*)
-
-lemma exec_1_exec_all_conf:
- "\<lbrakk> exec (P, \<sigma>) = Some \<sigma>'; P \<turnstile> \<sigma> -jvm\<rightarrow> \<sigma>''; \<sigma> \<noteq> \<sigma>'' \<rbrakk>
- \<Longrightarrow> P \<turnstile> \<sigma>' -jvm\<rightarrow> \<sigma>''"
- by(auto elim: converse_rtranclE simp: exec_1_eq exec_all_def)
-
-lemma exec_all_finalD: "P \<turnstile> (x, h, [], sh) -jvm\<rightarrow> \<sigma> \<Longrightarrow> \<sigma> = (x, h, [], sh)"
-(*<*)
-proof -
- assume "P \<turnstile> (x, h, [], sh) -jvm\<rightarrow> \<sigma>"
- then have "((x, h, [], sh), \<sigma>) \<in> {(\<sigma>, \<sigma>'). exec (P, \<sigma>) = \<lfloor>\<sigma>'\<rfloor>}\<^sup>*"
- by(simp only: exec_all_def)
- then show ?thesis proof(rule converse_rtranclE) qed simp+
-qed
-(*>*)
-
-lemma exec_all_deterministic:
- "\<lbrakk> P \<turnstile> \<sigma> -jvm\<rightarrow> (x,h,[],sh); P \<turnstile> \<sigma> -jvm\<rightarrow> \<sigma>' \<rbrakk> \<Longrightarrow> P \<turnstile> \<sigma>' -jvm\<rightarrow> (x,h,[],sh)"
-(*<*)
-proof -
- assume assms: "P \<turnstile> \<sigma> -jvm\<rightarrow> (x,h,[],sh)" "P \<turnstile> \<sigma> -jvm\<rightarrow> \<sigma>'"
- show ?thesis using exec_all_conf[OF assms]
- by(blast dest!: exec_all_finalD)
-qed
-(*>*)
-
-subsection "Preservation of preallocated"
-
-lemma exec_Calling_prealloc_pres:
-assumes "preallocated h"
- and "exec_Calling C Cs P h stk loc C\<^sub>0 M\<^sub>0 pc frs sh = (xp',h',frs',sh')"
-shows "preallocated h'"
-using assms
-proof(cases "sh C")
- case (Some a)
- then obtain sfs i where sfsi:"a = (sfs, i)" by(cases a)
- then show ?thesis using Some assms
- proof(cases i)
- case Prepared
- then show ?thesis using sfsi Some assms by(cases "method P C clinit", auto split: if_split_asm)
- next
- case Error
- then show ?thesis using sfsi Some assms by(cases "method P C clinit", auto)
- qed(auto)
-qed(auto)
-
-lemma exec_step_prealloc_pres:
-assumes pre: "preallocated h"
- and "exec_step P h stk loc C M pc ics frs sh = (xp',h',frs',sh')"
-shows "preallocated h'"
-proof(cases ics)
- case No_ics
- then show ?thesis using exec_instr_prealloc_pres assms by auto
-next
- case Calling
- then show ?thesis using exec_Calling_prealloc_pres assms by auto
-next
- case (Called Cs)
- then show ?thesis using exec_instr_prealloc_pres assms by(cases Cs, auto)
-next
- case (Throwing Cs a)
- then show ?thesis using assms by(cases Cs, auto)
-qed
-
-lemma exec_prealloc_pres:
-assumes pre: "preallocated h"
- and "exec (P, xp, h, frs, sh) = Some(xp',h',frs',sh')"
-shows "preallocated h'"
-using assms
-proof(cases "\<exists>x. xp = \<lfloor>x\<rfloor> \<or> frs = []")
- case False
- then obtain f1 frs1 where frs: "frs = f1#frs1" by(cases frs, simp+)
- then obtain stk1 loc1 C1 M1 pc1 ics1 where f1: "f1 = (stk1,loc1,C1,M1,pc1,ics1)" by(cases f1)
- let ?i = "instrs_of P C1 M1 ! pc1"
- obtain xp2 h2 frs2 sh2
- where exec_step: "exec_step P h stk1 loc1 C1 M1 pc1 ics1 frs1 sh = (xp2,h2,frs2,sh2)"
- by(cases "exec_step P h stk1 loc1 C1 M1 pc1 ics1 frs1 sh")
- then show ?thesis using exec_step_prealloc_pres[OF pre exec_step] f1 frs False assms
- proof(cases xp2)
- case (Some a)
- show ?thesis
- using find_handler_prealloc_pres[OF pre, where a=a]
- exec_step_prealloc_pres[OF pre]
- exec_step f1 frs Some False assms
- by(auto split: bool.splits init_call_status.splits list.splits)
- qed(auto split: init_call_status.splits)
-qed(auto)
-
-subsection "Start state"
-
-text \<open> The @{term Start} class is defined based on a given initial class
- and method. It has two methods: one that calls the initial method in the
- initial class, which is called by the starting program, and @{term clinit},
- as required for the class to be well-formed. \<close>
-definition start_method :: "cname \<Rightarrow> mname \<Rightarrow> jvm_method mdecl" where
-"start_method C M = (start_m, Static, [], Void, (1,0,[Invokestatic C M 0,Return],[]))"
-abbreviation start_clinit :: "jvm_method mdecl" where
-"start_clinit \<equiv> (clinit, Static, [], Void, (1,0,[Push Unit,Return],[]))"
-definition start_class :: "cname \<Rightarrow> mname \<Rightarrow> jvm_method cdecl" where
-"start_class C M = (Start, Object, [], [start_method C M, start_clinit])"
-
-text \<open>
- The start configuration of the JVM in program @{text P}:
- in the start heap, we call the ``start'' method of the
- ``Start''; this method performs @{text Invokestatic} on the
- class and method given to create the start program's @{term Start} class.
- This allows the initialization procedure to be called on the
- initial class in a natural way before the initial method is performed.
- There is no @{text this} pointer of the frame as @{term start} is @{term Static}.
- The start sheap has every class pre-prepared; this decision is not
- necessary.
- The starting program includes the added @{term Start} class, given a
- method @{text M} of class @{text C}, added to program @{text P}.
-\<close>
-definition start_state :: "jvm_prog \<Rightarrow> jvm_state" where
- "start_state P = (None, start_heap P, [([], [], Start, start_m, 0, No_ics)], start_sheap)"
-abbreviation start_prog :: "jvm_prog \<Rightarrow> cname \<Rightarrow> mname \<Rightarrow> jvm_prog" where
-"start_prog P C M \<equiv> start_class C M # P"
-
+
+(* Title: JinjaDCI/JVM/JVMExec.thy
+ Author: Cornelia Pusch, Gerwin Klein, Susannah Mansky
+ Copyright 1999 Technische Universitaet Muenchen, 2019-20
+
+ Based on the Jinja theory JVM/JVMExec.thy by Cornelia Pusch and Gerwin Klein
+*)
+
+section \<open> Program Execution in the JVM in full small step style \<close>
+
+theory JVMExec
+imports JVMExecInstr
+begin
+
+abbreviation
+ instrs_of :: "jvm_prog \<Rightarrow> cname \<Rightarrow> mname \<Rightarrow> instr list" where
+ "instrs_of P C M == fst(snd(snd(snd(snd(snd(snd(method P C M)))))))"
+
+fun curr_instr :: "jvm_prog \<Rightarrow> frame \<Rightarrow> instr" where
+"curr_instr P (stk,loc,C,M,pc,ics) = instrs_of P C M ! pc"
+
+\<comment> \<open> execution of single step of the initialization procedure \<close>
+fun exec_Calling :: "[cname, cname list, jvm_prog, heap, val list, val list,
+ cname, mname, pc, frame list, sheap] \<Rightarrow> jvm_state"
+where
+"exec_Calling C Cs P h stk loc C\<^sub>0 M\<^sub>0 pc frs sh =
+ (case sh C of
+ None \<Rightarrow> (None, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, Calling C Cs)#frs, sh(C := Some (sblank P C, Prepared)))
+ | Some (obj, iflag) \<Rightarrow>
+ (case iflag of
+ Done \<Rightarrow> (None, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, Called Cs)#frs, sh)
+ | Processing \<Rightarrow> (None, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, Called Cs)#frs, sh)
+ | Error \<Rightarrow> (None, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc,
+ Throwing Cs (addr_of_sys_xcpt NoClassDefFoundError))#frs, sh)
+ | Prepared \<Rightarrow>
+ let sh' = sh(C:=Some(fst(the(sh C)), Processing));
+ D = fst(the(class P C))
+ in if C = Object
+ then (None, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, Called (C#Cs))#frs, sh')
+ else (None, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, Calling D (C#Cs))#frs, sh')
+ )
+ )"
+
+\<comment> \<open> single step of execution without error handling \<close>
+fun exec_step :: "[jvm_prog, heap, val list, val list,
+ cname, mname, pc, init_call_status, frame list, sheap] \<Rightarrow> jvm_state"
+where
+"exec_step P h stk loc C M pc (Calling C' Cs) frs sh
+ = exec_Calling C' Cs P h stk loc C M pc frs sh" |
+"exec_step P h stk loc C M pc (Called (C'#Cs)) frs sh
+ = (None, h, create_init_frame P C'#(stk, loc, C, M, pc, Called Cs)#frs, sh)" |
+"exec_step P h stk loc C M pc (Throwing (C'#Cs) a) frs sh
+ = (None, h, (stk,loc,C,M,pc,Throwing Cs a)#frs, sh(C':=Some(fst(the(sh C')), Error)))" |
+"exec_step P h stk loc C M pc (Throwing [] a) frs sh
+ = (\<lfloor>a\<rfloor>, h, (stk,loc,C,M,pc,No_ics)#frs, sh)" |
+"exec_step P h stk loc C M pc ics frs sh
+ = exec_instr (instrs_of P C M ! pc) P h stk loc C M pc ics frs sh"
+
+\<comment> \<open> execution including error handling \<close>
+fun exec :: "jvm_prog \<times> jvm_state \<Rightarrow> jvm_state option" \<comment> \<open>single step execution\<close> where
+"exec (P, None, h, (stk,loc,C,M,pc,ics)#frs, sh) =
+ (let (xp', h', frs', sh') = exec_step P h stk loc C M pc ics frs sh
+ in case xp' of
+ None \<Rightarrow> Some (None,h',frs',sh')
+ | Some x \<Rightarrow> Some (find_handler P x h ((stk,loc,C,M,pc,ics)#frs) sh)
+ )"
+| "exec _ = None"
+
+\<comment> \<open>relational view\<close>
+inductive_set
+ exec_1 :: "jvm_prog \<Rightarrow> (jvm_state \<times> jvm_state) set"
+ and exec_1' :: "jvm_prog \<Rightarrow> jvm_state \<Rightarrow> jvm_state \<Rightarrow> bool"
+ ("_ \<turnstile>/ _ -jvm\<rightarrow>\<^sub>1/ _" [61,61,61] 60)
+ for P :: jvm_prog
+where
+ "P \<turnstile> \<sigma> -jvm\<rightarrow>\<^sub>1 \<sigma>' \<equiv> (\<sigma>,\<sigma>') \<in> exec_1 P"
+| exec_1I: "exec (P,\<sigma>) = Some \<sigma>' \<Longrightarrow> P \<turnstile> \<sigma> -jvm\<rightarrow>\<^sub>1 \<sigma>'"
+
+\<comment> \<open>reflexive transitive closure:\<close>
+definition exec_all :: "jvm_prog \<Rightarrow> jvm_state \<Rightarrow> jvm_state \<Rightarrow> bool"
+ ("(_ \<turnstile>/ _ -jvm\<rightarrow>/ _)" [61,61,61]60) where
+ exec_all_def1: "P \<turnstile> \<sigma> -jvm\<rightarrow> \<sigma>' \<longleftrightarrow> (\<sigma>,\<sigma>') \<in> (exec_1 P)\<^sup>*"
+
+notation (ASCII)
+ exec_all ("_ |-/ _ -jvm->/ _" [61,61,61]60)
+
+
+lemma exec_1_eq:
+ "exec_1 P = {(\<sigma>,\<sigma>'). exec (P,\<sigma>) = Some \<sigma>'}"
+(*<*)by (auto intro: exec_1I elim: exec_1.cases)(*>*)
+
+lemma exec_1_iff:
+ "P \<turnstile> \<sigma> -jvm\<rightarrow>\<^sub>1 \<sigma>' = (exec (P,\<sigma>) = Some \<sigma>')"
+(*<*)by (simp add: exec_1_eq)(*>*)
+
+lemma exec_all_def:
+ "P \<turnstile> \<sigma> -jvm\<rightarrow> \<sigma>' = ((\<sigma>,\<sigma>') \<in> {(\<sigma>,\<sigma>'). exec (P,\<sigma>) = Some \<sigma>'}\<^sup>*)"
+(*<*)by (simp add: exec_all_def1 exec_1_eq)(*>*)
+
+lemma jvm_refl[iff]: "P \<turnstile> \<sigma> -jvm\<rightarrow> \<sigma>"
+(*<*)by(simp add: exec_all_def)(*>*)
+
+lemma jvm_trans[trans]:
+ "\<lbrakk> P \<turnstile> \<sigma> -jvm\<rightarrow> \<sigma>'; P \<turnstile> \<sigma>' -jvm\<rightarrow> \<sigma>'' \<rbrakk> \<Longrightarrow> P \<turnstile> \<sigma> -jvm\<rightarrow> \<sigma>''"
+(*<*)by(simp add: exec_all_def)(*>*)
+
+lemma jvm_one_step1[trans]:
+ "\<lbrakk> P \<turnstile> \<sigma> -jvm\<rightarrow>\<^sub>1 \<sigma>'; P \<turnstile> \<sigma>' -jvm\<rightarrow> \<sigma>'' \<rbrakk> \<Longrightarrow> P \<turnstile> \<sigma> -jvm\<rightarrow> \<sigma>''"
+(*<*) by (simp add: exec_all_def1) (*>*)
+
+lemma jvm_one_step2[trans]:
+ "\<lbrakk> P \<turnstile> \<sigma> -jvm\<rightarrow> \<sigma>'; P \<turnstile> \<sigma>' -jvm\<rightarrow>\<^sub>1 \<sigma>'' \<rbrakk> \<Longrightarrow> P \<turnstile> \<sigma> -jvm\<rightarrow> \<sigma>''"
+(*<*) by (simp add: exec_all_def1) (*>*)
+
+lemma exec_all_conf:
+ "\<lbrakk> P \<turnstile> \<sigma> -jvm\<rightarrow> \<sigma>'; P \<turnstile> \<sigma> -jvm\<rightarrow> \<sigma>'' \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<sigma>' -jvm\<rightarrow> \<sigma>'' \<or> P \<turnstile> \<sigma>'' -jvm\<rightarrow> \<sigma>'"
+(*<*)by(simp add: exec_all_def single_valued_def single_valued_confluent)(*>*)
+
+lemma exec_1_exec_all_conf:
+ "\<lbrakk> exec (P, \<sigma>) = Some \<sigma>'; P \<turnstile> \<sigma> -jvm\<rightarrow> \<sigma>''; \<sigma> \<noteq> \<sigma>'' \<rbrakk>
+ \<Longrightarrow> P \<turnstile> \<sigma>' -jvm\<rightarrow> \<sigma>''"
+ by(auto elim: converse_rtranclE simp: exec_1_eq exec_all_def)
+
+lemma exec_all_finalD: "P \<turnstile> (x, h, [], sh) -jvm\<rightarrow> \<sigma> \<Longrightarrow> \<sigma> = (x, h, [], sh)"
+(*<*)
+proof -
+ assume "P \<turnstile> (x, h, [], sh) -jvm\<rightarrow> \<sigma>"
+ then have "((x, h, [], sh), \<sigma>) \<in> {(\<sigma>, \<sigma>'). exec (P, \<sigma>) = \<lfloor>\<sigma>'\<rfloor>}\<^sup>*"
+ by(simp only: exec_all_def)
+ then show ?thesis proof(rule converse_rtranclE) qed simp+
+qed
+(*>*)
+
+lemma exec_all_deterministic:
+ "\<lbrakk> P \<turnstile> \<sigma> -jvm\<rightarrow> (x,h,[],sh); P \<turnstile> \<sigma> -jvm\<rightarrow> \<sigma>' \<rbrakk> \<Longrightarrow> P \<turnstile> \<sigma>' -jvm\<rightarrow> (x,h,[],sh)"
+(*<*)
+proof -
+ assume assms: "P \<turnstile> \<sigma> -jvm\<rightarrow> (x,h,[],sh)" "P \<turnstile> \<sigma> -jvm\<rightarrow> \<sigma>'"
+ show ?thesis using exec_all_conf[OF assms]
+ by(blast dest!: exec_all_finalD)
+qed
+(*>*)
+
+subsection "Preservation of preallocated"
+
+lemma exec_Calling_prealloc_pres:
+assumes "preallocated h"
+ and "exec_Calling C Cs P h stk loc C\<^sub>0 M\<^sub>0 pc frs sh = (xp',h',frs',sh')"
+shows "preallocated h'"
+using assms
+proof(cases "sh C")
+ case (Some a)
+ then obtain sfs i where sfsi:"a = (sfs, i)" by(cases a)
+ then show ?thesis using Some assms
+ proof(cases i)
+ case Prepared
+ then show ?thesis using sfsi Some assms by(cases "method P C clinit", auto split: if_split_asm)
+ next
+ case Error
+ then show ?thesis using sfsi Some assms by(cases "method P C clinit", auto)
+ qed(auto)
+qed(auto)
+
+lemma exec_step_prealloc_pres:
+assumes pre: "preallocated h"
+ and "exec_step P h stk loc C M pc ics frs sh = (xp',h',frs',sh')"
+shows "preallocated h'"
+proof(cases ics)
+ case No_ics
+ then show ?thesis using exec_instr_prealloc_pres assms by auto
+next
+ case Calling
+ then show ?thesis using exec_Calling_prealloc_pres assms by auto
+next
+ case (Called Cs)
+ then show ?thesis using exec_instr_prealloc_pres assms by(cases Cs, auto)
+next
+ case (Throwing Cs a)
+ then show ?thesis using assms by(cases Cs, auto)
+qed
+
+lemma exec_prealloc_pres:
+assumes pre: "preallocated h"
+ and "exec (P, xp, h, frs, sh) = Some(xp',h',frs',sh')"
+shows "preallocated h'"
+using assms
+proof(cases "\<exists>x. xp = \<lfloor>x\<rfloor> \<or> frs = []")
+ case False
+ then obtain f1 frs1 where frs: "frs = f1#frs1" by(cases frs, simp+)
+ then obtain stk1 loc1 C1 M1 pc1 ics1 where f1: "f1 = (stk1,loc1,C1,M1,pc1,ics1)" by(cases f1)
+ let ?i = "instrs_of P C1 M1 ! pc1"
+ obtain xp2 h2 frs2 sh2
+ where exec_step: "exec_step P h stk1 loc1 C1 M1 pc1 ics1 frs1 sh = (xp2,h2,frs2,sh2)"
+ by(cases "exec_step P h stk1 loc1 C1 M1 pc1 ics1 frs1 sh")
+ then show ?thesis using exec_step_prealloc_pres[OF pre exec_step] f1 frs False assms
+ proof(cases xp2)
+ case (Some a)
+ show ?thesis
+ using find_handler_prealloc_pres[OF pre, where a=a]
+ exec_step_prealloc_pres[OF pre]
+ exec_step f1 frs Some False assms
+ by(auto split: bool.splits init_call_status.splits list.splits)
+ qed(auto split: init_call_status.splits)
+qed(auto)
+
+subsection "Start state"
+
+text \<open> The @{term Start} class is defined based on a given initial class
+ and method. It has two methods: one that calls the initial method in the
+ initial class, which is called by the starting program, and @{term clinit},
+ as required for the class to be well-formed. \<close>
+definition start_method :: "cname \<Rightarrow> mname \<Rightarrow> jvm_method mdecl" where
+"start_method C M = (start_m, Static, [], Void, (1,0,[Invokestatic C M 0,Return],[]))"
+abbreviation start_clinit :: "jvm_method mdecl" where
+"start_clinit \<equiv> (clinit, Static, [], Void, (1,0,[Push Unit,Return],[]))"
+definition start_class :: "cname \<Rightarrow> mname \<Rightarrow> jvm_method cdecl" where
+"start_class C M = (Start, Object, [], [start_method C M, start_clinit])"
+
+text \<open>
+ The start configuration of the JVM in program @{text P}:
+ in the start heap, we call the ``start'' method of the
+ ``Start''; this method performs @{text Invokestatic} on the
+ class and method given to create the start program's @{term Start} class.
+ This allows the initialization procedure to be called on the
+ initial class in a natural way before the initial method is performed.
+ There is no @{text this} pointer of the frame as @{term start} is @{term Static}.
+ The start sheap has every class pre-prepared; this decision is not
+ necessary.
+ The starting program includes the added @{term Start} class, given a
+ method @{text M} of class @{text C}, added to program @{text P}.
+\<close>
+definition start_state :: "jvm_prog \<Rightarrow> jvm_state" where
+ "start_state P = (None, start_heap P, [([], [], Start, start_m, 0, No_ics)], start_sheap)"
+abbreviation start_prog :: "jvm_prog \<Rightarrow> cname \<Rightarrow> mname \<Rightarrow> jvm_prog" where
+"start_prog P C M \<equiv> start_class C M # P"
+
end
\ No newline at end of file
diff --git a/thys/JinjaDCI/JVM/JVMExecInstr.thy b/thys/JinjaDCI/JVM/JVMExecInstr.thy
--- a/thys/JinjaDCI/JVM/JVMExecInstr.thy
+++ b/thys/JinjaDCI/JVM/JVMExecInstr.thy
@@ -1,393 +1,393 @@
-
-(* Title: JinjaDCI/JVM/JVMExecInstr.thy
- Author: Cornelia Pusch, Gerwin Klein, Susannah Mansky
- Copyright 1999 Technische Universitaet Muenchen
-
- Based on the Jinja theory JVM/JVMExecInstr.thy by Cornelia Pusch and Gerwin Klein
-*)
-
-section \<open> Program Execution in the JVM \<close>
-
-theory JVMExecInstr
-imports JVMInstructions JVMExceptions
-begin
-
- \<comment> \<open> frame calling the class initialization method for the given class
- in the given program \<close>
-fun create_init_frame :: "[jvm_prog, cname] \<Rightarrow> frame" where
-"create_init_frame P C =
- (let (D,b,Ts,T,(mxs,mxl\<^sub>0,ins,xt)) = method P C clinit
- in ([],(replicate mxl\<^sub>0 undefined),D,clinit,0,No_ics)
- )"
-
-primrec exec_instr :: "[instr, jvm_prog, heap, val list, val list,
- cname, mname, pc, init_call_status, frame list, sheap] \<Rightarrow> jvm_state"
-where
- exec_instr_Load:
-"exec_instr (Load n) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh =
- (None, h, ((loc ! n) # stk, loc, C\<^sub>0, M\<^sub>0, Suc pc, ics)#frs, sh)"
-
-| exec_instr_Store:
-"exec_instr (Store n) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh =
- (None, h, (tl stk, loc[n:=hd stk], C\<^sub>0, M\<^sub>0, Suc pc, ics)#frs, sh)"
-
-| exec_instr_Push:
-"exec_instr (Push v) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh =
- (None, h, (v # stk, loc, C\<^sub>0, M\<^sub>0, Suc pc, ics)#frs, sh)"
-
-| exec_instr_New:
-"exec_instr (New C) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh =
- (case (ics, sh C) of
- (Called Cs, _) \<Rightarrow>
- (case new_Addr h of
- None \<Rightarrow> (\<lfloor>addr_of_sys_xcpt OutOfMemory\<rfloor>, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, No_ics)#frs, sh)
- | Some a \<Rightarrow> (None, h(a\<mapsto>blank P C), (Addr a#stk, loc, C\<^sub>0, M\<^sub>0, Suc pc, No_ics)#frs, sh)
- )
- | (_, Some(obj, Done)) \<Rightarrow>
- (case new_Addr h of
- None \<Rightarrow> (\<lfloor>addr_of_sys_xcpt OutOfMemory\<rfloor>, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, ics)#frs, sh)
- | Some a \<Rightarrow> (None, h(a\<mapsto>blank P C), (Addr a#stk, loc, C\<^sub>0, M\<^sub>0, Suc pc, ics)#frs, sh)
- )
- | _ \<Rightarrow> (None, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, Calling C [])#frs, sh)
- )"
-
-| exec_instr_Getfield:
-"exec_instr (Getfield F C) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh =
- (let v = hd stk;
- (D,fs) = the(h(the_Addr v));
- (D',b,t) = field P C F;
- xp' = if v=Null then \<lfloor>addr_of_sys_xcpt NullPointer\<rfloor>
- else if \<not>(\<exists>t b. P \<turnstile> D has F,b:t in C)
- then \<lfloor>addr_of_sys_xcpt NoSuchFieldError\<rfloor>
- else case b of Static \<Rightarrow> \<lfloor>addr_of_sys_xcpt IncompatibleClassChangeError\<rfloor>
- | NonStatic \<Rightarrow> None
- in case xp' of None \<Rightarrow> (xp', h, (the(fs(F,C))#(tl stk), loc, C\<^sub>0, M\<^sub>0, pc+1, ics)#frs, sh)
- | Some x \<Rightarrow> (xp', h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, ics)#frs, sh))"
-
-| exec_instr_Getstatic:
-"exec_instr (Getstatic C F D) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh =
- (let (D',b,t) = field P D F;
- xp' = if \<not>(\<exists>t b. P \<turnstile> C has F,b:t in D)
- then \<lfloor>addr_of_sys_xcpt NoSuchFieldError\<rfloor>
- else case b of NonStatic \<Rightarrow> \<lfloor>addr_of_sys_xcpt IncompatibleClassChangeError\<rfloor>
- | Static \<Rightarrow> None
- in (case (xp', ics, sh D') of
- (Some a, _) \<Rightarrow> (xp', h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, ics)#frs, sh)
- | (_, Called Cs, _) \<Rightarrow> let (sfs, i) = the(sh D');
- v = the(sfs F)
- in (xp', h, (v#stk, loc, C\<^sub>0, M\<^sub>0, Suc pc, No_ics)#frs, sh)
- | (_, _, Some (sfs, Done)) \<Rightarrow> let v = the (sfs F)
- in (xp', h, (v#stk, loc, C\<^sub>0, M\<^sub>0, Suc pc, ics)#frs, sh)
- | _ \<Rightarrow> (xp', h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, Calling D' [])#frs, sh)
- )
- )"
-
-| exec_instr_Putfield:
-"exec_instr (Putfield F C) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh =
- (let v = hd stk;
- r = hd (tl stk);
- a = the_Addr r;
- (D,fs) = the (h a);
- (D',b,t) = field P C F;
- xp' = if r=Null then \<lfloor>addr_of_sys_xcpt NullPointer\<rfloor>
- else if \<not>(\<exists>t b. P \<turnstile> D has F,b:t in C)
- then \<lfloor>addr_of_sys_xcpt NoSuchFieldError\<rfloor>
- else case b of Static \<Rightarrow> \<lfloor>addr_of_sys_xcpt IncompatibleClassChangeError\<rfloor>
- | NonStatic \<Rightarrow> None;
- h' = h(a \<mapsto> (D, fs((F,C) \<mapsto> v)))
- in case xp' of None \<Rightarrow> (xp', h', (tl (tl stk), loc, C\<^sub>0, M\<^sub>0, pc+1, ics)#frs, sh)
- | Some x \<Rightarrow> (xp', h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, ics)#frs, sh)
- )"
-
-| exec_instr_Putstatic:
-"exec_instr (Putstatic C F D) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh =
- (let (D',b,t) = field P D F;
- xp' = if \<not>(\<exists>t b. P \<turnstile> C has F,b:t in D)
- then \<lfloor>addr_of_sys_xcpt NoSuchFieldError\<rfloor>
- else case b of NonStatic \<Rightarrow> \<lfloor>addr_of_sys_xcpt IncompatibleClassChangeError\<rfloor>
- | Static \<Rightarrow> None
- in (case (xp', ics, sh D') of
- (Some a, _) \<Rightarrow> (xp', h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, ics)#frs, sh)
- | (_, Called Cs, _)
- \<Rightarrow> let (sfs, i) = the(sh D')
- in (xp', h, (tl stk, loc, C\<^sub>0, M\<^sub>0, Suc pc, No_ics)#frs, sh(D':=Some ((sfs(F \<mapsto> hd stk)), i)))
- | (_, _, Some (sfs, Done))
- \<Rightarrow> (xp', h, (tl stk, loc, C\<^sub>0, M\<^sub>0, Suc pc, ics)#frs, sh(D':=Some ((sfs(F \<mapsto> hd stk)), Done)))
- | _ \<Rightarrow> (xp', h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, Calling D' [])#frs, sh)
- )
- )"
-
-| exec_instr_Checkcast:
-"exec_instr (Checkcast C) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh =
- (if cast_ok P C h (hd stk)
- then (None, h, (stk, loc, C\<^sub>0, M\<^sub>0, Suc pc, ics)#frs, sh)
- else (\<lfloor>addr_of_sys_xcpt ClassCast\<rfloor>, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, ics)#frs, sh)
- )"
-
-| exec_instr_Invoke:
-"exec_instr (Invoke M n) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh =
- (let ps = take n stk;
- r = stk!n;
- C = fst(the(h(the_Addr r)));
- (D,b,Ts,T,mxs,mxl\<^sub>0,ins,xt)= method P C M;
- xp' = if r=Null then \<lfloor>addr_of_sys_xcpt NullPointer\<rfloor>
- else if \<not>(\<exists>Ts T m D b. P \<turnstile> C sees M,b:Ts \<rightarrow> T = m in D)
- then \<lfloor>addr_of_sys_xcpt NoSuchMethodError\<rfloor>
- else case b of Static \<Rightarrow> \<lfloor>addr_of_sys_xcpt IncompatibleClassChangeError\<rfloor>
- | NonStatic \<Rightarrow> None;
- f' = ([],[r]@(rev ps)@(replicate mxl\<^sub>0 undefined),D,M,0,No_ics)
- in case xp' of None \<Rightarrow> (xp', h, f'#(stk, loc, C\<^sub>0, M\<^sub>0, pc, ics)#frs, sh)
- | Some a \<Rightarrow> (xp', h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, ics)#frs, sh)
- )"
-
-| exec_instr_Invokestatic:
-"exec_instr (Invokestatic C M n) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh =
- (let ps = take n stk;
- (D,b,Ts,T,mxs,mxl\<^sub>0,ins,xt)= method P C M;
- xp' = if \<not>(\<exists>Ts T m D b. P \<turnstile> C sees M,b:Ts \<rightarrow> T = m in D)
- then \<lfloor>addr_of_sys_xcpt NoSuchMethodError\<rfloor>
- else case b of NonStatic \<Rightarrow> \<lfloor>addr_of_sys_xcpt IncompatibleClassChangeError\<rfloor>
- | Static \<Rightarrow> None;
- f' = ([],(rev ps)@(replicate mxl\<^sub>0 undefined),D,M,0,No_ics)
- in (case (xp', ics, sh D) of
- (Some a, _) \<Rightarrow> (xp', h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, ics)#frs, sh)
- | (_, Called Cs, _) \<Rightarrow> (xp', h, f'#(stk, loc, C\<^sub>0, M\<^sub>0, pc, No_ics)#frs, sh)
- | (_, _, Some (sfs, Done)) \<Rightarrow> (xp', h, f'#(stk, loc, C\<^sub>0, M\<^sub>0, pc, ics)#frs, sh)
- | _ \<Rightarrow> (xp', h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, Calling D [])#frs, sh)
- )
- )"
-
-| exec_instr_Return:
- "exec_instr Return P h stk\<^sub>0 loc\<^sub>0 C\<^sub>0 M\<^sub>0 pc ics frs sh =
- (case frs of
- [] \<Rightarrow> let sh' = (case M\<^sub>0 = clinit of True \<Rightarrow> sh(C\<^sub>0:=Some(fst(the(sh C\<^sub>0)), Done))
- | _ \<Rightarrow> sh
- )
- in (None, h, [], sh')
- | (stk',loc',C',m',pc',ics')#frs'
- \<Rightarrow> let (D,b,Ts,T,(mxs,mxl\<^sub>0,ins,xt)) = method P C\<^sub>0 M\<^sub>0;
- offset = case b of NonStatic \<Rightarrow> 1 | Static \<Rightarrow> 0;
- (sh'', stk'', pc'') = (case M\<^sub>0 = clinit of True \<Rightarrow> (sh(C\<^sub>0:=Some(fst(the(sh C\<^sub>0)), Done)), stk', pc')
- | _ \<Rightarrow> (sh, (hd stk\<^sub>0)#(drop (length Ts + offset) stk'), Suc pc')
- )
- in (None, h, (stk'',loc',C',m',pc'',ics')#frs', sh'')
- )"
-
-| exec_instr_Pop:
-"exec_instr Pop P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh = (None, h, (tl stk, loc, C\<^sub>0, M\<^sub>0, Suc pc, ics)#frs, sh)"
-
-| exec_instr_IAdd:
-"exec_instr IAdd P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh =
- (None, h, (Intg (the_Intg (hd (tl stk)) + the_Intg (hd stk))#(tl (tl stk)), loc, C\<^sub>0, M\<^sub>0, Suc pc, ics)#frs, sh)"
-
-| exec_instr_IfFalse:
-"exec_instr (IfFalse i) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh =
- (let pc' = if hd stk = Bool False then nat(int pc+i) else pc+1
- in (None, h, (tl stk, loc, C\<^sub>0, M\<^sub>0, pc', ics)#frs, sh))"
-
-| exec_instr_CmpEq:
-"exec_instr CmpEq P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh =
- (None, h, (Bool (hd (tl stk) = hd stk) # tl (tl stk), loc, C\<^sub>0, M\<^sub>0, Suc pc, ics)#frs, sh)"
-
-| exec_instr_Goto:
-"exec_instr (Goto i) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh =
- (None, h, (stk, loc, C\<^sub>0, M\<^sub>0, nat(int pc+i), ics)#frs, sh)"
-
-| exec_instr_Throw:
-"exec_instr Throw P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh =
- (let xp' = if hd stk = Null then \<lfloor>addr_of_sys_xcpt NullPointer\<rfloor>
- else \<lfloor>the_Addr(hd stk)\<rfloor>
- in (xp', h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, ics)#frs, sh))"
-
-
-
-text "Given a preallocated heap, a thrown exception is either a system exception or
- thrown directly by @{term Throw}."
-lemma exec_instr_xcpts:
-assumes "\<sigma>' = exec_instr i P h stk loc C M pc ics' frs sh"
- and "fst \<sigma>' = Some a"
-shows "i = (JVMInstructions.Throw) \<or> a \<in> {a. \<exists>x \<in> sys_xcpts. a = addr_of_sys_xcpt x}"
-using assms
-proof(cases i)
- case (New C1) then show ?thesis using assms
- proof(cases "sh C1")
- case (Some a)
- then obtain sfs i where sfsi: "a = (sfs,i)" by(cases a)
- then show ?thesis using Some New assms
- proof(cases i) qed(cases ics', auto)+
- qed(cases ics', auto)
-next
- case (Getfield F1 C1)
- obtain D' b t where field: "field P C1 F1 = (D',b,t)" by(cases "field P C1 F1")
- obtain D fs where addr: "the (h (the_Addr (hd stk))) = (D,fs)" by(cases "the (h (the_Addr (hd stk)))")
- show ?thesis using addr field Getfield assms
- proof(cases "hd stk = Null")
- case nNull:False then show ?thesis using addr field Getfield assms
- proof(cases "\<nexists>t b. P \<turnstile> (cname_of h (the_Addr (hd stk))) has F1,b:t in C1")
- case exists:False show ?thesis
- proof(cases "fst(snd(field P C1 F1))")
- case Static
- then show ?thesis using exists nNull addr field Getfield assms
- by(auto simp: sys_xcpts_def split_beta split: staticb.splits)
- next
- case NonStatic
- then show ?thesis using exists nNull addr field Getfield assms
- by(auto simp: sys_xcpts_def split_beta split: staticb.splits)
- qed
- qed(auto)
- qed(auto)
-next
- case (Getstatic C1 F1 D1)
- obtain D' b t where field: "field P D1 F1 = (D',b,t)" by(cases "field P D1 F1")
- show ?thesis using field Getstatic assms
- proof(cases "\<nexists>t b. P \<turnstile> C1 has F1,b:t in D1")
- case exists:False then show ?thesis using field Getstatic assms
- proof(cases "fst(snd(field P D1 F1))")
- case Static
- then obtain sfs i where "the(sh D') = (sfs, i)" by(cases "the(sh D')")
- then show ?thesis using exists field Static Getstatic assms by(cases ics'; cases i, auto)
- qed(auto)
- qed(auto)
-next
- case (Putfield F1 C1)
- let ?r = "hd(tl stk)"
- obtain D' b t where field: "field P C1 F1 = (D',b,t)" by(cases "field P C1 F1")
- obtain D fs where addr: "the (h (the_Addr ?r)) = (D,fs)"
- by(cases "the (h (the_Addr ?r))")
- show ?thesis using addr field Putfield assms
- proof(cases "?r = Null")
- case nNull:False then show ?thesis using addr field Putfield assms
- proof(cases "\<nexists>t b. P \<turnstile> (cname_of h (the_Addr ?r)) has F1,b:t in C1")
- case exists:False show ?thesis
- proof(cases b)
- case Static
- then show ?thesis using exists nNull addr field Putfield assms
- by(auto simp: sys_xcpts_def split_beta split: staticb.splits)
- next
- case NonStatic
- then show ?thesis using exists nNull addr field Putfield assms
- by(auto simp: sys_xcpts_def split_beta split: staticb.splits)
- qed
- qed(auto)
- qed(auto)
-next
- case (Putstatic C1 F1 D1)
- obtain D' b t where field: "field P D1 F1 = (D',b,t)" by(cases "field P D1 F1")
- show ?thesis using field Putstatic assms
- proof(cases "\<nexists>t b. P \<turnstile> C1 has F1,b:t in D1")
- case exists:False then show ?thesis using field Putstatic assms
- proof(cases b)
- case Static
- then obtain sfs i where "the(sh D') = (sfs, i)" by(cases "the(sh D')")
- then show ?thesis using exists field Static Putstatic assms by(cases ics'; cases i, auto)
- qed(auto)
- qed(auto)
-next
- case (Checkcast C1) then show ?thesis using assms by(cases "cast_ok P C1 h (hd stk)", auto)
-next
- case (Invoke M n)
- let ?r = "stk!n"
- let ?C = "cname_of h (the_Addr ?r)"
- show ?thesis using Invoke assms
- proof(cases "?r = Null")
- case nNull:False then show ?thesis using Invoke assms
- proof(cases "\<not>(\<exists>Ts T m D b. P \<turnstile> ?C sees M,b:Ts \<rightarrow> T = m in D)")
- case exists:False then show ?thesis using nNull Invoke assms
- proof(cases "fst(snd(method P ?C M))")
- case Static
- then show ?thesis using exists nNull Invoke assms
- by(auto simp: sys_xcpts_def split_beta split: staticb.splits)
- next
- case NonStatic
- then show ?thesis using exists nNull Invoke assms
- by(auto simp: sys_xcpts_def split_beta split: staticb.splits)
- qed
- qed(auto)
- qed(auto)
-next
- case (Invokestatic C1 M n)
- show ?thesis using Invokestatic assms
- proof(cases "\<not>(\<exists>Ts T m D b. P \<turnstile> C1 sees M,b:Ts \<rightarrow> T = m in D)")
- case exists:False then show ?thesis using Invokestatic assms
- proof(cases "fst(snd(method P C1 M))")
- case Static
- then obtain sfs i where "the(sh (fst(method P C1 M))) = (sfs, i)"
- by(cases "the(sh (fst(method P C1 M)))")
- then show ?thesis using exists Static Invokestatic assms
- by(auto split: init_call_status.splits init_state.splits)
- qed(auto)
- qed(auto)
-next
- case Return then show ?thesis using assms
- proof(cases frs)
- case (Cons f frs') then show ?thesis using Return assms
- by(cases f, cases "method P C M", cases "M=clinit", auto)
- qed(auto)
-next
- case (IfFalse x17) then show ?thesis using assms
- proof(cases "hd stk")
- case (Bool b) then show ?thesis using IfFalse assms by(cases b, auto)
- qed(auto)
-qed(auto)
-
-lemma exec_instr_prealloc_pres:
-assumes "preallocated h"
- and "exec_instr i P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh = (xp',h',frs',sh')"
-shows "preallocated h'"
-using assms
-proof(cases i)
- case (New C1)
- then obtain sfs i where sfsi: "the(sh C1) = (sfs,i)" by(cases "the(sh C1)")
- then show ?thesis using preallocated_new[of h] New assms
- by(cases "blank P C1", auto dest: new_Addr_SomeD split: init_call_status.splits init_state.splits)
-next
- case (Getfield F1 C1) then show ?thesis using assms
- by(cases "the (h (the_Addr (hd stk)))", cases "field P C1 F1", auto)
-next
- case (Getstatic C1 F1 D1)
- then obtain sfs i where sfsi: "the(sh (fst (field P D1 F1))) = (sfs, i)"
- by(cases "the(sh (fst (field P D1 F1)))")
- then show ?thesis using Getstatic assms
- by(cases "field P D1 F1", auto split: init_call_status.splits init_state.splits)
-next
- case (Putfield F1 C1) then show ?thesis using preallocated_new preallocated_upd_obj assms
- by(cases "the (h (the_Addr (hd (tl stk))))", cases "field P C1 F1", auto, metis option.collapse)
-next
- case (Putstatic C1 F1 D1)
- then obtain sfs i where sfsi: "the(sh (fst (field P D1 F1))) = (sfs, i)"
- by(cases "the(sh (fst (field P D1 F1)))")
- then show ?thesis using Putstatic assms
- by(cases "field P D1 F1", auto split: init_call_status.splits init_state.splits)
-next
- case (Checkcast C1)
- then show ?thesis using assms
- proof(cases "hd stk = Null")
- case False then show ?thesis
- using Checkcast assms
- by(cases "P \<turnstile> cname_of h (the_Addr (hd stk)) \<preceq>\<^sup>* C1", auto simp: cast_ok_def)
- qed(simp add: cast_ok_def)
-next
- case (Invoke M n)
- then show ?thesis using assms by(cases "method P (cname_of h (the_Addr (stk ! n))) M", auto)
-next
- case (Invokestatic C1 M n)
- show ?thesis
- proof(cases "sh (fst (method P C1 M))")
- case None then show ?thesis using Invokestatic assms
- by(cases "method P C1 M", auto split: init_call_status.splits)
- next
- case (Some a)
- then obtain sfs i where sfsi: "a = (sfs, i)" by(cases a)
- then show ?thesis using Some Invokestatic assms
- proof(cases i) qed(cases "method P C1 M", auto split: init_call_status.splits)+
- qed
-next
- case Return
- then show ?thesis using assms by(cases "method P C\<^sub>0 M\<^sub>0", auto simp: split_beta split: list.splits)
-next
- case (IfFalse x17) then show ?thesis using assms by(auto split: val.splits bool.splits)
-next
- case Throw then show ?thesis using assms by(auto split: val.splits)
-qed(auto)
-
-end
+
+(* Title: JinjaDCI/JVM/JVMExecInstr.thy
+ Author: Cornelia Pusch, Gerwin Klein, Susannah Mansky
+ Copyright 1999 Technische Universitaet Muenchen
+
+ Based on the Jinja theory JVM/JVMExecInstr.thy by Cornelia Pusch and Gerwin Klein
+*)
+
+section \<open> Program Execution in the JVM \<close>
+
+theory JVMExecInstr
+imports JVMInstructions JVMExceptions
+begin
+
+ \<comment> \<open> frame calling the class initialization method for the given class
+ in the given program \<close>
+fun create_init_frame :: "[jvm_prog, cname] \<Rightarrow> frame" where
+"create_init_frame P C =
+ (let (D,b,Ts,T,(mxs,mxl\<^sub>0,ins,xt)) = method P C clinit
+ in ([],(replicate mxl\<^sub>0 undefined),D,clinit,0,No_ics)
+ )"
+
+primrec exec_instr :: "[instr, jvm_prog, heap, val list, val list,
+ cname, mname, pc, init_call_status, frame list, sheap] \<Rightarrow> jvm_state"
+where
+ exec_instr_Load:
+"exec_instr (Load n) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh =
+ (None, h, ((loc ! n) # stk, loc, C\<^sub>0, M\<^sub>0, Suc pc, ics)#frs, sh)"
+
+| exec_instr_Store:
+"exec_instr (Store n) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh =
+ (None, h, (tl stk, loc[n:=hd stk], C\<^sub>0, M\<^sub>0, Suc pc, ics)#frs, sh)"
+
+| exec_instr_Push:
+"exec_instr (Push v) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh =
+ (None, h, (v # stk, loc, C\<^sub>0, M\<^sub>0, Suc pc, ics)#frs, sh)"
+
+| exec_instr_New:
+"exec_instr (New C) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh =
+ (case (ics, sh C) of
+ (Called Cs, _) \<Rightarrow>
+ (case new_Addr h of
+ None \<Rightarrow> (\<lfloor>addr_of_sys_xcpt OutOfMemory\<rfloor>, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, No_ics)#frs, sh)
+ | Some a \<Rightarrow> (None, h(a\<mapsto>blank P C), (Addr a#stk, loc, C\<^sub>0, M\<^sub>0, Suc pc, No_ics)#frs, sh)
+ )
+ | (_, Some(obj, Done)) \<Rightarrow>
+ (case new_Addr h of
+ None \<Rightarrow> (\<lfloor>addr_of_sys_xcpt OutOfMemory\<rfloor>, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, ics)#frs, sh)
+ | Some a \<Rightarrow> (None, h(a\<mapsto>blank P C), (Addr a#stk, loc, C\<^sub>0, M\<^sub>0, Suc pc, ics)#frs, sh)
+ )
+ | _ \<Rightarrow> (None, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, Calling C [])#frs, sh)
+ )"
+
+| exec_instr_Getfield:
+"exec_instr (Getfield F C) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh =
+ (let v = hd stk;
+ (D,fs) = the(h(the_Addr v));
+ (D',b,t) = field P C F;
+ xp' = if v=Null then \<lfloor>addr_of_sys_xcpt NullPointer\<rfloor>
+ else if \<not>(\<exists>t b. P \<turnstile> D has F,b:t in C)
+ then \<lfloor>addr_of_sys_xcpt NoSuchFieldError\<rfloor>
+ else case b of Static \<Rightarrow> \<lfloor>addr_of_sys_xcpt IncompatibleClassChangeError\<rfloor>
+ | NonStatic \<Rightarrow> None
+ in case xp' of None \<Rightarrow> (xp', h, (the(fs(F,C))#(tl stk), loc, C\<^sub>0, M\<^sub>0, pc+1, ics)#frs, sh)
+ | Some x \<Rightarrow> (xp', h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, ics)#frs, sh))"
+
+| exec_instr_Getstatic:
+"exec_instr (Getstatic C F D) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh =
+ (let (D',b,t) = field P D F;
+ xp' = if \<not>(\<exists>t b. P \<turnstile> C has F,b:t in D)
+ then \<lfloor>addr_of_sys_xcpt NoSuchFieldError\<rfloor>
+ else case b of NonStatic \<Rightarrow> \<lfloor>addr_of_sys_xcpt IncompatibleClassChangeError\<rfloor>
+ | Static \<Rightarrow> None
+ in (case (xp', ics, sh D') of
+ (Some a, _) \<Rightarrow> (xp', h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, ics)#frs, sh)
+ | (_, Called Cs, _) \<Rightarrow> let (sfs, i) = the(sh D');
+ v = the(sfs F)
+ in (xp', h, (v#stk, loc, C\<^sub>0, M\<^sub>0, Suc pc, No_ics)#frs, sh)
+ | (_, _, Some (sfs, Done)) \<Rightarrow> let v = the (sfs F)
+ in (xp', h, (v#stk, loc, C\<^sub>0, M\<^sub>0, Suc pc, ics)#frs, sh)
+ | _ \<Rightarrow> (xp', h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, Calling D' [])#frs, sh)
+ )
+ )"
+
+| exec_instr_Putfield:
+"exec_instr (Putfield F C) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh =
+ (let v = hd stk;
+ r = hd (tl stk);
+ a = the_Addr r;
+ (D,fs) = the (h a);
+ (D',b,t) = field P C F;
+ xp' = if r=Null then \<lfloor>addr_of_sys_xcpt NullPointer\<rfloor>
+ else if \<not>(\<exists>t b. P \<turnstile> D has F,b:t in C)
+ then \<lfloor>addr_of_sys_xcpt NoSuchFieldError\<rfloor>
+ else case b of Static \<Rightarrow> \<lfloor>addr_of_sys_xcpt IncompatibleClassChangeError\<rfloor>
+ | NonStatic \<Rightarrow> None;
+ h' = h(a \<mapsto> (D, fs((F,C) \<mapsto> v)))
+ in case xp' of None \<Rightarrow> (xp', h', (tl (tl stk), loc, C\<^sub>0, M\<^sub>0, pc+1, ics)#frs, sh)
+ | Some x \<Rightarrow> (xp', h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, ics)#frs, sh)
+ )"
+
+| exec_instr_Putstatic:
+"exec_instr (Putstatic C F D) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh =
+ (let (D',b,t) = field P D F;
+ xp' = if \<not>(\<exists>t b. P \<turnstile> C has F,b:t in D)
+ then \<lfloor>addr_of_sys_xcpt NoSuchFieldError\<rfloor>
+ else case b of NonStatic \<Rightarrow> \<lfloor>addr_of_sys_xcpt IncompatibleClassChangeError\<rfloor>
+ | Static \<Rightarrow> None
+ in (case (xp', ics, sh D') of
+ (Some a, _) \<Rightarrow> (xp', h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, ics)#frs, sh)
+ | (_, Called Cs, _)
+ \<Rightarrow> let (sfs, i) = the(sh D')
+ in (xp', h, (tl stk, loc, C\<^sub>0, M\<^sub>0, Suc pc, No_ics)#frs, sh(D':=Some ((sfs(F \<mapsto> hd stk)), i)))
+ | (_, _, Some (sfs, Done))
+ \<Rightarrow> (xp', h, (tl stk, loc, C\<^sub>0, M\<^sub>0, Suc pc, ics)#frs, sh(D':=Some ((sfs(F \<mapsto> hd stk)), Done)))
+ | _ \<Rightarrow> (xp', h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, Calling D' [])#frs, sh)
+ )
+ )"
+
+| exec_instr_Checkcast:
+"exec_instr (Checkcast C) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh =
+ (if cast_ok P C h (hd stk)
+ then (None, h, (stk, loc, C\<^sub>0, M\<^sub>0, Suc pc, ics)#frs, sh)
+ else (\<lfloor>addr_of_sys_xcpt ClassCast\<rfloor>, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, ics)#frs, sh)
+ )"
+
+| exec_instr_Invoke:
+"exec_instr (Invoke M n) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh =
+ (let ps = take n stk;
+ r = stk!n;
+ C = fst(the(h(the_Addr r)));
+ (D,b,Ts,T,mxs,mxl\<^sub>0,ins,xt)= method P C M;
+ xp' = if r=Null then \<lfloor>addr_of_sys_xcpt NullPointer\<rfloor>
+ else if \<not>(\<exists>Ts T m D b. P \<turnstile> C sees M,b:Ts \<rightarrow> T = m in D)
+ then \<lfloor>addr_of_sys_xcpt NoSuchMethodError\<rfloor>
+ else case b of Static \<Rightarrow> \<lfloor>addr_of_sys_xcpt IncompatibleClassChangeError\<rfloor>
+ | NonStatic \<Rightarrow> None;
+ f' = ([],[r]@(rev ps)@(replicate mxl\<^sub>0 undefined),D,M,0,No_ics)
+ in case xp' of None \<Rightarrow> (xp', h, f'#(stk, loc, C\<^sub>0, M\<^sub>0, pc, ics)#frs, sh)
+ | Some a \<Rightarrow> (xp', h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, ics)#frs, sh)
+ )"
+
+| exec_instr_Invokestatic:
+"exec_instr (Invokestatic C M n) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh =
+ (let ps = take n stk;
+ (D,b,Ts,T,mxs,mxl\<^sub>0,ins,xt)= method P C M;
+ xp' = if \<not>(\<exists>Ts T m D b. P \<turnstile> C sees M,b:Ts \<rightarrow> T = m in D)
+ then \<lfloor>addr_of_sys_xcpt NoSuchMethodError\<rfloor>
+ else case b of NonStatic \<Rightarrow> \<lfloor>addr_of_sys_xcpt IncompatibleClassChangeError\<rfloor>
+ | Static \<Rightarrow> None;
+ f' = ([],(rev ps)@(replicate mxl\<^sub>0 undefined),D,M,0,No_ics)
+ in (case (xp', ics, sh D) of
+ (Some a, _) \<Rightarrow> (xp', h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, ics)#frs, sh)
+ | (_, Called Cs, _) \<Rightarrow> (xp', h, f'#(stk, loc, C\<^sub>0, M\<^sub>0, pc, No_ics)#frs, sh)
+ | (_, _, Some (sfs, Done)) \<Rightarrow> (xp', h, f'#(stk, loc, C\<^sub>0, M\<^sub>0, pc, ics)#frs, sh)
+ | _ \<Rightarrow> (xp', h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, Calling D [])#frs, sh)
+ )
+ )"
+
+| exec_instr_Return:
+ "exec_instr Return P h stk\<^sub>0 loc\<^sub>0 C\<^sub>0 M\<^sub>0 pc ics frs sh =
+ (case frs of
+ [] \<Rightarrow> let sh' = (case M\<^sub>0 = clinit of True \<Rightarrow> sh(C\<^sub>0:=Some(fst(the(sh C\<^sub>0)), Done))
+ | _ \<Rightarrow> sh
+ )
+ in (None, h, [], sh')
+ | (stk',loc',C',m',pc',ics')#frs'
+ \<Rightarrow> let (D,b,Ts,T,(mxs,mxl\<^sub>0,ins,xt)) = method P C\<^sub>0 M\<^sub>0;
+ offset = case b of NonStatic \<Rightarrow> 1 | Static \<Rightarrow> 0;
+ (sh'', stk'', pc'') = (case M\<^sub>0 = clinit of True \<Rightarrow> (sh(C\<^sub>0:=Some(fst(the(sh C\<^sub>0)), Done)), stk', pc')
+ | _ \<Rightarrow> (sh, (hd stk\<^sub>0)#(drop (length Ts + offset) stk'), Suc pc')
+ )
+ in (None, h, (stk'',loc',C',m',pc'',ics')#frs', sh'')
+ )"
+
+| exec_instr_Pop:
+"exec_instr Pop P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh = (None, h, (tl stk, loc, C\<^sub>0, M\<^sub>0, Suc pc, ics)#frs, sh)"
+
+| exec_instr_IAdd:
+"exec_instr IAdd P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh =
+ (None, h, (Intg (the_Intg (hd (tl stk)) + the_Intg (hd stk))#(tl (tl stk)), loc, C\<^sub>0, M\<^sub>0, Suc pc, ics)#frs, sh)"
+
+| exec_instr_IfFalse:
+"exec_instr (IfFalse i) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh =
+ (let pc' = if hd stk = Bool False then nat(int pc+i) else pc+1
+ in (None, h, (tl stk, loc, C\<^sub>0, M\<^sub>0, pc', ics)#frs, sh))"
+
+| exec_instr_CmpEq:
+"exec_instr CmpEq P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh =
+ (None, h, (Bool (hd (tl stk) = hd stk) # tl (tl stk), loc, C\<^sub>0, M\<^sub>0, Suc pc, ics)#frs, sh)"
+
+| exec_instr_Goto:
+"exec_instr (Goto i) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh =
+ (None, h, (stk, loc, C\<^sub>0, M\<^sub>0, nat(int pc+i), ics)#frs, sh)"
+
+| exec_instr_Throw:
+"exec_instr Throw P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh =
+ (let xp' = if hd stk = Null then \<lfloor>addr_of_sys_xcpt NullPointer\<rfloor>
+ else \<lfloor>the_Addr(hd stk)\<rfloor>
+ in (xp', h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, ics)#frs, sh))"
+
+
+
+text "Given a preallocated heap, a thrown exception is either a system exception or
+ thrown directly by @{term Throw}."
+lemma exec_instr_xcpts:
+assumes "\<sigma>' = exec_instr i P h stk loc C M pc ics' frs sh"
+ and "fst \<sigma>' = Some a"
+shows "i = (JVMInstructions.Throw) \<or> a \<in> {a. \<exists>x \<in> sys_xcpts. a = addr_of_sys_xcpt x}"
+using assms
+proof(cases i)
+ case (New C1) then show ?thesis using assms
+ proof(cases "sh C1")
+ case (Some a)
+ then obtain sfs i where sfsi: "a = (sfs,i)" by(cases a)
+ then show ?thesis using Some New assms
+ proof(cases i) qed(cases ics', auto)+
+ qed(cases ics', auto)
+next
+ case (Getfield F1 C1)
+ obtain D' b t where field: "field P C1 F1 = (D',b,t)" by(cases "field P C1 F1")
+ obtain D fs where addr: "the (h (the_Addr (hd stk))) = (D,fs)" by(cases "the (h (the_Addr (hd stk)))")
+ show ?thesis using addr field Getfield assms
+ proof(cases "hd stk = Null")
+ case nNull:False then show ?thesis using addr field Getfield assms
+ proof(cases "\<nexists>t b. P \<turnstile> (cname_of h (the_Addr (hd stk))) has F1,b:t in C1")
+ case exists:False show ?thesis
+ proof(cases "fst(snd(field P C1 F1))")
+ case Static
+ then show ?thesis using exists nNull addr field Getfield assms
+ by(auto simp: sys_xcpts_def split_beta split: staticb.splits)
+ next
+ case NonStatic
+ then show ?thesis using exists nNull addr field Getfield assms
+ by(auto simp: sys_xcpts_def split_beta split: staticb.splits)
+ qed
+ qed(auto)
+ qed(auto)
+next
+ case (Getstatic C1 F1 D1)
+ obtain D' b t where field: "field P D1 F1 = (D',b,t)" by(cases "field P D1 F1")
+ show ?thesis using field Getstatic assms
+ proof(cases "\<nexists>t b. P \<turnstile> C1 has F1,b:t in D1")
+ case exists:False then show ?thesis using field Getstatic assms
+ proof(cases "fst(snd(field P D1 F1))")
+ case Static
+ then obtain sfs i where "the(sh D') = (sfs, i)" by(cases "the(sh D')")
+ then show ?thesis using exists field Static Getstatic assms by(cases ics'; cases i, auto)
+ qed(auto)
+ qed(auto)
+next
+ case (Putfield F1 C1)
+ let ?r = "hd(tl stk)"
+ obtain D' b t where field: "field P C1 F1 = (D',b,t)" by(cases "field P C1 F1")
+ obtain D fs where addr: "the (h (the_Addr ?r)) = (D,fs)"
+ by(cases "the (h (the_Addr ?r))")
+ show ?thesis using addr field Putfield assms
+ proof(cases "?r = Null")
+ case nNull:False then show ?thesis using addr field Putfield assms
+ proof(cases "\<nexists>t b. P \<turnstile> (cname_of h (the_Addr ?r)) has F1,b:t in C1")
+ case exists:False show ?thesis
+ proof(cases b)
+ case Static
+ then show ?thesis using exists nNull addr field Putfield assms
+ by(auto simp: sys_xcpts_def split_beta split: staticb.splits)
+ next
+ case NonStatic
+ then show ?thesis using exists nNull addr field Putfield assms
+ by(auto simp: sys_xcpts_def split_beta split: staticb.splits)
+ qed
+ qed(auto)
+ qed(auto)
+next
+ case (Putstatic C1 F1 D1)
+ obtain D' b t where field: "field P D1 F1 = (D',b,t)" by(cases "field P D1 F1")
+ show ?thesis using field Putstatic assms
+ proof(cases "\<nexists>t b. P \<turnstile> C1 has F1,b:t in D1")
+ case exists:False then show ?thesis using field Putstatic assms
+ proof(cases b)
+ case Static
+ then obtain sfs i where "the(sh D') = (sfs, i)" by(cases "the(sh D')")
+ then show ?thesis using exists field Static Putstatic assms by(cases ics'; cases i, auto)
+ qed(auto)
+ qed(auto)
+next
+ case (Checkcast C1) then show ?thesis using assms by(cases "cast_ok P C1 h (hd stk)", auto)
+next
+ case (Invoke M n)
+ let ?r = "stk!n"
+ let ?C = "cname_of h (the_Addr ?r)"
+ show ?thesis using Invoke assms
+ proof(cases "?r = Null")
+ case nNull:False then show ?thesis using Invoke assms
+ proof(cases "\<not>(\<exists>Ts T m D b. P \<turnstile> ?C sees M,b:Ts \<rightarrow> T = m in D)")
+ case exists:False then show ?thesis using nNull Invoke assms
+ proof(cases "fst(snd(method P ?C M))")
+ case Static
+ then show ?thesis using exists nNull Invoke assms
+ by(auto simp: sys_xcpts_def split_beta split: staticb.splits)
+ next
+ case NonStatic
+ then show ?thesis using exists nNull Invoke assms
+ by(auto simp: sys_xcpts_def split_beta split: staticb.splits)
+ qed
+ qed(auto)
+ qed(auto)
+next
+ case (Invokestatic C1 M n)
+ show ?thesis using Invokestatic assms
+ proof(cases "\<not>(\<exists>Ts T m D b. P \<turnstile> C1 sees M,b:Ts \<rightarrow> T = m in D)")
+ case exists:False then show ?thesis using Invokestatic assms
+ proof(cases "fst(snd(method P C1 M))")
+ case Static
+ then obtain sfs i where "the(sh (fst(method P C1 M))) = (sfs, i)"
+ by(cases "the(sh (fst(method P C1 M)))")
+ then show ?thesis using exists Static Invokestatic assms
+ by(auto split: init_call_status.splits init_state.splits)
+ qed(auto)
+ qed(auto)
+next
+ case Return then show ?thesis using assms
+ proof(cases frs)
+ case (Cons f frs') then show ?thesis using Return assms
+ by(cases f, cases "method P C M", cases "M=clinit", auto)
+ qed(auto)
+next
+ case (IfFalse x17) then show ?thesis using assms
+ proof(cases "hd stk")
+ case (Bool b) then show ?thesis using IfFalse assms by(cases b, auto)
+ qed(auto)
+qed(auto)
+
+lemma exec_instr_prealloc_pres:
+assumes "preallocated h"
+ and "exec_instr i P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh = (xp',h',frs',sh')"
+shows "preallocated h'"
+using assms
+proof(cases i)
+ case (New C1)
+ then obtain sfs i where sfsi: "the(sh C1) = (sfs,i)" by(cases "the(sh C1)")
+ then show ?thesis using preallocated_new[of h] New assms
+ by(cases "blank P C1", auto dest: new_Addr_SomeD split: init_call_status.splits init_state.splits)
+next
+ case (Getfield F1 C1) then show ?thesis using assms
+ by(cases "the (h (the_Addr (hd stk)))", cases "field P C1 F1", auto)
+next
+ case (Getstatic C1 F1 D1)
+ then obtain sfs i where sfsi: "the(sh (fst (field P D1 F1))) = (sfs, i)"
+ by(cases "the(sh (fst (field P D1 F1)))")
+ then show ?thesis using Getstatic assms
+ by(cases "field P D1 F1", auto split: init_call_status.splits init_state.splits)
+next
+ case (Putfield F1 C1) then show ?thesis using preallocated_new preallocated_upd_obj assms
+ by(cases "the (h (the_Addr (hd (tl stk))))", cases "field P C1 F1", auto, metis option.collapse)
+next
+ case (Putstatic C1 F1 D1)
+ then obtain sfs i where sfsi: "the(sh (fst (field P D1 F1))) = (sfs, i)"
+ by(cases "the(sh (fst (field P D1 F1)))")
+ then show ?thesis using Putstatic assms
+ by(cases "field P D1 F1", auto split: init_call_status.splits init_state.splits)
+next
+ case (Checkcast C1)
+ then show ?thesis using assms
+ proof(cases "hd stk = Null")
+ case False then show ?thesis
+ using Checkcast assms
+ by(cases "P \<turnstile> cname_of h (the_Addr (hd stk)) \<preceq>\<^sup>* C1", auto simp: cast_ok_def)
+ qed(simp add: cast_ok_def)
+next
+ case (Invoke M n)
+ then show ?thesis using assms by(cases "method P (cname_of h (the_Addr (stk ! n))) M", auto)
+next
+ case (Invokestatic C1 M n)
+ show ?thesis
+ proof(cases "sh (fst (method P C1 M))")
+ case None then show ?thesis using Invokestatic assms
+ by(cases "method P C1 M", auto split: init_call_status.splits)
+ next
+ case (Some a)
+ then obtain sfs i where sfsi: "a = (sfs, i)" by(cases a)
+ then show ?thesis using Some Invokestatic assms
+ proof(cases i) qed(cases "method P C1 M", auto split: init_call_status.splits)+
+ qed
+next
+ case Return
+ then show ?thesis using assms by(cases "method P C\<^sub>0 M\<^sub>0", auto simp: split_beta split: list.splits)
+next
+ case (IfFalse x17) then show ?thesis using assms by(auto split: val.splits bool.splits)
+next
+ case Throw then show ?thesis using assms by(auto split: val.splits)
+qed(auto)
+
+end
diff --git a/thys/JinjaDCI/JVM/JVMInstructions.thy b/thys/JinjaDCI/JVM/JVMInstructions.thy
--- a/thys/JinjaDCI/JVM/JVMInstructions.thy
+++ b/thys/JinjaDCI/JVM/JVMInstructions.thy
@@ -1,64 +1,64 @@
-(* Title: JinjaDCI/JVM/JVMInstructions.thy
-
- Author: Gerwin Klein, Susannah Mansky
- Copyright 2000 Technische Universitaet Muenchen, 2019-20 UIUC
-
- Based on the Jinja theory JVM/JVMInstructions.thy by Gerwin Klein
-*)
-
-section \<open> Instructions of the JVM \<close>
-
-
-theory JVMInstructions imports JVMState begin
-
-
-datatype
- instr = Load nat \<comment> \<open>load from local variable\<close>
- | Store nat \<comment> \<open>store into local variable\<close>
- | Push val \<comment> \<open>push a value (constant)\<close>
- | New cname \<comment> \<open>create object\<close>
- | Getfield vname cname \<comment> \<open>Fetch field from object\<close>
- | Getstatic cname vname cname \<comment> \<open>Fetch static field from class\<close>
- | Putfield vname cname \<comment> \<open>Set field in object \<close>
- | Putstatic cname vname cname \<comment> \<open>Set static field in class\<close>
- | Checkcast cname \<comment> \<open>Check whether object is of given type\<close>
- | Invoke mname nat \<comment> \<open>inv. instance meth of an object\<close>
- | Invokestatic cname mname nat \<comment> \<open>inv. static method of a class\<close>
- | Return \<comment> \<open>return from method\<close>
- | Pop \<comment> \<open>pop top element from opstack\<close>
- | IAdd \<comment> \<open>integer addition\<close>
- | Goto int \<comment> \<open>goto relative address\<close>
- | CmpEq \<comment> \<open>equality comparison\<close>
- | IfFalse int \<comment> \<open>branch if top of stack false\<close>
- | Throw \<comment> \<open>throw top of stack as exception\<close>
-
-type_synonym
- bytecode = "instr list"
-
-type_synonym
- ex_entry = "pc \<times> pc \<times> cname \<times> pc \<times> nat"
- \<comment> \<open>start-pc, end-pc, exception type, handler-pc, remaining stack depth\<close>
-
-type_synonym
- ex_table = "ex_entry list"
-
-type_synonym
- jvm_method = "nat \<times> nat \<times> bytecode \<times> ex_table"
- \<comment> \<open>max stacksize\<close>
- \<comment> \<open>number of local variables. Add 1 + no. of parameters to get no. of registers\<close>
- \<comment> \<open>instruction sequence\<close>
- \<comment> \<open>exception handler table\<close>
-
-type_synonym
- jvm_prog = "jvm_method prog"
-
-(*<*)
-translations
- (type) "bytecode" <= (type) "instr list"
- (type) "ex_entry" <= (type) "nat \<times> nat \<times> char list \<times> nat \<times> nat"
- (type) "ex_table" <= (type) "ex_entry list"
- (type) "jvm_method" <= (type) "nat \<times> nat \<times> bytecode \<times> ex_table"
- (type) "jvm_prog" <= (type) "jvm_method prog"
-(*>*)
-
-end
+(* Title: JinjaDCI/JVM/JVMInstructions.thy
+
+ Author: Gerwin Klein, Susannah Mansky
+ Copyright 2000 Technische Universitaet Muenchen, 2019-20 UIUC
+
+ Based on the Jinja theory JVM/JVMInstructions.thy by Gerwin Klein
+*)
+
+section \<open> Instructions of the JVM \<close>
+
+
+theory JVMInstructions imports JVMState begin
+
+
+datatype
+ instr = Load nat \<comment> \<open>load from local variable\<close>
+ | Store nat \<comment> \<open>store into local variable\<close>
+ | Push val \<comment> \<open>push a value (constant)\<close>
+ | New cname \<comment> \<open>create object\<close>
+ | Getfield vname cname \<comment> \<open>Fetch field from object\<close>
+ | Getstatic cname vname cname \<comment> \<open>Fetch static field from class\<close>
+ | Putfield vname cname \<comment> \<open>Set field in object \<close>
+ | Putstatic cname vname cname \<comment> \<open>Set static field in class\<close>
+ | Checkcast cname \<comment> \<open>Check whether object is of given type\<close>
+ | Invoke mname nat \<comment> \<open>inv. instance meth of an object\<close>
+ | Invokestatic cname mname nat \<comment> \<open>inv. static method of a class\<close>
+ | Return \<comment> \<open>return from method\<close>
+ | Pop \<comment> \<open>pop top element from opstack\<close>
+ | IAdd \<comment> \<open>integer addition\<close>
+ | Goto int \<comment> \<open>goto relative address\<close>
+ | CmpEq \<comment> \<open>equality comparison\<close>
+ | IfFalse int \<comment> \<open>branch if top of stack false\<close>
+ | Throw \<comment> \<open>throw top of stack as exception\<close>
+
+type_synonym
+ bytecode = "instr list"
+
+type_synonym
+ ex_entry = "pc \<times> pc \<times> cname \<times> pc \<times> nat"
+ \<comment> \<open>start-pc, end-pc, exception type, handler-pc, remaining stack depth\<close>
+
+type_synonym
+ ex_table = "ex_entry list"
+
+type_synonym
+ jvm_method = "nat \<times> nat \<times> bytecode \<times> ex_table"
+ \<comment> \<open>max stacksize\<close>
+ \<comment> \<open>number of local variables. Add 1 + no. of parameters to get no. of registers\<close>
+ \<comment> \<open>instruction sequence\<close>
+ \<comment> \<open>exception handler table\<close>
+
+type_synonym
+ jvm_prog = "jvm_method prog"
+
+(*<*)
+translations
+ (type) "bytecode" <= (type) "instr list"
+ (type) "ex_entry" <= (type) "nat \<times> nat \<times> char list \<times> nat \<times> nat"
+ (type) "ex_table" <= (type) "ex_entry list"
+ (type) "jvm_method" <= (type) "nat \<times> nat \<times> bytecode \<times> ex_table"
+ (type) "jvm_prog" <= (type) "jvm_method prog"
+(*>*)
+
+end
diff --git a/thys/JinjaDCI/JVM/JVMState.thy b/thys/JinjaDCI/JVM/JVMState.thy
--- a/thys/JinjaDCI/JVM/JVMState.thy
+++ b/thys/JinjaDCI/JVM/JVMState.thy
@@ -1,84 +1,84 @@
-(* Title: Jinja/JVM/JVMState.thy
-
- Author: Cornelia Pusch, Gerwin Klein, Susannah Mansky
- Copyright 1999 Technische Universitaet Muenchen, 2019-20 UIUC
-
- Based on the Jinja theory JVM/JVMState.thy by Cornelia Pusch and Gerwin Klein
-*)
-
-chapter \<open> Jinja Virtual Machine \label{cha:jvm} \<close>
-
-section \<open> State of the JVM \<close>
-
-theory JVMState imports "../Common/Objects" begin
-
-
-type_synonym
- pc = nat
-
-abbreviation start_sheap :: "sheap"
-where "start_sheap \<equiv> (\<lambda>x. None)(Start \<mapsto> (Map.empty,Done))"
-
-definition start_sheap_preloaded :: "'m prog \<Rightarrow> sheap"
-where
- "start_sheap_preloaded P \<equiv> fold (\<lambda>(C,cl) f. f(C := Some (sblank P C, Prepared))) P (\<lambda>x. None)"
-
-subsection \<open> Frame Stack \<close>
-
-datatype init_call_status = No_ics | Calling cname "cname list"
- | Called "cname list" | Throwing "cname list" addr
- \<comment> \<open>@{text "No_ics"} = not currently calling or waiting for the result of an initialization procedure call\<close>
- \<comment> \<open>@{text "Calling C Cs"} = current instruction is calling for initialization of classes @{text "C#Cs"} (last class
- is the original) -- still collecting classes to be initialized, @{text "C"} most recently collected\<close>
- \<comment> \<open>@{text "Called Cs"} = current instruction called initialization and is waiting for the result
- -- now initializing classes in the list\<close>
- \<comment> \<open>@{text "Throwing Cs a"} = frame threw or was thrown an error causing erroneous end of initialization
- procedure for classes @{text "Cs"}\<close>
-
-type_synonym
- frame = "val list \<times> val list \<times> cname \<times> mname \<times> pc \<times> init_call_status"
- \<comment> \<open>operand stack\<close>
- \<comment> \<open>registers (including this pointer, method parameters, and local variables)\<close>
- \<comment> \<open>name of class where current method is defined\<close>
- \<comment> \<open>current method\<close>
- \<comment> \<open>program counter within frame\<close>
- \<comment> \<open>indicates frame's initialization call status\<close>
-
-translations
- (type) "frame" <= (type) "val list \<times> val list \<times> char list \<times> char list \<times> nat \<times> init_call_status"
-
-fun curr_stk :: "frame \<Rightarrow> val list" where
-"curr_stk (stk, loc, C, M, pc, ics) = stk"
-
-fun curr_class :: "frame \<Rightarrow> cname" where
-"curr_class (stk, loc, C, M, pc, ics) = C"
-
-fun curr_method :: "frame \<Rightarrow> mname" where
-"curr_method (stk, loc, C, M, pc, ics) = M"
-
-fun curr_pc :: "frame \<Rightarrow> nat" where
-"curr_pc (stk, loc, C, M, pc, ics) = pc"
-
-fun init_status :: "frame \<Rightarrow> init_call_status" where
- "init_status (stk, loc, C, M, pc, ics) = ics"
-
-fun ics_of :: "frame \<Rightarrow> init_call_status" where
- "ics_of fr = snd(snd(snd(snd(snd fr))))"
-
-
-subsection \<open> Runtime State \<close>
-
-type_synonym
- jvm_state = "addr option \<times> heap \<times> frame list \<times> sheap"
- \<comment> \<open>exception flag, heap, frames, static heap\<close>
-
-translations
- (type) "jvm_state" <= (type) "nat option \<times> heap \<times> frame list \<times> sheap"
-
-fun frames_of :: "jvm_state \<Rightarrow> frame list" where
-"frames_of (xp, h, frs, sh) = frs"
-
-abbreviation sheap :: "jvm_state \<Rightarrow> sheap" where
-"sheap js \<equiv> snd (snd (snd js))"
-
-end
+(* Title: Jinja/JVM/JVMState.thy
+
+ Author: Cornelia Pusch, Gerwin Klein, Susannah Mansky
+ Copyright 1999 Technische Universitaet Muenchen, 2019-20 UIUC
+
+ Based on the Jinja theory JVM/JVMState.thy by Cornelia Pusch and Gerwin Klein
+*)
+
+chapter \<open> Jinja Virtual Machine \label{cha:jvm} \<close>
+
+section \<open> State of the JVM \<close>
+
+theory JVMState imports "../Common/Objects" begin
+
+
+type_synonym
+ pc = nat
+
+abbreviation start_sheap :: "sheap"
+where "start_sheap \<equiv> (\<lambda>x. None)(Start \<mapsto> (Map.empty,Done))"
+
+definition start_sheap_preloaded :: "'m prog \<Rightarrow> sheap"
+where
+ "start_sheap_preloaded P \<equiv> fold (\<lambda>(C,cl) f. f(C := Some (sblank P C, Prepared))) P (\<lambda>x. None)"
+
+subsection \<open> Frame Stack \<close>
+
+datatype init_call_status = No_ics | Calling cname "cname list"
+ | Called "cname list" | Throwing "cname list" addr
+ \<comment> \<open>@{text "No_ics"} = not currently calling or waiting for the result of an initialization procedure call\<close>
+ \<comment> \<open>@{text "Calling C Cs"} = current instruction is calling for initialization of classes @{text "C#Cs"} (last class
+ is the original) -- still collecting classes to be initialized, @{text "C"} most recently collected\<close>
+ \<comment> \<open>@{text "Called Cs"} = current instruction called initialization and is waiting for the result
+ -- now initializing classes in the list\<close>
+ \<comment> \<open>@{text "Throwing Cs a"} = frame threw or was thrown an error causing erroneous end of initialization
+ procedure for classes @{text "Cs"}\<close>
+
+type_synonym
+ frame = "val list \<times> val list \<times> cname \<times> mname \<times> pc \<times> init_call_status"
+ \<comment> \<open>operand stack\<close>
+ \<comment> \<open>registers (including this pointer, method parameters, and local variables)\<close>
+ \<comment> \<open>name of class where current method is defined\<close>
+ \<comment> \<open>current method\<close>
+ \<comment> \<open>program counter within frame\<close>
+ \<comment> \<open>indicates frame's initialization call status\<close>
+
+translations
+ (type) "frame" <= (type) "val list \<times> val list \<times> char list \<times> char list \<times> nat \<times> init_call_status"
+
+fun curr_stk :: "frame \<Rightarrow> val list" where
+"curr_stk (stk, loc, C, M, pc, ics) = stk"
+
+fun curr_class :: "frame \<Rightarrow> cname" where
+"curr_class (stk, loc, C, M, pc, ics) = C"
+
+fun curr_method :: "frame \<Rightarrow> mname" where
+"curr_method (stk, loc, C, M, pc, ics) = M"
+
+fun curr_pc :: "frame \<Rightarrow> nat" where
+"curr_pc (stk, loc, C, M, pc, ics) = pc"
+
+fun init_status :: "frame \<Rightarrow> init_call_status" where
+ "init_status (stk, loc, C, M, pc, ics) = ics"
+
+fun ics_of :: "frame \<Rightarrow> init_call_status" where
+ "ics_of fr = snd(snd(snd(snd(snd fr))))"
+
+
+subsection \<open> Runtime State \<close>
+
+type_synonym
+ jvm_state = "addr option \<times> heap \<times> frame list \<times> sheap"
+ \<comment> \<open>exception flag, heap, frames, static heap\<close>
+
+translations
+ (type) "jvm_state" <= (type) "nat option \<times> heap \<times> frame list \<times> sheap"
+
+fun frames_of :: "jvm_state \<Rightarrow> frame list" where
+"frames_of (xp, h, frs, sh) = frs"
+
+abbreviation sheap :: "jvm_state \<Rightarrow> sheap" where
+"sheap js \<equiv> snd (snd (snd js))"
+
+end
diff --git a/thys/JinjaDCI/JinjaDCI.thy b/thys/JinjaDCI/JinjaDCI.thy
--- a/thys/JinjaDCI/JinjaDCI.thy
+++ b/thys/JinjaDCI/JinjaDCI.thy
@@ -1,12 +1,12 @@
-theory JinjaDCI
-imports
- "J/Equivalence"
- "J/Annotate"
- "JVM/JVMDefensive"
- "BV/BVExec"
- "BV/LBVJVM"
- "BV/BVNoTypeError"
- "Compiler/TypeComp"
-begin
-
-end
+theory JinjaDCI
+imports
+ "J/Equivalence"
+ "J/Annotate"
+ "JVM/JVMDefensive"
+ "BV/BVExec"
+ "BV/LBVJVM"
+ "BV/BVNoTypeError"
+ "Compiler/TypeComp"
+begin
+
+end
diff --git a/thys/MiniSail/BTVSubst.thy b/thys/MiniSail/BTVSubst.thy
--- a/thys/MiniSail/BTVSubst.thy
+++ b/thys/MiniSail/BTVSubst.thy
@@ -1,1550 +1,1550 @@
-(*<*)
-theory BTVSubst
- imports Syntax
-begin
- (*>*)
-chapter \<open>Basic Type Variable Substitution\<close>
-
-section \<open>Class\<close>
-
-class has_subst_b = fs +
- fixes subst_b :: "'a::fs \<Rightarrow> bv \<Rightarrow> b \<Rightarrow> 'a::fs" ("_[_::=_]\<^sub>b" [1000,50,50] 1000)
-
-assumes fresh_subst_if: "j \<sharp> (t[i::=x]\<^sub>b ) \<longleftrightarrow> (atom i \<sharp> t \<and> j \<sharp> t) \<or> (j \<sharp> x \<and> (j \<sharp> t \<or> j = atom i))"
- and forget_subst[simp]: "atom a \<sharp> tm \<Longrightarrow> tm[a::=x]\<^sub>b = tm"
- and subst_id[simp]: "tm[a::=(B_var a)]\<^sub>b = tm"
- and eqvt[simp,eqvt]: "(p::perm) \<bullet> (subst_b t1 x1 v ) = (subst_b (p \<bullet>t1) (p \<bullet>x1) (p \<bullet>v) )"
- and flip_subst[simp]: "atom bv \<sharp> c \<Longrightarrow> ((bv \<leftrightarrow> z) \<bullet> c) = c[z::=B_var bv]\<^sub>b"
- and flip_subst_subst[simp]: "atom bv \<sharp> c \<Longrightarrow> ((bv \<leftrightarrow> z) \<bullet> c)[bv::=v]\<^sub>b = c[z::=v]\<^sub>b"
-begin
-
-lemmas flip_subst_b = flip_subst_subst
-
-lemma subst_b_simple_commute:
- fixes x::bv
- assumes "atom x \<sharp> c"
- shows "(c[z::=B_var x]\<^sub>b)[x::=b]\<^sub>b = c[z::=b]\<^sub>b"
-proof -
- have "(c[z::=B_var x]\<^sub>b)[x::=b]\<^sub>b = (( x \<leftrightarrow> z) \<bullet> c)[x::=b]\<^sub>b" using flip_subst assms by simp
- thus ?thesis using flip_subst_subst assms by simp
-qed
-
-lemma subst_b_flip_eq_one:
- fixes z1::bv and z2::bv and x1::bv and x2::bv
- assumes "[[atom z1]]lst. c1 = [[atom z2]]lst. c2"
- and "atom x1 \<sharp> (z1,z2,c1,c2)"
- shows "(c1[z1::=B_var x1]\<^sub>b) = (c2[z2::=B_var x1]\<^sub>b)"
-proof -
- have "(c1[z1::=B_var x1]\<^sub>b) = (x1 \<leftrightarrow> z1) \<bullet> c1" using assms flip_subst by auto
- moreover have "(c2[z2::=B_var x1]\<^sub>b) = (x1 \<leftrightarrow> z2) \<bullet> c2" using assms flip_subst by auto
- ultimately show ?thesis using Abs1_eq_iff_all(3)[of z1 c1 z2 c2 z1] assms
- by (metis Abs1_eq_iff_fresh(3) flip_commute)
-qed
-
-lemma subst_b_flip_eq_two:
- fixes z1::bv and z2::bv and x1::bv and x2::bv
- assumes "[[atom z1]]lst. c1 = [[atom z2]]lst. c2"
- shows "(c1[z1::=b]\<^sub>b) = (c2[z2::=b]\<^sub>b)"
-proof -
- obtain x::bv where *:"atom x \<sharp> (z1,z2,c1,c2)" using obtain_fresh by metis
- hence "(c1[z1::=B_var x]\<^sub>b) = (c2[z2::=B_var x]\<^sub>b)" using subst_b_flip_eq_one[OF assms, of x] by metis
- hence "(c1[z1::=B_var x]\<^sub>b)[x::=b]\<^sub>b = (c2[z2::=B_var x]\<^sub>b)[x::=b]\<^sub>b" by auto
- thus ?thesis using subst_b_simple_commute * fresh_prod4 by metis
-qed
-
-lemma subst_b_fresh_x:
- fixes tm::"'a::fs" and x::x
- shows "atom x \<sharp> tm = atom x \<sharp> tm[bv::=b']\<^sub>b"
- using fresh_subst_if[of "atom x" tm bv b'] using x_fresh_b by auto
-
-lemma subst_b_x_flip[simp]:
- fixes x'::x and x::x and bv::bv
- shows "((x' \<leftrightarrow> x) \<bullet> tm)[bv::=b']\<^sub>b = (x' \<leftrightarrow> x) \<bullet> (tm[bv::=b']\<^sub>b)"
-proof -
- have "(x' \<leftrightarrow> x) \<bullet> bv = bv" using pure_supp flip_fresh_fresh by force
- moreover have "(x' \<leftrightarrow> x) \<bullet> b' = b'" using x_fresh_b flip_fresh_fresh by auto
- ultimately show ?thesis using eqvt by simp
-qed
-
-end
-
-section \<open>Base Type\<close>
-
-nominal_function subst_bb :: "b \<Rightarrow> bv \<Rightarrow> b \<Rightarrow> b" where
- "subst_bb (B_var bv2) bv1 b = (if bv1 = bv2 then b else (B_var bv2))"
-| "subst_bb B_int bv1 b = B_int"
-| "subst_bb B_bool bv1 b = B_bool"
-| "subst_bb (B_id s) bv1 b = B_id s "
-| "subst_bb (B_pair b1 b2) bv1 b = B_pair (subst_bb b1 bv1 b) (subst_bb b2 bv1 b)"
-| "subst_bb B_unit bv1 b = B_unit "
-| "subst_bb B_bitvec bv1 b = B_bitvec "
-| "subst_bb (B_app s b2) bv1 b = B_app s (subst_bb b2 bv1 b)"
- apply (simp add: eqvt_def subst_bb_graph_aux_def )
- apply (simp add: eqvt_def subst_bb_graph_aux_def )
- by (auto,meson b.strong_exhaust)
-nominal_termination (eqvt) by lexicographic_order
-
-abbreviation
- subst_bb_abbrev :: "b \<Rightarrow> bv \<Rightarrow> b \<Rightarrow> b" ("_[_::=_]\<^sub>b\<^sub>b" [1000,50,50] 1000)
- where
- "b[bv::=b']\<^sub>b\<^sub>b \<equiv> subst_bb b bv b' "
-
-instantiation b :: has_subst_b
-begin
-definition "subst_b = subst_bb"
-
-instance proof
- fix j::atom and i::bv and x::b and t::b
- show "j \<sharp> subst_b t i x = (atom i \<sharp> t \<and> j \<sharp> t \<or> j \<sharp> x \<and> (j \<sharp> t \<or> j = atom i))"
- proof (induct t rule: b.induct)
- case (B_id x)
- then show ?case using subst_bb.simps fresh_def pure_fresh subst_b_b_def by auto
- next
- case (B_var x)
- then show ?case using subst_bb.simps fresh_def pure_fresh subst_b_b_def by auto
- next
- case (B_app x1 x2)
- then show ?case using subst_bb.simps fresh_def pure_fresh subst_b_b_def by auto
- qed(auto simp add: subst_bb.simps fresh_def pure_fresh subst_b_b_def)+
-
- fix a::bv and tm::b and x::b
- show "atom a \<sharp> tm \<Longrightarrow> tm[a::=x]\<^sub>b = tm"
- by (induct tm rule: b.induct, auto simp add: fresh_at_base subst_bb.simps subst_b_b_def)
-
- fix a::bv and tm::b
- show "subst_b tm a (B_var a) = tm" using subst_bb.simps subst_b_b_def
- by (induct tm rule: b.induct, auto simp add: fresh_at_base subst_bb.simps subst_b_b_def)
-
- fix p::perm and x1::bv and v::b and t1::b
- show "p \<bullet> subst_b t1 x1 v = subst_b (p \<bullet> t1) (p \<bullet> x1) (p \<bullet> v)"
- by (induct tm rule: b.induct, auto simp add: fresh_at_base subst_bb.simps subst_b_b_def)
-
- fix bv::bv and c::b and z::bv
- show "atom bv \<sharp> c \<Longrightarrow> ((bv \<leftrightarrow> z) \<bullet> c) = c[z::=B_var bv]\<^sub>b"
- by (induct c rule: b.induct, (auto simp add: fresh_at_base subst_bb.simps subst_b_b_def permute_pure pure_supp )+)
-
- fix bv::bv and c::b and z::bv and v::b
- show "atom bv \<sharp> c \<Longrightarrow> ((bv \<leftrightarrow> z) \<bullet> c)[bv::=v]\<^sub>b = c[z::=v]\<^sub>b"
- by (induct c rule: b.induct, (auto simp add: fresh_at_base subst_bb.simps subst_b_b_def permute_pure pure_supp )+)
-qed
-end
-
-lemma subst_bb_inject:
- assumes "b1 = b2[bv::=b]\<^sub>b\<^sub>b" and "b2 \<noteq> B_var bv"
- shows
- "b1 = B_int \<Longrightarrow> b2 = B_int" and
- "b1 = B_bool \<Longrightarrow> b2 = B_bool" and
- "b1 = B_unit \<Longrightarrow> b2 = B_unit" and
- "b1 = B_bitvec \<Longrightarrow> b2 = B_bitvec" and
- "b1 = B_pair b11 b12 \<Longrightarrow> (\<exists>b11' b12' . b11 = b11'[bv::=b]\<^sub>b\<^sub>b \<and> b12 = b12'[bv::=b]\<^sub>b\<^sub>b \<and> b2 = B_pair b11' b12')" and
- "b1 = B_var bv' \<Longrightarrow> b2 = B_var bv'" and
- "b1 = B_id tyid \<Longrightarrow> b2 = B_id tyid" and
- "b1 = B_app tyid b11 \<Longrightarrow> (\<exists>b11'. b11= b11'[bv::=b]\<^sub>b\<^sub>b \<and> b2 = B_app tyid b11')"
- using assms by(nominal_induct b2 rule:b.strong_induct,auto+)
-
-lemma flip_b_subst4:
- fixes b1::b and bv1::bv and c::bv and b::b
- assumes "atom c \<sharp> (b1,bv1)"
- shows "b1[bv1::=b]\<^sub>b\<^sub>b = ((bv1 \<leftrightarrow> c) \<bullet> b1)[c ::= b]\<^sub>b\<^sub>b"
- using assms proof(nominal_induct b1 rule: b.strong_induct)
- case B_int
- then show ?case using subst_bb.simps b.perm_simps by auto
-next
- case B_bool
- then show ?case using subst_bb.simps b.perm_simps by auto
-next
- case (B_id x)
- hence "atom bv1 \<sharp> x \<and> atom c \<sharp> x" using fresh_def pure_supp by auto
- hence "((bv1 \<leftrightarrow> c) \<bullet> B_id x) = B_id x" using fresh_Pair b.fresh(3) flip_fresh_fresh b.perm_simps fresh_def pure_supp by metis
- then show ?case using subst_bb.simps by simp
-next
- case (B_pair x1 x2)
- hence "x1[bv1::=b]\<^sub>b\<^sub>b = ((bv1 \<leftrightarrow> c) \<bullet> x1)[c::=b]\<^sub>b\<^sub>b" using b.perm_simps(4) b.fresh(4) fresh_Pair by metis
- moreover have "x2[bv1::=b]\<^sub>b\<^sub>b = ((bv1 \<leftrightarrow> c) \<bullet> x2)[c::=b]\<^sub>b\<^sub>b" using b.perm_simps(4) b.fresh(4) fresh_Pair B_pair by metis
- ultimately show ?case using subst_bb.simps(5) b.perm_simps(4) b.fresh(4) fresh_Pair by auto
-next
- case B_unit
- then show ?case using subst_bb.simps b.perm_simps by auto
-next
- case B_bitvec
- then show ?case using subst_bb.simps b.perm_simps by auto
-next
- case (B_var x)
- then show ?case proof(cases "x=bv1")
- case True
- then show ?thesis using B_var subst_bb.simps b.perm_simps by simp
- next
- case False
- moreover have "x\<noteq>c" using B_var b.fresh fresh_def supp_at_base fresh_Pair by fastforce
- ultimately show ?thesis using B_var subst_bb.simps(1) b.perm_simps(7) by simp
- qed
-next
- case (B_app x1 x2)
- hence "x2[bv1::=b]\<^sub>b\<^sub>b = ((bv1 \<leftrightarrow> c) \<bullet> x2)[c::=b]\<^sub>b\<^sub>b" using b.perm_simps b.fresh fresh_Pair by metis
- thus ?case using subst_bb.simps b.perm_simps b.fresh fresh_Pair B_app
- by (simp add: permute_pure)
-qed
-
-lemma subst_bb_flip_sym:
- fixes b1::b and b2::b
- assumes "atom c \<sharp> b" and "atom c \<sharp> (bv1,bv2, b1, b2)" and "(bv1 \<leftrightarrow> c) \<bullet> b1 = (bv2 \<leftrightarrow> c) \<bullet> b2"
- shows "b1[bv1::=b]\<^sub>b\<^sub>b = b2[bv2::=b]\<^sub>b\<^sub>b"
- using assms flip_b_subst4[of c b1 bv1 b] flip_b_subst4[of c b2 bv2 b] fresh_prod4 fresh_Pair by simp
-
-section \<open>Value\<close>
-
-nominal_function subst_vb :: "v \<Rightarrow> bv \<Rightarrow> b \<Rightarrow> v" where
- "subst_vb (V_lit l) x v = V_lit l"
-| "subst_vb (V_var y) x v = V_var y"
-| "subst_vb (V_cons tyid c v') x v = V_cons tyid c (subst_vb v' x v)"
-| "subst_vb (V_consp tyid c b v') x v = V_consp tyid c (subst_bb b x v) (subst_vb v' x v)"
-| "subst_vb (V_pair v1 v2) x v = V_pair (subst_vb v1 x v ) (subst_vb v2 x v )"
- apply (simp add: eqvt_def subst_vb_graph_aux_def)
- apply auto
- using v.strong_exhaust by meson
-nominal_termination (eqvt) by lexicographic_order
-
-abbreviation
- subst_vb_abbrev :: "v \<Rightarrow> bv \<Rightarrow> b \<Rightarrow> v" ("_[_::=_]\<^sub>v\<^sub>b" [1000,50,50] 500)
- where
- "e[bv::=b]\<^sub>v\<^sub>b \<equiv> subst_vb e bv b"
-
-instantiation v :: has_subst_b
-begin
-definition "subst_b = subst_vb"
-
-instance proof
- fix j::atom and i::bv and x::b and t::v
- show "j \<sharp> subst_b t i x = (atom i \<sharp> t \<and> j \<sharp> t \<or> j \<sharp> x \<and> (j \<sharp> t \<or> j = atom i))"
- proof (induct t rule: v.induct)
- case (V_lit l)
- have "j \<sharp> subst_b (V_lit l) i x = j \<sharp> (V_lit l)" using subst_vb.simps fresh_def pure_fresh
- subst_b_v_def v.supp v.fresh has_subst_b_class.fresh_subst_if subst_b_b_def subst_b_v_def by auto
- also have "... = True" using fresh_at_base v.fresh l.fresh supp_l_empty fresh_def by metis
- moreover have "(atom i \<sharp> (V_lit l) \<and> j \<sharp> (V_lit l) \<or> j \<sharp> x \<and> (j \<sharp> (V_lit l) \<or> j = atom i)) = True" using fresh_at_base v.fresh l.fresh supp_l_empty fresh_def by metis
- ultimately show ?case by simp
- next
- case (V_var y)
- then show ?case using subst_b_v_def subst_vb.simps pure_fresh by force
- next
- case (V_pair x1a x2a)
- then show ?case using subst_b_v_def subst_vb.simps by auto
- next
- case (V_cons x1a x2a x3)
- then show ?case using V_cons subst_b_v_def subst_vb.simps pure_fresh by force
- next
- case (V_consp x1a x2a x3 x4)
- then show ?case using subst_b_v_def subst_vb.simps pure_fresh has_subst_b_class.fresh_subst_if subst_b_b_def subst_b_v_def by fastforce
- qed
-
- fix a::bv and tm::v and x::b
- show "atom a \<sharp> tm \<Longrightarrow> subst_b tm a x = tm"
- apply(induct tm rule: v.induct)
- apply(auto simp add: fresh_at_base subst_vb.simps subst_b_v_def)
- using has_subst_b_class.fresh_subst_if subst_b_b_def e.fresh
- using has_subst_b_class.forget_subst by fastforce
-
- fix a::bv and tm::v
- show "subst_b tm a (B_var a) = tm" using subst_bb.simps subst_b_b_def
- apply (induct tm rule: v.induct)
- apply(auto simp add: fresh_at_base subst_vb.simps subst_b_v_def)
- using has_subst_b_class.fresh_subst_if subst_b_b_def e.fresh
- using has_subst_b_class.subst_id by metis
-
- fix p::perm and x1::bv and v::b and t1::v
- show "p \<bullet> subst_b t1 x1 v = subst_b (p \<bullet> t1) (p \<bullet> x1) (p \<bullet> v)"
- apply(induct tm rule: v.induct)
- apply(auto simp add: fresh_at_base subst_bb.simps subst_b_b_def )
- using has_subst_b_class.eqvt subst_b_b_def e.fresh
- using has_subst_b_class.eqvt
- by (simp add: subst_b_v_def)+
-
- fix bv::bv and c::v and z::bv
- show "atom bv \<sharp> c \<Longrightarrow> ((bv \<leftrightarrow> z) \<bullet> c) = c[z::=B_var bv]\<^sub>b"
- apply (induct c rule: v.induct, (auto simp add: fresh_at_base subst_vb.simps subst_b_v_def permute_pure pure_supp )+)
- apply (metis flip_fresh_fresh flip_l_eq permute_flip_cancel2)
- using fresh_at_base flip_fresh_fresh[of bv x z]
- apply (simp add: flip_fresh_fresh)
- using subst_b_b_def by argo
-
- fix bv::bv and c::v and z::bv and v::b
- show "atom bv \<sharp> c \<Longrightarrow> ((bv \<leftrightarrow> z) \<bullet> c)[bv::=v]\<^sub>b = c[z::=v]\<^sub>b"
- apply (induct c rule: v.induct, (auto simp add: fresh_at_base subst_vb.simps subst_b_v_def permute_pure pure_supp )+)
- apply (metis flip_fresh_fresh flip_l_eq permute_flip_cancel2)
- using fresh_at_base flip_fresh_fresh[of bv x z]
- apply (simp add: flip_fresh_fresh)
- using subst_b_b_def flip_subst_subst by fastforce
-
-qed
-end
-
-section \<open>Constraints Expressions\<close>
-
-nominal_function subst_ceb :: "ce \<Rightarrow> bv \<Rightarrow> b \<Rightarrow> ce" where
- "subst_ceb ( (CE_val v') ) bv b = ( CE_val (subst_vb v' bv b) )"
-| "subst_ceb ( (CE_op opp v1 v2) ) bv b = ( (CE_op opp (subst_ceb v1 bv b)(subst_ceb v2 bv b)) )"
-| "subst_ceb ( (CE_fst v')) bv b = CE_fst (subst_ceb v' bv b)"
-| "subst_ceb ( (CE_snd v')) bv b = CE_snd (subst_ceb v' bv b)"
-| "subst_ceb ( (CE_len v')) bv b = CE_len (subst_ceb v' bv b)"
-| "subst_ceb ( CE_concat v1 v2) bv b = CE_concat (subst_ceb v1 bv b) (subst_ceb v2 bv b)"
- apply (simp add: eqvt_def subst_ceb_graph_aux_def)
- apply auto
- by (meson ce.strong_exhaust)
-nominal_termination (eqvt) by lexicographic_order
-
-abbreviation
- subst_ceb_abbrev :: "ce \<Rightarrow> bv \<Rightarrow> b \<Rightarrow> ce" ("_[_::=_]\<^sub>c\<^sub>e\<^sub>b" [1000,50,50] 500)
- where
- "ce[bv::=b]\<^sub>c\<^sub>e\<^sub>b \<equiv> subst_ceb ce bv b"
-
-instantiation ce :: has_subst_b
-begin
-definition "subst_b = subst_ceb"
-
-instance proof
- fix j::atom and i::bv and x::b and t::ce
- show "j \<sharp> subst_b t i x = (atom i \<sharp> t \<and> j \<sharp> t \<or> j \<sharp> x \<and> (j \<sharp> t \<or> j = atom i))"
- proof (induct t rule: ce.induct)
- case (CE_val v)
- then show ?case using subst_ceb.simps fresh_def pure_fresh subst_b_ce_def ce.supp v.supp ce.fresh has_subst_b_class.fresh_subst_if subst_b_b_def subst_b_v_def
- by metis
- next
- case (CE_op opp v1 v2)
-
- have "(j \<sharp> v1[i::=x]\<^sub>c\<^sub>e\<^sub>b \<and> j \<sharp> v2[i::=x]\<^sub>c\<^sub>e\<^sub>b) = ((atom i \<sharp> v1 \<and> atom i \<sharp> v2) \<and> j \<sharp> v1 \<and> j \<sharp> v2 \<or> j \<sharp> x \<and> (j \<sharp> v1 \<and> j \<sharp> v2 \<or> j = atom i))"
- using has_subst_b_class.fresh_subst_if subst_b_v_def
- using CE_op.hyps(1) CE_op.hyps(2) subst_b_ce_def by auto
-
- thus ?case unfolding subst_ceb.simps subst_b_ce_def ce.fresh
- using fresh_def pure_fresh opp.fresh subst_b_v_def opp.exhaust fresh_e_opp_all
- by (metis (full_types))
- next
- case (CE_concat x1a x2)
- then show ?case using subst_ceb.simps subst_b_ce_def e.supp v.supp has_subst_b_class.fresh_subst_if subst_b_v_def ce.fresh by force
- next
- case (CE_fst x)
- then show ?case using subst_ceb.simps subst_b_ce_def e.supp v.supp has_subst_b_class.fresh_subst_if subst_b_v_def ce.fresh by metis
-
- next
- case (CE_snd x)
- then show ?case using subst_ceb.simps subst_b_ce_def e.supp v.supp has_subst_b_class.fresh_subst_if subst_b_v_def ce.fresh by metis
- next
- case (CE_len x)
- then show ?case using subst_ceb.simps subst_b_ce_def e.supp v.supp has_subst_b_class.fresh_subst_if subst_b_v_def ce.fresh by metis
- qed
-
- fix a::bv and tm::ce and x::b
- show "atom a \<sharp> tm \<Longrightarrow> subst_b tm a x = tm"
- apply(induct tm rule: ce.induct)
- apply( auto simp add: fresh_at_base subst_ceb.simps subst_b_ce_def)
- using has_subst_b_class.fresh_subst_if subst_b_b_def e.fresh
- using has_subst_b_class.forget_subst subst_b_v_def apply metis+
- done
-
- fix a::bv and tm::ce
- show "subst_b tm a (B_var a) = tm" using subst_bb.simps subst_b_b_def
- apply (induct tm rule: ce.induct)
- apply(auto simp add: fresh_at_base subst_ceb.simps subst_b_ce_def)
- using has_subst_b_class.fresh_subst_if subst_b_b_def e.fresh
- using has_subst_b_class.subst_id subst_b_v_def apply metis+
- done
-
- fix p::perm and x1::bv and v::b and t1::ce
- show "p \<bullet> subst_b t1 x1 v = subst_b (p \<bullet> t1) (p \<bullet> x1) (p \<bullet> v)"
- apply(induct tm rule: ce.induct)
- apply( auto simp add: fresh_at_base subst_bb.simps subst_b_b_def )
- using has_subst_b_class.eqvt subst_b_b_def ce.fresh
- using has_subst_b_class.eqvt
- by (simp add: subst_b_ce_def)+
-
- fix bv::bv and c::ce and z::bv
- show "atom bv \<sharp> c \<Longrightarrow> ((bv \<leftrightarrow> z) \<bullet> c) = c[z::=B_var bv]\<^sub>b"
- apply (induct c rule: ce.induct, (auto simp add: fresh_at_base subst_ceb.simps subst_b_ce_def permute_pure pure_supp )+)
- using flip_fresh_fresh flip_l_eq permute_flip_cancel2 has_subst_b_class.flip_subst subst_b_v_def apply metis
- using flip_fresh_fresh flip_l_eq permute_flip_cancel2 has_subst_b_class.flip_subst subst_b_v_def
- by (simp add: flip_fresh_fresh fresh_opp_all)
-
- fix bv::bv and c::ce and z::bv and v::b
- show "atom bv \<sharp> c \<Longrightarrow> ((bv \<leftrightarrow> z) \<bullet> c)[bv::=v]\<^sub>b = c[z::=v]\<^sub>b"
- proof (induct c rule: ce.induct)
- case (CE_val x)
- then show ?case using flip_subst_subst subst_b_v_def subst_ceb.simps using subst_b_ce_def by fastforce
- next
- case (CE_op x1a x2 x3)
- then show ?case unfolding subst_ceb.simps subst_b_ce_def ce.perm_simps using flip_subst_subst subst_b_v_def opp.perm_simps opp.strong_exhaust
- by (metis (full_types) ce.fresh(2))
- next
- case (CE_concat x1a x2)
- then show ?case using flip_subst_subst subst_b_v_def subst_ceb.simps using subst_b_ce_def by fastforce
- next
- case (CE_fst x)
- then show ?case using flip_subst_subst subst_b_v_def subst_ceb.simps using subst_b_ce_def by fastforce
- next
- case (CE_snd x)
- then show ?case using flip_subst_subst subst_b_v_def subst_ceb.simps using subst_b_ce_def by fastforce
- next
- case (CE_len x)
- then show ?case using flip_subst_subst subst_b_v_def subst_ceb.simps using subst_b_ce_def by fastforce
- qed
-qed
-end
-
-section \<open>Constraints\<close>
-
-nominal_function subst_cb :: "c \<Rightarrow> bv \<Rightarrow> b \<Rightarrow> c" where
- "subst_cb (C_true) x v = C_true"
-| "subst_cb (C_false) x v = C_false"
-| "subst_cb (C_conj c1 c2) x v = C_conj (subst_cb c1 x v ) (subst_cb c2 x v )"
-| "subst_cb (C_disj c1 c2) x v = C_disj (subst_cb c1 x v ) (subst_cb c2 x v )"
-| "subst_cb (C_imp c1 c2) x v = C_imp (subst_cb c1 x v ) (subst_cb c2 x v )"
-| "subst_cb (C_eq e1 e2) x v = C_eq (subst_ceb e1 x v ) (subst_ceb e2 x v )"
-| "subst_cb (C_not c) x v = C_not (subst_cb c x v )"
- apply (simp add: eqvt_def subst_cb_graph_aux_def)
- apply auto
- using c.strong_exhaust apply metis
- done
-nominal_termination (eqvt) by lexicographic_order
-
-abbreviation
- subst_cb_abbrev :: "c \<Rightarrow> bv \<Rightarrow> b \<Rightarrow> c" ("_[_::=_]\<^sub>c\<^sub>b" [1000,50,50] 500)
- where
- "c[bv::=b]\<^sub>c\<^sub>b \<equiv> subst_cb c bv b"
-
-instantiation c :: has_subst_b
-begin
-definition "subst_b = subst_cb"
-
-instance proof
- fix j::atom and i::bv and x::b and t::c
- show "j \<sharp> subst_b t i x = (atom i \<sharp> t \<and> j \<sharp> t \<or> j \<sharp> x \<and> (j \<sharp> t \<or> j = atom i))"
- by (induct t rule: c.induct, unfold subst_cb.simps subst_b_c_def c.fresh,
- (metis has_subst_b_class.fresh_subst_if subst_b_ce_def c.fresh)+
- )
-
- fix a::bv and tm::c and x::b
- show "atom a \<sharp> tm \<Longrightarrow> subst_b tm a x = tm"
- by(induct tm rule: c.induct, unfold subst_cb.simps subst_b_c_def c.fresh,
- (metis has_subst_b_class.forget_subst subst_b_ce_def)+)
-
- fix a::bv and tm::c
- show "subst_b tm a (B_var a) = tm" using subst_bb.simps subst_b_c_def
- by(induct tm rule: c.induct, unfold subst_cb.simps subst_b_c_def c.fresh,
- (metis has_subst_b_class.subst_id subst_b_ce_def)+)
-
- fix p::perm and x1::bv and v::b and t1::c
- show "p \<bullet> subst_b t1 x1 v = subst_b (p \<bullet> t1) (p \<bullet> x1) (p \<bullet> v)"
- apply(induct tm rule: c.induct,unfold subst_cb.simps subst_b_c_def c.fresh)
- by( auto simp add: fresh_at_base subst_bb.simps subst_b_b_def )
-
- fix bv::bv and c::c and z::bv
- show "atom bv \<sharp> c \<Longrightarrow> ((bv \<leftrightarrow> z) \<bullet> c) = c[z::=B_var bv]\<^sub>b"
- apply (induct c rule: c.induct, (auto simp add: fresh_at_base subst_cb.simps subst_b_c_def permute_pure pure_supp )+)
- using flip_fresh_fresh flip_l_eq permute_flip_cancel2 has_subst_b_class.flip_subst subst_b_ce_def apply metis
- using flip_fresh_fresh flip_l_eq permute_flip_cancel2 has_subst_b_class.flip_subst subst_b_ce_def
- apply (metis opp.perm_simps(2) opp.strong_exhaust)+
- done
-
- fix bv::bv and c::c and z::bv and v::b
- show "atom bv \<sharp> c \<Longrightarrow> ((bv \<leftrightarrow> z) \<bullet> c)[bv::=v]\<^sub>b = c[z::=v]\<^sub>b"
- apply (induct c rule: c.induct, (auto simp add: fresh_at_base subst_cb.simps subst_b_c_def permute_pure pure_supp )+)
- using flip_fresh_fresh flip_l_eq permute_flip_cancel2 has_subst_b_class.flip_subst subst_b_ce_def
- using flip_subst_subst apply fastforce
- using flip_fresh_fresh flip_l_eq permute_flip_cancel2 has_subst_b_class.flip_subst subst_b_ce_def
- opp.perm_simps(2) opp.strong_exhaust
- proof -
- fix x1a :: ce and x2 :: ce
- assume a1: "atom bv \<sharp> x2"
- then have "((bv \<leftrightarrow> z) \<bullet> x2)[bv::=v]\<^sub>b = x2[z::=v]\<^sub>b"
- by (metis flip_subst_subst) (* 0.0 ms *)
- then show "x2[z::=B_var bv]\<^sub>b[bv::=v]\<^sub>c\<^sub>e\<^sub>b = x2[z::=v]\<^sub>c\<^sub>e\<^sub>b"
- using a1 by (simp add: subst_b_ce_def) (* 0.0 ms *)
- qed
-
-qed
-end
-
-section \<open>Types\<close>
-
-nominal_function subst_tb :: "\<tau> \<Rightarrow> bv \<Rightarrow> b \<Rightarrow> \<tau>" where
- "subst_tb (\<lbrace> z : b2 | c \<rbrace>) bv1 b1 = \<lbrace> z : b2[bv1::=b1]\<^sub>b\<^sub>b | c[bv1::=b1]\<^sub>c\<^sub>b \<rbrace>"
-proof(goal_cases)
- case 1
- then show ?case using eqvt_def subst_tb_graph_aux_def by force
-next
- case (2 x y)
- then show ?case by auto
-next
- case (3 P x)
- then show ?case using eqvt_def subst_tb_graph_aux_def \<tau>.strong_exhaust
- by (metis b_of.cases prod_cases3)
-next
- case (4 z' b2' c' bv1' b1' z b2 c bv1 b1)
- show ?case unfolding \<tau>.eq_iff proof
- have *:"[[atom z']]lst. c' = [[atom z]]lst. c" using \<tau>.eq_iff 4 by auto
- show "[[atom z']]lst. c'[bv1'::=b1']\<^sub>c\<^sub>b = [[atom z]]lst. c[bv1::=b1]\<^sub>c\<^sub>b"
- proof(subst Abs1_eq_iff_all(3),rule,rule,rule)
- fix ca::x
- assume "atom ca \<sharp> z" and 1:"atom ca \<sharp> (z', z, c'[bv1'::=b1']\<^sub>c\<^sub>b, c[bv1::=b1]\<^sub>c\<^sub>b)"
- hence 2:"atom ca \<sharp> (c',c)" using fresh_subst_if subst_b_c_def fresh_Pair fresh_prod4 fresh_at_base subst_b_fresh_x by metis
- hence "(z' \<leftrightarrow> ca) \<bullet> c' = (z \<leftrightarrow> ca) \<bullet> c" using 1 2 * Abs1_eq_iff_all(3) by auto
- hence "((z' \<leftrightarrow> ca) \<bullet> c')[bv1'::=b1']\<^sub>c\<^sub>b = ((z \<leftrightarrow> ca) \<bullet> c)[bv1'::=b1']\<^sub>c\<^sub>b" by auto
- hence "(z' \<leftrightarrow> ca) \<bullet> c'[(z' \<leftrightarrow> ca) \<bullet> bv1'::=(z' \<leftrightarrow> ca) \<bullet> b1']\<^sub>c\<^sub>b = (z \<leftrightarrow> ca) \<bullet> c[(z \<leftrightarrow> ca) \<bullet> bv1'::=(z \<leftrightarrow> ca) \<bullet> b1']\<^sub>c\<^sub>b" by auto
- thus "(z' \<leftrightarrow> ca) \<bullet> c'[bv1'::=b1']\<^sub>c\<^sub>b = (z \<leftrightarrow> ca) \<bullet> c[bv1::=b1]\<^sub>c\<^sub>b" using 4 flip_x_b_cancel by simp
- qed
- show "b2'[bv1'::=b1']\<^sub>b\<^sub>b = b2[bv1::=b1]\<^sub>b\<^sub>b" using 4 by simp
- qed
-qed
-
-nominal_termination (eqvt) by lexicographic_order
-
-abbreviation
- subst_tb_abbrev :: "\<tau> \<Rightarrow> bv \<Rightarrow> b \<Rightarrow> \<tau>" ("_[_::=_]\<^sub>\<tau>\<^sub>b" [1000,50,50] 1000)
- where
- "t[bv::=b']\<^sub>\<tau>\<^sub>b \<equiv> subst_tb t bv b' "
-
-instantiation \<tau> :: has_subst_b
-begin
-definition "subst_b = subst_tb"
-
-instance proof
- fix j::atom and i::bv and x::b and t::\<tau>
- show "j \<sharp> subst_b t i x = (atom i \<sharp> t \<and> j \<sharp> t \<or> j \<sharp> x \<and> (j \<sharp> t \<or> j = atom i))"
- proof (nominal_induct t avoiding: i x j rule: \<tau>.strong_induct)
- case (T_refined_type z b c)
- then show ?case
- unfolding subst_b_\<tau>_def subst_tb.simps \<tau>.fresh
- using fresh_subst_if[of j b i x ] subst_b_b_def subst_b_c_def
- by (metis has_subst_b_class.fresh_subst_if list.distinct(1) list.set_cases not_self_fresh set_ConsD)
- qed
-
- fix a::bv and tm::\<tau> and x::b
- show "atom a \<sharp> tm \<Longrightarrow> subst_b tm a x = tm"
- proof (nominal_induct tm avoiding: a x rule: \<tau>.strong_induct)
- case (T_refined_type xx bb cc )
- moreover hence "atom a \<sharp> bb \<and> atom a \<sharp> cc" using \<tau>.fresh by auto
- ultimately show ?case
- unfolding subst_b_\<tau>_def subst_tb.simps
- using forget_subst subst_b_b_def subst_b_c_def forget_subst \<tau>.fresh by metis
- qed
-
- fix a::bv and tm::\<tau>
- show "subst_b tm a (B_var a) = tm"
- proof (nominal_induct tm rule: \<tau>.strong_induct)
- case (T_refined_type xx bb cc)
- thus ?case
- unfolding subst_b_\<tau>_def subst_tb.simps
- using subst_id subst_b_b_def subst_b_c_def by metis
- qed
-
- fix p::perm and x1::bv and v::b and t1::\<tau>
- show "p \<bullet> subst_b t1 x1 v = subst_b (p \<bullet> t1) (p \<bullet> x1) (p \<bullet> v) "
- by (induct tm rule: \<tau>.induct, auto simp add: fresh_at_base subst_tb.simps subst_b_\<tau>_def subst_bb.simps subst_b_b_def)
-
- fix bv::bv and c::\<tau> and z::bv
- show "atom bv \<sharp> c \<Longrightarrow> ((bv \<leftrightarrow> z) \<bullet> c) = c[z::=B_var bv]\<^sub>b"
- apply (induct c rule: \<tau>.induct, (auto simp add: fresh_at_base subst_ceb.simps subst_b_ce_def permute_pure pure_supp )+)
- using flip_fresh_fresh permute_flip_cancel2 has_subst_b_class.flip_subst subst_b_c_def subst_b_b_def
- by (simp add: flip_fresh_fresh subst_b_\<tau>_def)
-
- fix bv::bv and c::\<tau> and z::bv and v::b
- show "atom bv \<sharp> c \<Longrightarrow> ((bv \<leftrightarrow> z) \<bullet> c)[bv::=v]\<^sub>b = c[z::=v]\<^sub>b"
- proof (induct c rule: \<tau>.induct)
- case (T_refined_type x1a x2a x3a)
- hence "atom bv \<sharp> x2a \<and> atom bv \<sharp> x3a \<and> atom bv \<sharp> x1a" using fresh_at_base \<tau>.fresh by simp
- then show ?case
- unfolding subst_tb.simps subst_b_\<tau>_def \<tau>.perm_simps
- using fresh_at_base flip_fresh_fresh[of bv x1a z] flip_subst_subst subst_b_b_def subst_b_c_def T_refined_type
- proof -
- have "atom z \<sharp> x1a"
- by (metis b.fresh(7) fresh_at_base(2) x_fresh_b) (* 0.0 ms *)
- then show "\<lbrace> (bv \<leftrightarrow> z) \<bullet> x1a : ((bv \<leftrightarrow> z) \<bullet> x2a)[bv::=v]\<^sub>b\<^sub>b | ((bv \<leftrightarrow> z) \<bullet> x3a)[bv::=v]\<^sub>c\<^sub>b \<rbrace> = \<lbrace> x1a : x2a[z::=v]\<^sub>b\<^sub>b | x3a[z::=v]\<^sub>c\<^sub>b \<rbrace>"
- by (metis \<open>\<lbrakk>atom bv \<sharp> x1a; atom z \<sharp> x1a\<rbrakk> \<Longrightarrow> (bv \<leftrightarrow> z) \<bullet> x1a = x1a\<close> \<open>atom bv \<sharp> x2a \<and> atom bv \<sharp> x3a \<and> atom bv \<sharp> x1a\<close> flip_subst_subst subst_b_b_def subst_b_c_def) (* 78 ms *)
- qed
- qed
-
-qed
-end
-
-lemma subst_bb_commute [simp]:
- "atom j \<sharp> A \<Longrightarrow> (subst_bb (subst_bb A i t ) j u ) = subst_bb A i (subst_bb t j u) "
- by (nominal_induct A avoiding: i j t u rule: b.strong_induct) (auto simp: fresh_at_base)
-
-lemma subst_vb_commute [simp]:
- "atom j \<sharp> A \<Longrightarrow> (subst_vb (subst_vb A i t )) j u = subst_vb A i (subst_bb t j u ) "
- by (nominal_induct A avoiding: i j t u rule: v.strong_induct) (auto simp: fresh_at_base)
-
-lemma subst_ceb_commute [simp]:
- "atom j \<sharp> A \<Longrightarrow> (subst_ceb (subst_ceb A i t )) j u = subst_ceb A i (subst_bb t j u ) "
- by (nominal_induct A avoiding: i j t u rule: ce.strong_induct) (auto simp: fresh_at_base)
-
-lemma subst_cb_commute [simp]:
- "atom j \<sharp> A \<Longrightarrow> (subst_cb (subst_cb A i t )) j u = subst_cb A i (subst_bb t j u ) "
- by (nominal_induct A avoiding: i j t u rule: c.strong_induct) (auto simp: fresh_at_base)
-
-lemma subst_tb_commute [simp]:
- "atom j \<sharp> A \<Longrightarrow> (subst_tb (subst_tb A i t )) j u = subst_tb A i (subst_bb t j u ) "
-proof (nominal_induct A avoiding: i j t u rule: \<tau>.strong_induct)
- case (T_refined_type z b c)
- then show ?case using subst_tb.simps subst_bb_commute subst_cb_commute by simp
-qed
-
-section \<open>Expressions\<close>
-
-nominal_function subst_eb :: "e \<Rightarrow> bv \<Rightarrow> b \<Rightarrow> e" where
- "subst_eb ( (AE_val v')) bv b = ( AE_val (subst_vb v' bv b))"
-| "subst_eb ( (AE_app f v') ) bv b = ( (AE_app f (subst_vb v' bv b)) )"
-| "subst_eb ( (AE_appP f b' v') ) bv b = ( (AE_appP f (b'[bv::=b]\<^sub>b\<^sub>b) (subst_vb v' bv b)))"
-| "subst_eb ( (AE_op opp v1 v2) ) bv b = ( (AE_op opp (subst_vb v1 bv b) (subst_vb v2 bv b)) )"
-| "subst_eb ( (AE_fst v')) bv b = AE_fst (subst_vb v' bv b)"
-| "subst_eb ( (AE_snd v')) bv b = AE_snd (subst_vb v' bv b)"
-| "subst_eb ( (AE_mvar u)) bv b = AE_mvar u"
-| "subst_eb ( (AE_len v')) bv b = AE_len (subst_vb v' bv b)"
-| "subst_eb ( AE_concat v1 v2) bv b = AE_concat (subst_vb v1 bv b) (subst_vb v2 bv b)"
-| "subst_eb ( AE_split v1 v2) bv b = AE_split (subst_vb v1 bv b) (subst_vb v2 bv b)"
- apply (simp add: eqvt_def subst_eb_graph_aux_def)
- apply auto
- by (meson e.strong_exhaust)
-nominal_termination (eqvt) by lexicographic_order
-
-abbreviation
- subst_eb_abbrev :: "e \<Rightarrow> bv \<Rightarrow> b \<Rightarrow> e" ("_[_::=_]\<^sub>e\<^sub>b" [1000,50,50] 500)
- where
- "e[bv::=b]\<^sub>e\<^sub>b \<equiv> subst_eb e bv b"
-
-instantiation e :: has_subst_b
-begin
-definition "subst_b = subst_eb"
-
-instance proof
- fix j::atom and i::bv and x::b and t::e
- show "j \<sharp> subst_b t i x = (atom i \<sharp> t \<and> j \<sharp> t \<or> j \<sharp> x \<and> (j \<sharp> t \<or> j = atom i))"
- proof (induct t rule: e.induct)
- case (AE_val v)
- then show ?case using subst_eb.simps fresh_def pure_fresh subst_b_e_def e.supp v.supp
- e.fresh has_subst_b_class.fresh_subst_if subst_b_e_def subst_b_v_def
- by metis
- next
- case (AE_app f v)
- then show ?case using subst_eb.simps fresh_def pure_fresh subst_b_e_def
- e.supp v.supp has_subst_b_class.fresh_subst_if subst_b_v_def
- by (metis (mono_tags, opaque_lifting) e.fresh(2))
- next
- case (AE_appP f b' v)
- then show ?case unfolding subst_eb.simps subst_b_e_def e.fresh using
- fresh_def pure_fresh subst_b_e_def e.supp v.supp
- e.fresh has_subst_b_class.fresh_subst_if subst_b_b_def subst_vb_def by (metis subst_b_v_def)
- next
- case (AE_op opp v1 v2)
- then show ?case unfolding subst_eb.simps subst_b_e_def e.fresh using
- fresh_def pure_fresh subst_b_e_def e.supp v.supp fresh_e_opp_all
- e.fresh has_subst_b_class.fresh_subst_if subst_b_b_def subst_vb_def by (metis subst_b_v_def)
- next
- case (AE_concat x1a x2)
- then show ?case using subst_eb.simps fresh_def pure_fresh subst_b_e_def e.supp v.supp
- has_subst_b_class.fresh_subst_if subst_b_v_def
- by (metis subst_vb.simps(5))
- next
- case (AE_split x1a x2)
- then show ?case using subst_eb.simps fresh_def pure_fresh subst_b_e_def e.supp v.supp
- has_subst_b_class.fresh_subst_if subst_b_v_def
- by (metis subst_vb.simps(5))
- next
- case (AE_fst x)
- then show ?case using subst_eb.simps fresh_def pure_fresh subst_b_e_def e.supp v.supp has_subst_b_class.fresh_subst_if subst_b_v_def by metis
- next
- case (AE_snd x)
- then show ?case using subst_eb.simps fresh_def pure_fresh subst_b_e_def e.supp v.supp using has_subst_b_class.fresh_subst_if subst_b_v_def by metis
- next
- case (AE_mvar x)
- then show ?case using subst_eb.simps fresh_def pure_fresh subst_b_e_def e.supp v.supp by auto
- next
- case (AE_len x)
- then show ?case using subst_eb.simps fresh_def pure_fresh subst_b_e_def e.supp v.supp using has_subst_b_class.fresh_subst_if subst_b_v_def by metis
- qed
-
- fix a::bv and tm::e and x::b
- show "atom a \<sharp> tm \<Longrightarrow> subst_b tm a x = tm"
- apply(induct tm rule: e.induct)
- apply( auto simp add: fresh_at_base subst_eb.simps subst_b_e_def)
- using has_subst_b_class.fresh_subst_if subst_b_b_def e.fresh
- using has_subst_b_class.forget_subst subst_b_v_def apply metis+
- done
-
- fix a::bv and tm::e
- show "subst_b tm a (B_var a) = tm" using subst_bb.simps subst_b_b_def
- apply (induct tm rule: e.induct)
- apply(auto simp add: fresh_at_base subst_eb.simps subst_b_e_def)
- using has_subst_b_class.fresh_subst_if subst_b_b_def e.fresh
- using has_subst_b_class.subst_id subst_b_v_def apply metis+
- done
-
- fix p::perm and x1::bv and v::b and t1::e
- show "p \<bullet> subst_b t1 x1 v = subst_b (p \<bullet> t1) (p \<bullet> x1) (p \<bullet> v)"
- apply(induct tm rule: e.induct)
- apply( auto simp add: fresh_at_base subst_bb.simps subst_b_b_def )
- using has_subst_b_class.eqvt subst_b_b_def e.fresh
- using has_subst_b_class.eqvt
- by (simp add: subst_b_e_def)+
-
- fix bv::bv and c::e and z::bv
- show "atom bv \<sharp> c \<Longrightarrow> ((bv \<leftrightarrow> z) \<bullet> c) = c[z::=B_var bv]\<^sub>b"
- apply (induct c rule: e.induct)
- apply(auto simp add: fresh_at_base subst_eb.simps subst_b_e_def subst_b_v_def permute_pure pure_supp )
- using flip_fresh_fresh permute_flip_cancel2 has_subst_b_class.flip_subst subst_b_v_def subst_b_b_def
- flip_fresh_fresh subst_b_\<tau>_def apply metis
- apply (metis (full_types) opp.perm_simps opp.strong_exhaust)
- done
-
- fix bv::bv and c::e and z::bv and v::b
- show "atom bv \<sharp> c \<Longrightarrow> ((bv \<leftrightarrow> z) \<bullet> c)[bv::=v]\<^sub>b = c[z::=v]\<^sub>b"
- apply (induct c rule: e.induct)
- apply(auto simp add: fresh_at_base subst_eb.simps subst_b_e_def subst_b_v_def permute_pure pure_supp )
- using flip_fresh_fresh permute_flip_cancel2 has_subst_b_class.flip_subst subst_b_v_def subst_b_b_def
- flip_fresh_fresh subst_b_\<tau>_def apply simp
-
- apply (metis (full_types) opp.perm_simps opp.strong_exhaust)
- done
-qed
-end
-
-section \<open>Statements\<close>
-
-nominal_function (default "case_sum (\<lambda>x. Inl undefined) (case_sum (\<lambda>x. Inl undefined) (\<lambda>x. Inr undefined))")
- subst_sb :: "s \<Rightarrow> bv \<Rightarrow> b \<Rightarrow> s"
- and subst_branchb :: "branch_s \<Rightarrow> bv \<Rightarrow> b \<Rightarrow> branch_s"
- and subst_branchlb :: "branch_list \<Rightarrow> bv \<Rightarrow> b \<Rightarrow> branch_list"
- where
- "subst_sb (AS_val v') bv b = (AS_val (subst_vb v' bv b))"
- | "subst_sb (AS_let y e s) bv b = (AS_let y (e[bv::=b]\<^sub>e\<^sub>b) (subst_sb s bv b ))"
- | "subst_sb (AS_let2 y t s1 s2) bv b = (AS_let2 y (subst_tb t bv b) (subst_sb s1 bv b ) (subst_sb s2 bv b))"
- | "subst_sb (AS_match v' cs) bv b = AS_match (subst_vb v' bv b) (subst_branchlb cs bv b)"
- | "subst_sb (AS_assign y v') bv b = AS_assign y (subst_vb v' bv b)"
- | "subst_sb (AS_if v' s1 s2) bv b = (AS_if (subst_vb v' bv b) (subst_sb s1 bv b ) (subst_sb s2 bv b ) )"
- | "subst_sb (AS_var u \<tau> v' s) bv b = AS_var u (subst_tb \<tau> bv b) (subst_vb v' bv b) (subst_sb s bv b )"
- | "subst_sb (AS_while s1 s2) bv b = AS_while (subst_sb s1 bv b ) (subst_sb s2 bv b )"
- | "subst_sb (AS_seq s1 s2) bv b = AS_seq (subst_sb s1 bv b ) (subst_sb s2 bv b )"
- | "subst_sb (AS_assert c s) bv b = AS_assert (subst_cb c bv b ) (subst_sb s bv b )"
-
-| "subst_branchb (AS_branch dc x1 s') bv b = AS_branch dc x1 (subst_sb s' bv b)"
-
-| "subst_branchlb (AS_final sb) bv b = AS_final (subst_branchb sb bv b )"
-| "subst_branchlb (AS_cons sb ssb) bv b = AS_cons (subst_branchb sb bv b) (subst_branchlb ssb bv b)"
-
- apply (simp add: eqvt_def subst_sb_subst_branchb_subst_branchlb_graph_aux_def )
- apply (auto,metis s_branch_s_branch_list.exhaust s_branch_s_branch_list.exhaust(2) old.sum.exhaust surj_pair)
-
-proof(goal_cases)
-
- have eqvt_at_proj: "\<And> s xa va . eqvt_at subst_sb_subst_branchb_subst_branchlb_sumC (Inl (s, xa, va)) \<Longrightarrow>
- eqvt_at (\<lambda>a. projl (subst_sb_subst_branchb_subst_branchlb_sumC (Inl a))) (s, xa, va)"
- apply(simp only: eqvt_at_def)
- apply(rule)
- apply(subst Projl_permute)
- apply(thin_tac _)+
- apply(simp add: subst_sb_subst_branchb_subst_branchlb_sumC_def)
- apply(simp add: THE_default_def)
- apply(case_tac "Ex1 (subst_sb_subst_branchb_subst_branchlb_graph (Inl (s,xa,va)))")
- apply simp
- apply(auto)[1]
- apply(erule_tac x="x" in allE)
- apply simp
- apply(cases rule: subst_sb_subst_branchb_subst_branchlb_graph.cases)
- apply(assumption)
- apply(rule_tac x="Sum_Type.projl x" in exI,clarify,rule the1_equality,blast,simp (no_asm) only: sum.sel)+
- apply(blast)+
- apply(simp)+
- done
- {
- case (1 y s bv b ya sa c)
- moreover have "atom y \<sharp> (bv, b) \<and> atom ya \<sharp> (bv, b)" using x_fresh_b x_fresh_bv fresh_Pair by simp
- ultimately show ?case
- using eqvt_triple eqvt_at_proj by metis
- next
- case (2 y s1 s2 bv b ya s2a c)
- moreover have "atom y \<sharp> (bv, b) \<and> atom ya \<sharp> (bv, b)" using x_fresh_b x_fresh_bv fresh_Pair by simp
- ultimately show ?case
- using eqvt_triple eqvt_at_proj by metis
- next
- case (3 u s bv b ua sa c)
- moreover have "atom u \<sharp> (bv, b) \<and> atom ua \<sharp> (bv, b)" using x_fresh_b x_fresh_bv fresh_Pair by simp
- ultimately show ?case using eqvt_triple eqvt_at_proj by metis
- next
- case (4 x1 s' bv b x1a s'a c)
- moreover have "atom x1 \<sharp> (bv, b) \<and> atom x1a \<sharp> (bv, b)" using x_fresh_b x_fresh_bv fresh_Pair by simp
- ultimately show ?case using eqvt_triple eqvt_at_proj by metis
- }
-qed
-
-nominal_termination (eqvt) by lexicographic_order
-
-abbreviation
- subst_sb_abbrev :: "s \<Rightarrow> bv \<Rightarrow> b \<Rightarrow> s" ("_[_::=_]\<^sub>s\<^sub>b" [1000,50,50] 1000)
- where
- "b[bv::=b']\<^sub>s\<^sub>b \<equiv> subst_sb b bv b'"
-
-lemma fresh_subst_sb_if [simp]:
- "(j \<sharp> (subst_sb A i x )) = ((atom i \<sharp> A \<and> j \<sharp> A) \<or> (j \<sharp> x \<and> (j \<sharp> A \<or> j = atom i)))" and
- "(j \<sharp> (subst_branchb B i x )) = ((atom i \<sharp> B \<and> j \<sharp> B) \<or> (j \<sharp> x \<and> (j \<sharp> B \<or> j = atom i)))" and
- "(j \<sharp> (subst_branchlb C i x )) = ((atom i \<sharp> C \<and> j \<sharp> C) \<or> (j \<sharp> x \<and> (j \<sharp> C \<or> j = atom i)))"
-proof (nominal_induct A and B and C avoiding: i x rule: s_branch_s_branch_list.strong_induct)
- case (AS_branch x1 x2 x3)
- have " (j \<sharp> subst_branchb (AS_branch x1 x2 x3) i x ) = (j \<sharp> (AS_branch x1 x2 (subst_sb x3 i x)))" by auto
- also have "... = ((j \<sharp> x3[i::=x]\<^sub>s\<^sub>b \<or> j \<in> set [atom x2]) \<and> j \<sharp> x1)" using s_branch_s_branch_list.fresh by auto
- also have "... = ((atom i \<sharp> AS_branch x1 x2 x3 \<and> j \<sharp> AS_branch x1 x2 x3) \<or> j \<sharp> x \<and> (j \<sharp> AS_branch x1 x2 x3 \<or> j = atom i))"
- using subst_branchb.simps(1) s_branch_s_branch_list.fresh(1) fresh_at_base has_subst_b_class.fresh_subst_if list.distinct list.set_cases set_ConsD subst_b_\<tau>_def
- v.fresh AS_branch
- proof -
- have f1: "\<forall>cs b. atom (b::bv) \<sharp> (cs::char list)" using pure_fresh by auto
-
- then have "j \<sharp> x \<and> atom i = j \<longrightarrow> ((j \<sharp> x3[i::=x]\<^sub>s\<^sub>b \<or> j \<in> set [atom x2]) \<and> j \<sharp> x1) = (atom i \<sharp> AS_branch x1 x2 x3 \<and> j \<sharp> AS_branch x1 x2 x3 \<or> j \<sharp> x \<and> (j \<sharp> AS_branch x1 x2 x3 \<or> j = atom i))"
- by (metis (full_types) AS_branch.hyps(3))
- then have "j \<sharp> x \<longrightarrow> ((j \<sharp> x3[i::=x]\<^sub>s\<^sub>b \<or> j \<in> set [atom x2]) \<and> j \<sharp> x1) = (atom i \<sharp> AS_branch x1 x2 x3 \<and> j \<sharp> AS_branch x1 x2 x3 \<or> j \<sharp> x \<and> (j \<sharp> AS_branch x1 x2 x3 \<or> j = atom i))"
- using AS_branch.hyps s_branch_s_branch_list.fresh by metis
- moreover
- { assume "\<not> j \<sharp> x"
- have ?thesis
- using f1 AS_branch.hyps(2) AS_branch.hyps(3) by force }
- ultimately show ?thesis
- by satx
- qed
- finally show ?case by auto
-next
- case (AS_cons cs css i x)
- show ?case
- unfolding subst_branchlb.simps s_branch_s_branch_list.fresh
- using AS_cons by auto
-next
- case (AS_val xx)
- then show ?case using subst_sb.simps(1) s_branch_s_branch_list.fresh has_subst_b_class.fresh_subst_if subst_b_b_def subst_b_v_def by metis
-next
- case (AS_let x1 x2 x3)
- then show ?case using subst_sb.simps s_branch_s_branch_list.fresh fresh_at_base has_subst_b_class.fresh_subst_if list.distinct list.set_cases set_ConsD subst_b_e_def
- by fastforce
-next
- case (AS_let2 x1 x2 x3 x4)
- then show ?case using subst_sb.simps s_branch_s_branch_list.fresh fresh_at_base has_subst_b_class.fresh_subst_if list.distinct list.set_cases set_ConsD subst_b_\<tau>_def
- by fastforce
-next
- case (AS_if x1 x2 x3)
- then show ?case unfolding subst_sb.simps s_branch_s_branch_list.fresh using
- has_subst_b_class.fresh_subst_if subst_b_v_def by metis
-next
- case (AS_var u t v s)
-
- have "(((atom i \<sharp> s \<and> j \<sharp> s \<or> j \<sharp> x \<and> (j \<sharp> s \<or> j = atom i)) \<or> j \<in> set [atom u]) \<and> j \<sharp> t[i::=x]\<^sub>\<tau>\<^sub>b \<and> j \<sharp> v[i::=x]\<^sub>v\<^sub>b) =
- (((atom i \<sharp> s \<and> j \<sharp> s \<or> j \<sharp> x \<and> (j \<sharp> s \<or> j = atom i)) \<or> j \<in> set [atom u]) \<and>
- ((atom i \<sharp> t \<and> j \<sharp> t \<or> j \<sharp> x \<and> (j \<sharp> t \<or> j = atom i))) \<and>
- ((atom i \<sharp> v \<and> j \<sharp> v \<or> j \<sharp> x \<and> (j \<sharp> v \<or> j = atom i))))"
- using has_subst_b_class.fresh_subst_if subst_b_v_def subst_b_\<tau>_def by metis
- also have "... = (((atom i \<sharp> s \<or> atom i \<in> set [atom u]) \<and> atom i \<sharp> t \<and> atom i \<sharp> v) \<and>
- (j \<sharp> s \<or> j \<in> set [atom u]) \<and> j \<sharp> t \<and> j \<sharp> v \<or> j \<sharp> x \<and> ((j \<sharp> s \<or> j \<in> set [atom u]) \<and> j \<sharp> t \<and> j \<sharp> v \<or> j = atom i))"
- using u_fresh_b by auto
- finally show ?case using subst_sb.simps s_branch_s_branch_list.fresh AS_var
- by simp
-next
- case (AS_assign u v )
- then show ?case unfolding subst_sb.simps s_branch_s_branch_list.fresh using
- has_subst_b_class.fresh_subst_if subst_b_v_def by force
-next
- case (AS_match v cs)
- have " j \<sharp> (AS_match v cs)[i::=x]\<^sub>s\<^sub>b = j \<sharp> (AS_match (subst_vb v i x) (subst_branchlb cs i x ))" using subst_sb.simps by auto
- also have "... = (j \<sharp> (subst_vb v i x) \<and> j \<sharp> (subst_branchlb cs i x ))" using s_branch_s_branch_list.fresh by simp
- also have "... = (j \<sharp> (subst_vb v i x) \<and> ((atom i \<sharp> cs \<and> j \<sharp> cs) \<or> j \<sharp> x \<and> (j \<sharp> cs \<or> j = atom i)))" using AS_match[of i x] by auto
- also have "... = (atom i \<sharp> AS_match v cs \<and> j \<sharp> AS_match v cs \<or> j \<sharp> x \<and> (j \<sharp> AS_match v cs \<or> j = atom i))"
- by (metis (no_types) s_branch_s_branch_list.fresh has_subst_b_class.fresh_subst_if subst_b_v_def) (* 93 ms *)
- finally show ?case by auto
-next
- case (AS_while x1 x2)
- then show ?case by auto
-next
- case (AS_seq x1 x2)
- then show ?case by auto
-next
- case (AS_assert x1 x2)
- then show ?case unfolding subst_sb.simps s_branch_s_branch_list.fresh
- using fresh_at_base has_subst_b_class.fresh_subst_if list.distinct list.set_cases set_ConsD subst_b_e_def
- by (metis subst_b_c_def)
-qed(auto+)
-
-lemma
- forget_subst_sb[simp]: "atom a \<sharp> A \<Longrightarrow> subst_sb A a x = A" and
- forget_subst_branchb [simp]: "atom a \<sharp> B \<Longrightarrow> subst_branchb B a x = B" and
- forget_subst_branchlb[simp]: "atom a \<sharp> C \<Longrightarrow> subst_branchlb C a x = C"
-proof (nominal_induct A and B and C avoiding: a x rule: s_branch_s_branch_list.strong_induct)
- case (AS_let x1 x2 x3)
- then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_e_def has_subst_b_class.forget_subst subst_b_v_def by force
-next
- case (AS_let2 x1 x2 x3 x4)
- then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_e_def has_subst_b_class.forget_subst subst_b_\<tau>_def by force
-next
- case (AS_var x1 x2 x3 x4)
- then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_e_def has_subst_b_class.forget_subst subst_b_v_def using subst_b_\<tau>_def
- proof -
- have f1: "(atom a \<sharp> x4 \<or> atom a \<in> set [atom x1]) \<and> atom a \<sharp> x2 \<and> atom a \<sharp> x3"
- using AS_var.prems s_branch_s_branch_list.fresh by simp
- then have "atom a \<sharp> x4"
- by (metis (no_types) "Nominal-Utils.fresh_star_singleton" AS_var.hyps(1) empty_set fresh_star_def list.simps(15) not_self_fresh) (* 46 ms *)
- then show ?thesis
- using f1 by (metis AS_var.hyps(3) has_subst_b_class.forget_subst subst_b_\<tau>_def subst_b_v_def subst_sb.simps(7)) (* 62 ms *)
- qed
-next
- case (AS_branch x1 x2 x3)
- then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_e_def has_subst_b_class.forget_subst subst_b_v_def by force
-next
- case (AS_cons x1 x2 x3 x4)
- then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_e_def has_subst_b_class.forget_subst subst_b_v_def by force
-next
- case (AS_val x)
- then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_e_def has_subst_b_class.forget_subst subst_b_v_def by force
-next
- case (AS_if x1 x2 x3)
- then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_e_def has_subst_b_class.forget_subst subst_b_v_def by force
-next
- case (AS_assign x1 x2)
- then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_e_def has_subst_b_class.forget_subst subst_b_v_def by force
-next
- case (AS_match x1 x2)
- then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_e_def has_subst_b_class.forget_subst subst_b_v_def by force
-next
- case (AS_while x1 x2)
- then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_e_def has_subst_b_class.forget_subst subst_b_v_def by force
-next
- case (AS_seq x1 x2)
- then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_e_def has_subst_b_class.forget_subst subst_b_v_def by force
-next
- case (AS_assert c s)
- then show ?case unfolding subst_sb.simps using
- s_branch_s_branch_list.fresh subst_b_e_def has_subst_b_class.forget_subst subst_b_v_def subst_b_c_def subst_cb.simps by force
-qed(auto+)
-
-lemma subst_sb_id: "subst_sb A a (B_var a) = A" and
- subst_branchb_id [simp]: "subst_branchb B a (B_var a) = B" and
- subst_branchlb_id: "subst_branchlb C a (B_var a) = C"
-proof(nominal_induct A and B and C avoiding: a rule: s_branch_s_branch_list.strong_induct)
- case (AS_branch x1 x2 x3)
- then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_\<tau>_def has_subst_b_class.subst_id subst_b_v_def
- by simp
-next
- case (AS_cons x1 x2 )
- then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_\<tau>_def has_subst_b_class.subst_id subst_b_v_def by simp
-next
- case (AS_val x)
- then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_\<tau>_def has_subst_b_class.subst_id subst_b_v_def by metis
-next
- case (AS_if x1 x2 x3)
- then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_\<tau>_def has_subst_b_class.subst_id subst_b_v_def by metis
-next
- case (AS_assign x1 x2)
- then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_\<tau>_def has_subst_b_class.subst_id subst_b_v_def by metis
-next
- case (AS_match x1 x2)
- then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_\<tau>_def has_subst_b_class.subst_id subst_b_v_def by metis
-next
- case (AS_while x1 x2)
- then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_\<tau>_def has_subst_b_class.subst_id subst_b_v_def by metis
-next
- case (AS_seq x1 x2)
- then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_\<tau>_def has_subst_b_class.subst_id subst_b_v_def by metis
-next
- case (AS_let x1 x2 x3)
- then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_e_def has_subst_b_class.subst_id by metis
-next
- case (AS_let2 x1 x2 x3 x4)
- then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_\<tau>_def has_subst_b_class.subst_id by metis
-next
- case (AS_var x1 x2 x3 x4)
- then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_\<tau>_def has_subst_b_class.subst_id subst_b_v_def by metis
-next
- case (AS_assert c s )
- then show ?case unfolding subst_sb.simps using s_branch_s_branch_list.fresh subst_b_c_def has_subst_b_class.subst_id by metis
-qed (auto)
-
-lemma flip_subst_s:
- fixes bv::bv and s::s and cs::branch_s and z::bv
- shows "atom bv \<sharp> s \<Longrightarrow> ((bv \<leftrightarrow> z) \<bullet> s) = s[z::=B_var bv]\<^sub>s\<^sub>b" and
- "atom bv \<sharp> cs \<Longrightarrow> ((bv \<leftrightarrow> z) \<bullet> cs) = subst_branchb cs z (B_var bv) " and
- "atom bv \<sharp> css \<Longrightarrow> ((bv \<leftrightarrow> z) \<bullet> css) = subst_branchlb css z (B_var bv) "
-
-proof(nominal_induct s and cs and css rule: s_branch_s_branch_list.strong_induct)
- case (AS_branch x1 x2 x3)
- hence "((bv \<leftrightarrow> z) \<bullet> x1) = x1" using pure_fresh fresh_at_base flip_fresh_fresh by metis
- moreover have "((bv \<leftrightarrow> z) \<bullet> x2) = x2" using fresh_at_base flip_fresh_fresh[of bv x2 z] AS_branch by auto
- ultimately show ?case unfolding s_branch_s_branch_list.perm_simps subst_branchb.simps using s_branch_s_branch_list.fresh(1) AS_branch by auto
-next
- case (AS_cons x1 x2 )
- hence "((bv \<leftrightarrow> z) \<bullet> x1) = subst_branchb x1 z (B_var bv)" using pure_fresh fresh_at_base flip_fresh_fresh s_branch_s_branch_list.fresh(13) by metis
- moreover have "((bv \<leftrightarrow> z) \<bullet> x2) = subst_branchlb x2 z (B_var bv)" using fresh_at_base flip_fresh_fresh[of bv x2 z] AS_cons s_branch_s_branch_list.fresh by metis
- ultimately show ?case unfolding s_branch_s_branch_list.perm_simps subst_branchb.simps using s_branch_s_branch_list.fresh(1) AS_cons by auto
-next
- case (AS_val x)
- then show ?case unfolding s_branch_s_branch_list.perm_simps subst_branchb.simps using flip_subst subst_b_v_def by simp
-next
- case (AS_let x1 x2 x3)
- moreover hence "((bv \<leftrightarrow> z) \<bullet> x1) = x1" using fresh_at_base flip_fresh_fresh[of bv x1 z] by auto
- ultimately show ?case
- unfolding s_branch_s_branch_list.perm_simps subst_sb.simps
- using flip_subst subst_b_e_def s_branch_s_branch_list.fresh by auto
-next
- case (AS_let2 x1 x2 x3 x4)
- moreover hence "((bv \<leftrightarrow> z) \<bullet> x1) = x1" using fresh_at_base flip_fresh_fresh[of bv x1 z] by auto
- ultimately show ?case
- unfolding s_branch_s_branch_list.perm_simps subst_sb.simps
- using flip_subst s_branch_s_branch_list.fresh(5) subst_b_\<tau>_def by auto
-next
- case (AS_if x1 x2 x3)
- thus ?case
- unfolding s_branch_s_branch_list.perm_simps subst_sb.simps
- using flip_subst subst_b_e_def subst_b_v_def s_branch_s_branch_list.fresh by auto
-next
- case (AS_var x1 x2 x3 x4)
- thus ?case
- unfolding s_branch_s_branch_list.perm_simps subst_sb.simps
- using flip_subst subst_b_e_def subst_b_v_def subst_b_\<tau>_def s_branch_s_branch_list.fresh fresh_at_base flip_fresh_fresh[of bv x1 z] by auto
-next
- case (AS_assign x1 x2)
- thus ?case
- unfolding s_branch_s_branch_list.perm_simps subst_sb.simps
- using flip_subst subst_b_e_def subst_b_v_def s_branch_s_branch_list.fresh fresh_at_base flip_fresh_fresh[of bv x1 z] by auto
-next
- case (AS_match x1 x2)
- thus ?case
- unfolding s_branch_s_branch_list.perm_simps subst_sb.simps
- using flip_subst subst_b_e_def subst_b_v_def s_branch_s_branch_list.fresh by auto
-next
- case (AS_while x1 x2)
- thus ?case
- unfolding s_branch_s_branch_list.perm_simps subst_sb.simps
- using flip_subst subst_b_e_def subst_b_v_def s_branch_s_branch_list.fresh by auto
-next
- case (AS_seq x1 x2)
- thus ?case
- unfolding s_branch_s_branch_list.perm_simps subst_sb.simps
- using flip_subst subst_b_e_def subst_b_v_def s_branch_s_branch_list.fresh by auto
-next
- case (AS_assert x1 x2)
- thus ?case
- unfolding s_branch_s_branch_list.perm_simps subst_sb.simps
- using flip_subst subst_b_c_def subst_b_v_def s_branch_s_branch_list.fresh by simp
-qed(auto)
-
-lemma flip_subst_subst_s:
- fixes bv::bv and s::s and cs::branch_s and z::bv
- shows "atom bv \<sharp> s \<Longrightarrow> ((bv \<leftrightarrow> z) \<bullet> s)[bv::=v]\<^sub>s\<^sub>b = s[z::=v]\<^sub>s\<^sub>b" and
- "atom bv \<sharp> cs \<Longrightarrow> subst_branchb ((bv \<leftrightarrow> z) \<bullet> cs) bv v = subst_branchb cs z v" and
- "atom bv \<sharp> css \<Longrightarrow> subst_branchlb ((bv \<leftrightarrow> z) \<bullet> css) bv v = subst_branchlb css z v "
-proof(nominal_induct s and cs and css rule: s_branch_s_branch_list.strong_induct)
- case (AS_branch x1 x2 x3)
- hence "((bv \<leftrightarrow> z) \<bullet> x1) = x1" using pure_fresh fresh_at_base flip_fresh_fresh by metis
- moreover have "((bv \<leftrightarrow> z) \<bullet> x2) = x2" using fresh_at_base flip_fresh_fresh[of bv x2 z] AS_branch by auto
- ultimately show ?case unfolding s_branch_s_branch_list.perm_simps subst_branchb.simps using s_branch_s_branch_list.fresh(1) AS_branch by auto
-next
- case (AS_cons x1 x2 )
- thus ?case
- unfolding s_branch_s_branch_list.perm_simps subst_branchlb.simps
- using s_branch_s_branch_list.fresh(1) AS_cons by auto
-
-next
- case (AS_val x)
- then show ?case unfolding s_branch_s_branch_list.perm_simps subst_branchb.simps using flip_subst subst_b_v_def by simp
-next
- case (AS_let x1 x2 x3)
- moreover hence "((bv \<leftrightarrow> z) \<bullet> x1) = x1" using fresh_at_base flip_fresh_fresh[of bv x1 z] by auto
- ultimately show ?case
- unfolding s_branch_s_branch_list.perm_simps subst_sb.simps
- using flip_subst_subst subst_b_e_def s_branch_s_branch_list.fresh by force
-next
- case (AS_let2 x1 x2 x3 x4)
- moreover hence "((bv \<leftrightarrow> z) \<bullet> x1) = x1" using fresh_at_base flip_fresh_fresh[of bv x1 z] by auto
- ultimately show ?case
- unfolding s_branch_s_branch_list.perm_simps subst_sb.simps
- using flip_subst s_branch_s_branch_list.fresh(5) subst_b_\<tau>_def by auto
-next
- case (AS_if x1 x2 x3)
- thus ?case
- unfolding s_branch_s_branch_list.perm_simps subst_sb.simps
- using flip_subst subst_b_e_def subst_b_v_def s_branch_s_branch_list.fresh by auto
-next
- case (AS_var x1 x2 x3 x4)
- thus ?case
- unfolding s_branch_s_branch_list.perm_simps subst_sb.simps
- using flip_subst subst_b_e_def subst_b_v_def subst_b_\<tau>_def s_branch_s_branch_list.fresh fresh_at_base flip_fresh_fresh[of bv x1 z] by auto
-next
- case (AS_assign x1 x2)
- thus ?case
- unfolding s_branch_s_branch_list.perm_simps subst_sb.simps
- using flip_subst subst_b_e_def subst_b_v_def s_branch_s_branch_list.fresh fresh_at_base flip_fresh_fresh[of bv x1 z] by auto
-next
- case (AS_match x1 x2)
- thus ?case
- unfolding s_branch_s_branch_list.perm_simps subst_sb.simps
- using flip_subst subst_b_e_def subst_b_v_def s_branch_s_branch_list.fresh by auto
-next
- case (AS_while x1 x2)
- thus ?case
- unfolding s_branch_s_branch_list.perm_simps subst_sb.simps
- using flip_subst subst_b_e_def subst_b_v_def s_branch_s_branch_list.fresh by auto
-next
- case (AS_seq x1 x2)
- thus ?case
- unfolding s_branch_s_branch_list.perm_simps subst_sb.simps
- using flip_subst subst_b_e_def subst_b_v_def s_branch_s_branch_list.fresh by auto
-next
- case (AS_assert x1 x2)
- thus ?case
- unfolding s_branch_s_branch_list.perm_simps subst_sb.simps
- using flip_subst subst_b_e_def subst_b_c_def s_branch_s_branch_list.fresh by auto
-qed(auto)
-
-instantiation s :: has_subst_b
-begin
-definition "subst_b = (\<lambda>s bv b. subst_sb s bv b)"
-
-instance proof
- fix j::atom and i::bv and x::b and t::s
- show "j \<sharp> subst_b t i x = ((atom i \<sharp> t \<and> j \<sharp> t) \<or> (j \<sharp> x \<and> (j \<sharp> t \<or> j = atom i)))"
- using fresh_subst_sb_if subst_b_s_def by metis
-
- fix a::bv and tm::s and x::b
- show "atom a \<sharp> tm \<Longrightarrow> subst_b tm a x = tm" using subst_b_s_def forget_subst_sb by metis
-
- fix a::bv and tm::s
- show "subst_b tm a (B_var a) = tm" using subst_b_s_def subst_sb_id by metis
-
- fix p::perm and x1::bv and v::b and t1::s
- show "p \<bullet> subst_b t1 x1 v = subst_b (p \<bullet> t1) (p \<bullet> x1) (p \<bullet> v)" using subst_b_s_def subst_sb_subst_branchb_subst_branchlb.eqvt by metis
-
- fix bv::bv and c::s and z::bv
- show "atom bv \<sharp> c \<Longrightarrow> ((bv \<leftrightarrow> z) \<bullet> c) = c[z::=B_var bv]\<^sub>b"
- using subst_b_s_def flip_subst_s by metis
-
- fix bv::bv and c::s and z::bv and v::b
- show "atom bv \<sharp> c \<Longrightarrow> ((bv \<leftrightarrow> z) \<bullet> c)[bv::=v]\<^sub>b = c[z::=v]\<^sub>b"
- using flip_subst_subst_s subst_b_s_def by metis
-qed
-end
-
-section \<open>Function Type\<close>
-
-nominal_function subst_ft_b :: "fun_typ \<Rightarrow> bv \<Rightarrow> b \<Rightarrow> fun_typ" where
- "subst_ft_b ( AF_fun_typ z b c t (s::s)) x v = AF_fun_typ z (subst_bb b x v) (subst_cb c x v) t[x::=v]\<^sub>\<tau>\<^sub>b s[x::=v]\<^sub>s\<^sub>b"
- apply(simp add: eqvt_def subst_ft_b_graph_aux_def )
- apply(simp add:fun_typ.strong_exhaust,auto )
- apply(rule_tac y=a and c="(a,b)" in fun_typ.strong_exhaust)
- apply (auto simp: eqvt_at_def fresh_star_def fresh_Pair fresh_at_base)
- done
-
-nominal_termination (eqvt) by lexicographic_order
-
-nominal_function subst_ftq_b :: "fun_typ_q \<Rightarrow> bv \<Rightarrow> b \<Rightarrow> fun_typ_q" where
- "atom bv \<sharp> (x,v) \<Longrightarrow> subst_ftq_b (AF_fun_typ_some bv ft) x v = (AF_fun_typ_some bv (subst_ft_b ft x v))"
-| "subst_ftq_b (AF_fun_typ_none ft) x v = (AF_fun_typ_none (subst_ft_b ft x v))"
- apply(simp add: eqvt_def subst_ftq_b_graph_aux_def )
- apply(simp add:fun_typ_q.strong_exhaust,auto )
- apply(rule_tac y=a and c="(aa,b)" in fun_typ_q.strong_exhaust)
- by (auto simp: eqvt_at_def fresh_star_def fresh_Pair fresh_at_base)
-nominal_termination (eqvt) by lexicographic_order
-
-instantiation fun_typ :: has_subst_b
-begin
-definition "subst_b = subst_ft_b"
-
-text \<open>Note: Using just simp in the second apply unpacks and gives a single goal
-whereas auto gives 43 non-intuitive goals. These goals are easier to solve
-and tedious, however they that it clear if a mistake is made in the definition of the function.
-For example, I saw that one of the goals was
-going through with metis and the next wasn't.
-It turned out the definition of the function itself was wrong\<close>
-
-instance proof
- fix j::atom and i::bv and x::b and t::fun_typ
- show "j \<sharp> subst_b t i x = (atom i \<sharp> t \<and> j \<sharp> t \<or> j \<sharp> x \<and> (j \<sharp> t \<or> j = atom i))"
- apply(nominal_induct t avoiding: i x rule:fun_typ.strong_induct)
- apply(auto simp add: subst_b_fun_typ_def )
- by(metis fresh_subst_if subst_b_s_def subst_b_\<tau>_def subst_b_b_def subst_b_c_def)+
-
- fix a::bv and tm::fun_typ and x::b
- show "atom a \<sharp> tm \<Longrightarrow> subst_b tm a x = tm"
- apply (nominal_induct tm avoiding: a x rule: fun_typ.strong_induct)
- apply(simp add: subst_b_fun_typ_def Abs1_eq_iff')
- using subst_b_b_def subst_b_fun_typ_def subst_b_\<tau>_def subst_b_c_def subst_b_s_def
- forget_subst fresh_at_base list.set_cases neq_Nil_conv set_ConsD
- subst_ft_b.simps by metis
-
- fix a::bv and tm::fun_typ
- show "subst_b tm a (B_var a) = tm"
- apply (nominal_induct tm rule: fun_typ.strong_induct)
- apply(simp add: subst_b_fun_typ_def Abs1_eq_iff',auto)
- using subst_b_b_def subst_b_fun_typ_def subst_b_\<tau>_def subst_b_c_def subst_b_s_def
- forget_subst fresh_at_base list.set_cases neq_Nil_conv set_ConsD
- subst_ft_b.simps
- by (metis has_subst_b_class.subst_id)+
-
- fix p::perm and x1::bv and v::b and t1::fun_typ
- show "p \<bullet> subst_b t1 x1 v = subst_b (p \<bullet> t1) (p \<bullet> x1) (p \<bullet> v) "
- apply (nominal_induct t1 avoiding: x1 v rule: fun_typ.strong_induct)
- by(auto simp add: subst_b_fun_typ_def Abs1_eq_iff' fun_typ.perm_simps)
-
- fix bv::bv and c::fun_typ and z::bv
- show "atom bv \<sharp> c \<Longrightarrow> ((bv \<leftrightarrow> z) \<bullet> c) = c[z::=B_var bv]\<^sub>b"
- apply (nominal_induct c avoiding: z bv rule: fun_typ.strong_induct)
- by(auto simp add: subst_b_fun_typ_def Abs1_eq_iff' fun_typ.perm_simps subst_b_b_def subst_b_c_def subst_b_\<tau>_def subst_b_s_def)
-
- fix bv::bv and c::fun_typ and z::bv and v::b
- show "atom bv \<sharp> c \<Longrightarrow> ((bv \<leftrightarrow> z) \<bullet> c)[bv::=v]\<^sub>b = c[z::=v]\<^sub>b"
- apply (nominal_induct c avoiding: bv v z rule: fun_typ.strong_induct)
- apply(auto simp add: subst_b_fun_typ_def Abs1_eq_iff' fun_typ.perm_simps subst_b_b_def subst_b_c_def subst_b_\<tau>_def subst_b_s_def flip_subst_subst flip_subst)
- using subst_b_fun_typ_def Abs1_eq_iff' fun_typ.perm_simps subst_b_b_def subst_b_c_def subst_b_\<tau>_def subst_b_s_def flip_subst_subst flip_subst
- using flip_subst_s(1) flip_subst_subst_s(1) by auto
-qed
-end
-
-instantiation fun_typ_q :: has_subst_b
-begin
-definition "subst_b = subst_ftq_b"
-
-instance proof
- fix j::atom and i::bv and x::b and t::fun_typ_q
- show "j \<sharp> subst_b t i x = (atom i \<sharp> t \<and> j \<sharp> t \<or> j \<sharp> x \<and> (j \<sharp> t \<or> j = atom i))"
- apply (nominal_induct t avoiding: i x j rule: fun_typ_q.strong_induct,auto simp add: subst_b_fun_typ_q_def subst_ftq_b.simps)
- using fresh_subst_if subst_b_fun_typ_q_def subst_b_s_def subst_b_\<tau>_def subst_b_b_def subst_b_c_def subst_b_fun_typ_def apply metis+
- done
-
- fix a::bv and t::fun_typ_q and x::b
- show "atom a \<sharp> t \<Longrightarrow> subst_b t a x = t"
- apply (nominal_induct t avoiding: a x rule: fun_typ_q.strong_induct)
- apply(auto simp add: subst_b_fun_typ_q_def subst_ftq_b.simps Abs1_eq_iff')
- using forget_subst subst_b_fun_typ_q_def subst_b_s_def subst_b_\<tau>_def subst_b_b_def subst_b_c_def subst_b_fun_typ_def eqvt by metis+
-
- fix p::perm and x1::bv and v::b and t::fun_typ_q
- show "p \<bullet> subst_b t x1 v = subst_b (p \<bullet> t) (p \<bullet> x1) (p \<bullet> v) "
- apply (nominal_induct t avoiding: x1 v rule: fun_typ_q.strong_induct)
- by(auto simp add: subst_b_fun_typ_q_def subst_ftq_b.simps Abs1_eq_iff')
-
- fix a::bv and tm::fun_typ_q
- show "subst_b tm a (B_var a) = tm"
- apply (nominal_induct tm avoiding: a rule: fun_typ_q.strong_induct)
- apply(auto simp add: subst_b_fun_typ_q_def subst_ftq_b.simps Abs1_eq_iff')
- using subst_id subst_b_b_def subst_b_fun_typ_def subst_b_\<tau>_def subst_b_c_def subst_b_s_def
- forget_subst fresh_at_base list.set_cases neq_Nil_conv set_ConsD
- subst_ft_b.simps by metis+
-
- fix bv::bv and c::fun_typ_q and z::bv
- show "atom bv \<sharp> c \<Longrightarrow> ((bv \<leftrightarrow> z) \<bullet> c) = c[z::=B_var bv]\<^sub>b"
- apply (nominal_induct c avoiding: z bv rule: fun_typ_q.strong_induct)
- apply(auto simp add: subst_b_fun_typ_q_def subst_ftq_b.simps Abs1_eq_iff')
- using forget_subst subst_b_fun_typ_q_def subst_b_s_def subst_b_\<tau>_def subst_b_b_def subst_b_c_def subst_b_fun_typ_def eqvt by metis+
-
- fix bv::bv and c::fun_typ_q and z::bv and v::b
- show "atom bv \<sharp> c \<Longrightarrow> ((bv \<leftrightarrow> z) \<bullet> c)[bv::=v]\<^sub>b = c[z::=v]\<^sub>b"
- apply (nominal_induct c avoiding: z v bv rule: fun_typ_q.strong_induct)
- apply(auto simp add: subst_b_fun_typ_q_def subst_ftq_b.simps Abs1_eq_iff')
- using flip_subst flip_subst_subst forget_subst subst_b_fun_typ_q_def subst_b_s_def subst_b_\<tau>_def subst_b_b_def subst_b_c_def subst_b_fun_typ_def eqvt by metis+
-
-qed
-end
-
-section \<open>Contexts\<close>
-
-subsection \<open>Immutable Variables\<close>
-nominal_function subst_gb :: "\<Gamma> \<Rightarrow> bv \<Rightarrow> b \<Rightarrow> \<Gamma>" where
- "subst_gb GNil _ _ = GNil"
-| "subst_gb ((y,b',c)#\<^sub>\<Gamma>\<Gamma>) bv b = ((y,b'[bv::=b]\<^sub>b\<^sub>b,c[bv::=b]\<^sub>c\<^sub>b)#\<^sub>\<Gamma> (subst_gb \<Gamma> bv b))"
- apply (simp add: eqvt_def subst_gb_graph_aux_def )+
- apply auto
- apply (insert \<Gamma>.exhaust neq_GNil_conv, force)
- done
-nominal_termination (eqvt) by lexicographic_order
-
-abbreviation
- subst_gb_abbrev :: "\<Gamma> \<Rightarrow> bv \<Rightarrow> b \<Rightarrow> \<Gamma>" ("_[_::=_]\<^sub>\<Gamma>\<^sub>b" [1000,50,50] 1000)
- where
- "g[bv::=b']\<^sub>\<Gamma>\<^sub>b \<equiv> subst_gb g bv b'"
-
-instantiation \<Gamma> :: has_subst_b
-begin
-definition "subst_b = subst_gb"
-
-instance proof
- fix j::atom and i::bv and x::b and t::\<Gamma>
- show "j \<sharp> subst_b t i x = (atom i \<sharp> t \<and> j \<sharp> t \<or> j \<sharp> x \<and> (j \<sharp> t \<or> j = atom i))"
- proof(induct t rule: \<Gamma>_induct)
- case GNil
- then show ?case using fresh_GNil subst_gb.simps fresh_def pure_fresh subst_b_\<Gamma>_def has_subst_b_class.fresh_subst_if fresh_GNil fresh_GCons by metis
- next
- case (GCons x' b' c' \<Gamma>')
- have *: "atom i \<sharp> x' " using fresh_at_base by simp
-
- have "j \<sharp> subst_b ((x', b', c') #\<^sub>\<Gamma> \<Gamma>') i x = j \<sharp> ((x', b'[i::=x]\<^sub>b\<^sub>b, c'[i::=x]\<^sub>c\<^sub>b) #\<^sub>\<Gamma> (subst_b \<Gamma>' i x))" using subst_gb.simps subst_b_\<Gamma>_def by auto
- also have "... = (j \<sharp> ((x', b'[i::=x]\<^sub>b\<^sub>b, c'[i::=x]\<^sub>c\<^sub>b)) \<and> (j \<sharp> (subst_b \<Gamma>' i x)))" using fresh_GCons by auto
- also have "... = (((j \<sharp> x') \<and> (j \<sharp> b'[i::=x]\<^sub>b\<^sub>b) \<and> (j \<sharp> c'[i::=x]\<^sub>c\<^sub>b)) \<and> (j \<sharp> (subst_b \<Gamma>' i x)))" by auto
- also have "... = (((j \<sharp> x') \<and> ((atom i \<sharp> b' \<and> j \<sharp> b' \<or> j \<sharp> x \<and> (j \<sharp> b' \<or> j = atom i))) \<and>
- ((atom i \<sharp> c' \<and> j \<sharp> c' \<or> j \<sharp> x \<and> (j \<sharp> c' \<or> j = atom i))) \<and>
- ((atom i \<sharp> \<Gamma>' \<and> j \<sharp> \<Gamma>' \<or> j \<sharp> x \<and> (j \<sharp> \<Gamma>' \<or> j = atom i)))))"
- using fresh_subst_if[of j b' i x] fresh_subst_if[of j c' i x] GCons subst_b_b_def subst_b_c_def by simp
- also have "... = ((atom i \<sharp> (x', b', c') #\<^sub>\<Gamma> \<Gamma>' \<and> j \<sharp> (x', b', c') #\<^sub>\<Gamma> \<Gamma>') \<or> (j \<sharp> x \<and> (j \<sharp> (x', b', c') #\<^sub>\<Gamma> \<Gamma>' \<or> j = atom i)))" using * fresh_GCons fresh_prod3 by metis
-
- finally show ?case by auto
- qed
-
- fix a::bv and tm::\<Gamma> and x::b
- show "atom a \<sharp> tm \<Longrightarrow> subst_b tm a x = tm"
- proof (induct tm rule: \<Gamma>_induct)
- case GNil
- then show ?case using subst_gb.simps subst_b_\<Gamma>_def by auto
- next
- case (GCons x' b' c' \<Gamma>')
- have *:"b'[a::=x]\<^sub>b\<^sub>b = b' \<and> c'[a::=x]\<^sub>c\<^sub>b = c'" using GCons fresh_GCons[of "atom a"] fresh_prod3[of "atom a"] has_subst_b_class.forget_subst subst_b_b_def subst_b_c_def by metis
- have "subst_b ((x', b', c') #\<^sub>\<Gamma> \<Gamma>') a x = ((x', b'[a::=x]\<^sub>b\<^sub>b, c'[a::=x]\<^sub>c\<^sub>b) #\<^sub>\<Gamma> (subst_b \<Gamma>' a x))" using subst_b_\<Gamma>_def subst_gb.simps by auto
- also have "... = ((x', b', c') #\<^sub>\<Gamma> \<Gamma>')" using * GCons fresh_GCons[of "atom a"] by auto
- finally show ?case using has_subst_b_class.forget_subst fresh_GCons fresh_prod3 GCons subst_b_\<Gamma>_def has_subst_b_class.forget_subst[of a b' x] fresh_prod3[of "atom a"] by argo
- qed
-
- fix a::bv and tm::\<Gamma>
- show "subst_b tm a (B_var a) = tm"
- proof(induct tm rule: \<Gamma>_induct)
- case GNil
- then show ?case using subst_gb.simps subst_b_\<Gamma>_def by auto
- next
- case (GCons x' b' c' \<Gamma>')
- then show ?case using has_subst_b_class.subst_id subst_b_\<Gamma>_def subst_b_b_def subst_b_c_def subst_gb.simps by metis
- qed
-
- fix p::perm and x1::bv and v::b and t1::\<Gamma>
- show "p \<bullet> subst_b t1 x1 v = subst_b (p \<bullet> t1) (p \<bullet> x1) (p \<bullet> v)"
- proof (induct tm rule: \<Gamma>_induct)
- case GNil
- then show ?case using subst_b_\<Gamma>_def subst_gb.simps by simp
- next
- case (GCons x' b' c' \<Gamma>')
- then show ?case using subst_b_\<Gamma>_def subst_gb.simps has_subst_b_class.eqvt by argo
- qed
-
- fix bv::bv and c::\<Gamma> and z::bv
- show "atom bv \<sharp> c \<Longrightarrow> ((bv \<leftrightarrow> z) \<bullet> c) = c[z::=B_var bv]\<^sub>b"
- proof (induct c rule: \<Gamma>_induct)
- case GNil
- then show ?case using subst_b_\<Gamma>_def subst_gb.simps by auto
- next
- case (GCons x b c \<Gamma>')
- have *:"(bv \<leftrightarrow> z) \<bullet> (x, b, c) = (x, (bv \<leftrightarrow> z) \<bullet> b, (bv \<leftrightarrow> z) \<bullet> c)" using flip_bv_x_cancel by auto
- then show ?case
- unfolding subst_gb.simps subst_b_\<Gamma>_def permute_\<Gamma>.simps *
- using GCons subst_b_\<Gamma>_def subst_gb.simps flip_subst subst_b_b_def subst_b_c_def fresh_GCons by auto
- qed
-
- fix bv::bv and c::\<Gamma> and z::bv and v::b
- show "atom bv \<sharp> c \<Longrightarrow> ((bv \<leftrightarrow> z) \<bullet> c)[bv::=v]\<^sub>b = c[z::=v]\<^sub>b"
- proof (induct c rule: \<Gamma>_induct)
- case GNil
- then show ?case using subst_b_\<Gamma>_def subst_gb.simps by auto
- next
- case (GCons x b c \<Gamma>')
- have *:"(bv \<leftrightarrow> z) \<bullet> (x, b, c) = (x, (bv \<leftrightarrow> z) \<bullet> b, (bv \<leftrightarrow> z) \<bullet> c)" using flip_bv_x_cancel by auto
- then show ?case
- unfolding subst_gb.simps subst_b_\<Gamma>_def permute_\<Gamma>.simps *
- using GCons subst_b_\<Gamma>_def subst_gb.simps flip_subst subst_b_b_def subst_b_c_def fresh_GCons by auto
- qed
-qed
-end
-
-lemma subst_b_base_for_lit:
- "(base_for_lit l)[bv::=b]\<^sub>b\<^sub>b = base_for_lit l"
- using base_for_lit.simps l.strong_exhaust
- by (metis subst_bb.simps(2) subst_bb.simps(3) subst_bb.simps(6) subst_bb.simps(7))
-
-lemma subst_b_lookup:
- assumes "Some (b, c) = lookup \<Gamma> x"
- shows " Some (b[bv::=b']\<^sub>b\<^sub>b, c[bv::=b']\<^sub>c\<^sub>b) = lookup \<Gamma>[bv::=b']\<^sub>\<Gamma>\<^sub>b x"
- using assms by(induct \<Gamma> rule: \<Gamma>_induct, auto)
-
-lemma subst_g_b_x_fresh:
- fixes x::x and b::b and \<Gamma>::\<Gamma> and bv::bv
- assumes "atom x \<sharp> \<Gamma>"
- shows "atom x \<sharp> \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b"
- using subst_b_fresh_x subst_b_\<Gamma>_def assms by metis
-
-subsection \<open>Mutable Variables\<close>
-
-nominal_function subst_db :: "\<Delta> \<Rightarrow> bv \<Rightarrow> b \<Rightarrow> \<Delta>" where
- "subst_db []\<^sub>\<Delta> _ _ = []\<^sub>\<Delta>"
-| "subst_db ((u,t) #\<^sub>\<Delta> \<Delta>) bv b = ((u,t[bv::=b]\<^sub>\<tau>\<^sub>b) #\<^sub>\<Delta> (subst_db \<Delta> bv b))"
- apply (simp add: eqvt_def subst_db_graph_aux_def,auto )
- using list.exhaust delete_aux.elims
- using neq_DNil_conv by fastforce
-nominal_termination (eqvt) by lexicographic_order
-
-abbreviation
- subst_db_abbrev :: "\<Delta> \<Rightarrow> bv \<Rightarrow> b \<Rightarrow> \<Delta>" ("_[_::=_]\<^sub>\<Delta>\<^sub>b" [1000,50,50] 1000)
- where
- "\<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b \<equiv> subst_db \<Delta> bv b"
-
-instantiation \<Delta> :: has_subst_b
-begin
-definition "subst_b = subst_db"
-
-instance proof
- fix j::atom and i::bv and x::b and t::\<Delta>
- show "j \<sharp> subst_b t i x = (atom i \<sharp> t \<and> j \<sharp> t \<or> j \<sharp> x \<and> (j \<sharp> t \<or> j = atom i))"
- proof(induct t rule: \<Delta>_induct)
- case DNil
- then show ?case using fresh_DNil subst_db.simps fresh_def pure_fresh subst_b_\<Delta>_def has_subst_b_class.fresh_subst_if fresh_DNil fresh_DCons by metis
- next
- case (DCons u t \<Gamma>')
- have "j \<sharp> subst_b ((u, t) #\<^sub>\<Delta> \<Gamma>') i x = j \<sharp> ((u, t[i::=x]\<^sub>\<tau>\<^sub>b) #\<^sub>\<Delta> (subst_b \<Gamma>' i x))" using subst_db.simps subst_b_\<Delta>_def by auto
- also have "... = (j \<sharp> ((u, t[i::=x]\<^sub>\<tau>\<^sub>b)) \<and> (j \<sharp> (subst_b \<Gamma>' i x)))" using fresh_DCons by auto
- also have "... = (((j \<sharp> u) \<and> (j \<sharp> t[i::=x]\<^sub>\<tau>\<^sub>b)) \<and> (j \<sharp> (subst_b \<Gamma>' i x)))" by auto
- also have "... = ((j \<sharp> u) \<and> ((atom i \<sharp> t \<and> j \<sharp> t) \<or> (j \<sharp> x \<and> (j \<sharp> t \<or> j = atom i))) \<and> (atom i \<sharp> \<Gamma>' \<and> j \<sharp> \<Gamma>' \<or> j \<sharp> x \<and> (j \<sharp> \<Gamma>' \<or> j = atom i)))"
- using has_subst_b_class.fresh_subst_if[of j t i x] subst_b_\<tau>_def DCons subst_b_\<Delta>_def by auto
- also have "... = (atom i \<sharp> (u, t) #\<^sub>\<Delta> \<Gamma>' \<and> j \<sharp> (u, t) #\<^sub>\<Delta> \<Gamma>' \<or> j \<sharp> x \<and> (j \<sharp> (u, t) #\<^sub>\<Delta> \<Gamma>' \<or> j = atom i))"
- using DCons subst_db.simps(2) has_subst_b_class.fresh_subst_if fresh_DCons subst_b_\<Delta>_def pure_fresh fresh_at_base by auto
- finally show ?case by auto
- qed
-
- fix a::bv and tm::\<Delta> and x::b
- show "atom a \<sharp> tm \<Longrightarrow> subst_b tm a x = tm"
- proof (induct tm rule: \<Delta>_induct)
- case DNil
- then show ?case using subst_db.simps subst_b_\<Delta>_def by auto
- next
- case (DCons u t \<Gamma>')
- have *:"t[a::=x]\<^sub>\<tau>\<^sub>b = t" using DCons fresh_DCons[of "atom a"] fresh_prod2[of "atom a"] has_subst_b_class.forget_subst subst_b_\<tau>_def by metis
- have "subst_b ((u,t) #\<^sub>\<Delta> \<Gamma>') a x = ((u,t[a::=x]\<^sub>\<tau>\<^sub>b) #\<^sub>\<Delta> (subst_b \<Gamma>' a x))" using subst_b_\<Delta>_def subst_db.simps by auto
- also have "... = ((u, t) #\<^sub>\<Delta> \<Gamma>')" using * DCons fresh_DCons[of "atom a"] by auto
- finally show ?case using
- has_subst_b_class.forget_subst fresh_DCons fresh_prod3
- DCons subst_b_\<Delta>_def has_subst_b_class.forget_subst[of a t x] fresh_prod3[of "atom a"] by argo
- qed
-
- fix a::bv and tm::\<Delta>
- show "subst_b tm a (B_var a) = tm"
- proof(induct tm rule: \<Delta>_induct)
- case DNil
- then show ?case using subst_db.simps subst_b_\<Delta>_def by auto
- next
- case (DCons u t \<Gamma>')
- then show ?case using has_subst_b_class.subst_id subst_b_\<Delta>_def subst_b_\<tau>_def subst_db.simps by metis
- qed
-
- fix p::perm and x1::bv and v::b and t1::\<Delta>
- show "p \<bullet> subst_b t1 x1 v = subst_b (p \<bullet> t1) (p \<bullet> x1) (p \<bullet> v)"
- proof (induct tm rule: \<Delta>_induct)
- case DNil
- then show ?case using subst_b_\<Delta>_def subst_db.simps by simp
- next
- case (DCons x' b' \<Gamma>')
- then show ?case by argo
- qed
-
- fix bv::bv and c::\<Delta> and z::bv
- show "atom bv \<sharp> c \<Longrightarrow> ((bv \<leftrightarrow> z) \<bullet> c) = c[z::=B_var bv]\<^sub>b"
- proof (induct c rule: \<Delta>_induct)
- case DNil
- then show ?case using subst_b_\<Delta>_def subst_db.simps by auto
- next
- case (DCons u t')
- then show ?case
- unfolding subst_db.simps subst_b_\<Delta>_def permute_\<Delta>.simps
- using DCons subst_b_\<Delta>_def subst_db.simps flip_subst subst_b_\<tau>_def flip_fresh_fresh fresh_at_base fresh_DCons flip_bv_u_cancel by simp
- qed
-
- fix bv::bv and c::\<Delta> and z::bv and v::b
- show "atom bv \<sharp> c \<Longrightarrow> ((bv \<leftrightarrow> z) \<bullet> c)[bv::=v]\<^sub>b = c[z::=v]\<^sub>b"
- proof (induct c rule: \<Delta>_induct)
- case DNil
- then show ?case using subst_b_\<Delta>_def subst_db.simps by auto
- next
- case (DCons u t')
- then show ?case
- unfolding subst_db.simps subst_b_\<Delta>_def permute_\<Delta>.simps
- using DCons subst_b_\<Delta>_def subst_db.simps flip_subst subst_b_\<tau>_def flip_fresh_fresh fresh_at_base fresh_DCons flip_bv_u_cancel by simp
- qed
-qed
-end
-
-lemma subst_d_b_member:
- assumes "(u, \<tau>) \<in> setD \<Delta>"
- shows "(u, \<tau>[bv::=b]\<^sub>\<tau>\<^sub>b) \<in> setD \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b"
- using assms by (induct \<Delta>,auto)
-
-lemmas ms_fresh_all = e.fresh s_branch_s_branch_list.fresh \<tau>.fresh c.fresh ce.fresh v.fresh l.fresh fresh_at_base opp.fresh pure_fresh ms_fresh
-
-lemmas fresh_intros[intro] = fresh_GNil x_not_in_b_set x_not_in_u_atoms x_fresh_b u_not_in_x_atoms bv_not_in_x_atoms u_not_in_b_atoms
-
-lemmas subst_b_simps = subst_tb.simps subst_cb.simps subst_ceb.simps subst_vb.simps subst_bb.simps subst_eb.simps subst_branchb.simps subst_sb.simps
-
-lemma subst_d_b_x_fresh:
- fixes x::x and b::b and \<Delta>::\<Delta> and bv::bv
- assumes "atom x \<sharp> \<Delta>"
- shows "atom x \<sharp> \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b"
- using subst_b_fresh_x subst_b_\<Delta>_def assms by metis
-
-lemma subst_b_fresh_x:
- fixes x::x
- shows "atom x \<sharp> v \<Longrightarrow> atom x \<sharp> v[bv::=b']\<^sub>v\<^sub>b" and
- "atom x \<sharp> ce \<Longrightarrow> atom x \<sharp> ce[bv::=b']\<^sub>c\<^sub>e\<^sub>b" and
- "atom x \<sharp> e \<Longrightarrow> atom x \<sharp> e[bv::=b']\<^sub>e\<^sub>b" and
- "atom x \<sharp> c \<Longrightarrow> atom x \<sharp> c[bv::=b']\<^sub>c\<^sub>b" and
- "atom x \<sharp> t \<Longrightarrow> atom x \<sharp> t[bv::=b']\<^sub>\<tau>\<^sub>b" and
- "atom x \<sharp> d \<Longrightarrow> atom x \<sharp> d[bv::=b']\<^sub>\<Delta>\<^sub>b" and
- "atom x \<sharp> g \<Longrightarrow> atom x \<sharp> g[bv::=b']\<^sub>\<Gamma>\<^sub>b" and
- "atom x \<sharp> s \<Longrightarrow> atom x \<sharp> s[bv::=b']\<^sub>s\<^sub>b"
- using fresh_subst_if x_fresh_b subst_b_v_def subst_b_ce_def subst_b_e_def subst_b_c_def subst_b_\<tau>_def subst_b_s_def subst_g_b_x_fresh subst_d_b_x_fresh
- by metis+
-
-lemma subst_b_fresh_u_cls:
- fixes tm::"'a::has_subst_b" and x::u
- shows "atom x \<sharp> tm = atom x \<sharp> tm[bv::=b']\<^sub>b"
- using fresh_subst_if[of "atom x" tm bv b'] using u_fresh_b by auto
-
-lemma subst_g_b_u_fresh:
- fixes x::u and b::b and \<Gamma>::\<Gamma> and bv::bv
- assumes "atom x \<sharp> \<Gamma>"
- shows "atom x \<sharp> \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b"
- using subst_b_fresh_u_cls subst_b_\<Gamma>_def assms by metis
-
-lemma subst_d_b_u_fresh:
- fixes x::u and b::b and \<Gamma>::\<Delta> and bv::bv
- assumes "atom x \<sharp> \<Gamma>"
- shows "atom x \<sharp> \<Gamma>[bv::=b]\<^sub>\<Delta>\<^sub>b"
- using subst_b_fresh_u_cls subst_b_\<Delta>_def assms by metis
-
-lemma subst_b_fresh_u:
- fixes x::u
- shows "atom x \<sharp> v \<Longrightarrow> atom x \<sharp> v[bv::=b']\<^sub>v\<^sub>b" and
- "atom x \<sharp> ce \<Longrightarrow> atom x \<sharp> ce[bv::=b']\<^sub>c\<^sub>e\<^sub>b" and
- "atom x \<sharp> e \<Longrightarrow> atom x \<sharp> e[bv::=b']\<^sub>e\<^sub>b" and
- "atom x \<sharp> c \<Longrightarrow> atom x \<sharp> c[bv::=b']\<^sub>c\<^sub>b" and
- "atom x \<sharp> t \<Longrightarrow> atom x \<sharp> t[bv::=b']\<^sub>\<tau>\<^sub>b" and
- "atom x \<sharp> d \<Longrightarrow> atom x \<sharp> d[bv::=b']\<^sub>\<Delta>\<^sub>b" and
- "atom x \<sharp> g \<Longrightarrow> atom x \<sharp> g[bv::=b']\<^sub>\<Gamma>\<^sub>b" and
- "atom x \<sharp> s \<Longrightarrow> atom x \<sharp> s[bv::=b']\<^sub>s\<^sub>b"
- using fresh_subst_if u_fresh_b subst_b_v_def subst_b_ce_def subst_b_e_def subst_b_c_def subst_b_\<tau>_def subst_b_s_def subst_g_b_u_fresh subst_d_b_u_fresh
- by metis+
-
-lemma subst_db_u_fresh:
- fixes u::u and b::b and D::\<Delta>
- assumes "atom u \<sharp> D"
- shows "atom u \<sharp> D[bv::=b]\<^sub>\<Delta>\<^sub>b"
- using assms proof(induct D rule: \<Delta>_induct)
- case DNil
- then show ?case by auto
-next
- case (DCons u' t' D')
- then show ?case using subst_db.simps fresh_def fresh_DCons fresh_subst_if subst_b_\<tau>_def
- by (metis fresh_Pair u_not_in_b_atoms)
-qed
-
-lemma flip_bt_subst4:
- fixes t::\<tau> and bv::bv
- assumes "atom bv \<sharp> t"
- shows "t[bv'::=b]\<^sub>\<tau>\<^sub>b = ((bv' \<leftrightarrow> bv) \<bullet> t)[bv::=b]\<^sub>\<tau>\<^sub>b"
- using flip_subst_subst[OF assms,of bv' b]
- by (simp add: flip_commute subst_b_\<tau>_def)
-
-lemma subst_bt_flip_sym:
- fixes t1::\<tau> and t2::\<tau>
- assumes "atom bv \<sharp> b" and "atom bv \<sharp> (bv1, bv2, t1, t2)" and "(bv1 \<leftrightarrow> bv) \<bullet> t1 = (bv2 \<leftrightarrow> bv) \<bullet> t2"
- shows " t1[bv1::=b]\<^sub>\<tau>\<^sub>b = t2[bv2::=b]\<^sub>\<tau>\<^sub>b"
- using assms flip_bt_subst4[of bv t1 bv1 b ] flip_bt_subst4 fresh_prod4 fresh_Pair by metis
-
+(*<*)
+theory BTVSubst
+ imports Syntax
+begin
+ (*>*)
+chapter \<open>Basic Type Variable Substitution\<close>
+
+section \<open>Class\<close>
+
+class has_subst_b = fs +
+ fixes subst_b :: "'a::fs \<Rightarrow> bv \<Rightarrow> b \<Rightarrow> 'a::fs" ("_[_::=_]\<^sub>b" [1000,50,50] 1000)
+
+assumes fresh_subst_if: "j \<sharp> (t[i::=x]\<^sub>b ) \<longleftrightarrow> (atom i \<sharp> t \<and> j \<sharp> t) \<or> (j \<sharp> x \<and> (j \<sharp> t \<or> j = atom i))"
+ and forget_subst[simp]: "atom a \<sharp> tm \<Longrightarrow> tm[a::=x]\<^sub>b = tm"
+ and subst_id[simp]: "tm[a::=(B_var a)]\<^sub>b = tm"
+ and eqvt[simp,eqvt]: "(p::perm) \<bullet> (subst_b t1 x1 v ) = (subst_b (p \<bullet>t1) (p \<bullet>x1) (p \<bullet>v) )"
+ and flip_subst[simp]: "atom bv \<sharp> c \<Longrightarrow> ((bv \<leftrightarrow> z) \<bullet> c) = c[z::=B_var bv]\<^sub>b"
+ and flip_subst_subst[simp]: "atom bv \<sharp> c \<Longrightarrow> ((bv \<leftrightarrow> z) \<bullet> c)[bv::=v]\<^sub>b = c[z::=v]\<^sub>b"
+begin
+
+lemmas flip_subst_b = flip_subst_subst
+
+lemma subst_b_simple_commute:
+ fixes x::bv
+ assumes "atom x \<sharp> c"
+ shows "(c[z::=B_var x]\<^sub>b)[x::=b]\<^sub>b = c[z::=b]\<^sub>b"
+proof -
+ have "(c[z::=B_var x]\<^sub>b)[x::=b]\<^sub>b = (( x \<leftrightarrow> z) \<bullet> c)[x::=b]\<^sub>b" using flip_subst assms by simp
+ thus ?thesis using flip_subst_subst assms by simp
+qed
+
+lemma subst_b_flip_eq_one:
+ fixes z1::bv and z2::bv and x1::bv and x2::bv
+ assumes "[[atom z1]]lst. c1 = [[atom z2]]lst. c2"
+ and "atom x1 \<sharp> (z1,z2,c1,c2)"
+ shows "(c1[z1::=B_var x1]\<^sub>b) = (c2[z2::=B_var x1]\<^sub>b)"
+proof -
+ have "(c1[z1::=B_var x1]\<^sub>b) = (x1 \<leftrightarrow> z1) \<bullet> c1" using assms flip_subst by auto
+ moreover have "(c2[z2::=B_var x1]\<^sub>b) = (x1 \<leftrightarrow> z2) \<bullet> c2" using assms flip_subst by auto
+ ultimately show ?thesis using Abs1_eq_iff_all(3)[of z1 c1 z2 c2 z1] assms
+ by (metis Abs1_eq_iff_fresh(3) flip_commute)
+qed
+
+lemma subst_b_flip_eq_two:
+ fixes z1::bv and z2::bv and x1::bv and x2::bv
+ assumes "[[atom z1]]lst. c1 = [[atom z2]]lst. c2"
+ shows "(c1[z1::=b]\<^sub>b) = (c2[z2::=b]\<^sub>b)"
+proof -
+ obtain x::bv where *:"atom x \<sharp> (z1,z2,c1,c2)" using obtain_fresh by metis
+ hence "(c1[z1::=B_var x]\<^sub>b) = (c2[z2::=B_var x]\<^sub>b)" using subst_b_flip_eq_one[OF assms, of x] by metis
+ hence "(c1[z1::=B_var x]\<^sub>b)[x::=b]\<^sub>b = (c2[z2::=B_var x]\<^sub>b)[x::=b]\<^sub>b" by auto
+ thus ?thesis using subst_b_simple_commute * fresh_prod4 by metis
+qed
+
+lemma subst_b_fresh_x:
+ fixes tm::"'a::fs" and x::x
+ shows "atom x \<sharp> tm = atom x \<sharp> tm[bv::=b']\<^sub>b"
+ using fresh_subst_if[of "atom x" tm bv b'] using x_fresh_b by auto
+
+lemma subst_b_x_flip[simp]:
+ fixes x'::x and x::x and bv::bv
+ shows "((x' \<leftrightarrow> x) \<bullet> tm)[bv::=b']\<^sub>b = (x' \<leftrightarrow> x) \<bullet> (tm[bv::=b']\<^sub>b)"
+proof -
+ have "(x' \<leftrightarrow> x) \<bullet> bv = bv" using pure_supp flip_fresh_fresh by force
+ moreover have "(x' \<leftrightarrow> x) \<bullet> b' = b'" using x_fresh_b flip_fresh_fresh by auto
+ ultimately show ?thesis using eqvt by simp
+qed
+
+end
+
+section \<open>Base Type\<close>
+
+nominal_function subst_bb :: "b \<Rightarrow> bv \<Rightarrow> b \<Rightarrow> b" where
+ "subst_bb (B_var bv2) bv1 b = (if bv1 = bv2 then b else (B_var bv2))"
+| "subst_bb B_int bv1 b = B_int"
+| "subst_bb B_bool bv1 b = B_bool"
+| "subst_bb (B_id s) bv1 b = B_id s "
+| "subst_bb (B_pair b1 b2) bv1 b = B_pair (subst_bb b1 bv1 b) (subst_bb b2 bv1 b)"
+| "subst_bb B_unit bv1 b = B_unit "
+| "subst_bb B_bitvec bv1 b = B_bitvec "
+| "subst_bb (B_app s b2) bv1 b = B_app s (subst_bb b2 bv1 b)"
+ apply (simp add: eqvt_def subst_bb_graph_aux_def )
+ apply (simp add: eqvt_def subst_bb_graph_aux_def )
+ by (auto,meson b.strong_exhaust)
+nominal_termination (eqvt) by lexicographic_order
+
+abbreviation
+ subst_bb_abbrev :: "b \<Rightarrow> bv \<Rightarrow> b \<Rightarrow> b" ("_[_::=_]\<^sub>b\<^sub>b" [1000,50,50] 1000)
+ where
+ "b[bv::=b']\<^sub>b\<^sub>b \<equiv> subst_bb b bv b' "
+
+instantiation b :: has_subst_b
+begin
+definition "subst_b = subst_bb"
+
+instance proof
+ fix j::atom and i::bv and x::b and t::b
+ show "j \<sharp> subst_b t i x = (atom i \<sharp> t \<and> j \<sharp> t \<or> j \<sharp> x \<and> (j \<sharp> t \<or> j = atom i))"
+ proof (induct t rule: b.induct)
+ case (B_id x)
+ then show ?case using subst_bb.simps fresh_def pure_fresh subst_b_b_def by auto
+ next
+ case (B_var x)
+ then show ?case using subst_bb.simps fresh_def pure_fresh subst_b_b_def by auto
+ next
+ case (B_app x1 x2)
+ then show ?case using subst_bb.simps fresh_def pure_fresh subst_b_b_def by auto
+ qed(auto simp add: subst_bb.simps fresh_def pure_fresh subst_b_b_def)+
+
+ fix a::bv and tm::b and x::b
+ show "atom a \<sharp> tm \<Longrightarrow> tm[a::=x]\<^sub>b = tm"
+ by (induct tm rule: b.induct, auto simp add: fresh_at_base subst_bb.simps subst_b_b_def)
+
+ fix a::bv and tm::b
+ show "subst_b tm a (B_var a) = tm" using subst_bb.simps subst_b_b_def
+ by (induct tm rule: b.induct, auto simp add: fresh_at_base subst_bb.simps subst_b_b_def)
+
+ fix p::perm and x1::bv and v::b and t1::b
+ show "p \<bullet> subst_b t1 x1 v = subst_b (p \<bullet> t1) (p \<bullet> x1) (p \<bullet> v)"
+ by (induct tm rule: b.induct, auto simp add: fresh_at_base subst_bb.simps subst_b_b_def)
+
+ fix bv::bv and c::b and z::bv
+ show "atom bv \<sharp> c \<Longrightarrow> ((bv \<leftrightarrow> z) \<bullet> c) = c[z::=B_var bv]\<^sub>b"
+ by (induct c rule: b.induct, (auto simp add: fresh_at_base subst_bb.simps subst_b_b_def permute_pure pure_supp )+)
+
+ fix bv::bv and c::b and z::bv and v::b
+ show "atom bv \<sharp> c \<Longrightarrow> ((bv \<leftrightarrow> z) \<bullet> c)[bv::=v]\<^sub>b = c[z::=v]\<^sub>b"
+ by (induct c rule: b.induct, (auto simp add: fresh_at_base subst_bb.simps subst_b_b_def permute_pure pure_supp )+)
+qed
+end
+
+lemma subst_bb_inject:
+ assumes "b1 = b2[bv::=b]\<^sub>b\<^sub>b" and "b2 \<noteq> B_var bv"
+ shows
+ "b1 = B_int \<Longrightarrow> b2 = B_int" and
+ "b1 = B_bool \<Longrightarrow> b2 = B_bool" and
+ "b1 = B_unit \<Longrightarrow> b2 = B_unit" and
+ "b1 = B_bitvec \<Longrightarrow> b2 = B_bitvec" and
+ "b1 = B_pair b11 b12 \<Longrightarrow> (\<exists>b11' b12' . b11 = b11'[bv::=b]\<^sub>b\<^sub>b \<and> b12 = b12'[bv::=b]\<^sub>b\<^sub>b \<and> b2 = B_pair b11' b12')" and
+ "b1 = B_var bv' \<Longrightarrow> b2 = B_var bv'" and
+ "b1 = B_id tyid \<Longrightarrow> b2 = B_id tyid" and
+ "b1 = B_app tyid b11 \<Longrightarrow> (\<exists>b11'. b11= b11'[bv::=b]\<^sub>b\<^sub>b \<and> b2 = B_app tyid b11')"
+ using assms by(nominal_induct b2 rule:b.strong_induct,auto+)
+
+lemma flip_b_subst4:
+ fixes b1::b and bv1::bv and c::bv and b::b
+ assumes "atom c \<sharp> (b1,bv1)"
+ shows "b1[bv1::=b]\<^sub>b\<^sub>b = ((bv1 \<leftrightarrow> c) \<bullet> b1)[c ::= b]\<^sub>b\<^sub>b"
+ using assms proof(nominal_induct b1 rule: b.strong_induct)
+ case B_int
+ then show ?case using subst_bb.simps b.perm_simps by auto
+next
+ case B_bool
+ then show ?case using subst_bb.simps b.perm_simps by auto
+next
+ case (B_id x)
+ hence "atom bv1 \<sharp> x \<and> atom c \<sharp> x" using fresh_def pure_supp by auto
+ hence "((bv1 \<leftrightarrow> c) \<bullet> B_id x) = B_id x" using fresh_Pair b.fresh(3) flip_fresh_fresh b.perm_simps fresh_def pure_supp by metis
+ then show ?case using subst_bb.simps by simp
+next
+ case (B_pair x1 x2)
+ hence "x1[bv1::=b]\<^sub>b\<^sub>b = ((bv1 \<leftrightarrow> c) \<bullet> x1)[c::=b]\<^sub>b\<^sub>b" using b.perm_simps(4) b.fresh(4) fresh_Pair by metis
+ moreover have "x2[bv1::=b]\<^sub>b\<^sub>b = ((bv1 \<leftrightarrow> c) \<bullet> x2)[c::=b]\<^sub>b\<^sub>b" using b.perm_simps(4) b.fresh(4) fresh_Pair B_pair by metis
+ ultimately show ?case using subst_bb.simps(5) b.perm_simps(4) b.fresh(4) fresh_Pair by auto
+next
+ case B_unit
+ then show ?case using subst_bb.simps b.perm_simps by auto
+next
+ case B_bitvec
+ then show ?case using subst_bb.simps b.perm_simps by auto
+next
+ case (B_var x)
+ then show ?case proof(cases "x=bv1")
+ case True
+ then show ?thesis using B_var subst_bb.simps b.perm_simps by simp
+ next
+ case False
+ moreover have "x\<noteq>c" using B_var b.fresh fresh_def supp_at_base fresh_Pair by fastforce
+ ultimately show ?thesis using B_var subst_bb.simps(1) b.perm_simps(7) by simp
+ qed
+next
+ case (B_app x1 x2)
+ hence "x2[bv1::=b]\<^sub>b\<^sub>b = ((bv1 \<leftrightarrow> c) \<bullet> x2)[c::=b]\<^sub>b\<^sub>b" using b.perm_simps b.fresh fresh_Pair by metis
+ thus ?case using subst_bb.simps b.perm_simps b.fresh fresh_Pair B_app
+ by (simp add: permute_pure)
+qed
+
+lemma subst_bb_flip_sym:
+ fixes b1::b and b2::b
+ assumes "atom c \<sharp> b" and "atom c \<sharp> (bv1,bv2, b1, b2)" and "(bv1 \<leftrightarrow> c) \<bullet> b1 = (bv2 \<leftrightarrow> c) \<bullet> b2"
+ shows "b1[bv1::=b]\<^sub>b\<^sub>b = b2[bv2::=b]\<^sub>b\<^sub>b"
+ using assms flip_b_subst4[of c b1 bv1 b] flip_b_subst4[of c b2 bv2 b] fresh_prod4 fresh_Pair by simp
+
+section \<open>Value\<close>
+
+nominal_function subst_vb :: "v \<Rightarrow> bv \<Rightarrow> b \<Rightarrow> v" where
+ "subst_vb (V_lit l) x v = V_lit l"
+| "subst_vb (V_var y) x v = V_var y"
+| "subst_vb (V_cons tyid c v') x v = V_cons tyid c (subst_vb v' x v)"
+| "subst_vb (V_consp tyid c b v') x v = V_consp tyid c (subst_bb b x v) (subst_vb v' x v)"
+| "subst_vb (V_pair v1 v2) x v = V_pair (subst_vb v1 x v ) (subst_vb v2 x v )"
+ apply (simp add: eqvt_def subst_vb_graph_aux_def)
+ apply auto
+ using v.strong_exhaust by meson
+nominal_termination (eqvt) by lexicographic_order
+
+abbreviation
+ subst_vb_abbrev :: "v \<Rightarrow> bv \<Rightarrow> b \<Rightarrow> v" ("_[_::=_]\<^sub>v\<^sub>b" [1000,50,50] 500)
+ where
+ "e[bv::=b]\<^sub>v\<^sub>b \<equiv> subst_vb e bv b"
+
+instantiation v :: has_subst_b
+begin
+definition "subst_b = subst_vb"
+
+instance proof
+ fix j::atom and i::bv and x::b and t::v
+ show "j \<sharp> subst_b t i x = (atom i \<sharp> t \<and> j \<sharp> t \<or> j \<sharp> x \<and> (j \<sharp> t \<or> j = atom i))"
+ proof (induct t rule: v.induct)
+ case (V_lit l)
+ have "j \<sharp> subst_b (V_lit l) i x = j \<sharp> (V_lit l)" using subst_vb.simps fresh_def pure_fresh
+ subst_b_v_def v.supp v.fresh has_subst_b_class.fresh_subst_if subst_b_b_def subst_b_v_def by auto
+ also have "... = True" using fresh_at_base v.fresh l.fresh supp_l_empty fresh_def by metis
+ moreover have "(atom i \<sharp> (V_lit l) \<and> j \<sharp> (V_lit l) \<or> j \<sharp> x \<and> (j \<sharp> (V_lit l) \<or> j = atom i)) = True" using fresh_at_base v.fresh l.fresh supp_l_empty fresh_def by metis
+ ultimately show ?case by simp
+ next
+ case (V_var y)
+ then show ?case using subst_b_v_def subst_vb.simps pure_fresh by force
+ next
+ case (V_pair x1a x2a)
+ then show ?case using subst_b_v_def subst_vb.simps by auto
+ next
+ case (V_cons x1a x2a x3)
+ then show ?case using V_cons subst_b_v_def subst_vb.simps pure_fresh by force
+ next
+ case (V_consp x1a x2a x3 x4)
+ then show ?case using subst_b_v_def subst_vb.simps pure_fresh has_subst_b_class.fresh_subst_if subst_b_b_def subst_b_v_def by fastforce
+ qed
+
+ fix a::bv and tm::v and x::b
+ show "atom a \<sharp> tm \<Longrightarrow> subst_b tm a x = tm"
+ apply(induct tm rule: v.induct)
+ apply(auto simp add: fresh_at_base subst_vb.simps subst_b_v_def)
+ using has_subst_b_class.fresh_subst_if subst_b_b_def e.fresh
+ using has_subst_b_class.forget_subst by fastforce
+
+ fix a::bv and tm::v
+ show "subst_b tm a (B_var a) = tm" using subst_bb.simps subst_b_b_def
+ apply (induct tm rule: v.induct)
+ apply(auto simp add: fresh_at_base subst_vb.simps subst_b_v_def)
+ using has_subst_b_class.fresh_subst_if subst_b_b_def e.fresh
+ using has_subst_b_class.subst_id by metis
+
+ fix p::perm and x1::bv and v::b and t1::v
+ show "p \<bullet> subst_b t1 x1 v = subst_b (p \<bullet> t1) (p \<bullet> x1) (p \<bullet> v)"
+ apply(induct tm rule: v.induct)
+ apply(auto simp add: fresh_at_base subst_bb.simps subst_b_b_def )
+ using has_subst_b_class.eqvt subst_b_b_def e.fresh
+ using has_subst_b_class.eqvt
+ by (simp add: subst_b_v_def)+
+
+ fix bv::bv and c::v and z::bv
+ show "atom bv \<sharp> c \<Longrightarrow> ((bv \<leftrightarrow> z) \<bullet> c) = c[z::=B_var bv]\<^sub>b"
+ apply (induct c rule: v.induct, (auto simp add: fresh_at_base subst_vb.simps subst_b_v_def permute_pure pure_supp )+)
+ apply (metis flip_fresh_fresh flip_l_eq permute_flip_cancel2)
+ using fresh_at_base flip_fresh_fresh[of bv x z]
+ apply (simp add: flip_fresh_fresh)
+ using subst_b_b_def by argo
+
+ fix bv::bv and c::v and z::bv and v::b
+ show "atom bv \<sharp> c \<Longrightarrow> ((bv \<leftrightarrow> z) \<bullet> c)[bv::=v]\<^sub>b = c[z::=v]\<^sub>b"
+ apply (induct c rule: v.induct, (auto simp add: fresh_at_base subst_vb.simps subst_b_v_def permute_pure pure_supp )+)
+ apply (metis flip_fresh_fresh flip_l_eq permute_flip_cancel2)
+ using fresh_at_base flip_fresh_fresh[of bv x z]
+ apply (simp add: flip_fresh_fresh)
+ using subst_b_b_def flip_subst_subst by fastforce
+
+qed
+end
+
+section \<open>Constraints Expressions\<close>
+
+nominal_function subst_ceb :: "ce \<Rightarrow> bv \<Rightarrow> b \<Rightarrow> ce" where
+ "subst_ceb ( (CE_val v') ) bv b = ( CE_val (subst_vb v' bv b) )"
+| "subst_ceb ( (CE_op opp v1 v2) ) bv b = ( (CE_op opp (subst_ceb v1 bv b)(subst_ceb v2 bv b)) )"
+| "subst_ceb ( (CE_fst v')) bv b = CE_fst (subst_ceb v' bv b)"
+| "subst_ceb ( (CE_snd v')) bv b = CE_snd (subst_ceb v' bv b)"
+| "subst_ceb ( (CE_len v')) bv b = CE_len (subst_ceb v' bv b)"
+| "subst_ceb ( CE_concat v1 v2) bv b = CE_concat (subst_ceb v1 bv b) (subst_ceb v2 bv b)"
+ apply (simp add: eqvt_def subst_ceb_graph_aux_def)
+ apply auto
+ by (meson ce.strong_exhaust)
+nominal_termination (eqvt) by lexicographic_order
+
+abbreviation
+ subst_ceb_abbrev :: "ce \<Rightarrow> bv \<Rightarrow> b \<Rightarrow> ce" ("_[_::=_]\<^sub>c\<^sub>e\<^sub>b" [1000,50,50] 500)
+ where
+ "ce[bv::=b]\<^sub>c\<^sub>e\<^sub>b \<equiv> subst_ceb ce bv b"
+
+instantiation ce :: has_subst_b
+begin
+definition "subst_b = subst_ceb"
+
+instance proof
+ fix j::atom and i::bv and x::b and t::ce
+ show "j \<sharp> subst_b t i x = (atom i \<sharp> t \<and> j \<sharp> t \<or> j \<sharp> x \<and> (j \<sharp> t \<or> j = atom i))"
+ proof (induct t rule: ce.induct)
+ case (CE_val v)
+ then show ?case using subst_ceb.simps fresh_def pure_fresh subst_b_ce_def ce.supp v.supp ce.fresh has_subst_b_class.fresh_subst_if subst_b_b_def subst_b_v_def
+ by metis
+ next
+ case (CE_op opp v1 v2)
+
+ have "(j \<sharp> v1[i::=x]\<^sub>c\<^sub>e\<^sub>b \<and> j \<sharp> v2[i::=x]\<^sub>c\<^sub>e\<^sub>b) = ((atom i \<sharp> v1 \<and> atom i \<sharp> v2) \<and> j \<sharp> v1 \<and> j \<sharp> v2 \<or> j \<sharp> x \<and> (j \<sharp> v1 \<and> j \<sharp> v2 \<or> j = atom i))"
+ using has_subst_b_class.fresh_subst_if subst_b_v_def
+ using CE_op.hyps(1) CE_op.hyps(2) subst_b_ce_def by auto
+
+ thus ?case unfolding subst_ceb.simps subst_b_ce_def ce.fresh
+ using fresh_def pure_fresh opp.fresh subst_b_v_def opp.exhaust fresh_e_opp_all
+ by (metis (full_types))
+ next
+ case (CE_concat x1a x2)
+ then show ?case using subst_ceb.simps subst_b_ce_def e.supp v.supp has_subst_b_class.fresh_subst_if subst_b_v_def ce.fresh by force
+ next
+ case (CE_fst x)
+ then show ?case using subst_ceb.simps subst_b_ce_def e.supp v.supp has_subst_b_class.fresh_subst_if subst_b_v_def ce.fresh by metis
+
+ next
+ case (CE_snd x)
+ then show ?case using subst_ceb.simps subst_b_ce_def e.supp v.supp has_subst_b_class.fresh_subst_if subst_b_v_def ce.fresh by metis
+ next
+ case (CE_len x)
+ then show ?case using subst_ceb.simps subst_b_ce_def e.supp v.supp has_subst_b_class.fresh_subst_if subst_b_v_def ce.fresh by metis
+ qed
+
+ fix a::bv and tm::ce and x::b
+ show "atom a \<sharp> tm \<Longrightarrow> subst_b tm a x = tm"
+ apply(induct tm rule: ce.induct)
+ apply( auto simp add: fresh_at_base subst_ceb.simps subst_b_ce_def)
+ using has_subst_b_class.fresh_subst_if subst_b_b_def e.fresh
+ using has_subst_b_class.forget_subst subst_b_v_def apply metis+
+ done
+
+ fix a::bv and tm::ce
+ show "subst_b tm a (B_var a) = tm" using subst_bb.simps subst_b_b_def
+ apply (induct tm rule: ce.induct)
+ apply(auto simp add: fresh_at_base subst_ceb.simps subst_b_ce_def)
+ using has_subst_b_class.fresh_subst_if subst_b_b_def e.fresh
+ using has_subst_b_class.subst_id subst_b_v_def apply metis+
+ done
+
+ fix p::perm and x1::bv and v::b and t1::ce
+ show "p \<bullet> subst_b t1 x1 v = subst_b (p \<bullet> t1) (p \<bullet> x1) (p \<bullet> v)"
+ apply(induct tm rule: ce.induct)
+ apply( auto simp add: fresh_at_base subst_bb.simps subst_b_b_def )
+ using has_subst_b_class.eqvt subst_b_b_def ce.fresh
+ using has_subst_b_class.eqvt
+ by (simp add: subst_b_ce_def)+
+
+ fix bv::bv and c::ce and z::bv
+ show "atom bv \<sharp> c \<Longrightarrow> ((bv \<leftrightarrow> z) \<bullet> c) = c[z::=B_var bv]\<^sub>b"
+ apply (induct c rule: ce.induct, (auto simp add: fresh_at_base subst_ceb.simps subst_b_ce_def permute_pure pure_supp )+)
+ using flip_fresh_fresh flip_l_eq permute_flip_cancel2 has_subst_b_class.flip_subst subst_b_v_def apply metis
+ using flip_fresh_fresh flip_l_eq permute_flip_cancel2 has_subst_b_class.flip_subst subst_b_v_def
+ by (simp add: flip_fresh_fresh fresh_opp_all)
+
+ fix bv::bv and c::ce and z::bv and v::b
+ show "atom bv \<sharp> c \<Longrightarrow> ((bv \<leftrightarrow> z) \<bullet> c)[bv::=v]\<^sub>b = c[z::=v]\<^sub>b"
+ proof (induct c rule: ce.induct)
+ case (CE_val x)
+ then show ?case using flip_subst_subst subst_b_v_def subst_ceb.simps using subst_b_ce_def by fastforce
+ next
+ case (CE_op x1a x2 x3)
+ then show ?case unfolding subst_ceb.simps subst_b_ce_def ce.perm_simps using flip_subst_subst subst_b_v_def opp.perm_simps opp.strong_exhaust
+ by (metis (full_types) ce.fresh(2))
+ next
+ case (CE_concat x1a x2)
+ then show ?case using flip_subst_subst subst_b_v_def subst_ceb.simps using subst_b_ce_def by fastforce
+ next
+ case (CE_fst x)
+ then show ?case using flip_subst_subst subst_b_v_def subst_ceb.simps using subst_b_ce_def by fastforce
+ next
+ case (CE_snd x)
+ then show ?case using flip_subst_subst subst_b_v_def subst_ceb.simps using subst_b_ce_def by fastforce
+ next
+ case (CE_len x)
+ then show ?case using flip_subst_subst subst_b_v_def subst_ceb.simps using subst_b_ce_def by fastforce
+ qed
+qed
+end
+
+section \<open>Constraints\<close>
+
+nominal_function subst_cb :: "c \<Rightarrow> bv \<Rightarrow> b \<Rightarrow> c" where
+ "subst_cb (C_true) x v = C_true"
+| "subst_cb (C_false) x v = C_false"
+| "subst_cb (C_conj c1 c2) x v = C_conj (subst_cb c1 x v ) (subst_cb c2 x v )"
+| "subst_cb (C_disj c1 c2) x v = C_disj (subst_cb c1 x v ) (subst_cb c2 x v )"
+| "subst_cb (C_imp c1 c2) x v = C_imp (subst_cb c1 x v ) (subst_cb c2 x v )"
+| "subst_cb (C_eq e1 e2) x v = C_eq (subst_ceb e1 x v ) (subst_ceb e2 x v )"
+| "subst_cb (C_not c) x v = C_not (subst_cb c x v )"
+ apply (simp add: eqvt_def subst_cb_graph_aux_def)
+ apply auto
+ using c.strong_exhaust apply metis
+ done
+nominal_termination (eqvt) by lexicographic_order
+
+abbreviation
+ subst_cb_abbrev :: "c \<Rightarrow> bv \<Rightarrow> b \<Rightarrow> c" ("_[_::=_]\<^sub>c\<^sub>b" [1000,50,50] 500)
+ where
+ "c[bv::=b]\<^sub>c\<^sub>b \<equiv> subst_cb c bv b"
+
+instantiation c :: has_subst_b
+begin
+definition "subst_b = subst_cb"
+
+instance proof
+ fix j::atom and i::bv and x::b and t::c
+ show "j \<sharp> subst_b t i x = (atom i \<sharp> t \<and> j \<sharp> t \<or> j \<sharp> x \<and> (j \<sharp> t \<or> j = atom i))"
+ by (induct t rule: c.induct, unfold subst_cb.simps subst_b_c_def c.fresh,
+ (metis has_subst_b_class.fresh_subst_if subst_b_ce_def c.fresh)+
+ )
+
+ fix a::bv and tm::c and x::b
+ show "atom a \<sharp> tm \<Longrightarrow> subst_b tm a x = tm"
+ by(induct tm rule: c.induct, unfold subst_cb.simps subst_b_c_def c.fresh,
+ (metis has_subst_b_class.forget_subst subst_b_ce_def)+)
+
+ fix a::bv and tm::c
+ show "subst_b tm a (B_var a) = tm" using subst_bb.simps subst_b_c_def
+ by(induct tm rule: c.induct, unfold subst_cb.simps subst_b_c_def c.fresh,
+ (metis has_subst_b_class.subst_id subst_b_ce_def)+)
+
+ fix p::perm and x1::bv and v::b and t1::c
+ show "p \<bullet> subst_b t1 x1 v = subst_b (p \<bullet> t1) (p \<bullet> x1) (p \<bullet> v)"
+ apply(induct tm rule: c.induct,unfold subst_cb.simps subst_b_c_def c.fresh)
+ by( auto simp add: fresh_at_base subst_bb.simps subst_b_b_def )
+
+ fix bv::bv and c::c and z::bv
+ show "atom bv \<sharp> c \<Longrightarrow> ((bv \<leftrightarrow> z) \<bullet> c) = c[z::=B_var bv]\<^sub>b"
+ apply (induct c rule: c.induct, (auto simp add: fresh_at_base subst_cb.simps subst_b_c_def permute_pure pure_supp )+)
+ using flip_fresh_fresh flip_l_eq permute_flip_cancel2 has_subst_b_class.flip_subst subst_b_ce_def apply metis
+ using flip_fresh_fresh flip_l_eq permute_flip_cancel2 has_subst_b_class.flip_subst subst_b_ce_def
+ apply (metis opp.perm_simps(2) opp.strong_exhaust)+
+ done
+
+ fix bv::bv and c::c and z::bv and v::b
+ show "atom bv \<sharp> c \<Longrightarrow> ((bv \<leftrightarrow> z) \<bullet> c)[bv::=v]\<^sub>b = c[z::=v]\<^sub>b"
+ apply (induct c rule: c.induct, (auto simp add: fresh_at_base subst_cb.simps subst_b_c_def permute_pure pure_supp )+)
+ using flip_fresh_fresh flip_l_eq permute_flip_cancel2 has_subst_b_class.flip_subst subst_b_ce_def
+ using flip_subst_subst apply fastforce
+ using flip_fresh_fresh flip_l_eq permute_flip_cancel2 has_subst_b_class.flip_subst subst_b_ce_def
+ opp.perm_simps(2) opp.strong_exhaust
+ proof -
+ fix x1a :: ce and x2 :: ce
+ assume a1: "atom bv \<sharp> x2"
+ then have "((bv \<leftrightarrow> z) \<bullet> x2)[bv::=v]\<^sub>b = x2[z::=v]\<^sub>b"
+ by (metis flip_subst_subst) (* 0.0 ms *)
+ then show "x2[z::=B_var bv]\<^sub>b[bv::=v]\<^sub>c\<^sub>e\<^sub>b = x2[z::=v]\<^sub>c\<^sub>e\<^sub>b"
+ using a1 by (simp add: subst_b_ce_def) (* 0.0 ms *)
+ qed
+
+qed
+end
+
+section \<open>Types\<close>
+
+nominal_function subst_tb :: "\<tau> \<Rightarrow> bv \<Rightarrow> b \<Rightarrow> \<tau>" where
+ "subst_tb (\<lbrace> z : b2 | c \<rbrace>) bv1 b1 = \<lbrace> z : b2[bv1::=b1]\<^sub>b\<^sub>b | c[bv1::=b1]\<^sub>c\<^sub>b \<rbrace>"
+proof(goal_cases)
+ case 1
+ then show ?case using eqvt_def subst_tb_graph_aux_def by force
+next
+ case (2 x y)
+ then show ?case by auto
+next
+ case (3 P x)
+ then show ?case using eqvt_def subst_tb_graph_aux_def \<tau>.strong_exhaust
+ by (metis b_of.cases prod_cases3)
+next
+ case (4 z' b2' c' bv1' b1' z b2 c bv1 b1)
+ show ?case unfolding \<tau>.eq_iff proof
+ have *:"[[atom z']]lst. c' = [[atom z]]lst. c" using \<tau>.eq_iff 4 by auto
+ show "[[atom z']]lst. c'[bv1'::=b1']\<^sub>c\<^sub>b = [[atom z]]lst. c[bv1::=b1]\<^sub>c\<^sub>b"
+ proof(subst Abs1_eq_iff_all(3),rule,rule,rule)
+ fix ca::x
+ assume "atom ca \<sharp> z" and 1:"atom ca \<sharp> (z', z, c'[bv1'::=b1']\<^sub>c\<^sub>b, c[bv1::=b1]\<^sub>c\<^sub>b)"
+ hence 2:"atom ca \<sharp> (c',c)" using fresh_subst_if subst_b_c_def fresh_Pair fresh_prod4 fresh_at_base subst_b_fresh_x by metis
+ hence "(z' \<leftrightarrow> ca) \<bullet> c' = (z \<leftrightarrow> ca) \<bullet> c" using 1 2 * Abs1_eq_iff_all(3) by auto
+ hence "((z' \<leftrightarrow> ca) \<bullet> c')[bv1'::=b1']\<^sub>c\<^sub>b = ((z \<leftrightarrow> ca) \<bullet> c)[bv1'::=b1']\<^sub>c\<^sub>b" by auto
+ hence "(z' \<leftrightarrow> ca) \<bullet> c'[(z' \<leftrightarrow> ca) \<bullet> bv1'::=(z' \<leftrightarrow> ca) \<bullet> b1']\<^sub>c\<^sub>b = (z \<leftrightarrow> ca) \<bullet> c[(z \<leftrightarrow> ca) \<bullet> bv1'::=(z \<leftrightarrow> ca) \<bullet> b1']\<^sub>c\<^sub>b" by auto
+ thus "(z' \<leftrightarrow> ca) \<bullet> c'[bv1'::=b1']\<^sub>c\<^sub>b = (z \<leftrightarrow> ca) \<bullet> c[bv1::=b1]\<^sub>c\<^sub>b" using 4 flip_x_b_cancel by simp
+ qed
+ show "b2'[bv1'::=b1']\<^sub>b\<^sub>b = b2[bv1::=b1]\<^sub>b\<^sub>b" using 4 by simp
+ qed
+qed
+
+nominal_termination (eqvt) by lexicographic_order
+
+abbreviation
+ subst_tb_abbrev :: "\<tau> \<Rightarrow> bv \<Rightarrow> b \<Rightarrow> \<tau>" ("_[_::=_]\<^sub>\<tau>\<^sub>b" [1000,50,50] 1000)
+ where
+ "t[bv::=b']\<^sub>\<tau>\<^sub>b \<equiv> subst_tb t bv b' "
+
+instantiation \<tau> :: has_subst_b
+begin
+definition "subst_b = subst_tb"
+
+instance proof
+ fix j::atom and i::bv and x::b and t::\<tau>
+ show "j \<sharp> subst_b t i x = (atom i \<sharp> t \<and> j \<sharp> t \<or> j \<sharp> x \<and> (j \<sharp> t \<or> j = atom i))"
+ proof (nominal_induct t avoiding: i x j rule: \<tau>.strong_induct)
+ case (T_refined_type z b c)
+ then show ?case
+ unfolding subst_b_\<tau>_def subst_tb.simps \<tau>.fresh
+ using fresh_subst_if[of j b i x ] subst_b_b_def subst_b_c_def
+ by (metis has_subst_b_class.fresh_subst_if list.distinct(1) list.set_cases not_self_fresh set_ConsD)
+ qed
+
+ fix a::bv and tm::\<tau> and x::b
+ show "atom a \<sharp> tm \<Longrightarrow> subst_b tm a x = tm"
+ proof (nominal_induct tm avoiding: a x rule: \<tau>.strong_induct)
+ case (T_refined_type xx bb cc )
+ moreover hence "atom a \<sharp> bb \<and> atom a \<sharp> cc" using \<tau>.fresh by auto
+ ultimately show ?case
+ unfolding subst_b_\<tau>_def subst_tb.simps
+ using forget_subst subst_b_b_def subst_b_c_def forget_subst \<tau>.fresh by metis
+ qed
+
+ fix a::bv and tm::\<tau>
+ show "subst_b tm a (B_var a) = tm"
+ proof (nominal_induct tm rule: \<tau>.strong_induct)
+ case (T_refined_type xx bb cc)
+ thus ?case
+ unfolding subst_b_\<tau>_def subst_tb.simps
+ using subst_id subst_b_b_def subst_b_c_def by metis
+ qed
+
+ fix p::perm and x1::bv and v::b and t1::\<tau>
+ show "p \<bullet> subst_b t1 x1 v = subst_b (p \<bullet> t1) (p \<bullet> x1) (p \<bullet> v) "
+ by (induct tm rule: \<tau>.induct, auto simp add: fresh_at_base subst_tb.simps subst_b_\<tau>_def subst_bb.simps subst_b_b_def)
+
+ fix bv::bv and c::\<tau> and z::bv
+ show "atom bv \<sharp> c \<Longrightarrow> ((bv \<leftrightarrow> z) \<bullet> c) = c[z::=B_var bv]\<^sub>b"
+ apply (induct c rule: \<tau>.induct, (auto simp add: fresh_at_base subst_ceb.simps subst_b_ce_def permute_pure pure_supp )+)
+ using flip_fresh_fresh permute_flip_cancel2 has_subst_b_class.flip_subst subst_b_c_def subst_b_b_def
+ by (simp add: flip_fresh_fresh subst_b_\<tau>_def)
+
+ fix bv::bv and c::\<tau> and z::bv and v::b
+ show "atom bv \<sharp> c \<Longrightarrow> ((bv \<leftrightarrow> z) \<bullet> c)[bv::=v]\<^sub>b = c[z::=v]\<^sub>b"
+ proof (induct c rule: \<tau>.induct)
+ case (T_refined_type x1a x2a x3a)
+ hence "atom bv \<sharp> x2a \<and> atom bv \<sharp> x3a \<and> atom bv \<sharp> x1a" using fresh_at_base \<tau>.fresh by simp
+ then show ?case
+ unfolding subst_tb.simps subst_b_\<tau>_def \<tau>.perm_simps
+ using fresh_at_base flip_fresh_fresh[of bv x1a z] flip_subst_subst subst_b_b_def subst_b_c_def T_refined_type
+ proof -
+ have "atom z \<sharp> x1a"
+ by (metis b.fresh(7) fresh_at_base(2) x_fresh_b) (* 0.0 ms *)
+ then show "\<lbrace> (bv \<leftrightarrow> z) \<bullet> x1a : ((bv \<leftrightarrow> z) \<bullet> x2a)[bv::=v]\<^sub>b\<^sub>b | ((bv \<leftrightarrow> z) \<bullet> x3a)[bv::=v]\<^sub>c\<^sub>b \<rbrace> = \<lbrace> x1a : x2a[z::=v]\<^sub>b\<^sub>b | x3a[z::=v]\<^sub>c\<^sub>b \<rbrace>"
+ by (metis \<open>\<lbrakk>atom bv \<sharp> x1a; atom z \<sharp> x1a\<rbrakk> \<Longrightarrow> (bv \<leftrightarrow> z) \<bullet> x1a = x1a\<close> \<open>atom bv \<sharp> x2a \<and> atom bv \<sharp> x3a \<and> atom bv \<sharp> x1a\<close> flip_subst_subst subst_b_b_def subst_b_c_def) (* 78 ms *)
+ qed
+ qed
+
+qed
+end
+
+lemma subst_bb_commute [simp]:
+ "atom j \<sharp> A \<Longrightarrow> (subst_bb (subst_bb A i t ) j u ) = subst_bb A i (subst_bb t j u) "
+ by (nominal_induct A avoiding: i j t u rule: b.strong_induct) (auto simp: fresh_at_base)
+
+lemma subst_vb_commute [simp]:
+ "atom j \<sharp> A \<Longrightarrow> (subst_vb (subst_vb A i t )) j u = subst_vb A i (subst_bb t j u ) "
+ by (nominal_induct A avoiding: i j t u rule: v.strong_induct) (auto simp: fresh_at_base)
+
+lemma subst_ceb_commute [simp]:
+ "atom j \<sharp> A \<Longrightarrow> (subst_ceb (subst_ceb A i t )) j u = subst_ceb A i (subst_bb t j u ) "
+ by (nominal_induct A avoiding: i j t u rule: ce.strong_induct) (auto simp: fresh_at_base)
+
+lemma subst_cb_commute [simp]:
+ "atom j \<sharp> A \<Longrightarrow> (subst_cb (subst_cb A i t )) j u = subst_cb A i (subst_bb t j u ) "
+ by (nominal_induct A avoiding: i j t u rule: c.strong_induct) (auto simp: fresh_at_base)
+
+lemma subst_tb_commute [simp]:
+ "atom j \<sharp> A \<Longrightarrow> (subst_tb (subst_tb A i t )) j u = subst_tb A i (subst_bb t j u ) "
+proof (nominal_induct A avoiding: i j t u rule: \<tau>.strong_induct)
+ case (T_refined_type z b c)
+ then show ?case using subst_tb.simps subst_bb_commute subst_cb_commute by simp
+qed
+
+section \<open>Expressions\<close>
+
+nominal_function subst_eb :: "e \<Rightarrow> bv \<Rightarrow> b \<Rightarrow> e" where
+ "subst_eb ( (AE_val v')) bv b = ( AE_val (subst_vb v' bv b))"
+| "subst_eb ( (AE_app f v') ) bv b = ( (AE_app f (subst_vb v' bv b)) )"
+| "subst_eb ( (AE_appP f b' v') ) bv b = ( (AE_appP f (b'[bv::=b]\<^sub>b\<^sub>b) (subst_vb v' bv b)))"
+| "subst_eb ( (AE_op opp v1 v2) ) bv b = ( (AE_op opp (subst_vb v1 bv b) (subst_vb v2 bv b)) )"
+| "subst_eb ( (AE_fst v')) bv b = AE_fst (subst_vb v' bv b)"
+| "subst_eb ( (AE_snd v')) bv b = AE_snd (subst_vb v' bv b)"
+| "subst_eb ( (AE_mvar u)) bv b = AE_mvar u"
+| "subst_eb ( (AE_len v')) bv b = AE_len (subst_vb v' bv b)"
+| "subst_eb ( AE_concat v1 v2) bv b = AE_concat (subst_vb v1 bv b) (subst_vb v2 bv b)"
+| "subst_eb ( AE_split v1 v2) bv b = AE_split (subst_vb v1 bv b) (subst_vb v2 bv b)"
+ apply (simp add: eqvt_def subst_eb_graph_aux_def)
+ apply auto
+ by (meson e.strong_exhaust)
+nominal_termination (eqvt) by lexicographic_order
+
+abbreviation
+ subst_eb_abbrev :: "e \<Rightarrow> bv \<Rightarrow> b \<Rightarrow> e" ("_[_::=_]\<^sub>e\<^sub>b" [1000,50,50] 500)
+ where
+ "e[bv::=b]\<^sub>e\<^sub>b \<equiv> subst_eb e bv b"
+
+instantiation e :: has_subst_b
+begin
+definition "subst_b = subst_eb"
+
+instance proof
+ fix j::atom and i::bv and x::b and t::e
+ show "j \<sharp> subst_b t i x = (atom i \<sharp> t \<and> j \<sharp> t \<or> j \<sharp> x \<and> (j \<sharp> t \<or> j = atom i))"
+ proof (induct t rule: e.induct)
+ case (AE_val v)
+ then show ?case using subst_eb.simps fresh_def pure_fresh subst_b_e_def e.supp v.supp
+ e.fresh has_subst_b_class.fresh_subst_if subst_b_e_def subst_b_v_def
+ by metis
+ next
+ case (AE_app f v)
+ then show ?case using subst_eb.simps fresh_def pure_fresh subst_b_e_def
+ e.supp v.supp has_subst_b_class.fresh_subst_if subst_b_v_def
+ by (metis (mono_tags, opaque_lifting) e.fresh(2))
+ next
+ case (AE_appP f b' v)
+ then show ?case unfolding subst_eb.simps subst_b_e_def e.fresh using
+ fresh_def pure_fresh subst_b_e_def e.supp v.supp
+ e.fresh has_subst_b_class.fresh_subst_if subst_b_b_def subst_vb_def by (metis subst_b_v_def)
+ next
+ case (AE_op opp v1 v2)
+ then show ?case unfolding subst_eb.simps subst_b_e_def e.fresh using
+ fresh_def pure_fresh subst_b_e_def e.supp v.supp fresh_e_opp_all
+ e.fresh has_subst_b_class.fresh_subst_if subst_b_b_def subst_vb_def by (metis subst_b_v_def)
+ next
+ case (AE_concat x1a x2)
+ then show ?case using subst_eb.simps fresh_def pure_fresh subst_b_e_def e.supp v.supp
+ has_subst_b_class.fresh_subst_if subst_b_v_def
+ by (metis subst_vb.simps(5))
+ next
+ case (AE_split x1a x2)
+ then show ?case using subst_eb.simps fresh_def pure_fresh subst_b_e_def e.supp v.supp
+ has_subst_b_class.fresh_subst_if subst_b_v_def
+ by (metis subst_vb.simps(5))
+ next
+ case (AE_fst x)
+ then show ?case using subst_eb.simps fresh_def pure_fresh subst_b_e_def e.supp v.supp has_subst_b_class.fresh_subst_if subst_b_v_def by metis
+ next
+ case (AE_snd x)
+ then show ?case using subst_eb.simps fresh_def pure_fresh subst_b_e_def e.supp v.supp using has_subst_b_class.fresh_subst_if subst_b_v_def by metis
+ next
+ case (AE_mvar x)
+ then show ?case using subst_eb.simps fresh_def pure_fresh subst_b_e_def e.supp v.supp by auto
+ next
+ case (AE_len x)
+ then show ?case using subst_eb.simps fresh_def pure_fresh subst_b_e_def e.supp v.supp using has_subst_b_class.fresh_subst_if subst_b_v_def by metis
+ qed
+
+ fix a::bv and tm::e and x::b
+ show "atom a \<sharp> tm \<Longrightarrow> subst_b tm a x = tm"
+ apply(induct tm rule: e.induct)
+ apply( auto simp add: fresh_at_base subst_eb.simps subst_b_e_def)
+ using has_subst_b_class.fresh_subst_if subst_b_b_def e.fresh
+ using has_subst_b_class.forget_subst subst_b_v_def apply metis+
+ done
+
+ fix a::bv and tm::e
+ show "subst_b tm a (B_var a) = tm" using subst_bb.simps subst_b_b_def
+ apply (induct tm rule: e.induct)
+ apply(auto simp add: fresh_at_base subst_eb.simps subst_b_e_def)
+ using has_subst_b_class.fresh_subst_if subst_b_b_def e.fresh
+ using has_subst_b_class.subst_id subst_b_v_def apply metis+
+ done
+
+ fix p::perm and x1::bv and v::b and t1::e
+ show "p \<bullet> subst_b t1 x1 v = subst_b (p \<bullet> t1) (p \<bullet> x1) (p \<bullet> v)"
+ apply(induct tm rule: e.induct)
+ apply( auto simp add: fresh_at_base subst_bb.simps subst_b_b_def )
+ using has_subst_b_class.eqvt subst_b_b_def e.fresh
+ using has_subst_b_class.eqvt
+ by (simp add: subst_b_e_def)+
+
+ fix bv::bv and c::e and z::bv
+ show "atom bv \<sharp> c \<Longrightarrow> ((bv \<leftrightarrow> z) \<bullet> c) = c[z::=B_var bv]\<^sub>b"
+ apply (induct c rule: e.induct)
+ apply(auto simp add: fresh_at_base subst_eb.simps subst_b_e_def subst_b_v_def permute_pure pure_supp )
+ using flip_fresh_fresh permute_flip_cancel2 has_subst_b_class.flip_subst subst_b_v_def subst_b_b_def
+ flip_fresh_fresh subst_b_\<tau>_def apply metis
+ apply (metis (full_types) opp.perm_simps opp.strong_exhaust)
+ done
+
+ fix bv::bv and c::e and z::bv and v::b
+ show "atom bv \<sharp> c \<Longrightarrow> ((bv \<leftrightarrow> z) \<bullet> c)[bv::=v]\<^sub>b = c[z::=v]\<^sub>b"
+ apply (induct c rule: e.induct)
+ apply(auto simp add: fresh_at_base subst_eb.simps subst_b_e_def subst_b_v_def permute_pure pure_supp )
+ using flip_fresh_fresh permute_flip_cancel2 has_subst_b_class.flip_subst subst_b_v_def subst_b_b_def
+ flip_fresh_fresh subst_b_\<tau>_def apply simp
+
+ apply (metis (full_types) opp.perm_simps opp.strong_exhaust)
+ done
+qed
+end
+
+section \<open>Statements\<close>
+
+nominal_function (default "case_sum (\<lambda>x. Inl undefined) (case_sum (\<lambda>x. Inl undefined) (\<lambda>x. Inr undefined))")
+ subst_sb :: "s \<Rightarrow> bv \<Rightarrow> b \<Rightarrow> s"
+ and subst_branchb :: "branch_s \<Rightarrow> bv \<Rightarrow> b \<Rightarrow> branch_s"
+ and subst_branchlb :: "branch_list \<Rightarrow> bv \<Rightarrow> b \<Rightarrow> branch_list"
+ where
+ "subst_sb (AS_val v') bv b = (AS_val (subst_vb v' bv b))"
+ | "subst_sb (AS_let y e s) bv b = (AS_let y (e[bv::=b]\<^sub>e\<^sub>b) (subst_sb s bv b ))"
+ | "subst_sb (AS_let2 y t s1 s2) bv b = (AS_let2 y (subst_tb t bv b) (subst_sb s1 bv b ) (subst_sb s2 bv b))"
+ | "subst_sb (AS_match v' cs) bv b = AS_match (subst_vb v' bv b) (subst_branchlb cs bv b)"
+ | "subst_sb (AS_assign y v') bv b = AS_assign y (subst_vb v' bv b)"
+ | "subst_sb (AS_if v' s1 s2) bv b = (AS_if (subst_vb v' bv b) (subst_sb s1 bv b ) (subst_sb s2 bv b ) )"
+ | "subst_sb (AS_var u \<tau> v' s) bv b = AS_var u (subst_tb \<tau> bv b) (subst_vb v' bv b) (subst_sb s bv b )"
+ | "subst_sb (AS_while s1 s2) bv b = AS_while (subst_sb s1 bv b ) (subst_sb s2 bv b )"
+ | "subst_sb (AS_seq s1 s2) bv b = AS_seq (subst_sb s1 bv b ) (subst_sb s2 bv b )"
+ | "subst_sb (AS_assert c s) bv b = AS_assert (subst_cb c bv b ) (subst_sb s bv b )"
+
+| "subst_branchb (AS_branch dc x1 s') bv b = AS_branch dc x1 (subst_sb s' bv b)"
+
+| "subst_branchlb (AS_final sb) bv b = AS_final (subst_branchb sb bv b )"
+| "subst_branchlb (AS_cons sb ssb) bv b = AS_cons (subst_branchb sb bv b) (subst_branchlb ssb bv b)"
+
+ apply (simp add: eqvt_def subst_sb_subst_branchb_subst_branchlb_graph_aux_def )
+ apply (auto,metis s_branch_s_branch_list.exhaust s_branch_s_branch_list.exhaust(2) old.sum.exhaust surj_pair)
+
+proof(goal_cases)
+
+ have eqvt_at_proj: "\<And> s xa va . eqvt_at subst_sb_subst_branchb_subst_branchlb_sumC (Inl (s, xa, va)) \<Longrightarrow>
+ eqvt_at (\<lambda>a. projl (subst_sb_subst_branchb_subst_branchlb_sumC (Inl a))) (s, xa, va)"
+ apply(simp only: eqvt_at_def)
+ apply(rule)
+ apply(subst Projl_permute)
+ apply(thin_tac _)+
+ apply(simp add: subst_sb_subst_branchb_subst_branchlb_sumC_def)
+ apply(simp add: THE_default_def)
+ apply(case_tac "Ex1 (subst_sb_subst_branchb_subst_branchlb_graph (Inl (s,xa,va)))")
+ apply simp
+ apply(auto)[1]
+ apply(erule_tac x="x" in allE)
+ apply simp
+ apply(cases rule: subst_sb_subst_branchb_subst_branchlb_graph.cases)
+ apply(assumption)
+ apply(rule_tac x="Sum_Type.projl x" in exI,clarify,rule the1_equality,blast,simp (no_asm) only: sum.sel)+
+ apply(blast)+
+ apply(simp)+
+ done
+ {
+ case (1 y s bv b ya sa c)
+ moreover have "atom y \<sharp> (bv, b) \<and> atom ya \<sharp> (bv, b)" using x_fresh_b x_fresh_bv fresh_Pair by simp
+ ultimately show ?case
+ using eqvt_triple eqvt_at_proj by metis
+ next
+ case (2 y s1 s2 bv b ya s2a c)
+ moreover have "atom y \<sharp> (bv, b) \<and> atom ya \<sharp> (bv, b)" using x_fresh_b x_fresh_bv fresh_Pair by simp
+ ultimately show ?case
+ using eqvt_triple eqvt_at_proj by metis
+ next
+ case (3 u s bv b ua sa c)
+ moreover have "atom u \<sharp> (bv, b) \<and> atom ua \<sharp> (bv, b)" using x_fresh_b x_fresh_bv fresh_Pair by simp
+ ultimately show ?case using eqvt_triple eqvt_at_proj by metis
+ next
+ case (4 x1 s' bv b x1a s'a c)
+ moreover have "atom x1 \<sharp> (bv, b) \<and> atom x1a \<sharp> (bv, b)" using x_fresh_b x_fresh_bv fresh_Pair by simp
+ ultimately show ?case using eqvt_triple eqvt_at_proj by metis
+ }
+qed
+
+nominal_termination (eqvt) by lexicographic_order
+
+abbreviation
+ subst_sb_abbrev :: "s \<Rightarrow> bv \<Rightarrow> b \<Rightarrow> s" ("_[_::=_]\<^sub>s\<^sub>b" [1000,50,50] 1000)
+ where
+ "b[bv::=b']\<^sub>s\<^sub>b \<equiv> subst_sb b bv b'"
+
+lemma fresh_subst_sb_if [simp]:
+ "(j \<sharp> (subst_sb A i x )) = ((atom i \<sharp> A \<and> j \<sharp> A) \<or> (j \<sharp> x \<and> (j \<sharp> A \<or> j = atom i)))" and
+ "(j \<sharp> (subst_branchb B i x )) = ((atom i \<sharp> B \<and> j \<sharp> B) \<or> (j \<sharp> x \<and> (j \<sharp> B \<or> j = atom i)))" and
+ "(j \<sharp> (subst_branchlb C i x )) = ((atom i \<sharp> C \<and> j \<sharp> C) \<or> (j \<sharp> x \<and> (j \<sharp> C \<or> j = atom i)))"
+proof (nominal_induct A and B and C avoiding: i x rule: s_branch_s_branch_list.strong_induct)
+ case (AS_branch x1 x2 x3)
+ have " (j \<sharp> subst_branchb (AS_branch x1 x2 x3) i x ) = (j \<sharp> (AS_branch x1 x2 (subst_sb x3 i x)))" by auto
+ also have "... = ((j \<sharp> x3[i::=x]\<^sub>s\<^sub>b \<or> j \<in> set [atom x2]) \<and> j \<sharp> x1)" using s_branch_s_branch_list.fresh by auto
+ also have "... = ((atom i \<sharp> AS_branch x1 x2 x3 \<and> j \<sharp> AS_branch x1 x2 x3) \<or> j \<sharp> x \<and> (j \<sharp> AS_branch x1 x2 x3 \<or> j = atom i))"
+ using subst_branchb.simps(1) s_branch_s_branch_list.fresh(1) fresh_at_base has_subst_b_class.fresh_subst_if list.distinct list.set_cases set_ConsD subst_b_\<tau>_def
+ v.fresh AS_branch
+ proof -
+ have f1: "\<forall>cs b. atom (b::bv) \<sharp> (cs::char list)" using pure_fresh by auto
+
+ then have "j \<sharp> x \<and> atom i = j \<longrightarrow> ((j \<sharp> x3[i::=x]\<^sub>s\<^sub>b \<or> j \<in> set [atom x2]) \<and> j \<sharp> x1) = (atom i \<sharp> AS_branch x1 x2 x3 \<and> j \<sharp> AS_branch x1 x2 x3 \<or> j \<sharp> x \<and> (j \<sharp> AS_branch x1 x2 x3 \<or> j = atom i))"
+ by (metis (full_types) AS_branch.hyps(3))
+ then have "j \<sharp> x \<longrightarrow> ((j \<sharp> x3[i::=x]\<^sub>s\<^sub>b \<or> j \<in> set [atom x2]) \<and> j \<sharp> x1) = (atom i \<sharp> AS_branch x1 x2 x3 \<and> j \<sharp> AS_branch x1 x2 x3 \<or> j \<sharp> x \<and> (j \<sharp> AS_branch x1 x2 x3 \<or> j = atom i))"
+ using AS_branch.hyps s_branch_s_branch_list.fresh by metis
+ moreover
+ { assume "\<not> j \<sharp> x"
+ have ?thesis
+ using f1 AS_branch.hyps(2) AS_branch.hyps(3) by force }
+ ultimately show ?thesis
+ by satx
+ qed
+ finally show ?case by auto
+next
+ case (AS_cons cs css i x)
+ show ?case
+ unfolding subst_branchlb.simps s_branch_s_branch_list.fresh
+ using AS_cons by auto
+next
+ case (AS_val xx)
+ then show ?case using subst_sb.simps(1) s_branch_s_branch_list.fresh has_subst_b_class.fresh_subst_if subst_b_b_def subst_b_v_def by metis
+next
+ case (AS_let x1 x2 x3)
+ then show ?case using subst_sb.simps s_branch_s_branch_list.fresh fresh_at_base has_subst_b_class.fresh_subst_if list.distinct list.set_cases set_ConsD subst_b_e_def
+ by fastforce
+next
+ case (AS_let2 x1 x2 x3 x4)
+ then show ?case using subst_sb.simps s_branch_s_branch_list.fresh fresh_at_base has_subst_b_class.fresh_subst_if list.distinct list.set_cases set_ConsD subst_b_\<tau>_def
+ by fastforce
+next
+ case (AS_if x1 x2 x3)
+ then show ?case unfolding subst_sb.simps s_branch_s_branch_list.fresh using
+ has_subst_b_class.fresh_subst_if subst_b_v_def by metis
+next
+ case (AS_var u t v s)
+
+ have "(((atom i \<sharp> s \<and> j \<sharp> s \<or> j \<sharp> x \<and> (j \<sharp> s \<or> j = atom i)) \<or> j \<in> set [atom u]) \<and> j \<sharp> t[i::=x]\<^sub>\<tau>\<^sub>b \<and> j \<sharp> v[i::=x]\<^sub>v\<^sub>b) =
+ (((atom i \<sharp> s \<and> j \<sharp> s \<or> j \<sharp> x \<and> (j \<sharp> s \<or> j = atom i)) \<or> j \<in> set [atom u]) \<and>
+ ((atom i \<sharp> t \<and> j \<sharp> t \<or> j \<sharp> x \<and> (j \<sharp> t \<or> j = atom i))) \<and>
+ ((atom i \<sharp> v \<and> j \<sharp> v \<or> j \<sharp> x \<and> (j \<sharp> v \<or> j = atom i))))"
+ using has_subst_b_class.fresh_subst_if subst_b_v_def subst_b_\<tau>_def by metis
+ also have "... = (((atom i \<sharp> s \<or> atom i \<in> set [atom u]) \<and> atom i \<sharp> t \<and> atom i \<sharp> v) \<and>
+ (j \<sharp> s \<or> j \<in> set [atom u]) \<and> j \<sharp> t \<and> j \<sharp> v \<or> j \<sharp> x \<and> ((j \<sharp> s \<or> j \<in> set [atom u]) \<and> j \<sharp> t \<and> j \<sharp> v \<or> j = atom i))"
+ using u_fresh_b by auto
+ finally show ?case using subst_sb.simps s_branch_s_branch_list.fresh AS_var
+ by simp
+next
+ case (AS_assign u v )
+ then show ?case unfolding subst_sb.simps s_branch_s_branch_list.fresh using
+ has_subst_b_class.fresh_subst_if subst_b_v_def by force
+next
+ case (AS_match v cs)
+ have " j \<sharp> (AS_match v cs)[i::=x]\<^sub>s\<^sub>b = j \<sharp> (AS_match (subst_vb v i x) (subst_branchlb cs i x ))" using subst_sb.simps by auto
+ also have "... = (j \<sharp> (subst_vb v i x) \<and> j \<sharp> (subst_branchlb cs i x ))" using s_branch_s_branch_list.fresh by simp
+ also have "... = (j \<sharp> (subst_vb v i x) \<and> ((atom i \<sharp> cs \<and> j \<sharp> cs) \<or> j \<sharp> x \<and> (j \<sharp> cs \<or> j = atom i)))" using AS_match[of i x] by auto
+ also have "... = (atom i \<sharp> AS_match v cs \<and> j \<sharp> AS_match v cs \<or> j \<sharp> x \<and> (j \<sharp> AS_match v cs \<or> j = atom i))"
+ by (metis (no_types) s_branch_s_branch_list.fresh has_subst_b_class.fresh_subst_if subst_b_v_def) (* 93 ms *)
+ finally show ?case by auto
+next
+ case (AS_while x1 x2)
+ then show ?case by auto
+next
+ case (AS_seq x1 x2)
+ then show ?case by auto
+next
+ case (AS_assert x1 x2)
+ then show ?case unfolding subst_sb.simps s_branch_s_branch_list.fresh
+ using fresh_at_base has_subst_b_class.fresh_subst_if list.distinct list.set_cases set_ConsD subst_b_e_def
+ by (metis subst_b_c_def)
+qed(auto+)
+
+lemma
+ forget_subst_sb[simp]: "atom a \<sharp> A \<Longrightarrow> subst_sb A a x = A" and
+ forget_subst_branchb [simp]: "atom a \<sharp> B \<Longrightarrow> subst_branchb B a x = B" and
+ forget_subst_branchlb[simp]: "atom a \<sharp> C \<Longrightarrow> subst_branchlb C a x = C"
+proof (nominal_induct A and B and C avoiding: a x rule: s_branch_s_branch_list.strong_induct)
+ case (AS_let x1 x2 x3)
+ then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_e_def has_subst_b_class.forget_subst subst_b_v_def by force
+next
+ case (AS_let2 x1 x2 x3 x4)
+ then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_e_def has_subst_b_class.forget_subst subst_b_\<tau>_def by force
+next
+ case (AS_var x1 x2 x3 x4)
+ then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_e_def has_subst_b_class.forget_subst subst_b_v_def using subst_b_\<tau>_def
+ proof -
+ have f1: "(atom a \<sharp> x4 \<or> atom a \<in> set [atom x1]) \<and> atom a \<sharp> x2 \<and> atom a \<sharp> x3"
+ using AS_var.prems s_branch_s_branch_list.fresh by simp
+ then have "atom a \<sharp> x4"
+ by (metis (no_types) "Nominal-Utils.fresh_star_singleton" AS_var.hyps(1) empty_set fresh_star_def list.simps(15) not_self_fresh) (* 46 ms *)
+ then show ?thesis
+ using f1 by (metis AS_var.hyps(3) has_subst_b_class.forget_subst subst_b_\<tau>_def subst_b_v_def subst_sb.simps(7)) (* 62 ms *)
+ qed
+next
+ case (AS_branch x1 x2 x3)
+ then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_e_def has_subst_b_class.forget_subst subst_b_v_def by force
+next
+ case (AS_cons x1 x2 x3 x4)
+ then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_e_def has_subst_b_class.forget_subst subst_b_v_def by force
+next
+ case (AS_val x)
+ then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_e_def has_subst_b_class.forget_subst subst_b_v_def by force
+next
+ case (AS_if x1 x2 x3)
+ then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_e_def has_subst_b_class.forget_subst subst_b_v_def by force
+next
+ case (AS_assign x1 x2)
+ then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_e_def has_subst_b_class.forget_subst subst_b_v_def by force
+next
+ case (AS_match x1 x2)
+ then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_e_def has_subst_b_class.forget_subst subst_b_v_def by force
+next
+ case (AS_while x1 x2)
+ then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_e_def has_subst_b_class.forget_subst subst_b_v_def by force
+next
+ case (AS_seq x1 x2)
+ then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_e_def has_subst_b_class.forget_subst subst_b_v_def by force
+next
+ case (AS_assert c s)
+ then show ?case unfolding subst_sb.simps using
+ s_branch_s_branch_list.fresh subst_b_e_def has_subst_b_class.forget_subst subst_b_v_def subst_b_c_def subst_cb.simps by force
+qed(auto+)
+
+lemma subst_sb_id: "subst_sb A a (B_var a) = A" and
+ subst_branchb_id [simp]: "subst_branchb B a (B_var a) = B" and
+ subst_branchlb_id: "subst_branchlb C a (B_var a) = C"
+proof(nominal_induct A and B and C avoiding: a rule: s_branch_s_branch_list.strong_induct)
+ case (AS_branch x1 x2 x3)
+ then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_\<tau>_def has_subst_b_class.subst_id subst_b_v_def
+ by simp
+next
+ case (AS_cons x1 x2 )
+ then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_\<tau>_def has_subst_b_class.subst_id subst_b_v_def by simp
+next
+ case (AS_val x)
+ then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_\<tau>_def has_subst_b_class.subst_id subst_b_v_def by metis
+next
+ case (AS_if x1 x2 x3)
+ then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_\<tau>_def has_subst_b_class.subst_id subst_b_v_def by metis
+next
+ case (AS_assign x1 x2)
+ then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_\<tau>_def has_subst_b_class.subst_id subst_b_v_def by metis
+next
+ case (AS_match x1 x2)
+ then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_\<tau>_def has_subst_b_class.subst_id subst_b_v_def by metis
+next
+ case (AS_while x1 x2)
+ then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_\<tau>_def has_subst_b_class.subst_id subst_b_v_def by metis
+next
+ case (AS_seq x1 x2)
+ then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_\<tau>_def has_subst_b_class.subst_id subst_b_v_def by metis
+next
+ case (AS_let x1 x2 x3)
+ then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_e_def has_subst_b_class.subst_id by metis
+next
+ case (AS_let2 x1 x2 x3 x4)
+ then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_\<tau>_def has_subst_b_class.subst_id by metis
+next
+ case (AS_var x1 x2 x3 x4)
+ then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_\<tau>_def has_subst_b_class.subst_id subst_b_v_def by metis
+next
+ case (AS_assert c s )
+ then show ?case unfolding subst_sb.simps using s_branch_s_branch_list.fresh subst_b_c_def has_subst_b_class.subst_id by metis
+qed (auto)
+
+lemma flip_subst_s:
+ fixes bv::bv and s::s and cs::branch_s and z::bv
+ shows "atom bv \<sharp> s \<Longrightarrow> ((bv \<leftrightarrow> z) \<bullet> s) = s[z::=B_var bv]\<^sub>s\<^sub>b" and
+ "atom bv \<sharp> cs \<Longrightarrow> ((bv \<leftrightarrow> z) \<bullet> cs) = subst_branchb cs z (B_var bv) " and
+ "atom bv \<sharp> css \<Longrightarrow> ((bv \<leftrightarrow> z) \<bullet> css) = subst_branchlb css z (B_var bv) "
+
+proof(nominal_induct s and cs and css rule: s_branch_s_branch_list.strong_induct)
+ case (AS_branch x1 x2 x3)
+ hence "((bv \<leftrightarrow> z) \<bullet> x1) = x1" using pure_fresh fresh_at_base flip_fresh_fresh by metis
+ moreover have "((bv \<leftrightarrow> z) \<bullet> x2) = x2" using fresh_at_base flip_fresh_fresh[of bv x2 z] AS_branch by auto
+ ultimately show ?case unfolding s_branch_s_branch_list.perm_simps subst_branchb.simps using s_branch_s_branch_list.fresh(1) AS_branch by auto
+next
+ case (AS_cons x1 x2 )
+ hence "((bv \<leftrightarrow> z) \<bullet> x1) = subst_branchb x1 z (B_var bv)" using pure_fresh fresh_at_base flip_fresh_fresh s_branch_s_branch_list.fresh(13) by metis
+ moreover have "((bv \<leftrightarrow> z) \<bullet> x2) = subst_branchlb x2 z (B_var bv)" using fresh_at_base flip_fresh_fresh[of bv x2 z] AS_cons s_branch_s_branch_list.fresh by metis
+ ultimately show ?case unfolding s_branch_s_branch_list.perm_simps subst_branchb.simps using s_branch_s_branch_list.fresh(1) AS_cons by auto
+next
+ case (AS_val x)
+ then show ?case unfolding s_branch_s_branch_list.perm_simps subst_branchb.simps using flip_subst subst_b_v_def by simp
+next
+ case (AS_let x1 x2 x3)
+ moreover hence "((bv \<leftrightarrow> z) \<bullet> x1) = x1" using fresh_at_base flip_fresh_fresh[of bv x1 z] by auto
+ ultimately show ?case
+ unfolding s_branch_s_branch_list.perm_simps subst_sb.simps
+ using flip_subst subst_b_e_def s_branch_s_branch_list.fresh by auto
+next
+ case (AS_let2 x1 x2 x3 x4)
+ moreover hence "((bv \<leftrightarrow> z) \<bullet> x1) = x1" using fresh_at_base flip_fresh_fresh[of bv x1 z] by auto
+ ultimately show ?case
+ unfolding s_branch_s_branch_list.perm_simps subst_sb.simps
+ using flip_subst s_branch_s_branch_list.fresh(5) subst_b_\<tau>_def by auto
+next
+ case (AS_if x1 x2 x3)
+ thus ?case
+ unfolding s_branch_s_branch_list.perm_simps subst_sb.simps
+ using flip_subst subst_b_e_def subst_b_v_def s_branch_s_branch_list.fresh by auto
+next
+ case (AS_var x1 x2 x3 x4)
+ thus ?case
+ unfolding s_branch_s_branch_list.perm_simps subst_sb.simps
+ using flip_subst subst_b_e_def subst_b_v_def subst_b_\<tau>_def s_branch_s_branch_list.fresh fresh_at_base flip_fresh_fresh[of bv x1 z] by auto
+next
+ case (AS_assign x1 x2)
+ thus ?case
+ unfolding s_branch_s_branch_list.perm_simps subst_sb.simps
+ using flip_subst subst_b_e_def subst_b_v_def s_branch_s_branch_list.fresh fresh_at_base flip_fresh_fresh[of bv x1 z] by auto
+next
+ case (AS_match x1 x2)
+ thus ?case
+ unfolding s_branch_s_branch_list.perm_simps subst_sb.simps
+ using flip_subst subst_b_e_def subst_b_v_def s_branch_s_branch_list.fresh by auto
+next
+ case (AS_while x1 x2)
+ thus ?case
+ unfolding s_branch_s_branch_list.perm_simps subst_sb.simps
+ using flip_subst subst_b_e_def subst_b_v_def s_branch_s_branch_list.fresh by auto
+next
+ case (AS_seq x1 x2)
+ thus ?case
+ unfolding s_branch_s_branch_list.perm_simps subst_sb.simps
+ using flip_subst subst_b_e_def subst_b_v_def s_branch_s_branch_list.fresh by auto
+next
+ case (AS_assert x1 x2)
+ thus ?case
+ unfolding s_branch_s_branch_list.perm_simps subst_sb.simps
+ using flip_subst subst_b_c_def subst_b_v_def s_branch_s_branch_list.fresh by simp
+qed(auto)
+
+lemma flip_subst_subst_s:
+ fixes bv::bv and s::s and cs::branch_s and z::bv
+ shows "atom bv \<sharp> s \<Longrightarrow> ((bv \<leftrightarrow> z) \<bullet> s)[bv::=v]\<^sub>s\<^sub>b = s[z::=v]\<^sub>s\<^sub>b" and
+ "atom bv \<sharp> cs \<Longrightarrow> subst_branchb ((bv \<leftrightarrow> z) \<bullet> cs) bv v = subst_branchb cs z v" and
+ "atom bv \<sharp> css \<Longrightarrow> subst_branchlb ((bv \<leftrightarrow> z) \<bullet> css) bv v = subst_branchlb css z v "
+proof(nominal_induct s and cs and css rule: s_branch_s_branch_list.strong_induct)
+ case (AS_branch x1 x2 x3)
+ hence "((bv \<leftrightarrow> z) \<bullet> x1) = x1" using pure_fresh fresh_at_base flip_fresh_fresh by metis
+ moreover have "((bv \<leftrightarrow> z) \<bullet> x2) = x2" using fresh_at_base flip_fresh_fresh[of bv x2 z] AS_branch by auto
+ ultimately show ?case unfolding s_branch_s_branch_list.perm_simps subst_branchb.simps using s_branch_s_branch_list.fresh(1) AS_branch by auto
+next
+ case (AS_cons x1 x2 )
+ thus ?case
+ unfolding s_branch_s_branch_list.perm_simps subst_branchlb.simps
+ using s_branch_s_branch_list.fresh(1) AS_cons by auto
+
+next
+ case (AS_val x)
+ then show ?case unfolding s_branch_s_branch_list.perm_simps subst_branchb.simps using flip_subst subst_b_v_def by simp
+next
+ case (AS_let x1 x2 x3)
+ moreover hence "((bv \<leftrightarrow> z) \<bullet> x1) = x1" using fresh_at_base flip_fresh_fresh[of bv x1 z] by auto
+ ultimately show ?case
+ unfolding s_branch_s_branch_list.perm_simps subst_sb.simps
+ using flip_subst_subst subst_b_e_def s_branch_s_branch_list.fresh by force
+next
+ case (AS_let2 x1 x2 x3 x4)
+ moreover hence "((bv \<leftrightarrow> z) \<bullet> x1) = x1" using fresh_at_base flip_fresh_fresh[of bv x1 z] by auto
+ ultimately show ?case
+ unfolding s_branch_s_branch_list.perm_simps subst_sb.simps
+ using flip_subst s_branch_s_branch_list.fresh(5) subst_b_\<tau>_def by auto
+next
+ case (AS_if x1 x2 x3)
+ thus ?case
+ unfolding s_branch_s_branch_list.perm_simps subst_sb.simps
+ using flip_subst subst_b_e_def subst_b_v_def s_branch_s_branch_list.fresh by auto
+next
+ case (AS_var x1 x2 x3 x4)
+ thus ?case
+ unfolding s_branch_s_branch_list.perm_simps subst_sb.simps
+ using flip_subst subst_b_e_def subst_b_v_def subst_b_\<tau>_def s_branch_s_branch_list.fresh fresh_at_base flip_fresh_fresh[of bv x1 z] by auto
+next
+ case (AS_assign x1 x2)
+ thus ?case
+ unfolding s_branch_s_branch_list.perm_simps subst_sb.simps
+ using flip_subst subst_b_e_def subst_b_v_def s_branch_s_branch_list.fresh fresh_at_base flip_fresh_fresh[of bv x1 z] by auto
+next
+ case (AS_match x1 x2)
+ thus ?case
+ unfolding s_branch_s_branch_list.perm_simps subst_sb.simps
+ using flip_subst subst_b_e_def subst_b_v_def s_branch_s_branch_list.fresh by auto
+next
+ case (AS_while x1 x2)
+ thus ?case
+ unfolding s_branch_s_branch_list.perm_simps subst_sb.simps
+ using flip_subst subst_b_e_def subst_b_v_def s_branch_s_branch_list.fresh by auto
+next
+ case (AS_seq x1 x2)
+ thus ?case
+ unfolding s_branch_s_branch_list.perm_simps subst_sb.simps
+ using flip_subst subst_b_e_def subst_b_v_def s_branch_s_branch_list.fresh by auto
+next
+ case (AS_assert x1 x2)
+ thus ?case
+ unfolding s_branch_s_branch_list.perm_simps subst_sb.simps
+ using flip_subst subst_b_e_def subst_b_c_def s_branch_s_branch_list.fresh by auto
+qed(auto)
+
+instantiation s :: has_subst_b
+begin
+definition "subst_b = (\<lambda>s bv b. subst_sb s bv b)"
+
+instance proof
+ fix j::atom and i::bv and x::b and t::s
+ show "j \<sharp> subst_b t i x = ((atom i \<sharp> t \<and> j \<sharp> t) \<or> (j \<sharp> x \<and> (j \<sharp> t \<or> j = atom i)))"
+ using fresh_subst_sb_if subst_b_s_def by metis
+
+ fix a::bv and tm::s and x::b
+ show "atom a \<sharp> tm \<Longrightarrow> subst_b tm a x = tm" using subst_b_s_def forget_subst_sb by metis
+
+ fix a::bv and tm::s
+ show "subst_b tm a (B_var a) = tm" using subst_b_s_def subst_sb_id by metis
+
+ fix p::perm and x1::bv and v::b and t1::s
+ show "p \<bullet> subst_b t1 x1 v = subst_b (p \<bullet> t1) (p \<bullet> x1) (p \<bullet> v)" using subst_b_s_def subst_sb_subst_branchb_subst_branchlb.eqvt by metis
+
+ fix bv::bv and c::s and z::bv
+ show "atom bv \<sharp> c \<Longrightarrow> ((bv \<leftrightarrow> z) \<bullet> c) = c[z::=B_var bv]\<^sub>b"
+ using subst_b_s_def flip_subst_s by metis
+
+ fix bv::bv and c::s and z::bv and v::b
+ show "atom bv \<sharp> c \<Longrightarrow> ((bv \<leftrightarrow> z) \<bullet> c)[bv::=v]\<^sub>b = c[z::=v]\<^sub>b"
+ using flip_subst_subst_s subst_b_s_def by metis
+qed
+end
+
+section \<open>Function Type\<close>
+
+nominal_function subst_ft_b :: "fun_typ \<Rightarrow> bv \<Rightarrow> b \<Rightarrow> fun_typ" where
+ "subst_ft_b ( AF_fun_typ z b c t (s::s)) x v = AF_fun_typ z (subst_bb b x v) (subst_cb c x v) t[x::=v]\<^sub>\<tau>\<^sub>b s[x::=v]\<^sub>s\<^sub>b"
+ apply(simp add: eqvt_def subst_ft_b_graph_aux_def )
+ apply(simp add:fun_typ.strong_exhaust,auto )
+ apply(rule_tac y=a and c="(a,b)" in fun_typ.strong_exhaust)
+ apply (auto simp: eqvt_at_def fresh_star_def fresh_Pair fresh_at_base)
+ done
+
+nominal_termination (eqvt) by lexicographic_order
+
+nominal_function subst_ftq_b :: "fun_typ_q \<Rightarrow> bv \<Rightarrow> b \<Rightarrow> fun_typ_q" where
+ "atom bv \<sharp> (x,v) \<Longrightarrow> subst_ftq_b (AF_fun_typ_some bv ft) x v = (AF_fun_typ_some bv (subst_ft_b ft x v))"
+| "subst_ftq_b (AF_fun_typ_none ft) x v = (AF_fun_typ_none (subst_ft_b ft x v))"
+ apply(simp add: eqvt_def subst_ftq_b_graph_aux_def )
+ apply(simp add:fun_typ_q.strong_exhaust,auto )
+ apply(rule_tac y=a and c="(aa,b)" in fun_typ_q.strong_exhaust)
+ by (auto simp: eqvt_at_def fresh_star_def fresh_Pair fresh_at_base)
+nominal_termination (eqvt) by lexicographic_order
+
+instantiation fun_typ :: has_subst_b
+begin
+definition "subst_b = subst_ft_b"
+
+text \<open>Note: Using just simp in the second apply unpacks and gives a single goal
+whereas auto gives 43 non-intuitive goals. These goals are easier to solve
+and tedious, however they that it clear if a mistake is made in the definition of the function.
+For example, I saw that one of the goals was
+going through with metis and the next wasn't.
+It turned out the definition of the function itself was wrong\<close>
+
+instance proof
+ fix j::atom and i::bv and x::b and t::fun_typ
+ show "j \<sharp> subst_b t i x = (atom i \<sharp> t \<and> j \<sharp> t \<or> j \<sharp> x \<and> (j \<sharp> t \<or> j = atom i))"
+ apply(nominal_induct t avoiding: i x rule:fun_typ.strong_induct)
+ apply(auto simp add: subst_b_fun_typ_def )
+ by(metis fresh_subst_if subst_b_s_def subst_b_\<tau>_def subst_b_b_def subst_b_c_def)+
+
+ fix a::bv and tm::fun_typ and x::b
+ show "atom a \<sharp> tm \<Longrightarrow> subst_b tm a x = tm"
+ apply (nominal_induct tm avoiding: a x rule: fun_typ.strong_induct)
+ apply(simp add: subst_b_fun_typ_def Abs1_eq_iff')
+ using subst_b_b_def subst_b_fun_typ_def subst_b_\<tau>_def subst_b_c_def subst_b_s_def
+ forget_subst fresh_at_base list.set_cases neq_Nil_conv set_ConsD
+ subst_ft_b.simps by metis
+
+ fix a::bv and tm::fun_typ
+ show "subst_b tm a (B_var a) = tm"
+ apply (nominal_induct tm rule: fun_typ.strong_induct)
+ apply(simp add: subst_b_fun_typ_def Abs1_eq_iff',auto)
+ using subst_b_b_def subst_b_fun_typ_def subst_b_\<tau>_def subst_b_c_def subst_b_s_def
+ forget_subst fresh_at_base list.set_cases neq_Nil_conv set_ConsD
+ subst_ft_b.simps
+ by (metis has_subst_b_class.subst_id)+
+
+ fix p::perm and x1::bv and v::b and t1::fun_typ
+ show "p \<bullet> subst_b t1 x1 v = subst_b (p \<bullet> t1) (p \<bullet> x1) (p \<bullet> v) "
+ apply (nominal_induct t1 avoiding: x1 v rule: fun_typ.strong_induct)
+ by(auto simp add: subst_b_fun_typ_def Abs1_eq_iff' fun_typ.perm_simps)
+
+ fix bv::bv and c::fun_typ and z::bv
+ show "atom bv \<sharp> c \<Longrightarrow> ((bv \<leftrightarrow> z) \<bullet> c) = c[z::=B_var bv]\<^sub>b"
+ apply (nominal_induct c avoiding: z bv rule: fun_typ.strong_induct)
+ by(auto simp add: subst_b_fun_typ_def Abs1_eq_iff' fun_typ.perm_simps subst_b_b_def subst_b_c_def subst_b_\<tau>_def subst_b_s_def)
+
+ fix bv::bv and c::fun_typ and z::bv and v::b
+ show "atom bv \<sharp> c \<Longrightarrow> ((bv \<leftrightarrow> z) \<bullet> c)[bv::=v]\<^sub>b = c[z::=v]\<^sub>b"
+ apply (nominal_induct c avoiding: bv v z rule: fun_typ.strong_induct)
+ apply(auto simp add: subst_b_fun_typ_def Abs1_eq_iff' fun_typ.perm_simps subst_b_b_def subst_b_c_def subst_b_\<tau>_def subst_b_s_def flip_subst_subst flip_subst)
+ using subst_b_fun_typ_def Abs1_eq_iff' fun_typ.perm_simps subst_b_b_def subst_b_c_def subst_b_\<tau>_def subst_b_s_def flip_subst_subst flip_subst
+ using flip_subst_s(1) flip_subst_subst_s(1) by auto
+qed
+end
+
+instantiation fun_typ_q :: has_subst_b
+begin
+definition "subst_b = subst_ftq_b"
+
+instance proof
+ fix j::atom and i::bv and x::b and t::fun_typ_q
+ show "j \<sharp> subst_b t i x = (atom i \<sharp> t \<and> j \<sharp> t \<or> j \<sharp> x \<and> (j \<sharp> t \<or> j = atom i))"
+ apply (nominal_induct t avoiding: i x j rule: fun_typ_q.strong_induct,auto simp add: subst_b_fun_typ_q_def subst_ftq_b.simps)
+ using fresh_subst_if subst_b_fun_typ_q_def subst_b_s_def subst_b_\<tau>_def subst_b_b_def subst_b_c_def subst_b_fun_typ_def apply metis+
+ done
+
+ fix a::bv and t::fun_typ_q and x::b
+ show "atom a \<sharp> t \<Longrightarrow> subst_b t a x = t"
+ apply (nominal_induct t avoiding: a x rule: fun_typ_q.strong_induct)
+ apply(auto simp add: subst_b_fun_typ_q_def subst_ftq_b.simps Abs1_eq_iff')
+ using forget_subst subst_b_fun_typ_q_def subst_b_s_def subst_b_\<tau>_def subst_b_b_def subst_b_c_def subst_b_fun_typ_def eqvt by metis+
+
+ fix p::perm and x1::bv and v::b and t::fun_typ_q
+ show "p \<bullet> subst_b t x1 v = subst_b (p \<bullet> t) (p \<bullet> x1) (p \<bullet> v) "
+ apply (nominal_induct t avoiding: x1 v rule: fun_typ_q.strong_induct)
+ by(auto simp add: subst_b_fun_typ_q_def subst_ftq_b.simps Abs1_eq_iff')
+
+ fix a::bv and tm::fun_typ_q
+ show "subst_b tm a (B_var a) = tm"
+ apply (nominal_induct tm avoiding: a rule: fun_typ_q.strong_induct)
+ apply(auto simp add: subst_b_fun_typ_q_def subst_ftq_b.simps Abs1_eq_iff')
+ using subst_id subst_b_b_def subst_b_fun_typ_def subst_b_\<tau>_def subst_b_c_def subst_b_s_def
+ forget_subst fresh_at_base list.set_cases neq_Nil_conv set_ConsD
+ subst_ft_b.simps by metis+
+
+ fix bv::bv and c::fun_typ_q and z::bv
+ show "atom bv \<sharp> c \<Longrightarrow> ((bv \<leftrightarrow> z) \<bullet> c) = c[z::=B_var bv]\<^sub>b"
+ apply (nominal_induct c avoiding: z bv rule: fun_typ_q.strong_induct)
+ apply(auto simp add: subst_b_fun_typ_q_def subst_ftq_b.simps Abs1_eq_iff')
+ using forget_subst subst_b_fun_typ_q_def subst_b_s_def subst_b_\<tau>_def subst_b_b_def subst_b_c_def subst_b_fun_typ_def eqvt by metis+
+
+ fix bv::bv and c::fun_typ_q and z::bv and v::b
+ show "atom bv \<sharp> c \<Longrightarrow> ((bv \<leftrightarrow> z) \<bullet> c)[bv::=v]\<^sub>b = c[z::=v]\<^sub>b"
+ apply (nominal_induct c avoiding: z v bv rule: fun_typ_q.strong_induct)
+ apply(auto simp add: subst_b_fun_typ_q_def subst_ftq_b.simps Abs1_eq_iff')
+ using flip_subst flip_subst_subst forget_subst subst_b_fun_typ_q_def subst_b_s_def subst_b_\<tau>_def subst_b_b_def subst_b_c_def subst_b_fun_typ_def eqvt by metis+
+
+qed
+end
+
+section \<open>Contexts\<close>
+
+subsection \<open>Immutable Variables\<close>
+nominal_function subst_gb :: "\<Gamma> \<Rightarrow> bv \<Rightarrow> b \<Rightarrow> \<Gamma>" where
+ "subst_gb GNil _ _ = GNil"
+| "subst_gb ((y,b',c)#\<^sub>\<Gamma>\<Gamma>) bv b = ((y,b'[bv::=b]\<^sub>b\<^sub>b,c[bv::=b]\<^sub>c\<^sub>b)#\<^sub>\<Gamma> (subst_gb \<Gamma> bv b))"
+ apply (simp add: eqvt_def subst_gb_graph_aux_def )+
+ apply auto
+ apply (insert \<Gamma>.exhaust neq_GNil_conv, force)
+ done
+nominal_termination (eqvt) by lexicographic_order
+
+abbreviation
+ subst_gb_abbrev :: "\<Gamma> \<Rightarrow> bv \<Rightarrow> b \<Rightarrow> \<Gamma>" ("_[_::=_]\<^sub>\<Gamma>\<^sub>b" [1000,50,50] 1000)
+ where
+ "g[bv::=b']\<^sub>\<Gamma>\<^sub>b \<equiv> subst_gb g bv b'"
+
+instantiation \<Gamma> :: has_subst_b
+begin
+definition "subst_b = subst_gb"
+
+instance proof
+ fix j::atom and i::bv and x::b and t::\<Gamma>
+ show "j \<sharp> subst_b t i x = (atom i \<sharp> t \<and> j \<sharp> t \<or> j \<sharp> x \<and> (j \<sharp> t \<or> j = atom i))"
+ proof(induct t rule: \<Gamma>_induct)
+ case GNil
+ then show ?case using fresh_GNil subst_gb.simps fresh_def pure_fresh subst_b_\<Gamma>_def has_subst_b_class.fresh_subst_if fresh_GNil fresh_GCons by metis
+ next
+ case (GCons x' b' c' \<Gamma>')
+ have *: "atom i \<sharp> x' " using fresh_at_base by simp
+
+ have "j \<sharp> subst_b ((x', b', c') #\<^sub>\<Gamma> \<Gamma>') i x = j \<sharp> ((x', b'[i::=x]\<^sub>b\<^sub>b, c'[i::=x]\<^sub>c\<^sub>b) #\<^sub>\<Gamma> (subst_b \<Gamma>' i x))" using subst_gb.simps subst_b_\<Gamma>_def by auto
+ also have "... = (j \<sharp> ((x', b'[i::=x]\<^sub>b\<^sub>b, c'[i::=x]\<^sub>c\<^sub>b)) \<and> (j \<sharp> (subst_b \<Gamma>' i x)))" using fresh_GCons by auto
+ also have "... = (((j \<sharp> x') \<and> (j \<sharp> b'[i::=x]\<^sub>b\<^sub>b) \<and> (j \<sharp> c'[i::=x]\<^sub>c\<^sub>b)) \<and> (j \<sharp> (subst_b \<Gamma>' i x)))" by auto
+ also have "... = (((j \<sharp> x') \<and> ((atom i \<sharp> b' \<and> j \<sharp> b' \<or> j \<sharp> x \<and> (j \<sharp> b' \<or> j = atom i))) \<and>
+ ((atom i \<sharp> c' \<and> j \<sharp> c' \<or> j \<sharp> x \<and> (j \<sharp> c' \<or> j = atom i))) \<and>
+ ((atom i \<sharp> \<Gamma>' \<and> j \<sharp> \<Gamma>' \<or> j \<sharp> x \<and> (j \<sharp> \<Gamma>' \<or> j = atom i)))))"
+ using fresh_subst_if[of j b' i x] fresh_subst_if[of j c' i x] GCons subst_b_b_def subst_b_c_def by simp
+ also have "... = ((atom i \<sharp> (x', b', c') #\<^sub>\<Gamma> \<Gamma>' \<and> j \<sharp> (x', b', c') #\<^sub>\<Gamma> \<Gamma>') \<or> (j \<sharp> x \<and> (j \<sharp> (x', b', c') #\<^sub>\<Gamma> \<Gamma>' \<or> j = atom i)))" using * fresh_GCons fresh_prod3 by metis
+
+ finally show ?case by auto
+ qed
+
+ fix a::bv and tm::\<Gamma> and x::b
+ show "atom a \<sharp> tm \<Longrightarrow> subst_b tm a x = tm"
+ proof (induct tm rule: \<Gamma>_induct)
+ case GNil
+ then show ?case using subst_gb.simps subst_b_\<Gamma>_def by auto
+ next
+ case (GCons x' b' c' \<Gamma>')
+ have *:"b'[a::=x]\<^sub>b\<^sub>b = b' \<and> c'[a::=x]\<^sub>c\<^sub>b = c'" using GCons fresh_GCons[of "atom a"] fresh_prod3[of "atom a"] has_subst_b_class.forget_subst subst_b_b_def subst_b_c_def by metis
+ have "subst_b ((x', b', c') #\<^sub>\<Gamma> \<Gamma>') a x = ((x', b'[a::=x]\<^sub>b\<^sub>b, c'[a::=x]\<^sub>c\<^sub>b) #\<^sub>\<Gamma> (subst_b \<Gamma>' a x))" using subst_b_\<Gamma>_def subst_gb.simps by auto
+ also have "... = ((x', b', c') #\<^sub>\<Gamma> \<Gamma>')" using * GCons fresh_GCons[of "atom a"] by auto
+ finally show ?case using has_subst_b_class.forget_subst fresh_GCons fresh_prod3 GCons subst_b_\<Gamma>_def has_subst_b_class.forget_subst[of a b' x] fresh_prod3[of "atom a"] by argo
+ qed
+
+ fix a::bv and tm::\<Gamma>
+ show "subst_b tm a (B_var a) = tm"
+ proof(induct tm rule: \<Gamma>_induct)
+ case GNil
+ then show ?case using subst_gb.simps subst_b_\<Gamma>_def by auto
+ next
+ case (GCons x' b' c' \<Gamma>')
+ then show ?case using has_subst_b_class.subst_id subst_b_\<Gamma>_def subst_b_b_def subst_b_c_def subst_gb.simps by metis
+ qed
+
+ fix p::perm and x1::bv and v::b and t1::\<Gamma>
+ show "p \<bullet> subst_b t1 x1 v = subst_b (p \<bullet> t1) (p \<bullet> x1) (p \<bullet> v)"
+ proof (induct tm rule: \<Gamma>_induct)
+ case GNil
+ then show ?case using subst_b_\<Gamma>_def subst_gb.simps by simp
+ next
+ case (GCons x' b' c' \<Gamma>')
+ then show ?case using subst_b_\<Gamma>_def subst_gb.simps has_subst_b_class.eqvt by argo
+ qed
+
+ fix bv::bv and c::\<Gamma> and z::bv
+ show "atom bv \<sharp> c \<Longrightarrow> ((bv \<leftrightarrow> z) \<bullet> c) = c[z::=B_var bv]\<^sub>b"
+ proof (induct c rule: \<Gamma>_induct)
+ case GNil
+ then show ?case using subst_b_\<Gamma>_def subst_gb.simps by auto
+ next
+ case (GCons x b c \<Gamma>')
+ have *:"(bv \<leftrightarrow> z) \<bullet> (x, b, c) = (x, (bv \<leftrightarrow> z) \<bullet> b, (bv \<leftrightarrow> z) \<bullet> c)" using flip_bv_x_cancel by auto
+ then show ?case
+ unfolding subst_gb.simps subst_b_\<Gamma>_def permute_\<Gamma>.simps *
+ using GCons subst_b_\<Gamma>_def subst_gb.simps flip_subst subst_b_b_def subst_b_c_def fresh_GCons by auto
+ qed
+
+ fix bv::bv and c::\<Gamma> and z::bv and v::b
+ show "atom bv \<sharp> c \<Longrightarrow> ((bv \<leftrightarrow> z) \<bullet> c)[bv::=v]\<^sub>b = c[z::=v]\<^sub>b"
+ proof (induct c rule: \<Gamma>_induct)
+ case GNil
+ then show ?case using subst_b_\<Gamma>_def subst_gb.simps by auto
+ next
+ case (GCons x b c \<Gamma>')
+ have *:"(bv \<leftrightarrow> z) \<bullet> (x, b, c) = (x, (bv \<leftrightarrow> z) \<bullet> b, (bv \<leftrightarrow> z) \<bullet> c)" using flip_bv_x_cancel by auto
+ then show ?case
+ unfolding subst_gb.simps subst_b_\<Gamma>_def permute_\<Gamma>.simps *
+ using GCons subst_b_\<Gamma>_def subst_gb.simps flip_subst subst_b_b_def subst_b_c_def fresh_GCons by auto
+ qed
+qed
+end
+
+lemma subst_b_base_for_lit:
+ "(base_for_lit l)[bv::=b]\<^sub>b\<^sub>b = base_for_lit l"
+ using base_for_lit.simps l.strong_exhaust
+ by (metis subst_bb.simps(2) subst_bb.simps(3) subst_bb.simps(6) subst_bb.simps(7))
+
+lemma subst_b_lookup:
+ assumes "Some (b, c) = lookup \<Gamma> x"
+ shows " Some (b[bv::=b']\<^sub>b\<^sub>b, c[bv::=b']\<^sub>c\<^sub>b) = lookup \<Gamma>[bv::=b']\<^sub>\<Gamma>\<^sub>b x"
+ using assms by(induct \<Gamma> rule: \<Gamma>_induct, auto)
+
+lemma subst_g_b_x_fresh:
+ fixes x::x and b::b and \<Gamma>::\<Gamma> and bv::bv
+ assumes "atom x \<sharp> \<Gamma>"
+ shows "atom x \<sharp> \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b"
+ using subst_b_fresh_x subst_b_\<Gamma>_def assms by metis
+
+subsection \<open>Mutable Variables\<close>
+
+nominal_function subst_db :: "\<Delta> \<Rightarrow> bv \<Rightarrow> b \<Rightarrow> \<Delta>" where
+ "subst_db []\<^sub>\<Delta> _ _ = []\<^sub>\<Delta>"
+| "subst_db ((u,t) #\<^sub>\<Delta> \<Delta>) bv b = ((u,t[bv::=b]\<^sub>\<tau>\<^sub>b) #\<^sub>\<Delta> (subst_db \<Delta> bv b))"
+ apply (simp add: eqvt_def subst_db_graph_aux_def,auto )
+ using list.exhaust delete_aux.elims
+ using neq_DNil_conv by fastforce
+nominal_termination (eqvt) by lexicographic_order
+
+abbreviation
+ subst_db_abbrev :: "\<Delta> \<Rightarrow> bv \<Rightarrow> b \<Rightarrow> \<Delta>" ("_[_::=_]\<^sub>\<Delta>\<^sub>b" [1000,50,50] 1000)
+ where
+ "\<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b \<equiv> subst_db \<Delta> bv b"
+
+instantiation \<Delta> :: has_subst_b
+begin
+definition "subst_b = subst_db"
+
+instance proof
+ fix j::atom and i::bv and x::b and t::\<Delta>
+ show "j \<sharp> subst_b t i x = (atom i \<sharp> t \<and> j \<sharp> t \<or> j \<sharp> x \<and> (j \<sharp> t \<or> j = atom i))"
+ proof(induct t rule: \<Delta>_induct)
+ case DNil
+ then show ?case using fresh_DNil subst_db.simps fresh_def pure_fresh subst_b_\<Delta>_def has_subst_b_class.fresh_subst_if fresh_DNil fresh_DCons by metis
+ next
+ case (DCons u t \<Gamma>')
+ have "j \<sharp> subst_b ((u, t) #\<^sub>\<Delta> \<Gamma>') i x = j \<sharp> ((u, t[i::=x]\<^sub>\<tau>\<^sub>b) #\<^sub>\<Delta> (subst_b \<Gamma>' i x))" using subst_db.simps subst_b_\<Delta>_def by auto
+ also have "... = (j \<sharp> ((u, t[i::=x]\<^sub>\<tau>\<^sub>b)) \<and> (j \<sharp> (subst_b \<Gamma>' i x)))" using fresh_DCons by auto
+ also have "... = (((j \<sharp> u) \<and> (j \<sharp> t[i::=x]\<^sub>\<tau>\<^sub>b)) \<and> (j \<sharp> (subst_b \<Gamma>' i x)))" by auto
+ also have "... = ((j \<sharp> u) \<and> ((atom i \<sharp> t \<and> j \<sharp> t) \<or> (j \<sharp> x \<and> (j \<sharp> t \<or> j = atom i))) \<and> (atom i \<sharp> \<Gamma>' \<and> j \<sharp> \<Gamma>' \<or> j \<sharp> x \<and> (j \<sharp> \<Gamma>' \<or> j = atom i)))"
+ using has_subst_b_class.fresh_subst_if[of j t i x] subst_b_\<tau>_def DCons subst_b_\<Delta>_def by auto
+ also have "... = (atom i \<sharp> (u, t) #\<^sub>\<Delta> \<Gamma>' \<and> j \<sharp> (u, t) #\<^sub>\<Delta> \<Gamma>' \<or> j \<sharp> x \<and> (j \<sharp> (u, t) #\<^sub>\<Delta> \<Gamma>' \<or> j = atom i))"
+ using DCons subst_db.simps(2) has_subst_b_class.fresh_subst_if fresh_DCons subst_b_\<Delta>_def pure_fresh fresh_at_base by auto
+ finally show ?case by auto
+ qed
+
+ fix a::bv and tm::\<Delta> and x::b
+ show "atom a \<sharp> tm \<Longrightarrow> subst_b tm a x = tm"
+ proof (induct tm rule: \<Delta>_induct)
+ case DNil
+ then show ?case using subst_db.simps subst_b_\<Delta>_def by auto
+ next
+ case (DCons u t \<Gamma>')
+ have *:"t[a::=x]\<^sub>\<tau>\<^sub>b = t" using DCons fresh_DCons[of "atom a"] fresh_prod2[of "atom a"] has_subst_b_class.forget_subst subst_b_\<tau>_def by metis
+ have "subst_b ((u,t) #\<^sub>\<Delta> \<Gamma>') a x = ((u,t[a::=x]\<^sub>\<tau>\<^sub>b) #\<^sub>\<Delta> (subst_b \<Gamma>' a x))" using subst_b_\<Delta>_def subst_db.simps by auto
+ also have "... = ((u, t) #\<^sub>\<Delta> \<Gamma>')" using * DCons fresh_DCons[of "atom a"] by auto
+ finally show ?case using
+ has_subst_b_class.forget_subst fresh_DCons fresh_prod3
+ DCons subst_b_\<Delta>_def has_subst_b_class.forget_subst[of a t x] fresh_prod3[of "atom a"] by argo
+ qed
+
+ fix a::bv and tm::\<Delta>
+ show "subst_b tm a (B_var a) = tm"
+ proof(induct tm rule: \<Delta>_induct)
+ case DNil
+ then show ?case using subst_db.simps subst_b_\<Delta>_def by auto
+ next
+ case (DCons u t \<Gamma>')
+ then show ?case using has_subst_b_class.subst_id subst_b_\<Delta>_def subst_b_\<tau>_def subst_db.simps by metis
+ qed
+
+ fix p::perm and x1::bv and v::b and t1::\<Delta>
+ show "p \<bullet> subst_b t1 x1 v = subst_b (p \<bullet> t1) (p \<bullet> x1) (p \<bullet> v)"
+ proof (induct tm rule: \<Delta>_induct)
+ case DNil
+ then show ?case using subst_b_\<Delta>_def subst_db.simps by simp
+ next
+ case (DCons x' b' \<Gamma>')
+ then show ?case by argo
+ qed
+
+ fix bv::bv and c::\<Delta> and z::bv
+ show "atom bv \<sharp> c \<Longrightarrow> ((bv \<leftrightarrow> z) \<bullet> c) = c[z::=B_var bv]\<^sub>b"
+ proof (induct c rule: \<Delta>_induct)
+ case DNil
+ then show ?case using subst_b_\<Delta>_def subst_db.simps by auto
+ next
+ case (DCons u t')
+ then show ?case
+ unfolding subst_db.simps subst_b_\<Delta>_def permute_\<Delta>.simps
+ using DCons subst_b_\<Delta>_def subst_db.simps flip_subst subst_b_\<tau>_def flip_fresh_fresh fresh_at_base fresh_DCons flip_bv_u_cancel by simp
+ qed
+
+ fix bv::bv and c::\<Delta> and z::bv and v::b
+ show "atom bv \<sharp> c \<Longrightarrow> ((bv \<leftrightarrow> z) \<bullet> c)[bv::=v]\<^sub>b = c[z::=v]\<^sub>b"
+ proof (induct c rule: \<Delta>_induct)
+ case DNil
+ then show ?case using subst_b_\<Delta>_def subst_db.simps by auto
+ next
+ case (DCons u t')
+ then show ?case
+ unfolding subst_db.simps subst_b_\<Delta>_def permute_\<Delta>.simps
+ using DCons subst_b_\<Delta>_def subst_db.simps flip_subst subst_b_\<tau>_def flip_fresh_fresh fresh_at_base fresh_DCons flip_bv_u_cancel by simp
+ qed
+qed
+end
+
+lemma subst_d_b_member:
+ assumes "(u, \<tau>) \<in> setD \<Delta>"
+ shows "(u, \<tau>[bv::=b]\<^sub>\<tau>\<^sub>b) \<in> setD \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b"
+ using assms by (induct \<Delta>,auto)
+
+lemmas ms_fresh_all = e.fresh s_branch_s_branch_list.fresh \<tau>.fresh c.fresh ce.fresh v.fresh l.fresh fresh_at_base opp.fresh pure_fresh ms_fresh
+
+lemmas fresh_intros[intro] = fresh_GNil x_not_in_b_set x_not_in_u_atoms x_fresh_b u_not_in_x_atoms bv_not_in_x_atoms u_not_in_b_atoms
+
+lemmas subst_b_simps = subst_tb.simps subst_cb.simps subst_ceb.simps subst_vb.simps subst_bb.simps subst_eb.simps subst_branchb.simps subst_sb.simps
+
+lemma subst_d_b_x_fresh:
+ fixes x::x and b::b and \<Delta>::\<Delta> and bv::bv
+ assumes "atom x \<sharp> \<Delta>"
+ shows "atom x \<sharp> \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b"
+ using subst_b_fresh_x subst_b_\<Delta>_def assms by metis
+
+lemma subst_b_fresh_x:
+ fixes x::x
+ shows "atom x \<sharp> v \<Longrightarrow> atom x \<sharp> v[bv::=b']\<^sub>v\<^sub>b" and
+ "atom x \<sharp> ce \<Longrightarrow> atom x \<sharp> ce[bv::=b']\<^sub>c\<^sub>e\<^sub>b" and
+ "atom x \<sharp> e \<Longrightarrow> atom x \<sharp> e[bv::=b']\<^sub>e\<^sub>b" and
+ "atom x \<sharp> c \<Longrightarrow> atom x \<sharp> c[bv::=b']\<^sub>c\<^sub>b" and
+ "atom x \<sharp> t \<Longrightarrow> atom x \<sharp> t[bv::=b']\<^sub>\<tau>\<^sub>b" and
+ "atom x \<sharp> d \<Longrightarrow> atom x \<sharp> d[bv::=b']\<^sub>\<Delta>\<^sub>b" and
+ "atom x \<sharp> g \<Longrightarrow> atom x \<sharp> g[bv::=b']\<^sub>\<Gamma>\<^sub>b" and
+ "atom x \<sharp> s \<Longrightarrow> atom x \<sharp> s[bv::=b']\<^sub>s\<^sub>b"
+ using fresh_subst_if x_fresh_b subst_b_v_def subst_b_ce_def subst_b_e_def subst_b_c_def subst_b_\<tau>_def subst_b_s_def subst_g_b_x_fresh subst_d_b_x_fresh
+ by metis+
+
+lemma subst_b_fresh_u_cls:
+ fixes tm::"'a::has_subst_b" and x::u
+ shows "atom x \<sharp> tm = atom x \<sharp> tm[bv::=b']\<^sub>b"
+ using fresh_subst_if[of "atom x" tm bv b'] using u_fresh_b by auto
+
+lemma subst_g_b_u_fresh:
+ fixes x::u and b::b and \<Gamma>::\<Gamma> and bv::bv
+ assumes "atom x \<sharp> \<Gamma>"
+ shows "atom x \<sharp> \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b"
+ using subst_b_fresh_u_cls subst_b_\<Gamma>_def assms by metis
+
+lemma subst_d_b_u_fresh:
+ fixes x::u and b::b and \<Gamma>::\<Delta> and bv::bv
+ assumes "atom x \<sharp> \<Gamma>"
+ shows "atom x \<sharp> \<Gamma>[bv::=b]\<^sub>\<Delta>\<^sub>b"
+ using subst_b_fresh_u_cls subst_b_\<Delta>_def assms by metis
+
+lemma subst_b_fresh_u:
+ fixes x::u
+ shows "atom x \<sharp> v \<Longrightarrow> atom x \<sharp> v[bv::=b']\<^sub>v\<^sub>b" and
+ "atom x \<sharp> ce \<Longrightarrow> atom x \<sharp> ce[bv::=b']\<^sub>c\<^sub>e\<^sub>b" and
+ "atom x \<sharp> e \<Longrightarrow> atom x \<sharp> e[bv::=b']\<^sub>e\<^sub>b" and
+ "atom x \<sharp> c \<Longrightarrow> atom x \<sharp> c[bv::=b']\<^sub>c\<^sub>b" and
+ "atom x \<sharp> t \<Longrightarrow> atom x \<sharp> t[bv::=b']\<^sub>\<tau>\<^sub>b" and
+ "atom x \<sharp> d \<Longrightarrow> atom x \<sharp> d[bv::=b']\<^sub>\<Delta>\<^sub>b" and
+ "atom x \<sharp> g \<Longrightarrow> atom x \<sharp> g[bv::=b']\<^sub>\<Gamma>\<^sub>b" and
+ "atom x \<sharp> s \<Longrightarrow> atom x \<sharp> s[bv::=b']\<^sub>s\<^sub>b"
+ using fresh_subst_if u_fresh_b subst_b_v_def subst_b_ce_def subst_b_e_def subst_b_c_def subst_b_\<tau>_def subst_b_s_def subst_g_b_u_fresh subst_d_b_u_fresh
+ by metis+
+
+lemma subst_db_u_fresh:
+ fixes u::u and b::b and D::\<Delta>
+ assumes "atom u \<sharp> D"
+ shows "atom u \<sharp> D[bv::=b]\<^sub>\<Delta>\<^sub>b"
+ using assms proof(induct D rule: \<Delta>_induct)
+ case DNil
+ then show ?case by auto
+next
+ case (DCons u' t' D')
+ then show ?case using subst_db.simps fresh_def fresh_DCons fresh_subst_if subst_b_\<tau>_def
+ by (metis fresh_Pair u_not_in_b_atoms)
+qed
+
+lemma flip_bt_subst4:
+ fixes t::\<tau> and bv::bv
+ assumes "atom bv \<sharp> t"
+ shows "t[bv'::=b]\<^sub>\<tau>\<^sub>b = ((bv' \<leftrightarrow> bv) \<bullet> t)[bv::=b]\<^sub>\<tau>\<^sub>b"
+ using flip_subst_subst[OF assms,of bv' b]
+ by (simp add: flip_commute subst_b_\<tau>_def)
+
+lemma subst_bt_flip_sym:
+ fixes t1::\<tau> and t2::\<tau>
+ assumes "atom bv \<sharp> b" and "atom bv \<sharp> (bv1, bv2, t1, t2)" and "(bv1 \<leftrightarrow> bv) \<bullet> t1 = (bv2 \<leftrightarrow> bv) \<bullet> t2"
+ shows " t1[bv1::=b]\<^sub>\<tau>\<^sub>b = t2[bv2::=b]\<^sub>\<tau>\<^sub>b"
+ using assms flip_bt_subst4[of bv t1 bv1 b ] flip_bt_subst4 fresh_prod4 fresh_Pair by metis
+
end
\ No newline at end of file
diff --git a/thys/MiniSail/BTVSubstTypingL.thy b/thys/MiniSail/BTVSubstTypingL.thy
--- a/thys/MiniSail/BTVSubstTypingL.thy
+++ b/thys/MiniSail/BTVSubstTypingL.thy
@@ -1,635 +1,635 @@
-(*<*)
-theory BTVSubstTypingL
- imports "HOL-Eisbach.Eisbach_Tools" ContextSubtypingL SubstMethods
-begin
- (*>*)
-
-chapter \<open>Basic Type Variable Substitution Lemmas\<close>
-text \<open>Lemmas that show that types are preserved, in some way,
-under basic type variable substitution\<close>
-
-lemma subst_vv_subst_bb_commute:
- fixes bv::bv and b::b and x::x and v::v
- assumes "atom bv \<sharp> v"
- shows "(v'[x::=v]\<^sub>v\<^sub>v)[bv::=b]\<^sub>v\<^sub>b = (v'[bv::=b]\<^sub>v\<^sub>b)[x::=v]\<^sub>v\<^sub>v"
- using assms proof(nominal_induct v' rule: v.strong_induct)
- case (V_lit x)
- then show ?case using subst_vb.simps subst_vv.simps by simp
-next
- case (V_var y)
- hence "v[bv::=b]\<^sub>v\<^sub>b=v" using forget_subst subst_b_v_def by metis
- then show ?case unfolding subst_vb.simps(2) subst_vv.simps(2) using V_var by auto
-next
- case (V_pair x1a x2a)
- then show ?case using subst_vb.simps subst_vv.simps by simp
-next
- case (V_cons x1a x2a x3)
- then show ?case using subst_vb.simps subst_vv.simps by simp
-next
- case (V_consp x1a x2a x3 x4)
- then show ?case using subst_vb.simps subst_vv.simps by simp
-qed
-
-lemma subst_cev_subst_bb_commute:
- fixes bv::bv and b::b and x::x and v::v
- assumes "atom bv \<sharp> v"
- shows "(ce[x::=v]\<^sub>v)[bv::=b]\<^sub>c\<^sub>e\<^sub>b = (ce[bv::=b]\<^sub>c\<^sub>e\<^sub>b)[x::=v]\<^sub>v"
- using assms apply (nominal_induct ce rule: ce.strong_induct, (simp add: subst_vv_subst_bb_commute subst_ceb.simps subst_cv.simps))
- using assms subst_vv_subst_bb_commute subst_ceb.simps subst_cv.simps
- by (simp add: subst_v_ce_def)+
-
-lemma subst_cv_subst_bb_commute:
- fixes bv::bv and b::b and x::x and v::v
- assumes "atom bv \<sharp> v"
- shows "c[x::=v]\<^sub>c\<^sub>v[bv::=b]\<^sub>c\<^sub>b = (c[bv::=b]\<^sub>c\<^sub>b)[x::=v]\<^sub>c\<^sub>v"
- using assms apply (nominal_induct c rule: c.strong_induct )
- using assms subst_vv_subst_bb_commute subst_ceb.simps subst_cv.simps
- subst_v_c_def subst_b_c_def apply auto
- using subst_cev_subst_bb_commute subst_v_ce_def by auto+
-
-lemma subst_b_c_of:
- "(c_of \<tau> z)[bv::=b]\<^sub>c\<^sub>b = c_of (\<tau>[bv::=b]\<^sub>\<tau>\<^sub>b) z"
-proof(nominal_induct \<tau> avoiding: z rule:\<tau>.strong_induct)
- case (T_refined_type z' b' c')
- moreover have "atom bv \<sharp> [ z ]\<^sup>v " using fresh_at_base v.fresh by auto
- ultimately show ?case using subst_cv_subst_bb_commute[of bv "V_var z" c' z' b] c_of.simps subst_tb.simps by metis
-qed
-
-lemma subst_b_b_of:
- "(b_of \<tau>)[bv::=b]\<^sub>b\<^sub>b = b_of (\<tau>[bv::=b]\<^sub>\<tau>\<^sub>b)"
- by(nominal_induct \<tau> rule:\<tau>.strong_induct, simp add: b_of.simps subst_tb.simps )
-
-lemma subst_b_if:
- "\<lbrace> z : b_of \<tau>[bv::=b]\<^sub>\<tau>\<^sub>b | CE_val (v[bv::=b]\<^sub>v\<^sub>b) == CE_val (V_lit ll) IMP c_of \<tau>[bv::=b]\<^sub>\<tau>\<^sub>b z \<rbrace> = \<lbrace> z : b_of \<tau> | CE_val (v) == CE_val (V_lit ll) IMP c_of \<tau> z \<rbrace>[bv::=b]\<^sub>\<tau>\<^sub>b "
- unfolding subst_tb.simps subst_cb.simps subst_ceb.simps subst_vb.simps using subst_b_b_of subst_b_c_of by auto
-
-lemma subst_b_top_eq:
- "\<lbrace> z : B_unit | TRUE \<rbrace>[bv::=b]\<^sub>\<tau>\<^sub>b = \<lbrace> z : B_unit | TRUE \<rbrace>" and "\<lbrace> z : B_bool | TRUE \<rbrace>[bv::=b]\<^sub>\<tau>\<^sub>b = \<lbrace> z : B_bool | TRUE \<rbrace>" and "\<lbrace> z : B_id tid | TRUE \<rbrace> = \<lbrace> z : B_id tid | TRUE \<rbrace>[bv::=b]\<^sub>\<tau>\<^sub>b"
- unfolding subst_tb.simps subst_bb.simps subst_cb.simps by auto
-
-lemmas subst_b_eq = subst_b_c_of subst_b_b_of subst_b_if subst_b_top_eq
-
-lemma subst_cx_subst_bb_commute[simp]:
- fixes bv::bv and b::b and x::x and v'::c
- shows "(v'[x::=V_var y]\<^sub>c\<^sub>v)[bv::=b]\<^sub>c\<^sub>b = (v'[bv::=b]\<^sub>c\<^sub>b)[x::=V_var y]\<^sub>c\<^sub>v"
- using subst_cv_subst_bb_commute fresh_at_base v.fresh by auto
-
-lemma subst_b_infer_b:
- fixes l::l and b::b
- assumes " \<turnstile> l \<Rightarrow> \<tau>" and "\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f b" and "B = {|bv|}"
- shows "\<turnstile> l \<Rightarrow> (\<tau>[bv::=b]\<^sub>\<tau>\<^sub>b)"
- using assms infer_l_form3 infer_l_form4 wf_b_subst infer_l_wf subst_tb.simps base_for_lit.simps subst_tb.simps
- subst_b_base_for_lit subst_cb.simps(6) subst_ceb.simps(1) subst_vb.simps(1) subst_vb.simps(2) type_l_eq
- by metis
-
-lemma subst_b_subtype:
- fixes s::s and b'::b
- assumes "\<Theta> ; {|bv|} ; \<Gamma> \<turnstile> \<tau>1 \<lesssim> \<tau>2" and "\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f b'" and "B = {|bv|}"
- shows "\<Theta> ; {||} ; \<Gamma>[bv::=b']\<^sub>\<Gamma>\<^sub>b \<turnstile> \<tau>1[bv::=b']\<^sub>\<tau>\<^sub>b \<lesssim> \<tau>2[bv::=b']\<^sub>\<tau>\<^sub>b"
- using assms proof(nominal_induct "{|bv|}" \<Gamma> \<tau>1 \<tau>2 rule:subtype.strong_induct)
- case (subtype_baseI x \<Theta> \<Gamma> z c z' c' b)
-
- hence **: "\<Theta> ; {|bv|} ; (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<Turnstile> c'[z'::=V_var x]\<^sub>c\<^sub>v " using validI subst_defs by metis
-
- have "\<Theta> ; {||} ; \<Gamma>[bv::=b']\<^sub>\<Gamma>\<^sub>b \<turnstile> \<lbrace> z : b[bv::=b']\<^sub>b\<^sub>b | c[bv::=b']\<^sub>c\<^sub>b \<rbrace> \<lesssim> \<lbrace> z' : b[bv::=b']\<^sub>b\<^sub>b | c'[bv::=b']\<^sub>c\<^sub>b \<rbrace>" proof(rule Typing.subtype_baseI)
- show "\<Theta> ; {||} ; \<Gamma>[bv::=b']\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b[bv::=b']\<^sub>b\<^sub>b | c[bv::=b']\<^sub>c\<^sub>b \<rbrace> "
- using subtype_baseI assms wf_b_subst(4) subst_tb.simps subst_defs by metis
- show "\<Theta> ; {||} ; \<Gamma>[bv::=b']\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f \<lbrace> z' : b[bv::=b']\<^sub>b\<^sub>b | c'[bv::=b']\<^sub>c\<^sub>b \<rbrace> "
- using subtype_baseI assms wf_b_subst(4) subst_tb.simps by metis
- show "atom x \<sharp>(\<Theta>, {||}::bv fset, \<Gamma>[bv::=b']\<^sub>\<Gamma>\<^sub>b, z , c[bv::=b']\<^sub>c\<^sub>b , z' , c'[bv::=b']\<^sub>c\<^sub>b )"
- apply(unfold fresh_prodN,auto simp add: * fresh_prodN fresh_empty_fset)
- using subst_b_fresh_x * fresh_prodN \<open>atom x \<sharp> c\<close> \<open>atom x \<sharp> c'\<close> subst_defs subtype_baseI by metis+
- have "\<Theta> ; {||} ; (x, b[bv::=b']\<^sub>b\<^sub>b, c[z::=V_var x]\<^sub>v[bv::=b']\<^sub>c\<^sub>b) #\<^sub>\<Gamma> \<Gamma>[bv::=b']\<^sub>\<Gamma>\<^sub>b \<Turnstile> c'[z'::=V_var x]\<^sub>v[bv::=b']\<^sub>c\<^sub>b"
- using ** subst_b_valid subst_gb.simps assms subtype_baseI by metis
- thus "\<Theta> ; {||} ; (x, b[bv::=b']\<^sub>b\<^sub>b, (c[bv::=b']\<^sub>c\<^sub>b)[z::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>[bv::=b']\<^sub>\<Gamma>\<^sub>b \<Turnstile> (c'[bv::=b']\<^sub>c\<^sub>b)[z'::=V_var x]\<^sub>v"
- using subst_defs subst_cv_subst_bb_commute by (metis subst_cx_subst_bb_commute)
- qed
- thus ?case using subtype_baseI subst_tb.simps subst_defs by metis
-qed
-
-lemma b_of_subst_bv:
- "(b_of \<tau>)[x::=v]\<^sub>b\<^sub>b = b_of (\<tau>[x::=v]\<^sub>\<tau>\<^sub>b)"
-proof -
- obtain z b c where *:"\<tau> = \<lbrace> z : b | c \<rbrace> \<and> atom z \<sharp> (x,v)" using obtain_fresh_z by metis
- thus ?thesis using subst_tv.simps * by auto
-qed
-
-lemma subst_b_infer_v:
- fixes v::v and b::b
- assumes "\<Theta> ; B ; G \<turnstile> v \<Rightarrow> \<tau>" and "\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f b" and "B = {|bv|}"
- shows "\<Theta> ; {||} ; G[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> v[bv::=b]\<^sub>v\<^sub>b \<Rightarrow> (\<tau>[bv::=b]\<^sub>\<tau>\<^sub>b)"
- using assms proof(nominal_induct avoiding: b bv rule: infer_v.strong_induct)
- case (infer_v_varI \<Theta> \<B> \<Gamma> b' c x z)
- show ?case unfolding subst_b_simps proof
- show "\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b " using infer_v_varI wf_b_subst by metis
- show "Some (b'[bv::=b]\<^sub>b\<^sub>b, c[bv::=b]\<^sub>c\<^sub>b) = lookup \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b x" using subst_b_lookup infer_v_varI by metis
- show "atom z \<sharp> x" using infer_v_varI by auto
- show "atom z \<sharp> (\<Theta>, {||}, \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b) " by(fresh_mth add: infer_v_varI subst_b_fresh_x subst_b_\<Gamma>_def fresh_prodN fresh_empty_fset )
- qed
-next
- case (infer_v_litI \<Theta> \<B> \<Gamma> l \<tau>)
- then show ?case using Typing.infer_v_litI subst_b_infer_b
- using wf_b_subst1(3) by auto
-next
- case (infer_v_pairI z v1 v2 \<Theta> \<B> \<Gamma> t1 t2)
- show ?case unfolding subst_b_simps b_of_subst_bv proof
- show "atom z \<sharp> (v1[bv::=b]\<^sub>v\<^sub>b, v2[bv::=b]\<^sub>v\<^sub>b)" by(fresh_mth add: infer_v_pairI subst_b_fresh_x)
- show "atom z \<sharp> (\<Theta>, {||}, \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b)" by(fresh_mth add: infer_v_pairI subst_b_fresh_x subst_b_\<Gamma>_def fresh_empty_fset)
- show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> v1[bv::=b]\<^sub>v\<^sub>b \<Rightarrow> t1[bv::=b]\<^sub>\<tau>\<^sub>b" using infer_v_pairI by auto
- show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> v2[bv::=b]\<^sub>v\<^sub>b \<Rightarrow> t2[bv::=b]\<^sub>\<tau>\<^sub>b" using infer_v_pairI by auto
- qed
-next
- case (infer_v_consI s dclist \<Theta> dc tc \<B> \<Gamma> v tv z)
- show ?case unfolding subst_b_simps b_of_subst_bv proof
- show "AF_typedef s dclist \<in> set \<Theta>" using infer_v_consI by auto
- show "(dc, tc) \<in> set dclist" using infer_v_consI by auto
- show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> v[bv::=b]\<^sub>v\<^sub>b \<Rightarrow> tv[bv::=b]\<^sub>\<tau>\<^sub>b" using infer_v_consI by auto
- show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> tv[bv::=b]\<^sub>\<tau>\<^sub>b \<lesssim> tc" proof -
- have "atom bv \<sharp> tc" using wfTh_lookup_supp_empty fresh_def infer_v_consI infer_v_wf by fast
- moreover have "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> tv[bv::=b]\<^sub>\<tau>\<^sub>b \<lesssim> tc[bv::=b]\<^sub>\<tau>\<^sub>b"
- using subst_b_subtype infer_v_consI by simp
- ultimately show ?thesis using forget_subst subst_b_\<tau>_def by metis
- qed
- show "atom z \<sharp> v[bv::=b]\<^sub>v\<^sub>b" using infer_v_consI using subst_b_fresh_x subst_b_v_def by metis
- show "atom z \<sharp> (\<Theta>, {||}, \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b)" by(fresh_mth add: infer_v_consI subst_b_fresh_x subst_b_\<Gamma>_def fresh_empty_fset)
- qed
-next
- case (infer_v_conspI s bv2 dclist2 \<Theta> dc tc \<B> \<Gamma> v tv ba z)
-
- have "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> V_consp s dc (ba[bv::=b]\<^sub>b\<^sub>b) (v[bv::=b]\<^sub>v\<^sub>b) \<Rightarrow> \<lbrace> z : B_app s (ba[bv::=b]\<^sub>b\<^sub>b) | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ V_consp s dc (ba[bv::=b]\<^sub>b\<^sub>b) (v[bv::=b]\<^sub>v\<^sub>b) ]\<^sup>c\<^sup>e \<rbrace>"
- proof(rule Typing.infer_v_conspI)
- show "AF_typedef_poly s bv2 dclist2 \<in> set \<Theta>" using infer_v_conspI by auto
- show "(dc, tc) \<in> set dclist2" using infer_v_conspI by auto
- show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> v[bv::=b]\<^sub>v\<^sub>b \<Rightarrow> tv[bv::=b]\<^sub>\<tau>\<^sub>b"
- using infer_v_conspI subst_tb.simps by metis
-
- show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> tv[bv::=b]\<^sub>\<tau>\<^sub>b \<lesssim> tc[bv2::=ba[bv::=b]\<^sub>b\<^sub>b]\<^sub>\<tau>\<^sub>b" proof -
- have "supp tc \<subseteq> { atom bv2 }" using infer_v_conspI wfTh_poly_lookup_supp wfX_wfY by metis
- moreover have "bv2 \<noteq> bv" using \<open>atom bv2 \<sharp> \<B>\<close> \<open>\<B> = {|bv|} \<close> fresh_at_base fresh_def
- using fresh_finsert by fastforce
- ultimately have "atom bv \<sharp> tc" unfolding fresh_def by auto
- hence "tc[bv2::=ba[bv::=b]\<^sub>b\<^sub>b]\<^sub>\<tau>\<^sub>b = tc[bv2::=ba]\<^sub>\<tau>\<^sub>b[bv::=b]\<^sub>\<tau>\<^sub>b"
- using subst_tb_commute by metis
- moreover have "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> tv[bv::=b]\<^sub>\<tau>\<^sub>b \<lesssim> tc[bv2::=ba]\<^sub>\<tau>\<^sub>b[bv::=b]\<^sub>\<tau>\<^sub>b"
- using infer_v_conspI(7) subst_b_subtype infer_v_conspI by metis
- ultimately show ?thesis by auto
- qed
- show "atom z \<sharp> (\<Theta>, {||}, \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b, v[bv::=b]\<^sub>v\<^sub>b, ba[bv::=b]\<^sub>b\<^sub>b)"
- apply(unfold fresh_prodN, intro conjI, auto simp add: infer_v_conspI fresh_empty_fset)
- using \<open>atom z \<sharp> \<Gamma>\<close> fresh_subst_if subst_b_\<Gamma>_def x_fresh_b apply metis
- using \<open>atom z \<sharp> v\<close> fresh_subst_if subst_b_v_def x_fresh_b by metis
- show "atom bv2 \<sharp> (\<Theta>, {||}, \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b, v[bv::=b]\<^sub>v\<^sub>b, ba[bv::=b]\<^sub>b\<^sub>b)"
- apply(unfold fresh_prodN, intro conjI, auto simp add: infer_v_conspI fresh_empty_fset)
- using \<open>atom bv2 \<sharp> b\<close> \<open>atom bv2 \<sharp> \<Gamma>\<close> fresh_subst_if subst_b_\<Gamma>_def apply metis
- using \<open>atom bv2 \<sharp> b\<close> \<open>atom bv2 \<sharp> v\<close> fresh_subst_if subst_b_v_def apply metis
- using \<open>atom bv2 \<sharp> b\<close> \<open>atom bv2 \<sharp> ba\<close> fresh_subst_if subst_b_b_def by metis
- show "\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f ba[bv::=b]\<^sub>b\<^sub>b "
- using infer_v_conspI wf_b_subst by metis
- qed
- thus ?case using subst_vb.simps subst_tb.simps subst_bb.simps by simp
-
-qed
-
-lemma subst_b_check_v:
- fixes v::v and b::b
- assumes "\<Theta> ; B ; G \<turnstile> v \<Leftarrow> \<tau>" and "\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f b" and "B = {|bv|}"
- shows "\<Theta> ; {||} ; G[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> v[bv::=b]\<^sub>v\<^sub>b \<Leftarrow> (\<tau>[bv::=b]\<^sub>\<tau>\<^sub>b)"
-proof -
- obtain \<tau>' where "\<Theta> ; B ; G \<turnstile> v \<Rightarrow> \<tau>' \<and> \<Theta> ; B ; G \<turnstile> \<tau>' \<lesssim> \<tau>" using check_v_elims[OF assms(1)] by metis
- thus ?thesis using subst_b_subtype subst_b_infer_v assms
- by (metis (no_types) check_v_subtypeI subst_b_infer_v subst_b_subtype)
-qed
-
-lemma subst_vv_subst_vb_switch:
- shows "(v'[bv::=b']\<^sub>v\<^sub>b)[x::=v[bv::=b']\<^sub>v\<^sub>b]\<^sub>v\<^sub>v = v'[x::=v]\<^sub>v\<^sub>v[bv::=b']\<^sub>v\<^sub>b"
-proof(nominal_induct v' rule:v.strong_induct)
- case (V_lit x)
- then show ?case using subst_vv.simps subst_vb.simps by auto
-next
- case (V_var x)
- then show ?case using subst_vv.simps subst_vb.simps by auto
-next
- case (V_pair x1a x2a)
- then show ?case using subst_vv.simps subst_vb.simps v.fresh by auto
-next
- case (V_cons x1a x2a x3)
- then show ?case using subst_vv.simps subst_vb.simps v.fresh by auto
-next
- case (V_consp x1a x2a x3 x4)
- then show ?case using subst_vv.simps subst_vb.simps v.fresh pure_fresh
- by (metis forget_subst subst_b_b_def)
-qed
-
-lemma subst_cev_subst_vb_switch:
- shows "(ce[bv::=b']\<^sub>c\<^sub>e\<^sub>b)[x::=v[bv::=b']\<^sub>v\<^sub>b]\<^sub>c\<^sub>e\<^sub>v = (ce[x::=v]\<^sub>c\<^sub>e\<^sub>v)[bv::=b']\<^sub>c\<^sub>e\<^sub>b"
- by(nominal_induct ce rule:ce.strong_induct, auto simp add: subst_vv_subst_vb_switch ce.fresh)
-
-lemma subst_cv_subst_vb_switch:
- shows "(c[bv::=b']\<^sub>c\<^sub>b)[x::=v[bv::=b']\<^sub>v\<^sub>b]\<^sub>c\<^sub>v = c[x::=v]\<^sub>c\<^sub>v[bv::=b']\<^sub>c\<^sub>b"
- by(nominal_induct c rule:c.strong_induct, auto simp add: subst_cev_subst_vb_switch c.fresh)
-
-lemma subst_tv_subst_vb_switch:
- shows "(\<tau>[bv::=b']\<^sub>\<tau>\<^sub>b)[x::=v[bv::=b']\<^sub>v\<^sub>b]\<^sub>\<tau>\<^sub>v = \<tau>[x::=v]\<^sub>\<tau>\<^sub>v[bv::=b']\<^sub>\<tau>\<^sub>b"
-proof(nominal_induct \<tau> avoiding: x v rule:\<tau>.strong_induct)
- case (T_refined_type z b c )
- hence ceq: "(c[bv::=b']\<^sub>c\<^sub>b)[x::=v[bv::=b']\<^sub>v\<^sub>b]\<^sub>c\<^sub>v = c[x::=v]\<^sub>c\<^sub>v[bv::=b']\<^sub>c\<^sub>b" using subst_cv_subst_vb_switch by auto
-
- moreover have "atom z \<sharp> v[bv::=b']\<^sub>v\<^sub>b" using x_fresh_b fresh_subst_if subst_b_v_def T_refined_type by metis
-
- hence "\<lbrace> z : b | c \<rbrace>[bv::=b']\<^sub>\<tau>\<^sub>b[x::=v[bv::=b']\<^sub>v\<^sub>b]\<^sub>\<tau>\<^sub>v = \<lbrace> z : b[bv::=b']\<^sub>b\<^sub>b | (c[bv::=b']\<^sub>c\<^sub>b)[x::=v[bv::=b']\<^sub>v\<^sub>b]\<^sub>c\<^sub>v \<rbrace>"
- using subst_tv.simps subst_tb.simps T_refined_type fresh_Pair by metis
-
- moreover have "\<lbrace> z : b[bv::=b']\<^sub>b\<^sub>b | (c[bv::=b']\<^sub>c\<^sub>b)[x::=v[bv::=b']\<^sub>v\<^sub>b]\<^sub>c\<^sub>v \<rbrace> = \<lbrace> z : b | c[x::=v]\<^sub>c\<^sub>v \<rbrace>[bv::=b']\<^sub>\<tau>\<^sub>b"
- using subst_tv.simps subst_tb.simps ceq \<tau>.fresh forget_subst[of "bv" b b'] subst_b_b_def T_refined_type by metis
-
- ultimately show ?case using subst_tv.simps subst_tb.simps ceq T_refined_type by auto
-qed
-
-lemma subst_tb_triple:
- assumes "atom bv \<sharp> \<tau>'"
- shows "\<tau>'[bv'::=b'[bv::=b]\<^sub>b\<^sub>b]\<^sub>\<tau>\<^sub>b[x'::=v'[bv::=b]\<^sub>v\<^sub>b]\<^sub>\<tau>\<^sub>v = \<tau>'[bv'::=b']\<^sub>\<tau>\<^sub>b[x'::=v']\<^sub>\<tau>\<^sub>v[bv::=b]\<^sub>\<tau>\<^sub>b"
-proof -
- have "\<tau>'[bv'::=b'[bv::=b]\<^sub>b\<^sub>b]\<^sub>\<tau>\<^sub>b[x'::=v'[bv::=b]\<^sub>v\<^sub>b]\<^sub>\<tau>\<^sub>v = \<tau>'[bv'::=b']\<^sub>\<tau>\<^sub>b[bv::=b]\<^sub>\<tau>\<^sub>b [x'::=v'[bv::=b]\<^sub>v\<^sub>b]\<^sub>\<tau>\<^sub>v"
- using subst_tb_commute \<open>atom bv \<sharp> \<tau>'\<close> by auto
- also have "... = \<tau>'[bv'::=b']\<^sub>\<tau>\<^sub>b [x'::=v']\<^sub>\<tau>\<^sub>v[bv::=b]\<^sub>\<tau>\<^sub>b"
- using subst_tv_subst_vb_switch by auto
- finally show ?thesis using fresh_subst_if forget_subst by auto
-qed
-
-lemma subst_b_infer_e:
- fixes s::s and b::b
- assumes "\<Theta> ; \<Phi> ; B ; G; D \<turnstile> e \<Rightarrow> \<tau>" and "\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f b" and "B = {|bv|}"
- shows "\<Theta> ; \<Phi> ; {||} ; G[bv::=b]\<^sub>\<Gamma>\<^sub>b; D[bv::=b]\<^sub>\<Delta>\<^sub>b \<turnstile> (e[bv::=b]\<^sub>e\<^sub>b) \<Rightarrow> (\<tau>[bv::=b]\<^sub>\<tau>\<^sub>b)"
- using assms proof(nominal_induct avoiding: b rule: infer_e.strong_induct)
- case (infer_e_valI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v \<tau>)
- thus ?case using subst_eb.simps infer_e.intros wf_b_subst subst_db.simps wf_b_subst infer_v_wf subst_b_infer_v
- by (metis forget_subst ms_fresh_all(1) wfV_b_fresh)
-next
- case (infer_e_plusI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v1 z1 c1 v2 z2 c2 z3)
-
- show ?case unfolding subst_b_simps subst_eb.simps proof(rule Typing.infer_e_plusI)
- show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b " using wf_b_subst(10) subst_db.simps infer_e_plusI wfX_wfY
- by (metis wf_b_subst(15))
- show "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> " using infer_e_plusI by auto
- show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> v1[bv::=b]\<^sub>v\<^sub>b \<Rightarrow> \<lbrace> z1 : B_int | c1[bv::=b]\<^sub>c\<^sub>b \<rbrace>" using subst_b_infer_v infer_e_plusI subst_b_simps by force
- show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> v2[bv::=b]\<^sub>v\<^sub>b \<Rightarrow> \<lbrace> z2 : B_int | c2[bv::=b]\<^sub>c\<^sub>b \<rbrace>" using subst_b_infer_v infer_e_plusI subst_b_simps by force
- show "atom z3 \<sharp> AE_op Plus (v1[bv::=b]\<^sub>v\<^sub>b) (v2[bv::=b]\<^sub>v\<^sub>b)" using subst_b_simps infer_e_plusI subst_b_fresh_x subst_b_e_def by metis
- show "atom z3 \<sharp> \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b" using subst_g_b_x_fresh infer_e_plusI by auto
- qed
-next
- case (infer_e_leqI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v1 z1 c1 v2 z2 c2 z3)
- show ?case unfolding subst_b_simps proof(rule Typing.infer_e_leqI)
- show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b " using wf_b_subst(10) subst_db.simps infer_e_leqI wfX_wfY
- by (metis wf_b_subst(15))
- show "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> " using infer_e_leqI by auto
- show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> v1[bv::=b]\<^sub>v\<^sub>b \<Rightarrow> \<lbrace> z1 : B_int | c1[bv::=b]\<^sub>c\<^sub>b \<rbrace>" using subst_b_infer_v infer_e_leqI subst_b_simps by force
- show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> v2[bv::=b]\<^sub>v\<^sub>b \<Rightarrow> \<lbrace> z2 : B_int | c2[bv::=b]\<^sub>c\<^sub>b \<rbrace>" using subst_b_infer_v infer_e_leqI subst_b_simps by force
- show "atom z3 \<sharp> AE_op LEq (v1[bv::=b]\<^sub>v\<^sub>b) (v2[bv::=b]\<^sub>v\<^sub>b)" using subst_b_simps infer_e_leqI subst_b_fresh_x subst_b_e_def by metis
- show "atom z3 \<sharp> \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b" using subst_g_b_x_fresh infer_e_leqI by auto
- qed
-next
- case (infer_e_eqI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v1 z1 bb c1 v2 z2 c2 z3)
- show ?case unfolding subst_b_simps proof(rule Typing.infer_e_eqI)
- show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b " using wf_b_subst(10) subst_db.simps infer_e_eqI wfX_wfY
- by (metis wf_b_subst(15))
- show "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> " using infer_e_eqI by auto
- show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> v1[bv::=b]\<^sub>v\<^sub>b \<Rightarrow> \<lbrace> z1 : bb[bv::=b]\<^sub>b\<^sub>b | c1[bv::=b]\<^sub>c\<^sub>b \<rbrace>" using subst_b_infer_v infer_e_eqI subst_b_simps by force
- show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> v2[bv::=b]\<^sub>v\<^sub>b \<Rightarrow> \<lbrace> z2 : bb[bv::=b]\<^sub>b\<^sub>b | c2[bv::=b]\<^sub>c\<^sub>b \<rbrace>" using subst_b_infer_v infer_e_eqI subst_b_simps by force
- show "atom z3 \<sharp> AE_op Eq (v1[bv::=b]\<^sub>v\<^sub>b) (v2[bv::=b]\<^sub>v\<^sub>b)" using subst_b_simps infer_e_eqI subst_b_fresh_x subst_b_e_def by metis
- show "atom z3 \<sharp> \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b" using subst_g_b_x_fresh infer_e_eqI by auto
- show "bb[bv::=b]\<^sub>b\<^sub>b \<in> {B_bool, B_int, B_unit}" using infer_e_eqI by auto
- qed
-next
- case (infer_e_appI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> f x b' c \<tau>' s' v \<tau>)
- show ?case proof(subst subst_eb.simps, rule Typing.infer_e_appI)
- show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b" using wf_b_subst(10) subst_db.simps infer_e_appI wfX_wfY by (metis wf_b_subst(15))
- show "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> " using infer_e_appI by auto
- show "Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x b' c \<tau>' s'))) = lookup_fun \<Phi> f" using infer_e_appI by auto
- (*show "\<Theta> ; {||} ; (x, b', c) #\<^sub>\<Gamma> GNil \<turnstile>\<^sub>w\<^sub>f \<tau>'" using infer_e_appI sory*)
- have "atom bv \<sharp> b'" using \<open>\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> infer_e_appI wfPhi_f_supp fresh_def[of "atom bv" b'] by simp
- hence "b' = b'[bv::=b]\<^sub>b\<^sub>b" using subst_b_simps
- using has_subst_b_class.forget_subst subst_b_b_def by force
- moreover have ceq:"c = c[bv::=b]\<^sub>c\<^sub>b" using subst_b_simps proof -
- have "supp c \<subseteq> {atom x}" using infer_e_appI wfPhi_f_simple_supp_c[OF _ \<open>\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close>] by simp
- hence "atom bv \<sharp> c" using
- fresh_def[of "atom bv" c]
- using fresh_def fresh_finsert insert_absorb
- insert_subset ms_fresh_all supp_at_base x_not_in_b_set fresh_prodN by metis
- thus ?thesis
- using forget_subst subst_b_c_def fresh_def[of "atom bv" c] by metis
- qed
- show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> v[bv::=b]\<^sub>v\<^sub>b \<Leftarrow> \<lbrace> x : b' | c \<rbrace>"
- using subst_b_check_v subst_tb.simps subst_vb.simps infer_e_appI
- proof -
- have "\<Theta> ; {|bv|} ; \<Gamma> \<turnstile> v \<Leftarrow> \<lbrace> x : b' | c \<rbrace>"
- by (metis \<open>\<B> = {|bv|}\<close> \<open>\<Theta> ; \<B> ; \<Gamma> \<turnstile> v \<Leftarrow> \<lbrace> x : b' | c \<rbrace>\<close>) (* 0.0 ms *)
- then show ?thesis
- by (metis (no_types) \<open>\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f b\<close> \<open>b' = b'[bv::=b]\<^sub>b\<^sub>b\<close> subst_b_check_v subst_tb.simps ceq)
- qed
- show "atom x \<sharp> (\<Theta>, \<Phi>, {||}::bv fset, \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b, \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b, v[bv::=b]\<^sub>v\<^sub>b, \<tau>[bv::=b]\<^sub>\<tau>\<^sub>b)"
- apply (fresh_mth add: fresh_prodN subst_g_b_x_fresh infer_e_appI )
- using subst_b_fresh_x infer_e_appI apply metis+
- done
- have "supp \<tau>' \<subseteq> { atom x }" using wfPhi_f_simple_supp_t infer_e_appI by auto
- hence "atom bv \<sharp> \<tau>'" using fresh_def fresh_at_base by force
- then show "\<tau>'[x::=v[bv::=b]\<^sub>v\<^sub>b]\<^sub>v = \<tau>[bv::=b]\<^sub>\<tau>\<^sub>b" using infer_e_appI
- forget_subst subst_b_\<tau>_def subst_tv_subst_vb_switch subst_defs by metis
- qed
-next
- case (infer_e_appPI \<Theta>' \<B> \<Gamma>' \<Delta> \<Phi>' b' f' bv' x' b1 c \<tau>' s' v' \<tau>1)
-
- have beq: "b1[bv'::=b']\<^sub>b\<^sub>b[bv::=b]\<^sub>b\<^sub>b = b1[bv'::=b'[bv::=b]\<^sub>b\<^sub>b]\<^sub>b\<^sub>b"
- proof -
- have "supp b1 \<subseteq> { atom bv' }" using wfPhi_f_poly_supp_b infer_e_appPI
- using supp_at_base by blast
- moreover have "bv \<noteq> bv'" using infer_e_appPI fresh_def supp_at_base
- by (simp add: fresh_def supp_at_base)
- ultimately have "atom bv \<sharp> b1" using fresh_def fresh_at_base by force
- thus ?thesis by simp
- qed
-
- have ceq: "(c[bv'::=b']\<^sub>c\<^sub>b)[bv::=b]\<^sub>c\<^sub>b = c[bv'::=b'[bv::=b]\<^sub>b\<^sub>b]\<^sub>c\<^sub>b" proof -
- have "supp c \<subseteq> { atom bv', atom x' }" using wfPhi_f_poly_supp_c infer_e_appPI
- using supp_at_base by blast
- moreover have "bv \<noteq> bv'" using infer_e_appPI fresh_def supp_at_base
- by (simp add: fresh_def supp_at_base)
- moreover have "atom x' \<noteq> atom bv" by auto
- ultimately have "atom bv \<sharp> c" using fresh_def[of "atom bv" c] fresh_at_base by auto
- thus ?thesis by simp
- qed
-
- show ?case proof(subst subst_eb.simps, rule Typing.infer_e_appPI)
- show "\<Theta>' ; {||} ; \<Gamma>'[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b " using wf_b_subst subst_db.simps infer_e_appPI wfX_wfY by metis
- show "\<Theta>' \<turnstile>\<^sub>w\<^sub>f \<Phi>' " using infer_e_appPI by auto
- show "Some (AF_fundef f' (AF_fun_typ_some bv' (AF_fun_typ x' b1 c \<tau>' s'))) = lookup_fun \<Phi>' f'" using infer_e_appPI by auto
- thus "\<Theta>' ; {||} ; \<Gamma>'[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> v'[bv::=b]\<^sub>v\<^sub>b \<Leftarrow> \<lbrace> x' : b1[bv'::=b'[bv::=b]\<^sub>b\<^sub>b]\<^sub>b | c[bv'::=b'[bv::=b]\<^sub>b\<^sub>b]\<^sub>b \<rbrace>"
- using subst_b_check_v subst_tb.simps subst_b_simps infer_e_appPI
- proof -
- have "\<Theta>' ; {||} ; \<Gamma>'[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> v'[bv::=b]\<^sub>v\<^sub>b \<Leftarrow> \<lbrace> x' : b1[bv'::=b']\<^sub>b[bv::=b]\<^sub>b\<^sub>b | (c[bv'::=b']\<^sub>b)[bv::=b]\<^sub>c\<^sub>b \<rbrace>"
- using infer_e_appPI subst_b_check_v subst_tb.simps by metis
- thus ?thesis using beq ceq subst_defs by metis
- qed
- show "atom x' \<sharp> \<Gamma>'[bv::=b]\<^sub>\<Gamma>\<^sub>b" using subst_g_b_x_fresh infer_e_appPI by auto
- show "\<tau>'[bv'::=b'[bv::=b]\<^sub>b\<^sub>b]\<^sub>b[x'::=v'[bv::=b]\<^sub>v\<^sub>b]\<^sub>v = \<tau>1[bv::=b]\<^sub>\<tau>\<^sub>b" proof - (* \<tau>'[bv'::=b']\<^sub>\<tau>\<^sub>b[x'::=v']\<^sub>\<tau>\<^sub>v = \<tau>1 *)
-
- have "supp \<tau>' \<subseteq> { atom x', atom bv' }" using wfPhi_f_poly_supp_t infer_e_appPI by auto
- moreover hence "bv \<noteq> bv'" using infer_e_appPI fresh_def supp_at_base
- by (simp add: fresh_def supp_at_base)
- ultimately have "atom bv \<sharp> \<tau>'" using fresh_def by force
- hence "\<tau>'[bv'::=b'[bv::=b]\<^sub>b\<^sub>b]\<^sub>b[x'::=v'[bv::=b]\<^sub>v\<^sub>b]\<^sub>v = \<tau>'[bv'::=b']\<^sub>b[x'::=v']\<^sub>v[bv::=b]\<^sub>\<tau>\<^sub>b" using subst_tb_triple subst_defs by auto
- thus ?thesis using infer_e_appPI by metis
- qed
- show "atom bv' \<sharp> (\<Theta>', \<Phi>', {||}, \<Gamma>'[bv::=b]\<^sub>\<Gamma>\<^sub>b, \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b, b'[bv::=b]\<^sub>b\<^sub>b, v'[bv::=b]\<^sub>v\<^sub>b, \<tau>1[bv::=b]\<^sub>\<tau>\<^sub>b)"
- unfolding fresh_prodN apply( auto simp add: infer_e_appPI fresh_empty_fset)
- using fresh_subst_if subst_b_\<Gamma>_def subst_b_\<Delta>_def subst_b_b_def subst_b_v_def subst_b_\<tau>_def infer_e_appPI by metis+
- show "\<Theta>' ; {||} \<turnstile>\<^sub>w\<^sub>f b'[bv::=b]\<^sub>b\<^sub>b " using infer_e_appPI wf_b_subst by simp
- qed
-next
- case (infer_e_fstI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v z' b1 b2 c z)
- show ?case unfolding subst_b_simps proof(rule Typing.infer_e_fstI)
- show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b " using wf_b_subst(10) subst_db.simps infer_e_fstI wfX_wfY
- by (metis wf_b_subst(15))
- show "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> " using infer_e_fstI by auto
- show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> v[bv::=b]\<^sub>v\<^sub>b \<Rightarrow> \<lbrace> z' : B_pair b1[bv::=b]\<^sub>b\<^sub>b b2[bv::=b]\<^sub>b\<^sub>b | c[bv::=b]\<^sub>c\<^sub>b \<rbrace>"
- using subst_b_infer_v subst_tb.simps subst_b_simps infer_e_fstI by force
- show "atom z \<sharp> AE_fst (v[bv::=b]\<^sub>v\<^sub>b)" using infer_e_fstI subst_b_fresh_x subst_b_v_def e.fresh by metis
- show "atom z \<sharp> \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b" using subst_g_b_x_fresh infer_e_fstI by auto
- qed
-next
- case (infer_e_sndI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v z' b1 b2 c z)
- show ?case unfolding subst_b_simps proof(rule Typing.infer_e_sndI)
- show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b " using wf_b_subst(10) subst_db.simps infer_e_sndI wfX_wfY
- by (metis wf_b_subst(15))
- show "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> " using infer_e_sndI by auto
- show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> v[bv::=b]\<^sub>v\<^sub>b \<Rightarrow> \<lbrace> z' : B_pair b1[bv::=b]\<^sub>b\<^sub>b b2[bv::=b]\<^sub>b\<^sub>b | c[bv::=b]\<^sub>c\<^sub>b \<rbrace>"
- using subst_b_infer_v subst_tb.simps subst_b_simps infer_e_sndI by force
- show "atom z \<sharp> AE_snd (v[bv::=b]\<^sub>v\<^sub>b)" using infer_e_sndI subst_b_fresh_x subst_b_v_def e.fresh by metis
- show "atom z \<sharp> \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b" using subst_g_b_x_fresh infer_e_sndI by auto
- qed
-next
- case (infer_e_lenI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v z' c z)
- show ?case unfolding subst_b_simps proof(rule Typing.infer_e_lenI)
- show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b " using wf_b_subst(10) subst_db.simps infer_e_lenI wfX_wfY
- by (metis wf_b_subst(15))
- show "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> " using infer_e_lenI by auto
- show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> v[bv::=b]\<^sub>v\<^sub>b \<Rightarrow> \<lbrace> z' : B_bitvec | c[bv::=b]\<^sub>c\<^sub>b \<rbrace>"
- using subst_b_infer_v subst_tb.simps subst_b_simps infer_e_lenI by force
- show "atom z \<sharp> AE_len (v[bv::=b]\<^sub>v\<^sub>b)" using infer_e_lenI subst_b_fresh_x subst_b_v_def e.fresh by metis
- show "atom z \<sharp> \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b" using subst_g_b_x_fresh infer_e_lenI by auto
- qed
-next
- case (infer_e_mvarI \<Theta> \<B> \<Gamma> \<Phi> \<Delta> u \<tau>)
- show ?case proof(subst subst subst_eb.simps, rule Typing.infer_e_mvarI)
- show "\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b " using infer_e_mvarI wf_b_subst by auto
- show "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> " using infer_e_mvarI by auto
- show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b " using infer_e_mvarI using wf_b_subst(10) subst_db.simps infer_e_sndI wfX_wfY
- by (metis wf_b_subst(15))
- show "(u, \<tau>[bv::=b]\<^sub>\<tau>\<^sub>b) \<in> setD \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b" using infer_e_mvarI subst_db.simps set_insert
- subst_d_b_member by simp
- qed
-next
- case (infer_e_concatI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v1 z1 c1 v2 z2 c2 z3)
- show ?case unfolding subst_b_simps proof(rule Typing.infer_e_concatI)
- show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b " using wf_b_subst(10) subst_db.simps infer_e_concatI wfX_wfY
- by (metis wf_b_subst(15))
- show "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> " using infer_e_concatI by auto
- show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> v1[bv::=b]\<^sub>v\<^sub>b \<Rightarrow> \<lbrace> z1 : B_bitvec | c1[bv::=b]\<^sub>c\<^sub>b \<rbrace>"
- using subst_b_infer_v subst_tb.simps subst_b_simps infer_e_concatI by force
- show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> v2[bv::=b]\<^sub>v\<^sub>b \<Rightarrow> \<lbrace> z2 : B_bitvec | c2[bv::=b]\<^sub>c\<^sub>b \<rbrace>"
- using subst_b_infer_v subst_tb.simps subst_b_simps infer_e_concatI by force
- show "atom z3 \<sharp> AE_concat (v1[bv::=b]\<^sub>v\<^sub>b) (v2[bv::=b]\<^sub>v\<^sub>b)" using infer_e_concatI subst_b_fresh_x subst_b_v_def e.fresh by metis
- show "atom z3 \<sharp> \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b" using subst_g_b_x_fresh infer_e_concatI by auto
- qed
-next
- case (infer_e_splitI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v1 z1 c1 v2 z2 z3)
- show ?case unfolding subst_b_simps proof(rule Typing.infer_e_splitI)
- show \<open> \<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b \<close> using wf_b_subst(10) subst_db.simps infer_e_splitI wfX_wfY
- by (metis wf_b_subst(15))
- show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using infer_e_splitI by auto
- show \<open> \<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> v1[bv::=b]\<^sub>v\<^sub>b \<Rightarrow> \<lbrace> z1 : B_bitvec | c1[bv::=b]\<^sub>c\<^sub>b \<rbrace>\<close>
- using subst_b_infer_v subst_tb.simps subst_b_simps infer_e_splitI by force
- show \<open>\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> v2[bv::=b]\<^sub>v\<^sub>b \<Leftarrow> \<lbrace> z2 : B_int | [ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ z2 ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e AND
- [ leq [ [ z2 ]\<^sup>v ]\<^sup>c\<^sup>e [| [ v1[bv::=b]\<^sub>v\<^sub>b ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrace>\<close>
- using subst_b_check_v subst_tb.simps subst_b_simps infer_e_splitI
- proof -
- have "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> v2[bv::=b]\<^sub>v\<^sub>b \<Leftarrow> \<lbrace> z2 : B_int | [ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ z2 ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e AND [ leq [ [ z2 ]\<^sup>v ]\<^sup>c\<^sup>e [| [ v1 ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrace>[bv::=b]\<^sub>\<tau>\<^sub>b"
- using infer_e_splitI.hyps(7) infer_e_splitI.prems(1) infer_e_splitI.prems(2) subst_b_check_v by presburger (* 0.0 ms *)
- then show ?thesis
- by simp (* 0.0 ms *)
- qed
- show \<open>atom z1 \<sharp> AE_split (v1[bv::=b]\<^sub>v\<^sub>b) (v2[bv::=b]\<^sub>v\<^sub>b)\<close> using infer_e_splitI subst_b_fresh_x subst_b_v_def e.fresh by metis
- show \<open>atom z1 \<sharp> \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b\<close> using subst_g_b_x_fresh infer_e_splitI by auto
-
- show \<open>atom z2 \<sharp> AE_split (v1[bv::=b]\<^sub>v\<^sub>b) (v2[bv::=b]\<^sub>v\<^sub>b)\<close> using infer_e_splitI subst_b_fresh_x subst_b_v_def e.fresh by metis
-
- show \<open>atom z2 \<sharp> \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b\<close> using subst_g_b_x_fresh infer_e_splitI by auto
- show \<open>atom z3 \<sharp> AE_split (v1[bv::=b]\<^sub>v\<^sub>b) (v2[bv::=b]\<^sub>v\<^sub>b)\<close> using infer_e_splitI subst_b_fresh_x subst_b_v_def e.fresh by metis
- show \<open>atom z3 \<sharp> \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b\<close> using subst_g_b_x_fresh infer_e_splitI by auto
- qed
-qed
-
-text \<open>This is needed for preservation. When we apply a function "f [b] v" we need to
-substitute into the body of the function f replacing type-variable with b\<close>
-
-lemma subst_b_c_of_forget:
- assumes "atom bv \<sharp> const"
- shows "(c_of const x)[bv::=b]\<^sub>c\<^sub>b = c_of const x"
- using assms proof(nominal_induct const avoiding: x rule:\<tau>.strong_induct)
- case (T_refined_type x' b' c')
- hence "c_of \<lbrace> x' : b' | c' \<rbrace> x = c'[x'::=V_var x]\<^sub>c\<^sub>v" using c_of.simps by metis
- moreover have "atom bv \<sharp> c'[x'::=V_var x]\<^sub>c\<^sub>v" proof -
- have "atom bv \<sharp> c'" using T_refined_type \<tau>.fresh by simp
- moreover have "atom bv \<sharp> V_var x" using v.fresh by simp
- ultimately show ?thesis
- using T_refined_type \<tau>.fresh subst_b_c_def fresh_subst_if
- \<tau>_fresh_c fresh_subst_cv_if has_subst_b_class.subst_b_fresh_x ms_fresh_all(37) ms_fresh_all assms by metis
- qed
- ultimately show ?case using forget_subst subst_b_c_def by metis
-qed
-
-lemma subst_b_check_s:
- fixes s::s and b::b and cs::branch_s and css::branch_list and v::v and \<tau>::\<tau>
- assumes "\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f b" and "B = {|bv|}" (*and "D = []"*)
- shows "\<Theta> ; \<Phi> ; B ; G; D \<turnstile> s \<Leftarrow> \<tau> \<Longrightarrow> \<Theta> ; \<Phi> ; {||} ; G[bv::=b]\<^sub>\<Gamma>\<^sub>b; D[bv::=b]\<^sub>\<Delta>\<^sub>b \<turnstile> (s[bv::=b]\<^sub>s\<^sub>b) \<Leftarrow> (\<tau>[bv::=b]\<^sub>\<tau>\<^sub>b)" and
- "\<Theta> ; \<Phi> ; B ; G; D ; tid ; cons ; const ; v \<turnstile> cs \<Leftarrow> \<tau> \<Longrightarrow> \<Theta> ; \<Phi> ; {||} ; G[bv::=b]\<^sub>\<Gamma>\<^sub>b; D[bv::=b]\<^sub>\<Delta>\<^sub>b ; tid ; cons ; const ; v[bv::=b]\<^sub>v\<^sub>b \<turnstile> (subst_branchb cs bv b) \<Leftarrow> (\<tau>[bv::=b]\<^sub>\<tau>\<^sub>b)" and
- "\<Theta> ; \<Phi> ; B ; G; D ; tid ; dclist ; v \<turnstile> css \<Leftarrow> \<tau> \<Longrightarrow> \<Theta> ; \<Phi> ; {||} ; G[bv::=b]\<^sub>\<Gamma>\<^sub>b; D[bv::=b]\<^sub>\<Delta>\<^sub>b ; tid ; dclist ; v[bv::=b]\<^sub>v\<^sub>b \<turnstile> (subst_branchlb css bv b ) \<Leftarrow> (\<tau>[bv::=b]\<^sub>\<tau>\<^sub>b)"
- using assms proof(induct rule: check_s_check_branch_s_check_branch_list.inducts)
- note facts = wfD_emptyI wfX_wfY wf_b_subst subst_b_subtype subst_b_infer_v
- case (check_valI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v \<tau>' \<tau>)
- show ?case
- apply(subst subst_sb.simps, rule Typing.check_valI)
- using facts check_valI apply metis
- using check_valI subst_b_infer_v wf_b_subst subst_b_subtype apply blast
- using check_valI subst_b_infer_v wf_b_subst subst_b_subtype apply blast
- using check_valI subst_b_infer_v wf_b_subst subst_b_subtype by metis
-next
- case (check_letI x \<Theta> \<Phi> \<B> \<Gamma> \<Delta> e \<tau> z s b' c)
- show ?case proof(subst subst_sb.simps, rule Typing.check_letI)
-
- show "atom x \<sharp>(\<Theta>, \<Phi>, {||}, \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b, \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b, e[bv::=b]\<^sub>e\<^sub>b, \<tau>[bv::=b]\<^sub>\<tau>\<^sub>b)" (*using check_letI subst_b_\<tau>_def subst_b_\<Gamma>_def subst_b_fresh_x fresh_prod4 subst_b_c_def sory*)
- apply(unfold fresh_prodN,auto)
- apply(simp add: check_letI fresh_empty_fset)+
- apply(metis * subst_b_fresh_x check_letI fresh_prodN)+ done
- show "atom z \<sharp> (x, \<Theta>, \<Phi>, {||}, \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b, \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b, e[bv::=b]\<^sub>e\<^sub>b, \<tau>[bv::=b]\<^sub>\<tau>\<^sub>b, s[bv::=b]\<^sub>s\<^sub>b)"
- apply(unfold fresh_prodN,auto)
- apply(simp add: check_letI fresh_empty_fset)+
- apply(metis * subst_b_fresh_x check_letI fresh_prodN)+ done
- show "\<Theta> ; \<Phi> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b ; \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b \<turnstile> e[bv::=b]\<^sub>e\<^sub>b \<Rightarrow> \<lbrace> z : b'[bv::=b]\<^sub>b\<^sub>b | c[bv::=b]\<^sub>c\<^sub>b \<rbrace>"
- using check_letI subst_b_infer_e subst_tb.simps by metis
- have "c[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v[bv::=b]\<^sub>c\<^sub>b = (c[bv::=b]\<^sub>c\<^sub>b)[z::=V_var x]\<^sub>c\<^sub>v"
- using subst_cv_subst_bb_commute[of bv "V_var x" c z b] fresh_at_base by simp
- thus "\<Theta> ; \<Phi> ; {||} ; ((x, b'[bv::=b]\<^sub>b\<^sub>b , (c[bv::=b]\<^sub>c\<^sub>b)[z::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b) ; \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b \<turnstile> s[bv::=b]\<^sub>s\<^sub>b \<Leftarrow> \<tau>[bv::=b]\<^sub>\<tau>\<^sub>b"
- using check_letI subst_gb.simps subst_defs by metis
- qed
-next
- case (check_assertI x \<Theta> \<Phi> \<B> \<Gamma> \<Delta> c \<tau> s)
- show ?case proof(subst subst_sb.simps, rule Typing.check_assertI)
- show "atom x \<sharp> (\<Theta>, \<Phi>, {||}, \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b, \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b, c[bv::=b]\<^sub>c\<^sub>b, \<tau>[bv::=b]\<^sub>\<tau>\<^sub>b, s[bv::=b]\<^sub>s\<^sub>b)"
- apply(unfold fresh_prodN,auto)
- apply(simp add: check_assertI fresh_empty_fset)+
- apply(metis * subst_b_fresh_x check_assertI fresh_prodN)+ done
-
- have "\<Theta> ; \<Phi> ; {||} ; ((x, B_bool, c) #\<^sub>\<Gamma> \<Gamma>)[bv::=b]\<^sub>\<Gamma>\<^sub>b ; \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b \<turnstile> s[bv::=b]\<^sub>s\<^sub>b \<Leftarrow> \<tau>[bv::=b]\<^sub>\<tau>\<^sub>b" using check_assertI
- by metis
- thus "\<Theta> ; \<Phi> ; {||} ; (x, B_bool, c[bv::=b]\<^sub>c\<^sub>b) #\<^sub>\<Gamma> \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b ; \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b \<turnstile> s[bv::=b]\<^sub>s\<^sub>b \<Leftarrow> \<tau>[bv::=b]\<^sub>\<tau>\<^sub>b" using subst_gb.simps by auto
- show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<Turnstile> c[bv::=b]\<^sub>c\<^sub>b " using subst_b_valid check_assertI by simp
- show " \<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b " using wf_b_subst2(6) check_assertI by simp
- qed
-next
- case (check_branch_list_consI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist v cs \<tau> css)
- then show ?case unfolding subst_branchlb.simps using Typing.check_branch_list_consI by simp
-next
- case (check_branch_list_finalI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist v cs \<tau>)
- then show ?case unfolding subst_branchlb.simps using Typing.check_branch_list_finalI by simp
-next
- case (check_branch_s_branchI \<Theta> \<B> \<Gamma> \<Delta> \<tau> const x \<Phi> tid cons v s)
-
- show ?case unfolding subst_b_simps proof(rule Typing.check_branch_s_branchI)
- show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b " using check_branch_s_branchI wf_b_subst subst_db.simps by metis
- show "\<turnstile>\<^sub>w\<^sub>f \<Theta> " using check_branch_s_branchI by auto
- show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f \<tau>[bv::=b]\<^sub>\<tau>\<^sub>b " using check_branch_s_branchI wf_b_subst by metis
-
- show "atom x \<sharp>(\<Theta>, \<Phi>, {||}, \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b, \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b, tid, cons , const, v[bv::=b]\<^sub>v\<^sub>b, \<tau>[bv::=b]\<^sub>\<tau>\<^sub>b)"
- apply(unfold fresh_prodN,auto)
- apply(simp add: check_branch_s_branchI fresh_empty_fset)+
- apply(metis * subst_b_fresh_x check_branch_s_branchI fresh_prodN)+
- done
- show wft:"\<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f const" using check_branch_s_branchI by auto
- hence "(b_of const) = (b_of const)[bv::=b]\<^sub>b\<^sub>b"
- using wfT_nil_supp fresh_def[of "atom bv" ] forget_subst subst_b_b_def \<tau>.supp
- bot.extremum_uniqueI ex_in_conv fresh_def supp_empty_fset
- by (metis b_of_supp)
- moreover have "(c_of const x)[bv::=b]\<^sub>c\<^sub>b = c_of const x"
- using wft wfT_nil_supp fresh_def[of "atom bv" ] forget_subst subst_b_c_def \<tau>.supp
- bot.extremum_uniqueI ex_in_conv fresh_def supp_empty_fset subst_b_c_of_forget by metis
- ultimately show "\<Theta> ; \<Phi> ; {||} ; (x, b_of const, CE_val (v[bv::=b]\<^sub>v\<^sub>b) == CE_val(V_cons tid cons (V_var x)) AND c_of const x) #\<^sub>\<Gamma> \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b ; \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b \<turnstile> s[bv::=b]\<^sub>s\<^sub>b \<Leftarrow> \<tau>[bv::=b]\<^sub>\<tau>\<^sub>b"
- using check_branch_s_branchI subst_gb.simps by auto
- qed
-next
- case (check_ifI z \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v s1 s2 \<tau>)
- show ?case unfolding subst_b_simps proof(rule Typing.check_ifI)
- show \<open>atom z \<sharp> (\<Theta>, \<Phi>, {||}, \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b, \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b, v[bv::=b]\<^sub>v\<^sub>b, s1[bv::=b]\<^sub>s\<^sub>b, s2[bv::=b]\<^sub>s\<^sub>b, \<tau>[bv::=b]\<^sub>\<tau>\<^sub>b)\<close>
- by(unfold fresh_prodN,auto, auto simp add: check_ifI fresh_empty_fset subst_b_fresh_x )
- have "\<lbrace> z : B_bool | TRUE \<rbrace>[bv::=b]\<^sub>\<tau>\<^sub>b = \<lbrace> z : B_bool | TRUE \<rbrace>" by auto
- thus \<open>\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> v[bv::=b]\<^sub>v\<^sub>b \<Leftarrow> \<lbrace> z : B_bool | TRUE \<rbrace>\<close> using check_ifI subst_b_check_v by metis
- show \<open> \<Theta> ; \<Phi> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b ; \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b \<turnstile> s1[bv::=b]\<^sub>s\<^sub>b \<Leftarrow> \<lbrace> z : b_of \<tau>[bv::=b]\<^sub>\<tau>\<^sub>b | CE_val (v[bv::=b]\<^sub>v\<^sub>b) == CE_val (V_lit L_true) IMP c_of \<tau>[bv::=b]\<^sub>\<tau>\<^sub>b z \<rbrace>\<close>
- using subst_b_if check_ifI by metis
- show \<open> \<Theta> ; \<Phi> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b ; \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b \<turnstile> s2[bv::=b]\<^sub>s\<^sub>b \<Leftarrow> \<lbrace> z : b_of \<tau>[bv::=b]\<^sub>\<tau>\<^sub>b | CE_val (v[bv::=b]\<^sub>v\<^sub>b) == CE_val (V_lit L_false) IMP c_of \<tau>[bv::=b]\<^sub>\<tau>\<^sub>b z \<rbrace>\<close>
- using subst_b_if check_ifI by metis
- qed
-next
- case (check_let2I x \<Theta> \<Phi> \<B> G \<Delta> t s1 \<tau> s2 )
- show ?case unfolding subst_b_simps proof (rule Typing.check_let2I)
- have "atom x \<sharp> b" using x_fresh_b by auto
- show \<open>atom x \<sharp> (\<Theta>, \<Phi>, {||}, G[bv::=b]\<^sub>\<Gamma>\<^sub>b, \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b, t[bv::=b]\<^sub>\<tau>\<^sub>b, s1[bv::=b]\<^sub>s\<^sub>b, \<tau>[bv::=b]\<^sub>\<tau>\<^sub>b)\<close>
- apply(unfold fresh_prodN, auto, auto simp add: check_let2I fresh_prodN fresh_empty_fset)
- apply(metis subst_b_fresh_x check_let2I fresh_prodN)+
- done
-
- show \<open> \<Theta> ; \<Phi> ; {||} ; G[bv::=b]\<^sub>\<Gamma>\<^sub>b ; \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b \<turnstile> s1[bv::=b]\<^sub>s\<^sub>b \<Leftarrow> t[bv::=b]\<^sub>\<tau>\<^sub>b \<close> using check_let2I subst_tb.simps by auto
- show \<open> \<Theta> ; \<Phi> ; {||} ; (x, b_of t[bv::=b]\<^sub>\<tau>\<^sub>b, c_of t[bv::=b]\<^sub>\<tau>\<^sub>b x) #\<^sub>\<Gamma> G[bv::=b]\<^sub>\<Gamma>\<^sub>b ; \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b \<turnstile> s2[bv::=b]\<^sub>s\<^sub>b \<Leftarrow> \<tau>[bv::=b]\<^sub>\<tau>\<^sub>b\<close>
- using check_let2I subst_tb.simps subst_gb.simps b_of.simps subst_b_c_of subst_b_b_of by auto
- qed
-next
- case (check_varI u \<Theta> \<Phi> \<B> \<Gamma> \<Delta> \<tau>' v \<tau> s)
- show ?case unfolding subst_b_simps proof(rule Typing.check_varI)
- show "atom u \<sharp> (\<Theta>, \<Phi>, {||}, \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b, \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b, \<tau>'[bv::=b]\<^sub>\<tau>\<^sub>b, v[bv::=b]\<^sub>v\<^sub>b, \<tau>[bv::=b]\<^sub>\<tau>\<^sub>b) "
- by(unfold fresh_prodN,auto simp add: check_varI fresh_empty_fset subst_b_fresh_u )
- show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> v[bv::=b]\<^sub>v\<^sub>b \<Leftarrow> \<tau>'[bv::=b]\<^sub>\<tau>\<^sub>b" using check_varI subst_b_check_v by auto
- show "\<Theta> ; \<Phi> ; {||} ; (subst_gb \<Gamma> bv b) ; (u, (\<tau>'[bv::=b]\<^sub>\<tau>\<^sub>b)) #\<^sub>\<Delta> (subst_db \<Delta> bv b) \<turnstile> (s[bv::=b]\<^sub>s\<^sub>b) \<Leftarrow> (\<tau>[bv::=b]\<^sub>\<tau>\<^sub>b)" using check_varI by auto
- qed
-next
- case (check_assignI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> u \<tau> v z \<tau>')
- show ?case unfolding subst_b_simps proof( rule Typing.check_assignI)
- show "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> " using check_assignI by auto
- show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b " using wf_b_subst check_assignI by auto
- show "(u, \<tau>[bv::=b]\<^sub>\<tau>\<^sub>b) \<in> setD \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b" using check_assignI subst_d_b_member by simp
- show " \<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> v[bv::=b]\<^sub>v\<^sub>b \<Leftarrow> \<tau>[bv::=b]\<^sub>\<tau>\<^sub>b" using check_assignI subst_b_check_v by auto
- show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> \<lbrace> z : B_unit | TRUE \<rbrace> \<lesssim> \<tau>'[bv::=b]\<^sub>\<tau>\<^sub>b" using check_assignI subst_b_subtype subst_b_simps subst_tb.simps by fastforce
- qed
-next
- case (check_whileI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> s1 z s2 \<tau>')
- show ?case unfolding subst_b_simps proof(rule Typing.check_whileI)
- show "\<Theta> ; \<Phi> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b ; \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b \<turnstile> s1[bv::=b]\<^sub>s\<^sub>b \<Leftarrow> \<lbrace> z : B_bool | TRUE \<rbrace>" using check_whileI by auto
- show "\<Theta> ; \<Phi> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b ; \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b \<turnstile> s2[bv::=b]\<^sub>s\<^sub>b \<Leftarrow> \<lbrace> z : B_unit | TRUE \<rbrace>" using check_whileI by auto
- show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> \<lbrace> z : B_unit | TRUE \<rbrace> \<lesssim> \<tau>'[bv::=b]\<^sub>\<tau>\<^sub>b" using subst_b_subtype check_whileI by fastforce
- qed
-next
- case (check_seqI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> s1 z s2 \<tau>)
- then show ?case unfolding subst_sb.simps using check_seqI Typing.check_seqI subst_b_eq by metis
-next
- case (check_caseI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist v cs \<tau> z)
- show ?case unfolding subst_b_simps proof(rule Typing.check_caseI)
- show \<open> \<Theta> ; \<Phi> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b ; \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b ; tid ; dclist ; v[bv::=b]\<^sub>v\<^sub>b \<turnstile> subst_branchlb cs bv b \<Leftarrow> \<tau>[bv::=b]\<^sub>\<tau>\<^sub>b\<close> using check_caseI by auto
- show \<open>AF_typedef tid dclist \<in> set \<Theta>\<close> using check_caseI by auto
- show \<open>\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> v[bv::=b]\<^sub>v\<^sub>b \<Leftarrow> \<lbrace> z : B_id tid | TRUE \<rbrace>\<close> using check_caseI subst_b_check_v subst_b_simps subst_tb.simps subst_b_simps
- proof -
- have "\<lbrace> z : B_id tid | TRUE \<rbrace> = \<lbrace> z : B_id tid | TRUE \<rbrace>[bv::=b]\<^sub>\<tau>\<^sub>b" using subst_b_eq by auto
- then show ?thesis
- by (metis (no_types) check_caseI.hyps(4) check_caseI.prems(1) check_caseI.prems(2) subst_b_check_v) (* 31 ms *)
- qed
- show \<open> \<turnstile>\<^sub>w\<^sub>f \<Theta> \<close> using check_caseI by auto
- qed
-qed
-
-end
+(*<*)
+theory BTVSubstTypingL
+ imports "HOL-Eisbach.Eisbach_Tools" ContextSubtypingL SubstMethods
+begin
+ (*>*)
+
+chapter \<open>Basic Type Variable Substitution Lemmas\<close>
+text \<open>Lemmas that show that types are preserved, in some way,
+under basic type variable substitution\<close>
+
+lemma subst_vv_subst_bb_commute:
+ fixes bv::bv and b::b and x::x and v::v
+ assumes "atom bv \<sharp> v"
+ shows "(v'[x::=v]\<^sub>v\<^sub>v)[bv::=b]\<^sub>v\<^sub>b = (v'[bv::=b]\<^sub>v\<^sub>b)[x::=v]\<^sub>v\<^sub>v"
+ using assms proof(nominal_induct v' rule: v.strong_induct)
+ case (V_lit x)
+ then show ?case using subst_vb.simps subst_vv.simps by simp
+next
+ case (V_var y)
+ hence "v[bv::=b]\<^sub>v\<^sub>b=v" using forget_subst subst_b_v_def by metis
+ then show ?case unfolding subst_vb.simps(2) subst_vv.simps(2) using V_var by auto
+next
+ case (V_pair x1a x2a)
+ then show ?case using subst_vb.simps subst_vv.simps by simp
+next
+ case (V_cons x1a x2a x3)
+ then show ?case using subst_vb.simps subst_vv.simps by simp
+next
+ case (V_consp x1a x2a x3 x4)
+ then show ?case using subst_vb.simps subst_vv.simps by simp
+qed
+
+lemma subst_cev_subst_bb_commute:
+ fixes bv::bv and b::b and x::x and v::v
+ assumes "atom bv \<sharp> v"
+ shows "(ce[x::=v]\<^sub>v)[bv::=b]\<^sub>c\<^sub>e\<^sub>b = (ce[bv::=b]\<^sub>c\<^sub>e\<^sub>b)[x::=v]\<^sub>v"
+ using assms apply (nominal_induct ce rule: ce.strong_induct, (simp add: subst_vv_subst_bb_commute subst_ceb.simps subst_cv.simps))
+ using assms subst_vv_subst_bb_commute subst_ceb.simps subst_cv.simps
+ by (simp add: subst_v_ce_def)+
+
+lemma subst_cv_subst_bb_commute:
+ fixes bv::bv and b::b and x::x and v::v
+ assumes "atom bv \<sharp> v"
+ shows "c[x::=v]\<^sub>c\<^sub>v[bv::=b]\<^sub>c\<^sub>b = (c[bv::=b]\<^sub>c\<^sub>b)[x::=v]\<^sub>c\<^sub>v"
+ using assms apply (nominal_induct c rule: c.strong_induct )
+ using assms subst_vv_subst_bb_commute subst_ceb.simps subst_cv.simps
+ subst_v_c_def subst_b_c_def apply auto
+ using subst_cev_subst_bb_commute subst_v_ce_def by auto+
+
+lemma subst_b_c_of:
+ "(c_of \<tau> z)[bv::=b]\<^sub>c\<^sub>b = c_of (\<tau>[bv::=b]\<^sub>\<tau>\<^sub>b) z"
+proof(nominal_induct \<tau> avoiding: z rule:\<tau>.strong_induct)
+ case (T_refined_type z' b' c')
+ moreover have "atom bv \<sharp> [ z ]\<^sup>v " using fresh_at_base v.fresh by auto
+ ultimately show ?case using subst_cv_subst_bb_commute[of bv "V_var z" c' z' b] c_of.simps subst_tb.simps by metis
+qed
+
+lemma subst_b_b_of:
+ "(b_of \<tau>)[bv::=b]\<^sub>b\<^sub>b = b_of (\<tau>[bv::=b]\<^sub>\<tau>\<^sub>b)"
+ by(nominal_induct \<tau> rule:\<tau>.strong_induct, simp add: b_of.simps subst_tb.simps )
+
+lemma subst_b_if:
+ "\<lbrace> z : b_of \<tau>[bv::=b]\<^sub>\<tau>\<^sub>b | CE_val (v[bv::=b]\<^sub>v\<^sub>b) == CE_val (V_lit ll) IMP c_of \<tau>[bv::=b]\<^sub>\<tau>\<^sub>b z \<rbrace> = \<lbrace> z : b_of \<tau> | CE_val (v) == CE_val (V_lit ll) IMP c_of \<tau> z \<rbrace>[bv::=b]\<^sub>\<tau>\<^sub>b "
+ unfolding subst_tb.simps subst_cb.simps subst_ceb.simps subst_vb.simps using subst_b_b_of subst_b_c_of by auto
+
+lemma subst_b_top_eq:
+ "\<lbrace> z : B_unit | TRUE \<rbrace>[bv::=b]\<^sub>\<tau>\<^sub>b = \<lbrace> z : B_unit | TRUE \<rbrace>" and "\<lbrace> z : B_bool | TRUE \<rbrace>[bv::=b]\<^sub>\<tau>\<^sub>b = \<lbrace> z : B_bool | TRUE \<rbrace>" and "\<lbrace> z : B_id tid | TRUE \<rbrace> = \<lbrace> z : B_id tid | TRUE \<rbrace>[bv::=b]\<^sub>\<tau>\<^sub>b"
+ unfolding subst_tb.simps subst_bb.simps subst_cb.simps by auto
+
+lemmas subst_b_eq = subst_b_c_of subst_b_b_of subst_b_if subst_b_top_eq
+
+lemma subst_cx_subst_bb_commute[simp]:
+ fixes bv::bv and b::b and x::x and v'::c
+ shows "(v'[x::=V_var y]\<^sub>c\<^sub>v)[bv::=b]\<^sub>c\<^sub>b = (v'[bv::=b]\<^sub>c\<^sub>b)[x::=V_var y]\<^sub>c\<^sub>v"
+ using subst_cv_subst_bb_commute fresh_at_base v.fresh by auto
+
+lemma subst_b_infer_b:
+ fixes l::l and b::b
+ assumes " \<turnstile> l \<Rightarrow> \<tau>" and "\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f b" and "B = {|bv|}"
+ shows "\<turnstile> l \<Rightarrow> (\<tau>[bv::=b]\<^sub>\<tau>\<^sub>b)"
+ using assms infer_l_form3 infer_l_form4 wf_b_subst infer_l_wf subst_tb.simps base_for_lit.simps subst_tb.simps
+ subst_b_base_for_lit subst_cb.simps(6) subst_ceb.simps(1) subst_vb.simps(1) subst_vb.simps(2) type_l_eq
+ by metis
+
+lemma subst_b_subtype:
+ fixes s::s and b'::b
+ assumes "\<Theta> ; {|bv|} ; \<Gamma> \<turnstile> \<tau>1 \<lesssim> \<tau>2" and "\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f b'" and "B = {|bv|}"
+ shows "\<Theta> ; {||} ; \<Gamma>[bv::=b']\<^sub>\<Gamma>\<^sub>b \<turnstile> \<tau>1[bv::=b']\<^sub>\<tau>\<^sub>b \<lesssim> \<tau>2[bv::=b']\<^sub>\<tau>\<^sub>b"
+ using assms proof(nominal_induct "{|bv|}" \<Gamma> \<tau>1 \<tau>2 rule:subtype.strong_induct)
+ case (subtype_baseI x \<Theta> \<Gamma> z c z' c' b)
+
+ hence **: "\<Theta> ; {|bv|} ; (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<Turnstile> c'[z'::=V_var x]\<^sub>c\<^sub>v " using validI subst_defs by metis
+
+ have "\<Theta> ; {||} ; \<Gamma>[bv::=b']\<^sub>\<Gamma>\<^sub>b \<turnstile> \<lbrace> z : b[bv::=b']\<^sub>b\<^sub>b | c[bv::=b']\<^sub>c\<^sub>b \<rbrace> \<lesssim> \<lbrace> z' : b[bv::=b']\<^sub>b\<^sub>b | c'[bv::=b']\<^sub>c\<^sub>b \<rbrace>" proof(rule Typing.subtype_baseI)
+ show "\<Theta> ; {||} ; \<Gamma>[bv::=b']\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b[bv::=b']\<^sub>b\<^sub>b | c[bv::=b']\<^sub>c\<^sub>b \<rbrace> "
+ using subtype_baseI assms wf_b_subst(4) subst_tb.simps subst_defs by metis
+ show "\<Theta> ; {||} ; \<Gamma>[bv::=b']\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f \<lbrace> z' : b[bv::=b']\<^sub>b\<^sub>b | c'[bv::=b']\<^sub>c\<^sub>b \<rbrace> "
+ using subtype_baseI assms wf_b_subst(4) subst_tb.simps by metis
+ show "atom x \<sharp>(\<Theta>, {||}::bv fset, \<Gamma>[bv::=b']\<^sub>\<Gamma>\<^sub>b, z , c[bv::=b']\<^sub>c\<^sub>b , z' , c'[bv::=b']\<^sub>c\<^sub>b )"
+ apply(unfold fresh_prodN,auto simp add: * fresh_prodN fresh_empty_fset)
+ using subst_b_fresh_x * fresh_prodN \<open>atom x \<sharp> c\<close> \<open>atom x \<sharp> c'\<close> subst_defs subtype_baseI by metis+
+ have "\<Theta> ; {||} ; (x, b[bv::=b']\<^sub>b\<^sub>b, c[z::=V_var x]\<^sub>v[bv::=b']\<^sub>c\<^sub>b) #\<^sub>\<Gamma> \<Gamma>[bv::=b']\<^sub>\<Gamma>\<^sub>b \<Turnstile> c'[z'::=V_var x]\<^sub>v[bv::=b']\<^sub>c\<^sub>b"
+ using ** subst_b_valid subst_gb.simps assms subtype_baseI by metis
+ thus "\<Theta> ; {||} ; (x, b[bv::=b']\<^sub>b\<^sub>b, (c[bv::=b']\<^sub>c\<^sub>b)[z::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>[bv::=b']\<^sub>\<Gamma>\<^sub>b \<Turnstile> (c'[bv::=b']\<^sub>c\<^sub>b)[z'::=V_var x]\<^sub>v"
+ using subst_defs subst_cv_subst_bb_commute by (metis subst_cx_subst_bb_commute)
+ qed
+ thus ?case using subtype_baseI subst_tb.simps subst_defs by metis
+qed
+
+lemma b_of_subst_bv:
+ "(b_of \<tau>)[x::=v]\<^sub>b\<^sub>b = b_of (\<tau>[x::=v]\<^sub>\<tau>\<^sub>b)"
+proof -
+ obtain z b c where *:"\<tau> = \<lbrace> z : b | c \<rbrace> \<and> atom z \<sharp> (x,v)" using obtain_fresh_z by metis
+ thus ?thesis using subst_tv.simps * by auto
+qed
+
+lemma subst_b_infer_v:
+ fixes v::v and b::b
+ assumes "\<Theta> ; B ; G \<turnstile> v \<Rightarrow> \<tau>" and "\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f b" and "B = {|bv|}"
+ shows "\<Theta> ; {||} ; G[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> v[bv::=b]\<^sub>v\<^sub>b \<Rightarrow> (\<tau>[bv::=b]\<^sub>\<tau>\<^sub>b)"
+ using assms proof(nominal_induct avoiding: b bv rule: infer_v.strong_induct)
+ case (infer_v_varI \<Theta> \<B> \<Gamma> b' c x z)
+ show ?case unfolding subst_b_simps proof
+ show "\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b " using infer_v_varI wf_b_subst by metis
+ show "Some (b'[bv::=b]\<^sub>b\<^sub>b, c[bv::=b]\<^sub>c\<^sub>b) = lookup \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b x" using subst_b_lookup infer_v_varI by metis
+ show "atom z \<sharp> x" using infer_v_varI by auto
+ show "atom z \<sharp> (\<Theta>, {||}, \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b) " by(fresh_mth add: infer_v_varI subst_b_fresh_x subst_b_\<Gamma>_def fresh_prodN fresh_empty_fset )
+ qed
+next
+ case (infer_v_litI \<Theta> \<B> \<Gamma> l \<tau>)
+ then show ?case using Typing.infer_v_litI subst_b_infer_b
+ using wf_b_subst1(3) by auto
+next
+ case (infer_v_pairI z v1 v2 \<Theta> \<B> \<Gamma> t1 t2)
+ show ?case unfolding subst_b_simps b_of_subst_bv proof
+ show "atom z \<sharp> (v1[bv::=b]\<^sub>v\<^sub>b, v2[bv::=b]\<^sub>v\<^sub>b)" by(fresh_mth add: infer_v_pairI subst_b_fresh_x)
+ show "atom z \<sharp> (\<Theta>, {||}, \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b)" by(fresh_mth add: infer_v_pairI subst_b_fresh_x subst_b_\<Gamma>_def fresh_empty_fset)
+ show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> v1[bv::=b]\<^sub>v\<^sub>b \<Rightarrow> t1[bv::=b]\<^sub>\<tau>\<^sub>b" using infer_v_pairI by auto
+ show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> v2[bv::=b]\<^sub>v\<^sub>b \<Rightarrow> t2[bv::=b]\<^sub>\<tau>\<^sub>b" using infer_v_pairI by auto
+ qed
+next
+ case (infer_v_consI s dclist \<Theta> dc tc \<B> \<Gamma> v tv z)
+ show ?case unfolding subst_b_simps b_of_subst_bv proof
+ show "AF_typedef s dclist \<in> set \<Theta>" using infer_v_consI by auto
+ show "(dc, tc) \<in> set dclist" using infer_v_consI by auto
+ show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> v[bv::=b]\<^sub>v\<^sub>b \<Rightarrow> tv[bv::=b]\<^sub>\<tau>\<^sub>b" using infer_v_consI by auto
+ show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> tv[bv::=b]\<^sub>\<tau>\<^sub>b \<lesssim> tc" proof -
+ have "atom bv \<sharp> tc" using wfTh_lookup_supp_empty fresh_def infer_v_consI infer_v_wf by fast
+ moreover have "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> tv[bv::=b]\<^sub>\<tau>\<^sub>b \<lesssim> tc[bv::=b]\<^sub>\<tau>\<^sub>b"
+ using subst_b_subtype infer_v_consI by simp
+ ultimately show ?thesis using forget_subst subst_b_\<tau>_def by metis
+ qed
+ show "atom z \<sharp> v[bv::=b]\<^sub>v\<^sub>b" using infer_v_consI using subst_b_fresh_x subst_b_v_def by metis
+ show "atom z \<sharp> (\<Theta>, {||}, \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b)" by(fresh_mth add: infer_v_consI subst_b_fresh_x subst_b_\<Gamma>_def fresh_empty_fset)
+ qed
+next
+ case (infer_v_conspI s bv2 dclist2 \<Theta> dc tc \<B> \<Gamma> v tv ba z)
+
+ have "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> V_consp s dc (ba[bv::=b]\<^sub>b\<^sub>b) (v[bv::=b]\<^sub>v\<^sub>b) \<Rightarrow> \<lbrace> z : B_app s (ba[bv::=b]\<^sub>b\<^sub>b) | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ V_consp s dc (ba[bv::=b]\<^sub>b\<^sub>b) (v[bv::=b]\<^sub>v\<^sub>b) ]\<^sup>c\<^sup>e \<rbrace>"
+ proof(rule Typing.infer_v_conspI)
+ show "AF_typedef_poly s bv2 dclist2 \<in> set \<Theta>" using infer_v_conspI by auto
+ show "(dc, tc) \<in> set dclist2" using infer_v_conspI by auto
+ show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> v[bv::=b]\<^sub>v\<^sub>b \<Rightarrow> tv[bv::=b]\<^sub>\<tau>\<^sub>b"
+ using infer_v_conspI subst_tb.simps by metis
+
+ show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> tv[bv::=b]\<^sub>\<tau>\<^sub>b \<lesssim> tc[bv2::=ba[bv::=b]\<^sub>b\<^sub>b]\<^sub>\<tau>\<^sub>b" proof -
+ have "supp tc \<subseteq> { atom bv2 }" using infer_v_conspI wfTh_poly_lookup_supp wfX_wfY by metis
+ moreover have "bv2 \<noteq> bv" using \<open>atom bv2 \<sharp> \<B>\<close> \<open>\<B> = {|bv|} \<close> fresh_at_base fresh_def
+ using fresh_finsert by fastforce
+ ultimately have "atom bv \<sharp> tc" unfolding fresh_def by auto
+ hence "tc[bv2::=ba[bv::=b]\<^sub>b\<^sub>b]\<^sub>\<tau>\<^sub>b = tc[bv2::=ba]\<^sub>\<tau>\<^sub>b[bv::=b]\<^sub>\<tau>\<^sub>b"
+ using subst_tb_commute by metis
+ moreover have "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> tv[bv::=b]\<^sub>\<tau>\<^sub>b \<lesssim> tc[bv2::=ba]\<^sub>\<tau>\<^sub>b[bv::=b]\<^sub>\<tau>\<^sub>b"
+ using infer_v_conspI(7) subst_b_subtype infer_v_conspI by metis
+ ultimately show ?thesis by auto
+ qed
+ show "atom z \<sharp> (\<Theta>, {||}, \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b, v[bv::=b]\<^sub>v\<^sub>b, ba[bv::=b]\<^sub>b\<^sub>b)"
+ apply(unfold fresh_prodN, intro conjI, auto simp add: infer_v_conspI fresh_empty_fset)
+ using \<open>atom z \<sharp> \<Gamma>\<close> fresh_subst_if subst_b_\<Gamma>_def x_fresh_b apply metis
+ using \<open>atom z \<sharp> v\<close> fresh_subst_if subst_b_v_def x_fresh_b by metis
+ show "atom bv2 \<sharp> (\<Theta>, {||}, \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b, v[bv::=b]\<^sub>v\<^sub>b, ba[bv::=b]\<^sub>b\<^sub>b)"
+ apply(unfold fresh_prodN, intro conjI, auto simp add: infer_v_conspI fresh_empty_fset)
+ using \<open>atom bv2 \<sharp> b\<close> \<open>atom bv2 \<sharp> \<Gamma>\<close> fresh_subst_if subst_b_\<Gamma>_def apply metis
+ using \<open>atom bv2 \<sharp> b\<close> \<open>atom bv2 \<sharp> v\<close> fresh_subst_if subst_b_v_def apply metis
+ using \<open>atom bv2 \<sharp> b\<close> \<open>atom bv2 \<sharp> ba\<close> fresh_subst_if subst_b_b_def by metis
+ show "\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f ba[bv::=b]\<^sub>b\<^sub>b "
+ using infer_v_conspI wf_b_subst by metis
+ qed
+ thus ?case using subst_vb.simps subst_tb.simps subst_bb.simps by simp
+
+qed
+
+lemma subst_b_check_v:
+ fixes v::v and b::b
+ assumes "\<Theta> ; B ; G \<turnstile> v \<Leftarrow> \<tau>" and "\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f b" and "B = {|bv|}"
+ shows "\<Theta> ; {||} ; G[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> v[bv::=b]\<^sub>v\<^sub>b \<Leftarrow> (\<tau>[bv::=b]\<^sub>\<tau>\<^sub>b)"
+proof -
+ obtain \<tau>' where "\<Theta> ; B ; G \<turnstile> v \<Rightarrow> \<tau>' \<and> \<Theta> ; B ; G \<turnstile> \<tau>' \<lesssim> \<tau>" using check_v_elims[OF assms(1)] by metis
+ thus ?thesis using subst_b_subtype subst_b_infer_v assms
+ by (metis (no_types) check_v_subtypeI subst_b_infer_v subst_b_subtype)
+qed
+
+lemma subst_vv_subst_vb_switch:
+ shows "(v'[bv::=b']\<^sub>v\<^sub>b)[x::=v[bv::=b']\<^sub>v\<^sub>b]\<^sub>v\<^sub>v = v'[x::=v]\<^sub>v\<^sub>v[bv::=b']\<^sub>v\<^sub>b"
+proof(nominal_induct v' rule:v.strong_induct)
+ case (V_lit x)
+ then show ?case using subst_vv.simps subst_vb.simps by auto
+next
+ case (V_var x)
+ then show ?case using subst_vv.simps subst_vb.simps by auto
+next
+ case (V_pair x1a x2a)
+ then show ?case using subst_vv.simps subst_vb.simps v.fresh by auto
+next
+ case (V_cons x1a x2a x3)
+ then show ?case using subst_vv.simps subst_vb.simps v.fresh by auto
+next
+ case (V_consp x1a x2a x3 x4)
+ then show ?case using subst_vv.simps subst_vb.simps v.fresh pure_fresh
+ by (metis forget_subst subst_b_b_def)
+qed
+
+lemma subst_cev_subst_vb_switch:
+ shows "(ce[bv::=b']\<^sub>c\<^sub>e\<^sub>b)[x::=v[bv::=b']\<^sub>v\<^sub>b]\<^sub>c\<^sub>e\<^sub>v = (ce[x::=v]\<^sub>c\<^sub>e\<^sub>v)[bv::=b']\<^sub>c\<^sub>e\<^sub>b"
+ by(nominal_induct ce rule:ce.strong_induct, auto simp add: subst_vv_subst_vb_switch ce.fresh)
+
+lemma subst_cv_subst_vb_switch:
+ shows "(c[bv::=b']\<^sub>c\<^sub>b)[x::=v[bv::=b']\<^sub>v\<^sub>b]\<^sub>c\<^sub>v = c[x::=v]\<^sub>c\<^sub>v[bv::=b']\<^sub>c\<^sub>b"
+ by(nominal_induct c rule:c.strong_induct, auto simp add: subst_cev_subst_vb_switch c.fresh)
+
+lemma subst_tv_subst_vb_switch:
+ shows "(\<tau>[bv::=b']\<^sub>\<tau>\<^sub>b)[x::=v[bv::=b']\<^sub>v\<^sub>b]\<^sub>\<tau>\<^sub>v = \<tau>[x::=v]\<^sub>\<tau>\<^sub>v[bv::=b']\<^sub>\<tau>\<^sub>b"
+proof(nominal_induct \<tau> avoiding: x v rule:\<tau>.strong_induct)
+ case (T_refined_type z b c )
+ hence ceq: "(c[bv::=b']\<^sub>c\<^sub>b)[x::=v[bv::=b']\<^sub>v\<^sub>b]\<^sub>c\<^sub>v = c[x::=v]\<^sub>c\<^sub>v[bv::=b']\<^sub>c\<^sub>b" using subst_cv_subst_vb_switch by auto
+
+ moreover have "atom z \<sharp> v[bv::=b']\<^sub>v\<^sub>b" using x_fresh_b fresh_subst_if subst_b_v_def T_refined_type by metis
+
+ hence "\<lbrace> z : b | c \<rbrace>[bv::=b']\<^sub>\<tau>\<^sub>b[x::=v[bv::=b']\<^sub>v\<^sub>b]\<^sub>\<tau>\<^sub>v = \<lbrace> z : b[bv::=b']\<^sub>b\<^sub>b | (c[bv::=b']\<^sub>c\<^sub>b)[x::=v[bv::=b']\<^sub>v\<^sub>b]\<^sub>c\<^sub>v \<rbrace>"
+ using subst_tv.simps subst_tb.simps T_refined_type fresh_Pair by metis
+
+ moreover have "\<lbrace> z : b[bv::=b']\<^sub>b\<^sub>b | (c[bv::=b']\<^sub>c\<^sub>b)[x::=v[bv::=b']\<^sub>v\<^sub>b]\<^sub>c\<^sub>v \<rbrace> = \<lbrace> z : b | c[x::=v]\<^sub>c\<^sub>v \<rbrace>[bv::=b']\<^sub>\<tau>\<^sub>b"
+ using subst_tv.simps subst_tb.simps ceq \<tau>.fresh forget_subst[of "bv" b b'] subst_b_b_def T_refined_type by metis
+
+ ultimately show ?case using subst_tv.simps subst_tb.simps ceq T_refined_type by auto
+qed
+
+lemma subst_tb_triple:
+ assumes "atom bv \<sharp> \<tau>'"
+ shows "\<tau>'[bv'::=b'[bv::=b]\<^sub>b\<^sub>b]\<^sub>\<tau>\<^sub>b[x'::=v'[bv::=b]\<^sub>v\<^sub>b]\<^sub>\<tau>\<^sub>v = \<tau>'[bv'::=b']\<^sub>\<tau>\<^sub>b[x'::=v']\<^sub>\<tau>\<^sub>v[bv::=b]\<^sub>\<tau>\<^sub>b"
+proof -
+ have "\<tau>'[bv'::=b'[bv::=b]\<^sub>b\<^sub>b]\<^sub>\<tau>\<^sub>b[x'::=v'[bv::=b]\<^sub>v\<^sub>b]\<^sub>\<tau>\<^sub>v = \<tau>'[bv'::=b']\<^sub>\<tau>\<^sub>b[bv::=b]\<^sub>\<tau>\<^sub>b [x'::=v'[bv::=b]\<^sub>v\<^sub>b]\<^sub>\<tau>\<^sub>v"
+ using subst_tb_commute \<open>atom bv \<sharp> \<tau>'\<close> by auto
+ also have "... = \<tau>'[bv'::=b']\<^sub>\<tau>\<^sub>b [x'::=v']\<^sub>\<tau>\<^sub>v[bv::=b]\<^sub>\<tau>\<^sub>b"
+ using subst_tv_subst_vb_switch by auto
+ finally show ?thesis using fresh_subst_if forget_subst by auto
+qed
+
+lemma subst_b_infer_e:
+ fixes s::s and b::b
+ assumes "\<Theta> ; \<Phi> ; B ; G; D \<turnstile> e \<Rightarrow> \<tau>" and "\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f b" and "B = {|bv|}"
+ shows "\<Theta> ; \<Phi> ; {||} ; G[bv::=b]\<^sub>\<Gamma>\<^sub>b; D[bv::=b]\<^sub>\<Delta>\<^sub>b \<turnstile> (e[bv::=b]\<^sub>e\<^sub>b) \<Rightarrow> (\<tau>[bv::=b]\<^sub>\<tau>\<^sub>b)"
+ using assms proof(nominal_induct avoiding: b rule: infer_e.strong_induct)
+ case (infer_e_valI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v \<tau>)
+ thus ?case using subst_eb.simps infer_e.intros wf_b_subst subst_db.simps wf_b_subst infer_v_wf subst_b_infer_v
+ by (metis forget_subst ms_fresh_all(1) wfV_b_fresh)
+next
+ case (infer_e_plusI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v1 z1 c1 v2 z2 c2 z3)
+
+ show ?case unfolding subst_b_simps subst_eb.simps proof(rule Typing.infer_e_plusI)
+ show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b " using wf_b_subst(10) subst_db.simps infer_e_plusI wfX_wfY
+ by (metis wf_b_subst(15))
+ show "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> " using infer_e_plusI by auto
+ show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> v1[bv::=b]\<^sub>v\<^sub>b \<Rightarrow> \<lbrace> z1 : B_int | c1[bv::=b]\<^sub>c\<^sub>b \<rbrace>" using subst_b_infer_v infer_e_plusI subst_b_simps by force
+ show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> v2[bv::=b]\<^sub>v\<^sub>b \<Rightarrow> \<lbrace> z2 : B_int | c2[bv::=b]\<^sub>c\<^sub>b \<rbrace>" using subst_b_infer_v infer_e_plusI subst_b_simps by force
+ show "atom z3 \<sharp> AE_op Plus (v1[bv::=b]\<^sub>v\<^sub>b) (v2[bv::=b]\<^sub>v\<^sub>b)" using subst_b_simps infer_e_plusI subst_b_fresh_x subst_b_e_def by metis
+ show "atom z3 \<sharp> \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b" using subst_g_b_x_fresh infer_e_plusI by auto
+ qed
+next
+ case (infer_e_leqI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v1 z1 c1 v2 z2 c2 z3)
+ show ?case unfolding subst_b_simps proof(rule Typing.infer_e_leqI)
+ show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b " using wf_b_subst(10) subst_db.simps infer_e_leqI wfX_wfY
+ by (metis wf_b_subst(15))
+ show "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> " using infer_e_leqI by auto
+ show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> v1[bv::=b]\<^sub>v\<^sub>b \<Rightarrow> \<lbrace> z1 : B_int | c1[bv::=b]\<^sub>c\<^sub>b \<rbrace>" using subst_b_infer_v infer_e_leqI subst_b_simps by force
+ show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> v2[bv::=b]\<^sub>v\<^sub>b \<Rightarrow> \<lbrace> z2 : B_int | c2[bv::=b]\<^sub>c\<^sub>b \<rbrace>" using subst_b_infer_v infer_e_leqI subst_b_simps by force
+ show "atom z3 \<sharp> AE_op LEq (v1[bv::=b]\<^sub>v\<^sub>b) (v2[bv::=b]\<^sub>v\<^sub>b)" using subst_b_simps infer_e_leqI subst_b_fresh_x subst_b_e_def by metis
+ show "atom z3 \<sharp> \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b" using subst_g_b_x_fresh infer_e_leqI by auto
+ qed
+next
+ case (infer_e_eqI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v1 z1 bb c1 v2 z2 c2 z3)
+ show ?case unfolding subst_b_simps proof(rule Typing.infer_e_eqI)
+ show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b " using wf_b_subst(10) subst_db.simps infer_e_eqI wfX_wfY
+ by (metis wf_b_subst(15))
+ show "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> " using infer_e_eqI by auto
+ show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> v1[bv::=b]\<^sub>v\<^sub>b \<Rightarrow> \<lbrace> z1 : bb[bv::=b]\<^sub>b\<^sub>b | c1[bv::=b]\<^sub>c\<^sub>b \<rbrace>" using subst_b_infer_v infer_e_eqI subst_b_simps by force
+ show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> v2[bv::=b]\<^sub>v\<^sub>b \<Rightarrow> \<lbrace> z2 : bb[bv::=b]\<^sub>b\<^sub>b | c2[bv::=b]\<^sub>c\<^sub>b \<rbrace>" using subst_b_infer_v infer_e_eqI subst_b_simps by force
+ show "atom z3 \<sharp> AE_op Eq (v1[bv::=b]\<^sub>v\<^sub>b) (v2[bv::=b]\<^sub>v\<^sub>b)" using subst_b_simps infer_e_eqI subst_b_fresh_x subst_b_e_def by metis
+ show "atom z3 \<sharp> \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b" using subst_g_b_x_fresh infer_e_eqI by auto
+ show "bb[bv::=b]\<^sub>b\<^sub>b \<in> {B_bool, B_int, B_unit}" using infer_e_eqI by auto
+ qed
+next
+ case (infer_e_appI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> f x b' c \<tau>' s' v \<tau>)
+ show ?case proof(subst subst_eb.simps, rule Typing.infer_e_appI)
+ show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b" using wf_b_subst(10) subst_db.simps infer_e_appI wfX_wfY by (metis wf_b_subst(15))
+ show "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> " using infer_e_appI by auto
+ show "Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x b' c \<tau>' s'))) = lookup_fun \<Phi> f" using infer_e_appI by auto
+ (*show "\<Theta> ; {||} ; (x, b', c) #\<^sub>\<Gamma> GNil \<turnstile>\<^sub>w\<^sub>f \<tau>'" using infer_e_appI sory*)
+ have "atom bv \<sharp> b'" using \<open>\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> infer_e_appI wfPhi_f_supp fresh_def[of "atom bv" b'] by simp
+ hence "b' = b'[bv::=b]\<^sub>b\<^sub>b" using subst_b_simps
+ using has_subst_b_class.forget_subst subst_b_b_def by force
+ moreover have ceq:"c = c[bv::=b]\<^sub>c\<^sub>b" using subst_b_simps proof -
+ have "supp c \<subseteq> {atom x}" using infer_e_appI wfPhi_f_simple_supp_c[OF _ \<open>\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close>] by simp
+ hence "atom bv \<sharp> c" using
+ fresh_def[of "atom bv" c]
+ using fresh_def fresh_finsert insert_absorb
+ insert_subset ms_fresh_all supp_at_base x_not_in_b_set fresh_prodN by metis
+ thus ?thesis
+ using forget_subst subst_b_c_def fresh_def[of "atom bv" c] by metis
+ qed
+ show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> v[bv::=b]\<^sub>v\<^sub>b \<Leftarrow> \<lbrace> x : b' | c \<rbrace>"
+ using subst_b_check_v subst_tb.simps subst_vb.simps infer_e_appI
+ proof -
+ have "\<Theta> ; {|bv|} ; \<Gamma> \<turnstile> v \<Leftarrow> \<lbrace> x : b' | c \<rbrace>"
+ by (metis \<open>\<B> = {|bv|}\<close> \<open>\<Theta> ; \<B> ; \<Gamma> \<turnstile> v \<Leftarrow> \<lbrace> x : b' | c \<rbrace>\<close>) (* 0.0 ms *)
+ then show ?thesis
+ by (metis (no_types) \<open>\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f b\<close> \<open>b' = b'[bv::=b]\<^sub>b\<^sub>b\<close> subst_b_check_v subst_tb.simps ceq)
+ qed
+ show "atom x \<sharp> (\<Theta>, \<Phi>, {||}::bv fset, \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b, \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b, v[bv::=b]\<^sub>v\<^sub>b, \<tau>[bv::=b]\<^sub>\<tau>\<^sub>b)"
+ apply (fresh_mth add: fresh_prodN subst_g_b_x_fresh infer_e_appI )
+ using subst_b_fresh_x infer_e_appI apply metis+
+ done
+ have "supp \<tau>' \<subseteq> { atom x }" using wfPhi_f_simple_supp_t infer_e_appI by auto
+ hence "atom bv \<sharp> \<tau>'" using fresh_def fresh_at_base by force
+ then show "\<tau>'[x::=v[bv::=b]\<^sub>v\<^sub>b]\<^sub>v = \<tau>[bv::=b]\<^sub>\<tau>\<^sub>b" using infer_e_appI
+ forget_subst subst_b_\<tau>_def subst_tv_subst_vb_switch subst_defs by metis
+ qed
+next
+ case (infer_e_appPI \<Theta>' \<B> \<Gamma>' \<Delta> \<Phi>' b' f' bv' x' b1 c \<tau>' s' v' \<tau>1)
+
+ have beq: "b1[bv'::=b']\<^sub>b\<^sub>b[bv::=b]\<^sub>b\<^sub>b = b1[bv'::=b'[bv::=b]\<^sub>b\<^sub>b]\<^sub>b\<^sub>b"
+ proof -
+ have "supp b1 \<subseteq> { atom bv' }" using wfPhi_f_poly_supp_b infer_e_appPI
+ using supp_at_base by blast
+ moreover have "bv \<noteq> bv'" using infer_e_appPI fresh_def supp_at_base
+ by (simp add: fresh_def supp_at_base)
+ ultimately have "atom bv \<sharp> b1" using fresh_def fresh_at_base by force
+ thus ?thesis by simp
+ qed
+
+ have ceq: "(c[bv'::=b']\<^sub>c\<^sub>b)[bv::=b]\<^sub>c\<^sub>b = c[bv'::=b'[bv::=b]\<^sub>b\<^sub>b]\<^sub>c\<^sub>b" proof -
+ have "supp c \<subseteq> { atom bv', atom x' }" using wfPhi_f_poly_supp_c infer_e_appPI
+ using supp_at_base by blast
+ moreover have "bv \<noteq> bv'" using infer_e_appPI fresh_def supp_at_base
+ by (simp add: fresh_def supp_at_base)
+ moreover have "atom x' \<noteq> atom bv" by auto
+ ultimately have "atom bv \<sharp> c" using fresh_def[of "atom bv" c] fresh_at_base by auto
+ thus ?thesis by simp
+ qed
+
+ show ?case proof(subst subst_eb.simps, rule Typing.infer_e_appPI)
+ show "\<Theta>' ; {||} ; \<Gamma>'[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b " using wf_b_subst subst_db.simps infer_e_appPI wfX_wfY by metis
+ show "\<Theta>' \<turnstile>\<^sub>w\<^sub>f \<Phi>' " using infer_e_appPI by auto
+ show "Some (AF_fundef f' (AF_fun_typ_some bv' (AF_fun_typ x' b1 c \<tau>' s'))) = lookup_fun \<Phi>' f'" using infer_e_appPI by auto
+ thus "\<Theta>' ; {||} ; \<Gamma>'[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> v'[bv::=b]\<^sub>v\<^sub>b \<Leftarrow> \<lbrace> x' : b1[bv'::=b'[bv::=b]\<^sub>b\<^sub>b]\<^sub>b | c[bv'::=b'[bv::=b]\<^sub>b\<^sub>b]\<^sub>b \<rbrace>"
+ using subst_b_check_v subst_tb.simps subst_b_simps infer_e_appPI
+ proof -
+ have "\<Theta>' ; {||} ; \<Gamma>'[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> v'[bv::=b]\<^sub>v\<^sub>b \<Leftarrow> \<lbrace> x' : b1[bv'::=b']\<^sub>b[bv::=b]\<^sub>b\<^sub>b | (c[bv'::=b']\<^sub>b)[bv::=b]\<^sub>c\<^sub>b \<rbrace>"
+ using infer_e_appPI subst_b_check_v subst_tb.simps by metis
+ thus ?thesis using beq ceq subst_defs by metis
+ qed
+ show "atom x' \<sharp> \<Gamma>'[bv::=b]\<^sub>\<Gamma>\<^sub>b" using subst_g_b_x_fresh infer_e_appPI by auto
+ show "\<tau>'[bv'::=b'[bv::=b]\<^sub>b\<^sub>b]\<^sub>b[x'::=v'[bv::=b]\<^sub>v\<^sub>b]\<^sub>v = \<tau>1[bv::=b]\<^sub>\<tau>\<^sub>b" proof - (* \<tau>'[bv'::=b']\<^sub>\<tau>\<^sub>b[x'::=v']\<^sub>\<tau>\<^sub>v = \<tau>1 *)
+
+ have "supp \<tau>' \<subseteq> { atom x', atom bv' }" using wfPhi_f_poly_supp_t infer_e_appPI by auto
+ moreover hence "bv \<noteq> bv'" using infer_e_appPI fresh_def supp_at_base
+ by (simp add: fresh_def supp_at_base)
+ ultimately have "atom bv \<sharp> \<tau>'" using fresh_def by force
+ hence "\<tau>'[bv'::=b'[bv::=b]\<^sub>b\<^sub>b]\<^sub>b[x'::=v'[bv::=b]\<^sub>v\<^sub>b]\<^sub>v = \<tau>'[bv'::=b']\<^sub>b[x'::=v']\<^sub>v[bv::=b]\<^sub>\<tau>\<^sub>b" using subst_tb_triple subst_defs by auto
+ thus ?thesis using infer_e_appPI by metis
+ qed
+ show "atom bv' \<sharp> (\<Theta>', \<Phi>', {||}, \<Gamma>'[bv::=b]\<^sub>\<Gamma>\<^sub>b, \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b, b'[bv::=b]\<^sub>b\<^sub>b, v'[bv::=b]\<^sub>v\<^sub>b, \<tau>1[bv::=b]\<^sub>\<tau>\<^sub>b)"
+ unfolding fresh_prodN apply( auto simp add: infer_e_appPI fresh_empty_fset)
+ using fresh_subst_if subst_b_\<Gamma>_def subst_b_\<Delta>_def subst_b_b_def subst_b_v_def subst_b_\<tau>_def infer_e_appPI by metis+
+ show "\<Theta>' ; {||} \<turnstile>\<^sub>w\<^sub>f b'[bv::=b]\<^sub>b\<^sub>b " using infer_e_appPI wf_b_subst by simp
+ qed
+next
+ case (infer_e_fstI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v z' b1 b2 c z)
+ show ?case unfolding subst_b_simps proof(rule Typing.infer_e_fstI)
+ show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b " using wf_b_subst(10) subst_db.simps infer_e_fstI wfX_wfY
+ by (metis wf_b_subst(15))
+ show "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> " using infer_e_fstI by auto
+ show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> v[bv::=b]\<^sub>v\<^sub>b \<Rightarrow> \<lbrace> z' : B_pair b1[bv::=b]\<^sub>b\<^sub>b b2[bv::=b]\<^sub>b\<^sub>b | c[bv::=b]\<^sub>c\<^sub>b \<rbrace>"
+ using subst_b_infer_v subst_tb.simps subst_b_simps infer_e_fstI by force
+ show "atom z \<sharp> AE_fst (v[bv::=b]\<^sub>v\<^sub>b)" using infer_e_fstI subst_b_fresh_x subst_b_v_def e.fresh by metis
+ show "atom z \<sharp> \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b" using subst_g_b_x_fresh infer_e_fstI by auto
+ qed
+next
+ case (infer_e_sndI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v z' b1 b2 c z)
+ show ?case unfolding subst_b_simps proof(rule Typing.infer_e_sndI)
+ show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b " using wf_b_subst(10) subst_db.simps infer_e_sndI wfX_wfY
+ by (metis wf_b_subst(15))
+ show "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> " using infer_e_sndI by auto
+ show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> v[bv::=b]\<^sub>v\<^sub>b \<Rightarrow> \<lbrace> z' : B_pair b1[bv::=b]\<^sub>b\<^sub>b b2[bv::=b]\<^sub>b\<^sub>b | c[bv::=b]\<^sub>c\<^sub>b \<rbrace>"
+ using subst_b_infer_v subst_tb.simps subst_b_simps infer_e_sndI by force
+ show "atom z \<sharp> AE_snd (v[bv::=b]\<^sub>v\<^sub>b)" using infer_e_sndI subst_b_fresh_x subst_b_v_def e.fresh by metis
+ show "atom z \<sharp> \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b" using subst_g_b_x_fresh infer_e_sndI by auto
+ qed
+next
+ case (infer_e_lenI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v z' c z)
+ show ?case unfolding subst_b_simps proof(rule Typing.infer_e_lenI)
+ show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b " using wf_b_subst(10) subst_db.simps infer_e_lenI wfX_wfY
+ by (metis wf_b_subst(15))
+ show "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> " using infer_e_lenI by auto
+ show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> v[bv::=b]\<^sub>v\<^sub>b \<Rightarrow> \<lbrace> z' : B_bitvec | c[bv::=b]\<^sub>c\<^sub>b \<rbrace>"
+ using subst_b_infer_v subst_tb.simps subst_b_simps infer_e_lenI by force
+ show "atom z \<sharp> AE_len (v[bv::=b]\<^sub>v\<^sub>b)" using infer_e_lenI subst_b_fresh_x subst_b_v_def e.fresh by metis
+ show "atom z \<sharp> \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b" using subst_g_b_x_fresh infer_e_lenI by auto
+ qed
+next
+ case (infer_e_mvarI \<Theta> \<B> \<Gamma> \<Phi> \<Delta> u \<tau>)
+ show ?case proof(subst subst subst_eb.simps, rule Typing.infer_e_mvarI)
+ show "\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b " using infer_e_mvarI wf_b_subst by auto
+ show "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> " using infer_e_mvarI by auto
+ show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b " using infer_e_mvarI using wf_b_subst(10) subst_db.simps infer_e_sndI wfX_wfY
+ by (metis wf_b_subst(15))
+ show "(u, \<tau>[bv::=b]\<^sub>\<tau>\<^sub>b) \<in> setD \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b" using infer_e_mvarI subst_db.simps set_insert
+ subst_d_b_member by simp
+ qed
+next
+ case (infer_e_concatI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v1 z1 c1 v2 z2 c2 z3)
+ show ?case unfolding subst_b_simps proof(rule Typing.infer_e_concatI)
+ show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b " using wf_b_subst(10) subst_db.simps infer_e_concatI wfX_wfY
+ by (metis wf_b_subst(15))
+ show "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> " using infer_e_concatI by auto
+ show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> v1[bv::=b]\<^sub>v\<^sub>b \<Rightarrow> \<lbrace> z1 : B_bitvec | c1[bv::=b]\<^sub>c\<^sub>b \<rbrace>"
+ using subst_b_infer_v subst_tb.simps subst_b_simps infer_e_concatI by force
+ show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> v2[bv::=b]\<^sub>v\<^sub>b \<Rightarrow> \<lbrace> z2 : B_bitvec | c2[bv::=b]\<^sub>c\<^sub>b \<rbrace>"
+ using subst_b_infer_v subst_tb.simps subst_b_simps infer_e_concatI by force
+ show "atom z3 \<sharp> AE_concat (v1[bv::=b]\<^sub>v\<^sub>b) (v2[bv::=b]\<^sub>v\<^sub>b)" using infer_e_concatI subst_b_fresh_x subst_b_v_def e.fresh by metis
+ show "atom z3 \<sharp> \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b" using subst_g_b_x_fresh infer_e_concatI by auto
+ qed
+next
+ case (infer_e_splitI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v1 z1 c1 v2 z2 z3)
+ show ?case unfolding subst_b_simps proof(rule Typing.infer_e_splitI)
+ show \<open> \<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b \<close> using wf_b_subst(10) subst_db.simps infer_e_splitI wfX_wfY
+ by (metis wf_b_subst(15))
+ show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using infer_e_splitI by auto
+ show \<open> \<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> v1[bv::=b]\<^sub>v\<^sub>b \<Rightarrow> \<lbrace> z1 : B_bitvec | c1[bv::=b]\<^sub>c\<^sub>b \<rbrace>\<close>
+ using subst_b_infer_v subst_tb.simps subst_b_simps infer_e_splitI by force
+ show \<open>\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> v2[bv::=b]\<^sub>v\<^sub>b \<Leftarrow> \<lbrace> z2 : B_int | [ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ z2 ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e AND
+ [ leq [ [ z2 ]\<^sup>v ]\<^sup>c\<^sup>e [| [ v1[bv::=b]\<^sub>v\<^sub>b ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrace>\<close>
+ using subst_b_check_v subst_tb.simps subst_b_simps infer_e_splitI
+ proof -
+ have "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> v2[bv::=b]\<^sub>v\<^sub>b \<Leftarrow> \<lbrace> z2 : B_int | [ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ z2 ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e AND [ leq [ [ z2 ]\<^sup>v ]\<^sup>c\<^sup>e [| [ v1 ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrace>[bv::=b]\<^sub>\<tau>\<^sub>b"
+ using infer_e_splitI.hyps(7) infer_e_splitI.prems(1) infer_e_splitI.prems(2) subst_b_check_v by presburger (* 0.0 ms *)
+ then show ?thesis
+ by simp (* 0.0 ms *)
+ qed
+ show \<open>atom z1 \<sharp> AE_split (v1[bv::=b]\<^sub>v\<^sub>b) (v2[bv::=b]\<^sub>v\<^sub>b)\<close> using infer_e_splitI subst_b_fresh_x subst_b_v_def e.fresh by metis
+ show \<open>atom z1 \<sharp> \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b\<close> using subst_g_b_x_fresh infer_e_splitI by auto
+
+ show \<open>atom z2 \<sharp> AE_split (v1[bv::=b]\<^sub>v\<^sub>b) (v2[bv::=b]\<^sub>v\<^sub>b)\<close> using infer_e_splitI subst_b_fresh_x subst_b_v_def e.fresh by metis
+
+ show \<open>atom z2 \<sharp> \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b\<close> using subst_g_b_x_fresh infer_e_splitI by auto
+ show \<open>atom z3 \<sharp> AE_split (v1[bv::=b]\<^sub>v\<^sub>b) (v2[bv::=b]\<^sub>v\<^sub>b)\<close> using infer_e_splitI subst_b_fresh_x subst_b_v_def e.fresh by metis
+ show \<open>atom z3 \<sharp> \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b\<close> using subst_g_b_x_fresh infer_e_splitI by auto
+ qed
+qed
+
+text \<open>This is needed for preservation. When we apply a function "f [b] v" we need to
+substitute into the body of the function f replacing type-variable with b\<close>
+
+lemma subst_b_c_of_forget:
+ assumes "atom bv \<sharp> const"
+ shows "(c_of const x)[bv::=b]\<^sub>c\<^sub>b = c_of const x"
+ using assms proof(nominal_induct const avoiding: x rule:\<tau>.strong_induct)
+ case (T_refined_type x' b' c')
+ hence "c_of \<lbrace> x' : b' | c' \<rbrace> x = c'[x'::=V_var x]\<^sub>c\<^sub>v" using c_of.simps by metis
+ moreover have "atom bv \<sharp> c'[x'::=V_var x]\<^sub>c\<^sub>v" proof -
+ have "atom bv \<sharp> c'" using T_refined_type \<tau>.fresh by simp
+ moreover have "atom bv \<sharp> V_var x" using v.fresh by simp
+ ultimately show ?thesis
+ using T_refined_type \<tau>.fresh subst_b_c_def fresh_subst_if
+ \<tau>_fresh_c fresh_subst_cv_if has_subst_b_class.subst_b_fresh_x ms_fresh_all(37) ms_fresh_all assms by metis
+ qed
+ ultimately show ?case using forget_subst subst_b_c_def by metis
+qed
+
+lemma subst_b_check_s:
+ fixes s::s and b::b and cs::branch_s and css::branch_list and v::v and \<tau>::\<tau>
+ assumes "\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f b" and "B = {|bv|}" (*and "D = []"*)
+ shows "\<Theta> ; \<Phi> ; B ; G; D \<turnstile> s \<Leftarrow> \<tau> \<Longrightarrow> \<Theta> ; \<Phi> ; {||} ; G[bv::=b]\<^sub>\<Gamma>\<^sub>b; D[bv::=b]\<^sub>\<Delta>\<^sub>b \<turnstile> (s[bv::=b]\<^sub>s\<^sub>b) \<Leftarrow> (\<tau>[bv::=b]\<^sub>\<tau>\<^sub>b)" and
+ "\<Theta> ; \<Phi> ; B ; G; D ; tid ; cons ; const ; v \<turnstile> cs \<Leftarrow> \<tau> \<Longrightarrow> \<Theta> ; \<Phi> ; {||} ; G[bv::=b]\<^sub>\<Gamma>\<^sub>b; D[bv::=b]\<^sub>\<Delta>\<^sub>b ; tid ; cons ; const ; v[bv::=b]\<^sub>v\<^sub>b \<turnstile> (subst_branchb cs bv b) \<Leftarrow> (\<tau>[bv::=b]\<^sub>\<tau>\<^sub>b)" and
+ "\<Theta> ; \<Phi> ; B ; G; D ; tid ; dclist ; v \<turnstile> css \<Leftarrow> \<tau> \<Longrightarrow> \<Theta> ; \<Phi> ; {||} ; G[bv::=b]\<^sub>\<Gamma>\<^sub>b; D[bv::=b]\<^sub>\<Delta>\<^sub>b ; tid ; dclist ; v[bv::=b]\<^sub>v\<^sub>b \<turnstile> (subst_branchlb css bv b ) \<Leftarrow> (\<tau>[bv::=b]\<^sub>\<tau>\<^sub>b)"
+ using assms proof(induct rule: check_s_check_branch_s_check_branch_list.inducts)
+ note facts = wfD_emptyI wfX_wfY wf_b_subst subst_b_subtype subst_b_infer_v
+ case (check_valI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v \<tau>' \<tau>)
+ show ?case
+ apply(subst subst_sb.simps, rule Typing.check_valI)
+ using facts check_valI apply metis
+ using check_valI subst_b_infer_v wf_b_subst subst_b_subtype apply blast
+ using check_valI subst_b_infer_v wf_b_subst subst_b_subtype apply blast
+ using check_valI subst_b_infer_v wf_b_subst subst_b_subtype by metis
+next
+ case (check_letI x \<Theta> \<Phi> \<B> \<Gamma> \<Delta> e \<tau> z s b' c)
+ show ?case proof(subst subst_sb.simps, rule Typing.check_letI)
+
+ show "atom x \<sharp>(\<Theta>, \<Phi>, {||}, \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b, \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b, e[bv::=b]\<^sub>e\<^sub>b, \<tau>[bv::=b]\<^sub>\<tau>\<^sub>b)" (*using check_letI subst_b_\<tau>_def subst_b_\<Gamma>_def subst_b_fresh_x fresh_prod4 subst_b_c_def sory*)
+ apply(unfold fresh_prodN,auto)
+ apply(simp add: check_letI fresh_empty_fset)+
+ apply(metis * subst_b_fresh_x check_letI fresh_prodN)+ done
+ show "atom z \<sharp> (x, \<Theta>, \<Phi>, {||}, \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b, \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b, e[bv::=b]\<^sub>e\<^sub>b, \<tau>[bv::=b]\<^sub>\<tau>\<^sub>b, s[bv::=b]\<^sub>s\<^sub>b)"
+ apply(unfold fresh_prodN,auto)
+ apply(simp add: check_letI fresh_empty_fset)+
+ apply(metis * subst_b_fresh_x check_letI fresh_prodN)+ done
+ show "\<Theta> ; \<Phi> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b ; \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b \<turnstile> e[bv::=b]\<^sub>e\<^sub>b \<Rightarrow> \<lbrace> z : b'[bv::=b]\<^sub>b\<^sub>b | c[bv::=b]\<^sub>c\<^sub>b \<rbrace>"
+ using check_letI subst_b_infer_e subst_tb.simps by metis
+ have "c[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v[bv::=b]\<^sub>c\<^sub>b = (c[bv::=b]\<^sub>c\<^sub>b)[z::=V_var x]\<^sub>c\<^sub>v"
+ using subst_cv_subst_bb_commute[of bv "V_var x" c z b] fresh_at_base by simp
+ thus "\<Theta> ; \<Phi> ; {||} ; ((x, b'[bv::=b]\<^sub>b\<^sub>b , (c[bv::=b]\<^sub>c\<^sub>b)[z::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b) ; \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b \<turnstile> s[bv::=b]\<^sub>s\<^sub>b \<Leftarrow> \<tau>[bv::=b]\<^sub>\<tau>\<^sub>b"
+ using check_letI subst_gb.simps subst_defs by metis
+ qed
+next
+ case (check_assertI x \<Theta> \<Phi> \<B> \<Gamma> \<Delta> c \<tau> s)
+ show ?case proof(subst subst_sb.simps, rule Typing.check_assertI)
+ show "atom x \<sharp> (\<Theta>, \<Phi>, {||}, \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b, \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b, c[bv::=b]\<^sub>c\<^sub>b, \<tau>[bv::=b]\<^sub>\<tau>\<^sub>b, s[bv::=b]\<^sub>s\<^sub>b)"
+ apply(unfold fresh_prodN,auto)
+ apply(simp add: check_assertI fresh_empty_fset)+
+ apply(metis * subst_b_fresh_x check_assertI fresh_prodN)+ done
+
+ have "\<Theta> ; \<Phi> ; {||} ; ((x, B_bool, c) #\<^sub>\<Gamma> \<Gamma>)[bv::=b]\<^sub>\<Gamma>\<^sub>b ; \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b \<turnstile> s[bv::=b]\<^sub>s\<^sub>b \<Leftarrow> \<tau>[bv::=b]\<^sub>\<tau>\<^sub>b" using check_assertI
+ by metis
+ thus "\<Theta> ; \<Phi> ; {||} ; (x, B_bool, c[bv::=b]\<^sub>c\<^sub>b) #\<^sub>\<Gamma> \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b ; \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b \<turnstile> s[bv::=b]\<^sub>s\<^sub>b \<Leftarrow> \<tau>[bv::=b]\<^sub>\<tau>\<^sub>b" using subst_gb.simps by auto
+ show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<Turnstile> c[bv::=b]\<^sub>c\<^sub>b " using subst_b_valid check_assertI by simp
+ show " \<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b " using wf_b_subst2(6) check_assertI by simp
+ qed
+next
+ case (check_branch_list_consI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist v cs \<tau> css)
+ then show ?case unfolding subst_branchlb.simps using Typing.check_branch_list_consI by simp
+next
+ case (check_branch_list_finalI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist v cs \<tau>)
+ then show ?case unfolding subst_branchlb.simps using Typing.check_branch_list_finalI by simp
+next
+ case (check_branch_s_branchI \<Theta> \<B> \<Gamma> \<Delta> \<tau> const x \<Phi> tid cons v s)
+
+ show ?case unfolding subst_b_simps proof(rule Typing.check_branch_s_branchI)
+ show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b " using check_branch_s_branchI wf_b_subst subst_db.simps by metis
+ show "\<turnstile>\<^sub>w\<^sub>f \<Theta> " using check_branch_s_branchI by auto
+ show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f \<tau>[bv::=b]\<^sub>\<tau>\<^sub>b " using check_branch_s_branchI wf_b_subst by metis
+
+ show "atom x \<sharp>(\<Theta>, \<Phi>, {||}, \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b, \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b, tid, cons , const, v[bv::=b]\<^sub>v\<^sub>b, \<tau>[bv::=b]\<^sub>\<tau>\<^sub>b)"
+ apply(unfold fresh_prodN,auto)
+ apply(simp add: check_branch_s_branchI fresh_empty_fset)+
+ apply(metis * subst_b_fresh_x check_branch_s_branchI fresh_prodN)+
+ done
+ show wft:"\<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f const" using check_branch_s_branchI by auto
+ hence "(b_of const) = (b_of const)[bv::=b]\<^sub>b\<^sub>b"
+ using wfT_nil_supp fresh_def[of "atom bv" ] forget_subst subst_b_b_def \<tau>.supp
+ bot.extremum_uniqueI ex_in_conv fresh_def supp_empty_fset
+ by (metis b_of_supp)
+ moreover have "(c_of const x)[bv::=b]\<^sub>c\<^sub>b = c_of const x"
+ using wft wfT_nil_supp fresh_def[of "atom bv" ] forget_subst subst_b_c_def \<tau>.supp
+ bot.extremum_uniqueI ex_in_conv fresh_def supp_empty_fset subst_b_c_of_forget by metis
+ ultimately show "\<Theta> ; \<Phi> ; {||} ; (x, b_of const, CE_val (v[bv::=b]\<^sub>v\<^sub>b) == CE_val(V_cons tid cons (V_var x)) AND c_of const x) #\<^sub>\<Gamma> \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b ; \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b \<turnstile> s[bv::=b]\<^sub>s\<^sub>b \<Leftarrow> \<tau>[bv::=b]\<^sub>\<tau>\<^sub>b"
+ using check_branch_s_branchI subst_gb.simps by auto
+ qed
+next
+ case (check_ifI z \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v s1 s2 \<tau>)
+ show ?case unfolding subst_b_simps proof(rule Typing.check_ifI)
+ show \<open>atom z \<sharp> (\<Theta>, \<Phi>, {||}, \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b, \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b, v[bv::=b]\<^sub>v\<^sub>b, s1[bv::=b]\<^sub>s\<^sub>b, s2[bv::=b]\<^sub>s\<^sub>b, \<tau>[bv::=b]\<^sub>\<tau>\<^sub>b)\<close>
+ by(unfold fresh_prodN,auto, auto simp add: check_ifI fresh_empty_fset subst_b_fresh_x )
+ have "\<lbrace> z : B_bool | TRUE \<rbrace>[bv::=b]\<^sub>\<tau>\<^sub>b = \<lbrace> z : B_bool | TRUE \<rbrace>" by auto
+ thus \<open>\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> v[bv::=b]\<^sub>v\<^sub>b \<Leftarrow> \<lbrace> z : B_bool | TRUE \<rbrace>\<close> using check_ifI subst_b_check_v by metis
+ show \<open> \<Theta> ; \<Phi> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b ; \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b \<turnstile> s1[bv::=b]\<^sub>s\<^sub>b \<Leftarrow> \<lbrace> z : b_of \<tau>[bv::=b]\<^sub>\<tau>\<^sub>b | CE_val (v[bv::=b]\<^sub>v\<^sub>b) == CE_val (V_lit L_true) IMP c_of \<tau>[bv::=b]\<^sub>\<tau>\<^sub>b z \<rbrace>\<close>
+ using subst_b_if check_ifI by metis
+ show \<open> \<Theta> ; \<Phi> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b ; \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b \<turnstile> s2[bv::=b]\<^sub>s\<^sub>b \<Leftarrow> \<lbrace> z : b_of \<tau>[bv::=b]\<^sub>\<tau>\<^sub>b | CE_val (v[bv::=b]\<^sub>v\<^sub>b) == CE_val (V_lit L_false) IMP c_of \<tau>[bv::=b]\<^sub>\<tau>\<^sub>b z \<rbrace>\<close>
+ using subst_b_if check_ifI by metis
+ qed
+next
+ case (check_let2I x \<Theta> \<Phi> \<B> G \<Delta> t s1 \<tau> s2 )
+ show ?case unfolding subst_b_simps proof (rule Typing.check_let2I)
+ have "atom x \<sharp> b" using x_fresh_b by auto
+ show \<open>atom x \<sharp> (\<Theta>, \<Phi>, {||}, G[bv::=b]\<^sub>\<Gamma>\<^sub>b, \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b, t[bv::=b]\<^sub>\<tau>\<^sub>b, s1[bv::=b]\<^sub>s\<^sub>b, \<tau>[bv::=b]\<^sub>\<tau>\<^sub>b)\<close>
+ apply(unfold fresh_prodN, auto, auto simp add: check_let2I fresh_prodN fresh_empty_fset)
+ apply(metis subst_b_fresh_x check_let2I fresh_prodN)+
+ done
+
+ show \<open> \<Theta> ; \<Phi> ; {||} ; G[bv::=b]\<^sub>\<Gamma>\<^sub>b ; \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b \<turnstile> s1[bv::=b]\<^sub>s\<^sub>b \<Leftarrow> t[bv::=b]\<^sub>\<tau>\<^sub>b \<close> using check_let2I subst_tb.simps by auto
+ show \<open> \<Theta> ; \<Phi> ; {||} ; (x, b_of t[bv::=b]\<^sub>\<tau>\<^sub>b, c_of t[bv::=b]\<^sub>\<tau>\<^sub>b x) #\<^sub>\<Gamma> G[bv::=b]\<^sub>\<Gamma>\<^sub>b ; \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b \<turnstile> s2[bv::=b]\<^sub>s\<^sub>b \<Leftarrow> \<tau>[bv::=b]\<^sub>\<tau>\<^sub>b\<close>
+ using check_let2I subst_tb.simps subst_gb.simps b_of.simps subst_b_c_of subst_b_b_of by auto
+ qed
+next
+ case (check_varI u \<Theta> \<Phi> \<B> \<Gamma> \<Delta> \<tau>' v \<tau> s)
+ show ?case unfolding subst_b_simps proof(rule Typing.check_varI)
+ show "atom u \<sharp> (\<Theta>, \<Phi>, {||}, \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b, \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b, \<tau>'[bv::=b]\<^sub>\<tau>\<^sub>b, v[bv::=b]\<^sub>v\<^sub>b, \<tau>[bv::=b]\<^sub>\<tau>\<^sub>b) "
+ by(unfold fresh_prodN,auto simp add: check_varI fresh_empty_fset subst_b_fresh_u )
+ show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> v[bv::=b]\<^sub>v\<^sub>b \<Leftarrow> \<tau>'[bv::=b]\<^sub>\<tau>\<^sub>b" using check_varI subst_b_check_v by auto
+ show "\<Theta> ; \<Phi> ; {||} ; (subst_gb \<Gamma> bv b) ; (u, (\<tau>'[bv::=b]\<^sub>\<tau>\<^sub>b)) #\<^sub>\<Delta> (subst_db \<Delta> bv b) \<turnstile> (s[bv::=b]\<^sub>s\<^sub>b) \<Leftarrow> (\<tau>[bv::=b]\<^sub>\<tau>\<^sub>b)" using check_varI by auto
+ qed
+next
+ case (check_assignI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> u \<tau> v z \<tau>')
+ show ?case unfolding subst_b_simps proof( rule Typing.check_assignI)
+ show "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> " using check_assignI by auto
+ show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b " using wf_b_subst check_assignI by auto
+ show "(u, \<tau>[bv::=b]\<^sub>\<tau>\<^sub>b) \<in> setD \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b" using check_assignI subst_d_b_member by simp
+ show " \<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> v[bv::=b]\<^sub>v\<^sub>b \<Leftarrow> \<tau>[bv::=b]\<^sub>\<tau>\<^sub>b" using check_assignI subst_b_check_v by auto
+ show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> \<lbrace> z : B_unit | TRUE \<rbrace> \<lesssim> \<tau>'[bv::=b]\<^sub>\<tau>\<^sub>b" using check_assignI subst_b_subtype subst_b_simps subst_tb.simps by fastforce
+ qed
+next
+ case (check_whileI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> s1 z s2 \<tau>')
+ show ?case unfolding subst_b_simps proof(rule Typing.check_whileI)
+ show "\<Theta> ; \<Phi> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b ; \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b \<turnstile> s1[bv::=b]\<^sub>s\<^sub>b \<Leftarrow> \<lbrace> z : B_bool | TRUE \<rbrace>" using check_whileI by auto
+ show "\<Theta> ; \<Phi> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b ; \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b \<turnstile> s2[bv::=b]\<^sub>s\<^sub>b \<Leftarrow> \<lbrace> z : B_unit | TRUE \<rbrace>" using check_whileI by auto
+ show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> \<lbrace> z : B_unit | TRUE \<rbrace> \<lesssim> \<tau>'[bv::=b]\<^sub>\<tau>\<^sub>b" using subst_b_subtype check_whileI by fastforce
+ qed
+next
+ case (check_seqI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> s1 z s2 \<tau>)
+ then show ?case unfolding subst_sb.simps using check_seqI Typing.check_seqI subst_b_eq by metis
+next
+ case (check_caseI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist v cs \<tau> z)
+ show ?case unfolding subst_b_simps proof(rule Typing.check_caseI)
+ show \<open> \<Theta> ; \<Phi> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b ; \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b ; tid ; dclist ; v[bv::=b]\<^sub>v\<^sub>b \<turnstile> subst_branchlb cs bv b \<Leftarrow> \<tau>[bv::=b]\<^sub>\<tau>\<^sub>b\<close> using check_caseI by auto
+ show \<open>AF_typedef tid dclist \<in> set \<Theta>\<close> using check_caseI by auto
+ show \<open>\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> v[bv::=b]\<^sub>v\<^sub>b \<Leftarrow> \<lbrace> z : B_id tid | TRUE \<rbrace>\<close> using check_caseI subst_b_check_v subst_b_simps subst_tb.simps subst_b_simps
+ proof -
+ have "\<lbrace> z : B_id tid | TRUE \<rbrace> = \<lbrace> z : B_id tid | TRUE \<rbrace>[bv::=b]\<^sub>\<tau>\<^sub>b" using subst_b_eq by auto
+ then show ?thesis
+ by (metis (no_types) check_caseI.hyps(4) check_caseI.prems(1) check_caseI.prems(2) subst_b_check_v) (* 31 ms *)
+ qed
+ show \<open> \<turnstile>\<^sub>w\<^sub>f \<Theta> \<close> using check_caseI by auto
+ qed
+qed
+
+end
diff --git a/thys/MiniSail/ContextSubtypingL.thy b/thys/MiniSail/ContextSubtypingL.thy
--- a/thys/MiniSail/ContextSubtypingL.thy
+++ b/thys/MiniSail/ContextSubtypingL.thy
@@ -1,1068 +1,1068 @@
-(*<*)
-theory ContextSubtypingL
- imports TypingL "HOL-Eisbach.Eisbach_Tools" SubstMethods
-begin
- (*>*)
-
-declare freshers[simp del]
-
-chapter \<open>Context Subtyping Lemmas\<close>
-
-text \<open>Lemmas allowing us to replace the type of a variable in the context with a subtype
-and have the judgement remain valid. Also known as narrowing.\<close>
-
-section \<open>Replace or exchange type of variable in a context\<close>
-
-text \<open> Because the G-context is extended by the statements like let, we will need a generalised
-substitution lemma for statements.
-For this we setup a function that replaces in G (rig) for a particular x the constraint for it.
-We also define a well-formedness relation for RIGs that ensures that each new constraint
-implies the old one\<close>
-
-nominal_function replace_in_g_many :: "\<Gamma> \<Rightarrow> (x*c) list \<Rightarrow> \<Gamma>" where
- "replace_in_g_many G xcs = List.foldr (\<lambda>(x,c) G. G[x \<longmapsto> c]) xcs G"
- by(auto,simp add: eqvt_def replace_in_g_many_graph_aux_def)
-nominal_termination (eqvt) by lexicographic_order
-
-inductive replace_in_g_subtyped :: "\<Theta> \<Rightarrow> \<B> \<Rightarrow> \<Gamma> \<Rightarrow> (x*c) list \<Rightarrow> \<Gamma> \<Rightarrow> bool" (" _ ; _ \<turnstile> _ \<langle> _ \<rangle> \<leadsto> _" [100,50,50] 50) where
- replace_in_g_subtyped_nilI: "\<Theta>; \<B> \<turnstile> G \<langle> [] \<rangle> \<leadsto> G"
-| replace_in_g_subtyped_consI: "\<lbrakk>
- Some (b,c') = lookup G x ;
- \<Theta>; \<B>; G \<turnstile>\<^sub>w\<^sub>f c ;
- \<Theta>; \<B>; G[x\<longmapsto>c] \<Turnstile> c' ;
- \<Theta>; \<B> \<turnstile> G[x\<longmapsto>c] \<langle> xcs \<rangle> \<leadsto> G'; x \<notin> fst ` set xcs \<rbrakk> \<Longrightarrow>
- \<Theta>; \<B> \<turnstile> G \<langle> (x,c)#xcs \<rangle> \<leadsto> G'"
-equivariance replace_in_g_subtyped
-nominal_inductive replace_in_g_subtyped .
-
-inductive_cases replace_in_g_subtyped_elims[elim!]:
- "\<Theta>; \<B> \<turnstile> G \<langle> [] \<rangle> \<leadsto> G'"
- "\<Theta>; \<B> \<turnstile> ((x,b,c)#\<^sub>\<Gamma>\<Gamma> G) \<langle> acs \<rangle> \<leadsto> ((x,b,c)#\<^sub>\<Gamma>G')"
- "\<Theta>; \<B> \<turnstile> G' \<langle> (x,c)# acs \<rangle> \<leadsto> G"
-
-lemma rigs_atom_dom_eq:
- assumes "\<Theta>; \<B> \<turnstile> G \<langle> xcs \<rangle> \<leadsto> G'"
- shows "atom_dom G = atom_dom G'"
- using assms proof(induct rule: replace_in_g_subtyped.induct)
- case (replace_in_g_subtyped_nilI G)
- then show ?case by simp
-next
- case (replace_in_g_subtyped_consI b c' G x \<Theta> \<B> c xcs G')
- then show ?case using rig_dom_eq atom_dom.simps dom.simps by simp
-qed
-
-lemma replace_in_g_wfG:
- assumes "\<Theta>; \<B> \<turnstile> G \<langle> xcs \<rangle> \<leadsto> G'" and "wfG \<Theta> \<B> G "
- shows "wfG \<Theta> \<B> G'"
- using assms proof(induct rule: replace_in_g_subtyped.induct)
- case (replace_in_g_subtyped_nilI \<Theta> G)
- then show ?case by auto
-next
- case (replace_in_g_subtyped_consI b c' G x \<Theta> c xcs G')
- then show ?case using valid_g_wf by auto
-qed
-
-lemma wfD_rig_single:
- fixes \<Delta>::\<Delta> and x::x and c::c and G::\<Gamma>
- assumes "\<Theta>; \<B>; G \<turnstile>\<^sub>w\<^sub>f \<Delta> " and "wfG \<Theta> \<B> (G[x\<longmapsto>c])"
- shows "\<Theta>; \<B>; G[x\<longmapsto>c] \<turnstile>\<^sub>w\<^sub>f \<Delta>"
-proof(cases "atom x \<in> atom_dom G")
- case False
- hence "(G[x\<longmapsto>c]) = G" using assms replace_in_g_forget wfX_wfY by metis
- then show ?thesis using assms by auto
-next
- case True
- then obtain G1 G2 b c' where *: "G=G1@(x,b,c')#\<^sub>\<Gamma>G2" using split_G by fastforce
- hence **: "(G[x\<longmapsto>c]) = G1@(x,b,c)#\<^sub>\<Gamma>G2" using replace_in_g_inside wfD_wf assms wfD_wf by metis
-
- hence "wfG \<Theta> \<B> ((x,b,c)#\<^sub>\<Gamma>G2)" using wfG_suffix assms by auto
- hence "\<Theta>; \<B>; (x, b, TRUE) #\<^sub>\<Gamma> G2 \<turnstile>\<^sub>w\<^sub>f c" using wfG_elim2 by auto
-
- thus ?thesis using wf_replace_inside1 assms * **
- by (simp add: wf_replace_inside2(6))
-qed
-
-lemma wfD_rig:
- assumes "\<Theta>; \<B> \<turnstile> G \<langle> xcs \<rangle> \<leadsto> G'" and "wfD \<Theta> \<B> G \<Delta>"
- shows "wfD \<Theta> \<B> G' \<Delta>"
- using assms proof(induct rule: replace_in_g_subtyped.induct)
- case (replace_in_g_subtyped_nilI \<Theta> G)
- then show ?case by auto
-next
- case (replace_in_g_subtyped_consI b c' G x \<Theta> c xcs G')
- then show ?case using wfD_rig_single valid.simps wfC_wf by auto
-qed
-
-lemma replace_in_g_fresh:
- fixes x::x
- assumes "\<Theta>; \<B> \<turnstile> \<Gamma> \<langle> xcs \<rangle> \<leadsto> \<Gamma>'" and "wfG \<Theta> \<B> \<Gamma>" and "wfG \<Theta> \<B> \<Gamma>'" and "atom x \<sharp> \<Gamma>"
- shows "atom x \<sharp> \<Gamma>'"
- using wfG_dom_supp assms fresh_def rigs_atom_dom_eq by metis
-
-lemma replace_in_g_fresh1:
- fixes x::x
- assumes "\<Theta>; \<B> \<turnstile> \<Gamma> \<langle> xcs \<rangle> \<leadsto> \<Gamma>'" and "wfG \<Theta> \<B> \<Gamma>" and "atom x \<sharp> \<Gamma>"
- shows "atom x \<sharp> \<Gamma>'"
-proof -
- have "wfG \<Theta> \<B> \<Gamma>'" using replace_in_g_wfG assms by auto
- thus ?thesis using assms replace_in_g_fresh by metis
-qed
-
-text \<open> Wellscoping for an eXchange list\<close>
-
-inductive wsX:: "\<Gamma> \<Rightarrow> (x*c) list \<Rightarrow> bool" where
- wsX_NilI: "wsX G []"
-| wsX_ConsI: "\<lbrakk> wsX G xcs ; atom x \<in> atom_dom G ; x \<notin> fst ` set xcs \<rbrakk> \<Longrightarrow> wsX G ((x,c)#xcs)"
-equivariance wsX
-nominal_inductive wsX .
-
-lemma wsX_if1:
- assumes "wsX G xcs"
- shows "(( atom ` fst ` set xcs) \<subseteq> atom_dom G) \<and> List.distinct (List.map fst xcs)"
- using assms by(induct rule: wsX.induct,force+ )
-
-lemma wsX_if2:
- assumes "(( atom ` fst ` set xcs) \<subseteq> atom_dom G) \<and> List.distinct (List.map fst xcs)"
- shows "wsX G xcs"
- using assms proof(induct xcs)
- case Nil
- then show ?case using wsX_NilI by fast
-next
- case (Cons a xcs)
- then obtain x and c where xc: "a=(x,c)" by force
- have " wsX G xcs" proof -
- have "distinct (map fst xcs)" using Cons by force
- moreover have " atom ` fst ` set xcs \<subseteq> atom_dom G" using Cons by simp
- ultimately show ?thesis using Cons by fast
- qed
- moreover have "atom x \<in> atom_dom G" using Cons xc
- by simp
- moreover have "x \<notin> fst ` set xcs" using Cons xc
- by simp
- ultimately show ?case using wsX_ConsI xc by blast
-qed
-
-lemma wsX_iff:
- "wsX G xcs = ((( atom ` fst ` set xcs) \<subseteq> atom_dom G) \<and> List.distinct (List.map fst xcs))"
- using wsX_if1 wsX_if2 by meson
-
-inductive_cases wsX_elims[elim!]:
- "wsX G []"
- "wsX G ((x,c)#xcs)"
-
-lemma wsX_cons:
- assumes "wsX \<Gamma> xcs" and "x \<notin> fst ` set xcs"
- shows "wsX ((x, b, c1) #\<^sub>\<Gamma> \<Gamma>) ((x, c2) # xcs)"
- using assms proof(induct \<Gamma>)
- case GNil
- then show ?case using atom_dom.simps wsX_iff by auto
-next
- case (GCons xbc \<Gamma>)
- obtain x' and b' and c' where xbc: "xbc = (x',b',c')" using prod_cases3 by blast
- then have "atom ` fst ` set xcs \<subseteq> atom_dom (xbc #\<^sub>\<Gamma> \<Gamma>) \<and> distinct (map fst xcs)"
- using GCons.prems(1) wsX_iff by blast
- then have "wsX ((x, b, c1) #\<^sub>\<Gamma> xbc #\<^sub>\<Gamma> \<Gamma>) xcs"
- by (simp add: Un_commute subset_Un_eq wsX_if2)
- then show ?case by (simp add: GCons.prems(2) wsX_ConsI)
-qed
-
-lemma wsX_cons2:
- assumes "wsX \<Gamma> xcs" and "x \<notin> fst ` set xcs"
- shows "wsX ((x, b, c1) #\<^sub>\<Gamma> \<Gamma>) xcs"
- using assms proof(induct \<Gamma>)
- case GNil
- then show ?case using atom_dom.simps wsX_iff by auto
-next
- case (GCons xbc \<Gamma>)
- obtain x' and b' and c' where xbc: "xbc = (x',b',c')" using prod_cases3 by blast
- then have "atom ` fst ` set xcs \<subseteq> atom_dom (xbc #\<^sub>\<Gamma> \<Gamma>) \<and> distinct (map fst xcs)"
- using GCons.prems(1) wsX_iff by blast then show ?case by (simp add: Un_commute subset_Un_eq wsX_if2)
-qed
-
-lemma wsX_cons3:
- assumes "wsX \<Gamma> xcs"
- shows "wsX ((x, b, c1) #\<^sub>\<Gamma> \<Gamma>) xcs"
- using assms proof(induct \<Gamma>)
- case GNil
- then show ?case using atom_dom.simps wsX_iff by auto
-next
- case (GCons xbc \<Gamma>)
- obtain x' and b' and c' where xbc: "xbc = (x',b',c')" using prod_cases3 by blast
- then have "atom ` fst ` set xcs \<subseteq> atom_dom (xbc #\<^sub>\<Gamma> \<Gamma>) \<and> distinct (map fst xcs)"
- using GCons.prems(1) wsX_iff by blast then show ?case by (simp add: Un_commute subset_Un_eq wsX_if2)
-qed
-
-lemma wsX_fresh:
- assumes "wsX G xcs" and "atom x \<sharp> G" and "wfG \<Theta> \<B> G "
- shows "x \<notin> fst ` set xcs"
-proof -
- have "atom x \<notin> atom_dom G" using assms
- using fresh_def wfG_dom_supp by auto
- thus ?thesis using wsX_iff assms by blast
-qed
-
-lemma replace_in_g_dist:
- assumes "x' \<noteq> x"
- shows "replace_in_g ((x, b,c) #\<^sub>\<Gamma> G) x' c'' = ((x, b,c) #\<^sub>\<Gamma> (replace_in_g G x' c''))" using replace_in_g.simps assms by presburger
-
-lemma wfG_replace_inside_rig:
- fixes c''::c
- assumes \<open>\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f G[x'\<longmapsto>c'']\<close> \<open>\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (x, b, c) #\<^sub>\<Gamma> G \<close>
- shows "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (x, b, c) #\<^sub>\<Gamma> G[x'\<longmapsto>c'']"
-proof(rule wfG_consI)
-
- have "wfG \<Theta> \<B> G " using wfG_cons assms by auto
-
- show *:"\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f G[x'\<longmapsto>c'']" using assms by auto
- show "atom x \<sharp> G[x'\<longmapsto>c'']" using replace_in_g_fresh_single[OF *] assms wfG_elims assms by metis
- show **:"\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b " using wfG_elim2 assms by auto
- show "\<Theta>; \<B>; (x, b, TRUE) #\<^sub>\<Gamma> G[x'\<longmapsto>c''] \<turnstile>\<^sub>w\<^sub>f c "
- proof(cases "atom x' \<notin> atom_dom G")
- case True
- hence "G = G[x'\<longmapsto>c'']" using replace_in_g_forget \<open>wfG \<Theta> \<B> G\<close> by auto
- thus ?thesis using assms wfG_wfC by auto
- next
- case False
- then obtain G1 G2 b' c' where **:"G=G1@(x',b',c')#\<^sub>\<Gamma>G2"
- using split_G by fastforce
- hence ***: "(G[x'\<longmapsto>c'']) = G1@(x',b',c'')#\<^sub>\<Gamma>G2"
- using replace_in_g_inside \<open>wfG \<Theta> \<B> G \<close> by metis
- hence "\<Theta>; \<B>; (x, b, TRUE) #\<^sub>\<Gamma> G1@(x',b',c')#\<^sub>\<Gamma>G2 \<turnstile>\<^sub>w\<^sub>f c" using * ** assms wfG_wfC by auto
- hence "\<Theta>; \<B>; (x, b, TRUE) #\<^sub>\<Gamma> G1@(x',b',c'')#\<^sub>\<Gamma>G2 \<turnstile>\<^sub>w\<^sub>f c" using * *** wf_replace_inside assms
- by (metis "**" append_g.simps(2) wfG_elim2 wfG_suffix)
- thus ?thesis using ** * *** by auto
- qed
-qed
-
-lemma replace_in_g_valid_weakening:
- assumes "\<Theta>; \<B>; \<Gamma>[x'\<longmapsto>c''] \<Turnstile> c'" and "x' \<noteq> x" and "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (x, b, c) #\<^sub>\<Gamma> \<Gamma>[x'\<longmapsto>c'']"
- shows "\<Theta>; \<B>; ((x, b,c) #\<^sub>\<Gamma> \<Gamma>)[x'\<longmapsto> c''] \<Turnstile> c'"
- apply(subst replace_in_g_dist,simp add: assms,rule valid_weakening)
- using assms by auto+
-
-lemma replace_in_g_subtyped_cons:
- assumes "replace_in_g_subtyped \<Theta> \<B> G xcs G'" and "wfG \<Theta> \<B> ((x,b,c)#\<^sub>\<Gamma>G)"
- shows "x \<notin> fst ` set xcs \<Longrightarrow> replace_in_g_subtyped \<Theta> \<B> ((x,b,c)#\<^sub>\<Gamma>G) xcs ((x,b,c)#\<^sub>\<Gamma>G')"
- using assms proof(induct rule: replace_in_g_subtyped.induct)
- case (replace_in_g_subtyped_nilI G)
- then show ?case
- by (simp add: replace_in_g_subtyped.replace_in_g_subtyped_nilI)
-next
- case (replace_in_g_subtyped_consI b' c' G x' \<Theta> \<B> c'' xcs' G')
- hence "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f G[x'\<longmapsto>c'']" using valid.simps wfC_wf by auto
-
- show ?case proof(rule replace_in_g_subtyped.replace_in_g_subtyped_consI)
- show " Some (b', c') = lookup ((x, b,c) #\<^sub>\<Gamma> G) x'" using lookup.simps
- fst_conv image_iff \<Gamma>_set_intros surj_pair replace_in_g_subtyped_consI by force
- show wbc: " \<Theta>; \<B>; (x, b, c) #\<^sub>\<Gamma> G \<turnstile>\<^sub>w\<^sub>f c'' " using wf_weakening \<open> \<Theta>; \<B>; G \<turnstile>\<^sub>w\<^sub>f c''\<close> \<open>\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (x, b, c) #\<^sub>\<Gamma> G \<close> by fastforce
- have "x' \<noteq> x" using replace_in_g_subtyped_consI by auto
- have wbc1: "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (x, b, c) #\<^sub>\<Gamma> G[x'\<longmapsto>c'']" proof -
- have "(x, b, c) #\<^sub>\<Gamma> G[x'\<longmapsto>c''] = ((x, b, c) #\<^sub>\<Gamma> G)[x'\<longmapsto>c'']" using \<open>x' \<noteq> x\<close> using replace_in_g.simps by auto
- thus ?thesis using wfG_replace_inside_rig \<open>\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f G[x'\<longmapsto>c'']\<close> \<open>\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (x, b, c) #\<^sub>\<Gamma> G \<close> by fastforce
- qed
- show *: "\<Theta>; \<B>; replace_in_g ((x, b,c) #\<^sub>\<Gamma> G) x' c'' \<Turnstile> c'"
- proof -
- have "\<Theta>; \<B>; G[x'\<longmapsto>c''] \<Turnstile> c'" using replace_in_g_subtyped_consI by auto
- thus ?thesis using replace_in_g_valid_weakening wbc1 \<open>x'\<noteq>x\<close> by auto
- qed
-
- show "replace_in_g_subtyped \<Theta> \<B> (replace_in_g ((x, b,c) #\<^sub>\<Gamma> G) x' c'') xcs' ((x, b,c) #\<^sub>\<Gamma> G')"
- using replace_in_g_subtyped_consI wbc1 by auto
- show "x' \<notin> fst ` set xcs'"
- using replace_in_g_subtyped_consI by linarith
- qed
-qed
-
-lemma replace_in_g_split:
- fixes G::\<Gamma>
- assumes "\<Gamma> = replace_in_g \<Gamma>' x c" and "\<Gamma>' = G'@(x,b,c')#\<^sub>\<Gamma>G" and "wfG \<Theta> \<B> \<Gamma>'"
- shows "\<Gamma> = G'@(x,b,c)#\<^sub>\<Gamma>G"
- using assms proof(induct G' arbitrary: G \<Gamma> \<Gamma>' rule: \<Gamma>_induct)
- case GNil
- then show ?case by simp
-next
- case (GCons x1 b1 c1 \<Gamma>1)
- hence "x1 \<noteq> x"
- using wfG_cons_fresh2[of \<Theta> \<B> x1 b1 c1 \<Gamma>1 x b ]
- using GCons.prems(2) GCons.prems(3) append_g.simps(2) by auto
- moreover hence *: "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (\<Gamma>1 @ (x, b, c') #\<^sub>\<Gamma> G)" using GCons append_g.simps wfG_elims by metis
- moreover hence "replace_in_g (\<Gamma>1 @ (x, b, c') #\<^sub>\<Gamma> G) x c = \<Gamma>1 @ (x, b, c) #\<^sub>\<Gamma> G" using GCons replace_in_g_inside[OF *, of c] by auto
-
- ultimately show ?case using replace_in_g.simps(2)[of x1 b1 c1 " \<Gamma>1 @ (x, b, c') #\<^sub>\<Gamma> G" x c] GCons
- by (simp add: GCons.prems(1) GCons.prems(2))
-qed
-
-lemma replace_in_g_subtyped_split0:
- fixes G::\<Gamma>
- assumes "replace_in_g_subtyped \<Theta> \<B> \<Gamma>'[(x,c)] \<Gamma>" and "\<Gamma>' = G'@(x,b,c')#\<^sub>\<Gamma>G" and "wfG \<Theta> \<B> \<Gamma>'"
- shows "\<Gamma> = G'@(x,b,c)#\<^sub>\<Gamma>G"
-proof -
- have "\<Gamma> = replace_in_g \<Gamma>' x c" using assms replace_in_g_subtyped.simps
- by (metis Pair_inject list.distinct(1) list.inject)
- thus ?thesis using assms replace_in_g_split by blast
-qed
-
-lemma replace_in_g_subtyped_split:
- assumes "Some (b, c') = lookup G x" and "\<Theta>; \<B>; replace_in_g G x c \<Turnstile> c'" and "wfG \<Theta> \<B> G "
- shows "\<exists> \<Gamma> \<Gamma>'. G = \<Gamma>'@(x,b,c')#\<^sub>\<Gamma>\<Gamma> \<and> \<Theta>; \<B>; \<Gamma>'@(x,b,c)#\<^sub>\<Gamma>\<Gamma> \<Turnstile> c'"
-proof -
- obtain \<Gamma> and \<Gamma>' where "G = \<Gamma>'@(x,b,c')#\<^sub>\<Gamma>\<Gamma>" using assms lookup_split by blast
- moreover hence "replace_in_g G x c = \<Gamma>'@(x,b,c)#\<^sub>\<Gamma>\<Gamma>" using replace_in_g_split assms by blast
- ultimately show ?thesis by (metis assms(2))
-qed
-
-section \<open>Validity and Subtyping\<close>
-
-lemma wfC_replace_in_g:
- fixes c::c and c0::c
- assumes "\<Theta>; \<B>; \<Gamma>'@(x,b,c0')#\<^sub>\<Gamma>\<Gamma> \<turnstile>\<^sub>w\<^sub>f c" and "\<Theta>; \<B>; (x,b,TRUE)#\<^sub>\<Gamma>\<Gamma> \<turnstile>\<^sub>w\<^sub>f c0"
- shows "\<Theta>; \<B>; \<Gamma>' @ (x, b, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f c"
- using wf_replace_inside1(2) assms by auto
-
-
-lemma ctx_subtype_valid:
- assumes "\<Theta>; \<B>; \<Gamma>'@(x,b,c0')#\<^sub>\<Gamma>\<Gamma> \<Turnstile> c" and
- "\<Theta>; \<B>; \<Gamma>'@(x,b,c0)#\<^sub>\<Gamma>\<Gamma> \<Turnstile> c0'"
- shows "\<Theta>; \<B>; \<Gamma>'@(x,b,c0)#\<^sub>\<Gamma>\<Gamma> \<Turnstile> c"
-proof(rule validI)
- show "\<Theta>; \<B>; \<Gamma>' @ (x, b, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f c" proof -
- have "\<Theta>; \<B>; \<Gamma>'@(x,b,c0')#\<^sub>\<Gamma>\<Gamma> \<turnstile>\<^sub>w\<^sub>f c" using valid.simps assms by auto
- moreover have "\<Theta>; \<B>; (x,b,TRUE)#\<^sub>\<Gamma>\<Gamma> \<turnstile>\<^sub>w\<^sub>f c0" proof -
- have "wfG \<Theta> \<B> (\<Gamma>'@(x,b,c0)#\<^sub>\<Gamma>\<Gamma>)" using assms valid.simps wfC_wf by auto
- hence "wfG \<Theta> \<B> ((x,b,c0)#\<^sub>\<Gamma>\<Gamma>)" using wfG_suffix by auto
- thus ?thesis using wfG_wfC by auto
- qed
- ultimately show ?thesis using assms wfC_replace_in_g by auto
- qed
-
- show "\<forall>i. wfI \<Theta> (\<Gamma>' @ (x, b, c0) #\<^sub>\<Gamma> \<Gamma>) i \<and> is_satis_g i (\<Gamma>' @ (x, b, c0) #\<^sub>\<Gamma> \<Gamma>) \<longrightarrow> is_satis i c" proof(rule,rule)
- fix i
- assume * : "wfI \<Theta> (\<Gamma>' @ (x, b, c0) #\<^sub>\<Gamma> \<Gamma>) i \<and> is_satis_g i (\<Gamma>' @ (x, b, c0) #\<^sub>\<Gamma> \<Gamma>) "
-
- hence "is_satis_g i (\<Gamma>'@(x, b, c0) #\<^sub>\<Gamma> \<Gamma>) \<and> wfI \<Theta> (\<Gamma>'@(x, b, c0) #\<^sub>\<Gamma> \<Gamma>) i" using is_satis_g_append wfI_suffix by metis
- moreover hence "is_satis i c0'" using valid.simps assms by presburger
-
- moreover have "is_satis_g i \<Gamma>'" using is_satis_g_append * by simp
- ultimately have "is_satis_g i (\<Gamma>' @ (x, b, c0') #\<^sub>\<Gamma> \<Gamma>) " using is_satis_g_append by simp
- moreover have "wfI \<Theta> (\<Gamma>' @ (x, b, c0') #\<^sub>\<Gamma> \<Gamma>) i" using wfI_def wfI_suffix * wfI_def wfI_replace_inside by metis
- ultimately show "is_satis i c" using assms valid.simps by metis
- qed
-qed
-
-lemma ctx_subtype_subtype:
- fixes \<Gamma>::\<Gamma>
- shows "\<Theta>; \<B>; G \<turnstile> t1 \<lesssim> t2 \<Longrightarrow> G = \<Gamma>'@(x,b0,c0')#\<^sub>\<Gamma>\<Gamma> \<Longrightarrow> \<Theta>; \<B>; \<Gamma>'@(x,b0,c0)#\<^sub>\<Gamma>\<Gamma> \<Turnstile> c0' \<Longrightarrow> \<Theta>; \<B>; \<Gamma>'@(x,b0,c0)#\<^sub>\<Gamma>\<Gamma> \<turnstile> t1 \<lesssim> t2"
-proof(nominal_induct avoiding: c0 rule: subtype.strong_induct)
-
- case (subtype_baseI x' \<Theta> \<B> \<Gamma>'' z c z' c' b)
- let ?\<Gamma>c0 = "\<Gamma>'@(x,b0,c0)#\<^sub>\<Gamma>\<Gamma>"
- have wb1: "wfG \<Theta> \<B> ?\<Gamma>c0" using valid.simps wfC_wf subtype_baseI by metis
- show ?case proof
- show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b | c \<rbrace> \<close> using wfT_replace_inside2[OF _ wb1] subtype_baseI by metis
- show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> z' : b | c' \<rbrace> \<close> using wfT_replace_inside2[OF _ wb1] subtype_baseI by metis
- have "atom x' \<sharp> \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma>" using fresh_prodN subtype_baseI fresh_replace_inside wb1 subtype_wf wfX_wfY by metis
- thus \<open>atom x' \<sharp> (\<Theta>, \<B>, \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma>, z, c , z' , c' )\<close> using subtype_baseI fresh_prodN by metis
- have "\<Theta>; \<B>; ((x', b, c[z::=V_var x']\<^sub>v) #\<^sub>\<Gamma> \<Gamma>') @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<Turnstile> c'[z'::=V_var x']\<^sub>v " proof(rule ctx_subtype_valid)
- show 1: \<open>\<Theta>; \<B>; ((x', b, c[z::=V_var x']\<^sub>v) #\<^sub>\<Gamma> \<Gamma>') @ (x, b0, c0') #\<^sub>\<Gamma> \<Gamma> \<Turnstile> c'[z'::=V_var x']\<^sub>v \<close>
- using subtype_baseI append_g.simps subst_defs by metis
- have *:"\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f ((x', b, c[z::=V_var x']\<^sub>v) #\<^sub>\<Gamma> \<Gamma>') @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> " proof(rule wfG_replace_inside2)
- show "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f ((x', b, c[z::=V_var x']\<^sub>v) #\<^sub>\<Gamma> \<Gamma>') @ (x, b0, c0') #\<^sub>\<Gamma> \<Gamma>"
- using * valid_wf_all wfC_wf 1 append_g.simps by metis
- show "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (x, b0, c0) #\<^sub>\<Gamma> \<Gamma>" using wfG_suffix wb1 by auto
- qed
- moreover have "toSet (\<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma>) \<subseteq> toSet (((x', b, c[z::=V_var x']\<^sub>v) #\<^sub>\<Gamma> \<Gamma>') @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma>)" using toSet.simps append_g.simps by auto
- ultimately show \<open>\<Theta>; \<B>; ((x', b, c[z::=V_var x']\<^sub>v) #\<^sub>\<Gamma> \<Gamma>') @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<Turnstile> c0' \<close> using valid_weakening subtype_baseI * by blast
- qed
- thus \<open>\<Theta>; \<B>; (x', b, c[z::=V_var x']\<^sub>v) #\<^sub>\<Gamma> \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<Turnstile> c'[z'::=V_var x']\<^sub>v \<close> using append_g.simps subst_defs by simp
- qed
-qed
-
-lemma ctx_subtype_subtype_rig:
- assumes "replace_in_g_subtyped \<Theta> \<B> \<Gamma>' [(x,c0)] \<Gamma>" and "\<Theta>; \<B>; \<Gamma>' \<turnstile> t1 \<lesssim> t2"
- shows "\<Theta>; \<B>; \<Gamma> \<turnstile> t1 \<lesssim> t2"
-proof -
- have wf: "wfG \<Theta> \<B> \<Gamma>'" using subtype_g_wf assms by auto
- obtain b and c0' where "Some (b, c0') = lookup \<Gamma>' x \<and> (\<Theta>; \<B>; replace_in_g \<Gamma>' x c0 \<Turnstile> c0')" using
- replace_in_g_subtyped.simps[of \<Theta> \<B> \<Gamma>' "[(x, c0)]" \<Gamma>] assms(1)
-
- by (metis fst_conv list.inject list.set_intros(1) list.simps(15) not_Cons_self2 old.prod.exhaust prod.inject set_ConsD surj_pair)
- moreover then obtain G and G' where *: "\<Gamma>' = G'@(x,b,c0')#\<^sub>\<Gamma>G \<and> \<Theta>; \<B>; G'@(x,b,c0)#\<^sub>\<Gamma>G \<Turnstile> c0'"
- using replace_in_g_subtyped_split[of b c0' \<Gamma>' x \<Theta> \<B> c0] wf by metis
-
- ultimately show ?thesis using ctx_subtype_subtype
- assms(1) assms(2) replace_in_g_subtyped_split0 subtype_g_wf
- by (metis (no_types, lifting) local.wf replace_in_g_split)
-qed
-
-text \<open> We now prove versions of the @{text "ctx_subtype"} lemmas above using @{text "replace_in_g"}. First we do case where
-the replace is just for a single variable (indicated by suffix rig) and then the general case for
-multiple replacements (indicated by suffix rigs)\<close>
-
-lemma ctx_subtype_subtype_rigs:
- assumes "replace_in_g_subtyped \<Theta> \<B> \<Gamma>' xcs \<Gamma>" and "\<Theta>; \<B>; \<Gamma>' \<turnstile> t1 \<lesssim> t2"
- shows "\<Theta>; \<B>; \<Gamma> \<turnstile> t1 \<lesssim> t2"
- using assms proof(induct xcs arbitrary: \<Gamma> \<Gamma>' )
- case Nil
- moreover have "\<Gamma>' = \<Gamma>" using replace_in_g_subtyped_nilI
- using calculation(1) by blast
- ultimately show ?case by auto
-next
- case (Cons a xcs)
-
- then obtain x and c where "a=(x,c)" by fastforce
- then obtain b and c' where bc: "Some (b, c') = lookup \<Gamma>' x \<and>
- replace_in_g_subtyped \<Theta> \<B> (replace_in_g \<Gamma>' x c) xcs \<Gamma> \<and> \<Theta>; \<B>; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f c \<and>
- x \<notin> fst ` set xcs \<and> \<Theta>; \<B>; (replace_in_g \<Gamma>' x c) \<Turnstile> c' " using replace_in_g_subtyped_elims(3)[of \<Theta> \<B> \<Gamma>' x c xcs \<Gamma>] Cons
- by (metis valid.simps)
-
- hence *: "replace_in_g_subtyped \<Theta> \<B> \<Gamma>' [(x,c)] (replace_in_g \<Gamma>' x c)" using replace_in_g_subtyped_consI
- by (meson image_iff list.distinct(1) list.set_cases replace_in_g_subtyped_nilI)
-
- hence "\<Theta>; \<B>; (replace_in_g \<Gamma>' x c) \<turnstile> t1 \<lesssim> t2"
- using ctx_subtype_subtype_rig * assms Cons.prems(2) by auto
-
- moreover have "replace_in_g_subtyped \<Theta> \<B> (replace_in_g \<Gamma>' x c) xcs \<Gamma>" using Cons
- using bc by blast
-
- ultimately show ?case using Cons by blast
-qed
-
-lemma replace_in_g_inside_valid:
- assumes "replace_in_g_subtyped \<Theta> \<B> \<Gamma>' [(x,c0)] \<Gamma>" and "wfG \<Theta> \<B> \<Gamma>'"
- shows "\<exists>b c0' G G'. \<Gamma>' = G' @ (x,b,c0')#\<^sub>\<Gamma>G \<and> \<Gamma> = G' @ (x,b,c0)#\<^sub>\<Gamma>G \<and> \<Theta>; \<B>; G'@ (x,b,c0)#\<^sub>\<Gamma>G \<Turnstile> c0'"
-proof -
- obtain b and c0' where bc: "Some (b, c0') = lookup \<Gamma>' x \<and> \<Theta>; \<B>; replace_in_g \<Gamma>' x c0 \<Turnstile> c0'" using
- replace_in_g_subtyped.simps[of \<Theta> \<B> \<Gamma>' "[(x, c0)]" \<Gamma>] assms(1)
- by (metis fst_conv list.inject list.set_intros(1) list.simps(15) not_Cons_self2 old.prod.exhaust prod.inject set_ConsD surj_pair)
- then obtain G and G' where *: "\<Gamma>' = G'@(x,b,c0')#\<^sub>\<Gamma>G \<and> \<Theta>; \<B>; G'@(x,b,c0)#\<^sub>\<Gamma>G \<Turnstile> c0'" using replace_in_g_subtyped_split[of b c0' \<Gamma>' x \<Theta> \<B> c0] assms
- by metis
- thus ?thesis using replace_in_g_inside bc
- using assms(1) assms(2) by blast
-qed
-
-lemma replace_in_g_valid:
- assumes "\<Theta>; \<B> \<turnstile> G \<langle> xcs \<rangle> \<leadsto> G'" and "\<Theta>; \<B>; G \<Turnstile> c "
- shows \<open>\<Theta>; \<B>; G' \<Turnstile> c \<close>
- using assms proof(induct rule: replace_in_g_subtyped.inducts)
- case (replace_in_g_subtyped_nilI \<Theta> \<B> G)
- then show ?case by auto
-next
- case (replace_in_g_subtyped_consI b c1 G x \<Theta> \<B> c2 xcs G')
- hence "\<Theta>; \<B>; G[x\<longmapsto>c2] \<Turnstile> c"
- by (metis ctx_subtype_valid replace_in_g_split replace_in_g_subtyped_split valid_g_wf)
- then show ?case using replace_in_g_subtyped_consI by auto
-qed
-
-section \<open>Literals\<close>
-
-section \<open>Values\<close>
-
-lemma lookup_inside_unique_b[simp]:
- assumes "\<Theta> ; B \<turnstile>\<^sub>w\<^sub>f (\<Gamma>'@(x,b0,c0)#\<^sub>\<Gamma>\<Gamma>)" and "\<Theta> ; B \<turnstile>\<^sub>w\<^sub>f (\<Gamma>'@(x,b0,c0')#\<^sub>\<Gamma>\<Gamma>)"
- and "Some (b, c) = lookup (\<Gamma>' @ (x, b0, c0') #\<^sub>\<Gamma> \<Gamma>) y" and "Some (b0,c0) = lookup (\<Gamma>'@((x,b0,c0))#\<^sub>\<Gamma>\<Gamma>) x" and "x=y"
- shows "b = b0"
- by (metis assms(2) assms(3) assms(5) lookup_inside_wf old.prod.exhaust option.inject prod.inject)
-
-lemma ctx_subtype_v_aux:
- fixes v::v
- assumes "\<Theta>; \<B>; \<Gamma>'@((x,b0,c0')#\<^sub>\<Gamma>\<Gamma>) \<turnstile> v \<Rightarrow> t1" and "\<Theta>; \<B>; \<Gamma>'@(x,b0,c0)#\<^sub>\<Gamma>\<Gamma> \<Turnstile> c0'"
- shows "\<Theta>; \<B>; \<Gamma>'@((x,b0,c0)#\<^sub>\<Gamma>\<Gamma>) \<turnstile> v \<Rightarrow> t1"
- using assms proof(nominal_induct "\<Gamma>'@((x,b0,c0')#\<^sub>\<Gamma>\<Gamma>)" v t1 avoiding: c0 rule: infer_v.strong_induct)
- case (infer_v_varI \<Theta> \<B> b c xa z)
- have wf:\<open> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<close> using wfG_inside_valid2 infer_v_varI by metis
- have xf1:\<open>atom z \<sharp> xa\<close> using infer_v_varI by metis
- have xf2: \<open>atom z \<sharp> (\<Theta>, \<B>, \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma>)\<close> apply( fresh_mth add: infer_v_varI )
- using fresh_def infer_v_varI wfG_supp fresh_append_g fresh_GCons fresh_prodN by metis+
- show ?case proof (cases "x=xa")
- case True
- moreover have "b = b0" using infer_v_varI True by simp
- moreover hence \<open>Some (b, c0) = lookup (\<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma>) xa\<close> using lookup_inside_wf[OF wf] infer_v_varI True by auto
- ultimately show ?thesis using wf xf1 xf2 Typing.infer_v_varI by metis
- next
- case False
- moreover hence \<open>Some (b, c) = lookup (\<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma>) xa\<close> using lookup_inside2 infer_v_varI by metis
- ultimately show ?thesis using wf xf1 xf2 Typing.infer_v_varI by simp
- qed
-next
- case (infer_v_litI \<Theta> \<B> l \<tau>)
- thus ?case using Typing.infer_v_litI wfG_inside_valid2 by simp
-next
- case (infer_v_pairI z v1 v2 \<Theta> \<B> t1' t2' c0)
- show ?case proof
- show "atom z \<sharp> (v1, v2)" using infer_v_pairI fresh_Pair by simp
- show "atom z \<sharp> (\<Theta>, \<B>, \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma>)" apply( fresh_mth add: infer_v_pairI )
- using fresh_def infer_v_pairI wfG_supp fresh_append_g fresh_GCons fresh_prodN by metis+
- show "\<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile> v1 \<Rightarrow> t1'" using infer_v_pairI by simp
- show "\<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile> v2 \<Rightarrow> t2'" using infer_v_pairI by simp
- qed
-next
- case (infer_v_consI s dclist \<Theta> dc tc \<B> v tv z)
- show ?case proof
- show \<open>AF_typedef s dclist \<in> set \<Theta>\<close> using infer_v_consI by auto
- show \<open>(dc, tc) \<in> set dclist\<close> using infer_v_consI by auto
- show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile> v \<Rightarrow> tv\<close> using infer_v_consI by auto
- show \<open>\<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile> tv \<lesssim> tc\<close> using infer_v_consI ctx_subtype_subtype by auto
- show \<open>atom z \<sharp> v\<close> using infer_v_consI by auto
- show \<open>atom z \<sharp> (\<Theta>, \<B>, \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma>)\<close> apply( fresh_mth add: infer_v_consI )
- using fresh_def infer_v_consI wfG_supp fresh_append_g fresh_GCons fresh_prodN by metis+
- qed
-next
- case (infer_v_conspI s bv dclist \<Theta> dc tc \<B> v tv b z)
- show ?case proof
- show \<open>AF_typedef_poly s bv dclist \<in> set \<Theta>\<close> using infer_v_conspI by auto
- show \<open>(dc, tc) \<in> set dclist\<close> using infer_v_conspI by auto
- show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile> v \<Rightarrow> tv\<close> using infer_v_conspI by auto
- show \<open>\<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile> tv \<lesssim> tc[bv::=b]\<^sub>\<tau>\<^sub>b\<close> using infer_v_conspI ctx_subtype_subtype by auto
- show \<open>atom z \<sharp> (\<Theta>, \<B>, \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma>, v, b)\<close> apply( fresh_mth add: infer_v_conspI )
- using fresh_def infer_v_conspI wfG_supp fresh_append_g fresh_GCons fresh_prodN by metis+
- show \<open>atom bv \<sharp> (\<Theta>, \<B>, \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma>, v, b)\<close> apply( fresh_mth add: infer_v_conspI )
- using fresh_def infer_v_conspI wfG_supp fresh_append_g fresh_GCons fresh_prodN by metis+
- show \<open> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b \<close> using infer_v_conspI by auto
- qed
-qed
-
-lemma ctx_subtype_v:
- fixes v::v
- assumes "\<Theta>; \<B>; \<Gamma>'@((x,b0,c0')#\<^sub>\<Gamma>\<Gamma>) \<turnstile> v \<Rightarrow> t1" and "\<Theta>; \<B>; \<Gamma>'@(x,b0,c0)#\<^sub>\<Gamma>\<Gamma> \<Turnstile> c0'"
- shows "\<exists>t2. \<Theta>; \<B>; \<Gamma>'@((x,b0,c0)#\<^sub>\<Gamma>\<Gamma>) \<turnstile> v \<Rightarrow> t2 \<and> \<Theta>; \<B>; \<Gamma>'@((x,b0,c0)#\<^sub>\<Gamma>\<Gamma>) \<turnstile> t2 \<lesssim> t1"
-proof -
-
- have "\<Theta>; \<B>; \<Gamma>'@((x,b0,c0)#\<^sub>\<Gamma>\<Gamma>) \<turnstile> v \<Rightarrow> t1 " using ctx_subtype_v_aux assms by auto
- moreover hence "\<Theta>; \<B>; \<Gamma>'@((x,b0,c0)#\<^sub>\<Gamma>\<Gamma>) \<turnstile> t1 \<lesssim> t1" using subtype_reflI2 infer_v_wf by simp
- ultimately show ?thesis by auto
-qed
-
-lemma ctx_subtype_v_eq:
- fixes v::v
- assumes
- "\<Theta>; \<B>; \<Gamma>'@((x,b0,c0')#\<^sub>\<Gamma>\<Gamma>) \<turnstile> v \<Rightarrow> t1" and
- " \<Theta>; \<B>; \<Gamma>'@(x,b0,c0)#\<^sub>\<Gamma>\<Gamma> \<Turnstile> c0'"
- shows "\<Theta>; \<B>; \<Gamma>'@((x,b0,c0)#\<^sub>\<Gamma>\<Gamma>) \<turnstile> v \<Rightarrow> t1"
-proof -
- obtain t1' where "\<Theta>; \<B>; \<Gamma>'@((x,b0,c0)#\<^sub>\<Gamma>\<Gamma>) \<turnstile> v \<Rightarrow> t1'" using ctx_subtype_v assms by metis
- moreover have "replace_in_g (\<Gamma>'@((x,b0,c0')#\<^sub>\<Gamma>\<Gamma>)) x c0 = \<Gamma>'@((x,b0,c0)#\<^sub>\<Gamma>\<Gamma>)" using replace_in_g_inside infer_v_wf assms by metis
- ultimately show ?thesis using infer_v_uniqueness_rig assms by metis
-qed
-
-lemma ctx_subtype_check_v_eq:
- assumes "\<Theta>; \<B>; \<Gamma>'@((x,b0,c0')#\<^sub>\<Gamma>\<Gamma>) \<turnstile> v \<Leftarrow> t1" and " \<Theta>; \<B>; \<Gamma>'@(x,b0,c0)#\<^sub>\<Gamma>\<Gamma> \<Turnstile> c0'"
- shows "\<Theta>; \<B>; \<Gamma>'@((x,b0,c0)#\<^sub>\<Gamma>\<Gamma>) \<turnstile> v \<Leftarrow> t1"
-proof -
- obtain t2 where t2: "\<Theta>; \<B>; \<Gamma>'@((x,b0,c0')#\<^sub>\<Gamma>\<Gamma>) \<turnstile> v \<Rightarrow> t2 \<and> \<Theta>; \<B>; \<Gamma>'@((x,b0,c0')#\<^sub>\<Gamma>\<Gamma>) \<turnstile> t2 \<lesssim> t1"
- using check_v_elims assms by blast
- hence t3: "\<Theta>; \<B>; \<Gamma>'@((x,b0,c0)#\<^sub>\<Gamma>\<Gamma>) \<turnstile> v \<Rightarrow> t2"
- using assms ctx_subtype_v_eq by blast
-
- have "\<Theta>; \<B>; \<Gamma>'@((x,b0,c0)#\<^sub>\<Gamma>\<Gamma>) \<turnstile> v \<Rightarrow> t2" using t3 by auto
- moreover have " \<Theta>; \<B>; \<Gamma>'@((x,b0,c0)#\<^sub>\<Gamma>\<Gamma>) \<turnstile> t2 \<lesssim> t1" proof -
-
- have " \<Theta>; \<B>; \<Gamma>'@((x,b0,c0')#\<^sub>\<Gamma>\<Gamma>) \<turnstile> t2 \<lesssim> t1" using t2 by auto
- thus ?thesis using subtype_trans
- using assms(2) ctx_subtype_subtype by blast
- qed
- ultimately show ?thesis using check_v.intros by presburger
-qed
-
-text \<open> Basically the same as @{text "ctx_subtype_v_eq"} but in a different form\<close>
-lemma ctx_subtype_v_rig_eq:
- fixes v::v
- assumes "replace_in_g_subtyped \<Theta> \<B> \<Gamma>' [(x,c0)] \<Gamma>" and
- "\<Theta>; \<B>; \<Gamma>' \<turnstile> v \<Rightarrow> t1"
- shows "\<Theta>; \<B>; \<Gamma> \<turnstile> v \<Rightarrow> t1"
-proof -
- obtain b and c0' and G and G' where "\<Gamma>' = G' @ (x,b,c0')#\<^sub>\<Gamma>G \<and> \<Gamma> = G' @ (x,b,c0)#\<^sub>\<Gamma>G \<and> \<Theta>; \<B>; G'@ (x,b,c0)#\<^sub>\<Gamma>G \<Turnstile> c0'"
- using assms replace_in_g_inside_valid infer_v_wf by metis
- thus ?thesis using ctx_subtype_v_eq[of \<Theta> \<B> G' x b c0' G v t1 c0] assms by simp
-qed
-
-lemma ctx_subtype_v_rigs_eq:
- fixes v::v
- assumes "replace_in_g_subtyped \<Theta> \<B> \<Gamma>' xcs \<Gamma>" and
- "\<Theta>; \<B>; \<Gamma>' \<turnstile> v \<Rightarrow> t1"
- shows "\<Theta>; \<B>; \<Gamma> \<turnstile> v \<Rightarrow> t1"
- using assms proof(induct xcs arbitrary: \<Gamma> \<Gamma>' t1 )
- case Nil
- then show ?case by auto
-next
- case (Cons a xcs)
- then obtain x and c where "a=(x,c)" by fastforce
-
- then obtain b and c' where bc: "Some (b, c') = lookup \<Gamma>' x \<and>
- replace_in_g_subtyped \<Theta> \<B> (replace_in_g \<Gamma>' x c) xcs \<Gamma> \<and> \<Theta>; \<B>; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f c \<and>
- x \<notin> fst ` set xcs \<and> \<Theta>; \<B>; (replace_in_g \<Gamma>' x c) \<Turnstile> c' "
- using replace_in_g_subtyped_elims(3)[of \<Theta> \<B> \<Gamma>' x c xcs \<Gamma>] Cons by (metis valid.simps)
-
- hence *: "replace_in_g_subtyped \<Theta> \<B> \<Gamma>' [(x,c)] (replace_in_g \<Gamma>' x c)" using replace_in_g_subtyped_consI
- by (meson image_iff list.distinct(1) list.set_cases replace_in_g_subtyped_nilI)
- hence t2: "\<Theta>; \<B>; (replace_in_g \<Gamma>' x c) \<turnstile> v \<Rightarrow> t1 " using ctx_subtype_v_rig_eq[OF * Cons(3)] by blast
- moreover have **: "replace_in_g_subtyped \<Theta> \<B> (replace_in_g \<Gamma>' x c) xcs \<Gamma>" using bc by auto
- ultimately have t2': "\<Theta>; \<B>; \<Gamma> \<turnstile> v \<Rightarrow> t1" using Cons by blast
- thus ?case by blast
-qed
-
-lemma ctx_subtype_check_v_rigs_eq:
- assumes "replace_in_g_subtyped \<Theta> \<B> \<Gamma>' xcs \<Gamma>" and
- "\<Theta>; \<B>; \<Gamma>' \<turnstile> v \<Leftarrow> t1"
- shows "\<Theta>; \<B>; \<Gamma> \<turnstile> v \<Leftarrow> t1"
-proof -
- obtain t2 where "\<Theta>; \<B>; \<Gamma>' \<turnstile> v \<Rightarrow> t2 \<and> \<Theta>; \<B>; \<Gamma>' \<turnstile> t2 \<lesssim> t1" using check_v_elims assms by fast
- hence "\<Theta>; \<B>; \<Gamma> \<turnstile> v \<Rightarrow> t2 \<and> \<Theta>; \<B>; \<Gamma> \<turnstile> t2 \<lesssim> t1" using ctx_subtype_v_rigs_eq ctx_subtype_subtype_rigs
- using assms(1) by blast
- thus ?thesis
- using check_v_subtypeI by blast
-qed
-
-section \<open>Expressions\<close>
-
-lemma valid_wfC:
- fixes c0::c
- assumes "\<Theta>; \<B>; \<Gamma>'@(x,b0,c0)#\<^sub>\<Gamma>\<Gamma> \<Turnstile> c0'"
- shows "\<Theta>; \<B>; (x, b0, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f c0"
- using wfG_elim2 valid.simps wfG_suffix
- using assms valid_g_wf by metis
-
-lemma ctx_subtype_e_eq:
- fixes G::\<Gamma>
- assumes
- "\<Theta> ; \<Phi> ; \<B> ; G ; \<Delta> \<turnstile> e \<Rightarrow> t1" and "G = \<Gamma>'@((x,b0,c0')#\<^sub>\<Gamma>\<Gamma>)"
- "\<Theta>; \<B>; \<Gamma>'@(x,b0,c0)#\<^sub>\<Gamma>\<Gamma> \<Turnstile> c0'"
- shows "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>'@((x,b0,c0)#\<^sub>\<Gamma>\<Gamma>) ; \<Delta> \<turnstile> e \<Rightarrow> t1"
- using assms proof(nominal_induct t1 avoiding: c0 rule: infer_e.strong_induct)
- case (infer_e_valI \<Theta> \<B> \<Gamma>'' \<Delta> \<Phi> v \<tau>)
- show ?case proof
- show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wf_replace_inside2(6) valid_wfC infer_e_valI by auto
- show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using infer_e_valI by auto
- show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile> v \<Rightarrow> \<tau>\<close> using infer_e_valI ctx_subtype_v_eq by auto
- qed
-next
- case (infer_e_plusI \<Theta> \<B> \<Gamma>'' \<Delta> \<Phi> v1 z1 c1 v2 z2 c2 z3)
- show ?case proof
- show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wf_replace_inside2(6) valid_wfC infer_e_plusI by auto
- show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using infer_e_plusI by auto
- show *:\<open> \<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile> v1 \<Rightarrow> \<lbrace> z1 : B_int | c1 \<rbrace>\<close> using infer_e_plusI ctx_subtype_v_eq by auto
- show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile> v2 \<Rightarrow> \<lbrace> z2 : B_int | c2 \<rbrace>\<close> using infer_e_plusI ctx_subtype_v_eq by auto
- show \<open>atom z3 \<sharp> AE_op Plus v1 v2\<close> using infer_e_plusI by auto
- show \<open>atom z3 \<sharp> \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma>\<close> using * infer_e_plusI fresh_replace_inside infer_v_wf by metis
- qed
-next
- case (infer_e_leqI \<Theta> \<B> \<Gamma>'' \<Delta> \<Phi> v1 z1 c1 v2 z2 c2 z3)
- show ?case proof
- show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wf_replace_inside2(6) valid_wfC infer_e_leqI by auto
- show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using infer_e_leqI by auto
- show *:\<open> \<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile> v1 \<Rightarrow> \<lbrace> z1 : B_int | c1 \<rbrace>\<close> using infer_e_leqI ctx_subtype_v_eq by auto
- show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile> v2 \<Rightarrow> \<lbrace> z2 : B_int | c2 \<rbrace>\<close> using infer_e_leqI ctx_subtype_v_eq by auto
- show \<open>atom z3 \<sharp> AE_op LEq v1 v2\<close> using infer_e_leqI by auto
- show \<open>atom z3 \<sharp> \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma>\<close> using * infer_e_leqI fresh_replace_inside infer_v_wf by metis
- qed
-next
- case (infer_e_eqI \<Theta> \<B> \<Gamma>'' \<Delta> \<Phi> v1 z1 bb c1 v2 z2 c2 z3)
- show ?case proof
- show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wf_replace_inside2(6) valid_wfC infer_e_eqI by auto
- show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using infer_e_eqI by auto
- show *:\<open> \<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile> v1 \<Rightarrow> \<lbrace> z1 : bb | c1 \<rbrace>\<close> using infer_e_eqI ctx_subtype_v_eq by auto
- show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile> v2 \<Rightarrow> \<lbrace> z2 : bb | c2 \<rbrace>\<close> using infer_e_eqI ctx_subtype_v_eq by auto
- show \<open>atom z3 \<sharp> AE_op Eq v1 v2\<close> using infer_e_eqI by auto
- show \<open>atom z3 \<sharp> \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma>\<close> using * infer_e_eqI fresh_replace_inside infer_v_wf by metis
- show "bb \<in> {B_bool, B_int, B_unit}" using infer_e_eqI by auto
- qed
-next
- case (infer_e_appI \<Theta> \<B> \<Gamma>'' \<Delta> \<Phi> f x' b c \<tau>' s' v \<tau>)
- show ?case proof
- show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wf_replace_inside2(6) valid_wfC infer_e_appI by auto
- show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using infer_e_appI by auto
- show \<open>Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x' b c \<tau>' s'))) = lookup_fun \<Phi> f\<close> using infer_e_appI by auto
- show \<open>\<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile> v \<Leftarrow> \<lbrace> x' : b | c \<rbrace>\<close> using infer_e_appI ctx_subtype_check_v_eq by auto
- thus \<open>atom x' \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma>, \<Delta>, v, \<tau>)\<close>
- using infer_e_appI fresh_replace_inside[of \<Theta> \<B> \<Gamma>' x b0 c0' \<Gamma> c0 x'] infer_v_wf by auto
- show \<open>\<tau>'[x'::=v]\<^sub>v = \<tau>\<close> using infer_e_appI by auto
- qed
-next
- case (infer_e_appPI \<Theta> \<B> \<Gamma>1 \<Delta> \<Phi> b' f bv x1 b c \<tau>' s' v \<tau>)
- show ?case proof
- show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wf_replace_inside2(6) valid_wfC infer_e_appPI by auto
- show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using infer_e_appPI by auto
- show \<open> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b' \<close> using infer_e_appPI by auto
- show \<open>Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x1 b c \<tau>' s'))) = lookup_fun \<Phi> f\<close> using infer_e_appPI by auto
- show \<open>\<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile> v \<Leftarrow> \<lbrace> x1 : b[bv::=b']\<^sub>b | c[bv::=b']\<^sub>b \<rbrace>\<close> using infer_e_appPI ctx_subtype_check_v_eq subst_defs by auto
- thus \<open>atom x1 \<sharp> \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma>\<close> using fresh_replace_inside[of \<Theta> \<B> \<Gamma>' x b0 c0' \<Gamma> c0 x1 ] infer_v_wf infer_e_appPI by auto
- show \<open>\<tau>'[bv::=b']\<^sub>b[x1::=v]\<^sub>v = \<tau>\<close> using infer_e_appPI by auto
- have "atom bv \<sharp> \<Gamma>' @ (x, b0, c0') #\<^sub>\<Gamma> \<Gamma>" using infer_e_appPI by metis
- hence "atom bv \<sharp> \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma>"
- unfolding fresh_append_g fresh_GCons fresh_prod3 using \<open>atom bv \<sharp> c0 \<close> fresh_append_g by metis
- thus \<open>atom bv \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma>, \<Delta>, b', v, \<tau>)\<close> using infer_e_appPI by auto
- qed
-next
- case (infer_e_fstI \<Theta> \<B> \<Gamma>'' \<Delta> \<Phi> v z' b1 b2 c z)
- show ?case proof
- show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wf_replace_inside2(6) valid_wfC infer_e_fstI by auto
- show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using infer_e_fstI by auto
- show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile> v \<Rightarrow> \<lbrace> z' : B_pair b1 b2 | c \<rbrace>\<close> using infer_e_fstI ctx_subtype_v_eq by auto
- thus \<open>atom z \<sharp> \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma>\<close> using infer_e_fstI fresh_replace_inside[of \<Theta> \<B> \<Gamma>' x b0 c0' \<Gamma> c0 z] infer_v_wf by auto
- show \<open>atom z \<sharp> AE_fst v\<close> using infer_e_fstI by auto
- qed
-next
- case (infer_e_sndI \<Theta> \<B> \<Gamma>'' \<Delta> \<Phi> v z' b1 b2 c z)
- show ?case proof
- show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wf_replace_inside2(6) valid_wfC infer_e_sndI by auto
- show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using infer_e_sndI by auto
- show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile> v \<Rightarrow> \<lbrace> z' : B_pair b1 b2 | c \<rbrace>\<close> using infer_e_sndI ctx_subtype_v_eq by auto
- thus \<open>atom z \<sharp> \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma>\<close> using infer_e_sndI fresh_replace_inside[of \<Theta> \<B> \<Gamma>' x b0 c0' \<Gamma> c0 z] infer_v_wf by auto
- show \<open>atom z \<sharp> AE_snd v\<close> using infer_e_sndI by auto
- qed
-next
- case (infer_e_lenI \<Theta> \<B> \<Gamma>'' \<Delta> \<Phi> v z' c z)
- show ?case proof
- show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wf_replace_inside2(6) valid_wfC infer_e_lenI by auto
- show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using infer_e_lenI by auto
- show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile> v \<Rightarrow> \<lbrace> z' : B_bitvec | c \<rbrace>\<close> using infer_e_lenI ctx_subtype_v_eq by auto
- thus \<open>atom z \<sharp> \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma>\<close> using infer_e_lenI fresh_replace_inside[of \<Theta> \<B> \<Gamma>' x b0 c0' \<Gamma> c0 z] infer_v_wf by auto
- show \<open>atom z \<sharp> AE_len v\<close> using infer_e_lenI by auto
- qed
-next
- case (infer_e_mvarI \<Theta> \<B> \<Gamma>'' \<Phi> \<Delta> u \<tau>)
- show ?case proof
- show "\<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>" using wf_replace_inside2(6) valid_wfC infer_e_mvarI by auto
- thus "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma>" using infer_e_mvarI fresh_replace_inside wfD_wf by blast
- show "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> " using infer_e_mvarI by auto
- show "(u, \<tau>) \<in> setD \<Delta>" using infer_e_mvarI by auto
- qed
-next
- case (infer_e_concatI \<Theta> \<B> \<Gamma>'' \<Delta> \<Phi> v1 z1 c1 v2 z2 c2 z3)
- show ?case proof
- show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wf_replace_inside2(6) valid_wfC infer_e_concatI by auto
- thus \<open>atom z3 \<sharp> \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma>\<close> using infer_e_concatI fresh_replace_inside[of \<Theta> \<B> \<Gamma>' x b0 c0' \<Gamma> c0 z3] infer_v_wf wfX_wfY by metis
- show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using infer_e_concatI by auto
- show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile> v1 \<Rightarrow> \<lbrace> z1 : B_bitvec | c1 \<rbrace>\<close> using infer_e_concatI ctx_subtype_v_eq by auto
- show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile> v2 \<Rightarrow> \<lbrace> z2 : B_bitvec | c2 \<rbrace>\<close> using infer_e_concatI ctx_subtype_v_eq by auto
- show \<open>atom z3 \<sharp> AE_concat v1 v2\<close> using infer_e_concatI by auto
- qed
-next
- case (infer_e_splitI \<Theta> \<B> \<Gamma>'' \<Delta> \<Phi> v1 z1 c1 v2 z2 z3)
- show ?case proof
- show *:\<open> \<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wf_replace_inside2(6) valid_wfC infer_e_splitI by auto
- show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using infer_e_splitI by auto
- show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile> v1 \<Rightarrow> \<lbrace> z1 : B_bitvec | c1 \<rbrace>\<close> using infer_e_splitI ctx_subtype_v_eq by auto
- show \<open>\<Theta>; \<B>; \<Gamma>' @
- (x, b0, c0) #\<^sub>\<Gamma>
- \<Gamma> \<turnstile> v2 \<Leftarrow> \<lbrace> z2 : B_int | [ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ z2 ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e AND
- [ leq [ [ z2 ]\<^sup>v ]\<^sup>c\<^sup>e [| [ v1 ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrace>\<close>
- using infer_e_splitI ctx_subtype_check_v_eq by auto
-
- show \<open>atom z1 \<sharp> \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma>\<close> using fresh_replace_inside[of \<Theta> \<B> \<Gamma>' x b0 c0' \<Gamma> c0 z1] infer_e_splitI infer_v_wf wfX_wfY * by metis
- show \<open>atom z2 \<sharp> \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma>\<close> using fresh_replace_inside[of \<Theta> \<B> \<Gamma>' x b0 c0' \<Gamma> c0 ] infer_e_splitI infer_v_wf wfX_wfY * by metis
- show \<open>atom z3 \<sharp> \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma>\<close> using fresh_replace_inside[of \<Theta> \<B> \<Gamma>' x b0 c0' \<Gamma> c0 ] infer_e_splitI infer_v_wf wfX_wfY * by metis
- show \<open>atom z1 \<sharp> AE_split v1 v2\<close> using infer_e_splitI by auto
- show \<open>atom z2 \<sharp> AE_split v1 v2\<close> using infer_e_splitI by auto
- show \<open>atom z3 \<sharp> AE_split v1 v2\<close> using infer_e_splitI by auto
- qed
-qed
-
-lemma ctx_subtype_e_rig_eq:
- assumes "replace_in_g_subtyped \<Theta> \<B> \<Gamma>' [(x,c0)] \<Gamma>" and
- "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>' ; \<Delta> \<turnstile> e \<Rightarrow> t1"
- shows "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile> e \<Rightarrow> t1"
-proof -
- obtain b and c0' and G and G' where "\<Gamma>' = G' @ (x,b,c0')#\<^sub>\<Gamma>G \<and> \<Gamma> = G' @ (x,b,c0)#\<^sub>\<Gamma>G \<and> \<Theta>; \<B>; G'@ (x,b,c0)#\<^sub>\<Gamma>G \<Turnstile> c0'"
- using assms replace_in_g_inside_valid infer_e_wf by meson
- thus ?thesis
- using assms ctx_subtype_e_eq by presburger
-qed
-
-lemma ctx_subtype_e_rigs_eq:
- assumes "replace_in_g_subtyped \<Theta> \<B> \<Gamma>' xcs \<Gamma>" and
- "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>'; \<Delta> \<turnstile> e \<Rightarrow> t1"
- shows "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile> e \<Rightarrow> t1"
- using assms proof(induct xcs arbitrary: \<Gamma> \<Gamma>' t1 )
- case Nil
- moreover have "\<Gamma>' = \<Gamma>" using replace_in_g_subtyped_nilI
- using calculation(1) by blast
- moreover have "\<Theta>;\<B>;\<Gamma> \<turnstile> t1 \<lesssim> t1" using subtype_reflI2 Nil infer_e_t_wf by blast
- ultimately show ?case by blast
-next
- case (Cons a xcs)
-
- then obtain x and c where "a=(x,c)" by fastforce
- then obtain b and c' where bc: "Some (b, c') = lookup \<Gamma>' x \<and>
- replace_in_g_subtyped \<Theta> \<B> (replace_in_g \<Gamma>' x c) xcs \<Gamma> \<and> \<Theta>; \<B>; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f c \<and>
- x \<notin> fst ` set xcs \<and> \<Theta>; \<B>; (replace_in_g \<Gamma>' x c) \<Turnstile> c' " using replace_in_g_subtyped_elims(3)[of \<Theta> \<B> \<Gamma>' x c xcs \<Gamma>] Cons
- by (metis valid.simps)
-
- hence *: "replace_in_g_subtyped \<Theta> \<B> \<Gamma>' [(x,c)] (replace_in_g \<Gamma>' x c)" using replace_in_g_subtyped_consI
- by (meson image_iff list.distinct(1) list.set_cases replace_in_g_subtyped_nilI)
- hence t2: "\<Theta> ; \<Phi> ; \<B> ; (replace_in_g \<Gamma>' x c) ; \<Delta> \<turnstile> e \<Rightarrow> t1 " using ctx_subtype_e_rig_eq[OF * Cons(3)] by blast
- moreover have **: "replace_in_g_subtyped \<Theta> \<B> (replace_in_g \<Gamma>' x c) xcs \<Gamma>" using bc by auto
- ultimately have t2': "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile> e \<Rightarrow> t1" using Cons by blast
- thus ?case by blast
-qed
-
-section \<open>Statements\<close>
-
-lemma ctx_subtype_s_rigs:
- fixes c0::c and s::s and G'::\<Gamma> and xcs :: "(x*c) list" and css::branch_list
- shows
- "check_s \<Theta> \<Phi> \<B> G \<Delta> s t1 \<Longrightarrow> wsX G xcs \<Longrightarrow> replace_in_g_subtyped \<Theta> \<B> G xcs G' \<Longrightarrow> check_s \<Theta> \<Phi> \<B> G' \<Delta> s t1" and
- "check_branch_s \<Theta> \<Phi> \<B> G \<Delta> tid cons const v cs t1 \<Longrightarrow> wsX G xcs \<Longrightarrow> replace_in_g_subtyped \<Theta> \<B> G xcs G' \<Longrightarrow> check_branch_s \<Theta> \<Phi> \<B> G' \<Delta> tid cons const v cs t1"
- "check_branch_list \<Theta> \<Phi> \<B> G \<Delta> tid dclist v css t1 \<Longrightarrow> wsX G xcs \<Longrightarrow> replace_in_g_subtyped \<Theta> \<B> G xcs G' \<Longrightarrow> check_branch_list \<Theta> \<Phi> \<B> G' \<Delta> tid dclist v css t1"
-proof(induction arbitrary: xcs G' and xcs G' and xcs G' rule: check_s_check_branch_s_check_branch_list.inducts)
- case (check_valI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v \<tau>' \<tau>)
- hence *:"\<Theta>; \<B>; G' \<turnstile> v \<Rightarrow> \<tau>' \<and> \<Theta>; \<B>; G' \<turnstile> \<tau>' \<lesssim> \<tau>" using ctx_subtype_v_rigs_eq ctx_subtype_subtype_rigs
- by (meson check_v.simps)
- show ?case proof
- show \<open>\<Theta>; \<B>; G' \<turnstile>\<^sub>w\<^sub>f \<Delta>\<close> using check_valI wfD_rig by auto
- show \<open>\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using check_valI by auto
- show \<open>\<Theta>; \<B>; G' \<turnstile> v \<Rightarrow> \<tau>'\<close> using * by auto
- show \<open>\<Theta>; \<B>; G' \<turnstile> \<tau>' \<lesssim> \<tau>\<close> using * by auto
- qed
-next
- case (check_letI x \<Theta> \<Phi> \<B> \<Gamma> \<Delta> e \<tau> z' s b' c')
-
- show ?case proof
- have wfG: "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<and> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f G'" using infer_e_wf check_letI replace_in_g_wfG using infer_e_wf(2) by (auto simp add: freshers)
- hence "atom x \<sharp> G'" using check_letI replace_in_g_fresh replace_in_g_wfG by auto
- thus "atom x \<sharp> (\<Theta>, \<Phi>, \<B>, G', \<Delta>, e, \<tau>)" using check_letI by auto
- have "atom z' \<sharp> G'" apply(rule replace_in_g_fresh[OF check_letI(7)])
- using replace_in_g_wfG check_letI fresh_prodN infer_e_wf by metis+
- thus "atom z' \<sharp> (x, \<Theta>, \<Phi>, \<B>, G', \<Delta>, e, \<tau>, s)" using check_letI fresh_prodN by metis
-
- show "\<Theta> ; \<Phi> ; \<B> ; G' ; \<Delta> \<turnstile> e \<Rightarrow> \<lbrace> z' : b' | c' \<rbrace>"
- using check_letI ctx_subtype_e_rigs_eq by blast
- show "\<Theta> ; \<Phi> ; \<B> ; (x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\<Gamma> G' ; \<Delta> \<turnstile> s \<Leftarrow> \<tau>"
- proof(rule check_letI(5))
- have vld: "\<Theta>;\<B>;((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) \<Turnstile> c'[z'::=V_var x]\<^sub>c\<^sub>v" proof -
- have "wfG \<Theta> \<B> ((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)" using check_letI check_s_wf by metis
- hence "wfC \<Theta> \<B> ((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) (c'[z'::=V_var x]\<^sub>c\<^sub>v)" using wfC_refl subst_defs by auto
- thus ?thesis using valid_reflI[of \<Theta> \<B> x b' "c'[z'::=V_var x]\<^sub>v" \<Gamma> " c'[z'::=V_var x]\<^sub>v"] subst_defs by auto
- qed
- have xf: "x \<notin> fst ` set xcs" proof -
- have "atom ` fst ` set xcs \<subseteq> atom_dom \<Gamma>" using check_letI wsX_iff by meson
- moreover have "wfG \<Theta> \<B> \<Gamma>" using infer_e_wf check_letI by metis
- ultimately show ?thesis using fresh_def check_letI wfG_dom_supp
- using wsX_fresh by auto
- qed
- show "replace_in_g_subtyped \<Theta> \<B> ((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) ((x, c'[z'::=V_var x]\<^sub>v) # xcs) ((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\<Gamma> G')" proof -
-
- have "Some (b', c'[z'::=V_var x]\<^sub>v) = lookup ((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) x" by auto
-
- moreover have "\<Theta>; \<B>; replace_in_g ((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) x (c'[z'::=V_var x]\<^sub>v) \<Turnstile> c'[z'::=V_var x]\<^sub>v" proof -
- have "replace_in_g ((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) x (c'[z'::=V_var x]\<^sub>v) = ((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)"
- using replace_in_g.simps by presburger
- thus ?thesis using vld subst_defs by auto
- qed
-
- moreover have " replace_in_g_subtyped \<Theta> \<B> (replace_in_g ((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) x (c'[z'::=V_var x]\<^sub>v)) xcs ( ((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\<Gamma> G'))" proof -
- have "wfG \<Theta> \<B> ( ((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>))" using check_letI check_s_wf by metis
- hence "replace_in_g_subtyped \<Theta> \<B> ( ((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)) xcs ( ((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\<Gamma> G'))"
- using check_letI replace_in_g_subtyped_cons xf by meson
- moreover have "replace_in_g ((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) x (c'[z'::=V_var x]\<^sub>v) = ( ((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>))"
- using replace_in_g.simps by presburger
- ultimately show ?thesis by argo
- qed
- moreover have "\<Theta>; \<B>; (x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f c'[z'::=V_var x]\<^sub>v " using vld subst_defs by auto
- ultimately show ?thesis using replace_in_g_subtyped_consI xf replace_in_g.simps(2) by metis
- qed
-
- show " wsX ((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) ((x, c'[z'::=V_var x]\<^sub>v) # xcs)"
- using check_letI xf subst_defs by (simp add: wsX_cons)
- qed
- qed
-
-next
- case (check_branch_list_consI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist v cs \<tau> css)
- then show ?case using Typing.check_branch_list_consI by auto
-next
- case (check_branch_list_finalI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist v cs \<tau>)
- then show ?case using Typing.check_branch_list_finalI by auto
-next
- case (check_branch_s_branchI \<Theta> \<B> \<Gamma> \<Delta> \<tau> const x \<Phi> tid cons v s)
-
- have wfcons: "wfG \<Theta> \<B> ((x, b_of const, CE_val v == CE_val (V_cons tid cons (V_var x)) AND c_of const x) #\<^sub>\<Gamma> \<Gamma>)" using check_s_wf check_branch_s_branchI
- by meson
- hence wf: "wfG \<Theta> \<B> \<Gamma>" using wfG_cons by metis
-
- moreover have "atom x \<sharp> (const, G',v)" proof -
- have "atom x \<sharp> G'" using check_branch_s_branchI wf replace_in_g_fresh
- wfG_dom_supp replace_in_g_wfG by simp
- thus ?thesis using check_branch_s_branchI fresh_prodN by simp
- qed
-
- moreover have st: "\<Theta> ; \<Phi> ; \<B> ; (x, b_of const, CE_val v == CE_val(V_cons tid cons (V_var x)) AND c_of const x) #\<^sub>\<Gamma> G' ; \<Delta> \<turnstile> s \<Leftarrow> \<tau> " proof -
- have " wsX ((x, b_of const, CE_val v == CE_val(V_cons tid cons (V_var x)) AND c_of const x) #\<^sub>\<Gamma> \<Gamma>) xcs" using check_branch_s_branchI wsX_cons2 wsX_fresh wf by force
- moreover have "replace_in_g_subtyped \<Theta> \<B> ((x, b_of const, CE_val v == CE_val (V_cons tid cons (V_var x)) AND c_of const x) #\<^sub>\<Gamma> \<Gamma>) xcs ((x, b_of const, CE_val v == CE_val(V_cons tid cons (V_var x)) AND c_of const x) #\<^sub>\<Gamma> G' )"
- using replace_in_g_subtyped_cons wsX_fresh wf check_branch_s_branchI wfcons by auto
- thus ?thesis using check_branch_s_branchI calculation by meson
- qed
- moreover have wft: " wfT \<Theta> \<B> G' \<tau> " using
- check_branch_s_branchI ctx_subtype_subtype_rigs subtype_reflI2 subtype_wf by metis
- moreover have "wfD \<Theta> \<B> G' \<Delta>" using check_branch_s_branchI wfD_rig by presburger
- ultimately show ?case using
- Typing.check_branch_s_branchI
- using check_branch_s_branchI.hyps by simp
-
-next
- case (check_ifI z \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v s1 s2 \<tau>)
- hence wf:"wfG \<Theta> \<B> \<Gamma>" using check_s_wf by presburger
- show ?case proof(rule check_s_check_branch_s_check_branch_list.check_ifI)
- show \<open>atom z \<sharp> (\<Theta>, \<Phi>, \<B>, G', \<Delta>, v, s1, s2, \<tau>)\<close> using fresh_prodN replace_in_g_fresh1 wf check_ifI by auto
- show \<open>\<Theta>; \<B>; G' \<turnstile> v \<Leftarrow> \<lbrace> z : B_bool | TRUE \<rbrace>\<close> using ctx_subtype_check_v_rigs_eq check_ifI by presburger
- show \<open> \<Theta> ; \<Phi> ; \<B> ; G' ; \<Delta> \<turnstile> s1 \<Leftarrow> \<lbrace> z : b_of \<tau> | CE_val v == CE_val (V_lit L_true) IMP c_of \<tau> z \<rbrace>\<close> using check_ifI by auto
- show \<open> \<Theta> ; \<Phi> ; \<B> ; G' ; \<Delta> \<turnstile> s2 \<Leftarrow> \<lbrace> z : b_of \<tau> | CE_val v == CE_val (V_lit L_false) IMP c_of \<tau> z \<rbrace>\<close> using check_ifI by auto
- qed
-next
-
- case (check_let2I x P \<Phi> \<B> G \<Delta> t s1 \<tau> s2 )
- show ?case proof
- have "wfG P \<B> G" using check_let2I check_s_wf by metis
- show *: "P ; \<Phi> ; \<B> ; G' ; \<Delta> \<turnstile> s1 \<Leftarrow>t" using check_let2I by blast
- show "atom x \<sharp> (P, \<Phi>, \<B>, G', \<Delta>, t, s1, \<tau>)" proof -
- have "wfG P \<B> G'" using check_s_wf * by blast
- hence "atom_dom G = atom_dom G'" using check_let2I rigs_atom_dom_eq by presburger
- moreover have "atom x \<sharp> G" using check_let2I by auto
- moreover have "wfG P \<B> G" using check_s_wf * replace_in_g_wfG check_let2I by simp
- ultimately have "atom x \<sharp> G'" using wfG_dom_supp fresh_def \<open>wfG P \<B> G'\<close> by metis
- thus ?thesis using check_let2I by auto
- qed
- show "P ; \<Phi> ; \<B> ;(x, b_of t, c_of t x) #\<^sub>\<Gamma> G' ; \<Delta> \<turnstile> s2 \<Leftarrow> \<tau> " proof -
- have "wsX ((x, b_of t, c_of t x) #\<^sub>\<Gamma> G) xcs" using check_let2I wsX_cons2 wsX_fresh \<open>wfG P \<B> G\<close> by simp
- moreover have "replace_in_g_subtyped P \<B> ((x, b_of t, c_of t x) #\<^sub>\<Gamma> G) xcs ((x, b_of t, c_of t x) #\<^sub>\<Gamma> G')" proof(rule replace_in_g_subtyped_cons )
- show "replace_in_g_subtyped P \<B> G xcs G'" using check_let2I by auto
- have "atom x \<sharp> G" using check_let2I by auto
- moreover have "wfT P \<B> G t" using check_let2I check_s_wf by metis
-
- moreover have "atom x \<sharp> t" using check_let2I check_s_wf wfT_supp by auto
- ultimately show "wfG P \<B> ((x, b_of t, c_of t x) #\<^sub>\<Gamma> G)" using wfT_wf_cons b_of_c_of_eq[of x t] by auto
- show "x \<notin> fst ` set xcs" using check_let2I wsX_fresh \<open>wfG P \<B> G\<close> by simp
- qed
- ultimately show ?thesis using check_let2I by presburger
- qed
- qed
-next
- case (check_varI u \<Theta> \<Phi> \<B> \<Gamma> \<Delta> \<tau>' v \<tau> s)
- show ?case proof
- have "atom u \<sharp> G'" unfolding fresh_def
- apply(rule u_not_in_g , rule replace_in_g_wfG)
- using check_v_wf check_varI by simp+
- thus \<open>atom u \<sharp> (\<Theta>, \<Phi>, \<B>, G', \<Delta>, \<tau>', v, \<tau>)\<close> unfolding fresh_prodN using check_varI by simp
- show \<open>\<Theta>; \<B>; G' \<turnstile> v \<Leftarrow> \<tau>'\<close> using ctx_subtype_check_v_rigs_eq check_varI by auto
- show \<open> \<Theta> ; \<Phi> ; \<B> ; G' ; (u, \<tau>') #\<^sub>\<Delta> \<Delta> \<turnstile> s \<Leftarrow> \<tau>\<close> using check_varI by auto
- qed
-next
- case (check_assignI P \<Phi> \<B> G \<Delta> u \<tau> v z \<tau>')
- show ?case proof
- show \<open>P \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using check_assignI by auto
- show \<open>P ; \<B> ; G' \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using check_assignI wfD_rig by auto
- show \<open>(u, \<tau>) \<in> setD \<Delta>\<close> using check_assignI by auto
- show \<open>P ; \<B> ; G' \<turnstile> v \<Leftarrow> \<tau>\<close> using ctx_subtype_check_v_rigs_eq check_assignI by auto
- show \<open>P ; \<B> ; G' \<turnstile> \<lbrace> z : B_unit | TRUE \<rbrace> \<lesssim> \<tau>'\<close> using ctx_subtype_subtype_rigs check_assignI by auto
- qed
-next
- case (check_whileI \<Delta> G P s1 z s2 \<tau>')
- then show ?case using Typing.check_whileI
- by (meson ctx_subtype_subtype_rigs)
-next
- case (check_seqI \<Delta> G P s1 z s2 \<tau>)
- then show ?case
- using check_s_check_branch_s_check_branch_list.check_seqI by blast
-next
- case (check_caseI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist v cs \<tau> z)
- show ?case proof
- show " \<Theta> ; \<Phi> ; \<B> ; G' ; \<Delta> ; tid ; dclist ; v \<turnstile> cs \<Leftarrow> \<tau>" using check_caseI ctx_subtype_check_v_rigs_eq by auto
- show "AF_typedef tid dclist \<in> set \<Theta>" using check_caseI by auto
- show "\<Theta>; \<B>; G' \<turnstile> v \<Leftarrow> \<lbrace> z : B_id tid | TRUE \<rbrace>" using check_caseI ctx_subtype_check_v_rigs_eq by auto
- show "\<turnstile>\<^sub>w\<^sub>f \<Theta> " using check_caseI by auto
- qed
-next
- case (check_assertI x \<Theta> \<Phi> \<B> \<Gamma> \<Delta> c \<tau> s)
- show ?case proof
- have wfG: "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<and> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f G'" using check_s_wf check_assertI replace_in_g_wfG wfX_wfY by metis
- hence "atom x \<sharp> G'" using check_assertI replace_in_g_fresh replace_in_g_wfG by auto
- thus \<open>atom x \<sharp> (\<Theta>, \<Phi>, \<B>, G', \<Delta>, c, \<tau>, s)\<close> using check_assertI fresh_prodN by auto
- show \<open> \<Theta> ; \<Phi> ; \<B> ; (x, B_bool, c) #\<^sub>\<Gamma> G' ; \<Delta> \<turnstile> s \<Leftarrow> \<tau>\<close> proof(rule check_assertI(5) )
- show "wsX ((x, B_bool, c) #\<^sub>\<Gamma> \<Gamma>) xcs" using check_assertI wsX_cons3 by simp
- show "\<Theta>; \<B> \<turnstile> (x, B_bool, c) #\<^sub>\<Gamma> \<Gamma> \<langle> xcs \<rangle> \<leadsto> (x, B_bool, c) #\<^sub>\<Gamma> G'" proof(rule replace_in_g_subtyped_cons)
- show \<open> \<Theta>; \<B> \<turnstile> \<Gamma> \<langle> xcs \<rangle> \<leadsto> G'\<close> using check_assertI by auto
- show \<open> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (x, B_bool, c) #\<^sub>\<Gamma> \<Gamma> \<close> using check_assertI check_s_wf by metis
- thus \<open>x \<notin> fst ` set xcs\<close> using check_assertI wsX_fresh wfG_elims wfX_wfY by metis
- qed
- qed
- show \<open>\<Theta>; \<B>; G' \<Turnstile> c \<close> using check_assertI replace_in_g_valid by auto
- show \<open> \<Theta>; \<B>; G' \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using check_assertI wfD_rig by auto
- qed
-qed
-
-lemma replace_in_g_subtyped_empty:
- assumes "wfG \<Theta> \<B> (\<Gamma>' @ (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)"
- shows "replace_in_g_subtyped \<Theta> \<B> (replace_in_g (\<Gamma>' @ (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) x (c'[z'::=V_var x]\<^sub>c\<^sub>v)) [] (\<Gamma>' @ (x, b, c'[z'::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)"
-proof -
- have "replace_in_g (\<Gamma>' @ (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) x (c'[z'::=V_var x]\<^sub>c\<^sub>v) = (\<Gamma>' @ (x, b, c'[z'::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)"
- using assms proof(induct \<Gamma>' rule: \<Gamma>_induct)
- case GNil
- then show ?case using replace_in_g.simps by auto
- next
- case (GCons x1 b1 c1 \<Gamma>1)
- have "x \<notin> fst ` toSet ((x1,b1,c1)#\<^sub>\<Gamma>\<Gamma>1)" using GCons wfG_inside_fresh atom_dom.simps dom.simps toSet.simps append_g.simps by fast
- hence "x1 \<noteq> x" using assms wfG_inside_fresh GCons by force
- hence "((x1,b1,c1) #\<^sub>\<Gamma> (\<Gamma>1 @ (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>))[x\<longmapsto>c'[z'::=V_var x]\<^sub>c\<^sub>v] = (x1,b1,c1) #\<^sub>\<Gamma> (\<Gamma>1 @ (x, b, c'[z'::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)"
- using replace_in_g.simps GCons wfG_elims append_g.simps by metis
- thus ?case using append_g.simps by simp
- qed
- thus ?thesis using replace_in_g_subtyped_nilI by presburger
-qed
-
-lemma ctx_subtype_s:
- fixes s::s
- assumes "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>'@((x,b,c[z::=V_var x]\<^sub>c\<^sub>v)#\<^sub>\<Gamma>\<Gamma>) ; \<Delta> \<turnstile> s \<Leftarrow> \<tau>" and
- "\<Theta>; \<B>; \<Gamma> \<turnstile> \<lbrace> z' : b | c' \<rbrace> \<lesssim> \<lbrace> z : b | c \<rbrace>" and
- "atom x \<sharp> (z,z',c,c')"
- shows "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>'@(x,b,c'[z'::=V_var x]\<^sub>c\<^sub>v)#\<^sub>\<Gamma>\<Gamma> ; \<Delta> \<turnstile> s \<Leftarrow> \<tau>"
-proof -
-
- have wf: "wfG \<Theta> \<B> (\<Gamma>'@((x,b,c[z::=V_var x]\<^sub>c\<^sub>v)#\<^sub>\<Gamma>\<Gamma>))" using check_s_wf assms by meson
- hence *:"x \<notin> fst ` toSet \<Gamma>'" using wfG_inside_fresh by force
- have "wfG \<Theta> \<B> ((x,b,c[z::=V_var x]\<^sub>c\<^sub>v)#\<^sub>\<Gamma>\<Gamma>)" using wf wfG_suffix by metis
- hence xfg: "atom x \<sharp> \<Gamma>" using wfG_elims by metis
- have "x \<noteq> z'" using assms fresh_at_base fresh_prod4 by metis
- hence a2: "atom x \<sharp> c'" using assms fresh_prod4 by metis
-
- have "atom x \<sharp> (z', c', z, c, \<Gamma>)" proof -
- have "x \<noteq> z" using assms using assms fresh_at_base fresh_prod4 by metis
- hence a1 : "atom x \<sharp> c" using assms subtype_wf subtype_wf assms wfT_fresh_c xfg by meson
- thus ?thesis using a1 a2 \<open>atom x \<sharp> (z,z',c,c')\<close> fresh_prod4 fresh_Pair xfg by simp
- qed
- hence wc1:" \<Theta>; \<B>; (x, b, c'[z'::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<Turnstile> c[z::=V_var x]\<^sub>v"
- using subtype_valid assms fresh_prodN by metis
-
- have vld: "\<Theta>;\<B> ; (\<Gamma>'@(x, b, c'[z'::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) \<Turnstile> c[z::=V_var x]\<^sub>c\<^sub>v" proof -
-
- have "toSet ((x, b, c'[z'::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) \<subseteq> toSet (\<Gamma>'@(x, b, c'[z'::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)" by auto
- moreover have "wfG \<Theta> \<B> (\<Gamma>'@(x, b, c'[z'::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)" proof -
- have *:"wfT \<Theta> \<B> \<Gamma> (\<lbrace> z' : b | c' \<rbrace>)" using subtype_wf assms by meson
- moreover have "atom x \<sharp> (c',\<Gamma>)" using xfg a2 by simp
- ultimately have "wfG \<Theta> \<B> ((x, b, c'[z'::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)" using wfT_wf_cons_flip freshers by blast
- thus ?thesis using wfG_replace_inside2 check_s_wf assms by metis
- qed
- ultimately show ?thesis using wc1 valid_weakening subst_defs by metis
- qed
- hence wbc: "\<Theta>; \<B>; \<Gamma>' @ (x, b, c'[z'::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f c[z::=V_var x]\<^sub>c\<^sub>v" using valid.simps by auto
- have wbc1: "\<Theta>; \<B>; (x, b, c'[z'::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f c[z::=V_var x]\<^sub>c\<^sub>v" using wc1 valid.simps subst_defs by auto
- have "wsX (\<Gamma>'@((x,b,c[z::=V_var x]\<^sub>c\<^sub>v)#\<^sub>\<Gamma>\<Gamma>)) [(x, c'[z'::=V_var x]\<^sub>c\<^sub>v)]" proof
- show "wsX (\<Gamma>' @ (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) []" using wsX_NilI by auto
- show "atom x \<in> atom_dom (\<Gamma>' @ (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)" by simp
- show "x \<notin> fst ` set []" by auto
- qed
- moreover have "replace_in_g_subtyped \<Theta> \<B> (\<Gamma>'@((x,b,c[z::=V_var x]\<^sub>c\<^sub>v)#\<^sub>\<Gamma>\<Gamma>)) [(x, c'[z'::=V_var x]\<^sub>c\<^sub>v)] (\<Gamma>'@(x,b,c'[z'::=V_var x]\<^sub>c\<^sub>v)#\<^sub>\<Gamma>\<Gamma>)" proof
- show "Some (b, c[z::=V_var x]\<^sub>c\<^sub>v) = lookup (\<Gamma>' @ (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) x" using lookup_inside* by auto
- show "\<Theta>; \<B>; replace_in_g (\<Gamma>' @ (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) x (c'[z'::=V_var x]\<^sub>c\<^sub>v) \<Turnstile> c[z::=V_var x]\<^sub>c\<^sub>v" using vld replace_in_g_split wf by metis
- show "replace_in_g_subtyped \<Theta> \<B> (replace_in_g (\<Gamma>' @ (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) x (c'[z'::=V_var x]\<^sub>c\<^sub>v)) [] (\<Gamma>' @ (x, b, c'[z'::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)"
- using replace_in_g_subtyped_empty wf by presburger
- show "x \<notin> fst ` set []" by auto
- show "\<Theta>; \<B>; \<Gamma>' @ (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f c'[z'::=V_var x]\<^sub>c\<^sub>v"
- proof(rule wf_weakening)
- show \<open>\<Theta>; \<B>; (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f c'[z'::=[ x ]\<^sup>v]\<^sub>c\<^sub>v \<close> using wfC_cons_switch[OF wbc1] wf_weakening(6) check_s_wf assms toSet.simps by metis
- show \<open>\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>' @ (x, b, c[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<close> using wfC_cons_switch[OF wbc1] wf_weakening(6) check_s_wf assms toSet.simps by metis
- show \<open>toSet ((x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) \<subseteq> toSet (\<Gamma>' @ (x, b, c[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)\<close> using append_g.simps toSet.simps by auto
- qed
- qed
- ultimately show ?thesis using ctx_subtype_s_rigs(1)[OF assms(1)] by presburger
-qed
-
+(*<*)
+theory ContextSubtypingL
+ imports TypingL "HOL-Eisbach.Eisbach_Tools" SubstMethods
+begin
+ (*>*)
+
+declare freshers[simp del]
+
+chapter \<open>Context Subtyping Lemmas\<close>
+
+text \<open>Lemmas allowing us to replace the type of a variable in the context with a subtype
+and have the judgement remain valid. Also known as narrowing.\<close>
+
+section \<open>Replace or exchange type of variable in a context\<close>
+
+text \<open> Because the G-context is extended by the statements like let, we will need a generalised
+substitution lemma for statements.
+For this we setup a function that replaces in G (rig) for a particular x the constraint for it.
+We also define a well-formedness relation for RIGs that ensures that each new constraint
+implies the old one\<close>
+
+nominal_function replace_in_g_many :: "\<Gamma> \<Rightarrow> (x*c) list \<Rightarrow> \<Gamma>" where
+ "replace_in_g_many G xcs = List.foldr (\<lambda>(x,c) G. G[x \<longmapsto> c]) xcs G"
+ by(auto,simp add: eqvt_def replace_in_g_many_graph_aux_def)
+nominal_termination (eqvt) by lexicographic_order
+
+inductive replace_in_g_subtyped :: "\<Theta> \<Rightarrow> \<B> \<Rightarrow> \<Gamma> \<Rightarrow> (x*c) list \<Rightarrow> \<Gamma> \<Rightarrow> bool" (" _ ; _ \<turnstile> _ \<langle> _ \<rangle> \<leadsto> _" [100,50,50] 50) where
+ replace_in_g_subtyped_nilI: "\<Theta>; \<B> \<turnstile> G \<langle> [] \<rangle> \<leadsto> G"
+| replace_in_g_subtyped_consI: "\<lbrakk>
+ Some (b,c') = lookup G x ;
+ \<Theta>; \<B>; G \<turnstile>\<^sub>w\<^sub>f c ;
+ \<Theta>; \<B>; G[x\<longmapsto>c] \<Turnstile> c' ;
+ \<Theta>; \<B> \<turnstile> G[x\<longmapsto>c] \<langle> xcs \<rangle> \<leadsto> G'; x \<notin> fst ` set xcs \<rbrakk> \<Longrightarrow>
+ \<Theta>; \<B> \<turnstile> G \<langle> (x,c)#xcs \<rangle> \<leadsto> G'"
+equivariance replace_in_g_subtyped
+nominal_inductive replace_in_g_subtyped .
+
+inductive_cases replace_in_g_subtyped_elims[elim!]:
+ "\<Theta>; \<B> \<turnstile> G \<langle> [] \<rangle> \<leadsto> G'"
+ "\<Theta>; \<B> \<turnstile> ((x,b,c)#\<^sub>\<Gamma>\<Gamma> G) \<langle> acs \<rangle> \<leadsto> ((x,b,c)#\<^sub>\<Gamma>G')"
+ "\<Theta>; \<B> \<turnstile> G' \<langle> (x,c)# acs \<rangle> \<leadsto> G"
+
+lemma rigs_atom_dom_eq:
+ assumes "\<Theta>; \<B> \<turnstile> G \<langle> xcs \<rangle> \<leadsto> G'"
+ shows "atom_dom G = atom_dom G'"
+ using assms proof(induct rule: replace_in_g_subtyped.induct)
+ case (replace_in_g_subtyped_nilI G)
+ then show ?case by simp
+next
+ case (replace_in_g_subtyped_consI b c' G x \<Theta> \<B> c xcs G')
+ then show ?case using rig_dom_eq atom_dom.simps dom.simps by simp
+qed
+
+lemma replace_in_g_wfG:
+ assumes "\<Theta>; \<B> \<turnstile> G \<langle> xcs \<rangle> \<leadsto> G'" and "wfG \<Theta> \<B> G "
+ shows "wfG \<Theta> \<B> G'"
+ using assms proof(induct rule: replace_in_g_subtyped.induct)
+ case (replace_in_g_subtyped_nilI \<Theta> G)
+ then show ?case by auto
+next
+ case (replace_in_g_subtyped_consI b c' G x \<Theta> c xcs G')
+ then show ?case using valid_g_wf by auto
+qed
+
+lemma wfD_rig_single:
+ fixes \<Delta>::\<Delta> and x::x and c::c and G::\<Gamma>
+ assumes "\<Theta>; \<B>; G \<turnstile>\<^sub>w\<^sub>f \<Delta> " and "wfG \<Theta> \<B> (G[x\<longmapsto>c])"
+ shows "\<Theta>; \<B>; G[x\<longmapsto>c] \<turnstile>\<^sub>w\<^sub>f \<Delta>"
+proof(cases "atom x \<in> atom_dom G")
+ case False
+ hence "(G[x\<longmapsto>c]) = G" using assms replace_in_g_forget wfX_wfY by metis
+ then show ?thesis using assms by auto
+next
+ case True
+ then obtain G1 G2 b c' where *: "G=G1@(x,b,c')#\<^sub>\<Gamma>G2" using split_G by fastforce
+ hence **: "(G[x\<longmapsto>c]) = G1@(x,b,c)#\<^sub>\<Gamma>G2" using replace_in_g_inside wfD_wf assms wfD_wf by metis
+
+ hence "wfG \<Theta> \<B> ((x,b,c)#\<^sub>\<Gamma>G2)" using wfG_suffix assms by auto
+ hence "\<Theta>; \<B>; (x, b, TRUE) #\<^sub>\<Gamma> G2 \<turnstile>\<^sub>w\<^sub>f c" using wfG_elim2 by auto
+
+ thus ?thesis using wf_replace_inside1 assms * **
+ by (simp add: wf_replace_inside2(6))
+qed
+
+lemma wfD_rig:
+ assumes "\<Theta>; \<B> \<turnstile> G \<langle> xcs \<rangle> \<leadsto> G'" and "wfD \<Theta> \<B> G \<Delta>"
+ shows "wfD \<Theta> \<B> G' \<Delta>"
+ using assms proof(induct rule: replace_in_g_subtyped.induct)
+ case (replace_in_g_subtyped_nilI \<Theta> G)
+ then show ?case by auto
+next
+ case (replace_in_g_subtyped_consI b c' G x \<Theta> c xcs G')
+ then show ?case using wfD_rig_single valid.simps wfC_wf by auto
+qed
+
+lemma replace_in_g_fresh:
+ fixes x::x
+ assumes "\<Theta>; \<B> \<turnstile> \<Gamma> \<langle> xcs \<rangle> \<leadsto> \<Gamma>'" and "wfG \<Theta> \<B> \<Gamma>" and "wfG \<Theta> \<B> \<Gamma>'" and "atom x \<sharp> \<Gamma>"
+ shows "atom x \<sharp> \<Gamma>'"
+ using wfG_dom_supp assms fresh_def rigs_atom_dom_eq by metis
+
+lemma replace_in_g_fresh1:
+ fixes x::x
+ assumes "\<Theta>; \<B> \<turnstile> \<Gamma> \<langle> xcs \<rangle> \<leadsto> \<Gamma>'" and "wfG \<Theta> \<B> \<Gamma>" and "atom x \<sharp> \<Gamma>"
+ shows "atom x \<sharp> \<Gamma>'"
+proof -
+ have "wfG \<Theta> \<B> \<Gamma>'" using replace_in_g_wfG assms by auto
+ thus ?thesis using assms replace_in_g_fresh by metis
+qed
+
+text \<open> Wellscoping for an eXchange list\<close>
+
+inductive wsX:: "\<Gamma> \<Rightarrow> (x*c) list \<Rightarrow> bool" where
+ wsX_NilI: "wsX G []"
+| wsX_ConsI: "\<lbrakk> wsX G xcs ; atom x \<in> atom_dom G ; x \<notin> fst ` set xcs \<rbrakk> \<Longrightarrow> wsX G ((x,c)#xcs)"
+equivariance wsX
+nominal_inductive wsX .
+
+lemma wsX_if1:
+ assumes "wsX G xcs"
+ shows "(( atom ` fst ` set xcs) \<subseteq> atom_dom G) \<and> List.distinct (List.map fst xcs)"
+ using assms by(induct rule: wsX.induct,force+ )
+
+lemma wsX_if2:
+ assumes "(( atom ` fst ` set xcs) \<subseteq> atom_dom G) \<and> List.distinct (List.map fst xcs)"
+ shows "wsX G xcs"
+ using assms proof(induct xcs)
+ case Nil
+ then show ?case using wsX_NilI by fast
+next
+ case (Cons a xcs)
+ then obtain x and c where xc: "a=(x,c)" by force
+ have " wsX G xcs" proof -
+ have "distinct (map fst xcs)" using Cons by force
+ moreover have " atom ` fst ` set xcs \<subseteq> atom_dom G" using Cons by simp
+ ultimately show ?thesis using Cons by fast
+ qed
+ moreover have "atom x \<in> atom_dom G" using Cons xc
+ by simp
+ moreover have "x \<notin> fst ` set xcs" using Cons xc
+ by simp
+ ultimately show ?case using wsX_ConsI xc by blast
+qed
+
+lemma wsX_iff:
+ "wsX G xcs = ((( atom ` fst ` set xcs) \<subseteq> atom_dom G) \<and> List.distinct (List.map fst xcs))"
+ using wsX_if1 wsX_if2 by meson
+
+inductive_cases wsX_elims[elim!]:
+ "wsX G []"
+ "wsX G ((x,c)#xcs)"
+
+lemma wsX_cons:
+ assumes "wsX \<Gamma> xcs" and "x \<notin> fst ` set xcs"
+ shows "wsX ((x, b, c1) #\<^sub>\<Gamma> \<Gamma>) ((x, c2) # xcs)"
+ using assms proof(induct \<Gamma>)
+ case GNil
+ then show ?case using atom_dom.simps wsX_iff by auto
+next
+ case (GCons xbc \<Gamma>)
+ obtain x' and b' and c' where xbc: "xbc = (x',b',c')" using prod_cases3 by blast
+ then have "atom ` fst ` set xcs \<subseteq> atom_dom (xbc #\<^sub>\<Gamma> \<Gamma>) \<and> distinct (map fst xcs)"
+ using GCons.prems(1) wsX_iff by blast
+ then have "wsX ((x, b, c1) #\<^sub>\<Gamma> xbc #\<^sub>\<Gamma> \<Gamma>) xcs"
+ by (simp add: Un_commute subset_Un_eq wsX_if2)
+ then show ?case by (simp add: GCons.prems(2) wsX_ConsI)
+qed
+
+lemma wsX_cons2:
+ assumes "wsX \<Gamma> xcs" and "x \<notin> fst ` set xcs"
+ shows "wsX ((x, b, c1) #\<^sub>\<Gamma> \<Gamma>) xcs"
+ using assms proof(induct \<Gamma>)
+ case GNil
+ then show ?case using atom_dom.simps wsX_iff by auto
+next
+ case (GCons xbc \<Gamma>)
+ obtain x' and b' and c' where xbc: "xbc = (x',b',c')" using prod_cases3 by blast
+ then have "atom ` fst ` set xcs \<subseteq> atom_dom (xbc #\<^sub>\<Gamma> \<Gamma>) \<and> distinct (map fst xcs)"
+ using GCons.prems(1) wsX_iff by blast then show ?case by (simp add: Un_commute subset_Un_eq wsX_if2)
+qed
+
+lemma wsX_cons3:
+ assumes "wsX \<Gamma> xcs"
+ shows "wsX ((x, b, c1) #\<^sub>\<Gamma> \<Gamma>) xcs"
+ using assms proof(induct \<Gamma>)
+ case GNil
+ then show ?case using atom_dom.simps wsX_iff by auto
+next
+ case (GCons xbc \<Gamma>)
+ obtain x' and b' and c' where xbc: "xbc = (x',b',c')" using prod_cases3 by blast
+ then have "atom ` fst ` set xcs \<subseteq> atom_dom (xbc #\<^sub>\<Gamma> \<Gamma>) \<and> distinct (map fst xcs)"
+ using GCons.prems(1) wsX_iff by blast then show ?case by (simp add: Un_commute subset_Un_eq wsX_if2)
+qed
+
+lemma wsX_fresh:
+ assumes "wsX G xcs" and "atom x \<sharp> G" and "wfG \<Theta> \<B> G "
+ shows "x \<notin> fst ` set xcs"
+proof -
+ have "atom x \<notin> atom_dom G" using assms
+ using fresh_def wfG_dom_supp by auto
+ thus ?thesis using wsX_iff assms by blast
+qed
+
+lemma replace_in_g_dist:
+ assumes "x' \<noteq> x"
+ shows "replace_in_g ((x, b,c) #\<^sub>\<Gamma> G) x' c'' = ((x, b,c) #\<^sub>\<Gamma> (replace_in_g G x' c''))" using replace_in_g.simps assms by presburger
+
+lemma wfG_replace_inside_rig:
+ fixes c''::c
+ assumes \<open>\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f G[x'\<longmapsto>c'']\<close> \<open>\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (x, b, c) #\<^sub>\<Gamma> G \<close>
+ shows "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (x, b, c) #\<^sub>\<Gamma> G[x'\<longmapsto>c'']"
+proof(rule wfG_consI)
+
+ have "wfG \<Theta> \<B> G " using wfG_cons assms by auto
+
+ show *:"\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f G[x'\<longmapsto>c'']" using assms by auto
+ show "atom x \<sharp> G[x'\<longmapsto>c'']" using replace_in_g_fresh_single[OF *] assms wfG_elims assms by metis
+ show **:"\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b " using wfG_elim2 assms by auto
+ show "\<Theta>; \<B>; (x, b, TRUE) #\<^sub>\<Gamma> G[x'\<longmapsto>c''] \<turnstile>\<^sub>w\<^sub>f c "
+ proof(cases "atom x' \<notin> atom_dom G")
+ case True
+ hence "G = G[x'\<longmapsto>c'']" using replace_in_g_forget \<open>wfG \<Theta> \<B> G\<close> by auto
+ thus ?thesis using assms wfG_wfC by auto
+ next
+ case False
+ then obtain G1 G2 b' c' where **:"G=G1@(x',b',c')#\<^sub>\<Gamma>G2"
+ using split_G by fastforce
+ hence ***: "(G[x'\<longmapsto>c'']) = G1@(x',b',c'')#\<^sub>\<Gamma>G2"
+ using replace_in_g_inside \<open>wfG \<Theta> \<B> G \<close> by metis
+ hence "\<Theta>; \<B>; (x, b, TRUE) #\<^sub>\<Gamma> G1@(x',b',c')#\<^sub>\<Gamma>G2 \<turnstile>\<^sub>w\<^sub>f c" using * ** assms wfG_wfC by auto
+ hence "\<Theta>; \<B>; (x, b, TRUE) #\<^sub>\<Gamma> G1@(x',b',c'')#\<^sub>\<Gamma>G2 \<turnstile>\<^sub>w\<^sub>f c" using * *** wf_replace_inside assms
+ by (metis "**" append_g.simps(2) wfG_elim2 wfG_suffix)
+ thus ?thesis using ** * *** by auto
+ qed
+qed
+
+lemma replace_in_g_valid_weakening:
+ assumes "\<Theta>; \<B>; \<Gamma>[x'\<longmapsto>c''] \<Turnstile> c'" and "x' \<noteq> x" and "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (x, b, c) #\<^sub>\<Gamma> \<Gamma>[x'\<longmapsto>c'']"
+ shows "\<Theta>; \<B>; ((x, b,c) #\<^sub>\<Gamma> \<Gamma>)[x'\<longmapsto> c''] \<Turnstile> c'"
+ apply(subst replace_in_g_dist,simp add: assms,rule valid_weakening)
+ using assms by auto+
+
+lemma replace_in_g_subtyped_cons:
+ assumes "replace_in_g_subtyped \<Theta> \<B> G xcs G'" and "wfG \<Theta> \<B> ((x,b,c)#\<^sub>\<Gamma>G)"
+ shows "x \<notin> fst ` set xcs \<Longrightarrow> replace_in_g_subtyped \<Theta> \<B> ((x,b,c)#\<^sub>\<Gamma>G) xcs ((x,b,c)#\<^sub>\<Gamma>G')"
+ using assms proof(induct rule: replace_in_g_subtyped.induct)
+ case (replace_in_g_subtyped_nilI G)
+ then show ?case
+ by (simp add: replace_in_g_subtyped.replace_in_g_subtyped_nilI)
+next
+ case (replace_in_g_subtyped_consI b' c' G x' \<Theta> \<B> c'' xcs' G')
+ hence "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f G[x'\<longmapsto>c'']" using valid.simps wfC_wf by auto
+
+ show ?case proof(rule replace_in_g_subtyped.replace_in_g_subtyped_consI)
+ show " Some (b', c') = lookup ((x, b,c) #\<^sub>\<Gamma> G) x'" using lookup.simps
+ fst_conv image_iff \<Gamma>_set_intros surj_pair replace_in_g_subtyped_consI by force
+ show wbc: " \<Theta>; \<B>; (x, b, c) #\<^sub>\<Gamma> G \<turnstile>\<^sub>w\<^sub>f c'' " using wf_weakening \<open> \<Theta>; \<B>; G \<turnstile>\<^sub>w\<^sub>f c''\<close> \<open>\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (x, b, c) #\<^sub>\<Gamma> G \<close> by fastforce
+ have "x' \<noteq> x" using replace_in_g_subtyped_consI by auto
+ have wbc1: "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (x, b, c) #\<^sub>\<Gamma> G[x'\<longmapsto>c'']" proof -
+ have "(x, b, c) #\<^sub>\<Gamma> G[x'\<longmapsto>c''] = ((x, b, c) #\<^sub>\<Gamma> G)[x'\<longmapsto>c'']" using \<open>x' \<noteq> x\<close> using replace_in_g.simps by auto
+ thus ?thesis using wfG_replace_inside_rig \<open>\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f G[x'\<longmapsto>c'']\<close> \<open>\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (x, b, c) #\<^sub>\<Gamma> G \<close> by fastforce
+ qed
+ show *: "\<Theta>; \<B>; replace_in_g ((x, b,c) #\<^sub>\<Gamma> G) x' c'' \<Turnstile> c'"
+ proof -
+ have "\<Theta>; \<B>; G[x'\<longmapsto>c''] \<Turnstile> c'" using replace_in_g_subtyped_consI by auto
+ thus ?thesis using replace_in_g_valid_weakening wbc1 \<open>x'\<noteq>x\<close> by auto
+ qed
+
+ show "replace_in_g_subtyped \<Theta> \<B> (replace_in_g ((x, b,c) #\<^sub>\<Gamma> G) x' c'') xcs' ((x, b,c) #\<^sub>\<Gamma> G')"
+ using replace_in_g_subtyped_consI wbc1 by auto
+ show "x' \<notin> fst ` set xcs'"
+ using replace_in_g_subtyped_consI by linarith
+ qed
+qed
+
+lemma replace_in_g_split:
+ fixes G::\<Gamma>
+ assumes "\<Gamma> = replace_in_g \<Gamma>' x c" and "\<Gamma>' = G'@(x,b,c')#\<^sub>\<Gamma>G" and "wfG \<Theta> \<B> \<Gamma>'"
+ shows "\<Gamma> = G'@(x,b,c)#\<^sub>\<Gamma>G"
+ using assms proof(induct G' arbitrary: G \<Gamma> \<Gamma>' rule: \<Gamma>_induct)
+ case GNil
+ then show ?case by simp
+next
+ case (GCons x1 b1 c1 \<Gamma>1)
+ hence "x1 \<noteq> x"
+ using wfG_cons_fresh2[of \<Theta> \<B> x1 b1 c1 \<Gamma>1 x b ]
+ using GCons.prems(2) GCons.prems(3) append_g.simps(2) by auto
+ moreover hence *: "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (\<Gamma>1 @ (x, b, c') #\<^sub>\<Gamma> G)" using GCons append_g.simps wfG_elims by metis
+ moreover hence "replace_in_g (\<Gamma>1 @ (x, b, c') #\<^sub>\<Gamma> G) x c = \<Gamma>1 @ (x, b, c) #\<^sub>\<Gamma> G" using GCons replace_in_g_inside[OF *, of c] by auto
+
+ ultimately show ?case using replace_in_g.simps(2)[of x1 b1 c1 " \<Gamma>1 @ (x, b, c') #\<^sub>\<Gamma> G" x c] GCons
+ by (simp add: GCons.prems(1) GCons.prems(2))
+qed
+
+lemma replace_in_g_subtyped_split0:
+ fixes G::\<Gamma>
+ assumes "replace_in_g_subtyped \<Theta> \<B> \<Gamma>'[(x,c)] \<Gamma>" and "\<Gamma>' = G'@(x,b,c')#\<^sub>\<Gamma>G" and "wfG \<Theta> \<B> \<Gamma>'"
+ shows "\<Gamma> = G'@(x,b,c)#\<^sub>\<Gamma>G"
+proof -
+ have "\<Gamma> = replace_in_g \<Gamma>' x c" using assms replace_in_g_subtyped.simps
+ by (metis Pair_inject list.distinct(1) list.inject)
+ thus ?thesis using assms replace_in_g_split by blast
+qed
+
+lemma replace_in_g_subtyped_split:
+ assumes "Some (b, c') = lookup G x" and "\<Theta>; \<B>; replace_in_g G x c \<Turnstile> c'" and "wfG \<Theta> \<B> G "
+ shows "\<exists> \<Gamma> \<Gamma>'. G = \<Gamma>'@(x,b,c')#\<^sub>\<Gamma>\<Gamma> \<and> \<Theta>; \<B>; \<Gamma>'@(x,b,c)#\<^sub>\<Gamma>\<Gamma> \<Turnstile> c'"
+proof -
+ obtain \<Gamma> and \<Gamma>' where "G = \<Gamma>'@(x,b,c')#\<^sub>\<Gamma>\<Gamma>" using assms lookup_split by blast
+ moreover hence "replace_in_g G x c = \<Gamma>'@(x,b,c)#\<^sub>\<Gamma>\<Gamma>" using replace_in_g_split assms by blast
+ ultimately show ?thesis by (metis assms(2))
+qed
+
+section \<open>Validity and Subtyping\<close>
+
+lemma wfC_replace_in_g:
+ fixes c::c and c0::c
+ assumes "\<Theta>; \<B>; \<Gamma>'@(x,b,c0')#\<^sub>\<Gamma>\<Gamma> \<turnstile>\<^sub>w\<^sub>f c" and "\<Theta>; \<B>; (x,b,TRUE)#\<^sub>\<Gamma>\<Gamma> \<turnstile>\<^sub>w\<^sub>f c0"
+ shows "\<Theta>; \<B>; \<Gamma>' @ (x, b, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f c"
+ using wf_replace_inside1(2) assms by auto
+
+
+lemma ctx_subtype_valid:
+ assumes "\<Theta>; \<B>; \<Gamma>'@(x,b,c0')#\<^sub>\<Gamma>\<Gamma> \<Turnstile> c" and
+ "\<Theta>; \<B>; \<Gamma>'@(x,b,c0)#\<^sub>\<Gamma>\<Gamma> \<Turnstile> c0'"
+ shows "\<Theta>; \<B>; \<Gamma>'@(x,b,c0)#\<^sub>\<Gamma>\<Gamma> \<Turnstile> c"
+proof(rule validI)
+ show "\<Theta>; \<B>; \<Gamma>' @ (x, b, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f c" proof -
+ have "\<Theta>; \<B>; \<Gamma>'@(x,b,c0')#\<^sub>\<Gamma>\<Gamma> \<turnstile>\<^sub>w\<^sub>f c" using valid.simps assms by auto
+ moreover have "\<Theta>; \<B>; (x,b,TRUE)#\<^sub>\<Gamma>\<Gamma> \<turnstile>\<^sub>w\<^sub>f c0" proof -
+ have "wfG \<Theta> \<B> (\<Gamma>'@(x,b,c0)#\<^sub>\<Gamma>\<Gamma>)" using assms valid.simps wfC_wf by auto
+ hence "wfG \<Theta> \<B> ((x,b,c0)#\<^sub>\<Gamma>\<Gamma>)" using wfG_suffix by auto
+ thus ?thesis using wfG_wfC by auto
+ qed
+ ultimately show ?thesis using assms wfC_replace_in_g by auto
+ qed
+
+ show "\<forall>i. wfI \<Theta> (\<Gamma>' @ (x, b, c0) #\<^sub>\<Gamma> \<Gamma>) i \<and> is_satis_g i (\<Gamma>' @ (x, b, c0) #\<^sub>\<Gamma> \<Gamma>) \<longrightarrow> is_satis i c" proof(rule,rule)
+ fix i
+ assume * : "wfI \<Theta> (\<Gamma>' @ (x, b, c0) #\<^sub>\<Gamma> \<Gamma>) i \<and> is_satis_g i (\<Gamma>' @ (x, b, c0) #\<^sub>\<Gamma> \<Gamma>) "
+
+ hence "is_satis_g i (\<Gamma>'@(x, b, c0) #\<^sub>\<Gamma> \<Gamma>) \<and> wfI \<Theta> (\<Gamma>'@(x, b, c0) #\<^sub>\<Gamma> \<Gamma>) i" using is_satis_g_append wfI_suffix by metis
+ moreover hence "is_satis i c0'" using valid.simps assms by presburger
+
+ moreover have "is_satis_g i \<Gamma>'" using is_satis_g_append * by simp
+ ultimately have "is_satis_g i (\<Gamma>' @ (x, b, c0') #\<^sub>\<Gamma> \<Gamma>) " using is_satis_g_append by simp
+ moreover have "wfI \<Theta> (\<Gamma>' @ (x, b, c0') #\<^sub>\<Gamma> \<Gamma>) i" using wfI_def wfI_suffix * wfI_def wfI_replace_inside by metis
+ ultimately show "is_satis i c" using assms valid.simps by metis
+ qed
+qed
+
+lemma ctx_subtype_subtype:
+ fixes \<Gamma>::\<Gamma>
+ shows "\<Theta>; \<B>; G \<turnstile> t1 \<lesssim> t2 \<Longrightarrow> G = \<Gamma>'@(x,b0,c0')#\<^sub>\<Gamma>\<Gamma> \<Longrightarrow> \<Theta>; \<B>; \<Gamma>'@(x,b0,c0)#\<^sub>\<Gamma>\<Gamma> \<Turnstile> c0' \<Longrightarrow> \<Theta>; \<B>; \<Gamma>'@(x,b0,c0)#\<^sub>\<Gamma>\<Gamma> \<turnstile> t1 \<lesssim> t2"
+proof(nominal_induct avoiding: c0 rule: subtype.strong_induct)
+
+ case (subtype_baseI x' \<Theta> \<B> \<Gamma>'' z c z' c' b)
+ let ?\<Gamma>c0 = "\<Gamma>'@(x,b0,c0)#\<^sub>\<Gamma>\<Gamma>"
+ have wb1: "wfG \<Theta> \<B> ?\<Gamma>c0" using valid.simps wfC_wf subtype_baseI by metis
+ show ?case proof
+ show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b | c \<rbrace> \<close> using wfT_replace_inside2[OF _ wb1] subtype_baseI by metis
+ show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> z' : b | c' \<rbrace> \<close> using wfT_replace_inside2[OF _ wb1] subtype_baseI by metis
+ have "atom x' \<sharp> \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma>" using fresh_prodN subtype_baseI fresh_replace_inside wb1 subtype_wf wfX_wfY by metis
+ thus \<open>atom x' \<sharp> (\<Theta>, \<B>, \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma>, z, c , z' , c' )\<close> using subtype_baseI fresh_prodN by metis
+ have "\<Theta>; \<B>; ((x', b, c[z::=V_var x']\<^sub>v) #\<^sub>\<Gamma> \<Gamma>') @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<Turnstile> c'[z'::=V_var x']\<^sub>v " proof(rule ctx_subtype_valid)
+ show 1: \<open>\<Theta>; \<B>; ((x', b, c[z::=V_var x']\<^sub>v) #\<^sub>\<Gamma> \<Gamma>') @ (x, b0, c0') #\<^sub>\<Gamma> \<Gamma> \<Turnstile> c'[z'::=V_var x']\<^sub>v \<close>
+ using subtype_baseI append_g.simps subst_defs by metis
+ have *:"\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f ((x', b, c[z::=V_var x']\<^sub>v) #\<^sub>\<Gamma> \<Gamma>') @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> " proof(rule wfG_replace_inside2)
+ show "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f ((x', b, c[z::=V_var x']\<^sub>v) #\<^sub>\<Gamma> \<Gamma>') @ (x, b0, c0') #\<^sub>\<Gamma> \<Gamma>"
+ using * valid_wf_all wfC_wf 1 append_g.simps by metis
+ show "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (x, b0, c0) #\<^sub>\<Gamma> \<Gamma>" using wfG_suffix wb1 by auto
+ qed
+ moreover have "toSet (\<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma>) \<subseteq> toSet (((x', b, c[z::=V_var x']\<^sub>v) #\<^sub>\<Gamma> \<Gamma>') @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma>)" using toSet.simps append_g.simps by auto
+ ultimately show \<open>\<Theta>; \<B>; ((x', b, c[z::=V_var x']\<^sub>v) #\<^sub>\<Gamma> \<Gamma>') @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<Turnstile> c0' \<close> using valid_weakening subtype_baseI * by blast
+ qed
+ thus \<open>\<Theta>; \<B>; (x', b, c[z::=V_var x']\<^sub>v) #\<^sub>\<Gamma> \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<Turnstile> c'[z'::=V_var x']\<^sub>v \<close> using append_g.simps subst_defs by simp
+ qed
+qed
+
+lemma ctx_subtype_subtype_rig:
+ assumes "replace_in_g_subtyped \<Theta> \<B> \<Gamma>' [(x,c0)] \<Gamma>" and "\<Theta>; \<B>; \<Gamma>' \<turnstile> t1 \<lesssim> t2"
+ shows "\<Theta>; \<B>; \<Gamma> \<turnstile> t1 \<lesssim> t2"
+proof -
+ have wf: "wfG \<Theta> \<B> \<Gamma>'" using subtype_g_wf assms by auto
+ obtain b and c0' where "Some (b, c0') = lookup \<Gamma>' x \<and> (\<Theta>; \<B>; replace_in_g \<Gamma>' x c0 \<Turnstile> c0')" using
+ replace_in_g_subtyped.simps[of \<Theta> \<B> \<Gamma>' "[(x, c0)]" \<Gamma>] assms(1)
+
+ by (metis fst_conv list.inject list.set_intros(1) list.simps(15) not_Cons_self2 old.prod.exhaust prod.inject set_ConsD surj_pair)
+ moreover then obtain G and G' where *: "\<Gamma>' = G'@(x,b,c0')#\<^sub>\<Gamma>G \<and> \<Theta>; \<B>; G'@(x,b,c0)#\<^sub>\<Gamma>G \<Turnstile> c0'"
+ using replace_in_g_subtyped_split[of b c0' \<Gamma>' x \<Theta> \<B> c0] wf by metis
+
+ ultimately show ?thesis using ctx_subtype_subtype
+ assms(1) assms(2) replace_in_g_subtyped_split0 subtype_g_wf
+ by (metis (no_types, lifting) local.wf replace_in_g_split)
+qed
+
+text \<open> We now prove versions of the @{text "ctx_subtype"} lemmas above using @{text "replace_in_g"}. First we do case where
+the replace is just for a single variable (indicated by suffix rig) and then the general case for
+multiple replacements (indicated by suffix rigs)\<close>
+
+lemma ctx_subtype_subtype_rigs:
+ assumes "replace_in_g_subtyped \<Theta> \<B> \<Gamma>' xcs \<Gamma>" and "\<Theta>; \<B>; \<Gamma>' \<turnstile> t1 \<lesssim> t2"
+ shows "\<Theta>; \<B>; \<Gamma> \<turnstile> t1 \<lesssim> t2"
+ using assms proof(induct xcs arbitrary: \<Gamma> \<Gamma>' )
+ case Nil
+ moreover have "\<Gamma>' = \<Gamma>" using replace_in_g_subtyped_nilI
+ using calculation(1) by blast
+ ultimately show ?case by auto
+next
+ case (Cons a xcs)
+
+ then obtain x and c where "a=(x,c)" by fastforce
+ then obtain b and c' where bc: "Some (b, c') = lookup \<Gamma>' x \<and>
+ replace_in_g_subtyped \<Theta> \<B> (replace_in_g \<Gamma>' x c) xcs \<Gamma> \<and> \<Theta>; \<B>; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f c \<and>
+ x \<notin> fst ` set xcs \<and> \<Theta>; \<B>; (replace_in_g \<Gamma>' x c) \<Turnstile> c' " using replace_in_g_subtyped_elims(3)[of \<Theta> \<B> \<Gamma>' x c xcs \<Gamma>] Cons
+ by (metis valid.simps)
+
+ hence *: "replace_in_g_subtyped \<Theta> \<B> \<Gamma>' [(x,c)] (replace_in_g \<Gamma>' x c)" using replace_in_g_subtyped_consI
+ by (meson image_iff list.distinct(1) list.set_cases replace_in_g_subtyped_nilI)
+
+ hence "\<Theta>; \<B>; (replace_in_g \<Gamma>' x c) \<turnstile> t1 \<lesssim> t2"
+ using ctx_subtype_subtype_rig * assms Cons.prems(2) by auto
+
+ moreover have "replace_in_g_subtyped \<Theta> \<B> (replace_in_g \<Gamma>' x c) xcs \<Gamma>" using Cons
+ using bc by blast
+
+ ultimately show ?case using Cons by blast
+qed
+
+lemma replace_in_g_inside_valid:
+ assumes "replace_in_g_subtyped \<Theta> \<B> \<Gamma>' [(x,c0)] \<Gamma>" and "wfG \<Theta> \<B> \<Gamma>'"
+ shows "\<exists>b c0' G G'. \<Gamma>' = G' @ (x,b,c0')#\<^sub>\<Gamma>G \<and> \<Gamma> = G' @ (x,b,c0)#\<^sub>\<Gamma>G \<and> \<Theta>; \<B>; G'@ (x,b,c0)#\<^sub>\<Gamma>G \<Turnstile> c0'"
+proof -
+ obtain b and c0' where bc: "Some (b, c0') = lookup \<Gamma>' x \<and> \<Theta>; \<B>; replace_in_g \<Gamma>' x c0 \<Turnstile> c0'" using
+ replace_in_g_subtyped.simps[of \<Theta> \<B> \<Gamma>' "[(x, c0)]" \<Gamma>] assms(1)
+ by (metis fst_conv list.inject list.set_intros(1) list.simps(15) not_Cons_self2 old.prod.exhaust prod.inject set_ConsD surj_pair)
+ then obtain G and G' where *: "\<Gamma>' = G'@(x,b,c0')#\<^sub>\<Gamma>G \<and> \<Theta>; \<B>; G'@(x,b,c0)#\<^sub>\<Gamma>G \<Turnstile> c0'" using replace_in_g_subtyped_split[of b c0' \<Gamma>' x \<Theta> \<B> c0] assms
+ by metis
+ thus ?thesis using replace_in_g_inside bc
+ using assms(1) assms(2) by blast
+qed
+
+lemma replace_in_g_valid:
+ assumes "\<Theta>; \<B> \<turnstile> G \<langle> xcs \<rangle> \<leadsto> G'" and "\<Theta>; \<B>; G \<Turnstile> c "
+ shows \<open>\<Theta>; \<B>; G' \<Turnstile> c \<close>
+ using assms proof(induct rule: replace_in_g_subtyped.inducts)
+ case (replace_in_g_subtyped_nilI \<Theta> \<B> G)
+ then show ?case by auto
+next
+ case (replace_in_g_subtyped_consI b c1 G x \<Theta> \<B> c2 xcs G')
+ hence "\<Theta>; \<B>; G[x\<longmapsto>c2] \<Turnstile> c"
+ by (metis ctx_subtype_valid replace_in_g_split replace_in_g_subtyped_split valid_g_wf)
+ then show ?case using replace_in_g_subtyped_consI by auto
+qed
+
+section \<open>Literals\<close>
+
+section \<open>Values\<close>
+
+lemma lookup_inside_unique_b[simp]:
+ assumes "\<Theta> ; B \<turnstile>\<^sub>w\<^sub>f (\<Gamma>'@(x,b0,c0)#\<^sub>\<Gamma>\<Gamma>)" and "\<Theta> ; B \<turnstile>\<^sub>w\<^sub>f (\<Gamma>'@(x,b0,c0')#\<^sub>\<Gamma>\<Gamma>)"
+ and "Some (b, c) = lookup (\<Gamma>' @ (x, b0, c0') #\<^sub>\<Gamma> \<Gamma>) y" and "Some (b0,c0) = lookup (\<Gamma>'@((x,b0,c0))#\<^sub>\<Gamma>\<Gamma>) x" and "x=y"
+ shows "b = b0"
+ by (metis assms(2) assms(3) assms(5) lookup_inside_wf old.prod.exhaust option.inject prod.inject)
+
+lemma ctx_subtype_v_aux:
+ fixes v::v
+ assumes "\<Theta>; \<B>; \<Gamma>'@((x,b0,c0')#\<^sub>\<Gamma>\<Gamma>) \<turnstile> v \<Rightarrow> t1" and "\<Theta>; \<B>; \<Gamma>'@(x,b0,c0)#\<^sub>\<Gamma>\<Gamma> \<Turnstile> c0'"
+ shows "\<Theta>; \<B>; \<Gamma>'@((x,b0,c0)#\<^sub>\<Gamma>\<Gamma>) \<turnstile> v \<Rightarrow> t1"
+ using assms proof(nominal_induct "\<Gamma>'@((x,b0,c0')#\<^sub>\<Gamma>\<Gamma>)" v t1 avoiding: c0 rule: infer_v.strong_induct)
+ case (infer_v_varI \<Theta> \<B> b c xa z)
+ have wf:\<open> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<close> using wfG_inside_valid2 infer_v_varI by metis
+ have xf1:\<open>atom z \<sharp> xa\<close> using infer_v_varI by metis
+ have xf2: \<open>atom z \<sharp> (\<Theta>, \<B>, \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma>)\<close> apply( fresh_mth add: infer_v_varI )
+ using fresh_def infer_v_varI wfG_supp fresh_append_g fresh_GCons fresh_prodN by metis+
+ show ?case proof (cases "x=xa")
+ case True
+ moreover have "b = b0" using infer_v_varI True by simp
+ moreover hence \<open>Some (b, c0) = lookup (\<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma>) xa\<close> using lookup_inside_wf[OF wf] infer_v_varI True by auto
+ ultimately show ?thesis using wf xf1 xf2 Typing.infer_v_varI by metis
+ next
+ case False
+ moreover hence \<open>Some (b, c) = lookup (\<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma>) xa\<close> using lookup_inside2 infer_v_varI by metis
+ ultimately show ?thesis using wf xf1 xf2 Typing.infer_v_varI by simp
+ qed
+next
+ case (infer_v_litI \<Theta> \<B> l \<tau>)
+ thus ?case using Typing.infer_v_litI wfG_inside_valid2 by simp
+next
+ case (infer_v_pairI z v1 v2 \<Theta> \<B> t1' t2' c0)
+ show ?case proof
+ show "atom z \<sharp> (v1, v2)" using infer_v_pairI fresh_Pair by simp
+ show "atom z \<sharp> (\<Theta>, \<B>, \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma>)" apply( fresh_mth add: infer_v_pairI )
+ using fresh_def infer_v_pairI wfG_supp fresh_append_g fresh_GCons fresh_prodN by metis+
+ show "\<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile> v1 \<Rightarrow> t1'" using infer_v_pairI by simp
+ show "\<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile> v2 \<Rightarrow> t2'" using infer_v_pairI by simp
+ qed
+next
+ case (infer_v_consI s dclist \<Theta> dc tc \<B> v tv z)
+ show ?case proof
+ show \<open>AF_typedef s dclist \<in> set \<Theta>\<close> using infer_v_consI by auto
+ show \<open>(dc, tc) \<in> set dclist\<close> using infer_v_consI by auto
+ show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile> v \<Rightarrow> tv\<close> using infer_v_consI by auto
+ show \<open>\<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile> tv \<lesssim> tc\<close> using infer_v_consI ctx_subtype_subtype by auto
+ show \<open>atom z \<sharp> v\<close> using infer_v_consI by auto
+ show \<open>atom z \<sharp> (\<Theta>, \<B>, \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma>)\<close> apply( fresh_mth add: infer_v_consI )
+ using fresh_def infer_v_consI wfG_supp fresh_append_g fresh_GCons fresh_prodN by metis+
+ qed
+next
+ case (infer_v_conspI s bv dclist \<Theta> dc tc \<B> v tv b z)
+ show ?case proof
+ show \<open>AF_typedef_poly s bv dclist \<in> set \<Theta>\<close> using infer_v_conspI by auto
+ show \<open>(dc, tc) \<in> set dclist\<close> using infer_v_conspI by auto
+ show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile> v \<Rightarrow> tv\<close> using infer_v_conspI by auto
+ show \<open>\<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile> tv \<lesssim> tc[bv::=b]\<^sub>\<tau>\<^sub>b\<close> using infer_v_conspI ctx_subtype_subtype by auto
+ show \<open>atom z \<sharp> (\<Theta>, \<B>, \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma>, v, b)\<close> apply( fresh_mth add: infer_v_conspI )
+ using fresh_def infer_v_conspI wfG_supp fresh_append_g fresh_GCons fresh_prodN by metis+
+ show \<open>atom bv \<sharp> (\<Theta>, \<B>, \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma>, v, b)\<close> apply( fresh_mth add: infer_v_conspI )
+ using fresh_def infer_v_conspI wfG_supp fresh_append_g fresh_GCons fresh_prodN by metis+
+ show \<open> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b \<close> using infer_v_conspI by auto
+ qed
+qed
+
+lemma ctx_subtype_v:
+ fixes v::v
+ assumes "\<Theta>; \<B>; \<Gamma>'@((x,b0,c0')#\<^sub>\<Gamma>\<Gamma>) \<turnstile> v \<Rightarrow> t1" and "\<Theta>; \<B>; \<Gamma>'@(x,b0,c0)#\<^sub>\<Gamma>\<Gamma> \<Turnstile> c0'"
+ shows "\<exists>t2. \<Theta>; \<B>; \<Gamma>'@((x,b0,c0)#\<^sub>\<Gamma>\<Gamma>) \<turnstile> v \<Rightarrow> t2 \<and> \<Theta>; \<B>; \<Gamma>'@((x,b0,c0)#\<^sub>\<Gamma>\<Gamma>) \<turnstile> t2 \<lesssim> t1"
+proof -
+
+ have "\<Theta>; \<B>; \<Gamma>'@((x,b0,c0)#\<^sub>\<Gamma>\<Gamma>) \<turnstile> v \<Rightarrow> t1 " using ctx_subtype_v_aux assms by auto
+ moreover hence "\<Theta>; \<B>; \<Gamma>'@((x,b0,c0)#\<^sub>\<Gamma>\<Gamma>) \<turnstile> t1 \<lesssim> t1" using subtype_reflI2 infer_v_wf by simp
+ ultimately show ?thesis by auto
+qed
+
+lemma ctx_subtype_v_eq:
+ fixes v::v
+ assumes
+ "\<Theta>; \<B>; \<Gamma>'@((x,b0,c0')#\<^sub>\<Gamma>\<Gamma>) \<turnstile> v \<Rightarrow> t1" and
+ " \<Theta>; \<B>; \<Gamma>'@(x,b0,c0)#\<^sub>\<Gamma>\<Gamma> \<Turnstile> c0'"
+ shows "\<Theta>; \<B>; \<Gamma>'@((x,b0,c0)#\<^sub>\<Gamma>\<Gamma>) \<turnstile> v \<Rightarrow> t1"
+proof -
+ obtain t1' where "\<Theta>; \<B>; \<Gamma>'@((x,b0,c0)#\<^sub>\<Gamma>\<Gamma>) \<turnstile> v \<Rightarrow> t1'" using ctx_subtype_v assms by metis
+ moreover have "replace_in_g (\<Gamma>'@((x,b0,c0')#\<^sub>\<Gamma>\<Gamma>)) x c0 = \<Gamma>'@((x,b0,c0)#\<^sub>\<Gamma>\<Gamma>)" using replace_in_g_inside infer_v_wf assms by metis
+ ultimately show ?thesis using infer_v_uniqueness_rig assms by metis
+qed
+
+lemma ctx_subtype_check_v_eq:
+ assumes "\<Theta>; \<B>; \<Gamma>'@((x,b0,c0')#\<^sub>\<Gamma>\<Gamma>) \<turnstile> v \<Leftarrow> t1" and " \<Theta>; \<B>; \<Gamma>'@(x,b0,c0)#\<^sub>\<Gamma>\<Gamma> \<Turnstile> c0'"
+ shows "\<Theta>; \<B>; \<Gamma>'@((x,b0,c0)#\<^sub>\<Gamma>\<Gamma>) \<turnstile> v \<Leftarrow> t1"
+proof -
+ obtain t2 where t2: "\<Theta>; \<B>; \<Gamma>'@((x,b0,c0')#\<^sub>\<Gamma>\<Gamma>) \<turnstile> v \<Rightarrow> t2 \<and> \<Theta>; \<B>; \<Gamma>'@((x,b0,c0')#\<^sub>\<Gamma>\<Gamma>) \<turnstile> t2 \<lesssim> t1"
+ using check_v_elims assms by blast
+ hence t3: "\<Theta>; \<B>; \<Gamma>'@((x,b0,c0)#\<^sub>\<Gamma>\<Gamma>) \<turnstile> v \<Rightarrow> t2"
+ using assms ctx_subtype_v_eq by blast
+
+ have "\<Theta>; \<B>; \<Gamma>'@((x,b0,c0)#\<^sub>\<Gamma>\<Gamma>) \<turnstile> v \<Rightarrow> t2" using t3 by auto
+ moreover have " \<Theta>; \<B>; \<Gamma>'@((x,b0,c0)#\<^sub>\<Gamma>\<Gamma>) \<turnstile> t2 \<lesssim> t1" proof -
+
+ have " \<Theta>; \<B>; \<Gamma>'@((x,b0,c0')#\<^sub>\<Gamma>\<Gamma>) \<turnstile> t2 \<lesssim> t1" using t2 by auto
+ thus ?thesis using subtype_trans
+ using assms(2) ctx_subtype_subtype by blast
+ qed
+ ultimately show ?thesis using check_v.intros by presburger
+qed
+
+text \<open> Basically the same as @{text "ctx_subtype_v_eq"} but in a different form\<close>
+lemma ctx_subtype_v_rig_eq:
+ fixes v::v
+ assumes "replace_in_g_subtyped \<Theta> \<B> \<Gamma>' [(x,c0)] \<Gamma>" and
+ "\<Theta>; \<B>; \<Gamma>' \<turnstile> v \<Rightarrow> t1"
+ shows "\<Theta>; \<B>; \<Gamma> \<turnstile> v \<Rightarrow> t1"
+proof -
+ obtain b and c0' and G and G' where "\<Gamma>' = G' @ (x,b,c0')#\<^sub>\<Gamma>G \<and> \<Gamma> = G' @ (x,b,c0)#\<^sub>\<Gamma>G \<and> \<Theta>; \<B>; G'@ (x,b,c0)#\<^sub>\<Gamma>G \<Turnstile> c0'"
+ using assms replace_in_g_inside_valid infer_v_wf by metis
+ thus ?thesis using ctx_subtype_v_eq[of \<Theta> \<B> G' x b c0' G v t1 c0] assms by simp
+qed
+
+lemma ctx_subtype_v_rigs_eq:
+ fixes v::v
+ assumes "replace_in_g_subtyped \<Theta> \<B> \<Gamma>' xcs \<Gamma>" and
+ "\<Theta>; \<B>; \<Gamma>' \<turnstile> v \<Rightarrow> t1"
+ shows "\<Theta>; \<B>; \<Gamma> \<turnstile> v \<Rightarrow> t1"
+ using assms proof(induct xcs arbitrary: \<Gamma> \<Gamma>' t1 )
+ case Nil
+ then show ?case by auto
+next
+ case (Cons a xcs)
+ then obtain x and c where "a=(x,c)" by fastforce
+
+ then obtain b and c' where bc: "Some (b, c') = lookup \<Gamma>' x \<and>
+ replace_in_g_subtyped \<Theta> \<B> (replace_in_g \<Gamma>' x c) xcs \<Gamma> \<and> \<Theta>; \<B>; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f c \<and>
+ x \<notin> fst ` set xcs \<and> \<Theta>; \<B>; (replace_in_g \<Gamma>' x c) \<Turnstile> c' "
+ using replace_in_g_subtyped_elims(3)[of \<Theta> \<B> \<Gamma>' x c xcs \<Gamma>] Cons by (metis valid.simps)
+
+ hence *: "replace_in_g_subtyped \<Theta> \<B> \<Gamma>' [(x,c)] (replace_in_g \<Gamma>' x c)" using replace_in_g_subtyped_consI
+ by (meson image_iff list.distinct(1) list.set_cases replace_in_g_subtyped_nilI)
+ hence t2: "\<Theta>; \<B>; (replace_in_g \<Gamma>' x c) \<turnstile> v \<Rightarrow> t1 " using ctx_subtype_v_rig_eq[OF * Cons(3)] by blast
+ moreover have **: "replace_in_g_subtyped \<Theta> \<B> (replace_in_g \<Gamma>' x c) xcs \<Gamma>" using bc by auto
+ ultimately have t2': "\<Theta>; \<B>; \<Gamma> \<turnstile> v \<Rightarrow> t1" using Cons by blast
+ thus ?case by blast
+qed
+
+lemma ctx_subtype_check_v_rigs_eq:
+ assumes "replace_in_g_subtyped \<Theta> \<B> \<Gamma>' xcs \<Gamma>" and
+ "\<Theta>; \<B>; \<Gamma>' \<turnstile> v \<Leftarrow> t1"
+ shows "\<Theta>; \<B>; \<Gamma> \<turnstile> v \<Leftarrow> t1"
+proof -
+ obtain t2 where "\<Theta>; \<B>; \<Gamma>' \<turnstile> v \<Rightarrow> t2 \<and> \<Theta>; \<B>; \<Gamma>' \<turnstile> t2 \<lesssim> t1" using check_v_elims assms by fast
+ hence "\<Theta>; \<B>; \<Gamma> \<turnstile> v \<Rightarrow> t2 \<and> \<Theta>; \<B>; \<Gamma> \<turnstile> t2 \<lesssim> t1" using ctx_subtype_v_rigs_eq ctx_subtype_subtype_rigs
+ using assms(1) by blast
+ thus ?thesis
+ using check_v_subtypeI by blast
+qed
+
+section \<open>Expressions\<close>
+
+lemma valid_wfC:
+ fixes c0::c
+ assumes "\<Theta>; \<B>; \<Gamma>'@(x,b0,c0)#\<^sub>\<Gamma>\<Gamma> \<Turnstile> c0'"
+ shows "\<Theta>; \<B>; (x, b0, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f c0"
+ using wfG_elim2 valid.simps wfG_suffix
+ using assms valid_g_wf by metis
+
+lemma ctx_subtype_e_eq:
+ fixes G::\<Gamma>
+ assumes
+ "\<Theta> ; \<Phi> ; \<B> ; G ; \<Delta> \<turnstile> e \<Rightarrow> t1" and "G = \<Gamma>'@((x,b0,c0')#\<^sub>\<Gamma>\<Gamma>)"
+ "\<Theta>; \<B>; \<Gamma>'@(x,b0,c0)#\<^sub>\<Gamma>\<Gamma> \<Turnstile> c0'"
+ shows "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>'@((x,b0,c0)#\<^sub>\<Gamma>\<Gamma>) ; \<Delta> \<turnstile> e \<Rightarrow> t1"
+ using assms proof(nominal_induct t1 avoiding: c0 rule: infer_e.strong_induct)
+ case (infer_e_valI \<Theta> \<B> \<Gamma>'' \<Delta> \<Phi> v \<tau>)
+ show ?case proof
+ show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wf_replace_inside2(6) valid_wfC infer_e_valI by auto
+ show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using infer_e_valI by auto
+ show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile> v \<Rightarrow> \<tau>\<close> using infer_e_valI ctx_subtype_v_eq by auto
+ qed
+next
+ case (infer_e_plusI \<Theta> \<B> \<Gamma>'' \<Delta> \<Phi> v1 z1 c1 v2 z2 c2 z3)
+ show ?case proof
+ show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wf_replace_inside2(6) valid_wfC infer_e_plusI by auto
+ show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using infer_e_plusI by auto
+ show *:\<open> \<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile> v1 \<Rightarrow> \<lbrace> z1 : B_int | c1 \<rbrace>\<close> using infer_e_plusI ctx_subtype_v_eq by auto
+ show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile> v2 \<Rightarrow> \<lbrace> z2 : B_int | c2 \<rbrace>\<close> using infer_e_plusI ctx_subtype_v_eq by auto
+ show \<open>atom z3 \<sharp> AE_op Plus v1 v2\<close> using infer_e_plusI by auto
+ show \<open>atom z3 \<sharp> \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma>\<close> using * infer_e_plusI fresh_replace_inside infer_v_wf by metis
+ qed
+next
+ case (infer_e_leqI \<Theta> \<B> \<Gamma>'' \<Delta> \<Phi> v1 z1 c1 v2 z2 c2 z3)
+ show ?case proof
+ show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wf_replace_inside2(6) valid_wfC infer_e_leqI by auto
+ show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using infer_e_leqI by auto
+ show *:\<open> \<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile> v1 \<Rightarrow> \<lbrace> z1 : B_int | c1 \<rbrace>\<close> using infer_e_leqI ctx_subtype_v_eq by auto
+ show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile> v2 \<Rightarrow> \<lbrace> z2 : B_int | c2 \<rbrace>\<close> using infer_e_leqI ctx_subtype_v_eq by auto
+ show \<open>atom z3 \<sharp> AE_op LEq v1 v2\<close> using infer_e_leqI by auto
+ show \<open>atom z3 \<sharp> \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma>\<close> using * infer_e_leqI fresh_replace_inside infer_v_wf by metis
+ qed
+next
+ case (infer_e_eqI \<Theta> \<B> \<Gamma>'' \<Delta> \<Phi> v1 z1 bb c1 v2 z2 c2 z3)
+ show ?case proof
+ show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wf_replace_inside2(6) valid_wfC infer_e_eqI by auto
+ show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using infer_e_eqI by auto
+ show *:\<open> \<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile> v1 \<Rightarrow> \<lbrace> z1 : bb | c1 \<rbrace>\<close> using infer_e_eqI ctx_subtype_v_eq by auto
+ show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile> v2 \<Rightarrow> \<lbrace> z2 : bb | c2 \<rbrace>\<close> using infer_e_eqI ctx_subtype_v_eq by auto
+ show \<open>atom z3 \<sharp> AE_op Eq v1 v2\<close> using infer_e_eqI by auto
+ show \<open>atom z3 \<sharp> \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma>\<close> using * infer_e_eqI fresh_replace_inside infer_v_wf by metis
+ show "bb \<in> {B_bool, B_int, B_unit}" using infer_e_eqI by auto
+ qed
+next
+ case (infer_e_appI \<Theta> \<B> \<Gamma>'' \<Delta> \<Phi> f x' b c \<tau>' s' v \<tau>)
+ show ?case proof
+ show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wf_replace_inside2(6) valid_wfC infer_e_appI by auto
+ show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using infer_e_appI by auto
+ show \<open>Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x' b c \<tau>' s'))) = lookup_fun \<Phi> f\<close> using infer_e_appI by auto
+ show \<open>\<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile> v \<Leftarrow> \<lbrace> x' : b | c \<rbrace>\<close> using infer_e_appI ctx_subtype_check_v_eq by auto
+ thus \<open>atom x' \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma>, \<Delta>, v, \<tau>)\<close>
+ using infer_e_appI fresh_replace_inside[of \<Theta> \<B> \<Gamma>' x b0 c0' \<Gamma> c0 x'] infer_v_wf by auto
+ show \<open>\<tau>'[x'::=v]\<^sub>v = \<tau>\<close> using infer_e_appI by auto
+ qed
+next
+ case (infer_e_appPI \<Theta> \<B> \<Gamma>1 \<Delta> \<Phi> b' f bv x1 b c \<tau>' s' v \<tau>)
+ show ?case proof
+ show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wf_replace_inside2(6) valid_wfC infer_e_appPI by auto
+ show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using infer_e_appPI by auto
+ show \<open> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b' \<close> using infer_e_appPI by auto
+ show \<open>Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x1 b c \<tau>' s'))) = lookup_fun \<Phi> f\<close> using infer_e_appPI by auto
+ show \<open>\<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile> v \<Leftarrow> \<lbrace> x1 : b[bv::=b']\<^sub>b | c[bv::=b']\<^sub>b \<rbrace>\<close> using infer_e_appPI ctx_subtype_check_v_eq subst_defs by auto
+ thus \<open>atom x1 \<sharp> \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma>\<close> using fresh_replace_inside[of \<Theta> \<B> \<Gamma>' x b0 c0' \<Gamma> c0 x1 ] infer_v_wf infer_e_appPI by auto
+ show \<open>\<tau>'[bv::=b']\<^sub>b[x1::=v]\<^sub>v = \<tau>\<close> using infer_e_appPI by auto
+ have "atom bv \<sharp> \<Gamma>' @ (x, b0, c0') #\<^sub>\<Gamma> \<Gamma>" using infer_e_appPI by metis
+ hence "atom bv \<sharp> \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma>"
+ unfolding fresh_append_g fresh_GCons fresh_prod3 using \<open>atom bv \<sharp> c0 \<close> fresh_append_g by metis
+ thus \<open>atom bv \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma>, \<Delta>, b', v, \<tau>)\<close> using infer_e_appPI by auto
+ qed
+next
+ case (infer_e_fstI \<Theta> \<B> \<Gamma>'' \<Delta> \<Phi> v z' b1 b2 c z)
+ show ?case proof
+ show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wf_replace_inside2(6) valid_wfC infer_e_fstI by auto
+ show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using infer_e_fstI by auto
+ show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile> v \<Rightarrow> \<lbrace> z' : B_pair b1 b2 | c \<rbrace>\<close> using infer_e_fstI ctx_subtype_v_eq by auto
+ thus \<open>atom z \<sharp> \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma>\<close> using infer_e_fstI fresh_replace_inside[of \<Theta> \<B> \<Gamma>' x b0 c0' \<Gamma> c0 z] infer_v_wf by auto
+ show \<open>atom z \<sharp> AE_fst v\<close> using infer_e_fstI by auto
+ qed
+next
+ case (infer_e_sndI \<Theta> \<B> \<Gamma>'' \<Delta> \<Phi> v z' b1 b2 c z)
+ show ?case proof
+ show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wf_replace_inside2(6) valid_wfC infer_e_sndI by auto
+ show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using infer_e_sndI by auto
+ show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile> v \<Rightarrow> \<lbrace> z' : B_pair b1 b2 | c \<rbrace>\<close> using infer_e_sndI ctx_subtype_v_eq by auto
+ thus \<open>atom z \<sharp> \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma>\<close> using infer_e_sndI fresh_replace_inside[of \<Theta> \<B> \<Gamma>' x b0 c0' \<Gamma> c0 z] infer_v_wf by auto
+ show \<open>atom z \<sharp> AE_snd v\<close> using infer_e_sndI by auto
+ qed
+next
+ case (infer_e_lenI \<Theta> \<B> \<Gamma>'' \<Delta> \<Phi> v z' c z)
+ show ?case proof
+ show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wf_replace_inside2(6) valid_wfC infer_e_lenI by auto
+ show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using infer_e_lenI by auto
+ show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile> v \<Rightarrow> \<lbrace> z' : B_bitvec | c \<rbrace>\<close> using infer_e_lenI ctx_subtype_v_eq by auto
+ thus \<open>atom z \<sharp> \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma>\<close> using infer_e_lenI fresh_replace_inside[of \<Theta> \<B> \<Gamma>' x b0 c0' \<Gamma> c0 z] infer_v_wf by auto
+ show \<open>atom z \<sharp> AE_len v\<close> using infer_e_lenI by auto
+ qed
+next
+ case (infer_e_mvarI \<Theta> \<B> \<Gamma>'' \<Phi> \<Delta> u \<tau>)
+ show ?case proof
+ show "\<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>" using wf_replace_inside2(6) valid_wfC infer_e_mvarI by auto
+ thus "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma>" using infer_e_mvarI fresh_replace_inside wfD_wf by blast
+ show "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> " using infer_e_mvarI by auto
+ show "(u, \<tau>) \<in> setD \<Delta>" using infer_e_mvarI by auto
+ qed
+next
+ case (infer_e_concatI \<Theta> \<B> \<Gamma>'' \<Delta> \<Phi> v1 z1 c1 v2 z2 c2 z3)
+ show ?case proof
+ show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wf_replace_inside2(6) valid_wfC infer_e_concatI by auto
+ thus \<open>atom z3 \<sharp> \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma>\<close> using infer_e_concatI fresh_replace_inside[of \<Theta> \<B> \<Gamma>' x b0 c0' \<Gamma> c0 z3] infer_v_wf wfX_wfY by metis
+ show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using infer_e_concatI by auto
+ show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile> v1 \<Rightarrow> \<lbrace> z1 : B_bitvec | c1 \<rbrace>\<close> using infer_e_concatI ctx_subtype_v_eq by auto
+ show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile> v2 \<Rightarrow> \<lbrace> z2 : B_bitvec | c2 \<rbrace>\<close> using infer_e_concatI ctx_subtype_v_eq by auto
+ show \<open>atom z3 \<sharp> AE_concat v1 v2\<close> using infer_e_concatI by auto
+ qed
+next
+ case (infer_e_splitI \<Theta> \<B> \<Gamma>'' \<Delta> \<Phi> v1 z1 c1 v2 z2 z3)
+ show ?case proof
+ show *:\<open> \<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wf_replace_inside2(6) valid_wfC infer_e_splitI by auto
+ show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using infer_e_splitI by auto
+ show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> \<turnstile> v1 \<Rightarrow> \<lbrace> z1 : B_bitvec | c1 \<rbrace>\<close> using infer_e_splitI ctx_subtype_v_eq by auto
+ show \<open>\<Theta>; \<B>; \<Gamma>' @
+ (x, b0, c0) #\<^sub>\<Gamma>
+ \<Gamma> \<turnstile> v2 \<Leftarrow> \<lbrace> z2 : B_int | [ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ z2 ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e AND
+ [ leq [ [ z2 ]\<^sup>v ]\<^sup>c\<^sup>e [| [ v1 ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrace>\<close>
+ using infer_e_splitI ctx_subtype_check_v_eq by auto
+
+ show \<open>atom z1 \<sharp> \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma>\<close> using fresh_replace_inside[of \<Theta> \<B> \<Gamma>' x b0 c0' \<Gamma> c0 z1] infer_e_splitI infer_v_wf wfX_wfY * by metis
+ show \<open>atom z2 \<sharp> \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma>\<close> using fresh_replace_inside[of \<Theta> \<B> \<Gamma>' x b0 c0' \<Gamma> c0 ] infer_e_splitI infer_v_wf wfX_wfY * by metis
+ show \<open>atom z3 \<sharp> \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma>\<close> using fresh_replace_inside[of \<Theta> \<B> \<Gamma>' x b0 c0' \<Gamma> c0 ] infer_e_splitI infer_v_wf wfX_wfY * by metis
+ show \<open>atom z1 \<sharp> AE_split v1 v2\<close> using infer_e_splitI by auto
+ show \<open>atom z2 \<sharp> AE_split v1 v2\<close> using infer_e_splitI by auto
+ show \<open>atom z3 \<sharp> AE_split v1 v2\<close> using infer_e_splitI by auto
+ qed
+qed
+
+lemma ctx_subtype_e_rig_eq:
+ assumes "replace_in_g_subtyped \<Theta> \<B> \<Gamma>' [(x,c0)] \<Gamma>" and
+ "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>' ; \<Delta> \<turnstile> e \<Rightarrow> t1"
+ shows "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile> e \<Rightarrow> t1"
+proof -
+ obtain b and c0' and G and G' where "\<Gamma>' = G' @ (x,b,c0')#\<^sub>\<Gamma>G \<and> \<Gamma> = G' @ (x,b,c0)#\<^sub>\<Gamma>G \<and> \<Theta>; \<B>; G'@ (x,b,c0)#\<^sub>\<Gamma>G \<Turnstile> c0'"
+ using assms replace_in_g_inside_valid infer_e_wf by meson
+ thus ?thesis
+ using assms ctx_subtype_e_eq by presburger
+qed
+
+lemma ctx_subtype_e_rigs_eq:
+ assumes "replace_in_g_subtyped \<Theta> \<B> \<Gamma>' xcs \<Gamma>" and
+ "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>'; \<Delta> \<turnstile> e \<Rightarrow> t1"
+ shows "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile> e \<Rightarrow> t1"
+ using assms proof(induct xcs arbitrary: \<Gamma> \<Gamma>' t1 )
+ case Nil
+ moreover have "\<Gamma>' = \<Gamma>" using replace_in_g_subtyped_nilI
+ using calculation(1) by blast
+ moreover have "\<Theta>;\<B>;\<Gamma> \<turnstile> t1 \<lesssim> t1" using subtype_reflI2 Nil infer_e_t_wf by blast
+ ultimately show ?case by blast
+next
+ case (Cons a xcs)
+
+ then obtain x and c where "a=(x,c)" by fastforce
+ then obtain b and c' where bc: "Some (b, c') = lookup \<Gamma>' x \<and>
+ replace_in_g_subtyped \<Theta> \<B> (replace_in_g \<Gamma>' x c) xcs \<Gamma> \<and> \<Theta>; \<B>; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f c \<and>
+ x \<notin> fst ` set xcs \<and> \<Theta>; \<B>; (replace_in_g \<Gamma>' x c) \<Turnstile> c' " using replace_in_g_subtyped_elims(3)[of \<Theta> \<B> \<Gamma>' x c xcs \<Gamma>] Cons
+ by (metis valid.simps)
+
+ hence *: "replace_in_g_subtyped \<Theta> \<B> \<Gamma>' [(x,c)] (replace_in_g \<Gamma>' x c)" using replace_in_g_subtyped_consI
+ by (meson image_iff list.distinct(1) list.set_cases replace_in_g_subtyped_nilI)
+ hence t2: "\<Theta> ; \<Phi> ; \<B> ; (replace_in_g \<Gamma>' x c) ; \<Delta> \<turnstile> e \<Rightarrow> t1 " using ctx_subtype_e_rig_eq[OF * Cons(3)] by blast
+ moreover have **: "replace_in_g_subtyped \<Theta> \<B> (replace_in_g \<Gamma>' x c) xcs \<Gamma>" using bc by auto
+ ultimately have t2': "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile> e \<Rightarrow> t1" using Cons by blast
+ thus ?case by blast
+qed
+
+section \<open>Statements\<close>
+
+lemma ctx_subtype_s_rigs:
+ fixes c0::c and s::s and G'::\<Gamma> and xcs :: "(x*c) list" and css::branch_list
+ shows
+ "check_s \<Theta> \<Phi> \<B> G \<Delta> s t1 \<Longrightarrow> wsX G xcs \<Longrightarrow> replace_in_g_subtyped \<Theta> \<B> G xcs G' \<Longrightarrow> check_s \<Theta> \<Phi> \<B> G' \<Delta> s t1" and
+ "check_branch_s \<Theta> \<Phi> \<B> G \<Delta> tid cons const v cs t1 \<Longrightarrow> wsX G xcs \<Longrightarrow> replace_in_g_subtyped \<Theta> \<B> G xcs G' \<Longrightarrow> check_branch_s \<Theta> \<Phi> \<B> G' \<Delta> tid cons const v cs t1"
+ "check_branch_list \<Theta> \<Phi> \<B> G \<Delta> tid dclist v css t1 \<Longrightarrow> wsX G xcs \<Longrightarrow> replace_in_g_subtyped \<Theta> \<B> G xcs G' \<Longrightarrow> check_branch_list \<Theta> \<Phi> \<B> G' \<Delta> tid dclist v css t1"
+proof(induction arbitrary: xcs G' and xcs G' and xcs G' rule: check_s_check_branch_s_check_branch_list.inducts)
+ case (check_valI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v \<tau>' \<tau>)
+ hence *:"\<Theta>; \<B>; G' \<turnstile> v \<Rightarrow> \<tau>' \<and> \<Theta>; \<B>; G' \<turnstile> \<tau>' \<lesssim> \<tau>" using ctx_subtype_v_rigs_eq ctx_subtype_subtype_rigs
+ by (meson check_v.simps)
+ show ?case proof
+ show \<open>\<Theta>; \<B>; G' \<turnstile>\<^sub>w\<^sub>f \<Delta>\<close> using check_valI wfD_rig by auto
+ show \<open>\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using check_valI by auto
+ show \<open>\<Theta>; \<B>; G' \<turnstile> v \<Rightarrow> \<tau>'\<close> using * by auto
+ show \<open>\<Theta>; \<B>; G' \<turnstile> \<tau>' \<lesssim> \<tau>\<close> using * by auto
+ qed
+next
+ case (check_letI x \<Theta> \<Phi> \<B> \<Gamma> \<Delta> e \<tau> z' s b' c')
+
+ show ?case proof
+ have wfG: "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<and> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f G'" using infer_e_wf check_letI replace_in_g_wfG using infer_e_wf(2) by (auto simp add: freshers)
+ hence "atom x \<sharp> G'" using check_letI replace_in_g_fresh replace_in_g_wfG by auto
+ thus "atom x \<sharp> (\<Theta>, \<Phi>, \<B>, G', \<Delta>, e, \<tau>)" using check_letI by auto
+ have "atom z' \<sharp> G'" apply(rule replace_in_g_fresh[OF check_letI(7)])
+ using replace_in_g_wfG check_letI fresh_prodN infer_e_wf by metis+
+ thus "atom z' \<sharp> (x, \<Theta>, \<Phi>, \<B>, G', \<Delta>, e, \<tau>, s)" using check_letI fresh_prodN by metis
+
+ show "\<Theta> ; \<Phi> ; \<B> ; G' ; \<Delta> \<turnstile> e \<Rightarrow> \<lbrace> z' : b' | c' \<rbrace>"
+ using check_letI ctx_subtype_e_rigs_eq by blast
+ show "\<Theta> ; \<Phi> ; \<B> ; (x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\<Gamma> G' ; \<Delta> \<turnstile> s \<Leftarrow> \<tau>"
+ proof(rule check_letI(5))
+ have vld: "\<Theta>;\<B>;((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) \<Turnstile> c'[z'::=V_var x]\<^sub>c\<^sub>v" proof -
+ have "wfG \<Theta> \<B> ((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)" using check_letI check_s_wf by metis
+ hence "wfC \<Theta> \<B> ((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) (c'[z'::=V_var x]\<^sub>c\<^sub>v)" using wfC_refl subst_defs by auto
+ thus ?thesis using valid_reflI[of \<Theta> \<B> x b' "c'[z'::=V_var x]\<^sub>v" \<Gamma> " c'[z'::=V_var x]\<^sub>v"] subst_defs by auto
+ qed
+ have xf: "x \<notin> fst ` set xcs" proof -
+ have "atom ` fst ` set xcs \<subseteq> atom_dom \<Gamma>" using check_letI wsX_iff by meson
+ moreover have "wfG \<Theta> \<B> \<Gamma>" using infer_e_wf check_letI by metis
+ ultimately show ?thesis using fresh_def check_letI wfG_dom_supp
+ using wsX_fresh by auto
+ qed
+ show "replace_in_g_subtyped \<Theta> \<B> ((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) ((x, c'[z'::=V_var x]\<^sub>v) # xcs) ((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\<Gamma> G')" proof -
+
+ have "Some (b', c'[z'::=V_var x]\<^sub>v) = lookup ((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) x" by auto
+
+ moreover have "\<Theta>; \<B>; replace_in_g ((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) x (c'[z'::=V_var x]\<^sub>v) \<Turnstile> c'[z'::=V_var x]\<^sub>v" proof -
+ have "replace_in_g ((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) x (c'[z'::=V_var x]\<^sub>v) = ((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)"
+ using replace_in_g.simps by presburger
+ thus ?thesis using vld subst_defs by auto
+ qed
+
+ moreover have " replace_in_g_subtyped \<Theta> \<B> (replace_in_g ((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) x (c'[z'::=V_var x]\<^sub>v)) xcs ( ((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\<Gamma> G'))" proof -
+ have "wfG \<Theta> \<B> ( ((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>))" using check_letI check_s_wf by metis
+ hence "replace_in_g_subtyped \<Theta> \<B> ( ((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)) xcs ( ((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\<Gamma> G'))"
+ using check_letI replace_in_g_subtyped_cons xf by meson
+ moreover have "replace_in_g ((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) x (c'[z'::=V_var x]\<^sub>v) = ( ((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>))"
+ using replace_in_g.simps by presburger
+ ultimately show ?thesis by argo
+ qed
+ moreover have "\<Theta>; \<B>; (x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f c'[z'::=V_var x]\<^sub>v " using vld subst_defs by auto
+ ultimately show ?thesis using replace_in_g_subtyped_consI xf replace_in_g.simps(2) by metis
+ qed
+
+ show " wsX ((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) ((x, c'[z'::=V_var x]\<^sub>v) # xcs)"
+ using check_letI xf subst_defs by (simp add: wsX_cons)
+ qed
+ qed
+
+next
+ case (check_branch_list_consI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist v cs \<tau> css)
+ then show ?case using Typing.check_branch_list_consI by auto
+next
+ case (check_branch_list_finalI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist v cs \<tau>)
+ then show ?case using Typing.check_branch_list_finalI by auto
+next
+ case (check_branch_s_branchI \<Theta> \<B> \<Gamma> \<Delta> \<tau> const x \<Phi> tid cons v s)
+
+ have wfcons: "wfG \<Theta> \<B> ((x, b_of const, CE_val v == CE_val (V_cons tid cons (V_var x)) AND c_of const x) #\<^sub>\<Gamma> \<Gamma>)" using check_s_wf check_branch_s_branchI
+ by meson
+ hence wf: "wfG \<Theta> \<B> \<Gamma>" using wfG_cons by metis
+
+ moreover have "atom x \<sharp> (const, G',v)" proof -
+ have "atom x \<sharp> G'" using check_branch_s_branchI wf replace_in_g_fresh
+ wfG_dom_supp replace_in_g_wfG by simp
+ thus ?thesis using check_branch_s_branchI fresh_prodN by simp
+ qed
+
+ moreover have st: "\<Theta> ; \<Phi> ; \<B> ; (x, b_of const, CE_val v == CE_val(V_cons tid cons (V_var x)) AND c_of const x) #\<^sub>\<Gamma> G' ; \<Delta> \<turnstile> s \<Leftarrow> \<tau> " proof -
+ have " wsX ((x, b_of const, CE_val v == CE_val(V_cons tid cons (V_var x)) AND c_of const x) #\<^sub>\<Gamma> \<Gamma>) xcs" using check_branch_s_branchI wsX_cons2 wsX_fresh wf by force
+ moreover have "replace_in_g_subtyped \<Theta> \<B> ((x, b_of const, CE_val v == CE_val (V_cons tid cons (V_var x)) AND c_of const x) #\<^sub>\<Gamma> \<Gamma>) xcs ((x, b_of const, CE_val v == CE_val(V_cons tid cons (V_var x)) AND c_of const x) #\<^sub>\<Gamma> G' )"
+ using replace_in_g_subtyped_cons wsX_fresh wf check_branch_s_branchI wfcons by auto
+ thus ?thesis using check_branch_s_branchI calculation by meson
+ qed
+ moreover have wft: " wfT \<Theta> \<B> G' \<tau> " using
+ check_branch_s_branchI ctx_subtype_subtype_rigs subtype_reflI2 subtype_wf by metis
+ moreover have "wfD \<Theta> \<B> G' \<Delta>" using check_branch_s_branchI wfD_rig by presburger
+ ultimately show ?case using
+ Typing.check_branch_s_branchI
+ using check_branch_s_branchI.hyps by simp
+
+next
+ case (check_ifI z \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v s1 s2 \<tau>)
+ hence wf:"wfG \<Theta> \<B> \<Gamma>" using check_s_wf by presburger
+ show ?case proof(rule check_s_check_branch_s_check_branch_list.check_ifI)
+ show \<open>atom z \<sharp> (\<Theta>, \<Phi>, \<B>, G', \<Delta>, v, s1, s2, \<tau>)\<close> using fresh_prodN replace_in_g_fresh1 wf check_ifI by auto
+ show \<open>\<Theta>; \<B>; G' \<turnstile> v \<Leftarrow> \<lbrace> z : B_bool | TRUE \<rbrace>\<close> using ctx_subtype_check_v_rigs_eq check_ifI by presburger
+ show \<open> \<Theta> ; \<Phi> ; \<B> ; G' ; \<Delta> \<turnstile> s1 \<Leftarrow> \<lbrace> z : b_of \<tau> | CE_val v == CE_val (V_lit L_true) IMP c_of \<tau> z \<rbrace>\<close> using check_ifI by auto
+ show \<open> \<Theta> ; \<Phi> ; \<B> ; G' ; \<Delta> \<turnstile> s2 \<Leftarrow> \<lbrace> z : b_of \<tau> | CE_val v == CE_val (V_lit L_false) IMP c_of \<tau> z \<rbrace>\<close> using check_ifI by auto
+ qed
+next
+
+ case (check_let2I x P \<Phi> \<B> G \<Delta> t s1 \<tau> s2 )
+ show ?case proof
+ have "wfG P \<B> G" using check_let2I check_s_wf by metis
+ show *: "P ; \<Phi> ; \<B> ; G' ; \<Delta> \<turnstile> s1 \<Leftarrow>t" using check_let2I by blast
+ show "atom x \<sharp> (P, \<Phi>, \<B>, G', \<Delta>, t, s1, \<tau>)" proof -
+ have "wfG P \<B> G'" using check_s_wf * by blast
+ hence "atom_dom G = atom_dom G'" using check_let2I rigs_atom_dom_eq by presburger
+ moreover have "atom x \<sharp> G" using check_let2I by auto
+ moreover have "wfG P \<B> G" using check_s_wf * replace_in_g_wfG check_let2I by simp
+ ultimately have "atom x \<sharp> G'" using wfG_dom_supp fresh_def \<open>wfG P \<B> G'\<close> by metis
+ thus ?thesis using check_let2I by auto
+ qed
+ show "P ; \<Phi> ; \<B> ;(x, b_of t, c_of t x) #\<^sub>\<Gamma> G' ; \<Delta> \<turnstile> s2 \<Leftarrow> \<tau> " proof -
+ have "wsX ((x, b_of t, c_of t x) #\<^sub>\<Gamma> G) xcs" using check_let2I wsX_cons2 wsX_fresh \<open>wfG P \<B> G\<close> by simp
+ moreover have "replace_in_g_subtyped P \<B> ((x, b_of t, c_of t x) #\<^sub>\<Gamma> G) xcs ((x, b_of t, c_of t x) #\<^sub>\<Gamma> G')" proof(rule replace_in_g_subtyped_cons )
+ show "replace_in_g_subtyped P \<B> G xcs G'" using check_let2I by auto
+ have "atom x \<sharp> G" using check_let2I by auto
+ moreover have "wfT P \<B> G t" using check_let2I check_s_wf by metis
+
+ moreover have "atom x \<sharp> t" using check_let2I check_s_wf wfT_supp by auto
+ ultimately show "wfG P \<B> ((x, b_of t, c_of t x) #\<^sub>\<Gamma> G)" using wfT_wf_cons b_of_c_of_eq[of x t] by auto
+ show "x \<notin> fst ` set xcs" using check_let2I wsX_fresh \<open>wfG P \<B> G\<close> by simp
+ qed
+ ultimately show ?thesis using check_let2I by presburger
+ qed
+ qed
+next
+ case (check_varI u \<Theta> \<Phi> \<B> \<Gamma> \<Delta> \<tau>' v \<tau> s)
+ show ?case proof
+ have "atom u \<sharp> G'" unfolding fresh_def
+ apply(rule u_not_in_g , rule replace_in_g_wfG)
+ using check_v_wf check_varI by simp+
+ thus \<open>atom u \<sharp> (\<Theta>, \<Phi>, \<B>, G', \<Delta>, \<tau>', v, \<tau>)\<close> unfolding fresh_prodN using check_varI by simp
+ show \<open>\<Theta>; \<B>; G' \<turnstile> v \<Leftarrow> \<tau>'\<close> using ctx_subtype_check_v_rigs_eq check_varI by auto
+ show \<open> \<Theta> ; \<Phi> ; \<B> ; G' ; (u, \<tau>') #\<^sub>\<Delta> \<Delta> \<turnstile> s \<Leftarrow> \<tau>\<close> using check_varI by auto
+ qed
+next
+ case (check_assignI P \<Phi> \<B> G \<Delta> u \<tau> v z \<tau>')
+ show ?case proof
+ show \<open>P \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using check_assignI by auto
+ show \<open>P ; \<B> ; G' \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using check_assignI wfD_rig by auto
+ show \<open>(u, \<tau>) \<in> setD \<Delta>\<close> using check_assignI by auto
+ show \<open>P ; \<B> ; G' \<turnstile> v \<Leftarrow> \<tau>\<close> using ctx_subtype_check_v_rigs_eq check_assignI by auto
+ show \<open>P ; \<B> ; G' \<turnstile> \<lbrace> z : B_unit | TRUE \<rbrace> \<lesssim> \<tau>'\<close> using ctx_subtype_subtype_rigs check_assignI by auto
+ qed
+next
+ case (check_whileI \<Delta> G P s1 z s2 \<tau>')
+ then show ?case using Typing.check_whileI
+ by (meson ctx_subtype_subtype_rigs)
+next
+ case (check_seqI \<Delta> G P s1 z s2 \<tau>)
+ then show ?case
+ using check_s_check_branch_s_check_branch_list.check_seqI by blast
+next
+ case (check_caseI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist v cs \<tau> z)
+ show ?case proof
+ show " \<Theta> ; \<Phi> ; \<B> ; G' ; \<Delta> ; tid ; dclist ; v \<turnstile> cs \<Leftarrow> \<tau>" using check_caseI ctx_subtype_check_v_rigs_eq by auto
+ show "AF_typedef tid dclist \<in> set \<Theta>" using check_caseI by auto
+ show "\<Theta>; \<B>; G' \<turnstile> v \<Leftarrow> \<lbrace> z : B_id tid | TRUE \<rbrace>" using check_caseI ctx_subtype_check_v_rigs_eq by auto
+ show "\<turnstile>\<^sub>w\<^sub>f \<Theta> " using check_caseI by auto
+ qed
+next
+ case (check_assertI x \<Theta> \<Phi> \<B> \<Gamma> \<Delta> c \<tau> s)
+ show ?case proof
+ have wfG: "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<and> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f G'" using check_s_wf check_assertI replace_in_g_wfG wfX_wfY by metis
+ hence "atom x \<sharp> G'" using check_assertI replace_in_g_fresh replace_in_g_wfG by auto
+ thus \<open>atom x \<sharp> (\<Theta>, \<Phi>, \<B>, G', \<Delta>, c, \<tau>, s)\<close> using check_assertI fresh_prodN by auto
+ show \<open> \<Theta> ; \<Phi> ; \<B> ; (x, B_bool, c) #\<^sub>\<Gamma> G' ; \<Delta> \<turnstile> s \<Leftarrow> \<tau>\<close> proof(rule check_assertI(5) )
+ show "wsX ((x, B_bool, c) #\<^sub>\<Gamma> \<Gamma>) xcs" using check_assertI wsX_cons3 by simp
+ show "\<Theta>; \<B> \<turnstile> (x, B_bool, c) #\<^sub>\<Gamma> \<Gamma> \<langle> xcs \<rangle> \<leadsto> (x, B_bool, c) #\<^sub>\<Gamma> G'" proof(rule replace_in_g_subtyped_cons)
+ show \<open> \<Theta>; \<B> \<turnstile> \<Gamma> \<langle> xcs \<rangle> \<leadsto> G'\<close> using check_assertI by auto
+ show \<open> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (x, B_bool, c) #\<^sub>\<Gamma> \<Gamma> \<close> using check_assertI check_s_wf by metis
+ thus \<open>x \<notin> fst ` set xcs\<close> using check_assertI wsX_fresh wfG_elims wfX_wfY by metis
+ qed
+ qed
+ show \<open>\<Theta>; \<B>; G' \<Turnstile> c \<close> using check_assertI replace_in_g_valid by auto
+ show \<open> \<Theta>; \<B>; G' \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using check_assertI wfD_rig by auto
+ qed
+qed
+
+lemma replace_in_g_subtyped_empty:
+ assumes "wfG \<Theta> \<B> (\<Gamma>' @ (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)"
+ shows "replace_in_g_subtyped \<Theta> \<B> (replace_in_g (\<Gamma>' @ (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) x (c'[z'::=V_var x]\<^sub>c\<^sub>v)) [] (\<Gamma>' @ (x, b, c'[z'::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)"
+proof -
+ have "replace_in_g (\<Gamma>' @ (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) x (c'[z'::=V_var x]\<^sub>c\<^sub>v) = (\<Gamma>' @ (x, b, c'[z'::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)"
+ using assms proof(induct \<Gamma>' rule: \<Gamma>_induct)
+ case GNil
+ then show ?case using replace_in_g.simps by auto
+ next
+ case (GCons x1 b1 c1 \<Gamma>1)
+ have "x \<notin> fst ` toSet ((x1,b1,c1)#\<^sub>\<Gamma>\<Gamma>1)" using GCons wfG_inside_fresh atom_dom.simps dom.simps toSet.simps append_g.simps by fast
+ hence "x1 \<noteq> x" using assms wfG_inside_fresh GCons by force
+ hence "((x1,b1,c1) #\<^sub>\<Gamma> (\<Gamma>1 @ (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>))[x\<longmapsto>c'[z'::=V_var x]\<^sub>c\<^sub>v] = (x1,b1,c1) #\<^sub>\<Gamma> (\<Gamma>1 @ (x, b, c'[z'::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)"
+ using replace_in_g.simps GCons wfG_elims append_g.simps by metis
+ thus ?case using append_g.simps by simp
+ qed
+ thus ?thesis using replace_in_g_subtyped_nilI by presburger
+qed
+
+lemma ctx_subtype_s:
+ fixes s::s
+ assumes "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>'@((x,b,c[z::=V_var x]\<^sub>c\<^sub>v)#\<^sub>\<Gamma>\<Gamma>) ; \<Delta> \<turnstile> s \<Leftarrow> \<tau>" and
+ "\<Theta>; \<B>; \<Gamma> \<turnstile> \<lbrace> z' : b | c' \<rbrace> \<lesssim> \<lbrace> z : b | c \<rbrace>" and
+ "atom x \<sharp> (z,z',c,c')"
+ shows "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>'@(x,b,c'[z'::=V_var x]\<^sub>c\<^sub>v)#\<^sub>\<Gamma>\<Gamma> ; \<Delta> \<turnstile> s \<Leftarrow> \<tau>"
+proof -
+
+ have wf: "wfG \<Theta> \<B> (\<Gamma>'@((x,b,c[z::=V_var x]\<^sub>c\<^sub>v)#\<^sub>\<Gamma>\<Gamma>))" using check_s_wf assms by meson
+ hence *:"x \<notin> fst ` toSet \<Gamma>'" using wfG_inside_fresh by force
+ have "wfG \<Theta> \<B> ((x,b,c[z::=V_var x]\<^sub>c\<^sub>v)#\<^sub>\<Gamma>\<Gamma>)" using wf wfG_suffix by metis
+ hence xfg: "atom x \<sharp> \<Gamma>" using wfG_elims by metis
+ have "x \<noteq> z'" using assms fresh_at_base fresh_prod4 by metis
+ hence a2: "atom x \<sharp> c'" using assms fresh_prod4 by metis
+
+ have "atom x \<sharp> (z', c', z, c, \<Gamma>)" proof -
+ have "x \<noteq> z" using assms using assms fresh_at_base fresh_prod4 by metis
+ hence a1 : "atom x \<sharp> c" using assms subtype_wf subtype_wf assms wfT_fresh_c xfg by meson
+ thus ?thesis using a1 a2 \<open>atom x \<sharp> (z,z',c,c')\<close> fresh_prod4 fresh_Pair xfg by simp
+ qed
+ hence wc1:" \<Theta>; \<B>; (x, b, c'[z'::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<Turnstile> c[z::=V_var x]\<^sub>v"
+ using subtype_valid assms fresh_prodN by metis
+
+ have vld: "\<Theta>;\<B> ; (\<Gamma>'@(x, b, c'[z'::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) \<Turnstile> c[z::=V_var x]\<^sub>c\<^sub>v" proof -
+
+ have "toSet ((x, b, c'[z'::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) \<subseteq> toSet (\<Gamma>'@(x, b, c'[z'::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)" by auto
+ moreover have "wfG \<Theta> \<B> (\<Gamma>'@(x, b, c'[z'::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)" proof -
+ have *:"wfT \<Theta> \<B> \<Gamma> (\<lbrace> z' : b | c' \<rbrace>)" using subtype_wf assms by meson
+ moreover have "atom x \<sharp> (c',\<Gamma>)" using xfg a2 by simp
+ ultimately have "wfG \<Theta> \<B> ((x, b, c'[z'::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)" using wfT_wf_cons_flip freshers by blast
+ thus ?thesis using wfG_replace_inside2 check_s_wf assms by metis
+ qed
+ ultimately show ?thesis using wc1 valid_weakening subst_defs by metis
+ qed
+ hence wbc: "\<Theta>; \<B>; \<Gamma>' @ (x, b, c'[z'::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f c[z::=V_var x]\<^sub>c\<^sub>v" using valid.simps by auto
+ have wbc1: "\<Theta>; \<B>; (x, b, c'[z'::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f c[z::=V_var x]\<^sub>c\<^sub>v" using wc1 valid.simps subst_defs by auto
+ have "wsX (\<Gamma>'@((x,b,c[z::=V_var x]\<^sub>c\<^sub>v)#\<^sub>\<Gamma>\<Gamma>)) [(x, c'[z'::=V_var x]\<^sub>c\<^sub>v)]" proof
+ show "wsX (\<Gamma>' @ (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) []" using wsX_NilI by auto
+ show "atom x \<in> atom_dom (\<Gamma>' @ (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)" by simp
+ show "x \<notin> fst ` set []" by auto
+ qed
+ moreover have "replace_in_g_subtyped \<Theta> \<B> (\<Gamma>'@((x,b,c[z::=V_var x]\<^sub>c\<^sub>v)#\<^sub>\<Gamma>\<Gamma>)) [(x, c'[z'::=V_var x]\<^sub>c\<^sub>v)] (\<Gamma>'@(x,b,c'[z'::=V_var x]\<^sub>c\<^sub>v)#\<^sub>\<Gamma>\<Gamma>)" proof
+ show "Some (b, c[z::=V_var x]\<^sub>c\<^sub>v) = lookup (\<Gamma>' @ (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) x" using lookup_inside* by auto
+ show "\<Theta>; \<B>; replace_in_g (\<Gamma>' @ (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) x (c'[z'::=V_var x]\<^sub>c\<^sub>v) \<Turnstile> c[z::=V_var x]\<^sub>c\<^sub>v" using vld replace_in_g_split wf by metis
+ show "replace_in_g_subtyped \<Theta> \<B> (replace_in_g (\<Gamma>' @ (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) x (c'[z'::=V_var x]\<^sub>c\<^sub>v)) [] (\<Gamma>' @ (x, b, c'[z'::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)"
+ using replace_in_g_subtyped_empty wf by presburger
+ show "x \<notin> fst ` set []" by auto
+ show "\<Theta>; \<B>; \<Gamma>' @ (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f c'[z'::=V_var x]\<^sub>c\<^sub>v"
+ proof(rule wf_weakening)
+ show \<open>\<Theta>; \<B>; (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f c'[z'::=[ x ]\<^sup>v]\<^sub>c\<^sub>v \<close> using wfC_cons_switch[OF wbc1] wf_weakening(6) check_s_wf assms toSet.simps by metis
+ show \<open>\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>' @ (x, b, c[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<close> using wfC_cons_switch[OF wbc1] wf_weakening(6) check_s_wf assms toSet.simps by metis
+ show \<open>toSet ((x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) \<subseteq> toSet (\<Gamma>' @ (x, b, c[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)\<close> using append_g.simps toSet.simps by auto
+ qed
+ qed
+ ultimately show ?thesis using ctx_subtype_s_rigs(1)[OF assms(1)] by presburger
+qed
+
end
\ No newline at end of file
diff --git a/thys/MiniSail/IVSubst.thy b/thys/MiniSail/IVSubst.thy
--- a/thys/MiniSail/IVSubst.thy
+++ b/thys/MiniSail/IVSubst.thy
@@ -1,1446 +1,1446 @@
-(*<*)
-theory IVSubst
- imports Syntax
-begin
- (*>*)
-
-chapter \<open>Immutable Variable Substitution\<close>
-
-text \<open>Substitution involving immutable variables. We define a class and instances for all
-of the term forms\<close>
-
-section \<open>Class\<close>
-
-class has_subst_v = fs +
- fixes subst_v :: "'a::fs \<Rightarrow> x \<Rightarrow> v \<Rightarrow> 'a::fs" ("_[_::=_]\<^sub>v" [1000,50,50] 1000)
- assumes fresh_subst_v_if: "y \<sharp> (subst_v a x v) \<longleftrightarrow> (atom x \<sharp> a \<and> y \<sharp> a) \<or> (y \<sharp> v \<and> (y \<sharp> a \<or> y = atom x))"
- and forget_subst_v[simp]: "atom x \<sharp> a \<Longrightarrow> subst_v a x v = a"
- and subst_v_id[simp]: "subst_v a x (V_var x) = a"
- and eqvt[simp,eqvt]: "(p::perm) \<bullet> (subst_v a x v) = (subst_v (p \<bullet> a) (p \<bullet>x) (p \<bullet>v))"
- and flip_subst_v[simp]: "atom x \<sharp> c \<Longrightarrow> ((x \<leftrightarrow> z) \<bullet> c) = c[z::=[x]\<^sup>v]\<^sub>v"
- and subst_v_simple_commute[simp]: "atom x \<sharp> c \<Longrightarrow>(c[z::=[x]\<^sup>v]\<^sub>v)[x::=b]\<^sub>v = c[z::=b]\<^sub>v"
-begin
-
-lemma subst_v_flip_eq_one:
- fixes z1::x and z2::x and x1::x and x2::x
- assumes "[[atom z1]]lst. c1 = [[atom z2]]lst. c2"
- and "atom x1 \<sharp> (z1,z2,c1,c2)"
- shows "(c1[z1::=[x1]\<^sup>v]\<^sub>v) = (c2[z2::=[x1]\<^sup>v]\<^sub>v)"
-proof -
- have "(c1[z1::=[x1]\<^sup>v]\<^sub>v) = (x1 \<leftrightarrow> z1) \<bullet> c1" using assms flip_subst_v by auto
- moreover have "(c2[z2::=[x1]\<^sup>v]\<^sub>v) = (x1 \<leftrightarrow> z2) \<bullet> c2" using assms flip_subst_v by auto
- ultimately show ?thesis using Abs1_eq_iff_all(3)[of z1 c1 z2 c2 z1] assms
- by (metis Abs1_eq_iff_fresh(3) flip_commute)
-qed
-
-lemma subst_v_flip_eq_two:
- fixes z1::x and z2::x and x1::x and x2::x
- assumes "[[atom z1]]lst. c1 = [[atom z2]]lst. c2"
- shows "(c1[z1::=b]\<^sub>v) = (c2[z2::=b]\<^sub>v)"
-proof -
- obtain x::x where *:"atom x \<sharp> (z1,z2,c1,c2)" using obtain_fresh by metis
- hence "(c1[z1::=[x]\<^sup>v]\<^sub>v) = (c2[z2::=[x]\<^sup>v]\<^sub>v)" using subst_v_flip_eq_one[OF assms, of x] by metis
- hence "(c1[z1::=[x]\<^sup>v]\<^sub>v)[x::=b]\<^sub>v = (c2[z2::=[x]\<^sup>v]\<^sub>v)[x::=b]\<^sub>v" by auto
- thus ?thesis using subst_v_simple_commute * fresh_prod4 by metis
-qed
-
-lemma subst_v_flip_eq_three:
- assumes "[[atom z1]]lst. c1 = [[atom z1']]lst. c1'" and "atom x \<sharp> c1" and "atom x' \<sharp> (x,z1,z1', c1, c1')"
- shows "(x \<leftrightarrow> x') \<bullet> (c1[z1::=[x]\<^sup>v]\<^sub>v) = c1'[z1'::=[x']\<^sup>v]\<^sub>v"
-proof -
- have "atom x' \<sharp> c1[z1::=[x]\<^sup>v]\<^sub>v" using assms fresh_subst_v_if by simp
- hence "(x \<leftrightarrow> x') \<bullet> (c1[z1::=[x]\<^sup>v]\<^sub>v) = c1[z1::=[x]\<^sup>v]\<^sub>v[x::=[x']\<^sup>v]\<^sub>v" using flip_subst_v[of x' "c1[z1::=[x]\<^sup>v]\<^sub>v" x] flip_commute by metis
- also have "... = c1[z1::=[x']\<^sup>v]\<^sub>v" using subst_v_simple_commute fresh_prod4 assms by auto
- also have "... = c1'[z1'::=[x']\<^sup>v]\<^sub>v" using subst_v_flip_eq_one[of z1 c1 z1' c1' x'] using assms by auto
- finally show ?thesis by auto
-qed
-
-end
-
-section \<open>Values\<close>
-
-nominal_function
- subst_vv :: "v \<Rightarrow> x \<Rightarrow> v \<Rightarrow> v" where
- "subst_vv (V_lit l) x v = V_lit l"
-| "subst_vv (V_var y) x v = (if x = y then v else V_var y)"
-| "subst_vv (V_cons tyid c v') x v = V_cons tyid c (subst_vv v' x v)"
-| "subst_vv (V_consp tyid c b v') x v = V_consp tyid c b (subst_vv v' x v)"
-| "subst_vv (V_pair v1 v2) x v = V_pair (subst_vv v1 x v ) (subst_vv v2 x v )"
- by(auto simp: eqvt_def subst_vv_graph_aux_def, metis v.strong_exhaust)
-nominal_termination (eqvt) by lexicographic_order
-
-abbreviation
- subst_vv_abbrev :: "v \<Rightarrow> x \<Rightarrow> v \<Rightarrow> v" ("_[_::=_]\<^sub>v\<^sub>v" [1000,50,50] 1000)
- where
- "v[x::=v']\<^sub>v\<^sub>v \<equiv> subst_vv v x v'"
-
-lemma fresh_subst_vv_if [simp]:
- "j \<sharp> t[i::=x]\<^sub>v\<^sub>v = ((atom i \<sharp> t \<and> j \<sharp> t) \<or> (j \<sharp> x \<and> (j \<sharp> t \<or> j = atom i)))"
- using supp_l_empty apply (induct t rule: v.induct,auto simp add: subst_vv.simps fresh_def, auto)
- by (simp add: supp_at_base |metis b.supp supp_b_empty )+
-
-lemma forget_subst_vv [simp]: "atom a \<sharp> tm \<Longrightarrow> tm[a::=x]\<^sub>v\<^sub>v = tm"
- by (induct tm rule: v.induct) (simp_all add: fresh_at_base)
-
-lemma subst_vv_id [simp]: "tm[a::=V_var a]\<^sub>v\<^sub>v = tm"
- by (induct tm rule: v.induct) simp_all
-
-lemma subst_vv_commute [simp]:
- "atom j \<sharp> tm \<Longrightarrow> tm[i::=t]\<^sub>v\<^sub>v[j::=u]\<^sub>v\<^sub>v = tm[i::=t[j::=u]\<^sub>v\<^sub>v]\<^sub>v\<^sub>v "
- by (induct tm rule: v.induct) (auto simp: fresh_Pair)
-
-lemma subst_vv_commute_full [simp]:
- "atom j \<sharp> t \<Longrightarrow> atom i \<sharp> u \<Longrightarrow> i \<noteq> j \<Longrightarrow> tm[i::=t]\<^sub>v\<^sub>v[j::=u]\<^sub>v\<^sub>v = tm[j::=u]\<^sub>v\<^sub>v[i::=t]\<^sub>v\<^sub>v"
- by (induct tm rule: v.induct) auto
-
-lemma subst_vv_var_flip[simp]:
- fixes v::v
- assumes "atom y \<sharp> v"
- shows "(y \<leftrightarrow> x) \<bullet> v = v [x::=V_var y]\<^sub>v\<^sub>v"
- using assms apply(induct v rule:v.induct)
- apply auto
- using l.fresh l.perm_simps l.strong_exhaust supp_l_empty permute_pure permute_list.simps fresh_def flip_fresh_fresh apply fastforce
- using permute_pure apply blast+
- done
-
-instantiation v :: has_subst_v
-begin
-
-definition
- "subst_v = subst_vv"
-
-instance proof
- fix j::atom and i::x and x::v and t::v
- show "(j \<sharp> subst_v t i x) = ((atom i \<sharp> t \<and> j \<sharp> t) \<or> (j \<sharp> x \<and> (j \<sharp> t \<or> j = atom i)))"
- using fresh_subst_vv_if[of j t i x] subst_v_v_def by metis
-
- fix a::x and tm::v and x::v
- show "atom a \<sharp> tm \<Longrightarrow> subst_v tm a x = tm"
- using forget_subst_vv subst_v_v_def by simp
-
- fix a::x and tm::v
- show "subst_v tm a (V_var a) = tm" using subst_vv_id subst_v_v_def by simp
-
- fix p::perm and x1::x and v::v and t1::v
- show "p \<bullet> subst_v t1 x1 v = subst_v (p \<bullet> t1) (p \<bullet> x1) (p \<bullet> v)"
- using subst_v_v_def by simp
-
- fix x::x and c::v and z::x
- show "atom x \<sharp> c \<Longrightarrow> ((x \<leftrightarrow> z) \<bullet> c) = c[z::=[x]\<^sup>v]\<^sub>v"
- using subst_v_v_def by simp
-
- fix x::x and c::v and z::x
- show "atom x \<sharp> c \<Longrightarrow> c[z::=[x]\<^sup>v]\<^sub>v[x::=v]\<^sub>v = c[z::=v]\<^sub>v"
- using subst_v_v_def by simp
-qed
-
-end
-
-section \<open>Expressions\<close>
-
-nominal_function subst_ev :: "e \<Rightarrow> x \<Rightarrow> v \<Rightarrow> e" where
- "subst_ev ( (AE_val v') ) x v = ( (AE_val (subst_vv v' x v)) )"
-| "subst_ev ( (AE_app f v') ) x v = ( (AE_app f (subst_vv v' x v )) )"
-| "subst_ev ( (AE_appP f b v') ) x v = ( (AE_appP f b (subst_vv v' x v )) )"
-| "subst_ev ( (AE_op opp v1 v2) ) x v = ( (AE_op opp (subst_vv v1 x v ) (subst_vv v2 x v )) )"
-| "subst_ev [#1 v']\<^sup>e x v = [#1 (subst_vv v' x v )]\<^sup>e"
-| "subst_ev [#2 v']\<^sup>e x v = [#2 (subst_vv v' x v )]\<^sup>e"
-| "subst_ev ( (AE_mvar u)) x v = AE_mvar u"
-| "subst_ev [| v' |]\<^sup>e x v = [| (subst_vv v' x v ) |]\<^sup>e"
-| "subst_ev ( AE_concat v1 v2) x v = AE_concat (subst_vv v1 x v ) (subst_vv v2 x v )"
-| "subst_ev ( AE_split v1 v2) x v = AE_split (subst_vv v1 x v ) (subst_vv v2 x v )"
- by(simp add: eqvt_def subst_ev_graph_aux_def,auto)(meson e.strong_exhaust)
-
-nominal_termination (eqvt) by lexicographic_order
-
-abbreviation
- subst_ev_abbrev :: "e \<Rightarrow> x \<Rightarrow> v \<Rightarrow> e" ("_[_::=_]\<^sub>e\<^sub>v" [1000,50,50] 500)
- where
- "e[x::=v']\<^sub>e\<^sub>v \<equiv> subst_ev e x v' "
-
-lemma size_subst_ev [simp]: "size ( subst_ev A i x) = size A"
- apply (nominal_induct A avoiding: i x rule: e.strong_induct)
- by auto
-
-lemma forget_subst_ev [simp]: "atom a \<sharp> A \<Longrightarrow> subst_ev A a x = A"
- apply (nominal_induct A avoiding: a x rule: e.strong_induct)
- by (auto simp: fresh_at_base)
-
-lemma subst_ev_id [simp]: "subst_ev A a (V_var a) = A"
- by (nominal_induct A avoiding: a rule: e.strong_induct) (auto simp: fresh_at_base)
-
-lemma fresh_subst_ev_if [simp]:
- "j \<sharp> (subst_ev A i x ) = ((atom i \<sharp> A \<and> j \<sharp> A) \<or> (j \<sharp> x \<and> (j \<sharp> A \<or> j = atom i)))"
- apply (induct A rule: e.induct)
- unfolding subst_ev.simps fresh_subst_vv_if apply auto+
- using pure_fresh fresh_opp_all apply metis+
- done
-
-lemma subst_ev_commute [simp]:
- "atom j \<sharp> A \<Longrightarrow> (A[i::=t]\<^sub>e\<^sub>v)[j::=u]\<^sub>e\<^sub>v = A[i::=t[j::=u]\<^sub>v\<^sub>v]\<^sub>e\<^sub>v"
- by (nominal_induct A avoiding: i j t u rule: e.strong_induct) (auto simp: fresh_at_base)
-
-lemma subst_ev_var_flip[simp]:
- fixes e::e and y::x and x::x
- assumes "atom y \<sharp> e"
- shows "(y \<leftrightarrow> x) \<bullet> e = e [x::=V_var y]\<^sub>e\<^sub>v"
- using assms apply(nominal_induct e rule:e.strong_induct)
- apply (simp add: subst_v_v_def)
- apply (metis (mono_tags, lifting) b.eq_iff b.perm_simps e.fresh e.perm_simps flip_b_id subst_ev.simps subst_vv_var_flip)
- apply (metis (mono_tags, lifting) b.eq_iff b.perm_simps e.fresh e.perm_simps flip_b_id subst_ev.simps subst_vv_var_flip)
- subgoal for x
- apply (rule_tac y=x in opp.strong_exhaust)
- using subst_vv_var_flip flip_def by (simp add: flip_def permute_pure)+
- using subst_vv_var_flip flip_def by (simp add: flip_def permute_pure)+
-
-lemma subst_ev_flip:
- fixes e::e and ea::e and c::x
- assumes "atom c \<sharp> (e, ea)" and "atom c \<sharp> (x, xa, e, ea)" and "(x \<leftrightarrow> c) \<bullet> e = (xa \<leftrightarrow> c) \<bullet> ea"
- shows "e[x::=v']\<^sub>e\<^sub>v = ea[xa::=v']\<^sub>e\<^sub>v"
-proof -
- have "e[x::=v']\<^sub>e\<^sub>v = (e[x::=V_var c]\<^sub>e\<^sub>v)[c::=v']\<^sub>e\<^sub>v" using subst_ev_commute assms by simp
- also have "... = ((c \<leftrightarrow> x) \<bullet> e)[c::=v']\<^sub>e\<^sub>v" using subst_ev_var_flip assms by simp
- also have "... = ((c \<leftrightarrow> xa) \<bullet> ea)[c::=v']\<^sub>e\<^sub>v" using assms flip_commute by metis
- also have "... = ea[xa::=v']\<^sub>e\<^sub>v" using subst_ev_var_flip assms by simp
- finally show ?thesis by auto
-qed
-
-lemma subst_ev_var[simp]:
- "(AE_val (V_var x))[x::=[z]\<^sup>v]\<^sub>e\<^sub>v = AE_val (V_var z)"
- by auto
-
-instantiation e :: has_subst_v
-begin
-
-definition
- "subst_v = subst_ev"
-
-instance proof
- fix j::atom and i::x and x::v and t::e
- show "(j \<sharp> subst_v t i x) = ((atom i \<sharp> t \<and> j \<sharp> t) \<or> (j \<sharp> x \<and> (j \<sharp> t \<or> j = atom i)))"
- using fresh_subst_ev_if[of j t i x] subst_v_e_def by metis
-
- fix a::x and tm::e and x::v
- show "atom a \<sharp> tm \<Longrightarrow> subst_v tm a x = tm"
- using forget_subst_ev subst_v_e_def by simp
-
- fix a::x and tm::e
- show "subst_v tm a (V_var a) = tm" using subst_ev_id subst_v_e_def by simp
-
- fix p::perm and x1::x and v::v and t1::e
- show "p \<bullet> subst_v t1 x1 v = subst_v (p \<bullet> t1) (p \<bullet> x1) (p \<bullet> v)"
- using subst_ev_commute subst_v_e_def by simp
-
- fix x::x and c::e and z::x
- show "atom x \<sharp> c \<Longrightarrow> ((x \<leftrightarrow> z) \<bullet> c) = c[z::=[x]\<^sup>v]\<^sub>v"
- using subst_v_e_def by simp
-
- fix x::x and c::e and z::x
- show "atom x \<sharp> c \<Longrightarrow> c[z::=[x]\<^sup>v]\<^sub>v[x::=v]\<^sub>v = c[z::=v]\<^sub>v"
- using subst_v_e_def by simp
-qed
-end
-
-lemma subst_ev_commute_full:
- fixes e::e and w::v and v::v
- assumes "atom z \<sharp> v" and "atom x \<sharp> w" and "x \<noteq> z"
- shows "subst_ev (e[z::=w]\<^sub>e\<^sub>v) x v = subst_ev (e[x::=v]\<^sub>e\<^sub>v) z w"
- using assms by(nominal_induct e rule: e.strong_induct,simp+)
-
-lemma subst_ev_v_flip1[simp]:
- fixes e::e
- assumes "atom z1 \<sharp> (z,e)" and "atom z1' \<sharp> (z,e)"
- shows"(z1 \<leftrightarrow> z1') \<bullet> e[z::=v]\<^sub>e\<^sub>v = e[z::= ((z1 \<leftrightarrow> z1') \<bullet> v)]\<^sub>e\<^sub>v"
- using assms proof(nominal_induct e rule:e.strong_induct)
-qed (simp add: flip_def fresh_Pair swap_fresh_fresh)+
-
-section \<open>Expressions in Constraints\<close>
-
-nominal_function subst_cev :: "ce \<Rightarrow> x \<Rightarrow> v \<Rightarrow> ce" where
- "subst_cev ( (CE_val v') ) x v = ( (CE_val (subst_vv v' x v )) )"
-| "subst_cev ( (CE_op opp v1 v2) ) x v = ( (CE_op opp (subst_cev v1 x v ) (subst_cev v2 x v )) )"
-| "subst_cev ( (CE_fst v')) x v = CE_fst (subst_cev v' x v )"
-| "subst_cev ( (CE_snd v')) x v = CE_snd (subst_cev v' x v )"
-| "subst_cev ( (CE_len v')) x v = CE_len (subst_cev v' x v )"
-| "subst_cev ( CE_concat v1 v2) x v = CE_concat (subst_cev v1 x v ) (subst_cev v2 x v )"
- apply (simp add: eqvt_def subst_cev_graph_aux_def,auto)
- by (meson ce.strong_exhaust)
-
-nominal_termination (eqvt) by lexicographic_order
-
-abbreviation
- subst_cev_abbrev :: "ce \<Rightarrow> x \<Rightarrow> v \<Rightarrow> ce" ("_[_::=_]\<^sub>c\<^sub>e\<^sub>v" [1000,50,50] 500)
- where
- "e[x::=v']\<^sub>c\<^sub>e\<^sub>v \<equiv> subst_cev e x v'"
-
-lemma size_subst_cev [simp]: "size ( subst_cev A i x ) = size A"
- by (nominal_induct A avoiding: i x rule: ce.strong_induct,auto)
-
-lemma forget_subst_cev [simp]: "atom a \<sharp> A \<Longrightarrow> subst_cev A a x = A"
- by (nominal_induct A avoiding: a x rule: ce.strong_induct, auto simp: fresh_at_base)
-
-lemma subst_cev_id [simp]: "subst_cev A a (V_var a) = A"
- by (nominal_induct A avoiding: a rule: ce.strong_induct) (auto simp: fresh_at_base)
-
-lemma fresh_subst_cev_if [simp]:
- "j \<sharp> (subst_cev A i x ) = ((atom i \<sharp> A \<and> j \<sharp> A) \<or> (j \<sharp> x \<and> (j \<sharp> A \<or> j = atom i)))"
-proof(nominal_induct A avoiding: i x rule: ce.strong_induct)
- case (CE_op opp v1 v2)
- then show ?case using fresh_subst_vv_if subst_ev.simps e.supp pure_fresh opp.fresh
- fresh_e_opp
- using fresh_opp_all by auto
-qed(auto)+
-
-lemma subst_cev_commute [simp]:
- "atom j \<sharp> A \<Longrightarrow> (subst_cev (subst_cev A i t ) j u) = subst_cev A i (subst_vv t j u )"
- by (nominal_induct A avoiding: i j t u rule: ce.strong_induct) (auto simp: fresh_at_base)
-
-lemma subst_cev_var_flip[simp]:
- fixes e::ce and y::x and x::x
- assumes "atom y \<sharp> e"
- shows "(y \<leftrightarrow> x) \<bullet> e = e [x::=V_var y]\<^sub>c\<^sub>e\<^sub>v"
- using assms proof(nominal_induct e rule:ce.strong_induct)
- case (CE_val v)
- then show ?case using subst_vv_var_flip by auto
-next
- case (CE_op opp v1 v2)
- hence yf: "atom y \<sharp> v1 \<and> atom y \<sharp> v2" using ce.fresh by blast
- have " (y \<leftrightarrow> x) \<bullet> (CE_op opp v1 v2 ) = CE_op ((y \<leftrightarrow> x) \<bullet> opp) ( (y \<leftrightarrow> x) \<bullet> v1 ) ( (y \<leftrightarrow> x) \<bullet> v2)"
- using opp.perm_simps ce.perm_simps permute_pure ce.fresh opp.strong_exhaust by presburger
- also have "... = CE_op ((y \<leftrightarrow> x) \<bullet> opp) (v1[x::=V_var y]\<^sub>c\<^sub>e\<^sub>v) (v2 [x::=V_var y]\<^sub>c\<^sub>e\<^sub>v)" using yf
- by (simp add: CE_op.hyps(1) CE_op.hyps(2))
- finally show ?case using subst_cev.simps opp.perm_simps opp.strong_exhaust
- by (metis (full_types))
-qed( (auto simp add: permute_pure subst_vv_var_flip)+)
-
-lemma subst_cev_flip:
- fixes e::ce and ea::ce and c::x
- assumes "atom c \<sharp> (e, ea)" and "atom c \<sharp> (x, xa, e, ea)" and "(x \<leftrightarrow> c) \<bullet> e = (xa \<leftrightarrow> c) \<bullet> ea"
- shows "e[x::=v']\<^sub>c\<^sub>e\<^sub>v = ea[xa::=v']\<^sub>c\<^sub>e\<^sub>v"
-proof -
- have "e[x::=v']\<^sub>c\<^sub>e\<^sub>v = (e[x::=V_var c]\<^sub>c\<^sub>e\<^sub>v)[c::=v']\<^sub>c\<^sub>e\<^sub>v" using subst_ev_commute assms by simp
- also have "... = ((c \<leftrightarrow> x) \<bullet> e)[c::=v']\<^sub>c\<^sub>e\<^sub>v" using subst_ev_var_flip assms by simp
- also have "... = ((c \<leftrightarrow> xa) \<bullet> ea)[c::=v']\<^sub>c\<^sub>e\<^sub>v" using assms flip_commute by metis
- also have "... = ea[xa::=v']\<^sub>c\<^sub>e\<^sub>v" using subst_ev_var_flip assms by simp
- finally show ?thesis by auto
-qed
-
-lemma subst_cev_var[simp]:
- fixes z::x and x::x
- shows "[[x]\<^sup>v]\<^sup>c\<^sup>e [x::=[z]\<^sup>v]\<^sub>c\<^sub>e\<^sub>v = [[z]\<^sup>v]\<^sup>c\<^sup>e"
- by auto
-
-instantiation ce :: has_subst_v
-begin
-
-definition
- "subst_v = subst_cev"
-
-instance proof
- fix j::atom and i::x and x::v and t::ce
- show "(j \<sharp> subst_v t i x) = ((atom i \<sharp> t \<and> j \<sharp> t) \<or> (j \<sharp> x \<and> (j \<sharp> t \<or> j = atom i)))"
- using fresh_subst_cev_if[of j t i x] subst_v_ce_def by metis
-
- fix a::x and tm::ce and x::v
- show "atom a \<sharp> tm \<Longrightarrow> subst_v tm a x = tm"
- using forget_subst_cev subst_v_ce_def by simp
-
- fix a::x and tm::ce
- show "subst_v tm a (V_var a) = tm" using subst_cev_id subst_v_ce_def by simp
-
- fix p::perm and x1::x and v::v and t1::ce
- show "p \<bullet> subst_v t1 x1 v = subst_v (p \<bullet> t1) (p \<bullet> x1) (p \<bullet> v)"
- using subst_cev_commute subst_v_ce_def by simp
-
- fix x::x and c::ce and z::x
- show "atom x \<sharp> c \<Longrightarrow> ((x \<leftrightarrow> z) \<bullet> c) = c [z::=V_var x]\<^sub>v"
- using subst_v_ce_def by simp
-
- fix x::x and c::ce and z::x
- show "atom x \<sharp> c \<Longrightarrow> c [z::=V_var x]\<^sub>v[x::=v]\<^sub>v = c[z::=v]\<^sub>v"
- using subst_v_ce_def by simp
-qed
-
-end
-
-lemma subst_cev_commute_full:
- fixes e::ce and w::v and v::v
- assumes "atom z \<sharp> v" and "atom x \<sharp> w" and "x \<noteq> z"
- shows "subst_cev (e[z::=w]\<^sub>c\<^sub>e\<^sub>v) x v = subst_cev (e[x::=v]\<^sub>c\<^sub>e\<^sub>v) z w "
- using assms by(nominal_induct e rule: ce.strong_induct,simp+)
-
-
-lemma subst_cev_v_flip1[simp]:
- fixes e::ce
- assumes "atom z1 \<sharp> (z,e)" and "atom z1' \<sharp> (z,e)"
- shows"(z1 \<leftrightarrow> z1') \<bullet> e[z::=v]\<^sub>c\<^sub>e\<^sub>v = e[z::= ((z1 \<leftrightarrow> z1') \<bullet> v)]\<^sub>c\<^sub>e\<^sub>v"
- using assms apply(nominal_induct e rule:ce.strong_induct)
- by (simp add: flip_def fresh_Pair swap_fresh_fresh)+
-
-section \<open>Constraints\<close>
-
-nominal_function subst_cv :: "c \<Rightarrow> x \<Rightarrow> v \<Rightarrow> c" where
- "subst_cv (C_true) x v = C_true"
-| "subst_cv (C_false) x v = C_false"
-| "subst_cv (C_conj c1 c2) x v = C_conj (subst_cv c1 x v ) (subst_cv c2 x v )"
-| "subst_cv (C_disj c1 c2) x v = C_disj (subst_cv c1 x v ) (subst_cv c2 x v )"
-| "subst_cv (C_imp c1 c2) x v = C_imp (subst_cv c1 x v ) (subst_cv c2 x v )"
-| "subst_cv (e1 == e2) x v = ((subst_cev e1 x v ) == (subst_cev e2 x v ))"
-| "subst_cv (C_not c) x v = C_not (subst_cv c x v )"
- apply (simp add: eqvt_def subst_cv_graph_aux_def,auto)
- using c.strong_exhaust by metis
-nominal_termination (eqvt) by lexicographic_order
-
-abbreviation
- subst_cv_abbrev :: "c \<Rightarrow> x \<Rightarrow> v \<Rightarrow> c" ("_[_::=_]\<^sub>c\<^sub>v" [1000,50,50] 1000)
- where
- "c[x::=v']\<^sub>c\<^sub>v \<equiv> subst_cv c x v'"
-
-lemma size_subst_cv [simp]: "size ( subst_cv A i x ) = size A"
- by (nominal_induct A avoiding: i x rule: c.strong_induct,auto)
-
-lemma forget_subst_cv [simp]: "atom a \<sharp> A \<Longrightarrow> subst_cv A a x = A"
- by (nominal_induct A avoiding: a x rule: c.strong_induct, auto simp: fresh_at_base)
-
-lemma subst_cv_id [simp]: "subst_cv A a (V_var a) = A"
- by (nominal_induct A avoiding: a rule: c.strong_induct) (auto simp: fresh_at_base)
-
-lemma fresh_subst_cv_if [simp]:
- "j \<sharp> (subst_cv A i x ) \<longleftrightarrow> (atom i \<sharp> A \<and> j \<sharp> A) \<or> (j \<sharp> x \<and> (j \<sharp> A \<or> j = atom i))"
- by (nominal_induct A avoiding: i x rule: c.strong_induct, (auto simp add: pure_fresh)+)
-
-lemma subst_cv_commute [simp]:
- "atom j \<sharp> A \<Longrightarrow> (subst_cv (subst_cv A i t ) j u ) = subst_cv A i (subst_vv t j u )"
- by (nominal_induct A avoiding: i j t u rule: c.strong_induct) (auto simp: fresh_at_base)
-
-lemma let_s_size [simp]: "size s \<le> size (AS_let x e s)"
- apply (nominal_induct s avoiding: e x rule: s_branch_s_branch_list.strong_induct(1))
- apply auto
- done
-
-lemma subst_cv_var_flip[simp]:
- fixes c::c
- assumes "atom y \<sharp> c"
- shows "(y \<leftrightarrow> x) \<bullet> c = c[x::=V_var y]\<^sub>c\<^sub>v"
- using assms by(nominal_induct c rule:c.strong_induct,(simp add: flip_subst_v subst_v_ce_def)+)
-
-instantiation c :: has_subst_v
-begin
-
-definition
- "subst_v = subst_cv"
-
-instance proof
- fix j::atom and i::x and x::v and t::c
- show "(j \<sharp> subst_v t i x) = ((atom i \<sharp> t \<and> j \<sharp> t) \<or> (j \<sharp> x \<and> (j \<sharp> t \<or> j = atom i)))"
- using fresh_subst_cv_if[of j t i x] subst_v_c_def by metis
-
- fix a::x and tm::c and x::v
- show "atom a \<sharp> tm \<Longrightarrow> subst_v tm a x = tm"
- using forget_subst_cv subst_v_c_def by simp
-
- fix a::x and tm::c
- show "subst_v tm a (V_var a) = tm" using subst_cv_id subst_v_c_def by simp
-
- fix p::perm and x1::x and v::v and t1::c
- show "p \<bullet> subst_v t1 x1 v = subst_v (p \<bullet> t1) (p \<bullet> x1) (p \<bullet> v)"
- using subst_cv_commute subst_v_c_def by simp
-
- fix x::x and c::c and z::x
- show "atom x \<sharp> c \<Longrightarrow> ((x \<leftrightarrow> z) \<bullet> c) = c[z::=[x]\<^sup>v]\<^sub>v"
- using subst_cv_var_flip subst_v_c_def by simp
-
- fix x::x and c::c and z::x
- show "atom x \<sharp> c \<Longrightarrow> c[z::=[x]\<^sup>v]\<^sub>v[x::=v]\<^sub>v = c[z::=v]\<^sub>v"
- using subst_cv_var_flip subst_v_c_def by simp
-qed
-
-end
-
-lemma subst_cv_var_flip1[simp]:
- fixes c::c
- assumes "atom y \<sharp> c"
- shows "(x \<leftrightarrow> y) \<bullet> c = c[x::=V_var y]\<^sub>c\<^sub>v"
- using subst_cv_var_flip flip_commute
- by (metis assms)
-
-lemma subst_cv_v_flip3[simp]:
- fixes c::c
- assumes "atom z1 \<sharp> c" and "atom z1' \<sharp> c"
- shows"(z1 \<leftrightarrow> z1') \<bullet> c[z::=[z1]\<^sup>v]\<^sub>c\<^sub>v = c[z::=[z1']\<^sup>v]\<^sub>c\<^sub>v"
-proof -
- consider "z1' = z" | "z1 = z" | "atom z1 \<sharp> z \<and> atom z1' \<sharp> z" by force
- then show ?thesis proof(cases)
- case 1
- then show ?thesis using 1 assms by auto
- next
- case 2
- then show ?thesis using 2 assms by auto
- next
- case 3
- then show ?thesis using assms by auto
- qed
-qed
-
-lemma subst_cv_v_flip[simp]:
- fixes c::c
- assumes "atom x \<sharp> c"
- shows "((x \<leftrightarrow> z) \<bullet> c)[x::=v]\<^sub>c\<^sub>v = c [z::=v]\<^sub>c\<^sub>v"
- using assms subst_v_c_def by auto
-
-lemma subst_cv_commute_full:
- fixes c::c
- assumes "atom z \<sharp> v" and "atom x \<sharp> w" and "x\<noteq>z"
- shows "(c[z::=w]\<^sub>c\<^sub>v)[x::=v]\<^sub>c\<^sub>v = (c[x::=v]\<^sub>c\<^sub>v)[z::=w]\<^sub>c\<^sub>v"
- using assms proof(nominal_induct c rule: c.strong_induct)
- case (C_eq e1 e2)
- then show ?case using subst_cev_commute_full by simp
-qed(force+)
-
-lemma subst_cv_eq[simp]:
- assumes "atom z1 \<sharp> e1"
- shows "(CE_val (V_var z1) == e1 )[z1::=[x]\<^sup>v]\<^sub>c\<^sub>v = (CE_val (V_var x) == e1 )" (is "?A = ?B")
-proof -
- have "?A = (((CE_val (V_var z1))[z1::=[x]\<^sup>v]\<^sub>c\<^sub>e\<^sub>v) == e1)" using subst_cv.simps assms by simp
- thus ?thesis by simp
-qed
-
-section \<open>Variable Context\<close>
-
-text \<open>The idea of this substitution is to remove x from the context. We really want to add the condition
-that x is fresh in v but this causes problems with proofs.\<close>
-
-nominal_function subst_gv :: "\<Gamma> \<Rightarrow> x \<Rightarrow> v \<Rightarrow> \<Gamma>" where
- "subst_gv GNil x v = GNil"
-| "subst_gv ((y,b,c) #\<^sub>\<Gamma> \<Gamma>) x v = (if x = y then \<Gamma> else ((y,b,c[x::=v]\<^sub>c\<^sub>v)#\<^sub>\<Gamma> (subst_gv \<Gamma> x v )))"
-proof(goal_cases)
- case 1
- then show ?case by(simp add: eqvt_def subst_gv_graph_aux_def )
-next
- case (3 P x)
- then show ?case by (metis neq_GNil_conv prod_cases3)
-qed(fast+)
-nominal_termination (eqvt) by lexicographic_order
-
-abbreviation
- subst_gv_abbrev :: "\<Gamma> \<Rightarrow> x \<Rightarrow> v \<Rightarrow> \<Gamma>" ("_[_::=_]\<^sub>\<Gamma>\<^sub>v" [1000,50,50] 1000)
- where
- "g[x::=v]\<^sub>\<Gamma>\<^sub>v \<equiv> subst_gv g x v"
-
-lemma size_subst_gv [simp]: "size ( subst_gv G i x ) \<le> size G"
- by (induct G,auto)
-
-lemma forget_subst_gv [simp]: "atom a \<sharp> G \<Longrightarrow> subst_gv G a x = G"
- apply (induct G ,auto)
- using fresh_GCons fresh_PairD(1) not_self_fresh apply blast
- apply (simp add: fresh_GCons)+
- done
-
-lemma fresh_subst_gv: "atom a \<sharp> G \<Longrightarrow> atom a \<sharp> v \<Longrightarrow> atom a \<sharp> subst_gv G x v"
-proof(induct G)
- case GNil
- then show ?case by auto
-next
- case (GCons xbc G)
- obtain x' and b' and c' where xbc: "xbc = (x',b',c')" using prod_cases3 by blast
- show ?case proof(cases "x=x'")
- case True
- have "atom a \<sharp> G" using GCons fresh_GCons by blast
- thus ?thesis using subst_gv.simps(2)[of x' b' c' G] GCons xbc True by presburger
- next
- case False
- then show ?thesis using subst_gv.simps(2)[of x' b' c' G] GCons xbc False fresh_GCons by simp
- qed
-qed
-
-lemma subst_gv_flip:
- fixes x::x and xa::x and z::x and c::c and b::b and \<Gamma>::\<Gamma>
- assumes "atom xa \<sharp> ((x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)" and "atom xa \<sharp> \<Gamma>" and "atom x \<sharp> \<Gamma>" and "atom x \<sharp> (z, c)" and "atom xa \<sharp> (z, c)"
- shows "(x \<leftrightarrow> xa) \<bullet> ((x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) = (xa, b, c[z::=V_var xa]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>"
-proof -
- have "(x \<leftrightarrow> xa) \<bullet> ((x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) = (( (x \<leftrightarrow> xa) \<bullet> x, b, (x \<leftrightarrow> xa) \<bullet> c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> ((x \<leftrightarrow> xa) \<bullet> \<Gamma>))"
- using subst Cons_eqvt flip_fresh_fresh using G_cons_flip by simp
- also have "... = ((xa, b, (x \<leftrightarrow> xa) \<bullet> c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> ((x \<leftrightarrow> xa) \<bullet> \<Gamma>))" using assms by fastforce
- also have "... = ((xa, b, c[z::=V_var xa]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> ((x \<leftrightarrow> xa) \<bullet> \<Gamma>))" using assms subst_cv_var_flip by fastforce
- also have "... = ((xa, b, c[z::=V_var xa]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)" using assms flip_fresh_fresh by blast
- finally show ?thesis by simp
-qed
-
-section \<open>Types\<close>
-
-nominal_function subst_tv :: "\<tau> \<Rightarrow> x \<Rightarrow> v \<Rightarrow> \<tau>" where
- "atom z \<sharp> (x,v) \<Longrightarrow> subst_tv \<lbrace> z : b | c \<rbrace> x v = \<lbrace> z : b | c[x::=v]\<^sub>c\<^sub>v \<rbrace>"
- apply (simp add: eqvt_def subst_tv_graph_aux_def )
- apply auto
- subgoal for P a aa b
- apply(rule_tac y=a and c="(aa,b)" in \<tau>.strong_exhaust)
- by (auto simp: eqvt_at_def fresh_star_def fresh_Pair fresh_at_base)
- apply (auto simp: eqvt_at_def fresh_star_def fresh_Pair fresh_at_base)
-proof -
- fix z :: x and c :: c and za :: x and xa :: x and va :: v and ca :: c and cb :: x
- assume a1: "atom za \<sharp> va" and a2: "atom z \<sharp> va" and a3: "\<forall>cb. atom cb \<sharp> c \<and> atom cb \<sharp> ca \<longrightarrow> cb \<noteq> z \<and> cb \<noteq> za \<longrightarrow> c[z::=V_var cb]\<^sub>c\<^sub>v = ca[za::=V_var cb]\<^sub>c\<^sub>v"
- assume a4: "atom cb \<sharp> c" and a5: "atom cb \<sharp> ca" and a6: "cb \<noteq> z" and a7: "cb \<noteq> za" and "atom cb \<sharp> va" and a8: "za \<noteq> xa" and a9: "z \<noteq> xa"
- assume a10:"cb \<noteq> xa"
- note assms = a10 a9 a8 a7 a6 a5 a4 a3 a2 a1
-
- have "c[z::=V_var cb]\<^sub>c\<^sub>v = ca[za::=V_var cb]\<^sub>c\<^sub>v" using assms by auto
- hence "c[z::=V_var cb]\<^sub>c\<^sub>v[xa::=va]\<^sub>c\<^sub>v = ca[za::=V_var cb]\<^sub>c\<^sub>v[xa::=va]\<^sub>c\<^sub>v" by simp
- moreover have "c[z::=V_var cb]\<^sub>c\<^sub>v[xa::=va]\<^sub>c\<^sub>v = c[xa::=va]\<^sub>c\<^sub>v[z::=V_var cb]\<^sub>c\<^sub>v" using subst_cv_commute_full[of z va xa "V_var cb" ] assms fresh_def v.supp by fastforce
- moreover have "ca[za::=V_var cb]\<^sub>c\<^sub>v[xa::=va]\<^sub>c\<^sub>v = ca[xa::=va]\<^sub>c\<^sub>v[za::=V_var cb]\<^sub>c\<^sub>v"
- using subst_cv_commute_full[of za va xa "V_var cb" ] assms fresh_def v.supp by fastforce
-
- ultimately show "c[xa::=va]\<^sub>c\<^sub>v[z::=V_var cb]\<^sub>c\<^sub>v = ca[xa::=va]\<^sub>c\<^sub>v[za::=V_var cb]\<^sub>c\<^sub>v" by simp
-qed
-
-nominal_termination (eqvt) by lexicographic_order
-
-abbreviation
- subst_tv_abbrev :: "\<tau> \<Rightarrow> x \<Rightarrow> v \<Rightarrow> \<tau>" ("_[_::=_]\<^sub>\<tau>\<^sub>v" [1000,50,50] 1000)
- where
- "t[x::=v]\<^sub>\<tau>\<^sub>v \<equiv> subst_tv t x v"
-
-lemma size_subst_tv [simp]: "size ( subst_tv A i x ) = size A"
-proof (nominal_induct A avoiding: i x rule: \<tau>.strong_induct)
- case (T_refined_type x' b' c')
- then show ?case by auto
-qed
-
-lemma forget_subst_tv [simp]: "atom a \<sharp> A \<Longrightarrow> subst_tv A a x = A"
- apply (nominal_induct A avoiding: a x rule: \<tau>.strong_induct)
- apply(auto simp: fresh_at_base)
- done
-
-lemma subst_tv_id [simp]: "subst_tv A a (V_var a) = A"
- by (nominal_induct A avoiding: a rule: \<tau>.strong_induct) (auto simp: fresh_at_base)
-
-lemma fresh_subst_tv_if [simp]:
- "j \<sharp> (subst_tv A i x ) \<longleftrightarrow> (atom i \<sharp> A \<and> j \<sharp> A) \<or> (j \<sharp> x \<and> (j \<sharp> A \<or> j = atom i))"
- apply (nominal_induct A avoiding: i x rule: \<tau>.strong_induct)
- using fresh_def supp_b_empty x_fresh_b by auto
-
-lemma subst_tv_commute [simp]:
- "atom y \<sharp> \<tau> \<Longrightarrow> (\<tau>[x::= t]\<^sub>\<tau>\<^sub>v)[y::=v]\<^sub>\<tau>\<^sub>v = \<tau>[x::= t[y::=v]\<^sub>v\<^sub>v]\<^sub>\<tau>\<^sub>v "
- by (nominal_induct \<tau> avoiding: x y t v rule: \<tau>.strong_induct) (auto simp: fresh_at_base)
-
-lemma subst_tv_var_flip [simp]:
- fixes x::x and xa::x and \<tau>::\<tau>
- assumes "atom xa \<sharp> \<tau>"
- shows "(x \<leftrightarrow> xa) \<bullet> \<tau> = \<tau>[x::=V_var xa]\<^sub>\<tau>\<^sub>v"
-proof -
- obtain z::x and b and c where zbc: "atom z \<sharp> (x,xa, V_var xa) \<and> \<tau> = \<lbrace> z : b | c \<rbrace>"
- using obtain_fresh_z by (metis prod.inject subst_tv.cases)
- hence "atom xa \<notin> supp c - { atom z }" using \<tau>.supp[of z b c] fresh_def supp_b_empty assms
- by auto
- moreover have "xa \<noteq> z" using zbc fresh_prod3 by force
- ultimately have xaf: "atom xa \<sharp> c" using fresh_def by auto
- have "(x \<leftrightarrow> xa) \<bullet> \<tau> = \<lbrace> z : b | (x \<leftrightarrow> xa) \<bullet> c \<rbrace>"
- by (metis \<tau>.perm_simps empty_iff flip_at_base_simps(3) flip_fresh_fresh fresh_PairD(1) fresh_PairD(2) fresh_def not_self_fresh supp_b_empty v.fresh(2) zbc)
- also have "... = \<lbrace> z : b | c[x::=V_var xa]\<^sub>c\<^sub>v \<rbrace>" using subst_cv_v_flip xaf
- by (metis permute_flip_cancel permute_flip_cancel2 subst_cv_var_flip)
- finally show ?thesis using subst_tv.simps zbc
- using fresh_PairD(1) not_self_fresh by force
-qed
-
-instantiation \<tau> :: has_subst_v
-begin
-
-definition
- "subst_v = subst_tv"
-
-instance proof
- fix j::atom and i::x and x::v and t::\<tau>
- show "(j \<sharp> subst_v t i x) = ((atom i \<sharp> t \<and> j \<sharp> t) \<or> (j \<sharp> x \<and> (j \<sharp> t \<or> j = atom i)))"
-
- proof(nominal_induct t avoiding: i x rule:\<tau>.strong_induct)
- case (T_refined_type z b c)
- hence " j \<sharp> \<lbrace> z : b | c \<rbrace>[i::=x]\<^sub>v = j \<sharp> \<lbrace> z : b | c[i::=x]\<^sub>c\<^sub>v \<rbrace>" using subst_tv.simps subst_v_\<tau>_def fresh_Pair by simp
- also have "... = (atom i \<sharp> \<lbrace> z : b | c \<rbrace> \<and> j \<sharp> \<lbrace> z : b | c \<rbrace> \<or> j \<sharp> x \<and> (j \<sharp> \<lbrace> z : b | c \<rbrace> \<or> j = atom i))"
- unfolding \<tau>.fresh using subst_v_c_def fresh_subst_v_if
- using T_refined_type.hyps(1) T_refined_type.hyps(2) x_fresh_b by auto
- finally show ?case by auto
- qed
-
- fix a::x and tm::\<tau> and x::v
- show "atom a \<sharp> tm \<Longrightarrow> subst_v tm a x = tm"
- apply(nominal_induct tm avoiding: a x rule:\<tau>.strong_induct)
- using subst_v_c_def forget_subst_v subst_tv.simps subst_v_\<tau>_def fresh_Pair by simp
-
- fix a::x and tm::\<tau>
- show "subst_v tm a (V_var a) = tm"
- apply(nominal_induct tm avoiding: a rule:\<tau>.strong_induct)
- using subst_v_c_def forget_subst_v subst_tv.simps subst_v_\<tau>_def fresh_Pair by simp
-
- fix p::perm and x1::x and v::v and t1::\<tau>
- show "p \<bullet> subst_v t1 x1 v = subst_v (p \<bullet> t1) (p \<bullet> x1) (p \<bullet> v)"
- apply(nominal_induct tm avoiding: a x rule:\<tau>.strong_induct)
- using subst_v_c_def forget_subst_v subst_tv.simps subst_v_\<tau>_def fresh_Pair by simp
-
- fix x::x and c::\<tau> and z::x
- show "atom x \<sharp> c \<Longrightarrow> ((x \<leftrightarrow> z) \<bullet> c) = c[z::=[x]\<^sup>v]\<^sub>v"
- apply(nominal_induct c avoiding: z x rule:\<tau>.strong_induct)
- using subst_v_c_def flip_subst_v subst_tv.simps subst_v_\<tau>_def fresh_Pair by auto
-
- fix x::x and c::\<tau> and z::x
- show "atom x \<sharp> c \<Longrightarrow> c[z::=[x]\<^sup>v]\<^sub>v[x::=v]\<^sub>v = c[z::=v]\<^sub>v"
- apply(nominal_induct c avoiding: x v z rule:\<tau>.strong_induct)
- using subst_v_c_def subst_tv.simps subst_v_\<tau>_def fresh_Pair
- by (metis flip_commute subst_tv_commute subst_tv_var_flip subst_v_\<tau>_def subst_vv.simps(2))
-qed
-
-end
-
-lemma subst_tv_commute_full:
- fixes c::\<tau>
- assumes "atom z \<sharp> v" and "atom x \<sharp> w" and "x\<noteq>z"
- shows "(c[z::=w]\<^sub>\<tau>\<^sub>v)[x::=v]\<^sub>\<tau>\<^sub>v = (c[x::=v]\<^sub>\<tau>\<^sub>v)[z::=w]\<^sub>\<tau>\<^sub>v"
- using assms proof(nominal_induct c avoiding: x v z w rule: \<tau>.strong_induct)
- case (T_refined_type x1a x2a x3a)
- then show ?case using subst_cv_commute_full by simp
-qed
-
-lemma type_eq_subst_eq:
- fixes v::v and c1::c
- assumes "\<lbrace> z1 : b1 | c1 \<rbrace> = \<lbrace> z2 : b2 | c2 \<rbrace>"
- shows "c1[z1::=v]\<^sub>c\<^sub>v = c2[z2::=v]\<^sub>c\<^sub>v"
- using subst_v_flip_eq_two[of z1 c1 z2 c2 v] \<tau>.eq_iff assms subst_v_c_def by simp
-
-text \<open>Extract constraint from a type. We cannot just project out the constraint as this would
-mean alpha-equivalent types give different answers \<close>
-
-nominal_function c_of :: "\<tau> \<Rightarrow> x \<Rightarrow> c" where
- "atom z \<sharp> x \<Longrightarrow> c_of (T_refined_type z b c) x = c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v"
-proof(goal_cases)
- case 1
- then show ?case using eqvt_def c_of_graph_aux_def by force
-next
- case (2 x y)
- then show ?case using eqvt_def c_of_graph_aux_def by force
-next
- case (3 P x)
- then obtain x1::\<tau> and x2::x where *:"x = (x1,x2)" by force
- obtain z' and b' and c' where "x1 = \<lbrace> z' : b' | c' \<rbrace> \<and> atom z' \<sharp> x2" using obtain_fresh_z by metis
- then show ?case using 3 * by auto
-next
- case (4 z1 x1 b1 c1 z2 x2 b2 c2)
- then show ?case using subst_v_flip_eq_two \<tau>.eq_iff by (metis prod.inject type_eq_subst_eq)
-qed
-
-nominal_termination (eqvt) by lexicographic_order
-
-lemma c_of_eq:
- shows "c_of \<lbrace> x : b | c \<rbrace> x = c"
-proof(nominal_induct "\<lbrace> x : b | c \<rbrace>" avoiding: x rule: \<tau>.strong_induct)
- case (T_refined_type x' c')
- moreover hence "c_of \<lbrace> x' : b | c' \<rbrace> x = c'[x'::=V_var x]\<^sub>c\<^sub>v" using c_of.simps by auto
- moreover have "\<lbrace> x' : b | c' \<rbrace> = \<lbrace> x : b | c \<rbrace>" using T_refined_type \<tau>.eq_iff by metis
- moreover have "c'[x'::=V_var x]\<^sub>c\<^sub>v = c" using T_refined_type Abs1_eq_iff flip_subst_v subst_v_c_def
- by (metis subst_cv_id)
- ultimately show ?case by auto
-qed
-
-lemma obtain_fresh_z_c_of:
- fixes t::"'b::fs"
- obtains z where "atom z \<sharp> t \<and> \<tau> = \<lbrace> z : b_of \<tau> | c_of \<tau> z \<rbrace>"
-proof -
- obtain z and c where "atom z \<sharp> t \<and> \<tau> = \<lbrace> z : b_of \<tau> | c \<rbrace>" using obtain_fresh_z2 by metis
- moreover hence "c = c_of \<tau> z" using c_of.simps using c_of_eq by metis
- ultimately show ?thesis
- using that by auto
-qed
-
-lemma c_of_fresh:
- fixes x::x
- assumes "atom x \<sharp> (t,z)"
- shows "atom x \<sharp> c_of t z"
-proof -
- obtain z' and c' where z:"t = \<lbrace> z' : b_of t | c' \<rbrace> \<and> atom z' \<sharp> (x,z)" using obtain_fresh_z_c_of by metis
- hence *:"c_of t z = c'[z'::=V_var z]\<^sub>c\<^sub>v" using c_of.simps fresh_Pair by metis
- have "(atom x \<sharp> c' \<or> atom x \<in> set [atom z']) \<and> atom x \<sharp> b_of t" using \<tau>.fresh assms z fresh_Pair by metis
- hence "atom x \<sharp> c'" using fresh_Pair z fresh_at_base(2) by fastforce
- moreover have "atom x \<sharp> V_var z" using assms fresh_Pair v.fresh by metis
- ultimately show ?thesis using assms fresh_subst_v_if[of "atom x" c' z' "V_var z"] subst_v_c_def * by metis
-qed
-
-lemma c_of_switch:
- fixes z::x
- assumes "atom z \<sharp> t"
- shows "(c_of t z)[z::=V_var x]\<^sub>c\<^sub>v = c_of t x"
-proof -
- obtain z' and c' where z:"t = \<lbrace> z' : b_of t | c' \<rbrace> \<and> atom z' \<sharp> (x,z)" using obtain_fresh_z_c_of by metis
- hence "(atom z \<sharp> c' \<or> atom z \<in> set [atom z']) \<and> atom z \<sharp> b_of t" using \<tau>.fresh[of "atom z" z' "b_of t" c'] assms by metis
- moreover have " atom z \<notin> set [atom z']" using z fresh_Pair by force
- ultimately have **:"atom z \<sharp> c'" using fresh_Pair z fresh_at_base(2) by metis
-
- have "(c_of t z)[z::=V_var x]\<^sub>c\<^sub>v = c'[z'::=V_var z]\<^sub>c\<^sub>v[z::=V_var x]\<^sub>c\<^sub>v" using c_of.simps fresh_Pair z by metis
- also have "... = c'[z'::=V_var x]\<^sub>c\<^sub>v" using subst_v_simple_commute subst_v_c_def assms c_of.simps z ** by metis
- finally show ?thesis using c_of.simps[of z' x "b_of t" c'] fresh_Pair z by metis
-qed
-
-lemma type_eq_subst_eq1:
- fixes v::v and c1::c
- assumes "\<lbrace> z1 : b1 | c1 \<rbrace> = (\<lbrace> z2 : b2 | c2 \<rbrace>)" and "atom z1 \<sharp> c2"
- shows "c1[z1::=v]\<^sub>c\<^sub>v = c2[z2::=v]\<^sub>c\<^sub>v" and "b1=b2" and " c1 = (z1 \<leftrightarrow> z2) \<bullet> c2"
-proof -
- show "c1[z1::=v]\<^sub>c\<^sub>v = c2[z2::=v]\<^sub>c\<^sub>v" using type_eq_subst_eq assms by blast
- show "b1=b2" using \<tau>.eq_iff assms by blast
- have "z1 = z2 \<and> c1 = c2 \<or> z1 \<noteq> z2 \<and> c1 = (z1 \<leftrightarrow> z2) \<bullet> c2 \<and> atom z1 \<sharp> c2"
- using \<tau>.eq_iff Abs1_eq_iff[of z1 c1 z2 c2] assms by blast
- thus "c1 = (z1 \<leftrightarrow> z2) \<bullet> c2" by auto
-qed
-
-lemma type_eq_subst_eq2:
- fixes v::v and c1::c
- assumes "\<lbrace> z1 : b1 | c1 \<rbrace> = (\<lbrace> z2 : b2 | c2 \<rbrace>)"
- shows "c1[z1::=v]\<^sub>c\<^sub>v = c2[z2::=v]\<^sub>c\<^sub>v" and "b1=b2" and "[[atom z1]]lst. c1 = [[atom z2]]lst. c2"
-proof -
- show "c1[z1::=v]\<^sub>c\<^sub>v = c2[z2::=v]\<^sub>c\<^sub>v" using type_eq_subst_eq assms by blast
- show "b1=b2" using \<tau>.eq_iff assms by blast
- show "[[atom z1]]lst. c1 = [[atom z2]]lst. c2"
- using \<tau>.eq_iff assms by auto
-qed
-
-lemma type_eq_subst_eq3:
- fixes v::v and c1::c
- assumes "\<lbrace> z1 : b1 | c1 \<rbrace> = (\<lbrace> z2 : b2 | c2 \<rbrace>)" and "atom z1 \<sharp> c2"
- shows "c1 = c2[z2::=V_var z1]\<^sub>c\<^sub>v" and "b1=b2"
- using type_eq_subst_eq1 assms subst_v_c_def
- by (metis subst_cv_var_flip)+
-
-lemma type_eq_flip:
- assumes "atom x \<sharp> c"
- shows "\<lbrace> z : b | c \<rbrace> = \<lbrace> x : b | (x \<leftrightarrow> z ) \<bullet> c \<rbrace>"
- using \<tau>.eq_iff Abs1_eq_iff assms
- by (metis (no_types, lifting) flip_fresh_fresh)
-
-lemma c_of_true:
- "c_of \<lbrace> z' : B_bool | TRUE \<rbrace> x = C_true"
-proof(nominal_induct "\<lbrace> z' : B_bool | TRUE \<rbrace>" avoiding: x rule:\<tau>.strong_induct)
- case (T_refined_type x1a x3a)
- hence "\<lbrace> z' : B_bool | TRUE \<rbrace> = \<lbrace> x1a : B_bool | x3a \<rbrace>" using \<tau>.eq_iff by metis
- then show ?case using subst_cv.simps c_of.simps T_refined_type
- type_eq_subst_eq3
- by (metis type_eq_subst_eq)
-qed
-
-lemma type_eq_subst:
- assumes "atom x \<sharp> c"
- shows "\<lbrace> z : b | c \<rbrace> = \<lbrace> x : b | c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v \<rbrace>"
- using \<tau>.eq_iff Abs1_eq_iff assms
- using subst_cv_var_flip type_eq_flip by auto
-
-lemma type_e_subst_fresh:
- fixes x::x and z::x
- assumes "atom z \<sharp> (x,v)" and "atom x \<sharp> e"
- shows "\<lbrace> z : b | CE_val (V_var z) == e \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v = \<lbrace> z : b | CE_val (V_var z) == e \<rbrace>"
- using assms subst_tv.simps subst_cv.simps forget_subst_cev by simp
-
-lemma type_v_subst_fresh:
- fixes x::x and z::x
- assumes "atom z \<sharp> (x,v)" and "atom x \<sharp> v'"
- shows "\<lbrace> z : b | CE_val (V_var z) == CE_val v' \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v = \<lbrace> z : b | CE_val (V_var z) == CE_val v' \<rbrace>"
- using assms subst_tv.simps subst_cv.simps by simp
-
-lemma subst_tbase_eq:
- "b_of \<tau> = b_of \<tau>[x::=v]\<^sub>\<tau>\<^sub>v"
-proof -
- obtain z and b and c where zbc: "\<tau> = \<lbrace> z:b|c\<rbrace> \<and> atom z \<sharp> (x,v)" using \<tau>.exhaust
- by (metis prod.inject subst_tv.cases)
- hence "b_of \<lbrace> z:b|c\<rbrace> = b_of \<lbrace> z:b|c\<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v" using subst_tv.simps by simp
- thus ?thesis using zbc by blast
-qed
-
-lemma subst_tv_if:
- assumes "atom z1 \<sharp> (x,v)" and "atom z' \<sharp> (x,v)"
- shows "\<lbrace> z1 : b | CE_val (v'[x::=v]\<^sub>v\<^sub>v) == CE_val (V_lit l) IMP (c'[x::=v]\<^sub>c\<^sub>v)[z'::=[z1]\<^sup>v]\<^sub>c\<^sub>v \<rbrace> =
- \<lbrace> z1 : b | CE_val v' == CE_val (V_lit l) IMP c'[z'::=[z1]\<^sup>v]\<^sub>c\<^sub>v \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v"
- using subst_cv_commute_full[of z' v x "V_var z1" c'] subst_tv.simps subst_vv.simps(1) subst_ev.simps subst_cv.simps assms
- by simp
-
-lemma subst_tv_tid:
- assumes "atom za \<sharp> (x,v)"
- shows "\<lbrace> za : B_id tid | TRUE \<rbrace> = \<lbrace> za : B_id tid | TRUE \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v"
- using assms subst_tv.simps subst_cv.simps by presburger
-
-
-lemma b_of_subst:
- "b_of (\<tau>[x::=v]\<^sub>\<tau>\<^sub>v) = b_of \<tau>"
-proof -
- obtain z b c where *:"\<tau> = \<lbrace> z : b | c \<rbrace> \<and> atom z \<sharp> (x,v)" using obtain_fresh_z by metis
- thus ?thesis using subst_tv.simps * by auto
-qed
-
-lemma subst_tv_flip:
- assumes "\<tau>'[x::=v]\<^sub>\<tau>\<^sub>v = \<tau>" and "atom x \<sharp> (v,\<tau>)" and "atom x' \<sharp> (v,\<tau>)"
- shows "((x' \<leftrightarrow> x) \<bullet> \<tau>')[x'::=v]\<^sub>\<tau>\<^sub>v = \<tau>"
-proof -
- have "(x' \<leftrightarrow> x) \<bullet> v = v \<and> (x' \<leftrightarrow> x) \<bullet> \<tau> = \<tau>" using assms flip_fresh_fresh by auto
- thus ?thesis using subst_tv.eqvt[of "(x' \<leftrightarrow> x)" \<tau>' x v ] assms by auto
-qed
-
-lemma subst_cv_true:
- "\<lbrace> z : B_id tid | TRUE \<rbrace> = \<lbrace> z : B_id tid | TRUE \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v"
-proof -
- obtain za::x where "atom za \<sharp> (x,v)" using obtain_fresh by auto
- hence "\<lbrace> z : B_id tid | TRUE \<rbrace> = \<lbrace> za: B_id tid | TRUE \<rbrace>" using \<tau>.eq_iff Abs1_eq_iff by fastforce
- moreover have "\<lbrace> za : B_id tid | TRUE \<rbrace> = \<lbrace> za : B_id tid | TRUE \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v"
- using subst_cv.simps subst_tv.simps by (simp add: \<open>atom za \<sharp> (x, v)\<close>)
- ultimately show ?thesis by argo
-qed
-
-lemma t_eq_supp:
- assumes "(\<lbrace> z : b | c \<rbrace>) = (\<lbrace> z1 : b1 | c1 \<rbrace>)"
- shows "supp c - { atom z } = supp c1 - { atom z1 }"
-proof -
- have "supp c - { atom z } \<union> supp b = supp c1 - { atom z1 } \<union> supp b1" using \<tau>.supp assms
- by (metis list.set(1) list.simps(15) sup_bot.right_neutral supp_b_empty)
- moreover have "supp b = supp b1" using assms \<tau>.eq_iff by simp
- moreover have "atom z1 \<notin> supp b1 \<and> atom z \<notin> supp b" using supp_b_empty by simp
- ultimately show ?thesis
- by (metis \<tau>.eq_iff \<tau>.supp assms b.supp(1) list.set(1) list.set(2) sup_bot.right_neutral)
-qed
-
-lemma fresh_t_eq:
- fixes x::x
- assumes "(\<lbrace> z : b | c \<rbrace>) = (\<lbrace> zz : b | cc \<rbrace>)" and "atom x \<sharp> c" and "x \<noteq> zz"
- shows "atom x \<sharp> cc"
-proof -
- have "supp c - { atom z } \<union> supp b = supp cc - { atom zz } \<union> supp b" using \<tau>.supp assms
- by (metis list.set(1) list.simps(15) sup_bot.right_neutral supp_b_empty)
- moreover have "atom x \<notin> supp c" using assms fresh_def by blast
- ultimately have "atom x \<notin> supp cc - { atom zz } \<union> supp b" by force
- hence "atom x \<notin> supp cc" using assms by simp
- thus ?thesis using fresh_def by auto
-qed
-
-section \<open>Mutable Variable Context\<close>
-
-nominal_function subst_dv :: "\<Delta> \<Rightarrow> x \<Rightarrow> v \<Rightarrow> \<Delta>" where
- "subst_dv DNil x v = DNil"
-| "subst_dv ((u,t) #\<^sub>\<Delta> \<Delta>) x v = ((u,t[x::=v]\<^sub>\<tau>\<^sub>v) #\<^sub>\<Delta> (subst_dv \<Delta> x v ))"
- apply (simp add: eqvt_def subst_dv_graph_aux_def,auto )
- using delete_aux.elims by (metis \<Delta>.exhaust surj_pair)
-nominal_termination (eqvt) by lexicographic_order
-
-abbreviation
- subst_dv_abbrev :: "\<Delta> \<Rightarrow> x \<Rightarrow> v \<Rightarrow> \<Delta>" ("_[_::=_]\<^sub>\<Delta>\<^sub>v" [1000,50,50] 1000)
- where
- "\<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<equiv> subst_dv \<Delta> x v "
-
-nominal_function dmap :: "(u*\<tau> \<Rightarrow> u*\<tau>) \<Rightarrow> \<Delta> \<Rightarrow> \<Delta>" where
- "dmap f DNil = DNil"
-| "dmap f ((u,t)#\<^sub>\<Delta>\<Delta>) = (f (u,t) #\<^sub>\<Delta> (dmap f \<Delta> ))"
- apply (simp add: eqvt_def dmap_graph_aux_def,auto )
- using delete_aux.elims by (metis \<Delta>.exhaust surj_pair)
-nominal_termination (eqvt) by lexicographic_order
-
-lemma subst_dv_iff:
- "\<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v = dmap (\<lambda>(u,t). (u, t[x::=v]\<^sub>\<tau>\<^sub>v)) \<Delta>"
- by(induct \<Delta>, auto)
-
-lemma size_subst_dv [simp]: "size ( subst_dv G i x) \<le> size G"
- by (induct G,auto)
-
-lemma forget_subst_dv [simp]: "atom a \<sharp> G \<Longrightarrow> subst_dv G a x = G"
- apply (induct G ,auto)
- using fresh_DCons fresh_PairD(1) not_self_fresh apply fastforce
- apply (simp add: fresh_DCons)+
- done
-
-lemma subst_dv_member:
- assumes "(u,\<tau>) \<in> setD \<Delta>"
- shows "(u, \<tau>[x::=v]\<^sub>\<tau>\<^sub>v) \<in> setD (\<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v)"
- using assms by(induct \<Delta> rule: \<Delta>_induct,auto)
-
-lemma fresh_subst_dv:
- fixes x::x
- assumes "atom xa \<sharp> \<Delta>" and "atom xa \<sharp> v"
- shows "atom xa \<sharp>\<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v"
- using assms proof(induct \<Delta> rule:\<Delta>_induct)
- case DNil
- then show ?case by auto
-next
- case (DCons u t \<Delta>)
- then show ?case using subst_dv.simps subst_v_\<tau>_def fresh_DCons fresh_Pair by simp
-qed
-
-lemma fresh_subst_dv_if:
- fixes j::atom and i::x and x::v and t::\<Delta>
- assumes "j \<sharp> t \<and> j \<sharp> x"
- shows "(j \<sharp> subst_dv t i x)"
- using assms proof(induct t rule: \<Delta>_induct)
- case DNil
- then show ?case using subst_gv.simps fresh_GNil by auto
-next
- case (DCons u' t' D')
- then show ?case unfolding subst_dv.simps using fresh_DCons fresh_subst_tv_if fresh_Pair by metis
-qed
-
-section \<open>Statements\<close>
-
-text \<open> Using ideas from proofs at top of AFP/Launchbury/Substitution.thy.
- Subproofs borrowed from there; hence the apply style proofs. \<close>
-
-nominal_function (default "case_sum (\<lambda>x. Inl undefined) (case_sum (\<lambda>x. Inl undefined) (\<lambda>x. Inr undefined))")
- subst_sv :: "s \<Rightarrow> x \<Rightarrow> v \<Rightarrow> s"
- and subst_branchv :: "branch_s \<Rightarrow> x \<Rightarrow> v \<Rightarrow> branch_s"
- and subst_branchlv :: "branch_list \<Rightarrow> x \<Rightarrow> v \<Rightarrow> branch_list" where
- "subst_sv ( (AS_val v') ) x v = (AS_val (subst_vv v' x v ))"
-| "atom y \<sharp> (x,v) \<Longrightarrow> subst_sv (AS_let y e s) x v = (AS_let y (e[x::=v]\<^sub>e\<^sub>v) (subst_sv s x v ))"
-| "atom y \<sharp> (x,v) \<Longrightarrow> subst_sv (AS_let2 y t s1 s2) x v = (AS_let2 y (t[x::=v]\<^sub>\<tau>\<^sub>v) (subst_sv s1 x v ) (subst_sv s2 x v ))"
-| " subst_sv (AS_match v' cs) x v = AS_match (v'[x::=v]\<^sub>v\<^sub>v) (subst_branchlv cs x v )"
-| "subst_sv (AS_assign y v') x v = AS_assign y (subst_vv v' x v )"
-| "subst_sv ( (AS_if v' s1 s2) ) x v = (AS_if (subst_vv v' x v ) (subst_sv s1 x v ) (subst_sv s2 x v ) )"
-| "atom u \<sharp> (x,v) \<Longrightarrow> subst_sv (AS_var u \<tau> v' s) x v = AS_var u (subst_tv \<tau> x v ) (subst_vv v' x v ) (subst_sv s x v ) "
-| "subst_sv (AS_while s1 s2) x v = AS_while (subst_sv s1 x v ) (subst_sv s2 x v )"
-| "subst_sv (AS_seq s1 s2) x v = AS_seq (subst_sv s1 x v ) (subst_sv s2 x v )"
-| "subst_sv (AS_assert c s) x v = AS_assert (subst_cv c x v) (subst_sv s x v)"
-| "atom x1 \<sharp> (x,v) \<Longrightarrow> subst_branchv (AS_branch dc x1 s1 ) x v = AS_branch dc x1 (subst_sv s1 x v )"
-
-| "subst_branchlv (AS_final cs) x v = AS_final (subst_branchv cs x v )"
-| "subst_branchlv (AS_cons cs css) x v = AS_cons (subst_branchv cs x v ) (subst_branchlv css x v )"
- apply (auto,simp add: eqvt_def subst_sv_subst_branchv_subst_branchlv_graph_aux_def )
-proof(goal_cases)
-
- have eqvt_at_proj: "\<And> s xa va . eqvt_at subst_sv_subst_branchv_subst_branchlv_sumC (Inl (s, xa, va)) \<Longrightarrow>
- eqvt_at (\<lambda>a. projl (subst_sv_subst_branchv_subst_branchlv_sumC (Inl a))) (s, xa, va)"
- apply(simp add: eqvt_at_def)
- apply(rule)
- apply(subst Projl_permute)
- apply(thin_tac _)+
- apply (simp add: subst_sv_subst_branchv_subst_branchlv_sumC_def)
- apply (simp add: THE_default_def)
- apply (case_tac "Ex1 (subst_sv_subst_branchv_subst_branchlv_graph (Inl (s,xa,va)))")
- apply simp
- apply(auto)[1]
- apply (erule_tac x="x" in allE)
- apply simp
- apply(cases rule: subst_sv_subst_branchv_subst_branchlv_graph.cases)
- apply(assumption)
- apply(rule_tac x="Sum_Type.projl x" in exI,clarify,rule the1_equality,blast,simp (no_asm) only: sum.sel)+
- apply blast +
-
- apply(simp)+
- done
-
- {
- case (1 P x')
- then show ?case proof(cases x')
- case (Inl a) thus P
- proof(cases a)
- case (fields aa bb cc)
- thus P using Inl 1 s_branch_s_branch_list.strong_exhaust fresh_star_insert by metis
- qed
- next
- case (Inr b) thus P
- proof(cases b)
- case (Inl a) thus P proof(cases a)
- case (fields aa bb cc)
- then show ?thesis using Inr Inl 1 s_branch_s_branch_list.strong_exhaust fresh_star_insert by metis
- qed
- next
- case Inr2: (Inr b) thus P proof(cases b)
- case (fields aa bb cc)
- then show ?thesis using Inr Inr2 1 s_branch_s_branch_list.strong_exhaust fresh_star_insert by metis
- qed
- qed
- qed
- next
- case (2 y s ya xa va sa c)
- thus ?case using eqvt_triple eqvt_at_proj by blast
- next
- case (3 y s2 ya xa va s1a s2a c)
- thus ?case using eqvt_triple eqvt_at_proj by blast
- next
- case (4 u xa va s ua sa c)
- moreover have "atom u \<sharp> (xa, va) \<and> atom ua \<sharp> (xa, va)"
- using fresh_Pair u_fresh_xv by auto
- ultimately show ?case using eqvt_triple[of u xa va ua s sa] subst_sv_def eqvt_at_proj by metis
- next
- case (5 x1 s1 x1a xa va s1a c)
- thus ?case using eqvt_triple eqvt_at_proj by blast
- }
-qed
-nominal_termination (eqvt) by lexicographic_order
-
-abbreviation
- subst_sv_abbrev :: "s \<Rightarrow> x \<Rightarrow> v \<Rightarrow> s" ("_[_::=_]\<^sub>s\<^sub>v" [1000,50,50] 1000)
- where
- "s[x::=v]\<^sub>s\<^sub>v \<equiv> subst_sv s x v"
-
-abbreviation
- subst_branchv_abbrev :: "branch_s \<Rightarrow> x \<Rightarrow> v \<Rightarrow> branch_s" ("_[_::=_]\<^sub>s\<^sub>v" [1000,50,50] 1000)
- where
- "s[x::=v]\<^sub>s\<^sub>v \<equiv> subst_branchv s x v"
-
-lemma size_subst_sv [simp]: "size (subst_sv A i x ) = size A" and "size (subst_branchv B i x ) = size B" and "size (subst_branchlv C i x ) = size C"
- by(nominal_induct A and B and C avoiding: i x rule: s_branch_s_branch_list.strong_induct,auto)
-
-lemma forget_subst_sv [simp]: shows "atom a \<sharp> A \<Longrightarrow> subst_sv A a x = A" and "atom a \<sharp> B \<Longrightarrow> subst_branchv B a x = B" and "atom a \<sharp> C \<Longrightarrow> subst_branchlv C a x = C"
- by (nominal_induct A and B and C avoiding: a x rule: s_branch_s_branch_list.strong_induct,auto simp: fresh_at_base)
-
-lemma subst_sv_id [simp]: "subst_sv A a (V_var a) = A" and "subst_branchv B a (V_var a) = B" and "subst_branchlv C a (V_var a) = C"
-proof(nominal_induct A and B and C avoiding: a rule: s_branch_s_branch_list.strong_induct)
- case (AS_let x option e s)
- then show ?case
- by (metis (no_types, lifting) fresh_Pair not_None_eq subst_ev_id subst_sv.simps(2) subst_sv.simps(3) subst_tv_id v.fresh(2))
-next
- case (AS_match v branch_s)
- then show ?case using fresh_Pair not_None_eq subst_ev_id subst_sv.simps subst_sv.simps subst_tv_id v.fresh subst_vv_id
- by metis
-qed(auto)+
-
-lemma fresh_subst_sv_if_rl:
- shows
- "(atom x \<sharp> s \<and> j \<sharp> s) \<or> (j \<sharp> v \<and> (j \<sharp> s \<or> j = atom x)) \<Longrightarrow> j \<sharp> (subst_sv s x v )" and
- "(atom x \<sharp> cs \<and> j \<sharp> cs) \<or> (j \<sharp> v \<and> (j \<sharp> cs \<or> j = atom x)) \<Longrightarrow> j \<sharp> (subst_branchv cs x v)" and
- "(atom x \<sharp> css \<and> j \<sharp> css) \<or> (j \<sharp> v \<and> (j \<sharp> css \<or> j = atom x)) \<Longrightarrow> j \<sharp> (subst_branchlv css x v )"
- apply(nominal_induct s and cs and css avoiding: v x rule: s_branch_s_branch_list.strong_induct)
- using pure_fresh by force+
-
-lemma fresh_subst_sv_if_lr:
- shows "j \<sharp> (subst_sv s x v) \<Longrightarrow> (atom x \<sharp> s \<and> j \<sharp> s) \<or> (j \<sharp> v \<and> (j \<sharp> s \<or> j = atom x))" and
- "j \<sharp> (subst_branchv cs x v) \<Longrightarrow> (atom x \<sharp> cs \<and> j \<sharp> cs) \<or> (j \<sharp> v \<and> (j \<sharp> cs \<or> j = atom x))" and
- "j \<sharp> (subst_branchlv css x v ) \<Longrightarrow> (atom x \<sharp> css \<and> j \<sharp> css) \<or> (j \<sharp> v \<and> (j \<sharp> css \<or> j = atom x))"
-proof(nominal_induct s and cs and css avoiding: v x rule: s_branch_s_branch_list.strong_induct)
- case (AS_branch list x s )
- then show ?case using s_branch_s_branch_list.fresh fresh_Pair list.distinct(1) list.set_cases pure_fresh set_ConsD subst_branchv.simps by metis
-next
- case (AS_let y e s')
- thus ?case proof(cases "atom x \<sharp> (AS_let y e s')")
- case True
- hence "subst_sv (AS_let y e s') x v = (AS_let y e s')" using forget_subst_sv by simp
- hence "j \<sharp> (AS_let y e s')" using AS_let by argo
- then show ?thesis using True by blast
- next
- case False
- have "subst_sv (AS_let y e s') x v = AS_let y (e[x::=v]\<^sub>e\<^sub>v) (s'[x::=v]\<^sub>s\<^sub>v)" using subst_sv.simps(2) AS_let by force
- hence "((j \<sharp> s'[x::=v]\<^sub>s\<^sub>v \<or> j \<in> set [atom y]) \<and> j \<sharp> None \<and> j \<sharp> e[x::=v]\<^sub>e\<^sub>v)" using s_branch_s_branch_list.fresh AS_let
- by (simp add: fresh_None)
- then show ?thesis using AS_let fresh_None fresh_subst_ev_if list.discI list.set_cases s_branch_s_branch_list.fresh set_ConsD
- by metis
- qed
-next
- case (AS_let2 y \<tau> s1 s2)
- thus ?case proof(cases "atom x \<sharp> (AS_let2 y \<tau> s1 s2)")
- case True
- hence "subst_sv (AS_let2 y \<tau> s1 s2) x v = (AS_let2 y \<tau> s1 s2)" using forget_subst_sv by simp
- hence "j \<sharp> (AS_let2 y \<tau> s1 s2)" using AS_let2 by argo
- then show ?thesis using True by blast
- next
- case False
- have "subst_sv (AS_let2 y \<tau> s1 s2) x v = AS_let2 y (\<tau>[x::=v]\<^sub>\<tau>\<^sub>v) (s1[x::=v]\<^sub>s\<^sub>v) (s2[x::=v]\<^sub>s\<^sub>v)" using subst_sv.simps AS_let2 by force
- then show ?thesis using AS_let2
- fresh_subst_tv_if list.discI list.set_cases s_branch_s_branch_list.fresh(4) set_ConsD by auto
- qed
-qed(auto)+
-
-lemma fresh_subst_sv_if[simp]:
- fixes x::x and v::v
- shows "j \<sharp> (subst_sv s x v) \<longleftrightarrow> (atom x \<sharp> s \<and> j \<sharp> s) \<or> (j \<sharp> v \<and> (j \<sharp> s \<or> j = atom x))" and
- "j \<sharp> (subst_branchv cs x v) \<longleftrightarrow> (atom x \<sharp> cs \<and> j \<sharp> cs) \<or> (j \<sharp> v \<and> (j \<sharp> cs \<or> j = atom x))"
- using fresh_subst_sv_if_lr fresh_subst_sv_if_rl by metis+
-
-lemma subst_sv_commute [simp]:
- fixes A::s and t::v and j::x and i::x
- shows "atom j \<sharp> A \<Longrightarrow> (subst_sv (subst_sv A i t) j u ) = subst_sv A i (subst_vv t j u )" and
- "atom j \<sharp> B \<Longrightarrow> (subst_branchv (subst_branchv B i t ) j u ) = subst_branchv B i (subst_vv t j u )" and
- "atom j \<sharp> C \<Longrightarrow> (subst_branchlv (subst_branchlv C i t) j u ) = subst_branchlv C i (subst_vv t j u ) "
- apply(nominal_induct A and B and C avoiding: i j t u rule: s_branch_s_branch_list.strong_induct)
- by(auto simp: fresh_at_base)
-
-lemma c_eq_perm:
- assumes "( (atom z) \<rightleftharpoons> (atom z') ) \<bullet> c = c'" and "atom z' \<sharp> c"
- shows "\<lbrace> z : b | c \<rbrace> = \<lbrace> z' : b | c' \<rbrace>"
- using \<tau>.eq_iff Abs1_eq_iff(3)
- by (metis Nominal2_Base.swap_commute assms(1) assms(2) flip_def swap_fresh_fresh)
-
-lemma subst_sv_flip:
- fixes s::s and sa::s and v'::v
- assumes "atom c \<sharp> (s, sa)" and "atom c \<sharp> (v',x, xa, s, sa)" "atom x \<sharp> v'" and "atom xa \<sharp> v'" and "(x \<leftrightarrow> c) \<bullet> s = (xa \<leftrightarrow> c) \<bullet> sa"
- shows "s[x::=v']\<^sub>s\<^sub>v = sa[xa::=v']\<^sub>s\<^sub>v"
-proof -
- have "atom x \<sharp> (s[x::=v']\<^sub>s\<^sub>v)" and xafr: "atom xa \<sharp> (sa[xa::=v']\<^sub>s\<^sub>v)"
- and "atom c \<sharp> ( s[x::=v']\<^sub>s\<^sub>v, sa[xa::=v']\<^sub>s\<^sub>v)" using assms using fresh_subst_sv_if assms by( blast+ ,force)
-
- hence "s[x::=v']\<^sub>s\<^sub>v = (x \<leftrightarrow> c) \<bullet> (s[x::=v']\<^sub>s\<^sub>v)" by (simp add: flip_fresh_fresh fresh_Pair)
- also have " ... = ((x \<leftrightarrow> c) \<bullet> s)[ ((x \<leftrightarrow> c) \<bullet> x) ::= ((x \<leftrightarrow> c) \<bullet> v') ]\<^sub>s\<^sub>v" using subst_sv_subst_branchv_subst_branchlv.eqvt by blast
- also have "... = ((xa \<leftrightarrow> c) \<bullet> sa)[ ((x \<leftrightarrow> c) \<bullet> x) ::= ((x \<leftrightarrow> c) \<bullet> v') ]\<^sub>s\<^sub>v" using assms by presburger
- also have "... = ((xa \<leftrightarrow> c) \<bullet> sa)[ ((xa \<leftrightarrow> c) \<bullet> xa) ::= ((xa \<leftrightarrow> c) \<bullet> v') ]\<^sub>s\<^sub>v" using assms
- by (metis flip_at_simps(1) flip_fresh_fresh fresh_PairD(1))
- also have "... = (xa \<leftrightarrow> c) \<bullet> (sa[xa::=v']\<^sub>s\<^sub>v)" using subst_sv_subst_branchv_subst_branchlv.eqvt by presburger
- also have "... = sa[xa::=v']\<^sub>s\<^sub>v" using xafr assms by (simp add: flip_fresh_fresh fresh_Pair)
- finally show ?thesis by simp
-qed
-
-lemma if_type_eq:
- fixes \<Gamma>::\<Gamma> and v::v and z1::x
- assumes "atom z1' \<sharp> (v, ca, (x, b, c) #\<^sub>\<Gamma> \<Gamma>, (CE_val v == CE_val (V_lit ll) IMP ca[za::=[z1]\<^sup>v]\<^sub>c\<^sub>v ))" and "atom z1 \<sharp> v"
- and "atom z1 \<sharp> (za,ca)" and "atom z1' \<sharp> (za,ca)"
- shows "(\<lbrace> z1' : ba | CE_val v == CE_val (V_lit ll) IMP ca[za::=[z1']\<^sup>v]\<^sub>c\<^sub>v \<rbrace>) = \<lbrace> z1 : ba | CE_val v == CE_val (V_lit ll) IMP ca[za::=[z1]\<^sup>v]\<^sub>c\<^sub>v \<rbrace>"
-proof -
- have "atom z1' \<sharp> (CE_val v == CE_val (V_lit ll) IMP ca[za::=[z1]\<^sup>v]\<^sub>c\<^sub>v )" using assms fresh_prod4 by blast
- moreover hence "(CE_val v == CE_val (V_lit ll) IMP ca[za::=[z1']\<^sup>v]\<^sub>c\<^sub>v) = (z1' \<leftrightarrow> z1) \<bullet> (CE_val v == CE_val (V_lit ll) IMP ca[za::=[z1]\<^sup>v]\<^sub>c\<^sub>v )"
- proof -
- have "(z1' \<leftrightarrow> z1) \<bullet> (CE_val v == CE_val (V_lit ll) IMP ca[za::=[z1]\<^sup>v]\<^sub>c\<^sub>v ) = ( (z1' \<leftrightarrow> z1) \<bullet> (CE_val v == CE_val (V_lit ll)) IMP ((z1' \<leftrightarrow> z1) \<bullet> ca[za::=[z1]\<^sup>v]\<^sub>c\<^sub>v ))"
- by auto
- also have "... = ((CE_val v == CE_val (V_lit ll)) IMP ((z1' \<leftrightarrow> z1) \<bullet> ca[za::=[z1]\<^sup>v]\<^sub>c\<^sub>v ))"
- using \<open>atom z1 \<sharp> v\<close> assms
- by (metis (mono_tags) \<open>atom z1' \<sharp> (CE_val v == CE_val (V_lit ll) IMP ca[za::=[z1]\<^sup>v]\<^sub>c\<^sub>v )\<close> c.fresh(6) c.fresh(7) ce.fresh(1) flip_at_simps(2) flip_fresh_fresh fresh_at_base_permute_iff fresh_def supp_l_empty v.fresh(1))
- also have "... = ((CE_val v == CE_val (V_lit ll)) IMP (ca[za::=[z1']\<^sup>v]\<^sub>c\<^sub>v ))"
- using assms by fastforce
- finally show ?thesis by auto
- qed
- ultimately show ?thesis
- using \<tau>.eq_iff Abs1_eq_iff(3)[of z1' "CE_val v == CE_val (V_lit ll) IMP ca[za::=[z1']\<^sup>v]\<^sub>c\<^sub>v"
- z1 "CE_val v == CE_val (V_lit ll) IMP ca[za::=[z1]\<^sup>v]\<^sub>c\<^sub>v"] by blast
-qed
-
-lemma subst_sv_var_flip:
- fixes x::x and s::s and z::x
- shows "atom x \<sharp> s \<Longrightarrow> ((x \<leftrightarrow> z) \<bullet> s) = s[z::=[x]\<^sup>v]\<^sub>s\<^sub>v" and
- "atom x \<sharp> cs \<Longrightarrow> ((x \<leftrightarrow> z) \<bullet> cs) = subst_branchv cs z [x]\<^sup>v" and
- "atom x \<sharp> css \<Longrightarrow> ((x \<leftrightarrow> z) \<bullet> css) = subst_branchlv css z [x]\<^sup>v"
- apply(nominal_induct s and cs and css avoiding: z x rule: s_branch_s_branch_list.strong_induct)
- using [[simproc del: alpha_lst]]
- apply (auto ) (* This unpacks subst, perm *)
- using subst_tv_var_flip flip_fresh_fresh v.fresh s_branch_s_branch_list.fresh
- subst_v_\<tau>_def subst_v_v_def subst_vv_var_flip subst_v_e_def subst_ev_var_flip pure_fresh apply auto
- defer 1 (* Sometimes defering hard goals to the end makes it easier to finish *)
- using x_fresh_u apply blast (* Next two involve u and flipping with x *)
- defer 1
- using x_fresh_u apply blast
- defer 1
- using x_fresh_u Abs1_eq_iff'(3) flip_fresh_fresh
- apply (simp add: subst_v_c_def)
- using x_fresh_u Abs1_eq_iff'(3) flip_fresh_fresh
- by (simp add: flip_fresh_fresh)
-
-instantiation s :: has_subst_v
-begin
-
-definition
- "subst_v = subst_sv"
-
-instance proof
- fix j::atom and i::x and x::v and t::s
- show "(j \<sharp> subst_v t i x) = ((atom i \<sharp> t \<and> j \<sharp> t) \<or> (j \<sharp> x \<and> (j \<sharp> t \<or> j = atom i)))"
- using fresh_subst_sv_if subst_v_s_def by auto
-
- fix a::x and tm::s and x::v
- show "atom a \<sharp> tm \<Longrightarrow> subst_v tm a x = tm"
- using forget_subst_sv subst_v_s_def by simp
-
- fix a::x and tm::s
- show "subst_v tm a (V_var a) = tm" using subst_sv_id subst_v_s_def by simp
-
- fix p::perm and x1::x and v::v and t1::s
- show "p \<bullet> subst_v t1 x1 v = subst_v (p \<bullet> t1) (p \<bullet> x1) (p \<bullet> v)"
- using subst_sv_commute subst_v_s_def by simp
-
- fix x::x and c::s and z::x
- show "atom x \<sharp> c \<Longrightarrow> ((x \<leftrightarrow> z) \<bullet> c) = c[z::=[x]\<^sup>v]\<^sub>v"
- using subst_sv_var_flip subst_v_s_def by simp
-
- fix x::x and c::s and z::x
- show "atom x \<sharp> c \<Longrightarrow> c[z::=[x]\<^sup>v]\<^sub>v[x::=v]\<^sub>v = c[z::=v]\<^sub>v"
- using subst_sv_var_flip subst_v_s_def by simp
-qed
-end
-
-section \<open>Type Definition\<close>
-
-nominal_function subst_ft_v :: "fun_typ \<Rightarrow> x \<Rightarrow> v \<Rightarrow> fun_typ" where
- "atom z \<sharp> (x,v) \<Longrightarrow> subst_ft_v ( AF_fun_typ z b c t (s::s)) x v = AF_fun_typ z b c[x::=v]\<^sub>c\<^sub>v t[x::=v]\<^sub>\<tau>\<^sub>v s[x::=v]\<^sub>s\<^sub>v"
- apply(simp add: eqvt_def subst_ft_v_graph_aux_def )
- apply(simp add:fun_typ.strong_exhaust )
- apply(auto)
- apply(rule_tac y=a and c="(aa,b)" in fun_typ.strong_exhaust)
- apply (auto simp: eqvt_at_def fresh_star_def fresh_Pair fresh_at_base)
-
-proof(goal_cases)
- case (1 z xa va c t s za ca ta sa cb)
- hence "c[z::=[ cb ]\<^sup>v]\<^sub>c\<^sub>v = ca[za::=[ cb ]\<^sup>v]\<^sub>c\<^sub>v"
- by (metis flip_commute subst_cv_var_flip)
- hence "c[z::=[ cb ]\<^sup>v]\<^sub>c\<^sub>v[xa::=va]\<^sub>c\<^sub>v = ca[za::=[ cb ]\<^sup>v]\<^sub>c\<^sub>v[xa::=va]\<^sub>c\<^sub>v" by auto
- then show ?case using subst_cv_commute atom_eq_iff fresh_atom fresh_atom_at_base subst_cv_commute_full v.fresh
- using 1 subst_cv_var_flip flip_commute by metis
-next
- case (2 z xa va c t s za ca ta sa cb)
- hence "t[z::=[ cb ]\<^sup>v]\<^sub>\<tau>\<^sub>v = ta[za::=[ cb ]\<^sup>v]\<^sub>\<tau>\<^sub>v" by metis
- hence "t[z::=[ cb ]\<^sup>v]\<^sub>\<tau>\<^sub>v[xa::=va]\<^sub>\<tau>\<^sub>v = ta[za::=[ cb ]\<^sup>v]\<^sub>\<tau>\<^sub>v[xa::=va]\<^sub>\<tau>\<^sub>v" by auto
- then show ?case using subst_tv_commute_full 2
- by (metis atom_eq_iff fresh_atom fresh_atom_at_base v.fresh(2))
-qed
-
-nominal_termination (eqvt) by lexicographic_order
-
-nominal_function subst_ftq_v :: "fun_typ_q \<Rightarrow> x \<Rightarrow> v \<Rightarrow> fun_typ_q" where
- "atom bv \<sharp> (x,v) \<Longrightarrow> subst_ftq_v (AF_fun_typ_some bv ft) x v = (AF_fun_typ_some bv (subst_ft_v ft x v))"
-| "subst_ftq_v (AF_fun_typ_none ft) x v = (AF_fun_typ_none (subst_ft_v ft x v))"
- apply(simp add: eqvt_def subst_ftq_v_graph_aux_def )
- apply(simp add:fun_typ_q.strong_exhaust )
- apply(auto)
- apply(rule_tac y=a and c="(aa,b)" in fun_typ_q.strong_exhaust)
- apply (auto simp: eqvt_at_def fresh_star_def fresh_Pair fresh_at_base)
-proof(goal_cases)
- case (1 bv ft bva fta xa va c)
- then show ?case using subst_ft_v.simps by (simp add: flip_fresh_fresh)
-qed
-nominal_termination (eqvt) by lexicographic_order
-
-lemma size_subst_ft[simp]: "size (subst_ft_v A x v) = size A"
- by(nominal_induct A avoiding: x v rule: fun_typ.strong_induct,auto)
-
-lemma forget_subst_ft [simp]: shows "atom x \<sharp> A \<Longrightarrow> subst_ft_v A x a = A"
- by (nominal_induct A avoiding: a x rule: fun_typ.strong_induct,auto simp: fresh_at_base)
-
-lemma subst_ft_id [simp]: "subst_ft_v A a (V_var a) = A"
- by(nominal_induct A avoiding: a rule: fun_typ.strong_induct,auto)
-
-instantiation fun_typ :: has_subst_v
-begin
-
-definition
- "subst_v = subst_ft_v"
-
-instance proof
-
- fix j::atom and i::x and x::v and t::fun_typ
- show "(j \<sharp> subst_v t i x) = ((atom i \<sharp> t \<and> j \<sharp> t) \<or> (j \<sharp> x \<and> (j \<sharp> t \<or> j = atom i)))"
- apply(nominal_induct t avoiding: i x rule:fun_typ.strong_induct)
- apply(simp only: subst_v_fun_typ_def subst_ft_v.simps )
- using fun_typ.fresh fresh_subst_v_if apply simp
- by auto
-
- fix a::x and tm::fun_typ and x::v
- show "atom a \<sharp> tm \<Longrightarrow> subst_v tm a x = tm"
- proof(nominal_induct tm avoiding: a x rule:fun_typ.strong_induct)
- case (AF_fun_typ x1a x2a x3a x4a x5a)
- then show ?case unfolding subst_ft_v.simps subst_v_fun_typ_def fun_typ.fresh using forget_subst_v subst_ft_v.simps subst_v_c_def forget_subst_sv subst_v_\<tau>_def by fastforce
- qed
-
- fix a::x and tm::fun_typ
- show "subst_v tm a (V_var a) = tm"
- proof(nominal_induct tm avoiding: a x rule:fun_typ.strong_induct)
- case (AF_fun_typ x1a x2a x3a x4a x5a)
- then show ?case unfolding subst_ft_v.simps subst_v_fun_typ_def fun_typ.fresh using forget_subst_v subst_ft_v.simps subst_v_c_def forget_subst_sv subst_v_\<tau>_def by fastforce
- qed
-
- fix p::perm and x1::x and v::v and t1::fun_typ
- show "p \<bullet> subst_v t1 x1 v = subst_v (p \<bullet> t1) (p \<bullet> x1) (p \<bullet> v)"
- proof(nominal_induct t1 avoiding: x1 v rule:fun_typ.strong_induct)
- case (AF_fun_typ x1a x2a x3a x4a x5a)
- then show ?case unfolding subst_ft_v.simps subst_v_fun_typ_def fun_typ.fresh using forget_subst_v subst_ft_v.simps subst_v_c_def forget_subst_sv subst_v_\<tau>_def by fastforce
- qed
-
- fix x::x and c::fun_typ and z::x
- show "atom x \<sharp> c \<Longrightarrow> ((x \<leftrightarrow> z) \<bullet> c) = c[z::=[x]\<^sup>v]\<^sub>v"
- apply(nominal_induct c avoiding: x z rule:fun_typ.strong_induct)
- by (auto simp add: subst_v_c_def subst_v_s_def subst_v_\<tau>_def subst_v_fun_typ_def)
-
- fix x::x and c::fun_typ and z::x
- show "atom x \<sharp> c \<Longrightarrow> c[z::=[x]\<^sup>v]\<^sub>v[x::=v]\<^sub>v = c[z::=v]\<^sub>v"
- apply(nominal_induct c avoiding: z x v rule:fun_typ.strong_induct)
- apply auto
- by (auto simp add: subst_v_c_def subst_v_s_def subst_v_\<tau>_def subst_v_fun_typ_def )
-qed
-end
-
-instantiation fun_typ_q :: has_subst_v
-begin
-
-definition
- "subst_v = subst_ftq_v"
-
-instance proof
- fix j::atom and i::x and x::v and t::fun_typ_q
- show "(j \<sharp> subst_v t i x) = ((atom i \<sharp> t \<and> j \<sharp> t) \<or> (j \<sharp> x \<and> (j \<sharp> t \<or> j = atom i)))"
- apply(nominal_induct t avoiding: i x rule:fun_typ_q.strong_induct,auto)
- apply(auto simp add: subst_v_fun_typ_def subst_v_s_def subst_v_\<tau>_def subst_v_fun_typ_q_def fresh_subst_v_if )
- by (metis (no_types) fresh_subst_v_if subst_v_fun_typ_def)+
-
- fix i::x and t::fun_typ_q and x::v
- show "atom i \<sharp> t \<Longrightarrow> subst_v t i x = t"
- apply(nominal_induct t avoiding: i x rule:fun_typ_q.strong_induct,auto)
- by(auto simp add: subst_v_fun_typ_def subst_v_s_def subst_v_\<tau>_def subst_v_fun_typ_q_def fresh_subst_v_if )
-
- fix i::x and t::fun_typ_q
- show "subst_v t i (V_var i) = t" using subst_cv_id subst_v_fun_typ_def
- apply(nominal_induct t avoiding: i x rule:fun_typ_q.strong_induct,auto)
- by(auto simp add: subst_v_fun_typ_def subst_v_s_def subst_v_\<tau>_def subst_v_fun_typ_q_def fresh_subst_v_if )
-
- fix p::perm and x1::x and v::v and t1::fun_typ_q
- show "p \<bullet> subst_v t1 x1 v = subst_v (p \<bullet> t1) (p \<bullet> x1) (p \<bullet> v)"
- apply(nominal_induct t1 avoiding: v x1 rule:fun_typ_q.strong_induct,auto)
- by(auto simp add: subst_v_fun_typ_def subst_v_s_def subst_v_\<tau>_def subst_v_fun_typ_q_def fresh_subst_v_if )
-
- fix x::x and c::fun_typ_q and z::x
- show "atom x \<sharp> c \<Longrightarrow> ((x \<leftrightarrow> z) \<bullet> c) = c[z::=[x]\<^sup>v]\<^sub>v"
- apply(nominal_induct c avoiding: x z rule:fun_typ_q.strong_induct,auto)
- by(auto simp add: subst_v_fun_typ_def subst_v_s_def subst_v_\<tau>_def subst_v_fun_typ_q_def fresh_subst_v_if )
-
- fix x::x and c::fun_typ_q and z::x
- show "atom x \<sharp> c \<Longrightarrow> c[z::=[x]\<^sup>v]\<^sub>v[x::=v]\<^sub>v = c[z::=v]\<^sub>v"
- apply(nominal_induct c avoiding: z x v rule:fun_typ_q.strong_induct,auto)
- apply(auto simp add: subst_v_fun_typ_def subst_v_s_def subst_v_\<tau>_def subst_v_fun_typ_q_def fresh_subst_v_if )
- by (metis subst_v_fun_typ_def flip_bv_x_cancel subst_ft_v.eqvt subst_v_simple_commute v.perm_simps )+
-qed
-
-end
-
-section \<open>Variable Context\<close>
-
-lemma subst_dv_fst_eq:
- "fst ` setD (\<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v) = fst ` setD \<Delta>"
- by(induct \<Delta> rule: \<Delta>_induct,simp,force)
-
-lemma subst_gv_member_iff:
- fixes x'::x and x::x and v::v and c'::c
- assumes "(x',b',c') \<in> toSet \<Gamma>" and "atom x \<notin> atom_dom \<Gamma>"
- shows "(x',b',c'[x::=v]\<^sub>c\<^sub>v) \<in> toSet \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v"
-proof -
- have "x' \<noteq> x" using assms fresh_dom_free2 by metis
- then show ?thesis using assms proof(induct \<Gamma> rule: \<Gamma>_induct)
- case GNil
- then show ?case by auto
- next
- case (GCons x1 b1 c1 \<Gamma>')
- show ?case proof(cases "(x',b',c') = (x1,b1,c1)")
- case True
- hence "((x1, b1, c1) #\<^sub>\<Gamma> \<Gamma>')[x::=v]\<^sub>\<Gamma>\<^sub>v = ((x1, b1, c1[x::=v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> (\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v))" using subst_gv.simps \<open>x'\<noteq>x\<close> by auto
- then show ?thesis using True by auto
- next
- case False
- have "x1\<noteq>x" using fresh_def fresh_GCons fresh_Pair supp_at_base GCons fresh_dom_free2 by auto
- hence "(x', b', c') \<in> toSet \<Gamma>'" using GCons False toSet.simps by auto
- moreover have "atom x \<notin> atom_dom \<Gamma>'" using fresh_GCons GCons dom.simps toSet.simps by simp
- ultimately have "(x', b', c'[x::=v]\<^sub>c\<^sub>v) \<in> toSet \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v" using GCons by auto
- hence "(x', b', c'[x::=v]\<^sub>c\<^sub>v) \<in> toSet ((x1, b1, c1[x::=v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> (\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v))" by auto
- then show ?thesis using subst_gv.simps \<open>x1\<noteq>x\<close> by auto
- qed
- qed
-qed
-
-lemma fresh_subst_gv_if:
- fixes j::atom and i::x and x::v and t::\<Gamma>
- assumes "j \<sharp> t \<and> j \<sharp> x"
- shows "(j \<sharp> subst_gv t i x)"
- using assms proof(induct t rule: \<Gamma>_induct)
- case GNil
- then show ?case using subst_gv.simps fresh_GNil by auto
-next
- case (GCons x' b' c' \<Gamma>')
- then show ?case unfolding subst_gv.simps using fresh_GCons fresh_subst_cv_if by auto
-qed
-
-section \<open>Lookup\<close>
-
-lemma set_GConsD: "y \<in> toSet (x #\<^sub>\<Gamma> xs) \<Longrightarrow> y=x \<or> y \<in> toSet xs"
- by auto
-
-lemma subst_g_assoc_cons:
- assumes "x \<noteq> x'"
- shows "(((x', b', c') #\<^sub>\<Gamma> \<Gamma>')[x::=v]\<^sub>\<Gamma>\<^sub>v @ G) = ((x', b', c'[x::=v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> ((\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v) @ G))"
- using subst_gv.simps append_g.simps assms by auto
-
+(*<*)
+theory IVSubst
+ imports Syntax
+begin
+ (*>*)
+
+chapter \<open>Immutable Variable Substitution\<close>
+
+text \<open>Substitution involving immutable variables. We define a class and instances for all
+of the term forms\<close>
+
+section \<open>Class\<close>
+
+class has_subst_v = fs +
+ fixes subst_v :: "'a::fs \<Rightarrow> x \<Rightarrow> v \<Rightarrow> 'a::fs" ("_[_::=_]\<^sub>v" [1000,50,50] 1000)
+ assumes fresh_subst_v_if: "y \<sharp> (subst_v a x v) \<longleftrightarrow> (atom x \<sharp> a \<and> y \<sharp> a) \<or> (y \<sharp> v \<and> (y \<sharp> a \<or> y = atom x))"
+ and forget_subst_v[simp]: "atom x \<sharp> a \<Longrightarrow> subst_v a x v = a"
+ and subst_v_id[simp]: "subst_v a x (V_var x) = a"
+ and eqvt[simp,eqvt]: "(p::perm) \<bullet> (subst_v a x v) = (subst_v (p \<bullet> a) (p \<bullet>x) (p \<bullet>v))"
+ and flip_subst_v[simp]: "atom x \<sharp> c \<Longrightarrow> ((x \<leftrightarrow> z) \<bullet> c) = c[z::=[x]\<^sup>v]\<^sub>v"
+ and subst_v_simple_commute[simp]: "atom x \<sharp> c \<Longrightarrow>(c[z::=[x]\<^sup>v]\<^sub>v)[x::=b]\<^sub>v = c[z::=b]\<^sub>v"
+begin
+
+lemma subst_v_flip_eq_one:
+ fixes z1::x and z2::x and x1::x and x2::x
+ assumes "[[atom z1]]lst. c1 = [[atom z2]]lst. c2"
+ and "atom x1 \<sharp> (z1,z2,c1,c2)"
+ shows "(c1[z1::=[x1]\<^sup>v]\<^sub>v) = (c2[z2::=[x1]\<^sup>v]\<^sub>v)"
+proof -
+ have "(c1[z1::=[x1]\<^sup>v]\<^sub>v) = (x1 \<leftrightarrow> z1) \<bullet> c1" using assms flip_subst_v by auto
+ moreover have "(c2[z2::=[x1]\<^sup>v]\<^sub>v) = (x1 \<leftrightarrow> z2) \<bullet> c2" using assms flip_subst_v by auto
+ ultimately show ?thesis using Abs1_eq_iff_all(3)[of z1 c1 z2 c2 z1] assms
+ by (metis Abs1_eq_iff_fresh(3) flip_commute)
+qed
+
+lemma subst_v_flip_eq_two:
+ fixes z1::x and z2::x and x1::x and x2::x
+ assumes "[[atom z1]]lst. c1 = [[atom z2]]lst. c2"
+ shows "(c1[z1::=b]\<^sub>v) = (c2[z2::=b]\<^sub>v)"
+proof -
+ obtain x::x where *:"atom x \<sharp> (z1,z2,c1,c2)" using obtain_fresh by metis
+ hence "(c1[z1::=[x]\<^sup>v]\<^sub>v) = (c2[z2::=[x]\<^sup>v]\<^sub>v)" using subst_v_flip_eq_one[OF assms, of x] by metis
+ hence "(c1[z1::=[x]\<^sup>v]\<^sub>v)[x::=b]\<^sub>v = (c2[z2::=[x]\<^sup>v]\<^sub>v)[x::=b]\<^sub>v" by auto
+ thus ?thesis using subst_v_simple_commute * fresh_prod4 by metis
+qed
+
+lemma subst_v_flip_eq_three:
+ assumes "[[atom z1]]lst. c1 = [[atom z1']]lst. c1'" and "atom x \<sharp> c1" and "atom x' \<sharp> (x,z1,z1', c1, c1')"
+ shows "(x \<leftrightarrow> x') \<bullet> (c1[z1::=[x]\<^sup>v]\<^sub>v) = c1'[z1'::=[x']\<^sup>v]\<^sub>v"
+proof -
+ have "atom x' \<sharp> c1[z1::=[x]\<^sup>v]\<^sub>v" using assms fresh_subst_v_if by simp
+ hence "(x \<leftrightarrow> x') \<bullet> (c1[z1::=[x]\<^sup>v]\<^sub>v) = c1[z1::=[x]\<^sup>v]\<^sub>v[x::=[x']\<^sup>v]\<^sub>v" using flip_subst_v[of x' "c1[z1::=[x]\<^sup>v]\<^sub>v" x] flip_commute by metis
+ also have "... = c1[z1::=[x']\<^sup>v]\<^sub>v" using subst_v_simple_commute fresh_prod4 assms by auto
+ also have "... = c1'[z1'::=[x']\<^sup>v]\<^sub>v" using subst_v_flip_eq_one[of z1 c1 z1' c1' x'] using assms by auto
+ finally show ?thesis by auto
+qed
+
+end
+
+section \<open>Values\<close>
+
+nominal_function
+ subst_vv :: "v \<Rightarrow> x \<Rightarrow> v \<Rightarrow> v" where
+ "subst_vv (V_lit l) x v = V_lit l"
+| "subst_vv (V_var y) x v = (if x = y then v else V_var y)"
+| "subst_vv (V_cons tyid c v') x v = V_cons tyid c (subst_vv v' x v)"
+| "subst_vv (V_consp tyid c b v') x v = V_consp tyid c b (subst_vv v' x v)"
+| "subst_vv (V_pair v1 v2) x v = V_pair (subst_vv v1 x v ) (subst_vv v2 x v )"
+ by(auto simp: eqvt_def subst_vv_graph_aux_def, metis v.strong_exhaust)
+nominal_termination (eqvt) by lexicographic_order
+
+abbreviation
+ subst_vv_abbrev :: "v \<Rightarrow> x \<Rightarrow> v \<Rightarrow> v" ("_[_::=_]\<^sub>v\<^sub>v" [1000,50,50] 1000)
+ where
+ "v[x::=v']\<^sub>v\<^sub>v \<equiv> subst_vv v x v'"
+
+lemma fresh_subst_vv_if [simp]:
+ "j \<sharp> t[i::=x]\<^sub>v\<^sub>v = ((atom i \<sharp> t \<and> j \<sharp> t) \<or> (j \<sharp> x \<and> (j \<sharp> t \<or> j = atom i)))"
+ using supp_l_empty apply (induct t rule: v.induct,auto simp add: subst_vv.simps fresh_def, auto)
+ by (simp add: supp_at_base |metis b.supp supp_b_empty )+
+
+lemma forget_subst_vv [simp]: "atom a \<sharp> tm \<Longrightarrow> tm[a::=x]\<^sub>v\<^sub>v = tm"
+ by (induct tm rule: v.induct) (simp_all add: fresh_at_base)
+
+lemma subst_vv_id [simp]: "tm[a::=V_var a]\<^sub>v\<^sub>v = tm"
+ by (induct tm rule: v.induct) simp_all
+
+lemma subst_vv_commute [simp]:
+ "atom j \<sharp> tm \<Longrightarrow> tm[i::=t]\<^sub>v\<^sub>v[j::=u]\<^sub>v\<^sub>v = tm[i::=t[j::=u]\<^sub>v\<^sub>v]\<^sub>v\<^sub>v "
+ by (induct tm rule: v.induct) (auto simp: fresh_Pair)
+
+lemma subst_vv_commute_full [simp]:
+ "atom j \<sharp> t \<Longrightarrow> atom i \<sharp> u \<Longrightarrow> i \<noteq> j \<Longrightarrow> tm[i::=t]\<^sub>v\<^sub>v[j::=u]\<^sub>v\<^sub>v = tm[j::=u]\<^sub>v\<^sub>v[i::=t]\<^sub>v\<^sub>v"
+ by (induct tm rule: v.induct) auto
+
+lemma subst_vv_var_flip[simp]:
+ fixes v::v
+ assumes "atom y \<sharp> v"
+ shows "(y \<leftrightarrow> x) \<bullet> v = v [x::=V_var y]\<^sub>v\<^sub>v"
+ using assms apply(induct v rule:v.induct)
+ apply auto
+ using l.fresh l.perm_simps l.strong_exhaust supp_l_empty permute_pure permute_list.simps fresh_def flip_fresh_fresh apply fastforce
+ using permute_pure apply blast+
+ done
+
+instantiation v :: has_subst_v
+begin
+
+definition
+ "subst_v = subst_vv"
+
+instance proof
+ fix j::atom and i::x and x::v and t::v
+ show "(j \<sharp> subst_v t i x) = ((atom i \<sharp> t \<and> j \<sharp> t) \<or> (j \<sharp> x \<and> (j \<sharp> t \<or> j = atom i)))"
+ using fresh_subst_vv_if[of j t i x] subst_v_v_def by metis
+
+ fix a::x and tm::v and x::v
+ show "atom a \<sharp> tm \<Longrightarrow> subst_v tm a x = tm"
+ using forget_subst_vv subst_v_v_def by simp
+
+ fix a::x and tm::v
+ show "subst_v tm a (V_var a) = tm" using subst_vv_id subst_v_v_def by simp
+
+ fix p::perm and x1::x and v::v and t1::v
+ show "p \<bullet> subst_v t1 x1 v = subst_v (p \<bullet> t1) (p \<bullet> x1) (p \<bullet> v)"
+ using subst_v_v_def by simp
+
+ fix x::x and c::v and z::x
+ show "atom x \<sharp> c \<Longrightarrow> ((x \<leftrightarrow> z) \<bullet> c) = c[z::=[x]\<^sup>v]\<^sub>v"
+ using subst_v_v_def by simp
+
+ fix x::x and c::v and z::x
+ show "atom x \<sharp> c \<Longrightarrow> c[z::=[x]\<^sup>v]\<^sub>v[x::=v]\<^sub>v = c[z::=v]\<^sub>v"
+ using subst_v_v_def by simp
+qed
+
+end
+
+section \<open>Expressions\<close>
+
+nominal_function subst_ev :: "e \<Rightarrow> x \<Rightarrow> v \<Rightarrow> e" where
+ "subst_ev ( (AE_val v') ) x v = ( (AE_val (subst_vv v' x v)) )"
+| "subst_ev ( (AE_app f v') ) x v = ( (AE_app f (subst_vv v' x v )) )"
+| "subst_ev ( (AE_appP f b v') ) x v = ( (AE_appP f b (subst_vv v' x v )) )"
+| "subst_ev ( (AE_op opp v1 v2) ) x v = ( (AE_op opp (subst_vv v1 x v ) (subst_vv v2 x v )) )"
+| "subst_ev [#1 v']\<^sup>e x v = [#1 (subst_vv v' x v )]\<^sup>e"
+| "subst_ev [#2 v']\<^sup>e x v = [#2 (subst_vv v' x v )]\<^sup>e"
+| "subst_ev ( (AE_mvar u)) x v = AE_mvar u"
+| "subst_ev [| v' |]\<^sup>e x v = [| (subst_vv v' x v ) |]\<^sup>e"
+| "subst_ev ( AE_concat v1 v2) x v = AE_concat (subst_vv v1 x v ) (subst_vv v2 x v )"
+| "subst_ev ( AE_split v1 v2) x v = AE_split (subst_vv v1 x v ) (subst_vv v2 x v )"
+ by(simp add: eqvt_def subst_ev_graph_aux_def,auto)(meson e.strong_exhaust)
+
+nominal_termination (eqvt) by lexicographic_order
+
+abbreviation
+ subst_ev_abbrev :: "e \<Rightarrow> x \<Rightarrow> v \<Rightarrow> e" ("_[_::=_]\<^sub>e\<^sub>v" [1000,50,50] 500)
+ where
+ "e[x::=v']\<^sub>e\<^sub>v \<equiv> subst_ev e x v' "
+
+lemma size_subst_ev [simp]: "size ( subst_ev A i x) = size A"
+ apply (nominal_induct A avoiding: i x rule: e.strong_induct)
+ by auto
+
+lemma forget_subst_ev [simp]: "atom a \<sharp> A \<Longrightarrow> subst_ev A a x = A"
+ apply (nominal_induct A avoiding: a x rule: e.strong_induct)
+ by (auto simp: fresh_at_base)
+
+lemma subst_ev_id [simp]: "subst_ev A a (V_var a) = A"
+ by (nominal_induct A avoiding: a rule: e.strong_induct) (auto simp: fresh_at_base)
+
+lemma fresh_subst_ev_if [simp]:
+ "j \<sharp> (subst_ev A i x ) = ((atom i \<sharp> A \<and> j \<sharp> A) \<or> (j \<sharp> x \<and> (j \<sharp> A \<or> j = atom i)))"
+ apply (induct A rule: e.induct)
+ unfolding subst_ev.simps fresh_subst_vv_if apply auto+
+ using pure_fresh fresh_opp_all apply metis+
+ done
+
+lemma subst_ev_commute [simp]:
+ "atom j \<sharp> A \<Longrightarrow> (A[i::=t]\<^sub>e\<^sub>v)[j::=u]\<^sub>e\<^sub>v = A[i::=t[j::=u]\<^sub>v\<^sub>v]\<^sub>e\<^sub>v"
+ by (nominal_induct A avoiding: i j t u rule: e.strong_induct) (auto simp: fresh_at_base)
+
+lemma subst_ev_var_flip[simp]:
+ fixes e::e and y::x and x::x
+ assumes "atom y \<sharp> e"
+ shows "(y \<leftrightarrow> x) \<bullet> e = e [x::=V_var y]\<^sub>e\<^sub>v"
+ using assms apply(nominal_induct e rule:e.strong_induct)
+ apply (simp add: subst_v_v_def)
+ apply (metis (mono_tags, lifting) b.eq_iff b.perm_simps e.fresh e.perm_simps flip_b_id subst_ev.simps subst_vv_var_flip)
+ apply (metis (mono_tags, lifting) b.eq_iff b.perm_simps e.fresh e.perm_simps flip_b_id subst_ev.simps subst_vv_var_flip)
+ subgoal for x
+ apply (rule_tac y=x in opp.strong_exhaust)
+ using subst_vv_var_flip flip_def by (simp add: flip_def permute_pure)+
+ using subst_vv_var_flip flip_def by (simp add: flip_def permute_pure)+
+
+lemma subst_ev_flip:
+ fixes e::e and ea::e and c::x
+ assumes "atom c \<sharp> (e, ea)" and "atom c \<sharp> (x, xa, e, ea)" and "(x \<leftrightarrow> c) \<bullet> e = (xa \<leftrightarrow> c) \<bullet> ea"
+ shows "e[x::=v']\<^sub>e\<^sub>v = ea[xa::=v']\<^sub>e\<^sub>v"
+proof -
+ have "e[x::=v']\<^sub>e\<^sub>v = (e[x::=V_var c]\<^sub>e\<^sub>v)[c::=v']\<^sub>e\<^sub>v" using subst_ev_commute assms by simp
+ also have "... = ((c \<leftrightarrow> x) \<bullet> e)[c::=v']\<^sub>e\<^sub>v" using subst_ev_var_flip assms by simp
+ also have "... = ((c \<leftrightarrow> xa) \<bullet> ea)[c::=v']\<^sub>e\<^sub>v" using assms flip_commute by metis
+ also have "... = ea[xa::=v']\<^sub>e\<^sub>v" using subst_ev_var_flip assms by simp
+ finally show ?thesis by auto
+qed
+
+lemma subst_ev_var[simp]:
+ "(AE_val (V_var x))[x::=[z]\<^sup>v]\<^sub>e\<^sub>v = AE_val (V_var z)"
+ by auto
+
+instantiation e :: has_subst_v
+begin
+
+definition
+ "subst_v = subst_ev"
+
+instance proof
+ fix j::atom and i::x and x::v and t::e
+ show "(j \<sharp> subst_v t i x) = ((atom i \<sharp> t \<and> j \<sharp> t) \<or> (j \<sharp> x \<and> (j \<sharp> t \<or> j = atom i)))"
+ using fresh_subst_ev_if[of j t i x] subst_v_e_def by metis
+
+ fix a::x and tm::e and x::v
+ show "atom a \<sharp> tm \<Longrightarrow> subst_v tm a x = tm"
+ using forget_subst_ev subst_v_e_def by simp
+
+ fix a::x and tm::e
+ show "subst_v tm a (V_var a) = tm" using subst_ev_id subst_v_e_def by simp
+
+ fix p::perm and x1::x and v::v and t1::e
+ show "p \<bullet> subst_v t1 x1 v = subst_v (p \<bullet> t1) (p \<bullet> x1) (p \<bullet> v)"
+ using subst_ev_commute subst_v_e_def by simp
+
+ fix x::x and c::e and z::x
+ show "atom x \<sharp> c \<Longrightarrow> ((x \<leftrightarrow> z) \<bullet> c) = c[z::=[x]\<^sup>v]\<^sub>v"
+ using subst_v_e_def by simp
+
+ fix x::x and c::e and z::x
+ show "atom x \<sharp> c \<Longrightarrow> c[z::=[x]\<^sup>v]\<^sub>v[x::=v]\<^sub>v = c[z::=v]\<^sub>v"
+ using subst_v_e_def by simp
+qed
+end
+
+lemma subst_ev_commute_full:
+ fixes e::e and w::v and v::v
+ assumes "atom z \<sharp> v" and "atom x \<sharp> w" and "x \<noteq> z"
+ shows "subst_ev (e[z::=w]\<^sub>e\<^sub>v) x v = subst_ev (e[x::=v]\<^sub>e\<^sub>v) z w"
+ using assms by(nominal_induct e rule: e.strong_induct,simp+)
+
+lemma subst_ev_v_flip1[simp]:
+ fixes e::e
+ assumes "atom z1 \<sharp> (z,e)" and "atom z1' \<sharp> (z,e)"
+ shows"(z1 \<leftrightarrow> z1') \<bullet> e[z::=v]\<^sub>e\<^sub>v = e[z::= ((z1 \<leftrightarrow> z1') \<bullet> v)]\<^sub>e\<^sub>v"
+ using assms proof(nominal_induct e rule:e.strong_induct)
+qed (simp add: flip_def fresh_Pair swap_fresh_fresh)+
+
+section \<open>Expressions in Constraints\<close>
+
+nominal_function subst_cev :: "ce \<Rightarrow> x \<Rightarrow> v \<Rightarrow> ce" where
+ "subst_cev ( (CE_val v') ) x v = ( (CE_val (subst_vv v' x v )) )"
+| "subst_cev ( (CE_op opp v1 v2) ) x v = ( (CE_op opp (subst_cev v1 x v ) (subst_cev v2 x v )) )"
+| "subst_cev ( (CE_fst v')) x v = CE_fst (subst_cev v' x v )"
+| "subst_cev ( (CE_snd v')) x v = CE_snd (subst_cev v' x v )"
+| "subst_cev ( (CE_len v')) x v = CE_len (subst_cev v' x v )"
+| "subst_cev ( CE_concat v1 v2) x v = CE_concat (subst_cev v1 x v ) (subst_cev v2 x v )"
+ apply (simp add: eqvt_def subst_cev_graph_aux_def,auto)
+ by (meson ce.strong_exhaust)
+
+nominal_termination (eqvt) by lexicographic_order
+
+abbreviation
+ subst_cev_abbrev :: "ce \<Rightarrow> x \<Rightarrow> v \<Rightarrow> ce" ("_[_::=_]\<^sub>c\<^sub>e\<^sub>v" [1000,50,50] 500)
+ where
+ "e[x::=v']\<^sub>c\<^sub>e\<^sub>v \<equiv> subst_cev e x v'"
+
+lemma size_subst_cev [simp]: "size ( subst_cev A i x ) = size A"
+ by (nominal_induct A avoiding: i x rule: ce.strong_induct,auto)
+
+lemma forget_subst_cev [simp]: "atom a \<sharp> A \<Longrightarrow> subst_cev A a x = A"
+ by (nominal_induct A avoiding: a x rule: ce.strong_induct, auto simp: fresh_at_base)
+
+lemma subst_cev_id [simp]: "subst_cev A a (V_var a) = A"
+ by (nominal_induct A avoiding: a rule: ce.strong_induct) (auto simp: fresh_at_base)
+
+lemma fresh_subst_cev_if [simp]:
+ "j \<sharp> (subst_cev A i x ) = ((atom i \<sharp> A \<and> j \<sharp> A) \<or> (j \<sharp> x \<and> (j \<sharp> A \<or> j = atom i)))"
+proof(nominal_induct A avoiding: i x rule: ce.strong_induct)
+ case (CE_op opp v1 v2)
+ then show ?case using fresh_subst_vv_if subst_ev.simps e.supp pure_fresh opp.fresh
+ fresh_e_opp
+ using fresh_opp_all by auto
+qed(auto)+
+
+lemma subst_cev_commute [simp]:
+ "atom j \<sharp> A \<Longrightarrow> (subst_cev (subst_cev A i t ) j u) = subst_cev A i (subst_vv t j u )"
+ by (nominal_induct A avoiding: i j t u rule: ce.strong_induct) (auto simp: fresh_at_base)
+
+lemma subst_cev_var_flip[simp]:
+ fixes e::ce and y::x and x::x
+ assumes "atom y \<sharp> e"
+ shows "(y \<leftrightarrow> x) \<bullet> e = e [x::=V_var y]\<^sub>c\<^sub>e\<^sub>v"
+ using assms proof(nominal_induct e rule:ce.strong_induct)
+ case (CE_val v)
+ then show ?case using subst_vv_var_flip by auto
+next
+ case (CE_op opp v1 v2)
+ hence yf: "atom y \<sharp> v1 \<and> atom y \<sharp> v2" using ce.fresh by blast
+ have " (y \<leftrightarrow> x) \<bullet> (CE_op opp v1 v2 ) = CE_op ((y \<leftrightarrow> x) \<bullet> opp) ( (y \<leftrightarrow> x) \<bullet> v1 ) ( (y \<leftrightarrow> x) \<bullet> v2)"
+ using opp.perm_simps ce.perm_simps permute_pure ce.fresh opp.strong_exhaust by presburger
+ also have "... = CE_op ((y \<leftrightarrow> x) \<bullet> opp) (v1[x::=V_var y]\<^sub>c\<^sub>e\<^sub>v) (v2 [x::=V_var y]\<^sub>c\<^sub>e\<^sub>v)" using yf
+ by (simp add: CE_op.hyps(1) CE_op.hyps(2))
+ finally show ?case using subst_cev.simps opp.perm_simps opp.strong_exhaust
+ by (metis (full_types))
+qed( (auto simp add: permute_pure subst_vv_var_flip)+)
+
+lemma subst_cev_flip:
+ fixes e::ce and ea::ce and c::x
+ assumes "atom c \<sharp> (e, ea)" and "atom c \<sharp> (x, xa, e, ea)" and "(x \<leftrightarrow> c) \<bullet> e = (xa \<leftrightarrow> c) \<bullet> ea"
+ shows "e[x::=v']\<^sub>c\<^sub>e\<^sub>v = ea[xa::=v']\<^sub>c\<^sub>e\<^sub>v"
+proof -
+ have "e[x::=v']\<^sub>c\<^sub>e\<^sub>v = (e[x::=V_var c]\<^sub>c\<^sub>e\<^sub>v)[c::=v']\<^sub>c\<^sub>e\<^sub>v" using subst_ev_commute assms by simp
+ also have "... = ((c \<leftrightarrow> x) \<bullet> e)[c::=v']\<^sub>c\<^sub>e\<^sub>v" using subst_ev_var_flip assms by simp
+ also have "... = ((c \<leftrightarrow> xa) \<bullet> ea)[c::=v']\<^sub>c\<^sub>e\<^sub>v" using assms flip_commute by metis
+ also have "... = ea[xa::=v']\<^sub>c\<^sub>e\<^sub>v" using subst_ev_var_flip assms by simp
+ finally show ?thesis by auto
+qed
+
+lemma subst_cev_var[simp]:
+ fixes z::x and x::x
+ shows "[[x]\<^sup>v]\<^sup>c\<^sup>e [x::=[z]\<^sup>v]\<^sub>c\<^sub>e\<^sub>v = [[z]\<^sup>v]\<^sup>c\<^sup>e"
+ by auto
+
+instantiation ce :: has_subst_v
+begin
+
+definition
+ "subst_v = subst_cev"
+
+instance proof
+ fix j::atom and i::x and x::v and t::ce
+ show "(j \<sharp> subst_v t i x) = ((atom i \<sharp> t \<and> j \<sharp> t) \<or> (j \<sharp> x \<and> (j \<sharp> t \<or> j = atom i)))"
+ using fresh_subst_cev_if[of j t i x] subst_v_ce_def by metis
+
+ fix a::x and tm::ce and x::v
+ show "atom a \<sharp> tm \<Longrightarrow> subst_v tm a x = tm"
+ using forget_subst_cev subst_v_ce_def by simp
+
+ fix a::x and tm::ce
+ show "subst_v tm a (V_var a) = tm" using subst_cev_id subst_v_ce_def by simp
+
+ fix p::perm and x1::x and v::v and t1::ce
+ show "p \<bullet> subst_v t1 x1 v = subst_v (p \<bullet> t1) (p \<bullet> x1) (p \<bullet> v)"
+ using subst_cev_commute subst_v_ce_def by simp
+
+ fix x::x and c::ce and z::x
+ show "atom x \<sharp> c \<Longrightarrow> ((x \<leftrightarrow> z) \<bullet> c) = c [z::=V_var x]\<^sub>v"
+ using subst_v_ce_def by simp
+
+ fix x::x and c::ce and z::x
+ show "atom x \<sharp> c \<Longrightarrow> c [z::=V_var x]\<^sub>v[x::=v]\<^sub>v = c[z::=v]\<^sub>v"
+ using subst_v_ce_def by simp
+qed
+
+end
+
+lemma subst_cev_commute_full:
+ fixes e::ce and w::v and v::v
+ assumes "atom z \<sharp> v" and "atom x \<sharp> w" and "x \<noteq> z"
+ shows "subst_cev (e[z::=w]\<^sub>c\<^sub>e\<^sub>v) x v = subst_cev (e[x::=v]\<^sub>c\<^sub>e\<^sub>v) z w "
+ using assms by(nominal_induct e rule: ce.strong_induct,simp+)
+
+
+lemma subst_cev_v_flip1[simp]:
+ fixes e::ce
+ assumes "atom z1 \<sharp> (z,e)" and "atom z1' \<sharp> (z,e)"
+ shows"(z1 \<leftrightarrow> z1') \<bullet> e[z::=v]\<^sub>c\<^sub>e\<^sub>v = e[z::= ((z1 \<leftrightarrow> z1') \<bullet> v)]\<^sub>c\<^sub>e\<^sub>v"
+ using assms apply(nominal_induct e rule:ce.strong_induct)
+ by (simp add: flip_def fresh_Pair swap_fresh_fresh)+
+
+section \<open>Constraints\<close>
+
+nominal_function subst_cv :: "c \<Rightarrow> x \<Rightarrow> v \<Rightarrow> c" where
+ "subst_cv (C_true) x v = C_true"
+| "subst_cv (C_false) x v = C_false"
+| "subst_cv (C_conj c1 c2) x v = C_conj (subst_cv c1 x v ) (subst_cv c2 x v )"
+| "subst_cv (C_disj c1 c2) x v = C_disj (subst_cv c1 x v ) (subst_cv c2 x v )"
+| "subst_cv (C_imp c1 c2) x v = C_imp (subst_cv c1 x v ) (subst_cv c2 x v )"
+| "subst_cv (e1 == e2) x v = ((subst_cev e1 x v ) == (subst_cev e2 x v ))"
+| "subst_cv (C_not c) x v = C_not (subst_cv c x v )"
+ apply (simp add: eqvt_def subst_cv_graph_aux_def,auto)
+ using c.strong_exhaust by metis
+nominal_termination (eqvt) by lexicographic_order
+
+abbreviation
+ subst_cv_abbrev :: "c \<Rightarrow> x \<Rightarrow> v \<Rightarrow> c" ("_[_::=_]\<^sub>c\<^sub>v" [1000,50,50] 1000)
+ where
+ "c[x::=v']\<^sub>c\<^sub>v \<equiv> subst_cv c x v'"
+
+lemma size_subst_cv [simp]: "size ( subst_cv A i x ) = size A"
+ by (nominal_induct A avoiding: i x rule: c.strong_induct,auto)
+
+lemma forget_subst_cv [simp]: "atom a \<sharp> A \<Longrightarrow> subst_cv A a x = A"
+ by (nominal_induct A avoiding: a x rule: c.strong_induct, auto simp: fresh_at_base)
+
+lemma subst_cv_id [simp]: "subst_cv A a (V_var a) = A"
+ by (nominal_induct A avoiding: a rule: c.strong_induct) (auto simp: fresh_at_base)
+
+lemma fresh_subst_cv_if [simp]:
+ "j \<sharp> (subst_cv A i x ) \<longleftrightarrow> (atom i \<sharp> A \<and> j \<sharp> A) \<or> (j \<sharp> x \<and> (j \<sharp> A \<or> j = atom i))"
+ by (nominal_induct A avoiding: i x rule: c.strong_induct, (auto simp add: pure_fresh)+)
+
+lemma subst_cv_commute [simp]:
+ "atom j \<sharp> A \<Longrightarrow> (subst_cv (subst_cv A i t ) j u ) = subst_cv A i (subst_vv t j u )"
+ by (nominal_induct A avoiding: i j t u rule: c.strong_induct) (auto simp: fresh_at_base)
+
+lemma let_s_size [simp]: "size s \<le> size (AS_let x e s)"
+ apply (nominal_induct s avoiding: e x rule: s_branch_s_branch_list.strong_induct(1))
+ apply auto
+ done
+
+lemma subst_cv_var_flip[simp]:
+ fixes c::c
+ assumes "atom y \<sharp> c"
+ shows "(y \<leftrightarrow> x) \<bullet> c = c[x::=V_var y]\<^sub>c\<^sub>v"
+ using assms by(nominal_induct c rule:c.strong_induct,(simp add: flip_subst_v subst_v_ce_def)+)
+
+instantiation c :: has_subst_v
+begin
+
+definition
+ "subst_v = subst_cv"
+
+instance proof
+ fix j::atom and i::x and x::v and t::c
+ show "(j \<sharp> subst_v t i x) = ((atom i \<sharp> t \<and> j \<sharp> t) \<or> (j \<sharp> x \<and> (j \<sharp> t \<or> j = atom i)))"
+ using fresh_subst_cv_if[of j t i x] subst_v_c_def by metis
+
+ fix a::x and tm::c and x::v
+ show "atom a \<sharp> tm \<Longrightarrow> subst_v tm a x = tm"
+ using forget_subst_cv subst_v_c_def by simp
+
+ fix a::x and tm::c
+ show "subst_v tm a (V_var a) = tm" using subst_cv_id subst_v_c_def by simp
+
+ fix p::perm and x1::x and v::v and t1::c
+ show "p \<bullet> subst_v t1 x1 v = subst_v (p \<bullet> t1) (p \<bullet> x1) (p \<bullet> v)"
+ using subst_cv_commute subst_v_c_def by simp
+
+ fix x::x and c::c and z::x
+ show "atom x \<sharp> c \<Longrightarrow> ((x \<leftrightarrow> z) \<bullet> c) = c[z::=[x]\<^sup>v]\<^sub>v"
+ using subst_cv_var_flip subst_v_c_def by simp
+
+ fix x::x and c::c and z::x
+ show "atom x \<sharp> c \<Longrightarrow> c[z::=[x]\<^sup>v]\<^sub>v[x::=v]\<^sub>v = c[z::=v]\<^sub>v"
+ using subst_cv_var_flip subst_v_c_def by simp
+qed
+
+end
+
+lemma subst_cv_var_flip1[simp]:
+ fixes c::c
+ assumes "atom y \<sharp> c"
+ shows "(x \<leftrightarrow> y) \<bullet> c = c[x::=V_var y]\<^sub>c\<^sub>v"
+ using subst_cv_var_flip flip_commute
+ by (metis assms)
+
+lemma subst_cv_v_flip3[simp]:
+ fixes c::c
+ assumes "atom z1 \<sharp> c" and "atom z1' \<sharp> c"
+ shows"(z1 \<leftrightarrow> z1') \<bullet> c[z::=[z1]\<^sup>v]\<^sub>c\<^sub>v = c[z::=[z1']\<^sup>v]\<^sub>c\<^sub>v"
+proof -
+ consider "z1' = z" | "z1 = z" | "atom z1 \<sharp> z \<and> atom z1' \<sharp> z" by force
+ then show ?thesis proof(cases)
+ case 1
+ then show ?thesis using 1 assms by auto
+ next
+ case 2
+ then show ?thesis using 2 assms by auto
+ next
+ case 3
+ then show ?thesis using assms by auto
+ qed
+qed
+
+lemma subst_cv_v_flip[simp]:
+ fixes c::c
+ assumes "atom x \<sharp> c"
+ shows "((x \<leftrightarrow> z) \<bullet> c)[x::=v]\<^sub>c\<^sub>v = c [z::=v]\<^sub>c\<^sub>v"
+ using assms subst_v_c_def by auto
+
+lemma subst_cv_commute_full:
+ fixes c::c
+ assumes "atom z \<sharp> v" and "atom x \<sharp> w" and "x\<noteq>z"
+ shows "(c[z::=w]\<^sub>c\<^sub>v)[x::=v]\<^sub>c\<^sub>v = (c[x::=v]\<^sub>c\<^sub>v)[z::=w]\<^sub>c\<^sub>v"
+ using assms proof(nominal_induct c rule: c.strong_induct)
+ case (C_eq e1 e2)
+ then show ?case using subst_cev_commute_full by simp
+qed(force+)
+
+lemma subst_cv_eq[simp]:
+ assumes "atom z1 \<sharp> e1"
+ shows "(CE_val (V_var z1) == e1 )[z1::=[x]\<^sup>v]\<^sub>c\<^sub>v = (CE_val (V_var x) == e1 )" (is "?A = ?B")
+proof -
+ have "?A = (((CE_val (V_var z1))[z1::=[x]\<^sup>v]\<^sub>c\<^sub>e\<^sub>v) == e1)" using subst_cv.simps assms by simp
+ thus ?thesis by simp
+qed
+
+section \<open>Variable Context\<close>
+
+text \<open>The idea of this substitution is to remove x from the context. We really want to add the condition
+that x is fresh in v but this causes problems with proofs.\<close>
+
+nominal_function subst_gv :: "\<Gamma> \<Rightarrow> x \<Rightarrow> v \<Rightarrow> \<Gamma>" where
+ "subst_gv GNil x v = GNil"
+| "subst_gv ((y,b,c) #\<^sub>\<Gamma> \<Gamma>) x v = (if x = y then \<Gamma> else ((y,b,c[x::=v]\<^sub>c\<^sub>v)#\<^sub>\<Gamma> (subst_gv \<Gamma> x v )))"
+proof(goal_cases)
+ case 1
+ then show ?case by(simp add: eqvt_def subst_gv_graph_aux_def )
+next
+ case (3 P x)
+ then show ?case by (metis neq_GNil_conv prod_cases3)
+qed(fast+)
+nominal_termination (eqvt) by lexicographic_order
+
+abbreviation
+ subst_gv_abbrev :: "\<Gamma> \<Rightarrow> x \<Rightarrow> v \<Rightarrow> \<Gamma>" ("_[_::=_]\<^sub>\<Gamma>\<^sub>v" [1000,50,50] 1000)
+ where
+ "g[x::=v]\<^sub>\<Gamma>\<^sub>v \<equiv> subst_gv g x v"
+
+lemma size_subst_gv [simp]: "size ( subst_gv G i x ) \<le> size G"
+ by (induct G,auto)
+
+lemma forget_subst_gv [simp]: "atom a \<sharp> G \<Longrightarrow> subst_gv G a x = G"
+ apply (induct G ,auto)
+ using fresh_GCons fresh_PairD(1) not_self_fresh apply blast
+ apply (simp add: fresh_GCons)+
+ done
+
+lemma fresh_subst_gv: "atom a \<sharp> G \<Longrightarrow> atom a \<sharp> v \<Longrightarrow> atom a \<sharp> subst_gv G x v"
+proof(induct G)
+ case GNil
+ then show ?case by auto
+next
+ case (GCons xbc G)
+ obtain x' and b' and c' where xbc: "xbc = (x',b',c')" using prod_cases3 by blast
+ show ?case proof(cases "x=x'")
+ case True
+ have "atom a \<sharp> G" using GCons fresh_GCons by blast
+ thus ?thesis using subst_gv.simps(2)[of x' b' c' G] GCons xbc True by presburger
+ next
+ case False
+ then show ?thesis using subst_gv.simps(2)[of x' b' c' G] GCons xbc False fresh_GCons by simp
+ qed
+qed
+
+lemma subst_gv_flip:
+ fixes x::x and xa::x and z::x and c::c and b::b and \<Gamma>::\<Gamma>
+ assumes "atom xa \<sharp> ((x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)" and "atom xa \<sharp> \<Gamma>" and "atom x \<sharp> \<Gamma>" and "atom x \<sharp> (z, c)" and "atom xa \<sharp> (z, c)"
+ shows "(x \<leftrightarrow> xa) \<bullet> ((x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) = (xa, b, c[z::=V_var xa]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>"
+proof -
+ have "(x \<leftrightarrow> xa) \<bullet> ((x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) = (( (x \<leftrightarrow> xa) \<bullet> x, b, (x \<leftrightarrow> xa) \<bullet> c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> ((x \<leftrightarrow> xa) \<bullet> \<Gamma>))"
+ using subst Cons_eqvt flip_fresh_fresh using G_cons_flip by simp
+ also have "... = ((xa, b, (x \<leftrightarrow> xa) \<bullet> c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> ((x \<leftrightarrow> xa) \<bullet> \<Gamma>))" using assms by fastforce
+ also have "... = ((xa, b, c[z::=V_var xa]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> ((x \<leftrightarrow> xa) \<bullet> \<Gamma>))" using assms subst_cv_var_flip by fastforce
+ also have "... = ((xa, b, c[z::=V_var xa]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)" using assms flip_fresh_fresh by blast
+ finally show ?thesis by simp
+qed
+
+section \<open>Types\<close>
+
+nominal_function subst_tv :: "\<tau> \<Rightarrow> x \<Rightarrow> v \<Rightarrow> \<tau>" where
+ "atom z \<sharp> (x,v) \<Longrightarrow> subst_tv \<lbrace> z : b | c \<rbrace> x v = \<lbrace> z : b | c[x::=v]\<^sub>c\<^sub>v \<rbrace>"
+ apply (simp add: eqvt_def subst_tv_graph_aux_def )
+ apply auto
+ subgoal for P a aa b
+ apply(rule_tac y=a and c="(aa,b)" in \<tau>.strong_exhaust)
+ by (auto simp: eqvt_at_def fresh_star_def fresh_Pair fresh_at_base)
+ apply (auto simp: eqvt_at_def fresh_star_def fresh_Pair fresh_at_base)
+proof -
+ fix z :: x and c :: c and za :: x and xa :: x and va :: v and ca :: c and cb :: x
+ assume a1: "atom za \<sharp> va" and a2: "atom z \<sharp> va" and a3: "\<forall>cb. atom cb \<sharp> c \<and> atom cb \<sharp> ca \<longrightarrow> cb \<noteq> z \<and> cb \<noteq> za \<longrightarrow> c[z::=V_var cb]\<^sub>c\<^sub>v = ca[za::=V_var cb]\<^sub>c\<^sub>v"
+ assume a4: "atom cb \<sharp> c" and a5: "atom cb \<sharp> ca" and a6: "cb \<noteq> z" and a7: "cb \<noteq> za" and "atom cb \<sharp> va" and a8: "za \<noteq> xa" and a9: "z \<noteq> xa"
+ assume a10:"cb \<noteq> xa"
+ note assms = a10 a9 a8 a7 a6 a5 a4 a3 a2 a1
+
+ have "c[z::=V_var cb]\<^sub>c\<^sub>v = ca[za::=V_var cb]\<^sub>c\<^sub>v" using assms by auto
+ hence "c[z::=V_var cb]\<^sub>c\<^sub>v[xa::=va]\<^sub>c\<^sub>v = ca[za::=V_var cb]\<^sub>c\<^sub>v[xa::=va]\<^sub>c\<^sub>v" by simp
+ moreover have "c[z::=V_var cb]\<^sub>c\<^sub>v[xa::=va]\<^sub>c\<^sub>v = c[xa::=va]\<^sub>c\<^sub>v[z::=V_var cb]\<^sub>c\<^sub>v" using subst_cv_commute_full[of z va xa "V_var cb" ] assms fresh_def v.supp by fastforce
+ moreover have "ca[za::=V_var cb]\<^sub>c\<^sub>v[xa::=va]\<^sub>c\<^sub>v = ca[xa::=va]\<^sub>c\<^sub>v[za::=V_var cb]\<^sub>c\<^sub>v"
+ using subst_cv_commute_full[of za va xa "V_var cb" ] assms fresh_def v.supp by fastforce
+
+ ultimately show "c[xa::=va]\<^sub>c\<^sub>v[z::=V_var cb]\<^sub>c\<^sub>v = ca[xa::=va]\<^sub>c\<^sub>v[za::=V_var cb]\<^sub>c\<^sub>v" by simp
+qed
+
+nominal_termination (eqvt) by lexicographic_order
+
+abbreviation
+ subst_tv_abbrev :: "\<tau> \<Rightarrow> x \<Rightarrow> v \<Rightarrow> \<tau>" ("_[_::=_]\<^sub>\<tau>\<^sub>v" [1000,50,50] 1000)
+ where
+ "t[x::=v]\<^sub>\<tau>\<^sub>v \<equiv> subst_tv t x v"
+
+lemma size_subst_tv [simp]: "size ( subst_tv A i x ) = size A"
+proof (nominal_induct A avoiding: i x rule: \<tau>.strong_induct)
+ case (T_refined_type x' b' c')
+ then show ?case by auto
+qed
+
+lemma forget_subst_tv [simp]: "atom a \<sharp> A \<Longrightarrow> subst_tv A a x = A"
+ apply (nominal_induct A avoiding: a x rule: \<tau>.strong_induct)
+ apply(auto simp: fresh_at_base)
+ done
+
+lemma subst_tv_id [simp]: "subst_tv A a (V_var a) = A"
+ by (nominal_induct A avoiding: a rule: \<tau>.strong_induct) (auto simp: fresh_at_base)
+
+lemma fresh_subst_tv_if [simp]:
+ "j \<sharp> (subst_tv A i x ) \<longleftrightarrow> (atom i \<sharp> A \<and> j \<sharp> A) \<or> (j \<sharp> x \<and> (j \<sharp> A \<or> j = atom i))"
+ apply (nominal_induct A avoiding: i x rule: \<tau>.strong_induct)
+ using fresh_def supp_b_empty x_fresh_b by auto
+
+lemma subst_tv_commute [simp]:
+ "atom y \<sharp> \<tau> \<Longrightarrow> (\<tau>[x::= t]\<^sub>\<tau>\<^sub>v)[y::=v]\<^sub>\<tau>\<^sub>v = \<tau>[x::= t[y::=v]\<^sub>v\<^sub>v]\<^sub>\<tau>\<^sub>v "
+ by (nominal_induct \<tau> avoiding: x y t v rule: \<tau>.strong_induct) (auto simp: fresh_at_base)
+
+lemma subst_tv_var_flip [simp]:
+ fixes x::x and xa::x and \<tau>::\<tau>
+ assumes "atom xa \<sharp> \<tau>"
+ shows "(x \<leftrightarrow> xa) \<bullet> \<tau> = \<tau>[x::=V_var xa]\<^sub>\<tau>\<^sub>v"
+proof -
+ obtain z::x and b and c where zbc: "atom z \<sharp> (x,xa, V_var xa) \<and> \<tau> = \<lbrace> z : b | c \<rbrace>"
+ using obtain_fresh_z by (metis prod.inject subst_tv.cases)
+ hence "atom xa \<notin> supp c - { atom z }" using \<tau>.supp[of z b c] fresh_def supp_b_empty assms
+ by auto
+ moreover have "xa \<noteq> z" using zbc fresh_prod3 by force
+ ultimately have xaf: "atom xa \<sharp> c" using fresh_def by auto
+ have "(x \<leftrightarrow> xa) \<bullet> \<tau> = \<lbrace> z : b | (x \<leftrightarrow> xa) \<bullet> c \<rbrace>"
+ by (metis \<tau>.perm_simps empty_iff flip_at_base_simps(3) flip_fresh_fresh fresh_PairD(1) fresh_PairD(2) fresh_def not_self_fresh supp_b_empty v.fresh(2) zbc)
+ also have "... = \<lbrace> z : b | c[x::=V_var xa]\<^sub>c\<^sub>v \<rbrace>" using subst_cv_v_flip xaf
+ by (metis permute_flip_cancel permute_flip_cancel2 subst_cv_var_flip)
+ finally show ?thesis using subst_tv.simps zbc
+ using fresh_PairD(1) not_self_fresh by force
+qed
+
+instantiation \<tau> :: has_subst_v
+begin
+
+definition
+ "subst_v = subst_tv"
+
+instance proof
+ fix j::atom and i::x and x::v and t::\<tau>
+ show "(j \<sharp> subst_v t i x) = ((atom i \<sharp> t \<and> j \<sharp> t) \<or> (j \<sharp> x \<and> (j \<sharp> t \<or> j = atom i)))"
+
+ proof(nominal_induct t avoiding: i x rule:\<tau>.strong_induct)
+ case (T_refined_type z b c)
+ hence " j \<sharp> \<lbrace> z : b | c \<rbrace>[i::=x]\<^sub>v = j \<sharp> \<lbrace> z : b | c[i::=x]\<^sub>c\<^sub>v \<rbrace>" using subst_tv.simps subst_v_\<tau>_def fresh_Pair by simp
+ also have "... = (atom i \<sharp> \<lbrace> z : b | c \<rbrace> \<and> j \<sharp> \<lbrace> z : b | c \<rbrace> \<or> j \<sharp> x \<and> (j \<sharp> \<lbrace> z : b | c \<rbrace> \<or> j = atom i))"
+ unfolding \<tau>.fresh using subst_v_c_def fresh_subst_v_if
+ using T_refined_type.hyps(1) T_refined_type.hyps(2) x_fresh_b by auto
+ finally show ?case by auto
+ qed
+
+ fix a::x and tm::\<tau> and x::v
+ show "atom a \<sharp> tm \<Longrightarrow> subst_v tm a x = tm"
+ apply(nominal_induct tm avoiding: a x rule:\<tau>.strong_induct)
+ using subst_v_c_def forget_subst_v subst_tv.simps subst_v_\<tau>_def fresh_Pair by simp
+
+ fix a::x and tm::\<tau>
+ show "subst_v tm a (V_var a) = tm"
+ apply(nominal_induct tm avoiding: a rule:\<tau>.strong_induct)
+ using subst_v_c_def forget_subst_v subst_tv.simps subst_v_\<tau>_def fresh_Pair by simp
+
+ fix p::perm and x1::x and v::v and t1::\<tau>
+ show "p \<bullet> subst_v t1 x1 v = subst_v (p \<bullet> t1) (p \<bullet> x1) (p \<bullet> v)"
+ apply(nominal_induct tm avoiding: a x rule:\<tau>.strong_induct)
+ using subst_v_c_def forget_subst_v subst_tv.simps subst_v_\<tau>_def fresh_Pair by simp
+
+ fix x::x and c::\<tau> and z::x
+ show "atom x \<sharp> c \<Longrightarrow> ((x \<leftrightarrow> z) \<bullet> c) = c[z::=[x]\<^sup>v]\<^sub>v"
+ apply(nominal_induct c avoiding: z x rule:\<tau>.strong_induct)
+ using subst_v_c_def flip_subst_v subst_tv.simps subst_v_\<tau>_def fresh_Pair by auto
+
+ fix x::x and c::\<tau> and z::x
+ show "atom x \<sharp> c \<Longrightarrow> c[z::=[x]\<^sup>v]\<^sub>v[x::=v]\<^sub>v = c[z::=v]\<^sub>v"
+ apply(nominal_induct c avoiding: x v z rule:\<tau>.strong_induct)
+ using subst_v_c_def subst_tv.simps subst_v_\<tau>_def fresh_Pair
+ by (metis flip_commute subst_tv_commute subst_tv_var_flip subst_v_\<tau>_def subst_vv.simps(2))
+qed
+
+end
+
+lemma subst_tv_commute_full:
+ fixes c::\<tau>
+ assumes "atom z \<sharp> v" and "atom x \<sharp> w" and "x\<noteq>z"
+ shows "(c[z::=w]\<^sub>\<tau>\<^sub>v)[x::=v]\<^sub>\<tau>\<^sub>v = (c[x::=v]\<^sub>\<tau>\<^sub>v)[z::=w]\<^sub>\<tau>\<^sub>v"
+ using assms proof(nominal_induct c avoiding: x v z w rule: \<tau>.strong_induct)
+ case (T_refined_type x1a x2a x3a)
+ then show ?case using subst_cv_commute_full by simp
+qed
+
+lemma type_eq_subst_eq:
+ fixes v::v and c1::c
+ assumes "\<lbrace> z1 : b1 | c1 \<rbrace> = \<lbrace> z2 : b2 | c2 \<rbrace>"
+ shows "c1[z1::=v]\<^sub>c\<^sub>v = c2[z2::=v]\<^sub>c\<^sub>v"
+ using subst_v_flip_eq_two[of z1 c1 z2 c2 v] \<tau>.eq_iff assms subst_v_c_def by simp
+
+text \<open>Extract constraint from a type. We cannot just project out the constraint as this would
+mean alpha-equivalent types give different answers \<close>
+
+nominal_function c_of :: "\<tau> \<Rightarrow> x \<Rightarrow> c" where
+ "atom z \<sharp> x \<Longrightarrow> c_of (T_refined_type z b c) x = c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v"
+proof(goal_cases)
+ case 1
+ then show ?case using eqvt_def c_of_graph_aux_def by force
+next
+ case (2 x y)
+ then show ?case using eqvt_def c_of_graph_aux_def by force
+next
+ case (3 P x)
+ then obtain x1::\<tau> and x2::x where *:"x = (x1,x2)" by force
+ obtain z' and b' and c' where "x1 = \<lbrace> z' : b' | c' \<rbrace> \<and> atom z' \<sharp> x2" using obtain_fresh_z by metis
+ then show ?case using 3 * by auto
+next
+ case (4 z1 x1 b1 c1 z2 x2 b2 c2)
+ then show ?case using subst_v_flip_eq_two \<tau>.eq_iff by (metis prod.inject type_eq_subst_eq)
+qed
+
+nominal_termination (eqvt) by lexicographic_order
+
+lemma c_of_eq:
+ shows "c_of \<lbrace> x : b | c \<rbrace> x = c"
+proof(nominal_induct "\<lbrace> x : b | c \<rbrace>" avoiding: x rule: \<tau>.strong_induct)
+ case (T_refined_type x' c')
+ moreover hence "c_of \<lbrace> x' : b | c' \<rbrace> x = c'[x'::=V_var x]\<^sub>c\<^sub>v" using c_of.simps by auto
+ moreover have "\<lbrace> x' : b | c' \<rbrace> = \<lbrace> x : b | c \<rbrace>" using T_refined_type \<tau>.eq_iff by metis
+ moreover have "c'[x'::=V_var x]\<^sub>c\<^sub>v = c" using T_refined_type Abs1_eq_iff flip_subst_v subst_v_c_def
+ by (metis subst_cv_id)
+ ultimately show ?case by auto
+qed
+
+lemma obtain_fresh_z_c_of:
+ fixes t::"'b::fs"
+ obtains z where "atom z \<sharp> t \<and> \<tau> = \<lbrace> z : b_of \<tau> | c_of \<tau> z \<rbrace>"
+proof -
+ obtain z and c where "atom z \<sharp> t \<and> \<tau> = \<lbrace> z : b_of \<tau> | c \<rbrace>" using obtain_fresh_z2 by metis
+ moreover hence "c = c_of \<tau> z" using c_of.simps using c_of_eq by metis
+ ultimately show ?thesis
+ using that by auto
+qed
+
+lemma c_of_fresh:
+ fixes x::x
+ assumes "atom x \<sharp> (t,z)"
+ shows "atom x \<sharp> c_of t z"
+proof -
+ obtain z' and c' where z:"t = \<lbrace> z' : b_of t | c' \<rbrace> \<and> atom z' \<sharp> (x,z)" using obtain_fresh_z_c_of by metis
+ hence *:"c_of t z = c'[z'::=V_var z]\<^sub>c\<^sub>v" using c_of.simps fresh_Pair by metis
+ have "(atom x \<sharp> c' \<or> atom x \<in> set [atom z']) \<and> atom x \<sharp> b_of t" using \<tau>.fresh assms z fresh_Pair by metis
+ hence "atom x \<sharp> c'" using fresh_Pair z fresh_at_base(2) by fastforce
+ moreover have "atom x \<sharp> V_var z" using assms fresh_Pair v.fresh by metis
+ ultimately show ?thesis using assms fresh_subst_v_if[of "atom x" c' z' "V_var z"] subst_v_c_def * by metis
+qed
+
+lemma c_of_switch:
+ fixes z::x
+ assumes "atom z \<sharp> t"
+ shows "(c_of t z)[z::=V_var x]\<^sub>c\<^sub>v = c_of t x"
+proof -
+ obtain z' and c' where z:"t = \<lbrace> z' : b_of t | c' \<rbrace> \<and> atom z' \<sharp> (x,z)" using obtain_fresh_z_c_of by metis
+ hence "(atom z \<sharp> c' \<or> atom z \<in> set [atom z']) \<and> atom z \<sharp> b_of t" using \<tau>.fresh[of "atom z" z' "b_of t" c'] assms by metis
+ moreover have " atom z \<notin> set [atom z']" using z fresh_Pair by force
+ ultimately have **:"atom z \<sharp> c'" using fresh_Pair z fresh_at_base(2) by metis
+
+ have "(c_of t z)[z::=V_var x]\<^sub>c\<^sub>v = c'[z'::=V_var z]\<^sub>c\<^sub>v[z::=V_var x]\<^sub>c\<^sub>v" using c_of.simps fresh_Pair z by metis
+ also have "... = c'[z'::=V_var x]\<^sub>c\<^sub>v" using subst_v_simple_commute subst_v_c_def assms c_of.simps z ** by metis
+ finally show ?thesis using c_of.simps[of z' x "b_of t" c'] fresh_Pair z by metis
+qed
+
+lemma type_eq_subst_eq1:
+ fixes v::v and c1::c
+ assumes "\<lbrace> z1 : b1 | c1 \<rbrace> = (\<lbrace> z2 : b2 | c2 \<rbrace>)" and "atom z1 \<sharp> c2"
+ shows "c1[z1::=v]\<^sub>c\<^sub>v = c2[z2::=v]\<^sub>c\<^sub>v" and "b1=b2" and " c1 = (z1 \<leftrightarrow> z2) \<bullet> c2"
+proof -
+ show "c1[z1::=v]\<^sub>c\<^sub>v = c2[z2::=v]\<^sub>c\<^sub>v" using type_eq_subst_eq assms by blast
+ show "b1=b2" using \<tau>.eq_iff assms by blast
+ have "z1 = z2 \<and> c1 = c2 \<or> z1 \<noteq> z2 \<and> c1 = (z1 \<leftrightarrow> z2) \<bullet> c2 \<and> atom z1 \<sharp> c2"
+ using \<tau>.eq_iff Abs1_eq_iff[of z1 c1 z2 c2] assms by blast
+ thus "c1 = (z1 \<leftrightarrow> z2) \<bullet> c2" by auto
+qed
+
+lemma type_eq_subst_eq2:
+ fixes v::v and c1::c
+ assumes "\<lbrace> z1 : b1 | c1 \<rbrace> = (\<lbrace> z2 : b2 | c2 \<rbrace>)"
+ shows "c1[z1::=v]\<^sub>c\<^sub>v = c2[z2::=v]\<^sub>c\<^sub>v" and "b1=b2" and "[[atom z1]]lst. c1 = [[atom z2]]lst. c2"
+proof -
+ show "c1[z1::=v]\<^sub>c\<^sub>v = c2[z2::=v]\<^sub>c\<^sub>v" using type_eq_subst_eq assms by blast
+ show "b1=b2" using \<tau>.eq_iff assms by blast
+ show "[[atom z1]]lst. c1 = [[atom z2]]lst. c2"
+ using \<tau>.eq_iff assms by auto
+qed
+
+lemma type_eq_subst_eq3:
+ fixes v::v and c1::c
+ assumes "\<lbrace> z1 : b1 | c1 \<rbrace> = (\<lbrace> z2 : b2 | c2 \<rbrace>)" and "atom z1 \<sharp> c2"
+ shows "c1 = c2[z2::=V_var z1]\<^sub>c\<^sub>v" and "b1=b2"
+ using type_eq_subst_eq1 assms subst_v_c_def
+ by (metis subst_cv_var_flip)+
+
+lemma type_eq_flip:
+ assumes "atom x \<sharp> c"
+ shows "\<lbrace> z : b | c \<rbrace> = \<lbrace> x : b | (x \<leftrightarrow> z ) \<bullet> c \<rbrace>"
+ using \<tau>.eq_iff Abs1_eq_iff assms
+ by (metis (no_types, lifting) flip_fresh_fresh)
+
+lemma c_of_true:
+ "c_of \<lbrace> z' : B_bool | TRUE \<rbrace> x = C_true"
+proof(nominal_induct "\<lbrace> z' : B_bool | TRUE \<rbrace>" avoiding: x rule:\<tau>.strong_induct)
+ case (T_refined_type x1a x3a)
+ hence "\<lbrace> z' : B_bool | TRUE \<rbrace> = \<lbrace> x1a : B_bool | x3a \<rbrace>" using \<tau>.eq_iff by metis
+ then show ?case using subst_cv.simps c_of.simps T_refined_type
+ type_eq_subst_eq3
+ by (metis type_eq_subst_eq)
+qed
+
+lemma type_eq_subst:
+ assumes "atom x \<sharp> c"
+ shows "\<lbrace> z : b | c \<rbrace> = \<lbrace> x : b | c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v \<rbrace>"
+ using \<tau>.eq_iff Abs1_eq_iff assms
+ using subst_cv_var_flip type_eq_flip by auto
+
+lemma type_e_subst_fresh:
+ fixes x::x and z::x
+ assumes "atom z \<sharp> (x,v)" and "atom x \<sharp> e"
+ shows "\<lbrace> z : b | CE_val (V_var z) == e \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v = \<lbrace> z : b | CE_val (V_var z) == e \<rbrace>"
+ using assms subst_tv.simps subst_cv.simps forget_subst_cev by simp
+
+lemma type_v_subst_fresh:
+ fixes x::x and z::x
+ assumes "atom z \<sharp> (x,v)" and "atom x \<sharp> v'"
+ shows "\<lbrace> z : b | CE_val (V_var z) == CE_val v' \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v = \<lbrace> z : b | CE_val (V_var z) == CE_val v' \<rbrace>"
+ using assms subst_tv.simps subst_cv.simps by simp
+
+lemma subst_tbase_eq:
+ "b_of \<tau> = b_of \<tau>[x::=v]\<^sub>\<tau>\<^sub>v"
+proof -
+ obtain z and b and c where zbc: "\<tau> = \<lbrace> z:b|c\<rbrace> \<and> atom z \<sharp> (x,v)" using \<tau>.exhaust
+ by (metis prod.inject subst_tv.cases)
+ hence "b_of \<lbrace> z:b|c\<rbrace> = b_of \<lbrace> z:b|c\<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v" using subst_tv.simps by simp
+ thus ?thesis using zbc by blast
+qed
+
+lemma subst_tv_if:
+ assumes "atom z1 \<sharp> (x,v)" and "atom z' \<sharp> (x,v)"
+ shows "\<lbrace> z1 : b | CE_val (v'[x::=v]\<^sub>v\<^sub>v) == CE_val (V_lit l) IMP (c'[x::=v]\<^sub>c\<^sub>v)[z'::=[z1]\<^sup>v]\<^sub>c\<^sub>v \<rbrace> =
+ \<lbrace> z1 : b | CE_val v' == CE_val (V_lit l) IMP c'[z'::=[z1]\<^sup>v]\<^sub>c\<^sub>v \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v"
+ using subst_cv_commute_full[of z' v x "V_var z1" c'] subst_tv.simps subst_vv.simps(1) subst_ev.simps subst_cv.simps assms
+ by simp
+
+lemma subst_tv_tid:
+ assumes "atom za \<sharp> (x,v)"
+ shows "\<lbrace> za : B_id tid | TRUE \<rbrace> = \<lbrace> za : B_id tid | TRUE \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v"
+ using assms subst_tv.simps subst_cv.simps by presburger
+
+
+lemma b_of_subst:
+ "b_of (\<tau>[x::=v]\<^sub>\<tau>\<^sub>v) = b_of \<tau>"
+proof -
+ obtain z b c where *:"\<tau> = \<lbrace> z : b | c \<rbrace> \<and> atom z \<sharp> (x,v)" using obtain_fresh_z by metis
+ thus ?thesis using subst_tv.simps * by auto
+qed
+
+lemma subst_tv_flip:
+ assumes "\<tau>'[x::=v]\<^sub>\<tau>\<^sub>v = \<tau>" and "atom x \<sharp> (v,\<tau>)" and "atom x' \<sharp> (v,\<tau>)"
+ shows "((x' \<leftrightarrow> x) \<bullet> \<tau>')[x'::=v]\<^sub>\<tau>\<^sub>v = \<tau>"
+proof -
+ have "(x' \<leftrightarrow> x) \<bullet> v = v \<and> (x' \<leftrightarrow> x) \<bullet> \<tau> = \<tau>" using assms flip_fresh_fresh by auto
+ thus ?thesis using subst_tv.eqvt[of "(x' \<leftrightarrow> x)" \<tau>' x v ] assms by auto
+qed
+
+lemma subst_cv_true:
+ "\<lbrace> z : B_id tid | TRUE \<rbrace> = \<lbrace> z : B_id tid | TRUE \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v"
+proof -
+ obtain za::x where "atom za \<sharp> (x,v)" using obtain_fresh by auto
+ hence "\<lbrace> z : B_id tid | TRUE \<rbrace> = \<lbrace> za: B_id tid | TRUE \<rbrace>" using \<tau>.eq_iff Abs1_eq_iff by fastforce
+ moreover have "\<lbrace> za : B_id tid | TRUE \<rbrace> = \<lbrace> za : B_id tid | TRUE \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v"
+ using subst_cv.simps subst_tv.simps by (simp add: \<open>atom za \<sharp> (x, v)\<close>)
+ ultimately show ?thesis by argo
+qed
+
+lemma t_eq_supp:
+ assumes "(\<lbrace> z : b | c \<rbrace>) = (\<lbrace> z1 : b1 | c1 \<rbrace>)"
+ shows "supp c - { atom z } = supp c1 - { atom z1 }"
+proof -
+ have "supp c - { atom z } \<union> supp b = supp c1 - { atom z1 } \<union> supp b1" using \<tau>.supp assms
+ by (metis list.set(1) list.simps(15) sup_bot.right_neutral supp_b_empty)
+ moreover have "supp b = supp b1" using assms \<tau>.eq_iff by simp
+ moreover have "atom z1 \<notin> supp b1 \<and> atom z \<notin> supp b" using supp_b_empty by simp
+ ultimately show ?thesis
+ by (metis \<tau>.eq_iff \<tau>.supp assms b.supp(1) list.set(1) list.set(2) sup_bot.right_neutral)
+qed
+
+lemma fresh_t_eq:
+ fixes x::x
+ assumes "(\<lbrace> z : b | c \<rbrace>) = (\<lbrace> zz : b | cc \<rbrace>)" and "atom x \<sharp> c" and "x \<noteq> zz"
+ shows "atom x \<sharp> cc"
+proof -
+ have "supp c - { atom z } \<union> supp b = supp cc - { atom zz } \<union> supp b" using \<tau>.supp assms
+ by (metis list.set(1) list.simps(15) sup_bot.right_neutral supp_b_empty)
+ moreover have "atom x \<notin> supp c" using assms fresh_def by blast
+ ultimately have "atom x \<notin> supp cc - { atom zz } \<union> supp b" by force
+ hence "atom x \<notin> supp cc" using assms by simp
+ thus ?thesis using fresh_def by auto
+qed
+
+section \<open>Mutable Variable Context\<close>
+
+nominal_function subst_dv :: "\<Delta> \<Rightarrow> x \<Rightarrow> v \<Rightarrow> \<Delta>" where
+ "subst_dv DNil x v = DNil"
+| "subst_dv ((u,t) #\<^sub>\<Delta> \<Delta>) x v = ((u,t[x::=v]\<^sub>\<tau>\<^sub>v) #\<^sub>\<Delta> (subst_dv \<Delta> x v ))"
+ apply (simp add: eqvt_def subst_dv_graph_aux_def,auto )
+ using delete_aux.elims by (metis \<Delta>.exhaust surj_pair)
+nominal_termination (eqvt) by lexicographic_order
+
+abbreviation
+ subst_dv_abbrev :: "\<Delta> \<Rightarrow> x \<Rightarrow> v \<Rightarrow> \<Delta>" ("_[_::=_]\<^sub>\<Delta>\<^sub>v" [1000,50,50] 1000)
+ where
+ "\<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<equiv> subst_dv \<Delta> x v "
+
+nominal_function dmap :: "(u*\<tau> \<Rightarrow> u*\<tau>) \<Rightarrow> \<Delta> \<Rightarrow> \<Delta>" where
+ "dmap f DNil = DNil"
+| "dmap f ((u,t)#\<^sub>\<Delta>\<Delta>) = (f (u,t) #\<^sub>\<Delta> (dmap f \<Delta> ))"
+ apply (simp add: eqvt_def dmap_graph_aux_def,auto )
+ using delete_aux.elims by (metis \<Delta>.exhaust surj_pair)
+nominal_termination (eqvt) by lexicographic_order
+
+lemma subst_dv_iff:
+ "\<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v = dmap (\<lambda>(u,t). (u, t[x::=v]\<^sub>\<tau>\<^sub>v)) \<Delta>"
+ by(induct \<Delta>, auto)
+
+lemma size_subst_dv [simp]: "size ( subst_dv G i x) \<le> size G"
+ by (induct G,auto)
+
+lemma forget_subst_dv [simp]: "atom a \<sharp> G \<Longrightarrow> subst_dv G a x = G"
+ apply (induct G ,auto)
+ using fresh_DCons fresh_PairD(1) not_self_fresh apply fastforce
+ apply (simp add: fresh_DCons)+
+ done
+
+lemma subst_dv_member:
+ assumes "(u,\<tau>) \<in> setD \<Delta>"
+ shows "(u, \<tau>[x::=v]\<^sub>\<tau>\<^sub>v) \<in> setD (\<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v)"
+ using assms by(induct \<Delta> rule: \<Delta>_induct,auto)
+
+lemma fresh_subst_dv:
+ fixes x::x
+ assumes "atom xa \<sharp> \<Delta>" and "atom xa \<sharp> v"
+ shows "atom xa \<sharp>\<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v"
+ using assms proof(induct \<Delta> rule:\<Delta>_induct)
+ case DNil
+ then show ?case by auto
+next
+ case (DCons u t \<Delta>)
+ then show ?case using subst_dv.simps subst_v_\<tau>_def fresh_DCons fresh_Pair by simp
+qed
+
+lemma fresh_subst_dv_if:
+ fixes j::atom and i::x and x::v and t::\<Delta>
+ assumes "j \<sharp> t \<and> j \<sharp> x"
+ shows "(j \<sharp> subst_dv t i x)"
+ using assms proof(induct t rule: \<Delta>_induct)
+ case DNil
+ then show ?case using subst_gv.simps fresh_GNil by auto
+next
+ case (DCons u' t' D')
+ then show ?case unfolding subst_dv.simps using fresh_DCons fresh_subst_tv_if fresh_Pair by metis
+qed
+
+section \<open>Statements\<close>
+
+text \<open> Using ideas from proofs at top of AFP/Launchbury/Substitution.thy.
+ Subproofs borrowed from there; hence the apply style proofs. \<close>
+
+nominal_function (default "case_sum (\<lambda>x. Inl undefined) (case_sum (\<lambda>x. Inl undefined) (\<lambda>x. Inr undefined))")
+ subst_sv :: "s \<Rightarrow> x \<Rightarrow> v \<Rightarrow> s"
+ and subst_branchv :: "branch_s \<Rightarrow> x \<Rightarrow> v \<Rightarrow> branch_s"
+ and subst_branchlv :: "branch_list \<Rightarrow> x \<Rightarrow> v \<Rightarrow> branch_list" where
+ "subst_sv ( (AS_val v') ) x v = (AS_val (subst_vv v' x v ))"
+| "atom y \<sharp> (x,v) \<Longrightarrow> subst_sv (AS_let y e s) x v = (AS_let y (e[x::=v]\<^sub>e\<^sub>v) (subst_sv s x v ))"
+| "atom y \<sharp> (x,v) \<Longrightarrow> subst_sv (AS_let2 y t s1 s2) x v = (AS_let2 y (t[x::=v]\<^sub>\<tau>\<^sub>v) (subst_sv s1 x v ) (subst_sv s2 x v ))"
+| " subst_sv (AS_match v' cs) x v = AS_match (v'[x::=v]\<^sub>v\<^sub>v) (subst_branchlv cs x v )"
+| "subst_sv (AS_assign y v') x v = AS_assign y (subst_vv v' x v )"
+| "subst_sv ( (AS_if v' s1 s2) ) x v = (AS_if (subst_vv v' x v ) (subst_sv s1 x v ) (subst_sv s2 x v ) )"
+| "atom u \<sharp> (x,v) \<Longrightarrow> subst_sv (AS_var u \<tau> v' s) x v = AS_var u (subst_tv \<tau> x v ) (subst_vv v' x v ) (subst_sv s x v ) "
+| "subst_sv (AS_while s1 s2) x v = AS_while (subst_sv s1 x v ) (subst_sv s2 x v )"
+| "subst_sv (AS_seq s1 s2) x v = AS_seq (subst_sv s1 x v ) (subst_sv s2 x v )"
+| "subst_sv (AS_assert c s) x v = AS_assert (subst_cv c x v) (subst_sv s x v)"
+| "atom x1 \<sharp> (x,v) \<Longrightarrow> subst_branchv (AS_branch dc x1 s1 ) x v = AS_branch dc x1 (subst_sv s1 x v )"
+
+| "subst_branchlv (AS_final cs) x v = AS_final (subst_branchv cs x v )"
+| "subst_branchlv (AS_cons cs css) x v = AS_cons (subst_branchv cs x v ) (subst_branchlv css x v )"
+ apply (auto,simp add: eqvt_def subst_sv_subst_branchv_subst_branchlv_graph_aux_def )
+proof(goal_cases)
+
+ have eqvt_at_proj: "\<And> s xa va . eqvt_at subst_sv_subst_branchv_subst_branchlv_sumC (Inl (s, xa, va)) \<Longrightarrow>
+ eqvt_at (\<lambda>a. projl (subst_sv_subst_branchv_subst_branchlv_sumC (Inl a))) (s, xa, va)"
+ apply(simp add: eqvt_at_def)
+ apply(rule)
+ apply(subst Projl_permute)
+ apply(thin_tac _)+
+ apply (simp add: subst_sv_subst_branchv_subst_branchlv_sumC_def)
+ apply (simp add: THE_default_def)
+ apply (case_tac "Ex1 (subst_sv_subst_branchv_subst_branchlv_graph (Inl (s,xa,va)))")
+ apply simp
+ apply(auto)[1]
+ apply (erule_tac x="x" in allE)
+ apply simp
+ apply(cases rule: subst_sv_subst_branchv_subst_branchlv_graph.cases)
+ apply(assumption)
+ apply(rule_tac x="Sum_Type.projl x" in exI,clarify,rule the1_equality,blast,simp (no_asm) only: sum.sel)+
+ apply blast +
+
+ apply(simp)+
+ done
+
+ {
+ case (1 P x')
+ then show ?case proof(cases x')
+ case (Inl a) thus P
+ proof(cases a)
+ case (fields aa bb cc)
+ thus P using Inl 1 s_branch_s_branch_list.strong_exhaust fresh_star_insert by metis
+ qed
+ next
+ case (Inr b) thus P
+ proof(cases b)
+ case (Inl a) thus P proof(cases a)
+ case (fields aa bb cc)
+ then show ?thesis using Inr Inl 1 s_branch_s_branch_list.strong_exhaust fresh_star_insert by metis
+ qed
+ next
+ case Inr2: (Inr b) thus P proof(cases b)
+ case (fields aa bb cc)
+ then show ?thesis using Inr Inr2 1 s_branch_s_branch_list.strong_exhaust fresh_star_insert by metis
+ qed
+ qed
+ qed
+ next
+ case (2 y s ya xa va sa c)
+ thus ?case using eqvt_triple eqvt_at_proj by blast
+ next
+ case (3 y s2 ya xa va s1a s2a c)
+ thus ?case using eqvt_triple eqvt_at_proj by blast
+ next
+ case (4 u xa va s ua sa c)
+ moreover have "atom u \<sharp> (xa, va) \<and> atom ua \<sharp> (xa, va)"
+ using fresh_Pair u_fresh_xv by auto
+ ultimately show ?case using eqvt_triple[of u xa va ua s sa] subst_sv_def eqvt_at_proj by metis
+ next
+ case (5 x1 s1 x1a xa va s1a c)
+ thus ?case using eqvt_triple eqvt_at_proj by blast
+ }
+qed
+nominal_termination (eqvt) by lexicographic_order
+
+abbreviation
+ subst_sv_abbrev :: "s \<Rightarrow> x \<Rightarrow> v \<Rightarrow> s" ("_[_::=_]\<^sub>s\<^sub>v" [1000,50,50] 1000)
+ where
+ "s[x::=v]\<^sub>s\<^sub>v \<equiv> subst_sv s x v"
+
+abbreviation
+ subst_branchv_abbrev :: "branch_s \<Rightarrow> x \<Rightarrow> v \<Rightarrow> branch_s" ("_[_::=_]\<^sub>s\<^sub>v" [1000,50,50] 1000)
+ where
+ "s[x::=v]\<^sub>s\<^sub>v \<equiv> subst_branchv s x v"
+
+lemma size_subst_sv [simp]: "size (subst_sv A i x ) = size A" and "size (subst_branchv B i x ) = size B" and "size (subst_branchlv C i x ) = size C"
+ by(nominal_induct A and B and C avoiding: i x rule: s_branch_s_branch_list.strong_induct,auto)
+
+lemma forget_subst_sv [simp]: shows "atom a \<sharp> A \<Longrightarrow> subst_sv A a x = A" and "atom a \<sharp> B \<Longrightarrow> subst_branchv B a x = B" and "atom a \<sharp> C \<Longrightarrow> subst_branchlv C a x = C"
+ by (nominal_induct A and B and C avoiding: a x rule: s_branch_s_branch_list.strong_induct,auto simp: fresh_at_base)
+
+lemma subst_sv_id [simp]: "subst_sv A a (V_var a) = A" and "subst_branchv B a (V_var a) = B" and "subst_branchlv C a (V_var a) = C"
+proof(nominal_induct A and B and C avoiding: a rule: s_branch_s_branch_list.strong_induct)
+ case (AS_let x option e s)
+ then show ?case
+ by (metis (no_types, lifting) fresh_Pair not_None_eq subst_ev_id subst_sv.simps(2) subst_sv.simps(3) subst_tv_id v.fresh(2))
+next
+ case (AS_match v branch_s)
+ then show ?case using fresh_Pair not_None_eq subst_ev_id subst_sv.simps subst_sv.simps subst_tv_id v.fresh subst_vv_id
+ by metis
+qed(auto)+
+
+lemma fresh_subst_sv_if_rl:
+ shows
+ "(atom x \<sharp> s \<and> j \<sharp> s) \<or> (j \<sharp> v \<and> (j \<sharp> s \<or> j = atom x)) \<Longrightarrow> j \<sharp> (subst_sv s x v )" and
+ "(atom x \<sharp> cs \<and> j \<sharp> cs) \<or> (j \<sharp> v \<and> (j \<sharp> cs \<or> j = atom x)) \<Longrightarrow> j \<sharp> (subst_branchv cs x v)" and
+ "(atom x \<sharp> css \<and> j \<sharp> css) \<or> (j \<sharp> v \<and> (j \<sharp> css \<or> j = atom x)) \<Longrightarrow> j \<sharp> (subst_branchlv css x v )"
+ apply(nominal_induct s and cs and css avoiding: v x rule: s_branch_s_branch_list.strong_induct)
+ using pure_fresh by force+
+
+lemma fresh_subst_sv_if_lr:
+ shows "j \<sharp> (subst_sv s x v) \<Longrightarrow> (atom x \<sharp> s \<and> j \<sharp> s) \<or> (j \<sharp> v \<and> (j \<sharp> s \<or> j = atom x))" and
+ "j \<sharp> (subst_branchv cs x v) \<Longrightarrow> (atom x \<sharp> cs \<and> j \<sharp> cs) \<or> (j \<sharp> v \<and> (j \<sharp> cs \<or> j = atom x))" and
+ "j \<sharp> (subst_branchlv css x v ) \<Longrightarrow> (atom x \<sharp> css \<and> j \<sharp> css) \<or> (j \<sharp> v \<and> (j \<sharp> css \<or> j = atom x))"
+proof(nominal_induct s and cs and css avoiding: v x rule: s_branch_s_branch_list.strong_induct)
+ case (AS_branch list x s )
+ then show ?case using s_branch_s_branch_list.fresh fresh_Pair list.distinct(1) list.set_cases pure_fresh set_ConsD subst_branchv.simps by metis
+next
+ case (AS_let y e s')
+ thus ?case proof(cases "atom x \<sharp> (AS_let y e s')")
+ case True
+ hence "subst_sv (AS_let y e s') x v = (AS_let y e s')" using forget_subst_sv by simp
+ hence "j \<sharp> (AS_let y e s')" using AS_let by argo
+ then show ?thesis using True by blast
+ next
+ case False
+ have "subst_sv (AS_let y e s') x v = AS_let y (e[x::=v]\<^sub>e\<^sub>v) (s'[x::=v]\<^sub>s\<^sub>v)" using subst_sv.simps(2) AS_let by force
+ hence "((j \<sharp> s'[x::=v]\<^sub>s\<^sub>v \<or> j \<in> set [atom y]) \<and> j \<sharp> None \<and> j \<sharp> e[x::=v]\<^sub>e\<^sub>v)" using s_branch_s_branch_list.fresh AS_let
+ by (simp add: fresh_None)
+ then show ?thesis using AS_let fresh_None fresh_subst_ev_if list.discI list.set_cases s_branch_s_branch_list.fresh set_ConsD
+ by metis
+ qed
+next
+ case (AS_let2 y \<tau> s1 s2)
+ thus ?case proof(cases "atom x \<sharp> (AS_let2 y \<tau> s1 s2)")
+ case True
+ hence "subst_sv (AS_let2 y \<tau> s1 s2) x v = (AS_let2 y \<tau> s1 s2)" using forget_subst_sv by simp
+ hence "j \<sharp> (AS_let2 y \<tau> s1 s2)" using AS_let2 by argo
+ then show ?thesis using True by blast
+ next
+ case False
+ have "subst_sv (AS_let2 y \<tau> s1 s2) x v = AS_let2 y (\<tau>[x::=v]\<^sub>\<tau>\<^sub>v) (s1[x::=v]\<^sub>s\<^sub>v) (s2[x::=v]\<^sub>s\<^sub>v)" using subst_sv.simps AS_let2 by force
+ then show ?thesis using AS_let2
+ fresh_subst_tv_if list.discI list.set_cases s_branch_s_branch_list.fresh(4) set_ConsD by auto
+ qed
+qed(auto)+
+
+lemma fresh_subst_sv_if[simp]:
+ fixes x::x and v::v
+ shows "j \<sharp> (subst_sv s x v) \<longleftrightarrow> (atom x \<sharp> s \<and> j \<sharp> s) \<or> (j \<sharp> v \<and> (j \<sharp> s \<or> j = atom x))" and
+ "j \<sharp> (subst_branchv cs x v) \<longleftrightarrow> (atom x \<sharp> cs \<and> j \<sharp> cs) \<or> (j \<sharp> v \<and> (j \<sharp> cs \<or> j = atom x))"
+ using fresh_subst_sv_if_lr fresh_subst_sv_if_rl by metis+
+
+lemma subst_sv_commute [simp]:
+ fixes A::s and t::v and j::x and i::x
+ shows "atom j \<sharp> A \<Longrightarrow> (subst_sv (subst_sv A i t) j u ) = subst_sv A i (subst_vv t j u )" and
+ "atom j \<sharp> B \<Longrightarrow> (subst_branchv (subst_branchv B i t ) j u ) = subst_branchv B i (subst_vv t j u )" and
+ "atom j \<sharp> C \<Longrightarrow> (subst_branchlv (subst_branchlv C i t) j u ) = subst_branchlv C i (subst_vv t j u ) "
+ apply(nominal_induct A and B and C avoiding: i j t u rule: s_branch_s_branch_list.strong_induct)
+ by(auto simp: fresh_at_base)
+
+lemma c_eq_perm:
+ assumes "( (atom z) \<rightleftharpoons> (atom z') ) \<bullet> c = c'" and "atom z' \<sharp> c"
+ shows "\<lbrace> z : b | c \<rbrace> = \<lbrace> z' : b | c' \<rbrace>"
+ using \<tau>.eq_iff Abs1_eq_iff(3)
+ by (metis Nominal2_Base.swap_commute assms(1) assms(2) flip_def swap_fresh_fresh)
+
+lemma subst_sv_flip:
+ fixes s::s and sa::s and v'::v
+ assumes "atom c \<sharp> (s, sa)" and "atom c \<sharp> (v',x, xa, s, sa)" "atom x \<sharp> v'" and "atom xa \<sharp> v'" and "(x \<leftrightarrow> c) \<bullet> s = (xa \<leftrightarrow> c) \<bullet> sa"
+ shows "s[x::=v']\<^sub>s\<^sub>v = sa[xa::=v']\<^sub>s\<^sub>v"
+proof -
+ have "atom x \<sharp> (s[x::=v']\<^sub>s\<^sub>v)" and xafr: "atom xa \<sharp> (sa[xa::=v']\<^sub>s\<^sub>v)"
+ and "atom c \<sharp> ( s[x::=v']\<^sub>s\<^sub>v, sa[xa::=v']\<^sub>s\<^sub>v)" using assms using fresh_subst_sv_if assms by( blast+ ,force)
+
+ hence "s[x::=v']\<^sub>s\<^sub>v = (x \<leftrightarrow> c) \<bullet> (s[x::=v']\<^sub>s\<^sub>v)" by (simp add: flip_fresh_fresh fresh_Pair)
+ also have " ... = ((x \<leftrightarrow> c) \<bullet> s)[ ((x \<leftrightarrow> c) \<bullet> x) ::= ((x \<leftrightarrow> c) \<bullet> v') ]\<^sub>s\<^sub>v" using subst_sv_subst_branchv_subst_branchlv.eqvt by blast
+ also have "... = ((xa \<leftrightarrow> c) \<bullet> sa)[ ((x \<leftrightarrow> c) \<bullet> x) ::= ((x \<leftrightarrow> c) \<bullet> v') ]\<^sub>s\<^sub>v" using assms by presburger
+ also have "... = ((xa \<leftrightarrow> c) \<bullet> sa)[ ((xa \<leftrightarrow> c) \<bullet> xa) ::= ((xa \<leftrightarrow> c) \<bullet> v') ]\<^sub>s\<^sub>v" using assms
+ by (metis flip_at_simps(1) flip_fresh_fresh fresh_PairD(1))
+ also have "... = (xa \<leftrightarrow> c) \<bullet> (sa[xa::=v']\<^sub>s\<^sub>v)" using subst_sv_subst_branchv_subst_branchlv.eqvt by presburger
+ also have "... = sa[xa::=v']\<^sub>s\<^sub>v" using xafr assms by (simp add: flip_fresh_fresh fresh_Pair)
+ finally show ?thesis by simp
+qed
+
+lemma if_type_eq:
+ fixes \<Gamma>::\<Gamma> and v::v and z1::x
+ assumes "atom z1' \<sharp> (v, ca, (x, b, c) #\<^sub>\<Gamma> \<Gamma>, (CE_val v == CE_val (V_lit ll) IMP ca[za::=[z1]\<^sup>v]\<^sub>c\<^sub>v ))" and "atom z1 \<sharp> v"
+ and "atom z1 \<sharp> (za,ca)" and "atom z1' \<sharp> (za,ca)"
+ shows "(\<lbrace> z1' : ba | CE_val v == CE_val (V_lit ll) IMP ca[za::=[z1']\<^sup>v]\<^sub>c\<^sub>v \<rbrace>) = \<lbrace> z1 : ba | CE_val v == CE_val (V_lit ll) IMP ca[za::=[z1]\<^sup>v]\<^sub>c\<^sub>v \<rbrace>"
+proof -
+ have "atom z1' \<sharp> (CE_val v == CE_val (V_lit ll) IMP ca[za::=[z1]\<^sup>v]\<^sub>c\<^sub>v )" using assms fresh_prod4 by blast
+ moreover hence "(CE_val v == CE_val (V_lit ll) IMP ca[za::=[z1']\<^sup>v]\<^sub>c\<^sub>v) = (z1' \<leftrightarrow> z1) \<bullet> (CE_val v == CE_val (V_lit ll) IMP ca[za::=[z1]\<^sup>v]\<^sub>c\<^sub>v )"
+ proof -
+ have "(z1' \<leftrightarrow> z1) \<bullet> (CE_val v == CE_val (V_lit ll) IMP ca[za::=[z1]\<^sup>v]\<^sub>c\<^sub>v ) = ( (z1' \<leftrightarrow> z1) \<bullet> (CE_val v == CE_val (V_lit ll)) IMP ((z1' \<leftrightarrow> z1) \<bullet> ca[za::=[z1]\<^sup>v]\<^sub>c\<^sub>v ))"
+ by auto
+ also have "... = ((CE_val v == CE_val (V_lit ll)) IMP ((z1' \<leftrightarrow> z1) \<bullet> ca[za::=[z1]\<^sup>v]\<^sub>c\<^sub>v ))"
+ using \<open>atom z1 \<sharp> v\<close> assms
+ by (metis (mono_tags) \<open>atom z1' \<sharp> (CE_val v == CE_val (V_lit ll) IMP ca[za::=[z1]\<^sup>v]\<^sub>c\<^sub>v )\<close> c.fresh(6) c.fresh(7) ce.fresh(1) flip_at_simps(2) flip_fresh_fresh fresh_at_base_permute_iff fresh_def supp_l_empty v.fresh(1))
+ also have "... = ((CE_val v == CE_val (V_lit ll)) IMP (ca[za::=[z1']\<^sup>v]\<^sub>c\<^sub>v ))"
+ using assms by fastforce
+ finally show ?thesis by auto
+ qed
+ ultimately show ?thesis
+ using \<tau>.eq_iff Abs1_eq_iff(3)[of z1' "CE_val v == CE_val (V_lit ll) IMP ca[za::=[z1']\<^sup>v]\<^sub>c\<^sub>v"
+ z1 "CE_val v == CE_val (V_lit ll) IMP ca[za::=[z1]\<^sup>v]\<^sub>c\<^sub>v"] by blast
+qed
+
+lemma subst_sv_var_flip:
+ fixes x::x and s::s and z::x
+ shows "atom x \<sharp> s \<Longrightarrow> ((x \<leftrightarrow> z) \<bullet> s) = s[z::=[x]\<^sup>v]\<^sub>s\<^sub>v" and
+ "atom x \<sharp> cs \<Longrightarrow> ((x \<leftrightarrow> z) \<bullet> cs) = subst_branchv cs z [x]\<^sup>v" and
+ "atom x \<sharp> css \<Longrightarrow> ((x \<leftrightarrow> z) \<bullet> css) = subst_branchlv css z [x]\<^sup>v"
+ apply(nominal_induct s and cs and css avoiding: z x rule: s_branch_s_branch_list.strong_induct)
+ using [[simproc del: alpha_lst]]
+ apply (auto ) (* This unpacks subst, perm *)
+ using subst_tv_var_flip flip_fresh_fresh v.fresh s_branch_s_branch_list.fresh
+ subst_v_\<tau>_def subst_v_v_def subst_vv_var_flip subst_v_e_def subst_ev_var_flip pure_fresh apply auto
+ defer 1 (* Sometimes defering hard goals to the end makes it easier to finish *)
+ using x_fresh_u apply blast (* Next two involve u and flipping with x *)
+ defer 1
+ using x_fresh_u apply blast
+ defer 1
+ using x_fresh_u Abs1_eq_iff'(3) flip_fresh_fresh
+ apply (simp add: subst_v_c_def)
+ using x_fresh_u Abs1_eq_iff'(3) flip_fresh_fresh
+ by (simp add: flip_fresh_fresh)
+
+instantiation s :: has_subst_v
+begin
+
+definition
+ "subst_v = subst_sv"
+
+instance proof
+ fix j::atom and i::x and x::v and t::s
+ show "(j \<sharp> subst_v t i x) = ((atom i \<sharp> t \<and> j \<sharp> t) \<or> (j \<sharp> x \<and> (j \<sharp> t \<or> j = atom i)))"
+ using fresh_subst_sv_if subst_v_s_def by auto
+
+ fix a::x and tm::s and x::v
+ show "atom a \<sharp> tm \<Longrightarrow> subst_v tm a x = tm"
+ using forget_subst_sv subst_v_s_def by simp
+
+ fix a::x and tm::s
+ show "subst_v tm a (V_var a) = tm" using subst_sv_id subst_v_s_def by simp
+
+ fix p::perm and x1::x and v::v and t1::s
+ show "p \<bullet> subst_v t1 x1 v = subst_v (p \<bullet> t1) (p \<bullet> x1) (p \<bullet> v)"
+ using subst_sv_commute subst_v_s_def by simp
+
+ fix x::x and c::s and z::x
+ show "atom x \<sharp> c \<Longrightarrow> ((x \<leftrightarrow> z) \<bullet> c) = c[z::=[x]\<^sup>v]\<^sub>v"
+ using subst_sv_var_flip subst_v_s_def by simp
+
+ fix x::x and c::s and z::x
+ show "atom x \<sharp> c \<Longrightarrow> c[z::=[x]\<^sup>v]\<^sub>v[x::=v]\<^sub>v = c[z::=v]\<^sub>v"
+ using subst_sv_var_flip subst_v_s_def by simp
+qed
+end
+
+section \<open>Type Definition\<close>
+
+nominal_function subst_ft_v :: "fun_typ \<Rightarrow> x \<Rightarrow> v \<Rightarrow> fun_typ" where
+ "atom z \<sharp> (x,v) \<Longrightarrow> subst_ft_v ( AF_fun_typ z b c t (s::s)) x v = AF_fun_typ z b c[x::=v]\<^sub>c\<^sub>v t[x::=v]\<^sub>\<tau>\<^sub>v s[x::=v]\<^sub>s\<^sub>v"
+ apply(simp add: eqvt_def subst_ft_v_graph_aux_def )
+ apply(simp add:fun_typ.strong_exhaust )
+ apply(auto)
+ apply(rule_tac y=a and c="(aa,b)" in fun_typ.strong_exhaust)
+ apply (auto simp: eqvt_at_def fresh_star_def fresh_Pair fresh_at_base)
+
+proof(goal_cases)
+ case (1 z xa va c t s za ca ta sa cb)
+ hence "c[z::=[ cb ]\<^sup>v]\<^sub>c\<^sub>v = ca[za::=[ cb ]\<^sup>v]\<^sub>c\<^sub>v"
+ by (metis flip_commute subst_cv_var_flip)
+ hence "c[z::=[ cb ]\<^sup>v]\<^sub>c\<^sub>v[xa::=va]\<^sub>c\<^sub>v = ca[za::=[ cb ]\<^sup>v]\<^sub>c\<^sub>v[xa::=va]\<^sub>c\<^sub>v" by auto
+ then show ?case using subst_cv_commute atom_eq_iff fresh_atom fresh_atom_at_base subst_cv_commute_full v.fresh
+ using 1 subst_cv_var_flip flip_commute by metis
+next
+ case (2 z xa va c t s za ca ta sa cb)
+ hence "t[z::=[ cb ]\<^sup>v]\<^sub>\<tau>\<^sub>v = ta[za::=[ cb ]\<^sup>v]\<^sub>\<tau>\<^sub>v" by metis
+ hence "t[z::=[ cb ]\<^sup>v]\<^sub>\<tau>\<^sub>v[xa::=va]\<^sub>\<tau>\<^sub>v = ta[za::=[ cb ]\<^sup>v]\<^sub>\<tau>\<^sub>v[xa::=va]\<^sub>\<tau>\<^sub>v" by auto
+ then show ?case using subst_tv_commute_full 2
+ by (metis atom_eq_iff fresh_atom fresh_atom_at_base v.fresh(2))
+qed
+
+nominal_termination (eqvt) by lexicographic_order
+
+nominal_function subst_ftq_v :: "fun_typ_q \<Rightarrow> x \<Rightarrow> v \<Rightarrow> fun_typ_q" where
+ "atom bv \<sharp> (x,v) \<Longrightarrow> subst_ftq_v (AF_fun_typ_some bv ft) x v = (AF_fun_typ_some bv (subst_ft_v ft x v))"
+| "subst_ftq_v (AF_fun_typ_none ft) x v = (AF_fun_typ_none (subst_ft_v ft x v))"
+ apply(simp add: eqvt_def subst_ftq_v_graph_aux_def )
+ apply(simp add:fun_typ_q.strong_exhaust )
+ apply(auto)
+ apply(rule_tac y=a and c="(aa,b)" in fun_typ_q.strong_exhaust)
+ apply (auto simp: eqvt_at_def fresh_star_def fresh_Pair fresh_at_base)
+proof(goal_cases)
+ case (1 bv ft bva fta xa va c)
+ then show ?case using subst_ft_v.simps by (simp add: flip_fresh_fresh)
+qed
+nominal_termination (eqvt) by lexicographic_order
+
+lemma size_subst_ft[simp]: "size (subst_ft_v A x v) = size A"
+ by(nominal_induct A avoiding: x v rule: fun_typ.strong_induct,auto)
+
+lemma forget_subst_ft [simp]: shows "atom x \<sharp> A \<Longrightarrow> subst_ft_v A x a = A"
+ by (nominal_induct A avoiding: a x rule: fun_typ.strong_induct,auto simp: fresh_at_base)
+
+lemma subst_ft_id [simp]: "subst_ft_v A a (V_var a) = A"
+ by(nominal_induct A avoiding: a rule: fun_typ.strong_induct,auto)
+
+instantiation fun_typ :: has_subst_v
+begin
+
+definition
+ "subst_v = subst_ft_v"
+
+instance proof
+
+ fix j::atom and i::x and x::v and t::fun_typ
+ show "(j \<sharp> subst_v t i x) = ((atom i \<sharp> t \<and> j \<sharp> t) \<or> (j \<sharp> x \<and> (j \<sharp> t \<or> j = atom i)))"
+ apply(nominal_induct t avoiding: i x rule:fun_typ.strong_induct)
+ apply(simp only: subst_v_fun_typ_def subst_ft_v.simps )
+ using fun_typ.fresh fresh_subst_v_if apply simp
+ by auto
+
+ fix a::x and tm::fun_typ and x::v
+ show "atom a \<sharp> tm \<Longrightarrow> subst_v tm a x = tm"
+ proof(nominal_induct tm avoiding: a x rule:fun_typ.strong_induct)
+ case (AF_fun_typ x1a x2a x3a x4a x5a)
+ then show ?case unfolding subst_ft_v.simps subst_v_fun_typ_def fun_typ.fresh using forget_subst_v subst_ft_v.simps subst_v_c_def forget_subst_sv subst_v_\<tau>_def by fastforce
+ qed
+
+ fix a::x and tm::fun_typ
+ show "subst_v tm a (V_var a) = tm"
+ proof(nominal_induct tm avoiding: a x rule:fun_typ.strong_induct)
+ case (AF_fun_typ x1a x2a x3a x4a x5a)
+ then show ?case unfolding subst_ft_v.simps subst_v_fun_typ_def fun_typ.fresh using forget_subst_v subst_ft_v.simps subst_v_c_def forget_subst_sv subst_v_\<tau>_def by fastforce
+ qed
+
+ fix p::perm and x1::x and v::v and t1::fun_typ
+ show "p \<bullet> subst_v t1 x1 v = subst_v (p \<bullet> t1) (p \<bullet> x1) (p \<bullet> v)"
+ proof(nominal_induct t1 avoiding: x1 v rule:fun_typ.strong_induct)
+ case (AF_fun_typ x1a x2a x3a x4a x5a)
+ then show ?case unfolding subst_ft_v.simps subst_v_fun_typ_def fun_typ.fresh using forget_subst_v subst_ft_v.simps subst_v_c_def forget_subst_sv subst_v_\<tau>_def by fastforce
+ qed
+
+ fix x::x and c::fun_typ and z::x
+ show "atom x \<sharp> c \<Longrightarrow> ((x \<leftrightarrow> z) \<bullet> c) = c[z::=[x]\<^sup>v]\<^sub>v"
+ apply(nominal_induct c avoiding: x z rule:fun_typ.strong_induct)
+ by (auto simp add: subst_v_c_def subst_v_s_def subst_v_\<tau>_def subst_v_fun_typ_def)
+
+ fix x::x and c::fun_typ and z::x
+ show "atom x \<sharp> c \<Longrightarrow> c[z::=[x]\<^sup>v]\<^sub>v[x::=v]\<^sub>v = c[z::=v]\<^sub>v"
+ apply(nominal_induct c avoiding: z x v rule:fun_typ.strong_induct)
+ apply auto
+ by (auto simp add: subst_v_c_def subst_v_s_def subst_v_\<tau>_def subst_v_fun_typ_def )
+qed
+end
+
+instantiation fun_typ_q :: has_subst_v
+begin
+
+definition
+ "subst_v = subst_ftq_v"
+
+instance proof
+ fix j::atom and i::x and x::v and t::fun_typ_q
+ show "(j \<sharp> subst_v t i x) = ((atom i \<sharp> t \<and> j \<sharp> t) \<or> (j \<sharp> x \<and> (j \<sharp> t \<or> j = atom i)))"
+ apply(nominal_induct t avoiding: i x rule:fun_typ_q.strong_induct,auto)
+ apply(auto simp add: subst_v_fun_typ_def subst_v_s_def subst_v_\<tau>_def subst_v_fun_typ_q_def fresh_subst_v_if )
+ by (metis (no_types) fresh_subst_v_if subst_v_fun_typ_def)+
+
+ fix i::x and t::fun_typ_q and x::v
+ show "atom i \<sharp> t \<Longrightarrow> subst_v t i x = t"
+ apply(nominal_induct t avoiding: i x rule:fun_typ_q.strong_induct,auto)
+ by(auto simp add: subst_v_fun_typ_def subst_v_s_def subst_v_\<tau>_def subst_v_fun_typ_q_def fresh_subst_v_if )
+
+ fix i::x and t::fun_typ_q
+ show "subst_v t i (V_var i) = t" using subst_cv_id subst_v_fun_typ_def
+ apply(nominal_induct t avoiding: i x rule:fun_typ_q.strong_induct,auto)
+ by(auto simp add: subst_v_fun_typ_def subst_v_s_def subst_v_\<tau>_def subst_v_fun_typ_q_def fresh_subst_v_if )
+
+ fix p::perm and x1::x and v::v and t1::fun_typ_q
+ show "p \<bullet> subst_v t1 x1 v = subst_v (p \<bullet> t1) (p \<bullet> x1) (p \<bullet> v)"
+ apply(nominal_induct t1 avoiding: v x1 rule:fun_typ_q.strong_induct,auto)
+ by(auto simp add: subst_v_fun_typ_def subst_v_s_def subst_v_\<tau>_def subst_v_fun_typ_q_def fresh_subst_v_if )
+
+ fix x::x and c::fun_typ_q and z::x
+ show "atom x \<sharp> c \<Longrightarrow> ((x \<leftrightarrow> z) \<bullet> c) = c[z::=[x]\<^sup>v]\<^sub>v"
+ apply(nominal_induct c avoiding: x z rule:fun_typ_q.strong_induct,auto)
+ by(auto simp add: subst_v_fun_typ_def subst_v_s_def subst_v_\<tau>_def subst_v_fun_typ_q_def fresh_subst_v_if )
+
+ fix x::x and c::fun_typ_q and z::x
+ show "atom x \<sharp> c \<Longrightarrow> c[z::=[x]\<^sup>v]\<^sub>v[x::=v]\<^sub>v = c[z::=v]\<^sub>v"
+ apply(nominal_induct c avoiding: z x v rule:fun_typ_q.strong_induct,auto)
+ apply(auto simp add: subst_v_fun_typ_def subst_v_s_def subst_v_\<tau>_def subst_v_fun_typ_q_def fresh_subst_v_if )
+ by (metis subst_v_fun_typ_def flip_bv_x_cancel subst_ft_v.eqvt subst_v_simple_commute v.perm_simps )+
+qed
+
+end
+
+section \<open>Variable Context\<close>
+
+lemma subst_dv_fst_eq:
+ "fst ` setD (\<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v) = fst ` setD \<Delta>"
+ by(induct \<Delta> rule: \<Delta>_induct,simp,force)
+
+lemma subst_gv_member_iff:
+ fixes x'::x and x::x and v::v and c'::c
+ assumes "(x',b',c') \<in> toSet \<Gamma>" and "atom x \<notin> atom_dom \<Gamma>"
+ shows "(x',b',c'[x::=v]\<^sub>c\<^sub>v) \<in> toSet \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v"
+proof -
+ have "x' \<noteq> x" using assms fresh_dom_free2 by metis
+ then show ?thesis using assms proof(induct \<Gamma> rule: \<Gamma>_induct)
+ case GNil
+ then show ?case by auto
+ next
+ case (GCons x1 b1 c1 \<Gamma>')
+ show ?case proof(cases "(x',b',c') = (x1,b1,c1)")
+ case True
+ hence "((x1, b1, c1) #\<^sub>\<Gamma> \<Gamma>')[x::=v]\<^sub>\<Gamma>\<^sub>v = ((x1, b1, c1[x::=v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> (\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v))" using subst_gv.simps \<open>x'\<noteq>x\<close> by auto
+ then show ?thesis using True by auto
+ next
+ case False
+ have "x1\<noteq>x" using fresh_def fresh_GCons fresh_Pair supp_at_base GCons fresh_dom_free2 by auto
+ hence "(x', b', c') \<in> toSet \<Gamma>'" using GCons False toSet.simps by auto
+ moreover have "atom x \<notin> atom_dom \<Gamma>'" using fresh_GCons GCons dom.simps toSet.simps by simp
+ ultimately have "(x', b', c'[x::=v]\<^sub>c\<^sub>v) \<in> toSet \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v" using GCons by auto
+ hence "(x', b', c'[x::=v]\<^sub>c\<^sub>v) \<in> toSet ((x1, b1, c1[x::=v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> (\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v))" by auto
+ then show ?thesis using subst_gv.simps \<open>x1\<noteq>x\<close> by auto
+ qed
+ qed
+qed
+
+lemma fresh_subst_gv_if:
+ fixes j::atom and i::x and x::v and t::\<Gamma>
+ assumes "j \<sharp> t \<and> j \<sharp> x"
+ shows "(j \<sharp> subst_gv t i x)"
+ using assms proof(induct t rule: \<Gamma>_induct)
+ case GNil
+ then show ?case using subst_gv.simps fresh_GNil by auto
+next
+ case (GCons x' b' c' \<Gamma>')
+ then show ?case unfolding subst_gv.simps using fresh_GCons fresh_subst_cv_if by auto
+qed
+
+section \<open>Lookup\<close>
+
+lemma set_GConsD: "y \<in> toSet (x #\<^sub>\<Gamma> xs) \<Longrightarrow> y=x \<or> y \<in> toSet xs"
+ by auto
+
+lemma subst_g_assoc_cons:
+ assumes "x \<noteq> x'"
+ shows "(((x', b', c') #\<^sub>\<Gamma> \<Gamma>')[x::=v]\<^sub>\<Gamma>\<^sub>v @ G) = ((x', b', c'[x::=v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> ((\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v) @ G))"
+ using subst_gv.simps append_g.simps assms by auto
+
end
\ No newline at end of file
diff --git a/thys/MiniSail/IVSubstTypingL.thy b/thys/MiniSail/IVSubstTypingL.thy
--- a/thys/MiniSail/IVSubstTypingL.thy
+++ b/thys/MiniSail/IVSubstTypingL.thy
@@ -1,1534 +1,1534 @@
-(*<*)
-theory IVSubstTypingL
- imports SubstMethods ContextSubtypingL
-begin
- (*>*)
-
-chapter \<open>Immutable Variable Substitution Lemmas\<close>
-text \<open>Lemmas that show that types are preserved, in some way,
-under immutable variable substitution\<close>
-
-section \<open>Proof Methods\<close>
-
-method subst_mth = (metis subst_g_inside infer_e_wf infer_v_wf infer_v_wf)
-
-method subst_tuple_mth uses add = (
- (unfold fresh_prodN), (simp add: add )+,
- (rule,metis fresh_z_subst_g add fresh_Pair ),
- (metis fresh_subst_dv add fresh_Pair ) )
-
-section \<open>Prelude\<close>
-
-lemma subst_top_eq:
- "\<lbrace> z : b | TRUE \<rbrace> = \<lbrace> z : b | TRUE \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v"
-proof -
- obtain z'::x and c' where zeq: "\<lbrace> z : b | TRUE \<rbrace> = \<lbrace> z' : b | c' \<rbrace> \<and> atom z' \<sharp> (x,v)" using obtain_fresh_z2 b_of.simps by metis
- hence "\<lbrace> z' : b | TRUE \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v = \<lbrace> z' : b | TRUE \<rbrace>" using subst_tv.simps subst_cv.simps by metis
- moreover have "c' = C_true" using \<tau>.eq_iff Abs1_eq_iff(3) c.fresh flip_fresh_fresh by (metis zeq)
- ultimately show ?thesis using zeq by metis
-qed
-
-lemma wfD_subst:
- fixes \<tau>\<^sub>1::\<tau> and v::v and \<Delta>::\<Delta> and \<Theta>::\<Theta> and \<Gamma>::\<Gamma>
- assumes "\<Theta> ; \<B> ; \<Gamma> \<turnstile> v \<Rightarrow> \<tau>\<^sub>1" and "wfD \<Theta> \<B> (\<Gamma>'@((x,b\<^sub>1,c0[z0::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)) \<Delta>" and "b_of \<tau>\<^sub>1=b\<^sub>1"
- shows " \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v"
-proof -
- have "\<Theta> ; \<B> ; \<Gamma>\<turnstile>\<^sub>w\<^sub>f v : b\<^sub>1" using infer_v_v_wf assms by auto
- moreover have "(\<Gamma>'@((x,b\<^sub>1,c0[z0::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\<Gamma>\<Gamma>))[x::=v]\<^sub>\<Gamma>\<^sub>v = \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>" using subst_g_inside wfD_wf assms by metis
- ultimately show ?thesis using wf_subst assms by metis
-qed
-
-lemma subst_v_c_of:
- assumes "atom xa \<sharp> (v,x)"
- shows "c_of t[x::=v]\<^sub>\<tau>\<^sub>v xa = (c_of t xa)[x::=v]\<^sub>c\<^sub>v"
- using assms proof(nominal_induct t avoiding: x v xa rule:\<tau>.strong_induct)
- case (T_refined_type z' b' c')
- then have " c_of \<lbrace> z' : b' | c' \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v xa = c_of \<lbrace> z' : b' | c'[x::=v]\<^sub>c\<^sub>v \<rbrace> xa"
- using subst_tv.simps fresh_Pair by metis
- also have "... = c'[x::=v]\<^sub>c\<^sub>v [z'::=V_var xa]\<^sub>c\<^sub>v" using c_of.simps T_refined_type by metis
- also have "... = c' [z'::=V_var xa]\<^sub>c\<^sub>v[x::=v]\<^sub>c\<^sub>v"
- using subst_cv_commute_full[of z' v x "V_var xa" c'] subst_v_c_def T_refined_type fresh_Pair fresh_at_base v.fresh fresh_x_neq by metis
- finally show ?case using c_of.simps T_refined_type by metis
-qed
-
-section \<open>Context\<close>
-
-lemma subst_lookup:
- assumes "Some (b,c) = lookup (\<Gamma>'@((x,b\<^sub>1,c\<^sub>1)#\<^sub>\<Gamma>\<Gamma>)) y" and "x \<noteq> y" and "wfG \<Theta> \<B> (\<Gamma>'@((x,b\<^sub>1,c\<^sub>1)#\<^sub>\<Gamma>\<Gamma>))"
- shows "\<exists>d. Some (b,d) = lookup ((\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v)@\<Gamma>) y"
- using assms proof(induct \<Gamma>' rule: \<Gamma>_induct)
- case GNil
- hence "Some (b,c) = lookup \<Gamma> y" by (simp add: assms(1))
- then show ?case using subst_gv.simps by auto
-next
- case (GCons x1 b1 c1 \<Gamma>1)
- show ?case proof(cases "x1 = x")
- case True
- hence "atom x \<sharp> (\<Gamma>1 @ (x, b\<^sub>1, c\<^sub>1) #\<^sub>\<Gamma> \<Gamma>)" using GCons wfG_elims(2)
- append_g.simps by metis
- moreover have "atom x \<in> atom_dom (\<Gamma>1 @ (x, b\<^sub>1, c\<^sub>1) #\<^sub>\<Gamma> \<Gamma>)" by simp
- ultimately show ?thesis
- using forget_subst_gv not_GCons_self2 subst_gv.simps append_g.simps
- by (metis GCons.prems(3) True wfG_cons_fresh2 )
- next
- case False
- hence "((x1,b1,c1) #\<^sub>\<Gamma> \<Gamma>1)[x::=v]\<^sub>\<Gamma>\<^sub>v = (x1,b1,c1[x::=v]\<^sub>c\<^sub>v)#\<^sub>\<Gamma>\<Gamma>1[x::=v]\<^sub>\<Gamma>\<^sub>v" using subst_gv.simps by auto
- then show ?thesis proof(cases "x1=y")
- case True
- then show ?thesis using GCons using lookup.simps
- by (metis \<open>((x1, b1, c1) #\<^sub>\<Gamma> \<Gamma>1)[x::=v]\<^sub>\<Gamma>\<^sub>v = (x1, b1, c1[x::=v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>1[x::=v]\<^sub>\<Gamma>\<^sub>v\<close> append_g.simps fst_conv option.inject)
- next
- case False
- then show ?thesis using GCons using lookup.simps
- using \<open>((x1, b1, c1) #\<^sub>\<Gamma> \<Gamma>1)[x::=v]\<^sub>\<Gamma>\<^sub>v = (x1, b1, c1[x::=v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>1[x::=v]\<^sub>\<Gamma>\<^sub>v\<close> append_g.simps \<Gamma>.distinct \<Gamma>.inject wfG.simps wfG_elims by metis
- qed
- qed
-qed
-
-section \<open>Validity\<close>
-
-lemma subst_self_valid:
- fixes v::v
- assumes "\<Theta> ; \<B> ; G \<turnstile> v \<Rightarrow> \<lbrace> z : b | c \<rbrace>" and "atom z \<sharp> v"
- shows " \<Theta> ; \<B> ; G \<Turnstile> c[z::=v]\<^sub>c\<^sub>v"
-proof -
- have "c= (CE_val (V_var z) == CE_val v )" using infer_v_form2 assms by presburger
- hence "c[z::=v]\<^sub>c\<^sub>v = (CE_val (V_var z) == CE_val v )[z::=v]\<^sub>c\<^sub>v" by auto
- also have "... = (((CE_val (V_var z))[z::=v]\<^sub>c\<^sub>e\<^sub>v) == ((CE_val v)[z::=v]\<^sub>c\<^sub>e\<^sub>v))" by fastforce
- also have "... = ((CE_val v) == ((CE_val v)[z::=v]\<^sub>c\<^sub>e\<^sub>v))" using subst_cev.simps subst_vv.simps by presburger
- also have "... = (CE_val v == CE_val v )" using infer_v_form subst_cev.simps assms forget_subst_vv by presburger
- finally have *:"c[z::=v]\<^sub>c\<^sub>v = (CE_val v == CE_val v )" by auto
-
- have **:"\<Theta> ; \<B> ; G\<turnstile>\<^sub>w\<^sub>f CE_val v : b" using wfCE_valI assms infer_v_v_wf b_of.simps by metis
-
- show ?thesis proof(rule validI)
- show "\<Theta> ; \<B> ; G\<turnstile>\<^sub>w\<^sub>f c[z::=v]\<^sub>c\<^sub>v" proof -
- have "\<Theta> ; \<B> ; G\<turnstile>\<^sub>w\<^sub>f v : b" using infer_v_v_wf assms b_of.simps by metis
- moreover have "\<Theta> \<turnstile>\<^sub>w\<^sub>f ([]::\<Phi>) \<and> \<Theta> ; \<B> ; G\<turnstile>\<^sub>w\<^sub>f []\<^sub>\<Delta>" using wfD_emptyI wfPhi_emptyI infer_v_wf assms by auto
- ultimately show ?thesis using * wfCE_valI wfC_eqI by metis
- qed
- show "\<forall>i. wfI \<Theta> G i \<and> is_satis_g i G \<longrightarrow> is_satis i c[z::=v]\<^sub>c\<^sub>v" proof(rule,rule)
- fix i
- assume \<open>wfI \<Theta> G i \<and> is_satis_g i G\<close>
- thus \<open>is_satis i c[z::=v]\<^sub>c\<^sub>v\<close> using * ** is_satis_eq by auto
- qed
- qed
-qed
-
-lemma subst_valid_simple:
- fixes v::v
- assumes "\<Theta> ; \<B> ; G \<turnstile> v \<Rightarrow> \<lbrace> z0 : b | c0 \<rbrace>" and
- "atom z0 \<sharp> c" and "atom z0 \<sharp> v"
- "\<Theta>; \<B> ; (z0,b,c0)#\<^sub>\<Gamma>G \<Turnstile> c[z::=V_var z0]\<^sub>c\<^sub>v"
- shows " \<Theta> ; \<B> ; G \<Turnstile> c[z::=v]\<^sub>c\<^sub>v"
-proof -
- have " \<Theta> ; \<B> ; G \<Turnstile> c0[z0::=v]\<^sub>c\<^sub>v" using subst_self_valid assms by metis
- moreover have "atom z0 \<sharp> G" using assms valid_wf_all by meson
- moreover have "wfV \<Theta> \<B> G v b" using infer_v_v_wf assms b_of.simps by metis
- moreover have "(c[z::=V_var z0]\<^sub>c\<^sub>v)[z0::=v]\<^sub>c\<^sub>v = c[z::=v]\<^sub>c\<^sub>v" using subst_v_simple_commute assms subst_v_c_def by metis
- ultimately show ?thesis using valid_trans assms subst_defs by metis
-qed
-
-lemma wfI_subst1:
- assumes " wfI \<Theta> (G'[x::=v]\<^sub>\<Gamma>\<^sub>v @ G) i" and "wfG \<Theta> \<B> (G' @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> G)" and "eval_v i v sv" and "wfRCV \<Theta> sv b"
- shows "wfI \<Theta> (G' @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> G) ( i( x \<mapsto> sv))"
-proof -
- {
- fix xa::x and ba::b and ca::c
- assume as: "(xa,ba,ca) \<in> toSet ((G' @ ((x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> G)))"
- then have " \<exists>s. Some s = (i(x \<mapsto> sv)) xa \<and> wfRCV \<Theta> s ba"
- proof(cases "x=xa")
- case True
- have "Some sv = (i(x \<mapsto> sv)) x \<and> wfRCV \<Theta> sv b" using as assms wfI_def by auto
- moreover have "b=ba" using assms as True wfG_member_unique by metis
- ultimately show ?thesis using True by auto
- next
- case False
-
- then obtain ca' where "(xa, ba, ca') \<in> toSet (G'[x::=v]\<^sub>\<Gamma>\<^sub>v @ G)" using wfG_member_subst2 assms as by metis
- then obtain s where " Some s = i xa \<and> wfRCV \<Theta> s ba" using wfI_def assms False by blast
- thus ?thesis using False by auto
- qed
- }
- from this show ?thesis using wfI_def allI by blast
-qed
-
-lemma subst_valid:
- fixes v::v and c'::c and \<Gamma> ::\<Gamma>
- assumes "\<Theta> ; \<B> ; \<Gamma> \<Turnstile> c[z::=v]\<^sub>c\<^sub>v" and "\<Theta> ; \<B> ; \<Gamma>\<turnstile>\<^sub>w\<^sub>f v : b" and
- "\<Theta> ; \<B>\<turnstile>\<^sub>w\<^sub>f \<Gamma>" and "atom x \<sharp> c" and "atom x \<sharp> \<Gamma>" and
- "\<Theta> ; \<B>\<turnstile>\<^sub>w\<^sub>f (\<Gamma>'@(x,b,c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v ) #\<^sub>\<Gamma> \<Gamma>)" and
- "\<Theta> ; \<B> ; \<Gamma>'@(x,b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v ) #\<^sub>\<Gamma> \<Gamma> \<Turnstile> c'" (is " \<Theta> ; \<B>; ?G \<Turnstile> c'")
- shows "\<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v@\<Gamma> \<Turnstile> c'[x::=v]\<^sub>c\<^sub>v"
-proof -
- have *:"wfC \<Theta> \<B> (\<Gamma>'@(x,b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v ) #\<^sub>\<Gamma> \<Gamma>) c'" using valid.simps assms by metis
- hence "wfC \<Theta> \<B> (\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>) (c'[x::=v]\<^sub>c\<^sub>v)" using wf_subst(2)[OF *] b_of.simps assms subst_g_inside wfC_wf by metis
- moreover have "\<forall>i. wfI \<Theta> (\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>) i \<and> is_satis_g i (\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>) \<longrightarrow> is_satis i (c'[x::=v]\<^sub>c\<^sub>v)"
-
- proof(rule,rule)
- fix i
- assume as: " wfI \<Theta> (\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>) i \<and> is_satis_g i (\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>)"
-
- hence wfig: "wfI \<Theta> \<Gamma> i" using wfI_suffix infer_v_wf assms by metis
- then obtain s where s:"eval_v i v s" and b:"wfRCV \<Theta> s b" using eval_v_exist infer_v_v_wf b_of.simps assms by metis
-
- have is1: "is_satis_g ( i( x \<mapsto> s)) (\<Gamma>' @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)" proof(rule is_satis_g_i_upd2)
- show "is_satis (i(x \<mapsto> s)) (c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v)" proof -
- have "is_satis i (c[z::=v]\<^sub>c\<^sub>v)"
- using subst_valid_simple assms as valid.simps infer_v_wf assms
- is_satis_g_suffix wfI_suffix by metis
- hence "is_satis i ((c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v)[x::=v]\<^sub>c\<^sub>v)" using assms subst_v_simple_commute[of x c z v] subst_v_c_def by metis
- moreover have "\<Theta> ; \<B> ; (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v" using wfC_refl wfG_suffix assms by metis
- moreover have "\<Theta> ; \<B> ; \<Gamma>\<turnstile>\<^sub>w\<^sub>f v : b" using assms infer_v_v_wf b_of.simps by metis
- ultimately show ?thesis using subst_c_satis[OF s , of \<Theta> \<B> x b "c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v" \<Gamma> "c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v"] wfig by auto
- qed
- show "atom x \<sharp> \<Gamma>" using assms by metis
- show "wfG \<Theta> \<B> (\<Gamma>' @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)" using valid_wf_all assms by metis
- show "\<Theta> ; \<B> ; \<Gamma>\<turnstile>\<^sub>w\<^sub>f v : b" using assms infer_v_v_wf by force
- show "i \<lbrakk> v \<rbrakk> ~ s " using s by auto
- show "\<Theta> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> i" using as by auto
- show "i \<Turnstile> \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> " using as by auto
- qed
- hence "is_satis ( i( x \<mapsto> s)) c'" proof -
- have "wfI \<Theta> (\<Gamma>' @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) ( i( x \<mapsto> s))"
- using wfI_subst1[of \<Theta> \<Gamma>' x v \<Gamma> i \<B> b c z s] as b s assms by metis
- thus ?thesis using is1 valid.simps assms by presburger
- qed
-
- thus "is_satis i (c'[x::=v]\<^sub>c\<^sub>v)" using subst_c_satis_full[OF s] valid.simps as infer_v_v_wf b_of.simps assms by metis
-
- qed
- ultimately show ?thesis using valid.simps by auto
-qed
-
-lemma subst_valid_infer_v:
- fixes v::v and c'::c
- assumes "\<Theta> ; \<B> ; G \<turnstile> v \<Rightarrow> \<lbrace> z0 : b | c0 \<rbrace>" and "atom x \<sharp> c" and "atom x \<sharp> G" and "wfG \<Theta> \<B> (G'@(x,b,c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v ) #\<^sub>\<Gamma> G)" and "atom z0 \<sharp> v"
- " \<Theta>;\<B>;(z0,b,c0)#\<^sub>\<Gamma>G \<Turnstile> c[z::=V_var z0]\<^sub>c\<^sub>v" and "atom z0 \<sharp> c" and
- " \<Theta>;\<B>;G'@(x,b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v ) #\<^sub>\<Gamma> G \<Turnstile> c'" (is " \<Theta> ; \<B>; ?G \<Turnstile> c'")
- shows " \<Theta>;\<B>;G'[x::=v]\<^sub>\<Gamma>\<^sub>v@G \<Turnstile> c'[x::=v]\<^sub>c\<^sub>v"
-proof -
- have "\<Theta> ; \<B> ; G \<Turnstile> c[z::=v]\<^sub>c\<^sub>v"
- using infer_v_wf subst_valid_simple valid.simps assms using subst_valid_simple assms valid.simps infer_v_wf assms
- is_satis_g_suffix wfI_suffix by metis
- moreover have "wfV \<Theta> \<B> G v b" and "wfG \<Theta> \<B> G"
- using assms infer_v_wf b_of.simps apply metis using assms infer_v_wf by metis
- ultimately show ?thesis using assms subst_valid by metis
-qed
-
-section \<open>Subtyping\<close>
-
-lemma subst_subtype:
- fixes v::v
- assumes "\<Theta> ; \<B> ; \<Gamma> \<turnstile> v \<Rightarrow> (\<lbrace>z0:b|c0\<rbrace>)" and
- " \<Theta>;\<B>;\<Gamma> \<turnstile> (\<lbrace>z0:b|c0\<rbrace>) \<lesssim> (\<lbrace> z : b | c \<rbrace>)" and
- " \<Theta>;\<B>;\<Gamma>'@((x,b,c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\<Gamma>\<Gamma>) \<turnstile> (\<lbrace> z1 : b1 | c1 \<rbrace>) \<lesssim> (\<lbrace> z2 : b1 | c2 \<rbrace>)" (is " \<Theta> ; \<B>; ?G1 \<turnstile> ?t1 \<lesssim> ?t2" ) and
- "atom z \<sharp> (x,v) \<and> atom z0 \<sharp> (c,x,v,z,\<Gamma>) \<and> atom z1 \<sharp> (x,v) \<and> atom z2 \<sharp> (x,v)" and "wsV \<Theta> \<B> \<Gamma> v"
- shows " \<Theta>;\<B>;\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v@\<Gamma> \<turnstile> \<lbrace> z1 : b1 | c1 \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v \<lesssim> \<lbrace> z2 : b1 | c2 \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v"
-proof -
- have z2: "atom z2 \<sharp> (x,v) " using assms by auto
- hence "x \<noteq> z2" by auto
-
- obtain xx::x where xxf: "atom xx \<sharp> (x,z1, c1, z2, c2, \<Gamma>' @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>, c1[x::=v]\<^sub>c\<^sub>v, c2[x::=v]\<^sub>c\<^sub>v, \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>,
- (\<Theta> , \<B> , \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v@\<Gamma>, z1 , c1[x::=v]\<^sub>c\<^sub>v , z2 , c2[x::=v]\<^sub>c\<^sub>v ))" (is "atom xx \<sharp> ?tup")
- using obtain_fresh by blast
- hence xxf2: "atom xx \<sharp> (z1, c1, z2, c2, \<Gamma>' @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)" using fresh_prod9 fresh_prod5 by fast
-
- have vd1: " \<Theta>;\<B>;((xx, b1, c1[z1::=V_var xx]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>')[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<Turnstile> (c2[z2::=V_var xx]\<^sub>c\<^sub>v)[x::=v]\<^sub>c\<^sub>v"
- proof(rule subst_valid_infer_v[of \<Theta> _ _ _ z0 b c0 _ c, where z=z])
- show "\<Theta> ; \<B> ; \<Gamma> \<turnstile> v \<Rightarrow> \<lbrace> z0 : b | c0 \<rbrace>" using assms by auto
-
- show xf: "atom x \<sharp> \<Gamma>" using subtype_g_wf wfG_inside_fresh_suffix assms by metis
-
- show "atom x \<sharp> c" proof -
- have "wfT \<Theta> \<B> \<Gamma> (\<lbrace> z : b | c \<rbrace>)" using subtype_wf[OF assms(2)] by auto
- moreover have "x \<noteq> z" using assms(4)
- using fresh_Pair not_self_fresh by blast
- ultimately show ?thesis using xf wfT_fresh_c assms by presburger
- qed
-
- show " \<Theta> ; \<B>\<turnstile>\<^sub>w\<^sub>f ((xx, b1, c1[z1::=V_var xx]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>') @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma> "
- proof(subst append_g.simps,rule wfG_consI)
- show *: \<open> \<Theta> ; \<B>\<turnstile>\<^sub>w\<^sub>f \<Gamma>' @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<close> using subtype_g_wf assms by metis
- show \<open>atom xx \<sharp> \<Gamma>' @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>\<close> using xxf fresh_prod9 by metis
- show \<open> \<Theta> ; \<B>\<turnstile>\<^sub>w\<^sub>f b1 \<close> using subtype_elims[OF assms(3)] wfT_wfC wfC_wf wfG_cons by metis
- show "\<Theta> ; \<B> ; (xx, b1, TRUE) #\<^sub>\<Gamma> \<Gamma>' @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f c1[z1::=V_var xx]\<^sub>c\<^sub>v " proof(rule wfT_wfC)
- have "\<lbrace> z1 : b1 | c1 \<rbrace> = \<lbrace> xx : b1 | c1[z1::=V_var xx]\<^sub>c\<^sub>v \<rbrace> " using xxf fresh_prod9 type_eq_subst xxf2 fresh_prodN by metis
- thus "\<Theta> ; \<B> ; \<Gamma>' @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> xx : b1 | c1[z1::=V_var xx]\<^sub>c\<^sub>v \<rbrace> " using subtype_wfT[OF assms(3)] by metis
- show "atom xx \<sharp> \<Gamma>' @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>" using xxf fresh_prod9 by metis
- qed
- qed
-
- show "atom z0 \<sharp> v" using assms fresh_prod5 by auto
- have "\<Theta> ; \<B> ; (z0, b, c0) #\<^sub>\<Gamma> \<Gamma> \<Turnstile> c[z::=V_var z0]\<^sub>v "
- apply(rule obtain_fresh[of "(z0,c0, \<Gamma>, c, z)"],rule subtype_valid[OF assms(2), THEN valid_flip],
- (fastforce simp add: assms fresh_prodN)+) done
- thus "\<Theta> ; \<B> ; (z0, b, c0) #\<^sub>\<Gamma> \<Gamma> \<Turnstile> c[z::=V_var z0]\<^sub>c\<^sub>v " using subst_defs by auto
-
- show "atom z0 \<sharp> c" using assms fresh_prod5 by auto
- show "\<Theta> ; \<B> ; ((xx, b1, c1[z1::=V_var xx]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>') @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<Turnstile> c2[z2::=V_var xx]\<^sub>c\<^sub>v "
- using subtype_valid assms(3) xxf xxf2 fresh_prodN append_g.simps subst_defs by metis
- qed
-
- have xfw1: "atom z1 \<sharp> v \<and> atom x \<sharp> [ xx ]\<^sup>v \<and> x \<noteq> z1"
- apply(intro conjI)
- apply(simp add: assms xxf fresh_at_base fresh_prodN freshers fresh_x_neq)+
- using fresh_x_neq fresh_prodN xxf apply blast
- using fresh_x_neq fresh_prodN assms by blast
-
- have xfw2: "atom z2 \<sharp> v \<and> atom x \<sharp> [ xx ]\<^sup>v \<and> x \<noteq> z2"
- apply(auto simp add: assms xxf fresh_at_base fresh_prodN freshers)
- by(insert xxf fresh_at_base fresh_prodN assms, fast+)
-
- have wf1: "wfT \<Theta> \<B> (\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v@\<Gamma>) (\<lbrace> z1 : b1 | c1[x::=v]\<^sub>c\<^sub>v \<rbrace>)" proof -
- have "wfT \<Theta> \<B> (\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v@\<Gamma>) (\<lbrace> z1 : b1 | c1 \<rbrace>)[x::=v]\<^sub>\<tau>\<^sub>v"
- using wf_subst(4) assms b_of.simps infer_v_v_wf subtype_wf subst_tv.simps subst_g_inside wfT_wf by metis
- moreover have "atom z1 \<sharp> (x,v)" using assms by auto
- ultimately show ?thesis using subst_tv.simps by auto
- qed
- moreover have wf2: "wfT \<Theta> \<B> (\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v@\<Gamma>) (\<lbrace> z2 : b1 | c2[x::=v]\<^sub>c\<^sub>v \<rbrace>)" proof -
- have "wfT \<Theta> \<B> (\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v@\<Gamma>) (\<lbrace> z2 : b1 | c2 \<rbrace>)[x::=v]\<^sub>\<tau>\<^sub>v" using wf_subst(4) assms b_of.simps infer_v_v_wf subtype_wf subst_tv.simps subst_g_inside wfT_wf by metis
- moreover have "atom z2 \<sharp> (x,v)" using assms by auto
- ultimately show ?thesis using subst_tv.simps by auto
- qed
- moreover have "\<Theta> ; \<B> ; (xx, b1, c1[x::=v]\<^sub>c\<^sub>v[z1::=V_var xx]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> (\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> ) \<Turnstile> (c2[x::=v]\<^sub>c\<^sub>v)[z2::=V_var xx]\<^sub>c\<^sub>v" proof -
- have "xx \<noteq> x" using xxf fresh_Pair fresh_at_base by fast
- hence "((xx, b1, subst_cv c1 z1 (V_var xx) ) #\<^sub>\<Gamma> \<Gamma>')[x::=v]\<^sub>\<Gamma>\<^sub>v = (xx, b1, (subst_cv c1 z1 (V_var xx) )[x::=v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> (\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v)"
- using subst_gv.simps by auto
- moreover have "(c1[z1::=V_var xx]\<^sub>c\<^sub>v )[x::=v]\<^sub>c\<^sub>v = (c1[x::=v]\<^sub>c\<^sub>v) [z1::=V_var xx]\<^sub>c\<^sub>v" using subst_cv_commute_full xfw1 by metis
- moreover have "c2[z2::=[ xx ]\<^sup>v]\<^sub>c\<^sub>v[x::=v]\<^sub>c\<^sub>v = (c2[x::=v]\<^sub>c\<^sub>v)[z2::=V_var xx]\<^sub>c\<^sub>v" using subst_cv_commute_full xfw2 by metis
- ultimately show ?thesis using vd1 append_g.simps by metis
- qed
- moreover have "atom xx \<sharp> (\<Theta> , \<B> , \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v@\<Gamma>, z1 , c1[x::=v]\<^sub>c\<^sub>v , z2 ,c2[x::=v]\<^sub>c\<^sub>v )"
- using xxf fresh_prodN by metis
- ultimately have "\<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v@\<Gamma> \<turnstile> \<lbrace> z1 : b1 | c1[x::=v]\<^sub>c\<^sub>v \<rbrace> \<lesssim> \<lbrace> z2 : b1 | c2[x::=v]\<^sub>c\<^sub>v \<rbrace>"
- using subtype_baseI subst_defs by metis
- thus ?thesis using subst_tv.simps assms by presburger
-qed
-
-lemma subst_subtype_tau:
- fixes v::v
- assumes "\<Theta> ; \<B> ; \<Gamma> \<turnstile> v \<Rightarrow> \<tau>" and
- "\<Theta> ; \<B> ; \<Gamma> \<turnstile> \<tau> \<lesssim> (\<lbrace> z : b | c \<rbrace>)"
- "\<Theta> ; \<B> ; \<Gamma>'@((x,b,c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\<Gamma>\<Gamma>) \<turnstile> \<tau>1 \<lesssim> \<tau>2" and
- "atom z \<sharp> (x,v)"
- shows "\<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v@\<Gamma> \<turnstile> \<tau>1[x::=v]\<^sub>\<tau>\<^sub>v \<lesssim> \<tau>2[x::=v]\<^sub>\<tau>\<^sub>v"
-proof -
- obtain z0 and b0 and c0 where zbc0: "\<tau>=(\<lbrace> z0 : b0 | c0 \<rbrace>) \<and> atom z0 \<sharp> (c,x,v,z,\<Gamma>)"
- using obtain_fresh_z by metis
- obtain z1 and b1 and c1 where zbc1: "\<tau>1=(\<lbrace> z1 : b1 | c1 \<rbrace>) \<and> atom z1 \<sharp> (x,v)"
- using obtain_fresh_z by metis
- obtain z2 and b2 and c2 where zbc2: "\<tau>2=(\<lbrace> z2 : b2 | c2 \<rbrace>) \<and> atom z2 \<sharp> (x,v)"
- using obtain_fresh_z by metis
-
- have "b0=b" using subtype_eq_base zbc0 assms by blast
-
- hence vinf: "\<Theta> ; \<B> ; \<Gamma> \<turnstile> v \<Rightarrow> \<lbrace> z0 : b | c0 \<rbrace>" using assms zbc0 by blast
- have vsub: "\<Theta> ; \<B> ; \<Gamma> \<turnstile>\<lbrace> z0 : b | c0 \<rbrace> \<lesssim> \<lbrace> z : b | c \<rbrace>" using assms zbc0 \<open>b0=b\<close> by blast
- have beq:"b1=b2" using subtype_eq_base
- using zbc1 zbc2 assms by blast
- have "\<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> \<lbrace> z1 : b1 | c1 \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v \<lesssim> \<lbrace> z2 : b1 | c2 \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v"
- proof(rule subst_subtype[OF vinf vsub] )
- show "\<Theta> ; \<B> ; \<Gamma>'@((x,b,c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\<Gamma>\<Gamma>) \<turnstile> \<lbrace> z1 : b1 | c1 \<rbrace> \<lesssim> \<lbrace> z2 : b1 | c2 \<rbrace>"
- using beq assms zbc1 zbc2 by auto
- show "atom z \<sharp> (x, v) \<and> atom z0 \<sharp> (c, x, v, z, \<Gamma>) \<and> atom z1 \<sharp> (x, v) \<and> atom z2 \<sharp> (x, v)"
- using zbc0 zbc1 zbc2 assms by blast
- show "wfV \<Theta> \<B> \<Gamma> v (b_of \<tau>)" using infer_v_wf assms by simp
- qed
-
- thus ?thesis using zbc1 zbc2 \<open>b1=b2\<close> assms by blast
-qed
-
-lemma subtype_if1:
- fixes v::v
- assumes "P ; \<B> ; \<Gamma> \<turnstile> t1 \<lesssim> t2" and "wfV P \<B> \<Gamma> v (base_for_lit l)" and
- "atom z1 \<sharp> v" and "atom z2 \<sharp> v" and "atom z1 \<sharp> t1" and "atom z2 \<sharp> t2" and "atom z1 \<sharp> \<Gamma>" and "atom z2 \<sharp> \<Gamma>"
- shows "P ; \<B> ; \<Gamma> \<turnstile> \<lbrace> z1 : b_of t1 | CE_val v == CE_val (V_lit l) IMP (c_of t1 z1) \<rbrace> \<lesssim> \<lbrace> z2 : b_of t2 | CE_val v == CE_val (V_lit l) IMP (c_of t2 z2) \<rbrace>"
-proof -
- obtain z1' where t1: "t1 = \<lbrace> z1' : b_of t1 | c_of t1 z1'\<rbrace> \<and> atom z1' \<sharp> (z1,\<Gamma>,t1)" using obtain_fresh_z_c_of by metis
- obtain z2' where t2: "t2 = \<lbrace> z2' : b_of t2 | c_of t2 z2'\<rbrace> \<and> atom z2' \<sharp> (z2,t2) " using obtain_fresh_z_c_of by metis
- have beq:"b_of t1 = b_of t2" using subtype_eq_base2 assms by auto
-
- have c1: "(c_of t1 z1')[z1'::=[ z1 ]\<^sup>v]\<^sub>c\<^sub>v = c_of t1 z1" using c_of_switch t1 assms by simp
- have c2: "(c_of t2 z2')[z2'::=[ z2 ]\<^sup>v]\<^sub>c\<^sub>v = c_of t2 z2" using c_of_switch t2 assms by simp
-
- have "P ; \<B> ; \<Gamma> \<turnstile> \<lbrace> z1 : b_of t1 | [ v ]\<^sup>c\<^sup>e == [ [ l ]\<^sup>v ]\<^sup>c\<^sup>e IMP (c_of t1 z1')[z1'::=[z1]\<^sup>v]\<^sub>v \<rbrace> \<lesssim> \<lbrace> z2 : b_of t1 | [ v ]\<^sup>c\<^sup>e == [ [ l ]\<^sup>v ]\<^sup>c\<^sup>e IMP (c_of t2 z2')[z2'::=[z2]\<^sup>v]\<^sub>v \<rbrace>"
- proof(rule subtype_if)
- show \<open>P ; \<B> ; \<Gamma> \<turnstile> \<lbrace> z1' : b_of t1 | c_of t1 z1' \<rbrace> \<lesssim> \<lbrace> z2' : b_of t1 | c_of t2 z2' \<rbrace>\<close> using t1 t2 assms beq by auto
- show \<open> P ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> z1 : b_of t1 | [ v ]\<^sup>c\<^sup>e == [ [ l ]\<^sup>v ]\<^sup>c\<^sup>e IMP (c_of t1 z1')[z1'::=[ z1 ]\<^sup>v]\<^sub>v \<rbrace> \<close> using wfT_wfT_if_rev assms subtype_wfT c1 subst_defs by metis
- show \<open> P ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> z2 : b_of t1 | [ v ]\<^sup>c\<^sup>e == [ [ l ]\<^sup>v ]\<^sup>c\<^sup>e IMP (c_of t2 z2')[z2'::=[ z2 ]\<^sup>v]\<^sub>v \<rbrace> \<close> using wfT_wfT_if_rev assms subtype_wfT c2 subst_defs beq by metis
- show \<open>atom z1 \<sharp> v\<close> using assms by auto
- show \<open>atom z1' \<sharp> \<Gamma>\<close> using t1 by auto
- show \<open>atom z1 \<sharp> c_of t1 z1'\<close> using t1 assms c_of_fresh by force
- show \<open>atom z2 \<sharp> c_of t2 z2'\<close> using t2 assms c_of_fresh by force
- show \<open>atom z2 \<sharp> v\<close> using assms by auto
- qed
- then show ?thesis using t1 t2 assms c1 c2 beq subst_defs by metis
-qed
-
-section \<open>Values\<close>
-
-lemma subst_infer_aux:
- fixes \<tau>\<^sub>1::\<tau> and v'::v
- assumes "\<Theta> ; \<B> ; \<Gamma> \<turnstile> v'[x::=v]\<^sub>v\<^sub>v \<Rightarrow> \<tau>\<^sub>1" and "\<Theta> ; \<B> ; \<Gamma>' \<turnstile> v' \<Rightarrow> \<tau>\<^sub>2" and "b_of \<tau>\<^sub>1 = b_of \<tau>\<^sub>2"
- shows "\<tau>\<^sub>1 = (\<tau>\<^sub>2[x::=v]\<^sub>\<tau>\<^sub>v)"
-proof -
- obtain z1 and b1 where zb1: "\<tau>\<^sub>1 = (\<lbrace> z1 : b1 | C_eq (CE_val (V_var z1)) (CE_val (v'[x::=v]\<^sub>v\<^sub>v)) \<rbrace>) \<and> atom z1 \<sharp> ((CE_val (v'[x::=v]\<^sub>v\<^sub>v), CE_val v),v'[x::=v]\<^sub>v\<^sub>v)"
- using infer_v_form_fresh[OF assms(1)] by fastforce
- obtain z2 and b2 where zb2: "\<tau>\<^sub>2 = (\<lbrace> z2 : b2 | C_eq (CE_val (V_var z2)) (CE_val v') \<rbrace>) \<and> atom z2 \<sharp> ((CE_val (v'[x::=v]\<^sub>v\<^sub>v), CE_val v,x,v),v')"
- using infer_v_form_fresh [OF assms(2)] by fastforce
- have beq: "b1 = b2" using assms zb1 zb2 by simp
-
- hence "(\<lbrace> z2 : b2 | C_eq (CE_val (V_var z2)) (CE_val v') \<rbrace>)[x::=v]\<^sub>\<tau>\<^sub>v = (\<lbrace> z2 : b2 | C_eq (CE_val (V_var z2)) (CE_val (v'[x::=v]\<^sub>v\<^sub>v)) \<rbrace>)"
- using subst_tv.simps subst_cv.simps subst_ev.simps forget_subst_vv[of x "V_var z2"] zb2 by force
- also have "... = (\<lbrace> z1 : b1 | C_eq (CE_val (V_var z1)) (CE_val (v'[x::=v]\<^sub>v\<^sub>v)) \<rbrace>)"
- using type_e_eq[of z2 "CE_val (v'[x::=v]\<^sub>v\<^sub>v)"z1 b1 ] zb1 zb2 fresh_PairD(1) assms beq by metis
- finally show ?thesis using zb1 zb2 by argo
-qed
-
-lemma subst_t_b_eq:
- fixes x::x and v::v
- shows "b_of (\<tau>[x::=v]\<^sub>\<tau>\<^sub>v) = b_of \<tau>"
-proof -
- obtain z and b and c where "\<tau> = \<lbrace> z : b | c \<rbrace> \<and> atom z \<sharp> (x,v)"
- using has_fresh_z by blast
- thus ?thesis using subst_tv.simps by simp
-qed
-
-lemma fresh_g_fresh_v:
- fixes x::x
- assumes "atom x \<sharp> \<Gamma>" and "wfV \<Theta> \<B> \<Gamma> v b"
- shows "atom x \<sharp> v"
- using assms wfV_supp wfX_wfY wfG_atoms_supp_eq fresh_def
- by (metis wfV_x_fresh)
-
-lemma infer_v_fresh_g_fresh_v:
- fixes x::x and \<Gamma>::\<Gamma> and v::v
- assumes "atom x \<sharp> \<Gamma>'@\<Gamma>" and "\<Theta> ; \<B> ; \<Gamma> \<turnstile> v \<Rightarrow> \<tau>"
- shows "atom x \<sharp> v"
-proof -
- have "atom x \<sharp> \<Gamma>" using fresh_suffix assms by auto
- moreover have "wfV \<Theta> \<B> \<Gamma> v (b_of \<tau>)" using infer_v_wf assms by auto
- ultimately show ?thesis using fresh_g_fresh_v by metis
-qed
-
-lemma infer_v_fresh_g_fresh_xv:
- fixes xa::x and v::v and \<Gamma>::\<Gamma>
- assumes "atom xa \<sharp> \<Gamma>'@((x,b,c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\<Gamma>\<Gamma>)" and "\<Theta> ; \<B> ; \<Gamma> \<turnstile> v \<Rightarrow> \<tau>"
- shows "atom xa \<sharp> (x,v)"
-proof -
- have "atom xa \<sharp> x" using assms fresh_in_g fresh_def by blast
- moreover have "\<Gamma>'@((x,b,c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\<Gamma>\<Gamma>) = ((\<Gamma>'@(x,b,c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\<Gamma>GNil)@\<Gamma>)" using append_g.simps append_g_assoc by simp
- moreover hence "atom xa \<sharp> v" using infer_v_fresh_g_fresh_v assms by metis
- ultimately show ?thesis by auto
-qed
-
-lemma wfG_subst_infer_v:
- fixes v::v
- assumes "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>' @ (x, b, c0[z0::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>" and "\<Theta> ; \<B> ; \<Gamma> \<turnstile> v \<Rightarrow> \<tau>" and "b_of \<tau> = b"
- shows "\<Theta> ; \<B>\<turnstile>\<^sub>w\<^sub>f \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> "
- using wfG_subst_wfV infer_v_v_wf assms by auto
-
-lemma fresh_subst_gv_inside:
- fixes \<Gamma>::\<Gamma>
- assumes "atom z \<sharp> \<Gamma>' @ (x, b\<^sub>1, c0[z0::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>" and "atom z \<sharp> v"
- shows "atom z \<sharp> \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v@\<Gamma>"
- unfolding fresh_append_g using fresh_append_g assms fresh_subst_gv fresh_GCons by metis
-
-lemma subst_t:
- fixes x::x and b::b and xa::x
- assumes "atom z \<sharp> x" and "atom z \<sharp> v"
- shows "(\<lbrace> z : b | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ v'[x::=v]\<^sub>v\<^sub>v ]\<^sup>c\<^sup>e \<rbrace>) = (\<lbrace> z : b | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ v']\<^sup>c\<^sup>e \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v)"
- using assms subst_vv.simps subst_tv.simps subst_cv.simps subst_cev.simps by auto
-
-lemma infer_l_fresh:
- assumes "\<turnstile> l \<Rightarrow> \<tau>"
- shows "atom x \<sharp> \<tau>"
-proof -
- have "[] ; {||} \<turnstile>\<^sub>w\<^sub>f GNil" using wfG_nilI wfTh_emptyI by auto
- hence "[] ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f \<tau>" using assms infer_l_wf by auto
- thus ?thesis using fresh_def wfT_supp by force
-qed
-
-lemma subst_infer_v:
- fixes v::v and v'::v
- assumes "\<Theta> ; \<B> ; \<Gamma>'@((x,b\<^sub>1,c0[z0::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\<Gamma>\<Gamma>) \<turnstile> v' \<Rightarrow> \<tau>\<^sub>2" and
- "\<Theta> ; \<B> ; \<Gamma> \<turnstile> v \<Rightarrow> \<tau>\<^sub>1" and
- "\<Theta> ; \<B> ; \<Gamma> \<turnstile> \<tau>\<^sub>1 \<lesssim> (\<lbrace> z0 : b\<^sub>1 | c0 \<rbrace>)" and "atom z0 \<sharp> (x,v)"
- shows "\<Theta> ; \<B> ; (\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v)@\<Gamma> \<turnstile> v'[x::=v]\<^sub>v\<^sub>v \<Rightarrow> \<tau>\<^sub>2[x::=v]\<^sub>\<tau>\<^sub>v"
- using assms proof(nominal_induct "\<Gamma>'@((x,b\<^sub>1,c0[z0::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\<Gamma>\<Gamma>)" v' \<tau>\<^sub>2 avoiding: x v rule: infer_v.strong_induct)
- case (infer_v_varI \<Theta> \<B> b c xa z)
- have "\<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> [ xa ]\<^sup>v[x::=v]\<^sub>v\<^sub>v \<Rightarrow> \<lbrace> z : b | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ xa ]\<^sup>v[x::=v]\<^sub>v\<^sub>v ]\<^sup>c\<^sup>e \<rbrace>"
- proof(cases "x=xa")
- case True
- have "\<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> v \<Rightarrow> \<lbrace> z : b | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e \<rbrace>"
- proof(rule infer_v_g_weakening)
- show *:\<open> \<Theta> ; \<B> ; \<Gamma> \<turnstile> v \<Rightarrow> \<lbrace> z : b | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e \<rbrace>\<close>
- using infer_v_form infer_v_varI
- by (metis True lookup_inside_unique_b lookup_inside_wf ms_fresh_all(32) subtype_eq_base type_e_eq)
- show \<open>toSet \<Gamma> \<subseteq> toSet (\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>)\<close> by simp
- have "\<Theta> ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b\<^sub>1" using infer_v_wf subtype_eq_base2 b_of.simps infer_v_varI by metis
- thus \<open> \<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<close>
- using wfG_subst[OF infer_v_varI(3), of \<Gamma>' x b\<^sub>1 "c0[z0::=[ x ]\<^sup>v]\<^sub>c\<^sub>v" \<Gamma> v] subst_g_inside infer_v_varI by metis
- qed
-
- thus ?thesis using subst_vv.simps True by simp
- next
- case False
- then obtain c' where c: "Some (b, c') = lookup (\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>) xa" using lookup_subst2 infer_v_varI by metis
- have "\<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> [ xa ]\<^sup>v \<Rightarrow> \<lbrace> z : b | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ xa ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrace>"
- proof
- have "\<Theta> ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b\<^sub>1" using infer_v_wf subtype_eq_base2 b_of.simps infer_v_varI by metis
- thus "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>" using infer_v_varI
- using wfG_subst[OF infer_v_varI(3), of \<Gamma>' x b\<^sub>1 "c0[z0::=[ x ]\<^sup>v]\<^sub>c\<^sub>v" \<Gamma> v] subst_g_inside infer_v_varI by metis
- show "atom z \<sharp> xa" using infer_v_varI by auto
- show "Some (b, c') = lookup (\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>) xa" using c by auto
- show "atom z \<sharp> (\<Theta>, \<B>, \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>)" by (fresh_mth add: infer_v_varI fresh_subst_gv_inside)
- qed
- then show ?thesis using subst_vv.simps False by simp
- qed
- thus ?case using subst_t fresh_prodN infer_v_varI by metis
-next
- case (infer_v_litI \<Theta> \<B> l \<tau>)
- show ?case unfolding subst_vv.simps proof
- show "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>" using wfG_subst_infer_v infer_v_litI subtype_eq_base2 b_of.simps by metis
- have "atom x \<sharp> \<tau>" using infer_v_litI infer_l_fresh by metis
- thus "\<turnstile> l \<Rightarrow> \<tau>[x::=v]\<^sub>\<tau>\<^sub>v" using infer_v_litI type_v_subst_fresh by simp
- qed
-next
- case (infer_v_pairI z v1 v2 \<Theta> \<B> t1 t2)
- have " \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @
- \<Gamma> \<turnstile> [ v1[x::=v]\<^sub>v\<^sub>v , v2[x::=v]\<^sub>v\<^sub>v ]\<^sup>v \<Rightarrow> \<lbrace> z : [ b_of t1[x::=v]\<^sub>\<tau>\<^sub>v , b_of
- t2[x::=v]\<^sub>\<tau>\<^sub>v ]\<^sup>b | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ v1[x::=v]\<^sub>v\<^sub>v , v2[x::=v]\<^sub>v\<^sub>v ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrace>"
- proof
- show \<open>atom z \<sharp> (v1[x::=v]\<^sub>v\<^sub>v, v2[x::=v]\<^sub>v\<^sub>v)\<close> by (fresh_mth add: infer_v_pairI)
- show \<open>atom z \<sharp> (\<Theta>, \<B>, \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>)\<close> by (fresh_mth add: infer_v_pairI fresh_subst_gv_inside)
- show \<open> \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> v1[x::=v]\<^sub>v\<^sub>v \<Rightarrow> t1[x::=v]\<^sub>\<tau>\<^sub>v\<close> using infer_v_pairI by metis
- show \<open> \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> v2[x::=v]\<^sub>v\<^sub>v \<Rightarrow> t2[x::=v]\<^sub>\<tau>\<^sub>v\<close> using infer_v_pairI by metis
- qed
- then show ?case using subst_vv.simps subst_tv.simps infer_v_pairI b_of_subst by simp
-next
- case (infer_v_consI s dclist \<Theta> dc tc \<B> va tv z)
-
- have " \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> (V_cons s dc va[x::=v]\<^sub>v\<^sub>v) \<Rightarrow> \<lbrace> z : B_id s | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ V_cons s dc va[x::=v]\<^sub>v\<^sub>v ]\<^sup>c\<^sup>e \<rbrace>"
- proof
- show td:\<open>AF_typedef s dclist \<in> set \<Theta>\<close> using infer_v_consI by auto
- show dc:\<open>(dc, tc) \<in> set dclist\<close> using infer_v_consI by auto
- show \<open> \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> va[x::=v]\<^sub>v\<^sub>v \<Rightarrow> tv[x::=v]\<^sub>\<tau>\<^sub>v\<close> using infer_v_consI by auto
- have \<open>\<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> tv[x::=v]\<^sub>\<tau>\<^sub>v \<lesssim> tc[x::=v]\<^sub>\<tau>\<^sub>v\<close>
- using subst_subtype_tau infer_v_consI by metis
- moreover have "atom x \<sharp> tc" using wfTh_lookup_supp_empty[OF td dc] infer_v_wf infer_v_consI fresh_def by fast
- ultimately show \<open>\<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> tv[x::=v]\<^sub>\<tau>\<^sub>v \<lesssim> tc\<close> by simp
- show \<open>atom z \<sharp> va[x::=v]\<^sub>v\<^sub>v\<close> using infer_v_consI by auto
- show \<open>atom z \<sharp> (\<Theta>, \<B>, \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>)\<close> by (fresh_mth add: infer_v_consI fresh_subst_gv_inside)
- qed
- thus ?case using subst_vv.simps subst_t[of z x v ] infer_v_consI by metis
-
-next
- case (infer_v_conspI s bv dclist \<Theta> dc tc \<B> va tv b z)
- have "\<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> (V_consp s dc b va[x::=v]\<^sub>v\<^sub>v) \<Rightarrow> \<lbrace> z : B_app s b | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ V_consp s dc b va[x::=v]\<^sub>v\<^sub>v ]\<^sup>c\<^sup>e \<rbrace>"
- proof
- show td:\<open>AF_typedef_poly s bv dclist \<in> set \<Theta>\<close> using infer_v_conspI by auto
- show dc:\<open>(dc, tc) \<in> set dclist\<close> using infer_v_conspI by auto
- show \<open> \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> va[x::=v]\<^sub>v\<^sub>v \<Rightarrow> tv[x::=v]\<^sub>\<tau>\<^sub>v\<close> using infer_v_conspI by metis
- have \<open>\<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> tv[x::=v]\<^sub>\<tau>\<^sub>v \<lesssim> tc[bv::=b]\<^sub>\<tau>\<^sub>b[x::=v]\<^sub>\<tau>\<^sub>v\<close>
- using subst_subtype_tau infer_v_conspI by metis
- moreover have "atom x \<sharp> tc[bv::=b]\<^sub>\<tau>\<^sub>b" proof -
- have "supp tc \<subseteq> { atom bv }" using wfTh_poly_lookup_supp infer_v_conspI wfX_wfY by metis
- hence "atom x \<sharp> tc" using x_not_in_b_set
- using fresh_def by fastforce
- moreover have "atom x \<sharp> b" using x_fresh_b by auto
- ultimately show ?thesis using fresh_subst_if subst_b_\<tau>_def by metis
- qed
- ultimately show \<open>\<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> tv[x::=v]\<^sub>\<tau>\<^sub>v \<lesssim> tc[bv::=b]\<^sub>\<tau>\<^sub>b\<close> by simp
- show \<open>atom z \<sharp> (\<Theta>, \<B>, \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>, va[x::=v]\<^sub>v\<^sub>v, b)\<close> proof -
- have "atom z \<sharp> va[x::=v]\<^sub>v\<^sub>v" using fresh_subst_v_if infer_v_conspI subst_v_v_def by metis
- moreover have "atom z \<sharp> \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>" using fresh_subst_gv_inside infer_v_conspI by metis
- ultimately show ?thesis using fresh_prodN infer_v_conspI by metis
- qed
- show \<open>atom bv \<sharp> (\<Theta>, \<B>, \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>, va[x::=v]\<^sub>v\<^sub>v, b)\<close> proof -
- have "atom bv \<sharp> va[x::=v]\<^sub>v\<^sub>v" using fresh_subst_v_if infer_v_conspI subst_v_v_def by metis
- moreover have "atom bv \<sharp> \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>" using fresh_subst_gv_inside infer_v_conspI by metis
- ultimately show ?thesis using fresh_prodN infer_v_conspI by metis
- qed
- show "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f b" using infer_v_conspI by auto
- qed
- thus ?case using subst_vv.simps subst_t[of z x v ] infer_v_conspI by metis
-
-qed
-
-lemma subst_infer_check_v:
- fixes v::v and v'::v
- assumes "\<Theta> ; \<B> ; \<Gamma> \<turnstile> v \<Rightarrow> \<tau>\<^sub>1" and
- "check_v \<Theta> \<B> (\<Gamma>'@((x,b\<^sub>1,c0[z0::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\<Gamma>\<Gamma>)) v' \<tau>\<^sub>2" and
- "\<Theta> ; \<B> ; \<Gamma> \<turnstile> \<tau>\<^sub>1 \<lesssim> \<lbrace> z0 : b\<^sub>1 | c0 \<rbrace>" and "atom z0 \<sharp> (x,v)"
- shows "check_v \<Theta> \<B> ((\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v)@\<Gamma>) (v'[x::=v]\<^sub>v\<^sub>v) (\<tau>\<^sub>2[x::=v]\<^sub>\<tau>\<^sub>v)"
-proof -
- obtain \<tau>\<^sub>2' where t2: "infer_v \<Theta> \<B> (\<Gamma>' @ (x, b\<^sub>1, c0[z0::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) v' \<tau>\<^sub>2' \<and> \<Theta> ; \<B> ; (\<Gamma>' @ (x, b\<^sub>1, c0[z0::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) \<turnstile> \<tau>\<^sub>2' \<lesssim> \<tau>\<^sub>2"
- using check_v_elims assms by blast
- hence "infer_v \<Theta> \<B> ((\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v)@\<Gamma>) (v'[x::=v]\<^sub>v\<^sub>v) (\<tau>\<^sub>2'[x::=v]\<^sub>\<tau>\<^sub>v)"
- using subst_infer_v[OF _ assms(1) assms(3) assms(4)] by blast
- moreover hence "\<Theta>; \<B> ; ((\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v)@\<Gamma>) \<turnstile>\<tau>\<^sub>2'[x::=v]\<^sub>\<tau>\<^sub>v \<lesssim> \<tau>\<^sub>2[x::=v]\<^sub>\<tau>\<^sub>v"
- using subst_subtype assms t2 by (meson subst_subtype_tau subtype_trans)
- ultimately show ?thesis using check_v.intros by blast
-qed
-
-lemma type_veq_subst[simp]:
- assumes "atom z \<sharp> (x,v)"
- shows "\<lbrace> z : b | CE_val (V_var z) == CE_val v' \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v = \<lbrace> z : b | CE_val (V_var z) == CE_val v'[x::=v]\<^sub>v\<^sub>v \<rbrace>"
- using assms by auto
-
-lemma subst_infer_v_form:
- fixes v::v and v'::v and \<Gamma>::\<Gamma>
- assumes "\<Theta> ; \<B> ; \<Gamma> \<turnstile> v \<Rightarrow> \<tau>\<^sub>1" and
- "\<Theta> ; \<B> ; \<Gamma>'@((x,b\<^sub>1,c0[z0::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\<Gamma>\<Gamma>) \<turnstile> v' \<Rightarrow> \<tau>\<^sub>2" and "b= b_of \<tau>\<^sub>2"
- "\<Theta> ; \<B> ; \<Gamma> \<turnstile> \<tau>\<^sub>1 \<lesssim> (\<lbrace> z0 : b\<^sub>1 | c0 \<rbrace>)" and "atom z0 \<sharp> (x,v)" and "atom z3' \<sharp> (x,v,v', \<Gamma>'@((x,b\<^sub>1,c0[z0::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\<Gamma>\<Gamma>) )"
- shows \<open> \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> v'[x::=v]\<^sub>v\<^sub>v \<Rightarrow> \<lbrace> z3' : b | CE_val (V_var z3') == CE_val v'[x::=v]\<^sub>v\<^sub>v \<rbrace>\<close>
-proof -
- have "\<Theta> ; \<B> ; \<Gamma>'@((x,b\<^sub>1,c0[z0::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\<Gamma>\<Gamma>) \<turnstile> v' \<Rightarrow> \<lbrace> z3' : b_of \<tau>\<^sub>2 | C_eq (CE_val (V_var z3')) (CE_val v') \<rbrace>"
- proof(rule infer_v_form4)
- show \<open> \<Theta> ; \<B> ; \<Gamma>' @ (x, b\<^sub>1, c0[z0::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<turnstile> v' \<Rightarrow> \<tau>\<^sub>2\<close> using assms by metis
- show \<open>atom z3' \<sharp> (v', \<Gamma>' @ (x, b\<^sub>1, c0[z0::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)\<close> using assms fresh_prodN by metis
- show \<open>b_of \<tau>\<^sub>2 = b_of \<tau>\<^sub>2\<close> by auto
- qed
- hence \<open> \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> v'[x::=v]\<^sub>v\<^sub>v \<Rightarrow> \<lbrace> z3' : b_of \<tau>\<^sub>2 | CE_val (V_var z3') == CE_val v' \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v\<close>
- using subst_infer_v assms by metis
- thus ?thesis using type_veq_subst fresh_prodN assms by metis
-qed
-
-section \<open>Expressions\<close>
-
-text \<open>
- For operator, fst and snd cases, we use elimination to get one or more values, apply the substitution lemma for values. The types always have
-the same form and are equal under substitution.
- For function application, the subst value is a subtype of the value which is a subtype of the argument. The return of the function is the same
- under substitution.
-\<close>
-
-text \<open> Observe a similar pattern for each case \<close>
-
-lemma subst_infer_e:
- fixes v::v and e::e and \<Gamma>'::\<Gamma>
- assumes
- "\<Theta> ; \<Phi> ; \<B> ; G ; \<Delta> \<turnstile> e \<Rightarrow> \<tau>\<^sub>2" and "G = (\<Gamma>'@((x,b\<^sub>1,subst_cv c0 z0 (V_var x))#\<^sub>\<Gamma>\<Gamma>))"
- "\<Theta> ; \<B> ; \<Gamma> \<turnstile> v \<Rightarrow> \<tau>\<^sub>1" and
- "\<Theta>; \<B> ; \<Gamma> \<turnstile> \<tau>\<^sub>1 \<lesssim> \<lbrace> z0 : b\<^sub>1 | c0 \<rbrace>" and "atom z0 \<sharp> (x,v)"
- shows "\<Theta> ; \<Phi> ; \<B> ; ((\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v)@\<Gamma>) ; (\<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v) \<turnstile> (subst_ev e x v ) \<Rightarrow> \<tau>\<^sub>2[x::=v]\<^sub>\<tau>\<^sub>v"
- using assms proof(nominal_induct avoiding: x v rule: infer_e.strong_induct)
- case (infer_e_valI \<Theta> \<B> \<Gamma>'' \<Delta> \<Phi> v' \<tau>)
-
- have "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> ; \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> (AE_val (v'[x::=v]\<^sub>v\<^sub>v)) \<Rightarrow> \<tau>[x::=v]\<^sub>\<tau>\<^sub>v"
- proof
- show "\<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v" using wfD_subst infer_e_valI subtype_eq_base2
- by (metis b_of.simps infer_v_v_wf subst_g_inside_simple wfD_wf wf_subst(11))
- show "\<Theta>\<turnstile>\<^sub>w\<^sub>f \<Phi>" using infer_e_valI by auto
- show "\<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> v'[x::=v]\<^sub>v\<^sub>v \<Rightarrow> \<tau>[x::=v]\<^sub>\<tau>\<^sub>v" using subst_infer_v infer_e_valI using wfD_subst infer_e_valI subtype_eq_base2
- by metis
- qed
- thus ?case using subst_ev.simps by simp
-next
- case (infer_e_plusI \<Theta> \<B> \<Gamma>'' \<Delta> \<Phi> v1 z1 c1 v2 z2 c2 z3)
-
- hence z3f: "atom z3 \<sharp> CE_op Plus [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e" using e.fresh ce.fresh opp.fresh by metis
-
- obtain z3'::x where *:"atom z3' \<sharp> (x,v,AE_op Plus v1 v2, CE_op Plus [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e , AE_op Plus v1[x::=v]\<^sub>v\<^sub>v v2[x::=v]\<^sub>v\<^sub>v , CE_op Plus [v1[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e [v2[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e,\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> )"
- using obtain_fresh by metis
- hence **:"(\<lbrace> z3 : B_int | CE_val (V_var z3) == CE_op Plus [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \<rbrace>) = \<lbrace> z3' : B_int | CE_val (V_var z3') == CE_op Plus [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \<rbrace>"
- using type_e_eq infer_e_plusI fresh_Pair z3f by metis
-
- obtain z1' b1' c1' where z1:"atom z1' \<sharp> (x,v) \<and> \<lbrace> z1 : B_int | c1 \<rbrace> = \<lbrace> z1' : b1' | c1' \<rbrace>" using obtain_fresh_z by metis
- obtain z2' b2' c2' where z2:"atom z2' \<sharp> (x,v) \<and> \<lbrace> z2 : B_int | c2 \<rbrace> = \<lbrace> z2' : b2' | c2' \<rbrace>" using obtain_fresh_z by metis
-
- have bb:"b1' = B_int \<and> b2' = B_int" using z1 z2 \<tau>.eq_iff by metis
-
- have "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> ; \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> (AE_op Plus (v1[x::=v]\<^sub>v\<^sub>v) (v2[x::=v]\<^sub>v\<^sub>v)) \<Rightarrow> \<lbrace> z3' : B_int | CE_val (V_var z3') == CE_op Plus ([v1[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) ([v2[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) \<rbrace>"
- proof
- show \<open> \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<close>
- using infer_e_plusI wfD_subst subtype_eq_base2 b_of.simps by metis
- show \<open> \<Theta>\<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using infer_e_plusI by blast
- show \<open> \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> v1[x::=v]\<^sub>v\<^sub>v \<Rightarrow> \<lbrace> z1' : B_int | c1'[x::=v]\<^sub>c\<^sub>v \<rbrace>\<close> using subst_tv.simps subst_infer_v infer_e_plusI z1 bb by metis
- show \<open> \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> v2[x::=v]\<^sub>v\<^sub>v \<Rightarrow> \<lbrace> z2' : B_int | c2'[x::=v]\<^sub>c\<^sub>v \<rbrace>\<close> using subst_tv.simps subst_infer_v infer_e_plusI z2 bb by metis
- show \<open>atom z3' \<sharp> AE_op Plus v1[x::=v]\<^sub>v\<^sub>v v2[x::=v]\<^sub>v\<^sub>v\<close> using fresh_prod6 * by metis
- show \<open>atom z3' \<sharp> \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>\<close> using * by auto
- qed
- moreover have "\<lbrace> z3' : B_int | CE_val (V_var z3') == CE_op Plus ([v1[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) ([v2[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) \<rbrace> = \<lbrace> z3' : B_int | CE_val (V_var z3') == CE_op Plus [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v"
- by(subst subst_tv.simps,auto simp add: * )
- ultimately show ?case using subst_ev.simps * ** by metis
-next
- case (infer_e_leqI \<Theta> \<B> \<Gamma>'' \<Delta> \<Phi> v1 z1 c1 v2 z2 c2 z3)
-
- hence z3f: "atom z3 \<sharp> CE_op LEq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e" using e.fresh ce.fresh opp.fresh by metis
-
- obtain z3'::x where *:"atom z3' \<sharp> (x,v,AE_op LEq v1 v2, CE_op LEq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e , CE_op LEq [v1[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e [v2[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e , AE_op LEq v1[x::=v]\<^sub>v\<^sub>v v2[x::=v]\<^sub>v\<^sub>v,\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> )"
- using obtain_fresh by metis
- hence **:"(\<lbrace> z3 : B_bool | CE_val (V_var z3) == CE_op LEq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \<rbrace>) = \<lbrace> z3' : B_bool | CE_val (V_var z3') == CE_op LEq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \<rbrace>"
- using type_e_eq infer_e_leqI fresh_Pair z3f by metis
-
- obtain z1' b1' c1' where z1:"atom z1' \<sharp> (x,v) \<and> \<lbrace> z1 : B_int | c1 \<rbrace> = \<lbrace> z1' : b1' | c1' \<rbrace>" using obtain_fresh_z by metis
- obtain z2' b2' c2' where z2:"atom z2' \<sharp> (x,v) \<and> \<lbrace> z2 : B_int | c2 \<rbrace> = \<lbrace> z2' : b2' | c2' \<rbrace>" using obtain_fresh_z by metis
-
- have bb:"b1' = B_int \<and> b2' = B_int" using z1 z2 \<tau>.eq_iff by metis
-
- have "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> ; \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> (AE_op LEq (v1[x::=v]\<^sub>v\<^sub>v) (v2[x::=v]\<^sub>v\<^sub>v)) \<Rightarrow> \<lbrace> z3' : B_bool | CE_val (V_var z3') == CE_op LEq ([v1[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) ([v2[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) \<rbrace>"
- proof
- show \<open> \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<close> using wfD_subst infer_e_leqI subtype_eq_base2 b_of.simps by metis
- show \<open> \<Theta>\<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using infer_e_leqI(2) by auto
- show \<open> \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> v1[x::=v]\<^sub>v\<^sub>v \<Rightarrow> \<lbrace> z1' : B_int | c1'[x::=v]\<^sub>c\<^sub>v \<rbrace>\<close> using subst_tv.simps subst_infer_v infer_e_leqI z1 bb by metis
- show \<open> \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> v2[x::=v]\<^sub>v\<^sub>v \<Rightarrow> \<lbrace> z2' : B_int | c2'[x::=v]\<^sub>c\<^sub>v \<rbrace>\<close> using subst_tv.simps subst_infer_v infer_e_leqI z2 bb by metis
- show \<open>atom z3' \<sharp> AE_op LEq v1[x::=v]\<^sub>v\<^sub>v v2[x::=v]\<^sub>v\<^sub>v\<close> using fresh_Pair * by metis
- show \<open>atom z3' \<sharp> \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>\<close> using * by auto
- qed
- moreover have "\<lbrace> z3' : B_bool | CE_val (V_var z3') == CE_op LEq ([v1[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) ([v2[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) \<rbrace> = \<lbrace> z3' : B_bool | CE_val (V_var z3') == CE_op LEq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v"
- using subst_tv.simps subst_ev.simps * by auto
- ultimately show ?case using subst_ev.simps * ** by metis
-next
- case (infer_e_eqI \<Theta> \<B> \<Gamma>'' \<Delta> \<Phi> v1 z1 bb c1 v2 z2 c2 z3)
-
- hence z3f: "atom z3 \<sharp> CE_op Eq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e" using e.fresh ce.fresh opp.fresh by metis
-
- obtain z3'::x where *:"atom z3' \<sharp> (x,v,AE_op Eq v1 v2, CE_op Eq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e , CE_op Eq [v1[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e [v2[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e , AE_op Eq v1[x::=v]\<^sub>v\<^sub>v v2[x::=v]\<^sub>v\<^sub>v,\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> )"
- using obtain_fresh by metis
- hence **:"(\<lbrace> z3 : B_bool | CE_val (V_var z3) == CE_op Eq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \<rbrace>) = \<lbrace> z3' : B_bool | CE_val (V_var z3') == CE_op Eq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \<rbrace>"
- using type_e_eq infer_e_eqI fresh_Pair z3f by metis
-
- obtain z1' b1' c1' where z1:"atom z1' \<sharp> (x,v) \<and> \<lbrace> z1 : bb | c1 \<rbrace> = \<lbrace> z1' : b1' | c1' \<rbrace>" using obtain_fresh_z by metis
- obtain z2' b2' c2' where z2:"atom z2' \<sharp> (x,v) \<and> \<lbrace> z2 : bb | c2 \<rbrace> = \<lbrace> z2' : b2' | c2' \<rbrace>" using obtain_fresh_z by metis
-
- have bb:"b1' = bb \<and> b2' = bb" using z1 z2 \<tau>.eq_iff by metis
-
- have "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> ; \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> (AE_op Eq (v1[x::=v]\<^sub>v\<^sub>v) (v2[x::=v]\<^sub>v\<^sub>v)) \<Rightarrow> \<lbrace> z3' : B_bool | CE_val (V_var z3') == CE_op Eq ([v1[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) ([v2[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) \<rbrace>"
- proof
- show \<open> \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<close> using wfD_subst infer_e_eqI subtype_eq_base2 b_of.simps by metis
- show \<open> \<Theta>\<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using infer_e_eqI(2) by auto
- show \<open> \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> v1[x::=v]\<^sub>v\<^sub>v \<Rightarrow> \<lbrace> z1' : bb | c1'[x::=v]\<^sub>c\<^sub>v \<rbrace>\<close> using subst_tv.simps subst_infer_v infer_e_eqI z1 bb by metis
- show \<open> \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> v2[x::=v]\<^sub>v\<^sub>v \<Rightarrow> \<lbrace> z2' : bb | c2'[x::=v]\<^sub>c\<^sub>v \<rbrace>\<close> using subst_tv.simps subst_infer_v infer_e_eqI z2 bb by metis
- show \<open>atom z3' \<sharp> AE_op Eq v1[x::=v]\<^sub>v\<^sub>v v2[x::=v]\<^sub>v\<^sub>v\<close> using fresh_Pair * by metis
- show \<open>atom z3' \<sharp> \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>\<close> using * by auto
- show "bb \<in> {B_bool, B_int, B_unit}" using infer_e_eqI by auto
- qed
- moreover have "\<lbrace> z3' : B_bool | CE_val (V_var z3') == CE_op Eq ([v1[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) ([v2[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) \<rbrace> = \<lbrace> z3' : B_bool | CE_val (V_var z3') == CE_op Eq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v"
- using subst_tv.simps subst_ev.simps * by auto
- ultimately show ?case using subst_ev.simps * ** by metis
-next
- case (infer_e_appI \<Theta> \<B> \<Gamma>'' \<Delta> \<Phi> f x' b c \<tau>' s' v' \<tau>)
-
- hence "x \<noteq> x'" using \<open>atom x' \<sharp> \<Gamma>''\<close> using wfG_inside_x_neq wfX_wfY by metis
-
- show ?case proof(subst subst_ev.simps,rule)
- show \<open> \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<close> using infer_e_appI wfD_subst subtype_eq_base2 b_of.simps by metis
- show \<open> \<Theta>\<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using infer_e_appI by metis
- show \<open>Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x' b c \<tau>' s'))) = lookup_fun \<Phi> f\<close> using infer_e_appI by metis
-
- have \<open>\<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> v'[x::=v]\<^sub>v\<^sub>v \<Leftarrow> \<lbrace> x' : b | c \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v\<close> proof(rule subst_infer_check_v )
- show "\<Theta> ; \<B> ; \<Gamma> \<turnstile> v \<Rightarrow> \<tau>\<^sub>1" using infer_e_appI by metis
- show "\<Theta> ; \<B> ; \<Gamma>' @ (x, b\<^sub>1, c0[z0::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<turnstile> v' \<Leftarrow> \<lbrace> x' : b | c \<rbrace>" using infer_e_appI by metis
- show "\<Theta> ; \<B> ; \<Gamma> \<turnstile> \<tau>\<^sub>1 \<lesssim> \<lbrace> z0 : b\<^sub>1 | c0 \<rbrace>" using infer_e_appI by metis
- show "atom z0 \<sharp> (x, v)" using infer_e_appI by metis
- qed
- moreover have "atom x \<sharp> c" using wfPhi_f_simple_supp_c infer_e_appI fresh_def \<open>x\<noteq>x'\<close>
- atom_eq_iff empty_iff infer_e_appI.hyps insert_iff subset_singletonD by metis
-
- moreover hence "atom x \<sharp> \<lbrace> x' : b | c \<rbrace>" using \<tau>.fresh supp_b_empty fresh_def by blast
- ultimately show \<open>\<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> v'[x::=v]\<^sub>v\<^sub>v \<Leftarrow> \<lbrace> x' : b | c \<rbrace>\<close> using forget_subst_tv by metis
-
- have *: "atom x' \<sharp> (x,v)" using infer_v_fresh_g_fresh_xv infer_e_appI check_v_wf by blast
-
- show \<open>atom x' \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>, \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v, v'[x::=v]\<^sub>v\<^sub>v, \<tau>[x::=v]\<^sub>\<tau>\<^sub>v)\<close>
- apply(unfold fresh_prodN, intro conjI)
- apply (fresh_subst_mth_aux add: infer_e_appI fresh_subst_gv wfD_wf subst_g_inside)
- using infer_e_appI fresh_subst_gv wfD_wf subst_g_inside apply metis
- using infer_e_appI fresh_subst_dv_if apply metis
- done
-
- have "supp \<tau>' \<subseteq> { atom x' } \<union> supp \<B>" using infer_e_appI wfT_supp wfPhi_f_simple_wfT
- by (meson infer_e_appI.hyps(2) le_supI1 wfPhi_f_simple_supp_t)
- hence "atom x \<sharp> \<tau>'" using \<open>x\<noteq>x'\<close> fresh_def supp_at_base x_not_in_b_set by fastforce
- thus \<open>\<tau>'[x'::=v'[x::=v]\<^sub>v\<^sub>v]\<^sub>v = \<tau>[x::=v]\<^sub>\<tau>\<^sub>v\<close> using subst_tv_commute infer_e_appI subst_defs by metis
- qed
-next
- case (infer_e_appPI \<Theta> \<B> \<Gamma>'' \<Delta> \<Phi> b' f bv x' b c \<tau>' s' v' \<tau>)
-
- hence "x \<noteq> x'" using \<open>atom x' \<sharp> \<Gamma>''\<close> using wfG_inside_x_neq wfX_wfY by metis
-
- show ?case proof(subst subst_ev.simps,rule)
- show \<open> \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<close> using infer_e_appPI wfD_subst subtype_eq_base2 b_of.simps by metis
- show \<open> \<Theta>\<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using infer_e_appPI(4) by auto
- show "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f b'" using infer_e_appPI(5) by auto
- show "Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x' b c \<tau>' s'))) = lookup_fun \<Phi> f" using infer_e_appPI(6) by auto
-
- show "\<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> v'[x::=v]\<^sub>v\<^sub>v \<Leftarrow> \<lbrace> x' : b[bv::=b']\<^sub>b | c[bv::=b']\<^sub>b \<rbrace>" proof -
- have \<open>\<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> v'[x::=v]\<^sub>v\<^sub>v \<Leftarrow> \<lbrace> x' : b[bv::=b']\<^sub>b\<^sub>b | c[bv::=b']\<^sub>c\<^sub>b \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v\<close> proof(rule subst_infer_check_v )
- show "\<Theta> ; \<B> ; \<Gamma> \<turnstile> v \<Rightarrow> \<tau>\<^sub>1" using infer_e_appPI by metis
- show "\<Theta> ; \<B> ; \<Gamma>' @ (x, b\<^sub>1, c0[z0::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<turnstile> v' \<Leftarrow> \<lbrace> x' : b[bv::=b']\<^sub>b\<^sub>b | c[bv::=b']\<^sub>c\<^sub>b \<rbrace>" using infer_e_appPI subst_defs by metis
- show "\<Theta> ; \<B> ; \<Gamma> \<turnstile> \<tau>\<^sub>1 \<lesssim> \<lbrace> z0 : b\<^sub>1 | c0 \<rbrace>" using infer_e_appPI by metis
- show "atom z0 \<sharp> (x, v)" using infer_e_appPI by metis
- qed
- moreover have "atom x \<sharp> c" proof -
- have "supp c \<subseteq> {atom x', atom bv}" using wfPhi_f_poly_supp_c[OF infer_e_appPI(6)] infer_e_appPI by metis
- thus ?thesis unfolding fresh_def using \<open>x\<noteq>x'\<close> atom_eq_iff by auto
- qed
- moreover hence "atom x \<sharp> \<lbrace> x' : b[bv::=b']\<^sub>b\<^sub>b | c[bv::=b']\<^sub>c\<^sub>b \<rbrace>" using \<tau>.fresh supp_b_empty fresh_def subst_b_fresh_x
- by (metis subst_b_c_def)
- ultimately show ?thesis using forget_subst_tv subst_defs by metis
- qed
- have "supp \<tau>' \<subseteq> { atom x', atom bv }" using wfPhi_f_poly_supp_t infer_e_appPI by metis
- hence "atom x \<sharp> \<tau>'" using fresh_def \<open>x\<noteq>x'\<close> by force
- hence *:"atom x \<sharp> \<tau>'[bv::=b']\<^sub>\<tau>\<^sub>b" using subst_b_fresh_x subst_b_\<tau>_def by metis
- have "atom x' \<sharp> (x,v)" using infer_v_fresh_g_fresh_xv infer_e_appPI check_v_wf by blast
- thus "atom x' \<sharp> \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>" using infer_e_appPI fresh_subst_gv wfD_wf subst_g_inside fresh_Pair by metis
- show "\<tau>'[bv::=b']\<^sub>b[x'::=v'[x::=v]\<^sub>v\<^sub>v]\<^sub>v = \<tau>[x::=v]\<^sub>\<tau>\<^sub>v" using infer_e_appPI subst_tv_commute[OF * ] subst_defs by metis
- show "atom bv \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>, \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v, b', v'[x::=v]\<^sub>v\<^sub>v, \<tau>[x::=v]\<^sub>\<tau>\<^sub>v)"
- by (fresh_mth add: infer_e_appPI fresh_subst_gv_inside)
- qed
-next
- case (infer_e_fstI \<Theta> \<B> \<Gamma>'' \<Delta> \<Phi> v' z' b1 b2 c z)
-
- hence zf: "atom z \<sharp> CE_fst [v']\<^sup>c\<^sup>e" using ce.fresh e.fresh opp.fresh by metis
-
- obtain z3'::x where *:"atom z3' \<sharp> (x,v,AE_fst v', CE_fst [v']\<^sup>c\<^sup>e , AE_fst v'[x::=v]\<^sub>v\<^sub>v ,\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> )" using obtain_fresh by auto
- hence **:"(\<lbrace> z : b1 | CE_val (V_var z) == CE_fst [v']\<^sup>c\<^sup>e \<rbrace>) = \<lbrace> z3' : b1 | CE_val (V_var z3') == CE_fst [v']\<^sup>c\<^sup>e \<rbrace>"
- using type_e_eq infer_e_fstI(4) fresh_Pair zf by metis
-
- obtain z1' b1' c1' where z1:"atom z1' \<sharp> (x,v) \<and> \<lbrace> z' : B_pair b1 b2 | c \<rbrace> = \<lbrace> z1' : b1' | c1' \<rbrace>" using obtain_fresh_z by metis
-
- have bb:"b1' = B_pair b1 b2" using z1 \<tau>.eq_iff by metis
- have "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> ; \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> (AE_fst v'[x::=v]\<^sub>v\<^sub>v) \<Rightarrow> \<lbrace> z3' : b1 | CE_val (V_var z3') == CE_fst [v'[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e \<rbrace>"
- proof
- show \<open> \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<close> using wfD_subst infer_e_fstI subtype_eq_base2 b_of.simps by metis
- show \<open> \<Theta>\<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using infer_e_fstI by metis
- show \<open> \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> v'[x::=v]\<^sub>v\<^sub>v \<Rightarrow> \<lbrace> z1' : B_pair b1 b2 | c1'[x::=v]\<^sub>c\<^sub>v \<rbrace>\<close> using subst_tv.simps subst_infer_v infer_e_fstI z1 bb by metis
-
- show \<open>atom z3' \<sharp> AE_fst v'[x::=v]\<^sub>v\<^sub>v \<close> using fresh_Pair * by metis
- show \<open>atom z3' \<sharp> \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>\<close> using * by auto
- qed
- moreover have "\<lbrace> z3' : b1 | CE_val (V_var z3') == CE_fst [v'[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e \<rbrace> = \<lbrace> z3' : b1 | CE_val (V_var z3') == CE_fst [v']\<^sup>c\<^sup>e \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v"
- using subst_tv.simps subst_ev.simps * by auto
- ultimately show ?case using subst_ev.simps * ** by metis
-next
- case (infer_e_sndI \<Theta> \<B> \<Gamma>'' \<Delta> \<Phi> v' z' b1 b2 c z)
- hence zf: "atom z \<sharp> CE_snd [v']\<^sup>c\<^sup>e" using ce.fresh e.fresh opp.fresh by metis
-
- obtain z3'::x where *:"atom z3' \<sharp> (x,v,AE_snd v', CE_snd [v']\<^sup>c\<^sup>e , AE_snd v'[x::=v]\<^sub>v\<^sub>v ,\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> ,v', \<Gamma>'')" using obtain_fresh by auto
- hence **:"(\<lbrace> z : b2 | CE_val (V_var z) == CE_snd [v']\<^sup>c\<^sup>e \<rbrace>) = \<lbrace> z3' : b2 | CE_val (V_var z3') == CE_snd [v']\<^sup>c\<^sup>e \<rbrace>"
- using type_e_eq infer_e_sndI(4) fresh_Pair zf by metis
-
- obtain z1' b2' c1' where z1:"atom z1' \<sharp> (x,v) \<and> \<lbrace> z' : B_pair b1 b2 | c \<rbrace> = \<lbrace> z1' : b2' | c1' \<rbrace>" using obtain_fresh_z by metis
-
- have bb:"b2' = B_pair b1 b2" using z1 \<tau>.eq_iff by metis
-
- have "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> ; \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> (AE_snd (v'[x::=v]\<^sub>v\<^sub>v)) \<Rightarrow> \<lbrace> z3' : b2 | CE_val (V_var z3') == CE_snd ([v'[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) \<rbrace>"
- proof
- show \<open> \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<close> using wfD_subst infer_e_sndI subtype_eq_base2 b_of.simps by metis
- show \<open> \<Theta>\<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using infer_e_sndI by metis
- show \<open> \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> v'[x::=v]\<^sub>v\<^sub>v \<Rightarrow> \<lbrace> z1' : B_pair b1 b2 | c1'[x::=v]\<^sub>c\<^sub>v \<rbrace>\<close> using subst_tv.simps subst_infer_v infer_e_sndI z1 bb by metis
-
- show \<open>atom z3' \<sharp> AE_snd v'[x::=v]\<^sub>v\<^sub>v \<close> using fresh_Pair * by metis
- show \<open>atom z3' \<sharp> \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>\<close> using * by auto
- qed
- moreover have "\<lbrace> z3' : b2 | CE_val (V_var z3') == CE_snd ([v'[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) \<rbrace> = \<lbrace> z3' : b2 | CE_val (V_var z3') == CE_snd [v']\<^sup>c\<^sup>e \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v"
- by(subst subst_tv.simps, auto simp add: fresh_prodN *)
- ultimately show ?case using subst_ev.simps * ** by metis
-next
- case (infer_e_lenI \<Theta> \<B> \<Gamma>'' \<Delta> \<Phi> v' z' c z)
- hence zf: "atom z \<sharp> CE_len [v']\<^sup>c\<^sup>e" using ce.fresh e.fresh opp.fresh by metis
- obtain z3'::x where *:"atom z3' \<sharp> (x,v,AE_len v', CE_len [v']\<^sup>c\<^sup>e , AE_len v'[x::=v]\<^sub>v\<^sub>v ,\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> , \<Gamma>'',v')" using obtain_fresh by auto
- hence **:"(\<lbrace> z : B_int | CE_val (V_var z) == CE_len [v']\<^sup>c\<^sup>e \<rbrace>) = \<lbrace> z3' : B_int | CE_val (V_var z3') == CE_len [v']\<^sup>c\<^sup>e \<rbrace>"
- using type_e_eq infer_e_lenI fresh_Pair zf by metis
-
- have ***: "\<Theta> ; \<B> ; \<Gamma>'' \<turnstile> v' \<Rightarrow> \<lbrace> z3' : B_bitvec | CE_val (V_var z3') == CE_val v' \<rbrace>"
- using infer_e_lenI infer_v_form3[OF infer_e_lenI(3), of z3'] b_of.simps * fresh_Pair by metis
-
- have "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> ; \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> (AE_len (v'[x::=v]\<^sub>v\<^sub>v)) \<Rightarrow> \<lbrace> z3' : B_int | CE_val (V_var z3') == CE_len ([v'[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) \<rbrace>"
- proof
- show \<open> \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<close> using wfD_subst infer_e_lenI subtype_eq_base2 b_of.simps by metis
- show \<open> \<Theta>\<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using infer_e_lenI by metis
-
- have \<open> \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> v'[x::=v]\<^sub>v\<^sub>v \<Rightarrow> \<lbrace> z3' : B_bitvec | CE_val (V_var z3') == CE_val v' \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v\<close>
- proof(rule subst_infer_v)
- show \<open> \<Theta> ; \<B> ; \<Gamma> \<turnstile> v \<Rightarrow> \<tau>\<^sub>1\<close> using infer_e_lenI by metis
- show \<open> \<Theta> ; \<B> ; \<Gamma>' @ (x, b\<^sub>1, c0[z0::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<turnstile> v' \<Rightarrow> \<lbrace> z3' : B_bitvec | [ [ z3' ]\<^sup>v ]\<^sup>c\<^sup>e == [ v' ]\<^sup>c\<^sup>e \<rbrace>\<close> using *** infer_e_lenI by metis
- show "\<Theta> ; \<B> ; \<Gamma> \<turnstile> \<tau>\<^sub>1 \<lesssim> \<lbrace> z0 : b\<^sub>1 | c0 \<rbrace>" using infer_e_lenI by metis
- show "atom z0 \<sharp> (x, v)" using infer_e_lenI by metis
- qed
-
- thus \<open> \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> v'[x::=v]\<^sub>v\<^sub>v \<Rightarrow> \<lbrace> z3' : B_bitvec | CE_val (V_var z3') == CE_val v'[x::=v]\<^sub>v\<^sub>v \<rbrace>\<close>
- using subst_tv.simps subst_ev.simps fresh_Pair * fresh_prodN subst_vv.simps by auto
-
- show \<open>atom z3' \<sharp> AE_len v'[x::=v]\<^sub>v\<^sub>v\<close> using fresh_Pair * by metis
- show \<open>atom z3' \<sharp> \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>\<close> using fresh_Pair * by metis
- qed
-
- moreover have "\<lbrace> z3' : B_int | CE_val (V_var z3') == CE_len ([v'[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) \<rbrace> = \<lbrace> z3' : B_int | CE_val (V_var z3') == CE_len [v']\<^sup>c\<^sup>e \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v"
- using subst_tv.simps subst_ev.simps * by auto
-
- ultimately show ?case using subst_ev.simps * ** by metis
-next
- case (infer_e_mvarI \<Theta> \<B> \<Gamma>'' \<Phi> \<Delta> u \<tau>)
-
- have "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> ; \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> (AE_mvar u) \<Rightarrow> \<tau>[x::=v]\<^sub>\<tau>\<^sub>v"
- proof
- show \<open> \<Theta> ; \<B>\<turnstile>\<^sub>w\<^sub>f \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>\<close> proof -
- have "wfV \<Theta> \<B> \<Gamma> v (b_of \<tau>\<^sub>1)" using infer_v_wf infer_e_mvarI by auto
- moreover have "b_of \<tau>\<^sub>1 = b\<^sub>1" using subtype_eq_base2 infer_e_mvarI b_of.simps by simp
- ultimately show ?thesis using wf_subst(3)[OF infer_e_mvarI(1), of \<Gamma>' x b\<^sub>1 "c0[z0::=[x]\<^sup>v]\<^sub>c\<^sub>v" \<Gamma> v] infer_e_mvarI subst_g_inside by metis
- qed
- show \<open> \<Theta>\<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using infer_e_mvarI by auto
- show \<open> \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<close> using wfD_subst infer_e_mvarI subtype_eq_base2 b_of.simps by metis
- show \<open>(u, \<tau>[x::=v]\<^sub>\<tau>\<^sub>v) \<in> setD \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v\<close> using infer_e_mvarI subst_dv_member by metis
- qed
- moreover have " (AE_mvar u) = (AE_mvar u)[x::=v]\<^sub>e\<^sub>v" using subst_ev.simps by auto
- ultimately show ?case by auto
-
-next
- case (infer_e_concatI \<Theta> \<B> \<Gamma>'' \<Delta> \<Phi> v1 z1 c1 v2 z2 c2 z3)
-
- hence zf: "atom z3 \<sharp> CE_concat [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e" using ce.fresh e.fresh opp.fresh by metis
-
- obtain z3'::x where *:"atom z3' \<sharp> (x,v,v1,v2,AE_concat v1 v2, CE_concat [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e , AE_concat (v1[x::=v]\<^sub>v\<^sub>v) (v2[x::=v]\<^sub>v\<^sub>v) ,\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> , \<Gamma>'',v1 , v2)" using obtain_fresh by auto
-
- hence **:"(\<lbrace> z3 : B_bitvec | CE_val (V_var z3) == CE_concat [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \<rbrace>) = \<lbrace> z3' : B_bitvec | CE_val (V_var z3') == CE_concat [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \<rbrace>"
- using type_e_eq infer_e_concatI fresh_Pair zf by metis
- have zfx: "atom x \<sharp> z3'" using fresh_at_base fresh_prodN * by auto
-
- have v1: "\<Theta> ; \<B> ; \<Gamma>'' \<turnstile> v1 \<Rightarrow> \<lbrace> z3' : B_bitvec | CE_val (V_var z3') == CE_val v1 \<rbrace>"
- using infer_e_concatI infer_v_form3 b_of.simps * fresh_Pair by metis
- have v2: "\<Theta> ; \<B> ; \<Gamma>'' \<turnstile> v2 \<Rightarrow> \<lbrace> z3' : B_bitvec | CE_val (V_var z3') == CE_val v2 \<rbrace>"
- using infer_e_concatI infer_v_form3 b_of.simps * fresh_Pair by metis
-
- have "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> ; \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> (AE_concat (v1[x::=v]\<^sub>v\<^sub>v) (v2[x::=v]\<^sub>v\<^sub>v)) \<Rightarrow> \<lbrace> z3' : B_bitvec | CE_val (V_var z3') == CE_concat ([v1[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) ([v2[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) \<rbrace>"
- proof
- show \<open> \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<close> using wfD_subst infer_e_concatI subtype_eq_base2 b_of.simps by metis
- show \<open> \<Theta>\<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> by(simp add: infer_e_concatI)
- show \<open> \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> v1[x::=v]\<^sub>v\<^sub>v \<Rightarrow> \<lbrace> z3' : B_bitvec | CE_val (V_var z3') == CE_val (v1[x::=v]\<^sub>v\<^sub>v) \<rbrace>\<close>
- using subst_infer_v_form infer_e_concatI fresh_prodN * b_of.simps by metis
- show \<open> \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> v2[x::=v]\<^sub>v\<^sub>v \<Rightarrow> \<lbrace> z3' : B_bitvec | CE_val (V_var z3') == CE_val (v2[x::=v]\<^sub>v\<^sub>v) \<rbrace>\<close>
- using subst_infer_v_form infer_e_concatI fresh_prodN * b_of.simps by metis
- show \<open>atom z3' \<sharp> AE_concat v1[x::=v]\<^sub>v\<^sub>v v2[x::=v]\<^sub>v\<^sub>v\<close> using fresh_Pair * by metis
- show \<open>atom z3' \<sharp> \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>\<close> using fresh_Pair * by metis
- qed
-
- moreover have "\<lbrace> z3' : B_bitvec | CE_val (V_var z3') == CE_concat ([v1[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) ([v2[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) \<rbrace> = \<lbrace> z3' : B_bitvec | CE_val (V_var z3') == CE_concat [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v"
- using subst_tv.simps subst_ev.simps * by auto
-
- ultimately show ?case using subst_ev.simps ** * by metis
-next
- case (infer_e_splitI \<Theta> \<B> \<Gamma>'' \<Delta> \<Phi> v1 z1 c1 v2 z2 z3)
- hence *:"atom z3 \<sharp> (x,v)" using fresh_Pair by auto
- have \<open>x \<noteq> z3 \<close> using infer_e_splitI by force
- have "\<Theta> ; \<Phi> ; \<B> ; (\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>) ; \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> (AE_split v1[x::=v]\<^sub>v\<^sub>v v2[x::=v]\<^sub>v\<^sub>v) \<Rightarrow>
- \<lbrace> z3 : [ B_bitvec , B_bitvec ]\<^sup>b | [ v1[x::=v]\<^sub>v\<^sub>v ]\<^sup>c\<^sup>e == [ [#1[ [ z3 ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e @@ [#2[ [ z3 ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e AND
- [| [#1[ [ z3 ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e == [ v2[x::=v]\<^sub>v\<^sub>v ]\<^sup>c\<^sup>e \<rbrace>"
- proof
- show \<open> \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<close> using wfD_subst infer_e_splitI subtype_eq_base2 b_of.simps by metis
- show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using infer_e_splitI by auto
- have \<open> \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> v1[x::=v]\<^sub>v\<^sub>v \<Rightarrow> \<lbrace> z1 : B_bitvec | c1 \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v\<close>
- using subst_infer_v infer_e_splitI by metis
- thus \<open> \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> v1[x::=v]\<^sub>v\<^sub>v \<Rightarrow> \<lbrace> z1 : B_bitvec | c1[x::=v]\<^sub>c\<^sub>v \<rbrace>\<close>
- using infer_e_splitI subst_tv.simps fresh_Pair by metis
- have \<open>x \<noteq> z2 \<close> using infer_e_splitI by force
- have "(\<lbrace> z2 : B_int | ([ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ z2 ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e)
- AND ([ leq [ [ z2 ]\<^sup>v ]\<^sup>c\<^sup>e [| [ v1[x::=v]\<^sub>v\<^sub>v ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e ) \<rbrace>) =
- (\<lbrace> z2 : B_int | ([ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ z2 ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e )
- AND ([ leq [ [ z2 ]\<^sup>v ]\<^sup>c\<^sup>e [| [ v1 ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e ) \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v)"
- unfolding subst_cv.simps subst_cev.simps subst_vv.simps using \<open>x \<noteq> z2\<close> infer_e_splitI fresh_Pair by simp
- thus \<open>\<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @
- \<Gamma> \<turnstile> v2[x::=v]\<^sub>v\<^sub>v \<Leftarrow> \<lbrace> z2 : B_int | [ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ z2 ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e
- AND [ leq [ [ z2 ]\<^sup>v ]\<^sup>c\<^sup>e [| [ v1[x::=v]\<^sub>v\<^sub>v ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrace>\<close>
- using infer_e_splitI subst_infer_check_v fresh_Pair by metis
-
- show \<open>atom z1 \<sharp> AE_split v1[x::=v]\<^sub>v\<^sub>v v2[x::=v]\<^sub>v\<^sub>v\<close> using infer_e_splitI fresh_subst_vv_if by auto
- show \<open>atom z2 \<sharp> AE_split v1[x::=v]\<^sub>v\<^sub>v v2[x::=v]\<^sub>v\<^sub>v\<close> using infer_e_splitI fresh_subst_vv_if by auto
- show \<open>atom z3 \<sharp> AE_split v1[x::=v]\<^sub>v\<^sub>v v2[x::=v]\<^sub>v\<^sub>v\<close> using infer_e_splitI fresh_subst_vv_if by auto
-
- show \<open>atom z3 \<sharp> \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>\<close> using fresh_subst_gv_inside infer_e_splitI by metis
- show \<open>atom z2 \<sharp> \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>\<close> using fresh_subst_gv_inside infer_e_splitI by metis
- show \<open>atom z1 \<sharp> \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>\<close> using fresh_subst_gv_inside infer_e_splitI by metis
- qed
- thus ?case apply (subst subst_tv.simps)
- using infer_e_splitI fresh_Pair apply metis
- unfolding subst_tv.simps subst_ev.simps subst_cv.simps subst_cev.simps subst_vv.simps *
- using \<open>x \<noteq> z3\<close> by simp
-qed
-
-lemma infer_e_uniqueness:
- assumes "\<Theta> ; \<Phi> ; \<B> ; GNil ; \<Delta> \<turnstile> e\<^sub>1 \<Rightarrow> \<tau>\<^sub>1" and "\<Theta> ; \<Phi> ; \<B> ; GNil ; \<Delta> \<turnstile> e\<^sub>1 \<Rightarrow> \<tau>\<^sub>2"
- shows "\<tau>\<^sub>1 = \<tau>\<^sub>2"
- using assms proof(nominal_induct rule:e.strong_induct)
- case (AE_val x)
- then show ?case using infer_e_elims(7)[OF AE_val(1)] infer_e_elims(7)[OF AE_val(2)] infer_v_uniqueness by metis
-next
- case (AE_app f v)
-
- obtain x1 b1 c1 s1' \<tau>1' where t1: "Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x1 b1 c1 \<tau>1' s1'))) = lookup_fun \<Phi> f \<and> \<tau>\<^sub>1 = \<tau>1'[x1::=v]\<^sub>\<tau>\<^sub>v" using infer_e_app2E[OF AE_app(1)] by metis
- moreover obtain x2 b2 c2 s2' \<tau>2' where t2: "Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x2 b2 c2 \<tau>2' s2'))) = lookup_fun \<Phi> f \<and> \<tau>\<^sub>2 = \<tau>2'[x2::=v]\<^sub>\<tau>\<^sub>v" using infer_e_app2E[OF AE_app(2)] by metis
-
- have "\<tau>1'[x1::=v]\<^sub>\<tau>\<^sub>v = \<tau>2'[x2::=v]\<^sub>\<tau>\<^sub>v" using t1 and t2 fun_ret_unique by metis
- thus ?thesis using t1 t2 by auto
-next
- case (AE_appP f b v)
- obtain bv1 x1 b1 c1 s1' \<tau>1' where t1: "Some (AF_fundef f (AF_fun_typ_some bv1 (AF_fun_typ x1 b1 c1 \<tau>1' s1'))) = lookup_fun \<Phi> f \<and> \<tau>\<^sub>1 = \<tau>1'[bv1::=b]\<^sub>\<tau>\<^sub>b[x1::=v]\<^sub>\<tau>\<^sub>v" using infer_e_appP2E[OF AE_appP(1)] by metis
- moreover obtain bv2 x2 b2 c2 s2' \<tau>2' where t2: "Some (AF_fundef f (AF_fun_typ_some bv2 (AF_fun_typ x2 b2 c2 \<tau>2' s2'))) = lookup_fun \<Phi> f \<and> \<tau>\<^sub>2 = \<tau>2'[bv2::=b]\<^sub>\<tau>\<^sub>b[x2::=v]\<^sub>\<tau>\<^sub>v" using infer_e_appP2E[OF AE_appP(2)] by metis
-
- have "\<tau>1'[bv1::=b]\<^sub>\<tau>\<^sub>b[x1::=v]\<^sub>\<tau>\<^sub>v = \<tau>2'[bv2::=b]\<^sub>\<tau>\<^sub>b[x2::=v]\<^sub>\<tau>\<^sub>v" using t1 and t2 fun_poly_ret_unique by metis
- thus ?thesis using t1 t2 by auto
-next
- case (AE_op opp v1 v2)
- show ?case proof(rule opp.exhaust)
- assume "opp = plus"
- hence "\<Theta> ; \<Phi> ; \<B> ; GNil ; \<Delta> \<turnstile> AE_op Plus v1 v2 \<Rightarrow> \<tau>\<^sub>1" and "\<Theta> ; \<Phi> ; \<B> ; GNil ; \<Delta> \<turnstile> AE_op Plus v1 v2 \<Rightarrow> \<tau>\<^sub>2" using AE_op by auto
- thus ?thesis using infer_e_elims(11)[OF \<open>\<Theta> ; \<Phi> ; \<B> ; GNil ; \<Delta> \<turnstile> AE_op Plus v1 v2 \<Rightarrow> \<tau>\<^sub>1\<close> ] infer_e_elims(11)[OF \<open>\<Theta> ; \<Phi> ; \<B> ; GNil ; \<Delta> \<turnstile> AE_op Plus v1 v2 \<Rightarrow> \<tau>\<^sub>2\<close> ]
- by force
- next
- assume "opp = leq"
- hence "opp = LEq" using opp.strong_exhaust by auto
- hence "\<Theta> ; \<Phi> ; \<B> ; GNil ; \<Delta> \<turnstile> AE_op LEq v1 v2 \<Rightarrow> \<tau>\<^sub>1" and "\<Theta> ; \<Phi> ; \<B> ; GNil ; \<Delta> \<turnstile> AE_op LEq v1 v2 \<Rightarrow> \<tau>\<^sub>2" using AE_op by auto
- thus ?thesis using infer_e_elims(12)[OF \<open>\<Theta> ; \<Phi> ; \<B> ; GNil ; \<Delta> \<turnstile> AE_op LEq v1 v2 \<Rightarrow> \<tau>\<^sub>1\<close> ] infer_e_elims(12)[OF \<open>\<Theta> ; \<Phi> ; \<B> ; GNil ; \<Delta> \<turnstile> AE_op LEq v1 v2 \<Rightarrow> \<tau>\<^sub>2\<close> ]
- by force
- next
- assume "opp = eq"
- hence "opp = Eq" using opp.strong_exhaust by auto
- hence "\<Theta> ; \<Phi> ; \<B> ; GNil ; \<Delta> \<turnstile> AE_op Eq v1 v2 \<Rightarrow> \<tau>\<^sub>1" and "\<Theta> ; \<Phi> ; \<B> ; GNil ; \<Delta> \<turnstile> AE_op Eq v1 v2 \<Rightarrow> \<tau>\<^sub>2" using AE_op by auto
- thus ?thesis using infer_e_elims(25)[OF \<open>\<Theta> ; \<Phi> ; \<B> ; GNil ; \<Delta> \<turnstile> AE_op Eq v1 v2 \<Rightarrow> \<tau>\<^sub>1\<close> ] infer_e_elims(25)[OF \<open>\<Theta> ; \<Phi> ; \<B> ; GNil ; \<Delta> \<turnstile> AE_op Eq v1 v2 \<Rightarrow> \<tau>\<^sub>2\<close> ]
- by force
- qed
-next
- case (AE_concat v1 v2)
-
- obtain z3::x where t1:"\<tau>\<^sub>1 = \<lbrace> z3 : B_bitvec | [ [ z3 ]\<^sup>v ]\<^sup>c\<^sup>e == CE_concat [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \<rbrace> \<and> atom z3 \<sharp> v1 \<and> atom z3 \<sharp> v2 " using infer_e_elims(18)[OF AE_concat(1)] by metis
- obtain z3'::x where t2:"\<tau>\<^sub>2 = \<lbrace> z3' : B_bitvec | [ [ z3' ]\<^sup>v ]\<^sup>c\<^sup>e == CE_concat [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \<rbrace> \<and> atom z3' \<sharp> v1 \<and> atom z3' \<sharp> v2" using infer_e_elims(18)[OF AE_concat(2)] by metis
-
- thus ?case using t1 t2 type_e_eq ce.fresh by metis
-
-next
- case (AE_fst v)
-
- obtain z1 and b1 where "\<tau>\<^sub>1 = \<lbrace> z1 : b1 | CE_val (V_var z1) == (CE_fst [v]\<^sup>c\<^sup>e) \<rbrace>" using infer_v_form AE_fst by auto
-
- obtain xx :: x and bb :: b and xxa :: x and bba :: b and cc :: c where
- f1: "\<tau>\<^sub>2 = \<lbrace> xx : bb | CE_val (V_var xx) == CE_fst [v]\<^sup>c\<^sup>e \<rbrace> \<and> \<Theta> ; \<B> ; GNil\<turnstile>\<^sub>w\<^sub>f \<Delta> \<and> \<Theta> ; \<B> ; GNil \<turnstile> v \<Rightarrow> \<lbrace> xxa : B_pair bb bba | cc \<rbrace> \<and> atom xx \<sharp> v"
- using infer_e_elims(8)[OF AE_fst(2)] by metis
- obtain xxb :: x and bbb :: b and xxc :: x and bbc :: b and cca :: c where
- f2: "\<tau>\<^sub>1 = \<lbrace> xxb : bbb | CE_val (V_var xxb) == CE_fst [v]\<^sup>c\<^sup>e \<rbrace> \<and> \<Theta> ; \<B> ; GNil\<turnstile>\<^sub>w\<^sub>f \<Delta> \<and> \<Theta> ; \<B> ; GNil \<turnstile> v \<Rightarrow> \<lbrace> xxc : B_pair bbb bbc | cca \<rbrace> \<and> atom xxb \<sharp> v"
- using infer_e_elims(8)[OF AE_fst(1)] by metis
- then have "B_pair bb bba = B_pair bbb bbc"
- using f1 by (metis (no_types) b_of.simps infer_v_uniqueness)
- then have "\<lbrace> xx : bbb | CE_val (V_var xx) == CE_fst [v]\<^sup>c\<^sup>e \<rbrace> = \<tau>\<^sub>2"
- using f1 by auto
- then show ?thesis
- using f2 by (meson ce.fresh fresh_GNil type_e_eq wfG_x_fresh_in_v_simple)
-next
- case (AE_snd v)
- obtain xx :: x and bb :: b and xxa :: x and bba :: b and cc :: c where
- f1: "\<tau>\<^sub>2 = \<lbrace> xx : bba | CE_val (V_var xx) == CE_snd [v]\<^sup>c\<^sup>e \<rbrace> \<and> \<Theta> ; \<B> ; GNil\<turnstile>\<^sub>w\<^sub>f \<Delta> \<and> \<Theta> ; \<B> ; GNil \<turnstile> v \<Rightarrow> \<lbrace> xxa : B_pair bb bba | cc \<rbrace> \<and> atom xx \<sharp> v"
- using infer_e_elims(9)[OF AE_snd(2)] by metis
- obtain xxb :: x and bbb :: b and xxc :: x and bbc :: b and cca :: c where
- f2: "\<tau>\<^sub>1 = \<lbrace> xxb : bbc | CE_val (V_var xxb) == CE_snd [v]\<^sup>c\<^sup>e \<rbrace> \<and> \<Theta> ; \<B> ; GNil\<turnstile>\<^sub>w\<^sub>f \<Delta> \<and> \<Theta> ; \<B> ; GNil \<turnstile> v \<Rightarrow> \<lbrace> xxc : B_pair bbb bbc | cca \<rbrace> \<and> atom xxb \<sharp> v"
- using infer_e_elims(9)[OF AE_snd(1)] by metis
- then have "B_pair bb bba = B_pair bbb bbc"
- using f1 by (metis (no_types) b_of.simps infer_v_uniqueness)
- then have "\<lbrace> xx : bbc | CE_val (V_var xx) == CE_snd [v]\<^sup>c\<^sup>e \<rbrace> = \<tau>\<^sub>2"
- using f1 by auto
- then show ?thesis
- using f2 by (meson ce.fresh fresh_GNil type_e_eq wfG_x_fresh_in_v_simple)
-next
- case (AE_mvar x)
- then show ?case using infer_e_elims(10)[OF AE_mvar(1)] infer_e_elims(10)[OF AE_mvar(2)] wfD_unique by metis
-next
- case (AE_len x)
- then show ?case using infer_e_elims(16)[OF AE_len(1)] infer_e_elims(16)[OF AE_len(2)] by force
-next
- case (AE_split x1a x2)
- then show ?case using infer_e_elims(22)[OF AE_split(1)] infer_e_elims(22)[OF AE_split(2)] by force
-qed
-
-section \<open>Statements\<close>
-
-lemma subst_infer_check_v1:
- fixes v::v and v'::v and \<Gamma>::\<Gamma>
- assumes "\<Gamma> = \<Gamma>\<^sub>1@((x,b\<^sub>1,c0[z0::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\<Gamma>\<Gamma>\<^sub>2)" and
- "\<Theta> ; \<B> ; \<Gamma>\<^sub>2 \<turnstile> v \<Rightarrow> \<tau>\<^sub>1" and
- "\<Theta> ; \<B> ; \<Gamma> \<turnstile> v' \<Leftarrow> \<tau>\<^sub>2" and
- "\<Theta> ; \<B> ; \<Gamma>\<^sub>2 \<turnstile> \<tau>\<^sub>1 \<lesssim> \<lbrace> z0 : b\<^sub>1 | c0 \<rbrace>" and "atom z0 \<sharp> (x,v)"
- shows "\<Theta> ; \<B> ; \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v \<turnstile> v'[x::=v]\<^sub>v\<^sub>v \<Leftarrow> \<tau>\<^sub>2[x::=v]\<^sub>\<tau>\<^sub>v"
- using subst_g_inside check_v_wf assms subst_infer_check_v by metis
-
-lemma infer_v_c_valid:
- assumes " \<Theta> ; \<B> ; \<Gamma> \<turnstile> v \<Rightarrow> \<tau>" and "\<Theta> ; \<B> ; \<Gamma> \<turnstile> \<tau> \<lesssim> \<lbrace> z : b | c \<rbrace>"
- shows \<open>\<Theta> ; \<B> ; \<Gamma> \<Turnstile> c[z::=v]\<^sub>c\<^sub>v \<close>
-proof -
- obtain z1 and b1 and c1 where *:"\<tau> = \<lbrace> z1 : b1 | c1 \<rbrace> \<and> atom z1 \<sharp> (c,v,\<Gamma>)" using obtain_fresh_z by metis
- then have "b1 = b" using assms subtype_eq_base by metis
- moreover then have "\<Theta> ; \<B> ; \<Gamma> \<turnstile> v \<Rightarrow> \<lbrace> z1 : b | c1 \<rbrace>" using assms * by auto
-
- moreover have "\<Theta> ; \<B> ; (z1, b, c1) #\<^sub>\<Gamma> \<Gamma> \<Turnstile> c[z::=[ z1 ]\<^sup>v]\<^sub>c\<^sub>v " proof -
- have "\<Theta> ; \<B> ; (z1, b, c1[z1::=[ z1 ]\<^sup>v]\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<Turnstile> c[z::=[ z1 ]\<^sup>v]\<^sub>v "
- using subtype_valid[OF assms(2), of z1 z1 b c1 z c ] * fresh_prodN \<open>b1 = b\<close> by metis
- moreover have "c1[z1::=[ z1 ]\<^sup>v]\<^sub>v = c1" using subst_v_v_def by simp
- ultimately show ?thesis using subst_v_c_def by metis
- qed
- ultimately show ?thesis using * fresh_prodN subst_valid_simple by metis
-qed
-
-text \<open> Substitution Lemma for Statements \<close>
-
-lemma subst_infer_check_s:
- fixes v::v and s::s and cs::branch_s and x::x and c::c and b::b and
- \<Gamma>\<^sub>1::\<Gamma> and \<Gamma>\<^sub>2::\<Gamma> and css::branch_list
- assumes "\<Theta> ; \<B> ; \<Gamma>\<^sub>1 \<turnstile> v \<Rightarrow> \<tau>" and "\<Theta> ; \<B> ; \<Gamma>\<^sub>1 \<turnstile> \<tau> \<lesssim> \<lbrace> z : b | c \<rbrace>" and
- "atom z \<sharp> (x, v)"
- shows "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>; \<Delta> \<turnstile> s \<Leftarrow> \<tau>' \<Longrightarrow>
- \<Gamma> = (\<Gamma>\<^sub>2@((x,b,c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\<Gamma>\<Gamma>\<^sub>1)) \<Longrightarrow>
- \<Theta> ; \<Phi> ; \<B> ; \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> s[x::=v]\<^sub>s\<^sub>v \<Leftarrow> \<tau>'[x::=v]\<^sub>\<tau>\<^sub>v"
- and
- "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>; \<Delta>; tid ; cons ; const ; v' \<turnstile> cs \<Leftarrow> \<tau>' \<Longrightarrow>
- \<Gamma> = (\<Gamma>\<^sub>2@((x,b,c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\<Gamma>\<Gamma>\<^sub>1)) \<Longrightarrow>
- \<Theta> ; \<Phi> ; \<B> ; \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v;
- tid ; cons ; const ; v'[x::=v]\<^sub>v\<^sub>v \<turnstile> cs[x::=v]\<^sub>s\<^sub>v \<Leftarrow> \<tau>'[x::=v]\<^sub>\<tau>\<^sub>v"
- and
- "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>; \<Delta>; tid ; dclist ; v' \<turnstile> css \<Leftarrow> \<tau>' \<Longrightarrow>
- \<Gamma> = (\<Gamma>\<^sub>2@((x,b,c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\<Gamma>\<Gamma>\<^sub>1)) \<Longrightarrow>
- \<Theta> ; \<Phi> ; \<B> ; \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v; tid ; dclist ; v'[x::=v]\<^sub>v\<^sub>v \<turnstile>
- subst_branchlv css x v \<Leftarrow> \<tau>'[x::=v]\<^sub>\<tau>\<^sub>v"
-
- using assms proof(nominal_induct \<tau>' and \<tau>' and \<tau>' avoiding: x v arbitrary: \<Gamma>\<^sub>2 and \<Gamma>\<^sub>2 and \<Gamma>\<^sub>2
- rule: check_s_check_branch_s_check_branch_list.strong_induct)
- case (check_valI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v' \<tau>' \<tau>'')
-
- have sg: " \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v = \<Gamma>\<^sub>2[x::=v]\<^sub>\<Gamma>\<^sub>v@\<Gamma>\<^sub>1" using check_valI by subst_mth
- have "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> (AS_val (v'[x::=v]\<^sub>v\<^sub>v)) \<Leftarrow> \<tau>''[x::=v]\<^sub>\<tau>\<^sub>v" proof
- have "\<Theta> ; \<B> ; \<Gamma>\<^sub>1\<turnstile>\<^sub>w\<^sub>f v : b " using infer_v_v_wf subtype_eq_base2 b_of.simps check_valI by metis
- thus \<open>\<Theta> ; \<B> ; \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v\<turnstile>\<^sub>w\<^sub>f \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v\<close> using wf_subst(15) check_valI by auto
- show \<open> \<Theta>\<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using check_valI by auto
- show \<open> \<Theta> ; \<B> ; \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v \<turnstile> v'[x::=v]\<^sub>v\<^sub>v \<Rightarrow> \<tau>'[x::=v]\<^sub>\<tau>\<^sub>v\<close> proof(subst sg, rule subst_infer_v)
- show "\<Theta> ; \<B> ; \<Gamma>\<^sub>1 \<turnstile> v \<Rightarrow> \<tau>" using check_valI by auto
- show "\<Theta> ; \<B> ; \<Gamma>\<^sub>2 @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>\<^sub>1 \<turnstile> v' \<Rightarrow> \<tau>'" using check_valI by metis
- show "\<Theta> ; \<B> ; \<Gamma>\<^sub>1 \<turnstile> \<tau> \<lesssim> \<lbrace> z: b | c \<rbrace>" using check_valI by auto
- show "atom z \<sharp> (x, v)" using check_valI by auto
- qed
- show \<open>\<Theta> ; \<B> ; \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v \<turnstile> \<tau>'[x::=v]\<^sub>\<tau>\<^sub>v \<lesssim> \<tau>''[x::=v]\<^sub>\<tau>\<^sub>v\<close> using subst_subtype_tau check_valI sg by metis
- qed
-
- thus ?case using Typing.check_valI subst_sv.simps sg by auto
-next
- case (check_letI xa \<Theta> \<Phi> \<B> \<Gamma> \<Delta> ea \<tau>a za sa ba ca)
- have *:"(AS_let xa ea sa)[x::=v]\<^sub>s\<^sub>v=(AS_let xa (ea[x::=v]\<^sub>e\<^sub>v) sa[x::=v]\<^sub>s\<^sub>v)"
- using subst_sv.simps \<open> atom xa \<sharp> x\<close> \<open> atom xa \<sharp> v\<close> by auto
- show ?case unfolding * proof
-
- show "atom xa \<sharp> (\<Theta>,\<Phi>,\<B>,\<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v,\<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v,ea[x::=v]\<^sub>e\<^sub>v,\<tau>a[x::=v]\<^sub>\<tau>\<^sub>v)"
- by(subst_tuple_mth add: check_letI)
-
- show "atom za \<sharp> (xa,\<Theta>,\<Phi>,\<B>,\<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v, \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v,ea[x::=v]\<^sub>e\<^sub>v,
- \<tau>a[x::=v]\<^sub>\<tau>\<^sub>v,sa[x::=v]\<^sub>s\<^sub>v)"
- by(subst_tuple_mth add: check_letI)
-
- show "\<Theta>; \<Phi>; \<B>; \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v; \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile>
- ea[x::=v]\<^sub>e\<^sub>v \<Rightarrow> \<lbrace> za : ba | ca[x::=v]\<^sub>c\<^sub>v \<rbrace>"
- proof -
- have "\<Theta>; \<Phi>; \<B>; \<Gamma>\<^sub>2[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>\<^sub>1; \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile>
- ea[x::=v]\<^sub>e\<^sub>v \<Rightarrow> \<lbrace> za : ba | ca \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v"
- using check_letI subst_infer_e by metis
- thus ?thesis using check_letI subst_tv.simps
- by (metis fresh_prod2I infer_e_wf subst_g_inside_simple)
- qed
-
- show "\<Theta>; \<Phi>; \<B>; (xa, ba, ca[x::=v]\<^sub>c\<^sub>v[za::=V_var xa]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v;
- \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> sa[x::=v]\<^sub>s\<^sub>v \<Leftarrow> \<tau>a[x::=v]\<^sub>\<tau>\<^sub>v"
- proof -
- have "\<Theta>; \<Phi>; \<B>; ((xa, ba, ca[za::=V_var xa]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)[x::=v]\<^sub>\<Gamma>\<^sub>v ;
- \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> sa[x::=v]\<^sub>s\<^sub>v \<Leftarrow> \<tau>a[x::=v]\<^sub>\<tau>\<^sub>v"
- apply(rule check_letI(23)[of "(xa, ba, ca[za::=V_var xa]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>\<^sub>2"])
- by(metis check_letI append_g.simps subst_defs)+
-
- moreover have "(xa, ba, ca[x::=v]\<^sub>c\<^sub>v[za::=V_var xa]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v =
- ((xa, ba, ca[za::=V_var xa]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)[x::=v]\<^sub>\<Gamma>\<^sub>v"
- using subst_cv_commute subst_gv.simps check_letI
- by (metis ms_fresh_all(39) ms_fresh_all(49) subst_cv_commute_full)
- ultimately show ?thesis
- using subst_defs by auto
- qed
- qed
-next
- case (check_assertI xa \<Theta> \<Phi> \<B> \<Gamma> \<Delta> ca \<tau> s)
- show ?case unfolding subst_sv.simps proof
- show \<open>atom xa \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v, \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v, ca[x::=v]\<^sub>c\<^sub>v, \<tau>[x::=v]\<^sub>\<tau>\<^sub>v, s[x::=v]\<^sub>s\<^sub>v)\<close>
- by(subst_tuple_mth add: check_assertI)
- have "xa \<noteq> x" using check_assertI by fastforce
- thus \<open> \<Theta> ; \<Phi> ; \<B> ; (xa, B_bool, ca[x::=v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> s[x::=v]\<^sub>s\<^sub>v \<Leftarrow> \<tau>[x::=v]\<^sub>\<tau>\<^sub>v\<close>
- using check_assertI(12)[of "(xa, B_bool, c) #\<^sub>\<Gamma> \<Gamma>\<^sub>2" x v] check_assertI subst_gv.simps append_g.simps by metis
- have \<open>\<Theta> ; \<B> ; \<Gamma>\<^sub>2[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>\<^sub>1 \<Turnstile> ca[x::=v]\<^sub>c\<^sub>v \<close> proof(rule subst_valid )
- show \<open>\<Theta> ; \<B> ; \<Gamma>\<^sub>1 \<Turnstile> c[z::=v]\<^sub>c\<^sub>v \<close> using infer_v_c_valid check_assertI by metis
- show \<open> \<Theta> ; \<B> ; \<Gamma>\<^sub>1 \<turnstile>\<^sub>w\<^sub>f v : b \<close> using check_assertI infer_v_wf b_of.simps subtype_eq_base
- by (metis subtype_eq_base2)
- show \<open> \<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>\<^sub>1 \<close> using check_assertI infer_v_wf by metis
- have " \<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>\<^sub>2 @ (x, b, c[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>\<^sub>1" using check_assertI wfX_wfY by metis
- thus \<open>atom x \<sharp> \<Gamma>\<^sub>1\<close> using check_assertI wfG_suffix wfG_elims by metis
-
- moreover have "\<Theta> ; \<B> ; \<Gamma>\<^sub>1 \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b | c \<rbrace>" using subtype_wfT check_assertI by metis
- moreover have "x \<noteq> z" using fresh_Pair check_assertI fresh_x_neq by metis
- ultimately show \<open>atom x \<sharp> c\<close> using check_assertI wfT_fresh_c by metis
-
- show \<open> \<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>\<^sub>2 @ (x, b, c[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>\<^sub>1 \<close> using check_assertI wfX_wfY by metis
- show \<open>\<Theta> ; \<B> ; \<Gamma>\<^sub>2 @ (x, b, c[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>\<^sub>1 \<Turnstile> ca \<close> using check_assertI by auto
- qed
- thus \<open>\<Theta> ; \<B> ; \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v \<Turnstile> ca[x::=v]\<^sub>c\<^sub>v \<close> using check_assertI
- proof -
- show ?thesis
- by (metis (no_types) \<open>\<Gamma> = \<Gamma>\<^sub>2 @ (x, b, c[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>\<^sub>1\<close> \<open>\<Theta> ; \<B> ; \<Gamma> \<Turnstile> ca\<close> \<open>\<Theta> ; \<B> ; \<Gamma>\<^sub>2[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>\<^sub>1 \<Turnstile> ca[x::=v]\<^sub>c\<^sub>v\<close> subst_g_inside valid_g_wf) (* 93 ms *)
- qed
-
- have "\<Theta> ; \<B> ; \<Gamma>\<^sub>1 \<turnstile>\<^sub>w\<^sub>f v : b" using infer_v_wf b_of.simps check_assertI
- by (metis subtype_eq_base2)
- thus \<open> \<Theta> ; \<B> ; \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v \<turnstile>\<^sub>w\<^sub>f \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<close> using wf_subst2(6) check_assertI by metis
- qed
-next
- case (check_branch_list_consI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist vv cs \<tau> css)
- show ?case unfolding * using subst_sv.simps check_branch_list_consI by simp
-next
- case (check_branch_list_finalI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist v cs \<tau>)
- show ?case unfolding * using subst_sv.simps check_branch_list_finalI by simp
-next
- case (check_branch_s_branchI \<Theta> \<B> \<Gamma> \<Delta> \<tau> const xa \<Phi> tid cons va sa)
- hence *:"(AS_branch cons xa sa)[x::=v]\<^sub>s\<^sub>v = (AS_branch cons xa sa[x::=v]\<^sub>s\<^sub>v)" using subst_branchv.simps fresh_Pair by metis
- show ?case unfolding * proof
-
- show "\<Theta> ; \<B> ; \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v\<turnstile>\<^sub>w\<^sub>f \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v "
- using wf_subst check_branch_s_branchI subtype_eq_base2 b_of.simps infer_v_wf by metis
-
- show "\<turnstile>\<^sub>w\<^sub>f \<Theta> " using check_branch_s_branchI by metis
-
- show "\<Theta> ; \<B> ; \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v \<turnstile>\<^sub>w\<^sub>f \<tau>[x::=v]\<^sub>\<tau>\<^sub>v "
- using wf_subst check_branch_s_branchI subtype_eq_base2 b_of.simps infer_v_wf by metis
-
- show wft:"\<Theta> ; {||} ; GNil\<turnstile>\<^sub>w\<^sub>f const " using check_branch_s_branchI by metis
-
- show "atom xa \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v, \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v, tid, cons, const,va[x::=v]\<^sub>v\<^sub>v, \<tau>[x::=v]\<^sub>\<tau>\<^sub>v)"
- apply(unfold fresh_prodN, (simp add: check_branch_s_branchI )+)
- apply(rule,metis fresh_z_subst_g check_branch_s_branchI fresh_Pair )
- by(metis fresh_subst_dv check_branch_s_branchI fresh_Pair )
-
- have "\<Theta> ; \<Phi> ; \<B> ; ((xa, b_of const, CE_val va == CE_val(V_cons tid cons (V_var xa)) AND c_of const xa) #\<^sub>\<Gamma> \<Gamma>)[x::=v]\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> sa[x::=v]\<^sub>s\<^sub>v \<Leftarrow> \<tau>[x::=v]\<^sub>\<tau>\<^sub>v"
- using check_branch_s_branchI by (metis append_g.simps(2))
-
- moreover have "(xa, b_of const, CE_val va[x::=v]\<^sub>v\<^sub>v == CE_val (V_cons tid cons (V_var xa)) AND c_of (const) xa) #\<^sub>\<Gamma> \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v =
- ((xa, b_of const , CE_val va == CE_val (V_cons tid cons (V_var xa)) AND c_of const xa) #\<^sub>\<Gamma> \<Gamma>)[x::=v]\<^sub>\<Gamma>\<^sub>v"
- proof -
- have *:"xa \<noteq> x" using check_branch_s_branchI fresh_at_base by metis
- have "atom x \<sharp> const" using wfT_nil_supp[OF wft] fresh_def by auto
- hence "atom x \<sharp> (const,xa)" using fresh_at_base wfT_nil_supp[OF wft] fresh_Pair fresh_def * by auto
- moreover hence "(c_of (const) xa)[x::=v]\<^sub>c\<^sub>v = c_of (const) xa"
- using c_of_fresh[of x const xa] forget_subst_cv wfT_nil_supp wft by metis
- moreover hence "(V_cons tid cons (V_var xa))[x::=v]\<^sub>v\<^sub>v = (V_cons tid cons (V_var xa))" using check_branch_s_branchI subst_vv.simps * by metis
- ultimately show ?thesis using subst_gv.simps check_branch_s_branchI subst_cv.simps subst_cev.simps * by presburger
- qed
-
- ultimately show "\<Theta> ; \<Phi> ; \<B> ; (xa, b_of const, CE_val va[x::=v]\<^sub>v\<^sub>v == CE_val (V_cons tid cons (V_var xa)) AND c_of const xa) #\<^sub>\<Gamma> \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> sa[x::=v]\<^sub>s\<^sub>v \<Leftarrow> \<tau>[x::=v]\<^sub>\<tau>\<^sub>v"
- by metis
- qed
-
-next
- case (check_let2I xa \<Theta> \<Phi> \<B> G \<Delta> t s1 \<tau>a s2 )
- hence *:"(AS_let2 xa t s1 s2)[x::=v]\<^sub>s\<^sub>v = (AS_let2 xa t[x::=v]\<^sub>\<tau>\<^sub>v (s1[x::=v]\<^sub>s\<^sub>v) s2[x::=v]\<^sub>s\<^sub>v)" using subst_sv.simps fresh_Pair by metis
- have "xa \<noteq> x" using check_let2I fresh_at_base by metis
- show ?case unfolding * proof
- show "atom xa \<sharp> (\<Theta>, \<Phi>, \<B>, G[x::=v]\<^sub>\<Gamma>\<^sub>v, \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v, t[x::=v]\<^sub>\<tau>\<^sub>v, s1[x::=v]\<^sub>s\<^sub>v, \<tau>a[x::=v]\<^sub>\<tau>\<^sub>v)"
- by(subst_tuple_mth add: check_let2I)
- show "\<Theta> ; \<Phi> ; \<B> ; G[x::=v]\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> s1[x::=v]\<^sub>s\<^sub>v \<Leftarrow> t[x::=v]\<^sub>\<tau>\<^sub>v" using check_let2I by metis
-
- have "\<Theta> ; \<Phi> ; \<B> ; ((xa, b_of t, c_of t xa) #\<^sub>\<Gamma> G)[x::=v]\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> s2[x::=v]\<^sub>s\<^sub>v \<Leftarrow> \<tau>a[x::=v]\<^sub>\<tau>\<^sub>v"
- proof(rule check_let2I(14))
- show \<open>(xa, b_of t, c_of t xa) #\<^sub>\<Gamma> G = (((xa, b_of t, c_of t xa)#\<^sub>\<Gamma> \<Gamma>\<^sub>2)) @ (x, b, c[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>\<^sub>1\<close>
- using check_let2I append_g.simps by metis
- show \<open> \<Theta> ; \<B> ; \<Gamma>\<^sub>1 \<turnstile> v \<Rightarrow> \<tau>\<close> using check_let2I by metis
- show \<open>\<Theta> ; \<B> ; \<Gamma>\<^sub>1 \<turnstile> \<tau> \<lesssim> \<lbrace> z : b | c \<rbrace>\<close> using check_let2I by metis
- show \<open>atom z \<sharp> (x, v)\<close> using check_let2I by metis
- qed
- moreover have "c_of t[x::=v]\<^sub>\<tau>\<^sub>v xa = (c_of t xa)[x::=v]\<^sub>c\<^sub>v" using subst_v_c_of fresh_Pair check_let2I by metis
- moreover have "b_of t[x::=v]\<^sub>\<tau>\<^sub>v = b_of t" using b_of.simps subst_tv.simps b_of_subst by metis
- ultimately show " \<Theta> ; \<Phi> ; \<B> ; (xa, b_of t[x::=v]\<^sub>\<tau>\<^sub>v, c_of t[x::=v]\<^sub>\<tau>\<^sub>v xa) #\<^sub>\<Gamma> G[x::=v]\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> s2[x::=v]\<^sub>s\<^sub>v \<Leftarrow> \<tau>a[x::=v]\<^sub>\<tau>\<^sub>v"
- using check_let2I(14) subst_gv.simps \<open>xa \<noteq> x\<close> b_of.simps by metis
- qed
-
-next
-
- case (check_varI u \<Theta> \<Phi> \<B> \<Gamma> \<Delta> \<tau>' va \<tau>'' s)
- have **: "\<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v = \<Gamma>\<^sub>2[x::=v]\<^sub>\<Gamma>\<^sub>v@\<Gamma>\<^sub>1" using subst_g_inside check_s_wf check_varI by meson
-
- have "\<Theta> ; \<Phi> ;\<B> ; subst_gv \<Gamma> x v ; \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> AS_var u \<tau>'[x::=v]\<^sub>\<tau>\<^sub>v (va[x::=v]\<^sub>v\<^sub>v) (subst_sv s x v) \<Leftarrow> \<tau>''[x::=v]\<^sub>\<tau>\<^sub>v"
- proof(rule Typing.check_varI)
- show "atom u \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v, \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v, \<tau>'[x::=v]\<^sub>\<tau>\<^sub>v, va[x::=v]\<^sub>v\<^sub>v, \<tau>''[x::=v]\<^sub>\<tau>\<^sub>v)"
- by(subst_tuple_mth add: check_varI)
- show "\<Theta> ; \<B> ; \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v \<turnstile> va[x::=v]\<^sub>v\<^sub>v \<Leftarrow> \<tau>'[x::=v]\<^sub>\<tau>\<^sub>v" using check_varI subst_infer_check_v ** by metis
- show "\<Theta> ; \<Phi> ; \<B> ; subst_gv \<Gamma> x v ; (u, \<tau>'[x::=v]\<^sub>\<tau>\<^sub>v) #\<^sub>\<Delta> \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> s[x::=v]\<^sub>s\<^sub>v \<Leftarrow> \<tau>''[x::=v]\<^sub>\<tau>\<^sub>v" proof -
- have "wfD \<Theta> \<B> (\<Gamma>\<^sub>2 @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>\<^sub>1) ((u,\<tau>')#\<^sub>\<Delta> \<Delta>)" using check_varI check_s_wf by meson
- moreover have "wfV \<Theta> \<B> \<Gamma>\<^sub>1 v (b_of \<tau>)" using infer_v_wf check_varI(6) check_varI by metis
- have "wfD \<Theta> \<B> (\<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v) ((u, \<tau>'[x::=v]\<^sub>\<tau>\<^sub>v) #\<^sub>\<Delta> \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v)" proof(subst subst_dv.simps(2)[symmetric], subst **, rule wfD_subst)
- show "\<Theta> ; \<B> ; \<Gamma>\<^sub>1 \<turnstile> v \<Rightarrow> \<tau>" using check_varI by auto
- show "\<Theta> ; \<B> ; \<Gamma>\<^sub>2 @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>\<^sub>1\<turnstile>\<^sub>w\<^sub>f (u, \<tau>') #\<^sub>\<Delta> \<Delta>" using check_varI check_s_wf by simp
- show "b_of \<tau> = b" using check_varI subtype_eq_base2 b_of.simps by auto
- qed
- thus ?thesis using check_varI by auto
- qed
- qed
- moreover have "atom u \<sharp> (x,v)" using u_fresh_xv by auto
- ultimately show ?case using subst_sv.simps(7) by auto
-
-next
- case (check_assignI P \<Phi> \<B> \<Gamma> \<Delta> u \<tau>1 v' z1 \<tau>') (* may need to revisit subst in \<Delta> as well *)
-
- have "wfG P \<B> \<Gamma>" using check_v_wf check_assignI by simp
- hence gs: "\<Gamma>\<^sub>2[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>\<^sub>1 = \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v" using subst_g_inside check_assignI by simp
-
- have "P ; \<Phi> ; \<B> ; \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> AS_assign u (v'[x::=v]\<^sub>v\<^sub>v) \<Leftarrow> \<tau>'[x::=v]\<^sub>\<tau>\<^sub>v"
- proof(rule Typing.check_assignI)
- show "P \<turnstile>\<^sub>w\<^sub>f \<Phi> " using check_assignI by auto
- show "wfD P \<B> (\<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v) \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v" using wf_subst(15)[OF check_assignI(2)] gs infer_v_v_wf check_assignI b_of.simps subtype_eq_base2 by metis
- thus "(u, \<tau>1[x::=v]\<^sub>\<tau>\<^sub>v) \<in> setD \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v" using check_assignI subst_dv_member by metis
- thus "P ; \<B> ; \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v \<turnstile> v'[x::=v]\<^sub>v\<^sub>v \<Leftarrow> \<tau>1[x::=v]\<^sub>\<tau>\<^sub>v" using subst_infer_check_v check_assignI gs by metis
-
- have "P ; \<B> ; \<Gamma>\<^sub>2[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>\<^sub>1 \<turnstile> \<lbrace> z : B_unit | TRUE \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v \<lesssim> \<tau>'[x::=v]\<^sub>\<tau>\<^sub>v" proof(rule subst_subtype_tau)
- show "P ; \<B> ; \<Gamma>\<^sub>1 \<turnstile> v \<Rightarrow> \<tau>" using check_assignI by auto
- show "P ; \<B> ; \<Gamma>\<^sub>1 \<turnstile> \<tau> \<lesssim> \<lbrace> z : b | c \<rbrace>" using check_assignI by meson
- show "P ; \<B> ; \<Gamma>\<^sub>2 @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>\<^sub>1 \<turnstile> \<lbrace> z : B_unit | TRUE \<rbrace> \<lesssim> \<tau>'" using check_assignI
- by (metis Abs1_eq_iff(3) \<tau>.eq_iff c.fresh(1) c.perm_simps(1))
- show "atom z \<sharp> (x, v)" using check_assignI by auto
- qed
- moreover have "\<lbrace> z : B_unit | TRUE \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v = \<lbrace> z : B_unit | TRUE \<rbrace>" using subst_tv.simps subst_cv.simps check_assignI by presburger
- ultimately show "P ; \<B> ; \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v \<turnstile> \<lbrace> z : B_unit | TRUE \<rbrace> \<lesssim> \<tau>'[x::=v]\<^sub>\<tau>\<^sub>v" using gs by auto
- qed
- thus ?case using subst_sv.simps(5) by auto
-
-next
- case (check_whileI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> s1 z' s2 \<tau>')
- have " wfG \<Theta> \<B> (\<Gamma>\<^sub>2 @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>\<^sub>1)" using check_whileI check_s_wf by meson
- hence **: " \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v = \<Gamma>\<^sub>2[x::=v]\<^sub>\<Gamma>\<^sub>v@\<Gamma>\<^sub>1" using subst_g_inside wf check_whileI by auto
- have teq: "(\<lbrace> z : B_unit | TRUE \<rbrace>)[x::=v]\<^sub>\<tau>\<^sub>v = (\<lbrace> z : B_unit | TRUE \<rbrace>)" by(auto simp add: subst_sv.simps check_whileI)
- moreover have "(\<lbrace> z : B_unit | TRUE \<rbrace>) = (\<lbrace> z' : B_unit | TRUE \<rbrace>)" using type_eq_flip c.fresh flip_fresh_fresh by metis
- ultimately have teq2:"(\<lbrace> z' : B_unit | TRUE \<rbrace>)[x::=v]\<^sub>\<tau>\<^sub>v = (\<lbrace> z' : B_unit | TRUE \<rbrace>)" by metis
-
- hence "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> s1[x::=v]\<^sub>s\<^sub>v \<Leftarrow> \<lbrace> z' : B_bool | TRUE \<rbrace>" using check_whileI subst_sv.simps subst_top_eq by metis
- moreover have "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> s2[x::=v]\<^sub>s\<^sub>v \<Leftarrow> \<lbrace> z' : B_unit | TRUE \<rbrace>" using check_whileI subst_top_eq by metis
- moreover have "\<Theta> ; \<B> ; \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v \<turnstile> \<lbrace> z' : B_unit | TRUE \<rbrace> \<lesssim> \<tau>'[x::=v]\<^sub>\<tau>\<^sub>v" proof -
- have "\<Theta> ; \<B> ; \<Gamma>\<^sub>2[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>\<^sub>1 \<turnstile> \<lbrace> z' : B_unit | TRUE \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v \<lesssim> \<tau>'[x::=v]\<^sub>\<tau>\<^sub>v" proof(rule subst_subtype_tau)
- show "\<Theta> ; \<B> ; \<Gamma>\<^sub>1 \<turnstile> v \<Rightarrow> \<tau>" by(auto simp add: check_whileI)
- show "\<Theta> ; \<B> ; \<Gamma>\<^sub>1 \<turnstile> \<tau> \<lesssim> \<lbrace> z : b | c \<rbrace>" by(auto simp add: check_whileI)
- show "\<Theta> ; \<B> ; \<Gamma>\<^sub>2 @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>\<^sub>1 \<turnstile> \<lbrace> z' : B_unit | TRUE \<rbrace> \<lesssim> \<tau>'" using check_whileI by metis
- show "atom z \<sharp> (x, v)" by(auto simp add: check_whileI)
- qed
- thus ?thesis using teq2 ** by auto
- qed
-
- ultimately have " \<Theta> ; \<Phi> ; \<B> ; \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> AS_while s1[x::=v]\<^sub>s\<^sub>v s2[x::=v]\<^sub>s\<^sub>v \<Leftarrow> \<tau>'[x::=v]\<^sub>\<tau>\<^sub>v"
- using Typing.check_whileI by metis
- then show ?case using subst_sv.simps by metis
-next
- case (check_seqI P \<Phi> \<B> \<Gamma> \<Delta> s1 z s2 \<tau> )
- hence "P ; \<Phi>; \<B> ; \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> AS_seq (s1[x::=v]\<^sub>s\<^sub>v) (s2[x::=v]\<^sub>s\<^sub>v) \<Leftarrow> \<tau>[x::=v]\<^sub>\<tau>\<^sub>v" using Typing.check_seqI subst_top_eq check_seqI by metis
- then show ?case using subst_sv.simps by metis
-next
- case (check_caseI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist v' cs \<tau> za)
-
- have wf: "wfG \<Theta> \<B> \<Gamma>" using check_caseI check_v_wf by simp
- have **: "\<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v = \<Gamma>\<^sub>2[x::=v]\<^sub>\<Gamma>\<^sub>v@\<Gamma>\<^sub>1" using subst_g_inside wf check_caseI by auto
-
- have "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> AS_match (v'[x::=v]\<^sub>v\<^sub>v) (subst_branchlv cs x v) \<Leftarrow> \<tau>[x::=v]\<^sub>\<tau>\<^sub>v" proof(rule Typing.check_caseI )
- show "check_branch_list \<Theta> \<Phi> \<B> (\<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v) \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v tid dclist v'[x::=v]\<^sub>v\<^sub>v (subst_branchlv cs x v ) (\<tau>[x::=v]\<^sub>\<tau>\<^sub>v)" using check_caseI by auto
- show "AF_typedef tid dclist \<in> set \<Theta>" using check_caseI by auto
- show "\<Theta> ; \<B> ; \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v \<turnstile> v'[x::=v]\<^sub>v\<^sub>v \<Leftarrow> \<lbrace> za : B_id tid | TRUE \<rbrace>" proof -
- have "\<Theta> ; \<B> ; \<Gamma>\<^sub>2 @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>\<^sub>1 \<turnstile> v' \<Leftarrow> \<lbrace> za : B_id tid | TRUE \<rbrace>"
- using check_caseI by argo
- hence "\<Theta> ; \<B> ; \<Gamma>\<^sub>2[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>\<^sub>1 \<turnstile> v'[x::=v]\<^sub>v\<^sub>v \<Leftarrow> (\<lbrace> za : B_id tid | TRUE \<rbrace>)[x::=v]\<^sub>\<tau>\<^sub>v"
- using check_caseI subst_infer_check_v[OF check_caseI(7) _ check_caseI(8) check_caseI(9)] by meson
- moreover have "(\<lbrace> za : B_id tid | TRUE \<rbrace>) = ((\<lbrace> za : B_id tid | TRUE \<rbrace>)[x::=v]\<^sub>\<tau>\<^sub>v)"
- using subst_cv.simps subst_tv.simps subst_cv_true by fast
- ultimately show ?thesis using check_caseI ** by argo
- qed
- show "wfTh \<Theta>" using check_caseI by auto
- qed
- thus ?case using subst_branchlv.simps subst_sv.simps(4) by metis
-next
- case (check_ifI z' \<Theta> \<Phi> \<B> \<Gamma> \<Delta> va s1 s2 \<tau>')
- show ?case unfolding subst_sv.simps proof
- show \<open>atom z' \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v, \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v, va[x::=v]\<^sub>v\<^sub>v, s1[x::=v]\<^sub>s\<^sub>v, s2[x::=v]\<^sub>s\<^sub>v, \<tau>'[x::=v]\<^sub>\<tau>\<^sub>v)\<close>
- by(subst_tuple_mth add: check_ifI)
- have *:"\<lbrace> z' : B_bool | TRUE \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v = \<lbrace> z' : B_bool | TRUE \<rbrace>" using subst_tv.simps check_ifI
- by (metis freshers(19) subst_cv.simps(1) type_eq_subst)
- have **: "\<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v = \<Gamma>\<^sub>2[x::=v]\<^sub>\<Gamma>\<^sub>v@\<Gamma>\<^sub>1" using subst_g_inside wf check_ifI check_v_wf by metis
- show \<open>\<Theta> ; \<B> ; \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v \<turnstile> va[x::=v]\<^sub>v\<^sub>v \<Leftarrow> \<lbrace> z' : B_bool | TRUE \<rbrace>\<close>
- proof(subst *[symmetric], rule subst_infer_check_v1[where \<Gamma>\<^sub>1=\<Gamma>\<^sub>2 and \<Gamma>\<^sub>2=\<Gamma>\<^sub>1])
- show "\<Gamma> = \<Gamma>\<^sub>2 @ ((x, b ,c[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>\<^sub>1)" using check_ifI by metis
- show \<open> \<Theta> ; \<B> ; \<Gamma>\<^sub>1 \<turnstile> v \<Rightarrow> \<tau>\<close> using check_ifI by metis
- show \<open>\<Theta> ; \<B> ; \<Gamma> \<turnstile> va \<Leftarrow> \<lbrace> z' : B_bool | TRUE \<rbrace>\<close> using check_ifI by metis
- show \<open>\<Theta> ; \<B> ; \<Gamma>\<^sub>1 \<turnstile> \<tau> \<lesssim> \<lbrace> z : b | c \<rbrace>\<close> using check_ifI by metis
- show \<open>atom z \<sharp> (x, v)\<close> using check_ifI by metis
- qed
-
- have " \<lbrace> z' : b_of \<tau>'[x::=v]\<^sub>\<tau>\<^sub>v | [ va[x::=v]\<^sub>v\<^sub>v ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of \<tau>'[x::=v]\<^sub>\<tau>\<^sub>v z' \<rbrace> = \<lbrace> z' : b_of \<tau>' | [ va ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of \<tau>' z' \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v"
- by(simp add: subst_tv.simps fresh_Pair check_ifI b_of_subst subst_v_c_of)
-
- thus \<open> \<Theta> ; \<Phi> ; \<B> ; \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> s1[x::=v]\<^sub>s\<^sub>v \<Leftarrow> \<lbrace> z' : b_of \<tau>'[x::=v]\<^sub>\<tau>\<^sub>v | [ va[x::=v]\<^sub>v\<^sub>v ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of \<tau>'[x::=v]\<^sub>\<tau>\<^sub>v z' \<rbrace>\<close>
- using check_ifI by metis
- have " \<lbrace> z' : b_of \<tau>'[x::=v]\<^sub>\<tau>\<^sub>v | [ va[x::=v]\<^sub>v\<^sub>v ]\<^sup>c\<^sup>e == [ [ L_false ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of \<tau>'[x::=v]\<^sub>\<tau>\<^sub>v z' \<rbrace> = \<lbrace> z' : b_of \<tau>' | [ va ]\<^sup>c\<^sup>e == [ [ L_false ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of \<tau>' z' \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v"
- by(simp add: subst_tv.simps fresh_Pair check_ifI b_of_subst subst_v_c_of)
- thus \<open> \<Theta> ; \<Phi> ; \<B> ; \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> s2[x::=v]\<^sub>s\<^sub>v \<Leftarrow> \<lbrace> z' : b_of \<tau>'[x::=v]\<^sub>\<tau>\<^sub>v | [ va[x::=v]\<^sub>v\<^sub>v ]\<^sup>c\<^sup>e == [ [ L_false ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of \<tau>'[x::=v]\<^sub>\<tau>\<^sub>v z' \<rbrace>\<close>
- using check_ifI by metis
- qed
-qed
-
-lemma subst_check_check_s:
- fixes v::v and s::s and cs::branch_s and x::x and c::c and b::b and \<Gamma>\<^sub>1::\<Gamma> and \<Gamma>\<^sub>2::\<Gamma>
- assumes "\<Theta> ; \<B> ; \<Gamma>\<^sub>1 \<turnstile> v \<Leftarrow> \<lbrace> z : b | c \<rbrace>" and "atom z \<sharp> (x, v)"
- and "check_s \<Theta> \<Phi> \<B> \<Gamma> \<Delta> s \<tau>'" and "\<Gamma> = (\<Gamma>\<^sub>2@((x,b,c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\<Gamma>\<Gamma>\<^sub>1))"
- shows "check_s \<Theta> \<Phi> \<B> (subst_gv \<Gamma> x v) \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v (s[x::=v]\<^sub>s\<^sub>v) (subst_tv \<tau>' x v )"
-proof -
- obtain \<tau> where "\<Theta> ; \<B> ; \<Gamma>\<^sub>1 \<turnstile> v \<Rightarrow> \<tau> \<and> \<Theta> ; \<B> ; \<Gamma>\<^sub>1 \<turnstile> \<tau> \<lesssim> \<lbrace> z : b | c \<rbrace>" using check_v_elims assms by auto
- thus ?thesis using subst_infer_check_s assms by metis
-qed
-
-text \<open> If a statement checks against a type @{text "\<tau>"} then it checks against a supertype of @{text "\<tau>"} \<close>
-lemma check_s_supertype:
- fixes v::v and s::s and cs::branch_s and x::x and c::c and b::b and \<Gamma>::\<Gamma> and \<Gamma>'::\<Gamma> and css::branch_list
- shows "check_s \<Theta> \<Phi> \<B> G \<Delta> s t1 \<Longrightarrow> \<Theta> ; \<B> ; G \<turnstile> t1 \<lesssim> t2 \<Longrightarrow> check_s \<Theta> \<Phi> \<B> G \<Delta> s t2" and
- "check_branch_s \<Theta> \<Phi> \<B> G \<Delta> tid cons const v cs t1 \<Longrightarrow> \<Theta> ; \<B> ; G \<turnstile> t1 \<lesssim> t2 \<Longrightarrow> check_branch_s \<Theta> \<Phi> \<B> G \<Delta> tid cons const v cs t2" and
- "check_branch_list \<Theta> \<Phi> \<B> G \<Delta> tid dclist v css t1 \<Longrightarrow> \<Theta> ; \<B> ; G \<turnstile> t1 \<lesssim> t2 \<Longrightarrow> check_branch_list \<Theta> \<Phi> \<B> G \<Delta> tid dclist v css t2"
-proof(induct arbitrary: t2 and t2 and t2 rule: check_s_check_branch_s_check_branch_list.inducts)
- case (check_valI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v \<tau>' \<tau> )
- hence " \<Theta> ; \<B> ; \<Gamma> \<turnstile> \<tau>' \<lesssim> t2" using subtype_trans by meson
- then show ?case using subtype_trans Typing.check_valI check_valI by metis
-
-next
- case (check_letI x \<Theta> \<Phi> \<B> \<Gamma> \<Delta> e \<tau> z s b c)
- show ?case proof(rule Typing.check_letI)
- show "atom x \<sharp>(\<Theta>, \<Phi>, \<B>, \<Gamma>, \<Delta>, e, t2)" using check_letI subtype_fresh_tau fresh_prodN by metis
- show "atom z \<sharp> (x, \<Theta>, \<Phi>, \<B>, \<Gamma>, \<Delta>, e, t2, s)" using check_letI(2) subtype_fresh_tau[of z \<tau> \<Gamma> _ _ t2] fresh_prodN check_letI(6) by auto
- show "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile> e \<Rightarrow> \<lbrace> z : b | c \<rbrace>" using check_letI by meson
-
- have "wfG \<Theta> \<B> ((x, b, c[z::=[x]\<^sup>v]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)" using check_letI check_s_wf subst_defs by metis
- moreover have "toSet \<Gamma> \<subseteq> toSet ((x, b, c[z::=[x]\<^sup>v]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)" by auto
- ultimately have " \<Theta> ; \<B> ; (x, b, c[z::=[x]\<^sup>v]\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<turnstile> \<tau> \<lesssim> t2" using subtype_weakening[OF check_letI(6)] by auto
- thus "\<Theta> ; \<Phi> ; \<B> ; (x, b, c[z::=[x]\<^sup>v]\<^sub>v) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile> s \<Leftarrow> t2" using check_letI subst_defs by metis
- qed
-next
- case (check_branch_list_consI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist v cs \<tau> css)
- then show ?case using Typing.check_branch_list_consI by auto
-next
- case (check_branch_list_finalI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist v cs \<tau>)
- then show ?case using Typing.check_branch_list_finalI by auto
-next
- case (check_branch_s_branchI \<Theta> \<B> \<Gamma> \<Delta> \<tau> const x \<Phi> tid cons v s)
- show ?case proof
- have "atom x \<sharp> t2" using subtype_fresh_tau[of x \<tau> ] check_branch_s_branchI(5,8) fresh_prodN by metis
- thus "atom x \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>, \<Delta>, tid, cons, const, v, t2)" using check_branch_s_branchI fresh_prodN by metis
- show "wfT \<Theta> \<B> \<Gamma> t2" using subtype_wf check_branch_s_branchI by meson
- show "\<Theta> ; \<Phi> ; \<B> ; (x, b_of const, CE_val v == CE_val(V_cons tid cons (V_var x)) AND c_of const x) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile> s \<Leftarrow> t2" proof -
- have "wfG \<Theta> \<B> ((x, b_of const, CE_val v == CE_val(V_cons tid cons (V_var x)) AND c_of const x) #\<^sub>\<Gamma> \<Gamma>)" using check_s_wf check_branch_s_branchI by metis
- moreover have "toSet \<Gamma> \<subseteq> toSet ((x, b_of const, CE_val v == CE_val (V_cons tid cons (V_var x)) AND c_of const x) #\<^sub>\<Gamma> \<Gamma>)" by auto
- hence "\<Theta> ; \<B> ; ((x, b_of const, CE_val v == CE_val(V_cons tid cons (V_var x)) AND c_of const x) #\<^sub>\<Gamma> \<Gamma>) \<turnstile> \<tau> \<lesssim> t2"
- using check_branch_s_branchI subtype_weakening
- using calculation by presburger
- thus ?thesis using check_branch_s_branchI by presburger
- qed
- qed(auto simp add: check_branch_s_branchI)
-next
- case (check_ifI z \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v s1 s2 \<tau>)
- show ?case proof(rule Typing.check_ifI)
- have *:"atom z \<sharp> t2" using subtype_fresh_tau[of z \<tau> \<Gamma> ] check_ifI fresh_prodN by auto
- thus \<open>atom z \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>, \<Delta>, v, s1, s2, t2)\<close> using check_ifI fresh_prodN by auto
- show \<open>\<Theta> ; \<B> ; \<Gamma> \<turnstile> v \<Leftarrow> \<lbrace> z : B_bool | TRUE \<rbrace>\<close> using check_ifI by auto
- show \<open> \<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile> s1 \<Leftarrow> \<lbrace> z : b_of t2 | [ v ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of t2 z \<rbrace>\<close>
- using check_ifI subtype_if1 fresh_prodN base_for_lit.simps b_of.simps * check_v_wf by metis
-
- show \<open> \<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile> s2 \<Leftarrow> \<lbrace> z : b_of t2 | [ v ]\<^sup>c\<^sup>e == [ [ L_false ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of t2 z \<rbrace>\<close>
- using check_ifI subtype_if1 fresh_prodN base_for_lit.simps b_of.simps * check_v_wf by metis
- qed
-next
- case (check_assertI x \<Theta> \<Phi> \<B> \<Gamma> \<Delta> c \<tau> s)
-
- show ?case proof
- have "atom x \<sharp> t2" using subtype_fresh_tau[OF _ _ \<open>\<Theta> ; \<B> ; \<Gamma> \<turnstile> \<tau> \<lesssim> t2\<close>] check_assertI fresh_prodN by simp
- thus "atom x \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>, \<Delta>, c, t2, s)" using subtype_fresh_tau check_assertI fresh_prodN by simp
- have "\<Theta> ; \<B> ; (x, B_bool, c) #\<^sub>\<Gamma> \<Gamma> \<turnstile> \<tau> \<lesssim> t2" apply(rule subtype_weakening)
- using check_assertI apply simp
- using toSet.simps apply blast
- using check_assertI check_s_wf by simp
- thus "\<Theta> ; \<Phi> ; \<B> ; (x, B_bool, c) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile> s \<Leftarrow> t2" using check_assertI by simp
- show "\<Theta> ; \<B> ; \<Gamma> \<Turnstile> c " using check_assertI by auto
- show "\<Theta> ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> " using check_assertI by auto
- qed
-next
- case (check_let2I x P \<Phi> \<B> G \<Delta> t s1 \<tau> s2 )
- have "wfG P \<B> ((x, b_of t, c_of t x) #\<^sub>\<Gamma> G)"
- using check_let2I check_s_wf by metis
- moreover have "toSet G \<subseteq> toSet ((x, b_of t, c_of t x) #\<^sub>\<Gamma> G)" by auto
- ultimately have *:"P ; \<B> ; (x, b_of t, c_of t x) #\<^sub>\<Gamma> G \<turnstile> \<tau> \<lesssim> t2" using check_let2I subtype_weakening by metis
- show ?case proof(rule Typing.check_let2I)
- have "atom x \<sharp> t2" using subtype_fresh_tau[of x \<tau> ] check_let2I fresh_prodN by metis
- thus "atom x \<sharp> (P, \<Phi>, \<B>, G, \<Delta>, t, s1, t2)" using check_let2I fresh_prodN by metis
- show "P ; \<Phi> ; \<B> ; G ; \<Delta> \<turnstile> s1 \<Leftarrow> t" using check_let2I by blast
- show "P ; \<Phi> ; \<B> ;(x, b_of t, c_of t x ) #\<^sub>\<Gamma> G ; \<Delta> \<turnstile> s2 \<Leftarrow> t2" using check_let2I * by blast
- qed
-next
- case (check_varI u \<Theta> \<Phi> \<B> \<Gamma> \<Delta> \<tau>' v \<tau> s)
- show ?case proof(rule Typing.check_varI)
- have "atom u \<sharp> t2" using u_fresh_t by auto
- thus \<open>atom u \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>, \<Delta>, \<tau>', v, t2)\<close> using check_varI fresh_prodN by auto
- show \<open>\<Theta> ; \<B> ; \<Gamma> \<turnstile> v \<Leftarrow> \<tau>'\<close> using check_varI by auto
- show \<open> \<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; (u, \<tau>') #\<^sub>\<Delta> \<Delta> \<turnstile> s \<Leftarrow> t2\<close> using check_varI by auto
- qed
-next
- case (check_assignI \<Delta> u \<tau> P G v z \<tau>')
- then show ?case using Typing.check_assignI by (meson subtype_trans)
-next
- case (check_whileI \<Delta> G P s1 z s2 \<tau>')
- then show ?case using Typing.check_whileI by (meson subtype_trans)
-next
- case (check_seqI \<Delta> G P s1 z s2 \<tau>)
- then show ?case using Typing.check_seqI by blast
-next
- case (check_caseI \<Delta> \<Gamma> \<Theta> tid cs \<tau> v z)
- then show ?case using Typing.check_caseI subtype_trans by meson
-
-qed
-
-lemma subtype_let:
- fixes s'::s and cs::branch_s and css::branch_list and v::v
- shows "\<Theta> ; \<Phi> ; \<B> ; GNil ; \<Delta> \<turnstile> AS_let x e\<^sub>1 s \<Leftarrow> \<tau> \<Longrightarrow> \<Theta> ; \<Phi> ; \<B> ; GNil ; \<Delta> \<turnstile> e\<^sub>1 \<Rightarrow> \<tau>\<^sub>1 \<Longrightarrow>
- \<Theta> ; \<Phi> ; \<B> ; GNil ; \<Delta> \<turnstile> e\<^sub>2 \<Rightarrow> \<tau>\<^sub>2 \<Longrightarrow> \<Theta> ; \<B> ; GNil \<turnstile> \<tau>\<^sub>2 \<lesssim> \<tau>\<^sub>1 \<Longrightarrow> \<Theta> ; \<Phi> ; \<B> ; GNil ; \<Delta> \<turnstile> AS_let x e\<^sub>2 s \<Leftarrow> \<tau>" and
- "check_branch_s \<Theta> \<Phi> {||} GNil \<Delta> tid dc const v cs \<tau> \<Longrightarrow> True" and
- "check_branch_list \<Theta> \<Phi> {||} \<Gamma> \<Delta> tid dclist v css \<tau> \<Longrightarrow> True"
-proof(nominal_induct GNil \<Delta> "AS_let x e\<^sub>1 s" \<tau> and \<tau> and \<tau> avoiding: e\<^sub>2 \<tau>\<^sub>1 \<tau>\<^sub>2
- rule: check_s_check_branch_s_check_branch_list.strong_induct)
- case (check_letI x1 \<Theta> \<Phi> \<B> \<Delta> \<tau>1 z1 s1 b1 c1)
- obtain z2 and b2 and c2 where t2:"\<tau>\<^sub>2 = \<lbrace> z2 : b2 | c2 \<rbrace> \<and> atom z2 \<sharp> (x1, \<Theta>, \<Phi>, \<B>, GNil, \<Delta>, e\<^sub>2, \<tau>1, s1) "
- using obtain_fresh_z by metis
-
- obtain z1a and b1a and c1a where t1:"\<tau>\<^sub>1 = \<lbrace> z1a : b1a | c1a \<rbrace> \<and> atom z1a \<sharp> x1" using infer_e_uniqueness check_letI by metis
- hence t3: " \<lbrace> z1a : b1a | c1a \<rbrace> = \<lbrace> z1 : b1 | c1 \<rbrace> " using infer_e_uniqueness check_letI by metis
-
- have beq: "b1a = b2 \<and> b2 = b1" using check_letI subtype_eq_base t1 t2 t3 by metis
-
- have " \<Theta> ; \<Phi> ; \<B> ; GNil ; \<Delta> \<turnstile> AS_let x1 e\<^sub>2 s1 \<Leftarrow> \<tau>1" proof
- show \<open>atom x1 \<sharp> (\<Theta>, \<Phi>, \<B>, GNil, \<Delta>, e\<^sub>2, \<tau>1)\<close> using check_letI t2 fresh_prodN by metis
- show \<open>atom z2 \<sharp> (x1, \<Theta>, \<Phi>, \<B>, GNil, \<Delta>, e\<^sub>2, \<tau>1, s1)\<close> using check_letI t2 using check_letI t2 fresh_prodN by metis
- show \<open>\<Theta> ; \<Phi> ; \<B> ; GNil ; \<Delta> \<turnstile> e\<^sub>2 \<Rightarrow> \<lbrace> z2 : b2 | c2 \<rbrace>\<close> using check_letI t2 by metis
-
- have \<open> \<Theta> ; \<Phi> ; \<B> ; GNil@(x1, b2, c2[z2::=[ x1 ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> GNil ; \<Delta> \<turnstile> s1 \<Leftarrow> \<tau>1\<close>
- proof(rule ctx_subtype_s)
- have "c1a[z1a::=[ x1 ]\<^sup>v]\<^sub>c\<^sub>v = c1[z1::=[ x1 ]\<^sup>v]\<^sub>c\<^sub>v" using subst_v_flip_eq_two subst_v_c_def t3 \<tau>.eq_iff by metis
- thus \<open> \<Theta> ; \<Phi> ; \<B> ; GNil @ (x1, b2, c1a[z1a::=[ x1 ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> GNil ; \<Delta> \<turnstile> s1 \<Leftarrow> \<tau>1\<close> using check_letI beq append_g.simps subst_defs by metis
- show \<open>\<Theta> ; \<B> ; GNil \<turnstile> \<lbrace> z2 : b2 | c2 \<rbrace> \<lesssim> \<lbrace> z1a : b2 | c1a \<rbrace>\<close> using check_letI beq t1 t2 by metis
- have "atom x1 \<sharp> c2" using t2 check_letI \<tau>_fresh_c fresh_prodN by blast
- moreover have "atom x1 \<sharp> c1a" using t1 check_letI \<tau>_fresh_c fresh_prodN by blast
- ultimately show \<open>atom x1 \<sharp> (z1a, z2, c1a, c2)\<close> using t1 t2 fresh_prodN fresh_x_neq by metis
- qed
-
- thus \<open> \<Theta> ; \<Phi> ; \<B> ; (x1, b2, c2[z2::=[ x1 ]\<^sup>v]\<^sub>v) #\<^sub>\<Gamma> GNil ; \<Delta> \<turnstile> s1 \<Leftarrow> \<tau>1\<close> using append_g.simps subst_defs by metis
- qed
-
- moreover have "AS_let x1 e\<^sub>2 s1 = AS_let x e\<^sub>2 s" using check_letI s_branch_s_branch_list.eq_iff by metis
-
- ultimately show ?case by metis
-
-qed(auto+)
-
+(*<*)
+theory IVSubstTypingL
+ imports SubstMethods ContextSubtypingL
+begin
+ (*>*)
+
+chapter \<open>Immutable Variable Substitution Lemmas\<close>
+text \<open>Lemmas that show that types are preserved, in some way,
+under immutable variable substitution\<close>
+
+section \<open>Proof Methods\<close>
+
+method subst_mth = (metis subst_g_inside infer_e_wf infer_v_wf infer_v_wf)
+
+method subst_tuple_mth uses add = (
+ (unfold fresh_prodN), (simp add: add )+,
+ (rule,metis fresh_z_subst_g add fresh_Pair ),
+ (metis fresh_subst_dv add fresh_Pair ) )
+
+section \<open>Prelude\<close>
+
+lemma subst_top_eq:
+ "\<lbrace> z : b | TRUE \<rbrace> = \<lbrace> z : b | TRUE \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v"
+proof -
+ obtain z'::x and c' where zeq: "\<lbrace> z : b | TRUE \<rbrace> = \<lbrace> z' : b | c' \<rbrace> \<and> atom z' \<sharp> (x,v)" using obtain_fresh_z2 b_of.simps by metis
+ hence "\<lbrace> z' : b | TRUE \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v = \<lbrace> z' : b | TRUE \<rbrace>" using subst_tv.simps subst_cv.simps by metis
+ moreover have "c' = C_true" using \<tau>.eq_iff Abs1_eq_iff(3) c.fresh flip_fresh_fresh by (metis zeq)
+ ultimately show ?thesis using zeq by metis
+qed
+
+lemma wfD_subst:
+ fixes \<tau>\<^sub>1::\<tau> and v::v and \<Delta>::\<Delta> and \<Theta>::\<Theta> and \<Gamma>::\<Gamma>
+ assumes "\<Theta> ; \<B> ; \<Gamma> \<turnstile> v \<Rightarrow> \<tau>\<^sub>1" and "wfD \<Theta> \<B> (\<Gamma>'@((x,b\<^sub>1,c0[z0::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)) \<Delta>" and "b_of \<tau>\<^sub>1=b\<^sub>1"
+ shows " \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v"
+proof -
+ have "\<Theta> ; \<B> ; \<Gamma>\<turnstile>\<^sub>w\<^sub>f v : b\<^sub>1" using infer_v_v_wf assms by auto
+ moreover have "(\<Gamma>'@((x,b\<^sub>1,c0[z0::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\<Gamma>\<Gamma>))[x::=v]\<^sub>\<Gamma>\<^sub>v = \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>" using subst_g_inside wfD_wf assms by metis
+ ultimately show ?thesis using wf_subst assms by metis
+qed
+
+lemma subst_v_c_of:
+ assumes "atom xa \<sharp> (v,x)"
+ shows "c_of t[x::=v]\<^sub>\<tau>\<^sub>v xa = (c_of t xa)[x::=v]\<^sub>c\<^sub>v"
+ using assms proof(nominal_induct t avoiding: x v xa rule:\<tau>.strong_induct)
+ case (T_refined_type z' b' c')
+ then have " c_of \<lbrace> z' : b' | c' \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v xa = c_of \<lbrace> z' : b' | c'[x::=v]\<^sub>c\<^sub>v \<rbrace> xa"
+ using subst_tv.simps fresh_Pair by metis
+ also have "... = c'[x::=v]\<^sub>c\<^sub>v [z'::=V_var xa]\<^sub>c\<^sub>v" using c_of.simps T_refined_type by metis
+ also have "... = c' [z'::=V_var xa]\<^sub>c\<^sub>v[x::=v]\<^sub>c\<^sub>v"
+ using subst_cv_commute_full[of z' v x "V_var xa" c'] subst_v_c_def T_refined_type fresh_Pair fresh_at_base v.fresh fresh_x_neq by metis
+ finally show ?case using c_of.simps T_refined_type by metis
+qed
+
+section \<open>Context\<close>
+
+lemma subst_lookup:
+ assumes "Some (b,c) = lookup (\<Gamma>'@((x,b\<^sub>1,c\<^sub>1)#\<^sub>\<Gamma>\<Gamma>)) y" and "x \<noteq> y" and "wfG \<Theta> \<B> (\<Gamma>'@((x,b\<^sub>1,c\<^sub>1)#\<^sub>\<Gamma>\<Gamma>))"
+ shows "\<exists>d. Some (b,d) = lookup ((\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v)@\<Gamma>) y"
+ using assms proof(induct \<Gamma>' rule: \<Gamma>_induct)
+ case GNil
+ hence "Some (b,c) = lookup \<Gamma> y" by (simp add: assms(1))
+ then show ?case using subst_gv.simps by auto
+next
+ case (GCons x1 b1 c1 \<Gamma>1)
+ show ?case proof(cases "x1 = x")
+ case True
+ hence "atom x \<sharp> (\<Gamma>1 @ (x, b\<^sub>1, c\<^sub>1) #\<^sub>\<Gamma> \<Gamma>)" using GCons wfG_elims(2)
+ append_g.simps by metis
+ moreover have "atom x \<in> atom_dom (\<Gamma>1 @ (x, b\<^sub>1, c\<^sub>1) #\<^sub>\<Gamma> \<Gamma>)" by simp
+ ultimately show ?thesis
+ using forget_subst_gv not_GCons_self2 subst_gv.simps append_g.simps
+ by (metis GCons.prems(3) True wfG_cons_fresh2 )
+ next
+ case False
+ hence "((x1,b1,c1) #\<^sub>\<Gamma> \<Gamma>1)[x::=v]\<^sub>\<Gamma>\<^sub>v = (x1,b1,c1[x::=v]\<^sub>c\<^sub>v)#\<^sub>\<Gamma>\<Gamma>1[x::=v]\<^sub>\<Gamma>\<^sub>v" using subst_gv.simps by auto
+ then show ?thesis proof(cases "x1=y")
+ case True
+ then show ?thesis using GCons using lookup.simps
+ by (metis \<open>((x1, b1, c1) #\<^sub>\<Gamma> \<Gamma>1)[x::=v]\<^sub>\<Gamma>\<^sub>v = (x1, b1, c1[x::=v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>1[x::=v]\<^sub>\<Gamma>\<^sub>v\<close> append_g.simps fst_conv option.inject)
+ next
+ case False
+ then show ?thesis using GCons using lookup.simps
+ using \<open>((x1, b1, c1) #\<^sub>\<Gamma> \<Gamma>1)[x::=v]\<^sub>\<Gamma>\<^sub>v = (x1, b1, c1[x::=v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>1[x::=v]\<^sub>\<Gamma>\<^sub>v\<close> append_g.simps \<Gamma>.distinct \<Gamma>.inject wfG.simps wfG_elims by metis
+ qed
+ qed
+qed
+
+section \<open>Validity\<close>
+
+lemma subst_self_valid:
+ fixes v::v
+ assumes "\<Theta> ; \<B> ; G \<turnstile> v \<Rightarrow> \<lbrace> z : b | c \<rbrace>" and "atom z \<sharp> v"
+ shows " \<Theta> ; \<B> ; G \<Turnstile> c[z::=v]\<^sub>c\<^sub>v"
+proof -
+ have "c= (CE_val (V_var z) == CE_val v )" using infer_v_form2 assms by presburger
+ hence "c[z::=v]\<^sub>c\<^sub>v = (CE_val (V_var z) == CE_val v )[z::=v]\<^sub>c\<^sub>v" by auto
+ also have "... = (((CE_val (V_var z))[z::=v]\<^sub>c\<^sub>e\<^sub>v) == ((CE_val v)[z::=v]\<^sub>c\<^sub>e\<^sub>v))" by fastforce
+ also have "... = ((CE_val v) == ((CE_val v)[z::=v]\<^sub>c\<^sub>e\<^sub>v))" using subst_cev.simps subst_vv.simps by presburger
+ also have "... = (CE_val v == CE_val v )" using infer_v_form subst_cev.simps assms forget_subst_vv by presburger
+ finally have *:"c[z::=v]\<^sub>c\<^sub>v = (CE_val v == CE_val v )" by auto
+
+ have **:"\<Theta> ; \<B> ; G\<turnstile>\<^sub>w\<^sub>f CE_val v : b" using wfCE_valI assms infer_v_v_wf b_of.simps by metis
+
+ show ?thesis proof(rule validI)
+ show "\<Theta> ; \<B> ; G\<turnstile>\<^sub>w\<^sub>f c[z::=v]\<^sub>c\<^sub>v" proof -
+ have "\<Theta> ; \<B> ; G\<turnstile>\<^sub>w\<^sub>f v : b" using infer_v_v_wf assms b_of.simps by metis
+ moreover have "\<Theta> \<turnstile>\<^sub>w\<^sub>f ([]::\<Phi>) \<and> \<Theta> ; \<B> ; G\<turnstile>\<^sub>w\<^sub>f []\<^sub>\<Delta>" using wfD_emptyI wfPhi_emptyI infer_v_wf assms by auto
+ ultimately show ?thesis using * wfCE_valI wfC_eqI by metis
+ qed
+ show "\<forall>i. wfI \<Theta> G i \<and> is_satis_g i G \<longrightarrow> is_satis i c[z::=v]\<^sub>c\<^sub>v" proof(rule,rule)
+ fix i
+ assume \<open>wfI \<Theta> G i \<and> is_satis_g i G\<close>
+ thus \<open>is_satis i c[z::=v]\<^sub>c\<^sub>v\<close> using * ** is_satis_eq by auto
+ qed
+ qed
+qed
+
+lemma subst_valid_simple:
+ fixes v::v
+ assumes "\<Theta> ; \<B> ; G \<turnstile> v \<Rightarrow> \<lbrace> z0 : b | c0 \<rbrace>" and
+ "atom z0 \<sharp> c" and "atom z0 \<sharp> v"
+ "\<Theta>; \<B> ; (z0,b,c0)#\<^sub>\<Gamma>G \<Turnstile> c[z::=V_var z0]\<^sub>c\<^sub>v"
+ shows " \<Theta> ; \<B> ; G \<Turnstile> c[z::=v]\<^sub>c\<^sub>v"
+proof -
+ have " \<Theta> ; \<B> ; G \<Turnstile> c0[z0::=v]\<^sub>c\<^sub>v" using subst_self_valid assms by metis
+ moreover have "atom z0 \<sharp> G" using assms valid_wf_all by meson
+ moreover have "wfV \<Theta> \<B> G v b" using infer_v_v_wf assms b_of.simps by metis
+ moreover have "(c[z::=V_var z0]\<^sub>c\<^sub>v)[z0::=v]\<^sub>c\<^sub>v = c[z::=v]\<^sub>c\<^sub>v" using subst_v_simple_commute assms subst_v_c_def by metis
+ ultimately show ?thesis using valid_trans assms subst_defs by metis
+qed
+
+lemma wfI_subst1:
+ assumes " wfI \<Theta> (G'[x::=v]\<^sub>\<Gamma>\<^sub>v @ G) i" and "wfG \<Theta> \<B> (G' @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> G)" and "eval_v i v sv" and "wfRCV \<Theta> sv b"
+ shows "wfI \<Theta> (G' @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> G) ( i( x \<mapsto> sv))"
+proof -
+ {
+ fix xa::x and ba::b and ca::c
+ assume as: "(xa,ba,ca) \<in> toSet ((G' @ ((x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> G)))"
+ then have " \<exists>s. Some s = (i(x \<mapsto> sv)) xa \<and> wfRCV \<Theta> s ba"
+ proof(cases "x=xa")
+ case True
+ have "Some sv = (i(x \<mapsto> sv)) x \<and> wfRCV \<Theta> sv b" using as assms wfI_def by auto
+ moreover have "b=ba" using assms as True wfG_member_unique by metis
+ ultimately show ?thesis using True by auto
+ next
+ case False
+
+ then obtain ca' where "(xa, ba, ca') \<in> toSet (G'[x::=v]\<^sub>\<Gamma>\<^sub>v @ G)" using wfG_member_subst2 assms as by metis
+ then obtain s where " Some s = i xa \<and> wfRCV \<Theta> s ba" using wfI_def assms False by blast
+ thus ?thesis using False by auto
+ qed
+ }
+ from this show ?thesis using wfI_def allI by blast
+qed
+
+lemma subst_valid:
+ fixes v::v and c'::c and \<Gamma> ::\<Gamma>
+ assumes "\<Theta> ; \<B> ; \<Gamma> \<Turnstile> c[z::=v]\<^sub>c\<^sub>v" and "\<Theta> ; \<B> ; \<Gamma>\<turnstile>\<^sub>w\<^sub>f v : b" and
+ "\<Theta> ; \<B>\<turnstile>\<^sub>w\<^sub>f \<Gamma>" and "atom x \<sharp> c" and "atom x \<sharp> \<Gamma>" and
+ "\<Theta> ; \<B>\<turnstile>\<^sub>w\<^sub>f (\<Gamma>'@(x,b,c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v ) #\<^sub>\<Gamma> \<Gamma>)" and
+ "\<Theta> ; \<B> ; \<Gamma>'@(x,b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v ) #\<^sub>\<Gamma> \<Gamma> \<Turnstile> c'" (is " \<Theta> ; \<B>; ?G \<Turnstile> c'")
+ shows "\<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v@\<Gamma> \<Turnstile> c'[x::=v]\<^sub>c\<^sub>v"
+proof -
+ have *:"wfC \<Theta> \<B> (\<Gamma>'@(x,b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v ) #\<^sub>\<Gamma> \<Gamma>) c'" using valid.simps assms by metis
+ hence "wfC \<Theta> \<B> (\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>) (c'[x::=v]\<^sub>c\<^sub>v)" using wf_subst(2)[OF *] b_of.simps assms subst_g_inside wfC_wf by metis
+ moreover have "\<forall>i. wfI \<Theta> (\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>) i \<and> is_satis_g i (\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>) \<longrightarrow> is_satis i (c'[x::=v]\<^sub>c\<^sub>v)"
+
+ proof(rule,rule)
+ fix i
+ assume as: " wfI \<Theta> (\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>) i \<and> is_satis_g i (\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>)"
+
+ hence wfig: "wfI \<Theta> \<Gamma> i" using wfI_suffix infer_v_wf assms by metis
+ then obtain s where s:"eval_v i v s" and b:"wfRCV \<Theta> s b" using eval_v_exist infer_v_v_wf b_of.simps assms by metis
+
+ have is1: "is_satis_g ( i( x \<mapsto> s)) (\<Gamma>' @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)" proof(rule is_satis_g_i_upd2)
+ show "is_satis (i(x \<mapsto> s)) (c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v)" proof -
+ have "is_satis i (c[z::=v]\<^sub>c\<^sub>v)"
+ using subst_valid_simple assms as valid.simps infer_v_wf assms
+ is_satis_g_suffix wfI_suffix by metis
+ hence "is_satis i ((c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v)[x::=v]\<^sub>c\<^sub>v)" using assms subst_v_simple_commute[of x c z v] subst_v_c_def by metis
+ moreover have "\<Theta> ; \<B> ; (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v" using wfC_refl wfG_suffix assms by metis
+ moreover have "\<Theta> ; \<B> ; \<Gamma>\<turnstile>\<^sub>w\<^sub>f v : b" using assms infer_v_v_wf b_of.simps by metis
+ ultimately show ?thesis using subst_c_satis[OF s , of \<Theta> \<B> x b "c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v" \<Gamma> "c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v"] wfig by auto
+ qed
+ show "atom x \<sharp> \<Gamma>" using assms by metis
+ show "wfG \<Theta> \<B> (\<Gamma>' @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)" using valid_wf_all assms by metis
+ show "\<Theta> ; \<B> ; \<Gamma>\<turnstile>\<^sub>w\<^sub>f v : b" using assms infer_v_v_wf by force
+ show "i \<lbrakk> v \<rbrakk> ~ s " using s by auto
+ show "\<Theta> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> i" using as by auto
+ show "i \<Turnstile> \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> " using as by auto
+ qed
+ hence "is_satis ( i( x \<mapsto> s)) c'" proof -
+ have "wfI \<Theta> (\<Gamma>' @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) ( i( x \<mapsto> s))"
+ using wfI_subst1[of \<Theta> \<Gamma>' x v \<Gamma> i \<B> b c z s] as b s assms by metis
+ thus ?thesis using is1 valid.simps assms by presburger
+ qed
+
+ thus "is_satis i (c'[x::=v]\<^sub>c\<^sub>v)" using subst_c_satis_full[OF s] valid.simps as infer_v_v_wf b_of.simps assms by metis
+
+ qed
+ ultimately show ?thesis using valid.simps by auto
+qed
+
+lemma subst_valid_infer_v:
+ fixes v::v and c'::c
+ assumes "\<Theta> ; \<B> ; G \<turnstile> v \<Rightarrow> \<lbrace> z0 : b | c0 \<rbrace>" and "atom x \<sharp> c" and "atom x \<sharp> G" and "wfG \<Theta> \<B> (G'@(x,b,c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v ) #\<^sub>\<Gamma> G)" and "atom z0 \<sharp> v"
+ " \<Theta>;\<B>;(z0,b,c0)#\<^sub>\<Gamma>G \<Turnstile> c[z::=V_var z0]\<^sub>c\<^sub>v" and "atom z0 \<sharp> c" and
+ " \<Theta>;\<B>;G'@(x,b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v ) #\<^sub>\<Gamma> G \<Turnstile> c'" (is " \<Theta> ; \<B>; ?G \<Turnstile> c'")
+ shows " \<Theta>;\<B>;G'[x::=v]\<^sub>\<Gamma>\<^sub>v@G \<Turnstile> c'[x::=v]\<^sub>c\<^sub>v"
+proof -
+ have "\<Theta> ; \<B> ; G \<Turnstile> c[z::=v]\<^sub>c\<^sub>v"
+ using infer_v_wf subst_valid_simple valid.simps assms using subst_valid_simple assms valid.simps infer_v_wf assms
+ is_satis_g_suffix wfI_suffix by metis
+ moreover have "wfV \<Theta> \<B> G v b" and "wfG \<Theta> \<B> G"
+ using assms infer_v_wf b_of.simps apply metis using assms infer_v_wf by metis
+ ultimately show ?thesis using assms subst_valid by metis
+qed
+
+section \<open>Subtyping\<close>
+
+lemma subst_subtype:
+ fixes v::v
+ assumes "\<Theta> ; \<B> ; \<Gamma> \<turnstile> v \<Rightarrow> (\<lbrace>z0:b|c0\<rbrace>)" and
+ " \<Theta>;\<B>;\<Gamma> \<turnstile> (\<lbrace>z0:b|c0\<rbrace>) \<lesssim> (\<lbrace> z : b | c \<rbrace>)" and
+ " \<Theta>;\<B>;\<Gamma>'@((x,b,c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\<Gamma>\<Gamma>) \<turnstile> (\<lbrace> z1 : b1 | c1 \<rbrace>) \<lesssim> (\<lbrace> z2 : b1 | c2 \<rbrace>)" (is " \<Theta> ; \<B>; ?G1 \<turnstile> ?t1 \<lesssim> ?t2" ) and
+ "atom z \<sharp> (x,v) \<and> atom z0 \<sharp> (c,x,v,z,\<Gamma>) \<and> atom z1 \<sharp> (x,v) \<and> atom z2 \<sharp> (x,v)" and "wsV \<Theta> \<B> \<Gamma> v"
+ shows " \<Theta>;\<B>;\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v@\<Gamma> \<turnstile> \<lbrace> z1 : b1 | c1 \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v \<lesssim> \<lbrace> z2 : b1 | c2 \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v"
+proof -
+ have z2: "atom z2 \<sharp> (x,v) " using assms by auto
+ hence "x \<noteq> z2" by auto
+
+ obtain xx::x where xxf: "atom xx \<sharp> (x,z1, c1, z2, c2, \<Gamma>' @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>, c1[x::=v]\<^sub>c\<^sub>v, c2[x::=v]\<^sub>c\<^sub>v, \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>,
+ (\<Theta> , \<B> , \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v@\<Gamma>, z1 , c1[x::=v]\<^sub>c\<^sub>v , z2 , c2[x::=v]\<^sub>c\<^sub>v ))" (is "atom xx \<sharp> ?tup")
+ using obtain_fresh by blast
+ hence xxf2: "atom xx \<sharp> (z1, c1, z2, c2, \<Gamma>' @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)" using fresh_prod9 fresh_prod5 by fast
+
+ have vd1: " \<Theta>;\<B>;((xx, b1, c1[z1::=V_var xx]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>')[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<Turnstile> (c2[z2::=V_var xx]\<^sub>c\<^sub>v)[x::=v]\<^sub>c\<^sub>v"
+ proof(rule subst_valid_infer_v[of \<Theta> _ _ _ z0 b c0 _ c, where z=z])
+ show "\<Theta> ; \<B> ; \<Gamma> \<turnstile> v \<Rightarrow> \<lbrace> z0 : b | c0 \<rbrace>" using assms by auto
+
+ show xf: "atom x \<sharp> \<Gamma>" using subtype_g_wf wfG_inside_fresh_suffix assms by metis
+
+ show "atom x \<sharp> c" proof -
+ have "wfT \<Theta> \<B> \<Gamma> (\<lbrace> z : b | c \<rbrace>)" using subtype_wf[OF assms(2)] by auto
+ moreover have "x \<noteq> z" using assms(4)
+ using fresh_Pair not_self_fresh by blast
+ ultimately show ?thesis using xf wfT_fresh_c assms by presburger
+ qed
+
+ show " \<Theta> ; \<B>\<turnstile>\<^sub>w\<^sub>f ((xx, b1, c1[z1::=V_var xx]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>') @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma> "
+ proof(subst append_g.simps,rule wfG_consI)
+ show *: \<open> \<Theta> ; \<B>\<turnstile>\<^sub>w\<^sub>f \<Gamma>' @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<close> using subtype_g_wf assms by metis
+ show \<open>atom xx \<sharp> \<Gamma>' @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>\<close> using xxf fresh_prod9 by metis
+ show \<open> \<Theta> ; \<B>\<turnstile>\<^sub>w\<^sub>f b1 \<close> using subtype_elims[OF assms(3)] wfT_wfC wfC_wf wfG_cons by metis
+ show "\<Theta> ; \<B> ; (xx, b1, TRUE) #\<^sub>\<Gamma> \<Gamma>' @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f c1[z1::=V_var xx]\<^sub>c\<^sub>v " proof(rule wfT_wfC)
+ have "\<lbrace> z1 : b1 | c1 \<rbrace> = \<lbrace> xx : b1 | c1[z1::=V_var xx]\<^sub>c\<^sub>v \<rbrace> " using xxf fresh_prod9 type_eq_subst xxf2 fresh_prodN by metis
+ thus "\<Theta> ; \<B> ; \<Gamma>' @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> xx : b1 | c1[z1::=V_var xx]\<^sub>c\<^sub>v \<rbrace> " using subtype_wfT[OF assms(3)] by metis
+ show "atom xx \<sharp> \<Gamma>' @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>" using xxf fresh_prod9 by metis
+ qed
+ qed
+
+ show "atom z0 \<sharp> v" using assms fresh_prod5 by auto
+ have "\<Theta> ; \<B> ; (z0, b, c0) #\<^sub>\<Gamma> \<Gamma> \<Turnstile> c[z::=V_var z0]\<^sub>v "
+ apply(rule obtain_fresh[of "(z0,c0, \<Gamma>, c, z)"],rule subtype_valid[OF assms(2), THEN valid_flip],
+ (fastforce simp add: assms fresh_prodN)+) done
+ thus "\<Theta> ; \<B> ; (z0, b, c0) #\<^sub>\<Gamma> \<Gamma> \<Turnstile> c[z::=V_var z0]\<^sub>c\<^sub>v " using subst_defs by auto
+
+ show "atom z0 \<sharp> c" using assms fresh_prod5 by auto
+ show "\<Theta> ; \<B> ; ((xx, b1, c1[z1::=V_var xx]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>') @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<Turnstile> c2[z2::=V_var xx]\<^sub>c\<^sub>v "
+ using subtype_valid assms(3) xxf xxf2 fresh_prodN append_g.simps subst_defs by metis
+ qed
+
+ have xfw1: "atom z1 \<sharp> v \<and> atom x \<sharp> [ xx ]\<^sup>v \<and> x \<noteq> z1"
+ apply(intro conjI)
+ apply(simp add: assms xxf fresh_at_base fresh_prodN freshers fresh_x_neq)+
+ using fresh_x_neq fresh_prodN xxf apply blast
+ using fresh_x_neq fresh_prodN assms by blast
+
+ have xfw2: "atom z2 \<sharp> v \<and> atom x \<sharp> [ xx ]\<^sup>v \<and> x \<noteq> z2"
+ apply(auto simp add: assms xxf fresh_at_base fresh_prodN freshers)
+ by(insert xxf fresh_at_base fresh_prodN assms, fast+)
+
+ have wf1: "wfT \<Theta> \<B> (\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v@\<Gamma>) (\<lbrace> z1 : b1 | c1[x::=v]\<^sub>c\<^sub>v \<rbrace>)" proof -
+ have "wfT \<Theta> \<B> (\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v@\<Gamma>) (\<lbrace> z1 : b1 | c1 \<rbrace>)[x::=v]\<^sub>\<tau>\<^sub>v"
+ using wf_subst(4) assms b_of.simps infer_v_v_wf subtype_wf subst_tv.simps subst_g_inside wfT_wf by metis
+ moreover have "atom z1 \<sharp> (x,v)" using assms by auto
+ ultimately show ?thesis using subst_tv.simps by auto
+ qed
+ moreover have wf2: "wfT \<Theta> \<B> (\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v@\<Gamma>) (\<lbrace> z2 : b1 | c2[x::=v]\<^sub>c\<^sub>v \<rbrace>)" proof -
+ have "wfT \<Theta> \<B> (\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v@\<Gamma>) (\<lbrace> z2 : b1 | c2 \<rbrace>)[x::=v]\<^sub>\<tau>\<^sub>v" using wf_subst(4) assms b_of.simps infer_v_v_wf subtype_wf subst_tv.simps subst_g_inside wfT_wf by metis
+ moreover have "atom z2 \<sharp> (x,v)" using assms by auto
+ ultimately show ?thesis using subst_tv.simps by auto
+ qed
+ moreover have "\<Theta> ; \<B> ; (xx, b1, c1[x::=v]\<^sub>c\<^sub>v[z1::=V_var xx]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> (\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> ) \<Turnstile> (c2[x::=v]\<^sub>c\<^sub>v)[z2::=V_var xx]\<^sub>c\<^sub>v" proof -
+ have "xx \<noteq> x" using xxf fresh_Pair fresh_at_base by fast
+ hence "((xx, b1, subst_cv c1 z1 (V_var xx) ) #\<^sub>\<Gamma> \<Gamma>')[x::=v]\<^sub>\<Gamma>\<^sub>v = (xx, b1, (subst_cv c1 z1 (V_var xx) )[x::=v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> (\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v)"
+ using subst_gv.simps by auto
+ moreover have "(c1[z1::=V_var xx]\<^sub>c\<^sub>v )[x::=v]\<^sub>c\<^sub>v = (c1[x::=v]\<^sub>c\<^sub>v) [z1::=V_var xx]\<^sub>c\<^sub>v" using subst_cv_commute_full xfw1 by metis
+ moreover have "c2[z2::=[ xx ]\<^sup>v]\<^sub>c\<^sub>v[x::=v]\<^sub>c\<^sub>v = (c2[x::=v]\<^sub>c\<^sub>v)[z2::=V_var xx]\<^sub>c\<^sub>v" using subst_cv_commute_full xfw2 by metis
+ ultimately show ?thesis using vd1 append_g.simps by metis
+ qed
+ moreover have "atom xx \<sharp> (\<Theta> , \<B> , \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v@\<Gamma>, z1 , c1[x::=v]\<^sub>c\<^sub>v , z2 ,c2[x::=v]\<^sub>c\<^sub>v )"
+ using xxf fresh_prodN by metis
+ ultimately have "\<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v@\<Gamma> \<turnstile> \<lbrace> z1 : b1 | c1[x::=v]\<^sub>c\<^sub>v \<rbrace> \<lesssim> \<lbrace> z2 : b1 | c2[x::=v]\<^sub>c\<^sub>v \<rbrace>"
+ using subtype_baseI subst_defs by metis
+ thus ?thesis using subst_tv.simps assms by presburger
+qed
+
+lemma subst_subtype_tau:
+ fixes v::v
+ assumes "\<Theta> ; \<B> ; \<Gamma> \<turnstile> v \<Rightarrow> \<tau>" and
+ "\<Theta> ; \<B> ; \<Gamma> \<turnstile> \<tau> \<lesssim> (\<lbrace> z : b | c \<rbrace>)"
+ "\<Theta> ; \<B> ; \<Gamma>'@((x,b,c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\<Gamma>\<Gamma>) \<turnstile> \<tau>1 \<lesssim> \<tau>2" and
+ "atom z \<sharp> (x,v)"
+ shows "\<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v@\<Gamma> \<turnstile> \<tau>1[x::=v]\<^sub>\<tau>\<^sub>v \<lesssim> \<tau>2[x::=v]\<^sub>\<tau>\<^sub>v"
+proof -
+ obtain z0 and b0 and c0 where zbc0: "\<tau>=(\<lbrace> z0 : b0 | c0 \<rbrace>) \<and> atom z0 \<sharp> (c,x,v,z,\<Gamma>)"
+ using obtain_fresh_z by metis
+ obtain z1 and b1 and c1 where zbc1: "\<tau>1=(\<lbrace> z1 : b1 | c1 \<rbrace>) \<and> atom z1 \<sharp> (x,v)"
+ using obtain_fresh_z by metis
+ obtain z2 and b2 and c2 where zbc2: "\<tau>2=(\<lbrace> z2 : b2 | c2 \<rbrace>) \<and> atom z2 \<sharp> (x,v)"
+ using obtain_fresh_z by metis
+
+ have "b0=b" using subtype_eq_base zbc0 assms by blast
+
+ hence vinf: "\<Theta> ; \<B> ; \<Gamma> \<turnstile> v \<Rightarrow> \<lbrace> z0 : b | c0 \<rbrace>" using assms zbc0 by blast
+ have vsub: "\<Theta> ; \<B> ; \<Gamma> \<turnstile>\<lbrace> z0 : b | c0 \<rbrace> \<lesssim> \<lbrace> z : b | c \<rbrace>" using assms zbc0 \<open>b0=b\<close> by blast
+ have beq:"b1=b2" using subtype_eq_base
+ using zbc1 zbc2 assms by blast
+ have "\<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> \<lbrace> z1 : b1 | c1 \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v \<lesssim> \<lbrace> z2 : b1 | c2 \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v"
+ proof(rule subst_subtype[OF vinf vsub] )
+ show "\<Theta> ; \<B> ; \<Gamma>'@((x,b,c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\<Gamma>\<Gamma>) \<turnstile> \<lbrace> z1 : b1 | c1 \<rbrace> \<lesssim> \<lbrace> z2 : b1 | c2 \<rbrace>"
+ using beq assms zbc1 zbc2 by auto
+ show "atom z \<sharp> (x, v) \<and> atom z0 \<sharp> (c, x, v, z, \<Gamma>) \<and> atom z1 \<sharp> (x, v) \<and> atom z2 \<sharp> (x, v)"
+ using zbc0 zbc1 zbc2 assms by blast
+ show "wfV \<Theta> \<B> \<Gamma> v (b_of \<tau>)" using infer_v_wf assms by simp
+ qed
+
+ thus ?thesis using zbc1 zbc2 \<open>b1=b2\<close> assms by blast
+qed
+
+lemma subtype_if1:
+ fixes v::v
+ assumes "P ; \<B> ; \<Gamma> \<turnstile> t1 \<lesssim> t2" and "wfV P \<B> \<Gamma> v (base_for_lit l)" and
+ "atom z1 \<sharp> v" and "atom z2 \<sharp> v" and "atom z1 \<sharp> t1" and "atom z2 \<sharp> t2" and "atom z1 \<sharp> \<Gamma>" and "atom z2 \<sharp> \<Gamma>"
+ shows "P ; \<B> ; \<Gamma> \<turnstile> \<lbrace> z1 : b_of t1 | CE_val v == CE_val (V_lit l) IMP (c_of t1 z1) \<rbrace> \<lesssim> \<lbrace> z2 : b_of t2 | CE_val v == CE_val (V_lit l) IMP (c_of t2 z2) \<rbrace>"
+proof -
+ obtain z1' where t1: "t1 = \<lbrace> z1' : b_of t1 | c_of t1 z1'\<rbrace> \<and> atom z1' \<sharp> (z1,\<Gamma>,t1)" using obtain_fresh_z_c_of by metis
+ obtain z2' where t2: "t2 = \<lbrace> z2' : b_of t2 | c_of t2 z2'\<rbrace> \<and> atom z2' \<sharp> (z2,t2) " using obtain_fresh_z_c_of by metis
+ have beq:"b_of t1 = b_of t2" using subtype_eq_base2 assms by auto
+
+ have c1: "(c_of t1 z1')[z1'::=[ z1 ]\<^sup>v]\<^sub>c\<^sub>v = c_of t1 z1" using c_of_switch t1 assms by simp
+ have c2: "(c_of t2 z2')[z2'::=[ z2 ]\<^sup>v]\<^sub>c\<^sub>v = c_of t2 z2" using c_of_switch t2 assms by simp
+
+ have "P ; \<B> ; \<Gamma> \<turnstile> \<lbrace> z1 : b_of t1 | [ v ]\<^sup>c\<^sup>e == [ [ l ]\<^sup>v ]\<^sup>c\<^sup>e IMP (c_of t1 z1')[z1'::=[z1]\<^sup>v]\<^sub>v \<rbrace> \<lesssim> \<lbrace> z2 : b_of t1 | [ v ]\<^sup>c\<^sup>e == [ [ l ]\<^sup>v ]\<^sup>c\<^sup>e IMP (c_of t2 z2')[z2'::=[z2]\<^sup>v]\<^sub>v \<rbrace>"
+ proof(rule subtype_if)
+ show \<open>P ; \<B> ; \<Gamma> \<turnstile> \<lbrace> z1' : b_of t1 | c_of t1 z1' \<rbrace> \<lesssim> \<lbrace> z2' : b_of t1 | c_of t2 z2' \<rbrace>\<close> using t1 t2 assms beq by auto
+ show \<open> P ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> z1 : b_of t1 | [ v ]\<^sup>c\<^sup>e == [ [ l ]\<^sup>v ]\<^sup>c\<^sup>e IMP (c_of t1 z1')[z1'::=[ z1 ]\<^sup>v]\<^sub>v \<rbrace> \<close> using wfT_wfT_if_rev assms subtype_wfT c1 subst_defs by metis
+ show \<open> P ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> z2 : b_of t1 | [ v ]\<^sup>c\<^sup>e == [ [ l ]\<^sup>v ]\<^sup>c\<^sup>e IMP (c_of t2 z2')[z2'::=[ z2 ]\<^sup>v]\<^sub>v \<rbrace> \<close> using wfT_wfT_if_rev assms subtype_wfT c2 subst_defs beq by metis
+ show \<open>atom z1 \<sharp> v\<close> using assms by auto
+ show \<open>atom z1' \<sharp> \<Gamma>\<close> using t1 by auto
+ show \<open>atom z1 \<sharp> c_of t1 z1'\<close> using t1 assms c_of_fresh by force
+ show \<open>atom z2 \<sharp> c_of t2 z2'\<close> using t2 assms c_of_fresh by force
+ show \<open>atom z2 \<sharp> v\<close> using assms by auto
+ qed
+ then show ?thesis using t1 t2 assms c1 c2 beq subst_defs by metis
+qed
+
+section \<open>Values\<close>
+
+lemma subst_infer_aux:
+ fixes \<tau>\<^sub>1::\<tau> and v'::v
+ assumes "\<Theta> ; \<B> ; \<Gamma> \<turnstile> v'[x::=v]\<^sub>v\<^sub>v \<Rightarrow> \<tau>\<^sub>1" and "\<Theta> ; \<B> ; \<Gamma>' \<turnstile> v' \<Rightarrow> \<tau>\<^sub>2" and "b_of \<tau>\<^sub>1 = b_of \<tau>\<^sub>2"
+ shows "\<tau>\<^sub>1 = (\<tau>\<^sub>2[x::=v]\<^sub>\<tau>\<^sub>v)"
+proof -
+ obtain z1 and b1 where zb1: "\<tau>\<^sub>1 = (\<lbrace> z1 : b1 | C_eq (CE_val (V_var z1)) (CE_val (v'[x::=v]\<^sub>v\<^sub>v)) \<rbrace>) \<and> atom z1 \<sharp> ((CE_val (v'[x::=v]\<^sub>v\<^sub>v), CE_val v),v'[x::=v]\<^sub>v\<^sub>v)"
+ using infer_v_form_fresh[OF assms(1)] by fastforce
+ obtain z2 and b2 where zb2: "\<tau>\<^sub>2 = (\<lbrace> z2 : b2 | C_eq (CE_val (V_var z2)) (CE_val v') \<rbrace>) \<and> atom z2 \<sharp> ((CE_val (v'[x::=v]\<^sub>v\<^sub>v), CE_val v,x,v),v')"
+ using infer_v_form_fresh [OF assms(2)] by fastforce
+ have beq: "b1 = b2" using assms zb1 zb2 by simp
+
+ hence "(\<lbrace> z2 : b2 | C_eq (CE_val (V_var z2)) (CE_val v') \<rbrace>)[x::=v]\<^sub>\<tau>\<^sub>v = (\<lbrace> z2 : b2 | C_eq (CE_val (V_var z2)) (CE_val (v'[x::=v]\<^sub>v\<^sub>v)) \<rbrace>)"
+ using subst_tv.simps subst_cv.simps subst_ev.simps forget_subst_vv[of x "V_var z2"] zb2 by force
+ also have "... = (\<lbrace> z1 : b1 | C_eq (CE_val (V_var z1)) (CE_val (v'[x::=v]\<^sub>v\<^sub>v)) \<rbrace>)"
+ using type_e_eq[of z2 "CE_val (v'[x::=v]\<^sub>v\<^sub>v)"z1 b1 ] zb1 zb2 fresh_PairD(1) assms beq by metis
+ finally show ?thesis using zb1 zb2 by argo
+qed
+
+lemma subst_t_b_eq:
+ fixes x::x and v::v
+ shows "b_of (\<tau>[x::=v]\<^sub>\<tau>\<^sub>v) = b_of \<tau>"
+proof -
+ obtain z and b and c where "\<tau> = \<lbrace> z : b | c \<rbrace> \<and> atom z \<sharp> (x,v)"
+ using has_fresh_z by blast
+ thus ?thesis using subst_tv.simps by simp
+qed
+
+lemma fresh_g_fresh_v:
+ fixes x::x
+ assumes "atom x \<sharp> \<Gamma>" and "wfV \<Theta> \<B> \<Gamma> v b"
+ shows "atom x \<sharp> v"
+ using assms wfV_supp wfX_wfY wfG_atoms_supp_eq fresh_def
+ by (metis wfV_x_fresh)
+
+lemma infer_v_fresh_g_fresh_v:
+ fixes x::x and \<Gamma>::\<Gamma> and v::v
+ assumes "atom x \<sharp> \<Gamma>'@\<Gamma>" and "\<Theta> ; \<B> ; \<Gamma> \<turnstile> v \<Rightarrow> \<tau>"
+ shows "atom x \<sharp> v"
+proof -
+ have "atom x \<sharp> \<Gamma>" using fresh_suffix assms by auto
+ moreover have "wfV \<Theta> \<B> \<Gamma> v (b_of \<tau>)" using infer_v_wf assms by auto
+ ultimately show ?thesis using fresh_g_fresh_v by metis
+qed
+
+lemma infer_v_fresh_g_fresh_xv:
+ fixes xa::x and v::v and \<Gamma>::\<Gamma>
+ assumes "atom xa \<sharp> \<Gamma>'@((x,b,c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\<Gamma>\<Gamma>)" and "\<Theta> ; \<B> ; \<Gamma> \<turnstile> v \<Rightarrow> \<tau>"
+ shows "atom xa \<sharp> (x,v)"
+proof -
+ have "atom xa \<sharp> x" using assms fresh_in_g fresh_def by blast
+ moreover have "\<Gamma>'@((x,b,c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\<Gamma>\<Gamma>) = ((\<Gamma>'@(x,b,c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\<Gamma>GNil)@\<Gamma>)" using append_g.simps append_g_assoc by simp
+ moreover hence "atom xa \<sharp> v" using infer_v_fresh_g_fresh_v assms by metis
+ ultimately show ?thesis by auto
+qed
+
+lemma wfG_subst_infer_v:
+ fixes v::v
+ assumes "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>' @ (x, b, c0[z0::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>" and "\<Theta> ; \<B> ; \<Gamma> \<turnstile> v \<Rightarrow> \<tau>" and "b_of \<tau> = b"
+ shows "\<Theta> ; \<B>\<turnstile>\<^sub>w\<^sub>f \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> "
+ using wfG_subst_wfV infer_v_v_wf assms by auto
+
+lemma fresh_subst_gv_inside:
+ fixes \<Gamma>::\<Gamma>
+ assumes "atom z \<sharp> \<Gamma>' @ (x, b\<^sub>1, c0[z0::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>" and "atom z \<sharp> v"
+ shows "atom z \<sharp> \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v@\<Gamma>"
+ unfolding fresh_append_g using fresh_append_g assms fresh_subst_gv fresh_GCons by metis
+
+lemma subst_t:
+ fixes x::x and b::b and xa::x
+ assumes "atom z \<sharp> x" and "atom z \<sharp> v"
+ shows "(\<lbrace> z : b | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ v'[x::=v]\<^sub>v\<^sub>v ]\<^sup>c\<^sup>e \<rbrace>) = (\<lbrace> z : b | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ v']\<^sup>c\<^sup>e \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v)"
+ using assms subst_vv.simps subst_tv.simps subst_cv.simps subst_cev.simps by auto
+
+lemma infer_l_fresh:
+ assumes "\<turnstile> l \<Rightarrow> \<tau>"
+ shows "atom x \<sharp> \<tau>"
+proof -
+ have "[] ; {||} \<turnstile>\<^sub>w\<^sub>f GNil" using wfG_nilI wfTh_emptyI by auto
+ hence "[] ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f \<tau>" using assms infer_l_wf by auto
+ thus ?thesis using fresh_def wfT_supp by force
+qed
+
+lemma subst_infer_v:
+ fixes v::v and v'::v
+ assumes "\<Theta> ; \<B> ; \<Gamma>'@((x,b\<^sub>1,c0[z0::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\<Gamma>\<Gamma>) \<turnstile> v' \<Rightarrow> \<tau>\<^sub>2" and
+ "\<Theta> ; \<B> ; \<Gamma> \<turnstile> v \<Rightarrow> \<tau>\<^sub>1" and
+ "\<Theta> ; \<B> ; \<Gamma> \<turnstile> \<tau>\<^sub>1 \<lesssim> (\<lbrace> z0 : b\<^sub>1 | c0 \<rbrace>)" and "atom z0 \<sharp> (x,v)"
+ shows "\<Theta> ; \<B> ; (\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v)@\<Gamma> \<turnstile> v'[x::=v]\<^sub>v\<^sub>v \<Rightarrow> \<tau>\<^sub>2[x::=v]\<^sub>\<tau>\<^sub>v"
+ using assms proof(nominal_induct "\<Gamma>'@((x,b\<^sub>1,c0[z0::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\<Gamma>\<Gamma>)" v' \<tau>\<^sub>2 avoiding: x v rule: infer_v.strong_induct)
+ case (infer_v_varI \<Theta> \<B> b c xa z)
+ have "\<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> [ xa ]\<^sup>v[x::=v]\<^sub>v\<^sub>v \<Rightarrow> \<lbrace> z : b | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ xa ]\<^sup>v[x::=v]\<^sub>v\<^sub>v ]\<^sup>c\<^sup>e \<rbrace>"
+ proof(cases "x=xa")
+ case True
+ have "\<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> v \<Rightarrow> \<lbrace> z : b | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e \<rbrace>"
+ proof(rule infer_v_g_weakening)
+ show *:\<open> \<Theta> ; \<B> ; \<Gamma> \<turnstile> v \<Rightarrow> \<lbrace> z : b | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e \<rbrace>\<close>
+ using infer_v_form infer_v_varI
+ by (metis True lookup_inside_unique_b lookup_inside_wf ms_fresh_all(32) subtype_eq_base type_e_eq)
+ show \<open>toSet \<Gamma> \<subseteq> toSet (\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>)\<close> by simp
+ have "\<Theta> ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b\<^sub>1" using infer_v_wf subtype_eq_base2 b_of.simps infer_v_varI by metis
+ thus \<open> \<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<close>
+ using wfG_subst[OF infer_v_varI(3), of \<Gamma>' x b\<^sub>1 "c0[z0::=[ x ]\<^sup>v]\<^sub>c\<^sub>v" \<Gamma> v] subst_g_inside infer_v_varI by metis
+ qed
+
+ thus ?thesis using subst_vv.simps True by simp
+ next
+ case False
+ then obtain c' where c: "Some (b, c') = lookup (\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>) xa" using lookup_subst2 infer_v_varI by metis
+ have "\<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> [ xa ]\<^sup>v \<Rightarrow> \<lbrace> z : b | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ xa ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrace>"
+ proof
+ have "\<Theta> ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b\<^sub>1" using infer_v_wf subtype_eq_base2 b_of.simps infer_v_varI by metis
+ thus "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>" using infer_v_varI
+ using wfG_subst[OF infer_v_varI(3), of \<Gamma>' x b\<^sub>1 "c0[z0::=[ x ]\<^sup>v]\<^sub>c\<^sub>v" \<Gamma> v] subst_g_inside infer_v_varI by metis
+ show "atom z \<sharp> xa" using infer_v_varI by auto
+ show "Some (b, c') = lookup (\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>) xa" using c by auto
+ show "atom z \<sharp> (\<Theta>, \<B>, \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>)" by (fresh_mth add: infer_v_varI fresh_subst_gv_inside)
+ qed
+ then show ?thesis using subst_vv.simps False by simp
+ qed
+ thus ?case using subst_t fresh_prodN infer_v_varI by metis
+next
+ case (infer_v_litI \<Theta> \<B> l \<tau>)
+ show ?case unfolding subst_vv.simps proof
+ show "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>" using wfG_subst_infer_v infer_v_litI subtype_eq_base2 b_of.simps by metis
+ have "atom x \<sharp> \<tau>" using infer_v_litI infer_l_fresh by metis
+ thus "\<turnstile> l \<Rightarrow> \<tau>[x::=v]\<^sub>\<tau>\<^sub>v" using infer_v_litI type_v_subst_fresh by simp
+ qed
+next
+ case (infer_v_pairI z v1 v2 \<Theta> \<B> t1 t2)
+ have " \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @
+ \<Gamma> \<turnstile> [ v1[x::=v]\<^sub>v\<^sub>v , v2[x::=v]\<^sub>v\<^sub>v ]\<^sup>v \<Rightarrow> \<lbrace> z : [ b_of t1[x::=v]\<^sub>\<tau>\<^sub>v , b_of
+ t2[x::=v]\<^sub>\<tau>\<^sub>v ]\<^sup>b | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ v1[x::=v]\<^sub>v\<^sub>v , v2[x::=v]\<^sub>v\<^sub>v ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrace>"
+ proof
+ show \<open>atom z \<sharp> (v1[x::=v]\<^sub>v\<^sub>v, v2[x::=v]\<^sub>v\<^sub>v)\<close> by (fresh_mth add: infer_v_pairI)
+ show \<open>atom z \<sharp> (\<Theta>, \<B>, \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>)\<close> by (fresh_mth add: infer_v_pairI fresh_subst_gv_inside)
+ show \<open> \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> v1[x::=v]\<^sub>v\<^sub>v \<Rightarrow> t1[x::=v]\<^sub>\<tau>\<^sub>v\<close> using infer_v_pairI by metis
+ show \<open> \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> v2[x::=v]\<^sub>v\<^sub>v \<Rightarrow> t2[x::=v]\<^sub>\<tau>\<^sub>v\<close> using infer_v_pairI by metis
+ qed
+ then show ?case using subst_vv.simps subst_tv.simps infer_v_pairI b_of_subst by simp
+next
+ case (infer_v_consI s dclist \<Theta> dc tc \<B> va tv z)
+
+ have " \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> (V_cons s dc va[x::=v]\<^sub>v\<^sub>v) \<Rightarrow> \<lbrace> z : B_id s | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ V_cons s dc va[x::=v]\<^sub>v\<^sub>v ]\<^sup>c\<^sup>e \<rbrace>"
+ proof
+ show td:\<open>AF_typedef s dclist \<in> set \<Theta>\<close> using infer_v_consI by auto
+ show dc:\<open>(dc, tc) \<in> set dclist\<close> using infer_v_consI by auto
+ show \<open> \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> va[x::=v]\<^sub>v\<^sub>v \<Rightarrow> tv[x::=v]\<^sub>\<tau>\<^sub>v\<close> using infer_v_consI by auto
+ have \<open>\<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> tv[x::=v]\<^sub>\<tau>\<^sub>v \<lesssim> tc[x::=v]\<^sub>\<tau>\<^sub>v\<close>
+ using subst_subtype_tau infer_v_consI by metis
+ moreover have "atom x \<sharp> tc" using wfTh_lookup_supp_empty[OF td dc] infer_v_wf infer_v_consI fresh_def by fast
+ ultimately show \<open>\<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> tv[x::=v]\<^sub>\<tau>\<^sub>v \<lesssim> tc\<close> by simp
+ show \<open>atom z \<sharp> va[x::=v]\<^sub>v\<^sub>v\<close> using infer_v_consI by auto
+ show \<open>atom z \<sharp> (\<Theta>, \<B>, \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>)\<close> by (fresh_mth add: infer_v_consI fresh_subst_gv_inside)
+ qed
+ thus ?case using subst_vv.simps subst_t[of z x v ] infer_v_consI by metis
+
+next
+ case (infer_v_conspI s bv dclist \<Theta> dc tc \<B> va tv b z)
+ have "\<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> (V_consp s dc b va[x::=v]\<^sub>v\<^sub>v) \<Rightarrow> \<lbrace> z : B_app s b | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ V_consp s dc b va[x::=v]\<^sub>v\<^sub>v ]\<^sup>c\<^sup>e \<rbrace>"
+ proof
+ show td:\<open>AF_typedef_poly s bv dclist \<in> set \<Theta>\<close> using infer_v_conspI by auto
+ show dc:\<open>(dc, tc) \<in> set dclist\<close> using infer_v_conspI by auto
+ show \<open> \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> va[x::=v]\<^sub>v\<^sub>v \<Rightarrow> tv[x::=v]\<^sub>\<tau>\<^sub>v\<close> using infer_v_conspI by metis
+ have \<open>\<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> tv[x::=v]\<^sub>\<tau>\<^sub>v \<lesssim> tc[bv::=b]\<^sub>\<tau>\<^sub>b[x::=v]\<^sub>\<tau>\<^sub>v\<close>
+ using subst_subtype_tau infer_v_conspI by metis
+ moreover have "atom x \<sharp> tc[bv::=b]\<^sub>\<tau>\<^sub>b" proof -
+ have "supp tc \<subseteq> { atom bv }" using wfTh_poly_lookup_supp infer_v_conspI wfX_wfY by metis
+ hence "atom x \<sharp> tc" using x_not_in_b_set
+ using fresh_def by fastforce
+ moreover have "atom x \<sharp> b" using x_fresh_b by auto
+ ultimately show ?thesis using fresh_subst_if subst_b_\<tau>_def by metis
+ qed
+ ultimately show \<open>\<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> tv[x::=v]\<^sub>\<tau>\<^sub>v \<lesssim> tc[bv::=b]\<^sub>\<tau>\<^sub>b\<close> by simp
+ show \<open>atom z \<sharp> (\<Theta>, \<B>, \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>, va[x::=v]\<^sub>v\<^sub>v, b)\<close> proof -
+ have "atom z \<sharp> va[x::=v]\<^sub>v\<^sub>v" using fresh_subst_v_if infer_v_conspI subst_v_v_def by metis
+ moreover have "atom z \<sharp> \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>" using fresh_subst_gv_inside infer_v_conspI by metis
+ ultimately show ?thesis using fresh_prodN infer_v_conspI by metis
+ qed
+ show \<open>atom bv \<sharp> (\<Theta>, \<B>, \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>, va[x::=v]\<^sub>v\<^sub>v, b)\<close> proof -
+ have "atom bv \<sharp> va[x::=v]\<^sub>v\<^sub>v" using fresh_subst_v_if infer_v_conspI subst_v_v_def by metis
+ moreover have "atom bv \<sharp> \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>" using fresh_subst_gv_inside infer_v_conspI by metis
+ ultimately show ?thesis using fresh_prodN infer_v_conspI by metis
+ qed
+ show "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f b" using infer_v_conspI by auto
+ qed
+ thus ?case using subst_vv.simps subst_t[of z x v ] infer_v_conspI by metis
+
+qed
+
+lemma subst_infer_check_v:
+ fixes v::v and v'::v
+ assumes "\<Theta> ; \<B> ; \<Gamma> \<turnstile> v \<Rightarrow> \<tau>\<^sub>1" and
+ "check_v \<Theta> \<B> (\<Gamma>'@((x,b\<^sub>1,c0[z0::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\<Gamma>\<Gamma>)) v' \<tau>\<^sub>2" and
+ "\<Theta> ; \<B> ; \<Gamma> \<turnstile> \<tau>\<^sub>1 \<lesssim> \<lbrace> z0 : b\<^sub>1 | c0 \<rbrace>" and "atom z0 \<sharp> (x,v)"
+ shows "check_v \<Theta> \<B> ((\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v)@\<Gamma>) (v'[x::=v]\<^sub>v\<^sub>v) (\<tau>\<^sub>2[x::=v]\<^sub>\<tau>\<^sub>v)"
+proof -
+ obtain \<tau>\<^sub>2' where t2: "infer_v \<Theta> \<B> (\<Gamma>' @ (x, b\<^sub>1, c0[z0::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) v' \<tau>\<^sub>2' \<and> \<Theta> ; \<B> ; (\<Gamma>' @ (x, b\<^sub>1, c0[z0::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) \<turnstile> \<tau>\<^sub>2' \<lesssim> \<tau>\<^sub>2"
+ using check_v_elims assms by blast
+ hence "infer_v \<Theta> \<B> ((\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v)@\<Gamma>) (v'[x::=v]\<^sub>v\<^sub>v) (\<tau>\<^sub>2'[x::=v]\<^sub>\<tau>\<^sub>v)"
+ using subst_infer_v[OF _ assms(1) assms(3) assms(4)] by blast
+ moreover hence "\<Theta>; \<B> ; ((\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v)@\<Gamma>) \<turnstile>\<tau>\<^sub>2'[x::=v]\<^sub>\<tau>\<^sub>v \<lesssim> \<tau>\<^sub>2[x::=v]\<^sub>\<tau>\<^sub>v"
+ using subst_subtype assms t2 by (meson subst_subtype_tau subtype_trans)
+ ultimately show ?thesis using check_v.intros by blast
+qed
+
+lemma type_veq_subst[simp]:
+ assumes "atom z \<sharp> (x,v)"
+ shows "\<lbrace> z : b | CE_val (V_var z) == CE_val v' \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v = \<lbrace> z : b | CE_val (V_var z) == CE_val v'[x::=v]\<^sub>v\<^sub>v \<rbrace>"
+ using assms by auto
+
+lemma subst_infer_v_form:
+ fixes v::v and v'::v and \<Gamma>::\<Gamma>
+ assumes "\<Theta> ; \<B> ; \<Gamma> \<turnstile> v \<Rightarrow> \<tau>\<^sub>1" and
+ "\<Theta> ; \<B> ; \<Gamma>'@((x,b\<^sub>1,c0[z0::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\<Gamma>\<Gamma>) \<turnstile> v' \<Rightarrow> \<tau>\<^sub>2" and "b= b_of \<tau>\<^sub>2"
+ "\<Theta> ; \<B> ; \<Gamma> \<turnstile> \<tau>\<^sub>1 \<lesssim> (\<lbrace> z0 : b\<^sub>1 | c0 \<rbrace>)" and "atom z0 \<sharp> (x,v)" and "atom z3' \<sharp> (x,v,v', \<Gamma>'@((x,b\<^sub>1,c0[z0::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\<Gamma>\<Gamma>) )"
+ shows \<open> \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> v'[x::=v]\<^sub>v\<^sub>v \<Rightarrow> \<lbrace> z3' : b | CE_val (V_var z3') == CE_val v'[x::=v]\<^sub>v\<^sub>v \<rbrace>\<close>
+proof -
+ have "\<Theta> ; \<B> ; \<Gamma>'@((x,b\<^sub>1,c0[z0::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\<Gamma>\<Gamma>) \<turnstile> v' \<Rightarrow> \<lbrace> z3' : b_of \<tau>\<^sub>2 | C_eq (CE_val (V_var z3')) (CE_val v') \<rbrace>"
+ proof(rule infer_v_form4)
+ show \<open> \<Theta> ; \<B> ; \<Gamma>' @ (x, b\<^sub>1, c0[z0::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<turnstile> v' \<Rightarrow> \<tau>\<^sub>2\<close> using assms by metis
+ show \<open>atom z3' \<sharp> (v', \<Gamma>' @ (x, b\<^sub>1, c0[z0::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)\<close> using assms fresh_prodN by metis
+ show \<open>b_of \<tau>\<^sub>2 = b_of \<tau>\<^sub>2\<close> by auto
+ qed
+ hence \<open> \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> v'[x::=v]\<^sub>v\<^sub>v \<Rightarrow> \<lbrace> z3' : b_of \<tau>\<^sub>2 | CE_val (V_var z3') == CE_val v' \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v\<close>
+ using subst_infer_v assms by metis
+ thus ?thesis using type_veq_subst fresh_prodN assms by metis
+qed
+
+section \<open>Expressions\<close>
+
+text \<open>
+ For operator, fst and snd cases, we use elimination to get one or more values, apply the substitution lemma for values. The types always have
+the same form and are equal under substitution.
+ For function application, the subst value is a subtype of the value which is a subtype of the argument. The return of the function is the same
+ under substitution.
+\<close>
+
+text \<open> Observe a similar pattern for each case \<close>
+
+lemma subst_infer_e:
+ fixes v::v and e::e and \<Gamma>'::\<Gamma>
+ assumes
+ "\<Theta> ; \<Phi> ; \<B> ; G ; \<Delta> \<turnstile> e \<Rightarrow> \<tau>\<^sub>2" and "G = (\<Gamma>'@((x,b\<^sub>1,subst_cv c0 z0 (V_var x))#\<^sub>\<Gamma>\<Gamma>))"
+ "\<Theta> ; \<B> ; \<Gamma> \<turnstile> v \<Rightarrow> \<tau>\<^sub>1" and
+ "\<Theta>; \<B> ; \<Gamma> \<turnstile> \<tau>\<^sub>1 \<lesssim> \<lbrace> z0 : b\<^sub>1 | c0 \<rbrace>" and "atom z0 \<sharp> (x,v)"
+ shows "\<Theta> ; \<Phi> ; \<B> ; ((\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v)@\<Gamma>) ; (\<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v) \<turnstile> (subst_ev e x v ) \<Rightarrow> \<tau>\<^sub>2[x::=v]\<^sub>\<tau>\<^sub>v"
+ using assms proof(nominal_induct avoiding: x v rule: infer_e.strong_induct)
+ case (infer_e_valI \<Theta> \<B> \<Gamma>'' \<Delta> \<Phi> v' \<tau>)
+
+ have "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> ; \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> (AE_val (v'[x::=v]\<^sub>v\<^sub>v)) \<Rightarrow> \<tau>[x::=v]\<^sub>\<tau>\<^sub>v"
+ proof
+ show "\<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v" using wfD_subst infer_e_valI subtype_eq_base2
+ by (metis b_of.simps infer_v_v_wf subst_g_inside_simple wfD_wf wf_subst(11))
+ show "\<Theta>\<turnstile>\<^sub>w\<^sub>f \<Phi>" using infer_e_valI by auto
+ show "\<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> v'[x::=v]\<^sub>v\<^sub>v \<Rightarrow> \<tau>[x::=v]\<^sub>\<tau>\<^sub>v" using subst_infer_v infer_e_valI using wfD_subst infer_e_valI subtype_eq_base2
+ by metis
+ qed
+ thus ?case using subst_ev.simps by simp
+next
+ case (infer_e_plusI \<Theta> \<B> \<Gamma>'' \<Delta> \<Phi> v1 z1 c1 v2 z2 c2 z3)
+
+ hence z3f: "atom z3 \<sharp> CE_op Plus [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e" using e.fresh ce.fresh opp.fresh by metis
+
+ obtain z3'::x where *:"atom z3' \<sharp> (x,v,AE_op Plus v1 v2, CE_op Plus [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e , AE_op Plus v1[x::=v]\<^sub>v\<^sub>v v2[x::=v]\<^sub>v\<^sub>v , CE_op Plus [v1[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e [v2[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e,\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> )"
+ using obtain_fresh by metis
+ hence **:"(\<lbrace> z3 : B_int | CE_val (V_var z3) == CE_op Plus [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \<rbrace>) = \<lbrace> z3' : B_int | CE_val (V_var z3') == CE_op Plus [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \<rbrace>"
+ using type_e_eq infer_e_plusI fresh_Pair z3f by metis
+
+ obtain z1' b1' c1' where z1:"atom z1' \<sharp> (x,v) \<and> \<lbrace> z1 : B_int | c1 \<rbrace> = \<lbrace> z1' : b1' | c1' \<rbrace>" using obtain_fresh_z by metis
+ obtain z2' b2' c2' where z2:"atom z2' \<sharp> (x,v) \<and> \<lbrace> z2 : B_int | c2 \<rbrace> = \<lbrace> z2' : b2' | c2' \<rbrace>" using obtain_fresh_z by metis
+
+ have bb:"b1' = B_int \<and> b2' = B_int" using z1 z2 \<tau>.eq_iff by metis
+
+ have "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> ; \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> (AE_op Plus (v1[x::=v]\<^sub>v\<^sub>v) (v2[x::=v]\<^sub>v\<^sub>v)) \<Rightarrow> \<lbrace> z3' : B_int | CE_val (V_var z3') == CE_op Plus ([v1[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) ([v2[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) \<rbrace>"
+ proof
+ show \<open> \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<close>
+ using infer_e_plusI wfD_subst subtype_eq_base2 b_of.simps by metis
+ show \<open> \<Theta>\<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using infer_e_plusI by blast
+ show \<open> \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> v1[x::=v]\<^sub>v\<^sub>v \<Rightarrow> \<lbrace> z1' : B_int | c1'[x::=v]\<^sub>c\<^sub>v \<rbrace>\<close> using subst_tv.simps subst_infer_v infer_e_plusI z1 bb by metis
+ show \<open> \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> v2[x::=v]\<^sub>v\<^sub>v \<Rightarrow> \<lbrace> z2' : B_int | c2'[x::=v]\<^sub>c\<^sub>v \<rbrace>\<close> using subst_tv.simps subst_infer_v infer_e_plusI z2 bb by metis
+ show \<open>atom z3' \<sharp> AE_op Plus v1[x::=v]\<^sub>v\<^sub>v v2[x::=v]\<^sub>v\<^sub>v\<close> using fresh_prod6 * by metis
+ show \<open>atom z3' \<sharp> \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>\<close> using * by auto
+ qed
+ moreover have "\<lbrace> z3' : B_int | CE_val (V_var z3') == CE_op Plus ([v1[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) ([v2[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) \<rbrace> = \<lbrace> z3' : B_int | CE_val (V_var z3') == CE_op Plus [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v"
+ by(subst subst_tv.simps,auto simp add: * )
+ ultimately show ?case using subst_ev.simps * ** by metis
+next
+ case (infer_e_leqI \<Theta> \<B> \<Gamma>'' \<Delta> \<Phi> v1 z1 c1 v2 z2 c2 z3)
+
+ hence z3f: "atom z3 \<sharp> CE_op LEq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e" using e.fresh ce.fresh opp.fresh by metis
+
+ obtain z3'::x where *:"atom z3' \<sharp> (x,v,AE_op LEq v1 v2, CE_op LEq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e , CE_op LEq [v1[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e [v2[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e , AE_op LEq v1[x::=v]\<^sub>v\<^sub>v v2[x::=v]\<^sub>v\<^sub>v,\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> )"
+ using obtain_fresh by metis
+ hence **:"(\<lbrace> z3 : B_bool | CE_val (V_var z3) == CE_op LEq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \<rbrace>) = \<lbrace> z3' : B_bool | CE_val (V_var z3') == CE_op LEq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \<rbrace>"
+ using type_e_eq infer_e_leqI fresh_Pair z3f by metis
+
+ obtain z1' b1' c1' where z1:"atom z1' \<sharp> (x,v) \<and> \<lbrace> z1 : B_int | c1 \<rbrace> = \<lbrace> z1' : b1' | c1' \<rbrace>" using obtain_fresh_z by metis
+ obtain z2' b2' c2' where z2:"atom z2' \<sharp> (x,v) \<and> \<lbrace> z2 : B_int | c2 \<rbrace> = \<lbrace> z2' : b2' | c2' \<rbrace>" using obtain_fresh_z by metis
+
+ have bb:"b1' = B_int \<and> b2' = B_int" using z1 z2 \<tau>.eq_iff by metis
+
+ have "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> ; \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> (AE_op LEq (v1[x::=v]\<^sub>v\<^sub>v) (v2[x::=v]\<^sub>v\<^sub>v)) \<Rightarrow> \<lbrace> z3' : B_bool | CE_val (V_var z3') == CE_op LEq ([v1[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) ([v2[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) \<rbrace>"
+ proof
+ show \<open> \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<close> using wfD_subst infer_e_leqI subtype_eq_base2 b_of.simps by metis
+ show \<open> \<Theta>\<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using infer_e_leqI(2) by auto
+ show \<open> \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> v1[x::=v]\<^sub>v\<^sub>v \<Rightarrow> \<lbrace> z1' : B_int | c1'[x::=v]\<^sub>c\<^sub>v \<rbrace>\<close> using subst_tv.simps subst_infer_v infer_e_leqI z1 bb by metis
+ show \<open> \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> v2[x::=v]\<^sub>v\<^sub>v \<Rightarrow> \<lbrace> z2' : B_int | c2'[x::=v]\<^sub>c\<^sub>v \<rbrace>\<close> using subst_tv.simps subst_infer_v infer_e_leqI z2 bb by metis
+ show \<open>atom z3' \<sharp> AE_op LEq v1[x::=v]\<^sub>v\<^sub>v v2[x::=v]\<^sub>v\<^sub>v\<close> using fresh_Pair * by metis
+ show \<open>atom z3' \<sharp> \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>\<close> using * by auto
+ qed
+ moreover have "\<lbrace> z3' : B_bool | CE_val (V_var z3') == CE_op LEq ([v1[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) ([v2[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) \<rbrace> = \<lbrace> z3' : B_bool | CE_val (V_var z3') == CE_op LEq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v"
+ using subst_tv.simps subst_ev.simps * by auto
+ ultimately show ?case using subst_ev.simps * ** by metis
+next
+ case (infer_e_eqI \<Theta> \<B> \<Gamma>'' \<Delta> \<Phi> v1 z1 bb c1 v2 z2 c2 z3)
+
+ hence z3f: "atom z3 \<sharp> CE_op Eq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e" using e.fresh ce.fresh opp.fresh by metis
+
+ obtain z3'::x where *:"atom z3' \<sharp> (x,v,AE_op Eq v1 v2, CE_op Eq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e , CE_op Eq [v1[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e [v2[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e , AE_op Eq v1[x::=v]\<^sub>v\<^sub>v v2[x::=v]\<^sub>v\<^sub>v,\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> )"
+ using obtain_fresh by metis
+ hence **:"(\<lbrace> z3 : B_bool | CE_val (V_var z3) == CE_op Eq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \<rbrace>) = \<lbrace> z3' : B_bool | CE_val (V_var z3') == CE_op Eq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \<rbrace>"
+ using type_e_eq infer_e_eqI fresh_Pair z3f by metis
+
+ obtain z1' b1' c1' where z1:"atom z1' \<sharp> (x,v) \<and> \<lbrace> z1 : bb | c1 \<rbrace> = \<lbrace> z1' : b1' | c1' \<rbrace>" using obtain_fresh_z by metis
+ obtain z2' b2' c2' where z2:"atom z2' \<sharp> (x,v) \<and> \<lbrace> z2 : bb | c2 \<rbrace> = \<lbrace> z2' : b2' | c2' \<rbrace>" using obtain_fresh_z by metis
+
+ have bb:"b1' = bb \<and> b2' = bb" using z1 z2 \<tau>.eq_iff by metis
+
+ have "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> ; \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> (AE_op Eq (v1[x::=v]\<^sub>v\<^sub>v) (v2[x::=v]\<^sub>v\<^sub>v)) \<Rightarrow> \<lbrace> z3' : B_bool | CE_val (V_var z3') == CE_op Eq ([v1[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) ([v2[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) \<rbrace>"
+ proof
+ show \<open> \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<close> using wfD_subst infer_e_eqI subtype_eq_base2 b_of.simps by metis
+ show \<open> \<Theta>\<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using infer_e_eqI(2) by auto
+ show \<open> \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> v1[x::=v]\<^sub>v\<^sub>v \<Rightarrow> \<lbrace> z1' : bb | c1'[x::=v]\<^sub>c\<^sub>v \<rbrace>\<close> using subst_tv.simps subst_infer_v infer_e_eqI z1 bb by metis
+ show \<open> \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> v2[x::=v]\<^sub>v\<^sub>v \<Rightarrow> \<lbrace> z2' : bb | c2'[x::=v]\<^sub>c\<^sub>v \<rbrace>\<close> using subst_tv.simps subst_infer_v infer_e_eqI z2 bb by metis
+ show \<open>atom z3' \<sharp> AE_op Eq v1[x::=v]\<^sub>v\<^sub>v v2[x::=v]\<^sub>v\<^sub>v\<close> using fresh_Pair * by metis
+ show \<open>atom z3' \<sharp> \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>\<close> using * by auto
+ show "bb \<in> {B_bool, B_int, B_unit}" using infer_e_eqI by auto
+ qed
+ moreover have "\<lbrace> z3' : B_bool | CE_val (V_var z3') == CE_op Eq ([v1[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) ([v2[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) \<rbrace> = \<lbrace> z3' : B_bool | CE_val (V_var z3') == CE_op Eq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v"
+ using subst_tv.simps subst_ev.simps * by auto
+ ultimately show ?case using subst_ev.simps * ** by metis
+next
+ case (infer_e_appI \<Theta> \<B> \<Gamma>'' \<Delta> \<Phi> f x' b c \<tau>' s' v' \<tau>)
+
+ hence "x \<noteq> x'" using \<open>atom x' \<sharp> \<Gamma>''\<close> using wfG_inside_x_neq wfX_wfY by metis
+
+ show ?case proof(subst subst_ev.simps,rule)
+ show \<open> \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<close> using infer_e_appI wfD_subst subtype_eq_base2 b_of.simps by metis
+ show \<open> \<Theta>\<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using infer_e_appI by metis
+ show \<open>Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x' b c \<tau>' s'))) = lookup_fun \<Phi> f\<close> using infer_e_appI by metis
+
+ have \<open>\<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> v'[x::=v]\<^sub>v\<^sub>v \<Leftarrow> \<lbrace> x' : b | c \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v\<close> proof(rule subst_infer_check_v )
+ show "\<Theta> ; \<B> ; \<Gamma> \<turnstile> v \<Rightarrow> \<tau>\<^sub>1" using infer_e_appI by metis
+ show "\<Theta> ; \<B> ; \<Gamma>' @ (x, b\<^sub>1, c0[z0::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<turnstile> v' \<Leftarrow> \<lbrace> x' : b | c \<rbrace>" using infer_e_appI by metis
+ show "\<Theta> ; \<B> ; \<Gamma> \<turnstile> \<tau>\<^sub>1 \<lesssim> \<lbrace> z0 : b\<^sub>1 | c0 \<rbrace>" using infer_e_appI by metis
+ show "atom z0 \<sharp> (x, v)" using infer_e_appI by metis
+ qed
+ moreover have "atom x \<sharp> c" using wfPhi_f_simple_supp_c infer_e_appI fresh_def \<open>x\<noteq>x'\<close>
+ atom_eq_iff empty_iff infer_e_appI.hyps insert_iff subset_singletonD by metis
+
+ moreover hence "atom x \<sharp> \<lbrace> x' : b | c \<rbrace>" using \<tau>.fresh supp_b_empty fresh_def by blast
+ ultimately show \<open>\<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> v'[x::=v]\<^sub>v\<^sub>v \<Leftarrow> \<lbrace> x' : b | c \<rbrace>\<close> using forget_subst_tv by metis
+
+ have *: "atom x' \<sharp> (x,v)" using infer_v_fresh_g_fresh_xv infer_e_appI check_v_wf by blast
+
+ show \<open>atom x' \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>, \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v, v'[x::=v]\<^sub>v\<^sub>v, \<tau>[x::=v]\<^sub>\<tau>\<^sub>v)\<close>
+ apply(unfold fresh_prodN, intro conjI)
+ apply (fresh_subst_mth_aux add: infer_e_appI fresh_subst_gv wfD_wf subst_g_inside)
+ using infer_e_appI fresh_subst_gv wfD_wf subst_g_inside apply metis
+ using infer_e_appI fresh_subst_dv_if apply metis
+ done
+
+ have "supp \<tau>' \<subseteq> { atom x' } \<union> supp \<B>" using infer_e_appI wfT_supp wfPhi_f_simple_wfT
+ by (meson infer_e_appI.hyps(2) le_supI1 wfPhi_f_simple_supp_t)
+ hence "atom x \<sharp> \<tau>'" using \<open>x\<noteq>x'\<close> fresh_def supp_at_base x_not_in_b_set by fastforce
+ thus \<open>\<tau>'[x'::=v'[x::=v]\<^sub>v\<^sub>v]\<^sub>v = \<tau>[x::=v]\<^sub>\<tau>\<^sub>v\<close> using subst_tv_commute infer_e_appI subst_defs by metis
+ qed
+next
+ case (infer_e_appPI \<Theta> \<B> \<Gamma>'' \<Delta> \<Phi> b' f bv x' b c \<tau>' s' v' \<tau>)
+
+ hence "x \<noteq> x'" using \<open>atom x' \<sharp> \<Gamma>''\<close> using wfG_inside_x_neq wfX_wfY by metis
+
+ show ?case proof(subst subst_ev.simps,rule)
+ show \<open> \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<close> using infer_e_appPI wfD_subst subtype_eq_base2 b_of.simps by metis
+ show \<open> \<Theta>\<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using infer_e_appPI(4) by auto
+ show "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f b'" using infer_e_appPI(5) by auto
+ show "Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x' b c \<tau>' s'))) = lookup_fun \<Phi> f" using infer_e_appPI(6) by auto
+
+ show "\<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> v'[x::=v]\<^sub>v\<^sub>v \<Leftarrow> \<lbrace> x' : b[bv::=b']\<^sub>b | c[bv::=b']\<^sub>b \<rbrace>" proof -
+ have \<open>\<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> v'[x::=v]\<^sub>v\<^sub>v \<Leftarrow> \<lbrace> x' : b[bv::=b']\<^sub>b\<^sub>b | c[bv::=b']\<^sub>c\<^sub>b \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v\<close> proof(rule subst_infer_check_v )
+ show "\<Theta> ; \<B> ; \<Gamma> \<turnstile> v \<Rightarrow> \<tau>\<^sub>1" using infer_e_appPI by metis
+ show "\<Theta> ; \<B> ; \<Gamma>' @ (x, b\<^sub>1, c0[z0::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<turnstile> v' \<Leftarrow> \<lbrace> x' : b[bv::=b']\<^sub>b\<^sub>b | c[bv::=b']\<^sub>c\<^sub>b \<rbrace>" using infer_e_appPI subst_defs by metis
+ show "\<Theta> ; \<B> ; \<Gamma> \<turnstile> \<tau>\<^sub>1 \<lesssim> \<lbrace> z0 : b\<^sub>1 | c0 \<rbrace>" using infer_e_appPI by metis
+ show "atom z0 \<sharp> (x, v)" using infer_e_appPI by metis
+ qed
+ moreover have "atom x \<sharp> c" proof -
+ have "supp c \<subseteq> {atom x', atom bv}" using wfPhi_f_poly_supp_c[OF infer_e_appPI(6)] infer_e_appPI by metis
+ thus ?thesis unfolding fresh_def using \<open>x\<noteq>x'\<close> atom_eq_iff by auto
+ qed
+ moreover hence "atom x \<sharp> \<lbrace> x' : b[bv::=b']\<^sub>b\<^sub>b | c[bv::=b']\<^sub>c\<^sub>b \<rbrace>" using \<tau>.fresh supp_b_empty fresh_def subst_b_fresh_x
+ by (metis subst_b_c_def)
+ ultimately show ?thesis using forget_subst_tv subst_defs by metis
+ qed
+ have "supp \<tau>' \<subseteq> { atom x', atom bv }" using wfPhi_f_poly_supp_t infer_e_appPI by metis
+ hence "atom x \<sharp> \<tau>'" using fresh_def \<open>x\<noteq>x'\<close> by force
+ hence *:"atom x \<sharp> \<tau>'[bv::=b']\<^sub>\<tau>\<^sub>b" using subst_b_fresh_x subst_b_\<tau>_def by metis
+ have "atom x' \<sharp> (x,v)" using infer_v_fresh_g_fresh_xv infer_e_appPI check_v_wf by blast
+ thus "atom x' \<sharp> \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>" using infer_e_appPI fresh_subst_gv wfD_wf subst_g_inside fresh_Pair by metis
+ show "\<tau>'[bv::=b']\<^sub>b[x'::=v'[x::=v]\<^sub>v\<^sub>v]\<^sub>v = \<tau>[x::=v]\<^sub>\<tau>\<^sub>v" using infer_e_appPI subst_tv_commute[OF * ] subst_defs by metis
+ show "atom bv \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>, \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v, b', v'[x::=v]\<^sub>v\<^sub>v, \<tau>[x::=v]\<^sub>\<tau>\<^sub>v)"
+ by (fresh_mth add: infer_e_appPI fresh_subst_gv_inside)
+ qed
+next
+ case (infer_e_fstI \<Theta> \<B> \<Gamma>'' \<Delta> \<Phi> v' z' b1 b2 c z)
+
+ hence zf: "atom z \<sharp> CE_fst [v']\<^sup>c\<^sup>e" using ce.fresh e.fresh opp.fresh by metis
+
+ obtain z3'::x where *:"atom z3' \<sharp> (x,v,AE_fst v', CE_fst [v']\<^sup>c\<^sup>e , AE_fst v'[x::=v]\<^sub>v\<^sub>v ,\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> )" using obtain_fresh by auto
+ hence **:"(\<lbrace> z : b1 | CE_val (V_var z) == CE_fst [v']\<^sup>c\<^sup>e \<rbrace>) = \<lbrace> z3' : b1 | CE_val (V_var z3') == CE_fst [v']\<^sup>c\<^sup>e \<rbrace>"
+ using type_e_eq infer_e_fstI(4) fresh_Pair zf by metis
+
+ obtain z1' b1' c1' where z1:"atom z1' \<sharp> (x,v) \<and> \<lbrace> z' : B_pair b1 b2 | c \<rbrace> = \<lbrace> z1' : b1' | c1' \<rbrace>" using obtain_fresh_z by metis
+
+ have bb:"b1' = B_pair b1 b2" using z1 \<tau>.eq_iff by metis
+ have "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> ; \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> (AE_fst v'[x::=v]\<^sub>v\<^sub>v) \<Rightarrow> \<lbrace> z3' : b1 | CE_val (V_var z3') == CE_fst [v'[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e \<rbrace>"
+ proof
+ show \<open> \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<close> using wfD_subst infer_e_fstI subtype_eq_base2 b_of.simps by metis
+ show \<open> \<Theta>\<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using infer_e_fstI by metis
+ show \<open> \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> v'[x::=v]\<^sub>v\<^sub>v \<Rightarrow> \<lbrace> z1' : B_pair b1 b2 | c1'[x::=v]\<^sub>c\<^sub>v \<rbrace>\<close> using subst_tv.simps subst_infer_v infer_e_fstI z1 bb by metis
+
+ show \<open>atom z3' \<sharp> AE_fst v'[x::=v]\<^sub>v\<^sub>v \<close> using fresh_Pair * by metis
+ show \<open>atom z3' \<sharp> \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>\<close> using * by auto
+ qed
+ moreover have "\<lbrace> z3' : b1 | CE_val (V_var z3') == CE_fst [v'[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e \<rbrace> = \<lbrace> z3' : b1 | CE_val (V_var z3') == CE_fst [v']\<^sup>c\<^sup>e \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v"
+ using subst_tv.simps subst_ev.simps * by auto
+ ultimately show ?case using subst_ev.simps * ** by metis
+next
+ case (infer_e_sndI \<Theta> \<B> \<Gamma>'' \<Delta> \<Phi> v' z' b1 b2 c z)
+ hence zf: "atom z \<sharp> CE_snd [v']\<^sup>c\<^sup>e" using ce.fresh e.fresh opp.fresh by metis
+
+ obtain z3'::x where *:"atom z3' \<sharp> (x,v,AE_snd v', CE_snd [v']\<^sup>c\<^sup>e , AE_snd v'[x::=v]\<^sub>v\<^sub>v ,\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> ,v', \<Gamma>'')" using obtain_fresh by auto
+ hence **:"(\<lbrace> z : b2 | CE_val (V_var z) == CE_snd [v']\<^sup>c\<^sup>e \<rbrace>) = \<lbrace> z3' : b2 | CE_val (V_var z3') == CE_snd [v']\<^sup>c\<^sup>e \<rbrace>"
+ using type_e_eq infer_e_sndI(4) fresh_Pair zf by metis
+
+ obtain z1' b2' c1' where z1:"atom z1' \<sharp> (x,v) \<and> \<lbrace> z' : B_pair b1 b2 | c \<rbrace> = \<lbrace> z1' : b2' | c1' \<rbrace>" using obtain_fresh_z by metis
+
+ have bb:"b2' = B_pair b1 b2" using z1 \<tau>.eq_iff by metis
+
+ have "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> ; \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> (AE_snd (v'[x::=v]\<^sub>v\<^sub>v)) \<Rightarrow> \<lbrace> z3' : b2 | CE_val (V_var z3') == CE_snd ([v'[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) \<rbrace>"
+ proof
+ show \<open> \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<close> using wfD_subst infer_e_sndI subtype_eq_base2 b_of.simps by metis
+ show \<open> \<Theta>\<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using infer_e_sndI by metis
+ show \<open> \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> v'[x::=v]\<^sub>v\<^sub>v \<Rightarrow> \<lbrace> z1' : B_pair b1 b2 | c1'[x::=v]\<^sub>c\<^sub>v \<rbrace>\<close> using subst_tv.simps subst_infer_v infer_e_sndI z1 bb by metis
+
+ show \<open>atom z3' \<sharp> AE_snd v'[x::=v]\<^sub>v\<^sub>v \<close> using fresh_Pair * by metis
+ show \<open>atom z3' \<sharp> \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>\<close> using * by auto
+ qed
+ moreover have "\<lbrace> z3' : b2 | CE_val (V_var z3') == CE_snd ([v'[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) \<rbrace> = \<lbrace> z3' : b2 | CE_val (V_var z3') == CE_snd [v']\<^sup>c\<^sup>e \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v"
+ by(subst subst_tv.simps, auto simp add: fresh_prodN *)
+ ultimately show ?case using subst_ev.simps * ** by metis
+next
+ case (infer_e_lenI \<Theta> \<B> \<Gamma>'' \<Delta> \<Phi> v' z' c z)
+ hence zf: "atom z \<sharp> CE_len [v']\<^sup>c\<^sup>e" using ce.fresh e.fresh opp.fresh by metis
+ obtain z3'::x where *:"atom z3' \<sharp> (x,v,AE_len v', CE_len [v']\<^sup>c\<^sup>e , AE_len v'[x::=v]\<^sub>v\<^sub>v ,\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> , \<Gamma>'',v')" using obtain_fresh by auto
+ hence **:"(\<lbrace> z : B_int | CE_val (V_var z) == CE_len [v']\<^sup>c\<^sup>e \<rbrace>) = \<lbrace> z3' : B_int | CE_val (V_var z3') == CE_len [v']\<^sup>c\<^sup>e \<rbrace>"
+ using type_e_eq infer_e_lenI fresh_Pair zf by metis
+
+ have ***: "\<Theta> ; \<B> ; \<Gamma>'' \<turnstile> v' \<Rightarrow> \<lbrace> z3' : B_bitvec | CE_val (V_var z3') == CE_val v' \<rbrace>"
+ using infer_e_lenI infer_v_form3[OF infer_e_lenI(3), of z3'] b_of.simps * fresh_Pair by metis
+
+ have "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> ; \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> (AE_len (v'[x::=v]\<^sub>v\<^sub>v)) \<Rightarrow> \<lbrace> z3' : B_int | CE_val (V_var z3') == CE_len ([v'[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) \<rbrace>"
+ proof
+ show \<open> \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<close> using wfD_subst infer_e_lenI subtype_eq_base2 b_of.simps by metis
+ show \<open> \<Theta>\<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using infer_e_lenI by metis
+
+ have \<open> \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> v'[x::=v]\<^sub>v\<^sub>v \<Rightarrow> \<lbrace> z3' : B_bitvec | CE_val (V_var z3') == CE_val v' \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v\<close>
+ proof(rule subst_infer_v)
+ show \<open> \<Theta> ; \<B> ; \<Gamma> \<turnstile> v \<Rightarrow> \<tau>\<^sub>1\<close> using infer_e_lenI by metis
+ show \<open> \<Theta> ; \<B> ; \<Gamma>' @ (x, b\<^sub>1, c0[z0::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<turnstile> v' \<Rightarrow> \<lbrace> z3' : B_bitvec | [ [ z3' ]\<^sup>v ]\<^sup>c\<^sup>e == [ v' ]\<^sup>c\<^sup>e \<rbrace>\<close> using *** infer_e_lenI by metis
+ show "\<Theta> ; \<B> ; \<Gamma> \<turnstile> \<tau>\<^sub>1 \<lesssim> \<lbrace> z0 : b\<^sub>1 | c0 \<rbrace>" using infer_e_lenI by metis
+ show "atom z0 \<sharp> (x, v)" using infer_e_lenI by metis
+ qed
+
+ thus \<open> \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> v'[x::=v]\<^sub>v\<^sub>v \<Rightarrow> \<lbrace> z3' : B_bitvec | CE_val (V_var z3') == CE_val v'[x::=v]\<^sub>v\<^sub>v \<rbrace>\<close>
+ using subst_tv.simps subst_ev.simps fresh_Pair * fresh_prodN subst_vv.simps by auto
+
+ show \<open>atom z3' \<sharp> AE_len v'[x::=v]\<^sub>v\<^sub>v\<close> using fresh_Pair * by metis
+ show \<open>atom z3' \<sharp> \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>\<close> using fresh_Pair * by metis
+ qed
+
+ moreover have "\<lbrace> z3' : B_int | CE_val (V_var z3') == CE_len ([v'[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) \<rbrace> = \<lbrace> z3' : B_int | CE_val (V_var z3') == CE_len [v']\<^sup>c\<^sup>e \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v"
+ using subst_tv.simps subst_ev.simps * by auto
+
+ ultimately show ?case using subst_ev.simps * ** by metis
+next
+ case (infer_e_mvarI \<Theta> \<B> \<Gamma>'' \<Phi> \<Delta> u \<tau>)
+
+ have "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> ; \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> (AE_mvar u) \<Rightarrow> \<tau>[x::=v]\<^sub>\<tau>\<^sub>v"
+ proof
+ show \<open> \<Theta> ; \<B>\<turnstile>\<^sub>w\<^sub>f \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>\<close> proof -
+ have "wfV \<Theta> \<B> \<Gamma> v (b_of \<tau>\<^sub>1)" using infer_v_wf infer_e_mvarI by auto
+ moreover have "b_of \<tau>\<^sub>1 = b\<^sub>1" using subtype_eq_base2 infer_e_mvarI b_of.simps by simp
+ ultimately show ?thesis using wf_subst(3)[OF infer_e_mvarI(1), of \<Gamma>' x b\<^sub>1 "c0[z0::=[x]\<^sup>v]\<^sub>c\<^sub>v" \<Gamma> v] infer_e_mvarI subst_g_inside by metis
+ qed
+ show \<open> \<Theta>\<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using infer_e_mvarI by auto
+ show \<open> \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<close> using wfD_subst infer_e_mvarI subtype_eq_base2 b_of.simps by metis
+ show \<open>(u, \<tau>[x::=v]\<^sub>\<tau>\<^sub>v) \<in> setD \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v\<close> using infer_e_mvarI subst_dv_member by metis
+ qed
+ moreover have " (AE_mvar u) = (AE_mvar u)[x::=v]\<^sub>e\<^sub>v" using subst_ev.simps by auto
+ ultimately show ?case by auto
+
+next
+ case (infer_e_concatI \<Theta> \<B> \<Gamma>'' \<Delta> \<Phi> v1 z1 c1 v2 z2 c2 z3)
+
+ hence zf: "atom z3 \<sharp> CE_concat [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e" using ce.fresh e.fresh opp.fresh by metis
+
+ obtain z3'::x where *:"atom z3' \<sharp> (x,v,v1,v2,AE_concat v1 v2, CE_concat [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e , AE_concat (v1[x::=v]\<^sub>v\<^sub>v) (v2[x::=v]\<^sub>v\<^sub>v) ,\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> , \<Gamma>'',v1 , v2)" using obtain_fresh by auto
+
+ hence **:"(\<lbrace> z3 : B_bitvec | CE_val (V_var z3) == CE_concat [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \<rbrace>) = \<lbrace> z3' : B_bitvec | CE_val (V_var z3') == CE_concat [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \<rbrace>"
+ using type_e_eq infer_e_concatI fresh_Pair zf by metis
+ have zfx: "atom x \<sharp> z3'" using fresh_at_base fresh_prodN * by auto
+
+ have v1: "\<Theta> ; \<B> ; \<Gamma>'' \<turnstile> v1 \<Rightarrow> \<lbrace> z3' : B_bitvec | CE_val (V_var z3') == CE_val v1 \<rbrace>"
+ using infer_e_concatI infer_v_form3 b_of.simps * fresh_Pair by metis
+ have v2: "\<Theta> ; \<B> ; \<Gamma>'' \<turnstile> v2 \<Rightarrow> \<lbrace> z3' : B_bitvec | CE_val (V_var z3') == CE_val v2 \<rbrace>"
+ using infer_e_concatI infer_v_form3 b_of.simps * fresh_Pair by metis
+
+ have "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> ; \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> (AE_concat (v1[x::=v]\<^sub>v\<^sub>v) (v2[x::=v]\<^sub>v\<^sub>v)) \<Rightarrow> \<lbrace> z3' : B_bitvec | CE_val (V_var z3') == CE_concat ([v1[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) ([v2[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) \<rbrace>"
+ proof
+ show \<open> \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<close> using wfD_subst infer_e_concatI subtype_eq_base2 b_of.simps by metis
+ show \<open> \<Theta>\<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> by(simp add: infer_e_concatI)
+ show \<open> \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> v1[x::=v]\<^sub>v\<^sub>v \<Rightarrow> \<lbrace> z3' : B_bitvec | CE_val (V_var z3') == CE_val (v1[x::=v]\<^sub>v\<^sub>v) \<rbrace>\<close>
+ using subst_infer_v_form infer_e_concatI fresh_prodN * b_of.simps by metis
+ show \<open> \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> v2[x::=v]\<^sub>v\<^sub>v \<Rightarrow> \<lbrace> z3' : B_bitvec | CE_val (V_var z3') == CE_val (v2[x::=v]\<^sub>v\<^sub>v) \<rbrace>\<close>
+ using subst_infer_v_form infer_e_concatI fresh_prodN * b_of.simps by metis
+ show \<open>atom z3' \<sharp> AE_concat v1[x::=v]\<^sub>v\<^sub>v v2[x::=v]\<^sub>v\<^sub>v\<close> using fresh_Pair * by metis
+ show \<open>atom z3' \<sharp> \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>\<close> using fresh_Pair * by metis
+ qed
+
+ moreover have "\<lbrace> z3' : B_bitvec | CE_val (V_var z3') == CE_concat ([v1[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) ([v2[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) \<rbrace> = \<lbrace> z3' : B_bitvec | CE_val (V_var z3') == CE_concat [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v"
+ using subst_tv.simps subst_ev.simps * by auto
+
+ ultimately show ?case using subst_ev.simps ** * by metis
+next
+ case (infer_e_splitI \<Theta> \<B> \<Gamma>'' \<Delta> \<Phi> v1 z1 c1 v2 z2 z3)
+ hence *:"atom z3 \<sharp> (x,v)" using fresh_Pair by auto
+ have \<open>x \<noteq> z3 \<close> using infer_e_splitI by force
+ have "\<Theta> ; \<Phi> ; \<B> ; (\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>) ; \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> (AE_split v1[x::=v]\<^sub>v\<^sub>v v2[x::=v]\<^sub>v\<^sub>v) \<Rightarrow>
+ \<lbrace> z3 : [ B_bitvec , B_bitvec ]\<^sup>b | [ v1[x::=v]\<^sub>v\<^sub>v ]\<^sup>c\<^sup>e == [ [#1[ [ z3 ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e @@ [#2[ [ z3 ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e AND
+ [| [#1[ [ z3 ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e == [ v2[x::=v]\<^sub>v\<^sub>v ]\<^sup>c\<^sup>e \<rbrace>"
+ proof
+ show \<open> \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<close> using wfD_subst infer_e_splitI subtype_eq_base2 b_of.simps by metis
+ show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using infer_e_splitI by auto
+ have \<open> \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> v1[x::=v]\<^sub>v\<^sub>v \<Rightarrow> \<lbrace> z1 : B_bitvec | c1 \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v\<close>
+ using subst_infer_v infer_e_splitI by metis
+ thus \<open> \<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> \<turnstile> v1[x::=v]\<^sub>v\<^sub>v \<Rightarrow> \<lbrace> z1 : B_bitvec | c1[x::=v]\<^sub>c\<^sub>v \<rbrace>\<close>
+ using infer_e_splitI subst_tv.simps fresh_Pair by metis
+ have \<open>x \<noteq> z2 \<close> using infer_e_splitI by force
+ have "(\<lbrace> z2 : B_int | ([ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ z2 ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e)
+ AND ([ leq [ [ z2 ]\<^sup>v ]\<^sup>c\<^sup>e [| [ v1[x::=v]\<^sub>v\<^sub>v ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e ) \<rbrace>) =
+ (\<lbrace> z2 : B_int | ([ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ z2 ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e )
+ AND ([ leq [ [ z2 ]\<^sup>v ]\<^sup>c\<^sup>e [| [ v1 ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e ) \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v)"
+ unfolding subst_cv.simps subst_cev.simps subst_vv.simps using \<open>x \<noteq> z2\<close> infer_e_splitI fresh_Pair by simp
+ thus \<open>\<Theta> ; \<B> ; \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @
+ \<Gamma> \<turnstile> v2[x::=v]\<^sub>v\<^sub>v \<Leftarrow> \<lbrace> z2 : B_int | [ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ z2 ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e
+ AND [ leq [ [ z2 ]\<^sup>v ]\<^sup>c\<^sup>e [| [ v1[x::=v]\<^sub>v\<^sub>v ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrace>\<close>
+ using infer_e_splitI subst_infer_check_v fresh_Pair by metis
+
+ show \<open>atom z1 \<sharp> AE_split v1[x::=v]\<^sub>v\<^sub>v v2[x::=v]\<^sub>v\<^sub>v\<close> using infer_e_splitI fresh_subst_vv_if by auto
+ show \<open>atom z2 \<sharp> AE_split v1[x::=v]\<^sub>v\<^sub>v v2[x::=v]\<^sub>v\<^sub>v\<close> using infer_e_splitI fresh_subst_vv_if by auto
+ show \<open>atom z3 \<sharp> AE_split v1[x::=v]\<^sub>v\<^sub>v v2[x::=v]\<^sub>v\<^sub>v\<close> using infer_e_splitI fresh_subst_vv_if by auto
+
+ show \<open>atom z3 \<sharp> \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>\<close> using fresh_subst_gv_inside infer_e_splitI by metis
+ show \<open>atom z2 \<sharp> \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>\<close> using fresh_subst_gv_inside infer_e_splitI by metis
+ show \<open>atom z1 \<sharp> \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>\<close> using fresh_subst_gv_inside infer_e_splitI by metis
+ qed
+ thus ?case apply (subst subst_tv.simps)
+ using infer_e_splitI fresh_Pair apply metis
+ unfolding subst_tv.simps subst_ev.simps subst_cv.simps subst_cev.simps subst_vv.simps *
+ using \<open>x \<noteq> z3\<close> by simp
+qed
+
+lemma infer_e_uniqueness:
+ assumes "\<Theta> ; \<Phi> ; \<B> ; GNil ; \<Delta> \<turnstile> e\<^sub>1 \<Rightarrow> \<tau>\<^sub>1" and "\<Theta> ; \<Phi> ; \<B> ; GNil ; \<Delta> \<turnstile> e\<^sub>1 \<Rightarrow> \<tau>\<^sub>2"
+ shows "\<tau>\<^sub>1 = \<tau>\<^sub>2"
+ using assms proof(nominal_induct rule:e.strong_induct)
+ case (AE_val x)
+ then show ?case using infer_e_elims(7)[OF AE_val(1)] infer_e_elims(7)[OF AE_val(2)] infer_v_uniqueness by metis
+next
+ case (AE_app f v)
+
+ obtain x1 b1 c1 s1' \<tau>1' where t1: "Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x1 b1 c1 \<tau>1' s1'))) = lookup_fun \<Phi> f \<and> \<tau>\<^sub>1 = \<tau>1'[x1::=v]\<^sub>\<tau>\<^sub>v" using infer_e_app2E[OF AE_app(1)] by metis
+ moreover obtain x2 b2 c2 s2' \<tau>2' where t2: "Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x2 b2 c2 \<tau>2' s2'))) = lookup_fun \<Phi> f \<and> \<tau>\<^sub>2 = \<tau>2'[x2::=v]\<^sub>\<tau>\<^sub>v" using infer_e_app2E[OF AE_app(2)] by metis
+
+ have "\<tau>1'[x1::=v]\<^sub>\<tau>\<^sub>v = \<tau>2'[x2::=v]\<^sub>\<tau>\<^sub>v" using t1 and t2 fun_ret_unique by metis
+ thus ?thesis using t1 t2 by auto
+next
+ case (AE_appP f b v)
+ obtain bv1 x1 b1 c1 s1' \<tau>1' where t1: "Some (AF_fundef f (AF_fun_typ_some bv1 (AF_fun_typ x1 b1 c1 \<tau>1' s1'))) = lookup_fun \<Phi> f \<and> \<tau>\<^sub>1 = \<tau>1'[bv1::=b]\<^sub>\<tau>\<^sub>b[x1::=v]\<^sub>\<tau>\<^sub>v" using infer_e_appP2E[OF AE_appP(1)] by metis
+ moreover obtain bv2 x2 b2 c2 s2' \<tau>2' where t2: "Some (AF_fundef f (AF_fun_typ_some bv2 (AF_fun_typ x2 b2 c2 \<tau>2' s2'))) = lookup_fun \<Phi> f \<and> \<tau>\<^sub>2 = \<tau>2'[bv2::=b]\<^sub>\<tau>\<^sub>b[x2::=v]\<^sub>\<tau>\<^sub>v" using infer_e_appP2E[OF AE_appP(2)] by metis
+
+ have "\<tau>1'[bv1::=b]\<^sub>\<tau>\<^sub>b[x1::=v]\<^sub>\<tau>\<^sub>v = \<tau>2'[bv2::=b]\<^sub>\<tau>\<^sub>b[x2::=v]\<^sub>\<tau>\<^sub>v" using t1 and t2 fun_poly_ret_unique by metis
+ thus ?thesis using t1 t2 by auto
+next
+ case (AE_op opp v1 v2)
+ show ?case proof(rule opp.exhaust)
+ assume "opp = plus"
+ hence "\<Theta> ; \<Phi> ; \<B> ; GNil ; \<Delta> \<turnstile> AE_op Plus v1 v2 \<Rightarrow> \<tau>\<^sub>1" and "\<Theta> ; \<Phi> ; \<B> ; GNil ; \<Delta> \<turnstile> AE_op Plus v1 v2 \<Rightarrow> \<tau>\<^sub>2" using AE_op by auto
+ thus ?thesis using infer_e_elims(11)[OF \<open>\<Theta> ; \<Phi> ; \<B> ; GNil ; \<Delta> \<turnstile> AE_op Plus v1 v2 \<Rightarrow> \<tau>\<^sub>1\<close> ] infer_e_elims(11)[OF \<open>\<Theta> ; \<Phi> ; \<B> ; GNil ; \<Delta> \<turnstile> AE_op Plus v1 v2 \<Rightarrow> \<tau>\<^sub>2\<close> ]
+ by force
+ next
+ assume "opp = leq"
+ hence "opp = LEq" using opp.strong_exhaust by auto
+ hence "\<Theta> ; \<Phi> ; \<B> ; GNil ; \<Delta> \<turnstile> AE_op LEq v1 v2 \<Rightarrow> \<tau>\<^sub>1" and "\<Theta> ; \<Phi> ; \<B> ; GNil ; \<Delta> \<turnstile> AE_op LEq v1 v2 \<Rightarrow> \<tau>\<^sub>2" using AE_op by auto
+ thus ?thesis using infer_e_elims(12)[OF \<open>\<Theta> ; \<Phi> ; \<B> ; GNil ; \<Delta> \<turnstile> AE_op LEq v1 v2 \<Rightarrow> \<tau>\<^sub>1\<close> ] infer_e_elims(12)[OF \<open>\<Theta> ; \<Phi> ; \<B> ; GNil ; \<Delta> \<turnstile> AE_op LEq v1 v2 \<Rightarrow> \<tau>\<^sub>2\<close> ]
+ by force
+ next
+ assume "opp = eq"
+ hence "opp = Eq" using opp.strong_exhaust by auto
+ hence "\<Theta> ; \<Phi> ; \<B> ; GNil ; \<Delta> \<turnstile> AE_op Eq v1 v2 \<Rightarrow> \<tau>\<^sub>1" and "\<Theta> ; \<Phi> ; \<B> ; GNil ; \<Delta> \<turnstile> AE_op Eq v1 v2 \<Rightarrow> \<tau>\<^sub>2" using AE_op by auto
+ thus ?thesis using infer_e_elims(25)[OF \<open>\<Theta> ; \<Phi> ; \<B> ; GNil ; \<Delta> \<turnstile> AE_op Eq v1 v2 \<Rightarrow> \<tau>\<^sub>1\<close> ] infer_e_elims(25)[OF \<open>\<Theta> ; \<Phi> ; \<B> ; GNil ; \<Delta> \<turnstile> AE_op Eq v1 v2 \<Rightarrow> \<tau>\<^sub>2\<close> ]
+ by force
+ qed
+next
+ case (AE_concat v1 v2)
+
+ obtain z3::x where t1:"\<tau>\<^sub>1 = \<lbrace> z3 : B_bitvec | [ [ z3 ]\<^sup>v ]\<^sup>c\<^sup>e == CE_concat [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \<rbrace> \<and> atom z3 \<sharp> v1 \<and> atom z3 \<sharp> v2 " using infer_e_elims(18)[OF AE_concat(1)] by metis
+ obtain z3'::x where t2:"\<tau>\<^sub>2 = \<lbrace> z3' : B_bitvec | [ [ z3' ]\<^sup>v ]\<^sup>c\<^sup>e == CE_concat [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \<rbrace> \<and> atom z3' \<sharp> v1 \<and> atom z3' \<sharp> v2" using infer_e_elims(18)[OF AE_concat(2)] by metis
+
+ thus ?case using t1 t2 type_e_eq ce.fresh by metis
+
+next
+ case (AE_fst v)
+
+ obtain z1 and b1 where "\<tau>\<^sub>1 = \<lbrace> z1 : b1 | CE_val (V_var z1) == (CE_fst [v]\<^sup>c\<^sup>e) \<rbrace>" using infer_v_form AE_fst by auto
+
+ obtain xx :: x and bb :: b and xxa :: x and bba :: b and cc :: c where
+ f1: "\<tau>\<^sub>2 = \<lbrace> xx : bb | CE_val (V_var xx) == CE_fst [v]\<^sup>c\<^sup>e \<rbrace> \<and> \<Theta> ; \<B> ; GNil\<turnstile>\<^sub>w\<^sub>f \<Delta> \<and> \<Theta> ; \<B> ; GNil \<turnstile> v \<Rightarrow> \<lbrace> xxa : B_pair bb bba | cc \<rbrace> \<and> atom xx \<sharp> v"
+ using infer_e_elims(8)[OF AE_fst(2)] by metis
+ obtain xxb :: x and bbb :: b and xxc :: x and bbc :: b and cca :: c where
+ f2: "\<tau>\<^sub>1 = \<lbrace> xxb : bbb | CE_val (V_var xxb) == CE_fst [v]\<^sup>c\<^sup>e \<rbrace> \<and> \<Theta> ; \<B> ; GNil\<turnstile>\<^sub>w\<^sub>f \<Delta> \<and> \<Theta> ; \<B> ; GNil \<turnstile> v \<Rightarrow> \<lbrace> xxc : B_pair bbb bbc | cca \<rbrace> \<and> atom xxb \<sharp> v"
+ using infer_e_elims(8)[OF AE_fst(1)] by metis
+ then have "B_pair bb bba = B_pair bbb bbc"
+ using f1 by (metis (no_types) b_of.simps infer_v_uniqueness)
+ then have "\<lbrace> xx : bbb | CE_val (V_var xx) == CE_fst [v]\<^sup>c\<^sup>e \<rbrace> = \<tau>\<^sub>2"
+ using f1 by auto
+ then show ?thesis
+ using f2 by (meson ce.fresh fresh_GNil type_e_eq wfG_x_fresh_in_v_simple)
+next
+ case (AE_snd v)
+ obtain xx :: x and bb :: b and xxa :: x and bba :: b and cc :: c where
+ f1: "\<tau>\<^sub>2 = \<lbrace> xx : bba | CE_val (V_var xx) == CE_snd [v]\<^sup>c\<^sup>e \<rbrace> \<and> \<Theta> ; \<B> ; GNil\<turnstile>\<^sub>w\<^sub>f \<Delta> \<and> \<Theta> ; \<B> ; GNil \<turnstile> v \<Rightarrow> \<lbrace> xxa : B_pair bb bba | cc \<rbrace> \<and> atom xx \<sharp> v"
+ using infer_e_elims(9)[OF AE_snd(2)] by metis
+ obtain xxb :: x and bbb :: b and xxc :: x and bbc :: b and cca :: c where
+ f2: "\<tau>\<^sub>1 = \<lbrace> xxb : bbc | CE_val (V_var xxb) == CE_snd [v]\<^sup>c\<^sup>e \<rbrace> \<and> \<Theta> ; \<B> ; GNil\<turnstile>\<^sub>w\<^sub>f \<Delta> \<and> \<Theta> ; \<B> ; GNil \<turnstile> v \<Rightarrow> \<lbrace> xxc : B_pair bbb bbc | cca \<rbrace> \<and> atom xxb \<sharp> v"
+ using infer_e_elims(9)[OF AE_snd(1)] by metis
+ then have "B_pair bb bba = B_pair bbb bbc"
+ using f1 by (metis (no_types) b_of.simps infer_v_uniqueness)
+ then have "\<lbrace> xx : bbc | CE_val (V_var xx) == CE_snd [v]\<^sup>c\<^sup>e \<rbrace> = \<tau>\<^sub>2"
+ using f1 by auto
+ then show ?thesis
+ using f2 by (meson ce.fresh fresh_GNil type_e_eq wfG_x_fresh_in_v_simple)
+next
+ case (AE_mvar x)
+ then show ?case using infer_e_elims(10)[OF AE_mvar(1)] infer_e_elims(10)[OF AE_mvar(2)] wfD_unique by metis
+next
+ case (AE_len x)
+ then show ?case using infer_e_elims(16)[OF AE_len(1)] infer_e_elims(16)[OF AE_len(2)] by force
+next
+ case (AE_split x1a x2)
+ then show ?case using infer_e_elims(22)[OF AE_split(1)] infer_e_elims(22)[OF AE_split(2)] by force
+qed
+
+section \<open>Statements\<close>
+
+lemma subst_infer_check_v1:
+ fixes v::v and v'::v and \<Gamma>::\<Gamma>
+ assumes "\<Gamma> = \<Gamma>\<^sub>1@((x,b\<^sub>1,c0[z0::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\<Gamma>\<Gamma>\<^sub>2)" and
+ "\<Theta> ; \<B> ; \<Gamma>\<^sub>2 \<turnstile> v \<Rightarrow> \<tau>\<^sub>1" and
+ "\<Theta> ; \<B> ; \<Gamma> \<turnstile> v' \<Leftarrow> \<tau>\<^sub>2" and
+ "\<Theta> ; \<B> ; \<Gamma>\<^sub>2 \<turnstile> \<tau>\<^sub>1 \<lesssim> \<lbrace> z0 : b\<^sub>1 | c0 \<rbrace>" and "atom z0 \<sharp> (x,v)"
+ shows "\<Theta> ; \<B> ; \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v \<turnstile> v'[x::=v]\<^sub>v\<^sub>v \<Leftarrow> \<tau>\<^sub>2[x::=v]\<^sub>\<tau>\<^sub>v"
+ using subst_g_inside check_v_wf assms subst_infer_check_v by metis
+
+lemma infer_v_c_valid:
+ assumes " \<Theta> ; \<B> ; \<Gamma> \<turnstile> v \<Rightarrow> \<tau>" and "\<Theta> ; \<B> ; \<Gamma> \<turnstile> \<tau> \<lesssim> \<lbrace> z : b | c \<rbrace>"
+ shows \<open>\<Theta> ; \<B> ; \<Gamma> \<Turnstile> c[z::=v]\<^sub>c\<^sub>v \<close>
+proof -
+ obtain z1 and b1 and c1 where *:"\<tau> = \<lbrace> z1 : b1 | c1 \<rbrace> \<and> atom z1 \<sharp> (c,v,\<Gamma>)" using obtain_fresh_z by metis
+ then have "b1 = b" using assms subtype_eq_base by metis
+ moreover then have "\<Theta> ; \<B> ; \<Gamma> \<turnstile> v \<Rightarrow> \<lbrace> z1 : b | c1 \<rbrace>" using assms * by auto
+
+ moreover have "\<Theta> ; \<B> ; (z1, b, c1) #\<^sub>\<Gamma> \<Gamma> \<Turnstile> c[z::=[ z1 ]\<^sup>v]\<^sub>c\<^sub>v " proof -
+ have "\<Theta> ; \<B> ; (z1, b, c1[z1::=[ z1 ]\<^sup>v]\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<Turnstile> c[z::=[ z1 ]\<^sup>v]\<^sub>v "
+ using subtype_valid[OF assms(2), of z1 z1 b c1 z c ] * fresh_prodN \<open>b1 = b\<close> by metis
+ moreover have "c1[z1::=[ z1 ]\<^sup>v]\<^sub>v = c1" using subst_v_v_def by simp
+ ultimately show ?thesis using subst_v_c_def by metis
+ qed
+ ultimately show ?thesis using * fresh_prodN subst_valid_simple by metis
+qed
+
+text \<open> Substitution Lemma for Statements \<close>
+
+lemma subst_infer_check_s:
+ fixes v::v and s::s and cs::branch_s and x::x and c::c and b::b and
+ \<Gamma>\<^sub>1::\<Gamma> and \<Gamma>\<^sub>2::\<Gamma> and css::branch_list
+ assumes "\<Theta> ; \<B> ; \<Gamma>\<^sub>1 \<turnstile> v \<Rightarrow> \<tau>" and "\<Theta> ; \<B> ; \<Gamma>\<^sub>1 \<turnstile> \<tau> \<lesssim> \<lbrace> z : b | c \<rbrace>" and
+ "atom z \<sharp> (x, v)"
+ shows "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>; \<Delta> \<turnstile> s \<Leftarrow> \<tau>' \<Longrightarrow>
+ \<Gamma> = (\<Gamma>\<^sub>2@((x,b,c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\<Gamma>\<Gamma>\<^sub>1)) \<Longrightarrow>
+ \<Theta> ; \<Phi> ; \<B> ; \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> s[x::=v]\<^sub>s\<^sub>v \<Leftarrow> \<tau>'[x::=v]\<^sub>\<tau>\<^sub>v"
+ and
+ "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>; \<Delta>; tid ; cons ; const ; v' \<turnstile> cs \<Leftarrow> \<tau>' \<Longrightarrow>
+ \<Gamma> = (\<Gamma>\<^sub>2@((x,b,c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\<Gamma>\<Gamma>\<^sub>1)) \<Longrightarrow>
+ \<Theta> ; \<Phi> ; \<B> ; \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v;
+ tid ; cons ; const ; v'[x::=v]\<^sub>v\<^sub>v \<turnstile> cs[x::=v]\<^sub>s\<^sub>v \<Leftarrow> \<tau>'[x::=v]\<^sub>\<tau>\<^sub>v"
+ and
+ "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>; \<Delta>; tid ; dclist ; v' \<turnstile> css \<Leftarrow> \<tau>' \<Longrightarrow>
+ \<Gamma> = (\<Gamma>\<^sub>2@((x,b,c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\<Gamma>\<Gamma>\<^sub>1)) \<Longrightarrow>
+ \<Theta> ; \<Phi> ; \<B> ; \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v; tid ; dclist ; v'[x::=v]\<^sub>v\<^sub>v \<turnstile>
+ subst_branchlv css x v \<Leftarrow> \<tau>'[x::=v]\<^sub>\<tau>\<^sub>v"
+
+ using assms proof(nominal_induct \<tau>' and \<tau>' and \<tau>' avoiding: x v arbitrary: \<Gamma>\<^sub>2 and \<Gamma>\<^sub>2 and \<Gamma>\<^sub>2
+ rule: check_s_check_branch_s_check_branch_list.strong_induct)
+ case (check_valI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v' \<tau>' \<tau>'')
+
+ have sg: " \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v = \<Gamma>\<^sub>2[x::=v]\<^sub>\<Gamma>\<^sub>v@\<Gamma>\<^sub>1" using check_valI by subst_mth
+ have "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> (AS_val (v'[x::=v]\<^sub>v\<^sub>v)) \<Leftarrow> \<tau>''[x::=v]\<^sub>\<tau>\<^sub>v" proof
+ have "\<Theta> ; \<B> ; \<Gamma>\<^sub>1\<turnstile>\<^sub>w\<^sub>f v : b " using infer_v_v_wf subtype_eq_base2 b_of.simps check_valI by metis
+ thus \<open>\<Theta> ; \<B> ; \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v\<turnstile>\<^sub>w\<^sub>f \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v\<close> using wf_subst(15) check_valI by auto
+ show \<open> \<Theta>\<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using check_valI by auto
+ show \<open> \<Theta> ; \<B> ; \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v \<turnstile> v'[x::=v]\<^sub>v\<^sub>v \<Rightarrow> \<tau>'[x::=v]\<^sub>\<tau>\<^sub>v\<close> proof(subst sg, rule subst_infer_v)
+ show "\<Theta> ; \<B> ; \<Gamma>\<^sub>1 \<turnstile> v \<Rightarrow> \<tau>" using check_valI by auto
+ show "\<Theta> ; \<B> ; \<Gamma>\<^sub>2 @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>\<^sub>1 \<turnstile> v' \<Rightarrow> \<tau>'" using check_valI by metis
+ show "\<Theta> ; \<B> ; \<Gamma>\<^sub>1 \<turnstile> \<tau> \<lesssim> \<lbrace> z: b | c \<rbrace>" using check_valI by auto
+ show "atom z \<sharp> (x, v)" using check_valI by auto
+ qed
+ show \<open>\<Theta> ; \<B> ; \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v \<turnstile> \<tau>'[x::=v]\<^sub>\<tau>\<^sub>v \<lesssim> \<tau>''[x::=v]\<^sub>\<tau>\<^sub>v\<close> using subst_subtype_tau check_valI sg by metis
+ qed
+
+ thus ?case using Typing.check_valI subst_sv.simps sg by auto
+next
+ case (check_letI xa \<Theta> \<Phi> \<B> \<Gamma> \<Delta> ea \<tau>a za sa ba ca)
+ have *:"(AS_let xa ea sa)[x::=v]\<^sub>s\<^sub>v=(AS_let xa (ea[x::=v]\<^sub>e\<^sub>v) sa[x::=v]\<^sub>s\<^sub>v)"
+ using subst_sv.simps \<open> atom xa \<sharp> x\<close> \<open> atom xa \<sharp> v\<close> by auto
+ show ?case unfolding * proof
+
+ show "atom xa \<sharp> (\<Theta>,\<Phi>,\<B>,\<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v,\<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v,ea[x::=v]\<^sub>e\<^sub>v,\<tau>a[x::=v]\<^sub>\<tau>\<^sub>v)"
+ by(subst_tuple_mth add: check_letI)
+
+ show "atom za \<sharp> (xa,\<Theta>,\<Phi>,\<B>,\<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v, \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v,ea[x::=v]\<^sub>e\<^sub>v,
+ \<tau>a[x::=v]\<^sub>\<tau>\<^sub>v,sa[x::=v]\<^sub>s\<^sub>v)"
+ by(subst_tuple_mth add: check_letI)
+
+ show "\<Theta>; \<Phi>; \<B>; \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v; \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile>
+ ea[x::=v]\<^sub>e\<^sub>v \<Rightarrow> \<lbrace> za : ba | ca[x::=v]\<^sub>c\<^sub>v \<rbrace>"
+ proof -
+ have "\<Theta>; \<Phi>; \<B>; \<Gamma>\<^sub>2[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>\<^sub>1; \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile>
+ ea[x::=v]\<^sub>e\<^sub>v \<Rightarrow> \<lbrace> za : ba | ca \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v"
+ using check_letI subst_infer_e by metis
+ thus ?thesis using check_letI subst_tv.simps
+ by (metis fresh_prod2I infer_e_wf subst_g_inside_simple)
+ qed
+
+ show "\<Theta>; \<Phi>; \<B>; (xa, ba, ca[x::=v]\<^sub>c\<^sub>v[za::=V_var xa]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v;
+ \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> sa[x::=v]\<^sub>s\<^sub>v \<Leftarrow> \<tau>a[x::=v]\<^sub>\<tau>\<^sub>v"
+ proof -
+ have "\<Theta>; \<Phi>; \<B>; ((xa, ba, ca[za::=V_var xa]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)[x::=v]\<^sub>\<Gamma>\<^sub>v ;
+ \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> sa[x::=v]\<^sub>s\<^sub>v \<Leftarrow> \<tau>a[x::=v]\<^sub>\<tau>\<^sub>v"
+ apply(rule check_letI(23)[of "(xa, ba, ca[za::=V_var xa]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>\<^sub>2"])
+ by(metis check_letI append_g.simps subst_defs)+
+
+ moreover have "(xa, ba, ca[x::=v]\<^sub>c\<^sub>v[za::=V_var xa]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v =
+ ((xa, ba, ca[za::=V_var xa]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)[x::=v]\<^sub>\<Gamma>\<^sub>v"
+ using subst_cv_commute subst_gv.simps check_letI
+ by (metis ms_fresh_all(39) ms_fresh_all(49) subst_cv_commute_full)
+ ultimately show ?thesis
+ using subst_defs by auto
+ qed
+ qed
+next
+ case (check_assertI xa \<Theta> \<Phi> \<B> \<Gamma> \<Delta> ca \<tau> s)
+ show ?case unfolding subst_sv.simps proof
+ show \<open>atom xa \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v, \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v, ca[x::=v]\<^sub>c\<^sub>v, \<tau>[x::=v]\<^sub>\<tau>\<^sub>v, s[x::=v]\<^sub>s\<^sub>v)\<close>
+ by(subst_tuple_mth add: check_assertI)
+ have "xa \<noteq> x" using check_assertI by fastforce
+ thus \<open> \<Theta> ; \<Phi> ; \<B> ; (xa, B_bool, ca[x::=v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> s[x::=v]\<^sub>s\<^sub>v \<Leftarrow> \<tau>[x::=v]\<^sub>\<tau>\<^sub>v\<close>
+ using check_assertI(12)[of "(xa, B_bool, c) #\<^sub>\<Gamma> \<Gamma>\<^sub>2" x v] check_assertI subst_gv.simps append_g.simps by metis
+ have \<open>\<Theta> ; \<B> ; \<Gamma>\<^sub>2[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>\<^sub>1 \<Turnstile> ca[x::=v]\<^sub>c\<^sub>v \<close> proof(rule subst_valid )
+ show \<open>\<Theta> ; \<B> ; \<Gamma>\<^sub>1 \<Turnstile> c[z::=v]\<^sub>c\<^sub>v \<close> using infer_v_c_valid check_assertI by metis
+ show \<open> \<Theta> ; \<B> ; \<Gamma>\<^sub>1 \<turnstile>\<^sub>w\<^sub>f v : b \<close> using check_assertI infer_v_wf b_of.simps subtype_eq_base
+ by (metis subtype_eq_base2)
+ show \<open> \<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>\<^sub>1 \<close> using check_assertI infer_v_wf by metis
+ have " \<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>\<^sub>2 @ (x, b, c[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>\<^sub>1" using check_assertI wfX_wfY by metis
+ thus \<open>atom x \<sharp> \<Gamma>\<^sub>1\<close> using check_assertI wfG_suffix wfG_elims by metis
+
+ moreover have "\<Theta> ; \<B> ; \<Gamma>\<^sub>1 \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b | c \<rbrace>" using subtype_wfT check_assertI by metis
+ moreover have "x \<noteq> z" using fresh_Pair check_assertI fresh_x_neq by metis
+ ultimately show \<open>atom x \<sharp> c\<close> using check_assertI wfT_fresh_c by metis
+
+ show \<open> \<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>\<^sub>2 @ (x, b, c[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>\<^sub>1 \<close> using check_assertI wfX_wfY by metis
+ show \<open>\<Theta> ; \<B> ; \<Gamma>\<^sub>2 @ (x, b, c[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>\<^sub>1 \<Turnstile> ca \<close> using check_assertI by auto
+ qed
+ thus \<open>\<Theta> ; \<B> ; \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v \<Turnstile> ca[x::=v]\<^sub>c\<^sub>v \<close> using check_assertI
+ proof -
+ show ?thesis
+ by (metis (no_types) \<open>\<Gamma> = \<Gamma>\<^sub>2 @ (x, b, c[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>\<^sub>1\<close> \<open>\<Theta> ; \<B> ; \<Gamma> \<Turnstile> ca\<close> \<open>\<Theta> ; \<B> ; \<Gamma>\<^sub>2[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>\<^sub>1 \<Turnstile> ca[x::=v]\<^sub>c\<^sub>v\<close> subst_g_inside valid_g_wf) (* 93 ms *)
+ qed
+
+ have "\<Theta> ; \<B> ; \<Gamma>\<^sub>1 \<turnstile>\<^sub>w\<^sub>f v : b" using infer_v_wf b_of.simps check_assertI
+ by (metis subtype_eq_base2)
+ thus \<open> \<Theta> ; \<B> ; \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v \<turnstile>\<^sub>w\<^sub>f \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<close> using wf_subst2(6) check_assertI by metis
+ qed
+next
+ case (check_branch_list_consI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist vv cs \<tau> css)
+ show ?case unfolding * using subst_sv.simps check_branch_list_consI by simp
+next
+ case (check_branch_list_finalI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist v cs \<tau>)
+ show ?case unfolding * using subst_sv.simps check_branch_list_finalI by simp
+next
+ case (check_branch_s_branchI \<Theta> \<B> \<Gamma> \<Delta> \<tau> const xa \<Phi> tid cons va sa)
+ hence *:"(AS_branch cons xa sa)[x::=v]\<^sub>s\<^sub>v = (AS_branch cons xa sa[x::=v]\<^sub>s\<^sub>v)" using subst_branchv.simps fresh_Pair by metis
+ show ?case unfolding * proof
+
+ show "\<Theta> ; \<B> ; \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v\<turnstile>\<^sub>w\<^sub>f \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v "
+ using wf_subst check_branch_s_branchI subtype_eq_base2 b_of.simps infer_v_wf by metis
+
+ show "\<turnstile>\<^sub>w\<^sub>f \<Theta> " using check_branch_s_branchI by metis
+
+ show "\<Theta> ; \<B> ; \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v \<turnstile>\<^sub>w\<^sub>f \<tau>[x::=v]\<^sub>\<tau>\<^sub>v "
+ using wf_subst check_branch_s_branchI subtype_eq_base2 b_of.simps infer_v_wf by metis
+
+ show wft:"\<Theta> ; {||} ; GNil\<turnstile>\<^sub>w\<^sub>f const " using check_branch_s_branchI by metis
+
+ show "atom xa \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v, \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v, tid, cons, const,va[x::=v]\<^sub>v\<^sub>v, \<tau>[x::=v]\<^sub>\<tau>\<^sub>v)"
+ apply(unfold fresh_prodN, (simp add: check_branch_s_branchI )+)
+ apply(rule,metis fresh_z_subst_g check_branch_s_branchI fresh_Pair )
+ by(metis fresh_subst_dv check_branch_s_branchI fresh_Pair )
+
+ have "\<Theta> ; \<Phi> ; \<B> ; ((xa, b_of const, CE_val va == CE_val(V_cons tid cons (V_var xa)) AND c_of const xa) #\<^sub>\<Gamma> \<Gamma>)[x::=v]\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> sa[x::=v]\<^sub>s\<^sub>v \<Leftarrow> \<tau>[x::=v]\<^sub>\<tau>\<^sub>v"
+ using check_branch_s_branchI by (metis append_g.simps(2))
+
+ moreover have "(xa, b_of const, CE_val va[x::=v]\<^sub>v\<^sub>v == CE_val (V_cons tid cons (V_var xa)) AND c_of (const) xa) #\<^sub>\<Gamma> \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v =
+ ((xa, b_of const , CE_val va == CE_val (V_cons tid cons (V_var xa)) AND c_of const xa) #\<^sub>\<Gamma> \<Gamma>)[x::=v]\<^sub>\<Gamma>\<^sub>v"
+ proof -
+ have *:"xa \<noteq> x" using check_branch_s_branchI fresh_at_base by metis
+ have "atom x \<sharp> const" using wfT_nil_supp[OF wft] fresh_def by auto
+ hence "atom x \<sharp> (const,xa)" using fresh_at_base wfT_nil_supp[OF wft] fresh_Pair fresh_def * by auto
+ moreover hence "(c_of (const) xa)[x::=v]\<^sub>c\<^sub>v = c_of (const) xa"
+ using c_of_fresh[of x const xa] forget_subst_cv wfT_nil_supp wft by metis
+ moreover hence "(V_cons tid cons (V_var xa))[x::=v]\<^sub>v\<^sub>v = (V_cons tid cons (V_var xa))" using check_branch_s_branchI subst_vv.simps * by metis
+ ultimately show ?thesis using subst_gv.simps check_branch_s_branchI subst_cv.simps subst_cev.simps * by presburger
+ qed
+
+ ultimately show "\<Theta> ; \<Phi> ; \<B> ; (xa, b_of const, CE_val va[x::=v]\<^sub>v\<^sub>v == CE_val (V_cons tid cons (V_var xa)) AND c_of const xa) #\<^sub>\<Gamma> \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> sa[x::=v]\<^sub>s\<^sub>v \<Leftarrow> \<tau>[x::=v]\<^sub>\<tau>\<^sub>v"
+ by metis
+ qed
+
+next
+ case (check_let2I xa \<Theta> \<Phi> \<B> G \<Delta> t s1 \<tau>a s2 )
+ hence *:"(AS_let2 xa t s1 s2)[x::=v]\<^sub>s\<^sub>v = (AS_let2 xa t[x::=v]\<^sub>\<tau>\<^sub>v (s1[x::=v]\<^sub>s\<^sub>v) s2[x::=v]\<^sub>s\<^sub>v)" using subst_sv.simps fresh_Pair by metis
+ have "xa \<noteq> x" using check_let2I fresh_at_base by metis
+ show ?case unfolding * proof
+ show "atom xa \<sharp> (\<Theta>, \<Phi>, \<B>, G[x::=v]\<^sub>\<Gamma>\<^sub>v, \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v, t[x::=v]\<^sub>\<tau>\<^sub>v, s1[x::=v]\<^sub>s\<^sub>v, \<tau>a[x::=v]\<^sub>\<tau>\<^sub>v)"
+ by(subst_tuple_mth add: check_let2I)
+ show "\<Theta> ; \<Phi> ; \<B> ; G[x::=v]\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> s1[x::=v]\<^sub>s\<^sub>v \<Leftarrow> t[x::=v]\<^sub>\<tau>\<^sub>v" using check_let2I by metis
+
+ have "\<Theta> ; \<Phi> ; \<B> ; ((xa, b_of t, c_of t xa) #\<^sub>\<Gamma> G)[x::=v]\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> s2[x::=v]\<^sub>s\<^sub>v \<Leftarrow> \<tau>a[x::=v]\<^sub>\<tau>\<^sub>v"
+ proof(rule check_let2I(14))
+ show \<open>(xa, b_of t, c_of t xa) #\<^sub>\<Gamma> G = (((xa, b_of t, c_of t xa)#\<^sub>\<Gamma> \<Gamma>\<^sub>2)) @ (x, b, c[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>\<^sub>1\<close>
+ using check_let2I append_g.simps by metis
+ show \<open> \<Theta> ; \<B> ; \<Gamma>\<^sub>1 \<turnstile> v \<Rightarrow> \<tau>\<close> using check_let2I by metis
+ show \<open>\<Theta> ; \<B> ; \<Gamma>\<^sub>1 \<turnstile> \<tau> \<lesssim> \<lbrace> z : b | c \<rbrace>\<close> using check_let2I by metis
+ show \<open>atom z \<sharp> (x, v)\<close> using check_let2I by metis
+ qed
+ moreover have "c_of t[x::=v]\<^sub>\<tau>\<^sub>v xa = (c_of t xa)[x::=v]\<^sub>c\<^sub>v" using subst_v_c_of fresh_Pair check_let2I by metis
+ moreover have "b_of t[x::=v]\<^sub>\<tau>\<^sub>v = b_of t" using b_of.simps subst_tv.simps b_of_subst by metis
+ ultimately show " \<Theta> ; \<Phi> ; \<B> ; (xa, b_of t[x::=v]\<^sub>\<tau>\<^sub>v, c_of t[x::=v]\<^sub>\<tau>\<^sub>v xa) #\<^sub>\<Gamma> G[x::=v]\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> s2[x::=v]\<^sub>s\<^sub>v \<Leftarrow> \<tau>a[x::=v]\<^sub>\<tau>\<^sub>v"
+ using check_let2I(14) subst_gv.simps \<open>xa \<noteq> x\<close> b_of.simps by metis
+ qed
+
+next
+
+ case (check_varI u \<Theta> \<Phi> \<B> \<Gamma> \<Delta> \<tau>' va \<tau>'' s)
+ have **: "\<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v = \<Gamma>\<^sub>2[x::=v]\<^sub>\<Gamma>\<^sub>v@\<Gamma>\<^sub>1" using subst_g_inside check_s_wf check_varI by meson
+
+ have "\<Theta> ; \<Phi> ;\<B> ; subst_gv \<Gamma> x v ; \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> AS_var u \<tau>'[x::=v]\<^sub>\<tau>\<^sub>v (va[x::=v]\<^sub>v\<^sub>v) (subst_sv s x v) \<Leftarrow> \<tau>''[x::=v]\<^sub>\<tau>\<^sub>v"
+ proof(rule Typing.check_varI)
+ show "atom u \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v, \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v, \<tau>'[x::=v]\<^sub>\<tau>\<^sub>v, va[x::=v]\<^sub>v\<^sub>v, \<tau>''[x::=v]\<^sub>\<tau>\<^sub>v)"
+ by(subst_tuple_mth add: check_varI)
+ show "\<Theta> ; \<B> ; \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v \<turnstile> va[x::=v]\<^sub>v\<^sub>v \<Leftarrow> \<tau>'[x::=v]\<^sub>\<tau>\<^sub>v" using check_varI subst_infer_check_v ** by metis
+ show "\<Theta> ; \<Phi> ; \<B> ; subst_gv \<Gamma> x v ; (u, \<tau>'[x::=v]\<^sub>\<tau>\<^sub>v) #\<^sub>\<Delta> \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> s[x::=v]\<^sub>s\<^sub>v \<Leftarrow> \<tau>''[x::=v]\<^sub>\<tau>\<^sub>v" proof -
+ have "wfD \<Theta> \<B> (\<Gamma>\<^sub>2 @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>\<^sub>1) ((u,\<tau>')#\<^sub>\<Delta> \<Delta>)" using check_varI check_s_wf by meson
+ moreover have "wfV \<Theta> \<B> \<Gamma>\<^sub>1 v (b_of \<tau>)" using infer_v_wf check_varI(6) check_varI by metis
+ have "wfD \<Theta> \<B> (\<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v) ((u, \<tau>'[x::=v]\<^sub>\<tau>\<^sub>v) #\<^sub>\<Delta> \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v)" proof(subst subst_dv.simps(2)[symmetric], subst **, rule wfD_subst)
+ show "\<Theta> ; \<B> ; \<Gamma>\<^sub>1 \<turnstile> v \<Rightarrow> \<tau>" using check_varI by auto
+ show "\<Theta> ; \<B> ; \<Gamma>\<^sub>2 @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>\<^sub>1\<turnstile>\<^sub>w\<^sub>f (u, \<tau>') #\<^sub>\<Delta> \<Delta>" using check_varI check_s_wf by simp
+ show "b_of \<tau> = b" using check_varI subtype_eq_base2 b_of.simps by auto
+ qed
+ thus ?thesis using check_varI by auto
+ qed
+ qed
+ moreover have "atom u \<sharp> (x,v)" using u_fresh_xv by auto
+ ultimately show ?case using subst_sv.simps(7) by auto
+
+next
+ case (check_assignI P \<Phi> \<B> \<Gamma> \<Delta> u \<tau>1 v' z1 \<tau>') (* may need to revisit subst in \<Delta> as well *)
+
+ have "wfG P \<B> \<Gamma>" using check_v_wf check_assignI by simp
+ hence gs: "\<Gamma>\<^sub>2[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>\<^sub>1 = \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v" using subst_g_inside check_assignI by simp
+
+ have "P ; \<Phi> ; \<B> ; \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> AS_assign u (v'[x::=v]\<^sub>v\<^sub>v) \<Leftarrow> \<tau>'[x::=v]\<^sub>\<tau>\<^sub>v"
+ proof(rule Typing.check_assignI)
+ show "P \<turnstile>\<^sub>w\<^sub>f \<Phi> " using check_assignI by auto
+ show "wfD P \<B> (\<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v) \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v" using wf_subst(15)[OF check_assignI(2)] gs infer_v_v_wf check_assignI b_of.simps subtype_eq_base2 by metis
+ thus "(u, \<tau>1[x::=v]\<^sub>\<tau>\<^sub>v) \<in> setD \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v" using check_assignI subst_dv_member by metis
+ thus "P ; \<B> ; \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v \<turnstile> v'[x::=v]\<^sub>v\<^sub>v \<Leftarrow> \<tau>1[x::=v]\<^sub>\<tau>\<^sub>v" using subst_infer_check_v check_assignI gs by metis
+
+ have "P ; \<B> ; \<Gamma>\<^sub>2[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>\<^sub>1 \<turnstile> \<lbrace> z : B_unit | TRUE \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v \<lesssim> \<tau>'[x::=v]\<^sub>\<tau>\<^sub>v" proof(rule subst_subtype_tau)
+ show "P ; \<B> ; \<Gamma>\<^sub>1 \<turnstile> v \<Rightarrow> \<tau>" using check_assignI by auto
+ show "P ; \<B> ; \<Gamma>\<^sub>1 \<turnstile> \<tau> \<lesssim> \<lbrace> z : b | c \<rbrace>" using check_assignI by meson
+ show "P ; \<B> ; \<Gamma>\<^sub>2 @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>\<^sub>1 \<turnstile> \<lbrace> z : B_unit | TRUE \<rbrace> \<lesssim> \<tau>'" using check_assignI
+ by (metis Abs1_eq_iff(3) \<tau>.eq_iff c.fresh(1) c.perm_simps(1))
+ show "atom z \<sharp> (x, v)" using check_assignI by auto
+ qed
+ moreover have "\<lbrace> z : B_unit | TRUE \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v = \<lbrace> z : B_unit | TRUE \<rbrace>" using subst_tv.simps subst_cv.simps check_assignI by presburger
+ ultimately show "P ; \<B> ; \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v \<turnstile> \<lbrace> z : B_unit | TRUE \<rbrace> \<lesssim> \<tau>'[x::=v]\<^sub>\<tau>\<^sub>v" using gs by auto
+ qed
+ thus ?case using subst_sv.simps(5) by auto
+
+next
+ case (check_whileI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> s1 z' s2 \<tau>')
+ have " wfG \<Theta> \<B> (\<Gamma>\<^sub>2 @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>\<^sub>1)" using check_whileI check_s_wf by meson
+ hence **: " \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v = \<Gamma>\<^sub>2[x::=v]\<^sub>\<Gamma>\<^sub>v@\<Gamma>\<^sub>1" using subst_g_inside wf check_whileI by auto
+ have teq: "(\<lbrace> z : B_unit | TRUE \<rbrace>)[x::=v]\<^sub>\<tau>\<^sub>v = (\<lbrace> z : B_unit | TRUE \<rbrace>)" by(auto simp add: subst_sv.simps check_whileI)
+ moreover have "(\<lbrace> z : B_unit | TRUE \<rbrace>) = (\<lbrace> z' : B_unit | TRUE \<rbrace>)" using type_eq_flip c.fresh flip_fresh_fresh by metis
+ ultimately have teq2:"(\<lbrace> z' : B_unit | TRUE \<rbrace>)[x::=v]\<^sub>\<tau>\<^sub>v = (\<lbrace> z' : B_unit | TRUE \<rbrace>)" by metis
+
+ hence "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> s1[x::=v]\<^sub>s\<^sub>v \<Leftarrow> \<lbrace> z' : B_bool | TRUE \<rbrace>" using check_whileI subst_sv.simps subst_top_eq by metis
+ moreover have "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> s2[x::=v]\<^sub>s\<^sub>v \<Leftarrow> \<lbrace> z' : B_unit | TRUE \<rbrace>" using check_whileI subst_top_eq by metis
+ moreover have "\<Theta> ; \<B> ; \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v \<turnstile> \<lbrace> z' : B_unit | TRUE \<rbrace> \<lesssim> \<tau>'[x::=v]\<^sub>\<tau>\<^sub>v" proof -
+ have "\<Theta> ; \<B> ; \<Gamma>\<^sub>2[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>\<^sub>1 \<turnstile> \<lbrace> z' : B_unit | TRUE \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v \<lesssim> \<tau>'[x::=v]\<^sub>\<tau>\<^sub>v" proof(rule subst_subtype_tau)
+ show "\<Theta> ; \<B> ; \<Gamma>\<^sub>1 \<turnstile> v \<Rightarrow> \<tau>" by(auto simp add: check_whileI)
+ show "\<Theta> ; \<B> ; \<Gamma>\<^sub>1 \<turnstile> \<tau> \<lesssim> \<lbrace> z : b | c \<rbrace>" by(auto simp add: check_whileI)
+ show "\<Theta> ; \<B> ; \<Gamma>\<^sub>2 @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>\<^sub>1 \<turnstile> \<lbrace> z' : B_unit | TRUE \<rbrace> \<lesssim> \<tau>'" using check_whileI by metis
+ show "atom z \<sharp> (x, v)" by(auto simp add: check_whileI)
+ qed
+ thus ?thesis using teq2 ** by auto
+ qed
+
+ ultimately have " \<Theta> ; \<Phi> ; \<B> ; \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> AS_while s1[x::=v]\<^sub>s\<^sub>v s2[x::=v]\<^sub>s\<^sub>v \<Leftarrow> \<tau>'[x::=v]\<^sub>\<tau>\<^sub>v"
+ using Typing.check_whileI by metis
+ then show ?case using subst_sv.simps by metis
+next
+ case (check_seqI P \<Phi> \<B> \<Gamma> \<Delta> s1 z s2 \<tau> )
+ hence "P ; \<Phi>; \<B> ; \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> AS_seq (s1[x::=v]\<^sub>s\<^sub>v) (s2[x::=v]\<^sub>s\<^sub>v) \<Leftarrow> \<tau>[x::=v]\<^sub>\<tau>\<^sub>v" using Typing.check_seqI subst_top_eq check_seqI by metis
+ then show ?case using subst_sv.simps by metis
+next
+ case (check_caseI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist v' cs \<tau> za)
+
+ have wf: "wfG \<Theta> \<B> \<Gamma>" using check_caseI check_v_wf by simp
+ have **: "\<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v = \<Gamma>\<^sub>2[x::=v]\<^sub>\<Gamma>\<^sub>v@\<Gamma>\<^sub>1" using subst_g_inside wf check_caseI by auto
+
+ have "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> AS_match (v'[x::=v]\<^sub>v\<^sub>v) (subst_branchlv cs x v) \<Leftarrow> \<tau>[x::=v]\<^sub>\<tau>\<^sub>v" proof(rule Typing.check_caseI )
+ show "check_branch_list \<Theta> \<Phi> \<B> (\<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v) \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v tid dclist v'[x::=v]\<^sub>v\<^sub>v (subst_branchlv cs x v ) (\<tau>[x::=v]\<^sub>\<tau>\<^sub>v)" using check_caseI by auto
+ show "AF_typedef tid dclist \<in> set \<Theta>" using check_caseI by auto
+ show "\<Theta> ; \<B> ; \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v \<turnstile> v'[x::=v]\<^sub>v\<^sub>v \<Leftarrow> \<lbrace> za : B_id tid | TRUE \<rbrace>" proof -
+ have "\<Theta> ; \<B> ; \<Gamma>\<^sub>2 @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>\<^sub>1 \<turnstile> v' \<Leftarrow> \<lbrace> za : B_id tid | TRUE \<rbrace>"
+ using check_caseI by argo
+ hence "\<Theta> ; \<B> ; \<Gamma>\<^sub>2[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>\<^sub>1 \<turnstile> v'[x::=v]\<^sub>v\<^sub>v \<Leftarrow> (\<lbrace> za : B_id tid | TRUE \<rbrace>)[x::=v]\<^sub>\<tau>\<^sub>v"
+ using check_caseI subst_infer_check_v[OF check_caseI(7) _ check_caseI(8) check_caseI(9)] by meson
+ moreover have "(\<lbrace> za : B_id tid | TRUE \<rbrace>) = ((\<lbrace> za : B_id tid | TRUE \<rbrace>)[x::=v]\<^sub>\<tau>\<^sub>v)"
+ using subst_cv.simps subst_tv.simps subst_cv_true by fast
+ ultimately show ?thesis using check_caseI ** by argo
+ qed
+ show "wfTh \<Theta>" using check_caseI by auto
+ qed
+ thus ?case using subst_branchlv.simps subst_sv.simps(4) by metis
+next
+ case (check_ifI z' \<Theta> \<Phi> \<B> \<Gamma> \<Delta> va s1 s2 \<tau>')
+ show ?case unfolding subst_sv.simps proof
+ show \<open>atom z' \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v, \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v, va[x::=v]\<^sub>v\<^sub>v, s1[x::=v]\<^sub>s\<^sub>v, s2[x::=v]\<^sub>s\<^sub>v, \<tau>'[x::=v]\<^sub>\<tau>\<^sub>v)\<close>
+ by(subst_tuple_mth add: check_ifI)
+ have *:"\<lbrace> z' : B_bool | TRUE \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v = \<lbrace> z' : B_bool | TRUE \<rbrace>" using subst_tv.simps check_ifI
+ by (metis freshers(19) subst_cv.simps(1) type_eq_subst)
+ have **: "\<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v = \<Gamma>\<^sub>2[x::=v]\<^sub>\<Gamma>\<^sub>v@\<Gamma>\<^sub>1" using subst_g_inside wf check_ifI check_v_wf by metis
+ show \<open>\<Theta> ; \<B> ; \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v \<turnstile> va[x::=v]\<^sub>v\<^sub>v \<Leftarrow> \<lbrace> z' : B_bool | TRUE \<rbrace>\<close>
+ proof(subst *[symmetric], rule subst_infer_check_v1[where \<Gamma>\<^sub>1=\<Gamma>\<^sub>2 and \<Gamma>\<^sub>2=\<Gamma>\<^sub>1])
+ show "\<Gamma> = \<Gamma>\<^sub>2 @ ((x, b ,c[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>\<^sub>1)" using check_ifI by metis
+ show \<open> \<Theta> ; \<B> ; \<Gamma>\<^sub>1 \<turnstile> v \<Rightarrow> \<tau>\<close> using check_ifI by metis
+ show \<open>\<Theta> ; \<B> ; \<Gamma> \<turnstile> va \<Leftarrow> \<lbrace> z' : B_bool | TRUE \<rbrace>\<close> using check_ifI by metis
+ show \<open>\<Theta> ; \<B> ; \<Gamma>\<^sub>1 \<turnstile> \<tau> \<lesssim> \<lbrace> z : b | c \<rbrace>\<close> using check_ifI by metis
+ show \<open>atom z \<sharp> (x, v)\<close> using check_ifI by metis
+ qed
+
+ have " \<lbrace> z' : b_of \<tau>'[x::=v]\<^sub>\<tau>\<^sub>v | [ va[x::=v]\<^sub>v\<^sub>v ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of \<tau>'[x::=v]\<^sub>\<tau>\<^sub>v z' \<rbrace> = \<lbrace> z' : b_of \<tau>' | [ va ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of \<tau>' z' \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v"
+ by(simp add: subst_tv.simps fresh_Pair check_ifI b_of_subst subst_v_c_of)
+
+ thus \<open> \<Theta> ; \<Phi> ; \<B> ; \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> s1[x::=v]\<^sub>s\<^sub>v \<Leftarrow> \<lbrace> z' : b_of \<tau>'[x::=v]\<^sub>\<tau>\<^sub>v | [ va[x::=v]\<^sub>v\<^sub>v ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of \<tau>'[x::=v]\<^sub>\<tau>\<^sub>v z' \<rbrace>\<close>
+ using check_ifI by metis
+ have " \<lbrace> z' : b_of \<tau>'[x::=v]\<^sub>\<tau>\<^sub>v | [ va[x::=v]\<^sub>v\<^sub>v ]\<^sup>c\<^sup>e == [ [ L_false ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of \<tau>'[x::=v]\<^sub>\<tau>\<^sub>v z' \<rbrace> = \<lbrace> z' : b_of \<tau>' | [ va ]\<^sup>c\<^sup>e == [ [ L_false ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of \<tau>' z' \<rbrace>[x::=v]\<^sub>\<tau>\<^sub>v"
+ by(simp add: subst_tv.simps fresh_Pair check_ifI b_of_subst subst_v_c_of)
+ thus \<open> \<Theta> ; \<Phi> ; \<B> ; \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> s2[x::=v]\<^sub>s\<^sub>v \<Leftarrow> \<lbrace> z' : b_of \<tau>'[x::=v]\<^sub>\<tau>\<^sub>v | [ va[x::=v]\<^sub>v\<^sub>v ]\<^sup>c\<^sup>e == [ [ L_false ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of \<tau>'[x::=v]\<^sub>\<tau>\<^sub>v z' \<rbrace>\<close>
+ using check_ifI by metis
+ qed
+qed
+
+lemma subst_check_check_s:
+ fixes v::v and s::s and cs::branch_s and x::x and c::c and b::b and \<Gamma>\<^sub>1::\<Gamma> and \<Gamma>\<^sub>2::\<Gamma>
+ assumes "\<Theta> ; \<B> ; \<Gamma>\<^sub>1 \<turnstile> v \<Leftarrow> \<lbrace> z : b | c \<rbrace>" and "atom z \<sharp> (x, v)"
+ and "check_s \<Theta> \<Phi> \<B> \<Gamma> \<Delta> s \<tau>'" and "\<Gamma> = (\<Gamma>\<^sub>2@((x,b,c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\<Gamma>\<Gamma>\<^sub>1))"
+ shows "check_s \<Theta> \<Phi> \<B> (subst_gv \<Gamma> x v) \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v (s[x::=v]\<^sub>s\<^sub>v) (subst_tv \<tau>' x v )"
+proof -
+ obtain \<tau> where "\<Theta> ; \<B> ; \<Gamma>\<^sub>1 \<turnstile> v \<Rightarrow> \<tau> \<and> \<Theta> ; \<B> ; \<Gamma>\<^sub>1 \<turnstile> \<tau> \<lesssim> \<lbrace> z : b | c \<rbrace>" using check_v_elims assms by auto
+ thus ?thesis using subst_infer_check_s assms by metis
+qed
+
+text \<open> If a statement checks against a type @{text "\<tau>"} then it checks against a supertype of @{text "\<tau>"} \<close>
+lemma check_s_supertype:
+ fixes v::v and s::s and cs::branch_s and x::x and c::c and b::b and \<Gamma>::\<Gamma> and \<Gamma>'::\<Gamma> and css::branch_list
+ shows "check_s \<Theta> \<Phi> \<B> G \<Delta> s t1 \<Longrightarrow> \<Theta> ; \<B> ; G \<turnstile> t1 \<lesssim> t2 \<Longrightarrow> check_s \<Theta> \<Phi> \<B> G \<Delta> s t2" and
+ "check_branch_s \<Theta> \<Phi> \<B> G \<Delta> tid cons const v cs t1 \<Longrightarrow> \<Theta> ; \<B> ; G \<turnstile> t1 \<lesssim> t2 \<Longrightarrow> check_branch_s \<Theta> \<Phi> \<B> G \<Delta> tid cons const v cs t2" and
+ "check_branch_list \<Theta> \<Phi> \<B> G \<Delta> tid dclist v css t1 \<Longrightarrow> \<Theta> ; \<B> ; G \<turnstile> t1 \<lesssim> t2 \<Longrightarrow> check_branch_list \<Theta> \<Phi> \<B> G \<Delta> tid dclist v css t2"
+proof(induct arbitrary: t2 and t2 and t2 rule: check_s_check_branch_s_check_branch_list.inducts)
+ case (check_valI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v \<tau>' \<tau> )
+ hence " \<Theta> ; \<B> ; \<Gamma> \<turnstile> \<tau>' \<lesssim> t2" using subtype_trans by meson
+ then show ?case using subtype_trans Typing.check_valI check_valI by metis
+
+next
+ case (check_letI x \<Theta> \<Phi> \<B> \<Gamma> \<Delta> e \<tau> z s b c)
+ show ?case proof(rule Typing.check_letI)
+ show "atom x \<sharp>(\<Theta>, \<Phi>, \<B>, \<Gamma>, \<Delta>, e, t2)" using check_letI subtype_fresh_tau fresh_prodN by metis
+ show "atom z \<sharp> (x, \<Theta>, \<Phi>, \<B>, \<Gamma>, \<Delta>, e, t2, s)" using check_letI(2) subtype_fresh_tau[of z \<tau> \<Gamma> _ _ t2] fresh_prodN check_letI(6) by auto
+ show "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile> e \<Rightarrow> \<lbrace> z : b | c \<rbrace>" using check_letI by meson
+
+ have "wfG \<Theta> \<B> ((x, b, c[z::=[x]\<^sup>v]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)" using check_letI check_s_wf subst_defs by metis
+ moreover have "toSet \<Gamma> \<subseteq> toSet ((x, b, c[z::=[x]\<^sup>v]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)" by auto
+ ultimately have " \<Theta> ; \<B> ; (x, b, c[z::=[x]\<^sup>v]\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<turnstile> \<tau> \<lesssim> t2" using subtype_weakening[OF check_letI(6)] by auto
+ thus "\<Theta> ; \<Phi> ; \<B> ; (x, b, c[z::=[x]\<^sup>v]\<^sub>v) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile> s \<Leftarrow> t2" using check_letI subst_defs by metis
+ qed
+next
+ case (check_branch_list_consI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist v cs \<tau> css)
+ then show ?case using Typing.check_branch_list_consI by auto
+next
+ case (check_branch_list_finalI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist v cs \<tau>)
+ then show ?case using Typing.check_branch_list_finalI by auto
+next
+ case (check_branch_s_branchI \<Theta> \<B> \<Gamma> \<Delta> \<tau> const x \<Phi> tid cons v s)
+ show ?case proof
+ have "atom x \<sharp> t2" using subtype_fresh_tau[of x \<tau> ] check_branch_s_branchI(5,8) fresh_prodN by metis
+ thus "atom x \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>, \<Delta>, tid, cons, const, v, t2)" using check_branch_s_branchI fresh_prodN by metis
+ show "wfT \<Theta> \<B> \<Gamma> t2" using subtype_wf check_branch_s_branchI by meson
+ show "\<Theta> ; \<Phi> ; \<B> ; (x, b_of const, CE_val v == CE_val(V_cons tid cons (V_var x)) AND c_of const x) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile> s \<Leftarrow> t2" proof -
+ have "wfG \<Theta> \<B> ((x, b_of const, CE_val v == CE_val(V_cons tid cons (V_var x)) AND c_of const x) #\<^sub>\<Gamma> \<Gamma>)" using check_s_wf check_branch_s_branchI by metis
+ moreover have "toSet \<Gamma> \<subseteq> toSet ((x, b_of const, CE_val v == CE_val (V_cons tid cons (V_var x)) AND c_of const x) #\<^sub>\<Gamma> \<Gamma>)" by auto
+ hence "\<Theta> ; \<B> ; ((x, b_of const, CE_val v == CE_val(V_cons tid cons (V_var x)) AND c_of const x) #\<^sub>\<Gamma> \<Gamma>) \<turnstile> \<tau> \<lesssim> t2"
+ using check_branch_s_branchI subtype_weakening
+ using calculation by presburger
+ thus ?thesis using check_branch_s_branchI by presburger
+ qed
+ qed(auto simp add: check_branch_s_branchI)
+next
+ case (check_ifI z \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v s1 s2 \<tau>)
+ show ?case proof(rule Typing.check_ifI)
+ have *:"atom z \<sharp> t2" using subtype_fresh_tau[of z \<tau> \<Gamma> ] check_ifI fresh_prodN by auto
+ thus \<open>atom z \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>, \<Delta>, v, s1, s2, t2)\<close> using check_ifI fresh_prodN by auto
+ show \<open>\<Theta> ; \<B> ; \<Gamma> \<turnstile> v \<Leftarrow> \<lbrace> z : B_bool | TRUE \<rbrace>\<close> using check_ifI by auto
+ show \<open> \<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile> s1 \<Leftarrow> \<lbrace> z : b_of t2 | [ v ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of t2 z \<rbrace>\<close>
+ using check_ifI subtype_if1 fresh_prodN base_for_lit.simps b_of.simps * check_v_wf by metis
+
+ show \<open> \<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile> s2 \<Leftarrow> \<lbrace> z : b_of t2 | [ v ]\<^sup>c\<^sup>e == [ [ L_false ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of t2 z \<rbrace>\<close>
+ using check_ifI subtype_if1 fresh_prodN base_for_lit.simps b_of.simps * check_v_wf by metis
+ qed
+next
+ case (check_assertI x \<Theta> \<Phi> \<B> \<Gamma> \<Delta> c \<tau> s)
+
+ show ?case proof
+ have "atom x \<sharp> t2" using subtype_fresh_tau[OF _ _ \<open>\<Theta> ; \<B> ; \<Gamma> \<turnstile> \<tau> \<lesssim> t2\<close>] check_assertI fresh_prodN by simp
+ thus "atom x \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>, \<Delta>, c, t2, s)" using subtype_fresh_tau check_assertI fresh_prodN by simp
+ have "\<Theta> ; \<B> ; (x, B_bool, c) #\<^sub>\<Gamma> \<Gamma> \<turnstile> \<tau> \<lesssim> t2" apply(rule subtype_weakening)
+ using check_assertI apply simp
+ using toSet.simps apply blast
+ using check_assertI check_s_wf by simp
+ thus "\<Theta> ; \<Phi> ; \<B> ; (x, B_bool, c) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile> s \<Leftarrow> t2" using check_assertI by simp
+ show "\<Theta> ; \<B> ; \<Gamma> \<Turnstile> c " using check_assertI by auto
+ show "\<Theta> ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> " using check_assertI by auto
+ qed
+next
+ case (check_let2I x P \<Phi> \<B> G \<Delta> t s1 \<tau> s2 )
+ have "wfG P \<B> ((x, b_of t, c_of t x) #\<^sub>\<Gamma> G)"
+ using check_let2I check_s_wf by metis
+ moreover have "toSet G \<subseteq> toSet ((x, b_of t, c_of t x) #\<^sub>\<Gamma> G)" by auto
+ ultimately have *:"P ; \<B> ; (x, b_of t, c_of t x) #\<^sub>\<Gamma> G \<turnstile> \<tau> \<lesssim> t2" using check_let2I subtype_weakening by metis
+ show ?case proof(rule Typing.check_let2I)
+ have "atom x \<sharp> t2" using subtype_fresh_tau[of x \<tau> ] check_let2I fresh_prodN by metis
+ thus "atom x \<sharp> (P, \<Phi>, \<B>, G, \<Delta>, t, s1, t2)" using check_let2I fresh_prodN by metis
+ show "P ; \<Phi> ; \<B> ; G ; \<Delta> \<turnstile> s1 \<Leftarrow> t" using check_let2I by blast
+ show "P ; \<Phi> ; \<B> ;(x, b_of t, c_of t x ) #\<^sub>\<Gamma> G ; \<Delta> \<turnstile> s2 \<Leftarrow> t2" using check_let2I * by blast
+ qed
+next
+ case (check_varI u \<Theta> \<Phi> \<B> \<Gamma> \<Delta> \<tau>' v \<tau> s)
+ show ?case proof(rule Typing.check_varI)
+ have "atom u \<sharp> t2" using u_fresh_t by auto
+ thus \<open>atom u \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>, \<Delta>, \<tau>', v, t2)\<close> using check_varI fresh_prodN by auto
+ show \<open>\<Theta> ; \<B> ; \<Gamma> \<turnstile> v \<Leftarrow> \<tau>'\<close> using check_varI by auto
+ show \<open> \<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; (u, \<tau>') #\<^sub>\<Delta> \<Delta> \<turnstile> s \<Leftarrow> t2\<close> using check_varI by auto
+ qed
+next
+ case (check_assignI \<Delta> u \<tau> P G v z \<tau>')
+ then show ?case using Typing.check_assignI by (meson subtype_trans)
+next
+ case (check_whileI \<Delta> G P s1 z s2 \<tau>')
+ then show ?case using Typing.check_whileI by (meson subtype_trans)
+next
+ case (check_seqI \<Delta> G P s1 z s2 \<tau>)
+ then show ?case using Typing.check_seqI by blast
+next
+ case (check_caseI \<Delta> \<Gamma> \<Theta> tid cs \<tau> v z)
+ then show ?case using Typing.check_caseI subtype_trans by meson
+
+qed
+
+lemma subtype_let:
+ fixes s'::s and cs::branch_s and css::branch_list and v::v
+ shows "\<Theta> ; \<Phi> ; \<B> ; GNil ; \<Delta> \<turnstile> AS_let x e\<^sub>1 s \<Leftarrow> \<tau> \<Longrightarrow> \<Theta> ; \<Phi> ; \<B> ; GNil ; \<Delta> \<turnstile> e\<^sub>1 \<Rightarrow> \<tau>\<^sub>1 \<Longrightarrow>
+ \<Theta> ; \<Phi> ; \<B> ; GNil ; \<Delta> \<turnstile> e\<^sub>2 \<Rightarrow> \<tau>\<^sub>2 \<Longrightarrow> \<Theta> ; \<B> ; GNil \<turnstile> \<tau>\<^sub>2 \<lesssim> \<tau>\<^sub>1 \<Longrightarrow> \<Theta> ; \<Phi> ; \<B> ; GNil ; \<Delta> \<turnstile> AS_let x e\<^sub>2 s \<Leftarrow> \<tau>" and
+ "check_branch_s \<Theta> \<Phi> {||} GNil \<Delta> tid dc const v cs \<tau> \<Longrightarrow> True" and
+ "check_branch_list \<Theta> \<Phi> {||} \<Gamma> \<Delta> tid dclist v css \<tau> \<Longrightarrow> True"
+proof(nominal_induct GNil \<Delta> "AS_let x e\<^sub>1 s" \<tau> and \<tau> and \<tau> avoiding: e\<^sub>2 \<tau>\<^sub>1 \<tau>\<^sub>2
+ rule: check_s_check_branch_s_check_branch_list.strong_induct)
+ case (check_letI x1 \<Theta> \<Phi> \<B> \<Delta> \<tau>1 z1 s1 b1 c1)
+ obtain z2 and b2 and c2 where t2:"\<tau>\<^sub>2 = \<lbrace> z2 : b2 | c2 \<rbrace> \<and> atom z2 \<sharp> (x1, \<Theta>, \<Phi>, \<B>, GNil, \<Delta>, e\<^sub>2, \<tau>1, s1) "
+ using obtain_fresh_z by metis
+
+ obtain z1a and b1a and c1a where t1:"\<tau>\<^sub>1 = \<lbrace> z1a : b1a | c1a \<rbrace> \<and> atom z1a \<sharp> x1" using infer_e_uniqueness check_letI by metis
+ hence t3: " \<lbrace> z1a : b1a | c1a \<rbrace> = \<lbrace> z1 : b1 | c1 \<rbrace> " using infer_e_uniqueness check_letI by metis
+
+ have beq: "b1a = b2 \<and> b2 = b1" using check_letI subtype_eq_base t1 t2 t3 by metis
+
+ have " \<Theta> ; \<Phi> ; \<B> ; GNil ; \<Delta> \<turnstile> AS_let x1 e\<^sub>2 s1 \<Leftarrow> \<tau>1" proof
+ show \<open>atom x1 \<sharp> (\<Theta>, \<Phi>, \<B>, GNil, \<Delta>, e\<^sub>2, \<tau>1)\<close> using check_letI t2 fresh_prodN by metis
+ show \<open>atom z2 \<sharp> (x1, \<Theta>, \<Phi>, \<B>, GNil, \<Delta>, e\<^sub>2, \<tau>1, s1)\<close> using check_letI t2 using check_letI t2 fresh_prodN by metis
+ show \<open>\<Theta> ; \<Phi> ; \<B> ; GNil ; \<Delta> \<turnstile> e\<^sub>2 \<Rightarrow> \<lbrace> z2 : b2 | c2 \<rbrace>\<close> using check_letI t2 by metis
+
+ have \<open> \<Theta> ; \<Phi> ; \<B> ; GNil@(x1, b2, c2[z2::=[ x1 ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> GNil ; \<Delta> \<turnstile> s1 \<Leftarrow> \<tau>1\<close>
+ proof(rule ctx_subtype_s)
+ have "c1a[z1a::=[ x1 ]\<^sup>v]\<^sub>c\<^sub>v = c1[z1::=[ x1 ]\<^sup>v]\<^sub>c\<^sub>v" using subst_v_flip_eq_two subst_v_c_def t3 \<tau>.eq_iff by metis
+ thus \<open> \<Theta> ; \<Phi> ; \<B> ; GNil @ (x1, b2, c1a[z1a::=[ x1 ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> GNil ; \<Delta> \<turnstile> s1 \<Leftarrow> \<tau>1\<close> using check_letI beq append_g.simps subst_defs by metis
+ show \<open>\<Theta> ; \<B> ; GNil \<turnstile> \<lbrace> z2 : b2 | c2 \<rbrace> \<lesssim> \<lbrace> z1a : b2 | c1a \<rbrace>\<close> using check_letI beq t1 t2 by metis
+ have "atom x1 \<sharp> c2" using t2 check_letI \<tau>_fresh_c fresh_prodN by blast
+ moreover have "atom x1 \<sharp> c1a" using t1 check_letI \<tau>_fresh_c fresh_prodN by blast
+ ultimately show \<open>atom x1 \<sharp> (z1a, z2, c1a, c2)\<close> using t1 t2 fresh_prodN fresh_x_neq by metis
+ qed
+
+ thus \<open> \<Theta> ; \<Phi> ; \<B> ; (x1, b2, c2[z2::=[ x1 ]\<^sup>v]\<^sub>v) #\<^sub>\<Gamma> GNil ; \<Delta> \<turnstile> s1 \<Leftarrow> \<tau>1\<close> using append_g.simps subst_defs by metis
+ qed
+
+ moreover have "AS_let x1 e\<^sub>2 s1 = AS_let x e\<^sub>2 s" using check_letI s_branch_s_branch_list.eq_iff by metis
+
+ ultimately show ?case by metis
+
+qed(auto+)
+
end
\ No newline at end of file
diff --git a/thys/MiniSail/MiniSail.thy b/thys/MiniSail/MiniSail.thy
--- a/thys/MiniSail/MiniSail.thy
+++ b/thys/MiniSail/MiniSail.thy
@@ -1,10 +1,10 @@
-(*<*)
-theory MiniSail
- imports
- "Safety"
-begin
-(*>*)
-
-(*<*)
-end
+(*<*)
+theory MiniSail
+ imports
+ "Safety"
+begin
+(*>*)
+
+(*<*)
+end
(*>*)
\ No newline at end of file
diff --git a/thys/MiniSail/Nominal-Utils.thy b/thys/MiniSail/Nominal-Utils.thy
--- a/thys/MiniSail/Nominal-Utils.thy
+++ b/thys/MiniSail/Nominal-Utils.thy
@@ -1,586 +1,586 @@
-(*<*)
-theory "Nominal-Utils"
-imports Nominal2.Nominal2 "HOL-Library.AList"
-begin
-(*>*)
-
-chapter \<open>Prelude\<close>
-text \<open>Some useful Nominal lemmas. Many of these are from Launchbury.Nominal-Utils.\<close>
-
-section \<open>Lemmas helping with equivariance proofs\<close>
-
-lemma perm_rel_lemma:
- assumes "\<And> \<pi> x y. r (\<pi> \<bullet> x) (\<pi> \<bullet> y) \<Longrightarrow> r x y"
- shows "r (\<pi> \<bullet> x) (\<pi> \<bullet> y) \<longleftrightarrow> r x y" (is "?l \<longleftrightarrow> ?r")
-by (metis (full_types) assms permute_minus_cancel(2))
-
-lemma perm_rel_lemma2:
- assumes "\<And> \<pi> x y. r x y \<Longrightarrow> r (\<pi> \<bullet> x) (\<pi> \<bullet> y)"
- shows "r x y \<longleftrightarrow> r (\<pi> \<bullet> x) (\<pi> \<bullet> y)" (is "?l \<longleftrightarrow> ?r")
-by (metis (full_types) assms permute_minus_cancel(2))
-
-lemma fun_eqvtI:
- assumes f_eqvt[eqvt]: "(\<And> p x. p \<bullet> (f x) = f (p \<bullet> x))"
- shows "p \<bullet> f = f" by perm_simp rule
-
-lemma eqvt_at_apply:
- assumes "eqvt_at f x"
- shows "(p \<bullet> f) x = f x"
-by (metis (opaque_lifting, no_types) assms eqvt_at_def permute_fun_def permute_minus_cancel(1))
-
-lemma eqvt_at_apply':
- assumes "eqvt_at f x"
- shows "p \<bullet> f x = f (p \<bullet> x)"
-by (metis (opaque_lifting, no_types) assms eqvt_at_def)
-
-lemma eqvt_at_apply'':
- assumes "eqvt_at f x"
- shows "(p \<bullet> f) (p \<bullet> x) = f (p \<bullet> x)"
-by (metis (opaque_lifting, no_types) assms eqvt_at_def permute_fun_def permute_minus_cancel(1))
-
-lemma size_list_eqvt[eqvt]: "p \<bullet> size_list f x = size_list (p \<bullet> f) (p \<bullet> x)"
-proof (induction x)
- case (Cons x xs)
- have "f x = p \<bullet> (f x)" by (simp add: permute_pure)
- also have "... = (p \<bullet> f) (p \<bullet> x)" by simp
- with Cons
- show ?case by (auto simp add: permute_pure)
-qed simp
-
-section \<open> Freshness via equivariance \<close>
-
-lemma eqvt_fresh_cong1: "(\<And>p x. p \<bullet> (f x) = f (p \<bullet> x)) \<Longrightarrow> a \<sharp> x \<Longrightarrow> a \<sharp> f x "
- apply (rule fresh_fun_eqvt_app[of f])
- apply (rule eqvtI)
- apply (rule eq_reflection)
- apply (rule ext)
- apply (metis permute_fun_def permute_minus_cancel(1))
- apply assumption
- done
-
-lemma eqvt_fresh_cong2:
- assumes eqvt: "(\<And>p x y. p \<bullet> (f x y) = f (p \<bullet> x) (p \<bullet> y))"
- and fresh1: "a \<sharp> x" and fresh2: "a \<sharp> y"
- shows "a \<sharp> f x y"
-proof-
- have "eqvt (\<lambda> (x,y). f x y)"
- using eqvt
- apply (- , auto simp add: eqvt_def)
- by (rule ext, auto, metis permute_minus_cancel(1))
- moreover
- have "a \<sharp> (x, y)" using fresh1 fresh2 by auto
- ultimately
- have "a \<sharp> (\<lambda> (x,y). f x y) (x, y)" by (rule fresh_fun_eqvt_app)
- thus ?thesis by simp
-qed
-
-lemma eqvt_fresh_star_cong1:
- assumes eqvt: "(\<And>p x. p \<bullet> (f x) = f (p \<bullet> x))"
- and fresh1: "a \<sharp>* x"
- shows "a \<sharp>* f x"
- by (metis fresh_star_def eqvt_fresh_cong1 assms)
-
-lemma eqvt_fresh_star_cong2:
- assumes eqvt: "(\<And>p x y. p \<bullet> (f x y) = f (p \<bullet> x) (p \<bullet> y))"
- and fresh1: "a \<sharp>* x" and fresh2: "a \<sharp>* y"
- shows "a \<sharp>* f x y"
- by (metis fresh_star_def eqvt_fresh_cong2 assms)
-
-lemma eqvt_fresh_cong3:
- assumes eqvt: "(\<And>p x y z. p \<bullet> (f x y z) = f (p \<bullet> x) (p \<bullet> y) (p \<bullet> z))"
- and fresh1: "a \<sharp> x" and fresh2: "a \<sharp> y" and fresh3: "a \<sharp> z"
- shows "a \<sharp> f x y z"
-proof-
- have "eqvt (\<lambda> (x,y,z). f x y z)"
- using eqvt
- apply (- , auto simp add: eqvt_def)
- by(rule ext, auto, metis permute_minus_cancel(1))
- moreover
- have "a \<sharp> (x, y, z)" using fresh1 fresh2 fresh3 by auto
- ultimately
- have "a \<sharp> (\<lambda> (x,y,z). f x y z) (x, y, z)" by (rule fresh_fun_eqvt_app)
- thus ?thesis by simp
-qed
-
-lemma eqvt_fresh_star_cong3:
- assumes eqvt: "(\<And>p x y z. p \<bullet> (f x y z) = f (p \<bullet> x) (p \<bullet> y) (p \<bullet> z))"
- and fresh1: "a \<sharp>* x" and fresh2: "a \<sharp>* y" and fresh3: "a \<sharp>* z"
- shows "a \<sharp>* f x y z"
- by (metis fresh_star_def eqvt_fresh_cong3 assms)
-
-section \<open> Additional simplification rules \<close>
-
-lemma not_self_fresh[simp]: "atom x \<sharp> x \<longleftrightarrow> False"
- by (metis fresh_at_base(2))
-
-lemma fresh_star_singleton: "{ x } \<sharp>* e \<longleftrightarrow> x \<sharp> e"
- by (simp add: fresh_star_def)
-
-section \<open> Additional equivariance lemmas \<close>
-
-lemma eqvt_cases:
- fixes f x \<pi>
- assumes eqvt: "\<And>x. \<pi> \<bullet> f x = f (\<pi> \<bullet> x)"
- obtains "f x" "f (\<pi> \<bullet> x)" | "\<not> f x " " \<not> f (\<pi> \<bullet> x)"
- using assms[symmetric]
- by (cases "f x") auto
-
-lemma range_eqvt: "\<pi> \<bullet> range Y = range (\<pi> \<bullet> Y)"
- unfolding image_eqvt UNIV_eqvt ..
-
-lemma case_option_eqvt[eqvt]:
- "\<pi> \<bullet> case_option d f x = case_option (\<pi> \<bullet> d) (\<pi> \<bullet> f) (\<pi> \<bullet> x)"
- by(cases x)(simp_all)
-
-lemma supp_option_eqvt:
- "supp (case_option d f x) \<subseteq> supp d \<union> supp f \<union> supp x"
- apply (cases x)
- apply (auto simp add: supp_Some )
- apply (metis (mono_tags) Un_iff subsetCE supp_fun_app)
- done
-
-lemma funpow_eqvt[simp,eqvt]:
- "\<pi> \<bullet> ((f :: 'a \<Rightarrow> 'a::pt) ^^ n) = (\<pi> \<bullet> f) ^^ (\<pi> \<bullet> n)"
- by (induct n,simp, rule ext, simp, perm_simp,simp)
-
-lemma delete_eqvt[eqvt]:
- "\<pi> \<bullet> AList.delete x \<Gamma> = AList.delete (\<pi> \<bullet> x) (\<pi> \<bullet> \<Gamma>)"
-by (induct \<Gamma>, auto)
-
-lemma restrict_eqvt[eqvt]:
- "\<pi> \<bullet> AList.restrict S \<Gamma> = AList.restrict (\<pi> \<bullet> S) (\<pi> \<bullet> \<Gamma>)"
-unfolding AList.restrict_eq by perm_simp rule
-
-lemma supp_restrict:
- "supp (AList.restrict S \<Gamma>) \<subseteq> supp \<Gamma>"
- by (induction \<Gamma>) (auto simp add: supp_Pair supp_Cons)
-
-lemma clearjunk_eqvt[eqvt]:
- "\<pi> \<bullet> AList.clearjunk \<Gamma> = AList.clearjunk (\<pi> \<bullet> \<Gamma>)"
- by (induction \<Gamma> rule: clearjunk.induct) auto
-
-lemma map_ran_eqvt[eqvt]:
- "\<pi> \<bullet> map_ran f \<Gamma> = map_ran (\<pi> \<bullet> f) (\<pi> \<bullet> \<Gamma>)"
-by (induct \<Gamma>, auto)
-
-lemma dom_perm:
- "dom (\<pi> \<bullet> f) = \<pi> \<bullet> (dom f)"
- unfolding dom_def by (perm_simp) (simp)
-
-lemmas dom_perm_rev[simp,eqvt] = dom_perm[symmetric]
-
-lemma ran_perm[simp]:
- "\<pi> \<bullet> (ran f) = ran (\<pi> \<bullet> f)"
- unfolding ran_def by (perm_simp) (simp)
-
-lemma map_add_eqvt[eqvt]:
- "\<pi> \<bullet> (m1 ++ m2) = (\<pi> \<bullet> m1) ++ (\<pi> \<bullet> m2)"
- unfolding map_add_def
- by (perm_simp, rule)
-
-lemma map_of_eqvt[eqvt]:
- "\<pi> \<bullet> map_of l = map_of (\<pi> \<bullet> l)"
- by (induct l, simp add: permute_fun_def,simp,perm_simp,auto)
-
-lemma concat_eqvt[eqvt]: "\<pi> \<bullet> concat l = concat (\<pi> \<bullet> l)"
- by (induction l)(auto simp add: append_eqvt)
-
-lemma tranclp_eqvt[eqvt]: "\<pi> \<bullet> tranclp P v\<^sub>1 v\<^sub>2 = tranclp (\<pi> \<bullet> P) (\<pi> \<bullet> v\<^sub>1) (\<pi> \<bullet> v\<^sub>2)"
- unfolding tranclp_def by perm_simp rule
-
-lemma rtranclp_eqvt[eqvt]: "\<pi> \<bullet> rtranclp P v\<^sub>1 v\<^sub>2 = rtranclp (\<pi> \<bullet> P) (\<pi> \<bullet> v\<^sub>1) (\<pi> \<bullet> v\<^sub>2)"
- unfolding rtranclp_def by perm_simp rule
-
-lemma Set_filter_eqvt[eqvt]: "\<pi> \<bullet> Set.filter P S = Set.filter (\<pi> \<bullet> P) (\<pi> \<bullet> S)"
- unfolding Set.filter_def
- by perm_simp rule
-
-lemma Sigma_eqvt'[eqvt]: "\<pi> \<bullet> Sigma = Sigma"
- apply (rule ext)
- apply (rule ext)
- apply (subst permute_fun_def)
- apply (subst permute_fun_def)
- unfolding Sigma_def
- apply perm_simp
- apply (simp add: permute_self)
- done
-
-lemma override_on_eqvt[eqvt]:
- "\<pi> \<bullet> (override_on m1 m2 S) = override_on (\<pi> \<bullet> m1) (\<pi> \<bullet> m2) (\<pi> \<bullet> S)"
- by (auto simp add: override_on_def )
-
-lemma card_eqvt[eqvt]:
- "\<pi> \<bullet> (card S) = card (\<pi> \<bullet> S)"
-by (cases "finite S", induct rule: finite_induct) (auto simp add: card_insert_if mem_permute_iff permute_pure)
-
-(* Helper lemmas provided by Christian Urban *)
-
-lemma Projl_permute:
- assumes a: "\<exists>y. f = Inl y"
- shows "(p \<bullet> (Sum_Type.projl f)) = Sum_Type.projl (p \<bullet> f)"
-using a by auto
-
-lemma Projr_permute:
- assumes a: "\<exists>y. f = Inr y"
- shows "(p \<bullet> (Sum_Type.projr f)) = Sum_Type.projr (p \<bullet> f)"
-using a by auto
-
-section \<open> Freshness lemmas \<close>
-
-lemma fresh_list_elem:
- assumes "a \<sharp> \<Gamma>"
- and "e \<in> set \<Gamma>"
- shows "a \<sharp> e"
-using assms
-by(induct \<Gamma>)(auto simp add: fresh_Cons)
-
-lemma set_not_fresh:
- "x \<in> set L \<Longrightarrow> \<not>(atom x \<sharp> L)"
- by (metis fresh_list_elem not_self_fresh)
-
-lemma pure_fresh_star[simp]: "a \<sharp>* (x :: 'a :: pure)"
- by (simp add: fresh_star_def pure_fresh)
-
-lemma supp_set_mem: "x \<in> set L \<Longrightarrow> supp x \<subseteq> supp L"
- by (induct L) (auto simp add: supp_Cons)
-
-lemma set_supp_mono: "set L \<subseteq> set L2 \<Longrightarrow> supp L \<subseteq> supp L2"
- by (induct L)(auto simp add: supp_Cons supp_Nil dest:supp_set_mem)
-
-lemma fresh_star_at_base:
- fixes x :: "'a :: at_base"
- shows "S \<sharp>* x \<longleftrightarrow> atom x \<notin> S"
- by (metis fresh_at_base(2) fresh_star_def)
-
-section \<open> Freshness and support for subsets of variables \<close>
-
-lemma supp_mono: "finite (B::'a::fs set) \<Longrightarrow> A \<subseteq> B \<Longrightarrow> supp A \<subseteq> supp B"
- by (metis infinite_super subset_Un_eq supp_of_finite_union)
-
-lemma fresh_subset:
- "finite B \<Longrightarrow> x \<sharp> (B :: 'a::at_base set) \<Longrightarrow> A \<subseteq> B \<Longrightarrow> x \<sharp> A"
- by (auto dest:supp_mono simp add: fresh_def)
-
-lemma fresh_star_subset:
- "finite B \<Longrightarrow> x \<sharp>* (B :: 'a::at_base set) \<Longrightarrow> A \<subseteq> B \<Longrightarrow> x \<sharp>* A"
- by (metis fresh_star_def fresh_subset)
-
-lemma fresh_star_set_subset:
- "x \<sharp>* (B :: 'a::at_base list) \<Longrightarrow> set A \<subseteq> set B \<Longrightarrow> x \<sharp>* A"
- by (metis fresh_star_set fresh_star_subset[OF finite_set])
-
-section \<open> The set of free variables of an expression \<close>
-
-definition fv :: "'a::pt \<Rightarrow> 'b::at_base set"
- where "fv e = {v. atom v \<in> supp e}"
-
-lemma fv_eqvt[simp,eqvt]: "\<pi> \<bullet> (fv e) = fv (\<pi> \<bullet> e)"
- unfolding fv_def by simp
-
-lemma fv_Nil[simp]: "fv [] = {}"
- by (auto simp add: fv_def supp_Nil)
-lemma fv_Cons[simp]: "fv (x # xs) = fv x \<union> fv xs"
- by (auto simp add: fv_def supp_Cons)
-lemma fv_Pair[simp]: "fv (x, y) = fv x \<union> fv y"
- by (auto simp add: fv_def supp_Pair)
-lemma fv_append[simp]: "fv (x @ y) = fv x \<union> fv y"
- by (auto simp add: fv_def supp_append)
-lemma fv_at_base[simp]: "fv a = {a::'a::at_base}"
- by (auto simp add: fv_def supp_at_base)
-lemma fv_pure[simp]: "fv (a::'a::pure) = {}"
- by (auto simp add: fv_def pure_supp)
-
-lemma fv_set_at_base[simp]: "fv (l :: ('a :: at_base) list) = set l"
- by (induction l) auto
-
-lemma flip_not_fv: "a \<notin> fv x \<Longrightarrow> b \<notin> fv x \<Longrightarrow> (a \<leftrightarrow> b) \<bullet> x = x"
- by (metis flip_def fresh_def fv_def mem_Collect_eq swap_fresh_fresh)
-
-lemma fv_not_fresh: "atom x \<sharp> e \<longleftrightarrow> x \<notin> fv e"
- unfolding fv_def fresh_def by blast
-
-lemma fresh_fv: "finite (fv e :: 'a set) \<Longrightarrow> atom (x :: ('a::at_base)) \<sharp> (fv e :: 'a set) \<longleftrightarrow> atom x \<sharp> e"
- unfolding fv_def fresh_def
- by (auto simp add: supp_finite_set_at_base)
-
-lemma finite_fv[simp]: "finite (fv (e::'a::fs) :: ('b::at_base) set)"
-proof-
- have "finite (supp e)" by (metis finite_supp)
- hence "finite (atom -` supp e :: 'b set)"
- apply (rule finite_vimageI)
- apply (rule inj_onI)
- apply (simp)
- done
- moreover
- have "(atom -` supp e :: 'b set) = fv e" unfolding fv_def by auto
- ultimately
- show ?thesis by simp
-qed
-
-definition fv_list :: "'a::fs \<Rightarrow> 'b::at_base list"
- where "fv_list e = (SOME l. set l = fv e)"
-
-lemma set_fv_list[simp]: "set (fv_list e) = (fv e :: ('b::at_base) set)"
-proof-
- have "finite (fv e :: 'b set)" by (rule finite_fv)
- from finite_list[OF finite_fv]
- obtain l where "set l = (fv e :: 'b set)"..
- thus ?thesis
- unfolding fv_list_def by (rule someI)
-qed
-
-lemma fresh_fv_list[simp]:
- "a \<sharp> (fv_list e :: 'b::at_base list) \<longleftrightarrow> a \<sharp> (fv e :: 'b::at_base set)"
-proof-
- have "a \<sharp> (fv_list e :: 'b::at_base list) \<longleftrightarrow> a \<sharp> set (fv_list e :: 'b::at_base list)"
- by (rule fresh_set[symmetric])
- also have "\<dots> \<longleftrightarrow> a \<sharp> (fv e :: 'b::at_base set)" by simp
- finally show ?thesis.
-qed
-
-section \<open> Other useful lemmas \<close>
-
-lemma pure_permute_id: "permute p = (\<lambda> x. (x::'a::pure))"
- by rule (simp add: permute_pure)
-
-lemma supp_set_elem_finite:
- assumes "finite S"
- and "(m::'a::fs) \<in> S"
- and "y \<in> supp m"
- shows "y \<in> supp S"
- using assms supp_of_finite_sets
- by auto
-
-lemmas fresh_star_Cons = fresh_star_list(2)
-
-lemma mem_permute_set:
- shows "x \<in> p \<bullet> S \<longleftrightarrow> (- p \<bullet> x) \<in> S"
- by (metis mem_permute_iff permute_minus_cancel(2))
-
-lemma flip_set_both_not_in:
- assumes "x \<notin> S" and "x' \<notin> S"
- shows "((x' \<leftrightarrow> x) \<bullet> S) = S"
- unfolding permute_set_def
- by (auto) (metis assms flip_at_base_simps(3))+
-
-lemma inj_atom: "inj atom" by (metis atom_eq_iff injI)
-
-lemmas image_Int[OF inj_atom, simp]
-
-lemma eqvt_uncurry: "eqvt f \<Longrightarrow> eqvt (case_prod f)"
- unfolding eqvt_def
- by perm_simp simp
-
-lemma supp_fun_app_eqvt2:
- assumes a: "eqvt f"
- shows "supp (f x y) \<subseteq> supp x \<union> supp y"
-proof-
- from supp_fun_app_eqvt[OF eqvt_uncurry [OF a]]
- have "supp (case_prod f (x,y)) \<subseteq> supp (x,y)".
- thus ?thesis by (simp add: supp_Pair)
-qed
-
-lemma supp_fun_app_eqvt3:
- assumes a: "eqvt f"
- shows "supp (f x y z) \<subseteq> supp x \<union> supp y \<union> supp z"
-proof-
- from supp_fun_app_eqvt2[OF eqvt_uncurry [OF a]]
- have "supp (case_prod f (x,y) z) \<subseteq> supp (x,y) \<union> supp z".
- thus ?thesis by (simp add: supp_Pair)
-qed
-
-(* Fighting eta-contraction *)
-lemma permute_0[simp]: "permute 0 = (\<lambda> x. x)"
- by auto
-lemma permute_comp[simp]: "permute x \<circ> permute y = permute (x + y)" by auto
-
-lemma map_permute: "map (permute p) = permute p"
- apply rule
- apply (induct_tac x)
- apply auto
- done
-
-lemma fresh_star_restrictA[intro]: "a \<sharp>* \<Gamma> \<Longrightarrow> a \<sharp>* AList.restrict V \<Gamma>"
- by (induction \<Gamma>) (auto simp add: fresh_star_Cons)
-
-lemma Abs_lst_Nil_eq[simp]: "[[]]lst. (x::'a::fs) = [xs]lst. x' \<longleftrightarrow> (([],x) = (xs, x'))"
- apply rule
- apply (frule Abs_lst_fcb2[where f = "\<lambda> x y _ . (x,y)" and as = "[]" and bs = "xs" and c = "()"])
- apply (auto simp add: fresh_star_def)
- done
-
-lemma Abs_lst_Nil_eq2[simp]: "[xs]lst. (x::'a::fs) = [[]]lst. x' \<longleftrightarrow> ((xs,x) = ([], x'))"
- by (subst eq_commute) auto
-
-lemma prod_cases8 [cases type]:
- obtains (fields) a b c d e f g h where "y = (a, b, c, d, e, f, g,h)"
- by (cases y, case_tac g) blast
-
-lemma prod_induct8 [case_names fields, induct type]:
- "(\<And>a b c d e f g h. P (a, b, c, d, e, f, g, h)) \<Longrightarrow> P x"
- by (cases x) blast
-
-lemma prod_cases9 [cases type]:
- obtains (fields) a b c d e f g h i where "y = (a, b, c, d, e, f, g,h,i)"
- by (cases y, case_tac h) blast
-
-lemma prod_induct9 [case_names fields, induct type]:
- "(\<And>a b c d e f g h i. P (a, b, c, d, e, f, g, h, i)) \<Longrightarrow> P x"
- by (cases x) blast
-
-named_theorems nominal_prod_simps
-
-named_theorems ms_fresh "Facts for helping with freshness proofs"
-
-lemma fresh_prod2[nominal_prod_simps,ms_fresh]: "x \<sharp> (a,b) = (x \<sharp> a \<and> x \<sharp> b )"
- using fresh_def supp_Pair by fastforce
-
-lemma fresh_prod3[nominal_prod_simps,ms_fresh]: "x \<sharp> (a,b,c) = (x \<sharp> a \<and> x \<sharp> b \<and> x \<sharp> c)"
- using fresh_def supp_Pair by fastforce
-
-lemma fresh_prod4[nominal_prod_simps,ms_fresh]: "x \<sharp> (a,b,c,d) = (x \<sharp> a \<and> x \<sharp> b \<and> x \<sharp> c \<and> x \<sharp> d)"
- using fresh_def supp_Pair by fastforce
-
-lemma fresh_prod5[nominal_prod_simps,ms_fresh]: "x \<sharp> (a,b,c,d,e) = (x \<sharp> a \<and> x \<sharp> b \<and> x \<sharp> c \<and> x \<sharp> d \<and> x \<sharp> e)"
- using fresh_def supp_Pair by fastforce
-
-lemma fresh_prod6[nominal_prod_simps,ms_fresh]: "x \<sharp> (a,b,c,d,e,f) = (x \<sharp> a \<and> x \<sharp> b \<and> x \<sharp> c \<and> x \<sharp> d \<and> x \<sharp> e \<and> x \<sharp> f)"
- using fresh_def supp_Pair by fastforce
-
-lemma fresh_prod7[nominal_prod_simps,ms_fresh]: "x \<sharp> (a,b,c,d,e,f,g) = (x \<sharp> a \<and> x \<sharp> b \<and> x \<sharp> c \<and> x \<sharp> d \<and> x \<sharp> e \<and> x \<sharp> f \<and> x \<sharp> g)"
- using fresh_def supp_Pair by fastforce
-
-lemma fresh_prod8[nominal_prod_simps,ms_fresh]: "x \<sharp> (a,b,c,d,e,f,g,h) = (x \<sharp> a \<and> x \<sharp> b \<and> x \<sharp> c \<and> x \<sharp> d \<and> x \<sharp> e \<and> x \<sharp> f \<and> x \<sharp> g \<and> x \<sharp> h )"
- using fresh_def supp_Pair by fastforce
-
-lemma fresh_prod9[nominal_prod_simps,ms_fresh]: "x \<sharp> (a,b,c,d,e,f,g,h,i) = (x \<sharp> a \<and> x \<sharp> b \<and> x \<sharp> c \<and> x \<sharp> d \<and> x \<sharp> e \<and> x \<sharp> f \<and> x \<sharp> g \<and> x \<sharp> h \<and> x \<sharp> i)"
- using fresh_def supp_Pair by fastforce
-
-lemma fresh_prod10[nominal_prod_simps,ms_fresh]: "x \<sharp> (a,b,c,d,e,f,g,h,i,j) = (x \<sharp> a \<and> x \<sharp> b \<and> x \<sharp> c \<and> x \<sharp> d \<and> x \<sharp> e \<and> x \<sharp> f \<and> x \<sharp> g \<and> x \<sharp> h \<and> x \<sharp> i \<and> x \<sharp> j)"
- using fresh_def supp_Pair by fastforce
-
-lemma fresh_prod12[nominal_prod_simps,ms_fresh]: "x \<sharp> (a,b,c,d,e,f,g,h,i,j,k,l) = (x \<sharp> a \<and> x \<sharp> b \<and> x \<sharp> c \<and> x \<sharp> d \<and> x \<sharp> e \<and> x \<sharp> f \<and> x \<sharp> g \<and> x \<sharp> h \<and> x \<sharp> i \<and> x \<sharp> j \<and> x \<sharp> k \<and> x \<sharp> l)"
- using fresh_def supp_Pair by fastforce
-
-lemmas fresh_prodN = fresh_Pair fresh_prod3 fresh_prod4 fresh_prod5 fresh_prod6 fresh_prod7 fresh_prod8 fresh_prod9 fresh_prod10 fresh_prod12
-
-lemma fresh_prod2I:
- fixes x and x1 and x2
- assumes "x \<sharp> x1" and "x \<sharp> x2"
- shows "x \<sharp> (x1,x2)" using fresh_prod2 assms by auto
-
-lemma fresh_prod3I:
- fixes x and x1 and x2 and x3
- assumes "x \<sharp> x1" and "x \<sharp> x2" and "x \<sharp> x3"
- shows "x \<sharp> (x1,x2,x3)" using fresh_prod3 assms by auto
-
-lemma fresh_prod4I:
- fixes x and x1 and x2 and x3 and x4
- assumes "x \<sharp> x1" and "x \<sharp> x2" and "x \<sharp> x3" and "x \<sharp> x4"
- shows "x \<sharp> (x1,x2,x3,x4)" using fresh_prod4 assms by auto
-
-lemma fresh_prod5I:
- fixes x and x1 and x2 and x3 and x4 and x5
- assumes "x \<sharp> x1" and "x \<sharp> x2" and "x \<sharp> x3" and "x \<sharp> x4" and "x \<sharp> x5"
- shows "x \<sharp> (x1,x2,x3,x4,x5)" using fresh_prod5 assms by auto
-
-lemma flip_collapse[simp]:
- fixes b1::"'a::pt" and bv1::"'b::at" and bv2::"'b::at"
- assumes "atom bv2 \<sharp> b1" and "atom c \<sharp> (bv1,bv2,b1)" and "bv1 \<noteq> bv2"
- shows "(bv2 \<leftrightarrow> c) \<bullet> (bv1 \<leftrightarrow> bv2) \<bullet> b1 = (bv1 \<leftrightarrow> c) \<bullet> b1"
-proof -
- have "c \<noteq> bv1" and "bv2 \<noteq> bv1" using assms by auto+
- hence "(bv2 \<leftrightarrow> c) + (bv1 \<leftrightarrow> bv2) + (bv2 \<leftrightarrow> c) = (bv1 \<leftrightarrow> c)" using flip_triple[of c bv1 bv2] flip_commute by metis
- hence "(bv2 \<leftrightarrow> c) \<bullet> (bv1 \<leftrightarrow> bv2) \<bullet> (bv2 \<leftrightarrow> c) \<bullet> b1 = (bv1 \<leftrightarrow> c) \<bullet> b1" using permute_plus by metis
- thus ?thesis using assms flip_fresh_fresh by force
-qed
-
-lemma triple_eqvt[simp]:
- "p \<bullet> (x, b,c) = (p \<bullet> x, p \<bullet> b , p \<bullet> c)"
-proof -
- have "(x,b,c) = (x,(b,c))" by simp
- thus ?thesis using Pair_eqvt by simp
-qed
-
-lemma lst_fst:
- fixes x::"'a::at" and t1::"'b::fs" and x'::"'a::at" and t2::"'c::fs"
- assumes " ([[atom x]]lst. (t1,t2) = [[atom x']]lst. (t1',t2'))"
- shows " ([[atom x]]lst. t1 = [[atom x']]lst. t1')"
-proof -
- have "(\<forall>c. atom c \<sharp> (t2,t2') \<longrightarrow> atom c \<sharp> (x, x', t1, t1') \<longrightarrow> (x \<leftrightarrow> c) \<bullet> t1 = (x' \<leftrightarrow> c) \<bullet> t1')"
- proof(rule,rule,rule)
- fix c::'a
- assume "atom c \<sharp> (t2,t2')" and "atom c \<sharp> (x, x', t1, t1')"
- hence "atom c \<sharp> (x, x', (t1,t2), (t1',t2'))" using fresh_prod2 by simp
- thus " (x \<leftrightarrow> c) \<bullet> t1 = (x' \<leftrightarrow> c) \<bullet> t1'" using assms Abs1_eq_iff_all(3) Pair_eqvt by simp
- qed
- thus ?thesis using Abs1_eq_iff_all(3)[of x t1 x' t1' "(t2,t2')"] by simp
-qed
-
-lemma lst_snd:
- fixes x::"'a::at" and t1::"'b::fs" and x'::"'a::at" and t2::"'c::fs"
- assumes " ([[atom x]]lst. (t1,t2) = [[atom x']]lst. (t1',t2'))"
- shows " ([[atom x]]lst. t2 = [[atom x']]lst. t2')"
-proof -
- have "(\<forall>c. atom c \<sharp> (t1,t1') \<longrightarrow> atom c \<sharp> (x, x', t2, t2') \<longrightarrow> (x \<leftrightarrow> c) \<bullet> t2 = (x' \<leftrightarrow> c) \<bullet> t2')"
- proof(rule,rule,rule)
- fix c::'a
- assume "atom c \<sharp> (t1,t1')" and "atom c \<sharp> (x, x', t2, t2')"
- hence "atom c \<sharp> (x, x', (t1,t2), (t1',t2'))" using fresh_prod2 by simp
- thus " (x \<leftrightarrow> c) \<bullet> t2 = (x' \<leftrightarrow> c) \<bullet> t2'" using assms Abs1_eq_iff_all(3) Pair_eqvt by simp
- qed
- thus ?thesis using Abs1_eq_iff_all(3)[of x t2 x' t2' "(t1,t1')"] by simp
-qed
-
-lemma lst_head_cons_pair:
- fixes y1::"'a ::at" and y2::"'a::at" and x1::"'b::fs" and x2::"'b::fs" and xs1::"('b::fs) list" and xs2::"('b::fs) list"
- assumes "[[atom y1]]lst. (x1 # xs1) = [[atom y2]]lst. (x2 # xs2)"
- shows "[[atom y1]]lst. (x1,xs1) = [[atom y2]]lst. (x2,xs2)"
-proof(subst Abs1_eq_iff_all(3)[of y1 "(x1,xs1)" y2 "(x2,xs2)"],rule,rule,rule)
- fix c::'a
- assume "atom c \<sharp> (x1#xs1,x2#xs2)" and "atom c \<sharp> (y1, y2, (x1, xs1), x2, xs2)"
- thus "(y1 \<leftrightarrow> c) \<bullet> (x1, xs1) = (y2 \<leftrightarrow> c) \<bullet> (x2, xs2)" using assms Abs1_eq_iff_all(3) by auto
-qed
-
-lemma lst_head_cons_neq_nil:
- fixes y1::"'a ::at" and y2::"'a::at" and x1::"'b::fs" and x2::"'b::fs" and xs1::"('b::fs) list" and xs2::"('b::fs) list"
- assumes "[[atom y1]]lst. (x1 # xs1) = [[atom y2]]lst. (xs2)"
- shows "xs2 \<noteq> []"
-proof
- assume as:"xs2 = []"
- thus False using Abs1_eq_iff(3)[of y1 "x1#xs1" y2 Nil] assms as by auto
-qed
-
-lemma lst_head_cons:
- fixes y1::"'a ::at" and y2::"'a::at" and x1::"'b::fs" and x2::"'b::fs" and xs1::"('b::fs) list" and xs2::"('b::fs) list"
- assumes "[[atom y1]]lst. (x1 # xs1) = [[atom y2]]lst. (x2 # xs2)"
- shows "[[atom y1]]lst. x1 = [[atom y2]]lst. x2" and "[[atom y1]]lst. xs1 = [[atom y2]]lst. xs2"
- using lst_head_cons_pair lst_fst lst_snd assms by metis+
-
-lemma lst_pure:
- fixes x1::"'a ::at" and t1::"'b::pure" and x2::"'a ::at" and t2::"'b::pure"
- assumes "[[atom x1]]lst. t1 = [[atom x2]]lst. t2"
- shows "t1=t2"
- using assms Abs1_eq_iff_all(3) pure_fresh flip_fresh_fresh
- by (metis Abs1_eq(3) permute_pure)
-
-lemma lst_supp:
- assumes "[[atom x1]]lst. t1 = [[atom x2]]lst. t2"
- shows "supp t1 - {atom x1} = supp t2 - {atom x2}"
-proof -
- have "supp ([[atom x1]]lst.t1) = supp ([[atom x2]]lst.t2)" using assms by auto
- thus ?thesis using Abs_finite_supp
- by (metis assms empty_set list.simps(15) supp_lst.simps)
-qed
-
-lemma lst_supp_subset:
- assumes "[[atom x1]]lst. t1 = [[atom x2]]lst. t2" and "supp t1 \<subseteq> {atom x1} \<union> B"
- shows "supp t2 \<subseteq> {atom x2} \<union> B"
- using assms lst_supp by fast
-
-lemma projl_inl_eqvt:
- fixes \<pi> :: perm
- shows "\<pi> \<bullet> (projl (Inl x)) = projl (Inl (\<pi> \<bullet> x))"
-unfolding projl_def Inl_eqvt by simp
-
-end
+(*<*)
+theory "Nominal-Utils"
+imports Nominal2.Nominal2 "HOL-Library.AList"
+begin
+(*>*)
+
+chapter \<open>Prelude\<close>
+text \<open>Some useful Nominal lemmas. Many of these are from Launchbury.Nominal-Utils.\<close>
+
+section \<open>Lemmas helping with equivariance proofs\<close>
+
+lemma perm_rel_lemma:
+ assumes "\<And> \<pi> x y. r (\<pi> \<bullet> x) (\<pi> \<bullet> y) \<Longrightarrow> r x y"
+ shows "r (\<pi> \<bullet> x) (\<pi> \<bullet> y) \<longleftrightarrow> r x y" (is "?l \<longleftrightarrow> ?r")
+by (metis (full_types) assms permute_minus_cancel(2))
+
+lemma perm_rel_lemma2:
+ assumes "\<And> \<pi> x y. r x y \<Longrightarrow> r (\<pi> \<bullet> x) (\<pi> \<bullet> y)"
+ shows "r x y \<longleftrightarrow> r (\<pi> \<bullet> x) (\<pi> \<bullet> y)" (is "?l \<longleftrightarrow> ?r")
+by (metis (full_types) assms permute_minus_cancel(2))
+
+lemma fun_eqvtI:
+ assumes f_eqvt[eqvt]: "(\<And> p x. p \<bullet> (f x) = f (p \<bullet> x))"
+ shows "p \<bullet> f = f" by perm_simp rule
+
+lemma eqvt_at_apply:
+ assumes "eqvt_at f x"
+ shows "(p \<bullet> f) x = f x"
+by (metis (opaque_lifting, no_types) assms eqvt_at_def permute_fun_def permute_minus_cancel(1))
+
+lemma eqvt_at_apply':
+ assumes "eqvt_at f x"
+ shows "p \<bullet> f x = f (p \<bullet> x)"
+by (metis (opaque_lifting, no_types) assms eqvt_at_def)
+
+lemma eqvt_at_apply'':
+ assumes "eqvt_at f x"
+ shows "(p \<bullet> f) (p \<bullet> x) = f (p \<bullet> x)"
+by (metis (opaque_lifting, no_types) assms eqvt_at_def permute_fun_def permute_minus_cancel(1))
+
+lemma size_list_eqvt[eqvt]: "p \<bullet> size_list f x = size_list (p \<bullet> f) (p \<bullet> x)"
+proof (induction x)
+ case (Cons x xs)
+ have "f x = p \<bullet> (f x)" by (simp add: permute_pure)
+ also have "... = (p \<bullet> f) (p \<bullet> x)" by simp
+ with Cons
+ show ?case by (auto simp add: permute_pure)
+qed simp
+
+section \<open> Freshness via equivariance \<close>
+
+lemma eqvt_fresh_cong1: "(\<And>p x. p \<bullet> (f x) = f (p \<bullet> x)) \<Longrightarrow> a \<sharp> x \<Longrightarrow> a \<sharp> f x "
+ apply (rule fresh_fun_eqvt_app[of f])
+ apply (rule eqvtI)
+ apply (rule eq_reflection)
+ apply (rule ext)
+ apply (metis permute_fun_def permute_minus_cancel(1))
+ apply assumption
+ done
+
+lemma eqvt_fresh_cong2:
+ assumes eqvt: "(\<And>p x y. p \<bullet> (f x y) = f (p \<bullet> x) (p \<bullet> y))"
+ and fresh1: "a \<sharp> x" and fresh2: "a \<sharp> y"
+ shows "a \<sharp> f x y"
+proof-
+ have "eqvt (\<lambda> (x,y). f x y)"
+ using eqvt
+ apply (- , auto simp add: eqvt_def)
+ by (rule ext, auto, metis permute_minus_cancel(1))
+ moreover
+ have "a \<sharp> (x, y)" using fresh1 fresh2 by auto
+ ultimately
+ have "a \<sharp> (\<lambda> (x,y). f x y) (x, y)" by (rule fresh_fun_eqvt_app)
+ thus ?thesis by simp
+qed
+
+lemma eqvt_fresh_star_cong1:
+ assumes eqvt: "(\<And>p x. p \<bullet> (f x) = f (p \<bullet> x))"
+ and fresh1: "a \<sharp>* x"
+ shows "a \<sharp>* f x"
+ by (metis fresh_star_def eqvt_fresh_cong1 assms)
+
+lemma eqvt_fresh_star_cong2:
+ assumes eqvt: "(\<And>p x y. p \<bullet> (f x y) = f (p \<bullet> x) (p \<bullet> y))"
+ and fresh1: "a \<sharp>* x" and fresh2: "a \<sharp>* y"
+ shows "a \<sharp>* f x y"
+ by (metis fresh_star_def eqvt_fresh_cong2 assms)
+
+lemma eqvt_fresh_cong3:
+ assumes eqvt: "(\<And>p x y z. p \<bullet> (f x y z) = f (p \<bullet> x) (p \<bullet> y) (p \<bullet> z))"
+ and fresh1: "a \<sharp> x" and fresh2: "a \<sharp> y" and fresh3: "a \<sharp> z"
+ shows "a \<sharp> f x y z"
+proof-
+ have "eqvt (\<lambda> (x,y,z). f x y z)"
+ using eqvt
+ apply (- , auto simp add: eqvt_def)
+ by(rule ext, auto, metis permute_minus_cancel(1))
+ moreover
+ have "a \<sharp> (x, y, z)" using fresh1 fresh2 fresh3 by auto
+ ultimately
+ have "a \<sharp> (\<lambda> (x,y,z). f x y z) (x, y, z)" by (rule fresh_fun_eqvt_app)
+ thus ?thesis by simp
+qed
+
+lemma eqvt_fresh_star_cong3:
+ assumes eqvt: "(\<And>p x y z. p \<bullet> (f x y z) = f (p \<bullet> x) (p \<bullet> y) (p \<bullet> z))"
+ and fresh1: "a \<sharp>* x" and fresh2: "a \<sharp>* y" and fresh3: "a \<sharp>* z"
+ shows "a \<sharp>* f x y z"
+ by (metis fresh_star_def eqvt_fresh_cong3 assms)
+
+section \<open> Additional simplification rules \<close>
+
+lemma not_self_fresh[simp]: "atom x \<sharp> x \<longleftrightarrow> False"
+ by (metis fresh_at_base(2))
+
+lemma fresh_star_singleton: "{ x } \<sharp>* e \<longleftrightarrow> x \<sharp> e"
+ by (simp add: fresh_star_def)
+
+section \<open> Additional equivariance lemmas \<close>
+
+lemma eqvt_cases:
+ fixes f x \<pi>
+ assumes eqvt: "\<And>x. \<pi> \<bullet> f x = f (\<pi> \<bullet> x)"
+ obtains "f x" "f (\<pi> \<bullet> x)" | "\<not> f x " " \<not> f (\<pi> \<bullet> x)"
+ using assms[symmetric]
+ by (cases "f x") auto
+
+lemma range_eqvt: "\<pi> \<bullet> range Y = range (\<pi> \<bullet> Y)"
+ unfolding image_eqvt UNIV_eqvt ..
+
+lemma case_option_eqvt[eqvt]:
+ "\<pi> \<bullet> case_option d f x = case_option (\<pi> \<bullet> d) (\<pi> \<bullet> f) (\<pi> \<bullet> x)"
+ by(cases x)(simp_all)
+
+lemma supp_option_eqvt:
+ "supp (case_option d f x) \<subseteq> supp d \<union> supp f \<union> supp x"
+ apply (cases x)
+ apply (auto simp add: supp_Some )
+ apply (metis (mono_tags) Un_iff subsetCE supp_fun_app)
+ done
+
+lemma funpow_eqvt[simp,eqvt]:
+ "\<pi> \<bullet> ((f :: 'a \<Rightarrow> 'a::pt) ^^ n) = (\<pi> \<bullet> f) ^^ (\<pi> \<bullet> n)"
+ by (induct n,simp, rule ext, simp, perm_simp,simp)
+
+lemma delete_eqvt[eqvt]:
+ "\<pi> \<bullet> AList.delete x \<Gamma> = AList.delete (\<pi> \<bullet> x) (\<pi> \<bullet> \<Gamma>)"
+by (induct \<Gamma>, auto)
+
+lemma restrict_eqvt[eqvt]:
+ "\<pi> \<bullet> AList.restrict S \<Gamma> = AList.restrict (\<pi> \<bullet> S) (\<pi> \<bullet> \<Gamma>)"
+unfolding AList.restrict_eq by perm_simp rule
+
+lemma supp_restrict:
+ "supp (AList.restrict S \<Gamma>) \<subseteq> supp \<Gamma>"
+ by (induction \<Gamma>) (auto simp add: supp_Pair supp_Cons)
+
+lemma clearjunk_eqvt[eqvt]:
+ "\<pi> \<bullet> AList.clearjunk \<Gamma> = AList.clearjunk (\<pi> \<bullet> \<Gamma>)"
+ by (induction \<Gamma> rule: clearjunk.induct) auto
+
+lemma map_ran_eqvt[eqvt]:
+ "\<pi> \<bullet> map_ran f \<Gamma> = map_ran (\<pi> \<bullet> f) (\<pi> \<bullet> \<Gamma>)"
+by (induct \<Gamma>, auto)
+
+lemma dom_perm:
+ "dom (\<pi> \<bullet> f) = \<pi> \<bullet> (dom f)"
+ unfolding dom_def by (perm_simp) (simp)
+
+lemmas dom_perm_rev[simp,eqvt] = dom_perm[symmetric]
+
+lemma ran_perm[simp]:
+ "\<pi> \<bullet> (ran f) = ran (\<pi> \<bullet> f)"
+ unfolding ran_def by (perm_simp) (simp)
+
+lemma map_add_eqvt[eqvt]:
+ "\<pi> \<bullet> (m1 ++ m2) = (\<pi> \<bullet> m1) ++ (\<pi> \<bullet> m2)"
+ unfolding map_add_def
+ by (perm_simp, rule)
+
+lemma map_of_eqvt[eqvt]:
+ "\<pi> \<bullet> map_of l = map_of (\<pi> \<bullet> l)"
+ by (induct l, simp add: permute_fun_def,simp,perm_simp,auto)
+
+lemma concat_eqvt[eqvt]: "\<pi> \<bullet> concat l = concat (\<pi> \<bullet> l)"
+ by (induction l)(auto simp add: append_eqvt)
+
+lemma tranclp_eqvt[eqvt]: "\<pi> \<bullet> tranclp P v\<^sub>1 v\<^sub>2 = tranclp (\<pi> \<bullet> P) (\<pi> \<bullet> v\<^sub>1) (\<pi> \<bullet> v\<^sub>2)"
+ unfolding tranclp_def by perm_simp rule
+
+lemma rtranclp_eqvt[eqvt]: "\<pi> \<bullet> rtranclp P v\<^sub>1 v\<^sub>2 = rtranclp (\<pi> \<bullet> P) (\<pi> \<bullet> v\<^sub>1) (\<pi> \<bullet> v\<^sub>2)"
+ unfolding rtranclp_def by perm_simp rule
+
+lemma Set_filter_eqvt[eqvt]: "\<pi> \<bullet> Set.filter P S = Set.filter (\<pi> \<bullet> P) (\<pi> \<bullet> S)"
+ unfolding Set.filter_def
+ by perm_simp rule
+
+lemma Sigma_eqvt'[eqvt]: "\<pi> \<bullet> Sigma = Sigma"
+ apply (rule ext)
+ apply (rule ext)
+ apply (subst permute_fun_def)
+ apply (subst permute_fun_def)
+ unfolding Sigma_def
+ apply perm_simp
+ apply (simp add: permute_self)
+ done
+
+lemma override_on_eqvt[eqvt]:
+ "\<pi> \<bullet> (override_on m1 m2 S) = override_on (\<pi> \<bullet> m1) (\<pi> \<bullet> m2) (\<pi> \<bullet> S)"
+ by (auto simp add: override_on_def )
+
+lemma card_eqvt[eqvt]:
+ "\<pi> \<bullet> (card S) = card (\<pi> \<bullet> S)"
+by (cases "finite S", induct rule: finite_induct) (auto simp add: card_insert_if mem_permute_iff permute_pure)
+
+(* Helper lemmas provided by Christian Urban *)
+
+lemma Projl_permute:
+ assumes a: "\<exists>y. f = Inl y"
+ shows "(p \<bullet> (Sum_Type.projl f)) = Sum_Type.projl (p \<bullet> f)"
+using a by auto
+
+lemma Projr_permute:
+ assumes a: "\<exists>y. f = Inr y"
+ shows "(p \<bullet> (Sum_Type.projr f)) = Sum_Type.projr (p \<bullet> f)"
+using a by auto
+
+section \<open> Freshness lemmas \<close>
+
+lemma fresh_list_elem:
+ assumes "a \<sharp> \<Gamma>"
+ and "e \<in> set \<Gamma>"
+ shows "a \<sharp> e"
+using assms
+by(induct \<Gamma>)(auto simp add: fresh_Cons)
+
+lemma set_not_fresh:
+ "x \<in> set L \<Longrightarrow> \<not>(atom x \<sharp> L)"
+ by (metis fresh_list_elem not_self_fresh)
+
+lemma pure_fresh_star[simp]: "a \<sharp>* (x :: 'a :: pure)"
+ by (simp add: fresh_star_def pure_fresh)
+
+lemma supp_set_mem: "x \<in> set L \<Longrightarrow> supp x \<subseteq> supp L"
+ by (induct L) (auto simp add: supp_Cons)
+
+lemma set_supp_mono: "set L \<subseteq> set L2 \<Longrightarrow> supp L \<subseteq> supp L2"
+ by (induct L)(auto simp add: supp_Cons supp_Nil dest:supp_set_mem)
+
+lemma fresh_star_at_base:
+ fixes x :: "'a :: at_base"
+ shows "S \<sharp>* x \<longleftrightarrow> atom x \<notin> S"
+ by (metis fresh_at_base(2) fresh_star_def)
+
+section \<open> Freshness and support for subsets of variables \<close>
+
+lemma supp_mono: "finite (B::'a::fs set) \<Longrightarrow> A \<subseteq> B \<Longrightarrow> supp A \<subseteq> supp B"
+ by (metis infinite_super subset_Un_eq supp_of_finite_union)
+
+lemma fresh_subset:
+ "finite B \<Longrightarrow> x \<sharp> (B :: 'a::at_base set) \<Longrightarrow> A \<subseteq> B \<Longrightarrow> x \<sharp> A"
+ by (auto dest:supp_mono simp add: fresh_def)
+
+lemma fresh_star_subset:
+ "finite B \<Longrightarrow> x \<sharp>* (B :: 'a::at_base set) \<Longrightarrow> A \<subseteq> B \<Longrightarrow> x \<sharp>* A"
+ by (metis fresh_star_def fresh_subset)
+
+lemma fresh_star_set_subset:
+ "x \<sharp>* (B :: 'a::at_base list) \<Longrightarrow> set A \<subseteq> set B \<Longrightarrow> x \<sharp>* A"
+ by (metis fresh_star_set fresh_star_subset[OF finite_set])
+
+section \<open> The set of free variables of an expression \<close>
+
+definition fv :: "'a::pt \<Rightarrow> 'b::at_base set"
+ where "fv e = {v. atom v \<in> supp e}"
+
+lemma fv_eqvt[simp,eqvt]: "\<pi> \<bullet> (fv e) = fv (\<pi> \<bullet> e)"
+ unfolding fv_def by simp
+
+lemma fv_Nil[simp]: "fv [] = {}"
+ by (auto simp add: fv_def supp_Nil)
+lemma fv_Cons[simp]: "fv (x # xs) = fv x \<union> fv xs"
+ by (auto simp add: fv_def supp_Cons)
+lemma fv_Pair[simp]: "fv (x, y) = fv x \<union> fv y"
+ by (auto simp add: fv_def supp_Pair)
+lemma fv_append[simp]: "fv (x @ y) = fv x \<union> fv y"
+ by (auto simp add: fv_def supp_append)
+lemma fv_at_base[simp]: "fv a = {a::'a::at_base}"
+ by (auto simp add: fv_def supp_at_base)
+lemma fv_pure[simp]: "fv (a::'a::pure) = {}"
+ by (auto simp add: fv_def pure_supp)
+
+lemma fv_set_at_base[simp]: "fv (l :: ('a :: at_base) list) = set l"
+ by (induction l) auto
+
+lemma flip_not_fv: "a \<notin> fv x \<Longrightarrow> b \<notin> fv x \<Longrightarrow> (a \<leftrightarrow> b) \<bullet> x = x"
+ by (metis flip_def fresh_def fv_def mem_Collect_eq swap_fresh_fresh)
+
+lemma fv_not_fresh: "atom x \<sharp> e \<longleftrightarrow> x \<notin> fv e"
+ unfolding fv_def fresh_def by blast
+
+lemma fresh_fv: "finite (fv e :: 'a set) \<Longrightarrow> atom (x :: ('a::at_base)) \<sharp> (fv e :: 'a set) \<longleftrightarrow> atom x \<sharp> e"
+ unfolding fv_def fresh_def
+ by (auto simp add: supp_finite_set_at_base)
+
+lemma finite_fv[simp]: "finite (fv (e::'a::fs) :: ('b::at_base) set)"
+proof-
+ have "finite (supp e)" by (metis finite_supp)
+ hence "finite (atom -` supp e :: 'b set)"
+ apply (rule finite_vimageI)
+ apply (rule inj_onI)
+ apply (simp)
+ done
+ moreover
+ have "(atom -` supp e :: 'b set) = fv e" unfolding fv_def by auto
+ ultimately
+ show ?thesis by simp
+qed
+
+definition fv_list :: "'a::fs \<Rightarrow> 'b::at_base list"
+ where "fv_list e = (SOME l. set l = fv e)"
+
+lemma set_fv_list[simp]: "set (fv_list e) = (fv e :: ('b::at_base) set)"
+proof-
+ have "finite (fv e :: 'b set)" by (rule finite_fv)
+ from finite_list[OF finite_fv]
+ obtain l where "set l = (fv e :: 'b set)"..
+ thus ?thesis
+ unfolding fv_list_def by (rule someI)
+qed
+
+lemma fresh_fv_list[simp]:
+ "a \<sharp> (fv_list e :: 'b::at_base list) \<longleftrightarrow> a \<sharp> (fv e :: 'b::at_base set)"
+proof-
+ have "a \<sharp> (fv_list e :: 'b::at_base list) \<longleftrightarrow> a \<sharp> set (fv_list e :: 'b::at_base list)"
+ by (rule fresh_set[symmetric])
+ also have "\<dots> \<longleftrightarrow> a \<sharp> (fv e :: 'b::at_base set)" by simp
+ finally show ?thesis.
+qed
+
+section \<open> Other useful lemmas \<close>
+
+lemma pure_permute_id: "permute p = (\<lambda> x. (x::'a::pure))"
+ by rule (simp add: permute_pure)
+
+lemma supp_set_elem_finite:
+ assumes "finite S"
+ and "(m::'a::fs) \<in> S"
+ and "y \<in> supp m"
+ shows "y \<in> supp S"
+ using assms supp_of_finite_sets
+ by auto
+
+lemmas fresh_star_Cons = fresh_star_list(2)
+
+lemma mem_permute_set:
+ shows "x \<in> p \<bullet> S \<longleftrightarrow> (- p \<bullet> x) \<in> S"
+ by (metis mem_permute_iff permute_minus_cancel(2))
+
+lemma flip_set_both_not_in:
+ assumes "x \<notin> S" and "x' \<notin> S"
+ shows "((x' \<leftrightarrow> x) \<bullet> S) = S"
+ unfolding permute_set_def
+ by (auto) (metis assms flip_at_base_simps(3))+
+
+lemma inj_atom: "inj atom" by (metis atom_eq_iff injI)
+
+lemmas image_Int[OF inj_atom, simp]
+
+lemma eqvt_uncurry: "eqvt f \<Longrightarrow> eqvt (case_prod f)"
+ unfolding eqvt_def
+ by perm_simp simp
+
+lemma supp_fun_app_eqvt2:
+ assumes a: "eqvt f"
+ shows "supp (f x y) \<subseteq> supp x \<union> supp y"
+proof-
+ from supp_fun_app_eqvt[OF eqvt_uncurry [OF a]]
+ have "supp (case_prod f (x,y)) \<subseteq> supp (x,y)".
+ thus ?thesis by (simp add: supp_Pair)
+qed
+
+lemma supp_fun_app_eqvt3:
+ assumes a: "eqvt f"
+ shows "supp (f x y z) \<subseteq> supp x \<union> supp y \<union> supp z"
+proof-
+ from supp_fun_app_eqvt2[OF eqvt_uncurry [OF a]]
+ have "supp (case_prod f (x,y) z) \<subseteq> supp (x,y) \<union> supp z".
+ thus ?thesis by (simp add: supp_Pair)
+qed
+
+(* Fighting eta-contraction *)
+lemma permute_0[simp]: "permute 0 = (\<lambda> x. x)"
+ by auto
+lemma permute_comp[simp]: "permute x \<circ> permute y = permute (x + y)" by auto
+
+lemma map_permute: "map (permute p) = permute p"
+ apply rule
+ apply (induct_tac x)
+ apply auto
+ done
+
+lemma fresh_star_restrictA[intro]: "a \<sharp>* \<Gamma> \<Longrightarrow> a \<sharp>* AList.restrict V \<Gamma>"
+ by (induction \<Gamma>) (auto simp add: fresh_star_Cons)
+
+lemma Abs_lst_Nil_eq[simp]: "[[]]lst. (x::'a::fs) = [xs]lst. x' \<longleftrightarrow> (([],x) = (xs, x'))"
+ apply rule
+ apply (frule Abs_lst_fcb2[where f = "\<lambda> x y _ . (x,y)" and as = "[]" and bs = "xs" and c = "()"])
+ apply (auto simp add: fresh_star_def)
+ done
+
+lemma Abs_lst_Nil_eq2[simp]: "[xs]lst. (x::'a::fs) = [[]]lst. x' \<longleftrightarrow> ((xs,x) = ([], x'))"
+ by (subst eq_commute) auto
+
+lemma prod_cases8 [cases type]:
+ obtains (fields) a b c d e f g h where "y = (a, b, c, d, e, f, g,h)"
+ by (cases y, case_tac g) blast
+
+lemma prod_induct8 [case_names fields, induct type]:
+ "(\<And>a b c d e f g h. P (a, b, c, d, e, f, g, h)) \<Longrightarrow> P x"
+ by (cases x) blast
+
+lemma prod_cases9 [cases type]:
+ obtains (fields) a b c d e f g h i where "y = (a, b, c, d, e, f, g,h,i)"
+ by (cases y, case_tac h) blast
+
+lemma prod_induct9 [case_names fields, induct type]:
+ "(\<And>a b c d e f g h i. P (a, b, c, d, e, f, g, h, i)) \<Longrightarrow> P x"
+ by (cases x) blast
+
+named_theorems nominal_prod_simps
+
+named_theorems ms_fresh "Facts for helping with freshness proofs"
+
+lemma fresh_prod2[nominal_prod_simps,ms_fresh]: "x \<sharp> (a,b) = (x \<sharp> a \<and> x \<sharp> b )"
+ using fresh_def supp_Pair by fastforce
+
+lemma fresh_prod3[nominal_prod_simps,ms_fresh]: "x \<sharp> (a,b,c) = (x \<sharp> a \<and> x \<sharp> b \<and> x \<sharp> c)"
+ using fresh_def supp_Pair by fastforce
+
+lemma fresh_prod4[nominal_prod_simps,ms_fresh]: "x \<sharp> (a,b,c,d) = (x \<sharp> a \<and> x \<sharp> b \<and> x \<sharp> c \<and> x \<sharp> d)"
+ using fresh_def supp_Pair by fastforce
+
+lemma fresh_prod5[nominal_prod_simps,ms_fresh]: "x \<sharp> (a,b,c,d,e) = (x \<sharp> a \<and> x \<sharp> b \<and> x \<sharp> c \<and> x \<sharp> d \<and> x \<sharp> e)"
+ using fresh_def supp_Pair by fastforce
+
+lemma fresh_prod6[nominal_prod_simps,ms_fresh]: "x \<sharp> (a,b,c,d,e,f) = (x \<sharp> a \<and> x \<sharp> b \<and> x \<sharp> c \<and> x \<sharp> d \<and> x \<sharp> e \<and> x \<sharp> f)"
+ using fresh_def supp_Pair by fastforce
+
+lemma fresh_prod7[nominal_prod_simps,ms_fresh]: "x \<sharp> (a,b,c,d,e,f,g) = (x \<sharp> a \<and> x \<sharp> b \<and> x \<sharp> c \<and> x \<sharp> d \<and> x \<sharp> e \<and> x \<sharp> f \<and> x \<sharp> g)"
+ using fresh_def supp_Pair by fastforce
+
+lemma fresh_prod8[nominal_prod_simps,ms_fresh]: "x \<sharp> (a,b,c,d,e,f,g,h) = (x \<sharp> a \<and> x \<sharp> b \<and> x \<sharp> c \<and> x \<sharp> d \<and> x \<sharp> e \<and> x \<sharp> f \<and> x \<sharp> g \<and> x \<sharp> h )"
+ using fresh_def supp_Pair by fastforce
+
+lemma fresh_prod9[nominal_prod_simps,ms_fresh]: "x \<sharp> (a,b,c,d,e,f,g,h,i) = (x \<sharp> a \<and> x \<sharp> b \<and> x \<sharp> c \<and> x \<sharp> d \<and> x \<sharp> e \<and> x \<sharp> f \<and> x \<sharp> g \<and> x \<sharp> h \<and> x \<sharp> i)"
+ using fresh_def supp_Pair by fastforce
+
+lemma fresh_prod10[nominal_prod_simps,ms_fresh]: "x \<sharp> (a,b,c,d,e,f,g,h,i,j) = (x \<sharp> a \<and> x \<sharp> b \<and> x \<sharp> c \<and> x \<sharp> d \<and> x \<sharp> e \<and> x \<sharp> f \<and> x \<sharp> g \<and> x \<sharp> h \<and> x \<sharp> i \<and> x \<sharp> j)"
+ using fresh_def supp_Pair by fastforce
+
+lemma fresh_prod12[nominal_prod_simps,ms_fresh]: "x \<sharp> (a,b,c,d,e,f,g,h,i,j,k,l) = (x \<sharp> a \<and> x \<sharp> b \<and> x \<sharp> c \<and> x \<sharp> d \<and> x \<sharp> e \<and> x \<sharp> f \<and> x \<sharp> g \<and> x \<sharp> h \<and> x \<sharp> i \<and> x \<sharp> j \<and> x \<sharp> k \<and> x \<sharp> l)"
+ using fresh_def supp_Pair by fastforce
+
+lemmas fresh_prodN = fresh_Pair fresh_prod3 fresh_prod4 fresh_prod5 fresh_prod6 fresh_prod7 fresh_prod8 fresh_prod9 fresh_prod10 fresh_prod12
+
+lemma fresh_prod2I:
+ fixes x and x1 and x2
+ assumes "x \<sharp> x1" and "x \<sharp> x2"
+ shows "x \<sharp> (x1,x2)" using fresh_prod2 assms by auto
+
+lemma fresh_prod3I:
+ fixes x and x1 and x2 and x3
+ assumes "x \<sharp> x1" and "x \<sharp> x2" and "x \<sharp> x3"
+ shows "x \<sharp> (x1,x2,x3)" using fresh_prod3 assms by auto
+
+lemma fresh_prod4I:
+ fixes x and x1 and x2 and x3 and x4
+ assumes "x \<sharp> x1" and "x \<sharp> x2" and "x \<sharp> x3" and "x \<sharp> x4"
+ shows "x \<sharp> (x1,x2,x3,x4)" using fresh_prod4 assms by auto
+
+lemma fresh_prod5I:
+ fixes x and x1 and x2 and x3 and x4 and x5
+ assumes "x \<sharp> x1" and "x \<sharp> x2" and "x \<sharp> x3" and "x \<sharp> x4" and "x \<sharp> x5"
+ shows "x \<sharp> (x1,x2,x3,x4,x5)" using fresh_prod5 assms by auto
+
+lemma flip_collapse[simp]:
+ fixes b1::"'a::pt" and bv1::"'b::at" and bv2::"'b::at"
+ assumes "atom bv2 \<sharp> b1" and "atom c \<sharp> (bv1,bv2,b1)" and "bv1 \<noteq> bv2"
+ shows "(bv2 \<leftrightarrow> c) \<bullet> (bv1 \<leftrightarrow> bv2) \<bullet> b1 = (bv1 \<leftrightarrow> c) \<bullet> b1"
+proof -
+ have "c \<noteq> bv1" and "bv2 \<noteq> bv1" using assms by auto+
+ hence "(bv2 \<leftrightarrow> c) + (bv1 \<leftrightarrow> bv2) + (bv2 \<leftrightarrow> c) = (bv1 \<leftrightarrow> c)" using flip_triple[of c bv1 bv2] flip_commute by metis
+ hence "(bv2 \<leftrightarrow> c) \<bullet> (bv1 \<leftrightarrow> bv2) \<bullet> (bv2 \<leftrightarrow> c) \<bullet> b1 = (bv1 \<leftrightarrow> c) \<bullet> b1" using permute_plus by metis
+ thus ?thesis using assms flip_fresh_fresh by force
+qed
+
+lemma triple_eqvt[simp]:
+ "p \<bullet> (x, b,c) = (p \<bullet> x, p \<bullet> b , p \<bullet> c)"
+proof -
+ have "(x,b,c) = (x,(b,c))" by simp
+ thus ?thesis using Pair_eqvt by simp
+qed
+
+lemma lst_fst:
+ fixes x::"'a::at" and t1::"'b::fs" and x'::"'a::at" and t2::"'c::fs"
+ assumes " ([[atom x]]lst. (t1,t2) = [[atom x']]lst. (t1',t2'))"
+ shows " ([[atom x]]lst. t1 = [[atom x']]lst. t1')"
+proof -
+ have "(\<forall>c. atom c \<sharp> (t2,t2') \<longrightarrow> atom c \<sharp> (x, x', t1, t1') \<longrightarrow> (x \<leftrightarrow> c) \<bullet> t1 = (x' \<leftrightarrow> c) \<bullet> t1')"
+ proof(rule,rule,rule)
+ fix c::'a
+ assume "atom c \<sharp> (t2,t2')" and "atom c \<sharp> (x, x', t1, t1')"
+ hence "atom c \<sharp> (x, x', (t1,t2), (t1',t2'))" using fresh_prod2 by simp
+ thus " (x \<leftrightarrow> c) \<bullet> t1 = (x' \<leftrightarrow> c) \<bullet> t1'" using assms Abs1_eq_iff_all(3) Pair_eqvt by simp
+ qed
+ thus ?thesis using Abs1_eq_iff_all(3)[of x t1 x' t1' "(t2,t2')"] by simp
+qed
+
+lemma lst_snd:
+ fixes x::"'a::at" and t1::"'b::fs" and x'::"'a::at" and t2::"'c::fs"
+ assumes " ([[atom x]]lst. (t1,t2) = [[atom x']]lst. (t1',t2'))"
+ shows " ([[atom x]]lst. t2 = [[atom x']]lst. t2')"
+proof -
+ have "(\<forall>c. atom c \<sharp> (t1,t1') \<longrightarrow> atom c \<sharp> (x, x', t2, t2') \<longrightarrow> (x \<leftrightarrow> c) \<bullet> t2 = (x' \<leftrightarrow> c) \<bullet> t2')"
+ proof(rule,rule,rule)
+ fix c::'a
+ assume "atom c \<sharp> (t1,t1')" and "atom c \<sharp> (x, x', t2, t2')"
+ hence "atom c \<sharp> (x, x', (t1,t2), (t1',t2'))" using fresh_prod2 by simp
+ thus " (x \<leftrightarrow> c) \<bullet> t2 = (x' \<leftrightarrow> c) \<bullet> t2'" using assms Abs1_eq_iff_all(3) Pair_eqvt by simp
+ qed
+ thus ?thesis using Abs1_eq_iff_all(3)[of x t2 x' t2' "(t1,t1')"] by simp
+qed
+
+lemma lst_head_cons_pair:
+ fixes y1::"'a ::at" and y2::"'a::at" and x1::"'b::fs" and x2::"'b::fs" and xs1::"('b::fs) list" and xs2::"('b::fs) list"
+ assumes "[[atom y1]]lst. (x1 # xs1) = [[atom y2]]lst. (x2 # xs2)"
+ shows "[[atom y1]]lst. (x1,xs1) = [[atom y2]]lst. (x2,xs2)"
+proof(subst Abs1_eq_iff_all(3)[of y1 "(x1,xs1)" y2 "(x2,xs2)"],rule,rule,rule)
+ fix c::'a
+ assume "atom c \<sharp> (x1#xs1,x2#xs2)" and "atom c \<sharp> (y1, y2, (x1, xs1), x2, xs2)"
+ thus "(y1 \<leftrightarrow> c) \<bullet> (x1, xs1) = (y2 \<leftrightarrow> c) \<bullet> (x2, xs2)" using assms Abs1_eq_iff_all(3) by auto
+qed
+
+lemma lst_head_cons_neq_nil:
+ fixes y1::"'a ::at" and y2::"'a::at" and x1::"'b::fs" and x2::"'b::fs" and xs1::"('b::fs) list" and xs2::"('b::fs) list"
+ assumes "[[atom y1]]lst. (x1 # xs1) = [[atom y2]]lst. (xs2)"
+ shows "xs2 \<noteq> []"
+proof
+ assume as:"xs2 = []"
+ thus False using Abs1_eq_iff(3)[of y1 "x1#xs1" y2 Nil] assms as by auto
+qed
+
+lemma lst_head_cons:
+ fixes y1::"'a ::at" and y2::"'a::at" and x1::"'b::fs" and x2::"'b::fs" and xs1::"('b::fs) list" and xs2::"('b::fs) list"
+ assumes "[[atom y1]]lst. (x1 # xs1) = [[atom y2]]lst. (x2 # xs2)"
+ shows "[[atom y1]]lst. x1 = [[atom y2]]lst. x2" and "[[atom y1]]lst. xs1 = [[atom y2]]lst. xs2"
+ using lst_head_cons_pair lst_fst lst_snd assms by metis+
+
+lemma lst_pure:
+ fixes x1::"'a ::at" and t1::"'b::pure" and x2::"'a ::at" and t2::"'b::pure"
+ assumes "[[atom x1]]lst. t1 = [[atom x2]]lst. t2"
+ shows "t1=t2"
+ using assms Abs1_eq_iff_all(3) pure_fresh flip_fresh_fresh
+ by (metis Abs1_eq(3) permute_pure)
+
+lemma lst_supp:
+ assumes "[[atom x1]]lst. t1 = [[atom x2]]lst. t2"
+ shows "supp t1 - {atom x1} = supp t2 - {atom x2}"
+proof -
+ have "supp ([[atom x1]]lst.t1) = supp ([[atom x2]]lst.t2)" using assms by auto
+ thus ?thesis using Abs_finite_supp
+ by (metis assms empty_set list.simps(15) supp_lst.simps)
+qed
+
+lemma lst_supp_subset:
+ assumes "[[atom x1]]lst. t1 = [[atom x2]]lst. t2" and "supp t1 \<subseteq> {atom x1} \<union> B"
+ shows "supp t2 \<subseteq> {atom x2} \<union> B"
+ using assms lst_supp by fast
+
+lemma projl_inl_eqvt:
+ fixes \<pi> :: perm
+ shows "\<pi> \<bullet> (projl (Inl x)) = projl (Inl (\<pi> \<bullet> x))"
+unfolding projl_def Inl_eqvt by simp
+
+end
diff --git a/thys/MiniSail/Operational.thy b/thys/MiniSail/Operational.thy
--- a/thys/MiniSail/Operational.thy
+++ b/thys/MiniSail/Operational.thy
@@ -1,171 +1,171 @@
-(*<*)
-theory Operational
- imports Typing
-begin
- (*>*)
-
-chapter \<open>Operational Semantics\<close>
-
-text \<open> Here we define the operational semantics in terms of a small-step reduction relation. \<close>
-
-section \<open>Reduction Rules\<close>
-
-text \<open> The store for mutable variables \<close>
-type_synonym \<delta> = "(u*v) list"
-
-nominal_function update_d :: "\<delta> \<Rightarrow> u \<Rightarrow> v \<Rightarrow> \<delta>" where
- "update_d [] _ _ = []"
-| "update_d ((u',v')#\<delta>) u v = (if u = u' then ((u,v)#\<delta>) else ((u',v')# (update_d \<delta> u v)))"
- by(auto,simp add: eqvt_def update_d_graph_aux_def ,metis neq_Nil_conv old.prod.exhaust)
-nominal_termination (eqvt) by lexicographic_order
-
-text \<open> Relates constructor to the branch in the case and binding variable and statement \<close>
-inductive find_branch :: "dc \<Rightarrow> branch_list \<Rightarrow> branch_s \<Rightarrow> bool" where
- find_branch_finalI: "dc' = dc \<Longrightarrow> find_branch dc' (AS_final (AS_branch dc x s )) (AS_branch dc x s)"
-| find_branch_branch_eqI: "dc' = dc \<Longrightarrow> find_branch dc' (AS_cons (AS_branch dc x s) css) (AS_branch dc x s)"
-| find_branch_branch_neqI: "\<lbrakk> dc \<noteq> dc'; find_branch dc' css cs \<rbrakk> \<Longrightarrow> find_branch dc' (AS_cons (AS_branch dc x s) css) cs"
-equivariance find_branch
-nominal_inductive find_branch .
-
-inductive_cases find_branch_elims[elim!]:
- "find_branch dc (AS_final cs') cs"
- "find_branch dc (AS_cons cs' css) cs"
-
-nominal_function lookup_branch :: "dc \<Rightarrow> branch_list \<Rightarrow> branch_s option" where
- "lookup_branch dc (AS_final (AS_branch dc' x s)) = (if dc = dc' then (Some (AS_branch dc' x s)) else None)"
-| "lookup_branch dc (AS_cons (AS_branch dc' x s) css) = (if dc = dc' then (Some (AS_branch dc' x s)) else lookup_branch dc css)"
- apply(auto,simp add: eqvt_def lookup_branch_graph_aux_def)
- by(metis neq_Nil_conv old.prod.exhaust s_branch_s_branch_list.strong_exhaust)
-nominal_termination (eqvt) by lexicographic_order
-
-text \<open> Reduction rules \<close>
-inductive reduce_stmt :: "\<Phi> \<Rightarrow> \<delta> \<Rightarrow> s \<Rightarrow> \<delta> \<Rightarrow> s \<Rightarrow> bool" (" _ \<turnstile> \<langle> _ , _\<rangle> \<longrightarrow> \<langle> _ , _\<rangle>" [50, 50, 50] 50) where
- reduce_if_trueI: " \<Phi> \<turnstile> \<langle>\<delta>, AS_if [L_true]\<^sup>v s1 s2\<rangle> \<longrightarrow> \<langle>\<delta>, s1\<rangle> "
-| reduce_if_falseI: " \<Phi> \<turnstile> \<langle>\<delta>, AS_if [L_false]\<^sup>v s1 s2\<rangle> \<longrightarrow> \<langle>\<delta>, s2\<rangle> "
-| reduce_let_valI: " \<Phi> \<turnstile> \<langle>\<delta>, AS_let x (AE_val v) s\<rangle> \<longrightarrow> \<langle>\<delta>, s[x::=v]\<^sub>s\<^sub>v\<rangle>"
-| reduce_let_plusI: " \<Phi> \<turnstile> \<langle>\<delta>, AS_let x (AE_op Plus ((V_lit (L_num n1))) ((V_lit (L_num n2)))) s\<rangle> \<longrightarrow>
- \<langle>\<delta>, AS_let x (AE_val (V_lit (L_num ( (( n1)+(n2)))))) s \<rangle> "
-| reduce_let_leqI: "b = (if (n1 \<le> n2) then L_true else L_false) \<Longrightarrow>
- \<Phi> \<turnstile> \<langle>\<delta>, AS_let x ((AE_op LEq (V_lit (L_num n1)) (V_lit (L_num n2)))) s\<rangle> \<longrightarrow>
- \<langle>\<delta>, AS_let x (AE_val (V_lit b)) s\<rangle>"
-| reduce_let_eqI: "b = (if (n1 = n2) then L_true else L_false) \<Longrightarrow>
- \<Phi> \<turnstile> \<langle>\<delta>, AS_let x ((AE_op Eq (V_lit n1) (V_lit n2))) s\<rangle> \<longrightarrow>
- \<langle>\<delta>, AS_let x (AE_val (V_lit b)) s\<rangle>"
-| reduce_let_appI: "Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ z b c \<tau> s'))) = lookup_fun \<Phi> f \<Longrightarrow>
- \<Phi> \<turnstile> \<langle>\<delta>, AS_let x ((AE_app f v)) s\<rangle> \<longrightarrow> \<langle>\<delta>, AS_let2 x \<tau>[z::=v]\<^sub>\<tau>\<^sub>v s'[z::=v]\<^sub>s\<^sub>v s\<rangle> "
-| reduce_let_appPI: "Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ z b c \<tau> s'))) = lookup_fun \<Phi> f \<Longrightarrow>
- \<Phi> \<turnstile> \<langle>\<delta>, AS_let x ((AE_appP f b' v)) s\<rangle> \<longrightarrow> \<langle>\<delta>, AS_let2 x \<tau>[bv::=b']\<^sub>\<tau>\<^sub>b[z::=v]\<^sub>\<tau>\<^sub>v s'[bv::=b']\<^sub>s\<^sub>b[z::=v]\<^sub>s\<^sub>v s\<rangle> "
-| reduce_let_fstI: "\<Phi> \<turnstile> \<langle>\<delta>, AS_let x (AE_fst (V_pair v1 v2)) s\<rangle> \<longrightarrow> \<langle>\<delta>, AS_let x (AE_val v1) s\<rangle>"
-| reduce_let_sndI: "\<Phi> \<turnstile> \<langle>\<delta>, AS_let x (AE_snd (V_pair v1 v2)) s\<rangle> \<longrightarrow> \<langle>\<delta>, AS_let x (AE_val v2) s\<rangle>"
-| reduce_let_concatI: "\<Phi> \<turnstile> \<langle>\<delta>, AS_let x (AE_concat (V_lit (L_bitvec v1)) (V_lit (L_bitvec v2))) s\<rangle> \<longrightarrow>
- \<langle>\<delta>, AS_let x (AE_val (V_lit (L_bitvec (v1@v2)))) s\<rangle>"
-| reduce_let_splitI: " split n v (v1 , v2 ) \<Longrightarrow> \<Phi> \<turnstile> \<langle>\<delta>, AS_let x (AE_split (V_lit (L_bitvec v)) (V_lit (L_num n))) s\<rangle> \<longrightarrow>
- \<langle>\<delta>, AS_let x (AE_val (V_pair (V_lit (L_bitvec v1)) (V_lit (L_bitvec v2)))) s\<rangle>"
-| reduce_let_lenI: "\<Phi> \<turnstile> \<langle>\<delta>, AS_let x (AE_len (V_lit (L_bitvec v))) s\<rangle> \<longrightarrow>
- \<langle>\<delta>, AS_let x (AE_val (V_lit (L_num (int (List.length v))))) s\<rangle>"
-| reduce_let_mvar: "(u,v) \<in> set \<delta> \<Longrightarrow> \<Phi> \<turnstile> \<langle>\<delta>, AS_let x (AE_mvar u) s\<rangle> \<longrightarrow> \<langle>\<delta>, AS_let x (AE_val v) s \<rangle>"
-| reduce_assert1I: "\<Phi> \<turnstile> \<langle>\<delta>, AS_assert c (AS_val v)\<rangle> \<longrightarrow> \<langle>\<delta>, AS_val v\<rangle>"
-| reduce_assert2I: " \<Phi> \<turnstile> \<langle>\<delta>, s \<rangle> \<longrightarrow> \<langle> \<delta>', s'\<rangle> \<Longrightarrow> \<Phi> \<turnstile> \<langle>\<delta>, AS_assert c s\<rangle> \<longrightarrow> \<langle> \<delta>', AS_assert c s'\<rangle>"
-| reduce_varI: "atom u \<sharp> \<delta> \<Longrightarrow> \<Phi> \<turnstile> \<langle>\<delta>, AS_var u \<tau> v s \<rangle> \<longrightarrow> \<langle> ((u,v)#\<delta>) , s\<rangle>"
-| reduce_assignI: " \<Phi> \<turnstile> \<langle>\<delta>, AS_assign u v \<rangle> \<longrightarrow> \<langle> update_d \<delta> u v , AS_val (V_lit L_unit)\<rangle>"
-| reduce_seq1I: "\<Phi> \<turnstile> \<langle>\<delta>, AS_seq (AS_val (V_lit L_unit )) s \<rangle> \<longrightarrow> \<langle>\<delta>, s\<rangle>"
-| reduce_seq2I: "\<lbrakk>s1 \<noteq> AS_val v ; \<Phi> \<turnstile> \<langle>\<delta>, s1\<rangle> \<longrightarrow> \<langle> \<delta>', s1'\<rangle> \<rbrakk> \<Longrightarrow>
- \<Phi> \<turnstile> \<langle>\<delta>, AS_seq s1 s2\<rangle> \<longrightarrow> \<langle> \<delta>', AS_seq s1' s2\<rangle>"
-| reduce_let2_valI: "\<Phi> \<turnstile> \<langle>\<delta>, AS_let2 x t (AS_val v) s\<rangle> \<longrightarrow> \<langle>\<delta>, s[x::=v]\<^sub>s\<^sub>v\<rangle>"
-| reduce_let2I: " \<Phi> \<turnstile> \<langle>\<delta>, s1 \<rangle> \<longrightarrow> \<langle> \<delta>', s1'\<rangle> \<Longrightarrow> \<Phi> \<turnstile> \<langle>\<delta>, AS_let2 x t s1 s2\<rangle> \<longrightarrow> \<langle> \<delta>', AS_let2 x t s1' s2\<rangle>"
-
-| reduce_caseI: "\<lbrakk> Some (AS_branch dc x' s') = lookup_branch dc cs \<rbrakk> \<Longrightarrow> \<Phi> \<turnstile> \<langle>\<delta>, AS_match (V_cons tyid dc v) cs\<rangle> \<longrightarrow> \<langle>\<delta>, s'[x'::=v]\<^sub>s\<^sub>v\<rangle> "
-| reduce_whileI: "\<lbrakk> atom x \<sharp> (s1,s2); atom z \<sharp> x \<rbrakk> \<Longrightarrow>
- \<Phi> \<turnstile> \<langle>\<delta>, AS_while s1 s2 \<rangle> \<longrightarrow>
- \<langle>\<delta>, AS_let2 x (\<lbrace> z : B_bool | TRUE \<rbrace>) s1 (AS_if (V_var x) (AS_seq s2 (AS_while s1 s2)) (AS_val (V_lit L_unit))) \<rangle>"
-
-equivariance reduce_stmt
-nominal_inductive reduce_stmt .
-
-inductive_cases reduce_stmt_elims[elim!]:
- "\<Phi> \<turnstile> \<langle>\<delta>, AS_if (V_lit L_true) s1 s2\<rangle> \<longrightarrow> \<langle>\<delta> , s1\<rangle>"
- "\<Phi> \<turnstile> \<langle>\<delta>, AS_if (V_lit L_false) s1 s2\<rangle> \<longrightarrow> \<langle>\<delta>,s2\<rangle>"
- "\<Phi> \<turnstile> \<langle>\<delta>, AS_let x (AE_val v) s\<rangle> \<longrightarrow> \<langle>\<delta>,s'\<rangle>"
- "\<Phi> \<turnstile> \<langle>\<delta>, AS_let x (AE_op Plus ((V_lit (L_num n1))) ((V_lit (L_num n2)))) s\<rangle> \<longrightarrow>
- \<langle>\<delta>, AS_let x (AE_val (V_lit (L_num ( (( n1)+(n2)))))) s \<rangle>"
- "\<Phi> \<turnstile> \<langle>\<delta>, AS_let x ((AE_op LEq (V_lit (L_num n1)) (V_lit (L_num n2)))) s\<rangle> \<longrightarrow> \<langle>\<delta>, AS_let x (AE_val (V_lit b)) s\<rangle>"
- "\<Phi> \<turnstile> \<langle>\<delta>, AS_let x ((AE_app f v)) s\<rangle> \<longrightarrow> \<langle>\<delta>, AS_let2 x \<tau> (subst_sv s' x v ) s\<rangle> "
- "\<Phi> \<turnstile> \<langle>\<delta>, AS_let x ((AE_len v)) s\<rangle> \<longrightarrow> \<langle>\<delta>, AS_let x v' s\<rangle> "
- "\<Phi> \<turnstile> \<langle>\<delta>, AS_let x ((AE_concat v1 v2)) s\<rangle> \<longrightarrow> \<langle>\<delta>, AS_let x v' s\<rangle> "
- "\<Phi> \<turnstile> \<langle>\<delta>, AS_seq s1 s2\<rangle> \<longrightarrow> \<langle> \<delta>' ,s'\<rangle>"
- "\<Phi> \<turnstile> \<langle>\<delta>, AS_let x ((AE_appP f b v)) s\<rangle> \<longrightarrow> \<langle>\<delta>, AS_let2 x \<tau> (subst_sv s' z v) s\<rangle> "
- "\<Phi> \<turnstile> \<langle>\<delta>, AS_let x ((AE_split v1 v2)) s\<rangle> \<longrightarrow> \<langle>\<delta>, AS_let x v' s\<rangle> "
- "\<Phi> \<turnstile> \<langle>\<delta>, AS_assert c s \<rangle> \<longrightarrow> \<langle>\<delta>, s'\<rangle> "
- "\<Phi> \<turnstile> \<langle>\<delta>, AS_let x ((AE_op Eq (V_lit (n1)) (V_lit (n2)))) s\<rangle> \<longrightarrow> \<langle>\<delta>, AS_let x (AE_val (V_lit b)) s\<rangle>"
-
-inductive reduce_stmt_many :: "\<Phi> \<Rightarrow> \<delta> \<Rightarrow> s \<Rightarrow> \<delta> \<Rightarrow> s \<Rightarrow> bool" ("_ \<turnstile> \<langle> _ , _\<rangle> \<longrightarrow>\<^sup>* \<langle> _ , _\<rangle>" [50, 50, 50] 50) where
- reduce_stmt_many_oneI: "\<Phi> \<turnstile> \<langle>\<delta>, s\<rangle> \<longrightarrow> \<langle>\<delta>', s'\<rangle> \<Longrightarrow> \<Phi> \<turnstile> \<langle>\<delta> , s\<rangle> \<longrightarrow>\<^sup>* \<langle>\<delta>', s'\<rangle> "
-| reduce_stmt_many_manyI: "\<lbrakk> \<Phi> \<turnstile> \<langle>\<delta>, s\<rangle> \<longrightarrow> \<langle>\<delta>', s'\<rangle> ; \<Phi> \<turnstile> \<langle>\<delta>', s'\<rangle> \<longrightarrow>\<^sup>* \<langle>\<delta>'', s''\<rangle> \<rbrakk> \<Longrightarrow> \<Phi> \<turnstile> \<langle>\<delta>, s\<rangle> \<longrightarrow>\<^sup>* \<langle>\<delta>'', s''\<rangle>"
-
-nominal_function convert_fds :: "fun_def list \<Rightarrow> (f*fun_def) list" where
- "convert_fds [] = []"
-| "convert_fds ((AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \<tau> s)))#fs) = ((f,AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \<tau> s)))#convert_fds fs)"
-| "convert_fds ((AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \<tau> s)))#fs) = ((f,AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \<tau> s)))#convert_fds fs)"
- apply(auto)
- apply (simp add: eqvt_def convert_fds_graph_aux_def )
- using fun_def.exhaust fun_typ.exhaust fun_typ_q.exhaust neq_Nil_conv
- by metis
-nominal_termination (eqvt) by lexicographic_order
-
-nominal_function convert_tds :: "type_def list \<Rightarrow> (f*type_def) list" where
- "convert_tds [] = []"
-| "convert_tds ((AF_typedef s dclist)#fs) = ((s,AF_typedef s dclist)#convert_tds fs)"
-| "convert_tds ((AF_typedef_poly s bv dclist)#fs) = ((s,AF_typedef_poly s bv dclist)#convert_tds fs)"
- apply(auto)
- apply (simp add: eqvt_def convert_tds_graph_aux_def )
- by (metis type_def.exhaust neq_Nil_conv)
-nominal_termination (eqvt) by lexicographic_order
-
-inductive reduce_prog :: "p \<Rightarrow> v \<Rightarrow> bool" where
- "\<lbrakk> reduce_stmt_many \<Phi> [] s \<delta> (AS_val v) \<rbrakk> \<Longrightarrow> reduce_prog (AP_prog \<Theta> \<Phi> [] s) v"
-
-section \<open>Reduction Typing\<close>
-
-text \<open> Checks that the store is consistent with @{typ \<Delta>} \<close>
-inductive delta_sim :: "\<Theta> \<Rightarrow> \<delta> \<Rightarrow> \<Delta> \<Rightarrow> bool" ( "_ \<turnstile> _ \<sim> _ " [50,50] 50 ) where
- delta_sim_nilI: "\<Theta> \<turnstile> [] \<sim> []\<^sub>\<Delta> "
-| delta_sim_consI: "\<lbrakk> \<Theta> \<turnstile> \<delta> \<sim> \<Delta> ; \<Theta> ; {||} ; GNil \<turnstile> v \<Leftarrow> \<tau> ; u \<notin> fst ` set \<delta> \<rbrakk> \<Longrightarrow> \<Theta> \<turnstile> ((u,v)#\<delta>) \<sim> ((u,\<tau>)#\<^sub>\<Delta>\<Delta>)"
-
-equivariance delta_sim
-nominal_inductive delta_sim .
-
-inductive_cases delta_sim_elims[elim!]:
- "\<Theta> \<turnstile> [] \<sim> []\<^sub>\<Delta>"
- "\<Theta> \<turnstile> ((u,v)#ds) \<sim> (u,\<tau>) #\<^sub>\<Delta> D"
- "\<Theta> \<turnstile> ((u,v)#ds) \<sim> D"
-
-text \<open>A typing judgement that combines typing of the statement, the store and the condition that definitions are well-typed\<close>
-inductive config_type :: "\<Theta> \<Rightarrow> \<Phi> \<Rightarrow> \<Delta> \<Rightarrow> \<delta> \<Rightarrow> s \<Rightarrow> \<tau> \<Rightarrow> bool" ("_ ; _ ; _ \<turnstile> \<langle> _ , _\<rangle> \<Leftarrow> _ " [50, 50, 50] 50) where
- config_typeI: "\<lbrakk> \<Theta> ; \<Phi> ; {||} ; GNil ; \<Delta> \<turnstile> s \<Leftarrow> \<tau>;
- (\<forall> fd \<in> set \<Phi>. \<Theta> ; \<Phi> \<turnstile> fd) ;
- \<Theta> \<turnstile> \<delta> \<sim> \<Delta> \<rbrakk>
- \<Longrightarrow> \<Theta> ; \<Phi> ; \<Delta> \<turnstile> \<langle> \<delta> , s\<rangle> \<Leftarrow> \<tau>"
-equivariance config_type
-nominal_inductive config_type .
-
-inductive_cases config_type_elims [elim!]:
- " \<Theta> ; \<Phi> ; \<Delta> \<turnstile> \<langle> \<delta> , s\<rangle> \<Leftarrow> \<tau>"
-
-nominal_function \<delta>_of :: "var_def list \<Rightarrow> \<delta>" where
- "\<delta>_of [] = []"
-| "\<delta>_of ((AV_def u t v)#vs) = (u,v) # (\<delta>_of vs)"
- apply auto
- using eqvt_def \<delta>_of_graph_aux_def neq_Nil_conv old.prod.exhaust apply force
- using eqvt_def \<delta>_of_graph_aux_def neq_Nil_conv old.prod.exhaust
- by (metis var_def.strong_exhaust)
-nominal_termination (eqvt) by lexicographic_order
-
-inductive config_type_prog :: "p \<Rightarrow> \<tau> \<Rightarrow> bool" (" \<turnstile> \<langle> _\<rangle> \<Leftarrow> _") where
- "\<lbrakk>
- \<Theta> ; \<Phi> ; \<Delta>_of \<G> \<turnstile> \<langle> \<delta>_of \<G> , s\<rangle> \<Leftarrow> \<tau>
-\<rbrakk> \<Longrightarrow> \<turnstile> \<langle> AP_prog \<Theta> \<Phi> \<G> s\<rangle> \<Leftarrow> \<tau>"
-
-inductive_cases config_type_prog_elims [elim!]:
- "\<turnstile> \<langle> AP_prog \<Theta> \<Phi> \<G> s\<rangle> \<Leftarrow> \<tau>"
-
-end
+(*<*)
+theory Operational
+ imports Typing
+begin
+ (*>*)
+
+chapter \<open>Operational Semantics\<close>
+
+text \<open> Here we define the operational semantics in terms of a small-step reduction relation. \<close>
+
+section \<open>Reduction Rules\<close>
+
+text \<open> The store for mutable variables \<close>
+type_synonym \<delta> = "(u*v) list"
+
+nominal_function update_d :: "\<delta> \<Rightarrow> u \<Rightarrow> v \<Rightarrow> \<delta>" where
+ "update_d [] _ _ = []"
+| "update_d ((u',v')#\<delta>) u v = (if u = u' then ((u,v)#\<delta>) else ((u',v')# (update_d \<delta> u v)))"
+ by(auto,simp add: eqvt_def update_d_graph_aux_def ,metis neq_Nil_conv old.prod.exhaust)
+nominal_termination (eqvt) by lexicographic_order
+
+text \<open> Relates constructor to the branch in the case and binding variable and statement \<close>
+inductive find_branch :: "dc \<Rightarrow> branch_list \<Rightarrow> branch_s \<Rightarrow> bool" where
+ find_branch_finalI: "dc' = dc \<Longrightarrow> find_branch dc' (AS_final (AS_branch dc x s )) (AS_branch dc x s)"
+| find_branch_branch_eqI: "dc' = dc \<Longrightarrow> find_branch dc' (AS_cons (AS_branch dc x s) css) (AS_branch dc x s)"
+| find_branch_branch_neqI: "\<lbrakk> dc \<noteq> dc'; find_branch dc' css cs \<rbrakk> \<Longrightarrow> find_branch dc' (AS_cons (AS_branch dc x s) css) cs"
+equivariance find_branch
+nominal_inductive find_branch .
+
+inductive_cases find_branch_elims[elim!]:
+ "find_branch dc (AS_final cs') cs"
+ "find_branch dc (AS_cons cs' css) cs"
+
+nominal_function lookup_branch :: "dc \<Rightarrow> branch_list \<Rightarrow> branch_s option" where
+ "lookup_branch dc (AS_final (AS_branch dc' x s)) = (if dc = dc' then (Some (AS_branch dc' x s)) else None)"
+| "lookup_branch dc (AS_cons (AS_branch dc' x s) css) = (if dc = dc' then (Some (AS_branch dc' x s)) else lookup_branch dc css)"
+ apply(auto,simp add: eqvt_def lookup_branch_graph_aux_def)
+ by(metis neq_Nil_conv old.prod.exhaust s_branch_s_branch_list.strong_exhaust)
+nominal_termination (eqvt) by lexicographic_order
+
+text \<open> Reduction rules \<close>
+inductive reduce_stmt :: "\<Phi> \<Rightarrow> \<delta> \<Rightarrow> s \<Rightarrow> \<delta> \<Rightarrow> s \<Rightarrow> bool" (" _ \<turnstile> \<langle> _ , _\<rangle> \<longrightarrow> \<langle> _ , _\<rangle>" [50, 50, 50] 50) where
+ reduce_if_trueI: " \<Phi> \<turnstile> \<langle>\<delta>, AS_if [L_true]\<^sup>v s1 s2\<rangle> \<longrightarrow> \<langle>\<delta>, s1\<rangle> "
+| reduce_if_falseI: " \<Phi> \<turnstile> \<langle>\<delta>, AS_if [L_false]\<^sup>v s1 s2\<rangle> \<longrightarrow> \<langle>\<delta>, s2\<rangle> "
+| reduce_let_valI: " \<Phi> \<turnstile> \<langle>\<delta>, AS_let x (AE_val v) s\<rangle> \<longrightarrow> \<langle>\<delta>, s[x::=v]\<^sub>s\<^sub>v\<rangle>"
+| reduce_let_plusI: " \<Phi> \<turnstile> \<langle>\<delta>, AS_let x (AE_op Plus ((V_lit (L_num n1))) ((V_lit (L_num n2)))) s\<rangle> \<longrightarrow>
+ \<langle>\<delta>, AS_let x (AE_val (V_lit (L_num ( (( n1)+(n2)))))) s \<rangle> "
+| reduce_let_leqI: "b = (if (n1 \<le> n2) then L_true else L_false) \<Longrightarrow>
+ \<Phi> \<turnstile> \<langle>\<delta>, AS_let x ((AE_op LEq (V_lit (L_num n1)) (V_lit (L_num n2)))) s\<rangle> \<longrightarrow>
+ \<langle>\<delta>, AS_let x (AE_val (V_lit b)) s\<rangle>"
+| reduce_let_eqI: "b = (if (n1 = n2) then L_true else L_false) \<Longrightarrow>
+ \<Phi> \<turnstile> \<langle>\<delta>, AS_let x ((AE_op Eq (V_lit n1) (V_lit n2))) s\<rangle> \<longrightarrow>
+ \<langle>\<delta>, AS_let x (AE_val (V_lit b)) s\<rangle>"
+| reduce_let_appI: "Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ z b c \<tau> s'))) = lookup_fun \<Phi> f \<Longrightarrow>
+ \<Phi> \<turnstile> \<langle>\<delta>, AS_let x ((AE_app f v)) s\<rangle> \<longrightarrow> \<langle>\<delta>, AS_let2 x \<tau>[z::=v]\<^sub>\<tau>\<^sub>v s'[z::=v]\<^sub>s\<^sub>v s\<rangle> "
+| reduce_let_appPI: "Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ z b c \<tau> s'))) = lookup_fun \<Phi> f \<Longrightarrow>
+ \<Phi> \<turnstile> \<langle>\<delta>, AS_let x ((AE_appP f b' v)) s\<rangle> \<longrightarrow> \<langle>\<delta>, AS_let2 x \<tau>[bv::=b']\<^sub>\<tau>\<^sub>b[z::=v]\<^sub>\<tau>\<^sub>v s'[bv::=b']\<^sub>s\<^sub>b[z::=v]\<^sub>s\<^sub>v s\<rangle> "
+| reduce_let_fstI: "\<Phi> \<turnstile> \<langle>\<delta>, AS_let x (AE_fst (V_pair v1 v2)) s\<rangle> \<longrightarrow> \<langle>\<delta>, AS_let x (AE_val v1) s\<rangle>"
+| reduce_let_sndI: "\<Phi> \<turnstile> \<langle>\<delta>, AS_let x (AE_snd (V_pair v1 v2)) s\<rangle> \<longrightarrow> \<langle>\<delta>, AS_let x (AE_val v2) s\<rangle>"
+| reduce_let_concatI: "\<Phi> \<turnstile> \<langle>\<delta>, AS_let x (AE_concat (V_lit (L_bitvec v1)) (V_lit (L_bitvec v2))) s\<rangle> \<longrightarrow>
+ \<langle>\<delta>, AS_let x (AE_val (V_lit (L_bitvec (v1@v2)))) s\<rangle>"
+| reduce_let_splitI: " split n v (v1 , v2 ) \<Longrightarrow> \<Phi> \<turnstile> \<langle>\<delta>, AS_let x (AE_split (V_lit (L_bitvec v)) (V_lit (L_num n))) s\<rangle> \<longrightarrow>
+ \<langle>\<delta>, AS_let x (AE_val (V_pair (V_lit (L_bitvec v1)) (V_lit (L_bitvec v2)))) s\<rangle>"
+| reduce_let_lenI: "\<Phi> \<turnstile> \<langle>\<delta>, AS_let x (AE_len (V_lit (L_bitvec v))) s\<rangle> \<longrightarrow>
+ \<langle>\<delta>, AS_let x (AE_val (V_lit (L_num (int (List.length v))))) s\<rangle>"
+| reduce_let_mvar: "(u,v) \<in> set \<delta> \<Longrightarrow> \<Phi> \<turnstile> \<langle>\<delta>, AS_let x (AE_mvar u) s\<rangle> \<longrightarrow> \<langle>\<delta>, AS_let x (AE_val v) s \<rangle>"
+| reduce_assert1I: "\<Phi> \<turnstile> \<langle>\<delta>, AS_assert c (AS_val v)\<rangle> \<longrightarrow> \<langle>\<delta>, AS_val v\<rangle>"
+| reduce_assert2I: " \<Phi> \<turnstile> \<langle>\<delta>, s \<rangle> \<longrightarrow> \<langle> \<delta>', s'\<rangle> \<Longrightarrow> \<Phi> \<turnstile> \<langle>\<delta>, AS_assert c s\<rangle> \<longrightarrow> \<langle> \<delta>', AS_assert c s'\<rangle>"
+| reduce_varI: "atom u \<sharp> \<delta> \<Longrightarrow> \<Phi> \<turnstile> \<langle>\<delta>, AS_var u \<tau> v s \<rangle> \<longrightarrow> \<langle> ((u,v)#\<delta>) , s\<rangle>"
+| reduce_assignI: " \<Phi> \<turnstile> \<langle>\<delta>, AS_assign u v \<rangle> \<longrightarrow> \<langle> update_d \<delta> u v , AS_val (V_lit L_unit)\<rangle>"
+| reduce_seq1I: "\<Phi> \<turnstile> \<langle>\<delta>, AS_seq (AS_val (V_lit L_unit )) s \<rangle> \<longrightarrow> \<langle>\<delta>, s\<rangle>"
+| reduce_seq2I: "\<lbrakk>s1 \<noteq> AS_val v ; \<Phi> \<turnstile> \<langle>\<delta>, s1\<rangle> \<longrightarrow> \<langle> \<delta>', s1'\<rangle> \<rbrakk> \<Longrightarrow>
+ \<Phi> \<turnstile> \<langle>\<delta>, AS_seq s1 s2\<rangle> \<longrightarrow> \<langle> \<delta>', AS_seq s1' s2\<rangle>"
+| reduce_let2_valI: "\<Phi> \<turnstile> \<langle>\<delta>, AS_let2 x t (AS_val v) s\<rangle> \<longrightarrow> \<langle>\<delta>, s[x::=v]\<^sub>s\<^sub>v\<rangle>"
+| reduce_let2I: " \<Phi> \<turnstile> \<langle>\<delta>, s1 \<rangle> \<longrightarrow> \<langle> \<delta>', s1'\<rangle> \<Longrightarrow> \<Phi> \<turnstile> \<langle>\<delta>, AS_let2 x t s1 s2\<rangle> \<longrightarrow> \<langle> \<delta>', AS_let2 x t s1' s2\<rangle>"
+
+| reduce_caseI: "\<lbrakk> Some (AS_branch dc x' s') = lookup_branch dc cs \<rbrakk> \<Longrightarrow> \<Phi> \<turnstile> \<langle>\<delta>, AS_match (V_cons tyid dc v) cs\<rangle> \<longrightarrow> \<langle>\<delta>, s'[x'::=v]\<^sub>s\<^sub>v\<rangle> "
+| reduce_whileI: "\<lbrakk> atom x \<sharp> (s1,s2); atom z \<sharp> x \<rbrakk> \<Longrightarrow>
+ \<Phi> \<turnstile> \<langle>\<delta>, AS_while s1 s2 \<rangle> \<longrightarrow>
+ \<langle>\<delta>, AS_let2 x (\<lbrace> z : B_bool | TRUE \<rbrace>) s1 (AS_if (V_var x) (AS_seq s2 (AS_while s1 s2)) (AS_val (V_lit L_unit))) \<rangle>"
+
+equivariance reduce_stmt
+nominal_inductive reduce_stmt .
+
+inductive_cases reduce_stmt_elims[elim!]:
+ "\<Phi> \<turnstile> \<langle>\<delta>, AS_if (V_lit L_true) s1 s2\<rangle> \<longrightarrow> \<langle>\<delta> , s1\<rangle>"
+ "\<Phi> \<turnstile> \<langle>\<delta>, AS_if (V_lit L_false) s1 s2\<rangle> \<longrightarrow> \<langle>\<delta>,s2\<rangle>"
+ "\<Phi> \<turnstile> \<langle>\<delta>, AS_let x (AE_val v) s\<rangle> \<longrightarrow> \<langle>\<delta>,s'\<rangle>"
+ "\<Phi> \<turnstile> \<langle>\<delta>, AS_let x (AE_op Plus ((V_lit (L_num n1))) ((V_lit (L_num n2)))) s\<rangle> \<longrightarrow>
+ \<langle>\<delta>, AS_let x (AE_val (V_lit (L_num ( (( n1)+(n2)))))) s \<rangle>"
+ "\<Phi> \<turnstile> \<langle>\<delta>, AS_let x ((AE_op LEq (V_lit (L_num n1)) (V_lit (L_num n2)))) s\<rangle> \<longrightarrow> \<langle>\<delta>, AS_let x (AE_val (V_lit b)) s\<rangle>"
+ "\<Phi> \<turnstile> \<langle>\<delta>, AS_let x ((AE_app f v)) s\<rangle> \<longrightarrow> \<langle>\<delta>, AS_let2 x \<tau> (subst_sv s' x v ) s\<rangle> "
+ "\<Phi> \<turnstile> \<langle>\<delta>, AS_let x ((AE_len v)) s\<rangle> \<longrightarrow> \<langle>\<delta>, AS_let x v' s\<rangle> "
+ "\<Phi> \<turnstile> \<langle>\<delta>, AS_let x ((AE_concat v1 v2)) s\<rangle> \<longrightarrow> \<langle>\<delta>, AS_let x v' s\<rangle> "
+ "\<Phi> \<turnstile> \<langle>\<delta>, AS_seq s1 s2\<rangle> \<longrightarrow> \<langle> \<delta>' ,s'\<rangle>"
+ "\<Phi> \<turnstile> \<langle>\<delta>, AS_let x ((AE_appP f b v)) s\<rangle> \<longrightarrow> \<langle>\<delta>, AS_let2 x \<tau> (subst_sv s' z v) s\<rangle> "
+ "\<Phi> \<turnstile> \<langle>\<delta>, AS_let x ((AE_split v1 v2)) s\<rangle> \<longrightarrow> \<langle>\<delta>, AS_let x v' s\<rangle> "
+ "\<Phi> \<turnstile> \<langle>\<delta>, AS_assert c s \<rangle> \<longrightarrow> \<langle>\<delta>, s'\<rangle> "
+ "\<Phi> \<turnstile> \<langle>\<delta>, AS_let x ((AE_op Eq (V_lit (n1)) (V_lit (n2)))) s\<rangle> \<longrightarrow> \<langle>\<delta>, AS_let x (AE_val (V_lit b)) s\<rangle>"
+
+inductive reduce_stmt_many :: "\<Phi> \<Rightarrow> \<delta> \<Rightarrow> s \<Rightarrow> \<delta> \<Rightarrow> s \<Rightarrow> bool" ("_ \<turnstile> \<langle> _ , _\<rangle> \<longrightarrow>\<^sup>* \<langle> _ , _\<rangle>" [50, 50, 50] 50) where
+ reduce_stmt_many_oneI: "\<Phi> \<turnstile> \<langle>\<delta>, s\<rangle> \<longrightarrow> \<langle>\<delta>', s'\<rangle> \<Longrightarrow> \<Phi> \<turnstile> \<langle>\<delta> , s\<rangle> \<longrightarrow>\<^sup>* \<langle>\<delta>', s'\<rangle> "
+| reduce_stmt_many_manyI: "\<lbrakk> \<Phi> \<turnstile> \<langle>\<delta>, s\<rangle> \<longrightarrow> \<langle>\<delta>', s'\<rangle> ; \<Phi> \<turnstile> \<langle>\<delta>', s'\<rangle> \<longrightarrow>\<^sup>* \<langle>\<delta>'', s''\<rangle> \<rbrakk> \<Longrightarrow> \<Phi> \<turnstile> \<langle>\<delta>, s\<rangle> \<longrightarrow>\<^sup>* \<langle>\<delta>'', s''\<rangle>"
+
+nominal_function convert_fds :: "fun_def list \<Rightarrow> (f*fun_def) list" where
+ "convert_fds [] = []"
+| "convert_fds ((AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \<tau> s)))#fs) = ((f,AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \<tau> s)))#convert_fds fs)"
+| "convert_fds ((AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \<tau> s)))#fs) = ((f,AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \<tau> s)))#convert_fds fs)"
+ apply(auto)
+ apply (simp add: eqvt_def convert_fds_graph_aux_def )
+ using fun_def.exhaust fun_typ.exhaust fun_typ_q.exhaust neq_Nil_conv
+ by metis
+nominal_termination (eqvt) by lexicographic_order
+
+nominal_function convert_tds :: "type_def list \<Rightarrow> (f*type_def) list" where
+ "convert_tds [] = []"
+| "convert_tds ((AF_typedef s dclist)#fs) = ((s,AF_typedef s dclist)#convert_tds fs)"
+| "convert_tds ((AF_typedef_poly s bv dclist)#fs) = ((s,AF_typedef_poly s bv dclist)#convert_tds fs)"
+ apply(auto)
+ apply (simp add: eqvt_def convert_tds_graph_aux_def )
+ by (metis type_def.exhaust neq_Nil_conv)
+nominal_termination (eqvt) by lexicographic_order
+
+inductive reduce_prog :: "p \<Rightarrow> v \<Rightarrow> bool" where
+ "\<lbrakk> reduce_stmt_many \<Phi> [] s \<delta> (AS_val v) \<rbrakk> \<Longrightarrow> reduce_prog (AP_prog \<Theta> \<Phi> [] s) v"
+
+section \<open>Reduction Typing\<close>
+
+text \<open> Checks that the store is consistent with @{typ \<Delta>} \<close>
+inductive delta_sim :: "\<Theta> \<Rightarrow> \<delta> \<Rightarrow> \<Delta> \<Rightarrow> bool" ( "_ \<turnstile> _ \<sim> _ " [50,50] 50 ) where
+ delta_sim_nilI: "\<Theta> \<turnstile> [] \<sim> []\<^sub>\<Delta> "
+| delta_sim_consI: "\<lbrakk> \<Theta> \<turnstile> \<delta> \<sim> \<Delta> ; \<Theta> ; {||} ; GNil \<turnstile> v \<Leftarrow> \<tau> ; u \<notin> fst ` set \<delta> \<rbrakk> \<Longrightarrow> \<Theta> \<turnstile> ((u,v)#\<delta>) \<sim> ((u,\<tau>)#\<^sub>\<Delta>\<Delta>)"
+
+equivariance delta_sim
+nominal_inductive delta_sim .
+
+inductive_cases delta_sim_elims[elim!]:
+ "\<Theta> \<turnstile> [] \<sim> []\<^sub>\<Delta>"
+ "\<Theta> \<turnstile> ((u,v)#ds) \<sim> (u,\<tau>) #\<^sub>\<Delta> D"
+ "\<Theta> \<turnstile> ((u,v)#ds) \<sim> D"
+
+text \<open>A typing judgement that combines typing of the statement, the store and the condition that definitions are well-typed\<close>
+inductive config_type :: "\<Theta> \<Rightarrow> \<Phi> \<Rightarrow> \<Delta> \<Rightarrow> \<delta> \<Rightarrow> s \<Rightarrow> \<tau> \<Rightarrow> bool" ("_ ; _ ; _ \<turnstile> \<langle> _ , _\<rangle> \<Leftarrow> _ " [50, 50, 50] 50) where
+ config_typeI: "\<lbrakk> \<Theta> ; \<Phi> ; {||} ; GNil ; \<Delta> \<turnstile> s \<Leftarrow> \<tau>;
+ (\<forall> fd \<in> set \<Phi>. \<Theta> ; \<Phi> \<turnstile> fd) ;
+ \<Theta> \<turnstile> \<delta> \<sim> \<Delta> \<rbrakk>
+ \<Longrightarrow> \<Theta> ; \<Phi> ; \<Delta> \<turnstile> \<langle> \<delta> , s\<rangle> \<Leftarrow> \<tau>"
+equivariance config_type
+nominal_inductive config_type .
+
+inductive_cases config_type_elims [elim!]:
+ " \<Theta> ; \<Phi> ; \<Delta> \<turnstile> \<langle> \<delta> , s\<rangle> \<Leftarrow> \<tau>"
+
+nominal_function \<delta>_of :: "var_def list \<Rightarrow> \<delta>" where
+ "\<delta>_of [] = []"
+| "\<delta>_of ((AV_def u t v)#vs) = (u,v) # (\<delta>_of vs)"
+ apply auto
+ using eqvt_def \<delta>_of_graph_aux_def neq_Nil_conv old.prod.exhaust apply force
+ using eqvt_def \<delta>_of_graph_aux_def neq_Nil_conv old.prod.exhaust
+ by (metis var_def.strong_exhaust)
+nominal_termination (eqvt) by lexicographic_order
+
+inductive config_type_prog :: "p \<Rightarrow> \<tau> \<Rightarrow> bool" (" \<turnstile> \<langle> _\<rangle> \<Leftarrow> _") where
+ "\<lbrakk>
+ \<Theta> ; \<Phi> ; \<Delta>_of \<G> \<turnstile> \<langle> \<delta>_of \<G> , s\<rangle> \<Leftarrow> \<tau>
+\<rbrakk> \<Longrightarrow> \<turnstile> \<langle> AP_prog \<Theta> \<Phi> \<G> s\<rangle> \<Leftarrow> \<tau>"
+
+inductive_cases config_type_prog_elims [elim!]:
+ "\<turnstile> \<langle> AP_prog \<Theta> \<Phi> \<G> s\<rangle> \<Leftarrow> \<tau>"
+
+end
diff --git a/thys/MiniSail/RCLogic.thy b/thys/MiniSail/RCLogic.thy
--- a/thys/MiniSail/RCLogic.thy
+++ b/thys/MiniSail/RCLogic.thy
@@ -1,222 +1,222 @@
-(*<*)
-theory RCLogic
- imports Wellformed
-begin
- (*>*)
-
-hide_const Syntax.dom
-
-chapter \<open>Refinement Constraint Logic\<close>
-
-text \<open> Semantics for the logic we use in the refinement constraints. It is a multi-sorted, quantifier free
-logic with polymorphic datatypes and linear arithmetic. We could have modelled by using one of the
-encodings to FOL however we wanted to explore using a more direct model. \<close>
-
-section \<open>Evaluation and Satisfiability\<close>
-
-subsection \<open>Valuation\<close>
-
-text \<open> Refinement constraint logic values. SUt is a value for the uninterpreted
- sort that corresponds to basic type variables. For now we only need one of these universes.
- We wrap an smt\_val inside it during a process we call 'boxing'
- which is introduced in the RCLogicL theory \<close>
-nominal_datatype rcl_val = SBitvec "bit list" | SNum int | SBool bool | SPair rcl_val rcl_val |
- SCons tyid string rcl_val | SConsp tyid string b rcl_val |
- SUnit | SUt rcl_val
-
-text \<open> RCL sorts. Represent our domains. The universe is the union of all of the these.
- S\_Ut is the single uninterpreted sort. These map almost directly to basic type
- but we have them to distinguish syntax (basic types) and semantics (RCL sorts) \<close>
-nominal_datatype rcl_sort = S_bool | S_int | S_unit | S_pair rcl_sort rcl_sort | S_id tyid | S_app tyid rcl_sort | S_bitvec | S_ut
-
-type_synonym valuation = "(x,rcl_val) map"
-
-type_synonym type_valuation = "(bv,rcl_sort) map"
-
-text \<open>Well-sortedness for RCL values\<close>
-inductive wfRCV:: "\<Theta> \<Rightarrow> rcl_val \<Rightarrow> b \<Rightarrow> bool" ( " _ \<turnstile> _ : _" [50,50] 50) where
- wfRCV_BBitvecI: "P \<turnstile> (SBitvec bv) : B_bitvec"
-| wfRCV_BIntI: "P \<turnstile> (SNum n) : B_int"
-| wfRCV_BBoolI: "P \<turnstile> (SBool b) : B_bool"
-| wfRCV_BPairI: "\<lbrakk> P \<turnstile> s1 : b1 ; P \<turnstile> s2 : b2 \<rbrakk> \<Longrightarrow> P \<turnstile> (SPair s1 s2) : (B_pair b1 b2)"
-| wfRCV_BConsI: "\<lbrakk> AF_typedef s dclist \<in> set \<Theta>;
- (dc, \<lbrace> x : b | c \<rbrace>) \<in> set dclist ;
- \<Theta> \<turnstile> s1 : b \<rbrakk> \<Longrightarrow> \<Theta> \<turnstile>(SCons s dc s1) : (B_id s)"
-| wfRCV_BConsPI:"\<lbrakk> AF_typedef_poly s bv dclist \<in> set \<Theta>;
- (dc, \<lbrace> x : b | c \<rbrace>) \<in> set dclist ;
- atom bv \<sharp> (\<Theta>, SConsp s dc b' s1, B_app s b');
- \<Theta> \<turnstile> s1 : b[bv::=b']\<^sub>b\<^sub>b \<rbrakk> \<Longrightarrow> \<Theta> \<turnstile>(SConsp s dc b' s1) : (B_app s b')"
-| wfRCV_BUnitI: "P \<turnstile> SUnit : B_unit"
-| wfRCV_BVarI: "P \<turnstile> (SUt n) : (B_var bv)"
-equivariance wfRCV
-nominal_inductive wfRCV
- avoids wfRCV_BConsPI: bv
-proof(goal_cases)
- case (1 s bv dclist \<Theta> dc x b c b' s1)
- then show ?case using fresh_star_def by auto
-next
- case (2 s bv dclist \<Theta> dc x b c s1 b')
- then show ?case by auto
-qed
-
-inductive_cases wfRCV_elims :
- "wfRCV P s B_bitvec"
- "wfRCV P s (B_pair b1 b2)"
- "wfRCV P s (B_int)"
- "wfRCV P s (B_bool)"
- "wfRCV P s (B_id ss)"
- "wfRCV P s (B_var bv)"
- "wfRCV P s (B_unit)"
- "wfRCV P s (B_app tyid b)"
- "wfRCV P (SBitvec bv) b"
- "wfRCV P (SNum n) b"
- "wfRCV P (SBool n) b"
- "wfRCV P (SPair s1 s2) b"
- "wfRCV P (SCons s dc s1) b"
- "wfRCV P (SConsp s dc b' s1) b"
- "wfRCV P SUnit b"
- "wfRCV P (SUt s1) b"
-
-text \<open> Sometimes we want to assert @{text "P \<turnstile> s ~ b[bv=b']"} and we want to know what b is
-however substitution is not injective so we can't write this in terms of @{text "wfRCV"}.
-So we define a relation that makes the components of the substitution explicit. \<close>
-
-inductive wfRCV_subst:: "\<Theta> \<Rightarrow> rcl_val \<Rightarrow> b \<Rightarrow> (bv*b) option \<Rightarrow> bool" where
- wfRCV_subst_BBitvecI: "wfRCV_subst P (SBitvec bv) B_bitvec sub "
-| wfRCV_subst_BIntI: "wfRCV_subst P (SNum n) B_int sub "
-| wfRCV_subst_BBoolI: "wfRCV_subst P (SBool b) B_bool sub "
-| wfRCV_subst_BPairI: "\<lbrakk> wfRCV_subst P s1 b1 sub ; wfRCV_subst P s2 b2 sub \<rbrakk> \<Longrightarrow> wfRCV_subst P (SPair s1 s2) (B_pair b1 b2) sub"
-| wfRCV_subst_BConsI: "\<lbrakk> AF_typedef s dclist \<in> set \<Theta>;
- (dc, \<lbrace> x : b | c \<rbrace>) \<in> set dclist ;
- wfRCV_subst \<Theta> s1 b None \<rbrakk> \<Longrightarrow> wfRCV_subst \<Theta> (SCons s dc s1) (B_id s) sub"
-| wfRCV_subst_BConspI: "\<lbrakk> AF_typedef_poly s bv dclist \<in> set \<Theta>;
- (dc, \<lbrace> x : b | c \<rbrace>) \<in> set dclist ;
- wfRCV_subst \<Theta> s1 (b[bv::=b']\<^sub>b\<^sub>b) sub \<rbrakk> \<Longrightarrow> wfRCV_subst \<Theta> (SConsp s dc b' s1) (B_app s b') sub"
-| wfRCV_subst_BUnitI: "wfRCV_subst P SUnit B_unit sub "
-| wfRCV_subst_BVar1I: "bvar \<noteq> bv \<Longrightarrow> wfRCV_subst P (SUt n) (B_var bv) (Some (bvar,bin))"
-| wfRCV_subst_BVar2I: "\<lbrakk> bvar = bv; wfRCV_subst P s bin None \<rbrakk> \<Longrightarrow> wfRCV_subst P s (B_var bv) (Some (bvar,bin))"
-| wfRCV_subst_BVar3I: "wfRCV_subst P (SUt n) (B_var bv) None"
-equivariance wfRCV_subst
-nominal_inductive wfRCV_subst .
-
-subsection \<open>Evaluation base-types\<close>
-
-inductive eval_b :: "type_valuation \<Rightarrow> b \<Rightarrow> rcl_sort \<Rightarrow> bool" ( "_ \<lbrakk> _ \<rbrakk> ~ _ " ) where
- "v \<lbrakk> B_bool \<rbrakk> ~ S_bool"
-| "v \<lbrakk> B_int \<rbrakk> ~ S_int"
-| "Some s = v bv \<Longrightarrow> v \<lbrakk> B_var bv \<rbrakk> ~ s"
-equivariance eval_b
-nominal_inductive eval_b .
-
-subsection \<open>Wellformed vvaluations\<close>
-
-definition wfI :: "\<Theta> \<Rightarrow> \<Gamma> \<Rightarrow> valuation \<Rightarrow> bool" ( " _ ; _ \<turnstile> _" ) where
- "\<Theta> ; \<Gamma> \<turnstile> i = (\<forall> (x,b,c) \<in> toSet \<Gamma>. \<exists>s. Some s = i x \<and> \<Theta> \<turnstile> s : b)"
-
-subsection \<open>Evaluating Terms\<close>
-
-nominal_function eval_l :: "l \<Rightarrow> rcl_val" ( "\<lbrakk> _ \<rbrakk> " ) where
- "\<lbrakk> L_true \<rbrakk> = SBool True"
-| "\<lbrakk> L_false \<rbrakk> = SBool False"
-| "\<lbrakk> L_num n \<rbrakk> = SNum n"
-| "\<lbrakk> L_unit \<rbrakk> = SUnit"
-| "\<lbrakk> L_bitvec n \<rbrakk> = SBitvec n"
- apply(auto simp: eqvt_def eval_l_graph_aux_def)
- by (metis l.exhaust)
-nominal_termination (eqvt) by lexicographic_order
-
-inductive eval_v :: "valuation \<Rightarrow> v \<Rightarrow> rcl_val \<Rightarrow> bool" ( "_ \<lbrakk> _ \<rbrakk> ~ _ " ) where
- eval_v_litI: "i \<lbrakk> V_lit l \<rbrakk> ~ \<lbrakk> l \<rbrakk> "
-| eval_v_varI: "Some sv = i x \<Longrightarrow> i \<lbrakk> V_var x \<rbrakk> ~ sv"
-| eval_v_pairI: "\<lbrakk> i \<lbrakk> v1 \<rbrakk> ~ s1 ; i \<lbrakk> v2 \<rbrakk> ~ s2 \<rbrakk> \<Longrightarrow> i \<lbrakk> V_pair v1 v2 \<rbrakk> ~ SPair s1 s2"
-| eval_v_consI: "i \<lbrakk> v \<rbrakk> ~ s \<Longrightarrow> i \<lbrakk> V_cons tyid dc v \<rbrakk> ~ SCons tyid dc s"
-| eval_v_conspI: "i \<lbrakk> v \<rbrakk> ~ s \<Longrightarrow> i \<lbrakk> V_consp tyid dc b v \<rbrakk> ~ SConsp tyid dc b s"
-equivariance eval_v
-nominal_inductive eval_v .
-
-inductive_cases eval_v_elims:
- "i \<lbrakk> V_lit l \<rbrakk> ~ s"
- "i \<lbrakk> V_var x \<rbrakk> ~ s"
- "i \<lbrakk> V_pair v1 v2 \<rbrakk> ~ s"
- "i \<lbrakk> V_cons tyid dc v \<rbrakk> ~ s"
- "i \<lbrakk> V_consp tyid dc b v \<rbrakk> ~ s"
-
-inductive eval_e :: "valuation \<Rightarrow> ce \<Rightarrow> rcl_val \<Rightarrow> bool" ( "_ \<lbrakk> _ \<rbrakk> ~ _ " ) where
- eval_e_valI: "i \<lbrakk> v \<rbrakk> ~ sv \<Longrightarrow> i \<lbrakk> CE_val v \<rbrakk> ~ sv"
-| eval_e_plusI: "\<lbrakk> i \<lbrakk> v1 \<rbrakk> ~ SNum n1; i \<lbrakk> v2 \<rbrakk> ~ SNum n2 \<rbrakk> \<Longrightarrow> i \<lbrakk> (CE_op Plus v1 v2) \<rbrakk> ~ (SNum (n1+n2))"
-| eval_e_leqI: "\<lbrakk> i \<lbrakk> v1 \<rbrakk> ~ (SNum n1); i \<lbrakk> v2 \<rbrakk> ~ (SNum n2) \<rbrakk> \<Longrightarrow> i \<lbrakk> (CE_op LEq v1 v2) \<rbrakk> ~ (SBool (n1 \<le> n2))"
-| eval_e_eqI: "\<lbrakk> i \<lbrakk> v1 \<rbrakk> ~ s1; i \<lbrakk> v2 \<rbrakk> ~ s2 \<rbrakk> \<Longrightarrow> i \<lbrakk> (CE_op Eq v1 v2) \<rbrakk> ~ (SBool (s1 = s2))"
-| eval_e_fstI: "\<lbrakk> i \<lbrakk> v \<rbrakk> ~ SPair v1 v2 \<rbrakk> \<Longrightarrow> i \<lbrakk> (CE_fst v) \<rbrakk> ~ v1"
-| eval_e_sndI: "\<lbrakk> i \<lbrakk> v \<rbrakk> ~ SPair v1 v2 \<rbrakk> \<Longrightarrow> i \<lbrakk> (CE_snd v) \<rbrakk> ~ v2"
-| eval_e_concatI:"\<lbrakk> i \<lbrakk> v1 \<rbrakk> ~ (SBitvec bv1); i \<lbrakk> v2 \<rbrakk> ~ (SBitvec bv2) \<rbrakk> \<Longrightarrow> i \<lbrakk> (CE_concat v1 v2) \<rbrakk> ~ (SBitvec (bv1@bv2))"
-| eval_e_lenI:"\<lbrakk> i \<lbrakk> v \<rbrakk> ~ (SBitvec bv) \<rbrakk> \<Longrightarrow> i \<lbrakk> (CE_len v) \<rbrakk> ~ (SNum (int (List.length bv)))"
-equivariance eval_e
-nominal_inductive eval_e .
-
-inductive_cases eval_e_elims:
- "i \<lbrakk> (CE_val v) \<rbrakk> ~ s"
- "i \<lbrakk> (CE_op Plus v1 v2) \<rbrakk> ~ s"
- "i \<lbrakk> (CE_op LEq v1 v2) \<rbrakk> ~ s"
- "i \<lbrakk> (CE_op Eq v1 v2) \<rbrakk> ~ s"
- "i \<lbrakk> (CE_fst v) \<rbrakk> ~ s"
- "i \<lbrakk> (CE_snd v) \<rbrakk> ~ s"
- "i \<lbrakk> (CE_concat v1 v2) \<rbrakk> ~ s"
- "i \<lbrakk> (CE_len v) \<rbrakk> ~ s"
-
-inductive eval_c :: "valuation \<Rightarrow> c \<Rightarrow> bool \<Rightarrow> bool" ( " _ \<lbrakk> _ \<rbrakk> ~ _ ") where
- eval_c_trueI: "i \<lbrakk> C_true \<rbrakk> ~ True"
-| eval_c_falseI:"i \<lbrakk> C_false \<rbrakk> ~ False"
-| eval_c_conjI: "\<lbrakk> i \<lbrakk> c1 \<rbrakk> ~ b1 ; i \<lbrakk> c2 \<rbrakk> ~ b2 \<rbrakk> \<Longrightarrow> i \<lbrakk> (C_conj c1 c2) \<rbrakk> ~ (b1 \<and> b2)"
-| eval_c_disjI: "\<lbrakk> i \<lbrakk> c1 \<rbrakk> ~ b1 ; i \<lbrakk> c2 \<rbrakk> ~ b2 \<rbrakk> \<Longrightarrow> i \<lbrakk> (C_disj c1 c2) \<rbrakk> ~ (b1 \<or> b2)"
-| eval_c_impI:"\<lbrakk> i \<lbrakk> c1 \<rbrakk> ~ b1 ; i \<lbrakk> c2 \<rbrakk> ~ b2 \<rbrakk> \<Longrightarrow> i \<lbrakk> (C_imp c1 c2) \<rbrakk> ~ (b1 \<longrightarrow> b2)"
-| eval_c_notI:"\<lbrakk> i \<lbrakk> c \<rbrakk> ~ b \<rbrakk> \<Longrightarrow> i \<lbrakk> (C_not c) \<rbrakk> ~ (\<not> b)"
-| eval_c_eqI:"\<lbrakk> i \<lbrakk> e1 \<rbrakk> ~ sv1; i \<lbrakk> e2 \<rbrakk> ~ sv2 \<rbrakk> \<Longrightarrow> i \<lbrakk> (C_eq e1 e2) \<rbrakk> ~ (sv1=sv2)"
-equivariance eval_c
-nominal_inductive eval_c .
-
-inductive_cases eval_c_elims:
- "i \<lbrakk> C_true \<rbrakk> ~ True"
- "i \<lbrakk> C_false \<rbrakk> ~ False"
- "i \<lbrakk> (C_conj c1 c2)\<rbrakk> ~ s"
- "i \<lbrakk> (C_disj c1 c2)\<rbrakk> ~ s"
- "i \<lbrakk> (C_imp c1 c2)\<rbrakk> ~ s"
- "i \<lbrakk> (C_not c) \<rbrakk> ~ s"
- "i \<lbrakk> (C_eq e1 e2)\<rbrakk> ~ s"
- "i \<lbrakk> C_true \<rbrakk> ~ s"
- "i \<lbrakk> C_false \<rbrakk> ~ s"
-
-subsection \<open>Satisfiability\<close>
-
-inductive is_satis :: "valuation \<Rightarrow> c \<Rightarrow> bool" ( " _ \<Turnstile> _ " ) where
- "i \<lbrakk> c \<rbrakk> ~ True \<Longrightarrow> i \<Turnstile> c"
-equivariance is_satis
-nominal_inductive is_satis .
-
-nominal_function is_satis_g :: "valuation \<Rightarrow> \<Gamma> \<Rightarrow> bool" ( " _ \<Turnstile> _ " ) where
- "i \<Turnstile> GNil = True"
-| "i \<Turnstile> ((x,b,c) #\<^sub>\<Gamma> G) = ( i \<Turnstile> c \<and> i \<Turnstile> G)"
- apply(auto simp: eqvt_def is_satis_g_graph_aux_def)
- by (metis \<Gamma>.exhaust old.prod.exhaust)
-nominal_termination (eqvt) by lexicographic_order
-
-section \<open>Validity\<close>
-
-nominal_function valid :: "\<Theta> \<Rightarrow> \<B> \<Rightarrow> \<Gamma> \<Rightarrow> c \<Rightarrow> bool" ("_ ; _ ; _ \<Turnstile> _ " [50, 50] 50) where
- "P ; B ; G \<Turnstile> c = ( (P ; B ; G \<turnstile>\<^sub>w\<^sub>f c) \<and> (\<forall>i. (P ; G \<turnstile> i) \<and> i \<Turnstile> G \<longrightarrow> i \<Turnstile> c))"
- by (auto simp: eqvt_def wfI_def valid_graph_aux_def)
-nominal_termination (eqvt) by lexicographic_order
-
-section \<open>Lemmas\<close>
-text \<open>Lemmas needed for Examples\<close>
-
-lemma valid_trueI [intro]:
- fixes G::\<Gamma>
- assumes "P ; B \<turnstile>\<^sub>w\<^sub>f G"
- shows "P ; B ; G \<Turnstile> C_true"
-proof -
- have "\<forall>i. i \<Turnstile> C_true" using is_satis.simps eval_c_trueI by simp
- moreover have "P ; B ; G \<turnstile>\<^sub>w\<^sub>f C_true" using wfC_trueI assms by simp
- ultimately show ?thesis using valid.simps by simp
-qed
-
+(*<*)
+theory RCLogic
+ imports Wellformed
+begin
+ (*>*)
+
+hide_const Syntax.dom
+
+chapter \<open>Refinement Constraint Logic\<close>
+
+text \<open> Semantics for the logic we use in the refinement constraints. It is a multi-sorted, quantifier free
+logic with polymorphic datatypes and linear arithmetic. We could have modelled by using one of the
+encodings to FOL however we wanted to explore using a more direct model. \<close>
+
+section \<open>Evaluation and Satisfiability\<close>
+
+subsection \<open>Valuation\<close>
+
+text \<open> Refinement constraint logic values. SUt is a value for the uninterpreted
+ sort that corresponds to basic type variables. For now we only need one of these universes.
+ We wrap an smt\_val inside it during a process we call 'boxing'
+ which is introduced in the RCLogicL theory \<close>
+nominal_datatype rcl_val = SBitvec "bit list" | SNum int | SBool bool | SPair rcl_val rcl_val |
+ SCons tyid string rcl_val | SConsp tyid string b rcl_val |
+ SUnit | SUt rcl_val
+
+text \<open> RCL sorts. Represent our domains. The universe is the union of all of the these.
+ S\_Ut is the single uninterpreted sort. These map almost directly to basic type
+ but we have them to distinguish syntax (basic types) and semantics (RCL sorts) \<close>
+nominal_datatype rcl_sort = S_bool | S_int | S_unit | S_pair rcl_sort rcl_sort | S_id tyid | S_app tyid rcl_sort | S_bitvec | S_ut
+
+type_synonym valuation = "(x,rcl_val) map"
+
+type_synonym type_valuation = "(bv,rcl_sort) map"
+
+text \<open>Well-sortedness for RCL values\<close>
+inductive wfRCV:: "\<Theta> \<Rightarrow> rcl_val \<Rightarrow> b \<Rightarrow> bool" ( " _ \<turnstile> _ : _" [50,50] 50) where
+ wfRCV_BBitvecI: "P \<turnstile> (SBitvec bv) : B_bitvec"
+| wfRCV_BIntI: "P \<turnstile> (SNum n) : B_int"
+| wfRCV_BBoolI: "P \<turnstile> (SBool b) : B_bool"
+| wfRCV_BPairI: "\<lbrakk> P \<turnstile> s1 : b1 ; P \<turnstile> s2 : b2 \<rbrakk> \<Longrightarrow> P \<turnstile> (SPair s1 s2) : (B_pair b1 b2)"
+| wfRCV_BConsI: "\<lbrakk> AF_typedef s dclist \<in> set \<Theta>;
+ (dc, \<lbrace> x : b | c \<rbrace>) \<in> set dclist ;
+ \<Theta> \<turnstile> s1 : b \<rbrakk> \<Longrightarrow> \<Theta> \<turnstile>(SCons s dc s1) : (B_id s)"
+| wfRCV_BConsPI:"\<lbrakk> AF_typedef_poly s bv dclist \<in> set \<Theta>;
+ (dc, \<lbrace> x : b | c \<rbrace>) \<in> set dclist ;
+ atom bv \<sharp> (\<Theta>, SConsp s dc b' s1, B_app s b');
+ \<Theta> \<turnstile> s1 : b[bv::=b']\<^sub>b\<^sub>b \<rbrakk> \<Longrightarrow> \<Theta> \<turnstile>(SConsp s dc b' s1) : (B_app s b')"
+| wfRCV_BUnitI: "P \<turnstile> SUnit : B_unit"
+| wfRCV_BVarI: "P \<turnstile> (SUt n) : (B_var bv)"
+equivariance wfRCV
+nominal_inductive wfRCV
+ avoids wfRCV_BConsPI: bv
+proof(goal_cases)
+ case (1 s bv dclist \<Theta> dc x b c b' s1)
+ then show ?case using fresh_star_def by auto
+next
+ case (2 s bv dclist \<Theta> dc x b c s1 b')
+ then show ?case by auto
+qed
+
+inductive_cases wfRCV_elims :
+ "wfRCV P s B_bitvec"
+ "wfRCV P s (B_pair b1 b2)"
+ "wfRCV P s (B_int)"
+ "wfRCV P s (B_bool)"
+ "wfRCV P s (B_id ss)"
+ "wfRCV P s (B_var bv)"
+ "wfRCV P s (B_unit)"
+ "wfRCV P s (B_app tyid b)"
+ "wfRCV P (SBitvec bv) b"
+ "wfRCV P (SNum n) b"
+ "wfRCV P (SBool n) b"
+ "wfRCV P (SPair s1 s2) b"
+ "wfRCV P (SCons s dc s1) b"
+ "wfRCV P (SConsp s dc b' s1) b"
+ "wfRCV P SUnit b"
+ "wfRCV P (SUt s1) b"
+
+text \<open> Sometimes we want to assert @{text "P \<turnstile> s ~ b[bv=b']"} and we want to know what b is
+however substitution is not injective so we can't write this in terms of @{text "wfRCV"}.
+So we define a relation that makes the components of the substitution explicit. \<close>
+
+inductive wfRCV_subst:: "\<Theta> \<Rightarrow> rcl_val \<Rightarrow> b \<Rightarrow> (bv*b) option \<Rightarrow> bool" where
+ wfRCV_subst_BBitvecI: "wfRCV_subst P (SBitvec bv) B_bitvec sub "
+| wfRCV_subst_BIntI: "wfRCV_subst P (SNum n) B_int sub "
+| wfRCV_subst_BBoolI: "wfRCV_subst P (SBool b) B_bool sub "
+| wfRCV_subst_BPairI: "\<lbrakk> wfRCV_subst P s1 b1 sub ; wfRCV_subst P s2 b2 sub \<rbrakk> \<Longrightarrow> wfRCV_subst P (SPair s1 s2) (B_pair b1 b2) sub"
+| wfRCV_subst_BConsI: "\<lbrakk> AF_typedef s dclist \<in> set \<Theta>;
+ (dc, \<lbrace> x : b | c \<rbrace>) \<in> set dclist ;
+ wfRCV_subst \<Theta> s1 b None \<rbrakk> \<Longrightarrow> wfRCV_subst \<Theta> (SCons s dc s1) (B_id s) sub"
+| wfRCV_subst_BConspI: "\<lbrakk> AF_typedef_poly s bv dclist \<in> set \<Theta>;
+ (dc, \<lbrace> x : b | c \<rbrace>) \<in> set dclist ;
+ wfRCV_subst \<Theta> s1 (b[bv::=b']\<^sub>b\<^sub>b) sub \<rbrakk> \<Longrightarrow> wfRCV_subst \<Theta> (SConsp s dc b' s1) (B_app s b') sub"
+| wfRCV_subst_BUnitI: "wfRCV_subst P SUnit B_unit sub "
+| wfRCV_subst_BVar1I: "bvar \<noteq> bv \<Longrightarrow> wfRCV_subst P (SUt n) (B_var bv) (Some (bvar,bin))"
+| wfRCV_subst_BVar2I: "\<lbrakk> bvar = bv; wfRCV_subst P s bin None \<rbrakk> \<Longrightarrow> wfRCV_subst P s (B_var bv) (Some (bvar,bin))"
+| wfRCV_subst_BVar3I: "wfRCV_subst P (SUt n) (B_var bv) None"
+equivariance wfRCV_subst
+nominal_inductive wfRCV_subst .
+
+subsection \<open>Evaluation base-types\<close>
+
+inductive eval_b :: "type_valuation \<Rightarrow> b \<Rightarrow> rcl_sort \<Rightarrow> bool" ( "_ \<lbrakk> _ \<rbrakk> ~ _ " ) where
+ "v \<lbrakk> B_bool \<rbrakk> ~ S_bool"
+| "v \<lbrakk> B_int \<rbrakk> ~ S_int"
+| "Some s = v bv \<Longrightarrow> v \<lbrakk> B_var bv \<rbrakk> ~ s"
+equivariance eval_b
+nominal_inductive eval_b .
+
+subsection \<open>Wellformed vvaluations\<close>
+
+definition wfI :: "\<Theta> \<Rightarrow> \<Gamma> \<Rightarrow> valuation \<Rightarrow> bool" ( " _ ; _ \<turnstile> _" ) where
+ "\<Theta> ; \<Gamma> \<turnstile> i = (\<forall> (x,b,c) \<in> toSet \<Gamma>. \<exists>s. Some s = i x \<and> \<Theta> \<turnstile> s : b)"
+
+subsection \<open>Evaluating Terms\<close>
+
+nominal_function eval_l :: "l \<Rightarrow> rcl_val" ( "\<lbrakk> _ \<rbrakk> " ) where
+ "\<lbrakk> L_true \<rbrakk> = SBool True"
+| "\<lbrakk> L_false \<rbrakk> = SBool False"
+| "\<lbrakk> L_num n \<rbrakk> = SNum n"
+| "\<lbrakk> L_unit \<rbrakk> = SUnit"
+| "\<lbrakk> L_bitvec n \<rbrakk> = SBitvec n"
+ apply(auto simp: eqvt_def eval_l_graph_aux_def)
+ by (metis l.exhaust)
+nominal_termination (eqvt) by lexicographic_order
+
+inductive eval_v :: "valuation \<Rightarrow> v \<Rightarrow> rcl_val \<Rightarrow> bool" ( "_ \<lbrakk> _ \<rbrakk> ~ _ " ) where
+ eval_v_litI: "i \<lbrakk> V_lit l \<rbrakk> ~ \<lbrakk> l \<rbrakk> "
+| eval_v_varI: "Some sv = i x \<Longrightarrow> i \<lbrakk> V_var x \<rbrakk> ~ sv"
+| eval_v_pairI: "\<lbrakk> i \<lbrakk> v1 \<rbrakk> ~ s1 ; i \<lbrakk> v2 \<rbrakk> ~ s2 \<rbrakk> \<Longrightarrow> i \<lbrakk> V_pair v1 v2 \<rbrakk> ~ SPair s1 s2"
+| eval_v_consI: "i \<lbrakk> v \<rbrakk> ~ s \<Longrightarrow> i \<lbrakk> V_cons tyid dc v \<rbrakk> ~ SCons tyid dc s"
+| eval_v_conspI: "i \<lbrakk> v \<rbrakk> ~ s \<Longrightarrow> i \<lbrakk> V_consp tyid dc b v \<rbrakk> ~ SConsp tyid dc b s"
+equivariance eval_v
+nominal_inductive eval_v .
+
+inductive_cases eval_v_elims:
+ "i \<lbrakk> V_lit l \<rbrakk> ~ s"
+ "i \<lbrakk> V_var x \<rbrakk> ~ s"
+ "i \<lbrakk> V_pair v1 v2 \<rbrakk> ~ s"
+ "i \<lbrakk> V_cons tyid dc v \<rbrakk> ~ s"
+ "i \<lbrakk> V_consp tyid dc b v \<rbrakk> ~ s"
+
+inductive eval_e :: "valuation \<Rightarrow> ce \<Rightarrow> rcl_val \<Rightarrow> bool" ( "_ \<lbrakk> _ \<rbrakk> ~ _ " ) where
+ eval_e_valI: "i \<lbrakk> v \<rbrakk> ~ sv \<Longrightarrow> i \<lbrakk> CE_val v \<rbrakk> ~ sv"
+| eval_e_plusI: "\<lbrakk> i \<lbrakk> v1 \<rbrakk> ~ SNum n1; i \<lbrakk> v2 \<rbrakk> ~ SNum n2 \<rbrakk> \<Longrightarrow> i \<lbrakk> (CE_op Plus v1 v2) \<rbrakk> ~ (SNum (n1+n2))"
+| eval_e_leqI: "\<lbrakk> i \<lbrakk> v1 \<rbrakk> ~ (SNum n1); i \<lbrakk> v2 \<rbrakk> ~ (SNum n2) \<rbrakk> \<Longrightarrow> i \<lbrakk> (CE_op LEq v1 v2) \<rbrakk> ~ (SBool (n1 \<le> n2))"
+| eval_e_eqI: "\<lbrakk> i \<lbrakk> v1 \<rbrakk> ~ s1; i \<lbrakk> v2 \<rbrakk> ~ s2 \<rbrakk> \<Longrightarrow> i \<lbrakk> (CE_op Eq v1 v2) \<rbrakk> ~ (SBool (s1 = s2))"
+| eval_e_fstI: "\<lbrakk> i \<lbrakk> v \<rbrakk> ~ SPair v1 v2 \<rbrakk> \<Longrightarrow> i \<lbrakk> (CE_fst v) \<rbrakk> ~ v1"
+| eval_e_sndI: "\<lbrakk> i \<lbrakk> v \<rbrakk> ~ SPair v1 v2 \<rbrakk> \<Longrightarrow> i \<lbrakk> (CE_snd v) \<rbrakk> ~ v2"
+| eval_e_concatI:"\<lbrakk> i \<lbrakk> v1 \<rbrakk> ~ (SBitvec bv1); i \<lbrakk> v2 \<rbrakk> ~ (SBitvec bv2) \<rbrakk> \<Longrightarrow> i \<lbrakk> (CE_concat v1 v2) \<rbrakk> ~ (SBitvec (bv1@bv2))"
+| eval_e_lenI:"\<lbrakk> i \<lbrakk> v \<rbrakk> ~ (SBitvec bv) \<rbrakk> \<Longrightarrow> i \<lbrakk> (CE_len v) \<rbrakk> ~ (SNum (int (List.length bv)))"
+equivariance eval_e
+nominal_inductive eval_e .
+
+inductive_cases eval_e_elims:
+ "i \<lbrakk> (CE_val v) \<rbrakk> ~ s"
+ "i \<lbrakk> (CE_op Plus v1 v2) \<rbrakk> ~ s"
+ "i \<lbrakk> (CE_op LEq v1 v2) \<rbrakk> ~ s"
+ "i \<lbrakk> (CE_op Eq v1 v2) \<rbrakk> ~ s"
+ "i \<lbrakk> (CE_fst v) \<rbrakk> ~ s"
+ "i \<lbrakk> (CE_snd v) \<rbrakk> ~ s"
+ "i \<lbrakk> (CE_concat v1 v2) \<rbrakk> ~ s"
+ "i \<lbrakk> (CE_len v) \<rbrakk> ~ s"
+
+inductive eval_c :: "valuation \<Rightarrow> c \<Rightarrow> bool \<Rightarrow> bool" ( " _ \<lbrakk> _ \<rbrakk> ~ _ ") where
+ eval_c_trueI: "i \<lbrakk> C_true \<rbrakk> ~ True"
+| eval_c_falseI:"i \<lbrakk> C_false \<rbrakk> ~ False"
+| eval_c_conjI: "\<lbrakk> i \<lbrakk> c1 \<rbrakk> ~ b1 ; i \<lbrakk> c2 \<rbrakk> ~ b2 \<rbrakk> \<Longrightarrow> i \<lbrakk> (C_conj c1 c2) \<rbrakk> ~ (b1 \<and> b2)"
+| eval_c_disjI: "\<lbrakk> i \<lbrakk> c1 \<rbrakk> ~ b1 ; i \<lbrakk> c2 \<rbrakk> ~ b2 \<rbrakk> \<Longrightarrow> i \<lbrakk> (C_disj c1 c2) \<rbrakk> ~ (b1 \<or> b2)"
+| eval_c_impI:"\<lbrakk> i \<lbrakk> c1 \<rbrakk> ~ b1 ; i \<lbrakk> c2 \<rbrakk> ~ b2 \<rbrakk> \<Longrightarrow> i \<lbrakk> (C_imp c1 c2) \<rbrakk> ~ (b1 \<longrightarrow> b2)"
+| eval_c_notI:"\<lbrakk> i \<lbrakk> c \<rbrakk> ~ b \<rbrakk> \<Longrightarrow> i \<lbrakk> (C_not c) \<rbrakk> ~ (\<not> b)"
+| eval_c_eqI:"\<lbrakk> i \<lbrakk> e1 \<rbrakk> ~ sv1; i \<lbrakk> e2 \<rbrakk> ~ sv2 \<rbrakk> \<Longrightarrow> i \<lbrakk> (C_eq e1 e2) \<rbrakk> ~ (sv1=sv2)"
+equivariance eval_c
+nominal_inductive eval_c .
+
+inductive_cases eval_c_elims:
+ "i \<lbrakk> C_true \<rbrakk> ~ True"
+ "i \<lbrakk> C_false \<rbrakk> ~ False"
+ "i \<lbrakk> (C_conj c1 c2)\<rbrakk> ~ s"
+ "i \<lbrakk> (C_disj c1 c2)\<rbrakk> ~ s"
+ "i \<lbrakk> (C_imp c1 c2)\<rbrakk> ~ s"
+ "i \<lbrakk> (C_not c) \<rbrakk> ~ s"
+ "i \<lbrakk> (C_eq e1 e2)\<rbrakk> ~ s"
+ "i \<lbrakk> C_true \<rbrakk> ~ s"
+ "i \<lbrakk> C_false \<rbrakk> ~ s"
+
+subsection \<open>Satisfiability\<close>
+
+inductive is_satis :: "valuation \<Rightarrow> c \<Rightarrow> bool" ( " _ \<Turnstile> _ " ) where
+ "i \<lbrakk> c \<rbrakk> ~ True \<Longrightarrow> i \<Turnstile> c"
+equivariance is_satis
+nominal_inductive is_satis .
+
+nominal_function is_satis_g :: "valuation \<Rightarrow> \<Gamma> \<Rightarrow> bool" ( " _ \<Turnstile> _ " ) where
+ "i \<Turnstile> GNil = True"
+| "i \<Turnstile> ((x,b,c) #\<^sub>\<Gamma> G) = ( i \<Turnstile> c \<and> i \<Turnstile> G)"
+ apply(auto simp: eqvt_def is_satis_g_graph_aux_def)
+ by (metis \<Gamma>.exhaust old.prod.exhaust)
+nominal_termination (eqvt) by lexicographic_order
+
+section \<open>Validity\<close>
+
+nominal_function valid :: "\<Theta> \<Rightarrow> \<B> \<Rightarrow> \<Gamma> \<Rightarrow> c \<Rightarrow> bool" ("_ ; _ ; _ \<Turnstile> _ " [50, 50] 50) where
+ "P ; B ; G \<Turnstile> c = ( (P ; B ; G \<turnstile>\<^sub>w\<^sub>f c) \<and> (\<forall>i. (P ; G \<turnstile> i) \<and> i \<Turnstile> G \<longrightarrow> i \<Turnstile> c))"
+ by (auto simp: eqvt_def wfI_def valid_graph_aux_def)
+nominal_termination (eqvt) by lexicographic_order
+
+section \<open>Lemmas\<close>
+text \<open>Lemmas needed for Examples\<close>
+
+lemma valid_trueI [intro]:
+ fixes G::\<Gamma>
+ assumes "P ; B \<turnstile>\<^sub>w\<^sub>f G"
+ shows "P ; B ; G \<Turnstile> C_true"
+proof -
+ have "\<forall>i. i \<Turnstile> C_true" using is_satis.simps eval_c_trueI by simp
+ moreover have "P ; B ; G \<turnstile>\<^sub>w\<^sub>f C_true" using wfC_trueI assms by simp
+ ultimately show ?thesis using valid.simps by simp
+qed
+
end
\ No newline at end of file
diff --git a/thys/MiniSail/RCLogicL.thy b/thys/MiniSail/RCLogicL.thy
--- a/thys/MiniSail/RCLogicL.thy
+++ b/thys/MiniSail/RCLogicL.thy
@@ -1,2729 +1,2729 @@
-(*<*)
-theory RCLogicL
- imports RCLogic WellformedL
-begin
- (*>*)
-
-hide_const Syntax.dom
-
-chapter \<open>Refinement Constraint Logic Lemmas\<close>
-
-section \<open>Lemmas\<close>
-
-lemma wfI_domi:
- assumes "\<Theta> ; \<Gamma> \<turnstile> i"
- shows "fst ` toSet \<Gamma> \<subseteq> dom i"
- using wfI_def toSet.simps assms by fastforce
-
-lemma wfI_lookup:
- fixes G::\<Gamma> and b::b
- assumes "Some (b,c) = lookup G x" and "P ; G \<turnstile> i" and "Some s = i x" and "P ; B \<turnstile>\<^sub>w\<^sub>f G"
- shows "P \<turnstile> s : b"
-proof -
- have "(x,b,c) \<in> toSet G" using lookup.simps assms
- using lookup_in_g by blast
- then obtain s' where *:"Some s' = i x \<and> wfRCV P s' b" using wfI_def assms by auto
- hence "s' = s" using assms by (metis option.inject)
- thus ?thesis using * by auto
-qed
-
-lemma wfI_restrict_weakening:
- assumes "wfI \<Theta> \<Gamma>' i'" and "i = restrict_map i' (fst `toSet \<Gamma>)" and "toSet \<Gamma> \<subseteq> toSet \<Gamma>'"
- shows "\<Theta> ; \<Gamma> \<turnstile> i"
-proof -
- { fix x
- assume "x \<in> toSet \<Gamma>"
- have "case x of (x, b, c) \<Rightarrow> \<exists>s. Some s = i x \<and> \<Theta> \<turnstile> s : b" using assms wfI_def
- proof -
- have "case x of (x, b, c) \<Rightarrow> \<exists>s. Some s = i' x \<and> \<Theta> \<turnstile> s : b"
- using \<open>x \<in> toSet \<Gamma>\<close> assms wfI_def by auto
- then have "\<exists>s. Some s = i (fst x) \<and> wfRCV \<Theta> s (fst (snd x))"
- by (simp add: \<open>x \<in> toSet \<Gamma>\<close> assms(2) case_prod_unfold)
- then show ?thesis
- by (simp add: case_prod_unfold)
- qed
- }
- thus ?thesis using wfI_def assms by auto
-qed
-
-lemma wfI_suffix:
- fixes G::\<Gamma>
- assumes "wfI P (G'@G) i" and "P ; B \<turnstile>\<^sub>w\<^sub>f G"
- shows "P ; G \<turnstile> i"
- using wfI_def append_g.simps assms toSet.simps by simp
-
-lemma wfI_replace_inside:
- assumes "wfI \<Theta> (\<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>) i"
- shows "wfI \<Theta> (\<Gamma>' @ (x, b, c') #\<^sub>\<Gamma> \<Gamma>) i"
- using wfI_def toSet_splitP assms by simp
-
-section \<open>Existence of evaluation\<close>
-
-lemma eval_l_base:
- "\<Theta> \<turnstile> \<lbrakk> l \<rbrakk> : (base_for_lit l)"
- apply(nominal_induct l rule:l.strong_induct)
- using wfRCV.intros eval_l.simps base_for_lit.simps by auto+
-
-lemma obtain_fresh_bv_dclist:
- fixes tm::"'a::fs"
- assumes "(dc, \<lbrace> x : b | c \<rbrace>) \<in> set dclist"
- obtains bv1::bv and dclist1 x1 b1 c1 where "AF_typedef_poly tyid bv dclist = AF_typedef_poly tyid bv1 dclist1
- \<and> (dc, \<lbrace> x1 : b1 | c1 \<rbrace>) \<in> set dclist1 \<and> atom bv1 \<sharp> tm"
-proof -
- obtain bv1 dclist1 where "AF_typedef_poly tyid bv dclist = AF_typedef_poly tyid bv1 dclist1 \<and> atom bv1 \<sharp> tm"
- using obtain_fresh_bv by metis
- moreover hence "[[atom bv]]lst. dclist = [[atom bv1]]lst. dclist1" using type_def.eq_iff by metis
- moreover then obtain x1 b1 c1 where "(dc, \<lbrace> x1 : b1 | c1 \<rbrace>) \<in> set dclist1" using td_lookup_eq_iff assms by metis
- ultimately show ?thesis using that by blast
-qed
-
-lemma obtain_fresh_bv_dclist_b_iff:
- fixes tm::"'a::fs"
- assumes "(dc, \<lbrace> x : b | c \<rbrace>) \<in> set dclist" and "AF_typedef_poly tyid bv dclist \<in> set P" and "\<turnstile>\<^sub>w\<^sub>f P"
- obtains bv1::bv and dclist1 x1 b1 c1 where "AF_typedef_poly tyid bv dclist = AF_typedef_poly tyid bv1 dclist1
- \<and> (dc, \<lbrace> x1 : b1 | c1 \<rbrace>) \<in> set dclist1 \<and> atom bv1 \<sharp> tm \<and> b[bv::=b']\<^sub>b\<^sub>b=b1[bv1::=b']\<^sub>b\<^sub>b"
-proof -
- obtain bv1 dclist1 x1 b1 c1 where *:"AF_typedef_poly tyid bv dclist = AF_typedef_poly tyid bv1 dclist1 \<and> atom bv1 \<sharp> tm
- \<and> (dc, \<lbrace> x1 : b1 | c1 \<rbrace>) \<in> set dclist1" using obtain_fresh_bv_dclist assms by metis
-
- hence "AF_typedef_poly tyid bv1 dclist1 \<in> set P" using assms by metis
-
- hence "b[bv::=b']\<^sub>b\<^sub>b = b1[bv1::=b']\<^sub>b\<^sub>b"
- using wfTh_typedef_poly_b_eq_iff[OF assms(2) assms(1) _ _ assms(3),of bv1 dclist1 x1 b1 c1 b'] * by metis
-
- from this that show ?thesis using * by metis
-qed
-
-lemma eval_v_exist:
- fixes \<Gamma>::\<Gamma> and v::v and b::b
- assumes "P ; \<Gamma> \<turnstile> i" and "P ; B ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b"
- shows "\<exists>s. i \<lbrakk> v \<rbrakk> ~ s \<and> P \<turnstile> s : b "
- using assms proof(nominal_induct v arbitrary: b rule:v.strong_induct)
- case (V_lit x)
- then show ?case using eval_l_base eval_v.intros eval_l.simps wfV_elims rcl_val.supp pure_supp by metis
-next
- case (V_var x)
- then obtain c where *:"Some (b,c) = lookup \<Gamma> x" using wfV_elims by metis
- hence "x \<in> fst ` toSet \<Gamma>" using lookup_x by blast
- hence "x \<in> dom i" using wfI_domi using assms by blast
- then obtain s where "i x = Some s" by auto
- moreover hence "P \<turnstile> s : b" using wfRCV.intros wfI_lookup * V_var
- by (metis wfV_wf)
-
- ultimately show ?case using eval_v.intros rcl_val.supp b.supp by metis
-next
- case (V_pair v1 v2)
- then obtain b1 and b2 where *:"P ; B ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v1 : b1 \<and> P ; B ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v2 : b2 \<and> b = B_pair b1 b2" using wfV_elims by metis
- then obtain s1 and s2 where "eval_v i v1 s1 \<and> wfRCV P s1 b1 \<and> eval_v i v2 s2 \<and> wfRCV P s2 b2" using V_pair by metis
- thus ?case using eval_v.intros wfRCV.intros * by metis
-next
- case (V_cons tyid dc v)
- then obtain s and b'::b and dclist and x::x and c::c where "(wfV P B \<Gamma> v b') \<and> i \<lbrakk> v \<rbrakk> ~ s \<and> P \<turnstile> s : b' \<and> b = B_id tyid \<and>
- AF_typedef tyid dclist \<in> set P \<and> (dc, \<lbrace> x : b' | c \<rbrace>) \<in> set dclist" using wfV_elims(4) by metis
- then show ?case using eval_v.intros(4) wfRCV.intros(5) V_cons by metis
-next
- case (V_consp tyid dc b' v)
-
- obtain b'a::b and bv and dclist and x::x and c::c where *:"(wfV P B \<Gamma> v b'a[bv::=b']\<^sub>b\<^sub>b) \<and> b = B_app tyid b' \<and>
- AF_typedef_poly tyid bv dclist \<in> set P \<and> (dc, \<lbrace> x : b'a | c \<rbrace>) \<in> set dclist \<and>
- atom bv \<sharp> (P, B_app tyid b',B) " using wf_strong_elim(1)[OF V_consp(3)] by metis
-
- then obtain s where **:"i \<lbrakk> v \<rbrakk> ~ s \<and> P \<turnstile> s : b'a[bv::=b']\<^sub>b\<^sub>b" using V_consp by auto
-
- have " \<turnstile>\<^sub>w\<^sub>f P" using wfX_wfY V_consp by metis
- then obtain bv1::bv and dclist1 x1 b1 c1 where 3:"AF_typedef_poly tyid bv dclist = AF_typedef_poly tyid bv1 dclist1
- \<and> (dc, \<lbrace> x1 : b1 | c1 \<rbrace>) \<in> set dclist1 \<and> atom bv1 \<sharp> (P, SConsp tyid dc b' s, B_app tyid b')
- \<and> b'a[bv::=b']\<^sub>b\<^sub>b = b1[bv1::=b']\<^sub>b\<^sub>b"
- using * obtain_fresh_bv_dclist_b_iff by blast
-
- have " i \<lbrakk> V_consp tyid dc b' v \<rbrakk> ~ SConsp tyid dc b' s" using eval_v.intros by (simp add: "**")
-
- moreover have " P \<turnstile> SConsp tyid dc b' s : B_app tyid b'" proof
- show \<open>AF_typedef_poly tyid bv1 dclist1 \<in> set P\<close> using 3 * by metis
- show \<open>(dc, \<lbrace> x1 : b1 | c1 \<rbrace>) \<in> set dclist1\<close> using 3 by auto
- thus \<open>atom bv1 \<sharp> (P, SConsp tyid dc b' s, B_app tyid b')\<close> using * 3 fresh_prodN by metis
- show \<open> P \<turnstile> s : b1[bv1::=b']\<^sub>b\<^sub>b\<close> using 3 ** by auto
- qed
-
- ultimately show ?case using eval_v.intros wfRCV.intros V_consp * by auto
-qed
-
-lemma eval_v_uniqueness:
- fixes v::v
- assumes "i \<lbrakk> v \<rbrakk> ~ s" and "i \<lbrakk> v \<rbrakk> ~ s'"
- shows "s=s'"
- using assms proof(nominal_induct v arbitrary: s s' rule:v.strong_induct)
- case (V_lit x)
- then show ?case using eval_v_elims eval_l.simps by metis
-next
- case (V_var x)
- then show ?case using eval_v_elims by (metis option.inject)
-next
- case (V_pair v1 v2)
- obtain s1 and s2 where s:"i \<lbrakk> v1 \<rbrakk> ~ s1 \<and> i \<lbrakk> v2 \<rbrakk> ~ s2 \<and> s = SPair s1 s2" using eval_v_elims V_pair by metis
- obtain s1' and s2' where s':"i \<lbrakk> v1 \<rbrakk> ~ s1' \<and> i \<lbrakk> v2 \<rbrakk> ~ s2' \<and> s' = SPair s1' s2'" using eval_v_elims V_pair by metis
- then show ?case using eval_v_elims using V_pair s s' by auto
-next
- case (V_cons tyid dc v1)
- obtain sv1 where 1:"i \<lbrakk> v1 \<rbrakk> ~ sv1 \<and> s = SCons tyid dc sv1" using eval_v_elims V_cons by metis
- moreover obtain sv2 where 2:"i \<lbrakk> v1 \<rbrakk> ~ sv2 \<and> s' = SCons tyid dc sv2" using eval_v_elims V_cons by metis
- ultimately have "sv1 = sv2" using V_cons by auto
- then show ?case using 1 2 by auto
-next
- case (V_consp tyid dc b v1)
- then show ?case using eval_v_elims by metis
-
-qed
-
-lemma eval_v_base:
- fixes \<Gamma>::\<Gamma> and v::v and b::b
- assumes "P ; \<Gamma> \<turnstile> i" and "P ; B ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b" and "i \<lbrakk> v \<rbrakk> ~ s"
- shows "P \<turnstile> s : b"
- using eval_v_exist eval_v_uniqueness assms by metis
-
-lemma eval_e_uniqueness:
- fixes e::ce
- assumes "i \<lbrakk> e \<rbrakk> ~ s" and "i \<lbrakk> e \<rbrakk> ~ s'"
- shows "s=s'"
- using assms proof(nominal_induct e arbitrary: s s' rule:ce.strong_induct)
- case (CE_val x)
- then show ?case using eval_v_uniqueness eval_e_elims by metis
-next
- case (CE_op opp x1 x2)
- consider "opp = Plus" | "opp = LEq" | "opp = Eq" using opp.exhaust by metis
- thus ?case proof(cases)
- case 1
- hence a1:"eval_e i (CE_op Plus x1 x2) s" and a2:"eval_e i (CE_op Plus x1 x2) s'" using CE_op by auto
- then show ?thesis using eval_e_elims(2)[OF a1] eval_e_elims(2)[OF a2]
- CE_op eval_e_plusI
- by (metis rcl_val.eq_iff(2))
- next
- case 2
- hence a1:"eval_e i (CE_op LEq x1 x2) s" and a2:"eval_e i (CE_op LEq x1 x2) s'" using CE_op by auto
- then show ?thesis using eval_v_uniqueness eval_e_elims(3)[OF a1] eval_e_elims(3)[OF a2]
- CE_op eval_e_plusI
- by (metis rcl_val.eq_iff(2))
- next
- case 3
- hence a1:"eval_e i (CE_op Eq x1 x2) s" and a2:"eval_e i (CE_op Eq x1 x2) s'" using CE_op by auto
- then show ?thesis using eval_v_uniqueness eval_e_elims(4)[OF a1] eval_e_elims(4)[OF a2]
- CE_op eval_e_plusI
- by (metis rcl_val.eq_iff(2))
- qed
-next
- case (CE_concat x1 x2)
- hence a1:"eval_e i (CE_concat x1 x2) s" and a2:"eval_e i (CE_concat x1 x2) s'" using CE_concat by auto
- show ?case using eval_e_elims(7)[OF a1] eval_e_elims(7)[OF a2] CE_concat eval_e_concatI rcl_val.eq_iff
- proof -
- assume "\<And>P. (\<And>bv1 bv2. \<lbrakk>s' = SBitvec (bv1 @ bv2); i \<lbrakk> x1 \<rbrakk> ~ SBitvec bv1 ; i \<lbrakk> x2 \<rbrakk> ~ SBitvec bv2 \<rbrakk> \<Longrightarrow> P) \<Longrightarrow> P"
- obtain bbs :: "bit list" and bbsa :: "bit list" where
- "i \<lbrakk> x2 \<rbrakk> ~ SBitvec bbs \<and> i \<lbrakk> x1 \<rbrakk> ~ SBitvec bbsa \<and> SBitvec (bbsa @ bbs) = s'"
- by (metis \<open>\<And>P. (\<And>bv1 bv2. \<lbrakk>s' = SBitvec (bv1 @ bv2); i \<lbrakk> x1 \<rbrakk> ~ SBitvec bv1 ; i \<lbrakk> x2 \<rbrakk> ~ SBitvec bv2 \<rbrakk> \<Longrightarrow> P) \<Longrightarrow> P\<close>) (* 93 ms *)
- then have "s' = s"
- by (metis (no_types) \<open>\<And>P. (\<And>bv1 bv2. \<lbrakk>s = SBitvec (bv1 @ bv2); i \<lbrakk> x1 \<rbrakk> ~ SBitvec bv1 ; i \<lbrakk> x2 \<rbrakk> ~ SBitvec bv2 \<rbrakk> \<Longrightarrow> P) \<Longrightarrow> P\<close> \<open>\<And>s' s. \<lbrakk>i \<lbrakk> x1 \<rbrakk> ~ s ; i \<lbrakk> x1 \<rbrakk> ~ s' \<rbrakk> \<Longrightarrow> s = s'\<close> \<open>\<And>s' s. \<lbrakk>i \<lbrakk> x2 \<rbrakk> ~ s ; i \<lbrakk> x2 \<rbrakk> ~ s' \<rbrakk> \<Longrightarrow> s = s'\<close> rcl_val.eq_iff(1)) (* 125 ms *)
- then show ?thesis
- by metis (* 0.0 ms *)
- qed
-next
- case (CE_fst x)
- then show ?case using eval_v_uniqueness by (meson eval_e_elims rcl_val.eq_iff)
-next
- case (CE_snd x)
- then show ?case using eval_v_uniqueness by (meson eval_e_elims rcl_val.eq_iff)
-next
- case (CE_len x)
- then show ?case using eval_e_elims rcl_val.eq_iff
- proof -
- obtain bbs :: "rcl_val \<Rightarrow> ce \<Rightarrow> (x \<Rightarrow> rcl_val option) \<Rightarrow> bit list" where
- "\<forall>x0 x1 x2. (\<exists>v3. x0 = SNum (int (length v3)) \<and> x2 \<lbrakk> x1 \<rbrakk> ~ SBitvec v3 ) = (x0 = SNum (int (length (bbs x0 x1 x2))) \<and> x2 \<lbrakk> x1 \<rbrakk> ~ SBitvec (bbs x0 x1 x2) )"
- by moura (* 0.0 ms *)
- then have "\<forall>f c r. \<not> f \<lbrakk> [| c |]\<^sup>c\<^sup>e \<rbrakk> ~ r \<or> r = SNum (int (length (bbs r c f))) \<and> f \<lbrakk> c \<rbrakk> ~ SBitvec (bbs r c f)"
- by (meson eval_e_elims(8)) (* 46 ms *)
- then show ?thesis
- by (metis (no_types) CE_len.hyps CE_len.prems(1) CE_len.prems(2) rcl_val.eq_iff(1)) (* 31 ms *)
- qed
-
-qed
-
-lemma wfV_eval_bitvec:
- fixes v::v
- assumes "P ; B ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : B_bitvec" and "P ; \<Gamma> \<turnstile> i"
- shows "\<exists>bv. eval_v i v (SBitvec bv)"
-proof -
- obtain s where "i \<lbrakk> v \<rbrakk> ~ s \<and> wfRCV P s B_bitvec" using eval_v_exist assms by metis
- moreover then obtain bv where "s = SBitvec bv" using wfRCV_elims(1)[of P s] by metis
- ultimately show ?thesis by metis
-qed
-
-lemma wfV_eval_pair:
- fixes v::v
- assumes "P ; B ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : B_pair b1 b2" and "P ; \<Gamma> \<turnstile> i"
- shows "\<exists>s1 s2. eval_v i v (SPair s1 s2)"
-proof -
- obtain s where "i \<lbrakk> v \<rbrakk> ~ s \<and> wfRCV P s (B_pair b1 b2)" using eval_v_exist assms by metis
- moreover then obtain s1 and s2 where "s = SPair s1 s2" using wfRCV_elims(2)[of P s] by metis
- ultimately show ?thesis by metis
-qed
-
-lemma wfV_eval_int:
- fixes v::v
- assumes "P ; B ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : B_int" and "P ; \<Gamma> \<turnstile> i"
- shows "\<exists>n. eval_v i v (SNum n)"
-proof -
- obtain s where "i \<lbrakk> v \<rbrakk> ~ s \<and> wfRCV P s (B_int)" using eval_v_exist assms by metis
- moreover then obtain n where "s = SNum n" using wfRCV_elims(3)[of P s] by metis
- ultimately show ?thesis by metis
-qed
-
-text \<open>Well sorted value with a well sorted valuation evaluates\<close>
-lemma wfI_wfV_eval_v:
- fixes v::v and b::b
- assumes "\<Theta> ; B ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b" and "wfI \<Theta> \<Gamma> i"
- shows "\<exists>s. i \<lbrakk> v \<rbrakk> ~ s \<and> \<Theta> \<turnstile> s : b"
- using eval_v_exist assms by auto
-
-lemma wfI_wfCE_eval_e:
- fixes e::ce and b::b
- assumes "wfCE P B G e b" and "P ; G \<turnstile> i"
- shows "\<exists>s. i \<lbrakk> e \<rbrakk> ~ s \<and> P \<turnstile> s : b"
- using assms proof(nominal_induct e arbitrary: b rule: ce.strong_induct)
- case (CE_val v)
- obtain s where "i \<lbrakk> v \<rbrakk> ~ s \<and> P \<turnstile> s : b" using wfI_wfV_eval_v[of P B G v b i] assms wfCE_elims(1) [of P B G v b] CE_val by auto
- then show ?case using CE_val eval_e.intros(1)[of i v s ] by auto
-next
- case (CE_op opp v1 v2)
-
- consider "opp =Plus" | "opp=LEq" | "opp=Eq" using opp.exhaust by auto
-
- thus ?case proof(cases)
- case 1
- hence "wfCE P B G v1 B_int \<and> wfCE P B G v2 B_int" using wfCE_elims(2) CE_op
-
- by blast
- then obtain s1 and s2 where *: "eval_e i v1 s1 \<and> wfRCV P s1 B_int \<and> eval_e i v2 s2 \<and> wfRCV P s2 B_int"
- using wfI_wfV_eval_v CE_op by metis
- then obtain n1 and n2 where **:"s2=SNum n2 \<and> s1 = SNum n1" using wfRCV_elims by meson
- hence "eval_e i (CE_op Plus v1 v2) (SNum (n1+n2))" using eval_e_plusI * ** by simp
- moreover have "wfRCV P (SNum (n1+n2)) B_int" using wfRCV.intros by auto
- ultimately show ?thesis using 1
- using CE_op.prems(1) wfCE_elims(2) by blast
- next
- case 2
- hence "wfCE P B G v1 B_int \<and> wfCE P B G v2 B_int" using wfCE_elims(3) CE_op
- by blast
- then obtain s1 and s2 where *: "eval_e i v1 s1 \<and> wfRCV P s1 B_int \<and> eval_e i v2 s2 \<and> wfRCV P s2 B_int"
- using wfI_wfV_eval_v CE_op by metis
- then obtain n1 and n2 where **:"s2=SNum n2 \<and> s1 = SNum n1" using wfRCV_elims by meson
- hence "eval_e i (CE_op LEq v1 v2) (SBool (n1 \<le> n2))" using eval_e_leqI * ** by simp
- moreover have "wfRCV P (SBool (n1\<le>n2)) B_bool" using wfRCV.intros by auto
- ultimately show ?thesis using 2
- using CE_op.prems wfCE_elims by metis
- next
- case 3
- then obtain b2 where "wfCE P B G v1 b2 \<and> wfCE P B G v2 b2" using wfCE_elims(9) CE_op
- by blast
- then obtain s1 and s2 where *: "eval_e i v1 s1 \<and> wfRCV P s1 b2 \<and> eval_e i v2 s2 \<and> wfRCV P s2 b2"
- using wfI_wfV_eval_v CE_op by metis
- hence "eval_e i (CE_op Eq v1 v2) (SBool (s1 = s2))" using eval_e_leqI *
- by (simp add: eval_e_eqI)
- moreover have "wfRCV P (SBool (s1 = s2)) B_bool" using wfRCV.intros by auto
- ultimately show ?thesis using 3
- using CE_op.prems wfCE_elims by metis
- qed
-next
- case (CE_concat v1 v2)
- then obtain s1 and s2 where *:"b = B_bitvec \<and> eval_e i v1 s1 \<and> eval_e i v2 s2 \<and>
- wfRCV P s1 B_bitvec \<and> wfRCV P s2 B_bitvec" using
- CE_concat
- by (meson wfCE_elims(6))
- thus ?case using eval_e_concatI wfRCV.intros(1) wfRCV_elims
- proof -
- obtain bbs :: "type_def list \<Rightarrow> rcl_val \<Rightarrow> bit list" where
- "\<forall>ts s. \<not> ts \<turnstile> s : B_bitvec \<or> s = SBitvec (bbs ts s)"
- using wfRCV_elims(1) by moura (* 156 ms *)
- then show ?thesis
- by (metis (no_types) "local.*" wfRCV_BBitvecI eval_e_concatI) (* 78 ms *)
- qed
-next
- case (CE_fst v1)
- thus ?case using eval_e_fstI wfRCV.intros wfCE_elims wfI_wfV_eval_v
- by (metis wfRCV_elims(2) rcl_val.eq_iff(4))
-next
- case (CE_snd v1)
- thus ?case using eval_e_sndI wfRCV.intros wfCE_elims wfI_wfV_eval_v
- by (metis wfRCV_elims(2) rcl_val.eq_iff(4))
-next
- case (CE_len x)
- thus ?case using eval_e_lenI wfRCV.intros wfCE_elims wfI_wfV_eval_v wfV_eval_bitvec
- by (metis wfRCV_elims(1))
-qed
-
-lemma eval_e_exist:
- fixes \<Gamma>::\<Gamma> and e::ce
- assumes "P ; \<Gamma> \<turnstile> i" and "P ; B ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f e : b"
- shows "\<exists>s. i \<lbrakk> e \<rbrakk> ~ s"
- using assms proof(nominal_induct e arbitrary: b rule:ce.strong_induct)
- case (CE_val v)
- then show ?case using eval_v_exist wfCE_elims eval_e.intros by metis
-next
- case (CE_op op v1 v2)
-
- show ?case proof(rule opp.exhaust)
- assume \<open>op = Plus\<close>
- hence "P ; B ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v1 : B_int \<and> P ; B ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v2 : B_int \<and> b = B_int" using wfCE_elims CE_op by metis
- then obtain n1 and n2 where "eval_e i v1 (SNum n1) \<and> eval_e i v2 (SNum n2)" using CE_op eval_v_exist wfV_eval_int
- by (metis wfI_wfCE_eval_e wfRCV_elims(3))
- then show \<open>\<exists>a. eval_e i (CE_op op v1 v2) a\<close> using eval_e_plusI[of i v1 _ v2] \<open>op=Plus\<close> by auto
- next
- assume \<open>op = LEq\<close>
- hence "P ; B ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v1 : B_int \<and> P ; B ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v2 : B_int \<and> b = B_bool" using wfCE_elims CE_op by metis
- then obtain n1 and n2 where "eval_e i v1 (SNum n1) \<and> eval_e i v2 (SNum n2)" using CE_op eval_v_exist wfV_eval_int
- by (metis wfI_wfCE_eval_e wfRCV_elims(3))
- then show \<open>\<exists>a. eval_e i (CE_op op v1 v2) a\<close> using eval_e_leqI[of i v1 _ v2] eval_v_exist \<open>op=LEq\<close> CE_op by auto
- next
- assume \<open>op = Eq\<close>
- then obtain b1 where "P ; B ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v1 : b1 \<and> P ; B ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v2 : b1 \<and> b = B_bool" using wfCE_elims CE_op by metis
- then obtain s1 and s2 where "eval_e i v1 s1 \<and> eval_e i v2 s2" using CE_op eval_v_exist wfV_eval_int
- by (metis wfI_wfCE_eval_e wfRCV_elims(3))
- then show \<open>\<exists>a. eval_e i (CE_op op v1 v2) a\<close> using eval_e_eqI[of i v1 _ v2] eval_v_exist \<open>op=Eq\<close> CE_op by auto
- qed
-next
- case (CE_concat v1 v2)
- then obtain bv1 and bv2 where "eval_e i v1 (SBitvec bv1) \<and> eval_e i v2 (SBitvec bv2)"
- using wfV_eval_bitvec wfCE_elims(6)
- by (meson eval_e_elims(7) wfI_wfCE_eval_e)
- then show ?case using eval_e.intros by metis
-next
- case (CE_fst ce)
- then obtain b2 where b:"P ; B ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ce : B_pair b b2" using wfCE_elims by metis
- then obtain s where s:"i \<lbrakk> ce \<rbrakk> ~ s" using CE_fst by auto
- then obtain s1 and s2 where "s = (SPair s1 s2)" using eval_e_elims(4) CE_fst wfI_wfCE_eval_e[of P B \<Gamma> ce "B_pair b b2" i,OF b] wfRCV_elims(2)[of P s b b2]
- by (metis eval_e_uniqueness)
- then show ?case using s eval_e.intros by metis
-next
- case (CE_snd ce)
- then obtain b1 where b:"P ; B ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ce : B_pair b1 b" using wfCE_elims by metis
- then obtain s where s:"i \<lbrakk> ce \<rbrakk> ~ s" using CE_snd by auto
- then obtain s1 and s2 where "s = (SPair s1 s2)"
- using eval_e_elims(5) CE_snd wfI_wfCE_eval_e[of P B \<Gamma> ce "B_pair b1 b" i,OF b] wfRCV_elims(2)[of P s b b1]
- eval_e_uniqueness
- by (metis wfRCV_elims(2))
- then show ?case using s eval_e.intros by metis
-next
- case (CE_len v1)
- then obtain bv1 where "eval_e i v1 (SBitvec bv1)"
- using wfV_eval_bitvec CE_len wfCE_elims eval_e_uniqueness
- by (metis eval_e_elims(7) wfCE_concatI wfI_wfCE_eval_e)
- then show ?case using eval_e.intros by metis
-qed
-
-lemma eval_c_exist:
- fixes \<Gamma>::\<Gamma> and c::c
- assumes "P ; \<Gamma> \<turnstile> i" and "P ; B ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f c"
- shows "\<exists>s. i \<lbrakk> c \<rbrakk> ~ s"
- using assms proof(nominal_induct c rule: c.strong_induct)
- case C_true
- then show ?case using eval_c.intros wfC_elims by metis
-next
- case C_false
- then show ?case using eval_c.intros wfC_elims by metis
-next
- case (C_conj c1 c2)
- then show ?case using eval_c.intros wfC_elims by metis
-next
- case (C_disj x1 x2)
- then show ?case using eval_c.intros wfC_elims by metis
-next
- case (C_not x)
- then show ?case using eval_c.intros wfC_elims by metis
-next
- case (C_imp x1 x2)
- then show ?case using eval_c.intros eval_e_exist wfC_elims by metis
-next
- case (C_eq x1 x2)
- then show ?case using eval_c.intros eval_e_exist wfC_elims by metis
-qed
-
-lemma eval_c_uniqueness:
- fixes c::c
- assumes "i \<lbrakk> c \<rbrakk> ~ s" and "i \<lbrakk> c \<rbrakk> ~ s'"
- shows "s=s'"
- using assms proof(nominal_induct c arbitrary: s s' rule:c.strong_induct)
- case C_true
- then show ?case using eval_c_elims by metis
-next
- case C_false
- then show ?case using eval_c_elims by metis
-next
- case (C_conj x1 x2)
- then show ?case using eval_c_elims(3) by (metis (full_types))
-next
- case (C_disj x1 x2)
- then show ?case using eval_c_elims(4) by (metis (full_types))
-next
- case (C_not x)
- then show ?case using eval_c_elims(6) by (metis (full_types))
-next
- case (C_imp x1 x2)
- then show ?case using eval_c_elims(5) by (metis (full_types))
-next
- case (C_eq x1 x2)
- then show ?case using eval_e_uniqueness eval_c_elims(7) by metis
-qed
-
-lemma wfI_wfC_eval_c:
- fixes c::c
- assumes "wfC P B G c" and "P ; G \<turnstile> i"
- shows "\<exists>s. i \<lbrakk> c \<rbrakk> ~ s"
- using assms proof(nominal_induct c rule: c.strong_induct)
-qed(metis wfC_elims wfI_wfCE_eval_e eval_c.intros)+
-
-section \<open>Satisfiability\<close>
-
-lemma satis_reflI:
- fixes c::c
- assumes "i \<Turnstile> ((x, b, c) #\<^sub>\<Gamma> G)"
- shows "i \<Turnstile> c"
- using assms by auto
-
-lemma is_satis_mp:
- fixes c1::c and c2::c
- assumes "i \<Turnstile> (c1 IMP c2)" and "i \<Turnstile> c1"
- shows "i \<Turnstile> c2"
- using assms proof -
- have "eval_c i (c1 IMP c2) True" using is_satis.simps using assms by blast
- then obtain b1 and b2 where "True = (b1 \<longrightarrow> b2) \<and> eval_c i c1 b1 \<and> eval_c i c2 b2"
- using eval_c_elims(5) by metis
- moreover have "eval_c i c1 True" using is_satis.simps using assms by blast
- moreover have "b1 = True" using calculation eval_c_uniqueness by blast
- ultimately have "eval_c i c2 True" by auto
- thus ?thesis using is_satis.intros by auto
-qed
-
-lemma is_satis_imp:
- fixes c1::c and c2::c
- assumes "i \<Turnstile> c1 \<longrightarrow> i \<Turnstile> c2" and "i \<lbrakk> c1 \<rbrakk> ~ b1" and "i \<lbrakk> c2 \<rbrakk> ~ b2"
- shows "i \<Turnstile> (c1 IMP c2)"
-proof(cases b1)
- case True
- hence "i \<Turnstile> c2" using assms is_satis.simps by simp
- hence "b2 = True" using is_satis.simps assms
- using eval_c_uniqueness by blast
- then show ?thesis using eval_c_impI is_satis.simps assms by force
-next
- case False
- then show ?thesis using assms eval_c_impI is_satis.simps by metis
-qed
-
-lemma is_satis_iff:
- "i \<Turnstile> G = (\<forall>x b c. (x,b,c) \<in> toSet G \<longrightarrow> i \<Turnstile> c)"
- by(induct G,auto)
-
-lemma is_satis_g_append:
- "i \<Turnstile> (G1@G2) = (i \<Turnstile> G1 \<and> i \<Turnstile> G2)"
- using is_satis_g.simps is_satis_iff by auto
-
-section \<open>Substitution for Evaluation\<close>
-
-lemma eval_v_i_upd:
- fixes v::v
- assumes "atom x \<sharp> v" and "i \<lbrakk> v \<rbrakk> ~ s'"
- shows "eval_v ((i ( x \<mapsto>s))) v s' "
- using assms proof(nominal_induct v arbitrary: s' rule:v.strong_induct)
- case (V_lit x)
- then show ?case by (metis eval_v_elims(1) eval_v_litI)
-next
- case (V_var y)
- then obtain s where *: "Some s = i y \<and> s=s'" using eval_v_elims by metis
- moreover have "x \<noteq> y" using \<open>atom x \<sharp> V_var y\<close> v.supp by simp
- ultimately have "(i ( x \<mapsto>s)) y = Some s"
- by (simp add: \<open>Some s = i y \<and> s = s'\<close>)
- then show ?case using eval_v_varI * \<open>x \<noteq> y\<close>
- by (simp add: eval_v.eval_v_varI)
-next
- case (V_pair v1 v2)
- hence "atom x \<sharp> v1 \<and> atom x \<sharp> v2" using v.supp by simp
- moreover obtain s1 and s2 where *: "eval_v i v1 s1 \<and> eval_v i v2 s2 \<and> s' = SPair s1 s2" using eval_v_elims V_pair by metis
- ultimately have "eval_v ((i ( x \<mapsto>s))) v1 s1 \<and> eval_v ((i ( x \<mapsto>s))) v2 s2" using V_pair by blast
- thus ?case using eval_v_pairI * by meson
-next
- case (V_cons tyid dc v1)
- hence "atom x \<sharp> v1" using v.supp by simp
- moreover obtain s1 where *: "eval_v i v1 s1 \<and> s' = SCons tyid dc s1" using eval_v_elims V_cons by metis
- ultimately have "eval_v ((i ( x \<mapsto>s))) v1 s1" using V_cons by blast
- thus ?case using eval_v_consI * by meson
-next
- case (V_consp tyid dc b1 v1)
-
- hence "atom x \<sharp> v1" using v.supp by simp
- moreover obtain s1 where *: "eval_v i v1 s1 \<and> s' = SConsp tyid dc b1 s1" using eval_v_elims V_consp by metis
- ultimately have "eval_v ((i ( x \<mapsto>s))) v1 s1" using V_consp by blast
- thus ?case using eval_v_conspI * by meson
-qed
-
-lemma eval_e_i_upd:
- fixes e::ce
- assumes "i \<lbrakk> e \<rbrakk> ~ s'" and "atom x \<sharp> e"
- shows " (i ( x \<mapsto>s)) \<lbrakk> e \<rbrakk> ~ s'"
- using assms apply(induct rule: eval_e.induct) using eval_v_i_upd eval_e_elims
- by (meson ce.fresh eval_e.intros)+
-
-lemma eval_c_i_upd:
- fixes c::c
- assumes "i \<lbrakk> c \<rbrakk> ~ s'" and "atom x \<sharp> c"
- shows "((i ( x \<mapsto>s))) \<lbrakk> c \<rbrakk> ~ s' "
- using assms proof(induct rule:eval_c.induct)
- case (eval_c_eqI i e1 sv1 e2 sv2)
- then show ?case using RCLogic.eval_c_eqI eval_e_i_upd c.fresh by metis
-qed(simp add: eval_c.intros)+
-
-lemma subst_v_eval_v[simp]:
- fixes v::v and v'::v
- assumes "i \<lbrakk> v \<rbrakk> ~ s" and "i \<lbrakk> (v'[x::=v]\<^sub>v\<^sub>v) \<rbrakk> ~ s'"
- shows "(i ( x \<mapsto> s )) \<lbrakk> v' \<rbrakk> ~ s'"
- using assms proof(nominal_induct v' arbitrary: s' rule:v.strong_induct)
- case (V_lit x)
- then show ?case using subst_vv.simps
- by (metis eval_v_elims(1) eval_v_litI)
-next
- case (V_var x')
- then show ?case proof(cases "x=x'")
- case True
- hence "(V_var x')[x::=v]\<^sub>v\<^sub>v = v" using subst_vv.simps by auto
- then show ?thesis using V_var eval_v_elims eval_v_varI eval_v_uniqueness True
- by (simp add: eval_v.eval_v_varI)
- next
- case False
- hence "atom x \<sharp> (V_var x')" by simp
- then show ?thesis using eval_v_i_upd False V_var by fastforce
- qed
-next
- case (V_pair v1 v2)
- then obtain s1 and s2 where *:"eval_v i (v1[x::=v]\<^sub>v\<^sub>v) s1 \<and> eval_v i (v2[x::=v]\<^sub>v\<^sub>v) s2 \<and> s' = SPair s1 s2" using V_pair eval_v_elims subst_vv.simps by metis
- hence "(i ( x \<mapsto> s )) \<lbrakk> v1 \<rbrakk> ~ s1 \<and> (i ( x \<mapsto> s )) \<lbrakk> v2 \<rbrakk> ~ s2" using V_pair by metis
- thus ?case using eval_v_pairI subst_vv.simps * V_pair by metis
-next
- case (V_cons tyid dc v1)
- then obtain s1 where "eval_v i (v1[x::=v]\<^sub>v\<^sub>v) s1" using eval_v_elims subst_vv.simps by metis
- thus ?case using eval_v_consI V_cons
- by (metis eval_v_elims subst_vv.simps)
-next
- case (V_consp tyid dc b1 v1)
-
- then obtain s1 where *:"eval_v i (v1[x::=v]\<^sub>v\<^sub>v) s1 \<and> s' = SConsp tyid dc b1 s1" using eval_v_elims subst_vv.simps by metis
- hence "i ( x \<mapsto> s ) \<lbrakk> v1 \<rbrakk> ~ s1" using V_consp by metis
- thus ?case using * eval_v_conspI by metis
-qed
-
-lemma subst_e_eval_v[simp]:
- fixes y::x and e::ce and v::v and e'::ce
- assumes "i \<lbrakk> e' \<rbrakk> ~ s'" and "e'=(e[y::=v]\<^sub>c\<^sub>e\<^sub>v)" and "i \<lbrakk> v \<rbrakk> ~ s"
- shows "(i ( y \<mapsto> s )) \<lbrakk> e \<rbrakk> ~ s'"
- using assms proof(induct arbitrary: e rule: eval_e.induct)
- case (eval_e_valI i v1 sv)
- then obtain v1' where *:"e = CE_val v1' \<and> v1 = v1'[y::=v]\<^sub>v\<^sub>v"
- using assms by(nominal_induct e rule:ce.strong_induct,simp+)
- hence "eval_v i (v1'[y::=v]\<^sub>v\<^sub>v) sv" using eval_e_valI by simp
- hence "eval_v (i ( y \<mapsto> s )) v1' sv" using subst_v_eval_v eval_e_valI by simp
- then show ?case using RCLogic.eval_e_valI * by meson
-next
- case (eval_e_plusI i v1 n1 v2 n2)
- then obtain v1' and v2' where *:"e = CE_op Plus v1' v2' \<and> v1 = v1'[y::=v]\<^sub>c\<^sub>e\<^sub>v \<and> v2 = v2'[y::=v]\<^sub>c\<^sub>e\<^sub>v"
- using assms by(nominal_induct e rule:ce.strong_induct,simp+)
- hence "eval_e i (v1'[y::=v]\<^sub>c\<^sub>e\<^sub>v) (SNum n1) \<and> eval_e i (v2'[y::=v]\<^sub>c\<^sub>e\<^sub>v) (SNum n2)" using eval_e_plusI by simp
- hence "eval_e (i ( y \<mapsto> s )) v1' (SNum n1) \<and> eval_e (i ( y \<mapsto> s )) v2' (SNum n2)" using subst_v_eval_v eval_e_plusI
- using "local.*" by blast
- then show ?case using RCLogic.eval_e_plusI * by meson
-next
- case (eval_e_leqI i v1 n1 v2 n2)
- then obtain v1' and v2' where *:"e = CE_op LEq v1' v2' \<and> v1 = v1'[y::=v]\<^sub>c\<^sub>e\<^sub>v \<and> v2 = v2'[y::=v]\<^sub>c\<^sub>e\<^sub>v"
- using assms by(nominal_induct e rule:ce.strong_induct,simp+)
- hence "eval_e i (v1'[y::=v]\<^sub>c\<^sub>e\<^sub>v) (SNum n1) \<and> eval_e i (v2'[y::=v]\<^sub>c\<^sub>e\<^sub>v) (SNum n2)" using eval_e_leqI by simp
- hence "eval_e (i ( y \<mapsto> s )) v1' (SNum n1) \<and> eval_e (i ( y \<mapsto> s )) v2' (SNum n2)" using subst_v_eval_v eval_e_leqI
- using * by blast
- then show ?case using RCLogic.eval_e_leqI * by meson
-next
- case (eval_e_eqI i v1 n1 v2 n2)
- then obtain v1' and v2' where *:"e = CE_op Eq v1' v2' \<and> v1 = v1'[y::=v]\<^sub>c\<^sub>e\<^sub>v \<and> v2 = v2'[y::=v]\<^sub>c\<^sub>e\<^sub>v"
- using assms by(nominal_induct e rule:ce.strong_induct,simp+)
- hence "eval_e i (v1'[y::=v]\<^sub>c\<^sub>e\<^sub>v) n1 \<and> eval_e i (v2'[y::=v]\<^sub>c\<^sub>e\<^sub>v) n2" using eval_e_eqI by simp
- hence "eval_e (i ( y \<mapsto> s )) v1' n1 \<and> eval_e (i ( y \<mapsto> s )) v2' n2" using subst_v_eval_v eval_e_eqI
- using * by blast
- then show ?case using RCLogic.eval_e_eqI * by meson
-next
- case (eval_e_fstI i v1 s1 s2)
- then obtain v1' and v2' where *:"e = CE_fst v1' \<and> v1 = v1'[y::=v]\<^sub>c\<^sub>e\<^sub>v"
- using assms by(nominal_induct e rule:ce.strong_induct,simp+)
- hence "eval_e i (v1'[y::=v]\<^sub>c\<^sub>e\<^sub>v) (SPair s1 s2)" using eval_e_fstI by simp
- hence "eval_e (i ( y \<mapsto> s )) v1' (SPair s1 s2)" using eval_e_fstI * by metis
- then show ?case using RCLogic.eval_e_fstI * by meson
-next
- case (eval_e_sndI i v1 s1 s2)
- then obtain v1' and v2' where *:"e = CE_snd v1' \<and> v1 = v1'[y::=v]\<^sub>c\<^sub>e\<^sub>v"
- using assms by(nominal_induct e rule:ce.strong_induct,simp+)
- hence "eval_e i (v1'[y::=v]\<^sub>c\<^sub>e\<^sub>v) (SPair s1 s2)" using eval_e_sndI by simp
- hence "eval_e (i ( y \<mapsto> s )) v1' (SPair s1 s2)" using subst_v_eval_v eval_e_sndI * by blast
- then show ?case using RCLogic.eval_e_sndI * by meson
-next
- case (eval_e_concatI i v1 bv1 v2 bv2)
- then obtain v1' and v2' where *:"e = CE_concat v1' v2' \<and> v1 = v1'[y::=v]\<^sub>c\<^sub>e\<^sub>v \<and> v2 = v2'[y::=v]\<^sub>c\<^sub>e\<^sub>v"
- using assms by(nominal_induct e rule:ce.strong_induct,simp+)
- hence "eval_e i (v1'[y::=v]\<^sub>c\<^sub>e\<^sub>v) (SBitvec bv1) \<and> eval_e i (v2'[y::=v]\<^sub>c\<^sub>e\<^sub>v) (SBitvec bv2)" using eval_e_concatI by simp
- moreover hence "eval_e (i ( y \<mapsto> s )) v1' (SBitvec bv1) \<and> eval_e (i ( y \<mapsto> s )) v2' (SBitvec bv2)"
- using subst_v_eval_v eval_e_concatI * by blast
- ultimately show ?case using RCLogic.eval_e_concatI * eval_v_uniqueness by (metis eval_e_concatI.hyps(1))
-next
- case (eval_e_lenI i v1 bv)
- then obtain v1' where *:"e = CE_len v1' \<and> v1 = v1'[y::=v]\<^sub>c\<^sub>e\<^sub>v"
- using assms by(nominal_induct e rule:ce.strong_induct,simp+)
- hence "eval_e i (v1'[y::=v]\<^sub>c\<^sub>e\<^sub>v) (SBitvec bv)" using eval_e_lenI by simp
- hence "eval_e (i ( y \<mapsto> s )) v1' (SBitvec bv)" using subst_v_eval_v eval_e_lenI * by blast
- then show ?case using RCLogic.eval_e_lenI * by meson
-qed
-
-lemma subst_c_eval_v[simp]:
- fixes v::v and c :: c
- assumes "i \<lbrakk> v \<rbrakk> ~ s" and "i \<lbrakk> c[x::=v]\<^sub>c\<^sub>v \<rbrakk> ~ s1" and
- "(i ( x \<mapsto> s)) \<lbrakk> c \<rbrakk> ~ s2"
- shows "s1 = s2"
- using assms proof(nominal_induct c arbitrary: s1 s2 rule: c.strong_induct)
- case C_true
- hence "s1 = True \<and> s2 = True" using eval_c_elims subst_cv.simps by auto
- then show ?case by auto
-next
- case C_false
- hence "s1 = False \<and> s2 = False" using eval_c_elims subst_cv.simps by metis
- then show ?case by auto
-next
- case (C_conj c1 c2)
- hence *:"eval_c i (c1[x::=v]\<^sub>c\<^sub>v AND c2[x::=v]\<^sub>c\<^sub>v) s1" using subst_cv.simps by auto
- then obtain s11 and s12 where "(s1 = (s11 \<and> s12)) \<and> eval_c i c1[x::=v]\<^sub>c\<^sub>v s11 \<and> eval_c i c2[x::=v]\<^sub>c\<^sub>v s12" using
- eval_c_elims(3) by metis
- moreover obtain s21 and s22 where "eval_c (i ( x \<mapsto> s)) c1 s21 \<and> eval_c (i ( x \<mapsto> s)) c2 s22 \<and> (s2 = (s21 \<and> s22))" using
- eval_c_elims(3) C_conj by metis
- ultimately show ?case using C_conj by (meson eval_c_elims)
-next
- case (C_disj c1 c2)
- hence *:"eval_c i (c1[x::=v]\<^sub>c\<^sub>v OR c2[x::=v]\<^sub>c\<^sub>v) s1" using subst_cv.simps by auto
- then obtain s11 and s12 where "(s1 = (s11 \<or> s12)) \<and> eval_c i c1[x::=v]\<^sub>c\<^sub>v s11 \<and> eval_c i c2[x::=v]\<^sub>c\<^sub>v s12" using
- eval_c_elims(4) by metis
- moreover obtain s21 and s22 where "eval_c (i ( x \<mapsto> s)) c1 s21 \<and> eval_c (i ( x \<mapsto> s)) c2 s22 \<and> (s2 = (s21 \<or> s22))" using
- eval_c_elims(4) C_disj by metis
- ultimately show ?case using C_disj by (meson eval_c_elims)
-next
- case (C_not c1)
- then obtain s11 where "(s1 = (\<not> s11)) \<and> eval_c i c1[x::=v]\<^sub>c\<^sub>v s11" using
- eval_c_elims(6) by (metis subst_cv.simps(7))
- moreover obtain s21 where "eval_c (i ( x \<mapsto> s)) c1 s21 \<and> (s2 = (\<not>s21))" using
- eval_c_elims(6) C_not by metis
- ultimately show ?case using C_not by (meson eval_c_elims)
-next
- case (C_imp c1 c2)
- hence *:"eval_c i (c1[x::=v]\<^sub>c\<^sub>v IMP c2[x::=v]\<^sub>c\<^sub>v) s1" using subst_cv.simps by auto
- then obtain s11 and s12 where "(s1 = (s11 \<longrightarrow> s12)) \<and> eval_c i c1[x::=v]\<^sub>c\<^sub>v s11 \<and> eval_c i c2[x::=v]\<^sub>c\<^sub>v s12" using
- eval_c_elims(5) by metis
- moreover obtain s21 and s22 where "eval_c (i ( x \<mapsto> s)) c1 s21 \<and> eval_c (i ( x \<mapsto> s)) c2 s22 \<and> (s2 = (s21 \<longrightarrow> s22))" using
- eval_c_elims(5) C_imp by metis
- ultimately show ?case using C_imp by (meson eval_c_elims)
-next
- case (C_eq e1 e2)
- hence *:"eval_c i (e1[x::=v]\<^sub>c\<^sub>e\<^sub>v == e2[x::=v]\<^sub>c\<^sub>e\<^sub>v) s1" using subst_cv.simps by auto
- then obtain s11 and s12 where "(s1 = (s11 = s12)) \<and> eval_e i (e1[x::=v]\<^sub>c\<^sub>e\<^sub>v) s11 \<and> eval_e i (e2[x::=v]\<^sub>c\<^sub>e\<^sub>v) s12" using
- eval_c_elims(7) by metis
- moreover obtain s21 and s22 where "eval_e (i ( x \<mapsto> s)) e1 s21 \<and> eval_e (i ( x \<mapsto> s)) e2 s22 \<and> (s2 = (s21 = s22))" using
- eval_c_elims(7) C_eq by metis
- ultimately show ?case using C_eq subst_e_eval_v by (metis eval_e_uniqueness)
-qed
-
-lemma wfI_upd:
- assumes "wfI \<Theta> \<Gamma> i" and "wfRCV \<Theta> s b" and "wfG \<Theta> B ((x, b, c) #\<^sub>\<Gamma> \<Gamma>)"
- shows "wfI \<Theta> ((x, b, c) #\<^sub>\<Gamma> \<Gamma>) (i(x \<mapsto> s))"
-proof(subst wfI_def,rule)
- fix xa
- assume as:"xa \<in> toSet ((x, b, c) #\<^sub>\<Gamma> \<Gamma>)"
-
- then obtain x1::x and b1::b and c1::c where xa: "xa = (x1,b1,c1)" using toSet.simps
- using prod_cases3 by blast
-
- have "\<exists>sa. Some sa = (i(x \<mapsto> s)) x1 \<and> wfRCV \<Theta> sa b1" proof(cases "x=x1")
- case True
- hence "b=b1" using as xa wfG_unique assms by metis
- hence "Some s = (i(x \<mapsto> s)) x1 \<and> wfRCV \<Theta> s b1" using assms True by simp
- then show ?thesis by auto
- next
- case False
- hence "(x1,b1,c1) \<in> toSet \<Gamma>" using xa as by auto
- then obtain sa where "Some sa = i x1 \<and> wfRCV \<Theta> sa b1" using assms wfI_def as xa by auto
- hence "Some sa = (i(x \<mapsto> s)) x1 \<and> wfRCV \<Theta> sa b1" using False by auto
- then show ?thesis by auto
- qed
-
- thus "case xa of (xa, ba, ca) \<Rightarrow> \<exists>sa. Some sa = (i(x \<mapsto> s)) xa \<and> wfRCV \<Theta> sa ba" using xa by auto
-qed
-
-lemma wfI_upd_full:
- fixes v::v
- assumes "wfI \<Theta> G i" and "G = ((\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v)@\<Gamma>)" and "wfRCV \<Theta> s b" and "wfG \<Theta> B (\<Gamma>'@((x,b,c)#\<^sub>\<Gamma>\<Gamma>))" and "\<Theta> ; B ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b"
- shows "wfI \<Theta> (\<Gamma>'@((x,b,c)#\<^sub>\<Gamma>\<Gamma>)) (i(x \<mapsto> s))"
-proof(subst wfI_def,rule)
- fix xa
- assume as:"xa \<in> toSet (\<Gamma>'@((x,b,c)#\<^sub>\<Gamma>\<Gamma>))"
-
- then obtain x1::x and b1::b and c1::c where xa: "xa = (x1,b1,c1)" using toSet.simps
- using prod_cases3 by blast
-
- have "\<exists>sa. Some sa = (i(x \<mapsto> s)) x1 \<and> wfRCV \<Theta> sa b1"
- proof(cases "x=x1")
- case True
- hence "b=b1" using as xa wfG_unique_full assms by metis
- hence "Some s = (i(x \<mapsto> s)) x1 \<and> wfRCV \<Theta> s b1" using assms True by simp
- then show ?thesis by auto
- next
- case False
- hence "(x1,b1,c1) \<in> toSet (\<Gamma>'@\<Gamma>)" using as xa by auto
- then obtain c1' where "(x1,b1,c1') \<in> toSet (\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v@\<Gamma>)" using xa as wfG_member_subst assms False by metis
- then obtain sa where "Some sa = i x1 \<and> wfRCV \<Theta> sa b1" using assms wfI_def as xa by blast
- hence "Some sa = (i(x \<mapsto> s)) x1 \<and> wfRCV \<Theta> sa b1" using False by auto
- then show ?thesis by auto
- qed
-
- thus "case xa of (xa, ba, ca) \<Rightarrow> \<exists>sa. Some sa = (i(x \<mapsto> s)) xa \<and> wfRCV \<Theta> sa ba" using xa by auto
-qed
-
-lemma subst_c_satis[simp]:
- fixes v::v
- assumes "i \<lbrakk> v \<rbrakk> ~ s" and "wfC \<Theta> B ((x,b,c')#\<^sub>\<Gamma>\<Gamma>) c" and "wfI \<Theta> \<Gamma> i" and "\<Theta> ; B ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b"
- shows "i \<Turnstile> (c[x::=v]\<^sub>c\<^sub>v) \<longleftrightarrow> (i ( x \<mapsto> s)) \<Turnstile> c"
-proof -
- have "wfI \<Theta> ((x, b, c') #\<^sub>\<Gamma> \<Gamma>) (i(x \<mapsto> s))" using wfI_upd assms wfC_wf eval_v_base by blast
- then obtain s1 where s1:"eval_c (i(x \<mapsto> s)) c s1" using eval_c_exist[of \<Theta> "((x,b,c')#\<^sub>\<Gamma>\<Gamma>)" "(i ( x \<mapsto> s))" B c ] assms by auto
-
- have "\<Theta> ; B ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f c[x::=v]\<^sub>c\<^sub>v" using wf_subst1(2)[OF assms(2) _ assms(4) , of GNil x ] subst_gv.simps by simp
- then obtain s2 where s2:"eval_c i c[x::=v]\<^sub>c\<^sub>v s2" using eval_c_exist[of \<Theta> "\<Gamma>" "i" B "c[x::=v]\<^sub>c\<^sub>v" ] assms by auto
-
- show ?thesis using s1 s2 subst_c_eval_v[OF assms(1) s2 s1] is_satis.cases
- using eval_c_uniqueness is_satis.simps by auto
-qed
-
-text \<open> Key theorem telling us we can replace a substitution with an update to the valuation \<close>
-lemma subst_c_satis_full:
- fixes v::v and \<Gamma>'::\<Gamma>
- assumes "i \<lbrakk> v \<rbrakk> ~ s" and "wfC \<Theta> B (\<Gamma>'@((x,b,c')#\<^sub>\<Gamma>\<Gamma>)) c" and "wfI \<Theta> ((\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v)@\<Gamma>) i" and "\<Theta> ; B ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b"
- shows "i \<Turnstile> (c[x::=v]\<^sub>c\<^sub>v) \<longleftrightarrow> (i ( x \<mapsto> s)) \<Turnstile> c"
-proof -
- have "wfI \<Theta> (\<Gamma>'@((x, b, c')) #\<^sub>\<Gamma> \<Gamma>) (i(x \<mapsto> s))" using wfI_upd_full assms wfC_wf eval_v_base wfI_suffix wfI_def wfV_wf by fast
- then obtain s1 where s1:"eval_c (i(x \<mapsto> s)) c s1" using eval_c_exist[of \<Theta> "(\<Gamma>'@(x,b,c')#\<^sub>\<Gamma>\<Gamma>)" "(i ( x \<mapsto> s))" B c ] assms by auto
-
- have "\<Theta> ; B ; ((\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v)@\<Gamma>) \<turnstile>\<^sub>w\<^sub>f c[x::=v]\<^sub>c\<^sub>v" using wbc_subst assms by auto
-
- then obtain s2 where s2:"eval_c i c[x::=v]\<^sub>c\<^sub>v s2" using eval_c_exist[of \<Theta> "((\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v)@\<Gamma>)" "i" B "c[x::=v]\<^sub>c\<^sub>v" ] assms by auto
-
- show ?thesis using s1 s2 subst_c_eval_v[OF assms(1) s2 s1] is_satis.cases
- using eval_c_uniqueness is_satis.simps by auto
-qed
-
-section \<open>Validity\<close>
-
-lemma validI[intro]:
- fixes c::c
- assumes "wfC P B G c" and "\<forall>i. P ; G \<turnstile> i \<and> i \<Turnstile> G \<longrightarrow> i \<Turnstile> c"
- shows "P ; B ; G \<Turnstile> c"
- using assms valid.simps by presburger
-
-lemma valid_g_wf:
- fixes c::c and G::\<Gamma>
- assumes "P ; B ; G \<Turnstile> c"
- shows "P ; B \<turnstile>\<^sub>w\<^sub>f G"
- using assms wfC_wf valid.simps by blast
-
-lemma valid_reflI [intro]:
- fixes b::b
- assumes "P ; B ; ((x,b,c1)#\<^sub>\<Gamma>G) \<turnstile>\<^sub>w\<^sub>f c1" and "c1 = c2"
- shows "P ; B ; ((x,b,c1)#\<^sub>\<Gamma>G) \<Turnstile> c2"
- using satis_reflI assms by simp
-
-subsection \<open>Weakening and Strengthening\<close>
-
-text \<open> Adding to the domain of a valuation doesn't change the result \<close>
-
-lemma eval_v_weakening:
- fixes c::v and B::"bv fset"
- assumes "i = i'|` d" and "supp c \<subseteq> atom ` d \<union> supp B " and "i \<lbrakk> c \<rbrakk> ~ s"
- shows "i' \<lbrakk> c \<rbrakk> ~ s"
- using assms proof(nominal_induct c arbitrary:s rule: v.strong_induct)
- case (V_lit x)
- then show ?case using eval_v_elims eval_v_litI by metis
-next
- case (V_var x)
- have "atom x \<in> atom ` d" using x_not_in_b_set[of x B] assms v.supp(2) supp_at_base
- proof -
- show ?thesis
- by (metis UnE V_var.prems(2) \<open>atom x \<notin> supp B\<close> singletonI subset_iff supp_at_base v.supp(2)) (* 46 ms *)
- qed
- moreover have "Some s = i x" using assms eval_v_elims(2)
- using V_var.prems(3) by blast
- hence "Some s= i' x" using assms insert_subset restrict_in
- proof -
- show ?thesis
- by (metis (no_types) \<open>i = i' |` d\<close> \<open>Some s = i x\<close> atom_eq_iff calculation imageE restrict_in) (* 31 ms *)
- qed
- thus ?case using eval_v.eval_v_varI by simp
-
-next
- case (V_pair v1 v2)
- then show ?case using eval_v_elims(3) eval_v_pairI v.supp
- by (metis assms le_sup_iff)
-next
- case (V_cons dc v1)
- then show ?case using eval_v_elims(4) eval_v_consI v.supp
- by (metis assms le_sup_iff)
-next
- case (V_consp tyid dc b1 v1)
-
- then obtain sv1 where *:"i \<lbrakk> v1 \<rbrakk> ~ sv1 \<and> s = SConsp tyid dc b1 sv1" using eval_v_elims by metis
- hence "i' \<lbrakk> v1 \<rbrakk> ~ sv1" using V_consp by auto
- then show ?case using * eval_v_conspI v.supp eval_v.simps assms le_sup_iff by metis
-qed
-
-lemma eval_v_restrict:
- fixes c::v and B::"bv fset"
- assumes "i = i' |` d" and "supp c \<subseteq> atom ` d \<union> supp B " and "i' \<lbrakk> c \<rbrakk> ~ s"
- shows "i \<lbrakk> c \<rbrakk> ~ s"
- using assms proof(nominal_induct c arbitrary:s rule: v.strong_induct)
- case (V_lit x)
- then show ?case using eval_v_elims eval_v_litI by metis
-next
- case (V_var x)
- have "atom x \<in> atom ` d" using x_not_in_b_set[of x B] assms v.supp(2) supp_at_base
- proof -
- show ?thesis
- by (metis UnE V_var.prems(2) \<open>atom x \<notin> supp B\<close> singletonI subset_iff supp_at_base v.supp(2)) (* 46 ms *)
- qed
- moreover have "Some s = i' x" using assms eval_v_elims(2)
- using V_var.prems(3) by blast
- hence "Some s= i x" using assms insert_subset restrict_in
- proof -
- show ?thesis
- by (metis (no_types) \<open>i = i' |` d\<close> \<open>Some s = i' x\<close> atom_eq_iff calculation imageE restrict_in) (* 31 ms *)
- qed
- thus ?case using eval_v.eval_v_varI by simp
-next
- case (V_pair v1 v2)
- then show ?case using eval_v_elims(3) eval_v_pairI v.supp
- by (metis assms le_sup_iff)
-next
- case (V_cons dc v1)
- then show ?case using eval_v_elims(4) eval_v_consI v.supp
- by (metis assms le_sup_iff)
-next
- case (V_consp tyid dc b1 v1)
- then obtain sv1 where *:"i' \<lbrakk> v1 \<rbrakk> ~ sv1 \<and> s = SConsp tyid dc b1 sv1" using eval_v_elims by metis
- hence "i \<lbrakk> v1 \<rbrakk> ~ sv1" using V_consp by auto
- then show ?case using * eval_v_conspI v.supp eval_v.simps assms le_sup_iff by metis
-qed
-
-lemma eval_e_weakening:
- fixes e::ce and B::"bv fset"
- assumes "i \<lbrakk> e \<rbrakk> ~ s" and "i = i' |` d" and "supp e \<subseteq> atom ` d \<union> supp B "
- shows "i' \<lbrakk> e \<rbrakk> ~ s"
- using assms proof(induct rule: eval_e.induct)
- case (eval_e_valI i v sv)
- then show ?case using ce.supp eval_e.intros
- using eval_v_weakening by auto
-next
- case (eval_e_plusI i v1 n1 v2 n2)
- then show ?case using ce.supp eval_e.intros
- using eval_v_weakening by auto
-next
- case (eval_e_leqI i v1 n1 v2 n2)
- then show ?case using ce.supp eval_e.intros
- using eval_v_weakening by auto
-next
- case (eval_e_eqI i v1 n1 v2 n2)
- then show ?case using ce.supp eval_e.intros
- using eval_v_weakening by auto
-next
- case (eval_e_fstI i v v1 v2)
- then show ?case using ce.supp eval_e.intros
- using eval_v_weakening by metis
-next
- case (eval_e_sndI i v v1 v2)
- then show ?case using ce.supp eval_e.intros
- using eval_v_weakening by metis
-next
- case (eval_e_concatI i v1 bv2 v2 bv1)
- then show ?case using ce.supp eval_e.intros
- using eval_v_weakening by auto
-next
- case (eval_e_lenI i v bv)
- then show ?case using ce.supp eval_e.intros
- using eval_v_weakening by auto
-qed
-
-lemma eval_e_restrict :
- fixes e::ce and B::"bv fset"
- assumes "i' \<lbrakk> e \<rbrakk> ~ s" and "i = i' |` d" and "supp e \<subseteq> atom ` d \<union> supp B "
- shows "i \<lbrakk> e \<rbrakk> ~ s"
- using assms proof(induct rule: eval_e.induct)
- case (eval_e_valI i v sv)
- then show ?case using ce.supp eval_e.intros
- using eval_v_restrict by auto
-next
- case (eval_e_plusI i v1 n1 v2 n2)
- then show ?case using ce.supp eval_e.intros
- using eval_v_restrict by auto
-next
- case (eval_e_leqI i v1 n1 v2 n2)
- then show ?case using ce.supp eval_e.intros
- using eval_v_restrict by auto
-next
- case (eval_e_eqI i v1 n1 v2 n2)
- then show ?case using ce.supp eval_e.intros
- using eval_v_restrict by auto
-next
- case (eval_e_fstI i v v1 v2)
- then show ?case using ce.supp eval_e.intros
- using eval_v_restrict by metis
-next
- case (eval_e_sndI i v v1 v2)
- then show ?case using ce.supp eval_e.intros
- using eval_v_restrict by metis
-next
- case (eval_e_concatI i v1 bv2 v2 bv1)
- then show ?case using ce.supp eval_e.intros
- using eval_v_restrict by auto
-next
- case (eval_e_lenI i v bv)
- then show ?case using ce.supp eval_e.intros
- using eval_v_restrict by auto
-qed
-
-lemma eval_c_i_weakening:
- fixes c::c and B::"bv fset"
- assumes "i \<lbrakk> c \<rbrakk> ~ s" and "i = i' |` d" and "supp c \<subseteq> atom ` d \<union> supp B"
- shows "i' \<lbrakk> c \<rbrakk> ~ s"
- using assms proof(induct rule:eval_c.induct)
- case (eval_c_eqI i e1 sv1 e2 sv2)
- then show ?case using eval_c.intros eval_e_weakening by auto
-qed(auto simp add: eval_c.intros)+
-
-lemma eval_c_i_restrict:
- fixes c::c and B::"bv fset"
- assumes "i' \<lbrakk> c \<rbrakk> ~ s" and "i = i' |` d" and "supp c \<subseteq> atom ` d \<union> supp B"
- shows "i \<lbrakk> c \<rbrakk> ~ s"
- using assms proof(induct rule:eval_c.induct)
- case (eval_c_eqI i e1 sv1 e2 sv2)
- then show ?case using eval_c.intros eval_e_restrict by auto
-qed(auto simp add: eval_c.intros)+
-
-lemma is_satis_i_weakening:
- fixes c::c and B::"bv fset"
- assumes "i = i' |` d" and "supp c \<subseteq> atom ` d \<union> supp B " and "i \<Turnstile> c"
- shows "i' \<Turnstile> c"
- using is_satis.simps eval_c_i_weakening[OF _ assms(1) assms(2) ]
- using assms(3) by auto
-
-lemma is_satis_i_restrict:
- fixes c::c and B::"bv fset"
- assumes "i = i' |` d" and "supp c \<subseteq> atom ` d \<union> supp B" and "i' \<Turnstile> c"
- shows "i \<Turnstile> c"
- using is_satis.simps eval_c_i_restrict[OF _ assms(1) assms(2) ]
- using assms(3) by auto
-
-lemma is_satis_g_restrict1:
- fixes \<Gamma>'::\<Gamma> and \<Gamma>::\<Gamma>
- assumes "toSet \<Gamma> \<subseteq> toSet \<Gamma>'" and "i \<Turnstile> \<Gamma>'"
- shows "i \<Turnstile> \<Gamma>"
- using assms proof(induct \<Gamma> rule: \<Gamma>.induct)
- case GNil
- then show ?case by auto
-next
- case (GCons xbc G)
- obtain x and b and c::c where xbc: "xbc=(x,b,c)"
- using prod_cases3 by blast
- hence "i \<Turnstile> G" using GCons by auto
- moreover have "i \<Turnstile> c" using GCons
- is_satis_iff toSet.simps subset_iff
- using xbc by blast
- ultimately show ?case using is_satis_g.simps GCons
- by (simp add: xbc)
-qed
-
-lemma is_satis_g_restrict2:
- fixes \<Gamma>'::\<Gamma> and \<Gamma>::\<Gamma>
- assumes "i \<Turnstile> \<Gamma>" and "i' = i |` d" and "atom_dom \<Gamma> \<subseteq> atom ` d" and "\<Theta> ; B \<turnstile>\<^sub>w\<^sub>f \<Gamma>"
- shows "i' \<Turnstile> \<Gamma>"
- using assms proof(induct \<Gamma> rule: \<Gamma>_induct)
- case GNil
- then show ?case by auto
-next
- case (GCons x b c G)
-
- hence "i' \<Turnstile> G" proof -
- have "i \<Turnstile> G" using GCons by simp
- moreover have "atom_dom G \<subseteq> atom ` d" using GCons by simp
- ultimately show ?thesis using GCons wfG_cons2 by blast
- qed
-
- moreover have "i' \<Turnstile> c" proof -
- have "i \<Turnstile> c" using GCons by auto
- moreover have "\<Theta> ; B ; (x, b, TRUE) #\<^sub>\<Gamma> G \<turnstile>\<^sub>w\<^sub>f c" using wfG_wfC GCons by simp
- moreover hence "supp c \<subseteq> atom ` d \<union> supp B" using wfC_supp GCons atom_dom_eq by blast
- ultimately show ?thesis using is_satis_i_restrict[of i' i d c] GCons by simp
- qed
-
- ultimately show ?case by auto
-qed
-
-lemma is_satis_g_restrict:
- fixes \<Gamma>'::\<Gamma> and \<Gamma>::\<Gamma>
- assumes "toSet \<Gamma> \<subseteq> toSet \<Gamma>'" and "i' \<Turnstile> \<Gamma>'" and "i = i' |` (fst ` toSet \<Gamma>)" and "\<Theta> ; B \<turnstile>\<^sub>w\<^sub>f \<Gamma>"
- shows "i \<Turnstile> \<Gamma>"
- using assms is_satis_g_restrict1[OF assms(1) assms(2)] is_satis_g_restrict2[OF _ assms(3)] by simp
-
-subsection \<open>Updating valuation\<close>
-
-lemma is_satis_c_i_upd:
- fixes c::c
- assumes "atom x \<sharp> c" and "i \<Turnstile> c"
- shows "((i ( x \<mapsto>s))) \<Turnstile> c"
- using assms eval_c_i_upd is_satis.simps by simp
-
-lemma is_satis_g_i_upd:
- fixes G::\<Gamma>
- assumes "atom x \<sharp> G" and "i \<Turnstile> G"
- shows "((i ( x \<mapsto>s))) \<Turnstile> G"
- using assms proof(induct G rule: \<Gamma>_induct)
- case GNil
- then show ?case by auto
-next
- case (GCons x' b' c' G')
-
- hence *:"atom x \<sharp> G' \<and> atom x \<sharp> c'"
- using fresh_def fresh_GCons GCons by force
- moreover hence "is_satis ((i ( x \<mapsto>s))) c'"
- using is_satis_c_i_upd GCons is_satis_g.simps by auto
- moreover have " is_satis_g (i(x \<mapsto> s)) G'" using GCons * by fastforce
- ultimately show ?case
- using GCons is_satis_g.simps(2) by metis
-qed
-
-lemma valid_weakening:
- assumes "\<Theta> ; B ; \<Gamma> \<Turnstile> c" and "toSet \<Gamma> \<subseteq> toSet \<Gamma>'" and "wfG \<Theta> B \<Gamma>'"
- shows "\<Theta> ; B ; \<Gamma>' \<Turnstile> c"
-proof -
- have "wfC \<Theta> B \<Gamma> c" using assms valid.simps by auto
- hence sp: "supp c \<subseteq> atom `(fst `toSet \<Gamma>) \<union> supp B" using wfX_wfY wfG_elims
- using atom_dom.simps dom.simps wf_supp(2) by metis
- have wfg: "wfG \<Theta> B \<Gamma>" using assms valid.simps wfC_wf by auto
-
- moreover have a1: "(\<forall>i. wfI \<Theta> \<Gamma>' i \<and> i \<Turnstile> \<Gamma>' \<longrightarrow> i \<Turnstile> c)" proof(rule allI, rule impI)
- fix i
- assume as: "wfI \<Theta> \<Gamma>' i \<and> i \<Turnstile> \<Gamma>'"
- hence as1: "fst ` toSet \<Gamma> \<subseteq> dom i" using assms wfI_domi by blast
- obtain i' where idash: "i' = restrict_map i (fst `toSet \<Gamma>)" by blast
- hence as2: "dom i' = (fst `toSet \<Gamma>)" using dom_restrict as1 by auto
-
- have id2: "\<Theta> ; \<Gamma> \<turnstile> i' \<and> i' \<Turnstile> \<Gamma>" proof -
- have "wfI \<Theta> \<Gamma> i'" using as2 wfI_restrict_weakening[of \<Theta> \<Gamma>' i i' \<Gamma>] as assms
- using idash by blast
- moreover have "i' \<Turnstile> \<Gamma>" using is_satis_g_restrict[OF assms(2)] wfg as idash by auto
- ultimately show ?thesis using idash by auto
- qed
- hence "i' \<Turnstile> c" using assms valid.simps by auto
- thus "i \<Turnstile> c" using assms valid.simps is_satis_i_weakening idash sp by blast
- qed
- moreover have "wfC \<Theta> B \<Gamma>' c" using wf_weakening assms valid.simps
- by (meson wfg)
- ultimately show ?thesis using assms valid.simps by auto
-qed
-
-lemma is_satis_g_suffix:
- fixes G::\<Gamma>
- assumes "i \<Turnstile> (G'@G)"
- shows "i \<Turnstile> G"
- using assms proof(induct G' rule:\<Gamma>.induct)
- case GNil
- then show ?case by auto
-next
- case (GCons xbc x2)
- obtain x and b and c::c where xbc: "xbc=(x,b,c)"
- using prod_cases3 by blast
- hence " i \<Turnstile> (xbc #\<^sub>\<Gamma> (x2 @ G))" using append_g.simps GCons by fastforce
- then show ?case using is_satis_g.simps GCons xbc by blast
-qed
-
-lemma wfG_inside_valid2:
- fixes x::x and \<Gamma>::\<Gamma> and c0::c and c0'::c
- assumes "wfG \<Theta> B (\<Gamma>'@((x,b0,c0')#\<^sub>\<Gamma>\<Gamma>))" and
- "\<Theta> ; B ; \<Gamma>'@(x,b0,c0)#\<^sub>\<Gamma>\<Gamma> \<Turnstile> c0'"
- shows "wfG \<Theta> B (\<Gamma>'@((x,b0,c0)#\<^sub>\<Gamma>\<Gamma>))"
-proof -
- have "wfG \<Theta> B (\<Gamma>'@(x,b0,c0)#\<^sub>\<Gamma>\<Gamma>)" using valid.simps wfC_wf assms by auto
- thus ?thesis using wfG_replace_inside_full assms by auto
-qed
-
-lemma valid_trans:
- assumes " \<Theta> ; \<B> ; \<Gamma> \<Turnstile> c0[z::=v]\<^sub>v" and " \<Theta> ; \<B> ; (z,b,c0)#\<^sub>\<Gamma>\<Gamma> \<Turnstile> c1" and "atom z \<sharp> \<Gamma>" and "wfV \<Theta> \<B> \<Gamma> v b"
- shows " \<Theta> ; \<B> ; \<Gamma> \<Turnstile> c1[z::=v]\<^sub>v"
-proof -
- have *:"wfC \<Theta> \<B> ((z,b,c0)#\<^sub>\<Gamma>\<Gamma>) c1" using valid.simps assms by auto
- hence "wfC \<Theta> \<B> \<Gamma> (c1[z::=v]\<^sub>v)" using wf_subst1(2)[OF * , of GNil ] assms subst_gv.simps subst_v_c_def by force
-
- moreover have "\<forall>i. wfI \<Theta> \<Gamma> i \<and> is_satis_g i \<Gamma> \<longrightarrow> is_satis i (c1[z::=v]\<^sub>v)"
- proof(rule,rule)
- fix i
- assume as: "wfI \<Theta> \<Gamma> i \<and> is_satis_g i \<Gamma>"
- then obtain sv where sv: "eval_v i v sv \<and> wfRCV \<Theta> sv b" using eval_v_exist assms by metis
- hence "is_satis i (c0[z::=v]\<^sub>v)" using assms valid.simps as by metis
- hence "is_satis (i(z \<mapsto> sv)) c0" using subst_c_satis sv as assms valid.simps wfC_wf wfG_elim2 subst_v_c_def by metis
- moreover have "is_satis_g (i(z \<mapsto> sv)) \<Gamma>"
- using is_satis_g_i_upd assms by (simp add: as)
- ultimately have "is_satis_g (i(z \<mapsto> sv)) ((z,b,c0)#\<^sub>\<Gamma>\<Gamma>)"
- using is_satis_g.simps by simp
- moreover have "wfI \<Theta> ((z,b,c0)#\<^sub>\<Gamma>\<Gamma>) (i(z \<mapsto> sv))" using as wfI_upd sv assms valid.simps wfC_wf by metis
- ultimately have "is_satis (i(z \<mapsto> sv)) c1" using assms valid.simps by auto
- thus "is_satis i (c1[z::=v]\<^sub>v)" using subst_c_satis sv as assms valid.simps wfC_wf wfG_elim2 subst_v_c_def by metis
- qed
-
- ultimately show ?thesis using valid.simps by auto
-qed
-
-lemma valid_trans_full:
- assumes "\<Theta> ; \<B> ; ((x, b, c1[z1::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) \<Turnstile> c2[z2::=V_var x]\<^sub>v" and
- "\<Theta> ; \<B> ; ((x, b, c2[z2::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) \<Turnstile> c3[z3::=V_var x]\<^sub>v"
- shows "\<Theta> ; \<B> ; ((x, b, c1[z1::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) \<Turnstile> c3[z3::=V_var x]\<^sub>v"
- unfolding valid.simps proof
- show "\<Theta> ; \<B> ; (x, b, c1[z1::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f c3[z3::=V_var x]\<^sub>v" using wf_trans valid.simps assms by metis
-
- show "\<forall>i. ( wfI \<Theta> ((x, b, c1[z1::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) i \<and> (is_satis_g i ((x, b, c1[z1::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)) \<longrightarrow> (is_satis i (c3[z3::=V_var x]\<^sub>v)) ) "
- proof(rule,rule)
- fix i
- assume as: "\<Theta> ; (x, b, c1[z1::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<turnstile> i \<and> i \<Turnstile> (x, b, c1[z1::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>"
- have "i \<Turnstile> c2[z2::=V_var x]\<^sub>v" using is_satis_g.simps as assms by simp
- moreover have "i \<Turnstile> \<Gamma>" using is_satis_g.simps as by simp
- ultimately show "i \<Turnstile> c3[z3::=V_var x]\<^sub>v " using assms is_satis_g.simps valid.simps
- by (metis append_g.simps(1) as wfI_replace_inside)
- qed
-qed
-
-lemma eval_v_weakening_x:
- fixes c::v
- assumes "i' \<lbrakk> c \<rbrakk> ~ s" and "atom x \<sharp> c" and "i = i' (x \<mapsto> s')"
- shows "i \<lbrakk> c \<rbrakk> ~ s"
- using assms proof(induct rule: eval_v.induct)
- case (eval_v_litI i l)
- then show ?case using eval_v.intros by auto
-next
- case (eval_v_varI sv i1 x1)
- hence "x \<noteq> x1" using v.fresh fresh_at_base by auto
- hence "i x1 = Some sv" using eval_v_varI by simp
- then show ?case using eval_v.intros by auto
-next
- case (eval_v_pairI i v1 s1 v2 s2)
- then show ?case using eval_v.intros by auto
-next
- case (eval_v_consI i v s tyid dc)
- then show ?case using eval_v.intros by auto
-next
- case (eval_v_conspI i v s tyid dc b)
- then show ?case using eval_v.intros by auto
-qed
-
-lemma eval_e_weakening_x:
- fixes c::ce
- assumes "i' \<lbrakk> c \<rbrakk> ~ s" and "atom x \<sharp> c" and "i = i' (x \<mapsto> s')"
- shows "i \<lbrakk> c \<rbrakk> ~ s"
- using assms proof(induct rule: eval_e.induct)
- case (eval_e_valI i v sv)
- then show ?case using eval_v_weakening_x eval_e.intros ce.fresh by metis
-next
- case (eval_e_plusI i v1 n1 v2 n2)
- then show ?case using eval_v_weakening_x eval_e.intros ce.fresh by metis
-next
- case (eval_e_leqI i v1 n1 v2 n2)
- then show ?case using eval_v_weakening_x eval_e.intros ce.fresh by metis
-next
- case (eval_e_eqI i v1 n1 v2 n2)
- then show ?case using eval_v_weakening_x eval_e.intros ce.fresh by metis
-next
- case (eval_e_fstI i v v1 v2)
- then show ?case using eval_v_weakening_x eval_e.intros ce.fresh by metis
-next
- case (eval_e_sndI i v v1 v2)
- then show ?case using eval_v_weakening_x eval_e.intros ce.fresh by metis
-next
- case (eval_e_concatI i v1 bv1 v2 bv2)
- then show ?case using eval_v_weakening_x eval_e.intros ce.fresh by metis
-next
- case (eval_e_lenI i v bv)
- then show ?case using eval_v_weakening_x eval_e.intros ce.fresh by metis
-qed
-
-lemma eval_c_weakening_x:
- fixes c::c
- assumes "i' \<lbrakk> c \<rbrakk> ~ s" and "atom x \<sharp> c" and "i = i' (x \<mapsto> s')"
- shows "i \<lbrakk> c \<rbrakk> ~ s"
- using assms proof(induct rule: eval_c.induct)
- case (eval_c_trueI i)
- then show ?case using eval_c.intros by auto
-next
- case (eval_c_falseI i)
- then show ?case using eval_c.intros by auto
-next
- case (eval_c_conjI i c1 b1 c2 b2)
- then show ?case using eval_c.intros by auto
-next
- case (eval_c_disjI i c1 b1 c2 b2)
- then show ?case using eval_c.intros by auto
-next
- case (eval_c_impI i c1 b1 c2 b2)
- then show ?case using eval_c.intros by auto
-next
- case (eval_c_notI i c b)
- then show ?case using eval_c.intros by auto
-next
- case (eval_c_eqI i e1 sv1 e2 sv2)
- then show ?case using eval_e_weakening_x c.fresh eval_c.intros by metis
-qed
-
-lemma is_satis_weakening_x:
- fixes c::c
- assumes "i' \<Turnstile> c" and "atom x \<sharp> c" and "i = i' (x \<mapsto> s)"
- shows "i \<Turnstile> c"
- using eval_c_weakening_x assms is_satis.simps by simp
-
-lemma is_satis_g_weakening_x:
- fixes G::\<Gamma>
- assumes "i' \<Turnstile> G" and "atom x \<sharp> G" and "i = i' (x \<mapsto> s)"
- shows "i \<Turnstile> G"
- using assms proof(induct G rule: \<Gamma>_induct)
- case GNil
- then show ?case by auto
-next
- case (GCons x' b' c' \<Gamma>')
- hence "atom x \<sharp> c'" using fresh_GCons fresh_prodN by simp
- moreover hence "i \<Turnstile> c'" using is_satis_weakening_x is_satis_g.simps(2) GCons by metis
- then show ?case using is_satis_g.simps(2)[of i x' b' c' \<Gamma>'] GCons fresh_GCons by simp
-qed
-
-section \<open>Base Type Substitution\<close>
-
-text \<open>The idea of boxing is to take an smt val and its base type and at nodes in the smt val that correspond to type variables we
-wrap them in an SUt smt val node. Another way of looking at it is that s' where the node for the base type variable is an 'any node'.
-It is needed to prove subst\_b\_valid - the base-type variable substitution lemma for validity.
-
-The first @{text "rcl_val"} is the expanded form (has type with base-variables replaced with base-type terms)
- ; the second is its corresponding form
-
-We only have one variable so we need to ensure that in all of the @{text "bs_boxed_BVarI"} cases, the s has the same
- base type.
-
-For example is an SMT value is (SPair (SInt 1) (SBool true)) and it has sort (BPair (BVar x) BBool)[x::=BInt] then
-the boxed version is SPair (SUt (SInt 1)) (SBool true) and is has sort (BPair (BVar x) BBool). We need to do this
-so that we can obtain from a valuation i, that gives values like the first smt value, to a valuation i' that gives values like
-the second.
-\<close>
-
-inductive boxed_b :: "\<Theta> \<Rightarrow> rcl_val \<Rightarrow> b \<Rightarrow> bv \<Rightarrow> b \<Rightarrow> rcl_val \<Rightarrow> bool" ( " _ \<turnstile> _ ~ _ [ _ ::= _ ] \<setminus> _ " [50,50] 50) where
- boxed_b_BVar1I: "\<lbrakk> bv = bv' ; wfRCV P s b \<rbrakk> \<Longrightarrow> boxed_b P s (B_var bv') bv b (SUt s)"
-| boxed_b_BVar2I: "\<lbrakk> bv \<noteq> bv'; wfRCV P s (B_var bv') \<rbrakk> \<Longrightarrow> boxed_b P s (B_var bv') bv b s"
-| boxed_b_BIntI:"wfRCV P s B_int \<Longrightarrow> boxed_b P s B_int _ _ s"
-| boxed_b_BBoolI:"wfRCV P s B_bool \<Longrightarrow> boxed_b P s B_bool _ _ s "
-| boxed_b_BUnitI:"wfRCV P s B_unit \<Longrightarrow> boxed_b P s B_unit _ _ s"
-| boxed_b_BPairI:"\<lbrakk> boxed_b P s1 b1 bv b s1' ; boxed_b P s2 b2 bv b s2' \<rbrakk> \<Longrightarrow> boxed_b P (SPair s1 s2) (B_pair b1 b2) bv b (SPair s1' s2')"
-
-| boxed_b_BConsI:"\<lbrakk>
- AF_typedef tyid dclist \<in> set P;
- (dc, \<lbrace> x : b | c \<rbrace>) \<in> set dclist ;
- boxed_b P s1 b bv b' s1'
- \<rbrakk> \<Longrightarrow>
- boxed_b P (SCons tyid dc s1) (B_id tyid) bv b' (SCons tyid dc s1')"
-
-| boxed_b_BConspI:"\<lbrakk> AF_typedef_poly tyid bva dclist \<in> set P;
- atom bva \<sharp> (b1,bv,b',s1,s1');
- (dc, \<lbrace> x : b | c \<rbrace>) \<in> set dclist ;
- boxed_b P s1 (b[bva::=b1]\<^sub>b\<^sub>b) bv b' s1'
- \<rbrakk> \<Longrightarrow>
- boxed_b P (SConsp tyid dc b1[bv::=b']\<^sub>b\<^sub>b s1) (B_app tyid b1) bv b' (SConsp tyid dc b1 s1')"
-
-| boxed_b_Bbitvec: "wfRCV P s B_bitvec \<Longrightarrow> boxed_b P s B_bitvec bv b s"
-
-equivariance boxed_b
-nominal_inductive boxed_b .
-
-inductive_cases boxed_b_elims:
- "boxed_b P s (B_var bv) bv' b s'"
- "boxed_b P s B_int bv b s'"
- "boxed_b P s B_bool bv b s'"
- "boxed_b P s B_unit bv b s'"
- "boxed_b P s (B_pair b1 b2) bv b s'"
- "boxed_b P s (B_id dc) bv b s'"
- "boxed_b P s B_bitvec bv b s'"
- "boxed_b P s (B_app dc b') bv b s'"
-
-lemma boxed_b_wfRCV:
- assumes "boxed_b P s b bv b' s'" (*and "supp s = {}"*) and "\<turnstile>\<^sub>w\<^sub>f P"
- shows "wfRCV P s b[bv::=b']\<^sub>b\<^sub>b \<and> wfRCV P s' b"
- using assms proof(induct rule: boxed_b.inducts)
- case (boxed_b_BVar1I bv bv' P s b )
- then show ?case using wfRCV.intros by auto
-next
- case (boxed_b_BVar2I bv bv' P s )
- then show ?case using wfRCV.intros by auto
-next
- case (boxed_b_BPairI P s1 b1 bv b s1' s2 b2 s2')
- then show ?case using wfRCV.intros rcl_val.supp by simp
-next
- case (boxed_b_BConsI tyid dclist P dc x b c s1 bv b' s1')
- hence "supp b = {}" using wfTh_supp_b by metis
- hence "b [ bv ::= b' ]\<^sub>b\<^sub>b = b" using fresh_def subst_b_b_def forget_subst[of "bv" b b'] by auto
- hence " P \<turnstile> SCons tyid dc s1 : (B_id tyid)" using wfRCV.intros rcl_val.supp subst_bb.simps boxed_b_BConsI by metis
- moreover have "P \<turnstile> SCons tyid dc s1' : B_id tyid" using boxed_b_BConsI
- using wfRCV.intros rcl_val.supp subst_bb.simps boxed_b_BConsI by metis
- ultimately show ?case using subst_bb.simps by metis
-next
- case (boxed_b_BConspI tyid bva dclist P b1 bv b' s1 s1' dc x b c)
-
- obtain bva2 and dclist2 where *:"AF_typedef_poly tyid bva dclist = AF_typedef_poly tyid bva2 dclist2 \<and>
- atom bva2 \<sharp> (bv,(P, SConsp tyid dc b1[bv::=b']\<^sub>b\<^sub>b s1, B_app tyid b1[bv::=b']\<^sub>b\<^sub>b))"
- using obtain_fresh_bv by metis
-
- then obtain x2 and b2 and c2 where **:\<open>(dc, \<lbrace> x2 : b2 | c2 \<rbrace>) \<in> set dclist2\<close>
- using boxed_b_BConspI td_lookup_eq_iff type_def.eq_iff by metis
-
- have "P \<turnstile> SConsp tyid dc b1[bv::=b']\<^sub>b\<^sub>b s1 : (B_app tyid b1[bv::=b']\<^sub>b\<^sub>b)" proof
- show 1: \<open>AF_typedef_poly tyid bva2 dclist2 \<in> set P\<close> using boxed_b_BConspI * by auto
- show 2: \<open>(dc, \<lbrace> x2 : b2 | c2 \<rbrace>) \<in> set dclist2\<close> using boxed_b_BConspI using ** by simp
-
- hence "atom bv \<sharp> b2" proof -
- have "supp b2 \<subseteq> { atom bva2 }" using wfTh_poly_supp_b 1 2 boxed_b_BConspI by auto
- moreover have "bv \<noteq> bva2" using * fresh_Pair fresh_at_base by metis
- ultimately show ?thesis using fresh_def by force
- qed
- moreover have "b[bva::=b1]\<^sub>b\<^sub>b = b2[bva2::=b1]\<^sub>b\<^sub>b" using wfTh_typedef_poly_b_eq_iff * 2 boxed_b_BConspI by metis
- ultimately show \<open> P \<turnstile> s1 : b2[bva2::=b1[bv::=b']\<^sub>b\<^sub>b]\<^sub>b\<^sub>b\<close> using boxed_b_BConspI subst_b_b_def subst_bb_commute by auto
- show "atom bva2 \<sharp> (P, SConsp tyid dc b1[bv::=b']\<^sub>b\<^sub>b s1, B_app tyid b1[bv::=b']\<^sub>b\<^sub>b)" using * fresh_Pair by metis
- qed
-
- moreover have "P \<turnstile> SConsp tyid dc b1 s1' : B_app tyid b1" proof
- show \<open>AF_typedef_poly tyid bva dclist \<in> set P\<close> using boxed_b_BConspI by auto
- show \<open>(dc, \<lbrace> x : b | c \<rbrace>) \<in> set dclist\<close> using boxed_b_BConspI by auto
- show \<open> P \<turnstile> s1' : b[bva::=b1]\<^sub>b\<^sub>b\<close> using boxed_b_BConspI by auto
- have "atom bva \<sharp> P" using boxed_b_BConspI wfTh_fresh by metis
- thus "atom bva \<sharp> (P, SConsp tyid dc b1 s1', B_app tyid b1)" using boxed_b_BConspI rcl_val.fresh b.fresh pure_fresh fresh_prodN by metis
- qed
-
- ultimately show ?case using subst_bb.simps by simp
-qed(auto)+
-
-lemma subst_b_var:
- assumes "B_var bv2 = b[bv::=b']\<^sub>b\<^sub>b"
- shows "(b = B_var bv \<and> b' = B_var bv2) \<or> (b=B_var bv2 \<and> bv \<noteq> bv2)"
- using assms by(nominal_induct b rule: b.strong_induct,auto+)
-
-text \<open>Here the valuation i' is the conv wrap version of i. For every x in G, i' x is the conv wrap version of i x.
-The next lemma for a clearer explanation of what this is. i produces values of sort b[bv::=b'] and i' produces values of sort b\<close>
-
-inductive boxed_i :: "\<Theta> \<Rightarrow> \<Gamma> \<Rightarrow> b \<Rightarrow> bv \<Rightarrow> valuation \<Rightarrow> valuation \<Rightarrow> bool" ( " _ ; _ ; _ , _ \<turnstile> _ \<approx> _" [50,50] 50) where
- boxed_i_GNilI: "\<Theta> ; GNil ; b , bv \<turnstile> i \<approx> i"
-| boxed_i_GConsI: "\<lbrakk> Some s = i x; boxed_b \<Theta> s b bv b' s' ; \<Theta> ; \<Gamma> ; b' , bv \<turnstile> i \<approx> i' \<rbrakk> \<Longrightarrow> \<Theta> ; ((x,b,c)#\<^sub>\<Gamma>\<Gamma>) ; b' , bv \<turnstile> i \<approx> (i'(x \<mapsto> s'))"
-equivariance boxed_i
-nominal_inductive boxed_i .
-
-inductive_cases boxed_i_elims:
- "\<Theta> ;GNil ; b , bv \<turnstile> i \<approx> i'"
- "\<Theta> ; ((x,b,c)#\<^sub>\<Gamma>\<Gamma>) ; b' , bv \<turnstile> i \<approx> i'"
-
-lemma wfRCV_poly_elims:
- fixes tm::"'a::fs" and b::b
- assumes "T \<turnstile> SConsp typid dc bdc s : b"
- obtains bva dclist x1 b1 c1 where "b = B_app typid bdc \<and>
- AF_typedef_poly typid bva dclist \<in> set T \<and> (dc, \<lbrace> x1 : b1 | c1 \<rbrace>) \<in> set dclist \<and> T \<turnstile> s : b1[bva::=bdc]\<^sub>b\<^sub>b \<and> atom bva \<sharp> tm"
- using assms proof(nominal_induct "SConsp typid dc bdc s" b avoiding: tm rule:wfRCV.strong_induct)
- case (wfRCV_BConsPI bv dclist \<Theta> x b c)
- then show ?case by simp
-qed
-
-lemma boxed_b_ex:
- assumes "wfRCV T s b[bv::=b']\<^sub>b\<^sub>b" and "wfTh T"
- shows "\<exists>s'. boxed_b T s b bv b' s'"
- using assms proof(nominal_induct s arbitrary: b rule: rcl_val.strong_induct)
- case (SBitvec x)
- have *:"b[bv::=b']\<^sub>b\<^sub>b = B_bitvec" using wfRCV_elims(9)[OF SBitvec(1)] by metis
- show ?case proof (cases "b = B_var bv")
- case True
- moreover have "T \<turnstile> SBitvec x : B_bitvec" using wfRCV.intros by simp
- moreover hence "b' = B_bitvec" using True SBitvec subst_bb.simps * by simp
- ultimately show ?thesis using boxed_b.intros wfRCV.intros by metis
- next
- case False
- hence "b = B_bitvec" using subst_bb_inject * by metis
- then show ?thesis using * SBitvec boxed_b.intros by metis
- qed
-next
- case (SNum x)
- have *:"b[bv::=b']\<^sub>b\<^sub>b = B_int" using wfRCV_elims(10)[OF SNum(1)] by metis
- show ?case proof (cases "b = B_var bv")
- case True
- moreover have "T \<turnstile> SNum x : B_int" using wfRCV.intros by simp
- moreover hence "b' = B_int" using True SNum subst_bb.simps(1) * by simp
- ultimately show ?thesis using boxed_b_BVar1I wfRCV.intros by metis
- next
- case False
- hence "b = B_int" using subst_bb_inject(1) * by metis
- then show ?thesis using * SNum boxed_b_BIntI by metis
- qed
-next
- case (SBool x)
- have *:"b[bv::=b']\<^sub>b\<^sub>b = B_bool" using wfRCV_elims(11)[OF SBool(1)] by metis
- show ?case proof (cases "b = B_var bv")
- case True
- moreover have "T \<turnstile> SBool x : B_bool" using wfRCV.intros by simp
- moreover hence "b' = B_bool" using True SBool subst_bb.simps * by simp
- ultimately show ?thesis using boxed_b.intros wfRCV.intros by metis
- next
- case False
- hence "b = B_bool" using subst_bb_inject * by metis
- then show ?thesis using * SBool boxed_b.intros by metis
- qed
-next
- case (SPair s1 s2)
- then obtain b1 and b2 where *:"b[bv::=b']\<^sub>b\<^sub>b = B_pair b1 b2 \<and> wfRCV T s1 b1 \<and> wfRCV T s2 b2" using wfRCV_elims(12) by metis
- show ?case proof (cases "b = B_var bv")
- case True
- moreover have "T \<turnstile> SPair s1 s2 : B_pair b1 b2" using wfRCV.intros * by simp
- moreover hence "b' = B_pair b1 b2" using True SPair subst_bb.simps(1) * by simp
- ultimately show ?thesis using boxed_b_BVar1I by metis
- next
- case False
- then obtain b1' and b2' where "b = B_pair b1' b2' \<and> b1=b1'[bv::=b']\<^sub>b\<^sub>b \<and> b2=b2'[bv::=b']\<^sub>b\<^sub>b" using subst_bb_inject(5)[OF _ False ] * by metis
- then show ?thesis using * SPair boxed_b_BPairI by blast
- qed
-next
- case (SCons tyid dc s1)
- have *:"b[bv::=b']\<^sub>b\<^sub>b = B_id tyid" using wfRCV_elims(13)[OF SCons(2)] by metis
- show ?case proof (cases "b = B_var bv")
- case True
- moreover have "T \<turnstile> SCons tyid dc s1 : B_id tyid" using wfRCV.intros
- using "local.*" SCons.prems by auto
- moreover hence "b' = B_id tyid" using True SCons subst_bb.simps(1) * by simp
- ultimately show ?thesis using boxed_b_BVar1I wfRCV.intros by metis
- next
- case False
- then obtain b1' where beq: "b = B_id tyid" using subst_bb_inject * by metis
- then obtain b2 dclist x c where **:"AF_typedef tyid dclist \<in> set T \<and> (dc, \<lbrace> x : b2 | c \<rbrace>) \<in> set dclist \<and> wfRCV T s1 b2" using wfRCV_elims(13) * SCons by metis
- then have "atom bv \<sharp> b2" using \<open>wfTh T\<close> wfTh_lookup_supp_empty[of tyid dclist T dc "\<lbrace> x : b2 | c \<rbrace>"] \<tau>.fresh fresh_def by auto
- then have "b2 = b2[ bv ::= b']\<^sub>b\<^sub>b" using forget_subst subst_b_b_def by metis
- then obtain s1' where s1:"T \<turnstile> s1 ~ b2 [ bv ::= b' ] \<setminus> s1'" using SCons ** by metis
-
- have "T \<turnstile> SCons tyid dc s1 ~ (B_id tyid) [ bv ::= b' ] \<setminus> SCons tyid dc s1'" proof(rule boxed_b_BConsI)
- show "AF_typedef tyid dclist \<in> set T" using ** by auto
- show "(dc, \<lbrace> x : b2 | c \<rbrace>) \<in> set dclist" using ** by auto
- show "T \<turnstile> s1 ~ b2 [ bv ::= b' ] \<setminus> s1' " using s1 ** by auto
-
- qed
- thus ?thesis using beq by metis
- qed
-next
- case (SConsp typid dc bdc s)
-
- obtain bva dclist x1 b1 c1 where **:"b[bv::=b']\<^sub>b\<^sub>b = B_app typid bdc \<and>
- AF_typedef_poly typid bva dclist \<in> set T \<and> (dc, \<lbrace> x1 : b1 | c1 \<rbrace>) \<in> set dclist \<and> T \<turnstile> s : b1[bva::=bdc]\<^sub>b\<^sub>b \<and> atom bva \<sharp> bv"
- using wfRCV_poly_elims[OF SConsp(2)] by metis
-
- then have *:"B_app typid bdc = b[bv::=b']\<^sub>b\<^sub>b" using wfRCV_elims(14)[OF SConsp(2)] by metis
- show ?case proof (cases "b = B_var bv")
- case True
- moreover have "T \<turnstile> SConsp typid dc bdc s : B_app typid bdc" using wfRCV.intros
- using "local.*" SConsp.prems(1) by auto
- moreover hence "b' = B_app typid bdc" using True SConsp subst_bb.simps * by simp
- ultimately show ?thesis using boxed_b.intros wfRCV.intros by metis
- next
- case False
- then obtain bdc' where bdc: "b = B_app typid bdc' \<and> bdc = bdc'[bv::=b']\<^sub>b\<^sub>b" using * subst_bb_inject(8)[OF *] by metis
- (*hence beq:"b = B_app typid bdc" using subst_bb_inject * sory (* going to be like the above as bdc is closed *)*)
- have "atom bv \<sharp> b1" proof -
- have "supp b1 \<subseteq> { atom bva }" using wfTh_poly_supp_b ** SConsp by metis
- moreover have "bv \<noteq> bva" using ** by auto
- ultimately show ?thesis using fresh_def by force
- qed
- have "T \<turnstile> s : b1[bva::=bdc]\<^sub>b\<^sub>b" using ** by auto
- moreover have "b1[bva::=bdc']\<^sub>b\<^sub>b[bv::=b']\<^sub>b\<^sub>b = b1[bva::=bdc]\<^sub>b\<^sub>b" using bdc subst_bb_commute \<open>atom bv \<sharp> b1\<close> by auto
- ultimately obtain s' where s':"T \<turnstile> s ~ b1[bva::=bdc']\<^sub>b\<^sub>b [ bv ::= b' ] \<setminus> s'"
- using SConsp(1)[of "b1[bva::=bdc']\<^sub>b\<^sub>b"] bdc SConsp by metis
- have "T \<turnstile> SConsp typid dc bdc'[bv::=b']\<^sub>b\<^sub>b s ~ (B_app typid bdc') [ bv ::= b' ] \<setminus> SConsp typid dc bdc' s'"
- proof -
-
- obtain bva3 and dclist3 where 3:"AF_typedef_poly typid bva3 dclist3 = AF_typedef_poly typid bva dclist \<and>
- atom bva3 \<sharp> (bdc', bv, b', s, s')" using obtain_fresh_bv by metis
- then obtain x3 b3 c3 where 4:"(dc, \<lbrace> x3 : b3 | c3 \<rbrace>) \<in> set dclist3"
- using boxed_b_BConspI td_lookup_eq_iff type_def.eq_iff
- by (metis "**")
-
- show ?thesis proof
- show \<open>AF_typedef_poly typid bva3 dclist3 \<in> set T\<close> using 3 ** by metis
- show \<open>atom bva3 \<sharp> (bdc', bv, b', s, s')\<close> using 3 by metis
- show 4:\<open>(dc, \<lbrace> x3 : b3 | c3 \<rbrace>) \<in> set dclist3\<close> using 4 by auto
- have "b3[bva3::=bdc']\<^sub>b\<^sub>b = b1[bva::=bdc']\<^sub>b\<^sub>b" proof(rule wfTh_typedef_poly_b_eq_iff)
- show \<open>AF_typedef_poly typid bva3 dclist3 \<in> set T\<close> using 3 ** by metis
- show \<open>(dc, \<lbrace> x3 : b3 | c3 \<rbrace>) \<in> set dclist3\<close> using 4 by auto
- show \<open>AF_typedef_poly typid bva dclist \<in> set T\<close> using ** by auto
- show \<open>(dc, \<lbrace> x1 : b1 | c1 \<rbrace>) \<in> set dclist\<close> using ** by auto
- qed(simp add: ** SConsp)
- thus \<open> T \<turnstile> s ~ b3[bva3::=bdc']\<^sub>b\<^sub>b [ bv ::= b' ] \<setminus> s' \<close> using s' by auto
- qed
- qed
- then show ?thesis using bdc by auto
-
- qed
-next
- case SUnit
- have *:"b[bv::=b']\<^sub>b\<^sub>b = B_unit" using wfRCV_elims SUnit by metis
- show ?case proof (cases "b = B_var bv")
- case True
- moreover have "T \<turnstile> SUnit : B_unit" using wfRCV.intros by simp
- moreover hence "b' = B_unit" using True SUnit subst_bb.simps * by simp
- ultimately show ?thesis using boxed_b.intros wfRCV.intros by metis
- next
- case False
- hence "b = B_unit" using subst_bb_inject * by metis
- then show ?thesis using * SUnit boxed_b.intros by metis
- qed
-next
- case (SUt x)
- then obtain bv' where *:"b[bv::=b']\<^sub>b\<^sub>b= B_var bv'" using wfRCV_elims by metis
- show ?case proof (cases "b = B_var bv")
- case True
- then show ?thesis using boxed_b_BVar1I
- using "local.*" wfRCV_BVarI by fastforce
- next
- case False
- then show ?thesis using boxed_b_BVar1I boxed_b_BVar2I
- using "local.*" wfRCV_BVarI by (metis subst_b_var)
- qed
-qed
-
-lemma boxed_i_ex:
- assumes "wfI T \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b i" and "wfTh T"
- shows "\<exists>i'. T ; \<Gamma> ; b , bv \<turnstile> i \<approx> i'"
- using assms proof(induct \<Gamma> arbitrary: i rule:\<Gamma>_induct)
- case GNil
- then show ?case using boxed_i_GNilI by metis
-next
- case (GCons x' b' c' \<Gamma>')
- then obtain s where 1:"Some s = i x' \<and> wfRCV T s b'[bv::=b]\<^sub>b\<^sub>b" using wfI_def subst_gb.simps by auto
- then obtain s' where 2: "boxed_b T s b' bv b s'" using boxed_b_ex GCons by metis
- then obtain i' where 3: "boxed_i T \<Gamma>' b bv i i'" using GCons wfI_def subst_gb.simps by force
- have "boxed_i T ((x', b', c') #\<^sub>\<Gamma> \<Gamma>') b bv i (i'(x' \<mapsto> s'))" proof
- show "Some s = i x'" using 1 by auto
- show "boxed_b T s b' bv b s'" using 2 by auto
- show "T ; \<Gamma>' ; b , bv \<turnstile> i \<approx> i'" using "3" by auto
- qed
- thus ?case by auto
-qed
-
-lemma boxed_b_eq:
- assumes "boxed_b \<Theta> s1 b bv b' s1'" and "\<turnstile>\<^sub>w\<^sub>f \<Theta>"
- shows "wfTh \<Theta> \<Longrightarrow> boxed_b \<Theta> s2 b bv b' s2' \<Longrightarrow> ( s1 = s2 ) = ( s1' = s2' )"
- using assms proof(induct arbitrary: s2 s2' rule: boxed_b.inducts )
- case (boxed_b_BVar1I bv bv' P s b )
- then show ?case
- using boxed_b_elims(1) rcl_val.eq_iff by metis
-next
- case (boxed_b_BVar2I bv bv' P s b)
- then show ?case using boxed_b_elims(1) by metis
-next
- case (boxed_b_BIntI P s uu uv)
- hence "s2 = s2'" using boxed_b_elims by metis
- then show ?case by auto
-next
- case (boxed_b_BBoolI P s uw ux)
- hence "s2 = s2'" using boxed_b_elims by metis
- then show ?case by auto
-next
- case (boxed_b_BUnitI P s uy uz)
- hence "s2 = s2'" using boxed_b_elims by metis
- then show ?case by auto
-next
- case (boxed_b_BPairI P s1 b1 bv b s1' s2a b2 s2a')
- then show ?case
- by (metis boxed_b_elims(5) rcl_val.eq_iff(4))
-next
- case (boxed_b_BConsI tyid dclist P dc x b c s1 bv b' s1')
- obtain s22 and s22' dclist2 dc2 x2 b2 c2 where *:"s2 = SCons tyid dc2 s22 \<and> s2' = SCons tyid dc2 s22' \<and> boxed_b P s22 b2 bv b' s22'
- \<and> AF_typedef tyid dclist2 \<in> set P \<and> (dc2, \<lbrace> x2 : b2 | c2 \<rbrace>) \<in> set dclist2" using boxed_b_elims(6)[OF boxed_b_BConsI(6)] by metis
- show ?case proof(cases "dc = dc2")
- case True
- hence "b = b2" using wfTh_ctor_unique \<tau>.eq_iff wfTh_dclist_unique wf boxed_b_BConsI * by metis
- then show ?thesis using boxed_b_BConsI True * by auto
- next
- case False
- then show ?thesis using * boxed_b_BConsI by simp
- qed
-next
- case (boxed_b_Bbitvec P s bv b)
- hence "s2 = s2'" using boxed_b_elims by metis
- then show ?case by auto
-next
- case (boxed_b_BConspI tyid bva dclist P b1 bv b' s1 s1' dc x b c)
- obtain bva2 s22 s22' dclist2 dc2 x2 b2 c2 where *:"
- s2 = SConsp tyid dc2 b1[bv::=b']\<^sub>b\<^sub>b s22 \<and>
- s2' = SConsp tyid dc2 b1 s22' \<and>
- boxed_b P s22 b2[bva2::=b1]\<^sub>b\<^sub>b bv b' s22' \<and>
- AF_typedef_poly tyid bva2 dclist2 \<in> set P \<and> (dc2, \<lbrace> x2 : b2 | c2 \<rbrace>) \<in> set dclist2" using boxed_b_elims(8)[OF boxed_b_BConspI(7)] by metis
- show ?case proof(cases "dc = dc2")
- case True
- hence "AF_typedef_poly tyid bva2 dclist2 \<in> set P \<and> (dc, \<lbrace> x2 : b2 | c2 \<rbrace>) \<in> set dclist2" using * by auto
- hence "b[bva::=b1]\<^sub>b\<^sub>b = b2[bva2::=b1]\<^sub>b\<^sub>b" using wfTh_typedef_poly_b_eq_iff[OF boxed_b_BConspI(1) boxed_b_BConspI(3)] * boxed_b_BConspI by metis
- then show ?thesis using boxed_b_BConspI True * by auto
- next
- case False
- then show ?thesis using * boxed_b_BConspI by simp
- qed
-qed
-
-lemma bs_boxed_var:
- assumes "boxed_i \<Theta> \<Gamma> b' bv i i'"
- shows "Some (b,c) = lookup \<Gamma> x \<Longrightarrow> Some s = i x \<Longrightarrow> Some s' = i' x \<Longrightarrow> boxed_b \<Theta> s b bv b' s'"
- using assms proof(induct rule: boxed_i.inducts)
- case (boxed_i_GNilI T i)
- then show ?case using lookup.simps by auto
-next
- case (boxed_i_GConsI s i x1 \<Theta> b1 bv b' s' \<Gamma> i' c)
- show ?case proof (cases "x=x1")
- case True
- then show ?thesis using boxed_i_GConsI
- fun_upd_same lookup.simps(2) option.inject prod.inject by metis
- next
- case False
- then show ?thesis using boxed_i_GConsI
- fun_upd_same lookup.simps option.inject prod.inject by auto
- qed
-qed
-
-lemma eval_l_boxed_b:
- assumes "\<lbrakk> l \<rbrakk> = s"
- shows "boxed_b \<Theta> s (base_for_lit l) bv b' s"
- using assms proof(nominal_induct l arbitrary: s rule:l.strong_induct)
-qed(auto simp add: boxed_b.intros wfRCV.intros )+
-
-lemma boxed_i_eval_v_boxed_b:
- fixes v::v
- assumes "boxed_i \<Theta> \<Gamma> b' bv i i'" and "i \<lbrakk> v[bv::=b']\<^sub>v\<^sub>b \<rbrakk> ~ s" and "i' \<lbrakk> v \<rbrakk> ~ s'" and "wfV \<Theta> B \<Gamma> v b" and "wfI \<Theta> \<Gamma> i'"
- shows "boxed_b \<Theta> s b bv b' s'"
- using assms proof(nominal_induct v arbitrary: s s' b rule:v.strong_induct)
- case (V_lit l)
- hence "\<lbrakk> l \<rbrakk> = s \<and> \<lbrakk> l \<rbrakk> = s'" using eval_v_elims by auto
- moreover have "b = base_for_lit l" using wfV_elims(2) V_lit by metis
- ultimately show ?case using V_lit using eval_l_boxed_b subst_b_base_for_lit by metis
-next
- case (V_var x) (* look at defn of bs_boxed *)
- hence "Some s = i x \<and> Some s' = i' x" using eval_v_elims subst_vb.simps by metis
- moreover obtain c1 where bc:"Some (b,c1) = lookup \<Gamma> x" using wfV_elims V_var by metis
- ultimately show ?case using bs_boxed_var V_var by metis
-
-next
- case (V_pair v1 v2)
- then obtain b1 and b2 where b:"b=B_pair b1 b2" using wfV_elims subst_vb.simps by metis
- obtain s1 and s2 where s: "eval_v i (v1[bv::=b']\<^sub>v\<^sub>b) s1 \<and> eval_v i (v2[bv::=b']\<^sub>v\<^sub>b) s2 \<and> s = SPair s1 s2" using eval_v_elims V_pair subst_vb.simps by metis
- obtain s1' and s2' where s': "eval_v i' v1 s1' \<and> eval_v i' v2 s2' \<and> s' = SPair s1' s2'" using eval_v_elims V_pair by metis
- have "boxed_b \<Theta> (SPair s1 s2) (B_pair b1 b2) bv b' (SPair s1' s2') " proof(rule boxed_b_BPairI)
- show "boxed_b \<Theta> s1 b1 bv b' s1'" using V_pair eval_v_elims wfV_elims b s s' b.eq_iff by metis
- show "boxed_b \<Theta> s2 b2 bv b' s2'" using V_pair eval_v_elims wfV_elims b s s' b.eq_iff by metis
- qed
- then show ?case using s s' b by auto
-next
- case (V_cons tyid dc v1)
-
- obtain dclist x b1 c where *: "b = B_id tyid \<and> AF_typedef tyid dclist \<in> set \<Theta> \<and> (dc, \<lbrace> x : b1 | c \<rbrace>) \<in> set dclist \<and> \<Theta> ; B ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v1 : b1"
- using wfV_elims(4)[OF V_cons(5)] V_cons by metis
- obtain s2 where s2: "s = SCons tyid dc s2 \<and> i \<lbrakk> (v1[bv::=b']\<^sub>v\<^sub>b) \<rbrakk> ~ s2" using eval_v_elims V_cons subst_vb.simps by metis
- obtain s2' where s2': "s' = SCons tyid dc s2' \<and> i' \<lbrakk> v1 \<rbrakk> ~ s2'" using eval_v_elims V_cons by metis
- have sp: "supp \<lbrace> x : b1 | c \<rbrace> = {}" using wfTh_lookup_supp_empty * wfX_wfY by metis
-
- have "boxed_b \<Theta> (SCons tyid dc s2) (B_id tyid) bv b' (SCons tyid dc s2')"
- proof(rule boxed_b_BConsI)
- show 1:"AF_typedef tyid dclist \<in> set \<Theta>" using * by auto
- show 2:"(dc, \<lbrace> x : b1 | c \<rbrace>) \<in> set dclist" using * by auto
- have bvf:"atom bv \<sharp> b1" using sp \<tau>.fresh fresh_def by auto
- show "\<Theta> \<turnstile> s2 ~ b1 [ bv ::= b' ] \<setminus> s2'" using V_cons s2 s2' * by metis
- qed
- then show ?case using * s2 s2' by simp
-next
- case (V_consp tyid dc b1 v1)
-
- obtain bv2 dclist x2 b2 c2 where *: "b = B_app tyid b1 \<and> AF_typedef_poly tyid bv2 dclist \<in> set \<Theta> \<and>
- (dc, \<lbrace> x2 : b2 | c2 \<rbrace>) \<in> set dclist \<and> \<Theta> ; B ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v1 : b2[bv2::=b1]\<^sub>b\<^sub>b"
- using wf_strong_elim(1)[OF V_consp (5)] by metis
-
- obtain s2 where s2: "s = SConsp tyid dc b1[bv::=b']\<^sub>b\<^sub>b s2 \<and> i \<lbrakk> (v1[bv::=b']\<^sub>v\<^sub>b) \<rbrakk> ~ s2"
- using eval_v_elims V_consp subst_vb.simps by metis
-
- obtain s2' where s2': "s' = SConsp tyid dc b1 s2' \<and> i' \<lbrakk> v1 \<rbrakk> ~ s2'"
- using eval_v_elims V_consp by metis
-
- have "\<turnstile>\<^sub>w\<^sub>f \<Theta>" using V_consp wfX_wfY by metis
- then obtain bv3::bv and dclist3 x3 b3 c3 where **:" AF_typedef_poly tyid bv2 dclist = AF_typedef_poly tyid bv3 dclist3 \<and>
- (dc, \<lbrace> x3 : b3 | c3 \<rbrace>) \<in> set dclist3 \<and> atom bv3 \<sharp> (b1, bv, b', s2, s2') \<and> b2[bv2::=b1]\<^sub>b\<^sub>b = b3[bv3::=b1]\<^sub>b\<^sub>b"
- using * obtain_fresh_bv_dclist_b_iff[where tm="(b1, bv, b', s2, s2')"] by metis
-
- have "boxed_b \<Theta> (SConsp tyid dc b1[bv::=b']\<^sub>b\<^sub>b s2) (B_app tyid b1) bv b' (SConsp tyid dc b1 s2')"
- proof(rule boxed_b_BConspI[of tyid bv3 dclist3 \<Theta>, where x=x3 and b=b3 and c=c3])
- show 1:"AF_typedef_poly tyid bv3 dclist3 \<in> set \<Theta>" using * ** by auto
- show 2:"(dc, \<lbrace> x3 : b3 | c3 \<rbrace>) \<in> set dclist3" using ** by auto
- show "atom bv3 \<sharp> (b1, bv, b', s2, s2')" using ** by auto
- show " \<Theta> \<turnstile> s2 ~ b3[bv3::=b1]\<^sub>b\<^sub>b [ bv ::= b' ] \<setminus> s2'" using V_consp s2 s2' * ** by metis
- qed
- then show ?case using * s2 s2' by simp
-qed
-
-lemma boxed_b_eq_eq:
- assumes "boxed_b \<Theta> n1 b1 bv b' n1'" and "boxed_b \<Theta> n2 b1 bv b' n2'" and "s = SBool (n1 = n2)" and "\<turnstile>\<^sub>w\<^sub>f \<Theta>"
- "s' = SBool (n1' = n2')"
- shows "s=s'"
- using boxed_b_eq assms by auto
-
-lemma boxed_i_eval_ce_boxed_b:
- fixes e::ce
- assumes "i' \<lbrakk> e \<rbrakk> ~ s'" and "i \<lbrakk> e[bv::=b']\<^sub>c\<^sub>e\<^sub>b \<rbrakk> ~ s" and "wfCE \<Theta> B \<Gamma> e b" and "boxed_i \<Theta> \<Gamma> b' bv i i'" and "wfI \<Theta> \<Gamma> i'"
- shows "boxed_b \<Theta> s b bv b' s'"
- using assms proof(nominal_induct e arbitrary: s s' b b' rule: ce.strong_induct)
- case (CE_val x)
- then show ?case using boxed_i_eval_v_boxed_b eval_e_elims wfCE_elims subst_ceb.simps by metis
-next
- case (CE_op opp v1 v2)
-
- show ?case proof(rule opp.exhaust)
- assume \<open>opp = Plus\<close>
- have 1:"wfCE \<Theta> B \<Gamma> v1 (B_int)" using wfCE_elims CE_op \<open>opp = Plus\<close> by metis
- have 2:"wfCE \<Theta> B \<Gamma> v2 (B_int)" using wfCE_elims CE_op \<open>opp = Plus\<close> by metis
- have *:"b = B_int" using CE_op wfCE_elims
- by (metis \<open>opp = plus\<close>)
-
- obtain n1 and n2 where n:"s = SNum (n1 + n2) \<and> i \<lbrakk> v1[bv::=b']\<^sub>c\<^sub>e\<^sub>b \<rbrakk> ~ SNum n1 \<and> i \<lbrakk> v2[bv::=b']\<^sub>c\<^sub>e\<^sub>b \<rbrakk> ~ SNum n2" using eval_e_elims CE_op subst_ceb.simps \<open>opp = plus\<close> by metis
- obtain n1' and n2' where n':"s' = SNum (n1' + n2') \<and> i' \<lbrakk> v1 \<rbrakk> ~ SNum n1' \<and> i' \<lbrakk> v2 \<rbrakk> ~ SNum n2'" using eval_e_elims Plus CE_op \<open>opp = plus\<close> by metis
-
- have "boxed_b \<Theta> (SNum n1) B_int bv b' (SNum n1')" using boxed_i_eval_v_boxed_b 1 2 n n' CE_op \<open>opp = plus\<close> by metis
- moreover have "boxed_b \<Theta> (SNum n2) B_int bv b' (SNum n2')" using boxed_i_eval_v_boxed_b 1 2 n n' CE_op by metis
- ultimately have "s=s'" using n' n boxed_b_elims(2)
- by (metis rcl_val.eq_iff(2))
- thus ?thesis using * n n' boxed_b_BIntI CE_op wfRCV.intros Plus by simp
- next
- assume \<open>opp = LEq\<close>
- have 1:"wfCE \<Theta> B \<Gamma> v1 (B_int)" using wfCE_elims CE_op \<open>opp = LEq\<close> by metis
- have 2:"wfCE \<Theta> B \<Gamma> v2 (B_int)" using wfCE_elims CE_op \<open>opp = LEq\<close> by metis
- hence *:"b = B_bool" using CE_op wfCE_elims \<open>opp = LEq\<close> by metis
- obtain n1 and n2 where n:"s = SBool (n1 \<le> n2) \<and> i \<lbrakk> v1[bv::=b']\<^sub>c\<^sub>e\<^sub>b \<rbrakk> ~ SNum n1 \<and> i \<lbrakk> v2[bv::=b']\<^sub>c\<^sub>e\<^sub>b \<rbrakk> ~ SNum n2" using eval_e_elims subst_ceb.simps CE_op \<open>opp = LEq\<close> by metis
- obtain n1' and n2' where n':"s' = SBool (n1' \<le> n2') \<and> i' \<lbrakk> v1 \<rbrakk> ~ SNum n1' \<and> i' \<lbrakk> v2 \<rbrakk> ~ SNum n2'" using eval_e_elims CE_op \<open>opp = LEq\<close> by metis
-
- have "boxed_b \<Theta> (SNum n1) B_int bv b' (SNum n1')" using boxed_i_eval_v_boxed_b 1 2 n n' CE_op by metis
- moreover have "boxed_b \<Theta> (SNum n2) B_int bv b' (SNum n2')" using boxed_i_eval_v_boxed_b 1 2 n n' CE_op by metis
- ultimately have "s=s'" using n' n boxed_b_elims(2)
- by (metis rcl_val.eq_iff(2))
- thus ?thesis using * n n' boxed_b_BBoolI CE_op wfRCV.intros \<open>opp = LEq\<close> by simp
- next
- assume \<open>opp = Eq\<close>
- obtain b1 where b1:"wfCE \<Theta> B \<Gamma> v1 b1 \<and> wfCE \<Theta> B \<Gamma> v2 b1" using wfCE_elims CE_op \<open>opp = Eq\<close> by metis
-
- hence *:"b = B_bool" using CE_op wfCE_elims \<open>opp = Eq\<close> by metis
- obtain n1 and n2 where n:"s = SBool (n1 = n2) \<and> i \<lbrakk> v1[bv::=b']\<^sub>c\<^sub>e\<^sub>b \<rbrakk> ~ n1 \<and> i \<lbrakk> v2[bv::=b']\<^sub>c\<^sub>e\<^sub>b \<rbrakk> ~ n2" using eval_e_elims subst_ceb.simps CE_op \<open>opp = Eq\<close> by metis
- obtain n1' and n2' where n':"s' = SBool (n1' = n2') \<and> i' \<lbrakk> v1 \<rbrakk> ~ n1' \<and> i' \<lbrakk> v2 \<rbrakk> ~ n2'" using eval_e_elims CE_op \<open>opp = Eq\<close> by metis
-
- have "boxed_b \<Theta> n1 b1 bv b' n1'" using boxed_i_eval_v_boxed_b b1 n n' CE_op by metis
- moreover have "boxed_b \<Theta> n2 b1 bv b' n2'" using boxed_i_eval_v_boxed_b b1 n n' CE_op by metis
- moreover have "\<turnstile>\<^sub>w\<^sub>f \<Theta>" using b1 wfX_wfY by metis
- ultimately have "s=s'" using n' n boxed_b_elims
- boxed_b_eq_eq by metis
- thus ?thesis using * n n' boxed_b_BBoolI CE_op wfRCV.intros \<open>opp = Eq\<close> by simp
- qed
-
-next
- case (CE_concat v1 v2)
-
- obtain bv1 and bv2 where s : "s = SBitvec (bv1 @ bv2) \<and> (i \<lbrakk> v1[bv::=b']\<^sub>c\<^sub>e\<^sub>b \<rbrakk> ~ SBitvec bv1) \<and> i \<lbrakk> v2[bv::=b']\<^sub>c\<^sub>e\<^sub>b \<rbrakk> ~ SBitvec bv2"
- using eval_e_elims(7) subst_ceb.simps CE_concat.prems(2) eval_e_elims(6) subst_ceb.simps(6) by metis
- obtain bv1' and bv2' where s' : "s' = SBitvec (bv1' @ bv2') \<and> i' \<lbrakk> v1 \<rbrakk> ~ SBitvec bv1' \<and> i' \<lbrakk> v2 \<rbrakk> ~ SBitvec bv2'"
- using eval_e_elims(7) CE_concat by metis
-
- then show ?case using boxed_i_eval_v_boxed_b wfCE_elims s s' CE_concat
- by (metis CE_concat.prems(3) assms assms(5) wfRCV_BBitvecI boxed_b_Bbitvec boxed_b_elims(7) eval_e_concatI eval_e_uniqueness)
-next
- case (CE_fst ce)
- obtain s2 where 1:"i \<lbrakk> ce[bv::=b']\<^sub>c\<^sub>e\<^sub>b \<rbrakk> ~ SPair s s2" using CE_fst eval_e_elims subst_ceb.simps by metis
- obtain s2' where 2:"i' \<lbrakk> ce \<rbrakk> ~ SPair s' s2'" using CE_fst eval_e_elims by metis
- obtain b2 where 3:"wfCE \<Theta> B \<Gamma> ce (B_pair b b2)" using wfCE_elims(4) CE_fst by metis
-
- have "boxed_b \<Theta> (SPair s s2) (B_pair b b2) bv b' (SPair s' s2')"
- using 1 2 3 CE_fst boxed_i_eval_v_boxed_b boxed_b_BPairI by auto
- thus ?case using boxed_b_elims(5) by force
-next
- case (CE_snd v)
- obtain s1 where 1:"i \<lbrakk> v[bv::=b']\<^sub>c\<^sub>e\<^sub>b \<rbrakk> ~ SPair s1 s" using CE_snd eval_e_elims subst_ceb.simps by metis
- obtain s1' where 2:"i' \<lbrakk> v \<rbrakk> ~ SPair s1' s'" using CE_snd eval_e_elims by metis
- obtain b1 where 3:"wfCE \<Theta> B \<Gamma> v (B_pair b1 b )" using wfCE_elims(5) CE_snd by metis
-
- have "boxed_b \<Theta> (SPair s1 s ) (B_pair b1 b ) bv b' (SPair s1' s')" using 1 2 3 CE_snd boxed_i_eval_v_boxed_b by simp
- thus ?case using boxed_b_elims(5) by force
-next
- case (CE_len v)
- obtain s1 where s: "i \<lbrakk> v[bv::=b']\<^sub>c\<^sub>e\<^sub>b \<rbrakk> ~ SBitvec s1" using CE_len eval_e_elims subst_ceb.simps by metis
- obtain s1' where s': "i' \<lbrakk> v \<rbrakk> ~ SBitvec s1'" using CE_len eval_e_elims by metis
-
- have "\<Theta> ; B ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : B_bitvec \<and> b = B_int" using wfCE_elims CE_len by metis
- then show ?case using boxed_i_eval_v_boxed_b s s' CE_len
- by (metis boxed_b_BIntI boxed_b_elims(7) eval_e_lenI eval_e_uniqueness subst_ceb.simps(5) wfI_wfCE_eval_e)
-qed
-
-lemma eval_c_eq_bs_boxed:
- fixes c::c
- assumes "i \<lbrakk> c[bv::=b]\<^sub>c\<^sub>b \<rbrakk> ~ s" and "i' \<lbrakk> c \<rbrakk> ~ s'" and "wfC \<Theta> B \<Gamma> c" and "wfI \<Theta> \<Gamma> i'" and "\<Theta> ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> i "
- and "boxed_i \<Theta> \<Gamma> b bv i i'"
- shows "s = s'"
- using assms proof(nominal_induct c arbitrary: s s' rule:c.strong_induct)
- case C_true
- then show ?case using eval_c_elims subst_cb.simps by metis
-next
- case C_false
- then show ?case using eval_c_elims subst_cb.simps by metis
-next
- case (C_conj c1 c2)
- obtain s1 and s2 where 1: "eval_c i (c1[bv::=b]\<^sub>c\<^sub>b) s1 \<and> eval_c i (c2[bv::=b]\<^sub>c\<^sub>b) s2 \<and> s = (s1\<and>s2)" using C_conj eval_c_elims(3) subst_cb.simps(3) by metis
- obtain s1' and s2' where 2:"eval_c i' c1 s1' \<and> eval_c i' c2 s2' \<and> s' = (s1'\<and>s2')" using C_conj eval_c_elims(3) by metis
- then show ?case using 1 2 wfC_elims C_conj by metis
-next
- case (C_disj c1 c2)
-
- obtain s1 and s2 where 1: "eval_c i (c1[bv::=b]\<^sub>c\<^sub>b) s1 \<and> eval_c i (c2[bv::=b]\<^sub>c\<^sub>b) s2 \<and> s = (s1\<or>s2)" using C_disj eval_c_elims(4) subst_cb.simps(4) by metis
- obtain s1' and s2' where 2:"eval_c i' c1 s1' \<and> eval_c i' c2 s2' \<and> s' = (s1'\<or>s2')" using C_disj eval_c_elims(4) by metis
- then show ?case using 1 2 wfC_elims C_disj by metis
-next
- case (C_not c)
- obtain s1::bool where 1: "(i \<lbrakk> c[bv::=b]\<^sub>c\<^sub>b \<rbrakk> ~ s1) \<and> (s = (\<not> s1))" using C_not eval_c_elims(6) subst_cb.simps(7) by metis
- obtain s1'::bool where 2: "(i' \<lbrakk> c \<rbrakk> ~ s1') \<and> (s' = (\<not> s1'))" using C_not eval_c_elims(6) by metis
- then show ?case using 1 2 wfC_elims C_not by metis
-next
- case (C_imp c1 c2)
- obtain s1 and s2 where 1: "eval_c i (c1[bv::=b]\<^sub>c\<^sub>b) s1 \<and> eval_c i (c2[bv::=b]\<^sub>c\<^sub>b) s2 \<and> s = (s1 \<longrightarrow> s2)" using C_imp eval_c_elims(5) subst_cb.simps(5) by metis
- obtain s1' and s2' where 2:"eval_c i' c1 s1' \<and> eval_c i' c2 s2' \<and> s' = (s1' \<longrightarrow> s2')" using C_imp eval_c_elims(5) by metis
- then show ?case using 1 2 wfC_elims C_imp by metis
-next
- case (C_eq e1 e2)
- obtain be where be: "wfCE \<Theta> B \<Gamma> e1 be \<and> wfCE \<Theta> B \<Gamma> e2 be" using C_eq wfC_elims by metis
- obtain s1 and s2 where 1: "eval_e i (e1[bv::=b]\<^sub>c\<^sub>e\<^sub>b) s1 \<and> eval_e i (e2[bv::=b]\<^sub>c\<^sub>e\<^sub>b) s2 \<and> s = (s1 = s2)" using C_eq eval_c_elims(7) subst_cb.simps(6) by metis
- obtain s1' and s2' where 2:"eval_e i' e1 s1' \<and> eval_e i' e2 s2' \<and> s' = (s1' = s2' )" using C_eq eval_c_elims(7) by metis
- have "\<turnstile>\<^sub>w\<^sub>f \<Theta>" using C_eq wfX_wfY by metis
- moreover have "\<Theta> ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> i " using C_eq by auto
- ultimately show ?case using boxed_b_eq[of \<Theta> s1 be bv b s1' s2 s2'] 1 2 boxed_i_eval_ce_boxed_b C_eq wfC_elims subst_cb.simps 1 2 be by auto
-qed
-
-lemma is_satis_bs_boxed:
- fixes c::c
- assumes "boxed_i \<Theta> \<Gamma> b bv i i'" and "wfC \<Theta> B \<Gamma> c" and "wfI \<Theta> \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b i" and "\<Theta> ; \<Gamma> \<turnstile> i'"
- and "(i \<Turnstile> c[bv::=b]\<^sub>c\<^sub>b)"
- shows "(i' \<Turnstile> c)"
-proof -
- have "eval_c i (c[bv::=b]\<^sub>c\<^sub>b) True" using is_satis.simps assms by auto
- moreover obtain s where "i' \<lbrakk> c \<rbrakk> ~ s" using eval_c_exist assms by metis
- ultimately show ?thesis using eval_c_eq_bs_boxed assms is_satis.simps by metis
-qed
-
-lemma is_satis_bs_boxed_rev:
- fixes c::c
- assumes "boxed_i \<Theta> \<Gamma> b bv i i'" and "wfC \<Theta> B \<Gamma> c" and "wfI \<Theta> \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b i" and "\<Theta> ; \<Gamma> \<turnstile> i'" and "wfC \<Theta> {||} \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b (c[bv::=b]\<^sub>c\<^sub>b)"
- and "(i' \<Turnstile> c)"
- shows "(i \<Turnstile> c[bv::=b]\<^sub>c\<^sub>b)"
-proof -
- have "eval_c i' c True" using is_satis.simps assms by auto
- moreover obtain s where "i \<lbrakk> c[bv::=b]\<^sub>c\<^sub>b \<rbrakk> ~ s" using eval_c_exist assms by metis
- ultimately show ?thesis using eval_c_eq_bs_boxed assms is_satis.simps by metis
-qed
-
-lemma bs_boxed_wfi_aux:
- fixes b::b and bv::bv and \<Theta>::\<Theta> and B::\<B>
- assumes "boxed_i \<Theta> \<Gamma> b bv i i'" and "wfI \<Theta> \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b i" and "\<turnstile>\<^sub>w\<^sub>f \<Theta>" and "wfG \<Theta> B \<Gamma>"
- shows "\<Theta> ; \<Gamma> \<turnstile> i'"
- using assms proof(induct rule: boxed_i.inducts)
- case (boxed_i_GNilI T i)
- then show ?case using wfI_def by auto
-next
- case (boxed_i_GConsI s i x1 T b1 bv b s' G i' c1)
- {
- fix x2 b2 c2
- assume as : "(x2,b2,c2) \<in> toSet ((x1, b1, c1) #\<^sub>\<Gamma> G)"
-
- then consider (hd) "(x2,b2,c2) = (x1, b1, c1)" | (tail) "(x2,b2,c2) \<in> toSet G" using toSet.simps by auto
- hence "\<exists>s. Some s = (i'(x1 \<mapsto> s')) x2 \<and> wfRCV T s b2" proof(cases)
- case hd
- hence "b1=b2" by auto
- moreover have "(x2,b2[bv::=b]\<^sub>b\<^sub>b,c2[bv::=b]\<^sub>c\<^sub>b) \<in> toSet ((x1, b1, c1) #\<^sub>\<Gamma> G)[bv::=b]\<^sub>\<Gamma>\<^sub>b" using hd subst_gb.simps by simp
- moreover hence "wfRCV T s b2[bv::=b]\<^sub>b\<^sub>b" using wfI_def boxed_i_GConsI hd
- proof -
- obtain ss :: "b \<Rightarrow> x \<Rightarrow> (x \<Rightarrow> rcl_val option) \<Rightarrow> type_def list \<Rightarrow> rcl_val" where
- "\<forall>x1a x2a x3 x4. (\<exists>v5. Some v5 = x3 x2a \<and> wfRCV x4 v5 x1a) = (Some (ss x1a x2a x3 x4) = x3 x2a \<and> wfRCV x4 (ss x1a x2a x3 x4) x1a)"
- by moura (* 0.0 ms *)
- then have f1: "Some (ss b2[bv::=b]\<^sub>b\<^sub>b x1 i T) = i x1 \<and> wfRCV T (ss b2[bv::=b]\<^sub>b\<^sub>b x1 i T) b2[bv::=b]\<^sub>b\<^sub>b"
- using boxed_i_GConsI.prems(1) hd wfI_def by auto (* 31 ms *)
- then have "ss b2[bv::=b]\<^sub>b\<^sub>b x1 i T = s"
- by (metis (no_types) boxed_i_GConsI.hyps(1) option.inject) (* 0.0 ms *)
- then show ?thesis
- using f1 by blast (* 0.0 ms *)
- qed
- ultimately have "wfRCV T s' b2" using boxed_i_GConsI boxed_b_wfRCV by metis
-
- then show ?thesis using hd by simp
- next
- case tail
- hence "wfI T G i'" using boxed_i_GConsI wfI_suffix wfG_suffix subst_gb.simps
- by (metis (no_types, lifting) Un_iff toSet.simps(2) wfG_cons2 wfI_def)
- then show ?thesis using wfI_def[of T G i'] tail
- using boxed_i_GConsI.prems(3) split_G wfG_cons_fresh2 by fastforce
- qed
- }
- thus ?case using wfI_def by fast
-
-qed
-
-lemma is_satis_g_bs_boxed_aux:
- fixes G::\<Gamma>
- assumes "boxed_i \<Theta> G1 b bv i i'" and "wfI \<Theta> G1[bv::=b]\<^sub>\<Gamma>\<^sub>b i" and "wfI \<Theta> G1 i'" and "G1 = (G2@G)" and "wfG \<Theta> B G1"
- and "(i \<Turnstile> G[bv::=b]\<^sub>\<Gamma>\<^sub>b) "
- shows "(i' \<Turnstile> G)"
- using assms proof(induct G arbitrary: G2 rule: \<Gamma>_induct)
- case GNil
- then show ?case by auto
-next
- case (GCons x' b' c' \<Gamma>' G2)
- show ?case proof(subst is_satis_g.simps,rule)
- have *:"wfC \<Theta> B G1 c'" using GCons wfG_wfC_inside by force
- show "i' \<Turnstile> c'" using is_satis_bs_boxed[OF assms(1) * ] GCons by auto
- obtain G3 where "G1 = G3 @ \<Gamma>'" using GCons append_g.simps
- by (metis append_g_assoc)
- then show "i' \<Turnstile> \<Gamma>'" using GCons append_g.simps by simp
- qed
-qed
-
-lemma is_satis_g_bs_boxed:
- fixes G::\<Gamma>
- assumes "boxed_i \<Theta> G b bv i i'" and "wfI \<Theta> G[bv::=b]\<^sub>\<Gamma>\<^sub>b i" and "wfI \<Theta> G i'" and "wfG \<Theta> B G"
- and "(i \<Turnstile> G[bv::=b]\<^sub>\<Gamma>\<^sub>b) "
- shows "(i' \<Turnstile> G)"
- using is_satis_g_bs_boxed_aux assms
- by (metis (full_types) append_g.simps(1))
-
-lemma subst_b_valid:
- fixes s::s and b::b
- assumes "\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f b" and "B = {|bv|}" and "\<Theta> ; {|bv|} ;\<Gamma> \<Turnstile> c"
- shows "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<Turnstile> c[bv::=b]\<^sub>c\<^sub>b "
-proof(rule validI)
-
- show **:"\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f c[bv::=b]\<^sub>c\<^sub>b " using assms valid.simps wf_b_subst subst_gb.simps by metis
- show "\<forall>i. (wfI \<Theta> \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b i \<and> i \<Turnstile> \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b) \<longrightarrow> i \<Turnstile> c[bv::=b]\<^sub>c\<^sub>b "
- proof(rule,rule)
- fix i
- assume *:"wfI \<Theta> \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b i \<and> i \<Turnstile> \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b"
-
- obtain i' where idash: "boxed_i \<Theta> \<Gamma> b bv i i'" using boxed_i_ex wfX_wfY assms * by fastforce
-
- have wfc: "\<Theta> ; {|bv|} ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f c" using valid.simps assms by simp
- have wfg: "\<Theta> ; {|bv|} \<turnstile>\<^sub>w\<^sub>f \<Gamma>" using valid.simps wfX_wfY assms by metis
- hence wfi: "wfI \<Theta> \<Gamma> i'" using idash * bs_boxed_wfi_aux subst_gb.simps wfX_wfY by metis
- moreover have "i' \<Turnstile> \<Gamma>" proof (rule is_satis_g_bs_boxed[OF idash ] wfX_wfY(2)[OF wfc])
- show "wfI \<Theta> \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b i" using subst_gb.simps * by simp
- show "wfI \<Theta> \<Gamma> i'" using wfi by auto
- show "\<Theta> ; B \<turnstile>\<^sub>w\<^sub>f \<Gamma> " using wfg assms by auto
- show "i \<Turnstile> \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b" using subst_gb.simps * by simp
- qed
- ultimately have ic:"i' \<Turnstile> c" using assms valid_def using valid.simps by blast
-
- show "i \<Turnstile> c[bv::=b]\<^sub>c\<^sub>b" proof(rule is_satis_bs_boxed_rev)
- show "\<Theta> ; \<Gamma> ; b , bv \<turnstile> i \<approx> i'" using idash by auto
- show "\<Theta> ; B ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f c " using wfc assms by auto
- show "\<Theta> ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> i" using subst_gb.simps * by simp
- show "\<Theta> ; \<Gamma> \<turnstile> i'" using wfi by auto
- show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f c[bv::=b]\<^sub>c\<^sub>b " using ** by auto
- show "i' \<Turnstile> c" using ic by auto
- qed
-
- qed
-qed
-
-section \<open>Expression Operator Lemmas\<close>
-
-lemma is_satis_len_imp:
- assumes "i \<Turnstile> (CE_val (V_var x) == CE_val (V_lit (L_num (int (length v)))) )" (is "is_satis i ?c1")
- shows "i \<Turnstile> (CE_val (V_var x) == CE_len [V_lit (L_bitvec v)]\<^sup>c\<^sup>e)"
-proof -
- have *:"eval_c i ?c1 True" using assms is_satis.simps by blast
- then have "eval_e i (CE_val (V_lit (L_num (int (length v))))) (SNum (int (length v)))"
- using eval_e_elims(1) eval_v_elims eval_l.simps by (metis eval_e.intros(1) eval_v_litI)
- hence "eval_e i (CE_val (V_var x)) (SNum (int (length v)))" using eval_c_elims(7)[OF *]
- by (metis eval_e_elims(1) eval_v_elims(1))
- moreover have "eval_e i (CE_len [V_lit (L_bitvec v)]\<^sup>c\<^sup>e) (SNum (int (length v)))"
- using eval_e_elims(7) eval_v_elims eval_l.simps by (metis eval_e.intros eval_v_litI)
- ultimately show ?thesis using eval_c.intros is_satis.simps by fastforce
-qed
-
-lemma is_satis_plus_imp:
- assumes "i \<Turnstile> (CE_val (V_var x) == CE_val (V_lit (L_num (n1+n2))))" (is "is_satis i ?c1")
- shows "i \<Turnstile> (CE_val (V_var x) == CE_op Plus ([V_lit (L_num n1)]\<^sup>c\<^sup>e) ([V_lit (L_num n2)]\<^sup>c\<^sup>e))"
-proof -
- have *:"eval_c i ?c1 True" using assms is_satis.simps by blast
- then have "eval_e i (CE_val (V_lit (L_num (n1+n2)))) (SNum (n1+n2))"
- using eval_e_elims(1) eval_v_elims eval_l.simps by (metis eval_e.intros(1) eval_v_litI)
- hence "eval_e i (CE_val (V_var x)) (SNum (n1+n2))" using eval_c_elims(7)[OF *]
- by (metis eval_e_elims(1) eval_v_elims(1))
- moreover have "eval_e i (CE_op Plus ([V_lit (L_num n1)]\<^sup>c\<^sup>e) ([V_lit (L_num n2)]\<^sup>c\<^sup>e)) (SNum (n1+n2))"
- using eval_e_elims(7) eval_v_elims eval_l.simps by (metis eval_e.intros eval_v_litI)
- ultimately show ?thesis using eval_c.intros is_satis.simps by fastforce
-qed
-
-lemma is_satis_leq_imp:
- assumes "i \<Turnstile> (CE_val (V_var x) == CE_val (V_lit (if (n1 \<le> n2) then L_true else L_false)))" (is "is_satis i ?c1")
- shows "i \<Turnstile> (CE_val (V_var x) == CE_op LEq [(V_lit (L_num n1))]\<^sup>c\<^sup>e [(V_lit (L_num n2))]\<^sup>c\<^sup>e)"
-proof -
- have *:"eval_c i ?c1 True" using assms is_satis.simps by blast
- then have "eval_e i (CE_val (V_lit ((if (n1 \<le> n2) then L_true else L_false)))) (SBool (n1\<le>n2))"
- using eval_e_elims(1) eval_v_elims eval_l.simps
- by (metis (full_types) eval_e.intros(1) eval_v_litI)
- hence "eval_e i (CE_val (V_var x)) (SBool (n1\<le>n2))" using eval_c_elims(7)[OF *]
- by (metis eval_e_elims(1) eval_v_elims(1))
- moreover have "eval_e i (CE_op LEq [(V_lit (L_num n1))]\<^sup>c\<^sup>e [(V_lit (L_num n2) )]\<^sup>c\<^sup>e) (SBool (n1\<le>n2))"
- using eval_e_elims(3) eval_v_elims eval_l.simps by (metis eval_e.intros eval_v_litI)
- ultimately show ?thesis using eval_c.intros is_satis.simps by fastforce
-qed
-
-lemma eval_lit_inj:
- fixes n1::l and n2::l
- assumes "\<lbrakk> n1 \<rbrakk> = s" and "\<lbrakk> n2 \<rbrakk> = s"
- shows "n1=n2"
- using assms proof(nominal_induct s rule: rcl_val.strong_induct)
- case (SBitvec x)
- then show ?case using eval_l.simps
- by (metis l.strong_exhaust rcl_val.distinct rcl_val.eq_iff)
-next
- case (SNum x)
- then show ?case using eval_l.simps
- by (metis l.strong_exhaust rcl_val.distinct rcl_val.eq_iff)
-next
- case (SBool x)
- then show ?case using eval_l.simps
- by (metis l.strong_exhaust rcl_val.distinct rcl_val.eq_iff)
-next
- case (SPair x1a x2a)
- then show ?case using eval_l.simps
- by (metis l.strong_exhaust rcl_val.distinct rcl_val.eq_iff)
-next
- case (SCons x1a x2a x3a)
- then show ?case using eval_l.simps
- by (metis l.strong_exhaust rcl_val.distinct rcl_val.eq_iff)
-next
- case (SConsp x1a x2a x3a x4)
- then show ?case using eval_l.simps
- by (metis l.strong_exhaust rcl_val.distinct rcl_val.eq_iff)
-next
- case SUnit
- then show ?case using eval_l.simps
- by (metis l.strong_exhaust rcl_val.distinct rcl_val.eq_iff)
-next
- case (SUt x)
- then show ?case using eval_l.simps
- by (metis l.strong_exhaust rcl_val.distinct rcl_val.eq_iff)
-qed
-
-lemma eval_e_lit_inj:
- fixes n1::l and n2::l
- assumes "i \<lbrakk> [ [ n1 ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrakk> ~ s" and "i \<lbrakk> [ [ n2 ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrakk> ~ s"
- shows "n1=n2"
- using eval_lit_inj assms eval_e_elims eval_v_elims by metis
-
-lemma is_satis_eq_imp:
- assumes "i \<Turnstile> (CE_val (V_var x) == CE_val (V_lit (if (n1 = n2) then L_true else L_false)))" (is "is_satis i ?c1")
- shows "i \<Turnstile> (CE_val (V_var x) == CE_op Eq [(V_lit (n1))]\<^sup>c\<^sup>e [(V_lit (n2))]\<^sup>c\<^sup>e)"
-proof -
- have *:"eval_c i ?c1 True" using assms is_satis.simps by blast
- then have "eval_e i (CE_val (V_lit ((if (n1=n2) then L_true else L_false)))) (SBool (n1=n2))"
- using eval_e_elims(1) eval_v_elims eval_l.simps
- by (metis (full_types) eval_e.intros(1) eval_v_litI)
- hence "eval_e i (CE_val (V_var x)) (SBool (n1=n2))" using eval_c_elims(7)[OF *]
- by (metis eval_e_elims(1) eval_v_elims(1))
- moreover have "eval_e i (CE_op Eq [(V_lit (n1))]\<^sup>c\<^sup>e [(V_lit (n2) )]\<^sup>c\<^sup>e) (SBool (n1=n2))"
- proof -
- obtain s1 and s2 where *:"i \<lbrakk> [ [ n1 ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrakk> ~ s1 \<and> i \<lbrakk> [ [ n2 ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrakk> ~ s2" using eval_l.simps eval_e.intros eval_v_litI by metis
- moreover have " SBool (n1 = n2) = SBool (s1 = s2)" proof(cases "n1=n2")
- case True
- then show ?thesis using *
- by (simp add: calculation eval_e_uniqueness)
- next
- case False
- then show ?thesis using * eval_e_lit_inj by auto
- qed
- ultimately show ?thesis using eval_e_eqI[of i "[(V_lit (n1))]\<^sup>c\<^sup>e" s1 "[(V_lit (n2))]\<^sup>c\<^sup>e" s2 ] by auto
- qed
- ultimately show ?thesis using eval_c.intros is_satis.simps by fastforce
-qed
-
-lemma valid_eq_e:
- assumes "\<forall>i s1 s2. wfG P \<B> GNil \<and> wfI P GNil i \<and> eval_e i e1 s1 \<and> eval_e i e2 s2 \<longrightarrow> s1 = s2"
- and "wfCE P \<B> GNil e1 b" and "wfCE P \<B> GNil e2 b"
- shows "P ; \<B> ; (x, b , CE_val (V_var x) == e1 )#\<^sub>\<Gamma> GNil \<Turnstile> CE_val (V_var x) == e2"
- unfolding valid.simps
-proof(intro conjI)
- show \<open> P ; \<B> ; (x, b, [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == e1 ) #\<^sub>\<Gamma> GNil \<turnstile>\<^sub>w\<^sub>f [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == e2 \<close>
- using assms wf_intros wfX_wfY b.eq_iff fresh_GNil wfC_e_eq2 wfV_elims by meson
- show \<open>\<forall>i. ((P ; (x, b, [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == e1 ) #\<^sub>\<Gamma> GNil \<turnstile> i) \<and> (i \<Turnstile> (x, b, [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == e1 ) #\<^sub>\<Gamma> GNil) \<longrightarrow>
- (i \<Turnstile> [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == e2)) \<close> proof(rule+)
- fix i
- assume as:"P ; (x, b, [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == e1 ) #\<^sub>\<Gamma> GNil \<turnstile> i \<and> i \<Turnstile> (x, b, [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == e1 ) #\<^sub>\<Gamma> GNil"
-
- have *: "P ; GNil \<turnstile> i " using wfI_def by auto
-
- then obtain s1 where s1:"eval_e i e1 s1" using assms eval_e_exist by metis
- obtain s2 where s2:"eval_e i e2 s2" using assms eval_e_exist * by metis
- moreover have "i x = Some s1" proof -
- have "i \<Turnstile> [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == e1" using as is_satis_g.simps by auto
- thus ?thesis using s1
- by (metis eval_c_elims(7) eval_e_elims(1) eval_e_uniqueness eval_v_elims(2) is_satis.cases)
- qed
- moreover have "s1 = s2" using s1 s2 * assms wfG_nilI wfX_wfY by metis
-
- ultimately show "i \<lbrakk> [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == e2 \<rbrakk> ~ True"
- using eval_c.intros eval_e.intros eval_v.intros
- proof -
- have "i \<lbrakk> e2 \<rbrakk> ~ s1"
- by (metis \<open>s1 = s2\<close> s2) (* 0.0 ms *)
- then show ?thesis
- by (metis (full_types) \<open>i x = Some s1\<close> eval_c_eqI eval_e_valI eval_v_varI) (* 31 ms *)
- qed
- qed
-qed
-
-lemma valid_len:
- assumes " \<turnstile>\<^sub>w\<^sub>f \<Theta>"
- shows "\<Theta> ; \<B> ; (x, B_int, [[x]\<^sup>v]\<^sup>c\<^sup>e == [[ L_num (int (length v)) ]\<^sup>v]\<^sup>c\<^sup>e) #\<^sub>\<Gamma> GNil \<Turnstile> [[x]\<^sup>v]\<^sup>c\<^sup>e == CE_len [[ L_bitvec v ]\<^sup>v]\<^sup>c\<^sup>e" (is "\<Theta> ; \<B> ; ?G \<Turnstile> ?c" )
-proof -
- have *:"\<Theta> \<turnstile>\<^sub>w\<^sub>f ([]::\<Phi>) \<and> \<Theta> ; \<B> ; GNil \<turnstile>\<^sub>w\<^sub>f []\<^sub>\<Delta> " using assms wfG_nilI wfD_emptyI wfPhi_emptyI by auto
-
- moreover hence "\<Theta> ; \<B> ; GNil \<turnstile>\<^sub>w\<^sub>f CE_val (V_lit (L_num (int (length v)))) : B_int"
- using wfCE_valI * wfV_litI base_for_lit.simps
- by (metis wfE_valI wfX_wfY)
-
- moreover have "\<Theta> ; \<B> ; GNil \<turnstile>\<^sub>w\<^sub>f CE_len [(V_lit (L_bitvec v))]\<^sup>c\<^sup>e : B_int"
- using wfE_valI * wfV_litI base_for_lit.simps wfE_valI wfX_wfY wfCE_valI
- by (metis wfCE_lenI)
- moreover have "atom x \<sharp> GNil" by auto
- ultimately have "\<Theta> ; \<B> ; ?G \<turnstile>\<^sub>w\<^sub>f ?c" using wfC_e_eq2 assms by simp
- moreover have "(\<forall>i. wfI \<Theta> ?G i \<and> is_satis_g i ?G \<longrightarrow> is_satis i ?c)" using is_satis_len_imp by auto
- ultimately show ?thesis using valid.simps by auto
-qed
-
-lemma valid_arith_bop:
- assumes "wfG \<Theta> \<B> \<Gamma>" and "opp = Plus \<and> ll = (L_num (n1+n2)) \<or> (opp = LEq \<and> ll = ( if n1\<le>n2 then L_true else L_false))"
- and "(opp = Plus \<longrightarrow> b = B_int) \<and> (opp = LEq \<longrightarrow> b = B_bool)" and
- "atom x \<sharp> \<Gamma>"
- shows "\<Theta>; \<B> ; (x, b, (CE_val (V_var x) == CE_val (V_lit (ll)) )) #\<^sub>\<Gamma> \<Gamma>
- \<Turnstile> (CE_val (V_var x) == CE_op opp ([V_lit (L_num n1)]\<^sup>c\<^sup>e) ([V_lit (L_num n2)]\<^sup>c\<^sup>e ))" (is "\<Theta> ; \<B> ; ?G \<Turnstile> ?c")
-proof -
- have "wfC \<Theta> \<B> ?G ?c" proof(rule wfC_e_eq2)
- show "\<Theta> ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_val (V_lit ll) : b" using wfCE_valI wfV_litI assms base_for_lit.simps by metis
- show "\<Theta> ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_op opp ([V_lit (L_num n1)]\<^sup>c\<^sup>e) ([V_lit (L_num n2)]\<^sup>c\<^sup>e) : b "
- using wfCE_plusI wfCE_leqI wfCE_eqI wfV_litI wfCE_valI base_for_lit.simps assms by metis
- show "\<turnstile>\<^sub>w\<^sub>f \<Theta>" using assms wfX_wfY by auto
- show "atom x \<sharp> \<Gamma>" using assms by auto
- qed
-
- moreover have "\<forall>i. wfI \<Theta> ?G i \<and> is_satis_g i ?G \<longrightarrow> is_satis i ?c" proof(rule allI , rule impI)
- fix i
- assume "wfI \<Theta> ?G i \<and> is_satis_g i ?G"
-
- hence "is_satis i ((CE_val (V_var x) == CE_val (V_lit (ll)) ))" by auto
- thus "is_satis i ((CE_val (V_var x) == CE_op opp ([V_lit (L_num n1)]\<^sup>c\<^sup>e) ([V_lit (L_num n2)]\<^sup>c\<^sup>e)))"
- using is_satis_plus_imp assms opp.exhaust is_satis_leq_imp by auto
- qed
- ultimately show ?thesis using valid.simps by metis
-qed
-
-lemma valid_eq_bop:
- assumes "wfG \<Theta> \<B> \<Gamma>" and "atom x \<sharp> \<Gamma>" and "base_for_lit l1 = base_for_lit l2"
- shows "\<Theta>; \<B> ; (x, B_bool, (CE_val (V_var x) == CE_val (V_lit (if l1 = l2 then L_true else L_false)) )) #\<^sub>\<Gamma> \<Gamma>
- \<Turnstile> (CE_val (V_var x) == CE_op Eq ([V_lit (l1)]\<^sup>c\<^sup>e) ([V_lit (l2)]\<^sup>c\<^sup>e ))" (is "\<Theta> ; \<B> ; ?G \<Turnstile> ?c")
-proof -
- let ?ll = "(if l1 = l2 then L_true else L_false)"
- have "wfC \<Theta> \<B> ?G ?c" proof(rule wfC_e_eq2)
- show "\<Theta> ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_val (V_lit ?ll) : B_bool" using wfCE_valI wfV_litI assms base_for_lit.simps by metis
- show "\<Theta> ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_op Eq ([V_lit (l1)]\<^sup>c\<^sup>e) ([V_lit (l2)]\<^sup>c\<^sup>e) : B_bool "
- using wfCE_eqI wfCE_leqI wfCE_eqI wfV_litI wfCE_valI base_for_lit.simps assms by metis
- show "\<turnstile>\<^sub>w\<^sub>f \<Theta>" using assms wfX_wfY by auto
- show "atom x \<sharp> \<Gamma>" using assms by auto
- qed
-
- moreover have "\<forall>i. wfI \<Theta> ?G i \<and> is_satis_g i ?G \<longrightarrow> is_satis i ?c" proof(rule allI , rule impI)
- fix i
- assume "wfI \<Theta> ?G i \<and> is_satis_g i ?G"
-
- hence "is_satis i ((CE_val (V_var x) == CE_val (V_lit (?ll)) ))" by auto
- thus "is_satis i ((CE_val (V_var x) == CE_op Eq ([V_lit (l1)]\<^sup>c\<^sup>e) ([V_lit (l2)]\<^sup>c\<^sup>e)))"
- using is_satis_eq_imp assms by auto
- qed
- ultimately show ?thesis using valid.simps by metis
-qed
-
-lemma valid_fst:
- fixes x::x and v\<^sub>1::v and v\<^sub>2::v
- assumes "wfTh \<Theta>" and "wfV \<Theta> \<B> GNil (V_pair v\<^sub>1 v\<^sub>2) (B_pair b\<^sub>1 b\<^sub>2)"
- shows "\<Theta> ; \<B> ; (x, b\<^sub>1, [[x]\<^sup>v]\<^sup>c\<^sup>e == [v\<^sub>1]\<^sup>c\<^sup>e) #\<^sub>\<Gamma> GNil \<Turnstile> [[x]\<^sup>v]\<^sup>c\<^sup>e == [#1[[v\<^sub>1,v\<^sub>2]\<^sup>v]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e"
-proof(rule valid_eq_e)
- show \<open>\<forall>i s1 s2. (\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f GNil) \<and> (\<Theta> ; GNil \<turnstile> i) \<and> (i \<lbrakk> [ v\<^sub>1 ]\<^sup>c\<^sup>e \<rbrakk> ~ s1) \<and> (i \<lbrakk> [#1[[ v\<^sub>1 , v\<^sub>2 ]\<^sup>v]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e \<rbrakk> ~ s2) \<longrightarrow> s1 = s2\<close>
- proof(rule+)
- fix i s1 s2
- assume as:"\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f GNil \<and> \<Theta> ; GNil \<turnstile> i \<and> (i \<lbrakk> [ v\<^sub>1 ]\<^sup>c\<^sup>e \<rbrakk> ~ s1) \<and> (i \<lbrakk> [#1[[ v\<^sub>1 , v\<^sub>2 ]\<^sup>v]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e \<rbrakk> ~ s2)"
- then obtain s2' where *:"i \<lbrakk> [ v\<^sub>1 , v\<^sub>2 ]\<^sup>v \<rbrakk> ~ SPair s2 s2'"
- using eval_e_elims(5)[of i "[[ v\<^sub>1 , v\<^sub>2 ]\<^sup>v]\<^sup>c\<^sup>e" s2] eval_e_elims
- by meson
- then have " i \<lbrakk> v\<^sub>1 \<rbrakk> ~ s2" using eval_v_elims(3)[OF *] by auto
- then show "s1 = s2" using eval_v_uniqueness as
- using eval_e_uniqueness eval_e_valI by blast
- qed
-
- show \<open> \<Theta> ; \<B> ; GNil \<turnstile>\<^sub>w\<^sub>f [ v\<^sub>1 ]\<^sup>c\<^sup>e : b\<^sub>1 \<close> using assms
- by (metis b.eq_iff(4) wfV_elims(3) wfV_wfCE)
- show \<open> \<Theta> ; \<B> ; GNil \<turnstile>\<^sub>w\<^sub>f [#1[[ v\<^sub>1 , v\<^sub>2 ]\<^sup>v]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e : b\<^sub>1 \<close> using assms using wfCE_fstI
- using wfCE_valI by blast
-qed
-
-lemma valid_snd:
- fixes x::x and v\<^sub>1::v and v\<^sub>2::v
- assumes "wfTh \<Theta>" and "wfV \<Theta> \<B> GNil (V_pair v\<^sub>1 v\<^sub>2) (B_pair b\<^sub>1 b\<^sub>2)"
- shows "\<Theta> ; \<B> ; (x, b\<^sub>2, [[x]\<^sup>v]\<^sup>c\<^sup>e == [v\<^sub>2]\<^sup>c\<^sup>e) #\<^sub>\<Gamma> GNil \<Turnstile> [[x]\<^sup>v]\<^sup>c\<^sup>e == [#2[[v\<^sub>1,v\<^sub>2]\<^sup>v]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e"
-proof(rule valid_eq_e)
- show \<open>\<forall>i s1 s2. (\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f GNil) \<and> (\<Theta> ; GNil \<turnstile> i) \<and> (i \<lbrakk> [ v\<^sub>2 ]\<^sup>c\<^sup>e \<rbrakk> ~ s1) \<and>
-(i \<lbrakk> [#2[[ v\<^sub>1 , v\<^sub>2 ]\<^sup>v]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e \<rbrakk> ~ s2) \<longrightarrow> s1 = s2\<close>
- proof(rule+)
- fix i s1 s2
- assume as:"\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f GNil \<and> \<Theta> ; GNil \<turnstile> i \<and> (i \<lbrakk> [ v\<^sub>2 ]\<^sup>c\<^sup>e \<rbrakk> ~ s1) \<and> (i \<lbrakk> [#2[[ v\<^sub>1 , v\<^sub>2 ]\<^sup>v]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e \<rbrakk> ~ s2)"
- then obtain s2' where *:"i \<lbrakk> [ v\<^sub>1 , v\<^sub>2 ]\<^sup>v \<rbrakk> ~ SPair s2' s2"
- using eval_e_elims(5)[of i "[[ v\<^sub>1 , v\<^sub>2 ]\<^sup>v]\<^sup>c\<^sup>e" s2] eval_e_elims
- by meson
- then have " i \<lbrakk> v\<^sub>2 \<rbrakk> ~ s2" using eval_v_elims(3)[OF *] by auto
- then show "s1 = s2" using eval_v_uniqueness as
- using eval_e_uniqueness eval_e_valI by blast
- qed
-
- show \<open> \<Theta> ; \<B> ; GNil \<turnstile>\<^sub>w\<^sub>f [ v\<^sub>2 ]\<^sup>c\<^sup>e : b\<^sub>2 \<close> using assms
- by (metis b.eq_iff wfV_elims wfV_wfCE)
- show \<open> \<Theta> ; \<B> ; GNil \<turnstile>\<^sub>w\<^sub>f [#2[[ v\<^sub>1 , v\<^sub>2 ]\<^sup>v]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e : b\<^sub>2 \<close> using assms using wfCE_sndI wfCE_valI by blast
-qed
-
-lemma valid_concat:
- fixes v1::"bit list" and v2::"bit list"
- assumes " \<turnstile>\<^sub>w\<^sub>f \<Pi>"
- shows "\<Pi> ; \<B> ; (x, B_bitvec, (CE_val (V_var x) == CE_val (V_lit (L_bitvec (v1@ v2))))) #\<^sub>\<Gamma> GNil \<Turnstile>
- (CE_val (V_var x) == CE_concat ([V_lit (L_bitvec v1)]\<^sup>c\<^sup>e ) ([V_lit (L_bitvec v2)]\<^sup>c\<^sup>e) )"
-proof(rule valid_eq_e)
- show \<open>\<forall>i s1 s2. ((\<Pi> ; \<B> \<turnstile>\<^sub>w\<^sub>f GNil) \<and> (\<Pi> ; GNil \<turnstile> i) \<and>
- (i \<lbrakk> [ [ L_bitvec (v1 @ v2) ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrakk> ~ s1) \<and> (i \<lbrakk> [[[ L_bitvec v1 ]\<^sup>v]\<^sup>c\<^sup>e @@ [[ L_bitvec v2 ]\<^sup>v]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e \<rbrakk> ~ s2) \<longrightarrow>
- s1 = s2)\<close>
- proof(rule+)
- fix i s1 s2
- assume as: "(\<Pi> ; \<B> \<turnstile>\<^sub>w\<^sub>f GNil) \<and> (\<Pi> ; GNil \<turnstile> i) \<and> (i \<lbrakk> [ [ L_bitvec (v1 @ v2) ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrakk> ~ s1) \<and>
- (i \<lbrakk> [[[ L_bitvec v1 ]\<^sup>v]\<^sup>c\<^sup>e @@ [[ L_bitvec v2 ]\<^sup>v]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e \<rbrakk> ~ s2) "
-
- hence *: "i \<lbrakk> [[[ L_bitvec v1 ]\<^sup>v]\<^sup>c\<^sup>e @@ [[ L_bitvec v2 ]\<^sup>v]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e \<rbrakk> ~ s2" by auto
- obtain bv1 bv2 where s2:"s2 = SBitvec (bv1 @ bv2) \<and> i \<lbrakk> [ L_bitvec v1 ]\<^sup>v \<rbrakk> ~ SBitvec bv1 \<and> (i \<lbrakk> [ L_bitvec v2 ]\<^sup>v \<rbrakk> ~ SBitvec bv2)"
- using eval_e_elims(7)[OF *] eval_e_elims(1) by metis
- hence "v1 = bv1 \<and> v2 = bv2" using eval_v_elims(1) eval_l.simps(5) by force
- moreover then have "s1 = SBitvec (bv1 @ bv2)" using s2 using eval_v_elims(1) eval_l.simps(5)
- by (metis as eval_e_elims(1))
-
- then show "s1 = s2" using s2 by auto
- qed
-
- show \<open> \<Pi> ; \<B> ; GNil \<turnstile>\<^sub>w\<^sub>f [ [ L_bitvec (v1 @ v2) ]\<^sup>v ]\<^sup>c\<^sup>e : B_bitvec \<close>
- by (metis assms base_for_lit.simps(5) wfG_nilI wfV_litI wfV_wfCE)
- show \<open> \<Pi> ; \<B> ; GNil \<turnstile>\<^sub>w\<^sub>f [[[ L_bitvec v1 ]\<^sup>v]\<^sup>c\<^sup>e @@ [[ L_bitvec v2 ]\<^sup>v]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e : B_bitvec \<close>
- by (metis assms base_for_lit.simps(5) wfCE_concatI wfG_nilI wfV_litI wfCE_valI)
-qed
-
-lemma valid_ce_eq:
- fixes ce::ce
- assumes "\<Theta> ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ce : b"
- shows \<open>\<Theta> ; \<B> ; \<Gamma> \<Turnstile> ce == ce \<close>
- unfolding valid.simps proof
- show \<open> \<Theta> ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ce == ce \<close> using assms wfC_eqI by auto
- show \<open>\<forall>i. \<Theta> ; \<Gamma> \<turnstile> i \<and> i \<Turnstile> \<Gamma> \<longrightarrow> i \<Turnstile> ce == ce \<close> proof(rule+)
- fix i
- assume "\<Theta> ; \<Gamma> \<turnstile> i \<and> i \<Turnstile> \<Gamma>"
- then obtain s where "i\<lbrakk> ce \<rbrakk> ~ s" using assms eval_e_exist by metis
- then show "i \<lbrakk> ce == ce \<rbrakk> ~ True " using eval_c_eqI by metis
- qed
-qed
-
-lemma valid_eq_imp:
- fixes c1::c and c2::c
- assumes " \<Theta> ; \<B> ; (x, b, c2) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f c1 IMP c2"
- shows " \<Theta> ; \<B> ; (x, b, c2) #\<^sub>\<Gamma> \<Gamma> \<Turnstile> c1 IMP c2 "
-proof -
- have "\<forall>i. (\<Theta> ; (x, b, c2) #\<^sub>\<Gamma> \<Gamma> \<turnstile> i \<and> i \<Turnstile> (x, b, c2) #\<^sub>\<Gamma> \<Gamma>) \<longrightarrow> i \<Turnstile> ( c1 IMP c2 )"
- proof(rule,rule)
- fix i
- assume as:"\<Theta> ; (x, b, c2) #\<^sub>\<Gamma> \<Gamma> \<turnstile> i \<and> i \<Turnstile> (x, b, c2) #\<^sub>\<Gamma> \<Gamma>"
-
- have "\<Theta> ; \<B> ; (x, b, c2) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f c1" using wfC_elims assms by metis
-
- then obtain sc where "i \<lbrakk> c1 \<rbrakk> ~ sc" using eval_c_exist assms as by metis
- moreover have "i \<lbrakk> c2 \<rbrakk> ~ True" using as is_satis_g.simps is_satis.simps by auto
-
- ultimately have "i \<lbrakk> c1 IMP c2 \<rbrakk> ~ True" using eval_c_impI by metis
-
- thus "i \<Turnstile> c1 IMP c2" using is_satis.simps by auto
- qed
- thus ?thesis using assms by auto
-qed
-
-lemma valid_range:
- assumes "0 \<le> n \<and> n \<le> m" and "\<turnstile>\<^sub>w\<^sub>f \<Theta>"
- shows "\<Theta> ; {||} ; (x, B_int , (C_eq (CE_val (V_var x)) (CE_val (V_lit (L_num n))))) #\<^sub>\<Gamma> GNil \<Turnstile>
- (C_eq (CE_op LEq (CE_val (V_var x)) (CE_val (V_lit (L_num m)))) [[ L_true ]\<^sup>v ]\<^sup>c\<^sup>e) AND
- (C_eq (CE_op LEq (CE_val (V_lit (L_num 0))) (CE_val (V_var x))) [[ L_true ]\<^sup>v ]\<^sup>c\<^sup>e)"
- (is "\<Theta> ; {||} ; ?G \<Turnstile> ?c1 AND ?c2")
-proof(rule validI)
- have wfg: " \<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f (x, B_int, [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_num n ]\<^sup>v ]\<^sup>c\<^sup>e ) #\<^sub>\<Gamma> GNil "
- using assms base_for_lit.simps wfG_nilI wfV_litI fresh_GNil wfB_intI wfC_v_eq wfG_cons1I wfG_cons2I by metis
-
- show "\<Theta> ; {||} ; ?G \<turnstile>\<^sub>w\<^sub>f ?c1 AND ?c2"
- using wfC_conjI wfC_eqI wfCE_leqI wfCE_valI wfV_varI wfg lookup.simps base_for_lit.simps wfV_litI wfB_intI wfB_boolI
- by metis
-
- show "\<forall>i. \<Theta> ; ?G \<turnstile> i \<and> i \<Turnstile> ?G \<longrightarrow> i \<Turnstile> ?c1 AND ?c2" proof(rule,rule)
- fix i
- assume a:"\<Theta> ; ?G \<turnstile> i \<and> i \<Turnstile> ?G"
- hence *:"i \<lbrakk> V_var x \<rbrakk> ~ SNum n"
- proof -
- obtain sv where sv: "i x = Some sv \<and> \<Theta> \<turnstile> sv : B_int" using a wfI_def by force
- have "i \<lbrakk> (C_eq (CE_val (V_var x)) (CE_val (V_lit (L_num n)))) \<rbrakk> ~ True"
- using a is_satis_g.simps
- using is_satis.cases by blast
- hence "i x = Some(SNum n)" using sv
- by (metis eval_c_elims(7) eval_e_elims(1) eval_l.simps(3) eval_v_elims(1) eval_v_elims(2))
- thus ?thesis using eval_v_varI by auto
- qed
-
- show "i \<Turnstile> ?c1 AND ?c2"
- proof -
- have "i \<lbrakk> ?c1 \<rbrakk> ~ True"
- proof -
- have "i \<lbrakk> [ leq [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e [ [ L_num m ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e\<rbrakk> ~ SBool True"
- using eval_e_leqI assms eval_v_litI eval_l.simps *
- by (metis (full_types) eval_e_valI)
- moreover have "i \<lbrakk> [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrakk> ~ SBool True"
- using eval_v_litI eval_e_valI eval_l.simps by metis
- ultimately show ?thesis using eval_c_eqI by metis
- qed
-
- moreover have "i \<lbrakk> ?c2 \<rbrakk> ~ True"
- proof -
- have "i \<lbrakk> [ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e \<rbrakk> ~ SBool True"
- using eval_e_leqI assms eval_v_litI eval_l.simps *
- by (metis (full_types) eval_e_valI)
- moreover have "i \<lbrakk> [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrakk> ~ SBool True"
- using eval_v_litI eval_e_valI eval_l.simps by metis
- ultimately show ?thesis using eval_c_eqI by metis
- qed
- ultimately show ?thesis using eval_c_conjI is_satis.simps by metis
- qed
- qed
-qed
-
-lemma valid_range_length:
- fixes \<Gamma>::\<Gamma>
- assumes "0 \<le> n \<and> n \<le> int (length v)" and " \<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f \<Gamma>" and "atom x \<sharp> \<Gamma>"
- shows "\<Theta> ; {||} ; (x, B_int , (C_eq (CE_val (V_var x)) (CE_val (V_lit (L_num n))))) #\<^sub>\<Gamma> \<Gamma> \<Turnstile>
- (C_eq (CE_op LEq (CE_val (V_lit (L_num 0))) (CE_val (V_var x))) [[ L_true ]\<^sup>v ]\<^sup>c\<^sup>e) AND
- (C_eq (CE_op LEq (CE_val (V_var x)) ([| [ [ L_bitvec v ]\<^sup>v ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e )) [[ L_true ]\<^sup>v ]\<^sup>c\<^sup>e)
- "
- (is "\<Theta> ; {||} ; ?G \<Turnstile> ?c1 AND ?c2")
-proof(rule validI)
- have wfg: " \<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f (x, B_int, [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_num n ]\<^sup>v ]\<^sup>c\<^sup>e ) #\<^sub>\<Gamma> \<Gamma> " apply(rule wfG_cons1I)
- apply simp
- using assms apply simp+
- using assms base_for_lit.simps wfG_nilI wfV_litI wfB_intI wfC_v_eq wfB_intI wfX_wfY assms by metis+
-
- show "\<Theta> ; {||} ; ?G \<turnstile>\<^sub>w\<^sub>f ?c1 AND ?c2"
- using wfC_conjI wfC_eqI wfCE_leqI wfCE_valI wfV_varI wfg lookup.simps base_for_lit.simps wfV_litI wfB_intI wfB_boolI
- by (metis (full_types) wfCE_lenI)
-
- show "\<forall>i. \<Theta> ; ?G \<turnstile> i \<and> i \<Turnstile> ?G \<longrightarrow> i \<Turnstile> ?c1 AND ?c2" proof(rule,rule)
- fix i
- assume a:"\<Theta> ; ?G \<turnstile> i \<and> i \<Turnstile> ?G"
- hence *:"i \<lbrakk> V_var x \<rbrakk> ~ SNum n"
- proof -
- obtain sv where sv: "i x = Some sv \<and> \<Theta> \<turnstile> sv : B_int" using a wfI_def by force
- have "i \<lbrakk> (C_eq (CE_val (V_var x)) (CE_val (V_lit (L_num n)))) \<rbrakk> ~ True"
- using a is_satis_g.simps
- using is_satis.cases by blast
- hence "i x = Some(SNum n)" using sv
- by (metis eval_c_elims(7) eval_e_elims(1) eval_l.simps(3) eval_v_elims(1) eval_v_elims(2))
- thus ?thesis using eval_v_varI by auto
- qed
-
- show "i \<Turnstile> ?c1 AND ?c2"
- proof -
- have "i \<lbrakk> ?c2 \<rbrakk> ~ True"
- proof -
- have "i \<lbrakk> [ leq [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e [| [ [ L_bitvec v ]\<^sup>v ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e\<rbrakk> ~ SBool True"
- using eval_e_leqI assms eval_v_litI eval_l.simps *
- by (metis (full_types) eval_e_lenI eval_e_valI)
- moreover have "i \<lbrakk> [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrakk> ~ SBool True"
- using eval_v_litI eval_e_valI eval_l.simps by metis
- ultimately show ?thesis using eval_c_eqI by metis
- qed
-
- moreover have "i \<lbrakk> ?c1 \<rbrakk> ~ True"
- proof -
- have "i \<lbrakk> [ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e \<rbrakk> ~ SBool True"
- using eval_e_leqI assms eval_v_litI eval_l.simps *
- by (metis (full_types) eval_e_valI)
- moreover have "i \<lbrakk> [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrakk> ~ SBool True"
- using eval_v_litI eval_e_valI eval_l.simps by metis
- ultimately show ?thesis using eval_c_eqI by metis
- qed
- ultimately show ?thesis using eval_c_conjI is_satis.simps by metis
- qed
- qed
-qed
-
-lemma valid_range_length_inv_gnil:
- fixes \<Gamma>::\<Gamma>
- assumes "\<turnstile>\<^sub>w\<^sub>f \<Theta> "
- and "\<Theta> ; {||} ; (x, B_int , (C_eq (CE_val (V_var x)) (CE_val (V_lit (L_num n))))) #\<^sub>\<Gamma> GNil \<Turnstile>
- (C_eq (CE_op LEq (CE_val (V_lit (L_num 0))) (CE_val (V_var x))) [[ L_true ]\<^sup>v ]\<^sup>c\<^sup>e) AND
- (C_eq (CE_op LEq (CE_val (V_var x)) ([| [ [ L_bitvec v ]\<^sup>v ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e )) [[ L_true ]\<^sup>v ]\<^sup>c\<^sup>e)
- "
- (is "\<Theta> ; {||} ; ?G \<Turnstile> ?c1 AND ?c2")
- shows "0 \<le> n \<and> n \<le> int (length v)"
-proof -
- have *:"\<forall>i. \<Theta> ; ?G \<turnstile> i \<and> i \<Turnstile> ?G \<longrightarrow> i \<Turnstile> ?c1 AND ?c2" using assms valid.simps by simp
-
- obtain i where i: "i x = Some (SNum n)" by auto
- have "\<Theta> ; ?G \<turnstile> i \<and> i \<Turnstile> ?G" proof
- show "\<Theta> ; ?G \<turnstile> i" unfolding wfI_def using wfRCV_BIntI i * by auto
- have "i \<lbrakk> ([ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_num n ]\<^sup>v ]\<^sup>c\<^sup>e ) \<rbrakk> ~ True"
- using * eval_c.intros(7) eval_e.intros eval_v.intros eval_l.simps
- by (metis (full_types) i)
- thus "i \<Turnstile> ?G" unfolding is_satis_g.simps is_satis.simps by auto
- qed
- hence **:"i \<Turnstile> ?c1 AND ?c2" using * by auto
-
- hence 1: "i \<lbrakk> ?c1 \<rbrakk> ~ True" using eval_c_elims(3) is_satis.simps
- by fastforce
- then obtain sv1 and sv2 where "(sv1 = sv2) = True \<and> i \<lbrakk> [ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e \<rbrakk> ~ sv1 \<and> i \<lbrakk> [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrakk> ~ sv2"
- using eval_c_elims(7) by metis
- hence "sv1 = SBool True" using eval_e_elims eval_v_elims eval_l.simps i by metis
- obtain n1 and n2 where "SBool True = SBool (n1 \<le> n2) \<and> (i \<lbrakk> [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrakk> ~ SNum n1) \<and> (i \<lbrakk> [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrakk> ~ SNum n2)"
- using eval_e_elims(3)[of i " [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e" "[ [ x ]\<^sup>v ]\<^sup>c\<^sup>e " "SBool True"]
- using \<open>(sv1 = sv2) = True \<and> i \<lbrakk> [ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e \<rbrakk> ~ sv1 \<and> i \<lbrakk> [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrakk> ~ sv2\<close> \<open>sv1 = SBool True\<close> by fastforce
- moreover hence "n1 = 0" and "n2 = n" using eval_e_elims eval_v_elims i
- apply (metis eval_l.simps(3) rcl_val.eq_iff(2))
- using eval_e_elims eval_v_elims i
- by (metis calculation option.inject rcl_val.eq_iff(2))
- ultimately have le1: "0 \<le> n " by simp
-
- hence 2: "i \<lbrakk> ?c2 \<rbrakk> ~ True" using ** eval_c_elims(3) is_satis.simps
- by fastforce
- then obtain sv1 and sv2 where sv: "(sv1 = sv2) = True \<and> i \<lbrakk> [ leq [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e [| [ [ L_bitvec v ]\<^sup>v ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e \<rbrakk> ~ sv1 \<and> i \<lbrakk> [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrakk> ~ sv2"
- using eval_c_elims(7) by metis
- hence "sv1 = SBool True" using eval_e_elims eval_v_elims eval_l.simps i by metis
- obtain n1 and n2 where ***:"SBool True = SBool (n1 \<le> n2) \<and> (i \<lbrakk> [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrakk> ~ SNum n1) \<and> (i \<lbrakk> [| [ [ L_bitvec v ]\<^sup>v ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e \<rbrakk> ~ SNum n2)"
- using eval_e_elims(3)
- using sv \<open>sv1 = SBool True\<close> by metis
- moreover hence "n1 = n" using eval_e_elims(1)[of i] eval_v_elims(2)[of i x "SNum n1"] i by auto
- moreover have "n2 = int (length v)" using eval_e_elims(8) eval_v_elims(1) eval_l.simps i
- by (metis "***" eval_e_elims(1) rcl_val.eq_iff(1) rcl_val.eq_iff(2))
- ultimately have le2: "n \<le> int (length v) " by simp
-
- show ?thesis using le1 le2 by auto
-qed
-
-lemma wfI_cons:
- fixes i::valuation and \<Gamma>::\<Gamma>
- assumes "i' \<Turnstile> \<Gamma>" and "\<Theta> ; \<Gamma> \<turnstile> i'" and "i = i' ( x \<mapsto> s)" and "\<Theta> \<turnstile> s : b" and "atom x \<sharp> \<Gamma>"
- shows "\<Theta> ; (x,b,c) #\<^sub>\<Gamma> \<Gamma> \<turnstile> i"
- unfolding wfI_def proof -
- {
- fix x' b' c'
- assume "(x',b',c') \<in> toSet ((x, b, c) #\<^sub>\<Gamma> \<Gamma>)"
- then consider "(x',b',c') = (x,b,c)" | "(x',b',c') \<in> toSet \<Gamma>" using toSet.simps by auto
- then have "\<exists>s. Some s = i x' \<and> \<Theta> \<turnstile> s : b'" proof(cases)
- case 1
- then show ?thesis using assms by auto
- next
- case 2
- then obtain s where s:"Some s = i' x' \<and> \<Theta> \<turnstile> s : b'" using assms wfI_def by auto
- moreover have "x' \<noteq> x" using assms 2 fresh_dom_free by auto
- ultimately have "Some s = i x'" using assms by auto
- then show ?thesis using s wfI_def by auto
- qed
- }
- thus "\<forall>(x, b, c)\<in>toSet ((x, b, c) #\<^sub>\<Gamma> \<Gamma>). \<exists>s. Some s = i x \<and> \<Theta> \<turnstile> s : b" by auto
-qed
-
-lemma valid_range_length_inv:
- fixes \<Gamma>::\<Gamma>
- assumes "\<Theta> ; B \<turnstile>\<^sub>w\<^sub>f \<Gamma> " and "atom x \<sharp> \<Gamma>" and "\<exists>i. i \<Turnstile> \<Gamma> \<and> \<Theta> ; \<Gamma> \<turnstile> i"
- and "\<Theta> ; B ; (x, B_int , (C_eq (CE_val (V_var x)) (CE_val (V_lit (L_num n))))) #\<^sub>\<Gamma> \<Gamma> \<Turnstile>
- (C_eq (CE_op LEq (CE_val (V_lit (L_num 0))) (CE_val (V_var x))) [[ L_true ]\<^sup>v ]\<^sup>c\<^sup>e) AND
- (C_eq (CE_op LEq (CE_val (V_var x)) ([| [ [ L_bitvec v ]\<^sup>v ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e )) [[ L_true ]\<^sup>v ]\<^sup>c\<^sup>e)
- "
- (is "\<Theta> ; ?B ; ?G \<Turnstile> ?c1 AND ?c2")
- shows "0 \<le> n \<and> n \<le> int (length v)"
-proof -
- have *:"\<forall>i. \<Theta> ; ?G \<turnstile> i \<and> i \<Turnstile> ?G \<longrightarrow> i \<Turnstile> ?c1 AND ?c2" using assms valid.simps by simp
-
- obtain i' where idash: "is_satis_g i' \<Gamma> \<and> \<Theta> ; \<Gamma> \<turnstile> i'" using assms by auto
- obtain i where i: "i = i' ( x \<mapsto> SNum n)" by auto
- hence ix: "i x = Some (SNum n)" by auto
- have "\<Theta> ; ?G \<turnstile> i \<and> i \<Turnstile> ?G" proof
- show "\<Theta> ; ?G \<turnstile> i" using wfI_cons i idash ix wfRCV_BIntI assms by simp
-
- have **:"i \<lbrakk> ([ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_num n ]\<^sup>v ]\<^sup>c\<^sup>e ) \<rbrakk> ~ True"
- using * eval_c.intros(7) eval_e.intros eval_v.intros eval_l.simps i
- by (metis (full_types) ix)
-
- show "i \<Turnstile> ?G" unfolding is_satis_g.simps proof
- show \<open> i \<Turnstile> [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_num n ]\<^sup>v ]\<^sup>c\<^sup>e \<close> using ** is_satis.simps by auto
- show \<open> i \<Turnstile> \<Gamma> \<close> using idash i assms is_satis_g_i_upd by metis
- qed
- qed
- hence **:"i \<Turnstile> ?c1 AND ?c2" using * by auto
-
- hence 1: "i \<lbrakk> ?c1 \<rbrakk> ~ True" using eval_c_elims(3) is_satis.simps
- by fastforce
- then obtain sv1 and sv2 where "(sv1 = sv2) = True \<and> i \<lbrakk> [ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e \<rbrakk> ~ sv1 \<and> i \<lbrakk> [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrakk> ~ sv2"
- using eval_c_elims(7) by metis
- hence "sv1 = SBool True" using eval_e_elims eval_v_elims eval_l.simps i by metis
- obtain n1 and n2 where "SBool True = SBool (n1 \<le> n2) \<and> (i \<lbrakk> [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrakk> ~ SNum n1) \<and> (i \<lbrakk> [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrakk> ~ SNum n2)"
- using eval_e_elims(3)[of i " [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e" "[ [ x ]\<^sup>v ]\<^sup>c\<^sup>e " "SBool True"]
- using \<open>(sv1 = sv2) = True \<and> i \<lbrakk> [ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e \<rbrakk> ~ sv1 \<and> i \<lbrakk> [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrakk> ~ sv2\<close> \<open>sv1 = SBool True\<close> by fastforce
- moreover hence "n1 = 0" and "n2 = n" using eval_e_elims eval_v_elims i
- apply (metis eval_l.simps(3) rcl_val.eq_iff(2))
- using eval_e_elims eval_v_elims i
- calculation option.inject rcl_val.eq_iff(2)
- by (metis ix)
- ultimately have le1: "0 \<le> n " by simp
-
- hence 2: "i \<lbrakk> ?c2 \<rbrakk> ~ True" using ** eval_c_elims(3) is_satis.simps
- by fastforce
- then obtain sv1 and sv2 where sv: "(sv1 = sv2) = True \<and> i \<lbrakk> [ leq [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e [| [ [ L_bitvec v ]\<^sup>v ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e \<rbrakk> ~ sv1 \<and> i \<lbrakk> [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrakk> ~ sv2"
- using eval_c_elims(7) by metis
- hence "sv1 = SBool True" using eval_e_elims eval_v_elims eval_l.simps i by metis
- obtain n1 and n2 where ***:"SBool True = SBool (n1 \<le> n2) \<and> (i \<lbrakk> [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrakk> ~ SNum n1) \<and> (i \<lbrakk> [| [ [ L_bitvec v ]\<^sup>v ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e \<rbrakk> ~ SNum n2)"
- using eval_e_elims(3)
- using sv \<open>sv1 = SBool True\<close> by metis
- moreover hence "n1 = n" using eval_e_elims(1)[of i] eval_v_elims(2)[of i x "SNum n1"] i by auto
- moreover have "n2 = int (length v)" using eval_e_elims(8) eval_v_elims(1) eval_l.simps i
- by (metis "***" eval_e_elims(1) rcl_val.eq_iff(1) rcl_val.eq_iff(2))
- ultimately have le2: "n \<le> int (length v) " by simp
-
- show ?thesis using le1 le2 by auto
-qed
-
-lemma eval_c_conj2I[intro]:
- assumes "i \<lbrakk> c1 \<rbrakk> ~ True" and "i \<lbrakk> c2 \<rbrakk> ~ True"
- shows "i \<lbrakk> (C_conj c1 c2) \<rbrakk> ~ True"
- using assms eval_c_conjI by metis
-
-lemma valid_split:
- assumes "split n v (v1,v2)" and "\<turnstile>\<^sub>w\<^sub>f \<Theta>"
- shows "\<Theta> ; {||} ; (z , [B_bitvec , B_bitvec ]\<^sup>b , [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ [ L_bitvec v1 ]\<^sup>v , [ L_bitvec v2 ]\<^sup>v ]\<^sup>v ]\<^sup>c\<^sup>e) #\<^sub>\<Gamma> GNil
-\<Turnstile> ([ [ L_bitvec v ]\<^sup>v ]\<^sup>c\<^sup>e == [ [#1[ [ z ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e @@ [#2[ [ z ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e) AND ([| [#1[ [ z ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e == [ [ L_num n ]\<^sup>v ]\<^sup>c\<^sup>e)"
- (is "\<Theta> ; {||} ; ?G \<Turnstile> ?c1 AND ?c2")
- unfolding valid.simps proof
-
- have wfg: " \<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f (z, [B_bitvec , B_bitvec ]\<^sup>b , [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ [ L_bitvec v1 ]\<^sup>v , [ L_bitvec v2 ]\<^sup>v ]\<^sup>v ]\<^sup>c\<^sup>e) #\<^sub>\<Gamma> GNil"
- using wf_intros assms base_for_lit.simps fresh_GNil wfC_v_eq wfG_intros2 by metis
-
- show "\<Theta> ; {||} ; ?G \<turnstile>\<^sub>w\<^sub>f ?c1 AND ?c2"
- apply(rule wfC_conjI)
- apply(rule wfC_eqI)
- apply(rule wfCE_valI)
- apply(rule wfV_litI)
- using wf_intros wfg lookup.simps base_for_lit.simps wfC_v_eq
- apply (metis )+
- done
-
- have len:"int (length v1) = n" using assms split_length by auto
-
- show "\<forall>i. \<Theta> ; ?G \<turnstile> i \<and> i \<Turnstile> ?G \<longrightarrow> i \<Turnstile> (?c1 AND ?c2)"
- proof(rule,rule)
- fix i
- assume a:"\<Theta> ; ?G \<turnstile> i \<and> i \<Turnstile> ?G"
- hence "i \<lbrakk> [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ [ L_bitvec v1 ]\<^sup>v , [ L_bitvec v2 ]\<^sup>v ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrakk> ~ True"
- using is_satis_g.simps is_satis.simps by simp
- then obtain sv where "i \<lbrakk> [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrakk> ~ sv \<and> i \<lbrakk> [ [ [ L_bitvec v1 ]\<^sup>v , [ L_bitvec v2 ]\<^sup>v ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrakk> ~ sv"
- using eval_c_elims by metis
- hence "i \<lbrakk> [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrakk> ~ (SPair (SBitvec v1) (SBitvec v2))" using eval_c_eqI eval_v.intros eval_l.simps
- by (metis eval_e_elims(1) eval_v_uniqueness)
- hence b:"i z = Some (SPair (SBitvec v1) (SBitvec v2))" using a eval_e_elims eval_v_elims by metis
-
- have v1: "i \<lbrakk> [#1[ [ z ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e \<rbrakk> ~ SBitvec v1 "
- using eval_e_fstI eval_e_valI eval_v_varI b by metis
- have v2: "i \<lbrakk> [#2[ [ z ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e \<rbrakk> ~ SBitvec v2"
- using eval_e_sndI eval_e_valI eval_v_varI b by metis
-
- have "i \<lbrakk> [ [ L_bitvec v ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrakk> ~ SBitvec v" using eval_e.intros eval_v.intros eval_l.simps by metis
- moreover have "i \<lbrakk> [ [#1[ [ z ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e @@ [#2[ [ z ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e \<rbrakk> ~ SBitvec v"
- using assms split_concat v1 v2 eval_e_concatI by metis
- moreover have "i \<lbrakk> [| [#1[ [ z ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e \<rbrakk> ~ SNum (int (length v1))"
- using v1 eval_e_lenI by auto
- moreover have "i \<lbrakk> [ [ L_num n ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrakk> ~ SNum n" using eval_e.intros eval_v.intros eval_l.simps by metis
- ultimately show "i \<Turnstile> ?c1 AND ?c2" using is_satis.intros eval_c_conj2I eval_c_eqI len by metis
- qed
-qed
-
-
-lemma is_satis_eq:
- assumes "wfI \<Theta> G i" and "wfCE \<Theta> \<B> G e b"
- shows "is_satis i (e == e)"
-proof(rule)
- obtain s where "eval_e i e s" using eval_e_exist assms by metis
- thus "eval_c i (e == e ) True" using eval_c_eqI by metis
-qed
-
-lemma is_satis_g_i_upd2:
- assumes "eval_v i v s" and "is_satis ((i ( x \<mapsto> s))) c0" and "atom x \<sharp> G" and "wfG \<Theta> \<B> (G3@((x,b,c0)#\<^sub>\<Gamma>G))" and "wfV \<Theta> \<B> G v b" and "wfI \<Theta> (G3[x::=v]\<^sub>\<Gamma>\<^sub>v@G) i"
- and "is_satis_g i (G3[x::=v]\<^sub>\<Gamma>\<^sub>v@G)"
- shows "is_satis_g (i ( x \<mapsto> s)) (G3@((x,b,c0)#\<^sub>\<Gamma>G))"
- using assms proof(induct G3 rule: \<Gamma>_induct)
- case GNil
- hence "is_satis_g (i(x \<mapsto> s)) G" using is_satis_g_i_upd by auto
- then show ?case using GNil using is_satis_g.simps append_g.simps by metis
-next
- case (GCons x' b' c' \<Gamma>')
- hence "x\<noteq>x'" using wfG_cons_append by metis
- hence "is_satis_g i (((x', b', c'[x::=v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> (\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v) @ G))" using subst_gv.simps GCons by auto
- hence *:"is_satis i c'[x::=v]\<^sub>c\<^sub>v \<and> is_satis_g i ((\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v) @ G)" using subst_gv.simps by auto
-
- have "is_satis_g (i(x \<mapsto> s)) ((x', b', c') #\<^sub>\<Gamma> (\<Gamma>'@ (x, b, c0) #\<^sub>\<Gamma> G))" proof(subst is_satis_g.simps,rule)
- show "is_satis (i(x \<mapsto> s)) c'" proof(subst subst_c_satis_full[symmetric])
- show \<open>eval_v i v s\<close> using GCons by auto
- show \<open> \<Theta> ; \<B> ; ((x', b', c') #\<^sub>\<Gamma>\<Gamma>')@(x, b, c0) #\<^sub>\<Gamma> G \<turnstile>\<^sub>w\<^sub>f c' \<close> using GCons wfC_refl by auto
- show \<open>wfI \<Theta> ((((x', b', c') #\<^sub>\<Gamma> \<Gamma>')[x::=v]\<^sub>\<Gamma>\<^sub>v) @ G) i\<close> using GCons by auto
- show \<open> \<Theta> ; \<B> ; G \<turnstile>\<^sub>w\<^sub>f v : b \<close> using GCons by auto
- show \<open>is_satis i c'[x::=v]\<^sub>c\<^sub>v\<close> using * by auto
- qed
- show "is_satis_g (i(x \<mapsto> s)) (\<Gamma>' @ (x, b, c0) #\<^sub>\<Gamma> G)" proof(rule GCons(1))
- show \<open>eval_v i v s\<close> using GCons by auto
- show \<open>is_satis (i(x \<mapsto> s)) c0\<close> using GCons by metis
- show \<open>atom x \<sharp> G\<close> using GCons by auto
- show \<open> \<Theta> ; \<B>\<turnstile>\<^sub>w\<^sub>f \<Gamma>' @ (x, b, c0) #\<^sub>\<Gamma> G \<close> using GCons wfG_elims append_g.simps by metis
- show \<open>is_satis_g i (\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ G)\<close> using * by auto
- show "wfI \<Theta> (\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ G) i" using GCons wfI_def subst_g_assoc_cons \<open>x\<noteq>x'\<close> by auto
- show "\<Theta> ; \<B> ; G \<turnstile>\<^sub>w\<^sub>f v : b " using GCons by auto
- qed
- qed
- moreover have "((x', b', c') #\<^sub>\<Gamma> \<Gamma>' @ (x, b, c0) #\<^sub>\<Gamma> G) = (((x', b', c') #\<^sub>\<Gamma> \<Gamma>') @ (x, b, c0) #\<^sub>\<Gamma> G)" by auto
- ultimately show ?case using GCons by metis
-qed
-
-
+(*<*)
+theory RCLogicL
+ imports RCLogic WellformedL
+begin
+ (*>*)
+
+hide_const Syntax.dom
+
+chapter \<open>Refinement Constraint Logic Lemmas\<close>
+
+section \<open>Lemmas\<close>
+
+lemma wfI_domi:
+ assumes "\<Theta> ; \<Gamma> \<turnstile> i"
+ shows "fst ` toSet \<Gamma> \<subseteq> dom i"
+ using wfI_def toSet.simps assms by fastforce
+
+lemma wfI_lookup:
+ fixes G::\<Gamma> and b::b
+ assumes "Some (b,c) = lookup G x" and "P ; G \<turnstile> i" and "Some s = i x" and "P ; B \<turnstile>\<^sub>w\<^sub>f G"
+ shows "P \<turnstile> s : b"
+proof -
+ have "(x,b,c) \<in> toSet G" using lookup.simps assms
+ using lookup_in_g by blast
+ then obtain s' where *:"Some s' = i x \<and> wfRCV P s' b" using wfI_def assms by auto
+ hence "s' = s" using assms by (metis option.inject)
+ thus ?thesis using * by auto
+qed
+
+lemma wfI_restrict_weakening:
+ assumes "wfI \<Theta> \<Gamma>' i'" and "i = restrict_map i' (fst `toSet \<Gamma>)" and "toSet \<Gamma> \<subseteq> toSet \<Gamma>'"
+ shows "\<Theta> ; \<Gamma> \<turnstile> i"
+proof -
+ { fix x
+ assume "x \<in> toSet \<Gamma>"
+ have "case x of (x, b, c) \<Rightarrow> \<exists>s. Some s = i x \<and> \<Theta> \<turnstile> s : b" using assms wfI_def
+ proof -
+ have "case x of (x, b, c) \<Rightarrow> \<exists>s. Some s = i' x \<and> \<Theta> \<turnstile> s : b"
+ using \<open>x \<in> toSet \<Gamma>\<close> assms wfI_def by auto
+ then have "\<exists>s. Some s = i (fst x) \<and> wfRCV \<Theta> s (fst (snd x))"
+ by (simp add: \<open>x \<in> toSet \<Gamma>\<close> assms(2) case_prod_unfold)
+ then show ?thesis
+ by (simp add: case_prod_unfold)
+ qed
+ }
+ thus ?thesis using wfI_def assms by auto
+qed
+
+lemma wfI_suffix:
+ fixes G::\<Gamma>
+ assumes "wfI P (G'@G) i" and "P ; B \<turnstile>\<^sub>w\<^sub>f G"
+ shows "P ; G \<turnstile> i"
+ using wfI_def append_g.simps assms toSet.simps by simp
+
+lemma wfI_replace_inside:
+ assumes "wfI \<Theta> (\<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>) i"
+ shows "wfI \<Theta> (\<Gamma>' @ (x, b, c') #\<^sub>\<Gamma> \<Gamma>) i"
+ using wfI_def toSet_splitP assms by simp
+
+section \<open>Existence of evaluation\<close>
+
+lemma eval_l_base:
+ "\<Theta> \<turnstile> \<lbrakk> l \<rbrakk> : (base_for_lit l)"
+ apply(nominal_induct l rule:l.strong_induct)
+ using wfRCV.intros eval_l.simps base_for_lit.simps by auto+
+
+lemma obtain_fresh_bv_dclist:
+ fixes tm::"'a::fs"
+ assumes "(dc, \<lbrace> x : b | c \<rbrace>) \<in> set dclist"
+ obtains bv1::bv and dclist1 x1 b1 c1 where "AF_typedef_poly tyid bv dclist = AF_typedef_poly tyid bv1 dclist1
+ \<and> (dc, \<lbrace> x1 : b1 | c1 \<rbrace>) \<in> set dclist1 \<and> atom bv1 \<sharp> tm"
+proof -
+ obtain bv1 dclist1 where "AF_typedef_poly tyid bv dclist = AF_typedef_poly tyid bv1 dclist1 \<and> atom bv1 \<sharp> tm"
+ using obtain_fresh_bv by metis
+ moreover hence "[[atom bv]]lst. dclist = [[atom bv1]]lst. dclist1" using type_def.eq_iff by metis
+ moreover then obtain x1 b1 c1 where "(dc, \<lbrace> x1 : b1 | c1 \<rbrace>) \<in> set dclist1" using td_lookup_eq_iff assms by metis
+ ultimately show ?thesis using that by blast
+qed
+
+lemma obtain_fresh_bv_dclist_b_iff:
+ fixes tm::"'a::fs"
+ assumes "(dc, \<lbrace> x : b | c \<rbrace>) \<in> set dclist" and "AF_typedef_poly tyid bv dclist \<in> set P" and "\<turnstile>\<^sub>w\<^sub>f P"
+ obtains bv1::bv and dclist1 x1 b1 c1 where "AF_typedef_poly tyid bv dclist = AF_typedef_poly tyid bv1 dclist1
+ \<and> (dc, \<lbrace> x1 : b1 | c1 \<rbrace>) \<in> set dclist1 \<and> atom bv1 \<sharp> tm \<and> b[bv::=b']\<^sub>b\<^sub>b=b1[bv1::=b']\<^sub>b\<^sub>b"
+proof -
+ obtain bv1 dclist1 x1 b1 c1 where *:"AF_typedef_poly tyid bv dclist = AF_typedef_poly tyid bv1 dclist1 \<and> atom bv1 \<sharp> tm
+ \<and> (dc, \<lbrace> x1 : b1 | c1 \<rbrace>) \<in> set dclist1" using obtain_fresh_bv_dclist assms by metis
+
+ hence "AF_typedef_poly tyid bv1 dclist1 \<in> set P" using assms by metis
+
+ hence "b[bv::=b']\<^sub>b\<^sub>b = b1[bv1::=b']\<^sub>b\<^sub>b"
+ using wfTh_typedef_poly_b_eq_iff[OF assms(2) assms(1) _ _ assms(3),of bv1 dclist1 x1 b1 c1 b'] * by metis
+
+ from this that show ?thesis using * by metis
+qed
+
+lemma eval_v_exist:
+ fixes \<Gamma>::\<Gamma> and v::v and b::b
+ assumes "P ; \<Gamma> \<turnstile> i" and "P ; B ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b"
+ shows "\<exists>s. i \<lbrakk> v \<rbrakk> ~ s \<and> P \<turnstile> s : b "
+ using assms proof(nominal_induct v arbitrary: b rule:v.strong_induct)
+ case (V_lit x)
+ then show ?case using eval_l_base eval_v.intros eval_l.simps wfV_elims rcl_val.supp pure_supp by metis
+next
+ case (V_var x)
+ then obtain c where *:"Some (b,c) = lookup \<Gamma> x" using wfV_elims by metis
+ hence "x \<in> fst ` toSet \<Gamma>" using lookup_x by blast
+ hence "x \<in> dom i" using wfI_domi using assms by blast
+ then obtain s where "i x = Some s" by auto
+ moreover hence "P \<turnstile> s : b" using wfRCV.intros wfI_lookup * V_var
+ by (metis wfV_wf)
+
+ ultimately show ?case using eval_v.intros rcl_val.supp b.supp by metis
+next
+ case (V_pair v1 v2)
+ then obtain b1 and b2 where *:"P ; B ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v1 : b1 \<and> P ; B ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v2 : b2 \<and> b = B_pair b1 b2" using wfV_elims by metis
+ then obtain s1 and s2 where "eval_v i v1 s1 \<and> wfRCV P s1 b1 \<and> eval_v i v2 s2 \<and> wfRCV P s2 b2" using V_pair by metis
+ thus ?case using eval_v.intros wfRCV.intros * by metis
+next
+ case (V_cons tyid dc v)
+ then obtain s and b'::b and dclist and x::x and c::c where "(wfV P B \<Gamma> v b') \<and> i \<lbrakk> v \<rbrakk> ~ s \<and> P \<turnstile> s : b' \<and> b = B_id tyid \<and>
+ AF_typedef tyid dclist \<in> set P \<and> (dc, \<lbrace> x : b' | c \<rbrace>) \<in> set dclist" using wfV_elims(4) by metis
+ then show ?case using eval_v.intros(4) wfRCV.intros(5) V_cons by metis
+next
+ case (V_consp tyid dc b' v)
+
+ obtain b'a::b and bv and dclist and x::x and c::c where *:"(wfV P B \<Gamma> v b'a[bv::=b']\<^sub>b\<^sub>b) \<and> b = B_app tyid b' \<and>
+ AF_typedef_poly tyid bv dclist \<in> set P \<and> (dc, \<lbrace> x : b'a | c \<rbrace>) \<in> set dclist \<and>
+ atom bv \<sharp> (P, B_app tyid b',B) " using wf_strong_elim(1)[OF V_consp(3)] by metis
+
+ then obtain s where **:"i \<lbrakk> v \<rbrakk> ~ s \<and> P \<turnstile> s : b'a[bv::=b']\<^sub>b\<^sub>b" using V_consp by auto
+
+ have " \<turnstile>\<^sub>w\<^sub>f P" using wfX_wfY V_consp by metis
+ then obtain bv1::bv and dclist1 x1 b1 c1 where 3:"AF_typedef_poly tyid bv dclist = AF_typedef_poly tyid bv1 dclist1
+ \<and> (dc, \<lbrace> x1 : b1 | c1 \<rbrace>) \<in> set dclist1 \<and> atom bv1 \<sharp> (P, SConsp tyid dc b' s, B_app tyid b')
+ \<and> b'a[bv::=b']\<^sub>b\<^sub>b = b1[bv1::=b']\<^sub>b\<^sub>b"
+ using * obtain_fresh_bv_dclist_b_iff by blast
+
+ have " i \<lbrakk> V_consp tyid dc b' v \<rbrakk> ~ SConsp tyid dc b' s" using eval_v.intros by (simp add: "**")
+
+ moreover have " P \<turnstile> SConsp tyid dc b' s : B_app tyid b'" proof
+ show \<open>AF_typedef_poly tyid bv1 dclist1 \<in> set P\<close> using 3 * by metis
+ show \<open>(dc, \<lbrace> x1 : b1 | c1 \<rbrace>) \<in> set dclist1\<close> using 3 by auto
+ thus \<open>atom bv1 \<sharp> (P, SConsp tyid dc b' s, B_app tyid b')\<close> using * 3 fresh_prodN by metis
+ show \<open> P \<turnstile> s : b1[bv1::=b']\<^sub>b\<^sub>b\<close> using 3 ** by auto
+ qed
+
+ ultimately show ?case using eval_v.intros wfRCV.intros V_consp * by auto
+qed
+
+lemma eval_v_uniqueness:
+ fixes v::v
+ assumes "i \<lbrakk> v \<rbrakk> ~ s" and "i \<lbrakk> v \<rbrakk> ~ s'"
+ shows "s=s'"
+ using assms proof(nominal_induct v arbitrary: s s' rule:v.strong_induct)
+ case (V_lit x)
+ then show ?case using eval_v_elims eval_l.simps by metis
+next
+ case (V_var x)
+ then show ?case using eval_v_elims by (metis option.inject)
+next
+ case (V_pair v1 v2)
+ obtain s1 and s2 where s:"i \<lbrakk> v1 \<rbrakk> ~ s1 \<and> i \<lbrakk> v2 \<rbrakk> ~ s2 \<and> s = SPair s1 s2" using eval_v_elims V_pair by metis
+ obtain s1' and s2' where s':"i \<lbrakk> v1 \<rbrakk> ~ s1' \<and> i \<lbrakk> v2 \<rbrakk> ~ s2' \<and> s' = SPair s1' s2'" using eval_v_elims V_pair by metis
+ then show ?case using eval_v_elims using V_pair s s' by auto
+next
+ case (V_cons tyid dc v1)
+ obtain sv1 where 1:"i \<lbrakk> v1 \<rbrakk> ~ sv1 \<and> s = SCons tyid dc sv1" using eval_v_elims V_cons by metis
+ moreover obtain sv2 where 2:"i \<lbrakk> v1 \<rbrakk> ~ sv2 \<and> s' = SCons tyid dc sv2" using eval_v_elims V_cons by metis
+ ultimately have "sv1 = sv2" using V_cons by auto
+ then show ?case using 1 2 by auto
+next
+ case (V_consp tyid dc b v1)
+ then show ?case using eval_v_elims by metis
+
+qed
+
+lemma eval_v_base:
+ fixes \<Gamma>::\<Gamma> and v::v and b::b
+ assumes "P ; \<Gamma> \<turnstile> i" and "P ; B ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b" and "i \<lbrakk> v \<rbrakk> ~ s"
+ shows "P \<turnstile> s : b"
+ using eval_v_exist eval_v_uniqueness assms by metis
+
+lemma eval_e_uniqueness:
+ fixes e::ce
+ assumes "i \<lbrakk> e \<rbrakk> ~ s" and "i \<lbrakk> e \<rbrakk> ~ s'"
+ shows "s=s'"
+ using assms proof(nominal_induct e arbitrary: s s' rule:ce.strong_induct)
+ case (CE_val x)
+ then show ?case using eval_v_uniqueness eval_e_elims by metis
+next
+ case (CE_op opp x1 x2)
+ consider "opp = Plus" | "opp = LEq" | "opp = Eq" using opp.exhaust by metis
+ thus ?case proof(cases)
+ case 1
+ hence a1:"eval_e i (CE_op Plus x1 x2) s" and a2:"eval_e i (CE_op Plus x1 x2) s'" using CE_op by auto
+ then show ?thesis using eval_e_elims(2)[OF a1] eval_e_elims(2)[OF a2]
+ CE_op eval_e_plusI
+ by (metis rcl_val.eq_iff(2))
+ next
+ case 2
+ hence a1:"eval_e i (CE_op LEq x1 x2) s" and a2:"eval_e i (CE_op LEq x1 x2) s'" using CE_op by auto
+ then show ?thesis using eval_v_uniqueness eval_e_elims(3)[OF a1] eval_e_elims(3)[OF a2]
+ CE_op eval_e_plusI
+ by (metis rcl_val.eq_iff(2))
+ next
+ case 3
+ hence a1:"eval_e i (CE_op Eq x1 x2) s" and a2:"eval_e i (CE_op Eq x1 x2) s'" using CE_op by auto
+ then show ?thesis using eval_v_uniqueness eval_e_elims(4)[OF a1] eval_e_elims(4)[OF a2]
+ CE_op eval_e_plusI
+ by (metis rcl_val.eq_iff(2))
+ qed
+next
+ case (CE_concat x1 x2)
+ hence a1:"eval_e i (CE_concat x1 x2) s" and a2:"eval_e i (CE_concat x1 x2) s'" using CE_concat by auto
+ show ?case using eval_e_elims(7)[OF a1] eval_e_elims(7)[OF a2] CE_concat eval_e_concatI rcl_val.eq_iff
+ proof -
+ assume "\<And>P. (\<And>bv1 bv2. \<lbrakk>s' = SBitvec (bv1 @ bv2); i \<lbrakk> x1 \<rbrakk> ~ SBitvec bv1 ; i \<lbrakk> x2 \<rbrakk> ~ SBitvec bv2 \<rbrakk> \<Longrightarrow> P) \<Longrightarrow> P"
+ obtain bbs :: "bit list" and bbsa :: "bit list" where
+ "i \<lbrakk> x2 \<rbrakk> ~ SBitvec bbs \<and> i \<lbrakk> x1 \<rbrakk> ~ SBitvec bbsa \<and> SBitvec (bbsa @ bbs) = s'"
+ by (metis \<open>\<And>P. (\<And>bv1 bv2. \<lbrakk>s' = SBitvec (bv1 @ bv2); i \<lbrakk> x1 \<rbrakk> ~ SBitvec bv1 ; i \<lbrakk> x2 \<rbrakk> ~ SBitvec bv2 \<rbrakk> \<Longrightarrow> P) \<Longrightarrow> P\<close>) (* 93 ms *)
+ then have "s' = s"
+ by (metis (no_types) \<open>\<And>P. (\<And>bv1 bv2. \<lbrakk>s = SBitvec (bv1 @ bv2); i \<lbrakk> x1 \<rbrakk> ~ SBitvec bv1 ; i \<lbrakk> x2 \<rbrakk> ~ SBitvec bv2 \<rbrakk> \<Longrightarrow> P) \<Longrightarrow> P\<close> \<open>\<And>s' s. \<lbrakk>i \<lbrakk> x1 \<rbrakk> ~ s ; i \<lbrakk> x1 \<rbrakk> ~ s' \<rbrakk> \<Longrightarrow> s = s'\<close> \<open>\<And>s' s. \<lbrakk>i \<lbrakk> x2 \<rbrakk> ~ s ; i \<lbrakk> x2 \<rbrakk> ~ s' \<rbrakk> \<Longrightarrow> s = s'\<close> rcl_val.eq_iff(1)) (* 125 ms *)
+ then show ?thesis
+ by metis (* 0.0 ms *)
+ qed
+next
+ case (CE_fst x)
+ then show ?case using eval_v_uniqueness by (meson eval_e_elims rcl_val.eq_iff)
+next
+ case (CE_snd x)
+ then show ?case using eval_v_uniqueness by (meson eval_e_elims rcl_val.eq_iff)
+next
+ case (CE_len x)
+ then show ?case using eval_e_elims rcl_val.eq_iff
+ proof -
+ obtain bbs :: "rcl_val \<Rightarrow> ce \<Rightarrow> (x \<Rightarrow> rcl_val option) \<Rightarrow> bit list" where
+ "\<forall>x0 x1 x2. (\<exists>v3. x0 = SNum (int (length v3)) \<and> x2 \<lbrakk> x1 \<rbrakk> ~ SBitvec v3 ) = (x0 = SNum (int (length (bbs x0 x1 x2))) \<and> x2 \<lbrakk> x1 \<rbrakk> ~ SBitvec (bbs x0 x1 x2) )"
+ by moura (* 0.0 ms *)
+ then have "\<forall>f c r. \<not> f \<lbrakk> [| c |]\<^sup>c\<^sup>e \<rbrakk> ~ r \<or> r = SNum (int (length (bbs r c f))) \<and> f \<lbrakk> c \<rbrakk> ~ SBitvec (bbs r c f)"
+ by (meson eval_e_elims(8)) (* 46 ms *)
+ then show ?thesis
+ by (metis (no_types) CE_len.hyps CE_len.prems(1) CE_len.prems(2) rcl_val.eq_iff(1)) (* 31 ms *)
+ qed
+
+qed
+
+lemma wfV_eval_bitvec:
+ fixes v::v
+ assumes "P ; B ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : B_bitvec" and "P ; \<Gamma> \<turnstile> i"
+ shows "\<exists>bv. eval_v i v (SBitvec bv)"
+proof -
+ obtain s where "i \<lbrakk> v \<rbrakk> ~ s \<and> wfRCV P s B_bitvec" using eval_v_exist assms by metis
+ moreover then obtain bv where "s = SBitvec bv" using wfRCV_elims(1)[of P s] by metis
+ ultimately show ?thesis by metis
+qed
+
+lemma wfV_eval_pair:
+ fixes v::v
+ assumes "P ; B ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : B_pair b1 b2" and "P ; \<Gamma> \<turnstile> i"
+ shows "\<exists>s1 s2. eval_v i v (SPair s1 s2)"
+proof -
+ obtain s where "i \<lbrakk> v \<rbrakk> ~ s \<and> wfRCV P s (B_pair b1 b2)" using eval_v_exist assms by metis
+ moreover then obtain s1 and s2 where "s = SPair s1 s2" using wfRCV_elims(2)[of P s] by metis
+ ultimately show ?thesis by metis
+qed
+
+lemma wfV_eval_int:
+ fixes v::v
+ assumes "P ; B ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : B_int" and "P ; \<Gamma> \<turnstile> i"
+ shows "\<exists>n. eval_v i v (SNum n)"
+proof -
+ obtain s where "i \<lbrakk> v \<rbrakk> ~ s \<and> wfRCV P s (B_int)" using eval_v_exist assms by metis
+ moreover then obtain n where "s = SNum n" using wfRCV_elims(3)[of P s] by metis
+ ultimately show ?thesis by metis
+qed
+
+text \<open>Well sorted value with a well sorted valuation evaluates\<close>
+lemma wfI_wfV_eval_v:
+ fixes v::v and b::b
+ assumes "\<Theta> ; B ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b" and "wfI \<Theta> \<Gamma> i"
+ shows "\<exists>s. i \<lbrakk> v \<rbrakk> ~ s \<and> \<Theta> \<turnstile> s : b"
+ using eval_v_exist assms by auto
+
+lemma wfI_wfCE_eval_e:
+ fixes e::ce and b::b
+ assumes "wfCE P B G e b" and "P ; G \<turnstile> i"
+ shows "\<exists>s. i \<lbrakk> e \<rbrakk> ~ s \<and> P \<turnstile> s : b"
+ using assms proof(nominal_induct e arbitrary: b rule: ce.strong_induct)
+ case (CE_val v)
+ obtain s where "i \<lbrakk> v \<rbrakk> ~ s \<and> P \<turnstile> s : b" using wfI_wfV_eval_v[of P B G v b i] assms wfCE_elims(1) [of P B G v b] CE_val by auto
+ then show ?case using CE_val eval_e.intros(1)[of i v s ] by auto
+next
+ case (CE_op opp v1 v2)
+
+ consider "opp =Plus" | "opp=LEq" | "opp=Eq" using opp.exhaust by auto
+
+ thus ?case proof(cases)
+ case 1
+ hence "wfCE P B G v1 B_int \<and> wfCE P B G v2 B_int" using wfCE_elims(2) CE_op
+
+ by blast
+ then obtain s1 and s2 where *: "eval_e i v1 s1 \<and> wfRCV P s1 B_int \<and> eval_e i v2 s2 \<and> wfRCV P s2 B_int"
+ using wfI_wfV_eval_v CE_op by metis
+ then obtain n1 and n2 where **:"s2=SNum n2 \<and> s1 = SNum n1" using wfRCV_elims by meson
+ hence "eval_e i (CE_op Plus v1 v2) (SNum (n1+n2))" using eval_e_plusI * ** by simp
+ moreover have "wfRCV P (SNum (n1+n2)) B_int" using wfRCV.intros by auto
+ ultimately show ?thesis using 1
+ using CE_op.prems(1) wfCE_elims(2) by blast
+ next
+ case 2
+ hence "wfCE P B G v1 B_int \<and> wfCE P B G v2 B_int" using wfCE_elims(3) CE_op
+ by blast
+ then obtain s1 and s2 where *: "eval_e i v1 s1 \<and> wfRCV P s1 B_int \<and> eval_e i v2 s2 \<and> wfRCV P s2 B_int"
+ using wfI_wfV_eval_v CE_op by metis
+ then obtain n1 and n2 where **:"s2=SNum n2 \<and> s1 = SNum n1" using wfRCV_elims by meson
+ hence "eval_e i (CE_op LEq v1 v2) (SBool (n1 \<le> n2))" using eval_e_leqI * ** by simp
+ moreover have "wfRCV P (SBool (n1\<le>n2)) B_bool" using wfRCV.intros by auto
+ ultimately show ?thesis using 2
+ using CE_op.prems wfCE_elims by metis
+ next
+ case 3
+ then obtain b2 where "wfCE P B G v1 b2 \<and> wfCE P B G v2 b2" using wfCE_elims(9) CE_op
+ by blast
+ then obtain s1 and s2 where *: "eval_e i v1 s1 \<and> wfRCV P s1 b2 \<and> eval_e i v2 s2 \<and> wfRCV P s2 b2"
+ using wfI_wfV_eval_v CE_op by metis
+ hence "eval_e i (CE_op Eq v1 v2) (SBool (s1 = s2))" using eval_e_leqI *
+ by (simp add: eval_e_eqI)
+ moreover have "wfRCV P (SBool (s1 = s2)) B_bool" using wfRCV.intros by auto
+ ultimately show ?thesis using 3
+ using CE_op.prems wfCE_elims by metis
+ qed
+next
+ case (CE_concat v1 v2)
+ then obtain s1 and s2 where *:"b = B_bitvec \<and> eval_e i v1 s1 \<and> eval_e i v2 s2 \<and>
+ wfRCV P s1 B_bitvec \<and> wfRCV P s2 B_bitvec" using
+ CE_concat
+ by (meson wfCE_elims(6))
+ thus ?case using eval_e_concatI wfRCV.intros(1) wfRCV_elims
+ proof -
+ obtain bbs :: "type_def list \<Rightarrow> rcl_val \<Rightarrow> bit list" where
+ "\<forall>ts s. \<not> ts \<turnstile> s : B_bitvec \<or> s = SBitvec (bbs ts s)"
+ using wfRCV_elims(1) by moura (* 156 ms *)
+ then show ?thesis
+ by (metis (no_types) "local.*" wfRCV_BBitvecI eval_e_concatI) (* 78 ms *)
+ qed
+next
+ case (CE_fst v1)
+ thus ?case using eval_e_fstI wfRCV.intros wfCE_elims wfI_wfV_eval_v
+ by (metis wfRCV_elims(2) rcl_val.eq_iff(4))
+next
+ case (CE_snd v1)
+ thus ?case using eval_e_sndI wfRCV.intros wfCE_elims wfI_wfV_eval_v
+ by (metis wfRCV_elims(2) rcl_val.eq_iff(4))
+next
+ case (CE_len x)
+ thus ?case using eval_e_lenI wfRCV.intros wfCE_elims wfI_wfV_eval_v wfV_eval_bitvec
+ by (metis wfRCV_elims(1))
+qed
+
+lemma eval_e_exist:
+ fixes \<Gamma>::\<Gamma> and e::ce
+ assumes "P ; \<Gamma> \<turnstile> i" and "P ; B ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f e : b"
+ shows "\<exists>s. i \<lbrakk> e \<rbrakk> ~ s"
+ using assms proof(nominal_induct e arbitrary: b rule:ce.strong_induct)
+ case (CE_val v)
+ then show ?case using eval_v_exist wfCE_elims eval_e.intros by metis
+next
+ case (CE_op op v1 v2)
+
+ show ?case proof(rule opp.exhaust)
+ assume \<open>op = Plus\<close>
+ hence "P ; B ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v1 : B_int \<and> P ; B ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v2 : B_int \<and> b = B_int" using wfCE_elims CE_op by metis
+ then obtain n1 and n2 where "eval_e i v1 (SNum n1) \<and> eval_e i v2 (SNum n2)" using CE_op eval_v_exist wfV_eval_int
+ by (metis wfI_wfCE_eval_e wfRCV_elims(3))
+ then show \<open>\<exists>a. eval_e i (CE_op op v1 v2) a\<close> using eval_e_plusI[of i v1 _ v2] \<open>op=Plus\<close> by auto
+ next
+ assume \<open>op = LEq\<close>
+ hence "P ; B ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v1 : B_int \<and> P ; B ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v2 : B_int \<and> b = B_bool" using wfCE_elims CE_op by metis
+ then obtain n1 and n2 where "eval_e i v1 (SNum n1) \<and> eval_e i v2 (SNum n2)" using CE_op eval_v_exist wfV_eval_int
+ by (metis wfI_wfCE_eval_e wfRCV_elims(3))
+ then show \<open>\<exists>a. eval_e i (CE_op op v1 v2) a\<close> using eval_e_leqI[of i v1 _ v2] eval_v_exist \<open>op=LEq\<close> CE_op by auto
+ next
+ assume \<open>op = Eq\<close>
+ then obtain b1 where "P ; B ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v1 : b1 \<and> P ; B ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v2 : b1 \<and> b = B_bool" using wfCE_elims CE_op by metis
+ then obtain s1 and s2 where "eval_e i v1 s1 \<and> eval_e i v2 s2" using CE_op eval_v_exist wfV_eval_int
+ by (metis wfI_wfCE_eval_e wfRCV_elims(3))
+ then show \<open>\<exists>a. eval_e i (CE_op op v1 v2) a\<close> using eval_e_eqI[of i v1 _ v2] eval_v_exist \<open>op=Eq\<close> CE_op by auto
+ qed
+next
+ case (CE_concat v1 v2)
+ then obtain bv1 and bv2 where "eval_e i v1 (SBitvec bv1) \<and> eval_e i v2 (SBitvec bv2)"
+ using wfV_eval_bitvec wfCE_elims(6)
+ by (meson eval_e_elims(7) wfI_wfCE_eval_e)
+ then show ?case using eval_e.intros by metis
+next
+ case (CE_fst ce)
+ then obtain b2 where b:"P ; B ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ce : B_pair b b2" using wfCE_elims by metis
+ then obtain s where s:"i \<lbrakk> ce \<rbrakk> ~ s" using CE_fst by auto
+ then obtain s1 and s2 where "s = (SPair s1 s2)" using eval_e_elims(4) CE_fst wfI_wfCE_eval_e[of P B \<Gamma> ce "B_pair b b2" i,OF b] wfRCV_elims(2)[of P s b b2]
+ by (metis eval_e_uniqueness)
+ then show ?case using s eval_e.intros by metis
+next
+ case (CE_snd ce)
+ then obtain b1 where b:"P ; B ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ce : B_pair b1 b" using wfCE_elims by metis
+ then obtain s where s:"i \<lbrakk> ce \<rbrakk> ~ s" using CE_snd by auto
+ then obtain s1 and s2 where "s = (SPair s1 s2)"
+ using eval_e_elims(5) CE_snd wfI_wfCE_eval_e[of P B \<Gamma> ce "B_pair b1 b" i,OF b] wfRCV_elims(2)[of P s b b1]
+ eval_e_uniqueness
+ by (metis wfRCV_elims(2))
+ then show ?case using s eval_e.intros by metis
+next
+ case (CE_len v1)
+ then obtain bv1 where "eval_e i v1 (SBitvec bv1)"
+ using wfV_eval_bitvec CE_len wfCE_elims eval_e_uniqueness
+ by (metis eval_e_elims(7) wfCE_concatI wfI_wfCE_eval_e)
+ then show ?case using eval_e.intros by metis
+qed
+
+lemma eval_c_exist:
+ fixes \<Gamma>::\<Gamma> and c::c
+ assumes "P ; \<Gamma> \<turnstile> i" and "P ; B ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f c"
+ shows "\<exists>s. i \<lbrakk> c \<rbrakk> ~ s"
+ using assms proof(nominal_induct c rule: c.strong_induct)
+ case C_true
+ then show ?case using eval_c.intros wfC_elims by metis
+next
+ case C_false
+ then show ?case using eval_c.intros wfC_elims by metis
+next
+ case (C_conj c1 c2)
+ then show ?case using eval_c.intros wfC_elims by metis
+next
+ case (C_disj x1 x2)
+ then show ?case using eval_c.intros wfC_elims by metis
+next
+ case (C_not x)
+ then show ?case using eval_c.intros wfC_elims by metis
+next
+ case (C_imp x1 x2)
+ then show ?case using eval_c.intros eval_e_exist wfC_elims by metis
+next
+ case (C_eq x1 x2)
+ then show ?case using eval_c.intros eval_e_exist wfC_elims by metis
+qed
+
+lemma eval_c_uniqueness:
+ fixes c::c
+ assumes "i \<lbrakk> c \<rbrakk> ~ s" and "i \<lbrakk> c \<rbrakk> ~ s'"
+ shows "s=s'"
+ using assms proof(nominal_induct c arbitrary: s s' rule:c.strong_induct)
+ case C_true
+ then show ?case using eval_c_elims by metis
+next
+ case C_false
+ then show ?case using eval_c_elims by metis
+next
+ case (C_conj x1 x2)
+ then show ?case using eval_c_elims(3) by (metis (full_types))
+next
+ case (C_disj x1 x2)
+ then show ?case using eval_c_elims(4) by (metis (full_types))
+next
+ case (C_not x)
+ then show ?case using eval_c_elims(6) by (metis (full_types))
+next
+ case (C_imp x1 x2)
+ then show ?case using eval_c_elims(5) by (metis (full_types))
+next
+ case (C_eq x1 x2)
+ then show ?case using eval_e_uniqueness eval_c_elims(7) by metis
+qed
+
+lemma wfI_wfC_eval_c:
+ fixes c::c
+ assumes "wfC P B G c" and "P ; G \<turnstile> i"
+ shows "\<exists>s. i \<lbrakk> c \<rbrakk> ~ s"
+ using assms proof(nominal_induct c rule: c.strong_induct)
+qed(metis wfC_elims wfI_wfCE_eval_e eval_c.intros)+
+
+section \<open>Satisfiability\<close>
+
+lemma satis_reflI:
+ fixes c::c
+ assumes "i \<Turnstile> ((x, b, c) #\<^sub>\<Gamma> G)"
+ shows "i \<Turnstile> c"
+ using assms by auto
+
+lemma is_satis_mp:
+ fixes c1::c and c2::c
+ assumes "i \<Turnstile> (c1 IMP c2)" and "i \<Turnstile> c1"
+ shows "i \<Turnstile> c2"
+ using assms proof -
+ have "eval_c i (c1 IMP c2) True" using is_satis.simps using assms by blast
+ then obtain b1 and b2 where "True = (b1 \<longrightarrow> b2) \<and> eval_c i c1 b1 \<and> eval_c i c2 b2"
+ using eval_c_elims(5) by metis
+ moreover have "eval_c i c1 True" using is_satis.simps using assms by blast
+ moreover have "b1 = True" using calculation eval_c_uniqueness by blast
+ ultimately have "eval_c i c2 True" by auto
+ thus ?thesis using is_satis.intros by auto
+qed
+
+lemma is_satis_imp:
+ fixes c1::c and c2::c
+ assumes "i \<Turnstile> c1 \<longrightarrow> i \<Turnstile> c2" and "i \<lbrakk> c1 \<rbrakk> ~ b1" and "i \<lbrakk> c2 \<rbrakk> ~ b2"
+ shows "i \<Turnstile> (c1 IMP c2)"
+proof(cases b1)
+ case True
+ hence "i \<Turnstile> c2" using assms is_satis.simps by simp
+ hence "b2 = True" using is_satis.simps assms
+ using eval_c_uniqueness by blast
+ then show ?thesis using eval_c_impI is_satis.simps assms by force
+next
+ case False
+ then show ?thesis using assms eval_c_impI is_satis.simps by metis
+qed
+
+lemma is_satis_iff:
+ "i \<Turnstile> G = (\<forall>x b c. (x,b,c) \<in> toSet G \<longrightarrow> i \<Turnstile> c)"
+ by(induct G,auto)
+
+lemma is_satis_g_append:
+ "i \<Turnstile> (G1@G2) = (i \<Turnstile> G1 \<and> i \<Turnstile> G2)"
+ using is_satis_g.simps is_satis_iff by auto
+
+section \<open>Substitution for Evaluation\<close>
+
+lemma eval_v_i_upd:
+ fixes v::v
+ assumes "atom x \<sharp> v" and "i \<lbrakk> v \<rbrakk> ~ s'"
+ shows "eval_v ((i ( x \<mapsto>s))) v s' "
+ using assms proof(nominal_induct v arbitrary: s' rule:v.strong_induct)
+ case (V_lit x)
+ then show ?case by (metis eval_v_elims(1) eval_v_litI)
+next
+ case (V_var y)
+ then obtain s where *: "Some s = i y \<and> s=s'" using eval_v_elims by metis
+ moreover have "x \<noteq> y" using \<open>atom x \<sharp> V_var y\<close> v.supp by simp
+ ultimately have "(i ( x \<mapsto>s)) y = Some s"
+ by (simp add: \<open>Some s = i y \<and> s = s'\<close>)
+ then show ?case using eval_v_varI * \<open>x \<noteq> y\<close>
+ by (simp add: eval_v.eval_v_varI)
+next
+ case (V_pair v1 v2)
+ hence "atom x \<sharp> v1 \<and> atom x \<sharp> v2" using v.supp by simp
+ moreover obtain s1 and s2 where *: "eval_v i v1 s1 \<and> eval_v i v2 s2 \<and> s' = SPair s1 s2" using eval_v_elims V_pair by metis
+ ultimately have "eval_v ((i ( x \<mapsto>s))) v1 s1 \<and> eval_v ((i ( x \<mapsto>s))) v2 s2" using V_pair by blast
+ thus ?case using eval_v_pairI * by meson
+next
+ case (V_cons tyid dc v1)
+ hence "atom x \<sharp> v1" using v.supp by simp
+ moreover obtain s1 where *: "eval_v i v1 s1 \<and> s' = SCons tyid dc s1" using eval_v_elims V_cons by metis
+ ultimately have "eval_v ((i ( x \<mapsto>s))) v1 s1" using V_cons by blast
+ thus ?case using eval_v_consI * by meson
+next
+ case (V_consp tyid dc b1 v1)
+
+ hence "atom x \<sharp> v1" using v.supp by simp
+ moreover obtain s1 where *: "eval_v i v1 s1 \<and> s' = SConsp tyid dc b1 s1" using eval_v_elims V_consp by metis
+ ultimately have "eval_v ((i ( x \<mapsto>s))) v1 s1" using V_consp by blast
+ thus ?case using eval_v_conspI * by meson
+qed
+
+lemma eval_e_i_upd:
+ fixes e::ce
+ assumes "i \<lbrakk> e \<rbrakk> ~ s'" and "atom x \<sharp> e"
+ shows " (i ( x \<mapsto>s)) \<lbrakk> e \<rbrakk> ~ s'"
+ using assms apply(induct rule: eval_e.induct) using eval_v_i_upd eval_e_elims
+ by (meson ce.fresh eval_e.intros)+
+
+lemma eval_c_i_upd:
+ fixes c::c
+ assumes "i \<lbrakk> c \<rbrakk> ~ s'" and "atom x \<sharp> c"
+ shows "((i ( x \<mapsto>s))) \<lbrakk> c \<rbrakk> ~ s' "
+ using assms proof(induct rule:eval_c.induct)
+ case (eval_c_eqI i e1 sv1 e2 sv2)
+ then show ?case using RCLogic.eval_c_eqI eval_e_i_upd c.fresh by metis
+qed(simp add: eval_c.intros)+
+
+lemma subst_v_eval_v[simp]:
+ fixes v::v and v'::v
+ assumes "i \<lbrakk> v \<rbrakk> ~ s" and "i \<lbrakk> (v'[x::=v]\<^sub>v\<^sub>v) \<rbrakk> ~ s'"
+ shows "(i ( x \<mapsto> s )) \<lbrakk> v' \<rbrakk> ~ s'"
+ using assms proof(nominal_induct v' arbitrary: s' rule:v.strong_induct)
+ case (V_lit x)
+ then show ?case using subst_vv.simps
+ by (metis eval_v_elims(1) eval_v_litI)
+next
+ case (V_var x')
+ then show ?case proof(cases "x=x'")
+ case True
+ hence "(V_var x')[x::=v]\<^sub>v\<^sub>v = v" using subst_vv.simps by auto
+ then show ?thesis using V_var eval_v_elims eval_v_varI eval_v_uniqueness True
+ by (simp add: eval_v.eval_v_varI)
+ next
+ case False
+ hence "atom x \<sharp> (V_var x')" by simp
+ then show ?thesis using eval_v_i_upd False V_var by fastforce
+ qed
+next
+ case (V_pair v1 v2)
+ then obtain s1 and s2 where *:"eval_v i (v1[x::=v]\<^sub>v\<^sub>v) s1 \<and> eval_v i (v2[x::=v]\<^sub>v\<^sub>v) s2 \<and> s' = SPair s1 s2" using V_pair eval_v_elims subst_vv.simps by metis
+ hence "(i ( x \<mapsto> s )) \<lbrakk> v1 \<rbrakk> ~ s1 \<and> (i ( x \<mapsto> s )) \<lbrakk> v2 \<rbrakk> ~ s2" using V_pair by metis
+ thus ?case using eval_v_pairI subst_vv.simps * V_pair by metis
+next
+ case (V_cons tyid dc v1)
+ then obtain s1 where "eval_v i (v1[x::=v]\<^sub>v\<^sub>v) s1" using eval_v_elims subst_vv.simps by metis
+ thus ?case using eval_v_consI V_cons
+ by (metis eval_v_elims subst_vv.simps)
+next
+ case (V_consp tyid dc b1 v1)
+
+ then obtain s1 where *:"eval_v i (v1[x::=v]\<^sub>v\<^sub>v) s1 \<and> s' = SConsp tyid dc b1 s1" using eval_v_elims subst_vv.simps by metis
+ hence "i ( x \<mapsto> s ) \<lbrakk> v1 \<rbrakk> ~ s1" using V_consp by metis
+ thus ?case using * eval_v_conspI by metis
+qed
+
+lemma subst_e_eval_v[simp]:
+ fixes y::x and e::ce and v::v and e'::ce
+ assumes "i \<lbrakk> e' \<rbrakk> ~ s'" and "e'=(e[y::=v]\<^sub>c\<^sub>e\<^sub>v)" and "i \<lbrakk> v \<rbrakk> ~ s"
+ shows "(i ( y \<mapsto> s )) \<lbrakk> e \<rbrakk> ~ s'"
+ using assms proof(induct arbitrary: e rule: eval_e.induct)
+ case (eval_e_valI i v1 sv)
+ then obtain v1' where *:"e = CE_val v1' \<and> v1 = v1'[y::=v]\<^sub>v\<^sub>v"
+ using assms by(nominal_induct e rule:ce.strong_induct,simp+)
+ hence "eval_v i (v1'[y::=v]\<^sub>v\<^sub>v) sv" using eval_e_valI by simp
+ hence "eval_v (i ( y \<mapsto> s )) v1' sv" using subst_v_eval_v eval_e_valI by simp
+ then show ?case using RCLogic.eval_e_valI * by meson
+next
+ case (eval_e_plusI i v1 n1 v2 n2)
+ then obtain v1' and v2' where *:"e = CE_op Plus v1' v2' \<and> v1 = v1'[y::=v]\<^sub>c\<^sub>e\<^sub>v \<and> v2 = v2'[y::=v]\<^sub>c\<^sub>e\<^sub>v"
+ using assms by(nominal_induct e rule:ce.strong_induct,simp+)
+ hence "eval_e i (v1'[y::=v]\<^sub>c\<^sub>e\<^sub>v) (SNum n1) \<and> eval_e i (v2'[y::=v]\<^sub>c\<^sub>e\<^sub>v) (SNum n2)" using eval_e_plusI by simp
+ hence "eval_e (i ( y \<mapsto> s )) v1' (SNum n1) \<and> eval_e (i ( y \<mapsto> s )) v2' (SNum n2)" using subst_v_eval_v eval_e_plusI
+ using "local.*" by blast
+ then show ?case using RCLogic.eval_e_plusI * by meson
+next
+ case (eval_e_leqI i v1 n1 v2 n2)
+ then obtain v1' and v2' where *:"e = CE_op LEq v1' v2' \<and> v1 = v1'[y::=v]\<^sub>c\<^sub>e\<^sub>v \<and> v2 = v2'[y::=v]\<^sub>c\<^sub>e\<^sub>v"
+ using assms by(nominal_induct e rule:ce.strong_induct,simp+)
+ hence "eval_e i (v1'[y::=v]\<^sub>c\<^sub>e\<^sub>v) (SNum n1) \<and> eval_e i (v2'[y::=v]\<^sub>c\<^sub>e\<^sub>v) (SNum n2)" using eval_e_leqI by simp
+ hence "eval_e (i ( y \<mapsto> s )) v1' (SNum n1) \<and> eval_e (i ( y \<mapsto> s )) v2' (SNum n2)" using subst_v_eval_v eval_e_leqI
+ using * by blast
+ then show ?case using RCLogic.eval_e_leqI * by meson
+next
+ case (eval_e_eqI i v1 n1 v2 n2)
+ then obtain v1' and v2' where *:"e = CE_op Eq v1' v2' \<and> v1 = v1'[y::=v]\<^sub>c\<^sub>e\<^sub>v \<and> v2 = v2'[y::=v]\<^sub>c\<^sub>e\<^sub>v"
+ using assms by(nominal_induct e rule:ce.strong_induct,simp+)
+ hence "eval_e i (v1'[y::=v]\<^sub>c\<^sub>e\<^sub>v) n1 \<and> eval_e i (v2'[y::=v]\<^sub>c\<^sub>e\<^sub>v) n2" using eval_e_eqI by simp
+ hence "eval_e (i ( y \<mapsto> s )) v1' n1 \<and> eval_e (i ( y \<mapsto> s )) v2' n2" using subst_v_eval_v eval_e_eqI
+ using * by blast
+ then show ?case using RCLogic.eval_e_eqI * by meson
+next
+ case (eval_e_fstI i v1 s1 s2)
+ then obtain v1' and v2' where *:"e = CE_fst v1' \<and> v1 = v1'[y::=v]\<^sub>c\<^sub>e\<^sub>v"
+ using assms by(nominal_induct e rule:ce.strong_induct,simp+)
+ hence "eval_e i (v1'[y::=v]\<^sub>c\<^sub>e\<^sub>v) (SPair s1 s2)" using eval_e_fstI by simp
+ hence "eval_e (i ( y \<mapsto> s )) v1' (SPair s1 s2)" using eval_e_fstI * by metis
+ then show ?case using RCLogic.eval_e_fstI * by meson
+next
+ case (eval_e_sndI i v1 s1 s2)
+ then obtain v1' and v2' where *:"e = CE_snd v1' \<and> v1 = v1'[y::=v]\<^sub>c\<^sub>e\<^sub>v"
+ using assms by(nominal_induct e rule:ce.strong_induct,simp+)
+ hence "eval_e i (v1'[y::=v]\<^sub>c\<^sub>e\<^sub>v) (SPair s1 s2)" using eval_e_sndI by simp
+ hence "eval_e (i ( y \<mapsto> s )) v1' (SPair s1 s2)" using subst_v_eval_v eval_e_sndI * by blast
+ then show ?case using RCLogic.eval_e_sndI * by meson
+next
+ case (eval_e_concatI i v1 bv1 v2 bv2)
+ then obtain v1' and v2' where *:"e = CE_concat v1' v2' \<and> v1 = v1'[y::=v]\<^sub>c\<^sub>e\<^sub>v \<and> v2 = v2'[y::=v]\<^sub>c\<^sub>e\<^sub>v"
+ using assms by(nominal_induct e rule:ce.strong_induct,simp+)
+ hence "eval_e i (v1'[y::=v]\<^sub>c\<^sub>e\<^sub>v) (SBitvec bv1) \<and> eval_e i (v2'[y::=v]\<^sub>c\<^sub>e\<^sub>v) (SBitvec bv2)" using eval_e_concatI by simp
+ moreover hence "eval_e (i ( y \<mapsto> s )) v1' (SBitvec bv1) \<and> eval_e (i ( y \<mapsto> s )) v2' (SBitvec bv2)"
+ using subst_v_eval_v eval_e_concatI * by blast
+ ultimately show ?case using RCLogic.eval_e_concatI * eval_v_uniqueness by (metis eval_e_concatI.hyps(1))
+next
+ case (eval_e_lenI i v1 bv)
+ then obtain v1' where *:"e = CE_len v1' \<and> v1 = v1'[y::=v]\<^sub>c\<^sub>e\<^sub>v"
+ using assms by(nominal_induct e rule:ce.strong_induct,simp+)
+ hence "eval_e i (v1'[y::=v]\<^sub>c\<^sub>e\<^sub>v) (SBitvec bv)" using eval_e_lenI by simp
+ hence "eval_e (i ( y \<mapsto> s )) v1' (SBitvec bv)" using subst_v_eval_v eval_e_lenI * by blast
+ then show ?case using RCLogic.eval_e_lenI * by meson
+qed
+
+lemma subst_c_eval_v[simp]:
+ fixes v::v and c :: c
+ assumes "i \<lbrakk> v \<rbrakk> ~ s" and "i \<lbrakk> c[x::=v]\<^sub>c\<^sub>v \<rbrakk> ~ s1" and
+ "(i ( x \<mapsto> s)) \<lbrakk> c \<rbrakk> ~ s2"
+ shows "s1 = s2"
+ using assms proof(nominal_induct c arbitrary: s1 s2 rule: c.strong_induct)
+ case C_true
+ hence "s1 = True \<and> s2 = True" using eval_c_elims subst_cv.simps by auto
+ then show ?case by auto
+next
+ case C_false
+ hence "s1 = False \<and> s2 = False" using eval_c_elims subst_cv.simps by metis
+ then show ?case by auto
+next
+ case (C_conj c1 c2)
+ hence *:"eval_c i (c1[x::=v]\<^sub>c\<^sub>v AND c2[x::=v]\<^sub>c\<^sub>v) s1" using subst_cv.simps by auto
+ then obtain s11 and s12 where "(s1 = (s11 \<and> s12)) \<and> eval_c i c1[x::=v]\<^sub>c\<^sub>v s11 \<and> eval_c i c2[x::=v]\<^sub>c\<^sub>v s12" using
+ eval_c_elims(3) by metis
+ moreover obtain s21 and s22 where "eval_c (i ( x \<mapsto> s)) c1 s21 \<and> eval_c (i ( x \<mapsto> s)) c2 s22 \<and> (s2 = (s21 \<and> s22))" using
+ eval_c_elims(3) C_conj by metis
+ ultimately show ?case using C_conj by (meson eval_c_elims)
+next
+ case (C_disj c1 c2)
+ hence *:"eval_c i (c1[x::=v]\<^sub>c\<^sub>v OR c2[x::=v]\<^sub>c\<^sub>v) s1" using subst_cv.simps by auto
+ then obtain s11 and s12 where "(s1 = (s11 \<or> s12)) \<and> eval_c i c1[x::=v]\<^sub>c\<^sub>v s11 \<and> eval_c i c2[x::=v]\<^sub>c\<^sub>v s12" using
+ eval_c_elims(4) by metis
+ moreover obtain s21 and s22 where "eval_c (i ( x \<mapsto> s)) c1 s21 \<and> eval_c (i ( x \<mapsto> s)) c2 s22 \<and> (s2 = (s21 \<or> s22))" using
+ eval_c_elims(4) C_disj by metis
+ ultimately show ?case using C_disj by (meson eval_c_elims)
+next
+ case (C_not c1)
+ then obtain s11 where "(s1 = (\<not> s11)) \<and> eval_c i c1[x::=v]\<^sub>c\<^sub>v s11" using
+ eval_c_elims(6) by (metis subst_cv.simps(7))
+ moreover obtain s21 where "eval_c (i ( x \<mapsto> s)) c1 s21 \<and> (s2 = (\<not>s21))" using
+ eval_c_elims(6) C_not by metis
+ ultimately show ?case using C_not by (meson eval_c_elims)
+next
+ case (C_imp c1 c2)
+ hence *:"eval_c i (c1[x::=v]\<^sub>c\<^sub>v IMP c2[x::=v]\<^sub>c\<^sub>v) s1" using subst_cv.simps by auto
+ then obtain s11 and s12 where "(s1 = (s11 \<longrightarrow> s12)) \<and> eval_c i c1[x::=v]\<^sub>c\<^sub>v s11 \<and> eval_c i c2[x::=v]\<^sub>c\<^sub>v s12" using
+ eval_c_elims(5) by metis
+ moreover obtain s21 and s22 where "eval_c (i ( x \<mapsto> s)) c1 s21 \<and> eval_c (i ( x \<mapsto> s)) c2 s22 \<and> (s2 = (s21 \<longrightarrow> s22))" using
+ eval_c_elims(5) C_imp by metis
+ ultimately show ?case using C_imp by (meson eval_c_elims)
+next
+ case (C_eq e1 e2)
+ hence *:"eval_c i (e1[x::=v]\<^sub>c\<^sub>e\<^sub>v == e2[x::=v]\<^sub>c\<^sub>e\<^sub>v) s1" using subst_cv.simps by auto
+ then obtain s11 and s12 where "(s1 = (s11 = s12)) \<and> eval_e i (e1[x::=v]\<^sub>c\<^sub>e\<^sub>v) s11 \<and> eval_e i (e2[x::=v]\<^sub>c\<^sub>e\<^sub>v) s12" using
+ eval_c_elims(7) by metis
+ moreover obtain s21 and s22 where "eval_e (i ( x \<mapsto> s)) e1 s21 \<and> eval_e (i ( x \<mapsto> s)) e2 s22 \<and> (s2 = (s21 = s22))" using
+ eval_c_elims(7) C_eq by metis
+ ultimately show ?case using C_eq subst_e_eval_v by (metis eval_e_uniqueness)
+qed
+
+lemma wfI_upd:
+ assumes "wfI \<Theta> \<Gamma> i" and "wfRCV \<Theta> s b" and "wfG \<Theta> B ((x, b, c) #\<^sub>\<Gamma> \<Gamma>)"
+ shows "wfI \<Theta> ((x, b, c) #\<^sub>\<Gamma> \<Gamma>) (i(x \<mapsto> s))"
+proof(subst wfI_def,rule)
+ fix xa
+ assume as:"xa \<in> toSet ((x, b, c) #\<^sub>\<Gamma> \<Gamma>)"
+
+ then obtain x1::x and b1::b and c1::c where xa: "xa = (x1,b1,c1)" using toSet.simps
+ using prod_cases3 by blast
+
+ have "\<exists>sa. Some sa = (i(x \<mapsto> s)) x1 \<and> wfRCV \<Theta> sa b1" proof(cases "x=x1")
+ case True
+ hence "b=b1" using as xa wfG_unique assms by metis
+ hence "Some s = (i(x \<mapsto> s)) x1 \<and> wfRCV \<Theta> s b1" using assms True by simp
+ then show ?thesis by auto
+ next
+ case False
+ hence "(x1,b1,c1) \<in> toSet \<Gamma>" using xa as by auto
+ then obtain sa where "Some sa = i x1 \<and> wfRCV \<Theta> sa b1" using assms wfI_def as xa by auto
+ hence "Some sa = (i(x \<mapsto> s)) x1 \<and> wfRCV \<Theta> sa b1" using False by auto
+ then show ?thesis by auto
+ qed
+
+ thus "case xa of (xa, ba, ca) \<Rightarrow> \<exists>sa. Some sa = (i(x \<mapsto> s)) xa \<and> wfRCV \<Theta> sa ba" using xa by auto
+qed
+
+lemma wfI_upd_full:
+ fixes v::v
+ assumes "wfI \<Theta> G i" and "G = ((\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v)@\<Gamma>)" and "wfRCV \<Theta> s b" and "wfG \<Theta> B (\<Gamma>'@((x,b,c)#\<^sub>\<Gamma>\<Gamma>))" and "\<Theta> ; B ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b"
+ shows "wfI \<Theta> (\<Gamma>'@((x,b,c)#\<^sub>\<Gamma>\<Gamma>)) (i(x \<mapsto> s))"
+proof(subst wfI_def,rule)
+ fix xa
+ assume as:"xa \<in> toSet (\<Gamma>'@((x,b,c)#\<^sub>\<Gamma>\<Gamma>))"
+
+ then obtain x1::x and b1::b and c1::c where xa: "xa = (x1,b1,c1)" using toSet.simps
+ using prod_cases3 by blast
+
+ have "\<exists>sa. Some sa = (i(x \<mapsto> s)) x1 \<and> wfRCV \<Theta> sa b1"
+ proof(cases "x=x1")
+ case True
+ hence "b=b1" using as xa wfG_unique_full assms by metis
+ hence "Some s = (i(x \<mapsto> s)) x1 \<and> wfRCV \<Theta> s b1" using assms True by simp
+ then show ?thesis by auto
+ next
+ case False
+ hence "(x1,b1,c1) \<in> toSet (\<Gamma>'@\<Gamma>)" using as xa by auto
+ then obtain c1' where "(x1,b1,c1') \<in> toSet (\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v@\<Gamma>)" using xa as wfG_member_subst assms False by metis
+ then obtain sa where "Some sa = i x1 \<and> wfRCV \<Theta> sa b1" using assms wfI_def as xa by blast
+ hence "Some sa = (i(x \<mapsto> s)) x1 \<and> wfRCV \<Theta> sa b1" using False by auto
+ then show ?thesis by auto
+ qed
+
+ thus "case xa of (xa, ba, ca) \<Rightarrow> \<exists>sa. Some sa = (i(x \<mapsto> s)) xa \<and> wfRCV \<Theta> sa ba" using xa by auto
+qed
+
+lemma subst_c_satis[simp]:
+ fixes v::v
+ assumes "i \<lbrakk> v \<rbrakk> ~ s" and "wfC \<Theta> B ((x,b,c')#\<^sub>\<Gamma>\<Gamma>) c" and "wfI \<Theta> \<Gamma> i" and "\<Theta> ; B ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b"
+ shows "i \<Turnstile> (c[x::=v]\<^sub>c\<^sub>v) \<longleftrightarrow> (i ( x \<mapsto> s)) \<Turnstile> c"
+proof -
+ have "wfI \<Theta> ((x, b, c') #\<^sub>\<Gamma> \<Gamma>) (i(x \<mapsto> s))" using wfI_upd assms wfC_wf eval_v_base by blast
+ then obtain s1 where s1:"eval_c (i(x \<mapsto> s)) c s1" using eval_c_exist[of \<Theta> "((x,b,c')#\<^sub>\<Gamma>\<Gamma>)" "(i ( x \<mapsto> s))" B c ] assms by auto
+
+ have "\<Theta> ; B ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f c[x::=v]\<^sub>c\<^sub>v" using wf_subst1(2)[OF assms(2) _ assms(4) , of GNil x ] subst_gv.simps by simp
+ then obtain s2 where s2:"eval_c i c[x::=v]\<^sub>c\<^sub>v s2" using eval_c_exist[of \<Theta> "\<Gamma>" "i" B "c[x::=v]\<^sub>c\<^sub>v" ] assms by auto
+
+ show ?thesis using s1 s2 subst_c_eval_v[OF assms(1) s2 s1] is_satis.cases
+ using eval_c_uniqueness is_satis.simps by auto
+qed
+
+text \<open> Key theorem telling us we can replace a substitution with an update to the valuation \<close>
+lemma subst_c_satis_full:
+ fixes v::v and \<Gamma>'::\<Gamma>
+ assumes "i \<lbrakk> v \<rbrakk> ~ s" and "wfC \<Theta> B (\<Gamma>'@((x,b,c')#\<^sub>\<Gamma>\<Gamma>)) c" and "wfI \<Theta> ((\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v)@\<Gamma>) i" and "\<Theta> ; B ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b"
+ shows "i \<Turnstile> (c[x::=v]\<^sub>c\<^sub>v) \<longleftrightarrow> (i ( x \<mapsto> s)) \<Turnstile> c"
+proof -
+ have "wfI \<Theta> (\<Gamma>'@((x, b, c')) #\<^sub>\<Gamma> \<Gamma>) (i(x \<mapsto> s))" using wfI_upd_full assms wfC_wf eval_v_base wfI_suffix wfI_def wfV_wf by fast
+ then obtain s1 where s1:"eval_c (i(x \<mapsto> s)) c s1" using eval_c_exist[of \<Theta> "(\<Gamma>'@(x,b,c')#\<^sub>\<Gamma>\<Gamma>)" "(i ( x \<mapsto> s))" B c ] assms by auto
+
+ have "\<Theta> ; B ; ((\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v)@\<Gamma>) \<turnstile>\<^sub>w\<^sub>f c[x::=v]\<^sub>c\<^sub>v" using wbc_subst assms by auto
+
+ then obtain s2 where s2:"eval_c i c[x::=v]\<^sub>c\<^sub>v s2" using eval_c_exist[of \<Theta> "((\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v)@\<Gamma>)" "i" B "c[x::=v]\<^sub>c\<^sub>v" ] assms by auto
+
+ show ?thesis using s1 s2 subst_c_eval_v[OF assms(1) s2 s1] is_satis.cases
+ using eval_c_uniqueness is_satis.simps by auto
+qed
+
+section \<open>Validity\<close>
+
+lemma validI[intro]:
+ fixes c::c
+ assumes "wfC P B G c" and "\<forall>i. P ; G \<turnstile> i \<and> i \<Turnstile> G \<longrightarrow> i \<Turnstile> c"
+ shows "P ; B ; G \<Turnstile> c"
+ using assms valid.simps by presburger
+
+lemma valid_g_wf:
+ fixes c::c and G::\<Gamma>
+ assumes "P ; B ; G \<Turnstile> c"
+ shows "P ; B \<turnstile>\<^sub>w\<^sub>f G"
+ using assms wfC_wf valid.simps by blast
+
+lemma valid_reflI [intro]:
+ fixes b::b
+ assumes "P ; B ; ((x,b,c1)#\<^sub>\<Gamma>G) \<turnstile>\<^sub>w\<^sub>f c1" and "c1 = c2"
+ shows "P ; B ; ((x,b,c1)#\<^sub>\<Gamma>G) \<Turnstile> c2"
+ using satis_reflI assms by simp
+
+subsection \<open>Weakening and Strengthening\<close>
+
+text \<open> Adding to the domain of a valuation doesn't change the result \<close>
+
+lemma eval_v_weakening:
+ fixes c::v and B::"bv fset"
+ assumes "i = i'|` d" and "supp c \<subseteq> atom ` d \<union> supp B " and "i \<lbrakk> c \<rbrakk> ~ s"
+ shows "i' \<lbrakk> c \<rbrakk> ~ s"
+ using assms proof(nominal_induct c arbitrary:s rule: v.strong_induct)
+ case (V_lit x)
+ then show ?case using eval_v_elims eval_v_litI by metis
+next
+ case (V_var x)
+ have "atom x \<in> atom ` d" using x_not_in_b_set[of x B] assms v.supp(2) supp_at_base
+ proof -
+ show ?thesis
+ by (metis UnE V_var.prems(2) \<open>atom x \<notin> supp B\<close> singletonI subset_iff supp_at_base v.supp(2)) (* 46 ms *)
+ qed
+ moreover have "Some s = i x" using assms eval_v_elims(2)
+ using V_var.prems(3) by blast
+ hence "Some s= i' x" using assms insert_subset restrict_in
+ proof -
+ show ?thesis
+ by (metis (no_types) \<open>i = i' |` d\<close> \<open>Some s = i x\<close> atom_eq_iff calculation imageE restrict_in) (* 31 ms *)
+ qed
+ thus ?case using eval_v.eval_v_varI by simp
+
+next
+ case (V_pair v1 v2)
+ then show ?case using eval_v_elims(3) eval_v_pairI v.supp
+ by (metis assms le_sup_iff)
+next
+ case (V_cons dc v1)
+ then show ?case using eval_v_elims(4) eval_v_consI v.supp
+ by (metis assms le_sup_iff)
+next
+ case (V_consp tyid dc b1 v1)
+
+ then obtain sv1 where *:"i \<lbrakk> v1 \<rbrakk> ~ sv1 \<and> s = SConsp tyid dc b1 sv1" using eval_v_elims by metis
+ hence "i' \<lbrakk> v1 \<rbrakk> ~ sv1" using V_consp by auto
+ then show ?case using * eval_v_conspI v.supp eval_v.simps assms le_sup_iff by metis
+qed
+
+lemma eval_v_restrict:
+ fixes c::v and B::"bv fset"
+ assumes "i = i' |` d" and "supp c \<subseteq> atom ` d \<union> supp B " and "i' \<lbrakk> c \<rbrakk> ~ s"
+ shows "i \<lbrakk> c \<rbrakk> ~ s"
+ using assms proof(nominal_induct c arbitrary:s rule: v.strong_induct)
+ case (V_lit x)
+ then show ?case using eval_v_elims eval_v_litI by metis
+next
+ case (V_var x)
+ have "atom x \<in> atom ` d" using x_not_in_b_set[of x B] assms v.supp(2) supp_at_base
+ proof -
+ show ?thesis
+ by (metis UnE V_var.prems(2) \<open>atom x \<notin> supp B\<close> singletonI subset_iff supp_at_base v.supp(2)) (* 46 ms *)
+ qed
+ moreover have "Some s = i' x" using assms eval_v_elims(2)
+ using V_var.prems(3) by blast
+ hence "Some s= i x" using assms insert_subset restrict_in
+ proof -
+ show ?thesis
+ by (metis (no_types) \<open>i = i' |` d\<close> \<open>Some s = i' x\<close> atom_eq_iff calculation imageE restrict_in) (* 31 ms *)
+ qed
+ thus ?case using eval_v.eval_v_varI by simp
+next
+ case (V_pair v1 v2)
+ then show ?case using eval_v_elims(3) eval_v_pairI v.supp
+ by (metis assms le_sup_iff)
+next
+ case (V_cons dc v1)
+ then show ?case using eval_v_elims(4) eval_v_consI v.supp
+ by (metis assms le_sup_iff)
+next
+ case (V_consp tyid dc b1 v1)
+ then obtain sv1 where *:"i' \<lbrakk> v1 \<rbrakk> ~ sv1 \<and> s = SConsp tyid dc b1 sv1" using eval_v_elims by metis
+ hence "i \<lbrakk> v1 \<rbrakk> ~ sv1" using V_consp by auto
+ then show ?case using * eval_v_conspI v.supp eval_v.simps assms le_sup_iff by metis
+qed
+
+lemma eval_e_weakening:
+ fixes e::ce and B::"bv fset"
+ assumes "i \<lbrakk> e \<rbrakk> ~ s" and "i = i' |` d" and "supp e \<subseteq> atom ` d \<union> supp B "
+ shows "i' \<lbrakk> e \<rbrakk> ~ s"
+ using assms proof(induct rule: eval_e.induct)
+ case (eval_e_valI i v sv)
+ then show ?case using ce.supp eval_e.intros
+ using eval_v_weakening by auto
+next
+ case (eval_e_plusI i v1 n1 v2 n2)
+ then show ?case using ce.supp eval_e.intros
+ using eval_v_weakening by auto
+next
+ case (eval_e_leqI i v1 n1 v2 n2)
+ then show ?case using ce.supp eval_e.intros
+ using eval_v_weakening by auto
+next
+ case (eval_e_eqI i v1 n1 v2 n2)
+ then show ?case using ce.supp eval_e.intros
+ using eval_v_weakening by auto
+next
+ case (eval_e_fstI i v v1 v2)
+ then show ?case using ce.supp eval_e.intros
+ using eval_v_weakening by metis
+next
+ case (eval_e_sndI i v v1 v2)
+ then show ?case using ce.supp eval_e.intros
+ using eval_v_weakening by metis
+next
+ case (eval_e_concatI i v1 bv2 v2 bv1)
+ then show ?case using ce.supp eval_e.intros
+ using eval_v_weakening by auto
+next
+ case (eval_e_lenI i v bv)
+ then show ?case using ce.supp eval_e.intros
+ using eval_v_weakening by auto
+qed
+
+lemma eval_e_restrict :
+ fixes e::ce and B::"bv fset"
+ assumes "i' \<lbrakk> e \<rbrakk> ~ s" and "i = i' |` d" and "supp e \<subseteq> atom ` d \<union> supp B "
+ shows "i \<lbrakk> e \<rbrakk> ~ s"
+ using assms proof(induct rule: eval_e.induct)
+ case (eval_e_valI i v sv)
+ then show ?case using ce.supp eval_e.intros
+ using eval_v_restrict by auto
+next
+ case (eval_e_plusI i v1 n1 v2 n2)
+ then show ?case using ce.supp eval_e.intros
+ using eval_v_restrict by auto
+next
+ case (eval_e_leqI i v1 n1 v2 n2)
+ then show ?case using ce.supp eval_e.intros
+ using eval_v_restrict by auto
+next
+ case (eval_e_eqI i v1 n1 v2 n2)
+ then show ?case using ce.supp eval_e.intros
+ using eval_v_restrict by auto
+next
+ case (eval_e_fstI i v v1 v2)
+ then show ?case using ce.supp eval_e.intros
+ using eval_v_restrict by metis
+next
+ case (eval_e_sndI i v v1 v2)
+ then show ?case using ce.supp eval_e.intros
+ using eval_v_restrict by metis
+next
+ case (eval_e_concatI i v1 bv2 v2 bv1)
+ then show ?case using ce.supp eval_e.intros
+ using eval_v_restrict by auto
+next
+ case (eval_e_lenI i v bv)
+ then show ?case using ce.supp eval_e.intros
+ using eval_v_restrict by auto
+qed
+
+lemma eval_c_i_weakening:
+ fixes c::c and B::"bv fset"
+ assumes "i \<lbrakk> c \<rbrakk> ~ s" and "i = i' |` d" and "supp c \<subseteq> atom ` d \<union> supp B"
+ shows "i' \<lbrakk> c \<rbrakk> ~ s"
+ using assms proof(induct rule:eval_c.induct)
+ case (eval_c_eqI i e1 sv1 e2 sv2)
+ then show ?case using eval_c.intros eval_e_weakening by auto
+qed(auto simp add: eval_c.intros)+
+
+lemma eval_c_i_restrict:
+ fixes c::c and B::"bv fset"
+ assumes "i' \<lbrakk> c \<rbrakk> ~ s" and "i = i' |` d" and "supp c \<subseteq> atom ` d \<union> supp B"
+ shows "i \<lbrakk> c \<rbrakk> ~ s"
+ using assms proof(induct rule:eval_c.induct)
+ case (eval_c_eqI i e1 sv1 e2 sv2)
+ then show ?case using eval_c.intros eval_e_restrict by auto
+qed(auto simp add: eval_c.intros)+
+
+lemma is_satis_i_weakening:
+ fixes c::c and B::"bv fset"
+ assumes "i = i' |` d" and "supp c \<subseteq> atom ` d \<union> supp B " and "i \<Turnstile> c"
+ shows "i' \<Turnstile> c"
+ using is_satis.simps eval_c_i_weakening[OF _ assms(1) assms(2) ]
+ using assms(3) by auto
+
+lemma is_satis_i_restrict:
+ fixes c::c and B::"bv fset"
+ assumes "i = i' |` d" and "supp c \<subseteq> atom ` d \<union> supp B" and "i' \<Turnstile> c"
+ shows "i \<Turnstile> c"
+ using is_satis.simps eval_c_i_restrict[OF _ assms(1) assms(2) ]
+ using assms(3) by auto
+
+lemma is_satis_g_restrict1:
+ fixes \<Gamma>'::\<Gamma> and \<Gamma>::\<Gamma>
+ assumes "toSet \<Gamma> \<subseteq> toSet \<Gamma>'" and "i \<Turnstile> \<Gamma>'"
+ shows "i \<Turnstile> \<Gamma>"
+ using assms proof(induct \<Gamma> rule: \<Gamma>.induct)
+ case GNil
+ then show ?case by auto
+next
+ case (GCons xbc G)
+ obtain x and b and c::c where xbc: "xbc=(x,b,c)"
+ using prod_cases3 by blast
+ hence "i \<Turnstile> G" using GCons by auto
+ moreover have "i \<Turnstile> c" using GCons
+ is_satis_iff toSet.simps subset_iff
+ using xbc by blast
+ ultimately show ?case using is_satis_g.simps GCons
+ by (simp add: xbc)
+qed
+
+lemma is_satis_g_restrict2:
+ fixes \<Gamma>'::\<Gamma> and \<Gamma>::\<Gamma>
+ assumes "i \<Turnstile> \<Gamma>" and "i' = i |` d" and "atom_dom \<Gamma> \<subseteq> atom ` d" and "\<Theta> ; B \<turnstile>\<^sub>w\<^sub>f \<Gamma>"
+ shows "i' \<Turnstile> \<Gamma>"
+ using assms proof(induct \<Gamma> rule: \<Gamma>_induct)
+ case GNil
+ then show ?case by auto
+next
+ case (GCons x b c G)
+
+ hence "i' \<Turnstile> G" proof -
+ have "i \<Turnstile> G" using GCons by simp
+ moreover have "atom_dom G \<subseteq> atom ` d" using GCons by simp
+ ultimately show ?thesis using GCons wfG_cons2 by blast
+ qed
+
+ moreover have "i' \<Turnstile> c" proof -
+ have "i \<Turnstile> c" using GCons by auto
+ moreover have "\<Theta> ; B ; (x, b, TRUE) #\<^sub>\<Gamma> G \<turnstile>\<^sub>w\<^sub>f c" using wfG_wfC GCons by simp
+ moreover hence "supp c \<subseteq> atom ` d \<union> supp B" using wfC_supp GCons atom_dom_eq by blast
+ ultimately show ?thesis using is_satis_i_restrict[of i' i d c] GCons by simp
+ qed
+
+ ultimately show ?case by auto
+qed
+
+lemma is_satis_g_restrict:
+ fixes \<Gamma>'::\<Gamma> and \<Gamma>::\<Gamma>
+ assumes "toSet \<Gamma> \<subseteq> toSet \<Gamma>'" and "i' \<Turnstile> \<Gamma>'" and "i = i' |` (fst ` toSet \<Gamma>)" and "\<Theta> ; B \<turnstile>\<^sub>w\<^sub>f \<Gamma>"
+ shows "i \<Turnstile> \<Gamma>"
+ using assms is_satis_g_restrict1[OF assms(1) assms(2)] is_satis_g_restrict2[OF _ assms(3)] by simp
+
+subsection \<open>Updating valuation\<close>
+
+lemma is_satis_c_i_upd:
+ fixes c::c
+ assumes "atom x \<sharp> c" and "i \<Turnstile> c"
+ shows "((i ( x \<mapsto>s))) \<Turnstile> c"
+ using assms eval_c_i_upd is_satis.simps by simp
+
+lemma is_satis_g_i_upd:
+ fixes G::\<Gamma>
+ assumes "atom x \<sharp> G" and "i \<Turnstile> G"
+ shows "((i ( x \<mapsto>s))) \<Turnstile> G"
+ using assms proof(induct G rule: \<Gamma>_induct)
+ case GNil
+ then show ?case by auto
+next
+ case (GCons x' b' c' G')
+
+ hence *:"atom x \<sharp> G' \<and> atom x \<sharp> c'"
+ using fresh_def fresh_GCons GCons by force
+ moreover hence "is_satis ((i ( x \<mapsto>s))) c'"
+ using is_satis_c_i_upd GCons is_satis_g.simps by auto
+ moreover have " is_satis_g (i(x \<mapsto> s)) G'" using GCons * by fastforce
+ ultimately show ?case
+ using GCons is_satis_g.simps(2) by metis
+qed
+
+lemma valid_weakening:
+ assumes "\<Theta> ; B ; \<Gamma> \<Turnstile> c" and "toSet \<Gamma> \<subseteq> toSet \<Gamma>'" and "wfG \<Theta> B \<Gamma>'"
+ shows "\<Theta> ; B ; \<Gamma>' \<Turnstile> c"
+proof -
+ have "wfC \<Theta> B \<Gamma> c" using assms valid.simps by auto
+ hence sp: "supp c \<subseteq> atom `(fst `toSet \<Gamma>) \<union> supp B" using wfX_wfY wfG_elims
+ using atom_dom.simps dom.simps wf_supp(2) by metis
+ have wfg: "wfG \<Theta> B \<Gamma>" using assms valid.simps wfC_wf by auto
+
+ moreover have a1: "(\<forall>i. wfI \<Theta> \<Gamma>' i \<and> i \<Turnstile> \<Gamma>' \<longrightarrow> i \<Turnstile> c)" proof(rule allI, rule impI)
+ fix i
+ assume as: "wfI \<Theta> \<Gamma>' i \<and> i \<Turnstile> \<Gamma>'"
+ hence as1: "fst ` toSet \<Gamma> \<subseteq> dom i" using assms wfI_domi by blast
+ obtain i' where idash: "i' = restrict_map i (fst `toSet \<Gamma>)" by blast
+ hence as2: "dom i' = (fst `toSet \<Gamma>)" using dom_restrict as1 by auto
+
+ have id2: "\<Theta> ; \<Gamma> \<turnstile> i' \<and> i' \<Turnstile> \<Gamma>" proof -
+ have "wfI \<Theta> \<Gamma> i'" using as2 wfI_restrict_weakening[of \<Theta> \<Gamma>' i i' \<Gamma>] as assms
+ using idash by blast
+ moreover have "i' \<Turnstile> \<Gamma>" using is_satis_g_restrict[OF assms(2)] wfg as idash by auto
+ ultimately show ?thesis using idash by auto
+ qed
+ hence "i' \<Turnstile> c" using assms valid.simps by auto
+ thus "i \<Turnstile> c" using assms valid.simps is_satis_i_weakening idash sp by blast
+ qed
+ moreover have "wfC \<Theta> B \<Gamma>' c" using wf_weakening assms valid.simps
+ by (meson wfg)
+ ultimately show ?thesis using assms valid.simps by auto
+qed
+
+lemma is_satis_g_suffix:
+ fixes G::\<Gamma>
+ assumes "i \<Turnstile> (G'@G)"
+ shows "i \<Turnstile> G"
+ using assms proof(induct G' rule:\<Gamma>.induct)
+ case GNil
+ then show ?case by auto
+next
+ case (GCons xbc x2)
+ obtain x and b and c::c where xbc: "xbc=(x,b,c)"
+ using prod_cases3 by blast
+ hence " i \<Turnstile> (xbc #\<^sub>\<Gamma> (x2 @ G))" using append_g.simps GCons by fastforce
+ then show ?case using is_satis_g.simps GCons xbc by blast
+qed
+
+lemma wfG_inside_valid2:
+ fixes x::x and \<Gamma>::\<Gamma> and c0::c and c0'::c
+ assumes "wfG \<Theta> B (\<Gamma>'@((x,b0,c0')#\<^sub>\<Gamma>\<Gamma>))" and
+ "\<Theta> ; B ; \<Gamma>'@(x,b0,c0)#\<^sub>\<Gamma>\<Gamma> \<Turnstile> c0'"
+ shows "wfG \<Theta> B (\<Gamma>'@((x,b0,c0)#\<^sub>\<Gamma>\<Gamma>))"
+proof -
+ have "wfG \<Theta> B (\<Gamma>'@(x,b0,c0)#\<^sub>\<Gamma>\<Gamma>)" using valid.simps wfC_wf assms by auto
+ thus ?thesis using wfG_replace_inside_full assms by auto
+qed
+
+lemma valid_trans:
+ assumes " \<Theta> ; \<B> ; \<Gamma> \<Turnstile> c0[z::=v]\<^sub>v" and " \<Theta> ; \<B> ; (z,b,c0)#\<^sub>\<Gamma>\<Gamma> \<Turnstile> c1" and "atom z \<sharp> \<Gamma>" and "wfV \<Theta> \<B> \<Gamma> v b"
+ shows " \<Theta> ; \<B> ; \<Gamma> \<Turnstile> c1[z::=v]\<^sub>v"
+proof -
+ have *:"wfC \<Theta> \<B> ((z,b,c0)#\<^sub>\<Gamma>\<Gamma>) c1" using valid.simps assms by auto
+ hence "wfC \<Theta> \<B> \<Gamma> (c1[z::=v]\<^sub>v)" using wf_subst1(2)[OF * , of GNil ] assms subst_gv.simps subst_v_c_def by force
+
+ moreover have "\<forall>i. wfI \<Theta> \<Gamma> i \<and> is_satis_g i \<Gamma> \<longrightarrow> is_satis i (c1[z::=v]\<^sub>v)"
+ proof(rule,rule)
+ fix i
+ assume as: "wfI \<Theta> \<Gamma> i \<and> is_satis_g i \<Gamma>"
+ then obtain sv where sv: "eval_v i v sv \<and> wfRCV \<Theta> sv b" using eval_v_exist assms by metis
+ hence "is_satis i (c0[z::=v]\<^sub>v)" using assms valid.simps as by metis
+ hence "is_satis (i(z \<mapsto> sv)) c0" using subst_c_satis sv as assms valid.simps wfC_wf wfG_elim2 subst_v_c_def by metis
+ moreover have "is_satis_g (i(z \<mapsto> sv)) \<Gamma>"
+ using is_satis_g_i_upd assms by (simp add: as)
+ ultimately have "is_satis_g (i(z \<mapsto> sv)) ((z,b,c0)#\<^sub>\<Gamma>\<Gamma>)"
+ using is_satis_g.simps by simp
+ moreover have "wfI \<Theta> ((z,b,c0)#\<^sub>\<Gamma>\<Gamma>) (i(z \<mapsto> sv))" using as wfI_upd sv assms valid.simps wfC_wf by metis
+ ultimately have "is_satis (i(z \<mapsto> sv)) c1" using assms valid.simps by auto
+ thus "is_satis i (c1[z::=v]\<^sub>v)" using subst_c_satis sv as assms valid.simps wfC_wf wfG_elim2 subst_v_c_def by metis
+ qed
+
+ ultimately show ?thesis using valid.simps by auto
+qed
+
+lemma valid_trans_full:
+ assumes "\<Theta> ; \<B> ; ((x, b, c1[z1::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) \<Turnstile> c2[z2::=V_var x]\<^sub>v" and
+ "\<Theta> ; \<B> ; ((x, b, c2[z2::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) \<Turnstile> c3[z3::=V_var x]\<^sub>v"
+ shows "\<Theta> ; \<B> ; ((x, b, c1[z1::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) \<Turnstile> c3[z3::=V_var x]\<^sub>v"
+ unfolding valid.simps proof
+ show "\<Theta> ; \<B> ; (x, b, c1[z1::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f c3[z3::=V_var x]\<^sub>v" using wf_trans valid.simps assms by metis
+
+ show "\<forall>i. ( wfI \<Theta> ((x, b, c1[z1::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) i \<and> (is_satis_g i ((x, b, c1[z1::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)) \<longrightarrow> (is_satis i (c3[z3::=V_var x]\<^sub>v)) ) "
+ proof(rule,rule)
+ fix i
+ assume as: "\<Theta> ; (x, b, c1[z1::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<turnstile> i \<and> i \<Turnstile> (x, b, c1[z1::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>"
+ have "i \<Turnstile> c2[z2::=V_var x]\<^sub>v" using is_satis_g.simps as assms by simp
+ moreover have "i \<Turnstile> \<Gamma>" using is_satis_g.simps as by simp
+ ultimately show "i \<Turnstile> c3[z3::=V_var x]\<^sub>v " using assms is_satis_g.simps valid.simps
+ by (metis append_g.simps(1) as wfI_replace_inside)
+ qed
+qed
+
+lemma eval_v_weakening_x:
+ fixes c::v
+ assumes "i' \<lbrakk> c \<rbrakk> ~ s" and "atom x \<sharp> c" and "i = i' (x \<mapsto> s')"
+ shows "i \<lbrakk> c \<rbrakk> ~ s"
+ using assms proof(induct rule: eval_v.induct)
+ case (eval_v_litI i l)
+ then show ?case using eval_v.intros by auto
+next
+ case (eval_v_varI sv i1 x1)
+ hence "x \<noteq> x1" using v.fresh fresh_at_base by auto
+ hence "i x1 = Some sv" using eval_v_varI by simp
+ then show ?case using eval_v.intros by auto
+next
+ case (eval_v_pairI i v1 s1 v2 s2)
+ then show ?case using eval_v.intros by auto
+next
+ case (eval_v_consI i v s tyid dc)
+ then show ?case using eval_v.intros by auto
+next
+ case (eval_v_conspI i v s tyid dc b)
+ then show ?case using eval_v.intros by auto
+qed
+
+lemma eval_e_weakening_x:
+ fixes c::ce
+ assumes "i' \<lbrakk> c \<rbrakk> ~ s" and "atom x \<sharp> c" and "i = i' (x \<mapsto> s')"
+ shows "i \<lbrakk> c \<rbrakk> ~ s"
+ using assms proof(induct rule: eval_e.induct)
+ case (eval_e_valI i v sv)
+ then show ?case using eval_v_weakening_x eval_e.intros ce.fresh by metis
+next
+ case (eval_e_plusI i v1 n1 v2 n2)
+ then show ?case using eval_v_weakening_x eval_e.intros ce.fresh by metis
+next
+ case (eval_e_leqI i v1 n1 v2 n2)
+ then show ?case using eval_v_weakening_x eval_e.intros ce.fresh by metis
+next
+ case (eval_e_eqI i v1 n1 v2 n2)
+ then show ?case using eval_v_weakening_x eval_e.intros ce.fresh by metis
+next
+ case (eval_e_fstI i v v1 v2)
+ then show ?case using eval_v_weakening_x eval_e.intros ce.fresh by metis
+next
+ case (eval_e_sndI i v v1 v2)
+ then show ?case using eval_v_weakening_x eval_e.intros ce.fresh by metis
+next
+ case (eval_e_concatI i v1 bv1 v2 bv2)
+ then show ?case using eval_v_weakening_x eval_e.intros ce.fresh by metis
+next
+ case (eval_e_lenI i v bv)
+ then show ?case using eval_v_weakening_x eval_e.intros ce.fresh by metis
+qed
+
+lemma eval_c_weakening_x:
+ fixes c::c
+ assumes "i' \<lbrakk> c \<rbrakk> ~ s" and "atom x \<sharp> c" and "i = i' (x \<mapsto> s')"
+ shows "i \<lbrakk> c \<rbrakk> ~ s"
+ using assms proof(induct rule: eval_c.induct)
+ case (eval_c_trueI i)
+ then show ?case using eval_c.intros by auto
+next
+ case (eval_c_falseI i)
+ then show ?case using eval_c.intros by auto
+next
+ case (eval_c_conjI i c1 b1 c2 b2)
+ then show ?case using eval_c.intros by auto
+next
+ case (eval_c_disjI i c1 b1 c2 b2)
+ then show ?case using eval_c.intros by auto
+next
+ case (eval_c_impI i c1 b1 c2 b2)
+ then show ?case using eval_c.intros by auto
+next
+ case (eval_c_notI i c b)
+ then show ?case using eval_c.intros by auto
+next
+ case (eval_c_eqI i e1 sv1 e2 sv2)
+ then show ?case using eval_e_weakening_x c.fresh eval_c.intros by metis
+qed
+
+lemma is_satis_weakening_x:
+ fixes c::c
+ assumes "i' \<Turnstile> c" and "atom x \<sharp> c" and "i = i' (x \<mapsto> s)"
+ shows "i \<Turnstile> c"
+ using eval_c_weakening_x assms is_satis.simps by simp
+
+lemma is_satis_g_weakening_x:
+ fixes G::\<Gamma>
+ assumes "i' \<Turnstile> G" and "atom x \<sharp> G" and "i = i' (x \<mapsto> s)"
+ shows "i \<Turnstile> G"
+ using assms proof(induct G rule: \<Gamma>_induct)
+ case GNil
+ then show ?case by auto
+next
+ case (GCons x' b' c' \<Gamma>')
+ hence "atom x \<sharp> c'" using fresh_GCons fresh_prodN by simp
+ moreover hence "i \<Turnstile> c'" using is_satis_weakening_x is_satis_g.simps(2) GCons by metis
+ then show ?case using is_satis_g.simps(2)[of i x' b' c' \<Gamma>'] GCons fresh_GCons by simp
+qed
+
+section \<open>Base Type Substitution\<close>
+
+text \<open>The idea of boxing is to take an smt val and its base type and at nodes in the smt val that correspond to type variables we
+wrap them in an SUt smt val node. Another way of looking at it is that s' where the node for the base type variable is an 'any node'.
+It is needed to prove subst\_b\_valid - the base-type variable substitution lemma for validity.
+
+The first @{text "rcl_val"} is the expanded form (has type with base-variables replaced with base-type terms)
+ ; the second is its corresponding form
+
+We only have one variable so we need to ensure that in all of the @{text "bs_boxed_BVarI"} cases, the s has the same
+ base type.
+
+For example is an SMT value is (SPair (SInt 1) (SBool true)) and it has sort (BPair (BVar x) BBool)[x::=BInt] then
+the boxed version is SPair (SUt (SInt 1)) (SBool true) and is has sort (BPair (BVar x) BBool). We need to do this
+so that we can obtain from a valuation i, that gives values like the first smt value, to a valuation i' that gives values like
+the second.
+\<close>
+
+inductive boxed_b :: "\<Theta> \<Rightarrow> rcl_val \<Rightarrow> b \<Rightarrow> bv \<Rightarrow> b \<Rightarrow> rcl_val \<Rightarrow> bool" ( " _ \<turnstile> _ ~ _ [ _ ::= _ ] \<setminus> _ " [50,50] 50) where
+ boxed_b_BVar1I: "\<lbrakk> bv = bv' ; wfRCV P s b \<rbrakk> \<Longrightarrow> boxed_b P s (B_var bv') bv b (SUt s)"
+| boxed_b_BVar2I: "\<lbrakk> bv \<noteq> bv'; wfRCV P s (B_var bv') \<rbrakk> \<Longrightarrow> boxed_b P s (B_var bv') bv b s"
+| boxed_b_BIntI:"wfRCV P s B_int \<Longrightarrow> boxed_b P s B_int _ _ s"
+| boxed_b_BBoolI:"wfRCV P s B_bool \<Longrightarrow> boxed_b P s B_bool _ _ s "
+| boxed_b_BUnitI:"wfRCV P s B_unit \<Longrightarrow> boxed_b P s B_unit _ _ s"
+| boxed_b_BPairI:"\<lbrakk> boxed_b P s1 b1 bv b s1' ; boxed_b P s2 b2 bv b s2' \<rbrakk> \<Longrightarrow> boxed_b P (SPair s1 s2) (B_pair b1 b2) bv b (SPair s1' s2')"
+
+| boxed_b_BConsI:"\<lbrakk>
+ AF_typedef tyid dclist \<in> set P;
+ (dc, \<lbrace> x : b | c \<rbrace>) \<in> set dclist ;
+ boxed_b P s1 b bv b' s1'
+ \<rbrakk> \<Longrightarrow>
+ boxed_b P (SCons tyid dc s1) (B_id tyid) bv b' (SCons tyid dc s1')"
+
+| boxed_b_BConspI:"\<lbrakk> AF_typedef_poly tyid bva dclist \<in> set P;
+ atom bva \<sharp> (b1,bv,b',s1,s1');
+ (dc, \<lbrace> x : b | c \<rbrace>) \<in> set dclist ;
+ boxed_b P s1 (b[bva::=b1]\<^sub>b\<^sub>b) bv b' s1'
+ \<rbrakk> \<Longrightarrow>
+ boxed_b P (SConsp tyid dc b1[bv::=b']\<^sub>b\<^sub>b s1) (B_app tyid b1) bv b' (SConsp tyid dc b1 s1')"
+
+| boxed_b_Bbitvec: "wfRCV P s B_bitvec \<Longrightarrow> boxed_b P s B_bitvec bv b s"
+
+equivariance boxed_b
+nominal_inductive boxed_b .
+
+inductive_cases boxed_b_elims:
+ "boxed_b P s (B_var bv) bv' b s'"
+ "boxed_b P s B_int bv b s'"
+ "boxed_b P s B_bool bv b s'"
+ "boxed_b P s B_unit bv b s'"
+ "boxed_b P s (B_pair b1 b2) bv b s'"
+ "boxed_b P s (B_id dc) bv b s'"
+ "boxed_b P s B_bitvec bv b s'"
+ "boxed_b P s (B_app dc b') bv b s'"
+
+lemma boxed_b_wfRCV:
+ assumes "boxed_b P s b bv b' s'" (*and "supp s = {}"*) and "\<turnstile>\<^sub>w\<^sub>f P"
+ shows "wfRCV P s b[bv::=b']\<^sub>b\<^sub>b \<and> wfRCV P s' b"
+ using assms proof(induct rule: boxed_b.inducts)
+ case (boxed_b_BVar1I bv bv' P s b )
+ then show ?case using wfRCV.intros by auto
+next
+ case (boxed_b_BVar2I bv bv' P s )
+ then show ?case using wfRCV.intros by auto
+next
+ case (boxed_b_BPairI P s1 b1 bv b s1' s2 b2 s2')
+ then show ?case using wfRCV.intros rcl_val.supp by simp
+next
+ case (boxed_b_BConsI tyid dclist P dc x b c s1 bv b' s1')
+ hence "supp b = {}" using wfTh_supp_b by metis
+ hence "b [ bv ::= b' ]\<^sub>b\<^sub>b = b" using fresh_def subst_b_b_def forget_subst[of "bv" b b'] by auto
+ hence " P \<turnstile> SCons tyid dc s1 : (B_id tyid)" using wfRCV.intros rcl_val.supp subst_bb.simps boxed_b_BConsI by metis
+ moreover have "P \<turnstile> SCons tyid dc s1' : B_id tyid" using boxed_b_BConsI
+ using wfRCV.intros rcl_val.supp subst_bb.simps boxed_b_BConsI by metis
+ ultimately show ?case using subst_bb.simps by metis
+next
+ case (boxed_b_BConspI tyid bva dclist P b1 bv b' s1 s1' dc x b c)
+
+ obtain bva2 and dclist2 where *:"AF_typedef_poly tyid bva dclist = AF_typedef_poly tyid bva2 dclist2 \<and>
+ atom bva2 \<sharp> (bv,(P, SConsp tyid dc b1[bv::=b']\<^sub>b\<^sub>b s1, B_app tyid b1[bv::=b']\<^sub>b\<^sub>b))"
+ using obtain_fresh_bv by metis
+
+ then obtain x2 and b2 and c2 where **:\<open>(dc, \<lbrace> x2 : b2 | c2 \<rbrace>) \<in> set dclist2\<close>
+ using boxed_b_BConspI td_lookup_eq_iff type_def.eq_iff by metis
+
+ have "P \<turnstile> SConsp tyid dc b1[bv::=b']\<^sub>b\<^sub>b s1 : (B_app tyid b1[bv::=b']\<^sub>b\<^sub>b)" proof
+ show 1: \<open>AF_typedef_poly tyid bva2 dclist2 \<in> set P\<close> using boxed_b_BConspI * by auto
+ show 2: \<open>(dc, \<lbrace> x2 : b2 | c2 \<rbrace>) \<in> set dclist2\<close> using boxed_b_BConspI using ** by simp
+
+ hence "atom bv \<sharp> b2" proof -
+ have "supp b2 \<subseteq> { atom bva2 }" using wfTh_poly_supp_b 1 2 boxed_b_BConspI by auto
+ moreover have "bv \<noteq> bva2" using * fresh_Pair fresh_at_base by metis
+ ultimately show ?thesis using fresh_def by force
+ qed
+ moreover have "b[bva::=b1]\<^sub>b\<^sub>b = b2[bva2::=b1]\<^sub>b\<^sub>b" using wfTh_typedef_poly_b_eq_iff * 2 boxed_b_BConspI by metis
+ ultimately show \<open> P \<turnstile> s1 : b2[bva2::=b1[bv::=b']\<^sub>b\<^sub>b]\<^sub>b\<^sub>b\<close> using boxed_b_BConspI subst_b_b_def subst_bb_commute by auto
+ show "atom bva2 \<sharp> (P, SConsp tyid dc b1[bv::=b']\<^sub>b\<^sub>b s1, B_app tyid b1[bv::=b']\<^sub>b\<^sub>b)" using * fresh_Pair by metis
+ qed
+
+ moreover have "P \<turnstile> SConsp tyid dc b1 s1' : B_app tyid b1" proof
+ show \<open>AF_typedef_poly tyid bva dclist \<in> set P\<close> using boxed_b_BConspI by auto
+ show \<open>(dc, \<lbrace> x : b | c \<rbrace>) \<in> set dclist\<close> using boxed_b_BConspI by auto
+ show \<open> P \<turnstile> s1' : b[bva::=b1]\<^sub>b\<^sub>b\<close> using boxed_b_BConspI by auto
+ have "atom bva \<sharp> P" using boxed_b_BConspI wfTh_fresh by metis
+ thus "atom bva \<sharp> (P, SConsp tyid dc b1 s1', B_app tyid b1)" using boxed_b_BConspI rcl_val.fresh b.fresh pure_fresh fresh_prodN by metis
+ qed
+
+ ultimately show ?case using subst_bb.simps by simp
+qed(auto)+
+
+lemma subst_b_var:
+ assumes "B_var bv2 = b[bv::=b']\<^sub>b\<^sub>b"
+ shows "(b = B_var bv \<and> b' = B_var bv2) \<or> (b=B_var bv2 \<and> bv \<noteq> bv2)"
+ using assms by(nominal_induct b rule: b.strong_induct,auto+)
+
+text \<open>Here the valuation i' is the conv wrap version of i. For every x in G, i' x is the conv wrap version of i x.
+The next lemma for a clearer explanation of what this is. i produces values of sort b[bv::=b'] and i' produces values of sort b\<close>
+
+inductive boxed_i :: "\<Theta> \<Rightarrow> \<Gamma> \<Rightarrow> b \<Rightarrow> bv \<Rightarrow> valuation \<Rightarrow> valuation \<Rightarrow> bool" ( " _ ; _ ; _ , _ \<turnstile> _ \<approx> _" [50,50] 50) where
+ boxed_i_GNilI: "\<Theta> ; GNil ; b , bv \<turnstile> i \<approx> i"
+| boxed_i_GConsI: "\<lbrakk> Some s = i x; boxed_b \<Theta> s b bv b' s' ; \<Theta> ; \<Gamma> ; b' , bv \<turnstile> i \<approx> i' \<rbrakk> \<Longrightarrow> \<Theta> ; ((x,b,c)#\<^sub>\<Gamma>\<Gamma>) ; b' , bv \<turnstile> i \<approx> (i'(x \<mapsto> s'))"
+equivariance boxed_i
+nominal_inductive boxed_i .
+
+inductive_cases boxed_i_elims:
+ "\<Theta> ;GNil ; b , bv \<turnstile> i \<approx> i'"
+ "\<Theta> ; ((x,b,c)#\<^sub>\<Gamma>\<Gamma>) ; b' , bv \<turnstile> i \<approx> i'"
+
+lemma wfRCV_poly_elims:
+ fixes tm::"'a::fs" and b::b
+ assumes "T \<turnstile> SConsp typid dc bdc s : b"
+ obtains bva dclist x1 b1 c1 where "b = B_app typid bdc \<and>
+ AF_typedef_poly typid bva dclist \<in> set T \<and> (dc, \<lbrace> x1 : b1 | c1 \<rbrace>) \<in> set dclist \<and> T \<turnstile> s : b1[bva::=bdc]\<^sub>b\<^sub>b \<and> atom bva \<sharp> tm"
+ using assms proof(nominal_induct "SConsp typid dc bdc s" b avoiding: tm rule:wfRCV.strong_induct)
+ case (wfRCV_BConsPI bv dclist \<Theta> x b c)
+ then show ?case by simp
+qed
+
+lemma boxed_b_ex:
+ assumes "wfRCV T s b[bv::=b']\<^sub>b\<^sub>b" and "wfTh T"
+ shows "\<exists>s'. boxed_b T s b bv b' s'"
+ using assms proof(nominal_induct s arbitrary: b rule: rcl_val.strong_induct)
+ case (SBitvec x)
+ have *:"b[bv::=b']\<^sub>b\<^sub>b = B_bitvec" using wfRCV_elims(9)[OF SBitvec(1)] by metis
+ show ?case proof (cases "b = B_var bv")
+ case True
+ moreover have "T \<turnstile> SBitvec x : B_bitvec" using wfRCV.intros by simp
+ moreover hence "b' = B_bitvec" using True SBitvec subst_bb.simps * by simp
+ ultimately show ?thesis using boxed_b.intros wfRCV.intros by metis
+ next
+ case False
+ hence "b = B_bitvec" using subst_bb_inject * by metis
+ then show ?thesis using * SBitvec boxed_b.intros by metis
+ qed
+next
+ case (SNum x)
+ have *:"b[bv::=b']\<^sub>b\<^sub>b = B_int" using wfRCV_elims(10)[OF SNum(1)] by metis
+ show ?case proof (cases "b = B_var bv")
+ case True
+ moreover have "T \<turnstile> SNum x : B_int" using wfRCV.intros by simp
+ moreover hence "b' = B_int" using True SNum subst_bb.simps(1) * by simp
+ ultimately show ?thesis using boxed_b_BVar1I wfRCV.intros by metis
+ next
+ case False
+ hence "b = B_int" using subst_bb_inject(1) * by metis
+ then show ?thesis using * SNum boxed_b_BIntI by metis
+ qed
+next
+ case (SBool x)
+ have *:"b[bv::=b']\<^sub>b\<^sub>b = B_bool" using wfRCV_elims(11)[OF SBool(1)] by metis
+ show ?case proof (cases "b = B_var bv")
+ case True
+ moreover have "T \<turnstile> SBool x : B_bool" using wfRCV.intros by simp
+ moreover hence "b' = B_bool" using True SBool subst_bb.simps * by simp
+ ultimately show ?thesis using boxed_b.intros wfRCV.intros by metis
+ next
+ case False
+ hence "b = B_bool" using subst_bb_inject * by metis
+ then show ?thesis using * SBool boxed_b.intros by metis
+ qed
+next
+ case (SPair s1 s2)
+ then obtain b1 and b2 where *:"b[bv::=b']\<^sub>b\<^sub>b = B_pair b1 b2 \<and> wfRCV T s1 b1 \<and> wfRCV T s2 b2" using wfRCV_elims(12) by metis
+ show ?case proof (cases "b = B_var bv")
+ case True
+ moreover have "T \<turnstile> SPair s1 s2 : B_pair b1 b2" using wfRCV.intros * by simp
+ moreover hence "b' = B_pair b1 b2" using True SPair subst_bb.simps(1) * by simp
+ ultimately show ?thesis using boxed_b_BVar1I by metis
+ next
+ case False
+ then obtain b1' and b2' where "b = B_pair b1' b2' \<and> b1=b1'[bv::=b']\<^sub>b\<^sub>b \<and> b2=b2'[bv::=b']\<^sub>b\<^sub>b" using subst_bb_inject(5)[OF _ False ] * by metis
+ then show ?thesis using * SPair boxed_b_BPairI by blast
+ qed
+next
+ case (SCons tyid dc s1)
+ have *:"b[bv::=b']\<^sub>b\<^sub>b = B_id tyid" using wfRCV_elims(13)[OF SCons(2)] by metis
+ show ?case proof (cases "b = B_var bv")
+ case True
+ moreover have "T \<turnstile> SCons tyid dc s1 : B_id tyid" using wfRCV.intros
+ using "local.*" SCons.prems by auto
+ moreover hence "b' = B_id tyid" using True SCons subst_bb.simps(1) * by simp
+ ultimately show ?thesis using boxed_b_BVar1I wfRCV.intros by metis
+ next
+ case False
+ then obtain b1' where beq: "b = B_id tyid" using subst_bb_inject * by metis
+ then obtain b2 dclist x c where **:"AF_typedef tyid dclist \<in> set T \<and> (dc, \<lbrace> x : b2 | c \<rbrace>) \<in> set dclist \<and> wfRCV T s1 b2" using wfRCV_elims(13) * SCons by metis
+ then have "atom bv \<sharp> b2" using \<open>wfTh T\<close> wfTh_lookup_supp_empty[of tyid dclist T dc "\<lbrace> x : b2 | c \<rbrace>"] \<tau>.fresh fresh_def by auto
+ then have "b2 = b2[ bv ::= b']\<^sub>b\<^sub>b" using forget_subst subst_b_b_def by metis
+ then obtain s1' where s1:"T \<turnstile> s1 ~ b2 [ bv ::= b' ] \<setminus> s1'" using SCons ** by metis
+
+ have "T \<turnstile> SCons tyid dc s1 ~ (B_id tyid) [ bv ::= b' ] \<setminus> SCons tyid dc s1'" proof(rule boxed_b_BConsI)
+ show "AF_typedef tyid dclist \<in> set T" using ** by auto
+ show "(dc, \<lbrace> x : b2 | c \<rbrace>) \<in> set dclist" using ** by auto
+ show "T \<turnstile> s1 ~ b2 [ bv ::= b' ] \<setminus> s1' " using s1 ** by auto
+
+ qed
+ thus ?thesis using beq by metis
+ qed
+next
+ case (SConsp typid dc bdc s)
+
+ obtain bva dclist x1 b1 c1 where **:"b[bv::=b']\<^sub>b\<^sub>b = B_app typid bdc \<and>
+ AF_typedef_poly typid bva dclist \<in> set T \<and> (dc, \<lbrace> x1 : b1 | c1 \<rbrace>) \<in> set dclist \<and> T \<turnstile> s : b1[bva::=bdc]\<^sub>b\<^sub>b \<and> atom bva \<sharp> bv"
+ using wfRCV_poly_elims[OF SConsp(2)] by metis
+
+ then have *:"B_app typid bdc = b[bv::=b']\<^sub>b\<^sub>b" using wfRCV_elims(14)[OF SConsp(2)] by metis
+ show ?case proof (cases "b = B_var bv")
+ case True
+ moreover have "T \<turnstile> SConsp typid dc bdc s : B_app typid bdc" using wfRCV.intros
+ using "local.*" SConsp.prems(1) by auto
+ moreover hence "b' = B_app typid bdc" using True SConsp subst_bb.simps * by simp
+ ultimately show ?thesis using boxed_b.intros wfRCV.intros by metis
+ next
+ case False
+ then obtain bdc' where bdc: "b = B_app typid bdc' \<and> bdc = bdc'[bv::=b']\<^sub>b\<^sub>b" using * subst_bb_inject(8)[OF *] by metis
+ (*hence beq:"b = B_app typid bdc" using subst_bb_inject * sory (* going to be like the above as bdc is closed *)*)
+ have "atom bv \<sharp> b1" proof -
+ have "supp b1 \<subseteq> { atom bva }" using wfTh_poly_supp_b ** SConsp by metis
+ moreover have "bv \<noteq> bva" using ** by auto
+ ultimately show ?thesis using fresh_def by force
+ qed
+ have "T \<turnstile> s : b1[bva::=bdc]\<^sub>b\<^sub>b" using ** by auto
+ moreover have "b1[bva::=bdc']\<^sub>b\<^sub>b[bv::=b']\<^sub>b\<^sub>b = b1[bva::=bdc]\<^sub>b\<^sub>b" using bdc subst_bb_commute \<open>atom bv \<sharp> b1\<close> by auto
+ ultimately obtain s' where s':"T \<turnstile> s ~ b1[bva::=bdc']\<^sub>b\<^sub>b [ bv ::= b' ] \<setminus> s'"
+ using SConsp(1)[of "b1[bva::=bdc']\<^sub>b\<^sub>b"] bdc SConsp by metis
+ have "T \<turnstile> SConsp typid dc bdc'[bv::=b']\<^sub>b\<^sub>b s ~ (B_app typid bdc') [ bv ::= b' ] \<setminus> SConsp typid dc bdc' s'"
+ proof -
+
+ obtain bva3 and dclist3 where 3:"AF_typedef_poly typid bva3 dclist3 = AF_typedef_poly typid bva dclist \<and>
+ atom bva3 \<sharp> (bdc', bv, b', s, s')" using obtain_fresh_bv by metis
+ then obtain x3 b3 c3 where 4:"(dc, \<lbrace> x3 : b3 | c3 \<rbrace>) \<in> set dclist3"
+ using boxed_b_BConspI td_lookup_eq_iff type_def.eq_iff
+ by (metis "**")
+
+ show ?thesis proof
+ show \<open>AF_typedef_poly typid bva3 dclist3 \<in> set T\<close> using 3 ** by metis
+ show \<open>atom bva3 \<sharp> (bdc', bv, b', s, s')\<close> using 3 by metis
+ show 4:\<open>(dc, \<lbrace> x3 : b3 | c3 \<rbrace>) \<in> set dclist3\<close> using 4 by auto
+ have "b3[bva3::=bdc']\<^sub>b\<^sub>b = b1[bva::=bdc']\<^sub>b\<^sub>b" proof(rule wfTh_typedef_poly_b_eq_iff)
+ show \<open>AF_typedef_poly typid bva3 dclist3 \<in> set T\<close> using 3 ** by metis
+ show \<open>(dc, \<lbrace> x3 : b3 | c3 \<rbrace>) \<in> set dclist3\<close> using 4 by auto
+ show \<open>AF_typedef_poly typid bva dclist \<in> set T\<close> using ** by auto
+ show \<open>(dc, \<lbrace> x1 : b1 | c1 \<rbrace>) \<in> set dclist\<close> using ** by auto
+ qed(simp add: ** SConsp)
+ thus \<open> T \<turnstile> s ~ b3[bva3::=bdc']\<^sub>b\<^sub>b [ bv ::= b' ] \<setminus> s' \<close> using s' by auto
+ qed
+ qed
+ then show ?thesis using bdc by auto
+
+ qed
+next
+ case SUnit
+ have *:"b[bv::=b']\<^sub>b\<^sub>b = B_unit" using wfRCV_elims SUnit by metis
+ show ?case proof (cases "b = B_var bv")
+ case True
+ moreover have "T \<turnstile> SUnit : B_unit" using wfRCV.intros by simp
+ moreover hence "b' = B_unit" using True SUnit subst_bb.simps * by simp
+ ultimately show ?thesis using boxed_b.intros wfRCV.intros by metis
+ next
+ case False
+ hence "b = B_unit" using subst_bb_inject * by metis
+ then show ?thesis using * SUnit boxed_b.intros by metis
+ qed
+next
+ case (SUt x)
+ then obtain bv' where *:"b[bv::=b']\<^sub>b\<^sub>b= B_var bv'" using wfRCV_elims by metis
+ show ?case proof (cases "b = B_var bv")
+ case True
+ then show ?thesis using boxed_b_BVar1I
+ using "local.*" wfRCV_BVarI by fastforce
+ next
+ case False
+ then show ?thesis using boxed_b_BVar1I boxed_b_BVar2I
+ using "local.*" wfRCV_BVarI by (metis subst_b_var)
+ qed
+qed
+
+lemma boxed_i_ex:
+ assumes "wfI T \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b i" and "wfTh T"
+ shows "\<exists>i'. T ; \<Gamma> ; b , bv \<turnstile> i \<approx> i'"
+ using assms proof(induct \<Gamma> arbitrary: i rule:\<Gamma>_induct)
+ case GNil
+ then show ?case using boxed_i_GNilI by metis
+next
+ case (GCons x' b' c' \<Gamma>')
+ then obtain s where 1:"Some s = i x' \<and> wfRCV T s b'[bv::=b]\<^sub>b\<^sub>b" using wfI_def subst_gb.simps by auto
+ then obtain s' where 2: "boxed_b T s b' bv b s'" using boxed_b_ex GCons by metis
+ then obtain i' where 3: "boxed_i T \<Gamma>' b bv i i'" using GCons wfI_def subst_gb.simps by force
+ have "boxed_i T ((x', b', c') #\<^sub>\<Gamma> \<Gamma>') b bv i (i'(x' \<mapsto> s'))" proof
+ show "Some s = i x'" using 1 by auto
+ show "boxed_b T s b' bv b s'" using 2 by auto
+ show "T ; \<Gamma>' ; b , bv \<turnstile> i \<approx> i'" using "3" by auto
+ qed
+ thus ?case by auto
+qed
+
+lemma boxed_b_eq:
+ assumes "boxed_b \<Theta> s1 b bv b' s1'" and "\<turnstile>\<^sub>w\<^sub>f \<Theta>"
+ shows "wfTh \<Theta> \<Longrightarrow> boxed_b \<Theta> s2 b bv b' s2' \<Longrightarrow> ( s1 = s2 ) = ( s1' = s2' )"
+ using assms proof(induct arbitrary: s2 s2' rule: boxed_b.inducts )
+ case (boxed_b_BVar1I bv bv' P s b )
+ then show ?case
+ using boxed_b_elims(1) rcl_val.eq_iff by metis
+next
+ case (boxed_b_BVar2I bv bv' P s b)
+ then show ?case using boxed_b_elims(1) by metis
+next
+ case (boxed_b_BIntI P s uu uv)
+ hence "s2 = s2'" using boxed_b_elims by metis
+ then show ?case by auto
+next
+ case (boxed_b_BBoolI P s uw ux)
+ hence "s2 = s2'" using boxed_b_elims by metis
+ then show ?case by auto
+next
+ case (boxed_b_BUnitI P s uy uz)
+ hence "s2 = s2'" using boxed_b_elims by metis
+ then show ?case by auto
+next
+ case (boxed_b_BPairI P s1 b1 bv b s1' s2a b2 s2a')
+ then show ?case
+ by (metis boxed_b_elims(5) rcl_val.eq_iff(4))
+next
+ case (boxed_b_BConsI tyid dclist P dc x b c s1 bv b' s1')
+ obtain s22 and s22' dclist2 dc2 x2 b2 c2 where *:"s2 = SCons tyid dc2 s22 \<and> s2' = SCons tyid dc2 s22' \<and> boxed_b P s22 b2 bv b' s22'
+ \<and> AF_typedef tyid dclist2 \<in> set P \<and> (dc2, \<lbrace> x2 : b2 | c2 \<rbrace>) \<in> set dclist2" using boxed_b_elims(6)[OF boxed_b_BConsI(6)] by metis
+ show ?case proof(cases "dc = dc2")
+ case True
+ hence "b = b2" using wfTh_ctor_unique \<tau>.eq_iff wfTh_dclist_unique wf boxed_b_BConsI * by metis
+ then show ?thesis using boxed_b_BConsI True * by auto
+ next
+ case False
+ then show ?thesis using * boxed_b_BConsI by simp
+ qed
+next
+ case (boxed_b_Bbitvec P s bv b)
+ hence "s2 = s2'" using boxed_b_elims by metis
+ then show ?case by auto
+next
+ case (boxed_b_BConspI tyid bva dclist P b1 bv b' s1 s1' dc x b c)
+ obtain bva2 s22 s22' dclist2 dc2 x2 b2 c2 where *:"
+ s2 = SConsp tyid dc2 b1[bv::=b']\<^sub>b\<^sub>b s22 \<and>
+ s2' = SConsp tyid dc2 b1 s22' \<and>
+ boxed_b P s22 b2[bva2::=b1]\<^sub>b\<^sub>b bv b' s22' \<and>
+ AF_typedef_poly tyid bva2 dclist2 \<in> set P \<and> (dc2, \<lbrace> x2 : b2 | c2 \<rbrace>) \<in> set dclist2" using boxed_b_elims(8)[OF boxed_b_BConspI(7)] by metis
+ show ?case proof(cases "dc = dc2")
+ case True
+ hence "AF_typedef_poly tyid bva2 dclist2 \<in> set P \<and> (dc, \<lbrace> x2 : b2 | c2 \<rbrace>) \<in> set dclist2" using * by auto
+ hence "b[bva::=b1]\<^sub>b\<^sub>b = b2[bva2::=b1]\<^sub>b\<^sub>b" using wfTh_typedef_poly_b_eq_iff[OF boxed_b_BConspI(1) boxed_b_BConspI(3)] * boxed_b_BConspI by metis
+ then show ?thesis using boxed_b_BConspI True * by auto
+ next
+ case False
+ then show ?thesis using * boxed_b_BConspI by simp
+ qed
+qed
+
+lemma bs_boxed_var:
+ assumes "boxed_i \<Theta> \<Gamma> b' bv i i'"
+ shows "Some (b,c) = lookup \<Gamma> x \<Longrightarrow> Some s = i x \<Longrightarrow> Some s' = i' x \<Longrightarrow> boxed_b \<Theta> s b bv b' s'"
+ using assms proof(induct rule: boxed_i.inducts)
+ case (boxed_i_GNilI T i)
+ then show ?case using lookup.simps by auto
+next
+ case (boxed_i_GConsI s i x1 \<Theta> b1 bv b' s' \<Gamma> i' c)
+ show ?case proof (cases "x=x1")
+ case True
+ then show ?thesis using boxed_i_GConsI
+ fun_upd_same lookup.simps(2) option.inject prod.inject by metis
+ next
+ case False
+ then show ?thesis using boxed_i_GConsI
+ fun_upd_same lookup.simps option.inject prod.inject by auto
+ qed
+qed
+
+lemma eval_l_boxed_b:
+ assumes "\<lbrakk> l \<rbrakk> = s"
+ shows "boxed_b \<Theta> s (base_for_lit l) bv b' s"
+ using assms proof(nominal_induct l arbitrary: s rule:l.strong_induct)
+qed(auto simp add: boxed_b.intros wfRCV.intros )+
+
+lemma boxed_i_eval_v_boxed_b:
+ fixes v::v
+ assumes "boxed_i \<Theta> \<Gamma> b' bv i i'" and "i \<lbrakk> v[bv::=b']\<^sub>v\<^sub>b \<rbrakk> ~ s" and "i' \<lbrakk> v \<rbrakk> ~ s'" and "wfV \<Theta> B \<Gamma> v b" and "wfI \<Theta> \<Gamma> i'"
+ shows "boxed_b \<Theta> s b bv b' s'"
+ using assms proof(nominal_induct v arbitrary: s s' b rule:v.strong_induct)
+ case (V_lit l)
+ hence "\<lbrakk> l \<rbrakk> = s \<and> \<lbrakk> l \<rbrakk> = s'" using eval_v_elims by auto
+ moreover have "b = base_for_lit l" using wfV_elims(2) V_lit by metis
+ ultimately show ?case using V_lit using eval_l_boxed_b subst_b_base_for_lit by metis
+next
+ case (V_var x) (* look at defn of bs_boxed *)
+ hence "Some s = i x \<and> Some s' = i' x" using eval_v_elims subst_vb.simps by metis
+ moreover obtain c1 where bc:"Some (b,c1) = lookup \<Gamma> x" using wfV_elims V_var by metis
+ ultimately show ?case using bs_boxed_var V_var by metis
+
+next
+ case (V_pair v1 v2)
+ then obtain b1 and b2 where b:"b=B_pair b1 b2" using wfV_elims subst_vb.simps by metis
+ obtain s1 and s2 where s: "eval_v i (v1[bv::=b']\<^sub>v\<^sub>b) s1 \<and> eval_v i (v2[bv::=b']\<^sub>v\<^sub>b) s2 \<and> s = SPair s1 s2" using eval_v_elims V_pair subst_vb.simps by metis
+ obtain s1' and s2' where s': "eval_v i' v1 s1' \<and> eval_v i' v2 s2' \<and> s' = SPair s1' s2'" using eval_v_elims V_pair by metis
+ have "boxed_b \<Theta> (SPair s1 s2) (B_pair b1 b2) bv b' (SPair s1' s2') " proof(rule boxed_b_BPairI)
+ show "boxed_b \<Theta> s1 b1 bv b' s1'" using V_pair eval_v_elims wfV_elims b s s' b.eq_iff by metis
+ show "boxed_b \<Theta> s2 b2 bv b' s2'" using V_pair eval_v_elims wfV_elims b s s' b.eq_iff by metis
+ qed
+ then show ?case using s s' b by auto
+next
+ case (V_cons tyid dc v1)
+
+ obtain dclist x b1 c where *: "b = B_id tyid \<and> AF_typedef tyid dclist \<in> set \<Theta> \<and> (dc, \<lbrace> x : b1 | c \<rbrace>) \<in> set dclist \<and> \<Theta> ; B ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v1 : b1"
+ using wfV_elims(4)[OF V_cons(5)] V_cons by metis
+ obtain s2 where s2: "s = SCons tyid dc s2 \<and> i \<lbrakk> (v1[bv::=b']\<^sub>v\<^sub>b) \<rbrakk> ~ s2" using eval_v_elims V_cons subst_vb.simps by metis
+ obtain s2' where s2': "s' = SCons tyid dc s2' \<and> i' \<lbrakk> v1 \<rbrakk> ~ s2'" using eval_v_elims V_cons by metis
+ have sp: "supp \<lbrace> x : b1 | c \<rbrace> = {}" using wfTh_lookup_supp_empty * wfX_wfY by metis
+
+ have "boxed_b \<Theta> (SCons tyid dc s2) (B_id tyid) bv b' (SCons tyid dc s2')"
+ proof(rule boxed_b_BConsI)
+ show 1:"AF_typedef tyid dclist \<in> set \<Theta>" using * by auto
+ show 2:"(dc, \<lbrace> x : b1 | c \<rbrace>) \<in> set dclist" using * by auto
+ have bvf:"atom bv \<sharp> b1" using sp \<tau>.fresh fresh_def by auto
+ show "\<Theta> \<turnstile> s2 ~ b1 [ bv ::= b' ] \<setminus> s2'" using V_cons s2 s2' * by metis
+ qed
+ then show ?case using * s2 s2' by simp
+next
+ case (V_consp tyid dc b1 v1)
+
+ obtain bv2 dclist x2 b2 c2 where *: "b = B_app tyid b1 \<and> AF_typedef_poly tyid bv2 dclist \<in> set \<Theta> \<and>
+ (dc, \<lbrace> x2 : b2 | c2 \<rbrace>) \<in> set dclist \<and> \<Theta> ; B ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v1 : b2[bv2::=b1]\<^sub>b\<^sub>b"
+ using wf_strong_elim(1)[OF V_consp (5)] by metis
+
+ obtain s2 where s2: "s = SConsp tyid dc b1[bv::=b']\<^sub>b\<^sub>b s2 \<and> i \<lbrakk> (v1[bv::=b']\<^sub>v\<^sub>b) \<rbrakk> ~ s2"
+ using eval_v_elims V_consp subst_vb.simps by metis
+
+ obtain s2' where s2': "s' = SConsp tyid dc b1 s2' \<and> i' \<lbrakk> v1 \<rbrakk> ~ s2'"
+ using eval_v_elims V_consp by metis
+
+ have "\<turnstile>\<^sub>w\<^sub>f \<Theta>" using V_consp wfX_wfY by metis
+ then obtain bv3::bv and dclist3 x3 b3 c3 where **:" AF_typedef_poly tyid bv2 dclist = AF_typedef_poly tyid bv3 dclist3 \<and>
+ (dc, \<lbrace> x3 : b3 | c3 \<rbrace>) \<in> set dclist3 \<and> atom bv3 \<sharp> (b1, bv, b', s2, s2') \<and> b2[bv2::=b1]\<^sub>b\<^sub>b = b3[bv3::=b1]\<^sub>b\<^sub>b"
+ using * obtain_fresh_bv_dclist_b_iff[where tm="(b1, bv, b', s2, s2')"] by metis
+
+ have "boxed_b \<Theta> (SConsp tyid dc b1[bv::=b']\<^sub>b\<^sub>b s2) (B_app tyid b1) bv b' (SConsp tyid dc b1 s2')"
+ proof(rule boxed_b_BConspI[of tyid bv3 dclist3 \<Theta>, where x=x3 and b=b3 and c=c3])
+ show 1:"AF_typedef_poly tyid bv3 dclist3 \<in> set \<Theta>" using * ** by auto
+ show 2:"(dc, \<lbrace> x3 : b3 | c3 \<rbrace>) \<in> set dclist3" using ** by auto
+ show "atom bv3 \<sharp> (b1, bv, b', s2, s2')" using ** by auto
+ show " \<Theta> \<turnstile> s2 ~ b3[bv3::=b1]\<^sub>b\<^sub>b [ bv ::= b' ] \<setminus> s2'" using V_consp s2 s2' * ** by metis
+ qed
+ then show ?case using * s2 s2' by simp
+qed
+
+lemma boxed_b_eq_eq:
+ assumes "boxed_b \<Theta> n1 b1 bv b' n1'" and "boxed_b \<Theta> n2 b1 bv b' n2'" and "s = SBool (n1 = n2)" and "\<turnstile>\<^sub>w\<^sub>f \<Theta>"
+ "s' = SBool (n1' = n2')"
+ shows "s=s'"
+ using boxed_b_eq assms by auto
+
+lemma boxed_i_eval_ce_boxed_b:
+ fixes e::ce
+ assumes "i' \<lbrakk> e \<rbrakk> ~ s'" and "i \<lbrakk> e[bv::=b']\<^sub>c\<^sub>e\<^sub>b \<rbrakk> ~ s" and "wfCE \<Theta> B \<Gamma> e b" and "boxed_i \<Theta> \<Gamma> b' bv i i'" and "wfI \<Theta> \<Gamma> i'"
+ shows "boxed_b \<Theta> s b bv b' s'"
+ using assms proof(nominal_induct e arbitrary: s s' b b' rule: ce.strong_induct)
+ case (CE_val x)
+ then show ?case using boxed_i_eval_v_boxed_b eval_e_elims wfCE_elims subst_ceb.simps by metis
+next
+ case (CE_op opp v1 v2)
+
+ show ?case proof(rule opp.exhaust)
+ assume \<open>opp = Plus\<close>
+ have 1:"wfCE \<Theta> B \<Gamma> v1 (B_int)" using wfCE_elims CE_op \<open>opp = Plus\<close> by metis
+ have 2:"wfCE \<Theta> B \<Gamma> v2 (B_int)" using wfCE_elims CE_op \<open>opp = Plus\<close> by metis
+ have *:"b = B_int" using CE_op wfCE_elims
+ by (metis \<open>opp = plus\<close>)
+
+ obtain n1 and n2 where n:"s = SNum (n1 + n2) \<and> i \<lbrakk> v1[bv::=b']\<^sub>c\<^sub>e\<^sub>b \<rbrakk> ~ SNum n1 \<and> i \<lbrakk> v2[bv::=b']\<^sub>c\<^sub>e\<^sub>b \<rbrakk> ~ SNum n2" using eval_e_elims CE_op subst_ceb.simps \<open>opp = plus\<close> by metis
+ obtain n1' and n2' where n':"s' = SNum (n1' + n2') \<and> i' \<lbrakk> v1 \<rbrakk> ~ SNum n1' \<and> i' \<lbrakk> v2 \<rbrakk> ~ SNum n2'" using eval_e_elims Plus CE_op \<open>opp = plus\<close> by metis
+
+ have "boxed_b \<Theta> (SNum n1) B_int bv b' (SNum n1')" using boxed_i_eval_v_boxed_b 1 2 n n' CE_op \<open>opp = plus\<close> by metis
+ moreover have "boxed_b \<Theta> (SNum n2) B_int bv b' (SNum n2')" using boxed_i_eval_v_boxed_b 1 2 n n' CE_op by metis
+ ultimately have "s=s'" using n' n boxed_b_elims(2)
+ by (metis rcl_val.eq_iff(2))
+ thus ?thesis using * n n' boxed_b_BIntI CE_op wfRCV.intros Plus by simp
+ next
+ assume \<open>opp = LEq\<close>
+ have 1:"wfCE \<Theta> B \<Gamma> v1 (B_int)" using wfCE_elims CE_op \<open>opp = LEq\<close> by metis
+ have 2:"wfCE \<Theta> B \<Gamma> v2 (B_int)" using wfCE_elims CE_op \<open>opp = LEq\<close> by metis
+ hence *:"b = B_bool" using CE_op wfCE_elims \<open>opp = LEq\<close> by metis
+ obtain n1 and n2 where n:"s = SBool (n1 \<le> n2) \<and> i \<lbrakk> v1[bv::=b']\<^sub>c\<^sub>e\<^sub>b \<rbrakk> ~ SNum n1 \<and> i \<lbrakk> v2[bv::=b']\<^sub>c\<^sub>e\<^sub>b \<rbrakk> ~ SNum n2" using eval_e_elims subst_ceb.simps CE_op \<open>opp = LEq\<close> by metis
+ obtain n1' and n2' where n':"s' = SBool (n1' \<le> n2') \<and> i' \<lbrakk> v1 \<rbrakk> ~ SNum n1' \<and> i' \<lbrakk> v2 \<rbrakk> ~ SNum n2'" using eval_e_elims CE_op \<open>opp = LEq\<close> by metis
+
+ have "boxed_b \<Theta> (SNum n1) B_int bv b' (SNum n1')" using boxed_i_eval_v_boxed_b 1 2 n n' CE_op by metis
+ moreover have "boxed_b \<Theta> (SNum n2) B_int bv b' (SNum n2')" using boxed_i_eval_v_boxed_b 1 2 n n' CE_op by metis
+ ultimately have "s=s'" using n' n boxed_b_elims(2)
+ by (metis rcl_val.eq_iff(2))
+ thus ?thesis using * n n' boxed_b_BBoolI CE_op wfRCV.intros \<open>opp = LEq\<close> by simp
+ next
+ assume \<open>opp = Eq\<close>
+ obtain b1 where b1:"wfCE \<Theta> B \<Gamma> v1 b1 \<and> wfCE \<Theta> B \<Gamma> v2 b1" using wfCE_elims CE_op \<open>opp = Eq\<close> by metis
+
+ hence *:"b = B_bool" using CE_op wfCE_elims \<open>opp = Eq\<close> by metis
+ obtain n1 and n2 where n:"s = SBool (n1 = n2) \<and> i \<lbrakk> v1[bv::=b']\<^sub>c\<^sub>e\<^sub>b \<rbrakk> ~ n1 \<and> i \<lbrakk> v2[bv::=b']\<^sub>c\<^sub>e\<^sub>b \<rbrakk> ~ n2" using eval_e_elims subst_ceb.simps CE_op \<open>opp = Eq\<close> by metis
+ obtain n1' and n2' where n':"s' = SBool (n1' = n2') \<and> i' \<lbrakk> v1 \<rbrakk> ~ n1' \<and> i' \<lbrakk> v2 \<rbrakk> ~ n2'" using eval_e_elims CE_op \<open>opp = Eq\<close> by metis
+
+ have "boxed_b \<Theta> n1 b1 bv b' n1'" using boxed_i_eval_v_boxed_b b1 n n' CE_op by metis
+ moreover have "boxed_b \<Theta> n2 b1 bv b' n2'" using boxed_i_eval_v_boxed_b b1 n n' CE_op by metis
+ moreover have "\<turnstile>\<^sub>w\<^sub>f \<Theta>" using b1 wfX_wfY by metis
+ ultimately have "s=s'" using n' n boxed_b_elims
+ boxed_b_eq_eq by metis
+ thus ?thesis using * n n' boxed_b_BBoolI CE_op wfRCV.intros \<open>opp = Eq\<close> by simp
+ qed
+
+next
+ case (CE_concat v1 v2)
+
+ obtain bv1 and bv2 where s : "s = SBitvec (bv1 @ bv2) \<and> (i \<lbrakk> v1[bv::=b']\<^sub>c\<^sub>e\<^sub>b \<rbrakk> ~ SBitvec bv1) \<and> i \<lbrakk> v2[bv::=b']\<^sub>c\<^sub>e\<^sub>b \<rbrakk> ~ SBitvec bv2"
+ using eval_e_elims(7) subst_ceb.simps CE_concat.prems(2) eval_e_elims(6) subst_ceb.simps(6) by metis
+ obtain bv1' and bv2' where s' : "s' = SBitvec (bv1' @ bv2') \<and> i' \<lbrakk> v1 \<rbrakk> ~ SBitvec bv1' \<and> i' \<lbrakk> v2 \<rbrakk> ~ SBitvec bv2'"
+ using eval_e_elims(7) CE_concat by metis
+
+ then show ?case using boxed_i_eval_v_boxed_b wfCE_elims s s' CE_concat
+ by (metis CE_concat.prems(3) assms assms(5) wfRCV_BBitvecI boxed_b_Bbitvec boxed_b_elims(7) eval_e_concatI eval_e_uniqueness)
+next
+ case (CE_fst ce)
+ obtain s2 where 1:"i \<lbrakk> ce[bv::=b']\<^sub>c\<^sub>e\<^sub>b \<rbrakk> ~ SPair s s2" using CE_fst eval_e_elims subst_ceb.simps by metis
+ obtain s2' where 2:"i' \<lbrakk> ce \<rbrakk> ~ SPair s' s2'" using CE_fst eval_e_elims by metis
+ obtain b2 where 3:"wfCE \<Theta> B \<Gamma> ce (B_pair b b2)" using wfCE_elims(4) CE_fst by metis
+
+ have "boxed_b \<Theta> (SPair s s2) (B_pair b b2) bv b' (SPair s' s2')"
+ using 1 2 3 CE_fst boxed_i_eval_v_boxed_b boxed_b_BPairI by auto
+ thus ?case using boxed_b_elims(5) by force
+next
+ case (CE_snd v)
+ obtain s1 where 1:"i \<lbrakk> v[bv::=b']\<^sub>c\<^sub>e\<^sub>b \<rbrakk> ~ SPair s1 s" using CE_snd eval_e_elims subst_ceb.simps by metis
+ obtain s1' where 2:"i' \<lbrakk> v \<rbrakk> ~ SPair s1' s'" using CE_snd eval_e_elims by metis
+ obtain b1 where 3:"wfCE \<Theta> B \<Gamma> v (B_pair b1 b )" using wfCE_elims(5) CE_snd by metis
+
+ have "boxed_b \<Theta> (SPair s1 s ) (B_pair b1 b ) bv b' (SPair s1' s')" using 1 2 3 CE_snd boxed_i_eval_v_boxed_b by simp
+ thus ?case using boxed_b_elims(5) by force
+next
+ case (CE_len v)
+ obtain s1 where s: "i \<lbrakk> v[bv::=b']\<^sub>c\<^sub>e\<^sub>b \<rbrakk> ~ SBitvec s1" using CE_len eval_e_elims subst_ceb.simps by metis
+ obtain s1' where s': "i' \<lbrakk> v \<rbrakk> ~ SBitvec s1'" using CE_len eval_e_elims by metis
+
+ have "\<Theta> ; B ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : B_bitvec \<and> b = B_int" using wfCE_elims CE_len by metis
+ then show ?case using boxed_i_eval_v_boxed_b s s' CE_len
+ by (metis boxed_b_BIntI boxed_b_elims(7) eval_e_lenI eval_e_uniqueness subst_ceb.simps(5) wfI_wfCE_eval_e)
+qed
+
+lemma eval_c_eq_bs_boxed:
+ fixes c::c
+ assumes "i \<lbrakk> c[bv::=b]\<^sub>c\<^sub>b \<rbrakk> ~ s" and "i' \<lbrakk> c \<rbrakk> ~ s'" and "wfC \<Theta> B \<Gamma> c" and "wfI \<Theta> \<Gamma> i'" and "\<Theta> ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> i "
+ and "boxed_i \<Theta> \<Gamma> b bv i i'"
+ shows "s = s'"
+ using assms proof(nominal_induct c arbitrary: s s' rule:c.strong_induct)
+ case C_true
+ then show ?case using eval_c_elims subst_cb.simps by metis
+next
+ case C_false
+ then show ?case using eval_c_elims subst_cb.simps by metis
+next
+ case (C_conj c1 c2)
+ obtain s1 and s2 where 1: "eval_c i (c1[bv::=b]\<^sub>c\<^sub>b) s1 \<and> eval_c i (c2[bv::=b]\<^sub>c\<^sub>b) s2 \<and> s = (s1\<and>s2)" using C_conj eval_c_elims(3) subst_cb.simps(3) by metis
+ obtain s1' and s2' where 2:"eval_c i' c1 s1' \<and> eval_c i' c2 s2' \<and> s' = (s1'\<and>s2')" using C_conj eval_c_elims(3) by metis
+ then show ?case using 1 2 wfC_elims C_conj by metis
+next
+ case (C_disj c1 c2)
+
+ obtain s1 and s2 where 1: "eval_c i (c1[bv::=b]\<^sub>c\<^sub>b) s1 \<and> eval_c i (c2[bv::=b]\<^sub>c\<^sub>b) s2 \<and> s = (s1\<or>s2)" using C_disj eval_c_elims(4) subst_cb.simps(4) by metis
+ obtain s1' and s2' where 2:"eval_c i' c1 s1' \<and> eval_c i' c2 s2' \<and> s' = (s1'\<or>s2')" using C_disj eval_c_elims(4) by metis
+ then show ?case using 1 2 wfC_elims C_disj by metis
+next
+ case (C_not c)
+ obtain s1::bool where 1: "(i \<lbrakk> c[bv::=b]\<^sub>c\<^sub>b \<rbrakk> ~ s1) \<and> (s = (\<not> s1))" using C_not eval_c_elims(6) subst_cb.simps(7) by metis
+ obtain s1'::bool where 2: "(i' \<lbrakk> c \<rbrakk> ~ s1') \<and> (s' = (\<not> s1'))" using C_not eval_c_elims(6) by metis
+ then show ?case using 1 2 wfC_elims C_not by metis
+next
+ case (C_imp c1 c2)
+ obtain s1 and s2 where 1: "eval_c i (c1[bv::=b]\<^sub>c\<^sub>b) s1 \<and> eval_c i (c2[bv::=b]\<^sub>c\<^sub>b) s2 \<and> s = (s1 \<longrightarrow> s2)" using C_imp eval_c_elims(5) subst_cb.simps(5) by metis
+ obtain s1' and s2' where 2:"eval_c i' c1 s1' \<and> eval_c i' c2 s2' \<and> s' = (s1' \<longrightarrow> s2')" using C_imp eval_c_elims(5) by metis
+ then show ?case using 1 2 wfC_elims C_imp by metis
+next
+ case (C_eq e1 e2)
+ obtain be where be: "wfCE \<Theta> B \<Gamma> e1 be \<and> wfCE \<Theta> B \<Gamma> e2 be" using C_eq wfC_elims by metis
+ obtain s1 and s2 where 1: "eval_e i (e1[bv::=b]\<^sub>c\<^sub>e\<^sub>b) s1 \<and> eval_e i (e2[bv::=b]\<^sub>c\<^sub>e\<^sub>b) s2 \<and> s = (s1 = s2)" using C_eq eval_c_elims(7) subst_cb.simps(6) by metis
+ obtain s1' and s2' where 2:"eval_e i' e1 s1' \<and> eval_e i' e2 s2' \<and> s' = (s1' = s2' )" using C_eq eval_c_elims(7) by metis
+ have "\<turnstile>\<^sub>w\<^sub>f \<Theta>" using C_eq wfX_wfY by metis
+ moreover have "\<Theta> ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> i " using C_eq by auto
+ ultimately show ?case using boxed_b_eq[of \<Theta> s1 be bv b s1' s2 s2'] 1 2 boxed_i_eval_ce_boxed_b C_eq wfC_elims subst_cb.simps 1 2 be by auto
+qed
+
+lemma is_satis_bs_boxed:
+ fixes c::c
+ assumes "boxed_i \<Theta> \<Gamma> b bv i i'" and "wfC \<Theta> B \<Gamma> c" and "wfI \<Theta> \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b i" and "\<Theta> ; \<Gamma> \<turnstile> i'"
+ and "(i \<Turnstile> c[bv::=b]\<^sub>c\<^sub>b)"
+ shows "(i' \<Turnstile> c)"
+proof -
+ have "eval_c i (c[bv::=b]\<^sub>c\<^sub>b) True" using is_satis.simps assms by auto
+ moreover obtain s where "i' \<lbrakk> c \<rbrakk> ~ s" using eval_c_exist assms by metis
+ ultimately show ?thesis using eval_c_eq_bs_boxed assms is_satis.simps by metis
+qed
+
+lemma is_satis_bs_boxed_rev:
+ fixes c::c
+ assumes "boxed_i \<Theta> \<Gamma> b bv i i'" and "wfC \<Theta> B \<Gamma> c" and "wfI \<Theta> \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b i" and "\<Theta> ; \<Gamma> \<turnstile> i'" and "wfC \<Theta> {||} \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b (c[bv::=b]\<^sub>c\<^sub>b)"
+ and "(i' \<Turnstile> c)"
+ shows "(i \<Turnstile> c[bv::=b]\<^sub>c\<^sub>b)"
+proof -
+ have "eval_c i' c True" using is_satis.simps assms by auto
+ moreover obtain s where "i \<lbrakk> c[bv::=b]\<^sub>c\<^sub>b \<rbrakk> ~ s" using eval_c_exist assms by metis
+ ultimately show ?thesis using eval_c_eq_bs_boxed assms is_satis.simps by metis
+qed
+
+lemma bs_boxed_wfi_aux:
+ fixes b::b and bv::bv and \<Theta>::\<Theta> and B::\<B>
+ assumes "boxed_i \<Theta> \<Gamma> b bv i i'" and "wfI \<Theta> \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b i" and "\<turnstile>\<^sub>w\<^sub>f \<Theta>" and "wfG \<Theta> B \<Gamma>"
+ shows "\<Theta> ; \<Gamma> \<turnstile> i'"
+ using assms proof(induct rule: boxed_i.inducts)
+ case (boxed_i_GNilI T i)
+ then show ?case using wfI_def by auto
+next
+ case (boxed_i_GConsI s i x1 T b1 bv b s' G i' c1)
+ {
+ fix x2 b2 c2
+ assume as : "(x2,b2,c2) \<in> toSet ((x1, b1, c1) #\<^sub>\<Gamma> G)"
+
+ then consider (hd) "(x2,b2,c2) = (x1, b1, c1)" | (tail) "(x2,b2,c2) \<in> toSet G" using toSet.simps by auto
+ hence "\<exists>s. Some s = (i'(x1 \<mapsto> s')) x2 \<and> wfRCV T s b2" proof(cases)
+ case hd
+ hence "b1=b2" by auto
+ moreover have "(x2,b2[bv::=b]\<^sub>b\<^sub>b,c2[bv::=b]\<^sub>c\<^sub>b) \<in> toSet ((x1, b1, c1) #\<^sub>\<Gamma> G)[bv::=b]\<^sub>\<Gamma>\<^sub>b" using hd subst_gb.simps by simp
+ moreover hence "wfRCV T s b2[bv::=b]\<^sub>b\<^sub>b" using wfI_def boxed_i_GConsI hd
+ proof -
+ obtain ss :: "b \<Rightarrow> x \<Rightarrow> (x \<Rightarrow> rcl_val option) \<Rightarrow> type_def list \<Rightarrow> rcl_val" where
+ "\<forall>x1a x2a x3 x4. (\<exists>v5. Some v5 = x3 x2a \<and> wfRCV x4 v5 x1a) = (Some (ss x1a x2a x3 x4) = x3 x2a \<and> wfRCV x4 (ss x1a x2a x3 x4) x1a)"
+ by moura (* 0.0 ms *)
+ then have f1: "Some (ss b2[bv::=b]\<^sub>b\<^sub>b x1 i T) = i x1 \<and> wfRCV T (ss b2[bv::=b]\<^sub>b\<^sub>b x1 i T) b2[bv::=b]\<^sub>b\<^sub>b"
+ using boxed_i_GConsI.prems(1) hd wfI_def by auto (* 31 ms *)
+ then have "ss b2[bv::=b]\<^sub>b\<^sub>b x1 i T = s"
+ by (metis (no_types) boxed_i_GConsI.hyps(1) option.inject) (* 0.0 ms *)
+ then show ?thesis
+ using f1 by blast (* 0.0 ms *)
+ qed
+ ultimately have "wfRCV T s' b2" using boxed_i_GConsI boxed_b_wfRCV by metis
+
+ then show ?thesis using hd by simp
+ next
+ case tail
+ hence "wfI T G i'" using boxed_i_GConsI wfI_suffix wfG_suffix subst_gb.simps
+ by (metis (no_types, lifting) Un_iff toSet.simps(2) wfG_cons2 wfI_def)
+ then show ?thesis using wfI_def[of T G i'] tail
+ using boxed_i_GConsI.prems(3) split_G wfG_cons_fresh2 by fastforce
+ qed
+ }
+ thus ?case using wfI_def by fast
+
+qed
+
+lemma is_satis_g_bs_boxed_aux:
+ fixes G::\<Gamma>
+ assumes "boxed_i \<Theta> G1 b bv i i'" and "wfI \<Theta> G1[bv::=b]\<^sub>\<Gamma>\<^sub>b i" and "wfI \<Theta> G1 i'" and "G1 = (G2@G)" and "wfG \<Theta> B G1"
+ and "(i \<Turnstile> G[bv::=b]\<^sub>\<Gamma>\<^sub>b) "
+ shows "(i' \<Turnstile> G)"
+ using assms proof(induct G arbitrary: G2 rule: \<Gamma>_induct)
+ case GNil
+ then show ?case by auto
+next
+ case (GCons x' b' c' \<Gamma>' G2)
+ show ?case proof(subst is_satis_g.simps,rule)
+ have *:"wfC \<Theta> B G1 c'" using GCons wfG_wfC_inside by force
+ show "i' \<Turnstile> c'" using is_satis_bs_boxed[OF assms(1) * ] GCons by auto
+ obtain G3 where "G1 = G3 @ \<Gamma>'" using GCons append_g.simps
+ by (metis append_g_assoc)
+ then show "i' \<Turnstile> \<Gamma>'" using GCons append_g.simps by simp
+ qed
+qed
+
+lemma is_satis_g_bs_boxed:
+ fixes G::\<Gamma>
+ assumes "boxed_i \<Theta> G b bv i i'" and "wfI \<Theta> G[bv::=b]\<^sub>\<Gamma>\<^sub>b i" and "wfI \<Theta> G i'" and "wfG \<Theta> B G"
+ and "(i \<Turnstile> G[bv::=b]\<^sub>\<Gamma>\<^sub>b) "
+ shows "(i' \<Turnstile> G)"
+ using is_satis_g_bs_boxed_aux assms
+ by (metis (full_types) append_g.simps(1))
+
+lemma subst_b_valid:
+ fixes s::s and b::b
+ assumes "\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f b" and "B = {|bv|}" and "\<Theta> ; {|bv|} ;\<Gamma> \<Turnstile> c"
+ shows "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<Turnstile> c[bv::=b]\<^sub>c\<^sub>b "
+proof(rule validI)
+
+ show **:"\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f c[bv::=b]\<^sub>c\<^sub>b " using assms valid.simps wf_b_subst subst_gb.simps by metis
+ show "\<forall>i. (wfI \<Theta> \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b i \<and> i \<Turnstile> \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b) \<longrightarrow> i \<Turnstile> c[bv::=b]\<^sub>c\<^sub>b "
+ proof(rule,rule)
+ fix i
+ assume *:"wfI \<Theta> \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b i \<and> i \<Turnstile> \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b"
+
+ obtain i' where idash: "boxed_i \<Theta> \<Gamma> b bv i i'" using boxed_i_ex wfX_wfY assms * by fastforce
+
+ have wfc: "\<Theta> ; {|bv|} ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f c" using valid.simps assms by simp
+ have wfg: "\<Theta> ; {|bv|} \<turnstile>\<^sub>w\<^sub>f \<Gamma>" using valid.simps wfX_wfY assms by metis
+ hence wfi: "wfI \<Theta> \<Gamma> i'" using idash * bs_boxed_wfi_aux subst_gb.simps wfX_wfY by metis
+ moreover have "i' \<Turnstile> \<Gamma>" proof (rule is_satis_g_bs_boxed[OF idash ] wfX_wfY(2)[OF wfc])
+ show "wfI \<Theta> \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b i" using subst_gb.simps * by simp
+ show "wfI \<Theta> \<Gamma> i'" using wfi by auto
+ show "\<Theta> ; B \<turnstile>\<^sub>w\<^sub>f \<Gamma> " using wfg assms by auto
+ show "i \<Turnstile> \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b" using subst_gb.simps * by simp
+ qed
+ ultimately have ic:"i' \<Turnstile> c" using assms valid_def using valid.simps by blast
+
+ show "i \<Turnstile> c[bv::=b]\<^sub>c\<^sub>b" proof(rule is_satis_bs_boxed_rev)
+ show "\<Theta> ; \<Gamma> ; b , bv \<turnstile> i \<approx> i'" using idash by auto
+ show "\<Theta> ; B ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f c " using wfc assms by auto
+ show "\<Theta> ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile> i" using subst_gb.simps * by simp
+ show "\<Theta> ; \<Gamma> \<turnstile> i'" using wfi by auto
+ show "\<Theta> ; {||} ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f c[bv::=b]\<^sub>c\<^sub>b " using ** by auto
+ show "i' \<Turnstile> c" using ic by auto
+ qed
+
+ qed
+qed
+
+section \<open>Expression Operator Lemmas\<close>
+
+lemma is_satis_len_imp:
+ assumes "i \<Turnstile> (CE_val (V_var x) == CE_val (V_lit (L_num (int (length v)))) )" (is "is_satis i ?c1")
+ shows "i \<Turnstile> (CE_val (V_var x) == CE_len [V_lit (L_bitvec v)]\<^sup>c\<^sup>e)"
+proof -
+ have *:"eval_c i ?c1 True" using assms is_satis.simps by blast
+ then have "eval_e i (CE_val (V_lit (L_num (int (length v))))) (SNum (int (length v)))"
+ using eval_e_elims(1) eval_v_elims eval_l.simps by (metis eval_e.intros(1) eval_v_litI)
+ hence "eval_e i (CE_val (V_var x)) (SNum (int (length v)))" using eval_c_elims(7)[OF *]
+ by (metis eval_e_elims(1) eval_v_elims(1))
+ moreover have "eval_e i (CE_len [V_lit (L_bitvec v)]\<^sup>c\<^sup>e) (SNum (int (length v)))"
+ using eval_e_elims(7) eval_v_elims eval_l.simps by (metis eval_e.intros eval_v_litI)
+ ultimately show ?thesis using eval_c.intros is_satis.simps by fastforce
+qed
+
+lemma is_satis_plus_imp:
+ assumes "i \<Turnstile> (CE_val (V_var x) == CE_val (V_lit (L_num (n1+n2))))" (is "is_satis i ?c1")
+ shows "i \<Turnstile> (CE_val (V_var x) == CE_op Plus ([V_lit (L_num n1)]\<^sup>c\<^sup>e) ([V_lit (L_num n2)]\<^sup>c\<^sup>e))"
+proof -
+ have *:"eval_c i ?c1 True" using assms is_satis.simps by blast
+ then have "eval_e i (CE_val (V_lit (L_num (n1+n2)))) (SNum (n1+n2))"
+ using eval_e_elims(1) eval_v_elims eval_l.simps by (metis eval_e.intros(1) eval_v_litI)
+ hence "eval_e i (CE_val (V_var x)) (SNum (n1+n2))" using eval_c_elims(7)[OF *]
+ by (metis eval_e_elims(1) eval_v_elims(1))
+ moreover have "eval_e i (CE_op Plus ([V_lit (L_num n1)]\<^sup>c\<^sup>e) ([V_lit (L_num n2)]\<^sup>c\<^sup>e)) (SNum (n1+n2))"
+ using eval_e_elims(7) eval_v_elims eval_l.simps by (metis eval_e.intros eval_v_litI)
+ ultimately show ?thesis using eval_c.intros is_satis.simps by fastforce
+qed
+
+lemma is_satis_leq_imp:
+ assumes "i \<Turnstile> (CE_val (V_var x) == CE_val (V_lit (if (n1 \<le> n2) then L_true else L_false)))" (is "is_satis i ?c1")
+ shows "i \<Turnstile> (CE_val (V_var x) == CE_op LEq [(V_lit (L_num n1))]\<^sup>c\<^sup>e [(V_lit (L_num n2))]\<^sup>c\<^sup>e)"
+proof -
+ have *:"eval_c i ?c1 True" using assms is_satis.simps by blast
+ then have "eval_e i (CE_val (V_lit ((if (n1 \<le> n2) then L_true else L_false)))) (SBool (n1\<le>n2))"
+ using eval_e_elims(1) eval_v_elims eval_l.simps
+ by (metis (full_types) eval_e.intros(1) eval_v_litI)
+ hence "eval_e i (CE_val (V_var x)) (SBool (n1\<le>n2))" using eval_c_elims(7)[OF *]
+ by (metis eval_e_elims(1) eval_v_elims(1))
+ moreover have "eval_e i (CE_op LEq [(V_lit (L_num n1))]\<^sup>c\<^sup>e [(V_lit (L_num n2) )]\<^sup>c\<^sup>e) (SBool (n1\<le>n2))"
+ using eval_e_elims(3) eval_v_elims eval_l.simps by (metis eval_e.intros eval_v_litI)
+ ultimately show ?thesis using eval_c.intros is_satis.simps by fastforce
+qed
+
+lemma eval_lit_inj:
+ fixes n1::l and n2::l
+ assumes "\<lbrakk> n1 \<rbrakk> = s" and "\<lbrakk> n2 \<rbrakk> = s"
+ shows "n1=n2"
+ using assms proof(nominal_induct s rule: rcl_val.strong_induct)
+ case (SBitvec x)
+ then show ?case using eval_l.simps
+ by (metis l.strong_exhaust rcl_val.distinct rcl_val.eq_iff)
+next
+ case (SNum x)
+ then show ?case using eval_l.simps
+ by (metis l.strong_exhaust rcl_val.distinct rcl_val.eq_iff)
+next
+ case (SBool x)
+ then show ?case using eval_l.simps
+ by (metis l.strong_exhaust rcl_val.distinct rcl_val.eq_iff)
+next
+ case (SPair x1a x2a)
+ then show ?case using eval_l.simps
+ by (metis l.strong_exhaust rcl_val.distinct rcl_val.eq_iff)
+next
+ case (SCons x1a x2a x3a)
+ then show ?case using eval_l.simps
+ by (metis l.strong_exhaust rcl_val.distinct rcl_val.eq_iff)
+next
+ case (SConsp x1a x2a x3a x4)
+ then show ?case using eval_l.simps
+ by (metis l.strong_exhaust rcl_val.distinct rcl_val.eq_iff)
+next
+ case SUnit
+ then show ?case using eval_l.simps
+ by (metis l.strong_exhaust rcl_val.distinct rcl_val.eq_iff)
+next
+ case (SUt x)
+ then show ?case using eval_l.simps
+ by (metis l.strong_exhaust rcl_val.distinct rcl_val.eq_iff)
+qed
+
+lemma eval_e_lit_inj:
+ fixes n1::l and n2::l
+ assumes "i \<lbrakk> [ [ n1 ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrakk> ~ s" and "i \<lbrakk> [ [ n2 ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrakk> ~ s"
+ shows "n1=n2"
+ using eval_lit_inj assms eval_e_elims eval_v_elims by metis
+
+lemma is_satis_eq_imp:
+ assumes "i \<Turnstile> (CE_val (V_var x) == CE_val (V_lit (if (n1 = n2) then L_true else L_false)))" (is "is_satis i ?c1")
+ shows "i \<Turnstile> (CE_val (V_var x) == CE_op Eq [(V_lit (n1))]\<^sup>c\<^sup>e [(V_lit (n2))]\<^sup>c\<^sup>e)"
+proof -
+ have *:"eval_c i ?c1 True" using assms is_satis.simps by blast
+ then have "eval_e i (CE_val (V_lit ((if (n1=n2) then L_true else L_false)))) (SBool (n1=n2))"
+ using eval_e_elims(1) eval_v_elims eval_l.simps
+ by (metis (full_types) eval_e.intros(1) eval_v_litI)
+ hence "eval_e i (CE_val (V_var x)) (SBool (n1=n2))" using eval_c_elims(7)[OF *]
+ by (metis eval_e_elims(1) eval_v_elims(1))
+ moreover have "eval_e i (CE_op Eq [(V_lit (n1))]\<^sup>c\<^sup>e [(V_lit (n2) )]\<^sup>c\<^sup>e) (SBool (n1=n2))"
+ proof -
+ obtain s1 and s2 where *:"i \<lbrakk> [ [ n1 ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrakk> ~ s1 \<and> i \<lbrakk> [ [ n2 ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrakk> ~ s2" using eval_l.simps eval_e.intros eval_v_litI by metis
+ moreover have " SBool (n1 = n2) = SBool (s1 = s2)" proof(cases "n1=n2")
+ case True
+ then show ?thesis using *
+ by (simp add: calculation eval_e_uniqueness)
+ next
+ case False
+ then show ?thesis using * eval_e_lit_inj by auto
+ qed
+ ultimately show ?thesis using eval_e_eqI[of i "[(V_lit (n1))]\<^sup>c\<^sup>e" s1 "[(V_lit (n2))]\<^sup>c\<^sup>e" s2 ] by auto
+ qed
+ ultimately show ?thesis using eval_c.intros is_satis.simps by fastforce
+qed
+
+lemma valid_eq_e:
+ assumes "\<forall>i s1 s2. wfG P \<B> GNil \<and> wfI P GNil i \<and> eval_e i e1 s1 \<and> eval_e i e2 s2 \<longrightarrow> s1 = s2"
+ and "wfCE P \<B> GNil e1 b" and "wfCE P \<B> GNil e2 b"
+ shows "P ; \<B> ; (x, b , CE_val (V_var x) == e1 )#\<^sub>\<Gamma> GNil \<Turnstile> CE_val (V_var x) == e2"
+ unfolding valid.simps
+proof(intro conjI)
+ show \<open> P ; \<B> ; (x, b, [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == e1 ) #\<^sub>\<Gamma> GNil \<turnstile>\<^sub>w\<^sub>f [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == e2 \<close>
+ using assms wf_intros wfX_wfY b.eq_iff fresh_GNil wfC_e_eq2 wfV_elims by meson
+ show \<open>\<forall>i. ((P ; (x, b, [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == e1 ) #\<^sub>\<Gamma> GNil \<turnstile> i) \<and> (i \<Turnstile> (x, b, [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == e1 ) #\<^sub>\<Gamma> GNil) \<longrightarrow>
+ (i \<Turnstile> [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == e2)) \<close> proof(rule+)
+ fix i
+ assume as:"P ; (x, b, [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == e1 ) #\<^sub>\<Gamma> GNil \<turnstile> i \<and> i \<Turnstile> (x, b, [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == e1 ) #\<^sub>\<Gamma> GNil"
+
+ have *: "P ; GNil \<turnstile> i " using wfI_def by auto
+
+ then obtain s1 where s1:"eval_e i e1 s1" using assms eval_e_exist by metis
+ obtain s2 where s2:"eval_e i e2 s2" using assms eval_e_exist * by metis
+ moreover have "i x = Some s1" proof -
+ have "i \<Turnstile> [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == e1" using as is_satis_g.simps by auto
+ thus ?thesis using s1
+ by (metis eval_c_elims(7) eval_e_elims(1) eval_e_uniqueness eval_v_elims(2) is_satis.cases)
+ qed
+ moreover have "s1 = s2" using s1 s2 * assms wfG_nilI wfX_wfY by metis
+
+ ultimately show "i \<lbrakk> [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == e2 \<rbrakk> ~ True"
+ using eval_c.intros eval_e.intros eval_v.intros
+ proof -
+ have "i \<lbrakk> e2 \<rbrakk> ~ s1"
+ by (metis \<open>s1 = s2\<close> s2) (* 0.0 ms *)
+ then show ?thesis
+ by (metis (full_types) \<open>i x = Some s1\<close> eval_c_eqI eval_e_valI eval_v_varI) (* 31 ms *)
+ qed
+ qed
+qed
+
+lemma valid_len:
+ assumes " \<turnstile>\<^sub>w\<^sub>f \<Theta>"
+ shows "\<Theta> ; \<B> ; (x, B_int, [[x]\<^sup>v]\<^sup>c\<^sup>e == [[ L_num (int (length v)) ]\<^sup>v]\<^sup>c\<^sup>e) #\<^sub>\<Gamma> GNil \<Turnstile> [[x]\<^sup>v]\<^sup>c\<^sup>e == CE_len [[ L_bitvec v ]\<^sup>v]\<^sup>c\<^sup>e" (is "\<Theta> ; \<B> ; ?G \<Turnstile> ?c" )
+proof -
+ have *:"\<Theta> \<turnstile>\<^sub>w\<^sub>f ([]::\<Phi>) \<and> \<Theta> ; \<B> ; GNil \<turnstile>\<^sub>w\<^sub>f []\<^sub>\<Delta> " using assms wfG_nilI wfD_emptyI wfPhi_emptyI by auto
+
+ moreover hence "\<Theta> ; \<B> ; GNil \<turnstile>\<^sub>w\<^sub>f CE_val (V_lit (L_num (int (length v)))) : B_int"
+ using wfCE_valI * wfV_litI base_for_lit.simps
+ by (metis wfE_valI wfX_wfY)
+
+ moreover have "\<Theta> ; \<B> ; GNil \<turnstile>\<^sub>w\<^sub>f CE_len [(V_lit (L_bitvec v))]\<^sup>c\<^sup>e : B_int"
+ using wfE_valI * wfV_litI base_for_lit.simps wfE_valI wfX_wfY wfCE_valI
+ by (metis wfCE_lenI)
+ moreover have "atom x \<sharp> GNil" by auto
+ ultimately have "\<Theta> ; \<B> ; ?G \<turnstile>\<^sub>w\<^sub>f ?c" using wfC_e_eq2 assms by simp
+ moreover have "(\<forall>i. wfI \<Theta> ?G i \<and> is_satis_g i ?G \<longrightarrow> is_satis i ?c)" using is_satis_len_imp by auto
+ ultimately show ?thesis using valid.simps by auto
+qed
+
+lemma valid_arith_bop:
+ assumes "wfG \<Theta> \<B> \<Gamma>" and "opp = Plus \<and> ll = (L_num (n1+n2)) \<or> (opp = LEq \<and> ll = ( if n1\<le>n2 then L_true else L_false))"
+ and "(opp = Plus \<longrightarrow> b = B_int) \<and> (opp = LEq \<longrightarrow> b = B_bool)" and
+ "atom x \<sharp> \<Gamma>"
+ shows "\<Theta>; \<B> ; (x, b, (CE_val (V_var x) == CE_val (V_lit (ll)) )) #\<^sub>\<Gamma> \<Gamma>
+ \<Turnstile> (CE_val (V_var x) == CE_op opp ([V_lit (L_num n1)]\<^sup>c\<^sup>e) ([V_lit (L_num n2)]\<^sup>c\<^sup>e ))" (is "\<Theta> ; \<B> ; ?G \<Turnstile> ?c")
+proof -
+ have "wfC \<Theta> \<B> ?G ?c" proof(rule wfC_e_eq2)
+ show "\<Theta> ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_val (V_lit ll) : b" using wfCE_valI wfV_litI assms base_for_lit.simps by metis
+ show "\<Theta> ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_op opp ([V_lit (L_num n1)]\<^sup>c\<^sup>e) ([V_lit (L_num n2)]\<^sup>c\<^sup>e) : b "
+ using wfCE_plusI wfCE_leqI wfCE_eqI wfV_litI wfCE_valI base_for_lit.simps assms by metis
+ show "\<turnstile>\<^sub>w\<^sub>f \<Theta>" using assms wfX_wfY by auto
+ show "atom x \<sharp> \<Gamma>" using assms by auto
+ qed
+
+ moreover have "\<forall>i. wfI \<Theta> ?G i \<and> is_satis_g i ?G \<longrightarrow> is_satis i ?c" proof(rule allI , rule impI)
+ fix i
+ assume "wfI \<Theta> ?G i \<and> is_satis_g i ?G"
+
+ hence "is_satis i ((CE_val (V_var x) == CE_val (V_lit (ll)) ))" by auto
+ thus "is_satis i ((CE_val (V_var x) == CE_op opp ([V_lit (L_num n1)]\<^sup>c\<^sup>e) ([V_lit (L_num n2)]\<^sup>c\<^sup>e)))"
+ using is_satis_plus_imp assms opp.exhaust is_satis_leq_imp by auto
+ qed
+ ultimately show ?thesis using valid.simps by metis
+qed
+
+lemma valid_eq_bop:
+ assumes "wfG \<Theta> \<B> \<Gamma>" and "atom x \<sharp> \<Gamma>" and "base_for_lit l1 = base_for_lit l2"
+ shows "\<Theta>; \<B> ; (x, B_bool, (CE_val (V_var x) == CE_val (V_lit (if l1 = l2 then L_true else L_false)) )) #\<^sub>\<Gamma> \<Gamma>
+ \<Turnstile> (CE_val (V_var x) == CE_op Eq ([V_lit (l1)]\<^sup>c\<^sup>e) ([V_lit (l2)]\<^sup>c\<^sup>e ))" (is "\<Theta> ; \<B> ; ?G \<Turnstile> ?c")
+proof -
+ let ?ll = "(if l1 = l2 then L_true else L_false)"
+ have "wfC \<Theta> \<B> ?G ?c" proof(rule wfC_e_eq2)
+ show "\<Theta> ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_val (V_lit ?ll) : B_bool" using wfCE_valI wfV_litI assms base_for_lit.simps by metis
+ show "\<Theta> ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_op Eq ([V_lit (l1)]\<^sup>c\<^sup>e) ([V_lit (l2)]\<^sup>c\<^sup>e) : B_bool "
+ using wfCE_eqI wfCE_leqI wfCE_eqI wfV_litI wfCE_valI base_for_lit.simps assms by metis
+ show "\<turnstile>\<^sub>w\<^sub>f \<Theta>" using assms wfX_wfY by auto
+ show "atom x \<sharp> \<Gamma>" using assms by auto
+ qed
+
+ moreover have "\<forall>i. wfI \<Theta> ?G i \<and> is_satis_g i ?G \<longrightarrow> is_satis i ?c" proof(rule allI , rule impI)
+ fix i
+ assume "wfI \<Theta> ?G i \<and> is_satis_g i ?G"
+
+ hence "is_satis i ((CE_val (V_var x) == CE_val (V_lit (?ll)) ))" by auto
+ thus "is_satis i ((CE_val (V_var x) == CE_op Eq ([V_lit (l1)]\<^sup>c\<^sup>e) ([V_lit (l2)]\<^sup>c\<^sup>e)))"
+ using is_satis_eq_imp assms by auto
+ qed
+ ultimately show ?thesis using valid.simps by metis
+qed
+
+lemma valid_fst:
+ fixes x::x and v\<^sub>1::v and v\<^sub>2::v
+ assumes "wfTh \<Theta>" and "wfV \<Theta> \<B> GNil (V_pair v\<^sub>1 v\<^sub>2) (B_pair b\<^sub>1 b\<^sub>2)"
+ shows "\<Theta> ; \<B> ; (x, b\<^sub>1, [[x]\<^sup>v]\<^sup>c\<^sup>e == [v\<^sub>1]\<^sup>c\<^sup>e) #\<^sub>\<Gamma> GNil \<Turnstile> [[x]\<^sup>v]\<^sup>c\<^sup>e == [#1[[v\<^sub>1,v\<^sub>2]\<^sup>v]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e"
+proof(rule valid_eq_e)
+ show \<open>\<forall>i s1 s2. (\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f GNil) \<and> (\<Theta> ; GNil \<turnstile> i) \<and> (i \<lbrakk> [ v\<^sub>1 ]\<^sup>c\<^sup>e \<rbrakk> ~ s1) \<and> (i \<lbrakk> [#1[[ v\<^sub>1 , v\<^sub>2 ]\<^sup>v]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e \<rbrakk> ~ s2) \<longrightarrow> s1 = s2\<close>
+ proof(rule+)
+ fix i s1 s2
+ assume as:"\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f GNil \<and> \<Theta> ; GNil \<turnstile> i \<and> (i \<lbrakk> [ v\<^sub>1 ]\<^sup>c\<^sup>e \<rbrakk> ~ s1) \<and> (i \<lbrakk> [#1[[ v\<^sub>1 , v\<^sub>2 ]\<^sup>v]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e \<rbrakk> ~ s2)"
+ then obtain s2' where *:"i \<lbrakk> [ v\<^sub>1 , v\<^sub>2 ]\<^sup>v \<rbrakk> ~ SPair s2 s2'"
+ using eval_e_elims(5)[of i "[[ v\<^sub>1 , v\<^sub>2 ]\<^sup>v]\<^sup>c\<^sup>e" s2] eval_e_elims
+ by meson
+ then have " i \<lbrakk> v\<^sub>1 \<rbrakk> ~ s2" using eval_v_elims(3)[OF *] by auto
+ then show "s1 = s2" using eval_v_uniqueness as
+ using eval_e_uniqueness eval_e_valI by blast
+ qed
+
+ show \<open> \<Theta> ; \<B> ; GNil \<turnstile>\<^sub>w\<^sub>f [ v\<^sub>1 ]\<^sup>c\<^sup>e : b\<^sub>1 \<close> using assms
+ by (metis b.eq_iff(4) wfV_elims(3) wfV_wfCE)
+ show \<open> \<Theta> ; \<B> ; GNil \<turnstile>\<^sub>w\<^sub>f [#1[[ v\<^sub>1 , v\<^sub>2 ]\<^sup>v]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e : b\<^sub>1 \<close> using assms using wfCE_fstI
+ using wfCE_valI by blast
+qed
+
+lemma valid_snd:
+ fixes x::x and v\<^sub>1::v and v\<^sub>2::v
+ assumes "wfTh \<Theta>" and "wfV \<Theta> \<B> GNil (V_pair v\<^sub>1 v\<^sub>2) (B_pair b\<^sub>1 b\<^sub>2)"
+ shows "\<Theta> ; \<B> ; (x, b\<^sub>2, [[x]\<^sup>v]\<^sup>c\<^sup>e == [v\<^sub>2]\<^sup>c\<^sup>e) #\<^sub>\<Gamma> GNil \<Turnstile> [[x]\<^sup>v]\<^sup>c\<^sup>e == [#2[[v\<^sub>1,v\<^sub>2]\<^sup>v]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e"
+proof(rule valid_eq_e)
+ show \<open>\<forall>i s1 s2. (\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f GNil) \<and> (\<Theta> ; GNil \<turnstile> i) \<and> (i \<lbrakk> [ v\<^sub>2 ]\<^sup>c\<^sup>e \<rbrakk> ~ s1) \<and>
+(i \<lbrakk> [#2[[ v\<^sub>1 , v\<^sub>2 ]\<^sup>v]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e \<rbrakk> ~ s2) \<longrightarrow> s1 = s2\<close>
+ proof(rule+)
+ fix i s1 s2
+ assume as:"\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f GNil \<and> \<Theta> ; GNil \<turnstile> i \<and> (i \<lbrakk> [ v\<^sub>2 ]\<^sup>c\<^sup>e \<rbrakk> ~ s1) \<and> (i \<lbrakk> [#2[[ v\<^sub>1 , v\<^sub>2 ]\<^sup>v]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e \<rbrakk> ~ s2)"
+ then obtain s2' where *:"i \<lbrakk> [ v\<^sub>1 , v\<^sub>2 ]\<^sup>v \<rbrakk> ~ SPair s2' s2"
+ using eval_e_elims(5)[of i "[[ v\<^sub>1 , v\<^sub>2 ]\<^sup>v]\<^sup>c\<^sup>e" s2] eval_e_elims
+ by meson
+ then have " i \<lbrakk> v\<^sub>2 \<rbrakk> ~ s2" using eval_v_elims(3)[OF *] by auto
+ then show "s1 = s2" using eval_v_uniqueness as
+ using eval_e_uniqueness eval_e_valI by blast
+ qed
+
+ show \<open> \<Theta> ; \<B> ; GNil \<turnstile>\<^sub>w\<^sub>f [ v\<^sub>2 ]\<^sup>c\<^sup>e : b\<^sub>2 \<close> using assms
+ by (metis b.eq_iff wfV_elims wfV_wfCE)
+ show \<open> \<Theta> ; \<B> ; GNil \<turnstile>\<^sub>w\<^sub>f [#2[[ v\<^sub>1 , v\<^sub>2 ]\<^sup>v]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e : b\<^sub>2 \<close> using assms using wfCE_sndI wfCE_valI by blast
+qed
+
+lemma valid_concat:
+ fixes v1::"bit list" and v2::"bit list"
+ assumes " \<turnstile>\<^sub>w\<^sub>f \<Pi>"
+ shows "\<Pi> ; \<B> ; (x, B_bitvec, (CE_val (V_var x) == CE_val (V_lit (L_bitvec (v1@ v2))))) #\<^sub>\<Gamma> GNil \<Turnstile>
+ (CE_val (V_var x) == CE_concat ([V_lit (L_bitvec v1)]\<^sup>c\<^sup>e ) ([V_lit (L_bitvec v2)]\<^sup>c\<^sup>e) )"
+proof(rule valid_eq_e)
+ show \<open>\<forall>i s1 s2. ((\<Pi> ; \<B> \<turnstile>\<^sub>w\<^sub>f GNil) \<and> (\<Pi> ; GNil \<turnstile> i) \<and>
+ (i \<lbrakk> [ [ L_bitvec (v1 @ v2) ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrakk> ~ s1) \<and> (i \<lbrakk> [[[ L_bitvec v1 ]\<^sup>v]\<^sup>c\<^sup>e @@ [[ L_bitvec v2 ]\<^sup>v]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e \<rbrakk> ~ s2) \<longrightarrow>
+ s1 = s2)\<close>
+ proof(rule+)
+ fix i s1 s2
+ assume as: "(\<Pi> ; \<B> \<turnstile>\<^sub>w\<^sub>f GNil) \<and> (\<Pi> ; GNil \<turnstile> i) \<and> (i \<lbrakk> [ [ L_bitvec (v1 @ v2) ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrakk> ~ s1) \<and>
+ (i \<lbrakk> [[[ L_bitvec v1 ]\<^sup>v]\<^sup>c\<^sup>e @@ [[ L_bitvec v2 ]\<^sup>v]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e \<rbrakk> ~ s2) "
+
+ hence *: "i \<lbrakk> [[[ L_bitvec v1 ]\<^sup>v]\<^sup>c\<^sup>e @@ [[ L_bitvec v2 ]\<^sup>v]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e \<rbrakk> ~ s2" by auto
+ obtain bv1 bv2 where s2:"s2 = SBitvec (bv1 @ bv2) \<and> i \<lbrakk> [ L_bitvec v1 ]\<^sup>v \<rbrakk> ~ SBitvec bv1 \<and> (i \<lbrakk> [ L_bitvec v2 ]\<^sup>v \<rbrakk> ~ SBitvec bv2)"
+ using eval_e_elims(7)[OF *] eval_e_elims(1) by metis
+ hence "v1 = bv1 \<and> v2 = bv2" using eval_v_elims(1) eval_l.simps(5) by force
+ moreover then have "s1 = SBitvec (bv1 @ bv2)" using s2 using eval_v_elims(1) eval_l.simps(5)
+ by (metis as eval_e_elims(1))
+
+ then show "s1 = s2" using s2 by auto
+ qed
+
+ show \<open> \<Pi> ; \<B> ; GNil \<turnstile>\<^sub>w\<^sub>f [ [ L_bitvec (v1 @ v2) ]\<^sup>v ]\<^sup>c\<^sup>e : B_bitvec \<close>
+ by (metis assms base_for_lit.simps(5) wfG_nilI wfV_litI wfV_wfCE)
+ show \<open> \<Pi> ; \<B> ; GNil \<turnstile>\<^sub>w\<^sub>f [[[ L_bitvec v1 ]\<^sup>v]\<^sup>c\<^sup>e @@ [[ L_bitvec v2 ]\<^sup>v]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e : B_bitvec \<close>
+ by (metis assms base_for_lit.simps(5) wfCE_concatI wfG_nilI wfV_litI wfCE_valI)
+qed
+
+lemma valid_ce_eq:
+ fixes ce::ce
+ assumes "\<Theta> ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ce : b"
+ shows \<open>\<Theta> ; \<B> ; \<Gamma> \<Turnstile> ce == ce \<close>
+ unfolding valid.simps proof
+ show \<open> \<Theta> ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ce == ce \<close> using assms wfC_eqI by auto
+ show \<open>\<forall>i. \<Theta> ; \<Gamma> \<turnstile> i \<and> i \<Turnstile> \<Gamma> \<longrightarrow> i \<Turnstile> ce == ce \<close> proof(rule+)
+ fix i
+ assume "\<Theta> ; \<Gamma> \<turnstile> i \<and> i \<Turnstile> \<Gamma>"
+ then obtain s where "i\<lbrakk> ce \<rbrakk> ~ s" using assms eval_e_exist by metis
+ then show "i \<lbrakk> ce == ce \<rbrakk> ~ True " using eval_c_eqI by metis
+ qed
+qed
+
+lemma valid_eq_imp:
+ fixes c1::c and c2::c
+ assumes " \<Theta> ; \<B> ; (x, b, c2) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f c1 IMP c2"
+ shows " \<Theta> ; \<B> ; (x, b, c2) #\<^sub>\<Gamma> \<Gamma> \<Turnstile> c1 IMP c2 "
+proof -
+ have "\<forall>i. (\<Theta> ; (x, b, c2) #\<^sub>\<Gamma> \<Gamma> \<turnstile> i \<and> i \<Turnstile> (x, b, c2) #\<^sub>\<Gamma> \<Gamma>) \<longrightarrow> i \<Turnstile> ( c1 IMP c2 )"
+ proof(rule,rule)
+ fix i
+ assume as:"\<Theta> ; (x, b, c2) #\<^sub>\<Gamma> \<Gamma> \<turnstile> i \<and> i \<Turnstile> (x, b, c2) #\<^sub>\<Gamma> \<Gamma>"
+
+ have "\<Theta> ; \<B> ; (x, b, c2) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f c1" using wfC_elims assms by metis
+
+ then obtain sc where "i \<lbrakk> c1 \<rbrakk> ~ sc" using eval_c_exist assms as by metis
+ moreover have "i \<lbrakk> c2 \<rbrakk> ~ True" using as is_satis_g.simps is_satis.simps by auto
+
+ ultimately have "i \<lbrakk> c1 IMP c2 \<rbrakk> ~ True" using eval_c_impI by metis
+
+ thus "i \<Turnstile> c1 IMP c2" using is_satis.simps by auto
+ qed
+ thus ?thesis using assms by auto
+qed
+
+lemma valid_range:
+ assumes "0 \<le> n \<and> n \<le> m" and "\<turnstile>\<^sub>w\<^sub>f \<Theta>"
+ shows "\<Theta> ; {||} ; (x, B_int , (C_eq (CE_val (V_var x)) (CE_val (V_lit (L_num n))))) #\<^sub>\<Gamma> GNil \<Turnstile>
+ (C_eq (CE_op LEq (CE_val (V_var x)) (CE_val (V_lit (L_num m)))) [[ L_true ]\<^sup>v ]\<^sup>c\<^sup>e) AND
+ (C_eq (CE_op LEq (CE_val (V_lit (L_num 0))) (CE_val (V_var x))) [[ L_true ]\<^sup>v ]\<^sup>c\<^sup>e)"
+ (is "\<Theta> ; {||} ; ?G \<Turnstile> ?c1 AND ?c2")
+proof(rule validI)
+ have wfg: " \<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f (x, B_int, [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_num n ]\<^sup>v ]\<^sup>c\<^sup>e ) #\<^sub>\<Gamma> GNil "
+ using assms base_for_lit.simps wfG_nilI wfV_litI fresh_GNil wfB_intI wfC_v_eq wfG_cons1I wfG_cons2I by metis
+
+ show "\<Theta> ; {||} ; ?G \<turnstile>\<^sub>w\<^sub>f ?c1 AND ?c2"
+ using wfC_conjI wfC_eqI wfCE_leqI wfCE_valI wfV_varI wfg lookup.simps base_for_lit.simps wfV_litI wfB_intI wfB_boolI
+ by metis
+
+ show "\<forall>i. \<Theta> ; ?G \<turnstile> i \<and> i \<Turnstile> ?G \<longrightarrow> i \<Turnstile> ?c1 AND ?c2" proof(rule,rule)
+ fix i
+ assume a:"\<Theta> ; ?G \<turnstile> i \<and> i \<Turnstile> ?G"
+ hence *:"i \<lbrakk> V_var x \<rbrakk> ~ SNum n"
+ proof -
+ obtain sv where sv: "i x = Some sv \<and> \<Theta> \<turnstile> sv : B_int" using a wfI_def by force
+ have "i \<lbrakk> (C_eq (CE_val (V_var x)) (CE_val (V_lit (L_num n)))) \<rbrakk> ~ True"
+ using a is_satis_g.simps
+ using is_satis.cases by blast
+ hence "i x = Some(SNum n)" using sv
+ by (metis eval_c_elims(7) eval_e_elims(1) eval_l.simps(3) eval_v_elims(1) eval_v_elims(2))
+ thus ?thesis using eval_v_varI by auto
+ qed
+
+ show "i \<Turnstile> ?c1 AND ?c2"
+ proof -
+ have "i \<lbrakk> ?c1 \<rbrakk> ~ True"
+ proof -
+ have "i \<lbrakk> [ leq [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e [ [ L_num m ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e\<rbrakk> ~ SBool True"
+ using eval_e_leqI assms eval_v_litI eval_l.simps *
+ by (metis (full_types) eval_e_valI)
+ moreover have "i \<lbrakk> [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrakk> ~ SBool True"
+ using eval_v_litI eval_e_valI eval_l.simps by metis
+ ultimately show ?thesis using eval_c_eqI by metis
+ qed
+
+ moreover have "i \<lbrakk> ?c2 \<rbrakk> ~ True"
+ proof -
+ have "i \<lbrakk> [ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e \<rbrakk> ~ SBool True"
+ using eval_e_leqI assms eval_v_litI eval_l.simps *
+ by (metis (full_types) eval_e_valI)
+ moreover have "i \<lbrakk> [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrakk> ~ SBool True"
+ using eval_v_litI eval_e_valI eval_l.simps by metis
+ ultimately show ?thesis using eval_c_eqI by metis
+ qed
+ ultimately show ?thesis using eval_c_conjI is_satis.simps by metis
+ qed
+ qed
+qed
+
+lemma valid_range_length:
+ fixes \<Gamma>::\<Gamma>
+ assumes "0 \<le> n \<and> n \<le> int (length v)" and " \<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f \<Gamma>" and "atom x \<sharp> \<Gamma>"
+ shows "\<Theta> ; {||} ; (x, B_int , (C_eq (CE_val (V_var x)) (CE_val (V_lit (L_num n))))) #\<^sub>\<Gamma> \<Gamma> \<Turnstile>
+ (C_eq (CE_op LEq (CE_val (V_lit (L_num 0))) (CE_val (V_var x))) [[ L_true ]\<^sup>v ]\<^sup>c\<^sup>e) AND
+ (C_eq (CE_op LEq (CE_val (V_var x)) ([| [ [ L_bitvec v ]\<^sup>v ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e )) [[ L_true ]\<^sup>v ]\<^sup>c\<^sup>e)
+ "
+ (is "\<Theta> ; {||} ; ?G \<Turnstile> ?c1 AND ?c2")
+proof(rule validI)
+ have wfg: " \<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f (x, B_int, [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_num n ]\<^sup>v ]\<^sup>c\<^sup>e ) #\<^sub>\<Gamma> \<Gamma> " apply(rule wfG_cons1I)
+ apply simp
+ using assms apply simp+
+ using assms base_for_lit.simps wfG_nilI wfV_litI wfB_intI wfC_v_eq wfB_intI wfX_wfY assms by metis+
+
+ show "\<Theta> ; {||} ; ?G \<turnstile>\<^sub>w\<^sub>f ?c1 AND ?c2"
+ using wfC_conjI wfC_eqI wfCE_leqI wfCE_valI wfV_varI wfg lookup.simps base_for_lit.simps wfV_litI wfB_intI wfB_boolI
+ by (metis (full_types) wfCE_lenI)
+
+ show "\<forall>i. \<Theta> ; ?G \<turnstile> i \<and> i \<Turnstile> ?G \<longrightarrow> i \<Turnstile> ?c1 AND ?c2" proof(rule,rule)
+ fix i
+ assume a:"\<Theta> ; ?G \<turnstile> i \<and> i \<Turnstile> ?G"
+ hence *:"i \<lbrakk> V_var x \<rbrakk> ~ SNum n"
+ proof -
+ obtain sv where sv: "i x = Some sv \<and> \<Theta> \<turnstile> sv : B_int" using a wfI_def by force
+ have "i \<lbrakk> (C_eq (CE_val (V_var x)) (CE_val (V_lit (L_num n)))) \<rbrakk> ~ True"
+ using a is_satis_g.simps
+ using is_satis.cases by blast
+ hence "i x = Some(SNum n)" using sv
+ by (metis eval_c_elims(7) eval_e_elims(1) eval_l.simps(3) eval_v_elims(1) eval_v_elims(2))
+ thus ?thesis using eval_v_varI by auto
+ qed
+
+ show "i \<Turnstile> ?c1 AND ?c2"
+ proof -
+ have "i \<lbrakk> ?c2 \<rbrakk> ~ True"
+ proof -
+ have "i \<lbrakk> [ leq [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e [| [ [ L_bitvec v ]\<^sup>v ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e\<rbrakk> ~ SBool True"
+ using eval_e_leqI assms eval_v_litI eval_l.simps *
+ by (metis (full_types) eval_e_lenI eval_e_valI)
+ moreover have "i \<lbrakk> [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrakk> ~ SBool True"
+ using eval_v_litI eval_e_valI eval_l.simps by metis
+ ultimately show ?thesis using eval_c_eqI by metis
+ qed
+
+ moreover have "i \<lbrakk> ?c1 \<rbrakk> ~ True"
+ proof -
+ have "i \<lbrakk> [ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e \<rbrakk> ~ SBool True"
+ using eval_e_leqI assms eval_v_litI eval_l.simps *
+ by (metis (full_types) eval_e_valI)
+ moreover have "i \<lbrakk> [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrakk> ~ SBool True"
+ using eval_v_litI eval_e_valI eval_l.simps by metis
+ ultimately show ?thesis using eval_c_eqI by metis
+ qed
+ ultimately show ?thesis using eval_c_conjI is_satis.simps by metis
+ qed
+ qed
+qed
+
+lemma valid_range_length_inv_gnil:
+ fixes \<Gamma>::\<Gamma>
+ assumes "\<turnstile>\<^sub>w\<^sub>f \<Theta> "
+ and "\<Theta> ; {||} ; (x, B_int , (C_eq (CE_val (V_var x)) (CE_val (V_lit (L_num n))))) #\<^sub>\<Gamma> GNil \<Turnstile>
+ (C_eq (CE_op LEq (CE_val (V_lit (L_num 0))) (CE_val (V_var x))) [[ L_true ]\<^sup>v ]\<^sup>c\<^sup>e) AND
+ (C_eq (CE_op LEq (CE_val (V_var x)) ([| [ [ L_bitvec v ]\<^sup>v ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e )) [[ L_true ]\<^sup>v ]\<^sup>c\<^sup>e)
+ "
+ (is "\<Theta> ; {||} ; ?G \<Turnstile> ?c1 AND ?c2")
+ shows "0 \<le> n \<and> n \<le> int (length v)"
+proof -
+ have *:"\<forall>i. \<Theta> ; ?G \<turnstile> i \<and> i \<Turnstile> ?G \<longrightarrow> i \<Turnstile> ?c1 AND ?c2" using assms valid.simps by simp
+
+ obtain i where i: "i x = Some (SNum n)" by auto
+ have "\<Theta> ; ?G \<turnstile> i \<and> i \<Turnstile> ?G" proof
+ show "\<Theta> ; ?G \<turnstile> i" unfolding wfI_def using wfRCV_BIntI i * by auto
+ have "i \<lbrakk> ([ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_num n ]\<^sup>v ]\<^sup>c\<^sup>e ) \<rbrakk> ~ True"
+ using * eval_c.intros(7) eval_e.intros eval_v.intros eval_l.simps
+ by (metis (full_types) i)
+ thus "i \<Turnstile> ?G" unfolding is_satis_g.simps is_satis.simps by auto
+ qed
+ hence **:"i \<Turnstile> ?c1 AND ?c2" using * by auto
+
+ hence 1: "i \<lbrakk> ?c1 \<rbrakk> ~ True" using eval_c_elims(3) is_satis.simps
+ by fastforce
+ then obtain sv1 and sv2 where "(sv1 = sv2) = True \<and> i \<lbrakk> [ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e \<rbrakk> ~ sv1 \<and> i \<lbrakk> [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrakk> ~ sv2"
+ using eval_c_elims(7) by metis
+ hence "sv1 = SBool True" using eval_e_elims eval_v_elims eval_l.simps i by metis
+ obtain n1 and n2 where "SBool True = SBool (n1 \<le> n2) \<and> (i \<lbrakk> [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrakk> ~ SNum n1) \<and> (i \<lbrakk> [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrakk> ~ SNum n2)"
+ using eval_e_elims(3)[of i " [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e" "[ [ x ]\<^sup>v ]\<^sup>c\<^sup>e " "SBool True"]
+ using \<open>(sv1 = sv2) = True \<and> i \<lbrakk> [ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e \<rbrakk> ~ sv1 \<and> i \<lbrakk> [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrakk> ~ sv2\<close> \<open>sv1 = SBool True\<close> by fastforce
+ moreover hence "n1 = 0" and "n2 = n" using eval_e_elims eval_v_elims i
+ apply (metis eval_l.simps(3) rcl_val.eq_iff(2))
+ using eval_e_elims eval_v_elims i
+ by (metis calculation option.inject rcl_val.eq_iff(2))
+ ultimately have le1: "0 \<le> n " by simp
+
+ hence 2: "i \<lbrakk> ?c2 \<rbrakk> ~ True" using ** eval_c_elims(3) is_satis.simps
+ by fastforce
+ then obtain sv1 and sv2 where sv: "(sv1 = sv2) = True \<and> i \<lbrakk> [ leq [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e [| [ [ L_bitvec v ]\<^sup>v ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e \<rbrakk> ~ sv1 \<and> i \<lbrakk> [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrakk> ~ sv2"
+ using eval_c_elims(7) by metis
+ hence "sv1 = SBool True" using eval_e_elims eval_v_elims eval_l.simps i by metis
+ obtain n1 and n2 where ***:"SBool True = SBool (n1 \<le> n2) \<and> (i \<lbrakk> [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrakk> ~ SNum n1) \<and> (i \<lbrakk> [| [ [ L_bitvec v ]\<^sup>v ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e \<rbrakk> ~ SNum n2)"
+ using eval_e_elims(3)
+ using sv \<open>sv1 = SBool True\<close> by metis
+ moreover hence "n1 = n" using eval_e_elims(1)[of i] eval_v_elims(2)[of i x "SNum n1"] i by auto
+ moreover have "n2 = int (length v)" using eval_e_elims(8) eval_v_elims(1) eval_l.simps i
+ by (metis "***" eval_e_elims(1) rcl_val.eq_iff(1) rcl_val.eq_iff(2))
+ ultimately have le2: "n \<le> int (length v) " by simp
+
+ show ?thesis using le1 le2 by auto
+qed
+
+lemma wfI_cons:
+ fixes i::valuation and \<Gamma>::\<Gamma>
+ assumes "i' \<Turnstile> \<Gamma>" and "\<Theta> ; \<Gamma> \<turnstile> i'" and "i = i' ( x \<mapsto> s)" and "\<Theta> \<turnstile> s : b" and "atom x \<sharp> \<Gamma>"
+ shows "\<Theta> ; (x,b,c) #\<^sub>\<Gamma> \<Gamma> \<turnstile> i"
+ unfolding wfI_def proof -
+ {
+ fix x' b' c'
+ assume "(x',b',c') \<in> toSet ((x, b, c) #\<^sub>\<Gamma> \<Gamma>)"
+ then consider "(x',b',c') = (x,b,c)" | "(x',b',c') \<in> toSet \<Gamma>" using toSet.simps by auto
+ then have "\<exists>s. Some s = i x' \<and> \<Theta> \<turnstile> s : b'" proof(cases)
+ case 1
+ then show ?thesis using assms by auto
+ next
+ case 2
+ then obtain s where s:"Some s = i' x' \<and> \<Theta> \<turnstile> s : b'" using assms wfI_def by auto
+ moreover have "x' \<noteq> x" using assms 2 fresh_dom_free by auto
+ ultimately have "Some s = i x'" using assms by auto
+ then show ?thesis using s wfI_def by auto
+ qed
+ }
+ thus "\<forall>(x, b, c)\<in>toSet ((x, b, c) #\<^sub>\<Gamma> \<Gamma>). \<exists>s. Some s = i x \<and> \<Theta> \<turnstile> s : b" by auto
+qed
+
+lemma valid_range_length_inv:
+ fixes \<Gamma>::\<Gamma>
+ assumes "\<Theta> ; B \<turnstile>\<^sub>w\<^sub>f \<Gamma> " and "atom x \<sharp> \<Gamma>" and "\<exists>i. i \<Turnstile> \<Gamma> \<and> \<Theta> ; \<Gamma> \<turnstile> i"
+ and "\<Theta> ; B ; (x, B_int , (C_eq (CE_val (V_var x)) (CE_val (V_lit (L_num n))))) #\<^sub>\<Gamma> \<Gamma> \<Turnstile>
+ (C_eq (CE_op LEq (CE_val (V_lit (L_num 0))) (CE_val (V_var x))) [[ L_true ]\<^sup>v ]\<^sup>c\<^sup>e) AND
+ (C_eq (CE_op LEq (CE_val (V_var x)) ([| [ [ L_bitvec v ]\<^sup>v ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e )) [[ L_true ]\<^sup>v ]\<^sup>c\<^sup>e)
+ "
+ (is "\<Theta> ; ?B ; ?G \<Turnstile> ?c1 AND ?c2")
+ shows "0 \<le> n \<and> n \<le> int (length v)"
+proof -
+ have *:"\<forall>i. \<Theta> ; ?G \<turnstile> i \<and> i \<Turnstile> ?G \<longrightarrow> i \<Turnstile> ?c1 AND ?c2" using assms valid.simps by simp
+
+ obtain i' where idash: "is_satis_g i' \<Gamma> \<and> \<Theta> ; \<Gamma> \<turnstile> i'" using assms by auto
+ obtain i where i: "i = i' ( x \<mapsto> SNum n)" by auto
+ hence ix: "i x = Some (SNum n)" by auto
+ have "\<Theta> ; ?G \<turnstile> i \<and> i \<Turnstile> ?G" proof
+ show "\<Theta> ; ?G \<turnstile> i" using wfI_cons i idash ix wfRCV_BIntI assms by simp
+
+ have **:"i \<lbrakk> ([ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_num n ]\<^sup>v ]\<^sup>c\<^sup>e ) \<rbrakk> ~ True"
+ using * eval_c.intros(7) eval_e.intros eval_v.intros eval_l.simps i
+ by (metis (full_types) ix)
+
+ show "i \<Turnstile> ?G" unfolding is_satis_g.simps proof
+ show \<open> i \<Turnstile> [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_num n ]\<^sup>v ]\<^sup>c\<^sup>e \<close> using ** is_satis.simps by auto
+ show \<open> i \<Turnstile> \<Gamma> \<close> using idash i assms is_satis_g_i_upd by metis
+ qed
+ qed
+ hence **:"i \<Turnstile> ?c1 AND ?c2" using * by auto
+
+ hence 1: "i \<lbrakk> ?c1 \<rbrakk> ~ True" using eval_c_elims(3) is_satis.simps
+ by fastforce
+ then obtain sv1 and sv2 where "(sv1 = sv2) = True \<and> i \<lbrakk> [ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e \<rbrakk> ~ sv1 \<and> i \<lbrakk> [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrakk> ~ sv2"
+ using eval_c_elims(7) by metis
+ hence "sv1 = SBool True" using eval_e_elims eval_v_elims eval_l.simps i by metis
+ obtain n1 and n2 where "SBool True = SBool (n1 \<le> n2) \<and> (i \<lbrakk> [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrakk> ~ SNum n1) \<and> (i \<lbrakk> [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrakk> ~ SNum n2)"
+ using eval_e_elims(3)[of i " [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e" "[ [ x ]\<^sup>v ]\<^sup>c\<^sup>e " "SBool True"]
+ using \<open>(sv1 = sv2) = True \<and> i \<lbrakk> [ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e \<rbrakk> ~ sv1 \<and> i \<lbrakk> [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrakk> ~ sv2\<close> \<open>sv1 = SBool True\<close> by fastforce
+ moreover hence "n1 = 0" and "n2 = n" using eval_e_elims eval_v_elims i
+ apply (metis eval_l.simps(3) rcl_val.eq_iff(2))
+ using eval_e_elims eval_v_elims i
+ calculation option.inject rcl_val.eq_iff(2)
+ by (metis ix)
+ ultimately have le1: "0 \<le> n " by simp
+
+ hence 2: "i \<lbrakk> ?c2 \<rbrakk> ~ True" using ** eval_c_elims(3) is_satis.simps
+ by fastforce
+ then obtain sv1 and sv2 where sv: "(sv1 = sv2) = True \<and> i \<lbrakk> [ leq [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e [| [ [ L_bitvec v ]\<^sup>v ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e \<rbrakk> ~ sv1 \<and> i \<lbrakk> [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrakk> ~ sv2"
+ using eval_c_elims(7) by metis
+ hence "sv1 = SBool True" using eval_e_elims eval_v_elims eval_l.simps i by metis
+ obtain n1 and n2 where ***:"SBool True = SBool (n1 \<le> n2) \<and> (i \<lbrakk> [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrakk> ~ SNum n1) \<and> (i \<lbrakk> [| [ [ L_bitvec v ]\<^sup>v ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e \<rbrakk> ~ SNum n2)"
+ using eval_e_elims(3)
+ using sv \<open>sv1 = SBool True\<close> by metis
+ moreover hence "n1 = n" using eval_e_elims(1)[of i] eval_v_elims(2)[of i x "SNum n1"] i by auto
+ moreover have "n2 = int (length v)" using eval_e_elims(8) eval_v_elims(1) eval_l.simps i
+ by (metis "***" eval_e_elims(1) rcl_val.eq_iff(1) rcl_val.eq_iff(2))
+ ultimately have le2: "n \<le> int (length v) " by simp
+
+ show ?thesis using le1 le2 by auto
+qed
+
+lemma eval_c_conj2I[intro]:
+ assumes "i \<lbrakk> c1 \<rbrakk> ~ True" and "i \<lbrakk> c2 \<rbrakk> ~ True"
+ shows "i \<lbrakk> (C_conj c1 c2) \<rbrakk> ~ True"
+ using assms eval_c_conjI by metis
+
+lemma valid_split:
+ assumes "split n v (v1,v2)" and "\<turnstile>\<^sub>w\<^sub>f \<Theta>"
+ shows "\<Theta> ; {||} ; (z , [B_bitvec , B_bitvec ]\<^sup>b , [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ [ L_bitvec v1 ]\<^sup>v , [ L_bitvec v2 ]\<^sup>v ]\<^sup>v ]\<^sup>c\<^sup>e) #\<^sub>\<Gamma> GNil
+\<Turnstile> ([ [ L_bitvec v ]\<^sup>v ]\<^sup>c\<^sup>e == [ [#1[ [ z ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e @@ [#2[ [ z ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e) AND ([| [#1[ [ z ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e == [ [ L_num n ]\<^sup>v ]\<^sup>c\<^sup>e)"
+ (is "\<Theta> ; {||} ; ?G \<Turnstile> ?c1 AND ?c2")
+ unfolding valid.simps proof
+
+ have wfg: " \<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f (z, [B_bitvec , B_bitvec ]\<^sup>b , [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ [ L_bitvec v1 ]\<^sup>v , [ L_bitvec v2 ]\<^sup>v ]\<^sup>v ]\<^sup>c\<^sup>e) #\<^sub>\<Gamma> GNil"
+ using wf_intros assms base_for_lit.simps fresh_GNil wfC_v_eq wfG_intros2 by metis
+
+ show "\<Theta> ; {||} ; ?G \<turnstile>\<^sub>w\<^sub>f ?c1 AND ?c2"
+ apply(rule wfC_conjI)
+ apply(rule wfC_eqI)
+ apply(rule wfCE_valI)
+ apply(rule wfV_litI)
+ using wf_intros wfg lookup.simps base_for_lit.simps wfC_v_eq
+ apply (metis )+
+ done
+
+ have len:"int (length v1) = n" using assms split_length by auto
+
+ show "\<forall>i. \<Theta> ; ?G \<turnstile> i \<and> i \<Turnstile> ?G \<longrightarrow> i \<Turnstile> (?c1 AND ?c2)"
+ proof(rule,rule)
+ fix i
+ assume a:"\<Theta> ; ?G \<turnstile> i \<and> i \<Turnstile> ?G"
+ hence "i \<lbrakk> [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ [ L_bitvec v1 ]\<^sup>v , [ L_bitvec v2 ]\<^sup>v ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrakk> ~ True"
+ using is_satis_g.simps is_satis.simps by simp
+ then obtain sv where "i \<lbrakk> [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrakk> ~ sv \<and> i \<lbrakk> [ [ [ L_bitvec v1 ]\<^sup>v , [ L_bitvec v2 ]\<^sup>v ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrakk> ~ sv"
+ using eval_c_elims by metis
+ hence "i \<lbrakk> [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrakk> ~ (SPair (SBitvec v1) (SBitvec v2))" using eval_c_eqI eval_v.intros eval_l.simps
+ by (metis eval_e_elims(1) eval_v_uniqueness)
+ hence b:"i z = Some (SPair (SBitvec v1) (SBitvec v2))" using a eval_e_elims eval_v_elims by metis
+
+ have v1: "i \<lbrakk> [#1[ [ z ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e \<rbrakk> ~ SBitvec v1 "
+ using eval_e_fstI eval_e_valI eval_v_varI b by metis
+ have v2: "i \<lbrakk> [#2[ [ z ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e \<rbrakk> ~ SBitvec v2"
+ using eval_e_sndI eval_e_valI eval_v_varI b by metis
+
+ have "i \<lbrakk> [ [ L_bitvec v ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrakk> ~ SBitvec v" using eval_e.intros eval_v.intros eval_l.simps by metis
+ moreover have "i \<lbrakk> [ [#1[ [ z ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e @@ [#2[ [ z ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e \<rbrakk> ~ SBitvec v"
+ using assms split_concat v1 v2 eval_e_concatI by metis
+ moreover have "i \<lbrakk> [| [#1[ [ z ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e \<rbrakk> ~ SNum (int (length v1))"
+ using v1 eval_e_lenI by auto
+ moreover have "i \<lbrakk> [ [ L_num n ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrakk> ~ SNum n" using eval_e.intros eval_v.intros eval_l.simps by metis
+ ultimately show "i \<Turnstile> ?c1 AND ?c2" using is_satis.intros eval_c_conj2I eval_c_eqI len by metis
+ qed
+qed
+
+
+lemma is_satis_eq:
+ assumes "wfI \<Theta> G i" and "wfCE \<Theta> \<B> G e b"
+ shows "is_satis i (e == e)"
+proof(rule)
+ obtain s where "eval_e i e s" using eval_e_exist assms by metis
+ thus "eval_c i (e == e ) True" using eval_c_eqI by metis
+qed
+
+lemma is_satis_g_i_upd2:
+ assumes "eval_v i v s" and "is_satis ((i ( x \<mapsto> s))) c0" and "atom x \<sharp> G" and "wfG \<Theta> \<B> (G3@((x,b,c0)#\<^sub>\<Gamma>G))" and "wfV \<Theta> \<B> G v b" and "wfI \<Theta> (G3[x::=v]\<^sub>\<Gamma>\<^sub>v@G) i"
+ and "is_satis_g i (G3[x::=v]\<^sub>\<Gamma>\<^sub>v@G)"
+ shows "is_satis_g (i ( x \<mapsto> s)) (G3@((x,b,c0)#\<^sub>\<Gamma>G))"
+ using assms proof(induct G3 rule: \<Gamma>_induct)
+ case GNil
+ hence "is_satis_g (i(x \<mapsto> s)) G" using is_satis_g_i_upd by auto
+ then show ?case using GNil using is_satis_g.simps append_g.simps by metis
+next
+ case (GCons x' b' c' \<Gamma>')
+ hence "x\<noteq>x'" using wfG_cons_append by metis
+ hence "is_satis_g i (((x', b', c'[x::=v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> (\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v) @ G))" using subst_gv.simps GCons by auto
+ hence *:"is_satis i c'[x::=v]\<^sub>c\<^sub>v \<and> is_satis_g i ((\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v) @ G)" using subst_gv.simps by auto
+
+ have "is_satis_g (i(x \<mapsto> s)) ((x', b', c') #\<^sub>\<Gamma> (\<Gamma>'@ (x, b, c0) #\<^sub>\<Gamma> G))" proof(subst is_satis_g.simps,rule)
+ show "is_satis (i(x \<mapsto> s)) c'" proof(subst subst_c_satis_full[symmetric])
+ show \<open>eval_v i v s\<close> using GCons by auto
+ show \<open> \<Theta> ; \<B> ; ((x', b', c') #\<^sub>\<Gamma>\<Gamma>')@(x, b, c0) #\<^sub>\<Gamma> G \<turnstile>\<^sub>w\<^sub>f c' \<close> using GCons wfC_refl by auto
+ show \<open>wfI \<Theta> ((((x', b', c') #\<^sub>\<Gamma> \<Gamma>')[x::=v]\<^sub>\<Gamma>\<^sub>v) @ G) i\<close> using GCons by auto
+ show \<open> \<Theta> ; \<B> ; G \<turnstile>\<^sub>w\<^sub>f v : b \<close> using GCons by auto
+ show \<open>is_satis i c'[x::=v]\<^sub>c\<^sub>v\<close> using * by auto
+ qed
+ show "is_satis_g (i(x \<mapsto> s)) (\<Gamma>' @ (x, b, c0) #\<^sub>\<Gamma> G)" proof(rule GCons(1))
+ show \<open>eval_v i v s\<close> using GCons by auto
+ show \<open>is_satis (i(x \<mapsto> s)) c0\<close> using GCons by metis
+ show \<open>atom x \<sharp> G\<close> using GCons by auto
+ show \<open> \<Theta> ; \<B>\<turnstile>\<^sub>w\<^sub>f \<Gamma>' @ (x, b, c0) #\<^sub>\<Gamma> G \<close> using GCons wfG_elims append_g.simps by metis
+ show \<open>is_satis_g i (\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ G)\<close> using * by auto
+ show "wfI \<Theta> (\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ G) i" using GCons wfI_def subst_g_assoc_cons \<open>x\<noteq>x'\<close> by auto
+ show "\<Theta> ; \<B> ; G \<turnstile>\<^sub>w\<^sub>f v : b " using GCons by auto
+ qed
+ qed
+ moreover have "((x', b', c') #\<^sub>\<Gamma> \<Gamma>' @ (x, b, c0) #\<^sub>\<Gamma> G) = (((x', b', c') #\<^sub>\<Gamma> \<Gamma>') @ (x, b, c0) #\<^sub>\<Gamma> G)" by auto
+ ultimately show ?case using GCons by metis
+qed
+
+
end
\ No newline at end of file
diff --git a/thys/MiniSail/ROOT b/thys/MiniSail/ROOT
--- a/thys/MiniSail/ROOT
+++ b/thys/MiniSail/ROOT
@@ -1,15 +1,15 @@
-chapter AFP
-
-session "MiniSail" (AFP) = "HOL-Library" +
- description \<open>Formalisation of MiniSail\<close>
- options [timeout = 4800]
- sessions
- "Nominal2"
- "HOL-Eisbach"
- "Native_Word"
- "Show"
- theories
- "MiniSail"
- document_files
- "root.tex"
- "root.bib"
+chapter AFP
+
+session "MiniSail" (AFP) = "HOL-Library" +
+ description \<open>Formalisation of MiniSail\<close>
+ options [timeout = 4800]
+ sessions
+ "Nominal2"
+ "HOL-Eisbach"
+ "Native_Word"
+ "Show"
+ theories
+ "MiniSail"
+ document_files
+ "root.tex"
+ "root.bib"
diff --git a/thys/MiniSail/Safety.thy b/thys/MiniSail/Safety.thy
--- a/thys/MiniSail/Safety.thy
+++ b/thys/MiniSail/Safety.thy
@@ -1,1812 +1,1812 @@
-(*<*)
-theory Safety
- imports Operational IVSubstTypingL BTVSubstTypingL
-begin
- (*>*)
-
-method supp_calc = (metis (mono_tags, opaque_lifting) pure_supp c.supp e.supp v.supp supp_l_empty opp.supp sup_bot.right_neutral supp_at_base)
-declare infer_e.intros[simp]
-declare infer_e.intros[intro]
-
-chapter \<open>Safety\<close>
-
-text \<open>Lemmas about the operational semantics leading up to progress and preservation and then
-safety.\<close>
-
-section \<open>Store Lemmas\<close>
-
-abbreviation delta_ext (" _ \<sqsubseteq> _ ") where
- "delta_ext \<Delta> \<Delta>' \<equiv> (setD \<Delta> \<subseteq> setD \<Delta>')"
-
-nominal_function dc_of :: "branch_s \<Rightarrow> string" where
- "dc_of (AS_branch dc _ _) = dc"
- apply(auto,simp add: eqvt_def dc_of_graph_aux_def)
- using s_branch_s_branch_list.exhaust by metis
-nominal_termination (eqvt) by lexicographic_order
-
-lemma delta_sim_fresh:
- assumes "\<Theta> \<turnstile> \<delta> \<sim> \<Delta>" and "atom u \<sharp> \<delta>"
- shows "atom u \<sharp> \<Delta>"
- using assms proof(induct rule : delta_sim.inducts)
- case (delta_sim_nilI \<Theta>)
- then show ?case using fresh_def supp_DNil by blast
-next
- case (delta_sim_consI \<Theta> \<delta> \<Delta> v \<tau> u')
- hence "\<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f \<tau>" using check_v_wf by meson
- hence "supp \<tau> = {}" using wfT_supp by fastforce
- moreover have "atom u \<sharp> u'" using delta_sim_consI fresh_Cons fresh_Pair by blast
- moreover have "atom u \<sharp> \<Delta>" using delta_sim_consI fresh_Cons by blast
- ultimately show ?case using fresh_Pair fresh_DCons fresh_def by blast
-qed
-
-lemma delta_sim_v:
- fixes \<Delta>::\<Delta>
- assumes "\<Theta> \<turnstile> \<delta> \<sim> \<Delta> " and "(u,v) \<in> set \<delta>" and "(u,\<tau>) \<in> setD \<Delta>" and "\<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f \<Delta>"
- shows "\<Theta> ; {||} ; GNil \<turnstile> v \<Leftarrow> \<tau>"
- using assms proof(induct \<delta> arbitrary: \<Delta>)
- case Nil
- then show ?case by auto
-next
- case (Cons uv \<delta>)
- obtain u' and v' where uv : "uv=(u',v')" by fastforce
- show ?case proof(cases "u'=u")
- case True
- hence *:"\<Theta> \<turnstile> ((u,v')#\<delta>) \<sim> \<Delta>" using uv Cons by blast
- then obtain \<tau>' and \<Delta>' where tt: "\<Theta> ; {||} ; GNil \<turnstile> v' \<Leftarrow> \<tau>' \<and> u \<notin> fst ` set \<delta> \<and> \<Delta> = (u,\<tau>')#\<^sub>\<Delta>\<Delta>'" using delta_sim_elims(3)[OF *] by metis
- moreover hence "v'=v" using Cons True
- by (metis Pair_inject fst_conv image_eqI set_ConsD uv)
- moreover have "\<tau>=\<tau>'" using wfD_unique tt Cons
- setD.simps list.set_intros by blast
- ultimately show ?thesis by metis
- next
- case False
- hence *:"\<Theta> \<turnstile> ((u',v')#\<delta>) \<sim> \<Delta>" using uv Cons by blast
- then obtain \<tau>' and \<Delta>' where tt: "\<Theta> \<turnstile> \<delta> \<sim> \<Delta>' \<and> \<Theta> ; {||} ; GNil \<turnstile> v' \<Leftarrow> \<tau>' \<and> u' \<notin> fst ` set \<delta> \<and> \<Delta> = (u',\<tau>')#\<^sub>\<Delta>\<Delta>'" using delta_sim_elims(3)[OF *] by metis
-
- moreover hence "\<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f \<Delta>'" using wfD_elims Cons delta_sim_elims by metis
- ultimately show ?thesis using Cons
- using False by auto
- qed
-qed
-
-lemma delta_sim_delta_lookup:
- assumes "\<Theta> \<turnstile> \<delta> \<sim> \<Delta> " and "(u, \<lbrace> z : b | c \<rbrace>) \<in> setD \<Delta>"
- shows "\<exists>v. (u,v) \<in> set \<delta>"
- using assms by(induct rule: delta_sim.inducts,auto+)
-
-lemma update_d_stable:
- "fst ` set \<delta> = fst ` set (update_d \<delta> u v)"
-proof(induct \<delta>)
- case Nil
- then show ?case by auto
-next
- case (Cons a \<delta>)
- then show ?case using update_d.simps
- by (metis (no_types, lifting) eq_fst_iff image_cong image_insert list.simps(15) prod.exhaust_sel)
-qed
-
-lemma update_d_sim:
- fixes \<Delta>::\<Delta>
- assumes "\<Theta> \<turnstile> \<delta> \<sim> \<Delta>" and "\<Theta> ; {||} ; GNil \<turnstile> v \<Leftarrow> \<tau>" and "(u,\<tau>) \<in> setD \<Delta>" and "\<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f \<Delta>"
- shows "\<Theta> \<turnstile> (update_d \<delta> u v) \<sim> \<Delta>"
- using assms proof(induct \<delta> arbitrary: \<Delta>)
- case Nil
- then show ?case using delta_sim_consI by simp
-next
- case (Cons uv \<delta>)
- obtain u' and v' where uv : "uv=(u',v')" by fastforce
-
- hence *:"\<Theta> \<turnstile> ((u',v')#\<delta>) \<sim> \<Delta>" using uv Cons by blast
- then obtain \<tau>' and \<Delta>' where tt: "\<Theta> \<turnstile> \<delta> \<sim> \<Delta>' \<and> \<Theta> ; {||} ; GNil \<turnstile> v' \<Leftarrow> \<tau>' \<and> u' \<notin> fst ` set \<delta> \<and> \<Delta> = (u',\<tau>')#\<^sub>\<Delta>\<Delta>'" using delta_sim_elims * by metis
-
- show ?case proof(cases "u=u'")
- case True
- then have "(u,\<tau>') \<in> setD \<Delta>" using tt by auto
- then have "\<tau> = \<tau>'" using Cons wfD_unique by metis
- moreover have "update_d ((u',v')#\<delta>) u v = ((u',v)#\<delta>)" using update_d.simps True by presburger
- ultimately show ?thesis using delta_sim_consI tt Cons True
- by (simp add: tt uv)
- next
- case False
- have "\<Theta> \<turnstile> (u',v') # (update_d \<delta> u v) \<sim> (u',\<tau>')#\<^sub>\<Delta>\<Delta>'"
- proof(rule delta_sim_consI)
- show "\<Theta> \<turnstile> update_d \<delta> u v \<sim> \<Delta>' " using Cons using delta_sim_consI
- delta_sim.simps update_d.simps Cons delta_sim_elims uv tt
- False fst_conv set_ConsD wfG_elims wfD_elims by (metis setD_ConsD)
- show "\<Theta> ; {||} ; GNil \<turnstile> v' \<Leftarrow> \<tau>'" using tt by auto
- show "u' \<notin> fst ` set (update_d \<delta> u v)" using update_d.simps Cons update_d_stable tt by auto
- qed
- thus ?thesis using False update_d.simps uv
- by (simp add: tt)
- qed
-qed
-
-section \<open>Preservation\<close>
-text \<open>Types are preserved under reduction step. Broken down into lemmas about different operations\<close>
-
-subsection \<open>Function Application\<close>
-
-lemma check_s_x_fresh:
- fixes x::x and s::s
- assumes "\<Theta> ; \<Phi> ; B ; GNil ; D \<turnstile> s \<Leftarrow> \<tau>"
- shows "atom x \<sharp> s \<and> atom x \<sharp> \<tau> \<and> atom x \<sharp> D"
-proof -
- have "\<Theta> ; \<Phi> ; B ; GNil ; D \<turnstile>\<^sub>w\<^sub>f s : b_of \<tau>" using check_s_wf[OF assms] by auto
- moreover have "\<Theta> ; B ; GNil \<turnstile>\<^sub>w\<^sub>f \<tau> " using check_s_wf assms by auto
- moreover have "\<Theta> ; B ; GNil \<turnstile>\<^sub>w\<^sub>f D" using check_s_wf assms by auto
- ultimately show ?thesis using wf_supp x_fresh_u
- by (meson fresh_GNil wfS_x_fresh wfT_x_fresh wfD_x_fresh)
-qed
-
-lemma check_funtyp_subst_b:
- fixes b'::b
- assumes "check_funtyp \<Theta> \<Phi> {|bv|} (AF_fun_typ x b c \<tau> s)" and \<open> \<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f b' \<close>
- shows "check_funtyp \<Theta> \<Phi> {||} (AF_fun_typ x b[bv::=b']\<^sub>b\<^sub>b (c[bv::=b']\<^sub>c\<^sub>b) \<tau>[bv::=b']\<^sub>\<tau>\<^sub>b s[bv::=b']\<^sub>s\<^sub>b)"
- using assms proof (nominal_induct "{|bv|}" "AF_fun_typ x b c \<tau> s" rule: check_funtyp.strong_induct)
- case (check_funtypI x' \<Theta> \<Phi> c' s' \<tau>')
- have "check_funtyp \<Theta> \<Phi> {||} (AF_fun_typ x' b[bv::=b']\<^sub>b\<^sub>b (c'[bv::=b']\<^sub>c\<^sub>b) \<tau>'[bv::=b']\<^sub>\<tau>\<^sub>b s'[bv::=b']\<^sub>s\<^sub>b)" proof
- show \<open>atom x' \<sharp> (\<Theta>, \<Phi>, {||}::bv fset, b[bv::=b']\<^sub>b\<^sub>b)\<close> using check_funtypI fresh_prodN x_fresh_b fresh_empty_fset by metis
-
- have \<open> \<Theta> ; \<Phi> ; {||} ; ((x', b, c') #\<^sub>\<Gamma> GNil)[bv::=b']\<^sub>\<Gamma>\<^sub>b ; []\<^sub>\<Delta>[bv::=b']\<^sub>\<Delta>\<^sub>b \<turnstile> s'[bv::=b']\<^sub>s\<^sub>b \<Leftarrow> \<tau>'[bv::=b']\<^sub>\<tau>\<^sub>b\<close> proof(rule subst_b_check_s)
- show \<open> \<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f b' \<close> using check_funtypI by metis
- show \<open>{|bv|} = {|bv|}\<close> by auto
- show \<open> \<Theta> ; \<Phi> ; {|bv|} ; (x', b, c') #\<^sub>\<Gamma> GNil ; []\<^sub>\<Delta> \<turnstile> s' \<Leftarrow> \<tau>'\<close> using check_funtypI by metis
- qed
-
- thus \<open> \<Theta> ; \<Phi> ; {||} ; (x', b[bv::=b']\<^sub>b\<^sub>b, c'[bv::=b']\<^sub>c\<^sub>b) #\<^sub>\<Gamma> GNil ; []\<^sub>\<Delta> \<turnstile> s'[bv::=b']\<^sub>s\<^sub>b \<Leftarrow> \<tau>'[bv::=b']\<^sub>\<tau>\<^sub>b\<close>
- using subst_gb.simps subst_db.simps by simp
- qed
-
- moreover have "(AF_fun_typ x b c \<tau> s) = (AF_fun_typ x' b c' \<tau>' s')" using fun_typ.eq_iff check_funtypI by metis
- moreover hence "(AF_fun_typ x b[bv::=b']\<^sub>b\<^sub>b (c[bv::=b']\<^sub>c\<^sub>b) \<tau>[bv::=b']\<^sub>\<tau>\<^sub>b s[bv::=b']\<^sub>s\<^sub>b) = (AF_fun_typ x' b[bv::=b']\<^sub>b\<^sub>b (c'[bv::=b']\<^sub>c\<^sub>b) \<tau>'[bv::=b']\<^sub>\<tau>\<^sub>b s'[bv::=b']\<^sub>s\<^sub>b)"
- using subst_ft_b.simps by metis
- ultimately show ?case by metis
-qed
-
-lemma funtyp_simple_check:
- fixes s::s and \<Delta>::\<Delta> and \<tau>::\<tau> and v::v
- assumes "check_funtyp \<Theta> \<Phi> ({||}::bv fset) (AF_fun_typ x b c \<tau> s)" and
- "\<Theta> ; {||} ; GNil \<turnstile> v \<Leftarrow> \<lbrace> x : b | c \<rbrace>"
- shows "\<Theta> ; \<Phi> ; {||} ; GNil ; DNil \<turnstile> s[x::=v]\<^sub>s\<^sub>v \<Leftarrow> \<tau>[x::=v]\<^sub>\<tau>\<^sub>v"
- using assms proof(nominal_induct " ({||}::bv fset)" "AF_fun_typ x b c \<tau> s" avoiding: v x rule: check_funtyp.strong_induct)
- case (check_funtypI x' \<Theta> \<Phi> c' s' \<tau>')
-
- hence eq1: "\<lbrace> x' : b | c' \<rbrace> = \<lbrace> x : b | c \<rbrace>" using funtyp_eq_iff_equalities by metis
-
- obtain x'' and c'' where xf:"\<lbrace> x : b | c \<rbrace> = \<lbrace> x'' : b | c'' \<rbrace> \<and> atom x'' \<sharp> (x',v) \<and> atom x'' \<sharp> (x,c)" using obtain_fresh_z3 by metis
- moreover have "atom x' \<sharp> c''" proof -
- have "supp \<lbrace> x'' : b | c'' \<rbrace> = {}" using eq1 check_funtypI xf check_v_wf wfT_nil_supp by metis
- hence "supp c'' \<subseteq> { atom x'' }" using \<tau>.supp eq1 xf by (auto simp add: freshers)
- moreover have "atom x' \<noteq> atom x''" using xf fresh_Pair fresh_x_neq by metis
- ultimately show ?thesis using xf fresh_Pair fresh_x_neq fresh_def fresh_at_base by blast
- qed
- ultimately have eq2: "c''[x''::=[ x' ]\<^sup>v]\<^sub>c\<^sub>v = c'" using eq1 type_eq_subst_eq3(1)[of x' b c' x'' b c''] by metis
-
- have "atom x' \<sharp> c" proof -
- have "supp \<lbrace> x : b | c \<rbrace> = {}" using eq1 check_funtypI xf check_v_wf wfT_nil_supp by metis
- hence "supp c \<subseteq> { atom x }" using \<tau>.supp by auto
- moreover have "atom x \<noteq> atom x'" using check_funtypI fresh_Pair fresh_x_neq by metis
- ultimately show ?thesis using fresh_def by force
- qed
- hence eq: " c[x::=[ x' ]\<^sup>v]\<^sub>c\<^sub>v = c' \<and> s'[x'::=v]\<^sub>s\<^sub>v = s[x::=v]\<^sub>s\<^sub>v \<and> \<tau>'[x'::=v]\<^sub>\<tau>\<^sub>v = \<tau>[x::=v]\<^sub>\<tau>\<^sub>v"
- using funtyp_eq_iff_equalities type_eq_subst_eq3 check_funtypI by metis
-
- have " \<Theta> ; \<Phi> ; {||} ; ((x', b, c''[x''::=[ x' ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> GNil)[x'::=v]\<^sub>\<Gamma>\<^sub>v ; []\<^sub>\<Delta>[x'::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> s'[x'::=v]\<^sub>s\<^sub>v \<Leftarrow> \<tau>'[x'::=v]\<^sub>\<tau>\<^sub>v"
- proof(rule subst_check_check_s )
- show \<open>\<Theta> ; {||} ; GNil \<turnstile> v \<Leftarrow> \<lbrace> x'' : b | c'' \<rbrace>\<close> using check_funtypI eq1 xf by metis
- show \<open>atom x'' \<sharp> (x', v)\<close> using check_funtypI fresh_x_neq fresh_Pair xf by metis
- show \<open> \<Theta> ; \<Phi> ; {||} ; (x', b, c''[x''::=[ x' ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> GNil ; []\<^sub>\<Delta> \<turnstile> s' \<Leftarrow> \<tau>'\<close> using check_funtypI eq2 by metis
- show \<open> (x', b, c''[x''::=[ x' ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> GNil = GNil @ (x', b, c''[x''::=[ x' ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> GNil\<close> using append_g.simps by auto
- qed
- hence " \<Theta>; \<Phi>; {||}; GNil; []\<^sub>\<Delta> \<turnstile> s'[x'::=v]\<^sub>s\<^sub>v \<Leftarrow> \<tau>'[x'::=v]\<^sub>\<tau>\<^sub>v" using subst_gv.simps subst_dv.simps by auto
- thus ?case using eq by auto
-qed
-
-lemma funtypq_simple_check:
- fixes s::s and \<Delta>::\<Delta> and \<tau>::\<tau> and v::v
- assumes "check_funtypq \<Theta> \<Phi> (AF_fun_typ_none (AF_fun_typ x b c t s))" and
- "\<Theta> ; {||} ; GNil \<turnstile> v \<Leftarrow> \<lbrace> x : b | c \<rbrace>"
- shows "\<Theta>; \<Phi>; {||}; GNil; DNil \<turnstile> s[x::=v]\<^sub>s\<^sub>v \<Leftarrow> t[x::=v]\<^sub>\<tau>\<^sub>v"
- using assms proof(nominal_induct "(AF_fun_typ_none (AF_fun_typ x b c t s))" avoiding: v rule: check_funtypq.strong_induct)
- case (check_fundefq_simpleI \<Theta> \<Phi> x' c' t' s')
- hence eq: "\<lbrace> x : b | c \<rbrace> = \<lbrace> x' : b | c' \<rbrace> \<and> s'[x'::=v]\<^sub>s\<^sub>v = s[x::=v]\<^sub>s\<^sub>v \<and> t[x::=v]\<^sub>\<tau>\<^sub>v = t'[x'::=v]\<^sub>\<tau>\<^sub>v"
- using funtyp_eq_iff_equalities by metis
- hence "\<Theta>; \<Phi>; {||}; GNil; []\<^sub>\<Delta> \<turnstile> s'[x'::=v]\<^sub>s\<^sub>v \<Leftarrow> t'[x'::=v]\<^sub>\<tau>\<^sub>v"
- using funtyp_simple_check[OF check_fundefq_simpleI(1)] check_fundefq_simpleI by metis
- thus ?case using eq by metis
-qed
-
-lemma funtyp_poly_eq_iff_equalities:
- assumes "[[atom bv']]lst. AF_fun_typ x' b'' c' t' s' = [[atom bv]]lst. AF_fun_typ x b c t s"
- shows "\<lbrace> x' : b''[bv'::=b']\<^sub>b\<^sub>b | c'[bv'::=b']\<^sub>c\<^sub>b \<rbrace> = \<lbrace> x : b[bv::=b']\<^sub>b\<^sub>b | c[bv::=b']\<^sub>c\<^sub>b \<rbrace> \<and>
- s'[bv'::=b']\<^sub>s\<^sub>b[x'::=v]\<^sub>s\<^sub>v = s[bv::=b']\<^sub>s\<^sub>b[x::=v]\<^sub>s\<^sub>v \<and> t'[bv'::=b']\<^sub>\<tau>\<^sub>b[x'::=v]\<^sub>\<tau>\<^sub>v = t[bv::=b']\<^sub>\<tau>\<^sub>b[x::=v]\<^sub>\<tau>\<^sub>v"
-proof -
- have "subst_ft_b (AF_fun_typ x' b'' c' t' s') bv' b' = subst_ft_b (AF_fun_typ x b c t s) bv b'"
- using subst_b_flip_eq_two subst_b_fun_typ_def assms by metis
- thus ?thesis using fun_typ.eq_iff subst_ft_b.simps funtyp_eq_iff_equalities subst_tb.simps
- by (metis (full_types) assms fun_poly_arg_unique)
-
-qed
-
-lemma funtypq_poly_check:
- fixes s::s and \<Delta>::\<Delta> and \<tau>::\<tau> and v::v and b'::b
- assumes "check_funtypq \<Theta> \<Phi> (AF_fun_typ_some bv (AF_fun_typ x b c t s))" and
- "\<Theta> ; {||} ; GNil \<turnstile> v \<Leftarrow> \<lbrace> x : b[bv::=b']\<^sub>b\<^sub>b | c[bv::=b']\<^sub>c\<^sub>b \<rbrace>" and
- "\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f b'"
- shows "\<Theta>; \<Phi>; {||}; GNil; DNil \<turnstile> s[bv::=b']\<^sub>s\<^sub>b[x::=v]\<^sub>s\<^sub>v \<Leftarrow> t[bv::=b']\<^sub>\<tau>\<^sub>b[x::=v]\<^sub>\<tau>\<^sub>v"
- using assms proof(nominal_induct "(AF_fun_typ_some bv (AF_fun_typ x b c t s))" avoiding: v rule: check_funtypq.strong_induct)
- case (check_funtypq_polyI bv' \<Theta> \<Phi> x' b'' c' t' s')
-
- hence **:"\<lbrace> x' : b''[bv'::=b']\<^sub>b\<^sub>b | c'[bv'::=b']\<^sub>c\<^sub>b \<rbrace> = \<lbrace> x : b[bv::=b']\<^sub>b\<^sub>b | c[bv::=b']\<^sub>c\<^sub>b \<rbrace> \<and>
- s'[bv'::=b']\<^sub>s\<^sub>b[x'::=v]\<^sub>s\<^sub>v = s[bv::=b']\<^sub>s\<^sub>b[x::=v]\<^sub>s\<^sub>v \<and> t'[bv'::=b']\<^sub>\<tau>\<^sub>b[x'::=v]\<^sub>\<tau>\<^sub>v = t[bv::=b']\<^sub>\<tau>\<^sub>b[x::=v]\<^sub>\<tau>\<^sub>v"
- using funtyp_poly_eq_iff_equalities by metis
-
- have *:"check_funtyp \<Theta> \<Phi> {||} (AF_fun_typ x' b''[bv'::=b']\<^sub>b\<^sub>b (c'[bv'::=b']\<^sub>c\<^sub>b) (t'[bv'::=b']\<^sub>\<tau>\<^sub>b) s'[bv'::=b']\<^sub>s\<^sub>b)"
- using check_funtyp_subst_b[OF check_funtypq_polyI(5) check_funtypq_polyI(8)] by metis
- moreover have "\<Theta> ; {||} ; GNil \<turnstile> v \<Leftarrow> \<lbrace> x' : b''[bv'::=b']\<^sub>b\<^sub>b | c'[bv'::=b']\<^sub>c\<^sub>b \<rbrace>" using ** check_funtypq_polyI by metis
- ultimately have "\<Theta>; \<Phi>; {||}; GNil; []\<^sub>\<Delta> \<turnstile> s'[bv'::=b']\<^sub>s\<^sub>b[x'::=v]\<^sub>s\<^sub>v \<Leftarrow> t'[bv'::=b']\<^sub>\<tau>\<^sub>b[x'::=v]\<^sub>\<tau>\<^sub>v"
- using funtyp_simple_check[OF *] check_funtypq_polyI by metis
- thus ?case using ** by metis
-
-qed
-
-lemma fundef_simple_check:
- fixes s::s and \<Delta>::\<Delta> and \<tau>::\<tau> and v::v
- assumes "check_fundef \<Theta> \<Phi> (AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c t s)))" and
- "\<Theta> ; {||} ; GNil \<turnstile> v \<Leftarrow> \<lbrace> x : b | c \<rbrace>" and "\<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f \<Delta>"
- shows "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> s[x::=v]\<^sub>s\<^sub>v \<Leftarrow> t[x::=v]\<^sub>\<tau>\<^sub>v"
- using assms proof(nominal_induct "(AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c t s)))" avoiding: v rule: check_fundef.strong_induct)
- case (check_fundefI \<Theta> \<Phi>)
- then show ?case using funtypq_simple_check[THEN check_s_d_weakening(1) ] setD.simps by auto
-qed
-
-lemma fundef_poly_check:
- fixes s::s and \<Delta>::\<Delta> and \<tau>::\<tau> and v::v and b'::b
- assumes "check_fundef \<Theta> \<Phi> (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c t s)))" and
- "\<Theta> ; {||} ; GNil \<turnstile> v \<Leftarrow> \<lbrace> x : b[bv::=b']\<^sub>b\<^sub>b | c[bv::=b']\<^sub>c\<^sub>b \<rbrace>" and "\<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f \<Delta>" and "\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f b'"
- shows "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> s[bv::=b']\<^sub>s\<^sub>b[x::=v]\<^sub>s\<^sub>v \<Leftarrow> t[bv::=b']\<^sub>\<tau>\<^sub>b[x::=v]\<^sub>\<tau>\<^sub>v"
- using assms proof(nominal_induct "(AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c t s)))" avoiding: v rule: check_fundef.strong_induct)
- case (check_fundefI \<Theta> \<Phi>)
- then show ?case using funtypq_poly_check[THEN check_s_d_weakening(1) ] setD.simps by auto
-qed
-
-lemma preservation_app:
- assumes
- "Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x1 b1 c1 \<tau>1' s1'))) = lookup_fun \<Phi> f" and "(\<forall>fd\<in>set \<Phi>. check_fundef \<Theta> \<Phi> fd)"
- shows "\<Theta> ; \<Phi> ; B ; G ; \<Delta> \<turnstile> ss \<Leftarrow> \<tau> \<Longrightarrow> B = {||} \<Longrightarrow> G = GNil \<Longrightarrow> ss = LET x = (AE_app f v) IN s \<Longrightarrow>
- \<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> LET x : (\<tau>1'[x1::=v]\<^sub>\<tau>\<^sub>v) = (s1'[x1::=v]\<^sub>s\<^sub>v) IN s \<Leftarrow> \<tau>" and
- "check_branch_s \<Theta> \<Phi> \<B> GNil \<Delta> tid dc const v cs \<tau> \<Longrightarrow> True" and
- "check_branch_list \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist v css \<tau> \<Longrightarrow> True"
- using assms proof(nominal_induct \<tau> and \<tau> and \<tau> avoiding: v rule: check_s_check_branch_s_check_branch_list.strong_induct)
- case (check_letI x2 \<Theta> \<Phi> \<B> \<Gamma> \<Delta> e \<tau> z s2 b c)
-
- hence eq: "e = (AE_app f v)" by simp
- hence *:"\<Theta> ; \<Phi> ; {||} ;GNil ; \<Delta> \<turnstile> (AE_app f v) \<Rightarrow> \<lbrace> z : b | c \<rbrace>" using check_letI by auto
-
- then obtain x3 b3 c3 \<tau>3 s3 where
- **:"\<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f \<Delta> \<and> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<and> Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x3 b3 c3 \<tau>3 s3))) = lookup_fun \<Phi> f \<and>
- \<Theta> ; {||} ; GNil \<turnstile> v \<Leftarrow> \<lbrace> x3 : b3 | c3 \<rbrace> \<and> atom x3 \<sharp> (\<Theta>, \<Phi>, ({||}::bv fset), GNil, \<Delta>, v, \<lbrace> z : b | c \<rbrace>) \<and> \<tau>3[x3::=v]\<^sub>\<tau>\<^sub>v = \<lbrace> z : b | c \<rbrace>"
- using infer_e_elims(6)[OF *] subst_defs by metis
-
- obtain z3 where z3:"\<lbrace> x3 : b3 | c3 \<rbrace> = \<lbrace> z3 : b3 | c3[x3::=V_var z3]\<^sub>c\<^sub>v \<rbrace> \<and> atom z3 \<sharp> (x3, v,c3,x1,c1)" using obtain_fresh_z3 by metis
-
- have seq:"[[atom x3]]lst. s3 = [[atom x1]]lst. s1'" using fun_def_eq check_letI ** option.inject by metis
-
- let ?ft = "AF_fun_typ x3 b3 c3 \<tau>3 s3"
-
-
- have sup: "supp \<tau>3 \<subseteq> { atom x3} \<and> supp s3 \<subseteq> { atom x3}" using wfPhi_f_supp ** by metis
-
- have "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let2 x2 \<tau>3[x3::=v]\<^sub>\<tau>\<^sub>v (s3[x3::=v]\<^sub>s\<^sub>v) s2 \<Leftarrow> \<tau>" proof
- show \<open>atom x2 \<sharp> (\<Theta>, \<Phi>, {||}::bv fset, GNil, \<Delta>, \<tau>3[x3::=v]\<^sub>\<tau>\<^sub>v, s3[x3::=v]\<^sub>s\<^sub>v, \<tau>)\<close>
- unfolding fresh_prodN using check_letI fresh_subst_v_if subst_v_\<tau>_def sup
- by (metis all_not_in_conv fresh_def fresh_empty_fset fresh_subst_sv_if fresh_subst_tv_if singleton_iff subset_singleton_iff)
-
- show \<open> \<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> s3[x3::=v]\<^sub>s\<^sub>v \<Leftarrow> \<tau>3[x3::=v]\<^sub>\<tau>\<^sub>v\<close> proof(rule fundef_simple_check)
- show \<open>check_fundef \<Theta> \<Phi> (AF_fundef f (AF_fun_typ_none (AF_fun_typ x3 b3 c3 \<tau>3 s3)))\<close> using ** check_letI lookup_fun_member by metis
- show \<open>\<Theta> ; {||} ; GNil \<turnstile> v \<Leftarrow> \<lbrace> x3 : b3 | c3 \<rbrace>\<close> using ** by auto
- show \<open> \<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using ** by auto
- qed
- show \<open> \<Theta> ; \<Phi> ; {||} ; (x2, b_of \<tau>3[x3::=v]\<^sub>\<tau>\<^sub>v, c_of \<tau>3[x3::=v]\<^sub>\<tau>\<^sub>v x2) #\<^sub>\<Gamma> GNil ; \<Delta> \<turnstile> s2 \<Leftarrow> \<tau>\<close>
- using check_letI ** b_of.simps c_of.simps subst_defs by metis
- qed
-
- moreover have "AS_let2 x2 \<tau>3[x3::=v]\<^sub>\<tau>\<^sub>v (s3[x3::=v]\<^sub>s\<^sub>v) s2 = AS_let2 x (\<tau>1'[x1::=v]\<^sub>\<tau>\<^sub>v) (s1'[x1::=v]\<^sub>s\<^sub>v) s" proof -
- have *: "[[atom x2]]lst. s2 = [[atom x]]lst. s" using check_letI s_branch_s_branch_list.eq_iff by auto
- moreover have " \<tau>3[x3::=v]\<^sub>\<tau>\<^sub>v = \<tau>1'[x1::=v]\<^sub>\<tau>\<^sub>v" using fun_ret_unique ** check_letI by metis
- moreover have "s3[x3::=v]\<^sub>s\<^sub>v = (s1'[x1::=v]\<^sub>s\<^sub>v)" using subst_v_flip_eq_two subst_v_s_def seq by metis
- ultimately show ?thesis using s_branch_s_branch_list.eq_iff by metis
- qed
-
- ultimately show ?case using check_letI by auto
-qed(auto+)
-
-lemma fresh_subst_v_subst_b:
- fixes x2::x and tm::"'a::{has_subst_v,has_subst_b}" and x::x
- assumes "supp tm \<subseteq> { atom bv, atom x }" and "atom x2 \<sharp> v"
- shows "atom x2 \<sharp> tm[bv::=b]\<^sub>b[x::=v]\<^sub>v"
- using assms proof(cases "x2=x")
- case True
- then show ?thesis using fresh_subst_v_if assms by blast
-next
- case False
- hence "atom x2 \<sharp> tm" using assms fresh_def fresh_at_base by force
- hence "atom x2 \<sharp> tm[bv::=b]\<^sub>b" using assms fresh_subst_if x_fresh_b False by force
- then show ?thesis using fresh_subst_v_if assms by auto
-qed
-
-lemma preservation_poly_app:
- assumes
- "Some (AF_fundef f (AF_fun_typ_some bv1 (AF_fun_typ x1 b1 c1 \<tau>1' s1'))) = lookup_fun \<Phi> f" and "(\<forall>fd\<in>set \<Phi>. check_fundef \<Theta> \<Phi> fd)"
- shows "\<Theta> ; \<Phi> ; B ; G ; \<Delta> \<turnstile> ss \<Leftarrow> \<tau> \<Longrightarrow> B = {||} \<Longrightarrow> G = GNil \<Longrightarrow> ss = LET x = (AE_appP f b' v) IN s \<Longrightarrow> \<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f b' \<Longrightarrow>
- \<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> LET x : (\<tau>1'[bv1::=b']\<^sub>\<tau>\<^sub>b[x1::=v]\<^sub>\<tau>\<^sub>v) = (s1'[bv1::=b']\<^sub>s\<^sub>b[x1::=v]\<^sub>s\<^sub>v) IN s \<Leftarrow> \<tau>" and
- "check_branch_s \<Theta> \<Phi> \<B> GNil \<Delta> tid dc const v cs \<tau> \<Longrightarrow> True" and
- "check_branch_list \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist v css \<tau> \<Longrightarrow> True"
- using assms proof(nominal_induct \<tau> and \<tau> and \<tau> avoiding: v x1 rule: check_s_check_branch_s_check_branch_list.strong_induct)
- case (check_letI x2 \<Theta> \<Phi> \<B> \<Gamma> \<Delta> e \<tau> z s2 b c)
-
- hence eq: "e = (AE_appP f b' v)" by simp
- hence *:"\<Theta> ; \<Phi> ; {||} ;GNil ; \<Delta> \<turnstile> (AE_appP f b' v) \<Rightarrow> \<lbrace> z : b | c \<rbrace>" using check_letI by auto
-
- then obtain x3 b3 c3 \<tau>3 s3 bv3 where
- **:"\<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f \<Delta> \<and> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<and> Some (AF_fundef f (AF_fun_typ_some bv3 (AF_fun_typ x3 b3 c3 \<tau>3 s3))) = lookup_fun \<Phi> f \<and>
- \<Theta> ; {||} ; GNil \<turnstile> v \<Leftarrow> \<lbrace> x3 : b3[bv3::=b']\<^sub>b\<^sub>b | c3[bv3::=b']\<^sub>c\<^sub>b \<rbrace> \<and> atom x3 \<sharp> GNil \<and> \<tau>3[bv3::=b']\<^sub>\<tau>\<^sub>b[x3::=v]\<^sub>\<tau>\<^sub>v = \<lbrace> z : b | c \<rbrace>
- \<and> \<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f b'"
- using infer_e_elims(21)[OF *] subst_defs by metis
-
- obtain z3 where z3:"\<lbrace> x3 : b3 | c3 \<rbrace> = \<lbrace> z3 : b3 | c3[x3::=V_var z3]\<^sub>c\<^sub>v \<rbrace> \<and> atom z3 \<sharp> (x3, v,c3,x1,c1)" using obtain_fresh_z3 by metis
-
- let ?ft = "(AF_fun_typ x3 (b3[bv3::=b']\<^sub>b\<^sub>b) (c3[bv3::=b']\<^sub>c\<^sub>b) (\<tau>3[bv3::=b']\<^sub>\<tau>\<^sub>b) (s3[bv3::=b']\<^sub>s\<^sub>b))"
-
- have *:"check_fundef \<Theta> \<Phi> (AF_fundef f (AF_fun_typ_some bv3 (AF_fun_typ x3 b3 c3 \<tau>3 s3)))" using ** check_letI lookup_fun_member by metis
-
- hence ftq:"check_funtypq \<Theta> \<Phi> (AF_fun_typ_some bv3 (AF_fun_typ x3 b3 c3 \<tau>3 s3))" using check_fundef_elims by auto
-
- let ?ft = "AF_fun_typ_some bv3 (AF_fun_typ x3 b3 c3 \<tau>3 s3)"
-
- have sup: "supp \<tau>3 \<subseteq> { atom x3, atom bv3} \<and> supp s3 \<subseteq> { atom x3, atom bv3 }"
- using wfPhi_f_poly_supp_t wfPhi_f_poly_supp_s ** by metis
-
- have "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let2 x2 \<tau>3[bv3::=b']\<^sub>\<tau>\<^sub>b[x3::=v]\<^sub>\<tau>\<^sub>v (s3[bv3::=b']\<^sub>s\<^sub>b[x3::=v]\<^sub>s\<^sub>v) s2 \<Leftarrow> \<tau>"
- proof
- show \<open>atom x2 \<sharp> (\<Theta>, \<Phi>, {||}::bv fset, GNil, \<Delta>, \<tau>3[bv3::=b']\<^sub>\<tau>\<^sub>b[x3::=v]\<^sub>\<tau>\<^sub>v, s3[bv3::=b']\<^sub>s\<^sub>b[x3::=v]\<^sub>s\<^sub>v, \<tau>)\<close>
- proof -
-
- have "atom x2 \<sharp> \<tau>3[bv3::=b']\<^sub>\<tau>\<^sub>b[x3::=v]\<^sub>\<tau>\<^sub>v"
- using fresh_subst_v_subst_b subst_v_\<tau>_def subst_b_\<tau>_def \<open> atom x2 \<sharp> v\<close> sup by fastforce
- moreover have "atom x2 \<sharp> s3[bv3::=b']\<^sub>s\<^sub>b[x3::=v]\<^sub>s\<^sub>v"
- using fresh_subst_v_subst_b subst_v_s_def subst_b_s_def \<open> atom x2 \<sharp> v\<close> sup
- proof -
- have "\<forall>b. atom x2 = atom x3 \<or> atom x2 \<sharp> s3[bv3::=b]\<^sub>b"
- by (metis (no_types) check_letI.hyps(1) fresh_subst_sv_if(1) fresh_subst_v_subst_b insert_commute subst_v_s_def sup) (* 140 ms *)
- then show ?thesis
- by (metis check_letI.hyps(1) fresh_subst_sb_if fresh_subst_sv_if(1) has_subst_b_class.subst_b_fresh_x x_fresh_b) (* 78 ms *)
- qed
- ultimately show ?thesis using fresh_prodN check_letI by metis
- qed
-
- show \<open> \<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> s3[bv3::=b']\<^sub>s\<^sub>b[x3::=v]\<^sub>s\<^sub>v \<Leftarrow> \<tau>3[bv3::=b']\<^sub>\<tau>\<^sub>b[x3::=v]\<^sub>\<tau>\<^sub>v\<close> proof( rule fundef_poly_check)
- show \<open>check_fundef \<Theta> \<Phi> (AF_fundef f (AF_fun_typ_some bv3 (AF_fun_typ x3 b3 c3 \<tau>3 s3)))\<close>
- using ** lookup_fun_member check_letI by metis
- show \<open>\<Theta> ; {||} ; GNil \<turnstile> v \<Leftarrow> \<lbrace> x3 : b3[bv3::=b']\<^sub>b\<^sub>b | c3[bv3::=b']\<^sub>c\<^sub>b \<rbrace>\<close> using ** by metis
- show \<open> \<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using ** by metis
- show \<open> \<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f b' \<close> using ** by metis
- qed
- show \<open> \<Theta> ; \<Phi> ; {||} ; (x2, b_of \<tau>3[bv3::=b']\<^sub>\<tau>\<^sub>b[x3::=v]\<^sub>\<tau>\<^sub>v, c_of \<tau>3[bv3::=b']\<^sub>\<tau>\<^sub>b[x3::=v]\<^sub>\<tau>\<^sub>v x2) #\<^sub>\<Gamma> GNil ; \<Delta> \<turnstile> s2 \<Leftarrow> \<tau>\<close>
- using check_letI ** b_of.simps c_of.simps subst_defs by metis
- qed
-
- moreover have "AS_let2 x2 \<tau>3[bv3::=b']\<^sub>\<tau>\<^sub>b[x3::=v]\<^sub>\<tau>\<^sub>v (s3[bv3::=b']\<^sub>s\<^sub>b[x3::=v]\<^sub>s\<^sub>v) s2 = AS_let2 x (\<tau>1'[bv1::=b']\<^sub>\<tau>\<^sub>b[x1::=v]\<^sub>\<tau>\<^sub>v) (s1'[bv1::=b']\<^sub>s\<^sub>b[x1::=v]\<^sub>s\<^sub>v) s" proof -
- have *: "[[atom x2]]lst. s2 = [[atom x]]lst. s" using check_letI s_branch_s_branch_list.eq_iff by auto
- moreover have " \<tau>3[bv3::=b']\<^sub>\<tau>\<^sub>b[x3::=v]\<^sub>\<tau>\<^sub>v = \<tau>1'[bv1::=b']\<^sub>\<tau>\<^sub>b[x1::=v]\<^sub>\<tau>\<^sub>v" using fun_poly_ret_unique ** check_letI by metis
- moreover have "s3[bv3::=b']\<^sub>s\<^sub>b[x3::=v]\<^sub>s\<^sub>v = (s1'[bv1::=b']\<^sub>s\<^sub>b[x1::=v]\<^sub>s\<^sub>v)" using subst_v_flip_eq_two subst_v_s_def fun_poly_body_unique ** check_letI by metis
- ultimately show ?thesis using s_branch_s_branch_list.eq_iff by metis
- qed
-
- ultimately show ?case using check_letI by auto
-qed(auto+)
-
-lemma check_s_plus:
- assumes "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> LET x = (AE_op Plus (V_lit (L_num n1)) (V_lit (L_num n2))) IN s' \<Leftarrow> \<tau>"
- shows "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> LET x = (AE_val (V_lit (L_num (n1+n2)))) IN s' \<Leftarrow> \<tau>"
-proof -
- obtain t1 where 1: "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AE_op Plus (V_lit (L_num n1)) (V_lit (L_num n2)) \<Rightarrow> t1"
- using assms check_s_elims by metis
- then obtain z1 where 2: "t1 = \<lbrace> z1 : B_int | CE_val (V_var z1) == CE_op Plus ([V_lit (L_num n1)]\<^sup>c\<^sup>e) ([V_lit (L_num n2)]\<^sup>c\<^sup>e) \<rbrace>"
- using infer_e_plus by metis
-
- obtain z2 where 3: \<open>\<Theta> ; \<Phi> ; {||} ; GNil ; \<Delta> \<turnstile> AE_val (V_lit (L_num (n1+n2))) \<Rightarrow> \<lbrace> z2 : B_int | CE_val (V_var z2) == CE_val (V_lit (L_num (n1+n2))) \<rbrace>\<close>
- using infer_v_form infer_e_valI infer_v_litI infer_l.intros infer_e_wf 1
- by (simp add: fresh_GNil)
-
- let ?e = " (AE_op Plus (V_lit (L_num n1)) (V_lit (L_num n2)))"
-
- show ?thesis proof(rule subtype_let)
- show "\<Theta> ; \<Phi> ; {||} ; GNil ; \<Delta> \<turnstile> LET x = ?e IN s' \<Leftarrow> \<tau>" using assms by auto
- show "\<Theta> ; \<Phi> ; {||} ; GNil ; \<Delta> \<turnstile> ?e \<Rightarrow> t1" using 1 by auto
- show "\<Theta> ; \<Phi> ; {||} ; GNil ; \<Delta> \<turnstile> [ [ L_num (n1 + n2) ]\<^sup>v ]\<^sup>e \<Rightarrow> \<lbrace> z2 : B_int | CE_val (V_var z2) == CE_val (V_lit (L_num (n1+n2))) \<rbrace>" using 3 by auto
- show "\<Theta> ; {||} ; GNil \<turnstile> \<lbrace> z2 : B_int | CE_val (V_var z2) == CE_val (V_lit (L_num (n1+n2))) \<rbrace> \<lesssim> t1" using subtype_bop_arith
- by (metis "1" \<open>\<And>thesis. (\<And>z1. t1 = \<lbrace> z1 : B_int | [ [ z1 ]\<^sup>v ]\<^sup>c\<^sup>e == [ plus [ [ L_num n1 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ L_num n2 ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e \<rbrace> \<Longrightarrow> thesis) \<Longrightarrow> thesis\<close> infer_e_wf(2) opp.distinct(1) type_for_lit.simps(3))
- qed
-
-qed
-
-lemma check_s_leq:
- assumes "\<Theta> ; \<Phi> ; {||} ; GNil ; \<Delta> \<turnstile> LET x = (AE_op LEq (V_lit (L_num n1)) (V_lit (L_num n2))) IN s' \<Leftarrow> \<tau>"
- shows "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> LET x = (AE_val (V_lit (if (n1 \<le> n2) then L_true else L_false))) IN s' \<Leftarrow> \<tau>"
-proof -
- obtain t1 where 1: "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AE_op LEq (V_lit (L_num n1)) (V_lit (L_num n2)) \<Rightarrow> t1"
- using assms check_s_elims by metis
- then obtain z1 where 2: "t1 = \<lbrace> z1 : B_bool | CE_val (V_var z1) == CE_op LEq ([V_lit (L_num n1)]\<^sup>c\<^sup>e) ([V_lit (L_num n2)]\<^sup>c\<^sup>e) \<rbrace>"
- using infer_e_leq by auto
-
- obtain z2 where 3: \<open>\<Theta> ; \<Phi> ; {||} ; GNil ; \<Delta> \<turnstile> AE_val (V_lit ((if (n1 \<le> n2) then L_true else L_false))) \<Rightarrow> \<lbrace> z2 : B_bool | CE_val (V_var z2) == CE_val (V_lit ((if (n1 \<le> n2) then L_true else L_false))) \<rbrace>\<close>
- using infer_v_form infer_e_valI infer_v_litI infer_l.intros infer_e_wf 1
- fresh_GNil
- by simp
-
- show ?thesis proof(rule subtype_let)
- show \<open> \<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let x (AE_op LEq [ L_num n1 ]\<^sup>v [ L_num n2 ]\<^sup>v) s' \<Leftarrow> \<tau>\<close> using assms by auto
- show \<open>\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AE_op LEq [ L_num n1 ]\<^sup>v [ L_num n2 ]\<^sup>v \<Rightarrow> t1\<close> using 1 by auto
- show \<open>\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> [ [ if n1 \<le> n2 then L_true else L_false ]\<^sup>v ]\<^sup>e \<Rightarrow> \<lbrace> z2 : B_bool | CE_val (V_var z2) == CE_val (V_lit ((if (n1 \<le> n2) then L_true else L_false))) \<rbrace>\<close> using 3 by auto
- show \<open>\<Theta> ; {||} ; GNil \<turnstile> \<lbrace> z2 : B_bool | CE_val (V_var z2) == CE_val (V_lit ((if (n1 \<le> n2) then L_true else L_false))) \<rbrace> \<lesssim> t1\<close>
- using subtype_bop_arith[where opp=LEq] check_s_wf assms 2
- by (metis opp.distinct(1) subtype_bop_arith type_l_eq)
- qed
-qed
-
-lemma check_s_eq:
- assumes "\<Theta> ; \<Phi> ; {||} ; GNil ; \<Delta> \<turnstile> LET x = (AE_op Eq (V_lit (n1)) (V_lit ( n2))) IN s' \<Leftarrow> \<tau>"
- shows "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> LET x = (AE_val (V_lit (if (n1 = n2) then L_true else L_false))) IN s' \<Leftarrow> \<tau>"
-proof -
- obtain t1 where 1: "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AE_op Eq (V_lit (n1)) (V_lit (n2)) \<Rightarrow> t1"
- using assms check_s_elims by metis
- then obtain z1 where 2: "t1 = \<lbrace> z1 : B_bool | CE_val (V_var z1) == CE_op Eq ([V_lit (n1)]\<^sup>c\<^sup>e) ([V_lit (n2)]\<^sup>c\<^sup>e) \<rbrace>"
- using infer_e_leq by auto
-
- obtain z2 where 3: \<open>\<Theta> ; \<Phi> ; {||} ; GNil ; \<Delta> \<turnstile> AE_val (V_lit ((if (n1 = n2) then L_true else L_false))) \<Rightarrow> \<lbrace> z2 : B_bool | CE_val (V_var z2) == CE_val (V_lit ((if (n1 = n2) then L_true else L_false))) \<rbrace>\<close>
- using infer_v_form infer_e_valI infer_v_litI infer_l.intros infer_e_wf 1
- fresh_GNil
- by simp
-
- show ?thesis proof(rule subtype_let)
- show \<open> \<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let x (AE_op Eq [ n1 ]\<^sup>v [ n2 ]\<^sup>v) s' \<Leftarrow> \<tau>\<close> using assms by auto
- show \<open>\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AE_op Eq [ n1 ]\<^sup>v [ n2 ]\<^sup>v \<Rightarrow> t1\<close> using 1 by auto
- show \<open>\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> [ [ if n1 = n2 then L_true else L_false ]\<^sup>v ]\<^sup>e \<Rightarrow> \<lbrace> z2 : B_bool | CE_val (V_var z2) == CE_val (V_lit ((if (n1 = n2) then L_true else L_false))) \<rbrace>\<close> using 3 by auto
- show \<open>\<Theta> ; {||} ; GNil \<turnstile> \<lbrace> z2 : B_bool | CE_val (V_var z2) == CE_val (V_lit ((if (n1 = n2) then L_true else L_false))) \<rbrace> \<lesssim> t1\<close>
- proof -
- have " \<lbrace> z2 : B_bool | [ [ z2 ]\<^sup>v ]\<^sup>c\<^sup>e == [ eq [ [ n1 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ n2 ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e \<rbrace> = t1" using 2
- by (metis \<tau>_fresh_c fresh_opp_all infer_l_form2 infer_l_fresh ms_fresh_all(31) ms_fresh_all(33) obtain_fresh_z type_e_eq type_l_eq)
- moreover have "\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f GNil" using assms wfX_wfY by fastforce
- moreover have "base_for_lit n1 = base_for_lit n2" using 1 infer_e_wf wfE_elims(12) wfV_elims
- by metis
- ultimately show ?thesis using subtype_bop_eq[OF \<open>\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f GNil\<close>, of n1 n2 z2] by auto
- qed
- qed
-qed
-
-subsection \<open>Operators\<close>
-
-lemma preservation_plus:
- assumes "\<Theta>; \<Phi>; \<Delta> \<turnstile> \<langle> \<delta> , LET x = (AE_op Plus (V_lit (L_num n1)) (V_lit (L_num n2))) IN s' \<rangle> \<Leftarrow> \<tau>"
- shows "\<Theta>; \<Phi>; \<Delta> \<turnstile> \<langle> \<delta> , LET x = (AE_val (V_lit (L_num (n1+n2)))) IN s' \<rangle> \<Leftarrow> \<tau>"
-proof -
-
- have tt: "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let x (AE_op Plus (V_lit (L_num n1)) (V_lit (L_num n2))) s' \<Leftarrow> \<tau>" and dsim: "\<Theta> \<turnstile> \<delta> \<sim> \<Delta>" and fd:"(\<forall>fd\<in>set \<Phi>. check_fundef \<Theta> \<Phi> fd)"
- using assms config_type_elims by blast+
-
- hence "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile>AS_let x (AE_val (V_lit (L_num (n1+n2)))) s' \<Leftarrow> \<tau>" using check_s_plus assms by auto
-
- hence "\<Theta>; \<Phi>; \<Delta> \<turnstile> \<langle> \<delta> , AS_let x (AE_val (V_lit ( (L_num (n1+n2))))) s' \<rangle> \<Leftarrow> \<tau>" using dsim config_typeI fd by presburger
- then show ?thesis using dsim config_typeI
- by (meson order_refl)
-qed
-
-lemma preservation_leq:
- assumes "\<Theta>; \<Phi>; \<Delta> \<turnstile> \<langle> \<delta> , AS_let x (AE_op LEq (V_lit (L_num n1)) (V_lit (L_num n2))) s' \<rangle> \<Leftarrow> \<tau>"
- shows "\<Theta>; \<Phi>; \<Delta> \<turnstile> \<langle> \<delta> , AS_let x (AE_val (V_lit (((if (n1 \<le> n2) then L_true else L_false))))) s' \<rangle> \<Leftarrow> \<tau>"
-proof -
-
- have tt: "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let x (AE_op LEq (V_lit (L_num n1)) (V_lit (L_num n2))) s' \<Leftarrow> \<tau>" and dsim: "\<Theta> \<turnstile> \<delta> \<sim> \<Delta>" and fd:"(\<forall>fd\<in>set \<Phi>. check_fundef \<Theta> \<Phi> fd)"
- using assms config_type_elims by blast+
-
- hence "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile>AS_let x (AE_val (V_lit ( ((if (n1 \<le> n2) then L_true else L_false))))) s' \<Leftarrow> \<tau>" using check_s_leq assms by auto
-
- hence "\<Theta>; \<Phi>; \<Delta> \<turnstile> \<langle> \<delta> , AS_let x (AE_val (V_lit ( (((if (n1 \<le> n2) then L_true else L_false)))))) s' \<rangle> \<Leftarrow> \<tau>" using dsim config_typeI fd by presburger
- then show ?thesis using dsim config_typeI
- by (meson order_refl)
-qed
-
-lemma preservation_eq:
- assumes "\<Theta>; \<Phi>; \<Delta> \<turnstile> \<langle> \<delta> , AS_let x (AE_op Eq (V_lit (n1)) (V_lit (n2))) s' \<rangle> \<Leftarrow> \<tau>"
- shows "\<Theta>; \<Phi>; \<Delta> \<turnstile> \<langle> \<delta> , AS_let x (AE_val (V_lit (((if (n1 = n2) then L_true else L_false))))) s' \<rangle> \<Leftarrow> \<tau>"
-proof -
-
- have tt: "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let x (AE_op Eq (V_lit (n1)) (V_lit (n2))) s' \<Leftarrow> \<tau>" and dsim: "\<Theta> \<turnstile> \<delta> \<sim> \<Delta>" and fd:"(\<forall>fd\<in>set \<Phi>. check_fundef \<Theta> \<Phi> fd)"
- using assms config_type_elims by blast+
-
- hence "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile>AS_let x (AE_val (V_lit ( ((if (n1 = n2) then L_true else L_false))))) s' \<Leftarrow> \<tau>" using check_s_eq assms by auto
-
- hence "\<Theta>; \<Phi>; \<Delta> \<turnstile> \<langle> \<delta> , AS_let x (AE_val (V_lit ( (((if (n1 = n2) then L_true else L_false)))))) s' \<rangle> \<Leftarrow> \<tau>" using dsim config_typeI fd by presburger
- then show ?thesis using dsim config_typeI
- by (meson order_refl)
-qed
-
-subsection \<open>Let Statements\<close>
-
-lemma subst_s_abs_lst:
- fixes s::s and sa::s and v'::v
- assumes "[[atom x]]lst. s = [[atom xa]]lst. sa" and "atom xa \<sharp> v \<and> atom x \<sharp> v"
- shows "s[x::=v]\<^sub>s\<^sub>v = sa[xa::=v]\<^sub>s\<^sub>v"
-proof -
- obtain c'::x where cdash: "atom c' \<sharp> (v, x, xa, s, sa)" using obtain_fresh by blast
- moreover have " (x \<leftrightarrow> c') \<bullet> s = (xa \<leftrightarrow> c') \<bullet> sa" proof -
- have "atom c' \<sharp> (s, sa) \<and> atom c' \<sharp> (x, xa, s, sa)" using cdash by auto
- thus ?thesis using assms by auto
- qed
- ultimately show ?thesis using assms
- using subst_sv_flip by auto
-qed
-
-lemma check_let_val:
- fixes v::v and s::s
- shows "\<Theta> ; \<Phi> ; B ; G ; \<Delta> \<turnstile> ss \<Leftarrow> \<tau> \<Longrightarrow> B = {||} \<Longrightarrow> G = GNil \<Longrightarrow>
- ss = AS_let x (AE_val v) s \<or> ss = AS_let2 x t (AS_val v) s \<Longrightarrow> \<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> (s[x::=v]\<^sub>s\<^sub>v) \<Leftarrow> \<tau>" and
- "check_branch_s \<Theta> \<Phi> \<B> GNil \<Delta> tid dc const v cs \<tau> \<Longrightarrow> True" and
- "check_branch_list \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist v css \<tau> \<Longrightarrow> True"
-proof(nominal_induct \<tau> and \<tau> and \<tau> avoiding: v rule: check_s_check_branch_s_check_branch_list.strong_induct)
- case (check_letI x1 \<Theta> \<Phi> \<B> \<Gamma> \<Delta> e \<tau> z s1 b c)
- hence *:"e = AE_val v" by auto
- let ?G = "(x1, b, c[z::=V_var x1]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>"
- have "\<Theta> ; \<Phi> ; \<B> ; ?G[x1::=v]\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x1::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> s1[x1::=v]\<^sub>s\<^sub>v \<Leftarrow> \<tau>[x1::=v]\<^sub>\<tau>\<^sub>v"
- proof(rule subst_infer_check_s(1))
- show **:\<open> \<Theta> ; \<B> ; \<Gamma> \<turnstile> v \<Rightarrow> \<lbrace> z : b | c \<rbrace>\<close> using infer_e_elims check_letI * by fast
- thus \<open>\<Theta> ; \<B> ; \<Gamma> \<turnstile> \<lbrace> z : b | c \<rbrace> \<lesssim> \<lbrace> z : b | c \<rbrace>\<close> using subtype_reflI infer_v_wf by metis
- show \<open>atom z \<sharp> (x1, v)\<close> using check_letI fresh_Pair by auto
- show \<open>\<Theta> ; \<Phi> ; \<B> ; (x1, b, c[z::=V_var x1]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile> s1 \<Leftarrow> \<tau>\<close> using check_letI subst_defs by auto
- show "(x1, b, c[z::=V_var x1]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma> = GNil @ (x1, b, c[z::=V_var x1]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>" by auto
- qed
-
- hence "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile> s1[x1::=v]\<^sub>s\<^sub>v \<Leftarrow> \<tau>" using check_letI by auto
- moreover have " s1[x1::=v]\<^sub>s\<^sub>v = s[x::=v]\<^sub>s\<^sub>v"
- by (metis (full_types) check_letI fresh_GNil infer_e_elims(7) s_branch_s_branch_list.distinct s_branch_s_branch_list.eq_iff(2)
- subst_s_abs_lst wfG_x_fresh_in_v_simple)
-
- ultimately show ?case using check_letI by simp
-next
- case (check_let2I x1 \<Theta> \<Phi> \<B> \<Gamma> \<Delta> t s1 \<tau> s2 )
-
- hence s1eq:"s1 = AS_val v" by auto
- let ?G = "(x1, b_of t, c_of t x1) #\<^sub>\<Gamma> \<Gamma>"
- obtain z::x where *:"atom z \<sharp> (x1 , v,t)" using obtain_fresh_z by metis
- hence teq:"t = \<lbrace> z : b_of t | c_of t z \<rbrace>" using b_of_c_of_eq by auto
- have "\<Theta> ; \<Phi> ; \<B> ; ?G[x1::=v]\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x1::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> s2[x1::=v]\<^sub>s\<^sub>v \<Leftarrow> \<tau>[x1::=v]\<^sub>\<tau>\<^sub>v"
- proof(rule subst_check_check_s(1))
- obtain t' where "\<Theta> ; \<B> ; \<Gamma> \<turnstile> v \<Rightarrow> t' \<and> \<Theta> ; \<B> ; \<Gamma> \<turnstile> t' \<lesssim> t" using check_s_elims(1) check_let2I(10) s1eq by auto
- thus **:\<open> \<Theta> ; \<B> ; \<Gamma> \<turnstile> v \<Leftarrow> \<lbrace> z : b_of t | c_of t z \<rbrace>\<close> using check_v.intros teq by auto
- show "atom z \<sharp> (x1, v)" using * by auto
- show \<open> \<Theta> ; \<Phi> ; \<B> ; (x1, b_of t, c_of t x1) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile> s2 \<Leftarrow> \<tau>\<close> using check_let2I by auto
- show "(x1, b_of t , c_of t x1) #\<^sub>\<Gamma> \<Gamma> = GNil @ (x1, b_of t, (c_of t z)[z::=V_var x1]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>" using append_g.simps c_of_switch * fresh_prodN by metis
- qed
-
- hence "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile> s2[x1::=v]\<^sub>s\<^sub>v \<Leftarrow> \<tau>" using check_let2I by auto
- moreover have " s2[x1::=v]\<^sub>s\<^sub>v = s[x::=v]\<^sub>s\<^sub>v" using
- check_let2I fresh_GNil check_s_elims s_branch_s_branch_list.distinct s_branch_s_branch_list.eq_iff
- subst_s_abs_lst wfG_x_fresh_in_v_simple
- proof -
- have "AS_let2 x t (AS_val v) s = AS_let2 x1 t s1 s2"
- by (metis check_let2I.prems(3) s_branch_s_branch_list.distinct s_branch_s_branch_list.eq_iff(3)) (* 62 ms *)
- then show ?thesis
- by (metis (no_types) check_let2I check_let2I.prems(2) check_s_elims(1) fresh_GNil s_branch_s_branch_list.eq_iff(3) subst_s_abs_lst wfG_x_fresh_in_v_simple) (* 375 ms *)
- qed
-
- ultimately show ?case using check_let2I by simp
-qed(auto+)
-
-lemma preservation_let_val:
- assumes "\<Theta>; \<Phi>; \<Delta> \<turnstile> \<langle> \<delta> , AS_let x (AE_val v) s \<rangle> \<Leftarrow> \<tau> \<or> \<Theta>; \<Phi>; \<Delta> \<turnstile> \<langle> \<delta> , AS_let2 x t (AS_val v) s \<rangle> \<Leftarrow> \<tau>" (is "?A \<or> ?B")
- shows " \<exists>\<Delta>'. \<Theta>; \<Phi>; \<Delta>' \<turnstile> \<langle> \<delta> , s[x::=v]\<^sub>s\<^sub>v \<rangle> \<Leftarrow> \<tau> \<and> \<Delta> \<sqsubseteq> \<Delta>'"
-proof -
- have tt: "\<Theta> \<turnstile> \<delta> \<sim> \<Delta>" and fd: "(\<forall>fd\<in>set \<Phi>. check_fundef \<Theta> \<Phi> fd)"
- using assms by blast+
-
- have "?A \<or> ?B" using assms by auto
- then have "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> s[x::=v]\<^sub>s\<^sub>v \<Leftarrow> \<tau>"
- proof
- assume "\<Theta>; \<Phi>; \<Delta> \<turnstile> \<langle> \<delta> , AS_let x (AE_val v) s \<rangle> \<Leftarrow> \<tau>"
- hence * : " \<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let x (AE_val v) s \<Leftarrow> \<tau>" by blast
- thus ?thesis using check_let_val by simp
- next
- assume "\<Theta>; \<Phi>; \<Delta> \<turnstile> \<langle> \<delta> , AS_let2 x t (AS_val v) s \<rangle> \<Leftarrow> \<tau>"
- hence * : "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let2 x t (AS_val v) s \<Leftarrow> \<tau>" by blast
- thus ?thesis using check_let_val by simp
- qed
-
- thus ?thesis using tt config_typeI fd
- order_refl by metis
-qed
-
-lemma check_s_fst_snd:
- assumes "fst_snd = AE_fst \<and> v=v1 \<or> fst_snd = AE_snd \<and> v=v2"
- and "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let x (fst_snd (V_pair v1 v2)) s' \<Leftarrow> \<tau>"
- shows "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let x ( AE_val v) s' \<Leftarrow> \<tau>"
-proof -
- have 1: \<open> \<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let x (fst_snd (V_pair v1 v2)) s' \<Leftarrow> \<tau>\<close> using assms by auto
-
- then obtain t1 where 2:"\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> (fst_snd (V_pair v1 v2)) \<Rightarrow> t1" using check_s_elims by auto
-
- show ?thesis using subtype_let 1 2 assms
- by (meson infer_e_fst_pair infer_e_snd_pair)
-qed
-
-lemma preservation_fst_snd:
- assumes "\<Theta>; \<Phi>; \<Delta> \<turnstile> \<langle> \<delta> , LET x = (fst_snd (V_pair v1 v2)) IN s' \<rangle> \<Leftarrow> \<tau>" and
- "fst_snd = AE_fst \<and> v=v1 \<or> fst_snd = AE_snd \<and> v=v2"
- shows "\<exists>\<Delta>'. \<Theta>; \<Phi>; \<Delta> \<turnstile> \<langle> \<delta> , LET x = (AE_val v) IN s' \<rangle> \<Leftarrow> \<tau> \<and> \<Delta> \<sqsubseteq> \<Delta>'"
-proof -
- have tt: "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let x (fst_snd (V_pair v1 v2)) s' \<Leftarrow> \<tau> \<and> \<Theta> \<turnstile> \<delta> \<sim> \<Delta>" using assms by blast
- hence t2: "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let x (fst_snd (V_pair v1 v2)) s' \<Leftarrow> \<tau>" by auto
-
- moreover have "\<forall>fd\<in>set \<Phi>. check_fundef \<Theta> \<Phi> fd" using assms config_type_elims by auto
- ultimately show ?thesis using config_typeI order_refl tt assms check_s_fst_snd by metis
-qed
-
-inductive_cases check_branch_s_elims2[elim!]:
- "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta>; tid ; cons ; const ; v \<turnstile> cs \<Leftarrow> \<tau>"
-
-lemmas freshers = freshers atom_dom.simps toSet.simps fresh_def x_not_in_b_set
-declare freshers [simp]
-
-lemma subtype_eq_if:
- fixes t::\<tau> and va::v
- assumes "\<Theta> ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b_of t | c_of t z \<rbrace>" and "\<Theta> ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b_of t | c IMP c_of t z \<rbrace> "
- shows "\<Theta> ; \<B> ; \<Gamma> \<turnstile> \<lbrace> z : b_of t | c_of t z \<rbrace> \<lesssim> \<lbrace> z : b_of t | c IMP c_of t z \<rbrace> "
-proof -
- obtain x::x where xf:"atom x \<sharp> ((\<Theta>, \<B>, \<Gamma>, z, c_of t z, z, c IMP c_of t z ),c)" using obtain_fresh by metis
-
- moreover have "\<Theta> ; \<B> ; (x, b_of t, (c_of t z)[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<Turnstile> (c IMP c_of t z )[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v "
- unfolding subst_cv.simps
- proof(rule valid_eq_imp)
-
- have "\<Theta> ; \<B> ; (x, b_of t, (c_of t z)[z::=[ x ]\<^sup>v]\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f (c IMP (c_of t z))[z::=[ x ]\<^sup>v]\<^sub>v "
- apply(rule wfT_wfC_cons)
- apply(simp add: assms, simp add: assms, unfold fresh_prodN )
- using xf fresh_prodN by metis+
- thus "\<Theta> ; \<B> ; (x, b_of t, (c_of t z)[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f c[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v IMP (c_of t z)[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v "
- using subst_cv.simps subst_defs by auto
- qed
-
- ultimately show ?thesis using subtype_baseI assms fresh_Pair subst_defs by metis
-qed
-
-lemma subtype_eq_if_\<tau>:
- fixes t::\<tau> and va::v
- assumes "\<Theta> ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f t" and "\<Theta> ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b_of t | c IMP c_of t z \<rbrace> " and "atom z \<sharp> t"
- shows "\<Theta> ; \<B> ; \<Gamma> \<turnstile> t \<lesssim> \<lbrace> z : b_of t | c IMP c_of t z \<rbrace> "
-proof -
- have "t = \<lbrace> z : b_of t | c_of t z \<rbrace>" using b_of_c_of_eq assms by auto
- thus ?thesis using subtype_eq_if assms c_of.simps b_of.simps by metis
-qed
-
-lemma valid_conj:
- assumes "\<Theta> ; \<B> ; \<Gamma> \<Turnstile> c1" and "\<Theta> ; \<B> ; \<Gamma> \<Turnstile> c2"
- shows "\<Theta> ; \<B> ; \<Gamma> \<Turnstile> c1 AND c2"
-proof
- show \<open> \<Theta> ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f c1 AND c2 \<close> using valid.simps wfC_conjI assms by auto
- show \<open>\<forall>i. \<Theta> ; \<Gamma> \<turnstile> i \<and> i \<Turnstile> \<Gamma> \<longrightarrow> i \<Turnstile> c1 AND c2 \<close>
- proof(rule+)
- fix i
- assume *:"\<Theta> ; \<Gamma> \<turnstile> i \<and> i \<Turnstile> \<Gamma> "
- thus "i \<lbrakk> c1 \<rbrakk> ~ True" using assms valid.simps
- using is_satis.cases by blast
- show "i \<lbrakk> c2 \<rbrakk> ~ True" using assms valid.simps
- using is_satis.cases * by blast
- qed
-qed
-
-subsection \<open>Other Statements\<close>
-
-lemma check_if:
- fixes s'::s and cs::branch_s and css::branch_list and v::v
- shows "\<Theta>; \<Phi>; B; G; \<Delta> \<turnstile> s' \<Leftarrow> \<tau> \<Longrightarrow> s' = IF (V_lit ll) THEN s1 ELSE s2 \<Longrightarrow>
- \<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f \<tau> \<Longrightarrow> G = GNil \<Longrightarrow> B = {||} \<Longrightarrow> ll = L_true \<and> s = s1 \<or> ll = L_false \<and> s = s2 \<Longrightarrow>
- \<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> s \<Leftarrow> \<tau>" and
- "check_branch_s \<Theta> \<Phi> {||} GNil \<Delta> tid dc const v cs \<tau> \<Longrightarrow> True" and
- "check_branch_list \<Theta> \<Phi> {||} \<Gamma> \<Delta> tid dclist v css \<tau> \<Longrightarrow> True"
-proof(nominal_induct \<tau> and \<tau> and \<tau> rule: check_s_check_branch_s_check_branch_list.strong_induct)
- case (check_ifI z \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v s1 s2 \<tau>)
- obtain z' where teq: "\<tau> = \<lbrace> z' : b_of \<tau> | c_of \<tau> z' \<rbrace> \<and> atom z' \<sharp> (z,\<tau>)" using obtain_fresh_z_c_of by metis
- hence ceq: "(c_of \<tau> z')[z'::=[ z ]\<^sup>v]\<^sub>c\<^sub>v = (c_of \<tau> z)" using c_of_switch fresh_Pair by metis
- have zf: " atom z \<sharp> c_of \<tau> z'"
- using c_of_fresh check_ifI teq fresh_Pair fresh_at_base by metis
- hence 1:"\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> s \<Leftarrow> \<lbrace> z : b_of \<tau> | CE_val (V_lit ll) == CE_val (V_lit ll) IMP c_of \<tau> z \<rbrace>" using check_ifI by auto
- moreover have 2:"\<Theta> ; {||} ; GNil \<turnstile> (\<lbrace> z : b_of \<tau> | CE_val (V_lit ll) == CE_val (V_lit ll) IMP c_of \<tau> z \<rbrace>) \<lesssim> \<tau>"
- proof -
- have "\<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f (\<lbrace> z : b_of \<tau> | CE_val (V_lit ll ) == CE_val (V_lit ll) IMP c_of \<tau> z \<rbrace>)" using check_ifI check_s_wf by auto
- moreover have "\<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f \<tau>" using check_s_wf check_ifI by auto
- ultimately show ?thesis using subtype_if_simp[of \<Theta> " {||}" z "b_of \<tau>" ll "c_of \<tau> z'" z'] using teq ceq zf subst_defs by metis
- qed
- ultimately show ?case using check_s_supertype(1) check_ifI by metis
-qed(auto+)
-
-lemma preservation_if:
- assumes "\<Theta>; \<Phi>; \<Delta> \<turnstile> \<langle> \<delta> , IF (V_lit ll) THEN s1 ELSE s2 \<rangle> \<Leftarrow> \<tau>" and
- "ll = L_true \<and> s = s1 \<or> ll = L_false \<and> s = s2"
- shows "\<Theta>; \<Phi>; \<Delta> \<turnstile> \<langle>\<delta>, s\<rangle> \<Leftarrow> \<tau> \<and> setD \<Delta> \<subseteq> setD \<Delta>"
-proof -
- have *: "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_if (V_lit ll) s1 s2 \<Leftarrow> \<tau> \<and> (\<forall>fd\<in>set \<Phi>. check_fundef \<Theta> \<Phi> fd)"
- using assms config_type_elims by metis
- hence "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> s \<Leftarrow> \<tau>" using check_s_wf check_if assms by metis
- hence "\<Theta>; \<Phi>; \<Delta> \<turnstile> \<langle>\<delta>, s\<rangle> \<Leftarrow> \<tau> \<and> setD \<Delta> \<subseteq> setD \<Delta>" using config_typeI *
- using assms(1) by blast
- thus ?thesis by blast
-qed
-
-lemma wfT_conj:
- assumes "\<Theta> ; \<B> ; GNil \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b | c1 \<rbrace>" and "\<Theta> ; \<B> ; GNil \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b | c2 \<rbrace>"
- shows "\<Theta> ; \<B> ; GNil \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b | c1 AND c2\<rbrace>"
-proof
- show \<open>atom z \<sharp> (\<Theta>, \<B>, GNil)\<close>
- apply(unfold fresh_prodN, intro conjI)
- using wfTh_fresh wfT_wf assms apply metis
- using fresh_GNil x_not_in_b_set fresh_def by metis+
- show \<open> \<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f b \<close> using wfT_elims assms by metis
- have "\<Theta> ; \<B> ; (z, b, TRUE) #\<^sub>\<Gamma> GNil \<turnstile>\<^sub>w\<^sub>f c1" using wfT_wfC fresh_GNil assms by auto
- moreover have "\<Theta> ; \<B> ; (z, b, TRUE) #\<^sub>\<Gamma> GNil \<turnstile>\<^sub>w\<^sub>f c2" using wfT_wfC fresh_GNil assms by auto
- ultimately show \<open> \<Theta> ; \<B> ; (z, b, TRUE) #\<^sub>\<Gamma> GNil \<turnstile>\<^sub>w\<^sub>f c1 AND c2 \<close> using wfC_conjI by auto
-qed
-
-lemma subtype_conj:
- assumes "\<Theta> ; \<B> ; GNil \<turnstile> t \<lesssim> \<lbrace> z : b | c1 \<rbrace>" and "\<Theta> ; \<B> ; GNil \<turnstile> t \<lesssim> \<lbrace> z : b | c2 \<rbrace>"
- shows "\<Theta> ; \<B> ; GNil \<turnstile> \<lbrace> z : b | c_of t z \<rbrace> \<lesssim> \<lbrace> z : b | c1 AND c2 \<rbrace>"
-proof -
- have beq: "b_of t = b" using subtype_eq_base2 b_of.simps assms by metis
- obtain x::x where x:\<open>atom x \<sharp> (\<Theta>, \<B>, GNil, z, c_of t z, z, c1 AND c2 )\<close> using obtain_fresh by metis
- thus ?thesis proof
- have "atom z \<sharp> t" using subtype_wf wfT_supp fresh_def x_not_in_b_set atom_dom.simps toSet.simps assms dom.simps by fastforce
- hence t:"t = \<lbrace> z : b_of t | c_of t z \<rbrace>" using b_of_c_of_eq by auto
- thus \<open> \<Theta> ; \<B> ; GNil \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b | c_of t z \<rbrace> \<close> using subtype_wf beq assms by auto
-
- show \<open>\<Theta> ; \<B> ; (x, b, (c_of t z)[z::=[ x ]\<^sup>v]\<^sub>v) #\<^sub>\<Gamma> GNil \<Turnstile> (c1 AND c2 )[z::=[ x ]\<^sup>v]\<^sub>v \<close>
- proof -
- have \<open>\<Theta> ; \<B> ; (x, b, (c_of t z)[z::=[ x ]\<^sup>v]\<^sub>v) #\<^sub>\<Gamma> GNil \<Turnstile> c1[z::=[ x ]\<^sup>v]\<^sub>v \<close>
- proof(rule subtype_valid)
- show \<open>\<Theta> ; \<B> ; GNil \<turnstile> t \<lesssim> \<lbrace> z : b | c1 \<rbrace>\<close> using assms by auto
- show \<open>atom x \<sharp> GNil\<close> using fresh_GNil by auto
- show \<open>t = \<lbrace> z : b | c_of t z \<rbrace>\<close> using t beq by auto
- show \<open>\<lbrace> z : b | c1 \<rbrace> = \<lbrace> z : b | c1 \<rbrace>\<close> by auto
- qed
- moreover have \<open>\<Theta> ; \<B> ; (x, b, (c_of t z)[z::=[ x ]\<^sup>v]\<^sub>v) #\<^sub>\<Gamma> GNil \<Turnstile> c2[z::=[ x ]\<^sup>v]\<^sub>v \<close>
- proof(rule subtype_valid)
- show \<open>\<Theta> ; \<B> ; GNil \<turnstile> t \<lesssim> \<lbrace> z : b | c2 \<rbrace>\<close> using assms by auto
- show \<open>atom x \<sharp> GNil\<close> using fresh_GNil by auto
- show \<open>t = \<lbrace> z : b | c_of t z \<rbrace>\<close> using t beq by auto
- show \<open>\<lbrace> z : b | c2 \<rbrace> = \<lbrace> z : b | c2 \<rbrace>\<close> by auto
- qed
- ultimately show ?thesis unfolding subst_cv.simps subst_v_c_def using valid_conj by metis
- qed
- thus \<open> \<Theta> ; \<B> ; GNil \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b | c1 AND c2 \<rbrace> \<close> using subtype_wf wfT_conj assms by auto
- qed
-qed
-
-lemma infer_v_conj:
- assumes "\<Theta> ; \<B> ; GNil \<turnstile> v \<Leftarrow> \<lbrace> z : b | c1 \<rbrace>" and "\<Theta> ; \<B> ; GNil \<turnstile> v \<Leftarrow> \<lbrace> z : b | c2 \<rbrace>"
- shows "\<Theta> ; \<B> ; GNil \<turnstile> v \<Leftarrow> \<lbrace> z : b | c1 AND c2 \<rbrace>"
-proof -
- obtain t1 where t1: "\<Theta> ; \<B> ; GNil \<turnstile> v \<Rightarrow> t1 \<and> \<Theta> ; \<B> ; GNil \<turnstile> t1 \<lesssim> \<lbrace> z : b | c1 \<rbrace>"
- using assms check_v_elims by metis
- obtain t2 where t2: "\<Theta> ; \<B> ; GNil \<turnstile> v \<Rightarrow> t2 \<and> \<Theta> ; \<B> ; GNil \<turnstile> t2 \<lesssim> \<lbrace> z : b | c2 \<rbrace>"
- using assms check_v_elims by metis
- have teq: "t1 = \<lbrace> z : b | c_of t1 z \<rbrace>" using subtype_eq_base2 b_of.simps
- by (metis (full_types) b_of_c_of_eq fresh_GNil infer_v_t_wf t1 wfT_x_fresh)
- have "t1 = t2" using infer_v_uniqueness t1 t2 by auto
- hence " \<Theta> ; \<B> ; GNil \<turnstile> \<lbrace> z : b | c_of t1 z \<rbrace> \<lesssim> \<lbrace> z : b | c1 AND c2 \<rbrace>" using subtype_conj t1 t2 by simp
- hence " \<Theta> ; \<B> ; GNil \<turnstile> t1 \<lesssim> \<lbrace> z : b | c1 AND c2 \<rbrace>" using teq by auto
- thus ?thesis using t1 using check_v.intros by auto
-qed
-
-lemma wfG_conj:
- fixes c1::c
- assumes "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f (x, b, c1 AND c2) #\<^sub>\<Gamma> \<Gamma>"
- shows "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f (x, b, c1) #\<^sub>\<Gamma> \<Gamma>"
-proof(cases "c1 \<in> {TRUE, FALSE}")
- case True
- then show ?thesis using assms wfG_cons2I wfG_elims wfX_wfY by metis
-next
- case False
- then show ?thesis using assms wfG_cons1I wfG_elims wfX_wfY wfC_elims
- by (metis wfG_elim2)
-qed
-
-lemma check_match:
- fixes s'::s and s::s and css::branch_list and cs::branch_s
- shows "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile> s \<Leftarrow> \<tau> \<Longrightarrow> True" and
- "\<Theta> ; \<Phi> ; B ; G ; \<Delta> ; tid ; dc ; const ; vcons \<turnstile> cs \<Leftarrow> \<tau> \<Longrightarrow>
- vcons = V_cons tid dc v \<Longrightarrow> B = {||} \<Longrightarrow> G = GNil \<Longrightarrow> cs = (dc x' \<Rightarrow> s') \<Longrightarrow>
- \<Theta> ; {||} ; GNil \<turnstile> v \<Leftarrow> const \<Longrightarrow>
- \<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> s'[x'::=v]\<^sub>s\<^sub>v \<Leftarrow> \<tau>" and
- "\<Theta> ; \<Phi> ; B ; G ; \<Delta> ; tid ; dclist ; vcons \<turnstile> css \<Leftarrow> \<tau> \<Longrightarrow> distinct (map fst dclist) \<Longrightarrow>
- vcons = V_cons tid dc v \<Longrightarrow> B = {||} \<Longrightarrow> (dc, const) \<in> set dclist \<Longrightarrow> G = GNil \<Longrightarrow>
- Some (AS_branch dc x' s') = lookup_branch dc css \<Longrightarrow> \<Theta> ; {||} ; GNil \<turnstile> v \<Leftarrow> const \<Longrightarrow>
- \<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> s'[x'::=v]\<^sub>s\<^sub>v \<Leftarrow> \<tau>"
-proof(nominal_induct \<tau> and \<tau> and \<tau> avoiding: x' v rule: check_s_check_branch_s_check_branch_list.strong_induct)
- case (check_branch_list_consI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid consa consta va cs \<tau> dclist cssa)
-
- then obtain xa and sa where cseq:"cs = AS_branch consa xa sa" using check_branch_s_elims2[OF check_branch_list_consI(1)] by metis
-
- show ?case proof(cases "dc = consa")
- case True
- hence "cs = AS_branch consa x' s'" using check_branch_list_consI cseq
- by (metis lookup_branch.simps(2) option.inject)
- moreover have "const = consta" using check_branch_list_consI distinct.simps
- by (metis True dclist_distinct_unique list.set_intros(1))
- moreover have "va = V_cons tid consa v" using check_branch_list_consI True by auto
- ultimately show ?thesis using check_branch_list_consI by auto
- next
- case False
- hence "Some (AS_branch dc x' s') = lookup_branch dc cssa" using lookup_branch.simps(2) check_branch_list_consI(10) cseq by auto
- moreover have "(dc, const) \<in> set dclist" using check_branch_list_consI False by simp
- ultimately show ?thesis using check_branch_list_consI by auto
- qed
-
-next
- case (check_branch_list_finalI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid cons const va cs \<tau>)
- hence " cs = AS_branch cons x' s'" using lookup.simps check_branch_list_finalI lookup_branch.simps option.inject
- by (metis map_of.simps(1) map_of_Cons_code(2) option.distinct(1) s_branch_s_branch_list.exhaust(2) weak_map_of_SomeI)
- then show ?case using check_branch_list_finalI by auto
-next
- case (check_branch_s_branchI \<Theta> \<B> \<Gamma> \<Delta> \<tau> const x \<Phi> tid cons va s)
-
- text \<open>Supporting facts here to make the main body of the proof concise\<close>
- have xf:"atom x \<sharp> \<tau>" proof -
- have "supp \<tau> \<subseteq> supp \<B>" using wf_supp(4) check_branch_s_branchI atom_dom.simps toSet.simps dom.simps by fastforce
- thus ?thesis using fresh_def x_not_in_b_set by blast
- qed
-
- hence "\<tau>[x::=v]\<^sub>\<tau>\<^sub>v = \<tau>" using forget_subst_v subst_v_\<tau>_def by auto
- have "\<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v = \<Delta>" using forget_subst_dv wfD_x_fresh fresh_GNil check_branch_s_branchI by metis
-
- have "supp v = {}" using check_branch_s_branchI check_v_wf wfV_supp_nil by metis
- hence "supp va = {}" using \<open> va = V_cons tid cons v\<close> v.supp pure_supp by auto
-
- let ?G = "(x, b_of const, [ va ]\<^sup>c\<^sup>e == [ V_cons tid cons [ x ]\<^sup>v ]\<^sup>c\<^sup>e AND c_of const x ) #\<^sub>\<Gamma> \<Gamma>"
- obtain z::x where z: "const = \<lbrace> z : b_of const | c_of const z \<rbrace> \<and> atom z \<sharp> (x', v,x,const,va)"
- using obtain_fresh_z_c_of by metis
-
- have vt: \<open>\<Theta> ; \<B> ; GNil \<turnstile> v \<Leftarrow> \<lbrace> z : b_of const | [ va ]\<^sup>c\<^sup>e == [ V_cons tid cons [ z ]\<^sup>v ]\<^sup>c\<^sup>e AND c_of const z \<rbrace>\<close>
- proof(rule infer_v_conj)
- obtain t' where t: "\<Theta> ; \<B> ; GNil \<turnstile> v \<Rightarrow> t' \<and> \<Theta> ; \<B> ; GNil \<turnstile> t' \<lesssim> const"
- using check_v_elims check_branch_s_branchI by metis
- show "\<Theta> ; \<B> ; GNil \<turnstile> v \<Leftarrow> \<lbrace> z : b_of const | [ va ]\<^sup>c\<^sup>e == [ V_cons tid cons [ z ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrace>"
- proof(rule check_v_top)
-
- show "\<Theta> ; \<B> ; GNil \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b_of const | [ va ]\<^sup>c\<^sup>e == [ V_cons tid cons [ z ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrace> "
-
- proof(rule wfG_wfT)
- show \<open> \<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f (x, b_of const, ([ va ]\<^sup>c\<^sup>e == [ V_cons tid cons [ z ]\<^sup>v ]\<^sup>c\<^sup>e )[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> GNil \<close>
- proof -
- have 1: "va[z::=[ x ]\<^sup>v]\<^sub>v\<^sub>v = va" using forget_subst_v subst_v_v_def z fresh_prodN by metis
- moreover have 2: "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f (x, b_of const, [ va ]\<^sup>c\<^sup>e == [ V_cons tid cons [ x ]\<^sup>v ]\<^sup>c\<^sup>e AND c_of const x ) #\<^sub>\<Gamma> GNil "
- using check_branch_s_branchI(17)[THEN check_s_wf] \<open>\<Gamma> = GNil\<close> by auto
- moreover hence "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f (x, b_of const, [ va ]\<^sup>c\<^sup>e == [ V_cons tid cons [ x ]\<^sup>v ]\<^sup>c\<^sup>e ) #\<^sub>\<Gamma> GNil"
- using wfG_conj by metis
- ultimately show ?thesis
- unfolding subst_cv.simps subst_cev.simps subst_vv.simps by auto
- qed
- show \<open>atom x \<sharp> ([ va ]\<^sup>c\<^sup>e == [ V_cons tid cons [ z ]\<^sup>v ]\<^sup>c\<^sup>e )\<close> unfolding c.fresh ce.fresh v.fresh
- apply(intro conjI )
- using check_branch_s_branchI fresh_at_base z freshers apply simp
- using check_branch_s_branchI fresh_at_base z freshers apply simp
- using pure_supp apply force
- using z fresh_x_neq fresh_prod5 by metis
- qed
- show \<open>[ va ]\<^sup>c\<^sup>e = [ V_cons tid cons [ z ]\<^sup>v ]\<^sup>c\<^sup>e[z::=v]\<^sub>c\<^sub>e\<^sub>v\<close>
- using \<open> va = V_cons tid cons v\<close> subst_cev.simps subst_vv.simps by auto
- show \<open> \<Theta> ; \<B> ; GNil \<turnstile> v \<Leftarrow> const \<close> using check_branch_s_branchI by auto
- show "supp [ va ]\<^sup>c\<^sup>e \<subseteq> supp \<B>" using \<open>supp va = {}\<close> ce.supp by simp
- qed
- show "\<Theta> ; \<B> ; GNil \<turnstile> v \<Leftarrow> \<lbrace> z : b_of const | c_of const z \<rbrace>"
- using check_branch_s_branchI z by auto
- qed
-
- text \<open>Main body of proof for this case\<close>
- have "\<Theta> ; \<Phi> ; \<B> ; (?G)[x::=v]\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> s[x::=v]\<^sub>s\<^sub>v \<Leftarrow> \<tau>[x::=v]\<^sub>\<tau>\<^sub>v"
- proof(rule subst_check_check_s)
- show \<open>\<Theta> ; \<B> ; GNil \<turnstile> v \<Leftarrow> \<lbrace> z : b_of const | [ va ]\<^sup>c\<^sup>e == [ V_cons tid cons [ z ]\<^sup>v ]\<^sup>c\<^sup>e AND c_of const z \<rbrace>\<close> using vt by auto
- show \<open>atom z \<sharp> (x, v)\<close> using z fresh_prodN by auto
- show \<open> \<Theta> ; \<Phi> ; \<B> ; ?G ; \<Delta> \<turnstile> s \<Leftarrow> \<tau>\<close>
- using check_branch_s_branchI by auto
-
- show \<open> ?G = GNil @ (x, b_of const, ([ va ]\<^sup>c\<^sup>e == [ V_cons tid cons [ z ]\<^sup>v ]\<^sup>c\<^sup>e AND c_of const z)[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> GNil\<close>
- proof -
- have "va[z::=[ x ]\<^sup>v]\<^sub>v\<^sub>v = va" using forget_subst_v subst_v_v_def z fresh_prodN by metis
- moreover have "(c_of const z)[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v = c_of const x"
- using c_of_switch[of z const x] z fresh_prodN by metis
- ultimately show ?thesis
- unfolding subst_cv.simps subst_cev.simps subst_vv.simps append_g.simps
- using c_of_switch[of z const x] z fresh_prodN z fresh_prodN check_branch_s_branchI by argo
- qed
- qed
- moreover have "s[x::=v]\<^sub>s\<^sub>v = s'[x'::=v]\<^sub>s\<^sub>v"
- using check_branch_s_branchI subst_v_flip_eq_two subst_v_s_def s_branch_s_branch_list.eq_iff by metis
- ultimately show ?case using check_branch_s_branchI \<open>\<tau>[x::=v]\<^sub>\<tau>\<^sub>v = \<tau>\<close> \<open>\<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v = \<Delta>\<close> by auto
-qed(auto+)
-
-text \<open>Lemmas for while reduction. Making these separate lemmas allows flexibility in wiring them into the main proof and robustness
-if we change it\<close>
-
-lemma check_unit:
- fixes \<tau>::\<tau> and \<Phi>::\<Phi> and \<Delta>::\<Delta> and G::\<Gamma>
- assumes "\<Theta> ; {||} ; GNil \<turnstile> \<lbrace> z : B_unit | TRUE \<rbrace> \<lesssim> \<tau>'" and "\<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f \<Delta>" and "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>" and "\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f G"
- shows \<open> \<Theta> ; \<Phi> ; {||} ; G ; \<Delta> \<turnstile> [[ L_unit ]\<^sup>v]\<^sup>s \<Leftarrow> \<tau>'\<close>
-proof -
- have *:"\<Theta> ; {||} ; GNil \<turnstile> [L_unit]\<^sup>v \<Rightarrow> \<lbrace> z : B_unit | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_unit ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrace>"
- using infer_l.intros(4) infer_v_litI fresh_GNil assms wfX_wfY by (meson subtype_g_wf)
- moreover have "\<Theta> ; {||} ; GNil \<turnstile> \<lbrace> z : B_unit | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_unit ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrace> \<lesssim> \<tau>'"
- using subtype_top subtype_trans * infer_v_wf
- by (meson assms(1))
- ultimately show ?thesis (*apply(rule check_valI, auto simp add: assms,rule * )*)
- using subtype_top subtype_trans fresh_GNil assms check_valI assms check_s_g_weakening assms toSet.simps
- by (metis bot.extremum infer_v_g_weakening subtype_weakening wfD_wf)
-qed
-
-lemma preservation_var:
- shows "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> VAR u : \<tau>' = v IN s \<Leftarrow> \<tau> \<Longrightarrow> \<Theta> \<turnstile> \<delta> \<sim> \<Delta> \<Longrightarrow> atom u \<sharp> \<delta> \<Longrightarrow> atom u \<sharp> \<Delta> \<Longrightarrow>
- \<Theta>; \<Phi>; {||}; GNil; (u,\<tau>')#\<^sub>\<Delta>\<Delta> \<turnstile> s \<Leftarrow> \<tau> \<and> \<Theta> \<turnstile> (u,v)#\<delta> \<sim> (u,\<tau>')#\<^sub>\<Delta>\<Delta>"
- and
- "check_branch_s \<Theta> \<Phi> {||} GNil \<Delta> tid dc const v cs \<tau> \<Longrightarrow> True" and
- "check_branch_list \<Theta> \<Phi> {||} \<Gamma> \<Delta> tid dclist v css \<tau> \<Longrightarrow> True"
-proof(nominal_induct "{||}::bv fset" GNil \<Delta> "VAR u : \<tau>' = v IN s" \<tau> and \<tau> and \<tau> rule: check_s_check_branch_s_check_branch_list.strong_induct)
- case (check_varI u' \<Theta> \<Phi> \<Delta> \<tau> s')
- hence "\<Theta>; \<Phi>; {||}; GNil; (u, \<tau>') #\<^sub>\<Delta> \<Delta> \<turnstile> s \<Leftarrow> \<tau>" using check_s_abs_u check_v_wf by metis
-
- moreover have "\<Theta> \<turnstile> ((u,v)#\<delta>) \<sim> ((u,\<tau>')#\<^sub>\<Delta>\<Delta>)" proof
- show \<open>\<Theta> \<turnstile> \<delta> \<sim> \<Delta> \<close> using check_varI by auto
- show \<open>\<Theta> ; {||} ; GNil \<turnstile> v \<Leftarrow> \<tau>'\<close> using check_varI by auto
- show \<open>u \<notin> fst ` set \<delta>\<close> using check_varI fresh_d_fst_d by auto
- qed
-
- ultimately show ?case by simp
-qed(auto+)
-
-lemma check_while:
- shows "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> WHILE s1 DO { s2 } \<Leftarrow> \<tau> \<Longrightarrow> atom x \<sharp> (s1, s2) \<Longrightarrow> atom z' \<sharp> x \<Longrightarrow>
- \<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> LET x : (\<lbrace> z' : B_bool | TRUE \<rbrace>) = s1 IN (IF (V_var x) THEN (s2 ;; (WHILE s1 DO {s2}))
- ELSE ([ V_lit L_unit]\<^sup>s)) \<Leftarrow> \<tau>" and
- "check_branch_s \<Theta> \<Phi> {||} GNil \<Delta> tid dc const v cs \<tau> \<Longrightarrow> True" and
- "check_branch_list \<Theta> \<Phi> {||} \<Gamma> \<Delta> tid dclist v css \<tau> \<Longrightarrow> True"
-proof(nominal_induct "{||}::bv fset" GNil \<Delta> "AS_while s1 s2" \<tau> and \<tau> and \<tau> avoiding: s1 s2 x z' rule: check_s_check_branch_s_check_branch_list.strong_induct)
- case (check_whileI \<Theta> \<Phi> \<Delta> s1 z s2 \<tau>')
- have teq:"\<lbrace> z' : B_bool | TRUE \<rbrace> = \<lbrace> z : B_bool | TRUE \<rbrace>" using \<tau>.eq_iff by auto
-
- show ?case proof
- have " atom x \<sharp> \<tau>' " using wfT_nil_supp fresh_def subtype_wfT check_whileI(5) by fast
- moreover have "atom x \<sharp> \<lbrace> z' : B_bool | TRUE \<rbrace> " using \<tau>.fresh c.fresh b.fresh by metis
- ultimately show \<open>atom x \<sharp> (\<Theta>, \<Phi>, {||}, GNil, \<Delta>, \<lbrace> z' : B_bool | TRUE \<rbrace>, s1, \<tau>')\<close>
- apply(unfold fresh_prodN)
- using check_whileI wb_x_fresh check_s_wf wfD_x_fresh fresh_empty_fset fresh_GNil fresh_Pair \<open>atom x \<sharp> \<tau>'\<close> by metis
-
- show \<open> \<Theta> ; \<Phi> ; {||} ; GNil ; \<Delta> \<turnstile> s1 \<Leftarrow> \<lbrace> z' : B_bool | TRUE \<rbrace>\<close> using check_whileI teq by metis
-
- let ?G = "(x, b_of \<lbrace> z' : B_bool | TRUE \<rbrace>, c_of \<lbrace> z' : B_bool | TRUE \<rbrace> x) #\<^sub>\<Gamma> GNil"
-
- have cof:"(c_of \<lbrace> z' : B_bool | TRUE \<rbrace> x) = C_true" using c_of.simps check_whileI subst_cv.simps by metis
- have wfg: "\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f ?G" proof
- show "c_of \<lbrace> z' : B_bool | TRUE \<rbrace> x \<in> {TRUE, FALSE}" using subst_cv.simps cof by auto
- show "\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f GNil " using wfG_nilI check_whileI wfX_wfY check_s_wf by metis
- show "atom x \<sharp> GNil" using fresh_GNil by auto
- show "\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f b_of \<lbrace> z' : B_bool | TRUE \<rbrace> " using wfB_boolI wfX_wfY check_s_wf b_of.simps
- by (metis \<open>\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f GNil\<close>)
- qed
-
- obtain zz::x where zf:\<open>atom zz \<sharp> ((\<Theta>, \<Phi>, {||}::bv fset, ?G , \<Delta>, [ x ]\<^sup>v,
- AS_seq s2 (AS_while s1 s2), AS_val [ L_unit ]\<^sup>v, \<tau>'),x,?G)\<close>
- using obtain_fresh by blast
- show \<open> \<Theta> ; \<Phi> ; {||} ; ?G ; \<Delta> \<turnstile>
- AS_if [ x ]\<^sup>v (AS_seq s2 (AS_while s1 s2)) (AS_val [ L_unit ]\<^sup>v) \<Leftarrow> \<tau>'\<close>
- proof
- show "atom zz \<sharp> (\<Theta>, \<Phi>, {||}::bv fset, ?G , \<Delta>, [ x ]\<^sup>v, AS_seq s2 (AS_while s1 s2), AS_val [ L_unit ]\<^sup>v, \<tau>')" using zf by auto
- show \<open>\<Theta> ; {||} ; ?G \<turnstile> [ x ]\<^sup>v \<Leftarrow> \<lbrace> zz : B_bool | TRUE \<rbrace>\<close> proof
- have "atom zz \<sharp> x \<and> atom zz \<sharp> (\<Theta>, {||}::bv fset, ?G)" using zf fresh_prodN by metis
- thus \<open> \<Theta> ; {||} ; ?G \<turnstile> [ x ]\<^sup>v \<Rightarrow>\<lbrace> zz : B_bool | [[zz]\<^sup>v]\<^sup>c\<^sup>e == [[ x ]\<^sup>v]\<^sup>c\<^sup>e \<rbrace>\<close>
- using infer_v_varI lookup.simps wfg b_of.simps by metis
- thus \<open>\<Theta> ; {||} ; ?G \<turnstile> \<lbrace> zz : B_bool | [[ zz ]\<^sup>v]\<^sup>c\<^sup>e == [[ x ]\<^sup>v]\<^sup>c\<^sup>e \<rbrace> \<lesssim> \<lbrace> zz : B_bool | TRUE \<rbrace>\<close>
- using subtype_top infer_v_wf by metis
- qed
- show \<open> \<Theta> ; \<Phi> ; {||} ; ?G ; \<Delta> \<turnstile> AS_seq s2 (AS_while s1 s2) \<Leftarrow> \<lbrace> zz : b_of \<tau>' | [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of \<tau>' zz \<rbrace>\<close>
- proof
- have "\<lbrace> zz : B_unit | TRUE \<rbrace> = \<lbrace> z : B_unit | TRUE \<rbrace>" using \<tau>.eq_iff by auto
- thus \<open> \<Theta> ; \<Phi> ; {||} ; ?G ; \<Delta> \<turnstile> s2 \<Leftarrow> \<lbrace> zz : B_unit | TRUE \<rbrace>\<close> using check_s_g_weakening(1) [OF check_whileI(3) _ wfg] toSet.simps
- by (simp add: \<open>\<lbrace> zz : B_unit | TRUE \<rbrace> = \<lbrace> z : B_unit | TRUE \<rbrace>\<close>)
-
- show \<open> \<Theta> ; \<Phi> ; {||} ; ?G ; \<Delta> \<turnstile> AS_while s1 s2 \<Leftarrow> \<lbrace> zz : b_of \<tau>' | [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of \<tau>' zz \<rbrace>\<close>
- proof(rule check_s_supertype(1))
- have \<open> \<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_while s1 s2 \<Leftarrow> \<tau>'\<close> using check_whileI by auto
- thus *:\<open> \<Theta> ; \<Phi> ; {||} ; ?G ; \<Delta> \<turnstile> AS_while s1 s2 \<Leftarrow> \<tau>' \<close> using check_s_g_weakening(1)[OF _ _ wfg] toSet.simps by auto
-
- show \<open>\<Theta> ; {||} ; ?G \<turnstile> \<tau>' \<lesssim> \<lbrace> zz : b_of \<tau>' | [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of \<tau>' zz \<rbrace>\<close>
- proof(rule subtype_eq_if_\<tau>)
- show \<open> \<Theta> ; {||} ; ?G \<turnstile>\<^sub>w\<^sub>f \<tau>' \<close> using * check_s_wf by auto
- show \<open> \<Theta> ; {||} ; ?G \<turnstile>\<^sub>w\<^sub>f \<lbrace> zz : b_of \<tau>' | [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of \<tau>' zz \<rbrace> \<close>
- apply(rule wfT_eq_imp, simp add: base_for_lit.simps)
- using subtype_wf check_whileI wfg zf fresh_prodN by metis+
- show \<open>atom zz \<sharp> \<tau>'\<close> using zf fresh_prodN by metis
- qed
- qed
-
- qed
- show \<open> \<Theta> ; \<Phi> ; {||} ; ?G ; \<Delta> \<turnstile> AS_val [ L_unit ]\<^sup>v \<Leftarrow> \<lbrace> zz : b_of \<tau>' | [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_false ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of \<tau>' zz \<rbrace>\<close>
- proof(rule check_s_supertype(1))
-
- show *:\<open> \<Theta> ; \<Phi> ; {||} ; ?G ; \<Delta> \<turnstile> AS_val [ L_unit ]\<^sup>v \<Leftarrow> \<tau>'\<close>
- using check_unit[OF check_whileI(5) _ _ wfg] using check_whileI wfg wfX_wfY check_s_wf by metis
- show \<open>\<Theta> ; {||} ; ?G \<turnstile> \<tau>' \<lesssim> \<lbrace> zz : b_of \<tau>' | [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_false ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of \<tau>' zz \<rbrace>\<close>
- proof(rule subtype_eq_if_\<tau>)
- show \<open> \<Theta> ; {||} ; ?G \<turnstile>\<^sub>w\<^sub>f \<tau>' \<close> using * check_s_wf by metis
- show \<open> \<Theta> ; {||} ; ?G \<turnstile>\<^sub>w\<^sub>f \<lbrace> zz : b_of \<tau>' | [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_false ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of \<tau>' zz \<rbrace> \<close>
- apply(rule wfT_eq_imp, simp add: base_for_lit.simps)
- using subtype_wf check_whileI wfg zf fresh_prodN by metis+
- show \<open>atom zz \<sharp> \<tau>'\<close> using zf fresh_prodN by metis
- qed
- qed
- qed
- qed
-qed(auto+)
-
-lemma check_s_narrow:
- fixes s::s and x::x
- assumes "atom x \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>, \<Delta>, c, \<tau>, s)" and "\<Theta> ; \<Phi> ; \<B> ; (x, B_bool, c) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile> s \<Leftarrow> \<tau>" and
- "\<Theta> ; \<B> ; \<Gamma> \<Turnstile> c "
- shows "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile> s \<Leftarrow> \<tau>"
-proof -
- let ?B = "({||}::bv fset)"
- let ?v = "V_lit L_true"
-
- obtain z::x where z: "atom z \<sharp> (x, [ L_true ]\<^sup>v,c)" using obtain_fresh by metis
-
- have "atom z \<sharp> c" using z fresh_prodN by auto
- hence c:" c[x::=[ z ]\<^sup>v]\<^sub>v[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v = c" using subst_v_c_def by simp
-
- have "\<Theta> ; \<Phi> ; \<B> ; ((x,B_bool, c) #\<^sub>\<Gamma> \<Gamma>)[x::=?v]\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=?v]\<^sub>\<Delta>\<^sub>v \<turnstile> s[x::=?v]\<^sub>s\<^sub>v \<Leftarrow> \<tau>[x::=?v]\<^sub>\<tau>\<^sub>v" proof(rule subst_infer_check_s(1) )
- show vt: \<open> \<Theta> ; \<B> ; \<Gamma> \<turnstile> [ L_true ]\<^sup>v \<Rightarrow> \<lbrace> z : B_bool | (CE_val (V_var z)) == (CE_val (V_lit L_true )) \<rbrace> \<close>
- using infer_v_litI check_s_wf wfG_elims(2) infer_trueI assms by metis
- show \<open>\<Theta> ; \<B> ; \<Gamma> \<turnstile> \<lbrace> z : B_bool | (CE_val (V_var z)) == (CE_val (V_lit L_true )) \<rbrace> \<lesssim> \<lbrace> z : B_bool | c[x::=[ z ]\<^sup>v]\<^sub>v \<rbrace>\<close> proof
- show \<open>atom x \<sharp> (\<Theta>, \<B>, \<Gamma>, z, [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e , z, c[x::=[ z ]\<^sup>v]\<^sub>v)\<close>
- apply(unfold fresh_prodN, intro conjI)
- prefer 5
- using c.fresh ce.fresh v.fresh z fresh_prodN apply auto[1]
- prefer 6
- using fresh_subst_v_if[of "atom x" c x] assms fresh_prodN apply simp
- using z assms fresh_prodN apply metis
- using z assms fresh_prodN apply metis
- using z assms fresh_prodN apply metis
- using z fresh_prodN assms fresh_at_base by metis+
- show \<open> \<Theta> ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : B_bool | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrace> \<close> using vt infer_v_wf by simp
- show \<open> \<Theta> ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : B_bool | c[x::=[ z ]\<^sup>v]\<^sub>v \<rbrace> \<close> proof(rule wfG_wfT)
- show \<open> \<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f (x, B_bool, c[x::=[ z ]\<^sup>v]\<^sub>v[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<close> using c check_s_wf assms by metis
- have " atom x \<sharp> [ z ]\<^sup>v" using v.fresh z fresh_at_base by auto
- thus \<open>atom x \<sharp> c[x::=[ z ]\<^sup>v]\<^sub>v\<close> using fresh_subst_v_if[of "atom x" c ] by auto
- qed
- have wfg: "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f(x, B_bool, ([ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e )[z::=[ x ]\<^sup>v]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>"
- using wfT_wfG vt infer_v_wf fresh_prodN assms by simp
- show \<open>\<Theta> ; \<B> ; (x, B_bool, ([ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e )[z::=[ x ]\<^sup>v]\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<Turnstile> c[x::=[ z ]\<^sup>v]\<^sub>v[z::=[ x ]\<^sup>v]\<^sub>v \<close>
- using c valid_weakening[OF assms(3) _ wfg ] toSet.simps
- using subst_v_c_def by auto
- qed
- show \<open>atom z \<sharp> (x, [ L_true ]\<^sup>v)\<close> using z fresh_prodN by metis
- show \<open> \<Theta> ; \<Phi> ; \<B> ; (x, B_bool, c) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile> s \<Leftarrow> \<tau>\<close> using assms by auto
-
- thus \<open>(x, B_bool, c) #\<^sub>\<Gamma> \<Gamma> = GNil @ (x, B_bool, c[x::=[ z ]\<^sup>v]\<^sub>v[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>\<close> using append_g.simps c by auto
- qed
-
- moreover have "((x,B_bool, c) #\<^sub>\<Gamma> \<Gamma>)[x::=?v]\<^sub>\<Gamma>\<^sub>v = \<Gamma> " using subst_gv.simps by auto
- ultimately show ?thesis using assms forget_subst_dv forget_subst_sv forget_subst_tv fresh_prodN by metis
-qed
-
-lemma check_assert_s:
- fixes s::s and x::x
- assumes "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_assert c s \<Leftarrow> \<tau>"
- shows "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> s \<Leftarrow> \<tau> \<and> \<Theta> ; {||} ; GNil \<Turnstile> c "
-proof -
- let ?B = "({||}::bv fset)"
- let ?v = "V_lit L_true"
-
- obtain x where x: "\<Theta> ; \<Phi> ; ?B ; (x,B_bool, c) #\<^sub>\<Gamma> GNil ; \<Delta> \<turnstile> s \<Leftarrow> \<tau> \<and> atom x \<sharp> (\<Theta>, \<Phi>, ?B, GNil, \<Delta>, c, \<tau>, s ) \<and> \<Theta> ; ?B ; GNil \<Turnstile> c "
- using check_s_elims(10)[OF \<open>\<Theta> ; \<Phi> ; ?B ; GNil ; \<Delta> \<turnstile> AS_assert c s \<Leftarrow> \<tau>\<close>] valid.simps by metis
-
- show ?thesis using assms check_s_narrow x by metis
-qed
-
-lemma infer_v_pair2I:
- "atom z \<sharp> (v1, v2) \<Longrightarrow>
- atom z \<sharp> (\<Theta>, \<B>, \<Gamma>) \<Longrightarrow>
- \<Theta> ; \<B> ; \<Gamma> \<turnstile> v1 \<Rightarrow> t1 \<Longrightarrow>
- \<Theta> ; \<B> ; \<Gamma> \<turnstile> v2 \<Rightarrow> t2 \<Longrightarrow>
- b1 = b_of t1 \<Longrightarrow> b2 = b_of t2 \<Longrightarrow>
- \<Theta> ; \<B> ; \<Gamma> \<turnstile> [ v1 , v2 ]\<^sup>v \<Rightarrow> \<lbrace> z : [ b1 , b2 ]\<^sup>b | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ v1 , v2 ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrace>"
- using infer_v_pairI by simp
-
-subsection \<open>Main Lemma\<close>
-
-lemma preservation:
- assumes "\<Phi> \<turnstile> \<langle>\<delta>, s\<rangle> \<longrightarrow> \<langle>\<delta>', s'\<rangle>" and "\<Theta>; \<Phi>; \<Delta> \<turnstile> \<langle>\<delta>, s\<rangle> \<Leftarrow> \<tau>"
- shows "\<exists>\<Delta>'. \<Theta>; \<Phi>; \<Delta>' \<turnstile> \<langle>\<delta>', s'\<rangle> \<Leftarrow> \<tau> \<and> \<Delta> \<sqsubseteq> \<Delta>'"
- using assms
-proof(induct arbitrary: \<tau> rule: reduce_stmt.induct)
- case (reduce_let_plusI \<delta> x n1 n2 s')
- then show ?case using preservation_plus
- by (metis order_refl)
-next
- case (reduce_let_leqI b n1 n2 \<delta> x s)
- then show ?case using preservation_leq by (metis order_refl)
-next
- case (reduce_let_eqI b n1 n2 \<Phi> \<delta> x s)
- then show ?case using preservation_eq[OF reduce_let_eqI(2)] order_refl by metis
-next
- case (reduce_let_appI f z b c \<tau>' s' \<Phi> \<delta> x v s)
- hence tt: "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let x (AE_app f v) s \<Leftarrow> \<tau> \<and> \<Theta> \<turnstile> \<delta> \<sim> \<Delta> \<and> (\<forall>fd\<in>set \<Phi>. check_fundef \<Theta> \<Phi> fd)" using config_type_elims[OF reduce_let_appI(2)] by metis
- hence *:"\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let x (AE_app f v) s \<Leftarrow> \<tau>" by auto
-
- hence "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let2 x (\<tau>'[z::=v]\<^sub>\<tau>\<^sub>v) (s'[z::=v]\<^sub>s\<^sub>v) s \<Leftarrow> \<tau>"
- using preservation_app reduce_let_appI tt by auto
-
- hence "\<Theta>; \<Phi>; \<Delta> \<turnstile> \<langle> \<delta> , AS_let2 x (\<tau>'[z::=v]\<^sub>\<tau>\<^sub>v) s'[z::=v]\<^sub>s\<^sub>v s \<rangle> \<Leftarrow> \<tau>" using config_typeI tt by auto
- then show ?case using tt order_refl reduce_let_appI by metis
-
-next
- case (reduce_let_appPI f bv z b c \<tau>' s' \<Phi> \<delta> x b' v s)
-
- hence tt: "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let x (AE_appP f b' v) s \<Leftarrow> \<tau> \<and> \<Theta> \<turnstile> \<delta> \<sim> \<Delta> \<and> (\<forall>fd\<in>set \<Phi>. check_fundef \<Theta> \<Phi> fd)" using config_type_elims[OF reduce_let_appPI(2)] by metis
- hence *:"\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let x (AE_appP f b' v) s \<Leftarrow> \<tau>" by auto
-
- have "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let2 x (\<tau>'[bv::=b']\<^sub>\<tau>\<^sub>b[z::=v]\<^sub>\<tau>\<^sub>v) (s'[bv::=b']\<^sub>s\<^sub>b[z::=v]\<^sub>s\<^sub>v) s \<Leftarrow> \<tau>"
- proof(rule preservation_poly_app)
- show \<open>Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ z b c \<tau>' s'))) = lookup_fun \<Phi> f\<close> using reduce_let_appPI by metis
- show \<open>\<forall>fd\<in>set \<Phi>. check_fundef \<Theta> \<Phi> fd\<close> using tt lookup_fun_member by metis
- show \<open> \<Theta> ; \<Phi> ;{||} ; GNil ; \<Delta> \<turnstile> AS_let x (AE_appP f b' v) s \<Leftarrow> \<tau>\<close> using * by auto
- show \<open> \<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f b' \<close> using check_s_elims infer_e_wf wfE_elims * by metis
- qed(auto+)
-
- hence "\<Theta>; \<Phi>; \<Delta> \<turnstile> \<langle> \<delta> , AS_let2 x (\<tau>'[bv::=b']\<^sub>\<tau>\<^sub>b[z::=v]\<^sub>\<tau>\<^sub>v) s'[bv::=b']\<^sub>s\<^sub>b[z::=v]\<^sub>s\<^sub>v s \<rangle> \<Leftarrow> \<tau>" using config_typeI tt by auto
- then show ?case using tt order_refl reduce_let_appI by metis
-
-next
- case (reduce_if_trueI \<delta> s1 s2)
- then show ?case using preservation_if by metis
-next
- case (reduce_if_falseI uw \<delta> s1 s2)
- then show ?case using preservation_if by metis
-next
- case (reduce_let_valI \<delta> x v s)
- then show ?case using preservation_let_val by presburger
-next
- case (reduce_let_mvar u v \<delta> \<Phi> x s)
- hence *:"\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let x (AE_mvar u) s \<Leftarrow> \<tau> \<and> \<Theta> \<turnstile> \<delta> \<sim> \<Delta> \<and> (\<forall>fd\<in>set \<Phi>. check_fundef \<Theta> \<Phi> fd)"
- using config_type_elims by blast
-
- hence **: "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let x (AE_mvar u) s \<Leftarrow> \<tau>" by auto
- obtain xa::x and za::x and ca::c and ba::b and sa::s where
- sa1: "atom xa \<sharp> (\<Theta>, \<Phi>, {||}::bv fset, GNil, \<Delta>, AE_mvar u, \<tau>) \<and> atom za \<sharp> (xa, \<Theta>, \<Phi>, {||}::bv fset, GNil, \<Delta>, AE_mvar u, \<tau>, sa) \<and>
- \<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AE_mvar u \<Rightarrow> \<lbrace> za : ba | ca \<rbrace> \<and>
- \<Theta> ; \<Phi> ; {||} ; (xa, ba, ca[za::=V_var xa]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> GNil ; \<Delta> \<turnstile> sa \<Leftarrow> \<tau> \<and>
- (\<forall>c. atom c \<sharp> (s, sa) \<longrightarrow> atom c \<sharp> (x, xa, s, sa) \<longrightarrow> (x \<leftrightarrow> c) \<bullet> s = (xa \<leftrightarrow> c) \<bullet> sa)"
- using check_s_elims(2)[OF **] subst_defs by metis
-
- have "\<Theta> ; {||} ; GNil \<turnstile> v \<Leftarrow> \<lbrace> za : ba | ca \<rbrace>" proof -
- have " (u , \<lbrace> za : ba | ca \<rbrace>) \<in> setD \<Delta>" using infer_e_elims(11) sa1 by fast
- thus ?thesis using delta_sim_v reduce_let_mvar config_type_elims check_s_wf by metis
- qed
-
- then obtain \<tau>' where vst: "\<Theta> ; {||} ; GNil \<turnstile> v \<Rightarrow> \<tau>' \<and>
- \<Theta> ; {||} ; GNil \<turnstile> \<tau>' \<lesssim> \<lbrace> za : ba | ca \<rbrace>" using check_v_elims by blast
-
- obtain za2 and ba2 and ca2 where zbc: "\<tau>' = (\<lbrace> za2 : ba2 | ca2 \<rbrace>) \<and> atom za2 \<sharp> (xa, (xa, \<Theta>, \<Phi>, {||}::bv fset, GNil, \<Delta>, AE_val v, \<tau>, sa))"
- using obtain_fresh_z by blast
- have beq: "ba=ba2" using subtype_eq_base vst zbc by blast
-
- moreover have xaf: "atom xa \<sharp> (za, za2)"
- apply(unfold fresh_prodN, intro conjI)
- using sa1 zbc fresh_prodN fresh_x_neq by metis+
-
- have sat2: " \<Theta> ; \<Phi> ; {||} ; GNil@(xa, ba, ca2[za2::=V_var xa]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> GNil ; \<Delta> \<turnstile> sa \<Leftarrow> \<tau>" proof(rule ctx_subtype_s)
- show "\<Theta> ; \<Phi> ; {||} ; GNil @ (xa, ba, ca[za::=V_var xa]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> GNil ; \<Delta> \<turnstile> sa \<Leftarrow> \<tau>" using sa1 by auto
- show "\<Theta> ; {||} ; GNil \<turnstile> \<lbrace> za2 : ba | ca2 \<rbrace> \<lesssim> \<lbrace> za : ba | ca \<rbrace>" using beq zbc vst by fast
- show "atom xa \<sharp> (za, za2, ca, ca2)" proof -
- have *:"\<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f (\<lbrace> za2 : ba2 | ca2 \<rbrace>) " using zbc vst subtype_wf by auto
- hence "supp ca2 \<subseteq> { atom za2 }" using wfT_supp_c[OF *] supp_GNil by simp
- moreover have "atom za2 \<sharp> xa" using zbc fresh_Pair fresh_x_neq by metis
- ultimately have "atom xa \<sharp> ca2" using zbc supp_at_base fresh_def
- by (metis empty_iff singleton_iff subset_singletonD)
- moreover have "atom xa \<sharp> ca" proof -
- have *:"\<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f (\<lbrace> za : ba | ca \<rbrace>) " using zbc vst subtype_wf by auto
- hence "supp ca \<subseteq> { atom za }" using wfT_supp \<tau>.supp by force
- moreover have "xa \<noteq> za" using fresh_def fresh_x_neq xaf fresh_Pair by metis
- ultimately show ?thesis using fresh_def by auto
- qed
- ultimately show ?thesis using xaf sa1 fresh_prod4 fresh_Pair by metis
- qed
- qed
- hence dwf: "\<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f \<Delta>" using sa1 infer_e_wf by meson
-
- have "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let xa (AE_val v) sa \<Leftarrow> \<tau>" proof
- have "atom xa \<sharp> (AE_val v)" using infer_v_wf(1) wfV_supp fresh_def e.fresh x_not_in_b_set vst by fastforce
- thus "atom xa \<sharp> (\<Theta>, \<Phi>, {||}::bv fset, GNil, \<Delta>, AE_val v, \<tau>)" using sa1 freshers by simp
- have "atom za2 \<sharp> (AE_val v)" using infer_v_wf(1) wfV_supp fresh_def e.fresh x_not_in_b_set vst by fastforce
- thus "atom za2 \<sharp> (xa, \<Theta>, \<Phi>, {||}::bv fset, GNil, \<Delta>, AE_val v, \<tau>, sa)" using zbc freshers fresh_prodN by auto
- have " \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>" using sa1 infer_e_wf by auto
- thus "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AE_val v \<Rightarrow> \<lbrace> za2 : ba | ca2 \<rbrace>"
- using zbc vst beq dwf infer_e_valI by blast
- show "\<Theta> ; \<Phi> ; {||} ; (xa, ba, ca2[za2::=V_var xa]\<^sub>v) #\<^sub>\<Gamma> GNil ; \<Delta> \<turnstile> sa \<Leftarrow> \<tau>" using sat2 append_g.simps subst_defs by metis
- qed
- moreover have "AS_let xa (AE_val v) sa = AS_let x (AE_val v) s" proof -
- have "[[atom x]]lst. s = [[atom xa]]lst. sa"
- using sa1 Abs1_eq_iff_all(3)[where z=" (s, sa)"] by metis
- thus ?thesis using s_branch_s_branch_list.eq_iff(2) by metis
- qed
- ultimately have "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let x (AE_val v) s \<Leftarrow> \<tau>" by auto
-
- then show ?case using reduce_let_mvar * config_typeI
- by (meson order_refl)
-next
- case (reduce_let2I \<Phi> \<delta> s1 \<delta>' s1' x t s2)
- hence **: "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let2 x t s1 s2 \<Leftarrow> \<tau> \<and> \<Theta> \<turnstile> \<delta> \<sim> \<Delta> \<and> (\<forall>fd\<in>set \<Phi>. check_fundef \<Theta> \<Phi> fd)" using config_type_elims[OF reduce_let2I(3)] by blast
- hence *:"\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let2 x t s1 s2 \<Leftarrow> \<tau>" by auto
-
- obtain xa::x and z::x and c and b and s2a::s where st: "atom xa \<sharp> (\<Theta>, \<Phi>, {||}::bv fset, GNil, \<Delta>, t, s1, \<tau>) \<and>
- \<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> s1 \<Leftarrow> t \<and>
- \<Theta> ; \<Phi> ; {||} ; (xa, b_of t, c_of t xa) #\<^sub>\<Gamma> GNil ; \<Delta> \<turnstile> s2a \<Leftarrow> \<tau> \<and> ([[atom x]]lst. s2 = [[atom xa]]lst. s2a) "
- using check_s_elims(4)[OF *] Abs1_eq_iff_all(3) by metis
-
- hence "\<Theta>; \<Phi>; \<Delta> \<turnstile> \<langle> \<delta> , s1 \<rangle> \<Leftarrow> t" using config_typeI ** by auto
- then obtain \<Delta>' where s1r: "\<Theta>; \<Phi>; \<Delta>' \<turnstile> \<langle> \<delta>' , s1' \<rangle> \<Leftarrow> t \<and> \<Delta> \<sqsubseteq> \<Delta>'" using reduce_let2I by presburger
-
- have "\<Theta>; \<Phi>; {||}; GNil; \<Delta>' \<turnstile> AS_let2 xa t s1' s2a \<Leftarrow> \<tau>"
- proof(rule check_let2I)
- show *:"\<Theta>; \<Phi>; {||}; GNil; \<Delta>' \<turnstile> s1' \<Leftarrow> t" using config_type_elims st s1r by metis
- show "atom xa \<sharp> (\<Theta>, \<Phi>, {||}::bv fset, GNil, \<Delta>',t, s1', \<tau>)" proof -
- have "atom xa \<sharp> s1'" using check_s_x_fresh * by auto
- moreover have "atom xa \<sharp> \<Delta>'" using check_s_x_fresh * by auto
- ultimately show ?thesis using st fresh_prodN by metis
- qed
-
- show "\<Theta> ; \<Phi> ; {||} ; (xa, b_of t, c_of t xa) #\<^sub>\<Gamma> GNil ; \<Delta>' \<turnstile> s2a \<Leftarrow> \<tau>" proof -
- have "\<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f \<Delta>'" using * check_s_wf by auto
- moreover have "\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f ((xa, b_of t, c_of t xa) #\<^sub>\<Gamma> GNil)" using st check_s_wf by auto
- ultimately have "\<Theta> ; {||} ; ((xa, b_of t , c_of t xa) #\<^sub>\<Gamma> GNil) \<turnstile>\<^sub>w\<^sub>f \<Delta>'" using wf_weakening by auto
- thus ?thesis using check_s_d_weakening check_s_wf st s1r by metis
- qed
- qed
- moreover have "AS_let2 xa t s1' s2a = AS_let2 x t s1' s2" using st s_branch_s_branch_list.eq_iff by metis
- ultimately have "\<Theta>; \<Phi>; {||}; GNil; \<Delta>' \<turnstile> AS_let2 x t s1' s2 \<Leftarrow> \<tau>" using st by argo
- moreover have "\<Theta> \<turnstile> \<delta>' \<sim> \<Delta>'" using config_type_elims s1r by fast
- ultimately show ?case using config_typeI **
- by (meson s1r)
-next
- case (reduce_let2_valI vb \<delta> x t v s)
- then show ?case using preservation_let_val by meson
-next
- case (reduce_varI u \<delta> \<Phi> \<tau>' v s)
-
- hence ** : "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_var u \<tau>' v s \<Leftarrow> \<tau> \<and> \<Theta> \<turnstile> \<delta> \<sim> \<Delta> \<and> (\<forall>fd\<in>set \<Phi>. check_fundef \<Theta> \<Phi> fd)"
- using config_type_elims by meson
- have uf: "atom u \<sharp> \<Delta>" using reduce_varI delta_sim_fresh by force
- hence *: "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_var u \<tau>' v s \<Leftarrow> \<tau>" and " \<Theta> \<turnstile> \<delta> \<sim> \<Delta>" using ** by auto
-
- thus ?case using preservation_var reduce_varI config_typeI ** set_subset_Cons
- setD_ConsD subsetI by (metis delta_sim_fresh)
-
-next
- case (reduce_assignI \<Phi> \<delta> u v )
- hence *: "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_assign u v \<Leftarrow> \<tau> \<and> \<Theta> \<turnstile> \<delta> \<sim> \<Delta> \<and> (\<forall>fd\<in>set \<Phi>. check_fundef \<Theta> \<Phi> fd)"
- using config_type_elims by meson
- then obtain z and \<tau>' where zt: "\<Theta> ; {||} ; GNil \<turnstile> (\<lbrace> z : B_unit | TRUE \<rbrace>) \<lesssim> \<tau> \<and> (u,\<tau>') \<in> setD \<Delta> \<and> \<Theta> ; {||} ; GNil \<turnstile> v \<Leftarrow> \<tau>' \<and> \<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f \<Delta>"
- using check_s_elims(8) by metis
- hence "\<Theta> \<turnstile> update_d \<delta> u v \<sim> \<Delta>" using update_d_sim * by metis
- moreover have "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_val (V_lit L_unit ) \<Leftarrow> \<tau>" using zt * check_s_v_unit check_s_wf
- by auto
- ultimately show ?case using config_typeI * by (meson order_refl)
-next
- case (reduce_seq1I \<Phi> \<delta> s)
- hence "\<Theta> ; \<Phi> ; {||} ; GNil ; \<Delta> \<turnstile> s \<Leftarrow> \<tau> \<and> \<Theta> \<turnstile> \<delta> \<sim> \<Delta> \<and> (\<forall>fd\<in>set \<Phi>. check_fundef \<Theta> \<Phi> fd)"
- using check_s_elims config_type_elims by force
- then show ?case using config_typeI by blast
-next
- case (reduce_seq2I s1 v \<Phi> \<delta> \<delta>' s1' s2)
- hence tt: "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_seq s1 s2 \<Leftarrow> \<tau> \<and> \<Theta> \<turnstile> \<delta> \<sim> \<Delta> \<and> (\<forall>fd\<in>set \<Phi>. check_fundef \<Theta> \<Phi> fd)"
- using config_type_elims by blast
- then obtain z where zz: "\<Theta> ; \<Phi> ; {||} ; GNil; \<Delta> \<turnstile> s1 \<Leftarrow> (\<lbrace> z : B_unit | TRUE \<rbrace>) \<and> \<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> s2 \<Leftarrow> \<tau>"
- using check_s_elims by blast
- hence "\<Theta>; \<Phi>; \<Delta> \<turnstile> \<langle> \<delta> , s1 \<rangle> \<Leftarrow> (\<lbrace> z : B_unit | TRUE \<rbrace>)"
- using tt config_typeI tt by simp
- then obtain \<Delta>' where *: "\<Theta>; \<Phi>; \<Delta>' \<turnstile> \<langle> \<delta>' , s1' \<rangle> \<Leftarrow> (\<lbrace> z : B_unit | TRUE \<rbrace>) \<and> \<Delta> \<sqsubseteq> \<Delta>'"
- using reduce_seq2I by meson
- moreover hence s't: "\<Theta>; \<Phi>; {||}; GNil; \<Delta>' \<turnstile> s1' \<Leftarrow> (\<lbrace> z : B_unit | TRUE \<rbrace>) \<and> \<Theta> \<turnstile> \<delta>' \<sim> \<Delta>'"
- using config_type_elims by force
- moreover hence "\<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f \<Delta>'" using check_s_wf by meson
- moreover hence "\<Theta>; \<Phi>; {||}; GNil; \<Delta>' \<turnstile> s2 \<Leftarrow> \<tau>"
- using calculation(1) zz check_s_d_weakening * by metis
- moreover hence "\<Theta>; \<Phi>; {||}; GNil; \<Delta>' \<turnstile> (AS_seq s1' s2) \<Leftarrow> \<tau>"
- using check_seqI zz s't by meson
- ultimately have "\<Theta>; \<Phi>; \<Delta>' \<turnstile> \<langle> \<delta>' , AS_seq s1' s2 \<rangle> \<Leftarrow> \<tau> \<and> \<Delta> \<sqsubseteq> \<Delta>'"
- using zz config_typeI tt by meson
- then show ?case by meson
-next
- case (reduce_whileI x s1 s2 z' \<Phi> \<delta> )
-
- hence *: "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_while s1 s2 \<Leftarrow> \<tau> \<and> \<Theta> \<turnstile> \<delta> \<sim> \<Delta> \<and> (\<forall>fd\<in>set \<Phi>. check_fundef \<Theta> \<Phi> fd)"
- using config_type_elims by meson
-
- hence **:"\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_while s1 s2 \<Leftarrow> \<tau>" by auto
- hence "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let2 x (\<lbrace> z' : B_bool | TRUE \<rbrace>) s1 (AS_if (V_var x) (AS_seq s2 (AS_while s1 s2)) (AS_val (V_lit L_unit)) ) \<Leftarrow> \<tau>"
- using check_while reduce_whileI by auto
- thus ?case using config_typeI * by (meson subset_refl)
-
-next
- case (reduce_caseI dc x' s' css \<Phi> \<delta> tyid v)
-
- hence **: "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_match (V_cons tyid dc v) css \<Leftarrow> \<tau> \<and> \<Theta> \<turnstile> \<delta> \<sim> \<Delta> \<and> (\<forall>fd\<in>set \<Phi>. check_fundef \<Theta> \<Phi> fd)"
- using config_type_elims[OF reduce_caseI(2)] by metis
- hence ***: "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_match (V_cons tyid dc v) css \<Leftarrow> \<tau>" by auto
-
- let ?vcons = "V_cons tyid dc v"
-
- obtain dclist tid and z::x where cv: "\<Theta>; {||} ; GNil \<turnstile> (V_cons tyid dc v) \<Leftarrow> (\<lbrace> z : B_id tid | TRUE \<rbrace>) \<and>
- \<Theta>; \<Phi>; {||}; GNil; \<Delta> ; tid ; dclist ; (V_cons tyid dc v) \<turnstile> css \<Leftarrow> \<tau> \<and> AF_typedef tid dclist \<in> set \<Theta> \<and>
- \<Theta> ; {||} ; GNil \<turnstile> V_cons tyid dc v \<Leftarrow> \<lbrace> z : B_id tid | TRUE \<rbrace>"
- using check_s_elims(9)[OF ***] by metis
-
- hence vi:" \<Theta> ; {||} ; GNil \<turnstile> V_cons tyid dc v \<Leftarrow> \<lbrace> z : B_id tid | TRUE \<rbrace>" by auto
- obtain tcons where vi2:" \<Theta> ; {||} ; GNil \<turnstile> V_cons tyid dc v \<Rightarrow> tcons \<and> \<Theta> ; {||} ; GNil \<turnstile> tcons \<lesssim> \<lbrace> z : B_id tid | TRUE \<rbrace>"
- using check_v_elims(1)[OF vi] by metis
- hence vi1: "\<Theta> ; {||} ; GNil \<turnstile> V_cons tyid dc v \<Rightarrow> tcons" by auto
-
- show ?case proof(rule infer_v_elims(4)[OF vi1],goal_cases)
- case (1 dclist2 tc tv z2)
- have "tyid = tid" using \<tau>.eq_iff using subtype_eq_base vi2 1 by fastforce
- hence deq:"dclist = dclist2" using check_v_wf wfX_wfY cv 1 wfTh_dclist_unique by metis
- have "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> s'[x'::=v]\<^sub>s\<^sub>v \<Leftarrow> \<tau>" proof(rule check_match(3))
- show \<open> \<Theta> ; \<Phi> ; {||} ; GNil ; \<Delta> ; tyid ; dclist ; ?vcons \<turnstile> css \<Leftarrow> \<tau>\<close> using \<open>tyid = tid\<close> cv by auto
- show "distinct (map fst dclist)" using wfTh_dclist_distinct check_v_wf wfX_wfY cv by metis
- show \<open>?vcons = V_cons tyid dc v\<close> by auto
- show \<open>{||} = {||}\<close> by auto
- show \<open>(dc, tc) \<in> set dclist\<close> using 1 deq by auto
- show \<open>GNil = GNil\<close> by auto
- show \<open>Some (AS_branch dc x' s') = lookup_branch dc css\<close> using reduce_caseI by auto
- show \<open>\<Theta> ; {||} ; GNil \<turnstile> v \<Leftarrow> tc\<close> using 1 check_v.intros by auto
- qed
- thus ?case using config_typeI ** by blast
- qed
-
-next
- case (reduce_let_fstI \<Phi> \<delta> x v1 v2 s)
- thus ?case using preservation_fst_snd order_refl by metis
-next
- case (reduce_let_sndI \<Phi> \<delta> x v1 v2 s)
- thus ?case using preservation_fst_snd order_refl by metis
-next
- case (reduce_let_concatI \<Phi> \<delta> x v1 v2 s)
- hence elim: "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let x (AE_concat (V_lit (L_bitvec v1)) (V_lit (L_bitvec v2))) s \<Leftarrow> \<tau> \<and>
- \<Theta> \<turnstile> \<delta> \<sim> \<Delta> \<and> (\<forall>fd\<in>set \<Phi>. check_fundef \<Theta> \<Phi> fd)"
- using config_type_elims by metis
-
- obtain z::x where z: "atom z \<sharp> (AE_concat (V_lit (L_bitvec v1)) (V_lit (L_bitvec v2)), GNil, CE_val (V_lit (L_bitvec (v1 @ v2))))"
- using obtain_fresh by metis
-
- have "\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f GNil" using check_s_wf elim by auto
-
- have "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let x (AE_val (V_lit (L_bitvec (v1 @ v2)))) s \<Leftarrow> \<tau>" proof(rule subtype_let)
- show \<open> \<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let x (AE_concat (V_lit (L_bitvec v1)) (V_lit (L_bitvec v2))) s \<Leftarrow> \<tau>\<close> using elim by auto
- show \<open>\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> (AE_concat (V_lit (L_bitvec v1)) (V_lit (L_bitvec v2))) \<Rightarrow> \<lbrace> z : B_bitvec | CE_val (V_var z) == (CE_concat ([V_lit (L_bitvec v1)]\<^sup>c\<^sup>e) ([V_lit (L_bitvec v2)]\<^sup>c\<^sup>e))\<rbrace> \<close>
- (is "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> ?e1 \<Rightarrow> ?t1")
- proof
- show \<open> \<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using check_s_wf elim by auto
- show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using check_s_wf elim by auto
- show \<open> \<Theta> ; {||} ; GNil \<turnstile> V_lit (L_bitvec v1) \<Rightarrow> \<lbrace> z : B_bitvec | CE_val (V_var z) == CE_val (V_lit (L_bitvec v1)) \<rbrace>\<close>
- using infer_v_litI infer_l.intros \<open>\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f GNil\<close> fresh_GNil by auto
- show \<open> \<Theta> ; {||} ; GNil \<turnstile> V_lit (L_bitvec v2) \<Rightarrow> \<lbrace> z : B_bitvec | CE_val (V_var z) == CE_val (V_lit (L_bitvec v2)) \<rbrace>\<close>
- using infer_v_litI infer_l.intros \<open>\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f GNil\<close> fresh_GNil by auto
- show \<open>atom z \<sharp> AE_concat (V_lit (L_bitvec v1)) (V_lit (L_bitvec v2))\<close> using z fresh_Pair by metis
- show \<open>atom z \<sharp> GNil\<close> using z fresh_Pair by auto
- qed
- show \<open>\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AE_val (V_lit (L_bitvec (v1 @ v2))) \<Rightarrow> \<lbrace> z : B_bitvec | CE_val (V_var z) == CE_val (V_lit (L_bitvec (v1 @ v2))) \<rbrace> \<close>
- (is "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> ?e2 \<Rightarrow> ?t2")
- using infer_e_valI infer_v_litI infer_l.intros \<open>\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f GNil\<close> fresh_GNil check_s_wf elim by metis
- show \<open>\<Theta> ; {||} ; GNil \<turnstile> ?t2 \<lesssim> ?t1\<close> using subtype_concat check_s_wf elim by auto
- qed
-
- thus ?case using config_typeI elim by (meson order_refl)
-next
- case (reduce_let_lenI \<Phi> \<delta> x v s)
- hence elim: "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let x (AE_len (V_lit (L_bitvec v))) s \<Leftarrow> \<tau> \<and> \<Theta> \<turnstile> \<delta> \<sim> \<Delta> \<and> (\<forall>fd\<in>set \<Phi>. check_fundef \<Theta> \<Phi> fd)"
- using check_s_elims config_type_elims by metis
-
- then obtain t where t: "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AE_len (V_lit (L_bitvec v)) \<Rightarrow> t" using check_s_elims by meson
-
- moreover then obtain z::x where "t = \<lbrace> z : B_int | CE_val (V_var z) == CE_len ([V_lit (L_bitvec v)]\<^sup>c\<^sup>e) \<rbrace>" using infer_e_elims by meson
-
- moreover obtain z'::x where "atom z' \<sharp> v" using obtain_fresh by metis
- moreover have "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AE_val (V_lit (L_num (int (length v)))) \<Rightarrow> \<lbrace> z' : B_int | CE_val (V_var z') == CE_val (V_lit (L_num (int (length v)))) \<rbrace>"
- using infer_e_valI infer_v_litI infer_l.intros(3) t check_s_wf elim
- by (metis infer_l_form2 type_for_lit.simps(3))
-
- moreover have "\<Theta> ; {||} ; GNil \<turnstile> \<lbrace> z' : B_int | CE_val (V_var z') == CE_val (V_lit (L_num (int (length v)))) \<rbrace> \<lesssim>
- \<lbrace> z : B_int | CE_val (V_var z) == CE_len [(V_lit (L_bitvec v))]\<^sup>c\<^sup>e \<rbrace>" using subtype_len check_s_wf elim by auto
-
- ultimately have "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let x (AE_val (V_lit (L_num (int (length v))))) s \<Leftarrow> \<tau>" using subtype_let by (meson elim)
- thus ?case using config_typeI elim by (meson order_refl)
-next
- case (reduce_let_splitI n v v1 v2 \<Phi> \<delta> x s)
- hence elim: "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let x (AE_split (V_lit (L_bitvec v)) (V_lit (L_num n))) s \<Leftarrow> \<tau> \<and>
- \<Theta> \<turnstile> \<delta> \<sim> \<Delta> \<and> (\<forall>fd\<in>set \<Phi>. check_fundef \<Theta> \<Phi> fd)"
- using config_type_elims by metis
-
- obtain z::x where z: "atom z \<sharp> (AE_split (V_lit (L_bitvec v)) (V_lit (L_num n)), GNil, CE_val (V_lit (L_bitvec (v1 @ v2))),
-([ L_bitvec v1 ]\<^sup>v, [ L_bitvec v2 ]\<^sup>v), \<Theta>, {||}::bv fset)"
- using obtain_fresh by metis
-
- have *:"\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f GNil" using check_s_wf elim by auto
-
- have "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let x (AE_val (V_pair (V_lit (L_bitvec v1)) (V_lit (L_bitvec v2)))) s \<Leftarrow> \<tau>" proof(rule subtype_let)
-
- show \<open> \<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let x (AE_split (V_lit (L_bitvec v)) (V_lit (L_num n))) s \<Leftarrow> \<tau>\<close> using elim by auto
- show \<open>\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> (AE_split (V_lit (L_bitvec v)) (V_lit (L_num n))) \<Rightarrow> \<lbrace> z : B_pair B_bitvec B_bitvec
- | ((CE_val (V_lit (L_bitvec v))) == (CE_concat (CE_fst (CE_val (V_var z))) (CE_snd (CE_val (V_var z)))))
- AND (((CE_len (CE_fst (CE_val (V_var z))))) == (CE_val (V_lit (L_num n)))) \<rbrace> \<close>
- (is "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> ?e1 \<Rightarrow> ?t1")
- proof
- show \<open> \<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using check_s_wf elim by auto
- show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using check_s_wf elim by auto
- show \<open> \<Theta> ; {||} ; GNil \<turnstile> V_lit (L_bitvec v) \<Rightarrow> \<lbrace> z : B_bitvec | CE_val (V_var z) == CE_val (V_lit (L_bitvec v)) \<rbrace>\<close>
- using infer_v_litI infer_l.intros \<open>\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f GNil\<close> fresh_GNil by auto
- show "\<Theta> ; {||} ; GNil \<turnstile> ([ L_num
- n ]\<^sup>v) \<Leftarrow> \<lbrace> z : B_int | (([ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e) == ([ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e)) AND [ leq [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e [| [ [ L_bitvec
- v ]\<^sup>v ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrace>" using split_n reduce_let_splitI check_v_num_leq * wfX_wfY by metis
- show \<open>atom z \<sharp> AE_split [ L_bitvec v ]\<^sup>v [ L_num n ]\<^sup>v\<close> using z fresh_Pair by auto
- show \<open>atom z \<sharp> GNil\<close> using z fresh_Pair by auto
- show \<open>atom z \<sharp> AE_split [ L_bitvec v ]\<^sup>v [ L_num n ]\<^sup>v\<close> using z fresh_Pair by auto
- show \<open>atom z \<sharp> GNil\<close> using z fresh_Pair by auto
- show \<open>atom z \<sharp> AE_split [ L_bitvec v ]\<^sup>v [ L_num n ]\<^sup>v\<close> using z fresh_Pair by auto
- show \<open>atom z \<sharp> GNil\<close> using z fresh_Pair by auto
- qed
-
- show \<open>\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AE_val (V_pair (V_lit (L_bitvec v1)) (V_lit (L_bitvec v2))) \<Rightarrow> \<lbrace> z : B_pair B_bitvec B_bitvec | CE_val (V_var z) == CE_val ((V_pair (V_lit (L_bitvec v1)) (V_lit (L_bitvec v2)))) \<rbrace> \<close>
- (is "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> ?e2 \<Rightarrow> ?t2")
- apply(rule infer_e_valI)
- using check_s_wf elim apply metis
- using check_s_wf elim apply metis
- apply(rule infer_v_pair2I)
- using z fresh_prodN apply metis
- using z fresh_GNil fresh_prodN apply metis
- using infer_v_litI infer_l.intros \<open>\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f GNil\<close> b_of.simps apply blast+
- using b_of.simps apply simp+
- done
- show \<open>\<Theta> ; {||} ; GNil \<turnstile> ?t2 \<lesssim> ?t1\<close> using subtype_split check_s_wf elim reduce_let_splitI by auto
- qed
-
- thus ?case using config_typeI elim by (meson order_refl)
-next
- case (reduce_assert1I \<Phi> \<delta> c v)
-
- hence elim: "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_assert c [v]\<^sup>s \<Leftarrow> \<tau> \<and>
- \<Theta> \<turnstile> \<delta> \<sim> \<Delta> \<and> (\<forall>fd\<in>set \<Phi>. check_fundef \<Theta> \<Phi> fd)"
- using config_type_elims reduce_assert1I by metis
- hence *:"\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_assert c [v]\<^sup>s \<Leftarrow> \<tau>" by auto
-
- have "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> [v]\<^sup>s \<Leftarrow> \<tau>" using check_assert_s * by metis
- thus ?case using elim config_typeI by blast
-next
- case (reduce_assert2I \<Phi> \<delta> s \<delta>' s' c)
-
- hence elim: "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_assert c s \<Leftarrow> \<tau> \<and>
- \<Theta> \<turnstile> \<delta> \<sim> \<Delta> \<and> (\<forall>fd\<in>set \<Phi>. check_fundef \<Theta> \<Phi> fd)"
- using config_type_elims by metis
- hence *:"\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_assert c s \<Leftarrow> \<tau>" by auto
-
- have cv: "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> s \<Leftarrow> \<tau> \<and> \<Theta> ; {||} ; GNil \<Turnstile> c " using check_assert_s * by metis
-
- hence "\<Theta>; \<Phi>; \<Delta> \<turnstile> \<langle>\<delta>, s\<rangle> \<Leftarrow> \<tau>" using elim config_typeI by simp
- then obtain \<Delta>' where D: "\<Theta>; \<Phi>; \<Delta>' \<turnstile> \<langle> \<delta>' , s' \<rangle> \<Leftarrow> \<tau> \<and> \<Delta> \<sqsubseteq> \<Delta>'" using reduce_assert2I by metis
- hence **:"\<Theta>; \<Phi>; {||}; GNil; \<Delta>' \<turnstile> s' \<Leftarrow> \<tau> \<and> \<Theta> \<turnstile> \<delta>' \<sim> \<Delta>'" using config_type_elims by metis
-
- obtain x::x where x:"atom x \<sharp> (\<Theta>, \<Phi>, ({||}::bv fset), GNil, \<Delta>', c, \<tau>, s')" using obtain_fresh by metis
-
- have *:"\<Theta>; \<Phi>; {||}; GNil; \<Delta>' \<turnstile> AS_assert c s' \<Leftarrow> \<tau>" proof
- show "atom x \<sharp> (\<Theta>, \<Phi>, {||}, GNil, \<Delta>', c, \<tau>, s')" using x by auto
- have "\<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f c" using * check_s_wf by auto
- hence wfg:"\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f (x, B_bool, c) #\<^sub>\<Gamma> GNil" using wfC_wfG wfB_boolI check_s_wf * fresh_GNil by auto
- moreover have cs: "\<Theta>; \<Phi>; {||}; GNil; \<Delta>' \<turnstile> s' \<Leftarrow> \<tau>" using ** by auto
- ultimately show "\<Theta> ; \<Phi> ; {||} ; (x, B_bool, c) #\<^sub>\<Gamma> GNil ; \<Delta>' \<turnstile> s' \<Leftarrow> \<tau>" using check_s_g_weakening(1)[OF cs _ wfg] toSet.simps by simp
- show "\<Theta> ; {||} ; GNil \<Turnstile> c " using cv by auto
- show "\<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f \<Delta>' " using check_s_wf ** by auto
- qed
-
- thus ?case using elim config_typeI D ** by metis
-qed
-
-lemma preservation_many:
- assumes " \<Phi> \<turnstile> \<langle>\<delta>, s\<rangle> \<longrightarrow>\<^sup>* \<langle> \<delta>' , s' \<rangle>"
- shows "\<Theta>; \<Phi>; \<Delta> \<turnstile> \<langle>\<delta>, s\<rangle> \<Leftarrow> \<tau> \<Longrightarrow> \<exists>\<Delta>'. \<Theta>; \<Phi>; \<Delta>' \<turnstile> \<langle> \<delta>' , s' \<rangle> \<Leftarrow> \<tau> \<and> \<Delta> \<sqsubseteq> \<Delta>'"
- using assms proof(induct arbitrary: \<Delta> rule: reduce_stmt_many.induct)
- case (reduce_stmt_many_oneI \<Phi> \<delta> s \<delta>' s')
- then show ?case using preservation by simp
-next
- case (reduce_stmt_many_manyI \<Phi> \<delta> s \<delta>' s' \<delta>'' s'')
- then show ?case using preservation subset_trans by metis
-qed
-
-section \<open>Progress\<close>
-text \<open>We prove that a well typed program is either a value or we can make a step\<close>
-
-lemma check_let_op_infer:
- assumes "\<Theta>; \<Phi>; {||}; \<Gamma>; \<Delta> \<turnstile> LET x = (AE_op opp v1 v2) IN s \<Leftarrow> \<tau>" and "supp ( LET x = (AE_op opp v1 v2) IN s) \<subseteq> atom`fst`setD \<Delta>"
- shows "\<exists> z b c. \<Theta>; \<Phi>; {||}; \<Gamma>; \<Delta> \<turnstile> (AE_op opp v1 v2) \<Rightarrow> \<lbrace>z:b|c\<rbrace>"
-proof -
- have xx: "\<Theta>; \<Phi>; {||}; \<Gamma>; \<Delta> \<turnstile> LET x = (AE_op opp v1 v2) IN s \<Leftarrow> \<tau>" using assms by simp
- then show ?thesis using check_s_elims(2)[OF xx] by meson
-qed
-
-lemma infer_pair:
- assumes "\<Theta> ; B; \<Gamma> \<turnstile> v \<Rightarrow> \<lbrace> z : B_pair b1 b2 | c \<rbrace>" and "supp v = {}"
- obtains v1 and v2 where "v = V_pair v1 v2"
- using assms proof(nominal_induct v rule: v.strong_induct)
- case (V_lit x)
- then show ?case by auto
-next
- case (V_var x)
- then show ?case using v.supp supp_at_base by auto
-next
- case (V_pair x1a x2a)
- then show ?case by auto
-next
- case (V_cons x1a x2a x3)
- then show ?case by auto
-next
- case (V_consp x1a x2a x3 x4)
- then show ?case by auto
-qed
-
-lemma progress_fst:
- assumes "\<Theta>; \<Phi>; {||}; \<Gamma>; \<Delta> \<turnstile> LET x = (AE_fst v) IN s \<Leftarrow> \<tau>" and "\<Theta> \<turnstile> \<delta> \<sim> \<Delta>" and
- "supp (LET x = (AE_fst v) IN s) \<subseteq> atom`fst`setD \<Delta>"
- shows "\<exists>\<delta>' s'. \<Phi> \<turnstile> \<langle> \<delta> , LET x = (AE_fst v) IN s \<rangle> \<longrightarrow> \<langle>\<delta>', s'\<rangle>"
-proof -
- have *:"supp v = {}" using assms s_branch_s_branch_list.supp by auto
- obtain z and b and c where "\<Theta>; \<Phi>; {||}; \<Gamma>; \<Delta> \<turnstile> (AE_fst v ) \<Rightarrow> \<lbrace> z : b | c \<rbrace>"
- using check_s_elims(2) using assms by meson
- moreover obtain z' and b' and c' where "\<Theta> ; {||} ; \<Gamma> \<turnstile> v \<Rightarrow> \<lbrace> z' : B_pair b b' | c' \<rbrace>"
- using infer_e_elims(8) using calculation by auto
- moreover then obtain v1 and v2 where "V_pair v1 v2 = v"
- using * infer_pair by metis
- ultimately show ?thesis using reduce_let_fstI assms by metis
-qed
-
-lemma progress_let:
- assumes "\<Theta>; \<Phi>; {||}; \<Gamma>; \<Delta> \<turnstile> LET x = e IN s \<Leftarrow> \<tau>" and "\<Theta> \<turnstile> \<delta> \<sim> \<Delta>" and
- "supp (LET x = e IN s) \<subseteq> atom ` fst ` setD \<Delta>" and "sble \<Theta> \<Gamma>"
- shows "\<exists>\<delta>' s'. \<Phi> \<turnstile> \<langle> \<delta> , LET x = e IN s\<rangle> \<longrightarrow> \<langle> \<delta>' , s'\<rangle>"
-proof -
- obtain z b c where *: "\<Theta>; \<Phi>; {||}; \<Gamma>; \<Delta> \<turnstile> e \<Rightarrow> \<lbrace> z : b | c \<rbrace> " using check_s_elims(2)[OF assms(1)] by metis
- have **: "supp e \<subseteq> atom ` fst ` setD \<Delta>" using assms s_branch_s_branch_list.supp by auto
- from * ** assms show ?thesis proof(nominal_induct "\<lbrace> z : b | c \<rbrace>" rule: infer_e.strong_induct)
- case (infer_e_valI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v)
- then show ?case using reduce_stmt_elims reduce_let_valI by metis
- next
- case (infer_e_plusI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v1 z1 c1 v2 z2 c2 z3)
- hence vf: "supp v1 = {} \<and> supp v2 = {}" by force
- then obtain n1 and n2 where *: "v1 = V_lit (L_num n1) \<and> v2 = (V_lit (L_num n2))" using infer_int infer_e_plusI by metis
- then show ?case using reduce_let_plusI * by metis
- next
- case (infer_e_leqI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v1 z1 c1 v2 z2 c2 z3)
- hence vf: "supp v1 = {} \<and> supp v2 = {}" by force
- then obtain n1 and n2 where *: "v1 = V_lit (L_num n1) \<and> v2 = (V_lit (L_num n2))" using infer_int infer_e_leqI by metis
- then show ?case using reduce_let_leqI * by metis
- next
- case (infer_e_eqI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v1 z1 bb c1 v2 z2 c2 z3)
- hence vf: "supp v1 = {} \<and> supp v2 = {}" by force
- then obtain n1 and n2 where *: "v1 = V_lit n1 \<and> v2 = (V_lit n2)" using infer_lit infer_e_eqI by metis
- then show ?case using reduce_let_eqI by blast
- next
- case (infer_e_appI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> f x b c \<tau>' s' v)
- then show ?case using reduce_let_appI by metis
- next
- case (infer_e_appPI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> b' f bv x b c \<tau>' s' v)
- then show ?case using reduce_let_appPI by metis
- next
- case (infer_e_fstI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v z' b2 c z)
- hence "supp v = {}" by force
- then obtain v1 and v2 where "v = V_pair v1 v2" using infer_e_fstI infer_pair by metis
- then show ?case using reduce_let_fstI * by metis
- next
- case (infer_e_sndI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v z' b1 c z)
- hence "supp v = {}" by force
- then obtain v1 and v2 where "v = V_pair v1 v2" using infer_e_sndI infer_pair by metis
- then show ?case using reduce_let_sndI * by metis
- next
- case (infer_e_lenI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v z' c za)
- hence "supp v = {}" by force
- then obtain bvec where "v = V_lit (L_bitvec bvec)" using infer_e_lenI infer_bitvec by metis
- then show ?case using reduce_let_lenI * by metis
- next
- case (infer_e_mvarI \<Theta> \<B> \<Gamma> \<Phi> \<Delta> u)
- hence "(u, \<lbrace> z : b | c \<rbrace>) \<in> setD \<Delta>" using infer_e_elims(10) by meson
- then obtain v where "(u,v) \<in> set \<delta>" using infer_e_mvarI delta_sim_delta_lookup by meson
- then show ?case using reduce_let_mvar by metis
- next
- case (infer_e_concatI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v1 z1 c1 v2 z2 c2 z3)
- hence vf: "supp v1 = {} \<and> supp v2 = {}" by force
- then obtain n1 and n2 where *: "v1 = V_lit (L_bitvec n1) \<and> v2 = (V_lit (L_bitvec n2))" using infer_bitvec infer_e_concatI by metis
- then show ?case using reduce_let_concatI * by metis
- next
- case (infer_e_splitI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v1 z1 c1 v2 z2 z3)
- hence vf: "supp v1 = {} \<and> supp v2 = {}" by force
- then obtain n1 and n2 where *: "v1 = V_lit (L_bitvec n1) \<and> v2 = (V_lit (L_num n2))" using infer_bitvec infer_e_splitI check_int by metis
-
- have "0 \<le> n2 \<and> n2 \<le> int (length n1)" using check_v_range[OF _ * ] infer_e_splitI by simp
- then obtain bv1 and bv2 where "split n2 n1 (bv1 , bv2)" using obtain_split by metis
- then show ?case using reduce_let_splitI * by metis
- qed
-qed
-
-lemma check_css_lookup_branch_exist:
- fixes s::s and cs::branch_s and css::branch_list and v::v
- shows
- "\<Theta>; \<Phi>; B; G; \<Delta> \<turnstile> s \<Leftarrow> \<tau> \<Longrightarrow> True" and
- "check_branch_s \<Theta> \<Phi> {||} GNil \<Delta> tid dc const v cs \<tau> \<Longrightarrow> True" and
- "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta>; tid; dclist; v \<turnstile> css \<Leftarrow> \<tau> \<Longrightarrow> (dc, t) \<in> set dclist \<Longrightarrow>
- \<exists>x' s'. Some (AS_branch dc x' s') = lookup_branch dc css"
-proof(nominal_induct \<tau> and \<tau> and \<tau> rule: check_s_check_branch_s_check_branch_list.strong_induct)
- case (check_branch_list_consI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid cons const v cs \<tau> dclist css)
- then show ?case using lookup_branch.simps check_branch_list_finalI by force
-next
- case (check_branch_list_finalI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid cons const v cs \<tau>)
- then show ?case using lookup_branch.simps check_branch_list_finalI by force
-qed(auto+)
-
-lemma progress_aux:
- shows "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> s \<Leftarrow> \<tau> \<Longrightarrow> \<B> = {||} \<Longrightarrow> sble \<Theta> \<Gamma> \<Longrightarrow> supp s \<subseteq> atom ` fst ` setD \<Delta> \<Longrightarrow> \<Theta> \<turnstile> \<delta> \<sim> \<Delta> \<Longrightarrow>
- (\<exists>v. s = [v]\<^sup>s) \<or> (\<exists>\<delta>' s'. \<Phi> \<turnstile> \<langle>\<delta>, s\<rangle> \<longrightarrow> \<langle>\<delta>', s'\<rangle>)" and
- "\<Theta>; \<Phi>; {||}; \<Gamma>; \<Delta>; tid; dc; const; v2 \<turnstile> cs \<Leftarrow> \<tau> \<Longrightarrow> supp cs = {} \<Longrightarrow> True "
- "\<Theta>; \<Phi>; {||}; \<Gamma>; \<Delta>; tid; dclist; v2 \<turnstile> css \<Leftarrow> \<tau> \<Longrightarrow> supp css = {} \<Longrightarrow> True "
-proof(induct rule: check_s_check_branch_s_check_branch_list.inducts)
- case (check_valI \<Delta> \<Theta> \<Gamma> v \<tau>' \<tau>)
- then show ?case by auto
-next
- case (check_letI x \<Theta> \<Phi> \<B> \<Gamma> \<Delta> e \<tau> z s b c)
- hence "\<Theta>; \<Phi>; {||}; \<Gamma>; \<Delta> \<turnstile> AS_let x e s \<Leftarrow> \<tau>" using Typing.check_letI by meson
- then show ?case using progress_let check_letI by metis
-next
- case (check_branch_s_branchI \<Theta> \<B> \<Gamma> \<Delta> \<tau> const x \<Phi> tid cons v s)
- then show ?case by auto
-next
- case (check_branch_list_consI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist v cs \<tau> css)
- then show ?case by auto
-next
- case (check_branch_list_finalI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist v cs \<tau>)
- then show ?case by auto
-next
- case (check_ifI z \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v s1 s2 \<tau>)
- have "supp v = {}" using check_ifI s_branch_s_branch_list.supp by auto
- hence "v = V_lit L_true \<or> v = V_lit L_false" using check_bool_options check_ifI by auto
- then show ?case using reduce_if_falseI reduce_if_trueI check_ifI by meson
-next
- case (check_let2I x \<Theta> \<Phi> \<B> G \<Delta> t s1 \<tau> s2 )
- then consider " (\<exists>v. s1 = AS_val v)" | "(\<exists>\<delta>' a. \<Phi> \<turnstile> \<langle>\<delta>, s1\<rangle> \<longrightarrow> \<langle>\<delta>', a\<rangle>)" by auto
- then show ?case proof(cases)
- case 1
- then show ?thesis using reduce_let2_valI by fast
- next
- case 2
- then show ?thesis using reduce_let2I check_let2I by meson
- qed
-next
- case (check_varI u \<Theta> \<Phi> \<B> \<Gamma> \<Delta> \<tau>' v \<tau> s)
-
- obtain uu::u where uf: "atom uu \<sharp> (u,\<delta>,s) " using obtain_fresh by blast
- obtain sa where " (uu \<leftrightarrow> u ) \<bullet> s = sa" by presburger
- moreover have "atom uu \<sharp> s" using uf fresh_prod3 by auto
- ultimately have "AS_var uu \<tau>' v sa = AS_var u \<tau>' v s" using s_branch_s_branch_list.eq_iff(7) Abs1_eq_iff(3)[of uu sa u s] by auto
-
- moreover have "atom uu \<sharp> \<delta>" using uf fresh_prod3 by auto
- ultimately have "\<Phi> \<turnstile> \<langle>\<delta>, AS_var u \<tau>' v s\<rangle> \<longrightarrow> \<langle>(uu, v) # \<delta>, sa\<rangle>"
- using reduce_varI uf by metis
- then show ?case by auto
-next
- case (check_assignI \<Delta> u \<tau> P G v z \<tau>')
- then show ?case using reduce_assignI by blast
-next
- case (check_whileI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> s1 z s2 \<tau>')
- obtain x::x where "atom x \<sharp> (s1,s2)" using obtain_fresh by metis
- moreover obtain z::x where "atom z \<sharp> x" using obtain_fresh by metis
- ultimately show ?case using reduce_whileI by fast
-next
- case (check_seqI P \<Phi> \<B> G \<Delta> s1 z s2 \<tau>)
- thus ?case proof(cases "\<exists>v. s1 = AS_val v")
- case True
- then obtain v where v: "s1 = AS_val v" by blast
- hence "supp v = {}" using check_seqI by auto
- have "\<exists>z1 c1. P; \<B>; G \<turnstile> v \<Rightarrow> (\<lbrace> z1 : B_unit | c1 \<rbrace>)" proof -
- obtain t where t:"P; \<B>; G \<turnstile> v \<Rightarrow> t \<and> P; \<B> ; G \<turnstile> t \<lesssim> (\<lbrace> z : B_unit | TRUE \<rbrace>)"
- using v check_seqI(1) check_s_elims(1) by blast
- obtain z1 and b1 and c1 where teq: "t = (\<lbrace> z1 : b1 | c1 \<rbrace>)" using obtain_fresh_z by meson
- hence "b1 = B_unit" using subtype_eq_base t by meson
- thus ?thesis using t teq by fast
- qed
- then obtain z1 and c1 where "P ; \<B> ; G \<turnstile> v \<Rightarrow> (\<lbrace> z1 : B_unit | c1 \<rbrace>)" by auto
- hence "v = V_lit L_unit" using infer_v_unit_form \<open>supp v = {}\<close> by simp
- hence "s1 = AS_val (V_lit L_unit)" using v by auto
- then show ?thesis using check_seqI reduce_seq1I by meson
- next
- case False
- then show ?thesis using check_seqI reduce_seq2I
- by (metis Un_subset_iff s_branch_s_branch_list.supp(9))
- qed
-
-next
- case (check_caseI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist v cs \<tau> z)
- hence "supp v = {}" by auto
-
- then obtain v' and dc and t::\<tau> where v: "v = V_cons tid dc v' \<and> (dc, t) \<in> set dclist"
- using check_v_tid_form check_caseI by metis
- obtain z and b and c where teq: "t = (\<lbrace> z : b | c \<rbrace>)" using obtain_fresh_z by meson
-
- moreover then obtain x' s' where "Some (AS_branch dc x' s') = lookup_branch dc cs" using v teq check_caseI check_css_lookup_branch_exist by metis
- ultimately show ?case using reduce_caseI v check_caseI dc_of.cases by metis
-next
- case (check_assertI x \<Theta> \<Phi> \<B> \<Gamma> \<Delta> c \<tau> s)
- hence sps: "supp s \<subseteq> atom ` fst ` setD \<Delta>" by auto
- have "atom x \<sharp> c " using check_assertI by auto
- have "atom x \<sharp> \<Gamma>" using check_assertI check_s_wf wfG_elims by metis
- have "sble \<Theta> ((x, B_bool, c) #\<^sub>\<Gamma> \<Gamma>)" proof -
- obtain i' where i':" i' \<Turnstile> \<Gamma> \<and> \<Theta>; \<Gamma> \<turnstile> i'" using check_assertI sble_def by metis
- obtain i::valuation where i:"i = i' ( x \<mapsto> SBool True)" by auto
-
- have "i \<Turnstile> (x, B_bool, c) #\<^sub>\<Gamma> \<Gamma>" proof -
- have "i' \<Turnstile> c" using valid.simps i' check_assertI by metis
- hence "i \<Turnstile> c" using is_satis_weakening_x i \<open>atom x \<sharp> c\<close> by auto
- moreover have "i \<Turnstile> \<Gamma>" using is_satis_g_weakening_x i' i check_assertI \<open>atom x \<sharp> \<Gamma>\<close> by metis
- ultimately show ?thesis using is_satis_g.simps i by auto
- qed
- moreover have "\<Theta> ; ((x, B_bool, c) #\<^sub>\<Gamma> \<Gamma>) \<turnstile> i" proof(rule wfI_cons)
- show \<open> i' \<Turnstile> \<Gamma> \<close> using i' by auto
- show \<open> \<Theta> ; \<Gamma> \<turnstile> i'\<close> using i' by auto
- show \<open>i = i'(x \<mapsto> SBool True)\<close> using i by auto
- show \<open> \<Theta> \<turnstile> SBool True: B_bool\<close> using wfRCV_BBoolI by auto
- show \<open>atom x \<sharp> \<Gamma>\<close> using check_assertI check_s_wf wfG_elims by auto
- qed
- ultimately show ?thesis using sble_def by auto
- qed
- then consider "(\<exists>v. s = [v]\<^sup>s)" | "(\<exists>\<delta>' a. \<Phi> \<turnstile> \<langle>\<delta>, s\<rangle> \<longrightarrow> \<langle> \<delta>', a\<rangle>)" using check_assertI sps by metis
- hence "(\<exists>\<delta>' a. \<Phi> \<turnstile> \<langle>\<delta>, ASSERT c IN s\<rangle> \<longrightarrow> \<langle>\<delta>', a\<rangle>)" proof(cases)
- case 1
- then show ?thesis using reduce_assert1I by metis
- next
- case 2
- then show ?thesis using reduce_assert2I by metis
- qed
- thus ?case by auto
-qed
-
-lemma progress:
- assumes "\<Theta>; \<Phi>; \<Delta> \<turnstile> \<langle>\<delta>, s\<rangle> \<Leftarrow> \<tau>"
- shows "(\<exists>v. s = [v]\<^sup>s) \<or> (\<exists>\<delta>' s'. \<Phi> \<turnstile> \<langle>\<delta>, s\<rangle> \<longrightarrow> \<langle>\<delta>', s'\<rangle>)"
-proof -
- have "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> s \<Leftarrow> \<tau>" and "\<Theta> \<turnstile> \<delta> \<sim> \<Delta>"
- using config_type_elims[OF assms(1)] by auto+
- moreover hence "supp s \<subseteq> atom ` fst ` setD \<Delta>" using check_s_wf wfS_supp by fastforce
- moreover have "sble \<Theta> GNil" using sble_def wfI_def is_satis_g.simps by simp
- ultimately show ?thesis using progress_aux by blast
-qed
-
-section \<open>Safety\<close>
-
-lemma safety_stmt:
- assumes "\<Phi> \<turnstile> \<langle>\<delta>, s\<rangle> \<longrightarrow>\<^sup>* \<langle>\<delta>', s'\<rangle>" and "\<Theta>; \<Phi>; \<Delta> \<turnstile> \<langle>\<delta>, s\<rangle> \<Leftarrow> \<tau>"
- shows "(\<exists>v. s' = [v]\<^sup>s) \<or> (\<exists>\<delta>'' s''. \<Phi> \<turnstile> \<langle>\<delta>', s'\<rangle> \<longrightarrow> \<langle>\<delta>'', s''\<rangle>)"
- using preservation_many progress assms by meson
-
-lemma safety:
- assumes "\<turnstile> \<langle>PROG \<Theta> \<Phi> \<G> s\<rangle> \<Leftarrow> \<tau>" and "\<Phi> \<turnstile> \<langle>\<delta>_of \<G>, s\<rangle> \<longrightarrow>\<^sup>* \<langle>\<delta>', s'\<rangle>"
- shows "(\<exists>v. s' = [v]\<^sup>s) \<or> (\<exists>\<delta>'' s''. \<Phi> \<turnstile> \<langle>\<delta>', s'\<rangle> \<longrightarrow> \<langle>\<delta>'', s''\<rangle>)"
- using assms config_type_prog_elims safety_stmt by metis
-
+(*<*)
+theory Safety
+ imports Operational IVSubstTypingL BTVSubstTypingL
+begin
+ (*>*)
+
+method supp_calc = (metis (mono_tags, opaque_lifting) pure_supp c.supp e.supp v.supp supp_l_empty opp.supp sup_bot.right_neutral supp_at_base)
+declare infer_e.intros[simp]
+declare infer_e.intros[intro]
+
+chapter \<open>Safety\<close>
+
+text \<open>Lemmas about the operational semantics leading up to progress and preservation and then
+safety.\<close>
+
+section \<open>Store Lemmas\<close>
+
+abbreviation delta_ext (" _ \<sqsubseteq> _ ") where
+ "delta_ext \<Delta> \<Delta>' \<equiv> (setD \<Delta> \<subseteq> setD \<Delta>')"
+
+nominal_function dc_of :: "branch_s \<Rightarrow> string" where
+ "dc_of (AS_branch dc _ _) = dc"
+ apply(auto,simp add: eqvt_def dc_of_graph_aux_def)
+ using s_branch_s_branch_list.exhaust by metis
+nominal_termination (eqvt) by lexicographic_order
+
+lemma delta_sim_fresh:
+ assumes "\<Theta> \<turnstile> \<delta> \<sim> \<Delta>" and "atom u \<sharp> \<delta>"
+ shows "atom u \<sharp> \<Delta>"
+ using assms proof(induct rule : delta_sim.inducts)
+ case (delta_sim_nilI \<Theta>)
+ then show ?case using fresh_def supp_DNil by blast
+next
+ case (delta_sim_consI \<Theta> \<delta> \<Delta> v \<tau> u')
+ hence "\<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f \<tau>" using check_v_wf by meson
+ hence "supp \<tau> = {}" using wfT_supp by fastforce
+ moreover have "atom u \<sharp> u'" using delta_sim_consI fresh_Cons fresh_Pair by blast
+ moreover have "atom u \<sharp> \<Delta>" using delta_sim_consI fresh_Cons by blast
+ ultimately show ?case using fresh_Pair fresh_DCons fresh_def by blast
+qed
+
+lemma delta_sim_v:
+ fixes \<Delta>::\<Delta>
+ assumes "\<Theta> \<turnstile> \<delta> \<sim> \<Delta> " and "(u,v) \<in> set \<delta>" and "(u,\<tau>) \<in> setD \<Delta>" and "\<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f \<Delta>"
+ shows "\<Theta> ; {||} ; GNil \<turnstile> v \<Leftarrow> \<tau>"
+ using assms proof(induct \<delta> arbitrary: \<Delta>)
+ case Nil
+ then show ?case by auto
+next
+ case (Cons uv \<delta>)
+ obtain u' and v' where uv : "uv=(u',v')" by fastforce
+ show ?case proof(cases "u'=u")
+ case True
+ hence *:"\<Theta> \<turnstile> ((u,v')#\<delta>) \<sim> \<Delta>" using uv Cons by blast
+ then obtain \<tau>' and \<Delta>' where tt: "\<Theta> ; {||} ; GNil \<turnstile> v' \<Leftarrow> \<tau>' \<and> u \<notin> fst ` set \<delta> \<and> \<Delta> = (u,\<tau>')#\<^sub>\<Delta>\<Delta>'" using delta_sim_elims(3)[OF *] by metis
+ moreover hence "v'=v" using Cons True
+ by (metis Pair_inject fst_conv image_eqI set_ConsD uv)
+ moreover have "\<tau>=\<tau>'" using wfD_unique tt Cons
+ setD.simps list.set_intros by blast
+ ultimately show ?thesis by metis
+ next
+ case False
+ hence *:"\<Theta> \<turnstile> ((u',v')#\<delta>) \<sim> \<Delta>" using uv Cons by blast
+ then obtain \<tau>' and \<Delta>' where tt: "\<Theta> \<turnstile> \<delta> \<sim> \<Delta>' \<and> \<Theta> ; {||} ; GNil \<turnstile> v' \<Leftarrow> \<tau>' \<and> u' \<notin> fst ` set \<delta> \<and> \<Delta> = (u',\<tau>')#\<^sub>\<Delta>\<Delta>'" using delta_sim_elims(3)[OF *] by metis
+
+ moreover hence "\<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f \<Delta>'" using wfD_elims Cons delta_sim_elims by metis
+ ultimately show ?thesis using Cons
+ using False by auto
+ qed
+qed
+
+lemma delta_sim_delta_lookup:
+ assumes "\<Theta> \<turnstile> \<delta> \<sim> \<Delta> " and "(u, \<lbrace> z : b | c \<rbrace>) \<in> setD \<Delta>"
+ shows "\<exists>v. (u,v) \<in> set \<delta>"
+ using assms by(induct rule: delta_sim.inducts,auto+)
+
+lemma update_d_stable:
+ "fst ` set \<delta> = fst ` set (update_d \<delta> u v)"
+proof(induct \<delta>)
+ case Nil
+ then show ?case by auto
+next
+ case (Cons a \<delta>)
+ then show ?case using update_d.simps
+ by (metis (no_types, lifting) eq_fst_iff image_cong image_insert list.simps(15) prod.exhaust_sel)
+qed
+
+lemma update_d_sim:
+ fixes \<Delta>::\<Delta>
+ assumes "\<Theta> \<turnstile> \<delta> \<sim> \<Delta>" and "\<Theta> ; {||} ; GNil \<turnstile> v \<Leftarrow> \<tau>" and "(u,\<tau>) \<in> setD \<Delta>" and "\<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f \<Delta>"
+ shows "\<Theta> \<turnstile> (update_d \<delta> u v) \<sim> \<Delta>"
+ using assms proof(induct \<delta> arbitrary: \<Delta>)
+ case Nil
+ then show ?case using delta_sim_consI by simp
+next
+ case (Cons uv \<delta>)
+ obtain u' and v' where uv : "uv=(u',v')" by fastforce
+
+ hence *:"\<Theta> \<turnstile> ((u',v')#\<delta>) \<sim> \<Delta>" using uv Cons by blast
+ then obtain \<tau>' and \<Delta>' where tt: "\<Theta> \<turnstile> \<delta> \<sim> \<Delta>' \<and> \<Theta> ; {||} ; GNil \<turnstile> v' \<Leftarrow> \<tau>' \<and> u' \<notin> fst ` set \<delta> \<and> \<Delta> = (u',\<tau>')#\<^sub>\<Delta>\<Delta>'" using delta_sim_elims * by metis
+
+ show ?case proof(cases "u=u'")
+ case True
+ then have "(u,\<tau>') \<in> setD \<Delta>" using tt by auto
+ then have "\<tau> = \<tau>'" using Cons wfD_unique by metis
+ moreover have "update_d ((u',v')#\<delta>) u v = ((u',v)#\<delta>)" using update_d.simps True by presburger
+ ultimately show ?thesis using delta_sim_consI tt Cons True
+ by (simp add: tt uv)
+ next
+ case False
+ have "\<Theta> \<turnstile> (u',v') # (update_d \<delta> u v) \<sim> (u',\<tau>')#\<^sub>\<Delta>\<Delta>'"
+ proof(rule delta_sim_consI)
+ show "\<Theta> \<turnstile> update_d \<delta> u v \<sim> \<Delta>' " using Cons using delta_sim_consI
+ delta_sim.simps update_d.simps Cons delta_sim_elims uv tt
+ False fst_conv set_ConsD wfG_elims wfD_elims by (metis setD_ConsD)
+ show "\<Theta> ; {||} ; GNil \<turnstile> v' \<Leftarrow> \<tau>'" using tt by auto
+ show "u' \<notin> fst ` set (update_d \<delta> u v)" using update_d.simps Cons update_d_stable tt by auto
+ qed
+ thus ?thesis using False update_d.simps uv
+ by (simp add: tt)
+ qed
+qed
+
+section \<open>Preservation\<close>
+text \<open>Types are preserved under reduction step. Broken down into lemmas about different operations\<close>
+
+subsection \<open>Function Application\<close>
+
+lemma check_s_x_fresh:
+ fixes x::x and s::s
+ assumes "\<Theta> ; \<Phi> ; B ; GNil ; D \<turnstile> s \<Leftarrow> \<tau>"
+ shows "atom x \<sharp> s \<and> atom x \<sharp> \<tau> \<and> atom x \<sharp> D"
+proof -
+ have "\<Theta> ; \<Phi> ; B ; GNil ; D \<turnstile>\<^sub>w\<^sub>f s : b_of \<tau>" using check_s_wf[OF assms] by auto
+ moreover have "\<Theta> ; B ; GNil \<turnstile>\<^sub>w\<^sub>f \<tau> " using check_s_wf assms by auto
+ moreover have "\<Theta> ; B ; GNil \<turnstile>\<^sub>w\<^sub>f D" using check_s_wf assms by auto
+ ultimately show ?thesis using wf_supp x_fresh_u
+ by (meson fresh_GNil wfS_x_fresh wfT_x_fresh wfD_x_fresh)
+qed
+
+lemma check_funtyp_subst_b:
+ fixes b'::b
+ assumes "check_funtyp \<Theta> \<Phi> {|bv|} (AF_fun_typ x b c \<tau> s)" and \<open> \<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f b' \<close>
+ shows "check_funtyp \<Theta> \<Phi> {||} (AF_fun_typ x b[bv::=b']\<^sub>b\<^sub>b (c[bv::=b']\<^sub>c\<^sub>b) \<tau>[bv::=b']\<^sub>\<tau>\<^sub>b s[bv::=b']\<^sub>s\<^sub>b)"
+ using assms proof (nominal_induct "{|bv|}" "AF_fun_typ x b c \<tau> s" rule: check_funtyp.strong_induct)
+ case (check_funtypI x' \<Theta> \<Phi> c' s' \<tau>')
+ have "check_funtyp \<Theta> \<Phi> {||} (AF_fun_typ x' b[bv::=b']\<^sub>b\<^sub>b (c'[bv::=b']\<^sub>c\<^sub>b) \<tau>'[bv::=b']\<^sub>\<tau>\<^sub>b s'[bv::=b']\<^sub>s\<^sub>b)" proof
+ show \<open>atom x' \<sharp> (\<Theta>, \<Phi>, {||}::bv fset, b[bv::=b']\<^sub>b\<^sub>b)\<close> using check_funtypI fresh_prodN x_fresh_b fresh_empty_fset by metis
+
+ have \<open> \<Theta> ; \<Phi> ; {||} ; ((x', b, c') #\<^sub>\<Gamma> GNil)[bv::=b']\<^sub>\<Gamma>\<^sub>b ; []\<^sub>\<Delta>[bv::=b']\<^sub>\<Delta>\<^sub>b \<turnstile> s'[bv::=b']\<^sub>s\<^sub>b \<Leftarrow> \<tau>'[bv::=b']\<^sub>\<tau>\<^sub>b\<close> proof(rule subst_b_check_s)
+ show \<open> \<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f b' \<close> using check_funtypI by metis
+ show \<open>{|bv|} = {|bv|}\<close> by auto
+ show \<open> \<Theta> ; \<Phi> ; {|bv|} ; (x', b, c') #\<^sub>\<Gamma> GNil ; []\<^sub>\<Delta> \<turnstile> s' \<Leftarrow> \<tau>'\<close> using check_funtypI by metis
+ qed
+
+ thus \<open> \<Theta> ; \<Phi> ; {||} ; (x', b[bv::=b']\<^sub>b\<^sub>b, c'[bv::=b']\<^sub>c\<^sub>b) #\<^sub>\<Gamma> GNil ; []\<^sub>\<Delta> \<turnstile> s'[bv::=b']\<^sub>s\<^sub>b \<Leftarrow> \<tau>'[bv::=b']\<^sub>\<tau>\<^sub>b\<close>
+ using subst_gb.simps subst_db.simps by simp
+ qed
+
+ moreover have "(AF_fun_typ x b c \<tau> s) = (AF_fun_typ x' b c' \<tau>' s')" using fun_typ.eq_iff check_funtypI by metis
+ moreover hence "(AF_fun_typ x b[bv::=b']\<^sub>b\<^sub>b (c[bv::=b']\<^sub>c\<^sub>b) \<tau>[bv::=b']\<^sub>\<tau>\<^sub>b s[bv::=b']\<^sub>s\<^sub>b) = (AF_fun_typ x' b[bv::=b']\<^sub>b\<^sub>b (c'[bv::=b']\<^sub>c\<^sub>b) \<tau>'[bv::=b']\<^sub>\<tau>\<^sub>b s'[bv::=b']\<^sub>s\<^sub>b)"
+ using subst_ft_b.simps by metis
+ ultimately show ?case by metis
+qed
+
+lemma funtyp_simple_check:
+ fixes s::s and \<Delta>::\<Delta> and \<tau>::\<tau> and v::v
+ assumes "check_funtyp \<Theta> \<Phi> ({||}::bv fset) (AF_fun_typ x b c \<tau> s)" and
+ "\<Theta> ; {||} ; GNil \<turnstile> v \<Leftarrow> \<lbrace> x : b | c \<rbrace>"
+ shows "\<Theta> ; \<Phi> ; {||} ; GNil ; DNil \<turnstile> s[x::=v]\<^sub>s\<^sub>v \<Leftarrow> \<tau>[x::=v]\<^sub>\<tau>\<^sub>v"
+ using assms proof(nominal_induct " ({||}::bv fset)" "AF_fun_typ x b c \<tau> s" avoiding: v x rule: check_funtyp.strong_induct)
+ case (check_funtypI x' \<Theta> \<Phi> c' s' \<tau>')
+
+ hence eq1: "\<lbrace> x' : b | c' \<rbrace> = \<lbrace> x : b | c \<rbrace>" using funtyp_eq_iff_equalities by metis
+
+ obtain x'' and c'' where xf:"\<lbrace> x : b | c \<rbrace> = \<lbrace> x'' : b | c'' \<rbrace> \<and> atom x'' \<sharp> (x',v) \<and> atom x'' \<sharp> (x,c)" using obtain_fresh_z3 by metis
+ moreover have "atom x' \<sharp> c''" proof -
+ have "supp \<lbrace> x'' : b | c'' \<rbrace> = {}" using eq1 check_funtypI xf check_v_wf wfT_nil_supp by metis
+ hence "supp c'' \<subseteq> { atom x'' }" using \<tau>.supp eq1 xf by (auto simp add: freshers)
+ moreover have "atom x' \<noteq> atom x''" using xf fresh_Pair fresh_x_neq by metis
+ ultimately show ?thesis using xf fresh_Pair fresh_x_neq fresh_def fresh_at_base by blast
+ qed
+ ultimately have eq2: "c''[x''::=[ x' ]\<^sup>v]\<^sub>c\<^sub>v = c'" using eq1 type_eq_subst_eq3(1)[of x' b c' x'' b c''] by metis
+
+ have "atom x' \<sharp> c" proof -
+ have "supp \<lbrace> x : b | c \<rbrace> = {}" using eq1 check_funtypI xf check_v_wf wfT_nil_supp by metis
+ hence "supp c \<subseteq> { atom x }" using \<tau>.supp by auto
+ moreover have "atom x \<noteq> atom x'" using check_funtypI fresh_Pair fresh_x_neq by metis
+ ultimately show ?thesis using fresh_def by force
+ qed
+ hence eq: " c[x::=[ x' ]\<^sup>v]\<^sub>c\<^sub>v = c' \<and> s'[x'::=v]\<^sub>s\<^sub>v = s[x::=v]\<^sub>s\<^sub>v \<and> \<tau>'[x'::=v]\<^sub>\<tau>\<^sub>v = \<tau>[x::=v]\<^sub>\<tau>\<^sub>v"
+ using funtyp_eq_iff_equalities type_eq_subst_eq3 check_funtypI by metis
+
+ have " \<Theta> ; \<Phi> ; {||} ; ((x', b, c''[x''::=[ x' ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> GNil)[x'::=v]\<^sub>\<Gamma>\<^sub>v ; []\<^sub>\<Delta>[x'::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> s'[x'::=v]\<^sub>s\<^sub>v \<Leftarrow> \<tau>'[x'::=v]\<^sub>\<tau>\<^sub>v"
+ proof(rule subst_check_check_s )
+ show \<open>\<Theta> ; {||} ; GNil \<turnstile> v \<Leftarrow> \<lbrace> x'' : b | c'' \<rbrace>\<close> using check_funtypI eq1 xf by metis
+ show \<open>atom x'' \<sharp> (x', v)\<close> using check_funtypI fresh_x_neq fresh_Pair xf by metis
+ show \<open> \<Theta> ; \<Phi> ; {||} ; (x', b, c''[x''::=[ x' ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> GNil ; []\<^sub>\<Delta> \<turnstile> s' \<Leftarrow> \<tau>'\<close> using check_funtypI eq2 by metis
+ show \<open> (x', b, c''[x''::=[ x' ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> GNil = GNil @ (x', b, c''[x''::=[ x' ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> GNil\<close> using append_g.simps by auto
+ qed
+ hence " \<Theta>; \<Phi>; {||}; GNil; []\<^sub>\<Delta> \<turnstile> s'[x'::=v]\<^sub>s\<^sub>v \<Leftarrow> \<tau>'[x'::=v]\<^sub>\<tau>\<^sub>v" using subst_gv.simps subst_dv.simps by auto
+ thus ?case using eq by auto
+qed
+
+lemma funtypq_simple_check:
+ fixes s::s and \<Delta>::\<Delta> and \<tau>::\<tau> and v::v
+ assumes "check_funtypq \<Theta> \<Phi> (AF_fun_typ_none (AF_fun_typ x b c t s))" and
+ "\<Theta> ; {||} ; GNil \<turnstile> v \<Leftarrow> \<lbrace> x : b | c \<rbrace>"
+ shows "\<Theta>; \<Phi>; {||}; GNil; DNil \<turnstile> s[x::=v]\<^sub>s\<^sub>v \<Leftarrow> t[x::=v]\<^sub>\<tau>\<^sub>v"
+ using assms proof(nominal_induct "(AF_fun_typ_none (AF_fun_typ x b c t s))" avoiding: v rule: check_funtypq.strong_induct)
+ case (check_fundefq_simpleI \<Theta> \<Phi> x' c' t' s')
+ hence eq: "\<lbrace> x : b | c \<rbrace> = \<lbrace> x' : b | c' \<rbrace> \<and> s'[x'::=v]\<^sub>s\<^sub>v = s[x::=v]\<^sub>s\<^sub>v \<and> t[x::=v]\<^sub>\<tau>\<^sub>v = t'[x'::=v]\<^sub>\<tau>\<^sub>v"
+ using funtyp_eq_iff_equalities by metis
+ hence "\<Theta>; \<Phi>; {||}; GNil; []\<^sub>\<Delta> \<turnstile> s'[x'::=v]\<^sub>s\<^sub>v \<Leftarrow> t'[x'::=v]\<^sub>\<tau>\<^sub>v"
+ using funtyp_simple_check[OF check_fundefq_simpleI(1)] check_fundefq_simpleI by metis
+ thus ?case using eq by metis
+qed
+
+lemma funtyp_poly_eq_iff_equalities:
+ assumes "[[atom bv']]lst. AF_fun_typ x' b'' c' t' s' = [[atom bv]]lst. AF_fun_typ x b c t s"
+ shows "\<lbrace> x' : b''[bv'::=b']\<^sub>b\<^sub>b | c'[bv'::=b']\<^sub>c\<^sub>b \<rbrace> = \<lbrace> x : b[bv::=b']\<^sub>b\<^sub>b | c[bv::=b']\<^sub>c\<^sub>b \<rbrace> \<and>
+ s'[bv'::=b']\<^sub>s\<^sub>b[x'::=v]\<^sub>s\<^sub>v = s[bv::=b']\<^sub>s\<^sub>b[x::=v]\<^sub>s\<^sub>v \<and> t'[bv'::=b']\<^sub>\<tau>\<^sub>b[x'::=v]\<^sub>\<tau>\<^sub>v = t[bv::=b']\<^sub>\<tau>\<^sub>b[x::=v]\<^sub>\<tau>\<^sub>v"
+proof -
+ have "subst_ft_b (AF_fun_typ x' b'' c' t' s') bv' b' = subst_ft_b (AF_fun_typ x b c t s) bv b'"
+ using subst_b_flip_eq_two subst_b_fun_typ_def assms by metis
+ thus ?thesis using fun_typ.eq_iff subst_ft_b.simps funtyp_eq_iff_equalities subst_tb.simps
+ by (metis (full_types) assms fun_poly_arg_unique)
+
+qed
+
+lemma funtypq_poly_check:
+ fixes s::s and \<Delta>::\<Delta> and \<tau>::\<tau> and v::v and b'::b
+ assumes "check_funtypq \<Theta> \<Phi> (AF_fun_typ_some bv (AF_fun_typ x b c t s))" and
+ "\<Theta> ; {||} ; GNil \<turnstile> v \<Leftarrow> \<lbrace> x : b[bv::=b']\<^sub>b\<^sub>b | c[bv::=b']\<^sub>c\<^sub>b \<rbrace>" and
+ "\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f b'"
+ shows "\<Theta>; \<Phi>; {||}; GNil; DNil \<turnstile> s[bv::=b']\<^sub>s\<^sub>b[x::=v]\<^sub>s\<^sub>v \<Leftarrow> t[bv::=b']\<^sub>\<tau>\<^sub>b[x::=v]\<^sub>\<tau>\<^sub>v"
+ using assms proof(nominal_induct "(AF_fun_typ_some bv (AF_fun_typ x b c t s))" avoiding: v rule: check_funtypq.strong_induct)
+ case (check_funtypq_polyI bv' \<Theta> \<Phi> x' b'' c' t' s')
+
+ hence **:"\<lbrace> x' : b''[bv'::=b']\<^sub>b\<^sub>b | c'[bv'::=b']\<^sub>c\<^sub>b \<rbrace> = \<lbrace> x : b[bv::=b']\<^sub>b\<^sub>b | c[bv::=b']\<^sub>c\<^sub>b \<rbrace> \<and>
+ s'[bv'::=b']\<^sub>s\<^sub>b[x'::=v]\<^sub>s\<^sub>v = s[bv::=b']\<^sub>s\<^sub>b[x::=v]\<^sub>s\<^sub>v \<and> t'[bv'::=b']\<^sub>\<tau>\<^sub>b[x'::=v]\<^sub>\<tau>\<^sub>v = t[bv::=b']\<^sub>\<tau>\<^sub>b[x::=v]\<^sub>\<tau>\<^sub>v"
+ using funtyp_poly_eq_iff_equalities by metis
+
+ have *:"check_funtyp \<Theta> \<Phi> {||} (AF_fun_typ x' b''[bv'::=b']\<^sub>b\<^sub>b (c'[bv'::=b']\<^sub>c\<^sub>b) (t'[bv'::=b']\<^sub>\<tau>\<^sub>b) s'[bv'::=b']\<^sub>s\<^sub>b)"
+ using check_funtyp_subst_b[OF check_funtypq_polyI(5) check_funtypq_polyI(8)] by metis
+ moreover have "\<Theta> ; {||} ; GNil \<turnstile> v \<Leftarrow> \<lbrace> x' : b''[bv'::=b']\<^sub>b\<^sub>b | c'[bv'::=b']\<^sub>c\<^sub>b \<rbrace>" using ** check_funtypq_polyI by metis
+ ultimately have "\<Theta>; \<Phi>; {||}; GNil; []\<^sub>\<Delta> \<turnstile> s'[bv'::=b']\<^sub>s\<^sub>b[x'::=v]\<^sub>s\<^sub>v \<Leftarrow> t'[bv'::=b']\<^sub>\<tau>\<^sub>b[x'::=v]\<^sub>\<tau>\<^sub>v"
+ using funtyp_simple_check[OF *] check_funtypq_polyI by metis
+ thus ?case using ** by metis
+
+qed
+
+lemma fundef_simple_check:
+ fixes s::s and \<Delta>::\<Delta> and \<tau>::\<tau> and v::v
+ assumes "check_fundef \<Theta> \<Phi> (AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c t s)))" and
+ "\<Theta> ; {||} ; GNil \<turnstile> v \<Leftarrow> \<lbrace> x : b | c \<rbrace>" and "\<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f \<Delta>"
+ shows "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> s[x::=v]\<^sub>s\<^sub>v \<Leftarrow> t[x::=v]\<^sub>\<tau>\<^sub>v"
+ using assms proof(nominal_induct "(AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c t s)))" avoiding: v rule: check_fundef.strong_induct)
+ case (check_fundefI \<Theta> \<Phi>)
+ then show ?case using funtypq_simple_check[THEN check_s_d_weakening(1) ] setD.simps by auto
+qed
+
+lemma fundef_poly_check:
+ fixes s::s and \<Delta>::\<Delta> and \<tau>::\<tau> and v::v and b'::b
+ assumes "check_fundef \<Theta> \<Phi> (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c t s)))" and
+ "\<Theta> ; {||} ; GNil \<turnstile> v \<Leftarrow> \<lbrace> x : b[bv::=b']\<^sub>b\<^sub>b | c[bv::=b']\<^sub>c\<^sub>b \<rbrace>" and "\<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f \<Delta>" and "\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f b'"
+ shows "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> s[bv::=b']\<^sub>s\<^sub>b[x::=v]\<^sub>s\<^sub>v \<Leftarrow> t[bv::=b']\<^sub>\<tau>\<^sub>b[x::=v]\<^sub>\<tau>\<^sub>v"
+ using assms proof(nominal_induct "(AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c t s)))" avoiding: v rule: check_fundef.strong_induct)
+ case (check_fundefI \<Theta> \<Phi>)
+ then show ?case using funtypq_poly_check[THEN check_s_d_weakening(1) ] setD.simps by auto
+qed
+
+lemma preservation_app:
+ assumes
+ "Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x1 b1 c1 \<tau>1' s1'))) = lookup_fun \<Phi> f" and "(\<forall>fd\<in>set \<Phi>. check_fundef \<Theta> \<Phi> fd)"
+ shows "\<Theta> ; \<Phi> ; B ; G ; \<Delta> \<turnstile> ss \<Leftarrow> \<tau> \<Longrightarrow> B = {||} \<Longrightarrow> G = GNil \<Longrightarrow> ss = LET x = (AE_app f v) IN s \<Longrightarrow>
+ \<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> LET x : (\<tau>1'[x1::=v]\<^sub>\<tau>\<^sub>v) = (s1'[x1::=v]\<^sub>s\<^sub>v) IN s \<Leftarrow> \<tau>" and
+ "check_branch_s \<Theta> \<Phi> \<B> GNil \<Delta> tid dc const v cs \<tau> \<Longrightarrow> True" and
+ "check_branch_list \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist v css \<tau> \<Longrightarrow> True"
+ using assms proof(nominal_induct \<tau> and \<tau> and \<tau> avoiding: v rule: check_s_check_branch_s_check_branch_list.strong_induct)
+ case (check_letI x2 \<Theta> \<Phi> \<B> \<Gamma> \<Delta> e \<tau> z s2 b c)
+
+ hence eq: "e = (AE_app f v)" by simp
+ hence *:"\<Theta> ; \<Phi> ; {||} ;GNil ; \<Delta> \<turnstile> (AE_app f v) \<Rightarrow> \<lbrace> z : b | c \<rbrace>" using check_letI by auto
+
+ then obtain x3 b3 c3 \<tau>3 s3 where
+ **:"\<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f \<Delta> \<and> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<and> Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x3 b3 c3 \<tau>3 s3))) = lookup_fun \<Phi> f \<and>
+ \<Theta> ; {||} ; GNil \<turnstile> v \<Leftarrow> \<lbrace> x3 : b3 | c3 \<rbrace> \<and> atom x3 \<sharp> (\<Theta>, \<Phi>, ({||}::bv fset), GNil, \<Delta>, v, \<lbrace> z : b | c \<rbrace>) \<and> \<tau>3[x3::=v]\<^sub>\<tau>\<^sub>v = \<lbrace> z : b | c \<rbrace>"
+ using infer_e_elims(6)[OF *] subst_defs by metis
+
+ obtain z3 where z3:"\<lbrace> x3 : b3 | c3 \<rbrace> = \<lbrace> z3 : b3 | c3[x3::=V_var z3]\<^sub>c\<^sub>v \<rbrace> \<and> atom z3 \<sharp> (x3, v,c3,x1,c1)" using obtain_fresh_z3 by metis
+
+ have seq:"[[atom x3]]lst. s3 = [[atom x1]]lst. s1'" using fun_def_eq check_letI ** option.inject by metis
+
+ let ?ft = "AF_fun_typ x3 b3 c3 \<tau>3 s3"
+
+
+ have sup: "supp \<tau>3 \<subseteq> { atom x3} \<and> supp s3 \<subseteq> { atom x3}" using wfPhi_f_supp ** by metis
+
+ have "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let2 x2 \<tau>3[x3::=v]\<^sub>\<tau>\<^sub>v (s3[x3::=v]\<^sub>s\<^sub>v) s2 \<Leftarrow> \<tau>" proof
+ show \<open>atom x2 \<sharp> (\<Theta>, \<Phi>, {||}::bv fset, GNil, \<Delta>, \<tau>3[x3::=v]\<^sub>\<tau>\<^sub>v, s3[x3::=v]\<^sub>s\<^sub>v, \<tau>)\<close>
+ unfolding fresh_prodN using check_letI fresh_subst_v_if subst_v_\<tau>_def sup
+ by (metis all_not_in_conv fresh_def fresh_empty_fset fresh_subst_sv_if fresh_subst_tv_if singleton_iff subset_singleton_iff)
+
+ show \<open> \<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> s3[x3::=v]\<^sub>s\<^sub>v \<Leftarrow> \<tau>3[x3::=v]\<^sub>\<tau>\<^sub>v\<close> proof(rule fundef_simple_check)
+ show \<open>check_fundef \<Theta> \<Phi> (AF_fundef f (AF_fun_typ_none (AF_fun_typ x3 b3 c3 \<tau>3 s3)))\<close> using ** check_letI lookup_fun_member by metis
+ show \<open>\<Theta> ; {||} ; GNil \<turnstile> v \<Leftarrow> \<lbrace> x3 : b3 | c3 \<rbrace>\<close> using ** by auto
+ show \<open> \<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using ** by auto
+ qed
+ show \<open> \<Theta> ; \<Phi> ; {||} ; (x2, b_of \<tau>3[x3::=v]\<^sub>\<tau>\<^sub>v, c_of \<tau>3[x3::=v]\<^sub>\<tau>\<^sub>v x2) #\<^sub>\<Gamma> GNil ; \<Delta> \<turnstile> s2 \<Leftarrow> \<tau>\<close>
+ using check_letI ** b_of.simps c_of.simps subst_defs by metis
+ qed
+
+ moreover have "AS_let2 x2 \<tau>3[x3::=v]\<^sub>\<tau>\<^sub>v (s3[x3::=v]\<^sub>s\<^sub>v) s2 = AS_let2 x (\<tau>1'[x1::=v]\<^sub>\<tau>\<^sub>v) (s1'[x1::=v]\<^sub>s\<^sub>v) s" proof -
+ have *: "[[atom x2]]lst. s2 = [[atom x]]lst. s" using check_letI s_branch_s_branch_list.eq_iff by auto
+ moreover have " \<tau>3[x3::=v]\<^sub>\<tau>\<^sub>v = \<tau>1'[x1::=v]\<^sub>\<tau>\<^sub>v" using fun_ret_unique ** check_letI by metis
+ moreover have "s3[x3::=v]\<^sub>s\<^sub>v = (s1'[x1::=v]\<^sub>s\<^sub>v)" using subst_v_flip_eq_two subst_v_s_def seq by metis
+ ultimately show ?thesis using s_branch_s_branch_list.eq_iff by metis
+ qed
+
+ ultimately show ?case using check_letI by auto
+qed(auto+)
+
+lemma fresh_subst_v_subst_b:
+ fixes x2::x and tm::"'a::{has_subst_v,has_subst_b}" and x::x
+ assumes "supp tm \<subseteq> { atom bv, atom x }" and "atom x2 \<sharp> v"
+ shows "atom x2 \<sharp> tm[bv::=b]\<^sub>b[x::=v]\<^sub>v"
+ using assms proof(cases "x2=x")
+ case True
+ then show ?thesis using fresh_subst_v_if assms by blast
+next
+ case False
+ hence "atom x2 \<sharp> tm" using assms fresh_def fresh_at_base by force
+ hence "atom x2 \<sharp> tm[bv::=b]\<^sub>b" using assms fresh_subst_if x_fresh_b False by force
+ then show ?thesis using fresh_subst_v_if assms by auto
+qed
+
+lemma preservation_poly_app:
+ assumes
+ "Some (AF_fundef f (AF_fun_typ_some bv1 (AF_fun_typ x1 b1 c1 \<tau>1' s1'))) = lookup_fun \<Phi> f" and "(\<forall>fd\<in>set \<Phi>. check_fundef \<Theta> \<Phi> fd)"
+ shows "\<Theta> ; \<Phi> ; B ; G ; \<Delta> \<turnstile> ss \<Leftarrow> \<tau> \<Longrightarrow> B = {||} \<Longrightarrow> G = GNil \<Longrightarrow> ss = LET x = (AE_appP f b' v) IN s \<Longrightarrow> \<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f b' \<Longrightarrow>
+ \<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> LET x : (\<tau>1'[bv1::=b']\<^sub>\<tau>\<^sub>b[x1::=v]\<^sub>\<tau>\<^sub>v) = (s1'[bv1::=b']\<^sub>s\<^sub>b[x1::=v]\<^sub>s\<^sub>v) IN s \<Leftarrow> \<tau>" and
+ "check_branch_s \<Theta> \<Phi> \<B> GNil \<Delta> tid dc const v cs \<tau> \<Longrightarrow> True" and
+ "check_branch_list \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist v css \<tau> \<Longrightarrow> True"
+ using assms proof(nominal_induct \<tau> and \<tau> and \<tau> avoiding: v x1 rule: check_s_check_branch_s_check_branch_list.strong_induct)
+ case (check_letI x2 \<Theta> \<Phi> \<B> \<Gamma> \<Delta> e \<tau> z s2 b c)
+
+ hence eq: "e = (AE_appP f b' v)" by simp
+ hence *:"\<Theta> ; \<Phi> ; {||} ;GNil ; \<Delta> \<turnstile> (AE_appP f b' v) \<Rightarrow> \<lbrace> z : b | c \<rbrace>" using check_letI by auto
+
+ then obtain x3 b3 c3 \<tau>3 s3 bv3 where
+ **:"\<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f \<Delta> \<and> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<and> Some (AF_fundef f (AF_fun_typ_some bv3 (AF_fun_typ x3 b3 c3 \<tau>3 s3))) = lookup_fun \<Phi> f \<and>
+ \<Theta> ; {||} ; GNil \<turnstile> v \<Leftarrow> \<lbrace> x3 : b3[bv3::=b']\<^sub>b\<^sub>b | c3[bv3::=b']\<^sub>c\<^sub>b \<rbrace> \<and> atom x3 \<sharp> GNil \<and> \<tau>3[bv3::=b']\<^sub>\<tau>\<^sub>b[x3::=v]\<^sub>\<tau>\<^sub>v = \<lbrace> z : b | c \<rbrace>
+ \<and> \<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f b'"
+ using infer_e_elims(21)[OF *] subst_defs by metis
+
+ obtain z3 where z3:"\<lbrace> x3 : b3 | c3 \<rbrace> = \<lbrace> z3 : b3 | c3[x3::=V_var z3]\<^sub>c\<^sub>v \<rbrace> \<and> atom z3 \<sharp> (x3, v,c3,x1,c1)" using obtain_fresh_z3 by metis
+
+ let ?ft = "(AF_fun_typ x3 (b3[bv3::=b']\<^sub>b\<^sub>b) (c3[bv3::=b']\<^sub>c\<^sub>b) (\<tau>3[bv3::=b']\<^sub>\<tau>\<^sub>b) (s3[bv3::=b']\<^sub>s\<^sub>b))"
+
+ have *:"check_fundef \<Theta> \<Phi> (AF_fundef f (AF_fun_typ_some bv3 (AF_fun_typ x3 b3 c3 \<tau>3 s3)))" using ** check_letI lookup_fun_member by metis
+
+ hence ftq:"check_funtypq \<Theta> \<Phi> (AF_fun_typ_some bv3 (AF_fun_typ x3 b3 c3 \<tau>3 s3))" using check_fundef_elims by auto
+
+ let ?ft = "AF_fun_typ_some bv3 (AF_fun_typ x3 b3 c3 \<tau>3 s3)"
+
+ have sup: "supp \<tau>3 \<subseteq> { atom x3, atom bv3} \<and> supp s3 \<subseteq> { atom x3, atom bv3 }"
+ using wfPhi_f_poly_supp_t wfPhi_f_poly_supp_s ** by metis
+
+ have "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let2 x2 \<tau>3[bv3::=b']\<^sub>\<tau>\<^sub>b[x3::=v]\<^sub>\<tau>\<^sub>v (s3[bv3::=b']\<^sub>s\<^sub>b[x3::=v]\<^sub>s\<^sub>v) s2 \<Leftarrow> \<tau>"
+ proof
+ show \<open>atom x2 \<sharp> (\<Theta>, \<Phi>, {||}::bv fset, GNil, \<Delta>, \<tau>3[bv3::=b']\<^sub>\<tau>\<^sub>b[x3::=v]\<^sub>\<tau>\<^sub>v, s3[bv3::=b']\<^sub>s\<^sub>b[x3::=v]\<^sub>s\<^sub>v, \<tau>)\<close>
+ proof -
+
+ have "atom x2 \<sharp> \<tau>3[bv3::=b']\<^sub>\<tau>\<^sub>b[x3::=v]\<^sub>\<tau>\<^sub>v"
+ using fresh_subst_v_subst_b subst_v_\<tau>_def subst_b_\<tau>_def \<open> atom x2 \<sharp> v\<close> sup by fastforce
+ moreover have "atom x2 \<sharp> s3[bv3::=b']\<^sub>s\<^sub>b[x3::=v]\<^sub>s\<^sub>v"
+ using fresh_subst_v_subst_b subst_v_s_def subst_b_s_def \<open> atom x2 \<sharp> v\<close> sup
+ proof -
+ have "\<forall>b. atom x2 = atom x3 \<or> atom x2 \<sharp> s3[bv3::=b]\<^sub>b"
+ by (metis (no_types) check_letI.hyps(1) fresh_subst_sv_if(1) fresh_subst_v_subst_b insert_commute subst_v_s_def sup) (* 140 ms *)
+ then show ?thesis
+ by (metis check_letI.hyps(1) fresh_subst_sb_if fresh_subst_sv_if(1) has_subst_b_class.subst_b_fresh_x x_fresh_b) (* 78 ms *)
+ qed
+ ultimately show ?thesis using fresh_prodN check_letI by metis
+ qed
+
+ show \<open> \<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> s3[bv3::=b']\<^sub>s\<^sub>b[x3::=v]\<^sub>s\<^sub>v \<Leftarrow> \<tau>3[bv3::=b']\<^sub>\<tau>\<^sub>b[x3::=v]\<^sub>\<tau>\<^sub>v\<close> proof( rule fundef_poly_check)
+ show \<open>check_fundef \<Theta> \<Phi> (AF_fundef f (AF_fun_typ_some bv3 (AF_fun_typ x3 b3 c3 \<tau>3 s3)))\<close>
+ using ** lookup_fun_member check_letI by metis
+ show \<open>\<Theta> ; {||} ; GNil \<turnstile> v \<Leftarrow> \<lbrace> x3 : b3[bv3::=b']\<^sub>b\<^sub>b | c3[bv3::=b']\<^sub>c\<^sub>b \<rbrace>\<close> using ** by metis
+ show \<open> \<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using ** by metis
+ show \<open> \<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f b' \<close> using ** by metis
+ qed
+ show \<open> \<Theta> ; \<Phi> ; {||} ; (x2, b_of \<tau>3[bv3::=b']\<^sub>\<tau>\<^sub>b[x3::=v]\<^sub>\<tau>\<^sub>v, c_of \<tau>3[bv3::=b']\<^sub>\<tau>\<^sub>b[x3::=v]\<^sub>\<tau>\<^sub>v x2) #\<^sub>\<Gamma> GNil ; \<Delta> \<turnstile> s2 \<Leftarrow> \<tau>\<close>
+ using check_letI ** b_of.simps c_of.simps subst_defs by metis
+ qed
+
+ moreover have "AS_let2 x2 \<tau>3[bv3::=b']\<^sub>\<tau>\<^sub>b[x3::=v]\<^sub>\<tau>\<^sub>v (s3[bv3::=b']\<^sub>s\<^sub>b[x3::=v]\<^sub>s\<^sub>v) s2 = AS_let2 x (\<tau>1'[bv1::=b']\<^sub>\<tau>\<^sub>b[x1::=v]\<^sub>\<tau>\<^sub>v) (s1'[bv1::=b']\<^sub>s\<^sub>b[x1::=v]\<^sub>s\<^sub>v) s" proof -
+ have *: "[[atom x2]]lst. s2 = [[atom x]]lst. s" using check_letI s_branch_s_branch_list.eq_iff by auto
+ moreover have " \<tau>3[bv3::=b']\<^sub>\<tau>\<^sub>b[x3::=v]\<^sub>\<tau>\<^sub>v = \<tau>1'[bv1::=b']\<^sub>\<tau>\<^sub>b[x1::=v]\<^sub>\<tau>\<^sub>v" using fun_poly_ret_unique ** check_letI by metis
+ moreover have "s3[bv3::=b']\<^sub>s\<^sub>b[x3::=v]\<^sub>s\<^sub>v = (s1'[bv1::=b']\<^sub>s\<^sub>b[x1::=v]\<^sub>s\<^sub>v)" using subst_v_flip_eq_two subst_v_s_def fun_poly_body_unique ** check_letI by metis
+ ultimately show ?thesis using s_branch_s_branch_list.eq_iff by metis
+ qed
+
+ ultimately show ?case using check_letI by auto
+qed(auto+)
+
+lemma check_s_plus:
+ assumes "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> LET x = (AE_op Plus (V_lit (L_num n1)) (V_lit (L_num n2))) IN s' \<Leftarrow> \<tau>"
+ shows "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> LET x = (AE_val (V_lit (L_num (n1+n2)))) IN s' \<Leftarrow> \<tau>"
+proof -
+ obtain t1 where 1: "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AE_op Plus (V_lit (L_num n1)) (V_lit (L_num n2)) \<Rightarrow> t1"
+ using assms check_s_elims by metis
+ then obtain z1 where 2: "t1 = \<lbrace> z1 : B_int | CE_val (V_var z1) == CE_op Plus ([V_lit (L_num n1)]\<^sup>c\<^sup>e) ([V_lit (L_num n2)]\<^sup>c\<^sup>e) \<rbrace>"
+ using infer_e_plus by metis
+
+ obtain z2 where 3: \<open>\<Theta> ; \<Phi> ; {||} ; GNil ; \<Delta> \<turnstile> AE_val (V_lit (L_num (n1+n2))) \<Rightarrow> \<lbrace> z2 : B_int | CE_val (V_var z2) == CE_val (V_lit (L_num (n1+n2))) \<rbrace>\<close>
+ using infer_v_form infer_e_valI infer_v_litI infer_l.intros infer_e_wf 1
+ by (simp add: fresh_GNil)
+
+ let ?e = " (AE_op Plus (V_lit (L_num n1)) (V_lit (L_num n2)))"
+
+ show ?thesis proof(rule subtype_let)
+ show "\<Theta> ; \<Phi> ; {||} ; GNil ; \<Delta> \<turnstile> LET x = ?e IN s' \<Leftarrow> \<tau>" using assms by auto
+ show "\<Theta> ; \<Phi> ; {||} ; GNil ; \<Delta> \<turnstile> ?e \<Rightarrow> t1" using 1 by auto
+ show "\<Theta> ; \<Phi> ; {||} ; GNil ; \<Delta> \<turnstile> [ [ L_num (n1 + n2) ]\<^sup>v ]\<^sup>e \<Rightarrow> \<lbrace> z2 : B_int | CE_val (V_var z2) == CE_val (V_lit (L_num (n1+n2))) \<rbrace>" using 3 by auto
+ show "\<Theta> ; {||} ; GNil \<turnstile> \<lbrace> z2 : B_int | CE_val (V_var z2) == CE_val (V_lit (L_num (n1+n2))) \<rbrace> \<lesssim> t1" using subtype_bop_arith
+ by (metis "1" \<open>\<And>thesis. (\<And>z1. t1 = \<lbrace> z1 : B_int | [ [ z1 ]\<^sup>v ]\<^sup>c\<^sup>e == [ plus [ [ L_num n1 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ L_num n2 ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e \<rbrace> \<Longrightarrow> thesis) \<Longrightarrow> thesis\<close> infer_e_wf(2) opp.distinct(1) type_for_lit.simps(3))
+ qed
+
+qed
+
+lemma check_s_leq:
+ assumes "\<Theta> ; \<Phi> ; {||} ; GNil ; \<Delta> \<turnstile> LET x = (AE_op LEq (V_lit (L_num n1)) (V_lit (L_num n2))) IN s' \<Leftarrow> \<tau>"
+ shows "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> LET x = (AE_val (V_lit (if (n1 \<le> n2) then L_true else L_false))) IN s' \<Leftarrow> \<tau>"
+proof -
+ obtain t1 where 1: "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AE_op LEq (V_lit (L_num n1)) (V_lit (L_num n2)) \<Rightarrow> t1"
+ using assms check_s_elims by metis
+ then obtain z1 where 2: "t1 = \<lbrace> z1 : B_bool | CE_val (V_var z1) == CE_op LEq ([V_lit (L_num n1)]\<^sup>c\<^sup>e) ([V_lit (L_num n2)]\<^sup>c\<^sup>e) \<rbrace>"
+ using infer_e_leq by auto
+
+ obtain z2 where 3: \<open>\<Theta> ; \<Phi> ; {||} ; GNil ; \<Delta> \<turnstile> AE_val (V_lit ((if (n1 \<le> n2) then L_true else L_false))) \<Rightarrow> \<lbrace> z2 : B_bool | CE_val (V_var z2) == CE_val (V_lit ((if (n1 \<le> n2) then L_true else L_false))) \<rbrace>\<close>
+ using infer_v_form infer_e_valI infer_v_litI infer_l.intros infer_e_wf 1
+ fresh_GNil
+ by simp
+
+ show ?thesis proof(rule subtype_let)
+ show \<open> \<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let x (AE_op LEq [ L_num n1 ]\<^sup>v [ L_num n2 ]\<^sup>v) s' \<Leftarrow> \<tau>\<close> using assms by auto
+ show \<open>\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AE_op LEq [ L_num n1 ]\<^sup>v [ L_num n2 ]\<^sup>v \<Rightarrow> t1\<close> using 1 by auto
+ show \<open>\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> [ [ if n1 \<le> n2 then L_true else L_false ]\<^sup>v ]\<^sup>e \<Rightarrow> \<lbrace> z2 : B_bool | CE_val (V_var z2) == CE_val (V_lit ((if (n1 \<le> n2) then L_true else L_false))) \<rbrace>\<close> using 3 by auto
+ show \<open>\<Theta> ; {||} ; GNil \<turnstile> \<lbrace> z2 : B_bool | CE_val (V_var z2) == CE_val (V_lit ((if (n1 \<le> n2) then L_true else L_false))) \<rbrace> \<lesssim> t1\<close>
+ using subtype_bop_arith[where opp=LEq] check_s_wf assms 2
+ by (metis opp.distinct(1) subtype_bop_arith type_l_eq)
+ qed
+qed
+
+lemma check_s_eq:
+ assumes "\<Theta> ; \<Phi> ; {||} ; GNil ; \<Delta> \<turnstile> LET x = (AE_op Eq (V_lit (n1)) (V_lit ( n2))) IN s' \<Leftarrow> \<tau>"
+ shows "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> LET x = (AE_val (V_lit (if (n1 = n2) then L_true else L_false))) IN s' \<Leftarrow> \<tau>"
+proof -
+ obtain t1 where 1: "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AE_op Eq (V_lit (n1)) (V_lit (n2)) \<Rightarrow> t1"
+ using assms check_s_elims by metis
+ then obtain z1 where 2: "t1 = \<lbrace> z1 : B_bool | CE_val (V_var z1) == CE_op Eq ([V_lit (n1)]\<^sup>c\<^sup>e) ([V_lit (n2)]\<^sup>c\<^sup>e) \<rbrace>"
+ using infer_e_leq by auto
+
+ obtain z2 where 3: \<open>\<Theta> ; \<Phi> ; {||} ; GNil ; \<Delta> \<turnstile> AE_val (V_lit ((if (n1 = n2) then L_true else L_false))) \<Rightarrow> \<lbrace> z2 : B_bool | CE_val (V_var z2) == CE_val (V_lit ((if (n1 = n2) then L_true else L_false))) \<rbrace>\<close>
+ using infer_v_form infer_e_valI infer_v_litI infer_l.intros infer_e_wf 1
+ fresh_GNil
+ by simp
+
+ show ?thesis proof(rule subtype_let)
+ show \<open> \<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let x (AE_op Eq [ n1 ]\<^sup>v [ n2 ]\<^sup>v) s' \<Leftarrow> \<tau>\<close> using assms by auto
+ show \<open>\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AE_op Eq [ n1 ]\<^sup>v [ n2 ]\<^sup>v \<Rightarrow> t1\<close> using 1 by auto
+ show \<open>\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> [ [ if n1 = n2 then L_true else L_false ]\<^sup>v ]\<^sup>e \<Rightarrow> \<lbrace> z2 : B_bool | CE_val (V_var z2) == CE_val (V_lit ((if (n1 = n2) then L_true else L_false))) \<rbrace>\<close> using 3 by auto
+ show \<open>\<Theta> ; {||} ; GNil \<turnstile> \<lbrace> z2 : B_bool | CE_val (V_var z2) == CE_val (V_lit ((if (n1 = n2) then L_true else L_false))) \<rbrace> \<lesssim> t1\<close>
+ proof -
+ have " \<lbrace> z2 : B_bool | [ [ z2 ]\<^sup>v ]\<^sup>c\<^sup>e == [ eq [ [ n1 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ n2 ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e \<rbrace> = t1" using 2
+ by (metis \<tau>_fresh_c fresh_opp_all infer_l_form2 infer_l_fresh ms_fresh_all(31) ms_fresh_all(33) obtain_fresh_z type_e_eq type_l_eq)
+ moreover have "\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f GNil" using assms wfX_wfY by fastforce
+ moreover have "base_for_lit n1 = base_for_lit n2" using 1 infer_e_wf wfE_elims(12) wfV_elims
+ by metis
+ ultimately show ?thesis using subtype_bop_eq[OF \<open>\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f GNil\<close>, of n1 n2 z2] by auto
+ qed
+ qed
+qed
+
+subsection \<open>Operators\<close>
+
+lemma preservation_plus:
+ assumes "\<Theta>; \<Phi>; \<Delta> \<turnstile> \<langle> \<delta> , LET x = (AE_op Plus (V_lit (L_num n1)) (V_lit (L_num n2))) IN s' \<rangle> \<Leftarrow> \<tau>"
+ shows "\<Theta>; \<Phi>; \<Delta> \<turnstile> \<langle> \<delta> , LET x = (AE_val (V_lit (L_num (n1+n2)))) IN s' \<rangle> \<Leftarrow> \<tau>"
+proof -
+
+ have tt: "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let x (AE_op Plus (V_lit (L_num n1)) (V_lit (L_num n2))) s' \<Leftarrow> \<tau>" and dsim: "\<Theta> \<turnstile> \<delta> \<sim> \<Delta>" and fd:"(\<forall>fd\<in>set \<Phi>. check_fundef \<Theta> \<Phi> fd)"
+ using assms config_type_elims by blast+
+
+ hence "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile>AS_let x (AE_val (V_lit (L_num (n1+n2)))) s' \<Leftarrow> \<tau>" using check_s_plus assms by auto
+
+ hence "\<Theta>; \<Phi>; \<Delta> \<turnstile> \<langle> \<delta> , AS_let x (AE_val (V_lit ( (L_num (n1+n2))))) s' \<rangle> \<Leftarrow> \<tau>" using dsim config_typeI fd by presburger
+ then show ?thesis using dsim config_typeI
+ by (meson order_refl)
+qed
+
+lemma preservation_leq:
+ assumes "\<Theta>; \<Phi>; \<Delta> \<turnstile> \<langle> \<delta> , AS_let x (AE_op LEq (V_lit (L_num n1)) (V_lit (L_num n2))) s' \<rangle> \<Leftarrow> \<tau>"
+ shows "\<Theta>; \<Phi>; \<Delta> \<turnstile> \<langle> \<delta> , AS_let x (AE_val (V_lit (((if (n1 \<le> n2) then L_true else L_false))))) s' \<rangle> \<Leftarrow> \<tau>"
+proof -
+
+ have tt: "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let x (AE_op LEq (V_lit (L_num n1)) (V_lit (L_num n2))) s' \<Leftarrow> \<tau>" and dsim: "\<Theta> \<turnstile> \<delta> \<sim> \<Delta>" and fd:"(\<forall>fd\<in>set \<Phi>. check_fundef \<Theta> \<Phi> fd)"
+ using assms config_type_elims by blast+
+
+ hence "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile>AS_let x (AE_val (V_lit ( ((if (n1 \<le> n2) then L_true else L_false))))) s' \<Leftarrow> \<tau>" using check_s_leq assms by auto
+
+ hence "\<Theta>; \<Phi>; \<Delta> \<turnstile> \<langle> \<delta> , AS_let x (AE_val (V_lit ( (((if (n1 \<le> n2) then L_true else L_false)))))) s' \<rangle> \<Leftarrow> \<tau>" using dsim config_typeI fd by presburger
+ then show ?thesis using dsim config_typeI
+ by (meson order_refl)
+qed
+
+lemma preservation_eq:
+ assumes "\<Theta>; \<Phi>; \<Delta> \<turnstile> \<langle> \<delta> , AS_let x (AE_op Eq (V_lit (n1)) (V_lit (n2))) s' \<rangle> \<Leftarrow> \<tau>"
+ shows "\<Theta>; \<Phi>; \<Delta> \<turnstile> \<langle> \<delta> , AS_let x (AE_val (V_lit (((if (n1 = n2) then L_true else L_false))))) s' \<rangle> \<Leftarrow> \<tau>"
+proof -
+
+ have tt: "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let x (AE_op Eq (V_lit (n1)) (V_lit (n2))) s' \<Leftarrow> \<tau>" and dsim: "\<Theta> \<turnstile> \<delta> \<sim> \<Delta>" and fd:"(\<forall>fd\<in>set \<Phi>. check_fundef \<Theta> \<Phi> fd)"
+ using assms config_type_elims by blast+
+
+ hence "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile>AS_let x (AE_val (V_lit ( ((if (n1 = n2) then L_true else L_false))))) s' \<Leftarrow> \<tau>" using check_s_eq assms by auto
+
+ hence "\<Theta>; \<Phi>; \<Delta> \<turnstile> \<langle> \<delta> , AS_let x (AE_val (V_lit ( (((if (n1 = n2) then L_true else L_false)))))) s' \<rangle> \<Leftarrow> \<tau>" using dsim config_typeI fd by presburger
+ then show ?thesis using dsim config_typeI
+ by (meson order_refl)
+qed
+
+subsection \<open>Let Statements\<close>
+
+lemma subst_s_abs_lst:
+ fixes s::s and sa::s and v'::v
+ assumes "[[atom x]]lst. s = [[atom xa]]lst. sa" and "atom xa \<sharp> v \<and> atom x \<sharp> v"
+ shows "s[x::=v]\<^sub>s\<^sub>v = sa[xa::=v]\<^sub>s\<^sub>v"
+proof -
+ obtain c'::x where cdash: "atom c' \<sharp> (v, x, xa, s, sa)" using obtain_fresh by blast
+ moreover have " (x \<leftrightarrow> c') \<bullet> s = (xa \<leftrightarrow> c') \<bullet> sa" proof -
+ have "atom c' \<sharp> (s, sa) \<and> atom c' \<sharp> (x, xa, s, sa)" using cdash by auto
+ thus ?thesis using assms by auto
+ qed
+ ultimately show ?thesis using assms
+ using subst_sv_flip by auto
+qed
+
+lemma check_let_val:
+ fixes v::v and s::s
+ shows "\<Theta> ; \<Phi> ; B ; G ; \<Delta> \<turnstile> ss \<Leftarrow> \<tau> \<Longrightarrow> B = {||} \<Longrightarrow> G = GNil \<Longrightarrow>
+ ss = AS_let x (AE_val v) s \<or> ss = AS_let2 x t (AS_val v) s \<Longrightarrow> \<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> (s[x::=v]\<^sub>s\<^sub>v) \<Leftarrow> \<tau>" and
+ "check_branch_s \<Theta> \<Phi> \<B> GNil \<Delta> tid dc const v cs \<tau> \<Longrightarrow> True" and
+ "check_branch_list \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist v css \<tau> \<Longrightarrow> True"
+proof(nominal_induct \<tau> and \<tau> and \<tau> avoiding: v rule: check_s_check_branch_s_check_branch_list.strong_induct)
+ case (check_letI x1 \<Theta> \<Phi> \<B> \<Gamma> \<Delta> e \<tau> z s1 b c)
+ hence *:"e = AE_val v" by auto
+ let ?G = "(x1, b, c[z::=V_var x1]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>"
+ have "\<Theta> ; \<Phi> ; \<B> ; ?G[x1::=v]\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x1::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> s1[x1::=v]\<^sub>s\<^sub>v \<Leftarrow> \<tau>[x1::=v]\<^sub>\<tau>\<^sub>v"
+ proof(rule subst_infer_check_s(1))
+ show **:\<open> \<Theta> ; \<B> ; \<Gamma> \<turnstile> v \<Rightarrow> \<lbrace> z : b | c \<rbrace>\<close> using infer_e_elims check_letI * by fast
+ thus \<open>\<Theta> ; \<B> ; \<Gamma> \<turnstile> \<lbrace> z : b | c \<rbrace> \<lesssim> \<lbrace> z : b | c \<rbrace>\<close> using subtype_reflI infer_v_wf by metis
+ show \<open>atom z \<sharp> (x1, v)\<close> using check_letI fresh_Pair by auto
+ show \<open>\<Theta> ; \<Phi> ; \<B> ; (x1, b, c[z::=V_var x1]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile> s1 \<Leftarrow> \<tau>\<close> using check_letI subst_defs by auto
+ show "(x1, b, c[z::=V_var x1]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma> = GNil @ (x1, b, c[z::=V_var x1]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>" by auto
+ qed
+
+ hence "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile> s1[x1::=v]\<^sub>s\<^sub>v \<Leftarrow> \<tau>" using check_letI by auto
+ moreover have " s1[x1::=v]\<^sub>s\<^sub>v = s[x::=v]\<^sub>s\<^sub>v"
+ by (metis (full_types) check_letI fresh_GNil infer_e_elims(7) s_branch_s_branch_list.distinct s_branch_s_branch_list.eq_iff(2)
+ subst_s_abs_lst wfG_x_fresh_in_v_simple)
+
+ ultimately show ?case using check_letI by simp
+next
+ case (check_let2I x1 \<Theta> \<Phi> \<B> \<Gamma> \<Delta> t s1 \<tau> s2 )
+
+ hence s1eq:"s1 = AS_val v" by auto
+ let ?G = "(x1, b_of t, c_of t x1) #\<^sub>\<Gamma> \<Gamma>"
+ obtain z::x where *:"atom z \<sharp> (x1 , v,t)" using obtain_fresh_z by metis
+ hence teq:"t = \<lbrace> z : b_of t | c_of t z \<rbrace>" using b_of_c_of_eq by auto
+ have "\<Theta> ; \<Phi> ; \<B> ; ?G[x1::=v]\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x1::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> s2[x1::=v]\<^sub>s\<^sub>v \<Leftarrow> \<tau>[x1::=v]\<^sub>\<tau>\<^sub>v"
+ proof(rule subst_check_check_s(1))
+ obtain t' where "\<Theta> ; \<B> ; \<Gamma> \<turnstile> v \<Rightarrow> t' \<and> \<Theta> ; \<B> ; \<Gamma> \<turnstile> t' \<lesssim> t" using check_s_elims(1) check_let2I(10) s1eq by auto
+ thus **:\<open> \<Theta> ; \<B> ; \<Gamma> \<turnstile> v \<Leftarrow> \<lbrace> z : b_of t | c_of t z \<rbrace>\<close> using check_v.intros teq by auto
+ show "atom z \<sharp> (x1, v)" using * by auto
+ show \<open> \<Theta> ; \<Phi> ; \<B> ; (x1, b_of t, c_of t x1) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile> s2 \<Leftarrow> \<tau>\<close> using check_let2I by auto
+ show "(x1, b_of t , c_of t x1) #\<^sub>\<Gamma> \<Gamma> = GNil @ (x1, b_of t, (c_of t z)[z::=V_var x1]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>" using append_g.simps c_of_switch * fresh_prodN by metis
+ qed
+
+ hence "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile> s2[x1::=v]\<^sub>s\<^sub>v \<Leftarrow> \<tau>" using check_let2I by auto
+ moreover have " s2[x1::=v]\<^sub>s\<^sub>v = s[x::=v]\<^sub>s\<^sub>v" using
+ check_let2I fresh_GNil check_s_elims s_branch_s_branch_list.distinct s_branch_s_branch_list.eq_iff
+ subst_s_abs_lst wfG_x_fresh_in_v_simple
+ proof -
+ have "AS_let2 x t (AS_val v) s = AS_let2 x1 t s1 s2"
+ by (metis check_let2I.prems(3) s_branch_s_branch_list.distinct s_branch_s_branch_list.eq_iff(3)) (* 62 ms *)
+ then show ?thesis
+ by (metis (no_types) check_let2I check_let2I.prems(2) check_s_elims(1) fresh_GNil s_branch_s_branch_list.eq_iff(3) subst_s_abs_lst wfG_x_fresh_in_v_simple) (* 375 ms *)
+ qed
+
+ ultimately show ?case using check_let2I by simp
+qed(auto+)
+
+lemma preservation_let_val:
+ assumes "\<Theta>; \<Phi>; \<Delta> \<turnstile> \<langle> \<delta> , AS_let x (AE_val v) s \<rangle> \<Leftarrow> \<tau> \<or> \<Theta>; \<Phi>; \<Delta> \<turnstile> \<langle> \<delta> , AS_let2 x t (AS_val v) s \<rangle> \<Leftarrow> \<tau>" (is "?A \<or> ?B")
+ shows " \<exists>\<Delta>'. \<Theta>; \<Phi>; \<Delta>' \<turnstile> \<langle> \<delta> , s[x::=v]\<^sub>s\<^sub>v \<rangle> \<Leftarrow> \<tau> \<and> \<Delta> \<sqsubseteq> \<Delta>'"
+proof -
+ have tt: "\<Theta> \<turnstile> \<delta> \<sim> \<Delta>" and fd: "(\<forall>fd\<in>set \<Phi>. check_fundef \<Theta> \<Phi> fd)"
+ using assms by blast+
+
+ have "?A \<or> ?B" using assms by auto
+ then have "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> s[x::=v]\<^sub>s\<^sub>v \<Leftarrow> \<tau>"
+ proof
+ assume "\<Theta>; \<Phi>; \<Delta> \<turnstile> \<langle> \<delta> , AS_let x (AE_val v) s \<rangle> \<Leftarrow> \<tau>"
+ hence * : " \<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let x (AE_val v) s \<Leftarrow> \<tau>" by blast
+ thus ?thesis using check_let_val by simp
+ next
+ assume "\<Theta>; \<Phi>; \<Delta> \<turnstile> \<langle> \<delta> , AS_let2 x t (AS_val v) s \<rangle> \<Leftarrow> \<tau>"
+ hence * : "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let2 x t (AS_val v) s \<Leftarrow> \<tau>" by blast
+ thus ?thesis using check_let_val by simp
+ qed
+
+ thus ?thesis using tt config_typeI fd
+ order_refl by metis
+qed
+
+lemma check_s_fst_snd:
+ assumes "fst_snd = AE_fst \<and> v=v1 \<or> fst_snd = AE_snd \<and> v=v2"
+ and "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let x (fst_snd (V_pair v1 v2)) s' \<Leftarrow> \<tau>"
+ shows "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let x ( AE_val v) s' \<Leftarrow> \<tau>"
+proof -
+ have 1: \<open> \<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let x (fst_snd (V_pair v1 v2)) s' \<Leftarrow> \<tau>\<close> using assms by auto
+
+ then obtain t1 where 2:"\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> (fst_snd (V_pair v1 v2)) \<Rightarrow> t1" using check_s_elims by auto
+
+ show ?thesis using subtype_let 1 2 assms
+ by (meson infer_e_fst_pair infer_e_snd_pair)
+qed
+
+lemma preservation_fst_snd:
+ assumes "\<Theta>; \<Phi>; \<Delta> \<turnstile> \<langle> \<delta> , LET x = (fst_snd (V_pair v1 v2)) IN s' \<rangle> \<Leftarrow> \<tau>" and
+ "fst_snd = AE_fst \<and> v=v1 \<or> fst_snd = AE_snd \<and> v=v2"
+ shows "\<exists>\<Delta>'. \<Theta>; \<Phi>; \<Delta> \<turnstile> \<langle> \<delta> , LET x = (AE_val v) IN s' \<rangle> \<Leftarrow> \<tau> \<and> \<Delta> \<sqsubseteq> \<Delta>'"
+proof -
+ have tt: "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let x (fst_snd (V_pair v1 v2)) s' \<Leftarrow> \<tau> \<and> \<Theta> \<turnstile> \<delta> \<sim> \<Delta>" using assms by blast
+ hence t2: "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let x (fst_snd (V_pair v1 v2)) s' \<Leftarrow> \<tau>" by auto
+
+ moreover have "\<forall>fd\<in>set \<Phi>. check_fundef \<Theta> \<Phi> fd" using assms config_type_elims by auto
+ ultimately show ?thesis using config_typeI order_refl tt assms check_s_fst_snd by metis
+qed
+
+inductive_cases check_branch_s_elims2[elim!]:
+ "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta>; tid ; cons ; const ; v \<turnstile> cs \<Leftarrow> \<tau>"
+
+lemmas freshers = freshers atom_dom.simps toSet.simps fresh_def x_not_in_b_set
+declare freshers [simp]
+
+lemma subtype_eq_if:
+ fixes t::\<tau> and va::v
+ assumes "\<Theta> ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b_of t | c_of t z \<rbrace>" and "\<Theta> ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b_of t | c IMP c_of t z \<rbrace> "
+ shows "\<Theta> ; \<B> ; \<Gamma> \<turnstile> \<lbrace> z : b_of t | c_of t z \<rbrace> \<lesssim> \<lbrace> z : b_of t | c IMP c_of t z \<rbrace> "
+proof -
+ obtain x::x where xf:"atom x \<sharp> ((\<Theta>, \<B>, \<Gamma>, z, c_of t z, z, c IMP c_of t z ),c)" using obtain_fresh by metis
+
+ moreover have "\<Theta> ; \<B> ; (x, b_of t, (c_of t z)[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<Turnstile> (c IMP c_of t z )[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v "
+ unfolding subst_cv.simps
+ proof(rule valid_eq_imp)
+
+ have "\<Theta> ; \<B> ; (x, b_of t, (c_of t z)[z::=[ x ]\<^sup>v]\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f (c IMP (c_of t z))[z::=[ x ]\<^sup>v]\<^sub>v "
+ apply(rule wfT_wfC_cons)
+ apply(simp add: assms, simp add: assms, unfold fresh_prodN )
+ using xf fresh_prodN by metis+
+ thus "\<Theta> ; \<B> ; (x, b_of t, (c_of t z)[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f c[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v IMP (c_of t z)[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v "
+ using subst_cv.simps subst_defs by auto
+ qed
+
+ ultimately show ?thesis using subtype_baseI assms fresh_Pair subst_defs by metis
+qed
+
+lemma subtype_eq_if_\<tau>:
+ fixes t::\<tau> and va::v
+ assumes "\<Theta> ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f t" and "\<Theta> ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b_of t | c IMP c_of t z \<rbrace> " and "atom z \<sharp> t"
+ shows "\<Theta> ; \<B> ; \<Gamma> \<turnstile> t \<lesssim> \<lbrace> z : b_of t | c IMP c_of t z \<rbrace> "
+proof -
+ have "t = \<lbrace> z : b_of t | c_of t z \<rbrace>" using b_of_c_of_eq assms by auto
+ thus ?thesis using subtype_eq_if assms c_of.simps b_of.simps by metis
+qed
+
+lemma valid_conj:
+ assumes "\<Theta> ; \<B> ; \<Gamma> \<Turnstile> c1" and "\<Theta> ; \<B> ; \<Gamma> \<Turnstile> c2"
+ shows "\<Theta> ; \<B> ; \<Gamma> \<Turnstile> c1 AND c2"
+proof
+ show \<open> \<Theta> ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f c1 AND c2 \<close> using valid.simps wfC_conjI assms by auto
+ show \<open>\<forall>i. \<Theta> ; \<Gamma> \<turnstile> i \<and> i \<Turnstile> \<Gamma> \<longrightarrow> i \<Turnstile> c1 AND c2 \<close>
+ proof(rule+)
+ fix i
+ assume *:"\<Theta> ; \<Gamma> \<turnstile> i \<and> i \<Turnstile> \<Gamma> "
+ thus "i \<lbrakk> c1 \<rbrakk> ~ True" using assms valid.simps
+ using is_satis.cases by blast
+ show "i \<lbrakk> c2 \<rbrakk> ~ True" using assms valid.simps
+ using is_satis.cases * by blast
+ qed
+qed
+
+subsection \<open>Other Statements\<close>
+
+lemma check_if:
+ fixes s'::s and cs::branch_s and css::branch_list and v::v
+ shows "\<Theta>; \<Phi>; B; G; \<Delta> \<turnstile> s' \<Leftarrow> \<tau> \<Longrightarrow> s' = IF (V_lit ll) THEN s1 ELSE s2 \<Longrightarrow>
+ \<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f \<tau> \<Longrightarrow> G = GNil \<Longrightarrow> B = {||} \<Longrightarrow> ll = L_true \<and> s = s1 \<or> ll = L_false \<and> s = s2 \<Longrightarrow>
+ \<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> s \<Leftarrow> \<tau>" and
+ "check_branch_s \<Theta> \<Phi> {||} GNil \<Delta> tid dc const v cs \<tau> \<Longrightarrow> True" and
+ "check_branch_list \<Theta> \<Phi> {||} \<Gamma> \<Delta> tid dclist v css \<tau> \<Longrightarrow> True"
+proof(nominal_induct \<tau> and \<tau> and \<tau> rule: check_s_check_branch_s_check_branch_list.strong_induct)
+ case (check_ifI z \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v s1 s2 \<tau>)
+ obtain z' where teq: "\<tau> = \<lbrace> z' : b_of \<tau> | c_of \<tau> z' \<rbrace> \<and> atom z' \<sharp> (z,\<tau>)" using obtain_fresh_z_c_of by metis
+ hence ceq: "(c_of \<tau> z')[z'::=[ z ]\<^sup>v]\<^sub>c\<^sub>v = (c_of \<tau> z)" using c_of_switch fresh_Pair by metis
+ have zf: " atom z \<sharp> c_of \<tau> z'"
+ using c_of_fresh check_ifI teq fresh_Pair fresh_at_base by metis
+ hence 1:"\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> s \<Leftarrow> \<lbrace> z : b_of \<tau> | CE_val (V_lit ll) == CE_val (V_lit ll) IMP c_of \<tau> z \<rbrace>" using check_ifI by auto
+ moreover have 2:"\<Theta> ; {||} ; GNil \<turnstile> (\<lbrace> z : b_of \<tau> | CE_val (V_lit ll) == CE_val (V_lit ll) IMP c_of \<tau> z \<rbrace>) \<lesssim> \<tau>"
+ proof -
+ have "\<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f (\<lbrace> z : b_of \<tau> | CE_val (V_lit ll ) == CE_val (V_lit ll) IMP c_of \<tau> z \<rbrace>)" using check_ifI check_s_wf by auto
+ moreover have "\<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f \<tau>" using check_s_wf check_ifI by auto
+ ultimately show ?thesis using subtype_if_simp[of \<Theta> " {||}" z "b_of \<tau>" ll "c_of \<tau> z'" z'] using teq ceq zf subst_defs by metis
+ qed
+ ultimately show ?case using check_s_supertype(1) check_ifI by metis
+qed(auto+)
+
+lemma preservation_if:
+ assumes "\<Theta>; \<Phi>; \<Delta> \<turnstile> \<langle> \<delta> , IF (V_lit ll) THEN s1 ELSE s2 \<rangle> \<Leftarrow> \<tau>" and
+ "ll = L_true \<and> s = s1 \<or> ll = L_false \<and> s = s2"
+ shows "\<Theta>; \<Phi>; \<Delta> \<turnstile> \<langle>\<delta>, s\<rangle> \<Leftarrow> \<tau> \<and> setD \<Delta> \<subseteq> setD \<Delta>"
+proof -
+ have *: "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_if (V_lit ll) s1 s2 \<Leftarrow> \<tau> \<and> (\<forall>fd\<in>set \<Phi>. check_fundef \<Theta> \<Phi> fd)"
+ using assms config_type_elims by metis
+ hence "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> s \<Leftarrow> \<tau>" using check_s_wf check_if assms by metis
+ hence "\<Theta>; \<Phi>; \<Delta> \<turnstile> \<langle>\<delta>, s\<rangle> \<Leftarrow> \<tau> \<and> setD \<Delta> \<subseteq> setD \<Delta>" using config_typeI *
+ using assms(1) by blast
+ thus ?thesis by blast
+qed
+
+lemma wfT_conj:
+ assumes "\<Theta> ; \<B> ; GNil \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b | c1 \<rbrace>" and "\<Theta> ; \<B> ; GNil \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b | c2 \<rbrace>"
+ shows "\<Theta> ; \<B> ; GNil \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b | c1 AND c2\<rbrace>"
+proof
+ show \<open>atom z \<sharp> (\<Theta>, \<B>, GNil)\<close>
+ apply(unfold fresh_prodN, intro conjI)
+ using wfTh_fresh wfT_wf assms apply metis
+ using fresh_GNil x_not_in_b_set fresh_def by metis+
+ show \<open> \<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f b \<close> using wfT_elims assms by metis
+ have "\<Theta> ; \<B> ; (z, b, TRUE) #\<^sub>\<Gamma> GNil \<turnstile>\<^sub>w\<^sub>f c1" using wfT_wfC fresh_GNil assms by auto
+ moreover have "\<Theta> ; \<B> ; (z, b, TRUE) #\<^sub>\<Gamma> GNil \<turnstile>\<^sub>w\<^sub>f c2" using wfT_wfC fresh_GNil assms by auto
+ ultimately show \<open> \<Theta> ; \<B> ; (z, b, TRUE) #\<^sub>\<Gamma> GNil \<turnstile>\<^sub>w\<^sub>f c1 AND c2 \<close> using wfC_conjI by auto
+qed
+
+lemma subtype_conj:
+ assumes "\<Theta> ; \<B> ; GNil \<turnstile> t \<lesssim> \<lbrace> z : b | c1 \<rbrace>" and "\<Theta> ; \<B> ; GNil \<turnstile> t \<lesssim> \<lbrace> z : b | c2 \<rbrace>"
+ shows "\<Theta> ; \<B> ; GNil \<turnstile> \<lbrace> z : b | c_of t z \<rbrace> \<lesssim> \<lbrace> z : b | c1 AND c2 \<rbrace>"
+proof -
+ have beq: "b_of t = b" using subtype_eq_base2 b_of.simps assms by metis
+ obtain x::x where x:\<open>atom x \<sharp> (\<Theta>, \<B>, GNil, z, c_of t z, z, c1 AND c2 )\<close> using obtain_fresh by metis
+ thus ?thesis proof
+ have "atom z \<sharp> t" using subtype_wf wfT_supp fresh_def x_not_in_b_set atom_dom.simps toSet.simps assms dom.simps by fastforce
+ hence t:"t = \<lbrace> z : b_of t | c_of t z \<rbrace>" using b_of_c_of_eq by auto
+ thus \<open> \<Theta> ; \<B> ; GNil \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b | c_of t z \<rbrace> \<close> using subtype_wf beq assms by auto
+
+ show \<open>\<Theta> ; \<B> ; (x, b, (c_of t z)[z::=[ x ]\<^sup>v]\<^sub>v) #\<^sub>\<Gamma> GNil \<Turnstile> (c1 AND c2 )[z::=[ x ]\<^sup>v]\<^sub>v \<close>
+ proof -
+ have \<open>\<Theta> ; \<B> ; (x, b, (c_of t z)[z::=[ x ]\<^sup>v]\<^sub>v) #\<^sub>\<Gamma> GNil \<Turnstile> c1[z::=[ x ]\<^sup>v]\<^sub>v \<close>
+ proof(rule subtype_valid)
+ show \<open>\<Theta> ; \<B> ; GNil \<turnstile> t \<lesssim> \<lbrace> z : b | c1 \<rbrace>\<close> using assms by auto
+ show \<open>atom x \<sharp> GNil\<close> using fresh_GNil by auto
+ show \<open>t = \<lbrace> z : b | c_of t z \<rbrace>\<close> using t beq by auto
+ show \<open>\<lbrace> z : b | c1 \<rbrace> = \<lbrace> z : b | c1 \<rbrace>\<close> by auto
+ qed
+ moreover have \<open>\<Theta> ; \<B> ; (x, b, (c_of t z)[z::=[ x ]\<^sup>v]\<^sub>v) #\<^sub>\<Gamma> GNil \<Turnstile> c2[z::=[ x ]\<^sup>v]\<^sub>v \<close>
+ proof(rule subtype_valid)
+ show \<open>\<Theta> ; \<B> ; GNil \<turnstile> t \<lesssim> \<lbrace> z : b | c2 \<rbrace>\<close> using assms by auto
+ show \<open>atom x \<sharp> GNil\<close> using fresh_GNil by auto
+ show \<open>t = \<lbrace> z : b | c_of t z \<rbrace>\<close> using t beq by auto
+ show \<open>\<lbrace> z : b | c2 \<rbrace> = \<lbrace> z : b | c2 \<rbrace>\<close> by auto
+ qed
+ ultimately show ?thesis unfolding subst_cv.simps subst_v_c_def using valid_conj by metis
+ qed
+ thus \<open> \<Theta> ; \<B> ; GNil \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b | c1 AND c2 \<rbrace> \<close> using subtype_wf wfT_conj assms by auto
+ qed
+qed
+
+lemma infer_v_conj:
+ assumes "\<Theta> ; \<B> ; GNil \<turnstile> v \<Leftarrow> \<lbrace> z : b | c1 \<rbrace>" and "\<Theta> ; \<B> ; GNil \<turnstile> v \<Leftarrow> \<lbrace> z : b | c2 \<rbrace>"
+ shows "\<Theta> ; \<B> ; GNil \<turnstile> v \<Leftarrow> \<lbrace> z : b | c1 AND c2 \<rbrace>"
+proof -
+ obtain t1 where t1: "\<Theta> ; \<B> ; GNil \<turnstile> v \<Rightarrow> t1 \<and> \<Theta> ; \<B> ; GNil \<turnstile> t1 \<lesssim> \<lbrace> z : b | c1 \<rbrace>"
+ using assms check_v_elims by metis
+ obtain t2 where t2: "\<Theta> ; \<B> ; GNil \<turnstile> v \<Rightarrow> t2 \<and> \<Theta> ; \<B> ; GNil \<turnstile> t2 \<lesssim> \<lbrace> z : b | c2 \<rbrace>"
+ using assms check_v_elims by metis
+ have teq: "t1 = \<lbrace> z : b | c_of t1 z \<rbrace>" using subtype_eq_base2 b_of.simps
+ by (metis (full_types) b_of_c_of_eq fresh_GNil infer_v_t_wf t1 wfT_x_fresh)
+ have "t1 = t2" using infer_v_uniqueness t1 t2 by auto
+ hence " \<Theta> ; \<B> ; GNil \<turnstile> \<lbrace> z : b | c_of t1 z \<rbrace> \<lesssim> \<lbrace> z : b | c1 AND c2 \<rbrace>" using subtype_conj t1 t2 by simp
+ hence " \<Theta> ; \<B> ; GNil \<turnstile> t1 \<lesssim> \<lbrace> z : b | c1 AND c2 \<rbrace>" using teq by auto
+ thus ?thesis using t1 using check_v.intros by auto
+qed
+
+lemma wfG_conj:
+ fixes c1::c
+ assumes "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f (x, b, c1 AND c2) #\<^sub>\<Gamma> \<Gamma>"
+ shows "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f (x, b, c1) #\<^sub>\<Gamma> \<Gamma>"
+proof(cases "c1 \<in> {TRUE, FALSE}")
+ case True
+ then show ?thesis using assms wfG_cons2I wfG_elims wfX_wfY by metis
+next
+ case False
+ then show ?thesis using assms wfG_cons1I wfG_elims wfX_wfY wfC_elims
+ by (metis wfG_elim2)
+qed
+
+lemma check_match:
+ fixes s'::s and s::s and css::branch_list and cs::branch_s
+ shows "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile> s \<Leftarrow> \<tau> \<Longrightarrow> True" and
+ "\<Theta> ; \<Phi> ; B ; G ; \<Delta> ; tid ; dc ; const ; vcons \<turnstile> cs \<Leftarrow> \<tau> \<Longrightarrow>
+ vcons = V_cons tid dc v \<Longrightarrow> B = {||} \<Longrightarrow> G = GNil \<Longrightarrow> cs = (dc x' \<Rightarrow> s') \<Longrightarrow>
+ \<Theta> ; {||} ; GNil \<turnstile> v \<Leftarrow> const \<Longrightarrow>
+ \<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> s'[x'::=v]\<^sub>s\<^sub>v \<Leftarrow> \<tau>" and
+ "\<Theta> ; \<Phi> ; B ; G ; \<Delta> ; tid ; dclist ; vcons \<turnstile> css \<Leftarrow> \<tau> \<Longrightarrow> distinct (map fst dclist) \<Longrightarrow>
+ vcons = V_cons tid dc v \<Longrightarrow> B = {||} \<Longrightarrow> (dc, const) \<in> set dclist \<Longrightarrow> G = GNil \<Longrightarrow>
+ Some (AS_branch dc x' s') = lookup_branch dc css \<Longrightarrow> \<Theta> ; {||} ; GNil \<turnstile> v \<Leftarrow> const \<Longrightarrow>
+ \<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> s'[x'::=v]\<^sub>s\<^sub>v \<Leftarrow> \<tau>"
+proof(nominal_induct \<tau> and \<tau> and \<tau> avoiding: x' v rule: check_s_check_branch_s_check_branch_list.strong_induct)
+ case (check_branch_list_consI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid consa consta va cs \<tau> dclist cssa)
+
+ then obtain xa and sa where cseq:"cs = AS_branch consa xa sa" using check_branch_s_elims2[OF check_branch_list_consI(1)] by metis
+
+ show ?case proof(cases "dc = consa")
+ case True
+ hence "cs = AS_branch consa x' s'" using check_branch_list_consI cseq
+ by (metis lookup_branch.simps(2) option.inject)
+ moreover have "const = consta" using check_branch_list_consI distinct.simps
+ by (metis True dclist_distinct_unique list.set_intros(1))
+ moreover have "va = V_cons tid consa v" using check_branch_list_consI True by auto
+ ultimately show ?thesis using check_branch_list_consI by auto
+ next
+ case False
+ hence "Some (AS_branch dc x' s') = lookup_branch dc cssa" using lookup_branch.simps(2) check_branch_list_consI(10) cseq by auto
+ moreover have "(dc, const) \<in> set dclist" using check_branch_list_consI False by simp
+ ultimately show ?thesis using check_branch_list_consI by auto
+ qed
+
+next
+ case (check_branch_list_finalI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid cons const va cs \<tau>)
+ hence " cs = AS_branch cons x' s'" using lookup.simps check_branch_list_finalI lookup_branch.simps option.inject
+ by (metis map_of.simps(1) map_of_Cons_code(2) option.distinct(1) s_branch_s_branch_list.exhaust(2) weak_map_of_SomeI)
+ then show ?case using check_branch_list_finalI by auto
+next
+ case (check_branch_s_branchI \<Theta> \<B> \<Gamma> \<Delta> \<tau> const x \<Phi> tid cons va s)
+
+ text \<open>Supporting facts here to make the main body of the proof concise\<close>
+ have xf:"atom x \<sharp> \<tau>" proof -
+ have "supp \<tau> \<subseteq> supp \<B>" using wf_supp(4) check_branch_s_branchI atom_dom.simps toSet.simps dom.simps by fastforce
+ thus ?thesis using fresh_def x_not_in_b_set by blast
+ qed
+
+ hence "\<tau>[x::=v]\<^sub>\<tau>\<^sub>v = \<tau>" using forget_subst_v subst_v_\<tau>_def by auto
+ have "\<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v = \<Delta>" using forget_subst_dv wfD_x_fresh fresh_GNil check_branch_s_branchI by metis
+
+ have "supp v = {}" using check_branch_s_branchI check_v_wf wfV_supp_nil by metis
+ hence "supp va = {}" using \<open> va = V_cons tid cons v\<close> v.supp pure_supp by auto
+
+ let ?G = "(x, b_of const, [ va ]\<^sup>c\<^sup>e == [ V_cons tid cons [ x ]\<^sup>v ]\<^sup>c\<^sup>e AND c_of const x ) #\<^sub>\<Gamma> \<Gamma>"
+ obtain z::x where z: "const = \<lbrace> z : b_of const | c_of const z \<rbrace> \<and> atom z \<sharp> (x', v,x,const,va)"
+ using obtain_fresh_z_c_of by metis
+
+ have vt: \<open>\<Theta> ; \<B> ; GNil \<turnstile> v \<Leftarrow> \<lbrace> z : b_of const | [ va ]\<^sup>c\<^sup>e == [ V_cons tid cons [ z ]\<^sup>v ]\<^sup>c\<^sup>e AND c_of const z \<rbrace>\<close>
+ proof(rule infer_v_conj)
+ obtain t' where t: "\<Theta> ; \<B> ; GNil \<turnstile> v \<Rightarrow> t' \<and> \<Theta> ; \<B> ; GNil \<turnstile> t' \<lesssim> const"
+ using check_v_elims check_branch_s_branchI by metis
+ show "\<Theta> ; \<B> ; GNil \<turnstile> v \<Leftarrow> \<lbrace> z : b_of const | [ va ]\<^sup>c\<^sup>e == [ V_cons tid cons [ z ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrace>"
+ proof(rule check_v_top)
+
+ show "\<Theta> ; \<B> ; GNil \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b_of const | [ va ]\<^sup>c\<^sup>e == [ V_cons tid cons [ z ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrace> "
+
+ proof(rule wfG_wfT)
+ show \<open> \<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f (x, b_of const, ([ va ]\<^sup>c\<^sup>e == [ V_cons tid cons [ z ]\<^sup>v ]\<^sup>c\<^sup>e )[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> GNil \<close>
+ proof -
+ have 1: "va[z::=[ x ]\<^sup>v]\<^sub>v\<^sub>v = va" using forget_subst_v subst_v_v_def z fresh_prodN by metis
+ moreover have 2: "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f (x, b_of const, [ va ]\<^sup>c\<^sup>e == [ V_cons tid cons [ x ]\<^sup>v ]\<^sup>c\<^sup>e AND c_of const x ) #\<^sub>\<Gamma> GNil "
+ using check_branch_s_branchI(17)[THEN check_s_wf] \<open>\<Gamma> = GNil\<close> by auto
+ moreover hence "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f (x, b_of const, [ va ]\<^sup>c\<^sup>e == [ V_cons tid cons [ x ]\<^sup>v ]\<^sup>c\<^sup>e ) #\<^sub>\<Gamma> GNil"
+ using wfG_conj by metis
+ ultimately show ?thesis
+ unfolding subst_cv.simps subst_cev.simps subst_vv.simps by auto
+ qed
+ show \<open>atom x \<sharp> ([ va ]\<^sup>c\<^sup>e == [ V_cons tid cons [ z ]\<^sup>v ]\<^sup>c\<^sup>e )\<close> unfolding c.fresh ce.fresh v.fresh
+ apply(intro conjI )
+ using check_branch_s_branchI fresh_at_base z freshers apply simp
+ using check_branch_s_branchI fresh_at_base z freshers apply simp
+ using pure_supp apply force
+ using z fresh_x_neq fresh_prod5 by metis
+ qed
+ show \<open>[ va ]\<^sup>c\<^sup>e = [ V_cons tid cons [ z ]\<^sup>v ]\<^sup>c\<^sup>e[z::=v]\<^sub>c\<^sub>e\<^sub>v\<close>
+ using \<open> va = V_cons tid cons v\<close> subst_cev.simps subst_vv.simps by auto
+ show \<open> \<Theta> ; \<B> ; GNil \<turnstile> v \<Leftarrow> const \<close> using check_branch_s_branchI by auto
+ show "supp [ va ]\<^sup>c\<^sup>e \<subseteq> supp \<B>" using \<open>supp va = {}\<close> ce.supp by simp
+ qed
+ show "\<Theta> ; \<B> ; GNil \<turnstile> v \<Leftarrow> \<lbrace> z : b_of const | c_of const z \<rbrace>"
+ using check_branch_s_branchI z by auto
+ qed
+
+ text \<open>Main body of proof for this case\<close>
+ have "\<Theta> ; \<Phi> ; \<B> ; (?G)[x::=v]\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v \<turnstile> s[x::=v]\<^sub>s\<^sub>v \<Leftarrow> \<tau>[x::=v]\<^sub>\<tau>\<^sub>v"
+ proof(rule subst_check_check_s)
+ show \<open>\<Theta> ; \<B> ; GNil \<turnstile> v \<Leftarrow> \<lbrace> z : b_of const | [ va ]\<^sup>c\<^sup>e == [ V_cons tid cons [ z ]\<^sup>v ]\<^sup>c\<^sup>e AND c_of const z \<rbrace>\<close> using vt by auto
+ show \<open>atom z \<sharp> (x, v)\<close> using z fresh_prodN by auto
+ show \<open> \<Theta> ; \<Phi> ; \<B> ; ?G ; \<Delta> \<turnstile> s \<Leftarrow> \<tau>\<close>
+ using check_branch_s_branchI by auto
+
+ show \<open> ?G = GNil @ (x, b_of const, ([ va ]\<^sup>c\<^sup>e == [ V_cons tid cons [ z ]\<^sup>v ]\<^sup>c\<^sup>e AND c_of const z)[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> GNil\<close>
+ proof -
+ have "va[z::=[ x ]\<^sup>v]\<^sub>v\<^sub>v = va" using forget_subst_v subst_v_v_def z fresh_prodN by metis
+ moreover have "(c_of const z)[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v = c_of const x"
+ using c_of_switch[of z const x] z fresh_prodN by metis
+ ultimately show ?thesis
+ unfolding subst_cv.simps subst_cev.simps subst_vv.simps append_g.simps
+ using c_of_switch[of z const x] z fresh_prodN z fresh_prodN check_branch_s_branchI by argo
+ qed
+ qed
+ moreover have "s[x::=v]\<^sub>s\<^sub>v = s'[x'::=v]\<^sub>s\<^sub>v"
+ using check_branch_s_branchI subst_v_flip_eq_two subst_v_s_def s_branch_s_branch_list.eq_iff by metis
+ ultimately show ?case using check_branch_s_branchI \<open>\<tau>[x::=v]\<^sub>\<tau>\<^sub>v = \<tau>\<close> \<open>\<Delta>[x::=v]\<^sub>\<Delta>\<^sub>v = \<Delta>\<close> by auto
+qed(auto+)
+
+text \<open>Lemmas for while reduction. Making these separate lemmas allows flexibility in wiring them into the main proof and robustness
+if we change it\<close>
+
+lemma check_unit:
+ fixes \<tau>::\<tau> and \<Phi>::\<Phi> and \<Delta>::\<Delta> and G::\<Gamma>
+ assumes "\<Theta> ; {||} ; GNil \<turnstile> \<lbrace> z : B_unit | TRUE \<rbrace> \<lesssim> \<tau>'" and "\<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f \<Delta>" and "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>" and "\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f G"
+ shows \<open> \<Theta> ; \<Phi> ; {||} ; G ; \<Delta> \<turnstile> [[ L_unit ]\<^sup>v]\<^sup>s \<Leftarrow> \<tau>'\<close>
+proof -
+ have *:"\<Theta> ; {||} ; GNil \<turnstile> [L_unit]\<^sup>v \<Rightarrow> \<lbrace> z : B_unit | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_unit ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrace>"
+ using infer_l.intros(4) infer_v_litI fresh_GNil assms wfX_wfY by (meson subtype_g_wf)
+ moreover have "\<Theta> ; {||} ; GNil \<turnstile> \<lbrace> z : B_unit | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_unit ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrace> \<lesssim> \<tau>'"
+ using subtype_top subtype_trans * infer_v_wf
+ by (meson assms(1))
+ ultimately show ?thesis (*apply(rule check_valI, auto simp add: assms,rule * )*)
+ using subtype_top subtype_trans fresh_GNil assms check_valI assms check_s_g_weakening assms toSet.simps
+ by (metis bot.extremum infer_v_g_weakening subtype_weakening wfD_wf)
+qed
+
+lemma preservation_var:
+ shows "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> VAR u : \<tau>' = v IN s \<Leftarrow> \<tau> \<Longrightarrow> \<Theta> \<turnstile> \<delta> \<sim> \<Delta> \<Longrightarrow> atom u \<sharp> \<delta> \<Longrightarrow> atom u \<sharp> \<Delta> \<Longrightarrow>
+ \<Theta>; \<Phi>; {||}; GNil; (u,\<tau>')#\<^sub>\<Delta>\<Delta> \<turnstile> s \<Leftarrow> \<tau> \<and> \<Theta> \<turnstile> (u,v)#\<delta> \<sim> (u,\<tau>')#\<^sub>\<Delta>\<Delta>"
+ and
+ "check_branch_s \<Theta> \<Phi> {||} GNil \<Delta> tid dc const v cs \<tau> \<Longrightarrow> True" and
+ "check_branch_list \<Theta> \<Phi> {||} \<Gamma> \<Delta> tid dclist v css \<tau> \<Longrightarrow> True"
+proof(nominal_induct "{||}::bv fset" GNil \<Delta> "VAR u : \<tau>' = v IN s" \<tau> and \<tau> and \<tau> rule: check_s_check_branch_s_check_branch_list.strong_induct)
+ case (check_varI u' \<Theta> \<Phi> \<Delta> \<tau> s')
+ hence "\<Theta>; \<Phi>; {||}; GNil; (u, \<tau>') #\<^sub>\<Delta> \<Delta> \<turnstile> s \<Leftarrow> \<tau>" using check_s_abs_u check_v_wf by metis
+
+ moreover have "\<Theta> \<turnstile> ((u,v)#\<delta>) \<sim> ((u,\<tau>')#\<^sub>\<Delta>\<Delta>)" proof
+ show \<open>\<Theta> \<turnstile> \<delta> \<sim> \<Delta> \<close> using check_varI by auto
+ show \<open>\<Theta> ; {||} ; GNil \<turnstile> v \<Leftarrow> \<tau>'\<close> using check_varI by auto
+ show \<open>u \<notin> fst ` set \<delta>\<close> using check_varI fresh_d_fst_d by auto
+ qed
+
+ ultimately show ?case by simp
+qed(auto+)
+
+lemma check_while:
+ shows "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> WHILE s1 DO { s2 } \<Leftarrow> \<tau> \<Longrightarrow> atom x \<sharp> (s1, s2) \<Longrightarrow> atom z' \<sharp> x \<Longrightarrow>
+ \<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> LET x : (\<lbrace> z' : B_bool | TRUE \<rbrace>) = s1 IN (IF (V_var x) THEN (s2 ;; (WHILE s1 DO {s2}))
+ ELSE ([ V_lit L_unit]\<^sup>s)) \<Leftarrow> \<tau>" and
+ "check_branch_s \<Theta> \<Phi> {||} GNil \<Delta> tid dc const v cs \<tau> \<Longrightarrow> True" and
+ "check_branch_list \<Theta> \<Phi> {||} \<Gamma> \<Delta> tid dclist v css \<tau> \<Longrightarrow> True"
+proof(nominal_induct "{||}::bv fset" GNil \<Delta> "AS_while s1 s2" \<tau> and \<tau> and \<tau> avoiding: s1 s2 x z' rule: check_s_check_branch_s_check_branch_list.strong_induct)
+ case (check_whileI \<Theta> \<Phi> \<Delta> s1 z s2 \<tau>')
+ have teq:"\<lbrace> z' : B_bool | TRUE \<rbrace> = \<lbrace> z : B_bool | TRUE \<rbrace>" using \<tau>.eq_iff by auto
+
+ show ?case proof
+ have " atom x \<sharp> \<tau>' " using wfT_nil_supp fresh_def subtype_wfT check_whileI(5) by fast
+ moreover have "atom x \<sharp> \<lbrace> z' : B_bool | TRUE \<rbrace> " using \<tau>.fresh c.fresh b.fresh by metis
+ ultimately show \<open>atom x \<sharp> (\<Theta>, \<Phi>, {||}, GNil, \<Delta>, \<lbrace> z' : B_bool | TRUE \<rbrace>, s1, \<tau>')\<close>
+ apply(unfold fresh_prodN)
+ using check_whileI wb_x_fresh check_s_wf wfD_x_fresh fresh_empty_fset fresh_GNil fresh_Pair \<open>atom x \<sharp> \<tau>'\<close> by metis
+
+ show \<open> \<Theta> ; \<Phi> ; {||} ; GNil ; \<Delta> \<turnstile> s1 \<Leftarrow> \<lbrace> z' : B_bool | TRUE \<rbrace>\<close> using check_whileI teq by metis
+
+ let ?G = "(x, b_of \<lbrace> z' : B_bool | TRUE \<rbrace>, c_of \<lbrace> z' : B_bool | TRUE \<rbrace> x) #\<^sub>\<Gamma> GNil"
+
+ have cof:"(c_of \<lbrace> z' : B_bool | TRUE \<rbrace> x) = C_true" using c_of.simps check_whileI subst_cv.simps by metis
+ have wfg: "\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f ?G" proof
+ show "c_of \<lbrace> z' : B_bool | TRUE \<rbrace> x \<in> {TRUE, FALSE}" using subst_cv.simps cof by auto
+ show "\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f GNil " using wfG_nilI check_whileI wfX_wfY check_s_wf by metis
+ show "atom x \<sharp> GNil" using fresh_GNil by auto
+ show "\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f b_of \<lbrace> z' : B_bool | TRUE \<rbrace> " using wfB_boolI wfX_wfY check_s_wf b_of.simps
+ by (metis \<open>\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f GNil\<close>)
+ qed
+
+ obtain zz::x where zf:\<open>atom zz \<sharp> ((\<Theta>, \<Phi>, {||}::bv fset, ?G , \<Delta>, [ x ]\<^sup>v,
+ AS_seq s2 (AS_while s1 s2), AS_val [ L_unit ]\<^sup>v, \<tau>'),x,?G)\<close>
+ using obtain_fresh by blast
+ show \<open> \<Theta> ; \<Phi> ; {||} ; ?G ; \<Delta> \<turnstile>
+ AS_if [ x ]\<^sup>v (AS_seq s2 (AS_while s1 s2)) (AS_val [ L_unit ]\<^sup>v) \<Leftarrow> \<tau>'\<close>
+ proof
+ show "atom zz \<sharp> (\<Theta>, \<Phi>, {||}::bv fset, ?G , \<Delta>, [ x ]\<^sup>v, AS_seq s2 (AS_while s1 s2), AS_val [ L_unit ]\<^sup>v, \<tau>')" using zf by auto
+ show \<open>\<Theta> ; {||} ; ?G \<turnstile> [ x ]\<^sup>v \<Leftarrow> \<lbrace> zz : B_bool | TRUE \<rbrace>\<close> proof
+ have "atom zz \<sharp> x \<and> atom zz \<sharp> (\<Theta>, {||}::bv fset, ?G)" using zf fresh_prodN by metis
+ thus \<open> \<Theta> ; {||} ; ?G \<turnstile> [ x ]\<^sup>v \<Rightarrow>\<lbrace> zz : B_bool | [[zz]\<^sup>v]\<^sup>c\<^sup>e == [[ x ]\<^sup>v]\<^sup>c\<^sup>e \<rbrace>\<close>
+ using infer_v_varI lookup.simps wfg b_of.simps by metis
+ thus \<open>\<Theta> ; {||} ; ?G \<turnstile> \<lbrace> zz : B_bool | [[ zz ]\<^sup>v]\<^sup>c\<^sup>e == [[ x ]\<^sup>v]\<^sup>c\<^sup>e \<rbrace> \<lesssim> \<lbrace> zz : B_bool | TRUE \<rbrace>\<close>
+ using subtype_top infer_v_wf by metis
+ qed
+ show \<open> \<Theta> ; \<Phi> ; {||} ; ?G ; \<Delta> \<turnstile> AS_seq s2 (AS_while s1 s2) \<Leftarrow> \<lbrace> zz : b_of \<tau>' | [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of \<tau>' zz \<rbrace>\<close>
+ proof
+ have "\<lbrace> zz : B_unit | TRUE \<rbrace> = \<lbrace> z : B_unit | TRUE \<rbrace>" using \<tau>.eq_iff by auto
+ thus \<open> \<Theta> ; \<Phi> ; {||} ; ?G ; \<Delta> \<turnstile> s2 \<Leftarrow> \<lbrace> zz : B_unit | TRUE \<rbrace>\<close> using check_s_g_weakening(1) [OF check_whileI(3) _ wfg] toSet.simps
+ by (simp add: \<open>\<lbrace> zz : B_unit | TRUE \<rbrace> = \<lbrace> z : B_unit | TRUE \<rbrace>\<close>)
+
+ show \<open> \<Theta> ; \<Phi> ; {||} ; ?G ; \<Delta> \<turnstile> AS_while s1 s2 \<Leftarrow> \<lbrace> zz : b_of \<tau>' | [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of \<tau>' zz \<rbrace>\<close>
+ proof(rule check_s_supertype(1))
+ have \<open> \<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_while s1 s2 \<Leftarrow> \<tau>'\<close> using check_whileI by auto
+ thus *:\<open> \<Theta> ; \<Phi> ; {||} ; ?G ; \<Delta> \<turnstile> AS_while s1 s2 \<Leftarrow> \<tau>' \<close> using check_s_g_weakening(1)[OF _ _ wfg] toSet.simps by auto
+
+ show \<open>\<Theta> ; {||} ; ?G \<turnstile> \<tau>' \<lesssim> \<lbrace> zz : b_of \<tau>' | [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of \<tau>' zz \<rbrace>\<close>
+ proof(rule subtype_eq_if_\<tau>)
+ show \<open> \<Theta> ; {||} ; ?G \<turnstile>\<^sub>w\<^sub>f \<tau>' \<close> using * check_s_wf by auto
+ show \<open> \<Theta> ; {||} ; ?G \<turnstile>\<^sub>w\<^sub>f \<lbrace> zz : b_of \<tau>' | [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of \<tau>' zz \<rbrace> \<close>
+ apply(rule wfT_eq_imp, simp add: base_for_lit.simps)
+ using subtype_wf check_whileI wfg zf fresh_prodN by metis+
+ show \<open>atom zz \<sharp> \<tau>'\<close> using zf fresh_prodN by metis
+ qed
+ qed
+
+ qed
+ show \<open> \<Theta> ; \<Phi> ; {||} ; ?G ; \<Delta> \<turnstile> AS_val [ L_unit ]\<^sup>v \<Leftarrow> \<lbrace> zz : b_of \<tau>' | [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_false ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of \<tau>' zz \<rbrace>\<close>
+ proof(rule check_s_supertype(1))
+
+ show *:\<open> \<Theta> ; \<Phi> ; {||} ; ?G ; \<Delta> \<turnstile> AS_val [ L_unit ]\<^sup>v \<Leftarrow> \<tau>'\<close>
+ using check_unit[OF check_whileI(5) _ _ wfg] using check_whileI wfg wfX_wfY check_s_wf by metis
+ show \<open>\<Theta> ; {||} ; ?G \<turnstile> \<tau>' \<lesssim> \<lbrace> zz : b_of \<tau>' | [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_false ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of \<tau>' zz \<rbrace>\<close>
+ proof(rule subtype_eq_if_\<tau>)
+ show \<open> \<Theta> ; {||} ; ?G \<turnstile>\<^sub>w\<^sub>f \<tau>' \<close> using * check_s_wf by metis
+ show \<open> \<Theta> ; {||} ; ?G \<turnstile>\<^sub>w\<^sub>f \<lbrace> zz : b_of \<tau>' | [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_false ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of \<tau>' zz \<rbrace> \<close>
+ apply(rule wfT_eq_imp, simp add: base_for_lit.simps)
+ using subtype_wf check_whileI wfg zf fresh_prodN by metis+
+ show \<open>atom zz \<sharp> \<tau>'\<close> using zf fresh_prodN by metis
+ qed
+ qed
+ qed
+ qed
+qed(auto+)
+
+lemma check_s_narrow:
+ fixes s::s and x::x
+ assumes "atom x \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>, \<Delta>, c, \<tau>, s)" and "\<Theta> ; \<Phi> ; \<B> ; (x, B_bool, c) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile> s \<Leftarrow> \<tau>" and
+ "\<Theta> ; \<B> ; \<Gamma> \<Turnstile> c "
+ shows "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile> s \<Leftarrow> \<tau>"
+proof -
+ let ?B = "({||}::bv fset)"
+ let ?v = "V_lit L_true"
+
+ obtain z::x where z: "atom z \<sharp> (x, [ L_true ]\<^sup>v,c)" using obtain_fresh by metis
+
+ have "atom z \<sharp> c" using z fresh_prodN by auto
+ hence c:" c[x::=[ z ]\<^sup>v]\<^sub>v[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v = c" using subst_v_c_def by simp
+
+ have "\<Theta> ; \<Phi> ; \<B> ; ((x,B_bool, c) #\<^sub>\<Gamma> \<Gamma>)[x::=?v]\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=?v]\<^sub>\<Delta>\<^sub>v \<turnstile> s[x::=?v]\<^sub>s\<^sub>v \<Leftarrow> \<tau>[x::=?v]\<^sub>\<tau>\<^sub>v" proof(rule subst_infer_check_s(1) )
+ show vt: \<open> \<Theta> ; \<B> ; \<Gamma> \<turnstile> [ L_true ]\<^sup>v \<Rightarrow> \<lbrace> z : B_bool | (CE_val (V_var z)) == (CE_val (V_lit L_true )) \<rbrace> \<close>
+ using infer_v_litI check_s_wf wfG_elims(2) infer_trueI assms by metis
+ show \<open>\<Theta> ; \<B> ; \<Gamma> \<turnstile> \<lbrace> z : B_bool | (CE_val (V_var z)) == (CE_val (V_lit L_true )) \<rbrace> \<lesssim> \<lbrace> z : B_bool | c[x::=[ z ]\<^sup>v]\<^sub>v \<rbrace>\<close> proof
+ show \<open>atom x \<sharp> (\<Theta>, \<B>, \<Gamma>, z, [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e , z, c[x::=[ z ]\<^sup>v]\<^sub>v)\<close>
+ apply(unfold fresh_prodN, intro conjI)
+ prefer 5
+ using c.fresh ce.fresh v.fresh z fresh_prodN apply auto[1]
+ prefer 6
+ using fresh_subst_v_if[of "atom x" c x] assms fresh_prodN apply simp
+ using z assms fresh_prodN apply metis
+ using z assms fresh_prodN apply metis
+ using z assms fresh_prodN apply metis
+ using z fresh_prodN assms fresh_at_base by metis+
+ show \<open> \<Theta> ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : B_bool | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrace> \<close> using vt infer_v_wf by simp
+ show \<open> \<Theta> ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : B_bool | c[x::=[ z ]\<^sup>v]\<^sub>v \<rbrace> \<close> proof(rule wfG_wfT)
+ show \<open> \<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f (x, B_bool, c[x::=[ z ]\<^sup>v]\<^sub>v[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<close> using c check_s_wf assms by metis
+ have " atom x \<sharp> [ z ]\<^sup>v" using v.fresh z fresh_at_base by auto
+ thus \<open>atom x \<sharp> c[x::=[ z ]\<^sup>v]\<^sub>v\<close> using fresh_subst_v_if[of "atom x" c ] by auto
+ qed
+ have wfg: "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f(x, B_bool, ([ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e )[z::=[ x ]\<^sup>v]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>"
+ using wfT_wfG vt infer_v_wf fresh_prodN assms by simp
+ show \<open>\<Theta> ; \<B> ; (x, B_bool, ([ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e )[z::=[ x ]\<^sup>v]\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<Turnstile> c[x::=[ z ]\<^sup>v]\<^sub>v[z::=[ x ]\<^sup>v]\<^sub>v \<close>
+ using c valid_weakening[OF assms(3) _ wfg ] toSet.simps
+ using subst_v_c_def by auto
+ qed
+ show \<open>atom z \<sharp> (x, [ L_true ]\<^sup>v)\<close> using z fresh_prodN by metis
+ show \<open> \<Theta> ; \<Phi> ; \<B> ; (x, B_bool, c) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile> s \<Leftarrow> \<tau>\<close> using assms by auto
+
+ thus \<open>(x, B_bool, c) #\<^sub>\<Gamma> \<Gamma> = GNil @ (x, B_bool, c[x::=[ z ]\<^sup>v]\<^sub>v[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>\<close> using append_g.simps c by auto
+ qed
+
+ moreover have "((x,B_bool, c) #\<^sub>\<Gamma> \<Gamma>)[x::=?v]\<^sub>\<Gamma>\<^sub>v = \<Gamma> " using subst_gv.simps by auto
+ ultimately show ?thesis using assms forget_subst_dv forget_subst_sv forget_subst_tv fresh_prodN by metis
+qed
+
+lemma check_assert_s:
+ fixes s::s and x::x
+ assumes "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_assert c s \<Leftarrow> \<tau>"
+ shows "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> s \<Leftarrow> \<tau> \<and> \<Theta> ; {||} ; GNil \<Turnstile> c "
+proof -
+ let ?B = "({||}::bv fset)"
+ let ?v = "V_lit L_true"
+
+ obtain x where x: "\<Theta> ; \<Phi> ; ?B ; (x,B_bool, c) #\<^sub>\<Gamma> GNil ; \<Delta> \<turnstile> s \<Leftarrow> \<tau> \<and> atom x \<sharp> (\<Theta>, \<Phi>, ?B, GNil, \<Delta>, c, \<tau>, s ) \<and> \<Theta> ; ?B ; GNil \<Turnstile> c "
+ using check_s_elims(10)[OF \<open>\<Theta> ; \<Phi> ; ?B ; GNil ; \<Delta> \<turnstile> AS_assert c s \<Leftarrow> \<tau>\<close>] valid.simps by metis
+
+ show ?thesis using assms check_s_narrow x by metis
+qed
+
+lemma infer_v_pair2I:
+ "atom z \<sharp> (v1, v2) \<Longrightarrow>
+ atom z \<sharp> (\<Theta>, \<B>, \<Gamma>) \<Longrightarrow>
+ \<Theta> ; \<B> ; \<Gamma> \<turnstile> v1 \<Rightarrow> t1 \<Longrightarrow>
+ \<Theta> ; \<B> ; \<Gamma> \<turnstile> v2 \<Rightarrow> t2 \<Longrightarrow>
+ b1 = b_of t1 \<Longrightarrow> b2 = b_of t2 \<Longrightarrow>
+ \<Theta> ; \<B> ; \<Gamma> \<turnstile> [ v1 , v2 ]\<^sup>v \<Rightarrow> \<lbrace> z : [ b1 , b2 ]\<^sup>b | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ v1 , v2 ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrace>"
+ using infer_v_pairI by simp
+
+subsection \<open>Main Lemma\<close>
+
+lemma preservation:
+ assumes "\<Phi> \<turnstile> \<langle>\<delta>, s\<rangle> \<longrightarrow> \<langle>\<delta>', s'\<rangle>" and "\<Theta>; \<Phi>; \<Delta> \<turnstile> \<langle>\<delta>, s\<rangle> \<Leftarrow> \<tau>"
+ shows "\<exists>\<Delta>'. \<Theta>; \<Phi>; \<Delta>' \<turnstile> \<langle>\<delta>', s'\<rangle> \<Leftarrow> \<tau> \<and> \<Delta> \<sqsubseteq> \<Delta>'"
+ using assms
+proof(induct arbitrary: \<tau> rule: reduce_stmt.induct)
+ case (reduce_let_plusI \<delta> x n1 n2 s')
+ then show ?case using preservation_plus
+ by (metis order_refl)
+next
+ case (reduce_let_leqI b n1 n2 \<delta> x s)
+ then show ?case using preservation_leq by (metis order_refl)
+next
+ case (reduce_let_eqI b n1 n2 \<Phi> \<delta> x s)
+ then show ?case using preservation_eq[OF reduce_let_eqI(2)] order_refl by metis
+next
+ case (reduce_let_appI f z b c \<tau>' s' \<Phi> \<delta> x v s)
+ hence tt: "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let x (AE_app f v) s \<Leftarrow> \<tau> \<and> \<Theta> \<turnstile> \<delta> \<sim> \<Delta> \<and> (\<forall>fd\<in>set \<Phi>. check_fundef \<Theta> \<Phi> fd)" using config_type_elims[OF reduce_let_appI(2)] by metis
+ hence *:"\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let x (AE_app f v) s \<Leftarrow> \<tau>" by auto
+
+ hence "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let2 x (\<tau>'[z::=v]\<^sub>\<tau>\<^sub>v) (s'[z::=v]\<^sub>s\<^sub>v) s \<Leftarrow> \<tau>"
+ using preservation_app reduce_let_appI tt by auto
+
+ hence "\<Theta>; \<Phi>; \<Delta> \<turnstile> \<langle> \<delta> , AS_let2 x (\<tau>'[z::=v]\<^sub>\<tau>\<^sub>v) s'[z::=v]\<^sub>s\<^sub>v s \<rangle> \<Leftarrow> \<tau>" using config_typeI tt by auto
+ then show ?case using tt order_refl reduce_let_appI by metis
+
+next
+ case (reduce_let_appPI f bv z b c \<tau>' s' \<Phi> \<delta> x b' v s)
+
+ hence tt: "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let x (AE_appP f b' v) s \<Leftarrow> \<tau> \<and> \<Theta> \<turnstile> \<delta> \<sim> \<Delta> \<and> (\<forall>fd\<in>set \<Phi>. check_fundef \<Theta> \<Phi> fd)" using config_type_elims[OF reduce_let_appPI(2)] by metis
+ hence *:"\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let x (AE_appP f b' v) s \<Leftarrow> \<tau>" by auto
+
+ have "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let2 x (\<tau>'[bv::=b']\<^sub>\<tau>\<^sub>b[z::=v]\<^sub>\<tau>\<^sub>v) (s'[bv::=b']\<^sub>s\<^sub>b[z::=v]\<^sub>s\<^sub>v) s \<Leftarrow> \<tau>"
+ proof(rule preservation_poly_app)
+ show \<open>Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ z b c \<tau>' s'))) = lookup_fun \<Phi> f\<close> using reduce_let_appPI by metis
+ show \<open>\<forall>fd\<in>set \<Phi>. check_fundef \<Theta> \<Phi> fd\<close> using tt lookup_fun_member by metis
+ show \<open> \<Theta> ; \<Phi> ;{||} ; GNil ; \<Delta> \<turnstile> AS_let x (AE_appP f b' v) s \<Leftarrow> \<tau>\<close> using * by auto
+ show \<open> \<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f b' \<close> using check_s_elims infer_e_wf wfE_elims * by metis
+ qed(auto+)
+
+ hence "\<Theta>; \<Phi>; \<Delta> \<turnstile> \<langle> \<delta> , AS_let2 x (\<tau>'[bv::=b']\<^sub>\<tau>\<^sub>b[z::=v]\<^sub>\<tau>\<^sub>v) s'[bv::=b']\<^sub>s\<^sub>b[z::=v]\<^sub>s\<^sub>v s \<rangle> \<Leftarrow> \<tau>" using config_typeI tt by auto
+ then show ?case using tt order_refl reduce_let_appI by metis
+
+next
+ case (reduce_if_trueI \<delta> s1 s2)
+ then show ?case using preservation_if by metis
+next
+ case (reduce_if_falseI uw \<delta> s1 s2)
+ then show ?case using preservation_if by metis
+next
+ case (reduce_let_valI \<delta> x v s)
+ then show ?case using preservation_let_val by presburger
+next
+ case (reduce_let_mvar u v \<delta> \<Phi> x s)
+ hence *:"\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let x (AE_mvar u) s \<Leftarrow> \<tau> \<and> \<Theta> \<turnstile> \<delta> \<sim> \<Delta> \<and> (\<forall>fd\<in>set \<Phi>. check_fundef \<Theta> \<Phi> fd)"
+ using config_type_elims by blast
+
+ hence **: "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let x (AE_mvar u) s \<Leftarrow> \<tau>" by auto
+ obtain xa::x and za::x and ca::c and ba::b and sa::s where
+ sa1: "atom xa \<sharp> (\<Theta>, \<Phi>, {||}::bv fset, GNil, \<Delta>, AE_mvar u, \<tau>) \<and> atom za \<sharp> (xa, \<Theta>, \<Phi>, {||}::bv fset, GNil, \<Delta>, AE_mvar u, \<tau>, sa) \<and>
+ \<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AE_mvar u \<Rightarrow> \<lbrace> za : ba | ca \<rbrace> \<and>
+ \<Theta> ; \<Phi> ; {||} ; (xa, ba, ca[za::=V_var xa]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> GNil ; \<Delta> \<turnstile> sa \<Leftarrow> \<tau> \<and>
+ (\<forall>c. atom c \<sharp> (s, sa) \<longrightarrow> atom c \<sharp> (x, xa, s, sa) \<longrightarrow> (x \<leftrightarrow> c) \<bullet> s = (xa \<leftrightarrow> c) \<bullet> sa)"
+ using check_s_elims(2)[OF **] subst_defs by metis
+
+ have "\<Theta> ; {||} ; GNil \<turnstile> v \<Leftarrow> \<lbrace> za : ba | ca \<rbrace>" proof -
+ have " (u , \<lbrace> za : ba | ca \<rbrace>) \<in> setD \<Delta>" using infer_e_elims(11) sa1 by fast
+ thus ?thesis using delta_sim_v reduce_let_mvar config_type_elims check_s_wf by metis
+ qed
+
+ then obtain \<tau>' where vst: "\<Theta> ; {||} ; GNil \<turnstile> v \<Rightarrow> \<tau>' \<and>
+ \<Theta> ; {||} ; GNil \<turnstile> \<tau>' \<lesssim> \<lbrace> za : ba | ca \<rbrace>" using check_v_elims by blast
+
+ obtain za2 and ba2 and ca2 where zbc: "\<tau>' = (\<lbrace> za2 : ba2 | ca2 \<rbrace>) \<and> atom za2 \<sharp> (xa, (xa, \<Theta>, \<Phi>, {||}::bv fset, GNil, \<Delta>, AE_val v, \<tau>, sa))"
+ using obtain_fresh_z by blast
+ have beq: "ba=ba2" using subtype_eq_base vst zbc by blast
+
+ moreover have xaf: "atom xa \<sharp> (za, za2)"
+ apply(unfold fresh_prodN, intro conjI)
+ using sa1 zbc fresh_prodN fresh_x_neq by metis+
+
+ have sat2: " \<Theta> ; \<Phi> ; {||} ; GNil@(xa, ba, ca2[za2::=V_var xa]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> GNil ; \<Delta> \<turnstile> sa \<Leftarrow> \<tau>" proof(rule ctx_subtype_s)
+ show "\<Theta> ; \<Phi> ; {||} ; GNil @ (xa, ba, ca[za::=V_var xa]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> GNil ; \<Delta> \<turnstile> sa \<Leftarrow> \<tau>" using sa1 by auto
+ show "\<Theta> ; {||} ; GNil \<turnstile> \<lbrace> za2 : ba | ca2 \<rbrace> \<lesssim> \<lbrace> za : ba | ca \<rbrace>" using beq zbc vst by fast
+ show "atom xa \<sharp> (za, za2, ca, ca2)" proof -
+ have *:"\<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f (\<lbrace> za2 : ba2 | ca2 \<rbrace>) " using zbc vst subtype_wf by auto
+ hence "supp ca2 \<subseteq> { atom za2 }" using wfT_supp_c[OF *] supp_GNil by simp
+ moreover have "atom za2 \<sharp> xa" using zbc fresh_Pair fresh_x_neq by metis
+ ultimately have "atom xa \<sharp> ca2" using zbc supp_at_base fresh_def
+ by (metis empty_iff singleton_iff subset_singletonD)
+ moreover have "atom xa \<sharp> ca" proof -
+ have *:"\<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f (\<lbrace> za : ba | ca \<rbrace>) " using zbc vst subtype_wf by auto
+ hence "supp ca \<subseteq> { atom za }" using wfT_supp \<tau>.supp by force
+ moreover have "xa \<noteq> za" using fresh_def fresh_x_neq xaf fresh_Pair by metis
+ ultimately show ?thesis using fresh_def by auto
+ qed
+ ultimately show ?thesis using xaf sa1 fresh_prod4 fresh_Pair by metis
+ qed
+ qed
+ hence dwf: "\<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f \<Delta>" using sa1 infer_e_wf by meson
+
+ have "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let xa (AE_val v) sa \<Leftarrow> \<tau>" proof
+ have "atom xa \<sharp> (AE_val v)" using infer_v_wf(1) wfV_supp fresh_def e.fresh x_not_in_b_set vst by fastforce
+ thus "atom xa \<sharp> (\<Theta>, \<Phi>, {||}::bv fset, GNil, \<Delta>, AE_val v, \<tau>)" using sa1 freshers by simp
+ have "atom za2 \<sharp> (AE_val v)" using infer_v_wf(1) wfV_supp fresh_def e.fresh x_not_in_b_set vst by fastforce
+ thus "atom za2 \<sharp> (xa, \<Theta>, \<Phi>, {||}::bv fset, GNil, \<Delta>, AE_val v, \<tau>, sa)" using zbc freshers fresh_prodN by auto
+ have " \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>" using sa1 infer_e_wf by auto
+ thus "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AE_val v \<Rightarrow> \<lbrace> za2 : ba | ca2 \<rbrace>"
+ using zbc vst beq dwf infer_e_valI by blast
+ show "\<Theta> ; \<Phi> ; {||} ; (xa, ba, ca2[za2::=V_var xa]\<^sub>v) #\<^sub>\<Gamma> GNil ; \<Delta> \<turnstile> sa \<Leftarrow> \<tau>" using sat2 append_g.simps subst_defs by metis
+ qed
+ moreover have "AS_let xa (AE_val v) sa = AS_let x (AE_val v) s" proof -
+ have "[[atom x]]lst. s = [[atom xa]]lst. sa"
+ using sa1 Abs1_eq_iff_all(3)[where z=" (s, sa)"] by metis
+ thus ?thesis using s_branch_s_branch_list.eq_iff(2) by metis
+ qed
+ ultimately have "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let x (AE_val v) s \<Leftarrow> \<tau>" by auto
+
+ then show ?case using reduce_let_mvar * config_typeI
+ by (meson order_refl)
+next
+ case (reduce_let2I \<Phi> \<delta> s1 \<delta>' s1' x t s2)
+ hence **: "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let2 x t s1 s2 \<Leftarrow> \<tau> \<and> \<Theta> \<turnstile> \<delta> \<sim> \<Delta> \<and> (\<forall>fd\<in>set \<Phi>. check_fundef \<Theta> \<Phi> fd)" using config_type_elims[OF reduce_let2I(3)] by blast
+ hence *:"\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let2 x t s1 s2 \<Leftarrow> \<tau>" by auto
+
+ obtain xa::x and z::x and c and b and s2a::s where st: "atom xa \<sharp> (\<Theta>, \<Phi>, {||}::bv fset, GNil, \<Delta>, t, s1, \<tau>) \<and>
+ \<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> s1 \<Leftarrow> t \<and>
+ \<Theta> ; \<Phi> ; {||} ; (xa, b_of t, c_of t xa) #\<^sub>\<Gamma> GNil ; \<Delta> \<turnstile> s2a \<Leftarrow> \<tau> \<and> ([[atom x]]lst. s2 = [[atom xa]]lst. s2a) "
+ using check_s_elims(4)[OF *] Abs1_eq_iff_all(3) by metis
+
+ hence "\<Theta>; \<Phi>; \<Delta> \<turnstile> \<langle> \<delta> , s1 \<rangle> \<Leftarrow> t" using config_typeI ** by auto
+ then obtain \<Delta>' where s1r: "\<Theta>; \<Phi>; \<Delta>' \<turnstile> \<langle> \<delta>' , s1' \<rangle> \<Leftarrow> t \<and> \<Delta> \<sqsubseteq> \<Delta>'" using reduce_let2I by presburger
+
+ have "\<Theta>; \<Phi>; {||}; GNil; \<Delta>' \<turnstile> AS_let2 xa t s1' s2a \<Leftarrow> \<tau>"
+ proof(rule check_let2I)
+ show *:"\<Theta>; \<Phi>; {||}; GNil; \<Delta>' \<turnstile> s1' \<Leftarrow> t" using config_type_elims st s1r by metis
+ show "atom xa \<sharp> (\<Theta>, \<Phi>, {||}::bv fset, GNil, \<Delta>',t, s1', \<tau>)" proof -
+ have "atom xa \<sharp> s1'" using check_s_x_fresh * by auto
+ moreover have "atom xa \<sharp> \<Delta>'" using check_s_x_fresh * by auto
+ ultimately show ?thesis using st fresh_prodN by metis
+ qed
+
+ show "\<Theta> ; \<Phi> ; {||} ; (xa, b_of t, c_of t xa) #\<^sub>\<Gamma> GNil ; \<Delta>' \<turnstile> s2a \<Leftarrow> \<tau>" proof -
+ have "\<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f \<Delta>'" using * check_s_wf by auto
+ moreover have "\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f ((xa, b_of t, c_of t xa) #\<^sub>\<Gamma> GNil)" using st check_s_wf by auto
+ ultimately have "\<Theta> ; {||} ; ((xa, b_of t , c_of t xa) #\<^sub>\<Gamma> GNil) \<turnstile>\<^sub>w\<^sub>f \<Delta>'" using wf_weakening by auto
+ thus ?thesis using check_s_d_weakening check_s_wf st s1r by metis
+ qed
+ qed
+ moreover have "AS_let2 xa t s1' s2a = AS_let2 x t s1' s2" using st s_branch_s_branch_list.eq_iff by metis
+ ultimately have "\<Theta>; \<Phi>; {||}; GNil; \<Delta>' \<turnstile> AS_let2 x t s1' s2 \<Leftarrow> \<tau>" using st by argo
+ moreover have "\<Theta> \<turnstile> \<delta>' \<sim> \<Delta>'" using config_type_elims s1r by fast
+ ultimately show ?case using config_typeI **
+ by (meson s1r)
+next
+ case (reduce_let2_valI vb \<delta> x t v s)
+ then show ?case using preservation_let_val by meson
+next
+ case (reduce_varI u \<delta> \<Phi> \<tau>' v s)
+
+ hence ** : "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_var u \<tau>' v s \<Leftarrow> \<tau> \<and> \<Theta> \<turnstile> \<delta> \<sim> \<Delta> \<and> (\<forall>fd\<in>set \<Phi>. check_fundef \<Theta> \<Phi> fd)"
+ using config_type_elims by meson
+ have uf: "atom u \<sharp> \<Delta>" using reduce_varI delta_sim_fresh by force
+ hence *: "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_var u \<tau>' v s \<Leftarrow> \<tau>" and " \<Theta> \<turnstile> \<delta> \<sim> \<Delta>" using ** by auto
+
+ thus ?case using preservation_var reduce_varI config_typeI ** set_subset_Cons
+ setD_ConsD subsetI by (metis delta_sim_fresh)
+
+next
+ case (reduce_assignI \<Phi> \<delta> u v )
+ hence *: "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_assign u v \<Leftarrow> \<tau> \<and> \<Theta> \<turnstile> \<delta> \<sim> \<Delta> \<and> (\<forall>fd\<in>set \<Phi>. check_fundef \<Theta> \<Phi> fd)"
+ using config_type_elims by meson
+ then obtain z and \<tau>' where zt: "\<Theta> ; {||} ; GNil \<turnstile> (\<lbrace> z : B_unit | TRUE \<rbrace>) \<lesssim> \<tau> \<and> (u,\<tau>') \<in> setD \<Delta> \<and> \<Theta> ; {||} ; GNil \<turnstile> v \<Leftarrow> \<tau>' \<and> \<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f \<Delta>"
+ using check_s_elims(8) by metis
+ hence "\<Theta> \<turnstile> update_d \<delta> u v \<sim> \<Delta>" using update_d_sim * by metis
+ moreover have "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_val (V_lit L_unit ) \<Leftarrow> \<tau>" using zt * check_s_v_unit check_s_wf
+ by auto
+ ultimately show ?case using config_typeI * by (meson order_refl)
+next
+ case (reduce_seq1I \<Phi> \<delta> s)
+ hence "\<Theta> ; \<Phi> ; {||} ; GNil ; \<Delta> \<turnstile> s \<Leftarrow> \<tau> \<and> \<Theta> \<turnstile> \<delta> \<sim> \<Delta> \<and> (\<forall>fd\<in>set \<Phi>. check_fundef \<Theta> \<Phi> fd)"
+ using check_s_elims config_type_elims by force
+ then show ?case using config_typeI by blast
+next
+ case (reduce_seq2I s1 v \<Phi> \<delta> \<delta>' s1' s2)
+ hence tt: "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_seq s1 s2 \<Leftarrow> \<tau> \<and> \<Theta> \<turnstile> \<delta> \<sim> \<Delta> \<and> (\<forall>fd\<in>set \<Phi>. check_fundef \<Theta> \<Phi> fd)"
+ using config_type_elims by blast
+ then obtain z where zz: "\<Theta> ; \<Phi> ; {||} ; GNil; \<Delta> \<turnstile> s1 \<Leftarrow> (\<lbrace> z : B_unit | TRUE \<rbrace>) \<and> \<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> s2 \<Leftarrow> \<tau>"
+ using check_s_elims by blast
+ hence "\<Theta>; \<Phi>; \<Delta> \<turnstile> \<langle> \<delta> , s1 \<rangle> \<Leftarrow> (\<lbrace> z : B_unit | TRUE \<rbrace>)"
+ using tt config_typeI tt by simp
+ then obtain \<Delta>' where *: "\<Theta>; \<Phi>; \<Delta>' \<turnstile> \<langle> \<delta>' , s1' \<rangle> \<Leftarrow> (\<lbrace> z : B_unit | TRUE \<rbrace>) \<and> \<Delta> \<sqsubseteq> \<Delta>'"
+ using reduce_seq2I by meson
+ moreover hence s't: "\<Theta>; \<Phi>; {||}; GNil; \<Delta>' \<turnstile> s1' \<Leftarrow> (\<lbrace> z : B_unit | TRUE \<rbrace>) \<and> \<Theta> \<turnstile> \<delta>' \<sim> \<Delta>'"
+ using config_type_elims by force
+ moreover hence "\<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f \<Delta>'" using check_s_wf by meson
+ moreover hence "\<Theta>; \<Phi>; {||}; GNil; \<Delta>' \<turnstile> s2 \<Leftarrow> \<tau>"
+ using calculation(1) zz check_s_d_weakening * by metis
+ moreover hence "\<Theta>; \<Phi>; {||}; GNil; \<Delta>' \<turnstile> (AS_seq s1' s2) \<Leftarrow> \<tau>"
+ using check_seqI zz s't by meson
+ ultimately have "\<Theta>; \<Phi>; \<Delta>' \<turnstile> \<langle> \<delta>' , AS_seq s1' s2 \<rangle> \<Leftarrow> \<tau> \<and> \<Delta> \<sqsubseteq> \<Delta>'"
+ using zz config_typeI tt by meson
+ then show ?case by meson
+next
+ case (reduce_whileI x s1 s2 z' \<Phi> \<delta> )
+
+ hence *: "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_while s1 s2 \<Leftarrow> \<tau> \<and> \<Theta> \<turnstile> \<delta> \<sim> \<Delta> \<and> (\<forall>fd\<in>set \<Phi>. check_fundef \<Theta> \<Phi> fd)"
+ using config_type_elims by meson
+
+ hence **:"\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_while s1 s2 \<Leftarrow> \<tau>" by auto
+ hence "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let2 x (\<lbrace> z' : B_bool | TRUE \<rbrace>) s1 (AS_if (V_var x) (AS_seq s2 (AS_while s1 s2)) (AS_val (V_lit L_unit)) ) \<Leftarrow> \<tau>"
+ using check_while reduce_whileI by auto
+ thus ?case using config_typeI * by (meson subset_refl)
+
+next
+ case (reduce_caseI dc x' s' css \<Phi> \<delta> tyid v)
+
+ hence **: "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_match (V_cons tyid dc v) css \<Leftarrow> \<tau> \<and> \<Theta> \<turnstile> \<delta> \<sim> \<Delta> \<and> (\<forall>fd\<in>set \<Phi>. check_fundef \<Theta> \<Phi> fd)"
+ using config_type_elims[OF reduce_caseI(2)] by metis
+ hence ***: "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_match (V_cons tyid dc v) css \<Leftarrow> \<tau>" by auto
+
+ let ?vcons = "V_cons tyid dc v"
+
+ obtain dclist tid and z::x where cv: "\<Theta>; {||} ; GNil \<turnstile> (V_cons tyid dc v) \<Leftarrow> (\<lbrace> z : B_id tid | TRUE \<rbrace>) \<and>
+ \<Theta>; \<Phi>; {||}; GNil; \<Delta> ; tid ; dclist ; (V_cons tyid dc v) \<turnstile> css \<Leftarrow> \<tau> \<and> AF_typedef tid dclist \<in> set \<Theta> \<and>
+ \<Theta> ; {||} ; GNil \<turnstile> V_cons tyid dc v \<Leftarrow> \<lbrace> z : B_id tid | TRUE \<rbrace>"
+ using check_s_elims(9)[OF ***] by metis
+
+ hence vi:" \<Theta> ; {||} ; GNil \<turnstile> V_cons tyid dc v \<Leftarrow> \<lbrace> z : B_id tid | TRUE \<rbrace>" by auto
+ obtain tcons where vi2:" \<Theta> ; {||} ; GNil \<turnstile> V_cons tyid dc v \<Rightarrow> tcons \<and> \<Theta> ; {||} ; GNil \<turnstile> tcons \<lesssim> \<lbrace> z : B_id tid | TRUE \<rbrace>"
+ using check_v_elims(1)[OF vi] by metis
+ hence vi1: "\<Theta> ; {||} ; GNil \<turnstile> V_cons tyid dc v \<Rightarrow> tcons" by auto
+
+ show ?case proof(rule infer_v_elims(4)[OF vi1],goal_cases)
+ case (1 dclist2 tc tv z2)
+ have "tyid = tid" using \<tau>.eq_iff using subtype_eq_base vi2 1 by fastforce
+ hence deq:"dclist = dclist2" using check_v_wf wfX_wfY cv 1 wfTh_dclist_unique by metis
+ have "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> s'[x'::=v]\<^sub>s\<^sub>v \<Leftarrow> \<tau>" proof(rule check_match(3))
+ show \<open> \<Theta> ; \<Phi> ; {||} ; GNil ; \<Delta> ; tyid ; dclist ; ?vcons \<turnstile> css \<Leftarrow> \<tau>\<close> using \<open>tyid = tid\<close> cv by auto
+ show "distinct (map fst dclist)" using wfTh_dclist_distinct check_v_wf wfX_wfY cv by metis
+ show \<open>?vcons = V_cons tyid dc v\<close> by auto
+ show \<open>{||} = {||}\<close> by auto
+ show \<open>(dc, tc) \<in> set dclist\<close> using 1 deq by auto
+ show \<open>GNil = GNil\<close> by auto
+ show \<open>Some (AS_branch dc x' s') = lookup_branch dc css\<close> using reduce_caseI by auto
+ show \<open>\<Theta> ; {||} ; GNil \<turnstile> v \<Leftarrow> tc\<close> using 1 check_v.intros by auto
+ qed
+ thus ?case using config_typeI ** by blast
+ qed
+
+next
+ case (reduce_let_fstI \<Phi> \<delta> x v1 v2 s)
+ thus ?case using preservation_fst_snd order_refl by metis
+next
+ case (reduce_let_sndI \<Phi> \<delta> x v1 v2 s)
+ thus ?case using preservation_fst_snd order_refl by metis
+next
+ case (reduce_let_concatI \<Phi> \<delta> x v1 v2 s)
+ hence elim: "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let x (AE_concat (V_lit (L_bitvec v1)) (V_lit (L_bitvec v2))) s \<Leftarrow> \<tau> \<and>
+ \<Theta> \<turnstile> \<delta> \<sim> \<Delta> \<and> (\<forall>fd\<in>set \<Phi>. check_fundef \<Theta> \<Phi> fd)"
+ using config_type_elims by metis
+
+ obtain z::x where z: "atom z \<sharp> (AE_concat (V_lit (L_bitvec v1)) (V_lit (L_bitvec v2)), GNil, CE_val (V_lit (L_bitvec (v1 @ v2))))"
+ using obtain_fresh by metis
+
+ have "\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f GNil" using check_s_wf elim by auto
+
+ have "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let x (AE_val (V_lit (L_bitvec (v1 @ v2)))) s \<Leftarrow> \<tau>" proof(rule subtype_let)
+ show \<open> \<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let x (AE_concat (V_lit (L_bitvec v1)) (V_lit (L_bitvec v2))) s \<Leftarrow> \<tau>\<close> using elim by auto
+ show \<open>\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> (AE_concat (V_lit (L_bitvec v1)) (V_lit (L_bitvec v2))) \<Rightarrow> \<lbrace> z : B_bitvec | CE_val (V_var z) == (CE_concat ([V_lit (L_bitvec v1)]\<^sup>c\<^sup>e) ([V_lit (L_bitvec v2)]\<^sup>c\<^sup>e))\<rbrace> \<close>
+ (is "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> ?e1 \<Rightarrow> ?t1")
+ proof
+ show \<open> \<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using check_s_wf elim by auto
+ show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using check_s_wf elim by auto
+ show \<open> \<Theta> ; {||} ; GNil \<turnstile> V_lit (L_bitvec v1) \<Rightarrow> \<lbrace> z : B_bitvec | CE_val (V_var z) == CE_val (V_lit (L_bitvec v1)) \<rbrace>\<close>
+ using infer_v_litI infer_l.intros \<open>\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f GNil\<close> fresh_GNil by auto
+ show \<open> \<Theta> ; {||} ; GNil \<turnstile> V_lit (L_bitvec v2) \<Rightarrow> \<lbrace> z : B_bitvec | CE_val (V_var z) == CE_val (V_lit (L_bitvec v2)) \<rbrace>\<close>
+ using infer_v_litI infer_l.intros \<open>\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f GNil\<close> fresh_GNil by auto
+ show \<open>atom z \<sharp> AE_concat (V_lit (L_bitvec v1)) (V_lit (L_bitvec v2))\<close> using z fresh_Pair by metis
+ show \<open>atom z \<sharp> GNil\<close> using z fresh_Pair by auto
+ qed
+ show \<open>\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AE_val (V_lit (L_bitvec (v1 @ v2))) \<Rightarrow> \<lbrace> z : B_bitvec | CE_val (V_var z) == CE_val (V_lit (L_bitvec (v1 @ v2))) \<rbrace> \<close>
+ (is "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> ?e2 \<Rightarrow> ?t2")
+ using infer_e_valI infer_v_litI infer_l.intros \<open>\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f GNil\<close> fresh_GNil check_s_wf elim by metis
+ show \<open>\<Theta> ; {||} ; GNil \<turnstile> ?t2 \<lesssim> ?t1\<close> using subtype_concat check_s_wf elim by auto
+ qed
+
+ thus ?case using config_typeI elim by (meson order_refl)
+next
+ case (reduce_let_lenI \<Phi> \<delta> x v s)
+ hence elim: "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let x (AE_len (V_lit (L_bitvec v))) s \<Leftarrow> \<tau> \<and> \<Theta> \<turnstile> \<delta> \<sim> \<Delta> \<and> (\<forall>fd\<in>set \<Phi>. check_fundef \<Theta> \<Phi> fd)"
+ using check_s_elims config_type_elims by metis
+
+ then obtain t where t: "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AE_len (V_lit (L_bitvec v)) \<Rightarrow> t" using check_s_elims by meson
+
+ moreover then obtain z::x where "t = \<lbrace> z : B_int | CE_val (V_var z) == CE_len ([V_lit (L_bitvec v)]\<^sup>c\<^sup>e) \<rbrace>" using infer_e_elims by meson
+
+ moreover obtain z'::x where "atom z' \<sharp> v" using obtain_fresh by metis
+ moreover have "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AE_val (V_lit (L_num (int (length v)))) \<Rightarrow> \<lbrace> z' : B_int | CE_val (V_var z') == CE_val (V_lit (L_num (int (length v)))) \<rbrace>"
+ using infer_e_valI infer_v_litI infer_l.intros(3) t check_s_wf elim
+ by (metis infer_l_form2 type_for_lit.simps(3))
+
+ moreover have "\<Theta> ; {||} ; GNil \<turnstile> \<lbrace> z' : B_int | CE_val (V_var z') == CE_val (V_lit (L_num (int (length v)))) \<rbrace> \<lesssim>
+ \<lbrace> z : B_int | CE_val (V_var z) == CE_len [(V_lit (L_bitvec v))]\<^sup>c\<^sup>e \<rbrace>" using subtype_len check_s_wf elim by auto
+
+ ultimately have "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let x (AE_val (V_lit (L_num (int (length v))))) s \<Leftarrow> \<tau>" using subtype_let by (meson elim)
+ thus ?case using config_typeI elim by (meson order_refl)
+next
+ case (reduce_let_splitI n v v1 v2 \<Phi> \<delta> x s)
+ hence elim: "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let x (AE_split (V_lit (L_bitvec v)) (V_lit (L_num n))) s \<Leftarrow> \<tau> \<and>
+ \<Theta> \<turnstile> \<delta> \<sim> \<Delta> \<and> (\<forall>fd\<in>set \<Phi>. check_fundef \<Theta> \<Phi> fd)"
+ using config_type_elims by metis
+
+ obtain z::x where z: "atom z \<sharp> (AE_split (V_lit (L_bitvec v)) (V_lit (L_num n)), GNil, CE_val (V_lit (L_bitvec (v1 @ v2))),
+([ L_bitvec v1 ]\<^sup>v, [ L_bitvec v2 ]\<^sup>v), \<Theta>, {||}::bv fset)"
+ using obtain_fresh by metis
+
+ have *:"\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f GNil" using check_s_wf elim by auto
+
+ have "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let x (AE_val (V_pair (V_lit (L_bitvec v1)) (V_lit (L_bitvec v2)))) s \<Leftarrow> \<tau>" proof(rule subtype_let)
+
+ show \<open> \<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_let x (AE_split (V_lit (L_bitvec v)) (V_lit (L_num n))) s \<Leftarrow> \<tau>\<close> using elim by auto
+ show \<open>\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> (AE_split (V_lit (L_bitvec v)) (V_lit (L_num n))) \<Rightarrow> \<lbrace> z : B_pair B_bitvec B_bitvec
+ | ((CE_val (V_lit (L_bitvec v))) == (CE_concat (CE_fst (CE_val (V_var z))) (CE_snd (CE_val (V_var z)))))
+ AND (((CE_len (CE_fst (CE_val (V_var z))))) == (CE_val (V_lit (L_num n)))) \<rbrace> \<close>
+ (is "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> ?e1 \<Rightarrow> ?t1")
+ proof
+ show \<open> \<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using check_s_wf elim by auto
+ show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using check_s_wf elim by auto
+ show \<open> \<Theta> ; {||} ; GNil \<turnstile> V_lit (L_bitvec v) \<Rightarrow> \<lbrace> z : B_bitvec | CE_val (V_var z) == CE_val (V_lit (L_bitvec v)) \<rbrace>\<close>
+ using infer_v_litI infer_l.intros \<open>\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f GNil\<close> fresh_GNil by auto
+ show "\<Theta> ; {||} ; GNil \<turnstile> ([ L_num
+ n ]\<^sup>v) \<Leftarrow> \<lbrace> z : B_int | (([ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e) == ([ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e)) AND [ leq [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e [| [ [ L_bitvec
+ v ]\<^sup>v ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrace>" using split_n reduce_let_splitI check_v_num_leq * wfX_wfY by metis
+ show \<open>atom z \<sharp> AE_split [ L_bitvec v ]\<^sup>v [ L_num n ]\<^sup>v\<close> using z fresh_Pair by auto
+ show \<open>atom z \<sharp> GNil\<close> using z fresh_Pair by auto
+ show \<open>atom z \<sharp> AE_split [ L_bitvec v ]\<^sup>v [ L_num n ]\<^sup>v\<close> using z fresh_Pair by auto
+ show \<open>atom z \<sharp> GNil\<close> using z fresh_Pair by auto
+ show \<open>atom z \<sharp> AE_split [ L_bitvec v ]\<^sup>v [ L_num n ]\<^sup>v\<close> using z fresh_Pair by auto
+ show \<open>atom z \<sharp> GNil\<close> using z fresh_Pair by auto
+ qed
+
+ show \<open>\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AE_val (V_pair (V_lit (L_bitvec v1)) (V_lit (L_bitvec v2))) \<Rightarrow> \<lbrace> z : B_pair B_bitvec B_bitvec | CE_val (V_var z) == CE_val ((V_pair (V_lit (L_bitvec v1)) (V_lit (L_bitvec v2)))) \<rbrace> \<close>
+ (is "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> ?e2 \<Rightarrow> ?t2")
+ apply(rule infer_e_valI)
+ using check_s_wf elim apply metis
+ using check_s_wf elim apply metis
+ apply(rule infer_v_pair2I)
+ using z fresh_prodN apply metis
+ using z fresh_GNil fresh_prodN apply metis
+ using infer_v_litI infer_l.intros \<open>\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f GNil\<close> b_of.simps apply blast+
+ using b_of.simps apply simp+
+ done
+ show \<open>\<Theta> ; {||} ; GNil \<turnstile> ?t2 \<lesssim> ?t1\<close> using subtype_split check_s_wf elim reduce_let_splitI by auto
+ qed
+
+ thus ?case using config_typeI elim by (meson order_refl)
+next
+ case (reduce_assert1I \<Phi> \<delta> c v)
+
+ hence elim: "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_assert c [v]\<^sup>s \<Leftarrow> \<tau> \<and>
+ \<Theta> \<turnstile> \<delta> \<sim> \<Delta> \<and> (\<forall>fd\<in>set \<Phi>. check_fundef \<Theta> \<Phi> fd)"
+ using config_type_elims reduce_assert1I by metis
+ hence *:"\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_assert c [v]\<^sup>s \<Leftarrow> \<tau>" by auto
+
+ have "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> [v]\<^sup>s \<Leftarrow> \<tau>" using check_assert_s * by metis
+ thus ?case using elim config_typeI by blast
+next
+ case (reduce_assert2I \<Phi> \<delta> s \<delta>' s' c)
+
+ hence elim: "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_assert c s \<Leftarrow> \<tau> \<and>
+ \<Theta> \<turnstile> \<delta> \<sim> \<Delta> \<and> (\<forall>fd\<in>set \<Phi>. check_fundef \<Theta> \<Phi> fd)"
+ using config_type_elims by metis
+ hence *:"\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> AS_assert c s \<Leftarrow> \<tau>" by auto
+
+ have cv: "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> s \<Leftarrow> \<tau> \<and> \<Theta> ; {||} ; GNil \<Turnstile> c " using check_assert_s * by metis
+
+ hence "\<Theta>; \<Phi>; \<Delta> \<turnstile> \<langle>\<delta>, s\<rangle> \<Leftarrow> \<tau>" using elim config_typeI by simp
+ then obtain \<Delta>' where D: "\<Theta>; \<Phi>; \<Delta>' \<turnstile> \<langle> \<delta>' , s' \<rangle> \<Leftarrow> \<tau> \<and> \<Delta> \<sqsubseteq> \<Delta>'" using reduce_assert2I by metis
+ hence **:"\<Theta>; \<Phi>; {||}; GNil; \<Delta>' \<turnstile> s' \<Leftarrow> \<tau> \<and> \<Theta> \<turnstile> \<delta>' \<sim> \<Delta>'" using config_type_elims by metis
+
+ obtain x::x where x:"atom x \<sharp> (\<Theta>, \<Phi>, ({||}::bv fset), GNil, \<Delta>', c, \<tau>, s')" using obtain_fresh by metis
+
+ have *:"\<Theta>; \<Phi>; {||}; GNil; \<Delta>' \<turnstile> AS_assert c s' \<Leftarrow> \<tau>" proof
+ show "atom x \<sharp> (\<Theta>, \<Phi>, {||}, GNil, \<Delta>', c, \<tau>, s')" using x by auto
+ have "\<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f c" using * check_s_wf by auto
+ hence wfg:"\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f (x, B_bool, c) #\<^sub>\<Gamma> GNil" using wfC_wfG wfB_boolI check_s_wf * fresh_GNil by auto
+ moreover have cs: "\<Theta>; \<Phi>; {||}; GNil; \<Delta>' \<turnstile> s' \<Leftarrow> \<tau>" using ** by auto
+ ultimately show "\<Theta> ; \<Phi> ; {||} ; (x, B_bool, c) #\<^sub>\<Gamma> GNil ; \<Delta>' \<turnstile> s' \<Leftarrow> \<tau>" using check_s_g_weakening(1)[OF cs _ wfg] toSet.simps by simp
+ show "\<Theta> ; {||} ; GNil \<Turnstile> c " using cv by auto
+ show "\<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f \<Delta>' " using check_s_wf ** by auto
+ qed
+
+ thus ?case using elim config_typeI D ** by metis
+qed
+
+lemma preservation_many:
+ assumes " \<Phi> \<turnstile> \<langle>\<delta>, s\<rangle> \<longrightarrow>\<^sup>* \<langle> \<delta>' , s' \<rangle>"
+ shows "\<Theta>; \<Phi>; \<Delta> \<turnstile> \<langle>\<delta>, s\<rangle> \<Leftarrow> \<tau> \<Longrightarrow> \<exists>\<Delta>'. \<Theta>; \<Phi>; \<Delta>' \<turnstile> \<langle> \<delta>' , s' \<rangle> \<Leftarrow> \<tau> \<and> \<Delta> \<sqsubseteq> \<Delta>'"
+ using assms proof(induct arbitrary: \<Delta> rule: reduce_stmt_many.induct)
+ case (reduce_stmt_many_oneI \<Phi> \<delta> s \<delta>' s')
+ then show ?case using preservation by simp
+next
+ case (reduce_stmt_many_manyI \<Phi> \<delta> s \<delta>' s' \<delta>'' s'')
+ then show ?case using preservation subset_trans by metis
+qed
+
+section \<open>Progress\<close>
+text \<open>We prove that a well typed program is either a value or we can make a step\<close>
+
+lemma check_let_op_infer:
+ assumes "\<Theta>; \<Phi>; {||}; \<Gamma>; \<Delta> \<turnstile> LET x = (AE_op opp v1 v2) IN s \<Leftarrow> \<tau>" and "supp ( LET x = (AE_op opp v1 v2) IN s) \<subseteq> atom`fst`setD \<Delta>"
+ shows "\<exists> z b c. \<Theta>; \<Phi>; {||}; \<Gamma>; \<Delta> \<turnstile> (AE_op opp v1 v2) \<Rightarrow> \<lbrace>z:b|c\<rbrace>"
+proof -
+ have xx: "\<Theta>; \<Phi>; {||}; \<Gamma>; \<Delta> \<turnstile> LET x = (AE_op opp v1 v2) IN s \<Leftarrow> \<tau>" using assms by simp
+ then show ?thesis using check_s_elims(2)[OF xx] by meson
+qed
+
+lemma infer_pair:
+ assumes "\<Theta> ; B; \<Gamma> \<turnstile> v \<Rightarrow> \<lbrace> z : B_pair b1 b2 | c \<rbrace>" and "supp v = {}"
+ obtains v1 and v2 where "v = V_pair v1 v2"
+ using assms proof(nominal_induct v rule: v.strong_induct)
+ case (V_lit x)
+ then show ?case by auto
+next
+ case (V_var x)
+ then show ?case using v.supp supp_at_base by auto
+next
+ case (V_pair x1a x2a)
+ then show ?case by auto
+next
+ case (V_cons x1a x2a x3)
+ then show ?case by auto
+next
+ case (V_consp x1a x2a x3 x4)
+ then show ?case by auto
+qed
+
+lemma progress_fst:
+ assumes "\<Theta>; \<Phi>; {||}; \<Gamma>; \<Delta> \<turnstile> LET x = (AE_fst v) IN s \<Leftarrow> \<tau>" and "\<Theta> \<turnstile> \<delta> \<sim> \<Delta>" and
+ "supp (LET x = (AE_fst v) IN s) \<subseteq> atom`fst`setD \<Delta>"
+ shows "\<exists>\<delta>' s'. \<Phi> \<turnstile> \<langle> \<delta> , LET x = (AE_fst v) IN s \<rangle> \<longrightarrow> \<langle>\<delta>', s'\<rangle>"
+proof -
+ have *:"supp v = {}" using assms s_branch_s_branch_list.supp by auto
+ obtain z and b and c where "\<Theta>; \<Phi>; {||}; \<Gamma>; \<Delta> \<turnstile> (AE_fst v ) \<Rightarrow> \<lbrace> z : b | c \<rbrace>"
+ using check_s_elims(2) using assms by meson
+ moreover obtain z' and b' and c' where "\<Theta> ; {||} ; \<Gamma> \<turnstile> v \<Rightarrow> \<lbrace> z' : B_pair b b' | c' \<rbrace>"
+ using infer_e_elims(8) using calculation by auto
+ moreover then obtain v1 and v2 where "V_pair v1 v2 = v"
+ using * infer_pair by metis
+ ultimately show ?thesis using reduce_let_fstI assms by metis
+qed
+
+lemma progress_let:
+ assumes "\<Theta>; \<Phi>; {||}; \<Gamma>; \<Delta> \<turnstile> LET x = e IN s \<Leftarrow> \<tau>" and "\<Theta> \<turnstile> \<delta> \<sim> \<Delta>" and
+ "supp (LET x = e IN s) \<subseteq> atom ` fst ` setD \<Delta>" and "sble \<Theta> \<Gamma>"
+ shows "\<exists>\<delta>' s'. \<Phi> \<turnstile> \<langle> \<delta> , LET x = e IN s\<rangle> \<longrightarrow> \<langle> \<delta>' , s'\<rangle>"
+proof -
+ obtain z b c where *: "\<Theta>; \<Phi>; {||}; \<Gamma>; \<Delta> \<turnstile> e \<Rightarrow> \<lbrace> z : b | c \<rbrace> " using check_s_elims(2)[OF assms(1)] by metis
+ have **: "supp e \<subseteq> atom ` fst ` setD \<Delta>" using assms s_branch_s_branch_list.supp by auto
+ from * ** assms show ?thesis proof(nominal_induct "\<lbrace> z : b | c \<rbrace>" rule: infer_e.strong_induct)
+ case (infer_e_valI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v)
+ then show ?case using reduce_stmt_elims reduce_let_valI by metis
+ next
+ case (infer_e_plusI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v1 z1 c1 v2 z2 c2 z3)
+ hence vf: "supp v1 = {} \<and> supp v2 = {}" by force
+ then obtain n1 and n2 where *: "v1 = V_lit (L_num n1) \<and> v2 = (V_lit (L_num n2))" using infer_int infer_e_plusI by metis
+ then show ?case using reduce_let_plusI * by metis
+ next
+ case (infer_e_leqI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v1 z1 c1 v2 z2 c2 z3)
+ hence vf: "supp v1 = {} \<and> supp v2 = {}" by force
+ then obtain n1 and n2 where *: "v1 = V_lit (L_num n1) \<and> v2 = (V_lit (L_num n2))" using infer_int infer_e_leqI by metis
+ then show ?case using reduce_let_leqI * by metis
+ next
+ case (infer_e_eqI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v1 z1 bb c1 v2 z2 c2 z3)
+ hence vf: "supp v1 = {} \<and> supp v2 = {}" by force
+ then obtain n1 and n2 where *: "v1 = V_lit n1 \<and> v2 = (V_lit n2)" using infer_lit infer_e_eqI by metis
+ then show ?case using reduce_let_eqI by blast
+ next
+ case (infer_e_appI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> f x b c \<tau>' s' v)
+ then show ?case using reduce_let_appI by metis
+ next
+ case (infer_e_appPI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> b' f bv x b c \<tau>' s' v)
+ then show ?case using reduce_let_appPI by metis
+ next
+ case (infer_e_fstI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v z' b2 c z)
+ hence "supp v = {}" by force
+ then obtain v1 and v2 where "v = V_pair v1 v2" using infer_e_fstI infer_pair by metis
+ then show ?case using reduce_let_fstI * by metis
+ next
+ case (infer_e_sndI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v z' b1 c z)
+ hence "supp v = {}" by force
+ then obtain v1 and v2 where "v = V_pair v1 v2" using infer_e_sndI infer_pair by metis
+ then show ?case using reduce_let_sndI * by metis
+ next
+ case (infer_e_lenI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v z' c za)
+ hence "supp v = {}" by force
+ then obtain bvec where "v = V_lit (L_bitvec bvec)" using infer_e_lenI infer_bitvec by metis
+ then show ?case using reduce_let_lenI * by metis
+ next
+ case (infer_e_mvarI \<Theta> \<B> \<Gamma> \<Phi> \<Delta> u)
+ hence "(u, \<lbrace> z : b | c \<rbrace>) \<in> setD \<Delta>" using infer_e_elims(10) by meson
+ then obtain v where "(u,v) \<in> set \<delta>" using infer_e_mvarI delta_sim_delta_lookup by meson
+ then show ?case using reduce_let_mvar by metis
+ next
+ case (infer_e_concatI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v1 z1 c1 v2 z2 c2 z3)
+ hence vf: "supp v1 = {} \<and> supp v2 = {}" by force
+ then obtain n1 and n2 where *: "v1 = V_lit (L_bitvec n1) \<and> v2 = (V_lit (L_bitvec n2))" using infer_bitvec infer_e_concatI by metis
+ then show ?case using reduce_let_concatI * by metis
+ next
+ case (infer_e_splitI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v1 z1 c1 v2 z2 z3)
+ hence vf: "supp v1 = {} \<and> supp v2 = {}" by force
+ then obtain n1 and n2 where *: "v1 = V_lit (L_bitvec n1) \<and> v2 = (V_lit (L_num n2))" using infer_bitvec infer_e_splitI check_int by metis
+
+ have "0 \<le> n2 \<and> n2 \<le> int (length n1)" using check_v_range[OF _ * ] infer_e_splitI by simp
+ then obtain bv1 and bv2 where "split n2 n1 (bv1 , bv2)" using obtain_split by metis
+ then show ?case using reduce_let_splitI * by metis
+ qed
+qed
+
+lemma check_css_lookup_branch_exist:
+ fixes s::s and cs::branch_s and css::branch_list and v::v
+ shows
+ "\<Theta>; \<Phi>; B; G; \<Delta> \<turnstile> s \<Leftarrow> \<tau> \<Longrightarrow> True" and
+ "check_branch_s \<Theta> \<Phi> {||} GNil \<Delta> tid dc const v cs \<tau> \<Longrightarrow> True" and
+ "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta>; tid; dclist; v \<turnstile> css \<Leftarrow> \<tau> \<Longrightarrow> (dc, t) \<in> set dclist \<Longrightarrow>
+ \<exists>x' s'. Some (AS_branch dc x' s') = lookup_branch dc css"
+proof(nominal_induct \<tau> and \<tau> and \<tau> rule: check_s_check_branch_s_check_branch_list.strong_induct)
+ case (check_branch_list_consI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid cons const v cs \<tau> dclist css)
+ then show ?case using lookup_branch.simps check_branch_list_finalI by force
+next
+ case (check_branch_list_finalI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid cons const v cs \<tau>)
+ then show ?case using lookup_branch.simps check_branch_list_finalI by force
+qed(auto+)
+
+lemma progress_aux:
+ shows "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> s \<Leftarrow> \<tau> \<Longrightarrow> \<B> = {||} \<Longrightarrow> sble \<Theta> \<Gamma> \<Longrightarrow> supp s \<subseteq> atom ` fst ` setD \<Delta> \<Longrightarrow> \<Theta> \<turnstile> \<delta> \<sim> \<Delta> \<Longrightarrow>
+ (\<exists>v. s = [v]\<^sup>s) \<or> (\<exists>\<delta>' s'. \<Phi> \<turnstile> \<langle>\<delta>, s\<rangle> \<longrightarrow> \<langle>\<delta>', s'\<rangle>)" and
+ "\<Theta>; \<Phi>; {||}; \<Gamma>; \<Delta>; tid; dc; const; v2 \<turnstile> cs \<Leftarrow> \<tau> \<Longrightarrow> supp cs = {} \<Longrightarrow> True "
+ "\<Theta>; \<Phi>; {||}; \<Gamma>; \<Delta>; tid; dclist; v2 \<turnstile> css \<Leftarrow> \<tau> \<Longrightarrow> supp css = {} \<Longrightarrow> True "
+proof(induct rule: check_s_check_branch_s_check_branch_list.inducts)
+ case (check_valI \<Delta> \<Theta> \<Gamma> v \<tau>' \<tau>)
+ then show ?case by auto
+next
+ case (check_letI x \<Theta> \<Phi> \<B> \<Gamma> \<Delta> e \<tau> z s b c)
+ hence "\<Theta>; \<Phi>; {||}; \<Gamma>; \<Delta> \<turnstile> AS_let x e s \<Leftarrow> \<tau>" using Typing.check_letI by meson
+ then show ?case using progress_let check_letI by metis
+next
+ case (check_branch_s_branchI \<Theta> \<B> \<Gamma> \<Delta> \<tau> const x \<Phi> tid cons v s)
+ then show ?case by auto
+next
+ case (check_branch_list_consI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist v cs \<tau> css)
+ then show ?case by auto
+next
+ case (check_branch_list_finalI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist v cs \<tau>)
+ then show ?case by auto
+next
+ case (check_ifI z \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v s1 s2 \<tau>)
+ have "supp v = {}" using check_ifI s_branch_s_branch_list.supp by auto
+ hence "v = V_lit L_true \<or> v = V_lit L_false" using check_bool_options check_ifI by auto
+ then show ?case using reduce_if_falseI reduce_if_trueI check_ifI by meson
+next
+ case (check_let2I x \<Theta> \<Phi> \<B> G \<Delta> t s1 \<tau> s2 )
+ then consider " (\<exists>v. s1 = AS_val v)" | "(\<exists>\<delta>' a. \<Phi> \<turnstile> \<langle>\<delta>, s1\<rangle> \<longrightarrow> \<langle>\<delta>', a\<rangle>)" by auto
+ then show ?case proof(cases)
+ case 1
+ then show ?thesis using reduce_let2_valI by fast
+ next
+ case 2
+ then show ?thesis using reduce_let2I check_let2I by meson
+ qed
+next
+ case (check_varI u \<Theta> \<Phi> \<B> \<Gamma> \<Delta> \<tau>' v \<tau> s)
+
+ obtain uu::u where uf: "atom uu \<sharp> (u,\<delta>,s) " using obtain_fresh by blast
+ obtain sa where " (uu \<leftrightarrow> u ) \<bullet> s = sa" by presburger
+ moreover have "atom uu \<sharp> s" using uf fresh_prod3 by auto
+ ultimately have "AS_var uu \<tau>' v sa = AS_var u \<tau>' v s" using s_branch_s_branch_list.eq_iff(7) Abs1_eq_iff(3)[of uu sa u s] by auto
+
+ moreover have "atom uu \<sharp> \<delta>" using uf fresh_prod3 by auto
+ ultimately have "\<Phi> \<turnstile> \<langle>\<delta>, AS_var u \<tau>' v s\<rangle> \<longrightarrow> \<langle>(uu, v) # \<delta>, sa\<rangle>"
+ using reduce_varI uf by metis
+ then show ?case by auto
+next
+ case (check_assignI \<Delta> u \<tau> P G v z \<tau>')
+ then show ?case using reduce_assignI by blast
+next
+ case (check_whileI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> s1 z s2 \<tau>')
+ obtain x::x where "atom x \<sharp> (s1,s2)" using obtain_fresh by metis
+ moreover obtain z::x where "atom z \<sharp> x" using obtain_fresh by metis
+ ultimately show ?case using reduce_whileI by fast
+next
+ case (check_seqI P \<Phi> \<B> G \<Delta> s1 z s2 \<tau>)
+ thus ?case proof(cases "\<exists>v. s1 = AS_val v")
+ case True
+ then obtain v where v: "s1 = AS_val v" by blast
+ hence "supp v = {}" using check_seqI by auto
+ have "\<exists>z1 c1. P; \<B>; G \<turnstile> v \<Rightarrow> (\<lbrace> z1 : B_unit | c1 \<rbrace>)" proof -
+ obtain t where t:"P; \<B>; G \<turnstile> v \<Rightarrow> t \<and> P; \<B> ; G \<turnstile> t \<lesssim> (\<lbrace> z : B_unit | TRUE \<rbrace>)"
+ using v check_seqI(1) check_s_elims(1) by blast
+ obtain z1 and b1 and c1 where teq: "t = (\<lbrace> z1 : b1 | c1 \<rbrace>)" using obtain_fresh_z by meson
+ hence "b1 = B_unit" using subtype_eq_base t by meson
+ thus ?thesis using t teq by fast
+ qed
+ then obtain z1 and c1 where "P ; \<B> ; G \<turnstile> v \<Rightarrow> (\<lbrace> z1 : B_unit | c1 \<rbrace>)" by auto
+ hence "v = V_lit L_unit" using infer_v_unit_form \<open>supp v = {}\<close> by simp
+ hence "s1 = AS_val (V_lit L_unit)" using v by auto
+ then show ?thesis using check_seqI reduce_seq1I by meson
+ next
+ case False
+ then show ?thesis using check_seqI reduce_seq2I
+ by (metis Un_subset_iff s_branch_s_branch_list.supp(9))
+ qed
+
+next
+ case (check_caseI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist v cs \<tau> z)
+ hence "supp v = {}" by auto
+
+ then obtain v' and dc and t::\<tau> where v: "v = V_cons tid dc v' \<and> (dc, t) \<in> set dclist"
+ using check_v_tid_form check_caseI by metis
+ obtain z and b and c where teq: "t = (\<lbrace> z : b | c \<rbrace>)" using obtain_fresh_z by meson
+
+ moreover then obtain x' s' where "Some (AS_branch dc x' s') = lookup_branch dc cs" using v teq check_caseI check_css_lookup_branch_exist by metis
+ ultimately show ?case using reduce_caseI v check_caseI dc_of.cases by metis
+next
+ case (check_assertI x \<Theta> \<Phi> \<B> \<Gamma> \<Delta> c \<tau> s)
+ hence sps: "supp s \<subseteq> atom ` fst ` setD \<Delta>" by auto
+ have "atom x \<sharp> c " using check_assertI by auto
+ have "atom x \<sharp> \<Gamma>" using check_assertI check_s_wf wfG_elims by metis
+ have "sble \<Theta> ((x, B_bool, c) #\<^sub>\<Gamma> \<Gamma>)" proof -
+ obtain i' where i':" i' \<Turnstile> \<Gamma> \<and> \<Theta>; \<Gamma> \<turnstile> i'" using check_assertI sble_def by metis
+ obtain i::valuation where i:"i = i' ( x \<mapsto> SBool True)" by auto
+
+ have "i \<Turnstile> (x, B_bool, c) #\<^sub>\<Gamma> \<Gamma>" proof -
+ have "i' \<Turnstile> c" using valid.simps i' check_assertI by metis
+ hence "i \<Turnstile> c" using is_satis_weakening_x i \<open>atom x \<sharp> c\<close> by auto
+ moreover have "i \<Turnstile> \<Gamma>" using is_satis_g_weakening_x i' i check_assertI \<open>atom x \<sharp> \<Gamma>\<close> by metis
+ ultimately show ?thesis using is_satis_g.simps i by auto
+ qed
+ moreover have "\<Theta> ; ((x, B_bool, c) #\<^sub>\<Gamma> \<Gamma>) \<turnstile> i" proof(rule wfI_cons)
+ show \<open> i' \<Turnstile> \<Gamma> \<close> using i' by auto
+ show \<open> \<Theta> ; \<Gamma> \<turnstile> i'\<close> using i' by auto
+ show \<open>i = i'(x \<mapsto> SBool True)\<close> using i by auto
+ show \<open> \<Theta> \<turnstile> SBool True: B_bool\<close> using wfRCV_BBoolI by auto
+ show \<open>atom x \<sharp> \<Gamma>\<close> using check_assertI check_s_wf wfG_elims by auto
+ qed
+ ultimately show ?thesis using sble_def by auto
+ qed
+ then consider "(\<exists>v. s = [v]\<^sup>s)" | "(\<exists>\<delta>' a. \<Phi> \<turnstile> \<langle>\<delta>, s\<rangle> \<longrightarrow> \<langle> \<delta>', a\<rangle>)" using check_assertI sps by metis
+ hence "(\<exists>\<delta>' a. \<Phi> \<turnstile> \<langle>\<delta>, ASSERT c IN s\<rangle> \<longrightarrow> \<langle>\<delta>', a\<rangle>)" proof(cases)
+ case 1
+ then show ?thesis using reduce_assert1I by metis
+ next
+ case 2
+ then show ?thesis using reduce_assert2I by metis
+ qed
+ thus ?case by auto
+qed
+
+lemma progress:
+ assumes "\<Theta>; \<Phi>; \<Delta> \<turnstile> \<langle>\<delta>, s\<rangle> \<Leftarrow> \<tau>"
+ shows "(\<exists>v. s = [v]\<^sup>s) \<or> (\<exists>\<delta>' s'. \<Phi> \<turnstile> \<langle>\<delta>, s\<rangle> \<longrightarrow> \<langle>\<delta>', s'\<rangle>)"
+proof -
+ have "\<Theta>; \<Phi>; {||}; GNil; \<Delta> \<turnstile> s \<Leftarrow> \<tau>" and "\<Theta> \<turnstile> \<delta> \<sim> \<Delta>"
+ using config_type_elims[OF assms(1)] by auto+
+ moreover hence "supp s \<subseteq> atom ` fst ` setD \<Delta>" using check_s_wf wfS_supp by fastforce
+ moreover have "sble \<Theta> GNil" using sble_def wfI_def is_satis_g.simps by simp
+ ultimately show ?thesis using progress_aux by blast
+qed
+
+section \<open>Safety\<close>
+
+lemma safety_stmt:
+ assumes "\<Phi> \<turnstile> \<langle>\<delta>, s\<rangle> \<longrightarrow>\<^sup>* \<langle>\<delta>', s'\<rangle>" and "\<Theta>; \<Phi>; \<Delta> \<turnstile> \<langle>\<delta>, s\<rangle> \<Leftarrow> \<tau>"
+ shows "(\<exists>v. s' = [v]\<^sup>s) \<or> (\<exists>\<delta>'' s''. \<Phi> \<turnstile> \<langle>\<delta>', s'\<rangle> \<longrightarrow> \<langle>\<delta>'', s''\<rangle>)"
+ using preservation_many progress assms by meson
+
+lemma safety:
+ assumes "\<turnstile> \<langle>PROG \<Theta> \<Phi> \<G> s\<rangle> \<Leftarrow> \<tau>" and "\<Phi> \<turnstile> \<langle>\<delta>_of \<G>, s\<rangle> \<longrightarrow>\<^sup>* \<langle>\<delta>', s'\<rangle>"
+ shows "(\<exists>v. s' = [v]\<^sup>s) \<or> (\<exists>\<delta>'' s''. \<Phi> \<turnstile> \<langle>\<delta>', s'\<rangle> \<longrightarrow> \<langle>\<delta>'', s''\<rangle>)"
+ using assms config_type_prog_elims safety_stmt by metis
+
end
\ No newline at end of file
diff --git a/thys/MiniSail/SubstMethods.thy b/thys/MiniSail/SubstMethods.thy
--- a/thys/MiniSail/SubstMethods.thy
+++ b/thys/MiniSail/SubstMethods.thy
@@ -1,71 +1,71 @@
-theory SubstMethods
- (* Its seems that it's best to load the Eisbach tools last *)
- imports IVSubst WellformedL "HOL-Eisbach.Eisbach_Tools"
-begin
-
-text \<open>
- See Eisbach/Examples.thy as well as Eisbach User Manual.
-
-Freshness for various substitution situations. It seems that if undirected and we throw all the
-facts at them to try to solve in one shot, the automatic methods are *sometimes* unable
-to handle the different variants, so some guidance is needed.
-First we split into subgoals using fresh\_prodN and intro conjI
-
-The 'add', for example, will be induction premises that will contain freshness facts or freshness conditions from
-prior obtains
-
-Use different arguments for different things or just lump into one bucket\<close>
-
-method fresh_subst_mth_aux uses add = (
- (match conclusion in "atom z \<sharp> (\<Gamma>::\<Gamma>)[x::=v]\<^sub>\<Gamma>\<^sub>v" for z x v \<Gamma> \<Rightarrow> \<open>auto simp add: fresh_subst_gv_if[of "atom z" \<Gamma> v x] add\<close>)
- | (match conclusion in "atom z \<sharp> (v'::v)[x::=v]\<^sub>v\<^sub>v" for z x v v' \<Rightarrow> \<open>auto simp add: v.fresh fresh_subst_v_if pure_fresh subst_v_v_def add\<close> )
- | (match conclusion in "atom z \<sharp> (ce::ce)[x::=v]\<^sub>c\<^sub>e\<^sub>v" for z x v ce \<Rightarrow> \<open>auto simp add: fresh_subst_v_if subst_v_ce_def add\<close> )
- | (match conclusion in "atom z \<sharp> (\<Delta>::\<Delta>)[x::=v]\<^sub>\<Delta>\<^sub>v" for z x v \<Delta> \<Rightarrow> \<open>auto simp add: fresh_subst_v_if fresh_subst_dv_if add\<close> )
- | (match conclusion in "atom z \<sharp> \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>" for z x v \<Gamma>' \<Gamma> \<Rightarrow> \<open>metis add \<close> )
- | (match conclusion in "atom z \<sharp> (\<tau>::\<tau>)[x::=v]\<^sub>\<tau>\<^sub>v" for z x v \<tau> \<Rightarrow> \<open>auto simp add: v.fresh fresh_subst_v_if pure_fresh subst_v_\<tau>_def add\<close> )
- | (match conclusion in "atom z \<sharp> ({||} :: bv fset)" for z \<Rightarrow> \<open>auto simp add: fresh_empty_fset\<close>)
- (* tbc delta and types *)
- | (auto simp add: add x_fresh_b pure_fresh) (* Cases where there is no subst and so can most likely get what we want from induction premises *)
- )
-
-method fresh_mth uses add = (
- (unfold fresh_prodN, intro conjI)?,
- (fresh_subst_mth_aux add: add)+)
-
-
-notepad
-begin
- fix \<Gamma>::\<Gamma> and z::x and x::x and v::v and \<Theta>::\<Theta> and v'::v and w::x and tyid::string and dc::string and b::b and ce::ce and bv::bv
-
- assume as:"atom z \<sharp> (\<Gamma>,v',\<Theta>, v,w,ce) \<and> atom bv \<sharp> (\<Gamma>,v',\<Theta>, v,w,ce,b) "
-
- have "atom z \<sharp> \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v"
- by (fresh_mth add: as)
-
- hence "atom z \<sharp> v'[x::=v]\<^sub>v\<^sub>v"
- by (fresh_mth add: as)
-
- hence "atom z \<sharp> \<Gamma>"
- by (fresh_mth add: as)
-
- hence "atom z \<sharp> \<Theta>"
- by (fresh_mth add: as)
-
- hence "atom z \<sharp> (CE_val v == ce)[x::=v]\<^sub>c\<^sub>v"
- using as by auto
-
- hence "atom bv \<sharp> (CE_val v == ce)[x::=v]\<^sub>c\<^sub>v"
- using as by auto
-
- have "atom z \<sharp> (\<Theta>,\<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v,v'[x::=v]\<^sub>v\<^sub>v,w, V_pair v v, V_consp tyid dc b v, (CE_val v == ce)[x::=v]\<^sub>c\<^sub>v) "
- by (fresh_mth add: as)
-
- have "atom bv \<sharp> (\<Theta>,\<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v,v'[x::=v]\<^sub>v\<^sub>v,w, V_pair v v, V_consp tyid dc b v) "
- by (fresh_mth add: as)
-
-end
-
-
-
-
+theory SubstMethods
+ (* Its seems that it's best to load the Eisbach tools last *)
+ imports IVSubst WellformedL "HOL-Eisbach.Eisbach_Tools"
+begin
+
+text \<open>
+ See Eisbach/Examples.thy as well as Eisbach User Manual.
+
+Freshness for various substitution situations. It seems that if undirected and we throw all the
+facts at them to try to solve in one shot, the automatic methods are *sometimes* unable
+to handle the different variants, so some guidance is needed.
+First we split into subgoals using fresh\_prodN and intro conjI
+
+The 'add', for example, will be induction premises that will contain freshness facts or freshness conditions from
+prior obtains
+
+Use different arguments for different things or just lump into one bucket\<close>
+
+method fresh_subst_mth_aux uses add = (
+ (match conclusion in "atom z \<sharp> (\<Gamma>::\<Gamma>)[x::=v]\<^sub>\<Gamma>\<^sub>v" for z x v \<Gamma> \<Rightarrow> \<open>auto simp add: fresh_subst_gv_if[of "atom z" \<Gamma> v x] add\<close>)
+ | (match conclusion in "atom z \<sharp> (v'::v)[x::=v]\<^sub>v\<^sub>v" for z x v v' \<Rightarrow> \<open>auto simp add: v.fresh fresh_subst_v_if pure_fresh subst_v_v_def add\<close> )
+ | (match conclusion in "atom z \<sharp> (ce::ce)[x::=v]\<^sub>c\<^sub>e\<^sub>v" for z x v ce \<Rightarrow> \<open>auto simp add: fresh_subst_v_if subst_v_ce_def add\<close> )
+ | (match conclusion in "atom z \<sharp> (\<Delta>::\<Delta>)[x::=v]\<^sub>\<Delta>\<^sub>v" for z x v \<Delta> \<Rightarrow> \<open>auto simp add: fresh_subst_v_if fresh_subst_dv_if add\<close> )
+ | (match conclusion in "atom z \<sharp> \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>" for z x v \<Gamma>' \<Gamma> \<Rightarrow> \<open>metis add \<close> )
+ | (match conclusion in "atom z \<sharp> (\<tau>::\<tau>)[x::=v]\<^sub>\<tau>\<^sub>v" for z x v \<tau> \<Rightarrow> \<open>auto simp add: v.fresh fresh_subst_v_if pure_fresh subst_v_\<tau>_def add\<close> )
+ | (match conclusion in "atom z \<sharp> ({||} :: bv fset)" for z \<Rightarrow> \<open>auto simp add: fresh_empty_fset\<close>)
+ (* tbc delta and types *)
+ | (auto simp add: add x_fresh_b pure_fresh) (* Cases where there is no subst and so can most likely get what we want from induction premises *)
+ )
+
+method fresh_mth uses add = (
+ (unfold fresh_prodN, intro conjI)?,
+ (fresh_subst_mth_aux add: add)+)
+
+
+notepad
+begin
+ fix \<Gamma>::\<Gamma> and z::x and x::x and v::v and \<Theta>::\<Theta> and v'::v and w::x and tyid::string and dc::string and b::b and ce::ce and bv::bv
+
+ assume as:"atom z \<sharp> (\<Gamma>,v',\<Theta>, v,w,ce) \<and> atom bv \<sharp> (\<Gamma>,v',\<Theta>, v,w,ce,b) "
+
+ have "atom z \<sharp> \<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v"
+ by (fresh_mth add: as)
+
+ hence "atom z \<sharp> v'[x::=v]\<^sub>v\<^sub>v"
+ by (fresh_mth add: as)
+
+ hence "atom z \<sharp> \<Gamma>"
+ by (fresh_mth add: as)
+
+ hence "atom z \<sharp> \<Theta>"
+ by (fresh_mth add: as)
+
+ hence "atom z \<sharp> (CE_val v == ce)[x::=v]\<^sub>c\<^sub>v"
+ using as by auto
+
+ hence "atom bv \<sharp> (CE_val v == ce)[x::=v]\<^sub>c\<^sub>v"
+ using as by auto
+
+ have "atom z \<sharp> (\<Theta>,\<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v,v'[x::=v]\<^sub>v\<^sub>v,w, V_pair v v, V_consp tyid dc b v, (CE_val v == ce)[x::=v]\<^sub>c\<^sub>v) "
+ by (fresh_mth add: as)
+
+ have "atom bv \<sharp> (\<Theta>,\<Gamma>[x::=v]\<^sub>\<Gamma>\<^sub>v,v'[x::=v]\<^sub>v\<^sub>v,w, V_pair v v, V_consp tyid dc b v) "
+ by (fresh_mth add: as)
+
+end
+
+
+
+
end
\ No newline at end of file
diff --git a/thys/MiniSail/Syntax.thy b/thys/MiniSail/Syntax.thy
--- a/thys/MiniSail/Syntax.thy
+++ b/thys/MiniSail/Syntax.thy
@@ -1,1174 +1,1174 @@
-(*<*)
-theory Syntax
- imports "Nominal2.Nominal2" "Nominal-Utils"
-begin
- (*>*)
-
-chapter \<open>Syntax\<close>
-
-text \<open>Syntax of MiniSail programs and the contexts we use in judgements.\<close>
-
-section \<open>Program Syntax\<close>
-
-subsection \<open>AST Datatypes\<close>
-
-type_synonym num_nat = nat
-
-atom_decl x (* Immutable variables*)
-atom_decl u (* Mutable variables *)
-atom_decl bv (* Basic-type variables *)
-
-type_synonym f = string (* Function name *)
-type_synonym dc = string (* Data constructor *)
-type_synonym tyid = string
-
-text \<open>Basic types. Types without refinement constraints\<close>
-nominal_datatype "b" =
- B_int | B_bool | B_id tyid
-| B_pair b b ("[ _ , _ ]\<^sup>b")
-| B_unit | B_bitvec | B_var bv
-| B_app tyid b
-
-nominal_datatype bit = BitOne | BitZero
-
-text \<open> Literals \<close>
-nominal_datatype "l" =
- L_num "int" | L_true | L_false | L_unit | L_bitvec "bit list"
-
-text \<open> Values. We include a type identifier, tyid, in the literal for constructors
-to make typing and well-formedness checking easier \<close>
-
-nominal_datatype "v" =
- V_lit "l" ( "[ _ ]\<^sup>v")
- | V_var "x" ( "[ _ ]\<^sup>v")
- | V_pair "v" "v" ( "[ _ , _ ]\<^sup>v")
- | V_cons tyid dc "v"
- | V_consp tyid dc b "v"
-
-text \<open> Binary Operations \<close>
-nominal_datatype "opp" = Plus ( "plus") | LEq ("leq") | Eq ("eq")
-
-text \<open> Expressions \<close>
-nominal_datatype "e" =
- AE_val "v" ( "[ _ ]\<^sup>e" )
- | AE_app "f" "v" ( "[ _ ( _ ) ]\<^sup>e" )
- | AE_appP "f" "b" "v" ( "[_ [ _ ] ( _ )]\<^sup>e" )
- | AE_op "opp" "v" "v" ( "[ _ _ _ ]\<^sup>e" )
- | AE_concat v v ( "[ _ @@ _ ]\<^sup>e" )
- | AE_fst "v" ( "[#1_ ]\<^sup>e" )
- | AE_snd "v" ( "[#2_ ]\<^sup>e" )
- | AE_mvar "u" ( "[ _ ]\<^sup>e" )
- | AE_len "v" ( "[| _ |]\<^sup>e" )
- | AE_split "v" "v" ( "[ _ / _ ]\<^sup>e" )
-
-text \<open> Expressions for constraints\<close>
-nominal_datatype "ce" =
- CE_val "v" ( "[ _ ]\<^sup>c\<^sup>e" )
- | CE_op "opp" "ce" "ce" ( "[ _ _ _ ]\<^sup>c\<^sup>e" )
- | CE_concat ce ce ( "[ _ @@ _ ]\<^sup>c\<^sup>e" )
- | CE_fst "ce" ( "[#1_]\<^sup>c\<^sup>e" )
- | CE_snd "ce" ( "[#2_]\<^sup>c\<^sup>e" )
- | CE_len "ce" ( "[| _ |]\<^sup>c\<^sup>e" )
-
-text \<open> Constraints \<close>
-nominal_datatype "c" =
- C_true ( "TRUE" [] 50 )
- | C_false ( "FALSE" [] 50 )
- | C_conj "c" "c" ("_ AND _ " [50, 50] 50)
- | C_disj "c" "c" ("_ OR _ " [50,50] 50)
- | C_not "c" ( "\<not> _ " [] 50 )
- | C_imp "c" "c" ("_ IMP _ " [50, 50] 50)
- | C_eq "ce" "ce" ("_ == _ " [50, 50] 50)
-
-text \<open> Refined types \<close>
-nominal_datatype "\<tau>" =
- T_refined_type x::x b c::c binds x in c ("\<lbrace> _ : _ | _ \<rbrace>" [50, 50] 1000)
-
-text \<open> Statements \<close>
-
-nominal_datatype
- s =
- AS_val v ( "[_]\<^sup>s")
- | AS_let x::x e s::s binds x in s ( "(LET _ = _ IN _)")
- | AS_let2 x::x \<tau> s s::s binds x in s ( "(LET _ : _ = _ IN _)")
- | AS_if v s s ( "(IF _ THEN _ ELSE _)" [0, 61, 0] 61)
- | AS_var u::u \<tau> v s::s binds u in s ( "(VAR _ : _ = _ IN _)")
- | AS_assign u v ( "(_ ::= _)")
- | AS_match v branch_list ( "(MATCH _ WITH { _ })")
- | AS_while s s ( "(WHILE _ DO { _ } )" [0, 0] 61)
- | AS_seq s s ( "( _ ;; _ )" [1000, 61] 61)
- | AS_assert c s ( "(ASSERT _ IN _ )" )
- and branch_s =
- AS_branch dc x::x s::s binds x in s ( "( _ _ \<Rightarrow> _ )")
- and branch_list =
- AS_final branch_s ( "{ _ }" )
- | AS_cons branch_s branch_list ( "( _ | _ )")
-
-text \<open> Function and union type definitions \<close>
-
-nominal_datatype "fun_typ" =
- AF_fun_typ x::x "b" c::c \<tau>::\<tau> s::s binds x in c \<tau> s
-
-nominal_datatype "fun_typ_q" =
- AF_fun_typ_some bv::bv ft::fun_typ binds bv in ft
- | AF_fun_typ_none fun_typ
-
-nominal_datatype "fun_def" = AF_fundef f fun_typ_q
-
-nominal_datatype "type_def" =
- AF_typedef string "(string * \<tau>) list"
- | AF_typedef_poly string bv::bv dclist::"(string * \<tau>) list" binds bv in dclist
-
-lemma check_typedef_poly:
- "AF_typedef_poly ''option'' bv [ (''None'', \<lbrace> zz : B_unit | TRUE \<rbrace>), (''Some'', \<lbrace> zz : B_var bv | TRUE \<rbrace>) ] =
- AF_typedef_poly ''option'' bv2 [ (''None'', \<lbrace> zz : B_unit | TRUE \<rbrace>), (''Some'', \<lbrace> zz : B_var bv2 | TRUE \<rbrace>) ]"
- by auto
-
-nominal_datatype "var_def" = AV_def u \<tau> v
-
-text \<open> Programs \<close>
-nominal_datatype "p" =
- AP_prog "type_def list" "fun_def list" "var_def list" "s" ("PROG _ _ _ _")
-
-declare l.supp [simp] v.supp [simp] e.supp [simp] s_branch_s_branch_list.supp [simp] \<tau>.supp [simp] c.supp [simp] b.supp[simp]
-
-subsection \<open>Lemmas\<close>
-
-text \<open>These lemmas deal primarily with freshness and alpha-equivalence\<close>
-
-subsubsection \<open>Atoms\<close>
-
-lemma x_not_in_u_atoms[simp]:
- fixes u::u and x::x and us::"u set"
- shows "atom x \<notin> atom`us"
- by (simp add: image_iff)
-
-lemma x_fresh_u[simp]:
- fixes u::u and x::x
- shows "atom x \<sharp> u"
- by auto
-
-lemma x_not_in_b_set[simp]:
- fixes x::x and bs::"bv fset"
- shows "atom x \<notin> supp bs"
- by(induct bs,auto, simp add: supp_finsert supp_at_base)
-
-lemma x_fresh_b[simp]:
- fixes x::x and b::b
- shows "atom x \<sharp> b"
- apply (induct b rule: b.induct, auto simp: pure_supp)
- using pure_supp fresh_def by blast+
-
-lemma x_fresh_bv[simp]:
- fixes x::x and bv::bv
- shows "atom x \<sharp> bv"
- using fresh_def supp_at_base by auto
-
-lemma u_not_in_x_atoms[simp]:
- fixes u::u and x::x and xs::"x set"
- shows "atom u \<notin> atom`xs"
- by (simp add: image_iff)
-
-lemma bv_not_in_x_atoms[simp]:
- fixes bv::bv and x::x and xs::"x set"
- shows "atom bv \<notin> atom`xs"
- by (simp add: image_iff)
-
-lemma u_not_in_b_atoms[simp]:
- fixes b :: b and u::u
- shows "atom u \<notin> supp b"
- by (induct b rule: b.induct,auto simp: pure_supp supp_at_base)
-
-lemma u_not_in_b_set[simp]:
- fixes u::u and bs::"bv fset"
- shows "atom u \<notin> supp bs"
- by(induct bs, auto simp add: supp_at_base supp_finsert)
-
-lemma u_fresh_b[simp]:
- fixes x::u and b::b
- shows "atom x \<sharp> b"
- by(induct b rule: b.induct, auto simp: pure_fresh )
-
-lemma supp_b_v_disjoint:
- fixes x::x and bv::bv
- shows "supp (V_var x) \<inter> supp (B_var bv) = {}"
- by (simp add: supp_at_base)
-
-lemma supp_b_u_disjoint[simp]:
- fixes b::b and u::u
- shows "supp u \<inter> supp b = {}"
- by(nominal_induct b rule:b.strong_induct,(auto simp add: pure_supp b.supp supp_at_base)+)
-
-lemma u_fresh_bv[simp]:
- fixes u::u and b::bv
- shows "atom u \<sharp> b"
- using fresh_at_base by simp
-
-subsubsection \<open>Basic Types\<close>
-
-nominal_function b_of :: "\<tau> \<Rightarrow> b" where
- "b_of \<lbrace> z : b | c \<rbrace> = b"
- apply(auto,simp add: eqvt_def b_of_graph_aux_def )
- by (meson \<tau>.exhaust)
-nominal_termination (eqvt) by lexicographic_order
-
-lemma supp_b_empty[simp]:
- fixes b :: b and x::x
- shows "atom x \<notin> supp b"
- by (induct b rule: b.induct, auto simp: pure_supp supp_at_base x_not_in_b_set)
-
-lemma flip_b_id[simp]:
- fixes x::x and b::b
- shows "(x \<leftrightarrow> x') \<bullet> b = b"
- by(rule flip_fresh_fresh, auto simp add: fresh_def)
-
-lemma flip_x_b_cancel[simp]:
- fixes x::x and y::x and b::b and bv::bv
- shows "(x \<leftrightarrow> y) \<bullet> b = b" and "(x \<leftrightarrow> y) \<bullet> bv = bv"
- using flip_b_id apply simp
- by (metis b.eq_iff(7) b.perm_simps(7) flip_b_id)
-
-lemma flip_bv_x_cancel[simp]:
- fixes bv::bv and z::bv and x::x
- shows "(bv \<leftrightarrow> z) \<bullet> x = x" using flip_fresh_fresh[of bv x z] fresh_at_base by auto
-
-lemma flip_bv_u_cancel[simp]:
- fixes bv::bv and z::bv and x::u
- shows "(bv \<leftrightarrow> z) \<bullet> x = x" using flip_fresh_fresh[of bv x z] fresh_at_base by auto
-
-subsubsection \<open>Literals\<close>
-
-lemma supp_bitvec_empty:
- fixes bv::"bit list"
- shows "supp bv = {}"
-proof(induct bv)
- case Nil
- then show ?case using supp_Nil by auto
-next
- case (Cons a bv)
- then show ?case using supp_Cons bit.supp
- by (metis (mono_tags, opaque_lifting) bit.strong_exhaust l.supp(5) sup_bot.right_neutral)
-qed
-
-lemma bitvec_pure[simp]:
- fixes bv::"bit list" and x::x
- shows "atom x \<sharp> bv" using fresh_def supp_bitvec_empty by auto
-
-lemma supp_l_empty[simp]:
- fixes l:: l
- shows "supp (V_lit l) = {}"
- by(nominal_induct l rule: l.strong_induct,
- auto simp add: l.strong_exhaust pure_supp v.fv_defs supp_bitvec_empty)
-
-lemma type_l_nosupp[simp]:
- fixes x::x and l::l
- shows "atom x \<notin> supp (\<lbrace> z : b | [[z]\<^sup>v]\<^sup>c\<^sup>e == [[l]\<^sup>v]\<^sup>c\<^sup>e \<rbrace>)"
- using supp_at_base supp_l_empty ce.supp(1) c.supp \<tau>.supp by force
-
-lemma flip_bitvec0:
- fixes x::"bit list"
- assumes "atom c \<sharp> (z, x, z')"
- shows "(z \<leftrightarrow> c) \<bullet> x = (z' \<leftrightarrow> c) \<bullet> x"
-proof -
- have "atom z \<sharp> x" and "atom z' \<sharp> x"
- using flip_fresh_fresh assms supp_bitvec_empty fresh_def by blast+
- moreover have "atom c \<sharp> x" using supp_bitvec_empty fresh_def by auto
- ultimately show ?thesis using assms flip_fresh_fresh by metis
-qed
-
-lemma flip_bitvec:
- assumes "atom c \<sharp> (z, L_bitvec x, z')"
- shows "(z \<leftrightarrow> c) \<bullet> x = (z' \<leftrightarrow> c) \<bullet> x"
-proof -
- have "atom z \<sharp> x" and "atom z' \<sharp> x"
- using flip_fresh_fresh assms supp_bitvec_empty fresh_def by blast+
- moreover have "atom c \<sharp> x" using supp_bitvec_empty fresh_def by auto
- ultimately show ?thesis using assms flip_fresh_fresh by metis
-qed
-
-lemma type_l_eq:
- shows "\<lbrace> z : b | [[z]\<^sup>v]\<^sup>c\<^sup>e == [V_lit l]\<^sup>c\<^sup>e \<rbrace> = (\<lbrace> z' : b | [[z']\<^sup>v]\<^sup>c\<^sup>e == [V_lit l]\<^sup>c\<^sup>e \<rbrace>)"
- by(auto,nominal_induct l rule: l.strong_induct,auto, metis permute_pure, auto simp add: flip_bitvec)
-
-lemma flip_l_eq:
- fixes x::l
- shows "(z \<leftrightarrow> c) \<bullet> x = (z' \<leftrightarrow> c) \<bullet> x"
-proof -
- have "atom z \<sharp> x" and "atom c \<sharp> x" and "atom z' \<sharp> x"
- using flip_fresh_fresh fresh_def supp_l_empty by fastforce+
- thus ?thesis using flip_fresh_fresh by metis
-qed
-
-lemma flip_l_eq1:
- fixes x::l
- assumes "(z \<leftrightarrow> c) \<bullet> x = (z' \<leftrightarrow> c) \<bullet> x'"
- shows "x' = x"
-proof -
- have "atom z \<sharp> x" and "atom c \<sharp> x'" and "atom c \<sharp> x" and "atom z' \<sharp> x'"
- using flip_fresh_fresh fresh_def supp_l_empty by fastforce+
- thus ?thesis using flip_fresh_fresh assms by metis
-qed
-
-subsubsection \<open>Types\<close>
-
-lemma flip_base_eq:
- fixes b::b and x::x and y::x
- shows "(x \<leftrightarrow> y) \<bullet> b = b"
- using b.fresh by (simp add: flip_fresh_fresh fresh_def)
-
-text \<open> Obtain an alpha-equivalent type where the bound variable is fresh in some term t \<close>
-lemma has_fresh_z0:
- fixes t::"'b::fs"
- shows "\<exists>z. atom z \<sharp> (c',t) \<and> (\<lbrace>z' : b | c' \<rbrace>) = (\<lbrace> z : b | (z \<leftrightarrow> z' ) \<bullet> c' \<rbrace>)"
-proof -
- obtain z::x where fr: "atom z \<sharp> (c',t)" using obtain_fresh by blast
- moreover hence "(\<lbrace> z' : b | c' \<rbrace>) = (\<lbrace> z : b | (z \<leftrightarrow> z') \<bullet> c' \<rbrace>)"
- using \<tau>.eq_iff Abs1_eq_iff
- by (metis flip_commute flip_fresh_fresh fresh_PairD(1))
- ultimately show ?thesis by fastforce
-qed
-
-lemma has_fresh_z:
- fixes t::"'b::fs"
- shows "\<exists>z b c. atom z \<sharp> t \<and> \<tau> = \<lbrace> z : b | c \<rbrace>"
-proof -
- obtain z' and b and c' where teq: "\<tau> = (\<lbrace> z' : b | c' \<rbrace>)" using \<tau>.exhaust by blast
- obtain z::x where fr: "atom z \<sharp> (t,c')" using obtain_fresh by blast
- hence "(\<lbrace> z' : b | c' \<rbrace>) = (\<lbrace> z : b | (z \<leftrightarrow> z') \<bullet> c' \<rbrace>)" using \<tau>.eq_iff Abs1_eq_iff
- flip_commute flip_fresh_fresh fresh_PairD(1) by (metis fresh_PairD(2))
- hence "atom z \<sharp> t \<and> \<tau> = (\<lbrace> z : b | (z \<leftrightarrow> z') \<bullet> c' \<rbrace>)" using fr teq by force
- thus ?thesis using teq fr by fast
-qed
-
-lemma obtain_fresh_z:
- fixes t::"'b::fs"
- obtains z and b and c where "atom z \<sharp> t \<and> \<tau> = \<lbrace> z : b | c \<rbrace>"
- using has_fresh_z by blast
-
-lemma has_fresh_z2:
- fixes t::"'b::fs"
- shows "\<exists>z c. atom z \<sharp> t \<and> \<tau> = \<lbrace> z : b_of \<tau> | c \<rbrace>"
-proof -
- obtain z and b and c where "atom z \<sharp> t \<and> \<tau> = \<lbrace> z : b | c \<rbrace>" using obtain_fresh_z by metis
- moreover then have "b_of \<tau> = b" using \<tau>.eq_iff by simp
- ultimately show ?thesis using obtain_fresh_z \<tau>.eq_iff by auto
-qed
-
-lemma obtain_fresh_z2:
- fixes t::"'b::fs"
- obtains z and c where "atom z \<sharp> t \<and> \<tau> = \<lbrace> z : b_of \<tau> | c \<rbrace>"
- using has_fresh_z2 by blast
-
-subsubsection \<open>Values\<close>
-
-lemma u_notin_supp_v[simp]:
- fixes u::u and v::v
- shows "atom u \<notin> supp v"
-proof(nominal_induct v rule: v.strong_induct)
- case (V_lit l)
- then show ?case using supp_l_empty by auto
-next
- case (V_var x)
- then show ?case
- by (simp add: supp_at_base)
-next
- case (V_pair v1 v2)
- then show ?case by auto
-next
- case (V_cons tyid list v)
- then show ?case using pure_supp by auto
-next
- case (V_consp tyid list b v)
- then show ?case using pure_supp by auto
-qed
-
-lemma u_fresh_xv[simp]:
- fixes u::u and x::x and v::v
- shows "atom u \<sharp> (x,v)"
-proof -
- have "atom u \<sharp> x" using fresh_def by fastforce
- moreover have "atom u \<sharp> v" using fresh_def u_notin_supp_v by metis
- ultimately show ?thesis using fresh_prod2 by auto
-qed
-
-text \<open> Part of an effort to make the proofs across inductive cases more uniform by distilling
-the non-uniform parts into lemmas like this\<close>
-lemma v_flip_eq:
- fixes v::v and va::v and x::x and c::x
- assumes "atom c \<sharp> (v, va)" and "atom c \<sharp> (x, xa, v, va)" and "(x \<leftrightarrow> c) \<bullet> v = (xa \<leftrightarrow> c) \<bullet> va"
- shows "((v = V_lit l \<longrightarrow> (\<exists>l'. va = V_lit l' \<and> (x \<leftrightarrow> c) \<bullet> l = (xa \<leftrightarrow> c) \<bullet> l'))) \<and>
- ((v = V_var y \<longrightarrow> (\<exists>y'. va = V_var y' \<and> (x \<leftrightarrow> c) \<bullet> y = (xa \<leftrightarrow> c) \<bullet> y'))) \<and>
- ((v = V_pair vone vtwo \<longrightarrow> (\<exists>v1' v2'. va = V_pair v1' v2' \<and> (x \<leftrightarrow> c) \<bullet> vone = (xa \<leftrightarrow> c) \<bullet> v1' \<and> (x \<leftrightarrow> c) \<bullet> vtwo = (xa \<leftrightarrow> c) \<bullet> v2'))) \<and>
- ((v = V_cons tyid dc vone \<longrightarrow> (\<exists>v1'. va = V_cons tyid dc v1'\<and> (x \<leftrightarrow> c) \<bullet> vone = (xa \<leftrightarrow> c) \<bullet> v1'))) \<and>
- ((v = V_consp tyid dc b vone \<longrightarrow> (\<exists>v1'. va = V_consp tyid dc b v1'\<and> (x \<leftrightarrow> c) \<bullet> vone = (xa \<leftrightarrow> c) \<bullet> v1')))"
- using assms proof(nominal_induct v rule:v.strong_induct)
- case (V_lit l)
- then show ?case using assms v.perm_simps
- empty_iff flip_def fresh_def fresh_permute_iff supp_l_empty swap_fresh_fresh v.fresh
- by (metis permute_swap_cancel2 v.distinct)
-next
- case (V_var x)
- then show ?case using assms v.perm_simps
- empty_iff flip_def fresh_def fresh_permute_iff supp_l_empty swap_fresh_fresh v.fresh
- by (metis permute_swap_cancel2 v.distinct)
-next
- case (V_pair v1 v2)
- have " (V_pair v1 v2 = V_pair vone vtwo \<longrightarrow> (\<exists>v1' v2'. va = V_pair v1' v2' \<and> (x \<leftrightarrow> c) \<bullet> vone = (xa \<leftrightarrow> c) \<bullet> v1' \<and> (x \<leftrightarrow> c) \<bullet> vtwo = (xa \<leftrightarrow> c) \<bullet> v2'))" proof
- assume "V_pair v1 v2= V_pair vone vtwo"
- thus "(\<exists>v1' v2'. va = V_pair v1' v2' \<and> (x \<leftrightarrow> c) \<bullet> vone = (xa \<leftrightarrow> c) \<bullet> v1' \<and> (x \<leftrightarrow> c) \<bullet> vtwo = (xa \<leftrightarrow> c) \<bullet> v2')"
- using V_pair assms
- by (metis (no_types, opaque_lifting) flip_def permute_swap_cancel v.perm_simps(3))
- qed
- thus ?case using V_pair by auto
-next
- case (V_cons tyid dc v1)
- have " (V_cons tyid dc v1 = V_cons tyid dc vone \<longrightarrow> (\<exists> v1'. va = V_cons tyid dc v1' \<and> (x \<leftrightarrow> c) \<bullet> vone = (xa \<leftrightarrow> c) \<bullet> v1'))" proof
- assume as: "V_cons tyid dc v1 = V_cons tyid dc vone"
- hence "(x \<leftrightarrow> c) \<bullet> (V_cons tyid dc vone) = V_cons tyid dc ((x \<leftrightarrow> c) \<bullet> vone)" proof -
- have "(x \<leftrightarrow> c) \<bullet> dc = dc" using pure_permute_id by metis
- moreover have "(x \<leftrightarrow> c) \<bullet> tyid = tyid" using pure_permute_id by metis
- ultimately show ?thesis using v.perm_simps(4) by simp
- qed
- then obtain v1' where "(xa \<leftrightarrow> c) \<bullet> va = V_cons tyid dc v1' \<and> (x \<leftrightarrow> c) \<bullet> vone = v1'" using assms V_cons
- using as by fastforce
- hence " va = V_cons tyid dc ((xa \<leftrightarrow> c) \<bullet> v1') \<and> (x \<leftrightarrow> c) \<bullet> vone = v1'" using permute_flip_cancel empty_iff flip_def fresh_def supp_b_empty swap_fresh_fresh
- by (metis pure_fresh v.perm_simps(4))
-
- thus "(\<exists>v1'. va = V_cons tyid dc v1' \<and> (x \<leftrightarrow> c) \<bullet> vone = (xa \<leftrightarrow> c) \<bullet> v1')"
- using V_cons assms by simp
- qed
- thus ?case using V_cons by auto
-next
- case (V_consp tyid dc b v1)
- have " (V_consp tyid dc b v1 = V_consp tyid dc b vone \<longrightarrow> (\<exists> v1'. va = V_consp tyid dc b v1' \<and> (x \<leftrightarrow> c) \<bullet> vone = (xa \<leftrightarrow> c) \<bullet> v1'))" proof
- assume as: "V_consp tyid dc b v1 = V_consp tyid dc b vone"
- hence "(x \<leftrightarrow> c) \<bullet> (V_consp tyid dc b vone) = V_consp tyid dc b ((x \<leftrightarrow> c) \<bullet> vone)" proof -
- have "(x \<leftrightarrow> c) \<bullet> dc = dc" using pure_permute_id by metis
- moreover have "(x \<leftrightarrow> c) \<bullet> tyid = tyid" using pure_permute_id by metis
- ultimately show ?thesis using v.perm_simps(4) by simp
- qed
- then obtain v1' where "(xa \<leftrightarrow> c) \<bullet> va = V_consp tyid dc b v1' \<and> (x \<leftrightarrow> c) \<bullet> vone = v1'" using assms V_consp
- using as by fastforce
- hence " va = V_consp tyid dc b ((xa \<leftrightarrow> c) \<bullet> v1') \<and> (x \<leftrightarrow> c) \<bullet> vone = v1'" using permute_flip_cancel empty_iff flip_def fresh_def supp_b_empty swap_fresh_fresh
- pure_fresh v.perm_simps
- by (metis (mono_tags, opaque_lifting))
- thus "(\<exists>v1'. va = V_consp tyid dc b v1' \<and> (x \<leftrightarrow> c) \<bullet> vone = (xa \<leftrightarrow> c) \<bullet> v1')"
- using V_consp assms by simp
- qed
- thus ?case using V_consp by auto
-qed
-
-lemma flip_eq:
- fixes x::x and xa::x and s::"'a::fs" and sa::"'a::fs"
- assumes "(\<forall>c. atom c \<sharp> (s, sa) \<longrightarrow> atom c \<sharp> (x, xa, s, sa) \<longrightarrow> (x \<leftrightarrow> c) \<bullet> s = (xa \<leftrightarrow> c) \<bullet> sa)" and "x \<noteq> xa"
- shows "(x \<leftrightarrow> xa) \<bullet> s = sa"
-proof -
- have " ([[atom x]]lst. s = [[atom xa]]lst. sa)" using assms Abs1_eq_iff_all by simp
- hence "(xa = x \<and> sa = s \<or> xa \<noteq> x \<and> sa = (xa \<leftrightarrow> x) \<bullet> s \<and> atom xa \<sharp> s)" using assms Abs1_eq_iff[of xa sa x s] by simp
- thus ?thesis using assms
- by (metis flip_commute)
-qed
-
-lemma swap_v_supp:
- fixes v::v and d::x and z::x
- assumes "atom d \<sharp> v"
- shows "supp ((z \<leftrightarrow> d ) \<bullet> v) \<subseteq> supp v - { atom z } \<union> { atom d}"
- using assms
-proof(nominal_induct v rule:v.strong_induct)
- case (V_lit l)
- then show ?case using l.supp by (metis supp_l_empty empty_subsetI l.strong_exhaust pure_supp supp_eqvt v.supp)
-next
- case (V_var x)
- hence "d \<noteq> x" using fresh_def by fastforce
- thus ?case apply(cases "z = x") using supp_at_base V_var \<open>d\<noteq>x\<close> by fastforce+
-next
- case (V_cons tyid dc v)
- show ?case using v.supp(4) pure_supp
- using V_cons.hyps V_cons.prems fresh_def by auto
-next
- case (V_consp tyid dc b v)
- show ?case using v.supp(4) pure_supp
- using V_consp.hyps V_consp.prems fresh_def by auto
-qed(force+)
-
-subsubsection \<open>Expressions\<close>
-
-lemma swap_e_supp:
- fixes e::e and d::x and z::x
- assumes "atom d \<sharp> e"
- shows "supp ((z \<leftrightarrow> d ) \<bullet> e) \<subseteq> supp e - { atom z } \<union> { atom d}"
- using assms
-proof(nominal_induct e rule:e.strong_induct)
- case (AE_val v)
- then show ?case using swap_v_supp by simp
-next
- case (AE_app f v)
- then show ?case using swap_v_supp by (simp add: pure_supp)
-next
- case (AE_appP b f v)
- hence df: "atom d \<sharp> v" using fresh_def e.supp by force
- have "supp ((z \<leftrightarrow> d ) \<bullet> (AE_appP b f v)) = supp (AE_appP b f ((z \<leftrightarrow> d ) \<bullet> v))" using e.supp
- by (metis b.eq_iff(3) b.perm_simps(3) e.perm_simps(3) flip_b_id)
- also have "... = supp b \<union> supp f \<union> supp ((z \<leftrightarrow> d ) \<bullet> v)" using e.supp by auto
- also have "... \<subseteq> supp b \<union> supp f \<union> supp v - { atom z } \<union> { atom d}" using swap_v_supp[OF df] pure_supp by auto
- finally show ?case using e.supp by auto
-next
- case (AE_op opp v1 v2)
- hence df: "atom d \<sharp> v1 \<and> atom d \<sharp> v2" using fresh_def e.supp by force
- have "((z \<leftrightarrow> d ) \<bullet> (AE_op opp v1 v2)) = AE_op opp ((z \<leftrightarrow> d ) \<bullet> v1) ((z \<leftrightarrow> d ) \<bullet> v2)" using
- e.perm_simps flip_commute opp.perm_simps AE_op opp.strong_exhaust pure_supp
- by (metis (full_types))
-
- hence "supp ((z \<leftrightarrow> d) \<bullet> AE_op opp v1 v2) = supp (AE_op opp ((z \<leftrightarrow> d) \<bullet>v1) ((z \<leftrightarrow> d) \<bullet>v2))" by simp
- also have "... = supp ((z \<leftrightarrow> d) \<bullet>v1) \<union> supp ((z \<leftrightarrow> d) \<bullet>v2)" using e.supp
- by (metis (mono_tags, opaque_lifting) opp.strong_exhaust opp.supp sup_bot.left_neutral)
- also have "... \<subseteq> (supp v1 - { atom z } \<union> { atom d}) \<union> (supp v2 - { atom z } \<union> { atom d})" using swap_v_supp AE_op df by blast
- finally show ?case using e.supp opp.supp by blast
-next
- case (AE_fst v)
- then show ?case using swap_v_supp by auto
-next
- case (AE_snd v)
- then show ?case using swap_v_supp by auto
-next
- case (AE_mvar u)
- then show ?case using
- Diff_empty Diff_insert0 Un_upper1 atom_x_sort flip_def flip_fresh_fresh fresh_def set_eq_subset supp_eqvt swap_set_in_eq
- by (metis sort_of_atom_eq)
-next
- case (AE_len v)
- then show ?case using swap_v_supp by auto
-next
- case (AE_concat v1 v2)
- then show ?case using swap_v_supp by auto
-next
- case (AE_split v1 v2)
- then show ?case using swap_v_supp by auto
-qed
-
-lemma swap_ce_supp:
- fixes e::ce and d::x and z::x
- assumes "atom d \<sharp> e"
- shows "supp ((z \<leftrightarrow> d ) \<bullet> e) \<subseteq> supp e - { atom z } \<union> { atom d}"
- using assms
-proof(nominal_induct e rule:ce.strong_induct)
- case (CE_val v)
- then show ?case using swap_v_supp ce.fresh ce.supp by simp
-next
- case (CE_op opp v1 v2)
- hence df: "atom d \<sharp> v1 \<and> atom d \<sharp> v2" using fresh_def e.supp by force
- have "((z \<leftrightarrow> d ) \<bullet> (CE_op opp v1 v2)) = CE_op opp ((z \<leftrightarrow> d ) \<bullet> v1) ((z \<leftrightarrow> d ) \<bullet> v2)" using
- ce.perm_simps flip_commute opp.perm_simps CE_op opp.strong_exhaust x_fresh_b pure_supp
- by (metis (full_types))
-
- hence "supp ((z \<leftrightarrow> d) \<bullet> CE_op opp v1 v2) = supp (CE_op opp ((z \<leftrightarrow> d) \<bullet>v1) ((z \<leftrightarrow> d) \<bullet>v2))" by simp
- also have "... = supp ((z \<leftrightarrow> d) \<bullet>v1) \<union> supp ((z \<leftrightarrow> d) \<bullet>v2)" using ce.supp
- by (metis (mono_tags, opaque_lifting) opp.strong_exhaust opp.supp sup_bot.left_neutral)
- also have "... \<subseteq> (supp v1 - { atom z } \<union> { atom d}) \<union> (supp v2 - { atom z } \<union> { atom d})" using swap_v_supp CE_op df by blast
- finally show ?case using ce.supp opp.supp by blast
-next
- case (CE_fst v)
- then show ?case using ce.supp ce.fresh swap_v_supp by auto
-next
- case (CE_snd v)
- then show ?case using ce.supp ce.fresh swap_v_supp by auto
-next
- case (CE_len v)
- then show ?case using ce.supp ce.fresh swap_v_supp by auto
-next
- case (CE_concat v1 v2)
- then show ?case using ce.supp ce.fresh swap_v_supp ce.perm_simps
- proof -
- have "\<forall>x v xa. \<not> atom (x::x) \<sharp> (v::v) \<or> supp ((xa \<leftrightarrow> x) \<bullet> v) \<subseteq> supp v - {atom xa} \<union> {atom x}"
- by (meson swap_v_supp) (* 0.0 ms *)
- then show ?thesis
- using CE_concat ce.supp by auto
- qed
-qed
-
-lemma swap_c_supp:
- fixes c::c and d::x and z::x
- assumes "atom d \<sharp> c"
- shows "supp ((z \<leftrightarrow> d ) \<bullet> c) \<subseteq> supp c - { atom z } \<union> { atom d}"
- using assms
-proof(nominal_induct c rule:c.strong_induct)
- case (C_eq e1 e2)
- then show ?case using swap_ce_supp by auto
-qed(auto+)
-
-lemma type_e_eq:
- assumes "atom z \<sharp> e" and "atom z' \<sharp> e"
- shows "\<lbrace> z : b | [[z]\<^sup>v]\<^sup>c\<^sup>e == e \<rbrace> = (\<lbrace> z' : b | [[z']\<^sup>v]\<^sup>c\<^sup>e == e \<rbrace>)"
- by (auto,metis (full_types) assms(1) assms(2) flip_fresh_fresh fresh_PairD(1) fresh_PairD(2))
-
-lemma type_e_eq2:
- assumes "atom z \<sharp> e" and "atom z' \<sharp> e" and "b=b'"
- shows "\<lbrace> z : b | [[z]\<^sup>v]\<^sup>c\<^sup>e == e \<rbrace> = (\<lbrace> z' : b' | [[z']\<^sup>v]\<^sup>c\<^sup>e == e \<rbrace>)"
- using assms type_e_eq by fast
-
-lemma e_flip_eq:
- fixes e::e and ea::e
- assumes "atom c \<sharp> (e, ea)" and "atom c \<sharp> (x, xa, e, ea)" and "(x \<leftrightarrow> c) \<bullet> e = (xa \<leftrightarrow> c) \<bullet> ea"
- shows "(e = AE_val w \<longrightarrow> (\<exists>w'. ea = AE_val w' \<and> (x \<leftrightarrow> c) \<bullet> w = (xa \<leftrightarrow> c) \<bullet> w')) \<or>
- (e = AE_op opp v1 v2 \<longrightarrow> (\<exists>v1' v2'. ea = AS_op opp v1' v2' \<and> (x \<leftrightarrow> c) \<bullet> v1 = (xa \<leftrightarrow> c) \<bullet> v1') \<and> (x \<leftrightarrow> c) \<bullet> v2 = (xa \<leftrightarrow> c) \<bullet> v2') \<or>
- (e = AE_fst v \<longrightarrow> (\<exists>v'. ea = AE_fst v' \<and> (x \<leftrightarrow> c) \<bullet> v = (xa \<leftrightarrow> c) \<bullet> v')) \<or>
- (e = AE_snd v \<longrightarrow> (\<exists>v'. ea = AE_snd v' \<and> (x \<leftrightarrow> c) \<bullet> v = (xa \<leftrightarrow> c) \<bullet> v')) \<or>
- (e = AE_len v \<longrightarrow> (\<exists>v'. ea = AE_len v' \<and> (x \<leftrightarrow> c) \<bullet> v = (xa \<leftrightarrow> c) \<bullet> v')) \<or>
- (e = AE_concat v1 v2 \<longrightarrow> (\<exists>v1' v2'. ea = AS_concat v1' v2' \<and> (x \<leftrightarrow> c) \<bullet> v1 = (xa \<leftrightarrow> c) \<bullet> v1') \<and> (x \<leftrightarrow> c) \<bullet> v2 = (xa \<leftrightarrow> c) \<bullet> v2') \<or>
- (e = AE_app f v \<longrightarrow> (\<exists>v'. ea = AE_app f v' \<and> (x \<leftrightarrow> c) \<bullet> v = (xa \<leftrightarrow> c) \<bullet> v'))"
- by (metis assms e.perm_simps permute_flip_cancel2)
-
-lemma fresh_opp_all:
- fixes opp::opp
- shows "z \<sharp> opp"
- using e.fresh opp.exhaust opp.fresh by metis
-
-lemma fresh_e_opp_all:
- shows "(z \<sharp> v1 \<and> z \<sharp> v2) = z \<sharp> AE_op opp v1 v2"
- using e.fresh opp.exhaust opp.fresh fresh_opp_all by simp
-
-lemma fresh_e_opp:
- fixes z::x
- assumes "atom z \<sharp> v1 \<and> atom z \<sharp> v2"
- shows "atom z \<sharp> AE_op opp v1 v2"
- using e.fresh opp.exhaust opp.fresh opp.supp by (metis assms)
-
-subsubsection \<open>Statements\<close>
-
-lemma branch_s_flip_eq:
- fixes v::v and va::v
- assumes "atom c \<sharp> (v, va)" and "atom c \<sharp> (x, xa, v, va)" and "(x \<leftrightarrow> c) \<bullet> s = (xa \<leftrightarrow> c) \<bullet> sa"
- shows "(s = AS_val w \<longrightarrow> (\<exists>w'. sa = AS_val w' \<and> (x \<leftrightarrow> c) \<bullet> w = (xa \<leftrightarrow> c) \<bullet> w')) \<or>
- (s = AS_seq s1 s2 \<longrightarrow> (\<exists>s1' s2'. sa = AS_seq s1' s2' \<and> (x \<leftrightarrow> c) \<bullet> s1 = (xa \<leftrightarrow> c) \<bullet> s1') \<and> (x \<leftrightarrow> c) \<bullet> s2 = (xa \<leftrightarrow> c) \<bullet> s2') \<or>
- (s = AS_if v s1 s2 \<longrightarrow> (\<exists>v' s1' s2'. sa = AS_if seq s1' s2' \<and> (x \<leftrightarrow> c) \<bullet> s1 = (xa \<leftrightarrow> c) \<bullet> s1') \<and> (x \<leftrightarrow> c) \<bullet> s2 = (xa \<leftrightarrow> c) \<bullet> s2' \<and> (x \<leftrightarrow> c) \<bullet> c = (xa \<leftrightarrow> c) \<bullet> v')"
- by (metis assms s_branch_s_branch_list.perm_simps permute_flip_cancel2)
-
-section \<open>Context Syntax\<close>
-
-subsection \<open>Datatypes\<close>
-
-text \<open>Type and function/type definition contexts\<close>
-type_synonym \<Phi> = "fun_def list"
-type_synonym \<Theta> = "type_def list"
-type_synonym \<B> = "bv fset"
-
-datatype \<Gamma> =
- GNil
- | GCons "x*b*c" \<Gamma> (infixr "#\<^sub>\<Gamma>" 65)
-
-datatype \<Delta> =
- DNil ("[]\<^sub>\<Delta>")
- | DCons "u*\<tau>" \<Delta> (infixr "#\<^sub>\<Delta>" 65)
-
-subsection \<open>Functions and Lemmas\<close>
-
-lemma \<Gamma>_induct [case_names GNil GCons] : "P GNil \<Longrightarrow> (\<And>x b c \<Gamma>'. P \<Gamma>' \<Longrightarrow> P ((x,b,c) #\<^sub>\<Gamma> \<Gamma>')) \<Longrightarrow> P \<Gamma>"
-proof(induct \<Gamma> rule:\<Gamma>.induct)
- case GNil
- then show ?case by auto
-next
- case (GCons x1 x2)
- then obtain x and b and c where "x1=(x,b,c)" using prod_cases3 by blast
- then show ?case using GCons by presburger
-qed
-
-instantiation \<Delta> :: pt
-begin
-
-primrec permute_\<Delta>
- where
- DNil_eqvt: "p \<bullet> DNil = DNil"
- | DCons_eqvt: "p \<bullet> (x #\<^sub>\<Delta> xs) = p \<bullet> x #\<^sub>\<Delta> p \<bullet> (xs::\<Delta>)"
-
-instance by standard (induct_tac [!] x, simp_all)
-end
-
-lemmas [eqvt] = permute_\<Delta>.simps
-
-lemma \<Delta>_induct [case_names DNil DCons] : "P DNil \<Longrightarrow> (\<And>u t \<Delta>'. P \<Delta>' \<Longrightarrow> P ((u,t) #\<^sub>\<Delta> \<Delta>')) \<Longrightarrow> P \<Delta>"
-proof(induct \<Delta> rule: \<Delta>.induct)
- case DNil
- then show ?case by auto
-next
- case (DCons x1 x2)
- then obtain u and t where "x1=(u,t)" by fastforce
- then show ?case using DCons by presburger
-qed
-
-lemma \<Phi>_induct [case_names PNil PConsNone PConsSome] : "P [] \<Longrightarrow> (\<And>f x b c \<tau> s' \<Phi>'. P \<Phi>' \<Longrightarrow> P ((AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \<tau> s'))) # \<Phi>')) \<Longrightarrow>
- (\<And>f bv x b c \<tau> s' \<Phi>'. P \<Phi>' \<Longrightarrow> P ((AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \<tau> s'))) # \<Phi>')) \<Longrightarrow> P \<Phi>"
-proof(induct \<Phi> rule: list.induct)
- case Nil
- then show ?case by auto
-next
- case (Cons x1 x2)
- then obtain f and t where ft: "x1 = (AF_fundef f t)"
- by (meson fun_def.exhaust)
- then show ?case proof(nominal_induct t rule:fun_typ_q.strong_induct)
- case (AF_fun_typ_some bv ft)
- then show ?case using Cons ft
- by (metis fun_typ.exhaust)
- next
- case (AF_fun_typ_none ft)
- then show ?case using Cons ft
- by (metis fun_typ.exhaust)
- qed
-qed
-
-lemma \<Theta>_induct [case_names TNil AF_typedef AF_typedef_poly] : "P [] \<Longrightarrow> (\<And>tid dclist \<Theta>'. P \<Theta>' \<Longrightarrow> P ((AF_typedef tid dclist) # \<Theta>')) \<Longrightarrow>
- (\<And>tid bv dclist \<Theta>'. P \<Theta>' \<Longrightarrow> P ((AF_typedef_poly tid bv dclist ) # \<Theta>')) \<Longrightarrow> P \<Theta>"
-proof(induct \<Theta> rule: list.induct)
- case Nil
- then show ?case by auto
-next
- case (Cons td T)
- show ?case by(cases td rule: type_def.exhaust, (simp add: Cons)+)
-qed
-
-instantiation \<Gamma> :: pt
-begin
-
-primrec permute_\<Gamma>
- where
- GNil_eqvt: "p \<bullet> GNil = GNil"
- | GCons_eqvt: "p \<bullet> (x #\<^sub>\<Gamma> xs) = p \<bullet> x #\<^sub>\<Gamma> p \<bullet> (xs::\<Gamma>)"
-
-instance by standard (induct_tac [!] x, simp_all)
-end
-
-lemmas [eqvt] = permute_\<Gamma>.simps
-
-lemma G_cons_eqvt[simp]:
- fixes \<Gamma>::\<Gamma>
- shows "p \<bullet> ((x,b,c) #\<^sub>\<Gamma> \<Gamma>) = ((p \<bullet> x, p \<bullet> b , p \<bullet> c) #\<^sub>\<Gamma> (p \<bullet> \<Gamma> ))" (is "?A = ?B" )
- using Cons_eqvt triple_eqvt supp_b_empty by simp
-
-lemma G_cons_flip[simp]:
- fixes x::x and \<Gamma>::\<Gamma>
- shows "(x\<leftrightarrow>x') \<bullet> ((x'',b,c) #\<^sub>\<Gamma> \<Gamma>) = (((x\<leftrightarrow>x') \<bullet> x'', b , (x\<leftrightarrow>x') \<bullet> c) #\<^sub>\<Gamma> ((x\<leftrightarrow>x') \<bullet> \<Gamma> ))"
- using Cons_eqvt triple_eqvt supp_b_empty by auto
-
-lemma G_cons_flip_fresh[simp]:
- fixes x::x and \<Gamma>::\<Gamma>
- assumes "atom x \<sharp> (c,\<Gamma>)" and "atom x' \<sharp> (c,\<Gamma>)"
- shows "(x\<leftrightarrow>x') \<bullet> ((x',b,c) #\<^sub>\<Gamma> \<Gamma>) = ((x, b , c) #\<^sub>\<Gamma> \<Gamma> )"
- using G_cons_flip flip_fresh_fresh assms by force
-
-lemma G_cons_flip_fresh2[simp]:
- fixes x::x and \<Gamma>::\<Gamma>
- assumes "atom x \<sharp> (c,\<Gamma>)" and "atom x' \<sharp> (c,\<Gamma>)"
- shows "(x\<leftrightarrow>x') \<bullet> ((x,b,c) #\<^sub>\<Gamma> \<Gamma>) = ((x', b , c) #\<^sub>\<Gamma> \<Gamma> )"
- using G_cons_flip flip_fresh_fresh assms by force
-
-lemma G_cons_flip_fresh3[simp]:
- fixes x::x and \<Gamma>::\<Gamma>
- assumes "atom x \<sharp> \<Gamma>" and "atom x' \<sharp> \<Gamma>"
- shows "(x\<leftrightarrow>x') \<bullet> ((x',b,c) #\<^sub>\<Gamma> \<Gamma>) = ((x, b , (x\<leftrightarrow>x') \<bullet> c) #\<^sub>\<Gamma> \<Gamma> )"
- using G_cons_flip flip_fresh_fresh assms by force
-
-lemma neq_GNil_conv: "(xs \<noteq> GNil) = (\<exists>y ys. xs = y #\<^sub>\<Gamma> ys)"
- by (induct xs) auto
-
-nominal_function toList :: "\<Gamma> \<Rightarrow> (x*b*c) list" where
- "toList GNil = []"
-| "toList (GCons xbc G) = xbc#(toList G)"
- apply (auto, simp add: eqvt_def toList_graph_aux_def )
- using neq_GNil_conv surj_pair by metis
-nominal_termination (eqvt)
- by lexicographic_order
-
-nominal_function toSet :: "\<Gamma> \<Rightarrow> (x*b*c) set" where
- "toSet GNil = {}"
-| "toSet (GCons xbc G) = {xbc} \<union> (toSet G)"
- apply (auto,simp add: eqvt_def toSet_graph_aux_def )
- using neq_GNil_conv surj_pair by metis
-nominal_termination (eqvt)
- by lexicographic_order
-
-nominal_function append_g :: "\<Gamma> \<Rightarrow> \<Gamma> \<Rightarrow> \<Gamma>" (infixr "@" 65) where
- "append_g GNil g = g"
-| "append_g (xbc #\<^sub>\<Gamma> g1) g2 = (xbc #\<^sub>\<Gamma> (g1@g2))"
- apply (auto,simp add: eqvt_def append_g_graph_aux_def )
- using neq_GNil_conv surj_pair by metis
-nominal_termination (eqvt) by lexicographic_order
-
-nominal_function dom :: "\<Gamma> \<Rightarrow> x set" where
- "dom \<Gamma> = (fst` (toSet \<Gamma>))"
- apply auto
- unfolding eqvt_def dom_graph_aux_def lfp_eqvt toSet.eqvt by simp
-nominal_termination (eqvt) by lexicographic_order
-
-text \<open> Use of this is sometimes mixed in with use of freshness and support for the context however it makes it clear
-that for immutable variables, the context is `self-supporting'\<close>
-
-nominal_function atom_dom :: "\<Gamma> \<Rightarrow> atom set" where
- "atom_dom \<Gamma> = atom`(dom \<Gamma>)"
- apply auto
- unfolding eqvt_def atom_dom_graph_aux_def lfp_eqvt toSet.eqvt by simp
-nominal_termination (eqvt) by lexicographic_order
-
-subsection \<open>Immutable Variable Context Lemmas\<close>
-
-lemma append_GNil[simp]:
- "GNil @ G = G"
- by simp
-
-lemma append_g_toSetU [simp]: "toSet (G1@G2) = toSet G1 \<union> toSet G2"
- by(induct G1, auto+)
-
-lemma supp_GNil:
- shows "supp GNil = {}"
- by (simp add: supp_def)
-
-lemma supp_GCons:
- fixes xs::\<Gamma>
- shows "supp (x #\<^sub>\<Gamma> xs) = supp x \<union> supp xs"
- by (simp add: supp_def Collect_imp_eq Collect_neg_eq)
-
-lemma atom_dom_eq[simp]:
- fixes G::\<Gamma>
- shows "atom_dom ((x, b, c) #\<^sub>\<Gamma> G) = atom_dom ((x, b, c') #\<^sub>\<Gamma> G)"
- using atom_dom.simps toSet.simps by simp
-
-lemma dom_append[simp]:
- "atom_dom (\<Gamma>@\<Gamma>') = atom_dom \<Gamma> \<union> atom_dom \<Gamma>'"
- using image_Un append_g_toSetU atom_dom.simps dom.simps by metis
-
-lemma dom_cons[simp]:
- "atom_dom ((x,b,c) #\<^sub>\<Gamma> G) = { atom x } \<union> atom_dom G"
- using image_Un append_g_toSetU atom_dom.simps by auto
-
-lemma fresh_GNil[ms_fresh]:
- shows "a \<sharp> GNil"
- by (simp add: fresh_def supp_GNil)
-
-lemma fresh_GCons[ms_fresh]:
- fixes xs::\<Gamma>
- shows "a \<sharp> (x #\<^sub>\<Gamma> xs) \<longleftrightarrow> a \<sharp> x \<and> a \<sharp> xs"
- by (simp add: fresh_def supp_GCons)
-
-lemma dom_supp_g[simp]:
- "atom_dom G \<subseteq> supp G"
- apply(induct G rule: \<Gamma>_induct,simp)
- using supp_at_base supp_Pair atom_dom.simps supp_GCons by fastforce
-
-lemma fresh_append_g[ms_fresh]:
- fixes xs::\<Gamma>
- shows "a \<sharp> (xs @ ys) \<longleftrightarrow> a \<sharp> xs \<and> a \<sharp> ys"
- by (induct xs) (simp_all add: fresh_GNil fresh_GCons)
-
-lemma append_g_assoc:
- fixes xs::\<Gamma>
- shows "(xs @ ys) @ zs = xs @ (ys @ zs)"
- by (induct xs) simp_all
-
-lemma append_g_inside:
- fixes xs::\<Gamma>
- shows "xs @ (x #\<^sub>\<Gamma> ys) = (xs @ (x #\<^sub>\<Gamma> GNil)) @ ys"
- by(induct xs,auto+)
-
-lemma finite_\<Gamma>:
- "finite (toSet \<Gamma>)"
- by(induct \<Gamma> rule: \<Gamma>_induct,auto)
-
-lemma supp_\<Gamma>:
- "supp \<Gamma> = supp (toSet \<Gamma>)"
-proof(induct \<Gamma> rule: \<Gamma>_induct)
- case GNil
- then show ?case using supp_GNil toSet.simps
- by (simp add: supp_set_empty)
-next
- case (GCons x b c \<Gamma>')
- then show ?case using supp_GCons toSet.simps finite_\<Gamma> supp_of_finite_union
- using supp_of_finite_insert by fastforce
-qed
-
-lemma supp_of_subset:
- fixes G::"('a::fs set)"
- assumes "finite G" and "finite G'" and "G \<subseteq> G'"
- shows "supp G \<subseteq> supp G'"
- using supp_of_finite_sets assms by (metis subset_Un_eq supp_of_finite_union)
-
-lemma supp_weakening:
- assumes "toSet G \<subseteq> toSet G'"
- shows "supp G \<subseteq> supp G'"
- using supp_\<Gamma> finite_\<Gamma> by (simp add: supp_of_subset assms)
-
-lemma fresh_weakening[ms_fresh]:
- assumes "toSet G \<subseteq> toSet G'" and "x \<sharp> G'"
- shows "x \<sharp> G"
-proof(rule ccontr)
- assume "\<not> x \<sharp> G"
- hence "x \<in> supp G" using fresh_def by auto
- hence "x \<in> supp G'" using supp_weakening assms by auto
- thus False using fresh_def assms by auto
-qed
-
-instance \<Gamma> :: fs
- by (standard, induct_tac x, simp_all add: supp_GNil supp_GCons finite_supp)
-
-lemma fresh_gamma_elem:
- fixes \<Gamma>::\<Gamma>
- assumes "a \<sharp> \<Gamma>"
- and "e \<in> toSet \<Gamma>"
- shows "a \<sharp> e"
- using assms by(induct \<Gamma>,auto simp add: fresh_GCons)
-
-lemma fresh_gamma_append:
- fixes xs::\<Gamma>
- shows "a \<sharp> (xs @ ys) \<longleftrightarrow> a \<sharp> xs \<and> a \<sharp> ys"
- by (induct xs, simp_all add: fresh_GNil fresh_GCons)
-
-lemma supp_triple[simp]:
- shows "supp (x, y, z) = supp x \<union> supp y \<union> supp z"
-proof -
- have "supp (x,y,z) = supp (x,(y,z))" by auto
- hence "supp (x,y,z) = supp x \<union> (supp y \<union> supp z)" using supp_Pair by metis
- thus ?thesis by auto
-qed
-
-lemma supp_append_g:
- fixes xs::\<Gamma>
- shows "supp (xs @ ys) = supp xs \<union> supp ys"
- by(induct xs,auto simp add: supp_GNil supp_GCons )
-
-lemma fresh_in_g[simp]:
- fixes \<Gamma>::\<Gamma> and x'::x
- shows "atom x' \<sharp> \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> = (atom x' \<notin> supp \<Gamma>' \<union> supp x \<union> supp b0 \<union> supp c0 \<union> supp \<Gamma>)"
-proof -
- have "atom x' \<sharp> \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> = (atom x' \<notin> supp (\<Gamma>' @((x,b0,c0) #\<^sub>\<Gamma> \<Gamma>)))"
- using fresh_def by auto
- also have "... = (atom x' \<notin> supp \<Gamma>' \<union> supp ((x,b0,c0) #\<^sub>\<Gamma> \<Gamma>))" using supp_append_g by fast
- also have "... = (atom x' \<notin> supp \<Gamma>' \<union> supp x \<union> supp b0 \<union> supp c0 \<union> supp \<Gamma>)" using supp_GCons supp_append_g supp_triple by auto
- finally show ?thesis by fast
-qed
-
-lemma fresh_suffix[ms_fresh]:
- fixes \<Gamma>::\<Gamma>
- assumes "atom x \<sharp> \<Gamma>'@\<Gamma>"
- shows "atom x \<sharp> \<Gamma>"
- using assms by(induct \<Gamma>' rule: \<Gamma>_induct, auto simp add: append_g.simps fresh_GCons)
-
-lemma not_GCons_self [simp]:
- fixes xs::\<Gamma>
- shows "xs \<noteq> x #\<^sub>\<Gamma> xs"
- by (induct xs) auto
-
-lemma not_GCons_self2 [simp]:
- fixes xs::\<Gamma>
- shows "x #\<^sub>\<Gamma> xs \<noteq> xs"
- by (rule not_GCons_self [symmetric])
-
-lemma fresh_restrict:
- fixes y::x and \<Gamma>::\<Gamma>
- assumes "atom y \<sharp> (\<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>)"
- shows "atom y \<sharp> (\<Gamma>'@\<Gamma>)"
- using assms by(induct \<Gamma>' rule: \<Gamma>_induct, auto simp add:fresh_GCons fresh_GNil )
-
-lemma fresh_dom_free:
- assumes "atom x \<sharp> \<Gamma>"
- shows "(x,b,c) \<notin> toSet \<Gamma>"
- using assms proof(induct \<Gamma> rule: \<Gamma>_induct)
- case GNil
- then show ?case by auto
-next
- case (GCons x' b' c' \<Gamma>')
- hence "x\<noteq>x'" using fresh_def fresh_GCons fresh_Pair supp_at_base by blast
- moreover have "atom x \<sharp> \<Gamma>'" using fresh_GCons GCons by auto
- ultimately show ?case using toSet.simps GCons by auto
-qed
-
-lemma \<Gamma>_set_intros: "x \<in> toSet ( x #\<^sub>\<Gamma> xs)" and "y \<in> toSet xs \<Longrightarrow> y \<in> toSet (x #\<^sub>\<Gamma> xs)"
- by simp+
-
-lemma fresh_dom_free2:
- assumes "atom x \<notin> atom_dom \<Gamma>"
- shows "(x,b,c) \<notin> toSet \<Gamma>"
- using assms proof(induct \<Gamma> rule: \<Gamma>_induct)
- case GNil
- then show ?case by auto
-next
- case (GCons x' b' c' \<Gamma>')
- hence "x\<noteq>x'" using fresh_def fresh_GCons fresh_Pair supp_at_base by auto
- moreover have "atom x \<notin> atom_dom \<Gamma>'" using fresh_GCons GCons by auto
- ultimately show ?case using toSet.simps GCons by auto
-qed
-
-subsection \<open>Mutable Variable Context Lemmas\<close>
-
-lemma supp_DNil:
- shows "supp DNil = {}"
- by (simp add: supp_def)
-
-lemma supp_DCons:
- fixes xs::\<Delta>
- shows "supp (x #\<^sub>\<Delta> xs) = supp x \<union> supp xs"
- by (simp add: supp_def Collect_imp_eq Collect_neg_eq)
-
-lemma fresh_DNil[ms_fresh]:
- shows "a \<sharp> DNil"
- by (simp add: fresh_def supp_DNil)
-
-lemma fresh_DCons[ms_fresh]:
- fixes xs::\<Delta>
- shows "a \<sharp> (x #\<^sub>\<Delta> xs) \<longleftrightarrow> a \<sharp> x \<and> a \<sharp> xs"
- by (simp add: fresh_def supp_DCons)
-
-instance \<Delta> :: fs
- by (standard, induct_tac x, simp_all add: supp_DNil supp_DCons finite_supp)
-
-subsection \<open>Lookup Functions\<close>
-
-nominal_function lookup :: "\<Gamma> \<Rightarrow> x \<Rightarrow> (b*c) option" where
- "lookup GNil x = None"
-| "lookup ((x,b,c)#\<^sub>\<Gamma>G) y = (if x=y then Some (b,c) else lookup G y)"
- by (auto,simp add: eqvt_def lookup_graph_aux_def, metis neq_GNil_conv surj_pair)
-nominal_termination (eqvt) by lexicographic_order
-
-nominal_function replace_in_g :: "\<Gamma> \<Rightarrow> x \<Rightarrow> c \<Rightarrow> \<Gamma>" ("_[_\<longmapsto>_]" [1000,0,0] 200) where
- "replace_in_g GNil _ _ = GNil"
-| "replace_in_g ((x,b,c)#\<^sub>\<Gamma>G) x' c' = (if x=x' then ((x,b,c')#\<^sub>\<Gamma>G) else (x,b,c)#\<^sub>\<Gamma>(replace_in_g G x' c'))"
- apply(auto,simp add: eqvt_def replace_in_g_graph_aux_def)
- using surj_pair \<Gamma>.exhaust by metis
-nominal_termination (eqvt) by lexicographic_order
-
-text \<open> Functions for looking up data-constructors in the Pi context \<close>
-
-nominal_function lookup_fun :: "\<Phi> \<Rightarrow> f \<Rightarrow> fun_def option" where
- "lookup_fun [] g = None"
-| "lookup_fun ((AF_fundef f ft)#\<Pi>) g = (if (f=g) then Some (AF_fundef f ft) else lookup_fun \<Pi> g)"
- apply(auto,simp add: eqvt_def lookup_fun_graph_aux_def )
- by (metis fun_def.exhaust neq_Nil_conv)
-nominal_termination (eqvt) by lexicographic_order
-
-nominal_function lookup_td :: "\<Theta> \<Rightarrow> string \<Rightarrow> type_def option" where
- "lookup_td [] g = None"
-| "lookup_td ((AF_typedef s lst ) # (\<Theta>::\<Theta>)) g = (if (s = g) then Some (AF_typedef s lst ) else lookup_td \<Theta> g)"
-| "lookup_td ((AF_typedef_poly s bv lst ) # (\<Theta>::\<Theta>)) g = (if (s = g) then Some (AF_typedef_poly s bv lst ) else lookup_td \<Theta> g)"
- apply(auto,simp add: eqvt_def lookup_td_graph_aux_def )
- by (metis type_def.exhaust neq_Nil_conv)
-nominal_termination (eqvt) by lexicographic_order
-
-nominal_function name_of_type ::"type_def \<Rightarrow> f " where
- "name_of_type (AF_typedef f _ ) = f"
-| "name_of_type (AF_typedef_poly f _ _) = f"
- apply(auto,simp add: eqvt_def name_of_type_graph_aux_def )
- using type_def.exhaust by blast
-nominal_termination (eqvt) by lexicographic_order
-
-nominal_function name_of_fun ::"fun_def \<Rightarrow> f " where
- "name_of_fun (AF_fundef f ft) = f"
- apply(auto,simp add: eqvt_def name_of_fun_graph_aux_def )
- using fun_def.exhaust by blast
-nominal_termination (eqvt) by lexicographic_order
-
-nominal_function remove2 :: "'a::pt \<Rightarrow> 'a list \<Rightarrow> 'a list" where
- "remove2 x [] = []" |
- "remove2 x (y # xs) = (if x = y then xs else y # remove2 x xs)"
- by (simp add: eqvt_def remove2_graph_aux_def,auto+,meson list.exhaust)
-nominal_termination (eqvt) by lexicographic_order
-
-nominal_function base_for_lit :: "l \<Rightarrow> b" where
- "base_for_lit (L_true) = B_bool "
-| "base_for_lit (L_false) = B_bool "
-| "base_for_lit (L_num n) = B_int "
-| "base_for_lit (L_unit) = B_unit "
-| "base_for_lit (L_bitvec v) = B_bitvec"
- apply (auto simp: eqvt_def base_for_lit_graph_aux_def )
- using l.strong_exhaust by blast
-nominal_termination (eqvt) by lexicographic_order
-
-lemma neq_DNil_conv: "(xs \<noteq> DNil) = (\<exists>y ys. xs = y #\<^sub>\<Delta> ys)"
- by (induct xs) auto
-
-nominal_function setD :: "\<Delta> \<Rightarrow> (u*\<tau>) set" where
- "setD DNil = {}"
-| "setD (DCons xbc G) = {xbc} \<union> (setD G)"
- apply (auto,simp add: eqvt_def setD_graph_aux_def )
- using neq_DNil_conv surj_pair by metis
-nominal_termination (eqvt) by lexicographic_order
-
-lemma eqvt_triple:
- fixes y::"'a::at" and ya::"'a::at" and xa::"'c::at" and va::"'d::fs" and s::s and sa::s and f::"s*'c*'d \<Rightarrow> s"
- assumes "atom y \<sharp> (xa, va)" and "atom ya \<sharp> (xa, va)" and
- "\<forall>c. atom c \<sharp> (s, sa) \<longrightarrow> atom c \<sharp> (y, ya, s, sa) \<longrightarrow> (y \<leftrightarrow> c) \<bullet> s = (ya \<leftrightarrow> c) \<bullet> sa"
- and "eqvt_at f (s,xa,va)" and "eqvt_at f (sa,xa,va)" and
- "atom c \<sharp> (s, va, xa, sa)" and "atom c \<sharp> (y, ya, f (s, xa, va), f (sa, xa, va))"
- shows "(y \<leftrightarrow> c) \<bullet> f (s, xa, va) = (ya \<leftrightarrow> c) \<bullet> f (sa, xa, va)"
-proof -
- have " (y \<leftrightarrow> c) \<bullet> f (s, xa, va) = f ( (y \<leftrightarrow> c) \<bullet> (s,xa,va))" using assms eqvt_at_def by metis
- also have "... = f ( (y \<leftrightarrow> c) \<bullet> s, (y \<leftrightarrow> c) \<bullet> xa ,(y \<leftrightarrow> c) \<bullet> va)" by auto
- also have "... = f ( (ya \<leftrightarrow> c) \<bullet> sa, (ya \<leftrightarrow> c) \<bullet> xa ,(ya \<leftrightarrow> c) \<bullet> va)" proof -
- have " (y \<leftrightarrow> c) \<bullet> s = (ya \<leftrightarrow> c) \<bullet> sa" using assms Abs1_eq_iff_all by auto
- moreover have "((y \<leftrightarrow> c) \<bullet> xa) = ((ya \<leftrightarrow> c) \<bullet> xa)" using assms flip_fresh_fresh fresh_prodN by metis
- moreover have "((y \<leftrightarrow> c) \<bullet> va) = ((ya \<leftrightarrow> c) \<bullet> va)" using assms flip_fresh_fresh fresh_prodN by metis
- ultimately show ?thesis by auto
- qed
- also have "... = f ( (ya \<leftrightarrow> c) \<bullet> (sa,xa,va))" by auto
- finally show ?thesis using assms eqvt_at_def by metis
-qed
-
-section \<open>Functions for bit list/vectors\<close>
-
-inductive split :: "int \<Rightarrow> bit list \<Rightarrow> bit list * bit list \<Rightarrow> bool" where
- "split 0 xs ([], xs)"
-| "split m xs (ys,zs) \<Longrightarrow> split (m+1) (x#xs) ((x # ys), zs)"
-equivariance split
-nominal_inductive split .
-
-lemma split_concat:
- assumes "split n v (v1,v2)"
- shows "v = append v1 v2"
- using assms proof(induct "(v1,v2)" arbitrary: v1 v2 rule: split.inducts)
- case 1
- then show ?case by auto
-next
- case (2 m xs ys zs x)
- then show ?case by auto
-qed
-
-lemma split_n:
- assumes "split n v (v1,v2)"
- shows "0 \<le> n \<and> n \<le> int (length v)"
- using assms proof(induct rule: split.inducts)
- case (1 xs)
- then show ?case by auto
-next
- case (2 m xs ys zs x)
- then show ?case by auto
-qed
-
-lemma split_length:
- assumes "split n v (v1,v2)"
- shows "n = int (length v1)"
- using assms proof(induct "(v1,v2)" arbitrary: v1 v2 rule: split.inducts)
- case (1 xs)
- then show ?case by auto
-next
- case (2 m xs ys zs x)
- then show ?case by auto
-qed
-
-lemma obtain_split:
- assumes "0 \<le> n" and "n \<le> int (length bv)"
- shows "\<exists> bv1 bv2. split n bv (bv1 , bv2)"
- using assms proof(induct bv arbitrary: n)
- case Nil
- then show ?case using split.intros by auto
-next
- case (Cons b bv)
- show ?case proof(cases "n = 0")
- case True
- then show ?thesis using split.intros by auto
- next
- case False
- then obtain m where m:"n=m+1" using Cons
- by (metis add.commute add_minus_cancel)
- moreover have "0 \<le> m" using False m Cons by linarith
- then obtain bv1 and bv2 where "split m bv (bv1 , bv2)" using Cons m by force
- hence "split n (b # bv) ((b#bv1), bv2)" using m split.intros by auto
- then show ?thesis by auto
- qed
-qed
-
+(*<*)
+theory Syntax
+ imports "Nominal2.Nominal2" "Nominal-Utils"
+begin
+ (*>*)
+
+chapter \<open>Syntax\<close>
+
+text \<open>Syntax of MiniSail programs and the contexts we use in judgements.\<close>
+
+section \<open>Program Syntax\<close>
+
+subsection \<open>AST Datatypes\<close>
+
+type_synonym num_nat = nat
+
+atom_decl x (* Immutable variables*)
+atom_decl u (* Mutable variables *)
+atom_decl bv (* Basic-type variables *)
+
+type_synonym f = string (* Function name *)
+type_synonym dc = string (* Data constructor *)
+type_synonym tyid = string
+
+text \<open>Basic types. Types without refinement constraints\<close>
+nominal_datatype "b" =
+ B_int | B_bool | B_id tyid
+| B_pair b b ("[ _ , _ ]\<^sup>b")
+| B_unit | B_bitvec | B_var bv
+| B_app tyid b
+
+nominal_datatype bit = BitOne | BitZero
+
+text \<open> Literals \<close>
+nominal_datatype "l" =
+ L_num "int" | L_true | L_false | L_unit | L_bitvec "bit list"
+
+text \<open> Values. We include a type identifier, tyid, in the literal for constructors
+to make typing and well-formedness checking easier \<close>
+
+nominal_datatype "v" =
+ V_lit "l" ( "[ _ ]\<^sup>v")
+ | V_var "x" ( "[ _ ]\<^sup>v")
+ | V_pair "v" "v" ( "[ _ , _ ]\<^sup>v")
+ | V_cons tyid dc "v"
+ | V_consp tyid dc b "v"
+
+text \<open> Binary Operations \<close>
+nominal_datatype "opp" = Plus ( "plus") | LEq ("leq") | Eq ("eq")
+
+text \<open> Expressions \<close>
+nominal_datatype "e" =
+ AE_val "v" ( "[ _ ]\<^sup>e" )
+ | AE_app "f" "v" ( "[ _ ( _ ) ]\<^sup>e" )
+ | AE_appP "f" "b" "v" ( "[_ [ _ ] ( _ )]\<^sup>e" )
+ | AE_op "opp" "v" "v" ( "[ _ _ _ ]\<^sup>e" )
+ | AE_concat v v ( "[ _ @@ _ ]\<^sup>e" )
+ | AE_fst "v" ( "[#1_ ]\<^sup>e" )
+ | AE_snd "v" ( "[#2_ ]\<^sup>e" )
+ | AE_mvar "u" ( "[ _ ]\<^sup>e" )
+ | AE_len "v" ( "[| _ |]\<^sup>e" )
+ | AE_split "v" "v" ( "[ _ / _ ]\<^sup>e" )
+
+text \<open> Expressions for constraints\<close>
+nominal_datatype "ce" =
+ CE_val "v" ( "[ _ ]\<^sup>c\<^sup>e" )
+ | CE_op "opp" "ce" "ce" ( "[ _ _ _ ]\<^sup>c\<^sup>e" )
+ | CE_concat ce ce ( "[ _ @@ _ ]\<^sup>c\<^sup>e" )
+ | CE_fst "ce" ( "[#1_]\<^sup>c\<^sup>e" )
+ | CE_snd "ce" ( "[#2_]\<^sup>c\<^sup>e" )
+ | CE_len "ce" ( "[| _ |]\<^sup>c\<^sup>e" )
+
+text \<open> Constraints \<close>
+nominal_datatype "c" =
+ C_true ( "TRUE" [] 50 )
+ | C_false ( "FALSE" [] 50 )
+ | C_conj "c" "c" ("_ AND _ " [50, 50] 50)
+ | C_disj "c" "c" ("_ OR _ " [50,50] 50)
+ | C_not "c" ( "\<not> _ " [] 50 )
+ | C_imp "c" "c" ("_ IMP _ " [50, 50] 50)
+ | C_eq "ce" "ce" ("_ == _ " [50, 50] 50)
+
+text \<open> Refined types \<close>
+nominal_datatype "\<tau>" =
+ T_refined_type x::x b c::c binds x in c ("\<lbrace> _ : _ | _ \<rbrace>" [50, 50] 1000)
+
+text \<open> Statements \<close>
+
+nominal_datatype
+ s =
+ AS_val v ( "[_]\<^sup>s")
+ | AS_let x::x e s::s binds x in s ( "(LET _ = _ IN _)")
+ | AS_let2 x::x \<tau> s s::s binds x in s ( "(LET _ : _ = _ IN _)")
+ | AS_if v s s ( "(IF _ THEN _ ELSE _)" [0, 61, 0] 61)
+ | AS_var u::u \<tau> v s::s binds u in s ( "(VAR _ : _ = _ IN _)")
+ | AS_assign u v ( "(_ ::= _)")
+ | AS_match v branch_list ( "(MATCH _ WITH { _ })")
+ | AS_while s s ( "(WHILE _ DO { _ } )" [0, 0] 61)
+ | AS_seq s s ( "( _ ;; _ )" [1000, 61] 61)
+ | AS_assert c s ( "(ASSERT _ IN _ )" )
+ and branch_s =
+ AS_branch dc x::x s::s binds x in s ( "( _ _ \<Rightarrow> _ )")
+ and branch_list =
+ AS_final branch_s ( "{ _ }" )
+ | AS_cons branch_s branch_list ( "( _ | _ )")
+
+text \<open> Function and union type definitions \<close>
+
+nominal_datatype "fun_typ" =
+ AF_fun_typ x::x "b" c::c \<tau>::\<tau> s::s binds x in c \<tau> s
+
+nominal_datatype "fun_typ_q" =
+ AF_fun_typ_some bv::bv ft::fun_typ binds bv in ft
+ | AF_fun_typ_none fun_typ
+
+nominal_datatype "fun_def" = AF_fundef f fun_typ_q
+
+nominal_datatype "type_def" =
+ AF_typedef string "(string * \<tau>) list"
+ | AF_typedef_poly string bv::bv dclist::"(string * \<tau>) list" binds bv in dclist
+
+lemma check_typedef_poly:
+ "AF_typedef_poly ''option'' bv [ (''None'', \<lbrace> zz : B_unit | TRUE \<rbrace>), (''Some'', \<lbrace> zz : B_var bv | TRUE \<rbrace>) ] =
+ AF_typedef_poly ''option'' bv2 [ (''None'', \<lbrace> zz : B_unit | TRUE \<rbrace>), (''Some'', \<lbrace> zz : B_var bv2 | TRUE \<rbrace>) ]"
+ by auto
+
+nominal_datatype "var_def" = AV_def u \<tau> v
+
+text \<open> Programs \<close>
+nominal_datatype "p" =
+ AP_prog "type_def list" "fun_def list" "var_def list" "s" ("PROG _ _ _ _")
+
+declare l.supp [simp] v.supp [simp] e.supp [simp] s_branch_s_branch_list.supp [simp] \<tau>.supp [simp] c.supp [simp] b.supp[simp]
+
+subsection \<open>Lemmas\<close>
+
+text \<open>These lemmas deal primarily with freshness and alpha-equivalence\<close>
+
+subsubsection \<open>Atoms\<close>
+
+lemma x_not_in_u_atoms[simp]:
+ fixes u::u and x::x and us::"u set"
+ shows "atom x \<notin> atom`us"
+ by (simp add: image_iff)
+
+lemma x_fresh_u[simp]:
+ fixes u::u and x::x
+ shows "atom x \<sharp> u"
+ by auto
+
+lemma x_not_in_b_set[simp]:
+ fixes x::x and bs::"bv fset"
+ shows "atom x \<notin> supp bs"
+ by(induct bs,auto, simp add: supp_finsert supp_at_base)
+
+lemma x_fresh_b[simp]:
+ fixes x::x and b::b
+ shows "atom x \<sharp> b"
+ apply (induct b rule: b.induct, auto simp: pure_supp)
+ using pure_supp fresh_def by blast+
+
+lemma x_fresh_bv[simp]:
+ fixes x::x and bv::bv
+ shows "atom x \<sharp> bv"
+ using fresh_def supp_at_base by auto
+
+lemma u_not_in_x_atoms[simp]:
+ fixes u::u and x::x and xs::"x set"
+ shows "atom u \<notin> atom`xs"
+ by (simp add: image_iff)
+
+lemma bv_not_in_x_atoms[simp]:
+ fixes bv::bv and x::x and xs::"x set"
+ shows "atom bv \<notin> atom`xs"
+ by (simp add: image_iff)
+
+lemma u_not_in_b_atoms[simp]:
+ fixes b :: b and u::u
+ shows "atom u \<notin> supp b"
+ by (induct b rule: b.induct,auto simp: pure_supp supp_at_base)
+
+lemma u_not_in_b_set[simp]:
+ fixes u::u and bs::"bv fset"
+ shows "atom u \<notin> supp bs"
+ by(induct bs, auto simp add: supp_at_base supp_finsert)
+
+lemma u_fresh_b[simp]:
+ fixes x::u and b::b
+ shows "atom x \<sharp> b"
+ by(induct b rule: b.induct, auto simp: pure_fresh )
+
+lemma supp_b_v_disjoint:
+ fixes x::x and bv::bv
+ shows "supp (V_var x) \<inter> supp (B_var bv) = {}"
+ by (simp add: supp_at_base)
+
+lemma supp_b_u_disjoint[simp]:
+ fixes b::b and u::u
+ shows "supp u \<inter> supp b = {}"
+ by(nominal_induct b rule:b.strong_induct,(auto simp add: pure_supp b.supp supp_at_base)+)
+
+lemma u_fresh_bv[simp]:
+ fixes u::u and b::bv
+ shows "atom u \<sharp> b"
+ using fresh_at_base by simp
+
+subsubsection \<open>Basic Types\<close>
+
+nominal_function b_of :: "\<tau> \<Rightarrow> b" where
+ "b_of \<lbrace> z : b | c \<rbrace> = b"
+ apply(auto,simp add: eqvt_def b_of_graph_aux_def )
+ by (meson \<tau>.exhaust)
+nominal_termination (eqvt) by lexicographic_order
+
+lemma supp_b_empty[simp]:
+ fixes b :: b and x::x
+ shows "atom x \<notin> supp b"
+ by (induct b rule: b.induct, auto simp: pure_supp supp_at_base x_not_in_b_set)
+
+lemma flip_b_id[simp]:
+ fixes x::x and b::b
+ shows "(x \<leftrightarrow> x') \<bullet> b = b"
+ by(rule flip_fresh_fresh, auto simp add: fresh_def)
+
+lemma flip_x_b_cancel[simp]:
+ fixes x::x and y::x and b::b and bv::bv
+ shows "(x \<leftrightarrow> y) \<bullet> b = b" and "(x \<leftrightarrow> y) \<bullet> bv = bv"
+ using flip_b_id apply simp
+ by (metis b.eq_iff(7) b.perm_simps(7) flip_b_id)
+
+lemma flip_bv_x_cancel[simp]:
+ fixes bv::bv and z::bv and x::x
+ shows "(bv \<leftrightarrow> z) \<bullet> x = x" using flip_fresh_fresh[of bv x z] fresh_at_base by auto
+
+lemma flip_bv_u_cancel[simp]:
+ fixes bv::bv and z::bv and x::u
+ shows "(bv \<leftrightarrow> z) \<bullet> x = x" using flip_fresh_fresh[of bv x z] fresh_at_base by auto
+
+subsubsection \<open>Literals\<close>
+
+lemma supp_bitvec_empty:
+ fixes bv::"bit list"
+ shows "supp bv = {}"
+proof(induct bv)
+ case Nil
+ then show ?case using supp_Nil by auto
+next
+ case (Cons a bv)
+ then show ?case using supp_Cons bit.supp
+ by (metis (mono_tags, opaque_lifting) bit.strong_exhaust l.supp(5) sup_bot.right_neutral)
+qed
+
+lemma bitvec_pure[simp]:
+ fixes bv::"bit list" and x::x
+ shows "atom x \<sharp> bv" using fresh_def supp_bitvec_empty by auto
+
+lemma supp_l_empty[simp]:
+ fixes l:: l
+ shows "supp (V_lit l) = {}"
+ by(nominal_induct l rule: l.strong_induct,
+ auto simp add: l.strong_exhaust pure_supp v.fv_defs supp_bitvec_empty)
+
+lemma type_l_nosupp[simp]:
+ fixes x::x and l::l
+ shows "atom x \<notin> supp (\<lbrace> z : b | [[z]\<^sup>v]\<^sup>c\<^sup>e == [[l]\<^sup>v]\<^sup>c\<^sup>e \<rbrace>)"
+ using supp_at_base supp_l_empty ce.supp(1) c.supp \<tau>.supp by force
+
+lemma flip_bitvec0:
+ fixes x::"bit list"
+ assumes "atom c \<sharp> (z, x, z')"
+ shows "(z \<leftrightarrow> c) \<bullet> x = (z' \<leftrightarrow> c) \<bullet> x"
+proof -
+ have "atom z \<sharp> x" and "atom z' \<sharp> x"
+ using flip_fresh_fresh assms supp_bitvec_empty fresh_def by blast+
+ moreover have "atom c \<sharp> x" using supp_bitvec_empty fresh_def by auto
+ ultimately show ?thesis using assms flip_fresh_fresh by metis
+qed
+
+lemma flip_bitvec:
+ assumes "atom c \<sharp> (z, L_bitvec x, z')"
+ shows "(z \<leftrightarrow> c) \<bullet> x = (z' \<leftrightarrow> c) \<bullet> x"
+proof -
+ have "atom z \<sharp> x" and "atom z' \<sharp> x"
+ using flip_fresh_fresh assms supp_bitvec_empty fresh_def by blast+
+ moreover have "atom c \<sharp> x" using supp_bitvec_empty fresh_def by auto
+ ultimately show ?thesis using assms flip_fresh_fresh by metis
+qed
+
+lemma type_l_eq:
+ shows "\<lbrace> z : b | [[z]\<^sup>v]\<^sup>c\<^sup>e == [V_lit l]\<^sup>c\<^sup>e \<rbrace> = (\<lbrace> z' : b | [[z']\<^sup>v]\<^sup>c\<^sup>e == [V_lit l]\<^sup>c\<^sup>e \<rbrace>)"
+ by(auto,nominal_induct l rule: l.strong_induct,auto, metis permute_pure, auto simp add: flip_bitvec)
+
+lemma flip_l_eq:
+ fixes x::l
+ shows "(z \<leftrightarrow> c) \<bullet> x = (z' \<leftrightarrow> c) \<bullet> x"
+proof -
+ have "atom z \<sharp> x" and "atom c \<sharp> x" and "atom z' \<sharp> x"
+ using flip_fresh_fresh fresh_def supp_l_empty by fastforce+
+ thus ?thesis using flip_fresh_fresh by metis
+qed
+
+lemma flip_l_eq1:
+ fixes x::l
+ assumes "(z \<leftrightarrow> c) \<bullet> x = (z' \<leftrightarrow> c) \<bullet> x'"
+ shows "x' = x"
+proof -
+ have "atom z \<sharp> x" and "atom c \<sharp> x'" and "atom c \<sharp> x" and "atom z' \<sharp> x'"
+ using flip_fresh_fresh fresh_def supp_l_empty by fastforce+
+ thus ?thesis using flip_fresh_fresh assms by metis
+qed
+
+subsubsection \<open>Types\<close>
+
+lemma flip_base_eq:
+ fixes b::b and x::x and y::x
+ shows "(x \<leftrightarrow> y) \<bullet> b = b"
+ using b.fresh by (simp add: flip_fresh_fresh fresh_def)
+
+text \<open> Obtain an alpha-equivalent type where the bound variable is fresh in some term t \<close>
+lemma has_fresh_z0:
+ fixes t::"'b::fs"
+ shows "\<exists>z. atom z \<sharp> (c',t) \<and> (\<lbrace>z' : b | c' \<rbrace>) = (\<lbrace> z : b | (z \<leftrightarrow> z' ) \<bullet> c' \<rbrace>)"
+proof -
+ obtain z::x where fr: "atom z \<sharp> (c',t)" using obtain_fresh by blast
+ moreover hence "(\<lbrace> z' : b | c' \<rbrace>) = (\<lbrace> z : b | (z \<leftrightarrow> z') \<bullet> c' \<rbrace>)"
+ using \<tau>.eq_iff Abs1_eq_iff
+ by (metis flip_commute flip_fresh_fresh fresh_PairD(1))
+ ultimately show ?thesis by fastforce
+qed
+
+lemma has_fresh_z:
+ fixes t::"'b::fs"
+ shows "\<exists>z b c. atom z \<sharp> t \<and> \<tau> = \<lbrace> z : b | c \<rbrace>"
+proof -
+ obtain z' and b and c' where teq: "\<tau> = (\<lbrace> z' : b | c' \<rbrace>)" using \<tau>.exhaust by blast
+ obtain z::x where fr: "atom z \<sharp> (t,c')" using obtain_fresh by blast
+ hence "(\<lbrace> z' : b | c' \<rbrace>) = (\<lbrace> z : b | (z \<leftrightarrow> z') \<bullet> c' \<rbrace>)" using \<tau>.eq_iff Abs1_eq_iff
+ flip_commute flip_fresh_fresh fresh_PairD(1) by (metis fresh_PairD(2))
+ hence "atom z \<sharp> t \<and> \<tau> = (\<lbrace> z : b | (z \<leftrightarrow> z') \<bullet> c' \<rbrace>)" using fr teq by force
+ thus ?thesis using teq fr by fast
+qed
+
+lemma obtain_fresh_z:
+ fixes t::"'b::fs"
+ obtains z and b and c where "atom z \<sharp> t \<and> \<tau> = \<lbrace> z : b | c \<rbrace>"
+ using has_fresh_z by blast
+
+lemma has_fresh_z2:
+ fixes t::"'b::fs"
+ shows "\<exists>z c. atom z \<sharp> t \<and> \<tau> = \<lbrace> z : b_of \<tau> | c \<rbrace>"
+proof -
+ obtain z and b and c where "atom z \<sharp> t \<and> \<tau> = \<lbrace> z : b | c \<rbrace>" using obtain_fresh_z by metis
+ moreover then have "b_of \<tau> = b" using \<tau>.eq_iff by simp
+ ultimately show ?thesis using obtain_fresh_z \<tau>.eq_iff by auto
+qed
+
+lemma obtain_fresh_z2:
+ fixes t::"'b::fs"
+ obtains z and c where "atom z \<sharp> t \<and> \<tau> = \<lbrace> z : b_of \<tau> | c \<rbrace>"
+ using has_fresh_z2 by blast
+
+subsubsection \<open>Values\<close>
+
+lemma u_notin_supp_v[simp]:
+ fixes u::u and v::v
+ shows "atom u \<notin> supp v"
+proof(nominal_induct v rule: v.strong_induct)
+ case (V_lit l)
+ then show ?case using supp_l_empty by auto
+next
+ case (V_var x)
+ then show ?case
+ by (simp add: supp_at_base)
+next
+ case (V_pair v1 v2)
+ then show ?case by auto
+next
+ case (V_cons tyid list v)
+ then show ?case using pure_supp by auto
+next
+ case (V_consp tyid list b v)
+ then show ?case using pure_supp by auto
+qed
+
+lemma u_fresh_xv[simp]:
+ fixes u::u and x::x and v::v
+ shows "atom u \<sharp> (x,v)"
+proof -
+ have "atom u \<sharp> x" using fresh_def by fastforce
+ moreover have "atom u \<sharp> v" using fresh_def u_notin_supp_v by metis
+ ultimately show ?thesis using fresh_prod2 by auto
+qed
+
+text \<open> Part of an effort to make the proofs across inductive cases more uniform by distilling
+the non-uniform parts into lemmas like this\<close>
+lemma v_flip_eq:
+ fixes v::v and va::v and x::x and c::x
+ assumes "atom c \<sharp> (v, va)" and "atom c \<sharp> (x, xa, v, va)" and "(x \<leftrightarrow> c) \<bullet> v = (xa \<leftrightarrow> c) \<bullet> va"
+ shows "((v = V_lit l \<longrightarrow> (\<exists>l'. va = V_lit l' \<and> (x \<leftrightarrow> c) \<bullet> l = (xa \<leftrightarrow> c) \<bullet> l'))) \<and>
+ ((v = V_var y \<longrightarrow> (\<exists>y'. va = V_var y' \<and> (x \<leftrightarrow> c) \<bullet> y = (xa \<leftrightarrow> c) \<bullet> y'))) \<and>
+ ((v = V_pair vone vtwo \<longrightarrow> (\<exists>v1' v2'. va = V_pair v1' v2' \<and> (x \<leftrightarrow> c) \<bullet> vone = (xa \<leftrightarrow> c) \<bullet> v1' \<and> (x \<leftrightarrow> c) \<bullet> vtwo = (xa \<leftrightarrow> c) \<bullet> v2'))) \<and>
+ ((v = V_cons tyid dc vone \<longrightarrow> (\<exists>v1'. va = V_cons tyid dc v1'\<and> (x \<leftrightarrow> c) \<bullet> vone = (xa \<leftrightarrow> c) \<bullet> v1'))) \<and>
+ ((v = V_consp tyid dc b vone \<longrightarrow> (\<exists>v1'. va = V_consp tyid dc b v1'\<and> (x \<leftrightarrow> c) \<bullet> vone = (xa \<leftrightarrow> c) \<bullet> v1')))"
+ using assms proof(nominal_induct v rule:v.strong_induct)
+ case (V_lit l)
+ then show ?case using assms v.perm_simps
+ empty_iff flip_def fresh_def fresh_permute_iff supp_l_empty swap_fresh_fresh v.fresh
+ by (metis permute_swap_cancel2 v.distinct)
+next
+ case (V_var x)
+ then show ?case using assms v.perm_simps
+ empty_iff flip_def fresh_def fresh_permute_iff supp_l_empty swap_fresh_fresh v.fresh
+ by (metis permute_swap_cancel2 v.distinct)
+next
+ case (V_pair v1 v2)
+ have " (V_pair v1 v2 = V_pair vone vtwo \<longrightarrow> (\<exists>v1' v2'. va = V_pair v1' v2' \<and> (x \<leftrightarrow> c) \<bullet> vone = (xa \<leftrightarrow> c) \<bullet> v1' \<and> (x \<leftrightarrow> c) \<bullet> vtwo = (xa \<leftrightarrow> c) \<bullet> v2'))" proof
+ assume "V_pair v1 v2= V_pair vone vtwo"
+ thus "(\<exists>v1' v2'. va = V_pair v1' v2' \<and> (x \<leftrightarrow> c) \<bullet> vone = (xa \<leftrightarrow> c) \<bullet> v1' \<and> (x \<leftrightarrow> c) \<bullet> vtwo = (xa \<leftrightarrow> c) \<bullet> v2')"
+ using V_pair assms
+ by (metis (no_types, opaque_lifting) flip_def permute_swap_cancel v.perm_simps(3))
+ qed
+ thus ?case using V_pair by auto
+next
+ case (V_cons tyid dc v1)
+ have " (V_cons tyid dc v1 = V_cons tyid dc vone \<longrightarrow> (\<exists> v1'. va = V_cons tyid dc v1' \<and> (x \<leftrightarrow> c) \<bullet> vone = (xa \<leftrightarrow> c) \<bullet> v1'))" proof
+ assume as: "V_cons tyid dc v1 = V_cons tyid dc vone"
+ hence "(x \<leftrightarrow> c) \<bullet> (V_cons tyid dc vone) = V_cons tyid dc ((x \<leftrightarrow> c) \<bullet> vone)" proof -
+ have "(x \<leftrightarrow> c) \<bullet> dc = dc" using pure_permute_id by metis
+ moreover have "(x \<leftrightarrow> c) \<bullet> tyid = tyid" using pure_permute_id by metis
+ ultimately show ?thesis using v.perm_simps(4) by simp
+ qed
+ then obtain v1' where "(xa \<leftrightarrow> c) \<bullet> va = V_cons tyid dc v1' \<and> (x \<leftrightarrow> c) \<bullet> vone = v1'" using assms V_cons
+ using as by fastforce
+ hence " va = V_cons tyid dc ((xa \<leftrightarrow> c) \<bullet> v1') \<and> (x \<leftrightarrow> c) \<bullet> vone = v1'" using permute_flip_cancel empty_iff flip_def fresh_def supp_b_empty swap_fresh_fresh
+ by (metis pure_fresh v.perm_simps(4))
+
+ thus "(\<exists>v1'. va = V_cons tyid dc v1' \<and> (x \<leftrightarrow> c) \<bullet> vone = (xa \<leftrightarrow> c) \<bullet> v1')"
+ using V_cons assms by simp
+ qed
+ thus ?case using V_cons by auto
+next
+ case (V_consp tyid dc b v1)
+ have " (V_consp tyid dc b v1 = V_consp tyid dc b vone \<longrightarrow> (\<exists> v1'. va = V_consp tyid dc b v1' \<and> (x \<leftrightarrow> c) \<bullet> vone = (xa \<leftrightarrow> c) \<bullet> v1'))" proof
+ assume as: "V_consp tyid dc b v1 = V_consp tyid dc b vone"
+ hence "(x \<leftrightarrow> c) \<bullet> (V_consp tyid dc b vone) = V_consp tyid dc b ((x \<leftrightarrow> c) \<bullet> vone)" proof -
+ have "(x \<leftrightarrow> c) \<bullet> dc = dc" using pure_permute_id by metis
+ moreover have "(x \<leftrightarrow> c) \<bullet> tyid = tyid" using pure_permute_id by metis
+ ultimately show ?thesis using v.perm_simps(4) by simp
+ qed
+ then obtain v1' where "(xa \<leftrightarrow> c) \<bullet> va = V_consp tyid dc b v1' \<and> (x \<leftrightarrow> c) \<bullet> vone = v1'" using assms V_consp
+ using as by fastforce
+ hence " va = V_consp tyid dc b ((xa \<leftrightarrow> c) \<bullet> v1') \<and> (x \<leftrightarrow> c) \<bullet> vone = v1'" using permute_flip_cancel empty_iff flip_def fresh_def supp_b_empty swap_fresh_fresh
+ pure_fresh v.perm_simps
+ by (metis (mono_tags, opaque_lifting))
+ thus "(\<exists>v1'. va = V_consp tyid dc b v1' \<and> (x \<leftrightarrow> c) \<bullet> vone = (xa \<leftrightarrow> c) \<bullet> v1')"
+ using V_consp assms by simp
+ qed
+ thus ?case using V_consp by auto
+qed
+
+lemma flip_eq:
+ fixes x::x and xa::x and s::"'a::fs" and sa::"'a::fs"
+ assumes "(\<forall>c. atom c \<sharp> (s, sa) \<longrightarrow> atom c \<sharp> (x, xa, s, sa) \<longrightarrow> (x \<leftrightarrow> c) \<bullet> s = (xa \<leftrightarrow> c) \<bullet> sa)" and "x \<noteq> xa"
+ shows "(x \<leftrightarrow> xa) \<bullet> s = sa"
+proof -
+ have " ([[atom x]]lst. s = [[atom xa]]lst. sa)" using assms Abs1_eq_iff_all by simp
+ hence "(xa = x \<and> sa = s \<or> xa \<noteq> x \<and> sa = (xa \<leftrightarrow> x) \<bullet> s \<and> atom xa \<sharp> s)" using assms Abs1_eq_iff[of xa sa x s] by simp
+ thus ?thesis using assms
+ by (metis flip_commute)
+qed
+
+lemma swap_v_supp:
+ fixes v::v and d::x and z::x
+ assumes "atom d \<sharp> v"
+ shows "supp ((z \<leftrightarrow> d ) \<bullet> v) \<subseteq> supp v - { atom z } \<union> { atom d}"
+ using assms
+proof(nominal_induct v rule:v.strong_induct)
+ case (V_lit l)
+ then show ?case using l.supp by (metis supp_l_empty empty_subsetI l.strong_exhaust pure_supp supp_eqvt v.supp)
+next
+ case (V_var x)
+ hence "d \<noteq> x" using fresh_def by fastforce
+ thus ?case apply(cases "z = x") using supp_at_base V_var \<open>d\<noteq>x\<close> by fastforce+
+next
+ case (V_cons tyid dc v)
+ show ?case using v.supp(4) pure_supp
+ using V_cons.hyps V_cons.prems fresh_def by auto
+next
+ case (V_consp tyid dc b v)
+ show ?case using v.supp(4) pure_supp
+ using V_consp.hyps V_consp.prems fresh_def by auto
+qed(force+)
+
+subsubsection \<open>Expressions\<close>
+
+lemma swap_e_supp:
+ fixes e::e and d::x and z::x
+ assumes "atom d \<sharp> e"
+ shows "supp ((z \<leftrightarrow> d ) \<bullet> e) \<subseteq> supp e - { atom z } \<union> { atom d}"
+ using assms
+proof(nominal_induct e rule:e.strong_induct)
+ case (AE_val v)
+ then show ?case using swap_v_supp by simp
+next
+ case (AE_app f v)
+ then show ?case using swap_v_supp by (simp add: pure_supp)
+next
+ case (AE_appP b f v)
+ hence df: "atom d \<sharp> v" using fresh_def e.supp by force
+ have "supp ((z \<leftrightarrow> d ) \<bullet> (AE_appP b f v)) = supp (AE_appP b f ((z \<leftrightarrow> d ) \<bullet> v))" using e.supp
+ by (metis b.eq_iff(3) b.perm_simps(3) e.perm_simps(3) flip_b_id)
+ also have "... = supp b \<union> supp f \<union> supp ((z \<leftrightarrow> d ) \<bullet> v)" using e.supp by auto
+ also have "... \<subseteq> supp b \<union> supp f \<union> supp v - { atom z } \<union> { atom d}" using swap_v_supp[OF df] pure_supp by auto
+ finally show ?case using e.supp by auto
+next
+ case (AE_op opp v1 v2)
+ hence df: "atom d \<sharp> v1 \<and> atom d \<sharp> v2" using fresh_def e.supp by force
+ have "((z \<leftrightarrow> d ) \<bullet> (AE_op opp v1 v2)) = AE_op opp ((z \<leftrightarrow> d ) \<bullet> v1) ((z \<leftrightarrow> d ) \<bullet> v2)" using
+ e.perm_simps flip_commute opp.perm_simps AE_op opp.strong_exhaust pure_supp
+ by (metis (full_types))
+
+ hence "supp ((z \<leftrightarrow> d) \<bullet> AE_op opp v1 v2) = supp (AE_op opp ((z \<leftrightarrow> d) \<bullet>v1) ((z \<leftrightarrow> d) \<bullet>v2))" by simp
+ also have "... = supp ((z \<leftrightarrow> d) \<bullet>v1) \<union> supp ((z \<leftrightarrow> d) \<bullet>v2)" using e.supp
+ by (metis (mono_tags, opaque_lifting) opp.strong_exhaust opp.supp sup_bot.left_neutral)
+ also have "... \<subseteq> (supp v1 - { atom z } \<union> { atom d}) \<union> (supp v2 - { atom z } \<union> { atom d})" using swap_v_supp AE_op df by blast
+ finally show ?case using e.supp opp.supp by blast
+next
+ case (AE_fst v)
+ then show ?case using swap_v_supp by auto
+next
+ case (AE_snd v)
+ then show ?case using swap_v_supp by auto
+next
+ case (AE_mvar u)
+ then show ?case using
+ Diff_empty Diff_insert0 Un_upper1 atom_x_sort flip_def flip_fresh_fresh fresh_def set_eq_subset supp_eqvt swap_set_in_eq
+ by (metis sort_of_atom_eq)
+next
+ case (AE_len v)
+ then show ?case using swap_v_supp by auto
+next
+ case (AE_concat v1 v2)
+ then show ?case using swap_v_supp by auto
+next
+ case (AE_split v1 v2)
+ then show ?case using swap_v_supp by auto
+qed
+
+lemma swap_ce_supp:
+ fixes e::ce and d::x and z::x
+ assumes "atom d \<sharp> e"
+ shows "supp ((z \<leftrightarrow> d ) \<bullet> e) \<subseteq> supp e - { atom z } \<union> { atom d}"
+ using assms
+proof(nominal_induct e rule:ce.strong_induct)
+ case (CE_val v)
+ then show ?case using swap_v_supp ce.fresh ce.supp by simp
+next
+ case (CE_op opp v1 v2)
+ hence df: "atom d \<sharp> v1 \<and> atom d \<sharp> v2" using fresh_def e.supp by force
+ have "((z \<leftrightarrow> d ) \<bullet> (CE_op opp v1 v2)) = CE_op opp ((z \<leftrightarrow> d ) \<bullet> v1) ((z \<leftrightarrow> d ) \<bullet> v2)" using
+ ce.perm_simps flip_commute opp.perm_simps CE_op opp.strong_exhaust x_fresh_b pure_supp
+ by (metis (full_types))
+
+ hence "supp ((z \<leftrightarrow> d) \<bullet> CE_op opp v1 v2) = supp (CE_op opp ((z \<leftrightarrow> d) \<bullet>v1) ((z \<leftrightarrow> d) \<bullet>v2))" by simp
+ also have "... = supp ((z \<leftrightarrow> d) \<bullet>v1) \<union> supp ((z \<leftrightarrow> d) \<bullet>v2)" using ce.supp
+ by (metis (mono_tags, opaque_lifting) opp.strong_exhaust opp.supp sup_bot.left_neutral)
+ also have "... \<subseteq> (supp v1 - { atom z } \<union> { atom d}) \<union> (supp v2 - { atom z } \<union> { atom d})" using swap_v_supp CE_op df by blast
+ finally show ?case using ce.supp opp.supp by blast
+next
+ case (CE_fst v)
+ then show ?case using ce.supp ce.fresh swap_v_supp by auto
+next
+ case (CE_snd v)
+ then show ?case using ce.supp ce.fresh swap_v_supp by auto
+next
+ case (CE_len v)
+ then show ?case using ce.supp ce.fresh swap_v_supp by auto
+next
+ case (CE_concat v1 v2)
+ then show ?case using ce.supp ce.fresh swap_v_supp ce.perm_simps
+ proof -
+ have "\<forall>x v xa. \<not> atom (x::x) \<sharp> (v::v) \<or> supp ((xa \<leftrightarrow> x) \<bullet> v) \<subseteq> supp v - {atom xa} \<union> {atom x}"
+ by (meson swap_v_supp) (* 0.0 ms *)
+ then show ?thesis
+ using CE_concat ce.supp by auto
+ qed
+qed
+
+lemma swap_c_supp:
+ fixes c::c and d::x and z::x
+ assumes "atom d \<sharp> c"
+ shows "supp ((z \<leftrightarrow> d ) \<bullet> c) \<subseteq> supp c - { atom z } \<union> { atom d}"
+ using assms
+proof(nominal_induct c rule:c.strong_induct)
+ case (C_eq e1 e2)
+ then show ?case using swap_ce_supp by auto
+qed(auto+)
+
+lemma type_e_eq:
+ assumes "atom z \<sharp> e" and "atom z' \<sharp> e"
+ shows "\<lbrace> z : b | [[z]\<^sup>v]\<^sup>c\<^sup>e == e \<rbrace> = (\<lbrace> z' : b | [[z']\<^sup>v]\<^sup>c\<^sup>e == e \<rbrace>)"
+ by (auto,metis (full_types) assms(1) assms(2) flip_fresh_fresh fresh_PairD(1) fresh_PairD(2))
+
+lemma type_e_eq2:
+ assumes "atom z \<sharp> e" and "atom z' \<sharp> e" and "b=b'"
+ shows "\<lbrace> z : b | [[z]\<^sup>v]\<^sup>c\<^sup>e == e \<rbrace> = (\<lbrace> z' : b' | [[z']\<^sup>v]\<^sup>c\<^sup>e == e \<rbrace>)"
+ using assms type_e_eq by fast
+
+lemma e_flip_eq:
+ fixes e::e and ea::e
+ assumes "atom c \<sharp> (e, ea)" and "atom c \<sharp> (x, xa, e, ea)" and "(x \<leftrightarrow> c) \<bullet> e = (xa \<leftrightarrow> c) \<bullet> ea"
+ shows "(e = AE_val w \<longrightarrow> (\<exists>w'. ea = AE_val w' \<and> (x \<leftrightarrow> c) \<bullet> w = (xa \<leftrightarrow> c) \<bullet> w')) \<or>
+ (e = AE_op opp v1 v2 \<longrightarrow> (\<exists>v1' v2'. ea = AS_op opp v1' v2' \<and> (x \<leftrightarrow> c) \<bullet> v1 = (xa \<leftrightarrow> c) \<bullet> v1') \<and> (x \<leftrightarrow> c) \<bullet> v2 = (xa \<leftrightarrow> c) \<bullet> v2') \<or>
+ (e = AE_fst v \<longrightarrow> (\<exists>v'. ea = AE_fst v' \<and> (x \<leftrightarrow> c) \<bullet> v = (xa \<leftrightarrow> c) \<bullet> v')) \<or>
+ (e = AE_snd v \<longrightarrow> (\<exists>v'. ea = AE_snd v' \<and> (x \<leftrightarrow> c) \<bullet> v = (xa \<leftrightarrow> c) \<bullet> v')) \<or>
+ (e = AE_len v \<longrightarrow> (\<exists>v'. ea = AE_len v' \<and> (x \<leftrightarrow> c) \<bullet> v = (xa \<leftrightarrow> c) \<bullet> v')) \<or>
+ (e = AE_concat v1 v2 \<longrightarrow> (\<exists>v1' v2'. ea = AS_concat v1' v2' \<and> (x \<leftrightarrow> c) \<bullet> v1 = (xa \<leftrightarrow> c) \<bullet> v1') \<and> (x \<leftrightarrow> c) \<bullet> v2 = (xa \<leftrightarrow> c) \<bullet> v2') \<or>
+ (e = AE_app f v \<longrightarrow> (\<exists>v'. ea = AE_app f v' \<and> (x \<leftrightarrow> c) \<bullet> v = (xa \<leftrightarrow> c) \<bullet> v'))"
+ by (metis assms e.perm_simps permute_flip_cancel2)
+
+lemma fresh_opp_all:
+ fixes opp::opp
+ shows "z \<sharp> opp"
+ using e.fresh opp.exhaust opp.fresh by metis
+
+lemma fresh_e_opp_all:
+ shows "(z \<sharp> v1 \<and> z \<sharp> v2) = z \<sharp> AE_op opp v1 v2"
+ using e.fresh opp.exhaust opp.fresh fresh_opp_all by simp
+
+lemma fresh_e_opp:
+ fixes z::x
+ assumes "atom z \<sharp> v1 \<and> atom z \<sharp> v2"
+ shows "atom z \<sharp> AE_op opp v1 v2"
+ using e.fresh opp.exhaust opp.fresh opp.supp by (metis assms)
+
+subsubsection \<open>Statements\<close>
+
+lemma branch_s_flip_eq:
+ fixes v::v and va::v
+ assumes "atom c \<sharp> (v, va)" and "atom c \<sharp> (x, xa, v, va)" and "(x \<leftrightarrow> c) \<bullet> s = (xa \<leftrightarrow> c) \<bullet> sa"
+ shows "(s = AS_val w \<longrightarrow> (\<exists>w'. sa = AS_val w' \<and> (x \<leftrightarrow> c) \<bullet> w = (xa \<leftrightarrow> c) \<bullet> w')) \<or>
+ (s = AS_seq s1 s2 \<longrightarrow> (\<exists>s1' s2'. sa = AS_seq s1' s2' \<and> (x \<leftrightarrow> c) \<bullet> s1 = (xa \<leftrightarrow> c) \<bullet> s1') \<and> (x \<leftrightarrow> c) \<bullet> s2 = (xa \<leftrightarrow> c) \<bullet> s2') \<or>
+ (s = AS_if v s1 s2 \<longrightarrow> (\<exists>v' s1' s2'. sa = AS_if seq s1' s2' \<and> (x \<leftrightarrow> c) \<bullet> s1 = (xa \<leftrightarrow> c) \<bullet> s1') \<and> (x \<leftrightarrow> c) \<bullet> s2 = (xa \<leftrightarrow> c) \<bullet> s2' \<and> (x \<leftrightarrow> c) \<bullet> c = (xa \<leftrightarrow> c) \<bullet> v')"
+ by (metis assms s_branch_s_branch_list.perm_simps permute_flip_cancel2)
+
+section \<open>Context Syntax\<close>
+
+subsection \<open>Datatypes\<close>
+
+text \<open>Type and function/type definition contexts\<close>
+type_synonym \<Phi> = "fun_def list"
+type_synonym \<Theta> = "type_def list"
+type_synonym \<B> = "bv fset"
+
+datatype \<Gamma> =
+ GNil
+ | GCons "x*b*c" \<Gamma> (infixr "#\<^sub>\<Gamma>" 65)
+
+datatype \<Delta> =
+ DNil ("[]\<^sub>\<Delta>")
+ | DCons "u*\<tau>" \<Delta> (infixr "#\<^sub>\<Delta>" 65)
+
+subsection \<open>Functions and Lemmas\<close>
+
+lemma \<Gamma>_induct [case_names GNil GCons] : "P GNil \<Longrightarrow> (\<And>x b c \<Gamma>'. P \<Gamma>' \<Longrightarrow> P ((x,b,c) #\<^sub>\<Gamma> \<Gamma>')) \<Longrightarrow> P \<Gamma>"
+proof(induct \<Gamma> rule:\<Gamma>.induct)
+ case GNil
+ then show ?case by auto
+next
+ case (GCons x1 x2)
+ then obtain x and b and c where "x1=(x,b,c)" using prod_cases3 by blast
+ then show ?case using GCons by presburger
+qed
+
+instantiation \<Delta> :: pt
+begin
+
+primrec permute_\<Delta>
+ where
+ DNil_eqvt: "p \<bullet> DNil = DNil"
+ | DCons_eqvt: "p \<bullet> (x #\<^sub>\<Delta> xs) = p \<bullet> x #\<^sub>\<Delta> p \<bullet> (xs::\<Delta>)"
+
+instance by standard (induct_tac [!] x, simp_all)
+end
+
+lemmas [eqvt] = permute_\<Delta>.simps
+
+lemma \<Delta>_induct [case_names DNil DCons] : "P DNil \<Longrightarrow> (\<And>u t \<Delta>'. P \<Delta>' \<Longrightarrow> P ((u,t) #\<^sub>\<Delta> \<Delta>')) \<Longrightarrow> P \<Delta>"
+proof(induct \<Delta> rule: \<Delta>.induct)
+ case DNil
+ then show ?case by auto
+next
+ case (DCons x1 x2)
+ then obtain u and t where "x1=(u,t)" by fastforce
+ then show ?case using DCons by presburger
+qed
+
+lemma \<Phi>_induct [case_names PNil PConsNone PConsSome] : "P [] \<Longrightarrow> (\<And>f x b c \<tau> s' \<Phi>'. P \<Phi>' \<Longrightarrow> P ((AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \<tau> s'))) # \<Phi>')) \<Longrightarrow>
+ (\<And>f bv x b c \<tau> s' \<Phi>'. P \<Phi>' \<Longrightarrow> P ((AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \<tau> s'))) # \<Phi>')) \<Longrightarrow> P \<Phi>"
+proof(induct \<Phi> rule: list.induct)
+ case Nil
+ then show ?case by auto
+next
+ case (Cons x1 x2)
+ then obtain f and t where ft: "x1 = (AF_fundef f t)"
+ by (meson fun_def.exhaust)
+ then show ?case proof(nominal_induct t rule:fun_typ_q.strong_induct)
+ case (AF_fun_typ_some bv ft)
+ then show ?case using Cons ft
+ by (metis fun_typ.exhaust)
+ next
+ case (AF_fun_typ_none ft)
+ then show ?case using Cons ft
+ by (metis fun_typ.exhaust)
+ qed
+qed
+
+lemma \<Theta>_induct [case_names TNil AF_typedef AF_typedef_poly] : "P [] \<Longrightarrow> (\<And>tid dclist \<Theta>'. P \<Theta>' \<Longrightarrow> P ((AF_typedef tid dclist) # \<Theta>')) \<Longrightarrow>
+ (\<And>tid bv dclist \<Theta>'. P \<Theta>' \<Longrightarrow> P ((AF_typedef_poly tid bv dclist ) # \<Theta>')) \<Longrightarrow> P \<Theta>"
+proof(induct \<Theta> rule: list.induct)
+ case Nil
+ then show ?case by auto
+next
+ case (Cons td T)
+ show ?case by(cases td rule: type_def.exhaust, (simp add: Cons)+)
+qed
+
+instantiation \<Gamma> :: pt
+begin
+
+primrec permute_\<Gamma>
+ where
+ GNil_eqvt: "p \<bullet> GNil = GNil"
+ | GCons_eqvt: "p \<bullet> (x #\<^sub>\<Gamma> xs) = p \<bullet> x #\<^sub>\<Gamma> p \<bullet> (xs::\<Gamma>)"
+
+instance by standard (induct_tac [!] x, simp_all)
+end
+
+lemmas [eqvt] = permute_\<Gamma>.simps
+
+lemma G_cons_eqvt[simp]:
+ fixes \<Gamma>::\<Gamma>
+ shows "p \<bullet> ((x,b,c) #\<^sub>\<Gamma> \<Gamma>) = ((p \<bullet> x, p \<bullet> b , p \<bullet> c) #\<^sub>\<Gamma> (p \<bullet> \<Gamma> ))" (is "?A = ?B" )
+ using Cons_eqvt triple_eqvt supp_b_empty by simp
+
+lemma G_cons_flip[simp]:
+ fixes x::x and \<Gamma>::\<Gamma>
+ shows "(x\<leftrightarrow>x') \<bullet> ((x'',b,c) #\<^sub>\<Gamma> \<Gamma>) = (((x\<leftrightarrow>x') \<bullet> x'', b , (x\<leftrightarrow>x') \<bullet> c) #\<^sub>\<Gamma> ((x\<leftrightarrow>x') \<bullet> \<Gamma> ))"
+ using Cons_eqvt triple_eqvt supp_b_empty by auto
+
+lemma G_cons_flip_fresh[simp]:
+ fixes x::x and \<Gamma>::\<Gamma>
+ assumes "atom x \<sharp> (c,\<Gamma>)" and "atom x' \<sharp> (c,\<Gamma>)"
+ shows "(x\<leftrightarrow>x') \<bullet> ((x',b,c) #\<^sub>\<Gamma> \<Gamma>) = ((x, b , c) #\<^sub>\<Gamma> \<Gamma> )"
+ using G_cons_flip flip_fresh_fresh assms by force
+
+lemma G_cons_flip_fresh2[simp]:
+ fixes x::x and \<Gamma>::\<Gamma>
+ assumes "atom x \<sharp> (c,\<Gamma>)" and "atom x' \<sharp> (c,\<Gamma>)"
+ shows "(x\<leftrightarrow>x') \<bullet> ((x,b,c) #\<^sub>\<Gamma> \<Gamma>) = ((x', b , c) #\<^sub>\<Gamma> \<Gamma> )"
+ using G_cons_flip flip_fresh_fresh assms by force
+
+lemma G_cons_flip_fresh3[simp]:
+ fixes x::x and \<Gamma>::\<Gamma>
+ assumes "atom x \<sharp> \<Gamma>" and "atom x' \<sharp> \<Gamma>"
+ shows "(x\<leftrightarrow>x') \<bullet> ((x',b,c) #\<^sub>\<Gamma> \<Gamma>) = ((x, b , (x\<leftrightarrow>x') \<bullet> c) #\<^sub>\<Gamma> \<Gamma> )"
+ using G_cons_flip flip_fresh_fresh assms by force
+
+lemma neq_GNil_conv: "(xs \<noteq> GNil) = (\<exists>y ys. xs = y #\<^sub>\<Gamma> ys)"
+ by (induct xs) auto
+
+nominal_function toList :: "\<Gamma> \<Rightarrow> (x*b*c) list" where
+ "toList GNil = []"
+| "toList (GCons xbc G) = xbc#(toList G)"
+ apply (auto, simp add: eqvt_def toList_graph_aux_def )
+ using neq_GNil_conv surj_pair by metis
+nominal_termination (eqvt)
+ by lexicographic_order
+
+nominal_function toSet :: "\<Gamma> \<Rightarrow> (x*b*c) set" where
+ "toSet GNil = {}"
+| "toSet (GCons xbc G) = {xbc} \<union> (toSet G)"
+ apply (auto,simp add: eqvt_def toSet_graph_aux_def )
+ using neq_GNil_conv surj_pair by metis
+nominal_termination (eqvt)
+ by lexicographic_order
+
+nominal_function append_g :: "\<Gamma> \<Rightarrow> \<Gamma> \<Rightarrow> \<Gamma>" (infixr "@" 65) where
+ "append_g GNil g = g"
+| "append_g (xbc #\<^sub>\<Gamma> g1) g2 = (xbc #\<^sub>\<Gamma> (g1@g2))"
+ apply (auto,simp add: eqvt_def append_g_graph_aux_def )
+ using neq_GNil_conv surj_pair by metis
+nominal_termination (eqvt) by lexicographic_order
+
+nominal_function dom :: "\<Gamma> \<Rightarrow> x set" where
+ "dom \<Gamma> = (fst` (toSet \<Gamma>))"
+ apply auto
+ unfolding eqvt_def dom_graph_aux_def lfp_eqvt toSet.eqvt by simp
+nominal_termination (eqvt) by lexicographic_order
+
+text \<open> Use of this is sometimes mixed in with use of freshness and support for the context however it makes it clear
+that for immutable variables, the context is `self-supporting'\<close>
+
+nominal_function atom_dom :: "\<Gamma> \<Rightarrow> atom set" where
+ "atom_dom \<Gamma> = atom`(dom \<Gamma>)"
+ apply auto
+ unfolding eqvt_def atom_dom_graph_aux_def lfp_eqvt toSet.eqvt by simp
+nominal_termination (eqvt) by lexicographic_order
+
+subsection \<open>Immutable Variable Context Lemmas\<close>
+
+lemma append_GNil[simp]:
+ "GNil @ G = G"
+ by simp
+
+lemma append_g_toSetU [simp]: "toSet (G1@G2) = toSet G1 \<union> toSet G2"
+ by(induct G1, auto+)
+
+lemma supp_GNil:
+ shows "supp GNil = {}"
+ by (simp add: supp_def)
+
+lemma supp_GCons:
+ fixes xs::\<Gamma>
+ shows "supp (x #\<^sub>\<Gamma> xs) = supp x \<union> supp xs"
+ by (simp add: supp_def Collect_imp_eq Collect_neg_eq)
+
+lemma atom_dom_eq[simp]:
+ fixes G::\<Gamma>
+ shows "atom_dom ((x, b, c) #\<^sub>\<Gamma> G) = atom_dom ((x, b, c') #\<^sub>\<Gamma> G)"
+ using atom_dom.simps toSet.simps by simp
+
+lemma dom_append[simp]:
+ "atom_dom (\<Gamma>@\<Gamma>') = atom_dom \<Gamma> \<union> atom_dom \<Gamma>'"
+ using image_Un append_g_toSetU atom_dom.simps dom.simps by metis
+
+lemma dom_cons[simp]:
+ "atom_dom ((x,b,c) #\<^sub>\<Gamma> G) = { atom x } \<union> atom_dom G"
+ using image_Un append_g_toSetU atom_dom.simps by auto
+
+lemma fresh_GNil[ms_fresh]:
+ shows "a \<sharp> GNil"
+ by (simp add: fresh_def supp_GNil)
+
+lemma fresh_GCons[ms_fresh]:
+ fixes xs::\<Gamma>
+ shows "a \<sharp> (x #\<^sub>\<Gamma> xs) \<longleftrightarrow> a \<sharp> x \<and> a \<sharp> xs"
+ by (simp add: fresh_def supp_GCons)
+
+lemma dom_supp_g[simp]:
+ "atom_dom G \<subseteq> supp G"
+ apply(induct G rule: \<Gamma>_induct,simp)
+ using supp_at_base supp_Pair atom_dom.simps supp_GCons by fastforce
+
+lemma fresh_append_g[ms_fresh]:
+ fixes xs::\<Gamma>
+ shows "a \<sharp> (xs @ ys) \<longleftrightarrow> a \<sharp> xs \<and> a \<sharp> ys"
+ by (induct xs) (simp_all add: fresh_GNil fresh_GCons)
+
+lemma append_g_assoc:
+ fixes xs::\<Gamma>
+ shows "(xs @ ys) @ zs = xs @ (ys @ zs)"
+ by (induct xs) simp_all
+
+lemma append_g_inside:
+ fixes xs::\<Gamma>
+ shows "xs @ (x #\<^sub>\<Gamma> ys) = (xs @ (x #\<^sub>\<Gamma> GNil)) @ ys"
+ by(induct xs,auto+)
+
+lemma finite_\<Gamma>:
+ "finite (toSet \<Gamma>)"
+ by(induct \<Gamma> rule: \<Gamma>_induct,auto)
+
+lemma supp_\<Gamma>:
+ "supp \<Gamma> = supp (toSet \<Gamma>)"
+proof(induct \<Gamma> rule: \<Gamma>_induct)
+ case GNil
+ then show ?case using supp_GNil toSet.simps
+ by (simp add: supp_set_empty)
+next
+ case (GCons x b c \<Gamma>')
+ then show ?case using supp_GCons toSet.simps finite_\<Gamma> supp_of_finite_union
+ using supp_of_finite_insert by fastforce
+qed
+
+lemma supp_of_subset:
+ fixes G::"('a::fs set)"
+ assumes "finite G" and "finite G'" and "G \<subseteq> G'"
+ shows "supp G \<subseteq> supp G'"
+ using supp_of_finite_sets assms by (metis subset_Un_eq supp_of_finite_union)
+
+lemma supp_weakening:
+ assumes "toSet G \<subseteq> toSet G'"
+ shows "supp G \<subseteq> supp G'"
+ using supp_\<Gamma> finite_\<Gamma> by (simp add: supp_of_subset assms)
+
+lemma fresh_weakening[ms_fresh]:
+ assumes "toSet G \<subseteq> toSet G'" and "x \<sharp> G'"
+ shows "x \<sharp> G"
+proof(rule ccontr)
+ assume "\<not> x \<sharp> G"
+ hence "x \<in> supp G" using fresh_def by auto
+ hence "x \<in> supp G'" using supp_weakening assms by auto
+ thus False using fresh_def assms by auto
+qed
+
+instance \<Gamma> :: fs
+ by (standard, induct_tac x, simp_all add: supp_GNil supp_GCons finite_supp)
+
+lemma fresh_gamma_elem:
+ fixes \<Gamma>::\<Gamma>
+ assumes "a \<sharp> \<Gamma>"
+ and "e \<in> toSet \<Gamma>"
+ shows "a \<sharp> e"
+ using assms by(induct \<Gamma>,auto simp add: fresh_GCons)
+
+lemma fresh_gamma_append:
+ fixes xs::\<Gamma>
+ shows "a \<sharp> (xs @ ys) \<longleftrightarrow> a \<sharp> xs \<and> a \<sharp> ys"
+ by (induct xs, simp_all add: fresh_GNil fresh_GCons)
+
+lemma supp_triple[simp]:
+ shows "supp (x, y, z) = supp x \<union> supp y \<union> supp z"
+proof -
+ have "supp (x,y,z) = supp (x,(y,z))" by auto
+ hence "supp (x,y,z) = supp x \<union> (supp y \<union> supp z)" using supp_Pair by metis
+ thus ?thesis by auto
+qed
+
+lemma supp_append_g:
+ fixes xs::\<Gamma>
+ shows "supp (xs @ ys) = supp xs \<union> supp ys"
+ by(induct xs,auto simp add: supp_GNil supp_GCons )
+
+lemma fresh_in_g[simp]:
+ fixes \<Gamma>::\<Gamma> and x'::x
+ shows "atom x' \<sharp> \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> = (atom x' \<notin> supp \<Gamma>' \<union> supp x \<union> supp b0 \<union> supp c0 \<union> supp \<Gamma>)"
+proof -
+ have "atom x' \<sharp> \<Gamma>' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> = (atom x' \<notin> supp (\<Gamma>' @((x,b0,c0) #\<^sub>\<Gamma> \<Gamma>)))"
+ using fresh_def by auto
+ also have "... = (atom x' \<notin> supp \<Gamma>' \<union> supp ((x,b0,c0) #\<^sub>\<Gamma> \<Gamma>))" using supp_append_g by fast
+ also have "... = (atom x' \<notin> supp \<Gamma>' \<union> supp x \<union> supp b0 \<union> supp c0 \<union> supp \<Gamma>)" using supp_GCons supp_append_g supp_triple by auto
+ finally show ?thesis by fast
+qed
+
+lemma fresh_suffix[ms_fresh]:
+ fixes \<Gamma>::\<Gamma>
+ assumes "atom x \<sharp> \<Gamma>'@\<Gamma>"
+ shows "atom x \<sharp> \<Gamma>"
+ using assms by(induct \<Gamma>' rule: \<Gamma>_induct, auto simp add: append_g.simps fresh_GCons)
+
+lemma not_GCons_self [simp]:
+ fixes xs::\<Gamma>
+ shows "xs \<noteq> x #\<^sub>\<Gamma> xs"
+ by (induct xs) auto
+
+lemma not_GCons_self2 [simp]:
+ fixes xs::\<Gamma>
+ shows "x #\<^sub>\<Gamma> xs \<noteq> xs"
+ by (rule not_GCons_self [symmetric])
+
+lemma fresh_restrict:
+ fixes y::x and \<Gamma>::\<Gamma>
+ assumes "atom y \<sharp> (\<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>)"
+ shows "atom y \<sharp> (\<Gamma>'@\<Gamma>)"
+ using assms by(induct \<Gamma>' rule: \<Gamma>_induct, auto simp add:fresh_GCons fresh_GNil )
+
+lemma fresh_dom_free:
+ assumes "atom x \<sharp> \<Gamma>"
+ shows "(x,b,c) \<notin> toSet \<Gamma>"
+ using assms proof(induct \<Gamma> rule: \<Gamma>_induct)
+ case GNil
+ then show ?case by auto
+next
+ case (GCons x' b' c' \<Gamma>')
+ hence "x\<noteq>x'" using fresh_def fresh_GCons fresh_Pair supp_at_base by blast
+ moreover have "atom x \<sharp> \<Gamma>'" using fresh_GCons GCons by auto
+ ultimately show ?case using toSet.simps GCons by auto
+qed
+
+lemma \<Gamma>_set_intros: "x \<in> toSet ( x #\<^sub>\<Gamma> xs)" and "y \<in> toSet xs \<Longrightarrow> y \<in> toSet (x #\<^sub>\<Gamma> xs)"
+ by simp+
+
+lemma fresh_dom_free2:
+ assumes "atom x \<notin> atom_dom \<Gamma>"
+ shows "(x,b,c) \<notin> toSet \<Gamma>"
+ using assms proof(induct \<Gamma> rule: \<Gamma>_induct)
+ case GNil
+ then show ?case by auto
+next
+ case (GCons x' b' c' \<Gamma>')
+ hence "x\<noteq>x'" using fresh_def fresh_GCons fresh_Pair supp_at_base by auto
+ moreover have "atom x \<notin> atom_dom \<Gamma>'" using fresh_GCons GCons by auto
+ ultimately show ?case using toSet.simps GCons by auto
+qed
+
+subsection \<open>Mutable Variable Context Lemmas\<close>
+
+lemma supp_DNil:
+ shows "supp DNil = {}"
+ by (simp add: supp_def)
+
+lemma supp_DCons:
+ fixes xs::\<Delta>
+ shows "supp (x #\<^sub>\<Delta> xs) = supp x \<union> supp xs"
+ by (simp add: supp_def Collect_imp_eq Collect_neg_eq)
+
+lemma fresh_DNil[ms_fresh]:
+ shows "a \<sharp> DNil"
+ by (simp add: fresh_def supp_DNil)
+
+lemma fresh_DCons[ms_fresh]:
+ fixes xs::\<Delta>
+ shows "a \<sharp> (x #\<^sub>\<Delta> xs) \<longleftrightarrow> a \<sharp> x \<and> a \<sharp> xs"
+ by (simp add: fresh_def supp_DCons)
+
+instance \<Delta> :: fs
+ by (standard, induct_tac x, simp_all add: supp_DNil supp_DCons finite_supp)
+
+subsection \<open>Lookup Functions\<close>
+
+nominal_function lookup :: "\<Gamma> \<Rightarrow> x \<Rightarrow> (b*c) option" where
+ "lookup GNil x = None"
+| "lookup ((x,b,c)#\<^sub>\<Gamma>G) y = (if x=y then Some (b,c) else lookup G y)"
+ by (auto,simp add: eqvt_def lookup_graph_aux_def, metis neq_GNil_conv surj_pair)
+nominal_termination (eqvt) by lexicographic_order
+
+nominal_function replace_in_g :: "\<Gamma> \<Rightarrow> x \<Rightarrow> c \<Rightarrow> \<Gamma>" ("_[_\<longmapsto>_]" [1000,0,0] 200) where
+ "replace_in_g GNil _ _ = GNil"
+| "replace_in_g ((x,b,c)#\<^sub>\<Gamma>G) x' c' = (if x=x' then ((x,b,c')#\<^sub>\<Gamma>G) else (x,b,c)#\<^sub>\<Gamma>(replace_in_g G x' c'))"
+ apply(auto,simp add: eqvt_def replace_in_g_graph_aux_def)
+ using surj_pair \<Gamma>.exhaust by metis
+nominal_termination (eqvt) by lexicographic_order
+
+text \<open> Functions for looking up data-constructors in the Pi context \<close>
+
+nominal_function lookup_fun :: "\<Phi> \<Rightarrow> f \<Rightarrow> fun_def option" where
+ "lookup_fun [] g = None"
+| "lookup_fun ((AF_fundef f ft)#\<Pi>) g = (if (f=g) then Some (AF_fundef f ft) else lookup_fun \<Pi> g)"
+ apply(auto,simp add: eqvt_def lookup_fun_graph_aux_def )
+ by (metis fun_def.exhaust neq_Nil_conv)
+nominal_termination (eqvt) by lexicographic_order
+
+nominal_function lookup_td :: "\<Theta> \<Rightarrow> string \<Rightarrow> type_def option" where
+ "lookup_td [] g = None"
+| "lookup_td ((AF_typedef s lst ) # (\<Theta>::\<Theta>)) g = (if (s = g) then Some (AF_typedef s lst ) else lookup_td \<Theta> g)"
+| "lookup_td ((AF_typedef_poly s bv lst ) # (\<Theta>::\<Theta>)) g = (if (s = g) then Some (AF_typedef_poly s bv lst ) else lookup_td \<Theta> g)"
+ apply(auto,simp add: eqvt_def lookup_td_graph_aux_def )
+ by (metis type_def.exhaust neq_Nil_conv)
+nominal_termination (eqvt) by lexicographic_order
+
+nominal_function name_of_type ::"type_def \<Rightarrow> f " where
+ "name_of_type (AF_typedef f _ ) = f"
+| "name_of_type (AF_typedef_poly f _ _) = f"
+ apply(auto,simp add: eqvt_def name_of_type_graph_aux_def )
+ using type_def.exhaust by blast
+nominal_termination (eqvt) by lexicographic_order
+
+nominal_function name_of_fun ::"fun_def \<Rightarrow> f " where
+ "name_of_fun (AF_fundef f ft) = f"
+ apply(auto,simp add: eqvt_def name_of_fun_graph_aux_def )
+ using fun_def.exhaust by blast
+nominal_termination (eqvt) by lexicographic_order
+
+nominal_function remove2 :: "'a::pt \<Rightarrow> 'a list \<Rightarrow> 'a list" where
+ "remove2 x [] = []" |
+ "remove2 x (y # xs) = (if x = y then xs else y # remove2 x xs)"
+ by (simp add: eqvt_def remove2_graph_aux_def,auto+,meson list.exhaust)
+nominal_termination (eqvt) by lexicographic_order
+
+nominal_function base_for_lit :: "l \<Rightarrow> b" where
+ "base_for_lit (L_true) = B_bool "
+| "base_for_lit (L_false) = B_bool "
+| "base_for_lit (L_num n) = B_int "
+| "base_for_lit (L_unit) = B_unit "
+| "base_for_lit (L_bitvec v) = B_bitvec"
+ apply (auto simp: eqvt_def base_for_lit_graph_aux_def )
+ using l.strong_exhaust by blast
+nominal_termination (eqvt) by lexicographic_order
+
+lemma neq_DNil_conv: "(xs \<noteq> DNil) = (\<exists>y ys. xs = y #\<^sub>\<Delta> ys)"
+ by (induct xs) auto
+
+nominal_function setD :: "\<Delta> \<Rightarrow> (u*\<tau>) set" where
+ "setD DNil = {}"
+| "setD (DCons xbc G) = {xbc} \<union> (setD G)"
+ apply (auto,simp add: eqvt_def setD_graph_aux_def )
+ using neq_DNil_conv surj_pair by metis
+nominal_termination (eqvt) by lexicographic_order
+
+lemma eqvt_triple:
+ fixes y::"'a::at" and ya::"'a::at" and xa::"'c::at" and va::"'d::fs" and s::s and sa::s and f::"s*'c*'d \<Rightarrow> s"
+ assumes "atom y \<sharp> (xa, va)" and "atom ya \<sharp> (xa, va)" and
+ "\<forall>c. atom c \<sharp> (s, sa) \<longrightarrow> atom c \<sharp> (y, ya, s, sa) \<longrightarrow> (y \<leftrightarrow> c) \<bullet> s = (ya \<leftrightarrow> c) \<bullet> sa"
+ and "eqvt_at f (s,xa,va)" and "eqvt_at f (sa,xa,va)" and
+ "atom c \<sharp> (s, va, xa, sa)" and "atom c \<sharp> (y, ya, f (s, xa, va), f (sa, xa, va))"
+ shows "(y \<leftrightarrow> c) \<bullet> f (s, xa, va) = (ya \<leftrightarrow> c) \<bullet> f (sa, xa, va)"
+proof -
+ have " (y \<leftrightarrow> c) \<bullet> f (s, xa, va) = f ( (y \<leftrightarrow> c) \<bullet> (s,xa,va))" using assms eqvt_at_def by metis
+ also have "... = f ( (y \<leftrightarrow> c) \<bullet> s, (y \<leftrightarrow> c) \<bullet> xa ,(y \<leftrightarrow> c) \<bullet> va)" by auto
+ also have "... = f ( (ya \<leftrightarrow> c) \<bullet> sa, (ya \<leftrightarrow> c) \<bullet> xa ,(ya \<leftrightarrow> c) \<bullet> va)" proof -
+ have " (y \<leftrightarrow> c) \<bullet> s = (ya \<leftrightarrow> c) \<bullet> sa" using assms Abs1_eq_iff_all by auto
+ moreover have "((y \<leftrightarrow> c) \<bullet> xa) = ((ya \<leftrightarrow> c) \<bullet> xa)" using assms flip_fresh_fresh fresh_prodN by metis
+ moreover have "((y \<leftrightarrow> c) \<bullet> va) = ((ya \<leftrightarrow> c) \<bullet> va)" using assms flip_fresh_fresh fresh_prodN by metis
+ ultimately show ?thesis by auto
+ qed
+ also have "... = f ( (ya \<leftrightarrow> c) \<bullet> (sa,xa,va))" by auto
+ finally show ?thesis using assms eqvt_at_def by metis
+qed
+
+section \<open>Functions for bit list/vectors\<close>
+
+inductive split :: "int \<Rightarrow> bit list \<Rightarrow> bit list * bit list \<Rightarrow> bool" where
+ "split 0 xs ([], xs)"
+| "split m xs (ys,zs) \<Longrightarrow> split (m+1) (x#xs) ((x # ys), zs)"
+equivariance split
+nominal_inductive split .
+
+lemma split_concat:
+ assumes "split n v (v1,v2)"
+ shows "v = append v1 v2"
+ using assms proof(induct "(v1,v2)" arbitrary: v1 v2 rule: split.inducts)
+ case 1
+ then show ?case by auto
+next
+ case (2 m xs ys zs x)
+ then show ?case by auto
+qed
+
+lemma split_n:
+ assumes "split n v (v1,v2)"
+ shows "0 \<le> n \<and> n \<le> int (length v)"
+ using assms proof(induct rule: split.inducts)
+ case (1 xs)
+ then show ?case by auto
+next
+ case (2 m xs ys zs x)
+ then show ?case by auto
+qed
+
+lemma split_length:
+ assumes "split n v (v1,v2)"
+ shows "n = int (length v1)"
+ using assms proof(induct "(v1,v2)" arbitrary: v1 v2 rule: split.inducts)
+ case (1 xs)
+ then show ?case by auto
+next
+ case (2 m xs ys zs x)
+ then show ?case by auto
+qed
+
+lemma obtain_split:
+ assumes "0 \<le> n" and "n \<le> int (length bv)"
+ shows "\<exists> bv1 bv2. split n bv (bv1 , bv2)"
+ using assms proof(induct bv arbitrary: n)
+ case Nil
+ then show ?case using split.intros by auto
+next
+ case (Cons b bv)
+ show ?case proof(cases "n = 0")
+ case True
+ then show ?thesis using split.intros by auto
+ next
+ case False
+ then obtain m where m:"n=m+1" using Cons
+ by (metis add.commute add_minus_cancel)
+ moreover have "0 \<le> m" using False m Cons by linarith
+ then obtain bv1 and bv2 where "split m bv (bv1 , bv2)" using Cons m by force
+ hence "split n (b # bv) ((b#bv1), bv2)" using m split.intros by auto
+ then show ?thesis by auto
+ qed
+qed
+
end
\ No newline at end of file
diff --git a/thys/MiniSail/SyntaxL.thy b/thys/MiniSail/SyntaxL.thy
--- a/thys/MiniSail/SyntaxL.thy
+++ b/thys/MiniSail/SyntaxL.thy
@@ -1,383 +1,383 @@
-(*<*)
-theory SyntaxL
- imports Syntax IVSubst
-begin
- (*>*)
-
-chapter \<open>Syntax Lemmas\<close>
-
-section \<open>Support, lookup and contexts\<close>
-
-lemma supp_v_tau [simp]:
- assumes "atom z \<sharp> v"
- shows "supp (\<lbrace> z : b | CE_val (V_var z) == CE_val v \<rbrace>) = supp v \<union> supp b"
- using assms \<tau>.supp c.supp ce.supp
- by (simp add: fresh_def supp_at_base)
-
-lemma supp_v_var_tau [simp]:
- assumes "z \<noteq> x"
- shows "supp (\<lbrace> z : b | CE_val (V_var z) == CE_val (V_var x) \<rbrace>) = { atom x } \<union> supp b"
- using supp_v_tau assms
- using supp_at_base by fastforce
-
-text \<open> Sometimes we need to work with a version of a binder where the variable is fresh
-in something else, such as a bigger context. I think these could be generated automatically \<close>
-
-lemma obtain_fresh_fun_def:
- fixes t::"'b::fs"
- shows "\<exists>y::x. atom y \<sharp> (s,c,\<tau>,t) \<and> (AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \<tau> s)) = AF_fundef f (AF_fun_typ_none (AF_fun_typ y b ((y \<leftrightarrow> x) \<bullet> c ) ((y \<leftrightarrow> x) \<bullet> \<tau>) ((y \<leftrightarrow> x) \<bullet> s))))"
-proof -
- obtain y::x where y: "atom y \<sharp> (s,c,\<tau>,t)" using obtain_fresh by blast
- moreover have "AF_fundef f (AF_fun_typ_none (AF_fun_typ y b ((y \<leftrightarrow> x) \<bullet> c ) ((y \<leftrightarrow> x) \<bullet> \<tau>) ((y \<leftrightarrow> x) \<bullet> s))) = (AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \<tau> s)))"
- proof(cases "x=y")
- case True
- then show ?thesis using fun_def.eq_iff Abs1_eq_iff(3) flip_commute flip_fresh_fresh fresh_PairD by auto
- next
- case False
-
- have "(AF_fun_typ y b ((y \<leftrightarrow> x) \<bullet> c) ((y \<leftrightarrow> x) \<bullet> \<tau>) ((y \<leftrightarrow> x) \<bullet> s)) =(AF_fun_typ x b c \<tau> s)" proof(subst fun_typ.eq_iff, subst Abs1_eq_iff(3))
- show \<open>(y = x \<and> (((y \<leftrightarrow> x) \<bullet> c, (y \<leftrightarrow> x) \<bullet> \<tau>), (y \<leftrightarrow> x) \<bullet> s) = ((c, \<tau>), s) \<or>
- y \<noteq> x \<and> (((y \<leftrightarrow> x) \<bullet> c, (y \<leftrightarrow> x) \<bullet> \<tau>), (y \<leftrightarrow> x) \<bullet> s) = (y \<leftrightarrow> x) \<bullet> ((c, \<tau>), s) \<and> atom y \<sharp> ((c, \<tau>), s)) \<and>
- b = b\<close> using False flip_commute flip_fresh_fresh fresh_PairD y by auto
- qed
- thus ?thesis by metis
- qed
- ultimately show ?thesis using y fresh_Pair by metis
-qed
-
-lemma lookup_fun_member:
- assumes "Some (AF_fundef f ft) = lookup_fun \<Phi> f"
- shows "AF_fundef f ft \<in> set \<Phi>"
- using assms proof (induct \<Phi>)
- case Nil
- then show ?case by auto
-next
- case (Cons a \<Phi>)
- then show ?case using lookup_fun.simps
- by (metis fun_def.exhaust insert_iff list.simps(15) option.inject)
-qed
-
-lemma rig_dom_eq:
- "dom (G[x \<longmapsto> c]) = dom G"
-proof(induct G rule: \<Gamma>.induct)
- case GNil
- then show ?case using replace_in_g.simps by presburger
-next
- case (GCons xbc \<Gamma>')
- obtain x' and b' and c' where xbc: "xbc=(x',b',c')" using prod_cases3 by blast
- then show ?case using replace_in_g.simps GCons by simp
-qed
-
-lemma lookup_in_rig_eq:
- assumes "Some (b,c) = lookup \<Gamma> x"
- shows "Some (b,c') = lookup (\<Gamma>[x\<longmapsto>c']) x"
- using assms proof(induct \<Gamma> rule: \<Gamma>_induct)
- case GNil
- then show ?case by auto
-next
- case (GCons x b c \<Gamma>')
- then show ?case using replace_in_g.simps lookup.simps by auto
-qed
-
-lemma lookup_in_rig_neq:
- assumes "Some (b,c) = lookup \<Gamma> y" and "x\<noteq>y"
- shows "Some (b,c) = lookup (\<Gamma>[x\<longmapsto>c']) y"
- using assms proof(induct \<Gamma> rule: \<Gamma>_induct)
- case GNil
- then show ?case by auto
-next
- case (GCons x' b' c' \<Gamma>')
- then show ?case using replace_in_g.simps lookup.simps by auto
-qed
-
-lemma lookup_in_rig:
- assumes "Some (b,c) = lookup \<Gamma> y"
- shows "\<exists>c''. Some (b,c'') = lookup (\<Gamma>[x\<longmapsto>c']) y"
-proof(cases "x=y")
- case True
- then show ?thesis using lookup_in_rig_eq using assms by blast
-next
- case False
- then show ?thesis using lookup_in_rig_neq using assms by blast
-qed
-
-lemma lookup_inside[simp]:
- assumes "x \<notin> fst ` toSet \<Gamma>'"
- shows "Some (b1,c1) = lookup (\<Gamma>'@(x,b1,c1)#\<^sub>\<Gamma>\<Gamma>) x"
- using assms by(induct \<Gamma>',auto)
-
-lemma lookup_inside2:
- assumes "Some (b1,c1) = lookup (\<Gamma>'@((x,b0,c0)#\<^sub>\<Gamma>\<Gamma>)) y" and "x\<noteq>y"
- shows "Some (b1,c1) = lookup (\<Gamma>'@((x,b0,c0')#\<^sub>\<Gamma>\<Gamma>)) y"
- using assms by(induct \<Gamma>' rule: \<Gamma>.induct,auto+)
-
-fun tail:: "'a list \<Rightarrow> 'a list" where
- "tail [] = []"
-| "tail (x#xs) = xs"
-
-lemma lookup_options:
- assumes "Some (b,c) = lookup (xt#\<^sub>\<Gamma>G) x"
- shows "((x,b,c) = xt) \<or> (Some (b,c) = lookup G x)"
- by (metis assms lookup.simps(2) option.inject surj_pair)
-
-lemma lookup_x:
- assumes "Some (b,c) = lookup G x"
- shows "x \<in> fst ` toSet G"
- using assms
- by(induct "G" rule: \<Gamma>.induct ,auto+)
-
-lemma GCons_eq_appendI:
- fixes xs1::\<Gamma>
- shows "[| x #\<^sub>\<Gamma> xs1 = ys; xs = xs1 @ zs |] ==> x #\<^sub>\<Gamma> xs = ys @ zs"
- by (drule sym) simp
-
-lemma split_G: "x : toSet xs \<Longrightarrow> \<exists>ys zs. xs = ys @ x #\<^sub>\<Gamma> zs"
-proof (induct xs)
- case GNil thus ?case by simp
-next
- case GCons thus ?case using GCons_eq_appendI
- by (metis Un_iff append_g.simps(1) singletonD toSet.simps(2))
-qed
-
-lemma lookup_not_empty:
- assumes "Some \<tau> = lookup G x"
- shows "G \<noteq> GNil"
- using assms by auto
-
-lemma lookup_in_g:
- assumes "Some (b,c) = lookup \<Gamma> x"
- shows "(x,b,c) \<in> toSet \<Gamma>"
- using assms apply(induct \<Gamma>, simp)
- using lookup_options by fastforce
-
-lemma lookup_split:
- fixes \<Gamma>::\<Gamma>
- assumes "Some (b,c) = lookup \<Gamma> x"
- shows "\<exists>G G' . \<Gamma> = G'@(x,b,c)#\<^sub>\<Gamma>G"
- by (meson assms(1) lookup_in_g split_G)
-
-lemma toSet_splitU[simp]:
- "(x',b',c') \<in> toSet (\<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>) \<longleftrightarrow> (x',b',c') \<in> (toSet \<Gamma>' \<union> {(x, b, c)} \<union> toSet \<Gamma>)"
- using append_g_toSetU toSet.simps by auto
-
-lemma toSet_splitP[simp]:
- "(\<forall>(x', b', c')\<in>toSet (\<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>). P x' b' c') \<longleftrightarrow> (\<forall> (x', b', c')\<in>toSet \<Gamma>'. P x' b' c') \<and> P x b c \<and> (\<forall> (x', b', c')\<in>toSet \<Gamma>. P x' b' c')" (is "?A \<longleftrightarrow> ?B")
- using toSet_splitU by force
-
-lemma lookup_restrict:
- assumes "Some (b',c') = lookup (\<Gamma>'@(x,b,c)#\<^sub>\<Gamma>\<Gamma>) y" and "x \<noteq> y"
- shows "Some (b',c') = lookup (\<Gamma>'@\<Gamma>) y"
- using assms proof(induct \<Gamma>' rule:\<Gamma>_induct)
- case GNil
- then show ?case by auto
-next
- case (GCons x1 b1 c1 \<Gamma>')
- then show ?case by auto
-qed
-
-lemma supp_list_member:
- fixes x::"'a::fs" and l::"'a list"
- assumes "x \<in> set l"
- shows "supp x \<subseteq> supp l"
- using assms apply(induct l, auto)
- using supp_Cons by auto
-
-lemma GNil_append:
- assumes "GNil = G1@G2"
- shows "G1 = GNil \<and> G2 = GNil"
-proof(rule ccontr)
- assume "\<not> (G1 = GNil \<and> G2 = GNil)"
- hence "G1@G2 \<noteq> GNil" using append_g.simps by (metis \<Gamma>.distinct(1) \<Gamma>.exhaust)
- thus False using assms by auto
-qed
-
-lemma GCons_eq_append_conv:
- fixes xs::\<Gamma>
- shows "x#\<^sub>\<Gamma>xs = ys@zs = (ys = GNil \<and> x#\<^sub>\<Gamma>xs = zs \<or> (\<exists>ys'. x#\<^sub>\<Gamma>ys' = ys \<and> xs = ys'@zs))"
- by(cases ys) auto
-
-lemma dclist_distinct_unique:
- assumes "(dc, const) \<in> set dclist2" and "(cons, const1) \<in> set dclist2" and "dc=cons" and "distinct (List.map fst dclist2)"
- shows "(const) = const1"
-proof -
- have "(cons, const) = (dc, const1)"
- using assms by (metis (no_types, lifting) assms(3) assms(4) distinct.simps(1) distinct.simps(2) empty_iff insert_iff list.set(1) list.simps(15) list.simps(8) list.simps(9) map_of_eq_Some_iff)
- thus ?thesis by auto
-qed
-
-lemma fresh_d_fst_d:
- assumes "atom u \<sharp> \<delta>"
- shows "u \<notin> fst ` set \<delta>"
- using assms proof(induct \<delta>)
- case Nil
- then show ?case by auto
-next
- case (Cons ut \<delta>')
- obtain u' and t' where *:"ut = (u',t') " by fastforce
- hence "atom u \<sharp> ut \<and> atom u \<sharp> \<delta>'" using fresh_Cons Cons by auto
- moreover hence "atom u \<sharp> fst ut" using * fresh_Pair[of "atom u" u' t'] Cons by auto
- ultimately show ?case using Cons by auto
-qed
-
-lemma bv_not_in_bset_supp:
- fixes bv::bv
- assumes "bv |\<notin>| B"
- shows "atom bv \<notin> supp B"
-proof -
- have *:"supp B = fset (fimage atom B)"
- by (metis fimage.rep_eq finite_fset supp_finite_set_at_base supp_fset)
- thus ?thesis using assms
- using notin_fset by fastforce
-qed
-
-lemma u_fresh_d:
- assumes "atom u \<sharp> D"
- shows "u \<notin> fst ` setD D"
- using assms proof(induct D rule: \<Delta>_induct)
- case DNil
- then show ?case by auto
-next
- case (DCons u' t' \<Delta>')
- then show ?case unfolding setD.simps
- using fresh_DCons fresh_Pair by (simp add: fresh_Pair fresh_at_base(2))
-qed
-
-section \<open>Type Definitions\<close>
-
-lemma exist_fresh_bv:
- fixes tm::"'a::fs"
- shows "\<exists>bva2 dclist2. AF_typedef_poly tyid bva dclist = AF_typedef_poly tyid bva2 dclist2 \<and>
- atom bva2 \<sharp> tm"
-proof -
- obtain bva2::bv where *:"atom bva2 \<sharp> (bva, dclist,tyid,tm)" using obtain_fresh by metis
- moreover hence "bva2 \<noteq> bva" using fresh_at_base by auto
- moreover have " dclist = (bva \<leftrightarrow> bva2) \<bullet> (bva2 \<leftrightarrow> bva) \<bullet> dclist" by simp
- moreover have "atom bva \<sharp> (bva2 \<leftrightarrow> bva) \<bullet> dclist" proof -
- have "atom bva2 \<sharp> dclist" using * fresh_prodN by auto
- hence "atom ((bva2 \<leftrightarrow> bva) \<bullet> bva2) \<sharp> (bva2 \<leftrightarrow> bva) \<bullet> dclist" using fresh_eqvt True_eqvt
- proof -
- have "(bva2 \<leftrightarrow> bva) \<bullet> atom bva2 \<sharp> (bva2 \<leftrightarrow> bva) \<bullet> dclist"
- by (metis True_eqvt \<open>atom bva2 \<sharp> dclist\<close> fresh_eqvt) (* 62 ms *)
- then show ?thesis
- by simp (* 125 ms *)
- qed
- thus ?thesis by auto
- qed
- ultimately have "AF_typedef_poly tyid bva dclist = AF_typedef_poly tyid bva2 ((bva2 \<leftrightarrow> bva ) \<bullet> dclist)"
- unfolding type_def.eq_iff Abs1_eq_iff by metis
- thus ?thesis using * fresh_prodN by metis
-qed
-
-lemma obtain_fresh_bv:
- fixes tm::"'a::fs"
- obtains bva2::bv and dclist2 where "AF_typedef_poly tyid bva dclist = AF_typedef_poly tyid bva2 dclist2 \<and>
- atom bva2 \<sharp> tm"
- using exist_fresh_bv by metis
-
-section \<open>Function Definitions\<close>
-
-lemma fun_typ_flip:
- fixes bv1::bv and c::bv
- shows "(bv1 \<leftrightarrow> c) \<bullet> AF_fun_typ x1 b1 c1 \<tau>1 s1 = AF_fun_typ x1 ((bv1 \<leftrightarrow> c) \<bullet> b1) ((bv1 \<leftrightarrow> c) \<bullet> c1) ((bv1 \<leftrightarrow> c) \<bullet> \<tau>1) ((bv1 \<leftrightarrow> c) \<bullet> s1)"
- using fun_typ.perm_simps flip_fresh_fresh supp_at_base fresh_def
- flip_fresh_fresh fresh_def supp_at_base
- by (simp add: flip_fresh_fresh)
-
-lemma fun_def_eq:
- assumes "AF_fundef fa (AF_fun_typ_none (AF_fun_typ xa ba ca \<tau>a sa)) = AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \<tau> s))"
- shows "f = fa" and "b = ba" and "[[atom xa]]lst. sa = [[atom x]]lst. s" and "[[atom xa]]lst. \<tau>a = [[atom x]]lst. \<tau>" and
- " [[atom xa]]lst. ca = [[atom x]]lst. c"
- using fun_def.eq_iff fun_typ_q.eq_iff fun_typ.eq_iff lst_snd lst_fst using assms apply metis
- using fun_def.eq_iff fun_typ_q.eq_iff fun_typ.eq_iff lst_snd lst_fst using assms apply metis
-proof -
- have "([[atom xa]]lst. ((ca, \<tau>a), sa) = [[atom x]]lst. ((c, \<tau>), s))" using assms fun_def.eq_iff fun_typ_q.eq_iff fun_typ.eq_iff by auto
- thus "[[atom xa]]lst. sa = [[atom x]]lst. s" and "[[atom xa]]lst. \<tau>a = [[atom x]]lst. \<tau>" and
- " [[atom xa]]lst. ca = [[atom x]]lst. c" using lst_snd lst_fst by metis+
-qed
-
-lemma fun_arg_unique_aux:
- assumes "AF_fun_typ x1 b1 c1 \<tau>1' s1' = AF_fun_typ x2 b2 c2 \<tau>2' s2'"
- shows "\<lbrace> x1 : b1 | c1 \<rbrace> = \<lbrace> x2 : b2 | c2\<rbrace>"
-proof -
- have " ([[atom x1]]lst. c1 = [[atom x2]]lst. c2)" using fun_def_eq assms by metis
- moreover have "b1 = b2" using fun_typ.eq_iff assms by metis
- ultimately show ?thesis using \<tau>.eq_iff by fast
-qed
-
-lemma fresh_x_neq:
- fixes x::x and y::x
- shows "atom x \<sharp> y = (x \<noteq> y)"
- using fresh_at_base fresh_def by auto
-
-lemma obtain_fresh_z3:
- fixes tm::"'b::fs"
- obtains z::x where "\<lbrace> x : b | c \<rbrace> = \<lbrace> z : b | c[x::=V_var z]\<^sub>c\<^sub>v \<rbrace> \<and> atom z \<sharp> tm \<and> atom z \<sharp> (x,c)"
-proof -
- obtain z::x and c'::c where z:"\<lbrace> x : b | c \<rbrace> = \<lbrace> z : b | c' \<rbrace> \<and> atom z \<sharp> (tm,x,c)" using obtain_fresh_z2 b_of.simps by metis
- hence "c' = c[x::=V_var z]\<^sub>c\<^sub>v" proof -
- have "([[atom z]]lst. c' = [[atom x]]lst. c)" using z \<tau>.eq_iff by metis
- hence "c' = (z \<leftrightarrow> x) \<bullet> c" using Abs1_eq_iff[of z c' x c] fresh_x_neq fresh_prodN by fastforce
- also have "... = c[x::=V_var z]\<^sub>c\<^sub>v"
- using subst_v_c_def flip_subst_v[of z c x] z fresh_prod3 by metis
- finally show ?thesis by auto
- qed
- thus ?thesis using z fresh_prodN that by metis
-qed
-
-lemma u_fresh_v:
- fixes u::u and t::v
- shows "atom u \<sharp> t"
- by(nominal_induct t rule:v.strong_induct,auto)
-
-lemma u_fresh_ce:
- fixes u::u and t::ce
- shows "atom u \<sharp> t"
- apply(nominal_induct t rule:ce.strong_induct)
- using u_fresh_v pure_fresh
- apply (auto simp add: opp.fresh ce.fresh opp.fresh opp.exhaust)
- unfolding ce.fresh opp.fresh opp.exhaust by (simp add: fresh_opp_all)
-
-lemma u_fresh_c:
- fixes u::u and t::c
- shows "atom u \<sharp> t"
- by(nominal_induct t rule:c.strong_induct,auto simp add: c.fresh u_fresh_ce)
-
-lemma u_fresh_g:
- fixes u::u and t::\<Gamma>
- shows "atom u \<sharp> t"
- by(induct t rule:\<Gamma>_induct, auto simp add: u_fresh_b u_fresh_c fresh_GCons fresh_GNil)
-
-lemma u_fresh_t:
- fixes u::u and t::\<tau>
- shows "atom u \<sharp> t"
- by(nominal_induct t rule:\<tau>.strong_induct,auto simp add: \<tau>.fresh u_fresh_c u_fresh_b)
-
-lemma b_of_c_of_eq:
- assumes "atom z \<sharp> \<tau>"
- shows "\<lbrace> z : b_of \<tau> | c_of \<tau> z \<rbrace> = \<tau>"
- using assms proof(nominal_induct \<tau> avoiding: z rule: \<tau>.strong_induct)
- case (T_refined_type x1a x2a x3a)
-
- hence " \<lbrace> z : b_of \<lbrace> x1a : x2a | x3a \<rbrace> | c_of \<lbrace> x1a : x2a | x3a \<rbrace> z \<rbrace> = \<lbrace> z : x2a | x3a[x1a::=V_var z]\<^sub>c\<^sub>v \<rbrace>"
- using b_of.simps c_of.simps c_of_eq by auto
- moreover have "\<lbrace> z : x2a | x3a[x1a::=V_var z]\<^sub>c\<^sub>v \<rbrace> = \<lbrace> x1a : x2a | x3a \<rbrace>" using T_refined_type \<tau>.fresh by auto
- ultimately show ?case by auto
-qed
-
-lemma fresh_d_not_in:
- assumes "atom u2 \<sharp> \<Delta>'"
- shows "u2 \<notin> fst ` setD \<Delta>'"
- using assms proof(induct \<Delta>' rule: \<Delta>_induct)
- case DNil
- then show ?case by simp
-next
- case (DCons u t \<Delta>')
- hence *: "atom u2 \<sharp> \<Delta>' \<and> atom u2 \<sharp> (u,t)"
- by (simp add: fresh_def supp_DCons)
- hence "u2 \<notin> fst ` setD \<Delta>'" using DCons by auto
- moreover have "u2 \<noteq> u" using * fresh_Pair
- by (metis eq_fst_iff not_self_fresh)
- ultimately show ?case by simp
-qed
-
+(*<*)
+theory SyntaxL
+ imports Syntax IVSubst
+begin
+ (*>*)
+
+chapter \<open>Syntax Lemmas\<close>
+
+section \<open>Support, lookup and contexts\<close>
+
+lemma supp_v_tau [simp]:
+ assumes "atom z \<sharp> v"
+ shows "supp (\<lbrace> z : b | CE_val (V_var z) == CE_val v \<rbrace>) = supp v \<union> supp b"
+ using assms \<tau>.supp c.supp ce.supp
+ by (simp add: fresh_def supp_at_base)
+
+lemma supp_v_var_tau [simp]:
+ assumes "z \<noteq> x"
+ shows "supp (\<lbrace> z : b | CE_val (V_var z) == CE_val (V_var x) \<rbrace>) = { atom x } \<union> supp b"
+ using supp_v_tau assms
+ using supp_at_base by fastforce
+
+text \<open> Sometimes we need to work with a version of a binder where the variable is fresh
+in something else, such as a bigger context. I think these could be generated automatically \<close>
+
+lemma obtain_fresh_fun_def:
+ fixes t::"'b::fs"
+ shows "\<exists>y::x. atom y \<sharp> (s,c,\<tau>,t) \<and> (AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \<tau> s)) = AF_fundef f (AF_fun_typ_none (AF_fun_typ y b ((y \<leftrightarrow> x) \<bullet> c ) ((y \<leftrightarrow> x) \<bullet> \<tau>) ((y \<leftrightarrow> x) \<bullet> s))))"
+proof -
+ obtain y::x where y: "atom y \<sharp> (s,c,\<tau>,t)" using obtain_fresh by blast
+ moreover have "AF_fundef f (AF_fun_typ_none (AF_fun_typ y b ((y \<leftrightarrow> x) \<bullet> c ) ((y \<leftrightarrow> x) \<bullet> \<tau>) ((y \<leftrightarrow> x) \<bullet> s))) = (AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \<tau> s)))"
+ proof(cases "x=y")
+ case True
+ then show ?thesis using fun_def.eq_iff Abs1_eq_iff(3) flip_commute flip_fresh_fresh fresh_PairD by auto
+ next
+ case False
+
+ have "(AF_fun_typ y b ((y \<leftrightarrow> x) \<bullet> c) ((y \<leftrightarrow> x) \<bullet> \<tau>) ((y \<leftrightarrow> x) \<bullet> s)) =(AF_fun_typ x b c \<tau> s)" proof(subst fun_typ.eq_iff, subst Abs1_eq_iff(3))
+ show \<open>(y = x \<and> (((y \<leftrightarrow> x) \<bullet> c, (y \<leftrightarrow> x) \<bullet> \<tau>), (y \<leftrightarrow> x) \<bullet> s) = ((c, \<tau>), s) \<or>
+ y \<noteq> x \<and> (((y \<leftrightarrow> x) \<bullet> c, (y \<leftrightarrow> x) \<bullet> \<tau>), (y \<leftrightarrow> x) \<bullet> s) = (y \<leftrightarrow> x) \<bullet> ((c, \<tau>), s) \<and> atom y \<sharp> ((c, \<tau>), s)) \<and>
+ b = b\<close> using False flip_commute flip_fresh_fresh fresh_PairD y by auto
+ qed
+ thus ?thesis by metis
+ qed
+ ultimately show ?thesis using y fresh_Pair by metis
+qed
+
+lemma lookup_fun_member:
+ assumes "Some (AF_fundef f ft) = lookup_fun \<Phi> f"
+ shows "AF_fundef f ft \<in> set \<Phi>"
+ using assms proof (induct \<Phi>)
+ case Nil
+ then show ?case by auto
+next
+ case (Cons a \<Phi>)
+ then show ?case using lookup_fun.simps
+ by (metis fun_def.exhaust insert_iff list.simps(15) option.inject)
+qed
+
+lemma rig_dom_eq:
+ "dom (G[x \<longmapsto> c]) = dom G"
+proof(induct G rule: \<Gamma>.induct)
+ case GNil
+ then show ?case using replace_in_g.simps by presburger
+next
+ case (GCons xbc \<Gamma>')
+ obtain x' and b' and c' where xbc: "xbc=(x',b',c')" using prod_cases3 by blast
+ then show ?case using replace_in_g.simps GCons by simp
+qed
+
+lemma lookup_in_rig_eq:
+ assumes "Some (b,c) = lookup \<Gamma> x"
+ shows "Some (b,c') = lookup (\<Gamma>[x\<longmapsto>c']) x"
+ using assms proof(induct \<Gamma> rule: \<Gamma>_induct)
+ case GNil
+ then show ?case by auto
+next
+ case (GCons x b c \<Gamma>')
+ then show ?case using replace_in_g.simps lookup.simps by auto
+qed
+
+lemma lookup_in_rig_neq:
+ assumes "Some (b,c) = lookup \<Gamma> y" and "x\<noteq>y"
+ shows "Some (b,c) = lookup (\<Gamma>[x\<longmapsto>c']) y"
+ using assms proof(induct \<Gamma> rule: \<Gamma>_induct)
+ case GNil
+ then show ?case by auto
+next
+ case (GCons x' b' c' \<Gamma>')
+ then show ?case using replace_in_g.simps lookup.simps by auto
+qed
+
+lemma lookup_in_rig:
+ assumes "Some (b,c) = lookup \<Gamma> y"
+ shows "\<exists>c''. Some (b,c'') = lookup (\<Gamma>[x\<longmapsto>c']) y"
+proof(cases "x=y")
+ case True
+ then show ?thesis using lookup_in_rig_eq using assms by blast
+next
+ case False
+ then show ?thesis using lookup_in_rig_neq using assms by blast
+qed
+
+lemma lookup_inside[simp]:
+ assumes "x \<notin> fst ` toSet \<Gamma>'"
+ shows "Some (b1,c1) = lookup (\<Gamma>'@(x,b1,c1)#\<^sub>\<Gamma>\<Gamma>) x"
+ using assms by(induct \<Gamma>',auto)
+
+lemma lookup_inside2:
+ assumes "Some (b1,c1) = lookup (\<Gamma>'@((x,b0,c0)#\<^sub>\<Gamma>\<Gamma>)) y" and "x\<noteq>y"
+ shows "Some (b1,c1) = lookup (\<Gamma>'@((x,b0,c0')#\<^sub>\<Gamma>\<Gamma>)) y"
+ using assms by(induct \<Gamma>' rule: \<Gamma>.induct,auto+)
+
+fun tail:: "'a list \<Rightarrow> 'a list" where
+ "tail [] = []"
+| "tail (x#xs) = xs"
+
+lemma lookup_options:
+ assumes "Some (b,c) = lookup (xt#\<^sub>\<Gamma>G) x"
+ shows "((x,b,c) = xt) \<or> (Some (b,c) = lookup G x)"
+ by (metis assms lookup.simps(2) option.inject surj_pair)
+
+lemma lookup_x:
+ assumes "Some (b,c) = lookup G x"
+ shows "x \<in> fst ` toSet G"
+ using assms
+ by(induct "G" rule: \<Gamma>.induct ,auto+)
+
+lemma GCons_eq_appendI:
+ fixes xs1::\<Gamma>
+ shows "[| x #\<^sub>\<Gamma> xs1 = ys; xs = xs1 @ zs |] ==> x #\<^sub>\<Gamma> xs = ys @ zs"
+ by (drule sym) simp
+
+lemma split_G: "x : toSet xs \<Longrightarrow> \<exists>ys zs. xs = ys @ x #\<^sub>\<Gamma> zs"
+proof (induct xs)
+ case GNil thus ?case by simp
+next
+ case GCons thus ?case using GCons_eq_appendI
+ by (metis Un_iff append_g.simps(1) singletonD toSet.simps(2))
+qed
+
+lemma lookup_not_empty:
+ assumes "Some \<tau> = lookup G x"
+ shows "G \<noteq> GNil"
+ using assms by auto
+
+lemma lookup_in_g:
+ assumes "Some (b,c) = lookup \<Gamma> x"
+ shows "(x,b,c) \<in> toSet \<Gamma>"
+ using assms apply(induct \<Gamma>, simp)
+ using lookup_options by fastforce
+
+lemma lookup_split:
+ fixes \<Gamma>::\<Gamma>
+ assumes "Some (b,c) = lookup \<Gamma> x"
+ shows "\<exists>G G' . \<Gamma> = G'@(x,b,c)#\<^sub>\<Gamma>G"
+ by (meson assms(1) lookup_in_g split_G)
+
+lemma toSet_splitU[simp]:
+ "(x',b',c') \<in> toSet (\<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>) \<longleftrightarrow> (x',b',c') \<in> (toSet \<Gamma>' \<union> {(x, b, c)} \<union> toSet \<Gamma>)"
+ using append_g_toSetU toSet.simps by auto
+
+lemma toSet_splitP[simp]:
+ "(\<forall>(x', b', c')\<in>toSet (\<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>). P x' b' c') \<longleftrightarrow> (\<forall> (x', b', c')\<in>toSet \<Gamma>'. P x' b' c') \<and> P x b c \<and> (\<forall> (x', b', c')\<in>toSet \<Gamma>. P x' b' c')" (is "?A \<longleftrightarrow> ?B")
+ using toSet_splitU by force
+
+lemma lookup_restrict:
+ assumes "Some (b',c') = lookup (\<Gamma>'@(x,b,c)#\<^sub>\<Gamma>\<Gamma>) y" and "x \<noteq> y"
+ shows "Some (b',c') = lookup (\<Gamma>'@\<Gamma>) y"
+ using assms proof(induct \<Gamma>' rule:\<Gamma>_induct)
+ case GNil
+ then show ?case by auto
+next
+ case (GCons x1 b1 c1 \<Gamma>')
+ then show ?case by auto
+qed
+
+lemma supp_list_member:
+ fixes x::"'a::fs" and l::"'a list"
+ assumes "x \<in> set l"
+ shows "supp x \<subseteq> supp l"
+ using assms apply(induct l, auto)
+ using supp_Cons by auto
+
+lemma GNil_append:
+ assumes "GNil = G1@G2"
+ shows "G1 = GNil \<and> G2 = GNil"
+proof(rule ccontr)
+ assume "\<not> (G1 = GNil \<and> G2 = GNil)"
+ hence "G1@G2 \<noteq> GNil" using append_g.simps by (metis \<Gamma>.distinct(1) \<Gamma>.exhaust)
+ thus False using assms by auto
+qed
+
+lemma GCons_eq_append_conv:
+ fixes xs::\<Gamma>
+ shows "x#\<^sub>\<Gamma>xs = ys@zs = (ys = GNil \<and> x#\<^sub>\<Gamma>xs = zs \<or> (\<exists>ys'. x#\<^sub>\<Gamma>ys' = ys \<and> xs = ys'@zs))"
+ by(cases ys) auto
+
+lemma dclist_distinct_unique:
+ assumes "(dc, const) \<in> set dclist2" and "(cons, const1) \<in> set dclist2" and "dc=cons" and "distinct (List.map fst dclist2)"
+ shows "(const) = const1"
+proof -
+ have "(cons, const) = (dc, const1)"
+ using assms by (metis (no_types, lifting) assms(3) assms(4) distinct.simps(1) distinct.simps(2) empty_iff insert_iff list.set(1) list.simps(15) list.simps(8) list.simps(9) map_of_eq_Some_iff)
+ thus ?thesis by auto
+qed
+
+lemma fresh_d_fst_d:
+ assumes "atom u \<sharp> \<delta>"
+ shows "u \<notin> fst ` set \<delta>"
+ using assms proof(induct \<delta>)
+ case Nil
+ then show ?case by auto
+next
+ case (Cons ut \<delta>')
+ obtain u' and t' where *:"ut = (u',t') " by fastforce
+ hence "atom u \<sharp> ut \<and> atom u \<sharp> \<delta>'" using fresh_Cons Cons by auto
+ moreover hence "atom u \<sharp> fst ut" using * fresh_Pair[of "atom u" u' t'] Cons by auto
+ ultimately show ?case using Cons by auto
+qed
+
+lemma bv_not_in_bset_supp:
+ fixes bv::bv
+ assumes "bv |\<notin>| B"
+ shows "atom bv \<notin> supp B"
+proof -
+ have *:"supp B = fset (fimage atom B)"
+ by (metis fimage.rep_eq finite_fset supp_finite_set_at_base supp_fset)
+ thus ?thesis using assms
+ using notin_fset by fastforce
+qed
+
+lemma u_fresh_d:
+ assumes "atom u \<sharp> D"
+ shows "u \<notin> fst ` setD D"
+ using assms proof(induct D rule: \<Delta>_induct)
+ case DNil
+ then show ?case by auto
+next
+ case (DCons u' t' \<Delta>')
+ then show ?case unfolding setD.simps
+ using fresh_DCons fresh_Pair by (simp add: fresh_Pair fresh_at_base(2))
+qed
+
+section \<open>Type Definitions\<close>
+
+lemma exist_fresh_bv:
+ fixes tm::"'a::fs"
+ shows "\<exists>bva2 dclist2. AF_typedef_poly tyid bva dclist = AF_typedef_poly tyid bva2 dclist2 \<and>
+ atom bva2 \<sharp> tm"
+proof -
+ obtain bva2::bv where *:"atom bva2 \<sharp> (bva, dclist,tyid,tm)" using obtain_fresh by metis
+ moreover hence "bva2 \<noteq> bva" using fresh_at_base by auto
+ moreover have " dclist = (bva \<leftrightarrow> bva2) \<bullet> (bva2 \<leftrightarrow> bva) \<bullet> dclist" by simp
+ moreover have "atom bva \<sharp> (bva2 \<leftrightarrow> bva) \<bullet> dclist" proof -
+ have "atom bva2 \<sharp> dclist" using * fresh_prodN by auto
+ hence "atom ((bva2 \<leftrightarrow> bva) \<bullet> bva2) \<sharp> (bva2 \<leftrightarrow> bva) \<bullet> dclist" using fresh_eqvt True_eqvt
+ proof -
+ have "(bva2 \<leftrightarrow> bva) \<bullet> atom bva2 \<sharp> (bva2 \<leftrightarrow> bva) \<bullet> dclist"
+ by (metis True_eqvt \<open>atom bva2 \<sharp> dclist\<close> fresh_eqvt) (* 62 ms *)
+ then show ?thesis
+ by simp (* 125 ms *)
+ qed
+ thus ?thesis by auto
+ qed
+ ultimately have "AF_typedef_poly tyid bva dclist = AF_typedef_poly tyid bva2 ((bva2 \<leftrightarrow> bva ) \<bullet> dclist)"
+ unfolding type_def.eq_iff Abs1_eq_iff by metis
+ thus ?thesis using * fresh_prodN by metis
+qed
+
+lemma obtain_fresh_bv:
+ fixes tm::"'a::fs"
+ obtains bva2::bv and dclist2 where "AF_typedef_poly tyid bva dclist = AF_typedef_poly tyid bva2 dclist2 \<and>
+ atom bva2 \<sharp> tm"
+ using exist_fresh_bv by metis
+
+section \<open>Function Definitions\<close>
+
+lemma fun_typ_flip:
+ fixes bv1::bv and c::bv
+ shows "(bv1 \<leftrightarrow> c) \<bullet> AF_fun_typ x1 b1 c1 \<tau>1 s1 = AF_fun_typ x1 ((bv1 \<leftrightarrow> c) \<bullet> b1) ((bv1 \<leftrightarrow> c) \<bullet> c1) ((bv1 \<leftrightarrow> c) \<bullet> \<tau>1) ((bv1 \<leftrightarrow> c) \<bullet> s1)"
+ using fun_typ.perm_simps flip_fresh_fresh supp_at_base fresh_def
+ flip_fresh_fresh fresh_def supp_at_base
+ by (simp add: flip_fresh_fresh)
+
+lemma fun_def_eq:
+ assumes "AF_fundef fa (AF_fun_typ_none (AF_fun_typ xa ba ca \<tau>a sa)) = AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \<tau> s))"
+ shows "f = fa" and "b = ba" and "[[atom xa]]lst. sa = [[atom x]]lst. s" and "[[atom xa]]lst. \<tau>a = [[atom x]]lst. \<tau>" and
+ " [[atom xa]]lst. ca = [[atom x]]lst. c"
+ using fun_def.eq_iff fun_typ_q.eq_iff fun_typ.eq_iff lst_snd lst_fst using assms apply metis
+ using fun_def.eq_iff fun_typ_q.eq_iff fun_typ.eq_iff lst_snd lst_fst using assms apply metis
+proof -
+ have "([[atom xa]]lst. ((ca, \<tau>a), sa) = [[atom x]]lst. ((c, \<tau>), s))" using assms fun_def.eq_iff fun_typ_q.eq_iff fun_typ.eq_iff by auto
+ thus "[[atom xa]]lst. sa = [[atom x]]lst. s" and "[[atom xa]]lst. \<tau>a = [[atom x]]lst. \<tau>" and
+ " [[atom xa]]lst. ca = [[atom x]]lst. c" using lst_snd lst_fst by metis+
+qed
+
+lemma fun_arg_unique_aux:
+ assumes "AF_fun_typ x1 b1 c1 \<tau>1' s1' = AF_fun_typ x2 b2 c2 \<tau>2' s2'"
+ shows "\<lbrace> x1 : b1 | c1 \<rbrace> = \<lbrace> x2 : b2 | c2\<rbrace>"
+proof -
+ have " ([[atom x1]]lst. c1 = [[atom x2]]lst. c2)" using fun_def_eq assms by metis
+ moreover have "b1 = b2" using fun_typ.eq_iff assms by metis
+ ultimately show ?thesis using \<tau>.eq_iff by fast
+qed
+
+lemma fresh_x_neq:
+ fixes x::x and y::x
+ shows "atom x \<sharp> y = (x \<noteq> y)"
+ using fresh_at_base fresh_def by auto
+
+lemma obtain_fresh_z3:
+ fixes tm::"'b::fs"
+ obtains z::x where "\<lbrace> x : b | c \<rbrace> = \<lbrace> z : b | c[x::=V_var z]\<^sub>c\<^sub>v \<rbrace> \<and> atom z \<sharp> tm \<and> atom z \<sharp> (x,c)"
+proof -
+ obtain z::x and c'::c where z:"\<lbrace> x : b | c \<rbrace> = \<lbrace> z : b | c' \<rbrace> \<and> atom z \<sharp> (tm,x,c)" using obtain_fresh_z2 b_of.simps by metis
+ hence "c' = c[x::=V_var z]\<^sub>c\<^sub>v" proof -
+ have "([[atom z]]lst. c' = [[atom x]]lst. c)" using z \<tau>.eq_iff by metis
+ hence "c' = (z \<leftrightarrow> x) \<bullet> c" using Abs1_eq_iff[of z c' x c] fresh_x_neq fresh_prodN by fastforce
+ also have "... = c[x::=V_var z]\<^sub>c\<^sub>v"
+ using subst_v_c_def flip_subst_v[of z c x] z fresh_prod3 by metis
+ finally show ?thesis by auto
+ qed
+ thus ?thesis using z fresh_prodN that by metis
+qed
+
+lemma u_fresh_v:
+ fixes u::u and t::v
+ shows "atom u \<sharp> t"
+ by(nominal_induct t rule:v.strong_induct,auto)
+
+lemma u_fresh_ce:
+ fixes u::u and t::ce
+ shows "atom u \<sharp> t"
+ apply(nominal_induct t rule:ce.strong_induct)
+ using u_fresh_v pure_fresh
+ apply (auto simp add: opp.fresh ce.fresh opp.fresh opp.exhaust)
+ unfolding ce.fresh opp.fresh opp.exhaust by (simp add: fresh_opp_all)
+
+lemma u_fresh_c:
+ fixes u::u and t::c
+ shows "atom u \<sharp> t"
+ by(nominal_induct t rule:c.strong_induct,auto simp add: c.fresh u_fresh_ce)
+
+lemma u_fresh_g:
+ fixes u::u and t::\<Gamma>
+ shows "atom u \<sharp> t"
+ by(induct t rule:\<Gamma>_induct, auto simp add: u_fresh_b u_fresh_c fresh_GCons fresh_GNil)
+
+lemma u_fresh_t:
+ fixes u::u and t::\<tau>
+ shows "atom u \<sharp> t"
+ by(nominal_induct t rule:\<tau>.strong_induct,auto simp add: \<tau>.fresh u_fresh_c u_fresh_b)
+
+lemma b_of_c_of_eq:
+ assumes "atom z \<sharp> \<tau>"
+ shows "\<lbrace> z : b_of \<tau> | c_of \<tau> z \<rbrace> = \<tau>"
+ using assms proof(nominal_induct \<tau> avoiding: z rule: \<tau>.strong_induct)
+ case (T_refined_type x1a x2a x3a)
+
+ hence " \<lbrace> z : b_of \<lbrace> x1a : x2a | x3a \<rbrace> | c_of \<lbrace> x1a : x2a | x3a \<rbrace> z \<rbrace> = \<lbrace> z : x2a | x3a[x1a::=V_var z]\<^sub>c\<^sub>v \<rbrace>"
+ using b_of.simps c_of.simps c_of_eq by auto
+ moreover have "\<lbrace> z : x2a | x3a[x1a::=V_var z]\<^sub>c\<^sub>v \<rbrace> = \<lbrace> x1a : x2a | x3a \<rbrace>" using T_refined_type \<tau>.fresh by auto
+ ultimately show ?case by auto
+qed
+
+lemma fresh_d_not_in:
+ assumes "atom u2 \<sharp> \<Delta>'"
+ shows "u2 \<notin> fst ` setD \<Delta>'"
+ using assms proof(induct \<Delta>' rule: \<Delta>_induct)
+ case DNil
+ then show ?case by simp
+next
+ case (DCons u t \<Delta>')
+ hence *: "atom u2 \<sharp> \<Delta>' \<and> atom u2 \<sharp> (u,t)"
+ by (simp add: fresh_def supp_DCons)
+ hence "u2 \<notin> fst ` setD \<Delta>'" using DCons by auto
+ moreover have "u2 \<noteq> u" using * fresh_Pair
+ by (metis eq_fst_iff not_self_fresh)
+ ultimately show ?case by simp
+qed
+
end
\ No newline at end of file
diff --git a/thys/MiniSail/Typing.thy b/thys/MiniSail/Typing.thy
--- a/thys/MiniSail/Typing.thy
+++ b/thys/MiniSail/Typing.thy
@@ -1,586 +1,586 @@
-(*<*)
-theory Typing
- imports RCLogic WellformedL
-begin
- (*>*)
-
-chapter \<open>Type System\<close>
-
-text \<open>The MiniSail type system. We define subtyping judgement first and then typing judgement
-for the term forms\<close>
-
-section \<open>Subtyping\<close>
-
-text \<open> Subtyping is defined on top of refinement constraint logic (RCL).
-A subtyping check is converted into an RCL validity check. \<close>
-
-inductive subtype :: "\<Theta> \<Rightarrow> \<B> \<Rightarrow> \<Gamma> \<Rightarrow> \<tau> \<Rightarrow> \<tau> \<Rightarrow> bool" ("_ ; _ ; _ \<turnstile> _ \<lesssim> _" [50, 50, 50] 50) where
- subtype_baseI: "\<lbrakk>
- atom x \<sharp> (\<Theta>, \<B>, \<Gamma>, z,c,z',c') ;
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b | c \<rbrace>;
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> z' : b | c' \<rbrace>;
- \<Theta>; \<B> ; (x,b, c[z::=[x]\<^sup>v]\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<Turnstile> c'[z'::=[x]\<^sup>v]\<^sub>v
-\<rbrakk> \<Longrightarrow>
- \<Theta>; \<B>; \<Gamma> \<turnstile> \<lbrace> z : b | c \<rbrace> \<lesssim> \<lbrace> z' : b | c' \<rbrace>"
-
-equivariance subtype
-nominal_inductive subtype
- avoids subtype_baseI: x
-proof(goal_cases)
- case (1 \<Theta> \<B> \<Gamma> z b c z' c' x)
- then show ?case using fresh_star_def 1 by force
-next
- case (2 \<Theta> \<B> \<Gamma> z b c z' c' x)
- then show ?case by auto
-qed
-
-inductive_cases subtype_elims:
- "\<Theta>; \<B>; \<Gamma> \<turnstile> \<lbrace> z : b | c \<rbrace> \<lesssim> \<lbrace> z' : b | c' \<rbrace>"
- "\<Theta>; \<B>; \<Gamma> \<turnstile> \<tau>\<^sub>1 \<lesssim> \<tau>\<^sub>2"
-
-section \<open>Literals\<close>
-
-text \<open>The type synthesised has the constraint that z equates to the literal\<close>
-
-inductive infer_l :: "l \<Rightarrow> \<tau> \<Rightarrow> bool" (" \<turnstile> _ \<Rightarrow> _" [50, 50] 50) where
- infer_trueI: " \<turnstile> L_true \<Rightarrow> \<lbrace> z : B_bool | [[z]\<^sup>v]\<^sup>c\<^sup>e == [[L_true]\<^sup>v]\<^sup>c\<^sup>e \<rbrace>"
-| infer_falseI: " \<turnstile> L_false \<Rightarrow> \<lbrace> z : B_bool | [[z]\<^sup>v]\<^sup>c\<^sup>e == [[L_false]\<^sup>v]\<^sup>c\<^sup>e \<rbrace>"
-| infer_natI: " \<turnstile> L_num n \<Rightarrow> \<lbrace> z : B_int | [[z]\<^sup>v]\<^sup>c\<^sup>e == [[L_num n]\<^sup>v]\<^sup>c\<^sup>e \<rbrace>"
-| infer_unitI: " \<turnstile> L_unit \<Rightarrow> \<lbrace> z : B_unit | [[z]\<^sup>v]\<^sup>c\<^sup>e == [[L_unit]\<^sup>v]\<^sup>c\<^sup>e \<rbrace>"
-| infer_bitvecI: " \<turnstile> L_bitvec bv \<Rightarrow> \<lbrace> z : B_bitvec | [[z]\<^sup>v]\<^sup>c\<^sup>e == [[L_bitvec bv]\<^sup>v]\<^sup>c\<^sup>e \<rbrace>"
-
-nominal_inductive infer_l .
-equivariance infer_l
-
-inductive_cases infer_l_elims[elim!]:
- "\<turnstile> L_true \<Rightarrow> \<tau>"
- "\<turnstile> L_false \<Rightarrow> \<tau>"
- "\<turnstile> L_num n \<Rightarrow> \<tau>"
- "\<turnstile> L_unit \<Rightarrow> \<tau>"
- "\<turnstile> L_bitvec x \<Rightarrow> \<tau>"
- "\<turnstile> l \<Rightarrow> \<tau>"
-
-lemma infer_l_form2[simp]:
- shows "\<exists>z. \<turnstile> l \<Rightarrow> (\<lbrace> z : base_for_lit l | [[z]\<^sup>v]\<^sup>c\<^sup>e == [[l]\<^sup>v]\<^sup>c\<^sup>e \<rbrace>)"
-proof (nominal_induct l rule: l.strong_induct)
- case (L_num x)
- then show ?case using infer_l.intros base_for_lit.simps has_fresh_z by metis
-next
- case L_true
- then show ?case using infer_l.intros base_for_lit.simps has_fresh_z by metis
-next
- case L_false
- then show ?case using infer_l.intros base_for_lit.simps has_fresh_z by metis
-next
- case L_unit
- then show ?case using infer_l.intros base_for_lit.simps has_fresh_z by metis
-next
- case (L_bitvec x)
- then show ?case using infer_l.intros base_for_lit.simps has_fresh_z by metis
-qed
-
-section \<open>Values\<close>
-
-inductive infer_v :: "\<Theta> \<Rightarrow> \<B> \<Rightarrow> \<Gamma> \<Rightarrow> v \<Rightarrow> \<tau> \<Rightarrow> bool" (" _ ; _ ; _ \<turnstile> _ \<Rightarrow> _" [50, 50, 50] 50) where
-
-infer_v_varI: "\<lbrakk>
- \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> ;
- Some (b,c) = lookup \<Gamma> x;
- atom z \<sharp> x ; atom z \<sharp> (\<Theta>, \<B>, \<Gamma>)
-\<rbrakk> \<Longrightarrow>
- \<Theta>; \<B>; \<Gamma> \<turnstile> [x]\<^sup>v \<Rightarrow> \<lbrace> z : b | [[z]\<^sup>v]\<^sup>c\<^sup>e == [[x]\<^sup>v]\<^sup>c\<^sup>e \<rbrace>"
-
-| infer_v_litI: "\<lbrakk>
- \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> ;
- \<turnstile> l \<Rightarrow> \<tau>
-\<rbrakk> \<Longrightarrow>
- \<Theta>; \<B>; \<Gamma> \<turnstile> [l]\<^sup>v \<Rightarrow> \<tau>"
-
-| infer_v_pairI: "\<lbrakk>
- atom z \<sharp> (v1, v2); atom z \<sharp> (\<Theta>, \<B>, \<Gamma>) ;
- \<Theta>; \<B>; \<Gamma> \<turnstile> (v1::v) \<Rightarrow> t1 ;
- \<Theta>; \<B> ; \<Gamma> \<turnstile> (v2::v) \<Rightarrow> t2
-\<rbrakk> \<Longrightarrow>
- \<Theta>; \<B>; \<Gamma> \<turnstile> V_pair v1 v2 \<Rightarrow> (\<lbrace> z : B_pair (b_of t1) (b_of t2) | [[z]\<^sup>v]\<^sup>c\<^sup>e == [[v1,v2]\<^sup>v]\<^sup>c\<^sup>e \<rbrace>) "
-
-| infer_v_consI: "\<lbrakk>
- AF_typedef s dclist \<in> set \<Theta>;
- (dc, tc) \<in> set dclist ;
- \<Theta>; \<B>; \<Gamma> \<turnstile> v \<Rightarrow> tv ;
- \<Theta>; \<B>; \<Gamma> \<turnstile> tv \<lesssim> tc ;
- atom z \<sharp> v ; atom z \<sharp> (\<Theta>, \<B>, \<Gamma>)
-\<rbrakk> \<Longrightarrow>
- \<Theta>; \<B>; \<Gamma> \<turnstile> V_cons s dc v \<Rightarrow> (\<lbrace> z : B_id s | [[z]\<^sup>v]\<^sup>c\<^sup>e == [ V_cons s dc v ]\<^sup>c\<^sup>e \<rbrace>)"
-
-| infer_v_conspI: "\<lbrakk>
- AF_typedef_poly s bv dclist \<in> set \<Theta>;
- (dc, tc) \<in> set dclist ;
- \<Theta>; \<B>; \<Gamma> \<turnstile> v \<Rightarrow> tv;
- \<Theta>; \<B>; \<Gamma> \<turnstile> tv \<lesssim> tc[bv::=b]\<^sub>\<tau>\<^sub>b ;
- atom z \<sharp> (\<Theta>, \<B>, \<Gamma>, v, b);
- atom bv \<sharp> (\<Theta>, \<B>, \<Gamma>, v, b);
- \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b
-\<rbrakk> \<Longrightarrow>
- \<Theta>; \<B>; \<Gamma> \<turnstile> V_consp s dc b v \<Rightarrow> (\<lbrace> z : B_app s b | [[z]\<^sup>v]\<^sup>c\<^sup>e == (CE_val (V_consp s dc b v)) \<rbrace>)"
-
-equivariance infer_v
-nominal_inductive infer_v
- avoids infer_v_conspI: bv and z | infer_v_varI: z | infer_v_pairI: z | infer_v_consI: z
-proof(goal_cases)
- case (1 \<Theta> \<B> \<Gamma> b c x z)
- hence "atom z \<sharp> \<lbrace> z : b | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrace>" using \<tau>.fresh by simp
- then show ?case unfolding fresh_star_def using 1 by simp
-next
- case (2 \<Theta> \<B> \<Gamma> b c x z)
- then show ?case by auto
-next
- case (3 z v1 v2 \<Theta> \<B> \<Gamma> t1 t2)
- hence "atom z \<sharp> \<lbrace> z : [ b_of t1 , b_of t2 ]\<^sup>b | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ v1 , v2 ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrace>" using \<tau>.fresh by simp
- then show ?case unfolding fresh_star_def using 3 by simp
-next
- case (4 z v1 v2 \<Theta> \<B> \<Gamma> t1 t2)
- then show ?case by auto
-next
- case (5 s dclist \<Theta> dc tc \<B> \<Gamma> v tv z)
- hence "atom z \<sharp> \<lbrace> z : B_id s | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ V_cons s dc v ]\<^sup>c\<^sup>e \<rbrace>" using \<tau>.fresh b.fresh pure_fresh by auto
- moreover have "atom z \<sharp> V_cons s dc v" using v.fresh 5 using v.fresh fresh_prodN pure_fresh by metis
- then show ?case unfolding fresh_star_def using 5 by simp
-next
- case (6 s dclist \<Theta> dc tc \<B> \<Gamma> v tv z)
- then show ?case by auto
-next
- case (7 s bv dclist \<Theta> dc tc \<B> \<Gamma> v tv b z)
- hence "atom bv \<sharp> V_consp s dc b v" using v.fresh fresh_prodN pure_fresh by metis
- moreover then have "atom bv \<sharp> \<lbrace> z : B_id s | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ V_consp s dc b v ]\<^sup>c\<^sup>e \<rbrace>"
- using \<tau>.fresh ce.fresh v.fresh by auto
- moreover have "atom z \<sharp> V_consp s dc b v" using v.fresh fresh_prodN pure_fresh 7 by metis
- moreover then have "atom z \<sharp> \<lbrace> z : B_id s | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ V_consp s dc b v ]\<^sup>c\<^sup>e \<rbrace>"
- using \<tau>.fresh ce.fresh v.fresh by auto
- ultimately show ?case using fresh_star_def 7 by force
-next
- case (8 s bv dclist \<Theta> dc tc \<B> \<Gamma> v tv b z)
- then show ?case by auto
-qed
-
-inductive_cases infer_v_elims[elim!]:
- "\<Theta>; \<B>; \<Gamma> \<turnstile> V_var x \<Rightarrow> \<tau>"
- "\<Theta>; \<B>; \<Gamma> \<turnstile> V_lit l \<Rightarrow> \<tau>"
- "\<Theta>; \<B>; \<Gamma> \<turnstile> V_pair v1 v2 \<Rightarrow> \<tau>"
- "\<Theta>; \<B>; \<Gamma> \<turnstile> V_cons s dc v \<Rightarrow> \<tau>"
- "\<Theta>; \<B>; \<Gamma> \<turnstile> V_pair v1 v2 \<Rightarrow> (\<lbrace> z : b | c \<rbrace>) "
- "\<Theta>; \<B>; \<Gamma> \<turnstile> V_pair v1 v2 \<Rightarrow> (\<lbrace> z : [ b1 , b2 ]\<^sup>b | [[z]\<^sup>v]\<^sup>c\<^sup>e == [[v1,v2]\<^sup>v]\<^sup>c\<^sup>e \<rbrace>) "
- "\<Theta>; \<B>; \<Gamma> \<turnstile> V_consp s dc b v \<Rightarrow> \<tau> "
-
-inductive check_v :: "\<Theta> \<Rightarrow> \<B> \<Rightarrow> \<Gamma> \<Rightarrow> v \<Rightarrow> \<tau> \<Rightarrow> bool" ("_ ; _ ; _ \<turnstile> _ \<Leftarrow> _" [50, 50, 50] 50) where
- check_v_subtypeI: "\<lbrakk> \<Theta>; \<B>; \<Gamma> \<turnstile> \<tau>1 \<lesssim> \<tau>2; \<Theta>; \<B>; \<Gamma> \<turnstile> v \<Rightarrow> \<tau>1 \<rbrakk> \<Longrightarrow> \<Theta>; \<B> ; \<Gamma> \<turnstile> v \<Leftarrow> \<tau>2"
-equivariance check_v
-nominal_inductive check_v .
-
-inductive_cases check_v_elims[elim!]:
- "\<Theta>; \<B> ; \<Gamma> \<turnstile> v \<Leftarrow> \<tau>"
-
-section \<open>Expressions\<close>
-
-text \<open> Type synthesis for expressions \<close>
-
-inductive infer_e :: "\<Theta> \<Rightarrow> \<Phi> \<Rightarrow> \<B> \<Rightarrow> \<Gamma> \<Rightarrow> \<Delta> \<Rightarrow> e \<Rightarrow> \<tau> \<Rightarrow> bool" ("_ ; _ ; _ ; _ ; _ \<turnstile> _ \<Rightarrow> _" [50, 50, 50,50] 50) where
-
-infer_e_valI: "\<lbrakk>
- (\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>) ;
- (\<Theta> \<turnstile>\<^sub>w\<^sub>f (\<Phi>::\<Phi>)) ;
- (\<Theta>; \<B>; \<Gamma> \<turnstile> v \<Rightarrow> \<tau>) \<rbrakk> \<Longrightarrow>
- \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> (AE_val v) \<Rightarrow> \<tau>"
-
-| infer_e_plusI: "\<lbrakk>
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> ;
- \<Theta> \<turnstile>\<^sub>w\<^sub>f (\<Phi>::\<Phi>) ;
- \<Theta>; \<B>; \<Gamma> \<turnstile> v1 \<Rightarrow> \<lbrace> z1 : B_int | c1 \<rbrace> ;
- \<Theta>; \<B>; \<Gamma> \<turnstile> v2 \<Rightarrow> \<lbrace> z2 : B_int | c2 \<rbrace>;
- atom z3 \<sharp> (AE_op Plus v1 v2); atom z3 \<sharp> \<Gamma> \<rbrakk> \<Longrightarrow>
- \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> AE_op Plus v1 v2 \<Rightarrow> \<lbrace> z3 : B_int | [[z3]\<^sup>v]\<^sup>c\<^sup>e == (CE_op Plus [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e) \<rbrace>"
-
-| infer_e_leqI: "\<lbrakk>
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>;
- \<Theta> \<turnstile>\<^sub>w\<^sub>f (\<Phi>::\<Phi>) ;
- \<Theta>; \<B>; \<Gamma> \<turnstile> v1 \<Rightarrow> \<lbrace> z1 : B_int | c1 \<rbrace> ;
- \<Theta>; \<B>; \<Gamma> \<turnstile> v2 \<Rightarrow> \<lbrace> z2 : B_int | c2 \<rbrace>;
- atom z3 \<sharp> (AE_op LEq v1 v2); atom z3 \<sharp> \<Gamma>
-\<rbrakk> \<Longrightarrow>
- \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> AE_op LEq v1 v2 \<Rightarrow> \<lbrace> z3 : B_bool | [[z3]\<^sup>v]\<^sup>c\<^sup>e == (CE_op LEq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e) \<rbrace>"
-
-| infer_e_eqI: "\<lbrakk>
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>;
- \<Theta> \<turnstile>\<^sub>w\<^sub>f (\<Phi>::\<Phi>) ;
- \<Theta>; \<B>; \<Gamma> \<turnstile> v1 \<Rightarrow> \<lbrace> z1 : b | c1 \<rbrace> ;
- \<Theta>; \<B>; \<Gamma> \<turnstile> v2 \<Rightarrow> \<lbrace> z2 : b | c2 \<rbrace>;
- atom z3 \<sharp> (AE_op Eq v1 v2); atom z3 \<sharp> \<Gamma> ;
- b \<in> { B_bool, B_int, B_unit }
-\<rbrakk> \<Longrightarrow>
- \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> AE_op Eq v1 v2 \<Rightarrow> \<lbrace> z3 : B_bool | [[z3]\<^sup>v]\<^sup>c\<^sup>e == (CE_op Eq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e) \<rbrace>"
-
-| infer_e_appI: "\<lbrakk>
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> ;
- \<Theta> \<turnstile>\<^sub>w\<^sub>f (\<Phi>::\<Phi>) ;
- Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \<tau>' s'))) = lookup_fun \<Phi> f;
- \<Theta>; \<B>; \<Gamma> \<turnstile> v \<Leftarrow> \<lbrace> x : b | c \<rbrace>;
- atom x \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>, \<Delta>,v , \<tau>);
- \<tau>'[x::=v]\<^sub>v = \<tau>
-\<rbrakk> \<Longrightarrow>
- \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> AE_app f v \<Rightarrow> \<tau>"
-
-| infer_e_appPI: "\<lbrakk>
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> ;
- \<Theta> \<turnstile>\<^sub>w\<^sub>f (\<Phi>::\<Phi>) ;
- \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b' ;
- Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \<tau>' s'))) = lookup_fun \<Phi> f;
- \<Theta>; \<B>; \<Gamma> \<turnstile> v \<Leftarrow> \<lbrace> x : b[bv::=b']\<^sub>b | c[bv::=b']\<^sub>b \<rbrace>; atom x \<sharp> \<Gamma>;
- (\<tau>'[bv::=b']\<^sub>b[x::=v]\<^sub>v) = \<tau> ;
- atom bv \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>, \<Delta>, b', v, \<tau>)
-\<rbrakk> \<Longrightarrow>
- \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> AE_appP f b' v \<Rightarrow> \<tau>"
-
-| infer_e_fstI: "\<lbrakk>
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> ;
- \<Theta> \<turnstile>\<^sub>w\<^sub>f (\<Phi>::\<Phi>) ;
- \<Theta>; \<B>; \<Gamma> \<turnstile> v \<Rightarrow> \<lbrace> z' : [b1,b2]\<^sup>b | c \<rbrace>;
- atom z \<sharp> AE_fst v ; atom z \<sharp> \<Gamma> \<rbrakk> \<Longrightarrow>
- \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> AE_fst v \<Rightarrow> \<lbrace> z : b1 | [[z]\<^sup>v]\<^sup>c\<^sup>e == ((CE_fst [v]\<^sup>c\<^sup>e)) \<rbrace>"
-
-| infer_e_sndI: "\<lbrakk>
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> ;
- \<Theta> \<turnstile>\<^sub>w\<^sub>f (\<Phi>::\<Phi>) ;
- \<Theta>; \<B>; \<Gamma> \<turnstile> v \<Rightarrow> \<lbrace> z' : [ b1, b2]\<^sup>b | c \<rbrace>;
- atom z \<sharp> AE_snd v ; atom z \<sharp> \<Gamma> \<rbrakk> \<Longrightarrow>
- \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> AE_snd v \<Rightarrow> \<lbrace> z : b2 | [[z]\<^sup>v]\<^sup>c\<^sup>e == ((CE_snd [v]\<^sup>c\<^sup>e)) \<rbrace>"
-
-| infer_e_lenI: "\<lbrakk>
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> ;
- \<Theta> \<turnstile>\<^sub>w\<^sub>f (\<Phi>::\<Phi>) ;
- \<Theta>; \<B>; \<Gamma> \<turnstile> v \<Rightarrow> \<lbrace> z' : B_bitvec | c \<rbrace>;
- atom z \<sharp> AE_len v ; atom z \<sharp> \<Gamma>\<rbrakk> \<Longrightarrow>
- \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> AE_len v \<Rightarrow> \<lbrace> z : B_int | [[z]\<^sup>v]\<^sup>c\<^sup>e == ((CE_len [v]\<^sup>c\<^sup>e)) \<rbrace>"
-
-| infer_e_mvarI: "\<lbrakk>
- \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> ;
- \<Theta> \<turnstile>\<^sub>w\<^sub>f (\<Phi>::\<Phi>) ;
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>;
- (u,\<tau>) \<in> setD \<Delta> \<rbrakk> \<Longrightarrow>
- \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> AE_mvar u \<Rightarrow> \<tau>"
-
-| infer_e_concatI: "\<lbrakk>
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> ;
- \<Theta> \<turnstile>\<^sub>w\<^sub>f (\<Phi>::\<Phi>) ;
- \<Theta>; \<B>; \<Gamma> \<turnstile> v1 \<Rightarrow> \<lbrace> z1 : B_bitvec | c1 \<rbrace> ;
- \<Theta>; \<B>; \<Gamma> \<turnstile> v2 \<Rightarrow> \<lbrace> z2 : B_bitvec | c2 \<rbrace>;
- atom z3 \<sharp> (AE_concat v1 v2); atom z3 \<sharp> \<Gamma> \<rbrakk> \<Longrightarrow>
- \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> AE_concat v1 v2 \<Rightarrow> \<lbrace> z3 : B_bitvec | [[z3]\<^sup>v]\<^sup>c\<^sup>e == (CE_concat [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e) \<rbrace>"
-
-| infer_e_splitI: "\<lbrakk>
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> ;
- \<Theta> \<turnstile>\<^sub>w\<^sub>f (\<Phi>::\<Phi>);
- \<Theta> ; \<B> ; \<Gamma> \<turnstile> v1 \<Rightarrow> \<lbrace> z1 : B_bitvec | c1 \<rbrace> ;
- \<Theta> ; \<B> ; \<Gamma> \<turnstile> v2 \<Leftarrow> \<lbrace> z2 : B_int | (CE_op LEq (CE_val (V_lit (L_num 0))) (CE_val (V_var z2))) == (CE_val (V_lit L_true)) AND
- (CE_op LEq (CE_val (V_var z2)) (CE_len (CE_val (v1)))) == (CE_val (V_lit L_true)) \<rbrace>;
- atom z1 \<sharp> (AE_split v1 v2); atom z1 \<sharp> \<Gamma>;
- atom z2 \<sharp> (AE_split v1 v2); atom z2 \<sharp> \<Gamma>;
- atom z3 \<sharp> (AE_split v1 v2); atom z3 \<sharp> \<Gamma>
-\<rbrakk> \<Longrightarrow>
- \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> (AE_split v1 v2) \<Rightarrow> \<lbrace> z3 : B_pair B_bitvec B_bitvec |
- ((CE_val v1) == (CE_concat (CE_fst (CE_val (V_var z3))) (CE_snd (CE_val (V_var z3)))))
- AND (((CE_len (CE_fst (CE_val (V_var z3))))) == (CE_val ( v2))) \<rbrace>"
-
-equivariance infer_e
-nominal_inductive infer_e
- avoids infer_e_appI: x |infer_e_appPI: bv | infer_e_splitI: z3 and z1 and z2
-proof(goal_cases)
- case (1 \<Theta> \<B> \<Gamma> \<Delta> \<Phi> f x b c \<tau>' s' v \<tau>)
- moreover hence "atom x \<sharp> [ f v ]\<^sup>e" using fresh_prodN pure_fresh e.fresh by force
- ultimately show ?case unfolding fresh_star_def using fresh_prodN e.fresh pure_fresh by simp
-next
- case (2 \<Theta> \<B> \<Gamma> \<Delta> \<Phi> f x b c \<tau>' s' v \<tau>)
- then show ?case by auto
-next
- case (3 \<Theta> \<B> \<Gamma> \<Delta> \<Phi> b' f bv x b c \<tau>' s' v \<tau>)
- moreover hence "atom bv \<sharp> AE_appP f b' v" using fresh_prodN pure_fresh e.fresh by force
- ultimately show ?case unfolding fresh_star_def using fresh_prodN e.fresh pure_fresh fresh_Pair by auto
-next
- case (4 \<Theta> \<B> \<Gamma> \<Delta> \<Phi> b' f bv x b c \<tau>' s' v \<tau>)
- then show ?case by auto
-next
- case (5 \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v1 z1 c1 v2 z2 z3)
- have "atom z3 \<sharp> \<lbrace> z3 : [ B_bitvec , B_bitvec ]\<^sup>b | [ v1 ]\<^sup>c\<^sup>e == [ [#1[ [ z3 ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e @@ [#2[ [ z3 ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e AND [| [#1[ [ z3 ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e == [ v2 ]\<^sup>c\<^sup>e \<rbrace>"
- using \<tau>.fresh by simp
- then show ?case unfolding fresh_star_def fresh_prod7 using wfG_fresh_x2 5 by auto
-next
- case (6 \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v1 z1 c1 v2 z2 z3)
- then show ?case by auto
-qed
-
-inductive_cases infer_e_elims[elim!]:
- "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> (AE_op Plus v1 v2) \<Rightarrow> \<lbrace> z3 : B_int | [[z3]\<^sup>v]\<^sup>c\<^sup>e == (CE_op Plus [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e) \<rbrace>"
- "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> (AE_op LEq v1 v2) \<Rightarrow> \<lbrace> z3 : B_bool | [[z3]\<^sup>v]\<^sup>c\<^sup>e == (CE_op LEq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e) \<rbrace>"
- "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> (AE_op Plus v1 v2) \<Rightarrow> \<lbrace> z3 : B_int | c \<rbrace>"
- "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> (AE_op Plus v1 v2) \<Rightarrow> \<lbrace> z3 : b | c \<rbrace>"
- "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> (AE_op LEq v1 v2) \<Rightarrow> \<lbrace> z3 : b | c \<rbrace>"
- "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> (AE_app f v ) \<Rightarrow> \<tau>"
- "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> (AE_val v) \<Rightarrow> \<tau>"
- "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> (AE_fst v) \<Rightarrow> \<tau>"
- "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> (AE_snd v) \<Rightarrow> \<tau>"
- "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> (AE_mvar u) \<Rightarrow> \<tau>"
- "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> (AE_op Plus v1 v2) \<Rightarrow> \<tau>"
- "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> (AE_op LEq v1 v2) \<Rightarrow> \<tau>"
- "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> (AE_op LEq v1 v2) \<Rightarrow> \<lbrace> z3 : B_bool | c \<rbrace>"
- "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> (AE_app f v ) \<Rightarrow> \<tau>[x::=v]\<^sub>\<tau>\<^sub>v"
- "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> (AE_op opp v1 v2) \<Rightarrow> \<tau>"
- "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> (AE_len v) \<Rightarrow> \<tau>"
- "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> (AE_len v) \<Rightarrow> \<lbrace> z : B_int | [[z]\<^sup>v]\<^sup>c\<^sup>e == ((CE_len [v]\<^sup>c\<^sup>e))\<rbrace> "
- "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> AE_concat v1 v2 \<Rightarrow> \<tau>"
- "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> AE_concat v1 v2 \<Rightarrow> (\<lbrace> z : b | c \<rbrace>) "
- "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> AE_concat v1 v2 \<Rightarrow> (\<lbrace> z : B_bitvec | [[z]\<^sup>v]\<^sup>c\<^sup>e == (CE_concat [v1]\<^sup>c\<^sup>e [v1]\<^sup>c\<^sup>e) \<rbrace>) "
- "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> (AE_appP f b v ) \<Rightarrow> \<tau>"
- "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> AE_split v1 v2 \<Rightarrow> \<tau>"
- "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> (AE_op Eq v1 v2) \<Rightarrow> \<lbrace> z3 : b | c \<rbrace>"
- "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> (AE_op Eq v1 v2) \<Rightarrow> \<lbrace> z3 : B_bool | c \<rbrace>"
- "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> (AE_op Eq v1 v2) \<Rightarrow> \<tau>"
-nominal_termination (eqvt) by lexicographic_order
-
-section \<open>Statements\<close>
-
-inductive check_s :: "\<Theta> \<Rightarrow> \<Phi> \<Rightarrow> \<B> \<Rightarrow> \<Gamma> \<Rightarrow> \<Delta> \<Rightarrow> s \<Rightarrow> \<tau> \<Rightarrow> bool" (" _ ; _ ; _ ; _ ; _ \<turnstile> _ \<Leftarrow> _" [50, 50, 50,50,50] 50) and
- check_branch_s :: "\<Theta> \<Rightarrow> \<Phi> \<Rightarrow> \<B> \<Rightarrow> \<Gamma> \<Rightarrow> \<Delta> \<Rightarrow> tyid \<Rightarrow> string \<Rightarrow> \<tau> \<Rightarrow> v \<Rightarrow> branch_s \<Rightarrow> \<tau> \<Rightarrow> bool" (" _ ; _ ; _ ; _ ; _ ; _ ; _ ; _ ; _ \<turnstile> _ \<Leftarrow> _" [50, 50, 50,50,50] 50) and
- check_branch_list :: "\<Theta> \<Rightarrow> \<Phi> \<Rightarrow> \<B> \<Rightarrow> \<Gamma> \<Rightarrow> \<Delta> \<Rightarrow> tyid \<Rightarrow> (string * \<tau>) list \<Rightarrow> v \<Rightarrow> branch_list \<Rightarrow> \<tau> \<Rightarrow> bool" (" _ ; _ ; _ ; _ ; _ ; _ ; _ ; _ \<turnstile> _ \<Leftarrow> _" [50, 50, 50,50,50] 50) where
- check_valI: "\<lbrakk>
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> ;
- \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> ;
- \<Theta>; \<B>; \<Gamma> \<turnstile> v \<Rightarrow> \<tau>';
- \<Theta>; \<B>; \<Gamma> \<turnstile> \<tau>' \<lesssim> \<tau> \<rbrakk> \<Longrightarrow>
- \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> (AS_val v) \<Leftarrow> \<tau>"
-
-| check_letI: "\<lbrakk>
- atom x \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>, \<Delta>, e, \<tau>);
- atom z \<sharp> (x, \<Theta>, \<Phi>, \<B>, \<Gamma>, \<Delta>, e, \<tau>, s);
- \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> e \<Rightarrow> \<lbrace> z : b | c \<rbrace> ;
- \<Theta>; \<Phi> ; \<B> ; ((x,b,c[z::=V_var x]\<^sub>v)#\<^sub>\<Gamma>\<Gamma>) ; \<Delta> \<turnstile> s \<Leftarrow> \<tau>
-\<rbrakk> \<Longrightarrow>
- \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> (AS_let x e s) \<Leftarrow> \<tau>"
-
-| check_assertI: "\<lbrakk>
- atom x \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>, \<Delta>, c, \<tau>, s);
- \<Theta>; \<Phi> ; \<B> ; ((x,B_bool,c)#\<^sub>\<Gamma>\<Gamma>) ; \<Delta> \<turnstile> s \<Leftarrow> \<tau> ;
- \<Theta>; \<B>; \<Gamma> \<Turnstile> c;
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>
-\<rbrakk> \<Longrightarrow>
- \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> (AS_assert c s) \<Leftarrow> \<tau>"
-
-|check_branch_s_branchI: "\<lbrakk>
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> ;
- \<turnstile>\<^sub>w\<^sub>f \<Theta> ;
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau> ;
- \<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f const;
- atom x \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>, \<Delta>, tid, cons , const, v, \<tau>);
- \<Theta>; \<Phi>; \<B>; ((x,b_of const, ([v]\<^sup>c\<^sup>e == [ V_cons tid cons [x]\<^sup>v]\<^sup>c\<^sup>e ) AND (c_of const x))#\<^sub>\<Gamma>\<Gamma>) ; \<Delta> \<turnstile> s \<Leftarrow> \<tau>
-\<rbrakk> \<Longrightarrow>
- \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> ; tid ; cons ; const ; v \<turnstile> (AS_branch cons x s) \<Leftarrow> \<tau>"
-
-|check_branch_list_consI: "\<lbrakk>
- \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta>; tid; cons; const; v \<turnstile> cs \<Leftarrow> \<tau> ;
- \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta>; tid; dclist; v \<turnstile> css \<Leftarrow> \<tau>
-\<rbrakk> \<Longrightarrow>
- \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> ; tid ; (cons,const)#dclist ; v \<turnstile> AS_cons cs css \<Leftarrow> \<tau>"
-
-|check_branch_list_finalI: "\<lbrakk>
- \<Theta>; \<Phi>;\<B>; \<Gamma>; \<Delta>; tid ; cons ; const ; v \<turnstile> cs \<Leftarrow> \<tau>
-\<rbrakk> \<Longrightarrow>
- \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> ; tid ; [(cons,const)] ; v \<turnstile> AS_final cs \<Leftarrow> \<tau>"
-
-| check_ifI: "\<lbrakk>
- atom z \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>, \<Delta>, v , s1 , s2 , \<tau> );
- (\<Theta>; \<B>; \<Gamma> \<turnstile> v \<Leftarrow> (\<lbrace> z : B_bool | TRUE \<rbrace>)) ;
- \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> s1 \<Leftarrow> (\<lbrace> z : b_of \<tau> | ([v]\<^sup>c\<^sup>e == [[L_true]\<^sup>v]\<^sup>c\<^sup>e) IMP (c_of \<tau> z) \<rbrace>) ;
- \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> s2 \<Leftarrow> (\<lbrace> z : b_of \<tau> | ([v]\<^sup>c\<^sup>e == [[L_false]\<^sup>v]\<^sup>c\<^sup>e) IMP (c_of \<tau> z) \<rbrace>)
-\<rbrakk> \<Longrightarrow>
- \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> IF v THEN s1 ELSE s2 \<Leftarrow> \<tau>"
-
-| check_let2I: "\<lbrakk>
- atom x \<sharp> (\<Theta>, \<Phi>, \<B>, G, \<Delta>, t, s1, \<tau>) ;
- \<Theta>; \<Phi> ; \<B> ; G; \<Delta> \<turnstile> s1 \<Leftarrow> t;
- \<Theta>; \<Phi> ; \<B> ; ((x,b_of t,c_of t x)#\<^sub>\<Gamma>G) ; \<Delta> \<turnstile> s2 \<Leftarrow> \<tau>
-\<rbrakk> \<Longrightarrow>
- \<Theta>; \<Phi> ; \<B> ; G ; \<Delta> \<turnstile> (LET x : t = s1 IN s2) \<Leftarrow> \<tau>"
-
-| check_varI: "\<lbrakk>
- atom u \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>, \<Delta>, \<tau>', v, \<tau>) ;
- \<Theta>; \<B>; \<Gamma> \<turnstile> v \<Leftarrow> \<tau>';
- \<Theta>; \<Phi>; \<B>; \<Gamma> ; ((u,\<tau>') #\<^sub>\<Delta> \<Delta>) \<turnstile> s \<Leftarrow> \<tau>
-\<rbrakk> \<Longrightarrow>
- \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> (VAR u : \<tau>' = v IN s) \<Leftarrow> \<tau>"
-
-| check_assignI: "\<lbrakk>
- \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> ;
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> ;
- (u,\<tau>) \<in> setD \<Delta> ;
- \<Theta>; \<B>; \<Gamma> \<turnstile> v \<Leftarrow> \<tau>;
- \<Theta>; \<B>; \<Gamma> \<turnstile> (\<lbrace> z : B_unit | TRUE \<rbrace>) \<lesssim> \<tau>'
-\<rbrakk> \<Longrightarrow>
- \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> (u ::= v) \<Leftarrow> \<tau>'"
-
-| check_whileI: "\<lbrakk>
- \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> s1 \<Leftarrow> \<lbrace> z : B_bool | TRUE \<rbrace>;
- \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> s2 \<Leftarrow> \<lbrace> z : B_unit | TRUE \<rbrace>;
- \<Theta>; \<B>; \<Gamma> \<turnstile> (\<lbrace> z : B_unit | TRUE \<rbrace>) \<lesssim> \<tau>'
-\<rbrakk> \<Longrightarrow>
- \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> WHILE s1 DO { s2 } \<Leftarrow> \<tau>'"
-
-| check_seqI: "\<lbrakk>
- \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> s1 \<Leftarrow> \<lbrace> z : B_unit | TRUE \<rbrace>;
- \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> s2 \<Leftarrow> \<tau>
-\<rbrakk> \<Longrightarrow>
- \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> s1 ;; s2 \<Leftarrow> \<tau>"
-
-| check_caseI: "\<lbrakk>
- \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta>; tid ; dclist ; v \<turnstile> cs \<Leftarrow> \<tau> ;
- (AF_typedef tid dclist ) \<in> set \<Theta> ;
- \<Theta>; \<B>; \<Gamma> \<turnstile> v \<Leftarrow> \<lbrace> z : B_id tid | TRUE \<rbrace>;
- \<turnstile>\<^sub>w\<^sub>f \<Theta>
-\<rbrakk> \<Longrightarrow>
- \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> AS_match v cs \<Leftarrow> \<tau>"
-
-equivariance check_s
-
-text \<open>We only need avoidance for cases where a variable is added to a context\<close>
-nominal_inductive check_s
- avoids check_letI: x and z | check_branch_s_branchI: x | check_let2I: x | check_varI: u | check_ifI: z | check_assertI: x
-proof(goal_cases)
- case (1 x \<Theta> \<Phi> \<B> \<Gamma> \<Delta> e \<tau> z s b c)
- hence "atom x \<sharp> AS_let x e s" using s_branch_s_branch_list.fresh(2) by auto
- moreover have "atom z \<sharp> AS_let x e s" using s_branch_s_branch_list.fresh(2) 1 fresh_prod8 by auto
- then show ?case using fresh_star_def 1 by force
-next
- case (3 x \<Theta> \<Phi> \<B> \<Gamma> \<Delta> c \<tau> s)
- hence "atom x \<sharp> AS_assert c s" using fresh_prodN s_branch_s_branch_list.fresh pure_fresh by auto
- then show ?case using fresh_star_def 3 by force
-next
- case (5 \<Theta> \<B> \<Gamma> \<Delta> \<tau> const x \<Phi> tid cons v s)
- hence "atom x \<sharp> AS_branch cons x s" using fresh_prodN s_branch_s_branch_list.fresh pure_fresh by auto
- then show ?case using fresh_star_def 5 by force
-next
- case (7 z \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v s1 s2 \<tau>)
- hence "atom z \<sharp> AS_if v s1 s2 " using s_branch_s_branch_list.fresh by auto
- then show ?case using 7 fresh_prodN fresh_star_def by fastforce
-next
- case (9 x \<Theta> \<Phi> \<B> G \<Delta> t s1 \<tau> s2)
- hence "atom x \<sharp> AS_let2 x t s1 s2" using s_branch_s_branch_list.fresh by auto
- thus ?case using fresh_star_def 9 by force
-next
- case (11 u \<Theta> \<Phi> \<B> \<Gamma> \<Delta> \<tau>' v \<tau> s)
- hence "atom u \<sharp> AS_var u \<tau>' v s" using s_branch_s_branch_list.fresh by auto
- then show ?case using fresh_star_def 11 by force
-
-qed(auto+)
-
-inductive_cases check_s_elims[elim!]:
- "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> AS_val v \<Leftarrow> \<tau>"
- "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> AS_let x e s \<Leftarrow> \<tau>"
- "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> AS_if v s1 s2 \<Leftarrow> \<tau>"
- "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> AS_let2 x t s1 s2 \<Leftarrow> \<tau>"
- "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> AS_while s1 s2 \<Leftarrow> \<tau>"
- "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> AS_var u t v s \<Leftarrow> \<tau>"
- "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> AS_seq s1 s2 \<Leftarrow> \<tau>"
- "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> AS_assign u v \<Leftarrow> \<tau>"
- "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> AS_match v cs \<Leftarrow> \<tau>"
- "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> AS_assert c s \<Leftarrow> \<tau>"
-
-inductive_cases check_branch_s_elims[elim!]:
- "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta>; tid ; dclist ; v \<turnstile> (AS_final cs) \<Leftarrow> \<tau>"
- "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta>; tid ; dclist ; v \<turnstile> (AS_cons cs css) \<Leftarrow> \<tau>"
- "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta>; tid ; cons ; const ; v \<turnstile> (AS_branch dc x s ) \<Leftarrow> \<tau>"
-
-section \<open>Programs\<close>
-
-text \<open>Type check function bodies\<close>
-
-inductive check_funtyp :: "\<Theta> \<Rightarrow> \<Phi> \<Rightarrow> \<B> \<Rightarrow> fun_typ \<Rightarrow> bool" ( " _ ; _ ; _ \<turnstile> _ " ) where
- check_funtypI: "\<lbrakk>
- atom x \<sharp> (\<Theta>, \<Phi>, B , b );
- \<Theta>; \<Phi> ; B ; ((x,b,c) #\<^sub>\<Gamma> GNil) ; []\<^sub>\<Delta> \<turnstile> s \<Leftarrow> \<tau>
-\<rbrakk> \<Longrightarrow>
- \<Theta>; \<Phi> ; B \<turnstile> (AF_fun_typ x b c \<tau> s)"
-
-equivariance check_funtyp
-nominal_inductive check_funtyp
- avoids check_funtypI: x
-proof(goal_cases)
- case (1 x \<Theta> \<Phi> B b c s \<tau> )
- hence "atom x \<sharp> (AF_fun_typ x b c \<tau> s)" using fun_def.fresh fun_typ_q.fresh fun_typ.fresh by simp
- then show ?case using fresh_star_def 1 fresh_prodN by fastforce
-next
- case (2 \<Theta> \<Phi> x b c s \<tau> f)
- then show ?case by auto
-qed
-
-inductive check_funtypq :: "\<Theta> \<Rightarrow> \<Phi> \<Rightarrow> fun_typ_q \<Rightarrow> bool" ( " _ ; _ \<turnstile> _ " ) where
- check_fundefq_simpleI: "\<lbrakk>
- \<Theta>; \<Phi> ; {||} \<turnstile> (AF_fun_typ x b c t s)
-\<rbrakk> \<Longrightarrow>
- \<Theta>; \<Phi> \<turnstile> ((AF_fun_typ_none (AF_fun_typ x b c t s)))"
-
-|check_funtypq_polyI: "\<lbrakk>
- atom bv \<sharp> (\<Theta>, \<Phi>, (AF_fun_typ x b c t s));
- \<Theta>; \<Phi>; {|bv|} \<turnstile> (AF_fun_typ x b c t s)
-\<rbrakk> \<Longrightarrow>
- \<Theta>; \<Phi> \<turnstile> (AF_fun_typ_some bv (AF_fun_typ x b c t s))"
-
-equivariance check_funtypq
-nominal_inductive check_funtypq
- avoids check_funtypq_polyI: bv
-proof(goal_cases)
- case (1 bv \<Theta> \<Phi> x b c t s )
- hence "atom bv \<sharp> (AF_fun_typ_some bv (AF_fun_typ x b c t s))" using fun_def.fresh fun_typ_q.fresh fun_typ.fresh by simp
- thus ?case using fresh_star_def 1 fresh_prodN by fastforce
-next
- case (2 bv \<Theta> \<Phi> ft )
- then show ?case by auto
-qed
-
-inductive check_fundef :: "\<Theta> \<Rightarrow> \<Phi> \<Rightarrow> fun_def \<Rightarrow> bool" ( " _ ; _ \<turnstile> _ " ) where
- check_fundefI: "\<lbrakk>
- \<Theta>; \<Phi> \<turnstile> ft
-\<rbrakk> \<Longrightarrow>
- \<Theta>; \<Phi> \<turnstile> (AF_fundef f ft)"
-
-equivariance check_fundef
-nominal_inductive check_fundef .
-
-text \<open>Temporarily remove this simproc as it produces untidy eliminations\<close>
-declare[[ simproc del: alpha_lst]]
-
-inductive_cases check_funtyp_elims[elim!]:
- "check_funtyp \<Theta> \<Phi> B ft"
-
-inductive_cases check_funtypq_elims[elim!]:
- "check_funtypq \<Theta> \<Phi> (AF_fun_typ_none (AF_fun_typ x b c \<tau> s))"
- "check_funtypq \<Theta> \<Phi> (AF_fun_typ_some bv (AF_fun_typ x b c \<tau> s))"
-
-inductive_cases check_fundef_elims[elim!]:
- "check_fundef \<Theta> \<Phi> (AF_fundef f ftq)"
-
-declare[[ simproc add: alpha_lst]]
-
-nominal_function \<Delta>_of :: "var_def list \<Rightarrow> \<Delta>" where
- "\<Delta>_of [] = DNil"
-| "\<Delta>_of ((AV_def u t v)#vs) = (u,t) #\<^sub>\<Delta> (\<Delta>_of vs)"
- apply auto
- using eqvt_def \<Delta>_of_graph_aux_def neq_Nil_conv old.prod.exhaust apply force
- using eqvt_def \<Delta>_of_graph_aux_def neq_Nil_conv old.prod.exhaust
- by (metis var_def.strong_exhaust)
-nominal_termination (eqvt) by lexicographic_order
-
-inductive check_prog :: "p \<Rightarrow> \<tau> \<Rightarrow> bool" ( "\<turnstile> _ \<Leftarrow> _") where
- "\<lbrakk>
- \<Theta>; \<Phi>; {||}; GNil ; \<Delta>_of \<G> \<turnstile> s \<Leftarrow> \<tau>
-\<rbrakk> \<Longrightarrow> \<turnstile> (AP_prog \<Theta> \<Phi> \<G> s) \<Leftarrow> \<tau>"
-
-inductive_cases check_prog_elims[elim!]:
- "\<turnstile> (AP_prog \<Theta> \<Phi> \<G> s) \<Leftarrow> \<tau>"
-
+(*<*)
+theory Typing
+ imports RCLogic WellformedL
+begin
+ (*>*)
+
+chapter \<open>Type System\<close>
+
+text \<open>The MiniSail type system. We define subtyping judgement first and then typing judgement
+for the term forms\<close>
+
+section \<open>Subtyping\<close>
+
+text \<open> Subtyping is defined on top of refinement constraint logic (RCL).
+A subtyping check is converted into an RCL validity check. \<close>
+
+inductive subtype :: "\<Theta> \<Rightarrow> \<B> \<Rightarrow> \<Gamma> \<Rightarrow> \<tau> \<Rightarrow> \<tau> \<Rightarrow> bool" ("_ ; _ ; _ \<turnstile> _ \<lesssim> _" [50, 50, 50] 50) where
+ subtype_baseI: "\<lbrakk>
+ atom x \<sharp> (\<Theta>, \<B>, \<Gamma>, z,c,z',c') ;
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b | c \<rbrace>;
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> z' : b | c' \<rbrace>;
+ \<Theta>; \<B> ; (x,b, c[z::=[x]\<^sup>v]\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<Turnstile> c'[z'::=[x]\<^sup>v]\<^sub>v
+\<rbrakk> \<Longrightarrow>
+ \<Theta>; \<B>; \<Gamma> \<turnstile> \<lbrace> z : b | c \<rbrace> \<lesssim> \<lbrace> z' : b | c' \<rbrace>"
+
+equivariance subtype
+nominal_inductive subtype
+ avoids subtype_baseI: x
+proof(goal_cases)
+ case (1 \<Theta> \<B> \<Gamma> z b c z' c' x)
+ then show ?case using fresh_star_def 1 by force
+next
+ case (2 \<Theta> \<B> \<Gamma> z b c z' c' x)
+ then show ?case by auto
+qed
+
+inductive_cases subtype_elims:
+ "\<Theta>; \<B>; \<Gamma> \<turnstile> \<lbrace> z : b | c \<rbrace> \<lesssim> \<lbrace> z' : b | c' \<rbrace>"
+ "\<Theta>; \<B>; \<Gamma> \<turnstile> \<tau>\<^sub>1 \<lesssim> \<tau>\<^sub>2"
+
+section \<open>Literals\<close>
+
+text \<open>The type synthesised has the constraint that z equates to the literal\<close>
+
+inductive infer_l :: "l \<Rightarrow> \<tau> \<Rightarrow> bool" (" \<turnstile> _ \<Rightarrow> _" [50, 50] 50) where
+ infer_trueI: " \<turnstile> L_true \<Rightarrow> \<lbrace> z : B_bool | [[z]\<^sup>v]\<^sup>c\<^sup>e == [[L_true]\<^sup>v]\<^sup>c\<^sup>e \<rbrace>"
+| infer_falseI: " \<turnstile> L_false \<Rightarrow> \<lbrace> z : B_bool | [[z]\<^sup>v]\<^sup>c\<^sup>e == [[L_false]\<^sup>v]\<^sup>c\<^sup>e \<rbrace>"
+| infer_natI: " \<turnstile> L_num n \<Rightarrow> \<lbrace> z : B_int | [[z]\<^sup>v]\<^sup>c\<^sup>e == [[L_num n]\<^sup>v]\<^sup>c\<^sup>e \<rbrace>"
+| infer_unitI: " \<turnstile> L_unit \<Rightarrow> \<lbrace> z : B_unit | [[z]\<^sup>v]\<^sup>c\<^sup>e == [[L_unit]\<^sup>v]\<^sup>c\<^sup>e \<rbrace>"
+| infer_bitvecI: " \<turnstile> L_bitvec bv \<Rightarrow> \<lbrace> z : B_bitvec | [[z]\<^sup>v]\<^sup>c\<^sup>e == [[L_bitvec bv]\<^sup>v]\<^sup>c\<^sup>e \<rbrace>"
+
+nominal_inductive infer_l .
+equivariance infer_l
+
+inductive_cases infer_l_elims[elim!]:
+ "\<turnstile> L_true \<Rightarrow> \<tau>"
+ "\<turnstile> L_false \<Rightarrow> \<tau>"
+ "\<turnstile> L_num n \<Rightarrow> \<tau>"
+ "\<turnstile> L_unit \<Rightarrow> \<tau>"
+ "\<turnstile> L_bitvec x \<Rightarrow> \<tau>"
+ "\<turnstile> l \<Rightarrow> \<tau>"
+
+lemma infer_l_form2[simp]:
+ shows "\<exists>z. \<turnstile> l \<Rightarrow> (\<lbrace> z : base_for_lit l | [[z]\<^sup>v]\<^sup>c\<^sup>e == [[l]\<^sup>v]\<^sup>c\<^sup>e \<rbrace>)"
+proof (nominal_induct l rule: l.strong_induct)
+ case (L_num x)
+ then show ?case using infer_l.intros base_for_lit.simps has_fresh_z by metis
+next
+ case L_true
+ then show ?case using infer_l.intros base_for_lit.simps has_fresh_z by metis
+next
+ case L_false
+ then show ?case using infer_l.intros base_for_lit.simps has_fresh_z by metis
+next
+ case L_unit
+ then show ?case using infer_l.intros base_for_lit.simps has_fresh_z by metis
+next
+ case (L_bitvec x)
+ then show ?case using infer_l.intros base_for_lit.simps has_fresh_z by metis
+qed
+
+section \<open>Values\<close>
+
+inductive infer_v :: "\<Theta> \<Rightarrow> \<B> \<Rightarrow> \<Gamma> \<Rightarrow> v \<Rightarrow> \<tau> \<Rightarrow> bool" (" _ ; _ ; _ \<turnstile> _ \<Rightarrow> _" [50, 50, 50] 50) where
+
+infer_v_varI: "\<lbrakk>
+ \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> ;
+ Some (b,c) = lookup \<Gamma> x;
+ atom z \<sharp> x ; atom z \<sharp> (\<Theta>, \<B>, \<Gamma>)
+\<rbrakk> \<Longrightarrow>
+ \<Theta>; \<B>; \<Gamma> \<turnstile> [x]\<^sup>v \<Rightarrow> \<lbrace> z : b | [[z]\<^sup>v]\<^sup>c\<^sup>e == [[x]\<^sup>v]\<^sup>c\<^sup>e \<rbrace>"
+
+| infer_v_litI: "\<lbrakk>
+ \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> ;
+ \<turnstile> l \<Rightarrow> \<tau>
+\<rbrakk> \<Longrightarrow>
+ \<Theta>; \<B>; \<Gamma> \<turnstile> [l]\<^sup>v \<Rightarrow> \<tau>"
+
+| infer_v_pairI: "\<lbrakk>
+ atom z \<sharp> (v1, v2); atom z \<sharp> (\<Theta>, \<B>, \<Gamma>) ;
+ \<Theta>; \<B>; \<Gamma> \<turnstile> (v1::v) \<Rightarrow> t1 ;
+ \<Theta>; \<B> ; \<Gamma> \<turnstile> (v2::v) \<Rightarrow> t2
+\<rbrakk> \<Longrightarrow>
+ \<Theta>; \<B>; \<Gamma> \<turnstile> V_pair v1 v2 \<Rightarrow> (\<lbrace> z : B_pair (b_of t1) (b_of t2) | [[z]\<^sup>v]\<^sup>c\<^sup>e == [[v1,v2]\<^sup>v]\<^sup>c\<^sup>e \<rbrace>) "
+
+| infer_v_consI: "\<lbrakk>
+ AF_typedef s dclist \<in> set \<Theta>;
+ (dc, tc) \<in> set dclist ;
+ \<Theta>; \<B>; \<Gamma> \<turnstile> v \<Rightarrow> tv ;
+ \<Theta>; \<B>; \<Gamma> \<turnstile> tv \<lesssim> tc ;
+ atom z \<sharp> v ; atom z \<sharp> (\<Theta>, \<B>, \<Gamma>)
+\<rbrakk> \<Longrightarrow>
+ \<Theta>; \<B>; \<Gamma> \<turnstile> V_cons s dc v \<Rightarrow> (\<lbrace> z : B_id s | [[z]\<^sup>v]\<^sup>c\<^sup>e == [ V_cons s dc v ]\<^sup>c\<^sup>e \<rbrace>)"
+
+| infer_v_conspI: "\<lbrakk>
+ AF_typedef_poly s bv dclist \<in> set \<Theta>;
+ (dc, tc) \<in> set dclist ;
+ \<Theta>; \<B>; \<Gamma> \<turnstile> v \<Rightarrow> tv;
+ \<Theta>; \<B>; \<Gamma> \<turnstile> tv \<lesssim> tc[bv::=b]\<^sub>\<tau>\<^sub>b ;
+ atom z \<sharp> (\<Theta>, \<B>, \<Gamma>, v, b);
+ atom bv \<sharp> (\<Theta>, \<B>, \<Gamma>, v, b);
+ \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b
+\<rbrakk> \<Longrightarrow>
+ \<Theta>; \<B>; \<Gamma> \<turnstile> V_consp s dc b v \<Rightarrow> (\<lbrace> z : B_app s b | [[z]\<^sup>v]\<^sup>c\<^sup>e == (CE_val (V_consp s dc b v)) \<rbrace>)"
+
+equivariance infer_v
+nominal_inductive infer_v
+ avoids infer_v_conspI: bv and z | infer_v_varI: z | infer_v_pairI: z | infer_v_consI: z
+proof(goal_cases)
+ case (1 \<Theta> \<B> \<Gamma> b c x z)
+ hence "atom z \<sharp> \<lbrace> z : b | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrace>" using \<tau>.fresh by simp
+ then show ?case unfolding fresh_star_def using 1 by simp
+next
+ case (2 \<Theta> \<B> \<Gamma> b c x z)
+ then show ?case by auto
+next
+ case (3 z v1 v2 \<Theta> \<B> \<Gamma> t1 t2)
+ hence "atom z \<sharp> \<lbrace> z : [ b_of t1 , b_of t2 ]\<^sup>b | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ v1 , v2 ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrace>" using \<tau>.fresh by simp
+ then show ?case unfolding fresh_star_def using 3 by simp
+next
+ case (4 z v1 v2 \<Theta> \<B> \<Gamma> t1 t2)
+ then show ?case by auto
+next
+ case (5 s dclist \<Theta> dc tc \<B> \<Gamma> v tv z)
+ hence "atom z \<sharp> \<lbrace> z : B_id s | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ V_cons s dc v ]\<^sup>c\<^sup>e \<rbrace>" using \<tau>.fresh b.fresh pure_fresh by auto
+ moreover have "atom z \<sharp> V_cons s dc v" using v.fresh 5 using v.fresh fresh_prodN pure_fresh by metis
+ then show ?case unfolding fresh_star_def using 5 by simp
+next
+ case (6 s dclist \<Theta> dc tc \<B> \<Gamma> v tv z)
+ then show ?case by auto
+next
+ case (7 s bv dclist \<Theta> dc tc \<B> \<Gamma> v tv b z)
+ hence "atom bv \<sharp> V_consp s dc b v" using v.fresh fresh_prodN pure_fresh by metis
+ moreover then have "atom bv \<sharp> \<lbrace> z : B_id s | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ V_consp s dc b v ]\<^sup>c\<^sup>e \<rbrace>"
+ using \<tau>.fresh ce.fresh v.fresh by auto
+ moreover have "atom z \<sharp> V_consp s dc b v" using v.fresh fresh_prodN pure_fresh 7 by metis
+ moreover then have "atom z \<sharp> \<lbrace> z : B_id s | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ V_consp s dc b v ]\<^sup>c\<^sup>e \<rbrace>"
+ using \<tau>.fresh ce.fresh v.fresh by auto
+ ultimately show ?case using fresh_star_def 7 by force
+next
+ case (8 s bv dclist \<Theta> dc tc \<B> \<Gamma> v tv b z)
+ then show ?case by auto
+qed
+
+inductive_cases infer_v_elims[elim!]:
+ "\<Theta>; \<B>; \<Gamma> \<turnstile> V_var x \<Rightarrow> \<tau>"
+ "\<Theta>; \<B>; \<Gamma> \<turnstile> V_lit l \<Rightarrow> \<tau>"
+ "\<Theta>; \<B>; \<Gamma> \<turnstile> V_pair v1 v2 \<Rightarrow> \<tau>"
+ "\<Theta>; \<B>; \<Gamma> \<turnstile> V_cons s dc v \<Rightarrow> \<tau>"
+ "\<Theta>; \<B>; \<Gamma> \<turnstile> V_pair v1 v2 \<Rightarrow> (\<lbrace> z : b | c \<rbrace>) "
+ "\<Theta>; \<B>; \<Gamma> \<turnstile> V_pair v1 v2 \<Rightarrow> (\<lbrace> z : [ b1 , b2 ]\<^sup>b | [[z]\<^sup>v]\<^sup>c\<^sup>e == [[v1,v2]\<^sup>v]\<^sup>c\<^sup>e \<rbrace>) "
+ "\<Theta>; \<B>; \<Gamma> \<turnstile> V_consp s dc b v \<Rightarrow> \<tau> "
+
+inductive check_v :: "\<Theta> \<Rightarrow> \<B> \<Rightarrow> \<Gamma> \<Rightarrow> v \<Rightarrow> \<tau> \<Rightarrow> bool" ("_ ; _ ; _ \<turnstile> _ \<Leftarrow> _" [50, 50, 50] 50) where
+ check_v_subtypeI: "\<lbrakk> \<Theta>; \<B>; \<Gamma> \<turnstile> \<tau>1 \<lesssim> \<tau>2; \<Theta>; \<B>; \<Gamma> \<turnstile> v \<Rightarrow> \<tau>1 \<rbrakk> \<Longrightarrow> \<Theta>; \<B> ; \<Gamma> \<turnstile> v \<Leftarrow> \<tau>2"
+equivariance check_v
+nominal_inductive check_v .
+
+inductive_cases check_v_elims[elim!]:
+ "\<Theta>; \<B> ; \<Gamma> \<turnstile> v \<Leftarrow> \<tau>"
+
+section \<open>Expressions\<close>
+
+text \<open> Type synthesis for expressions \<close>
+
+inductive infer_e :: "\<Theta> \<Rightarrow> \<Phi> \<Rightarrow> \<B> \<Rightarrow> \<Gamma> \<Rightarrow> \<Delta> \<Rightarrow> e \<Rightarrow> \<tau> \<Rightarrow> bool" ("_ ; _ ; _ ; _ ; _ \<turnstile> _ \<Rightarrow> _" [50, 50, 50,50] 50) where
+
+infer_e_valI: "\<lbrakk>
+ (\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>) ;
+ (\<Theta> \<turnstile>\<^sub>w\<^sub>f (\<Phi>::\<Phi>)) ;
+ (\<Theta>; \<B>; \<Gamma> \<turnstile> v \<Rightarrow> \<tau>) \<rbrakk> \<Longrightarrow>
+ \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> (AE_val v) \<Rightarrow> \<tau>"
+
+| infer_e_plusI: "\<lbrakk>
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> ;
+ \<Theta> \<turnstile>\<^sub>w\<^sub>f (\<Phi>::\<Phi>) ;
+ \<Theta>; \<B>; \<Gamma> \<turnstile> v1 \<Rightarrow> \<lbrace> z1 : B_int | c1 \<rbrace> ;
+ \<Theta>; \<B>; \<Gamma> \<turnstile> v2 \<Rightarrow> \<lbrace> z2 : B_int | c2 \<rbrace>;
+ atom z3 \<sharp> (AE_op Plus v1 v2); atom z3 \<sharp> \<Gamma> \<rbrakk> \<Longrightarrow>
+ \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> AE_op Plus v1 v2 \<Rightarrow> \<lbrace> z3 : B_int | [[z3]\<^sup>v]\<^sup>c\<^sup>e == (CE_op Plus [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e) \<rbrace>"
+
+| infer_e_leqI: "\<lbrakk>
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>;
+ \<Theta> \<turnstile>\<^sub>w\<^sub>f (\<Phi>::\<Phi>) ;
+ \<Theta>; \<B>; \<Gamma> \<turnstile> v1 \<Rightarrow> \<lbrace> z1 : B_int | c1 \<rbrace> ;
+ \<Theta>; \<B>; \<Gamma> \<turnstile> v2 \<Rightarrow> \<lbrace> z2 : B_int | c2 \<rbrace>;
+ atom z3 \<sharp> (AE_op LEq v1 v2); atom z3 \<sharp> \<Gamma>
+\<rbrakk> \<Longrightarrow>
+ \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> AE_op LEq v1 v2 \<Rightarrow> \<lbrace> z3 : B_bool | [[z3]\<^sup>v]\<^sup>c\<^sup>e == (CE_op LEq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e) \<rbrace>"
+
+| infer_e_eqI: "\<lbrakk>
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>;
+ \<Theta> \<turnstile>\<^sub>w\<^sub>f (\<Phi>::\<Phi>) ;
+ \<Theta>; \<B>; \<Gamma> \<turnstile> v1 \<Rightarrow> \<lbrace> z1 : b | c1 \<rbrace> ;
+ \<Theta>; \<B>; \<Gamma> \<turnstile> v2 \<Rightarrow> \<lbrace> z2 : b | c2 \<rbrace>;
+ atom z3 \<sharp> (AE_op Eq v1 v2); atom z3 \<sharp> \<Gamma> ;
+ b \<in> { B_bool, B_int, B_unit }
+\<rbrakk> \<Longrightarrow>
+ \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> AE_op Eq v1 v2 \<Rightarrow> \<lbrace> z3 : B_bool | [[z3]\<^sup>v]\<^sup>c\<^sup>e == (CE_op Eq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e) \<rbrace>"
+
+| infer_e_appI: "\<lbrakk>
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> ;
+ \<Theta> \<turnstile>\<^sub>w\<^sub>f (\<Phi>::\<Phi>) ;
+ Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \<tau>' s'))) = lookup_fun \<Phi> f;
+ \<Theta>; \<B>; \<Gamma> \<turnstile> v \<Leftarrow> \<lbrace> x : b | c \<rbrace>;
+ atom x \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>, \<Delta>,v , \<tau>);
+ \<tau>'[x::=v]\<^sub>v = \<tau>
+\<rbrakk> \<Longrightarrow>
+ \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> AE_app f v \<Rightarrow> \<tau>"
+
+| infer_e_appPI: "\<lbrakk>
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> ;
+ \<Theta> \<turnstile>\<^sub>w\<^sub>f (\<Phi>::\<Phi>) ;
+ \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b' ;
+ Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \<tau>' s'))) = lookup_fun \<Phi> f;
+ \<Theta>; \<B>; \<Gamma> \<turnstile> v \<Leftarrow> \<lbrace> x : b[bv::=b']\<^sub>b | c[bv::=b']\<^sub>b \<rbrace>; atom x \<sharp> \<Gamma>;
+ (\<tau>'[bv::=b']\<^sub>b[x::=v]\<^sub>v) = \<tau> ;
+ atom bv \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>, \<Delta>, b', v, \<tau>)
+\<rbrakk> \<Longrightarrow>
+ \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> AE_appP f b' v \<Rightarrow> \<tau>"
+
+| infer_e_fstI: "\<lbrakk>
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> ;
+ \<Theta> \<turnstile>\<^sub>w\<^sub>f (\<Phi>::\<Phi>) ;
+ \<Theta>; \<B>; \<Gamma> \<turnstile> v \<Rightarrow> \<lbrace> z' : [b1,b2]\<^sup>b | c \<rbrace>;
+ atom z \<sharp> AE_fst v ; atom z \<sharp> \<Gamma> \<rbrakk> \<Longrightarrow>
+ \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> AE_fst v \<Rightarrow> \<lbrace> z : b1 | [[z]\<^sup>v]\<^sup>c\<^sup>e == ((CE_fst [v]\<^sup>c\<^sup>e)) \<rbrace>"
+
+| infer_e_sndI: "\<lbrakk>
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> ;
+ \<Theta> \<turnstile>\<^sub>w\<^sub>f (\<Phi>::\<Phi>) ;
+ \<Theta>; \<B>; \<Gamma> \<turnstile> v \<Rightarrow> \<lbrace> z' : [ b1, b2]\<^sup>b | c \<rbrace>;
+ atom z \<sharp> AE_snd v ; atom z \<sharp> \<Gamma> \<rbrakk> \<Longrightarrow>
+ \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> AE_snd v \<Rightarrow> \<lbrace> z : b2 | [[z]\<^sup>v]\<^sup>c\<^sup>e == ((CE_snd [v]\<^sup>c\<^sup>e)) \<rbrace>"
+
+| infer_e_lenI: "\<lbrakk>
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> ;
+ \<Theta> \<turnstile>\<^sub>w\<^sub>f (\<Phi>::\<Phi>) ;
+ \<Theta>; \<B>; \<Gamma> \<turnstile> v \<Rightarrow> \<lbrace> z' : B_bitvec | c \<rbrace>;
+ atom z \<sharp> AE_len v ; atom z \<sharp> \<Gamma>\<rbrakk> \<Longrightarrow>
+ \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> AE_len v \<Rightarrow> \<lbrace> z : B_int | [[z]\<^sup>v]\<^sup>c\<^sup>e == ((CE_len [v]\<^sup>c\<^sup>e)) \<rbrace>"
+
+| infer_e_mvarI: "\<lbrakk>
+ \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> ;
+ \<Theta> \<turnstile>\<^sub>w\<^sub>f (\<Phi>::\<Phi>) ;
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>;
+ (u,\<tau>) \<in> setD \<Delta> \<rbrakk> \<Longrightarrow>
+ \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> AE_mvar u \<Rightarrow> \<tau>"
+
+| infer_e_concatI: "\<lbrakk>
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> ;
+ \<Theta> \<turnstile>\<^sub>w\<^sub>f (\<Phi>::\<Phi>) ;
+ \<Theta>; \<B>; \<Gamma> \<turnstile> v1 \<Rightarrow> \<lbrace> z1 : B_bitvec | c1 \<rbrace> ;
+ \<Theta>; \<B>; \<Gamma> \<turnstile> v2 \<Rightarrow> \<lbrace> z2 : B_bitvec | c2 \<rbrace>;
+ atom z3 \<sharp> (AE_concat v1 v2); atom z3 \<sharp> \<Gamma> \<rbrakk> \<Longrightarrow>
+ \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> AE_concat v1 v2 \<Rightarrow> \<lbrace> z3 : B_bitvec | [[z3]\<^sup>v]\<^sup>c\<^sup>e == (CE_concat [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e) \<rbrace>"
+
+| infer_e_splitI: "\<lbrakk>
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> ;
+ \<Theta> \<turnstile>\<^sub>w\<^sub>f (\<Phi>::\<Phi>);
+ \<Theta> ; \<B> ; \<Gamma> \<turnstile> v1 \<Rightarrow> \<lbrace> z1 : B_bitvec | c1 \<rbrace> ;
+ \<Theta> ; \<B> ; \<Gamma> \<turnstile> v2 \<Leftarrow> \<lbrace> z2 : B_int | (CE_op LEq (CE_val (V_lit (L_num 0))) (CE_val (V_var z2))) == (CE_val (V_lit L_true)) AND
+ (CE_op LEq (CE_val (V_var z2)) (CE_len (CE_val (v1)))) == (CE_val (V_lit L_true)) \<rbrace>;
+ atom z1 \<sharp> (AE_split v1 v2); atom z1 \<sharp> \<Gamma>;
+ atom z2 \<sharp> (AE_split v1 v2); atom z2 \<sharp> \<Gamma>;
+ atom z3 \<sharp> (AE_split v1 v2); atom z3 \<sharp> \<Gamma>
+\<rbrakk> \<Longrightarrow>
+ \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> (AE_split v1 v2) \<Rightarrow> \<lbrace> z3 : B_pair B_bitvec B_bitvec |
+ ((CE_val v1) == (CE_concat (CE_fst (CE_val (V_var z3))) (CE_snd (CE_val (V_var z3)))))
+ AND (((CE_len (CE_fst (CE_val (V_var z3))))) == (CE_val ( v2))) \<rbrace>"
+
+equivariance infer_e
+nominal_inductive infer_e
+ avoids infer_e_appI: x |infer_e_appPI: bv | infer_e_splitI: z3 and z1 and z2
+proof(goal_cases)
+ case (1 \<Theta> \<B> \<Gamma> \<Delta> \<Phi> f x b c \<tau>' s' v \<tau>)
+ moreover hence "atom x \<sharp> [ f v ]\<^sup>e" using fresh_prodN pure_fresh e.fresh by force
+ ultimately show ?case unfolding fresh_star_def using fresh_prodN e.fresh pure_fresh by simp
+next
+ case (2 \<Theta> \<B> \<Gamma> \<Delta> \<Phi> f x b c \<tau>' s' v \<tau>)
+ then show ?case by auto
+next
+ case (3 \<Theta> \<B> \<Gamma> \<Delta> \<Phi> b' f bv x b c \<tau>' s' v \<tau>)
+ moreover hence "atom bv \<sharp> AE_appP f b' v" using fresh_prodN pure_fresh e.fresh by force
+ ultimately show ?case unfolding fresh_star_def using fresh_prodN e.fresh pure_fresh fresh_Pair by auto
+next
+ case (4 \<Theta> \<B> \<Gamma> \<Delta> \<Phi> b' f bv x b c \<tau>' s' v \<tau>)
+ then show ?case by auto
+next
+ case (5 \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v1 z1 c1 v2 z2 z3)
+ have "atom z3 \<sharp> \<lbrace> z3 : [ B_bitvec , B_bitvec ]\<^sup>b | [ v1 ]\<^sup>c\<^sup>e == [ [#1[ [ z3 ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e @@ [#2[ [ z3 ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e AND [| [#1[ [ z3 ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e == [ v2 ]\<^sup>c\<^sup>e \<rbrace>"
+ using \<tau>.fresh by simp
+ then show ?case unfolding fresh_star_def fresh_prod7 using wfG_fresh_x2 5 by auto
+next
+ case (6 \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v1 z1 c1 v2 z2 z3)
+ then show ?case by auto
+qed
+
+inductive_cases infer_e_elims[elim!]:
+ "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> (AE_op Plus v1 v2) \<Rightarrow> \<lbrace> z3 : B_int | [[z3]\<^sup>v]\<^sup>c\<^sup>e == (CE_op Plus [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e) \<rbrace>"
+ "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> (AE_op LEq v1 v2) \<Rightarrow> \<lbrace> z3 : B_bool | [[z3]\<^sup>v]\<^sup>c\<^sup>e == (CE_op LEq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e) \<rbrace>"
+ "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> (AE_op Plus v1 v2) \<Rightarrow> \<lbrace> z3 : B_int | c \<rbrace>"
+ "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> (AE_op Plus v1 v2) \<Rightarrow> \<lbrace> z3 : b | c \<rbrace>"
+ "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> (AE_op LEq v1 v2) \<Rightarrow> \<lbrace> z3 : b | c \<rbrace>"
+ "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> (AE_app f v ) \<Rightarrow> \<tau>"
+ "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> (AE_val v) \<Rightarrow> \<tau>"
+ "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> (AE_fst v) \<Rightarrow> \<tau>"
+ "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> (AE_snd v) \<Rightarrow> \<tau>"
+ "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> (AE_mvar u) \<Rightarrow> \<tau>"
+ "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> (AE_op Plus v1 v2) \<Rightarrow> \<tau>"
+ "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> (AE_op LEq v1 v2) \<Rightarrow> \<tau>"
+ "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> (AE_op LEq v1 v2) \<Rightarrow> \<lbrace> z3 : B_bool | c \<rbrace>"
+ "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> (AE_app f v ) \<Rightarrow> \<tau>[x::=v]\<^sub>\<tau>\<^sub>v"
+ "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> (AE_op opp v1 v2) \<Rightarrow> \<tau>"
+ "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> (AE_len v) \<Rightarrow> \<tau>"
+ "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> (AE_len v) \<Rightarrow> \<lbrace> z : B_int | [[z]\<^sup>v]\<^sup>c\<^sup>e == ((CE_len [v]\<^sup>c\<^sup>e))\<rbrace> "
+ "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> AE_concat v1 v2 \<Rightarrow> \<tau>"
+ "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> AE_concat v1 v2 \<Rightarrow> (\<lbrace> z : b | c \<rbrace>) "
+ "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> AE_concat v1 v2 \<Rightarrow> (\<lbrace> z : B_bitvec | [[z]\<^sup>v]\<^sup>c\<^sup>e == (CE_concat [v1]\<^sup>c\<^sup>e [v1]\<^sup>c\<^sup>e) \<rbrace>) "
+ "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> (AE_appP f b v ) \<Rightarrow> \<tau>"
+ "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> AE_split v1 v2 \<Rightarrow> \<tau>"
+ "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> (AE_op Eq v1 v2) \<Rightarrow> \<lbrace> z3 : b | c \<rbrace>"
+ "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> (AE_op Eq v1 v2) \<Rightarrow> \<lbrace> z3 : B_bool | c \<rbrace>"
+ "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> (AE_op Eq v1 v2) \<Rightarrow> \<tau>"
+nominal_termination (eqvt) by lexicographic_order
+
+section \<open>Statements\<close>
+
+inductive check_s :: "\<Theta> \<Rightarrow> \<Phi> \<Rightarrow> \<B> \<Rightarrow> \<Gamma> \<Rightarrow> \<Delta> \<Rightarrow> s \<Rightarrow> \<tau> \<Rightarrow> bool" (" _ ; _ ; _ ; _ ; _ \<turnstile> _ \<Leftarrow> _" [50, 50, 50,50,50] 50) and
+ check_branch_s :: "\<Theta> \<Rightarrow> \<Phi> \<Rightarrow> \<B> \<Rightarrow> \<Gamma> \<Rightarrow> \<Delta> \<Rightarrow> tyid \<Rightarrow> string \<Rightarrow> \<tau> \<Rightarrow> v \<Rightarrow> branch_s \<Rightarrow> \<tau> \<Rightarrow> bool" (" _ ; _ ; _ ; _ ; _ ; _ ; _ ; _ ; _ \<turnstile> _ \<Leftarrow> _" [50, 50, 50,50,50] 50) and
+ check_branch_list :: "\<Theta> \<Rightarrow> \<Phi> \<Rightarrow> \<B> \<Rightarrow> \<Gamma> \<Rightarrow> \<Delta> \<Rightarrow> tyid \<Rightarrow> (string * \<tau>) list \<Rightarrow> v \<Rightarrow> branch_list \<Rightarrow> \<tau> \<Rightarrow> bool" (" _ ; _ ; _ ; _ ; _ ; _ ; _ ; _ \<turnstile> _ \<Leftarrow> _" [50, 50, 50,50,50] 50) where
+ check_valI: "\<lbrakk>
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> ;
+ \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> ;
+ \<Theta>; \<B>; \<Gamma> \<turnstile> v \<Rightarrow> \<tau>';
+ \<Theta>; \<B>; \<Gamma> \<turnstile> \<tau>' \<lesssim> \<tau> \<rbrakk> \<Longrightarrow>
+ \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> (AS_val v) \<Leftarrow> \<tau>"
+
+| check_letI: "\<lbrakk>
+ atom x \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>, \<Delta>, e, \<tau>);
+ atom z \<sharp> (x, \<Theta>, \<Phi>, \<B>, \<Gamma>, \<Delta>, e, \<tau>, s);
+ \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> e \<Rightarrow> \<lbrace> z : b | c \<rbrace> ;
+ \<Theta>; \<Phi> ; \<B> ; ((x,b,c[z::=V_var x]\<^sub>v)#\<^sub>\<Gamma>\<Gamma>) ; \<Delta> \<turnstile> s \<Leftarrow> \<tau>
+\<rbrakk> \<Longrightarrow>
+ \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> (AS_let x e s) \<Leftarrow> \<tau>"
+
+| check_assertI: "\<lbrakk>
+ atom x \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>, \<Delta>, c, \<tau>, s);
+ \<Theta>; \<Phi> ; \<B> ; ((x,B_bool,c)#\<^sub>\<Gamma>\<Gamma>) ; \<Delta> \<turnstile> s \<Leftarrow> \<tau> ;
+ \<Theta>; \<B>; \<Gamma> \<Turnstile> c;
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>
+\<rbrakk> \<Longrightarrow>
+ \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> (AS_assert c s) \<Leftarrow> \<tau>"
+
+|check_branch_s_branchI: "\<lbrakk>
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> ;
+ \<turnstile>\<^sub>w\<^sub>f \<Theta> ;
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau> ;
+ \<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f const;
+ atom x \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>, \<Delta>, tid, cons , const, v, \<tau>);
+ \<Theta>; \<Phi>; \<B>; ((x,b_of const, ([v]\<^sup>c\<^sup>e == [ V_cons tid cons [x]\<^sup>v]\<^sup>c\<^sup>e ) AND (c_of const x))#\<^sub>\<Gamma>\<Gamma>) ; \<Delta> \<turnstile> s \<Leftarrow> \<tau>
+\<rbrakk> \<Longrightarrow>
+ \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> ; tid ; cons ; const ; v \<turnstile> (AS_branch cons x s) \<Leftarrow> \<tau>"
+
+|check_branch_list_consI: "\<lbrakk>
+ \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta>; tid; cons; const; v \<turnstile> cs \<Leftarrow> \<tau> ;
+ \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta>; tid; dclist; v \<turnstile> css \<Leftarrow> \<tau>
+\<rbrakk> \<Longrightarrow>
+ \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> ; tid ; (cons,const)#dclist ; v \<turnstile> AS_cons cs css \<Leftarrow> \<tau>"
+
+|check_branch_list_finalI: "\<lbrakk>
+ \<Theta>; \<Phi>;\<B>; \<Gamma>; \<Delta>; tid ; cons ; const ; v \<turnstile> cs \<Leftarrow> \<tau>
+\<rbrakk> \<Longrightarrow>
+ \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> ; tid ; [(cons,const)] ; v \<turnstile> AS_final cs \<Leftarrow> \<tau>"
+
+| check_ifI: "\<lbrakk>
+ atom z \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>, \<Delta>, v , s1 , s2 , \<tau> );
+ (\<Theta>; \<B>; \<Gamma> \<turnstile> v \<Leftarrow> (\<lbrace> z : B_bool | TRUE \<rbrace>)) ;
+ \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> s1 \<Leftarrow> (\<lbrace> z : b_of \<tau> | ([v]\<^sup>c\<^sup>e == [[L_true]\<^sup>v]\<^sup>c\<^sup>e) IMP (c_of \<tau> z) \<rbrace>) ;
+ \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> s2 \<Leftarrow> (\<lbrace> z : b_of \<tau> | ([v]\<^sup>c\<^sup>e == [[L_false]\<^sup>v]\<^sup>c\<^sup>e) IMP (c_of \<tau> z) \<rbrace>)
+\<rbrakk> \<Longrightarrow>
+ \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> IF v THEN s1 ELSE s2 \<Leftarrow> \<tau>"
+
+| check_let2I: "\<lbrakk>
+ atom x \<sharp> (\<Theta>, \<Phi>, \<B>, G, \<Delta>, t, s1, \<tau>) ;
+ \<Theta>; \<Phi> ; \<B> ; G; \<Delta> \<turnstile> s1 \<Leftarrow> t;
+ \<Theta>; \<Phi> ; \<B> ; ((x,b_of t,c_of t x)#\<^sub>\<Gamma>G) ; \<Delta> \<turnstile> s2 \<Leftarrow> \<tau>
+\<rbrakk> \<Longrightarrow>
+ \<Theta>; \<Phi> ; \<B> ; G ; \<Delta> \<turnstile> (LET x : t = s1 IN s2) \<Leftarrow> \<tau>"
+
+| check_varI: "\<lbrakk>
+ atom u \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>, \<Delta>, \<tau>', v, \<tau>) ;
+ \<Theta>; \<B>; \<Gamma> \<turnstile> v \<Leftarrow> \<tau>';
+ \<Theta>; \<Phi>; \<B>; \<Gamma> ; ((u,\<tau>') #\<^sub>\<Delta> \<Delta>) \<turnstile> s \<Leftarrow> \<tau>
+\<rbrakk> \<Longrightarrow>
+ \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> (VAR u : \<tau>' = v IN s) \<Leftarrow> \<tau>"
+
+| check_assignI: "\<lbrakk>
+ \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> ;
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> ;
+ (u,\<tau>) \<in> setD \<Delta> ;
+ \<Theta>; \<B>; \<Gamma> \<turnstile> v \<Leftarrow> \<tau>;
+ \<Theta>; \<B>; \<Gamma> \<turnstile> (\<lbrace> z : B_unit | TRUE \<rbrace>) \<lesssim> \<tau>'
+\<rbrakk> \<Longrightarrow>
+ \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> (u ::= v) \<Leftarrow> \<tau>'"
+
+| check_whileI: "\<lbrakk>
+ \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> s1 \<Leftarrow> \<lbrace> z : B_bool | TRUE \<rbrace>;
+ \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> s2 \<Leftarrow> \<lbrace> z : B_unit | TRUE \<rbrace>;
+ \<Theta>; \<B>; \<Gamma> \<turnstile> (\<lbrace> z : B_unit | TRUE \<rbrace>) \<lesssim> \<tau>'
+\<rbrakk> \<Longrightarrow>
+ \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> WHILE s1 DO { s2 } \<Leftarrow> \<tau>'"
+
+| check_seqI: "\<lbrakk>
+ \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> s1 \<Leftarrow> \<lbrace> z : B_unit | TRUE \<rbrace>;
+ \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> s2 \<Leftarrow> \<tau>
+\<rbrakk> \<Longrightarrow>
+ \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> s1 ;; s2 \<Leftarrow> \<tau>"
+
+| check_caseI: "\<lbrakk>
+ \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta>; tid ; dclist ; v \<turnstile> cs \<Leftarrow> \<tau> ;
+ (AF_typedef tid dclist ) \<in> set \<Theta> ;
+ \<Theta>; \<B>; \<Gamma> \<turnstile> v \<Leftarrow> \<lbrace> z : B_id tid | TRUE \<rbrace>;
+ \<turnstile>\<^sub>w\<^sub>f \<Theta>
+\<rbrakk> \<Longrightarrow>
+ \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> AS_match v cs \<Leftarrow> \<tau>"
+
+equivariance check_s
+
+text \<open>We only need avoidance for cases where a variable is added to a context\<close>
+nominal_inductive check_s
+ avoids check_letI: x and z | check_branch_s_branchI: x | check_let2I: x | check_varI: u | check_ifI: z | check_assertI: x
+proof(goal_cases)
+ case (1 x \<Theta> \<Phi> \<B> \<Gamma> \<Delta> e \<tau> z s b c)
+ hence "atom x \<sharp> AS_let x e s" using s_branch_s_branch_list.fresh(2) by auto
+ moreover have "atom z \<sharp> AS_let x e s" using s_branch_s_branch_list.fresh(2) 1 fresh_prod8 by auto
+ then show ?case using fresh_star_def 1 by force
+next
+ case (3 x \<Theta> \<Phi> \<B> \<Gamma> \<Delta> c \<tau> s)
+ hence "atom x \<sharp> AS_assert c s" using fresh_prodN s_branch_s_branch_list.fresh pure_fresh by auto
+ then show ?case using fresh_star_def 3 by force
+next
+ case (5 \<Theta> \<B> \<Gamma> \<Delta> \<tau> const x \<Phi> tid cons v s)
+ hence "atom x \<sharp> AS_branch cons x s" using fresh_prodN s_branch_s_branch_list.fresh pure_fresh by auto
+ then show ?case using fresh_star_def 5 by force
+next
+ case (7 z \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v s1 s2 \<tau>)
+ hence "atom z \<sharp> AS_if v s1 s2 " using s_branch_s_branch_list.fresh by auto
+ then show ?case using 7 fresh_prodN fresh_star_def by fastforce
+next
+ case (9 x \<Theta> \<Phi> \<B> G \<Delta> t s1 \<tau> s2)
+ hence "atom x \<sharp> AS_let2 x t s1 s2" using s_branch_s_branch_list.fresh by auto
+ thus ?case using fresh_star_def 9 by force
+next
+ case (11 u \<Theta> \<Phi> \<B> \<Gamma> \<Delta> \<tau>' v \<tau> s)
+ hence "atom u \<sharp> AS_var u \<tau>' v s" using s_branch_s_branch_list.fresh by auto
+ then show ?case using fresh_star_def 11 by force
+
+qed(auto+)
+
+inductive_cases check_s_elims[elim!]:
+ "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> AS_val v \<Leftarrow> \<tau>"
+ "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> AS_let x e s \<Leftarrow> \<tau>"
+ "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> AS_if v s1 s2 \<Leftarrow> \<tau>"
+ "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> AS_let2 x t s1 s2 \<Leftarrow> \<tau>"
+ "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> AS_while s1 s2 \<Leftarrow> \<tau>"
+ "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> AS_var u t v s \<Leftarrow> \<tau>"
+ "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> AS_seq s1 s2 \<Leftarrow> \<tau>"
+ "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> AS_assign u v \<Leftarrow> \<tau>"
+ "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> AS_match v cs \<Leftarrow> \<tau>"
+ "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile> AS_assert c s \<Leftarrow> \<tau>"
+
+inductive_cases check_branch_s_elims[elim!]:
+ "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta>; tid ; dclist ; v \<turnstile> (AS_final cs) \<Leftarrow> \<tau>"
+ "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta>; tid ; dclist ; v \<turnstile> (AS_cons cs css) \<Leftarrow> \<tau>"
+ "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta>; tid ; cons ; const ; v \<turnstile> (AS_branch dc x s ) \<Leftarrow> \<tau>"
+
+section \<open>Programs\<close>
+
+text \<open>Type check function bodies\<close>
+
+inductive check_funtyp :: "\<Theta> \<Rightarrow> \<Phi> \<Rightarrow> \<B> \<Rightarrow> fun_typ \<Rightarrow> bool" ( " _ ; _ ; _ \<turnstile> _ " ) where
+ check_funtypI: "\<lbrakk>
+ atom x \<sharp> (\<Theta>, \<Phi>, B , b );
+ \<Theta>; \<Phi> ; B ; ((x,b,c) #\<^sub>\<Gamma> GNil) ; []\<^sub>\<Delta> \<turnstile> s \<Leftarrow> \<tau>
+\<rbrakk> \<Longrightarrow>
+ \<Theta>; \<Phi> ; B \<turnstile> (AF_fun_typ x b c \<tau> s)"
+
+equivariance check_funtyp
+nominal_inductive check_funtyp
+ avoids check_funtypI: x
+proof(goal_cases)
+ case (1 x \<Theta> \<Phi> B b c s \<tau> )
+ hence "atom x \<sharp> (AF_fun_typ x b c \<tau> s)" using fun_def.fresh fun_typ_q.fresh fun_typ.fresh by simp
+ then show ?case using fresh_star_def 1 fresh_prodN by fastforce
+next
+ case (2 \<Theta> \<Phi> x b c s \<tau> f)
+ then show ?case by auto
+qed
+
+inductive check_funtypq :: "\<Theta> \<Rightarrow> \<Phi> \<Rightarrow> fun_typ_q \<Rightarrow> bool" ( " _ ; _ \<turnstile> _ " ) where
+ check_fundefq_simpleI: "\<lbrakk>
+ \<Theta>; \<Phi> ; {||} \<turnstile> (AF_fun_typ x b c t s)
+\<rbrakk> \<Longrightarrow>
+ \<Theta>; \<Phi> \<turnstile> ((AF_fun_typ_none (AF_fun_typ x b c t s)))"
+
+|check_funtypq_polyI: "\<lbrakk>
+ atom bv \<sharp> (\<Theta>, \<Phi>, (AF_fun_typ x b c t s));
+ \<Theta>; \<Phi>; {|bv|} \<turnstile> (AF_fun_typ x b c t s)
+\<rbrakk> \<Longrightarrow>
+ \<Theta>; \<Phi> \<turnstile> (AF_fun_typ_some bv (AF_fun_typ x b c t s))"
+
+equivariance check_funtypq
+nominal_inductive check_funtypq
+ avoids check_funtypq_polyI: bv
+proof(goal_cases)
+ case (1 bv \<Theta> \<Phi> x b c t s )
+ hence "atom bv \<sharp> (AF_fun_typ_some bv (AF_fun_typ x b c t s))" using fun_def.fresh fun_typ_q.fresh fun_typ.fresh by simp
+ thus ?case using fresh_star_def 1 fresh_prodN by fastforce
+next
+ case (2 bv \<Theta> \<Phi> ft )
+ then show ?case by auto
+qed
+
+inductive check_fundef :: "\<Theta> \<Rightarrow> \<Phi> \<Rightarrow> fun_def \<Rightarrow> bool" ( " _ ; _ \<turnstile> _ " ) where
+ check_fundefI: "\<lbrakk>
+ \<Theta>; \<Phi> \<turnstile> ft
+\<rbrakk> \<Longrightarrow>
+ \<Theta>; \<Phi> \<turnstile> (AF_fundef f ft)"
+
+equivariance check_fundef
+nominal_inductive check_fundef .
+
+text \<open>Temporarily remove this simproc as it produces untidy eliminations\<close>
+declare[[ simproc del: alpha_lst]]
+
+inductive_cases check_funtyp_elims[elim!]:
+ "check_funtyp \<Theta> \<Phi> B ft"
+
+inductive_cases check_funtypq_elims[elim!]:
+ "check_funtypq \<Theta> \<Phi> (AF_fun_typ_none (AF_fun_typ x b c \<tau> s))"
+ "check_funtypq \<Theta> \<Phi> (AF_fun_typ_some bv (AF_fun_typ x b c \<tau> s))"
+
+inductive_cases check_fundef_elims[elim!]:
+ "check_fundef \<Theta> \<Phi> (AF_fundef f ftq)"
+
+declare[[ simproc add: alpha_lst]]
+
+nominal_function \<Delta>_of :: "var_def list \<Rightarrow> \<Delta>" where
+ "\<Delta>_of [] = DNil"
+| "\<Delta>_of ((AV_def u t v)#vs) = (u,t) #\<^sub>\<Delta> (\<Delta>_of vs)"
+ apply auto
+ using eqvt_def \<Delta>_of_graph_aux_def neq_Nil_conv old.prod.exhaust apply force
+ using eqvt_def \<Delta>_of_graph_aux_def neq_Nil_conv old.prod.exhaust
+ by (metis var_def.strong_exhaust)
+nominal_termination (eqvt) by lexicographic_order
+
+inductive check_prog :: "p \<Rightarrow> \<tau> \<Rightarrow> bool" ( "\<turnstile> _ \<Leftarrow> _") where
+ "\<lbrakk>
+ \<Theta>; \<Phi>; {||}; GNil ; \<Delta>_of \<G> \<turnstile> s \<Leftarrow> \<tau>
+\<rbrakk> \<Longrightarrow> \<turnstile> (AP_prog \<Theta> \<Phi> \<G> s) \<Leftarrow> \<tau>"
+
+inductive_cases check_prog_elims[elim!]:
+ "\<turnstile> (AP_prog \<Theta> \<Phi> \<G> s) \<Leftarrow> \<tau>"
+
end
\ No newline at end of file
diff --git a/thys/MiniSail/TypingL.thy b/thys/MiniSail/TypingL.thy
--- a/thys/MiniSail/TypingL.thy
+++ b/thys/MiniSail/TypingL.thy
@@ -1,2800 +1,2800 @@
-(*<*)
-theory TypingL
- imports Typing RCLogicL "HOL-Eisbach.Eisbach"
-begin
- (*>*)
-
-chapter \<open>Typing Lemmas\<close>
-
-section \<open>Prelude\<close>
-
-text \<open>Needed as the typing elimination rules give us facts for an alpha-equivalent version of a term
- and so need to be able to 'jump back' to a typing judgement for the orginal term\<close>
-
-lemma \<tau>_fresh_c[simp]:
- assumes "atom x \<sharp> \<lbrace> z : b | c \<rbrace>" and "atom z \<sharp> x"
- shows "atom x \<sharp> c"
- using \<tau>.fresh assms fresh_at_base
- by (simp add: fresh_at_base(2))
-
-lemmas subst_defs = subst_b_b_def subst_b_c_def subst_b_\<tau>_def subst_v_v_def subst_v_c_def subst_v_\<tau>_def
-
-lemma wfT_wfT_if1:
- assumes "wfT \<Theta> \<B> \<Gamma> (\<lbrace> z : b_of t | CE_val v == CE_val (V_lit L_false) IMP c_of t z \<rbrace>)" and "atom z \<sharp> (\<Gamma>,t)"
- shows "wfT \<Theta> \<B> \<Gamma> t"
- using assms proof(nominal_induct t avoiding: \<Gamma> z rule: \<tau>.strong_induct)
- case (T_refined_type z' b' c')
- show ?case proof(rule wfT_wfT_if)
- show \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b' | [ v ]\<^sup>c\<^sup>e == [ [ L_false ]\<^sup>v ]\<^sup>c\<^sup>e IMP c'[z'::=[ z]\<^sup>v]\<^sub>c\<^sub>v \<rbrace> \<close>
- using T_refined_type b_of.simps c_of.simps subst_defs by metis
- show \<open>atom z \<sharp> (c', \<Gamma>)\<close> using T_refined_type fresh_prodN \<tau>_fresh_c by metis
- qed
-qed
-
-lemma fresh_u_replace_true:
- fixes bv::bv and \<Gamma>::\<Gamma>
- assumes "atom bv \<sharp> \<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>"
- shows "atom bv \<sharp> \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>"
- using fresh_append_g fresh_GCons assms fresh_Pair c.fresh(1) by auto
-
-lemma wf_replace_true1:
- fixes \<Gamma>::\<Gamma> and \<Phi>::\<Phi> and \<Theta>::\<Theta> and \<Gamma>'::\<Gamma> and v::v and e::e and c::c and c''::c and c'::c and \<tau>::\<tau> and ts::"(string*\<tau>) list" and \<Delta>::\<Delta> and b'::b and b::b and s::s
- and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def and cs::branch_s and css::branch_list
-
-shows "\<Theta>; \<B>; G \<turnstile>\<^sub>w\<^sub>f v : b' \<Longrightarrow> G = \<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma> \<Longrightarrow> \<Theta> ; \<B> ; \<Gamma>' @ ((x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>) \<turnstile>\<^sub>w\<^sub>f v : b'" and
- "\<Theta>; \<B>; G \<turnstile>\<^sub>w\<^sub>f c'' \<Longrightarrow> G = \<Gamma>' @(x, b, c) #\<^sub>\<Gamma> \<Gamma> \<Longrightarrow> \<Theta> ; \<B> ; \<Gamma>' @ ((x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>) \<turnstile>\<^sub>w\<^sub>f c''" and
- "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f G \<Longrightarrow> G = \<Gamma>' @(x, b, c) #\<^sub>\<Gamma> \<Gamma> \<Longrightarrow> \<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>' @ ((x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>) " and
- "\<Theta>; \<B>; G \<turnstile>\<^sub>w\<^sub>f \<tau> \<Longrightarrow> G = \<Gamma>' @(x, b, c) #\<^sub>\<Gamma> \<Gamma> \<Longrightarrow> \<Theta> ; \<B> ; \<Gamma>' @ ((x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>) \<turnstile>\<^sub>w\<^sub>f \<tau>" and
- "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ts \<Longrightarrow> True" and
- "\<turnstile>\<^sub>w\<^sub>f P \<Longrightarrow> True" and
- "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f b \<Longrightarrow> True" and
- "\<Theta> ; \<B> ; G \<turnstile>\<^sub>w\<^sub>f ce : b' \<Longrightarrow> G = \<Gamma>' @(x, b, c) #\<^sub>\<Gamma> \<Gamma> \<Longrightarrow> \<Theta> ; \<B> ; \<Gamma>' @ ((x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>) \<turnstile>\<^sub>w\<^sub>f ce : b'" and
- "\<Theta> \<turnstile>\<^sub>w\<^sub>f td \<Longrightarrow> True"
-proof(nominal_induct
- b' and c'' and G and \<tau> and ts and P and b and b' and td
- arbitrary: \<Gamma> \<Gamma>' and \<Gamma> \<Gamma>' and \<Gamma> \<Gamma>' and \<Gamma> \<Gamma>' and \<Gamma> \<Gamma>' and \<Gamma> \<Gamma>' and \<Gamma> \<Gamma>' and \<Gamma> \<Gamma>' and \<Gamma> \<Gamma>' and \<Gamma> \<Gamma>' and \<Gamma> \<Gamma>' and \<Gamma> \<Gamma>' and \<Gamma> \<Gamma>' and \<Gamma> \<Gamma>' and \<Gamma> \<Gamma>' and \<Gamma> \<Gamma>' and \<Gamma> \<Gamma>'
- rule:wfV_wfC_wfG_wfT_wfTs_wfTh_wfB_wfCE_wfTD.strong_induct)
- case (wfB_intI \<Theta> \<B>)
- then show ?case using wf_intros by metis
-next
- case (wfB_boolI \<Theta> \<B>)
- then show ?case using wf_intros by metis
-next
- case (wfB_unitI \<Theta> \<B>)
- then show ?case using wf_intros by metis
-next
- case (wfB_bitvecI \<Theta> \<B>)
- then show ?case using wf_intros by metis
-next
- case (wfB_pairI \<Theta> \<B> b1 b2)
- then show ?case using wf_intros by metis
-next
- case (wfB_consI \<Theta> s dclist \<B>)
- then show ?case using wf_intros by metis
-next
- case (wfB_appI \<Theta> b s bv dclist \<B>)
- then show ?case using wf_intros by metis
-next
- case (wfV_varI \<Theta> \<B> \<Gamma>'' b' c x')
- hence wfg: \<open> \<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<close> by auto
- show ?case proof(cases "x=x'")
- case True
- hence "Some (b, TRUE) = lookup (\<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>) x'" using lookup.simps lookup_inside_wf wfg by simp
- thus ?thesis using Wellformed.wfV_varI[OF wfg]
- by (metis True lookup_inside_wf old.prod.inject option.inject wfV_varI.hyps(1) wfV_varI.hyps(3) wfV_varI.prems)
- next
- case False
- hence "Some (b', c) = lookup (\<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>) x'" using lookup_inside2 wfV_varI by metis
- then show ?thesis using Wellformed.wfV_varI[OF wfg]
- by (metis wfG_elim2 wfG_suffix wfV_varI.hyps(1) wfV_varI.hyps(2) wfV_varI.hyps(3)
- wfV_varI.prems Wellformed.wfV_varI wf_replace_inside(1))
- qed
-next
- case (wfV_litI \<Theta> \<B> \<Gamma> l)
- then show ?case using wf_intros using wf_intros by metis
-next
- case (wfV_pairI \<Theta> \<B> \<Gamma> v1 b1 v2 b2)
- then show ?case using wf_intros by metis
-next
- case (wfV_consI s dclist \<Theta> dc x b' c \<B> \<Gamma> v)
- then show ?case using wf_intros by metis
-next
- case (wfV_conspI s bv dclist \<Theta> dc xc bc cc \<B> b' \<Gamma>'' v)
- show ?case proof
- show \<open>AF_typedef_poly s bv dclist \<in> set \<Theta>\<close> using wfV_conspI by metis
- show \<open>(dc, \<lbrace> xc : bc | cc \<rbrace>) \<in> set dclist\<close> using wfV_conspI by metis
- show \<open> \<Theta> ;\<B> \<turnstile>\<^sub>w\<^sub>f b' \<close> using wfV_conspI by metis
- show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : bc[bv::=b']\<^sub>b\<^sub>b \<close> using wfV_conspI by metis
- have "atom bv \<sharp> \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>" using fresh_u_replace_true wfV_conspI by metis
- thus \<open>atom bv \<sharp> (\<Theta>, \<B>, \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>, b', v)\<close> using wfV_conspI fresh_prodN by metis
- qed
-next
- case (wfCE_valI \<Theta> \<B> \<Gamma> v b)
- then show ?case using wf_intros by metis
-next
- case (wfCE_plusI \<Theta> \<B> \<Gamma> v1 v2)
- then show ?case using wf_intros by metis
-next
- case (wfCE_leqI \<Theta> \<B> \<Gamma> v1 v2)
- then show ?case using wf_intros by metis
-next
- case (wfCE_eqI \<Theta> \<B> \<Gamma> v1 v2)
- then show ?case using wf_intros by metis
-next
- case (wfCE_fstI \<Theta> \<B> \<Gamma> v1 b1 b2)
- then show ?case using wf_intros by metis
-next
- case (wfCE_sndI \<Theta> \<B> \<Gamma> v1 b1 b2)
- then show ?case using wf_intros by metis
-next
- case (wfCE_concatI \<Theta> \<B> \<Gamma> v1 v2)
- then show ?case using wf_intros by metis
-next
- case (wfCE_lenI \<Theta> \<B> \<Gamma> v1)
- then show ?case using wf_intros by metis
-next
- case (wfTI z \<Theta> \<B> \<Gamma>'' b' c')
- show ?case proof
- show \<open>atom z \<sharp> (\<Theta>, \<B>, \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>)\<close> using wfTI fresh_append_g fresh_GCons fresh_prodN by auto
- show \<open> \<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f b' \<close> using wfTI by metis
- show \<open> \<Theta>; \<B>; (z, b', TRUE) #\<^sub>\<Gamma> \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f c' \<close> using wfTI append_g.simps by metis
- qed
-next
- case (wfC_eqI \<Theta> \<B> \<Gamma> e1 b e2)
- then show ?case using wf_intros by metis
-next
- case (wfC_trueI \<Theta> \<B> \<Gamma>)
- then show ?case using wf_intros by metis
-next
- case (wfC_falseI \<Theta> \<B> \<Gamma>)
- then show ?case using wf_intros by metis
-next
- case (wfC_conjI \<Theta> \<B> \<Gamma> c1 c2)
- then show ?case using wf_intros by metis
-next
- case (wfC_disjI \<Theta> \<B> \<Gamma> c1 c2)
- then show ?case using wf_intros by metis
-next
- case (wfC_notI \<Theta> \<B> \<Gamma> c1)
- then show ?case using wf_intros by metis
-next
- case (wfC_impI \<Theta> \<B> \<Gamma> c1 c2)
- then show ?case using wf_intros by metis
-next
- case (wfG_nilI \<Theta> \<B>)
- then show ?case using GNil_append by blast
-next
- case (wfG_cons1I c \<Theta> \<B> \<Gamma>'' x b)
- then show ?case using wf_intros wfG_cons_TRUE2 wfG_elims(2) wfG_replace_inside wfG_suffix
- by (metis (no_types, lifting))
-next
- case (wfG_cons2I c \<Theta> \<B> \<Gamma>'' x' b)
- then show ?case using wf_intros
- by (metis wfG_cons_TRUE2 wfG_elims(2) wfG_replace_inside wfG_suffix)
-next
- case wfTh_emptyI
- then show ?case using wf_intros by metis
-next
- case (wfTh_consI tdef \<Theta>)
- then show ?case using wf_intros by metis
-next
- case (wfTD_simpleI \<Theta> lst s)
- then show ?case using wf_intros by metis
-next
- case (wfTD_poly \<Theta> bv lst s)
- then show ?case using wf_intros by metis
-next
- case (wfTs_nil \<Theta> \<B> \<Gamma>)
- then show ?case using wf_intros by metis
-next
- case (wfTs_cons \<Theta> \<B> \<Gamma> \<tau> dc ts)
- then show ?case using wf_intros by metis
-qed
-
-lemma wf_replace_true2:
- fixes \<Gamma>::\<Gamma> and \<Phi>::\<Phi> and \<Theta>::\<Theta> and \<Gamma>'::\<Gamma> and v::v and e::e and c::c and c''::c and c'::c and \<tau>::\<tau> and ts::"(string*\<tau>) list" and \<Delta>::\<Delta> and b'::b and b::b and s::s
- and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def and cs::branch_s and css::branch_list
-
-shows "\<Theta> ; \<Phi> ; \<B> ; G ; D \<turnstile>\<^sub>w\<^sub>f e : b' \<Longrightarrow> G = \<Gamma>' @(x, b, c) #\<^sub>\<Gamma> \<Gamma> \<Longrightarrow> \<Theta> ; \<Phi> ; \<B> ; \<Gamma>' @ ((x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>); D \<turnstile>\<^sub>w\<^sub>f e : b'" and
- "\<Theta> ; \<Phi> ; \<B> ; G ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b' \<Longrightarrow> G = \<Gamma>' @(x, b, c) #\<^sub>\<Gamma> \<Gamma> \<Longrightarrow> \<Theta> ; \<Phi> ; \<B> ; \<Gamma>' @ ((x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>) ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b'" and
- "\<Theta> ; \<Phi> ; \<B> ; G ; \<Delta> ; tid ; dc ; t \<turnstile>\<^sub>w\<^sub>f cs : b' \<Longrightarrow> G = \<Gamma>' @(x, b, c) #\<^sub>\<Gamma> \<Gamma> \<Longrightarrow> \<Theta> ; \<Phi> ; \<B> ; \<Gamma>' @ ((x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>) ; \<Delta> ; tid ; dc ; t \<turnstile>\<^sub>w\<^sub>f cs : b'" and
- "\<Theta> ; \<Phi> ; \<B> ; G ; \<Delta> ; tid ; dclist \<turnstile>\<^sub>w\<^sub>f css : b' \<Longrightarrow> G = \<Gamma>' @(x, b, c) #\<^sub>\<Gamma> \<Gamma> \<Longrightarrow> \<Theta> ; \<Phi> ; \<B> ; \<Gamma>' @ ((x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>) ; \<Delta> ; tid ; dclist \<turnstile>\<^sub>w\<^sub>f css : b'" and
-
-"\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<Longrightarrow> True" and
-"\<Theta>; \<B>; G \<turnstile>\<^sub>w\<^sub>f \<Delta> \<Longrightarrow> G = \<Gamma>' @(x, b, c) #\<^sub>\<Gamma> \<Gamma> \<Longrightarrow> \<Theta> ; \<B> ; \<Gamma>' @ ((x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>) \<turnstile>\<^sub>w\<^sub>f \<Delta>" and
-
-"\<Theta> ; \<Phi> \<turnstile>\<^sub>w\<^sub>f ftq \<Longrightarrow> True" and
-"\<Theta> ; \<Phi> ; \<B> \<turnstile>\<^sub>w\<^sub>f ft \<Longrightarrow> True"
-proof(nominal_induct
- b' and b' and b' and b' and \<Phi> and \<Delta> and ftq and ft
- arbitrary: \<Gamma> \<Gamma>' and \<Gamma> \<Gamma>' and \<Gamma> \<Gamma>' and \<Gamma> \<Gamma>' and \<Gamma> \<Gamma>' and \<Gamma> \<Gamma>' and \<Gamma> \<Gamma>' and \<Gamma> \<Gamma>' and \<Gamma> \<Gamma>' and \<Gamma> \<Gamma>' and \<Gamma> \<Gamma>' and \<Gamma> \<Gamma>' and \<Gamma> \<Gamma>' and \<Gamma> \<Gamma>' and \<Gamma> \<Gamma>' and \<Gamma> \<Gamma>' and \<Gamma> \<Gamma>'
- rule:wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.strong_induct)
-
- case (wfE_valI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v b)
- then show ?case using wf_intros using wf_intros wf_replace_true1 by metis
-next
- case (wfE_plusI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 v2)
- then show ?case using wf_intros wf_replace_true1 by metis
-next
- case (wfE_leqI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 v2)
- then show ?case using wf_intros wf_replace_true1 by metis
-next
- case (wfE_eqI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 b v2)
- then show ?case using wf_intros wf_replace_true1 by metis
-next
- case (wfE_fstI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 b1 b2)
- then show ?case using wf_intros wf_replace_true1 by metis
-next
- case (wfE_sndI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 b1 b2)
- then show ?case using wf_intros wf_replace_true1 by metis
-next
- case (wfE_concatI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 v2)
- then show ?case using wf_intros wf_replace_true1 by metis
-next
- case (wfE_splitI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 v2)
- then show ?case using wf_intros wf_replace_true1 by metis
-next
- case (wfE_lenI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1)
- then show ?case using wf_intros wf_replace_true1 by metis
-next
- case (wfE_appI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> f x b c \<tau> s v)
- then show ?case using wf_intros wf_replace_true1 by metis
-next
- case (wfE_appPI \<Theta> \<Phi> \<B> \<Gamma>'' \<Delta> b' bv v \<tau> f x1 b1 c1 s)
- show ?case proof
- show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using wfE_appPI wf_replace_true1 by metis
- show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wfE_appPI by metis
- show \<open> \<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f b' \<close> using wfE_appPI by metis
- have "atom bv \<sharp> \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>" using fresh_u_replace_true wfE_appPI fresh_prodN by metis
- thus \<open>atom bv \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>, \<Delta>, b', v, (b_of \<tau>)[bv::=b']\<^sub>b)\<close>
- using wfE_appPI fresh_prodN by auto
- show \<open>Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x1 b1 c1 \<tau> s))) = lookup_fun \<Phi> f\<close> using wfE_appPI by metis
- show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b1[bv::=b']\<^sub>b \<close> using wfE_appPI wf_replace_true1 by metis
- qed
-next
- case (wfE_mvarI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> u \<tau>)
- then show ?case using wf_intros wf_replace_true1 by metis
-next
-
- case (wfS_valI \<Theta> \<Phi> \<B> \<Gamma> v b \<Delta>)
- then show ?case using wf_intros wf_replace_true1 by metis
-next
- case (wfS_letI \<Theta> \<Phi> \<B> \<Gamma>'' \<Delta> e b' x1 s b1)
- show ?case proof
- show \<open> \<Theta> ; \<Phi> ; \<B> ; \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f e : b' \<close> using wfS_letI wf_replace_true1 by metis
- have \<open> \<Theta> ; \<Phi> ; \<B> ; ((x1, b', TRUE) #\<^sub>\<Gamma> \<Gamma>') @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b1 \<close> apply(rule wfS_letI(4))
- using wfS_letI append_g.simps by simp
- thus \<open> \<Theta> ; \<Phi> ; \<B> ; (x1, b', TRUE) #\<^sub>\<Gamma> \<Gamma>'@ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b1 \<close> using append_g.simps by auto
- show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wfS_letI by metis
- show "atom x1 \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>, \<Delta>, e, b1)" using fresh_append_g fresh_GCons fresh_prodN wfS_letI by auto
- qed
-next
- case (wfS_assertI \<Theta> \<Phi> \<B> x' c \<Gamma>'' \<Delta> s b')
- show ?case proof
- show \<open> \<Theta> ; \<Phi> ; \<B> ; (x', B_bool, c) #\<^sub>\<Gamma> \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b' \<close>
- using wfS_assertI (2)[of "(x', B_bool, c) #\<^sub>\<Gamma> \<Gamma>'" \<Gamma>] wfS_assertI by simp
- show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f c \<close> using wfS_assertI wf_replace_true1 by metis
- show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wfS_assertI by metis
- show \<open>atom x' \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>, \<Delta>, c, b', s)\<close> using wfS_assertI fresh_prodN by simp
- qed
-next
- case (wfS_let2I \<Theta> \<Phi> \<B> \<Gamma>'' \<Delta> s1 \<tau> x' s2 ba')
- show ?case proof
- show \<open> \<Theta> ; \<Phi> ; \<B> ; \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s1 : b_of \<tau> \<close> using wfS_let2I wf_replace_true1 by metis
- show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau> \<close> using wfS_let2I wf_replace_true1 by metis
- have \<open> \<Theta> ; \<Phi> ; \<B> ; ((x', b_of \<tau>, TRUE) #\<^sub>\<Gamma> \<Gamma>') @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s2 : ba' \<close>
- apply(rule wfS_let2I(5))
- using wfS_let2I append_g.simps by auto
- thus \<open> \<Theta> ; \<Phi> ; \<B> ; (x', b_of \<tau>, TRUE) #\<^sub>\<Gamma> \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s2 : ba' \<close> using wfS_let2I append_g.simps by auto
- show \<open>atom x' \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>, \<Delta>, s1, ba', \<tau>)\<close> using fresh_append_g fresh_GCons fresh_prodN wfS_let2I by auto
- qed
-next
- case (wfS_ifI \<Theta> \<B> \<Gamma> v \<Phi> \<Delta> s1 b s2)
- then show ?case using wf_intros wf_replace_true1 by metis
-next
- case (wfS_varI \<Theta> \<B> \<Gamma>'' \<tau> v u \<Phi> \<Delta> b' s)
- show ?case proof
- show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau> \<close> using wfS_varI wf_replace_true1 by metis
- show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b_of \<tau> \<close> using wfS_varI wf_replace_true1 by metis
- show \<open>atom u \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>, \<Delta>, \<tau>, v, b')\<close> using wfS_varI u_fresh_g fresh_prodN by auto
- show \<open> \<Theta> ; \<Phi> ; \<B> ; \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> ; (u, \<tau>) #\<^sub>\<Delta> \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b' \<close> using wfS_varI by metis
- qed
-
-next
- case (wfS_assignI u \<tau> \<Delta> \<Theta> \<B> \<Gamma> \<Phi> v)
- then show ?case using wf_intros wf_replace_true1 by metis
-next
- case (wfS_whileI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> s1 s2 b)
- then show ?case using wf_intros wf_replace_true1 by metis
-next
- case (wfS_seqI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> s1 s2 b)
- then show ?case using wf_intros by metis
-next
- case (wfS_matchI \<Theta> \<B> \<Gamma>'' v tid dclist \<Delta> \<Phi> cs b')
- show ?case proof
- show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : B_id tid \<close> using wfS_matchI wf_replace_true1 by auto
- show \<open>AF_typedef tid dclist \<in> set \<Theta>\<close> using wfS_matchI by auto
- show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wfS_matchI by auto
- show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using wfS_matchI by auto
- show \<open> \<Theta> ; \<Phi> ; \<B> ; \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> ; tid ; dclist \<turnstile>\<^sub>w\<^sub>f cs : b' \<close> using wfS_matchI by auto
- qed
-next
- case (wfS_branchI \<Theta> \<Phi> \<B> x' \<tau> \<Gamma>'' \<Delta> s b' tid dc)
- show ?case proof
- have \<open> \<Theta> ; \<Phi> ; \<B> ; ((x', b_of \<tau>, TRUE) #\<^sub>\<Gamma> \<Gamma>') @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b' \<close> using wfS_branchI append_g.simps by metis
- thus \<open> \<Theta> ; \<Phi> ; \<B> ; (x', b_of \<tau>, TRUE) #\<^sub>\<Gamma> \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b' \<close> using wfS_branchI append_g.simps append_g.simps by metis
- show \<open>atom x' \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>, \<Delta>, \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>, \<tau>)\<close> using wfS_branchI by auto
- show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wfS_branchI by auto
- qed
-next
- case (wfS_finalI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dc t cs b)
- then show ?case using wf_intros by metis
-next
- case (wfS_cons \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dc t cs b dclist css)
- then show ?case using wf_intros by metis
-next
- case (wfD_emptyI \<Theta> \<B> \<Gamma>)
- then show ?case using wf_intros wf_replace_true1 by metis
-next
- case (wfD_cons \<Theta> \<B> \<Gamma> \<Delta> \<tau> u)
- then show ?case using wf_intros wf_replace_true1 by metis
-next
- case (wfPhi_emptyI \<Theta>)
- then show ?case using wf_intros by metis
-next
- case (wfPhi_consI f \<Theta> \<Phi> ft)
- then show ?case using wf_intros by metis
-next
- case (wfFTNone \<Theta> \<Phi> ft)
- then show ?case using wf_intros by metis
-next
- case (wfFTSome \<Theta> \<Phi> bv ft)
- then show ?case using wf_intros by metis
-next
- case (wfFTI \<Theta> B b \<Phi> x c s \<tau>)
- then show ?case using wf_intros by metis
-qed
-
-lemmas wf_replace_true = wf_replace_true1 wf_replace_true2
-
-section \<open>Subtyping\<close>
-
-lemma subtype_reflI2:
- fixes \<tau>::\<tau>
- assumes "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau>"
- shows "\<Theta>; \<B>; \<Gamma> \<turnstile> \<tau> \<lesssim> \<tau>"
-proof -
- obtain z b c where *:"\<tau> = \<lbrace> z : b | c \<rbrace> \<and> atom z \<sharp> (\<Theta>,\<B>,\<Gamma>) \<and> \<Theta>; \<B>; (z, b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f c"
- using wfT_elims(1)[OF assms] by metis
- obtain x::x where **: "atom x \<sharp> (\<Theta>, \<B>, \<Gamma>, c, z ,c, z , c )" using obtain_fresh by metis
- have "\<Theta>; \<B>; \<Gamma> \<turnstile> \<lbrace> z : b | c \<rbrace> \<lesssim> \<lbrace> z : b | c \<rbrace>" proof
- show "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b | c \<rbrace>" using * assms by auto
- show "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b | c \<rbrace>" using * assms by auto
- show "atom x \<sharp> (\<Theta>, \<B>, \<Gamma>, z , c , z , c )" using fresh_prod6 fresh_prod5 ** by metis
- thus "\<Theta>; \<B>; (x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<Turnstile> c[z::=V_var x]\<^sub>v " using wfT_wfC_cons assms * ** subst_v_c_def by simp
- qed
- thus ?thesis using * by auto
-qed
-
-lemma subtype_reflI:
- assumes "\<lbrace> z1 : b | c1 \<rbrace> = \<lbrace> z2 : b | c2 \<rbrace>" and wf1: "\<Theta>; \<B>;\<Gamma> \<turnstile>\<^sub>w\<^sub>f (\<lbrace> z1 : b | c1 \<rbrace>)"
- shows "\<Theta>; \<B>; \<Gamma> \<turnstile> (\<lbrace> z1 : b | c1 \<rbrace>) \<lesssim> (\<lbrace> z2 : b | c2 \<rbrace>)"
- using assms subtype_reflI2 by metis
-
-nominal_function base_eq :: "\<Gamma> \<Rightarrow> \<tau> \<Rightarrow> \<tau> \<Rightarrow> bool" where
- "base_eq _ \<lbrace> z1 : b1| c1 \<rbrace> \<lbrace> z2 : b2 | c2 \<rbrace> = (b1 = b2)"
- apply(auto,simp add: eqvt_def base_eq_graph_aux_def )
- by (meson \<tau>.exhaust)
-nominal_termination (eqvt) by lexicographic_order
-
-lemma subtype_wfT:
- fixes t1::\<tau> and t2::\<tau>
- assumes "\<Theta>; \<B>; \<Gamma> \<turnstile> t1 \<lesssim> t2"
- shows "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f t1 \<and> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f t2"
- using assms subtype_elims by metis
-
-lemma subtype_eq_base:
- assumes "\<Theta>; \<B>; \<Gamma> \<turnstile> (\<lbrace> z1 : b1| c1 \<rbrace>) \<lesssim> (\<lbrace> z2 : b2 | c2 \<rbrace>)"
- shows "b1=b2"
- using subtype.simps assms by auto
-
-lemma subtype_eq_base2:
- assumes "\<Theta>; \<B>; \<Gamma> \<turnstile> t1 \<lesssim> t2"
- shows "b_of t1 = b_of t2"
- using assms proof(rule subtype.induct[of \<Theta> \<B> \<Gamma> t1 t2],goal_cases)
- case (1 \<Theta> \<Gamma> z1 b c1 z2 c2 x)
- then show ?case using subtype_eq_base by auto
-qed
-
-lemma subtype_wf:
- fixes \<tau>1::\<tau> and \<tau>2::\<tau>
- assumes "\<Theta>; \<B>; \<Gamma> \<turnstile> \<tau>1 \<lesssim> \<tau>2"
- shows "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau>1 \<and>\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau>2"
- using assms
-proof(rule subtype.induct[of \<Theta> \<B> \<Gamma> \<tau>1 \<tau>2],goal_cases)
- case (1 \<Theta> \<Gamma>G z1 b c1 z2 c2 x)
- then show ?case by blast
-qed
-
-lemma subtype_g_wf:
- fixes \<tau>1::\<tau> and \<tau>2::\<tau> and \<Gamma>::\<Gamma>
- assumes "\<Theta>; \<B>; \<Gamma> \<turnstile> \<tau>1 \<lesssim> \<tau>2"
- shows "\<Theta> ; \<B>\<turnstile>\<^sub>w\<^sub>f \<Gamma>"
- using assms
-proof(rule subtype.induct[of \<Theta> \<B> \<Gamma> \<tau>1 \<tau>2],goal_cases)
- case (1 \<Theta> \<B> \<Gamma> z1 b c1 z2 c2 x)
- then show ?case using wfX_wfY by auto
-qed
-
-text \<open> For when we have a particular y that satisfies the freshness conditions that we want the validity check to use \<close>
-
-lemma valid_flip_simple:
- assumes "\<Theta>; \<B>; (z, b, c) #\<^sub>\<Gamma> \<Gamma> \<Turnstile> c'" and "atom z \<sharp> \<Gamma>" and "atom x \<sharp> (z, c, z, c', \<Gamma>)"
- shows "\<Theta>; \<B>; (x, b, (z \<leftrightarrow> x ) \<bullet> c) #\<^sub>\<Gamma> \<Gamma> \<Turnstile> (z \<leftrightarrow> x ) \<bullet> c'"
-proof -
- have "(z \<leftrightarrow> x ) \<bullet> \<Theta>; \<B>; (z \<leftrightarrow> x ) \<bullet> ((z, b, c) #\<^sub>\<Gamma> \<Gamma>) \<Turnstile> (z \<leftrightarrow> x ) \<bullet> c'"
- using True_eqvt valid.eqvt assms beta_flip_eq wfX_wfY by metis
- moreover have " \<turnstile>\<^sub>w\<^sub>f \<Theta>" using valid.simps wfC_wf wfG_wf assms by metis
- ultimately show ?thesis
- using theta_flip_eq G_cons_flip_fresh3[of x \<Gamma> z b c] assms fresh_Pair flip_commute by metis
-qed
-
-lemma valid_wf_all:
- assumes " \<Theta>; \<B>; (z0,b,c0)#\<^sub>\<Gamma>G \<Turnstile> c"
- shows "wfG \<Theta> \<B> G" and "wfC \<Theta> \<B> ((z0,b,c0)#\<^sub>\<Gamma>G) c" and "atom z0 \<sharp> G"
- using valid.simps wfC_wf wfG_cons assms by metis+
-
-lemma valid_wfT:
- fixes z::x
- assumes " \<Theta>; \<B>; (z0,b,c0[z::=V_var z0]\<^sub>v)#\<^sub>\<Gamma>G \<Turnstile> c[z::=V_var z0]\<^sub>v" and "atom z0 \<sharp> (\<Theta>, \<B>, G,c,c0)"
- shows "\<Theta>; \<B>; G \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b | c0 \<rbrace>" and "\<Theta>; \<B>; G \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b | c \<rbrace>"
-proof -
- have "atom z0 \<sharp> c0" using assms fresh_Pair by auto
- moreover have *:" \<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f (z0,b,c0[z::=V_var z0]\<^sub>c\<^sub>v)#\<^sub>\<Gamma>G" using valid_wf_all wfX_wfY assms subst_v_c_def by metis
- ultimately show wft: "\<Theta>; \<B>; G \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b | c0 \<rbrace>" using wfG_wfT[OF *] by auto
-
- have "atom z0 \<sharp> c" using assms fresh_Pair by auto
- moreover have wfc: "\<Theta>; \<B>; (z0,b,c0[z::=V_var z0]\<^sub>v)#\<^sub>\<Gamma>G \<turnstile>\<^sub>w\<^sub>f c[z::=V_var z0]\<^sub>v" using valid_wf_all assms by metis
- have "\<Theta>; \<B>; G \<turnstile>\<^sub>w\<^sub>f \<lbrace> z0 : b | c[z::=V_var z0]\<^sub>v \<rbrace>" proof
- show \<open>atom z0 \<sharp> (\<Theta>, \<B>, G)\<close> using assms fresh_prodN by simp
- show \<open> \<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f b \<close> using wft wfT_wfB by force
- show \<open> \<Theta>; \<B>; (z0, b, TRUE) #\<^sub>\<Gamma> G \<turnstile>\<^sub>w\<^sub>f c[z::=[ z0 ]\<^sup>v]\<^sub>v \<close> using wfc wfC_replace_inside[OF wfc, of GNil z0 b "c0[z::=[ z0 ]\<^sup>v]\<^sub>v" G C_true] wfC_trueI
- append_g.simps
- by (metis "local.*" wfG_elim2 wf_trans(2))
- qed
- moreover have "\<lbrace> z0 : b | c[z::=V_var z0]\<^sub>v \<rbrace> = \<lbrace> z : b | c \<rbrace>" using \<open>atom z0 \<sharp> c0\<close> \<tau>.eq_iff Abs1_eq_iff(3)
- using calculation(1) subst_v_c_def by auto
- ultimately show "\<Theta>; \<B>; G \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b | c \<rbrace>" by auto
-qed
-
-lemma valid_flip:
- fixes c::c and z::x and z0::x and xx2::x
- assumes " \<Theta>; \<B>; (xx2, b, c0[z0::=V_var xx2]\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<Turnstile> c[z::=V_var xx2]\<^sub>v" and
- "atom xx2 \<sharp> (c0,\<Gamma>,c,z) " and "atom z0 \<sharp> (\<Gamma>,c,z)"
- shows " \<Theta>; \<B>; (z0, b, c0) #\<^sub>\<Gamma> \<Gamma> \<Turnstile> c[z::=V_var z0]\<^sub>v"
-proof -
-
- have " \<turnstile>\<^sub>w\<^sub>f \<Theta>" using assms valid_wf_all wfX_wfY by metis
- hence " \<Theta> ; \<B> ; (xx2 \<leftrightarrow> z0 ) \<bullet> ((xx2, b, c0[z0::=V_var xx2]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) \<Turnstile> ((xx2 \<leftrightarrow> z0 ) \<bullet> c[z::=V_var xx2]\<^sub>v)"
- using valid.eqvt True_eqvt assms beta_flip_eq theta_flip_eq by metis
- hence " \<Theta>; \<B>; (((xx2 \<leftrightarrow> z0 ) \<bullet> xx2, b, (xx2 \<leftrightarrow> z0 ) \<bullet> c0[z0::=V_var xx2]\<^sub>v) #\<^sub>\<Gamma> (xx2 \<leftrightarrow> z0 ) \<bullet> \<Gamma>) \<Turnstile> ((xx2 \<leftrightarrow> z0 ) \<bullet>(c[z::=V_var xx2]\<^sub>v))"
- using G_cons_flip[of xx2 z0 xx2 b "c0[z0::=V_var xx2]\<^sub>v" \<Gamma>] by auto
- moreover have "(xx2 \<leftrightarrow> z0 ) \<bullet> xx2 = z0" by simp
- moreover have "(xx2 \<leftrightarrow> z0 ) \<bullet> c0[z0::=V_var xx2]\<^sub>v = c0"
- using assms subst_cv_v_flip[of xx2 c0 z0 "V_var z0"] assms fresh_prod4 by auto
- moreover have "(xx2 \<leftrightarrow> z0 ) \<bullet> \<Gamma> = \<Gamma>" proof -
- have "atom xx2 \<sharp> \<Gamma>" using assms by auto
- moreover have "atom z0 \<sharp> \<Gamma>" using assms by auto
- ultimately show ?thesis using flip_fresh_fresh by auto
- qed
- moreover have "(xx2 \<leftrightarrow> z0 ) \<bullet> (c[z::=V_var xx2]\<^sub>v) =c[z::=V_var z0]\<^sub>v"
- using subst_cv_v_flip3 assms by simp
- ultimately show ?thesis by auto
-qed
-
-lemma subtype_valid:
- assumes "\<Theta>; \<B>; \<Gamma> \<turnstile> t1 \<lesssim> t2" and "atom y \<sharp> \<Gamma>" and "t1 = \<lbrace> z1 : b | c1 \<rbrace>" and "t2 = \<lbrace> z2 : b | c2 \<rbrace>"
- shows "\<Theta>; \<B>; ((y, b, c1[z1::=V_var y]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) \<Turnstile> c2[z2::=V_var y]\<^sub>v"
- using assms proof(nominal_induct t2 avoiding: y rule: subtype.strong_induct)
- case (subtype_baseI x \<Theta> \<B> \<Gamma> z c z' c' ba)
-
- hence "(x \<leftrightarrow> y) \<bullet> \<Theta> ; (x \<leftrightarrow> y) \<bullet> \<B> ; (x \<leftrightarrow> y) \<bullet> ((x, ba, c[z::=[ x ]\<^sup>v]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) \<Turnstile> (x \<leftrightarrow> y) \<bullet> c'[z'::=[ x ]\<^sup>v]\<^sub>v" using valid.eqvt
- using permute_boolI by blast
- moreover have " \<turnstile>\<^sub>w\<^sub>f \<Theta>" using valid.simps wfC_wf wfG_wf subtype_baseI by metis
- ultimately have "\<Theta>; \<B>; ((y, ba, (x \<leftrightarrow> y) \<bullet> c[z::=[ x ]\<^sup>v]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) \<Turnstile> (x \<leftrightarrow> y) \<bullet> c'[z'::=[ x ]\<^sup>v]\<^sub>v"
- using subtype_baseI theta_flip_eq beta_flip_eq \<tau>.eq_iff G_cons_flip_fresh3[of y \<Gamma> x ba] by (metis flip_commute)
- moreover have "(x \<leftrightarrow> y) \<bullet> c[z::=[ x ]\<^sup>v]\<^sub>v = c1[z1::=[ y ]\<^sup>v]\<^sub>v"
- by (metis subtype_baseI permute_flip_cancel subst_cv_id subst_cv_v_flip3 subst_cv_var_flip type_eq_subst_eq wfT_fresh_c subst_v_c_def)
- moreover have "(x \<leftrightarrow> y) \<bullet> c'[z'::=[ x ]\<^sup>v]\<^sub>v = c2[z2::=[ y ]\<^sup>v]\<^sub>v"
- by (metis subtype_baseI permute_flip_cancel subst_cv_id subst_cv_v_flip3 subst_cv_var_flip type_eq_subst_eq wfT_fresh_c subst_v_c_def)
-
- ultimately show ?case using subtype_baseI \<tau>.eq_iff by metis
-qed
-
-lemma subtype_valid_simple:
- assumes "\<Theta>; \<B>; \<Gamma> \<turnstile> t1 \<lesssim> t2" and "atom z \<sharp> \<Gamma>" and "t1 = \<lbrace> z : b | c1 \<rbrace>" and "t2 = \<lbrace> z : b | c2 \<rbrace>"
- shows "\<Theta>; \<B>; ((z, b, c1) #\<^sub>\<Gamma> \<Gamma>) \<Turnstile> c2"
- using subst_v_c_def subst_v_id assms subtype_valid[OF assms] by simp
-
-lemma obtain_for_t_with_fresh:
- assumes "atom x \<sharp> t"
- shows "\<exists>b c. t = \<lbrace> x : b | c \<rbrace>"
-proof -
- obtain z1 b1 c1 where *:" t = \<lbrace> z1 : b1 | c1 \<rbrace> \<and> atom z1 \<sharp> t" using obtain_fresh_z by metis
- then have "t = (x \<leftrightarrow> z1) \<bullet> t" using flip_fresh_fresh assms by metis
- also have "... = \<lbrace> (x \<leftrightarrow> z1) \<bullet> z1 : (x \<leftrightarrow> z1) \<bullet> b1 | (x \<leftrightarrow> z1) \<bullet> c1 \<rbrace>" using * assms by simp
- also have "... = \<lbrace> x : b1 | (x \<leftrightarrow> z1) \<bullet> c1 \<rbrace>" using * assms by auto
- finally show ?thesis by auto
-qed
-
-lemma subtype_trans:
- assumes "\<Theta>; \<B>; \<Gamma> \<turnstile> \<tau>1 \<lesssim> \<tau>2" and "\<Theta>; \<B>; \<Gamma> \<turnstile> \<tau>2 \<lesssim> \<tau>3"
- shows "\<Theta>; \<B>; \<Gamma> \<turnstile> \<tau>1 \<lesssim> \<tau>3"
- using assms proof(nominal_induct avoiding: \<tau>3 rule: subtype.strong_induct)
- case (subtype_baseI x \<Theta> \<B> \<Gamma> z c z' c' b)
- hence "b_of \<tau>3 = b" using subtype_eq_base2 b_of.simps by metis
- then obtain z'' c'' where t3: "\<tau>3 = \<lbrace> z'' : b | c''\<rbrace> \<and> atom z'' \<sharp> x"
- using obtain_fresh_z2 by metis
- hence xf: "atom x \<sharp> (z'', c'')" using fresh_prodN subtype_baseI \<tau>.fresh by auto
- have "\<Theta>; \<B>; \<Gamma> \<turnstile> \<lbrace> z : b | c \<rbrace> \<lesssim> \<lbrace> z'' : b | c''\<rbrace>"
- proof(rule Typing.subtype_baseI)
- show \<open>atom x \<sharp> (\<Theta>, \<B>, \<Gamma>, z, c, z'', c'')\<close> using t3 fresh_prodN subtype_baseI xf by simp
- show \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b | c \<rbrace> \<close> using subtype_baseI by auto
- show \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> z'' : b | c'' \<rbrace> \<close> using subtype_baseI t3 subtype_elims by metis
- have " \<Theta>; \<B>; (x, b, c'[z'::=[ x ]\<^sup>v]\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<Turnstile> c''[z''::=[ x ]\<^sup>v]\<^sub>v "
- using subtype_valid[OF \<open>\<Theta>; \<B>; \<Gamma> \<turnstile> \<lbrace> z' : b | c' \<rbrace> \<lesssim> \<tau>3\<close> , of x z' b c' z'' c''] subtype_baseI
- t3 by simp
- thus \<open>\<Theta>; \<B>; (x, b, c[z::=[ x ]\<^sup>v]\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<Turnstile> c''[z''::=[ x ]\<^sup>v]\<^sub>v \<close>
- using valid_trans_full[of \<Theta> \<B> x b c z \<Gamma> c' z' c'' z'' ] subtype_baseI t3 by simp
- qed
- thus ?case using t3 by simp
-qed
-
-lemma subtype_eq_e:
- assumes "\<forall>i s1 s2 G. wfG P \<B> G \<and> wfI P G i \<and> eval_e i e1 s1 \<and> eval_e i e2 s2 \<longrightarrow> s1 = s2" and "atom z1 \<sharp> e1" and "atom z2 \<sharp> e2" and "atom z1 \<sharp> \<Gamma>" and "atom z2 \<sharp> \<Gamma>"
- and "wfCE P \<B> \<Gamma> e1 b" and "wfCE P \<B> \<Gamma> e2 b"
- shows "P; \<B>; \<Gamma> \<turnstile> \<lbrace> z1 : b | CE_val (V_var z1) == e1 \<rbrace> \<lesssim> (\<lbrace> z2 : b | CE_val (V_var z2) == e2 \<rbrace>)"
-proof -
-
- have "wfCE P \<B> \<Gamma> e1 b" and "wfCE P \<B> \<Gamma> e2 b" using assms by auto
-
- have wst1: "wfT P \<B> \<Gamma> (\<lbrace> z1 : b | CE_val (V_var z1) == e1 \<rbrace>)"
- using wfC_e_eq wfTI assms wfX_wfB wfG_fresh_x
- by (simp add: wfT_e_eq)
-
- moreover have wst2:"wfT P \<B> \<Gamma> (\<lbrace> z2 : b | CE_val (V_var z2) == e2 \<rbrace>)"
- using wfC_e_eq wfX_wfB wfTI assms wfG_fresh_x
- by (simp add: wfT_e_eq)
-
- moreover obtain x::x where xf: "atom x \<sharp> (P, \<B> , z1, CE_val (V_var z1) == e1 , z2, CE_val (V_var z2) == e2 , \<Gamma>)" using obtain_fresh by blast
- moreover have vld: "P; \<B> ; (x, b, (CE_val (V_var z1) == e1 )[z1::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<Turnstile> (CE_val (V_var z2) == e2 )[z2::=V_var x]\<^sub>v " (is "P; \<B> ; ?G \<Turnstile> ?c")
- proof -
-
- have wbg: "P; \<B> \<turnstile>\<^sub>w\<^sub>f ?G \<and> P ; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<and> toSet \<Gamma> \<subseteq> toSet ?G" proof -
- have "P; \<B> \<turnstile>\<^sub>w\<^sub>f ?G" proof(rule wfG_consI)
- show "P; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>" using assms wfX_wfY by metis
- show "atom x \<sharp> \<Gamma>" using xf by auto
- show "P; \<B> \<turnstile>\<^sub>w\<^sub>f b " using assms(6) wfX_wfB by auto
- show "P; \<B> ; (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f (CE_val (V_var z1) == e1 )[z1::=V_var x]\<^sub>v "
- using wfC_e_eq[OF assms(6)] wf_subst(2)
- by (simp add: \<open>atom x \<sharp> \<Gamma>\<close> assms(2) subst_v_c_def)
- qed
- moreover hence "P; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>" using wfG_elims by metis
- ultimately show ?thesis using toSet.simps by auto
- qed
-
- have wsc: "wfC P \<B> ?G ?c" proof -
- have "wfCE P \<B> ?G (CE_val (V_var x)) b" proof
- show \<open> P; \<B> ; (x, b, (CE_val (V_var z1) == e1 )[z1::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f V_var x : b \<close> using wfV_varI lookup.simps wbg by auto
- qed
- moreover have "wfCE P \<B> ?G e2 b" using wf_weakening assms wbg by metis
- ultimately have "wfC P \<B> ?G (CE_val (V_var x) == e2 )" using wfC_eqI by simp
- thus ?thesis using subst_cv.simps(6) \<open>atom z2 \<sharp> e2\<close> subst_v_c_def by simp
- qed
-
- moreover have "\<forall>i. wfI P ?G i \<and> is_satis_g i ?G \<longrightarrow> is_satis i ?c" proof(rule allI , rule impI)
- fix i
- assume as: "wfI P ?G i \<and> is_satis_g i ?G"
- hence "is_satis i ((CE_val (V_var z1) == e1 )[z1::=V_var x]\<^sub>v)"
- by (simp add: is_satis_g.simps(2))
- hence "is_satis i (CE_val (V_var x) == e1 )" using subst_cv.simps assms subst_v_c_def by auto
- then obtain s1 and s2 where *:"eval_e i (CE_val (V_var x)) s1 \<and> eval_e i e1 s2 \<and> s1=s2" using is_satis.simps eval_c_elims by metis
- moreover hence "eval_e i e2 s1" proof -
- have **:"wfI P ?G i" using as by auto
- moreover have "wfCE P \<B> ?G e1 b \<and> wfCE P \<B> ?G e2 b" using assms xf wf_weakening wbg by metis
- moreover then obtain s2' where "eval_e i e2 s2'" using assms wfI_wfCE_eval_e ** by metis
- ultimately show ?thesis using * assms(1) wfX_wfY by metis
- qed
- ultimately have "is_satis i (CE_val (V_var x) == e2 )" using is_satis.simps eval_c_eqI by force
- thus "is_satis i ((CE_val (V_var z2) == e2 )[z2::=V_var x]\<^sub>v)" using is_satis.simps eval_c_eqI assms subst_cv.simps subst_v_c_def by auto
- qed
- ultimately show ?thesis using valid.simps by simp
- qed
- moreover have "atom x \<sharp> (P, \<B>, \<Gamma>, z1 , CE_val (V_var z1) == e1, z2, CE_val (V_var z2) == e2 ) "
- unfolding fresh_prodN using xf fresh_prod7 \<tau>.fresh by fast
- ultimately show ?thesis using subtype_baseI[OF _ wst1 wst2 vld] xf by simp
-qed
-
-lemma subtype_eq_e_nil:
- assumes "\<forall>i s1 s2 G. wfG P \<B> G \<and> wfI P G i \<and> eval_e i e1 s1 \<and> eval_e i e2 s2 \<longrightarrow> s1 = s2" and "supp e1 = {}" and "supp e2 = {}" and "wfTh P"
- and "wfCE P \<B> GNil e1 b" and "wfCE P \<B> GNil e2 b" and "atom z1 \<sharp> GNil" and "atom z2 \<sharp> GNil"
- shows "P; \<B> ; GNil \<turnstile> \<lbrace> z1 : b | CE_val (V_var z1) == e1 \<rbrace> \<lesssim> (\<lbrace> z2 : b | CE_val (V_var z2) == e2 \<rbrace>)"
- apply(rule subtype_eq_e,auto simp add: assms e.fresh)
- using assms fresh_def e.fresh supp_GNil by metis+
-
-lemma subtype_gnil_fst_aux:
- assumes "supp v\<^sub>1 = {}" and "supp (V_pair v\<^sub>1 v\<^sub>2) = {}" and "wfTh P" and "wfCE P \<B> GNil (CE_val v\<^sub>1) b" and "wfCE P \<B> GNil (CE_fst [V_pair v\<^sub>1 v\<^sub>2]\<^sup>c\<^sup>e) b" and
- "wfCE P \<B> GNil (CE_val v\<^sub>2) b2" and "atom z1 \<sharp> GNil" and "atom z2 \<sharp> GNil"
- shows "P; \<B> ; GNil \<turnstile> (\<lbrace> z1 : b | CE_val (V_var z1) == CE_val v\<^sub>1 \<rbrace>) \<lesssim> (\<lbrace> z2 : b | CE_val (V_var z2) == CE_fst [V_pair v\<^sub>1 v\<^sub>2]\<^sup>c\<^sup>e \<rbrace>)"
-proof -
- have "\<forall>i s1 s2 G . wfG P \<B> G \<and> wfI P G i \<and> eval_e i ( CE_val v\<^sub>1) s1 \<and> eval_e i (CE_fst [V_pair v\<^sub>1 v\<^sub>2]\<^sup>c\<^sup>e) s2 \<longrightarrow> s1 = s2" proof(rule+)
- fix i s1 s2 G
- assume as: "wfG P \<B> G \<and> wfI P G i \<and> eval_e i (CE_val v\<^sub>1) s1 \<and> eval_e i (CE_fst [V_pair v\<^sub>1 v\<^sub>2]\<^sup>c\<^sup>e) s2"
- hence "wfCE P \<B> G (CE_val v\<^sub>2) b2" using assms wf_weakening
- by (metis empty_subsetI toSet.simps(1))
- then obtain s3 where "eval_e i (CE_val v\<^sub>2) s3" using wfI_wfCE_eval_e as by metis
- hence "eval_v i ((V_pair v\<^sub>1 v\<^sub>2)) (SPair s1 s3)"
- by (meson as eval_e_elims(1) eval_v_pairI)
- hence "eval_e i (CE_fst [V_pair v\<^sub>1 v\<^sub>2]\<^sup>c\<^sup>e) s1" using eval_e_fstI eval_e_valI by metis
- show "s1 = s2" using as eval_e_uniqueness
- using \<open>eval_e i (CE_fst [V_pair v\<^sub>1 v\<^sub>2]\<^sup>c\<^sup>e) s1\<close> by auto
- qed
- thus ?thesis using subtype_eq_e_nil ce.supp assms by auto
-qed
-
-lemma subtype_gnil_snd_aux:
- assumes "supp v\<^sub>2 = {}" and "supp (V_pair v\<^sub>1 v\<^sub>2) = {}" and "wfTh P" and "wfCE P \<B> GNil (CE_val v\<^sub>2) b" and
- "wfCE P \<B> GNil (CE_snd [(V_pair v\<^sub>1 v\<^sub>2)]\<^sup>c\<^sup>e) b" and
- "wfCE P \<B> GNil (CE_val v\<^sub>1) b2" and "atom z1 \<sharp> GNil" and "atom z2 \<sharp> GNil"
- shows "P; \<B> ; GNil \<turnstile> (\<lbrace> z1 : b | CE_val (V_var z1) == CE_val v\<^sub>2 \<rbrace>) \<lesssim> (\<lbrace> z2 : b | CE_val (V_var z2) == CE_snd [(V_pair v\<^sub>1 v\<^sub>2)]\<^sup>c\<^sup>e \<rbrace>)"
-proof -
- have "\<forall>i s1 s2 G. wfG P \<B> G \<and> wfI P G i \<and> eval_e i ( CE_val v\<^sub>2) s1 \<and> eval_e i (CE_snd [(V_pair v\<^sub>1 v\<^sub>2)]\<^sup>c\<^sup>e) s2 \<longrightarrow> s1 = s2" proof(rule+)
- fix i s1 s2 G
- assume as: " wfG P \<B> G \<and> wfI P G i \<and> eval_e i (CE_val v\<^sub>2) s1 \<and> eval_e i (CE_snd [(V_pair v\<^sub>1 v\<^sub>2)]\<^sup>c\<^sup>e) s2"
- hence "wfCE P \<B> G (CE_val v\<^sub>1) b2" using assms wf_weakening
- by (metis empty_subsetI toSet.simps(1))
- then obtain s3 where "eval_e i (CE_val v\<^sub>1) s3" using wfI_wfCE_eval_e as by metis
- hence "eval_v i ((V_pair v\<^sub>1 v\<^sub>2)) (SPair s3 s1)"
- by (meson as eval_e_elims eval_v_pairI)
- hence "eval_e i (CE_snd [(V_pair v\<^sub>1 v\<^sub>2)]\<^sup>c\<^sup>e) s1" using eval_e_sndI eval_e_valI by metis
- show "s1 = s2" using as eval_e_uniqueness
- using \<open>eval_e i (CE_snd [V_pair v\<^sub>1 v\<^sub>2]\<^sup>c\<^sup>e) s1\<close> by auto
- qed
- thus ?thesis using assms subtype_eq_e_nil by (simp add: ce.supp ce.supp)
-qed
-
-lemma subtype_gnil_fst:
- assumes "\<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f [#1[[v\<^sub>1,v\<^sub>2]\<^sup>v]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e : b"
- shows "\<Theta> ; {||} ; GNil \<turnstile> (\<lbrace> z\<^sub>1 : b | [[z\<^sub>1]\<^sup>v]\<^sup>c\<^sup>e == [v\<^sub>1]\<^sup>c\<^sup>e \<rbrace>) \<lesssim>
- (\<lbrace> z\<^sub>2 : b | [[z\<^sub>2]\<^sup>v]\<^sup>c\<^sup>e == [#1[[v\<^sub>1, v\<^sub>2]\<^sup>v]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e \<rbrace>)"
-proof -
- obtain b2 where **:" \<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f V_pair v\<^sub>1 v\<^sub>2 : B_pair b b2" using wfCE_elims(4)[OF assms ] wfCE_elims by metis
- obtain b1' b2' where *:"B_pair b b2 = B_pair b1' b2' \<and> \<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f v\<^sub>1 : b1' \<and> \<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f v\<^sub>2 : b2'"
- using wfV_elims(3)[OF **] by metis
-
- show ?thesis proof(rule subtype_gnil_fst_aux)
- show \<open>supp v\<^sub>1 = {}\<close> using * wfV_supp_nil by auto
- show \<open>supp (V_pair v\<^sub>1 v\<^sub>2) = {}\<close> using ** wfV_supp_nil e.supp by metis
- show \<open>\<turnstile>\<^sub>w\<^sub>f \<Theta>\<close> using assms wfX_wfY by metis
- show \<open>\<Theta>; {||}; GNil \<turnstile>\<^sub>w\<^sub>f CE_val v\<^sub>1 : b \<close> using wfCE_valI "*" by auto
- show \<open>\<Theta>; {||}; GNil \<turnstile>\<^sub>w\<^sub>f CE_fst [V_pair v\<^sub>1 v\<^sub>2]\<^sup>c\<^sup>e : b \<close> using assms by auto
- show \<open>\<Theta>; {||}; GNil \<turnstile>\<^sub>w\<^sub>f CE_val v\<^sub>2 : b2 \<close>using wfCE_valI "*" by auto
- show \<open>atom z\<^sub>1 \<sharp> GNil\<close> using fresh_GNil by metis
- show \<open>atom z\<^sub>2 \<sharp> GNil\<close> using fresh_GNil by metis
- qed
-qed
-
-lemma subtype_gnil_snd:
- assumes "wfCE P {||} GNil (CE_snd ([V_pair v\<^sub>1 v\<^sub>2]\<^sup>c\<^sup>e)) b"
- shows "P ; {||} ; GNil \<turnstile> (\<lbrace> z1 : b | CE_val (V_var z1) == CE_val v\<^sub>2 \<rbrace>) \<lesssim> (\<lbrace> z2 : b | CE_val (V_var z2) == CE_snd [(V_pair v\<^sub>1 v\<^sub>2)]\<^sup>c\<^sup>e \<rbrace>)"
-proof -
- obtain b1 where **:" P ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f V_pair v\<^sub>1 v\<^sub>2 : B_pair b1 b " using wfCE_elims assms by metis
- obtain b1' b2' where *:"B_pair b1 b = B_pair b1' b2' \<and> P ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f v\<^sub>1 : b1' \<and> P ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f v\<^sub>2 : b2'" using wfV_elims(3)[OF **] by metis
-
- show ?thesis proof(rule subtype_gnil_snd_aux)
- show \<open>supp v\<^sub>2 = {}\<close> using * wfV_supp_nil by auto
- show \<open>supp (V_pair v\<^sub>1 v\<^sub>2) = {}\<close> using ** wfV_supp_nil e.supp by metis
- show \<open>\<turnstile>\<^sub>w\<^sub>f P\<close> using assms wfX_wfY by metis
- show \<open>P; {||}; GNil \<turnstile>\<^sub>w\<^sub>f CE_val v\<^sub>1 : b1 \<close> using wfCE_valI "*" by simp
- show \<open>P; {||}; GNil \<turnstile>\<^sub>w\<^sub>f CE_snd [(V_pair v\<^sub>1 v\<^sub>2)]\<^sup>c\<^sup>e : b \<close> using assms by auto
- show \<open>P; {||}; GNil \<turnstile>\<^sub>w\<^sub>f CE_val v\<^sub>2 : b \<close>using wfCE_valI "*" by simp
- show \<open>atom z1 \<sharp> GNil\<close> using fresh_GNil by metis
- show \<open>atom z2 \<sharp> GNil\<close> using fresh_GNil by metis
- qed
-qed
-
-lemma subtype_fresh_tau:
- fixes x::x
- assumes "atom x \<sharp> t1" and "atom x \<sharp> \<Gamma>" and "P; \<B>; \<Gamma> \<turnstile> t1 \<lesssim> t2"
- shows "atom x \<sharp> t2"
-proof -
- have wfg: "P; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>" using subtype_wf wfX_wfY assms by metis
- have wft: "wfT P \<B> \<Gamma> t2" using subtype_wf wfX_wfY assms by blast
- hence "supp t2 \<subseteq> atom_dom \<Gamma> \<union> supp \<B>" using wf_supp
- using atom_dom.simps by auto
- moreover have "atom x \<notin> atom_dom \<Gamma>" using \<open>atom x \<sharp> \<Gamma>\<close> wfG_atoms_supp_eq wfg fresh_def by blast
- ultimately show "atom x \<sharp> t2" using fresh_def
- by (metis Un_iff contra_subsetD x_not_in_b_set)
-qed
-
-lemma subtype_if_simp:
- assumes "wfT P \<B> GNil (\<lbrace> z1 : b | CE_val (V_lit l ) == CE_val (V_lit l) IMP c[z::=V_var z1]\<^sub>v \<rbrace>)" and
- "wfT P \<B> GNil (\<lbrace> z : b | c \<rbrace>)" and "atom z1 \<sharp> c"
- shows "P; \<B> ; GNil \<turnstile> (\<lbrace> z1 : b | CE_val (V_lit l) == CE_val (V_lit l) IMP c[z::=V_var z1]\<^sub>v \<rbrace>) \<lesssim> (\<lbrace> z : b | c \<rbrace>)"
-proof -
- obtain x::x where xx: "atom x \<sharp> ( P , \<B> , z1, CE_val (V_lit l) == CE_val (V_lit l) IMP c[z::=V_var z1]\<^sub>v , z, c, GNil)" using obtain_fresh_z by blast
- hence xx2: " atom x \<sharp> (CE_val (V_lit l) == CE_val (V_lit l) IMP c[z::=V_var z1]\<^sub>v , c, GNil)" using fresh_prod7 fresh_prod3 by fast
- have *:"P; \<B> ; (x, b, (CE_val (V_lit l) == CE_val (V_lit l) IMP c[z::=V_var z1]\<^sub>v )[z1::=V_var x]\<^sub>v) #\<^sub>\<Gamma> GNil \<Turnstile> c[z::=V_var x]\<^sub>v " (is "P; \<B> ; ?G \<Turnstile> ?c" ) proof -
- have "wfC P \<B> ?G ?c" using wfT_wfC_cons[OF assms(1) assms(2),of x] xx fresh_prod5 fresh_prod3 subst_v_c_def by metis
- moreover have "(\<forall>i. wfI P ?G i \<and> is_satis_g i ?G \<longrightarrow> is_satis i ?c)" proof(rule allI, rule impI)
- fix i
- assume as1: "wfI P ?G i \<and> is_satis_g i ?G"
- have "((CE_val (V_lit l) == CE_val (V_lit l) IMP c[z::=V_var z1]\<^sub>v )[z1::=V_var x]\<^sub>v) = ((CE_val (V_lit l) == CE_val (V_lit l) IMP c[z::=V_var x]\<^sub>v ))"
- using assms subst_v_c_def by auto
- hence "is_satis i ((CE_val (V_lit l) == CE_val (V_lit l) IMP c[z::=V_var x]\<^sub>v ))" using is_satis_g.simps as1 by presburger
- moreover have "is_satis i ((CE_val (V_lit l) == CE_val (V_lit l)))" using is_satis.simps eval_c_eqI[of i "(CE_val (V_lit l))" "eval_l l"] eval_e_uniqueness
- eval_e_valI eval_v_litI by metis
- ultimately show "is_satis i ?c" using is_satis_mp[of i] by metis
- qed
- ultimately show ?thesis using valid.simps by simp
- qed
- moreover have "atom x \<sharp> (P, \<B>, GNil, z1 , CE_val (V_lit l) == CE_val (V_lit l) IMP c[z::=V_var z1]\<^sub>v , z, c )"
- unfolding fresh_prod5 \<tau>.fresh using xx fresh_prodN x_fresh_b by metis
- ultimately show ?thesis using subtype_baseI assms xx xx2 by metis
-qed
-
-lemma subtype_if:
- assumes "P; \<B>; \<Gamma> \<turnstile> \<lbrace> z : b | c \<rbrace> \<lesssim> \<lbrace> z' : b | c' \<rbrace>" and
- "wfT P \<B> \<Gamma> (\<lbrace> z1 : b | CE_val v == CE_val (V_lit l) IMP c[z::=V_var z1]\<^sub>v \<rbrace>)" and
- "wfT P \<B> \<Gamma> (\<lbrace> z2 : b | CE_val v == CE_val (V_lit l) IMP c'[z'::=V_var z2]\<^sub>v \<rbrace>)" and
- "atom z1 \<sharp> v" and "atom z \<sharp> \<Gamma>" and "atom z1 \<sharp> c" and "atom z2 \<sharp> c'" and "atom z2 \<sharp> v"
- shows "P; \<B> ; \<Gamma> \<turnstile> \<lbrace> z1 : b | CE_val v == CE_val (V_lit l) IMP c[z::=V_var z1]\<^sub>v \<rbrace> \<lesssim> \<lbrace> z2 : b | CE_val v == CE_val (V_lit l) IMP c'[z'::=V_var z2]\<^sub>v \<rbrace>"
-proof -
- obtain x::x where xx: "atom x \<sharp> (P,\<B>,z,c,z', c', z1, CE_val v == CE_val (V_lit l) IMP c[z::=V_var z1]\<^sub>v , z2, CE_val v == CE_val (V_lit l) IMP c'[z'::=V_var z2]\<^sub>v , \<Gamma>)"
- using obtain_fresh_z by blast
- hence xf: "atom x \<sharp> (z, c, z', c', \<Gamma>)" by simp
- have xf2: "atom x \<sharp> (z1, CE_val v == CE_val (V_lit l) IMP c[z::=V_var z1]\<^sub>v , z2, CE_val v == CE_val (V_lit l) IMP c'[z'::=V_var z2]\<^sub>v , \<Gamma>)"
- using xx fresh_prod4 fresh_prodN by metis
-
- moreover have "P; \<B> ; (x, b, (CE_val v == CE_val (V_lit l) IMP c[z::=V_var z1]\<^sub>v )[z1::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<Turnstile> (CE_val v == CE_val (V_lit l) IMP c'[z'::=V_var z2]\<^sub>v )[z2::=V_var x]\<^sub>v"
- (is "P; \<B> ; ?G \<Turnstile> ?c" )
- proof -
- have wbc: "wfC P \<B> ?G ?c" using assms xx fresh_prod4 fresh_prod2 wfT_wfC_cons assms subst_v_c_def by metis
-
- moreover have "\<forall>i. wfI P ?G i \<and> is_satis_g i ?G \<longrightarrow> is_satis i ?c" proof(rule allI, rule impI)
- fix i
- assume a1: "wfI P ?G i \<and> is_satis_g i ?G"
-
- have *: " is_satis i ((CE_val v == CE_val (V_lit l))) \<longrightarrow> is_satis i ((c'[z'::=V_var z2]\<^sub>v )[z2::=V_var x]\<^sub>v)"
- proof
- assume a2: "is_satis i ((CE_val v == CE_val (V_lit l)))"
-
- have "is_satis i ((CE_val v == CE_val (V_lit l) IMP (c[z::=V_var z1]\<^sub>v ))[z1::=V_var x]\<^sub>v)"
- using a1 is_satis_g.simps by simp
- moreover have "((CE_val v == CE_val (V_lit l) IMP (c[z::=V_var z1]\<^sub>v ))[z1::=V_var x]\<^sub>v) = (CE_val v == CE_val (V_lit l) IMP ((c[z::=V_var z1]\<^sub>v )[z1::=V_var x]\<^sub>v))"
- using assms subst_v_c_def by simp
- ultimately have "is_satis i (CE_val v == CE_val (V_lit l) IMP ((c[z::=V_var z1]\<^sub>v )[z1::=V_var x]\<^sub>v))" by argo
-
- hence "is_satis i ((c[z::=V_var z1]\<^sub>v )[z1::=V_var x]\<^sub>v)" using a2 is_satis_mp by auto
- moreover have "((c[z::=V_var z1]\<^sub>v )[z1::=V_var x]\<^sub>v) = ((c[z::=V_var x]\<^sub>v ))" using assms by auto
- ultimately have "is_satis i ((c[z::=V_var x]\<^sub>v ))" using a2 is_satis.simps by auto
-
- hence "is_satis_g i ((x,b,(c[z::=V_var x]\<^sub>v )) #\<^sub>\<Gamma> \<Gamma>)" using a1 is_satis_g.simps by meson
- moreover have "wfI P ((x,b,(c[z::=V_var x]\<^sub>v )) #\<^sub>\<Gamma> \<Gamma>) i" proof -
- obtain s where "Some s = i x \<and> wfRCV P s b \<and> wfI P \<Gamma> i" using wfI_def a1 by auto
- thus ?thesis using wfI_def by auto
- qed
- ultimately have "is_satis i ((c'[z'::=V_var x]\<^sub>v))" using subtype_valid assms(1) xf valid.simps by simp
-
- moreover have "(c'[z'::=V_var x]\<^sub>v) = ((c'[z'::=V_var z2]\<^sub>v )[z2::=V_var x]\<^sub>v)" using assms by auto
- ultimately show "is_satis i ((c'[z'::=V_var z2]\<^sub>v )[z2::=V_var x]\<^sub>v)" by auto
- qed
-
- moreover have "?c = ((CE_val v == CE_val (V_lit l)) IMP ((c'[z'::=V_var z2]\<^sub>v )[z2::=V_var x]\<^sub>v))"
- using assms subst_v_c_def by simp
-
- moreover have "\<exists>b1 b2. eval_c i (CE_val v == CE_val (V_lit l) ) b1 \<and>
- eval_c i c'[z'::=V_var z2]\<^sub>v[z2::=V_var x]\<^sub>v b2" proof -
-
- have "wfC P \<B> ?G (CE_val v == CE_val (V_lit l))" using wbc wfC_elims(7) assms subst_cv.simps subst_v_c_def by fastforce
-
- moreover have "wfC P \<B> ?G (c'[z'::=V_var z2]\<^sub>v[z2::=V_var x]\<^sub>v)" proof(rule wfT_wfC_cons)
- show \<open> P; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> z1 : b | CE_val v == CE_val (V_lit l) IMP (c[z::=V_var z1]\<^sub>v) \<rbrace> \<close> using assms subst_v_c_def by auto
- have " \<lbrace> z2 : b | c'[z'::=V_var z2]\<^sub>v \<rbrace> = \<lbrace> z' : b | c' \<rbrace>" using assms subst_v_c_def by auto
- thus \<open> P; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> z2 : b | c'[z'::=V_var z2]\<^sub>v \<rbrace> \<close> using assms subtype_elims by metis
- show \<open>atom x \<sharp> (CE_val v == CE_val (V_lit l) IMP c[z::=V_var z1]\<^sub>v , c'[z'::=V_var z2]\<^sub>v, \<Gamma>)\<close> using xx fresh_Pair c.fresh by metis
- qed
-
- ultimately show ?thesis using wfI_wfC_eval_c a1 subst_v_c_def by simp
- qed
-
- ultimately show "is_satis i ?c" using is_satis_imp[OF *] by auto
- qed
- ultimately show ?thesis using valid.simps by simp
- qed
- moreover have "atom x \<sharp> (P, \<B>, \<Gamma>, z1 , CE_val v == CE_val (V_lit l) IMP c[z::=V_var z1]\<^sub>v , z2 , CE_val v == CE_val (V_lit l) IMP c'[z'::=V_var z2]\<^sub>v )"
- unfolding fresh_prod5 \<tau>.fresh using xx xf2 fresh_prodN x_fresh_b by metis
- ultimately show ?thesis using subtype_baseI assms xf2 by metis
-qed
-
-lemma eval_e_concat_eq:
- assumes "wfI \<Theta> \<Gamma> i"
- shows "\<exists>s. eval_e i (CE_val (V_lit (L_bitvec (v1 @ v2))) ) s \<and> eval_e i (CE_concat [(V_lit (L_bitvec v1))]\<^sup>c\<^sup>e [(V_lit (L_bitvec v2))]\<^sup>c\<^sup>e) s"
- using eval_e_valI eval_e_concatI eval_v_litI eval_l.simps by metis
-
-lemma is_satis_eval_e_eq_imp:
- assumes "wfI \<Theta> \<Gamma> i" and "eval_e i e1 s" and "eval_e i e2 s"
- and "is_satis i (CE_val (V_var x) == e1)" (is "is_satis i ?c1")
- shows "is_satis i (CE_val (V_var x) == e2)"
-proof -
- have *:"eval_c i ?c1 True" using assms is_satis.simps by blast
- hence "eval_e i (CE_val (V_var x)) s" using assms is_satis.simps eval_c_elims
- by (metis (full_types) eval_e_uniqueness)
- thus ?thesis using is_satis.simps eval_c.intros assms by fastforce
-qed
-
-lemma valid_eval_e_eq:
- fixes e1::ce and e2::ce
- assumes "\<forall>\<Gamma> i. wfI \<Theta> \<Gamma> i \<longrightarrow> (\<exists>s. eval_e i e1 s \<and> eval_e i e2 s)" and "\<Theta>; \<B>; GNil \<turnstile>\<^sub>w\<^sub>f e1 : b " and "\<Theta>; \<B>; GNil \<turnstile>\<^sub>w\<^sub>f e2 : b"
- shows "\<Theta>; \<B>; (x, b, (CE_val (V_var x) == e1 )) #\<^sub>\<Gamma> GNil \<Turnstile> (CE_val (V_var x) == e2) "
-proof(rule validI)
- show "\<Theta>; \<B>; (x, b, CE_val (V_var x) == e1 ) #\<^sub>\<Gamma> GNil \<turnstile>\<^sub>w\<^sub>f CE_val (V_var x) == e2"
- proof
- have " \<Theta> ; \<B> ; (x, b, TRUE )#\<^sub>\<Gamma>GNil \<turnstile>\<^sub>w\<^sub>f CE_val (V_var x) == e1" using assms wfC_eqI wfE_valI wfV_varI wfX_wfY
- by (simp add: fresh_GNil wfC_e_eq)
- hence "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f (x, b, CE_val (V_var x) == e1 ) #\<^sub>\<Gamma> GNil" using wfG_consI fresh_GNil wfX_wfY assms wfX_wfB by metis
- thus "\<Theta>; \<B>; (x, b, CE_val (V_var x) == e1 ) #\<^sub>\<Gamma> GNil \<turnstile>\<^sub>w\<^sub>f CE_val (V_var x) : b " using wfCE_valI wfV_varI wfX_wfY
- lookup.simps assms wfX_wfY by simp
- show "\<Theta>; \<B>; (x, b, CE_val (V_var x) == e1 ) #\<^sub>\<Gamma> GNil \<turnstile>\<^sub>w\<^sub>f e2 : b " using assms wf_weakening wfX_wfY
- by (metis (full_types) \<open>\<Theta>; \<B>; (x, b, CE_val (V_var x) == e1 ) #\<^sub>\<Gamma> GNil \<turnstile>\<^sub>w\<^sub>f CE_val (V_var x) : b\<close> empty_iff subsetI toSet.simps(1))
- qed
- show " \<forall>i. wfI \<Theta> ((x, b, CE_val (V_var x) == e1 ) #\<^sub>\<Gamma> GNil) i \<and> is_satis_g i ((x, b, CE_val (V_var x) == e1 ) #\<^sub>\<Gamma> GNil) \<longrightarrow> is_satis i (CE_val (V_var x) == e2 )"
- proof(rule,rule)
- fix i
- assume "wfI \<Theta> ((x, b, CE_val (V_var x) == e1 ) #\<^sub>\<Gamma> GNil) i \<and> is_satis_g i ((x, b, CE_val (V_var x) == e1 ) #\<^sub>\<Gamma> GNil)"
- moreover then obtain s where "eval_e i e1 s \<and> eval_e i e2 s" using assms by auto
- ultimately show "is_satis i (CE_val (V_var x) == e2 )" using assms is_satis_eval_e_eq_imp is_satis_g.simps by meson
- qed
-qed
-
-lemma subtype_concat:
- assumes " \<turnstile>\<^sub>w\<^sub>f \<Theta>"
- shows "\<Theta>; \<B>; GNil \<turnstile> \<lbrace> z : B_bitvec | CE_val (V_var z) == CE_val (V_lit (L_bitvec (v1 @ v2))) \<rbrace> \<lesssim>
- \<lbrace> z : B_bitvec | CE_val (V_var z) == CE_concat [(V_lit (L_bitvec v1))]\<^sup>c\<^sup>e [(V_lit (L_bitvec v2))]\<^sup>c\<^sup>e \<rbrace>" (is "\<Theta>; \<B>; GNil \<turnstile> ?t1 \<lesssim> ?t2")
-proof -
- obtain x::x where x: "atom x \<sharp> (\<Theta>, \<B>, GNil, z , CE_val (V_var z) == CE_val (V_lit (L_bitvec (v1 @ v2))),
- z , CE_val (V_var z) == CE_concat [V_lit (L_bitvec v1)]\<^sup>c\<^sup>e [V_lit (L_bitvec v2)]\<^sup>c\<^sup>e )"
- (is "?xfree" )
- using obtain_fresh by auto
-
- have wb1: "\<Theta>; \<B>; GNil \<turnstile>\<^sub>w\<^sub>f CE_val (V_lit (L_bitvec (v1 @ v2))) : B_bitvec" using wfX_wfY wfCE_valI wfV_litI assms base_for_lit.simps wfG_nilI by metis
- hence wb2: "\<Theta>; \<B>; GNil \<turnstile>\<^sub>w\<^sub>f CE_concat [(V_lit (L_bitvec v1))]\<^sup>c\<^sup>e [(V_lit (L_bitvec v2))]\<^sup>c\<^sup>e : B_bitvec"
- using wfCE_concatI wfX_wfY wfV_litI base_for_lit.simps wfCE_valI by metis
-
- show ?thesis proof
- show " \<Theta>; \<B>; GNil \<turnstile>\<^sub>w\<^sub>f ?t1" using wfT_e_eq fresh_GNil wb1 wb2 by metis
- show " \<Theta>; \<B>; GNil \<turnstile>\<^sub>w\<^sub>f ?t2" using wfT_e_eq fresh_GNil wb1 wb2 by metis
- show ?xfree using x by auto
- show "\<Theta>; \<B>; (x, B_bitvec, (CE_val (V_var z) == CE_val (V_lit (L_bitvec (v1 @ v2))) )[z::=V_var x]\<^sub>v) #\<^sub>\<Gamma>
- GNil \<Turnstile> (CE_val (V_var z) == CE_concat [(V_lit (L_bitvec v1))]\<^sup>c\<^sup>e [(V_lit (L_bitvec v2))]\<^sup>c\<^sup>e )[z::=V_var x]\<^sub>v "
- using valid_eval_e_eq eval_e_concat_eq wb1 wb2 subst_v_c_def by fastforce
- qed
-qed
-
-lemma subtype_len:
- assumes " \<turnstile>\<^sub>w\<^sub>f \<Theta>"
- shows "\<Theta>; \<B>; GNil \<turnstile> \<lbrace> z' : B_int | CE_val (V_var z') == CE_val (V_lit (L_num (int (length v)))) \<rbrace> \<lesssim>
- \<lbrace> z : B_int | CE_val (V_var z) == CE_len [(V_lit (L_bitvec v))]\<^sup>c\<^sup>e \<rbrace>" (is "\<Theta>; \<B>; GNil \<turnstile> ?t1 \<lesssim> ?t2")
-proof -
-
- have *: "\<Theta> \<turnstile>\<^sub>w\<^sub>f [] \<and> \<Theta>; \<B>; GNil \<turnstile>\<^sub>w\<^sub>f []\<^sub>\<Delta> " using assms wfG_nilI wfD_emptyI wfPhi_emptyI by auto
- obtain x::x where x: "atom x \<sharp> (\<Theta>, \<B>, GNil, z' , CE_val (V_var z') ==
- CE_val (V_lit (L_num (int (length v)))) , z, CE_val (V_var z) == CE_len [(V_lit (L_bitvec v))]\<^sup>c\<^sup>e )"
- (is "atom x \<sharp> ?F")
- using obtain_fresh by metis
- then show ?thesis proof
- have "\<Theta> ; \<B> ; GNil \<turnstile>\<^sub>w\<^sub>f CE_val (V_lit (L_num (int (length v)))) : B_int"
- using wfCE_valI * wfV_litI base_for_lit.simps
- by (metis wfE_valI wfX_wfY)
-
- thus "\<Theta>; \<B>; GNil \<turnstile>\<^sub>w\<^sub>f ?t1" using wfT_e_eq fresh_GNil by auto
-
- have "\<Theta> ; \<B> ; GNil \<turnstile>\<^sub>w\<^sub>f CE_len [(V_lit (L_bitvec v))]\<^sup>c\<^sup>e : B_int"
- using wfE_valI * wfV_litI base_for_lit.simps wfE_valI wfX_wfY
- by (metis wfCE_lenI wfCE_valI)
-
- thus "\<Theta>; \<B>; GNil \<turnstile>\<^sub>w\<^sub>f ?t2" using wfT_e_eq fresh_GNil by auto
-
- show "\<Theta>; \<B>; (x, B_int, (CE_val (V_var z') == CE_val (V_lit (L_num (int (length v)))) )[z'::=V_var x]\<^sub>v) #\<^sub>\<Gamma> GNil \<Turnstile> (CE_val (V_var z) == CE_len [(V_lit (L_bitvec v))]\<^sup>c\<^sup>e )[z::=V_var x]\<^sub>v"
- (is "\<Theta>; \<B>; ?G \<Turnstile> ?c" ) using valid_len assms subst_v_c_def by auto
- qed
-qed
-
-lemma subtype_base_fresh:
- assumes "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b | c \<rbrace>" and "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b | c' \<rbrace> " and
- "atom z \<sharp> \<Gamma>" and "\<Theta>; \<B>; (z, b, c) #\<^sub>\<Gamma> \<Gamma> \<Turnstile> c'"
- shows "\<Theta>; \<B>; \<Gamma> \<turnstile> \<lbrace> z : b | c \<rbrace> \<lesssim> \<lbrace> z : b | c' \<rbrace>"
-proof -
- obtain x::x where *:"atom x \<sharp> ((\<Theta> , \<B> , z, c, z, c', \<Gamma>) , (\<Theta>, \<B>, \<Gamma>, \<lbrace> z : b | c \<rbrace>, \<lbrace> z : b | c' \<rbrace>))" using obtain_fresh by metis
- moreover hence "atom x \<sharp> \<Gamma>" using fresh_Pair by auto
- moreover hence "\<Theta>; \<B>; (x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<Turnstile> c'[z::=V_var x]\<^sub>v" using assms valid_flip_simple * subst_v_c_def by auto
- ultimately show ?thesis using subtype_baseI assms \<tau>.fresh fresh_Pair by metis
-qed
-
-lemma subtype_bop_arith:
- assumes "wfG \<Theta> \<B> \<Gamma>" and "(opp = Plus \<and> ll = (L_num (n1+n2))) \<or> (opp = LEq \<and> ll = ( if n1\<le>n2 then L_true else L_false))"
- and "(opp = Plus \<longrightarrow> b = B_int) \<and> (opp = LEq \<longrightarrow> b = B_bool)"
- shows "\<Theta>; \<B>; \<Gamma> \<turnstile> (\<lbrace> z : b | C_eq (CE_val (V_var z)) (CE_val (V_lit (ll))) \<rbrace>) \<lesssim>
- \<lbrace> z : b | C_eq (CE_val (V_var z)) (CE_op opp [(V_lit (L_num n1))]\<^sup>c\<^sup>e [(V_lit (L_num n2))]\<^sup>c\<^sup>e) \<rbrace>" (is "\<Theta>; \<B>; \<Gamma> \<turnstile> ?T1 \<lesssim> ?T2")
-proof -
- obtain x::x where xf: "atom x \<sharp> (z, CE_val (V_var z) == CE_val (V_lit (ll)) , z, CE_val (V_var z) == CE_op opp [(V_lit (L_num n1))]\<^sup>c\<^sup>e [(V_lit (L_num n2))]\<^sup>c\<^sup>e , \<Gamma>)"
- using obtain_fresh by blast
-
- have "\<Theta>; \<B>; \<Gamma> \<turnstile> (\<lbrace> x : b | C_eq (CE_val (V_var x)) (CE_val (V_lit (ll))) \<rbrace>) \<lesssim>
- \<lbrace> x : b | C_eq (CE_val (V_var x)) (CE_op opp [(V_lit (L_num n1))]\<^sup>c\<^sup>e [(V_lit (L_num n2))]\<^sup>c\<^sup>e) \<rbrace>" (is "\<Theta>; \<B>; \<Gamma> \<turnstile> ?S1 \<lesssim> ?S2")
- proof(rule subtype_base_fresh)
-
- show "atom x \<sharp> \<Gamma>" using xf fresh_Pair by auto
-
- show "wfT \<Theta> \<B> \<Gamma> (\<lbrace> x : b | CE_val (V_var x) == CE_val (V_lit ll) \<rbrace>)" (is "wfT \<Theta> \<B> ?A ?B")
- proof(rule wfT_e_eq)
- have " \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f (V_lit ll) : b" using wfV_litI base_for_lit.simps assms by metis
- thus " \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_val (V_lit ll) : b" using wfCE_valI by auto
- show "atom x \<sharp> \<Gamma>" using xf fresh_Pair by auto
- qed
-
- consider "opp = Plus" | "opp = LEq" using opp.exhaust assms by blast
- then show "wfT \<Theta> \<B> \<Gamma> (\<lbrace> x : b | CE_val (V_var x) == CE_op opp [(V_lit (L_num n1))]\<^sup>c\<^sup>e [(V_lit (L_num n2))]\<^sup>c\<^sup>e \<rbrace>)" (is "wfT \<Theta> \<B> ?A ?C")
- proof(cases)
- case 1
- then show "\<Theta> ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> x : b | [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ opp [ [ L_num n1 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ L_num n2 ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e \<rbrace>"
- using wfCE_valI wfCE_plusI assms wfV_litI base_for_lit.simps assms
- by (metis \<open>atom x \<sharp> \<Gamma>\<close> wfT_e_eq)
- next
- case 2
- then show "\<Theta> ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> x : b | [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ opp [ [ L_num n1 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ L_num n2 ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e \<rbrace> "
- using wfCE_valI wfCE_plusI assms wfV_litI base_for_lit.simps assms
-
- by (metis \<open>atom x \<sharp> \<Gamma>\<close> wfCE_leqI wfT_e_eq)
- qed
-
- show "\<Theta>; \<B> ; (x, b, (CE_val (V_var x) == CE_val (V_lit (ll)) )) #\<^sub>\<Gamma> \<Gamma>
- \<Turnstile> (CE_val (V_var x) == CE_op opp [V_lit (L_num n1)]\<^sup>c\<^sup>e [V_lit (L_num n2)]\<^sup>c\<^sup>e)" (is "\<Theta>; \<B>; ?G \<Turnstile> ?c")
- using valid_arith_bop assms xf by simp
-
- qed
- moreover have "?S1 = ?T1 " using type_l_eq by auto
- moreover have "?S2 = ?T2" using type_e_eq ce.fresh v.fresh supp_l_empty fresh_def empty_iff fresh_e_opp
- by (metis ms_fresh_all(4))
- ultimately show ?thesis by auto
-
-qed
-
-lemma subtype_bop_eq:
- assumes "wfG \<Theta> \<B> \<Gamma>" and "base_for_lit l1 = base_for_lit l2"
- shows "\<Theta>; \<B>; \<Gamma> \<turnstile> (\<lbrace> z : B_bool | C_eq (CE_val (V_var z)) (CE_val (V_lit (if l1 = l2 then L_true else L_false))) \<rbrace>) \<lesssim>
- \<lbrace> z : B_bool | C_eq (CE_val (V_var z)) (CE_op Eq [(V_lit l1)]\<^sup>c\<^sup>e [(V_lit l2)]\<^sup>c\<^sup>e) \<rbrace>" (is "\<Theta>; \<B>; \<Gamma> \<turnstile> ?T1 \<lesssim> ?T2")
-proof -
- let ?ll = "if l1 = l2 then L_true else L_false"
- obtain x::x where xf: "atom x \<sharp> (z, CE_val (V_var z) == CE_val (V_lit (if l1 = l2 then L_true else L_false)) , z, CE_val (V_var z) == CE_op Eq [(V_lit l1)]\<^sup>c\<^sup>e [(V_lit l2)]\<^sup>c\<^sup>e , \<Gamma>, (\<Theta>, \<B>, \<Gamma>))"
- using obtain_fresh by blast
-
- have "\<Theta>; \<B>; \<Gamma> \<turnstile> (\<lbrace> x : B_bool | C_eq (CE_val (V_var x)) (CE_val (V_lit (?ll))) \<rbrace>) \<lesssim>
- \<lbrace> x : B_bool | C_eq (CE_val (V_var x)) (CE_op Eq [(V_lit (l1))]\<^sup>c\<^sup>e [(V_lit (l2))]\<^sup>c\<^sup>e) \<rbrace>" (is "\<Theta>; \<B>; \<Gamma> \<turnstile> ?S1 \<lesssim> ?S2")
- proof(rule subtype_base_fresh)
-
- show "atom x \<sharp> \<Gamma>" using xf fresh_Pair by auto
-
- show "wfT \<Theta> \<B> \<Gamma> (\<lbrace> x : B_bool | CE_val (V_var x) == CE_val (V_lit ?ll) \<rbrace>)" (is "wfT \<Theta> \<B> ?A ?B")
- proof(rule wfT_e_eq)
- have " \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f (V_lit ?ll) : B_bool" using wfV_litI base_for_lit.simps assms by metis
- thus " \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_val (V_lit ?ll) : B_bool" using wfCE_valI by auto
- show "atom x \<sharp> \<Gamma>" using xf fresh_Pair by auto
- qed
-
- show " \<Theta> ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> x : B_bool | [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ eq [ [ l1 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ l2 ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e \<rbrace> "
- proof(rule wfT_e_eq)
- show "\<Theta> ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f [ eq [ [ l1 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ l2 ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e : B_bool"
- apply(rule wfCE_eqI, rule wfCE_valI)
- apply(rule wfV_litI, simp add: assms)
- using wfV_litI assms wfCE_valI by auto
- show "atom x \<sharp> \<Gamma>" using xf fresh_Pair by auto
- qed
-
- show "\<Theta>; \<B> ; (x, B_bool, (CE_val (V_var x) == CE_val (V_lit (?ll)) )) #\<^sub>\<Gamma> \<Gamma>
- \<Turnstile> (CE_val (V_var x) == CE_op Eq [V_lit (l1)]\<^sup>c\<^sup>e [V_lit (l2)]\<^sup>c\<^sup>e)" (is "\<Theta>; \<B>; ?G \<Turnstile> ?c")
- using valid_eq_bop assms xf by auto
-
- qed
- moreover have "?S1 = ?T1 " using type_l_eq by auto
- moreover have "?S2 = ?T2" using type_e_eq ce.fresh v.fresh supp_l_empty fresh_def empty_iff fresh_e_opp
- by (metis ms_fresh_all(4))
- ultimately show ?thesis by auto
-
-qed
-
-lemma subtype_top:
- assumes "wfT \<Theta> \<B> G (\<lbrace> z : b | c \<rbrace>)"
- shows "\<Theta>; \<B>; G \<turnstile> (\<lbrace> z : b | c \<rbrace>) \<lesssim> (\<lbrace> z : b | TRUE \<rbrace>)"
-proof -
- obtain x::x where *: "atom x \<sharp> (\<Theta>, \<B>, G, z , c, z , TRUE)" using obtain_fresh by blast
- then show ?thesis proof(rule subtype_baseI)
- show \<open> \<Theta>; \<B>; G \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b | c \<rbrace> \<close> using assms by auto
- show \<open> \<Theta>; \<B>;G \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b | TRUE \<rbrace> \<close> using wfT_TRUE assms wfX_wfY b_of.simps wfT_wf
- by (metis wfX_wfB(8))
- hence "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f (x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\<Gamma> G" using wfT_wf_cons3 assms fresh_Pair * subst_v_c_def by auto
- thus \<open>\<Theta>; \<B>; (x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\<Gamma> G \<Turnstile> (TRUE)[z::=V_var x]\<^sub>v \<close> using valid_trueI subst_cv.simps subst_v_c_def by metis
- qed
-qed
-
-lemma if_simp:
- "(if x = x then e1 else e2) = e1"
- by auto
-
-lemma subtype_split:
- assumes "split n v (v1,v2)" and "\<turnstile>\<^sub>w\<^sub>f \<Theta>"
- shows "\<Theta> ; {||} ; GNil \<turnstile> \<lbrace> z : [ B_bitvec , B_bitvec ]\<^sup>b | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ [ L_bitvec
- v1 ]\<^sup>v , [ L_bitvec
- v2 ]\<^sup>v ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrace> \<lesssim> \<lbrace> z : [ B_bitvec , B_bitvec ]\<^sup>b | [ [ L_bitvec
- v ]\<^sup>v ]\<^sup>c\<^sup>e == [ [#1[ [ z ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e @@ [#2[ [ z ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e AND [| [#1[ [ z ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e == [ [ L_num
- n ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrace>"
- (is "\<Theta> ;?B ; GNil \<turnstile> \<lbrace> z : [ B_bitvec , B_bitvec ]\<^sup>b | ?c1 \<rbrace> \<lesssim> \<lbrace> z : [ B_bitvec , B_bitvec ]\<^sup>b | ?c2 \<rbrace>")
-proof -
- obtain x::x where xf:"atom x \<sharp> (\<Theta>, ?B, GNil, z, ?c1,z, ?c2 )" using obtain_fresh by auto
- then show ?thesis proof(rule subtype_baseI)
- show *: \<open>\<Theta> ; ?B ; (x, [ B_bitvec , B_bitvec ]\<^sup>b, (?c1)[z::=[ x ]\<^sup>v]\<^sub>v) #\<^sub>\<Gamma>
- GNil \<Turnstile> (?c2)[z::=[ x ]\<^sup>v]\<^sub>v \<close>
- unfolding subst_v_c_def subst_cv.simps subst_cev.simps subst_vv.simps if_simp
- using valid_split[OF assms, of x] by simp
- show \<open> \<Theta> ; ?B ; GNil \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : [ B_bitvec , B_bitvec ]\<^sup>b| ?c1 \<rbrace> \<close> using valid_wfT[OF *] xf fresh_prodN by metis
- show \<open> \<Theta> ; ?B ; GNil \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : [ B_bitvec , B_bitvec ]\<^sup>b| ?c2 \<rbrace> \<close> using valid_wfT[OF *] xf fresh_prodN by metis
- qed
-qed
-
-lemma subtype_range:
- fixes n::int and \<Gamma>::\<Gamma>
- assumes "0 \<le> n \<and> n \<le> int (length v)" and "\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f \<Gamma>"
- shows "\<Theta> ; {||} ; \<Gamma> \<turnstile> \<lbrace> z : B_int | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_num n ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrace> \<lesssim>
- \<lbrace> z : B_int | ([ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e ) AND (
- [ leq [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e [| [ [ L_bitvec v ]\<^sup>v ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e) \<rbrace>"
- (is "\<Theta> ; ?B ; \<Gamma> \<turnstile> \<lbrace> z : B_int | ?c1 \<rbrace> \<lesssim> \<lbrace> z : B_int | ?c2 AND ?c3 \<rbrace>")
-proof -
- obtain x::x where *:\<open>atom x \<sharp> (\<Theta>, ?B, \<Gamma>, z, ?c1 , z, ?c2 AND ?c3)\<close> using obtain_fresh by auto
- moreover have **:\<open>\<Theta> ; ?B ; (x, B_int, (?c1)[z::=[ x ]\<^sup>v]\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<Turnstile> (?c2 AND ?c3)[z::=[ x ]\<^sup>v]\<^sub>v\<close>
- unfolding subst_v_c_def subst_cv.simps subst_cev.simps subst_vv.simps if_simp using valid_range_length[OF assms(1)] assms fresh_prodN * by simp
-
- moreover hence \<open> \<Theta> ; ?B ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : B_int | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_num n ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrace> \<close> using
- valid_wfT * fresh_prodN by metis
- moreover have \<open> \<Theta> ; ?B ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : B_int | ?c2 AND ?c3 \<rbrace> \<close>
- using valid_wfT[OF **] * fresh_prodN by metis
- ultimately show ?thesis using subtype_baseI by auto
-qed
-
-lemma check_num_range:
- assumes "0 \<le> n \<and> n \<le> int (length v)" and "\<turnstile>\<^sub>w\<^sub>f \<Theta>"
- shows "\<Theta> ; {||} ; GNil \<turnstile> ([ L_num n ]\<^sup>v) \<Leftarrow> \<lbrace> z : B_int | ([ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e) AND
- [ leq [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e [| [ [ L_bitvec v ]\<^sup>v ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrace>"
- using assms subtype_range check_v.intros infer_v_litI wfG_nilI
- by (meson infer_natI)
-
-section \<open>Literals\<close>
-
-nominal_function type_for_lit :: "l \<Rightarrow> \<tau>" where
- "type_for_lit (L_true) = (\<lbrace> z : B_bool | [[z]\<^sup>v]\<^sup>c\<^sup>e == [V_lit L_true]\<^sup>c\<^sup>e \<rbrace>)"
-| "type_for_lit (L_false) = (\<lbrace> z : B_bool | [[z]\<^sup>v]\<^sup>c\<^sup>e == [V_lit L_false]\<^sup>c\<^sup>e \<rbrace>)"
-| "type_for_lit (L_num n) = (\<lbrace> z : B_int | [[z]\<^sup>v]\<^sup>c\<^sup>e == [V_lit (L_num n)]\<^sup>c\<^sup>e \<rbrace>)"
-| "type_for_lit (L_unit) = (\<lbrace> z : B_unit | [[z]\<^sup>v]\<^sup>c\<^sup>e == [V_lit (L_unit )]\<^sup>c\<^sup>e \<rbrace>)" (* could have done { z : unit | True } but want the uniformity *)
-| "type_for_lit (L_bitvec v) = (\<lbrace> z : B_bitvec | [[z]\<^sup>v]\<^sup>c\<^sup>e == [V_lit (L_bitvec v)]\<^sup>c\<^sup>e \<rbrace>)"
- by (auto simp: eqvt_def type_for_lit_graph_aux_def, metis l.strong_exhaust,(simp add: permute_int_def flip_bitvec0)+ )
-nominal_termination (eqvt) by lexicographic_order
-
-
-nominal_function type_for_var :: "\<Gamma> \<Rightarrow> \<tau> \<Rightarrow> x \<Rightarrow> \<tau>" where
- "type_for_var G \<tau> x = (case lookup G x of
- None \<Rightarrow> \<tau>
- | Some (b,c) \<Rightarrow> (\<lbrace> x : b | c \<rbrace>)) "
- apply auto unfolding eqvt_def apply(rule allI) unfolding type_for_var_graph_aux_def eqvt_def by simp
-nominal_termination (eqvt) by lexicographic_order
-
-lemma infer_l_form:
- fixes l::l and tm::"'a::fs"
- assumes "\<turnstile> l \<Rightarrow> \<tau>"
- shows "\<exists>z b. \<tau> = (\<lbrace> z : b | C_eq (CE_val (V_var z)) (CE_val (V_lit l)) \<rbrace>) \<and> atom z \<sharp> tm"
-proof -
- obtain z' and b where t:"\<tau> = (\<lbrace> z' : b | C_eq (CE_val (V_var z')) (CE_val (V_lit l)) \<rbrace>)" using infer_l_elims assms using infer_l.simps type_for_lit.simps
- type_for_lit.cases by blast
- obtain z::x where zf: "atom z \<sharp> tm" using obtain_fresh by metis
- have "\<tau> = \<lbrace> z : b | C_eq (CE_val (V_var z)) (CE_val (V_lit l)) \<rbrace>" using type_e_eq ce.fresh v.fresh l.fresh
- by (metis t type_l_eq)
- thus ?thesis using zf by auto
-qed
-
-lemma infer_l_form3:
- fixes l::l
- assumes "\<turnstile> l \<Rightarrow> \<tau>"
- shows "\<exists>z. \<tau> = (\<lbrace> z : base_for_lit l | C_eq (CE_val (V_var z)) (CE_val (V_lit l)) \<rbrace>)"
- using infer_l_elims using assms using infer_l.simps type_for_lit.simps base_for_lit.simps by auto
-
-lemma infer_l_form4[simp]:
- fixes \<Gamma>::\<Gamma>
- assumes "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> "
- shows "\<exists>z. \<turnstile> l \<Rightarrow> (\<lbrace> z : base_for_lit l | C_eq (CE_val (V_var z)) (CE_val (V_lit l)) \<rbrace>)"
- using assms infer_l_form2 infer_l_form3 by metis
-
-lemma infer_v_unit_form:
- fixes v::v
- assumes "P ; \<B> ; \<Gamma> \<turnstile> v \<Rightarrow> (\<lbrace> z1 : B_unit | c1 \<rbrace>)" and "supp v = {}"
- shows "v = V_lit L_unit"
- using assms proof(nominal_induct \<Gamma> v "\<lbrace> z1 : B_unit | c1 \<rbrace>" rule: infer_v.strong_induct)
- case (infer_v_varI \<Theta> \<B> c x z)
- then show ?case using supp_at_base by auto
-next
- case (infer_v_litI \<Theta> \<B> \<Gamma> l)
- from \<open>\<turnstile> l \<Rightarrow> \<lbrace> z1 : B_unit | c1 \<rbrace>\<close> show ?case by(nominal_induct "\<lbrace> z1 : B_unit | c1 \<rbrace>" rule: infer_l.strong_induct,auto)
-qed
-
-lemma base_for_lit_wf:
- assumes "\<turnstile>\<^sub>w\<^sub>f \<Theta>"
- shows "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f base_for_lit l"
- using base_for_lit.simps using wfV_elims wf_intros assms l.exhaust by metis
-
-lemma infer_l_t_wf:
- fixes \<Gamma>::\<Gamma>
- assumes "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<and> atom z \<sharp> \<Gamma>"
- shows "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : base_for_lit l | C_eq (CE_val (V_var z)) (CE_val (V_lit l)) \<rbrace>"
-proof
- show "atom z \<sharp> (\<Theta>, \<B>, \<Gamma>)" using wfG_fresh_x assms by auto
- show "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f base_for_lit l" using base_for_lit_wf assms wfX_wfY by metis
- thus "\<Theta>; \<B>; (z, base_for_lit l, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_val (V_var z) == CE_val (V_lit l)" using wfC_v_eq wfV_litI assms wfX_wfY by metis
-qed
-
-lemma infer_l_wf:
- fixes l::l and \<Gamma>::\<Gamma> and \<tau>::\<tau> and \<Theta>::\<Theta>
- assumes "\<turnstile> l \<Rightarrow> \<tau>" and "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>"
- shows "\<turnstile>\<^sub>w\<^sub>f \<Theta>" and "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>" and "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau>"
-proof -
- show *:"\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>" using assms infer_l_elims by auto
- thus "\<turnstile>\<^sub>w\<^sub>f \<Theta>" using wfX_wfY by auto
- show *:"\<Theta> ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau>" using infer_l_t_wf assms infer_l_form3 *
- by (metis \<open>\<turnstile>\<^sub>w\<^sub>f \<Theta>\<close> fresh_GNil wfG_nilI wfT_weakening_nil)
-qed
-
-lemma infer_l_uniqueness:
- fixes l::l
- assumes "\<turnstile> l \<Rightarrow> \<tau>" and "\<turnstile> l \<Rightarrow> \<tau>'"
- shows "\<tau> = \<tau>'"
- using assms
-proof -
- obtain z and b where zt: "\<tau> = (\<lbrace> z : b | C_eq (CE_val (V_var z)) (CE_val (V_lit l)) \<rbrace>)" using infer_l_form assms by blast
- obtain z' and b where z't: "\<tau>' = (\<lbrace> z' : b | C_eq (CE_val (V_var z')) (CE_val (V_lit l)) \<rbrace>)" using infer_l_form assms by blast
- thus ?thesis using type_l_eq zt z't assms infer_l.simps infer_l_elims l.distinct
- by (metis infer_l_form3)
-qed
-
-section \<open>Values\<close>
-
-lemma type_v_eq:
- assumes "\<lbrace> z1 : b1 | c1 \<rbrace> = \<lbrace> z : b | C_eq (CE_val (V_var z)) (CE_val (V_var x)) \<rbrace>" and "atom z \<sharp> x"
- shows "b = b1" and "c1 = C_eq (CE_val (V_var z1)) (CE_val (V_var x))"
- using assms by (auto,metis Abs1_eq_iff \<tau>.eq_iff assms c.fresh ce.fresh type_e_eq v.fresh)
-
-lemma infer_var2 [elim]:
- assumes "P; \<B> ; G \<turnstile> V_var x \<Rightarrow> \<tau>"
- shows "\<exists>b c. Some (b,c) = lookup G x"
- using assms infer_v_elims lookup_iff by (metis (no_types, lifting))
-
-lemma infer_var3 [elim]:
- assumes "\<Theta>; \<B>; \<Gamma> \<turnstile> V_var x \<Rightarrow> \<tau>"
- shows "\<exists>z b c. Some (b,c) = lookup \<Gamma> x \<and> \<tau> = (\<lbrace> z : b | C_eq (CE_val (V_var z)) (CE_val (V_var x)) \<rbrace>) \<and> atom z \<sharp> x \<and> atom z \<sharp> (\<Theta>, \<B>, \<Gamma>)"
- using infer_v_elims(1)[OF assms(1)] by metis
-
-lemma infer_bool_options2:
- fixes v::v
- assumes "\<Theta>; \<B>; \<Gamma> \<turnstile> v \<Rightarrow> \<lbrace> z : b | c \<rbrace>" and "supp v = {} \<and> b = B_bool"
- shows "v = V_lit L_true \<or> (v = (V_lit L_false))"
- using assms
-proof(nominal_induct "\<lbrace> z : b | c \<rbrace>" rule: infer_v.strong_induct)
- case (infer_v_varI \<Theta> \<B> \<Gamma> c x z)
- then show ?case using v.supp supp_at_base by auto
-next
- case (infer_v_litI \<Theta> \<B> \<Gamma> l)
- from \<open> \<turnstile> l \<Rightarrow> \<lbrace> z : b | c \<rbrace>\<close> show ?case proof(nominal_induct "\<lbrace> z : b | c \<rbrace>" rule: infer_l.strong_induct)
- case (infer_trueI z)
- then show ?case by auto
- next
- case (infer_falseI z)
- then show ?case by auto
- next
- case (infer_natI n z)
- then show ?case using infer_v_litI by simp
- next
- case (infer_unitI z)
- then show ?case using infer_v_litI by simp
- next
- case (infer_bitvecI bv z)
- then show ?case using infer_v_litI by simp
- qed
-qed(auto+)
-
-lemma infer_bool_options:
- fixes v::v
- assumes "\<Theta>; \<B>; \<Gamma> \<turnstile> v \<Rightarrow> \<lbrace> z : B_bool | c \<rbrace>" and "supp v = {}"
- shows "v = V_lit L_true \<or> (v = (V_lit L_false))"
- using infer_bool_options2 assms by blast
-
-lemma infer_int2:
- fixes v::v
- assumes "\<Theta>; \<B>; \<Gamma> \<turnstile> v \<Rightarrow> \<lbrace> z : b | c \<rbrace>"
- shows "supp v = {} \<and> b = B_int \<longrightarrow> (\<exists>n. v = V_lit (L_num n))"
- using assms
-proof(nominal_induct "\<lbrace> z : b | c \<rbrace>" rule: infer_v.strong_induct)
- case (infer_v_varI \<Theta> \<B> \<Gamma> c x z)
- then show ?case using v.supp supp_at_base by auto
-next
- case (infer_v_litI \<Theta> \<B> \<Gamma> l)
- from \<open> \<turnstile> l \<Rightarrow> \<lbrace> z : b | c \<rbrace>\<close> show ?case proof(nominal_induct "\<lbrace> z : b | c \<rbrace>" rule: infer_l.strong_induct)
- case (infer_trueI z)
- then show ?case by auto
- next
- case (infer_falseI z)
- then show ?case by auto
- next
- case (infer_natI n z)
- then show ?case using infer_v_litI by simp
- next
- case (infer_unitI z)
- then show ?case using infer_v_litI by simp
- next
- case (infer_bitvecI bv z)
- then show ?case using infer_v_litI by simp
- qed
-qed(auto+)
-
-lemma infer_bitvec:
- fixes \<Theta>::\<Theta> and v::v
- assumes "\<Theta>; \<B>; \<Gamma> \<turnstile> v \<Rightarrow> \<lbrace> z' : B_bitvec | c' \<rbrace>" and "supp v = {}"
- shows "\<exists>bv. v = V_lit (L_bitvec bv)"
- using assms proof(nominal_induct v rule: v.strong_induct)
- case (V_lit l)
- then show ?case by(nominal_induct l rule: l.strong_induct,force+)
-next
- case (V_consp s dc b v)
- then show ?case using infer_v_elims(7)[OF V_consp(2)] \<tau>.eq_iff by auto
-next
- case (V_var x)
- then show ?case using supp_at_base by auto
-qed(force+)
-
-lemma infer_int:
- assumes "infer_v \<Theta> \<B> \<Gamma> v (\<lbrace> z : B_int | c \<rbrace>)" and "supp v= {}"
- shows "\<exists>n. V_lit (L_num n) = v"
- using assms infer_int2 by (metis (no_types, lifting))
-
-lemma infer_lit:
- assumes "infer_v \<Theta> \<B> \<Gamma> v (\<lbrace> z : b | c \<rbrace>)" and "supp v= {}" and "b \<in> { B_bool , B_int , B_unit }"
- shows "\<exists>l. V_lit l = v"
- using assms proof(nominal_induct v rule: v.strong_induct)
- case (V_lit x)
- then show ?case by (simp add: supp_at_base)
-next
- case (V_var x)
- then show ?case
- by (simp add: supp_at_base)
-next
- case (V_pair x1a x2a)
- then show ?case using supp_at_base by auto
-next
- case (V_cons x1a x2a x3)
- then show ?case using supp_at_base by auto
-next
- case (V_consp x1a x2a x3 x4)
- then show ?case using supp_at_base by auto
-qed
-
-lemma infer_v_form[simp]:
- fixes v::v
- assumes "\<Theta>; \<B>; \<Gamma> \<turnstile> v \<Rightarrow> \<tau>"
- shows "\<exists>z b. \<tau> = (\<lbrace> z : b | C_eq (CE_val (V_var z)) (CE_val v)\<rbrace>) \<and> atom z \<sharp> v \<and> atom z \<sharp> (\<Theta>, \<B>, \<Gamma>)"
- using assms
-proof(nominal_induct rule: infer_v.strong_induct)
- case (infer_v_varI \<Theta> \<B> \<Gamma> b c x z)
- then show ?case by force
-next
- case (infer_v_litI \<Theta> \<B> \<Gamma> l \<tau>)
- then obtain z and b where "\<tau> = \<lbrace> z : b | CE_val (V_var z) == CE_val (V_lit l) \<rbrace> \<and>atom z \<sharp> (\<Theta>, \<B>, \<Gamma>) "
- using infer_l_form by metis
- moreover hence "atom z \<sharp> (V_lit l)" using supp_l_empty v.fresh(1) fresh_prod2 fresh_def by blast
- ultimately show ?case by metis
-next
- case (infer_v_pairI z v1 v2 \<Theta> \<B> \<Gamma> t1 t2)
- then show ?case by force
-next
- case (infer_v_consI s dclist \<Theta> dc tc \<B> \<Gamma> v tv z)
- moreover hence "atom z \<sharp> (V_cons s dc v)" using
- Un_commute b.supp(3) fresh_def sup_bot.right_neutral supp_b_empty v.supp(4) pure_supp by metis
- ultimately show ?case using fresh_prodN by metis
-next
- case (infer_v_conspI s bv dclist \<Theta> dc tc \<B> \<Gamma> v tv b z)
- moreover hence "atom z \<sharp> (V_consp s dc b v)" unfolding v.fresh using pure_fresh fresh_prodN * by metis
- ultimately show ?case using fresh_prodN by metis
-qed
-
-lemma infer_v_form2:
- fixes v::v
- assumes "\<Theta>; \<B>; \<Gamma> \<turnstile> v \<Rightarrow> (\<lbrace> z : b | c \<rbrace>)" and "atom z \<sharp> v"
- shows "c = C_eq (CE_val (V_var z)) (CE_val v)"
- using assms
-proof -
- obtain z' and b' where "(\<lbrace> z : b | c \<rbrace>) = (\<lbrace> z' : b' | CE_val (V_var z') == CE_val v \<rbrace>) \<and> atom z' \<sharp> v"
- using infer_v_form assms by meson
- thus ?thesis using Abs1_eq_iff(3) \<tau>.eq_iff type_e_eq
- by (metis assms(2) ce.fresh(1))
-qed
-
-lemma infer_v_form3:
- fixes v::v
- assumes "\<Theta>; \<B>; \<Gamma> \<turnstile> v \<Rightarrow> \<tau>" and "atom z \<sharp> (v,\<Gamma>)"
- shows "\<Theta>; \<B>; \<Gamma> \<turnstile> v \<Rightarrow> \<lbrace> z : b_of \<tau> | C_eq (CE_val (V_var z)) (CE_val v)\<rbrace>"
-proof -
- obtain z' and b' where "\<tau> = \<lbrace> z' : b' | C_eq (CE_val (V_var z')) (CE_val v)\<rbrace> \<and> atom z' \<sharp> v \<and> atom z' \<sharp> (\<Theta>, \<B>, \<Gamma>)"
- using infer_v_form assms by metis
- moreover hence "\<lbrace> z' : b' | C_eq (CE_val (V_var z')) (CE_val v)\<rbrace> = \<lbrace> z : b' | C_eq (CE_val (V_var z)) (CE_val v)\<rbrace>"
- using assms type_e_eq fresh_Pair ce.fresh by auto
- ultimately show ?thesis using b_of.simps assms by auto
-qed
-
-lemma infer_v_form4:
- fixes v::v
- assumes "\<Theta>; \<B>; \<Gamma> \<turnstile> v \<Rightarrow> \<tau>" and "atom z \<sharp> (v,\<Gamma>)" and "b = b_of \<tau>"
- shows "\<Theta>; \<B>; \<Gamma> \<turnstile> v \<Rightarrow> \<lbrace> z : b | C_eq (CE_val (V_var z)) (CE_val v)\<rbrace>"
- using assms infer_v_form3 by simp
-
-lemma infer_v_v_wf:
- fixes v::v
- shows "\<Theta>; \<B> ; G \<turnstile> v \<Rightarrow> \<tau> \<Longrightarrow> \<Theta>; \<B>; G \<turnstile>\<^sub>w\<^sub>f v : (b_of \<tau>)"
-proof(induct rule: infer_v.induct)
- case (infer_v_varI \<Theta> \<B> \<Gamma> b c x z)
- then show ?case using wfC_elims wf_intros by auto
-next
- case (infer_v_pairI z v1 v2 \<Theta> \<B> \<Gamma> t1 t2)
- then show ?case using wfC_elims wf_intros by auto
-next
- case (infer_v_litI \<Theta> \<B> \<Gamma> l \<tau>)
- hence "b_of \<tau> = base_for_lit l" using infer_l_form3 b_of.simps by metis
- then show ?case using wfV_litI infer_l_wf infer_v_litI wfG_b_weakening
- by (metis fempty_fsubsetI)
-next
- case (infer_v_consI s dclist \<Theta> dc tc \<B> \<Gamma> v tv z)
- then show ?case using wfC_elims wf_intros
- by (metis (no_types, lifting) b_of.simps has_fresh_z2 subtype_eq_base2)
-next
- case (infer_v_conspI s bv dclist \<Theta> dc tc \<B> \<Gamma> v tv b z)
- obtain z1 b1 c1 where t:"tc = \<lbrace> z1 : b1 | c1 \<rbrace>" using obtain_fresh_z by metis
- show ?case unfolding b_of.simps proof(rule wfV_conspI)
- show \<open>AF_typedef_poly s bv dclist \<in> set \<Theta>\<close> using infer_v_conspI by auto
- show \<open>(dc, \<lbrace> z1 : b1 | c1 \<rbrace> ) \<in> set dclist\<close> using infer_v_conspI t by auto
- show \<open> \<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f b \<close> using infer_v_conspI by auto
- show \<open>atom bv \<sharp> (\<Theta>, \<B>, \<Gamma>, b, v)\<close> using infer_v_conspI by auto
- have " b1[bv::=b]\<^sub>b\<^sub>b = b_of tv" using subtype_eq_base2[OF infer_v_conspI(5)] b_of.simps t subst_tb.simps by auto
- thus \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b1[bv::=b]\<^sub>b\<^sub>b \<close> using infer_v_conspI by auto
- qed
-qed
-
-lemma infer_v_t_form_wf:
- assumes " wfB \<Theta> \<B> b" and "wfV \<Theta> \<B> \<Gamma> v b" and "atom z \<sharp> \<Gamma>"
- shows "wfT \<Theta> \<B> \<Gamma> \<lbrace> z : b | C_eq (CE_val (V_var z)) (CE_val v)\<rbrace>"
- using wfT_v_eq assms by auto
-
-lemma infer_v_t_wf:
- fixes v::v
- assumes "\<Theta>; \<B>; G \<turnstile> v \<Rightarrow> \<tau>"
- shows "wfT \<Theta> \<B> G \<tau> \<and> wfB \<Theta> \<B> (b_of \<tau>) "
-proof -
- obtain z and b where "\<tau> = \<lbrace> z : b | CE_val (V_var z) == CE_val v \<rbrace> \<and> atom z \<sharp> v \<and> atom z \<sharp> (\<Theta>, \<B>, G)" using infer_v_form assms by metis
- moreover have "wfB \<Theta> \<B> b" using infer_v_v_wf b_of.simps wfX_wfB(1) assms
- using calculation by fastforce
- ultimately show "wfT \<Theta> \<B> G \<tau> \<and> wfB \<Theta> \<B> (b_of \<tau>)" using infer_v_v_wf infer_v_t_form_wf assms by fastforce
-qed
-
-lemma infer_v_wf:
- fixes v::v
- assumes "\<Theta>; \<B>; G \<turnstile> v \<Rightarrow> \<tau>"
- shows "\<Theta>; \<B>; G \<turnstile>\<^sub>w\<^sub>f v : (b_of \<tau>)" and "wfT \<Theta> \<B> G \<tau>" and "wfTh \<Theta>" and "wfG \<Theta> \<B> G"
-proof -
- show "\<Theta>; \<B>; G \<turnstile>\<^sub>w\<^sub>f v : b_of \<tau> " using infer_v_v_wf assms by auto
- show "\<Theta>; \<B>; G \<turnstile>\<^sub>w\<^sub>f \<tau>" using infer_v_t_wf assms by auto
- thus "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f G" using wfX_wfY by auto
- thus "\<turnstile>\<^sub>w\<^sub>f \<Theta>" using wfX_wfY by auto
-qed
-
-lemma check_bool_options:
- assumes "\<Theta>; \<B>; \<Gamma> \<turnstile> v \<Leftarrow> \<lbrace> z : B_bool | TRUE \<rbrace>" and "supp v = {}"
- shows "v = V_lit L_true \<or> v = V_lit L_false"
-proof -
- obtain t1 where " \<Theta>; \<B>; \<Gamma> \<turnstile> t1 \<lesssim> \<lbrace> z : B_bool | TRUE \<rbrace> \<and> \<Theta>; \<B>; \<Gamma> \<turnstile> v \<Rightarrow> t1" using check_v_elims
- using assms by blast
- thus ?thesis using infer_bool_options assms
- by (metis \<tau>.exhaust b_of.simps subtype_eq_base2)
-qed
-
-lemma check_v_wf:
- fixes v::v and \<Gamma>::\<Gamma> and \<tau>::\<tau>
- assumes "\<Theta>; \<B>; \<Gamma> \<turnstile> v \<Leftarrow> \<tau>"
- shows " \<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>" and "\<Theta>; \<B>;\<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b_of \<tau>" and "\<Theta>; \<B>;\<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau>"
-proof -
- obtain \<tau>' where *: "\<Theta>; \<B>; \<Gamma> \<turnstile> \<tau>' \<lesssim> \<tau> \<and> \<Theta>; \<B>; \<Gamma> \<turnstile> v \<Rightarrow> \<tau>'" using check_v_elims assms by auto
- thus "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> " and "\<Theta>; \<B>;\<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b_of \<tau>" and "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau>"
- using infer_v_wf infer_v_v_wf subtype_eq_base2 * subtype_wf by metis+
-qed
-
-lemma infer_v_form_fresh:
- fixes v::v and t::"'a::fs"
- assumes "\<Theta>; \<B>; \<Gamma> \<turnstile> v \<Rightarrow> \<tau>"
- shows "\<exists>z b. \<tau> = \<lbrace> z : b | C_eq (CE_val (V_var z)) (CE_val v)\<rbrace> \<and> atom z \<sharp> (t,v)"
-proof -
- obtain z' and b' where "\<tau> = \<lbrace> z' : b' | C_eq (CE_val (V_var z')) (CE_val v)\<rbrace>" using infer_v_form assms by blast
- moreover then obtain z and b and c where "\<tau> = \<lbrace> z : b | c \<rbrace> \<and> atom z \<sharp> (t,v)" using obtain_fresh_z by metis
- ultimately have "\<tau> = \<lbrace> z : b | C_eq (CE_val (V_var z)) (CE_val v)\<rbrace> \<and> atom z \<sharp> (t,v) "
- using assms infer_v_form2 by auto
- thus ?thesis by blast
-qed
-
-text \<open> More generally, if support of a term is empty then any G will do \<close>
-
-lemma infer_v_form_consp:
- assumes "\<Theta>; \<B>; \<Gamma> \<turnstile> V_consp s dc b v \<Rightarrow> \<tau>"
- shows "b_of \<tau> = B_app s b"
- using assms proof(nominal_induct "V_consp s dc b v" \<tau> rule: infer_v.strong_induct)
- case (infer_v_conspI bv dclist \<Theta> tc \<B> \<Gamma> tv z)
- then show ?case using b_of.simps by metis
-qed
-
-lemma lookup_in_rig_b:
- assumes "Some (b2, c2) = lookup (\<Gamma>[x\<longmapsto>c']) x'" and
- "Some (b1, c1) = lookup \<Gamma> x'"
- shows "b1 = b2"
- using assms lookup_in_rig[OF assms(2)]
- by (metis option.inject prod.inject)
-
-lemma infer_v_uniqueness_rig:
- fixes x::x and c::c
- assumes "infer_v P B G v \<tau>" and "infer_v P B (replace_in_g G x c') v \<tau>'"
- shows "\<tau> = \<tau>'"
- using assms
-proof(nominal_induct "v" arbitrary: \<tau>' \<tau> rule: v.strong_induct)
- case (V_lit l)
- hence "infer_l l \<tau>" and "infer_l l \<tau>'" using assms(1) infer_v_elims(2) by auto
- then show ?case using infer_l_uniqueness by presburger
-next
- case (V_var y)
-
- obtain b and c where bc: "Some (b,c) = lookup G y"
- using assms(1) infer_v_elims(2) using V_var.prems(1) lookup_iff by force
- then obtain c'' where bc':"Some (b,c'') = lookup (replace_in_g G x c') y"
- using lookup_in_rig by blast
-
- obtain z where "\<tau> = (\<lbrace> z : b | C_eq (CE_val (V_var z)) (CE_val (V_var y)) \<rbrace>)" using infer_v_elims(1)[of P B G y \<tau>] V_var
- bc option.inject prod.inject lookup_in_g by metis
- moreover obtain z' where "\<tau>' = (\<lbrace> z' : b | C_eq (CE_val (V_var z')) (CE_val (V_var y)) \<rbrace>)" using infer_v_elims(1)[of P B _ y \<tau>'] V_var
- option.inject prod.inject lookup_in_rig by (metis bc')
- ultimately show ?case using type_e_eq
- by (metis V_var.prems(1) V_var.prems(2) \<tau>.eq_iff ce.fresh(1) finite.emptyI fresh_atom_at_base
- fresh_finite_insert infer_v_elims(1) v.fresh(2))
-next
- case (V_pair v1 v2)
- obtain z and z1 and z2 and t1 and t2 and c1 and c2 where
- t1: "\<tau> = (\<lbrace> z : [ b_of t1 , b_of t2 ]\<^sup>b | CE_val (V_var z) == CE_val (V_pair v1 v2) \<rbrace>) \<and>
- atom z \<sharp> (v1, v2) \<and> P ; B ; G \<turnstile> v1 \<Rightarrow> t1 \<and> P ; B ; G \<turnstile> v2 \<Rightarrow> t2"
- using infer_v_elims(3)[OF V_pair(3)] by metis
- moreover obtain z' and z1' and z2' and t1' and t2' and c1' and c2' where
- t2: "\<tau>' = (\<lbrace> z' : [ b_of t1' , b_of t2' ]\<^sup>b | CE_val (V_var z') == CE_val (V_pair v1 v2) \<rbrace>) \<and>
- atom z' \<sharp> (v1, v2) \<and> P ; B ; (replace_in_g G x c') \<turnstile> v1 \<Rightarrow> t1' \<and>
- P ; B ; (replace_in_g G x c') \<turnstile> v2 \<Rightarrow> t2'"
- using infer_v_elims(3)[OF V_pair(4)] by metis
- ultimately have "t1 = t1' \<and> t2 = t2'" using V_pair.hyps(1) V_pair.hyps(2) \<tau>.eq_iff by blast
- then show ?case using t1 t2 by simp
-next
- case (V_cons s dc v)
- obtain x and z and tc and dclist where t1: "\<tau> = (\<lbrace> z : B_id s | CE_val (V_var z) == CE_val (V_cons s dc v) \<rbrace>) \<and>
- AF_typedef s dclist \<in> set P \<and>
- (dc, tc) \<in> set dclist \<and> atom z \<sharp> v"
- using infer_v_elims(4)[OF V_cons(2)] by metis
- moreover obtain x' and z' and tc' and dclist' where t2: "\<tau>' = (\<lbrace> z' : B_id s | CE_val (V_var z') == CE_val (V_cons s dc v) \<rbrace>)
- \<and> AF_typedef s dclist' \<in> set P \<and> (dc, tc') \<in> set dclist' \<and> atom z' \<sharp> v"
- using infer_v_elims(4)[OF V_cons(3)] by metis
- moreover have a: "AF_typedef s dclist' \<in> set P" and b:"(dc,tc') \<in> set dclist'" and c:"AF_typedef s dclist \<in> set P" and
- d:"(dc, tc) \<in> set dclist" using t1 t2 by auto
- ultimately have "tc = tc'" using wfTh_dc_t_unique2 infer_v_wf(3)[OF V_cons(2)] by metis
-
- moreover have "atom z \<sharp> CE_val (V_cons s dc v) \<and> atom z' \<sharp> CE_val (V_cons s dc v)"
- using e.fresh(1) v.fresh(4) t1 t2 pure_fresh by auto
- ultimately have "(\<lbrace> z : B_id s | CE_val (V_var z) == CE_val (V_cons s dc v) \<rbrace>) = (\<lbrace> z' : B_id s | CE_val (V_var z') == CE_val (V_cons s dc v) \<rbrace>)"
- using type_e_eq by metis
- thus ?case using t1 t2 by simp
-next
- case (V_consp s dc b v)
- from V_consp(2) V_consp show ?case proof(nominal_induct "V_consp s dc b v" \<tau> arbitrary: v rule:infer_v.strong_induct)
-
- case (infer_v_conspI bv dclist \<Theta> tc \<B> \<Gamma> v tv z)
-
- obtain z3 and b3 where *:"\<tau>' = \<lbrace> z3 : b3 | [ [ z3 ]\<^sup>v ]\<^sup>c\<^sup>e == [ V_consp s dc b v ]\<^sup>c\<^sup>e \<rbrace> \<and> atom z3 \<sharp> V_consp s dc b v"
- using infer_v_form[OF \<open>\<Theta>; \<B>; \<Gamma>[x\<longmapsto>c'] \<turnstile> V_consp s dc b v \<Rightarrow> \<tau>'\<close> ] by metis
- moreover then have "b3 = B_app s b" using infer_v_form_consp b_of.simps * infer_v_conspI by metis
-
- moreover have "\<lbrace> z3 : B_app s b | [ [ z3 ]\<^sup>v ]\<^sup>c\<^sup>e == [ V_consp s dc b v ]\<^sup>c\<^sup>e \<rbrace> = \<lbrace> z : B_app s b | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ V_consp s dc b v ]\<^sup>c\<^sup>e \<rbrace>"
- proof -
- have "atom z3 \<sharp> [V_consp s dc b v]\<^sup>c\<^sup>e" using * ce.fresh by auto
- moreover have "atom z \<sharp> [V_consp s dc b v]\<^sup>c\<^sup>e" using * infer_v_conspI ce.fresh v.fresh pure_fresh by metis
- ultimately show ?thesis using type_e_eq infer_v_conspI v.fresh ce.fresh by metis
- qed
- ultimately show ?case using * by auto
- qed
-qed
-
-lemma infer_v_uniqueness:
- assumes "infer_v P B G v \<tau>" and "infer_v P B G v \<tau>'"
- shows "\<tau> = \<tau>'"
-proof -
- obtain x::x where "atom x \<sharp> G" using obtain_fresh by metis
- hence "G [ x \<longmapsto> C_true ] = G" using replace_in_g_forget assms infer_v_wf by fast
- thus ?thesis using infer_v_uniqueness_rig assms by metis
-qed
-
-lemma infer_v_tid_form:
- fixes v::v
- assumes "\<Theta> ; B ; \<Gamma> \<turnstile> v \<Rightarrow> \<lbrace> z : B_id tid | c \<rbrace>" and "AF_typedef tid dclist \<in> set \<Theta>" and "supp v = {}"
- shows "\<exists>dc v' t. v = V_cons tid dc v' \<and> (dc , t ) \<in> set dclist"
- using assms proof(nominal_induct v "\<lbrace> z : B_id tid | c \<rbrace>" rule: infer_v.strong_induct)
- case (infer_v_varI \<Theta> \<B> c x z)
- then show ?case using v.supp supp_at_base by auto
-next
- case (infer_v_litI \<Theta> \<B> l)
- then show ?case by auto
-next
- case (infer_v_consI dclist1 \<Theta> dc tc \<B> \<Gamma> v tv z)
- hence "supp v = {}" using v.supp by simp
- then obtain dca and v' where *:"V_cons tid dc v = V_cons tid dca v'" using infer_v_consI by auto
- hence "dca = dc" using v.eq_iff(4) by auto
- hence "V_cons tid dc v = V_cons tid dca v' \<and> (dca, tc) \<in> set dclist1" using infer_v_consI * by auto
- moreover have "dclist = dclist1" using wfTh_dclist_unique infer_v_consI wfX_wfY \<open>dca=dc\<close>
- proof -
- show ?thesis
- by (meson \<open>AF_typedef tid dclist1 \<in> set \<Theta>\<close> \<open>\<Theta>; \<B>; \<Gamma> \<turnstile> v \<Rightarrow> tv\<close> infer_v_consI.prems infer_v_wf(4) wfTh_dclist_unique wfX_wfY)
- qed
- ultimately show ?case by auto
-qed
-
-lemma check_v_tid_form:
- assumes "\<Theta> ; B ; \<Gamma> \<turnstile> v \<Leftarrow> \<lbrace> z : B_id tid | TRUE \<rbrace>" and "AF_typedef tid dclist \<in> set \<Theta>" and "supp v = {}"
- shows "\<exists>dc v' t. v = V_cons tid dc v' \<and> (dc , t ) \<in> set dclist"
- using assms proof(nominal_induct v "\<lbrace> z : B_id tid | TRUE \<rbrace>" rule: check_v.strong_induct)
- case (check_v_subtypeI \<Theta> \<B> \<Gamma> \<tau>1 v)
- then obtain z and c where "\<tau>1 = \<lbrace> z : B_id tid | c \<rbrace>" using subtype_eq_base2 b_of.simps
- by (metis obtain_fresh_z2)
- then show ?case using infer_v_tid_form check_v_subtypeI by simp
-qed
-
-lemma check_v_num_leq:
- fixes n::int and \<Gamma>::\<Gamma>
- assumes "0 \<le> n \<and> n \<le> int (length v)" and " \<turnstile>\<^sub>w\<^sub>f \<Theta> " and "\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f \<Gamma>"
- shows "\<Theta> ; {||} ; \<Gamma> \<turnstile> [ L_num n ]\<^sup>v \<Leftarrow> \<lbrace> z : B_int | ([ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e)
- AND ([ leq [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e [| [ [ L_bitvec v ]\<^sup>v ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e ) \<rbrace>"
-proof -
- have "\<Theta> ; {||} ; \<Gamma> \<turnstile> [ L_num n ]\<^sup>v \<Rightarrow> \<lbrace> z : B_int | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_num n ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrace>"
- using infer_v_litI infer_natI wfG_nilI assms by auto
- thus ?thesis using subtype_range[OF assms(1) ] assms check_v_subtypeI by metis
-qed
-
-lemma check_int:
- assumes "check_v \<Theta> \<B> \<Gamma> v (\<lbrace> z : B_int | c \<rbrace>)" and "supp v = {}"
- shows "\<exists>n. V_lit (L_num n) = v"
- using assms infer_int check_v_elims by (metis b_of.simps infer_v_form subtype_eq_base2)
-
-definition sble :: "\<Theta> \<Rightarrow> \<Gamma> \<Rightarrow> bool" where
- "sble \<Theta> \<Gamma> = (\<exists>i. i \<Turnstile> \<Gamma> \<and> \<Theta> ; \<Gamma> \<turnstile> i)"
-
-lemma check_v_range:
- assumes "\<Theta> ; B ; \<Gamma> \<turnstile> v2 \<Leftarrow> \<lbrace> z : B_int | [ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e AND
- [ leq [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e [| [ v1 ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrace> "
- (is "\<Theta> ; ?B ; \<Gamma> \<turnstile> v2 \<Leftarrow> \<lbrace> z : B_int | ?c1 \<rbrace>")
- and "v1 = V_lit (L_bitvec bv) \<and> v2 = V_lit (L_num n) " and "atom z \<sharp> \<Gamma>" and "sble \<Theta> \<Gamma>"
- shows "0 \<le> n \<and> n \<le> int (length bv)"
-proof -
- have "\<Theta> ; ?B ; \<Gamma> \<turnstile> \<lbrace> z : B_int | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_num n ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrace> \<lesssim> \<lbrace> z : B_int | ?c1 \<rbrace>"
- using check_v_elims assms
- by (metis infer_l_uniqueness infer_natI infer_v_elims(2))
- moreover have "atom z \<sharp> \<Gamma>" using fresh_GNil assms by simp
- ultimately have "\<Theta> ; ?B ; ((z, B_int, [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_num n ]\<^sup>v ]\<^sup>c\<^sup>e ) #\<^sub>\<Gamma> \<Gamma>) \<Turnstile> ?c1"
- using subtype_valid_simple by auto
- thus ?thesis using assms valid_range_length_inv check_v_wf wfX_wfY sble_def by metis
-qed
-
-section \<open>Expressions\<close>
-
-lemma infer_e_plus[elim]:
- fixes v1::v and v2::v
- assumes "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile> AE_op Plus v1 v2 \<Rightarrow> \<tau>"
- shows "\<exists>z . (\<lbrace> z : B_int | C_eq (CE_val (V_var z)) (CE_op Plus [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e) \<rbrace> = \<tau>)"
- using infer_e_elims assms by metis
-
-lemma infer_e_leq[elim]:
- assumes "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile> AE_op LEq v1 v2 \<Rightarrow> \<tau>"
- shows "\<exists>z . (\<lbrace> z : B_bool | C_eq (CE_val (V_var z)) (CE_op LEq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e) \<rbrace> = \<tau>)"
- using infer_e_elims assms by metis
-
-lemma infer_e_eq[elim]:
- assumes "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile> AE_op Eq v1 v2 \<Rightarrow> \<tau>"
- shows "\<exists>z . (\<lbrace> z : B_bool | C_eq (CE_val (V_var z)) (CE_op Eq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e) \<rbrace> = \<tau>)"
- using infer_e_elims(25)[OF assms] by metis
-
-lemma infer_e_e_wf:
- fixes e::e
- assumes "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile> e \<Rightarrow> \<tau>"
- shows "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f e : b_of \<tau>"
- using assms proof(nominal_induct \<tau> avoiding: \<tau> rule: infer_e.strong_induct)
- case (infer_e_valI \<Theta> \<B> \<Gamma> \<Delta>' \<Phi> v \<tau>)
- then show ?case using infer_v_v_wf wf_intros by metis
-next
- case (infer_e_plusI \<Theta> \<B> \<Gamma> \<Delta>' \<Phi> v1 z1 c1 v2 z2 c2 z3)
- then show ?case using b_of.simps infer_v_v_wf wf_intros by metis
-next
- case (infer_e_leqI \<Theta> \<B> \<Gamma> \<Delta>' v1 z1 c1 v2 z2 c2 z3)
- then show ?case using b_of.simps infer_v_v_wf wf_intros by metis
-next
- case (infer_e_eqI \<Theta> \<B> \<Gamma> \<Delta>' v1 z1 c1 v2 z2 c2 z3)
- then show ?case using b_of.simps infer_v_v_wf wf_intros by metis
-next
- case (infer_e_appI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> f x b c \<tau>' s' v \<tau>'')
- have "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f AE_app f v : b_of \<tau>' " proof
- show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using infer_e_appI by auto
- show \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using infer_e_appI by auto
- show \<open>Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \<tau>' s'))) = lookup_fun \<Phi> f\<close> using infer_e_appI by auto
- show "\<Theta>; \<B>;\<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b" using infer_e_appI check_v_wf b_of.simps by metis
- qed
- moreover have "b_of \<tau>' = b_of (\<tau>'[x::=v]\<^sub>v)" using subst_tbase_eq subst_v_\<tau>_def by auto
- ultimately show ?case using infer_e_appI subst_v_c_def subst_b_\<tau>_def by auto
-next
- case (infer_e_appPI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> b' f bv x b c \<tau>'' s' v \<tau>')
-
- have "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f AE_appP f b' v : (b_of \<tau>'')[bv::=b']\<^sub>b " proof
- show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using infer_e_appPI by auto
- show \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using infer_e_appPI by auto
- show \<open>Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \<tau>'' s'))) = lookup_fun \<Phi> f\<close> using * infer_e_appPI by metis
- show "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f b'" using infer_e_appPI by auto
- show "\<Theta>; \<B>;\<Gamma> \<turnstile>\<^sub>w\<^sub>f v : (b[bv::=b']\<^sub>b)" using infer_e_appPI check_v_wf b_of.simps subst_b_b_def by metis
- have "atom bv \<sharp> (b_of \<tau>'')[bv::=b']\<^sub>b\<^sub>b" using fresh_subst_if subst_b_b_def infer_e_appPI by metis
- thus "atom bv \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>, \<Delta>, b', v, (b_of \<tau>'')[bv::=b']\<^sub>b)" using infer_e_appPI fresh_prodN subst_b_b_def by metis
- qed
- moreover have "b_of \<tau>' = (b_of \<tau>'')[bv::=b']\<^sub>b"
- using \<open>\<tau>''[bv::=b']\<^sub>b[x::=v]\<^sub>v = \<tau>'\<close> b_of_subst_bb_commute subst_tbase_eq subst_b_b_def subst_v_\<tau>_def subst_b_\<tau>_def by auto
- ultimately show ?case using infer_e_appI by auto
-next
- case (infer_e_fstI \<Theta> \<B> \<Gamma> \<Delta>' \<Phi> v z' b1 b2 c z)
- then show ?case using b_of.simps infer_v_v_wf wf_intros by metis
-next
- case (infer_e_sndI \<Theta> \<B> \<Gamma> \<Delta>' \<Phi> v z' b1 b2 c z)
- then show ?case using b_of.simps infer_v_v_wf wf_intros by metis
-next
- case (infer_e_lenI \<Theta> \<B> \<Gamma> \<Delta>' \<Phi> v z' c z)
- then show ?case using b_of.simps infer_v_v_wf wf_intros by metis
-next
- case (infer_e_mvarI \<Theta> \<Gamma> \<Phi> \<Delta> u \<tau>)
- then show ?case using b_of.simps infer_v_v_wf wf_intros by metis
-next
- case (infer_e_concatI \<Theta> \<B> \<Gamma> \<Delta>' \<Phi> v1 z1 c1 v2 z2 c2 z3)
- then show ?case using b_of.simps infer_v_v_wf wf_intros by metis
-next
- case (infer_e_splitI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v1 z1 c1 v2 z2 z3)
- have " \<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f AE_split v1 v2 : B_pair B_bitvec B_bitvec"
- proof
- show "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>" using infer_e_splitI by auto
- show "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>" using infer_e_splitI by auto
- show "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v1 : B_bitvec" using infer_e_splitI b_of.simps infer_v_wf by metis
- show "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v2 : B_int" using infer_e_splitI b_of.simps check_v_wf by metis
- qed
- then show ?case using b_of.simps by auto
-qed
-
-lemma infer_e_t_wf:
- fixes e::e and \<Gamma>::\<Gamma> and \<tau>::\<tau> and \<Delta>::\<Delta> and \<Phi>::\<Phi>
- assumes "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile> e \<Rightarrow> \<tau>"
- shows "\<Theta>; \<B>;\<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau> \<and> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>"
- using assms proof(induct rule: infer_e.induct)
- case (infer_e_valI \<Theta> \<B> \<Gamma> \<Delta>' \<Phi> v \<tau>)
- then show ?case using infer_v_t_wf by auto
-next
- case (infer_e_plusI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v1 z1 c1 v2 z2 c2 z3)
- hence "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_op Plus [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e : B_int" using wfCE_plusI wfD_emptyI wfPhi_emptyI infer_v_v_wf wfCE_valI
- by (metis b_of.simps infer_v_wf)
- then show ?case using wfT_e_eq infer_e_plusI by auto
-next
- case (infer_e_leqI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v1 z1 c1 v2 z2 c2 z3)
- hence " \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_op LEq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e : B_bool" using wfCE_leqI wfD_emptyI wfPhi_emptyI infer_v_v_wf wfCE_valI
- by (metis b_of.simps infer_v_wf)
- then show ?case using wfT_e_eq infer_e_leqI by auto
-next
- case (infer_e_eqI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v1 z1 b c1 v2 z2 c2 z3)
- hence " \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_op Eq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e : B_bool" using wfCE_eqI wfD_emptyI wfPhi_emptyI infer_v_v_wf wfCE_valI
- by (metis b_of.simps infer_v_wf)
- then show ?case using wfT_e_eq infer_e_eqI by auto
-next
- case (infer_e_appI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> f x b c \<tau> s' v \<tau>')
- show ?case proof
- show "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>" using infer_e_appI by auto
- show "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau>'" proof -
- have *:" \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b " using infer_e_appI check_v_wf(2) b_of.simps by metis
- moreover have *:"\<Theta>; \<B>; (x, b, c) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau>" proof(rule wf_weakening1(4))
- show \<open> \<Theta>; \<B>; (x,b,c)#\<^sub>\<Gamma>GNil \<turnstile>\<^sub>w\<^sub>f \<tau> \<close> using wfPhi_f_simple_wfT wfD_wf infer_e_appI wb_b_weakening by fastforce
- have "\<Theta> ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> x : b | c \<rbrace>" using infer_e_appI check_v_wf(3) by auto
- thus \<open> \<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f (x, b, c) #\<^sub>\<Gamma> \<Gamma> \<close> using infer_e_appI
- wfT_wfC[THEN wfG_consI[rotated 3] ] * wfT_wf_cons fresh_prodN by simp
- show \<open>toSet ((x,b,c)#\<^sub>\<Gamma>GNil) \<subseteq> toSet ((x, b, c) #\<^sub>\<Gamma> \<Gamma>)\<close> using toSet.simps by auto
- qed
- moreover have "((x, b, c) #\<^sub>\<Gamma> \<Gamma>)[x::=v]\<^sub>\<Gamma>\<^sub>v = \<Gamma>" using subst_gv.simps by auto
-
- ultimately show ?thesis using infer_e_appI wf_subst1(4)[OF *, of GNil x b c \<Gamma> v] subst_v_\<tau>_def by auto
- qed
- qed
-next
- case (infer_e_appPI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> b' f bv x b c \<tau>' s' v \<tau>)
-
- have "\<Theta>; \<B>; ((x, b[bv::=b']\<^sub>b\<^sub>b, c[bv::=b']\<^sub>c\<^sub>b) #\<^sub>\<Gamma> \<Gamma>)[x::=v]\<^sub>\<Gamma>\<^sub>v \<turnstile>\<^sub>w\<^sub>f (\<tau>'[bv::=b']\<^sub>b)[x::=v]\<^sub>\<tau>\<^sub>v"
- proof(rule wf_subst(4))
- show \<open> \<Theta>; \<B>; (x, b[bv::=b']\<^sub>b\<^sub>b, c[bv::=b']\<^sub>c\<^sub>b) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau>'[bv::=b']\<^sub>b \<close>
- proof(rule wf_weakening1(4))
- have \<open> \<Theta> ; {|bv|} ; (x,b,c)#\<^sub>\<Gamma>GNil \<turnstile>\<^sub>w\<^sub>f \<tau>' \<close> using wfPhi_f_poly_wfT infer_e_appI infer_e_appPI by simp
- thus \<open> \<Theta>; \<B>; (x,b[bv::=b']\<^sub>b\<^sub>b,c[bv::=b']\<^sub>c\<^sub>b)#\<^sub>\<Gamma>GNil \<turnstile>\<^sub>w\<^sub>f \<tau>'[bv::=b']\<^sub>b \<close>
- using wfT_subst_wfT infer_e_appPI wb_b_weakening subst_b_\<tau>_def subst_v_\<tau>_def by presburger
- have "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> x : b[bv::=b']\<^sub>b\<^sub>b | c[bv::=b']\<^sub>c\<^sub>b \<rbrace>"
- using infer_e_appPI check_v_wf(3) subst_b_b_def subst_b_c_def by metis
- thus \<open> \<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f (x, b[bv::=b']\<^sub>b\<^sub>b, c[bv::=b']\<^sub>c\<^sub>b) #\<^sub>\<Gamma> \<Gamma> \<close>
- using infer_e_appPI wfT_wfC[THEN wfG_consI[rotated 3] ] * wfX_wfY wfT_wf_cons wb_b_weakening by metis
- show \<open>toSet ((x,b[bv::=b']\<^sub>b\<^sub>b,c[bv::=b']\<^sub>c\<^sub>b)#\<^sub>\<Gamma>GNil) \<subseteq> toSet ((x, b[bv::=b']\<^sub>b\<^sub>b, c[bv::=b']\<^sub>c\<^sub>b) #\<^sub>\<Gamma> \<Gamma>)\<close> using toSet.simps by auto
- qed
- show \<open>(x, b[bv::=b']\<^sub>b\<^sub>b, c[bv::=b']\<^sub>c\<^sub>b) #\<^sub>\<Gamma> \<Gamma> = GNil @ (x, b[bv::=b']\<^sub>b\<^sub>b, c[bv::=b']\<^sub>c\<^sub>b) #\<^sub>\<Gamma> \<Gamma>\<close> using append_g.simps by auto
- show \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v :b[bv::=b']\<^sub>b\<^sub>b \<close> using infer_e_appPI check_v_wf(2) b_of.simps subst_b_b_def by metis
- qed
- moreover have "((x, b[bv::=b']\<^sub>b\<^sub>b, c[bv::=b']\<^sub>c\<^sub>b) #\<^sub>\<Gamma> \<Gamma>)[x::=v]\<^sub>\<Gamma>\<^sub>v = \<Gamma>" using subst_gv.simps by auto
- ultimately show ?case using infer_e_appPI subst_v_\<tau>_def by simp
-next
- case (infer_e_fstI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v z' b1 b2 c z)
- hence "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_fst [v]\<^sup>c\<^sup>e: b1" using wfCE_fstI wfD_emptyI wfPhi_emptyI infer_v_v_wf
- b_of.simps using wfCE_valI by fastforce
- then show ?case using wfT_e_eq infer_e_fstI by auto
-next
- case (infer_e_sndI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v z' b1 b2 c z)
- hence "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_snd [v]\<^sup>c\<^sup>e: b2" using wfCE_sndI wfD_emptyI wfPhi_emptyI infer_v_v_wf wfCE_valI
- by (metis b_of.simps infer_v_wf)
- then show ?case using wfT_e_eq infer_e_sndI by auto
-next
- case (infer_e_lenI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v z' c z)
- hence "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_len [v]\<^sup>c\<^sup>e: B_int" using wfCE_lenI wfD_emptyI wfPhi_emptyI infer_v_v_wf wfCE_valI
- by (metis b_of.simps infer_v_wf)
- then show ?case using wfT_e_eq infer_e_lenI by auto
-next
- case (infer_e_mvarI \<Theta> \<Gamma> \<Phi> \<Delta> u \<tau>)
- then show ?case using wfD_wfT by blast
-next
- case (infer_e_concatI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v1 z1 c1 v2 z2 c2 z3)
- hence "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_concat [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e: B_bitvec" using wfCE_concatI wfD_emptyI wfPhi_emptyI infer_v_v_wf wfCE_valI
- by (metis b_of.simps infer_v_wf)
- then show ?case using wfT_e_eq infer_e_concatI by auto
-next
- case (infer_e_splitI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v1 z1 c1 v2 z2 z3)
-
- hence wfg: " \<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f (z3, [ B_bitvec , B_bitvec ]\<^sup>b, TRUE) #\<^sub>\<Gamma> \<Gamma>"
- using infer_v_wf wfG_cons2I wfB_pairI wfB_bitvecI by simp
- have wfz: "\<Theta>; \<B>; (z3, [ B_bitvec , B_bitvec ]\<^sup>b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f [ [ z3 ]\<^sup>v ]\<^sup>c\<^sup>e : [ B_bitvec , B_bitvec ]\<^sup>b "
- apply(rule wfCE_valI , rule wfV_varI)
- using wfg apply simp
- using lookup.simps(2)[of z3 "[ B_bitvec , B_bitvec ]\<^sup>b" TRUE \<Gamma> z3] by simp
- have 1: "\<Theta>; \<B>; (z3, [ B_bitvec , B_bitvec ]\<^sup>b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f [ v2 ]\<^sup>c\<^sup>e : B_int "
- using check_v_wf[OF infer_e_splitI(4)] wf_weakening(1)[OF _ wfg] b_of.simps toSet.simps wfCE_valI by fastforce
- have 2: "\<Theta>; \<B>; (z3, [ B_bitvec , B_bitvec ]\<^sup>b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f [ v1 ]\<^sup>c\<^sup>e : B_bitvec"
- using infer_v_wf[OF infer_e_splitI(3)] wf_weakening(1)[OF _ wfg] b_of.simps toSet.simps wfCE_valI by fastforce
-
- have "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> z3 : [ B_bitvec , B_bitvec ]\<^sup>b | [ v1 ]\<^sup>c\<^sup>e == [ [#1[ [ z3 ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e @@ [#2[ [ z3 ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e AND [| [#1[ [ z3 ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e == [ v2 ]\<^sup>c\<^sup>e \<rbrace>"
- proof
- show "atom z3 \<sharp> (\<Theta>, \<B>, \<Gamma>)" using infer_e_splitI wfTh_x_fresh wfX_wfY fresh_prod3 wfG_fresh_x by metis
- show "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f [ B_bitvec , B_bitvec ]\<^sup>b " using wfB_pairI wfB_bitvecI infer_e_splitI wfX_wfY by metis
- show "\<Theta>; \<B>; (z3, [ B_bitvec , B_bitvec ]\<^sup>b, TRUE) #\<^sub>\<Gamma>
- \<Gamma> \<turnstile>\<^sub>w\<^sub>f [ v1 ]\<^sup>c\<^sup>e == [ [#1[ [ z3 ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e @@ [#2[ [ z3 ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e AND [| [#1[ [ z3 ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e == [ v2 ]\<^sup>c\<^sup>e"
- using wfg wfz 1 2 wf_intros by meson
- qed
- thus ?case using infer_e_splitI by auto
-qed
-
-lemma infer_e_wf:
- fixes e::e and \<Gamma>::\<Gamma> and \<tau>::\<tau> and \<Delta>::\<Delta> and \<Phi>::\<Phi>
- assumes "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile> e \<Rightarrow> \<tau>"
- shows "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau>" and "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>" and "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>" and "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>" and "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f e : (b_of \<tau>)"
- using infer_e_t_wf infer_e_e_wf wfE_wf assms by metis+
-
-lemma infer_e_fresh:
- fixes x::x
- assumes "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile> e \<Rightarrow> \<tau>" and "atom x \<sharp> \<Gamma>"
- shows "atom x \<sharp> (e,\<tau>)"
-proof -
- have "atom x \<sharp> e" using infer_e_e_wf[THEN wfE_x_fresh,OF assms(1)] assms(2) by auto
- moreover have "atom x \<sharp> \<tau>" using assms infer_e_wf wfT_x_fresh by metis
- ultimately show ?thesis using fresh_Pair by auto
-qed
-
-inductive check_e :: "\<Theta> \<Rightarrow> \<Phi> \<Rightarrow> \<B> \<Rightarrow> \<Gamma> \<Rightarrow> \<Delta> \<Rightarrow> e \<Rightarrow> \<tau> \<Rightarrow> bool" (" _ ; _ ; _ ; _ ; _ \<turnstile> _ \<Leftarrow> _" [50, 50, 50] 50) where
- check_e_subtypeI: "\<lbrakk> infer_e T P B G D e \<tau>' ; subtype T B G \<tau>' \<tau> \<rbrakk> \<Longrightarrow> check_e T P B G D e \<tau>"
-equivariance check_e
-nominal_inductive check_e .
-
-inductive_cases check_e_elims[elim!]:
- "check_e F D B G \<Theta> (AE_val v) \<tau>"
- "check_e F D B G \<Theta> e \<tau>"
-
-lemma infer_e_fst_pair:
- fixes v1::v
- assumes "\<Theta> ; \<Phi> ; {||} ; GNil ; \<Delta> \<turnstile> [#1[ v1 , v2 ]\<^sup>v]\<^sup>e \<Rightarrow> \<tau>"
- shows "\<exists>\<tau>'. \<Theta> ; \<Phi> ; {||} ; GNil ; \<Delta> \<turnstile> [v1]\<^sup>e \<Rightarrow> \<tau>' \<and>
- \<Theta> ; {||} ; GNil \<turnstile> \<tau>' \<lesssim> \<tau>"
-proof -
- obtain z' and b1 and b2 and c and z where ** : "\<tau> = (\<lbrace> z : b1 | C_eq (CE_val (V_var z)) (CE_fst [(V_pair v1 v2)]\<^sup>c\<^sup>e) \<rbrace>) \<and>
- wfD \<Theta> {||} GNil \<Delta> \<and> wfPhi \<Theta> \<Phi> \<and>
- (\<Theta> ; {||} ; GNil \<turnstile> V_pair v1 v2 \<Rightarrow> \<lbrace> z' : B_pair b1 b2 | c \<rbrace>) \<and> atom z \<sharp> V_pair v1 v2 "
- using infer_e_elims assms by metis
- hence *:" \<Theta> ; {||} ; GNil \<turnstile> V_pair v1 v2 \<Rightarrow> \<lbrace> z' : B_pair b1 b2 | c \<rbrace>" by auto
-
- obtain t1a and t2a where
- *: "\<Theta> ; {||} ; GNil \<turnstile> v1 \<Rightarrow> t1a \<and> \<Theta> ; {||} ; GNil \<turnstile> v2 \<Rightarrow> t2a \<and> B_pair b1 b2 = B_pair (b_of t1a) (b_of t2a)"
- using infer_v_elims(5)[OF *] by metis
-
- hence suppv: "supp v1 = {} \<and> supp v2 = {} \<and> supp (V_pair v1 v2) = {}" using ** infer_v_v_wf wfV_supp atom_dom.simps toSet.simps supp_GNil
- by (meson wfV_supp_nil)
-
- obtain z1 and b1' where "t1a = \<lbrace> z1 : b1' | [ [ z1 ]\<^sup>v ]\<^sup>c\<^sup>e == [ v1 ]\<^sup>c\<^sup>e \<rbrace> "
- using infer_v_form[of \<Theta> "{||}" GNil v1 t1a] * by auto
- moreover hence "b1' = b1" using * b_of.simps by simp
- ultimately have "\<Theta> ; {||} ; GNil \<turnstile> v1 \<Rightarrow> \<lbrace> z1 : b1 | CE_val (V_var z1) == CE_val v1 \<rbrace>" using * by auto
- moreover have "\<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f CE_fst [V_pair v1 v2]\<^sup>c\<^sup>e : b1" using wfCE_fstI infer_v_wf(1) ** b_of.simps wfCE_valI by metis
- moreover hence st: "\<Theta> ; {||} ; GNil \<turnstile> \<lbrace> z1 : b1 | CE_val (V_var z1) == CE_val v1 \<rbrace> \<lesssim> (\<lbrace> z : b1 | CE_val (V_var z) == CE_fst [V_pair v1 v2]\<^sup>c\<^sup>e \<rbrace>)"
- using subtype_gnil_fst infer_v_v_wf by auto
- moreover have "wfD \<Theta> {||} GNil \<Delta> \<and> wfPhi \<Theta> \<Phi>" using ** by auto
- ultimately show ?thesis using wfX_wfY ** infer_e_valI by metis
-qed
-
-lemma infer_e_snd_pair:
- assumes "\<Theta> ; \<Phi> ; {||} ; GNil ; \<Delta> \<turnstile> AE_snd (V_pair v1 v2) \<Rightarrow> \<tau>"
- shows "\<exists>\<tau>'. \<Theta> ; \<Phi> ; {||} ; GNil ; \<Delta> \<turnstile> AE_val v2 \<Rightarrow> \<tau>' \<and> \<Theta> ; {||} ; GNil \<turnstile> \<tau>' \<lesssim> \<tau>"
-proof -
- obtain z' and b1 and b2 and c and z where ** : "(\<tau> = (\<lbrace> z : b2 | C_eq (CE_val (V_var z)) (CE_snd [(V_pair v1 v2)]\<^sup>c\<^sup>e) \<rbrace>)) \<and>
- (wfD \<Theta> {||} GNil \<Delta>) \<and> (wfPhi \<Theta> \<Phi>) \<and>
- \<Theta> ; {||} ; GNil \<turnstile> V_pair v1 v2 \<Rightarrow> \<lbrace> z' : B_pair b1 b2 | c \<rbrace> \<and> atom z \<sharp> V_pair v1 v2 "
- using infer_e_elims(9)[OF assms(1)] by metis
- hence *:" \<Theta> ; {||} ; GNil \<turnstile> V_pair v1 v2 \<Rightarrow> \<lbrace> z' : B_pair b1 b2 | c \<rbrace>" by auto
-
- obtain t1a and t2a where
- *: "\<Theta> ; {||} ; GNil \<turnstile> v1 \<Rightarrow> t1a \<and> \<Theta> ; {||} ; GNil \<turnstile> v2 \<Rightarrow> t2a \<and> B_pair b1 b2 = B_pair (b_of t1a) (b_of t2a)"
- using infer_v_elims(5)[OF *] by metis
-
- hence suppv: "supp v1 = {} \<and> supp v2 = {} \<and> supp (V_pair v1 v2) = {}" using infer_v_v_wf wfV.simps v.supp by (meson "**" wfV_supp_nil)
-
- obtain z2 and b2' where "t2a = \<lbrace> z2 : b2' | [ [ z2 ]\<^sup>v ]\<^sup>c\<^sup>e == [ v2 ]\<^sup>c\<^sup>e \<rbrace> "
- using infer_v_form[of \<Theta> "{||}" GNil v2 t2a] * by auto
- moreover hence "b2' = b2" using * b_of.simps by simp
-
- ultimately have "\<Theta> ; {||} ; GNil \<turnstile> v2 \<Rightarrow> \<lbrace> z2 : b2 | CE_val (V_var z2) == CE_val v2 \<rbrace>" using * by auto
- moreover have "\<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f CE_snd [V_pair v1 v2]\<^sup>c\<^sup>e : b2" using wfCE_sndI infer_v_wf(1) ** b_of.simps wfCE_valI by metis
- moreover hence st: "\<Theta> ; {||} ; GNil \<turnstile> \<lbrace> z2 : b2 | CE_val (V_var z2) == CE_val v2 \<rbrace> \<lesssim> (\<lbrace> z : b2 | CE_val (V_var z) == CE_snd [V_pair v1 v2]\<^sup>c\<^sup>e \<rbrace>)"
- using subtype_gnil_snd infer_v_v_wf by auto
- moreover have "wfD \<Theta> {||} GNil \<Delta> \<and> wfPhi \<Theta> \<Phi>" using ** by metis
- ultimately show ?thesis using wfX_wfY ** infer_e_valI by metis
-qed
-
-section \<open>Statements\<close>
-
-lemma check_s_v_unit:
- assumes "\<Theta>; \<B>; \<Gamma> \<turnstile> (\<lbrace> z : B_unit | TRUE \<rbrace>) \<lesssim> \<tau>" and "wfD \<Theta> \<B> \<Gamma> \<Delta>" and "wfPhi \<Theta> \<Phi>"
- shows "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile> AS_val (V_lit L_unit ) \<Leftarrow> \<tau>"
-proof -
- have "wfG \<Theta> \<B> \<Gamma>" using assms subtype_g_wf by meson
- moreover hence "wfTh \<Theta>" using wfG_wf by simp
- moreover obtain z'::x where "atom z' \<sharp> \<Gamma>" using obtain_fresh by auto
- ultimately have *:"\<Theta>; \<B>; \<Gamma> \<turnstile> V_lit L_unit \<Rightarrow> \<lbrace> z' : B_unit | CE_val (V_var z') == CE_val (V_lit L_unit) \<rbrace>"
- using infer_v_litI infer_unitI by simp
- moreover have "wfT \<Theta> \<B> \<Gamma> (\<lbrace> z' : B_unit | CE_val (V_var z') == CE_val (V_lit L_unit) \<rbrace>)" using infer_v_t_wf
- by (meson calculation)
- moreover then have "\<Theta>; \<B>; \<Gamma> \<turnstile> (\<lbrace> z' : B_unit | CE_val (V_var z') == CE_val (V_lit L_unit) \<rbrace>) \<lesssim> \<tau>" using subtype_trans subtype_top assms
- type_for_lit.simps(4) wfX_wfY by metis
- ultimately show ?thesis using check_valI assms * by auto
-qed
-
-lemma check_s_check_branch_s_wf:
- fixes s::s and cs::branch_s and \<Theta>::\<Theta> and \<Phi>::\<Phi> and \<Gamma>::\<Gamma> and \<Delta>::\<Delta> and v::v and \<tau>::\<tau> and css::branch_list
- shows "\<Theta> ; \<Phi> ; B ; \<Gamma> ; \<Delta> \<turnstile> s \<Leftarrow> \<tau> \<Longrightarrow> \<Theta> ; B \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<and> wfTh \<Theta> \<and> wfD \<Theta> B \<Gamma> \<Delta> \<and> wfT \<Theta> B \<Gamma> \<tau> \<and> wfPhi \<Theta> \<Phi> " and
- "check_branch_s \<Theta> \<Phi> B \<Gamma> \<Delta> tid cons const v cs \<tau> \<Longrightarrow> \<Theta> ; B \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<and> wfTh \<Theta> \<and> wfD \<Theta> B \<Gamma> \<Delta> \<and> wfT \<Theta> B \<Gamma> \<tau> \<and> wfPhi \<Theta> \<Phi> "
- "check_branch_list \<Theta> \<Phi> B \<Gamma> \<Delta> tid dclist v css \<tau> \<Longrightarrow> \<Theta> ; B \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<and> wfTh \<Theta> \<and> wfD \<Theta> B \<Gamma> \<Delta> \<and> wfT \<Theta> B \<Gamma> \<tau> \<and> wfPhi \<Theta> \<Phi> "
-proof(induct rule: check_s_check_branch_s_check_branch_list.inducts)
- case (check_valI \<Theta> B \<Gamma> \<Delta> \<Phi> v \<tau>' \<tau>)
- then show ?case using infer_v_wf infer_v_wf subtype_wf wfX_wfY wfS_valI
- by (metis subtype_eq_base2)
-next
- case (check_letI x \<Theta> \<Phi> \<B> \<Gamma> \<Delta> e \<tau> z s b c)
- then have *:"wfT \<Theta> \<B> ((x, b , c[z::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) \<tau>" by force
- moreover have "atom x \<sharp> \<tau>" using check_letI fresh_prodN by force
- ultimately have "\<Theta>; \<B>;\<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau>" using wfT_restrict2 by force
- then show ?case using check_letI infer_e_wf wfS_letI wfX_wfY wfG_elims by metis
-next
- case (check_assertI x \<Theta> \<Phi> \<B> \<Gamma> \<Delta> c \<tau> s)
- then have *:"wfT \<Theta> \<B> ((x, B_bool , c) #\<^sub>\<Gamma> \<Gamma>) \<tau>" by force
- moreover have "atom x \<sharp> \<tau>" using check_assertI fresh_prodN by force
- ultimately have "\<Theta>; \<B>;\<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau>" using wfT_restrict2 by force
- then show ?case using check_assertI wfS_assertI wfX_wfY wfG_elims by metis
-next
- case (check_branch_s_branchI \<Theta> \<B> \<Gamma> \<Delta> \<tau> cons const x v \<Phi> s tid)
- then show ?case using wfX_wfY by metis
-next
- case (check_branch_list_consI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist' v cs \<tau> css )
- then show ?case using wfX_wfY by metis
-next
- case (check_branch_list_finalI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist' v cs \<tau> )
- then show ?case using wfX_wfY by metis
-next
- case (check_ifI z \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v s1 s2 \<tau>)
- hence *:"wfT \<Theta> \<B> \<Gamma> (\<lbrace> z : b_of \<tau> | CE_val v == CE_val (V_lit L_false) IMP c_of \<tau> z \<rbrace>)" (is "wfT \<Theta> \<B> \<Gamma> ?tau") by auto
- hence "wfT \<Theta> \<B> \<Gamma> \<tau>" using wfT_wfT_if1 check_ifI fresh_prodN by metis
- hence " \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau> " using check_ifI b_of_c_of_eq fresh_prodN by auto
- thus ?case using check_ifI by metis
-next
- case (check_let2I x \<Theta> \<Phi> \<B> G \<Delta> t s1 \<tau> s2)
- then have "wfT \<Theta> \<B> ((x, b_of t, (c_of t x)) #\<^sub>\<Gamma> G) \<tau>" by fastforce
- moreover have "atom x \<sharp> \<tau>" using check_let2I by force
- ultimately have "wfT \<Theta> \<B> G \<tau>" using wfT_restrict2 by metis
- then show ?case using check_let2I by argo
-next
- case (check_varI u \<Delta> P G v \<tau>' \<Phi> s \<tau>)
- then show ?case using wfG_elims wfD_elims
- list.distinct list.inject by metis
-next
- case (check_assignI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> u \<tau> v z \<tau>')
- obtain z'::x where *:"atom z' \<sharp> \<Gamma>" using obtain_fresh by metis
- moreover have "\<lbrace> z : B_unit | TRUE \<rbrace> = \<lbrace> z' : B_unit | TRUE \<rbrace>" by auto
- moreover hence "wfT \<Theta> \<B> \<Gamma> \<lbrace> z' : B_unit |TRUE \<rbrace>" using wfT_TRUE check_assignI check_v_wf * wfB_unitI wfG_wf by metis
- ultimately show ?case using check_v.cases infer_v_wf subtype_wf check_assignI wfT_wf check_v_wf wfG_wf
- by (meson subtype_wf)
-next
- case (check_whileI \<Phi> \<Delta> G P s1 z s2 \<tau>')
- then show ?case using subtype_wf subtype_wf by auto
-next
- case (check_seqI \<Delta> G P s1 z s2 \<tau>)
- then show ?case by fast
-next
- case (check_caseI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> dclist cs \<tau> tid v z)
- then show ?case by fast
-qed
-
-lemma check_s_check_branch_s_wfS:
- fixes s::s and cs::branch_s and \<Theta>::\<Theta> and \<Phi>::\<Phi> and \<Gamma>::\<Gamma> and \<Delta>::\<Delta> and v::v and \<tau>::\<tau> and css::branch_list
- shows "\<Theta> ; \<Phi> ; B ; \<Gamma> ; \<Delta> \<turnstile> s \<Leftarrow> \<tau> \<Longrightarrow> \<Theta> ; \<Phi> ; B ; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b_of \<tau> " and
- "check_branch_s \<Theta> \<Phi> B \<Gamma> \<Delta> tid cons const v cs \<tau> \<Longrightarrow> wfCS \<Theta> \<Phi> B \<Gamma> \<Delta> tid cons const cs (b_of \<tau>)"
- "check_branch_list \<Theta> \<Phi> B \<Gamma> \<Delta> tid dclist v css \<tau> \<Longrightarrow> wfCSS \<Theta> \<Phi> B \<Gamma> \<Delta> tid dclist css (b_of \<tau>)"
-proof(induct rule: check_s_check_branch_s_check_branch_list.inducts)
- case (check_valI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v \<tau>' \<tau>)
- then show ?case using infer_v_wf infer_v_wf subtype_wf wfX_wfY wfS_valI
- by (metis subtype_eq_base2)
-next
- case (check_letI x \<Theta> \<Phi> \<B> \<Gamma> \<Delta> e \<tau> z s b c)
- show ?case proof
- show \<open> \<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f e : b \<close> using infer_e_wf check_letI b_of.simps by metis
- show \<open> \<Theta> ; \<Phi> ; \<B> ; (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b_of \<tau> \<close>
- using check_letI b_of.simps wf_replace_true2(2)[OF check_letI(5)] append_g.simps by metis
- show \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using infer_e_wf check_letI b_of.simps by metis
- show \<open>atom x \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>, \<Delta>, e, b_of \<tau>)\<close>
- apply(simp add: fresh_prodN, intro conjI)
- using check_letI(1) fresh_prod7 by simp+
- qed
-next
- case (check_assertI x \<Theta> \<Phi> \<B> \<Gamma> \<Delta> c \<tau> s)
- show ?case proof
-
- show \<open> \<Theta> ; \<Phi> ; \<B> ; (x, B_bool, c) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b_of \<tau> \<close> using check_assertI by auto
- next
- show \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f c \<close> using check_assertI by auto
- next
- show \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using check_assertI by auto
- next
- show \<open>atom x \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>, \<Delta>, c, b_of \<tau>, s)\<close> using check_assertI by auto
- qed
-next
- case (check_branch_s_branchI \<Theta> \<B> \<Gamma> \<Delta> \<tau> const x \<Phi> tid cons v s)
- show ?case proof
- show \<open> \<Theta> ; \<Phi> ; \<B> ; (x, b_of const, TRUE) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b_of \<tau> \<close>
- using wf_replace_true append_g.simps check_branch_s_branchI by metis
- show \<open>atom x \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>, \<Delta>, \<Gamma>, const)\<close>
- using wf_replace_true append_g.simps check_branch_s_branchI fresh_prodN by metis
- show \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wf_replace_true append_g.simps check_branch_s_branchI by metis
- qed
-next
- case (check_branch_list_consI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid cons const v cs \<tau> dclist css)
- then show ?case using wf_intros by metis
-next
- case (check_branch_list_finalI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid cons const v cs \<tau>)
- then show ?case using wf_intros by metis
-next
- case (check_ifI z \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v s1 s2 \<tau>)
- show ?case using wfS_ifI check_v_wf check_ifI b_of.simps by metis
-next
- case (check_let2I x \<Theta> \<Phi> \<B> G \<Delta> t s1 \<tau> s2)
- show ?case proof
- show \<open> \<Theta> ; \<Phi> ; \<B> ; G ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s1 : b_of t \<close> using check_let2I b_of.simps by metis
- show \<open> \<Theta>; \<B>; G \<turnstile>\<^sub>w\<^sub>f t \<close> using check_let2I check_s_check_branch_s_wf by metis
- show \<open> \<Theta> ; \<Phi> ; \<B> ; (x, b_of t, TRUE) #\<^sub>\<Gamma> G ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s2 : b_of \<tau> \<close>
- using check_let2I(5) wf_replace_true2(2) append_g.simps check_let2I by metis
- show \<open>atom x \<sharp> (\<Phi>, \<Theta>, \<B>, G, \<Delta>, s1, b_of \<tau>, t)\<close>
- apply(simp add: fresh_prodN, intro conjI)
- using check_let2I(1) fresh_prod7 by simp+
- qed
-next
- case (check_varI u \<Theta> \<Phi> \<B> \<Gamma> \<Delta> \<tau>' v \<tau> s)
- show ?case proof
- show \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau>' \<close> using check_v_wf check_varI by metis
- show \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b_of \<tau>' \<close> using check_v_wf check_varI by metis
- show \<open>atom u \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>, \<Delta>, \<tau>', v, b_of \<tau>)\<close> using check_varI fresh_prodN u_fresh_b by metis
- show \<open> \<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; (u, \<tau>') #\<^sub>\<Delta>\<Delta> \<turnstile>\<^sub>w\<^sub>f s : b_of \<tau> \<close> using check_varI by metis
- qed
-next
- case (check_assignI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> u \<tau> v z \<tau>')
- then show ?case using wf_intros check_v_wf subtype_eq_base2 b_of.simps by metis
-next
- case (check_whileI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> s1 z s2 \<tau>')
- thus ?case using wf_intros b_of.simps check_v_wf subtype_eq_base2 b_of.simps by metis
-next
- case (check_seqI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> s1 z s2 \<tau>)
- thus ?case using wf_intros b_of.simps by metis
-next
- case (check_caseI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist v cs \<tau> z)
- show ?case proof
- show \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : B_id tid \<close> using check_caseI check_v_wf b_of.simps by metis
- show \<open>AF_typedef tid dclist \<in> set \<Theta>\<close> using check_caseI by metis
- show \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using check_caseI check_s_check_branch_s_wf by metis
- show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using check_caseI check_s_check_branch_s_wf by metis
- show \<open> \<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> ; tid ; dclist \<turnstile>\<^sub>w\<^sub>f cs : b_of \<tau> \<close> using check_caseI by metis
- qed
-qed
-
-lemma check_s_wf:
- fixes s::s
- assumes "\<Theta> ; \<Phi> ; B ; \<Gamma> ; \<Delta> \<turnstile> s \<Leftarrow> \<tau>"
- shows "\<Theta> ; B \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<and> wfT \<Theta> B \<Gamma> \<tau> \<and> wfPhi \<Theta> \<Phi> \<and> wfTh \<Theta> \<and> wfD \<Theta> B \<Gamma> \<Delta> \<and> wfS \<Theta> \<Phi> B \<Gamma> \<Delta> s (b_of \<tau>)"
- using check_s_check_branch_s_wf check_s_check_branch_s_wfS assms by meson
-
-lemma check_s_flip_u1:
- fixes s::s and u::u and u'::u
- assumes "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile> s \<Leftarrow> \<tau>"
- shows "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; ( u \<leftrightarrow> u') \<bullet> \<Delta> \<turnstile> (u \<leftrightarrow> u') \<bullet> s \<Leftarrow> \<tau>"
-proof -
- have "(u \<leftrightarrow> u') \<bullet> \<Theta> ; (u \<leftrightarrow> u') \<bullet> \<Phi> ; (u \<leftrightarrow> u') \<bullet> \<B> ; (u \<leftrightarrow> u') \<bullet> \<Gamma> ; (u \<leftrightarrow> u') \<bullet> \<Delta> \<turnstile> (u \<leftrightarrow> u') \<bullet> s \<Leftarrow> (u \<leftrightarrow> u') \<bullet> \<tau>"
- using check_s.eqvt assms by blast
- thus ?thesis using check_s_wf[OF assms] flip_u_eq phi_flip_eq by metis
-qed
-
-lemma check_s_flip_u2:
- fixes s::s and u::u and u'::u
- assumes "\<Theta> ; \<Phi> ; B ; \<Gamma> ; ( u \<leftrightarrow> u') \<bullet> \<Delta> \<turnstile> (u \<leftrightarrow> u') \<bullet> s \<Leftarrow> \<tau>"
- shows "\<Theta> ; \<Phi> ; B ; \<Gamma> ; \<Delta> \<turnstile> s \<Leftarrow> \<tau>"
-proof -
- have "\<Theta> ; \<Phi> ; B ; \<Gamma> ; ( u \<leftrightarrow> u') \<bullet> ( u \<leftrightarrow> u') \<bullet> \<Delta> \<turnstile> ( u \<leftrightarrow> u') \<bullet> (u \<leftrightarrow> u') \<bullet> s \<Leftarrow> \<tau>"
- using check_s_flip_u1 assms by blast
- thus ?thesis using permute_flip_cancel by force
-qed
-
-lemma check_s_flip_u:
- fixes s::s and u::u and u'::u
- shows "\<Theta> ; \<Phi> ; B ; \<Gamma> ; ( u \<leftrightarrow> u') \<bullet> \<Delta> \<turnstile> (u \<leftrightarrow> u') \<bullet> s \<Leftarrow> \<tau> = (\<Theta> ; \<Phi> ; B ; \<Gamma> ; \<Delta> \<turnstile> s \<Leftarrow> \<tau>)"
- using check_s_flip_u1 check_s_flip_u2 by metis
-
-lemma check_s_abs_u:
- fixes s::s and s'::s and u::u and u'::u and \<tau>'::\<tau>
- assumes "[[atom u]]lst. s = [[atom u']]lst. s'" and "atom u \<sharp> \<Delta>" and "atom u' \<sharp> \<Delta>"
- and "\<Theta> ; B ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau>'"
- and "\<Theta> ; \<Phi> ; B ; \<Gamma> ; ( u , \<tau>') #\<^sub>\<Delta>\<Delta> \<turnstile> s \<Leftarrow> \<tau>"
- shows "\<Theta> ; \<Phi> ; B ; \<Gamma> ; ( u', \<tau>' ) #\<^sub>\<Delta>\<Delta> \<turnstile> s' \<Leftarrow> \<tau>"
-proof -
- have "\<Theta> ; \<Phi> ; B ; \<Gamma> ; ( u' \<leftrightarrow> u) \<bullet> (( u , \<tau>') #\<^sub>\<Delta>\<Delta>) \<turnstile> ( u' \<leftrightarrow> u) \<bullet> s \<Leftarrow> \<tau>"
- using assms check_s_flip_u by metis
- moreover have " ( u' \<leftrightarrow> u) \<bullet> (( u , \<tau>') #\<^sub>\<Delta> \<Delta>) = ( u' , \<tau>') #\<^sub>\<Delta>\<Delta>" proof -
- have " ( u' \<leftrightarrow> u) \<bullet> (( u , \<tau>') #\<^sub>\<Delta> \<Delta>) = ((u' \<leftrightarrow> u) \<bullet> u, (u' \<leftrightarrow> u) \<bullet> \<tau>') #\<^sub>\<Delta> (u' \<leftrightarrow> u) \<bullet> \<Delta>"
- using DCons_eqvt Pair_eqvt by auto
- also have "... = ( u' , (u' \<leftrightarrow> u) \<bullet> \<tau>') #\<^sub>\<Delta> \<Delta>"
- using assms flip_fresh_fresh by auto
- also have "... = ( u' , \<tau>') #\<^sub>\<Delta> \<Delta>" using
- u_not_in_t fresh_def flip_fresh_fresh assms by metis
- finally show ?thesis by auto
- qed
- moreover have "( u' \<leftrightarrow> u) \<bullet> s = s'" using assms Abs1_eq_iff(3)[of u' s' u s] by auto
- ultimately show ?thesis by auto
-qed
-
-section \<open>Additional Elimination and Intros\<close>
-
-subsection \<open>Values\<close>
-
-nominal_function b_for :: "opp \<Rightarrow> b" where
- "b_for Plus = B_int"
-| "b_for LEq = B_bool" | "b_for Eq = B_bool"
- apply(auto,simp add: eqvt_def b_for_graph_aux_def )
- by (meson opp.exhaust)
-nominal_termination (eqvt) by lexicographic_order
-
-
-lemma infer_v_pair2I:
- fixes v\<^sub>1::v and v\<^sub>2::v
- assumes "\<Theta>; \<B>; \<Gamma> \<turnstile> v\<^sub>1 \<Rightarrow> \<tau>\<^sub>1" and "\<Theta>; \<B>; \<Gamma> \<turnstile> v\<^sub>2 \<Rightarrow> \<tau>\<^sub>2"
- shows "\<exists>\<tau>. \<Theta>; \<B>; \<Gamma> \<turnstile> V_pair v\<^sub>1 v\<^sub>2 \<Rightarrow> \<tau> \<and> b_of \<tau> = B_pair (b_of \<tau>\<^sub>1) (b_of \<tau>\<^sub>2)"
-proof -
- obtain z1 and b1 and c1 and z2 and b2 and c2 where zbc: "\<tau>\<^sub>1 = (\<lbrace> z1 : b1 | c1 \<rbrace>) \<and> \<tau>\<^sub>2 = (\<lbrace> z2 : b2 | c2 \<rbrace>)"
- using \<tau>.exhaust by meson
- obtain z::x where "atom z \<sharp> ( v\<^sub>1, v\<^sub>2,\<Theta>, \<B>,\<Gamma>)" using obtain_fresh
- by blast
- hence "atom z \<sharp> ( v\<^sub>1, v\<^sub>2) \<and> atom z \<sharp> (\<Theta>, \<B>,\<Gamma>)" using fresh_prodN by metis
- hence " \<Theta>; \<B>; \<Gamma> \<turnstile> V_pair v\<^sub>1 v\<^sub>2 \<Rightarrow> \<lbrace> z : [ b_of \<tau>\<^sub>1 , b_of \<tau>\<^sub>2 ]\<^sup>b | CE_val (V_var z) == CE_val (V_pair v\<^sub>1 v\<^sub>2) \<rbrace>"
- using assms infer_v_pairI zbc by metis
- moreover obtain \<tau> where "\<tau> = (\<lbrace> z : B_pair b1 b2 | CE_val (V_var z) == CE_val (V_pair v\<^sub>1 v\<^sub>2) \<rbrace>)" by blast
- moreover hence "b_of \<tau> = B_pair (b_of \<tau>\<^sub>1) (b_of \<tau>\<^sub>2)" using b_of.simps zbc by presburger
- ultimately show ?thesis using b_of.simps by metis
-qed
-
-lemma infer_v_pair2I_zbc:
- fixes v\<^sub>1::v and v\<^sub>2::v
- assumes "\<Theta>; \<B>; \<Gamma> \<turnstile> v\<^sub>1 \<Rightarrow> \<tau>\<^sub>1" and "\<Theta>; \<B>; \<Gamma> \<turnstile> v\<^sub>2 \<Rightarrow> \<tau>\<^sub>2"
- shows "\<exists>z \<tau>. \<Theta>; \<B>; \<Gamma> \<turnstile> V_pair v\<^sub>1 v\<^sub>2 \<Rightarrow> \<tau> \<and> \<tau> = (\<lbrace> z : B_pair (b_of \<tau>\<^sub>1) (b_of \<tau>\<^sub>2) | C_eq (CE_val (V_var z)) (CE_val (V_pair v\<^sub>1 v\<^sub>2)) \<rbrace>) \<and> atom z \<sharp> (v\<^sub>1,v\<^sub>2) \<and> atom z \<sharp> \<Gamma>"
-proof -
- obtain z1 and b1 and c1 and z2 and b2 and c2 where zbc: "\<tau>\<^sub>1 = (\<lbrace> z1 : b1 | c1 \<rbrace>) \<and> \<tau>\<^sub>2 = (\<lbrace> z2 : b2 | c2 \<rbrace>)"
- using \<tau>.exhaust by meson
- obtain z::x where * : "atom z \<sharp> ( v\<^sub>1, v\<^sub>2,\<Gamma>,\<Theta> , \<B> )" using obtain_fresh
- by blast
- hence vinf: " \<Theta>; \<B>; \<Gamma> \<turnstile> V_pair v\<^sub>1 v\<^sub>2 \<Rightarrow> \<lbrace> z :[ b_of \<tau>\<^sub>1 , b_of \<tau>\<^sub>2 ]\<^sup>b | CE_val (V_var z) == CE_val (V_pair v\<^sub>1 v\<^sub>2) \<rbrace>"
- using assms infer_v_pairI by simp
- moreover obtain \<tau> where "\<tau> = (\<lbrace> z : B_pair b1 b2 | CE_val (V_var z) == CE_val (V_pair v\<^sub>1 v\<^sub>2) \<rbrace>)" by blast
- moreover have "b_of \<tau>\<^sub>1 = b1 \<and> b_of \<tau>\<^sub>2 = b2" using zbc b_of.simps by auto
- ultimately have "\<Theta>; \<B>; \<Gamma> \<turnstile> V_pair v\<^sub>1 v\<^sub>2 \<Rightarrow> \<tau> \<and> \<tau> = (\<lbrace> z : B_pair (b_of \<tau>\<^sub>1) (b_of \<tau>\<^sub>2) | CE_val (V_var z) == CE_val (V_pair v\<^sub>1 v\<^sub>2) \<rbrace>)" by auto
- thus ?thesis using * fresh_prod2 fresh_prod3 by metis
-qed
-
-lemma infer_v_pair2E:
- assumes "\<Theta>; \<B>; \<Gamma> \<turnstile> V_pair v\<^sub>1 v\<^sub>2 \<Rightarrow> \<tau>"
- shows "\<exists>\<tau>\<^sub>1 \<tau>\<^sub>2 z . \<Theta>; \<B>; \<Gamma> \<turnstile> v\<^sub>1 \<Rightarrow> \<tau>\<^sub>1 \<and> \<Theta>; \<B>; \<Gamma> \<turnstile> v\<^sub>2 \<Rightarrow> \<tau>\<^sub>2 \<and>
- \<tau> = (\<lbrace> z : B_pair (b_of \<tau>\<^sub>1) (b_of \<tau>\<^sub>2) | C_eq (CE_val (V_var z)) (CE_val (V_pair v\<^sub>1 v\<^sub>2)) \<rbrace>) \<and> atom z \<sharp> (v\<^sub>1, v\<^sub>2)"
-proof -
- obtain z and t1 and t2 where
- "\<tau> = (\<lbrace> z : B_pair (b_of t1) (b_of t2) | CE_val (V_var z) == CE_val (V_pair v\<^sub>1 v\<^sub>2) \<rbrace>) \<and>
- atom z \<sharp> (v\<^sub>1, v\<^sub>2) \<and> \<Theta>; \<B>; \<Gamma> \<turnstile> v\<^sub>1 \<Rightarrow> t1 \<and> \<Theta>; \<B>; \<Gamma> \<turnstile> v\<^sub>2 \<Rightarrow> t2 " using infer_v_elims(3)[OF assms ] by metis
- thus ?thesis using b_of.simps by metis
-qed
-
-subsection \<open>Expressions\<close>
-
-lemma infer_e_app2E:
- fixes \<Phi>::\<Phi> and \<Theta>::\<Theta>
- assumes "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile> AE_app f v \<Rightarrow> \<tau>"
- shows "\<exists>x b c s' \<tau>'. wfD \<Theta> \<B> \<Gamma> \<Delta> \<and> Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \<tau>' s'))) = lookup_fun \<Phi> f \<and> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<and>
- \<Theta>; \<B>; \<Gamma> \<turnstile> v \<Leftarrow> \<lbrace> x : b | c \<rbrace> \<and> \<tau> = \<tau>'[x::=v]\<^sub>\<tau>\<^sub>v \<and> atom x \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>, \<Delta>, v, \<tau>)"
- using infer_e_elims(6)[OF assms] b_of.simps subst_defs by metis
-
-lemma infer_e_appP2E:
- fixes \<Phi>::\<Phi> and \<Theta>::\<Theta>
- assumes "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile> AE_appP f b v \<Rightarrow> \<tau>"
- shows "\<exists>bv x ba c s' \<tau>'. wfD \<Theta> \<B> \<Gamma> \<Delta> \<and> Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x ba c \<tau>' s'))) = lookup_fun \<Phi> f \<and> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<and> \<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f b \<and>
- (\<Theta>; \<B>; \<Gamma> \<turnstile> v \<Leftarrow> \<lbrace> x : ba[bv::=b]\<^sub>b\<^sub>b | c[bv::=b]\<^sub>c\<^sub>b \<rbrace>) \<and> (\<tau> = \<tau>'[bv::=b]\<^sub>\<tau>\<^sub>b[x::=v]\<^sub>\<tau>\<^sub>v) \<and> atom x \<sharp> \<Gamma> \<and> atom bv \<sharp> v"
-proof -
- obtain bv x ba c s' \<tau>' where *:"wfD \<Theta> \<B> \<Gamma> \<Delta> \<and> Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x ba c \<tau>' s'))) = lookup_fun \<Phi> f \<and> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<and> \<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f b \<and>
- (\<Theta>; \<B>; \<Gamma> \<turnstile> v \<Leftarrow> \<lbrace> x : ba[bv::=b]\<^sub>b\<^sub>b | c[bv::=b]\<^sub>c\<^sub>b \<rbrace>) \<and> (\<tau> = \<tau>'[bv::=b]\<^sub>\<tau>\<^sub>b[x::=v]\<^sub>\<tau>\<^sub>v) \<and> atom x \<sharp> \<Gamma> \<and> atom bv \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>, \<Delta>, b, v, \<tau>)"
- using infer_e_elims(21)[OF assms] subst_defs by metis
- moreover then have "atom bv \<sharp> v" using fresh_prodN by metis
- ultimately show ?thesis by metis
-qed
-
-section \<open>Weakening\<close>
-
-text \<open> Lemmas showing that typing judgements hold when a context is extended \<close>
-
-lemma subtype_weakening:
- fixes \<Gamma>'::\<Gamma>
- assumes "\<Theta>; \<B>; \<Gamma> \<turnstile> \<tau>1 \<lesssim> \<tau>2" and "toSet \<Gamma> \<subseteq> toSet \<Gamma>'" and "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>'"
- shows "\<Theta>; \<B>; \<Gamma>' \<turnstile> \<tau>1 \<lesssim> \<tau>2"
- using assms proof(nominal_induct \<tau>2 avoiding: \<Gamma>' rule: subtype.strong_induct)
-
- case (subtype_baseI x \<Theta> \<B> \<Gamma> z c z' c' b)
- show ?case proof
- show *:"\<Theta>; \<B>; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b | c \<rbrace>" using wfT_weakening subtype_baseI by metis
- show "\<Theta>; \<B>; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f \<lbrace> z' : b | c' \<rbrace>" using wfT_weakening subtype_baseI by metis
- show "atom x \<sharp> (\<Theta>, \<B>, \<Gamma>', z , c , z' , c' )" using subtype_baseI fresh_Pair by metis
- have "toSet ((x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) \<subseteq> toSet ((x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>')" using subtype_baseI toSet.simps by blast
- moreover have "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f (x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>'" using wfT_wf_cons3[OF *, of x] subtype_baseI fresh_Pair subst_defs by metis
- ultimately show "\<Theta>; \<B>; (x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>' \<Turnstile> c'[z'::=V_var x]\<^sub>v" using valid_weakening subtype_baseI by metis
- qed
-qed
-
-method many_rules uses add = ( (rule+),((simp add: add)+)?)
-
-lemma infer_v_g_weakening:
- fixes e::e and \<Gamma>'::\<Gamma> and v::v
- assumes "\<Theta>; \<B> ; \<Gamma> \<turnstile> v \<Rightarrow> \<tau>" and "toSet \<Gamma> \<subseteq> toSet \<Gamma>'" and "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>'"
- shows "\<Theta>; \<B> ; \<Gamma>' \<turnstile> v \<Rightarrow> \<tau>"
- using assms proof(nominal_induct avoiding: \<Gamma>' rule: infer_v.strong_induct)
- case (infer_v_varI \<Theta> \<B> \<Gamma> b c x' z)
- show ?case proof
- show \<open> \<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>' \<close> using infer_v_varI by auto
- show \<open>Some (b, c) = lookup \<Gamma>' x'\<close> using infer_v_varI lookup_weakening by metis
- show \<open>atom z \<sharp> x'\<close> using infer_v_varI by auto
- show \<open>atom z \<sharp> (\<Theta>, \<B>, \<Gamma>')\<close> using infer_v_varI by auto
- qed
-next
- case (infer_v_litI \<Theta> \<B> \<Gamma> l \<tau>)
- then show ?case using infer_v.intros by simp
-next
- case (infer_v_pairI z v1 v2 \<Theta> \<B> \<Gamma> t1 t2)
- then show ?case using infer_v.intros by simp
-next
- case (infer_v_consI s dclist \<Theta> dc tc \<B> \<Gamma> v tv z)
- show ?case proof
- show \<open>AF_typedef s dclist \<in> set \<Theta>\<close> using infer_v_consI by auto
- show \<open>(dc, tc) \<in> set dclist\<close> using infer_v_consI by auto
- show \<open> \<Theta>; \<B>; \<Gamma>' \<turnstile> v \<Rightarrow> tv\<close> using infer_v_consI by auto
- show \<open>\<Theta>; \<B>; \<Gamma>' \<turnstile> tv \<lesssim> tc\<close> using infer_v_consI subtype_weakening by auto
- show \<open>atom z \<sharp> v\<close> using infer_v_consI by auto
- show \<open>atom z \<sharp> (\<Theta>, \<B>, \<Gamma>')\<close> using infer_v_consI by auto
- qed
-next
- case (infer_v_conspI s bv dclist \<Theta> dc tc \<B> \<Gamma> v tv b z)
- show ?case proof
- show \<open>AF_typedef_poly s bv dclist \<in> set \<Theta>\<close> using infer_v_conspI by auto
- show \<open>(dc, tc) \<in> set dclist\<close> using infer_v_conspI by auto
- show \<open> \<Theta>; \<B>; \<Gamma>' \<turnstile> v \<Rightarrow> tv\<close> using infer_v_conspI by auto
- show \<open>\<Theta>; \<B>; \<Gamma>' \<turnstile> tv \<lesssim> tc[bv::=b]\<^sub>\<tau>\<^sub>b\<close> using infer_v_conspI subtype_weakening by auto
- show \<open>atom z \<sharp> (\<Theta>, \<B>, \<Gamma>', v, b)\<close> using infer_v_conspI by auto
- show \<open>atom bv \<sharp> (\<Theta>, \<B>, \<Gamma>', v, b)\<close> using infer_v_conspI by auto
- show \<open> \<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f b \<close> using infer_v_conspI by auto
- qed
-qed
-
-lemma check_v_g_weakening:
- fixes e::e and \<Gamma>'::\<Gamma>
- assumes "\<Theta>; \<B> ; \<Gamma> \<turnstile> v \<Leftarrow> \<tau>" and "toSet \<Gamma> \<subseteq> toSet \<Gamma>'" and "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>'"
- shows "\<Theta>; \<B> ; \<Gamma>' \<turnstile> v \<Leftarrow> \<tau>"
- using subtype_weakening infer_v_g_weakening check_v_elims check_v_subtypeI assms by metis
-
-lemma infer_e_g_weakening:
- fixes e::e and \<Gamma>'::\<Gamma>
- assumes "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile> e \<Rightarrow> \<tau>" and "toSet \<Gamma> \<subseteq> toSet \<Gamma>'" and "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>'"
- shows "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>'; \<Delta> \<turnstile> e \<Rightarrow> \<tau>"
- using assms proof(nominal_induct \<tau> avoiding: \<Gamma>' rule: infer_e.strong_induct)
- case (infer_e_valI \<Theta> \<B> \<Gamma> \<Delta>' \<Phi> v \<tau>)
- then show ?case using infer_v_g_weakening wf_weakening infer_e.intros by metis
-next
- case (infer_e_plusI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v1 z1 c1 v2 z2 c2 z3)
-
- obtain z'::x where z': "atom z' \<sharp> v1 \<and> atom z' \<sharp> v2 \<and> atom z' \<sharp> \<Gamma>'" using obtain_fresh fresh_prod3 by metis
- moreover hence *:"\<lbrace> z3 : B_int | CE_val (V_var z3) == CE_op Plus [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \<rbrace> = (\<lbrace> z' : B_int | CE_val (V_var z') == CE_op Plus [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \<rbrace>)"
- using infer_e_plusI type_e_eq ce.fresh fresh_e_opp by auto
-
- have "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>' ; \<Delta> \<turnstile> AE_op Plus v1 v2 \<Rightarrow> \<lbrace> z' : B_int | CE_val (V_var z') == CE_op Plus [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \<rbrace>" proof
- show \<open> \<Theta> ; \<B> ; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wf_weakening infer_e_plusI by auto
- show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using infer_e_plusI by auto
- show \<open> \<Theta> ; \<B> ; \<Gamma>' \<turnstile> v1 \<Rightarrow> \<lbrace> z1 : B_int | c1 \<rbrace>\<close> using infer_v_g_weakening infer_e_plusI by auto
- show \<open> \<Theta> ; \<B> ; \<Gamma>' \<turnstile> v2 \<Rightarrow> \<lbrace> z2 : B_int | c2 \<rbrace>\<close> using infer_v_g_weakening infer_e_plusI by auto
- show \<open>atom z' \<sharp> AE_op Plus v1 v2\<close> using z' by auto
- show \<open>atom z' \<sharp> \<Gamma>'\<close> using z' by auto
- qed
- thus ?case using * by metis
-
-next
- case (infer_e_leqI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v1 z1 c1 v2 z2 c2 z3)
- obtain z'::x where z': "atom z' \<sharp> v1 \<and> atom z' \<sharp> v2 \<and> atom z' \<sharp> \<Gamma>'" using obtain_fresh fresh_prod3 by metis
-
- moreover hence *:"\<lbrace> z3 : B_bool | CE_val (V_var z3) == CE_op LEq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \<rbrace> = (\<lbrace> z' : B_bool | CE_val (V_var z') == CE_op LEq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \<rbrace>)"
- using infer_e_leqI type_e_eq ce.fresh fresh_e_opp by auto
-
- have "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>' ; \<Delta> \<turnstile> AE_op LEq v1 v2 \<Rightarrow> \<lbrace> z' : B_bool | CE_val (V_var z') == CE_op LEq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \<rbrace>" proof
- show \<open> \<Theta> ; \<B> ; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wf_weakening infer_e_leqI by auto
- show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using infer_e_leqI by auto
- show \<open> \<Theta> ; \<B> ; \<Gamma>' \<turnstile> v1 \<Rightarrow> \<lbrace> z1 : B_int | c1 \<rbrace>\<close> using infer_v_g_weakening infer_e_leqI by auto
- show \<open> \<Theta> ; \<B> ; \<Gamma>' \<turnstile> v2 \<Rightarrow> \<lbrace> z2 : B_int | c2 \<rbrace>\<close> using infer_v_g_weakening infer_e_leqI by auto
- show \<open>atom z' \<sharp> AE_op LEq v1 v2\<close> using z' by auto
- show \<open>atom z' \<sharp> \<Gamma>'\<close> using z' by auto
- qed
- thus ?case using * by metis
-next
- case (infer_e_eqI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v1 z1 bb c1 v2 z2 c2 z3)
- obtain z'::x where z': "atom z' \<sharp> v1 \<and> atom z' \<sharp> v2 \<and> atom z' \<sharp> \<Gamma>'" using obtain_fresh fresh_prod3 by metis
-
- moreover hence *:"\<lbrace> z3 : B_bool | CE_val (V_var z3) == CE_op Eq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \<rbrace> = (\<lbrace> z' : B_bool | CE_val (V_var z') == CE_op Eq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \<rbrace>)"
- using infer_e_eqI type_e_eq ce.fresh fresh_e_opp by auto
-
- have "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>' ; \<Delta> \<turnstile> AE_op Eq v1 v2 \<Rightarrow> \<lbrace> z' : B_bool | CE_val (V_var z') == CE_op Eq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \<rbrace>" proof
- show \<open> \<Theta> ; \<B> ; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wf_weakening infer_e_eqI by auto
- show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using infer_e_eqI by auto
- show \<open> \<Theta> ; \<B> ; \<Gamma>' \<turnstile> v1 \<Rightarrow> \<lbrace> z1 : bb | c1 \<rbrace>\<close> using infer_v_g_weakening infer_e_eqI by auto
- show \<open> \<Theta> ; \<B> ; \<Gamma>' \<turnstile> v2 \<Rightarrow> \<lbrace> z2 : bb | c2 \<rbrace>\<close> using infer_v_g_weakening infer_e_eqI by auto
- show \<open>atom z' \<sharp> AE_op Eq v1 v2\<close> using z' by auto
- show \<open>atom z' \<sharp> \<Gamma>'\<close> using z' by auto
- show "bb \<in> {B_bool, B_int, B_unit}" using infer_e_eqI by auto
- qed
- thus ?case using * by metis
-next
- case (infer_e_appI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> f x b c \<tau>' s' v \<tau>)
- show ?case proof
- show "\<Theta>; \<B>; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f \<Delta> " using wf_weakening infer_e_appI by auto
- show "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>" using wf_weakening infer_e_appI by auto
- show "Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \<tau>' s'))) = lookup_fun \<Phi> f" using wf_weakening infer_e_appI by auto
- show "\<Theta>; \<B>; \<Gamma>' \<turnstile> v \<Leftarrow> \<lbrace> x : b | c \<rbrace>" using wf_weakening infer_e_appI check_v_g_weakening by auto
- show "atom x \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>', \<Delta>, v, \<tau>)" using wf_weakening infer_e_appI by auto
- show "\<tau>'[x::=v]\<^sub>v = \<tau>" using wf_weakening infer_e_appI by auto
- qed
-
-next
- case (infer_e_appPI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> b' f bv x b c \<tau>' s' v \<tau>)
-
- hence *:"\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile> AE_appP f b' v \<Rightarrow> \<tau>" using Typing.infer_e_appPI by auto
-
- obtain x'::x where x':"atom x' \<sharp> (s', c, \<tau>', (\<Gamma>',v,\<tau>)) \<and> (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \<tau>' s'))) = (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x' b ((x' \<leftrightarrow> x) \<bullet> c) ((x' \<leftrightarrow> x) \<bullet> \<tau>') ((x' \<leftrightarrow> x) \<bullet> s'))))"
- using obtain_fresh_fun_def[of s' c \<tau>' "(\<Gamma>',v,\<tau>)" f x b]
- by (metis fun_def.eq_iff fun_typ_q.eq_iff(2))
-
- hence **: " \<lbrace> x : b | c \<rbrace> = \<lbrace> x' : b | (x' \<leftrightarrow> x) \<bullet> c \<rbrace>"
- using fresh_PairD(1) fresh_PairD(2) type_eq_flip by blast
- have "atom x' \<sharp> \<Gamma>" using x' infer_e_appPI fresh_weakening fresh_Pair by metis
-
- show ?case proof(rule Typing.infer_e_appPI[where x=x' and bv=bv and b=b and c="(x' \<leftrightarrow> x) \<bullet> c" and \<tau>'="(x' \<leftrightarrow> x) \<bullet> \<tau>'"and s'="((x' \<leftrightarrow> x) \<bullet> s')"])
- show \<open> \<Theta> ; \<B> ; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wf_weakening infer_e_appPI by auto
- show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using wf_weakening infer_e_appPI by auto
- show "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f b'" using infer_e_appPI by auto
- show "Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x' b ((x' \<leftrightarrow> x) \<bullet> c) ((x' \<leftrightarrow> x) \<bullet> \<tau>') ((x' \<leftrightarrow> x) \<bullet> s')))) = lookup_fun \<Phi> f" using x' infer_e_appPI by argo
- show "\<Theta>; \<B>; \<Gamma>' \<turnstile> v \<Leftarrow> \<lbrace> x' : b[bv::=b']\<^sub>b | ((x' \<leftrightarrow> x) \<bullet> c)[bv::=b']\<^sub>b \<rbrace>" using **
- \<tau>.eq_iff check_v_g_weakening infer_e_appPI.hyps infer_e_appPI.prems(1) infer_e_appPI.prems subst_defs
- subst_tb.simps by metis
- show "atom x' \<sharp> \<Gamma>'" using x' fresh_prodN by metis
-
- have "atom x \<sharp> (v, \<tau>) \<and> atom x' \<sharp> (v, \<tau>)" using x' infer_e_fresh[OF *] e.fresh(2) fresh_Pair infer_e_appPI \<open>atom x' \<sharp> \<Gamma>\<close> e.fresh by metis
- moreover then have "((x' \<leftrightarrow> x) \<bullet> \<tau>')[bv::=b']\<^sub>\<tau>\<^sub>b = (x' \<leftrightarrow> x) \<bullet> (\<tau>'[bv::=b']\<^sub>\<tau>\<^sub>b)" using subst_b_x_flip
- by (metis subst_b_\<tau>_def)
- ultimately show "((x' \<leftrightarrow> x) \<bullet> \<tau>')[bv::=b']\<^sub>b[x'::=v]\<^sub>v = \<tau>"
- using infer_e_appPI subst_tv_flip subst_defs by metis
-
- show "atom bv \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>', \<Delta>, b', v, \<tau>)" using infer_e_appPI fresh_prodN by metis
- qed
-
-next
- case (infer_e_fstI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v z'' b1 b2 c z)
-
- obtain z'::x where *: "atom z' \<sharp> \<Gamma>' \<and> atom z' \<sharp> v \<and> atom z' \<sharp> c" using obtain_fresh_z fresh_Pair by metis
- hence **:"\<lbrace> z : b1 | CE_val (V_var z) == CE_fst [v]\<^sup>c\<^sup>e \<rbrace> = \<lbrace> z' : b1 | CE_val (V_var z') == CE_fst [v]\<^sup>c\<^sup>e \<rbrace>"
- using type_e_eq infer_e_fstI v.fresh e.fresh ce.fresh obtain_fresh_z fresh_Pair by metis
-
- have "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>' ; \<Delta> \<turnstile> AE_fst v \<Rightarrow> \<lbrace> z' : b1 | CE_val (V_var z') == CE_fst [v]\<^sup>c\<^sup>e \<rbrace>" proof
- show \<open> \<Theta> ; \<B> ; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wf_weakening infer_e_fstI by auto
- show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using wf_weakening infer_e_fstI by auto
- show "\<Theta> ; \<B> ; \<Gamma>' \<turnstile> v \<Rightarrow> \<lbrace> z'' : B_pair b1 b2 | c \<rbrace>" using infer_v_g_weakening infer_e_fstI by metis
- show "atom z' \<sharp> AE_fst v" using * ** e.supp by auto
- show "atom z' \<sharp> \<Gamma>'" using * by auto
- qed
- thus ?case using * ** by metis
-next
- case (infer_e_sndI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v z'' b1 b2 c z)
-
- obtain z'::x where *: "atom z' \<sharp> \<Gamma>' \<and> atom z' \<sharp> v \<and> atom z' \<sharp> c" using obtain_fresh_z fresh_Pair by metis
- hence **:"\<lbrace> z : b2 | CE_val (V_var z) == CE_snd [v]\<^sup>c\<^sup>e \<rbrace> = \<lbrace> z' : b2 | CE_val (V_var z') == CE_snd [v]\<^sup>c\<^sup>e \<rbrace>"
- using type_e_eq infer_e_sndI e.fresh ce.fresh obtain_fresh_z fresh_Pair by metis
-
- have "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>' ; \<Delta> \<turnstile> AE_snd v \<Rightarrow> \<lbrace> z' : b2 | CE_val (V_var z') == CE_snd [v]\<^sup>c\<^sup>e \<rbrace>" proof
- show \<open> \<Theta> ; \<B> ;\<Gamma>' \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wf_weakening infer_e_sndI by auto
- show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using wf_weakening infer_e_sndI by auto
- show "\<Theta> ; \<B> ; \<Gamma>' \<turnstile> v \<Rightarrow> \<lbrace> z'' : B_pair b1 b2 | c \<rbrace>" using infer_v_g_weakening infer_e_sndI by metis
- show "atom z' \<sharp> AE_snd v" using * e.supp by auto
- show "atom z' \<sharp> \<Gamma>'" using * by auto
- qed
- thus ?case using ** by metis
-next
- case (infer_e_lenI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v z'' c z)
-
- obtain z'::x where *: "atom z' \<sharp> \<Gamma>' \<and> atom z' \<sharp> v \<and> atom z' \<sharp> c" using obtain_fresh_z fresh_Pair by metis
- hence **:"\<lbrace> z : B_int | CE_val (V_var z) == CE_len [v]\<^sup>c\<^sup>e \<rbrace> = \<lbrace> z' : B_int | CE_val (V_var z') == CE_len [v]\<^sup>c\<^sup>e \<rbrace>"
- using type_e_eq infer_e_lenI e.fresh ce.fresh obtain_fresh_z fresh_Pair by metis
-
- have "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>' ; \<Delta> \<turnstile> AE_len v \<Rightarrow> \<lbrace> z' : B_int | CE_val (V_var z') == CE_len [v]\<^sup>c\<^sup>e \<rbrace>" proof
- show \<open> \<Theta>; \<B>; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wf_weakening infer_e_lenI by auto
- show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using wf_weakening infer_e_lenI by auto
- show "\<Theta> ; \<B> ; \<Gamma>' \<turnstile> v \<Rightarrow> \<lbrace> z'' : B_bitvec | c \<rbrace>" using infer_v_g_weakening infer_e_lenI by metis
- show "atom z' \<sharp> AE_len v" using * e.supp by auto
- show "atom z' \<sharp> \<Gamma>'" using * by auto
- qed
- thus ?case using * ** by metis
-next
- case (infer_e_mvarI \<Theta> \<Gamma> \<Phi> \<Delta> u \<tau>)
- then show ?case using wf_weakening infer_e.intros by metis
-next
- case (infer_e_concatI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v1 z1 c1 v2 z2 c2 z3)
-
- obtain z'::x where *: "atom z' \<sharp> \<Gamma>' \<and> atom z' \<sharp> v1 \<and> atom z' \<sharp> v2" using obtain_fresh_z fresh_Pair by metis
- hence **:"\<lbrace> z3 : B_bitvec | CE_val (V_var z3) == CE_concat [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \<rbrace> = \<lbrace> z' : B_bitvec | CE_val (V_var z') == CE_concat [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \<rbrace>"
- using type_e_eq infer_e_concatI e.fresh ce.fresh obtain_fresh_z fresh_Pair by metis
-
- have "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>' ; \<Delta> \<turnstile> AE_concat v1 v2 \<Rightarrow> \<lbrace> z' : B_bitvec | CE_val (V_var z') == CE_concat [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \<rbrace>" proof
- show \<open> \<Theta>; \<B>; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wf_weakening infer_e_concatI by auto
- show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using wf_weakening infer_e_concatI by auto
- show "\<Theta> ; \<B> ; \<Gamma>' \<turnstile> v1 \<Rightarrow> \<lbrace> z1 : B_bitvec | c1 \<rbrace>" using infer_v_g_weakening infer_e_concatI by metis
- show "\<Theta> ; \<B> ; \<Gamma>' \<turnstile> v2 \<Rightarrow> \<lbrace> z2 : B_bitvec | c2 \<rbrace>" using infer_v_g_weakening infer_e_concatI by metis
- show "atom z' \<sharp> AE_concat v1 v2" using * e.supp by auto
- show "atom z' \<sharp> \<Gamma>'" using * by auto
- qed
- thus ?case using * ** by metis
-next
- case (infer_e_splitI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v1 z1 c1 v2 z2 z3)
- show ?case proof
- show "\<Theta>; \<B>; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f \<Delta>" using infer_e_splitI wf_weakening by auto
- show "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> " using infer_e_splitI wf_weakening by auto
- show "\<Theta>; \<B>; \<Gamma>' \<turnstile> v1 \<Rightarrow> \<lbrace> z1 : B_bitvec | c1 \<rbrace>" using infer_v_g_weakening infer_e_splitI by metis
- show "\<Theta>; \<B>; \<Gamma>' \<turnstile> v2 \<Leftarrow> \<lbrace> z2 : B_int | [ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ z2 ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e
- AND [ leq [ [ z2 ]\<^sup>v ]\<^sup>c\<^sup>e [| [ v1 ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrace>"
- using check_v_g_weakening infer_e_splitI by metis
- show "atom z1 \<sharp> AE_split v1 v2" using infer_e_splitI by auto
- show "atom z1 \<sharp> \<Gamma>'" using infer_e_splitI by auto
- show "atom z2 \<sharp> AE_split v1 v2" using infer_e_splitI by auto
- show "atom z2 \<sharp> \<Gamma>'" using infer_e_splitI by auto
- show "atom z3 \<sharp> AE_split v1 v2" using infer_e_splitI by auto
- show "atom z3 \<sharp> \<Gamma>'" using infer_e_splitI by auto
- qed
-qed
-
-text \<open> Special cases proved explicitly, other cases at the end with method + \<close>
-
-lemma infer_e_d_weakening:
- fixes e::e
- assumes "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile> e \<Rightarrow> \<tau>" and "setD \<Delta> \<subseteq> setD \<Delta>'" and "wfD \<Theta> \<B> \<Gamma> \<Delta>'"
- shows "\<Theta>; \<Phi> ; \<B> ; \<Gamma> ; \<Delta>' \<turnstile> e \<Rightarrow> \<tau>"
- using assms by(nominal_induct \<tau> avoiding: \<Delta>' rule: infer_e.strong_induct,auto simp add:infer_e.intros)
-
-lemma wfG_x_fresh_in_v_simple:
- fixes x::x and v :: v
- assumes "\<Theta>; \<B>; \<Gamma> \<turnstile> v \<Rightarrow> \<tau>" and "atom x \<sharp> \<Gamma>"
- shows "atom x \<sharp> v"
- using wfV_x_fresh infer_v_wf assms by metis
-
-lemma check_s_g_weakening:
- fixes v::v and s::s and cs::branch_s and x::x and c::c and b::b and \<Gamma>'::\<Gamma> and \<Theta>::\<Theta> and css::branch_list
- shows "check_s \<Theta> \<Phi> \<B> \<Gamma> \<Delta> s t \<Longrightarrow> toSet \<Gamma> \<subseteq> toSet \<Gamma>' \<Longrightarrow> \<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>' \<Longrightarrow> check_s \<Theta> \<Phi> \<B> \<Gamma>' \<Delta> s t" and
- "check_branch_s \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid cons const v cs t \<Longrightarrow> toSet \<Gamma> \<subseteq> toSet \<Gamma>' \<Longrightarrow> \<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>' \<Longrightarrow> check_branch_s \<Theta> \<Phi> \<B> \<Gamma>' \<Delta> tid cons const v cs t" and
- "check_branch_list \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist v css t \<Longrightarrow> toSet \<Gamma> \<subseteq> toSet \<Gamma>' \<Longrightarrow> \<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>' \<Longrightarrow> check_branch_list \<Theta> \<Phi> \<B> \<Gamma>' \<Delta> tid dclist v css t"
-proof(nominal_induct t and t and t avoiding: \<Gamma>' rule: check_s_check_branch_s_check_branch_list.strong_induct)
- case (check_valI \<Theta> \<B> \<Gamma> \<Delta>' \<Phi> v \<tau>' \<tau>)
- then show ?case using Typing.check_valI infer_v_g_weakening wf_weakening subtype_weakening by metis
-next
- case (check_letI x \<Theta> \<Phi> \<B> \<Gamma> \<Delta> e \<tau> z s b c)
- hence xf:"atom x \<sharp> \<Gamma>'" by metis
- show ?case proof
- show "atom x \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>', \<Delta>, e, \<tau>)" using check_letI using fresh_prod4 xf by metis
- show "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>' ; \<Delta> \<turnstile> e \<Rightarrow> \<lbrace> z : b | c \<rbrace>" using infer_e_g_weakening check_letI by metis
- show "atom z \<sharp> (x, \<Theta>, \<Phi>, \<B>, \<Gamma>', \<Delta>, e, \<tau>, s)"
- by(unfold fresh_prodN,auto simp add: check_letI fresh_prodN)
- have "toSet ((x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) \<subseteq> toSet ((x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>')" using check_letI toSet.simps
- by (metis Un_commute Un_empty_right Un_insert_right insert_mono)
- moreover hence "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f ((x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>')" using check_letI wfG_cons_weakening check_s_wf by metis
- ultimately show "\<Theta> ; \<Phi> ; \<B> ; (x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>' ; \<Delta> \<turnstile> s \<Leftarrow> \<tau>" using check_letI by metis
- qed
-next
- case (check_let2I x \<Theta> \<Phi> \<B> G \<Delta> t s1 \<tau> s2)
- show ?case proof
- show "atom x \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>', \<Delta>, t, s1, \<tau>)" using check_let2I using fresh_prod4 by auto
- show "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>' ; \<Delta> \<turnstile> s1 \<Leftarrow> t" using check_let2I by metis
- have "toSet ((x, b_of t, c_of t x) #\<^sub>\<Gamma> G) \<subseteq> toSet ((x, b_of t, c_of t x ) #\<^sub>\<Gamma> \<Gamma>')" using check_let2I by auto
- moreover hence "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f ((x, b_of t, c_of t x) #\<^sub>\<Gamma> \<Gamma>')" using check_let2I wfG_cons_weakening check_s_wf by metis
- ultimately show "\<Theta> ; \<Phi> ; \<B> ; (x, b_of t, c_of t x) #\<^sub>\<Gamma> \<Gamma>' ; \<Delta> \<turnstile> s2 \<Leftarrow> \<tau>" using check_let2I by metis
- qed
-next
- case (check_branch_list_consI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist' v cs \<tau> css dclist)
- thus ?case using Typing.check_branch_list_consI by metis
-next
- case (check_branch_list_finalI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist' v cs \<tau> dclist)
- thus ?case using Typing.check_branch_list_finalI by metis
-next
- case (check_branch_s_branchI \<Theta> \<B> \<Gamma> \<Delta> \<tau> const x \<Phi> tid cons v s)
- show ?case proof
- show "\<Theta>; \<B>; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f \<Delta> " using wf_weakening2(6) check_branch_s_branchI by metis
- show "\<turnstile>\<^sub>w\<^sub>f \<Theta>" using check_branch_s_branchI by auto
- show "\<Theta>; \<B>; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f \<tau> " using check_branch_s_branchI wfT_weakening \<open>wfG \<Theta> \<B> \<Gamma>'\<close> by presburger
-
- show "\<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f const " using check_branch_s_branchI by auto
- show "atom x \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>', \<Delta>, tid, cons, const, v, \<tau>)" using check_branch_s_branchI by auto
- have "toSet ((x, b_of const, CE_val v == CE_val(V_cons tid cons (V_var x)) AND c_of const x) #\<^sub>\<Gamma> \<Gamma>) \<subseteq> toSet ((x, b_of const, CE_val v == CE_val (V_cons tid cons (V_var x)) AND c_of const x) #\<^sub>\<Gamma> \<Gamma>')"
- using check_branch_s_branchI by auto
- moreover hence "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f ((x, b_of const, CE_val v == CE_val (V_cons tid cons (V_var x)) AND c_of const x ) #\<^sub>\<Gamma> \<Gamma>')"
- using check_branch_s_branchI wfG_cons_weakening check_s_wf by metis
- ultimately show "\<Theta> ; \<Phi> ; \<B> ; (x, b_of const, CE_val v == CE_val (V_cons tid cons (V_var x)) AND c_of const x ) #\<^sub>\<Gamma> \<Gamma>' ; \<Delta> \<turnstile> s \<Leftarrow> \<tau>"
- using check_branch_s_branchI using fresh_dom_free by auto
- qed
-next
- case (check_ifI z \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v s1 s2 \<tau>)
- show ?case proof
- show \<open>atom z \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>', \<Delta>, v, s1, s2, \<tau>)\<close> using fresh_prodN check_ifI by auto
- show \<open>\<Theta>; \<B>; \<Gamma>' \<turnstile> v \<Leftarrow> \<lbrace> z : B_bool | TRUE \<rbrace>\<close> using check_v_g_weakening check_ifI by auto
- show \<open> \<Theta> ; \<Phi> ; \<B> ; \<Gamma>' ; \<Delta> \<turnstile> s1 \<Leftarrow> \<lbrace> z : b_of \<tau> | CE_val v == CE_val (V_lit L_true) IMP c_of \<tau> z \<rbrace>\<close> using check_ifI by auto
- show \<open> \<Theta> ; \<Phi> ; \<B> ; \<Gamma>' ; \<Delta> \<turnstile> s2 \<Leftarrow> \<lbrace> z : b_of \<tau> | CE_val v == CE_val (V_lit L_false) IMP c_of \<tau> z \<rbrace>\<close> using check_ifI by auto
- qed
-next
- case (check_whileI \<Delta> G P s1 z s2 \<tau>')
- then show ?case using check_s_check_branch_s_check_branch_list.intros check_v_g_weakening subtype_weakening wf_weakening
- by (meson infer_v_g_weakening)
-next
- case (check_seqI \<Delta> G P s1 z s2 \<tau>)
- then show ?case using check_s_check_branch_s_check_branch_list.intros check_v_g_weakening subtype_weakening wf_weakening
- by (meson infer_v_g_weakening)
-next
- case (check_varI u \<Theta> \<Phi> \<B> \<Gamma> \<Delta> \<tau>' v \<tau> s)
- thus ?case using check_v_g_weakening check_s_check_branch_s_check_branch_list.intros by auto
-next
- case (check_assignI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> u \<tau> v z \<tau>')
- show ?case proof
- show \<open>\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using check_assignI by auto
- show \<open>\<Theta>; \<B>; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f \<Delta>\<close> using check_assignI wf_weakening by auto
- show \<open>(u, \<tau>) \<in> setD \<Delta>\<close> using check_assignI by auto
- show \<open>\<Theta>; \<B>; \<Gamma>' \<turnstile> v \<Leftarrow> \<tau>\<close> using check_assignI check_v_g_weakening by auto
- show \<open>\<Theta>; \<B>; \<Gamma>' \<turnstile> \<lbrace> z : B_unit | TRUE \<rbrace> \<lesssim> \<tau>'\<close> using subtype_weakening check_assignI by auto
- qed
-next
- case (check_caseI \<Delta> \<Gamma> \<Theta> dclist cs \<tau> tid v z)
-
- then show ?case using check_s_check_branch_s_check_branch_list.intros check_v_g_weakening subtype_weakening wf_weakening
- by (meson infer_v_g_weakening)
-next
- case (check_assertI x \<Theta> \<Phi> \<B> \<Gamma> \<Delta> c \<tau> s)
- show ?case proof
- show \<open>atom x \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>', \<Delta>, c, \<tau>, s)\<close> using check_assertI by auto
-
- have " \<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f (x, B_bool, c) #\<^sub>\<Gamma> \<Gamma>" using check_assertI check_s_wf by metis
- hence *:" \<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f (x, B_bool, c) #\<^sub>\<Gamma> \<Gamma>'" using wfG_cons_weakening check_assertI by metis
- moreover have "toSet ((x, B_bool, c) #\<^sub>\<Gamma> \<Gamma>) \<subseteq> toSet ((x, B_bool, c) #\<^sub>\<Gamma> \<Gamma>')" using check_assertI by auto
- thus \<open> \<Theta> ; \<Phi> ; \<B> ; (x, B_bool, c) #\<^sub>\<Gamma> \<Gamma>' ; \<Delta> \<turnstile> s \<Leftarrow> \<tau>\<close> using check_assertI(11) [OF _ *] by auto
-
- show \<open>\<Theta>; \<B>; \<Gamma>' \<Turnstile> c \<close> using check_assertI valid_weakening by metis
- show \<open> \<Theta>; \<B>; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using check_assertI wf_weakening by metis
- qed
-qed
-
-lemma wfG_xa_fresh_in_v:
- fixes c::c and \<Gamma>::\<Gamma> and G::\<Gamma> and v::v and xa::x
- assumes "\<Theta>; \<B>; \<Gamma> \<turnstile> v \<Rightarrow> \<tau>" and "G=( \<Gamma>'@ (x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)" and "atom xa \<sharp> G" and "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f G"
- shows "atom xa \<sharp> v"
-proof -
- have "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b_of \<tau>" using infer_v_wf assms by metis
- hence "supp v \<subseteq> atom_dom \<Gamma> \<union> supp \<B>" using wfV_supp by simp
- moreover have "atom xa \<notin> atom_dom G"
- using assms wfG_atoms_supp_eq[OF assms(4)] fresh_def by metis
- ultimately show ?thesis using fresh_def
- using assms infer_v_wf wfG_atoms_supp_eq
- fresh_GCons fresh_append_g subsetCE
- by (metis wfG_x_fresh_in_v_simple)
-qed
-
-lemma fresh_z_subst_g:
- fixes G::\<Gamma>
- assumes "atom z' \<sharp> (x,v)" and \<open>atom z' \<sharp> G\<close>
- shows "atom z' \<sharp> G[x::=v]\<^sub>\<Gamma>\<^sub>v"
-proof -
- have "atom z' \<sharp> v" using assms fresh_prod2 by auto
- thus ?thesis using fresh_subst_gv assms by metis
-qed
-
-lemma wfG_xa_fresh_in_subst_v:
- fixes c::c and v::v and x::x and \<Gamma>::\<Gamma> and G::\<Gamma> and xa::x
- assumes "\<Theta>; \<B>; \<Gamma> \<turnstile> v \<Rightarrow> \<tau>" and "G=( \<Gamma>'@ (x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)" and "atom xa \<sharp> G" and "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f G"
- shows "atom xa \<sharp> (subst_gv G x v)"
-proof -
- have "atom xa \<sharp> v" using wfG_xa_fresh_in_v assms by metis
- thus ?thesis using fresh_subst_gv assms by metis
-qed
-
-subsection \<open>Weakening Immutable Variable Context\<close>
-
-declare check_s_check_branch_s_check_branch_list.intros[simp]
-declare check_s_check_branch_s_check_branch_list.intros[intro]
-
-lemma check_s_d_weakening:
- fixes s::s and v::v and cs::branch_s and css::branch_list
- shows " \<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile> s \<Leftarrow> \<tau> \<Longrightarrow> setD \<Delta> \<subseteq> setD \<Delta>' \<Longrightarrow> wfD \<Theta> \<B> \<Gamma> \<Delta>' \<Longrightarrow> \<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta>' \<turnstile> s \<Leftarrow> \<tau>" and
- "check_branch_s \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid cons const v cs \<tau> \<Longrightarrow> setD \<Delta> \<subseteq> setD \<Delta>' \<Longrightarrow> wfD \<Theta> \<B> \<Gamma> \<Delta>' \<Longrightarrow> check_branch_s \<Theta> \<Phi> \<B> \<Gamma> \<Delta>' tid cons const v cs \<tau>" and
- "check_branch_list \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist v css \<tau> \<Longrightarrow> setD \<Delta> \<subseteq> setD \<Delta>' \<Longrightarrow> wfD \<Theta> \<B> \<Gamma> \<Delta>' \<Longrightarrow> check_branch_list \<Theta> \<Phi> \<B> \<Gamma> \<Delta>' tid dclist v css \<tau>"
-proof(nominal_induct \<tau> and \<tau> and \<tau> avoiding: \<Delta>' arbitrary: v rule: check_s_check_branch_s_check_branch_list.strong_induct)
- case (check_valI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v \<tau>' \<tau>)
- then show ?case using check_s_check_branch_s_check_branch_list.intros by blast
-next
- case (check_letI x \<Theta> \<Phi> \<B> \<Gamma> \<Delta> e \<tau> z s b c)
- show ?case proof
- show "atom x \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>, \<Delta>', e, \<tau>)" using check_letI by auto
- show "atom z \<sharp> (x, \<Theta>, \<Phi>, \<B>, \<Gamma>, \<Delta>', e, \<tau>, s)" using check_letI by auto
- show "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta>' \<turnstile> e \<Rightarrow> \<lbrace> z : b | c \<rbrace>" using check_letI infer_e_d_weakening by auto
- have "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f (x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>" using check_letI check_s_wf by metis
- moreover have "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>'" using check_letI check_s_wf by metis
- ultimately have "\<Theta>; \<B>; (x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>'" using wf_weakening2(6) toSet.simps by fast
- thus "\<Theta> ; \<Phi> ; \<B> ; (x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma> ; \<Delta>' \<turnstile> s \<Leftarrow> \<tau>" using check_letI by simp
- qed
-next
- case (check_branch_s_branchI \<Theta> \<B> \<Gamma> \<Delta> \<tau> const x \<Phi> tid cons v s)
- moreover have "\<Theta> ;\<B> \<turnstile>\<^sub>w\<^sub>f (x, b_of const, CE_val v == CE_val (V_cons tid cons (V_var x)) AND c_of const x ) #\<^sub>\<Gamma> \<Gamma>"
- using check_s_wf[OF check_branch_s_branchI(16) ] by metis
- moreover hence " \<Theta> ;\<B> ; (x, b_of const, CE_val v == CE_val (V_cons tid cons (V_var x)) AND c_of const x ) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>'"
- using wf_weakening2(6) check_branch_s_branchI by fastforce
- ultimately show ?case
- using check_s_check_branch_s_check_branch_list.intros by simp
-next
- case (check_branch_list_consI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist v cs \<tau> css)
- then show ?case using check_s_check_branch_s_check_branch_list.intros by meson
-next
- case (check_branch_list_finalI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist v cs \<tau>)
- then show ?case using check_s_check_branch_s_check_branch_list.intros by meson
-next
- case (check_ifI z \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v s1 s2 \<tau>)
- show ?case proof
- show \<open>atom z \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>, \<Delta>', v, s1, s2, \<tau>)\<close> using fresh_prodN check_ifI by auto
- show \<open>\<Theta>; \<B>; \<Gamma> \<turnstile> v \<Leftarrow> \<lbrace> z : B_bool | TRUE \<rbrace>\<close> using check_ifI by auto
- show \<open> \<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta>' \<turnstile> s1 \<Leftarrow> \<lbrace> z : b_of \<tau> | CE_val v == CE_val (V_lit L_true) IMP c_of \<tau> z \<rbrace>\<close> using check_ifI by auto
- show \<open> \<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta>' \<turnstile> s2 \<Leftarrow> \<lbrace> z : b_of \<tau> | CE_val v == CE_val (V_lit L_false) IMP c_of \<tau> z \<rbrace>\<close> using check_ifI by auto
- qed
-next
- case (check_assertI x \<Theta> \<Phi> \<B> \<Gamma> \<Delta> c \<tau> s)
- show ?case proof
- show "atom x \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>, \<Delta>', c, \<tau>,s)" using fresh_prodN check_assertI by auto
- show *:" \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>' " using check_assertI by auto
- hence "\<Theta>; \<B>; (x, B_bool, c) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>' " using wf_weakening2(6)[OF *, of " (x, B_bool, c) #\<^sub>\<Gamma> \<Gamma>"] check_assertI check_s_wf toSet.simps by auto
- thus "\<Theta> ; \<Phi> ; \<B> ; (x, B_bool, c) #\<^sub>\<Gamma> \<Gamma> ; \<Delta>' \<turnstile> s \<Leftarrow> \<tau>"
- using check_assertI(11)[OF \<open>setD \<Delta> \<subseteq> setD \<Delta>'\<close>] by simp
- (* wf_weakening2(6) check_s_wf check_assertI *)
- show "\<Theta>; \<B>; \<Gamma> \<Turnstile> c " using fresh_prodN check_assertI by auto
-
- qed
-next
- case (check_let2I x \<Theta> \<Phi> \<B> G \<Delta> t s1 \<tau> s2)
- show ?case proof
- show "atom x \<sharp> (\<Theta>, \<Phi>, \<B>, G, \<Delta>', t , s1, \<tau>)" using check_let2I by auto
-
- show "\<Theta> ; \<Phi> ; \<B> ; G ; \<Delta>' \<turnstile> s1 \<Leftarrow> t" using check_let2I infer_e_d_weakening by auto
- have "\<Theta>; \<B>; (x, b_of t, c_of t x ) #\<^sub>\<Gamma> G \<turnstile>\<^sub>w\<^sub>f \<Delta>'" using check_let2I wf_weakening2(6) check_s_wf by fastforce
- thus "\<Theta> ; \<Phi> ; \<B> ; (x, b_of t, c_of t x) #\<^sub>\<Gamma> G ; \<Delta>' \<turnstile> s2 \<Leftarrow> \<tau>" using check_let2I by simp
- qed
-next
- case (check_varI u \<Theta> \<Phi> \<B> \<Gamma> \<Delta> \<tau>' v \<tau> s)
- show ?case proof
- show "atom u \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>, \<Delta>', \<tau>', v, \<tau>)" using check_varI by auto
- show "\<Theta>; \<B>; \<Gamma> \<turnstile> v \<Leftarrow> \<tau>'" using check_varI by auto
- have "setD ((u, \<tau>') #\<^sub>\<Delta>\<Delta>) \<subseteq> setD ((u, \<tau>') #\<^sub>\<Delta>\<Delta>')" using setD.simps check_varI by auto
- moreover have "u \<notin> fst ` setD \<Delta>'" using check_varI(1) setD.simps fresh_DCons by (simp add: fresh_d_not_in)
- moreover hence "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f (u, \<tau>') #\<^sub>\<Delta>\<Delta>' " using wfD_cons fresh_DCons setD.simps check_varI check_v_wf by metis
- ultimately show "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; (u, \<tau>') #\<^sub>\<Delta>\<Delta>' \<turnstile> s \<Leftarrow> \<tau>" using check_varI by auto
- qed
-next
- case (check_assignI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> u \<tau> v z \<tau>')
- moreover hence "(u, \<tau>) \<in> setD \<Delta>'" by auto
- ultimately show ?case using Typing.check_assignI by simp
-next
- case (check_whileI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> s1 z s2 \<tau>')
- then show ?case using check_s_check_branch_s_check_branch_list.intros by meson
-next
- case (check_seqI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> s1 z s2 \<tau>)
- then show ?case using check_s_check_branch_s_check_branch_list.intros by meson
-next
- case (check_caseI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist v cs \<tau> z)
- then show ?case using check_s_check_branch_s_check_branch_list.intros by meson
-
-qed
-
-lemma valid_ce_eq:
- fixes v::v and ce2::ce
- assumes "ce1 = ce2[x::=v]\<^sub>c\<^sub>e\<^sub>v" and "wfV \<Theta> \<B> GNil v b" and "wfCE \<Theta> \<B> ((x, b, TRUE) #\<^sub>\<Gamma> GNil) ce2 b'" and "wfCE \<Theta> \<B> GNil ce1 b'"
- shows \<open>\<Theta>; \<B>; (x, b, ([[x ]\<^sup>v]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e )) #\<^sub>\<Gamma> GNil \<Turnstile> ce1 == ce2 \<close>
- unfolding valid.simps proof
- have wfg: "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f (x, b, [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e ) #\<^sub>\<Gamma> GNil"
- using wfG_cons1I wfG_nilI wfX_wfY assms wf_intros
- by (meson fresh_GNil wfC_e_eq wfG_intros2)
-
- show wf: \<open>\<Theta>; \<B>; (x, b, [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e ) #\<^sub>\<Gamma> GNil \<turnstile>\<^sub>w\<^sub>f ce1 == ce2 \<close>
- apply(rule wfC_eqI[where b=b'])
- using wfg toSet.simps assms wfCE_weakening apply simp
-
- using wfg assms wf_replace_inside1(8) assms
- using wfC_trueI wf_trans(8) by auto
-
- show \<open>\<forall>i. ((\<Theta> ; (x, b, [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e ) #\<^sub>\<Gamma> GNil \<turnstile> i) \<and> (i \<Turnstile> (x, b, [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e ) #\<^sub>\<Gamma> GNil) \<longrightarrow>
- (i \<Turnstile> (ce1 == ce2 )))\<close> proof(rule+,goal_cases)
- fix i
- assume as:"\<Theta> ; (x, b, [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e ) #\<^sub>\<Gamma> GNil \<turnstile> i \<and> i \<Turnstile> (x, b, [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e ) #\<^sub>\<Gamma> GNil"
- have 1:"wfV \<Theta> \<B> ((x, b, [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e ) #\<^sub>\<Gamma> GNil) v b"
- using wf_weakening assms append_g.simps toSet.simps wf wfX_wfY
- by (metis empty_subsetI)
- hence "\<exists>s. i \<lbrakk> v \<rbrakk> ~ s" using eval_v_exist[OF _ 1] as by auto
- then obtain s where iv:"i\<lbrakk> v \<rbrakk> ~ s" ..
-
- hence ix:"i x = Some s" proof -
- have "i \<Turnstile> [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e" using is_satis_g.simps as by auto
- hence "i \<lbrakk> [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e \<rbrakk> ~ True" using is_satis.simps by auto
- hence "i \<lbrakk> [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrakk> ~ s" using
- iv eval_e_elims
- by (metis eval_c_elims(7) eval_e_uniqueness eval_e_valI)
- thus ?thesis using eval_v_elims(2) eval_e_elims(1) by metis
- qed
-
- have 1:"wfCE \<Theta> \<B> ((x, b, [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e ) #\<^sub>\<Gamma> GNil) ce1 b'"
- using wfCE_weakening assms append_g.simps toSet.simps wf wfX_wfY
- by (metis empty_subsetI)
- hence "\<exists>s1. i \<lbrakk> ce1 \<rbrakk> ~ s1" using eval_e_exist assms as by auto
- then obtain s1 where s1: "i\<lbrakk>ce1\<rbrakk> ~ s1" ..
-
- moreover have "i\<lbrakk> ce2 \<rbrakk> ~ s1" proof -
- have "i\<lbrakk> ce2[x::=v]\<^sub>c\<^sub>e\<^sub>v \<rbrakk> ~ s1" using assms s1 by auto
- moreover have "ce1 = ce2[x::=v]\<^sub>c\<^sub>e\<^sub>v" using subst_v_ce_def assms subst_v_simple_commute by auto
- ultimately have "i(x \<mapsto> s) \<lbrakk> ce2 \<rbrakk> ~ s1"
- using ix subst_e_eval_v[of i ce1 s1 "ce2[z::=[ x ]\<^sup>v]\<^sub>v" x v s] iv s1 by auto
- moreover have "i(x \<mapsto> s) = i" using ix by auto
- ultimately show ?thesis by auto
- qed
- ultimately show "i \<lbrakk> ce1 == ce2 \<rbrakk> ~ True " using eval_c_eqI by metis
- qed
-qed
-
-
-lemma check_v_top:
- fixes v::v
- assumes "\<Theta>; \<B>; GNil \<turnstile> v \<Leftarrow> \<tau>" and "ce1 = ce2[z::=v]\<^sub>c\<^sub>e\<^sub>v" and "\<Theta>; \<B>; GNil \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b_of \<tau> | ce1 == ce2 \<rbrace>"
- and "supp ce1 \<subseteq> supp \<B>"
- shows "\<Theta>; \<B>; GNil \<turnstile> v \<Leftarrow> \<lbrace> z : b_of \<tau> | ce1 == ce2 \<rbrace>"
-proof -
- obtain t where t: "\<Theta>; \<B>; GNil \<turnstile> v \<Rightarrow> t \<and> \<Theta>; \<B>; GNil \<turnstile> t \<lesssim> \<tau>"
- using assms check_v_elims by metis
-
- then obtain z' and b' where *:"t = \<lbrace> z' : b' | [ [ z' ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e \<rbrace> \<and> atom z' \<sharp> v \<and> atom z' \<sharp> (\<Theta>, \<B>,GNil)"
- using assms infer_v_form by metis
- have beq: "b_of t = b_of \<tau>" using subtype_eq_base2 b_of.simps t by auto
- obtain x::x where xf: \<open>atom x \<sharp> (\<Theta>, \<B>, GNil, z', [ [ z' ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e , z, ce1 == ce2 )\<close>
- using obtain_fresh by metis
-
- have "\<Theta>; \<B>; (x, b_of \<tau>, TRUE) #\<^sub>\<Gamma> GNil \<turnstile>\<^sub>w\<^sub>f (ce1[z::=[ x ]\<^sup>v]\<^sub>v == ce2[z::=[ x ]\<^sup>v]\<^sub>v )"
- using wfT_wfC2[OF assms(3), of x] subst_cv.simps(6) subst_v_c_def subst_v_ce_def fresh_GNil by simp
-
- then obtain b2 where b2: "\<Theta>; \<B>; (x, b_of t, TRUE) #\<^sub>\<Gamma> GNil \<turnstile>\<^sub>w\<^sub>f ce1[z::=[ x ]\<^sup>v]\<^sub>v : b2 \<and>
- \<Theta>; \<B>; (x, b_of t, TRUE) #\<^sub>\<Gamma> GNil \<turnstile>\<^sub>w\<^sub>f ce2[z::=[ x ]\<^sup>v]\<^sub>v : b2" using wfC_elims(3)
- beq by metis
-
- from xf have "\<Theta>; \<B>; GNil \<turnstile> \<lbrace> z' : b_of t | [ [ z' ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e \<rbrace> \<lesssim> \<lbrace> z : b_of t | ce1 == ce2 \<rbrace>"
- proof
- show \<open> \<Theta>; \<B>; GNil \<turnstile>\<^sub>w\<^sub>f \<lbrace> z' : b_of t | [ [ z' ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e \<rbrace> \<close> using b_of.simps assms infer_v_wf t * by auto
- show \<open> \<Theta>; \<B>; GNil \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b_of t | ce1 == ce2 \<rbrace> \<close> using beq assms by auto
- have \<open>\<Theta>; \<B>; (x, b_of t, ([ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e )) #\<^sub>\<Gamma> GNil \<Turnstile> (ce1[z::=[ x ]\<^sup>v]\<^sub>v == ce2[z::=[ x ]\<^sup>v]\<^sub>v ) \<close>
- proof(rule valid_ce_eq)
- show \<open>ce1[z::=[ x ]\<^sup>v]\<^sub>v = ce2[z::=[ x ]\<^sup>v]\<^sub>v[x::=v]\<^sub>c\<^sub>e\<^sub>v\<close> proof -
- have "atom z \<sharp> ce1" using assms fresh_def x_not_in_b_set by fast
- hence "ce1[z::=[ x ]\<^sup>v]\<^sub>v = ce1"
- using forget_subst_v by auto
- also have "... = ce2[z::=v]\<^sub>c\<^sub>e\<^sub>v" using assms by auto
- also have "... = ce2[z::=[ x ]\<^sup>v]\<^sub>v[x::=v]\<^sub>c\<^sub>e\<^sub>v" proof -
- have "atom x \<sharp> ce2" using xf fresh_prodN c.fresh by metis
- thus ?thesis using subst_v_simple_commute subst_v_ce_def by simp
- qed
- finally show ?thesis by auto
- qed
- show \<open> \<Theta>; \<B>; GNil \<turnstile>\<^sub>w\<^sub>f v : b_of t \<close> using infer_v_wf t by simp
- show \<open> \<Theta>; \<B>; (x, b_of t, TRUE) #\<^sub>\<Gamma> GNil \<turnstile>\<^sub>w\<^sub>f ce2[z::=[ x ]\<^sup>v]\<^sub>v : b2 \<close> using b2 by auto
-
- have " \<Theta>; \<B>; (x, b_of t, TRUE) #\<^sub>\<Gamma> GNil \<turnstile>\<^sub>w\<^sub>f ce1[z::=[ x ]\<^sup>v]\<^sub>v : b2" using b2 by auto
- moreover have "atom x \<sharp> ce1[z::=[ x ]\<^sup>v]\<^sub>v"
- using fresh_subst_v_if assms fresh_def
- using \<open>\<Theta>; \<B>; GNil \<turnstile>\<^sub>w\<^sub>f v : b_of t\<close> \<open>ce1[z::=[ x ]\<^sup>v]\<^sub>v = ce2[z::=[ x ]\<^sup>v]\<^sub>v[x::=v]\<^sub>c\<^sub>e\<^sub>v\<close>
- fresh_GNil subst_v_ce_def wfV_x_fresh by auto
- ultimately show \<open> \<Theta>; \<B>; GNil \<turnstile>\<^sub>w\<^sub>f ce1[z::=[ x ]\<^sup>v]\<^sub>v : b2 \<close> using
- wf_restrict(8) by force
- qed
- moreover have v: "v[z'::=[ x ]\<^sup>v]\<^sub>v\<^sub>v = v"
- using forget_subst assms infer_v_wf wfV_supp x_not_in_b_set
- by (simp add: "local.*")
- ultimately show "\<Theta>; \<B>; (x, b_of t, ([ [ z' ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e )[z'::=[ x ]\<^sup>v]\<^sub>v) #\<^sub>\<Gamma> GNil \<Turnstile> (ce1 == ce2 )[z::=[ x ]\<^sup>v]\<^sub>v"
- unfolding subst_cv.simps subst_v_c_def subst_cev.simps subst_vv.simps
- using subst_v_ce_def by simp
- qed
- thus ?thesis using b_of.simps assms * check_v_subtypeI t b_of.simps subtype_eq_base2 by metis
-qed
-
-end
-
+(*<*)
+theory TypingL
+ imports Typing RCLogicL "HOL-Eisbach.Eisbach"
+begin
+ (*>*)
+
+chapter \<open>Typing Lemmas\<close>
+
+section \<open>Prelude\<close>
+
+text \<open>Needed as the typing elimination rules give us facts for an alpha-equivalent version of a term
+ and so need to be able to 'jump back' to a typing judgement for the orginal term\<close>
+
+lemma \<tau>_fresh_c[simp]:
+ assumes "atom x \<sharp> \<lbrace> z : b | c \<rbrace>" and "atom z \<sharp> x"
+ shows "atom x \<sharp> c"
+ using \<tau>.fresh assms fresh_at_base
+ by (simp add: fresh_at_base(2))
+
+lemmas subst_defs = subst_b_b_def subst_b_c_def subst_b_\<tau>_def subst_v_v_def subst_v_c_def subst_v_\<tau>_def
+
+lemma wfT_wfT_if1:
+ assumes "wfT \<Theta> \<B> \<Gamma> (\<lbrace> z : b_of t | CE_val v == CE_val (V_lit L_false) IMP c_of t z \<rbrace>)" and "atom z \<sharp> (\<Gamma>,t)"
+ shows "wfT \<Theta> \<B> \<Gamma> t"
+ using assms proof(nominal_induct t avoiding: \<Gamma> z rule: \<tau>.strong_induct)
+ case (T_refined_type z' b' c')
+ show ?case proof(rule wfT_wfT_if)
+ show \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b' | [ v ]\<^sup>c\<^sup>e == [ [ L_false ]\<^sup>v ]\<^sup>c\<^sup>e IMP c'[z'::=[ z]\<^sup>v]\<^sub>c\<^sub>v \<rbrace> \<close>
+ using T_refined_type b_of.simps c_of.simps subst_defs by metis
+ show \<open>atom z \<sharp> (c', \<Gamma>)\<close> using T_refined_type fresh_prodN \<tau>_fresh_c by metis
+ qed
+qed
+
+lemma fresh_u_replace_true:
+ fixes bv::bv and \<Gamma>::\<Gamma>
+ assumes "atom bv \<sharp> \<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>"
+ shows "atom bv \<sharp> \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>"
+ using fresh_append_g fresh_GCons assms fresh_Pair c.fresh(1) by auto
+
+lemma wf_replace_true1:
+ fixes \<Gamma>::\<Gamma> and \<Phi>::\<Phi> and \<Theta>::\<Theta> and \<Gamma>'::\<Gamma> and v::v and e::e and c::c and c''::c and c'::c and \<tau>::\<tau> and ts::"(string*\<tau>) list" and \<Delta>::\<Delta> and b'::b and b::b and s::s
+ and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def and cs::branch_s and css::branch_list
+
+shows "\<Theta>; \<B>; G \<turnstile>\<^sub>w\<^sub>f v : b' \<Longrightarrow> G = \<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma> \<Longrightarrow> \<Theta> ; \<B> ; \<Gamma>' @ ((x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>) \<turnstile>\<^sub>w\<^sub>f v : b'" and
+ "\<Theta>; \<B>; G \<turnstile>\<^sub>w\<^sub>f c'' \<Longrightarrow> G = \<Gamma>' @(x, b, c) #\<^sub>\<Gamma> \<Gamma> \<Longrightarrow> \<Theta> ; \<B> ; \<Gamma>' @ ((x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>) \<turnstile>\<^sub>w\<^sub>f c''" and
+ "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f G \<Longrightarrow> G = \<Gamma>' @(x, b, c) #\<^sub>\<Gamma> \<Gamma> \<Longrightarrow> \<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>' @ ((x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>) " and
+ "\<Theta>; \<B>; G \<turnstile>\<^sub>w\<^sub>f \<tau> \<Longrightarrow> G = \<Gamma>' @(x, b, c) #\<^sub>\<Gamma> \<Gamma> \<Longrightarrow> \<Theta> ; \<B> ; \<Gamma>' @ ((x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>) \<turnstile>\<^sub>w\<^sub>f \<tau>" and
+ "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ts \<Longrightarrow> True" and
+ "\<turnstile>\<^sub>w\<^sub>f P \<Longrightarrow> True" and
+ "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f b \<Longrightarrow> True" and
+ "\<Theta> ; \<B> ; G \<turnstile>\<^sub>w\<^sub>f ce : b' \<Longrightarrow> G = \<Gamma>' @(x, b, c) #\<^sub>\<Gamma> \<Gamma> \<Longrightarrow> \<Theta> ; \<B> ; \<Gamma>' @ ((x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>) \<turnstile>\<^sub>w\<^sub>f ce : b'" and
+ "\<Theta> \<turnstile>\<^sub>w\<^sub>f td \<Longrightarrow> True"
+proof(nominal_induct
+ b' and c'' and G and \<tau> and ts and P and b and b' and td
+ arbitrary: \<Gamma> \<Gamma>' and \<Gamma> \<Gamma>' and \<Gamma> \<Gamma>' and \<Gamma> \<Gamma>' and \<Gamma> \<Gamma>' and \<Gamma> \<Gamma>' and \<Gamma> \<Gamma>' and \<Gamma> \<Gamma>' and \<Gamma> \<Gamma>' and \<Gamma> \<Gamma>' and \<Gamma> \<Gamma>' and \<Gamma> \<Gamma>' and \<Gamma> \<Gamma>' and \<Gamma> \<Gamma>' and \<Gamma> \<Gamma>' and \<Gamma> \<Gamma>' and \<Gamma> \<Gamma>'
+ rule:wfV_wfC_wfG_wfT_wfTs_wfTh_wfB_wfCE_wfTD.strong_induct)
+ case (wfB_intI \<Theta> \<B>)
+ then show ?case using wf_intros by metis
+next
+ case (wfB_boolI \<Theta> \<B>)
+ then show ?case using wf_intros by metis
+next
+ case (wfB_unitI \<Theta> \<B>)
+ then show ?case using wf_intros by metis
+next
+ case (wfB_bitvecI \<Theta> \<B>)
+ then show ?case using wf_intros by metis
+next
+ case (wfB_pairI \<Theta> \<B> b1 b2)
+ then show ?case using wf_intros by metis
+next
+ case (wfB_consI \<Theta> s dclist \<B>)
+ then show ?case using wf_intros by metis
+next
+ case (wfB_appI \<Theta> b s bv dclist \<B>)
+ then show ?case using wf_intros by metis
+next
+ case (wfV_varI \<Theta> \<B> \<Gamma>'' b' c x')
+ hence wfg: \<open> \<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<close> by auto
+ show ?case proof(cases "x=x'")
+ case True
+ hence "Some (b, TRUE) = lookup (\<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>) x'" using lookup.simps lookup_inside_wf wfg by simp
+ thus ?thesis using Wellformed.wfV_varI[OF wfg]
+ by (metis True lookup_inside_wf old.prod.inject option.inject wfV_varI.hyps(1) wfV_varI.hyps(3) wfV_varI.prems)
+ next
+ case False
+ hence "Some (b', c) = lookup (\<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>) x'" using lookup_inside2 wfV_varI by metis
+ then show ?thesis using Wellformed.wfV_varI[OF wfg]
+ by (metis wfG_elim2 wfG_suffix wfV_varI.hyps(1) wfV_varI.hyps(2) wfV_varI.hyps(3)
+ wfV_varI.prems Wellformed.wfV_varI wf_replace_inside(1))
+ qed
+next
+ case (wfV_litI \<Theta> \<B> \<Gamma> l)
+ then show ?case using wf_intros using wf_intros by metis
+next
+ case (wfV_pairI \<Theta> \<B> \<Gamma> v1 b1 v2 b2)
+ then show ?case using wf_intros by metis
+next
+ case (wfV_consI s dclist \<Theta> dc x b' c \<B> \<Gamma> v)
+ then show ?case using wf_intros by metis
+next
+ case (wfV_conspI s bv dclist \<Theta> dc xc bc cc \<B> b' \<Gamma>'' v)
+ show ?case proof
+ show \<open>AF_typedef_poly s bv dclist \<in> set \<Theta>\<close> using wfV_conspI by metis
+ show \<open>(dc, \<lbrace> xc : bc | cc \<rbrace>) \<in> set dclist\<close> using wfV_conspI by metis
+ show \<open> \<Theta> ;\<B> \<turnstile>\<^sub>w\<^sub>f b' \<close> using wfV_conspI by metis
+ show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : bc[bv::=b']\<^sub>b\<^sub>b \<close> using wfV_conspI by metis
+ have "atom bv \<sharp> \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>" using fresh_u_replace_true wfV_conspI by metis
+ thus \<open>atom bv \<sharp> (\<Theta>, \<B>, \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>, b', v)\<close> using wfV_conspI fresh_prodN by metis
+ qed
+next
+ case (wfCE_valI \<Theta> \<B> \<Gamma> v b)
+ then show ?case using wf_intros by metis
+next
+ case (wfCE_plusI \<Theta> \<B> \<Gamma> v1 v2)
+ then show ?case using wf_intros by metis
+next
+ case (wfCE_leqI \<Theta> \<B> \<Gamma> v1 v2)
+ then show ?case using wf_intros by metis
+next
+ case (wfCE_eqI \<Theta> \<B> \<Gamma> v1 v2)
+ then show ?case using wf_intros by metis
+next
+ case (wfCE_fstI \<Theta> \<B> \<Gamma> v1 b1 b2)
+ then show ?case using wf_intros by metis
+next
+ case (wfCE_sndI \<Theta> \<B> \<Gamma> v1 b1 b2)
+ then show ?case using wf_intros by metis
+next
+ case (wfCE_concatI \<Theta> \<B> \<Gamma> v1 v2)
+ then show ?case using wf_intros by metis
+next
+ case (wfCE_lenI \<Theta> \<B> \<Gamma> v1)
+ then show ?case using wf_intros by metis
+next
+ case (wfTI z \<Theta> \<B> \<Gamma>'' b' c')
+ show ?case proof
+ show \<open>atom z \<sharp> (\<Theta>, \<B>, \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>)\<close> using wfTI fresh_append_g fresh_GCons fresh_prodN by auto
+ show \<open> \<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f b' \<close> using wfTI by metis
+ show \<open> \<Theta>; \<B>; (z, b', TRUE) #\<^sub>\<Gamma> \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f c' \<close> using wfTI append_g.simps by metis
+ qed
+next
+ case (wfC_eqI \<Theta> \<B> \<Gamma> e1 b e2)
+ then show ?case using wf_intros by metis
+next
+ case (wfC_trueI \<Theta> \<B> \<Gamma>)
+ then show ?case using wf_intros by metis
+next
+ case (wfC_falseI \<Theta> \<B> \<Gamma>)
+ then show ?case using wf_intros by metis
+next
+ case (wfC_conjI \<Theta> \<B> \<Gamma> c1 c2)
+ then show ?case using wf_intros by metis
+next
+ case (wfC_disjI \<Theta> \<B> \<Gamma> c1 c2)
+ then show ?case using wf_intros by metis
+next
+ case (wfC_notI \<Theta> \<B> \<Gamma> c1)
+ then show ?case using wf_intros by metis
+next
+ case (wfC_impI \<Theta> \<B> \<Gamma> c1 c2)
+ then show ?case using wf_intros by metis
+next
+ case (wfG_nilI \<Theta> \<B>)
+ then show ?case using GNil_append by blast
+next
+ case (wfG_cons1I c \<Theta> \<B> \<Gamma>'' x b)
+ then show ?case using wf_intros wfG_cons_TRUE2 wfG_elims(2) wfG_replace_inside wfG_suffix
+ by (metis (no_types, lifting))
+next
+ case (wfG_cons2I c \<Theta> \<B> \<Gamma>'' x' b)
+ then show ?case using wf_intros
+ by (metis wfG_cons_TRUE2 wfG_elims(2) wfG_replace_inside wfG_suffix)
+next
+ case wfTh_emptyI
+ then show ?case using wf_intros by metis
+next
+ case (wfTh_consI tdef \<Theta>)
+ then show ?case using wf_intros by metis
+next
+ case (wfTD_simpleI \<Theta> lst s)
+ then show ?case using wf_intros by metis
+next
+ case (wfTD_poly \<Theta> bv lst s)
+ then show ?case using wf_intros by metis
+next
+ case (wfTs_nil \<Theta> \<B> \<Gamma>)
+ then show ?case using wf_intros by metis
+next
+ case (wfTs_cons \<Theta> \<B> \<Gamma> \<tau> dc ts)
+ then show ?case using wf_intros by metis
+qed
+
+lemma wf_replace_true2:
+ fixes \<Gamma>::\<Gamma> and \<Phi>::\<Phi> and \<Theta>::\<Theta> and \<Gamma>'::\<Gamma> and v::v and e::e and c::c and c''::c and c'::c and \<tau>::\<tau> and ts::"(string*\<tau>) list" and \<Delta>::\<Delta> and b'::b and b::b and s::s
+ and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def and cs::branch_s and css::branch_list
+
+shows "\<Theta> ; \<Phi> ; \<B> ; G ; D \<turnstile>\<^sub>w\<^sub>f e : b' \<Longrightarrow> G = \<Gamma>' @(x, b, c) #\<^sub>\<Gamma> \<Gamma> \<Longrightarrow> \<Theta> ; \<Phi> ; \<B> ; \<Gamma>' @ ((x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>); D \<turnstile>\<^sub>w\<^sub>f e : b'" and
+ "\<Theta> ; \<Phi> ; \<B> ; G ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b' \<Longrightarrow> G = \<Gamma>' @(x, b, c) #\<^sub>\<Gamma> \<Gamma> \<Longrightarrow> \<Theta> ; \<Phi> ; \<B> ; \<Gamma>' @ ((x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>) ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b'" and
+ "\<Theta> ; \<Phi> ; \<B> ; G ; \<Delta> ; tid ; dc ; t \<turnstile>\<^sub>w\<^sub>f cs : b' \<Longrightarrow> G = \<Gamma>' @(x, b, c) #\<^sub>\<Gamma> \<Gamma> \<Longrightarrow> \<Theta> ; \<Phi> ; \<B> ; \<Gamma>' @ ((x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>) ; \<Delta> ; tid ; dc ; t \<turnstile>\<^sub>w\<^sub>f cs : b'" and
+ "\<Theta> ; \<Phi> ; \<B> ; G ; \<Delta> ; tid ; dclist \<turnstile>\<^sub>w\<^sub>f css : b' \<Longrightarrow> G = \<Gamma>' @(x, b, c) #\<^sub>\<Gamma> \<Gamma> \<Longrightarrow> \<Theta> ; \<Phi> ; \<B> ; \<Gamma>' @ ((x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>) ; \<Delta> ; tid ; dclist \<turnstile>\<^sub>w\<^sub>f css : b'" and
+
+"\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<Longrightarrow> True" and
+"\<Theta>; \<B>; G \<turnstile>\<^sub>w\<^sub>f \<Delta> \<Longrightarrow> G = \<Gamma>' @(x, b, c) #\<^sub>\<Gamma> \<Gamma> \<Longrightarrow> \<Theta> ; \<B> ; \<Gamma>' @ ((x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>) \<turnstile>\<^sub>w\<^sub>f \<Delta>" and
+
+"\<Theta> ; \<Phi> \<turnstile>\<^sub>w\<^sub>f ftq \<Longrightarrow> True" and
+"\<Theta> ; \<Phi> ; \<B> \<turnstile>\<^sub>w\<^sub>f ft \<Longrightarrow> True"
+proof(nominal_induct
+ b' and b' and b' and b' and \<Phi> and \<Delta> and ftq and ft
+ arbitrary: \<Gamma> \<Gamma>' and \<Gamma> \<Gamma>' and \<Gamma> \<Gamma>' and \<Gamma> \<Gamma>' and \<Gamma> \<Gamma>' and \<Gamma> \<Gamma>' and \<Gamma> \<Gamma>' and \<Gamma> \<Gamma>' and \<Gamma> \<Gamma>' and \<Gamma> \<Gamma>' and \<Gamma> \<Gamma>' and \<Gamma> \<Gamma>' and \<Gamma> \<Gamma>' and \<Gamma> \<Gamma>' and \<Gamma> \<Gamma>' and \<Gamma> \<Gamma>' and \<Gamma> \<Gamma>'
+ rule:wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.strong_induct)
+
+ case (wfE_valI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v b)
+ then show ?case using wf_intros using wf_intros wf_replace_true1 by metis
+next
+ case (wfE_plusI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 v2)
+ then show ?case using wf_intros wf_replace_true1 by metis
+next
+ case (wfE_leqI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 v2)
+ then show ?case using wf_intros wf_replace_true1 by metis
+next
+ case (wfE_eqI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 b v2)
+ then show ?case using wf_intros wf_replace_true1 by metis
+next
+ case (wfE_fstI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 b1 b2)
+ then show ?case using wf_intros wf_replace_true1 by metis
+next
+ case (wfE_sndI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 b1 b2)
+ then show ?case using wf_intros wf_replace_true1 by metis
+next
+ case (wfE_concatI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 v2)
+ then show ?case using wf_intros wf_replace_true1 by metis
+next
+ case (wfE_splitI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 v2)
+ then show ?case using wf_intros wf_replace_true1 by metis
+next
+ case (wfE_lenI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1)
+ then show ?case using wf_intros wf_replace_true1 by metis
+next
+ case (wfE_appI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> f x b c \<tau> s v)
+ then show ?case using wf_intros wf_replace_true1 by metis
+next
+ case (wfE_appPI \<Theta> \<Phi> \<B> \<Gamma>'' \<Delta> b' bv v \<tau> f x1 b1 c1 s)
+ show ?case proof
+ show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using wfE_appPI wf_replace_true1 by metis
+ show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wfE_appPI by metis
+ show \<open> \<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f b' \<close> using wfE_appPI by metis
+ have "atom bv \<sharp> \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>" using fresh_u_replace_true wfE_appPI fresh_prodN by metis
+ thus \<open>atom bv \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>, \<Delta>, b', v, (b_of \<tau>)[bv::=b']\<^sub>b)\<close>
+ using wfE_appPI fresh_prodN by auto
+ show \<open>Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x1 b1 c1 \<tau> s))) = lookup_fun \<Phi> f\<close> using wfE_appPI by metis
+ show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b1[bv::=b']\<^sub>b \<close> using wfE_appPI wf_replace_true1 by metis
+ qed
+next
+ case (wfE_mvarI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> u \<tau>)
+ then show ?case using wf_intros wf_replace_true1 by metis
+next
+
+ case (wfS_valI \<Theta> \<Phi> \<B> \<Gamma> v b \<Delta>)
+ then show ?case using wf_intros wf_replace_true1 by metis
+next
+ case (wfS_letI \<Theta> \<Phi> \<B> \<Gamma>'' \<Delta> e b' x1 s b1)
+ show ?case proof
+ show \<open> \<Theta> ; \<Phi> ; \<B> ; \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f e : b' \<close> using wfS_letI wf_replace_true1 by metis
+ have \<open> \<Theta> ; \<Phi> ; \<B> ; ((x1, b', TRUE) #\<^sub>\<Gamma> \<Gamma>') @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b1 \<close> apply(rule wfS_letI(4))
+ using wfS_letI append_g.simps by simp
+ thus \<open> \<Theta> ; \<Phi> ; \<B> ; (x1, b', TRUE) #\<^sub>\<Gamma> \<Gamma>'@ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b1 \<close> using append_g.simps by auto
+ show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wfS_letI by metis
+ show "atom x1 \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>, \<Delta>, e, b1)" using fresh_append_g fresh_GCons fresh_prodN wfS_letI by auto
+ qed
+next
+ case (wfS_assertI \<Theta> \<Phi> \<B> x' c \<Gamma>'' \<Delta> s b')
+ show ?case proof
+ show \<open> \<Theta> ; \<Phi> ; \<B> ; (x', B_bool, c) #\<^sub>\<Gamma> \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b' \<close>
+ using wfS_assertI (2)[of "(x', B_bool, c) #\<^sub>\<Gamma> \<Gamma>'" \<Gamma>] wfS_assertI by simp
+ show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f c \<close> using wfS_assertI wf_replace_true1 by metis
+ show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wfS_assertI by metis
+ show \<open>atom x' \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>, \<Delta>, c, b', s)\<close> using wfS_assertI fresh_prodN by simp
+ qed
+next
+ case (wfS_let2I \<Theta> \<Phi> \<B> \<Gamma>'' \<Delta> s1 \<tau> x' s2 ba')
+ show ?case proof
+ show \<open> \<Theta> ; \<Phi> ; \<B> ; \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s1 : b_of \<tau> \<close> using wfS_let2I wf_replace_true1 by metis
+ show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau> \<close> using wfS_let2I wf_replace_true1 by metis
+ have \<open> \<Theta> ; \<Phi> ; \<B> ; ((x', b_of \<tau>, TRUE) #\<^sub>\<Gamma> \<Gamma>') @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s2 : ba' \<close>
+ apply(rule wfS_let2I(5))
+ using wfS_let2I append_g.simps by auto
+ thus \<open> \<Theta> ; \<Phi> ; \<B> ; (x', b_of \<tau>, TRUE) #\<^sub>\<Gamma> \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s2 : ba' \<close> using wfS_let2I append_g.simps by auto
+ show \<open>atom x' \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>, \<Delta>, s1, ba', \<tau>)\<close> using fresh_append_g fresh_GCons fresh_prodN wfS_let2I by auto
+ qed
+next
+ case (wfS_ifI \<Theta> \<B> \<Gamma> v \<Phi> \<Delta> s1 b s2)
+ then show ?case using wf_intros wf_replace_true1 by metis
+next
+ case (wfS_varI \<Theta> \<B> \<Gamma>'' \<tau> v u \<Phi> \<Delta> b' s)
+ show ?case proof
+ show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau> \<close> using wfS_varI wf_replace_true1 by metis
+ show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b_of \<tau> \<close> using wfS_varI wf_replace_true1 by metis
+ show \<open>atom u \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>, \<Delta>, \<tau>, v, b')\<close> using wfS_varI u_fresh_g fresh_prodN by auto
+ show \<open> \<Theta> ; \<Phi> ; \<B> ; \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> ; (u, \<tau>) #\<^sub>\<Delta> \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b' \<close> using wfS_varI by metis
+ qed
+
+next
+ case (wfS_assignI u \<tau> \<Delta> \<Theta> \<B> \<Gamma> \<Phi> v)
+ then show ?case using wf_intros wf_replace_true1 by metis
+next
+ case (wfS_whileI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> s1 s2 b)
+ then show ?case using wf_intros wf_replace_true1 by metis
+next
+ case (wfS_seqI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> s1 s2 b)
+ then show ?case using wf_intros by metis
+next
+ case (wfS_matchI \<Theta> \<B> \<Gamma>'' v tid dclist \<Delta> \<Phi> cs b')
+ show ?case proof
+ show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : B_id tid \<close> using wfS_matchI wf_replace_true1 by auto
+ show \<open>AF_typedef tid dclist \<in> set \<Theta>\<close> using wfS_matchI by auto
+ show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wfS_matchI by auto
+ show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using wfS_matchI by auto
+ show \<open> \<Theta> ; \<Phi> ; \<B> ; \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> ; tid ; dclist \<turnstile>\<^sub>w\<^sub>f cs : b' \<close> using wfS_matchI by auto
+ qed
+next
+ case (wfS_branchI \<Theta> \<Phi> \<B> x' \<tau> \<Gamma>'' \<Delta> s b' tid dc)
+ show ?case proof
+ have \<open> \<Theta> ; \<Phi> ; \<B> ; ((x', b_of \<tau>, TRUE) #\<^sub>\<Gamma> \<Gamma>') @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b' \<close> using wfS_branchI append_g.simps by metis
+ thus \<open> \<Theta> ; \<Phi> ; \<B> ; (x', b_of \<tau>, TRUE) #\<^sub>\<Gamma> \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b' \<close> using wfS_branchI append_g.simps append_g.simps by metis
+ show \<open>atom x' \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>, \<Delta>, \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>, \<tau>)\<close> using wfS_branchI by auto
+ show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wfS_branchI by auto
+ qed
+next
+ case (wfS_finalI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dc t cs b)
+ then show ?case using wf_intros by metis
+next
+ case (wfS_cons \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dc t cs b dclist css)
+ then show ?case using wf_intros by metis
+next
+ case (wfD_emptyI \<Theta> \<B> \<Gamma>)
+ then show ?case using wf_intros wf_replace_true1 by metis
+next
+ case (wfD_cons \<Theta> \<B> \<Gamma> \<Delta> \<tau> u)
+ then show ?case using wf_intros wf_replace_true1 by metis
+next
+ case (wfPhi_emptyI \<Theta>)
+ then show ?case using wf_intros by metis
+next
+ case (wfPhi_consI f \<Theta> \<Phi> ft)
+ then show ?case using wf_intros by metis
+next
+ case (wfFTNone \<Theta> \<Phi> ft)
+ then show ?case using wf_intros by metis
+next
+ case (wfFTSome \<Theta> \<Phi> bv ft)
+ then show ?case using wf_intros by metis
+next
+ case (wfFTI \<Theta> B b \<Phi> x c s \<tau>)
+ then show ?case using wf_intros by metis
+qed
+
+lemmas wf_replace_true = wf_replace_true1 wf_replace_true2
+
+section \<open>Subtyping\<close>
+
+lemma subtype_reflI2:
+ fixes \<tau>::\<tau>
+ assumes "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau>"
+ shows "\<Theta>; \<B>; \<Gamma> \<turnstile> \<tau> \<lesssim> \<tau>"
+proof -
+ obtain z b c where *:"\<tau> = \<lbrace> z : b | c \<rbrace> \<and> atom z \<sharp> (\<Theta>,\<B>,\<Gamma>) \<and> \<Theta>; \<B>; (z, b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f c"
+ using wfT_elims(1)[OF assms] by metis
+ obtain x::x where **: "atom x \<sharp> (\<Theta>, \<B>, \<Gamma>, c, z ,c, z , c )" using obtain_fresh by metis
+ have "\<Theta>; \<B>; \<Gamma> \<turnstile> \<lbrace> z : b | c \<rbrace> \<lesssim> \<lbrace> z : b | c \<rbrace>" proof
+ show "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b | c \<rbrace>" using * assms by auto
+ show "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b | c \<rbrace>" using * assms by auto
+ show "atom x \<sharp> (\<Theta>, \<B>, \<Gamma>, z , c , z , c )" using fresh_prod6 fresh_prod5 ** by metis
+ thus "\<Theta>; \<B>; (x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<Turnstile> c[z::=V_var x]\<^sub>v " using wfT_wfC_cons assms * ** subst_v_c_def by simp
+ qed
+ thus ?thesis using * by auto
+qed
+
+lemma subtype_reflI:
+ assumes "\<lbrace> z1 : b | c1 \<rbrace> = \<lbrace> z2 : b | c2 \<rbrace>" and wf1: "\<Theta>; \<B>;\<Gamma> \<turnstile>\<^sub>w\<^sub>f (\<lbrace> z1 : b | c1 \<rbrace>)"
+ shows "\<Theta>; \<B>; \<Gamma> \<turnstile> (\<lbrace> z1 : b | c1 \<rbrace>) \<lesssim> (\<lbrace> z2 : b | c2 \<rbrace>)"
+ using assms subtype_reflI2 by metis
+
+nominal_function base_eq :: "\<Gamma> \<Rightarrow> \<tau> \<Rightarrow> \<tau> \<Rightarrow> bool" where
+ "base_eq _ \<lbrace> z1 : b1| c1 \<rbrace> \<lbrace> z2 : b2 | c2 \<rbrace> = (b1 = b2)"
+ apply(auto,simp add: eqvt_def base_eq_graph_aux_def )
+ by (meson \<tau>.exhaust)
+nominal_termination (eqvt) by lexicographic_order
+
+lemma subtype_wfT:
+ fixes t1::\<tau> and t2::\<tau>
+ assumes "\<Theta>; \<B>; \<Gamma> \<turnstile> t1 \<lesssim> t2"
+ shows "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f t1 \<and> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f t2"
+ using assms subtype_elims by metis
+
+lemma subtype_eq_base:
+ assumes "\<Theta>; \<B>; \<Gamma> \<turnstile> (\<lbrace> z1 : b1| c1 \<rbrace>) \<lesssim> (\<lbrace> z2 : b2 | c2 \<rbrace>)"
+ shows "b1=b2"
+ using subtype.simps assms by auto
+
+lemma subtype_eq_base2:
+ assumes "\<Theta>; \<B>; \<Gamma> \<turnstile> t1 \<lesssim> t2"
+ shows "b_of t1 = b_of t2"
+ using assms proof(rule subtype.induct[of \<Theta> \<B> \<Gamma> t1 t2],goal_cases)
+ case (1 \<Theta> \<Gamma> z1 b c1 z2 c2 x)
+ then show ?case using subtype_eq_base by auto
+qed
+
+lemma subtype_wf:
+ fixes \<tau>1::\<tau> and \<tau>2::\<tau>
+ assumes "\<Theta>; \<B>; \<Gamma> \<turnstile> \<tau>1 \<lesssim> \<tau>2"
+ shows "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau>1 \<and>\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau>2"
+ using assms
+proof(rule subtype.induct[of \<Theta> \<B> \<Gamma> \<tau>1 \<tau>2],goal_cases)
+ case (1 \<Theta> \<Gamma>G z1 b c1 z2 c2 x)
+ then show ?case by blast
+qed
+
+lemma subtype_g_wf:
+ fixes \<tau>1::\<tau> and \<tau>2::\<tau> and \<Gamma>::\<Gamma>
+ assumes "\<Theta>; \<B>; \<Gamma> \<turnstile> \<tau>1 \<lesssim> \<tau>2"
+ shows "\<Theta> ; \<B>\<turnstile>\<^sub>w\<^sub>f \<Gamma>"
+ using assms
+proof(rule subtype.induct[of \<Theta> \<B> \<Gamma> \<tau>1 \<tau>2],goal_cases)
+ case (1 \<Theta> \<B> \<Gamma> z1 b c1 z2 c2 x)
+ then show ?case using wfX_wfY by auto
+qed
+
+text \<open> For when we have a particular y that satisfies the freshness conditions that we want the validity check to use \<close>
+
+lemma valid_flip_simple:
+ assumes "\<Theta>; \<B>; (z, b, c) #\<^sub>\<Gamma> \<Gamma> \<Turnstile> c'" and "atom z \<sharp> \<Gamma>" and "atom x \<sharp> (z, c, z, c', \<Gamma>)"
+ shows "\<Theta>; \<B>; (x, b, (z \<leftrightarrow> x ) \<bullet> c) #\<^sub>\<Gamma> \<Gamma> \<Turnstile> (z \<leftrightarrow> x ) \<bullet> c'"
+proof -
+ have "(z \<leftrightarrow> x ) \<bullet> \<Theta>; \<B>; (z \<leftrightarrow> x ) \<bullet> ((z, b, c) #\<^sub>\<Gamma> \<Gamma>) \<Turnstile> (z \<leftrightarrow> x ) \<bullet> c'"
+ using True_eqvt valid.eqvt assms beta_flip_eq wfX_wfY by metis
+ moreover have " \<turnstile>\<^sub>w\<^sub>f \<Theta>" using valid.simps wfC_wf wfG_wf assms by metis
+ ultimately show ?thesis
+ using theta_flip_eq G_cons_flip_fresh3[of x \<Gamma> z b c] assms fresh_Pair flip_commute by metis
+qed
+
+lemma valid_wf_all:
+ assumes " \<Theta>; \<B>; (z0,b,c0)#\<^sub>\<Gamma>G \<Turnstile> c"
+ shows "wfG \<Theta> \<B> G" and "wfC \<Theta> \<B> ((z0,b,c0)#\<^sub>\<Gamma>G) c" and "atom z0 \<sharp> G"
+ using valid.simps wfC_wf wfG_cons assms by metis+
+
+lemma valid_wfT:
+ fixes z::x
+ assumes " \<Theta>; \<B>; (z0,b,c0[z::=V_var z0]\<^sub>v)#\<^sub>\<Gamma>G \<Turnstile> c[z::=V_var z0]\<^sub>v" and "atom z0 \<sharp> (\<Theta>, \<B>, G,c,c0)"
+ shows "\<Theta>; \<B>; G \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b | c0 \<rbrace>" and "\<Theta>; \<B>; G \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b | c \<rbrace>"
+proof -
+ have "atom z0 \<sharp> c0" using assms fresh_Pair by auto
+ moreover have *:" \<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f (z0,b,c0[z::=V_var z0]\<^sub>c\<^sub>v)#\<^sub>\<Gamma>G" using valid_wf_all wfX_wfY assms subst_v_c_def by metis
+ ultimately show wft: "\<Theta>; \<B>; G \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b | c0 \<rbrace>" using wfG_wfT[OF *] by auto
+
+ have "atom z0 \<sharp> c" using assms fresh_Pair by auto
+ moreover have wfc: "\<Theta>; \<B>; (z0,b,c0[z::=V_var z0]\<^sub>v)#\<^sub>\<Gamma>G \<turnstile>\<^sub>w\<^sub>f c[z::=V_var z0]\<^sub>v" using valid_wf_all assms by metis
+ have "\<Theta>; \<B>; G \<turnstile>\<^sub>w\<^sub>f \<lbrace> z0 : b | c[z::=V_var z0]\<^sub>v \<rbrace>" proof
+ show \<open>atom z0 \<sharp> (\<Theta>, \<B>, G)\<close> using assms fresh_prodN by simp
+ show \<open> \<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f b \<close> using wft wfT_wfB by force
+ show \<open> \<Theta>; \<B>; (z0, b, TRUE) #\<^sub>\<Gamma> G \<turnstile>\<^sub>w\<^sub>f c[z::=[ z0 ]\<^sup>v]\<^sub>v \<close> using wfc wfC_replace_inside[OF wfc, of GNil z0 b "c0[z::=[ z0 ]\<^sup>v]\<^sub>v" G C_true] wfC_trueI
+ append_g.simps
+ by (metis "local.*" wfG_elim2 wf_trans(2))
+ qed
+ moreover have "\<lbrace> z0 : b | c[z::=V_var z0]\<^sub>v \<rbrace> = \<lbrace> z : b | c \<rbrace>" using \<open>atom z0 \<sharp> c0\<close> \<tau>.eq_iff Abs1_eq_iff(3)
+ using calculation(1) subst_v_c_def by auto
+ ultimately show "\<Theta>; \<B>; G \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b | c \<rbrace>" by auto
+qed
+
+lemma valid_flip:
+ fixes c::c and z::x and z0::x and xx2::x
+ assumes " \<Theta>; \<B>; (xx2, b, c0[z0::=V_var xx2]\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<Turnstile> c[z::=V_var xx2]\<^sub>v" and
+ "atom xx2 \<sharp> (c0,\<Gamma>,c,z) " and "atom z0 \<sharp> (\<Gamma>,c,z)"
+ shows " \<Theta>; \<B>; (z0, b, c0) #\<^sub>\<Gamma> \<Gamma> \<Turnstile> c[z::=V_var z0]\<^sub>v"
+proof -
+
+ have " \<turnstile>\<^sub>w\<^sub>f \<Theta>" using assms valid_wf_all wfX_wfY by metis
+ hence " \<Theta> ; \<B> ; (xx2 \<leftrightarrow> z0 ) \<bullet> ((xx2, b, c0[z0::=V_var xx2]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) \<Turnstile> ((xx2 \<leftrightarrow> z0 ) \<bullet> c[z::=V_var xx2]\<^sub>v)"
+ using valid.eqvt True_eqvt assms beta_flip_eq theta_flip_eq by metis
+ hence " \<Theta>; \<B>; (((xx2 \<leftrightarrow> z0 ) \<bullet> xx2, b, (xx2 \<leftrightarrow> z0 ) \<bullet> c0[z0::=V_var xx2]\<^sub>v) #\<^sub>\<Gamma> (xx2 \<leftrightarrow> z0 ) \<bullet> \<Gamma>) \<Turnstile> ((xx2 \<leftrightarrow> z0 ) \<bullet>(c[z::=V_var xx2]\<^sub>v))"
+ using G_cons_flip[of xx2 z0 xx2 b "c0[z0::=V_var xx2]\<^sub>v" \<Gamma>] by auto
+ moreover have "(xx2 \<leftrightarrow> z0 ) \<bullet> xx2 = z0" by simp
+ moreover have "(xx2 \<leftrightarrow> z0 ) \<bullet> c0[z0::=V_var xx2]\<^sub>v = c0"
+ using assms subst_cv_v_flip[of xx2 c0 z0 "V_var z0"] assms fresh_prod4 by auto
+ moreover have "(xx2 \<leftrightarrow> z0 ) \<bullet> \<Gamma> = \<Gamma>" proof -
+ have "atom xx2 \<sharp> \<Gamma>" using assms by auto
+ moreover have "atom z0 \<sharp> \<Gamma>" using assms by auto
+ ultimately show ?thesis using flip_fresh_fresh by auto
+ qed
+ moreover have "(xx2 \<leftrightarrow> z0 ) \<bullet> (c[z::=V_var xx2]\<^sub>v) =c[z::=V_var z0]\<^sub>v"
+ using subst_cv_v_flip3 assms by simp
+ ultimately show ?thesis by auto
+qed
+
+lemma subtype_valid:
+ assumes "\<Theta>; \<B>; \<Gamma> \<turnstile> t1 \<lesssim> t2" and "atom y \<sharp> \<Gamma>" and "t1 = \<lbrace> z1 : b | c1 \<rbrace>" and "t2 = \<lbrace> z2 : b | c2 \<rbrace>"
+ shows "\<Theta>; \<B>; ((y, b, c1[z1::=V_var y]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) \<Turnstile> c2[z2::=V_var y]\<^sub>v"
+ using assms proof(nominal_induct t2 avoiding: y rule: subtype.strong_induct)
+ case (subtype_baseI x \<Theta> \<B> \<Gamma> z c z' c' ba)
+
+ hence "(x \<leftrightarrow> y) \<bullet> \<Theta> ; (x \<leftrightarrow> y) \<bullet> \<B> ; (x \<leftrightarrow> y) \<bullet> ((x, ba, c[z::=[ x ]\<^sup>v]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) \<Turnstile> (x \<leftrightarrow> y) \<bullet> c'[z'::=[ x ]\<^sup>v]\<^sub>v" using valid.eqvt
+ using permute_boolI by blast
+ moreover have " \<turnstile>\<^sub>w\<^sub>f \<Theta>" using valid.simps wfC_wf wfG_wf subtype_baseI by metis
+ ultimately have "\<Theta>; \<B>; ((y, ba, (x \<leftrightarrow> y) \<bullet> c[z::=[ x ]\<^sup>v]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) \<Turnstile> (x \<leftrightarrow> y) \<bullet> c'[z'::=[ x ]\<^sup>v]\<^sub>v"
+ using subtype_baseI theta_flip_eq beta_flip_eq \<tau>.eq_iff G_cons_flip_fresh3[of y \<Gamma> x ba] by (metis flip_commute)
+ moreover have "(x \<leftrightarrow> y) \<bullet> c[z::=[ x ]\<^sup>v]\<^sub>v = c1[z1::=[ y ]\<^sup>v]\<^sub>v"
+ by (metis subtype_baseI permute_flip_cancel subst_cv_id subst_cv_v_flip3 subst_cv_var_flip type_eq_subst_eq wfT_fresh_c subst_v_c_def)
+ moreover have "(x \<leftrightarrow> y) \<bullet> c'[z'::=[ x ]\<^sup>v]\<^sub>v = c2[z2::=[ y ]\<^sup>v]\<^sub>v"
+ by (metis subtype_baseI permute_flip_cancel subst_cv_id subst_cv_v_flip3 subst_cv_var_flip type_eq_subst_eq wfT_fresh_c subst_v_c_def)
+
+ ultimately show ?case using subtype_baseI \<tau>.eq_iff by metis
+qed
+
+lemma subtype_valid_simple:
+ assumes "\<Theta>; \<B>; \<Gamma> \<turnstile> t1 \<lesssim> t2" and "atom z \<sharp> \<Gamma>" and "t1 = \<lbrace> z : b | c1 \<rbrace>" and "t2 = \<lbrace> z : b | c2 \<rbrace>"
+ shows "\<Theta>; \<B>; ((z, b, c1) #\<^sub>\<Gamma> \<Gamma>) \<Turnstile> c2"
+ using subst_v_c_def subst_v_id assms subtype_valid[OF assms] by simp
+
+lemma obtain_for_t_with_fresh:
+ assumes "atom x \<sharp> t"
+ shows "\<exists>b c. t = \<lbrace> x : b | c \<rbrace>"
+proof -
+ obtain z1 b1 c1 where *:" t = \<lbrace> z1 : b1 | c1 \<rbrace> \<and> atom z1 \<sharp> t" using obtain_fresh_z by metis
+ then have "t = (x \<leftrightarrow> z1) \<bullet> t" using flip_fresh_fresh assms by metis
+ also have "... = \<lbrace> (x \<leftrightarrow> z1) \<bullet> z1 : (x \<leftrightarrow> z1) \<bullet> b1 | (x \<leftrightarrow> z1) \<bullet> c1 \<rbrace>" using * assms by simp
+ also have "... = \<lbrace> x : b1 | (x \<leftrightarrow> z1) \<bullet> c1 \<rbrace>" using * assms by auto
+ finally show ?thesis by auto
+qed
+
+lemma subtype_trans:
+ assumes "\<Theta>; \<B>; \<Gamma> \<turnstile> \<tau>1 \<lesssim> \<tau>2" and "\<Theta>; \<B>; \<Gamma> \<turnstile> \<tau>2 \<lesssim> \<tau>3"
+ shows "\<Theta>; \<B>; \<Gamma> \<turnstile> \<tau>1 \<lesssim> \<tau>3"
+ using assms proof(nominal_induct avoiding: \<tau>3 rule: subtype.strong_induct)
+ case (subtype_baseI x \<Theta> \<B> \<Gamma> z c z' c' b)
+ hence "b_of \<tau>3 = b" using subtype_eq_base2 b_of.simps by metis
+ then obtain z'' c'' where t3: "\<tau>3 = \<lbrace> z'' : b | c''\<rbrace> \<and> atom z'' \<sharp> x"
+ using obtain_fresh_z2 by metis
+ hence xf: "atom x \<sharp> (z'', c'')" using fresh_prodN subtype_baseI \<tau>.fresh by auto
+ have "\<Theta>; \<B>; \<Gamma> \<turnstile> \<lbrace> z : b | c \<rbrace> \<lesssim> \<lbrace> z'' : b | c''\<rbrace>"
+ proof(rule Typing.subtype_baseI)
+ show \<open>atom x \<sharp> (\<Theta>, \<B>, \<Gamma>, z, c, z'', c'')\<close> using t3 fresh_prodN subtype_baseI xf by simp
+ show \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b | c \<rbrace> \<close> using subtype_baseI by auto
+ show \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> z'' : b | c'' \<rbrace> \<close> using subtype_baseI t3 subtype_elims by metis
+ have " \<Theta>; \<B>; (x, b, c'[z'::=[ x ]\<^sup>v]\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<Turnstile> c''[z''::=[ x ]\<^sup>v]\<^sub>v "
+ using subtype_valid[OF \<open>\<Theta>; \<B>; \<Gamma> \<turnstile> \<lbrace> z' : b | c' \<rbrace> \<lesssim> \<tau>3\<close> , of x z' b c' z'' c''] subtype_baseI
+ t3 by simp
+ thus \<open>\<Theta>; \<B>; (x, b, c[z::=[ x ]\<^sup>v]\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<Turnstile> c''[z''::=[ x ]\<^sup>v]\<^sub>v \<close>
+ using valid_trans_full[of \<Theta> \<B> x b c z \<Gamma> c' z' c'' z'' ] subtype_baseI t3 by simp
+ qed
+ thus ?case using t3 by simp
+qed
+
+lemma subtype_eq_e:
+ assumes "\<forall>i s1 s2 G. wfG P \<B> G \<and> wfI P G i \<and> eval_e i e1 s1 \<and> eval_e i e2 s2 \<longrightarrow> s1 = s2" and "atom z1 \<sharp> e1" and "atom z2 \<sharp> e2" and "atom z1 \<sharp> \<Gamma>" and "atom z2 \<sharp> \<Gamma>"
+ and "wfCE P \<B> \<Gamma> e1 b" and "wfCE P \<B> \<Gamma> e2 b"
+ shows "P; \<B>; \<Gamma> \<turnstile> \<lbrace> z1 : b | CE_val (V_var z1) == e1 \<rbrace> \<lesssim> (\<lbrace> z2 : b | CE_val (V_var z2) == e2 \<rbrace>)"
+proof -
+
+ have "wfCE P \<B> \<Gamma> e1 b" and "wfCE P \<B> \<Gamma> e2 b" using assms by auto
+
+ have wst1: "wfT P \<B> \<Gamma> (\<lbrace> z1 : b | CE_val (V_var z1) == e1 \<rbrace>)"
+ using wfC_e_eq wfTI assms wfX_wfB wfG_fresh_x
+ by (simp add: wfT_e_eq)
+
+ moreover have wst2:"wfT P \<B> \<Gamma> (\<lbrace> z2 : b | CE_val (V_var z2) == e2 \<rbrace>)"
+ using wfC_e_eq wfX_wfB wfTI assms wfG_fresh_x
+ by (simp add: wfT_e_eq)
+
+ moreover obtain x::x where xf: "atom x \<sharp> (P, \<B> , z1, CE_val (V_var z1) == e1 , z2, CE_val (V_var z2) == e2 , \<Gamma>)" using obtain_fresh by blast
+ moreover have vld: "P; \<B> ; (x, b, (CE_val (V_var z1) == e1 )[z1::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<Turnstile> (CE_val (V_var z2) == e2 )[z2::=V_var x]\<^sub>v " (is "P; \<B> ; ?G \<Turnstile> ?c")
+ proof -
+
+ have wbg: "P; \<B> \<turnstile>\<^sub>w\<^sub>f ?G \<and> P ; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<and> toSet \<Gamma> \<subseteq> toSet ?G" proof -
+ have "P; \<B> \<turnstile>\<^sub>w\<^sub>f ?G" proof(rule wfG_consI)
+ show "P; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>" using assms wfX_wfY by metis
+ show "atom x \<sharp> \<Gamma>" using xf by auto
+ show "P; \<B> \<turnstile>\<^sub>w\<^sub>f b " using assms(6) wfX_wfB by auto
+ show "P; \<B> ; (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f (CE_val (V_var z1) == e1 )[z1::=V_var x]\<^sub>v "
+ using wfC_e_eq[OF assms(6)] wf_subst(2)
+ by (simp add: \<open>atom x \<sharp> \<Gamma>\<close> assms(2) subst_v_c_def)
+ qed
+ moreover hence "P; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>" using wfG_elims by metis
+ ultimately show ?thesis using toSet.simps by auto
+ qed
+
+ have wsc: "wfC P \<B> ?G ?c" proof -
+ have "wfCE P \<B> ?G (CE_val (V_var x)) b" proof
+ show \<open> P; \<B> ; (x, b, (CE_val (V_var z1) == e1 )[z1::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f V_var x : b \<close> using wfV_varI lookup.simps wbg by auto
+ qed
+ moreover have "wfCE P \<B> ?G e2 b" using wf_weakening assms wbg by metis
+ ultimately have "wfC P \<B> ?G (CE_val (V_var x) == e2 )" using wfC_eqI by simp
+ thus ?thesis using subst_cv.simps(6) \<open>atom z2 \<sharp> e2\<close> subst_v_c_def by simp
+ qed
+
+ moreover have "\<forall>i. wfI P ?G i \<and> is_satis_g i ?G \<longrightarrow> is_satis i ?c" proof(rule allI , rule impI)
+ fix i
+ assume as: "wfI P ?G i \<and> is_satis_g i ?G"
+ hence "is_satis i ((CE_val (V_var z1) == e1 )[z1::=V_var x]\<^sub>v)"
+ by (simp add: is_satis_g.simps(2))
+ hence "is_satis i (CE_val (V_var x) == e1 )" using subst_cv.simps assms subst_v_c_def by auto
+ then obtain s1 and s2 where *:"eval_e i (CE_val (V_var x)) s1 \<and> eval_e i e1 s2 \<and> s1=s2" using is_satis.simps eval_c_elims by metis
+ moreover hence "eval_e i e2 s1" proof -
+ have **:"wfI P ?G i" using as by auto
+ moreover have "wfCE P \<B> ?G e1 b \<and> wfCE P \<B> ?G e2 b" using assms xf wf_weakening wbg by metis
+ moreover then obtain s2' where "eval_e i e2 s2'" using assms wfI_wfCE_eval_e ** by metis
+ ultimately show ?thesis using * assms(1) wfX_wfY by metis
+ qed
+ ultimately have "is_satis i (CE_val (V_var x) == e2 )" using is_satis.simps eval_c_eqI by force
+ thus "is_satis i ((CE_val (V_var z2) == e2 )[z2::=V_var x]\<^sub>v)" using is_satis.simps eval_c_eqI assms subst_cv.simps subst_v_c_def by auto
+ qed
+ ultimately show ?thesis using valid.simps by simp
+ qed
+ moreover have "atom x \<sharp> (P, \<B>, \<Gamma>, z1 , CE_val (V_var z1) == e1, z2, CE_val (V_var z2) == e2 ) "
+ unfolding fresh_prodN using xf fresh_prod7 \<tau>.fresh by fast
+ ultimately show ?thesis using subtype_baseI[OF _ wst1 wst2 vld] xf by simp
+qed
+
+lemma subtype_eq_e_nil:
+ assumes "\<forall>i s1 s2 G. wfG P \<B> G \<and> wfI P G i \<and> eval_e i e1 s1 \<and> eval_e i e2 s2 \<longrightarrow> s1 = s2" and "supp e1 = {}" and "supp e2 = {}" and "wfTh P"
+ and "wfCE P \<B> GNil e1 b" and "wfCE P \<B> GNil e2 b" and "atom z1 \<sharp> GNil" and "atom z2 \<sharp> GNil"
+ shows "P; \<B> ; GNil \<turnstile> \<lbrace> z1 : b | CE_val (V_var z1) == e1 \<rbrace> \<lesssim> (\<lbrace> z2 : b | CE_val (V_var z2) == e2 \<rbrace>)"
+ apply(rule subtype_eq_e,auto simp add: assms e.fresh)
+ using assms fresh_def e.fresh supp_GNil by metis+
+
+lemma subtype_gnil_fst_aux:
+ assumes "supp v\<^sub>1 = {}" and "supp (V_pair v\<^sub>1 v\<^sub>2) = {}" and "wfTh P" and "wfCE P \<B> GNil (CE_val v\<^sub>1) b" and "wfCE P \<B> GNil (CE_fst [V_pair v\<^sub>1 v\<^sub>2]\<^sup>c\<^sup>e) b" and
+ "wfCE P \<B> GNil (CE_val v\<^sub>2) b2" and "atom z1 \<sharp> GNil" and "atom z2 \<sharp> GNil"
+ shows "P; \<B> ; GNil \<turnstile> (\<lbrace> z1 : b | CE_val (V_var z1) == CE_val v\<^sub>1 \<rbrace>) \<lesssim> (\<lbrace> z2 : b | CE_val (V_var z2) == CE_fst [V_pair v\<^sub>1 v\<^sub>2]\<^sup>c\<^sup>e \<rbrace>)"
+proof -
+ have "\<forall>i s1 s2 G . wfG P \<B> G \<and> wfI P G i \<and> eval_e i ( CE_val v\<^sub>1) s1 \<and> eval_e i (CE_fst [V_pair v\<^sub>1 v\<^sub>2]\<^sup>c\<^sup>e) s2 \<longrightarrow> s1 = s2" proof(rule+)
+ fix i s1 s2 G
+ assume as: "wfG P \<B> G \<and> wfI P G i \<and> eval_e i (CE_val v\<^sub>1) s1 \<and> eval_e i (CE_fst [V_pair v\<^sub>1 v\<^sub>2]\<^sup>c\<^sup>e) s2"
+ hence "wfCE P \<B> G (CE_val v\<^sub>2) b2" using assms wf_weakening
+ by (metis empty_subsetI toSet.simps(1))
+ then obtain s3 where "eval_e i (CE_val v\<^sub>2) s3" using wfI_wfCE_eval_e as by metis
+ hence "eval_v i ((V_pair v\<^sub>1 v\<^sub>2)) (SPair s1 s3)"
+ by (meson as eval_e_elims(1) eval_v_pairI)
+ hence "eval_e i (CE_fst [V_pair v\<^sub>1 v\<^sub>2]\<^sup>c\<^sup>e) s1" using eval_e_fstI eval_e_valI by metis
+ show "s1 = s2" using as eval_e_uniqueness
+ using \<open>eval_e i (CE_fst [V_pair v\<^sub>1 v\<^sub>2]\<^sup>c\<^sup>e) s1\<close> by auto
+ qed
+ thus ?thesis using subtype_eq_e_nil ce.supp assms by auto
+qed
+
+lemma subtype_gnil_snd_aux:
+ assumes "supp v\<^sub>2 = {}" and "supp (V_pair v\<^sub>1 v\<^sub>2) = {}" and "wfTh P" and "wfCE P \<B> GNil (CE_val v\<^sub>2) b" and
+ "wfCE P \<B> GNil (CE_snd [(V_pair v\<^sub>1 v\<^sub>2)]\<^sup>c\<^sup>e) b" and
+ "wfCE P \<B> GNil (CE_val v\<^sub>1) b2" and "atom z1 \<sharp> GNil" and "atom z2 \<sharp> GNil"
+ shows "P; \<B> ; GNil \<turnstile> (\<lbrace> z1 : b | CE_val (V_var z1) == CE_val v\<^sub>2 \<rbrace>) \<lesssim> (\<lbrace> z2 : b | CE_val (V_var z2) == CE_snd [(V_pair v\<^sub>1 v\<^sub>2)]\<^sup>c\<^sup>e \<rbrace>)"
+proof -
+ have "\<forall>i s1 s2 G. wfG P \<B> G \<and> wfI P G i \<and> eval_e i ( CE_val v\<^sub>2) s1 \<and> eval_e i (CE_snd [(V_pair v\<^sub>1 v\<^sub>2)]\<^sup>c\<^sup>e) s2 \<longrightarrow> s1 = s2" proof(rule+)
+ fix i s1 s2 G
+ assume as: " wfG P \<B> G \<and> wfI P G i \<and> eval_e i (CE_val v\<^sub>2) s1 \<and> eval_e i (CE_snd [(V_pair v\<^sub>1 v\<^sub>2)]\<^sup>c\<^sup>e) s2"
+ hence "wfCE P \<B> G (CE_val v\<^sub>1) b2" using assms wf_weakening
+ by (metis empty_subsetI toSet.simps(1))
+ then obtain s3 where "eval_e i (CE_val v\<^sub>1) s3" using wfI_wfCE_eval_e as by metis
+ hence "eval_v i ((V_pair v\<^sub>1 v\<^sub>2)) (SPair s3 s1)"
+ by (meson as eval_e_elims eval_v_pairI)
+ hence "eval_e i (CE_snd [(V_pair v\<^sub>1 v\<^sub>2)]\<^sup>c\<^sup>e) s1" using eval_e_sndI eval_e_valI by metis
+ show "s1 = s2" using as eval_e_uniqueness
+ using \<open>eval_e i (CE_snd [V_pair v\<^sub>1 v\<^sub>2]\<^sup>c\<^sup>e) s1\<close> by auto
+ qed
+ thus ?thesis using assms subtype_eq_e_nil by (simp add: ce.supp ce.supp)
+qed
+
+lemma subtype_gnil_fst:
+ assumes "\<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f [#1[[v\<^sub>1,v\<^sub>2]\<^sup>v]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e : b"
+ shows "\<Theta> ; {||} ; GNil \<turnstile> (\<lbrace> z\<^sub>1 : b | [[z\<^sub>1]\<^sup>v]\<^sup>c\<^sup>e == [v\<^sub>1]\<^sup>c\<^sup>e \<rbrace>) \<lesssim>
+ (\<lbrace> z\<^sub>2 : b | [[z\<^sub>2]\<^sup>v]\<^sup>c\<^sup>e == [#1[[v\<^sub>1, v\<^sub>2]\<^sup>v]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e \<rbrace>)"
+proof -
+ obtain b2 where **:" \<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f V_pair v\<^sub>1 v\<^sub>2 : B_pair b b2" using wfCE_elims(4)[OF assms ] wfCE_elims by metis
+ obtain b1' b2' where *:"B_pair b b2 = B_pair b1' b2' \<and> \<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f v\<^sub>1 : b1' \<and> \<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f v\<^sub>2 : b2'"
+ using wfV_elims(3)[OF **] by metis
+
+ show ?thesis proof(rule subtype_gnil_fst_aux)
+ show \<open>supp v\<^sub>1 = {}\<close> using * wfV_supp_nil by auto
+ show \<open>supp (V_pair v\<^sub>1 v\<^sub>2) = {}\<close> using ** wfV_supp_nil e.supp by metis
+ show \<open>\<turnstile>\<^sub>w\<^sub>f \<Theta>\<close> using assms wfX_wfY by metis
+ show \<open>\<Theta>; {||}; GNil \<turnstile>\<^sub>w\<^sub>f CE_val v\<^sub>1 : b \<close> using wfCE_valI "*" by auto
+ show \<open>\<Theta>; {||}; GNil \<turnstile>\<^sub>w\<^sub>f CE_fst [V_pair v\<^sub>1 v\<^sub>2]\<^sup>c\<^sup>e : b \<close> using assms by auto
+ show \<open>\<Theta>; {||}; GNil \<turnstile>\<^sub>w\<^sub>f CE_val v\<^sub>2 : b2 \<close>using wfCE_valI "*" by auto
+ show \<open>atom z\<^sub>1 \<sharp> GNil\<close> using fresh_GNil by metis
+ show \<open>atom z\<^sub>2 \<sharp> GNil\<close> using fresh_GNil by metis
+ qed
+qed
+
+lemma subtype_gnil_snd:
+ assumes "wfCE P {||} GNil (CE_snd ([V_pair v\<^sub>1 v\<^sub>2]\<^sup>c\<^sup>e)) b"
+ shows "P ; {||} ; GNil \<turnstile> (\<lbrace> z1 : b | CE_val (V_var z1) == CE_val v\<^sub>2 \<rbrace>) \<lesssim> (\<lbrace> z2 : b | CE_val (V_var z2) == CE_snd [(V_pair v\<^sub>1 v\<^sub>2)]\<^sup>c\<^sup>e \<rbrace>)"
+proof -
+ obtain b1 where **:" P ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f V_pair v\<^sub>1 v\<^sub>2 : B_pair b1 b " using wfCE_elims assms by metis
+ obtain b1' b2' where *:"B_pair b1 b = B_pair b1' b2' \<and> P ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f v\<^sub>1 : b1' \<and> P ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f v\<^sub>2 : b2'" using wfV_elims(3)[OF **] by metis
+
+ show ?thesis proof(rule subtype_gnil_snd_aux)
+ show \<open>supp v\<^sub>2 = {}\<close> using * wfV_supp_nil by auto
+ show \<open>supp (V_pair v\<^sub>1 v\<^sub>2) = {}\<close> using ** wfV_supp_nil e.supp by metis
+ show \<open>\<turnstile>\<^sub>w\<^sub>f P\<close> using assms wfX_wfY by metis
+ show \<open>P; {||}; GNil \<turnstile>\<^sub>w\<^sub>f CE_val v\<^sub>1 : b1 \<close> using wfCE_valI "*" by simp
+ show \<open>P; {||}; GNil \<turnstile>\<^sub>w\<^sub>f CE_snd [(V_pair v\<^sub>1 v\<^sub>2)]\<^sup>c\<^sup>e : b \<close> using assms by auto
+ show \<open>P; {||}; GNil \<turnstile>\<^sub>w\<^sub>f CE_val v\<^sub>2 : b \<close>using wfCE_valI "*" by simp
+ show \<open>atom z1 \<sharp> GNil\<close> using fresh_GNil by metis
+ show \<open>atom z2 \<sharp> GNil\<close> using fresh_GNil by metis
+ qed
+qed
+
+lemma subtype_fresh_tau:
+ fixes x::x
+ assumes "atom x \<sharp> t1" and "atom x \<sharp> \<Gamma>" and "P; \<B>; \<Gamma> \<turnstile> t1 \<lesssim> t2"
+ shows "atom x \<sharp> t2"
+proof -
+ have wfg: "P; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>" using subtype_wf wfX_wfY assms by metis
+ have wft: "wfT P \<B> \<Gamma> t2" using subtype_wf wfX_wfY assms by blast
+ hence "supp t2 \<subseteq> atom_dom \<Gamma> \<union> supp \<B>" using wf_supp
+ using atom_dom.simps by auto
+ moreover have "atom x \<notin> atom_dom \<Gamma>" using \<open>atom x \<sharp> \<Gamma>\<close> wfG_atoms_supp_eq wfg fresh_def by blast
+ ultimately show "atom x \<sharp> t2" using fresh_def
+ by (metis Un_iff contra_subsetD x_not_in_b_set)
+qed
+
+lemma subtype_if_simp:
+ assumes "wfT P \<B> GNil (\<lbrace> z1 : b | CE_val (V_lit l ) == CE_val (V_lit l) IMP c[z::=V_var z1]\<^sub>v \<rbrace>)" and
+ "wfT P \<B> GNil (\<lbrace> z : b | c \<rbrace>)" and "atom z1 \<sharp> c"
+ shows "P; \<B> ; GNil \<turnstile> (\<lbrace> z1 : b | CE_val (V_lit l) == CE_val (V_lit l) IMP c[z::=V_var z1]\<^sub>v \<rbrace>) \<lesssim> (\<lbrace> z : b | c \<rbrace>)"
+proof -
+ obtain x::x where xx: "atom x \<sharp> ( P , \<B> , z1, CE_val (V_lit l) == CE_val (V_lit l) IMP c[z::=V_var z1]\<^sub>v , z, c, GNil)" using obtain_fresh_z by blast
+ hence xx2: " atom x \<sharp> (CE_val (V_lit l) == CE_val (V_lit l) IMP c[z::=V_var z1]\<^sub>v , c, GNil)" using fresh_prod7 fresh_prod3 by fast
+ have *:"P; \<B> ; (x, b, (CE_val (V_lit l) == CE_val (V_lit l) IMP c[z::=V_var z1]\<^sub>v )[z1::=V_var x]\<^sub>v) #\<^sub>\<Gamma> GNil \<Turnstile> c[z::=V_var x]\<^sub>v " (is "P; \<B> ; ?G \<Turnstile> ?c" ) proof -
+ have "wfC P \<B> ?G ?c" using wfT_wfC_cons[OF assms(1) assms(2),of x] xx fresh_prod5 fresh_prod3 subst_v_c_def by metis
+ moreover have "(\<forall>i. wfI P ?G i \<and> is_satis_g i ?G \<longrightarrow> is_satis i ?c)" proof(rule allI, rule impI)
+ fix i
+ assume as1: "wfI P ?G i \<and> is_satis_g i ?G"
+ have "((CE_val (V_lit l) == CE_val (V_lit l) IMP c[z::=V_var z1]\<^sub>v )[z1::=V_var x]\<^sub>v) = ((CE_val (V_lit l) == CE_val (V_lit l) IMP c[z::=V_var x]\<^sub>v ))"
+ using assms subst_v_c_def by auto
+ hence "is_satis i ((CE_val (V_lit l) == CE_val (V_lit l) IMP c[z::=V_var x]\<^sub>v ))" using is_satis_g.simps as1 by presburger
+ moreover have "is_satis i ((CE_val (V_lit l) == CE_val (V_lit l)))" using is_satis.simps eval_c_eqI[of i "(CE_val (V_lit l))" "eval_l l"] eval_e_uniqueness
+ eval_e_valI eval_v_litI by metis
+ ultimately show "is_satis i ?c" using is_satis_mp[of i] by metis
+ qed
+ ultimately show ?thesis using valid.simps by simp
+ qed
+ moreover have "atom x \<sharp> (P, \<B>, GNil, z1 , CE_val (V_lit l) == CE_val (V_lit l) IMP c[z::=V_var z1]\<^sub>v , z, c )"
+ unfolding fresh_prod5 \<tau>.fresh using xx fresh_prodN x_fresh_b by metis
+ ultimately show ?thesis using subtype_baseI assms xx xx2 by metis
+qed
+
+lemma subtype_if:
+ assumes "P; \<B>; \<Gamma> \<turnstile> \<lbrace> z : b | c \<rbrace> \<lesssim> \<lbrace> z' : b | c' \<rbrace>" and
+ "wfT P \<B> \<Gamma> (\<lbrace> z1 : b | CE_val v == CE_val (V_lit l) IMP c[z::=V_var z1]\<^sub>v \<rbrace>)" and
+ "wfT P \<B> \<Gamma> (\<lbrace> z2 : b | CE_val v == CE_val (V_lit l) IMP c'[z'::=V_var z2]\<^sub>v \<rbrace>)" and
+ "atom z1 \<sharp> v" and "atom z \<sharp> \<Gamma>" and "atom z1 \<sharp> c" and "atom z2 \<sharp> c'" and "atom z2 \<sharp> v"
+ shows "P; \<B> ; \<Gamma> \<turnstile> \<lbrace> z1 : b | CE_val v == CE_val (V_lit l) IMP c[z::=V_var z1]\<^sub>v \<rbrace> \<lesssim> \<lbrace> z2 : b | CE_val v == CE_val (V_lit l) IMP c'[z'::=V_var z2]\<^sub>v \<rbrace>"
+proof -
+ obtain x::x where xx: "atom x \<sharp> (P,\<B>,z,c,z', c', z1, CE_val v == CE_val (V_lit l) IMP c[z::=V_var z1]\<^sub>v , z2, CE_val v == CE_val (V_lit l) IMP c'[z'::=V_var z2]\<^sub>v , \<Gamma>)"
+ using obtain_fresh_z by blast
+ hence xf: "atom x \<sharp> (z, c, z', c', \<Gamma>)" by simp
+ have xf2: "atom x \<sharp> (z1, CE_val v == CE_val (V_lit l) IMP c[z::=V_var z1]\<^sub>v , z2, CE_val v == CE_val (V_lit l) IMP c'[z'::=V_var z2]\<^sub>v , \<Gamma>)"
+ using xx fresh_prod4 fresh_prodN by metis
+
+ moreover have "P; \<B> ; (x, b, (CE_val v == CE_val (V_lit l) IMP c[z::=V_var z1]\<^sub>v )[z1::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<Turnstile> (CE_val v == CE_val (V_lit l) IMP c'[z'::=V_var z2]\<^sub>v )[z2::=V_var x]\<^sub>v"
+ (is "P; \<B> ; ?G \<Turnstile> ?c" )
+ proof -
+ have wbc: "wfC P \<B> ?G ?c" using assms xx fresh_prod4 fresh_prod2 wfT_wfC_cons assms subst_v_c_def by metis
+
+ moreover have "\<forall>i. wfI P ?G i \<and> is_satis_g i ?G \<longrightarrow> is_satis i ?c" proof(rule allI, rule impI)
+ fix i
+ assume a1: "wfI P ?G i \<and> is_satis_g i ?G"
+
+ have *: " is_satis i ((CE_val v == CE_val (V_lit l))) \<longrightarrow> is_satis i ((c'[z'::=V_var z2]\<^sub>v )[z2::=V_var x]\<^sub>v)"
+ proof
+ assume a2: "is_satis i ((CE_val v == CE_val (V_lit l)))"
+
+ have "is_satis i ((CE_val v == CE_val (V_lit l) IMP (c[z::=V_var z1]\<^sub>v ))[z1::=V_var x]\<^sub>v)"
+ using a1 is_satis_g.simps by simp
+ moreover have "((CE_val v == CE_val (V_lit l) IMP (c[z::=V_var z1]\<^sub>v ))[z1::=V_var x]\<^sub>v) = (CE_val v == CE_val (V_lit l) IMP ((c[z::=V_var z1]\<^sub>v )[z1::=V_var x]\<^sub>v))"
+ using assms subst_v_c_def by simp
+ ultimately have "is_satis i (CE_val v == CE_val (V_lit l) IMP ((c[z::=V_var z1]\<^sub>v )[z1::=V_var x]\<^sub>v))" by argo
+
+ hence "is_satis i ((c[z::=V_var z1]\<^sub>v )[z1::=V_var x]\<^sub>v)" using a2 is_satis_mp by auto
+ moreover have "((c[z::=V_var z1]\<^sub>v )[z1::=V_var x]\<^sub>v) = ((c[z::=V_var x]\<^sub>v ))" using assms by auto
+ ultimately have "is_satis i ((c[z::=V_var x]\<^sub>v ))" using a2 is_satis.simps by auto
+
+ hence "is_satis_g i ((x,b,(c[z::=V_var x]\<^sub>v )) #\<^sub>\<Gamma> \<Gamma>)" using a1 is_satis_g.simps by meson
+ moreover have "wfI P ((x,b,(c[z::=V_var x]\<^sub>v )) #\<^sub>\<Gamma> \<Gamma>) i" proof -
+ obtain s where "Some s = i x \<and> wfRCV P s b \<and> wfI P \<Gamma> i" using wfI_def a1 by auto
+ thus ?thesis using wfI_def by auto
+ qed
+ ultimately have "is_satis i ((c'[z'::=V_var x]\<^sub>v))" using subtype_valid assms(1) xf valid.simps by simp
+
+ moreover have "(c'[z'::=V_var x]\<^sub>v) = ((c'[z'::=V_var z2]\<^sub>v )[z2::=V_var x]\<^sub>v)" using assms by auto
+ ultimately show "is_satis i ((c'[z'::=V_var z2]\<^sub>v )[z2::=V_var x]\<^sub>v)" by auto
+ qed
+
+ moreover have "?c = ((CE_val v == CE_val (V_lit l)) IMP ((c'[z'::=V_var z2]\<^sub>v )[z2::=V_var x]\<^sub>v))"
+ using assms subst_v_c_def by simp
+
+ moreover have "\<exists>b1 b2. eval_c i (CE_val v == CE_val (V_lit l) ) b1 \<and>
+ eval_c i c'[z'::=V_var z2]\<^sub>v[z2::=V_var x]\<^sub>v b2" proof -
+
+ have "wfC P \<B> ?G (CE_val v == CE_val (V_lit l))" using wbc wfC_elims(7) assms subst_cv.simps subst_v_c_def by fastforce
+
+ moreover have "wfC P \<B> ?G (c'[z'::=V_var z2]\<^sub>v[z2::=V_var x]\<^sub>v)" proof(rule wfT_wfC_cons)
+ show \<open> P; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> z1 : b | CE_val v == CE_val (V_lit l) IMP (c[z::=V_var z1]\<^sub>v) \<rbrace> \<close> using assms subst_v_c_def by auto
+ have " \<lbrace> z2 : b | c'[z'::=V_var z2]\<^sub>v \<rbrace> = \<lbrace> z' : b | c' \<rbrace>" using assms subst_v_c_def by auto
+ thus \<open> P; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> z2 : b | c'[z'::=V_var z2]\<^sub>v \<rbrace> \<close> using assms subtype_elims by metis
+ show \<open>atom x \<sharp> (CE_val v == CE_val (V_lit l) IMP c[z::=V_var z1]\<^sub>v , c'[z'::=V_var z2]\<^sub>v, \<Gamma>)\<close> using xx fresh_Pair c.fresh by metis
+ qed
+
+ ultimately show ?thesis using wfI_wfC_eval_c a1 subst_v_c_def by simp
+ qed
+
+ ultimately show "is_satis i ?c" using is_satis_imp[OF *] by auto
+ qed
+ ultimately show ?thesis using valid.simps by simp
+ qed
+ moreover have "atom x \<sharp> (P, \<B>, \<Gamma>, z1 , CE_val v == CE_val (V_lit l) IMP c[z::=V_var z1]\<^sub>v , z2 , CE_val v == CE_val (V_lit l) IMP c'[z'::=V_var z2]\<^sub>v )"
+ unfolding fresh_prod5 \<tau>.fresh using xx xf2 fresh_prodN x_fresh_b by metis
+ ultimately show ?thesis using subtype_baseI assms xf2 by metis
+qed
+
+lemma eval_e_concat_eq:
+ assumes "wfI \<Theta> \<Gamma> i"
+ shows "\<exists>s. eval_e i (CE_val (V_lit (L_bitvec (v1 @ v2))) ) s \<and> eval_e i (CE_concat [(V_lit (L_bitvec v1))]\<^sup>c\<^sup>e [(V_lit (L_bitvec v2))]\<^sup>c\<^sup>e) s"
+ using eval_e_valI eval_e_concatI eval_v_litI eval_l.simps by metis
+
+lemma is_satis_eval_e_eq_imp:
+ assumes "wfI \<Theta> \<Gamma> i" and "eval_e i e1 s" and "eval_e i e2 s"
+ and "is_satis i (CE_val (V_var x) == e1)" (is "is_satis i ?c1")
+ shows "is_satis i (CE_val (V_var x) == e2)"
+proof -
+ have *:"eval_c i ?c1 True" using assms is_satis.simps by blast
+ hence "eval_e i (CE_val (V_var x)) s" using assms is_satis.simps eval_c_elims
+ by (metis (full_types) eval_e_uniqueness)
+ thus ?thesis using is_satis.simps eval_c.intros assms by fastforce
+qed
+
+lemma valid_eval_e_eq:
+ fixes e1::ce and e2::ce
+ assumes "\<forall>\<Gamma> i. wfI \<Theta> \<Gamma> i \<longrightarrow> (\<exists>s. eval_e i e1 s \<and> eval_e i e2 s)" and "\<Theta>; \<B>; GNil \<turnstile>\<^sub>w\<^sub>f e1 : b " and "\<Theta>; \<B>; GNil \<turnstile>\<^sub>w\<^sub>f e2 : b"
+ shows "\<Theta>; \<B>; (x, b, (CE_val (V_var x) == e1 )) #\<^sub>\<Gamma> GNil \<Turnstile> (CE_val (V_var x) == e2) "
+proof(rule validI)
+ show "\<Theta>; \<B>; (x, b, CE_val (V_var x) == e1 ) #\<^sub>\<Gamma> GNil \<turnstile>\<^sub>w\<^sub>f CE_val (V_var x) == e2"
+ proof
+ have " \<Theta> ; \<B> ; (x, b, TRUE )#\<^sub>\<Gamma>GNil \<turnstile>\<^sub>w\<^sub>f CE_val (V_var x) == e1" using assms wfC_eqI wfE_valI wfV_varI wfX_wfY
+ by (simp add: fresh_GNil wfC_e_eq)
+ hence "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f (x, b, CE_val (V_var x) == e1 ) #\<^sub>\<Gamma> GNil" using wfG_consI fresh_GNil wfX_wfY assms wfX_wfB by metis
+ thus "\<Theta>; \<B>; (x, b, CE_val (V_var x) == e1 ) #\<^sub>\<Gamma> GNil \<turnstile>\<^sub>w\<^sub>f CE_val (V_var x) : b " using wfCE_valI wfV_varI wfX_wfY
+ lookup.simps assms wfX_wfY by simp
+ show "\<Theta>; \<B>; (x, b, CE_val (V_var x) == e1 ) #\<^sub>\<Gamma> GNil \<turnstile>\<^sub>w\<^sub>f e2 : b " using assms wf_weakening wfX_wfY
+ by (metis (full_types) \<open>\<Theta>; \<B>; (x, b, CE_val (V_var x) == e1 ) #\<^sub>\<Gamma> GNil \<turnstile>\<^sub>w\<^sub>f CE_val (V_var x) : b\<close> empty_iff subsetI toSet.simps(1))
+ qed
+ show " \<forall>i. wfI \<Theta> ((x, b, CE_val (V_var x) == e1 ) #\<^sub>\<Gamma> GNil) i \<and> is_satis_g i ((x, b, CE_val (V_var x) == e1 ) #\<^sub>\<Gamma> GNil) \<longrightarrow> is_satis i (CE_val (V_var x) == e2 )"
+ proof(rule,rule)
+ fix i
+ assume "wfI \<Theta> ((x, b, CE_val (V_var x) == e1 ) #\<^sub>\<Gamma> GNil) i \<and> is_satis_g i ((x, b, CE_val (V_var x) == e1 ) #\<^sub>\<Gamma> GNil)"
+ moreover then obtain s where "eval_e i e1 s \<and> eval_e i e2 s" using assms by auto
+ ultimately show "is_satis i (CE_val (V_var x) == e2 )" using assms is_satis_eval_e_eq_imp is_satis_g.simps by meson
+ qed
+qed
+
+lemma subtype_concat:
+ assumes " \<turnstile>\<^sub>w\<^sub>f \<Theta>"
+ shows "\<Theta>; \<B>; GNil \<turnstile> \<lbrace> z : B_bitvec | CE_val (V_var z) == CE_val (V_lit (L_bitvec (v1 @ v2))) \<rbrace> \<lesssim>
+ \<lbrace> z : B_bitvec | CE_val (V_var z) == CE_concat [(V_lit (L_bitvec v1))]\<^sup>c\<^sup>e [(V_lit (L_bitvec v2))]\<^sup>c\<^sup>e \<rbrace>" (is "\<Theta>; \<B>; GNil \<turnstile> ?t1 \<lesssim> ?t2")
+proof -
+ obtain x::x where x: "atom x \<sharp> (\<Theta>, \<B>, GNil, z , CE_val (V_var z) == CE_val (V_lit (L_bitvec (v1 @ v2))),
+ z , CE_val (V_var z) == CE_concat [V_lit (L_bitvec v1)]\<^sup>c\<^sup>e [V_lit (L_bitvec v2)]\<^sup>c\<^sup>e )"
+ (is "?xfree" )
+ using obtain_fresh by auto
+
+ have wb1: "\<Theta>; \<B>; GNil \<turnstile>\<^sub>w\<^sub>f CE_val (V_lit (L_bitvec (v1 @ v2))) : B_bitvec" using wfX_wfY wfCE_valI wfV_litI assms base_for_lit.simps wfG_nilI by metis
+ hence wb2: "\<Theta>; \<B>; GNil \<turnstile>\<^sub>w\<^sub>f CE_concat [(V_lit (L_bitvec v1))]\<^sup>c\<^sup>e [(V_lit (L_bitvec v2))]\<^sup>c\<^sup>e : B_bitvec"
+ using wfCE_concatI wfX_wfY wfV_litI base_for_lit.simps wfCE_valI by metis
+
+ show ?thesis proof
+ show " \<Theta>; \<B>; GNil \<turnstile>\<^sub>w\<^sub>f ?t1" using wfT_e_eq fresh_GNil wb1 wb2 by metis
+ show " \<Theta>; \<B>; GNil \<turnstile>\<^sub>w\<^sub>f ?t2" using wfT_e_eq fresh_GNil wb1 wb2 by metis
+ show ?xfree using x by auto
+ show "\<Theta>; \<B>; (x, B_bitvec, (CE_val (V_var z) == CE_val (V_lit (L_bitvec (v1 @ v2))) )[z::=V_var x]\<^sub>v) #\<^sub>\<Gamma>
+ GNil \<Turnstile> (CE_val (V_var z) == CE_concat [(V_lit (L_bitvec v1))]\<^sup>c\<^sup>e [(V_lit (L_bitvec v2))]\<^sup>c\<^sup>e )[z::=V_var x]\<^sub>v "
+ using valid_eval_e_eq eval_e_concat_eq wb1 wb2 subst_v_c_def by fastforce
+ qed
+qed
+
+lemma subtype_len:
+ assumes " \<turnstile>\<^sub>w\<^sub>f \<Theta>"
+ shows "\<Theta>; \<B>; GNil \<turnstile> \<lbrace> z' : B_int | CE_val (V_var z') == CE_val (V_lit (L_num (int (length v)))) \<rbrace> \<lesssim>
+ \<lbrace> z : B_int | CE_val (V_var z) == CE_len [(V_lit (L_bitvec v))]\<^sup>c\<^sup>e \<rbrace>" (is "\<Theta>; \<B>; GNil \<turnstile> ?t1 \<lesssim> ?t2")
+proof -
+
+ have *: "\<Theta> \<turnstile>\<^sub>w\<^sub>f [] \<and> \<Theta>; \<B>; GNil \<turnstile>\<^sub>w\<^sub>f []\<^sub>\<Delta> " using assms wfG_nilI wfD_emptyI wfPhi_emptyI by auto
+ obtain x::x where x: "atom x \<sharp> (\<Theta>, \<B>, GNil, z' , CE_val (V_var z') ==
+ CE_val (V_lit (L_num (int (length v)))) , z, CE_val (V_var z) == CE_len [(V_lit (L_bitvec v))]\<^sup>c\<^sup>e )"
+ (is "atom x \<sharp> ?F")
+ using obtain_fresh by metis
+ then show ?thesis proof
+ have "\<Theta> ; \<B> ; GNil \<turnstile>\<^sub>w\<^sub>f CE_val (V_lit (L_num (int (length v)))) : B_int"
+ using wfCE_valI * wfV_litI base_for_lit.simps
+ by (metis wfE_valI wfX_wfY)
+
+ thus "\<Theta>; \<B>; GNil \<turnstile>\<^sub>w\<^sub>f ?t1" using wfT_e_eq fresh_GNil by auto
+
+ have "\<Theta> ; \<B> ; GNil \<turnstile>\<^sub>w\<^sub>f CE_len [(V_lit (L_bitvec v))]\<^sup>c\<^sup>e : B_int"
+ using wfE_valI * wfV_litI base_for_lit.simps wfE_valI wfX_wfY
+ by (metis wfCE_lenI wfCE_valI)
+
+ thus "\<Theta>; \<B>; GNil \<turnstile>\<^sub>w\<^sub>f ?t2" using wfT_e_eq fresh_GNil by auto
+
+ show "\<Theta>; \<B>; (x, B_int, (CE_val (V_var z') == CE_val (V_lit (L_num (int (length v)))) )[z'::=V_var x]\<^sub>v) #\<^sub>\<Gamma> GNil \<Turnstile> (CE_val (V_var z) == CE_len [(V_lit (L_bitvec v))]\<^sup>c\<^sup>e )[z::=V_var x]\<^sub>v"
+ (is "\<Theta>; \<B>; ?G \<Turnstile> ?c" ) using valid_len assms subst_v_c_def by auto
+ qed
+qed
+
+lemma subtype_base_fresh:
+ assumes "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b | c \<rbrace>" and "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b | c' \<rbrace> " and
+ "atom z \<sharp> \<Gamma>" and "\<Theta>; \<B>; (z, b, c) #\<^sub>\<Gamma> \<Gamma> \<Turnstile> c'"
+ shows "\<Theta>; \<B>; \<Gamma> \<turnstile> \<lbrace> z : b | c \<rbrace> \<lesssim> \<lbrace> z : b | c' \<rbrace>"
+proof -
+ obtain x::x where *:"atom x \<sharp> ((\<Theta> , \<B> , z, c, z, c', \<Gamma>) , (\<Theta>, \<B>, \<Gamma>, \<lbrace> z : b | c \<rbrace>, \<lbrace> z : b | c' \<rbrace>))" using obtain_fresh by metis
+ moreover hence "atom x \<sharp> \<Gamma>" using fresh_Pair by auto
+ moreover hence "\<Theta>; \<B>; (x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<Turnstile> c'[z::=V_var x]\<^sub>v" using assms valid_flip_simple * subst_v_c_def by auto
+ ultimately show ?thesis using subtype_baseI assms \<tau>.fresh fresh_Pair by metis
+qed
+
+lemma subtype_bop_arith:
+ assumes "wfG \<Theta> \<B> \<Gamma>" and "(opp = Plus \<and> ll = (L_num (n1+n2))) \<or> (opp = LEq \<and> ll = ( if n1\<le>n2 then L_true else L_false))"
+ and "(opp = Plus \<longrightarrow> b = B_int) \<and> (opp = LEq \<longrightarrow> b = B_bool)"
+ shows "\<Theta>; \<B>; \<Gamma> \<turnstile> (\<lbrace> z : b | C_eq (CE_val (V_var z)) (CE_val (V_lit (ll))) \<rbrace>) \<lesssim>
+ \<lbrace> z : b | C_eq (CE_val (V_var z)) (CE_op opp [(V_lit (L_num n1))]\<^sup>c\<^sup>e [(V_lit (L_num n2))]\<^sup>c\<^sup>e) \<rbrace>" (is "\<Theta>; \<B>; \<Gamma> \<turnstile> ?T1 \<lesssim> ?T2")
+proof -
+ obtain x::x where xf: "atom x \<sharp> (z, CE_val (V_var z) == CE_val (V_lit (ll)) , z, CE_val (V_var z) == CE_op opp [(V_lit (L_num n1))]\<^sup>c\<^sup>e [(V_lit (L_num n2))]\<^sup>c\<^sup>e , \<Gamma>)"
+ using obtain_fresh by blast
+
+ have "\<Theta>; \<B>; \<Gamma> \<turnstile> (\<lbrace> x : b | C_eq (CE_val (V_var x)) (CE_val (V_lit (ll))) \<rbrace>) \<lesssim>
+ \<lbrace> x : b | C_eq (CE_val (V_var x)) (CE_op opp [(V_lit (L_num n1))]\<^sup>c\<^sup>e [(V_lit (L_num n2))]\<^sup>c\<^sup>e) \<rbrace>" (is "\<Theta>; \<B>; \<Gamma> \<turnstile> ?S1 \<lesssim> ?S2")
+ proof(rule subtype_base_fresh)
+
+ show "atom x \<sharp> \<Gamma>" using xf fresh_Pair by auto
+
+ show "wfT \<Theta> \<B> \<Gamma> (\<lbrace> x : b | CE_val (V_var x) == CE_val (V_lit ll) \<rbrace>)" (is "wfT \<Theta> \<B> ?A ?B")
+ proof(rule wfT_e_eq)
+ have " \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f (V_lit ll) : b" using wfV_litI base_for_lit.simps assms by metis
+ thus " \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_val (V_lit ll) : b" using wfCE_valI by auto
+ show "atom x \<sharp> \<Gamma>" using xf fresh_Pair by auto
+ qed
+
+ consider "opp = Plus" | "opp = LEq" using opp.exhaust assms by blast
+ then show "wfT \<Theta> \<B> \<Gamma> (\<lbrace> x : b | CE_val (V_var x) == CE_op opp [(V_lit (L_num n1))]\<^sup>c\<^sup>e [(V_lit (L_num n2))]\<^sup>c\<^sup>e \<rbrace>)" (is "wfT \<Theta> \<B> ?A ?C")
+ proof(cases)
+ case 1
+ then show "\<Theta> ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> x : b | [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ opp [ [ L_num n1 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ L_num n2 ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e \<rbrace>"
+ using wfCE_valI wfCE_plusI assms wfV_litI base_for_lit.simps assms
+ by (metis \<open>atom x \<sharp> \<Gamma>\<close> wfT_e_eq)
+ next
+ case 2
+ then show "\<Theta> ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> x : b | [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ opp [ [ L_num n1 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ L_num n2 ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e \<rbrace> "
+ using wfCE_valI wfCE_plusI assms wfV_litI base_for_lit.simps assms
+
+ by (metis \<open>atom x \<sharp> \<Gamma>\<close> wfCE_leqI wfT_e_eq)
+ qed
+
+ show "\<Theta>; \<B> ; (x, b, (CE_val (V_var x) == CE_val (V_lit (ll)) )) #\<^sub>\<Gamma> \<Gamma>
+ \<Turnstile> (CE_val (V_var x) == CE_op opp [V_lit (L_num n1)]\<^sup>c\<^sup>e [V_lit (L_num n2)]\<^sup>c\<^sup>e)" (is "\<Theta>; \<B>; ?G \<Turnstile> ?c")
+ using valid_arith_bop assms xf by simp
+
+ qed
+ moreover have "?S1 = ?T1 " using type_l_eq by auto
+ moreover have "?S2 = ?T2" using type_e_eq ce.fresh v.fresh supp_l_empty fresh_def empty_iff fresh_e_opp
+ by (metis ms_fresh_all(4))
+ ultimately show ?thesis by auto
+
+qed
+
+lemma subtype_bop_eq:
+ assumes "wfG \<Theta> \<B> \<Gamma>" and "base_for_lit l1 = base_for_lit l2"
+ shows "\<Theta>; \<B>; \<Gamma> \<turnstile> (\<lbrace> z : B_bool | C_eq (CE_val (V_var z)) (CE_val (V_lit (if l1 = l2 then L_true else L_false))) \<rbrace>) \<lesssim>
+ \<lbrace> z : B_bool | C_eq (CE_val (V_var z)) (CE_op Eq [(V_lit l1)]\<^sup>c\<^sup>e [(V_lit l2)]\<^sup>c\<^sup>e) \<rbrace>" (is "\<Theta>; \<B>; \<Gamma> \<turnstile> ?T1 \<lesssim> ?T2")
+proof -
+ let ?ll = "if l1 = l2 then L_true else L_false"
+ obtain x::x where xf: "atom x \<sharp> (z, CE_val (V_var z) == CE_val (V_lit (if l1 = l2 then L_true else L_false)) , z, CE_val (V_var z) == CE_op Eq [(V_lit l1)]\<^sup>c\<^sup>e [(V_lit l2)]\<^sup>c\<^sup>e , \<Gamma>, (\<Theta>, \<B>, \<Gamma>))"
+ using obtain_fresh by blast
+
+ have "\<Theta>; \<B>; \<Gamma> \<turnstile> (\<lbrace> x : B_bool | C_eq (CE_val (V_var x)) (CE_val (V_lit (?ll))) \<rbrace>) \<lesssim>
+ \<lbrace> x : B_bool | C_eq (CE_val (V_var x)) (CE_op Eq [(V_lit (l1))]\<^sup>c\<^sup>e [(V_lit (l2))]\<^sup>c\<^sup>e) \<rbrace>" (is "\<Theta>; \<B>; \<Gamma> \<turnstile> ?S1 \<lesssim> ?S2")
+ proof(rule subtype_base_fresh)
+
+ show "atom x \<sharp> \<Gamma>" using xf fresh_Pair by auto
+
+ show "wfT \<Theta> \<B> \<Gamma> (\<lbrace> x : B_bool | CE_val (V_var x) == CE_val (V_lit ?ll) \<rbrace>)" (is "wfT \<Theta> \<B> ?A ?B")
+ proof(rule wfT_e_eq)
+ have " \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f (V_lit ?ll) : B_bool" using wfV_litI base_for_lit.simps assms by metis
+ thus " \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_val (V_lit ?ll) : B_bool" using wfCE_valI by auto
+ show "atom x \<sharp> \<Gamma>" using xf fresh_Pair by auto
+ qed
+
+ show " \<Theta> ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> x : B_bool | [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ eq [ [ l1 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ l2 ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e \<rbrace> "
+ proof(rule wfT_e_eq)
+ show "\<Theta> ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f [ eq [ [ l1 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ l2 ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e : B_bool"
+ apply(rule wfCE_eqI, rule wfCE_valI)
+ apply(rule wfV_litI, simp add: assms)
+ using wfV_litI assms wfCE_valI by auto
+ show "atom x \<sharp> \<Gamma>" using xf fresh_Pair by auto
+ qed
+
+ show "\<Theta>; \<B> ; (x, B_bool, (CE_val (V_var x) == CE_val (V_lit (?ll)) )) #\<^sub>\<Gamma> \<Gamma>
+ \<Turnstile> (CE_val (V_var x) == CE_op Eq [V_lit (l1)]\<^sup>c\<^sup>e [V_lit (l2)]\<^sup>c\<^sup>e)" (is "\<Theta>; \<B>; ?G \<Turnstile> ?c")
+ using valid_eq_bop assms xf by auto
+
+ qed
+ moreover have "?S1 = ?T1 " using type_l_eq by auto
+ moreover have "?S2 = ?T2" using type_e_eq ce.fresh v.fresh supp_l_empty fresh_def empty_iff fresh_e_opp
+ by (metis ms_fresh_all(4))
+ ultimately show ?thesis by auto
+
+qed
+
+lemma subtype_top:
+ assumes "wfT \<Theta> \<B> G (\<lbrace> z : b | c \<rbrace>)"
+ shows "\<Theta>; \<B>; G \<turnstile> (\<lbrace> z : b | c \<rbrace>) \<lesssim> (\<lbrace> z : b | TRUE \<rbrace>)"
+proof -
+ obtain x::x where *: "atom x \<sharp> (\<Theta>, \<B>, G, z , c, z , TRUE)" using obtain_fresh by blast
+ then show ?thesis proof(rule subtype_baseI)
+ show \<open> \<Theta>; \<B>; G \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b | c \<rbrace> \<close> using assms by auto
+ show \<open> \<Theta>; \<B>;G \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b | TRUE \<rbrace> \<close> using wfT_TRUE assms wfX_wfY b_of.simps wfT_wf
+ by (metis wfX_wfB(8))
+ hence "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f (x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\<Gamma> G" using wfT_wf_cons3 assms fresh_Pair * subst_v_c_def by auto
+ thus \<open>\<Theta>; \<B>; (x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\<Gamma> G \<Turnstile> (TRUE)[z::=V_var x]\<^sub>v \<close> using valid_trueI subst_cv.simps subst_v_c_def by metis
+ qed
+qed
+
+lemma if_simp:
+ "(if x = x then e1 else e2) = e1"
+ by auto
+
+lemma subtype_split:
+ assumes "split n v (v1,v2)" and "\<turnstile>\<^sub>w\<^sub>f \<Theta>"
+ shows "\<Theta> ; {||} ; GNil \<turnstile> \<lbrace> z : [ B_bitvec , B_bitvec ]\<^sup>b | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ [ L_bitvec
+ v1 ]\<^sup>v , [ L_bitvec
+ v2 ]\<^sup>v ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrace> \<lesssim> \<lbrace> z : [ B_bitvec , B_bitvec ]\<^sup>b | [ [ L_bitvec
+ v ]\<^sup>v ]\<^sup>c\<^sup>e == [ [#1[ [ z ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e @@ [#2[ [ z ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e AND [| [#1[ [ z ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e == [ [ L_num
+ n ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrace>"
+ (is "\<Theta> ;?B ; GNil \<turnstile> \<lbrace> z : [ B_bitvec , B_bitvec ]\<^sup>b | ?c1 \<rbrace> \<lesssim> \<lbrace> z : [ B_bitvec , B_bitvec ]\<^sup>b | ?c2 \<rbrace>")
+proof -
+ obtain x::x where xf:"atom x \<sharp> (\<Theta>, ?B, GNil, z, ?c1,z, ?c2 )" using obtain_fresh by auto
+ then show ?thesis proof(rule subtype_baseI)
+ show *: \<open>\<Theta> ; ?B ; (x, [ B_bitvec , B_bitvec ]\<^sup>b, (?c1)[z::=[ x ]\<^sup>v]\<^sub>v) #\<^sub>\<Gamma>
+ GNil \<Turnstile> (?c2)[z::=[ x ]\<^sup>v]\<^sub>v \<close>
+ unfolding subst_v_c_def subst_cv.simps subst_cev.simps subst_vv.simps if_simp
+ using valid_split[OF assms, of x] by simp
+ show \<open> \<Theta> ; ?B ; GNil \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : [ B_bitvec , B_bitvec ]\<^sup>b| ?c1 \<rbrace> \<close> using valid_wfT[OF *] xf fresh_prodN by metis
+ show \<open> \<Theta> ; ?B ; GNil \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : [ B_bitvec , B_bitvec ]\<^sup>b| ?c2 \<rbrace> \<close> using valid_wfT[OF *] xf fresh_prodN by metis
+ qed
+qed
+
+lemma subtype_range:
+ fixes n::int and \<Gamma>::\<Gamma>
+ assumes "0 \<le> n \<and> n \<le> int (length v)" and "\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f \<Gamma>"
+ shows "\<Theta> ; {||} ; \<Gamma> \<turnstile> \<lbrace> z : B_int | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_num n ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrace> \<lesssim>
+ \<lbrace> z : B_int | ([ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e ) AND (
+ [ leq [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e [| [ [ L_bitvec v ]\<^sup>v ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e) \<rbrace>"
+ (is "\<Theta> ; ?B ; \<Gamma> \<turnstile> \<lbrace> z : B_int | ?c1 \<rbrace> \<lesssim> \<lbrace> z : B_int | ?c2 AND ?c3 \<rbrace>")
+proof -
+ obtain x::x where *:\<open>atom x \<sharp> (\<Theta>, ?B, \<Gamma>, z, ?c1 , z, ?c2 AND ?c3)\<close> using obtain_fresh by auto
+ moreover have **:\<open>\<Theta> ; ?B ; (x, B_int, (?c1)[z::=[ x ]\<^sup>v]\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<Turnstile> (?c2 AND ?c3)[z::=[ x ]\<^sup>v]\<^sub>v\<close>
+ unfolding subst_v_c_def subst_cv.simps subst_cev.simps subst_vv.simps if_simp using valid_range_length[OF assms(1)] assms fresh_prodN * by simp
+
+ moreover hence \<open> \<Theta> ; ?B ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : B_int | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_num n ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrace> \<close> using
+ valid_wfT * fresh_prodN by metis
+ moreover have \<open> \<Theta> ; ?B ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : B_int | ?c2 AND ?c3 \<rbrace> \<close>
+ using valid_wfT[OF **] * fresh_prodN by metis
+ ultimately show ?thesis using subtype_baseI by auto
+qed
+
+lemma check_num_range:
+ assumes "0 \<le> n \<and> n \<le> int (length v)" and "\<turnstile>\<^sub>w\<^sub>f \<Theta>"
+ shows "\<Theta> ; {||} ; GNil \<turnstile> ([ L_num n ]\<^sup>v) \<Leftarrow> \<lbrace> z : B_int | ([ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e) AND
+ [ leq [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e [| [ [ L_bitvec v ]\<^sup>v ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrace>"
+ using assms subtype_range check_v.intros infer_v_litI wfG_nilI
+ by (meson infer_natI)
+
+section \<open>Literals\<close>
+
+nominal_function type_for_lit :: "l \<Rightarrow> \<tau>" where
+ "type_for_lit (L_true) = (\<lbrace> z : B_bool | [[z]\<^sup>v]\<^sup>c\<^sup>e == [V_lit L_true]\<^sup>c\<^sup>e \<rbrace>)"
+| "type_for_lit (L_false) = (\<lbrace> z : B_bool | [[z]\<^sup>v]\<^sup>c\<^sup>e == [V_lit L_false]\<^sup>c\<^sup>e \<rbrace>)"
+| "type_for_lit (L_num n) = (\<lbrace> z : B_int | [[z]\<^sup>v]\<^sup>c\<^sup>e == [V_lit (L_num n)]\<^sup>c\<^sup>e \<rbrace>)"
+| "type_for_lit (L_unit) = (\<lbrace> z : B_unit | [[z]\<^sup>v]\<^sup>c\<^sup>e == [V_lit (L_unit )]\<^sup>c\<^sup>e \<rbrace>)" (* could have done { z : unit | True } but want the uniformity *)
+| "type_for_lit (L_bitvec v) = (\<lbrace> z : B_bitvec | [[z]\<^sup>v]\<^sup>c\<^sup>e == [V_lit (L_bitvec v)]\<^sup>c\<^sup>e \<rbrace>)"
+ by (auto simp: eqvt_def type_for_lit_graph_aux_def, metis l.strong_exhaust,(simp add: permute_int_def flip_bitvec0)+ )
+nominal_termination (eqvt) by lexicographic_order
+
+
+nominal_function type_for_var :: "\<Gamma> \<Rightarrow> \<tau> \<Rightarrow> x \<Rightarrow> \<tau>" where
+ "type_for_var G \<tau> x = (case lookup G x of
+ None \<Rightarrow> \<tau>
+ | Some (b,c) \<Rightarrow> (\<lbrace> x : b | c \<rbrace>)) "
+ apply auto unfolding eqvt_def apply(rule allI) unfolding type_for_var_graph_aux_def eqvt_def by simp
+nominal_termination (eqvt) by lexicographic_order
+
+lemma infer_l_form:
+ fixes l::l and tm::"'a::fs"
+ assumes "\<turnstile> l \<Rightarrow> \<tau>"
+ shows "\<exists>z b. \<tau> = (\<lbrace> z : b | C_eq (CE_val (V_var z)) (CE_val (V_lit l)) \<rbrace>) \<and> atom z \<sharp> tm"
+proof -
+ obtain z' and b where t:"\<tau> = (\<lbrace> z' : b | C_eq (CE_val (V_var z')) (CE_val (V_lit l)) \<rbrace>)" using infer_l_elims assms using infer_l.simps type_for_lit.simps
+ type_for_lit.cases by blast
+ obtain z::x where zf: "atom z \<sharp> tm" using obtain_fresh by metis
+ have "\<tau> = \<lbrace> z : b | C_eq (CE_val (V_var z)) (CE_val (V_lit l)) \<rbrace>" using type_e_eq ce.fresh v.fresh l.fresh
+ by (metis t type_l_eq)
+ thus ?thesis using zf by auto
+qed
+
+lemma infer_l_form3:
+ fixes l::l
+ assumes "\<turnstile> l \<Rightarrow> \<tau>"
+ shows "\<exists>z. \<tau> = (\<lbrace> z : base_for_lit l | C_eq (CE_val (V_var z)) (CE_val (V_lit l)) \<rbrace>)"
+ using infer_l_elims using assms using infer_l.simps type_for_lit.simps base_for_lit.simps by auto
+
+lemma infer_l_form4[simp]:
+ fixes \<Gamma>::\<Gamma>
+ assumes "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> "
+ shows "\<exists>z. \<turnstile> l \<Rightarrow> (\<lbrace> z : base_for_lit l | C_eq (CE_val (V_var z)) (CE_val (V_lit l)) \<rbrace>)"
+ using assms infer_l_form2 infer_l_form3 by metis
+
+lemma infer_v_unit_form:
+ fixes v::v
+ assumes "P ; \<B> ; \<Gamma> \<turnstile> v \<Rightarrow> (\<lbrace> z1 : B_unit | c1 \<rbrace>)" and "supp v = {}"
+ shows "v = V_lit L_unit"
+ using assms proof(nominal_induct \<Gamma> v "\<lbrace> z1 : B_unit | c1 \<rbrace>" rule: infer_v.strong_induct)
+ case (infer_v_varI \<Theta> \<B> c x z)
+ then show ?case using supp_at_base by auto
+next
+ case (infer_v_litI \<Theta> \<B> \<Gamma> l)
+ from \<open>\<turnstile> l \<Rightarrow> \<lbrace> z1 : B_unit | c1 \<rbrace>\<close> show ?case by(nominal_induct "\<lbrace> z1 : B_unit | c1 \<rbrace>" rule: infer_l.strong_induct,auto)
+qed
+
+lemma base_for_lit_wf:
+ assumes "\<turnstile>\<^sub>w\<^sub>f \<Theta>"
+ shows "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f base_for_lit l"
+ using base_for_lit.simps using wfV_elims wf_intros assms l.exhaust by metis
+
+lemma infer_l_t_wf:
+ fixes \<Gamma>::\<Gamma>
+ assumes "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<and> atom z \<sharp> \<Gamma>"
+ shows "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : base_for_lit l | C_eq (CE_val (V_var z)) (CE_val (V_lit l)) \<rbrace>"
+proof
+ show "atom z \<sharp> (\<Theta>, \<B>, \<Gamma>)" using wfG_fresh_x assms by auto
+ show "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f base_for_lit l" using base_for_lit_wf assms wfX_wfY by metis
+ thus "\<Theta>; \<B>; (z, base_for_lit l, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_val (V_var z) == CE_val (V_lit l)" using wfC_v_eq wfV_litI assms wfX_wfY by metis
+qed
+
+lemma infer_l_wf:
+ fixes l::l and \<Gamma>::\<Gamma> and \<tau>::\<tau> and \<Theta>::\<Theta>
+ assumes "\<turnstile> l \<Rightarrow> \<tau>" and "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>"
+ shows "\<turnstile>\<^sub>w\<^sub>f \<Theta>" and "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>" and "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau>"
+proof -
+ show *:"\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>" using assms infer_l_elims by auto
+ thus "\<turnstile>\<^sub>w\<^sub>f \<Theta>" using wfX_wfY by auto
+ show *:"\<Theta> ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau>" using infer_l_t_wf assms infer_l_form3 *
+ by (metis \<open>\<turnstile>\<^sub>w\<^sub>f \<Theta>\<close> fresh_GNil wfG_nilI wfT_weakening_nil)
+qed
+
+lemma infer_l_uniqueness:
+ fixes l::l
+ assumes "\<turnstile> l \<Rightarrow> \<tau>" and "\<turnstile> l \<Rightarrow> \<tau>'"
+ shows "\<tau> = \<tau>'"
+ using assms
+proof -
+ obtain z and b where zt: "\<tau> = (\<lbrace> z : b | C_eq (CE_val (V_var z)) (CE_val (V_lit l)) \<rbrace>)" using infer_l_form assms by blast
+ obtain z' and b where z't: "\<tau>' = (\<lbrace> z' : b | C_eq (CE_val (V_var z')) (CE_val (V_lit l)) \<rbrace>)" using infer_l_form assms by blast
+ thus ?thesis using type_l_eq zt z't assms infer_l.simps infer_l_elims l.distinct
+ by (metis infer_l_form3)
+qed
+
+section \<open>Values\<close>
+
+lemma type_v_eq:
+ assumes "\<lbrace> z1 : b1 | c1 \<rbrace> = \<lbrace> z : b | C_eq (CE_val (V_var z)) (CE_val (V_var x)) \<rbrace>" and "atom z \<sharp> x"
+ shows "b = b1" and "c1 = C_eq (CE_val (V_var z1)) (CE_val (V_var x))"
+ using assms by (auto,metis Abs1_eq_iff \<tau>.eq_iff assms c.fresh ce.fresh type_e_eq v.fresh)
+
+lemma infer_var2 [elim]:
+ assumes "P; \<B> ; G \<turnstile> V_var x \<Rightarrow> \<tau>"
+ shows "\<exists>b c. Some (b,c) = lookup G x"
+ using assms infer_v_elims lookup_iff by (metis (no_types, lifting))
+
+lemma infer_var3 [elim]:
+ assumes "\<Theta>; \<B>; \<Gamma> \<turnstile> V_var x \<Rightarrow> \<tau>"
+ shows "\<exists>z b c. Some (b,c) = lookup \<Gamma> x \<and> \<tau> = (\<lbrace> z : b | C_eq (CE_val (V_var z)) (CE_val (V_var x)) \<rbrace>) \<and> atom z \<sharp> x \<and> atom z \<sharp> (\<Theta>, \<B>, \<Gamma>)"
+ using infer_v_elims(1)[OF assms(1)] by metis
+
+lemma infer_bool_options2:
+ fixes v::v
+ assumes "\<Theta>; \<B>; \<Gamma> \<turnstile> v \<Rightarrow> \<lbrace> z : b | c \<rbrace>" and "supp v = {} \<and> b = B_bool"
+ shows "v = V_lit L_true \<or> (v = (V_lit L_false))"
+ using assms
+proof(nominal_induct "\<lbrace> z : b | c \<rbrace>" rule: infer_v.strong_induct)
+ case (infer_v_varI \<Theta> \<B> \<Gamma> c x z)
+ then show ?case using v.supp supp_at_base by auto
+next
+ case (infer_v_litI \<Theta> \<B> \<Gamma> l)
+ from \<open> \<turnstile> l \<Rightarrow> \<lbrace> z : b | c \<rbrace>\<close> show ?case proof(nominal_induct "\<lbrace> z : b | c \<rbrace>" rule: infer_l.strong_induct)
+ case (infer_trueI z)
+ then show ?case by auto
+ next
+ case (infer_falseI z)
+ then show ?case by auto
+ next
+ case (infer_natI n z)
+ then show ?case using infer_v_litI by simp
+ next
+ case (infer_unitI z)
+ then show ?case using infer_v_litI by simp
+ next
+ case (infer_bitvecI bv z)
+ then show ?case using infer_v_litI by simp
+ qed
+qed(auto+)
+
+lemma infer_bool_options:
+ fixes v::v
+ assumes "\<Theta>; \<B>; \<Gamma> \<turnstile> v \<Rightarrow> \<lbrace> z : B_bool | c \<rbrace>" and "supp v = {}"
+ shows "v = V_lit L_true \<or> (v = (V_lit L_false))"
+ using infer_bool_options2 assms by blast
+
+lemma infer_int2:
+ fixes v::v
+ assumes "\<Theta>; \<B>; \<Gamma> \<turnstile> v \<Rightarrow> \<lbrace> z : b | c \<rbrace>"
+ shows "supp v = {} \<and> b = B_int \<longrightarrow> (\<exists>n. v = V_lit (L_num n))"
+ using assms
+proof(nominal_induct "\<lbrace> z : b | c \<rbrace>" rule: infer_v.strong_induct)
+ case (infer_v_varI \<Theta> \<B> \<Gamma> c x z)
+ then show ?case using v.supp supp_at_base by auto
+next
+ case (infer_v_litI \<Theta> \<B> \<Gamma> l)
+ from \<open> \<turnstile> l \<Rightarrow> \<lbrace> z : b | c \<rbrace>\<close> show ?case proof(nominal_induct "\<lbrace> z : b | c \<rbrace>" rule: infer_l.strong_induct)
+ case (infer_trueI z)
+ then show ?case by auto
+ next
+ case (infer_falseI z)
+ then show ?case by auto
+ next
+ case (infer_natI n z)
+ then show ?case using infer_v_litI by simp
+ next
+ case (infer_unitI z)
+ then show ?case using infer_v_litI by simp
+ next
+ case (infer_bitvecI bv z)
+ then show ?case using infer_v_litI by simp
+ qed
+qed(auto+)
+
+lemma infer_bitvec:
+ fixes \<Theta>::\<Theta> and v::v
+ assumes "\<Theta>; \<B>; \<Gamma> \<turnstile> v \<Rightarrow> \<lbrace> z' : B_bitvec | c' \<rbrace>" and "supp v = {}"
+ shows "\<exists>bv. v = V_lit (L_bitvec bv)"
+ using assms proof(nominal_induct v rule: v.strong_induct)
+ case (V_lit l)
+ then show ?case by(nominal_induct l rule: l.strong_induct,force+)
+next
+ case (V_consp s dc b v)
+ then show ?case using infer_v_elims(7)[OF V_consp(2)] \<tau>.eq_iff by auto
+next
+ case (V_var x)
+ then show ?case using supp_at_base by auto
+qed(force+)
+
+lemma infer_int:
+ assumes "infer_v \<Theta> \<B> \<Gamma> v (\<lbrace> z : B_int | c \<rbrace>)" and "supp v= {}"
+ shows "\<exists>n. V_lit (L_num n) = v"
+ using assms infer_int2 by (metis (no_types, lifting))
+
+lemma infer_lit:
+ assumes "infer_v \<Theta> \<B> \<Gamma> v (\<lbrace> z : b | c \<rbrace>)" and "supp v= {}" and "b \<in> { B_bool , B_int , B_unit }"
+ shows "\<exists>l. V_lit l = v"
+ using assms proof(nominal_induct v rule: v.strong_induct)
+ case (V_lit x)
+ then show ?case by (simp add: supp_at_base)
+next
+ case (V_var x)
+ then show ?case
+ by (simp add: supp_at_base)
+next
+ case (V_pair x1a x2a)
+ then show ?case using supp_at_base by auto
+next
+ case (V_cons x1a x2a x3)
+ then show ?case using supp_at_base by auto
+next
+ case (V_consp x1a x2a x3 x4)
+ then show ?case using supp_at_base by auto
+qed
+
+lemma infer_v_form[simp]:
+ fixes v::v
+ assumes "\<Theta>; \<B>; \<Gamma> \<turnstile> v \<Rightarrow> \<tau>"
+ shows "\<exists>z b. \<tau> = (\<lbrace> z : b | C_eq (CE_val (V_var z)) (CE_val v)\<rbrace>) \<and> atom z \<sharp> v \<and> atom z \<sharp> (\<Theta>, \<B>, \<Gamma>)"
+ using assms
+proof(nominal_induct rule: infer_v.strong_induct)
+ case (infer_v_varI \<Theta> \<B> \<Gamma> b c x z)
+ then show ?case by force
+next
+ case (infer_v_litI \<Theta> \<B> \<Gamma> l \<tau>)
+ then obtain z and b where "\<tau> = \<lbrace> z : b | CE_val (V_var z) == CE_val (V_lit l) \<rbrace> \<and>atom z \<sharp> (\<Theta>, \<B>, \<Gamma>) "
+ using infer_l_form by metis
+ moreover hence "atom z \<sharp> (V_lit l)" using supp_l_empty v.fresh(1) fresh_prod2 fresh_def by blast
+ ultimately show ?case by metis
+next
+ case (infer_v_pairI z v1 v2 \<Theta> \<B> \<Gamma> t1 t2)
+ then show ?case by force
+next
+ case (infer_v_consI s dclist \<Theta> dc tc \<B> \<Gamma> v tv z)
+ moreover hence "atom z \<sharp> (V_cons s dc v)" using
+ Un_commute b.supp(3) fresh_def sup_bot.right_neutral supp_b_empty v.supp(4) pure_supp by metis
+ ultimately show ?case using fresh_prodN by metis
+next
+ case (infer_v_conspI s bv dclist \<Theta> dc tc \<B> \<Gamma> v tv b z)
+ moreover hence "atom z \<sharp> (V_consp s dc b v)" unfolding v.fresh using pure_fresh fresh_prodN * by metis
+ ultimately show ?case using fresh_prodN by metis
+qed
+
+lemma infer_v_form2:
+ fixes v::v
+ assumes "\<Theta>; \<B>; \<Gamma> \<turnstile> v \<Rightarrow> (\<lbrace> z : b | c \<rbrace>)" and "atom z \<sharp> v"
+ shows "c = C_eq (CE_val (V_var z)) (CE_val v)"
+ using assms
+proof -
+ obtain z' and b' where "(\<lbrace> z : b | c \<rbrace>) = (\<lbrace> z' : b' | CE_val (V_var z') == CE_val v \<rbrace>) \<and> atom z' \<sharp> v"
+ using infer_v_form assms by meson
+ thus ?thesis using Abs1_eq_iff(3) \<tau>.eq_iff type_e_eq
+ by (metis assms(2) ce.fresh(1))
+qed
+
+lemma infer_v_form3:
+ fixes v::v
+ assumes "\<Theta>; \<B>; \<Gamma> \<turnstile> v \<Rightarrow> \<tau>" and "atom z \<sharp> (v,\<Gamma>)"
+ shows "\<Theta>; \<B>; \<Gamma> \<turnstile> v \<Rightarrow> \<lbrace> z : b_of \<tau> | C_eq (CE_val (V_var z)) (CE_val v)\<rbrace>"
+proof -
+ obtain z' and b' where "\<tau> = \<lbrace> z' : b' | C_eq (CE_val (V_var z')) (CE_val v)\<rbrace> \<and> atom z' \<sharp> v \<and> atom z' \<sharp> (\<Theta>, \<B>, \<Gamma>)"
+ using infer_v_form assms by metis
+ moreover hence "\<lbrace> z' : b' | C_eq (CE_val (V_var z')) (CE_val v)\<rbrace> = \<lbrace> z : b' | C_eq (CE_val (V_var z)) (CE_val v)\<rbrace>"
+ using assms type_e_eq fresh_Pair ce.fresh by auto
+ ultimately show ?thesis using b_of.simps assms by auto
+qed
+
+lemma infer_v_form4:
+ fixes v::v
+ assumes "\<Theta>; \<B>; \<Gamma> \<turnstile> v \<Rightarrow> \<tau>" and "atom z \<sharp> (v,\<Gamma>)" and "b = b_of \<tau>"
+ shows "\<Theta>; \<B>; \<Gamma> \<turnstile> v \<Rightarrow> \<lbrace> z : b | C_eq (CE_val (V_var z)) (CE_val v)\<rbrace>"
+ using assms infer_v_form3 by simp
+
+lemma infer_v_v_wf:
+ fixes v::v
+ shows "\<Theta>; \<B> ; G \<turnstile> v \<Rightarrow> \<tau> \<Longrightarrow> \<Theta>; \<B>; G \<turnstile>\<^sub>w\<^sub>f v : (b_of \<tau>)"
+proof(induct rule: infer_v.induct)
+ case (infer_v_varI \<Theta> \<B> \<Gamma> b c x z)
+ then show ?case using wfC_elims wf_intros by auto
+next
+ case (infer_v_pairI z v1 v2 \<Theta> \<B> \<Gamma> t1 t2)
+ then show ?case using wfC_elims wf_intros by auto
+next
+ case (infer_v_litI \<Theta> \<B> \<Gamma> l \<tau>)
+ hence "b_of \<tau> = base_for_lit l" using infer_l_form3 b_of.simps by metis
+ then show ?case using wfV_litI infer_l_wf infer_v_litI wfG_b_weakening
+ by (metis fempty_fsubsetI)
+next
+ case (infer_v_consI s dclist \<Theta> dc tc \<B> \<Gamma> v tv z)
+ then show ?case using wfC_elims wf_intros
+ by (metis (no_types, lifting) b_of.simps has_fresh_z2 subtype_eq_base2)
+next
+ case (infer_v_conspI s bv dclist \<Theta> dc tc \<B> \<Gamma> v tv b z)
+ obtain z1 b1 c1 where t:"tc = \<lbrace> z1 : b1 | c1 \<rbrace>" using obtain_fresh_z by metis
+ show ?case unfolding b_of.simps proof(rule wfV_conspI)
+ show \<open>AF_typedef_poly s bv dclist \<in> set \<Theta>\<close> using infer_v_conspI by auto
+ show \<open>(dc, \<lbrace> z1 : b1 | c1 \<rbrace> ) \<in> set dclist\<close> using infer_v_conspI t by auto
+ show \<open> \<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f b \<close> using infer_v_conspI by auto
+ show \<open>atom bv \<sharp> (\<Theta>, \<B>, \<Gamma>, b, v)\<close> using infer_v_conspI by auto
+ have " b1[bv::=b]\<^sub>b\<^sub>b = b_of tv" using subtype_eq_base2[OF infer_v_conspI(5)] b_of.simps t subst_tb.simps by auto
+ thus \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b1[bv::=b]\<^sub>b\<^sub>b \<close> using infer_v_conspI by auto
+ qed
+qed
+
+lemma infer_v_t_form_wf:
+ assumes " wfB \<Theta> \<B> b" and "wfV \<Theta> \<B> \<Gamma> v b" and "atom z \<sharp> \<Gamma>"
+ shows "wfT \<Theta> \<B> \<Gamma> \<lbrace> z : b | C_eq (CE_val (V_var z)) (CE_val v)\<rbrace>"
+ using wfT_v_eq assms by auto
+
+lemma infer_v_t_wf:
+ fixes v::v
+ assumes "\<Theta>; \<B>; G \<turnstile> v \<Rightarrow> \<tau>"
+ shows "wfT \<Theta> \<B> G \<tau> \<and> wfB \<Theta> \<B> (b_of \<tau>) "
+proof -
+ obtain z and b where "\<tau> = \<lbrace> z : b | CE_val (V_var z) == CE_val v \<rbrace> \<and> atom z \<sharp> v \<and> atom z \<sharp> (\<Theta>, \<B>, G)" using infer_v_form assms by metis
+ moreover have "wfB \<Theta> \<B> b" using infer_v_v_wf b_of.simps wfX_wfB(1) assms
+ using calculation by fastforce
+ ultimately show "wfT \<Theta> \<B> G \<tau> \<and> wfB \<Theta> \<B> (b_of \<tau>)" using infer_v_v_wf infer_v_t_form_wf assms by fastforce
+qed
+
+lemma infer_v_wf:
+ fixes v::v
+ assumes "\<Theta>; \<B>; G \<turnstile> v \<Rightarrow> \<tau>"
+ shows "\<Theta>; \<B>; G \<turnstile>\<^sub>w\<^sub>f v : (b_of \<tau>)" and "wfT \<Theta> \<B> G \<tau>" and "wfTh \<Theta>" and "wfG \<Theta> \<B> G"
+proof -
+ show "\<Theta>; \<B>; G \<turnstile>\<^sub>w\<^sub>f v : b_of \<tau> " using infer_v_v_wf assms by auto
+ show "\<Theta>; \<B>; G \<turnstile>\<^sub>w\<^sub>f \<tau>" using infer_v_t_wf assms by auto
+ thus "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f G" using wfX_wfY by auto
+ thus "\<turnstile>\<^sub>w\<^sub>f \<Theta>" using wfX_wfY by auto
+qed
+
+lemma check_bool_options:
+ assumes "\<Theta>; \<B>; \<Gamma> \<turnstile> v \<Leftarrow> \<lbrace> z : B_bool | TRUE \<rbrace>" and "supp v = {}"
+ shows "v = V_lit L_true \<or> v = V_lit L_false"
+proof -
+ obtain t1 where " \<Theta>; \<B>; \<Gamma> \<turnstile> t1 \<lesssim> \<lbrace> z : B_bool | TRUE \<rbrace> \<and> \<Theta>; \<B>; \<Gamma> \<turnstile> v \<Rightarrow> t1" using check_v_elims
+ using assms by blast
+ thus ?thesis using infer_bool_options assms
+ by (metis \<tau>.exhaust b_of.simps subtype_eq_base2)
+qed
+
+lemma check_v_wf:
+ fixes v::v and \<Gamma>::\<Gamma> and \<tau>::\<tau>
+ assumes "\<Theta>; \<B>; \<Gamma> \<turnstile> v \<Leftarrow> \<tau>"
+ shows " \<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>" and "\<Theta>; \<B>;\<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b_of \<tau>" and "\<Theta>; \<B>;\<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau>"
+proof -
+ obtain \<tau>' where *: "\<Theta>; \<B>; \<Gamma> \<turnstile> \<tau>' \<lesssim> \<tau> \<and> \<Theta>; \<B>; \<Gamma> \<turnstile> v \<Rightarrow> \<tau>'" using check_v_elims assms by auto
+ thus "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> " and "\<Theta>; \<B>;\<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b_of \<tau>" and "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau>"
+ using infer_v_wf infer_v_v_wf subtype_eq_base2 * subtype_wf by metis+
+qed
+
+lemma infer_v_form_fresh:
+ fixes v::v and t::"'a::fs"
+ assumes "\<Theta>; \<B>; \<Gamma> \<turnstile> v \<Rightarrow> \<tau>"
+ shows "\<exists>z b. \<tau> = \<lbrace> z : b | C_eq (CE_val (V_var z)) (CE_val v)\<rbrace> \<and> atom z \<sharp> (t,v)"
+proof -
+ obtain z' and b' where "\<tau> = \<lbrace> z' : b' | C_eq (CE_val (V_var z')) (CE_val v)\<rbrace>" using infer_v_form assms by blast
+ moreover then obtain z and b and c where "\<tau> = \<lbrace> z : b | c \<rbrace> \<and> atom z \<sharp> (t,v)" using obtain_fresh_z by metis
+ ultimately have "\<tau> = \<lbrace> z : b | C_eq (CE_val (V_var z)) (CE_val v)\<rbrace> \<and> atom z \<sharp> (t,v) "
+ using assms infer_v_form2 by auto
+ thus ?thesis by blast
+qed
+
+text \<open> More generally, if support of a term is empty then any G will do \<close>
+
+lemma infer_v_form_consp:
+ assumes "\<Theta>; \<B>; \<Gamma> \<turnstile> V_consp s dc b v \<Rightarrow> \<tau>"
+ shows "b_of \<tau> = B_app s b"
+ using assms proof(nominal_induct "V_consp s dc b v" \<tau> rule: infer_v.strong_induct)
+ case (infer_v_conspI bv dclist \<Theta> tc \<B> \<Gamma> tv z)
+ then show ?case using b_of.simps by metis
+qed
+
+lemma lookup_in_rig_b:
+ assumes "Some (b2, c2) = lookup (\<Gamma>[x\<longmapsto>c']) x'" and
+ "Some (b1, c1) = lookup \<Gamma> x'"
+ shows "b1 = b2"
+ using assms lookup_in_rig[OF assms(2)]
+ by (metis option.inject prod.inject)
+
+lemma infer_v_uniqueness_rig:
+ fixes x::x and c::c
+ assumes "infer_v P B G v \<tau>" and "infer_v P B (replace_in_g G x c') v \<tau>'"
+ shows "\<tau> = \<tau>'"
+ using assms
+proof(nominal_induct "v" arbitrary: \<tau>' \<tau> rule: v.strong_induct)
+ case (V_lit l)
+ hence "infer_l l \<tau>" and "infer_l l \<tau>'" using assms(1) infer_v_elims(2) by auto
+ then show ?case using infer_l_uniqueness by presburger
+next
+ case (V_var y)
+
+ obtain b and c where bc: "Some (b,c) = lookup G y"
+ using assms(1) infer_v_elims(2) using V_var.prems(1) lookup_iff by force
+ then obtain c'' where bc':"Some (b,c'') = lookup (replace_in_g G x c') y"
+ using lookup_in_rig by blast
+
+ obtain z where "\<tau> = (\<lbrace> z : b | C_eq (CE_val (V_var z)) (CE_val (V_var y)) \<rbrace>)" using infer_v_elims(1)[of P B G y \<tau>] V_var
+ bc option.inject prod.inject lookup_in_g by metis
+ moreover obtain z' where "\<tau>' = (\<lbrace> z' : b | C_eq (CE_val (V_var z')) (CE_val (V_var y)) \<rbrace>)" using infer_v_elims(1)[of P B _ y \<tau>'] V_var
+ option.inject prod.inject lookup_in_rig by (metis bc')
+ ultimately show ?case using type_e_eq
+ by (metis V_var.prems(1) V_var.prems(2) \<tau>.eq_iff ce.fresh(1) finite.emptyI fresh_atom_at_base
+ fresh_finite_insert infer_v_elims(1) v.fresh(2))
+next
+ case (V_pair v1 v2)
+ obtain z and z1 and z2 and t1 and t2 and c1 and c2 where
+ t1: "\<tau> = (\<lbrace> z : [ b_of t1 , b_of t2 ]\<^sup>b | CE_val (V_var z) == CE_val (V_pair v1 v2) \<rbrace>) \<and>
+ atom z \<sharp> (v1, v2) \<and> P ; B ; G \<turnstile> v1 \<Rightarrow> t1 \<and> P ; B ; G \<turnstile> v2 \<Rightarrow> t2"
+ using infer_v_elims(3)[OF V_pair(3)] by metis
+ moreover obtain z' and z1' and z2' and t1' and t2' and c1' and c2' where
+ t2: "\<tau>' = (\<lbrace> z' : [ b_of t1' , b_of t2' ]\<^sup>b | CE_val (V_var z') == CE_val (V_pair v1 v2) \<rbrace>) \<and>
+ atom z' \<sharp> (v1, v2) \<and> P ; B ; (replace_in_g G x c') \<turnstile> v1 \<Rightarrow> t1' \<and>
+ P ; B ; (replace_in_g G x c') \<turnstile> v2 \<Rightarrow> t2'"
+ using infer_v_elims(3)[OF V_pair(4)] by metis
+ ultimately have "t1 = t1' \<and> t2 = t2'" using V_pair.hyps(1) V_pair.hyps(2) \<tau>.eq_iff by blast
+ then show ?case using t1 t2 by simp
+next
+ case (V_cons s dc v)
+ obtain x and z and tc and dclist where t1: "\<tau> = (\<lbrace> z : B_id s | CE_val (V_var z) == CE_val (V_cons s dc v) \<rbrace>) \<and>
+ AF_typedef s dclist \<in> set P \<and>
+ (dc, tc) \<in> set dclist \<and> atom z \<sharp> v"
+ using infer_v_elims(4)[OF V_cons(2)] by metis
+ moreover obtain x' and z' and tc' and dclist' where t2: "\<tau>' = (\<lbrace> z' : B_id s | CE_val (V_var z') == CE_val (V_cons s dc v) \<rbrace>)
+ \<and> AF_typedef s dclist' \<in> set P \<and> (dc, tc') \<in> set dclist' \<and> atom z' \<sharp> v"
+ using infer_v_elims(4)[OF V_cons(3)] by metis
+ moreover have a: "AF_typedef s dclist' \<in> set P" and b:"(dc,tc') \<in> set dclist'" and c:"AF_typedef s dclist \<in> set P" and
+ d:"(dc, tc) \<in> set dclist" using t1 t2 by auto
+ ultimately have "tc = tc'" using wfTh_dc_t_unique2 infer_v_wf(3)[OF V_cons(2)] by metis
+
+ moreover have "atom z \<sharp> CE_val (V_cons s dc v) \<and> atom z' \<sharp> CE_val (V_cons s dc v)"
+ using e.fresh(1) v.fresh(4) t1 t2 pure_fresh by auto
+ ultimately have "(\<lbrace> z : B_id s | CE_val (V_var z) == CE_val (V_cons s dc v) \<rbrace>) = (\<lbrace> z' : B_id s | CE_val (V_var z') == CE_val (V_cons s dc v) \<rbrace>)"
+ using type_e_eq by metis
+ thus ?case using t1 t2 by simp
+next
+ case (V_consp s dc b v)
+ from V_consp(2) V_consp show ?case proof(nominal_induct "V_consp s dc b v" \<tau> arbitrary: v rule:infer_v.strong_induct)
+
+ case (infer_v_conspI bv dclist \<Theta> tc \<B> \<Gamma> v tv z)
+
+ obtain z3 and b3 where *:"\<tau>' = \<lbrace> z3 : b3 | [ [ z3 ]\<^sup>v ]\<^sup>c\<^sup>e == [ V_consp s dc b v ]\<^sup>c\<^sup>e \<rbrace> \<and> atom z3 \<sharp> V_consp s dc b v"
+ using infer_v_form[OF \<open>\<Theta>; \<B>; \<Gamma>[x\<longmapsto>c'] \<turnstile> V_consp s dc b v \<Rightarrow> \<tau>'\<close> ] by metis
+ moreover then have "b3 = B_app s b" using infer_v_form_consp b_of.simps * infer_v_conspI by metis
+
+ moreover have "\<lbrace> z3 : B_app s b | [ [ z3 ]\<^sup>v ]\<^sup>c\<^sup>e == [ V_consp s dc b v ]\<^sup>c\<^sup>e \<rbrace> = \<lbrace> z : B_app s b | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ V_consp s dc b v ]\<^sup>c\<^sup>e \<rbrace>"
+ proof -
+ have "atom z3 \<sharp> [V_consp s dc b v]\<^sup>c\<^sup>e" using * ce.fresh by auto
+ moreover have "atom z \<sharp> [V_consp s dc b v]\<^sup>c\<^sup>e" using * infer_v_conspI ce.fresh v.fresh pure_fresh by metis
+ ultimately show ?thesis using type_e_eq infer_v_conspI v.fresh ce.fresh by metis
+ qed
+ ultimately show ?case using * by auto
+ qed
+qed
+
+lemma infer_v_uniqueness:
+ assumes "infer_v P B G v \<tau>" and "infer_v P B G v \<tau>'"
+ shows "\<tau> = \<tau>'"
+proof -
+ obtain x::x where "atom x \<sharp> G" using obtain_fresh by metis
+ hence "G [ x \<longmapsto> C_true ] = G" using replace_in_g_forget assms infer_v_wf by fast
+ thus ?thesis using infer_v_uniqueness_rig assms by metis
+qed
+
+lemma infer_v_tid_form:
+ fixes v::v
+ assumes "\<Theta> ; B ; \<Gamma> \<turnstile> v \<Rightarrow> \<lbrace> z : B_id tid | c \<rbrace>" and "AF_typedef tid dclist \<in> set \<Theta>" and "supp v = {}"
+ shows "\<exists>dc v' t. v = V_cons tid dc v' \<and> (dc , t ) \<in> set dclist"
+ using assms proof(nominal_induct v "\<lbrace> z : B_id tid | c \<rbrace>" rule: infer_v.strong_induct)
+ case (infer_v_varI \<Theta> \<B> c x z)
+ then show ?case using v.supp supp_at_base by auto
+next
+ case (infer_v_litI \<Theta> \<B> l)
+ then show ?case by auto
+next
+ case (infer_v_consI dclist1 \<Theta> dc tc \<B> \<Gamma> v tv z)
+ hence "supp v = {}" using v.supp by simp
+ then obtain dca and v' where *:"V_cons tid dc v = V_cons tid dca v'" using infer_v_consI by auto
+ hence "dca = dc" using v.eq_iff(4) by auto
+ hence "V_cons tid dc v = V_cons tid dca v' \<and> (dca, tc) \<in> set dclist1" using infer_v_consI * by auto
+ moreover have "dclist = dclist1" using wfTh_dclist_unique infer_v_consI wfX_wfY \<open>dca=dc\<close>
+ proof -
+ show ?thesis
+ by (meson \<open>AF_typedef tid dclist1 \<in> set \<Theta>\<close> \<open>\<Theta>; \<B>; \<Gamma> \<turnstile> v \<Rightarrow> tv\<close> infer_v_consI.prems infer_v_wf(4) wfTh_dclist_unique wfX_wfY)
+ qed
+ ultimately show ?case by auto
+qed
+
+lemma check_v_tid_form:
+ assumes "\<Theta> ; B ; \<Gamma> \<turnstile> v \<Leftarrow> \<lbrace> z : B_id tid | TRUE \<rbrace>" and "AF_typedef tid dclist \<in> set \<Theta>" and "supp v = {}"
+ shows "\<exists>dc v' t. v = V_cons tid dc v' \<and> (dc , t ) \<in> set dclist"
+ using assms proof(nominal_induct v "\<lbrace> z : B_id tid | TRUE \<rbrace>" rule: check_v.strong_induct)
+ case (check_v_subtypeI \<Theta> \<B> \<Gamma> \<tau>1 v)
+ then obtain z and c where "\<tau>1 = \<lbrace> z : B_id tid | c \<rbrace>" using subtype_eq_base2 b_of.simps
+ by (metis obtain_fresh_z2)
+ then show ?case using infer_v_tid_form check_v_subtypeI by simp
+qed
+
+lemma check_v_num_leq:
+ fixes n::int and \<Gamma>::\<Gamma>
+ assumes "0 \<le> n \<and> n \<le> int (length v)" and " \<turnstile>\<^sub>w\<^sub>f \<Theta> " and "\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f \<Gamma>"
+ shows "\<Theta> ; {||} ; \<Gamma> \<turnstile> [ L_num n ]\<^sup>v \<Leftarrow> \<lbrace> z : B_int | ([ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e)
+ AND ([ leq [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e [| [ [ L_bitvec v ]\<^sup>v ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e ) \<rbrace>"
+proof -
+ have "\<Theta> ; {||} ; \<Gamma> \<turnstile> [ L_num n ]\<^sup>v \<Rightarrow> \<lbrace> z : B_int | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_num n ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrace>"
+ using infer_v_litI infer_natI wfG_nilI assms by auto
+ thus ?thesis using subtype_range[OF assms(1) ] assms check_v_subtypeI by metis
+qed
+
+lemma check_int:
+ assumes "check_v \<Theta> \<B> \<Gamma> v (\<lbrace> z : B_int | c \<rbrace>)" and "supp v = {}"
+ shows "\<exists>n. V_lit (L_num n) = v"
+ using assms infer_int check_v_elims by (metis b_of.simps infer_v_form subtype_eq_base2)
+
+definition sble :: "\<Theta> \<Rightarrow> \<Gamma> \<Rightarrow> bool" where
+ "sble \<Theta> \<Gamma> = (\<exists>i. i \<Turnstile> \<Gamma> \<and> \<Theta> ; \<Gamma> \<turnstile> i)"
+
+lemma check_v_range:
+ assumes "\<Theta> ; B ; \<Gamma> \<turnstile> v2 \<Leftarrow> \<lbrace> z : B_int | [ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e AND
+ [ leq [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e [| [ v1 ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrace> "
+ (is "\<Theta> ; ?B ; \<Gamma> \<turnstile> v2 \<Leftarrow> \<lbrace> z : B_int | ?c1 \<rbrace>")
+ and "v1 = V_lit (L_bitvec bv) \<and> v2 = V_lit (L_num n) " and "atom z \<sharp> \<Gamma>" and "sble \<Theta> \<Gamma>"
+ shows "0 \<le> n \<and> n \<le> int (length bv)"
+proof -
+ have "\<Theta> ; ?B ; \<Gamma> \<turnstile> \<lbrace> z : B_int | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_num n ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrace> \<lesssim> \<lbrace> z : B_int | ?c1 \<rbrace>"
+ using check_v_elims assms
+ by (metis infer_l_uniqueness infer_natI infer_v_elims(2))
+ moreover have "atom z \<sharp> \<Gamma>" using fresh_GNil assms by simp
+ ultimately have "\<Theta> ; ?B ; ((z, B_int, [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_num n ]\<^sup>v ]\<^sup>c\<^sup>e ) #\<^sub>\<Gamma> \<Gamma>) \<Turnstile> ?c1"
+ using subtype_valid_simple by auto
+ thus ?thesis using assms valid_range_length_inv check_v_wf wfX_wfY sble_def by metis
+qed
+
+section \<open>Expressions\<close>
+
+lemma infer_e_plus[elim]:
+ fixes v1::v and v2::v
+ assumes "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile> AE_op Plus v1 v2 \<Rightarrow> \<tau>"
+ shows "\<exists>z . (\<lbrace> z : B_int | C_eq (CE_val (V_var z)) (CE_op Plus [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e) \<rbrace> = \<tau>)"
+ using infer_e_elims assms by metis
+
+lemma infer_e_leq[elim]:
+ assumes "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile> AE_op LEq v1 v2 \<Rightarrow> \<tau>"
+ shows "\<exists>z . (\<lbrace> z : B_bool | C_eq (CE_val (V_var z)) (CE_op LEq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e) \<rbrace> = \<tau>)"
+ using infer_e_elims assms by metis
+
+lemma infer_e_eq[elim]:
+ assumes "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile> AE_op Eq v1 v2 \<Rightarrow> \<tau>"
+ shows "\<exists>z . (\<lbrace> z : B_bool | C_eq (CE_val (V_var z)) (CE_op Eq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e) \<rbrace> = \<tau>)"
+ using infer_e_elims(25)[OF assms] by metis
+
+lemma infer_e_e_wf:
+ fixes e::e
+ assumes "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile> e \<Rightarrow> \<tau>"
+ shows "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f e : b_of \<tau>"
+ using assms proof(nominal_induct \<tau> avoiding: \<tau> rule: infer_e.strong_induct)
+ case (infer_e_valI \<Theta> \<B> \<Gamma> \<Delta>' \<Phi> v \<tau>)
+ then show ?case using infer_v_v_wf wf_intros by metis
+next
+ case (infer_e_plusI \<Theta> \<B> \<Gamma> \<Delta>' \<Phi> v1 z1 c1 v2 z2 c2 z3)
+ then show ?case using b_of.simps infer_v_v_wf wf_intros by metis
+next
+ case (infer_e_leqI \<Theta> \<B> \<Gamma> \<Delta>' v1 z1 c1 v2 z2 c2 z3)
+ then show ?case using b_of.simps infer_v_v_wf wf_intros by metis
+next
+ case (infer_e_eqI \<Theta> \<B> \<Gamma> \<Delta>' v1 z1 c1 v2 z2 c2 z3)
+ then show ?case using b_of.simps infer_v_v_wf wf_intros by metis
+next
+ case (infer_e_appI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> f x b c \<tau>' s' v \<tau>'')
+ have "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f AE_app f v : b_of \<tau>' " proof
+ show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using infer_e_appI by auto
+ show \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using infer_e_appI by auto
+ show \<open>Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \<tau>' s'))) = lookup_fun \<Phi> f\<close> using infer_e_appI by auto
+ show "\<Theta>; \<B>;\<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b" using infer_e_appI check_v_wf b_of.simps by metis
+ qed
+ moreover have "b_of \<tau>' = b_of (\<tau>'[x::=v]\<^sub>v)" using subst_tbase_eq subst_v_\<tau>_def by auto
+ ultimately show ?case using infer_e_appI subst_v_c_def subst_b_\<tau>_def by auto
+next
+ case (infer_e_appPI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> b' f bv x b c \<tau>'' s' v \<tau>')
+
+ have "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f AE_appP f b' v : (b_of \<tau>'')[bv::=b']\<^sub>b " proof
+ show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using infer_e_appPI by auto
+ show \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using infer_e_appPI by auto
+ show \<open>Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \<tau>'' s'))) = lookup_fun \<Phi> f\<close> using * infer_e_appPI by metis
+ show "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f b'" using infer_e_appPI by auto
+ show "\<Theta>; \<B>;\<Gamma> \<turnstile>\<^sub>w\<^sub>f v : (b[bv::=b']\<^sub>b)" using infer_e_appPI check_v_wf b_of.simps subst_b_b_def by metis
+ have "atom bv \<sharp> (b_of \<tau>'')[bv::=b']\<^sub>b\<^sub>b" using fresh_subst_if subst_b_b_def infer_e_appPI by metis
+ thus "atom bv \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>, \<Delta>, b', v, (b_of \<tau>'')[bv::=b']\<^sub>b)" using infer_e_appPI fresh_prodN subst_b_b_def by metis
+ qed
+ moreover have "b_of \<tau>' = (b_of \<tau>'')[bv::=b']\<^sub>b"
+ using \<open>\<tau>''[bv::=b']\<^sub>b[x::=v]\<^sub>v = \<tau>'\<close> b_of_subst_bb_commute subst_tbase_eq subst_b_b_def subst_v_\<tau>_def subst_b_\<tau>_def by auto
+ ultimately show ?case using infer_e_appI by auto
+next
+ case (infer_e_fstI \<Theta> \<B> \<Gamma> \<Delta>' \<Phi> v z' b1 b2 c z)
+ then show ?case using b_of.simps infer_v_v_wf wf_intros by metis
+next
+ case (infer_e_sndI \<Theta> \<B> \<Gamma> \<Delta>' \<Phi> v z' b1 b2 c z)
+ then show ?case using b_of.simps infer_v_v_wf wf_intros by metis
+next
+ case (infer_e_lenI \<Theta> \<B> \<Gamma> \<Delta>' \<Phi> v z' c z)
+ then show ?case using b_of.simps infer_v_v_wf wf_intros by metis
+next
+ case (infer_e_mvarI \<Theta> \<Gamma> \<Phi> \<Delta> u \<tau>)
+ then show ?case using b_of.simps infer_v_v_wf wf_intros by metis
+next
+ case (infer_e_concatI \<Theta> \<B> \<Gamma> \<Delta>' \<Phi> v1 z1 c1 v2 z2 c2 z3)
+ then show ?case using b_of.simps infer_v_v_wf wf_intros by metis
+next
+ case (infer_e_splitI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v1 z1 c1 v2 z2 z3)
+ have " \<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f AE_split v1 v2 : B_pair B_bitvec B_bitvec"
+ proof
+ show "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>" using infer_e_splitI by auto
+ show "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>" using infer_e_splitI by auto
+ show "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v1 : B_bitvec" using infer_e_splitI b_of.simps infer_v_wf by metis
+ show "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v2 : B_int" using infer_e_splitI b_of.simps check_v_wf by metis
+ qed
+ then show ?case using b_of.simps by auto
+qed
+
+lemma infer_e_t_wf:
+ fixes e::e and \<Gamma>::\<Gamma> and \<tau>::\<tau> and \<Delta>::\<Delta> and \<Phi>::\<Phi>
+ assumes "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile> e \<Rightarrow> \<tau>"
+ shows "\<Theta>; \<B>;\<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau> \<and> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>"
+ using assms proof(induct rule: infer_e.induct)
+ case (infer_e_valI \<Theta> \<B> \<Gamma> \<Delta>' \<Phi> v \<tau>)
+ then show ?case using infer_v_t_wf by auto
+next
+ case (infer_e_plusI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v1 z1 c1 v2 z2 c2 z3)
+ hence "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_op Plus [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e : B_int" using wfCE_plusI wfD_emptyI wfPhi_emptyI infer_v_v_wf wfCE_valI
+ by (metis b_of.simps infer_v_wf)
+ then show ?case using wfT_e_eq infer_e_plusI by auto
+next
+ case (infer_e_leqI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v1 z1 c1 v2 z2 c2 z3)
+ hence " \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_op LEq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e : B_bool" using wfCE_leqI wfD_emptyI wfPhi_emptyI infer_v_v_wf wfCE_valI
+ by (metis b_of.simps infer_v_wf)
+ then show ?case using wfT_e_eq infer_e_leqI by auto
+next
+ case (infer_e_eqI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v1 z1 b c1 v2 z2 c2 z3)
+ hence " \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_op Eq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e : B_bool" using wfCE_eqI wfD_emptyI wfPhi_emptyI infer_v_v_wf wfCE_valI
+ by (metis b_of.simps infer_v_wf)
+ then show ?case using wfT_e_eq infer_e_eqI by auto
+next
+ case (infer_e_appI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> f x b c \<tau> s' v \<tau>')
+ show ?case proof
+ show "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>" using infer_e_appI by auto
+ show "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau>'" proof -
+ have *:" \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b " using infer_e_appI check_v_wf(2) b_of.simps by metis
+ moreover have *:"\<Theta>; \<B>; (x, b, c) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau>" proof(rule wf_weakening1(4))
+ show \<open> \<Theta>; \<B>; (x,b,c)#\<^sub>\<Gamma>GNil \<turnstile>\<^sub>w\<^sub>f \<tau> \<close> using wfPhi_f_simple_wfT wfD_wf infer_e_appI wb_b_weakening by fastforce
+ have "\<Theta> ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> x : b | c \<rbrace>" using infer_e_appI check_v_wf(3) by auto
+ thus \<open> \<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f (x, b, c) #\<^sub>\<Gamma> \<Gamma> \<close> using infer_e_appI
+ wfT_wfC[THEN wfG_consI[rotated 3] ] * wfT_wf_cons fresh_prodN by simp
+ show \<open>toSet ((x,b,c)#\<^sub>\<Gamma>GNil) \<subseteq> toSet ((x, b, c) #\<^sub>\<Gamma> \<Gamma>)\<close> using toSet.simps by auto
+ qed
+ moreover have "((x, b, c) #\<^sub>\<Gamma> \<Gamma>)[x::=v]\<^sub>\<Gamma>\<^sub>v = \<Gamma>" using subst_gv.simps by auto
+
+ ultimately show ?thesis using infer_e_appI wf_subst1(4)[OF *, of GNil x b c \<Gamma> v] subst_v_\<tau>_def by auto
+ qed
+ qed
+next
+ case (infer_e_appPI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> b' f bv x b c \<tau>' s' v \<tau>)
+
+ have "\<Theta>; \<B>; ((x, b[bv::=b']\<^sub>b\<^sub>b, c[bv::=b']\<^sub>c\<^sub>b) #\<^sub>\<Gamma> \<Gamma>)[x::=v]\<^sub>\<Gamma>\<^sub>v \<turnstile>\<^sub>w\<^sub>f (\<tau>'[bv::=b']\<^sub>b)[x::=v]\<^sub>\<tau>\<^sub>v"
+ proof(rule wf_subst(4))
+ show \<open> \<Theta>; \<B>; (x, b[bv::=b']\<^sub>b\<^sub>b, c[bv::=b']\<^sub>c\<^sub>b) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau>'[bv::=b']\<^sub>b \<close>
+ proof(rule wf_weakening1(4))
+ have \<open> \<Theta> ; {|bv|} ; (x,b,c)#\<^sub>\<Gamma>GNil \<turnstile>\<^sub>w\<^sub>f \<tau>' \<close> using wfPhi_f_poly_wfT infer_e_appI infer_e_appPI by simp
+ thus \<open> \<Theta>; \<B>; (x,b[bv::=b']\<^sub>b\<^sub>b,c[bv::=b']\<^sub>c\<^sub>b)#\<^sub>\<Gamma>GNil \<turnstile>\<^sub>w\<^sub>f \<tau>'[bv::=b']\<^sub>b \<close>
+ using wfT_subst_wfT infer_e_appPI wb_b_weakening subst_b_\<tau>_def subst_v_\<tau>_def by presburger
+ have "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> x : b[bv::=b']\<^sub>b\<^sub>b | c[bv::=b']\<^sub>c\<^sub>b \<rbrace>"
+ using infer_e_appPI check_v_wf(3) subst_b_b_def subst_b_c_def by metis
+ thus \<open> \<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f (x, b[bv::=b']\<^sub>b\<^sub>b, c[bv::=b']\<^sub>c\<^sub>b) #\<^sub>\<Gamma> \<Gamma> \<close>
+ using infer_e_appPI wfT_wfC[THEN wfG_consI[rotated 3] ] * wfX_wfY wfT_wf_cons wb_b_weakening by metis
+ show \<open>toSet ((x,b[bv::=b']\<^sub>b\<^sub>b,c[bv::=b']\<^sub>c\<^sub>b)#\<^sub>\<Gamma>GNil) \<subseteq> toSet ((x, b[bv::=b']\<^sub>b\<^sub>b, c[bv::=b']\<^sub>c\<^sub>b) #\<^sub>\<Gamma> \<Gamma>)\<close> using toSet.simps by auto
+ qed
+ show \<open>(x, b[bv::=b']\<^sub>b\<^sub>b, c[bv::=b']\<^sub>c\<^sub>b) #\<^sub>\<Gamma> \<Gamma> = GNil @ (x, b[bv::=b']\<^sub>b\<^sub>b, c[bv::=b']\<^sub>c\<^sub>b) #\<^sub>\<Gamma> \<Gamma>\<close> using append_g.simps by auto
+ show \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v :b[bv::=b']\<^sub>b\<^sub>b \<close> using infer_e_appPI check_v_wf(2) b_of.simps subst_b_b_def by metis
+ qed
+ moreover have "((x, b[bv::=b']\<^sub>b\<^sub>b, c[bv::=b']\<^sub>c\<^sub>b) #\<^sub>\<Gamma> \<Gamma>)[x::=v]\<^sub>\<Gamma>\<^sub>v = \<Gamma>" using subst_gv.simps by auto
+ ultimately show ?case using infer_e_appPI subst_v_\<tau>_def by simp
+next
+ case (infer_e_fstI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v z' b1 b2 c z)
+ hence "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_fst [v]\<^sup>c\<^sup>e: b1" using wfCE_fstI wfD_emptyI wfPhi_emptyI infer_v_v_wf
+ b_of.simps using wfCE_valI by fastforce
+ then show ?case using wfT_e_eq infer_e_fstI by auto
+next
+ case (infer_e_sndI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v z' b1 b2 c z)
+ hence "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_snd [v]\<^sup>c\<^sup>e: b2" using wfCE_sndI wfD_emptyI wfPhi_emptyI infer_v_v_wf wfCE_valI
+ by (metis b_of.simps infer_v_wf)
+ then show ?case using wfT_e_eq infer_e_sndI by auto
+next
+ case (infer_e_lenI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v z' c z)
+ hence "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_len [v]\<^sup>c\<^sup>e: B_int" using wfCE_lenI wfD_emptyI wfPhi_emptyI infer_v_v_wf wfCE_valI
+ by (metis b_of.simps infer_v_wf)
+ then show ?case using wfT_e_eq infer_e_lenI by auto
+next
+ case (infer_e_mvarI \<Theta> \<Gamma> \<Phi> \<Delta> u \<tau>)
+ then show ?case using wfD_wfT by blast
+next
+ case (infer_e_concatI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v1 z1 c1 v2 z2 c2 z3)
+ hence "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_concat [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e: B_bitvec" using wfCE_concatI wfD_emptyI wfPhi_emptyI infer_v_v_wf wfCE_valI
+ by (metis b_of.simps infer_v_wf)
+ then show ?case using wfT_e_eq infer_e_concatI by auto
+next
+ case (infer_e_splitI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v1 z1 c1 v2 z2 z3)
+
+ hence wfg: " \<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f (z3, [ B_bitvec , B_bitvec ]\<^sup>b, TRUE) #\<^sub>\<Gamma> \<Gamma>"
+ using infer_v_wf wfG_cons2I wfB_pairI wfB_bitvecI by simp
+ have wfz: "\<Theta>; \<B>; (z3, [ B_bitvec , B_bitvec ]\<^sup>b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f [ [ z3 ]\<^sup>v ]\<^sup>c\<^sup>e : [ B_bitvec , B_bitvec ]\<^sup>b "
+ apply(rule wfCE_valI , rule wfV_varI)
+ using wfg apply simp
+ using lookup.simps(2)[of z3 "[ B_bitvec , B_bitvec ]\<^sup>b" TRUE \<Gamma> z3] by simp
+ have 1: "\<Theta>; \<B>; (z3, [ B_bitvec , B_bitvec ]\<^sup>b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f [ v2 ]\<^sup>c\<^sup>e : B_int "
+ using check_v_wf[OF infer_e_splitI(4)] wf_weakening(1)[OF _ wfg] b_of.simps toSet.simps wfCE_valI by fastforce
+ have 2: "\<Theta>; \<B>; (z3, [ B_bitvec , B_bitvec ]\<^sup>b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f [ v1 ]\<^sup>c\<^sup>e : B_bitvec"
+ using infer_v_wf[OF infer_e_splitI(3)] wf_weakening(1)[OF _ wfg] b_of.simps toSet.simps wfCE_valI by fastforce
+
+ have "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> z3 : [ B_bitvec , B_bitvec ]\<^sup>b | [ v1 ]\<^sup>c\<^sup>e == [ [#1[ [ z3 ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e @@ [#2[ [ z3 ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e AND [| [#1[ [ z3 ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e == [ v2 ]\<^sup>c\<^sup>e \<rbrace>"
+ proof
+ show "atom z3 \<sharp> (\<Theta>, \<B>, \<Gamma>)" using infer_e_splitI wfTh_x_fresh wfX_wfY fresh_prod3 wfG_fresh_x by metis
+ show "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f [ B_bitvec , B_bitvec ]\<^sup>b " using wfB_pairI wfB_bitvecI infer_e_splitI wfX_wfY by metis
+ show "\<Theta>; \<B>; (z3, [ B_bitvec , B_bitvec ]\<^sup>b, TRUE) #\<^sub>\<Gamma>
+ \<Gamma> \<turnstile>\<^sub>w\<^sub>f [ v1 ]\<^sup>c\<^sup>e == [ [#1[ [ z3 ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e @@ [#2[ [ z3 ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e AND [| [#1[ [ z3 ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e == [ v2 ]\<^sup>c\<^sup>e"
+ using wfg wfz 1 2 wf_intros by meson
+ qed
+ thus ?case using infer_e_splitI by auto
+qed
+
+lemma infer_e_wf:
+ fixes e::e and \<Gamma>::\<Gamma> and \<tau>::\<tau> and \<Delta>::\<Delta> and \<Phi>::\<Phi>
+ assumes "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile> e \<Rightarrow> \<tau>"
+ shows "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau>" and "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>" and "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>" and "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>" and "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f e : (b_of \<tau>)"
+ using infer_e_t_wf infer_e_e_wf wfE_wf assms by metis+
+
+lemma infer_e_fresh:
+ fixes x::x
+ assumes "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile> e \<Rightarrow> \<tau>" and "atom x \<sharp> \<Gamma>"
+ shows "atom x \<sharp> (e,\<tau>)"
+proof -
+ have "atom x \<sharp> e" using infer_e_e_wf[THEN wfE_x_fresh,OF assms(1)] assms(2) by auto
+ moreover have "atom x \<sharp> \<tau>" using assms infer_e_wf wfT_x_fresh by metis
+ ultimately show ?thesis using fresh_Pair by auto
+qed
+
+inductive check_e :: "\<Theta> \<Rightarrow> \<Phi> \<Rightarrow> \<B> \<Rightarrow> \<Gamma> \<Rightarrow> \<Delta> \<Rightarrow> e \<Rightarrow> \<tau> \<Rightarrow> bool" (" _ ; _ ; _ ; _ ; _ \<turnstile> _ \<Leftarrow> _" [50, 50, 50] 50) where
+ check_e_subtypeI: "\<lbrakk> infer_e T P B G D e \<tau>' ; subtype T B G \<tau>' \<tau> \<rbrakk> \<Longrightarrow> check_e T P B G D e \<tau>"
+equivariance check_e
+nominal_inductive check_e .
+
+inductive_cases check_e_elims[elim!]:
+ "check_e F D B G \<Theta> (AE_val v) \<tau>"
+ "check_e F D B G \<Theta> e \<tau>"
+
+lemma infer_e_fst_pair:
+ fixes v1::v
+ assumes "\<Theta> ; \<Phi> ; {||} ; GNil ; \<Delta> \<turnstile> [#1[ v1 , v2 ]\<^sup>v]\<^sup>e \<Rightarrow> \<tau>"
+ shows "\<exists>\<tau>'. \<Theta> ; \<Phi> ; {||} ; GNil ; \<Delta> \<turnstile> [v1]\<^sup>e \<Rightarrow> \<tau>' \<and>
+ \<Theta> ; {||} ; GNil \<turnstile> \<tau>' \<lesssim> \<tau>"
+proof -
+ obtain z' and b1 and b2 and c and z where ** : "\<tau> = (\<lbrace> z : b1 | C_eq (CE_val (V_var z)) (CE_fst [(V_pair v1 v2)]\<^sup>c\<^sup>e) \<rbrace>) \<and>
+ wfD \<Theta> {||} GNil \<Delta> \<and> wfPhi \<Theta> \<Phi> \<and>
+ (\<Theta> ; {||} ; GNil \<turnstile> V_pair v1 v2 \<Rightarrow> \<lbrace> z' : B_pair b1 b2 | c \<rbrace>) \<and> atom z \<sharp> V_pair v1 v2 "
+ using infer_e_elims assms by metis
+ hence *:" \<Theta> ; {||} ; GNil \<turnstile> V_pair v1 v2 \<Rightarrow> \<lbrace> z' : B_pair b1 b2 | c \<rbrace>" by auto
+
+ obtain t1a and t2a where
+ *: "\<Theta> ; {||} ; GNil \<turnstile> v1 \<Rightarrow> t1a \<and> \<Theta> ; {||} ; GNil \<turnstile> v2 \<Rightarrow> t2a \<and> B_pair b1 b2 = B_pair (b_of t1a) (b_of t2a)"
+ using infer_v_elims(5)[OF *] by metis
+
+ hence suppv: "supp v1 = {} \<and> supp v2 = {} \<and> supp (V_pair v1 v2) = {}" using ** infer_v_v_wf wfV_supp atom_dom.simps toSet.simps supp_GNil
+ by (meson wfV_supp_nil)
+
+ obtain z1 and b1' where "t1a = \<lbrace> z1 : b1' | [ [ z1 ]\<^sup>v ]\<^sup>c\<^sup>e == [ v1 ]\<^sup>c\<^sup>e \<rbrace> "
+ using infer_v_form[of \<Theta> "{||}" GNil v1 t1a] * by auto
+ moreover hence "b1' = b1" using * b_of.simps by simp
+ ultimately have "\<Theta> ; {||} ; GNil \<turnstile> v1 \<Rightarrow> \<lbrace> z1 : b1 | CE_val (V_var z1) == CE_val v1 \<rbrace>" using * by auto
+ moreover have "\<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f CE_fst [V_pair v1 v2]\<^sup>c\<^sup>e : b1" using wfCE_fstI infer_v_wf(1) ** b_of.simps wfCE_valI by metis
+ moreover hence st: "\<Theta> ; {||} ; GNil \<turnstile> \<lbrace> z1 : b1 | CE_val (V_var z1) == CE_val v1 \<rbrace> \<lesssim> (\<lbrace> z : b1 | CE_val (V_var z) == CE_fst [V_pair v1 v2]\<^sup>c\<^sup>e \<rbrace>)"
+ using subtype_gnil_fst infer_v_v_wf by auto
+ moreover have "wfD \<Theta> {||} GNil \<Delta> \<and> wfPhi \<Theta> \<Phi>" using ** by auto
+ ultimately show ?thesis using wfX_wfY ** infer_e_valI by metis
+qed
+
+lemma infer_e_snd_pair:
+ assumes "\<Theta> ; \<Phi> ; {||} ; GNil ; \<Delta> \<turnstile> AE_snd (V_pair v1 v2) \<Rightarrow> \<tau>"
+ shows "\<exists>\<tau>'. \<Theta> ; \<Phi> ; {||} ; GNil ; \<Delta> \<turnstile> AE_val v2 \<Rightarrow> \<tau>' \<and> \<Theta> ; {||} ; GNil \<turnstile> \<tau>' \<lesssim> \<tau>"
+proof -
+ obtain z' and b1 and b2 and c and z where ** : "(\<tau> = (\<lbrace> z : b2 | C_eq (CE_val (V_var z)) (CE_snd [(V_pair v1 v2)]\<^sup>c\<^sup>e) \<rbrace>)) \<and>
+ (wfD \<Theta> {||} GNil \<Delta>) \<and> (wfPhi \<Theta> \<Phi>) \<and>
+ \<Theta> ; {||} ; GNil \<turnstile> V_pair v1 v2 \<Rightarrow> \<lbrace> z' : B_pair b1 b2 | c \<rbrace> \<and> atom z \<sharp> V_pair v1 v2 "
+ using infer_e_elims(9)[OF assms(1)] by metis
+ hence *:" \<Theta> ; {||} ; GNil \<turnstile> V_pair v1 v2 \<Rightarrow> \<lbrace> z' : B_pair b1 b2 | c \<rbrace>" by auto
+
+ obtain t1a and t2a where
+ *: "\<Theta> ; {||} ; GNil \<turnstile> v1 \<Rightarrow> t1a \<and> \<Theta> ; {||} ; GNil \<turnstile> v2 \<Rightarrow> t2a \<and> B_pair b1 b2 = B_pair (b_of t1a) (b_of t2a)"
+ using infer_v_elims(5)[OF *] by metis
+
+ hence suppv: "supp v1 = {} \<and> supp v2 = {} \<and> supp (V_pair v1 v2) = {}" using infer_v_v_wf wfV.simps v.supp by (meson "**" wfV_supp_nil)
+
+ obtain z2 and b2' where "t2a = \<lbrace> z2 : b2' | [ [ z2 ]\<^sup>v ]\<^sup>c\<^sup>e == [ v2 ]\<^sup>c\<^sup>e \<rbrace> "
+ using infer_v_form[of \<Theta> "{||}" GNil v2 t2a] * by auto
+ moreover hence "b2' = b2" using * b_of.simps by simp
+
+ ultimately have "\<Theta> ; {||} ; GNil \<turnstile> v2 \<Rightarrow> \<lbrace> z2 : b2 | CE_val (V_var z2) == CE_val v2 \<rbrace>" using * by auto
+ moreover have "\<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f CE_snd [V_pair v1 v2]\<^sup>c\<^sup>e : b2" using wfCE_sndI infer_v_wf(1) ** b_of.simps wfCE_valI by metis
+ moreover hence st: "\<Theta> ; {||} ; GNil \<turnstile> \<lbrace> z2 : b2 | CE_val (V_var z2) == CE_val v2 \<rbrace> \<lesssim> (\<lbrace> z : b2 | CE_val (V_var z) == CE_snd [V_pair v1 v2]\<^sup>c\<^sup>e \<rbrace>)"
+ using subtype_gnil_snd infer_v_v_wf by auto
+ moreover have "wfD \<Theta> {||} GNil \<Delta> \<and> wfPhi \<Theta> \<Phi>" using ** by metis
+ ultimately show ?thesis using wfX_wfY ** infer_e_valI by metis
+qed
+
+section \<open>Statements\<close>
+
+lemma check_s_v_unit:
+ assumes "\<Theta>; \<B>; \<Gamma> \<turnstile> (\<lbrace> z : B_unit | TRUE \<rbrace>) \<lesssim> \<tau>" and "wfD \<Theta> \<B> \<Gamma> \<Delta>" and "wfPhi \<Theta> \<Phi>"
+ shows "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile> AS_val (V_lit L_unit ) \<Leftarrow> \<tau>"
+proof -
+ have "wfG \<Theta> \<B> \<Gamma>" using assms subtype_g_wf by meson
+ moreover hence "wfTh \<Theta>" using wfG_wf by simp
+ moreover obtain z'::x where "atom z' \<sharp> \<Gamma>" using obtain_fresh by auto
+ ultimately have *:"\<Theta>; \<B>; \<Gamma> \<turnstile> V_lit L_unit \<Rightarrow> \<lbrace> z' : B_unit | CE_val (V_var z') == CE_val (V_lit L_unit) \<rbrace>"
+ using infer_v_litI infer_unitI by simp
+ moreover have "wfT \<Theta> \<B> \<Gamma> (\<lbrace> z' : B_unit | CE_val (V_var z') == CE_val (V_lit L_unit) \<rbrace>)" using infer_v_t_wf
+ by (meson calculation)
+ moreover then have "\<Theta>; \<B>; \<Gamma> \<turnstile> (\<lbrace> z' : B_unit | CE_val (V_var z') == CE_val (V_lit L_unit) \<rbrace>) \<lesssim> \<tau>" using subtype_trans subtype_top assms
+ type_for_lit.simps(4) wfX_wfY by metis
+ ultimately show ?thesis using check_valI assms * by auto
+qed
+
+lemma check_s_check_branch_s_wf:
+ fixes s::s and cs::branch_s and \<Theta>::\<Theta> and \<Phi>::\<Phi> and \<Gamma>::\<Gamma> and \<Delta>::\<Delta> and v::v and \<tau>::\<tau> and css::branch_list
+ shows "\<Theta> ; \<Phi> ; B ; \<Gamma> ; \<Delta> \<turnstile> s \<Leftarrow> \<tau> \<Longrightarrow> \<Theta> ; B \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<and> wfTh \<Theta> \<and> wfD \<Theta> B \<Gamma> \<Delta> \<and> wfT \<Theta> B \<Gamma> \<tau> \<and> wfPhi \<Theta> \<Phi> " and
+ "check_branch_s \<Theta> \<Phi> B \<Gamma> \<Delta> tid cons const v cs \<tau> \<Longrightarrow> \<Theta> ; B \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<and> wfTh \<Theta> \<and> wfD \<Theta> B \<Gamma> \<Delta> \<and> wfT \<Theta> B \<Gamma> \<tau> \<and> wfPhi \<Theta> \<Phi> "
+ "check_branch_list \<Theta> \<Phi> B \<Gamma> \<Delta> tid dclist v css \<tau> \<Longrightarrow> \<Theta> ; B \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<and> wfTh \<Theta> \<and> wfD \<Theta> B \<Gamma> \<Delta> \<and> wfT \<Theta> B \<Gamma> \<tau> \<and> wfPhi \<Theta> \<Phi> "
+proof(induct rule: check_s_check_branch_s_check_branch_list.inducts)
+ case (check_valI \<Theta> B \<Gamma> \<Delta> \<Phi> v \<tau>' \<tau>)
+ then show ?case using infer_v_wf infer_v_wf subtype_wf wfX_wfY wfS_valI
+ by (metis subtype_eq_base2)
+next
+ case (check_letI x \<Theta> \<Phi> \<B> \<Gamma> \<Delta> e \<tau> z s b c)
+ then have *:"wfT \<Theta> \<B> ((x, b , c[z::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) \<tau>" by force
+ moreover have "atom x \<sharp> \<tau>" using check_letI fresh_prodN by force
+ ultimately have "\<Theta>; \<B>;\<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau>" using wfT_restrict2 by force
+ then show ?case using check_letI infer_e_wf wfS_letI wfX_wfY wfG_elims by metis
+next
+ case (check_assertI x \<Theta> \<Phi> \<B> \<Gamma> \<Delta> c \<tau> s)
+ then have *:"wfT \<Theta> \<B> ((x, B_bool , c) #\<^sub>\<Gamma> \<Gamma>) \<tau>" by force
+ moreover have "atom x \<sharp> \<tau>" using check_assertI fresh_prodN by force
+ ultimately have "\<Theta>; \<B>;\<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau>" using wfT_restrict2 by force
+ then show ?case using check_assertI wfS_assertI wfX_wfY wfG_elims by metis
+next
+ case (check_branch_s_branchI \<Theta> \<B> \<Gamma> \<Delta> \<tau> cons const x v \<Phi> s tid)
+ then show ?case using wfX_wfY by metis
+next
+ case (check_branch_list_consI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist' v cs \<tau> css )
+ then show ?case using wfX_wfY by metis
+next
+ case (check_branch_list_finalI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist' v cs \<tau> )
+ then show ?case using wfX_wfY by metis
+next
+ case (check_ifI z \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v s1 s2 \<tau>)
+ hence *:"wfT \<Theta> \<B> \<Gamma> (\<lbrace> z : b_of \<tau> | CE_val v == CE_val (V_lit L_false) IMP c_of \<tau> z \<rbrace>)" (is "wfT \<Theta> \<B> \<Gamma> ?tau") by auto
+ hence "wfT \<Theta> \<B> \<Gamma> \<tau>" using wfT_wfT_if1 check_ifI fresh_prodN by metis
+ hence " \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau> " using check_ifI b_of_c_of_eq fresh_prodN by auto
+ thus ?case using check_ifI by metis
+next
+ case (check_let2I x \<Theta> \<Phi> \<B> G \<Delta> t s1 \<tau> s2)
+ then have "wfT \<Theta> \<B> ((x, b_of t, (c_of t x)) #\<^sub>\<Gamma> G) \<tau>" by fastforce
+ moreover have "atom x \<sharp> \<tau>" using check_let2I by force
+ ultimately have "wfT \<Theta> \<B> G \<tau>" using wfT_restrict2 by metis
+ then show ?case using check_let2I by argo
+next
+ case (check_varI u \<Delta> P G v \<tau>' \<Phi> s \<tau>)
+ then show ?case using wfG_elims wfD_elims
+ list.distinct list.inject by metis
+next
+ case (check_assignI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> u \<tau> v z \<tau>')
+ obtain z'::x where *:"atom z' \<sharp> \<Gamma>" using obtain_fresh by metis
+ moreover have "\<lbrace> z : B_unit | TRUE \<rbrace> = \<lbrace> z' : B_unit | TRUE \<rbrace>" by auto
+ moreover hence "wfT \<Theta> \<B> \<Gamma> \<lbrace> z' : B_unit |TRUE \<rbrace>" using wfT_TRUE check_assignI check_v_wf * wfB_unitI wfG_wf by metis
+ ultimately show ?case using check_v.cases infer_v_wf subtype_wf check_assignI wfT_wf check_v_wf wfG_wf
+ by (meson subtype_wf)
+next
+ case (check_whileI \<Phi> \<Delta> G P s1 z s2 \<tau>')
+ then show ?case using subtype_wf subtype_wf by auto
+next
+ case (check_seqI \<Delta> G P s1 z s2 \<tau>)
+ then show ?case by fast
+next
+ case (check_caseI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> dclist cs \<tau> tid v z)
+ then show ?case by fast
+qed
+
+lemma check_s_check_branch_s_wfS:
+ fixes s::s and cs::branch_s and \<Theta>::\<Theta> and \<Phi>::\<Phi> and \<Gamma>::\<Gamma> and \<Delta>::\<Delta> and v::v and \<tau>::\<tau> and css::branch_list
+ shows "\<Theta> ; \<Phi> ; B ; \<Gamma> ; \<Delta> \<turnstile> s \<Leftarrow> \<tau> \<Longrightarrow> \<Theta> ; \<Phi> ; B ; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b_of \<tau> " and
+ "check_branch_s \<Theta> \<Phi> B \<Gamma> \<Delta> tid cons const v cs \<tau> \<Longrightarrow> wfCS \<Theta> \<Phi> B \<Gamma> \<Delta> tid cons const cs (b_of \<tau>)"
+ "check_branch_list \<Theta> \<Phi> B \<Gamma> \<Delta> tid dclist v css \<tau> \<Longrightarrow> wfCSS \<Theta> \<Phi> B \<Gamma> \<Delta> tid dclist css (b_of \<tau>)"
+proof(induct rule: check_s_check_branch_s_check_branch_list.inducts)
+ case (check_valI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v \<tau>' \<tau>)
+ then show ?case using infer_v_wf infer_v_wf subtype_wf wfX_wfY wfS_valI
+ by (metis subtype_eq_base2)
+next
+ case (check_letI x \<Theta> \<Phi> \<B> \<Gamma> \<Delta> e \<tau> z s b c)
+ show ?case proof
+ show \<open> \<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f e : b \<close> using infer_e_wf check_letI b_of.simps by metis
+ show \<open> \<Theta> ; \<Phi> ; \<B> ; (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b_of \<tau> \<close>
+ using check_letI b_of.simps wf_replace_true2(2)[OF check_letI(5)] append_g.simps by metis
+ show \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using infer_e_wf check_letI b_of.simps by metis
+ show \<open>atom x \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>, \<Delta>, e, b_of \<tau>)\<close>
+ apply(simp add: fresh_prodN, intro conjI)
+ using check_letI(1) fresh_prod7 by simp+
+ qed
+next
+ case (check_assertI x \<Theta> \<Phi> \<B> \<Gamma> \<Delta> c \<tau> s)
+ show ?case proof
+
+ show \<open> \<Theta> ; \<Phi> ; \<B> ; (x, B_bool, c) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b_of \<tau> \<close> using check_assertI by auto
+ next
+ show \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f c \<close> using check_assertI by auto
+ next
+ show \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using check_assertI by auto
+ next
+ show \<open>atom x \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>, \<Delta>, c, b_of \<tau>, s)\<close> using check_assertI by auto
+ qed
+next
+ case (check_branch_s_branchI \<Theta> \<B> \<Gamma> \<Delta> \<tau> const x \<Phi> tid cons v s)
+ show ?case proof
+ show \<open> \<Theta> ; \<Phi> ; \<B> ; (x, b_of const, TRUE) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b_of \<tau> \<close>
+ using wf_replace_true append_g.simps check_branch_s_branchI by metis
+ show \<open>atom x \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>, \<Delta>, \<Gamma>, const)\<close>
+ using wf_replace_true append_g.simps check_branch_s_branchI fresh_prodN by metis
+ show \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wf_replace_true append_g.simps check_branch_s_branchI by metis
+ qed
+next
+ case (check_branch_list_consI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid cons const v cs \<tau> dclist css)
+ then show ?case using wf_intros by metis
+next
+ case (check_branch_list_finalI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid cons const v cs \<tau>)
+ then show ?case using wf_intros by metis
+next
+ case (check_ifI z \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v s1 s2 \<tau>)
+ show ?case using wfS_ifI check_v_wf check_ifI b_of.simps by metis
+next
+ case (check_let2I x \<Theta> \<Phi> \<B> G \<Delta> t s1 \<tau> s2)
+ show ?case proof
+ show \<open> \<Theta> ; \<Phi> ; \<B> ; G ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s1 : b_of t \<close> using check_let2I b_of.simps by metis
+ show \<open> \<Theta>; \<B>; G \<turnstile>\<^sub>w\<^sub>f t \<close> using check_let2I check_s_check_branch_s_wf by metis
+ show \<open> \<Theta> ; \<Phi> ; \<B> ; (x, b_of t, TRUE) #\<^sub>\<Gamma> G ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s2 : b_of \<tau> \<close>
+ using check_let2I(5) wf_replace_true2(2) append_g.simps check_let2I by metis
+ show \<open>atom x \<sharp> (\<Phi>, \<Theta>, \<B>, G, \<Delta>, s1, b_of \<tau>, t)\<close>
+ apply(simp add: fresh_prodN, intro conjI)
+ using check_let2I(1) fresh_prod7 by simp+
+ qed
+next
+ case (check_varI u \<Theta> \<Phi> \<B> \<Gamma> \<Delta> \<tau>' v \<tau> s)
+ show ?case proof
+ show \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau>' \<close> using check_v_wf check_varI by metis
+ show \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b_of \<tau>' \<close> using check_v_wf check_varI by metis
+ show \<open>atom u \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>, \<Delta>, \<tau>', v, b_of \<tau>)\<close> using check_varI fresh_prodN u_fresh_b by metis
+ show \<open> \<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; (u, \<tau>') #\<^sub>\<Delta>\<Delta> \<turnstile>\<^sub>w\<^sub>f s : b_of \<tau> \<close> using check_varI by metis
+ qed
+next
+ case (check_assignI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> u \<tau> v z \<tau>')
+ then show ?case using wf_intros check_v_wf subtype_eq_base2 b_of.simps by metis
+next
+ case (check_whileI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> s1 z s2 \<tau>')
+ thus ?case using wf_intros b_of.simps check_v_wf subtype_eq_base2 b_of.simps by metis
+next
+ case (check_seqI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> s1 z s2 \<tau>)
+ thus ?case using wf_intros b_of.simps by metis
+next
+ case (check_caseI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist v cs \<tau> z)
+ show ?case proof
+ show \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : B_id tid \<close> using check_caseI check_v_wf b_of.simps by metis
+ show \<open>AF_typedef tid dclist \<in> set \<Theta>\<close> using check_caseI by metis
+ show \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using check_caseI check_s_check_branch_s_wf by metis
+ show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using check_caseI check_s_check_branch_s_wf by metis
+ show \<open> \<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> ; tid ; dclist \<turnstile>\<^sub>w\<^sub>f cs : b_of \<tau> \<close> using check_caseI by metis
+ qed
+qed
+
+lemma check_s_wf:
+ fixes s::s
+ assumes "\<Theta> ; \<Phi> ; B ; \<Gamma> ; \<Delta> \<turnstile> s \<Leftarrow> \<tau>"
+ shows "\<Theta> ; B \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<and> wfT \<Theta> B \<Gamma> \<tau> \<and> wfPhi \<Theta> \<Phi> \<and> wfTh \<Theta> \<and> wfD \<Theta> B \<Gamma> \<Delta> \<and> wfS \<Theta> \<Phi> B \<Gamma> \<Delta> s (b_of \<tau>)"
+ using check_s_check_branch_s_wf check_s_check_branch_s_wfS assms by meson
+
+lemma check_s_flip_u1:
+ fixes s::s and u::u and u'::u
+ assumes "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile> s \<Leftarrow> \<tau>"
+ shows "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; ( u \<leftrightarrow> u') \<bullet> \<Delta> \<turnstile> (u \<leftrightarrow> u') \<bullet> s \<Leftarrow> \<tau>"
+proof -
+ have "(u \<leftrightarrow> u') \<bullet> \<Theta> ; (u \<leftrightarrow> u') \<bullet> \<Phi> ; (u \<leftrightarrow> u') \<bullet> \<B> ; (u \<leftrightarrow> u') \<bullet> \<Gamma> ; (u \<leftrightarrow> u') \<bullet> \<Delta> \<turnstile> (u \<leftrightarrow> u') \<bullet> s \<Leftarrow> (u \<leftrightarrow> u') \<bullet> \<tau>"
+ using check_s.eqvt assms by blast
+ thus ?thesis using check_s_wf[OF assms] flip_u_eq phi_flip_eq by metis
+qed
+
+lemma check_s_flip_u2:
+ fixes s::s and u::u and u'::u
+ assumes "\<Theta> ; \<Phi> ; B ; \<Gamma> ; ( u \<leftrightarrow> u') \<bullet> \<Delta> \<turnstile> (u \<leftrightarrow> u') \<bullet> s \<Leftarrow> \<tau>"
+ shows "\<Theta> ; \<Phi> ; B ; \<Gamma> ; \<Delta> \<turnstile> s \<Leftarrow> \<tau>"
+proof -
+ have "\<Theta> ; \<Phi> ; B ; \<Gamma> ; ( u \<leftrightarrow> u') \<bullet> ( u \<leftrightarrow> u') \<bullet> \<Delta> \<turnstile> ( u \<leftrightarrow> u') \<bullet> (u \<leftrightarrow> u') \<bullet> s \<Leftarrow> \<tau>"
+ using check_s_flip_u1 assms by blast
+ thus ?thesis using permute_flip_cancel by force
+qed
+
+lemma check_s_flip_u:
+ fixes s::s and u::u and u'::u
+ shows "\<Theta> ; \<Phi> ; B ; \<Gamma> ; ( u \<leftrightarrow> u') \<bullet> \<Delta> \<turnstile> (u \<leftrightarrow> u') \<bullet> s \<Leftarrow> \<tau> = (\<Theta> ; \<Phi> ; B ; \<Gamma> ; \<Delta> \<turnstile> s \<Leftarrow> \<tau>)"
+ using check_s_flip_u1 check_s_flip_u2 by metis
+
+lemma check_s_abs_u:
+ fixes s::s and s'::s and u::u and u'::u and \<tau>'::\<tau>
+ assumes "[[atom u]]lst. s = [[atom u']]lst. s'" and "atom u \<sharp> \<Delta>" and "atom u' \<sharp> \<Delta>"
+ and "\<Theta> ; B ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau>'"
+ and "\<Theta> ; \<Phi> ; B ; \<Gamma> ; ( u , \<tau>') #\<^sub>\<Delta>\<Delta> \<turnstile> s \<Leftarrow> \<tau>"
+ shows "\<Theta> ; \<Phi> ; B ; \<Gamma> ; ( u', \<tau>' ) #\<^sub>\<Delta>\<Delta> \<turnstile> s' \<Leftarrow> \<tau>"
+proof -
+ have "\<Theta> ; \<Phi> ; B ; \<Gamma> ; ( u' \<leftrightarrow> u) \<bullet> (( u , \<tau>') #\<^sub>\<Delta>\<Delta>) \<turnstile> ( u' \<leftrightarrow> u) \<bullet> s \<Leftarrow> \<tau>"
+ using assms check_s_flip_u by metis
+ moreover have " ( u' \<leftrightarrow> u) \<bullet> (( u , \<tau>') #\<^sub>\<Delta> \<Delta>) = ( u' , \<tau>') #\<^sub>\<Delta>\<Delta>" proof -
+ have " ( u' \<leftrightarrow> u) \<bullet> (( u , \<tau>') #\<^sub>\<Delta> \<Delta>) = ((u' \<leftrightarrow> u) \<bullet> u, (u' \<leftrightarrow> u) \<bullet> \<tau>') #\<^sub>\<Delta> (u' \<leftrightarrow> u) \<bullet> \<Delta>"
+ using DCons_eqvt Pair_eqvt by auto
+ also have "... = ( u' , (u' \<leftrightarrow> u) \<bullet> \<tau>') #\<^sub>\<Delta> \<Delta>"
+ using assms flip_fresh_fresh by auto
+ also have "... = ( u' , \<tau>') #\<^sub>\<Delta> \<Delta>" using
+ u_not_in_t fresh_def flip_fresh_fresh assms by metis
+ finally show ?thesis by auto
+ qed
+ moreover have "( u' \<leftrightarrow> u) \<bullet> s = s'" using assms Abs1_eq_iff(3)[of u' s' u s] by auto
+ ultimately show ?thesis by auto
+qed
+
+section \<open>Additional Elimination and Intros\<close>
+
+subsection \<open>Values\<close>
+
+nominal_function b_for :: "opp \<Rightarrow> b" where
+ "b_for Plus = B_int"
+| "b_for LEq = B_bool" | "b_for Eq = B_bool"
+ apply(auto,simp add: eqvt_def b_for_graph_aux_def )
+ by (meson opp.exhaust)
+nominal_termination (eqvt) by lexicographic_order
+
+
+lemma infer_v_pair2I:
+ fixes v\<^sub>1::v and v\<^sub>2::v
+ assumes "\<Theta>; \<B>; \<Gamma> \<turnstile> v\<^sub>1 \<Rightarrow> \<tau>\<^sub>1" and "\<Theta>; \<B>; \<Gamma> \<turnstile> v\<^sub>2 \<Rightarrow> \<tau>\<^sub>2"
+ shows "\<exists>\<tau>. \<Theta>; \<B>; \<Gamma> \<turnstile> V_pair v\<^sub>1 v\<^sub>2 \<Rightarrow> \<tau> \<and> b_of \<tau> = B_pair (b_of \<tau>\<^sub>1) (b_of \<tau>\<^sub>2)"
+proof -
+ obtain z1 and b1 and c1 and z2 and b2 and c2 where zbc: "\<tau>\<^sub>1 = (\<lbrace> z1 : b1 | c1 \<rbrace>) \<and> \<tau>\<^sub>2 = (\<lbrace> z2 : b2 | c2 \<rbrace>)"
+ using \<tau>.exhaust by meson
+ obtain z::x where "atom z \<sharp> ( v\<^sub>1, v\<^sub>2,\<Theta>, \<B>,\<Gamma>)" using obtain_fresh
+ by blast
+ hence "atom z \<sharp> ( v\<^sub>1, v\<^sub>2) \<and> atom z \<sharp> (\<Theta>, \<B>,\<Gamma>)" using fresh_prodN by metis
+ hence " \<Theta>; \<B>; \<Gamma> \<turnstile> V_pair v\<^sub>1 v\<^sub>2 \<Rightarrow> \<lbrace> z : [ b_of \<tau>\<^sub>1 , b_of \<tau>\<^sub>2 ]\<^sup>b | CE_val (V_var z) == CE_val (V_pair v\<^sub>1 v\<^sub>2) \<rbrace>"
+ using assms infer_v_pairI zbc by metis
+ moreover obtain \<tau> where "\<tau> = (\<lbrace> z : B_pair b1 b2 | CE_val (V_var z) == CE_val (V_pair v\<^sub>1 v\<^sub>2) \<rbrace>)" by blast
+ moreover hence "b_of \<tau> = B_pair (b_of \<tau>\<^sub>1) (b_of \<tau>\<^sub>2)" using b_of.simps zbc by presburger
+ ultimately show ?thesis using b_of.simps by metis
+qed
+
+lemma infer_v_pair2I_zbc:
+ fixes v\<^sub>1::v and v\<^sub>2::v
+ assumes "\<Theta>; \<B>; \<Gamma> \<turnstile> v\<^sub>1 \<Rightarrow> \<tau>\<^sub>1" and "\<Theta>; \<B>; \<Gamma> \<turnstile> v\<^sub>2 \<Rightarrow> \<tau>\<^sub>2"
+ shows "\<exists>z \<tau>. \<Theta>; \<B>; \<Gamma> \<turnstile> V_pair v\<^sub>1 v\<^sub>2 \<Rightarrow> \<tau> \<and> \<tau> = (\<lbrace> z : B_pair (b_of \<tau>\<^sub>1) (b_of \<tau>\<^sub>2) | C_eq (CE_val (V_var z)) (CE_val (V_pair v\<^sub>1 v\<^sub>2)) \<rbrace>) \<and> atom z \<sharp> (v\<^sub>1,v\<^sub>2) \<and> atom z \<sharp> \<Gamma>"
+proof -
+ obtain z1 and b1 and c1 and z2 and b2 and c2 where zbc: "\<tau>\<^sub>1 = (\<lbrace> z1 : b1 | c1 \<rbrace>) \<and> \<tau>\<^sub>2 = (\<lbrace> z2 : b2 | c2 \<rbrace>)"
+ using \<tau>.exhaust by meson
+ obtain z::x where * : "atom z \<sharp> ( v\<^sub>1, v\<^sub>2,\<Gamma>,\<Theta> , \<B> )" using obtain_fresh
+ by blast
+ hence vinf: " \<Theta>; \<B>; \<Gamma> \<turnstile> V_pair v\<^sub>1 v\<^sub>2 \<Rightarrow> \<lbrace> z :[ b_of \<tau>\<^sub>1 , b_of \<tau>\<^sub>2 ]\<^sup>b | CE_val (V_var z) == CE_val (V_pair v\<^sub>1 v\<^sub>2) \<rbrace>"
+ using assms infer_v_pairI by simp
+ moreover obtain \<tau> where "\<tau> = (\<lbrace> z : B_pair b1 b2 | CE_val (V_var z) == CE_val (V_pair v\<^sub>1 v\<^sub>2) \<rbrace>)" by blast
+ moreover have "b_of \<tau>\<^sub>1 = b1 \<and> b_of \<tau>\<^sub>2 = b2" using zbc b_of.simps by auto
+ ultimately have "\<Theta>; \<B>; \<Gamma> \<turnstile> V_pair v\<^sub>1 v\<^sub>2 \<Rightarrow> \<tau> \<and> \<tau> = (\<lbrace> z : B_pair (b_of \<tau>\<^sub>1) (b_of \<tau>\<^sub>2) | CE_val (V_var z) == CE_val (V_pair v\<^sub>1 v\<^sub>2) \<rbrace>)" by auto
+ thus ?thesis using * fresh_prod2 fresh_prod3 by metis
+qed
+
+lemma infer_v_pair2E:
+ assumes "\<Theta>; \<B>; \<Gamma> \<turnstile> V_pair v\<^sub>1 v\<^sub>2 \<Rightarrow> \<tau>"
+ shows "\<exists>\<tau>\<^sub>1 \<tau>\<^sub>2 z . \<Theta>; \<B>; \<Gamma> \<turnstile> v\<^sub>1 \<Rightarrow> \<tau>\<^sub>1 \<and> \<Theta>; \<B>; \<Gamma> \<turnstile> v\<^sub>2 \<Rightarrow> \<tau>\<^sub>2 \<and>
+ \<tau> = (\<lbrace> z : B_pair (b_of \<tau>\<^sub>1) (b_of \<tau>\<^sub>2) | C_eq (CE_val (V_var z)) (CE_val (V_pair v\<^sub>1 v\<^sub>2)) \<rbrace>) \<and> atom z \<sharp> (v\<^sub>1, v\<^sub>2)"
+proof -
+ obtain z and t1 and t2 where
+ "\<tau> = (\<lbrace> z : B_pair (b_of t1) (b_of t2) | CE_val (V_var z) == CE_val (V_pair v\<^sub>1 v\<^sub>2) \<rbrace>) \<and>
+ atom z \<sharp> (v\<^sub>1, v\<^sub>2) \<and> \<Theta>; \<B>; \<Gamma> \<turnstile> v\<^sub>1 \<Rightarrow> t1 \<and> \<Theta>; \<B>; \<Gamma> \<turnstile> v\<^sub>2 \<Rightarrow> t2 " using infer_v_elims(3)[OF assms ] by metis
+ thus ?thesis using b_of.simps by metis
+qed
+
+subsection \<open>Expressions\<close>
+
+lemma infer_e_app2E:
+ fixes \<Phi>::\<Phi> and \<Theta>::\<Theta>
+ assumes "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile> AE_app f v \<Rightarrow> \<tau>"
+ shows "\<exists>x b c s' \<tau>'. wfD \<Theta> \<B> \<Gamma> \<Delta> \<and> Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \<tau>' s'))) = lookup_fun \<Phi> f \<and> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<and>
+ \<Theta>; \<B>; \<Gamma> \<turnstile> v \<Leftarrow> \<lbrace> x : b | c \<rbrace> \<and> \<tau> = \<tau>'[x::=v]\<^sub>\<tau>\<^sub>v \<and> atom x \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>, \<Delta>, v, \<tau>)"
+ using infer_e_elims(6)[OF assms] b_of.simps subst_defs by metis
+
+lemma infer_e_appP2E:
+ fixes \<Phi>::\<Phi> and \<Theta>::\<Theta>
+ assumes "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile> AE_appP f b v \<Rightarrow> \<tau>"
+ shows "\<exists>bv x ba c s' \<tau>'. wfD \<Theta> \<B> \<Gamma> \<Delta> \<and> Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x ba c \<tau>' s'))) = lookup_fun \<Phi> f \<and> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<and> \<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f b \<and>
+ (\<Theta>; \<B>; \<Gamma> \<turnstile> v \<Leftarrow> \<lbrace> x : ba[bv::=b]\<^sub>b\<^sub>b | c[bv::=b]\<^sub>c\<^sub>b \<rbrace>) \<and> (\<tau> = \<tau>'[bv::=b]\<^sub>\<tau>\<^sub>b[x::=v]\<^sub>\<tau>\<^sub>v) \<and> atom x \<sharp> \<Gamma> \<and> atom bv \<sharp> v"
+proof -
+ obtain bv x ba c s' \<tau>' where *:"wfD \<Theta> \<B> \<Gamma> \<Delta> \<and> Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x ba c \<tau>' s'))) = lookup_fun \<Phi> f \<and> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<and> \<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f b \<and>
+ (\<Theta>; \<B>; \<Gamma> \<turnstile> v \<Leftarrow> \<lbrace> x : ba[bv::=b]\<^sub>b\<^sub>b | c[bv::=b]\<^sub>c\<^sub>b \<rbrace>) \<and> (\<tau> = \<tau>'[bv::=b]\<^sub>\<tau>\<^sub>b[x::=v]\<^sub>\<tau>\<^sub>v) \<and> atom x \<sharp> \<Gamma> \<and> atom bv \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>, \<Delta>, b, v, \<tau>)"
+ using infer_e_elims(21)[OF assms] subst_defs by metis
+ moreover then have "atom bv \<sharp> v" using fresh_prodN by metis
+ ultimately show ?thesis by metis
+qed
+
+section \<open>Weakening\<close>
+
+text \<open> Lemmas showing that typing judgements hold when a context is extended \<close>
+
+lemma subtype_weakening:
+ fixes \<Gamma>'::\<Gamma>
+ assumes "\<Theta>; \<B>; \<Gamma> \<turnstile> \<tau>1 \<lesssim> \<tau>2" and "toSet \<Gamma> \<subseteq> toSet \<Gamma>'" and "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>'"
+ shows "\<Theta>; \<B>; \<Gamma>' \<turnstile> \<tau>1 \<lesssim> \<tau>2"
+ using assms proof(nominal_induct \<tau>2 avoiding: \<Gamma>' rule: subtype.strong_induct)
+
+ case (subtype_baseI x \<Theta> \<B> \<Gamma> z c z' c' b)
+ show ?case proof
+ show *:"\<Theta>; \<B>; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b | c \<rbrace>" using wfT_weakening subtype_baseI by metis
+ show "\<Theta>; \<B>; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f \<lbrace> z' : b | c' \<rbrace>" using wfT_weakening subtype_baseI by metis
+ show "atom x \<sharp> (\<Theta>, \<B>, \<Gamma>', z , c , z' , c' )" using subtype_baseI fresh_Pair by metis
+ have "toSet ((x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) \<subseteq> toSet ((x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>')" using subtype_baseI toSet.simps by blast
+ moreover have "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f (x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>'" using wfT_wf_cons3[OF *, of x] subtype_baseI fresh_Pair subst_defs by metis
+ ultimately show "\<Theta>; \<B>; (x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>' \<Turnstile> c'[z'::=V_var x]\<^sub>v" using valid_weakening subtype_baseI by metis
+ qed
+qed
+
+method many_rules uses add = ( (rule+),((simp add: add)+)?)
+
+lemma infer_v_g_weakening:
+ fixes e::e and \<Gamma>'::\<Gamma> and v::v
+ assumes "\<Theta>; \<B> ; \<Gamma> \<turnstile> v \<Rightarrow> \<tau>" and "toSet \<Gamma> \<subseteq> toSet \<Gamma>'" and "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>'"
+ shows "\<Theta>; \<B> ; \<Gamma>' \<turnstile> v \<Rightarrow> \<tau>"
+ using assms proof(nominal_induct avoiding: \<Gamma>' rule: infer_v.strong_induct)
+ case (infer_v_varI \<Theta> \<B> \<Gamma> b c x' z)
+ show ?case proof
+ show \<open> \<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>' \<close> using infer_v_varI by auto
+ show \<open>Some (b, c) = lookup \<Gamma>' x'\<close> using infer_v_varI lookup_weakening by metis
+ show \<open>atom z \<sharp> x'\<close> using infer_v_varI by auto
+ show \<open>atom z \<sharp> (\<Theta>, \<B>, \<Gamma>')\<close> using infer_v_varI by auto
+ qed
+next
+ case (infer_v_litI \<Theta> \<B> \<Gamma> l \<tau>)
+ then show ?case using infer_v.intros by simp
+next
+ case (infer_v_pairI z v1 v2 \<Theta> \<B> \<Gamma> t1 t2)
+ then show ?case using infer_v.intros by simp
+next
+ case (infer_v_consI s dclist \<Theta> dc tc \<B> \<Gamma> v tv z)
+ show ?case proof
+ show \<open>AF_typedef s dclist \<in> set \<Theta>\<close> using infer_v_consI by auto
+ show \<open>(dc, tc) \<in> set dclist\<close> using infer_v_consI by auto
+ show \<open> \<Theta>; \<B>; \<Gamma>' \<turnstile> v \<Rightarrow> tv\<close> using infer_v_consI by auto
+ show \<open>\<Theta>; \<B>; \<Gamma>' \<turnstile> tv \<lesssim> tc\<close> using infer_v_consI subtype_weakening by auto
+ show \<open>atom z \<sharp> v\<close> using infer_v_consI by auto
+ show \<open>atom z \<sharp> (\<Theta>, \<B>, \<Gamma>')\<close> using infer_v_consI by auto
+ qed
+next
+ case (infer_v_conspI s bv dclist \<Theta> dc tc \<B> \<Gamma> v tv b z)
+ show ?case proof
+ show \<open>AF_typedef_poly s bv dclist \<in> set \<Theta>\<close> using infer_v_conspI by auto
+ show \<open>(dc, tc) \<in> set dclist\<close> using infer_v_conspI by auto
+ show \<open> \<Theta>; \<B>; \<Gamma>' \<turnstile> v \<Rightarrow> tv\<close> using infer_v_conspI by auto
+ show \<open>\<Theta>; \<B>; \<Gamma>' \<turnstile> tv \<lesssim> tc[bv::=b]\<^sub>\<tau>\<^sub>b\<close> using infer_v_conspI subtype_weakening by auto
+ show \<open>atom z \<sharp> (\<Theta>, \<B>, \<Gamma>', v, b)\<close> using infer_v_conspI by auto
+ show \<open>atom bv \<sharp> (\<Theta>, \<B>, \<Gamma>', v, b)\<close> using infer_v_conspI by auto
+ show \<open> \<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f b \<close> using infer_v_conspI by auto
+ qed
+qed
+
+lemma check_v_g_weakening:
+ fixes e::e and \<Gamma>'::\<Gamma>
+ assumes "\<Theta>; \<B> ; \<Gamma> \<turnstile> v \<Leftarrow> \<tau>" and "toSet \<Gamma> \<subseteq> toSet \<Gamma>'" and "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>'"
+ shows "\<Theta>; \<B> ; \<Gamma>' \<turnstile> v \<Leftarrow> \<tau>"
+ using subtype_weakening infer_v_g_weakening check_v_elims check_v_subtypeI assms by metis
+
+lemma infer_e_g_weakening:
+ fixes e::e and \<Gamma>'::\<Gamma>
+ assumes "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile> e \<Rightarrow> \<tau>" and "toSet \<Gamma> \<subseteq> toSet \<Gamma>'" and "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>'"
+ shows "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>'; \<Delta> \<turnstile> e \<Rightarrow> \<tau>"
+ using assms proof(nominal_induct \<tau> avoiding: \<Gamma>' rule: infer_e.strong_induct)
+ case (infer_e_valI \<Theta> \<B> \<Gamma> \<Delta>' \<Phi> v \<tau>)
+ then show ?case using infer_v_g_weakening wf_weakening infer_e.intros by metis
+next
+ case (infer_e_plusI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v1 z1 c1 v2 z2 c2 z3)
+
+ obtain z'::x where z': "atom z' \<sharp> v1 \<and> atom z' \<sharp> v2 \<and> atom z' \<sharp> \<Gamma>'" using obtain_fresh fresh_prod3 by metis
+ moreover hence *:"\<lbrace> z3 : B_int | CE_val (V_var z3) == CE_op Plus [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \<rbrace> = (\<lbrace> z' : B_int | CE_val (V_var z') == CE_op Plus [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \<rbrace>)"
+ using infer_e_plusI type_e_eq ce.fresh fresh_e_opp by auto
+
+ have "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>' ; \<Delta> \<turnstile> AE_op Plus v1 v2 \<Rightarrow> \<lbrace> z' : B_int | CE_val (V_var z') == CE_op Plus [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \<rbrace>" proof
+ show \<open> \<Theta> ; \<B> ; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wf_weakening infer_e_plusI by auto
+ show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using infer_e_plusI by auto
+ show \<open> \<Theta> ; \<B> ; \<Gamma>' \<turnstile> v1 \<Rightarrow> \<lbrace> z1 : B_int | c1 \<rbrace>\<close> using infer_v_g_weakening infer_e_plusI by auto
+ show \<open> \<Theta> ; \<B> ; \<Gamma>' \<turnstile> v2 \<Rightarrow> \<lbrace> z2 : B_int | c2 \<rbrace>\<close> using infer_v_g_weakening infer_e_plusI by auto
+ show \<open>atom z' \<sharp> AE_op Plus v1 v2\<close> using z' by auto
+ show \<open>atom z' \<sharp> \<Gamma>'\<close> using z' by auto
+ qed
+ thus ?case using * by metis
+
+next
+ case (infer_e_leqI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v1 z1 c1 v2 z2 c2 z3)
+ obtain z'::x where z': "atom z' \<sharp> v1 \<and> atom z' \<sharp> v2 \<and> atom z' \<sharp> \<Gamma>'" using obtain_fresh fresh_prod3 by metis
+
+ moreover hence *:"\<lbrace> z3 : B_bool | CE_val (V_var z3) == CE_op LEq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \<rbrace> = (\<lbrace> z' : B_bool | CE_val (V_var z') == CE_op LEq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \<rbrace>)"
+ using infer_e_leqI type_e_eq ce.fresh fresh_e_opp by auto
+
+ have "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>' ; \<Delta> \<turnstile> AE_op LEq v1 v2 \<Rightarrow> \<lbrace> z' : B_bool | CE_val (V_var z') == CE_op LEq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \<rbrace>" proof
+ show \<open> \<Theta> ; \<B> ; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wf_weakening infer_e_leqI by auto
+ show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using infer_e_leqI by auto
+ show \<open> \<Theta> ; \<B> ; \<Gamma>' \<turnstile> v1 \<Rightarrow> \<lbrace> z1 : B_int | c1 \<rbrace>\<close> using infer_v_g_weakening infer_e_leqI by auto
+ show \<open> \<Theta> ; \<B> ; \<Gamma>' \<turnstile> v2 \<Rightarrow> \<lbrace> z2 : B_int | c2 \<rbrace>\<close> using infer_v_g_weakening infer_e_leqI by auto
+ show \<open>atom z' \<sharp> AE_op LEq v1 v2\<close> using z' by auto
+ show \<open>atom z' \<sharp> \<Gamma>'\<close> using z' by auto
+ qed
+ thus ?case using * by metis
+next
+ case (infer_e_eqI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v1 z1 bb c1 v2 z2 c2 z3)
+ obtain z'::x where z': "atom z' \<sharp> v1 \<and> atom z' \<sharp> v2 \<and> atom z' \<sharp> \<Gamma>'" using obtain_fresh fresh_prod3 by metis
+
+ moreover hence *:"\<lbrace> z3 : B_bool | CE_val (V_var z3) == CE_op Eq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \<rbrace> = (\<lbrace> z' : B_bool | CE_val (V_var z') == CE_op Eq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \<rbrace>)"
+ using infer_e_eqI type_e_eq ce.fresh fresh_e_opp by auto
+
+ have "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>' ; \<Delta> \<turnstile> AE_op Eq v1 v2 \<Rightarrow> \<lbrace> z' : B_bool | CE_val (V_var z') == CE_op Eq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \<rbrace>" proof
+ show \<open> \<Theta> ; \<B> ; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wf_weakening infer_e_eqI by auto
+ show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using infer_e_eqI by auto
+ show \<open> \<Theta> ; \<B> ; \<Gamma>' \<turnstile> v1 \<Rightarrow> \<lbrace> z1 : bb | c1 \<rbrace>\<close> using infer_v_g_weakening infer_e_eqI by auto
+ show \<open> \<Theta> ; \<B> ; \<Gamma>' \<turnstile> v2 \<Rightarrow> \<lbrace> z2 : bb | c2 \<rbrace>\<close> using infer_v_g_weakening infer_e_eqI by auto
+ show \<open>atom z' \<sharp> AE_op Eq v1 v2\<close> using z' by auto
+ show \<open>atom z' \<sharp> \<Gamma>'\<close> using z' by auto
+ show "bb \<in> {B_bool, B_int, B_unit}" using infer_e_eqI by auto
+ qed
+ thus ?case using * by metis
+next
+ case (infer_e_appI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> f x b c \<tau>' s' v \<tau>)
+ show ?case proof
+ show "\<Theta>; \<B>; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f \<Delta> " using wf_weakening infer_e_appI by auto
+ show "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>" using wf_weakening infer_e_appI by auto
+ show "Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \<tau>' s'))) = lookup_fun \<Phi> f" using wf_weakening infer_e_appI by auto
+ show "\<Theta>; \<B>; \<Gamma>' \<turnstile> v \<Leftarrow> \<lbrace> x : b | c \<rbrace>" using wf_weakening infer_e_appI check_v_g_weakening by auto
+ show "atom x \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>', \<Delta>, v, \<tau>)" using wf_weakening infer_e_appI by auto
+ show "\<tau>'[x::=v]\<^sub>v = \<tau>" using wf_weakening infer_e_appI by auto
+ qed
+
+next
+ case (infer_e_appPI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> b' f bv x b c \<tau>' s' v \<tau>)
+
+ hence *:"\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile> AE_appP f b' v \<Rightarrow> \<tau>" using Typing.infer_e_appPI by auto
+
+ obtain x'::x where x':"atom x' \<sharp> (s', c, \<tau>', (\<Gamma>',v,\<tau>)) \<and> (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \<tau>' s'))) = (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x' b ((x' \<leftrightarrow> x) \<bullet> c) ((x' \<leftrightarrow> x) \<bullet> \<tau>') ((x' \<leftrightarrow> x) \<bullet> s'))))"
+ using obtain_fresh_fun_def[of s' c \<tau>' "(\<Gamma>',v,\<tau>)" f x b]
+ by (metis fun_def.eq_iff fun_typ_q.eq_iff(2))
+
+ hence **: " \<lbrace> x : b | c \<rbrace> = \<lbrace> x' : b | (x' \<leftrightarrow> x) \<bullet> c \<rbrace>"
+ using fresh_PairD(1) fresh_PairD(2) type_eq_flip by blast
+ have "atom x' \<sharp> \<Gamma>" using x' infer_e_appPI fresh_weakening fresh_Pair by metis
+
+ show ?case proof(rule Typing.infer_e_appPI[where x=x' and bv=bv and b=b and c="(x' \<leftrightarrow> x) \<bullet> c" and \<tau>'="(x' \<leftrightarrow> x) \<bullet> \<tau>'"and s'="((x' \<leftrightarrow> x) \<bullet> s')"])
+ show \<open> \<Theta> ; \<B> ; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wf_weakening infer_e_appPI by auto
+ show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using wf_weakening infer_e_appPI by auto
+ show "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f b'" using infer_e_appPI by auto
+ show "Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x' b ((x' \<leftrightarrow> x) \<bullet> c) ((x' \<leftrightarrow> x) \<bullet> \<tau>') ((x' \<leftrightarrow> x) \<bullet> s')))) = lookup_fun \<Phi> f" using x' infer_e_appPI by argo
+ show "\<Theta>; \<B>; \<Gamma>' \<turnstile> v \<Leftarrow> \<lbrace> x' : b[bv::=b']\<^sub>b | ((x' \<leftrightarrow> x) \<bullet> c)[bv::=b']\<^sub>b \<rbrace>" using **
+ \<tau>.eq_iff check_v_g_weakening infer_e_appPI.hyps infer_e_appPI.prems(1) infer_e_appPI.prems subst_defs
+ subst_tb.simps by metis
+ show "atom x' \<sharp> \<Gamma>'" using x' fresh_prodN by metis
+
+ have "atom x \<sharp> (v, \<tau>) \<and> atom x' \<sharp> (v, \<tau>)" using x' infer_e_fresh[OF *] e.fresh(2) fresh_Pair infer_e_appPI \<open>atom x' \<sharp> \<Gamma>\<close> e.fresh by metis
+ moreover then have "((x' \<leftrightarrow> x) \<bullet> \<tau>')[bv::=b']\<^sub>\<tau>\<^sub>b = (x' \<leftrightarrow> x) \<bullet> (\<tau>'[bv::=b']\<^sub>\<tau>\<^sub>b)" using subst_b_x_flip
+ by (metis subst_b_\<tau>_def)
+ ultimately show "((x' \<leftrightarrow> x) \<bullet> \<tau>')[bv::=b']\<^sub>b[x'::=v]\<^sub>v = \<tau>"
+ using infer_e_appPI subst_tv_flip subst_defs by metis
+
+ show "atom bv \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>', \<Delta>, b', v, \<tau>)" using infer_e_appPI fresh_prodN by metis
+ qed
+
+next
+ case (infer_e_fstI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v z'' b1 b2 c z)
+
+ obtain z'::x where *: "atom z' \<sharp> \<Gamma>' \<and> atom z' \<sharp> v \<and> atom z' \<sharp> c" using obtain_fresh_z fresh_Pair by metis
+ hence **:"\<lbrace> z : b1 | CE_val (V_var z) == CE_fst [v]\<^sup>c\<^sup>e \<rbrace> = \<lbrace> z' : b1 | CE_val (V_var z') == CE_fst [v]\<^sup>c\<^sup>e \<rbrace>"
+ using type_e_eq infer_e_fstI v.fresh e.fresh ce.fresh obtain_fresh_z fresh_Pair by metis
+
+ have "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>' ; \<Delta> \<turnstile> AE_fst v \<Rightarrow> \<lbrace> z' : b1 | CE_val (V_var z') == CE_fst [v]\<^sup>c\<^sup>e \<rbrace>" proof
+ show \<open> \<Theta> ; \<B> ; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wf_weakening infer_e_fstI by auto
+ show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using wf_weakening infer_e_fstI by auto
+ show "\<Theta> ; \<B> ; \<Gamma>' \<turnstile> v \<Rightarrow> \<lbrace> z'' : B_pair b1 b2 | c \<rbrace>" using infer_v_g_weakening infer_e_fstI by metis
+ show "atom z' \<sharp> AE_fst v" using * ** e.supp by auto
+ show "atom z' \<sharp> \<Gamma>'" using * by auto
+ qed
+ thus ?case using * ** by metis
+next
+ case (infer_e_sndI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v z'' b1 b2 c z)
+
+ obtain z'::x where *: "atom z' \<sharp> \<Gamma>' \<and> atom z' \<sharp> v \<and> atom z' \<sharp> c" using obtain_fresh_z fresh_Pair by metis
+ hence **:"\<lbrace> z : b2 | CE_val (V_var z) == CE_snd [v]\<^sup>c\<^sup>e \<rbrace> = \<lbrace> z' : b2 | CE_val (V_var z') == CE_snd [v]\<^sup>c\<^sup>e \<rbrace>"
+ using type_e_eq infer_e_sndI e.fresh ce.fresh obtain_fresh_z fresh_Pair by metis
+
+ have "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>' ; \<Delta> \<turnstile> AE_snd v \<Rightarrow> \<lbrace> z' : b2 | CE_val (V_var z') == CE_snd [v]\<^sup>c\<^sup>e \<rbrace>" proof
+ show \<open> \<Theta> ; \<B> ;\<Gamma>' \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wf_weakening infer_e_sndI by auto
+ show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using wf_weakening infer_e_sndI by auto
+ show "\<Theta> ; \<B> ; \<Gamma>' \<turnstile> v \<Rightarrow> \<lbrace> z'' : B_pair b1 b2 | c \<rbrace>" using infer_v_g_weakening infer_e_sndI by metis
+ show "atom z' \<sharp> AE_snd v" using * e.supp by auto
+ show "atom z' \<sharp> \<Gamma>'" using * by auto
+ qed
+ thus ?case using ** by metis
+next
+ case (infer_e_lenI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v z'' c z)
+
+ obtain z'::x where *: "atom z' \<sharp> \<Gamma>' \<and> atom z' \<sharp> v \<and> atom z' \<sharp> c" using obtain_fresh_z fresh_Pair by metis
+ hence **:"\<lbrace> z : B_int | CE_val (V_var z) == CE_len [v]\<^sup>c\<^sup>e \<rbrace> = \<lbrace> z' : B_int | CE_val (V_var z') == CE_len [v]\<^sup>c\<^sup>e \<rbrace>"
+ using type_e_eq infer_e_lenI e.fresh ce.fresh obtain_fresh_z fresh_Pair by metis
+
+ have "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>' ; \<Delta> \<turnstile> AE_len v \<Rightarrow> \<lbrace> z' : B_int | CE_val (V_var z') == CE_len [v]\<^sup>c\<^sup>e \<rbrace>" proof
+ show \<open> \<Theta>; \<B>; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wf_weakening infer_e_lenI by auto
+ show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using wf_weakening infer_e_lenI by auto
+ show "\<Theta> ; \<B> ; \<Gamma>' \<turnstile> v \<Rightarrow> \<lbrace> z'' : B_bitvec | c \<rbrace>" using infer_v_g_weakening infer_e_lenI by metis
+ show "atom z' \<sharp> AE_len v" using * e.supp by auto
+ show "atom z' \<sharp> \<Gamma>'" using * by auto
+ qed
+ thus ?case using * ** by metis
+next
+ case (infer_e_mvarI \<Theta> \<Gamma> \<Phi> \<Delta> u \<tau>)
+ then show ?case using wf_weakening infer_e.intros by metis
+next
+ case (infer_e_concatI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v1 z1 c1 v2 z2 c2 z3)
+
+ obtain z'::x where *: "atom z' \<sharp> \<Gamma>' \<and> atom z' \<sharp> v1 \<and> atom z' \<sharp> v2" using obtain_fresh_z fresh_Pair by metis
+ hence **:"\<lbrace> z3 : B_bitvec | CE_val (V_var z3) == CE_concat [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \<rbrace> = \<lbrace> z' : B_bitvec | CE_val (V_var z') == CE_concat [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \<rbrace>"
+ using type_e_eq infer_e_concatI e.fresh ce.fresh obtain_fresh_z fresh_Pair by metis
+
+ have "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>' ; \<Delta> \<turnstile> AE_concat v1 v2 \<Rightarrow> \<lbrace> z' : B_bitvec | CE_val (V_var z') == CE_concat [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \<rbrace>" proof
+ show \<open> \<Theta>; \<B>; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wf_weakening infer_e_concatI by auto
+ show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using wf_weakening infer_e_concatI by auto
+ show "\<Theta> ; \<B> ; \<Gamma>' \<turnstile> v1 \<Rightarrow> \<lbrace> z1 : B_bitvec | c1 \<rbrace>" using infer_v_g_weakening infer_e_concatI by metis
+ show "\<Theta> ; \<B> ; \<Gamma>' \<turnstile> v2 \<Rightarrow> \<lbrace> z2 : B_bitvec | c2 \<rbrace>" using infer_v_g_weakening infer_e_concatI by metis
+ show "atom z' \<sharp> AE_concat v1 v2" using * e.supp by auto
+ show "atom z' \<sharp> \<Gamma>'" using * by auto
+ qed
+ thus ?case using * ** by metis
+next
+ case (infer_e_splitI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v1 z1 c1 v2 z2 z3)
+ show ?case proof
+ show "\<Theta>; \<B>; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f \<Delta>" using infer_e_splitI wf_weakening by auto
+ show "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> " using infer_e_splitI wf_weakening by auto
+ show "\<Theta>; \<B>; \<Gamma>' \<turnstile> v1 \<Rightarrow> \<lbrace> z1 : B_bitvec | c1 \<rbrace>" using infer_v_g_weakening infer_e_splitI by metis
+ show "\<Theta>; \<B>; \<Gamma>' \<turnstile> v2 \<Leftarrow> \<lbrace> z2 : B_int | [ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ z2 ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e
+ AND [ leq [ [ z2 ]\<^sup>v ]\<^sup>c\<^sup>e [| [ v1 ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrace>"
+ using check_v_g_weakening infer_e_splitI by metis
+ show "atom z1 \<sharp> AE_split v1 v2" using infer_e_splitI by auto
+ show "atom z1 \<sharp> \<Gamma>'" using infer_e_splitI by auto
+ show "atom z2 \<sharp> AE_split v1 v2" using infer_e_splitI by auto
+ show "atom z2 \<sharp> \<Gamma>'" using infer_e_splitI by auto
+ show "atom z3 \<sharp> AE_split v1 v2" using infer_e_splitI by auto
+ show "atom z3 \<sharp> \<Gamma>'" using infer_e_splitI by auto
+ qed
+qed
+
+text \<open> Special cases proved explicitly, other cases at the end with method + \<close>
+
+lemma infer_e_d_weakening:
+ fixes e::e
+ assumes "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile> e \<Rightarrow> \<tau>" and "setD \<Delta> \<subseteq> setD \<Delta>'" and "wfD \<Theta> \<B> \<Gamma> \<Delta>'"
+ shows "\<Theta>; \<Phi> ; \<B> ; \<Gamma> ; \<Delta>' \<turnstile> e \<Rightarrow> \<tau>"
+ using assms by(nominal_induct \<tau> avoiding: \<Delta>' rule: infer_e.strong_induct,auto simp add:infer_e.intros)
+
+lemma wfG_x_fresh_in_v_simple:
+ fixes x::x and v :: v
+ assumes "\<Theta>; \<B>; \<Gamma> \<turnstile> v \<Rightarrow> \<tau>" and "atom x \<sharp> \<Gamma>"
+ shows "atom x \<sharp> v"
+ using wfV_x_fresh infer_v_wf assms by metis
+
+lemma check_s_g_weakening:
+ fixes v::v and s::s and cs::branch_s and x::x and c::c and b::b and \<Gamma>'::\<Gamma> and \<Theta>::\<Theta> and css::branch_list
+ shows "check_s \<Theta> \<Phi> \<B> \<Gamma> \<Delta> s t \<Longrightarrow> toSet \<Gamma> \<subseteq> toSet \<Gamma>' \<Longrightarrow> \<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>' \<Longrightarrow> check_s \<Theta> \<Phi> \<B> \<Gamma>' \<Delta> s t" and
+ "check_branch_s \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid cons const v cs t \<Longrightarrow> toSet \<Gamma> \<subseteq> toSet \<Gamma>' \<Longrightarrow> \<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>' \<Longrightarrow> check_branch_s \<Theta> \<Phi> \<B> \<Gamma>' \<Delta> tid cons const v cs t" and
+ "check_branch_list \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist v css t \<Longrightarrow> toSet \<Gamma> \<subseteq> toSet \<Gamma>' \<Longrightarrow> \<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>' \<Longrightarrow> check_branch_list \<Theta> \<Phi> \<B> \<Gamma>' \<Delta> tid dclist v css t"
+proof(nominal_induct t and t and t avoiding: \<Gamma>' rule: check_s_check_branch_s_check_branch_list.strong_induct)
+ case (check_valI \<Theta> \<B> \<Gamma> \<Delta>' \<Phi> v \<tau>' \<tau>)
+ then show ?case using Typing.check_valI infer_v_g_weakening wf_weakening subtype_weakening by metis
+next
+ case (check_letI x \<Theta> \<Phi> \<B> \<Gamma> \<Delta> e \<tau> z s b c)
+ hence xf:"atom x \<sharp> \<Gamma>'" by metis
+ show ?case proof
+ show "atom x \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>', \<Delta>, e, \<tau>)" using check_letI using fresh_prod4 xf by metis
+ show "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>' ; \<Delta> \<turnstile> e \<Rightarrow> \<lbrace> z : b | c \<rbrace>" using infer_e_g_weakening check_letI by metis
+ show "atom z \<sharp> (x, \<Theta>, \<Phi>, \<B>, \<Gamma>', \<Delta>, e, \<tau>, s)"
+ by(unfold fresh_prodN,auto simp add: check_letI fresh_prodN)
+ have "toSet ((x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) \<subseteq> toSet ((x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>')" using check_letI toSet.simps
+ by (metis Un_commute Un_empty_right Un_insert_right insert_mono)
+ moreover hence "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f ((x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>')" using check_letI wfG_cons_weakening check_s_wf by metis
+ ultimately show "\<Theta> ; \<Phi> ; \<B> ; (x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>' ; \<Delta> \<turnstile> s \<Leftarrow> \<tau>" using check_letI by metis
+ qed
+next
+ case (check_let2I x \<Theta> \<Phi> \<B> G \<Delta> t s1 \<tau> s2)
+ show ?case proof
+ show "atom x \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>', \<Delta>, t, s1, \<tau>)" using check_let2I using fresh_prod4 by auto
+ show "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>' ; \<Delta> \<turnstile> s1 \<Leftarrow> t" using check_let2I by metis
+ have "toSet ((x, b_of t, c_of t x) #\<^sub>\<Gamma> G) \<subseteq> toSet ((x, b_of t, c_of t x ) #\<^sub>\<Gamma> \<Gamma>')" using check_let2I by auto
+ moreover hence "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f ((x, b_of t, c_of t x) #\<^sub>\<Gamma> \<Gamma>')" using check_let2I wfG_cons_weakening check_s_wf by metis
+ ultimately show "\<Theta> ; \<Phi> ; \<B> ; (x, b_of t, c_of t x) #\<^sub>\<Gamma> \<Gamma>' ; \<Delta> \<turnstile> s2 \<Leftarrow> \<tau>" using check_let2I by metis
+ qed
+next
+ case (check_branch_list_consI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist' v cs \<tau> css dclist)
+ thus ?case using Typing.check_branch_list_consI by metis
+next
+ case (check_branch_list_finalI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist' v cs \<tau> dclist)
+ thus ?case using Typing.check_branch_list_finalI by metis
+next
+ case (check_branch_s_branchI \<Theta> \<B> \<Gamma> \<Delta> \<tau> const x \<Phi> tid cons v s)
+ show ?case proof
+ show "\<Theta>; \<B>; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f \<Delta> " using wf_weakening2(6) check_branch_s_branchI by metis
+ show "\<turnstile>\<^sub>w\<^sub>f \<Theta>" using check_branch_s_branchI by auto
+ show "\<Theta>; \<B>; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f \<tau> " using check_branch_s_branchI wfT_weakening \<open>wfG \<Theta> \<B> \<Gamma>'\<close> by presburger
+
+ show "\<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f const " using check_branch_s_branchI by auto
+ show "atom x \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>', \<Delta>, tid, cons, const, v, \<tau>)" using check_branch_s_branchI by auto
+ have "toSet ((x, b_of const, CE_val v == CE_val(V_cons tid cons (V_var x)) AND c_of const x) #\<^sub>\<Gamma> \<Gamma>) \<subseteq> toSet ((x, b_of const, CE_val v == CE_val (V_cons tid cons (V_var x)) AND c_of const x) #\<^sub>\<Gamma> \<Gamma>')"
+ using check_branch_s_branchI by auto
+ moreover hence "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f ((x, b_of const, CE_val v == CE_val (V_cons tid cons (V_var x)) AND c_of const x ) #\<^sub>\<Gamma> \<Gamma>')"
+ using check_branch_s_branchI wfG_cons_weakening check_s_wf by metis
+ ultimately show "\<Theta> ; \<Phi> ; \<B> ; (x, b_of const, CE_val v == CE_val (V_cons tid cons (V_var x)) AND c_of const x ) #\<^sub>\<Gamma> \<Gamma>' ; \<Delta> \<turnstile> s \<Leftarrow> \<tau>"
+ using check_branch_s_branchI using fresh_dom_free by auto
+ qed
+next
+ case (check_ifI z \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v s1 s2 \<tau>)
+ show ?case proof
+ show \<open>atom z \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>', \<Delta>, v, s1, s2, \<tau>)\<close> using fresh_prodN check_ifI by auto
+ show \<open>\<Theta>; \<B>; \<Gamma>' \<turnstile> v \<Leftarrow> \<lbrace> z : B_bool | TRUE \<rbrace>\<close> using check_v_g_weakening check_ifI by auto
+ show \<open> \<Theta> ; \<Phi> ; \<B> ; \<Gamma>' ; \<Delta> \<turnstile> s1 \<Leftarrow> \<lbrace> z : b_of \<tau> | CE_val v == CE_val (V_lit L_true) IMP c_of \<tau> z \<rbrace>\<close> using check_ifI by auto
+ show \<open> \<Theta> ; \<Phi> ; \<B> ; \<Gamma>' ; \<Delta> \<turnstile> s2 \<Leftarrow> \<lbrace> z : b_of \<tau> | CE_val v == CE_val (V_lit L_false) IMP c_of \<tau> z \<rbrace>\<close> using check_ifI by auto
+ qed
+next
+ case (check_whileI \<Delta> G P s1 z s2 \<tau>')
+ then show ?case using check_s_check_branch_s_check_branch_list.intros check_v_g_weakening subtype_weakening wf_weakening
+ by (meson infer_v_g_weakening)
+next
+ case (check_seqI \<Delta> G P s1 z s2 \<tau>)
+ then show ?case using check_s_check_branch_s_check_branch_list.intros check_v_g_weakening subtype_weakening wf_weakening
+ by (meson infer_v_g_weakening)
+next
+ case (check_varI u \<Theta> \<Phi> \<B> \<Gamma> \<Delta> \<tau>' v \<tau> s)
+ thus ?case using check_v_g_weakening check_s_check_branch_s_check_branch_list.intros by auto
+next
+ case (check_assignI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> u \<tau> v z \<tau>')
+ show ?case proof
+ show \<open>\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using check_assignI by auto
+ show \<open>\<Theta>; \<B>; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f \<Delta>\<close> using check_assignI wf_weakening by auto
+ show \<open>(u, \<tau>) \<in> setD \<Delta>\<close> using check_assignI by auto
+ show \<open>\<Theta>; \<B>; \<Gamma>' \<turnstile> v \<Leftarrow> \<tau>\<close> using check_assignI check_v_g_weakening by auto
+ show \<open>\<Theta>; \<B>; \<Gamma>' \<turnstile> \<lbrace> z : B_unit | TRUE \<rbrace> \<lesssim> \<tau>'\<close> using subtype_weakening check_assignI by auto
+ qed
+next
+ case (check_caseI \<Delta> \<Gamma> \<Theta> dclist cs \<tau> tid v z)
+
+ then show ?case using check_s_check_branch_s_check_branch_list.intros check_v_g_weakening subtype_weakening wf_weakening
+ by (meson infer_v_g_weakening)
+next
+ case (check_assertI x \<Theta> \<Phi> \<B> \<Gamma> \<Delta> c \<tau> s)
+ show ?case proof
+ show \<open>atom x \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>', \<Delta>, c, \<tau>, s)\<close> using check_assertI by auto
+
+ have " \<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f (x, B_bool, c) #\<^sub>\<Gamma> \<Gamma>" using check_assertI check_s_wf by metis
+ hence *:" \<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f (x, B_bool, c) #\<^sub>\<Gamma> \<Gamma>'" using wfG_cons_weakening check_assertI by metis
+ moreover have "toSet ((x, B_bool, c) #\<^sub>\<Gamma> \<Gamma>) \<subseteq> toSet ((x, B_bool, c) #\<^sub>\<Gamma> \<Gamma>')" using check_assertI by auto
+ thus \<open> \<Theta> ; \<Phi> ; \<B> ; (x, B_bool, c) #\<^sub>\<Gamma> \<Gamma>' ; \<Delta> \<turnstile> s \<Leftarrow> \<tau>\<close> using check_assertI(11) [OF _ *] by auto
+
+ show \<open>\<Theta>; \<B>; \<Gamma>' \<Turnstile> c \<close> using check_assertI valid_weakening by metis
+ show \<open> \<Theta>; \<B>; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using check_assertI wf_weakening by metis
+ qed
+qed
+
+lemma wfG_xa_fresh_in_v:
+ fixes c::c and \<Gamma>::\<Gamma> and G::\<Gamma> and v::v and xa::x
+ assumes "\<Theta>; \<B>; \<Gamma> \<turnstile> v \<Rightarrow> \<tau>" and "G=( \<Gamma>'@ (x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)" and "atom xa \<sharp> G" and "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f G"
+ shows "atom xa \<sharp> v"
+proof -
+ have "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b_of \<tau>" using infer_v_wf assms by metis
+ hence "supp v \<subseteq> atom_dom \<Gamma> \<union> supp \<B>" using wfV_supp by simp
+ moreover have "atom xa \<notin> atom_dom G"
+ using assms wfG_atoms_supp_eq[OF assms(4)] fresh_def by metis
+ ultimately show ?thesis using fresh_def
+ using assms infer_v_wf wfG_atoms_supp_eq
+ fresh_GCons fresh_append_g subsetCE
+ by (metis wfG_x_fresh_in_v_simple)
+qed
+
+lemma fresh_z_subst_g:
+ fixes G::\<Gamma>
+ assumes "atom z' \<sharp> (x,v)" and \<open>atom z' \<sharp> G\<close>
+ shows "atom z' \<sharp> G[x::=v]\<^sub>\<Gamma>\<^sub>v"
+proof -
+ have "atom z' \<sharp> v" using assms fresh_prod2 by auto
+ thus ?thesis using fresh_subst_gv assms by metis
+qed
+
+lemma wfG_xa_fresh_in_subst_v:
+ fixes c::c and v::v and x::x and \<Gamma>::\<Gamma> and G::\<Gamma> and xa::x
+ assumes "\<Theta>; \<B>; \<Gamma> \<turnstile> v \<Rightarrow> \<tau>" and "G=( \<Gamma>'@ (x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)" and "atom xa \<sharp> G" and "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f G"
+ shows "atom xa \<sharp> (subst_gv G x v)"
+proof -
+ have "atom xa \<sharp> v" using wfG_xa_fresh_in_v assms by metis
+ thus ?thesis using fresh_subst_gv assms by metis
+qed
+
+subsection \<open>Weakening Immutable Variable Context\<close>
+
+declare check_s_check_branch_s_check_branch_list.intros[simp]
+declare check_s_check_branch_s_check_branch_list.intros[intro]
+
+lemma check_s_d_weakening:
+ fixes s::s and v::v and cs::branch_s and css::branch_list
+ shows " \<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile> s \<Leftarrow> \<tau> \<Longrightarrow> setD \<Delta> \<subseteq> setD \<Delta>' \<Longrightarrow> wfD \<Theta> \<B> \<Gamma> \<Delta>' \<Longrightarrow> \<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta>' \<turnstile> s \<Leftarrow> \<tau>" and
+ "check_branch_s \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid cons const v cs \<tau> \<Longrightarrow> setD \<Delta> \<subseteq> setD \<Delta>' \<Longrightarrow> wfD \<Theta> \<B> \<Gamma> \<Delta>' \<Longrightarrow> check_branch_s \<Theta> \<Phi> \<B> \<Gamma> \<Delta>' tid cons const v cs \<tau>" and
+ "check_branch_list \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist v css \<tau> \<Longrightarrow> setD \<Delta> \<subseteq> setD \<Delta>' \<Longrightarrow> wfD \<Theta> \<B> \<Gamma> \<Delta>' \<Longrightarrow> check_branch_list \<Theta> \<Phi> \<B> \<Gamma> \<Delta>' tid dclist v css \<tau>"
+proof(nominal_induct \<tau> and \<tau> and \<tau> avoiding: \<Delta>' arbitrary: v rule: check_s_check_branch_s_check_branch_list.strong_induct)
+ case (check_valI \<Theta> \<B> \<Gamma> \<Delta> \<Phi> v \<tau>' \<tau>)
+ then show ?case using check_s_check_branch_s_check_branch_list.intros by blast
+next
+ case (check_letI x \<Theta> \<Phi> \<B> \<Gamma> \<Delta> e \<tau> z s b c)
+ show ?case proof
+ show "atom x \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>, \<Delta>', e, \<tau>)" using check_letI by auto
+ show "atom z \<sharp> (x, \<Theta>, \<Phi>, \<B>, \<Gamma>, \<Delta>', e, \<tau>, s)" using check_letI by auto
+ show "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta>' \<turnstile> e \<Rightarrow> \<lbrace> z : b | c \<rbrace>" using check_letI infer_e_d_weakening by auto
+ have "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f (x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>" using check_letI check_s_wf by metis
+ moreover have "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>'" using check_letI check_s_wf by metis
+ ultimately have "\<Theta>; \<B>; (x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>'" using wf_weakening2(6) toSet.simps by fast
+ thus "\<Theta> ; \<Phi> ; \<B> ; (x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\<Gamma> \<Gamma> ; \<Delta>' \<turnstile> s \<Leftarrow> \<tau>" using check_letI by simp
+ qed
+next
+ case (check_branch_s_branchI \<Theta> \<B> \<Gamma> \<Delta> \<tau> const x \<Phi> tid cons v s)
+ moreover have "\<Theta> ;\<B> \<turnstile>\<^sub>w\<^sub>f (x, b_of const, CE_val v == CE_val (V_cons tid cons (V_var x)) AND c_of const x ) #\<^sub>\<Gamma> \<Gamma>"
+ using check_s_wf[OF check_branch_s_branchI(16) ] by metis
+ moreover hence " \<Theta> ;\<B> ; (x, b_of const, CE_val v == CE_val (V_cons tid cons (V_var x)) AND c_of const x ) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>'"
+ using wf_weakening2(6) check_branch_s_branchI by fastforce
+ ultimately show ?case
+ using check_s_check_branch_s_check_branch_list.intros by simp
+next
+ case (check_branch_list_consI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist v cs \<tau> css)
+ then show ?case using check_s_check_branch_s_check_branch_list.intros by meson
+next
+ case (check_branch_list_finalI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist v cs \<tau>)
+ then show ?case using check_s_check_branch_s_check_branch_list.intros by meson
+next
+ case (check_ifI z \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v s1 s2 \<tau>)
+ show ?case proof
+ show \<open>atom z \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>, \<Delta>', v, s1, s2, \<tau>)\<close> using fresh_prodN check_ifI by auto
+ show \<open>\<Theta>; \<B>; \<Gamma> \<turnstile> v \<Leftarrow> \<lbrace> z : B_bool | TRUE \<rbrace>\<close> using check_ifI by auto
+ show \<open> \<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta>' \<turnstile> s1 \<Leftarrow> \<lbrace> z : b_of \<tau> | CE_val v == CE_val (V_lit L_true) IMP c_of \<tau> z \<rbrace>\<close> using check_ifI by auto
+ show \<open> \<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta>' \<turnstile> s2 \<Leftarrow> \<lbrace> z : b_of \<tau> | CE_val v == CE_val (V_lit L_false) IMP c_of \<tau> z \<rbrace>\<close> using check_ifI by auto
+ qed
+next
+ case (check_assertI x \<Theta> \<Phi> \<B> \<Gamma> \<Delta> c \<tau> s)
+ show ?case proof
+ show "atom x \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>, \<Delta>', c, \<tau>,s)" using fresh_prodN check_assertI by auto
+ show *:" \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>' " using check_assertI by auto
+ hence "\<Theta>; \<B>; (x, B_bool, c) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>' " using wf_weakening2(6)[OF *, of " (x, B_bool, c) #\<^sub>\<Gamma> \<Gamma>"] check_assertI check_s_wf toSet.simps by auto
+ thus "\<Theta> ; \<Phi> ; \<B> ; (x, B_bool, c) #\<^sub>\<Gamma> \<Gamma> ; \<Delta>' \<turnstile> s \<Leftarrow> \<tau>"
+ using check_assertI(11)[OF \<open>setD \<Delta> \<subseteq> setD \<Delta>'\<close>] by simp
+ (* wf_weakening2(6) check_s_wf check_assertI *)
+ show "\<Theta>; \<B>; \<Gamma> \<Turnstile> c " using fresh_prodN check_assertI by auto
+
+ qed
+next
+ case (check_let2I x \<Theta> \<Phi> \<B> G \<Delta> t s1 \<tau> s2)
+ show ?case proof
+ show "atom x \<sharp> (\<Theta>, \<Phi>, \<B>, G, \<Delta>', t , s1, \<tau>)" using check_let2I by auto
+
+ show "\<Theta> ; \<Phi> ; \<B> ; G ; \<Delta>' \<turnstile> s1 \<Leftarrow> t" using check_let2I infer_e_d_weakening by auto
+ have "\<Theta>; \<B>; (x, b_of t, c_of t x ) #\<^sub>\<Gamma> G \<turnstile>\<^sub>w\<^sub>f \<Delta>'" using check_let2I wf_weakening2(6) check_s_wf by fastforce
+ thus "\<Theta> ; \<Phi> ; \<B> ; (x, b_of t, c_of t x) #\<^sub>\<Gamma> G ; \<Delta>' \<turnstile> s2 \<Leftarrow> \<tau>" using check_let2I by simp
+ qed
+next
+ case (check_varI u \<Theta> \<Phi> \<B> \<Gamma> \<Delta> \<tau>' v \<tau> s)
+ show ?case proof
+ show "atom u \<sharp> (\<Theta>, \<Phi>, \<B>, \<Gamma>, \<Delta>', \<tau>', v, \<tau>)" using check_varI by auto
+ show "\<Theta>; \<B>; \<Gamma> \<turnstile> v \<Leftarrow> \<tau>'" using check_varI by auto
+ have "setD ((u, \<tau>') #\<^sub>\<Delta>\<Delta>) \<subseteq> setD ((u, \<tau>') #\<^sub>\<Delta>\<Delta>')" using setD.simps check_varI by auto
+ moreover have "u \<notin> fst ` setD \<Delta>'" using check_varI(1) setD.simps fresh_DCons by (simp add: fresh_d_not_in)
+ moreover hence "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f (u, \<tau>') #\<^sub>\<Delta>\<Delta>' " using wfD_cons fresh_DCons setD.simps check_varI check_v_wf by metis
+ ultimately show "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; (u, \<tau>') #\<^sub>\<Delta>\<Delta>' \<turnstile> s \<Leftarrow> \<tau>" using check_varI by auto
+ qed
+next
+ case (check_assignI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> u \<tau> v z \<tau>')
+ moreover hence "(u, \<tau>) \<in> setD \<Delta>'" by auto
+ ultimately show ?case using Typing.check_assignI by simp
+next
+ case (check_whileI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> s1 z s2 \<tau>')
+ then show ?case using check_s_check_branch_s_check_branch_list.intros by meson
+next
+ case (check_seqI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> s1 z s2 \<tau>)
+ then show ?case using check_s_check_branch_s_check_branch_list.intros by meson
+next
+ case (check_caseI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist v cs \<tau> z)
+ then show ?case using check_s_check_branch_s_check_branch_list.intros by meson
+
+qed
+
+lemma valid_ce_eq:
+ fixes v::v and ce2::ce
+ assumes "ce1 = ce2[x::=v]\<^sub>c\<^sub>e\<^sub>v" and "wfV \<Theta> \<B> GNil v b" and "wfCE \<Theta> \<B> ((x, b, TRUE) #\<^sub>\<Gamma> GNil) ce2 b'" and "wfCE \<Theta> \<B> GNil ce1 b'"
+ shows \<open>\<Theta>; \<B>; (x, b, ([[x ]\<^sup>v]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e )) #\<^sub>\<Gamma> GNil \<Turnstile> ce1 == ce2 \<close>
+ unfolding valid.simps proof
+ have wfg: "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f (x, b, [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e ) #\<^sub>\<Gamma> GNil"
+ using wfG_cons1I wfG_nilI wfX_wfY assms wf_intros
+ by (meson fresh_GNil wfC_e_eq wfG_intros2)
+
+ show wf: \<open>\<Theta>; \<B>; (x, b, [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e ) #\<^sub>\<Gamma> GNil \<turnstile>\<^sub>w\<^sub>f ce1 == ce2 \<close>
+ apply(rule wfC_eqI[where b=b'])
+ using wfg toSet.simps assms wfCE_weakening apply simp
+
+ using wfg assms wf_replace_inside1(8) assms
+ using wfC_trueI wf_trans(8) by auto
+
+ show \<open>\<forall>i. ((\<Theta> ; (x, b, [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e ) #\<^sub>\<Gamma> GNil \<turnstile> i) \<and> (i \<Turnstile> (x, b, [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e ) #\<^sub>\<Gamma> GNil) \<longrightarrow>
+ (i \<Turnstile> (ce1 == ce2 )))\<close> proof(rule+,goal_cases)
+ fix i
+ assume as:"\<Theta> ; (x, b, [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e ) #\<^sub>\<Gamma> GNil \<turnstile> i \<and> i \<Turnstile> (x, b, [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e ) #\<^sub>\<Gamma> GNil"
+ have 1:"wfV \<Theta> \<B> ((x, b, [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e ) #\<^sub>\<Gamma> GNil) v b"
+ using wf_weakening assms append_g.simps toSet.simps wf wfX_wfY
+ by (metis empty_subsetI)
+ hence "\<exists>s. i \<lbrakk> v \<rbrakk> ~ s" using eval_v_exist[OF _ 1] as by auto
+ then obtain s where iv:"i\<lbrakk> v \<rbrakk> ~ s" ..
+
+ hence ix:"i x = Some s" proof -
+ have "i \<Turnstile> [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e" using is_satis_g.simps as by auto
+ hence "i \<lbrakk> [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e \<rbrakk> ~ True" using is_satis.simps by auto
+ hence "i \<lbrakk> [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e \<rbrakk> ~ s" using
+ iv eval_e_elims
+ by (metis eval_c_elims(7) eval_e_uniqueness eval_e_valI)
+ thus ?thesis using eval_v_elims(2) eval_e_elims(1) by metis
+ qed
+
+ have 1:"wfCE \<Theta> \<B> ((x, b, [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e ) #\<^sub>\<Gamma> GNil) ce1 b'"
+ using wfCE_weakening assms append_g.simps toSet.simps wf wfX_wfY
+ by (metis empty_subsetI)
+ hence "\<exists>s1. i \<lbrakk> ce1 \<rbrakk> ~ s1" using eval_e_exist assms as by auto
+ then obtain s1 where s1: "i\<lbrakk>ce1\<rbrakk> ~ s1" ..
+
+ moreover have "i\<lbrakk> ce2 \<rbrakk> ~ s1" proof -
+ have "i\<lbrakk> ce2[x::=v]\<^sub>c\<^sub>e\<^sub>v \<rbrakk> ~ s1" using assms s1 by auto
+ moreover have "ce1 = ce2[x::=v]\<^sub>c\<^sub>e\<^sub>v" using subst_v_ce_def assms subst_v_simple_commute by auto
+ ultimately have "i(x \<mapsto> s) \<lbrakk> ce2 \<rbrakk> ~ s1"
+ using ix subst_e_eval_v[of i ce1 s1 "ce2[z::=[ x ]\<^sup>v]\<^sub>v" x v s] iv s1 by auto
+ moreover have "i(x \<mapsto> s) = i" using ix by auto
+ ultimately show ?thesis by auto
+ qed
+ ultimately show "i \<lbrakk> ce1 == ce2 \<rbrakk> ~ True " using eval_c_eqI by metis
+ qed
+qed
+
+
+lemma check_v_top:
+ fixes v::v
+ assumes "\<Theta>; \<B>; GNil \<turnstile> v \<Leftarrow> \<tau>" and "ce1 = ce2[z::=v]\<^sub>c\<^sub>e\<^sub>v" and "\<Theta>; \<B>; GNil \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b_of \<tau> | ce1 == ce2 \<rbrace>"
+ and "supp ce1 \<subseteq> supp \<B>"
+ shows "\<Theta>; \<B>; GNil \<turnstile> v \<Leftarrow> \<lbrace> z : b_of \<tau> | ce1 == ce2 \<rbrace>"
+proof -
+ obtain t where t: "\<Theta>; \<B>; GNil \<turnstile> v \<Rightarrow> t \<and> \<Theta>; \<B>; GNil \<turnstile> t \<lesssim> \<tau>"
+ using assms check_v_elims by metis
+
+ then obtain z' and b' where *:"t = \<lbrace> z' : b' | [ [ z' ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e \<rbrace> \<and> atom z' \<sharp> v \<and> atom z' \<sharp> (\<Theta>, \<B>,GNil)"
+ using assms infer_v_form by metis
+ have beq: "b_of t = b_of \<tau>" using subtype_eq_base2 b_of.simps t by auto
+ obtain x::x where xf: \<open>atom x \<sharp> (\<Theta>, \<B>, GNil, z', [ [ z' ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e , z, ce1 == ce2 )\<close>
+ using obtain_fresh by metis
+
+ have "\<Theta>; \<B>; (x, b_of \<tau>, TRUE) #\<^sub>\<Gamma> GNil \<turnstile>\<^sub>w\<^sub>f (ce1[z::=[ x ]\<^sup>v]\<^sub>v == ce2[z::=[ x ]\<^sup>v]\<^sub>v )"
+ using wfT_wfC2[OF assms(3), of x] subst_cv.simps(6) subst_v_c_def subst_v_ce_def fresh_GNil by simp
+
+ then obtain b2 where b2: "\<Theta>; \<B>; (x, b_of t, TRUE) #\<^sub>\<Gamma> GNil \<turnstile>\<^sub>w\<^sub>f ce1[z::=[ x ]\<^sup>v]\<^sub>v : b2 \<and>
+ \<Theta>; \<B>; (x, b_of t, TRUE) #\<^sub>\<Gamma> GNil \<turnstile>\<^sub>w\<^sub>f ce2[z::=[ x ]\<^sup>v]\<^sub>v : b2" using wfC_elims(3)
+ beq by metis
+
+ from xf have "\<Theta>; \<B>; GNil \<turnstile> \<lbrace> z' : b_of t | [ [ z' ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e \<rbrace> \<lesssim> \<lbrace> z : b_of t | ce1 == ce2 \<rbrace>"
+ proof
+ show \<open> \<Theta>; \<B>; GNil \<turnstile>\<^sub>w\<^sub>f \<lbrace> z' : b_of t | [ [ z' ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e \<rbrace> \<close> using b_of.simps assms infer_v_wf t * by auto
+ show \<open> \<Theta>; \<B>; GNil \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b_of t | ce1 == ce2 \<rbrace> \<close> using beq assms by auto
+ have \<open>\<Theta>; \<B>; (x, b_of t, ([ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e )) #\<^sub>\<Gamma> GNil \<Turnstile> (ce1[z::=[ x ]\<^sup>v]\<^sub>v == ce2[z::=[ x ]\<^sup>v]\<^sub>v ) \<close>
+ proof(rule valid_ce_eq)
+ show \<open>ce1[z::=[ x ]\<^sup>v]\<^sub>v = ce2[z::=[ x ]\<^sup>v]\<^sub>v[x::=v]\<^sub>c\<^sub>e\<^sub>v\<close> proof -
+ have "atom z \<sharp> ce1" using assms fresh_def x_not_in_b_set by fast
+ hence "ce1[z::=[ x ]\<^sup>v]\<^sub>v = ce1"
+ using forget_subst_v by auto
+ also have "... = ce2[z::=v]\<^sub>c\<^sub>e\<^sub>v" using assms by auto
+ also have "... = ce2[z::=[ x ]\<^sup>v]\<^sub>v[x::=v]\<^sub>c\<^sub>e\<^sub>v" proof -
+ have "atom x \<sharp> ce2" using xf fresh_prodN c.fresh by metis
+ thus ?thesis using subst_v_simple_commute subst_v_ce_def by simp
+ qed
+ finally show ?thesis by auto
+ qed
+ show \<open> \<Theta>; \<B>; GNil \<turnstile>\<^sub>w\<^sub>f v : b_of t \<close> using infer_v_wf t by simp
+ show \<open> \<Theta>; \<B>; (x, b_of t, TRUE) #\<^sub>\<Gamma> GNil \<turnstile>\<^sub>w\<^sub>f ce2[z::=[ x ]\<^sup>v]\<^sub>v : b2 \<close> using b2 by auto
+
+ have " \<Theta>; \<B>; (x, b_of t, TRUE) #\<^sub>\<Gamma> GNil \<turnstile>\<^sub>w\<^sub>f ce1[z::=[ x ]\<^sup>v]\<^sub>v : b2" using b2 by auto
+ moreover have "atom x \<sharp> ce1[z::=[ x ]\<^sup>v]\<^sub>v"
+ using fresh_subst_v_if assms fresh_def
+ using \<open>\<Theta>; \<B>; GNil \<turnstile>\<^sub>w\<^sub>f v : b_of t\<close> \<open>ce1[z::=[ x ]\<^sup>v]\<^sub>v = ce2[z::=[ x ]\<^sup>v]\<^sub>v[x::=v]\<^sub>c\<^sub>e\<^sub>v\<close>
+ fresh_GNil subst_v_ce_def wfV_x_fresh by auto
+ ultimately show \<open> \<Theta>; \<B>; GNil \<turnstile>\<^sub>w\<^sub>f ce1[z::=[ x ]\<^sup>v]\<^sub>v : b2 \<close> using
+ wf_restrict(8) by force
+ qed
+ moreover have v: "v[z'::=[ x ]\<^sup>v]\<^sub>v\<^sub>v = v"
+ using forget_subst assms infer_v_wf wfV_supp x_not_in_b_set
+ by (simp add: "local.*")
+ ultimately show "\<Theta>; \<B>; (x, b_of t, ([ [ z' ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e )[z'::=[ x ]\<^sup>v]\<^sub>v) #\<^sub>\<Gamma> GNil \<Turnstile> (ce1 == ce2 )[z::=[ x ]\<^sup>v]\<^sub>v"
+ unfolding subst_cv.simps subst_v_c_def subst_cev.simps subst_vv.simps
+ using subst_v_ce_def by simp
+ qed
+ thus ?thesis using b_of.simps assms * check_v_subtypeI t b_of.simps subtype_eq_base2 by metis
+qed
+
+end
+
diff --git a/thys/MiniSail/Wellformed.thy b/thys/MiniSail/Wellformed.thy
--- a/thys/MiniSail/Wellformed.thy
+++ b/thys/MiniSail/Wellformed.thy
@@ -1,581 +1,581 @@
-(*<*)
-theory Wellformed
- imports IVSubst BTVSubst
-begin
- (*>*)
-
-chapter \<open>Wellformed Terms\<close>
-
-text \<open>We require that expressions and values are well-formed. This includes a notion
-of well-sortedness. We identify a sort with a basic type and
-define the judgement as two clusters of mutually recursive
-inductive predicates. Some of the proofs are across all of the predicates and
-although they seemed at first to be daunting, they have all
-worked out well.\<close>
-
-named_theorems ms_wb "Facts for helping with well-sortedness"
-
-section \<open>Definitions\<close>
-
-inductive wfV :: "\<Theta> \<Rightarrow> \<B> \<Rightarrow> \<Gamma> \<Rightarrow> v \<Rightarrow> b \<Rightarrow> bool" (" _ ; _ ; _ \<turnstile>\<^sub>w\<^sub>f _ : _ " [50,50,50] 50) and
- wfC :: "\<Theta> \<Rightarrow> \<B> \<Rightarrow> \<Gamma> \<Rightarrow> c \<Rightarrow> bool" (" _ ; _ ; _ \<turnstile>\<^sub>w\<^sub>f _ " [50,50] 50) and
- wfG :: "\<Theta> \<Rightarrow> \<B> \<Rightarrow> \<Gamma> \<Rightarrow> bool" (" _ ; _ \<turnstile>\<^sub>w\<^sub>f _ " [50,50] 50) and
- wfT :: "\<Theta> \<Rightarrow> \<B> \<Rightarrow> \<Gamma> \<Rightarrow> \<tau> \<Rightarrow> bool" (" _ ; _ ; _ \<turnstile>\<^sub>w\<^sub>f _ " [50,50] 50) and
- wfTs :: "\<Theta> \<Rightarrow> \<B> \<Rightarrow> \<Gamma> \<Rightarrow> (string*\<tau>) list \<Rightarrow> bool" (" _ ; _ ; _ \<turnstile>\<^sub>w\<^sub>f _ " [50,50] 50) and
- wfTh :: "\<Theta> \<Rightarrow> bool" (" \<turnstile>\<^sub>w\<^sub>f _ " [50] 50) and
- wfB :: "\<Theta> \<Rightarrow> \<B> \<Rightarrow> b \<Rightarrow> bool" (" _ ; _ \<turnstile>\<^sub>w\<^sub>f _ " [50,50] 50) and
- wfCE :: "\<Theta> \<Rightarrow> \<B> \<Rightarrow> \<Gamma> \<Rightarrow> ce \<Rightarrow> b \<Rightarrow> bool" (" _ ; _ ; _ \<turnstile>\<^sub>w\<^sub>f _ : _ " [50,50,50] 50) and
- wfTD :: "\<Theta> \<Rightarrow> type_def \<Rightarrow> bool" (" _ \<turnstile>\<^sub>w\<^sub>f _ " [50,50] 50)
- where
-
-wfB_intI: "\<turnstile>\<^sub>w\<^sub>f \<Theta> \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f B_int"
-| wfB_boolI: "\<turnstile>\<^sub>w\<^sub>f \<Theta> \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f B_bool"
-| wfB_unitI: "\<turnstile>\<^sub>w\<^sub>f \<Theta> \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f B_unit"
-| wfB_bitvecI: "\<turnstile>\<^sub>w\<^sub>f \<Theta> \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f B_bitvec"
-| wfB_pairI: "\<lbrakk> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b1 ; \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b2 \<rbrakk> \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f B_pair b1 b2"
-
-| wfB_consI: "\<lbrakk>
- \<turnstile>\<^sub>w\<^sub>f \<Theta>;
- (AF_typedef s dclist) \<in> set \<Theta>
-\<rbrakk> \<Longrightarrow>
- \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f B_id s"
-
-| wfB_appI: "\<lbrakk>
- \<turnstile>\<^sub>w\<^sub>f \<Theta>;
- \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b;
- (AF_typedef_poly s bv dclist) \<in> set \<Theta>
-\<rbrakk> \<Longrightarrow>
- \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f B_app s b"
-
-| wfV_varI: "\<lbrakk> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> ; Some (b,c) = lookup \<Gamma> x \<rbrakk> \<Longrightarrow> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f V_var x : b"
-| wfV_litI: "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<Longrightarrow> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f V_lit l : base_for_lit l"
-
-| wfV_pairI: "\<lbrakk>
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v1 : b1 ;
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v2 : b2
-\<rbrakk> \<Longrightarrow>
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f (V_pair v1 v2) : B_pair b1 b2"
-
-| wfV_consI: "\<lbrakk>
- AF_typedef s dclist \<in> set \<Theta>;
- (dc, \<lbrace> x : b' | c \<rbrace>) \<in> set dclist ;
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b'
-\<rbrakk> \<Longrightarrow>
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f V_cons s dc v : B_id s"
-
-| wfV_conspI: "\<lbrakk>
- AF_typedef_poly s bv dclist \<in> set \<Theta>;
- (dc, \<lbrace> x : b' | c \<rbrace>) \<in> set dclist ;
- \<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f b;
- atom bv \<sharp> (\<Theta>, \<B>, \<Gamma>, b , v);
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b'[bv::=b]\<^sub>b\<^sub>b
-\<rbrakk> \<Longrightarrow>
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f V_consp s dc b v : B_app s b"
-
-| wfCE_valI : "\<lbrakk>
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b
-\<rbrakk> \<Longrightarrow>
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_val v : b"
-
-| wfCE_plusI: "\<lbrakk>
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v1 : B_int;
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v2 : B_int
-\<rbrakk> \<Longrightarrow>
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_op Plus v1 v2 : B_int"
-
-| wfCE_leqI:"\<lbrakk>
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v1 : B_int;
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v2 : B_int
-\<rbrakk> \<Longrightarrow>
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_op LEq v1 v2 : B_bool"
-
-| wfCE_eqI:"\<lbrakk>
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v1 : b;
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v2 : b
-\<rbrakk> \<Longrightarrow>
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_op Eq v1 v2 : B_bool"
-
-| wfCE_fstI: "\<lbrakk>
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v1 : B_pair b1 b2
-\<rbrakk> \<Longrightarrow>
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_fst v1 : b1"
-
-| wfCE_sndI: "\<lbrakk>
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v1 : B_pair b1 b2
-\<rbrakk> \<Longrightarrow>
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_snd v1 : b2"
-
-| wfCE_concatI: "\<lbrakk>
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v1 : B_bitvec ;
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v2 : B_bitvec
-\<rbrakk> \<Longrightarrow>
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_concat v1 v2 : B_bitvec"
-
-| wfCE_lenI: "\<lbrakk>
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v1 : B_bitvec
-\<rbrakk> \<Longrightarrow>
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_len v1 : B_int"
-
-| wfTI : "\<lbrakk>
- atom z \<sharp> (\<Theta>, \<B>, \<Gamma>) ;
- \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b;
- \<Theta>; \<B> ; (z,b,C_true) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f c
-\<rbrakk> \<Longrightarrow>
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b | c \<rbrace>"
-
-| wfC_eqI: "\<lbrakk>
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f e1 : b ;
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f e2 : b \<rbrakk> \<Longrightarrow>
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f C_eq e1 e2"
-| wfC_trueI: " \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<Longrightarrow> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f C_true "
-| wfC_falseI: " \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<Longrightarrow> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f C_false "
-
-| wfC_conjI: "\<lbrakk> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f c1 ; \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f c2 \<rbrakk> \<Longrightarrow> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f C_conj c1 c2"
-| wfC_disjI: "\<lbrakk> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f c1 ; \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f c2 \<rbrakk> \<Longrightarrow> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f C_disj c1 c2"
-| wfC_notI: "\<lbrakk> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f c1 \<rbrakk> \<Longrightarrow> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f C_not c1"
-| wfC_impI: "\<lbrakk> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f c1 ;
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f c2 \<rbrakk> \<Longrightarrow> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f C_imp c1 c2"
-
-| wfG_nilI: " \<turnstile>\<^sub>w\<^sub>f \<Theta> \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f GNil"
-| wfG_cons1I: "\<lbrakk> c \<notin> { TRUE, FALSE } ;
- \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> ;
- atom x \<sharp> \<Gamma> ;
- \<Theta> ; \<B> ; (x,b,C_true)#\<^sub>\<Gamma>\<Gamma> \<turnstile>\<^sub>w\<^sub>f c ; wfB \<Theta> \<B> b
- \<rbrakk> \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f ((x,b,c)#\<^sub>\<Gamma>\<Gamma>)"
-| wfG_cons2I: "\<lbrakk> c \<in> { TRUE, FALSE } ;
- \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> ;
- atom x \<sharp> \<Gamma> ;
- wfB \<Theta> \<B> b
- \<rbrakk> \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f ((x,b,c)#\<^sub>\<Gamma>\<Gamma>)"
-
-| wfTh_emptyI: " \<turnstile>\<^sub>w\<^sub>f []"
-
-| wfTh_consI: "\<lbrakk>
- (name_of_type tdef) \<notin> name_of_type ` set \<Theta> ;
- \<turnstile>\<^sub>w\<^sub>f \<Theta> ;
- \<Theta> \<turnstile>\<^sub>w\<^sub>f tdef \<rbrakk> \<Longrightarrow> \<turnstile>\<^sub>w\<^sub>f tdef#\<Theta>"
-
-| wfTD_simpleI: "\<lbrakk>
- \<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f lst
- \<rbrakk> \<Longrightarrow>
- \<Theta> \<turnstile>\<^sub>w\<^sub>f (AF_typedef s lst )"
-
-| wfTD_poly: "\<lbrakk>
- \<Theta> ; {|bv|} ; GNil \<turnstile>\<^sub>w\<^sub>f lst
- \<rbrakk> \<Longrightarrow>
- \<Theta> \<turnstile>\<^sub>w\<^sub>f (AF_typedef_poly s bv lst)"
-
-| wfTs_nil: "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<Longrightarrow> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f []::(string*\<tau>) list"
-| wfTs_cons: "\<lbrakk> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau> ;
- dc \<notin> fst ` set ts;
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ts::(string*\<tau>) list \<rbrakk> \<Longrightarrow> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ((dc,\<tau>)#ts)"
-
-inductive_cases wfC_elims:
- "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f C_true"
- "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f C_false"
- "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f C_eq e1 e2"
- "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f C_conj c1 c2"
- "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f C_disj c1 c2"
- "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f C_not c1"
- "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f C_imp c1 c2"
-
-inductive_cases wfV_elims:
- "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f V_var x : b"
- "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f V_lit l : b"
- "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f V_pair v1 v2 : b"
- "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f V_cons tyid dc v : b"
- "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f V_consp tyid dc b v : b'"
-
-inductive_cases wfCE_elims:
- " \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_val v : b"
- " \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_op Plus v1 v2 : b"
- " \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_op LEq v1 v2 : b"
- " \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_fst v1 : b"
- " \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_snd v1 : b"
- " \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_concat v1 v2 : b"
- " \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_len v1 : b"
- " \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_op opp v1 v2 : b"
- " \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_op Eq v1 v2 : b"
-
-inductive_cases wfT_elims:
- "\<Theta>; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau>::\<tau>"
- "\<Theta>; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b | c \<rbrace>"
-
-inductive_cases wfG_elims:
- "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f GNil"
- "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (x,b,c)#\<^sub>\<Gamma>\<Gamma>"
- "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (x,b,TRUE)#\<^sub>\<Gamma>\<Gamma>"
- "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (x,b,FALSE)#\<^sub>\<Gamma>\<Gamma>"
-
-inductive_cases wfTh_elims:
- " \<turnstile>\<^sub>w\<^sub>f []"
- " \<turnstile>\<^sub>w\<^sub>f td#\<Pi>"
-
-inductive_cases wfTD_elims:
- "\<Theta> \<turnstile>\<^sub>w\<^sub>f (AF_typedef s lst )"
- "\<Theta> \<turnstile>\<^sub>w\<^sub>f (AF_typedef_poly s bv lst )"
-
-inductive_cases wfTs_elims:
- "\<Theta>; \<B> ; GNil \<turnstile>\<^sub>w\<^sub>f ([]::((string*\<tau>) list))"
- "\<Theta>; \<B> ; GNil \<turnstile>\<^sub>w\<^sub>f ((t#ts)::((string*\<tau>) list))"
-
-inductive_cases wfB_elims:
- " \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f B_pair b1 b2"
- " \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f B_id s"
- " \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f B_app s b"
-
-equivariance wfV
-
-text \<open>This setup of 'avoiding' is not complete and for some of lemmas, such as weakening,
-do it the hard way\<close>
-nominal_inductive wfV
- avoids wfV_conspI: bv | wfTI: z
-proof(goal_cases)
- case (1 s bv dclist \<Theta> dc x b' c \<B> b \<Gamma> v)
-
- moreover hence "atom bv \<sharp> V_consp s dc b v" using v.fresh fresh_prodN pure_fresh by metis
- moreover have "atom bv \<sharp> B_app s b" using b.fresh fresh_prodN pure_fresh 1 by metis
- ultimately show ?case using b.fresh v.fresh pure_fresh fresh_star_def fresh_prodN by fastforce
-next
- case (2 s bv dclist \<Theta> dc x b' c \<B> b \<Gamma> v)
- then show ?case by auto
-next
- case (3 z \<Gamma> \<Theta> \<B> b c)
- then show ?case using \<tau>.fresh fresh_star_def fresh_prodN by fastforce
-next
- case (4 z \<Gamma> \<Theta> \<B> b c)
- then show ?case by auto
-qed
-
-inductive
- wfE :: "\<Theta> \<Rightarrow> \<Phi> \<Rightarrow> \<B> \<Rightarrow> \<Gamma> \<Rightarrow> \<Delta> \<Rightarrow> e \<Rightarrow> b \<Rightarrow> bool" (" _ ; _ ; _ ; _ ; _ \<turnstile>\<^sub>w\<^sub>f _ : _ " [50,50,50] 50) and
- wfS :: "\<Theta> \<Rightarrow> \<Phi> \<Rightarrow> \<B> \<Rightarrow> \<Gamma> \<Rightarrow> \<Delta> \<Rightarrow> s \<Rightarrow> b \<Rightarrow> bool" (" _ ; _ ; _ ; _ ; _ \<turnstile>\<^sub>w\<^sub>f _ : _ " [50,50,50] 50) and
- wfCS :: "\<Theta> \<Rightarrow> \<Phi> \<Rightarrow> \<B> \<Rightarrow> \<Gamma> \<Rightarrow> \<Delta> \<Rightarrow> tyid \<Rightarrow> string \<Rightarrow> \<tau> \<Rightarrow> branch_s \<Rightarrow> b \<Rightarrow> bool" (" _ ; _ ; _ ; _ ; _ ; _ ; _ ; _ \<turnstile>\<^sub>w\<^sub>f _ : _ " [50,50,50,50,50] 50) and
- wfCSS :: "\<Theta> \<Rightarrow> \<Phi> \<Rightarrow> \<B> \<Rightarrow> \<Gamma> \<Rightarrow> \<Delta> \<Rightarrow> tyid \<Rightarrow> (string * \<tau>) list \<Rightarrow> branch_list \<Rightarrow> b \<Rightarrow> bool" (" _ ; _ ; _ ; _ ; _ ; _ ; _ \<turnstile>\<^sub>w\<^sub>f _ : _ " [50,50,50,50,50] 50) and
- wfPhi :: "\<Theta> \<Rightarrow> \<Phi> \<Rightarrow> bool" (" _ \<turnstile>\<^sub>w\<^sub>f _ " [50,50] 50) and
- wfD :: "\<Theta> \<Rightarrow> \<B> \<Rightarrow> \<Gamma> \<Rightarrow> \<Delta> \<Rightarrow> bool" (" _ ; _ ; _ \<turnstile>\<^sub>w\<^sub>f _ " [50,50] 50) and
- wfFTQ :: "\<Theta> \<Rightarrow> \<Phi> \<Rightarrow> fun_typ_q \<Rightarrow> bool" (" _ ; _ \<turnstile>\<^sub>w\<^sub>f _ " [50] 50) and
- wfFT :: "\<Theta> \<Rightarrow> \<Phi> \<Rightarrow> \<B> \<Rightarrow> fun_typ \<Rightarrow> bool" (" _ ; _ ; _ \<turnstile>\<^sub>w\<^sub>f _ " [50] 50) where
-
-wfE_valI : "\<lbrakk>
- \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>;
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>;
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b
-\<rbrakk> \<Longrightarrow>
- \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f AE_val v : b"
-
-| wfE_plusI: "\<lbrakk>
- \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>;
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>;
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v1 : B_int;
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v2 : B_int
-\<rbrakk> \<Longrightarrow>
- \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f AE_op Plus v1 v2 : B_int"
-
-| wfE_leqI:"\<lbrakk>
- \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> ;
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>;
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v1 : B_int;
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v2 : B_int
-\<rbrakk> \<Longrightarrow>
- \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f AE_op LEq v1 v2 : B_bool"
-
-| wfE_eqI:"\<lbrakk>
- \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> ;
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>;
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v1 : b;
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v2 : b;
- b \<in> { B_bool, B_int, B_unit }
-\<rbrakk> \<Longrightarrow>
- \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f AE_op Eq v1 v2 : B_bool"
-
-| wfE_fstI: "\<lbrakk>
- \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>;
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>;
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v1 : B_pair b1 b2
- \<rbrakk> \<Longrightarrow>
- \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f AE_fst v1 : b1"
-
-| wfE_sndI: "\<lbrakk>
- \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> ;
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>;
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v1 : B_pair b1 b2
-\<rbrakk> \<Longrightarrow>
- \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f AE_snd v1 : b2"
-
-| wfE_concatI: "\<lbrakk>
- \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> ;
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>;
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v1 : B_bitvec;
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v2 : B_bitvec
-\<rbrakk> \<Longrightarrow>
- \<Theta> ; \<Phi> ; \<B> ; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f AE_concat v1 v2 : B_bitvec"
-
-| wfE_splitI: "\<lbrakk>
- \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> ;
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>;
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v1 : B_bitvec;
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v2 : B_int
-\<rbrakk> \<Longrightarrow>
- \<Theta> ; \<Phi> ; \<B> ; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f AE_split v1 v2 : B_pair B_bitvec B_bitvec"
-
-| wfE_lenI: "\<lbrakk>
- \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> ;
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>;
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v1 : B_bitvec
-\<rbrakk> \<Longrightarrow>
- \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f AE_len v1 : B_int"
-
-| wfE_appI: "\<lbrakk>
- \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> ;
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>;
- Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \<tau> s))) = lookup_fun \<Phi> f ;
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b
-\<rbrakk> \<Longrightarrow>
- \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f AE_app f v : b_of \<tau>"
-
-| wfE_appPI: "\<lbrakk>
- \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> ;
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>;
- \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b';
- atom bv \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>, \<Delta>, b', v, (b_of \<tau>)[bv::=b']\<^sub>b);
- Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \<tau> s))) = lookup_fun \<Phi> f;
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : (b[bv::=b']\<^sub>b)
-\<rbrakk> \<Longrightarrow>
- \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f (AE_appP f b' v) : ((b_of \<tau>)[bv::=b']\<^sub>b)"
-
-| wfE_mvarI: "\<lbrakk>
- \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> ;
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>;
- (u,\<tau>) \<in> setD \<Delta>
-\<rbrakk> \<Longrightarrow>
- \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f AE_mvar u : b_of \<tau>"
-
-| wfS_valI: "\<lbrakk>
- \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> ;
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b ;
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>
-\<rbrakk> \<Longrightarrow>
- \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f (AS_val v) : b"
-
-| wfS_letI: "\<lbrakk>
- wfE \<Theta> \<Phi> \<B> \<Gamma> \<Delta> e b' ;
- \<Theta> ; \<Phi> ; \<B> ; (x,b',C_true) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b;
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> ;
- atom x \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>, \<Delta>, e, b)
-\<rbrakk> \<Longrightarrow>
- \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f LET x = e IN s : b"
-
-| wfS_assertI: "\<lbrakk>
- \<Theta> ; \<Phi> ; \<B> ; (x,B_bool,c) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b;
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f c ;
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> ;
- atom x \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>, \<Delta>, c, b, s)
-\<rbrakk> \<Longrightarrow>
- \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f ASSERT c IN s : b"
-
-| wfS_let2I: "\<lbrakk>
- \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f s1 : b_of \<tau> ;
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau>;
- \<Theta> ; \<Phi> ; \<B> ; (x,b_of \<tau>,C_true) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s2 : b ;
- atom x \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>, \<Delta>, s1, b,\<tau>)
-\<rbrakk> \<Longrightarrow>
- \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f LET x : \<tau> = s1 IN s2 : b"
-
-| wfS_ifI: "\<lbrakk>
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : B_bool;
- \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f s1 : b ;
- \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f s2 : b
-\<rbrakk> \<Longrightarrow>
- \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f IF v THEN s1 ELSE s2 : b"
-
-| wfS_varI : "\<lbrakk>
- wfT \<Theta> \<B> \<Gamma> \<tau> ;
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b_of \<tau>;
- atom u \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>, \<Delta>, \<tau>, v, b);
- \<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; (u,\<tau>)#\<^sub>\<Delta>\<Delta> \<turnstile>\<^sub>w\<^sub>f s : b
-\<rbrakk> \<Longrightarrow>
- \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f VAR u : \<tau> = v IN s : b "
-
-| wfS_assignI: "\<lbrakk>
- (u,\<tau>) \<in> setD \<Delta> ;
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> ;
- \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> ;
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b_of \<tau>
-\<rbrakk> \<Longrightarrow>
- \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f u ::= v : B_unit"
-
-| wfS_whileI: "\<lbrakk>
- \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f s1 : B_bool ;
- \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f s2 : b
-\<rbrakk> \<Longrightarrow>
- \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f WHILE s1 DO { s2 } : b"
-
-| wfS_seqI: "\<lbrakk>
- \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f s1 : B_unit ;
- \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f s2 : b
-\<rbrakk> \<Longrightarrow>
- \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f s1 ;; s2 : b"
-
-| wfS_matchI: "\<lbrakk>
- wfV \<Theta> \<B> \<Gamma> v (B_id tid) ;
- (AF_typedef tid dclist ) \<in> set \<Theta>;
- wfD \<Theta> \<B> \<Gamma> \<Delta> ;
- \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> ;
- \<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> ; tid ; dclist \<turnstile>\<^sub>w\<^sub>f cs : b
-\<rbrakk> \<Longrightarrow>
- \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f AS_match v cs : b "
-
-| wfS_branchI: "\<lbrakk>
- \<Theta> ; \<Phi> ; \<B> ; (x,b_of \<tau>,C_true) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b ;
- atom x \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>, \<Delta>, \<Gamma>,\<tau>);
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>
-\<rbrakk> \<Longrightarrow>
- \<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> ; tid ; dc ; \<tau> \<turnstile>\<^sub>w\<^sub>f dc x \<Rightarrow> s : b"
-
-| wfS_finalI: "\<lbrakk>
- \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> ; tid ; dc ; t \<turnstile>\<^sub>w\<^sub>f cs : b
- \<rbrakk> \<Longrightarrow>
- \<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> ; tid ; [(dc,t)] \<turnstile>\<^sub>w\<^sub>f AS_final cs : b "
-
-| wfS_cons: "\<lbrakk>
- \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> ; tid ; dc ; t \<turnstile>\<^sub>w\<^sub>f cs : b;
- \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> ; tid ; dclist \<turnstile>\<^sub>w\<^sub>f css : b
- \<rbrakk> \<Longrightarrow>
- \<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> ; tid ; (dc,t)#dclist \<turnstile>\<^sub>w\<^sub>f AS_cons cs css : b "
-
-| wfD_emptyI: "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<Longrightarrow> \<Theta> ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f []\<^sub>\<Delta>"
-
-| wfD_cons: "\<lbrakk>
- \<Theta> ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>::\<Delta> ;
- \<Theta> ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau>;
- u \<notin> fst ` setD \<Delta>
-\<rbrakk> \<Longrightarrow>
- \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ((u,\<tau>) #\<^sub>\<Delta> \<Delta>)"
-
-| wfPhi_emptyI: " \<turnstile>\<^sub>w\<^sub>f \<Theta> \<Longrightarrow> \<Theta> \<turnstile>\<^sub>w\<^sub>f []"
-
-| wfPhi_consI: "\<lbrakk>
- f \<notin> name_of_fun ` set \<Phi>;
- \<Theta> ; \<Phi> \<turnstile>\<^sub>w\<^sub>f ft;
- \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>
-\<rbrakk> \<Longrightarrow>
- \<Theta> \<turnstile>\<^sub>w\<^sub>f ((AF_fundef f ft)#\<Phi>)"
-
-| wfFTNone: " \<Theta> ; \<Phi> ; {||} \<turnstile>\<^sub>w\<^sub>f ft \<Longrightarrow> \<Theta> ; \<Phi> \<turnstile>\<^sub>w\<^sub>f AF_fun_typ_none ft"
-| wfFTSome: " \<Theta> ; \<Phi> ; {| bv |} \<turnstile>\<^sub>w\<^sub>f ft \<Longrightarrow> \<Theta> ; \<Phi> \<turnstile>\<^sub>w\<^sub>f AF_fun_typ_some bv ft"
-
-| wfFTI: "\<lbrakk>
- \<Theta> ; B \<turnstile>\<^sub>w\<^sub>f b;
- supp s \<subseteq> {atom x} \<union> supp B ;
- supp c \<subseteq> { atom x } ;
- \<Theta> ; B ; (x,b,c) #\<^sub>\<Gamma> GNil \<turnstile>\<^sub>w\<^sub>f \<tau>;
- \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>
-\<rbrakk> \<Longrightarrow>
- \<Theta> ; \<Phi> ; B \<turnstile>\<^sub>w\<^sub>f (AF_fun_typ x b c \<tau> s)"
-
-inductive_cases wfE_elims:
- "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f AE_val v : b"
- "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f AE_op Plus v1 v2 : b"
- "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f AE_op LEq v1 v2 : b"
- "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f AE_fst v1 : b"
- "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f AE_snd v1 : b"
- "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f AE_concat v1 v2 : b"
- "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f AE_len v1 : b"
- "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f AE_op opp v1 v2 : b"
- "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f AE_app f v: b"
- "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f AE_appP f b' v: b"
- "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f AE_mvar u : b"
- "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f AE_op Eq v1 v2 : b"
-
-inductive_cases wfCS_elims:
- "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> ; tid ; dc ; t \<turnstile>\<^sub>w\<^sub>f (cs::branch_s) : b"
- "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> ; tid ; dc \<turnstile>\<^sub>w\<^sub>f (cs::branch_list) : b"
-
-inductive_cases wfPhi_elims:
- " \<Theta> \<turnstile>\<^sub>w\<^sub>f []"
- " \<Theta> \<turnstile>\<^sub>w\<^sub>f ((AF_fundef f ft)#\<Pi>)"
- " \<Theta> \<turnstile>\<^sub>w\<^sub>f (fd#\<Phi>::\<Phi>)"
-
-declare[[ simproc del: alpha_lst]]
-
-inductive_cases wfFTQ_elims:
- "\<Theta> ; \<Phi> \<turnstile>\<^sub>w\<^sub>f AF_fun_typ_none ft"
- "\<Theta> ; \<Phi> \<turnstile>\<^sub>w\<^sub>f AF_fun_typ_some bv ft"
- "\<Theta> ; \<Phi> \<turnstile>\<^sub>w\<^sub>f AF_fun_typ_some bv (AF_fun_typ x b c \<tau> s)"
-
-inductive_cases wfFT_elims:
- "\<Theta> ; \<Phi> ; \<B> \<turnstile>\<^sub>w\<^sub>f AF_fun_typ x b c \<tau> s"
-
-declare[[ simproc add: alpha_lst]]
-
-inductive_cases wfD_elims:
- "\<Pi> ; \<B> ; (\<Gamma>::\<Gamma>) \<turnstile>\<^sub>w\<^sub>f []\<^sub>\<Delta>"
- "\<Pi> ; \<B> ; (\<Gamma>::\<Gamma>) \<turnstile>\<^sub>w\<^sub>f (u,\<tau>) #\<^sub>\<Delta> \<Delta>::\<Delta>"
-
-equivariance wfE
-
-nominal_inductive wfE
- avoids wfE_appPI: bv | wfS_varI: u | wfS_letI: x | wfS_let2I: x | wfS_branchI: x | wfS_assertI: x
-
-proof(goal_cases)
- case (1 \<Theta> \<Phi> \<B> \<Gamma> \<Delta> b' bv v \<tau> f x b c s)
- moreover hence "atom bv \<sharp> AE_appP f b' v" using pure_fresh fresh_prodN e.fresh by auto
- ultimately show ?case using fresh_star_def by fastforce
-next
- case (2 \<Theta> \<Phi> \<B> \<Gamma> \<Delta> b' bv v \<tau> f x b c s)
- then show ?case by auto
-next
- case (3 \<Phi> \<Theta> \<B> \<Gamma> \<Delta> e b' x s b)
- moreover hence "atom x \<sharp> LET x = e IN s" using fresh_prodN by auto
- ultimately show ?case using fresh_prodN fresh_star_def by fastforce
-next
- case (4 \<Phi> \<Theta> \<B> \<Gamma> \<Delta> e b' x s b)
- then show ?case by auto
-next
- case (5 \<Theta> \<Phi> \<B> x c \<Gamma> \<Delta> s b)
- hence "atom x \<sharp> ASSERT c IN s" using s_branch_s_branch_list.fresh by auto
- then show ?case using fresh_prodN fresh_star_def 5 by fastforce
-next
- case (6 \<Theta> \<Phi> \<B> x c \<Gamma> \<Delta> s b)
- then show ?case by auto
-next
- case (7 \<Phi> \<Theta> \<B> \<Gamma> \<Delta> s1 \<tau> x s2 b)
- hence "atom x \<sharp> \<tau> \<and> atom x \<sharp> s1" using fresh_prodN by metis
- moreover hence "atom x \<sharp> LET x : \<tau> = s1 IN s2"
- using s_branch_s_branch_list.fresh(3)[of "atom x" x "\<tau>" s1 s2 ] fresh_prodN by simp
- ultimately show ?case using fresh_prodN fresh_star_def 7 by fastforce
-next
- case (8 \<Phi> \<Theta> \<B> \<Gamma> \<Delta> s1 \<tau> x s2 b)
- then show ?case by auto
-next
- case (9 \<Theta> \<B> \<Gamma> \<tau> v u \<Phi> \<Delta> b s)
- moreover hence " atom u \<sharp> AS_var u \<tau> v s" using fresh_prodN s_branch_s_branch_list.fresh by simp
- ultimately show ?case using fresh_star_def fresh_prodN s_branch_s_branch_list.fresh by fastforce
-next
- case (10 \<Theta> \<B> \<Gamma> \<tau> v u \<Phi> \<Delta> b s)
- then show ?case by auto
-next
- case (11 \<Phi> \<Theta> \<B> x \<tau> \<Gamma> \<Delta> s b tid dc)
- moreover have "atom x \<sharp> (dc x \<Rightarrow> s)" using pure_fresh s_branch_s_branch_list.fresh by auto
- ultimately show ?case using fresh_prodN fresh_star_def pure_fresh by fastforce
-next
- case (12 \<Phi> \<Theta> \<B> x \<tau> \<Gamma> \<Delta> s b tid dc)
- then show ?case by auto
-qed
-
-inductive wfVDs :: "var_def list \<Rightarrow> bool" where
-
-wfVDs_nilI: "wfVDs []"
-
-| wfVDs_consI: "\<lbrakk>
- atom u \<sharp> ts;
- wfV ([]::\<Theta>) {||} GNil v (b_of \<tau>);
- wfT ([]::\<Theta>) {||} GNil \<tau>;
- wfVDs ts
-\<rbrakk> \<Longrightarrow> wfVDs ((AV_def u \<tau> v)#ts)"
-
-equivariance wfVDs
-nominal_inductive wfVDs .
-
+(*<*)
+theory Wellformed
+ imports IVSubst BTVSubst
+begin
+ (*>*)
+
+chapter \<open>Wellformed Terms\<close>
+
+text \<open>We require that expressions and values are well-formed. This includes a notion
+of well-sortedness. We identify a sort with a basic type and
+define the judgement as two clusters of mutually recursive
+inductive predicates. Some of the proofs are across all of the predicates and
+although they seemed at first to be daunting, they have all
+worked out well.\<close>
+
+named_theorems ms_wb "Facts for helping with well-sortedness"
+
+section \<open>Definitions\<close>
+
+inductive wfV :: "\<Theta> \<Rightarrow> \<B> \<Rightarrow> \<Gamma> \<Rightarrow> v \<Rightarrow> b \<Rightarrow> bool" (" _ ; _ ; _ \<turnstile>\<^sub>w\<^sub>f _ : _ " [50,50,50] 50) and
+ wfC :: "\<Theta> \<Rightarrow> \<B> \<Rightarrow> \<Gamma> \<Rightarrow> c \<Rightarrow> bool" (" _ ; _ ; _ \<turnstile>\<^sub>w\<^sub>f _ " [50,50] 50) and
+ wfG :: "\<Theta> \<Rightarrow> \<B> \<Rightarrow> \<Gamma> \<Rightarrow> bool" (" _ ; _ \<turnstile>\<^sub>w\<^sub>f _ " [50,50] 50) and
+ wfT :: "\<Theta> \<Rightarrow> \<B> \<Rightarrow> \<Gamma> \<Rightarrow> \<tau> \<Rightarrow> bool" (" _ ; _ ; _ \<turnstile>\<^sub>w\<^sub>f _ " [50,50] 50) and
+ wfTs :: "\<Theta> \<Rightarrow> \<B> \<Rightarrow> \<Gamma> \<Rightarrow> (string*\<tau>) list \<Rightarrow> bool" (" _ ; _ ; _ \<turnstile>\<^sub>w\<^sub>f _ " [50,50] 50) and
+ wfTh :: "\<Theta> \<Rightarrow> bool" (" \<turnstile>\<^sub>w\<^sub>f _ " [50] 50) and
+ wfB :: "\<Theta> \<Rightarrow> \<B> \<Rightarrow> b \<Rightarrow> bool" (" _ ; _ \<turnstile>\<^sub>w\<^sub>f _ " [50,50] 50) and
+ wfCE :: "\<Theta> \<Rightarrow> \<B> \<Rightarrow> \<Gamma> \<Rightarrow> ce \<Rightarrow> b \<Rightarrow> bool" (" _ ; _ ; _ \<turnstile>\<^sub>w\<^sub>f _ : _ " [50,50,50] 50) and
+ wfTD :: "\<Theta> \<Rightarrow> type_def \<Rightarrow> bool" (" _ \<turnstile>\<^sub>w\<^sub>f _ " [50,50] 50)
+ where
+
+wfB_intI: "\<turnstile>\<^sub>w\<^sub>f \<Theta> \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f B_int"
+| wfB_boolI: "\<turnstile>\<^sub>w\<^sub>f \<Theta> \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f B_bool"
+| wfB_unitI: "\<turnstile>\<^sub>w\<^sub>f \<Theta> \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f B_unit"
+| wfB_bitvecI: "\<turnstile>\<^sub>w\<^sub>f \<Theta> \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f B_bitvec"
+| wfB_pairI: "\<lbrakk> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b1 ; \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b2 \<rbrakk> \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f B_pair b1 b2"
+
+| wfB_consI: "\<lbrakk>
+ \<turnstile>\<^sub>w\<^sub>f \<Theta>;
+ (AF_typedef s dclist) \<in> set \<Theta>
+\<rbrakk> \<Longrightarrow>
+ \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f B_id s"
+
+| wfB_appI: "\<lbrakk>
+ \<turnstile>\<^sub>w\<^sub>f \<Theta>;
+ \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b;
+ (AF_typedef_poly s bv dclist) \<in> set \<Theta>
+\<rbrakk> \<Longrightarrow>
+ \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f B_app s b"
+
+| wfV_varI: "\<lbrakk> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> ; Some (b,c) = lookup \<Gamma> x \<rbrakk> \<Longrightarrow> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f V_var x : b"
+| wfV_litI: "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<Longrightarrow> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f V_lit l : base_for_lit l"
+
+| wfV_pairI: "\<lbrakk>
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v1 : b1 ;
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v2 : b2
+\<rbrakk> \<Longrightarrow>
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f (V_pair v1 v2) : B_pair b1 b2"
+
+| wfV_consI: "\<lbrakk>
+ AF_typedef s dclist \<in> set \<Theta>;
+ (dc, \<lbrace> x : b' | c \<rbrace>) \<in> set dclist ;
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b'
+\<rbrakk> \<Longrightarrow>
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f V_cons s dc v : B_id s"
+
+| wfV_conspI: "\<lbrakk>
+ AF_typedef_poly s bv dclist \<in> set \<Theta>;
+ (dc, \<lbrace> x : b' | c \<rbrace>) \<in> set dclist ;
+ \<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f b;
+ atom bv \<sharp> (\<Theta>, \<B>, \<Gamma>, b , v);
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b'[bv::=b]\<^sub>b\<^sub>b
+\<rbrakk> \<Longrightarrow>
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f V_consp s dc b v : B_app s b"
+
+| wfCE_valI : "\<lbrakk>
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b
+\<rbrakk> \<Longrightarrow>
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_val v : b"
+
+| wfCE_plusI: "\<lbrakk>
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v1 : B_int;
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v2 : B_int
+\<rbrakk> \<Longrightarrow>
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_op Plus v1 v2 : B_int"
+
+| wfCE_leqI:"\<lbrakk>
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v1 : B_int;
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v2 : B_int
+\<rbrakk> \<Longrightarrow>
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_op LEq v1 v2 : B_bool"
+
+| wfCE_eqI:"\<lbrakk>
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v1 : b;
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v2 : b
+\<rbrakk> \<Longrightarrow>
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_op Eq v1 v2 : B_bool"
+
+| wfCE_fstI: "\<lbrakk>
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v1 : B_pair b1 b2
+\<rbrakk> \<Longrightarrow>
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_fst v1 : b1"
+
+| wfCE_sndI: "\<lbrakk>
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v1 : B_pair b1 b2
+\<rbrakk> \<Longrightarrow>
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_snd v1 : b2"
+
+| wfCE_concatI: "\<lbrakk>
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v1 : B_bitvec ;
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v2 : B_bitvec
+\<rbrakk> \<Longrightarrow>
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_concat v1 v2 : B_bitvec"
+
+| wfCE_lenI: "\<lbrakk>
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v1 : B_bitvec
+\<rbrakk> \<Longrightarrow>
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_len v1 : B_int"
+
+| wfTI : "\<lbrakk>
+ atom z \<sharp> (\<Theta>, \<B>, \<Gamma>) ;
+ \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b;
+ \<Theta>; \<B> ; (z,b,C_true) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f c
+\<rbrakk> \<Longrightarrow>
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b | c \<rbrace>"
+
+| wfC_eqI: "\<lbrakk>
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f e1 : b ;
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f e2 : b \<rbrakk> \<Longrightarrow>
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f C_eq e1 e2"
+| wfC_trueI: " \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<Longrightarrow> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f C_true "
+| wfC_falseI: " \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<Longrightarrow> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f C_false "
+
+| wfC_conjI: "\<lbrakk> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f c1 ; \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f c2 \<rbrakk> \<Longrightarrow> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f C_conj c1 c2"
+| wfC_disjI: "\<lbrakk> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f c1 ; \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f c2 \<rbrakk> \<Longrightarrow> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f C_disj c1 c2"
+| wfC_notI: "\<lbrakk> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f c1 \<rbrakk> \<Longrightarrow> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f C_not c1"
+| wfC_impI: "\<lbrakk> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f c1 ;
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f c2 \<rbrakk> \<Longrightarrow> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f C_imp c1 c2"
+
+| wfG_nilI: " \<turnstile>\<^sub>w\<^sub>f \<Theta> \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f GNil"
+| wfG_cons1I: "\<lbrakk> c \<notin> { TRUE, FALSE } ;
+ \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> ;
+ atom x \<sharp> \<Gamma> ;
+ \<Theta> ; \<B> ; (x,b,C_true)#\<^sub>\<Gamma>\<Gamma> \<turnstile>\<^sub>w\<^sub>f c ; wfB \<Theta> \<B> b
+ \<rbrakk> \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f ((x,b,c)#\<^sub>\<Gamma>\<Gamma>)"
+| wfG_cons2I: "\<lbrakk> c \<in> { TRUE, FALSE } ;
+ \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> ;
+ atom x \<sharp> \<Gamma> ;
+ wfB \<Theta> \<B> b
+ \<rbrakk> \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f ((x,b,c)#\<^sub>\<Gamma>\<Gamma>)"
+
+| wfTh_emptyI: " \<turnstile>\<^sub>w\<^sub>f []"
+
+| wfTh_consI: "\<lbrakk>
+ (name_of_type tdef) \<notin> name_of_type ` set \<Theta> ;
+ \<turnstile>\<^sub>w\<^sub>f \<Theta> ;
+ \<Theta> \<turnstile>\<^sub>w\<^sub>f tdef \<rbrakk> \<Longrightarrow> \<turnstile>\<^sub>w\<^sub>f tdef#\<Theta>"
+
+| wfTD_simpleI: "\<lbrakk>
+ \<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f lst
+ \<rbrakk> \<Longrightarrow>
+ \<Theta> \<turnstile>\<^sub>w\<^sub>f (AF_typedef s lst )"
+
+| wfTD_poly: "\<lbrakk>
+ \<Theta> ; {|bv|} ; GNil \<turnstile>\<^sub>w\<^sub>f lst
+ \<rbrakk> \<Longrightarrow>
+ \<Theta> \<turnstile>\<^sub>w\<^sub>f (AF_typedef_poly s bv lst)"
+
+| wfTs_nil: "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<Longrightarrow> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f []::(string*\<tau>) list"
+| wfTs_cons: "\<lbrakk> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau> ;
+ dc \<notin> fst ` set ts;
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ts::(string*\<tau>) list \<rbrakk> \<Longrightarrow> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ((dc,\<tau>)#ts)"
+
+inductive_cases wfC_elims:
+ "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f C_true"
+ "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f C_false"
+ "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f C_eq e1 e2"
+ "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f C_conj c1 c2"
+ "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f C_disj c1 c2"
+ "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f C_not c1"
+ "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f C_imp c1 c2"
+
+inductive_cases wfV_elims:
+ "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f V_var x : b"
+ "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f V_lit l : b"
+ "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f V_pair v1 v2 : b"
+ "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f V_cons tyid dc v : b"
+ "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f V_consp tyid dc b v : b'"
+
+inductive_cases wfCE_elims:
+ " \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_val v : b"
+ " \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_op Plus v1 v2 : b"
+ " \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_op LEq v1 v2 : b"
+ " \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_fst v1 : b"
+ " \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_snd v1 : b"
+ " \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_concat v1 v2 : b"
+ " \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_len v1 : b"
+ " \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_op opp v1 v2 : b"
+ " \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_op Eq v1 v2 : b"
+
+inductive_cases wfT_elims:
+ "\<Theta>; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau>::\<tau>"
+ "\<Theta>; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b | c \<rbrace>"
+
+inductive_cases wfG_elims:
+ "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f GNil"
+ "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (x,b,c)#\<^sub>\<Gamma>\<Gamma>"
+ "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (x,b,TRUE)#\<^sub>\<Gamma>\<Gamma>"
+ "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (x,b,FALSE)#\<^sub>\<Gamma>\<Gamma>"
+
+inductive_cases wfTh_elims:
+ " \<turnstile>\<^sub>w\<^sub>f []"
+ " \<turnstile>\<^sub>w\<^sub>f td#\<Pi>"
+
+inductive_cases wfTD_elims:
+ "\<Theta> \<turnstile>\<^sub>w\<^sub>f (AF_typedef s lst )"
+ "\<Theta> \<turnstile>\<^sub>w\<^sub>f (AF_typedef_poly s bv lst )"
+
+inductive_cases wfTs_elims:
+ "\<Theta>; \<B> ; GNil \<turnstile>\<^sub>w\<^sub>f ([]::((string*\<tau>) list))"
+ "\<Theta>; \<B> ; GNil \<turnstile>\<^sub>w\<^sub>f ((t#ts)::((string*\<tau>) list))"
+
+inductive_cases wfB_elims:
+ " \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f B_pair b1 b2"
+ " \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f B_id s"
+ " \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f B_app s b"
+
+equivariance wfV
+
+text \<open>This setup of 'avoiding' is not complete and for some of lemmas, such as weakening,
+do it the hard way\<close>
+nominal_inductive wfV
+ avoids wfV_conspI: bv | wfTI: z
+proof(goal_cases)
+ case (1 s bv dclist \<Theta> dc x b' c \<B> b \<Gamma> v)
+
+ moreover hence "atom bv \<sharp> V_consp s dc b v" using v.fresh fresh_prodN pure_fresh by metis
+ moreover have "atom bv \<sharp> B_app s b" using b.fresh fresh_prodN pure_fresh 1 by metis
+ ultimately show ?case using b.fresh v.fresh pure_fresh fresh_star_def fresh_prodN by fastforce
+next
+ case (2 s bv dclist \<Theta> dc x b' c \<B> b \<Gamma> v)
+ then show ?case by auto
+next
+ case (3 z \<Gamma> \<Theta> \<B> b c)
+ then show ?case using \<tau>.fresh fresh_star_def fresh_prodN by fastforce
+next
+ case (4 z \<Gamma> \<Theta> \<B> b c)
+ then show ?case by auto
+qed
+
+inductive
+ wfE :: "\<Theta> \<Rightarrow> \<Phi> \<Rightarrow> \<B> \<Rightarrow> \<Gamma> \<Rightarrow> \<Delta> \<Rightarrow> e \<Rightarrow> b \<Rightarrow> bool" (" _ ; _ ; _ ; _ ; _ \<turnstile>\<^sub>w\<^sub>f _ : _ " [50,50,50] 50) and
+ wfS :: "\<Theta> \<Rightarrow> \<Phi> \<Rightarrow> \<B> \<Rightarrow> \<Gamma> \<Rightarrow> \<Delta> \<Rightarrow> s \<Rightarrow> b \<Rightarrow> bool" (" _ ; _ ; _ ; _ ; _ \<turnstile>\<^sub>w\<^sub>f _ : _ " [50,50,50] 50) and
+ wfCS :: "\<Theta> \<Rightarrow> \<Phi> \<Rightarrow> \<B> \<Rightarrow> \<Gamma> \<Rightarrow> \<Delta> \<Rightarrow> tyid \<Rightarrow> string \<Rightarrow> \<tau> \<Rightarrow> branch_s \<Rightarrow> b \<Rightarrow> bool" (" _ ; _ ; _ ; _ ; _ ; _ ; _ ; _ \<turnstile>\<^sub>w\<^sub>f _ : _ " [50,50,50,50,50] 50) and
+ wfCSS :: "\<Theta> \<Rightarrow> \<Phi> \<Rightarrow> \<B> \<Rightarrow> \<Gamma> \<Rightarrow> \<Delta> \<Rightarrow> tyid \<Rightarrow> (string * \<tau>) list \<Rightarrow> branch_list \<Rightarrow> b \<Rightarrow> bool" (" _ ; _ ; _ ; _ ; _ ; _ ; _ \<turnstile>\<^sub>w\<^sub>f _ : _ " [50,50,50,50,50] 50) and
+ wfPhi :: "\<Theta> \<Rightarrow> \<Phi> \<Rightarrow> bool" (" _ \<turnstile>\<^sub>w\<^sub>f _ " [50,50] 50) and
+ wfD :: "\<Theta> \<Rightarrow> \<B> \<Rightarrow> \<Gamma> \<Rightarrow> \<Delta> \<Rightarrow> bool" (" _ ; _ ; _ \<turnstile>\<^sub>w\<^sub>f _ " [50,50] 50) and
+ wfFTQ :: "\<Theta> \<Rightarrow> \<Phi> \<Rightarrow> fun_typ_q \<Rightarrow> bool" (" _ ; _ \<turnstile>\<^sub>w\<^sub>f _ " [50] 50) and
+ wfFT :: "\<Theta> \<Rightarrow> \<Phi> \<Rightarrow> \<B> \<Rightarrow> fun_typ \<Rightarrow> bool" (" _ ; _ ; _ \<turnstile>\<^sub>w\<^sub>f _ " [50] 50) where
+
+wfE_valI : "\<lbrakk>
+ \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>;
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>;
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b
+\<rbrakk> \<Longrightarrow>
+ \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f AE_val v : b"
+
+| wfE_plusI: "\<lbrakk>
+ \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>;
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>;
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v1 : B_int;
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v2 : B_int
+\<rbrakk> \<Longrightarrow>
+ \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f AE_op Plus v1 v2 : B_int"
+
+| wfE_leqI:"\<lbrakk>
+ \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> ;
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>;
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v1 : B_int;
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v2 : B_int
+\<rbrakk> \<Longrightarrow>
+ \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f AE_op LEq v1 v2 : B_bool"
+
+| wfE_eqI:"\<lbrakk>
+ \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> ;
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>;
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v1 : b;
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v2 : b;
+ b \<in> { B_bool, B_int, B_unit }
+\<rbrakk> \<Longrightarrow>
+ \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f AE_op Eq v1 v2 : B_bool"
+
+| wfE_fstI: "\<lbrakk>
+ \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>;
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>;
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v1 : B_pair b1 b2
+ \<rbrakk> \<Longrightarrow>
+ \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f AE_fst v1 : b1"
+
+| wfE_sndI: "\<lbrakk>
+ \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> ;
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>;
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v1 : B_pair b1 b2
+\<rbrakk> \<Longrightarrow>
+ \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f AE_snd v1 : b2"
+
+| wfE_concatI: "\<lbrakk>
+ \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> ;
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>;
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v1 : B_bitvec;
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v2 : B_bitvec
+\<rbrakk> \<Longrightarrow>
+ \<Theta> ; \<Phi> ; \<B> ; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f AE_concat v1 v2 : B_bitvec"
+
+| wfE_splitI: "\<lbrakk>
+ \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> ;
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>;
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v1 : B_bitvec;
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v2 : B_int
+\<rbrakk> \<Longrightarrow>
+ \<Theta> ; \<Phi> ; \<B> ; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f AE_split v1 v2 : B_pair B_bitvec B_bitvec"
+
+| wfE_lenI: "\<lbrakk>
+ \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> ;
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>;
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v1 : B_bitvec
+\<rbrakk> \<Longrightarrow>
+ \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f AE_len v1 : B_int"
+
+| wfE_appI: "\<lbrakk>
+ \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> ;
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>;
+ Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \<tau> s))) = lookup_fun \<Phi> f ;
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b
+\<rbrakk> \<Longrightarrow>
+ \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f AE_app f v : b_of \<tau>"
+
+| wfE_appPI: "\<lbrakk>
+ \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> ;
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>;
+ \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b';
+ atom bv \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>, \<Delta>, b', v, (b_of \<tau>)[bv::=b']\<^sub>b);
+ Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \<tau> s))) = lookup_fun \<Phi> f;
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : (b[bv::=b']\<^sub>b)
+\<rbrakk> \<Longrightarrow>
+ \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f (AE_appP f b' v) : ((b_of \<tau>)[bv::=b']\<^sub>b)"
+
+| wfE_mvarI: "\<lbrakk>
+ \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> ;
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>;
+ (u,\<tau>) \<in> setD \<Delta>
+\<rbrakk> \<Longrightarrow>
+ \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f AE_mvar u : b_of \<tau>"
+
+| wfS_valI: "\<lbrakk>
+ \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> ;
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b ;
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>
+\<rbrakk> \<Longrightarrow>
+ \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f (AS_val v) : b"
+
+| wfS_letI: "\<lbrakk>
+ wfE \<Theta> \<Phi> \<B> \<Gamma> \<Delta> e b' ;
+ \<Theta> ; \<Phi> ; \<B> ; (x,b',C_true) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b;
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> ;
+ atom x \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>, \<Delta>, e, b)
+\<rbrakk> \<Longrightarrow>
+ \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f LET x = e IN s : b"
+
+| wfS_assertI: "\<lbrakk>
+ \<Theta> ; \<Phi> ; \<B> ; (x,B_bool,c) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b;
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f c ;
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> ;
+ atom x \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>, \<Delta>, c, b, s)
+\<rbrakk> \<Longrightarrow>
+ \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f ASSERT c IN s : b"
+
+| wfS_let2I: "\<lbrakk>
+ \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f s1 : b_of \<tau> ;
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau>;
+ \<Theta> ; \<Phi> ; \<B> ; (x,b_of \<tau>,C_true) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s2 : b ;
+ atom x \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>, \<Delta>, s1, b,\<tau>)
+\<rbrakk> \<Longrightarrow>
+ \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f LET x : \<tau> = s1 IN s2 : b"
+
+| wfS_ifI: "\<lbrakk>
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : B_bool;
+ \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f s1 : b ;
+ \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f s2 : b
+\<rbrakk> \<Longrightarrow>
+ \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f IF v THEN s1 ELSE s2 : b"
+
+| wfS_varI : "\<lbrakk>
+ wfT \<Theta> \<B> \<Gamma> \<tau> ;
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b_of \<tau>;
+ atom u \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>, \<Delta>, \<tau>, v, b);
+ \<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; (u,\<tau>)#\<^sub>\<Delta>\<Delta> \<turnstile>\<^sub>w\<^sub>f s : b
+\<rbrakk> \<Longrightarrow>
+ \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f VAR u : \<tau> = v IN s : b "
+
+| wfS_assignI: "\<lbrakk>
+ (u,\<tau>) \<in> setD \<Delta> ;
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> ;
+ \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> ;
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b_of \<tau>
+\<rbrakk> \<Longrightarrow>
+ \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f u ::= v : B_unit"
+
+| wfS_whileI: "\<lbrakk>
+ \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f s1 : B_bool ;
+ \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f s2 : b
+\<rbrakk> \<Longrightarrow>
+ \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f WHILE s1 DO { s2 } : b"
+
+| wfS_seqI: "\<lbrakk>
+ \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f s1 : B_unit ;
+ \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f s2 : b
+\<rbrakk> \<Longrightarrow>
+ \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f s1 ;; s2 : b"
+
+| wfS_matchI: "\<lbrakk>
+ wfV \<Theta> \<B> \<Gamma> v (B_id tid) ;
+ (AF_typedef tid dclist ) \<in> set \<Theta>;
+ wfD \<Theta> \<B> \<Gamma> \<Delta> ;
+ \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> ;
+ \<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> ; tid ; dclist \<turnstile>\<^sub>w\<^sub>f cs : b
+\<rbrakk> \<Longrightarrow>
+ \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f AS_match v cs : b "
+
+| wfS_branchI: "\<lbrakk>
+ \<Theta> ; \<Phi> ; \<B> ; (x,b_of \<tau>,C_true) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b ;
+ atom x \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>, \<Delta>, \<Gamma>,\<tau>);
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>
+\<rbrakk> \<Longrightarrow>
+ \<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> ; tid ; dc ; \<tau> \<turnstile>\<^sub>w\<^sub>f dc x \<Rightarrow> s : b"
+
+| wfS_finalI: "\<lbrakk>
+ \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> ; tid ; dc ; t \<turnstile>\<^sub>w\<^sub>f cs : b
+ \<rbrakk> \<Longrightarrow>
+ \<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> ; tid ; [(dc,t)] \<turnstile>\<^sub>w\<^sub>f AS_final cs : b "
+
+| wfS_cons: "\<lbrakk>
+ \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> ; tid ; dc ; t \<turnstile>\<^sub>w\<^sub>f cs : b;
+ \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> ; tid ; dclist \<turnstile>\<^sub>w\<^sub>f css : b
+ \<rbrakk> \<Longrightarrow>
+ \<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> ; tid ; (dc,t)#dclist \<turnstile>\<^sub>w\<^sub>f AS_cons cs css : b "
+
+| wfD_emptyI: "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<Longrightarrow> \<Theta> ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f []\<^sub>\<Delta>"
+
+| wfD_cons: "\<lbrakk>
+ \<Theta> ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>::\<Delta> ;
+ \<Theta> ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau>;
+ u \<notin> fst ` setD \<Delta>
+\<rbrakk> \<Longrightarrow>
+ \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ((u,\<tau>) #\<^sub>\<Delta> \<Delta>)"
+
+| wfPhi_emptyI: " \<turnstile>\<^sub>w\<^sub>f \<Theta> \<Longrightarrow> \<Theta> \<turnstile>\<^sub>w\<^sub>f []"
+
+| wfPhi_consI: "\<lbrakk>
+ f \<notin> name_of_fun ` set \<Phi>;
+ \<Theta> ; \<Phi> \<turnstile>\<^sub>w\<^sub>f ft;
+ \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>
+\<rbrakk> \<Longrightarrow>
+ \<Theta> \<turnstile>\<^sub>w\<^sub>f ((AF_fundef f ft)#\<Phi>)"
+
+| wfFTNone: " \<Theta> ; \<Phi> ; {||} \<turnstile>\<^sub>w\<^sub>f ft \<Longrightarrow> \<Theta> ; \<Phi> \<turnstile>\<^sub>w\<^sub>f AF_fun_typ_none ft"
+| wfFTSome: " \<Theta> ; \<Phi> ; {| bv |} \<turnstile>\<^sub>w\<^sub>f ft \<Longrightarrow> \<Theta> ; \<Phi> \<turnstile>\<^sub>w\<^sub>f AF_fun_typ_some bv ft"
+
+| wfFTI: "\<lbrakk>
+ \<Theta> ; B \<turnstile>\<^sub>w\<^sub>f b;
+ supp s \<subseteq> {atom x} \<union> supp B ;
+ supp c \<subseteq> { atom x } ;
+ \<Theta> ; B ; (x,b,c) #\<^sub>\<Gamma> GNil \<turnstile>\<^sub>w\<^sub>f \<tau>;
+ \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>
+\<rbrakk> \<Longrightarrow>
+ \<Theta> ; \<Phi> ; B \<turnstile>\<^sub>w\<^sub>f (AF_fun_typ x b c \<tau> s)"
+
+inductive_cases wfE_elims:
+ "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f AE_val v : b"
+ "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f AE_op Plus v1 v2 : b"
+ "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f AE_op LEq v1 v2 : b"
+ "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f AE_fst v1 : b"
+ "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f AE_snd v1 : b"
+ "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f AE_concat v1 v2 : b"
+ "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f AE_len v1 : b"
+ "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f AE_op opp v1 v2 : b"
+ "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f AE_app f v: b"
+ "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f AE_appP f b' v: b"
+ "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f AE_mvar u : b"
+ "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f AE_op Eq v1 v2 : b"
+
+inductive_cases wfCS_elims:
+ "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> ; tid ; dc ; t \<turnstile>\<^sub>w\<^sub>f (cs::branch_s) : b"
+ "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> ; tid ; dc \<turnstile>\<^sub>w\<^sub>f (cs::branch_list) : b"
+
+inductive_cases wfPhi_elims:
+ " \<Theta> \<turnstile>\<^sub>w\<^sub>f []"
+ " \<Theta> \<turnstile>\<^sub>w\<^sub>f ((AF_fundef f ft)#\<Pi>)"
+ " \<Theta> \<turnstile>\<^sub>w\<^sub>f (fd#\<Phi>::\<Phi>)"
+
+declare[[ simproc del: alpha_lst]]
+
+inductive_cases wfFTQ_elims:
+ "\<Theta> ; \<Phi> \<turnstile>\<^sub>w\<^sub>f AF_fun_typ_none ft"
+ "\<Theta> ; \<Phi> \<turnstile>\<^sub>w\<^sub>f AF_fun_typ_some bv ft"
+ "\<Theta> ; \<Phi> \<turnstile>\<^sub>w\<^sub>f AF_fun_typ_some bv (AF_fun_typ x b c \<tau> s)"
+
+inductive_cases wfFT_elims:
+ "\<Theta> ; \<Phi> ; \<B> \<turnstile>\<^sub>w\<^sub>f AF_fun_typ x b c \<tau> s"
+
+declare[[ simproc add: alpha_lst]]
+
+inductive_cases wfD_elims:
+ "\<Pi> ; \<B> ; (\<Gamma>::\<Gamma>) \<turnstile>\<^sub>w\<^sub>f []\<^sub>\<Delta>"
+ "\<Pi> ; \<B> ; (\<Gamma>::\<Gamma>) \<turnstile>\<^sub>w\<^sub>f (u,\<tau>) #\<^sub>\<Delta> \<Delta>::\<Delta>"
+
+equivariance wfE
+
+nominal_inductive wfE
+ avoids wfE_appPI: bv | wfS_varI: u | wfS_letI: x | wfS_let2I: x | wfS_branchI: x | wfS_assertI: x
+
+proof(goal_cases)
+ case (1 \<Theta> \<Phi> \<B> \<Gamma> \<Delta> b' bv v \<tau> f x b c s)
+ moreover hence "atom bv \<sharp> AE_appP f b' v" using pure_fresh fresh_prodN e.fresh by auto
+ ultimately show ?case using fresh_star_def by fastforce
+next
+ case (2 \<Theta> \<Phi> \<B> \<Gamma> \<Delta> b' bv v \<tau> f x b c s)
+ then show ?case by auto
+next
+ case (3 \<Phi> \<Theta> \<B> \<Gamma> \<Delta> e b' x s b)
+ moreover hence "atom x \<sharp> LET x = e IN s" using fresh_prodN by auto
+ ultimately show ?case using fresh_prodN fresh_star_def by fastforce
+next
+ case (4 \<Phi> \<Theta> \<B> \<Gamma> \<Delta> e b' x s b)
+ then show ?case by auto
+next
+ case (5 \<Theta> \<Phi> \<B> x c \<Gamma> \<Delta> s b)
+ hence "atom x \<sharp> ASSERT c IN s" using s_branch_s_branch_list.fresh by auto
+ then show ?case using fresh_prodN fresh_star_def 5 by fastforce
+next
+ case (6 \<Theta> \<Phi> \<B> x c \<Gamma> \<Delta> s b)
+ then show ?case by auto
+next
+ case (7 \<Phi> \<Theta> \<B> \<Gamma> \<Delta> s1 \<tau> x s2 b)
+ hence "atom x \<sharp> \<tau> \<and> atom x \<sharp> s1" using fresh_prodN by metis
+ moreover hence "atom x \<sharp> LET x : \<tau> = s1 IN s2"
+ using s_branch_s_branch_list.fresh(3)[of "atom x" x "\<tau>" s1 s2 ] fresh_prodN by simp
+ ultimately show ?case using fresh_prodN fresh_star_def 7 by fastforce
+next
+ case (8 \<Phi> \<Theta> \<B> \<Gamma> \<Delta> s1 \<tau> x s2 b)
+ then show ?case by auto
+next
+ case (9 \<Theta> \<B> \<Gamma> \<tau> v u \<Phi> \<Delta> b s)
+ moreover hence " atom u \<sharp> AS_var u \<tau> v s" using fresh_prodN s_branch_s_branch_list.fresh by simp
+ ultimately show ?case using fresh_star_def fresh_prodN s_branch_s_branch_list.fresh by fastforce
+next
+ case (10 \<Theta> \<B> \<Gamma> \<tau> v u \<Phi> \<Delta> b s)
+ then show ?case by auto
+next
+ case (11 \<Phi> \<Theta> \<B> x \<tau> \<Gamma> \<Delta> s b tid dc)
+ moreover have "atom x \<sharp> (dc x \<Rightarrow> s)" using pure_fresh s_branch_s_branch_list.fresh by auto
+ ultimately show ?case using fresh_prodN fresh_star_def pure_fresh by fastforce
+next
+ case (12 \<Phi> \<Theta> \<B> x \<tau> \<Gamma> \<Delta> s b tid dc)
+ then show ?case by auto
+qed
+
+inductive wfVDs :: "var_def list \<Rightarrow> bool" where
+
+wfVDs_nilI: "wfVDs []"
+
+| wfVDs_consI: "\<lbrakk>
+ atom u \<sharp> ts;
+ wfV ([]::\<Theta>) {||} GNil v (b_of \<tau>);
+ wfT ([]::\<Theta>) {||} GNil \<tau>;
+ wfVDs ts
+\<rbrakk> \<Longrightarrow> wfVDs ((AV_def u \<tau> v)#ts)"
+
+equivariance wfVDs
+nominal_inductive wfVDs .
+
end
\ No newline at end of file
diff --git a/thys/MiniSail/WellformedL.thy b/thys/MiniSail/WellformedL.thy
--- a/thys/MiniSail/WellformedL.thy
+++ b/thys/MiniSail/WellformedL.thy
@@ -1,4755 +1,4755 @@
-(*<*)
-theory WellformedL
- imports Wellformed "SyntaxL"
-begin
-(*>*)
-
-chapter \<open>Wellformedness Lemmas\<close>
-
-section \<open>Prelude\<close>
-
-lemma b_of_subst_bb_commute:
- "(b_of (\<tau>[bv::=b]\<^sub>\<tau>\<^sub>b)) = (b_of \<tau>)[bv::=b]\<^sub>b\<^sub>b"
-proof -
- obtain z' and b' and c' where "\<tau> = \<lbrace> z' : b' | c' \<rbrace> " using obtain_fresh_z by metis
- moreover hence "(b_of (\<tau>[bv::=b]\<^sub>\<tau>\<^sub>b)) = b_of \<lbrace> z' : b'[bv::=b]\<^sub>b\<^sub>b | c' \<rbrace>" using subst_tb.simps by simp
- ultimately show ?thesis using subst_tv.simps subst_tb.simps by simp
-qed
-
-lemmas wf_intros = wfV_wfC_wfG_wfT_wfTs_wfTh_wfB_wfCE_wfTD.intros wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.intros
-
-lemmas freshers = fresh_prodN b.fresh c.fresh v.fresh ce.fresh fresh_GCons fresh_GNil fresh_at_base
-
-section \<open>Strong Elimination\<close>
-
-text \<open>Inversion/elimination for well-formed polymorphic constructors \<close>
-lemma wf_strong_elim:
- fixes \<Gamma>::\<Gamma> and \<Gamma>'::\<Gamma> and v::v and e::e and c::c and \<tau>::\<tau> and ts::"(string*\<tau>) list"
- and \<Delta>::\<Delta> and b::b and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def and s::s and tm::"'a::fs"
- and cs::branch_s and css::branch_list and \<Theta>::\<Theta>
- shows "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f (V_consp tyid dc b v) : b'' \<Longrightarrow> (\<exists> bv dclist x b' c. b'' = B_app tyid b \<and>
- AF_typedef_poly tyid bv dclist \<in> set \<Theta> \<and>
- (dc, \<lbrace> x : b' | c \<rbrace>) \<in> set dclist \<and>
- \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b \<and> atom bv \<sharp> (\<Theta>, \<B>, \<Gamma>, b, v) \<and> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b'[bv::=b]\<^sub>b\<^sub>b \<and> atom bv \<sharp> tm)" and
- "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f c \<Longrightarrow> True" and
- "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<Longrightarrow> True" and
- "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau> \<Longrightarrow> True" and
- "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ts \<Longrightarrow> True" and
- "\<turnstile>\<^sub>w\<^sub>f \<Theta> \<Longrightarrow>True" and
- "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b \<Longrightarrow> True " and
- "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ce : b' \<Longrightarrow> True" and
- "\<Theta> \<turnstile>\<^sub>w\<^sub>f td \<Longrightarrow> True"
-proof(nominal_induct
- "V_consp tyid dc b v" b'' and c and \<Gamma> and \<tau> and ts and \<Theta> and b and b' and td
- avoiding: tm
-
-rule:wfV_wfC_wfG_wfT_wfTs_wfTh_wfB_wfCE_wfTD.strong_induct)
- case (wfV_conspI bv dclist \<Theta> x b' c \<B> \<Gamma>)
- then show ?case by force
-qed(auto+)
-
-section \<open>Context Extension\<close>
-
-definition wfExt :: "\<Theta> \<Rightarrow> \<B> \<Rightarrow> \<Gamma> \<Rightarrow> \<Gamma> \<Rightarrow> bool" (" _ ; _ \<turnstile>\<^sub>w\<^sub>f _ < _ " [50,50,50] 50) where
- "wfExt T B G1 G2 = (wfG T B G2 \<and> wfG T B G1 \<and> toSet G1 \<subseteq> toSet G2)"
-
-section \<open>Context\<close>
-
-lemma wfG_cons[ms_wb]:
- fixes \<Gamma>::\<Gamma>
- assumes "P; \<B> \<turnstile>\<^sub>w\<^sub>f (z,b,c) #\<^sub>\<Gamma>\<Gamma>"
- shows "P; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<and> atom z \<sharp> \<Gamma> \<and> wfB P \<B> b"
- using wfG_elims(2)[OF assms] by metis
-
-lemma wfG_cons2[ms_wb]:
- fixes \<Gamma>::\<Gamma>
- assumes "P; \<B> \<turnstile>\<^sub>w\<^sub>f zbc #\<^sub>\<Gamma>\<Gamma>"
- shows "P; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>"
-proof -
- obtain z and b and c where zbc: "zbc=(z,b,c)" using prod_cases3 by blast
- hence "P; \<B> \<turnstile>\<^sub>w\<^sub>f (z,b,c) #\<^sub>\<Gamma>\<Gamma>" using assms by auto
- thus ?thesis using zbc wfG_cons assms by simp
-qed
-
-lemma wf_g_unique:
- fixes \<Gamma>::\<Gamma>
- assumes "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>" and "(x,b,c) \<in> toSet \<Gamma>" and "(x,b',c') \<in> toSet \<Gamma>"
- shows "b=b' \<and> c=c'"
-using assms proof(induct \<Gamma> rule: \<Gamma>.induct)
- case GNil
- then show ?case by simp
-next
- case (GCons a \<Gamma>)
- consider "(x,b,c)=a \<and> (x,b',c')=a" | "(x,b,c)=a \<and> (x,b',c')\<noteq>a" | "(x,b,c)\<noteq>a \<and> (x,b',c')=a" | "(x,b,c)\<noteq> a \<and> (x,b',c')\<noteq>a" by blast
- then show ?case proof(cases)
- case 1
- then show ?thesis by auto
- next
- case 2
- hence "atom x \<sharp> \<Gamma>" using wfG_elims(2) GCons by blast
- moreover have "(x,b',c') \<in> toSet \<Gamma>" using GCons 2 by force
- ultimately show ?thesis using forget_subst_gv fresh_GCons fresh_GNil fresh_gamma_elem \<Gamma>.distinct subst_gv.simps 2 GCons by metis
- next
- case 3
- hence "atom x \<sharp> \<Gamma>" using wfG_elims(2) GCons by blast
- moreover have "(x,b,c) \<in> toSet \<Gamma>" using GCons 3 by force
- ultimately show ?thesis
- using forget_subst_gv fresh_GCons fresh_GNil fresh_gamma_elem \<Gamma>.distinct subst_gv.simps 3 GCons by metis
- next
- case 4
- then obtain x'' and b'' and c''::c where xbc: "a=(x'',b'',c'')"
- using prod_cases3 by blast
- hence "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f ((x'',b'',c'') #\<^sub>\<Gamma>\<Gamma>)" using GCons wfG_elims by blast
- hence "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<and> (x, b, c) \<in> toSet \<Gamma> \<and> (x, b', c') \<in> toSet \<Gamma>" using GCons wfG_elims 4 xbc
- prod_cases3 set_GConsD using forget_subst_gv fresh_GCons fresh_GNil fresh_gamma_elem \<Gamma>.distinct subst_gv.simps 4 GCons by meson
- thus ?thesis using GCons by auto
- qed
-qed
-
-lemma lookup_if1:
- fixes \<Gamma>::\<Gamma>
- assumes "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>" and "Some (b,c) = lookup \<Gamma> x"
- shows "(x,b,c) \<in> toSet \<Gamma> \<and> (\<forall>b' c'. (x,b',c') \<in> toSet \<Gamma> \<longrightarrow> b'=b \<and> c'=c)"
-using assms proof(induct \<Gamma> rule: \<Gamma>.induct)
- case GNil
- then show ?case by auto
-next
- case (GCons xbc \<Gamma>)
- then obtain x' and b' and c'::c where xbc: "xbc=(x',b',c')"
- using prod_cases3 by blast
- then show ?case using wf_g_unique GCons lookup_in_g xbc
- lookup.simps set_GConsD wfG.cases
- insertE insert_is_Un toSet.simps wfG_elims by metis
-qed
-
-lemma lookup_if2:
- assumes "wfG P \<B> \<Gamma>" and "(x,b,c) \<in> toSet \<Gamma> \<and> (\<forall>b' c'. (x,b',c') \<in> toSet \<Gamma> \<longrightarrow> b'=b \<and> c'=c)"
- shows "Some (b,c) = lookup \<Gamma> x"
-using assms proof(induct \<Gamma> rule: \<Gamma>.induct)
- case GNil
- then show ?case by auto
-next
- case (GCons xbc \<Gamma>)
- then obtain x' and b' and c'::c where xbc: "xbc=(x',b',c')"
- using prod_cases3 by blast
- then show ?case proof(cases "x=x'")
- case True
- then show ?thesis using lookup.simps GCons xbc by simp
- next
- case False
- then show ?thesis using lookup.simps GCons xbc toSet.simps Un_iff set_GConsD wfG_cons2
- by (metis (full_types) Un_iff set_GConsD toSet.simps(2) wfG_cons2)
- qed
-qed
-
-lemma lookup_iff:
- fixes \<Theta>::\<Theta> and \<Gamma>::\<Gamma>
- assumes "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>"
- shows "Some (b,c) = lookup \<Gamma> x \<longleftrightarrow> (x,b,c) \<in> toSet \<Gamma> \<and> (\<forall>b' c'. (x,b',c') \<in> toSet \<Gamma> \<longrightarrow> b'=b \<and> c'=c)"
- using assms lookup_if1 lookup_if2 by meson
-
-lemma wfG_lookup_wf:
- fixes \<Theta>::\<Theta> and \<Gamma>::\<Gamma> and b::b and \<B>::\<B>
- assumes "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>" and "Some (b,c) = lookup \<Gamma> x"
- shows "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b"
-using assms proof(induct \<Gamma> rule: \<Gamma>_induct)
- case GNil
- then show ?case by auto
-next
- case (GCons x' b' c' \<Gamma>')
- then show ?case proof(cases "x=x'")
- case True
- then show ?thesis using lookup.simps wfG_elims(2) GCons by fastforce
- next
- case False
- then show ?thesis using lookup.simps wfG_elims(2) GCons by fastforce
- qed
-qed
-
-lemma wfG_unique:
- fixes \<Gamma>::\<Gamma>
- assumes "wfG B \<Theta> ((x, b, c) #\<^sub>\<Gamma> \<Gamma>)" and "(x1,b1,c1) \<in> toSet ((x, b, c) #\<^sub>\<Gamma> \<Gamma>)" and "x1=x"
- shows "b1 = b \<and> c1 = c"
-proof -
- have "(x, b, c) \<in> toSet ((x, b, c) #\<^sub>\<Gamma> \<Gamma>)" by simp
- thus ?thesis using wf_g_unique assms by blast
-qed
-
-lemma wfG_unique_full:
- fixes \<Gamma>::\<Gamma>
- assumes "wfG \<Theta> B (\<Gamma>'@(x, b, c) #\<^sub>\<Gamma> \<Gamma>)" and "(x1,b1,c1) \<in> toSet (\<Gamma>'@(x, b, c) #\<^sub>\<Gamma> \<Gamma>)" and "x1=x"
- shows "b1 = b \<and> c1 = c"
-proof -
- have "(x, b, c) \<in> toSet (\<Gamma>'@(x, b, c) #\<^sub>\<Gamma> \<Gamma>)" by simp
- thus ?thesis using wf_g_unique assms by blast
-qed
-
-section \<open>Converting between wb forms\<close>
-
-text \<open> We cannot prove wfB properties here for expressions and statements as need some more facts about @{term \<Phi>}
- context which we can prove without this lemma. Trying to cram everything into a single large
- mutually recursive lemma is not a good idea \<close>
-
-lemma wfX_wfY1:
- fixes \<Gamma>::\<Gamma> and \<Gamma>'::\<Gamma> and v::v and e::e and c::c and \<tau>::\<tau> and ts::"(string*\<tau>) list" and \<Delta>::\<Delta> and s::s and b::b and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def and cs::branch_s
- and css::branch_list
- shows wfV_wf: "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<and> \<turnstile>\<^sub>w\<^sub>f \<Theta> " and
- wfC_wf: "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f c \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<and> \<turnstile>\<^sub>w\<^sub>f \<Theta> " and
- wfG_wf :"\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<Longrightarrow> \<turnstile>\<^sub>w\<^sub>f \<Theta>" and
- wfT_wf: "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau> \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<and> \<turnstile>\<^sub>w\<^sub>f \<Theta> \<and> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b_of \<tau>" and
- wfTs_wf:"\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ts \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<and> \<turnstile>\<^sub>w\<^sub>f \<Theta>" and
- "\<turnstile>\<^sub>w\<^sub>f \<Theta> \<Longrightarrow> True" and
- wfB_wf: "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b \<Longrightarrow> \<turnstile>\<^sub>w\<^sub>f \<Theta>" and
- wfCE_wf: "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ce : b \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<and> \<turnstile>\<^sub>w\<^sub>f \<Theta> " and
- wfTD_wf: "\<Theta> \<turnstile>\<^sub>w\<^sub>f td \<Longrightarrow> \<turnstile>\<^sub>w\<^sub>f \<Theta>"
-proof(induct rule:wfV_wfC_wfG_wfT_wfTs_wfTh_wfB_wfCE_wfTD.inducts)
-
- case (wfV_varI \<Theta> \<B> \<Gamma> b c x)
- hence "(x,b,c) \<in> toSet \<Gamma>" using lookup_iff lookup_in_g by presburger
- hence "b \<in> fst`snd`toSet \<Gamma>" by force
- hence "wfB \<Theta> \<B> b" using wfV_varI using wfG_lookup_wf by auto
- then show ?case using wfV_varI wfV_elims wf_intros by metis
-next
- case (wfV_litI \<Theta> \<B> \<Gamma> l)
- moreover have "wfTh \<Theta>" using wfV_litI by metis
- ultimately show ?case using wf_intros base_for_lit.simps l.exhaust by metis
-next
- case (wfV_pairI \<Theta> \<B> \<Gamma> v1 b1 v2 b2)
- then show ?case using wfB_pairI by simp
-next
- case (wfV_consI s dclist \<Theta> dc x b c \<B> \<Gamma> v)
- then show ?case using wf_intros by metis
-next
- case (wfTI z \<Gamma> \<Theta> \<B> b c)
- then show ?case using wf_intros b_of.simps wfG_cons2 by metis
-qed(auto)
-
-lemma wfX_wfY2:
- fixes \<Gamma>::\<Gamma> and \<Gamma>'::\<Gamma> and v::v and e::e and c::c and \<tau>::\<tau> and ts::"(string*\<tau>) list" and \<Delta>::\<Delta> and s::s and b::b and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def and cs::branch_s
- and css::branch_list
- shows
- wfE_wf: "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f e : b \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<and> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<and> \<turnstile>\<^sub>w\<^sub>f \<Theta> \<and> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> " and
- wfS_wf: "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<and> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<and> \<turnstile>\<^sub>w\<^sub>f \<Theta> \<and> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> " and
- "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> ; tid ; dc ; t \<turnstile>\<^sub>w\<^sub>f cs : b \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<and> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<and> \<turnstile>\<^sub>w\<^sub>f \<Theta> \<and> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> " and
- "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> ; tid ; dclist \<turnstile>\<^sub>w\<^sub>f css : b \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<and> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<and> \<turnstile>\<^sub>w\<^sub>f \<Theta> \<and> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> " and
- wfPhi_wf: "\<Theta> \<turnstile>\<^sub>w\<^sub>f (\<Phi>::\<Phi>) \<Longrightarrow> \<turnstile>\<^sub>w\<^sub>f \<Theta>" and
- wfD_wf: "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<and> \<turnstile>\<^sub>w\<^sub>f \<Theta> " and
- wfFTQ_wf: "\<Theta> ; \<Phi> \<turnstile>\<^sub>w\<^sub>f ftq \<Longrightarrow> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<and> \<turnstile>\<^sub>w\<^sub>f \<Theta>" and
- wfFT_wf: "\<Theta> ; \<Phi> ; \<B> \<turnstile>\<^sub>w\<^sub>f ft \<Longrightarrow> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<and> \<turnstile>\<^sub>w\<^sub>f \<Theta>"
-proof(induct rule:wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.inducts)
- case (wfS_varI \<Theta> \<B> \<Gamma> \<tau> v u \<Delta> \<Phi> s b)
- then show ?case using wfD_elims by auto
-next
- case (wfS_assignI u \<tau> \<Delta> \<Theta> \<B> \<Gamma> \<Phi> v)
- then show ?case using wf_intros by metis
-next
- case (wfD_emptyI \<Theta> \<B> \<Gamma>)
- then show ?case using wfX_wfY1 by auto
-next
- case (wfS_assertI \<Theta> \<Phi> \<B> x c \<Gamma> \<Delta> s b)
- then have "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>" using wfX_wfY1 by auto
- moreover have "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>" using wfS_assertI by auto
- moreover have "\<turnstile>\<^sub>w\<^sub>f \<Theta> \<and> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> " using wfS_assertI by auto
- ultimately show ?case by auto
-qed(auto)
-
-lemmas wfX_wfY=wfX_wfY1 wfX_wfY2
-
-lemma setD_ConsD:
- "ut \<in> setD (ut' #\<^sub>\<Delta> D) = (ut = ut' \<or> ut \<in> setD D)"
-proof(induct D rule: \<Delta>_induct)
- case DNil
- then show ?case by auto
-next
- case (DCons u' t' x2)
- then show ?case using setD.simps by auto
-qed
-
-lemma wfD_wfT:
- fixes \<Delta>::\<Delta> and \<tau>::\<tau>
- assumes "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>"
- shows "\<forall>(u,\<tau>) \<in> setD \<Delta>. \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau>"
-using assms proof(induct \<Delta> rule: \<Delta>_induct)
- case DNil
- then show ?case by auto
-next
- case (DCons u' t' x2)
- then show ?case using wfD_elims DCons setD_ConsD
- by (metis case_prodI2 set_ConsD)
-qed
-
-lemma subst_b_lookup_d:
- assumes "u \<notin> fst ` setD \<Delta>"
- shows "u \<notin> fst ` setD \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b"
-using assms proof(induct \<Delta> rule: \<Delta>_induct)
- case DNil
- then show ?case by auto
-next
- case (DCons u' t' x2)
- hence "u\<noteq>u'" using DCons by simp
- show ?case using DCons subst_db.simps by simp
-qed
-
-lemma wfG_cons_splitI:
- fixes \<Phi>::\<Phi> and \<Gamma>::\<Gamma>
- assumes "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>" and "atom x \<sharp> \<Gamma>" and "wfB \<Theta> \<B> b" and
- "c \<in> { TRUE, FALSE } \<longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> " and
- "c \<notin> { TRUE, FALSE } \<longrightarrow> \<Theta> ;\<B> ; (x,b,C_true) #\<^sub>\<Gamma>\<Gamma> \<turnstile>\<^sub>w\<^sub>f c"
- shows "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f ((x,b,c) #\<^sub>\<Gamma>\<Gamma>)"
- using wfG_cons1I wfG_cons2I assms by metis
-
-lemma wfG_consI:
- fixes \<Phi>::\<Phi> and \<Gamma>::\<Gamma> and c::c
- assumes "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>" and "atom x \<sharp> \<Gamma>" and "wfB \<Theta> \<B> b" and
- "\<Theta> ; \<B> ; (x,b,C_true) #\<^sub>\<Gamma>\<Gamma> \<turnstile>\<^sub>w\<^sub>f c"
- shows "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f ((x,b,c) #\<^sub>\<Gamma>\<Gamma>)"
- using wfG_cons1I wfG_cons2I wfG_cons_splitI wfC_trueI assms by metis
-
-lemma wfG_elim2:
- fixes c::c
- assumes "wfG P \<B> ((x,b,c) #\<^sub>\<Gamma>\<Gamma>)"
- shows "P; \<B> ; (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f c \<and> wfB P \<B> b"
-proof(cases "c \<in> {TRUE,FALSE}")
- case True
- have "P; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<and> atom x \<sharp> \<Gamma> \<and> wfB P \<B> b" using wfG_elims(2)[OF assms] by auto
- hence "P; \<B> \<turnstile>\<^sub>w\<^sub>f ((x,b,TRUE) #\<^sub>\<Gamma>\<Gamma>) \<and> wfB P \<B> b" using wfG_cons2I by auto
- thus ?thesis using wfC_trueI wfC_falseI True by auto
-next
- case False
- then show ?thesis using wfG_elims(2)[OF assms] by auto
-qed
-
-lemma wfG_cons_wfC:
- fixes \<Gamma>::\<Gamma> and c::c
- assumes "\<Theta> ; B \<turnstile>\<^sub>w\<^sub>f (x, b, c) #\<^sub>\<Gamma> \<Gamma>"
- shows "\<Theta> ; B ; ((x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>) \<turnstile>\<^sub>w\<^sub>f c"
- using assms wfG_elim2 by auto
-
-lemma wfG_wfB:
- assumes "wfG P \<B> \<Gamma>" and "b \<in> fst`snd`toSet \<Gamma>"
- shows "wfB P \<B> b"
-using assms proof(induct \<Gamma> rule:\<Gamma>_induct)
-case GNil
- then show ?case by auto
-next
- case (GCons x' b' c' \<Gamma>')
- show ?case proof(cases "b=b'")
- case True
- then show ?thesis using wfG_elim2 GCons by auto
- next
- case False
- hence "b \<in> fst`snd`toSet \<Gamma>'" using GCons by auto
- moreover have "wfG P \<B> \<Gamma>'" using wfG_cons GCons by auto
- ultimately show ?thesis using GCons by auto
- qed
-qed
-
-lemma wfG_cons_TRUE:
- fixes \<Gamma>::\<Gamma> and b::b
- assumes "P; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>" and "atom z \<sharp> \<Gamma>" and "P; \<B> \<turnstile>\<^sub>w\<^sub>f b"
- shows "P ; \<B> \<turnstile>\<^sub>w\<^sub>f (z, b, TRUE) #\<^sub>\<Gamma> \<Gamma>"
- using wfG_cons2I wfG_wfB assms by simp
-
-lemma wfG_cons_TRUE2:
- assumes "P; \<B> \<turnstile>\<^sub>w\<^sub>f (z,b,c) #\<^sub>\<Gamma>\<Gamma>" and "atom z \<sharp> \<Gamma>"
- shows "P; \<B> \<turnstile>\<^sub>w\<^sub>f (z, b, TRUE) #\<^sub>\<Gamma> \<Gamma>"
- using wfG_cons wfG_cons2I assms by simp
-
-lemma wfG_suffix:
- fixes \<Gamma>::\<Gamma>
- assumes "wfG P \<B> (\<Gamma>'@\<Gamma>)"
- shows "wfG P \<B> \<Gamma>"
-using assms proof(induct \<Gamma>' rule: \<Gamma>_induct)
- case GNil
- then show ?case by auto
-next
- case (GCons x b c \<Gamma>')
- hence " P; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>' @ \<Gamma>" using wfG_elims by auto
- then show ?case using GCons wfG_elims by auto
-qed
-
-lemma wfV_wfCE:
- fixes v::v
- assumes "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b"
- shows " \<Theta> ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_val v : b"
-proof -
- have "\<Theta> \<turnstile>\<^sub>w\<^sub>f ([]::\<Phi>) " using wfPhi_emptyI wfV_wf wfG_wf assms by metis
- moreover have "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f []\<^sub>\<Delta>" using wfD_emptyI wfV_wf wfG_wf assms by metis
- ultimately show ?thesis using wfCE_valI assms by auto
-qed
-
-section \<open>Support\<close>
-
-lemma wf_supp1:
- fixes \<Gamma>::\<Gamma> and \<Gamma>'::\<Gamma> and v::v and e::e and c::c and \<tau>::\<tau> and ts::"(string*\<tau>) list" and \<Delta>::\<Delta> and s::s and b::b and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def and cs::branch_s and css ::branch_list
-
- shows wfV_supp: "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b \<Longrightarrow> supp v \<subseteq> atom_dom \<Gamma> \<union> supp \<B>" and
- wfC_supp: "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f c \<Longrightarrow> supp c \<subseteq> atom_dom \<Gamma> \<union> supp \<B>" and
- wfG_supp: "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<Longrightarrow> atom_dom \<Gamma> \<subseteq> supp \<Gamma>" and
- wfT_supp: "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau> \<Longrightarrow> supp \<tau> \<subseteq> atom_dom \<Gamma> \<union> supp \<B> " and
- wfTs_supp: "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ts \<Longrightarrow> supp ts \<subseteq> atom_dom \<Gamma> \<union> supp \<B>" and
- wfTh_supp: "\<turnstile>\<^sub>w\<^sub>f \<Theta> \<Longrightarrow> supp \<Theta> = {}" and
- wfB_supp: "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b \<Longrightarrow> supp b \<subseteq> supp \<B>" and
- wfCE_supp: "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ce : b \<Longrightarrow> supp ce \<subseteq> atom_dom \<Gamma> \<union> supp \<B>" and
- wfTD_supp: "\<Theta> \<turnstile>\<^sub>w\<^sub>f td \<Longrightarrow> supp td \<subseteq> {}"
-proof(induct rule:wfV_wfC_wfG_wfT_wfTs_wfTh_wfB_wfCE_wfTD.inducts)
- case (wfB_consI \<Theta> s dclist \<B>)
- then show ?case by(auto simp add: b.supp pure_supp)
-next
- case (wfB_appI \<Theta> \<B> b s bv dclist)
- then show ?case by(auto simp add: b.supp pure_supp)
-next
- case (wfV_varI \<Theta> \<B> \<Gamma> b c x)
- then show ?case using v.supp wfV_elims
- empty_subsetI insert_subset supp_at_base
- fresh_dom_free2 lookup_if1
- by (metis sup.coboundedI1)
-next
- case (wfV_litI \<Theta> \<B> \<Gamma> l)
- then show ?case using supp_l_empty v.supp by simp
-next
- case (wfV_pairI \<Theta> \<B> \<Gamma> v1 b1 v2 b2)
- then show ?case using v.supp wfV_elims by (metis Un_subset_iff)
-next
- case (wfV_consI s dclist \<Theta> dc x b c \<B> \<Gamma> v)
- then show ?case using v.supp wfV_elims
- Un_commute b.supp sup_bot.right_neutral supp_b_empty pure_supp by metis
-next
- case (wfV_conspI typid bv dclist \<Theta> dc x b' c \<B> \<Gamma> v b)
- then show ?case unfolding v.supp
- using wfV_elims
- Un_commute b.supp sup_bot.right_neutral supp_b_empty pure_supp
- by (simp add: Un_commute pure_supp sup.coboundedI1)
-next
- case (wfC_eqI \<Theta> \<B> \<Gamma> e1 b e2)
- hence "supp e1 \<subseteq> atom_dom \<Gamma> \<union> supp \<B>" using c.supp wfC_elims
- image_empty list.set(1) sup_bot.right_neutral by (metis IntI UnE empty_iff subsetCE subsetI)
- moreover have "supp e2 \<subseteq> atom_dom \<Gamma> \<union> supp \<B>" using c.supp wfC_elims
- image_empty list.set(1) sup_bot.right_neutral IntI UnE empty_iff subsetCE subsetI
- by (metis wfC_eqI.hyps(4))
- ultimately show ?case using c.supp by auto
-next
- case (wfG_cons1I c \<Theta> \<B> \<Gamma> x b)
- then show ?case using atom_dom.simps dom_supp_g supp_GCons by metis
-next
- case (wfG_cons2I c \<Theta> \<B> \<Gamma> x b)
- then show ?case using atom_dom.simps dom_supp_g supp_GCons by metis
-next
- case wfTh_emptyI
- then show ?case by (simp add: supp_Nil)
-next
- case (wfTh_consI \<Theta> lst)
- then show ?case using supp_Cons by fast
-next
- case (wfTD_simpleI \<Theta> lst s)
- then have "supp (AF_typedef s lst ) = supp lst \<union> supp s" using type_def.supp by auto
- then show ?case using wfTD_simpleI pure_supp
- by (simp add: pure_supp supp_Cons supp_at_base)
-next
- case (wfTD_poly \<Theta> bv lst s)
- then have "supp (AF_typedef_poly s bv lst ) = supp lst - { atom bv } \<union> supp s" using type_def.supp by auto
- then show ?case using wfTD_poly pure_supp
- by (simp add: pure_supp supp_Cons supp_at_base)
-next
- case (wfTs_nil \<Theta> \<B> \<Gamma>)
- then show ?case using supp_Nil by auto
-next
- case (wfTs_cons \<Theta> \<B> \<Gamma> \<tau> dc ts)
- then show ?case using supp_Cons supp_Pair pure_supp[of dc] by blast
-next
- case (wfCE_valI \<Theta> \<B> \<Gamma> v b)
- thus ?case using ce.supp wfCE_elims by simp
-next
- case (wfCE_plusI \<Theta> \<B> \<Gamma> v1 v2)
- hence "supp (CE_op Plus v1 v2) \<subseteq> atom_dom \<Gamma> \<union> supp \<B>" using ce.supp pure_supp
- by (simp add: wfCE_plusI opp.supp)
- then show ?case using ce.supp wfCE_elims UnCI subsetCE subsetI x_not_in_b_set by auto
-next
- case (wfCE_leqI \<Theta> \<B> \<Gamma> v1 v2)
- hence "supp (CE_op LEq v1 v2) \<subseteq> atom_dom \<Gamma> \<union> supp \<B>" using ce.supp pure_supp
- by (simp add: wfCE_plusI opp.supp)
- then show ?case using ce.supp wfE_elims UnCI subsetCE subsetI x_not_in_b_set by auto
-next
- case (wfCE_eqI \<Theta> \<B> \<Gamma> v1 b v2 )
- hence "supp (CE_op Eq v1 v2) \<subseteq> atom_dom \<Gamma> \<union> supp \<B>" using ce.supp pure_supp
- by (simp add: wfCE_eqI opp.supp)
- then show ?case using ce.supp wfE_elims UnCI subsetCE subsetI x_not_in_b_set by auto
-next
- case (wfCE_fstI \<Theta> \<B> \<Gamma> v1 b1 b2)
- thus ?case using ce.supp wfCE_elims by simp
-next
- case (wfCE_sndI \<Theta> \<B> \<Gamma> v1 b1 b2)
- thus ?case using ce.supp wfCE_elims by simp
-next
- case (wfCE_concatI \<Theta> \<B> \<Gamma> v1 v2)
- thus ?case using ce.supp wfCE_elims by simp
-next
- case (wfCE_lenI \<Theta> \<B> \<Gamma> v1)
- thus ?case using ce.supp wfCE_elims by simp
-next
- case (wfTI z \<Theta> \<B> \<Gamma> b c)
- hence "supp c \<subseteq> supp z \<union> atom_dom \<Gamma> \<union> supp \<B>" using supp_at_base dom_cons by metis
- moreover have "supp b \<subseteq> supp \<B>" using wfTI by auto
- ultimately have " supp \<lbrace> z : b | c \<rbrace> \<subseteq> atom_dom \<Gamma> \<union> supp \<B>" using \<tau>.supp supp_at_base by force
- thus ?case by auto
-qed(auto)
-
-lemma wf_supp2:
- fixes \<Gamma>::\<Gamma> and \<Gamma>'::\<Gamma> and v::v and e::e and c::c and \<tau>::\<tau> and
- ts::"(string*\<tau>) list" and \<Delta>::\<Delta> and s::s and b::b and ftq::fun_typ_q and
- ft::fun_typ and ce::ce and td::type_def and cs::branch_s and css ::branch_list
- shows
- wfE_supp: "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f e : b \<Longrightarrow> (supp e \<subseteq> atom_dom \<Gamma> \<union> supp \<B> \<union> atom ` fst ` setD \<Delta>)" and (* \<and> ( \<Phi> = [] \<longrightarrow> supp e \<inter> supp \<B> = {})" and*)
- wfS_supp: "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b \<Longrightarrow> supp s \<subseteq> atom_dom \<Gamma> \<union> atom ` fst ` setD \<Delta> \<union> supp \<B>" and
- "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> ; tid ; dc ; t \<turnstile>\<^sub>w\<^sub>f cs : b \<Longrightarrow> supp cs \<subseteq> atom_dom \<Gamma> \<union> atom ` fst ` setD \<Delta> \<union> supp \<B>" and
- "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> ; tid ; dclist \<turnstile>\<^sub>w\<^sub>f css : b \<Longrightarrow> supp css \<subseteq> atom_dom \<Gamma> \<union> atom ` fst ` setD \<Delta> \<union> supp \<B>" and
- wfPhi_supp: "\<Theta> \<turnstile>\<^sub>w\<^sub>f (\<Phi>::\<Phi>) \<Longrightarrow> supp \<Phi> = {}" and
- wfD_supp: "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<Longrightarrow> supp \<Delta> \<subseteq> atom`fst`(setD \<Delta>) \<union> atom_dom \<Gamma> \<union> supp \<B> " and
- "\<Theta> ; \<Phi> \<turnstile>\<^sub>w\<^sub>f ftq \<Longrightarrow> supp ftq = {}" and
- "\<Theta> ; \<Phi> ; \<B> \<turnstile>\<^sub>w\<^sub>f ft \<Longrightarrow> supp ft \<subseteq> supp \<B>"
-proof(induct rule:wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.inducts)
- case (wfE_valI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v b)
- hence "supp (AE_val v) \<subseteq> atom_dom \<Gamma> \<union> supp \<B>" using e.supp wf_supp1 by simp
- then show ?case using e.supp wfE_elims UnCI subsetCE subsetI x_not_in_b_set by metis
-next
- case (wfE_plusI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 v2)
- hence "supp (AE_op Plus v1 v2) \<subseteq> atom_dom \<Gamma> \<union> supp \<B>"
- using wfE_plusI opp.supp wf_supp1 e.supp pure_supp Un_least
- by (metis sup_bot.left_neutral)
-
- then show ?case using e.supp wfE_elims UnCI subsetCE subsetI x_not_in_b_set by metis
-next
- case (wfE_leqI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 v2)
- hence "supp (AE_op LEq v1 v2) \<subseteq> atom_dom \<Gamma> \<union> supp \<B>" using e.supp pure_supp Un_least
- sup_bot.left_neutral using opp.supp wf_supp1 by auto
- then show ?case using e.supp wfE_elims UnCI subsetCE subsetI x_not_in_b_set by metis
-next
- case (wfE_eqI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 b v2)
- hence "supp (AE_op Eq v1 v2) \<subseteq> atom_dom \<Gamma> \<union> supp \<B>" using e.supp pure_supp Un_least
- sup_bot.left_neutral using opp.supp wf_supp1 by auto
- then show ?case using e.supp wfE_elims UnCI subsetCE subsetI x_not_in_b_set by metis
-next
- case (wfE_fstI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 b1 b2)
- hence "supp (AE_fst v1 ) \<subseteq> atom_dom \<Gamma> \<union> supp \<B>" using e.supp pure_supp sup_bot.left_neutral using opp.supp wf_supp1 by auto
- then show ?case using e.supp wfE_elims UnCI subsetCE subsetI x_not_in_b_set by metis
-next
- case (wfE_sndI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 b1 b2)
- hence "supp (AE_snd v1 ) \<subseteq> atom_dom \<Gamma> \<union> supp \<B>" using e.supp pure_supp wfE_plusI opp.supp wf_supp1 by (metis Un_least)
- then show ?case using e.supp wfE_elims UnCI subsetCE subsetI x_not_in_b_set by metis
-next
- case (wfE_concatI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 v2)
- hence "supp (AE_concat v1 v2) \<subseteq> atom_dom \<Gamma> \<union> supp \<B>" using e.supp pure_supp
- wfE_plusI opp.supp wf_supp1 by (metis Un_least)
- then show ?case using e.supp wfE_elims UnCI subsetCE subsetI x_not_in_b_set by metis
-next
- case (wfE_splitI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 v2)
- hence "supp (AE_split v1 v2) \<subseteq> atom_dom \<Gamma> \<union> supp \<B>" using e.supp pure_supp
- wfE_plusI opp.supp wf_supp1 by (metis Un_least)
- then show ?case using e.supp wfE_elims UnCI subsetCE subsetI x_not_in_b_set by metis
-next
- case (wfE_lenI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1)
- hence "supp (AE_len v1 ) \<subseteq> atom_dom \<Gamma> \<union> supp \<B>" using e.supp pure_supp
- using e.supp pure_supp sup_bot.left_neutral using opp.supp wf_supp1 by auto
- then show ?case using e.supp wfE_elims UnCI subsetCE subsetI x_not_in_b_set by metis
-next
- case (wfE_appI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> f x b c \<tau> s v)
- then obtain b where "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b" using wfE_elims by metis
- hence "supp v \<subseteq> atom_dom \<Gamma> \<union> supp \<B>" using wfE_appI wf_supp1 by metis
- hence "supp (AE_app f v) \<subseteq> atom_dom \<Gamma> \<union> supp \<B>" using e.supp pure_supp by fast
- then show ?case using e.supp(2) UnCI subsetCE subsetI wfE_appI using b.supp(3) pure_supp x_not_in_b_set by metis
-next
- case (wfE_appPI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> b' bv v \<tau> f xa ba ca s)
- then obtain b where "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : ( b[bv::=b']\<^sub>b)" using wfE_elims by metis
- hence "supp v \<subseteq> atom_dom \<Gamma> \<union> supp \<B> " using wfE_appPI wf_supp1 by auto
- moreover have "supp b' \<subseteq> supp \<B>" using wf_supp1(7) wfE_appPI by simp
- ultimately show ?case unfolding e.supp using wfE_appPI pure_supp by fast
-next
- case (wfE_mvarI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> u \<tau>)
- then obtain \<tau> where "(u,\<tau>) \<in> setD \<Delta>" using wfE_elims(10) by metis
- hence "atom u \<in> atom`fst`setD \<Delta>" by force
- hence "supp (AE_mvar u ) \<subseteq> atom`fst`setD \<Delta>" using e.supp
- by (simp add: supp_at_base)
- thus ?case using UnCI subsetCE subsetI e.supp wfE_mvarI supp_at_base subsetCE supp_at_base u_not_in_b_set
- by (simp add: supp_at_base)
-next
- case (wfS_valI \<Theta> \<Phi> \<B> \<Gamma> v b \<Delta>)
- then show ?case using wf_supp1
- by (metis s_branch_s_branch_list.supp(1) sup.coboundedI2 sup_assoc sup_commute)
-next
- case (wfS_letI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> e b' x s b)
- then show ?case by auto
-next
- case (wfS_let2I \<Theta> \<Phi> \<B> \<Gamma> \<Delta> s1 \<tau> x s2 b)
- then show ?case unfolding s_branch_s_branch_list.supp (3) using wf_supp1(4)[OF wfS_let2I(3)] by auto
-next
- case (wfS_ifI \<Theta> \<B> \<Gamma> v \<Phi> \<Delta> s1 b s2)
- then show ?case using wf_supp1(1)[OF wfS_ifI(1)] by auto
-next
- case (wfS_varI \<Theta> \<B> \<Gamma> \<tau> v u \<Delta> \<Phi> s b)
- then show ?case using wf_supp1(1)[OF wfS_varI(2)] wf_supp1(4)[OF wfS_varI(1)] by auto
-next
-next
- case (wfS_assignI u \<tau> \<Delta> \<Theta> \<B> \<Gamma> \<Phi> v)
- hence "supp u \<subseteq> atom ` fst ` setD \<Delta>" proof(induct \<Delta> rule:\<Delta>_induct)
- case DNil
- then show ?case by auto
- next
- case (DCons u' t' \<Delta>')
- show ?case proof(cases "u=u'")
- case True
- then show ?thesis using toSet.simps DCons supp_at_base by fastforce
- next
- case False
- then show ?thesis using toSet.simps DCons supp_at_base wfS_assignI
- by (metis empty_subsetI fstI image_eqI insert_subset)
- qed
- qed
- then show ?case using s_branch_s_branch_list.supp(8) wfS_assignI wf_supp1(1)[OF wfS_assignI(6)] by auto
-next
- case (wfS_matchI \<Theta> \<B> \<Gamma> v tid dclist \<Delta> \<Phi> cs b)
- then show ?case using wf_supp1(1)[OF wfS_matchI(1)] by auto
-next
- case (wfS_branchI \<Theta> \<Phi> \<B> x \<tau> \<Gamma> \<Delta> s b tid dc)
- moreover have "supp s \<subseteq> supp x \<union> atom_dom \<Gamma> \<union> atom ` fst ` setD \<Delta> \<union> supp \<B>"
- using dom_cons supp_at_base wfS_branchI by auto
- moreover hence "supp s - set [atom x] \<subseteq> atom_dom \<Gamma> \<union> atom ` fst ` setD \<Delta> \<union> supp \<B>" using supp_at_base by force
- ultimately have
- "(supp s - set [atom x]) \<union> (supp dc ) \<subseteq> atom_dom \<Gamma> \<union> atom ` fst ` setD \<Delta> \<union> supp \<B>"
- by (simp add: pure_supp)
- thus ?case using s_branch_s_branch_list.supp(2) by auto
-next
- case (wfD_emptyI \<Theta> \<B> \<Gamma>)
- then show ?case using supp_DNil by auto
-next
- case (wfD_cons \<Theta> \<B> \<Gamma> \<Delta> \<tau> u)
- have "supp ((u, \<tau>) #\<^sub>\<Delta> \<Delta>) = supp u \<union> supp \<tau> \<union> supp \<Delta>" using supp_DCons supp_Pair by metis
- also have "... \<subseteq> supp u \<union> atom ` fst ` setD \<Delta> \<union> atom_dom \<Gamma> \<union> supp \<B>"
- using wfD_cons wf_supp1(4)[OF wfD_cons(3)] by auto
- also have "... \<subseteq> atom ` fst ` setD ((u, \<tau>) #\<^sub>\<Delta> \<Delta>) \<union> atom_dom \<Gamma> \<union> supp \<B>" using supp_at_base by auto
- finally show ?case by auto
-next
- case (wfPhi_emptyI \<Theta>)
- then show ?case using supp_Nil by auto
-next
- case (wfPhi_consI f \<Theta> \<Phi> ft)
- then show ?case using fun_def.supp
- by (simp add: pure_supp supp_Cons)
-next
- case (wfFTI \<Theta> B' b s x c \<tau> \<Phi>)
- have " supp (AF_fun_typ x b c \<tau> s) = supp c \<union> (supp \<tau> \<union> supp s) - set [atom x] \<union> supp b" using fun_typ.supp by auto
- thus ?case using wfFTI wf_supp1
- proof -
- have f1: "supp \<tau> \<subseteq> {atom x} \<union> atom_dom GNil \<union> supp B'"
- using dom_cons wfFTI.hyps wf_supp1(4) by blast (* 0.0 ms *)
- have "supp b \<subseteq> supp B'"
- using wfFTI.hyps(1) wf_supp1(7) by blast (* 0.0 ms *)
- then show ?thesis
- using f1 \<open>supp (AF_fun_typ x b c \<tau> s) = supp c \<union> (supp \<tau> \<union> supp s) - set [atom x] \<union> supp b\<close>
- wfFTI.hyps(4) wfFTI.hyps by auto (* 234 ms *)
- qed
-next
- case (wfFTNone \<Theta> \<Phi> ft)
- then show ?case by (simp add: fun_typ_q.supp(2))
-next
- case (wfFTSome \<Theta> \<Phi> bv ft)
- then show ?case using fun_typ_q.supp
- by (simp add: supp_at_base)
-next
- case (wfS_assertI \<Theta> \<Phi> \<B> x c \<Gamma> \<Delta> s b)
- then have "supp c \<subseteq> atom_dom \<Gamma> \<union> atom ` fst ` setD \<Delta> \<union> supp \<B>" using wf_supp1
- by (metis Un_assoc Un_commute le_supI2)
- moreover have "supp s \<subseteq> atom_dom \<Gamma> \<union> atom ` fst ` setD \<Delta> \<union> supp \<B>" proof
- fix z
- assume *:"z \<in> supp s"
- have **:"atom x \<notin> supp s" using wfS_assertI fresh_prodN fresh_def by metis
- have "z \<in> atom_dom ((x, B_bool, c) #\<^sub>\<Gamma> \<Gamma>) \<union> atom ` fst ` setD \<Delta> \<union> supp \<B>" using wfS_assertI * by blast
- have "z \<in> atom_dom ((x, B_bool, c) #\<^sub>\<Gamma> \<Gamma>) \<Longrightarrow> z \<in> atom_dom \<Gamma>" using * ** by auto
- thus "z \<in> atom_dom \<Gamma> \<union> atom ` fst ` setD \<Delta> \<union> supp \<B>" using * **
- using \<open>z \<in> atom_dom ((x, B_bool, c) #\<^sub>\<Gamma> \<Gamma>) \<union> atom ` fst ` setD \<Delta> \<union> supp \<B>\<close> by blast
- qed
- ultimately show ?case by auto
-qed(auto)
-
-lemmas wf_supp = wf_supp1 wf_supp2
-
-lemma wfV_supp_nil:
- fixes v::v
- assumes "P ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f v : b"
- shows "supp v = {}"
- using wfV_supp[of P " {||}" GNil v b] dom.simps toSet.simps
- using assms by auto
-
-lemma wfT_TRUE_aux:
- assumes "wfG P \<B> \<Gamma>" and "atom z \<sharp> (P, \<B>, \<Gamma>)" and "wfB P \<B> b"
- shows "wfT P \<B> \<Gamma> (\<lbrace> z : b | TRUE \<rbrace>)"
-proof (rule)
- show \<open> atom z \<sharp> (P, \<B>, \<Gamma>)\<close> using assms by auto
- show \<open> P; \<B> \<turnstile>\<^sub>w\<^sub>f b \<close> using assms by auto
- show \<open> P ;\<B> ; (z, b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f TRUE \<close> using wfG_cons2I wfC_trueI assms by auto
-qed
-
-lemma wfT_TRUE:
- assumes "wfG P \<B> \<Gamma>" and "wfB P \<B> b"
- shows "wfT P \<B> \<Gamma> (\<lbrace> z : b | TRUE \<rbrace>)"
-proof -
- obtain z'::x where *:"atom z' \<sharp> (P, \<B>, \<Gamma>)" using obtain_fresh by metis
- hence "\<lbrace> z : b | TRUE \<rbrace> = \<lbrace> z' : b | TRUE \<rbrace>" by auto
- thus ?thesis using wfT_TRUE_aux assms * by metis
-qed
-
-lemma phi_flip_eq:
- assumes "wfPhi T P"
- shows "(x \<leftrightarrow> xa) \<bullet> P = P"
- using wfPhi_supp[OF assms] flip_fresh_fresh fresh_def by blast
-
-lemma wfC_supp_cons:
- fixes c'::c and G::\<Gamma>
- assumes "P; \<B> ; (x', b' , TRUE) #\<^sub>\<Gamma>G \<turnstile>\<^sub>w\<^sub>f c'"
- shows "supp c' \<subseteq> atom_dom G \<union> supp x' \<union> supp \<B>" and "supp c' \<subseteq> supp G \<union> supp x' \<union> supp \<B>"
-proof -
- show "supp c' \<subseteq> atom_dom G \<union> supp x' \<union> supp \<B>"
- using wfC_supp[OF assms] dom_cons supp_at_base by blast
- moreover have "atom_dom G \<subseteq> supp G"
- by (meson assms wfC_wf wfG_cons wfG_supp)
- ultimately show "supp c' \<subseteq> supp G \<union> supp x' \<union> supp \<B>" using wfG_supp assms wfG_cons wfC_wf by fast
-qed
-
-lemma wfG_dom_supp:
- fixes x::x
- assumes "wfG P \<B> G"
- shows "atom x \<in> atom_dom G \<longleftrightarrow> atom x \<in> supp G"
-using assms proof(induct G rule: \<Gamma>_induct)
- case GNil
- then show ?case using dom.simps supp_of_atom_list
- using supp_GNil by auto
-next
- case (GCons x' b' c' G)
-
- show ?case proof(cases "x' = x")
- case True
- then show ?thesis using dom.simps supp_of_atom_list supp_at_base
- using supp_GCons by auto
- next
- case False
- have "(atom x \<in> atom_dom ((x', b', c') #\<^sub>\<Gamma> G)) = (atom x \<in> atom_dom G)" using atom_dom.simps False by simp
- also have "... = (atom x \<in> supp G)" using GCons wfG_elims by metis
- also have "... = (atom x \<in> (supp (x', b', c') \<union> supp G))" proof
- show "atom x \<in> supp G \<Longrightarrow> atom x \<in> supp (x', b', c') \<union> supp G" by auto
- assume "atom x \<in> supp (x', b', c') \<union> supp G"
- then consider "atom x \<in> supp (x', b', c')" | "atom x \<in> supp G" by auto
- then show "atom x \<in> supp G" proof(cases)
- case 1
- assume " atom x \<in> supp (x', b', c') "
- hence "atom x \<in> supp c'" using supp_triple False supp_b_empty supp_at_base by force
-
- moreover have "P; \<B> ; (x', b' , TRUE) #\<^sub>\<Gamma>G \<turnstile>\<^sub>w\<^sub>f c'" using wfG_elim2 GCons by simp
- moreover hence "supp c' \<subseteq> supp G \<union> supp x' \<union> supp \<B>" using wfC_supp_cons by auto
- ultimately have "atom x \<in> supp G \<union> supp x' " using x_not_in_b_set by auto
- then show ?thesis using False supp_at_base by (simp add: supp_at_base)
- next
- case 2
- then show ?thesis by simp
- qed
- qed
- also have "... = (atom x \<in> supp ((x', b', c') #\<^sub>\<Gamma> G))" using supp_at_base False supp_GCons by simp
- finally show ?thesis by simp
- qed
-qed
-
-lemma wfG_atoms_supp_eq :
- fixes x::x
- assumes "wfG P \<B> G"
- shows "atom x \<in> atom_dom G \<longleftrightarrow> atom x \<in> supp G"
- using wfG_dom_supp assms by auto
-
-lemma beta_flip_eq:
- fixes x::x and xa::x and \<B>::\<B>
- shows "(x \<leftrightarrow> xa) \<bullet> \<B> = \<B>"
-proof -
- have "atom x \<sharp> \<B> \<and> atom xa \<sharp> \<B>" using x_not_in_b_set fresh_def supp_set by metis
- thus ?thesis by (simp add: flip_fresh_fresh fresh_def)
-qed
-
-lemma theta_flip_eq2:
- assumes "\<turnstile>\<^sub>w\<^sub>f \<Theta>"
- shows " (z \<leftrightarrow> za ) \<bullet> \<Theta> = \<Theta>"
-proof -
- have "supp \<Theta> = {}" using wfTh_supp assms by simp
- thus ?thesis
- by (simp add: flip_fresh_fresh fresh_def)
- qed
-
-lemma theta_flip_eq:
- assumes "wfTh \<Theta>"
- shows "(x \<leftrightarrow> xa) \<bullet> \<Theta> = \<Theta>"
- using wfTh_supp flip_fresh_fresh fresh_def
- by (simp add: assms theta_flip_eq2)
-
-lemma wfT_wfC:
- fixes c::c
- assumes "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b | c \<rbrace>" and "atom z \<sharp> \<Gamma>"
- shows "\<Theta>; \<B>; (z,b,TRUE) #\<^sub>\<Gamma>\<Gamma> \<turnstile>\<^sub>w\<^sub>f c"
-proof -
- obtain za ba ca where *:"\<lbrace> z : b | c \<rbrace> = \<lbrace> za : ba | ca \<rbrace> \<and> atom za \<sharp> (\<Theta>,\<B>,\<Gamma>) \<and> \<Theta>; \<B>; (za, ba, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f ca"
- using wfT_elims[OF assms(1)] by metis
- hence c1: "[[atom z]]lst. c = [[atom za]]lst. ca" using \<tau>.eq_iff by meson
- show ?thesis proof(cases "z=za")
- case True
- hence "ca = c" using c1 by (simp add: Abs1_eq_iff(3))
- then show ?thesis using * True by simp
- next
- case False
- have " \<turnstile>\<^sub>w\<^sub>f \<Theta>" using wfT_wf wfG_wf assms by metis
- moreover have "atom za \<sharp> \<Gamma>" using * fresh_prodN by auto
- ultimately have "\<Theta>; \<B>; (z \<leftrightarrow> za ) \<bullet> (za, ba, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f (z \<leftrightarrow> za ) \<bullet> ca"
- using wfC.eqvt theta_flip_eq2 beta_flip_eq * GCons_eqvt assms flip_fresh_fresh by metis
- moreover have "atom z \<sharp> ca"
- proof -
- have "supp ca \<subseteq> atom_dom \<Gamma> \<union> { atom za } \<union> supp \<B>" using * wfC_supp atom_dom.simps toSet.simps by fastforce
- moreover have "atom z \<notin> atom_dom \<Gamma> " using assms fresh_def wfT_wf wfG_dom_supp wfC_supp by metis
- moreover hence "atom z \<notin> atom_dom \<Gamma> \<union> { atom za }" using False by simp
- moreover have "atom z \<notin> supp \<B>" using x_not_in_b_set by simp
- ultimately show ?thesis using fresh_def False by fast
- qed
- moreover hence "(z \<leftrightarrow> za ) \<bullet> ca = c" using type_eq_subst_eq1(3) * by metis
- ultimately show ?thesis using assms G_cons_flip_fresh * by auto
- qed
-qed
-
-lemma u_not_in_dom_g:
- fixes u::u
- shows "atom u \<notin> atom_dom G"
- using toSet.simps atom_dom.simps u_not_in_x_atoms by auto
-
-lemma bv_not_in_dom_g:
- fixes bv::bv
- shows "atom bv \<notin> atom_dom G"
- using toSet.simps atom_dom.simps u_not_in_x_atoms by auto
-
-text \<open>An important lemma that confirms that @{term \<Gamma>} does not rely on mutable variables\<close>
-lemma u_not_in_g:
- fixes u::u
- assumes "wfG \<Theta> B G"
- shows "atom u \<notin> supp G"
-using assms proof(induct G rule: \<Gamma>_induct)
-case GNil
- then show ?case using supp_GNil fresh_def
- using fresh_set_empty by fastforce
-next
- case (GCons x b c \<Gamma>')
- moreover hence "atom u \<notin> supp b" using
- wfB_supp wfC_supp u_not_in_x_atoms wfG_elims wfX_wfY by auto
- moreover hence "atom u \<notin> supp x" using u_not_in_x_atoms supp_at_base by blast
- moreover hence "atom u \<notin> supp c" proof -
- have "\<Theta> ; B ; (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>' \<turnstile>\<^sub>w\<^sub>f c" using wfG_cons_wfC GCons by simp
- hence "supp c \<subseteq> atom_dom ((x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>') \<union> supp B" using wfC_supp by blast
- thus ?thesis using u_not_in_dom_g u_not_in_b_atoms
- using u_not_in_b_set by auto
- qed
- ultimately have "atom u \<notin> supp (x,b,c)" using supp_Pair by simp
- thus ?case using supp_GCons GCons wfG_elims by blast
-qed
-
-text \<open>An important lemma that confirms that types only depend on immutable variables\<close>
-lemma u_not_in_t:
- fixes u::u
- assumes "wfT \<Theta> B G \<tau>"
- shows "atom u \<notin> supp \<tau>"
-proof -
- have "supp \<tau> \<subseteq> atom_dom G \<union> supp B" using wfT_supp assms by auto
- thus ?thesis using u_not_in_dom_g u_not_in_b_set by blast
-qed
-
-lemma wfT_supp_c:
- fixes \<B>::\<B> and z::x
- assumes "wfT P \<B> \<Gamma> (\<lbrace> z : b | c \<rbrace>)"
- shows "supp c - { atom z } \<subseteq> atom_dom \<Gamma> \<union> supp \<B>"
- using wf_supp \<tau>.supp assms
- by (metis Un_subset_iff empty_set list.simps(15))
-
-lemma wfG_wfC[ms_wb]:
- assumes "wfG P \<B> ((x,b,c) #\<^sub>\<Gamma>\<Gamma>)"
- shows "wfC P \<B> ((x,b,TRUE) #\<^sub>\<Gamma>\<Gamma>) c"
-using assms proof(cases "c \<in> {TRUE,FALSE}")
- case True
- have "atom x \<sharp> \<Gamma> \<and> wfG P \<B> \<Gamma> \<and> wfB P \<B> b" using wfG_cons assms by auto
- hence "wfG P \<B> ((x,b,TRUE) #\<^sub>\<Gamma>\<Gamma>)" using wfG_cons2I by auto
- then show ?thesis using wfC_trueI wfC_falseI True by auto
-next
- case False
- then show ?thesis using wfG_elims assms by blast
-qed
-
-lemma wfT_wf_cons:
- assumes "wfT P \<B> \<Gamma> \<lbrace> z : b | c \<rbrace>" and "atom z \<sharp> \<Gamma>"
- shows "wfG P \<B> ((z,b,c) #\<^sub>\<Gamma>\<Gamma>)"
-using assms proof(cases "c \<in> { TRUE,FALSE }")
- case True
- then show ?thesis using wfT_wfC wfC_wf wfG_wfB wfG_cons2I assms wfT_wf by fastforce
-next
- case False
- then show ?thesis using wfT_wfC wfC_wf wfG_wfB wfG_cons1I wfT_wf wfT_wfC assms by fastforce
-qed
-
-lemma wfV_b_fresh:
- fixes b::b and v::v and bv::bv
- assumes "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v: b" and "bv |\<notin>| \<B>"
- shows "atom bv \<sharp> v"
-using wfV_supp bv_not_in_dom_g fresh_def assms bv_not_in_bset_supp by blast
-
-lemma wfCE_b_fresh:
- fixes b::b and ce::ce and bv::bv
- assumes "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ce: b" and "bv |\<notin>| \<B>"
- shows "atom bv \<sharp> ce"
-using bv_not_in_dom_g fresh_def assms bv_not_in_bset_supp wf_supp1(8) by fast
-
-section \<open>Freshness\<close>
-
-lemma wfG_fresh_x:
- fixes \<Gamma>::\<Gamma> and z::x
- assumes "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>" and "atom z \<sharp> \<Gamma>"
- shows "atom z \<sharp> (\<Theta>,\<B>, \<Gamma>)"
-unfolding fresh_prodN apply(intro conjI)
- using wf_supp1 wfX_wfY assms fresh_def x_not_in_b_set by(metis empty_iff)+
-
-lemma wfG_wfT:
- assumes "wfG P \<B> ((x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> G)" and "atom x \<sharp> c"
- shows "P; \<B> ; G \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b | c \<rbrace>"
-proof -
- have " P; \<B> ; (x, b, TRUE) #\<^sub>\<Gamma> G \<turnstile>\<^sub>w\<^sub>f c[z::=V_var x]\<^sub>c\<^sub>v \<and> wfB P \<B> b" using assms
- using wfG_elim2 by auto
- moreover have "atom x \<sharp> (P ,\<B>,G)" using wfG_elims assms wfG_fresh_x by metis
- ultimately have "wfT P \<B> G \<lbrace> x : b | c[z::=V_var x]\<^sub>c\<^sub>v \<rbrace>" using wfTI assms by metis
- moreover have "\<lbrace> x : b | c[z::=V_var x]\<^sub>c\<^sub>v \<rbrace> = \<lbrace> z : b | c \<rbrace>" using type_eq_subst \<open>atom x \<sharp> c\<close> by auto
- ultimately show ?thesis by auto
-qed
-
-lemma wfT_wfT_if:
- assumes "wfT \<Theta> \<B> \<Gamma> (\<lbrace> z2 : b | CE_val v == CE_val (V_lit L_false) IMP c[z::=V_var z2]\<^sub>c\<^sub>v \<rbrace>)" and "atom z2 \<sharp> (c,\<Gamma>)"
- shows "wfT \<Theta> \<B> \<Gamma> \<lbrace> z : b | c \<rbrace>"
-proof -
- have *: "atom z2 \<sharp> (\<Theta>, \<B>, \<Gamma>)" using wfG_fresh_x wfX_wfY assms fresh_Pair by metis
- have "wfB \<Theta> \<B> b" using assms wfT_elims by metis
- have "\<Theta>; \<B>; (GCons (z2,b,TRUE) \<Gamma>) \<turnstile>\<^sub>w\<^sub>f (CE_val v == CE_val (V_lit L_false) IMP c[z::=V_var z2]\<^sub>c\<^sub>v)" using wfT_wfC assms fresh_Pair by auto
- hence "\<Theta>; \<B>; ((z2,b,TRUE) #\<^sub>\<Gamma>\<Gamma>) \<turnstile>\<^sub>w\<^sub>f c[z::=V_var z2]\<^sub>c\<^sub>v" using wfC_elims by metis
- hence "wfT \<Theta> \<B> \<Gamma> (\<lbrace> z2 : b | c[z::=V_var z2]\<^sub>c\<^sub>v\<rbrace>)" using assms fresh_Pair wfTI \<open>wfB \<Theta> \<B> b\<close> * by auto
- moreover have "\<lbrace> z : b | c \<rbrace> = \<lbrace> z2 : b | c[z::=V_var z2]\<^sub>c\<^sub>v \<rbrace>" using type_eq_subst assms fresh_Pair by auto
- ultimately show ?thesis using wfTI assms by argo
-qed
-
-lemma wfT_fresh_c:
- fixes x::x
- assumes "wfT P \<B> \<Gamma> \<lbrace> z : b | c \<rbrace>" and "atom x \<sharp> \<Gamma>" and "x \<noteq> z"
- shows "atom x \<sharp> c"
-proof(rule ccontr)
- assume "\<not> atom x \<sharp> c"
- hence *:"atom x \<in> supp c" using fresh_def by auto
- moreover have "supp c - set [atom z] \<union> supp b \<subseteq> atom_dom \<Gamma> \<union> supp \<B>"
- using assms wfT_supp \<tau>.supp by blast
- moreover hence "atom x \<in> supp c - set [atom z]" using assms * by auto
- ultimately have "atom x \<in> atom_dom \<Gamma>" using x_not_in_b_set by auto
- thus False using assms wfG_atoms_supp_eq wfT_wf fresh_def by metis
-qed
-
-lemma wfG_x_fresh [simp]:
- fixes x::x
- assumes "wfG P \<B> G"
- shows "atom x \<notin> atom_dom G \<longleftrightarrow> atom x \<sharp> G"
- using wfG_atoms_supp_eq assms fresh_def by metis
-
-lemma wfD_x_fresh:
- fixes x::x
- assumes "atom x \<sharp> \<Gamma>" and "wfD P B \<Gamma> \<Delta>"
- shows "atom x \<sharp> \<Delta>"
-using assms proof(induct \<Delta> rule: \<Delta>_induct)
- case DNil
- then show ?case using supp_DNil fresh_def by auto
-next
- case (DCons u' t' \<Delta>')
- have wfg: "wfG P B \<Gamma>" using wfD_wf DCons by blast
- hence wfd: "wfD P B \<Gamma> \<Delta>'" using wfD_elims DCons by blast
- have "supp t' \<subseteq> atom_dom \<Gamma> \<union> supp B" using wfT_supp DCons wfD_elims by metis
- moreover have "atom x \<notin> atom_dom \<Gamma>" using DCons(2) fresh_def wfG_supp wfg by blast
- ultimately have "atom x \<sharp> t'" using fresh_def DCons wfG_supp wfg x_not_in_b_set by blast
- moreover have "atom x \<sharp> u'" using supp_at_base fresh_def by fastforce
- ultimately have "atom x \<sharp> (u',t')" using supp_Pair by fastforce
- thus ?case using DCons fresh_DCons wfd by fast
-qed
-
-lemma wfG_fresh_x2:
- fixes \<Gamma>::\<Gamma> and z::x and \<Delta>::\<Delta> and \<Phi>::\<Phi>
- assumes "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>" and "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>" and "atom z \<sharp> \<Gamma>"
- shows "atom z \<sharp> (\<Theta>,\<Phi>,\<B>, \<Gamma>,\<Delta>)"
- unfolding fresh_prodN apply(intro conjI)
- using wfG_fresh_x assms fresh_prod3 wfX_wfY apply metis
- using wf_supp2(5) assms fresh_def apply blast
- using assms wfG_fresh_x wfX_wfY fresh_prod3 apply metis
- using assms wfG_fresh_x wfX_wfY fresh_prod3 apply metis
- using wf_supp2(6) assms fresh_def wfD_x_fresh by metis
-
-lemma wfV_x_fresh:
- fixes v::v and b::b and \<Gamma>::\<Gamma> and x::x
- assumes "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b" and "atom x \<sharp> \<Gamma>"
- shows "atom x \<sharp> v"
-proof -
- have "supp v \<subseteq> atom_dom \<Gamma> \<union> supp \<B> " using assms wfV_supp by auto
- moreover have "atom x \<notin> atom_dom \<Gamma>" using fresh_def assms
- dom.simps subsetCE wfG_elims wfG_supp by (metis dom_supp_g)
- moreover have "atom x \<notin> supp \<B>" using x_not_in_b_set by auto
- ultimately show ?thesis using fresh_def by fast
-qed
-
-lemma wfE_x_fresh:
- fixes e::e and b::b and \<Gamma>::\<Gamma> and \<Delta>::\<Delta> and \<Phi>::\<Phi> and x::x
- assumes "\<Theta>; \<Phi>; \<B>; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f e : b" and "atom x \<sharp> \<Gamma>"
- shows "atom x \<sharp> e"
-proof -
- have "wfG \<Theta> \<B> \<Gamma>" using assms wfE_wf by auto
- hence "supp e \<subseteq> atom_dom \<Gamma> \<union> supp \<B> \<union> atom`fst`setD \<Delta>" using wfE_supp dom.simps assms by auto
- moreover have "atom x \<notin> atom_dom \<Gamma>" using fresh_def assms
- dom.simps subsetCE \<open>wfG \<Theta> \<B> \<Gamma>\<close> wfG_supp by (metis dom_supp_g)
- moreover have "atom x \<notin> atom`fst`setD \<Delta>" by auto
- ultimately show ?thesis using fresh_def x_not_in_b_set by fast
-qed
-
-lemma wfT_x_fresh:
- fixes \<tau>::\<tau> and \<Gamma>::\<Gamma> and x::x
- assumes "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau>" and "atom x \<sharp> \<Gamma>"
- shows "atom x \<sharp> \<tau>"
-proof -
- have "wfG \<Theta> \<B> \<Gamma>" using assms wfX_wfY by auto
- hence "supp \<tau> \<subseteq> atom_dom \<Gamma> \<union> supp \<B>" using wfT_supp dom.simps assms by auto
- moreover have "atom x \<notin> atom_dom \<Gamma>" using fresh_def assms
- dom.simps subsetCE \<open>wfG \<Theta> \<B> \<Gamma>\<close> wfG_supp by (metis dom_supp_g)
- moreover have "atom x \<notin> supp \<B>" using x_not_in_b_set by simp
- ultimately show ?thesis using fresh_def by fast
-qed
-
-lemma wfS_x_fresh:
- fixes s::s and \<Delta>::\<Delta> and x::x
- assumes "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b" and "atom x \<sharp> \<Gamma>"
- shows "atom x \<sharp> s"
-proof -
- have "supp s \<subseteq> atom_dom \<Gamma> \<union> atom ` fst ` setD \<Delta> \<union> supp \<B>" using wf_supp assms by metis
- moreover have "atom x \<notin> atom ` fst ` setD \<Delta>" by auto
- moreover have "atom x \<notin> atom_dom \<Gamma>" using assms fresh_def wfG_dom_supp wfX_wfY by metis
- moreover have "atom x \<notin> supp \<B>" using supp_b_empty supp_fset
- by (simp add: x_not_in_b_set)
- ultimately show ?thesis using fresh_def by fast
-qed
-
-lemma wfTh_fresh:
- fixes x
- assumes "wfTh T"
- shows "atom x \<sharp> T"
- using wf_supp1 assms fresh_def by fastforce
-
-lemmas wfTh_x_fresh = wfTh_fresh
-
-lemma wfPhi_fresh:
- fixes x
- assumes "wfPhi T P"
- shows "atom x \<sharp> P"
- using wf_supp assms fresh_def by fastforce
-
-lemmas wfPhi_x_fresh = wfPhi_fresh
-lemmas wb_x_fresh = wfTh_x_fresh wfPhi_x_fresh wfD_x_fresh wfT_x_fresh wfV_x_fresh
-
-lemma wfG_inside_fresh[ms_fresh]:
- fixes \<Gamma>::\<Gamma> and x::x
- assumes "wfG P \<B> (\<Gamma>'@((x,b,c) #\<^sub>\<Gamma>\<Gamma>))"
- shows "atom x \<notin> atom_dom \<Gamma>'"
-using assms proof(induct \<Gamma>' rule: \<Gamma>_induct)
- case GNil
- then show ?case by auto
-next
- case (GCons x1 b1 c1 \<Gamma>1)
- moreover hence "atom x \<notin> atom ` fst `({(x1,b1,c1)})" proof -
- have *: "P; \<B> \<turnstile>\<^sub>w\<^sub>f (\<Gamma>1 @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>)" using wfG_elims append_g.simps GCons by metis
- have "atom x1 \<sharp> (\<Gamma>1 @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>)" using GCons wfG_elims append_g.simps by metis
- hence "atom x1 \<notin> atom_dom (\<Gamma>1 @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>)" using wfG_dom_supp fresh_def * by metis
- thus ?thesis by auto
- qed
- ultimately show ?case using append_g.simps atom_dom.simps toSet.simps wfG_elims dom.simps
- by (metis image_insert insert_iff insert_is_Un)
-qed
-
-lemma wfG_inside_x_in_atom_dom:
- fixes c::c and x::x and \<Gamma>::\<Gamma>
- shows "atom x \<in> atom_dom ( \<Gamma>'@ (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)"
- by(induct \<Gamma>' rule: \<Gamma>_induct, (simp add: toSet.simps atom_dom.simps)+)
-
-lemma wfG_inside_x_neq:
- fixes c::c and x::x and \<Gamma>::\<Gamma> and G::\<Gamma> and xa::x
- assumes "G=( \<Gamma>'@ (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)" and "atom xa \<sharp> G" and " \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f G"
- shows "xa \<noteq> x"
-proof -
- have "atom xa \<notin> atom_dom G" using fresh_def wfG_atoms_supp_eq assms by metis
- moreover have "atom x \<in> atom_dom G" using wfG_inside_x_in_atom_dom assms by simp
- ultimately show ?thesis by auto
-qed
-
-lemma wfG_inside_x_fresh:
- fixes c::c and x::x and \<Gamma>::\<Gamma> and G::\<Gamma> and xa::x
- assumes "G=( \<Gamma>'@ (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)" and "atom xa \<sharp> G" and " \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f G"
- shows "atom xa \<sharp> x"
- using fresh_def supp_at_base wfG_inside_x_neq assms by auto
-
-lemma wfT_nil_supp:
- fixes t::\<tau>
- assumes "\<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f t"
- shows "supp t = {}"
- using wfT_supp atom_dom.simps assms toSet.simps by force
-
-section \<open>Misc\<close>
-
-lemma wfG_cons_append:
- fixes b'::b
- assumes "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f ((x', b', c') #\<^sub>\<Gamma> \<Gamma>') @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>"
- shows "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (\<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>) \<and> atom x' \<sharp> (\<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>) \<and> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b' \<and> x' \<noteq> x"
-proof -
- have "((x', b', c') #\<^sub>\<Gamma> \<Gamma>') @ (x, b, c) #\<^sub>\<Gamma> \<Gamma> = (x', b', c') #\<^sub>\<Gamma> (\<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>)" using append_g.simps by auto
- hence *:"\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (\<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>) \<and> atom x' \<sharp> (\<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>) \<and> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b'" using assms wfG_cons by metis
- moreover have "atom x' \<sharp> x" proof(rule wfG_inside_x_fresh[of "(\<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>)"])
- show "\<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma> = \<Gamma>' @ (x, b, c[x::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>" by simp
- show " atom x' \<sharp> \<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>" using * by auto
- show "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma> " using * by auto
- qed
- ultimately show ?thesis by auto
-qed
-
-lemma flip_u_eq:
- fixes u::u and u'::u and \<Theta>::\<Theta> and \<tau>::\<tau>
- assumes "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau>"
- shows "(u \<leftrightarrow> u') \<bullet> \<tau> = \<tau>" and "(u \<leftrightarrow> u') \<bullet> \<Gamma> = \<Gamma>" and "(u \<leftrightarrow> u') \<bullet> \<Theta> = \<Theta>" and "(u \<leftrightarrow> u') \<bullet> \<B> = \<B>"
-proof -
- show "(u \<leftrightarrow> u') \<bullet> \<tau> = \<tau>" using wfT_supp flip_fresh_fresh
- by (metis assms(1) fresh_def u_not_in_t)
- show "(u \<leftrightarrow> u') \<bullet> \<Gamma> = \<Gamma>" using u_not_in_g wfX_wfY assms flip_fresh_fresh fresh_def by metis
- show "(u \<leftrightarrow> u') \<bullet> \<Theta> = \<Theta>" using theta_flip_eq assms wfX_wfY by metis
- show "(u \<leftrightarrow> u') \<bullet> \<B> = \<B>" using u_not_in_b_set flip_fresh_fresh fresh_def by metis
-qed
-
-lemma wfT_wf_cons_flip:
- fixes c::c and x::x
- assumes "wfT P \<B> \<Gamma> \<lbrace> z : b | c \<rbrace>" and "atom x \<sharp> (c,\<Gamma>)"
- shows "wfG P \<B> ((x,b,c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma>\<Gamma>)"
-proof -
- have "\<lbrace> x : b | c[z::=V_var x]\<^sub>c\<^sub>v \<rbrace> = \<lbrace> z : b | c \<rbrace>" using assms freshers type_eq_subst by metis
- hence *:"wfT P \<B> \<Gamma> \<lbrace> x : b | c[z::=V_var x]\<^sub>c\<^sub>v \<rbrace>" using assms by metis
- show ?thesis proof(rule wfG_consI)
- show \<open> P; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<close> using assms wfT_wf by auto
- show \<open>atom x \<sharp> \<Gamma>\<close> using assms by auto
- show \<open> P; \<B> \<turnstile>\<^sub>w\<^sub>f b \<close> using assms wfX_wfY b_of.simps by metis
- show \<open> P; \<B> ; (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f c[z::=V_var x]\<^sub>c\<^sub>v \<close> using wfT_wfC * assms fresh_Pair by metis
- qed
-qed
-
-section \<open>Context Strengthening\<close>
-
-text \<open>We can remove an entry for a variable from the context if the variable doesn't appear in the
-term and the variable is not used later in the context or any other context\<close>
-
-lemma fresh_restrict:
- fixes y::"'a::at_base" and \<Gamma>::\<Gamma>
- assumes "atom y \<sharp> (\<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>)"
- shows "atom y \<sharp> (\<Gamma>'@\<Gamma>)"
-using assms proof(induct \<Gamma>' rule: \<Gamma>_induct)
- case GNil
- then show ?case using fresh_GCons fresh_GNil by auto
-next
- case (GCons x' b' c' \<Gamma>'')
- then show ?case using fresh_GCons fresh_GNil by auto
-qed
-
-lemma wf_restrict1:
- fixes \<Gamma>::\<Gamma> and \<Gamma>'::\<Gamma> and v::v and e::e and c::c and \<tau>::\<tau> and ts::"(string*\<tau>) list" and \<Delta>::\<Delta> and s::s and b::b and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def
- and cs::branch_s and css::branch_list
- shows "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b \<Longrightarrow> \<Gamma>=\<Gamma>\<^sub>1@((x,b',c') #\<^sub>\<Gamma>\<Gamma>\<^sub>2) \<Longrightarrow> atom x \<sharp> v \<Longrightarrow> atom x \<sharp> \<Gamma>\<^sub>1 \<Longrightarrow> \<Theta>; \<B>; \<Gamma>\<^sub>1@\<Gamma>\<^sub>2 \<turnstile>\<^sub>w\<^sub>f v : b" and
-
- "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f c \<Longrightarrow> \<Gamma>=\<Gamma>\<^sub>1@((x,b',c') #\<^sub>\<Gamma>\<Gamma>\<^sub>2) \<Longrightarrow> atom x \<sharp> c\<Longrightarrow> atom x \<sharp> \<Gamma>\<^sub>1 \<Longrightarrow> \<Theta> ; \<B> ; \<Gamma>\<^sub>1@\<Gamma>\<^sub>2 \<turnstile>\<^sub>w\<^sub>f c" and
- "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<Longrightarrow> \<Gamma>=\<Gamma>\<^sub>1@((x,b',c') #\<^sub>\<Gamma>\<Gamma>\<^sub>2) \<Longrightarrow> atom x \<sharp> \<Gamma>\<^sub>1 \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>\<^sub>1@\<Gamma>\<^sub>2" and
- "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau> \<Longrightarrow> \<Gamma>=\<Gamma>\<^sub>1@((x,b',c') #\<^sub>\<Gamma>\<Gamma>\<^sub>2) \<Longrightarrow> atom x \<sharp> \<tau>\<Longrightarrow> atom x \<sharp> \<Gamma>\<^sub>1 \<Longrightarrow> \<Theta>; \<B>; \<Gamma>\<^sub>1@\<Gamma>\<^sub>2 \<turnstile>\<^sub>w\<^sub>f \<tau>" and
- "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ts \<Longrightarrow> True" and
- "\<turnstile>\<^sub>w\<^sub>f \<Theta> \<Longrightarrow>True" and
-
- "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b \<Longrightarrow> True" and
-
- "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ce : b \<Longrightarrow> \<Gamma>=\<Gamma>\<^sub>1@((x,b',c') #\<^sub>\<Gamma>\<Gamma>\<^sub>2) \<Longrightarrow> atom x \<sharp> ce \<Longrightarrow> atom x \<sharp> \<Gamma>\<^sub>1 \<Longrightarrow> \<Theta>; \<B>; \<Gamma>\<^sub>1@\<Gamma>\<^sub>2 \<turnstile>\<^sub>w\<^sub>f ce : b" and
- "\<Theta> \<turnstile>\<^sub>w\<^sub>f td \<Longrightarrow> True"
-proof(induct arbitrary: \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1
- rule:wfV_wfC_wfG_wfT_wfTs_wfTh_wfB_wfCE_wfTD.inducts)
- case (wfV_varI \<Theta> \<B> \<Gamma> b c y)
- hence "y\<noteq>x" using v.fresh by auto
- hence "Some (b, c) = lookup (\<Gamma>\<^sub>1@\<Gamma>\<^sub>2) y" using lookup_restrict wfV_varI by metis
- then show ?case using wfV_varI wf_intros by metis
-next
- case (wfV_litI \<Theta> \<Gamma> l)
- then show ?case using e.fresh wf_intros by metis
-next
- case (wfV_pairI \<Theta> \<B> \<Gamma> v1 b1 v2 b2)
- show ?case proof
- show "\<Theta>; \<B>; \<Gamma>\<^sub>1 @ \<Gamma>\<^sub>2 \<turnstile>\<^sub>w\<^sub>f v1 : b1" using wfV_pairI by auto
- show "\<Theta>; \<B>; \<Gamma>\<^sub>1 @ \<Gamma>\<^sub>2 \<turnstile>\<^sub>w\<^sub>f v2 : b2" using wfV_pairI by auto
- qed
-next
- case (wfV_consI s dclist \<Theta> dc x b c \<B> \<Gamma> v)
- show ?case proof
- show "AF_typedef s dclist \<in> set \<Theta>" using wfV_consI by auto
- show "(dc, \<lbrace> x : b | c \<rbrace>) \<in> set dclist" using wfV_consI by auto
- show "\<Theta>; \<B>; \<Gamma>\<^sub>1 @ \<Gamma>\<^sub>2 \<turnstile>\<^sub>w\<^sub>f v : b" using wfV_consI by auto
- qed
-next
- case (wfV_conspI s bv dclist \<Theta> dc x b' c \<B> b \<Gamma> v)
- show ?case proof
- show "AF_typedef_poly s bv dclist \<in> set \<Theta>" using wfV_conspI by auto
- show "(dc, \<lbrace> x : b' | c \<rbrace>) \<in> set dclist" using wfV_conspI by auto
- show "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b" using wfV_conspI by auto
- show " \<Theta>; \<B>; \<Gamma>\<^sub>1 @ \<Gamma>\<^sub>2 \<turnstile>\<^sub>w\<^sub>f v : b'[bv::=b]\<^sub>b\<^sub>b" using wfV_conspI by auto
- show "atom bv \<sharp> (\<Theta>, \<B>, \<Gamma>\<^sub>1 @ \<Gamma>\<^sub>2, b, v)" unfolding fresh_prodN fresh_append_g using wfV_conspI fresh_prodN fresh_GCons fresh_append_g by metis
- qed
-next
- case (wfCE_valI \<Theta> \<B> \<Gamma> v b)
- then show ?case using ce.fresh wf_intros by metis
-next
- case (wfCE_plusI \<Theta> \<B> \<Gamma> v1 v2)
- then show ?case using ce.fresh wf_intros by metis
-next
- case (wfCE_leqI \<Theta> \<B> \<Gamma> v1 v2)
- then show ?case using ce.fresh wf_intros by metis
-next
- case (wfCE_eqI \<Theta> \<B> \<Gamma> v1 v2)
- then show ?case using ce.fresh wf_intros by metis
-next
- case (wfCE_fstI \<Theta> \<B> \<Gamma> v1 b1 b2)
- then show ?case using ce.fresh wf_intros by metis
-next
- case (wfCE_sndI \<Theta> \<B> \<Gamma> v1 b1 b2)
- then show ?case using ce.fresh wf_intros by metis
-next
- case (wfCE_concatI \<Theta> \<B> \<Gamma> v1 v2)
- then show ?case using ce.fresh wf_intros by metis
-next
- case (wfCE_lenI \<Theta> \<B> \<Gamma> v1)
- then show ?case using ce.fresh wf_intros by metis
-next
- case (wfTI z \<Theta> \<B> \<Gamma> b c)
- hence "x \<noteq> z" using wfTI
- fresh_GCons fresh_prodN fresh_PairD(1) fresh_gamma_append not_self_fresh by metis
- show ?case proof
- show \<open>atom z \<sharp> (\<Theta>, \<B>, \<Gamma>\<^sub>1 @ \<Gamma>\<^sub>2)\<close> using wfTI fresh_restrict[of z] using wfG_fresh_x wfX_wfY wfTI fresh_prodN by metis
- show \<open> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b \<close> using wfTI by auto
- have "\<Theta>; \<B>; ((z, b, TRUE) #\<^sub>\<Gamma> \<Gamma>\<^sub>1) @ \<Gamma>\<^sub>2 \<turnstile>\<^sub>w\<^sub>f c " proof(rule wfTI(5)[of "(z, b, TRUE) #\<^sub>\<Gamma> \<Gamma>\<^sub>1" ])
- show \<open>(z, b, TRUE) #\<^sub>\<Gamma> \<Gamma> = ((z, b, TRUE) #\<^sub>\<Gamma> \<Gamma>\<^sub>1) @ (x, b', c') #\<^sub>\<Gamma> \<Gamma>\<^sub>2\<close> using wfTI by auto
- show \<open>atom x \<sharp> c\<close> using wfTI \<tau>.fresh \<open>x \<noteq> z\<close> by auto
- show \<open>atom x \<sharp> (z, b, TRUE) #\<^sub>\<Gamma> \<Gamma>\<^sub>1\<close> using wfTI \<open>x \<noteq> z\<close> fresh_GCons by simp
- qed
- thus \<open> \<Theta>; \<B>; (z, b, TRUE) #\<^sub>\<Gamma> \<Gamma>\<^sub>1 @ \<Gamma>\<^sub>2 \<turnstile>\<^sub>w\<^sub>f c \<close> by auto
- qed
-next
- case (wfC_eqI \<Theta> \<B> \<Gamma> e1 b e2)
- show ?case proof
- show "\<Theta>; \<B>; \<Gamma>\<^sub>1 @ \<Gamma>\<^sub>2 \<turnstile>\<^sub>w\<^sub>f e1 : b " using wfC_eqI c.fresh fresh_Nil by auto
- show "\<Theta>; \<B>; \<Gamma>\<^sub>1 @ \<Gamma>\<^sub>2 \<turnstile>\<^sub>w\<^sub>f e2 : b " using wfC_eqI c.fresh fresh_Nil by auto
- qed
-next
- case (wfC_trueI \<Theta> \<Gamma>)
- then show ?case using c.fresh wf_intros by metis
-next
- case (wfC_falseI \<Theta> \<Gamma>)
- then show ?case using c.fresh wf_intros by metis
-next
- case (wfC_conjI \<Theta> \<Gamma> c1 c2)
- then show ?case using c.fresh wf_intros by metis
-next
- case (wfC_disjI \<Theta> \<Gamma> c1 c2)
- then show ?case using c.fresh wf_intros by metis
-next
-case (wfC_notI \<Theta> \<Gamma> c1)
- then show ?case using c.fresh wf_intros by metis
-next
- case (wfC_impI \<Theta> \<Gamma> c1 c2)
- then show ?case using c.fresh wf_intros by metis
-next
- case (wfG_nilI \<Theta>)
- then show ?case using wfV_varI wf_intros
- by (meson GNil_append \<Gamma>.simps(3))
-next
- case (wfG_cons1I c1 \<Theta> \<B> G x1 b1)
- show ?case proof(cases "\<Gamma>\<^sub>1=GNil")
- case True
- then show ?thesis using wfG_cons1I wfG_consI by auto
- next
- case False
- then obtain G'::\<Gamma> where *:"(x1, b1, c1) #\<^sub>\<Gamma> G' = \<Gamma>\<^sub>1" using GCons_eq_append_conv wfG_cons1I by auto
- hence **:"G=G' @ (x, b', c') #\<^sub>\<Gamma> \<Gamma>\<^sub>2" using wfG_cons1I by auto
-
- have " \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (x1, b1, c1) #\<^sub>\<Gamma> (G' @ \<Gamma>\<^sub>2)" proof(rule Wellformed.wfG_cons1I)
- show \<open>c1 \<notin> {TRUE, FALSE}\<close> using wfG_cons1I by auto
- show \<open>atom x1 \<sharp> G' @ \<Gamma>\<^sub>2\<close> using wfG_cons1I(4) ** fresh_restrict by metis
- have " atom x \<sharp> G'" using wfG_cons1I * using fresh_GCons by blast
- thus \<open> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f G' @ \<Gamma>\<^sub>2 \<close> using wfG_cons1I(3)[of G'] ** by auto
- have "atom x \<sharp> c1 \<and> atom x \<sharp> (x1, b1, TRUE) #\<^sub>\<Gamma> G'" using fresh_GCons \<open>atom x \<sharp> \<Gamma>\<^sub>1\<close> * by auto
- thus \<open> \<Theta>; \<B>; (x1, b1, TRUE) #\<^sub>\<Gamma> G' @ \<Gamma>\<^sub>2 \<turnstile>\<^sub>w\<^sub>f c1 \<close> using wfG_cons1I(6)[of "(x1, b1, TRUE) #\<^sub>\<Gamma> G'"] ** * wfG_cons1I by auto
- show \<open> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b1 \<close> using wfG_cons1I by auto
- qed
- thus ?thesis using * by auto
- qed
-next
- case (wfG_cons2I c1 \<Theta> \<B> G x1 b1)
- show ?case proof(cases "\<Gamma>\<^sub>1=GNil")
- case True
- then show ?thesis using wfG_cons2I wfG_consI by auto
- next
- case False
- then obtain G'::\<Gamma> where *:"(x1, b1, c1) #\<^sub>\<Gamma> G' = \<Gamma>\<^sub>1" using GCons_eq_append_conv wfG_cons2I by auto
- hence **:"G=G' @ (x, b', c') #\<^sub>\<Gamma> \<Gamma>\<^sub>2" using wfG_cons2I by auto
-
- have " \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (x1, b1, c1) #\<^sub>\<Gamma> (G' @ \<Gamma>\<^sub>2)" proof(rule Wellformed.wfG_cons2I)
- show \<open>c1 \<in> {TRUE, FALSE}\<close> using wfG_cons2I by auto
- show \<open>atom x1 \<sharp> G' @ \<Gamma>\<^sub>2\<close> using wfG_cons2I ** fresh_restrict by metis
- have " atom x \<sharp> G'" using wfG_cons2I * using fresh_GCons by blast
- thus \<open> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f G' @ \<Gamma>\<^sub>2 \<close> using wfG_cons2I ** by auto
- show \<open> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b1 \<close> using wfG_cons2I by auto
- qed
- thus ?thesis using * by auto
- qed
-qed(auto)+
-
-lemma wf_restrict2:
- fixes \<Gamma>::\<Gamma> and \<Gamma>'::\<Gamma> and v::v and e::e and c::c and \<tau>::\<tau> and ts::"(string*\<tau>) list" and \<Delta>::\<Delta> and s::s and b::b and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def
- and cs::branch_s and css::branch_list
- shows "\<Theta>; \<Phi>; \<B>; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f e : b \<Longrightarrow> \<Gamma>=\<Gamma>\<^sub>1@((x,b',c') #\<^sub>\<Gamma>\<Gamma>\<^sub>2) \<Longrightarrow> atom x \<sharp> e \<Longrightarrow> atom x \<sharp> \<Gamma>\<^sub>1 \<Longrightarrow> atom x \<sharp> \<Delta> \<Longrightarrow> \<Theta>; \<Phi>; \<B>; \<Gamma>\<^sub>1@\<Gamma>\<^sub>2 ; \<Delta> \<turnstile>\<^sub>w\<^sub>f e : b" and
- "\<Theta>; \<Phi>; \<B>; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b \<Longrightarrow> True" and
- "\<Theta>; \<Phi>; \<B>; \<Gamma> ; \<Delta> ; tid ; dc ; t \<turnstile>\<^sub>w\<^sub>f cs : b \<Longrightarrow> True" and
- "\<Theta>; \<Phi>; \<B>; \<Gamma> ; \<Delta> ; tid ; dclist \<turnstile>\<^sub>w\<^sub>f css : b \<Longrightarrow> True" and
- "\<Theta> \<turnstile>\<^sub>w\<^sub>f (\<Phi>::\<Phi>) \<Longrightarrow> True " and
- "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<Longrightarrow> \<Gamma>=\<Gamma>\<^sub>1@((x,b',c') #\<^sub>\<Gamma>\<Gamma>\<^sub>2) \<Longrightarrow> atom x \<sharp> \<Gamma>\<^sub>1 \<Longrightarrow> atom x \<sharp> \<Delta> \<Longrightarrow> \<Theta>; \<B>; \<Gamma>\<^sub>1@\<Gamma>\<^sub>2 \<turnstile>\<^sub>w\<^sub>f \<Delta>" and
- "\<Theta> ; \<Phi> \<turnstile>\<^sub>w\<^sub>f ftq \<Longrightarrow> True" and
- "\<Theta> ; \<Phi> ; \<B> \<turnstile>\<^sub>w\<^sub>f ft \<Longrightarrow> True"
-
-proof(induct arbitrary: \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1
- rule:wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.inducts)
- case (wfE_valI \<Theta> \<Phi> \<Gamma> \<Delta> v b)
- then show ?case using e.fresh wf_intros wf_restrict1 by metis
-next
- case (wfE_plusI \<Theta> \<Phi> \<Gamma> \<Delta> v1 v2)
- then show ?case using e.fresh wf_intros wf_restrict1 by metis
-next
- case (wfE_leqI \<Theta> \<Phi> \<Gamma> \<Delta> v1 v2)
- then show ?case using e.fresh wf_intros wf_restrict1 by metis
-next
- case (wfE_eqI \<Theta> \<Phi> \<Gamma> \<Delta> v1 b v2)
- then show ?case using e.fresh wf_intros wf_restrict1 by metis
-next
- case (wfE_fstI \<Theta> \<Phi> \<Gamma> \<Delta> v1 b1 b2)
- then show ?case using e.fresh wf_intros wf_restrict1 by metis
-next
- case (wfE_sndI \<Theta> \<Phi> \<Gamma> \<Delta> v1 b1 b2)
- then show ?case using e.fresh wf_intros wf_restrict1 by metis
-next
- case (wfE_concatI \<Theta> \<Phi> \<Gamma> \<Delta> v1 v2)
- then show ?case using e.fresh wf_intros wf_restrict1 by metis
-next
- case (wfE_splitI \<Theta> \<Phi> \<Gamma> \<Delta> v1 v2)
- then show ?case using e.fresh wf_intros wf_restrict1 by metis
-next
- case (wfE_lenI \<Theta> \<Phi> \<Gamma> \<Delta> v1)
- then show ?case using e.fresh wf_intros wf_restrict1 by metis
-next
- case (wfE_appI \<Theta> \<Phi> \<Gamma> \<Delta> f x b c \<tau> s' v)
- then show ?case using e.fresh wf_intros wf_restrict1 by metis
-next
- case (wfE_appPI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> b' bv v \<tau> f x b c s)
- show ?case proof
- show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using wfE_appPI by auto
- show \<open> \<Theta>; \<B>; \<Gamma>\<^sub>1 @ \<Gamma>\<^sub>2 \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wfE_appPI by auto
- show \<open> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b' \<close> using wfE_appPI by auto
-
- have "atom bv \<sharp> \<Gamma>\<^sub>1 @ \<Gamma>\<^sub>2" using wfE_appPI fresh_prodN fresh_restrict by metis
- thus \<open>atom bv \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>\<^sub>1 @ \<Gamma>\<^sub>2, \<Delta>, b', v, (b_of \<tau>)[bv::=b']\<^sub>b)\<close>
- using wfE_appPI fresh_prodN by auto
-
- show \<open>Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \<tau> s))) = lookup_fun \<Phi> f\<close> using wfE_appPI by auto
- show \<open> \<Theta>; \<B>; \<Gamma>\<^sub>1 @ \<Gamma>\<^sub>2 \<turnstile>\<^sub>w\<^sub>f v : b[bv::=b']\<^sub>b \<close> using wfE_appPI wf_restrict1 by auto
- qed
-next
- case (wfE_mvarI \<Theta> \<Phi> \<Gamma> \<Delta> u \<tau>)
- then show ?case using e.fresh wf_intros by metis
-next
- case (wfD_emptyI \<Theta> \<Gamma>)
- then show ?case using c.fresh wf_intros wf_restrict1 by metis
-next
- case (wfD_cons \<Theta> \<B> \<Gamma> \<Delta> \<tau> u)
- show ?case proof
- show "\<Theta>; \<B>; \<Gamma>\<^sub>1 @ \<Gamma>\<^sub>2 \<turnstile>\<^sub>w\<^sub>f \<Delta>" using wfD_cons fresh_DCons by metis
- show "\<Theta>; \<B>; \<Gamma>\<^sub>1 @ \<Gamma>\<^sub>2 \<turnstile>\<^sub>w\<^sub>f \<tau> " using wfD_cons fresh_DCons fresh_Pair wf_restrict1 by metis
- show "u \<notin> fst ` setD \<Delta>" using wfD_cons by auto
- qed
-next
- case (wfFTNone \<Theta> ft)
- then show ?case by auto
-next
- case (wfFTSome \<Theta> bv ft)
- then show ?case by auto
-next
- case (wfFTI \<Theta> B b \<Phi> x c s \<tau>)
- then show ?case by auto
-qed(auto)+
-
-lemmas wf_restrict=wf_restrict1 wf_restrict2
-
-lemma wfT_restrict2:
- fixes \<tau>::\<tau>
- assumes "wfT \<Theta> \<B> ((x, b, c) #\<^sub>\<Gamma> \<Gamma>) \<tau>" and "atom x \<sharp> \<tau>"
- shows "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau>"
- using wf_restrict1(4)[of \<Theta> \<B> "((x, b, c) #\<^sub>\<Gamma> \<Gamma>)" \<tau> GNil x "b" "c" \<Gamma>] assms fresh_GNil append_g.simps by auto
-
-lemma wfG_intros2:
- assumes "wfC P \<B> ((x,b,TRUE) #\<^sub>\<Gamma>\<Gamma>) c"
- shows "wfG P \<B> ((x,b,c) #\<^sub>\<Gamma>\<Gamma>)"
-proof -
- have "wfG P \<B> ((x,b,TRUE) #\<^sub>\<Gamma>\<Gamma>)" using wfC_wf assms by auto
- hence *:"wfG P \<B> \<Gamma> \<and> atom x \<sharp> \<Gamma> \<and> wfB P \<B> b" using wfG_elims by metis
- show ?thesis using assms proof(cases "c \<in> {TRUE,FALSE}")
- case True
- then show ?thesis using wfG_cons2I * by auto
- next
- case False
- then show ?thesis using wfG_cons1I * assms by auto
- qed
-qed
-
-section \<open>Type Definitions\<close>
-
-lemma wf_theta_weakening1:
- fixes \<Gamma>::\<Gamma> and \<Gamma>'::\<Gamma> and v::v and e::e and c::c and \<tau>::\<tau> and ts::"(string*\<tau>) list" and \<Delta>::\<Delta> and s::s and b::b and \<B> :: \<B> and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def
- and cs::branch_s and css::branch_list and t::\<tau>
-
- shows "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b \<Longrightarrow> \<turnstile>\<^sub>w\<^sub>f \<Theta>' \<Longrightarrow> set \<Theta> \<subseteq> set \<Theta>' \<Longrightarrow> \<Theta>' ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b" and
- "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f c \<Longrightarrow> \<turnstile>\<^sub>w\<^sub>f \<Theta>' \<Longrightarrow> set \<Theta> \<subseteq> set \<Theta>' \<Longrightarrow> \<Theta>' ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f c" and
- "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<Longrightarrow> \<turnstile>\<^sub>w\<^sub>f \<Theta>' \<Longrightarrow> set \<Theta> \<subseteq> set \<Theta>' \<Longrightarrow> \<Theta>' ; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>" and
- "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau> \<Longrightarrow> \<turnstile>\<^sub>w\<^sub>f \<Theta>' \<Longrightarrow> set \<Theta> \<subseteq> set \<Theta>' \<Longrightarrow> \<Theta>' ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau>" and
- "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ts \<Longrightarrow> \<turnstile>\<^sub>w\<^sub>f \<Theta>' \<Longrightarrow> set \<Theta> \<subseteq> set \<Theta>' \<Longrightarrow> \<Theta>' ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ts" and
- "\<turnstile>\<^sub>w\<^sub>f P \<Longrightarrow> True " and
- "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b \<Longrightarrow> \<turnstile>\<^sub>w\<^sub>f \<Theta>' \<Longrightarrow> set \<Theta> \<subseteq> set \<Theta>' \<Longrightarrow> \<Theta>' ; \<B> \<turnstile>\<^sub>w\<^sub>f b" and
- "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ce : b \<Longrightarrow> \<turnstile>\<^sub>w\<^sub>f \<Theta>' \<Longrightarrow> set \<Theta> \<subseteq> set \<Theta>' \<Longrightarrow> \<Theta>' ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ce : b" and
- "\<Theta> \<turnstile>\<^sub>w\<^sub>f td \<Longrightarrow> \<turnstile>\<^sub>w\<^sub>f \<Theta>' \<Longrightarrow> set \<Theta> \<subseteq> set \<Theta>' \<Longrightarrow> \<Theta>' \<turnstile>\<^sub>w\<^sub>f td"
-proof(nominal_induct b and c and \<Gamma> and \<tau> and ts and P and b and b and td
- avoiding: \<Theta>'
- rule:wfV_wfC_wfG_wfT_wfTs_wfTh_wfB_wfCE_wfTD.strong_induct)
- case (wfV_consI s dclist \<Theta> dc x b c \<B> \<Gamma> v)
- show ?case proof
- show \<open>AF_typedef s dclist \<in> set \<Theta>'\<close> using wfV_consI by auto
- show \<open>(dc, \<lbrace> x : b | c \<rbrace>) \<in> set dclist\<close> using wfV_consI by auto
- show \<open> \<Theta>' ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b \<close> using wfV_consI by auto
- qed
-next
- case (wfV_conspI s bv dclist \<Theta> dc x b' c \<B> b \<Gamma> v)
- show ?case proof
- show \<open>AF_typedef_poly s bv dclist \<in> set \<Theta>'\<close> using wfV_conspI by auto
- show \<open>(dc, \<lbrace> x : b' | c \<rbrace>) \<in> set dclist\<close> using wfV_conspI by auto
- show \<open>\<Theta>' ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b'[bv::=b]\<^sub>b\<^sub>b \<close> using wfV_conspI by auto
- show "\<Theta>' ; \<B> \<turnstile>\<^sub>w\<^sub>f b " using wfV_conspI by auto
- show "atom bv \<sharp> (\<Theta>', \<B>, \<Gamma>, b, v)" using wfV_conspI fresh_prodN by auto
- qed
-next
- case (wfTI z \<Theta> \<B> \<Gamma> b c)
- thus ?case using Wellformed.wfTI by auto
-next
- case (wfB_consI \<Theta> s dclist)
- show ?case proof
- show \<open> \<turnstile>\<^sub>w\<^sub>f \<Theta>' \<close> using wfB_consI by auto
- show \<open>AF_typedef s dclist \<in> set \<Theta>'\<close> using wfB_consI by auto
- qed
-next
- case (wfB_appI \<Theta> \<B> b s bv dclist)
- show ?case proof
- show \<open> \<turnstile>\<^sub>w\<^sub>f \<Theta>' \<close> using wfB_appI by auto
- show \<open>AF_typedef_poly s bv dclist \<in> set \<Theta>'\<close> using wfB_appI by auto
- show "\<Theta>' ; \<B> \<turnstile>\<^sub>w\<^sub>f b" using wfB_appI by simp
- qed
-qed(metis wf_intros)+
-
-lemma wf_theta_weakening2:
- fixes \<Gamma>::\<Gamma> and \<Gamma>'::\<Gamma> and v::v and e::e and c::c and \<tau>::\<tau> and ts::"(string*\<tau>) list" and \<Delta>::\<Delta> and s::s and b::b and \<B> :: \<B> and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def
- and cs::branch_s and css::branch_list and t::\<tau>
- shows
- "\<Theta>; \<Phi>; \<B>; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f e : b \<Longrightarrow> \<turnstile>\<^sub>w\<^sub>f \<Theta>' \<Longrightarrow> set \<Theta> \<subseteq> set \<Theta>' \<Longrightarrow> \<Theta>' ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f e : b" and
- "\<Theta>; \<Phi>; \<B>; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b \<Longrightarrow> \<turnstile>\<^sub>w\<^sub>f \<Theta>' \<Longrightarrow> set \<Theta> \<subseteq> set \<Theta>' \<Longrightarrow> \<Theta>' ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b" and
- "\<Theta>; \<Phi>; \<B>; \<Gamma> ; \<Delta> ; tid ; dc ; t \<turnstile>\<^sub>w\<^sub>f cs : b \<Longrightarrow> \<turnstile>\<^sub>w\<^sub>f \<Theta>' \<Longrightarrow> set \<Theta> \<subseteq> set \<Theta>' \<Longrightarrow> \<Theta>' ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> ; tid ; dc ; t \<turnstile>\<^sub>w\<^sub>f cs : b" and
- "\<Theta>; \<Phi>; \<B>; \<Gamma> ; \<Delta> ; tid ; dclist \<turnstile>\<^sub>w\<^sub>f css : b \<Longrightarrow> \<turnstile>\<^sub>w\<^sub>f \<Theta>' \<Longrightarrow> set \<Theta> \<subseteq> set \<Theta>' \<Longrightarrow> \<Theta>' ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> ; tid ; dclist \<turnstile>\<^sub>w\<^sub>f css : b" and
- "\<Theta> \<turnstile>\<^sub>w\<^sub>f (\<Phi>::\<Phi>) \<Longrightarrow> \<turnstile>\<^sub>w\<^sub>f \<Theta>' \<Longrightarrow> set \<Theta> \<subseteq> set \<Theta>' \<Longrightarrow> \<Theta>' \<turnstile>\<^sub>w\<^sub>f (\<Phi>::\<Phi>)" and
- "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<Longrightarrow> \<turnstile>\<^sub>w\<^sub>f \<Theta>' \<Longrightarrow> set \<Theta> \<subseteq> set \<Theta>' \<Longrightarrow> \<Theta>' ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>" and
- "\<Theta> ; \<Phi> \<turnstile>\<^sub>w\<^sub>f ftq \<Longrightarrow> \<turnstile>\<^sub>w\<^sub>f \<Theta>' \<Longrightarrow> set \<Theta> \<subseteq> set \<Theta>' \<Longrightarrow> \<Theta>' ; \<Phi> \<turnstile>\<^sub>w\<^sub>f ftq" and
- "\<Theta> ; \<Phi> ; \<B> \<turnstile>\<^sub>w\<^sub>f ft \<Longrightarrow> \<turnstile>\<^sub>w\<^sub>f \<Theta>' \<Longrightarrow> set \<Theta> \<subseteq> set \<Theta>' \<Longrightarrow> \<Theta>' ; \<Phi> ; \<B> \<turnstile>\<^sub>w\<^sub>f ft"
-
-proof(nominal_induct b and b and b and b and \<Phi> and \<Delta> and ftq and ft
- avoiding: \<Theta>'
-rule:wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.strong_induct)
- case (wfE_appPI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> b' bv v \<tau> f x b c s)
- show ?case proof
- show \<open> \<Theta>' \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using wfE_appPI by auto
- show \<open> \<Theta>' ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wfE_appPI by auto
- show \<open> \<Theta>' ; \<B> \<turnstile>\<^sub>w\<^sub>f b' \<close> using wfE_appPI wf_theta_weakening1 by auto
- show \<open>atom bv \<sharp> (\<Phi>, \<Theta>', \<B>, \<Gamma>, \<Delta>, b', v, (b_of \<tau>)[bv::=b']\<^sub>b)\<close> using wfE_appPI by auto
- show \<open>Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \<tau> s))) = lookup_fun \<Phi> f\<close> using wfE_appPI by auto
- show \<open> \<Theta>' ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b[bv::=b']\<^sub>b \<close> using wfE_appPI wf_theta_weakening1 by auto
- qed
-next
- case (wfS_matchI \<Theta> \<B> \<Gamma> v tid dclist \<Delta> \<Phi> cs b)
- show ?case proof
- show \<open> \<Theta>' ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : B_id tid \<close> using wfS_matchI wf_theta_weakening1 by auto
- show \<open>AF_typedef tid dclist \<in> set \<Theta>'\<close> using wfS_matchI by auto
- show \<open> \<Theta>' ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wfS_matchI by auto
- show \<open> \<Theta>' \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using wfS_matchI by auto
- show \<open>\<Theta>'; \<Phi>; \<B>; \<Gamma>; \<Delta>; tid; dclist \<turnstile>\<^sub>w\<^sub>f cs : b \<close> using wfS_matchI by auto
- qed
-next
- case (wfS_varI \<Theta> \<B> \<Gamma> \<tau> v u \<Phi> \<Delta> b s)
- show ?case proof
- show \<open> \<Theta>' ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau> \<close> using wfS_varI wf_theta_weakening1 by auto
- show \<open> \<Theta>' ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b_of \<tau> \<close> using wfS_varI wf_theta_weakening1 by auto
- show \<open>atom u \<sharp> (\<Phi>, \<Theta>', \<B>, \<Gamma>, \<Delta>, \<tau>, v, b)\<close> using wfS_varI by auto
- show \<open> \<Theta>' ; \<Phi> ; \<B> ; \<Gamma> ; (u, \<tau>) #\<^sub>\<Delta> \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b \<close> using wfS_varI by auto
- qed
-next
- case (wfS_letI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> e b' x s b)
- show ?case proof
- show \<open> \<Theta>' ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f e : b' \<close> using wfS_letI by auto
- show \<open> \<Theta>' ; \<Phi> ; \<B> ; (x, b', TRUE) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b \<close> using wfS_letI by auto
- show \<open> \<Theta>' ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wfS_letI by auto
- show \<open>atom x \<sharp> (\<Phi>, \<Theta>', \<B>, \<Gamma>, \<Delta>, e, b)\<close> using wfS_letI by auto
- qed
-next
- case (wfS_let2I \<Theta> \<Phi> \<B> \<Gamma> \<Delta> s1 \<tau> x s2 b)
- show ?case proof
- show \<open> \<Theta>' ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s1 : b_of \<tau> \<close> using wfS_let2I by auto
- show \<open> \<Theta>' ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau> \<close> using wfS_let2I wf_theta_weakening1 by auto
- show \<open> \<Theta>' ; \<Phi> ; \<B> ; (x, b_of \<tau>, TRUE) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s2 : b \<close> using wfS_let2I by auto
- show \<open>atom x \<sharp> (\<Phi>, \<Theta>', \<B>, \<Gamma>, \<Delta>, s1, b, \<tau>)\<close> using wfS_let2I by auto
- qed
-next
- case (wfS_branchI \<Theta> \<Phi> \<B> x \<tau> \<Gamma> \<Delta> s b tid dc)
- show ?case proof
- show \<open> \<Theta>' ; \<Phi> ; \<B> ; (x, b_of \<tau>, TRUE) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b \<close> using wfS_branchI by auto
- show \<open>atom x \<sharp> (\<Phi>, \<Theta>', \<B>, \<Gamma>, \<Delta>, \<Gamma>, \<tau>)\<close> using wfS_branchI by auto
- show \<open> \<Theta>' ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wfS_branchI by auto
- qed
-next
- case (wfPhi_consI f \<Phi> \<Theta> ft)
- show ?case proof
- show "f \<notin> name_of_fun ` set \<Phi>" using wfPhi_consI by auto
- show "\<Theta>' ; \<Phi> \<turnstile>\<^sub>w\<^sub>f ft" using wfPhi_consI by auto
- show "\<Theta>' \<turnstile>\<^sub>w\<^sub>f \<Phi>" using wfPhi_consI by auto
- qed
-next
- case (wfFTNone \<Theta> ft)
- then show ?case using wf_intros by metis
-next
- case (wfFTSome \<Theta> bv ft)
- then show ?case using wf_intros by metis
-next
- case (wfFTI \<Theta> B b \<Phi> x c s \<tau>)
- thus ?case using Wellformed.wfFTI wf_theta_weakening1 by simp
-next
- case (wfS_assertI \<Theta> \<Phi> \<B> x c \<Gamma> \<Delta> s b)
- show ?case proof
- show \<open> \<Theta>' ; \<Phi> ; \<B> ; (x, B_bool, c) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b \<close> using wfS_assertI wf_theta_weakening1 by auto
- show \<open> \<Theta>' ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f c \<close> using wfS_assertI wf_theta_weakening1 by auto
- show \<open> \<Theta>' ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wfS_assertI wf_theta_weakening1 by auto
- have "atom x \<sharp> \<Theta>'" using wf_supp(6)[OF \<open>\<turnstile>\<^sub>w\<^sub>f \<Theta>' \<close>] fresh_def by auto
- thus \<open>atom x \<sharp> (\<Phi>, \<Theta>', \<B>, \<Gamma>, \<Delta>, c, b, s)\<close> using wfS_assertI fresh_prodN fresh_def by simp
- qed
-qed(metis wf_intros wf_theta_weakening1 )+
-
-lemmas wf_theta_weakening = wf_theta_weakening1 wf_theta_weakening2
-
-lemma lookup_wfTD:
- fixes td::type_def
- assumes "td \<in> set \<Theta>" and "\<turnstile>\<^sub>w\<^sub>f \<Theta>"
- shows "\<Theta> \<turnstile>\<^sub>w\<^sub>f td"
- using assms proof(induct \<Theta> )
- case Nil
- then show ?case by auto
-next
- case (Cons td' \<Theta>')
- then consider "td = td'" | "td \<in> set \<Theta>'" by auto
- then have "\<Theta>' \<turnstile>\<^sub>w\<^sub>f td" proof(cases)
- case 1
- then show ?thesis using Cons using wfTh_elims by auto
- next
- case 2
- then show ?thesis using Cons using wfTh_elims by auto
- qed
- then show ?case using wf_theta_weakening Cons by (meson set_subset_Cons)
-qed
-
-subsection \<open>Simple\<close>
-
-lemma wfTh_dclist_unique:
- assumes "wfTh \<Theta>" and "AF_typedef tid dclist1 \<in> set \<Theta>" and "AF_typedef tid dclist2 \<in> set \<Theta>"
- shows "dclist1 = dclist2"
-using assms proof(induct \<Theta> rule: \<Theta>_induct)
- case TNil
- then show ?case by auto
-next
- case (AF_typedef tid' dclist' \<Theta>')
- then show ?case using wfTh_elims
- by (metis image_eqI name_of_type.simps(1) set_ConsD type_def.eq_iff(1))
-next
- case (AF_typedef_poly tid bv dclist \<Theta>')
- then show ?case using wfTh_elims by auto
-qed
-
-lemma wfTs_ctor_unique:
- fixes dclist::"(string*\<tau>) list"
- assumes "\<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f dclist" and "(c, t1) \<in> set dclist" and "(c,t2) \<in> set dclist"
- shows "t1 = t2"
- using assms proof(induct dclist rule: list.inducts)
- case Nil
- then show ?case by auto
-next
- case (Cons x1 x2)
- consider "x1 = (c,t1)" | "x1 = (c,t2)" | "x1 \<noteq> (c,t1) \<and> x1 \<noteq> (c,t2)" by auto
- thus ?case proof(cases)
- case 1
- then show ?thesis using Cons wfTs_elims set_ConsD
- by (metis fst_conv image_eqI prod.inject)
- next
- case 2
- then show ?thesis using Cons wfTs_elims set_ConsD
- by (metis fst_conv image_eqI prod.inject)
- next
- case 3
- then show ?thesis using Cons wfTs_elims by (metis set_ConsD)
- qed
-qed
-
-lemma wfTD_ctor_unique:
- assumes "\<Theta> \<turnstile>\<^sub>w\<^sub>f (AF_typedef tid dclist)" and "(c, t1) \<in> set dclist" and "(c,t2) \<in> set dclist"
- shows "t1 = t2"
- using wfTD_elims wfTs_elims assms wfTs_ctor_unique by metis
-
-lemma wfTh_ctor_unique:
- assumes "wfTh \<Theta>" and "AF_typedef tid dclist \<in> set \<Theta>" and "(c, t1) \<in> set dclist" and "(c,t2) \<in> set dclist"
- shows "t1 = t2"
- using lookup_wfTD wfTD_ctor_unique assms by metis
-
-lemma wfTs_supp_t:
- fixes dclist::"(string*\<tau>) list"
- assumes "(c,t) \<in> set dclist" and "\<Theta> ; B ; GNil \<turnstile>\<^sub>w\<^sub>f dclist"
- shows "supp t \<subseteq> supp B"
-using assms proof(induct dclist arbitrary: c t rule:list.induct)
- case Nil
- then show ?case by auto
-next
- case (Cons ct dclist')
- then consider "ct = (c,t)" | "(c,t) \<in> set dclist'" by auto
- then show ?case proof(cases)
- case 1
- then have "\<Theta> ; B ; GNil \<turnstile>\<^sub>w\<^sub>f t" using Cons wfTs_elims by blast
- thus ?thesis using wfT_supp atom_dom.simps by force
- next
- case 2
- then show ?thesis using Cons wfTs_elims by metis
- qed
-qed
-
-lemma wfTh_lookup_supp_empty:
- fixes t::\<tau>
- assumes "AF_typedef tid dclist \<in> set \<Theta>" and "(c,t) \<in> set dclist" and "\<turnstile>\<^sub>w\<^sub>f \<Theta>"
- shows "supp t = {}"
-proof -
- have "\<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f dclist" using assms lookup_wfTD wfTD_elims by metis
- thus ?thesis using wfTs_supp_t assms by force
-qed
-
-lemma wfTh_supp_b:
- assumes "AF_typedef tid dclist \<in> set \<Theta>" and "(dc,\<lbrace> z : b | c \<rbrace> ) \<in> set dclist" and "\<turnstile>\<^sub>w\<^sub>f \<Theta>"
- shows "supp b = {}"
- using assms wfTh_lookup_supp_empty \<tau>.supp by blast
-
-lemma wfTh_b_eq_iff:
- fixes bva1::bv and bva2::bv and dc::string
- assumes "(dc, \<lbrace> x1 : b1 | c1 \<rbrace>) \<in> set dclist1" and "(dc, \<lbrace> x2 : b2 | c2 \<rbrace>) \<in> set dclist2" and
- "wfTs P {|bva1|} GNil dclist1" and "wfTs P {|bva2|} GNil dclist2"
- "[[atom bva1]]lst.dclist1 = [[atom bva2]]lst.dclist2"
- shows "[[atom bva1]]lst. (dc,\<lbrace> x1 : b1 | c1 \<rbrace>) = [[atom bva2]]lst. (dc,\<lbrace> x2 : b2 | c2 \<rbrace>)"
-using assms proof(induct dclist1 arbitrary: dclist2)
- case Nil
- then show ?case by auto
-next
- case (Cons dct1' dclist1')
- show ?case proof(cases "dclist2 = []")
- case True
- then show ?thesis using Cons by auto
- next
- case False
- then obtain dct2' and dclist2' where cons:"dct2' # dclist2' = dclist2" using list.exhaust by metis
- hence *:"[[atom bva1]]lst. dclist1' = [[atom bva2]]lst. dclist2' \<and> [[atom bva1]]lst. dct1' = [[atom bva2]]lst. dct2'"
- using Cons lst_head_cons Cons cons by metis
- hence **: "fst dct1' = fst dct2'" using lst_fst[THEN lst_pure]
- by (metis (no_types) \<open>[[atom bva1]]lst. dclist1' = [[atom bva2]]lst. dclist2' \<and> [[atom bva1]]lst. dct1' = [[atom bva2]]lst. dct2'\<close>
- \<open>\<And>x2 x1 t2' t2a t2 t1. [[atom x1]]lst. (t1, t2a) = [[atom x2]]lst. (t2, t2') \<Longrightarrow> t1 = t2\<close> fst_conv surj_pair)
- show ?thesis proof(cases "fst dct1' = dc")
- case True
- have "dc \<notin> fst ` set dclist1'" using wfTs_elims Cons by (metis True fstI)
- hence 1:"(dc, \<lbrace> x1 : b1 | c1 \<rbrace>) = dct1'" using Cons by (metis fstI image_iff set_ConsD)
- have "dc \<notin> fst ` set dclist2'" using wfTs_elims Cons cons
- by (metis "**" True fstI)
- hence 2:"(dc, \<lbrace> x2 : b2 | c2 \<rbrace>) = dct2' " using Cons cons by (metis fst_conv image_eqI set_ConsD)
- then show ?thesis using Cons * 1 2 by blast
- next
- case False
- hence "fst dct2' \<noteq> dc" using ** by auto
- hence "(dc, \<lbrace> x1 : b1 | c1 \<rbrace>) \<in> set dclist1' \<and> (dc, \<lbrace> x2 : b2 | c2 \<rbrace>) \<in> set dclist2' " using Cons cons False
- by (metis fstI set_ConsD)
- moreover have "[[atom bva1]]lst. dclist1' = [[atom bva2]]lst. dclist2'" using * False by metis
- ultimately show ?thesis using Cons ** *
- using cons wfTs_elims(2) by blast
- qed
- qed
-qed
-
-
-subsection \<open>Polymorphic\<close>
-
-lemma wfTh_wfTs_poly:
- fixes dclist::"(string * \<tau>) list"
- assumes "AF_typedef_poly tyid bva dclist \<in> set P" and "\<turnstile>\<^sub>w\<^sub>f P"
- shows "P ; {|bva|} ; GNil \<turnstile>\<^sub>w\<^sub>f dclist"
-proof -
- have *:"P \<turnstile>\<^sub>w\<^sub>f AF_typedef_poly tyid bva dclist" using lookup_wfTD assms by simp
-
- obtain bv lst where *:"P ; {|bv|} ; GNil \<turnstile>\<^sub>w\<^sub>f lst \<and>
- (\<forall>c. atom c \<sharp> (dclist, lst) \<longrightarrow> atom c \<sharp> (bva, bv, dclist, lst) \<longrightarrow> (bva \<leftrightarrow> c) \<bullet> dclist = (bv \<leftrightarrow> c) \<bullet> lst)"
- using wfTD_elims(2)[OF *] by metis
-
- obtain c::bv where **:"atom c \<sharp> ((dclist, lst),(bva, bv, dclist, lst))" using obtain_fresh by metis
- have "P ; {|bv|} ; GNil \<turnstile>\<^sub>w\<^sub>f lst" using * by metis
- hence "wfTs ((bv \<leftrightarrow> c) \<bullet> P) ((bv \<leftrightarrow> c) \<bullet> {|bv|}) ((bv \<leftrightarrow> c) \<bullet> GNil) ((bv \<leftrightarrow> c) \<bullet> lst)" using ** wfTs.eqvt by metis
- hence "wfTs P {|c|} GNil ((bva \<leftrightarrow> c) \<bullet> dclist)" using * theta_flip_eq fresh_GNil assms
- proof -
- have "\<forall>b ba. (ba::bv \<leftrightarrow> b) \<bullet> P = P" by (metis \<open>\<turnstile>\<^sub>w\<^sub>f P\<close> theta_flip_eq)
- then show ?thesis
- using "*" "**" \<open>(bv \<leftrightarrow> c) \<bullet> P ; (bv \<leftrightarrow> c) \<bullet> {|bv|} ; (bv \<leftrightarrow> c) \<bullet> GNil \<turnstile>\<^sub>w\<^sub>f (bv \<leftrightarrow> c) \<bullet> lst\<close> by fastforce
- qed
- hence "wfTs ((bva \<leftrightarrow> c) \<bullet> P) ((bva \<leftrightarrow> c) \<bullet> {|bva|}) ((bva \<leftrightarrow> c) \<bullet> GNil) ((bva \<leftrightarrow> c) \<bullet> dclist)"
- using wfTs.eqvt fresh_GNil
- by (simp add: assms(2) theta_flip_eq2)
-
- thus ?thesis using wfTs.eqvt permute_flip_cancel by metis
-qed
-
-lemma wfTh_dclist_poly_unique:
- assumes "wfTh \<Theta>" and "AF_typedef_poly tid bva dclist1 \<in> set \<Theta>" and "AF_typedef_poly tid bva2 dclist2 \<in> set \<Theta>"
- shows "[[atom bva]]lst. dclist1 = [[atom bva2]]lst.dclist2"
-using assms proof(induct \<Theta> rule: \<Theta>_induct)
- case TNil
- then show ?case by auto
-next
- case (AF_typedef tid' dclist' \<Theta>')
- then show ?case using wfTh_elims by auto
-next
- case (AF_typedef_poly tid bv dclist \<Theta>')
- then show ?case using wfTh_elims image_eqI name_of_type.simps set_ConsD type_def.eq_iff
- by (metis Abs1_eq(3))
-qed
-
-lemma wfTh_poly_lookup_supp:
- fixes t::\<tau>
- assumes "AF_typedef_poly tid bv dclist \<in> set \<Theta>" and "(c,t) \<in> set dclist" and "\<turnstile>\<^sub>w\<^sub>f \<Theta>"
- shows "supp t \<subseteq> {atom bv}"
-proof -
- have "supp dclist \<subseteq> {atom bv}" using assms lookup_wfTD wf_supp1 type_def.supp
- by (metis Diff_single_insert Un_subset_iff list.simps(15) supp_Nil supp_of_atom_list)
- then show ?thesis using assms(2) proof(induct dclist)
- case Nil
- then show ?case by auto
- next
- case (Cons a dclist)
- then show ?case using supp_Pair supp_Cons
- by (metis (mono_tags, opaque_lifting) Un_empty_left Un_empty_right pure_supp subset_Un_eq subset_singletonD supp_list_member)
- qed
-qed
-
-lemma wfTh_poly_supp_b:
- assumes "AF_typedef_poly tid bv dclist \<in> set \<Theta>" and "(dc,\<lbrace> z : b | c \<rbrace> ) \<in> set dclist" and "\<turnstile>\<^sub>w\<^sub>f \<Theta>"
- shows "supp b \<subseteq> {atom bv}"
- using assms wfTh_poly_lookup_supp \<tau>.supp by force
-
-lemma subst_g_inside:
- fixes x::x and c::c and \<Gamma>::\<Gamma> and \<Gamma>'::\<Gamma>
- assumes "wfG P \<B> (\<Gamma>' @ (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)"
- shows "(\<Gamma>' @ (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)[x::=v]\<^sub>\<Gamma>\<^sub>v = (\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v@\<Gamma>)"
-using assms proof(induct \<Gamma>' rule: \<Gamma>_induct)
- case GNil
- then show ?case using subst_gb.simps by simp
-next
- case (GCons x' b' c' G)
- hence wfg:"wfG P \<B> (G @ (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) \<and> atom x' \<sharp> (G @ (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)" using wfG_elims(2)
- using GCons.prems append_g.simps by metis
- hence "atom x \<notin> atom_dom ((x', b', c') #\<^sub>\<Gamma> G)" using GCons wfG_inside_fresh by fast
- hence "x\<noteq>x'"
- using GCons append_Cons wfG_inside_fresh atom_dom.simps toSet.simps by simp
- hence "((GCons (x', b', c') G) @ (GCons (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) \<Gamma>))[x::=v]\<^sub>\<Gamma>\<^sub>v =
- (GCons (x', b', c') (G @ (GCons (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) \<Gamma>)))[x::=v]\<^sub>\<Gamma>\<^sub>v" by auto
- also have "... = GCons (x', b', c'[x::=v]\<^sub>c\<^sub>v) ((G @ (GCons (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) \<Gamma>))[x::=v]\<^sub>\<Gamma>\<^sub>v)"
- using subst_gv.simps \<open>x\<noteq>x'\<close> by simp
- also have "... = (x', b', c'[x::=v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> (G[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>)" using GCons wfg by blast
- also have "... = ((x', b', c') #\<^sub>\<Gamma> G)[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>" using subst_gv.simps \<open>x\<noteq>x'\<close> by simp
- finally show ?case by auto
-qed
-
-lemma wfTh_td_eq:
- assumes "td1 \<in> set (td2 # P)" and "wfTh (td2 # P)" and "name_of_type td1 = name_of_type td2"
- shows "td1 = td2"
-proof(rule ccontr)
- assume as: "td1 \<noteq> td2"
- have "name_of_type td2 \<notin> name_of_type ` set P" using wfTh_elims(2)[OF assms(2)] by metis
- moreover have "td1 \<in> set P" using assms as by simp
- ultimately have "name_of_type td1 \<noteq> name_of_type td2"
- by (metis rev_image_eqI)
- thus False using assms by auto
-qed
-
-lemma wfTh_td_unique:
- assumes "td1 \<in> set P" and "td2 \<in> set P" and "wfTh P" and "name_of_type td1 = name_of_type td2"
- shows "td1 = td2"
-using assms proof(induct P rule: list.induct)
- case Nil
- then show ?case by auto
-next
- case (Cons td \<Theta>')
- consider "td = td1" | "td = td2" | "td \<noteq> td1 \<and> td \<noteq> td2" by auto
- then show ?case proof(cases)
- case 1
- then show ?thesis using Cons wfTh_elims wfTh_td_eq by metis
- next
- case 2
- then show ?thesis using Cons wfTh_elims wfTh_td_eq by metis
- next
- case 3
- then show ?thesis using Cons wfTh_elims by auto
- qed
-qed
-
-lemma wfTs_distinct:
- fixes dclist::"(string * \<tau>) list"
- assumes "\<Theta> ; B ; GNil \<turnstile>\<^sub>w\<^sub>f dclist"
- shows "distinct (map fst dclist)"
-using assms proof(induct dclist rule: list.induct)
- case Nil
- then show ?case by auto
-next
- case (Cons x1 x2)
- then show ?case
- by (metis Cons.hyps Cons.prems distinct.simps(2) fst_conv list.set_map list.simps(9) wfTs_elims(2))
-qed
-
-lemma wfTh_dclist_distinct:
- assumes "AF_typedef s dclist \<in> set P" and "wfTh P"
- shows "distinct (map fst dclist)"
-proof -
- have "wfTD P (AF_typedef s dclist)" using assms lookup_wfTD by auto
- hence "wfTs P {||} GNil dclist" using wfTD_elims by metis
- thus ?thesis using wfTs_distinct by metis
-qed
-
-lemma wfTh_dc_t_unique2:
- assumes "AF_typedef s dclist' \<in> set P" and "(dc,tc' ) \<in> set dclist'" and "AF_typedef s dclist \<in> set P" and "wfTh P" and
- "(dc, tc) \<in> set dclist"
- shows "tc= tc'"
-proof -
- have "dclist = dclist'" using assms wfTh_td_unique name_of_type.simps by force
- moreover have "distinct (map fst dclist)" using wfTh_dclist_distinct assms by auto
- ultimately show ?thesis using assms
- by (meson eq_key_imp_eq_value)
-qed
-
-lemma wfTh_dc_t_unique:
- assumes "AF_typedef s dclist' \<in> set P" and "(dc, \<lbrace> x' : b' | c' \<rbrace> ) \<in> set dclist'" and "AF_typedef s dclist \<in> set P" and "wfTh P" and
- "(dc, \<lbrace> x : b | c \<rbrace>) \<in> set dclist"
- shows "\<lbrace> x' : b' | c' \<rbrace>= \<lbrace> x : b | c \<rbrace>"
- using assms wfTh_dc_t_unique2 by metis
-
-lemma wfTs_wfT:
- fixes dclist::"(string *\<tau>) list" and t::\<tau>
- assumes "\<Theta>; \<B>; GNil \<turnstile>\<^sub>w\<^sub>f dclist" and "(dc,t) \<in> set dclist"
- shows "\<Theta>; \<B>; GNil \<turnstile>\<^sub>w\<^sub>f t"
-using assms proof(induct dclist rule:list.induct)
- case Nil
- then show ?case by auto
-next
- case (Cons x1 x2)
- thus ?case using wfTs_elims(2)[OF Cons(2)] by auto
-qed
-
-lemma wfTh_wfT:
- fixes t::\<tau>
- assumes "wfTh P" and "AF_typedef tid dclist \<in> set P" and "(dc,t) \<in> set dclist"
- shows "P ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f t"
-proof -
- have "P \<turnstile>\<^sub>w\<^sub>f AF_typedef tid dclist" using lookup_wfTD assms by auto
- hence "P ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f dclist" using wfTD_elims by auto
- thus ?thesis using wfTs_wfT assms by auto
-qed
-
-lemma td_lookup_eq_iff:
- fixes dc :: string and bva1::bv and bva2::bv
- assumes "[[atom bva1]]lst. dclist1 = [[atom bva2]]lst. dclist2" and "(dc, \<lbrace> x : b | c \<rbrace>) \<in> set dclist1"
- shows "\<exists>x2 b2 c2. (dc, \<lbrace> x2 : b2 | c2 \<rbrace>) \<in> set dclist2"
-using assms proof(induct dclist1 arbitrary: dclist2)
- case Nil
- then show ?case by auto
-next
- case (Cons dct1' dclist1')
- then obtain dct2' and dclist2' where cons:"dct2' # dclist2' = dclist2" using lst_head_cons_neq_nil[OF Cons(2)] list.exhaust by metis
- hence *:"[[atom bva1]]lst. dclist1' = [[atom bva2]]lst. dclist2' \<and> [[atom bva1]]lst. dct1' = [[atom bva2]]lst. dct2'"
- using Cons lst_head_cons Cons cons by metis
- show ?case proof(cases "dc=fst dct1'")
- case True
- hence "dc = fst dct2'" using * lst_fst[ THEN lst_pure ]
- proof -
- show ?thesis
- by (metis (no_types) "local.*" True \<open>\<And>x2 x1 t2' t2a t2 t1. [[atom x1]]lst. (t1, t2a) = [[atom x2]]lst. (t2, t2') \<Longrightarrow> t1 = t2\<close> prod.exhaust_sel) (* 31 ms *)
- qed
- obtain x2 b2 and c2 where "snd dct2' = \<lbrace> x2 : b2 | c2 \<rbrace>" using obtain_fresh_z by metis
- hence "(dc, \<lbrace> x2 : b2 | c2 \<rbrace>) = dct2'" using \<open>dc = fst dct2'\<close>
- by (metis prod.exhaust_sel)
- then show ?thesis using cons by force
- next
- case False
- hence "(dc, \<lbrace> x : b | c \<rbrace>) \<in> set dclist1'" using Cons by auto
- then show ?thesis using Cons
- by (metis "local.*" cons list.set_intros(2))
- qed
-qed
-
-lemma lst_t_b_eq_iff:
- fixes bva1::bv and bva2::bv
- assumes "[[atom bva1]]lst. \<lbrace> x1 : b1 | c1 \<rbrace> = [[atom bva2]]lst. \<lbrace> x2 : b2 | c2 \<rbrace>"
- shows "[[atom bva1]]lst. b1 = [[atom bva2]]lst.b2"
-proof(subst Abs1_eq_iff_all(3)[of bva1 b1 bva2 b2],rule,rule,rule)
- fix c::bv
- assume "atom c \<sharp> ( \<lbrace> x1 : b1 | c1 \<rbrace> , \<lbrace> x2 : b2 | c2 \<rbrace>)" and "atom c \<sharp> (bva1, bva2, b1, b2)"
-
- show "(bva1 \<leftrightarrow> c) \<bullet> b1 = (bva2 \<leftrightarrow> c) \<bullet> b2" using assms Abs1_eq_iff(3) assms
- by (metis Abs1_eq_iff_fresh(3) \<open>atom c \<sharp> (bva1, bva2, b1, b2)\<close> \<tau>.fresh \<tau>.perm_simps type_eq_subst_eq2(2))
-qed
-
-lemma wfTh_typedef_poly_b_eq_iff:
- assumes "AF_typedef_poly tyid bva1 dclist1 \<in> set P" and "(dc, \<lbrace> x1 : b1 | c1 \<rbrace>) \<in> set dclist1"
- and "AF_typedef_poly tyid bva2 dclist2 \<in> set P" and "(dc, \<lbrace> x2 : b2 | c2 \<rbrace>) \<in> set dclist2" and "\<turnstile>\<^sub>w\<^sub>f P"
-shows "b1[bva1::=b]\<^sub>b\<^sub>b = b2[bva2::=b]\<^sub>b\<^sub>b"
-proof -
- have "[[atom bva1]]lst. dclist1 = [[atom bva2]]lst.dclist2" using assms wfTh_dclist_poly_unique by metis
- hence "[[atom bva1]]lst. (dc,\<lbrace> x1 : b1 | c1 \<rbrace>) = [[atom bva2]]lst. (dc,\<lbrace> x2 : b2 | c2 \<rbrace>)" using wfTh_b_eq_iff assms wfTh_wfTs_poly by metis
- hence "[[atom bva1]]lst. \<lbrace> x1 : b1 | c1 \<rbrace> = [[atom bva2]]lst. \<lbrace> x2 : b2 | c2 \<rbrace>" using lst_snd by metis
- hence "[[atom bva1]]lst. b1 = [[atom bva2]]lst.b2" using lst_t_b_eq_iff by metis
- thus ?thesis using subst_b_flip_eq_two subst_b_b_def by metis
-qed
-
-section \<open>Equivariance Lemmas\<close>
-
-lemma x_not_in_u_set[simp]:
- fixes x::x and us::"u fset"
- shows "atom x \<notin> supp us"
- by(induct us,auto, simp add: supp_finsert supp_at_base)
-
-lemma wfS_flip_eq:
- fixes s1::s and x1::x and s2::s and x2::x and \<Delta>::\<Delta>
- assumes "[[atom x1]]lst. s1 = [[atom x2]]lst. s2" and "[[atom x1]]lst. t1 = [[atom x2]]lst. t2" and "[[atom x1]]lst. c1 = [[atom x2]]lst. c2" and "atom x2 \<sharp> \<Gamma>" and
- " \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>" and
- "\<Theta> ; \<Phi> ; \<B> ; (x1, b, c1) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s1 : b_of t1"
- shows "\<Theta> ; \<Phi> ; \<B> ; (x2, b, c2) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s2 : b_of t2"
-proof(cases "x1=x2")
- case True
- hence "s1 = s2 \<and> t1 = t2 \<and> c1 = c2" using assms Abs1_eq_iff by metis
- then show ?thesis using assms True by simp
-next
- case False
- have "\<turnstile>\<^sub>w\<^sub>f \<Theta> \<and> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<and> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>" using wfX_wfY assms by metis
- moreover have "atom x1 \<sharp> \<Gamma>" using wfX_wfY wfG_elims assms by metis
- moreover hence "atom x1 \<sharp> \<Delta> \<and> atom x2 \<sharp> \<Delta> " using wfD_x_fresh assms by auto
- ultimately have " \<Theta> ; \<Phi> ; \<B> ; (x2 \<leftrightarrow> x1) \<bullet> ((x1, b, c1) #\<^sub>\<Gamma> \<Gamma>) ; \<Delta> \<turnstile>\<^sub>w\<^sub>f (x2 \<leftrightarrow> x1) \<bullet> s1 : (x2 \<leftrightarrow> x1) \<bullet> b_of t1"
- using wfS.eqvt theta_flip_eq phi_flip_eq assms flip_base_eq beta_flip_eq flip_fresh_fresh supp_b_empty by metis
- hence " \<Theta> ; \<Phi> ; \<B> ; ((x2, b, (x2 \<leftrightarrow> x1) \<bullet> c1) #\<^sub>\<Gamma> ((x2 \<leftrightarrow> x1) \<bullet> \<Gamma>)) ; \<Delta> \<turnstile>\<^sub>w\<^sub>f (x2 \<leftrightarrow> x1) \<bullet> s1 : b_of ((x2 \<leftrightarrow> x2) \<bullet> t1)" by fastforce
- thus ?thesis using assms Abs1_eq_iff
- proof -
- have f1: "x2 = x1 \<and> t2 = t1 \<or> x2 \<noteq> x1 \<and> t2 = (x2 \<leftrightarrow> x1) \<bullet> t1 \<and> atom x2 \<sharp> t1"
- by (metis (full_types) Abs1_eq_iff(3) \<open>[[atom x1]]lst. t1 = [[atom x2]]lst. t2\<close>) (* 125 ms *)
- then have "x2 \<noteq> x1 \<and> s2 = (x2 \<leftrightarrow> x1) \<bullet> s1 \<and> atom x2 \<sharp> s1 \<longrightarrow> b_of t2 = (x2 \<leftrightarrow> x1) \<bullet> b_of t1"
- by (metis b_of.eqvt) (* 0.0 ms *)
- then show ?thesis
- using f1 by (metis (no_types) Abs1_eq_iff(3) G_cons_flip_fresh3 \<open>[[atom x1]]lst. c1 = [[atom x2]]lst. c2\<close> \<open>[[atom x1]]lst. s1 = [[atom x2]]lst. s2\<close> \<open>\<Theta> ; \<Phi> ; \<B> ; (x1, b, c1) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s1 : b_of t1\<close> \<open>\<Theta> ; \<Phi> ; \<B> ; (x2 \<leftrightarrow> x1) \<bullet> ((x1, b, c1) #\<^sub>\<Gamma> \<Gamma>) ; \<Delta> \<turnstile>\<^sub>w\<^sub>f (x2 \<leftrightarrow> x1) \<bullet> s1 : (x2 \<leftrightarrow> x1) \<bullet> b_of t1\<close> \<open>atom x1 \<sharp> \<Gamma>\<close> \<open>atom x2 \<sharp> \<Gamma>\<close>) (* 593 ms *)
- qed
-qed
-
-section \<open>Lookup\<close>
-
-lemma wf_not_in_prefix:
- assumes "\<Theta> ; B \<turnstile>\<^sub>w\<^sub>f (\<Gamma>'@(x,b1,c1) #\<^sub>\<Gamma>\<Gamma>)"
- shows "x \<notin> fst ` toSet \<Gamma>'"
-using assms proof(induct \<Gamma>' rule: \<Gamma>.induct)
- case GNil
- then show ?case by simp
-next
- case (GCons xbc \<Gamma>')
- then obtain x' and b' and c'::c where xbc: "xbc=(x',b',c')"
- using prod_cases3 by blast
- hence *:"(xbc #\<^sub>\<Gamma> \<Gamma>') @ (x, b1, c1) #\<^sub>\<Gamma> \<Gamma> = ((x',b',c') #\<^sub>\<Gamma>(\<Gamma>'@ ((x, b1, c1) #\<^sub>\<Gamma> \<Gamma>)))" by simp
- hence "atom x' \<sharp> (\<Gamma>'@(x,b1,c1) #\<^sub>\<Gamma>\<Gamma>)" using wfG_elims(2) GCons by metis
-
- moreover have "\<Theta> ; B \<turnstile>\<^sub>w\<^sub>f (\<Gamma>' @ (x, b1, c1) #\<^sub>\<Gamma> \<Gamma>)" using GCons wfG_elims * by metis
- ultimately have "atom x' \<notin> atom_dom (\<Gamma>'@(x,b1,c1) #\<^sub>\<Gamma>\<Gamma>)" using wfG_dom_supp GCons append_g.simps xbc fresh_def by fast
- hence "x' \<noteq> x" using GCons fresh_GCons xbc by fastforce
- then show ?case using GCons xbc toSet.simps
- using Un_commute \<open>\<Theta> ; B \<turnstile>\<^sub>w\<^sub>f \<Gamma>' @ (x, b1, c1) #\<^sub>\<Gamma> \<Gamma>\<close> atom_dom.simps by auto
-qed
-
-lemma lookup_inside_wf[simp]:
- assumes "\<Theta> ; B \<turnstile>\<^sub>w\<^sub>f (\<Gamma>'@(x,b1,c1) #\<^sub>\<Gamma>\<Gamma>)"
- shows "Some (b1,c1) = lookup (\<Gamma>'@(x,b1,c1) #\<^sub>\<Gamma>\<Gamma>) x"
- using wf_not_in_prefix lookup_inside assms by fast
-
-lemma lookup_weakening:
- fixes \<Theta>::\<Theta> and \<Gamma>::\<Gamma> and \<Gamma>'::\<Gamma>
- assumes "Some (b,c) = lookup \<Gamma> x" and "toSet \<Gamma> \<subseteq> toSet \<Gamma>'" and "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>'" and "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>"
- shows "Some (b,c) = lookup \<Gamma>' x"
-proof -
- have "(x,b,c) \<in> toSet \<Gamma> \<and> (\<forall>b' c'. (x,b',c') \<in> toSet \<Gamma> \<longrightarrow> b'=b \<and> c'=c)" using assms lookup_iff toSet.simps by force
- hence "(x,b,c) \<in> toSet \<Gamma>'" using assms by auto
- moreover have "(\<forall>b' c'. (x,b',c') \<in> toSet \<Gamma>' \<longrightarrow> b'=b \<and> c'=c)" using assms wf_g_unique
- using calculation by auto
- ultimately show ?thesis using lookup_iff
- using assms(3) by blast
-qed
-
-lemma wfPhi_lookup_fun_unique:
- fixes \<Phi>::\<Phi>
- assumes "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>" and "AF_fundef f fd \<in> set \<Phi>"
- shows "Some (AF_fundef f fd) = lookup_fun \<Phi> f"
-using assms proof(induct \<Phi> rule: list.induct )
- case Nil
- then show ?case using lookup_fun.simps by simp
-next
- case (Cons a \<Phi>')
- then obtain f' and fd' where a:"a = AF_fundef f' fd'" using fun_def.exhaust by auto
- have wf: "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>' \<and> f' \<notin> name_of_fun ` set \<Phi>' " using wfPhi_elims Cons a by metis
- then show ?case using Cons lookup_fun.simps using Cons lookup_fun.simps wf a
- by (metis image_eqI name_of_fun.simps set_ConsD)
-qed
-
-lemma lookup_fun_weakening:
- fixes \<Phi>'::\<Phi>
- assumes "Some fd = lookup_fun \<Phi> f" and "set \<Phi> \<subseteq> set \<Phi>'" and "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>'"
- shows "Some fd = lookup_fun \<Phi>' f"
-using assms proof(induct \<Phi> )
- case Nil
- then show ?case using lookup_fun.simps by simp
-next
- case (Cons a \<Phi>'')
- then obtain f' and fd' where a: "a = AF_fundef f' fd'" using fun_def.exhaust by auto
- then show ?case proof(cases "f=f'")
- case True
- then show ?thesis using lookup_fun.simps Cons wfPhi_lookup_fun_unique a
- by (metis lookup_fun_member subset_iff)
- next
- case False
- then show ?thesis using lookup_fun.simps Cons
- using \<open>a = AF_fundef f' fd'\<close> by auto
- qed
-qed
-
-lemma fundef_poly_fresh_bv:
- assumes "atom bv2 \<sharp> (bv1,b1,c1,\<tau>1,s1)"
- shows * : "(AF_fun_typ_some bv2 (AF_fun_typ x1 ((bv1\<leftrightarrow>bv2) \<bullet> b1) ((bv1\<leftrightarrow>bv2) \<bullet>c1) ((bv1\<leftrightarrow>bv2) \<bullet> \<tau>1) ((bv1\<leftrightarrow>bv2) \<bullet> s1)) = (AF_fun_typ_some bv1 (AF_fun_typ x1 b1 c1 \<tau>1 s1)))"
- (is "(AF_fun_typ_some ?bv ?fun_typ = AF_fun_typ_some ?bva ?fun_typa)")
-proof -
- have 1:"atom bv2 \<notin> set [atom x1]" using bv_not_in_x_atoms by simp
- have 2:"bv1 \<noteq> bv2" using assms by auto
- have 3:"(bv2 \<leftrightarrow> bv1) \<bullet> x1 = x1" using pure_fresh flip_fresh_fresh
- by (simp add: flip_fresh_fresh)
- have " AF_fun_typ x1 ((bv1 \<leftrightarrow> bv2) \<bullet> b1) ((bv1 \<leftrightarrow> bv2) \<bullet> c1) ((bv1 \<leftrightarrow> bv2) \<bullet> \<tau>1) ((bv1 \<leftrightarrow> bv2) \<bullet> s1) = (bv2 \<leftrightarrow> bv1) \<bullet> AF_fun_typ x1 b1 c1 \<tau>1 s1"
- using 1 2 3 assms by (simp add: flip_commute)
- moreover have "(atom bv2 \<sharp> c1 \<and> atom bv2 \<sharp> \<tau>1 \<and> atom bv2 \<sharp> s1 \<or> atom bv2 \<in> set [atom x1]) \<and> atom bv2 \<sharp> b1"
- using 1 2 3 assms fresh_prod5 by metis
- ultimately show ?thesis unfolding fun_typ_q.eq_iff Abs1_eq_iff(3) fun_typ.fresh 1 2 by fastforce
-qed
-
-
-text \<open>It is possible to collapse some of the easy to prove inductive cases into a single proof at the qed line
- but this makes it fragile under change. For example, changing the lemma statement might make one of the previously
- trivial cases non-trivial and so the collapsing needs to be unpacked. Is there a way to find which case
- has failed in the qed line?\<close>
-
-lemma wb_b_weakening1:
- fixes \<Gamma>::\<Gamma> and \<Gamma>'::\<Gamma> and v::v and e::e and c::c and \<tau>::\<tau> and ts::"(string*\<tau>) list" and \<Delta>::\<Delta> and s::s and \<B>::\<B> and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def
- and cs::branch_s and css::branch_list
-
- shows "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b \<Longrightarrow> \<B> |\<subseteq>| \<B>' \<Longrightarrow> \<Theta>; \<B>' ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b" and
- "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f c \<Longrightarrow>\<B> |\<subseteq>| \<B>' \<Longrightarrow> \<Theta>; \<B>' ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f c" and
- "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<Longrightarrow>\<B> |\<subseteq>| \<B>' \<Longrightarrow> \<Theta>; \<B>' \<turnstile>\<^sub>w\<^sub>f \<Gamma> " and
- "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau> \<Longrightarrow> \<B> |\<subseteq>| \<B>' \<Longrightarrow> \<Theta>; \<B>' ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau>" and
- "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ts \<Longrightarrow> \<B> |\<subseteq>| \<B>' \<Longrightarrow> \<Theta>; \<B>' ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ts" and
- "\<turnstile>\<^sub>w\<^sub>f P \<Longrightarrow> True " and
- "wfB \<Theta> \<B> b \<Longrightarrow> \<B> |\<subseteq>| \<B>' \<Longrightarrow> wfB \<Theta> \<B>' b" and
- "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ce : b \<Longrightarrow> \<B> |\<subseteq>| \<B>' \<Longrightarrow> \<Theta>; \<B>' ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ce : b" and
- "\<Theta> \<turnstile>\<^sub>w\<^sub>f td \<Longrightarrow> True"
-proof(nominal_induct b and c and \<Gamma> and \<tau> and ts and P and b and b and td
- avoiding: \<B>'
-rule:wfV_wfC_wfG_wfT_wfTs_wfTh_wfB_wfCE_wfTD.strong_induct)
- case (wfV_conspI s bv dclist \<Theta> dc x b' c \<B> b \<Gamma> v)
- show ?case proof
- show \<open>AF_typedef_poly s bv dclist \<in> set \<Theta>\<close> using wfV_conspI by metis
- show \<open>(dc, \<lbrace> x : b' | c \<rbrace>) \<in> set dclist\<close> using wfV_conspI by auto
- show \<open> \<Theta> ; \<B>' \<turnstile>\<^sub>w\<^sub>f b \<close> using wfV_conspI by auto
- show \<open>atom bv \<sharp> (\<Theta>, \<B>', \<Gamma>, b, v)\<close> using fresh_prodN wfV_conspI by auto
- thus \<open> \<Theta>; \<B>' ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b'[bv::=b]\<^sub>b\<^sub>b \<close> using wfV_conspI by simp
- qed
-next
- case (wfTI z \<Theta> \<B> \<Gamma> b c)
- show ?case proof
- show "atom z \<sharp> (\<Theta>, \<B>', \<Gamma>)" using wfTI by auto
- show "\<Theta>; \<B>' \<turnstile>\<^sub>w\<^sub>f b " using wfTI by auto
- show "\<Theta>; \<B>' ; (z, b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f c " using wfTI by auto
- qed
-qed( (auto simp add: wf_intros | metis wf_intros)+ )
-
-lemma wb_b_weakening2:
- fixes \<Gamma>::\<Gamma> and \<Gamma>'::\<Gamma> and v::v and e::e and c::c and \<tau>::\<tau> and ts::"(string*\<tau>) list" and \<Delta>::\<Delta> and s::s and \<B>::\<B> and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def
- and cs::branch_s and css::branch_list
-
- shows
- "\<Theta>; \<Phi>; \<B>; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f e : b \<Longrightarrow> \<B> |\<subseteq>| \<B>' \<Longrightarrow> \<Theta> ; \<Phi> ; \<B>' ; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f e : b" and
- "\<Theta>; \<Phi>; \<B>; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b \<Longrightarrow> \<B> |\<subseteq>| \<B>' \<Longrightarrow> \<Theta> ; \<Phi> ; \<B>' ; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b" and
- "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> ; tid ; dc ; t \<turnstile>\<^sub>w\<^sub>f cs : b \<Longrightarrow> \<B> |\<subseteq>| \<B>' \<Longrightarrow> \<Theta> ; \<Phi> ; \<B>' ; \<Gamma> ; \<Delta> ; tid ; dc ; t \<turnstile>\<^sub>w\<^sub>f cs : b" and
- "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> ; tid ; dclist \<turnstile>\<^sub>w\<^sub>f css : b \<Longrightarrow> \<B> |\<subseteq>| \<B>' \<Longrightarrow> \<Theta> ; \<Phi> ; \<B>' ; \<Gamma> ; \<Delta> ; tid ; dclist \<turnstile>\<^sub>w\<^sub>f css : b" and
- "\<Theta> \<turnstile>\<^sub>w\<^sub>f (\<Phi>::\<Phi>) \<Longrightarrow> True" and
- "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<Longrightarrow> \<B> |\<subseteq>| \<B>' \<Longrightarrow> \<Theta>; \<B>' ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>" and
- "\<Theta> ; \<Phi> \<turnstile>\<^sub>w\<^sub>f ftq \<Longrightarrow> True" and
- "\<Theta> ; \<Phi> ; \<B> \<turnstile>\<^sub>w\<^sub>f ft \<Longrightarrow> \<B> |\<subseteq>| \<B>' \<Longrightarrow> \<Theta> ; \<Phi> ; \<B>' \<turnstile>\<^sub>w\<^sub>f ft"
-proof(nominal_induct b and b and b and b and \<Phi> and \<Delta> and ftq and ft
- avoiding: \<B>'
- rule:wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.strong_induct)
- case (wfE_valI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v b)
- then show ?case using wf_intros wb_b_weakening1 by metis
-next
- case (wfE_plusI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 v2)
- then show ?case using wf_intros wb_b_weakening1 by metis
-next
- case (wfE_leqI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 v2)
- then show ?case using wf_intros wb_b_weakening1 by metis
-next
- case (wfE_eqI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 b v2)
- then show ?case using wf_intros wb_b_weakening1
- by meson
-next
- case (wfE_fstI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 b1 b2)
- then show ?case using Wellformed.wfE_fstI wb_b_weakening1 by metis
-next
- case (wfE_sndI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 b1 b2)
- then show ?case using wf_intros wb_b_weakening1 by metis
-next
- case (wfE_concatI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 v2)
- then show ?case using wf_intros wb_b_weakening1 by metis
-next
- case (wfE_splitI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 v2)
- then show ?case using wf_intros wb_b_weakening1 by metis
-next
- case (wfE_lenI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1)
- then show ?case using wf_intros wb_b_weakening1 by metis
-next
- case (wfE_appI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> f ft v)
- then show ?case using wf_intros using wb_b_weakening1 by meson
-next
- case (wfE_appPI \<Theta> \<Phi> \<B>1 \<Gamma> \<Delta> b' bv1 v1 \<tau>1 f1 x1 b1 c1 s1)
-
- have "\<Theta> ; \<Phi> ; \<B>' ; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f AE_appP f1 b' v1 : (b_of \<tau>1)[bv1::=b']\<^sub>b"
- proof
- show "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>" using wfE_appPI by auto
- show "\<Theta>; \<B>' ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> " using wfE_appPI by auto
- show "\<Theta>; \<B>' \<turnstile>\<^sub>w\<^sub>f b' " using wfE_appPI wb_b_weakening1 by auto
- thus " atom bv1 \<sharp> (\<Phi>, \<Theta>, \<B>', \<Gamma>, \<Delta>, b', v1, (b_of \<tau>1)[bv1::=b']\<^sub>b)"
- using wfE_appPI fresh_prodN by auto
-
- show "Some (AF_fundef f1 (AF_fun_typ_some bv1 (AF_fun_typ x1 b1 c1 \<tau>1 s1))) = lookup_fun \<Phi> f1" using wfE_appPI by auto
- show "\<Theta>; \<B>' ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v1 : b1[bv1::=b']\<^sub>b " using wfE_appPI wb_b_weakening1 by auto
- qed
- then show ?case by auto
-next
- case (wfE_mvarI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> u \<tau>)
- then show ?case using wf_intros wb_b_weakening1 by metis
-next
- case (wfS_valI \<Theta> \<Phi> \<B> \<Gamma> v b \<Delta>)
- then show ?case using wf_intros wb_b_weakening1 by metis
-next
- case (wfS_letI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> e b' x s b)
- show ?case proof
- show \<open> \<Theta> ; \<Phi> ; \<B>' ; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f e : b' \<close> using wfS_letI by auto
- show \<open> \<Theta> ; \<Phi> ; \<B>' ; (x, b', TRUE) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b \<close> using wfS_letI by auto
- show \<open> \<Theta>; \<B>' ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wfS_letI by auto
- show \<open>atom x \<sharp> (\<Phi>, \<Theta>, \<B>', \<Gamma>, \<Delta>, e, b)\<close> using wfS_letI by auto
- qed
-next
- case (wfS_let2I \<Theta> \<Phi> \<B> \<Gamma> \<Delta> s1 \<tau> x s2 b)
- then show ?case using wb_b_weakening1 Wellformed.wfS_let2I by simp
-next
- case (wfS_ifI \<Theta> \<B> \<Gamma> v \<Phi> \<Delta> s1 b s2)
- then show ?case using wb_b_weakening1 Wellformed.wfS_ifI by simp
-next
- case (wfS_varI \<Theta> \<B> \<Gamma> \<tau> v u \<Delta> \<Phi> s b)
- then show ?case using wb_b_weakening1 Wellformed.wfS_varI by simp
-next
- case (wfS_assignI u \<tau> \<Delta> \<Theta> \<B> \<Gamma> \<Phi> v)
- then show ?case using wb_b_weakening1 Wellformed.wfS_assignI by simp
-next
-case (wfS_whileI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> s1 s2 b)
- then show ?case using wb_b_weakening1 Wellformed.wfS_whileI by simp
-next
- case (wfS_seqI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> s1 s2 b)
- then show ?case using Wellformed.wfS_seqI by metis
-next
- case (wfS_matchI \<Theta> \<B> \<Gamma> v tid dclist \<Delta> \<Phi> cs b)
- then show ?case using wb_b_weakening1 Wellformed.wfS_matchI by metis
-next
- case (wfS_branchI \<Theta> \<Phi> \<B> x \<tau> \<Gamma> \<Delta> s b tid dc)
- then show ?case using Wellformed.wfS_branchI by auto
-next
- case (wfS_finalI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist' cs b dclist)
- then show ?case using wf_intros by metis
-next
- case (wfS_cons \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist' cs b css dclist)
- then show ?case using wf_intros by metis
-next
- case (wfD_emptyI \<Theta> \<B> \<Gamma>)
- then show ?case using wf_intros wb_b_weakening1 by metis
-next
- case (wfD_cons \<Theta> \<B> \<Gamma> \<Delta> \<tau> u)
- then show ?case using wf_intros wb_b_weakening1 by metis
-next
- case (wfPhi_emptyI \<Theta>)
- then show ?case using wf_intros wb_b_weakening1 by metis
-next
- case (wfPhi_consI f \<Theta> \<Phi> ft)
- then show ?case using wf_intros wb_b_weakening1 by metis
-next
- case (wfFTSome \<Theta> bv ft)
- then show ?case using wf_intros wb_b_weakening1 by metis
-next
- case (wfFTI \<Theta> B b s x c \<tau> \<Phi>)
- show ?case proof
- show "\<Theta>; \<B>' \<turnstile>\<^sub>w\<^sub>f b" using wfFTI wb_b_weakening1 by auto
-
- show "supp c \<subseteq> {atom x}" using wfFTI wb_b_weakening1 by auto
- show "\<Theta>; \<B>' ; (x, b, c) #\<^sub>\<Gamma> GNil \<turnstile>\<^sub>w\<^sub>f \<tau> " using wfFTI wb_b_weakening1 by auto
- show "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> " using wfFTI wb_b_weakening1 by auto
- from \<open> B |\<subseteq>| \<B>'\<close> have "supp B \<subseteq> supp \<B>'" proof(induct B)
- case empty
- then show ?case by auto
- next
- case (insert x B)
- then show ?case
- by (metis fsubset_funion_eq subset_Un_eq supp_union_fset)
- qed
- thus "supp s \<subseteq> {atom x} \<union> supp \<B>'" using wfFTI by auto
- qed
-next
- case (wfS_assertI \<Theta> \<Phi> \<B> x c \<Gamma> \<Delta> s b)
- show ?case proof
- show \<open> \<Theta> ; \<Phi> ; \<B>' ; (x, B_bool, c) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b \<close> using wb_b_weakening1 wfS_assertI by simp
- show \<open> \<Theta>; \<B>' ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f c \<close> using wb_b_weakening1 wfS_assertI by simp
- show \<open> \<Theta>; \<B>' ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wb_b_weakening1 wfS_assertI by simp
- have "atom x \<sharp> \<B>'" using x_not_in_b_set fresh_def by metis
- thus \<open>atom x \<sharp> (\<Phi>, \<Theta>, \<B>', \<Gamma>, \<Delta>, c, b, s)\<close> using wfS_assertI fresh_prodN by simp
- qed
-qed(auto)
-
-lemmas wb_b_weakening = wb_b_weakening1 wb_b_weakening2
-
-lemma wfG_b_weakening:
- fixes \<Gamma>::\<Gamma>
- assumes "\<B> |\<subseteq>| \<B>'" and "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>"
- shows "\<Theta>; \<B>' \<turnstile>\<^sub>w\<^sub>f \<Gamma> "
- using wb_b_weakening assms by auto
-
-lemma wfT_b_weakening:
- fixes \<Gamma>::\<Gamma> and \<Theta>::\<Theta> and \<tau>::\<tau>
- assumes "\<B> |\<subseteq>| \<B>'" and "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau>"
- shows "\<Theta>; \<B>' ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau> "
- using wb_b_weakening assms by auto
-
-lemma wfB_subst_wfB:
- fixes \<tau>::\<tau> and b'::b and b::b
- assumes "\<Theta> ; {|bv|} \<turnstile>\<^sub>w\<^sub>f b" and "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b'"
- shows "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b[bv::=b']\<^sub>b\<^sub>b "
-using assms proof(nominal_induct b rule:b.strong_induct)
- case B_int
- hence "\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f B_int" using wfB_intI wfX_wfY by fast
- then show ?case using subst_bb.simps wb_b_weakening by fastforce
-next
- case B_bool
- hence "\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f B_bool" using wfB_boolI wfX_wfY by fast
- then show ?case using subst_bb.simps wb_b_weakening by fastforce
-next
- case (B_id x )
- hence " \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (B_id x)" using wfB_consI wfB_elims wfX_wfY by metis
- then show ?case using subst_bb.simps(4) by auto
-next
- case (B_pair x1 x2)
- then show ?case using subst_bb.simps
- by (metis wfB_elims(1) wfB_pairI)
-next
- case B_unit
- hence "\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f B_unit" using wfB_unitI wfX_wfY by fast
- then show ?case using subst_bb.simps wb_b_weakening by fastforce
-next
- case B_bitvec
- hence "\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f B_bitvec" using wfB_bitvecI wfX_wfY by fast
- then show ?case using subst_bb.simps wb_b_weakening by fastforce
-next
- case (B_var x)
- then show ?case
- proof -
- have False
- using B_var.prems(1) wfB.cases by fastforce (* 781 ms *)
- then show ?thesis by metis
- qed
-next
- case (B_app s b)
- then obtain bv' dclist where *:"AF_typedef_poly s bv' dclist \<in> set \<Theta> \<and> \<Theta> ; {|bv|} \<turnstile>\<^sub>w\<^sub>f b" using wfB_elims by metis
- show ?case unfolding subst_b_simps proof
- show "\<turnstile>\<^sub>w\<^sub>f \<Theta> " using B_app wfX_wfY by metis
- show "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f b[bv::=b']\<^sub>b\<^sub>b " using * B_app forget_subst wfB_supp fresh_def
- by (metis ex_in_conv subset_empty subst_b_b_def supp_empty_fset)
- show "AF_typedef_poly s bv' dclist \<in> set \<Theta>" using * by auto
- qed
-qed
-
-lemma wfT_subst_wfB:
- fixes \<tau>::\<tau> and b'::b
- assumes "\<Theta> ; {|bv|} ; (x, b, c) #\<^sub>\<Gamma> GNil \<turnstile>\<^sub>w\<^sub>f \<tau>" and "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b'"
- shows "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (b_of \<tau>)[bv::=b']\<^sub>b\<^sub>b "
-proof -
- obtain b where "\<Theta> ; {|bv|} \<turnstile>\<^sub>w\<^sub>f b \<and> b_of \<tau> = b" using wfT_elims b_of.simps assms by metis
- thus ?thesis using wfB_subst_wfB assms by auto
-qed
-
-lemma wfG_cons_unique:
- assumes "(x1,b1,c1) \<in> toSet (((x,b,c) #\<^sub>\<Gamma>\<Gamma>))" and "wfG \<Theta> \<B> (((x,b,c) #\<^sub>\<Gamma>\<Gamma>))" and "x = x1"
- shows "b1 = b \<and> c1 = c"
-proof -
- have "x1 \<notin> fst ` toSet \<Gamma>"
- proof -
- have "atom x1 \<sharp> \<Gamma>" using assms wfG_cons by metis
- then show ?thesis
- using fresh_gamma_elem
- by (metis assms(2) atom_dom.simps dom.simps rev_image_eqI wfG_cons2 wfG_x_fresh)
- qed
- thus ?thesis using assms by force
-qed
-
-lemma wfG_member_unique:
- assumes "(x1,b1,c1) \<in> toSet (\<Gamma>'@((x,b,c) #\<^sub>\<Gamma>\<Gamma>))" and "wfG \<Theta> \<B> (\<Gamma>'@((x,b,c) #\<^sub>\<Gamma>\<Gamma>))" and "x = x1"
- shows "b1 = b \<and> c1 = c"
- using assms proof(induct \<Gamma>' rule: \<Gamma>_induct)
- case GNil
- then show ?case using wfG_suffix wfG_cons_unique append_g.simps by metis
-next
- case (GCons x' b' c' \<Gamma>')
- moreover hence "(x1, b1, c1) \<in> toSet (\<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>)" using wf_not_in_prefix by fastforce
- ultimately show ?case using wfG_cons by fastforce
-qed
-
-section \<open>Function Definitions\<close>
-
-lemma wb_phi_weakening:
- fixes \<Gamma>::\<Gamma> and \<Gamma>'::\<Gamma> and v::v and e::e and c::c and \<tau>::\<tau> and ts::"(string*\<tau>) list" and \<Delta>::\<Delta> and s::s and \<B>::\<B> and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def
- and cs::branch_s and css::branch_list and \<Phi>::\<Phi>
- shows
- "\<Theta>; \<Phi>; \<B>; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f e : b \<Longrightarrow> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>' \<Longrightarrow> set \<Phi> \<subseteq> set \<Phi>' \<Longrightarrow> \<Theta> ; \<Phi>' ; \<B> ; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f e : b" and
- "\<Theta>; \<Phi>; \<B>; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b \<Longrightarrow> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>' \<Longrightarrow> set \<Phi> \<subseteq> set \<Phi>' \<Longrightarrow> \<Theta> ; \<Phi>' ; \<B> ; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b" and
- "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> ; tid ; dc ; t \<turnstile>\<^sub>w\<^sub>f cs : b \<Longrightarrow> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>' \<Longrightarrow> set \<Phi> \<subseteq> set \<Phi>' \<Longrightarrow> \<Theta> ; \<Phi>' ; \<B> ; \<Gamma> ; \<Delta> ; tid ; dc ; t \<turnstile>\<^sub>w\<^sub>f cs : b" and
- "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> ; tid ; dclist \<turnstile>\<^sub>w\<^sub>f css : b \<Longrightarrow> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>' \<Longrightarrow> set \<Phi> \<subseteq> set \<Phi>' \<Longrightarrow> \<Theta> ; \<Phi>' ; \<B> ; \<Gamma> ; \<Delta> ; tid ; dclist \<turnstile>\<^sub>w\<^sub>f css : b" and
- "\<Theta> \<turnstile>\<^sub>w\<^sub>f (\<Phi>::\<Phi>) \<Longrightarrow> True" and
- "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<Longrightarrow> True" and
- "\<Theta> ; \<Phi> \<turnstile>\<^sub>w\<^sub>f ftq \<Longrightarrow> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>' \<Longrightarrow> set \<Phi> \<subseteq> set \<Phi>' \<Longrightarrow> \<Theta> ; \<Phi>' \<turnstile>\<^sub>w\<^sub>f ftq" and
- "\<Theta> ; \<Phi> ; \<B> \<turnstile>\<^sub>w\<^sub>f ft \<Longrightarrow> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>' \<Longrightarrow> set \<Phi> \<subseteq> set \<Phi>' \<Longrightarrow> \<Theta> ; \<Phi>' ; \<B> \<turnstile>\<^sub>w\<^sub>f ft"
-proof(nominal_induct
- b and b and b and b and \<Phi> and \<Delta> and ftq and ft
- avoiding: \<Phi>'
- rule:wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.strong_induct)
- case (wfE_valI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v b)
- then show ?case using wf_intros by metis
-next
- case (wfE_plusI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 v2)
- then show ?case using wf_intros by metis
-next
- case (wfE_leqI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 v2)
- then show ?case using wf_intros by metis
-next
- case (wfE_eqI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 b v2)
- then show ?case using wf_intros by metis
-next
- case (wfE_fstI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 b1 b2)
- then show ?case using wf_intros by metis
-next
- case (wfE_sndI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 b1 b2)
- then show ?case using wf_intros by metis
-next
- case (wfE_concatI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 v2)
- then show ?case using wf_intros by metis
-next
- case (wfE_splitI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 v2)
- then show ?case using wf_intros by metis
-next
- case (wfE_lenI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1)
- then show ?case using wf_intros by metis
-next
- case (wfE_appI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> f x b c \<tau> s v)
- then show ?case using wf_intros lookup_fun_weakening by metis
-next
- case (wfE_appPI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> b' bv v \<tau> f x b c s)
- show ?case proof
- show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>' \<close> using wfE_appPI by auto
- show \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wfE_appPI by auto
- show \<open> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b' \<close> using wfE_appPI by auto
- show \<open>atom bv \<sharp> (\<Phi>', \<Theta>, \<B>, \<Gamma>, \<Delta>, b', v, (b_of \<tau>)[bv::=b']\<^sub>b)\<close> using wfE_appPI by auto
- show \<open>Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \<tau> s))) = lookup_fun \<Phi>' f\<close>
- using wfE_appPI lookup_fun_weakening by metis
- show \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b[bv::=b']\<^sub>b \<close> using wfE_appPI by auto
- qed
-next
- case (wfE_mvarI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> u \<tau>)
- then show ?case using wf_intros by metis
-next
- case (wfS_valI \<Theta> \<Phi> \<B> \<Gamma> v b \<Delta>)
- then show ?case using wf_intros by metis
-next
- case (wfS_letI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> e b' x s b)
- then show ?case using Wellformed.wfS_letI by fastforce
-next
- case (wfS_let2I \<Theta> \<Phi> \<B> \<Gamma> \<Delta> s1 b' x s2 b)
- then show ?case using Wellformed.wfS_let2I by fastforce
-next
- case (wfS_ifI \<Theta> \<B> \<Gamma> v \<Phi> \<Delta> s1 b s2)
- then show ?case using wf_intros by metis
-next
- case (wfS_varI \<Theta> \<B> \<Gamma> \<tau> v u \<Phi> \<Delta> b s)
- show ?case proof
- show \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau> \<close> using wfS_varI by simp
- show \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b_of \<tau> \<close> using wfS_varI by simp
- show \<open>atom u \<sharp> (\<Phi>', \<Theta>, \<B>, \<Gamma>, \<Delta>, \<tau>, v, b)\<close> using wfS_varI by simp
- show \<open> \<Theta> ; \<Phi>' ; \<B> ; \<Gamma> ; (u, \<tau>) #\<^sub>\<Delta> \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b \<close> using wfS_varI by simp
- qed
-next
- case (wfS_assignI u \<tau> \<Delta> \<Theta> \<B> \<Gamma> \<Phi> v)
- then show ?case using wf_intros by metis
-next
- case (wfS_whileI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> s1 s2 b)
- then show ?case using wf_intros by metis
-next
- case (wfS_seqI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> s1 s2 b)
- then show ?case using wf_intros by metis
-next
- case (wfS_matchI \<Theta> \<B> \<Gamma> v tid dclist \<Delta> \<Phi> cs b)
- then show ?case using wf_intros by metis
-next
- case (wfS_branchI \<Theta> \<Phi> \<B> x \<tau> \<Gamma> \<Delta> s b tid dc)
- then show ?case using Wellformed.wfS_branchI by fastforce
-next
- case (wfS_assertI \<Theta> \<Phi> \<B> x c \<Gamma> \<Delta> s b)
- show ?case proof
- show \<open> \<Theta> ; \<Phi>' ; \<B> ; (x, B_bool, c) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b \<close> using wfS_assertI by auto
- next
- show \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f c \<close> using wfS_assertI by auto
- next
- show \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wfS_assertI by auto
- have "atom x \<sharp> \<Phi>'" using wfS_assertI wfPhi_supp fresh_def by blast
- thus \<open>atom x \<sharp> (\<Phi>', \<Theta>, \<B>, \<Gamma>, \<Delta>, c, b, s)\<close> using fresh_prodN wfS_assertI wfPhi_supp fresh_def by auto
- qed
-next
- case (wfFTI \<Theta> B b s x c \<tau> \<Phi>)
- show ?case proof
- show \<open> \<Theta> ; B \<turnstile>\<^sub>w\<^sub>f b \<close> using wfFTI by auto
- next
- show \<open>supp c \<subseteq> {atom x}\<close> using wfFTI by auto
- next
- show \<open> \<Theta> ; B ; (x, b, c) #\<^sub>\<Gamma> GNil \<turnstile>\<^sub>w\<^sub>f \<tau> \<close> using wfFTI by auto
- next
- show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>' \<close> using wfFTI by auto
- next
- show \<open>supp s \<subseteq> {atom x} \<union> supp B\<close> using wfFTI by auto
- qed
-qed(auto|metis wf_intros)+
-
-
-lemma wfT_fun_return_t:
- fixes \<tau>a'::\<tau> and \<tau>'::\<tau>
- assumes "\<Theta>; \<B>; (xa, b, ca) #\<^sub>\<Gamma> GNil \<turnstile>\<^sub>w\<^sub>f \<tau>a'" and "(AF_fun_typ x b c \<tau>' s') = (AF_fun_typ xa b ca \<tau>a' sa')"
- shows "\<Theta>; \<B>; (x, b, c) #\<^sub>\<Gamma> GNil \<turnstile>\<^sub>w\<^sub>f \<tau>'"
-proof -
- obtain cb::x where xf: "atom cb \<sharp> (c, \<tau>', s', sa', \<tau>a', ca, x , xa)" using obtain_fresh by blast
- hence "atom cb \<sharp> (c, \<tau>', s', sa', \<tau>a', ca) \<and> atom cb \<sharp> (x, xa, ((c, \<tau>'), s'), (ca, \<tau>a'), sa')" using fresh_prod6 fresh_prod4 fresh_prod8 by auto
- hence *:"c[x::=V_var cb]\<^sub>c\<^sub>v = ca[xa::=V_var cb]\<^sub>c\<^sub>v \<and> \<tau>'[x::=V_var cb]\<^sub>\<tau>\<^sub>v = \<tau>a'[xa::=V_var cb]\<^sub>\<tau>\<^sub>v" using assms \<tau>.eq_iff Abs1_eq_iff_all by auto
-
- have **: "\<Theta>; \<B>; (xa \<leftrightarrow> cb ) \<bullet> ((xa, b, ca) #\<^sub>\<Gamma> GNil) \<turnstile>\<^sub>w\<^sub>f (xa \<leftrightarrow> cb ) \<bullet> \<tau>a'" using assms True_eqvt beta_flip_eq theta_flip_eq wfG_wf
- by (metis GCons_eqvt GNil_eqvt wfT.eqvt wfT_wf)
-
- have "\<Theta>; \<B>; (x \<leftrightarrow> cb ) \<bullet> ((x, b, c) #\<^sub>\<Gamma> GNil) \<turnstile>\<^sub>w\<^sub>f (x \<leftrightarrow> cb ) \<bullet> \<tau>'" proof -
- have "(xa \<leftrightarrow> cb ) \<bullet> xa = (x \<leftrightarrow> cb ) \<bullet> x" using xf by auto
- hence "(x \<leftrightarrow> cb ) \<bullet> ((x, b, c) #\<^sub>\<Gamma> GNil) = (xa \<leftrightarrow> cb ) \<bullet> ((xa, b, ca) #\<^sub>\<Gamma> GNil)" using * ** xf G_cons_flip fresh_GNil by simp
- thus ?thesis using ** * xf by simp
- qed
- thus ?thesis using beta_flip_eq theta_flip_eq wfT_wf wfG_wf * ** True_eqvt wfT.eqvt permute_flip_cancel by metis
-qed
-
-lemma wfFT_wf_aux:
- fixes \<tau>::\<tau> and \<Theta>::\<Theta> and \<Phi>::\<Phi> and ft :: fun_typ_q and s::s and \<Delta>::\<Delta>
- assumes "\<Theta> ; \<Phi> ; B \<turnstile>\<^sub>w\<^sub>f (AF_fun_typ x b c \<tau> s)"
- shows "\<Theta> ; B ; (x,b,c) #\<^sub>\<Gamma> GNil \<turnstile>\<^sub>w\<^sub>f \<tau> \<and> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<and> supp s \<subseteq> { atom x } \<union> supp B"
-proof -
- obtain xa and ca and sa and \<tau>' where *:"\<Theta> ; B \<turnstile>\<^sub>w\<^sub>f b \<and> (\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> ) \<and>
- supp sa \<subseteq> {atom xa} \<union> supp B \<and> (\<Theta> ; B ; (xa, b, ca) #\<^sub>\<Gamma> GNil \<turnstile>\<^sub>w\<^sub>f \<tau>') \<and>
- AF_fun_typ x b c \<tau> s = AF_fun_typ xa b ca \<tau>' sa "
- using wfFT.simps[of \<Theta> \<Phi> B "AF_fun_typ x b c \<tau> s"] assms by auto
-
- moreover hence **: "(AF_fun_typ x b c \<tau> s) = (AF_fun_typ xa b ca \<tau>' sa)" by simp
- ultimately have "\<Theta> ; B ; (x,b,c) #\<^sub>\<Gamma>GNil \<turnstile>\<^sub>w\<^sub>f \<tau>" using wfT_fun_return_t by metis
- moreover have " (\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> ) " using * by auto
- moreover have "supp s \<subseteq> { atom x } \<union> supp B" proof -
- have "[[atom x]]lst.s = [[atom xa]]lst.sa" using ** fun_typ.eq_iff lst_fst lst_snd by metis
- thus ?thesis using lst_supp_subset * by metis
- qed
- ultimately show ?thesis by auto
-qed
-
-lemma wfFT_simple_wf:
- fixes \<tau>::\<tau> and \<Theta>::\<Theta> and \<Phi>::\<Phi> and ft :: fun_typ_q and s::s and \<Delta>::\<Delta>
- assumes "\<Theta> ; \<Phi> \<turnstile>\<^sub>w\<^sub>f (AF_fun_typ_none (AF_fun_typ x b c \<tau> s))"
- shows "\<Theta> ; {||} ; (x,b,c) #\<^sub>\<Gamma>GNil \<turnstile>\<^sub>w\<^sub>f \<tau> \<and> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<and> supp s \<subseteq> { atom x } "
-proof -
- have *:"\<Theta> ; \<Phi> ; {||} \<turnstile>\<^sub>w\<^sub>f (AF_fun_typ x b c \<tau> s)" using wfFTQ_elims assms by auto
- thus ?thesis using wfFT_wf_aux by force
-qed
-
-lemma wfFT_poly_wf:
- fixes \<tau>::\<tau> and \<Theta>::\<Theta> and \<Phi>::\<Phi> and ftq :: fun_typ_q and s::s and \<Delta>::\<Delta>
- assumes "\<Theta> ; \<Phi> \<turnstile>\<^sub>w\<^sub>f (AF_fun_typ_some bv (AF_fun_typ x b c \<tau> s))"
- shows "\<Theta> ; {|bv|} ; (x,b,c) #\<^sub>\<Gamma>GNil \<turnstile>\<^sub>w\<^sub>f \<tau> \<and> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<and> \<Theta> ; \<Phi> ; {|bv|} \<turnstile>\<^sub>w\<^sub>f (AF_fun_typ x b c \<tau> s)"
-proof -
- obtain bv1 ft1 where *:"\<Theta> ; \<Phi> ; {|bv1|} \<turnstile>\<^sub>w\<^sub>f ft1 \<and> [[atom bv1]]lst. ft1 = [[atom bv]]lst. AF_fun_typ x b c \<tau> s"
- using wfFTQ_elims(3)[OF assms] by metis
-
- show ?thesis proof(cases "bv1 = bv")
- case True
- then show ?thesis using * fun_typ_q.eq_iff Abs1_eq_iff by (metis (no_types, opaque_lifting) wfFT_wf_aux)
- next
- case False
- obtain x1 b1 c1 t1 s1 where **: "ft1 = AF_fun_typ x1 b1 c1 t1 s1" using fun_typ.eq_iff
- by (meson fun_typ.exhaust)
-
- hence eqv: "(bv \<leftrightarrow> bv1) \<bullet> AF_fun_typ x1 b1 c1 t1 s1 = AF_fun_typ x b c \<tau> s \<and> atom bv1 \<sharp> AF_fun_typ x b c \<tau> s" using
- Abs1_eq_iff(3) * False by metis
-
- have "(bv \<leftrightarrow> bv1) \<bullet> \<Theta> ; (bv \<leftrightarrow> bv1) \<bullet> \<Phi> ; (bv \<leftrightarrow> bv1) \<bullet> {|bv1|} \<turnstile>\<^sub>w\<^sub>f (bv \<leftrightarrow> bv1) \<bullet> ft1" using wfFT.eqvt * by metis
- moreover have "(bv \<leftrightarrow> bv1) \<bullet> \<Phi> = \<Phi>" using phi_flip_eq wfX_wfY * by metis
- moreover have "(bv \<leftrightarrow> bv1) \<bullet> \<Theta> =\<Theta>" using wfX_wfY * theta_flip_eq2 by metis
- moreover have "(bv \<leftrightarrow> bv1) \<bullet> ft1 = AF_fun_typ x b c \<tau> s" using eqv ** by metis
- ultimately have "\<Theta> ; \<Phi> ; {|bv|} \<turnstile>\<^sub>w\<^sub>f AF_fun_typ x b c \<tau> s" by auto
- thus ?thesis using wfFT_wf_aux by auto
- qed
-qed
-
-lemma wfFT_poly_wfT:
- fixes \<tau>::\<tau> and \<Theta>::\<Theta> and \<Phi>::\<Phi> and ft :: fun_typ_q
- assumes "\<Theta> ; \<Phi> \<turnstile>\<^sub>w\<^sub>f (AF_fun_typ_some bv (AF_fun_typ x b c \<tau> s))"
- shows "\<Theta> ; {| bv |} ; (x,b,c) #\<^sub>\<Gamma>GNil \<turnstile>\<^sub>w\<^sub>f \<tau>"
- using wfFT_poly_wf assms by simp
-
-lemma b_of_supp:
- "supp (b_of t) \<subseteq> supp t"
-proof(nominal_induct t rule:\<tau>.strong_induct)
- case (T_refined_type x b c)
- then show ?case by auto
-qed
-
-lemma wfPhi_f_simple_wf:
- fixes \<tau>::\<tau> and \<Theta>::\<Theta> and \<Phi>::\<Phi> and ft :: fun_typ_q and s::s and \<Phi>'::\<Phi>
- assumes "AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \<tau> s)) \<in> set \<Phi> " and "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>" and "set \<Phi> \<subseteq> set \<Phi>'" and "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>'"
- shows "\<Theta> ; {||} ; (x,b,c) #\<^sub>\<Gamma> GNil \<turnstile>\<^sub>w\<^sub>f \<tau> \<and> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<and> supp s \<subseteq> { atom x }"
-using assms proof(induct \<Phi> rule: \<Phi>_induct)
- case PNil
- then show ?case by auto
-next
- case (PConsSome f1 bv x1 b1 c1 \<tau>1 s' \<Phi>'')
- hence "AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \<tau> s)) \<in> set \<Phi>''" by auto
- moreover have " \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>'' \<and> set \<Phi>'' \<subseteq> set \<Phi>'" using wfPhi_elims(3) PConsSome by auto
- ultimately show ?case using PConsSome wfPhi_elims wfFT_simple_wf by auto
-next
- case (PConsNone f' x' b' c' \<tau>' s' \<Phi>'')
- show ?case proof(cases "f=f'")
- case True
- have "AF_fun_typ_none (AF_fun_typ x' b' c' \<tau>' s') = AF_fun_typ_none (AF_fun_typ x b c \<tau> s)"
- by (metis PConsNone.prems(1) PConsNone.prems(2) True fun_def.eq_iff image_eqI name_of_fun.simps set_ConsD wfPhi_elims(2))
- hence *:"\<Theta> ; \<Phi>'' \<turnstile>\<^sub>w\<^sub>f AF_fun_typ_none (AF_fun_typ x b c \<tau> s) " using wfPhi_elims(2)[OF PConsNone(3)] by metis
- hence "\<Theta> ; \<Phi>'' ; {||} \<turnstile>\<^sub>w\<^sub>f (AF_fun_typ x b c \<tau> s)" using wfFTQ_elims(1) by metis
- thus ?thesis using wfFT_simple_wf[OF *] wb_phi_weakening PConsNone by force
- next
- case False
- hence "AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \<tau> s)) \<in> set \<Phi>''" using PConsNone by simp
- moreover have " \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>'' \<and> set \<Phi>'' \<subseteq> set \<Phi>'" using wfPhi_elims(3) PConsNone by auto
- ultimately show ?thesis using PConsNone wfPhi_elims wfFT_simple_wf by auto
- qed
-qed
-
-lemma wfPhi_f_simple_wfT:
- fixes \<tau>::\<tau> and \<Theta>::\<Theta> and \<Phi>::\<Phi> and ft :: fun_typ_q
- assumes "Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \<tau> s))) = lookup_fun \<Phi> f" and "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>"
- shows "\<Theta> ; {||} ; (x,b,c) #\<^sub>\<Gamma>GNil \<turnstile>\<^sub>w\<^sub>f \<tau>"
- using wfPhi_f_simple_wf assms using lookup_fun_member by blast
-
-lemma wfPhi_f_simple_supp_b:
- fixes \<tau>::\<tau> and \<Theta>::\<Theta> and \<Phi>::\<Phi> and ft :: fun_typ_q
- assumes "Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \<tau> s))) = lookup_fun \<Phi> f" and "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>"
- shows "supp b = {}"
-proof -
- have "\<Theta> ; {||} ; (x,b,c) #\<^sub>\<Gamma>GNil \<turnstile>\<^sub>w\<^sub>f \<tau>" using wfPhi_f_simple_wfT assms by auto
- thus ?thesis using wfT_wf wfG_cons wfB_supp by fastforce
-qed
-
-lemma wfPhi_f_simple_supp_t:
- fixes \<tau>::\<tau> and \<Theta>::\<Theta> and \<Phi>::\<Phi> and ft :: fun_typ_q
- assumes "Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \<tau> s))) = lookup_fun \<Phi> f" and "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>"
- shows "supp \<tau> \<subseteq> { atom x }"
- using wfPhi_f_simple_wfT wfT_supp assms by fastforce
-
-lemma wfPhi_f_simple_supp_c:
- fixes \<tau>::\<tau> and \<Theta>::\<Theta> and \<Phi>::\<Phi> and ft :: fun_typ_q
- assumes "Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \<tau> s))) = lookup_fun \<Phi> f" and "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>"
- shows "supp c \<subseteq> { atom x }"
-proof -
- have "\<Theta> ; {||} ; (x,b,c) #\<^sub>\<Gamma>GNil \<turnstile>\<^sub>w\<^sub>f \<tau>" using wfPhi_f_simple_wfT assms by auto
- thus ?thesis using wfG_wfC wfC_supp wfT_wf by fastforce
-qed
-
-lemma wfPhi_f_simple_supp_s:
- fixes \<tau>::\<tau> and \<Theta>::\<Theta> and \<Phi>::\<Phi> and ft :: fun_typ_q
- assumes "Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \<tau> s))) = lookup_fun \<Phi> f" and "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>"
- shows "supp s \<subseteq> {atom x}"
-proof -
- have "AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \<tau> s)) \<in> set \<Phi>" using lookup_fun_member assms by auto
- hence "supp s \<subseteq> { atom x }" using wfPhi_f_simple_wf assms by blast
- thus ?thesis using wf_supp(3) atom_dom.simps toSet.simps x_not_in_u_set x_not_in_b_set setD.simps
- using wf_supp2(2) by fastforce
-qed
-
-lemma wfPhi_f_poly_wf:
- fixes \<tau>::\<tau> and \<Theta>::\<Theta> and \<Phi>::\<Phi> and ft :: fun_typ_q and s::s and \<Phi>'::\<Phi>
- assumes "AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \<tau> s)) \<in> set \<Phi> " and "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>" and "set \<Phi> \<subseteq> set \<Phi>'" and "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>'"
- shows "\<Theta> ; {|bv|} ; (x,b,c) #\<^sub>\<Gamma>GNil \<turnstile>\<^sub>w\<^sub>f \<tau> \<and> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>' \<and> \<Theta> ; \<Phi>' ; {|bv|} \<turnstile>\<^sub>w\<^sub>f (AF_fun_typ x b c \<tau> s)"
-using assms proof(induct \<Phi> rule: \<Phi>_induct)
- case PNil
- then show ?case by auto
-next
- case (PConsNone f x b c \<tau> s' \<Phi>'')
- moreover have " \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>'' \<and> set \<Phi>'' \<subseteq> set \<Phi>'" using wfPhi_elims(3) PConsNone by auto
- ultimately show ?case using PConsNone wfPhi_elims wfFT_poly_wf by auto
-next
- case (PConsSome f1 bv1 x1 b1 c1 \<tau>1 s1 \<Phi>'')
- show ?case proof(cases "f=f1")
- case True
- have "AF_fun_typ_some bv1 (AF_fun_typ x1 b1 c1 \<tau>1 s1) = AF_fun_typ_some bv (AF_fun_typ x b c \<tau> s)"
- by (metis PConsSome.prems(1) PConsSome.prems(2) True fun_def.eq_iff list.set_intros(1) option.inject wfPhi_lookup_fun_unique)
- hence *:"\<Theta> ; \<Phi>'' \<turnstile>\<^sub>w\<^sub>f AF_fun_typ_some bv (AF_fun_typ x b c \<tau> s) " using wfPhi_elims PConsSome by metis
- thus ?thesis using wfFT_poly_wf * wb_phi_weakening PConsSome
- by (meson set_subset_Cons)
- next
- case False
- hence "AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \<tau> s)) \<in> set \<Phi>''" using PConsSome
- by (meson fun_def.eq_iff set_ConsD)
- moreover have " \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>'' \<and> set \<Phi>'' \<subseteq> set \<Phi>'" using wfPhi_elims(3) PConsSome
- by (meson dual_order.trans set_subset_Cons)
- ultimately show ?thesis using PConsSome wfPhi_elims wfFT_poly_wf
- by blast
- qed
-qed
-
-lemma wfPhi_f_poly_wfT:
- fixes \<tau>::\<tau> and \<Theta>::\<Theta> and \<Phi>::\<Phi> and ft :: fun_typ_q
- assumes "Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \<tau> s))) = lookup_fun \<Phi> f" and "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>"
- shows "\<Theta> ; {| bv |} ; (x,b,c) #\<^sub>\<Gamma>GNil \<turnstile>\<^sub>w\<^sub>f \<tau>"
-using assms proof(induct \<Phi> rule: \<Phi>_induct)
- case PNil
- then show ?case by auto
-next
- case (PConsSome f1 bv1 x1 b1 c1 \<tau>1 s' \<Phi>')
- then show ?case proof(cases "f1=f")
- case True
- hence "lookup_fun (AF_fundef f1 (AF_fun_typ_some bv1 (AF_fun_typ x1 b1 c1 \<tau>1 s')) # \<Phi>') f = Some (AF_fundef f1 (AF_fun_typ_some bv1 (AF_fun_typ x1 b1 c1 \<tau>1 s')))" using
- lookup_fun.simps using PConsSome.prems by simp
- then show ?thesis using PConsSome.prems wfPhi_elims wfFT_poly_wfT
- by (metis option.inject)
- next
- case False
- then show ?thesis using PConsSome using lookup_fun.simps
- using wfPhi_elims(3) by auto
- qed
-next
- case (PConsNone f' x' b' c' \<tau>' s' \<Phi>')
- then show ?case proof(cases "f'=f")
- case True
- then have *:"\<Theta> ; \<Phi>' \<turnstile>\<^sub>w\<^sub>f AF_fun_typ_none (AF_fun_typ x' b' c' \<tau>' s') " using lookup_fun.simps PConsNone wfPhi_elims by metis
- thus ?thesis using PConsNone wfFT_poly_wfT wfPhi_elims lookup_fun.simps
- by (metis fun_def.eq_iff fun_typ_q.distinct(1) option.inject)
- next
- case False
- thus ?thesis using PConsNone wfPhi_elims
- by (metis False lookup_fun.simps(2))
- qed
-qed
-
-lemma wfPhi_f_poly_supp_b:
- fixes \<tau>::\<tau> and \<Theta>::\<Theta> and \<Phi>::\<Phi> and ft :: fun_typ_q
- assumes "Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \<tau> s))) = lookup_fun \<Phi> f" and "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>"
- shows "supp b \<subseteq> supp bv"
-proof -
- have "\<Theta> ; {|bv|} ; (x,b,c) #\<^sub>\<Gamma>GNil \<turnstile>\<^sub>w\<^sub>f \<tau>" using wfPhi_f_poly_wfT assms by auto
- thus ?thesis using wfT_wf wfG_cons wfB_supp by fastforce
-qed
-
-lemma wfPhi_f_poly_supp_t:
- fixes \<tau>::\<tau> and \<Theta>::\<Theta> and \<Phi>::\<Phi> and ft :: fun_typ_q
- assumes "Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \<tau> s))) = lookup_fun \<Phi> f" and "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>"
- shows "supp \<tau> \<subseteq> { atom x , atom bv }"
- using wfPhi_f_poly_wfT[OF assms, THEN wfT_supp] atom_dom.simps supp_at_base by auto
-
-
-lemma wfPhi_f_poly_supp_b_of_t:
- fixes \<tau>::\<tau> and \<Theta>::\<Theta> and \<Phi>::\<Phi> and ft :: fun_typ_q
- assumes "Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \<tau> s))) = lookup_fun \<Phi> f" and "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>"
- shows "supp (b_of \<tau>) \<subseteq> { atom bv }"
-proof -
- have "atom x \<notin> supp (b_of \<tau>)" using x_fresh_b by auto
- moreover have "supp (b_of \<tau>) \<subseteq> { atom x , atom bv }" using wfPhi_f_poly_supp_t
- using supp_at_base b_of.simps wfPhi_f_poly_supp_t \<tau>.supp b_of_supp assms by fast
- ultimately show ?thesis by blast
-qed
-
-lemma wfPhi_f_poly_supp_c:
- fixes \<tau>::\<tau> and \<Theta>::\<Theta> and \<Phi>::\<Phi> and ft :: fun_typ_q
- assumes "Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \<tau> s))) = lookup_fun \<Phi> f" and "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>"
- shows "supp c \<subseteq> { atom x, atom bv }"
-proof -
- have "\<Theta> ; {|bv|} ; (x,b,c) #\<^sub>\<Gamma>GNil \<turnstile>\<^sub>w\<^sub>f \<tau>" using wfPhi_f_poly_wfT assms by auto
- thus ?thesis using wfG_wfC wfC_supp wfT_wf
- using supp_at_base by fastforce
-qed
-
-lemma wfPhi_f_poly_supp_s:
- fixes \<tau>::\<tau> and \<Theta>::\<Theta> and \<Phi>::\<Phi> and ft :: fun_typ_q
- assumes "Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \<tau> s))) = lookup_fun \<Phi> f" and "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>"
- shows "supp s \<subseteq> {atom x, atom bv}"
-proof -
-
- have "AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \<tau> s)) \<in> set \<Phi>" using lookup_fun_member assms by auto
- hence *:"\<Theta> ; \<Phi> ; {|bv|} \<turnstile>\<^sub>w\<^sub>f (AF_fun_typ x b c \<tau> s)" using assms wfPhi_f_poly_wf by simp
-
- thus ?thesis using wfFT_wf_aux[OF *] using supp_at_base by auto
-qed
-
-lemmas wfPhi_f_supp = wfPhi_f_poly_supp_b wfPhi_f_simple_supp_b wfPhi_f_poly_supp_c
- wfPhi_f_simple_supp_t wfPhi_f_poly_supp_t wfPhi_f_simple_supp_t wfPhi_f_poly_wfT wfPhi_f_simple_wfT
- wfPhi_f_poly_supp_s wfPhi_f_simple_supp_s
-
-lemma fun_typ_eq_ret_unique:
- assumes "(AF_fun_typ x1 b1 c1 \<tau>1' s1') = (AF_fun_typ x2 b2 c2 \<tau>2' s2')"
- shows "\<tau>1'[x1::=v]\<^sub>\<tau>\<^sub>v = \<tau>2'[x2::=v]\<^sub>\<tau>\<^sub>v"
-proof -
- have "[[atom x1]]lst. \<tau>1' = [[atom x2]]lst. \<tau>2'" using assms lst_fst fun_typ.eq_iff lst_snd by metis
- thus ?thesis using subst_v_flip_eq_two[of x1 \<tau>1' x2 \<tau>2' v] subst_v_\<tau>_def by metis
-qed
-
-lemma fun_typ_eq_body_unique:
- fixes v::v and x1::x and x2::x and s1'::s and s2'::s
- assumes "(AF_fun_typ x1 b1 c1 \<tau>1' s1') = (AF_fun_typ x2 b2 c2 \<tau>2' s2')"
- shows "s1'[x1::=v]\<^sub>s\<^sub>v = s2'[x2::=v]\<^sub>s\<^sub>v"
-proof -
- have "[[atom x1]]lst. s1' = [[atom x2]]lst. s2'" using assms lst_fst fun_typ.eq_iff lst_snd by metis
- thus ?thesis using subst_v_flip_eq_two[of x1 s1' x2 s2' v] subst_v_s_def by metis
-qed
-
-lemma fun_ret_unique:
- assumes "Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x1 b1 c1 \<tau>1' s1'))) = lookup_fun \<Phi> f" and "Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x2 b2 c2 \<tau>2' s2'))) = lookup_fun \<Phi> f"
- shows "\<tau>1'[x1::=v]\<^sub>\<tau>\<^sub>v = \<tau>2'[x2::=v]\<^sub>\<tau>\<^sub>v"
-proof -
- have *: " (AF_fundef f (AF_fun_typ_none (AF_fun_typ x1 b1 c1 \<tau>1' s1'))) = (AF_fundef f (AF_fun_typ_none (AF_fun_typ x2 b2 c2 \<tau>2' s2')))" using option.inject assms by metis
- thus ?thesis using fun_typ_eq_ret_unique fun_def.eq_iff fun_typ_q.eq_iff by metis
-qed
-
-lemma fun_poly_arg_unique:
- fixes bv1::bv and bv2::bv and b::b and \<tau>1::\<tau> and \<tau>2::\<tau>
- assumes "[[atom bv1]]lst. (AF_fun_typ x1 b1 c1 \<tau>1 s1) = [[atom bv2]]lst. (AF_fun_typ x2 b2 c2 \<tau>2 s2)" (is "[[atom ?x]]lst. ?a = [[atom ?y]]lst. ?b")
- shows "\<lbrace> x1 : b1[bv1::=b]\<^sub>b\<^sub>b | c1[bv1::=b]\<^sub>c\<^sub>b \<rbrace> = \<lbrace> x2 : b2[bv2::=b]\<^sub>b\<^sub>b | c2[bv2::=b]\<^sub>c\<^sub>b \<rbrace>"
-proof -
- obtain c::bv where *:"atom c \<sharp> (b,b1,b2,c1,c2) \<and> atom c \<sharp> (bv1, bv2, AF_fun_typ x1 b1 c1 \<tau>1 s1, AF_fun_typ x2 b2 c2 \<tau>2 s2)" using obtain_fresh fresh_Pair by metis
- hence "(bv1 \<leftrightarrow> c) \<bullet> AF_fun_typ x1 b1 c1 \<tau>1 s1 = (bv2 \<leftrightarrow> c) \<bullet> AF_fun_typ x2 b2 c2 \<tau>2 s2" using Abs1_eq_iff_all(3)[of ?x ?a ?y ?b] assms by metis
- hence "AF_fun_typ x1 ((bv1 \<leftrightarrow> c) \<bullet> b1) ((bv1 \<leftrightarrow> c) \<bullet> c1) ((bv1 \<leftrightarrow> c) \<bullet> \<tau>1) ((bv1 \<leftrightarrow> c) \<bullet> s1) = AF_fun_typ x2 ((bv2 \<leftrightarrow> c) \<bullet> b2) ((bv2 \<leftrightarrow> c) \<bullet> c2) ((bv2 \<leftrightarrow> c) \<bullet> \<tau>2) ((bv2 \<leftrightarrow> c) \<bullet> s2)"
- using fun_typ_flip by metis
- hence **:"\<lbrace> x1 :((bv1 \<leftrightarrow> c) \<bullet> b1) | ((bv1 \<leftrightarrow> c) \<bullet> c1) \<rbrace> = \<lbrace> x2 : ((bv2 \<leftrightarrow> c) \<bullet> b2) | ((bv2 \<leftrightarrow> c) \<bullet> c2) \<rbrace>" (is "\<lbrace> x1 : ?b1 | ?c1 \<rbrace> = \<lbrace> x2 : ?b2 | ?c2 \<rbrace>") using fun_arg_unique_aux by metis
- hence "\<lbrace> x1 :((bv1 \<leftrightarrow> c) \<bullet> b1) | ((bv1 \<leftrightarrow> c) \<bullet> c1) \<rbrace>[c::=b]\<^sub>\<tau>\<^sub>b = \<lbrace> x2 : ((bv2 \<leftrightarrow> c) \<bullet> b2) | ((bv2 \<leftrightarrow> c) \<bullet> c2) \<rbrace>[c::=b]\<^sub>\<tau>\<^sub>b" by metis
- hence "\<lbrace> x1 :((bv1 \<leftrightarrow> c) \<bullet> b1)[c::=b]\<^sub>b\<^sub>b | ((bv1 \<leftrightarrow> c) \<bullet> c1)[c::=b]\<^sub>c\<^sub>b \<rbrace> = \<lbrace> x2 : ((bv2 \<leftrightarrow> c) \<bullet> b2)[c::=b]\<^sub>b\<^sub>b | ((bv2 \<leftrightarrow> c) \<bullet> c2)[c::=b]\<^sub>c\<^sub>b \<rbrace>" using subst_tb.simps by metis
- thus ?thesis using * flip_subst_subst subst_b_c_def subst_b_b_def fresh_prodN flip_commute by metis
-qed
-
-lemma fun_poly_ret_unique:
- assumes "Some (AF_fundef f (AF_fun_typ_some bv1 (AF_fun_typ x1 b1 c1 \<tau>1' s1'))) = lookup_fun \<Phi> f" and "Some (AF_fundef f (AF_fun_typ_some bv2 (AF_fun_typ x2 b2 c2 \<tau>2' s2'))) = lookup_fun \<Phi> f"
- shows "\<tau>1'[bv1::=b]\<^sub>\<tau>\<^sub>b[x1::=v]\<^sub>\<tau>\<^sub>v = \<tau>2'[bv2::=b]\<^sub>\<tau>\<^sub>b[x2::=v]\<^sub>\<tau>\<^sub>v"
-proof -
- have *: " (AF_fundef f (AF_fun_typ_some bv1 (AF_fun_typ x1 b1 c1 \<tau>1' s1'))) = (AF_fundef f (AF_fun_typ_some bv2 (AF_fun_typ x2 b2 c2 \<tau>2' s2')))" using option.inject assms by metis
- hence "AF_fun_typ_some bv1 (AF_fun_typ x1 b1 c1 \<tau>1' s1') = AF_fun_typ_some bv2 (AF_fun_typ x2 b2 c2 \<tau>2' s2')"
- (is "AF_fun_typ_some bv1 ?ft1 = AF_fun_typ_some bv2 ?ft2") using fun_def.eq_iff by metis
- hence **:"[[atom bv1]]lst. ?ft1 = [[atom bv2]]lst. ?ft2" using fun_typ_q.eq_iff(1) by metis
-
- hence *:"subst_ft_b ?ft1 bv1 b = subst_ft_b ?ft2 bv2 b" using subst_b_flip_eq_two subst_b_fun_typ_def by metis
- have "[[atom x1]]lst. \<tau>1'[bv1::=b]\<^sub>\<tau>\<^sub>b = [[atom x2]]lst. \<tau>2'[bv2::=b]\<^sub>\<tau>\<^sub>b"
- apply(rule lst_snd[of _ "c1[bv1::=b]\<^sub>c\<^sub>b" _ _ "c2[bv2::=b]\<^sub>c\<^sub>b"])
- apply(rule lst_fst[of _ _ "s1'[bv1::=b]\<^sub>s\<^sub>b" _ _ "s2'[bv2::=b]\<^sub>s\<^sub>b"])
- using * subst_ft_b.simps fun_typ.eq_iff by metis
- thus ?thesis using subst_v_flip_eq_two subst_v_\<tau>_def by metis
-qed
-
-lemma fun_poly_body_unique:
- assumes "Some (AF_fundef f (AF_fun_typ_some bv1 (AF_fun_typ x1 b1 c1 \<tau>1' s1'))) = lookup_fun \<Phi> f" and "Some (AF_fundef f (AF_fun_typ_some bv2 (AF_fun_typ x2 b2 c2 \<tau>2' s2'))) = lookup_fun \<Phi> f"
- shows "s1'[bv1::=b]\<^sub>s\<^sub>b[x1::=v]\<^sub>s\<^sub>v = s2'[bv2::=b]\<^sub>s\<^sub>b[x2::=v]\<^sub>s\<^sub>v"
-proof -
- have *: " (AF_fundef f (AF_fun_typ_some bv1 (AF_fun_typ x1 b1 c1 \<tau>1' s1'))) = (AF_fundef f (AF_fun_typ_some bv2 (AF_fun_typ x2 b2 c2 \<tau>2' s2')))"
- using option.inject assms by metis
- hence "AF_fun_typ_some bv1 (AF_fun_typ x1 b1 c1 \<tau>1' s1') = AF_fun_typ_some bv2 (AF_fun_typ x2 b2 c2 \<tau>2' s2')"
- (is "AF_fun_typ_some bv1 ?ft1 = AF_fun_typ_some bv2 ?ft2") using fun_def.eq_iff by metis
- hence **:"[[atom bv1]]lst. ?ft1 = [[atom bv2]]lst. ?ft2" using fun_typ_q.eq_iff(1) by metis
-
- hence *:"subst_ft_b ?ft1 bv1 b = subst_ft_b ?ft2 bv2 b" using subst_b_flip_eq_two subst_b_fun_typ_def by metis
- have "[[atom x1]]lst. s1'[bv1::=b]\<^sub>s\<^sub>b = [[atom x2]]lst. s2'[bv2::=b]\<^sub>s\<^sub>b"
- using lst_snd lst_fst subst_ft_b.simps fun_typ.eq_iff
- by (metis "local.*")
-
- thus ?thesis using subst_v_flip_eq_two subst_v_s_def by metis
-qed
-
-lemma funtyp_eq_iff_equalities:
- fixes s'::s and s::s
- assumes " [[atom x']]lst. ((c', \<tau>'), s') = [[atom x]]lst. ((c, \<tau>), s)"
- shows "\<lbrace> x' : b | c' \<rbrace> = \<lbrace> x : b | c \<rbrace> \<and> s'[x'::=v]\<^sub>s\<^sub>v = s[x::=v]\<^sub>s\<^sub>v \<and> \<tau>'[x'::=v]\<^sub>\<tau>\<^sub>v = \<tau>[x::=v]\<^sub>\<tau>\<^sub>v"
-proof -
- have "[[atom x']]lst. s' = [[atom x]]lst. s" and "[[atom x']]lst. \<tau>' = [[atom x]]lst. \<tau>" and
- " [[atom x']]lst. c' = [[atom x]]lst. c" using lst_snd lst_fst assms by metis+
- thus ?thesis using subst_v_flip_eq_two \<tau>.eq_iff
- by (metis assms fun_typ.eq_iff fun_typ_eq_body_unique fun_typ_eq_ret_unique)
-qed
-
-section \<open>Weakening\<close>
-
-lemma wfX_wfB1:
- fixes \<Gamma>::\<Gamma> and \<Gamma>'::\<Gamma> and v::v and e::e and c::c and \<tau>::\<tau> and ts::"(string*\<tau>) list" and \<Delta>::\<Delta> and s::s and b::b and \<B>::\<B> and \<Phi>::\<Phi> and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def
- and cs::branch_s and css::branch_list
- shows wfV_wfB: "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b" and
- "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f c \<Longrightarrow> True" and
- "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<Longrightarrow> True" and
- wfT_wfB: "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau> \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b_of \<tau> " and
- "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ts \<Longrightarrow> True" and
- "\<turnstile>\<^sub>w\<^sub>f \<Theta> \<Longrightarrow> True" and
- "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b \<Longrightarrow> True" and
- wfCE_wfB: "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ce : b \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b" and
- "\<Theta> \<turnstile>\<^sub>w\<^sub>f td \<Longrightarrow> True"
-proof(induct rule:wfV_wfC_wfG_wfT_wfTs_wfTh_wfB_wfCE_wfTD.inducts)
- case (wfV_varI \<Theta> \<B> \<Gamma> b c x)
- hence "(x,b,c) \<in> toSet \<Gamma>" using lookup_iff wfV_wf using lookup_in_g by presburger
- hence "b \<in> fst`snd`toSet \<Gamma>" by force
- hence "wfB \<Theta> \<B> b" using wfG_wfB wfV_varI by metis
- then show ?case using wfV_elims wfG_wf wf_intros by metis
-next
- case (wfV_litI \<Theta> \<Gamma> l)
- moreover have "wfTh \<Theta>" using wfV_wf wfG_wf wfV_litI by metis
- ultimately show ?case using wfV_wf wfG_wf wf_intros base_for_lit.simps l.exhaust by metis
-next
- case (wfV_pairI \<Theta> \<Gamma> v1 b1 v2 b2)
- then show ?case using wfG_wf wf_intros by metis
-next
- case (wfV_consI s dclist \<Theta> dc x b c B \<Gamma> v)
- then show ?case
- using wfV_wf wfG_wf wfB_consI by metis
-next
- case (wfV_conspI s bv dclist \<Theta> dc x b' c \<B> b \<Gamma> v)
- then show ?case
- using wfV_wf wfG_wf using wfB_appI by metis
-next
- case (wfCE_valI \<Theta> \<B> \<Gamma> v b)
- then show ?case using wfB_elims by auto
-next
- case (wfCE_plusI \<Theta> \<B> \<Gamma> v1 v2)
- then show ?case using wfB_elims by auto
-next
- case (wfCE_leqI \<Theta> \<B> \<Gamma> v1 v2)
- then show ?case using wfV_wf wfG_wf wf_intros wfX_wfY by metis
-next
- case (wfCE_eqI \<Theta> \<B> \<Gamma> v1 b v2)
- then show ?case using wfV_wf wfG_wf wf_intros wfX_wfY by metis
-next
- case (wfCE_fstI \<Theta> \<B> \<Gamma> v1 b1 b2)
- then show ?case using wfB_elims by metis
-next
- case (wfCE_sndI \<Theta> \<B> \<Gamma> v1 b1 b2)
- then show ?case using wfB_elims by metis
-next
- case (wfCE_concatI \<Theta> \<B> \<Gamma> v1 v2)
- then show ?case using wfB_elims by auto
-next
- case (wfCE_lenI \<Theta> \<B> \<Gamma> v1)
- then show ?case using wfV_wf wfG_wf wf_intros wfX_wfY by metis
-qed(auto | metis wfV_wf wfG_wf wf_intros )+
-
-lemma wfX_wfB2:
- fixes \<Gamma>::\<Gamma> and \<Gamma>'::\<Gamma> and v::v and e::e and c::c and \<tau>::\<tau> and ts::"(string*\<tau>) list" and \<Delta>::\<Delta> and s::s and b::b and \<B>::\<B> and \<Phi>::\<Phi> and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def
- and cs::branch_s and css::branch_list
- shows
- wfE_wfB: "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f e : b \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b" and
- wfS_wfB: "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b" and
- wfCS_wfB: "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> ; tid ; dc ; t \<turnstile>\<^sub>w\<^sub>f cs : b \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b" and
- wfCSS_wfB: "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> ; tid ; dclist \<turnstile>\<^sub>w\<^sub>f css : b \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b" and
- "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<Longrightarrow> True" and
- "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<Longrightarrow> True" and
- "\<Theta> ; \<Phi> \<turnstile>\<^sub>w\<^sub>f ftq \<Longrightarrow> True" and
- "\<Theta> ; \<Phi> ; \<B> \<turnstile>\<^sub>w\<^sub>f ft \<Longrightarrow> \<B> |\<subseteq>| \<B>' \<Longrightarrow> \<Theta> ; \<Phi> ; \<B>' \<turnstile>\<^sub>w\<^sub>f ft"
-proof(induct rule:wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.inducts)
- case (wfE_valI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v b)
- then show ?case using wfB_elims wfX_wfB1 by metis
-next
- case (wfE_plusI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 v2)
- then show ?case using wfB_elims wfX_wfB1 by metis
-next
- case (wfE_eqI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 b v2)
- then show ?case using wfB_boolI wfX_wfY by metis
-next
- case (wfE_fstI \<Theta> \<Phi> \<Gamma> \<Delta> v1 b1 b2)
- then show ?case using wfB_elims wfX_wfB1 by metis
-next
- case (wfE_sndI \<Theta> \<Phi> \<Gamma> \<Delta> v1 b1 b2)
- then show ?case using wfB_elims wfX_wfB1 by metis
-next
- case (wfE_concatI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 v2)
- then show ?case using wfB_elims wfX_wfB1 by metis
-next
- case (wfE_splitI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 v2)
- then show ?case using wfB_elims wfX_wfB1
- using wfB_pairI by auto
-next
- case (wfE_lenI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1)
- then show ?case using wfB_elims wfX_wfB1
- using wfB_intI wfX_wfY1(1) by auto
-next
- case (wfE_appI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> f x b c \<tau> s v)
- hence "\<Theta>; \<B>;(x,b,c) #\<^sub>\<Gamma> GNil \<turnstile>\<^sub>w\<^sub>f \<tau>" using wfPhi_f_simple_wfT wfT_b_weakening by fast
- then show ?case using b_of.simps using wfT_b_weakening
- by (metis b_of.cases bot.extremum wfT_elims(2))
-next
- case (wfE_appPI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> b' bv v \<tau> f x b c s)
- hence "\<Theta> ; {| bv |} ;(x,b,c) #\<^sub>\<Gamma> GNil \<turnstile>\<^sub>w\<^sub>f \<tau>" using wfPhi_f_poly_wfT wfX_wfY by blast
- then show ?case using wfE_appPI b_of.simps using wfT_b_weakening wfT_elims wfT_subst_wfB subst_b_b_def by metis
-next
- case (wfE_mvarI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> u \<tau>)
- hence "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau>" using wfD_wfT by fast
- then show ?case using wfT_elims b_of.simps by metis
-next
- case (wfFTNone \<Theta> ft)
- then show ?case by auto
-next
- case (wfFTSome \<Theta> bv ft)
- then show ?case by auto
-next
- case (wfS_valI \<Theta> \<Phi> \<B> \<Gamma> v b \<Delta>)
- then show ?case using wfX_wfB1 by auto
-next
- case (wfS_letI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> e b' x s b)
- then show ?case using wfX_wfB1 by auto
-next
- case (wfS_let2I \<Theta> \<Phi> \<B> \<Gamma> \<Delta> s1 \<tau> x s2 b)
- then show ?case using wfX_wfB1 by auto
-next
- case (wfS_ifI \<Theta> \<B> \<Gamma> v \<Phi> \<Delta> s1 b s2)
- then show ?case using wfX_wfB1 by auto
-next
- case (wfS_varI \<Theta> \<B> \<Gamma> \<tau> v u \<Phi> \<Delta> b s)
- then show ?case using wfX_wfB1 by auto
-next
- case (wfS_assignI u \<tau> \<Delta> \<Theta> \<B> \<Gamma> \<Phi> v)
- then show ?case using wfX_wfB1
- using wfB_unitI wfX_wfY2(5) by auto
-next
- case (wfS_whileI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> s1 s2 b)
- then show ?case using wfX_wfB1 by auto
-next
- case (wfS_seqI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> s1 s2 b)
- then show ?case using wfX_wfB1 by auto
-next
- case (wfS_matchI \<Theta> \<B> \<Gamma> v tid dclist \<Delta> \<Phi> cs b)
- then show ?case using wfX_wfB1 by auto
-next
- case (wfS_branchI \<Theta> \<Phi> \<B> x \<tau> \<Gamma> \<Delta> s b tid dc)
- then show ?case using wfX_wfB1 by auto
-next
- case (wfS_finalI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dc t cs b)
- then show ?case using wfX_wfB1 by auto
-next
- case (wfS_cons \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dc t cs b dclist css)
- then show ?case using wfX_wfB1 by auto
-next
- case (wfD_emptyI \<Theta> \<B> \<Gamma>)
- then show ?case using wfX_wfB1 by auto
-next
- case (wfD_cons \<Theta> \<B> \<Gamma> \<Delta> \<tau> u)
- then show ?case using wfX_wfB1 by auto
-next
- case (wfPhi_emptyI \<Theta>)
- then show ?case using wfX_wfB1 by auto
-next
- case (wfPhi_consI f \<Theta> \<Phi> ft)
- then show ?case using wfX_wfB1 by auto
-next
- case (wfFTI \<Theta> B b \<Phi> x c s \<tau>)
- then show ?case using wfX_wfB1
- by (meson Wellformed.wfFTI wb_b_weakening2(8))
-qed(metis wfV_wf wfG_wf wf_intros wfX_wfB1)
-
-lemmas wfX_wfB = wfX_wfB1 wfX_wfB2
-
-lemma wf_weakening1:
- fixes \<Gamma>::\<Gamma> and \<Gamma>'::\<Gamma> and v::v and e::e and c::c and \<tau>::\<tau> and ts::"(string*\<tau>) list" and \<Delta>::\<Delta> and s::s and \<B>::\<B> and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def
- and cs::branch_s and css::branch_list
-
- shows wfV_weakening: "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>' \<Longrightarrow> toSet \<Gamma> \<subseteq> toSet \<Gamma>' \<Longrightarrow> \<Theta>; \<B>; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f v : b" and
- wfC_weakening: "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f c \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>' \<Longrightarrow> toSet \<Gamma> \<subseteq> toSet \<Gamma>' \<Longrightarrow> \<Theta>; \<B>; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f c" and
- "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<Longrightarrow> True" and
- wfT_weakening: "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau> \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>' \<Longrightarrow> toSet \<Gamma> \<subseteq> toSet \<Gamma>' \<Longrightarrow> \<Theta>; \<B>; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f \<tau>" and
- "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ts \<Longrightarrow> True" and
- "\<turnstile>\<^sub>w\<^sub>f P \<Longrightarrow> True " and
- wfB_weakening: "wfB \<Theta> \<B> b \<Longrightarrow> \<B> |\<subseteq>| \<B>' \<Longrightarrow> wfB \<Theta> \<B> b" and
- wfCE_weakening: "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ce : b \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>' \<Longrightarrow> toSet \<Gamma> \<subseteq> toSet \<Gamma>' \<Longrightarrow> \<Theta>; \<B>; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f ce : b" and
- "\<Theta> \<turnstile>\<^sub>w\<^sub>f td \<Longrightarrow> True"
-proof(nominal_induct
- b and c and \<Gamma> and \<tau> and ts and P and b and b and td
- avoiding: \<Gamma>'
- rule:wfV_wfC_wfG_wfT_wfTs_wfTh_wfB_wfCE_wfTD.strong_induct)
- case (wfV_varI \<Theta> \<B> \<Gamma> b c x)
- hence "Some (b, c) = lookup \<Gamma>' x" using lookup_weakening by metis
- then show ?case using Wellformed.wfV_varI wfV_varI by metis
-next
- case (wfTI z \<Theta> \<B> \<Gamma> b c) (* This proof pattern is used elsewhere when proving weakening for typing predicates *)
- show ?case proof
- show \<open>atom z \<sharp> (\<Theta>, \<B>, \<Gamma>')\<close> using wfTI by auto
- show \<open> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b \<close> using wfTI by auto
- have *:"toSet ((z, b, TRUE) #\<^sub>\<Gamma> \<Gamma>) \<subseteq> toSet ((z, b, TRUE) #\<^sub>\<Gamma> \<Gamma>')" using toSet.simps wfTI by auto
- thus \<open> \<Theta>; \<B>; (z, b, TRUE) #\<^sub>\<Gamma> \<Gamma>' \<turnstile>\<^sub>w\<^sub>f c \<close> using wfTI(8)[OF _ *] wfTI wfX_wfY
- by (simp add: wfG_cons_TRUE)
- qed
-next
- case (wfV_conspI s bv dclist \<Theta> dc x b' c \<B> b \<Gamma> v)
- show ?case proof
- show \<open>AF_typedef_poly s bv dclist \<in> set \<Theta>\<close> using wfV_conspI by auto
- show \<open>(dc, \<lbrace> x : b' | c \<rbrace>) \<in> set dclist\<close> using wfV_conspI by auto
- show \<open> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b \<close> using wfV_conspI by auto
- show \<open>atom bv \<sharp> (\<Theta>, \<B>, \<Gamma>', b, v)\<close> using wfV_conspI by simp
- show \<open> \<Theta>; \<B>; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f v : b'[bv::=b]\<^sub>b\<^sub>b \<close> using wfV_conspI by auto
- qed
-qed(metis wf_intros)+
-
-lemma wf_weakening2:
- fixes \<Gamma>::\<Gamma> and \<Gamma>'::\<Gamma> and v::v and e::e and c::c and \<tau>::\<tau> and ts::"(string*\<tau>) list" and \<Delta>::\<Delta> and s::s and \<B>::\<B> and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def
- and cs::branch_s and css::branch_list
- shows
- wfE_weakening: "\<Theta>; \<Phi>; \<B>; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f e : b \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>' \<Longrightarrow> toSet \<Gamma> \<subseteq> toSet \<Gamma>' \<Longrightarrow> \<Theta>; \<Phi>; \<B>; \<Gamma>' ; \<Delta> \<turnstile>\<^sub>w\<^sub>f e : b" and
- wfS_weakening: "\<Theta>; \<Phi>; \<B>; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>' \<Longrightarrow> toSet \<Gamma> \<subseteq> toSet \<Gamma>' \<Longrightarrow> \<Theta>; \<Phi>; \<B>; \<Gamma>' ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b" and
- "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> ; tid ; dc ; t \<turnstile>\<^sub>w\<^sub>f cs : b \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>' \<Longrightarrow> toSet \<Gamma> \<subseteq> toSet \<Gamma>' \<Longrightarrow> \<Theta>; \<Phi>; \<B>; \<Gamma>' ; \<Delta> ; tid ; dc ; t \<turnstile>\<^sub>w\<^sub>f cs : b" and
- "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> ; tid ; dclist \<turnstile>\<^sub>w\<^sub>f css : b \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>' \<Longrightarrow> toSet \<Gamma> \<subseteq> toSet \<Gamma>' \<Longrightarrow> \<Theta>; \<Phi>; \<B>; \<Gamma>' ; \<Delta> ; tid ; dclist \<turnstile>\<^sub>w\<^sub>f css : b" and
- "\<Theta> \<turnstile>\<^sub>w\<^sub>f (\<Phi>::\<Phi>) \<Longrightarrow> True" and
- wfD_weakning: "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>' \<Longrightarrow> toSet \<Gamma> \<subseteq> toSet \<Gamma>' \<Longrightarrow> \<Theta>; \<B>; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f \<Delta>" and
- "\<Theta> ; \<Phi> \<turnstile>\<^sub>w\<^sub>f ftq \<Longrightarrow> True" and
- "\<Theta> ; \<Phi> ; \<B> \<turnstile>\<^sub>w\<^sub>f ft \<Longrightarrow> True"
-proof(nominal_induct
- b and b and b and b and \<Phi> and \<Delta> and ftq and ft
- avoiding: \<Gamma>'
- rule:wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.strong_induct)
- case (wfE_appPI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> b' bv v \<tau> f x b c s)
- show ?case proof
- show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using wfE_appPI by auto
- show \<open> \<Theta>; \<B>; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wfE_appPI by auto
- show \<open> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b' \<close> using wfE_appPI by auto
- show \<open>atom bv \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>', \<Delta>, b', v, (b_of \<tau>)[bv::=b']\<^sub>b)\<close> using wfE_appPI by auto
- show \<open>Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \<tau> s))) = lookup_fun \<Phi> f\<close> using wfE_appPI by auto
- show \<open> \<Theta>; \<B>; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f v : b[bv::=b']\<^sub>b \<close> using wfE_appPI wf_weakening1 by auto
- qed
-next
- case (wfS_letI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> e b' x s b)
- show ?case proof(rule)
- show \<open> \<Theta> ; \<Phi> ; \<B> ; \<Gamma>' ; \<Delta> \<turnstile>\<^sub>w\<^sub>f e : b' \<close> using wfS_letI by auto
- have "toSet ((x, b', TRUE) #\<^sub>\<Gamma> \<Gamma>) \<subseteq> toSet ((x, b', TRUE) #\<^sub>\<Gamma> \<Gamma>')" using wfS_letI by auto
- thus \<open> \<Theta> ; \<Phi> ; \<B> ; (x, b', TRUE) #\<^sub>\<Gamma> \<Gamma>' ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b \<close> using wfS_letI by (meson wfG_cons wfG_cons_TRUE wfS_wf)
- show \<open> \<Theta>; \<B>; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wfS_letI by auto
- show \<open>atom x \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>', \<Delta>, e, b)\<close> using wfS_letI by auto
- qed
-next
- case (wfS_let2I \<Theta> \<Phi> \<B> \<Gamma> \<Delta> s1 \<tau> x s2 b)
- show ?case proof
- show \<open> \<Theta> ; \<Phi> ; \<B> ; \<Gamma>' ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s1 : b_of \<tau> \<close> using wfS_let2I by auto
- show \<open> \<Theta>; \<B>; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f \<tau> \<close> using wfS_let2I wf_weakening1 by auto
- have "toSet ((x, b_of \<tau>, TRUE) #\<^sub>\<Gamma> \<Gamma>) \<subseteq> toSet ((x, b_of \<tau>, TRUE) #\<^sub>\<Gamma> \<Gamma>')" using wfS_let2I by auto
- thus \<open> \<Theta> ; \<Phi> ; \<B> ; (x, b_of \<tau>, TRUE) #\<^sub>\<Gamma> \<Gamma>' ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s2 : b \<close> using wfS_let2I by (meson wfG_cons wfG_cons_TRUE wfS_wf)
- show \<open>atom x \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>', \<Delta>, s1, b, \<tau>)\<close> using wfS_let2I by auto
- qed
-next
- case (wfS_varI \<Theta> \<B> \<Gamma> \<tau> v u \<Phi> \<Delta> b s)
- show ?case proof
- show "\<Theta>; \<B>; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f \<tau> " using wfS_varI wf_weakening1 by auto
- show "\<Theta>; \<B>; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f v : b_of \<tau> " using wfS_varI wf_weakening1 by auto
- show "atom u \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>', \<Delta>, \<tau>, v, b)" using wfS_varI by auto
- show "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>' ; (u, \<tau>) #\<^sub>\<Delta> \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b " using wfS_varI by auto
- qed
-next
- case (wfS_branchI \<Theta> \<Phi> \<B> x \<tau> \<Gamma> \<Delta> s b tid dc)
- show ?case proof
- have "toSet ((x, b_of \<tau>, TRUE) #\<^sub>\<Gamma> \<Gamma>) \<subseteq> toSet ((x, b_of \<tau>, TRUE) #\<^sub>\<Gamma> \<Gamma>')" using wfS_branchI by auto
- thus \<open> \<Theta> ; \<Phi> ; \<B> ; (x, b_of \<tau>, TRUE) #\<^sub>\<Gamma> \<Gamma>' ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b \<close> using wfS_branchI by (meson wfG_cons wfG_cons_TRUE wfS_wf)
- show \<open>atom x \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>', \<Delta>, \<Gamma>', \<tau>)\<close> using wfS_branchI by auto
- show \<open> \<Theta>; \<B>; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wfS_branchI by auto
- qed
-next
- case (wfS_finalI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist' cs b dclist)
- then show ?case using wf_intros by metis
-next
- case (wfS_cons \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist' cs b css dclist)
- then show ?case using wf_intros by metis
-next
- case (wfS_assertI \<Theta> \<Phi> \<B> x c \<Gamma> \<Delta> s b)
- show ?case proof(rule)
- show \<open> \<Theta>; \<B>; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f c \<close> using wfS_assertI wf_weakening1 by auto
- have "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (x, B_bool, c) #\<^sub>\<Gamma> \<Gamma>'" proof(rule wfG_consI)
- show \<open> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>' \<close> using wfS_assertI by auto
- show \<open>atom x \<sharp> \<Gamma>'\<close> using wfS_assertI by auto
- show \<open> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f B_bool \<close> using wfS_assertI wfB_boolI wfX_wfY by metis
- have "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (x, B_bool, TRUE) #\<^sub>\<Gamma> \<Gamma>'" proof
- show "(TRUE) \<in> {TRUE, FALSE}" by auto
- show \<open> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>' \<close> using wfS_assertI by auto
- show \<open>atom x \<sharp> \<Gamma>'\<close> using wfS_assertI by auto
- show \<open> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f B_bool \<close> using wfS_assertI wfB_boolI wfX_wfY by metis
- qed
- thus \<open> \<Theta>; \<B>; (x, B_bool, TRUE) #\<^sub>\<Gamma> \<Gamma>' \<turnstile>\<^sub>w\<^sub>f c \<close>
- using wf_weakening1(2)[OF \<open> \<Theta>; \<B>; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f c \<close> \<open> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (x, B_bool, TRUE) #\<^sub>\<Gamma> \<Gamma>' \<close>] by force
- qed
- thus \<open> \<Theta>; \<Phi>; \<B>; (x, B_bool, c) #\<^sub>\<Gamma> \<Gamma>' ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b \<close> using wfS_assertI by fastforce
- show \<open> \<Theta>; \<B>; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wfS_assertI by auto
- show \<open>atom x \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>', \<Delta>, c, b, s)\<close> using wfS_assertI by auto
- qed
-qed(metis wf_intros wf_weakening1)+
-
-lemmas wf_weakening = wf_weakening1 wf_weakening2
-
-lemma wfV_weakening_cons:
- fixes \<Gamma>::\<Gamma> and \<Gamma>'::\<Gamma> and v::v and c::c
- assumes "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b" and "atom y \<sharp> \<Gamma>" and "\<Theta>; \<B>; ((y,b',TRUE) #\<^sub>\<Gamma> \<Gamma>) \<turnstile>\<^sub>w\<^sub>f c"
- shows "\<Theta>; \<B>; (y,b',c) #\<^sub>\<Gamma>\<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b"
-proof -
- have "wfG \<Theta> \<B> ((y,b',c) #\<^sub>\<Gamma>\<Gamma>)" using wfG_intros2 assms by auto
- moreover have "toSet \<Gamma> \<subseteq> toSet ((y,b',c) #\<^sub>\<Gamma>\<Gamma>)" using toSet.simps by auto
- ultimately show ?thesis using wf_weakening using assms(1) by blast
-qed
-
-lemma wfG_cons_weakening:
- fixes \<Gamma>'::\<Gamma>
- assumes "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f ((x, b, c) #\<^sub>\<Gamma> \<Gamma>)" and "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>'" and "toSet \<Gamma> \<subseteq> toSet \<Gamma>'" and "atom x \<sharp> \<Gamma>'"
- shows "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f ((x, b, c) #\<^sub>\<Gamma> \<Gamma>')"
-proof(cases "c \<in> {TRUE,FALSE}")
- case True
- then show ?thesis using wfG_wfB wfG_cons2I assms by auto
-next
- case False
- hence *:"\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<and> atom x \<sharp> \<Gamma> \<and> \<Theta>; \<B>; (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f c"
- using wfG_elims(2)[OF assms(1)] by auto
- have a1:"\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>'" using wfG_wfB wfG_cons2I assms by simp
- moreover have a2:"toSet ((x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> ) \<subseteq> toSet ((x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>')" using toSet.simps assms by blast
- moreover have " \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>'" proof
- show "(TRUE) \<in> {TRUE, FALSE}" by auto
- show "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>'" using assms by auto
- show "atom x \<sharp> \<Gamma>'" using assms by auto
- show "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b" using assms wfG_elims by metis
- qed
- hence " \<Theta>; \<B>; (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>' \<turnstile>\<^sub>w\<^sub>f c" using wf_weakening a1 a2 * by auto
- then show ?thesis using wfG_cons1I[of c \<Theta> \<B> \<Gamma>' x b, OF False ] wfG_wfB assms by simp
-qed
-
-lemma wfT_weakening_aux:
- fixes \<Gamma>::\<Gamma> and \<Gamma>'::\<Gamma> and c::c
- assumes "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b | c \<rbrace>" and "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>'" and "toSet \<Gamma> \<subseteq> toSet \<Gamma>'" and "atom z \<sharp> \<Gamma>'"
- shows "\<Theta>; \<B>; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b | c \<rbrace>"
-proof
- show \<open>atom z \<sharp> (\<Theta>, \<B>, \<Gamma>')\<close>
- using wf_supp wfX_wfY assms fresh_prodN fresh_def x_not_in_b_set wfG_fresh_x by metis
- show \<open> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b \<close> using assms wfT_elims by metis
- show \<open> \<Theta>; \<B>; (z, b, TRUE) #\<^sub>\<Gamma> \<Gamma>' \<turnstile>\<^sub>w\<^sub>f c \<close> proof -
- have *:"\<Theta>; \<B>; (z,b,TRUE) #\<^sub>\<Gamma>\<Gamma> \<turnstile>\<^sub>w\<^sub>f c" using wfT_wfC fresh_weakening assms by auto
- moreover have a1:"\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (z, b, TRUE) #\<^sub>\<Gamma> \<Gamma>'" using wfG_cons2I assms \<open>\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b\<close> by simp
- moreover have a2:"toSet ((z, b, TRUE) #\<^sub>\<Gamma> \<Gamma> ) \<subseteq> toSet ((z, b, TRUE) #\<^sub>\<Gamma> \<Gamma>')" using toSet.simps assms by blast
- moreover have " \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (z, b, TRUE) #\<^sub>\<Gamma> \<Gamma>' " proof
- show "(TRUE) \<in> {TRUE, FALSE}" by auto
- show "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>'" using assms by auto
- show "atom z \<sharp> \<Gamma>'" using assms by auto
- show "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b" using assms wfT_elims by metis
- qed
- thus ?thesis using wf_weakening a1 a2 * by auto
- qed
-qed
-
-lemma wfT_weakening_all:
- fixes \<Gamma>::\<Gamma> and \<Gamma>'::\<Gamma> and \<tau>::\<tau>
- assumes "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau>" and "\<Theta>; \<B>' \<turnstile>\<^sub>w\<^sub>f \<Gamma>'" and "toSet \<Gamma> \<subseteq> toSet \<Gamma>'" and "\<B> |\<subseteq>| \<B>'"
- shows "\<Theta>; \<B>' ; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f \<tau>"
- using wb_b_weakening assms wfT_weakening by metis
-
-lemma wfT_weakening_nil:
- fixes \<Gamma>::\<Gamma> and \<Gamma>'::\<Gamma> and \<tau>::\<tau>
- assumes "\<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f \<tau>" and "\<Theta>; \<B>' \<turnstile>\<^sub>w\<^sub>f \<Gamma>'"
- shows "\<Theta>; \<B>' ; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f \<tau>"
- using wfT_weakening_all
- using assms(1) assms(2) toSet.simps(1) by blast
-
-lemma wfTh_wfT2:
- fixes x::x and v::v and \<tau>::\<tau> and G::\<Gamma>
- assumes "wfTh \<Theta>" and "AF_typedef s dclist \<in> set \<Theta>" and
- "(dc, \<tau>) \<in> set dclist" and "\<Theta> ; B \<turnstile>\<^sub>w\<^sub>f G"
- shows "supp \<tau> = {}" and "\<tau>[x::=v]\<^sub>\<tau>\<^sub>v = \<tau>" and "wfT \<Theta> B G \<tau>"
-proof -
- show "supp \<tau> = {}" proof(rule ccontr)
- assume a1: "supp \<tau> \<noteq> {}"
- have "supp \<Theta> \<noteq> {}" proof -
- obtain dclist where dc: "AF_typedef s dclist \<in> set \<Theta> \<and> (dc, \<tau>) \<in> set dclist"
- using assms by auto
- hence "supp (dc,\<tau>) \<noteq> {}"
- using a1 by (simp add: supp_Pair)
- hence "supp dclist \<noteq> {}" using dc supp_list_member by auto
- hence "supp (AF_typedef s dclist) \<noteq> {}" using type_def.supp by auto
- thus ?thesis using supp_list_member dc by auto
- qed
- thus False using assms wfTh_supp by simp
- qed
- thus "\<tau>[x::=v]\<^sub>\<tau>\<^sub>v = \<tau>" by (simp add: fresh_def)
- have "wfT \<Theta> {||} GNil \<tau>" using assms wfTh_wfT by auto
- thus "wfT \<Theta> B G \<tau>" using assms wfT_weakening_nil by simp
-qed
-
-lemma wf_d_weakening:
- fixes \<Gamma>::\<Gamma> and \<Gamma>'::\<Gamma> and v::v and e::e and c::c and \<tau>::\<tau> and ts::"(string*\<tau>) list" and \<Delta>::\<Delta> and s::s and \<B>::\<B> and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def
- and cs::branch_s and css::branch_list
- shows
- "\<Theta>; \<Phi>; \<B>; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f e : b \<Longrightarrow> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>' \<Longrightarrow> setD \<Delta> \<subseteq> setD \<Delta>' \<Longrightarrow> \<Theta>; \<Phi>; \<B>; \<Gamma> ; \<Delta>' \<turnstile>\<^sub>w\<^sub>f e : b" and
- "\<Theta>; \<Phi>; \<B>; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b \<Longrightarrow> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>' \<Longrightarrow> setD \<Delta> \<subseteq> setD \<Delta>' \<Longrightarrow> \<Theta>; \<Phi>; \<B>; \<Gamma> ; \<Delta>' \<turnstile>\<^sub>w\<^sub>f s : b" and
- "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> ; tid ; dc ; t \<turnstile>\<^sub>w\<^sub>f cs : b \<Longrightarrow> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>' \<Longrightarrow> setD \<Delta> \<subseteq> setD \<Delta>' \<Longrightarrow> \<Theta>; \<Phi>; \<B>; \<Gamma> ; \<Delta>' ; tid ; dc ; t \<turnstile>\<^sub>w\<^sub>f cs : b" and
- "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> ; tid ; dclist \<turnstile>\<^sub>w\<^sub>f css : b \<Longrightarrow> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>' \<Longrightarrow> setD \<Delta> \<subseteq> setD \<Delta>' \<Longrightarrow> \<Theta>; \<Phi>; \<B>; \<Gamma> ; \<Delta>' ; tid ; dclist \<turnstile>\<^sub>w\<^sub>f css : b" and
- "\<Theta> \<turnstile>\<^sub>w\<^sub>f (\<Phi>::\<Phi>) \<Longrightarrow> True" and
- "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<Longrightarrow> True" and
- "\<Theta> ; \<Phi> \<turnstile>\<^sub>w\<^sub>f ftq \<Longrightarrow> True" and
- "\<Theta> ; \<Phi> ; \<B> \<turnstile>\<^sub>w\<^sub>f ft \<Longrightarrow> True"
-proof(nominal_induct
- b and b and b and b and \<Phi> and \<Delta> and ftq and ft
- avoiding: \<Delta>'
- rule:wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.strong_induct)
- case (wfE_valI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v b)
- then show ?case using wf_intros by metis
-next
- case (wfE_plusI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 v2)
- then show ?case using wf_intros by metis
-next
- case (wfE_leqI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 v2)
- then show ?case using wf_intros by metis
-next
- case (wfE_eqI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 b v2)
- then show ?case using wf_intros by metis
-next
- case (wfE_fstI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 b1 b2)
- then show ?case using wf_intros by metis
-next
- case (wfE_sndI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 b1 b2)
- then show ?case using wf_intros by metis
-next
- case (wfE_concatI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 v2)
- then show ?case using wf_intros by metis
-next
- case (wfE_splitI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 v2)
- then show ?case using wf_intros by metis
-next
- case (wfE_lenI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1)
- then show ?case using wf_intros by metis
-next
- case (wfE_appI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> f x b c \<tau> s v)
- then show ?case using wf_intros by metis
-next
- case (wfE_appPI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> b' bv v \<tau> f x b c s)
- show ?case proof(rule, (rule wfE_appPI)+)
- show \<open>atom bv \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>, \<Delta>', b', v, (b_of \<tau>)[bv::=b']\<^sub>b)\<close> using wfE_appPI by auto
- show \<open>Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \<tau> s))) = lookup_fun \<Phi> f\<close> using wfE_appPI by auto
- show \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b[bv::=b']\<^sub>b \<close> using wfE_appPI by auto
- qed
-next
- case (wfE_mvarI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> u \<tau>)
- show ?case proof
- show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using wfE_mvarI by auto
- show \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>' \<close> using wfE_mvarI by auto
- show \<open>(u, \<tau>) \<in> setD \<Delta>'\<close> using wfE_mvarI by auto
- qed
-next
- case (wfS_valI \<Theta> \<Phi> \<B> \<Gamma> v b \<Delta>)
- then show ?case using wf_intros by metis
-next
- case (wfS_letI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> e b' x s b)
- show ?case proof(rule)
- show \<open> \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta>' \<turnstile>\<^sub>w\<^sub>f e : b' \<close> using wfS_letI by auto
- have "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (x, b', TRUE) #\<^sub>\<Gamma> \<Gamma>" using wfG_cons2I wfX_wfY wfS_letI by metis
- hence "\<Theta>; \<B>; (x, b', TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>'" using wf_weakening2(6) wfS_letI by force
- thus \<open> \<Theta> ; \<Phi> ; \<B> ; (x, b', TRUE) #\<^sub>\<Gamma> \<Gamma> ; \<Delta>' \<turnstile>\<^sub>w\<^sub>f s : b \<close> using wfS_letI by metis
- show \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>' \<close> using wfS_letI by auto
- show \<open>atom x \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>, \<Delta>', e, b)\<close> using wfS_letI by auto
- qed
-next
- case (wfS_assertI \<Theta> \<Phi> \<B> x c \<Gamma> \<Delta> s b)
- show ?case proof
- have "\<Theta>; \<B>; (x, B_bool, c) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>'" proof(rule wf_weakening2(6))
- show \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>' \<close> using wfS_assertI by auto
- next
- show \<open> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (x, B_bool, c) #\<^sub>\<Gamma> \<Gamma> \<close> using wfS_assertI wfX_wfY by metis
- next
- show \<open>toSet \<Gamma> \<subseteq> toSet ((x, B_bool, c) #\<^sub>\<Gamma> \<Gamma>)\<close> using wfS_assertI by auto
- qed
- thus \<open> \<Theta>; \<Phi>; \<B>; (x, B_bool, c) #\<^sub>\<Gamma> \<Gamma> ; \<Delta>' \<turnstile>\<^sub>w\<^sub>f s : b \<close> using wfS_assertI wfX_wfY by metis
- next
- show \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f c \<close> using wfS_assertI by auto
- next
- show \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>' \<close> using wfS_assertI by auto
- next
- show \<open>atom x \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>, \<Delta>', c, b, s)\<close> using wfS_assertI by auto
- qed
-next
- case (wfS_let2I \<Theta> \<Phi> \<B> \<Gamma> \<Delta> s1 \<tau> x s2 b)
- show ?case proof
- show \<open> \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta>' \<turnstile>\<^sub>w\<^sub>f s1 : b_of \<tau> \<close> using wfS_let2I by auto
- show \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau> \<close> using wfS_let2I by auto
- have "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (x, b_of \<tau>, TRUE) #\<^sub>\<Gamma> \<Gamma>" using wfG_cons2I wfX_wfY wfS_let2I by metis
- hence "\<Theta>; \<B>; (x, b_of \<tau>, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>'" using wf_weakening2(6) wfS_let2I by force
- thus \<open> \<Theta> ; \<Phi> ; \<B> ; (x, b_of \<tau>, TRUE) #\<^sub>\<Gamma> \<Gamma> ; \<Delta>' \<turnstile>\<^sub>w\<^sub>f s2 : b \<close> using wfS_let2I by metis
- show \<open>atom x \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>, \<Delta>', s1, b,\<tau>)\<close> using wfS_let2I by auto
- qed
-next
- case (wfS_ifI \<Theta> \<B> \<Gamma> v \<Phi> \<Delta> s1 b s2)
- then show ?case using wf_intros by metis
-next
- case (wfS_varI \<Theta> \<B> \<Gamma> \<tau> v u \<Phi> \<Delta> b s)
- show ?case proof
- show \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau> \<close> using wfS_varI by auto
- show \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b_of \<tau> \<close> using wfS_varI by auto
- show \<open>atom u \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>, \<Delta>', \<tau>, v, b)\<close> using wfS_varI setD.simps by auto
- have "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f (u, \<tau>) #\<^sub>\<Delta> \<Delta>'" using wfS_varI wfD_cons setD.simps u_fresh_d by metis
- thus \<open> \<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; (u, \<tau>) #\<^sub>\<Delta> \<Delta>' \<turnstile>\<^sub>w\<^sub>f s : b \<close> using wfS_varI setD.simps by blast
- qed
-next
- case (wfS_assignI u \<tau> \<Delta> \<Theta> \<B> \<Gamma> \<Phi> v)
- show ?case proof
- show \<open>(u, \<tau>) \<in> setD \<Delta>'\<close> using wfS_assignI setD.simps by auto
- show \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>' \<close> using wfS_assignI by auto
- show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using wfS_assignI by auto
- show \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b_of \<tau> \<close> using wfS_assignI by auto
- qed
-next
- case (wfS_whileI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> s1 s2 b)
- then show ?case using wf_intros by metis
-next
- case (wfS_seqI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> s1 s2 b)
- then show ?case using wf_intros by metis
-next
- case (wfS_matchI \<Theta> \<B> \<Gamma> v tid dclist \<Delta> \<Phi> cs b)
- then show ?case using wf_intros by metis
-next
- case (wfS_branchI \<Theta> \<Phi> \<B> x \<tau> \<Gamma> \<Delta> s b tid dc)
- show ?case proof
- have "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (x, b_of \<tau>, TRUE) #\<^sub>\<Gamma> \<Gamma>" using wfG_cons2I wfX_wfY wfS_branchI by metis
- hence "\<Theta>; \<B>; (x, b_of \<tau>, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>'" using wf_weakening2(6) wfS_branchI by force
- thus \<open> \<Theta> ; \<Phi> ; \<B> ; (x, b_of \<tau>, TRUE) #\<^sub>\<Gamma> \<Gamma> ; \<Delta>' \<turnstile>\<^sub>w\<^sub>f s : b \<close> using wfS_branchI by simp
- show \<open> atom x \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>, \<Delta>', \<Gamma>, \<tau>)\<close> using wfS_branchI by auto
- show \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>' \<close> using wfS_branchI by auto
- qed
-next
- case (wfS_finalI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist' cs b dclist)
- then show ?case using wf_intros by metis
-next
- case (wfS_cons \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist' cs b css dclist)
- then show ?case using wf_intros by metis
-qed(auto+)
-
-section \<open>Useful well-formedness instances\<close>
-
-text \<open>Well-formedness for particular constructs that we will need later\<close>
-
-lemma wfC_e_eq:
- fixes ce::ce and \<Gamma>::\<Gamma>
- assumes "\<Theta> ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ce : b" and "atom x \<sharp> \<Gamma> "
- shows "\<Theta> ; \<B> ; ((x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>) \<turnstile>\<^sub>w\<^sub>f (CE_val (V_var x) == ce )"
-proof -
- have "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b" using assms wfX_wfB by auto
- hence wbg: "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>" using wfX_wfY assms by auto
- show ?thesis proof
- show *:"\<Theta> ; \<B> ; (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_val (V_var x) : b"
- proof(rule)
- show "\<Theta> ; \<B> ; (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f V_var x : b " proof
- show "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> " using wfG_cons2I wfX_wfY assms \<open>\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b\<close> by auto
- show "Some (b, TRUE) = lookup ((x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>) x" using lookup.simps by auto
- qed
- qed
- show "\<Theta> ; \<B> ; (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f ce : b"
- using assms wf_weakening1(8)[OF assms(1), of "(x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> "] * toSet.simps wfX_wfY
- by (metis Un_subset_iff equalityE)
- qed
-qed
-
-lemma wfC_e_eq2:
- fixes e1::ce and e2::ce
- assumes "\<Theta> ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f e1 : b" and "\<Theta> ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f e2 : b" and " \<turnstile>\<^sub>w\<^sub>f \<Theta>" and "atom x \<sharp> \<Gamma>"
- shows "\<Theta>; \<B>; (x, b, (CE_val (V_var x)) == e1) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f (CE_val (V_var x)) == e2 "
-proof(rule wfC_eqI)
- have *: "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (x, b, CE_val (V_var x) == e1 ) #\<^sub>\<Gamma> \<Gamma>" proof(rule wfG_cons1I )
- show "(CE_val (V_var x) == e1 ) \<notin> {TRUE, FALSE}" by auto
- show "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>" using assms wfX_wfY by metis
- show *:"atom x \<sharp> \<Gamma>" using assms by auto
- show "\<Theta>; \<B>; (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_val (V_var x) == e1" using wfC_e_eq assms * by auto
- show "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b" using assms wfX_wfB by auto
- qed
- show "\<Theta>; \<B>; (x, b, CE_val (V_var x) == e1 ) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_val (V_var x) : b" using assms * wfCE_valI wfV_varI by auto
- show "\<Theta>; \<B>; (x, b, CE_val (V_var x) == e1 ) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f e2 : b" proof(rule wf_weakening1(8))
- show "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f e2 : b " using assms by auto
- show "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (x, b, CE_val (V_var x) == e1 ) #\<^sub>\<Gamma> \<Gamma>" using * by auto
- show "toSet \<Gamma> \<subseteq> toSet ((x, b, CE_val (V_var x) == e1 ) #\<^sub>\<Gamma> \<Gamma>)" by auto
- qed
-qed
-
-lemma wfT_wfT_if_rev:
- assumes "wfV P \<B> \<Gamma> v (base_for_lit l)" and "wfT P \<B> \<Gamma> t" and \<open>atom z1 \<sharp> \<Gamma>\<close>
- shows "wfT P \<B> \<Gamma> (\<lbrace> z1 : b_of t | CE_val v == CE_val (V_lit l) IMP (c_of t z1) \<rbrace>)"
-proof
- show \<open> P; \<B> \<turnstile>\<^sub>w\<^sub>f b_of t \<close> using wfX_wfY assms by meson
- have wfg: " P; \<B> \<turnstile>\<^sub>w\<^sub>f (z1, b_of t, TRUE) #\<^sub>\<Gamma> \<Gamma>" using assms wfV_wf wfG_cons2I wfX_wfY
- by (meson wfG_cons_TRUE)
- show \<open> P; \<B> ; (z1, b_of t, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f [ v ]\<^sup>c\<^sup>e == [ [ l ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of t z1 \<close> proof
- show *: \<open> P; \<B> ; (z1, b_of t, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f [ v ]\<^sup>c\<^sup>e == [ [ l ]\<^sup>v ]\<^sup>c\<^sup>e \<close>
- proof(rule wfC_eqI[where b="base_for_lit l"])
- show "P; \<B> ; (z1, b_of t, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f [ v ]\<^sup>c\<^sup>e : base_for_lit l"
- using assms wf_intros wf_weakening wfg by (meson wfV_weakening_cons)
- show "P; \<B> ; (z1, b_of t, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f [ [ l ]\<^sup>v ]\<^sup>c\<^sup>e : base_for_lit l" using wfg assms wf_intros wf_weakening wfV_weakening_cons by meson
- qed
- have " t = \<lbrace> z1 : b_of t | c_of t z1 \<rbrace>" using c_of_eq
- using assms(2) assms(3) b_of_c_of_eq wfT_x_fresh by auto
- thus \<open> P; \<B> ; (z1, b_of t, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f c_of t z1 \<close> using wfT_wfC assms wfG_elims * by simp
- qed
- show \<open>atom z1 \<sharp> (P, \<B>, \<Gamma>)\<close> using assms wfG_fresh_x wfX_wfY by metis
-qed
-
-lemma wfT_eq_imp:
- fixes zz::x and ll::l and \<tau>'::\<tau>
- assumes "base_for_lit ll = B_bool" and "\<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f \<tau>'" and
- "\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f (x, b_of \<lbrace> z' : B_bool | TRUE \<rbrace>, c_of \<lbrace> z' : B_bool | TRUE \<rbrace> x) #\<^sub>\<Gamma> GNil" and "atom zz \<sharp> x"
- shows "\<Theta> ; {||} ; (x, b_of \<lbrace> z' : B_bool | TRUE \<rbrace>, c_of \<lbrace> z' : B_bool | TRUE \<rbrace> x) #\<^sub>\<Gamma>
- GNil \<turnstile>\<^sub>w\<^sub>f \<lbrace> zz : b_of \<tau>' | [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ ll ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of \<tau>' zz \<rbrace>"
-proof(rule wfT_wfT_if_rev)
- show \<open> \<Theta> ; {||} ; (x, b_of \<lbrace> z' : B_bool | TRUE \<rbrace>, c_of \<lbrace> z' : B_bool | TRUE \<rbrace> x) #\<^sub>\<Gamma> GNil \<turnstile>\<^sub>w\<^sub>f [ x ]\<^sup>v : base_for_lit ll \<close>
- using wfV_varI lookup.simps base_for_lit.simps assms by simp
- show \<open> \<Theta> ; {||} ; (x, b_of \<lbrace> z' : B_bool | TRUE \<rbrace>, c_of \<lbrace> z' : B_bool | TRUE \<rbrace> x) #\<^sub>\<Gamma> GNil \<turnstile>\<^sub>w\<^sub>f \<tau>' \<close>
- using wf_weakening assms toSet.simps by auto
- show \<open>atom zz \<sharp> (x, b_of \<lbrace> z' : B_bool | TRUE \<rbrace>, c_of \<lbrace> z' : B_bool | TRUE \<rbrace> x) #\<^sub>\<Gamma> GNil\<close>
- unfolding fresh_GCons fresh_prod3 b_of.simps c_of_true
- using x_fresh_b fresh_GNil c_of_true c.fresh assms by metis
-qed
-
-lemma wfC_v_eq:
- fixes ce::ce and \<Gamma>::\<Gamma> and v::v
- assumes "\<Theta> ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b" and "atom x \<sharp> \<Gamma> "
- shows "\<Theta> ; \<B> ; ((x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>) \<turnstile>\<^sub>w\<^sub>f (CE_val (V_var x) == CE_val v )"
- using wfC_e_eq wfCE_valI assms wfX_wfY by auto
-
-lemma wfT_e_eq:
- fixes ce::ce
- assumes "\<Theta> ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ce : b" and "atom z \<sharp> \<Gamma>"
- shows "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b | CE_val (V_var z) == ce \<rbrace>"
-proof
- show "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b" using wfX_wfB assms by auto
- show " atom z \<sharp> (\<Theta>, \<B>, \<Gamma>)" using assms wfG_fresh_x wfX_wfY by metis
- show "\<Theta> ; \<B> ; (z, b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_val (V_var z) == ce "
- using wfTI wfC_e_eq assms wfTI by auto
-qed
-
-lemma wfT_v_eq:
- assumes " wfB \<Theta> \<B> b" and "wfV \<Theta> \<B> \<Gamma> v b" and "atom z \<sharp> \<Gamma>"
- shows "wfT \<Theta> \<B> \<Gamma> \<lbrace> z : b | C_eq (CE_val (V_var z)) (CE_val v)\<rbrace>"
- using wfT_e_eq wfE_valI assms wfX_wfY
- by (simp add: wfCE_valI)
-
-lemma wfC_wfG:
- fixes \<Gamma>::\<Gamma> and c::c and b::b
- assumes "\<Theta> ; B ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f c" and "\<Theta> ; B \<turnstile>\<^sub>w\<^sub>f b" and "atom x \<sharp> \<Gamma>"
- shows "\<Theta> ; B \<turnstile>\<^sub>w\<^sub>f (x,b,c)#\<^sub>\<Gamma> \<Gamma>"
-proof -
- have " \<Theta> ; B \<turnstile>\<^sub>w\<^sub>f (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>" using wfG_cons2I assms wfX_wfY by fast
- hence " \<Theta> ; B ; (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f c " using wfC_weakening assms by force
- thus ?thesis using wfG_consI assms wfX_wfY by metis
-qed
-
-section \<open>Replacing the constraint on a variable in a context\<close>
-
-lemma wfG_cons_fresh2:
- fixes \<Gamma>'::\<Gamma>
- assumes "wfG P \<B> (( (x',b',c') #\<^sub>\<Gamma> \<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>))"
- shows "x'\<noteq>x"
-proof -
- have "atom x' \<sharp> (\<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>)"
- using assms wfG_elims(2) by blast
- thus ?thesis
- using fresh_gamma_append[of "atom x'" \<Gamma>' "(x, b, c) #\<^sub>\<Gamma> \<Gamma>"] fresh_GCons fresh_prod3[of "atom x'" x b c] by auto
-qed
-
-lemma replace_in_g_inside:
- fixes \<Gamma>::\<Gamma>
- assumes "wfG P \<B> (\<Gamma>'@((x,b0,c0') #\<^sub>\<Gamma>\<Gamma>))"
- shows "replace_in_g (\<Gamma>'@((x,b0,c0') #\<^sub>\<Gamma>\<Gamma>)) x c0 = (\<Gamma>'@((x,b0,c0) #\<^sub>\<Gamma>\<Gamma>))"
-using assms proof(induct \<Gamma>' rule: \<Gamma>_induct)
- case GNil
- then show ?case using replace_in_g.simps by auto
-next
- case (GCons x' b' c' \<Gamma>'')
- hence "P; \<B> \<turnstile>\<^sub>w\<^sub>f ((x', b', c') #\<^sub>\<Gamma> (\<Gamma>''@ (x, b0, c0') #\<^sub>\<Gamma> \<Gamma> ))" by simp
- hence "x \<noteq> x'" using wfG_cons_fresh2 by metis
- then show ?case using replace_in_g.simps GCons by (simp add: wfG_cons)
-qed
-
-lemma wfG_supp_rig_eq:
- fixes \<Gamma>::\<Gamma>
- assumes "wfG P \<B> (\<Gamma>'' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma>)" and "wfG P \<B> (\<Gamma>'' @ (x, b0, c0') #\<^sub>\<Gamma> \<Gamma>)"
- shows "supp (\<Gamma>'' @ (x, b0, c0') #\<^sub>\<Gamma> \<Gamma>) \<union> supp \<B> = supp (\<Gamma>'' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma>) \<union> supp \<B>"
-using assms proof(induct \<Gamma>'')
- case GNil
- have "supp (GNil @ (x, b0, c0') #\<^sub>\<Gamma> \<Gamma>) \<union> supp \<B> = supp ((x, b0, c0') #\<^sub>\<Gamma> \<Gamma>) \<union> supp \<B>" using supp_Cons supp_GNil by auto
- also have "... = supp x \<union> supp b0 \<union> supp c0' \<union> supp \<Gamma> \<union> supp \<B> " using supp_GCons by auto
- also have "... = supp x \<union> supp b0 \<union> supp c0 \<union> supp \<Gamma> \<union> supp \<B> " using GNil wfG_wfC[THEN wfC_supp_cons(2) ] by fastforce
- also have "... = (supp ((x, b0, c0) #\<^sub>\<Gamma> \<Gamma>)) \<union> supp \<B> " using supp_GCons by auto
- finally have "supp (GNil @ (x, b0, c0') #\<^sub>\<Gamma> \<Gamma>) \<union> supp \<B> = supp (GNil @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma>) \<union> supp \<B>" using supp_Cons supp_GNil by auto
- then show ?case using supp_GCons wfG_cons2 by auto
-next
- case (GCons xbc \<Gamma>1)
- moreover have " (xbc #\<^sub>\<Gamma> \<Gamma>1) @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> = (xbc #\<^sub>\<Gamma> (\<Gamma>1 @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma>))" by simp
- moreover have " (xbc #\<^sub>\<Gamma> \<Gamma>1) @ (x, b0, c0') #\<^sub>\<Gamma> \<Gamma> = (xbc #\<^sub>\<Gamma> (\<Gamma>1 @ (x, b0, c0') #\<^sub>\<Gamma> \<Gamma>))" by simp
- ultimately have "(P; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>1 @ ((x, b0, c0) #\<^sub>\<Gamma> \<Gamma>)) \<and> P; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>1 @ ((x, b0, c0') #\<^sub>\<Gamma> \<Gamma>)" using wfG_cons2 by metis
- thus ?case using GCons supp_GCons by auto
-qed
-
-lemma fresh_replace_inside[ms_fresh]:
- fixes y::x and \<Gamma>::\<Gamma>
- assumes "wfG P \<B> (\<Gamma>'' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>)" and "wfG P \<B> (\<Gamma>'' @ (x, b, c') #\<^sub>\<Gamma> \<Gamma>)"
- shows "atom y \<sharp> (\<Gamma>'' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>) = atom y \<sharp> (\<Gamma>'' @ (x, b, c') #\<^sub>\<Gamma> \<Gamma>)"
- unfolding fresh_def using wfG_supp_rig_eq assms x_not_in_b_set by fast
-
-lemma wf_replace_inside1:
- fixes \<Gamma>::\<Gamma> and \<Phi>::\<Phi> and \<Theta>::\<Theta> and \<Gamma>'::\<Gamma> and v::v and e::e and c::c and c''::c and c'::c and \<tau>::\<tau> and ts::"(string*\<tau>) list" and \<Delta>::\<Delta> and b'::b and b::b and s::s
- and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def and cs::branch_s and css::branch_list
-
-shows wfV_replace_inside: "\<Theta>; \<B>; G \<turnstile>\<^sub>w\<^sub>f v : b' \<Longrightarrow> G = (\<Gamma>' @ (x, b, c') #\<^sub>\<Gamma> \<Gamma>) \<Longrightarrow> \<Theta>; \<B>; ((x,b,TRUE) #\<^sub>\<Gamma>\<Gamma>) \<turnstile>\<^sub>w\<^sub>f c \<Longrightarrow> \<Theta> ; \<B> ; (\<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>) \<turnstile>\<^sub>w\<^sub>f v : b'" and
- wfC_replace_inside: "\<Theta>; \<B>; G \<turnstile>\<^sub>w\<^sub>f c'' \<Longrightarrow> G = (\<Gamma>' @ (x, b, c') #\<^sub>\<Gamma> \<Gamma>) \<Longrightarrow> \<Theta>; \<B>; ((x,b,TRUE) #\<^sub>\<Gamma>\<Gamma>) \<turnstile>\<^sub>w\<^sub>f c \<Longrightarrow> \<Theta> ; \<B> ; (\<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>) \<turnstile>\<^sub>w\<^sub>f c''" and
- wfG_replace_inside: "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f G \<Longrightarrow> G = (\<Gamma>' @ (x, b, c') #\<^sub>\<Gamma> \<Gamma>) \<Longrightarrow> \<Theta>; \<B>; ((x,b,TRUE) #\<^sub>\<Gamma>\<Gamma>) \<turnstile>\<^sub>w\<^sub>f c \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (\<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>) " and
- wfT_replace_inside: "\<Theta>; \<B>; G \<turnstile>\<^sub>w\<^sub>f \<tau> \<Longrightarrow> G = (\<Gamma>' @ (x, b, c') #\<^sub>\<Gamma> \<Gamma>) \<Longrightarrow> \<Theta>; \<B>; ((x,b,TRUE) #\<^sub>\<Gamma>\<Gamma>) \<turnstile>\<^sub>w\<^sub>f c \<Longrightarrow> \<Theta> ; \<B> ; (\<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>) \<turnstile>\<^sub>w\<^sub>f \<tau>" and
- "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ts \<Longrightarrow> True" and
- "\<turnstile>\<^sub>w\<^sub>f P \<Longrightarrow> True" and
- "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b \<Longrightarrow> True" and
- wfCE_replace_inside: "\<Theta> ; \<B> ; G \<turnstile>\<^sub>w\<^sub>f ce : b' \<Longrightarrow> G = (\<Gamma>' @ (x, b, c') #\<^sub>\<Gamma> \<Gamma>) \<Longrightarrow> \<Theta>; \<B>; ((x,b,TRUE) #\<^sub>\<Gamma>\<Gamma>) \<turnstile>\<^sub>w\<^sub>f c \<Longrightarrow> \<Theta> ; \<B> ; (\<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>) \<turnstile>\<^sub>w\<^sub>f ce : b'" and
- "\<Theta> \<turnstile>\<^sub>w\<^sub>f td \<Longrightarrow> True"
-proof(nominal_induct
- b' and c'' and G and \<tau> and ts and P and b and b' and td
- avoiding: \<Gamma>' c'
-rule:wfV_wfC_wfG_wfT_wfTs_wfTh_wfB_wfCE_wfTD.strong_induct)
- case (wfV_varI \<Theta> \<B> \<Gamma>2 b2 c2 x2)
- then show ?case using wf_intros by (metis lookup_in_rig_eq lookup_in_rig_neq replace_in_g_inside)
-next
- case (wfV_conspI s bv dclist \<Theta> dc x1 b' c1 \<B> b1 \<Gamma>1 v)
- show ?case proof
- show \<open>AF_typedef_poly s bv dclist \<in> set \<Theta>\<close> using wfV_conspI by auto
- show \<open>(dc, \<lbrace> x1 : b' | c1 \<rbrace>) \<in> set dclist\<close> using wfV_conspI by auto
- show \<open> \<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f b1 \<close> using wfV_conspI by auto
- show *: \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b'[bv::=b1]\<^sub>b\<^sub>b \<close> using wfV_conspI by auto
- moreover have "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>' @ (x, b, c') #\<^sub>\<Gamma> \<Gamma>" using wfV_wf wfV_conspI by simp
- ultimately have "atom bv \<sharp> \<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>" unfolding fresh_def using wfV_wf wfG_supp_rig_eq wfV_conspI
- by (metis Un_iff fresh_def)
- thus \<open>atom bv \<sharp> (\<Theta>, \<B>, \<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>, b1, v)\<close>
- unfolding fresh_prodN using fresh_prodN wfV_conspI by metis
- qed
-next
- case (wfTI z \<Theta> \<B> G b1 c1)
- show ?case proof
- show \<open> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b1 \<close> using wfTI by auto
-
- have "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (x, b, c) #\<^sub>\<Gamma> \<Gamma>" using wfG_consI wfTI wfG_cons wfX_wfY by metis
- moreover hence *:"wfG \<Theta> \<B> (\<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>)" using wfX_wfY
- by (metis append_g.simps(2) wfG_cons2 wfTI.hyps wfTI.prems(1) wfTI.prems(2))
- hence \<open>atom z \<sharp> \<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>\<close>
- using fresh_replace_inside[of \<Theta> \<B> \<Gamma>' x b c \<Gamma> c' z,OF *] wfTI wfX_wfY wfG_elims by metis
- thus \<open>atom z \<sharp> (\<Theta>, \<B>, \<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>)\<close> using wfG_fresh_x[OF *] by auto
-
- have "(z, b1, TRUE) #\<^sub>\<Gamma> G = ((z, b1, TRUE) #\<^sub>\<Gamma> \<Gamma>') @ (x, b, c') #\<^sub>\<Gamma> \<Gamma>"
- using wfTI append_g.simps by metis
- thus \<open> \<Theta>; \<B>; (z, b1, TRUE) #\<^sub>\<Gamma> \<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f c1 \<close>
- using wfTI(9)[OF _ wfTI(11)] by fastforce
- qed
-next
- case (wfG_nilI \<Theta>)
- hence "GNil = (x, b, c') #\<^sub>\<Gamma> \<Gamma>" using append_g.simps \<Gamma>.distinct GNil_append by auto
- hence "False" using \<Gamma>.distinct by auto
- then show ?case by auto
-next
- case (wfG_cons1I c1 \<Theta> \<B> G x1 b1)
- show ?case proof(cases "\<Gamma>'=GNil")
- case True
- then show ?thesis using wfG_cons1I wfG_consI by auto
- next
- case False
- then obtain G'::\<Gamma> where *:"(x1, b1, c1) #\<^sub>\<Gamma> G' = \<Gamma>'" using wfG_cons1I wfG_cons1I(7) GCons_eq_append_conv by auto
- hence **:" G = G' @ (x, b, c') #\<^sub>\<Gamma> \<Gamma>" using wfG_cons1I by auto
- hence " \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f G' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>" using wfG_cons1I by auto
- have "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (x1, b1, c1) #\<^sub>\<Gamma> G' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>" proof(rule Wellformed.wfG_cons1I)
- show "c1 \<notin> {TRUE, FALSE}" using wfG_cons1I by auto
- show "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f G' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma> " using wfG_cons1I(3)[of G',OF **] wfG_cons1I by auto
- show "atom x1 \<sharp> G' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>" using wfG_cons1I * ** fresh_replace_inside by metis
- show "\<Theta>; \<B>; (x1, b1, TRUE) #\<^sub>\<Gamma> G' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f c1" using wfG_cons1I(6)[of " (x1, b1, TRUE) #\<^sub>\<Gamma> G'"] wfG_cons1I ** by auto
- show "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b1" using wfG_cons1I by auto
- qed
- thus ?thesis using * by auto
- qed
-next
- case (wfG_cons2I c1 \<Theta> \<B> G x1 b1)
- show ?case proof(cases "\<Gamma>'=GNil")
- case True
- then show ?thesis using wfG_cons2I wfG_consI by auto
- next
- case False
- then obtain G'::\<Gamma> where *:"(x1, b1, c1) #\<^sub>\<Gamma> G' = \<Gamma>'" using wfG_cons2I GCons_eq_append_conv by auto
- hence **:" G = G' @ (x, b, c') #\<^sub>\<Gamma> \<Gamma>" using wfG_cons2I by auto
- moreover have " \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f G' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>" using wfG_cons2I * ** by auto
- moreover hence "atom x1 \<sharp> G' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>" using wfG_cons2I * ** fresh_replace_inside by metis
- ultimately show ?thesis using Wellformed.wfG_cons2I[OF wfG_cons2I(1), of \<Theta> \<B> "G'@ (x, b, c) #\<^sub>\<Gamma> \<Gamma>" x1 b1] wfG_cons2I * ** by auto
- qed
-qed(metis wf_intros )+
-
-lemma wf_replace_inside2:
- fixes \<Gamma>::\<Gamma> and \<Phi>::\<Phi> and \<Theta>::\<Theta> and \<Gamma>'::\<Gamma> and v::v and e::e and c::c and c''::c and c'::c and \<tau>::\<tau> and ts::"(string*\<tau>) list" and \<Delta>::\<Delta> and b'::b and b::b and s::s
- and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def and cs::branch_s and css::branch_list
-shows
- "\<Theta> ; \<Phi> ; \<B> ; G ; D \<turnstile>\<^sub>w\<^sub>f e : b' \<Longrightarrow> G = (\<Gamma>' @ (x, b, c') #\<^sub>\<Gamma> \<Gamma>) \<Longrightarrow> \<Theta>; \<B>; ((x,b,TRUE) #\<^sub>\<Gamma>\<Gamma>) \<turnstile>\<^sub>w\<^sub>f c \<Longrightarrow> \<Theta> ; \<Phi> ; \<B> ; (\<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>); D \<turnstile>\<^sub>w\<^sub>f e : b'" and
- "\<Theta>; \<Phi>; \<B>; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b \<Longrightarrow> True" and
- "\<Theta>; \<Phi>; \<B>; \<Gamma> ; \<Delta> ; tid ; dc ; t \<turnstile>\<^sub>w\<^sub>f cs : b \<Longrightarrow> True" and
- "\<Theta>; \<Phi>; \<B>; \<Gamma> ; \<Delta> ; tid ; dclist \<turnstile>\<^sub>w\<^sub>f css : b \<Longrightarrow> True" and
- "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<Longrightarrow> True" and
- "\<Theta>; \<B>; G \<turnstile>\<^sub>w\<^sub>f \<Delta> \<Longrightarrow> G = (\<Gamma>' @ (x, b, c') #\<^sub>\<Gamma> \<Gamma>) \<Longrightarrow> \<Theta>; \<B>; ((x,b,TRUE) #\<^sub>\<Gamma>\<Gamma>) \<turnstile>\<^sub>w\<^sub>f c \<Longrightarrow> \<Theta> ; \<B> ; (\<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>) \<turnstile>\<^sub>w\<^sub>f \<Delta>" and
- "\<Theta> ; \<Phi> \<turnstile>\<^sub>w\<^sub>f ftq \<Longrightarrow> True" and
- "\<Theta> ; \<Phi> ; \<B> \<turnstile>\<^sub>w\<^sub>f ft \<Longrightarrow> True"
-proof(nominal_induct
- b' and b and b and b and \<Phi> and \<Delta> and ftq and ft
- avoiding: \<Gamma>' c'
- rule:wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.strong_induct)
-case (wfE_valI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v b)
- then show ?case using wf_replace_inside1 Wellformed.wfE_valI by auto
-next
- case (wfE_plusI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 v2)
- then show ?case using wf_replace_inside1 Wellformed.wfE_plusI by auto
-next
- case (wfE_leqI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 v2)
- then show ?case using wf_replace_inside1 Wellformed.wfE_leqI by auto
-next
- case (wfE_eqI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 b v2)
- then show ?case using wf_replace_inside1 Wellformed.wfE_eqI by metis
-next
- case (wfE_fstI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 b1 b2)
- then show ?case using wf_replace_inside1 Wellformed.wfE_fstI by metis
-next
- case (wfE_sndI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 b1 b2)
- then show ?case using wf_replace_inside1 Wellformed.wfE_sndI by metis
-next
- case (wfE_concatI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 v2)
- then show ?case using wf_replace_inside1 Wellformed.wfE_concatI by auto
-next
- case (wfE_splitI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 v2)
- then show ?case using wf_replace_inside1 Wellformed.wfE_splitI by auto
-next
- case (wfE_lenI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1)
- then show ?case using wf_replace_inside1 Wellformed.wfE_lenI by metis
-next
- case (wfE_appI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> f x b c \<tau> s v)
- then show ?case using wf_replace_inside1 Wellformed.wfE_appI by metis
-next
- case (wfE_appPI \<Theta> \<Phi> \<B> \<Gamma>'' \<Delta> b' bv v \<tau> f x1 b1 c1 s)
- show ?case proof
- show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using wfE_appPI by auto
- show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wfE_appPI by auto
- show \<open> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b' \<close> using wfE_appPI by auto
- show *:\<open> \<Theta>; \<B>; \<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b1[bv::=b']\<^sub>b \<close> using wfE_appPI wf_replace_inside1 by auto
-
- moreover have "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>' @ (x, b, c') #\<^sub>\<Gamma> \<Gamma>" using wfV_wf wfE_appPI by metis
- ultimately have "atom bv \<sharp> \<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>"
- unfolding fresh_def using wfV_wf wfG_supp_rig_eq wfE_appPI Un_iff fresh_def by metis
-
- thus \<open>atom bv \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>, \<Delta>, b', v, (b_of \<tau>)[bv::=b']\<^sub>b)\<close>
- using wfE_appPI fresh_prodN by metis
- show \<open>Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x1 b1 c1 \<tau> s))) = lookup_fun \<Phi> f\<close> using wfE_appPI by auto
- qed
-next
- case (wfE_mvarI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> u \<tau>)
- then show ?case using wf_replace_inside1 Wellformed.wfE_mvarI by metis
-next
- case (wfD_emptyI \<Theta> \<B> \<Gamma>)
- then show ?case using wf_replace_inside1 Wellformed.wfD_emptyI by metis
-next
- case (wfD_cons \<Theta> \<B> \<Gamma> \<Delta> \<tau> u)
- then show ?case using wf_replace_inside1 Wellformed.wfD_emptyI
- by (simp add: wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.wfD_cons)
-next
- case (wfFTNone \<Theta> \<Phi> ft)
- then show ?case using wf_replace_inside1 Wellformed.wfD_emptyI by metis
-next
- case (wfFTSome \<Theta> \<Phi> bv ft)
- then show ?case using wf_replace_inside1 Wellformed.wfD_emptyI by metis
-qed(auto)
-
-lemmas wf_replace_inside = wf_replace_inside1 wf_replace_inside2
-
-lemma wfC_replace_cons:
- assumes "wfG P \<B> ((x,b,c1) #\<^sub>\<Gamma>\<Gamma>)" and "wfC P \<B> ((x,b,TRUE) #\<^sub>\<Gamma>\<Gamma>) c2"
- shows "wfC P \<B> ((x,b,c1) #\<^sub>\<Gamma>\<Gamma>) c2"
-proof -
- have "wfC P \<B> (GNil@((x,b,c1) #\<^sub>\<Gamma>\<Gamma>)) c2" proof(rule wf_replace_inside1(2))
- show " P; \<B> ; (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f c2 " using wfG_elim2 assms by auto
- show \<open>(x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> = GNil @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>\<close> using append_g.simps by auto
- show \<open>P; \<B> ; (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f c1 \<close> using wfG_elim2 assms by auto
- qed
- thus ?thesis using append_g.simps by auto
-qed
-
-lemma wfC_refl:
- assumes "wfG \<Theta> \<B> ((x, b', c') #\<^sub>\<Gamma>\<Gamma>)"
- shows "wfC \<Theta> \<B> ((x, b', c') #\<^sub>\<Gamma>\<Gamma>) c'"
- using wfG_wfC assms wfC_replace_cons by auto
-
-lemma wfG_wfC_inside:
- assumes " (x, b, c) \<in> toSet G" and "wfG \<Theta> B G"
- shows "wfC \<Theta> B G c"
- using assms proof(induct G rule: \<Gamma>_induct)
- case GNil
- then show ?case by auto
-next
- case (GCons x' b' c' \<Gamma>')
- then consider (hd) "(x, b, c) = (x',b',c')" | (tail) "(x, b, c) \<in> toSet \<Gamma>'" using toSet.simps by auto
- then show ?case proof(cases)
- case hd
- then show ?thesis using GCons wf_weakening
- by (metis wfC_replace_cons wfG_cons_wfC)
- next
- case tail
- then show ?thesis using GCons wf_weakening
- by (metis insert_iff insert_is_Un subsetI toSet.simps(2) wfG_cons2)
- qed
-qed
-
-lemma wfT_wf_cons3:
- assumes "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b | c \<rbrace>" and "atom y \<sharp> (c,\<Gamma>)"
- shows "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (y, b, c[z::=V_var y]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>"
-proof -
- have "\<lbrace> z : b | c \<rbrace> = \<lbrace> y : b | (y \<leftrightarrow> z) \<bullet> c \<rbrace>" using type_eq_flip assms by auto
- moreover hence " (y \<leftrightarrow> z) \<bullet> c = c[z::=V_var y]\<^sub>c\<^sub>v" using assms subst_v_c_def by auto
- ultimately have "\<lbrace> z : b | c \<rbrace> = \<lbrace> y : b | c[z::=V_var y]\<^sub>c\<^sub>v \<rbrace>" by metis
- thus ?thesis using assms wfT_wf_cons[of \<Theta> \<B> \<Gamma> y b] fresh_Pair by metis
-qed
-
-lemma wfT_wfC_cons:
- assumes "wfT P \<B> \<Gamma> \<lbrace> z1 : b | c1 \<rbrace>" and "wfT P \<B> \<Gamma> \<lbrace> z2 : b | c2 \<rbrace>" and "atom x \<sharp> (c1,c2,\<Gamma>)"
- shows "wfC P \<B> ((x,b,c1[z1::=V_var x]\<^sub>v) #\<^sub>\<Gamma>\<Gamma>) (c2[z2::=V_var x]\<^sub>v)" (is "wfC P \<B> ?G ?c")
-proof -
- have eq: "\<lbrace> z2 : b | c2 \<rbrace> = \<lbrace> x : b | c2[z2::=V_var x]\<^sub>c\<^sub>v \<rbrace>" using type_eq_subst assms fresh_prod3 by simp
- have eq2: "\<lbrace> z1 : b | c1 \<rbrace> = \<lbrace> x : b | c1[z1::=V_var x]\<^sub>c\<^sub>v \<rbrace>" using type_eq_subst assms fresh_prod3 by simp
- moreover have "wfT P \<B> \<Gamma> \<lbrace> x : b | c1[z1::=V_var x]\<^sub>c\<^sub>v \<rbrace>" using assms eq2 by auto
- moreover hence "wfG P \<B> ((x,b,c1[z1::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma>\<Gamma>)" using wfT_wf_cons fresh_prod3 assms by auto
- moreover have "wfT P \<B> \<Gamma> \<lbrace> x : b | c2[z2::=V_var x]\<^sub>c\<^sub>v \<rbrace>" using assms eq by auto
- moreover hence "wfC P \<B> ((x,b,TRUE) #\<^sub>\<Gamma>\<Gamma>) (c2[z2::=V_var x]\<^sub>c\<^sub>v)" using wfT_wfC assms fresh_prod3 by simp
- ultimately show ?thesis using wfC_replace_cons subst_v_c_def by simp
-qed
-
-lemma wfT_wfC2:
- fixes c::c and x::x
- assumes "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b | c \<rbrace>" and "atom x \<sharp> \<Gamma>"
- shows "\<Theta>; \<B>; (x,b,TRUE)#\<^sub>\<Gamma>\<Gamma> \<turnstile>\<^sub>w\<^sub>f c[z::=[x]\<^sup>v]\<^sub>v"
-proof(cases "x=z")
- case True
- then show ?thesis using wfT_wfC assms by auto
-next
- case False
- hence "atom x \<sharp> c" using wfT_fresh_c assms by metis
- hence "\<lbrace> x : b | c[z::=[ x ]\<^sup>v]\<^sub>v \<rbrace> = \<lbrace> z : b | c \<rbrace>"
- using \<tau>.eq_iff Abs1_eq_iff(3)[of x "c[z::=[ x ]\<^sup>v]\<^sub>v" z c]
- by (metis flip_subst_v type_eq_flip)
- hence " \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> x : b | c[z::=[ x ]\<^sup>v]\<^sub>v \<rbrace>" using assms by metis
- thus ?thesis using wfT_wfC assms by auto
-qed
-
-lemma wfT_wfG:
- fixes x::x and \<Gamma>::\<Gamma> and z::x and c::c and b::b
- assumes "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b | c \<rbrace>" and "atom x \<sharp> \<Gamma>"
- shows "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (x,b, c[z::=[ x ]\<^sup>v]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>"
-proof -
- have "\<Theta>; \<B>; (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f c[z::=[ x ]\<^sup>v]\<^sub>v" using wfT_wfC2 assms by metis
- thus ?thesis using wfG_consI assms wfT_wfB b_of.simps wfX_wfY by metis
-qed
-
-lemma wfG_replace_inside2:
- fixes \<Gamma>::\<Gamma>
- assumes "wfG P \<B> (\<Gamma>' @ (x, b, c') #\<^sub>\<Gamma> \<Gamma>)" and "wfG P \<B> ((x,b,c) #\<^sub>\<Gamma>\<Gamma>)"
- shows "wfG P \<B> (\<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>)"
-proof -
- have "wfC P \<B> ((x,b,TRUE) #\<^sub>\<Gamma>\<Gamma>) c" using wfG_wfC assms by auto
- thus ?thesis using wf_replace_inside1(3)[OF assms(1)] by auto
-qed
-
-lemma wfG_replace_inside_full:
- fixes \<Gamma>::\<Gamma>
- assumes "wfG P \<B> (\<Gamma>' @ (x, b, c') #\<^sub>\<Gamma> \<Gamma>)" and "wfG P \<B> (\<Gamma>'@((x,b,c) #\<^sub>\<Gamma>\<Gamma>))"
- shows "wfG P \<B> (\<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>)"
-proof -
- have "wfG P \<B> ((x,b,c) #\<^sub>\<Gamma>\<Gamma>)" using wfG_suffix assms by auto
- thus ?thesis using wfG_replace_inside assms by auto
-qed
-
-lemma wfT_replace_inside2:
- assumes "wfT \<Theta> \<B> (\<Gamma>' @ (x, b, c') #\<^sub>\<Gamma> \<Gamma>) t" and "wfG \<Theta> \<B> (\<Gamma>'@((x,b,c) #\<^sub>\<Gamma>\<Gamma>))"
- shows "wfT \<Theta> \<B> (\<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>) t"
-proof -
- have "wfG \<Theta> \<B> (((x,b,c) #\<^sub>\<Gamma>\<Gamma>))" using wfG_suffix assms by auto
- hence "wfC \<Theta> \<B> ((x,b,TRUE) #\<^sub>\<Gamma>\<Gamma>) c" using wfG_wfC by auto
- thus ?thesis using wf_replace_inside assms by metis
-qed
-
-lemma wfD_unique:
- assumes "wfD P \<B> \<Gamma> \<Delta>" and " (u,\<tau>') \<in> setD \<Delta>" and "(u,\<tau>) \<in> setD \<Delta>"
- shows "\<tau>'=\<tau>"
-using assms proof(induct \<Delta> rule: \<Delta>_induct)
- case DNil
- then show ?case by auto
-next
- case (DCons u' t' D)
- hence *: "wfD P \<B> \<Gamma> ((u',t') #\<^sub>\<Delta> D)" using Cons by auto
- show ?case proof(cases "u=u'")
- case True
- then have "u \<notin> fst ` setD D" using wfD_elims * by blast
- then show ?thesis using DCons by force
- next
- case False
- then show ?thesis using DCons wfD_elims * by (metis fst_conv setD_ConsD)
- qed
-qed
-
-lemma replace_in_g_forget:
- fixes x::x
- assumes "wfG P B G"
- shows "atom x \<notin> atom_dom G \<Longrightarrow> (G[x\<longmapsto>c]) = G" and
- "atom x \<sharp> G \<Longrightarrow> (G[x\<longmapsto>c]) = G"
-proof -
- show "atom x \<notin> atom_dom G \<Longrightarrow> G[x\<longmapsto>c] = G" by (induct G rule: \<Gamma>_induct,auto)
- thus "atom x \<sharp> G \<Longrightarrow> (G[x\<longmapsto>c]) = G" using wfG_x_fresh assms by simp
-qed
-
-lemma replace_in_g_fresh_single:
- fixes G::\<Gamma> and x::x
- assumes \<open>\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f G[x'\<longmapsto>c'']\<close> and "atom x \<sharp> G" and \<open>\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f G \<close>
- shows "atom x \<sharp> G[x'\<longmapsto>c'']"
- using rig_dom_eq wfG_dom_supp assms fresh_def atom_dom.simps dom.simps by metis
-
-section \<open>Preservation of well-formedness under substitution\<close>
-
-lemma wfC_cons_switch:
- fixes c::c and c'::c
- assumes "\<Theta>; \<B>; (x, b, c) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f c'"
- shows "\<Theta>; \<B>; (x, b, c') #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f c"
-proof -
- have *:"\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (x, b, c) #\<^sub>\<Gamma> \<Gamma>" using wfC_wf assms by auto
- hence "atom x \<sharp> \<Gamma> \<and> wfG \<Theta> \<B> \<Gamma> \<and> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b" using wfG_cons by auto
- hence " \<Theta>; \<B>; (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f TRUE " using wfC_trueI wfG_cons2I by simp
- hence "\<Theta>; \<B>;(x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f c'"
- using wf_replace_inside1(2)[of \<Theta> \<B> "(x, b, c) #\<^sub>\<Gamma> \<Gamma>" c' GNil x b c \<Gamma> TRUE] assms by auto
- hence "wfG \<Theta> \<B> ((x,b,c') #\<^sub>\<Gamma>\<Gamma>)" using wf_replace_inside1(3)[OF *, of GNil x b c \<Gamma> c'] by auto
- moreover have "wfC \<Theta> \<B> ((x,b,TRUE) #\<^sub>\<Gamma>\<Gamma>) c" proof(cases "c \<in> { TRUE, FALSE }")
- case True
- have "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<and> atom x \<sharp> \<Gamma> \<and> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b" using wfG_elims(2)[OF *] by auto
- hence "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (x,b,TRUE) #\<^sub>\<Gamma> \<Gamma>" using wfG_cons_TRUE by auto
- then show ?thesis using wfC_trueI wfC_falseI True by auto
- next
- case False
- then show ?thesis using wfG_elims(2)[OF *] by auto
- qed
- ultimately show ?thesis using wfC_replace_cons by auto
-qed
-
-lemma subst_g_inside_simple:
- fixes \<Gamma>\<^sub>1::\<Gamma> and \<Gamma>\<^sub>2::\<Gamma>
- assumes "wfG P \<B> (\<Gamma>\<^sub>1@((x,b,c) #\<^sub>\<Gamma>\<Gamma>\<^sub>2))"
- shows "(\<Gamma>\<^sub>1@((x,b,c) #\<^sub>\<Gamma>\<Gamma>\<^sub>2))[x::=v]\<^sub>\<Gamma>\<^sub>v = \<Gamma>\<^sub>1[x::=v]\<^sub>\<Gamma>\<^sub>v@\<Gamma>\<^sub>2"
-using assms proof(induct \<Gamma>\<^sub>1 rule: \<Gamma>_induct)
- case GNil
- then show ?case using subst_gv.simps by simp
-next
- case (GCons x' b' c' G)
- hence *:"P; \<B> \<turnstile>\<^sub>w\<^sub>f (x', b', c') #\<^sub>\<Gamma> (G @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>\<^sub>2)" by auto
- hence "x\<noteq>x'"
- using GCons append_Cons wfG_cons_fresh2[OF *] by auto
- hence "((GCons (x', b', c') G) @ (GCons (x, b, c) \<Gamma>\<^sub>2))[x::=v]\<^sub>\<Gamma>\<^sub>v =
- (GCons (x', b', c') (G @ (GCons (x, b, c) \<Gamma>\<^sub>2)))[x::=v]\<^sub>\<Gamma>\<^sub>v" by auto
- also have "... = GCons (x', b', c'[x::=v]\<^sub>c\<^sub>v) ((G @ (GCons (x, b, c) \<Gamma>\<^sub>2))[x::=v]\<^sub>\<Gamma>\<^sub>v)"
- using subst_gv.simps \<open>x\<noteq>x'\<close> by simp
- also have "... = (x', b', c'[x::=v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> (G[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>\<^sub>2)" using GCons * wfG_elims by metis
- also have "... = ((x', b', c') #\<^sub>\<Gamma> G)[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>\<^sub>2" using subst_gv.simps \<open>x\<noteq>x'\<close> by simp
- finally show ?case by blast
-qed
-
-lemma subst_c_TRUE_FALSE:
- fixes c::c
- assumes "c \<notin> {TRUE,FALSE}"
- shows "c[x::=v']\<^sub>c\<^sub>v \<notin> {TRUE, FALSE}"
-using assms by(nominal_induct c rule:c.strong_induct,auto simp add: subst_cv.simps)
-
-lemma lookup_subst:
- assumes "Some (b, c) = lookup \<Gamma> x" and "x \<noteq> x'"
- shows "\<exists>c'. Some (b,c') = lookup \<Gamma>[x'::=v']\<^sub>\<Gamma>\<^sub>v x"
-using assms proof(induct \<Gamma> rule: \<Gamma>_induct)
-case GNil
- then show ?case by auto
-next
- case (GCons x1 b1 c1 \<Gamma>1)
- then show ?case proof(cases "x1=x'")
- case True
- then show ?thesis using subst_gv.simps GCons by auto
- next
- case False
- hence *:"((x1, b1, c1) #\<^sub>\<Gamma> \<Gamma>1)[x'::=v']\<^sub>\<Gamma>\<^sub>v = ((x1, b1, c1[x'::=v']\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>1[x'::=v']\<^sub>\<Gamma>\<^sub>v)" using subst_gv.simps by auto
- then show ?thesis proof(cases "x1=x")
- case True
- then show ?thesis using lookup.simps *
- using GCons.prems(1) by auto
- next
- case False
- then show ?thesis using lookup.simps *
- using GCons.prems(1) by (simp add: GCons.hyps assms(2))
- qed
- qed
-qed
-
-lemma lookup_subst2:
- assumes "Some (b, c) = lookup (\<Gamma>'@((x',b\<^sub>1,c0[z0::=[x']\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\<Gamma>\<Gamma>)) x" and "x \<noteq> x'" and
- "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (\<Gamma>'@((x',b\<^sub>1,c0[z0::=[x']\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\<Gamma>\<Gamma>))"
- shows "\<exists>c'. Some (b,c') = lookup (\<Gamma>'[x'::=v']\<^sub>\<Gamma>\<^sub>v@\<Gamma>) x"
- using assms lookup_subst subst_g_inside by metis
-
-lemma wf_subst1:
- fixes \<Gamma>::\<Gamma> and \<Gamma>'::\<Gamma> and v::v and e::e and c::c and \<tau>::\<tau> and ts::"(string*\<tau>) list" and \<Delta>::\<Delta> and b::b and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def
- shows wfV_subst: "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b \<Longrightarrow> \<Gamma>=\<Gamma>\<^sub>1@((x,b',c') #\<^sub>\<Gamma>\<Gamma>\<^sub>2) \<Longrightarrow> \<Theta>; \<B>;\<Gamma>\<^sub>2 \<turnstile>\<^sub>w\<^sub>f v' : b' \<Longrightarrow> \<Theta> ; \<B> ; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v \<turnstile>\<^sub>w\<^sub>f v[x::=v']\<^sub>v\<^sub>v : b" and
- wfC_subst: "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f c \<Longrightarrow> \<Gamma>=\<Gamma>\<^sub>1@((x,b',c') #\<^sub>\<Gamma>\<Gamma>\<^sub>2) \<Longrightarrow> \<Theta>; \<B>; \<Gamma>\<^sub>2 \<turnstile>\<^sub>w\<^sub>f v' : b' \<Longrightarrow> \<Theta>; \<B>; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v \<turnstile>\<^sub>w\<^sub>f c[x::=v']\<^sub>c\<^sub>v" and
- wfG_subst: "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<Longrightarrow> \<Gamma>=\<Gamma>\<^sub>1@((x,b',c') #\<^sub>\<Gamma>\<Gamma>\<^sub>2) \<Longrightarrow> \<Theta>; \<B> ; \<Gamma>\<^sub>2 \<turnstile>\<^sub>w\<^sub>f v' : b' \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v" and
- wfT_subst: "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau> \<Longrightarrow> \<Gamma>=\<Gamma>\<^sub>1@((x,b',c') #\<^sub>\<Gamma>\<Gamma>\<^sub>2) \<Longrightarrow> \<Theta>; \<B> ; \<Gamma>\<^sub>2 \<turnstile>\<^sub>w\<^sub>f v' : b' \<Longrightarrow> \<Theta>; \<B>; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v \<turnstile>\<^sub>w\<^sub>f \<tau>[x::=v']\<^sub>\<tau>\<^sub>v" and
- "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ts \<Longrightarrow> True" and
- "\<turnstile>\<^sub>w\<^sub>f \<Theta> \<Longrightarrow>True" and
- "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b \<Longrightarrow> True " and
- wfCE_subst: "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ce : b \<Longrightarrow> \<Gamma>=\<Gamma>\<^sub>1@((x,b',c') #\<^sub>\<Gamma>\<Gamma>\<^sub>2) \<Longrightarrow> \<Theta>; \<B> ; \<Gamma>\<^sub>2 \<turnstile>\<^sub>w\<^sub>f v' : b' \<Longrightarrow> \<Theta> ; \<B> ; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v \<turnstile>\<^sub>w\<^sub>f ce[x::=v']\<^sub>c\<^sub>e\<^sub>v : b" and
- "\<Theta> \<turnstile>\<^sub>w\<^sub>f td \<Longrightarrow> True"
-proof(nominal_induct
- b and c and \<Gamma> and \<tau> and ts and \<Theta> and b and b and td
- avoiding: x v'
- arbitrary: \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1
- rule:wfV_wfC_wfG_wfT_wfTs_wfTh_wfB_wfCE_wfTD.strong_induct)
- case (wfV_varI \<Theta> \<B> \<Gamma> b1 c1 x1)
-
- show ?case proof(cases "x1=x")
- case True
- hence "(V_var x1)[x::=v']\<^sub>v\<^sub>v = v' " using subst_vv.simps by auto
- moreover have "b' = b1" using wfV_varI True lookup_inside_wf
- by (metis option.inject prod.inject)
- moreover have " \<Theta>; \<B> ; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v \<turnstile>\<^sub>w\<^sub>f v' : b'" using wfV_varI subst_g_inside_simple wf_weakening
- append_g_toSetU sup_ge2 wfV_wf by metis
- ultimately show ?thesis by auto
- next
- case False
- hence "(V_var x1)[x::=v']\<^sub>v\<^sub>v = (V_var x1) " using subst_vv.simps by auto
- moreover have "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v" using wfV_varI by simp
- moreover obtain c1' where "Some (b1, c1') = lookup \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v x1" using wfV_varI False lookup_subst by metis
- ultimately show ?thesis using Wellformed.wfV_varI[of \<Theta> \<B> "\<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v" b1 c1' x1] by metis
- qed
-next
- case (wfV_litI \<Theta> \<Gamma> l)
- then show ?case using subst_vv.simps wf_intros by auto
-next
- case (wfV_pairI \<Theta> \<Gamma> v1 b1 v2 b2)
- then show ?case using subst_vv.simps wf_intros by auto
-next
- case (wfV_consI s dclist \<Theta> dc x b c \<Gamma> v)
- then show ?case using subst_vv.simps wf_intros by auto
-next
- case (wfV_conspI s bv dclist \<Theta> dc x' b' c \<B> b \<Gamma> va)
- show ?case unfolding subst_vv.simps proof
- show \<open>AF_typedef_poly s bv dclist \<in> set \<Theta>\<close> and \<open>(dc, \<lbrace> x' : b' | c \<rbrace>) \<in> set dclist\<close> using wfV_conspI by auto
- show \<open> \<Theta> ;\<B> \<turnstile>\<^sub>w\<^sub>f b \<close> using wfV_conspI by auto
- have "atom bv \<sharp> \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v" using fresh_subst_gv_if wfV_conspI by metis
- moreover have "atom bv \<sharp> va[x::=v']\<^sub>v\<^sub>v" using wfV_conspI fresh_subst_if by simp
- ultimately show \<open>atom bv \<sharp> (\<Theta>, \<B>, \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v, b, va[x::=v']\<^sub>v\<^sub>v)\<close> unfolding fresh_prodN using wfV_conspI by auto
- show \<open> \<Theta>; \<B>; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v \<turnstile>\<^sub>w\<^sub>f va[x::=v']\<^sub>v\<^sub>v : b'[bv::=b]\<^sub>b\<^sub>b \<close> using wfV_conspI by auto
- qed
-next
- case (wfTI z \<Theta> \<B> \<Gamma> b c)
- have " \<Theta>; \<B>; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b | c[x::=v']\<^sub>c\<^sub>v \<rbrace>" proof
- have \<open>\<Theta>; \<B>; ((z, b, TRUE) #\<^sub>\<Gamma> \<Gamma>)[x::=v']\<^sub>\<Gamma>\<^sub>v \<turnstile>\<^sub>w\<^sub>f c[x::=v']\<^sub>c\<^sub>v \<close>
- proof(rule wfTI(9))
- show \<open>(z, b, TRUE) #\<^sub>\<Gamma> \<Gamma> = ((z, b, TRUE) #\<^sub>\<Gamma> \<Gamma>\<^sub>1) @ (x, b', c') #\<^sub>\<Gamma> \<Gamma>\<^sub>2\<close> using wfTI append_g.simps by simp
- show \<open> \<Theta>; \<B>; \<Gamma>\<^sub>2 \<turnstile>\<^sub>w\<^sub>f v' : b' \<close> using wfTI by auto
- qed
- thus *:\<open>\<Theta>; \<B>; (z, b, TRUE) #\<^sub>\<Gamma> \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v \<turnstile>\<^sub>w\<^sub>f c[x::=v']\<^sub>c\<^sub>v \<close>
- using subst_gv.simps subst_cv.simps wfTI fresh_x_neq by auto
-
- have "atom z \<sharp> \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v" using fresh_subst_gv_if wfTI by metis
- moreover have "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v" using wfTI wfX_wfY wfG_elims subst_gv.simps * by metis
- ultimately show \<open>atom z \<sharp> (\<Theta>, \<B>, \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v)\<close> using wfG_fresh_x by metis
- show \<open> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b \<close> using wfTI by auto
- qed
- thus ?case using subst_tv.simps wfTI by auto
-next
- case (wfC_trueI \<Theta> \<Gamma>)
- then show ?case using subst_cv.simps wf_intros by auto
-next
- case (wfC_falseI \<Theta> \<Gamma>)
- then show ?case using subst_cv.simps wf_intros by auto
-next
- case (wfC_eqI \<Theta> \<B> \<Gamma> e1 b e2)
- show ?case proof(subst subst_cv.simps,rule)
- show "\<Theta>; \<B>; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v \<turnstile>\<^sub>w\<^sub>f e1[x::=v']\<^sub>c\<^sub>e\<^sub>v : b " using wfC_eqI subst_dv.simps by auto
- show "\<Theta>; \<B>; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v \<turnstile>\<^sub>w\<^sub>f e2[x::=v']\<^sub>c\<^sub>e\<^sub>v : b " using wfC_eqI by auto
- qed
-next
- case (wfC_conjI \<Theta> \<Gamma> c1 c2)
- then show ?case using subst_cv.simps wf_intros by auto
-next
- case (wfC_disjI \<Theta> \<Gamma> c1 c2)
- then show ?case using subst_cv.simps wf_intros by auto
-next
- case (wfC_notI \<Theta> \<Gamma> c1)
- then show ?case using subst_cv.simps wf_intros by auto
-next
- case (wfC_impI \<Theta> \<Gamma> c1 c2)
- then show ?case using subst_cv.simps wf_intros by auto
-next
- case (wfG_nilI \<Theta>)
- then show ?case using subst_cv.simps wf_intros by auto
-next
- case (wfG_cons1I c \<Theta> \<B> \<Gamma> y b)
-
- show ?case proof(cases "x=y")
- case True
- hence "((y, b, c) #\<^sub>\<Gamma> \<Gamma>)[x::=v']\<^sub>\<Gamma>\<^sub>v = \<Gamma>" using subst_gv.simps by auto
- moreover have "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>" using wfG_cons1I by auto
- ultimately show ?thesis by auto
- next
- case False
- have "\<Gamma>\<^sub>1 \<noteq> GNil" using wfG_cons1I False by auto
- then obtain G where "\<Gamma>\<^sub>1 = (y, b, c) #\<^sub>\<Gamma> G" using GCons_eq_append_conv wfG_cons1I by auto
- hence *:"\<Gamma> = G @ (x, b', c') #\<^sub>\<Gamma> \<Gamma>\<^sub>2" using wfG_cons1I by auto
- hence "((y, b, c) #\<^sub>\<Gamma> \<Gamma>)[x::=v']\<^sub>\<Gamma>\<^sub>v =(y, b, c[x::=v']\<^sub>c\<^sub>v) #\<^sub>\<Gamma>\<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v" using subst_gv.simps False by auto
- moreover have "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (y, b, c[x::=v']\<^sub>c\<^sub>v) #\<^sub>\<Gamma>\<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v" proof(rule Wellformed.wfG_cons1I)
- show \<open>c[x::=v']\<^sub>c\<^sub>v \<notin> {TRUE, FALSE}\<close> using wfG_cons1I subst_c_TRUE_FALSE by auto
- show \<open> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v \<close> using wfG_cons1I * by auto
- have "\<Gamma> = (G @ ((x, b', c') #\<^sub>\<Gamma>GNil)) @ \<Gamma>\<^sub>2" using * append_g_assoc by auto
- hence "atom y \<sharp> \<Gamma>\<^sub>2" using fresh_suffix \<open>atom y \<sharp> \<Gamma>\<close> by auto
- hence "atom y \<sharp> v'" using wfG_cons1I wfV_x_fresh by metis
- thus \<open>atom y \<sharp> \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v\<close> using fresh_subst_gv wfG_cons1I by auto
- have "((y, b, TRUE) #\<^sub>\<Gamma> \<Gamma>)[x::=v']\<^sub>\<Gamma>\<^sub>v = (y, b, TRUE) #\<^sub>\<Gamma> \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v" using subst_gv.simps subst_cv.simps False by auto
- thus \<open> \<Theta>; \<B>; (y, b, TRUE) #\<^sub>\<Gamma> \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v \<turnstile>\<^sub>w\<^sub>f c[x::=v']\<^sub>c\<^sub>v \<close> using wfG_cons1I(6)[of "(y,b,TRUE) #\<^sub>\<Gamma>G"] * subst_gv.simps
- wfG_cons1I by fastforce
- show "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b " using wfG_cons1I by auto
- qed
- ultimately show ?thesis by auto
- qed
-next
- case (wfG_cons2I c \<Theta> \<B> \<Gamma> y b)
- show ?case proof(cases "x=y")
- case True
- hence "((y, b, c) #\<^sub>\<Gamma> \<Gamma>)[x::=v']\<^sub>\<Gamma>\<^sub>v = \<Gamma>" using subst_gv.simps by auto
- moreover have "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>" using wfG_cons2I by auto
- ultimately show ?thesis by auto
- next
- case False
- have "\<Gamma>\<^sub>1 \<noteq> GNil" using wfG_cons2I False by auto
- then obtain G where "\<Gamma>\<^sub>1 = (y, b, c) #\<^sub>\<Gamma> G" using GCons_eq_append_conv wfG_cons2I by auto
- hence *:"\<Gamma> = G @ (x, b', c') #\<^sub>\<Gamma> \<Gamma>\<^sub>2" using wfG_cons2I by auto
- hence "((y, b, c) #\<^sub>\<Gamma> \<Gamma>)[x::=v']\<^sub>\<Gamma>\<^sub>v =(y, b, c[x::=v']\<^sub>c\<^sub>v) #\<^sub>\<Gamma>\<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v" using subst_gv.simps False by auto
- moreover have "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (y, b, c[x::=v']\<^sub>c\<^sub>v) #\<^sub>\<Gamma>\<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v" proof(rule Wellformed.wfG_cons2I)
- show \<open>c[x::=v']\<^sub>c\<^sub>v \<in> {TRUE, FALSE}\<close> using subst_cv.simps wfG_cons2I by auto
- show \<open> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v \<close> using wfG_cons2I * by auto
- have "\<Gamma> = (G @ ((x, b', c') #\<^sub>\<Gamma>GNil)) @ \<Gamma>\<^sub>2" using * append_g_assoc by auto
- hence "atom y \<sharp> \<Gamma>\<^sub>2" using fresh_suffix wfG_cons2I by metis
- hence "atom y \<sharp> v'" using wfG_cons2I wfV_x_fresh by metis
- thus \<open>atom y \<sharp> \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v\<close> using fresh_subst_gv wfG_cons2I by auto
- show "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b " using wfG_cons2I by auto
- qed
- ultimately show ?thesis by auto
- qed
-next
- case (wfCE_valI \<Theta> \<B> \<Gamma> v b)
- then show ?case using subst_vv.simps wf_intros by auto
-next
- case (wfCE_plusI \<Theta> \<B> \<Gamma> v1 v2)
- then show ?case using subst_vv.simps wf_intros by auto
-next
- case (wfCE_leqI \<Theta> \<B> \<Gamma> v1 v2)
- then show ?case using subst_vv.simps wf_intros by auto
-next
- case (wfCE_eqI \<Theta> \<B> \<Gamma> v1 b v2)
- then show ?case unfolding subst_cev.simps
- using Wellformed.wfCE_eqI by metis
-next
- case (wfCE_fstI \<Theta> \<B> \<Gamma> v1 b1 b2)
- then show ?case using Wellformed.wfCE_fstI subst_cev.simps by metis
-next
- case (wfCE_sndI \<Theta> \<B> \<Gamma> v1 b1 b2)
- then show ?case using subst_cev.simps wf_intros by metis
-next
- case (wfCE_concatI \<Theta> \<B> \<Gamma> v1 v2)
- then show ?case using subst_vv.simps wf_intros by auto
-next
- case (wfCE_lenI \<Theta> \<B> \<Gamma> v1)
- then show ?case using subst_vv.simps wf_intros by auto
-qed(metis subst_sv.simps wf_intros)+
-
-lemma wf_subst2:
- fixes \<Gamma>::\<Gamma> and \<Gamma>'::\<Gamma> and v::v and e::e and c::c and \<tau>::\<tau> and ts::"(string*\<tau>) list" and \<Delta>::\<Delta> and b::b and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def
- shows "\<Theta>; \<Phi>; \<B>; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f e : b \<Longrightarrow> \<Gamma>=\<Gamma>\<^sub>1@((x,b',c') #\<^sub>\<Gamma>\<Gamma>\<^sub>2) \<Longrightarrow> \<Theta>; \<B> ; \<Gamma>\<^sub>2 \<turnstile>\<^sub>w\<^sub>f v' : b' \<Longrightarrow> \<Theta> ; \<Phi> ; \<B> ; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v \<turnstile>\<^sub>w\<^sub>f e[x::=v']\<^sub>e\<^sub>v : b" and
- "\<Theta>; \<Phi>; \<B>; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b \<Longrightarrow> \<Gamma>=\<Gamma>\<^sub>1@((x,b',c') #\<^sub>\<Gamma>\<Gamma>\<^sub>2) \<Longrightarrow> \<Theta> ;\<B> ; \<Gamma>\<^sub>2 \<turnstile>\<^sub>w\<^sub>f v' : b' \<Longrightarrow> \<Theta> ; \<Phi> ; \<B> ; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v \<turnstile>\<^sub>w\<^sub>f s[x::=v']\<^sub>s\<^sub>v : b" and
- "\<Theta>; \<Phi>; \<B>; \<Gamma> ; \<Delta> ; tid ; dc ; t \<turnstile>\<^sub>w\<^sub>f cs : b \<Longrightarrow> \<Gamma>=\<Gamma>\<^sub>1@((x,b',c') #\<^sub>\<Gamma>\<Gamma>\<^sub>2) \<Longrightarrow> \<Theta>; \<B>; \<Gamma>\<^sub>2 \<turnstile>\<^sub>w\<^sub>f v' : b' \<Longrightarrow> \<Theta>; \<Phi>; \<B>; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v ; tid ; dc ; t \<turnstile>\<^sub>w\<^sub>f subst_branchv cs x v' : b" and
- "\<Theta>; \<Phi>; \<B>; \<Gamma> ; \<Delta> ; tid ; dclist \<turnstile>\<^sub>w\<^sub>f css : b \<Longrightarrow> \<Gamma>=\<Gamma>\<^sub>1@((x,b',c') #\<^sub>\<Gamma>\<Gamma>\<^sub>2) \<Longrightarrow> \<Theta>; \<B>; \<Gamma>\<^sub>2 \<turnstile>\<^sub>w\<^sub>f v' : b' \<Longrightarrow> \<Theta>; \<Phi>; \<B>; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v ; tid ; dclist \<turnstile>\<^sub>w\<^sub>f subst_branchlv css x v' : b" and
- "\<Theta> \<turnstile>\<^sub>w\<^sub>f (\<Phi>::\<Phi>) \<Longrightarrow> True " and
- "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<Longrightarrow> \<Gamma>=\<Gamma>\<^sub>1@((x,b',c') #\<^sub>\<Gamma>\<Gamma>\<^sub>2) \<Longrightarrow> \<Theta>; \<B> ; \<Gamma>\<^sub>2 \<turnstile>\<^sub>w\<^sub>f v' : b' \<Longrightarrow> \<Theta> ; \<B> ; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v \<turnstile>\<^sub>w\<^sub>f \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v" and
- "\<Theta> ; \<Phi> \<turnstile>\<^sub>w\<^sub>f ftq \<Longrightarrow> True" and
- "\<Theta> ; \<Phi> ; \<B> \<turnstile>\<^sub>w\<^sub>f ft \<Longrightarrow> True"
-proof(nominal_induct
- b and b and b and b and \<Phi> and \<Delta> and ftq and ft
- avoiding: x v'
- arbitrary: \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1
- rule:wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.strong_induct)
- case (wfE_valI \<Theta> \<Gamma> v b)
- then show ?case using subst_vv.simps wf_intros wf_subst1
- by (metis subst_ev.simps(1))
-next
- case (wfE_plusI \<Theta> \<Gamma> v1 v2)
- then show ?case using subst_vv.simps wf_intros wf_subst1 by auto
-next
- case (wfE_leqI \<Theta> \<Phi> \<Gamma> \<Delta> v1 v2)
- then show ?case
- using subst_vv.simps subst_ev.simps subst_ev.simps wf_subst1 Wellformed.wfE_leqI
- by auto
-next
- case (wfE_eqI \<Theta> \<Phi> \<Gamma> \<Delta> v1 b v2)
- then show ?case
- using subst_vv.simps subst_ev.simps subst_ev.simps wf_subst1 Wellformed.wfE_eqI
- proof -
- show ?thesis
- by (metis (no_types) subst_ev.simps(4) wfE_eqI.hyps(1) wfE_eqI.hyps(4) wfE_eqI.hyps(5) wfE_eqI.hyps(6) wfE_eqI.hyps(7) wfE_eqI.prems(1) wfE_eqI.prems(2) wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.wfE_eqI wfV_subst) (* 31 ms *)
- qed
-next
- case (wfE_fstI \<Theta> \<Gamma> v1 b1 b2)
- then show ?case using subst_vv.simps subst_ev.simps wf_subst1 Wellformed.wfE_fstI
- proof -
- show ?thesis
- by (metis (full_types) subst_ev.simps(5) wfE_fstI.hyps(1) wfE_fstI.hyps(4) wfE_fstI.hyps(5) wfE_fstI.prems(1) wfE_fstI.prems(2) wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.wfE_fstI wf_subst1(1)) (* 78 ms *)
- qed
-next
- case (wfE_sndI \<Theta> \<Gamma> v1 b1 b2)
- then show ?case
- by (metis (full_types) subst_ev.simps wfE_sndI Wellformed.wfE_sndI wf_subst1(1))
-next
- case (wfE_concatI \<Theta> \<Phi> \<Gamma> \<Delta> v1 v2)
- then show ?case
- by (metis (full_types) subst_ev.simps wfE_sndI Wellformed.wfE_concatI wf_subst1(1))
-next
- case (wfE_splitI \<Theta> \<Phi> \<Gamma> \<Delta> v1 v2)
- then show ?case
- by (metis (full_types) subst_ev.simps wfE_sndI Wellformed.wfE_splitI wf_subst1(1))
-next
- case (wfE_lenI \<Theta> \<Phi> \<Gamma> \<Delta> v1)
-then show ?case
- by (metis (full_types) subst_ev.simps wfE_sndI Wellformed.wfE_lenI wf_subst1(1))
-next
- case (wfE_appI \<Theta> \<Phi> \<Gamma> \<Delta> f x b c \<tau> s' v)
-then show ?case
- by (metis (full_types) subst_ev.simps wfE_sndI Wellformed.wfE_appI wf_subst1(1))
-next
- case (wfE_appPI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> b' bv1 v1 \<tau>1 f1 x1 b1 c1 s1)
- show ?case proof(subst subst_ev.simps, rule)
- show "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>" using wfE_appPI wfX_wfY by metis
- show "\<Theta>; \<B>; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v \<turnstile>\<^sub>w\<^sub>f \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v " using wfE_appPI by auto
- show "Some (AF_fundef f1 (AF_fun_typ_some bv1 (AF_fun_typ x1 b1 c1 \<tau>1 s1))) = lookup_fun \<Phi> f1" using wfE_appPI by auto
- show "\<Theta>; \<B>; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v \<turnstile>\<^sub>w\<^sub>f v1[x::=v']\<^sub>v\<^sub>v : b1[bv1::=b']\<^sub>b " using wfE_appPI wf_subst1 by auto
- show "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b' " using wfE_appPI by auto
- have "atom bv1 \<sharp> \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v" using fresh_subst_gv_if wfE_appPI by metis
- moreover have "atom bv1 \<sharp> v1[x::=v']\<^sub>v\<^sub>v" using wfE_appPI fresh_subst_if by simp
- moreover have "atom bv1 \<sharp> \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v" using wfE_appPI fresh_subst_dv_if by simp
- ultimately show "atom bv1 \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v, \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v, b', v1[x::=v']\<^sub>v\<^sub>v, (b_of \<tau>1)[bv1::=b']\<^sub>b)"
- using wfE_appPI fresh_prodN by metis
- qed
-next
- case (wfE_mvarI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> u \<tau>)
- have " \<Theta> ; \<Phi> ; \<B> ; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v \<turnstile>\<^sub>w\<^sub>f (AE_mvar u) : b_of \<tau>[x::=v']\<^sub>\<tau>\<^sub>v" proof
- show "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> " using wfE_mvarI by auto
- show "\<Theta>; \<B> ; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v \<turnstile>\<^sub>w\<^sub>f \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v " using wfE_mvarI by auto
- show "(u, \<tau>[x::=v']\<^sub>\<tau>\<^sub>v) \<in> setD \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v" using wfE_mvarI subst_dv_member by auto
- qed
- thus ?case using subst_ev.simps b_of_subst by auto
-next
- case (wfD_emptyI \<Theta> \<Gamma>)
- then show ?case using subst_dv.simps wf_intros wf_subst1 by auto
-next
- case (wfD_cons \<Theta> \<B> \<Gamma> \<Delta> \<tau> u)
- moreover hence "u \<notin> fst ` setD \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v" using subst_dv.simps subst_dv_iff using subst_dv_fst_eq by presburger
- ultimately show ?case using subst_dv.simps Wellformed.wfD_cons wf_subst1 by auto
-next
- case (wfPhi_emptyI \<Theta>)
- then show ?case by auto
-next
- case (wfPhi_consI f \<Theta> \<Phi> ft)
- then show ?case by auto
-next
- case (wfS_assertI \<Theta> \<Phi> \<B> x2 c \<Gamma> \<Delta> s b)
- show ?case unfolding subst_sv.simps proof
- show \<open> \<Theta>; \<Phi>; \<B>; (x2, B_bool, c[x::=v']\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v \<turnstile>\<^sub>w\<^sub>f s[x::=v']\<^sub>s\<^sub>v : b \<close>
- using wfS_assertI(4)[of "(x2, B_bool, c) #\<^sub>\<Gamma> \<Gamma>\<^sub>1" x ] wfS_assertI by auto
-
- show \<open> \<Theta>; \<B>; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v \<turnstile>\<^sub>w\<^sub>f c[x::=v']\<^sub>c\<^sub>v \<close> using wfS_assertI wf_subst1 by auto
- show \<open> \<Theta>; \<B>; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v \<turnstile>\<^sub>w\<^sub>f \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v \<close> using wfS_assertI wf_subst1 by auto
- show \<open>atom x2 \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v, \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v, c[x::=v']\<^sub>c\<^sub>v, b, s[x::=v']\<^sub>s\<^sub>v)\<close>
- apply(unfold fresh_prodN,intro conjI)
- apply(simp add: wfS_assertI )+
- apply(metis fresh_subst_gv_if wfS_assertI)
- apply(simp add: fresh_prodN fresh_subst_dv_if wfS_assertI)
- apply(simp add: fresh_prodN fresh_subst_v_if subst_v_e_def wfS_assertI)
- apply(simp add: fresh_prodN fresh_subst_v_if subst_v_\<tau>_def wfS_assertI)
- by(simp add: fresh_prodN fresh_subst_v_if subst_v_s_def wfS_assertI)
- qed
-next
- case (wfS_letI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> e b1 y s b2)
- have "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v \<turnstile>\<^sub>w\<^sub>f LET y = (e[x::=v']\<^sub>e\<^sub>v) IN (s[x::=v']\<^sub>s\<^sub>v) : b2"
- proof
- show \<open> \<Theta> ; \<Phi> ; \<B> ; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v \<turnstile>\<^sub>w\<^sub>f e[x::=v']\<^sub>e\<^sub>v : b1 \<close> using wfS_letI by auto
- have \<open> \<Theta> ; \<Phi> ; \<B> ; ((y, b1, TRUE) #\<^sub>\<Gamma> \<Gamma>)[x::=v']\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v \<turnstile>\<^sub>w\<^sub>f s[x::=v']\<^sub>s\<^sub>v : b2 \<close>
- using wfS_letI(6) wfS_letI append_g.simps by metis
- thus \<open> \<Theta> ; \<Phi> ; \<B> ; (y, b1, TRUE) #\<^sub>\<Gamma> \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v \<turnstile>\<^sub>w\<^sub>f s[x::=v']\<^sub>s\<^sub>v : b2 \<close>
- using wfS_letI subst_gv.simps by auto
- show \<open> \<Theta>; \<B>; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v \<turnstile>\<^sub>w\<^sub>f \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v \<close> using wfS_letI by auto
- show \<open>atom y \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v, \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v, e[x::=v']\<^sub>e\<^sub>v, b2)\<close>
- apply(unfold fresh_prodN,intro conjI)
- apply(simp add: wfS_letI )+
- apply(metis fresh_subst_gv_if wfS_letI)
- apply(simp add: fresh_prodN fresh_subst_dv_if wfS_letI)
- apply(simp add: fresh_prodN fresh_subst_v_if subst_v_e_def wfS_letI)
- apply(simp add: fresh_prodN fresh_subst_v_if subst_v_\<tau>_def wfS_letI)
- done
- qed
- thus ?case using subst_sv.simps wfS_letI by auto
-next
- case (wfS_let2I \<Theta> \<Phi> \<B> \<Gamma> \<Delta> s1 \<tau> y s2 b)
- have "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v \<turnstile>\<^sub>w\<^sub>f LET y : \<tau>[x::=v']\<^sub>\<tau>\<^sub>v = (s1[x::=v']\<^sub>s\<^sub>v) IN (s2[x::=v']\<^sub>s\<^sub>v) : b"
- proof
- show \<open> \<Theta> ; \<Phi> ; \<B> ; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v \<turnstile>\<^sub>w\<^sub>f s1[x::=v']\<^sub>s\<^sub>v : b_of (\<tau>[x::=v']\<^sub>\<tau>\<^sub>v) \<close> using wfS_let2I b_of_subst by simp
- have \<open> \<Theta> ; \<Phi> ; \<B> ; ((y, b_of \<tau>, TRUE) #\<^sub>\<Gamma> \<Gamma>)[x::=v']\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v \<turnstile>\<^sub>w\<^sub>f s2[x::=v']\<^sub>s\<^sub>v : b \<close>
- using wfS_let2I append_g.simps by metis
- thus \<open> \<Theta> ; \<Phi> ; \<B> ; (y, b_of \<tau>[x::=v']\<^sub>\<tau>\<^sub>v, TRUE) #\<^sub>\<Gamma> \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v \<turnstile>\<^sub>w\<^sub>f s2[x::=v']\<^sub>s\<^sub>v : b \<close>
- using wfS_let2I subst_gv.simps append_g.simps using b_of_subst by simp
- show \<open> \<Theta>; \<B>; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v \<turnstile>\<^sub>w\<^sub>f \<tau>[x::=v']\<^sub>\<tau>\<^sub>v \<close> using wfS_let2I wf_subst1 by metis
- show \<open>atom y \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v, \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v, s1[x::=v']\<^sub>s\<^sub>v, b, \<tau>[x::=v']\<^sub>\<tau>\<^sub>v)\<close>
- apply(unfold fresh_prodN,intro conjI)
- apply(simp add: wfS_let2I )+
- apply(metis fresh_subst_gv_if wfS_let2I)
- apply(simp add: fresh_prodN fresh_subst_dv_if wfS_let2I)
- apply(simp add: fresh_prodN fresh_subst_v_if subst_v_e_def wfS_let2I)
- apply(simp add: fresh_prodN fresh_subst_v_if subst_v_\<tau>_def wfS_let2I)+
- done
- qed
- thus ?case using subst_sv.simps(3) subst_tv.simps wfS_let2I by auto
-next
- case (wfS_varI \<Theta> \<B> \<Gamma> \<tau> v u \<Phi> \<Delta> b s)
- show ?case proof(subst subst_sv.simps, auto simp add: u_fresh_xv,rule)
- show \<open> \<Theta>; \<B>; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v \<turnstile>\<^sub>w\<^sub>f \<tau>[x::=v']\<^sub>\<tau>\<^sub>v \<close> using wfS_varI wf_subst1 by auto
- have "b_of (\<tau>[x::=v']\<^sub>\<tau>\<^sub>v) = b_of \<tau>" using b_of_subst by auto
- thus \<open> \<Theta>; \<B>; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v \<turnstile>\<^sub>w\<^sub>f v[x::=v']\<^sub>v\<^sub>v : b_of \<tau>[x::=v']\<^sub>\<tau>\<^sub>v \<close> using wfS_varI wf_subst1 by auto
- have *:"atom u \<sharp> v'" using wfV_supp wfS_varI fresh_def by metis
- show \<open>atom u \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v, \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v, \<tau>[x::=v']\<^sub>\<tau>\<^sub>v, v[x::=v']\<^sub>v\<^sub>v, b)\<close>
- unfolding fresh_prodN apply(auto simp add: wfS_varI)
- using wfS_varI fresh_subst_gv * fresh_subst_dv by metis+
- show \<open> \<Theta> ; \<Phi> ; \<B> ; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v ; (u, \<tau>[x::=v']\<^sub>\<tau>\<^sub>v) #\<^sub>\<Delta> \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v \<turnstile>\<^sub>w\<^sub>f s[x::=v']\<^sub>s\<^sub>v : b \<close> using wfS_varI by auto
- qed
-next
- case (wfS_assignI u \<tau> \<Delta> \<Theta> \<B> \<Gamma> \<Phi> v)
- show ?case proof(subst subst_sv.simps, rule wf_intros)
- show \<open>(u, \<tau>[x::=v']\<^sub>\<tau>\<^sub>v) \<in> setD \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v\<close> using subst_dv_iff wfS_assignI using subst_dv_fst_eq
- using subst_dv_member by auto
- show \<open> \<Theta>; \<B>; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v \<turnstile>\<^sub>w\<^sub>f \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v \<close> using wfS_assignI by auto
- show \<open> \<Theta>; \<B>; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v \<turnstile>\<^sub>w\<^sub>f v[x::=v']\<^sub>v\<^sub>v : b_of \<tau>[x::=v']\<^sub>\<tau>\<^sub>v \<close> using wfS_assignI b_of_subst wf_subst1 by auto
- show "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> " using wfS_assignI by auto
- qed
-next
- case (wfS_matchI \<Theta> \<B> \<Gamma> v tid dclist \<Delta> \<Phi> cs b)
- show ?case proof(subst subst_sv.simps, rule wf_intros)
- show \<open> \<Theta>; \<B>; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v \<turnstile>\<^sub>w\<^sub>f v[x::=v']\<^sub>v\<^sub>v : B_id tid \<close> using wfS_matchI wf_subst1 by auto
- show \<open>AF_typedef tid dclist \<in> set \<Theta>\<close> using wfS_matchI by auto
- show \<open> \<Theta> ; \<Phi> ; \<B> ; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v ; tid ; dclist \<turnstile>\<^sub>w\<^sub>f subst_branchlv cs x v' : b \<close> using wfS_matchI by simp
- show "\<Theta>; \<B>; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v \<turnstile>\<^sub>w\<^sub>f \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v " using wfS_matchI by auto
- show "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> " using wfS_matchI by auto
- qed
-next
- case (wfS_branchI \<Theta> \<Phi> \<B> y \<tau> \<Gamma> \<Delta> s b tid dc)
- have " \<Theta> ; \<Phi> ; \<B> ; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v ; tid ; dc ; \<tau> \<turnstile>\<^sub>w\<^sub>f dc y \<Rightarrow> (s[x::=v']\<^sub>s\<^sub>v) : b"
- proof
- have \<open> \<Theta> ; \<Phi> ; \<B> ; ((y, b_of \<tau>, TRUE) #\<^sub>\<Gamma> \<Gamma>)[x::=v']\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v \<turnstile>\<^sub>w\<^sub>f s[x::=v']\<^sub>s\<^sub>v : b \<close>
- using wfS_branchI append_g.simps by metis
- thus \<open> \<Theta> ; \<Phi> ; \<B> ; (y, b_of \<tau>, TRUE) #\<^sub>\<Gamma> \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v \<turnstile>\<^sub>w\<^sub>f s[x::=v']\<^sub>s\<^sub>v : b \<close>
- using subst_gv.simps b_of_subst wfS_branchI by simp
- show \<open>atom y \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v, \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v, \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v, \<tau>)\<close>
- apply(unfold fresh_prodN,intro conjI)
- apply(simp add: wfS_branchI )+
- apply(metis fresh_subst_gv_if wfS_branchI)
- apply(simp add: fresh_prodN fresh_subst_dv_if wfS_branchI)
- apply(metis fresh_subst_gv_if wfS_branchI)+
- done
- show \<open> \<Theta>; \<B>; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v \<turnstile>\<^sub>w\<^sub>f \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v \<close> using wfS_branchI by auto
- qed
- thus ?case using subst_branchv.simps wfS_branchI by auto
-next
- case (wfS_finalI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist' cs b dclist)
- then show ?case using subst_branchlv.simps wf_intros by metis
-next
- case (wfS_cons \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist' cs b css dclist)
- then show ?case using subst_branchlv.simps wf_intros by metis
-
-qed(metis subst_sv.simps wf_subst1 wf_intros)+
-
-lemmas wf_subst = wf_subst1 wf_subst2
-
-lemma wfG_subst_wfV:
- assumes "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>' @ (x, b, c0[z0::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>" and "wfV \<Theta> \<B> \<Gamma> v b"
- shows "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> "
- using assms wf_subst subst_g_inside_simple by auto
-
-lemma wfG_member_subst:
- assumes "(x1,b1,c1) \<in> toSet (\<Gamma>'@\<Gamma>)" and "wfG \<Theta> \<B> (\<Gamma>'@((x,b,c) #\<^sub>\<Gamma>\<Gamma>))" and "x \<noteq> x1"
- shows "\<exists>c1'. (x1,b1,c1') \<in> toSet ((\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v)@\<Gamma>)"
-proof -
- consider (lhs) "(x1,b1,c1) \<in> toSet \<Gamma>'" | (rhs) "(x1,b1,c1) \<in> toSet \<Gamma>" using append_g_toSetU assms by auto
- thus ?thesis proof(cases)
- case lhs
- hence "(x1,b1,c1[x::=v]\<^sub>c\<^sub>v) \<in> toSet (\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v)" using wfG_inside_fresh[THEN subst_gv_member_iff[OF lhs]] assms by metis
- hence "(x1,b1,c1[x::=v]\<^sub>c\<^sub>v) \<in> toSet (\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v@\<Gamma>)" using append_g_toSetU by auto
- then show ?thesis by auto
- next
- case rhs
- hence "(x1,b1,c1) \<in> toSet (\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v@\<Gamma>)" using append_g_toSetU by auto
- then show ?thesis by auto
- qed
-qed
-
-lemma wfG_member_subst2:
- assumes "(x1,b1,c1) \<in> toSet (\<Gamma>'@((x,b,c) #\<^sub>\<Gamma>\<Gamma>))" and "wfG \<Theta> \<B> (\<Gamma>'@((x,b,c) #\<^sub>\<Gamma>\<Gamma>))" and "x \<noteq> x1"
- shows "\<exists>c1'. (x1,b1,c1') \<in> toSet ((\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v)@\<Gamma>)"
-proof -
- consider (lhs) "(x1,b1,c1) \<in> toSet \<Gamma>'" | (rhs) "(x1,b1,c1) \<in> toSet \<Gamma>" using append_g_toSetU assms by auto
- thus ?thesis proof(cases)
- case lhs
- hence "(x1,b1,c1[x::=v]\<^sub>c\<^sub>v) \<in> toSet (\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v)" using wfG_inside_fresh[THEN subst_gv_member_iff[OF lhs]] assms by metis
- hence "(x1,b1,c1[x::=v]\<^sub>c\<^sub>v) \<in> toSet (\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v@\<Gamma>)" using append_g_toSetU by auto
- then show ?thesis by auto
- next
- case rhs
- hence "(x1,b1,c1) \<in> toSet (\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v@\<Gamma>)" using append_g_toSetU by auto
- then show ?thesis by auto
- qed
-qed
-
-lemma wbc_subst:
- fixes \<Gamma>::\<Gamma> and \<Gamma>'::\<Gamma> and v::v
- assumes "wfC \<Theta> \<B> (\<Gamma>'@((x,b,c') #\<^sub>\<Gamma>\<Gamma>)) c" and "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b"
- shows "\<Theta>; \<B>; ((\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v)@\<Gamma>) \<turnstile>\<^sub>w\<^sub>f c[x::=v]\<^sub>c\<^sub>v"
-proof -
- have "(\<Gamma>'@((x,b,c') #\<^sub>\<Gamma>\<Gamma>))[x::=v]\<^sub>\<Gamma>\<^sub>v = ((\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v)@\<Gamma>)" using assms subst_g_inside_simple wfC_wf by metis
- thus ?thesis using wf_subst1(2)[OF assms(1) _ assms(2)] by metis
-qed
-
-lemma wfG_inside_fresh_suffix:
- assumes "wfG P B (\<Gamma>'@(x,b,c) #\<^sub>\<Gamma>\<Gamma>)"
- shows "atom x \<sharp> \<Gamma>"
-proof -
- have "wfG P B ((x,b,c) #\<^sub>\<Gamma>\<Gamma>)" using wfG_suffix assms by auto
- thus ?thesis using wfG_elims by metis
-qed
-
-lemmas wf_b_subst_lemmas = subst_eb.simps wf_intros
- forget_subst subst_b_b_def subst_b_v_def subst_b_ce_def fresh_e_opp_all subst_bb.simps wfV_b_fresh ms_fresh_all(6)
-
-lemma wf_b_subst1:
- fixes \<Gamma>::\<Gamma> and \<Gamma>'::\<Gamma> and v::v and e::e and c::c and \<tau>::\<tau> and ts::"(string*\<tau>) list" and \<Delta>::\<Delta> and b::b and ftq::fun_typ_q and ft::fun_typ and s::s and b'::b and ce::ce and td::type_def
- and cs::branch_s and css::branch_list
- shows "\<Theta> ; B' ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b' \<Longrightarrow> {|bv|} = B' \<Longrightarrow> \<Theta> ; B \<turnstile>\<^sub>w\<^sub>f b \<Longrightarrow> \<Theta> ; B ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f v[bv::=b]\<^sub>v\<^sub>b : b'[bv::=b]\<^sub>b\<^sub>b" and
- "\<Theta> ; B' ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f c \<Longrightarrow> {|bv|} = B' \<Longrightarrow> \<Theta> ; B \<turnstile>\<^sub>w\<^sub>f b \<Longrightarrow> \<Theta> ; B ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f c[bv::=b]\<^sub>c\<^sub>b" and
- "\<Theta> ; B' \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<Longrightarrow> {|bv|} = B' \<Longrightarrow> \<Theta> ; B \<turnstile>\<^sub>w\<^sub>f b \<Longrightarrow> \<Theta> ; B \<turnstile>\<^sub>w\<^sub>f \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b" and
- "\<Theta> ; B' ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau> \<Longrightarrow> {|bv|} = B' \<Longrightarrow> \<Theta> ; B \<turnstile>\<^sub>w\<^sub>f b \<Longrightarrow> \<Theta> ; B ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f \<tau>[bv::=b]\<^sub>\<tau>\<^sub>b" and
- "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ts \<Longrightarrow> True" and
- "\<turnstile>\<^sub>w\<^sub>f \<Theta> \<Longrightarrow>True" and
- "\<Theta> ; B' \<turnstile>\<^sub>w\<^sub>f b' \<Longrightarrow> {|bv|} = B' \<Longrightarrow> \<Theta> ; B \<turnstile>\<^sub>w\<^sub>f b \<Longrightarrow> \<Theta> ; B \<turnstile>\<^sub>w\<^sub>f b'[bv::=b]\<^sub>b\<^sub>b " and
- "\<Theta> ; B' ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ce : b' \<Longrightarrow> {|bv|} = B' \<Longrightarrow> \<Theta> ; B \<turnstile>\<^sub>w\<^sub>f b \<Longrightarrow> \<Theta> ; B ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f ce[bv::=b]\<^sub>c\<^sub>e\<^sub>b : b'[bv::=b]\<^sub>b\<^sub>b" and
- "\<Theta> \<turnstile>\<^sub>w\<^sub>f td \<Longrightarrow> True"
-proof(nominal_induct
- b' and c and \<Gamma> and \<tau> and ts and \<Theta> and b' and b' and td
- avoiding: bv b B
- rule:wfV_wfC_wfG_wfT_wfTs_wfTh_wfB_wfCE_wfTD.strong_induct)
- case (wfB_intI \<Theta> \<B>)
- then show ?case using subst_bb.simps wf_intros wfX_wfY by metis
-next
- case (wfB_boolI \<Theta> \<B>)
- then show ?case using subst_bb.simps wf_intros wfX_wfY by metis
-next
- case (wfB_unitI \<Theta> \<B>)
- then show ?case using subst_bb.simps wf_intros wfX_wfY by metis
-next
- case (wfB_bitvecI \<Theta> \<B>)
- then show ?case using subst_bb.simps wf_intros wfX_wfY by metis
-next
- case (wfB_pairI \<Theta> \<B> b1 b2)
- then show ?case using subst_bb.simps wf_intros wfX_wfY by metis
-next
- case (wfB_consI \<Theta> s dclist \<B>)
- then show ?case using subst_bb.simps Wellformed.wfB_consI by simp
-next
- case (wfB_appI \<Theta> ba s bva dclist \<B>)
- then show ?case using subst_bb.simps Wellformed.wfB_appI forget_subst wfB_supp
- by (metis bot.extremum_uniqueI ex_in_conv fresh_def subst_b_b_def supp_empty_fset)
-next
- case (wfV_varI \<Theta> \<B>1 \<Gamma> b1 c x)
- show ?case unfolding subst_vb.simps proof
- show "\<Theta> ; B \<turnstile>\<^sub>w\<^sub>f \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b " using wfV_varI by auto
- show "Some (b1[bv::=b]\<^sub>b\<^sub>b, c[bv::=b]\<^sub>c\<^sub>b) = lookup \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b x" using subst_b_lookup wfV_varI by simp
- qed
-next
- case (wfV_litI \<Theta> \<B> \<Gamma> l)
- then show ?case using Wellformed.wfV_litI subst_b_base_for_lit by simp
-next
- case (wfV_pairI \<Theta> \<B>1 \<Gamma> v1 b1 v2 b2)
- show ?case unfolding subst_vb.simps proof(subst subst_bb.simps,rule)
- show "\<Theta> ; B ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f v1[bv::=b]\<^sub>v\<^sub>b : b1[bv::=b]\<^sub>b\<^sub>b" using wfV_pairI by simp
- show "\<Theta> ; B ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f v2[bv::=b]\<^sub>v\<^sub>b : b2[bv::=b]\<^sub>b\<^sub>b " using wfV_pairI by simp
- qed
-next
- case (wfV_consI s dclist \<Theta> dc x b' c \<B>' \<Gamma> v)
- show ?case unfolding subst_vb.simps proof(subst subst_bb.simps, rule Wellformed.wfV_consI)
- show 1:"AF_typedef s dclist \<in> set \<Theta>" using wfV_consI by auto
- show 2:"(dc, \<lbrace> x : b' | c \<rbrace>) \<in> set dclist" using wfV_consI by auto
- have "\<Theta> ; B ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f v[bv::=b]\<^sub>v\<^sub>b : b'[bv::=b]\<^sub>b\<^sub>b" using wfV_consI by auto
- moreover hence "supp b' = {}" using 1 2 wfTh_lookup_supp_empty \<tau>.supp wfX_wfY by blast
- moreover hence "b'[bv::=b]\<^sub>b\<^sub>b = b'" using forget_subst subst_bb_def fresh_def by (metis empty_iff subst_b_b_def)
- ultimately show "\<Theta> ; B ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f v[bv::=b]\<^sub>v\<^sub>b : b'" using wfV_consI by simp
- qed
-next
- case (wfV_conspI s bva dclist \<Theta> dc x b' c \<B>' ba \<Gamma> v)
- have *:"atom bv \<sharp> b'" using wfTh_poly_supp_b[of s bva dclist \<Theta> dc x b' c] fresh_def wfX_wfY \<open>atom bva \<sharp> bv\<close>
- by (metis insert_iff not_self_fresh singleton_insert_inj_eq' subsetI subset_antisym wfV_conspI wfV_conspI.hyps(4) wfV_conspI.prems(2))
- show ?case unfolding subst_vb.simps subst_bb.simps proof
- show \<open>AF_typedef_poly s bva dclist \<in> set \<Theta>\<close> using wfV_conspI by auto
- show \<open>(dc, \<lbrace> x : b' | c \<rbrace>) \<in> set dclist\<close> using wfV_conspI by auto
- thus \<open> \<Theta> ; B \<turnstile>\<^sub>w\<^sub>f ba[bv::=b]\<^sub>b\<^sub>b \<close> using wfV_conspI by metis
- have "atom bva \<sharp> \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b" using fresh_subst_if subst_b_\<Gamma>_def wfV_conspI by metis
- moreover have "atom bva \<sharp> ba[bv::=b]\<^sub>b\<^sub>b" using fresh_subst_if subst_b_b_def wfV_conspI by metis
- moreover have "atom bva \<sharp> v[bv::=b]\<^sub>v\<^sub>b" using fresh_subst_if subst_b_v_def wfV_conspI by metis
- ultimately show \<open>atom bva \<sharp> (\<Theta>, B, \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b, ba[bv::=b]\<^sub>b\<^sub>b, v[bv::=b]\<^sub>v\<^sub>b)\<close>
- unfolding fresh_prodN using wfV_conspI fresh_def supp_fset by auto
- show \<open> \<Theta> ; B ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f v[bv::=b]\<^sub>v\<^sub>b : b'[bva::=ba[bv::=b]\<^sub>b\<^sub>b]\<^sub>b\<^sub>b \<close>
- using wfV_conspI subst_bb_commute[of bv b' bva ba b] * wfV_conspI by metis
- qed
-next
- case (wfTI z \<Theta> \<B>' \<Gamma>' b' c)
- show ?case proof(subst subst_tb.simps, rule Wellformed.wfTI)
- show "atom z \<sharp> (\<Theta>, B, \<Gamma>'[bv::=b]\<^sub>\<Gamma>\<^sub>b)" using wfTI subst_g_b_x_fresh by simp
- show "\<Theta> ; B \<turnstile>\<^sub>w\<^sub>f b'[bv::=b]\<^sub>b\<^sub>b " using wfTI by auto
- show "\<Theta> ; B ; (z, b'[bv::=b]\<^sub>b\<^sub>b, TRUE) #\<^sub>\<Gamma> \<Gamma>'[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f c[bv::=b]\<^sub>c\<^sub>b " using wfTI by simp
- qed
-next
- case (wfC_eqI \<Theta> \<B>' \<Gamma> e1 b' e2)
- thus ?case using Wellformed.wfC_eqI subst_db.simps subst_cb.simps wfC_eqI by metis
-next
- case (wfG_nilI \<Theta> \<B>')
- then show ?case using Wellformed.wfG_nilI subst_gb.simps by simp
-next
- case (wfG_cons1I c' \<Theta> \<B>' \<Gamma>' x b')
- show ?case proof(subst subst_gb.simps, rule Wellformed.wfG_cons1I)
- show "c'[bv::=b]\<^sub>c\<^sub>b \<notin> {TRUE, FALSE}" using wfG_cons1I(1)
- by(nominal_induct c' rule: c.strong_induct,auto+)
- show "\<Theta> ; B \<turnstile>\<^sub>w\<^sub>f \<Gamma>'[bv::=b]\<^sub>\<Gamma>\<^sub>b " using wfG_cons1I by auto
- show "atom x \<sharp> \<Gamma>'[bv::=b]\<^sub>\<Gamma>\<^sub>b" using wfG_cons1I subst_g_b_x_fresh by auto
- show "\<Theta> ; B ; (x, b'[bv::=b]\<^sub>b\<^sub>b, TRUE) #\<^sub>\<Gamma> \<Gamma>'[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f c'[bv::=b]\<^sub>c\<^sub>b" using wfG_cons1I by auto
- show "\<Theta> ; B \<turnstile>\<^sub>w\<^sub>f b'[bv::=b]\<^sub>b\<^sub>b " using wfG_cons1I by auto
- qed
-next
- case (wfG_cons2I c' \<Theta> \<B>' \<Gamma>' x b')
- show ?case proof(subst subst_gb.simps, rule Wellformed.wfG_cons2I)
- show "c'[bv::=b]\<^sub>c\<^sub>b \<in> {TRUE, FALSE}" using wfG_cons2I by auto
- show "\<Theta> ; B \<turnstile>\<^sub>w\<^sub>f \<Gamma>'[bv::=b]\<^sub>\<Gamma>\<^sub>b " using wfG_cons2I by auto
- show "atom x \<sharp> \<Gamma>'[bv::=b]\<^sub>\<Gamma>\<^sub>b" using wfG_cons2I subst_g_b_x_fresh by auto
- show "\<Theta> ; B \<turnstile>\<^sub>w\<^sub>f b'[bv::=b]\<^sub>b\<^sub>b " using wfG_cons2I by auto
- qed
-next
- case (wfCE_valI \<Theta> \<B> \<Gamma> v b)
- then show ?case using subst_ceb.simps wf_intros wfX_wfY
- by (metis wf_b_subst_lemmas wfCE_b_fresh)
-next
- case (wfCE_plusI \<Theta> \<B> \<Gamma> v1 v2)
- then show ?case using subst_bb.simps subst_ceb.simps wf_intros wfX_wfY
- by metis
-next
- case (wfCE_leqI \<Theta> \<B> \<Gamma> v1 v2)
- then show ?case using subst_bb.simps subst_ceb.simps wf_intros wfX_wfY
- by metis
-next
- case (wfCE_eqI \<Theta> \<B> \<Gamma> v1 b v2)
- then show ?case using subst_bb.simps subst_ceb.simps wf_intros wfX_wfY
- by metis
-next
- case (wfCE_fstI \<Theta> \<B> \<Gamma> v1 b1 b2)
- then show ?case
- by (metis (no_types) subst_bb.simps(5) subst_ceb.simps(3) wfCE_fstI.hyps(2)
- wfCE_fstI.prems(1) wfCE_fstI.prems(2) Wellformed.wfCE_fstI)
-next
- case (wfCE_sndI \<Theta> \<B> \<Gamma> v1 b1 b2)
- then show ?case
- by (metis (no_types) subst_bb.simps(5) subst_ceb.simps wfCE_sndI.hyps(2)
- wfCE_sndI wfCE_sndI.prems(2) Wellformed.wfCE_sndI)
-next
- case (wfCE_concatI \<Theta> \<B> \<Gamma> v1 v2)
- then show ?case using subst_bb.simps subst_ceb.simps wf_intros wfX_wfY wf_b_subst_lemmas wfCE_b_fresh
- proof -
- show ?thesis
- using wfCE_concatI.hyps(2) wfCE_concatI.hyps(4) wfCE_concatI.prems(1) wfCE_concatI.prems(2)
- Wellformed.wfCE_concatI by auto (* 46 ms *)
- qed
-next
- case (wfCE_lenI \<Theta> \<B> \<Gamma> v1)
- then show ?case using subst_bb.simps subst_ceb.simps wf_intros wfX_wfY wf_b_subst_lemmas wfCE_b_fresh by metis
-qed(auto simp add: wf_intros)
-
-lemma wf_b_subst2:
- fixes \<Gamma>::\<Gamma> and \<Gamma>'::\<Gamma> and v::v and e::e and c::c and \<tau>::\<tau> and ts::"(string*\<tau>) list" and \<Delta>::\<Delta> and b::b and ftq::fun_typ_q and ft::fun_typ and s::s and b'::b and ce::ce and td::type_def
- and cs::branch_s and css::branch_list
- shows "\<Theta> ; \<Phi> ; B' ; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f e : b' \<Longrightarrow> {|bv|} = B' \<Longrightarrow> \<Theta> ; B \<turnstile>\<^sub>w\<^sub>f b \<Longrightarrow> \<Theta> ; \<Phi> ; B ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b ; \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b \<turnstile>\<^sub>w\<^sub>f e[bv::=b]\<^sub>e\<^sub>b : b'[bv::=b]\<^sub>b\<^sub>b" and
- "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b \<Longrightarrow> True" and
- "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> ; tid ; dc ; t \<turnstile>\<^sub>w\<^sub>f cs : b \<Longrightarrow> True" and
- "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> ; tid ; dclist \<turnstile>\<^sub>w\<^sub>f css : b \<Longrightarrow> True" and
- "\<Theta> \<turnstile>\<^sub>w\<^sub>f (\<Phi>::\<Phi>) \<Longrightarrow> True " and
- "\<Theta> ; B' ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<Longrightarrow> {|bv|} = B' \<Longrightarrow> \<Theta> ; B \<turnstile>\<^sub>w\<^sub>f b \<Longrightarrow> \<Theta> ; B ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b" and
- "\<Theta> ; \<Phi> \<turnstile>\<^sub>w\<^sub>f ftq \<Longrightarrow> True" and
- "\<Theta> ; \<Phi> ; \<B> \<turnstile>\<^sub>w\<^sub>f ft \<Longrightarrow> True"
-proof(nominal_induct
- b' and b and b and b and \<Phi> and \<Delta> and ftq and ft
- avoiding: bv b B
-rule:wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.strong_induct)
- case (wfE_valI \<Theta>' \<Phi>' \<B>' \<Gamma>' \<Delta>' v' b')
- then show ?case unfolding subst_vb.simps subst_eb.simps using wf_b_subst1(1) Wellformed.wfE_valI by auto
-next
- case (wfE_plusI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 v2)
- then show ?case unfolding subst_eb.simps
- using wf_b_subst_lemmas wf_b_subst1(1) Wellformed.wfE_plusI
- proof -
- have "\<forall>b ba v g f ts. (( ts ; f ; g[bv::=ba]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f v[bv::=ba]\<^sub>v\<^sub>b : b[bv::=ba]\<^sub>b\<^sub>b) \<or> \<not> ts ; \<B> ; g \<turnstile>\<^sub>w\<^sub>f v : b ) \<or> \<not> ts ; f \<turnstile>\<^sub>w\<^sub>f ba"
- using wfE_plusI.prems(1) wf_b_subst1(1) by force (* 0.0 ms *)
- then show "\<Theta> ; \<Phi> ; B ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b ; \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b \<turnstile>\<^sub>w\<^sub>f [ plus v1[bv::=b]\<^sub>v\<^sub>b v2[bv::=b]\<^sub>v\<^sub>b ]\<^sup>e : B_int[bv::=b]\<^sub>b\<^sub>b"
-
- by (metis wfE_plusI.hyps(1) wfE_plusI.hyps(4) wfE_plusI.hyps(5) wfE_plusI.hyps(6) wfE_plusI.prems(1) wfE_plusI.prems(2) wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.wfE_plusI wf_b_subst_lemmas(86))
- qed
-next
- case (wfE_leqI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 v2)
- then show ?case unfolding subst_eb.simps
- using wf_b_subst_lemmas wf_b_subst1 Wellformed.wfE_leqI
- proof -
- have "\<And>ts f b ba g v. \<not> (ts ; f \<turnstile>\<^sub>w\<^sub>f b) \<or> \<not> (ts ; {|ba|} ; g \<turnstile>\<^sub>w\<^sub>f v : B_int) \<or> (ts ; f ; g[ba::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f v[ba::=b]\<^sub>v\<^sub>b : B_int)"
- by (metis wf_b_subst1(1) wf_b_subst_lemmas(86)) (* 46 ms *)
- then show "\<Theta> ; \<Phi> ; B ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b ; \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b \<turnstile>\<^sub>w\<^sub>f [ leq v1[bv::=b]\<^sub>v\<^sub>b v2[bv::=b]\<^sub>v\<^sub>b ]\<^sup>e : B_bool[bv::=b]\<^sub>b\<^sub>b"
- by (metis (no_types) wfE_leqI.hyps(1) wfE_leqI.hyps(4) wfE_leqI.hyps(5) wfE_leqI.hyps(6) wfE_leqI.prems(1) wfE_leqI.prems(2) wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.wfE_leqI wf_b_subst_lemmas(87)) (* 46 ms *)
- qed
-next
- case (wfE_eqI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 bb v2)
- show ?case unfolding subst_eb.simps subst_bb.simps proof
- show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using wfX_wfY wfE_eqI by metis
- show \<open> \<Theta> ; B ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b \<close> using wfX_wfY wfE_eqI by metis
- show \<open> \<Theta> ; B ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f v1[bv::=b]\<^sub>v\<^sub>b : bb \<close> using subst_bb.simps wfE_eqI
- by (metis (no_types, opaque_lifting) empty_iff insert_iff wf_b_subst1(1))
- show \<open> \<Theta> ; B ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f v2[bv::=b]\<^sub>v\<^sub>b : bb \<close> using wfX_wfY wfE_eqI
- by (metis insert_iff singleton_iff wf_b_subst1(1) wf_b_subst_lemmas(86) wf_b_subst_lemmas(87) wf_b_subst_lemmas(90))
- show \<open>bb \<in> {B_bool, B_int, B_unit}\<close> using wfE_eqI by auto
- qed
-next
- case (wfE_fstI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 b1 b2)
- then show ?case unfolding subst_eb.simps using wf_b_subst_lemmas(84) wf_b_subst1(1) Wellformed.wfE_fstI
- by (metis wf_b_subst_lemmas(89))
-next
- case (wfE_sndI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 b1 b2)
- then show ?case unfolding subst_eb.simps using wf_b_subst_lemmas(86) wf_b_subst1(1) Wellformed.wfE_sndI
- by (metis wf_b_subst_lemmas(89))
-next
- case (wfE_concatI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 v2)
-then show ?case unfolding subst_eb.simps using wf_b_subst_lemmas(86) wf_b_subst1(1) Wellformed.wfE_concatI
- by (metis wf_b_subst_lemmas(91))
-next
- case (wfE_splitI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 v2)
- then show ?case unfolding subst_eb.simps using wf_b_subst_lemmas(86) wf_b_subst1(1) Wellformed.wfE_splitI
- by (metis wf_b_subst_lemmas(89) wf_b_subst_lemmas(91))
-next
- case (wfE_lenI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1)
- then show ?case unfolding subst_eb.simps using wf_b_subst_lemmas(86) wf_b_subst1(1) Wellformed.wfE_lenI
- by (metis wf_b_subst_lemmas(91) wf_b_subst_lemmas(89))
-next
- case (wfE_appI \<Theta> \<Phi> \<B>' \<Gamma> \<Delta> f x b' c \<tau> s v)
- hence bf: "atom bv \<sharp> b'" using wfPhi_f_simple_wfT wfT_supp bv_not_in_dom_g wfPhi_f_simple_supp_b fresh_def by fast
- hence bseq: "b'[bv::=b]\<^sub>b\<^sub>b = b'" using subst_bb.simps wf_b_subst_lemmas by metis
- have "\<Theta> ; \<Phi> ; B ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b ; \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b \<turnstile>\<^sub>w\<^sub>f (AE_app f (v[bv::=b]\<^sub>v\<^sub>b)) : (b_of (\<tau>[bv::=b]\<^sub>\<tau>\<^sub>b))"
- proof
- show "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>" using wfE_appI by auto
- show "\<Theta> ; B ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b " using wfE_appI by simp
- have "atom bv \<sharp> \<tau>" using wfPhi_f_simple_wfT[OF wfE_appI(5) wfE_appI(1),THEN wfT_supp] bv_not_in_dom_g fresh_def by force
- hence " \<tau>[bv::=b]\<^sub>\<tau>\<^sub>b = \<tau>" using forget_subst subst_b_\<tau>_def by metis
- thus "Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x b' c \<tau>[bv::=b]\<^sub>\<tau>\<^sub>b s))) = lookup_fun \<Phi> f" using wfE_appI by simp
- show "\<Theta> ; B ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f v[bv::=b]\<^sub>v\<^sub>b : b'" using wfE_appI bseq wf_b_subst1 by metis
- qed
- then show ?case using subst_eb.simps b_of_subst_bb_commute by simp
-next
- case (wfE_appPI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> b' bv1 v1 \<tau>1 f x1 b1 c1 s1)
- then have *: "atom bv \<sharp> b1" using wfPhi_f_supp(1) wfE_appPI(7,11)
- by (metis fresh_def fresh_finsert singleton_iff subsetD fresh_def supp_at_base wfE_appPI.hyps(1))
- have "\<Theta> ; \<Phi> ; B ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b ; \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b \<turnstile>\<^sub>w\<^sub>f AE_appP f b'[bv::=b]\<^sub>b\<^sub>b (v1[bv::=b]\<^sub>v\<^sub>b) : (b_of \<tau>1)[bv1::=b'[bv::=b]\<^sub>b\<^sub>b]\<^sub>b"
- proof
- show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using wfE_appPI by auto
- show \<open> \<Theta> ; B ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b \<close> using wfE_appPI by auto
- show \<open> \<Theta> ; B \<turnstile>\<^sub>w\<^sub>f b'[bv::=b]\<^sub>b\<^sub>b \<close> using wfE_appPI wf_b_subst1 by auto
- have "atom bv1 \<sharp> \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b" using fresh_subst_if subst_b_\<Gamma>_def wfE_appPI by metis
- moreover have "atom bv1 \<sharp> b'[bv::=b]\<^sub>b\<^sub>b" using fresh_subst_if subst_b_b_def wfE_appPI by metis
- moreover have "atom bv1 \<sharp> v1[bv::=b]\<^sub>v\<^sub>b" using fresh_subst_if subst_b_v_def wfE_appPI by metis
- moreover have "atom bv1 \<sharp> \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b" using fresh_subst_if subst_b_\<Delta>_def wfE_appPI by metis
- moreover have "atom bv1 \<sharp> (b_of \<tau>1)[bv1::=b'[bv::=b]\<^sub>b\<^sub>b]\<^sub>b\<^sub>b" using fresh_subst_if subst_b_b_def wfE_appPI by metis
- ultimately show "atom bv1 \<sharp> (\<Phi>, \<Theta>, B, \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b, \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b, b'[bv::=b]\<^sub>b\<^sub>b, v1[bv::=b]\<^sub>v\<^sub>b, (b_of \<tau>1)[bv1::=b'[bv::=b]\<^sub>b\<^sub>b]\<^sub>b)"
- using wfE_appPI using fresh_def fresh_prodN subst_b_b_def by metis
- show \<open>Some (AF_fundef f (AF_fun_typ_some bv1 (AF_fun_typ x1 b1 c1 \<tau>1 s1))) = lookup_fun \<Phi> f\<close> using wfE_appPI by auto
-
- have \<open> \<Theta> ; B ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f v1[bv::=b]\<^sub>v\<^sub>b : b1[bv1::=b']\<^sub>b[bv::=b]\<^sub>b\<^sub>b \<close>
- using wfE_appPI subst_b_b_def * wf_b_subst1 by metis
- thus \<open> \<Theta> ; B ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f v1[bv::=b]\<^sub>v\<^sub>b : b1[bv1::=b'[bv::=b]\<^sub>b\<^sub>b]\<^sub>b \<close>
- using subst_bb_commute subst_b_b_def * by auto
- qed
- moreover have "atom bv \<sharp> b_of \<tau>1" proof -
- have "supp (b_of \<tau>1) \<subseteq> { atom bv1 }" using wfPhi_f_poly_supp_b_of_t
- using b_of.simps wfE_appPI wfPhi_f_supp(5) by simp
- thus ?thesis using wfE_appPI
- fresh_def fresh_finsert singleton_iff subsetD fresh_def supp_at_base wfE_appPI.hyps by metis
- qed
- ultimately show ?case using subst_eb.simps(3) subst_bb_commute subst_b_b_def * by simp
-next
- case (wfE_mvarI \<Theta> \<Phi> \<B>' \<Gamma> \<Delta> u \<tau>)
-
- have "\<Theta> ; \<Phi> ; B ; subst_gb \<Gamma> bv b ; subst_db \<Delta> bv b \<turnstile>\<^sub>w\<^sub>f (AE_mvar u)[bv::=b]\<^sub>e\<^sub>b : (b_of (\<tau>[bv::=b]\<^sub>\<tau>\<^sub>b))"
- proof(subst subst_eb.simps,rule Wellformed.wfE_mvarI)
- show "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> " using wfE_mvarI by simp
- show "\<Theta> ; B ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b" using wfE_mvarI by metis
- show "(u, \<tau>[bv::=b]\<^sub>\<tau>\<^sub>b) \<in> setD \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b"
- using wfE_mvarI subst_db.simps set_insert subst_d_b_member by simp
- qed
- thus ?case using b_of_subst_bb_commute by auto
-
-next
- case (wfS_seqI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> s1 s2 b)
- then show ?case using subst_bb.simps wf_intros wfX_wfY by metis
-next
- case (wfD_emptyI \<Theta> \<B>' \<Gamma>)
- then show ?case using subst_db.simps Wellformed.wfD_emptyI wf_b_subst1 by simp
-next
- case (wfD_cons \<Theta> \<B>' \<Gamma>' \<Delta> \<tau> u)
- show ?case proof(subst subst_db.simps, rule Wellformed.wfD_cons )
- show "\<Theta> ; B ; \<Gamma>'[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b " using wfD_cons by auto
- show "\<Theta> ; B ; \<Gamma>'[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f \<tau>[bv::=b]\<^sub>\<tau>\<^sub>b " using wfD_cons wf_b_subst1 by auto
- show "u \<notin> fst ` setD \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b" using wfD_cons subst_b_lookup_d by metis
- qed
-next
- case (wfS_assertI \<Theta> \<Phi> \<B> x c \<Gamma> \<Delta> s b)
- show ?case by auto
-qed(auto)
-
-lemmas wf_b_subst = wf_b_subst1 wf_b_subst2
-
-lemma wfT_subst_wfT:
- fixes \<tau>::\<tau> and b'::b and bv::bv
- assumes "\<Theta> ; {|bv|} ; (x,b,c) #\<^sub>\<Gamma>GNil \<turnstile>\<^sub>w\<^sub>f \<tau>" and "\<Theta> ; B \<turnstile>\<^sub>w\<^sub>f b'"
- shows "\<Theta> ; B ; (x,b[bv::=b']\<^sub>b\<^sub>b,c[bv::=b']\<^sub>c\<^sub>b) #\<^sub>\<Gamma>GNil \<turnstile>\<^sub>w\<^sub>f (\<tau>[bv::=b']\<^sub>\<tau>\<^sub>b)"
-proof -
- have "\<Theta> ; B ; ((x,b,c) #\<^sub>\<Gamma>GNil)[bv::=b']\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f (\<tau>[bv::=b']\<^sub>\<tau>\<^sub>b)"
- using wf_b_subst assms by metis
- thus ?thesis using subst_gb.simps wf_b_subst_lemmas wfCE_b_fresh by metis
-qed
-
-lemma wf_trans:
- fixes \<Gamma>::\<Gamma> and \<Gamma>'::\<Gamma> and v::v and e::e and c::c and \<tau>::\<tau> and ts::"(string*\<tau>) list" and \<Delta>::\<Delta> and b::b and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def and s::s
- and cs::branch_s and css::branch_list and \<Theta>::\<Theta>
- shows "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b' \<Longrightarrow> \<Gamma> = (x, b, c2) #\<^sub>\<Gamma> G \<Longrightarrow> \<Theta>; \<B>; (x, b, c1) #\<^sub>\<Gamma> G \<turnstile>\<^sub>w\<^sub>f c2 \<Longrightarrow> \<Theta>; \<B>; (x, b, c1) #\<^sub>\<Gamma> G \<turnstile>\<^sub>w\<^sub>f v : b'" and
- "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f c \<Longrightarrow> \<Gamma> = (x, b, c2) #\<^sub>\<Gamma> G \<Longrightarrow> \<Theta>; \<B>; (x, b, c1) #\<^sub>\<Gamma> G \<turnstile>\<^sub>w\<^sub>f c2 \<Longrightarrow> \<Theta>; \<B>; (x, b, c1) #\<^sub>\<Gamma> G \<turnstile>\<^sub>w\<^sub>f c" and
- "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<Longrightarrow> True" and
- "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau> \<Longrightarrow> True" and
- "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ts \<Longrightarrow> True" and
- "\<turnstile>\<^sub>w\<^sub>f \<Theta> \<Longrightarrow>True" and
- "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b \<Longrightarrow> True " and
- "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ce : b' \<Longrightarrow> \<Gamma> = (x, b, c2) #\<^sub>\<Gamma> G \<Longrightarrow> \<Theta>; \<B>; (x, b, c1) #\<^sub>\<Gamma> G \<turnstile>\<^sub>w\<^sub>f c2 \<Longrightarrow> \<Theta>; \<B>; (x, b, c1) #\<^sub>\<Gamma> G \<turnstile>\<^sub>w\<^sub>f ce : b' " and
- "\<Theta> \<turnstile>\<^sub>w\<^sub>f td \<Longrightarrow> True"
-proof(nominal_induct
- b' and c and \<Gamma> and \<tau> and ts and \<Theta> and b and b' and td
- avoiding: c1
- arbitrary: \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1
- rule:wfV_wfC_wfG_wfT_wfTs_wfTh_wfB_wfCE_wfTD.strong_induct)
- case (wfV_varI \<Theta> \<B> \<Gamma> b' c' x')
- have wbg: "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (x, b, c1) #\<^sub>\<Gamma> G " using wfC_wf wfV_varI by simp
- show ?case proof(cases "x=x'")
- case True
- have "Some (b', c1) = lookup ((x, b, c1) #\<^sub>\<Gamma> G) x'" using lookup.simps wfV_varI using True by auto
- then show ?thesis using Wellformed.wfV_varI wbg by simp
- next
- case False
- then have "Some (b', c') = lookup ((x, b, c1) #\<^sub>\<Gamma> G) x'" using lookup.simps wfV_varI
- by simp
- then show ?thesis using Wellformed.wfV_varI wbg by simp
- qed
-next
- case (wfV_conspI s bv dclist \<Theta> dc x1 b' c \<B> b1 \<Gamma> v)
- show ?case proof
- show \<open>AF_typedef_poly s bv dclist \<in> set \<Theta>\<close> using wfV_conspI by auto
- show \<open>(dc, \<lbrace> x1 : b' | c \<rbrace>) \<in> set dclist\<close> using wfV_conspI by auto
- show \<open>\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b1 \<close> using wfV_conspI by auto
- show \<open>atom bv \<sharp> (\<Theta>, \<B>, (x, b, c1) #\<^sub>\<Gamma> G, b1, v)\<close> unfolding fresh_prodN fresh_GCons using wfV_conspI fresh_prodN fresh_GCons by simp
- show \<open>\<Theta>; \<B>; (x, b, c1) #\<^sub>\<Gamma> G \<turnstile>\<^sub>w\<^sub>f v : b'[bv::=b1]\<^sub>b\<^sub>b\<close> using wfV_conspI by auto
- qed
-qed( (auto | metis wfC_wf wf_intros) +)
-
-
+(*<*)
+theory WellformedL
+ imports Wellformed "SyntaxL"
+begin
+(*>*)
+
+chapter \<open>Wellformedness Lemmas\<close>
+
+section \<open>Prelude\<close>
+
+lemma b_of_subst_bb_commute:
+ "(b_of (\<tau>[bv::=b]\<^sub>\<tau>\<^sub>b)) = (b_of \<tau>)[bv::=b]\<^sub>b\<^sub>b"
+proof -
+ obtain z' and b' and c' where "\<tau> = \<lbrace> z' : b' | c' \<rbrace> " using obtain_fresh_z by metis
+ moreover hence "(b_of (\<tau>[bv::=b]\<^sub>\<tau>\<^sub>b)) = b_of \<lbrace> z' : b'[bv::=b]\<^sub>b\<^sub>b | c' \<rbrace>" using subst_tb.simps by simp
+ ultimately show ?thesis using subst_tv.simps subst_tb.simps by simp
+qed
+
+lemmas wf_intros = wfV_wfC_wfG_wfT_wfTs_wfTh_wfB_wfCE_wfTD.intros wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.intros
+
+lemmas freshers = fresh_prodN b.fresh c.fresh v.fresh ce.fresh fresh_GCons fresh_GNil fresh_at_base
+
+section \<open>Strong Elimination\<close>
+
+text \<open>Inversion/elimination for well-formed polymorphic constructors \<close>
+lemma wf_strong_elim:
+ fixes \<Gamma>::\<Gamma> and \<Gamma>'::\<Gamma> and v::v and e::e and c::c and \<tau>::\<tau> and ts::"(string*\<tau>) list"
+ and \<Delta>::\<Delta> and b::b and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def and s::s and tm::"'a::fs"
+ and cs::branch_s and css::branch_list and \<Theta>::\<Theta>
+ shows "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f (V_consp tyid dc b v) : b'' \<Longrightarrow> (\<exists> bv dclist x b' c. b'' = B_app tyid b \<and>
+ AF_typedef_poly tyid bv dclist \<in> set \<Theta> \<and>
+ (dc, \<lbrace> x : b' | c \<rbrace>) \<in> set dclist \<and>
+ \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b \<and> atom bv \<sharp> (\<Theta>, \<B>, \<Gamma>, b, v) \<and> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b'[bv::=b]\<^sub>b\<^sub>b \<and> atom bv \<sharp> tm)" and
+ "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f c \<Longrightarrow> True" and
+ "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<Longrightarrow> True" and
+ "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau> \<Longrightarrow> True" and
+ "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ts \<Longrightarrow> True" and
+ "\<turnstile>\<^sub>w\<^sub>f \<Theta> \<Longrightarrow>True" and
+ "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b \<Longrightarrow> True " and
+ "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ce : b' \<Longrightarrow> True" and
+ "\<Theta> \<turnstile>\<^sub>w\<^sub>f td \<Longrightarrow> True"
+proof(nominal_induct
+ "V_consp tyid dc b v" b'' and c and \<Gamma> and \<tau> and ts and \<Theta> and b and b' and td
+ avoiding: tm
+
+rule:wfV_wfC_wfG_wfT_wfTs_wfTh_wfB_wfCE_wfTD.strong_induct)
+ case (wfV_conspI bv dclist \<Theta> x b' c \<B> \<Gamma>)
+ then show ?case by force
+qed(auto+)
+
+section \<open>Context Extension\<close>
+
+definition wfExt :: "\<Theta> \<Rightarrow> \<B> \<Rightarrow> \<Gamma> \<Rightarrow> \<Gamma> \<Rightarrow> bool" (" _ ; _ \<turnstile>\<^sub>w\<^sub>f _ < _ " [50,50,50] 50) where
+ "wfExt T B G1 G2 = (wfG T B G2 \<and> wfG T B G1 \<and> toSet G1 \<subseteq> toSet G2)"
+
+section \<open>Context\<close>
+
+lemma wfG_cons[ms_wb]:
+ fixes \<Gamma>::\<Gamma>
+ assumes "P; \<B> \<turnstile>\<^sub>w\<^sub>f (z,b,c) #\<^sub>\<Gamma>\<Gamma>"
+ shows "P; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<and> atom z \<sharp> \<Gamma> \<and> wfB P \<B> b"
+ using wfG_elims(2)[OF assms] by metis
+
+lemma wfG_cons2[ms_wb]:
+ fixes \<Gamma>::\<Gamma>
+ assumes "P; \<B> \<turnstile>\<^sub>w\<^sub>f zbc #\<^sub>\<Gamma>\<Gamma>"
+ shows "P; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>"
+proof -
+ obtain z and b and c where zbc: "zbc=(z,b,c)" using prod_cases3 by blast
+ hence "P; \<B> \<turnstile>\<^sub>w\<^sub>f (z,b,c) #\<^sub>\<Gamma>\<Gamma>" using assms by auto
+ thus ?thesis using zbc wfG_cons assms by simp
+qed
+
+lemma wf_g_unique:
+ fixes \<Gamma>::\<Gamma>
+ assumes "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>" and "(x,b,c) \<in> toSet \<Gamma>" and "(x,b',c') \<in> toSet \<Gamma>"
+ shows "b=b' \<and> c=c'"
+using assms proof(induct \<Gamma> rule: \<Gamma>.induct)
+ case GNil
+ then show ?case by simp
+next
+ case (GCons a \<Gamma>)
+ consider "(x,b,c)=a \<and> (x,b',c')=a" | "(x,b,c)=a \<and> (x,b',c')\<noteq>a" | "(x,b,c)\<noteq>a \<and> (x,b',c')=a" | "(x,b,c)\<noteq> a \<and> (x,b',c')\<noteq>a" by blast
+ then show ?case proof(cases)
+ case 1
+ then show ?thesis by auto
+ next
+ case 2
+ hence "atom x \<sharp> \<Gamma>" using wfG_elims(2) GCons by blast
+ moreover have "(x,b',c') \<in> toSet \<Gamma>" using GCons 2 by force
+ ultimately show ?thesis using forget_subst_gv fresh_GCons fresh_GNil fresh_gamma_elem \<Gamma>.distinct subst_gv.simps 2 GCons by metis
+ next
+ case 3
+ hence "atom x \<sharp> \<Gamma>" using wfG_elims(2) GCons by blast
+ moreover have "(x,b,c) \<in> toSet \<Gamma>" using GCons 3 by force
+ ultimately show ?thesis
+ using forget_subst_gv fresh_GCons fresh_GNil fresh_gamma_elem \<Gamma>.distinct subst_gv.simps 3 GCons by metis
+ next
+ case 4
+ then obtain x'' and b'' and c''::c where xbc: "a=(x'',b'',c'')"
+ using prod_cases3 by blast
+ hence "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f ((x'',b'',c'') #\<^sub>\<Gamma>\<Gamma>)" using GCons wfG_elims by blast
+ hence "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<and> (x, b, c) \<in> toSet \<Gamma> \<and> (x, b', c') \<in> toSet \<Gamma>" using GCons wfG_elims 4 xbc
+ prod_cases3 set_GConsD using forget_subst_gv fresh_GCons fresh_GNil fresh_gamma_elem \<Gamma>.distinct subst_gv.simps 4 GCons by meson
+ thus ?thesis using GCons by auto
+ qed
+qed
+
+lemma lookup_if1:
+ fixes \<Gamma>::\<Gamma>
+ assumes "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>" and "Some (b,c) = lookup \<Gamma> x"
+ shows "(x,b,c) \<in> toSet \<Gamma> \<and> (\<forall>b' c'. (x,b',c') \<in> toSet \<Gamma> \<longrightarrow> b'=b \<and> c'=c)"
+using assms proof(induct \<Gamma> rule: \<Gamma>.induct)
+ case GNil
+ then show ?case by auto
+next
+ case (GCons xbc \<Gamma>)
+ then obtain x' and b' and c'::c where xbc: "xbc=(x',b',c')"
+ using prod_cases3 by blast
+ then show ?case using wf_g_unique GCons lookup_in_g xbc
+ lookup.simps set_GConsD wfG.cases
+ insertE insert_is_Un toSet.simps wfG_elims by metis
+qed
+
+lemma lookup_if2:
+ assumes "wfG P \<B> \<Gamma>" and "(x,b,c) \<in> toSet \<Gamma> \<and> (\<forall>b' c'. (x,b',c') \<in> toSet \<Gamma> \<longrightarrow> b'=b \<and> c'=c)"
+ shows "Some (b,c) = lookup \<Gamma> x"
+using assms proof(induct \<Gamma> rule: \<Gamma>.induct)
+ case GNil
+ then show ?case by auto
+next
+ case (GCons xbc \<Gamma>)
+ then obtain x' and b' and c'::c where xbc: "xbc=(x',b',c')"
+ using prod_cases3 by blast
+ then show ?case proof(cases "x=x'")
+ case True
+ then show ?thesis using lookup.simps GCons xbc by simp
+ next
+ case False
+ then show ?thesis using lookup.simps GCons xbc toSet.simps Un_iff set_GConsD wfG_cons2
+ by (metis (full_types) Un_iff set_GConsD toSet.simps(2) wfG_cons2)
+ qed
+qed
+
+lemma lookup_iff:
+ fixes \<Theta>::\<Theta> and \<Gamma>::\<Gamma>
+ assumes "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>"
+ shows "Some (b,c) = lookup \<Gamma> x \<longleftrightarrow> (x,b,c) \<in> toSet \<Gamma> \<and> (\<forall>b' c'. (x,b',c') \<in> toSet \<Gamma> \<longrightarrow> b'=b \<and> c'=c)"
+ using assms lookup_if1 lookup_if2 by meson
+
+lemma wfG_lookup_wf:
+ fixes \<Theta>::\<Theta> and \<Gamma>::\<Gamma> and b::b and \<B>::\<B>
+ assumes "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>" and "Some (b,c) = lookup \<Gamma> x"
+ shows "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b"
+using assms proof(induct \<Gamma> rule: \<Gamma>_induct)
+ case GNil
+ then show ?case by auto
+next
+ case (GCons x' b' c' \<Gamma>')
+ then show ?case proof(cases "x=x'")
+ case True
+ then show ?thesis using lookup.simps wfG_elims(2) GCons by fastforce
+ next
+ case False
+ then show ?thesis using lookup.simps wfG_elims(2) GCons by fastforce
+ qed
+qed
+
+lemma wfG_unique:
+ fixes \<Gamma>::\<Gamma>
+ assumes "wfG B \<Theta> ((x, b, c) #\<^sub>\<Gamma> \<Gamma>)" and "(x1,b1,c1) \<in> toSet ((x, b, c) #\<^sub>\<Gamma> \<Gamma>)" and "x1=x"
+ shows "b1 = b \<and> c1 = c"
+proof -
+ have "(x, b, c) \<in> toSet ((x, b, c) #\<^sub>\<Gamma> \<Gamma>)" by simp
+ thus ?thesis using wf_g_unique assms by blast
+qed
+
+lemma wfG_unique_full:
+ fixes \<Gamma>::\<Gamma>
+ assumes "wfG \<Theta> B (\<Gamma>'@(x, b, c) #\<^sub>\<Gamma> \<Gamma>)" and "(x1,b1,c1) \<in> toSet (\<Gamma>'@(x, b, c) #\<^sub>\<Gamma> \<Gamma>)" and "x1=x"
+ shows "b1 = b \<and> c1 = c"
+proof -
+ have "(x, b, c) \<in> toSet (\<Gamma>'@(x, b, c) #\<^sub>\<Gamma> \<Gamma>)" by simp
+ thus ?thesis using wf_g_unique assms by blast
+qed
+
+section \<open>Converting between wb forms\<close>
+
+text \<open> We cannot prove wfB properties here for expressions and statements as need some more facts about @{term \<Phi>}
+ context which we can prove without this lemma. Trying to cram everything into a single large
+ mutually recursive lemma is not a good idea \<close>
+
+lemma wfX_wfY1:
+ fixes \<Gamma>::\<Gamma> and \<Gamma>'::\<Gamma> and v::v and e::e and c::c and \<tau>::\<tau> and ts::"(string*\<tau>) list" and \<Delta>::\<Delta> and s::s and b::b and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def and cs::branch_s
+ and css::branch_list
+ shows wfV_wf: "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<and> \<turnstile>\<^sub>w\<^sub>f \<Theta> " and
+ wfC_wf: "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f c \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<and> \<turnstile>\<^sub>w\<^sub>f \<Theta> " and
+ wfG_wf :"\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<Longrightarrow> \<turnstile>\<^sub>w\<^sub>f \<Theta>" and
+ wfT_wf: "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau> \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<and> \<turnstile>\<^sub>w\<^sub>f \<Theta> \<and> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b_of \<tau>" and
+ wfTs_wf:"\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ts \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<and> \<turnstile>\<^sub>w\<^sub>f \<Theta>" and
+ "\<turnstile>\<^sub>w\<^sub>f \<Theta> \<Longrightarrow> True" and
+ wfB_wf: "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b \<Longrightarrow> \<turnstile>\<^sub>w\<^sub>f \<Theta>" and
+ wfCE_wf: "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ce : b \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<and> \<turnstile>\<^sub>w\<^sub>f \<Theta> " and
+ wfTD_wf: "\<Theta> \<turnstile>\<^sub>w\<^sub>f td \<Longrightarrow> \<turnstile>\<^sub>w\<^sub>f \<Theta>"
+proof(induct rule:wfV_wfC_wfG_wfT_wfTs_wfTh_wfB_wfCE_wfTD.inducts)
+
+ case (wfV_varI \<Theta> \<B> \<Gamma> b c x)
+ hence "(x,b,c) \<in> toSet \<Gamma>" using lookup_iff lookup_in_g by presburger
+ hence "b \<in> fst`snd`toSet \<Gamma>" by force
+ hence "wfB \<Theta> \<B> b" using wfV_varI using wfG_lookup_wf by auto
+ then show ?case using wfV_varI wfV_elims wf_intros by metis
+next
+ case (wfV_litI \<Theta> \<B> \<Gamma> l)
+ moreover have "wfTh \<Theta>" using wfV_litI by metis
+ ultimately show ?case using wf_intros base_for_lit.simps l.exhaust by metis
+next
+ case (wfV_pairI \<Theta> \<B> \<Gamma> v1 b1 v2 b2)
+ then show ?case using wfB_pairI by simp
+next
+ case (wfV_consI s dclist \<Theta> dc x b c \<B> \<Gamma> v)
+ then show ?case using wf_intros by metis
+next
+ case (wfTI z \<Gamma> \<Theta> \<B> b c)
+ then show ?case using wf_intros b_of.simps wfG_cons2 by metis
+qed(auto)
+
+lemma wfX_wfY2:
+ fixes \<Gamma>::\<Gamma> and \<Gamma>'::\<Gamma> and v::v and e::e and c::c and \<tau>::\<tau> and ts::"(string*\<tau>) list" and \<Delta>::\<Delta> and s::s and b::b and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def and cs::branch_s
+ and css::branch_list
+ shows
+ wfE_wf: "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f e : b \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<and> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<and> \<turnstile>\<^sub>w\<^sub>f \<Theta> \<and> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> " and
+ wfS_wf: "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<and> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<and> \<turnstile>\<^sub>w\<^sub>f \<Theta> \<and> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> " and
+ "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> ; tid ; dc ; t \<turnstile>\<^sub>w\<^sub>f cs : b \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<and> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<and> \<turnstile>\<^sub>w\<^sub>f \<Theta> \<and> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> " and
+ "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> ; tid ; dclist \<turnstile>\<^sub>w\<^sub>f css : b \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<and> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<and> \<turnstile>\<^sub>w\<^sub>f \<Theta> \<and> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> " and
+ wfPhi_wf: "\<Theta> \<turnstile>\<^sub>w\<^sub>f (\<Phi>::\<Phi>) \<Longrightarrow> \<turnstile>\<^sub>w\<^sub>f \<Theta>" and
+ wfD_wf: "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<and> \<turnstile>\<^sub>w\<^sub>f \<Theta> " and
+ wfFTQ_wf: "\<Theta> ; \<Phi> \<turnstile>\<^sub>w\<^sub>f ftq \<Longrightarrow> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<and> \<turnstile>\<^sub>w\<^sub>f \<Theta>" and
+ wfFT_wf: "\<Theta> ; \<Phi> ; \<B> \<turnstile>\<^sub>w\<^sub>f ft \<Longrightarrow> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<and> \<turnstile>\<^sub>w\<^sub>f \<Theta>"
+proof(induct rule:wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.inducts)
+ case (wfS_varI \<Theta> \<B> \<Gamma> \<tau> v u \<Delta> \<Phi> s b)
+ then show ?case using wfD_elims by auto
+next
+ case (wfS_assignI u \<tau> \<Delta> \<Theta> \<B> \<Gamma> \<Phi> v)
+ then show ?case using wf_intros by metis
+next
+ case (wfD_emptyI \<Theta> \<B> \<Gamma>)
+ then show ?case using wfX_wfY1 by auto
+next
+ case (wfS_assertI \<Theta> \<Phi> \<B> x c \<Gamma> \<Delta> s b)
+ then have "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>" using wfX_wfY1 by auto
+ moreover have "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>" using wfS_assertI by auto
+ moreover have "\<turnstile>\<^sub>w\<^sub>f \<Theta> \<and> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> " using wfS_assertI by auto
+ ultimately show ?case by auto
+qed(auto)
+
+lemmas wfX_wfY=wfX_wfY1 wfX_wfY2
+
+lemma setD_ConsD:
+ "ut \<in> setD (ut' #\<^sub>\<Delta> D) = (ut = ut' \<or> ut \<in> setD D)"
+proof(induct D rule: \<Delta>_induct)
+ case DNil
+ then show ?case by auto
+next
+ case (DCons u' t' x2)
+ then show ?case using setD.simps by auto
+qed
+
+lemma wfD_wfT:
+ fixes \<Delta>::\<Delta> and \<tau>::\<tau>
+ assumes "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>"
+ shows "\<forall>(u,\<tau>) \<in> setD \<Delta>. \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau>"
+using assms proof(induct \<Delta> rule: \<Delta>_induct)
+ case DNil
+ then show ?case by auto
+next
+ case (DCons u' t' x2)
+ then show ?case using wfD_elims DCons setD_ConsD
+ by (metis case_prodI2 set_ConsD)
+qed
+
+lemma subst_b_lookup_d:
+ assumes "u \<notin> fst ` setD \<Delta>"
+ shows "u \<notin> fst ` setD \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b"
+using assms proof(induct \<Delta> rule: \<Delta>_induct)
+ case DNil
+ then show ?case by auto
+next
+ case (DCons u' t' x2)
+ hence "u\<noteq>u'" using DCons by simp
+ show ?case using DCons subst_db.simps by simp
+qed
+
+lemma wfG_cons_splitI:
+ fixes \<Phi>::\<Phi> and \<Gamma>::\<Gamma>
+ assumes "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>" and "atom x \<sharp> \<Gamma>" and "wfB \<Theta> \<B> b" and
+ "c \<in> { TRUE, FALSE } \<longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> " and
+ "c \<notin> { TRUE, FALSE } \<longrightarrow> \<Theta> ;\<B> ; (x,b,C_true) #\<^sub>\<Gamma>\<Gamma> \<turnstile>\<^sub>w\<^sub>f c"
+ shows "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f ((x,b,c) #\<^sub>\<Gamma>\<Gamma>)"
+ using wfG_cons1I wfG_cons2I assms by metis
+
+lemma wfG_consI:
+ fixes \<Phi>::\<Phi> and \<Gamma>::\<Gamma> and c::c
+ assumes "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>" and "atom x \<sharp> \<Gamma>" and "wfB \<Theta> \<B> b" and
+ "\<Theta> ; \<B> ; (x,b,C_true) #\<^sub>\<Gamma>\<Gamma> \<turnstile>\<^sub>w\<^sub>f c"
+ shows "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f ((x,b,c) #\<^sub>\<Gamma>\<Gamma>)"
+ using wfG_cons1I wfG_cons2I wfG_cons_splitI wfC_trueI assms by metis
+
+lemma wfG_elim2:
+ fixes c::c
+ assumes "wfG P \<B> ((x,b,c) #\<^sub>\<Gamma>\<Gamma>)"
+ shows "P; \<B> ; (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f c \<and> wfB P \<B> b"
+proof(cases "c \<in> {TRUE,FALSE}")
+ case True
+ have "P; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<and> atom x \<sharp> \<Gamma> \<and> wfB P \<B> b" using wfG_elims(2)[OF assms] by auto
+ hence "P; \<B> \<turnstile>\<^sub>w\<^sub>f ((x,b,TRUE) #\<^sub>\<Gamma>\<Gamma>) \<and> wfB P \<B> b" using wfG_cons2I by auto
+ thus ?thesis using wfC_trueI wfC_falseI True by auto
+next
+ case False
+ then show ?thesis using wfG_elims(2)[OF assms] by auto
+qed
+
+lemma wfG_cons_wfC:
+ fixes \<Gamma>::\<Gamma> and c::c
+ assumes "\<Theta> ; B \<turnstile>\<^sub>w\<^sub>f (x, b, c) #\<^sub>\<Gamma> \<Gamma>"
+ shows "\<Theta> ; B ; ((x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>) \<turnstile>\<^sub>w\<^sub>f c"
+ using assms wfG_elim2 by auto
+
+lemma wfG_wfB:
+ assumes "wfG P \<B> \<Gamma>" and "b \<in> fst`snd`toSet \<Gamma>"
+ shows "wfB P \<B> b"
+using assms proof(induct \<Gamma> rule:\<Gamma>_induct)
+case GNil
+ then show ?case by auto
+next
+ case (GCons x' b' c' \<Gamma>')
+ show ?case proof(cases "b=b'")
+ case True
+ then show ?thesis using wfG_elim2 GCons by auto
+ next
+ case False
+ hence "b \<in> fst`snd`toSet \<Gamma>'" using GCons by auto
+ moreover have "wfG P \<B> \<Gamma>'" using wfG_cons GCons by auto
+ ultimately show ?thesis using GCons by auto
+ qed
+qed
+
+lemma wfG_cons_TRUE:
+ fixes \<Gamma>::\<Gamma> and b::b
+ assumes "P; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>" and "atom z \<sharp> \<Gamma>" and "P; \<B> \<turnstile>\<^sub>w\<^sub>f b"
+ shows "P ; \<B> \<turnstile>\<^sub>w\<^sub>f (z, b, TRUE) #\<^sub>\<Gamma> \<Gamma>"
+ using wfG_cons2I wfG_wfB assms by simp
+
+lemma wfG_cons_TRUE2:
+ assumes "P; \<B> \<turnstile>\<^sub>w\<^sub>f (z,b,c) #\<^sub>\<Gamma>\<Gamma>" and "atom z \<sharp> \<Gamma>"
+ shows "P; \<B> \<turnstile>\<^sub>w\<^sub>f (z, b, TRUE) #\<^sub>\<Gamma> \<Gamma>"
+ using wfG_cons wfG_cons2I assms by simp
+
+lemma wfG_suffix:
+ fixes \<Gamma>::\<Gamma>
+ assumes "wfG P \<B> (\<Gamma>'@\<Gamma>)"
+ shows "wfG P \<B> \<Gamma>"
+using assms proof(induct \<Gamma>' rule: \<Gamma>_induct)
+ case GNil
+ then show ?case by auto
+next
+ case (GCons x b c \<Gamma>')
+ hence " P; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>' @ \<Gamma>" using wfG_elims by auto
+ then show ?case using GCons wfG_elims by auto
+qed
+
+lemma wfV_wfCE:
+ fixes v::v
+ assumes "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b"
+ shows " \<Theta> ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_val v : b"
+proof -
+ have "\<Theta> \<turnstile>\<^sub>w\<^sub>f ([]::\<Phi>) " using wfPhi_emptyI wfV_wf wfG_wf assms by metis
+ moreover have "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f []\<^sub>\<Delta>" using wfD_emptyI wfV_wf wfG_wf assms by metis
+ ultimately show ?thesis using wfCE_valI assms by auto
+qed
+
+section \<open>Support\<close>
+
+lemma wf_supp1:
+ fixes \<Gamma>::\<Gamma> and \<Gamma>'::\<Gamma> and v::v and e::e and c::c and \<tau>::\<tau> and ts::"(string*\<tau>) list" and \<Delta>::\<Delta> and s::s and b::b and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def and cs::branch_s and css ::branch_list
+
+ shows wfV_supp: "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b \<Longrightarrow> supp v \<subseteq> atom_dom \<Gamma> \<union> supp \<B>" and
+ wfC_supp: "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f c \<Longrightarrow> supp c \<subseteq> atom_dom \<Gamma> \<union> supp \<B>" and
+ wfG_supp: "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<Longrightarrow> atom_dom \<Gamma> \<subseteq> supp \<Gamma>" and
+ wfT_supp: "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau> \<Longrightarrow> supp \<tau> \<subseteq> atom_dom \<Gamma> \<union> supp \<B> " and
+ wfTs_supp: "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ts \<Longrightarrow> supp ts \<subseteq> atom_dom \<Gamma> \<union> supp \<B>" and
+ wfTh_supp: "\<turnstile>\<^sub>w\<^sub>f \<Theta> \<Longrightarrow> supp \<Theta> = {}" and
+ wfB_supp: "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b \<Longrightarrow> supp b \<subseteq> supp \<B>" and
+ wfCE_supp: "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ce : b \<Longrightarrow> supp ce \<subseteq> atom_dom \<Gamma> \<union> supp \<B>" and
+ wfTD_supp: "\<Theta> \<turnstile>\<^sub>w\<^sub>f td \<Longrightarrow> supp td \<subseteq> {}"
+proof(induct rule:wfV_wfC_wfG_wfT_wfTs_wfTh_wfB_wfCE_wfTD.inducts)
+ case (wfB_consI \<Theta> s dclist \<B>)
+ then show ?case by(auto simp add: b.supp pure_supp)
+next
+ case (wfB_appI \<Theta> \<B> b s bv dclist)
+ then show ?case by(auto simp add: b.supp pure_supp)
+next
+ case (wfV_varI \<Theta> \<B> \<Gamma> b c x)
+ then show ?case using v.supp wfV_elims
+ empty_subsetI insert_subset supp_at_base
+ fresh_dom_free2 lookup_if1
+ by (metis sup.coboundedI1)
+next
+ case (wfV_litI \<Theta> \<B> \<Gamma> l)
+ then show ?case using supp_l_empty v.supp by simp
+next
+ case (wfV_pairI \<Theta> \<B> \<Gamma> v1 b1 v2 b2)
+ then show ?case using v.supp wfV_elims by (metis Un_subset_iff)
+next
+ case (wfV_consI s dclist \<Theta> dc x b c \<B> \<Gamma> v)
+ then show ?case using v.supp wfV_elims
+ Un_commute b.supp sup_bot.right_neutral supp_b_empty pure_supp by metis
+next
+ case (wfV_conspI typid bv dclist \<Theta> dc x b' c \<B> \<Gamma> v b)
+ then show ?case unfolding v.supp
+ using wfV_elims
+ Un_commute b.supp sup_bot.right_neutral supp_b_empty pure_supp
+ by (simp add: Un_commute pure_supp sup.coboundedI1)
+next
+ case (wfC_eqI \<Theta> \<B> \<Gamma> e1 b e2)
+ hence "supp e1 \<subseteq> atom_dom \<Gamma> \<union> supp \<B>" using c.supp wfC_elims
+ image_empty list.set(1) sup_bot.right_neutral by (metis IntI UnE empty_iff subsetCE subsetI)
+ moreover have "supp e2 \<subseteq> atom_dom \<Gamma> \<union> supp \<B>" using c.supp wfC_elims
+ image_empty list.set(1) sup_bot.right_neutral IntI UnE empty_iff subsetCE subsetI
+ by (metis wfC_eqI.hyps(4))
+ ultimately show ?case using c.supp by auto
+next
+ case (wfG_cons1I c \<Theta> \<B> \<Gamma> x b)
+ then show ?case using atom_dom.simps dom_supp_g supp_GCons by metis
+next
+ case (wfG_cons2I c \<Theta> \<B> \<Gamma> x b)
+ then show ?case using atom_dom.simps dom_supp_g supp_GCons by metis
+next
+ case wfTh_emptyI
+ then show ?case by (simp add: supp_Nil)
+next
+ case (wfTh_consI \<Theta> lst)
+ then show ?case using supp_Cons by fast
+next
+ case (wfTD_simpleI \<Theta> lst s)
+ then have "supp (AF_typedef s lst ) = supp lst \<union> supp s" using type_def.supp by auto
+ then show ?case using wfTD_simpleI pure_supp
+ by (simp add: pure_supp supp_Cons supp_at_base)
+next
+ case (wfTD_poly \<Theta> bv lst s)
+ then have "supp (AF_typedef_poly s bv lst ) = supp lst - { atom bv } \<union> supp s" using type_def.supp by auto
+ then show ?case using wfTD_poly pure_supp
+ by (simp add: pure_supp supp_Cons supp_at_base)
+next
+ case (wfTs_nil \<Theta> \<B> \<Gamma>)
+ then show ?case using supp_Nil by auto
+next
+ case (wfTs_cons \<Theta> \<B> \<Gamma> \<tau> dc ts)
+ then show ?case using supp_Cons supp_Pair pure_supp[of dc] by blast
+next
+ case (wfCE_valI \<Theta> \<B> \<Gamma> v b)
+ thus ?case using ce.supp wfCE_elims by simp
+next
+ case (wfCE_plusI \<Theta> \<B> \<Gamma> v1 v2)
+ hence "supp (CE_op Plus v1 v2) \<subseteq> atom_dom \<Gamma> \<union> supp \<B>" using ce.supp pure_supp
+ by (simp add: wfCE_plusI opp.supp)
+ then show ?case using ce.supp wfCE_elims UnCI subsetCE subsetI x_not_in_b_set by auto
+next
+ case (wfCE_leqI \<Theta> \<B> \<Gamma> v1 v2)
+ hence "supp (CE_op LEq v1 v2) \<subseteq> atom_dom \<Gamma> \<union> supp \<B>" using ce.supp pure_supp
+ by (simp add: wfCE_plusI opp.supp)
+ then show ?case using ce.supp wfE_elims UnCI subsetCE subsetI x_not_in_b_set by auto
+next
+ case (wfCE_eqI \<Theta> \<B> \<Gamma> v1 b v2 )
+ hence "supp (CE_op Eq v1 v2) \<subseteq> atom_dom \<Gamma> \<union> supp \<B>" using ce.supp pure_supp
+ by (simp add: wfCE_eqI opp.supp)
+ then show ?case using ce.supp wfE_elims UnCI subsetCE subsetI x_not_in_b_set by auto
+next
+ case (wfCE_fstI \<Theta> \<B> \<Gamma> v1 b1 b2)
+ thus ?case using ce.supp wfCE_elims by simp
+next
+ case (wfCE_sndI \<Theta> \<B> \<Gamma> v1 b1 b2)
+ thus ?case using ce.supp wfCE_elims by simp
+next
+ case (wfCE_concatI \<Theta> \<B> \<Gamma> v1 v2)
+ thus ?case using ce.supp wfCE_elims by simp
+next
+ case (wfCE_lenI \<Theta> \<B> \<Gamma> v1)
+ thus ?case using ce.supp wfCE_elims by simp
+next
+ case (wfTI z \<Theta> \<B> \<Gamma> b c)
+ hence "supp c \<subseteq> supp z \<union> atom_dom \<Gamma> \<union> supp \<B>" using supp_at_base dom_cons by metis
+ moreover have "supp b \<subseteq> supp \<B>" using wfTI by auto
+ ultimately have " supp \<lbrace> z : b | c \<rbrace> \<subseteq> atom_dom \<Gamma> \<union> supp \<B>" using \<tau>.supp supp_at_base by force
+ thus ?case by auto
+qed(auto)
+
+lemma wf_supp2:
+ fixes \<Gamma>::\<Gamma> and \<Gamma>'::\<Gamma> and v::v and e::e and c::c and \<tau>::\<tau> and
+ ts::"(string*\<tau>) list" and \<Delta>::\<Delta> and s::s and b::b and ftq::fun_typ_q and
+ ft::fun_typ and ce::ce and td::type_def and cs::branch_s and css ::branch_list
+ shows
+ wfE_supp: "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f e : b \<Longrightarrow> (supp e \<subseteq> atom_dom \<Gamma> \<union> supp \<B> \<union> atom ` fst ` setD \<Delta>)" and (* \<and> ( \<Phi> = [] \<longrightarrow> supp e \<inter> supp \<B> = {})" and*)
+ wfS_supp: "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b \<Longrightarrow> supp s \<subseteq> atom_dom \<Gamma> \<union> atom ` fst ` setD \<Delta> \<union> supp \<B>" and
+ "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> ; tid ; dc ; t \<turnstile>\<^sub>w\<^sub>f cs : b \<Longrightarrow> supp cs \<subseteq> atom_dom \<Gamma> \<union> atom ` fst ` setD \<Delta> \<union> supp \<B>" and
+ "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> ; tid ; dclist \<turnstile>\<^sub>w\<^sub>f css : b \<Longrightarrow> supp css \<subseteq> atom_dom \<Gamma> \<union> atom ` fst ` setD \<Delta> \<union> supp \<B>" and
+ wfPhi_supp: "\<Theta> \<turnstile>\<^sub>w\<^sub>f (\<Phi>::\<Phi>) \<Longrightarrow> supp \<Phi> = {}" and
+ wfD_supp: "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<Longrightarrow> supp \<Delta> \<subseteq> atom`fst`(setD \<Delta>) \<union> atom_dom \<Gamma> \<union> supp \<B> " and
+ "\<Theta> ; \<Phi> \<turnstile>\<^sub>w\<^sub>f ftq \<Longrightarrow> supp ftq = {}" and
+ "\<Theta> ; \<Phi> ; \<B> \<turnstile>\<^sub>w\<^sub>f ft \<Longrightarrow> supp ft \<subseteq> supp \<B>"
+proof(induct rule:wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.inducts)
+ case (wfE_valI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v b)
+ hence "supp (AE_val v) \<subseteq> atom_dom \<Gamma> \<union> supp \<B>" using e.supp wf_supp1 by simp
+ then show ?case using e.supp wfE_elims UnCI subsetCE subsetI x_not_in_b_set by metis
+next
+ case (wfE_plusI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 v2)
+ hence "supp (AE_op Plus v1 v2) \<subseteq> atom_dom \<Gamma> \<union> supp \<B>"
+ using wfE_plusI opp.supp wf_supp1 e.supp pure_supp Un_least
+ by (metis sup_bot.left_neutral)
+
+ then show ?case using e.supp wfE_elims UnCI subsetCE subsetI x_not_in_b_set by metis
+next
+ case (wfE_leqI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 v2)
+ hence "supp (AE_op LEq v1 v2) \<subseteq> atom_dom \<Gamma> \<union> supp \<B>" using e.supp pure_supp Un_least
+ sup_bot.left_neutral using opp.supp wf_supp1 by auto
+ then show ?case using e.supp wfE_elims UnCI subsetCE subsetI x_not_in_b_set by metis
+next
+ case (wfE_eqI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 b v2)
+ hence "supp (AE_op Eq v1 v2) \<subseteq> atom_dom \<Gamma> \<union> supp \<B>" using e.supp pure_supp Un_least
+ sup_bot.left_neutral using opp.supp wf_supp1 by auto
+ then show ?case using e.supp wfE_elims UnCI subsetCE subsetI x_not_in_b_set by metis
+next
+ case (wfE_fstI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 b1 b2)
+ hence "supp (AE_fst v1 ) \<subseteq> atom_dom \<Gamma> \<union> supp \<B>" using e.supp pure_supp sup_bot.left_neutral using opp.supp wf_supp1 by auto
+ then show ?case using e.supp wfE_elims UnCI subsetCE subsetI x_not_in_b_set by metis
+next
+ case (wfE_sndI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 b1 b2)
+ hence "supp (AE_snd v1 ) \<subseteq> atom_dom \<Gamma> \<union> supp \<B>" using e.supp pure_supp wfE_plusI opp.supp wf_supp1 by (metis Un_least)
+ then show ?case using e.supp wfE_elims UnCI subsetCE subsetI x_not_in_b_set by metis
+next
+ case (wfE_concatI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 v2)
+ hence "supp (AE_concat v1 v2) \<subseteq> atom_dom \<Gamma> \<union> supp \<B>" using e.supp pure_supp
+ wfE_plusI opp.supp wf_supp1 by (metis Un_least)
+ then show ?case using e.supp wfE_elims UnCI subsetCE subsetI x_not_in_b_set by metis
+next
+ case (wfE_splitI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 v2)
+ hence "supp (AE_split v1 v2) \<subseteq> atom_dom \<Gamma> \<union> supp \<B>" using e.supp pure_supp
+ wfE_plusI opp.supp wf_supp1 by (metis Un_least)
+ then show ?case using e.supp wfE_elims UnCI subsetCE subsetI x_not_in_b_set by metis
+next
+ case (wfE_lenI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1)
+ hence "supp (AE_len v1 ) \<subseteq> atom_dom \<Gamma> \<union> supp \<B>" using e.supp pure_supp
+ using e.supp pure_supp sup_bot.left_neutral using opp.supp wf_supp1 by auto
+ then show ?case using e.supp wfE_elims UnCI subsetCE subsetI x_not_in_b_set by metis
+next
+ case (wfE_appI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> f x b c \<tau> s v)
+ then obtain b where "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b" using wfE_elims by metis
+ hence "supp v \<subseteq> atom_dom \<Gamma> \<union> supp \<B>" using wfE_appI wf_supp1 by metis
+ hence "supp (AE_app f v) \<subseteq> atom_dom \<Gamma> \<union> supp \<B>" using e.supp pure_supp by fast
+ then show ?case using e.supp(2) UnCI subsetCE subsetI wfE_appI using b.supp(3) pure_supp x_not_in_b_set by metis
+next
+ case (wfE_appPI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> b' bv v \<tau> f xa ba ca s)
+ then obtain b where "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : ( b[bv::=b']\<^sub>b)" using wfE_elims by metis
+ hence "supp v \<subseteq> atom_dom \<Gamma> \<union> supp \<B> " using wfE_appPI wf_supp1 by auto
+ moreover have "supp b' \<subseteq> supp \<B>" using wf_supp1(7) wfE_appPI by simp
+ ultimately show ?case unfolding e.supp using wfE_appPI pure_supp by fast
+next
+ case (wfE_mvarI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> u \<tau>)
+ then obtain \<tau> where "(u,\<tau>) \<in> setD \<Delta>" using wfE_elims(10) by metis
+ hence "atom u \<in> atom`fst`setD \<Delta>" by force
+ hence "supp (AE_mvar u ) \<subseteq> atom`fst`setD \<Delta>" using e.supp
+ by (simp add: supp_at_base)
+ thus ?case using UnCI subsetCE subsetI e.supp wfE_mvarI supp_at_base subsetCE supp_at_base u_not_in_b_set
+ by (simp add: supp_at_base)
+next
+ case (wfS_valI \<Theta> \<Phi> \<B> \<Gamma> v b \<Delta>)
+ then show ?case using wf_supp1
+ by (metis s_branch_s_branch_list.supp(1) sup.coboundedI2 sup_assoc sup_commute)
+next
+ case (wfS_letI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> e b' x s b)
+ then show ?case by auto
+next
+ case (wfS_let2I \<Theta> \<Phi> \<B> \<Gamma> \<Delta> s1 \<tau> x s2 b)
+ then show ?case unfolding s_branch_s_branch_list.supp (3) using wf_supp1(4)[OF wfS_let2I(3)] by auto
+next
+ case (wfS_ifI \<Theta> \<B> \<Gamma> v \<Phi> \<Delta> s1 b s2)
+ then show ?case using wf_supp1(1)[OF wfS_ifI(1)] by auto
+next
+ case (wfS_varI \<Theta> \<B> \<Gamma> \<tau> v u \<Delta> \<Phi> s b)
+ then show ?case using wf_supp1(1)[OF wfS_varI(2)] wf_supp1(4)[OF wfS_varI(1)] by auto
+next
+next
+ case (wfS_assignI u \<tau> \<Delta> \<Theta> \<B> \<Gamma> \<Phi> v)
+ hence "supp u \<subseteq> atom ` fst ` setD \<Delta>" proof(induct \<Delta> rule:\<Delta>_induct)
+ case DNil
+ then show ?case by auto
+ next
+ case (DCons u' t' \<Delta>')
+ show ?case proof(cases "u=u'")
+ case True
+ then show ?thesis using toSet.simps DCons supp_at_base by fastforce
+ next
+ case False
+ then show ?thesis using toSet.simps DCons supp_at_base wfS_assignI
+ by (metis empty_subsetI fstI image_eqI insert_subset)
+ qed
+ qed
+ then show ?case using s_branch_s_branch_list.supp(8) wfS_assignI wf_supp1(1)[OF wfS_assignI(6)] by auto
+next
+ case (wfS_matchI \<Theta> \<B> \<Gamma> v tid dclist \<Delta> \<Phi> cs b)
+ then show ?case using wf_supp1(1)[OF wfS_matchI(1)] by auto
+next
+ case (wfS_branchI \<Theta> \<Phi> \<B> x \<tau> \<Gamma> \<Delta> s b tid dc)
+ moreover have "supp s \<subseteq> supp x \<union> atom_dom \<Gamma> \<union> atom ` fst ` setD \<Delta> \<union> supp \<B>"
+ using dom_cons supp_at_base wfS_branchI by auto
+ moreover hence "supp s - set [atom x] \<subseteq> atom_dom \<Gamma> \<union> atom ` fst ` setD \<Delta> \<union> supp \<B>" using supp_at_base by force
+ ultimately have
+ "(supp s - set [atom x]) \<union> (supp dc ) \<subseteq> atom_dom \<Gamma> \<union> atom ` fst ` setD \<Delta> \<union> supp \<B>"
+ by (simp add: pure_supp)
+ thus ?case using s_branch_s_branch_list.supp(2) by auto
+next
+ case (wfD_emptyI \<Theta> \<B> \<Gamma>)
+ then show ?case using supp_DNil by auto
+next
+ case (wfD_cons \<Theta> \<B> \<Gamma> \<Delta> \<tau> u)
+ have "supp ((u, \<tau>) #\<^sub>\<Delta> \<Delta>) = supp u \<union> supp \<tau> \<union> supp \<Delta>" using supp_DCons supp_Pair by metis
+ also have "... \<subseteq> supp u \<union> atom ` fst ` setD \<Delta> \<union> atom_dom \<Gamma> \<union> supp \<B>"
+ using wfD_cons wf_supp1(4)[OF wfD_cons(3)] by auto
+ also have "... \<subseteq> atom ` fst ` setD ((u, \<tau>) #\<^sub>\<Delta> \<Delta>) \<union> atom_dom \<Gamma> \<union> supp \<B>" using supp_at_base by auto
+ finally show ?case by auto
+next
+ case (wfPhi_emptyI \<Theta>)
+ then show ?case using supp_Nil by auto
+next
+ case (wfPhi_consI f \<Theta> \<Phi> ft)
+ then show ?case using fun_def.supp
+ by (simp add: pure_supp supp_Cons)
+next
+ case (wfFTI \<Theta> B' b s x c \<tau> \<Phi>)
+ have " supp (AF_fun_typ x b c \<tau> s) = supp c \<union> (supp \<tau> \<union> supp s) - set [atom x] \<union> supp b" using fun_typ.supp by auto
+ thus ?case using wfFTI wf_supp1
+ proof -
+ have f1: "supp \<tau> \<subseteq> {atom x} \<union> atom_dom GNil \<union> supp B'"
+ using dom_cons wfFTI.hyps wf_supp1(4) by blast (* 0.0 ms *)
+ have "supp b \<subseteq> supp B'"
+ using wfFTI.hyps(1) wf_supp1(7) by blast (* 0.0 ms *)
+ then show ?thesis
+ using f1 \<open>supp (AF_fun_typ x b c \<tau> s) = supp c \<union> (supp \<tau> \<union> supp s) - set [atom x] \<union> supp b\<close>
+ wfFTI.hyps(4) wfFTI.hyps by auto (* 234 ms *)
+ qed
+next
+ case (wfFTNone \<Theta> \<Phi> ft)
+ then show ?case by (simp add: fun_typ_q.supp(2))
+next
+ case (wfFTSome \<Theta> \<Phi> bv ft)
+ then show ?case using fun_typ_q.supp
+ by (simp add: supp_at_base)
+next
+ case (wfS_assertI \<Theta> \<Phi> \<B> x c \<Gamma> \<Delta> s b)
+ then have "supp c \<subseteq> atom_dom \<Gamma> \<union> atom ` fst ` setD \<Delta> \<union> supp \<B>" using wf_supp1
+ by (metis Un_assoc Un_commute le_supI2)
+ moreover have "supp s \<subseteq> atom_dom \<Gamma> \<union> atom ` fst ` setD \<Delta> \<union> supp \<B>" proof
+ fix z
+ assume *:"z \<in> supp s"
+ have **:"atom x \<notin> supp s" using wfS_assertI fresh_prodN fresh_def by metis
+ have "z \<in> atom_dom ((x, B_bool, c) #\<^sub>\<Gamma> \<Gamma>) \<union> atom ` fst ` setD \<Delta> \<union> supp \<B>" using wfS_assertI * by blast
+ have "z \<in> atom_dom ((x, B_bool, c) #\<^sub>\<Gamma> \<Gamma>) \<Longrightarrow> z \<in> atom_dom \<Gamma>" using * ** by auto
+ thus "z \<in> atom_dom \<Gamma> \<union> atom ` fst ` setD \<Delta> \<union> supp \<B>" using * **
+ using \<open>z \<in> atom_dom ((x, B_bool, c) #\<^sub>\<Gamma> \<Gamma>) \<union> atom ` fst ` setD \<Delta> \<union> supp \<B>\<close> by blast
+ qed
+ ultimately show ?case by auto
+qed(auto)
+
+lemmas wf_supp = wf_supp1 wf_supp2
+
+lemma wfV_supp_nil:
+ fixes v::v
+ assumes "P ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f v : b"
+ shows "supp v = {}"
+ using wfV_supp[of P " {||}" GNil v b] dom.simps toSet.simps
+ using assms by auto
+
+lemma wfT_TRUE_aux:
+ assumes "wfG P \<B> \<Gamma>" and "atom z \<sharp> (P, \<B>, \<Gamma>)" and "wfB P \<B> b"
+ shows "wfT P \<B> \<Gamma> (\<lbrace> z : b | TRUE \<rbrace>)"
+proof (rule)
+ show \<open> atom z \<sharp> (P, \<B>, \<Gamma>)\<close> using assms by auto
+ show \<open> P; \<B> \<turnstile>\<^sub>w\<^sub>f b \<close> using assms by auto
+ show \<open> P ;\<B> ; (z, b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f TRUE \<close> using wfG_cons2I wfC_trueI assms by auto
+qed
+
+lemma wfT_TRUE:
+ assumes "wfG P \<B> \<Gamma>" and "wfB P \<B> b"
+ shows "wfT P \<B> \<Gamma> (\<lbrace> z : b | TRUE \<rbrace>)"
+proof -
+ obtain z'::x where *:"atom z' \<sharp> (P, \<B>, \<Gamma>)" using obtain_fresh by metis
+ hence "\<lbrace> z : b | TRUE \<rbrace> = \<lbrace> z' : b | TRUE \<rbrace>" by auto
+ thus ?thesis using wfT_TRUE_aux assms * by metis
+qed
+
+lemma phi_flip_eq:
+ assumes "wfPhi T P"
+ shows "(x \<leftrightarrow> xa) \<bullet> P = P"
+ using wfPhi_supp[OF assms] flip_fresh_fresh fresh_def by blast
+
+lemma wfC_supp_cons:
+ fixes c'::c and G::\<Gamma>
+ assumes "P; \<B> ; (x', b' , TRUE) #\<^sub>\<Gamma>G \<turnstile>\<^sub>w\<^sub>f c'"
+ shows "supp c' \<subseteq> atom_dom G \<union> supp x' \<union> supp \<B>" and "supp c' \<subseteq> supp G \<union> supp x' \<union> supp \<B>"
+proof -
+ show "supp c' \<subseteq> atom_dom G \<union> supp x' \<union> supp \<B>"
+ using wfC_supp[OF assms] dom_cons supp_at_base by blast
+ moreover have "atom_dom G \<subseteq> supp G"
+ by (meson assms wfC_wf wfG_cons wfG_supp)
+ ultimately show "supp c' \<subseteq> supp G \<union> supp x' \<union> supp \<B>" using wfG_supp assms wfG_cons wfC_wf by fast
+qed
+
+lemma wfG_dom_supp:
+ fixes x::x
+ assumes "wfG P \<B> G"
+ shows "atom x \<in> atom_dom G \<longleftrightarrow> atom x \<in> supp G"
+using assms proof(induct G rule: \<Gamma>_induct)
+ case GNil
+ then show ?case using dom.simps supp_of_atom_list
+ using supp_GNil by auto
+next
+ case (GCons x' b' c' G)
+
+ show ?case proof(cases "x' = x")
+ case True
+ then show ?thesis using dom.simps supp_of_atom_list supp_at_base
+ using supp_GCons by auto
+ next
+ case False
+ have "(atom x \<in> atom_dom ((x', b', c') #\<^sub>\<Gamma> G)) = (atom x \<in> atom_dom G)" using atom_dom.simps False by simp
+ also have "... = (atom x \<in> supp G)" using GCons wfG_elims by metis
+ also have "... = (atom x \<in> (supp (x', b', c') \<union> supp G))" proof
+ show "atom x \<in> supp G \<Longrightarrow> atom x \<in> supp (x', b', c') \<union> supp G" by auto
+ assume "atom x \<in> supp (x', b', c') \<union> supp G"
+ then consider "atom x \<in> supp (x', b', c')" | "atom x \<in> supp G" by auto
+ then show "atom x \<in> supp G" proof(cases)
+ case 1
+ assume " atom x \<in> supp (x', b', c') "
+ hence "atom x \<in> supp c'" using supp_triple False supp_b_empty supp_at_base by force
+
+ moreover have "P; \<B> ; (x', b' , TRUE) #\<^sub>\<Gamma>G \<turnstile>\<^sub>w\<^sub>f c'" using wfG_elim2 GCons by simp
+ moreover hence "supp c' \<subseteq> supp G \<union> supp x' \<union> supp \<B>" using wfC_supp_cons by auto
+ ultimately have "atom x \<in> supp G \<union> supp x' " using x_not_in_b_set by auto
+ then show ?thesis using False supp_at_base by (simp add: supp_at_base)
+ next
+ case 2
+ then show ?thesis by simp
+ qed
+ qed
+ also have "... = (atom x \<in> supp ((x', b', c') #\<^sub>\<Gamma> G))" using supp_at_base False supp_GCons by simp
+ finally show ?thesis by simp
+ qed
+qed
+
+lemma wfG_atoms_supp_eq :
+ fixes x::x
+ assumes "wfG P \<B> G"
+ shows "atom x \<in> atom_dom G \<longleftrightarrow> atom x \<in> supp G"
+ using wfG_dom_supp assms by auto
+
+lemma beta_flip_eq:
+ fixes x::x and xa::x and \<B>::\<B>
+ shows "(x \<leftrightarrow> xa) \<bullet> \<B> = \<B>"
+proof -
+ have "atom x \<sharp> \<B> \<and> atom xa \<sharp> \<B>" using x_not_in_b_set fresh_def supp_set by metis
+ thus ?thesis by (simp add: flip_fresh_fresh fresh_def)
+qed
+
+lemma theta_flip_eq2:
+ assumes "\<turnstile>\<^sub>w\<^sub>f \<Theta>"
+ shows " (z \<leftrightarrow> za ) \<bullet> \<Theta> = \<Theta>"
+proof -
+ have "supp \<Theta> = {}" using wfTh_supp assms by simp
+ thus ?thesis
+ by (simp add: flip_fresh_fresh fresh_def)
+ qed
+
+lemma theta_flip_eq:
+ assumes "wfTh \<Theta>"
+ shows "(x \<leftrightarrow> xa) \<bullet> \<Theta> = \<Theta>"
+ using wfTh_supp flip_fresh_fresh fresh_def
+ by (simp add: assms theta_flip_eq2)
+
+lemma wfT_wfC:
+ fixes c::c
+ assumes "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b | c \<rbrace>" and "atom z \<sharp> \<Gamma>"
+ shows "\<Theta>; \<B>; (z,b,TRUE) #\<^sub>\<Gamma>\<Gamma> \<turnstile>\<^sub>w\<^sub>f c"
+proof -
+ obtain za ba ca where *:"\<lbrace> z : b | c \<rbrace> = \<lbrace> za : ba | ca \<rbrace> \<and> atom za \<sharp> (\<Theta>,\<B>,\<Gamma>) \<and> \<Theta>; \<B>; (za, ba, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f ca"
+ using wfT_elims[OF assms(1)] by metis
+ hence c1: "[[atom z]]lst. c = [[atom za]]lst. ca" using \<tau>.eq_iff by meson
+ show ?thesis proof(cases "z=za")
+ case True
+ hence "ca = c" using c1 by (simp add: Abs1_eq_iff(3))
+ then show ?thesis using * True by simp
+ next
+ case False
+ have " \<turnstile>\<^sub>w\<^sub>f \<Theta>" using wfT_wf wfG_wf assms by metis
+ moreover have "atom za \<sharp> \<Gamma>" using * fresh_prodN by auto
+ ultimately have "\<Theta>; \<B>; (z \<leftrightarrow> za ) \<bullet> (za, ba, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f (z \<leftrightarrow> za ) \<bullet> ca"
+ using wfC.eqvt theta_flip_eq2 beta_flip_eq * GCons_eqvt assms flip_fresh_fresh by metis
+ moreover have "atom z \<sharp> ca"
+ proof -
+ have "supp ca \<subseteq> atom_dom \<Gamma> \<union> { atom za } \<union> supp \<B>" using * wfC_supp atom_dom.simps toSet.simps by fastforce
+ moreover have "atom z \<notin> atom_dom \<Gamma> " using assms fresh_def wfT_wf wfG_dom_supp wfC_supp by metis
+ moreover hence "atom z \<notin> atom_dom \<Gamma> \<union> { atom za }" using False by simp
+ moreover have "atom z \<notin> supp \<B>" using x_not_in_b_set by simp
+ ultimately show ?thesis using fresh_def False by fast
+ qed
+ moreover hence "(z \<leftrightarrow> za ) \<bullet> ca = c" using type_eq_subst_eq1(3) * by metis
+ ultimately show ?thesis using assms G_cons_flip_fresh * by auto
+ qed
+qed
+
+lemma u_not_in_dom_g:
+ fixes u::u
+ shows "atom u \<notin> atom_dom G"
+ using toSet.simps atom_dom.simps u_not_in_x_atoms by auto
+
+lemma bv_not_in_dom_g:
+ fixes bv::bv
+ shows "atom bv \<notin> atom_dom G"
+ using toSet.simps atom_dom.simps u_not_in_x_atoms by auto
+
+text \<open>An important lemma that confirms that @{term \<Gamma>} does not rely on mutable variables\<close>
+lemma u_not_in_g:
+ fixes u::u
+ assumes "wfG \<Theta> B G"
+ shows "atom u \<notin> supp G"
+using assms proof(induct G rule: \<Gamma>_induct)
+case GNil
+ then show ?case using supp_GNil fresh_def
+ using fresh_set_empty by fastforce
+next
+ case (GCons x b c \<Gamma>')
+ moreover hence "atom u \<notin> supp b" using
+ wfB_supp wfC_supp u_not_in_x_atoms wfG_elims wfX_wfY by auto
+ moreover hence "atom u \<notin> supp x" using u_not_in_x_atoms supp_at_base by blast
+ moreover hence "atom u \<notin> supp c" proof -
+ have "\<Theta> ; B ; (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>' \<turnstile>\<^sub>w\<^sub>f c" using wfG_cons_wfC GCons by simp
+ hence "supp c \<subseteq> atom_dom ((x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>') \<union> supp B" using wfC_supp by blast
+ thus ?thesis using u_not_in_dom_g u_not_in_b_atoms
+ using u_not_in_b_set by auto
+ qed
+ ultimately have "atom u \<notin> supp (x,b,c)" using supp_Pair by simp
+ thus ?case using supp_GCons GCons wfG_elims by blast
+qed
+
+text \<open>An important lemma that confirms that types only depend on immutable variables\<close>
+lemma u_not_in_t:
+ fixes u::u
+ assumes "wfT \<Theta> B G \<tau>"
+ shows "atom u \<notin> supp \<tau>"
+proof -
+ have "supp \<tau> \<subseteq> atom_dom G \<union> supp B" using wfT_supp assms by auto
+ thus ?thesis using u_not_in_dom_g u_not_in_b_set by blast
+qed
+
+lemma wfT_supp_c:
+ fixes \<B>::\<B> and z::x
+ assumes "wfT P \<B> \<Gamma> (\<lbrace> z : b | c \<rbrace>)"
+ shows "supp c - { atom z } \<subseteq> atom_dom \<Gamma> \<union> supp \<B>"
+ using wf_supp \<tau>.supp assms
+ by (metis Un_subset_iff empty_set list.simps(15))
+
+lemma wfG_wfC[ms_wb]:
+ assumes "wfG P \<B> ((x,b,c) #\<^sub>\<Gamma>\<Gamma>)"
+ shows "wfC P \<B> ((x,b,TRUE) #\<^sub>\<Gamma>\<Gamma>) c"
+using assms proof(cases "c \<in> {TRUE,FALSE}")
+ case True
+ have "atom x \<sharp> \<Gamma> \<and> wfG P \<B> \<Gamma> \<and> wfB P \<B> b" using wfG_cons assms by auto
+ hence "wfG P \<B> ((x,b,TRUE) #\<^sub>\<Gamma>\<Gamma>)" using wfG_cons2I by auto
+ then show ?thesis using wfC_trueI wfC_falseI True by auto
+next
+ case False
+ then show ?thesis using wfG_elims assms by blast
+qed
+
+lemma wfT_wf_cons:
+ assumes "wfT P \<B> \<Gamma> \<lbrace> z : b | c \<rbrace>" and "atom z \<sharp> \<Gamma>"
+ shows "wfG P \<B> ((z,b,c) #\<^sub>\<Gamma>\<Gamma>)"
+using assms proof(cases "c \<in> { TRUE,FALSE }")
+ case True
+ then show ?thesis using wfT_wfC wfC_wf wfG_wfB wfG_cons2I assms wfT_wf by fastforce
+next
+ case False
+ then show ?thesis using wfT_wfC wfC_wf wfG_wfB wfG_cons1I wfT_wf wfT_wfC assms by fastforce
+qed
+
+lemma wfV_b_fresh:
+ fixes b::b and v::v and bv::bv
+ assumes "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v: b" and "bv |\<notin>| \<B>"
+ shows "atom bv \<sharp> v"
+using wfV_supp bv_not_in_dom_g fresh_def assms bv_not_in_bset_supp by blast
+
+lemma wfCE_b_fresh:
+ fixes b::b and ce::ce and bv::bv
+ assumes "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ce: b" and "bv |\<notin>| \<B>"
+ shows "atom bv \<sharp> ce"
+using bv_not_in_dom_g fresh_def assms bv_not_in_bset_supp wf_supp1(8) by fast
+
+section \<open>Freshness\<close>
+
+lemma wfG_fresh_x:
+ fixes \<Gamma>::\<Gamma> and z::x
+ assumes "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>" and "atom z \<sharp> \<Gamma>"
+ shows "atom z \<sharp> (\<Theta>,\<B>, \<Gamma>)"
+unfolding fresh_prodN apply(intro conjI)
+ using wf_supp1 wfX_wfY assms fresh_def x_not_in_b_set by(metis empty_iff)+
+
+lemma wfG_wfT:
+ assumes "wfG P \<B> ((x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> G)" and "atom x \<sharp> c"
+ shows "P; \<B> ; G \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b | c \<rbrace>"
+proof -
+ have " P; \<B> ; (x, b, TRUE) #\<^sub>\<Gamma> G \<turnstile>\<^sub>w\<^sub>f c[z::=V_var x]\<^sub>c\<^sub>v \<and> wfB P \<B> b" using assms
+ using wfG_elim2 by auto
+ moreover have "atom x \<sharp> (P ,\<B>,G)" using wfG_elims assms wfG_fresh_x by metis
+ ultimately have "wfT P \<B> G \<lbrace> x : b | c[z::=V_var x]\<^sub>c\<^sub>v \<rbrace>" using wfTI assms by metis
+ moreover have "\<lbrace> x : b | c[z::=V_var x]\<^sub>c\<^sub>v \<rbrace> = \<lbrace> z : b | c \<rbrace>" using type_eq_subst \<open>atom x \<sharp> c\<close> by auto
+ ultimately show ?thesis by auto
+qed
+
+lemma wfT_wfT_if:
+ assumes "wfT \<Theta> \<B> \<Gamma> (\<lbrace> z2 : b | CE_val v == CE_val (V_lit L_false) IMP c[z::=V_var z2]\<^sub>c\<^sub>v \<rbrace>)" and "atom z2 \<sharp> (c,\<Gamma>)"
+ shows "wfT \<Theta> \<B> \<Gamma> \<lbrace> z : b | c \<rbrace>"
+proof -
+ have *: "atom z2 \<sharp> (\<Theta>, \<B>, \<Gamma>)" using wfG_fresh_x wfX_wfY assms fresh_Pair by metis
+ have "wfB \<Theta> \<B> b" using assms wfT_elims by metis
+ have "\<Theta>; \<B>; (GCons (z2,b,TRUE) \<Gamma>) \<turnstile>\<^sub>w\<^sub>f (CE_val v == CE_val (V_lit L_false) IMP c[z::=V_var z2]\<^sub>c\<^sub>v)" using wfT_wfC assms fresh_Pair by auto
+ hence "\<Theta>; \<B>; ((z2,b,TRUE) #\<^sub>\<Gamma>\<Gamma>) \<turnstile>\<^sub>w\<^sub>f c[z::=V_var z2]\<^sub>c\<^sub>v" using wfC_elims by metis
+ hence "wfT \<Theta> \<B> \<Gamma> (\<lbrace> z2 : b | c[z::=V_var z2]\<^sub>c\<^sub>v\<rbrace>)" using assms fresh_Pair wfTI \<open>wfB \<Theta> \<B> b\<close> * by auto
+ moreover have "\<lbrace> z : b | c \<rbrace> = \<lbrace> z2 : b | c[z::=V_var z2]\<^sub>c\<^sub>v \<rbrace>" using type_eq_subst assms fresh_Pair by auto
+ ultimately show ?thesis using wfTI assms by argo
+qed
+
+lemma wfT_fresh_c:
+ fixes x::x
+ assumes "wfT P \<B> \<Gamma> \<lbrace> z : b | c \<rbrace>" and "atom x \<sharp> \<Gamma>" and "x \<noteq> z"
+ shows "atom x \<sharp> c"
+proof(rule ccontr)
+ assume "\<not> atom x \<sharp> c"
+ hence *:"atom x \<in> supp c" using fresh_def by auto
+ moreover have "supp c - set [atom z] \<union> supp b \<subseteq> atom_dom \<Gamma> \<union> supp \<B>"
+ using assms wfT_supp \<tau>.supp by blast
+ moreover hence "atom x \<in> supp c - set [atom z]" using assms * by auto
+ ultimately have "atom x \<in> atom_dom \<Gamma>" using x_not_in_b_set by auto
+ thus False using assms wfG_atoms_supp_eq wfT_wf fresh_def by metis
+qed
+
+lemma wfG_x_fresh [simp]:
+ fixes x::x
+ assumes "wfG P \<B> G"
+ shows "atom x \<notin> atom_dom G \<longleftrightarrow> atom x \<sharp> G"
+ using wfG_atoms_supp_eq assms fresh_def by metis
+
+lemma wfD_x_fresh:
+ fixes x::x
+ assumes "atom x \<sharp> \<Gamma>" and "wfD P B \<Gamma> \<Delta>"
+ shows "atom x \<sharp> \<Delta>"
+using assms proof(induct \<Delta> rule: \<Delta>_induct)
+ case DNil
+ then show ?case using supp_DNil fresh_def by auto
+next
+ case (DCons u' t' \<Delta>')
+ have wfg: "wfG P B \<Gamma>" using wfD_wf DCons by blast
+ hence wfd: "wfD P B \<Gamma> \<Delta>'" using wfD_elims DCons by blast
+ have "supp t' \<subseteq> atom_dom \<Gamma> \<union> supp B" using wfT_supp DCons wfD_elims by metis
+ moreover have "atom x \<notin> atom_dom \<Gamma>" using DCons(2) fresh_def wfG_supp wfg by blast
+ ultimately have "atom x \<sharp> t'" using fresh_def DCons wfG_supp wfg x_not_in_b_set by blast
+ moreover have "atom x \<sharp> u'" using supp_at_base fresh_def by fastforce
+ ultimately have "atom x \<sharp> (u',t')" using supp_Pair by fastforce
+ thus ?case using DCons fresh_DCons wfd by fast
+qed
+
+lemma wfG_fresh_x2:
+ fixes \<Gamma>::\<Gamma> and z::x and \<Delta>::\<Delta> and \<Phi>::\<Phi>
+ assumes "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>" and "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>" and "atom z \<sharp> \<Gamma>"
+ shows "atom z \<sharp> (\<Theta>,\<Phi>,\<B>, \<Gamma>,\<Delta>)"
+ unfolding fresh_prodN apply(intro conjI)
+ using wfG_fresh_x assms fresh_prod3 wfX_wfY apply metis
+ using wf_supp2(5) assms fresh_def apply blast
+ using assms wfG_fresh_x wfX_wfY fresh_prod3 apply metis
+ using assms wfG_fresh_x wfX_wfY fresh_prod3 apply metis
+ using wf_supp2(6) assms fresh_def wfD_x_fresh by metis
+
+lemma wfV_x_fresh:
+ fixes v::v and b::b and \<Gamma>::\<Gamma> and x::x
+ assumes "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b" and "atom x \<sharp> \<Gamma>"
+ shows "atom x \<sharp> v"
+proof -
+ have "supp v \<subseteq> atom_dom \<Gamma> \<union> supp \<B> " using assms wfV_supp by auto
+ moreover have "atom x \<notin> atom_dom \<Gamma>" using fresh_def assms
+ dom.simps subsetCE wfG_elims wfG_supp by (metis dom_supp_g)
+ moreover have "atom x \<notin> supp \<B>" using x_not_in_b_set by auto
+ ultimately show ?thesis using fresh_def by fast
+qed
+
+lemma wfE_x_fresh:
+ fixes e::e and b::b and \<Gamma>::\<Gamma> and \<Delta>::\<Delta> and \<Phi>::\<Phi> and x::x
+ assumes "\<Theta>; \<Phi>; \<B>; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f e : b" and "atom x \<sharp> \<Gamma>"
+ shows "atom x \<sharp> e"
+proof -
+ have "wfG \<Theta> \<B> \<Gamma>" using assms wfE_wf by auto
+ hence "supp e \<subseteq> atom_dom \<Gamma> \<union> supp \<B> \<union> atom`fst`setD \<Delta>" using wfE_supp dom.simps assms by auto
+ moreover have "atom x \<notin> atom_dom \<Gamma>" using fresh_def assms
+ dom.simps subsetCE \<open>wfG \<Theta> \<B> \<Gamma>\<close> wfG_supp by (metis dom_supp_g)
+ moreover have "atom x \<notin> atom`fst`setD \<Delta>" by auto
+ ultimately show ?thesis using fresh_def x_not_in_b_set by fast
+qed
+
+lemma wfT_x_fresh:
+ fixes \<tau>::\<tau> and \<Gamma>::\<Gamma> and x::x
+ assumes "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau>" and "atom x \<sharp> \<Gamma>"
+ shows "atom x \<sharp> \<tau>"
+proof -
+ have "wfG \<Theta> \<B> \<Gamma>" using assms wfX_wfY by auto
+ hence "supp \<tau> \<subseteq> atom_dom \<Gamma> \<union> supp \<B>" using wfT_supp dom.simps assms by auto
+ moreover have "atom x \<notin> atom_dom \<Gamma>" using fresh_def assms
+ dom.simps subsetCE \<open>wfG \<Theta> \<B> \<Gamma>\<close> wfG_supp by (metis dom_supp_g)
+ moreover have "atom x \<notin> supp \<B>" using x_not_in_b_set by simp
+ ultimately show ?thesis using fresh_def by fast
+qed
+
+lemma wfS_x_fresh:
+ fixes s::s and \<Delta>::\<Delta> and x::x
+ assumes "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b" and "atom x \<sharp> \<Gamma>"
+ shows "atom x \<sharp> s"
+proof -
+ have "supp s \<subseteq> atom_dom \<Gamma> \<union> atom ` fst ` setD \<Delta> \<union> supp \<B>" using wf_supp assms by metis
+ moreover have "atom x \<notin> atom ` fst ` setD \<Delta>" by auto
+ moreover have "atom x \<notin> atom_dom \<Gamma>" using assms fresh_def wfG_dom_supp wfX_wfY by metis
+ moreover have "atom x \<notin> supp \<B>" using supp_b_empty supp_fset
+ by (simp add: x_not_in_b_set)
+ ultimately show ?thesis using fresh_def by fast
+qed
+
+lemma wfTh_fresh:
+ fixes x
+ assumes "wfTh T"
+ shows "atom x \<sharp> T"
+ using wf_supp1 assms fresh_def by fastforce
+
+lemmas wfTh_x_fresh = wfTh_fresh
+
+lemma wfPhi_fresh:
+ fixes x
+ assumes "wfPhi T P"
+ shows "atom x \<sharp> P"
+ using wf_supp assms fresh_def by fastforce
+
+lemmas wfPhi_x_fresh = wfPhi_fresh
+lemmas wb_x_fresh = wfTh_x_fresh wfPhi_x_fresh wfD_x_fresh wfT_x_fresh wfV_x_fresh
+
+lemma wfG_inside_fresh[ms_fresh]:
+ fixes \<Gamma>::\<Gamma> and x::x
+ assumes "wfG P \<B> (\<Gamma>'@((x,b,c) #\<^sub>\<Gamma>\<Gamma>))"
+ shows "atom x \<notin> atom_dom \<Gamma>'"
+using assms proof(induct \<Gamma>' rule: \<Gamma>_induct)
+ case GNil
+ then show ?case by auto
+next
+ case (GCons x1 b1 c1 \<Gamma>1)
+ moreover hence "atom x \<notin> atom ` fst `({(x1,b1,c1)})" proof -
+ have *: "P; \<B> \<turnstile>\<^sub>w\<^sub>f (\<Gamma>1 @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>)" using wfG_elims append_g.simps GCons by metis
+ have "atom x1 \<sharp> (\<Gamma>1 @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>)" using GCons wfG_elims append_g.simps by metis
+ hence "atom x1 \<notin> atom_dom (\<Gamma>1 @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>)" using wfG_dom_supp fresh_def * by metis
+ thus ?thesis by auto
+ qed
+ ultimately show ?case using append_g.simps atom_dom.simps toSet.simps wfG_elims dom.simps
+ by (metis image_insert insert_iff insert_is_Un)
+qed
+
+lemma wfG_inside_x_in_atom_dom:
+ fixes c::c and x::x and \<Gamma>::\<Gamma>
+ shows "atom x \<in> atom_dom ( \<Gamma>'@ (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)"
+ by(induct \<Gamma>' rule: \<Gamma>_induct, (simp add: toSet.simps atom_dom.simps)+)
+
+lemma wfG_inside_x_neq:
+ fixes c::c and x::x and \<Gamma>::\<Gamma> and G::\<Gamma> and xa::x
+ assumes "G=( \<Gamma>'@ (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)" and "atom xa \<sharp> G" and " \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f G"
+ shows "xa \<noteq> x"
+proof -
+ have "atom xa \<notin> atom_dom G" using fresh_def wfG_atoms_supp_eq assms by metis
+ moreover have "atom x \<in> atom_dom G" using wfG_inside_x_in_atom_dom assms by simp
+ ultimately show ?thesis by auto
+qed
+
+lemma wfG_inside_x_fresh:
+ fixes c::c and x::x and \<Gamma>::\<Gamma> and G::\<Gamma> and xa::x
+ assumes "G=( \<Gamma>'@ (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)" and "atom xa \<sharp> G" and " \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f G"
+ shows "atom xa \<sharp> x"
+ using fresh_def supp_at_base wfG_inside_x_neq assms by auto
+
+lemma wfT_nil_supp:
+ fixes t::\<tau>
+ assumes "\<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f t"
+ shows "supp t = {}"
+ using wfT_supp atom_dom.simps assms toSet.simps by force
+
+section \<open>Misc\<close>
+
+lemma wfG_cons_append:
+ fixes b'::b
+ assumes "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f ((x', b', c') #\<^sub>\<Gamma> \<Gamma>') @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>"
+ shows "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (\<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>) \<and> atom x' \<sharp> (\<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>) \<and> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b' \<and> x' \<noteq> x"
+proof -
+ have "((x', b', c') #\<^sub>\<Gamma> \<Gamma>') @ (x, b, c) #\<^sub>\<Gamma> \<Gamma> = (x', b', c') #\<^sub>\<Gamma> (\<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>)" using append_g.simps by auto
+ hence *:"\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (\<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>) \<and> atom x' \<sharp> (\<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>) \<and> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b'" using assms wfG_cons by metis
+ moreover have "atom x' \<sharp> x" proof(rule wfG_inside_x_fresh[of "(\<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>)"])
+ show "\<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma> = \<Gamma>' @ (x, b, c[x::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>" by simp
+ show " atom x' \<sharp> \<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>" using * by auto
+ show "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma> " using * by auto
+ qed
+ ultimately show ?thesis by auto
+qed
+
+lemma flip_u_eq:
+ fixes u::u and u'::u and \<Theta>::\<Theta> and \<tau>::\<tau>
+ assumes "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau>"
+ shows "(u \<leftrightarrow> u') \<bullet> \<tau> = \<tau>" and "(u \<leftrightarrow> u') \<bullet> \<Gamma> = \<Gamma>" and "(u \<leftrightarrow> u') \<bullet> \<Theta> = \<Theta>" and "(u \<leftrightarrow> u') \<bullet> \<B> = \<B>"
+proof -
+ show "(u \<leftrightarrow> u') \<bullet> \<tau> = \<tau>" using wfT_supp flip_fresh_fresh
+ by (metis assms(1) fresh_def u_not_in_t)
+ show "(u \<leftrightarrow> u') \<bullet> \<Gamma> = \<Gamma>" using u_not_in_g wfX_wfY assms flip_fresh_fresh fresh_def by metis
+ show "(u \<leftrightarrow> u') \<bullet> \<Theta> = \<Theta>" using theta_flip_eq assms wfX_wfY by metis
+ show "(u \<leftrightarrow> u') \<bullet> \<B> = \<B>" using u_not_in_b_set flip_fresh_fresh fresh_def by metis
+qed
+
+lemma wfT_wf_cons_flip:
+ fixes c::c and x::x
+ assumes "wfT P \<B> \<Gamma> \<lbrace> z : b | c \<rbrace>" and "atom x \<sharp> (c,\<Gamma>)"
+ shows "wfG P \<B> ((x,b,c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma>\<Gamma>)"
+proof -
+ have "\<lbrace> x : b | c[z::=V_var x]\<^sub>c\<^sub>v \<rbrace> = \<lbrace> z : b | c \<rbrace>" using assms freshers type_eq_subst by metis
+ hence *:"wfT P \<B> \<Gamma> \<lbrace> x : b | c[z::=V_var x]\<^sub>c\<^sub>v \<rbrace>" using assms by metis
+ show ?thesis proof(rule wfG_consI)
+ show \<open> P; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<close> using assms wfT_wf by auto
+ show \<open>atom x \<sharp> \<Gamma>\<close> using assms by auto
+ show \<open> P; \<B> \<turnstile>\<^sub>w\<^sub>f b \<close> using assms wfX_wfY b_of.simps by metis
+ show \<open> P; \<B> ; (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f c[z::=V_var x]\<^sub>c\<^sub>v \<close> using wfT_wfC * assms fresh_Pair by metis
+ qed
+qed
+
+section \<open>Context Strengthening\<close>
+
+text \<open>We can remove an entry for a variable from the context if the variable doesn't appear in the
+term and the variable is not used later in the context or any other context\<close>
+
+lemma fresh_restrict:
+ fixes y::"'a::at_base" and \<Gamma>::\<Gamma>
+ assumes "atom y \<sharp> (\<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>)"
+ shows "atom y \<sharp> (\<Gamma>'@\<Gamma>)"
+using assms proof(induct \<Gamma>' rule: \<Gamma>_induct)
+ case GNil
+ then show ?case using fresh_GCons fresh_GNil by auto
+next
+ case (GCons x' b' c' \<Gamma>'')
+ then show ?case using fresh_GCons fresh_GNil by auto
+qed
+
+lemma wf_restrict1:
+ fixes \<Gamma>::\<Gamma> and \<Gamma>'::\<Gamma> and v::v and e::e and c::c and \<tau>::\<tau> and ts::"(string*\<tau>) list" and \<Delta>::\<Delta> and s::s and b::b and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def
+ and cs::branch_s and css::branch_list
+ shows "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b \<Longrightarrow> \<Gamma>=\<Gamma>\<^sub>1@((x,b',c') #\<^sub>\<Gamma>\<Gamma>\<^sub>2) \<Longrightarrow> atom x \<sharp> v \<Longrightarrow> atom x \<sharp> \<Gamma>\<^sub>1 \<Longrightarrow> \<Theta>; \<B>; \<Gamma>\<^sub>1@\<Gamma>\<^sub>2 \<turnstile>\<^sub>w\<^sub>f v : b" and
+
+ "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f c \<Longrightarrow> \<Gamma>=\<Gamma>\<^sub>1@((x,b',c') #\<^sub>\<Gamma>\<Gamma>\<^sub>2) \<Longrightarrow> atom x \<sharp> c\<Longrightarrow> atom x \<sharp> \<Gamma>\<^sub>1 \<Longrightarrow> \<Theta> ; \<B> ; \<Gamma>\<^sub>1@\<Gamma>\<^sub>2 \<turnstile>\<^sub>w\<^sub>f c" and
+ "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<Longrightarrow> \<Gamma>=\<Gamma>\<^sub>1@((x,b',c') #\<^sub>\<Gamma>\<Gamma>\<^sub>2) \<Longrightarrow> atom x \<sharp> \<Gamma>\<^sub>1 \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>\<^sub>1@\<Gamma>\<^sub>2" and
+ "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau> \<Longrightarrow> \<Gamma>=\<Gamma>\<^sub>1@((x,b',c') #\<^sub>\<Gamma>\<Gamma>\<^sub>2) \<Longrightarrow> atom x \<sharp> \<tau>\<Longrightarrow> atom x \<sharp> \<Gamma>\<^sub>1 \<Longrightarrow> \<Theta>; \<B>; \<Gamma>\<^sub>1@\<Gamma>\<^sub>2 \<turnstile>\<^sub>w\<^sub>f \<tau>" and
+ "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ts \<Longrightarrow> True" and
+ "\<turnstile>\<^sub>w\<^sub>f \<Theta> \<Longrightarrow>True" and
+
+ "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b \<Longrightarrow> True" and
+
+ "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ce : b \<Longrightarrow> \<Gamma>=\<Gamma>\<^sub>1@((x,b',c') #\<^sub>\<Gamma>\<Gamma>\<^sub>2) \<Longrightarrow> atom x \<sharp> ce \<Longrightarrow> atom x \<sharp> \<Gamma>\<^sub>1 \<Longrightarrow> \<Theta>; \<B>; \<Gamma>\<^sub>1@\<Gamma>\<^sub>2 \<turnstile>\<^sub>w\<^sub>f ce : b" and
+ "\<Theta> \<turnstile>\<^sub>w\<^sub>f td \<Longrightarrow> True"
+proof(induct arbitrary: \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1
+ rule:wfV_wfC_wfG_wfT_wfTs_wfTh_wfB_wfCE_wfTD.inducts)
+ case (wfV_varI \<Theta> \<B> \<Gamma> b c y)
+ hence "y\<noteq>x" using v.fresh by auto
+ hence "Some (b, c) = lookup (\<Gamma>\<^sub>1@\<Gamma>\<^sub>2) y" using lookup_restrict wfV_varI by metis
+ then show ?case using wfV_varI wf_intros by metis
+next
+ case (wfV_litI \<Theta> \<Gamma> l)
+ then show ?case using e.fresh wf_intros by metis
+next
+ case (wfV_pairI \<Theta> \<B> \<Gamma> v1 b1 v2 b2)
+ show ?case proof
+ show "\<Theta>; \<B>; \<Gamma>\<^sub>1 @ \<Gamma>\<^sub>2 \<turnstile>\<^sub>w\<^sub>f v1 : b1" using wfV_pairI by auto
+ show "\<Theta>; \<B>; \<Gamma>\<^sub>1 @ \<Gamma>\<^sub>2 \<turnstile>\<^sub>w\<^sub>f v2 : b2" using wfV_pairI by auto
+ qed
+next
+ case (wfV_consI s dclist \<Theta> dc x b c \<B> \<Gamma> v)
+ show ?case proof
+ show "AF_typedef s dclist \<in> set \<Theta>" using wfV_consI by auto
+ show "(dc, \<lbrace> x : b | c \<rbrace>) \<in> set dclist" using wfV_consI by auto
+ show "\<Theta>; \<B>; \<Gamma>\<^sub>1 @ \<Gamma>\<^sub>2 \<turnstile>\<^sub>w\<^sub>f v : b" using wfV_consI by auto
+ qed
+next
+ case (wfV_conspI s bv dclist \<Theta> dc x b' c \<B> b \<Gamma> v)
+ show ?case proof
+ show "AF_typedef_poly s bv dclist \<in> set \<Theta>" using wfV_conspI by auto
+ show "(dc, \<lbrace> x : b' | c \<rbrace>) \<in> set dclist" using wfV_conspI by auto
+ show "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b" using wfV_conspI by auto
+ show " \<Theta>; \<B>; \<Gamma>\<^sub>1 @ \<Gamma>\<^sub>2 \<turnstile>\<^sub>w\<^sub>f v : b'[bv::=b]\<^sub>b\<^sub>b" using wfV_conspI by auto
+ show "atom bv \<sharp> (\<Theta>, \<B>, \<Gamma>\<^sub>1 @ \<Gamma>\<^sub>2, b, v)" unfolding fresh_prodN fresh_append_g using wfV_conspI fresh_prodN fresh_GCons fresh_append_g by metis
+ qed
+next
+ case (wfCE_valI \<Theta> \<B> \<Gamma> v b)
+ then show ?case using ce.fresh wf_intros by metis
+next
+ case (wfCE_plusI \<Theta> \<B> \<Gamma> v1 v2)
+ then show ?case using ce.fresh wf_intros by metis
+next
+ case (wfCE_leqI \<Theta> \<B> \<Gamma> v1 v2)
+ then show ?case using ce.fresh wf_intros by metis
+next
+ case (wfCE_eqI \<Theta> \<B> \<Gamma> v1 v2)
+ then show ?case using ce.fresh wf_intros by metis
+next
+ case (wfCE_fstI \<Theta> \<B> \<Gamma> v1 b1 b2)
+ then show ?case using ce.fresh wf_intros by metis
+next
+ case (wfCE_sndI \<Theta> \<B> \<Gamma> v1 b1 b2)
+ then show ?case using ce.fresh wf_intros by metis
+next
+ case (wfCE_concatI \<Theta> \<B> \<Gamma> v1 v2)
+ then show ?case using ce.fresh wf_intros by metis
+next
+ case (wfCE_lenI \<Theta> \<B> \<Gamma> v1)
+ then show ?case using ce.fresh wf_intros by metis
+next
+ case (wfTI z \<Theta> \<B> \<Gamma> b c)
+ hence "x \<noteq> z" using wfTI
+ fresh_GCons fresh_prodN fresh_PairD(1) fresh_gamma_append not_self_fresh by metis
+ show ?case proof
+ show \<open>atom z \<sharp> (\<Theta>, \<B>, \<Gamma>\<^sub>1 @ \<Gamma>\<^sub>2)\<close> using wfTI fresh_restrict[of z] using wfG_fresh_x wfX_wfY wfTI fresh_prodN by metis
+ show \<open> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b \<close> using wfTI by auto
+ have "\<Theta>; \<B>; ((z, b, TRUE) #\<^sub>\<Gamma> \<Gamma>\<^sub>1) @ \<Gamma>\<^sub>2 \<turnstile>\<^sub>w\<^sub>f c " proof(rule wfTI(5)[of "(z, b, TRUE) #\<^sub>\<Gamma> \<Gamma>\<^sub>1" ])
+ show \<open>(z, b, TRUE) #\<^sub>\<Gamma> \<Gamma> = ((z, b, TRUE) #\<^sub>\<Gamma> \<Gamma>\<^sub>1) @ (x, b', c') #\<^sub>\<Gamma> \<Gamma>\<^sub>2\<close> using wfTI by auto
+ show \<open>atom x \<sharp> c\<close> using wfTI \<tau>.fresh \<open>x \<noteq> z\<close> by auto
+ show \<open>atom x \<sharp> (z, b, TRUE) #\<^sub>\<Gamma> \<Gamma>\<^sub>1\<close> using wfTI \<open>x \<noteq> z\<close> fresh_GCons by simp
+ qed
+ thus \<open> \<Theta>; \<B>; (z, b, TRUE) #\<^sub>\<Gamma> \<Gamma>\<^sub>1 @ \<Gamma>\<^sub>2 \<turnstile>\<^sub>w\<^sub>f c \<close> by auto
+ qed
+next
+ case (wfC_eqI \<Theta> \<B> \<Gamma> e1 b e2)
+ show ?case proof
+ show "\<Theta>; \<B>; \<Gamma>\<^sub>1 @ \<Gamma>\<^sub>2 \<turnstile>\<^sub>w\<^sub>f e1 : b " using wfC_eqI c.fresh fresh_Nil by auto
+ show "\<Theta>; \<B>; \<Gamma>\<^sub>1 @ \<Gamma>\<^sub>2 \<turnstile>\<^sub>w\<^sub>f e2 : b " using wfC_eqI c.fresh fresh_Nil by auto
+ qed
+next
+ case (wfC_trueI \<Theta> \<Gamma>)
+ then show ?case using c.fresh wf_intros by metis
+next
+ case (wfC_falseI \<Theta> \<Gamma>)
+ then show ?case using c.fresh wf_intros by metis
+next
+ case (wfC_conjI \<Theta> \<Gamma> c1 c2)
+ then show ?case using c.fresh wf_intros by metis
+next
+ case (wfC_disjI \<Theta> \<Gamma> c1 c2)
+ then show ?case using c.fresh wf_intros by metis
+next
+case (wfC_notI \<Theta> \<Gamma> c1)
+ then show ?case using c.fresh wf_intros by metis
+next
+ case (wfC_impI \<Theta> \<Gamma> c1 c2)
+ then show ?case using c.fresh wf_intros by metis
+next
+ case (wfG_nilI \<Theta>)
+ then show ?case using wfV_varI wf_intros
+ by (meson GNil_append \<Gamma>.simps(3))
+next
+ case (wfG_cons1I c1 \<Theta> \<B> G x1 b1)
+ show ?case proof(cases "\<Gamma>\<^sub>1=GNil")
+ case True
+ then show ?thesis using wfG_cons1I wfG_consI by auto
+ next
+ case False
+ then obtain G'::\<Gamma> where *:"(x1, b1, c1) #\<^sub>\<Gamma> G' = \<Gamma>\<^sub>1" using GCons_eq_append_conv wfG_cons1I by auto
+ hence **:"G=G' @ (x, b', c') #\<^sub>\<Gamma> \<Gamma>\<^sub>2" using wfG_cons1I by auto
+
+ have " \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (x1, b1, c1) #\<^sub>\<Gamma> (G' @ \<Gamma>\<^sub>2)" proof(rule Wellformed.wfG_cons1I)
+ show \<open>c1 \<notin> {TRUE, FALSE}\<close> using wfG_cons1I by auto
+ show \<open>atom x1 \<sharp> G' @ \<Gamma>\<^sub>2\<close> using wfG_cons1I(4) ** fresh_restrict by metis
+ have " atom x \<sharp> G'" using wfG_cons1I * using fresh_GCons by blast
+ thus \<open> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f G' @ \<Gamma>\<^sub>2 \<close> using wfG_cons1I(3)[of G'] ** by auto
+ have "atom x \<sharp> c1 \<and> atom x \<sharp> (x1, b1, TRUE) #\<^sub>\<Gamma> G'" using fresh_GCons \<open>atom x \<sharp> \<Gamma>\<^sub>1\<close> * by auto
+ thus \<open> \<Theta>; \<B>; (x1, b1, TRUE) #\<^sub>\<Gamma> G' @ \<Gamma>\<^sub>2 \<turnstile>\<^sub>w\<^sub>f c1 \<close> using wfG_cons1I(6)[of "(x1, b1, TRUE) #\<^sub>\<Gamma> G'"] ** * wfG_cons1I by auto
+ show \<open> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b1 \<close> using wfG_cons1I by auto
+ qed
+ thus ?thesis using * by auto
+ qed
+next
+ case (wfG_cons2I c1 \<Theta> \<B> G x1 b1)
+ show ?case proof(cases "\<Gamma>\<^sub>1=GNil")
+ case True
+ then show ?thesis using wfG_cons2I wfG_consI by auto
+ next
+ case False
+ then obtain G'::\<Gamma> where *:"(x1, b1, c1) #\<^sub>\<Gamma> G' = \<Gamma>\<^sub>1" using GCons_eq_append_conv wfG_cons2I by auto
+ hence **:"G=G' @ (x, b', c') #\<^sub>\<Gamma> \<Gamma>\<^sub>2" using wfG_cons2I by auto
+
+ have " \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (x1, b1, c1) #\<^sub>\<Gamma> (G' @ \<Gamma>\<^sub>2)" proof(rule Wellformed.wfG_cons2I)
+ show \<open>c1 \<in> {TRUE, FALSE}\<close> using wfG_cons2I by auto
+ show \<open>atom x1 \<sharp> G' @ \<Gamma>\<^sub>2\<close> using wfG_cons2I ** fresh_restrict by metis
+ have " atom x \<sharp> G'" using wfG_cons2I * using fresh_GCons by blast
+ thus \<open> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f G' @ \<Gamma>\<^sub>2 \<close> using wfG_cons2I ** by auto
+ show \<open> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b1 \<close> using wfG_cons2I by auto
+ qed
+ thus ?thesis using * by auto
+ qed
+qed(auto)+
+
+lemma wf_restrict2:
+ fixes \<Gamma>::\<Gamma> and \<Gamma>'::\<Gamma> and v::v and e::e and c::c and \<tau>::\<tau> and ts::"(string*\<tau>) list" and \<Delta>::\<Delta> and s::s and b::b and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def
+ and cs::branch_s and css::branch_list
+ shows "\<Theta>; \<Phi>; \<B>; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f e : b \<Longrightarrow> \<Gamma>=\<Gamma>\<^sub>1@((x,b',c') #\<^sub>\<Gamma>\<Gamma>\<^sub>2) \<Longrightarrow> atom x \<sharp> e \<Longrightarrow> atom x \<sharp> \<Gamma>\<^sub>1 \<Longrightarrow> atom x \<sharp> \<Delta> \<Longrightarrow> \<Theta>; \<Phi>; \<B>; \<Gamma>\<^sub>1@\<Gamma>\<^sub>2 ; \<Delta> \<turnstile>\<^sub>w\<^sub>f e : b" and
+ "\<Theta>; \<Phi>; \<B>; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b \<Longrightarrow> True" and
+ "\<Theta>; \<Phi>; \<B>; \<Gamma> ; \<Delta> ; tid ; dc ; t \<turnstile>\<^sub>w\<^sub>f cs : b \<Longrightarrow> True" and
+ "\<Theta>; \<Phi>; \<B>; \<Gamma> ; \<Delta> ; tid ; dclist \<turnstile>\<^sub>w\<^sub>f css : b \<Longrightarrow> True" and
+ "\<Theta> \<turnstile>\<^sub>w\<^sub>f (\<Phi>::\<Phi>) \<Longrightarrow> True " and
+ "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<Longrightarrow> \<Gamma>=\<Gamma>\<^sub>1@((x,b',c') #\<^sub>\<Gamma>\<Gamma>\<^sub>2) \<Longrightarrow> atom x \<sharp> \<Gamma>\<^sub>1 \<Longrightarrow> atom x \<sharp> \<Delta> \<Longrightarrow> \<Theta>; \<B>; \<Gamma>\<^sub>1@\<Gamma>\<^sub>2 \<turnstile>\<^sub>w\<^sub>f \<Delta>" and
+ "\<Theta> ; \<Phi> \<turnstile>\<^sub>w\<^sub>f ftq \<Longrightarrow> True" and
+ "\<Theta> ; \<Phi> ; \<B> \<turnstile>\<^sub>w\<^sub>f ft \<Longrightarrow> True"
+
+proof(induct arbitrary: \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1
+ rule:wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.inducts)
+ case (wfE_valI \<Theta> \<Phi> \<Gamma> \<Delta> v b)
+ then show ?case using e.fresh wf_intros wf_restrict1 by metis
+next
+ case (wfE_plusI \<Theta> \<Phi> \<Gamma> \<Delta> v1 v2)
+ then show ?case using e.fresh wf_intros wf_restrict1 by metis
+next
+ case (wfE_leqI \<Theta> \<Phi> \<Gamma> \<Delta> v1 v2)
+ then show ?case using e.fresh wf_intros wf_restrict1 by metis
+next
+ case (wfE_eqI \<Theta> \<Phi> \<Gamma> \<Delta> v1 b v2)
+ then show ?case using e.fresh wf_intros wf_restrict1 by metis
+next
+ case (wfE_fstI \<Theta> \<Phi> \<Gamma> \<Delta> v1 b1 b2)
+ then show ?case using e.fresh wf_intros wf_restrict1 by metis
+next
+ case (wfE_sndI \<Theta> \<Phi> \<Gamma> \<Delta> v1 b1 b2)
+ then show ?case using e.fresh wf_intros wf_restrict1 by metis
+next
+ case (wfE_concatI \<Theta> \<Phi> \<Gamma> \<Delta> v1 v2)
+ then show ?case using e.fresh wf_intros wf_restrict1 by metis
+next
+ case (wfE_splitI \<Theta> \<Phi> \<Gamma> \<Delta> v1 v2)
+ then show ?case using e.fresh wf_intros wf_restrict1 by metis
+next
+ case (wfE_lenI \<Theta> \<Phi> \<Gamma> \<Delta> v1)
+ then show ?case using e.fresh wf_intros wf_restrict1 by metis
+next
+ case (wfE_appI \<Theta> \<Phi> \<Gamma> \<Delta> f x b c \<tau> s' v)
+ then show ?case using e.fresh wf_intros wf_restrict1 by metis
+next
+ case (wfE_appPI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> b' bv v \<tau> f x b c s)
+ show ?case proof
+ show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using wfE_appPI by auto
+ show \<open> \<Theta>; \<B>; \<Gamma>\<^sub>1 @ \<Gamma>\<^sub>2 \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wfE_appPI by auto
+ show \<open> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b' \<close> using wfE_appPI by auto
+
+ have "atom bv \<sharp> \<Gamma>\<^sub>1 @ \<Gamma>\<^sub>2" using wfE_appPI fresh_prodN fresh_restrict by metis
+ thus \<open>atom bv \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>\<^sub>1 @ \<Gamma>\<^sub>2, \<Delta>, b', v, (b_of \<tau>)[bv::=b']\<^sub>b)\<close>
+ using wfE_appPI fresh_prodN by auto
+
+ show \<open>Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \<tau> s))) = lookup_fun \<Phi> f\<close> using wfE_appPI by auto
+ show \<open> \<Theta>; \<B>; \<Gamma>\<^sub>1 @ \<Gamma>\<^sub>2 \<turnstile>\<^sub>w\<^sub>f v : b[bv::=b']\<^sub>b \<close> using wfE_appPI wf_restrict1 by auto
+ qed
+next
+ case (wfE_mvarI \<Theta> \<Phi> \<Gamma> \<Delta> u \<tau>)
+ then show ?case using e.fresh wf_intros by metis
+next
+ case (wfD_emptyI \<Theta> \<Gamma>)
+ then show ?case using c.fresh wf_intros wf_restrict1 by metis
+next
+ case (wfD_cons \<Theta> \<B> \<Gamma> \<Delta> \<tau> u)
+ show ?case proof
+ show "\<Theta>; \<B>; \<Gamma>\<^sub>1 @ \<Gamma>\<^sub>2 \<turnstile>\<^sub>w\<^sub>f \<Delta>" using wfD_cons fresh_DCons by metis
+ show "\<Theta>; \<B>; \<Gamma>\<^sub>1 @ \<Gamma>\<^sub>2 \<turnstile>\<^sub>w\<^sub>f \<tau> " using wfD_cons fresh_DCons fresh_Pair wf_restrict1 by metis
+ show "u \<notin> fst ` setD \<Delta>" using wfD_cons by auto
+ qed
+next
+ case (wfFTNone \<Theta> ft)
+ then show ?case by auto
+next
+ case (wfFTSome \<Theta> bv ft)
+ then show ?case by auto
+next
+ case (wfFTI \<Theta> B b \<Phi> x c s \<tau>)
+ then show ?case by auto
+qed(auto)+
+
+lemmas wf_restrict=wf_restrict1 wf_restrict2
+
+lemma wfT_restrict2:
+ fixes \<tau>::\<tau>
+ assumes "wfT \<Theta> \<B> ((x, b, c) #\<^sub>\<Gamma> \<Gamma>) \<tau>" and "atom x \<sharp> \<tau>"
+ shows "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau>"
+ using wf_restrict1(4)[of \<Theta> \<B> "((x, b, c) #\<^sub>\<Gamma> \<Gamma>)" \<tau> GNil x "b" "c" \<Gamma>] assms fresh_GNil append_g.simps by auto
+
+lemma wfG_intros2:
+ assumes "wfC P \<B> ((x,b,TRUE) #\<^sub>\<Gamma>\<Gamma>) c"
+ shows "wfG P \<B> ((x,b,c) #\<^sub>\<Gamma>\<Gamma>)"
+proof -
+ have "wfG P \<B> ((x,b,TRUE) #\<^sub>\<Gamma>\<Gamma>)" using wfC_wf assms by auto
+ hence *:"wfG P \<B> \<Gamma> \<and> atom x \<sharp> \<Gamma> \<and> wfB P \<B> b" using wfG_elims by metis
+ show ?thesis using assms proof(cases "c \<in> {TRUE,FALSE}")
+ case True
+ then show ?thesis using wfG_cons2I * by auto
+ next
+ case False
+ then show ?thesis using wfG_cons1I * assms by auto
+ qed
+qed
+
+section \<open>Type Definitions\<close>
+
+lemma wf_theta_weakening1:
+ fixes \<Gamma>::\<Gamma> and \<Gamma>'::\<Gamma> and v::v and e::e and c::c and \<tau>::\<tau> and ts::"(string*\<tau>) list" and \<Delta>::\<Delta> and s::s and b::b and \<B> :: \<B> and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def
+ and cs::branch_s and css::branch_list and t::\<tau>
+
+ shows "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b \<Longrightarrow> \<turnstile>\<^sub>w\<^sub>f \<Theta>' \<Longrightarrow> set \<Theta> \<subseteq> set \<Theta>' \<Longrightarrow> \<Theta>' ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b" and
+ "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f c \<Longrightarrow> \<turnstile>\<^sub>w\<^sub>f \<Theta>' \<Longrightarrow> set \<Theta> \<subseteq> set \<Theta>' \<Longrightarrow> \<Theta>' ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f c" and
+ "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<Longrightarrow> \<turnstile>\<^sub>w\<^sub>f \<Theta>' \<Longrightarrow> set \<Theta> \<subseteq> set \<Theta>' \<Longrightarrow> \<Theta>' ; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>" and
+ "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau> \<Longrightarrow> \<turnstile>\<^sub>w\<^sub>f \<Theta>' \<Longrightarrow> set \<Theta> \<subseteq> set \<Theta>' \<Longrightarrow> \<Theta>' ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau>" and
+ "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ts \<Longrightarrow> \<turnstile>\<^sub>w\<^sub>f \<Theta>' \<Longrightarrow> set \<Theta> \<subseteq> set \<Theta>' \<Longrightarrow> \<Theta>' ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ts" and
+ "\<turnstile>\<^sub>w\<^sub>f P \<Longrightarrow> True " and
+ "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b \<Longrightarrow> \<turnstile>\<^sub>w\<^sub>f \<Theta>' \<Longrightarrow> set \<Theta> \<subseteq> set \<Theta>' \<Longrightarrow> \<Theta>' ; \<B> \<turnstile>\<^sub>w\<^sub>f b" and
+ "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ce : b \<Longrightarrow> \<turnstile>\<^sub>w\<^sub>f \<Theta>' \<Longrightarrow> set \<Theta> \<subseteq> set \<Theta>' \<Longrightarrow> \<Theta>' ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ce : b" and
+ "\<Theta> \<turnstile>\<^sub>w\<^sub>f td \<Longrightarrow> \<turnstile>\<^sub>w\<^sub>f \<Theta>' \<Longrightarrow> set \<Theta> \<subseteq> set \<Theta>' \<Longrightarrow> \<Theta>' \<turnstile>\<^sub>w\<^sub>f td"
+proof(nominal_induct b and c and \<Gamma> and \<tau> and ts and P and b and b and td
+ avoiding: \<Theta>'
+ rule:wfV_wfC_wfG_wfT_wfTs_wfTh_wfB_wfCE_wfTD.strong_induct)
+ case (wfV_consI s dclist \<Theta> dc x b c \<B> \<Gamma> v)
+ show ?case proof
+ show \<open>AF_typedef s dclist \<in> set \<Theta>'\<close> using wfV_consI by auto
+ show \<open>(dc, \<lbrace> x : b | c \<rbrace>) \<in> set dclist\<close> using wfV_consI by auto
+ show \<open> \<Theta>' ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b \<close> using wfV_consI by auto
+ qed
+next
+ case (wfV_conspI s bv dclist \<Theta> dc x b' c \<B> b \<Gamma> v)
+ show ?case proof
+ show \<open>AF_typedef_poly s bv dclist \<in> set \<Theta>'\<close> using wfV_conspI by auto
+ show \<open>(dc, \<lbrace> x : b' | c \<rbrace>) \<in> set dclist\<close> using wfV_conspI by auto
+ show \<open>\<Theta>' ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b'[bv::=b]\<^sub>b\<^sub>b \<close> using wfV_conspI by auto
+ show "\<Theta>' ; \<B> \<turnstile>\<^sub>w\<^sub>f b " using wfV_conspI by auto
+ show "atom bv \<sharp> (\<Theta>', \<B>, \<Gamma>, b, v)" using wfV_conspI fresh_prodN by auto
+ qed
+next
+ case (wfTI z \<Theta> \<B> \<Gamma> b c)
+ thus ?case using Wellformed.wfTI by auto
+next
+ case (wfB_consI \<Theta> s dclist)
+ show ?case proof
+ show \<open> \<turnstile>\<^sub>w\<^sub>f \<Theta>' \<close> using wfB_consI by auto
+ show \<open>AF_typedef s dclist \<in> set \<Theta>'\<close> using wfB_consI by auto
+ qed
+next
+ case (wfB_appI \<Theta> \<B> b s bv dclist)
+ show ?case proof
+ show \<open> \<turnstile>\<^sub>w\<^sub>f \<Theta>' \<close> using wfB_appI by auto
+ show \<open>AF_typedef_poly s bv dclist \<in> set \<Theta>'\<close> using wfB_appI by auto
+ show "\<Theta>' ; \<B> \<turnstile>\<^sub>w\<^sub>f b" using wfB_appI by simp
+ qed
+qed(metis wf_intros)+
+
+lemma wf_theta_weakening2:
+ fixes \<Gamma>::\<Gamma> and \<Gamma>'::\<Gamma> and v::v and e::e and c::c and \<tau>::\<tau> and ts::"(string*\<tau>) list" and \<Delta>::\<Delta> and s::s and b::b and \<B> :: \<B> and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def
+ and cs::branch_s and css::branch_list and t::\<tau>
+ shows
+ "\<Theta>; \<Phi>; \<B>; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f e : b \<Longrightarrow> \<turnstile>\<^sub>w\<^sub>f \<Theta>' \<Longrightarrow> set \<Theta> \<subseteq> set \<Theta>' \<Longrightarrow> \<Theta>' ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f e : b" and
+ "\<Theta>; \<Phi>; \<B>; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b \<Longrightarrow> \<turnstile>\<^sub>w\<^sub>f \<Theta>' \<Longrightarrow> set \<Theta> \<subseteq> set \<Theta>' \<Longrightarrow> \<Theta>' ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b" and
+ "\<Theta>; \<Phi>; \<B>; \<Gamma> ; \<Delta> ; tid ; dc ; t \<turnstile>\<^sub>w\<^sub>f cs : b \<Longrightarrow> \<turnstile>\<^sub>w\<^sub>f \<Theta>' \<Longrightarrow> set \<Theta> \<subseteq> set \<Theta>' \<Longrightarrow> \<Theta>' ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> ; tid ; dc ; t \<turnstile>\<^sub>w\<^sub>f cs : b" and
+ "\<Theta>; \<Phi>; \<B>; \<Gamma> ; \<Delta> ; tid ; dclist \<turnstile>\<^sub>w\<^sub>f css : b \<Longrightarrow> \<turnstile>\<^sub>w\<^sub>f \<Theta>' \<Longrightarrow> set \<Theta> \<subseteq> set \<Theta>' \<Longrightarrow> \<Theta>' ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> ; tid ; dclist \<turnstile>\<^sub>w\<^sub>f css : b" and
+ "\<Theta> \<turnstile>\<^sub>w\<^sub>f (\<Phi>::\<Phi>) \<Longrightarrow> \<turnstile>\<^sub>w\<^sub>f \<Theta>' \<Longrightarrow> set \<Theta> \<subseteq> set \<Theta>' \<Longrightarrow> \<Theta>' \<turnstile>\<^sub>w\<^sub>f (\<Phi>::\<Phi>)" and
+ "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<Longrightarrow> \<turnstile>\<^sub>w\<^sub>f \<Theta>' \<Longrightarrow> set \<Theta> \<subseteq> set \<Theta>' \<Longrightarrow> \<Theta>' ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>" and
+ "\<Theta> ; \<Phi> \<turnstile>\<^sub>w\<^sub>f ftq \<Longrightarrow> \<turnstile>\<^sub>w\<^sub>f \<Theta>' \<Longrightarrow> set \<Theta> \<subseteq> set \<Theta>' \<Longrightarrow> \<Theta>' ; \<Phi> \<turnstile>\<^sub>w\<^sub>f ftq" and
+ "\<Theta> ; \<Phi> ; \<B> \<turnstile>\<^sub>w\<^sub>f ft \<Longrightarrow> \<turnstile>\<^sub>w\<^sub>f \<Theta>' \<Longrightarrow> set \<Theta> \<subseteq> set \<Theta>' \<Longrightarrow> \<Theta>' ; \<Phi> ; \<B> \<turnstile>\<^sub>w\<^sub>f ft"
+
+proof(nominal_induct b and b and b and b and \<Phi> and \<Delta> and ftq and ft
+ avoiding: \<Theta>'
+rule:wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.strong_induct)
+ case (wfE_appPI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> b' bv v \<tau> f x b c s)
+ show ?case proof
+ show \<open> \<Theta>' \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using wfE_appPI by auto
+ show \<open> \<Theta>' ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wfE_appPI by auto
+ show \<open> \<Theta>' ; \<B> \<turnstile>\<^sub>w\<^sub>f b' \<close> using wfE_appPI wf_theta_weakening1 by auto
+ show \<open>atom bv \<sharp> (\<Phi>, \<Theta>', \<B>, \<Gamma>, \<Delta>, b', v, (b_of \<tau>)[bv::=b']\<^sub>b)\<close> using wfE_appPI by auto
+ show \<open>Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \<tau> s))) = lookup_fun \<Phi> f\<close> using wfE_appPI by auto
+ show \<open> \<Theta>' ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b[bv::=b']\<^sub>b \<close> using wfE_appPI wf_theta_weakening1 by auto
+ qed
+next
+ case (wfS_matchI \<Theta> \<B> \<Gamma> v tid dclist \<Delta> \<Phi> cs b)
+ show ?case proof
+ show \<open> \<Theta>' ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : B_id tid \<close> using wfS_matchI wf_theta_weakening1 by auto
+ show \<open>AF_typedef tid dclist \<in> set \<Theta>'\<close> using wfS_matchI by auto
+ show \<open> \<Theta>' ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wfS_matchI by auto
+ show \<open> \<Theta>' \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using wfS_matchI by auto
+ show \<open>\<Theta>'; \<Phi>; \<B>; \<Gamma>; \<Delta>; tid; dclist \<turnstile>\<^sub>w\<^sub>f cs : b \<close> using wfS_matchI by auto
+ qed
+next
+ case (wfS_varI \<Theta> \<B> \<Gamma> \<tau> v u \<Phi> \<Delta> b s)
+ show ?case proof
+ show \<open> \<Theta>' ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau> \<close> using wfS_varI wf_theta_weakening1 by auto
+ show \<open> \<Theta>' ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b_of \<tau> \<close> using wfS_varI wf_theta_weakening1 by auto
+ show \<open>atom u \<sharp> (\<Phi>, \<Theta>', \<B>, \<Gamma>, \<Delta>, \<tau>, v, b)\<close> using wfS_varI by auto
+ show \<open> \<Theta>' ; \<Phi> ; \<B> ; \<Gamma> ; (u, \<tau>) #\<^sub>\<Delta> \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b \<close> using wfS_varI by auto
+ qed
+next
+ case (wfS_letI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> e b' x s b)
+ show ?case proof
+ show \<open> \<Theta>' ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f e : b' \<close> using wfS_letI by auto
+ show \<open> \<Theta>' ; \<Phi> ; \<B> ; (x, b', TRUE) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b \<close> using wfS_letI by auto
+ show \<open> \<Theta>' ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wfS_letI by auto
+ show \<open>atom x \<sharp> (\<Phi>, \<Theta>', \<B>, \<Gamma>, \<Delta>, e, b)\<close> using wfS_letI by auto
+ qed
+next
+ case (wfS_let2I \<Theta> \<Phi> \<B> \<Gamma> \<Delta> s1 \<tau> x s2 b)
+ show ?case proof
+ show \<open> \<Theta>' ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s1 : b_of \<tau> \<close> using wfS_let2I by auto
+ show \<open> \<Theta>' ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau> \<close> using wfS_let2I wf_theta_weakening1 by auto
+ show \<open> \<Theta>' ; \<Phi> ; \<B> ; (x, b_of \<tau>, TRUE) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s2 : b \<close> using wfS_let2I by auto
+ show \<open>atom x \<sharp> (\<Phi>, \<Theta>', \<B>, \<Gamma>, \<Delta>, s1, b, \<tau>)\<close> using wfS_let2I by auto
+ qed
+next
+ case (wfS_branchI \<Theta> \<Phi> \<B> x \<tau> \<Gamma> \<Delta> s b tid dc)
+ show ?case proof
+ show \<open> \<Theta>' ; \<Phi> ; \<B> ; (x, b_of \<tau>, TRUE) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b \<close> using wfS_branchI by auto
+ show \<open>atom x \<sharp> (\<Phi>, \<Theta>', \<B>, \<Gamma>, \<Delta>, \<Gamma>, \<tau>)\<close> using wfS_branchI by auto
+ show \<open> \<Theta>' ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wfS_branchI by auto
+ qed
+next
+ case (wfPhi_consI f \<Phi> \<Theta> ft)
+ show ?case proof
+ show "f \<notin> name_of_fun ` set \<Phi>" using wfPhi_consI by auto
+ show "\<Theta>' ; \<Phi> \<turnstile>\<^sub>w\<^sub>f ft" using wfPhi_consI by auto
+ show "\<Theta>' \<turnstile>\<^sub>w\<^sub>f \<Phi>" using wfPhi_consI by auto
+ qed
+next
+ case (wfFTNone \<Theta> ft)
+ then show ?case using wf_intros by metis
+next
+ case (wfFTSome \<Theta> bv ft)
+ then show ?case using wf_intros by metis
+next
+ case (wfFTI \<Theta> B b \<Phi> x c s \<tau>)
+ thus ?case using Wellformed.wfFTI wf_theta_weakening1 by simp
+next
+ case (wfS_assertI \<Theta> \<Phi> \<B> x c \<Gamma> \<Delta> s b)
+ show ?case proof
+ show \<open> \<Theta>' ; \<Phi> ; \<B> ; (x, B_bool, c) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b \<close> using wfS_assertI wf_theta_weakening1 by auto
+ show \<open> \<Theta>' ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f c \<close> using wfS_assertI wf_theta_weakening1 by auto
+ show \<open> \<Theta>' ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wfS_assertI wf_theta_weakening1 by auto
+ have "atom x \<sharp> \<Theta>'" using wf_supp(6)[OF \<open>\<turnstile>\<^sub>w\<^sub>f \<Theta>' \<close>] fresh_def by auto
+ thus \<open>atom x \<sharp> (\<Phi>, \<Theta>', \<B>, \<Gamma>, \<Delta>, c, b, s)\<close> using wfS_assertI fresh_prodN fresh_def by simp
+ qed
+qed(metis wf_intros wf_theta_weakening1 )+
+
+lemmas wf_theta_weakening = wf_theta_weakening1 wf_theta_weakening2
+
+lemma lookup_wfTD:
+ fixes td::type_def
+ assumes "td \<in> set \<Theta>" and "\<turnstile>\<^sub>w\<^sub>f \<Theta>"
+ shows "\<Theta> \<turnstile>\<^sub>w\<^sub>f td"
+ using assms proof(induct \<Theta> )
+ case Nil
+ then show ?case by auto
+next
+ case (Cons td' \<Theta>')
+ then consider "td = td'" | "td \<in> set \<Theta>'" by auto
+ then have "\<Theta>' \<turnstile>\<^sub>w\<^sub>f td" proof(cases)
+ case 1
+ then show ?thesis using Cons using wfTh_elims by auto
+ next
+ case 2
+ then show ?thesis using Cons using wfTh_elims by auto
+ qed
+ then show ?case using wf_theta_weakening Cons by (meson set_subset_Cons)
+qed
+
+subsection \<open>Simple\<close>
+
+lemma wfTh_dclist_unique:
+ assumes "wfTh \<Theta>" and "AF_typedef tid dclist1 \<in> set \<Theta>" and "AF_typedef tid dclist2 \<in> set \<Theta>"
+ shows "dclist1 = dclist2"
+using assms proof(induct \<Theta> rule: \<Theta>_induct)
+ case TNil
+ then show ?case by auto
+next
+ case (AF_typedef tid' dclist' \<Theta>')
+ then show ?case using wfTh_elims
+ by (metis image_eqI name_of_type.simps(1) set_ConsD type_def.eq_iff(1))
+next
+ case (AF_typedef_poly tid bv dclist \<Theta>')
+ then show ?case using wfTh_elims by auto
+qed
+
+lemma wfTs_ctor_unique:
+ fixes dclist::"(string*\<tau>) list"
+ assumes "\<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f dclist" and "(c, t1) \<in> set dclist" and "(c,t2) \<in> set dclist"
+ shows "t1 = t2"
+ using assms proof(induct dclist rule: list.inducts)
+ case Nil
+ then show ?case by auto
+next
+ case (Cons x1 x2)
+ consider "x1 = (c,t1)" | "x1 = (c,t2)" | "x1 \<noteq> (c,t1) \<and> x1 \<noteq> (c,t2)" by auto
+ thus ?case proof(cases)
+ case 1
+ then show ?thesis using Cons wfTs_elims set_ConsD
+ by (metis fst_conv image_eqI prod.inject)
+ next
+ case 2
+ then show ?thesis using Cons wfTs_elims set_ConsD
+ by (metis fst_conv image_eqI prod.inject)
+ next
+ case 3
+ then show ?thesis using Cons wfTs_elims by (metis set_ConsD)
+ qed
+qed
+
+lemma wfTD_ctor_unique:
+ assumes "\<Theta> \<turnstile>\<^sub>w\<^sub>f (AF_typedef tid dclist)" and "(c, t1) \<in> set dclist" and "(c,t2) \<in> set dclist"
+ shows "t1 = t2"
+ using wfTD_elims wfTs_elims assms wfTs_ctor_unique by metis
+
+lemma wfTh_ctor_unique:
+ assumes "wfTh \<Theta>" and "AF_typedef tid dclist \<in> set \<Theta>" and "(c, t1) \<in> set dclist" and "(c,t2) \<in> set dclist"
+ shows "t1 = t2"
+ using lookup_wfTD wfTD_ctor_unique assms by metis
+
+lemma wfTs_supp_t:
+ fixes dclist::"(string*\<tau>) list"
+ assumes "(c,t) \<in> set dclist" and "\<Theta> ; B ; GNil \<turnstile>\<^sub>w\<^sub>f dclist"
+ shows "supp t \<subseteq> supp B"
+using assms proof(induct dclist arbitrary: c t rule:list.induct)
+ case Nil
+ then show ?case by auto
+next
+ case (Cons ct dclist')
+ then consider "ct = (c,t)" | "(c,t) \<in> set dclist'" by auto
+ then show ?case proof(cases)
+ case 1
+ then have "\<Theta> ; B ; GNil \<turnstile>\<^sub>w\<^sub>f t" using Cons wfTs_elims by blast
+ thus ?thesis using wfT_supp atom_dom.simps by force
+ next
+ case 2
+ then show ?thesis using Cons wfTs_elims by metis
+ qed
+qed
+
+lemma wfTh_lookup_supp_empty:
+ fixes t::\<tau>
+ assumes "AF_typedef tid dclist \<in> set \<Theta>" and "(c,t) \<in> set dclist" and "\<turnstile>\<^sub>w\<^sub>f \<Theta>"
+ shows "supp t = {}"
+proof -
+ have "\<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f dclist" using assms lookup_wfTD wfTD_elims by metis
+ thus ?thesis using wfTs_supp_t assms by force
+qed
+
+lemma wfTh_supp_b:
+ assumes "AF_typedef tid dclist \<in> set \<Theta>" and "(dc,\<lbrace> z : b | c \<rbrace> ) \<in> set dclist" and "\<turnstile>\<^sub>w\<^sub>f \<Theta>"
+ shows "supp b = {}"
+ using assms wfTh_lookup_supp_empty \<tau>.supp by blast
+
+lemma wfTh_b_eq_iff:
+ fixes bva1::bv and bva2::bv and dc::string
+ assumes "(dc, \<lbrace> x1 : b1 | c1 \<rbrace>) \<in> set dclist1" and "(dc, \<lbrace> x2 : b2 | c2 \<rbrace>) \<in> set dclist2" and
+ "wfTs P {|bva1|} GNil dclist1" and "wfTs P {|bva2|} GNil dclist2"
+ "[[atom bva1]]lst.dclist1 = [[atom bva2]]lst.dclist2"
+ shows "[[atom bva1]]lst. (dc,\<lbrace> x1 : b1 | c1 \<rbrace>) = [[atom bva2]]lst. (dc,\<lbrace> x2 : b2 | c2 \<rbrace>)"
+using assms proof(induct dclist1 arbitrary: dclist2)
+ case Nil
+ then show ?case by auto
+next
+ case (Cons dct1' dclist1')
+ show ?case proof(cases "dclist2 = []")
+ case True
+ then show ?thesis using Cons by auto
+ next
+ case False
+ then obtain dct2' and dclist2' where cons:"dct2' # dclist2' = dclist2" using list.exhaust by metis
+ hence *:"[[atom bva1]]lst. dclist1' = [[atom bva2]]lst. dclist2' \<and> [[atom bva1]]lst. dct1' = [[atom bva2]]lst. dct2'"
+ using Cons lst_head_cons Cons cons by metis
+ hence **: "fst dct1' = fst dct2'" using lst_fst[THEN lst_pure]
+ by (metis (no_types) \<open>[[atom bva1]]lst. dclist1' = [[atom bva2]]lst. dclist2' \<and> [[atom bva1]]lst. dct1' = [[atom bva2]]lst. dct2'\<close>
+ \<open>\<And>x2 x1 t2' t2a t2 t1. [[atom x1]]lst. (t1, t2a) = [[atom x2]]lst. (t2, t2') \<Longrightarrow> t1 = t2\<close> fst_conv surj_pair)
+ show ?thesis proof(cases "fst dct1' = dc")
+ case True
+ have "dc \<notin> fst ` set dclist1'" using wfTs_elims Cons by (metis True fstI)
+ hence 1:"(dc, \<lbrace> x1 : b1 | c1 \<rbrace>) = dct1'" using Cons by (metis fstI image_iff set_ConsD)
+ have "dc \<notin> fst ` set dclist2'" using wfTs_elims Cons cons
+ by (metis "**" True fstI)
+ hence 2:"(dc, \<lbrace> x2 : b2 | c2 \<rbrace>) = dct2' " using Cons cons by (metis fst_conv image_eqI set_ConsD)
+ then show ?thesis using Cons * 1 2 by blast
+ next
+ case False
+ hence "fst dct2' \<noteq> dc" using ** by auto
+ hence "(dc, \<lbrace> x1 : b1 | c1 \<rbrace>) \<in> set dclist1' \<and> (dc, \<lbrace> x2 : b2 | c2 \<rbrace>) \<in> set dclist2' " using Cons cons False
+ by (metis fstI set_ConsD)
+ moreover have "[[atom bva1]]lst. dclist1' = [[atom bva2]]lst. dclist2'" using * False by metis
+ ultimately show ?thesis using Cons ** *
+ using cons wfTs_elims(2) by blast
+ qed
+ qed
+qed
+
+
+subsection \<open>Polymorphic\<close>
+
+lemma wfTh_wfTs_poly:
+ fixes dclist::"(string * \<tau>) list"
+ assumes "AF_typedef_poly tyid bva dclist \<in> set P" and "\<turnstile>\<^sub>w\<^sub>f P"
+ shows "P ; {|bva|} ; GNil \<turnstile>\<^sub>w\<^sub>f dclist"
+proof -
+ have *:"P \<turnstile>\<^sub>w\<^sub>f AF_typedef_poly tyid bva dclist" using lookup_wfTD assms by simp
+
+ obtain bv lst where *:"P ; {|bv|} ; GNil \<turnstile>\<^sub>w\<^sub>f lst \<and>
+ (\<forall>c. atom c \<sharp> (dclist, lst) \<longrightarrow> atom c \<sharp> (bva, bv, dclist, lst) \<longrightarrow> (bva \<leftrightarrow> c) \<bullet> dclist = (bv \<leftrightarrow> c) \<bullet> lst)"
+ using wfTD_elims(2)[OF *] by metis
+
+ obtain c::bv where **:"atom c \<sharp> ((dclist, lst),(bva, bv, dclist, lst))" using obtain_fresh by metis
+ have "P ; {|bv|} ; GNil \<turnstile>\<^sub>w\<^sub>f lst" using * by metis
+ hence "wfTs ((bv \<leftrightarrow> c) \<bullet> P) ((bv \<leftrightarrow> c) \<bullet> {|bv|}) ((bv \<leftrightarrow> c) \<bullet> GNil) ((bv \<leftrightarrow> c) \<bullet> lst)" using ** wfTs.eqvt by metis
+ hence "wfTs P {|c|} GNil ((bva \<leftrightarrow> c) \<bullet> dclist)" using * theta_flip_eq fresh_GNil assms
+ proof -
+ have "\<forall>b ba. (ba::bv \<leftrightarrow> b) \<bullet> P = P" by (metis \<open>\<turnstile>\<^sub>w\<^sub>f P\<close> theta_flip_eq)
+ then show ?thesis
+ using "*" "**" \<open>(bv \<leftrightarrow> c) \<bullet> P ; (bv \<leftrightarrow> c) \<bullet> {|bv|} ; (bv \<leftrightarrow> c) \<bullet> GNil \<turnstile>\<^sub>w\<^sub>f (bv \<leftrightarrow> c) \<bullet> lst\<close> by fastforce
+ qed
+ hence "wfTs ((bva \<leftrightarrow> c) \<bullet> P) ((bva \<leftrightarrow> c) \<bullet> {|bva|}) ((bva \<leftrightarrow> c) \<bullet> GNil) ((bva \<leftrightarrow> c) \<bullet> dclist)"
+ using wfTs.eqvt fresh_GNil
+ by (simp add: assms(2) theta_flip_eq2)
+
+ thus ?thesis using wfTs.eqvt permute_flip_cancel by metis
+qed
+
+lemma wfTh_dclist_poly_unique:
+ assumes "wfTh \<Theta>" and "AF_typedef_poly tid bva dclist1 \<in> set \<Theta>" and "AF_typedef_poly tid bva2 dclist2 \<in> set \<Theta>"
+ shows "[[atom bva]]lst. dclist1 = [[atom bva2]]lst.dclist2"
+using assms proof(induct \<Theta> rule: \<Theta>_induct)
+ case TNil
+ then show ?case by auto
+next
+ case (AF_typedef tid' dclist' \<Theta>')
+ then show ?case using wfTh_elims by auto
+next
+ case (AF_typedef_poly tid bv dclist \<Theta>')
+ then show ?case using wfTh_elims image_eqI name_of_type.simps set_ConsD type_def.eq_iff
+ by (metis Abs1_eq(3))
+qed
+
+lemma wfTh_poly_lookup_supp:
+ fixes t::\<tau>
+ assumes "AF_typedef_poly tid bv dclist \<in> set \<Theta>" and "(c,t) \<in> set dclist" and "\<turnstile>\<^sub>w\<^sub>f \<Theta>"
+ shows "supp t \<subseteq> {atom bv}"
+proof -
+ have "supp dclist \<subseteq> {atom bv}" using assms lookup_wfTD wf_supp1 type_def.supp
+ by (metis Diff_single_insert Un_subset_iff list.simps(15) supp_Nil supp_of_atom_list)
+ then show ?thesis using assms(2) proof(induct dclist)
+ case Nil
+ then show ?case by auto
+ next
+ case (Cons a dclist)
+ then show ?case using supp_Pair supp_Cons
+ by (metis (mono_tags, opaque_lifting) Un_empty_left Un_empty_right pure_supp subset_Un_eq subset_singletonD supp_list_member)
+ qed
+qed
+
+lemma wfTh_poly_supp_b:
+ assumes "AF_typedef_poly tid bv dclist \<in> set \<Theta>" and "(dc,\<lbrace> z : b | c \<rbrace> ) \<in> set dclist" and "\<turnstile>\<^sub>w\<^sub>f \<Theta>"
+ shows "supp b \<subseteq> {atom bv}"
+ using assms wfTh_poly_lookup_supp \<tau>.supp by force
+
+lemma subst_g_inside:
+ fixes x::x and c::c and \<Gamma>::\<Gamma> and \<Gamma>'::\<Gamma>
+ assumes "wfG P \<B> (\<Gamma>' @ (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)"
+ shows "(\<Gamma>' @ (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)[x::=v]\<^sub>\<Gamma>\<^sub>v = (\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v@\<Gamma>)"
+using assms proof(induct \<Gamma>' rule: \<Gamma>_induct)
+ case GNil
+ then show ?case using subst_gb.simps by simp
+next
+ case (GCons x' b' c' G)
+ hence wfg:"wfG P \<B> (G @ (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>) \<and> atom x' \<sharp> (G @ (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>)" using wfG_elims(2)
+ using GCons.prems append_g.simps by metis
+ hence "atom x \<notin> atom_dom ((x', b', c') #\<^sub>\<Gamma> G)" using GCons wfG_inside_fresh by fast
+ hence "x\<noteq>x'"
+ using GCons append_Cons wfG_inside_fresh atom_dom.simps toSet.simps by simp
+ hence "((GCons (x', b', c') G) @ (GCons (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) \<Gamma>))[x::=v]\<^sub>\<Gamma>\<^sub>v =
+ (GCons (x', b', c') (G @ (GCons (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) \<Gamma>)))[x::=v]\<^sub>\<Gamma>\<^sub>v" by auto
+ also have "... = GCons (x', b', c'[x::=v]\<^sub>c\<^sub>v) ((G @ (GCons (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) \<Gamma>))[x::=v]\<^sub>\<Gamma>\<^sub>v)"
+ using subst_gv.simps \<open>x\<noteq>x'\<close> by simp
+ also have "... = (x', b', c'[x::=v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> (G[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>)" using GCons wfg by blast
+ also have "... = ((x', b', c') #\<^sub>\<Gamma> G)[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>" using subst_gv.simps \<open>x\<noteq>x'\<close> by simp
+ finally show ?case by auto
+qed
+
+lemma wfTh_td_eq:
+ assumes "td1 \<in> set (td2 # P)" and "wfTh (td2 # P)" and "name_of_type td1 = name_of_type td2"
+ shows "td1 = td2"
+proof(rule ccontr)
+ assume as: "td1 \<noteq> td2"
+ have "name_of_type td2 \<notin> name_of_type ` set P" using wfTh_elims(2)[OF assms(2)] by metis
+ moreover have "td1 \<in> set P" using assms as by simp
+ ultimately have "name_of_type td1 \<noteq> name_of_type td2"
+ by (metis rev_image_eqI)
+ thus False using assms by auto
+qed
+
+lemma wfTh_td_unique:
+ assumes "td1 \<in> set P" and "td2 \<in> set P" and "wfTh P" and "name_of_type td1 = name_of_type td2"
+ shows "td1 = td2"
+using assms proof(induct P rule: list.induct)
+ case Nil
+ then show ?case by auto
+next
+ case (Cons td \<Theta>')
+ consider "td = td1" | "td = td2" | "td \<noteq> td1 \<and> td \<noteq> td2" by auto
+ then show ?case proof(cases)
+ case 1
+ then show ?thesis using Cons wfTh_elims wfTh_td_eq by metis
+ next
+ case 2
+ then show ?thesis using Cons wfTh_elims wfTh_td_eq by metis
+ next
+ case 3
+ then show ?thesis using Cons wfTh_elims by auto
+ qed
+qed
+
+lemma wfTs_distinct:
+ fixes dclist::"(string * \<tau>) list"
+ assumes "\<Theta> ; B ; GNil \<turnstile>\<^sub>w\<^sub>f dclist"
+ shows "distinct (map fst dclist)"
+using assms proof(induct dclist rule: list.induct)
+ case Nil
+ then show ?case by auto
+next
+ case (Cons x1 x2)
+ then show ?case
+ by (metis Cons.hyps Cons.prems distinct.simps(2) fst_conv list.set_map list.simps(9) wfTs_elims(2))
+qed
+
+lemma wfTh_dclist_distinct:
+ assumes "AF_typedef s dclist \<in> set P" and "wfTh P"
+ shows "distinct (map fst dclist)"
+proof -
+ have "wfTD P (AF_typedef s dclist)" using assms lookup_wfTD by auto
+ hence "wfTs P {||} GNil dclist" using wfTD_elims by metis
+ thus ?thesis using wfTs_distinct by metis
+qed
+
+lemma wfTh_dc_t_unique2:
+ assumes "AF_typedef s dclist' \<in> set P" and "(dc,tc' ) \<in> set dclist'" and "AF_typedef s dclist \<in> set P" and "wfTh P" and
+ "(dc, tc) \<in> set dclist"
+ shows "tc= tc'"
+proof -
+ have "dclist = dclist'" using assms wfTh_td_unique name_of_type.simps by force
+ moreover have "distinct (map fst dclist)" using wfTh_dclist_distinct assms by auto
+ ultimately show ?thesis using assms
+ by (meson eq_key_imp_eq_value)
+qed
+
+lemma wfTh_dc_t_unique:
+ assumes "AF_typedef s dclist' \<in> set P" and "(dc, \<lbrace> x' : b' | c' \<rbrace> ) \<in> set dclist'" and "AF_typedef s dclist \<in> set P" and "wfTh P" and
+ "(dc, \<lbrace> x : b | c \<rbrace>) \<in> set dclist"
+ shows "\<lbrace> x' : b' | c' \<rbrace>= \<lbrace> x : b | c \<rbrace>"
+ using assms wfTh_dc_t_unique2 by metis
+
+lemma wfTs_wfT:
+ fixes dclist::"(string *\<tau>) list" and t::\<tau>
+ assumes "\<Theta>; \<B>; GNil \<turnstile>\<^sub>w\<^sub>f dclist" and "(dc,t) \<in> set dclist"
+ shows "\<Theta>; \<B>; GNil \<turnstile>\<^sub>w\<^sub>f t"
+using assms proof(induct dclist rule:list.induct)
+ case Nil
+ then show ?case by auto
+next
+ case (Cons x1 x2)
+ thus ?case using wfTs_elims(2)[OF Cons(2)] by auto
+qed
+
+lemma wfTh_wfT:
+ fixes t::\<tau>
+ assumes "wfTh P" and "AF_typedef tid dclist \<in> set P" and "(dc,t) \<in> set dclist"
+ shows "P ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f t"
+proof -
+ have "P \<turnstile>\<^sub>w\<^sub>f AF_typedef tid dclist" using lookup_wfTD assms by auto
+ hence "P ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f dclist" using wfTD_elims by auto
+ thus ?thesis using wfTs_wfT assms by auto
+qed
+
+lemma td_lookup_eq_iff:
+ fixes dc :: string and bva1::bv and bva2::bv
+ assumes "[[atom bva1]]lst. dclist1 = [[atom bva2]]lst. dclist2" and "(dc, \<lbrace> x : b | c \<rbrace>) \<in> set dclist1"
+ shows "\<exists>x2 b2 c2. (dc, \<lbrace> x2 : b2 | c2 \<rbrace>) \<in> set dclist2"
+using assms proof(induct dclist1 arbitrary: dclist2)
+ case Nil
+ then show ?case by auto
+next
+ case (Cons dct1' dclist1')
+ then obtain dct2' and dclist2' where cons:"dct2' # dclist2' = dclist2" using lst_head_cons_neq_nil[OF Cons(2)] list.exhaust by metis
+ hence *:"[[atom bva1]]lst. dclist1' = [[atom bva2]]lst. dclist2' \<and> [[atom bva1]]lst. dct1' = [[atom bva2]]lst. dct2'"
+ using Cons lst_head_cons Cons cons by metis
+ show ?case proof(cases "dc=fst dct1'")
+ case True
+ hence "dc = fst dct2'" using * lst_fst[ THEN lst_pure ]
+ proof -
+ show ?thesis
+ by (metis (no_types) "local.*" True \<open>\<And>x2 x1 t2' t2a t2 t1. [[atom x1]]lst. (t1, t2a) = [[atom x2]]lst. (t2, t2') \<Longrightarrow> t1 = t2\<close> prod.exhaust_sel) (* 31 ms *)
+ qed
+ obtain x2 b2 and c2 where "snd dct2' = \<lbrace> x2 : b2 | c2 \<rbrace>" using obtain_fresh_z by metis
+ hence "(dc, \<lbrace> x2 : b2 | c2 \<rbrace>) = dct2'" using \<open>dc = fst dct2'\<close>
+ by (metis prod.exhaust_sel)
+ then show ?thesis using cons by force
+ next
+ case False
+ hence "(dc, \<lbrace> x : b | c \<rbrace>) \<in> set dclist1'" using Cons by auto
+ then show ?thesis using Cons
+ by (metis "local.*" cons list.set_intros(2))
+ qed
+qed
+
+lemma lst_t_b_eq_iff:
+ fixes bva1::bv and bva2::bv
+ assumes "[[atom bva1]]lst. \<lbrace> x1 : b1 | c1 \<rbrace> = [[atom bva2]]lst. \<lbrace> x2 : b2 | c2 \<rbrace>"
+ shows "[[atom bva1]]lst. b1 = [[atom bva2]]lst.b2"
+proof(subst Abs1_eq_iff_all(3)[of bva1 b1 bva2 b2],rule,rule,rule)
+ fix c::bv
+ assume "atom c \<sharp> ( \<lbrace> x1 : b1 | c1 \<rbrace> , \<lbrace> x2 : b2 | c2 \<rbrace>)" and "atom c \<sharp> (bva1, bva2, b1, b2)"
+
+ show "(bva1 \<leftrightarrow> c) \<bullet> b1 = (bva2 \<leftrightarrow> c) \<bullet> b2" using assms Abs1_eq_iff(3) assms
+ by (metis Abs1_eq_iff_fresh(3) \<open>atom c \<sharp> (bva1, bva2, b1, b2)\<close> \<tau>.fresh \<tau>.perm_simps type_eq_subst_eq2(2))
+qed
+
+lemma wfTh_typedef_poly_b_eq_iff:
+ assumes "AF_typedef_poly tyid bva1 dclist1 \<in> set P" and "(dc, \<lbrace> x1 : b1 | c1 \<rbrace>) \<in> set dclist1"
+ and "AF_typedef_poly tyid bva2 dclist2 \<in> set P" and "(dc, \<lbrace> x2 : b2 | c2 \<rbrace>) \<in> set dclist2" and "\<turnstile>\<^sub>w\<^sub>f P"
+shows "b1[bva1::=b]\<^sub>b\<^sub>b = b2[bva2::=b]\<^sub>b\<^sub>b"
+proof -
+ have "[[atom bva1]]lst. dclist1 = [[atom bva2]]lst.dclist2" using assms wfTh_dclist_poly_unique by metis
+ hence "[[atom bva1]]lst. (dc,\<lbrace> x1 : b1 | c1 \<rbrace>) = [[atom bva2]]lst. (dc,\<lbrace> x2 : b2 | c2 \<rbrace>)" using wfTh_b_eq_iff assms wfTh_wfTs_poly by metis
+ hence "[[atom bva1]]lst. \<lbrace> x1 : b1 | c1 \<rbrace> = [[atom bva2]]lst. \<lbrace> x2 : b2 | c2 \<rbrace>" using lst_snd by metis
+ hence "[[atom bva1]]lst. b1 = [[atom bva2]]lst.b2" using lst_t_b_eq_iff by metis
+ thus ?thesis using subst_b_flip_eq_two subst_b_b_def by metis
+qed
+
+section \<open>Equivariance Lemmas\<close>
+
+lemma x_not_in_u_set[simp]:
+ fixes x::x and us::"u fset"
+ shows "atom x \<notin> supp us"
+ by(induct us,auto, simp add: supp_finsert supp_at_base)
+
+lemma wfS_flip_eq:
+ fixes s1::s and x1::x and s2::s and x2::x and \<Delta>::\<Delta>
+ assumes "[[atom x1]]lst. s1 = [[atom x2]]lst. s2" and "[[atom x1]]lst. t1 = [[atom x2]]lst. t2" and "[[atom x1]]lst. c1 = [[atom x2]]lst. c2" and "atom x2 \<sharp> \<Gamma>" and
+ " \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>" and
+ "\<Theta> ; \<Phi> ; \<B> ; (x1, b, c1) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s1 : b_of t1"
+ shows "\<Theta> ; \<Phi> ; \<B> ; (x2, b, c2) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s2 : b_of t2"
+proof(cases "x1=x2")
+ case True
+ hence "s1 = s2 \<and> t1 = t2 \<and> c1 = c2" using assms Abs1_eq_iff by metis
+ then show ?thesis using assms True by simp
+next
+ case False
+ have "\<turnstile>\<^sub>w\<^sub>f \<Theta> \<and> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<and> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>" using wfX_wfY assms by metis
+ moreover have "atom x1 \<sharp> \<Gamma>" using wfX_wfY wfG_elims assms by metis
+ moreover hence "atom x1 \<sharp> \<Delta> \<and> atom x2 \<sharp> \<Delta> " using wfD_x_fresh assms by auto
+ ultimately have " \<Theta> ; \<Phi> ; \<B> ; (x2 \<leftrightarrow> x1) \<bullet> ((x1, b, c1) #\<^sub>\<Gamma> \<Gamma>) ; \<Delta> \<turnstile>\<^sub>w\<^sub>f (x2 \<leftrightarrow> x1) \<bullet> s1 : (x2 \<leftrightarrow> x1) \<bullet> b_of t1"
+ using wfS.eqvt theta_flip_eq phi_flip_eq assms flip_base_eq beta_flip_eq flip_fresh_fresh supp_b_empty by metis
+ hence " \<Theta> ; \<Phi> ; \<B> ; ((x2, b, (x2 \<leftrightarrow> x1) \<bullet> c1) #\<^sub>\<Gamma> ((x2 \<leftrightarrow> x1) \<bullet> \<Gamma>)) ; \<Delta> \<turnstile>\<^sub>w\<^sub>f (x2 \<leftrightarrow> x1) \<bullet> s1 : b_of ((x2 \<leftrightarrow> x2) \<bullet> t1)" by fastforce
+ thus ?thesis using assms Abs1_eq_iff
+ proof -
+ have f1: "x2 = x1 \<and> t2 = t1 \<or> x2 \<noteq> x1 \<and> t2 = (x2 \<leftrightarrow> x1) \<bullet> t1 \<and> atom x2 \<sharp> t1"
+ by (metis (full_types) Abs1_eq_iff(3) \<open>[[atom x1]]lst. t1 = [[atom x2]]lst. t2\<close>) (* 125 ms *)
+ then have "x2 \<noteq> x1 \<and> s2 = (x2 \<leftrightarrow> x1) \<bullet> s1 \<and> atom x2 \<sharp> s1 \<longrightarrow> b_of t2 = (x2 \<leftrightarrow> x1) \<bullet> b_of t1"
+ by (metis b_of.eqvt) (* 0.0 ms *)
+ then show ?thesis
+ using f1 by (metis (no_types) Abs1_eq_iff(3) G_cons_flip_fresh3 \<open>[[atom x1]]lst. c1 = [[atom x2]]lst. c2\<close> \<open>[[atom x1]]lst. s1 = [[atom x2]]lst. s2\<close> \<open>\<Theta> ; \<Phi> ; \<B> ; (x1, b, c1) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s1 : b_of t1\<close> \<open>\<Theta> ; \<Phi> ; \<B> ; (x2 \<leftrightarrow> x1) \<bullet> ((x1, b, c1) #\<^sub>\<Gamma> \<Gamma>) ; \<Delta> \<turnstile>\<^sub>w\<^sub>f (x2 \<leftrightarrow> x1) \<bullet> s1 : (x2 \<leftrightarrow> x1) \<bullet> b_of t1\<close> \<open>atom x1 \<sharp> \<Gamma>\<close> \<open>atom x2 \<sharp> \<Gamma>\<close>) (* 593 ms *)
+ qed
+qed
+
+section \<open>Lookup\<close>
+
+lemma wf_not_in_prefix:
+ assumes "\<Theta> ; B \<turnstile>\<^sub>w\<^sub>f (\<Gamma>'@(x,b1,c1) #\<^sub>\<Gamma>\<Gamma>)"
+ shows "x \<notin> fst ` toSet \<Gamma>'"
+using assms proof(induct \<Gamma>' rule: \<Gamma>.induct)
+ case GNil
+ then show ?case by simp
+next
+ case (GCons xbc \<Gamma>')
+ then obtain x' and b' and c'::c where xbc: "xbc=(x',b',c')"
+ using prod_cases3 by blast
+ hence *:"(xbc #\<^sub>\<Gamma> \<Gamma>') @ (x, b1, c1) #\<^sub>\<Gamma> \<Gamma> = ((x',b',c') #\<^sub>\<Gamma>(\<Gamma>'@ ((x, b1, c1) #\<^sub>\<Gamma> \<Gamma>)))" by simp
+ hence "atom x' \<sharp> (\<Gamma>'@(x,b1,c1) #\<^sub>\<Gamma>\<Gamma>)" using wfG_elims(2) GCons by metis
+
+ moreover have "\<Theta> ; B \<turnstile>\<^sub>w\<^sub>f (\<Gamma>' @ (x, b1, c1) #\<^sub>\<Gamma> \<Gamma>)" using GCons wfG_elims * by metis
+ ultimately have "atom x' \<notin> atom_dom (\<Gamma>'@(x,b1,c1) #\<^sub>\<Gamma>\<Gamma>)" using wfG_dom_supp GCons append_g.simps xbc fresh_def by fast
+ hence "x' \<noteq> x" using GCons fresh_GCons xbc by fastforce
+ then show ?case using GCons xbc toSet.simps
+ using Un_commute \<open>\<Theta> ; B \<turnstile>\<^sub>w\<^sub>f \<Gamma>' @ (x, b1, c1) #\<^sub>\<Gamma> \<Gamma>\<close> atom_dom.simps by auto
+qed
+
+lemma lookup_inside_wf[simp]:
+ assumes "\<Theta> ; B \<turnstile>\<^sub>w\<^sub>f (\<Gamma>'@(x,b1,c1) #\<^sub>\<Gamma>\<Gamma>)"
+ shows "Some (b1,c1) = lookup (\<Gamma>'@(x,b1,c1) #\<^sub>\<Gamma>\<Gamma>) x"
+ using wf_not_in_prefix lookup_inside assms by fast
+
+lemma lookup_weakening:
+ fixes \<Theta>::\<Theta> and \<Gamma>::\<Gamma> and \<Gamma>'::\<Gamma>
+ assumes "Some (b,c) = lookup \<Gamma> x" and "toSet \<Gamma> \<subseteq> toSet \<Gamma>'" and "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>'" and "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>"
+ shows "Some (b,c) = lookup \<Gamma>' x"
+proof -
+ have "(x,b,c) \<in> toSet \<Gamma> \<and> (\<forall>b' c'. (x,b',c') \<in> toSet \<Gamma> \<longrightarrow> b'=b \<and> c'=c)" using assms lookup_iff toSet.simps by force
+ hence "(x,b,c) \<in> toSet \<Gamma>'" using assms by auto
+ moreover have "(\<forall>b' c'. (x,b',c') \<in> toSet \<Gamma>' \<longrightarrow> b'=b \<and> c'=c)" using assms wf_g_unique
+ using calculation by auto
+ ultimately show ?thesis using lookup_iff
+ using assms(3) by blast
+qed
+
+lemma wfPhi_lookup_fun_unique:
+ fixes \<Phi>::\<Phi>
+ assumes "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>" and "AF_fundef f fd \<in> set \<Phi>"
+ shows "Some (AF_fundef f fd) = lookup_fun \<Phi> f"
+using assms proof(induct \<Phi> rule: list.induct )
+ case Nil
+ then show ?case using lookup_fun.simps by simp
+next
+ case (Cons a \<Phi>')
+ then obtain f' and fd' where a:"a = AF_fundef f' fd'" using fun_def.exhaust by auto
+ have wf: "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>' \<and> f' \<notin> name_of_fun ` set \<Phi>' " using wfPhi_elims Cons a by metis
+ then show ?case using Cons lookup_fun.simps using Cons lookup_fun.simps wf a
+ by (metis image_eqI name_of_fun.simps set_ConsD)
+qed
+
+lemma lookup_fun_weakening:
+ fixes \<Phi>'::\<Phi>
+ assumes "Some fd = lookup_fun \<Phi> f" and "set \<Phi> \<subseteq> set \<Phi>'" and "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>'"
+ shows "Some fd = lookup_fun \<Phi>' f"
+using assms proof(induct \<Phi> )
+ case Nil
+ then show ?case using lookup_fun.simps by simp
+next
+ case (Cons a \<Phi>'')
+ then obtain f' and fd' where a: "a = AF_fundef f' fd'" using fun_def.exhaust by auto
+ then show ?case proof(cases "f=f'")
+ case True
+ then show ?thesis using lookup_fun.simps Cons wfPhi_lookup_fun_unique a
+ by (metis lookup_fun_member subset_iff)
+ next
+ case False
+ then show ?thesis using lookup_fun.simps Cons
+ using \<open>a = AF_fundef f' fd'\<close> by auto
+ qed
+qed
+
+lemma fundef_poly_fresh_bv:
+ assumes "atom bv2 \<sharp> (bv1,b1,c1,\<tau>1,s1)"
+ shows * : "(AF_fun_typ_some bv2 (AF_fun_typ x1 ((bv1\<leftrightarrow>bv2) \<bullet> b1) ((bv1\<leftrightarrow>bv2) \<bullet>c1) ((bv1\<leftrightarrow>bv2) \<bullet> \<tau>1) ((bv1\<leftrightarrow>bv2) \<bullet> s1)) = (AF_fun_typ_some bv1 (AF_fun_typ x1 b1 c1 \<tau>1 s1)))"
+ (is "(AF_fun_typ_some ?bv ?fun_typ = AF_fun_typ_some ?bva ?fun_typa)")
+proof -
+ have 1:"atom bv2 \<notin> set [atom x1]" using bv_not_in_x_atoms by simp
+ have 2:"bv1 \<noteq> bv2" using assms by auto
+ have 3:"(bv2 \<leftrightarrow> bv1) \<bullet> x1 = x1" using pure_fresh flip_fresh_fresh
+ by (simp add: flip_fresh_fresh)
+ have " AF_fun_typ x1 ((bv1 \<leftrightarrow> bv2) \<bullet> b1) ((bv1 \<leftrightarrow> bv2) \<bullet> c1) ((bv1 \<leftrightarrow> bv2) \<bullet> \<tau>1) ((bv1 \<leftrightarrow> bv2) \<bullet> s1) = (bv2 \<leftrightarrow> bv1) \<bullet> AF_fun_typ x1 b1 c1 \<tau>1 s1"
+ using 1 2 3 assms by (simp add: flip_commute)
+ moreover have "(atom bv2 \<sharp> c1 \<and> atom bv2 \<sharp> \<tau>1 \<and> atom bv2 \<sharp> s1 \<or> atom bv2 \<in> set [atom x1]) \<and> atom bv2 \<sharp> b1"
+ using 1 2 3 assms fresh_prod5 by metis
+ ultimately show ?thesis unfolding fun_typ_q.eq_iff Abs1_eq_iff(3) fun_typ.fresh 1 2 by fastforce
+qed
+
+
+text \<open>It is possible to collapse some of the easy to prove inductive cases into a single proof at the qed line
+ but this makes it fragile under change. For example, changing the lemma statement might make one of the previously
+ trivial cases non-trivial and so the collapsing needs to be unpacked. Is there a way to find which case
+ has failed in the qed line?\<close>
+
+lemma wb_b_weakening1:
+ fixes \<Gamma>::\<Gamma> and \<Gamma>'::\<Gamma> and v::v and e::e and c::c and \<tau>::\<tau> and ts::"(string*\<tau>) list" and \<Delta>::\<Delta> and s::s and \<B>::\<B> and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def
+ and cs::branch_s and css::branch_list
+
+ shows "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b \<Longrightarrow> \<B> |\<subseteq>| \<B>' \<Longrightarrow> \<Theta>; \<B>' ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b" and
+ "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f c \<Longrightarrow>\<B> |\<subseteq>| \<B>' \<Longrightarrow> \<Theta>; \<B>' ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f c" and
+ "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<Longrightarrow>\<B> |\<subseteq>| \<B>' \<Longrightarrow> \<Theta>; \<B>' \<turnstile>\<^sub>w\<^sub>f \<Gamma> " and
+ "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau> \<Longrightarrow> \<B> |\<subseteq>| \<B>' \<Longrightarrow> \<Theta>; \<B>' ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau>" and
+ "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ts \<Longrightarrow> \<B> |\<subseteq>| \<B>' \<Longrightarrow> \<Theta>; \<B>' ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ts" and
+ "\<turnstile>\<^sub>w\<^sub>f P \<Longrightarrow> True " and
+ "wfB \<Theta> \<B> b \<Longrightarrow> \<B> |\<subseteq>| \<B>' \<Longrightarrow> wfB \<Theta> \<B>' b" and
+ "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ce : b \<Longrightarrow> \<B> |\<subseteq>| \<B>' \<Longrightarrow> \<Theta>; \<B>' ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ce : b" and
+ "\<Theta> \<turnstile>\<^sub>w\<^sub>f td \<Longrightarrow> True"
+proof(nominal_induct b and c and \<Gamma> and \<tau> and ts and P and b and b and td
+ avoiding: \<B>'
+rule:wfV_wfC_wfG_wfT_wfTs_wfTh_wfB_wfCE_wfTD.strong_induct)
+ case (wfV_conspI s bv dclist \<Theta> dc x b' c \<B> b \<Gamma> v)
+ show ?case proof
+ show \<open>AF_typedef_poly s bv dclist \<in> set \<Theta>\<close> using wfV_conspI by metis
+ show \<open>(dc, \<lbrace> x : b' | c \<rbrace>) \<in> set dclist\<close> using wfV_conspI by auto
+ show \<open> \<Theta> ; \<B>' \<turnstile>\<^sub>w\<^sub>f b \<close> using wfV_conspI by auto
+ show \<open>atom bv \<sharp> (\<Theta>, \<B>', \<Gamma>, b, v)\<close> using fresh_prodN wfV_conspI by auto
+ thus \<open> \<Theta>; \<B>' ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b'[bv::=b]\<^sub>b\<^sub>b \<close> using wfV_conspI by simp
+ qed
+next
+ case (wfTI z \<Theta> \<B> \<Gamma> b c)
+ show ?case proof
+ show "atom z \<sharp> (\<Theta>, \<B>', \<Gamma>)" using wfTI by auto
+ show "\<Theta>; \<B>' \<turnstile>\<^sub>w\<^sub>f b " using wfTI by auto
+ show "\<Theta>; \<B>' ; (z, b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f c " using wfTI by auto
+ qed
+qed( (auto simp add: wf_intros | metis wf_intros)+ )
+
+lemma wb_b_weakening2:
+ fixes \<Gamma>::\<Gamma> and \<Gamma>'::\<Gamma> and v::v and e::e and c::c and \<tau>::\<tau> and ts::"(string*\<tau>) list" and \<Delta>::\<Delta> and s::s and \<B>::\<B> and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def
+ and cs::branch_s and css::branch_list
+
+ shows
+ "\<Theta>; \<Phi>; \<B>; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f e : b \<Longrightarrow> \<B> |\<subseteq>| \<B>' \<Longrightarrow> \<Theta> ; \<Phi> ; \<B>' ; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f e : b" and
+ "\<Theta>; \<Phi>; \<B>; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b \<Longrightarrow> \<B> |\<subseteq>| \<B>' \<Longrightarrow> \<Theta> ; \<Phi> ; \<B>' ; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b" and
+ "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> ; tid ; dc ; t \<turnstile>\<^sub>w\<^sub>f cs : b \<Longrightarrow> \<B> |\<subseteq>| \<B>' \<Longrightarrow> \<Theta> ; \<Phi> ; \<B>' ; \<Gamma> ; \<Delta> ; tid ; dc ; t \<turnstile>\<^sub>w\<^sub>f cs : b" and
+ "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> ; tid ; dclist \<turnstile>\<^sub>w\<^sub>f css : b \<Longrightarrow> \<B> |\<subseteq>| \<B>' \<Longrightarrow> \<Theta> ; \<Phi> ; \<B>' ; \<Gamma> ; \<Delta> ; tid ; dclist \<turnstile>\<^sub>w\<^sub>f css : b" and
+ "\<Theta> \<turnstile>\<^sub>w\<^sub>f (\<Phi>::\<Phi>) \<Longrightarrow> True" and
+ "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<Longrightarrow> \<B> |\<subseteq>| \<B>' \<Longrightarrow> \<Theta>; \<B>' ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>" and
+ "\<Theta> ; \<Phi> \<turnstile>\<^sub>w\<^sub>f ftq \<Longrightarrow> True" and
+ "\<Theta> ; \<Phi> ; \<B> \<turnstile>\<^sub>w\<^sub>f ft \<Longrightarrow> \<B> |\<subseteq>| \<B>' \<Longrightarrow> \<Theta> ; \<Phi> ; \<B>' \<turnstile>\<^sub>w\<^sub>f ft"
+proof(nominal_induct b and b and b and b and \<Phi> and \<Delta> and ftq and ft
+ avoiding: \<B>'
+ rule:wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.strong_induct)
+ case (wfE_valI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v b)
+ then show ?case using wf_intros wb_b_weakening1 by metis
+next
+ case (wfE_plusI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 v2)
+ then show ?case using wf_intros wb_b_weakening1 by metis
+next
+ case (wfE_leqI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 v2)
+ then show ?case using wf_intros wb_b_weakening1 by metis
+next
+ case (wfE_eqI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 b v2)
+ then show ?case using wf_intros wb_b_weakening1
+ by meson
+next
+ case (wfE_fstI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 b1 b2)
+ then show ?case using Wellformed.wfE_fstI wb_b_weakening1 by metis
+next
+ case (wfE_sndI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 b1 b2)
+ then show ?case using wf_intros wb_b_weakening1 by metis
+next
+ case (wfE_concatI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 v2)
+ then show ?case using wf_intros wb_b_weakening1 by metis
+next
+ case (wfE_splitI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 v2)
+ then show ?case using wf_intros wb_b_weakening1 by metis
+next
+ case (wfE_lenI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1)
+ then show ?case using wf_intros wb_b_weakening1 by metis
+next
+ case (wfE_appI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> f ft v)
+ then show ?case using wf_intros using wb_b_weakening1 by meson
+next
+ case (wfE_appPI \<Theta> \<Phi> \<B>1 \<Gamma> \<Delta> b' bv1 v1 \<tau>1 f1 x1 b1 c1 s1)
+
+ have "\<Theta> ; \<Phi> ; \<B>' ; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f AE_appP f1 b' v1 : (b_of \<tau>1)[bv1::=b']\<^sub>b"
+ proof
+ show "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>" using wfE_appPI by auto
+ show "\<Theta>; \<B>' ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> " using wfE_appPI by auto
+ show "\<Theta>; \<B>' \<turnstile>\<^sub>w\<^sub>f b' " using wfE_appPI wb_b_weakening1 by auto
+ thus " atom bv1 \<sharp> (\<Phi>, \<Theta>, \<B>', \<Gamma>, \<Delta>, b', v1, (b_of \<tau>1)[bv1::=b']\<^sub>b)"
+ using wfE_appPI fresh_prodN by auto
+
+ show "Some (AF_fundef f1 (AF_fun_typ_some bv1 (AF_fun_typ x1 b1 c1 \<tau>1 s1))) = lookup_fun \<Phi> f1" using wfE_appPI by auto
+ show "\<Theta>; \<B>' ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v1 : b1[bv1::=b']\<^sub>b " using wfE_appPI wb_b_weakening1 by auto
+ qed
+ then show ?case by auto
+next
+ case (wfE_mvarI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> u \<tau>)
+ then show ?case using wf_intros wb_b_weakening1 by metis
+next
+ case (wfS_valI \<Theta> \<Phi> \<B> \<Gamma> v b \<Delta>)
+ then show ?case using wf_intros wb_b_weakening1 by metis
+next
+ case (wfS_letI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> e b' x s b)
+ show ?case proof
+ show \<open> \<Theta> ; \<Phi> ; \<B>' ; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f e : b' \<close> using wfS_letI by auto
+ show \<open> \<Theta> ; \<Phi> ; \<B>' ; (x, b', TRUE) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b \<close> using wfS_letI by auto
+ show \<open> \<Theta>; \<B>' ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wfS_letI by auto
+ show \<open>atom x \<sharp> (\<Phi>, \<Theta>, \<B>', \<Gamma>, \<Delta>, e, b)\<close> using wfS_letI by auto
+ qed
+next
+ case (wfS_let2I \<Theta> \<Phi> \<B> \<Gamma> \<Delta> s1 \<tau> x s2 b)
+ then show ?case using wb_b_weakening1 Wellformed.wfS_let2I by simp
+next
+ case (wfS_ifI \<Theta> \<B> \<Gamma> v \<Phi> \<Delta> s1 b s2)
+ then show ?case using wb_b_weakening1 Wellformed.wfS_ifI by simp
+next
+ case (wfS_varI \<Theta> \<B> \<Gamma> \<tau> v u \<Delta> \<Phi> s b)
+ then show ?case using wb_b_weakening1 Wellformed.wfS_varI by simp
+next
+ case (wfS_assignI u \<tau> \<Delta> \<Theta> \<B> \<Gamma> \<Phi> v)
+ then show ?case using wb_b_weakening1 Wellformed.wfS_assignI by simp
+next
+case (wfS_whileI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> s1 s2 b)
+ then show ?case using wb_b_weakening1 Wellformed.wfS_whileI by simp
+next
+ case (wfS_seqI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> s1 s2 b)
+ then show ?case using Wellformed.wfS_seqI by metis
+next
+ case (wfS_matchI \<Theta> \<B> \<Gamma> v tid dclist \<Delta> \<Phi> cs b)
+ then show ?case using wb_b_weakening1 Wellformed.wfS_matchI by metis
+next
+ case (wfS_branchI \<Theta> \<Phi> \<B> x \<tau> \<Gamma> \<Delta> s b tid dc)
+ then show ?case using Wellformed.wfS_branchI by auto
+next
+ case (wfS_finalI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist' cs b dclist)
+ then show ?case using wf_intros by metis
+next
+ case (wfS_cons \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist' cs b css dclist)
+ then show ?case using wf_intros by metis
+next
+ case (wfD_emptyI \<Theta> \<B> \<Gamma>)
+ then show ?case using wf_intros wb_b_weakening1 by metis
+next
+ case (wfD_cons \<Theta> \<B> \<Gamma> \<Delta> \<tau> u)
+ then show ?case using wf_intros wb_b_weakening1 by metis
+next
+ case (wfPhi_emptyI \<Theta>)
+ then show ?case using wf_intros wb_b_weakening1 by metis
+next
+ case (wfPhi_consI f \<Theta> \<Phi> ft)
+ then show ?case using wf_intros wb_b_weakening1 by metis
+next
+ case (wfFTSome \<Theta> bv ft)
+ then show ?case using wf_intros wb_b_weakening1 by metis
+next
+ case (wfFTI \<Theta> B b s x c \<tau> \<Phi>)
+ show ?case proof
+ show "\<Theta>; \<B>' \<turnstile>\<^sub>w\<^sub>f b" using wfFTI wb_b_weakening1 by auto
+
+ show "supp c \<subseteq> {atom x}" using wfFTI wb_b_weakening1 by auto
+ show "\<Theta>; \<B>' ; (x, b, c) #\<^sub>\<Gamma> GNil \<turnstile>\<^sub>w\<^sub>f \<tau> " using wfFTI wb_b_weakening1 by auto
+ show "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> " using wfFTI wb_b_weakening1 by auto
+ from \<open> B |\<subseteq>| \<B>'\<close> have "supp B \<subseteq> supp \<B>'" proof(induct B)
+ case empty
+ then show ?case by auto
+ next
+ case (insert x B)
+ then show ?case
+ by (metis fsubset_funion_eq subset_Un_eq supp_union_fset)
+ qed
+ thus "supp s \<subseteq> {atom x} \<union> supp \<B>'" using wfFTI by auto
+ qed
+next
+ case (wfS_assertI \<Theta> \<Phi> \<B> x c \<Gamma> \<Delta> s b)
+ show ?case proof
+ show \<open> \<Theta> ; \<Phi> ; \<B>' ; (x, B_bool, c) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b \<close> using wb_b_weakening1 wfS_assertI by simp
+ show \<open> \<Theta>; \<B>' ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f c \<close> using wb_b_weakening1 wfS_assertI by simp
+ show \<open> \<Theta>; \<B>' ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wb_b_weakening1 wfS_assertI by simp
+ have "atom x \<sharp> \<B>'" using x_not_in_b_set fresh_def by metis
+ thus \<open>atom x \<sharp> (\<Phi>, \<Theta>, \<B>', \<Gamma>, \<Delta>, c, b, s)\<close> using wfS_assertI fresh_prodN by simp
+ qed
+qed(auto)
+
+lemmas wb_b_weakening = wb_b_weakening1 wb_b_weakening2
+
+lemma wfG_b_weakening:
+ fixes \<Gamma>::\<Gamma>
+ assumes "\<B> |\<subseteq>| \<B>'" and "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>"
+ shows "\<Theta>; \<B>' \<turnstile>\<^sub>w\<^sub>f \<Gamma> "
+ using wb_b_weakening assms by auto
+
+lemma wfT_b_weakening:
+ fixes \<Gamma>::\<Gamma> and \<Theta>::\<Theta> and \<tau>::\<tau>
+ assumes "\<B> |\<subseteq>| \<B>'" and "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau>"
+ shows "\<Theta>; \<B>' ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau> "
+ using wb_b_weakening assms by auto
+
+lemma wfB_subst_wfB:
+ fixes \<tau>::\<tau> and b'::b and b::b
+ assumes "\<Theta> ; {|bv|} \<turnstile>\<^sub>w\<^sub>f b" and "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b'"
+ shows "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b[bv::=b']\<^sub>b\<^sub>b "
+using assms proof(nominal_induct b rule:b.strong_induct)
+ case B_int
+ hence "\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f B_int" using wfB_intI wfX_wfY by fast
+ then show ?case using subst_bb.simps wb_b_weakening by fastforce
+next
+ case B_bool
+ hence "\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f B_bool" using wfB_boolI wfX_wfY by fast
+ then show ?case using subst_bb.simps wb_b_weakening by fastforce
+next
+ case (B_id x )
+ hence " \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (B_id x)" using wfB_consI wfB_elims wfX_wfY by metis
+ then show ?case using subst_bb.simps(4) by auto
+next
+ case (B_pair x1 x2)
+ then show ?case using subst_bb.simps
+ by (metis wfB_elims(1) wfB_pairI)
+next
+ case B_unit
+ hence "\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f B_unit" using wfB_unitI wfX_wfY by fast
+ then show ?case using subst_bb.simps wb_b_weakening by fastforce
+next
+ case B_bitvec
+ hence "\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f B_bitvec" using wfB_bitvecI wfX_wfY by fast
+ then show ?case using subst_bb.simps wb_b_weakening by fastforce
+next
+ case (B_var x)
+ then show ?case
+ proof -
+ have False
+ using B_var.prems(1) wfB.cases by fastforce (* 781 ms *)
+ then show ?thesis by metis
+ qed
+next
+ case (B_app s b)
+ then obtain bv' dclist where *:"AF_typedef_poly s bv' dclist \<in> set \<Theta> \<and> \<Theta> ; {|bv|} \<turnstile>\<^sub>w\<^sub>f b" using wfB_elims by metis
+ show ?case unfolding subst_b_simps proof
+ show "\<turnstile>\<^sub>w\<^sub>f \<Theta> " using B_app wfX_wfY by metis
+ show "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f b[bv::=b']\<^sub>b\<^sub>b " using * B_app forget_subst wfB_supp fresh_def
+ by (metis ex_in_conv subset_empty subst_b_b_def supp_empty_fset)
+ show "AF_typedef_poly s bv' dclist \<in> set \<Theta>" using * by auto
+ qed
+qed
+
+lemma wfT_subst_wfB:
+ fixes \<tau>::\<tau> and b'::b
+ assumes "\<Theta> ; {|bv|} ; (x, b, c) #\<^sub>\<Gamma> GNil \<turnstile>\<^sub>w\<^sub>f \<tau>" and "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b'"
+ shows "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (b_of \<tau>)[bv::=b']\<^sub>b\<^sub>b "
+proof -
+ obtain b where "\<Theta> ; {|bv|} \<turnstile>\<^sub>w\<^sub>f b \<and> b_of \<tau> = b" using wfT_elims b_of.simps assms by metis
+ thus ?thesis using wfB_subst_wfB assms by auto
+qed
+
+lemma wfG_cons_unique:
+ assumes "(x1,b1,c1) \<in> toSet (((x,b,c) #\<^sub>\<Gamma>\<Gamma>))" and "wfG \<Theta> \<B> (((x,b,c) #\<^sub>\<Gamma>\<Gamma>))" and "x = x1"
+ shows "b1 = b \<and> c1 = c"
+proof -
+ have "x1 \<notin> fst ` toSet \<Gamma>"
+ proof -
+ have "atom x1 \<sharp> \<Gamma>" using assms wfG_cons by metis
+ then show ?thesis
+ using fresh_gamma_elem
+ by (metis assms(2) atom_dom.simps dom.simps rev_image_eqI wfG_cons2 wfG_x_fresh)
+ qed
+ thus ?thesis using assms by force
+qed
+
+lemma wfG_member_unique:
+ assumes "(x1,b1,c1) \<in> toSet (\<Gamma>'@((x,b,c) #\<^sub>\<Gamma>\<Gamma>))" and "wfG \<Theta> \<B> (\<Gamma>'@((x,b,c) #\<^sub>\<Gamma>\<Gamma>))" and "x = x1"
+ shows "b1 = b \<and> c1 = c"
+ using assms proof(induct \<Gamma>' rule: \<Gamma>_induct)
+ case GNil
+ then show ?case using wfG_suffix wfG_cons_unique append_g.simps by metis
+next
+ case (GCons x' b' c' \<Gamma>')
+ moreover hence "(x1, b1, c1) \<in> toSet (\<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>)" using wf_not_in_prefix by fastforce
+ ultimately show ?case using wfG_cons by fastforce
+qed
+
+section \<open>Function Definitions\<close>
+
+lemma wb_phi_weakening:
+ fixes \<Gamma>::\<Gamma> and \<Gamma>'::\<Gamma> and v::v and e::e and c::c and \<tau>::\<tau> and ts::"(string*\<tau>) list" and \<Delta>::\<Delta> and s::s and \<B>::\<B> and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def
+ and cs::branch_s and css::branch_list and \<Phi>::\<Phi>
+ shows
+ "\<Theta>; \<Phi>; \<B>; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f e : b \<Longrightarrow> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>' \<Longrightarrow> set \<Phi> \<subseteq> set \<Phi>' \<Longrightarrow> \<Theta> ; \<Phi>' ; \<B> ; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f e : b" and
+ "\<Theta>; \<Phi>; \<B>; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b \<Longrightarrow> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>' \<Longrightarrow> set \<Phi> \<subseteq> set \<Phi>' \<Longrightarrow> \<Theta> ; \<Phi>' ; \<B> ; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b" and
+ "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> ; tid ; dc ; t \<turnstile>\<^sub>w\<^sub>f cs : b \<Longrightarrow> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>' \<Longrightarrow> set \<Phi> \<subseteq> set \<Phi>' \<Longrightarrow> \<Theta> ; \<Phi>' ; \<B> ; \<Gamma> ; \<Delta> ; tid ; dc ; t \<turnstile>\<^sub>w\<^sub>f cs : b" and
+ "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> ; tid ; dclist \<turnstile>\<^sub>w\<^sub>f css : b \<Longrightarrow> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>' \<Longrightarrow> set \<Phi> \<subseteq> set \<Phi>' \<Longrightarrow> \<Theta> ; \<Phi>' ; \<B> ; \<Gamma> ; \<Delta> ; tid ; dclist \<turnstile>\<^sub>w\<^sub>f css : b" and
+ "\<Theta> \<turnstile>\<^sub>w\<^sub>f (\<Phi>::\<Phi>) \<Longrightarrow> True" and
+ "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<Longrightarrow> True" and
+ "\<Theta> ; \<Phi> \<turnstile>\<^sub>w\<^sub>f ftq \<Longrightarrow> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>' \<Longrightarrow> set \<Phi> \<subseteq> set \<Phi>' \<Longrightarrow> \<Theta> ; \<Phi>' \<turnstile>\<^sub>w\<^sub>f ftq" and
+ "\<Theta> ; \<Phi> ; \<B> \<turnstile>\<^sub>w\<^sub>f ft \<Longrightarrow> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>' \<Longrightarrow> set \<Phi> \<subseteq> set \<Phi>' \<Longrightarrow> \<Theta> ; \<Phi>' ; \<B> \<turnstile>\<^sub>w\<^sub>f ft"
+proof(nominal_induct
+ b and b and b and b and \<Phi> and \<Delta> and ftq and ft
+ avoiding: \<Phi>'
+ rule:wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.strong_induct)
+ case (wfE_valI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v b)
+ then show ?case using wf_intros by metis
+next
+ case (wfE_plusI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 v2)
+ then show ?case using wf_intros by metis
+next
+ case (wfE_leqI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 v2)
+ then show ?case using wf_intros by metis
+next
+ case (wfE_eqI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 b v2)
+ then show ?case using wf_intros by metis
+next
+ case (wfE_fstI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 b1 b2)
+ then show ?case using wf_intros by metis
+next
+ case (wfE_sndI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 b1 b2)
+ then show ?case using wf_intros by metis
+next
+ case (wfE_concatI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 v2)
+ then show ?case using wf_intros by metis
+next
+ case (wfE_splitI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 v2)
+ then show ?case using wf_intros by metis
+next
+ case (wfE_lenI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1)
+ then show ?case using wf_intros by metis
+next
+ case (wfE_appI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> f x b c \<tau> s v)
+ then show ?case using wf_intros lookup_fun_weakening by metis
+next
+ case (wfE_appPI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> b' bv v \<tau> f x b c s)
+ show ?case proof
+ show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>' \<close> using wfE_appPI by auto
+ show \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wfE_appPI by auto
+ show \<open> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b' \<close> using wfE_appPI by auto
+ show \<open>atom bv \<sharp> (\<Phi>', \<Theta>, \<B>, \<Gamma>, \<Delta>, b', v, (b_of \<tau>)[bv::=b']\<^sub>b)\<close> using wfE_appPI by auto
+ show \<open>Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \<tau> s))) = lookup_fun \<Phi>' f\<close>
+ using wfE_appPI lookup_fun_weakening by metis
+ show \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b[bv::=b']\<^sub>b \<close> using wfE_appPI by auto
+ qed
+next
+ case (wfE_mvarI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> u \<tau>)
+ then show ?case using wf_intros by metis
+next
+ case (wfS_valI \<Theta> \<Phi> \<B> \<Gamma> v b \<Delta>)
+ then show ?case using wf_intros by metis
+next
+ case (wfS_letI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> e b' x s b)
+ then show ?case using Wellformed.wfS_letI by fastforce
+next
+ case (wfS_let2I \<Theta> \<Phi> \<B> \<Gamma> \<Delta> s1 b' x s2 b)
+ then show ?case using Wellformed.wfS_let2I by fastforce
+next
+ case (wfS_ifI \<Theta> \<B> \<Gamma> v \<Phi> \<Delta> s1 b s2)
+ then show ?case using wf_intros by metis
+next
+ case (wfS_varI \<Theta> \<B> \<Gamma> \<tau> v u \<Phi> \<Delta> b s)
+ show ?case proof
+ show \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau> \<close> using wfS_varI by simp
+ show \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b_of \<tau> \<close> using wfS_varI by simp
+ show \<open>atom u \<sharp> (\<Phi>', \<Theta>, \<B>, \<Gamma>, \<Delta>, \<tau>, v, b)\<close> using wfS_varI by simp
+ show \<open> \<Theta> ; \<Phi>' ; \<B> ; \<Gamma> ; (u, \<tau>) #\<^sub>\<Delta> \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b \<close> using wfS_varI by simp
+ qed
+next
+ case (wfS_assignI u \<tau> \<Delta> \<Theta> \<B> \<Gamma> \<Phi> v)
+ then show ?case using wf_intros by metis
+next
+ case (wfS_whileI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> s1 s2 b)
+ then show ?case using wf_intros by metis
+next
+ case (wfS_seqI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> s1 s2 b)
+ then show ?case using wf_intros by metis
+next
+ case (wfS_matchI \<Theta> \<B> \<Gamma> v tid dclist \<Delta> \<Phi> cs b)
+ then show ?case using wf_intros by metis
+next
+ case (wfS_branchI \<Theta> \<Phi> \<B> x \<tau> \<Gamma> \<Delta> s b tid dc)
+ then show ?case using Wellformed.wfS_branchI by fastforce
+next
+ case (wfS_assertI \<Theta> \<Phi> \<B> x c \<Gamma> \<Delta> s b)
+ show ?case proof
+ show \<open> \<Theta> ; \<Phi>' ; \<B> ; (x, B_bool, c) #\<^sub>\<Gamma> \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b \<close> using wfS_assertI by auto
+ next
+ show \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f c \<close> using wfS_assertI by auto
+ next
+ show \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wfS_assertI by auto
+ have "atom x \<sharp> \<Phi>'" using wfS_assertI wfPhi_supp fresh_def by blast
+ thus \<open>atom x \<sharp> (\<Phi>', \<Theta>, \<B>, \<Gamma>, \<Delta>, c, b, s)\<close> using fresh_prodN wfS_assertI wfPhi_supp fresh_def by auto
+ qed
+next
+ case (wfFTI \<Theta> B b s x c \<tau> \<Phi>)
+ show ?case proof
+ show \<open> \<Theta> ; B \<turnstile>\<^sub>w\<^sub>f b \<close> using wfFTI by auto
+ next
+ show \<open>supp c \<subseteq> {atom x}\<close> using wfFTI by auto
+ next
+ show \<open> \<Theta> ; B ; (x, b, c) #\<^sub>\<Gamma> GNil \<turnstile>\<^sub>w\<^sub>f \<tau> \<close> using wfFTI by auto
+ next
+ show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>' \<close> using wfFTI by auto
+ next
+ show \<open>supp s \<subseteq> {atom x} \<union> supp B\<close> using wfFTI by auto
+ qed
+qed(auto|metis wf_intros)+
+
+
+lemma wfT_fun_return_t:
+ fixes \<tau>a'::\<tau> and \<tau>'::\<tau>
+ assumes "\<Theta>; \<B>; (xa, b, ca) #\<^sub>\<Gamma> GNil \<turnstile>\<^sub>w\<^sub>f \<tau>a'" and "(AF_fun_typ x b c \<tau>' s') = (AF_fun_typ xa b ca \<tau>a' sa')"
+ shows "\<Theta>; \<B>; (x, b, c) #\<^sub>\<Gamma> GNil \<turnstile>\<^sub>w\<^sub>f \<tau>'"
+proof -
+ obtain cb::x where xf: "atom cb \<sharp> (c, \<tau>', s', sa', \<tau>a', ca, x , xa)" using obtain_fresh by blast
+ hence "atom cb \<sharp> (c, \<tau>', s', sa', \<tau>a', ca) \<and> atom cb \<sharp> (x, xa, ((c, \<tau>'), s'), (ca, \<tau>a'), sa')" using fresh_prod6 fresh_prod4 fresh_prod8 by auto
+ hence *:"c[x::=V_var cb]\<^sub>c\<^sub>v = ca[xa::=V_var cb]\<^sub>c\<^sub>v \<and> \<tau>'[x::=V_var cb]\<^sub>\<tau>\<^sub>v = \<tau>a'[xa::=V_var cb]\<^sub>\<tau>\<^sub>v" using assms \<tau>.eq_iff Abs1_eq_iff_all by auto
+
+ have **: "\<Theta>; \<B>; (xa \<leftrightarrow> cb ) \<bullet> ((xa, b, ca) #\<^sub>\<Gamma> GNil) \<turnstile>\<^sub>w\<^sub>f (xa \<leftrightarrow> cb ) \<bullet> \<tau>a'" using assms True_eqvt beta_flip_eq theta_flip_eq wfG_wf
+ by (metis GCons_eqvt GNil_eqvt wfT.eqvt wfT_wf)
+
+ have "\<Theta>; \<B>; (x \<leftrightarrow> cb ) \<bullet> ((x, b, c) #\<^sub>\<Gamma> GNil) \<turnstile>\<^sub>w\<^sub>f (x \<leftrightarrow> cb ) \<bullet> \<tau>'" proof -
+ have "(xa \<leftrightarrow> cb ) \<bullet> xa = (x \<leftrightarrow> cb ) \<bullet> x" using xf by auto
+ hence "(x \<leftrightarrow> cb ) \<bullet> ((x, b, c) #\<^sub>\<Gamma> GNil) = (xa \<leftrightarrow> cb ) \<bullet> ((xa, b, ca) #\<^sub>\<Gamma> GNil)" using * ** xf G_cons_flip fresh_GNil by simp
+ thus ?thesis using ** * xf by simp
+ qed
+ thus ?thesis using beta_flip_eq theta_flip_eq wfT_wf wfG_wf * ** True_eqvt wfT.eqvt permute_flip_cancel by metis
+qed
+
+lemma wfFT_wf_aux:
+ fixes \<tau>::\<tau> and \<Theta>::\<Theta> and \<Phi>::\<Phi> and ft :: fun_typ_q and s::s and \<Delta>::\<Delta>
+ assumes "\<Theta> ; \<Phi> ; B \<turnstile>\<^sub>w\<^sub>f (AF_fun_typ x b c \<tau> s)"
+ shows "\<Theta> ; B ; (x,b,c) #\<^sub>\<Gamma> GNil \<turnstile>\<^sub>w\<^sub>f \<tau> \<and> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<and> supp s \<subseteq> { atom x } \<union> supp B"
+proof -
+ obtain xa and ca and sa and \<tau>' where *:"\<Theta> ; B \<turnstile>\<^sub>w\<^sub>f b \<and> (\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> ) \<and>
+ supp sa \<subseteq> {atom xa} \<union> supp B \<and> (\<Theta> ; B ; (xa, b, ca) #\<^sub>\<Gamma> GNil \<turnstile>\<^sub>w\<^sub>f \<tau>') \<and>
+ AF_fun_typ x b c \<tau> s = AF_fun_typ xa b ca \<tau>' sa "
+ using wfFT.simps[of \<Theta> \<Phi> B "AF_fun_typ x b c \<tau> s"] assms by auto
+
+ moreover hence **: "(AF_fun_typ x b c \<tau> s) = (AF_fun_typ xa b ca \<tau>' sa)" by simp
+ ultimately have "\<Theta> ; B ; (x,b,c) #\<^sub>\<Gamma>GNil \<turnstile>\<^sub>w\<^sub>f \<tau>" using wfT_fun_return_t by metis
+ moreover have " (\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> ) " using * by auto
+ moreover have "supp s \<subseteq> { atom x } \<union> supp B" proof -
+ have "[[atom x]]lst.s = [[atom xa]]lst.sa" using ** fun_typ.eq_iff lst_fst lst_snd by metis
+ thus ?thesis using lst_supp_subset * by metis
+ qed
+ ultimately show ?thesis by auto
+qed
+
+lemma wfFT_simple_wf:
+ fixes \<tau>::\<tau> and \<Theta>::\<Theta> and \<Phi>::\<Phi> and ft :: fun_typ_q and s::s and \<Delta>::\<Delta>
+ assumes "\<Theta> ; \<Phi> \<turnstile>\<^sub>w\<^sub>f (AF_fun_typ_none (AF_fun_typ x b c \<tau> s))"
+ shows "\<Theta> ; {||} ; (x,b,c) #\<^sub>\<Gamma>GNil \<turnstile>\<^sub>w\<^sub>f \<tau> \<and> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<and> supp s \<subseteq> { atom x } "
+proof -
+ have *:"\<Theta> ; \<Phi> ; {||} \<turnstile>\<^sub>w\<^sub>f (AF_fun_typ x b c \<tau> s)" using wfFTQ_elims assms by auto
+ thus ?thesis using wfFT_wf_aux by force
+qed
+
+lemma wfFT_poly_wf:
+ fixes \<tau>::\<tau> and \<Theta>::\<Theta> and \<Phi>::\<Phi> and ftq :: fun_typ_q and s::s and \<Delta>::\<Delta>
+ assumes "\<Theta> ; \<Phi> \<turnstile>\<^sub>w\<^sub>f (AF_fun_typ_some bv (AF_fun_typ x b c \<tau> s))"
+ shows "\<Theta> ; {|bv|} ; (x,b,c) #\<^sub>\<Gamma>GNil \<turnstile>\<^sub>w\<^sub>f \<tau> \<and> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<and> \<Theta> ; \<Phi> ; {|bv|} \<turnstile>\<^sub>w\<^sub>f (AF_fun_typ x b c \<tau> s)"
+proof -
+ obtain bv1 ft1 where *:"\<Theta> ; \<Phi> ; {|bv1|} \<turnstile>\<^sub>w\<^sub>f ft1 \<and> [[atom bv1]]lst. ft1 = [[atom bv]]lst. AF_fun_typ x b c \<tau> s"
+ using wfFTQ_elims(3)[OF assms] by metis
+
+ show ?thesis proof(cases "bv1 = bv")
+ case True
+ then show ?thesis using * fun_typ_q.eq_iff Abs1_eq_iff by (metis (no_types, opaque_lifting) wfFT_wf_aux)
+ next
+ case False
+ obtain x1 b1 c1 t1 s1 where **: "ft1 = AF_fun_typ x1 b1 c1 t1 s1" using fun_typ.eq_iff
+ by (meson fun_typ.exhaust)
+
+ hence eqv: "(bv \<leftrightarrow> bv1) \<bullet> AF_fun_typ x1 b1 c1 t1 s1 = AF_fun_typ x b c \<tau> s \<and> atom bv1 \<sharp> AF_fun_typ x b c \<tau> s" using
+ Abs1_eq_iff(3) * False by metis
+
+ have "(bv \<leftrightarrow> bv1) \<bullet> \<Theta> ; (bv \<leftrightarrow> bv1) \<bullet> \<Phi> ; (bv \<leftrightarrow> bv1) \<bullet> {|bv1|} \<turnstile>\<^sub>w\<^sub>f (bv \<leftrightarrow> bv1) \<bullet> ft1" using wfFT.eqvt * by metis
+ moreover have "(bv \<leftrightarrow> bv1) \<bullet> \<Phi> = \<Phi>" using phi_flip_eq wfX_wfY * by metis
+ moreover have "(bv \<leftrightarrow> bv1) \<bullet> \<Theta> =\<Theta>" using wfX_wfY * theta_flip_eq2 by metis
+ moreover have "(bv \<leftrightarrow> bv1) \<bullet> ft1 = AF_fun_typ x b c \<tau> s" using eqv ** by metis
+ ultimately have "\<Theta> ; \<Phi> ; {|bv|} \<turnstile>\<^sub>w\<^sub>f AF_fun_typ x b c \<tau> s" by auto
+ thus ?thesis using wfFT_wf_aux by auto
+ qed
+qed
+
+lemma wfFT_poly_wfT:
+ fixes \<tau>::\<tau> and \<Theta>::\<Theta> and \<Phi>::\<Phi> and ft :: fun_typ_q
+ assumes "\<Theta> ; \<Phi> \<turnstile>\<^sub>w\<^sub>f (AF_fun_typ_some bv (AF_fun_typ x b c \<tau> s))"
+ shows "\<Theta> ; {| bv |} ; (x,b,c) #\<^sub>\<Gamma>GNil \<turnstile>\<^sub>w\<^sub>f \<tau>"
+ using wfFT_poly_wf assms by simp
+
+lemma b_of_supp:
+ "supp (b_of t) \<subseteq> supp t"
+proof(nominal_induct t rule:\<tau>.strong_induct)
+ case (T_refined_type x b c)
+ then show ?case by auto
+qed
+
+lemma wfPhi_f_simple_wf:
+ fixes \<tau>::\<tau> and \<Theta>::\<Theta> and \<Phi>::\<Phi> and ft :: fun_typ_q and s::s and \<Phi>'::\<Phi>
+ assumes "AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \<tau> s)) \<in> set \<Phi> " and "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>" and "set \<Phi> \<subseteq> set \<Phi>'" and "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>'"
+ shows "\<Theta> ; {||} ; (x,b,c) #\<^sub>\<Gamma> GNil \<turnstile>\<^sub>w\<^sub>f \<tau> \<and> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<and> supp s \<subseteq> { atom x }"
+using assms proof(induct \<Phi> rule: \<Phi>_induct)
+ case PNil
+ then show ?case by auto
+next
+ case (PConsSome f1 bv x1 b1 c1 \<tau>1 s' \<Phi>'')
+ hence "AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \<tau> s)) \<in> set \<Phi>''" by auto
+ moreover have " \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>'' \<and> set \<Phi>'' \<subseteq> set \<Phi>'" using wfPhi_elims(3) PConsSome by auto
+ ultimately show ?case using PConsSome wfPhi_elims wfFT_simple_wf by auto
+next
+ case (PConsNone f' x' b' c' \<tau>' s' \<Phi>'')
+ show ?case proof(cases "f=f'")
+ case True
+ have "AF_fun_typ_none (AF_fun_typ x' b' c' \<tau>' s') = AF_fun_typ_none (AF_fun_typ x b c \<tau> s)"
+ by (metis PConsNone.prems(1) PConsNone.prems(2) True fun_def.eq_iff image_eqI name_of_fun.simps set_ConsD wfPhi_elims(2))
+ hence *:"\<Theta> ; \<Phi>'' \<turnstile>\<^sub>w\<^sub>f AF_fun_typ_none (AF_fun_typ x b c \<tau> s) " using wfPhi_elims(2)[OF PConsNone(3)] by metis
+ hence "\<Theta> ; \<Phi>'' ; {||} \<turnstile>\<^sub>w\<^sub>f (AF_fun_typ x b c \<tau> s)" using wfFTQ_elims(1) by metis
+ thus ?thesis using wfFT_simple_wf[OF *] wb_phi_weakening PConsNone by force
+ next
+ case False
+ hence "AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \<tau> s)) \<in> set \<Phi>''" using PConsNone by simp
+ moreover have " \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>'' \<and> set \<Phi>'' \<subseteq> set \<Phi>'" using wfPhi_elims(3) PConsNone by auto
+ ultimately show ?thesis using PConsNone wfPhi_elims wfFT_simple_wf by auto
+ qed
+qed
+
+lemma wfPhi_f_simple_wfT:
+ fixes \<tau>::\<tau> and \<Theta>::\<Theta> and \<Phi>::\<Phi> and ft :: fun_typ_q
+ assumes "Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \<tau> s))) = lookup_fun \<Phi> f" and "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>"
+ shows "\<Theta> ; {||} ; (x,b,c) #\<^sub>\<Gamma>GNil \<turnstile>\<^sub>w\<^sub>f \<tau>"
+ using wfPhi_f_simple_wf assms using lookup_fun_member by blast
+
+lemma wfPhi_f_simple_supp_b:
+ fixes \<tau>::\<tau> and \<Theta>::\<Theta> and \<Phi>::\<Phi> and ft :: fun_typ_q
+ assumes "Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \<tau> s))) = lookup_fun \<Phi> f" and "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>"
+ shows "supp b = {}"
+proof -
+ have "\<Theta> ; {||} ; (x,b,c) #\<^sub>\<Gamma>GNil \<turnstile>\<^sub>w\<^sub>f \<tau>" using wfPhi_f_simple_wfT assms by auto
+ thus ?thesis using wfT_wf wfG_cons wfB_supp by fastforce
+qed
+
+lemma wfPhi_f_simple_supp_t:
+ fixes \<tau>::\<tau> and \<Theta>::\<Theta> and \<Phi>::\<Phi> and ft :: fun_typ_q
+ assumes "Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \<tau> s))) = lookup_fun \<Phi> f" and "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>"
+ shows "supp \<tau> \<subseteq> { atom x }"
+ using wfPhi_f_simple_wfT wfT_supp assms by fastforce
+
+lemma wfPhi_f_simple_supp_c:
+ fixes \<tau>::\<tau> and \<Theta>::\<Theta> and \<Phi>::\<Phi> and ft :: fun_typ_q
+ assumes "Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \<tau> s))) = lookup_fun \<Phi> f" and "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>"
+ shows "supp c \<subseteq> { atom x }"
+proof -
+ have "\<Theta> ; {||} ; (x,b,c) #\<^sub>\<Gamma>GNil \<turnstile>\<^sub>w\<^sub>f \<tau>" using wfPhi_f_simple_wfT assms by auto
+ thus ?thesis using wfG_wfC wfC_supp wfT_wf by fastforce
+qed
+
+lemma wfPhi_f_simple_supp_s:
+ fixes \<tau>::\<tau> and \<Theta>::\<Theta> and \<Phi>::\<Phi> and ft :: fun_typ_q
+ assumes "Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \<tau> s))) = lookup_fun \<Phi> f" and "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>"
+ shows "supp s \<subseteq> {atom x}"
+proof -
+ have "AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \<tau> s)) \<in> set \<Phi>" using lookup_fun_member assms by auto
+ hence "supp s \<subseteq> { atom x }" using wfPhi_f_simple_wf assms by blast
+ thus ?thesis using wf_supp(3) atom_dom.simps toSet.simps x_not_in_u_set x_not_in_b_set setD.simps
+ using wf_supp2(2) by fastforce
+qed
+
+lemma wfPhi_f_poly_wf:
+ fixes \<tau>::\<tau> and \<Theta>::\<Theta> and \<Phi>::\<Phi> and ft :: fun_typ_q and s::s and \<Phi>'::\<Phi>
+ assumes "AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \<tau> s)) \<in> set \<Phi> " and "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>" and "set \<Phi> \<subseteq> set \<Phi>'" and "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>'"
+ shows "\<Theta> ; {|bv|} ; (x,b,c) #\<^sub>\<Gamma>GNil \<turnstile>\<^sub>w\<^sub>f \<tau> \<and> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>' \<and> \<Theta> ; \<Phi>' ; {|bv|} \<turnstile>\<^sub>w\<^sub>f (AF_fun_typ x b c \<tau> s)"
+using assms proof(induct \<Phi> rule: \<Phi>_induct)
+ case PNil
+ then show ?case by auto
+next
+ case (PConsNone f x b c \<tau> s' \<Phi>'')
+ moreover have " \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>'' \<and> set \<Phi>'' \<subseteq> set \<Phi>'" using wfPhi_elims(3) PConsNone by auto
+ ultimately show ?case using PConsNone wfPhi_elims wfFT_poly_wf by auto
+next
+ case (PConsSome f1 bv1 x1 b1 c1 \<tau>1 s1 \<Phi>'')
+ show ?case proof(cases "f=f1")
+ case True
+ have "AF_fun_typ_some bv1 (AF_fun_typ x1 b1 c1 \<tau>1 s1) = AF_fun_typ_some bv (AF_fun_typ x b c \<tau> s)"
+ by (metis PConsSome.prems(1) PConsSome.prems(2) True fun_def.eq_iff list.set_intros(1) option.inject wfPhi_lookup_fun_unique)
+ hence *:"\<Theta> ; \<Phi>'' \<turnstile>\<^sub>w\<^sub>f AF_fun_typ_some bv (AF_fun_typ x b c \<tau> s) " using wfPhi_elims PConsSome by metis
+ thus ?thesis using wfFT_poly_wf * wb_phi_weakening PConsSome
+ by (meson set_subset_Cons)
+ next
+ case False
+ hence "AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \<tau> s)) \<in> set \<Phi>''" using PConsSome
+ by (meson fun_def.eq_iff set_ConsD)
+ moreover have " \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>'' \<and> set \<Phi>'' \<subseteq> set \<Phi>'" using wfPhi_elims(3) PConsSome
+ by (meson dual_order.trans set_subset_Cons)
+ ultimately show ?thesis using PConsSome wfPhi_elims wfFT_poly_wf
+ by blast
+ qed
+qed
+
+lemma wfPhi_f_poly_wfT:
+ fixes \<tau>::\<tau> and \<Theta>::\<Theta> and \<Phi>::\<Phi> and ft :: fun_typ_q
+ assumes "Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \<tau> s))) = lookup_fun \<Phi> f" and "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>"
+ shows "\<Theta> ; {| bv |} ; (x,b,c) #\<^sub>\<Gamma>GNil \<turnstile>\<^sub>w\<^sub>f \<tau>"
+using assms proof(induct \<Phi> rule: \<Phi>_induct)
+ case PNil
+ then show ?case by auto
+next
+ case (PConsSome f1 bv1 x1 b1 c1 \<tau>1 s' \<Phi>')
+ then show ?case proof(cases "f1=f")
+ case True
+ hence "lookup_fun (AF_fundef f1 (AF_fun_typ_some bv1 (AF_fun_typ x1 b1 c1 \<tau>1 s')) # \<Phi>') f = Some (AF_fundef f1 (AF_fun_typ_some bv1 (AF_fun_typ x1 b1 c1 \<tau>1 s')))" using
+ lookup_fun.simps using PConsSome.prems by simp
+ then show ?thesis using PConsSome.prems wfPhi_elims wfFT_poly_wfT
+ by (metis option.inject)
+ next
+ case False
+ then show ?thesis using PConsSome using lookup_fun.simps
+ using wfPhi_elims(3) by auto
+ qed
+next
+ case (PConsNone f' x' b' c' \<tau>' s' \<Phi>')
+ then show ?case proof(cases "f'=f")
+ case True
+ then have *:"\<Theta> ; \<Phi>' \<turnstile>\<^sub>w\<^sub>f AF_fun_typ_none (AF_fun_typ x' b' c' \<tau>' s') " using lookup_fun.simps PConsNone wfPhi_elims by metis
+ thus ?thesis using PConsNone wfFT_poly_wfT wfPhi_elims lookup_fun.simps
+ by (metis fun_def.eq_iff fun_typ_q.distinct(1) option.inject)
+ next
+ case False
+ thus ?thesis using PConsNone wfPhi_elims
+ by (metis False lookup_fun.simps(2))
+ qed
+qed
+
+lemma wfPhi_f_poly_supp_b:
+ fixes \<tau>::\<tau> and \<Theta>::\<Theta> and \<Phi>::\<Phi> and ft :: fun_typ_q
+ assumes "Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \<tau> s))) = lookup_fun \<Phi> f" and "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>"
+ shows "supp b \<subseteq> supp bv"
+proof -
+ have "\<Theta> ; {|bv|} ; (x,b,c) #\<^sub>\<Gamma>GNil \<turnstile>\<^sub>w\<^sub>f \<tau>" using wfPhi_f_poly_wfT assms by auto
+ thus ?thesis using wfT_wf wfG_cons wfB_supp by fastforce
+qed
+
+lemma wfPhi_f_poly_supp_t:
+ fixes \<tau>::\<tau> and \<Theta>::\<Theta> and \<Phi>::\<Phi> and ft :: fun_typ_q
+ assumes "Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \<tau> s))) = lookup_fun \<Phi> f" and "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>"
+ shows "supp \<tau> \<subseteq> { atom x , atom bv }"
+ using wfPhi_f_poly_wfT[OF assms, THEN wfT_supp] atom_dom.simps supp_at_base by auto
+
+
+lemma wfPhi_f_poly_supp_b_of_t:
+ fixes \<tau>::\<tau> and \<Theta>::\<Theta> and \<Phi>::\<Phi> and ft :: fun_typ_q
+ assumes "Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \<tau> s))) = lookup_fun \<Phi> f" and "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>"
+ shows "supp (b_of \<tau>) \<subseteq> { atom bv }"
+proof -
+ have "atom x \<notin> supp (b_of \<tau>)" using x_fresh_b by auto
+ moreover have "supp (b_of \<tau>) \<subseteq> { atom x , atom bv }" using wfPhi_f_poly_supp_t
+ using supp_at_base b_of.simps wfPhi_f_poly_supp_t \<tau>.supp b_of_supp assms by fast
+ ultimately show ?thesis by blast
+qed
+
+lemma wfPhi_f_poly_supp_c:
+ fixes \<tau>::\<tau> and \<Theta>::\<Theta> and \<Phi>::\<Phi> and ft :: fun_typ_q
+ assumes "Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \<tau> s))) = lookup_fun \<Phi> f" and "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>"
+ shows "supp c \<subseteq> { atom x, atom bv }"
+proof -
+ have "\<Theta> ; {|bv|} ; (x,b,c) #\<^sub>\<Gamma>GNil \<turnstile>\<^sub>w\<^sub>f \<tau>" using wfPhi_f_poly_wfT assms by auto
+ thus ?thesis using wfG_wfC wfC_supp wfT_wf
+ using supp_at_base by fastforce
+qed
+
+lemma wfPhi_f_poly_supp_s:
+ fixes \<tau>::\<tau> and \<Theta>::\<Theta> and \<Phi>::\<Phi> and ft :: fun_typ_q
+ assumes "Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \<tau> s))) = lookup_fun \<Phi> f" and "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>"
+ shows "supp s \<subseteq> {atom x, atom bv}"
+proof -
+
+ have "AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \<tau> s)) \<in> set \<Phi>" using lookup_fun_member assms by auto
+ hence *:"\<Theta> ; \<Phi> ; {|bv|} \<turnstile>\<^sub>w\<^sub>f (AF_fun_typ x b c \<tau> s)" using assms wfPhi_f_poly_wf by simp
+
+ thus ?thesis using wfFT_wf_aux[OF *] using supp_at_base by auto
+qed
+
+lemmas wfPhi_f_supp = wfPhi_f_poly_supp_b wfPhi_f_simple_supp_b wfPhi_f_poly_supp_c
+ wfPhi_f_simple_supp_t wfPhi_f_poly_supp_t wfPhi_f_simple_supp_t wfPhi_f_poly_wfT wfPhi_f_simple_wfT
+ wfPhi_f_poly_supp_s wfPhi_f_simple_supp_s
+
+lemma fun_typ_eq_ret_unique:
+ assumes "(AF_fun_typ x1 b1 c1 \<tau>1' s1') = (AF_fun_typ x2 b2 c2 \<tau>2' s2')"
+ shows "\<tau>1'[x1::=v]\<^sub>\<tau>\<^sub>v = \<tau>2'[x2::=v]\<^sub>\<tau>\<^sub>v"
+proof -
+ have "[[atom x1]]lst. \<tau>1' = [[atom x2]]lst. \<tau>2'" using assms lst_fst fun_typ.eq_iff lst_snd by metis
+ thus ?thesis using subst_v_flip_eq_two[of x1 \<tau>1' x2 \<tau>2' v] subst_v_\<tau>_def by metis
+qed
+
+lemma fun_typ_eq_body_unique:
+ fixes v::v and x1::x and x2::x and s1'::s and s2'::s
+ assumes "(AF_fun_typ x1 b1 c1 \<tau>1' s1') = (AF_fun_typ x2 b2 c2 \<tau>2' s2')"
+ shows "s1'[x1::=v]\<^sub>s\<^sub>v = s2'[x2::=v]\<^sub>s\<^sub>v"
+proof -
+ have "[[atom x1]]lst. s1' = [[atom x2]]lst. s2'" using assms lst_fst fun_typ.eq_iff lst_snd by metis
+ thus ?thesis using subst_v_flip_eq_two[of x1 s1' x2 s2' v] subst_v_s_def by metis
+qed
+
+lemma fun_ret_unique:
+ assumes "Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x1 b1 c1 \<tau>1' s1'))) = lookup_fun \<Phi> f" and "Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x2 b2 c2 \<tau>2' s2'))) = lookup_fun \<Phi> f"
+ shows "\<tau>1'[x1::=v]\<^sub>\<tau>\<^sub>v = \<tau>2'[x2::=v]\<^sub>\<tau>\<^sub>v"
+proof -
+ have *: " (AF_fundef f (AF_fun_typ_none (AF_fun_typ x1 b1 c1 \<tau>1' s1'))) = (AF_fundef f (AF_fun_typ_none (AF_fun_typ x2 b2 c2 \<tau>2' s2')))" using option.inject assms by metis
+ thus ?thesis using fun_typ_eq_ret_unique fun_def.eq_iff fun_typ_q.eq_iff by metis
+qed
+
+lemma fun_poly_arg_unique:
+ fixes bv1::bv and bv2::bv and b::b and \<tau>1::\<tau> and \<tau>2::\<tau>
+ assumes "[[atom bv1]]lst. (AF_fun_typ x1 b1 c1 \<tau>1 s1) = [[atom bv2]]lst. (AF_fun_typ x2 b2 c2 \<tau>2 s2)" (is "[[atom ?x]]lst. ?a = [[atom ?y]]lst. ?b")
+ shows "\<lbrace> x1 : b1[bv1::=b]\<^sub>b\<^sub>b | c1[bv1::=b]\<^sub>c\<^sub>b \<rbrace> = \<lbrace> x2 : b2[bv2::=b]\<^sub>b\<^sub>b | c2[bv2::=b]\<^sub>c\<^sub>b \<rbrace>"
+proof -
+ obtain c::bv where *:"atom c \<sharp> (b,b1,b2,c1,c2) \<and> atom c \<sharp> (bv1, bv2, AF_fun_typ x1 b1 c1 \<tau>1 s1, AF_fun_typ x2 b2 c2 \<tau>2 s2)" using obtain_fresh fresh_Pair by metis
+ hence "(bv1 \<leftrightarrow> c) \<bullet> AF_fun_typ x1 b1 c1 \<tau>1 s1 = (bv2 \<leftrightarrow> c) \<bullet> AF_fun_typ x2 b2 c2 \<tau>2 s2" using Abs1_eq_iff_all(3)[of ?x ?a ?y ?b] assms by metis
+ hence "AF_fun_typ x1 ((bv1 \<leftrightarrow> c) \<bullet> b1) ((bv1 \<leftrightarrow> c) \<bullet> c1) ((bv1 \<leftrightarrow> c) \<bullet> \<tau>1) ((bv1 \<leftrightarrow> c) \<bullet> s1) = AF_fun_typ x2 ((bv2 \<leftrightarrow> c) \<bullet> b2) ((bv2 \<leftrightarrow> c) \<bullet> c2) ((bv2 \<leftrightarrow> c) \<bullet> \<tau>2) ((bv2 \<leftrightarrow> c) \<bullet> s2)"
+ using fun_typ_flip by metis
+ hence **:"\<lbrace> x1 :((bv1 \<leftrightarrow> c) \<bullet> b1) | ((bv1 \<leftrightarrow> c) \<bullet> c1) \<rbrace> = \<lbrace> x2 : ((bv2 \<leftrightarrow> c) \<bullet> b2) | ((bv2 \<leftrightarrow> c) \<bullet> c2) \<rbrace>" (is "\<lbrace> x1 : ?b1 | ?c1 \<rbrace> = \<lbrace> x2 : ?b2 | ?c2 \<rbrace>") using fun_arg_unique_aux by metis
+ hence "\<lbrace> x1 :((bv1 \<leftrightarrow> c) \<bullet> b1) | ((bv1 \<leftrightarrow> c) \<bullet> c1) \<rbrace>[c::=b]\<^sub>\<tau>\<^sub>b = \<lbrace> x2 : ((bv2 \<leftrightarrow> c) \<bullet> b2) | ((bv2 \<leftrightarrow> c) \<bullet> c2) \<rbrace>[c::=b]\<^sub>\<tau>\<^sub>b" by metis
+ hence "\<lbrace> x1 :((bv1 \<leftrightarrow> c) \<bullet> b1)[c::=b]\<^sub>b\<^sub>b | ((bv1 \<leftrightarrow> c) \<bullet> c1)[c::=b]\<^sub>c\<^sub>b \<rbrace> = \<lbrace> x2 : ((bv2 \<leftrightarrow> c) \<bullet> b2)[c::=b]\<^sub>b\<^sub>b | ((bv2 \<leftrightarrow> c) \<bullet> c2)[c::=b]\<^sub>c\<^sub>b \<rbrace>" using subst_tb.simps by metis
+ thus ?thesis using * flip_subst_subst subst_b_c_def subst_b_b_def fresh_prodN flip_commute by metis
+qed
+
+lemma fun_poly_ret_unique:
+ assumes "Some (AF_fundef f (AF_fun_typ_some bv1 (AF_fun_typ x1 b1 c1 \<tau>1' s1'))) = lookup_fun \<Phi> f" and "Some (AF_fundef f (AF_fun_typ_some bv2 (AF_fun_typ x2 b2 c2 \<tau>2' s2'))) = lookup_fun \<Phi> f"
+ shows "\<tau>1'[bv1::=b]\<^sub>\<tau>\<^sub>b[x1::=v]\<^sub>\<tau>\<^sub>v = \<tau>2'[bv2::=b]\<^sub>\<tau>\<^sub>b[x2::=v]\<^sub>\<tau>\<^sub>v"
+proof -
+ have *: " (AF_fundef f (AF_fun_typ_some bv1 (AF_fun_typ x1 b1 c1 \<tau>1' s1'))) = (AF_fundef f (AF_fun_typ_some bv2 (AF_fun_typ x2 b2 c2 \<tau>2' s2')))" using option.inject assms by metis
+ hence "AF_fun_typ_some bv1 (AF_fun_typ x1 b1 c1 \<tau>1' s1') = AF_fun_typ_some bv2 (AF_fun_typ x2 b2 c2 \<tau>2' s2')"
+ (is "AF_fun_typ_some bv1 ?ft1 = AF_fun_typ_some bv2 ?ft2") using fun_def.eq_iff by metis
+ hence **:"[[atom bv1]]lst. ?ft1 = [[atom bv2]]lst. ?ft2" using fun_typ_q.eq_iff(1) by metis
+
+ hence *:"subst_ft_b ?ft1 bv1 b = subst_ft_b ?ft2 bv2 b" using subst_b_flip_eq_two subst_b_fun_typ_def by metis
+ have "[[atom x1]]lst. \<tau>1'[bv1::=b]\<^sub>\<tau>\<^sub>b = [[atom x2]]lst. \<tau>2'[bv2::=b]\<^sub>\<tau>\<^sub>b"
+ apply(rule lst_snd[of _ "c1[bv1::=b]\<^sub>c\<^sub>b" _ _ "c2[bv2::=b]\<^sub>c\<^sub>b"])
+ apply(rule lst_fst[of _ _ "s1'[bv1::=b]\<^sub>s\<^sub>b" _ _ "s2'[bv2::=b]\<^sub>s\<^sub>b"])
+ using * subst_ft_b.simps fun_typ.eq_iff by metis
+ thus ?thesis using subst_v_flip_eq_two subst_v_\<tau>_def by metis
+qed
+
+lemma fun_poly_body_unique:
+ assumes "Some (AF_fundef f (AF_fun_typ_some bv1 (AF_fun_typ x1 b1 c1 \<tau>1' s1'))) = lookup_fun \<Phi> f" and "Some (AF_fundef f (AF_fun_typ_some bv2 (AF_fun_typ x2 b2 c2 \<tau>2' s2'))) = lookup_fun \<Phi> f"
+ shows "s1'[bv1::=b]\<^sub>s\<^sub>b[x1::=v]\<^sub>s\<^sub>v = s2'[bv2::=b]\<^sub>s\<^sub>b[x2::=v]\<^sub>s\<^sub>v"
+proof -
+ have *: " (AF_fundef f (AF_fun_typ_some bv1 (AF_fun_typ x1 b1 c1 \<tau>1' s1'))) = (AF_fundef f (AF_fun_typ_some bv2 (AF_fun_typ x2 b2 c2 \<tau>2' s2')))"
+ using option.inject assms by metis
+ hence "AF_fun_typ_some bv1 (AF_fun_typ x1 b1 c1 \<tau>1' s1') = AF_fun_typ_some bv2 (AF_fun_typ x2 b2 c2 \<tau>2' s2')"
+ (is "AF_fun_typ_some bv1 ?ft1 = AF_fun_typ_some bv2 ?ft2") using fun_def.eq_iff by metis
+ hence **:"[[atom bv1]]lst. ?ft1 = [[atom bv2]]lst. ?ft2" using fun_typ_q.eq_iff(1) by metis
+
+ hence *:"subst_ft_b ?ft1 bv1 b = subst_ft_b ?ft2 bv2 b" using subst_b_flip_eq_two subst_b_fun_typ_def by metis
+ have "[[atom x1]]lst. s1'[bv1::=b]\<^sub>s\<^sub>b = [[atom x2]]lst. s2'[bv2::=b]\<^sub>s\<^sub>b"
+ using lst_snd lst_fst subst_ft_b.simps fun_typ.eq_iff
+ by (metis "local.*")
+
+ thus ?thesis using subst_v_flip_eq_two subst_v_s_def by metis
+qed
+
+lemma funtyp_eq_iff_equalities:
+ fixes s'::s and s::s
+ assumes " [[atom x']]lst. ((c', \<tau>'), s') = [[atom x]]lst. ((c, \<tau>), s)"
+ shows "\<lbrace> x' : b | c' \<rbrace> = \<lbrace> x : b | c \<rbrace> \<and> s'[x'::=v]\<^sub>s\<^sub>v = s[x::=v]\<^sub>s\<^sub>v \<and> \<tau>'[x'::=v]\<^sub>\<tau>\<^sub>v = \<tau>[x::=v]\<^sub>\<tau>\<^sub>v"
+proof -
+ have "[[atom x']]lst. s' = [[atom x]]lst. s" and "[[atom x']]lst. \<tau>' = [[atom x]]lst. \<tau>" and
+ " [[atom x']]lst. c' = [[atom x]]lst. c" using lst_snd lst_fst assms by metis+
+ thus ?thesis using subst_v_flip_eq_two \<tau>.eq_iff
+ by (metis assms fun_typ.eq_iff fun_typ_eq_body_unique fun_typ_eq_ret_unique)
+qed
+
+section \<open>Weakening\<close>
+
+lemma wfX_wfB1:
+ fixes \<Gamma>::\<Gamma> and \<Gamma>'::\<Gamma> and v::v and e::e and c::c and \<tau>::\<tau> and ts::"(string*\<tau>) list" and \<Delta>::\<Delta> and s::s and b::b and \<B>::\<B> and \<Phi>::\<Phi> and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def
+ and cs::branch_s and css::branch_list
+ shows wfV_wfB: "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b" and
+ "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f c \<Longrightarrow> True" and
+ "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<Longrightarrow> True" and
+ wfT_wfB: "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau> \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b_of \<tau> " and
+ "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ts \<Longrightarrow> True" and
+ "\<turnstile>\<^sub>w\<^sub>f \<Theta> \<Longrightarrow> True" and
+ "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b \<Longrightarrow> True" and
+ wfCE_wfB: "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ce : b \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b" and
+ "\<Theta> \<turnstile>\<^sub>w\<^sub>f td \<Longrightarrow> True"
+proof(induct rule:wfV_wfC_wfG_wfT_wfTs_wfTh_wfB_wfCE_wfTD.inducts)
+ case (wfV_varI \<Theta> \<B> \<Gamma> b c x)
+ hence "(x,b,c) \<in> toSet \<Gamma>" using lookup_iff wfV_wf using lookup_in_g by presburger
+ hence "b \<in> fst`snd`toSet \<Gamma>" by force
+ hence "wfB \<Theta> \<B> b" using wfG_wfB wfV_varI by metis
+ then show ?case using wfV_elims wfG_wf wf_intros by metis
+next
+ case (wfV_litI \<Theta> \<Gamma> l)
+ moreover have "wfTh \<Theta>" using wfV_wf wfG_wf wfV_litI by metis
+ ultimately show ?case using wfV_wf wfG_wf wf_intros base_for_lit.simps l.exhaust by metis
+next
+ case (wfV_pairI \<Theta> \<Gamma> v1 b1 v2 b2)
+ then show ?case using wfG_wf wf_intros by metis
+next
+ case (wfV_consI s dclist \<Theta> dc x b c B \<Gamma> v)
+ then show ?case
+ using wfV_wf wfG_wf wfB_consI by metis
+next
+ case (wfV_conspI s bv dclist \<Theta> dc x b' c \<B> b \<Gamma> v)
+ then show ?case
+ using wfV_wf wfG_wf using wfB_appI by metis
+next
+ case (wfCE_valI \<Theta> \<B> \<Gamma> v b)
+ then show ?case using wfB_elims by auto
+next
+ case (wfCE_plusI \<Theta> \<B> \<Gamma> v1 v2)
+ then show ?case using wfB_elims by auto
+next
+ case (wfCE_leqI \<Theta> \<B> \<Gamma> v1 v2)
+ then show ?case using wfV_wf wfG_wf wf_intros wfX_wfY by metis
+next
+ case (wfCE_eqI \<Theta> \<B> \<Gamma> v1 b v2)
+ then show ?case using wfV_wf wfG_wf wf_intros wfX_wfY by metis
+next
+ case (wfCE_fstI \<Theta> \<B> \<Gamma> v1 b1 b2)
+ then show ?case using wfB_elims by metis
+next
+ case (wfCE_sndI \<Theta> \<B> \<Gamma> v1 b1 b2)
+ then show ?case using wfB_elims by metis
+next
+ case (wfCE_concatI \<Theta> \<B> \<Gamma> v1 v2)
+ then show ?case using wfB_elims by auto
+next
+ case (wfCE_lenI \<Theta> \<B> \<Gamma> v1)
+ then show ?case using wfV_wf wfG_wf wf_intros wfX_wfY by metis
+qed(auto | metis wfV_wf wfG_wf wf_intros )+
+
+lemma wfX_wfB2:
+ fixes \<Gamma>::\<Gamma> and \<Gamma>'::\<Gamma> and v::v and e::e and c::c and \<tau>::\<tau> and ts::"(string*\<tau>) list" and \<Delta>::\<Delta> and s::s and b::b and \<B>::\<B> and \<Phi>::\<Phi> and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def
+ and cs::branch_s and css::branch_list
+ shows
+ wfE_wfB: "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f e : b \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b" and
+ wfS_wfB: "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b" and
+ wfCS_wfB: "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> ; tid ; dc ; t \<turnstile>\<^sub>w\<^sub>f cs : b \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b" and
+ wfCSS_wfB: "\<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta> ; tid ; dclist \<turnstile>\<^sub>w\<^sub>f css : b \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b" and
+ "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<Longrightarrow> True" and
+ "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<Longrightarrow> True" and
+ "\<Theta> ; \<Phi> \<turnstile>\<^sub>w\<^sub>f ftq \<Longrightarrow> True" and
+ "\<Theta> ; \<Phi> ; \<B> \<turnstile>\<^sub>w\<^sub>f ft \<Longrightarrow> \<B> |\<subseteq>| \<B>' \<Longrightarrow> \<Theta> ; \<Phi> ; \<B>' \<turnstile>\<^sub>w\<^sub>f ft"
+proof(induct rule:wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.inducts)
+ case (wfE_valI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v b)
+ then show ?case using wfB_elims wfX_wfB1 by metis
+next
+ case (wfE_plusI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 v2)
+ then show ?case using wfB_elims wfX_wfB1 by metis
+next
+ case (wfE_eqI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 b v2)
+ then show ?case using wfB_boolI wfX_wfY by metis
+next
+ case (wfE_fstI \<Theta> \<Phi> \<Gamma> \<Delta> v1 b1 b2)
+ then show ?case using wfB_elims wfX_wfB1 by metis
+next
+ case (wfE_sndI \<Theta> \<Phi> \<Gamma> \<Delta> v1 b1 b2)
+ then show ?case using wfB_elims wfX_wfB1 by metis
+next
+ case (wfE_concatI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 v2)
+ then show ?case using wfB_elims wfX_wfB1 by metis
+next
+ case (wfE_splitI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 v2)
+ then show ?case using wfB_elims wfX_wfB1
+ using wfB_pairI by auto
+next
+ case (wfE_lenI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1)
+ then show ?case using wfB_elims wfX_wfB1
+ using wfB_intI wfX_wfY1(1) by auto
+next
+ case (wfE_appI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> f x b c \<tau> s v)
+ hence "\<Theta>; \<B>;(x,b,c) #\<^sub>\<Gamma> GNil \<turnstile>\<^sub>w\<^sub>f \<tau>" using wfPhi_f_simple_wfT wfT_b_weakening by fast
+ then show ?case using b_of.simps using wfT_b_weakening
+ by (metis b_of.cases bot.extremum wfT_elims(2))
+next
+ case (wfE_appPI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> b' bv v \<tau> f x b c s)
+ hence "\<Theta> ; {| bv |} ;(x,b,c) #\<^sub>\<Gamma> GNil \<turnstile>\<^sub>w\<^sub>f \<tau>" using wfPhi_f_poly_wfT wfX_wfY by blast
+ then show ?case using wfE_appPI b_of.simps using wfT_b_weakening wfT_elims wfT_subst_wfB subst_b_b_def by metis
+next
+ case (wfE_mvarI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> u \<tau>)
+ hence "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau>" using wfD_wfT by fast
+ then show ?case using wfT_elims b_of.simps by metis
+next
+ case (wfFTNone \<Theta> ft)
+ then show ?case by auto
+next
+ case (wfFTSome \<Theta> bv ft)
+ then show ?case by auto
+next
+ case (wfS_valI \<Theta> \<Phi> \<B> \<Gamma> v b \<Delta>)
+ then show ?case using wfX_wfB1 by auto
+next
+ case (wfS_letI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> e b' x s b)
+ then show ?case using wfX_wfB1 by auto
+next
+ case (wfS_let2I \<Theta> \<Phi> \<B> \<Gamma> \<Delta> s1 \<tau> x s2 b)
+ then show ?case using wfX_wfB1 by auto
+next
+ case (wfS_ifI \<Theta> \<B> \<Gamma> v \<Phi> \<Delta> s1 b s2)
+ then show ?case using wfX_wfB1 by auto
+next
+ case (wfS_varI \<Theta> \<B> \<Gamma> \<tau> v u \<Phi> \<Delta> b s)
+ then show ?case using wfX_wfB1 by auto
+next
+ case (wfS_assignI u \<tau> \<Delta> \<Theta> \<B> \<Gamma> \<Phi> v)
+ then show ?case using wfX_wfB1
+ using wfB_unitI wfX_wfY2(5) by auto
+next
+ case (wfS_whileI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> s1 s2 b)
+ then show ?case using wfX_wfB1 by auto
+next
+ case (wfS_seqI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> s1 s2 b)
+ then show ?case using wfX_wfB1 by auto
+next
+ case (wfS_matchI \<Theta> \<B> \<Gamma> v tid dclist \<Delta> \<Phi> cs b)
+ then show ?case using wfX_wfB1 by auto
+next
+ case (wfS_branchI \<Theta> \<Phi> \<B> x \<tau> \<Gamma> \<Delta> s b tid dc)
+ then show ?case using wfX_wfB1 by auto
+next
+ case (wfS_finalI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dc t cs b)
+ then show ?case using wfX_wfB1 by auto
+next
+ case (wfS_cons \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dc t cs b dclist css)
+ then show ?case using wfX_wfB1 by auto
+next
+ case (wfD_emptyI \<Theta> \<B> \<Gamma>)
+ then show ?case using wfX_wfB1 by auto
+next
+ case (wfD_cons \<Theta> \<B> \<Gamma> \<Delta> \<tau> u)
+ then show ?case using wfX_wfB1 by auto
+next
+ case (wfPhi_emptyI \<Theta>)
+ then show ?case using wfX_wfB1 by auto
+next
+ case (wfPhi_consI f \<Theta> \<Phi> ft)
+ then show ?case using wfX_wfB1 by auto
+next
+ case (wfFTI \<Theta> B b \<Phi> x c s \<tau>)
+ then show ?case using wfX_wfB1
+ by (meson Wellformed.wfFTI wb_b_weakening2(8))
+qed(metis wfV_wf wfG_wf wf_intros wfX_wfB1)
+
+lemmas wfX_wfB = wfX_wfB1 wfX_wfB2
+
+lemma wf_weakening1:
+ fixes \<Gamma>::\<Gamma> and \<Gamma>'::\<Gamma> and v::v and e::e and c::c and \<tau>::\<tau> and ts::"(string*\<tau>) list" and \<Delta>::\<Delta> and s::s and \<B>::\<B> and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def
+ and cs::branch_s and css::branch_list
+
+ shows wfV_weakening: "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>' \<Longrightarrow> toSet \<Gamma> \<subseteq> toSet \<Gamma>' \<Longrightarrow> \<Theta>; \<B>; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f v : b" and
+ wfC_weakening: "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f c \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>' \<Longrightarrow> toSet \<Gamma> \<subseteq> toSet \<Gamma>' \<Longrightarrow> \<Theta>; \<B>; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f c" and
+ "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<Longrightarrow> True" and
+ wfT_weakening: "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau> \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>' \<Longrightarrow> toSet \<Gamma> \<subseteq> toSet \<Gamma>' \<Longrightarrow> \<Theta>; \<B>; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f \<tau>" and
+ "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ts \<Longrightarrow> True" and
+ "\<turnstile>\<^sub>w\<^sub>f P \<Longrightarrow> True " and
+ wfB_weakening: "wfB \<Theta> \<B> b \<Longrightarrow> \<B> |\<subseteq>| \<B>' \<Longrightarrow> wfB \<Theta> \<B> b" and
+ wfCE_weakening: "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ce : b \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>' \<Longrightarrow> toSet \<Gamma> \<subseteq> toSet \<Gamma>' \<Longrightarrow> \<Theta>; \<B>; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f ce : b" and
+ "\<Theta> \<turnstile>\<^sub>w\<^sub>f td \<Longrightarrow> True"
+proof(nominal_induct
+ b and c and \<Gamma> and \<tau> and ts and P and b and b and td
+ avoiding: \<Gamma>'
+ rule:wfV_wfC_wfG_wfT_wfTs_wfTh_wfB_wfCE_wfTD.strong_induct)
+ case (wfV_varI \<Theta> \<B> \<Gamma> b c x)
+ hence "Some (b, c) = lookup \<Gamma>' x" using lookup_weakening by metis
+ then show ?case using Wellformed.wfV_varI wfV_varI by metis
+next
+ case (wfTI z \<Theta> \<B> \<Gamma> b c) (* This proof pattern is used elsewhere when proving weakening for typing predicates *)
+ show ?case proof
+ show \<open>atom z \<sharp> (\<Theta>, \<B>, \<Gamma>')\<close> using wfTI by auto
+ show \<open> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b \<close> using wfTI by auto
+ have *:"toSet ((z, b, TRUE) #\<^sub>\<Gamma> \<Gamma>) \<subseteq> toSet ((z, b, TRUE) #\<^sub>\<Gamma> \<Gamma>')" using toSet.simps wfTI by auto
+ thus \<open> \<Theta>; \<B>; (z, b, TRUE) #\<^sub>\<Gamma> \<Gamma>' \<turnstile>\<^sub>w\<^sub>f c \<close> using wfTI(8)[OF _ *] wfTI wfX_wfY
+ by (simp add: wfG_cons_TRUE)
+ qed
+next
+ case (wfV_conspI s bv dclist \<Theta> dc x b' c \<B> b \<Gamma> v)
+ show ?case proof
+ show \<open>AF_typedef_poly s bv dclist \<in> set \<Theta>\<close> using wfV_conspI by auto
+ show \<open>(dc, \<lbrace> x : b' | c \<rbrace>) \<in> set dclist\<close> using wfV_conspI by auto
+ show \<open> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b \<close> using wfV_conspI by auto
+ show \<open>atom bv \<sharp> (\<Theta>, \<B>, \<Gamma>', b, v)\<close> using wfV_conspI by simp
+ show \<open> \<Theta>; \<B>; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f v : b'[bv::=b]\<^sub>b\<^sub>b \<close> using wfV_conspI by auto
+ qed
+qed(metis wf_intros)+
+
+lemma wf_weakening2:
+ fixes \<Gamma>::\<Gamma> and \<Gamma>'::\<Gamma> and v::v and e::e and c::c and \<tau>::\<tau> and ts::"(string*\<tau>) list" and \<Delta>::\<Delta> and s::s and \<B>::\<B> and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def
+ and cs::branch_s and css::branch_list
+ shows
+ wfE_weakening: "\<Theta>; \<Phi>; \<B>; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f e : b \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>' \<Longrightarrow> toSet \<Gamma> \<subseteq> toSet \<Gamma>' \<Longrightarrow> \<Theta>; \<Phi>; \<B>; \<Gamma>' ; \<Delta> \<turnstile>\<^sub>w\<^sub>f e : b" and
+ wfS_weakening: "\<Theta>; \<Phi>; \<B>; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>' \<Longrightarrow> toSet \<Gamma> \<subseteq> toSet \<Gamma>' \<Longrightarrow> \<Theta>; \<Phi>; \<B>; \<Gamma>' ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b" and
+ "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> ; tid ; dc ; t \<turnstile>\<^sub>w\<^sub>f cs : b \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>' \<Longrightarrow> toSet \<Gamma> \<subseteq> toSet \<Gamma>' \<Longrightarrow> \<Theta>; \<Phi>; \<B>; \<Gamma>' ; \<Delta> ; tid ; dc ; t \<turnstile>\<^sub>w\<^sub>f cs : b" and
+ "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> ; tid ; dclist \<turnstile>\<^sub>w\<^sub>f css : b \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>' \<Longrightarrow> toSet \<Gamma> \<subseteq> toSet \<Gamma>' \<Longrightarrow> \<Theta>; \<Phi>; \<B>; \<Gamma>' ; \<Delta> ; tid ; dclist \<turnstile>\<^sub>w\<^sub>f css : b" and
+ "\<Theta> \<turnstile>\<^sub>w\<^sub>f (\<Phi>::\<Phi>) \<Longrightarrow> True" and
+ wfD_weakning: "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>' \<Longrightarrow> toSet \<Gamma> \<subseteq> toSet \<Gamma>' \<Longrightarrow> \<Theta>; \<B>; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f \<Delta>" and
+ "\<Theta> ; \<Phi> \<turnstile>\<^sub>w\<^sub>f ftq \<Longrightarrow> True" and
+ "\<Theta> ; \<Phi> ; \<B> \<turnstile>\<^sub>w\<^sub>f ft \<Longrightarrow> True"
+proof(nominal_induct
+ b and b and b and b and \<Phi> and \<Delta> and ftq and ft
+ avoiding: \<Gamma>'
+ rule:wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.strong_induct)
+ case (wfE_appPI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> b' bv v \<tau> f x b c s)
+ show ?case proof
+ show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using wfE_appPI by auto
+ show \<open> \<Theta>; \<B>; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wfE_appPI by auto
+ show \<open> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b' \<close> using wfE_appPI by auto
+ show \<open>atom bv \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>', \<Delta>, b', v, (b_of \<tau>)[bv::=b']\<^sub>b)\<close> using wfE_appPI by auto
+ show \<open>Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \<tau> s))) = lookup_fun \<Phi> f\<close> using wfE_appPI by auto
+ show \<open> \<Theta>; \<B>; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f v : b[bv::=b']\<^sub>b \<close> using wfE_appPI wf_weakening1 by auto
+ qed
+next
+ case (wfS_letI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> e b' x s b)
+ show ?case proof(rule)
+ show \<open> \<Theta> ; \<Phi> ; \<B> ; \<Gamma>' ; \<Delta> \<turnstile>\<^sub>w\<^sub>f e : b' \<close> using wfS_letI by auto
+ have "toSet ((x, b', TRUE) #\<^sub>\<Gamma> \<Gamma>) \<subseteq> toSet ((x, b', TRUE) #\<^sub>\<Gamma> \<Gamma>')" using wfS_letI by auto
+ thus \<open> \<Theta> ; \<Phi> ; \<B> ; (x, b', TRUE) #\<^sub>\<Gamma> \<Gamma>' ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b \<close> using wfS_letI by (meson wfG_cons wfG_cons_TRUE wfS_wf)
+ show \<open> \<Theta>; \<B>; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wfS_letI by auto
+ show \<open>atom x \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>', \<Delta>, e, b)\<close> using wfS_letI by auto
+ qed
+next
+ case (wfS_let2I \<Theta> \<Phi> \<B> \<Gamma> \<Delta> s1 \<tau> x s2 b)
+ show ?case proof
+ show \<open> \<Theta> ; \<Phi> ; \<B> ; \<Gamma>' ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s1 : b_of \<tau> \<close> using wfS_let2I by auto
+ show \<open> \<Theta>; \<B>; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f \<tau> \<close> using wfS_let2I wf_weakening1 by auto
+ have "toSet ((x, b_of \<tau>, TRUE) #\<^sub>\<Gamma> \<Gamma>) \<subseteq> toSet ((x, b_of \<tau>, TRUE) #\<^sub>\<Gamma> \<Gamma>')" using wfS_let2I by auto
+ thus \<open> \<Theta> ; \<Phi> ; \<B> ; (x, b_of \<tau>, TRUE) #\<^sub>\<Gamma> \<Gamma>' ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s2 : b \<close> using wfS_let2I by (meson wfG_cons wfG_cons_TRUE wfS_wf)
+ show \<open>atom x \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>', \<Delta>, s1, b, \<tau>)\<close> using wfS_let2I by auto
+ qed
+next
+ case (wfS_varI \<Theta> \<B> \<Gamma> \<tau> v u \<Phi> \<Delta> b s)
+ show ?case proof
+ show "\<Theta>; \<B>; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f \<tau> " using wfS_varI wf_weakening1 by auto
+ show "\<Theta>; \<B>; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f v : b_of \<tau> " using wfS_varI wf_weakening1 by auto
+ show "atom u \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>', \<Delta>, \<tau>, v, b)" using wfS_varI by auto
+ show "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>' ; (u, \<tau>) #\<^sub>\<Delta> \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b " using wfS_varI by auto
+ qed
+next
+ case (wfS_branchI \<Theta> \<Phi> \<B> x \<tau> \<Gamma> \<Delta> s b tid dc)
+ show ?case proof
+ have "toSet ((x, b_of \<tau>, TRUE) #\<^sub>\<Gamma> \<Gamma>) \<subseteq> toSet ((x, b_of \<tau>, TRUE) #\<^sub>\<Gamma> \<Gamma>')" using wfS_branchI by auto
+ thus \<open> \<Theta> ; \<Phi> ; \<B> ; (x, b_of \<tau>, TRUE) #\<^sub>\<Gamma> \<Gamma>' ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b \<close> using wfS_branchI by (meson wfG_cons wfG_cons_TRUE wfS_wf)
+ show \<open>atom x \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>', \<Delta>, \<Gamma>', \<tau>)\<close> using wfS_branchI by auto
+ show \<open> \<Theta>; \<B>; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wfS_branchI by auto
+ qed
+next
+ case (wfS_finalI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist' cs b dclist)
+ then show ?case using wf_intros by metis
+next
+ case (wfS_cons \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist' cs b css dclist)
+ then show ?case using wf_intros by metis
+next
+ case (wfS_assertI \<Theta> \<Phi> \<B> x c \<Gamma> \<Delta> s b)
+ show ?case proof(rule)
+ show \<open> \<Theta>; \<B>; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f c \<close> using wfS_assertI wf_weakening1 by auto
+ have "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (x, B_bool, c) #\<^sub>\<Gamma> \<Gamma>'" proof(rule wfG_consI)
+ show \<open> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>' \<close> using wfS_assertI by auto
+ show \<open>atom x \<sharp> \<Gamma>'\<close> using wfS_assertI by auto
+ show \<open> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f B_bool \<close> using wfS_assertI wfB_boolI wfX_wfY by metis
+ have "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (x, B_bool, TRUE) #\<^sub>\<Gamma> \<Gamma>'" proof
+ show "(TRUE) \<in> {TRUE, FALSE}" by auto
+ show \<open> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>' \<close> using wfS_assertI by auto
+ show \<open>atom x \<sharp> \<Gamma>'\<close> using wfS_assertI by auto
+ show \<open> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f B_bool \<close> using wfS_assertI wfB_boolI wfX_wfY by metis
+ qed
+ thus \<open> \<Theta>; \<B>; (x, B_bool, TRUE) #\<^sub>\<Gamma> \<Gamma>' \<turnstile>\<^sub>w\<^sub>f c \<close>
+ using wf_weakening1(2)[OF \<open> \<Theta>; \<B>; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f c \<close> \<open> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (x, B_bool, TRUE) #\<^sub>\<Gamma> \<Gamma>' \<close>] by force
+ qed
+ thus \<open> \<Theta>; \<Phi>; \<B>; (x, B_bool, c) #\<^sub>\<Gamma> \<Gamma>' ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b \<close> using wfS_assertI by fastforce
+ show \<open> \<Theta>; \<B>; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wfS_assertI by auto
+ show \<open>atom x \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>', \<Delta>, c, b, s)\<close> using wfS_assertI by auto
+ qed
+qed(metis wf_intros wf_weakening1)+
+
+lemmas wf_weakening = wf_weakening1 wf_weakening2
+
+lemma wfV_weakening_cons:
+ fixes \<Gamma>::\<Gamma> and \<Gamma>'::\<Gamma> and v::v and c::c
+ assumes "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b" and "atom y \<sharp> \<Gamma>" and "\<Theta>; \<B>; ((y,b',TRUE) #\<^sub>\<Gamma> \<Gamma>) \<turnstile>\<^sub>w\<^sub>f c"
+ shows "\<Theta>; \<B>; (y,b',c) #\<^sub>\<Gamma>\<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b"
+proof -
+ have "wfG \<Theta> \<B> ((y,b',c) #\<^sub>\<Gamma>\<Gamma>)" using wfG_intros2 assms by auto
+ moreover have "toSet \<Gamma> \<subseteq> toSet ((y,b',c) #\<^sub>\<Gamma>\<Gamma>)" using toSet.simps by auto
+ ultimately show ?thesis using wf_weakening using assms(1) by blast
+qed
+
+lemma wfG_cons_weakening:
+ fixes \<Gamma>'::\<Gamma>
+ assumes "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f ((x, b, c) #\<^sub>\<Gamma> \<Gamma>)" and "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>'" and "toSet \<Gamma> \<subseteq> toSet \<Gamma>'" and "atom x \<sharp> \<Gamma>'"
+ shows "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f ((x, b, c) #\<^sub>\<Gamma> \<Gamma>')"
+proof(cases "c \<in> {TRUE,FALSE}")
+ case True
+ then show ?thesis using wfG_wfB wfG_cons2I assms by auto
+next
+ case False
+ hence *:"\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<and> atom x \<sharp> \<Gamma> \<and> \<Theta>; \<B>; (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f c"
+ using wfG_elims(2)[OF assms(1)] by auto
+ have a1:"\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>'" using wfG_wfB wfG_cons2I assms by simp
+ moreover have a2:"toSet ((x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> ) \<subseteq> toSet ((x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>')" using toSet.simps assms by blast
+ moreover have " \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>'" proof
+ show "(TRUE) \<in> {TRUE, FALSE}" by auto
+ show "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>'" using assms by auto
+ show "atom x \<sharp> \<Gamma>'" using assms by auto
+ show "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b" using assms wfG_elims by metis
+ qed
+ hence " \<Theta>; \<B>; (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>' \<turnstile>\<^sub>w\<^sub>f c" using wf_weakening a1 a2 * by auto
+ then show ?thesis using wfG_cons1I[of c \<Theta> \<B> \<Gamma>' x b, OF False ] wfG_wfB assms by simp
+qed
+
+lemma wfT_weakening_aux:
+ fixes \<Gamma>::\<Gamma> and \<Gamma>'::\<Gamma> and c::c
+ assumes "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b | c \<rbrace>" and "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>'" and "toSet \<Gamma> \<subseteq> toSet \<Gamma>'" and "atom z \<sharp> \<Gamma>'"
+ shows "\<Theta>; \<B>; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b | c \<rbrace>"
+proof
+ show \<open>atom z \<sharp> (\<Theta>, \<B>, \<Gamma>')\<close>
+ using wf_supp wfX_wfY assms fresh_prodN fresh_def x_not_in_b_set wfG_fresh_x by metis
+ show \<open> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b \<close> using assms wfT_elims by metis
+ show \<open> \<Theta>; \<B>; (z, b, TRUE) #\<^sub>\<Gamma> \<Gamma>' \<turnstile>\<^sub>w\<^sub>f c \<close> proof -
+ have *:"\<Theta>; \<B>; (z,b,TRUE) #\<^sub>\<Gamma>\<Gamma> \<turnstile>\<^sub>w\<^sub>f c" using wfT_wfC fresh_weakening assms by auto
+ moreover have a1:"\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (z, b, TRUE) #\<^sub>\<Gamma> \<Gamma>'" using wfG_cons2I assms \<open>\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b\<close> by simp
+ moreover have a2:"toSet ((z, b, TRUE) #\<^sub>\<Gamma> \<Gamma> ) \<subseteq> toSet ((z, b, TRUE) #\<^sub>\<Gamma> \<Gamma>')" using toSet.simps assms by blast
+ moreover have " \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (z, b, TRUE) #\<^sub>\<Gamma> \<Gamma>' " proof
+ show "(TRUE) \<in> {TRUE, FALSE}" by auto
+ show "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>'" using assms by auto
+ show "atom z \<sharp> \<Gamma>'" using assms by auto
+ show "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b" using assms wfT_elims by metis
+ qed
+ thus ?thesis using wf_weakening a1 a2 * by auto
+ qed
+qed
+
+lemma wfT_weakening_all:
+ fixes \<Gamma>::\<Gamma> and \<Gamma>'::\<Gamma> and \<tau>::\<tau>
+ assumes "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau>" and "\<Theta>; \<B>' \<turnstile>\<^sub>w\<^sub>f \<Gamma>'" and "toSet \<Gamma> \<subseteq> toSet \<Gamma>'" and "\<B> |\<subseteq>| \<B>'"
+ shows "\<Theta>; \<B>' ; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f \<tau>"
+ using wb_b_weakening assms wfT_weakening by metis
+
+lemma wfT_weakening_nil:
+ fixes \<Gamma>::\<Gamma> and \<Gamma>'::\<Gamma> and \<tau>::\<tau>
+ assumes "\<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f \<tau>" and "\<Theta>; \<B>' \<turnstile>\<^sub>w\<^sub>f \<Gamma>'"
+ shows "\<Theta>; \<B>' ; \<Gamma>' \<turnstile>\<^sub>w\<^sub>f \<tau>"
+ using wfT_weakening_all
+ using assms(1) assms(2) toSet.simps(1) by blast
+
+lemma wfTh_wfT2:
+ fixes x::x and v::v and \<tau>::\<tau> and G::\<Gamma>
+ assumes "wfTh \<Theta>" and "AF_typedef s dclist \<in> set \<Theta>" and
+ "(dc, \<tau>) \<in> set dclist" and "\<Theta> ; B \<turnstile>\<^sub>w\<^sub>f G"
+ shows "supp \<tau> = {}" and "\<tau>[x::=v]\<^sub>\<tau>\<^sub>v = \<tau>" and "wfT \<Theta> B G \<tau>"
+proof -
+ show "supp \<tau> = {}" proof(rule ccontr)
+ assume a1: "supp \<tau> \<noteq> {}"
+ have "supp \<Theta> \<noteq> {}" proof -
+ obtain dclist where dc: "AF_typedef s dclist \<in> set \<Theta> \<and> (dc, \<tau>) \<in> set dclist"
+ using assms by auto
+ hence "supp (dc,\<tau>) \<noteq> {}"
+ using a1 by (simp add: supp_Pair)
+ hence "supp dclist \<noteq> {}" using dc supp_list_member by auto
+ hence "supp (AF_typedef s dclist) \<noteq> {}" using type_def.supp by auto
+ thus ?thesis using supp_list_member dc by auto
+ qed
+ thus False using assms wfTh_supp by simp
+ qed
+ thus "\<tau>[x::=v]\<^sub>\<tau>\<^sub>v = \<tau>" by (simp add: fresh_def)
+ have "wfT \<Theta> {||} GNil \<tau>" using assms wfTh_wfT by auto
+ thus "wfT \<Theta> B G \<tau>" using assms wfT_weakening_nil by simp
+qed
+
+lemma wf_d_weakening:
+ fixes \<Gamma>::\<Gamma> and \<Gamma>'::\<Gamma> and v::v and e::e and c::c and \<tau>::\<tau> and ts::"(string*\<tau>) list" and \<Delta>::\<Delta> and s::s and \<B>::\<B> and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def
+ and cs::branch_s and css::branch_list
+ shows
+ "\<Theta>; \<Phi>; \<B>; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f e : b \<Longrightarrow> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>' \<Longrightarrow> setD \<Delta> \<subseteq> setD \<Delta>' \<Longrightarrow> \<Theta>; \<Phi>; \<B>; \<Gamma> ; \<Delta>' \<turnstile>\<^sub>w\<^sub>f e : b" and
+ "\<Theta>; \<Phi>; \<B>; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b \<Longrightarrow> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>' \<Longrightarrow> setD \<Delta> \<subseteq> setD \<Delta>' \<Longrightarrow> \<Theta>; \<Phi>; \<B>; \<Gamma> ; \<Delta>' \<turnstile>\<^sub>w\<^sub>f s : b" and
+ "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> ; tid ; dc ; t \<turnstile>\<^sub>w\<^sub>f cs : b \<Longrightarrow> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>' \<Longrightarrow> setD \<Delta> \<subseteq> setD \<Delta>' \<Longrightarrow> \<Theta>; \<Phi>; \<B>; \<Gamma> ; \<Delta>' ; tid ; dc ; t \<turnstile>\<^sub>w\<^sub>f cs : b" and
+ "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> ; tid ; dclist \<turnstile>\<^sub>w\<^sub>f css : b \<Longrightarrow> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>' \<Longrightarrow> setD \<Delta> \<subseteq> setD \<Delta>' \<Longrightarrow> \<Theta>; \<Phi>; \<B>; \<Gamma> ; \<Delta>' ; tid ; dclist \<turnstile>\<^sub>w\<^sub>f css : b" and
+ "\<Theta> \<turnstile>\<^sub>w\<^sub>f (\<Phi>::\<Phi>) \<Longrightarrow> True" and
+ "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<Longrightarrow> True" and
+ "\<Theta> ; \<Phi> \<turnstile>\<^sub>w\<^sub>f ftq \<Longrightarrow> True" and
+ "\<Theta> ; \<Phi> ; \<B> \<turnstile>\<^sub>w\<^sub>f ft \<Longrightarrow> True"
+proof(nominal_induct
+ b and b and b and b and \<Phi> and \<Delta> and ftq and ft
+ avoiding: \<Delta>'
+ rule:wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.strong_induct)
+ case (wfE_valI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v b)
+ then show ?case using wf_intros by metis
+next
+ case (wfE_plusI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 v2)
+ then show ?case using wf_intros by metis
+next
+ case (wfE_leqI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 v2)
+ then show ?case using wf_intros by metis
+next
+ case (wfE_eqI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 b v2)
+ then show ?case using wf_intros by metis
+next
+ case (wfE_fstI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 b1 b2)
+ then show ?case using wf_intros by metis
+next
+ case (wfE_sndI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 b1 b2)
+ then show ?case using wf_intros by metis
+next
+ case (wfE_concatI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 v2)
+ then show ?case using wf_intros by metis
+next
+ case (wfE_splitI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 v2)
+ then show ?case using wf_intros by metis
+next
+ case (wfE_lenI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1)
+ then show ?case using wf_intros by metis
+next
+ case (wfE_appI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> f x b c \<tau> s v)
+ then show ?case using wf_intros by metis
+next
+ case (wfE_appPI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> b' bv v \<tau> f x b c s)
+ show ?case proof(rule, (rule wfE_appPI)+)
+ show \<open>atom bv \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>, \<Delta>', b', v, (b_of \<tau>)[bv::=b']\<^sub>b)\<close> using wfE_appPI by auto
+ show \<open>Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \<tau> s))) = lookup_fun \<Phi> f\<close> using wfE_appPI by auto
+ show \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b[bv::=b']\<^sub>b \<close> using wfE_appPI by auto
+ qed
+next
+ case (wfE_mvarI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> u \<tau>)
+ show ?case proof
+ show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using wfE_mvarI by auto
+ show \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>' \<close> using wfE_mvarI by auto
+ show \<open>(u, \<tau>) \<in> setD \<Delta>'\<close> using wfE_mvarI by auto
+ qed
+next
+ case (wfS_valI \<Theta> \<Phi> \<B> \<Gamma> v b \<Delta>)
+ then show ?case using wf_intros by metis
+next
+ case (wfS_letI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> e b' x s b)
+ show ?case proof(rule)
+ show \<open> \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta>' \<turnstile>\<^sub>w\<^sub>f e : b' \<close> using wfS_letI by auto
+ have "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (x, b', TRUE) #\<^sub>\<Gamma> \<Gamma>" using wfG_cons2I wfX_wfY wfS_letI by metis
+ hence "\<Theta>; \<B>; (x, b', TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>'" using wf_weakening2(6) wfS_letI by force
+ thus \<open> \<Theta> ; \<Phi> ; \<B> ; (x, b', TRUE) #\<^sub>\<Gamma> \<Gamma> ; \<Delta>' \<turnstile>\<^sub>w\<^sub>f s : b \<close> using wfS_letI by metis
+ show \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>' \<close> using wfS_letI by auto
+ show \<open>atom x \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>, \<Delta>', e, b)\<close> using wfS_letI by auto
+ qed
+next
+ case (wfS_assertI \<Theta> \<Phi> \<B> x c \<Gamma> \<Delta> s b)
+ show ?case proof
+ have "\<Theta>; \<B>; (x, B_bool, c) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>'" proof(rule wf_weakening2(6))
+ show \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>' \<close> using wfS_assertI by auto
+ next
+ show \<open> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (x, B_bool, c) #\<^sub>\<Gamma> \<Gamma> \<close> using wfS_assertI wfX_wfY by metis
+ next
+ show \<open>toSet \<Gamma> \<subseteq> toSet ((x, B_bool, c) #\<^sub>\<Gamma> \<Gamma>)\<close> using wfS_assertI by auto
+ qed
+ thus \<open> \<Theta>; \<Phi>; \<B>; (x, B_bool, c) #\<^sub>\<Gamma> \<Gamma> ; \<Delta>' \<turnstile>\<^sub>w\<^sub>f s : b \<close> using wfS_assertI wfX_wfY by metis
+ next
+ show \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f c \<close> using wfS_assertI by auto
+ next
+ show \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>' \<close> using wfS_assertI by auto
+ next
+ show \<open>atom x \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>, \<Delta>', c, b, s)\<close> using wfS_assertI by auto
+ qed
+next
+ case (wfS_let2I \<Theta> \<Phi> \<B> \<Gamma> \<Delta> s1 \<tau> x s2 b)
+ show ?case proof
+ show \<open> \<Theta>; \<Phi>; \<B>; \<Gamma>; \<Delta>' \<turnstile>\<^sub>w\<^sub>f s1 : b_of \<tau> \<close> using wfS_let2I by auto
+ show \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau> \<close> using wfS_let2I by auto
+ have "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (x, b_of \<tau>, TRUE) #\<^sub>\<Gamma> \<Gamma>" using wfG_cons2I wfX_wfY wfS_let2I by metis
+ hence "\<Theta>; \<B>; (x, b_of \<tau>, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>'" using wf_weakening2(6) wfS_let2I by force
+ thus \<open> \<Theta> ; \<Phi> ; \<B> ; (x, b_of \<tau>, TRUE) #\<^sub>\<Gamma> \<Gamma> ; \<Delta>' \<turnstile>\<^sub>w\<^sub>f s2 : b \<close> using wfS_let2I by metis
+ show \<open>atom x \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>, \<Delta>', s1, b,\<tau>)\<close> using wfS_let2I by auto
+ qed
+next
+ case (wfS_ifI \<Theta> \<B> \<Gamma> v \<Phi> \<Delta> s1 b s2)
+ then show ?case using wf_intros by metis
+next
+ case (wfS_varI \<Theta> \<B> \<Gamma> \<tau> v u \<Phi> \<Delta> b s)
+ show ?case proof
+ show \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau> \<close> using wfS_varI by auto
+ show \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b_of \<tau> \<close> using wfS_varI by auto
+ show \<open>atom u \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>, \<Delta>', \<tau>, v, b)\<close> using wfS_varI setD.simps by auto
+ have "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f (u, \<tau>) #\<^sub>\<Delta> \<Delta>'" using wfS_varI wfD_cons setD.simps u_fresh_d by metis
+ thus \<open> \<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; (u, \<tau>) #\<^sub>\<Delta> \<Delta>' \<turnstile>\<^sub>w\<^sub>f s : b \<close> using wfS_varI setD.simps by blast
+ qed
+next
+ case (wfS_assignI u \<tau> \<Delta> \<Theta> \<B> \<Gamma> \<Phi> v)
+ show ?case proof
+ show \<open>(u, \<tau>) \<in> setD \<Delta>'\<close> using wfS_assignI setD.simps by auto
+ show \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>' \<close> using wfS_assignI by auto
+ show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using wfS_assignI by auto
+ show \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b_of \<tau> \<close> using wfS_assignI by auto
+ qed
+next
+ case (wfS_whileI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> s1 s2 b)
+ then show ?case using wf_intros by metis
+next
+ case (wfS_seqI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> s1 s2 b)
+ then show ?case using wf_intros by metis
+next
+ case (wfS_matchI \<Theta> \<B> \<Gamma> v tid dclist \<Delta> \<Phi> cs b)
+ then show ?case using wf_intros by metis
+next
+ case (wfS_branchI \<Theta> \<Phi> \<B> x \<tau> \<Gamma> \<Delta> s b tid dc)
+ show ?case proof
+ have "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (x, b_of \<tau>, TRUE) #\<^sub>\<Gamma> \<Gamma>" using wfG_cons2I wfX_wfY wfS_branchI by metis
+ hence "\<Theta>; \<B>; (x, b_of \<tau>, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>'" using wf_weakening2(6) wfS_branchI by force
+ thus \<open> \<Theta> ; \<Phi> ; \<B> ; (x, b_of \<tau>, TRUE) #\<^sub>\<Gamma> \<Gamma> ; \<Delta>' \<turnstile>\<^sub>w\<^sub>f s : b \<close> using wfS_branchI by simp
+ show \<open> atom x \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>, \<Delta>', \<Gamma>, \<tau>)\<close> using wfS_branchI by auto
+ show \<open> \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta>' \<close> using wfS_branchI by auto
+ qed
+next
+ case (wfS_finalI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist' cs b dclist)
+ then show ?case using wf_intros by metis
+next
+ case (wfS_cons \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist' cs b css dclist)
+ then show ?case using wf_intros by metis
+qed(auto+)
+
+section \<open>Useful well-formedness instances\<close>
+
+text \<open>Well-formedness for particular constructs that we will need later\<close>
+
+lemma wfC_e_eq:
+ fixes ce::ce and \<Gamma>::\<Gamma>
+ assumes "\<Theta> ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ce : b" and "atom x \<sharp> \<Gamma> "
+ shows "\<Theta> ; \<B> ; ((x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>) \<turnstile>\<^sub>w\<^sub>f (CE_val (V_var x) == ce )"
+proof -
+ have "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b" using assms wfX_wfB by auto
+ hence wbg: "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>" using wfX_wfY assms by auto
+ show ?thesis proof
+ show *:"\<Theta> ; \<B> ; (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_val (V_var x) : b"
+ proof(rule)
+ show "\<Theta> ; \<B> ; (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f V_var x : b " proof
+ show "\<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> " using wfG_cons2I wfX_wfY assms \<open>\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b\<close> by auto
+ show "Some (b, TRUE) = lookup ((x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>) x" using lookup.simps by auto
+ qed
+ qed
+ show "\<Theta> ; \<B> ; (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f ce : b"
+ using assms wf_weakening1(8)[OF assms(1), of "(x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> "] * toSet.simps wfX_wfY
+ by (metis Un_subset_iff equalityE)
+ qed
+qed
+
+lemma wfC_e_eq2:
+ fixes e1::ce and e2::ce
+ assumes "\<Theta> ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f e1 : b" and "\<Theta> ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f e2 : b" and " \<turnstile>\<^sub>w\<^sub>f \<Theta>" and "atom x \<sharp> \<Gamma>"
+ shows "\<Theta>; \<B>; (x, b, (CE_val (V_var x)) == e1) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f (CE_val (V_var x)) == e2 "
+proof(rule wfC_eqI)
+ have *: "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (x, b, CE_val (V_var x) == e1 ) #\<^sub>\<Gamma> \<Gamma>" proof(rule wfG_cons1I )
+ show "(CE_val (V_var x) == e1 ) \<notin> {TRUE, FALSE}" by auto
+ show "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>" using assms wfX_wfY by metis
+ show *:"atom x \<sharp> \<Gamma>" using assms by auto
+ show "\<Theta>; \<B>; (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_val (V_var x) == e1" using wfC_e_eq assms * by auto
+ show "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b" using assms wfX_wfB by auto
+ qed
+ show "\<Theta>; \<B>; (x, b, CE_val (V_var x) == e1 ) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_val (V_var x) : b" using assms * wfCE_valI wfV_varI by auto
+ show "\<Theta>; \<B>; (x, b, CE_val (V_var x) == e1 ) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f e2 : b" proof(rule wf_weakening1(8))
+ show "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f e2 : b " using assms by auto
+ show "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (x, b, CE_val (V_var x) == e1 ) #\<^sub>\<Gamma> \<Gamma>" using * by auto
+ show "toSet \<Gamma> \<subseteq> toSet ((x, b, CE_val (V_var x) == e1 ) #\<^sub>\<Gamma> \<Gamma>)" by auto
+ qed
+qed
+
+lemma wfT_wfT_if_rev:
+ assumes "wfV P \<B> \<Gamma> v (base_for_lit l)" and "wfT P \<B> \<Gamma> t" and \<open>atom z1 \<sharp> \<Gamma>\<close>
+ shows "wfT P \<B> \<Gamma> (\<lbrace> z1 : b_of t | CE_val v == CE_val (V_lit l) IMP (c_of t z1) \<rbrace>)"
+proof
+ show \<open> P; \<B> \<turnstile>\<^sub>w\<^sub>f b_of t \<close> using wfX_wfY assms by meson
+ have wfg: " P; \<B> \<turnstile>\<^sub>w\<^sub>f (z1, b_of t, TRUE) #\<^sub>\<Gamma> \<Gamma>" using assms wfV_wf wfG_cons2I wfX_wfY
+ by (meson wfG_cons_TRUE)
+ show \<open> P; \<B> ; (z1, b_of t, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f [ v ]\<^sup>c\<^sup>e == [ [ l ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of t z1 \<close> proof
+ show *: \<open> P; \<B> ; (z1, b_of t, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f [ v ]\<^sup>c\<^sup>e == [ [ l ]\<^sup>v ]\<^sup>c\<^sup>e \<close>
+ proof(rule wfC_eqI[where b="base_for_lit l"])
+ show "P; \<B> ; (z1, b_of t, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f [ v ]\<^sup>c\<^sup>e : base_for_lit l"
+ using assms wf_intros wf_weakening wfg by (meson wfV_weakening_cons)
+ show "P; \<B> ; (z1, b_of t, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f [ [ l ]\<^sup>v ]\<^sup>c\<^sup>e : base_for_lit l" using wfg assms wf_intros wf_weakening wfV_weakening_cons by meson
+ qed
+ have " t = \<lbrace> z1 : b_of t | c_of t z1 \<rbrace>" using c_of_eq
+ using assms(2) assms(3) b_of_c_of_eq wfT_x_fresh by auto
+ thus \<open> P; \<B> ; (z1, b_of t, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f c_of t z1 \<close> using wfT_wfC assms wfG_elims * by simp
+ qed
+ show \<open>atom z1 \<sharp> (P, \<B>, \<Gamma>)\<close> using assms wfG_fresh_x wfX_wfY by metis
+qed
+
+lemma wfT_eq_imp:
+ fixes zz::x and ll::l and \<tau>'::\<tau>
+ assumes "base_for_lit ll = B_bool" and "\<Theta> ; {||} ; GNil \<turnstile>\<^sub>w\<^sub>f \<tau>'" and
+ "\<Theta> ; {||} \<turnstile>\<^sub>w\<^sub>f (x, b_of \<lbrace> z' : B_bool | TRUE \<rbrace>, c_of \<lbrace> z' : B_bool | TRUE \<rbrace> x) #\<^sub>\<Gamma> GNil" and "atom zz \<sharp> x"
+ shows "\<Theta> ; {||} ; (x, b_of \<lbrace> z' : B_bool | TRUE \<rbrace>, c_of \<lbrace> z' : B_bool | TRUE \<rbrace> x) #\<^sub>\<Gamma>
+ GNil \<turnstile>\<^sub>w\<^sub>f \<lbrace> zz : b_of \<tau>' | [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ ll ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of \<tau>' zz \<rbrace>"
+proof(rule wfT_wfT_if_rev)
+ show \<open> \<Theta> ; {||} ; (x, b_of \<lbrace> z' : B_bool | TRUE \<rbrace>, c_of \<lbrace> z' : B_bool | TRUE \<rbrace> x) #\<^sub>\<Gamma> GNil \<turnstile>\<^sub>w\<^sub>f [ x ]\<^sup>v : base_for_lit ll \<close>
+ using wfV_varI lookup.simps base_for_lit.simps assms by simp
+ show \<open> \<Theta> ; {||} ; (x, b_of \<lbrace> z' : B_bool | TRUE \<rbrace>, c_of \<lbrace> z' : B_bool | TRUE \<rbrace> x) #\<^sub>\<Gamma> GNil \<turnstile>\<^sub>w\<^sub>f \<tau>' \<close>
+ using wf_weakening assms toSet.simps by auto
+ show \<open>atom zz \<sharp> (x, b_of \<lbrace> z' : B_bool | TRUE \<rbrace>, c_of \<lbrace> z' : B_bool | TRUE \<rbrace> x) #\<^sub>\<Gamma> GNil\<close>
+ unfolding fresh_GCons fresh_prod3 b_of.simps c_of_true
+ using x_fresh_b fresh_GNil c_of_true c.fresh assms by metis
+qed
+
+lemma wfC_v_eq:
+ fixes ce::ce and \<Gamma>::\<Gamma> and v::v
+ assumes "\<Theta> ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b" and "atom x \<sharp> \<Gamma> "
+ shows "\<Theta> ; \<B> ; ((x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>) \<turnstile>\<^sub>w\<^sub>f (CE_val (V_var x) == CE_val v )"
+ using wfC_e_eq wfCE_valI assms wfX_wfY by auto
+
+lemma wfT_e_eq:
+ fixes ce::ce
+ assumes "\<Theta> ; \<B> ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ce : b" and "atom z \<sharp> \<Gamma>"
+ shows "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b | CE_val (V_var z) == ce \<rbrace>"
+proof
+ show "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b" using wfX_wfB assms by auto
+ show " atom z \<sharp> (\<Theta>, \<B>, \<Gamma>)" using assms wfG_fresh_x wfX_wfY by metis
+ show "\<Theta> ; \<B> ; (z, b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f CE_val (V_var z) == ce "
+ using wfTI wfC_e_eq assms wfTI by auto
+qed
+
+lemma wfT_v_eq:
+ assumes " wfB \<Theta> \<B> b" and "wfV \<Theta> \<B> \<Gamma> v b" and "atom z \<sharp> \<Gamma>"
+ shows "wfT \<Theta> \<B> \<Gamma> \<lbrace> z : b | C_eq (CE_val (V_var z)) (CE_val v)\<rbrace>"
+ using wfT_e_eq wfE_valI assms wfX_wfY
+ by (simp add: wfCE_valI)
+
+lemma wfC_wfG:
+ fixes \<Gamma>::\<Gamma> and c::c and b::b
+ assumes "\<Theta> ; B ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f c" and "\<Theta> ; B \<turnstile>\<^sub>w\<^sub>f b" and "atom x \<sharp> \<Gamma>"
+ shows "\<Theta> ; B \<turnstile>\<^sub>w\<^sub>f (x,b,c)#\<^sub>\<Gamma> \<Gamma>"
+proof -
+ have " \<Theta> ; B \<turnstile>\<^sub>w\<^sub>f (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>" using wfG_cons2I assms wfX_wfY by fast
+ hence " \<Theta> ; B ; (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f c " using wfC_weakening assms by force
+ thus ?thesis using wfG_consI assms wfX_wfY by metis
+qed
+
+section \<open>Replacing the constraint on a variable in a context\<close>
+
+lemma wfG_cons_fresh2:
+ fixes \<Gamma>'::\<Gamma>
+ assumes "wfG P \<B> (( (x',b',c') #\<^sub>\<Gamma> \<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>))"
+ shows "x'\<noteq>x"
+proof -
+ have "atom x' \<sharp> (\<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>)"
+ using assms wfG_elims(2) by blast
+ thus ?thesis
+ using fresh_gamma_append[of "atom x'" \<Gamma>' "(x, b, c) #\<^sub>\<Gamma> \<Gamma>"] fresh_GCons fresh_prod3[of "atom x'" x b c] by auto
+qed
+
+lemma replace_in_g_inside:
+ fixes \<Gamma>::\<Gamma>
+ assumes "wfG P \<B> (\<Gamma>'@((x,b0,c0') #\<^sub>\<Gamma>\<Gamma>))"
+ shows "replace_in_g (\<Gamma>'@((x,b0,c0') #\<^sub>\<Gamma>\<Gamma>)) x c0 = (\<Gamma>'@((x,b0,c0) #\<^sub>\<Gamma>\<Gamma>))"
+using assms proof(induct \<Gamma>' rule: \<Gamma>_induct)
+ case GNil
+ then show ?case using replace_in_g.simps by auto
+next
+ case (GCons x' b' c' \<Gamma>'')
+ hence "P; \<B> \<turnstile>\<^sub>w\<^sub>f ((x', b', c') #\<^sub>\<Gamma> (\<Gamma>''@ (x, b0, c0') #\<^sub>\<Gamma> \<Gamma> ))" by simp
+ hence "x \<noteq> x'" using wfG_cons_fresh2 by metis
+ then show ?case using replace_in_g.simps GCons by (simp add: wfG_cons)
+qed
+
+lemma wfG_supp_rig_eq:
+ fixes \<Gamma>::\<Gamma>
+ assumes "wfG P \<B> (\<Gamma>'' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma>)" and "wfG P \<B> (\<Gamma>'' @ (x, b0, c0') #\<^sub>\<Gamma> \<Gamma>)"
+ shows "supp (\<Gamma>'' @ (x, b0, c0') #\<^sub>\<Gamma> \<Gamma>) \<union> supp \<B> = supp (\<Gamma>'' @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma>) \<union> supp \<B>"
+using assms proof(induct \<Gamma>'')
+ case GNil
+ have "supp (GNil @ (x, b0, c0') #\<^sub>\<Gamma> \<Gamma>) \<union> supp \<B> = supp ((x, b0, c0') #\<^sub>\<Gamma> \<Gamma>) \<union> supp \<B>" using supp_Cons supp_GNil by auto
+ also have "... = supp x \<union> supp b0 \<union> supp c0' \<union> supp \<Gamma> \<union> supp \<B> " using supp_GCons by auto
+ also have "... = supp x \<union> supp b0 \<union> supp c0 \<union> supp \<Gamma> \<union> supp \<B> " using GNil wfG_wfC[THEN wfC_supp_cons(2) ] by fastforce
+ also have "... = (supp ((x, b0, c0) #\<^sub>\<Gamma> \<Gamma>)) \<union> supp \<B> " using supp_GCons by auto
+ finally have "supp (GNil @ (x, b0, c0') #\<^sub>\<Gamma> \<Gamma>) \<union> supp \<B> = supp (GNil @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma>) \<union> supp \<B>" using supp_Cons supp_GNil by auto
+ then show ?case using supp_GCons wfG_cons2 by auto
+next
+ case (GCons xbc \<Gamma>1)
+ moreover have " (xbc #\<^sub>\<Gamma> \<Gamma>1) @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma> = (xbc #\<^sub>\<Gamma> (\<Gamma>1 @ (x, b0, c0) #\<^sub>\<Gamma> \<Gamma>))" by simp
+ moreover have " (xbc #\<^sub>\<Gamma> \<Gamma>1) @ (x, b0, c0') #\<^sub>\<Gamma> \<Gamma> = (xbc #\<^sub>\<Gamma> (\<Gamma>1 @ (x, b0, c0') #\<^sub>\<Gamma> \<Gamma>))" by simp
+ ultimately have "(P; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>1 @ ((x, b0, c0) #\<^sub>\<Gamma> \<Gamma>)) \<and> P; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>1 @ ((x, b0, c0') #\<^sub>\<Gamma> \<Gamma>)" using wfG_cons2 by metis
+ thus ?case using GCons supp_GCons by auto
+qed
+
+lemma fresh_replace_inside[ms_fresh]:
+ fixes y::x and \<Gamma>::\<Gamma>
+ assumes "wfG P \<B> (\<Gamma>'' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>)" and "wfG P \<B> (\<Gamma>'' @ (x, b, c') #\<^sub>\<Gamma> \<Gamma>)"
+ shows "atom y \<sharp> (\<Gamma>'' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>) = atom y \<sharp> (\<Gamma>'' @ (x, b, c') #\<^sub>\<Gamma> \<Gamma>)"
+ unfolding fresh_def using wfG_supp_rig_eq assms x_not_in_b_set by fast
+
+lemma wf_replace_inside1:
+ fixes \<Gamma>::\<Gamma> and \<Phi>::\<Phi> and \<Theta>::\<Theta> and \<Gamma>'::\<Gamma> and v::v and e::e and c::c and c''::c and c'::c and \<tau>::\<tau> and ts::"(string*\<tau>) list" and \<Delta>::\<Delta> and b'::b and b::b and s::s
+ and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def and cs::branch_s and css::branch_list
+
+shows wfV_replace_inside: "\<Theta>; \<B>; G \<turnstile>\<^sub>w\<^sub>f v : b' \<Longrightarrow> G = (\<Gamma>' @ (x, b, c') #\<^sub>\<Gamma> \<Gamma>) \<Longrightarrow> \<Theta>; \<B>; ((x,b,TRUE) #\<^sub>\<Gamma>\<Gamma>) \<turnstile>\<^sub>w\<^sub>f c \<Longrightarrow> \<Theta> ; \<B> ; (\<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>) \<turnstile>\<^sub>w\<^sub>f v : b'" and
+ wfC_replace_inside: "\<Theta>; \<B>; G \<turnstile>\<^sub>w\<^sub>f c'' \<Longrightarrow> G = (\<Gamma>' @ (x, b, c') #\<^sub>\<Gamma> \<Gamma>) \<Longrightarrow> \<Theta>; \<B>; ((x,b,TRUE) #\<^sub>\<Gamma>\<Gamma>) \<turnstile>\<^sub>w\<^sub>f c \<Longrightarrow> \<Theta> ; \<B> ; (\<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>) \<turnstile>\<^sub>w\<^sub>f c''" and
+ wfG_replace_inside: "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f G \<Longrightarrow> G = (\<Gamma>' @ (x, b, c') #\<^sub>\<Gamma> \<Gamma>) \<Longrightarrow> \<Theta>; \<B>; ((x,b,TRUE) #\<^sub>\<Gamma>\<Gamma>) \<turnstile>\<^sub>w\<^sub>f c \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (\<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>) " and
+ wfT_replace_inside: "\<Theta>; \<B>; G \<turnstile>\<^sub>w\<^sub>f \<tau> \<Longrightarrow> G = (\<Gamma>' @ (x, b, c') #\<^sub>\<Gamma> \<Gamma>) \<Longrightarrow> \<Theta>; \<B>; ((x,b,TRUE) #\<^sub>\<Gamma>\<Gamma>) \<turnstile>\<^sub>w\<^sub>f c \<Longrightarrow> \<Theta> ; \<B> ; (\<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>) \<turnstile>\<^sub>w\<^sub>f \<tau>" and
+ "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ts \<Longrightarrow> True" and
+ "\<turnstile>\<^sub>w\<^sub>f P \<Longrightarrow> True" and
+ "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b \<Longrightarrow> True" and
+ wfCE_replace_inside: "\<Theta> ; \<B> ; G \<turnstile>\<^sub>w\<^sub>f ce : b' \<Longrightarrow> G = (\<Gamma>' @ (x, b, c') #\<^sub>\<Gamma> \<Gamma>) \<Longrightarrow> \<Theta>; \<B>; ((x,b,TRUE) #\<^sub>\<Gamma>\<Gamma>) \<turnstile>\<^sub>w\<^sub>f c \<Longrightarrow> \<Theta> ; \<B> ; (\<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>) \<turnstile>\<^sub>w\<^sub>f ce : b'" and
+ "\<Theta> \<turnstile>\<^sub>w\<^sub>f td \<Longrightarrow> True"
+proof(nominal_induct
+ b' and c'' and G and \<tau> and ts and P and b and b' and td
+ avoiding: \<Gamma>' c'
+rule:wfV_wfC_wfG_wfT_wfTs_wfTh_wfB_wfCE_wfTD.strong_induct)
+ case (wfV_varI \<Theta> \<B> \<Gamma>2 b2 c2 x2)
+ then show ?case using wf_intros by (metis lookup_in_rig_eq lookup_in_rig_neq replace_in_g_inside)
+next
+ case (wfV_conspI s bv dclist \<Theta> dc x1 b' c1 \<B> b1 \<Gamma>1 v)
+ show ?case proof
+ show \<open>AF_typedef_poly s bv dclist \<in> set \<Theta>\<close> using wfV_conspI by auto
+ show \<open>(dc, \<lbrace> x1 : b' | c1 \<rbrace>) \<in> set dclist\<close> using wfV_conspI by auto
+ show \<open> \<Theta> ; \<B> \<turnstile>\<^sub>w\<^sub>f b1 \<close> using wfV_conspI by auto
+ show *: \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b'[bv::=b1]\<^sub>b\<^sub>b \<close> using wfV_conspI by auto
+ moreover have "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>' @ (x, b, c') #\<^sub>\<Gamma> \<Gamma>" using wfV_wf wfV_conspI by simp
+ ultimately have "atom bv \<sharp> \<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>" unfolding fresh_def using wfV_wf wfG_supp_rig_eq wfV_conspI
+ by (metis Un_iff fresh_def)
+ thus \<open>atom bv \<sharp> (\<Theta>, \<B>, \<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>, b1, v)\<close>
+ unfolding fresh_prodN using fresh_prodN wfV_conspI by metis
+ qed
+next
+ case (wfTI z \<Theta> \<B> G b1 c1)
+ show ?case proof
+ show \<open> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b1 \<close> using wfTI by auto
+
+ have "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (x, b, c) #\<^sub>\<Gamma> \<Gamma>" using wfG_consI wfTI wfG_cons wfX_wfY by metis
+ moreover hence *:"wfG \<Theta> \<B> (\<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>)" using wfX_wfY
+ by (metis append_g.simps(2) wfG_cons2 wfTI.hyps wfTI.prems(1) wfTI.prems(2))
+ hence \<open>atom z \<sharp> \<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>\<close>
+ using fresh_replace_inside[of \<Theta> \<B> \<Gamma>' x b c \<Gamma> c' z,OF *] wfTI wfX_wfY wfG_elims by metis
+ thus \<open>atom z \<sharp> (\<Theta>, \<B>, \<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>)\<close> using wfG_fresh_x[OF *] by auto
+
+ have "(z, b1, TRUE) #\<^sub>\<Gamma> G = ((z, b1, TRUE) #\<^sub>\<Gamma> \<Gamma>') @ (x, b, c') #\<^sub>\<Gamma> \<Gamma>"
+ using wfTI append_g.simps by metis
+ thus \<open> \<Theta>; \<B>; (z, b1, TRUE) #\<^sub>\<Gamma> \<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f c1 \<close>
+ using wfTI(9)[OF _ wfTI(11)] by fastforce
+ qed
+next
+ case (wfG_nilI \<Theta>)
+ hence "GNil = (x, b, c') #\<^sub>\<Gamma> \<Gamma>" using append_g.simps \<Gamma>.distinct GNil_append by auto
+ hence "False" using \<Gamma>.distinct by auto
+ then show ?case by auto
+next
+ case (wfG_cons1I c1 \<Theta> \<B> G x1 b1)
+ show ?case proof(cases "\<Gamma>'=GNil")
+ case True
+ then show ?thesis using wfG_cons1I wfG_consI by auto
+ next
+ case False
+ then obtain G'::\<Gamma> where *:"(x1, b1, c1) #\<^sub>\<Gamma> G' = \<Gamma>'" using wfG_cons1I wfG_cons1I(7) GCons_eq_append_conv by auto
+ hence **:" G = G' @ (x, b, c') #\<^sub>\<Gamma> \<Gamma>" using wfG_cons1I by auto
+ hence " \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f G' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>" using wfG_cons1I by auto
+ have "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (x1, b1, c1) #\<^sub>\<Gamma> G' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>" proof(rule Wellformed.wfG_cons1I)
+ show "c1 \<notin> {TRUE, FALSE}" using wfG_cons1I by auto
+ show "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f G' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma> " using wfG_cons1I(3)[of G',OF **] wfG_cons1I by auto
+ show "atom x1 \<sharp> G' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>" using wfG_cons1I * ** fresh_replace_inside by metis
+ show "\<Theta>; \<B>; (x1, b1, TRUE) #\<^sub>\<Gamma> G' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f c1" using wfG_cons1I(6)[of " (x1, b1, TRUE) #\<^sub>\<Gamma> G'"] wfG_cons1I ** by auto
+ show "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b1" using wfG_cons1I by auto
+ qed
+ thus ?thesis using * by auto
+ qed
+next
+ case (wfG_cons2I c1 \<Theta> \<B> G x1 b1)
+ show ?case proof(cases "\<Gamma>'=GNil")
+ case True
+ then show ?thesis using wfG_cons2I wfG_consI by auto
+ next
+ case False
+ then obtain G'::\<Gamma> where *:"(x1, b1, c1) #\<^sub>\<Gamma> G' = \<Gamma>'" using wfG_cons2I GCons_eq_append_conv by auto
+ hence **:" G = G' @ (x, b, c') #\<^sub>\<Gamma> \<Gamma>" using wfG_cons2I by auto
+ moreover have " \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f G' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>" using wfG_cons2I * ** by auto
+ moreover hence "atom x1 \<sharp> G' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>" using wfG_cons2I * ** fresh_replace_inside by metis
+ ultimately show ?thesis using Wellformed.wfG_cons2I[OF wfG_cons2I(1), of \<Theta> \<B> "G'@ (x, b, c) #\<^sub>\<Gamma> \<Gamma>" x1 b1] wfG_cons2I * ** by auto
+ qed
+qed(metis wf_intros )+
+
+lemma wf_replace_inside2:
+ fixes \<Gamma>::\<Gamma> and \<Phi>::\<Phi> and \<Theta>::\<Theta> and \<Gamma>'::\<Gamma> and v::v and e::e and c::c and c''::c and c'::c and \<tau>::\<tau> and ts::"(string*\<tau>) list" and \<Delta>::\<Delta> and b'::b and b::b and s::s
+ and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def and cs::branch_s and css::branch_list
+shows
+ "\<Theta> ; \<Phi> ; \<B> ; G ; D \<turnstile>\<^sub>w\<^sub>f e : b' \<Longrightarrow> G = (\<Gamma>' @ (x, b, c') #\<^sub>\<Gamma> \<Gamma>) \<Longrightarrow> \<Theta>; \<B>; ((x,b,TRUE) #\<^sub>\<Gamma>\<Gamma>) \<turnstile>\<^sub>w\<^sub>f c \<Longrightarrow> \<Theta> ; \<Phi> ; \<B> ; (\<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>); D \<turnstile>\<^sub>w\<^sub>f e : b'" and
+ "\<Theta>; \<Phi>; \<B>; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b \<Longrightarrow> True" and
+ "\<Theta>; \<Phi>; \<B>; \<Gamma> ; \<Delta> ; tid ; dc ; t \<turnstile>\<^sub>w\<^sub>f cs : b \<Longrightarrow> True" and
+ "\<Theta>; \<Phi>; \<B>; \<Gamma> ; \<Delta> ; tid ; dclist \<turnstile>\<^sub>w\<^sub>f css : b \<Longrightarrow> True" and
+ "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<Longrightarrow> True" and
+ "\<Theta>; \<B>; G \<turnstile>\<^sub>w\<^sub>f \<Delta> \<Longrightarrow> G = (\<Gamma>' @ (x, b, c') #\<^sub>\<Gamma> \<Gamma>) \<Longrightarrow> \<Theta>; \<B>; ((x,b,TRUE) #\<^sub>\<Gamma>\<Gamma>) \<turnstile>\<^sub>w\<^sub>f c \<Longrightarrow> \<Theta> ; \<B> ; (\<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>) \<turnstile>\<^sub>w\<^sub>f \<Delta>" and
+ "\<Theta> ; \<Phi> \<turnstile>\<^sub>w\<^sub>f ftq \<Longrightarrow> True" and
+ "\<Theta> ; \<Phi> ; \<B> \<turnstile>\<^sub>w\<^sub>f ft \<Longrightarrow> True"
+proof(nominal_induct
+ b' and b and b and b and \<Phi> and \<Delta> and ftq and ft
+ avoiding: \<Gamma>' c'
+ rule:wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.strong_induct)
+case (wfE_valI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v b)
+ then show ?case using wf_replace_inside1 Wellformed.wfE_valI by auto
+next
+ case (wfE_plusI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 v2)
+ then show ?case using wf_replace_inside1 Wellformed.wfE_plusI by auto
+next
+ case (wfE_leqI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 v2)
+ then show ?case using wf_replace_inside1 Wellformed.wfE_leqI by auto
+next
+ case (wfE_eqI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 b v2)
+ then show ?case using wf_replace_inside1 Wellformed.wfE_eqI by metis
+next
+ case (wfE_fstI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 b1 b2)
+ then show ?case using wf_replace_inside1 Wellformed.wfE_fstI by metis
+next
+ case (wfE_sndI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 b1 b2)
+ then show ?case using wf_replace_inside1 Wellformed.wfE_sndI by metis
+next
+ case (wfE_concatI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 v2)
+ then show ?case using wf_replace_inside1 Wellformed.wfE_concatI by auto
+next
+ case (wfE_splitI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 v2)
+ then show ?case using wf_replace_inside1 Wellformed.wfE_splitI by auto
+next
+ case (wfE_lenI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1)
+ then show ?case using wf_replace_inside1 Wellformed.wfE_lenI by metis
+next
+ case (wfE_appI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> f x b c \<tau> s v)
+ then show ?case using wf_replace_inside1 Wellformed.wfE_appI by metis
+next
+ case (wfE_appPI \<Theta> \<Phi> \<B> \<Gamma>'' \<Delta> b' bv v \<tau> f x1 b1 c1 s)
+ show ?case proof
+ show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using wfE_appPI by auto
+ show \<open> \<Theta>; \<B>; \<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<close> using wfE_appPI by auto
+ show \<open> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b' \<close> using wfE_appPI by auto
+ show *:\<open> \<Theta>; \<B>; \<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b1[bv::=b']\<^sub>b \<close> using wfE_appPI wf_replace_inside1 by auto
+
+ moreover have "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>' @ (x, b, c') #\<^sub>\<Gamma> \<Gamma>" using wfV_wf wfE_appPI by metis
+ ultimately have "atom bv \<sharp> \<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>"
+ unfolding fresh_def using wfV_wf wfG_supp_rig_eq wfE_appPI Un_iff fresh_def by metis
+
+ thus \<open>atom bv \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>, \<Delta>, b', v, (b_of \<tau>)[bv::=b']\<^sub>b)\<close>
+ using wfE_appPI fresh_prodN by metis
+ show \<open>Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x1 b1 c1 \<tau> s))) = lookup_fun \<Phi> f\<close> using wfE_appPI by auto
+ qed
+next
+ case (wfE_mvarI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> u \<tau>)
+ then show ?case using wf_replace_inside1 Wellformed.wfE_mvarI by metis
+next
+ case (wfD_emptyI \<Theta> \<B> \<Gamma>)
+ then show ?case using wf_replace_inside1 Wellformed.wfD_emptyI by metis
+next
+ case (wfD_cons \<Theta> \<B> \<Gamma> \<Delta> \<tau> u)
+ then show ?case using wf_replace_inside1 Wellformed.wfD_emptyI
+ by (simp add: wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.wfD_cons)
+next
+ case (wfFTNone \<Theta> \<Phi> ft)
+ then show ?case using wf_replace_inside1 Wellformed.wfD_emptyI by metis
+next
+ case (wfFTSome \<Theta> \<Phi> bv ft)
+ then show ?case using wf_replace_inside1 Wellformed.wfD_emptyI by metis
+qed(auto)
+
+lemmas wf_replace_inside = wf_replace_inside1 wf_replace_inside2
+
+lemma wfC_replace_cons:
+ assumes "wfG P \<B> ((x,b,c1) #\<^sub>\<Gamma>\<Gamma>)" and "wfC P \<B> ((x,b,TRUE) #\<^sub>\<Gamma>\<Gamma>) c2"
+ shows "wfC P \<B> ((x,b,c1) #\<^sub>\<Gamma>\<Gamma>) c2"
+proof -
+ have "wfC P \<B> (GNil@((x,b,c1) #\<^sub>\<Gamma>\<Gamma>)) c2" proof(rule wf_replace_inside1(2))
+ show " P; \<B> ; (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f c2 " using wfG_elim2 assms by auto
+ show \<open>(x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> = GNil @ (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma>\<close> using append_g.simps by auto
+ show \<open>P; \<B> ; (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f c1 \<close> using wfG_elim2 assms by auto
+ qed
+ thus ?thesis using append_g.simps by auto
+qed
+
+lemma wfC_refl:
+ assumes "wfG \<Theta> \<B> ((x, b', c') #\<^sub>\<Gamma>\<Gamma>)"
+ shows "wfC \<Theta> \<B> ((x, b', c') #\<^sub>\<Gamma>\<Gamma>) c'"
+ using wfG_wfC assms wfC_replace_cons by auto
+
+lemma wfG_wfC_inside:
+ assumes " (x, b, c) \<in> toSet G" and "wfG \<Theta> B G"
+ shows "wfC \<Theta> B G c"
+ using assms proof(induct G rule: \<Gamma>_induct)
+ case GNil
+ then show ?case by auto
+next
+ case (GCons x' b' c' \<Gamma>')
+ then consider (hd) "(x, b, c) = (x',b',c')" | (tail) "(x, b, c) \<in> toSet \<Gamma>'" using toSet.simps by auto
+ then show ?case proof(cases)
+ case hd
+ then show ?thesis using GCons wf_weakening
+ by (metis wfC_replace_cons wfG_cons_wfC)
+ next
+ case tail
+ then show ?thesis using GCons wf_weakening
+ by (metis insert_iff insert_is_Un subsetI toSet.simps(2) wfG_cons2)
+ qed
+qed
+
+lemma wfT_wf_cons3:
+ assumes "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b | c \<rbrace>" and "atom y \<sharp> (c,\<Gamma>)"
+ shows "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (y, b, c[z::=V_var y]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>"
+proof -
+ have "\<lbrace> z : b | c \<rbrace> = \<lbrace> y : b | (y \<leftrightarrow> z) \<bullet> c \<rbrace>" using type_eq_flip assms by auto
+ moreover hence " (y \<leftrightarrow> z) \<bullet> c = c[z::=V_var y]\<^sub>c\<^sub>v" using assms subst_v_c_def by auto
+ ultimately have "\<lbrace> z : b | c \<rbrace> = \<lbrace> y : b | c[z::=V_var y]\<^sub>c\<^sub>v \<rbrace>" by metis
+ thus ?thesis using assms wfT_wf_cons[of \<Theta> \<B> \<Gamma> y b] fresh_Pair by metis
+qed
+
+lemma wfT_wfC_cons:
+ assumes "wfT P \<B> \<Gamma> \<lbrace> z1 : b | c1 \<rbrace>" and "wfT P \<B> \<Gamma> \<lbrace> z2 : b | c2 \<rbrace>" and "atom x \<sharp> (c1,c2,\<Gamma>)"
+ shows "wfC P \<B> ((x,b,c1[z1::=V_var x]\<^sub>v) #\<^sub>\<Gamma>\<Gamma>) (c2[z2::=V_var x]\<^sub>v)" (is "wfC P \<B> ?G ?c")
+proof -
+ have eq: "\<lbrace> z2 : b | c2 \<rbrace> = \<lbrace> x : b | c2[z2::=V_var x]\<^sub>c\<^sub>v \<rbrace>" using type_eq_subst assms fresh_prod3 by simp
+ have eq2: "\<lbrace> z1 : b | c1 \<rbrace> = \<lbrace> x : b | c1[z1::=V_var x]\<^sub>c\<^sub>v \<rbrace>" using type_eq_subst assms fresh_prod3 by simp
+ moreover have "wfT P \<B> \<Gamma> \<lbrace> x : b | c1[z1::=V_var x]\<^sub>c\<^sub>v \<rbrace>" using assms eq2 by auto
+ moreover hence "wfG P \<B> ((x,b,c1[z1::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma>\<Gamma>)" using wfT_wf_cons fresh_prod3 assms by auto
+ moreover have "wfT P \<B> \<Gamma> \<lbrace> x : b | c2[z2::=V_var x]\<^sub>c\<^sub>v \<rbrace>" using assms eq by auto
+ moreover hence "wfC P \<B> ((x,b,TRUE) #\<^sub>\<Gamma>\<Gamma>) (c2[z2::=V_var x]\<^sub>c\<^sub>v)" using wfT_wfC assms fresh_prod3 by simp
+ ultimately show ?thesis using wfC_replace_cons subst_v_c_def by simp
+qed
+
+lemma wfT_wfC2:
+ fixes c::c and x::x
+ assumes "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b | c \<rbrace>" and "atom x \<sharp> \<Gamma>"
+ shows "\<Theta>; \<B>; (x,b,TRUE)#\<^sub>\<Gamma>\<Gamma> \<turnstile>\<^sub>w\<^sub>f c[z::=[x]\<^sup>v]\<^sub>v"
+proof(cases "x=z")
+ case True
+ then show ?thesis using wfT_wfC assms by auto
+next
+ case False
+ hence "atom x \<sharp> c" using wfT_fresh_c assms by metis
+ hence "\<lbrace> x : b | c[z::=[ x ]\<^sup>v]\<^sub>v \<rbrace> = \<lbrace> z : b | c \<rbrace>"
+ using \<tau>.eq_iff Abs1_eq_iff(3)[of x "c[z::=[ x ]\<^sup>v]\<^sub>v" z c]
+ by (metis flip_subst_v type_eq_flip)
+ hence " \<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> x : b | c[z::=[ x ]\<^sup>v]\<^sub>v \<rbrace>" using assms by metis
+ thus ?thesis using wfT_wfC assms by auto
+qed
+
+lemma wfT_wfG:
+ fixes x::x and \<Gamma>::\<Gamma> and z::x and c::c and b::b
+ assumes "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b | c \<rbrace>" and "atom x \<sharp> \<Gamma>"
+ shows "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (x,b, c[z::=[ x ]\<^sup>v]\<^sub>v) #\<^sub>\<Gamma> \<Gamma>"
+proof -
+ have "\<Theta>; \<B>; (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f c[z::=[ x ]\<^sup>v]\<^sub>v" using wfT_wfC2 assms by metis
+ thus ?thesis using wfG_consI assms wfT_wfB b_of.simps wfX_wfY by metis
+qed
+
+lemma wfG_replace_inside2:
+ fixes \<Gamma>::\<Gamma>
+ assumes "wfG P \<B> (\<Gamma>' @ (x, b, c') #\<^sub>\<Gamma> \<Gamma>)" and "wfG P \<B> ((x,b,c) #\<^sub>\<Gamma>\<Gamma>)"
+ shows "wfG P \<B> (\<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>)"
+proof -
+ have "wfC P \<B> ((x,b,TRUE) #\<^sub>\<Gamma>\<Gamma>) c" using wfG_wfC assms by auto
+ thus ?thesis using wf_replace_inside1(3)[OF assms(1)] by auto
+qed
+
+lemma wfG_replace_inside_full:
+ fixes \<Gamma>::\<Gamma>
+ assumes "wfG P \<B> (\<Gamma>' @ (x, b, c') #\<^sub>\<Gamma> \<Gamma>)" and "wfG P \<B> (\<Gamma>'@((x,b,c) #\<^sub>\<Gamma>\<Gamma>))"
+ shows "wfG P \<B> (\<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>)"
+proof -
+ have "wfG P \<B> ((x,b,c) #\<^sub>\<Gamma>\<Gamma>)" using wfG_suffix assms by auto
+ thus ?thesis using wfG_replace_inside assms by auto
+qed
+
+lemma wfT_replace_inside2:
+ assumes "wfT \<Theta> \<B> (\<Gamma>' @ (x, b, c') #\<^sub>\<Gamma> \<Gamma>) t" and "wfG \<Theta> \<B> (\<Gamma>'@((x,b,c) #\<^sub>\<Gamma>\<Gamma>))"
+ shows "wfT \<Theta> \<B> (\<Gamma>' @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>) t"
+proof -
+ have "wfG \<Theta> \<B> (((x,b,c) #\<^sub>\<Gamma>\<Gamma>))" using wfG_suffix assms by auto
+ hence "wfC \<Theta> \<B> ((x,b,TRUE) #\<^sub>\<Gamma>\<Gamma>) c" using wfG_wfC by auto
+ thus ?thesis using wf_replace_inside assms by metis
+qed
+
+lemma wfD_unique:
+ assumes "wfD P \<B> \<Gamma> \<Delta>" and " (u,\<tau>') \<in> setD \<Delta>" and "(u,\<tau>) \<in> setD \<Delta>"
+ shows "\<tau>'=\<tau>"
+using assms proof(induct \<Delta> rule: \<Delta>_induct)
+ case DNil
+ then show ?case by auto
+next
+ case (DCons u' t' D)
+ hence *: "wfD P \<B> \<Gamma> ((u',t') #\<^sub>\<Delta> D)" using Cons by auto
+ show ?case proof(cases "u=u'")
+ case True
+ then have "u \<notin> fst ` setD D" using wfD_elims * by blast
+ then show ?thesis using DCons by force
+ next
+ case False
+ then show ?thesis using DCons wfD_elims * by (metis fst_conv setD_ConsD)
+ qed
+qed
+
+lemma replace_in_g_forget:
+ fixes x::x
+ assumes "wfG P B G"
+ shows "atom x \<notin> atom_dom G \<Longrightarrow> (G[x\<longmapsto>c]) = G" and
+ "atom x \<sharp> G \<Longrightarrow> (G[x\<longmapsto>c]) = G"
+proof -
+ show "atom x \<notin> atom_dom G \<Longrightarrow> G[x\<longmapsto>c] = G" by (induct G rule: \<Gamma>_induct,auto)
+ thus "atom x \<sharp> G \<Longrightarrow> (G[x\<longmapsto>c]) = G" using wfG_x_fresh assms by simp
+qed
+
+lemma replace_in_g_fresh_single:
+ fixes G::\<Gamma> and x::x
+ assumes \<open>\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f G[x'\<longmapsto>c'']\<close> and "atom x \<sharp> G" and \<open>\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f G \<close>
+ shows "atom x \<sharp> G[x'\<longmapsto>c'']"
+ using rig_dom_eq wfG_dom_supp assms fresh_def atom_dom.simps dom.simps by metis
+
+section \<open>Preservation of well-formedness under substitution\<close>
+
+lemma wfC_cons_switch:
+ fixes c::c and c'::c
+ assumes "\<Theta>; \<B>; (x, b, c) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f c'"
+ shows "\<Theta>; \<B>; (x, b, c') #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f c"
+proof -
+ have *:"\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (x, b, c) #\<^sub>\<Gamma> \<Gamma>" using wfC_wf assms by auto
+ hence "atom x \<sharp> \<Gamma> \<and> wfG \<Theta> \<B> \<Gamma> \<and> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b" using wfG_cons by auto
+ hence " \<Theta>; \<B>; (x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f TRUE " using wfC_trueI wfG_cons2I by simp
+ hence "\<Theta>; \<B>;(x, b, TRUE) #\<^sub>\<Gamma> \<Gamma> \<turnstile>\<^sub>w\<^sub>f c'"
+ using wf_replace_inside1(2)[of \<Theta> \<B> "(x, b, c) #\<^sub>\<Gamma> \<Gamma>" c' GNil x b c \<Gamma> TRUE] assms by auto
+ hence "wfG \<Theta> \<B> ((x,b,c') #\<^sub>\<Gamma>\<Gamma>)" using wf_replace_inside1(3)[OF *, of GNil x b c \<Gamma> c'] by auto
+ moreover have "wfC \<Theta> \<B> ((x,b,TRUE) #\<^sub>\<Gamma>\<Gamma>) c" proof(cases "c \<in> { TRUE, FALSE }")
+ case True
+ have "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<and> atom x \<sharp> \<Gamma> \<and> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b" using wfG_elims(2)[OF *] by auto
+ hence "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (x,b,TRUE) #\<^sub>\<Gamma> \<Gamma>" using wfG_cons_TRUE by auto
+ then show ?thesis using wfC_trueI wfC_falseI True by auto
+ next
+ case False
+ then show ?thesis using wfG_elims(2)[OF *] by auto
+ qed
+ ultimately show ?thesis using wfC_replace_cons by auto
+qed
+
+lemma subst_g_inside_simple:
+ fixes \<Gamma>\<^sub>1::\<Gamma> and \<Gamma>\<^sub>2::\<Gamma>
+ assumes "wfG P \<B> (\<Gamma>\<^sub>1@((x,b,c) #\<^sub>\<Gamma>\<Gamma>\<^sub>2))"
+ shows "(\<Gamma>\<^sub>1@((x,b,c) #\<^sub>\<Gamma>\<Gamma>\<^sub>2))[x::=v]\<^sub>\<Gamma>\<^sub>v = \<Gamma>\<^sub>1[x::=v]\<^sub>\<Gamma>\<^sub>v@\<Gamma>\<^sub>2"
+using assms proof(induct \<Gamma>\<^sub>1 rule: \<Gamma>_induct)
+ case GNil
+ then show ?case using subst_gv.simps by simp
+next
+ case (GCons x' b' c' G)
+ hence *:"P; \<B> \<turnstile>\<^sub>w\<^sub>f (x', b', c') #\<^sub>\<Gamma> (G @ (x, b, c) #\<^sub>\<Gamma> \<Gamma>\<^sub>2)" by auto
+ hence "x\<noteq>x'"
+ using GCons append_Cons wfG_cons_fresh2[OF *] by auto
+ hence "((GCons (x', b', c') G) @ (GCons (x, b, c) \<Gamma>\<^sub>2))[x::=v]\<^sub>\<Gamma>\<^sub>v =
+ (GCons (x', b', c') (G @ (GCons (x, b, c) \<Gamma>\<^sub>2)))[x::=v]\<^sub>\<Gamma>\<^sub>v" by auto
+ also have "... = GCons (x', b', c'[x::=v]\<^sub>c\<^sub>v) ((G @ (GCons (x, b, c) \<Gamma>\<^sub>2))[x::=v]\<^sub>\<Gamma>\<^sub>v)"
+ using subst_gv.simps \<open>x\<noteq>x'\<close> by simp
+ also have "... = (x', b', c'[x::=v]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> (G[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>\<^sub>2)" using GCons * wfG_elims by metis
+ also have "... = ((x', b', c') #\<^sub>\<Gamma> G)[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma>\<^sub>2" using subst_gv.simps \<open>x\<noteq>x'\<close> by simp
+ finally show ?case by blast
+qed
+
+lemma subst_c_TRUE_FALSE:
+ fixes c::c
+ assumes "c \<notin> {TRUE,FALSE}"
+ shows "c[x::=v']\<^sub>c\<^sub>v \<notin> {TRUE, FALSE}"
+using assms by(nominal_induct c rule:c.strong_induct,auto simp add: subst_cv.simps)
+
+lemma lookup_subst:
+ assumes "Some (b, c) = lookup \<Gamma> x" and "x \<noteq> x'"
+ shows "\<exists>c'. Some (b,c') = lookup \<Gamma>[x'::=v']\<^sub>\<Gamma>\<^sub>v x"
+using assms proof(induct \<Gamma> rule: \<Gamma>_induct)
+case GNil
+ then show ?case by auto
+next
+ case (GCons x1 b1 c1 \<Gamma>1)
+ then show ?case proof(cases "x1=x'")
+ case True
+ then show ?thesis using subst_gv.simps GCons by auto
+ next
+ case False
+ hence *:"((x1, b1, c1) #\<^sub>\<Gamma> \<Gamma>1)[x'::=v']\<^sub>\<Gamma>\<^sub>v = ((x1, b1, c1[x'::=v']\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>1[x'::=v']\<^sub>\<Gamma>\<^sub>v)" using subst_gv.simps by auto
+ then show ?thesis proof(cases "x1=x")
+ case True
+ then show ?thesis using lookup.simps *
+ using GCons.prems(1) by auto
+ next
+ case False
+ then show ?thesis using lookup.simps *
+ using GCons.prems(1) by (simp add: GCons.hyps assms(2))
+ qed
+ qed
+qed
+
+lemma lookup_subst2:
+ assumes "Some (b, c) = lookup (\<Gamma>'@((x',b\<^sub>1,c0[z0::=[x']\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\<Gamma>\<Gamma>)) x" and "x \<noteq> x'" and
+ "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (\<Gamma>'@((x',b\<^sub>1,c0[z0::=[x']\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\<Gamma>\<Gamma>))"
+ shows "\<exists>c'. Some (b,c') = lookup (\<Gamma>'[x'::=v']\<^sub>\<Gamma>\<^sub>v@\<Gamma>) x"
+ using assms lookup_subst subst_g_inside by metis
+
+lemma wf_subst1:
+ fixes \<Gamma>::\<Gamma> and \<Gamma>'::\<Gamma> and v::v and e::e and c::c and \<tau>::\<tau> and ts::"(string*\<tau>) list" and \<Delta>::\<Delta> and b::b and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def
+ shows wfV_subst: "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b \<Longrightarrow> \<Gamma>=\<Gamma>\<^sub>1@((x,b',c') #\<^sub>\<Gamma>\<Gamma>\<^sub>2) \<Longrightarrow> \<Theta>; \<B>;\<Gamma>\<^sub>2 \<turnstile>\<^sub>w\<^sub>f v' : b' \<Longrightarrow> \<Theta> ; \<B> ; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v \<turnstile>\<^sub>w\<^sub>f v[x::=v']\<^sub>v\<^sub>v : b" and
+ wfC_subst: "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f c \<Longrightarrow> \<Gamma>=\<Gamma>\<^sub>1@((x,b',c') #\<^sub>\<Gamma>\<Gamma>\<^sub>2) \<Longrightarrow> \<Theta>; \<B>; \<Gamma>\<^sub>2 \<turnstile>\<^sub>w\<^sub>f v' : b' \<Longrightarrow> \<Theta>; \<B>; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v \<turnstile>\<^sub>w\<^sub>f c[x::=v']\<^sub>c\<^sub>v" and
+ wfG_subst: "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<Longrightarrow> \<Gamma>=\<Gamma>\<^sub>1@((x,b',c') #\<^sub>\<Gamma>\<Gamma>\<^sub>2) \<Longrightarrow> \<Theta>; \<B> ; \<Gamma>\<^sub>2 \<turnstile>\<^sub>w\<^sub>f v' : b' \<Longrightarrow> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v" and
+ wfT_subst: "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau> \<Longrightarrow> \<Gamma>=\<Gamma>\<^sub>1@((x,b',c') #\<^sub>\<Gamma>\<Gamma>\<^sub>2) \<Longrightarrow> \<Theta>; \<B> ; \<Gamma>\<^sub>2 \<turnstile>\<^sub>w\<^sub>f v' : b' \<Longrightarrow> \<Theta>; \<B>; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v \<turnstile>\<^sub>w\<^sub>f \<tau>[x::=v']\<^sub>\<tau>\<^sub>v" and
+ "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ts \<Longrightarrow> True" and
+ "\<turnstile>\<^sub>w\<^sub>f \<Theta> \<Longrightarrow>True" and
+ "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b \<Longrightarrow> True " and
+ wfCE_subst: "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ce : b \<Longrightarrow> \<Gamma>=\<Gamma>\<^sub>1@((x,b',c') #\<^sub>\<Gamma>\<Gamma>\<^sub>2) \<Longrightarrow> \<Theta>; \<B> ; \<Gamma>\<^sub>2 \<turnstile>\<^sub>w\<^sub>f v' : b' \<Longrightarrow> \<Theta> ; \<B> ; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v \<turnstile>\<^sub>w\<^sub>f ce[x::=v']\<^sub>c\<^sub>e\<^sub>v : b" and
+ "\<Theta> \<turnstile>\<^sub>w\<^sub>f td \<Longrightarrow> True"
+proof(nominal_induct
+ b and c and \<Gamma> and \<tau> and ts and \<Theta> and b and b and td
+ avoiding: x v'
+ arbitrary: \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1
+ rule:wfV_wfC_wfG_wfT_wfTs_wfTh_wfB_wfCE_wfTD.strong_induct)
+ case (wfV_varI \<Theta> \<B> \<Gamma> b1 c1 x1)
+
+ show ?case proof(cases "x1=x")
+ case True
+ hence "(V_var x1)[x::=v']\<^sub>v\<^sub>v = v' " using subst_vv.simps by auto
+ moreover have "b' = b1" using wfV_varI True lookup_inside_wf
+ by (metis option.inject prod.inject)
+ moreover have " \<Theta>; \<B> ; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v \<turnstile>\<^sub>w\<^sub>f v' : b'" using wfV_varI subst_g_inside_simple wf_weakening
+ append_g_toSetU sup_ge2 wfV_wf by metis
+ ultimately show ?thesis by auto
+ next
+ case False
+ hence "(V_var x1)[x::=v']\<^sub>v\<^sub>v = (V_var x1) " using subst_vv.simps by auto
+ moreover have "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v" using wfV_varI by simp
+ moreover obtain c1' where "Some (b1, c1') = lookup \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v x1" using wfV_varI False lookup_subst by metis
+ ultimately show ?thesis using Wellformed.wfV_varI[of \<Theta> \<B> "\<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v" b1 c1' x1] by metis
+ qed
+next
+ case (wfV_litI \<Theta> \<Gamma> l)
+ then show ?case using subst_vv.simps wf_intros by auto
+next
+ case (wfV_pairI \<Theta> \<Gamma> v1 b1 v2 b2)
+ then show ?case using subst_vv.simps wf_intros by auto
+next
+ case (wfV_consI s dclist \<Theta> dc x b c \<Gamma> v)
+ then show ?case using subst_vv.simps wf_intros by auto
+next
+ case (wfV_conspI s bv dclist \<Theta> dc x' b' c \<B> b \<Gamma> va)
+ show ?case unfolding subst_vv.simps proof
+ show \<open>AF_typedef_poly s bv dclist \<in> set \<Theta>\<close> and \<open>(dc, \<lbrace> x' : b' | c \<rbrace>) \<in> set dclist\<close> using wfV_conspI by auto
+ show \<open> \<Theta> ;\<B> \<turnstile>\<^sub>w\<^sub>f b \<close> using wfV_conspI by auto
+ have "atom bv \<sharp> \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v" using fresh_subst_gv_if wfV_conspI by metis
+ moreover have "atom bv \<sharp> va[x::=v']\<^sub>v\<^sub>v" using wfV_conspI fresh_subst_if by simp
+ ultimately show \<open>atom bv \<sharp> (\<Theta>, \<B>, \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v, b, va[x::=v']\<^sub>v\<^sub>v)\<close> unfolding fresh_prodN using wfV_conspI by auto
+ show \<open> \<Theta>; \<B>; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v \<turnstile>\<^sub>w\<^sub>f va[x::=v']\<^sub>v\<^sub>v : b'[bv::=b]\<^sub>b\<^sub>b \<close> using wfV_conspI by auto
+ qed
+next
+ case (wfTI z \<Theta> \<B> \<Gamma> b c)
+ have " \<Theta>; \<B>; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v \<turnstile>\<^sub>w\<^sub>f \<lbrace> z : b | c[x::=v']\<^sub>c\<^sub>v \<rbrace>" proof
+ have \<open>\<Theta>; \<B>; ((z, b, TRUE) #\<^sub>\<Gamma> \<Gamma>)[x::=v']\<^sub>\<Gamma>\<^sub>v \<turnstile>\<^sub>w\<^sub>f c[x::=v']\<^sub>c\<^sub>v \<close>
+ proof(rule wfTI(9))
+ show \<open>(z, b, TRUE) #\<^sub>\<Gamma> \<Gamma> = ((z, b, TRUE) #\<^sub>\<Gamma> \<Gamma>\<^sub>1) @ (x, b', c') #\<^sub>\<Gamma> \<Gamma>\<^sub>2\<close> using wfTI append_g.simps by simp
+ show \<open> \<Theta>; \<B>; \<Gamma>\<^sub>2 \<turnstile>\<^sub>w\<^sub>f v' : b' \<close> using wfTI by auto
+ qed
+ thus *:\<open>\<Theta>; \<B>; (z, b, TRUE) #\<^sub>\<Gamma> \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v \<turnstile>\<^sub>w\<^sub>f c[x::=v']\<^sub>c\<^sub>v \<close>
+ using subst_gv.simps subst_cv.simps wfTI fresh_x_neq by auto
+
+ have "atom z \<sharp> \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v" using fresh_subst_gv_if wfTI by metis
+ moreover have "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v" using wfTI wfX_wfY wfG_elims subst_gv.simps * by metis
+ ultimately show \<open>atom z \<sharp> (\<Theta>, \<B>, \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v)\<close> using wfG_fresh_x by metis
+ show \<open> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b \<close> using wfTI by auto
+ qed
+ thus ?case using subst_tv.simps wfTI by auto
+next
+ case (wfC_trueI \<Theta> \<Gamma>)
+ then show ?case using subst_cv.simps wf_intros by auto
+next
+ case (wfC_falseI \<Theta> \<Gamma>)
+ then show ?case using subst_cv.simps wf_intros by auto
+next
+ case (wfC_eqI \<Theta> \<B> \<Gamma> e1 b e2)
+ show ?case proof(subst subst_cv.simps,rule)
+ show "\<Theta>; \<B>; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v \<turnstile>\<^sub>w\<^sub>f e1[x::=v']\<^sub>c\<^sub>e\<^sub>v : b " using wfC_eqI subst_dv.simps by auto
+ show "\<Theta>; \<B>; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v \<turnstile>\<^sub>w\<^sub>f e2[x::=v']\<^sub>c\<^sub>e\<^sub>v : b " using wfC_eqI by auto
+ qed
+next
+ case (wfC_conjI \<Theta> \<Gamma> c1 c2)
+ then show ?case using subst_cv.simps wf_intros by auto
+next
+ case (wfC_disjI \<Theta> \<Gamma> c1 c2)
+ then show ?case using subst_cv.simps wf_intros by auto
+next
+ case (wfC_notI \<Theta> \<Gamma> c1)
+ then show ?case using subst_cv.simps wf_intros by auto
+next
+ case (wfC_impI \<Theta> \<Gamma> c1 c2)
+ then show ?case using subst_cv.simps wf_intros by auto
+next
+ case (wfG_nilI \<Theta>)
+ then show ?case using subst_cv.simps wf_intros by auto
+next
+ case (wfG_cons1I c \<Theta> \<B> \<Gamma> y b)
+
+ show ?case proof(cases "x=y")
+ case True
+ hence "((y, b, c) #\<^sub>\<Gamma> \<Gamma>)[x::=v']\<^sub>\<Gamma>\<^sub>v = \<Gamma>" using subst_gv.simps by auto
+ moreover have "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>" using wfG_cons1I by auto
+ ultimately show ?thesis by auto
+ next
+ case False
+ have "\<Gamma>\<^sub>1 \<noteq> GNil" using wfG_cons1I False by auto
+ then obtain G where "\<Gamma>\<^sub>1 = (y, b, c) #\<^sub>\<Gamma> G" using GCons_eq_append_conv wfG_cons1I by auto
+ hence *:"\<Gamma> = G @ (x, b', c') #\<^sub>\<Gamma> \<Gamma>\<^sub>2" using wfG_cons1I by auto
+ hence "((y, b, c) #\<^sub>\<Gamma> \<Gamma>)[x::=v']\<^sub>\<Gamma>\<^sub>v =(y, b, c[x::=v']\<^sub>c\<^sub>v) #\<^sub>\<Gamma>\<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v" using subst_gv.simps False by auto
+ moreover have "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (y, b, c[x::=v']\<^sub>c\<^sub>v) #\<^sub>\<Gamma>\<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v" proof(rule Wellformed.wfG_cons1I)
+ show \<open>c[x::=v']\<^sub>c\<^sub>v \<notin> {TRUE, FALSE}\<close> using wfG_cons1I subst_c_TRUE_FALSE by auto
+ show \<open> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v \<close> using wfG_cons1I * by auto
+ have "\<Gamma> = (G @ ((x, b', c') #\<^sub>\<Gamma>GNil)) @ \<Gamma>\<^sub>2" using * append_g_assoc by auto
+ hence "atom y \<sharp> \<Gamma>\<^sub>2" using fresh_suffix \<open>atom y \<sharp> \<Gamma>\<close> by auto
+ hence "atom y \<sharp> v'" using wfG_cons1I wfV_x_fresh by metis
+ thus \<open>atom y \<sharp> \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v\<close> using fresh_subst_gv wfG_cons1I by auto
+ have "((y, b, TRUE) #\<^sub>\<Gamma> \<Gamma>)[x::=v']\<^sub>\<Gamma>\<^sub>v = (y, b, TRUE) #\<^sub>\<Gamma> \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v" using subst_gv.simps subst_cv.simps False by auto
+ thus \<open> \<Theta>; \<B>; (y, b, TRUE) #\<^sub>\<Gamma> \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v \<turnstile>\<^sub>w\<^sub>f c[x::=v']\<^sub>c\<^sub>v \<close> using wfG_cons1I(6)[of "(y,b,TRUE) #\<^sub>\<Gamma>G"] * subst_gv.simps
+ wfG_cons1I by fastforce
+ show "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b " using wfG_cons1I by auto
+ qed
+ ultimately show ?thesis by auto
+ qed
+next
+ case (wfG_cons2I c \<Theta> \<B> \<Gamma> y b)
+ show ?case proof(cases "x=y")
+ case True
+ hence "((y, b, c) #\<^sub>\<Gamma> \<Gamma>)[x::=v']\<^sub>\<Gamma>\<^sub>v = \<Gamma>" using subst_gv.simps by auto
+ moreover have "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>" using wfG_cons2I by auto
+ ultimately show ?thesis by auto
+ next
+ case False
+ have "\<Gamma>\<^sub>1 \<noteq> GNil" using wfG_cons2I False by auto
+ then obtain G where "\<Gamma>\<^sub>1 = (y, b, c) #\<^sub>\<Gamma> G" using GCons_eq_append_conv wfG_cons2I by auto
+ hence *:"\<Gamma> = G @ (x, b', c') #\<^sub>\<Gamma> \<Gamma>\<^sub>2" using wfG_cons2I by auto
+ hence "((y, b, c) #\<^sub>\<Gamma> \<Gamma>)[x::=v']\<^sub>\<Gamma>\<^sub>v =(y, b, c[x::=v']\<^sub>c\<^sub>v) #\<^sub>\<Gamma>\<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v" using subst_gv.simps False by auto
+ moreover have "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (y, b, c[x::=v']\<^sub>c\<^sub>v) #\<^sub>\<Gamma>\<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v" proof(rule Wellformed.wfG_cons2I)
+ show \<open>c[x::=v']\<^sub>c\<^sub>v \<in> {TRUE, FALSE}\<close> using subst_cv.simps wfG_cons2I by auto
+ show \<open> \<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v \<close> using wfG_cons2I * by auto
+ have "\<Gamma> = (G @ ((x, b', c') #\<^sub>\<Gamma>GNil)) @ \<Gamma>\<^sub>2" using * append_g_assoc by auto
+ hence "atom y \<sharp> \<Gamma>\<^sub>2" using fresh_suffix wfG_cons2I by metis
+ hence "atom y \<sharp> v'" using wfG_cons2I wfV_x_fresh by metis
+ thus \<open>atom y \<sharp> \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v\<close> using fresh_subst_gv wfG_cons2I by auto
+ show "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b " using wfG_cons2I by auto
+ qed
+ ultimately show ?thesis by auto
+ qed
+next
+ case (wfCE_valI \<Theta> \<B> \<Gamma> v b)
+ then show ?case using subst_vv.simps wf_intros by auto
+next
+ case (wfCE_plusI \<Theta> \<B> \<Gamma> v1 v2)
+ then show ?case using subst_vv.simps wf_intros by auto
+next
+ case (wfCE_leqI \<Theta> \<B> \<Gamma> v1 v2)
+ then show ?case using subst_vv.simps wf_intros by auto
+next
+ case (wfCE_eqI \<Theta> \<B> \<Gamma> v1 b v2)
+ then show ?case unfolding subst_cev.simps
+ using Wellformed.wfCE_eqI by metis
+next
+ case (wfCE_fstI \<Theta> \<B> \<Gamma> v1 b1 b2)
+ then show ?case using Wellformed.wfCE_fstI subst_cev.simps by metis
+next
+ case (wfCE_sndI \<Theta> \<B> \<Gamma> v1 b1 b2)
+ then show ?case using subst_cev.simps wf_intros by metis
+next
+ case (wfCE_concatI \<Theta> \<B> \<Gamma> v1 v2)
+ then show ?case using subst_vv.simps wf_intros by auto
+next
+ case (wfCE_lenI \<Theta> \<B> \<Gamma> v1)
+ then show ?case using subst_vv.simps wf_intros by auto
+qed(metis subst_sv.simps wf_intros)+
+
+lemma wf_subst2:
+ fixes \<Gamma>::\<Gamma> and \<Gamma>'::\<Gamma> and v::v and e::e and c::c and \<tau>::\<tau> and ts::"(string*\<tau>) list" and \<Delta>::\<Delta> and b::b and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def
+ shows "\<Theta>; \<Phi>; \<B>; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f e : b \<Longrightarrow> \<Gamma>=\<Gamma>\<^sub>1@((x,b',c') #\<^sub>\<Gamma>\<Gamma>\<^sub>2) \<Longrightarrow> \<Theta>; \<B> ; \<Gamma>\<^sub>2 \<turnstile>\<^sub>w\<^sub>f v' : b' \<Longrightarrow> \<Theta> ; \<Phi> ; \<B> ; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v \<turnstile>\<^sub>w\<^sub>f e[x::=v']\<^sub>e\<^sub>v : b" and
+ "\<Theta>; \<Phi>; \<B>; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b \<Longrightarrow> \<Gamma>=\<Gamma>\<^sub>1@((x,b',c') #\<^sub>\<Gamma>\<Gamma>\<^sub>2) \<Longrightarrow> \<Theta> ;\<B> ; \<Gamma>\<^sub>2 \<turnstile>\<^sub>w\<^sub>f v' : b' \<Longrightarrow> \<Theta> ; \<Phi> ; \<B> ; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v \<turnstile>\<^sub>w\<^sub>f s[x::=v']\<^sub>s\<^sub>v : b" and
+ "\<Theta>; \<Phi>; \<B>; \<Gamma> ; \<Delta> ; tid ; dc ; t \<turnstile>\<^sub>w\<^sub>f cs : b \<Longrightarrow> \<Gamma>=\<Gamma>\<^sub>1@((x,b',c') #\<^sub>\<Gamma>\<Gamma>\<^sub>2) \<Longrightarrow> \<Theta>; \<B>; \<Gamma>\<^sub>2 \<turnstile>\<^sub>w\<^sub>f v' : b' \<Longrightarrow> \<Theta>; \<Phi>; \<B>; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v ; tid ; dc ; t \<turnstile>\<^sub>w\<^sub>f subst_branchv cs x v' : b" and
+ "\<Theta>; \<Phi>; \<B>; \<Gamma> ; \<Delta> ; tid ; dclist \<turnstile>\<^sub>w\<^sub>f css : b \<Longrightarrow> \<Gamma>=\<Gamma>\<^sub>1@((x,b',c') #\<^sub>\<Gamma>\<Gamma>\<^sub>2) \<Longrightarrow> \<Theta>; \<B>; \<Gamma>\<^sub>2 \<turnstile>\<^sub>w\<^sub>f v' : b' \<Longrightarrow> \<Theta>; \<Phi>; \<B>; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v ; tid ; dclist \<turnstile>\<^sub>w\<^sub>f subst_branchlv css x v' : b" and
+ "\<Theta> \<turnstile>\<^sub>w\<^sub>f (\<Phi>::\<Phi>) \<Longrightarrow> True " and
+ "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<Longrightarrow> \<Gamma>=\<Gamma>\<^sub>1@((x,b',c') #\<^sub>\<Gamma>\<Gamma>\<^sub>2) \<Longrightarrow> \<Theta>; \<B> ; \<Gamma>\<^sub>2 \<turnstile>\<^sub>w\<^sub>f v' : b' \<Longrightarrow> \<Theta> ; \<B> ; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v \<turnstile>\<^sub>w\<^sub>f \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v" and
+ "\<Theta> ; \<Phi> \<turnstile>\<^sub>w\<^sub>f ftq \<Longrightarrow> True" and
+ "\<Theta> ; \<Phi> ; \<B> \<turnstile>\<^sub>w\<^sub>f ft \<Longrightarrow> True"
+proof(nominal_induct
+ b and b and b and b and \<Phi> and \<Delta> and ftq and ft
+ avoiding: x v'
+ arbitrary: \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1
+ rule:wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.strong_induct)
+ case (wfE_valI \<Theta> \<Gamma> v b)
+ then show ?case using subst_vv.simps wf_intros wf_subst1
+ by (metis subst_ev.simps(1))
+next
+ case (wfE_plusI \<Theta> \<Gamma> v1 v2)
+ then show ?case using subst_vv.simps wf_intros wf_subst1 by auto
+next
+ case (wfE_leqI \<Theta> \<Phi> \<Gamma> \<Delta> v1 v2)
+ then show ?case
+ using subst_vv.simps subst_ev.simps subst_ev.simps wf_subst1 Wellformed.wfE_leqI
+ by auto
+next
+ case (wfE_eqI \<Theta> \<Phi> \<Gamma> \<Delta> v1 b v2)
+ then show ?case
+ using subst_vv.simps subst_ev.simps subst_ev.simps wf_subst1 Wellformed.wfE_eqI
+ proof -
+ show ?thesis
+ by (metis (no_types) subst_ev.simps(4) wfE_eqI.hyps(1) wfE_eqI.hyps(4) wfE_eqI.hyps(5) wfE_eqI.hyps(6) wfE_eqI.hyps(7) wfE_eqI.prems(1) wfE_eqI.prems(2) wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.wfE_eqI wfV_subst) (* 31 ms *)
+ qed
+next
+ case (wfE_fstI \<Theta> \<Gamma> v1 b1 b2)
+ then show ?case using subst_vv.simps subst_ev.simps wf_subst1 Wellformed.wfE_fstI
+ proof -
+ show ?thesis
+ by (metis (full_types) subst_ev.simps(5) wfE_fstI.hyps(1) wfE_fstI.hyps(4) wfE_fstI.hyps(5) wfE_fstI.prems(1) wfE_fstI.prems(2) wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.wfE_fstI wf_subst1(1)) (* 78 ms *)
+ qed
+next
+ case (wfE_sndI \<Theta> \<Gamma> v1 b1 b2)
+ then show ?case
+ by (metis (full_types) subst_ev.simps wfE_sndI Wellformed.wfE_sndI wf_subst1(1))
+next
+ case (wfE_concatI \<Theta> \<Phi> \<Gamma> \<Delta> v1 v2)
+ then show ?case
+ by (metis (full_types) subst_ev.simps wfE_sndI Wellformed.wfE_concatI wf_subst1(1))
+next
+ case (wfE_splitI \<Theta> \<Phi> \<Gamma> \<Delta> v1 v2)
+ then show ?case
+ by (metis (full_types) subst_ev.simps wfE_sndI Wellformed.wfE_splitI wf_subst1(1))
+next
+ case (wfE_lenI \<Theta> \<Phi> \<Gamma> \<Delta> v1)
+then show ?case
+ by (metis (full_types) subst_ev.simps wfE_sndI Wellformed.wfE_lenI wf_subst1(1))
+next
+ case (wfE_appI \<Theta> \<Phi> \<Gamma> \<Delta> f x b c \<tau> s' v)
+then show ?case
+ by (metis (full_types) subst_ev.simps wfE_sndI Wellformed.wfE_appI wf_subst1(1))
+next
+ case (wfE_appPI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> b' bv1 v1 \<tau>1 f1 x1 b1 c1 s1)
+ show ?case proof(subst subst_ev.simps, rule)
+ show "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>" using wfE_appPI wfX_wfY by metis
+ show "\<Theta>; \<B>; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v \<turnstile>\<^sub>w\<^sub>f \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v " using wfE_appPI by auto
+ show "Some (AF_fundef f1 (AF_fun_typ_some bv1 (AF_fun_typ x1 b1 c1 \<tau>1 s1))) = lookup_fun \<Phi> f1" using wfE_appPI by auto
+ show "\<Theta>; \<B>; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v \<turnstile>\<^sub>w\<^sub>f v1[x::=v']\<^sub>v\<^sub>v : b1[bv1::=b']\<^sub>b " using wfE_appPI wf_subst1 by auto
+ show "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b' " using wfE_appPI by auto
+ have "atom bv1 \<sharp> \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v" using fresh_subst_gv_if wfE_appPI by metis
+ moreover have "atom bv1 \<sharp> v1[x::=v']\<^sub>v\<^sub>v" using wfE_appPI fresh_subst_if by simp
+ moreover have "atom bv1 \<sharp> \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v" using wfE_appPI fresh_subst_dv_if by simp
+ ultimately show "atom bv1 \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v, \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v, b', v1[x::=v']\<^sub>v\<^sub>v, (b_of \<tau>1)[bv1::=b']\<^sub>b)"
+ using wfE_appPI fresh_prodN by metis
+ qed
+next
+ case (wfE_mvarI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> u \<tau>)
+ have " \<Theta> ; \<Phi> ; \<B> ; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v \<turnstile>\<^sub>w\<^sub>f (AE_mvar u) : b_of \<tau>[x::=v']\<^sub>\<tau>\<^sub>v" proof
+ show "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> " using wfE_mvarI by auto
+ show "\<Theta>; \<B> ; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v \<turnstile>\<^sub>w\<^sub>f \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v " using wfE_mvarI by auto
+ show "(u, \<tau>[x::=v']\<^sub>\<tau>\<^sub>v) \<in> setD \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v" using wfE_mvarI subst_dv_member by auto
+ qed
+ thus ?case using subst_ev.simps b_of_subst by auto
+next
+ case (wfD_emptyI \<Theta> \<Gamma>)
+ then show ?case using subst_dv.simps wf_intros wf_subst1 by auto
+next
+ case (wfD_cons \<Theta> \<B> \<Gamma> \<Delta> \<tau> u)
+ moreover hence "u \<notin> fst ` setD \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v" using subst_dv.simps subst_dv_iff using subst_dv_fst_eq by presburger
+ ultimately show ?case using subst_dv.simps Wellformed.wfD_cons wf_subst1 by auto
+next
+ case (wfPhi_emptyI \<Theta>)
+ then show ?case by auto
+next
+ case (wfPhi_consI f \<Theta> \<Phi> ft)
+ then show ?case by auto
+next
+ case (wfS_assertI \<Theta> \<Phi> \<B> x2 c \<Gamma> \<Delta> s b)
+ show ?case unfolding subst_sv.simps proof
+ show \<open> \<Theta>; \<Phi>; \<B>; (x2, B_bool, c[x::=v']\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v \<turnstile>\<^sub>w\<^sub>f s[x::=v']\<^sub>s\<^sub>v : b \<close>
+ using wfS_assertI(4)[of "(x2, B_bool, c) #\<^sub>\<Gamma> \<Gamma>\<^sub>1" x ] wfS_assertI by auto
+
+ show \<open> \<Theta>; \<B>; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v \<turnstile>\<^sub>w\<^sub>f c[x::=v']\<^sub>c\<^sub>v \<close> using wfS_assertI wf_subst1 by auto
+ show \<open> \<Theta>; \<B>; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v \<turnstile>\<^sub>w\<^sub>f \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v \<close> using wfS_assertI wf_subst1 by auto
+ show \<open>atom x2 \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v, \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v, c[x::=v']\<^sub>c\<^sub>v, b, s[x::=v']\<^sub>s\<^sub>v)\<close>
+ apply(unfold fresh_prodN,intro conjI)
+ apply(simp add: wfS_assertI )+
+ apply(metis fresh_subst_gv_if wfS_assertI)
+ apply(simp add: fresh_prodN fresh_subst_dv_if wfS_assertI)
+ apply(simp add: fresh_prodN fresh_subst_v_if subst_v_e_def wfS_assertI)
+ apply(simp add: fresh_prodN fresh_subst_v_if subst_v_\<tau>_def wfS_assertI)
+ by(simp add: fresh_prodN fresh_subst_v_if subst_v_s_def wfS_assertI)
+ qed
+next
+ case (wfS_letI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> e b1 y s b2)
+ have "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v \<turnstile>\<^sub>w\<^sub>f LET y = (e[x::=v']\<^sub>e\<^sub>v) IN (s[x::=v']\<^sub>s\<^sub>v) : b2"
+ proof
+ show \<open> \<Theta> ; \<Phi> ; \<B> ; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v \<turnstile>\<^sub>w\<^sub>f e[x::=v']\<^sub>e\<^sub>v : b1 \<close> using wfS_letI by auto
+ have \<open> \<Theta> ; \<Phi> ; \<B> ; ((y, b1, TRUE) #\<^sub>\<Gamma> \<Gamma>)[x::=v']\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v \<turnstile>\<^sub>w\<^sub>f s[x::=v']\<^sub>s\<^sub>v : b2 \<close>
+ using wfS_letI(6) wfS_letI append_g.simps by metis
+ thus \<open> \<Theta> ; \<Phi> ; \<B> ; (y, b1, TRUE) #\<^sub>\<Gamma> \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v \<turnstile>\<^sub>w\<^sub>f s[x::=v']\<^sub>s\<^sub>v : b2 \<close>
+ using wfS_letI subst_gv.simps by auto
+ show \<open> \<Theta>; \<B>; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v \<turnstile>\<^sub>w\<^sub>f \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v \<close> using wfS_letI by auto
+ show \<open>atom y \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v, \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v, e[x::=v']\<^sub>e\<^sub>v, b2)\<close>
+ apply(unfold fresh_prodN,intro conjI)
+ apply(simp add: wfS_letI )+
+ apply(metis fresh_subst_gv_if wfS_letI)
+ apply(simp add: fresh_prodN fresh_subst_dv_if wfS_letI)
+ apply(simp add: fresh_prodN fresh_subst_v_if subst_v_e_def wfS_letI)
+ apply(simp add: fresh_prodN fresh_subst_v_if subst_v_\<tau>_def wfS_letI)
+ done
+ qed
+ thus ?case using subst_sv.simps wfS_letI by auto
+next
+ case (wfS_let2I \<Theta> \<Phi> \<B> \<Gamma> \<Delta> s1 \<tau> y s2 b)
+ have "\<Theta> ; \<Phi> ; \<B> ; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v \<turnstile>\<^sub>w\<^sub>f LET y : \<tau>[x::=v']\<^sub>\<tau>\<^sub>v = (s1[x::=v']\<^sub>s\<^sub>v) IN (s2[x::=v']\<^sub>s\<^sub>v) : b"
+ proof
+ show \<open> \<Theta> ; \<Phi> ; \<B> ; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v \<turnstile>\<^sub>w\<^sub>f s1[x::=v']\<^sub>s\<^sub>v : b_of (\<tau>[x::=v']\<^sub>\<tau>\<^sub>v) \<close> using wfS_let2I b_of_subst by simp
+ have \<open> \<Theta> ; \<Phi> ; \<B> ; ((y, b_of \<tau>, TRUE) #\<^sub>\<Gamma> \<Gamma>)[x::=v']\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v \<turnstile>\<^sub>w\<^sub>f s2[x::=v']\<^sub>s\<^sub>v : b \<close>
+ using wfS_let2I append_g.simps by metis
+ thus \<open> \<Theta> ; \<Phi> ; \<B> ; (y, b_of \<tau>[x::=v']\<^sub>\<tau>\<^sub>v, TRUE) #\<^sub>\<Gamma> \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v \<turnstile>\<^sub>w\<^sub>f s2[x::=v']\<^sub>s\<^sub>v : b \<close>
+ using wfS_let2I subst_gv.simps append_g.simps using b_of_subst by simp
+ show \<open> \<Theta>; \<B>; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v \<turnstile>\<^sub>w\<^sub>f \<tau>[x::=v']\<^sub>\<tau>\<^sub>v \<close> using wfS_let2I wf_subst1 by metis
+ show \<open>atom y \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v, \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v, s1[x::=v']\<^sub>s\<^sub>v, b, \<tau>[x::=v']\<^sub>\<tau>\<^sub>v)\<close>
+ apply(unfold fresh_prodN,intro conjI)
+ apply(simp add: wfS_let2I )+
+ apply(metis fresh_subst_gv_if wfS_let2I)
+ apply(simp add: fresh_prodN fresh_subst_dv_if wfS_let2I)
+ apply(simp add: fresh_prodN fresh_subst_v_if subst_v_e_def wfS_let2I)
+ apply(simp add: fresh_prodN fresh_subst_v_if subst_v_\<tau>_def wfS_let2I)+
+ done
+ qed
+ thus ?case using subst_sv.simps(3) subst_tv.simps wfS_let2I by auto
+next
+ case (wfS_varI \<Theta> \<B> \<Gamma> \<tau> v u \<Phi> \<Delta> b s)
+ show ?case proof(subst subst_sv.simps, auto simp add: u_fresh_xv,rule)
+ show \<open> \<Theta>; \<B>; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v \<turnstile>\<^sub>w\<^sub>f \<tau>[x::=v']\<^sub>\<tau>\<^sub>v \<close> using wfS_varI wf_subst1 by auto
+ have "b_of (\<tau>[x::=v']\<^sub>\<tau>\<^sub>v) = b_of \<tau>" using b_of_subst by auto
+ thus \<open> \<Theta>; \<B>; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v \<turnstile>\<^sub>w\<^sub>f v[x::=v']\<^sub>v\<^sub>v : b_of \<tau>[x::=v']\<^sub>\<tau>\<^sub>v \<close> using wfS_varI wf_subst1 by auto
+ have *:"atom u \<sharp> v'" using wfV_supp wfS_varI fresh_def by metis
+ show \<open>atom u \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v, \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v, \<tau>[x::=v']\<^sub>\<tau>\<^sub>v, v[x::=v']\<^sub>v\<^sub>v, b)\<close>
+ unfolding fresh_prodN apply(auto simp add: wfS_varI)
+ using wfS_varI fresh_subst_gv * fresh_subst_dv by metis+
+ show \<open> \<Theta> ; \<Phi> ; \<B> ; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v ; (u, \<tau>[x::=v']\<^sub>\<tau>\<^sub>v) #\<^sub>\<Delta> \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v \<turnstile>\<^sub>w\<^sub>f s[x::=v']\<^sub>s\<^sub>v : b \<close> using wfS_varI by auto
+ qed
+next
+ case (wfS_assignI u \<tau> \<Delta> \<Theta> \<B> \<Gamma> \<Phi> v)
+ show ?case proof(subst subst_sv.simps, rule wf_intros)
+ show \<open>(u, \<tau>[x::=v']\<^sub>\<tau>\<^sub>v) \<in> setD \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v\<close> using subst_dv_iff wfS_assignI using subst_dv_fst_eq
+ using subst_dv_member by auto
+ show \<open> \<Theta>; \<B>; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v \<turnstile>\<^sub>w\<^sub>f \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v \<close> using wfS_assignI by auto
+ show \<open> \<Theta>; \<B>; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v \<turnstile>\<^sub>w\<^sub>f v[x::=v']\<^sub>v\<^sub>v : b_of \<tau>[x::=v']\<^sub>\<tau>\<^sub>v \<close> using wfS_assignI b_of_subst wf_subst1 by auto
+ show "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> " using wfS_assignI by auto
+ qed
+next
+ case (wfS_matchI \<Theta> \<B> \<Gamma> v tid dclist \<Delta> \<Phi> cs b)
+ show ?case proof(subst subst_sv.simps, rule wf_intros)
+ show \<open> \<Theta>; \<B>; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v \<turnstile>\<^sub>w\<^sub>f v[x::=v']\<^sub>v\<^sub>v : B_id tid \<close> using wfS_matchI wf_subst1 by auto
+ show \<open>AF_typedef tid dclist \<in> set \<Theta>\<close> using wfS_matchI by auto
+ show \<open> \<Theta> ; \<Phi> ; \<B> ; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v ; tid ; dclist \<turnstile>\<^sub>w\<^sub>f subst_branchlv cs x v' : b \<close> using wfS_matchI by simp
+ show "\<Theta>; \<B>; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v \<turnstile>\<^sub>w\<^sub>f \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v " using wfS_matchI by auto
+ show "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> " using wfS_matchI by auto
+ qed
+next
+ case (wfS_branchI \<Theta> \<Phi> \<B> y \<tau> \<Gamma> \<Delta> s b tid dc)
+ have " \<Theta> ; \<Phi> ; \<B> ; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v ; tid ; dc ; \<tau> \<turnstile>\<^sub>w\<^sub>f dc y \<Rightarrow> (s[x::=v']\<^sub>s\<^sub>v) : b"
+ proof
+ have \<open> \<Theta> ; \<Phi> ; \<B> ; ((y, b_of \<tau>, TRUE) #\<^sub>\<Gamma> \<Gamma>)[x::=v']\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v \<turnstile>\<^sub>w\<^sub>f s[x::=v']\<^sub>s\<^sub>v : b \<close>
+ using wfS_branchI append_g.simps by metis
+ thus \<open> \<Theta> ; \<Phi> ; \<B> ; (y, b_of \<tau>, TRUE) #\<^sub>\<Gamma> \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v ; \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v \<turnstile>\<^sub>w\<^sub>f s[x::=v']\<^sub>s\<^sub>v : b \<close>
+ using subst_gv.simps b_of_subst wfS_branchI by simp
+ show \<open>atom y \<sharp> (\<Phi>, \<Theta>, \<B>, \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v, \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v, \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v, \<tau>)\<close>
+ apply(unfold fresh_prodN,intro conjI)
+ apply(simp add: wfS_branchI )+
+ apply(metis fresh_subst_gv_if wfS_branchI)
+ apply(simp add: fresh_prodN fresh_subst_dv_if wfS_branchI)
+ apply(metis fresh_subst_gv_if wfS_branchI)+
+ done
+ show \<open> \<Theta>; \<B>; \<Gamma>[x::=v']\<^sub>\<Gamma>\<^sub>v \<turnstile>\<^sub>w\<^sub>f \<Delta>[x::=v']\<^sub>\<Delta>\<^sub>v \<close> using wfS_branchI by auto
+ qed
+ thus ?case using subst_branchv.simps wfS_branchI by auto
+next
+ case (wfS_finalI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist' cs b dclist)
+ then show ?case using subst_branchlv.simps wf_intros by metis
+next
+ case (wfS_cons \<Theta> \<Phi> \<B> \<Gamma> \<Delta> tid dclist' cs b css dclist)
+ then show ?case using subst_branchlv.simps wf_intros by metis
+
+qed(metis subst_sv.simps wf_subst1 wf_intros)+
+
+lemmas wf_subst = wf_subst1 wf_subst2
+
+lemma wfG_subst_wfV:
+ assumes "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>' @ (x, b, c0[z0::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\<Gamma> \<Gamma>" and "wfV \<Theta> \<B> \<Gamma> v b"
+ shows "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v @ \<Gamma> "
+ using assms wf_subst subst_g_inside_simple by auto
+
+lemma wfG_member_subst:
+ assumes "(x1,b1,c1) \<in> toSet (\<Gamma>'@\<Gamma>)" and "wfG \<Theta> \<B> (\<Gamma>'@((x,b,c) #\<^sub>\<Gamma>\<Gamma>))" and "x \<noteq> x1"
+ shows "\<exists>c1'. (x1,b1,c1') \<in> toSet ((\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v)@\<Gamma>)"
+proof -
+ consider (lhs) "(x1,b1,c1) \<in> toSet \<Gamma>'" | (rhs) "(x1,b1,c1) \<in> toSet \<Gamma>" using append_g_toSetU assms by auto
+ thus ?thesis proof(cases)
+ case lhs
+ hence "(x1,b1,c1[x::=v]\<^sub>c\<^sub>v) \<in> toSet (\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v)" using wfG_inside_fresh[THEN subst_gv_member_iff[OF lhs]] assms by metis
+ hence "(x1,b1,c1[x::=v]\<^sub>c\<^sub>v) \<in> toSet (\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v@\<Gamma>)" using append_g_toSetU by auto
+ then show ?thesis by auto
+ next
+ case rhs
+ hence "(x1,b1,c1) \<in> toSet (\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v@\<Gamma>)" using append_g_toSetU by auto
+ then show ?thesis by auto
+ qed
+qed
+
+lemma wfG_member_subst2:
+ assumes "(x1,b1,c1) \<in> toSet (\<Gamma>'@((x,b,c) #\<^sub>\<Gamma>\<Gamma>))" and "wfG \<Theta> \<B> (\<Gamma>'@((x,b,c) #\<^sub>\<Gamma>\<Gamma>))" and "x \<noteq> x1"
+ shows "\<exists>c1'. (x1,b1,c1') \<in> toSet ((\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v)@\<Gamma>)"
+proof -
+ consider (lhs) "(x1,b1,c1) \<in> toSet \<Gamma>'" | (rhs) "(x1,b1,c1) \<in> toSet \<Gamma>" using append_g_toSetU assms by auto
+ thus ?thesis proof(cases)
+ case lhs
+ hence "(x1,b1,c1[x::=v]\<^sub>c\<^sub>v) \<in> toSet (\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v)" using wfG_inside_fresh[THEN subst_gv_member_iff[OF lhs]] assms by metis
+ hence "(x1,b1,c1[x::=v]\<^sub>c\<^sub>v) \<in> toSet (\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v@\<Gamma>)" using append_g_toSetU by auto
+ then show ?thesis by auto
+ next
+ case rhs
+ hence "(x1,b1,c1) \<in> toSet (\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v@\<Gamma>)" using append_g_toSetU by auto
+ then show ?thesis by auto
+ qed
+qed
+
+lemma wbc_subst:
+ fixes \<Gamma>::\<Gamma> and \<Gamma>'::\<Gamma> and v::v
+ assumes "wfC \<Theta> \<B> (\<Gamma>'@((x,b,c') #\<^sub>\<Gamma>\<Gamma>)) c" and "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b"
+ shows "\<Theta>; \<B>; ((\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v)@\<Gamma>) \<turnstile>\<^sub>w\<^sub>f c[x::=v]\<^sub>c\<^sub>v"
+proof -
+ have "(\<Gamma>'@((x,b,c') #\<^sub>\<Gamma>\<Gamma>))[x::=v]\<^sub>\<Gamma>\<^sub>v = ((\<Gamma>'[x::=v]\<^sub>\<Gamma>\<^sub>v)@\<Gamma>)" using assms subst_g_inside_simple wfC_wf by metis
+ thus ?thesis using wf_subst1(2)[OF assms(1) _ assms(2)] by metis
+qed
+
+lemma wfG_inside_fresh_suffix:
+ assumes "wfG P B (\<Gamma>'@(x,b,c) #\<^sub>\<Gamma>\<Gamma>)"
+ shows "atom x \<sharp> \<Gamma>"
+proof -
+ have "wfG P B ((x,b,c) #\<^sub>\<Gamma>\<Gamma>)" using wfG_suffix assms by auto
+ thus ?thesis using wfG_elims by metis
+qed
+
+lemmas wf_b_subst_lemmas = subst_eb.simps wf_intros
+ forget_subst subst_b_b_def subst_b_v_def subst_b_ce_def fresh_e_opp_all subst_bb.simps wfV_b_fresh ms_fresh_all(6)
+
+lemma wf_b_subst1:
+ fixes \<Gamma>::\<Gamma> and \<Gamma>'::\<Gamma> and v::v and e::e and c::c and \<tau>::\<tau> and ts::"(string*\<tau>) list" and \<Delta>::\<Delta> and b::b and ftq::fun_typ_q and ft::fun_typ and s::s and b'::b and ce::ce and td::type_def
+ and cs::branch_s and css::branch_list
+ shows "\<Theta> ; B' ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b' \<Longrightarrow> {|bv|} = B' \<Longrightarrow> \<Theta> ; B \<turnstile>\<^sub>w\<^sub>f b \<Longrightarrow> \<Theta> ; B ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f v[bv::=b]\<^sub>v\<^sub>b : b'[bv::=b]\<^sub>b\<^sub>b" and
+ "\<Theta> ; B' ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f c \<Longrightarrow> {|bv|} = B' \<Longrightarrow> \<Theta> ; B \<turnstile>\<^sub>w\<^sub>f b \<Longrightarrow> \<Theta> ; B ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f c[bv::=b]\<^sub>c\<^sub>b" and
+ "\<Theta> ; B' \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<Longrightarrow> {|bv|} = B' \<Longrightarrow> \<Theta> ; B \<turnstile>\<^sub>w\<^sub>f b \<Longrightarrow> \<Theta> ; B \<turnstile>\<^sub>w\<^sub>f \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b" and
+ "\<Theta> ; B' ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau> \<Longrightarrow> {|bv|} = B' \<Longrightarrow> \<Theta> ; B \<turnstile>\<^sub>w\<^sub>f b \<Longrightarrow> \<Theta> ; B ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f \<tau>[bv::=b]\<^sub>\<tau>\<^sub>b" and
+ "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ts \<Longrightarrow> True" and
+ "\<turnstile>\<^sub>w\<^sub>f \<Theta> \<Longrightarrow>True" and
+ "\<Theta> ; B' \<turnstile>\<^sub>w\<^sub>f b' \<Longrightarrow> {|bv|} = B' \<Longrightarrow> \<Theta> ; B \<turnstile>\<^sub>w\<^sub>f b \<Longrightarrow> \<Theta> ; B \<turnstile>\<^sub>w\<^sub>f b'[bv::=b]\<^sub>b\<^sub>b " and
+ "\<Theta> ; B' ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ce : b' \<Longrightarrow> {|bv|} = B' \<Longrightarrow> \<Theta> ; B \<turnstile>\<^sub>w\<^sub>f b \<Longrightarrow> \<Theta> ; B ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f ce[bv::=b]\<^sub>c\<^sub>e\<^sub>b : b'[bv::=b]\<^sub>b\<^sub>b" and
+ "\<Theta> \<turnstile>\<^sub>w\<^sub>f td \<Longrightarrow> True"
+proof(nominal_induct
+ b' and c and \<Gamma> and \<tau> and ts and \<Theta> and b' and b' and td
+ avoiding: bv b B
+ rule:wfV_wfC_wfG_wfT_wfTs_wfTh_wfB_wfCE_wfTD.strong_induct)
+ case (wfB_intI \<Theta> \<B>)
+ then show ?case using subst_bb.simps wf_intros wfX_wfY by metis
+next
+ case (wfB_boolI \<Theta> \<B>)
+ then show ?case using subst_bb.simps wf_intros wfX_wfY by metis
+next
+ case (wfB_unitI \<Theta> \<B>)
+ then show ?case using subst_bb.simps wf_intros wfX_wfY by metis
+next
+ case (wfB_bitvecI \<Theta> \<B>)
+ then show ?case using subst_bb.simps wf_intros wfX_wfY by metis
+next
+ case (wfB_pairI \<Theta> \<B> b1 b2)
+ then show ?case using subst_bb.simps wf_intros wfX_wfY by metis
+next
+ case (wfB_consI \<Theta> s dclist \<B>)
+ then show ?case using subst_bb.simps Wellformed.wfB_consI by simp
+next
+ case (wfB_appI \<Theta> ba s bva dclist \<B>)
+ then show ?case using subst_bb.simps Wellformed.wfB_appI forget_subst wfB_supp
+ by (metis bot.extremum_uniqueI ex_in_conv fresh_def subst_b_b_def supp_empty_fset)
+next
+ case (wfV_varI \<Theta> \<B>1 \<Gamma> b1 c x)
+ show ?case unfolding subst_vb.simps proof
+ show "\<Theta> ; B \<turnstile>\<^sub>w\<^sub>f \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b " using wfV_varI by auto
+ show "Some (b1[bv::=b]\<^sub>b\<^sub>b, c[bv::=b]\<^sub>c\<^sub>b) = lookup \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b x" using subst_b_lookup wfV_varI by simp
+ qed
+next
+ case (wfV_litI \<Theta> \<B> \<Gamma> l)
+ then show ?case using Wellformed.wfV_litI subst_b_base_for_lit by simp
+next
+ case (wfV_pairI \<Theta> \<B>1 \<Gamma> v1 b1 v2 b2)
+ show ?case unfolding subst_vb.simps proof(subst subst_bb.simps,rule)
+ show "\<Theta> ; B ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f v1[bv::=b]\<^sub>v\<^sub>b : b1[bv::=b]\<^sub>b\<^sub>b" using wfV_pairI by simp
+ show "\<Theta> ; B ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f v2[bv::=b]\<^sub>v\<^sub>b : b2[bv::=b]\<^sub>b\<^sub>b " using wfV_pairI by simp
+ qed
+next
+ case (wfV_consI s dclist \<Theta> dc x b' c \<B>' \<Gamma> v)
+ show ?case unfolding subst_vb.simps proof(subst subst_bb.simps, rule Wellformed.wfV_consI)
+ show 1:"AF_typedef s dclist \<in> set \<Theta>" using wfV_consI by auto
+ show 2:"(dc, \<lbrace> x : b' | c \<rbrace>) \<in> set dclist" using wfV_consI by auto
+ have "\<Theta> ; B ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f v[bv::=b]\<^sub>v\<^sub>b : b'[bv::=b]\<^sub>b\<^sub>b" using wfV_consI by auto
+ moreover hence "supp b' = {}" using 1 2 wfTh_lookup_supp_empty \<tau>.supp wfX_wfY by blast
+ moreover hence "b'[bv::=b]\<^sub>b\<^sub>b = b'" using forget_subst subst_bb_def fresh_def by (metis empty_iff subst_b_b_def)
+ ultimately show "\<Theta> ; B ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f v[bv::=b]\<^sub>v\<^sub>b : b'" using wfV_consI by simp
+ qed
+next
+ case (wfV_conspI s bva dclist \<Theta> dc x b' c \<B>' ba \<Gamma> v)
+ have *:"atom bv \<sharp> b'" using wfTh_poly_supp_b[of s bva dclist \<Theta> dc x b' c] fresh_def wfX_wfY \<open>atom bva \<sharp> bv\<close>
+ by (metis insert_iff not_self_fresh singleton_insert_inj_eq' subsetI subset_antisym wfV_conspI wfV_conspI.hyps(4) wfV_conspI.prems(2))
+ show ?case unfolding subst_vb.simps subst_bb.simps proof
+ show \<open>AF_typedef_poly s bva dclist \<in> set \<Theta>\<close> using wfV_conspI by auto
+ show \<open>(dc, \<lbrace> x : b' | c \<rbrace>) \<in> set dclist\<close> using wfV_conspI by auto
+ thus \<open> \<Theta> ; B \<turnstile>\<^sub>w\<^sub>f ba[bv::=b]\<^sub>b\<^sub>b \<close> using wfV_conspI by metis
+ have "atom bva \<sharp> \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b" using fresh_subst_if subst_b_\<Gamma>_def wfV_conspI by metis
+ moreover have "atom bva \<sharp> ba[bv::=b]\<^sub>b\<^sub>b" using fresh_subst_if subst_b_b_def wfV_conspI by metis
+ moreover have "atom bva \<sharp> v[bv::=b]\<^sub>v\<^sub>b" using fresh_subst_if subst_b_v_def wfV_conspI by metis
+ ultimately show \<open>atom bva \<sharp> (\<Theta>, B, \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b, ba[bv::=b]\<^sub>b\<^sub>b, v[bv::=b]\<^sub>v\<^sub>b)\<close>
+ unfolding fresh_prodN using wfV_conspI fresh_def supp_fset by auto
+ show \<open> \<Theta> ; B ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f v[bv::=b]\<^sub>v\<^sub>b : b'[bva::=ba[bv::=b]\<^sub>b\<^sub>b]\<^sub>b\<^sub>b \<close>
+ using wfV_conspI subst_bb_commute[of bv b' bva ba b] * wfV_conspI by metis
+ qed
+next
+ case (wfTI z \<Theta> \<B>' \<Gamma>' b' c)
+ show ?case proof(subst subst_tb.simps, rule Wellformed.wfTI)
+ show "atom z \<sharp> (\<Theta>, B, \<Gamma>'[bv::=b]\<^sub>\<Gamma>\<^sub>b)" using wfTI subst_g_b_x_fresh by simp
+ show "\<Theta> ; B \<turnstile>\<^sub>w\<^sub>f b'[bv::=b]\<^sub>b\<^sub>b " using wfTI by auto
+ show "\<Theta> ; B ; (z, b'[bv::=b]\<^sub>b\<^sub>b, TRUE) #\<^sub>\<Gamma> \<Gamma>'[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f c[bv::=b]\<^sub>c\<^sub>b " using wfTI by simp
+ qed
+next
+ case (wfC_eqI \<Theta> \<B>' \<Gamma> e1 b' e2)
+ thus ?case using Wellformed.wfC_eqI subst_db.simps subst_cb.simps wfC_eqI by metis
+next
+ case (wfG_nilI \<Theta> \<B>')
+ then show ?case using Wellformed.wfG_nilI subst_gb.simps by simp
+next
+ case (wfG_cons1I c' \<Theta> \<B>' \<Gamma>' x b')
+ show ?case proof(subst subst_gb.simps, rule Wellformed.wfG_cons1I)
+ show "c'[bv::=b]\<^sub>c\<^sub>b \<notin> {TRUE, FALSE}" using wfG_cons1I(1)
+ by(nominal_induct c' rule: c.strong_induct,auto+)
+ show "\<Theta> ; B \<turnstile>\<^sub>w\<^sub>f \<Gamma>'[bv::=b]\<^sub>\<Gamma>\<^sub>b " using wfG_cons1I by auto
+ show "atom x \<sharp> \<Gamma>'[bv::=b]\<^sub>\<Gamma>\<^sub>b" using wfG_cons1I subst_g_b_x_fresh by auto
+ show "\<Theta> ; B ; (x, b'[bv::=b]\<^sub>b\<^sub>b, TRUE) #\<^sub>\<Gamma> \<Gamma>'[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f c'[bv::=b]\<^sub>c\<^sub>b" using wfG_cons1I by auto
+ show "\<Theta> ; B \<turnstile>\<^sub>w\<^sub>f b'[bv::=b]\<^sub>b\<^sub>b " using wfG_cons1I by auto
+ qed
+next
+ case (wfG_cons2I c' \<Theta> \<B>' \<Gamma>' x b')
+ show ?case proof(subst subst_gb.simps, rule Wellformed.wfG_cons2I)
+ show "c'[bv::=b]\<^sub>c\<^sub>b \<in> {TRUE, FALSE}" using wfG_cons2I by auto
+ show "\<Theta> ; B \<turnstile>\<^sub>w\<^sub>f \<Gamma>'[bv::=b]\<^sub>\<Gamma>\<^sub>b " using wfG_cons2I by auto
+ show "atom x \<sharp> \<Gamma>'[bv::=b]\<^sub>\<Gamma>\<^sub>b" using wfG_cons2I subst_g_b_x_fresh by auto
+ show "\<Theta> ; B \<turnstile>\<^sub>w\<^sub>f b'[bv::=b]\<^sub>b\<^sub>b " using wfG_cons2I by auto
+ qed
+next
+ case (wfCE_valI \<Theta> \<B> \<Gamma> v b)
+ then show ?case using subst_ceb.simps wf_intros wfX_wfY
+ by (metis wf_b_subst_lemmas wfCE_b_fresh)
+next
+ case (wfCE_plusI \<Theta> \<B> \<Gamma> v1 v2)
+ then show ?case using subst_bb.simps subst_ceb.simps wf_intros wfX_wfY
+ by metis
+next
+ case (wfCE_leqI \<Theta> \<B> \<Gamma> v1 v2)
+ then show ?case using subst_bb.simps subst_ceb.simps wf_intros wfX_wfY
+ by metis
+next
+ case (wfCE_eqI \<Theta> \<B> \<Gamma> v1 b v2)
+ then show ?case using subst_bb.simps subst_ceb.simps wf_intros wfX_wfY
+ by metis
+next
+ case (wfCE_fstI \<Theta> \<B> \<Gamma> v1 b1 b2)
+ then show ?case
+ by (metis (no_types) subst_bb.simps(5) subst_ceb.simps(3) wfCE_fstI.hyps(2)
+ wfCE_fstI.prems(1) wfCE_fstI.prems(2) Wellformed.wfCE_fstI)
+next
+ case (wfCE_sndI \<Theta> \<B> \<Gamma> v1 b1 b2)
+ then show ?case
+ by (metis (no_types) subst_bb.simps(5) subst_ceb.simps wfCE_sndI.hyps(2)
+ wfCE_sndI wfCE_sndI.prems(2) Wellformed.wfCE_sndI)
+next
+ case (wfCE_concatI \<Theta> \<B> \<Gamma> v1 v2)
+ then show ?case using subst_bb.simps subst_ceb.simps wf_intros wfX_wfY wf_b_subst_lemmas wfCE_b_fresh
+ proof -
+ show ?thesis
+ using wfCE_concatI.hyps(2) wfCE_concatI.hyps(4) wfCE_concatI.prems(1) wfCE_concatI.prems(2)
+ Wellformed.wfCE_concatI by auto (* 46 ms *)
+ qed
+next
+ case (wfCE_lenI \<Theta> \<B> \<Gamma> v1)
+ then show ?case using subst_bb.simps subst_ceb.simps wf_intros wfX_wfY wf_b_subst_lemmas wfCE_b_fresh by metis
+qed(auto simp add: wf_intros)
+
+lemma wf_b_subst2:
+ fixes \<Gamma>::\<Gamma> and \<Gamma>'::\<Gamma> and v::v and e::e and c::c and \<tau>::\<tau> and ts::"(string*\<tau>) list" and \<Delta>::\<Delta> and b::b and ftq::fun_typ_q and ft::fun_typ and s::s and b'::b and ce::ce and td::type_def
+ and cs::branch_s and css::branch_list
+ shows "\<Theta> ; \<Phi> ; B' ; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f e : b' \<Longrightarrow> {|bv|} = B' \<Longrightarrow> \<Theta> ; B \<turnstile>\<^sub>w\<^sub>f b \<Longrightarrow> \<Theta> ; \<Phi> ; B ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b ; \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b \<turnstile>\<^sub>w\<^sub>f e[bv::=b]\<^sub>e\<^sub>b : b'[bv::=b]\<^sub>b\<^sub>b" and
+ "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> \<turnstile>\<^sub>w\<^sub>f s : b \<Longrightarrow> True" and
+ "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> ; tid ; dc ; t \<turnstile>\<^sub>w\<^sub>f cs : b \<Longrightarrow> True" and
+ "\<Theta> ; \<Phi> ; \<B> ; \<Gamma> ; \<Delta> ; tid ; dclist \<turnstile>\<^sub>w\<^sub>f css : b \<Longrightarrow> True" and
+ "\<Theta> \<turnstile>\<^sub>w\<^sub>f (\<Phi>::\<Phi>) \<Longrightarrow> True " and
+ "\<Theta> ; B' ; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Delta> \<Longrightarrow> {|bv|} = B' \<Longrightarrow> \<Theta> ; B \<turnstile>\<^sub>w\<^sub>f b \<Longrightarrow> \<Theta> ; B ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b" and
+ "\<Theta> ; \<Phi> \<turnstile>\<^sub>w\<^sub>f ftq \<Longrightarrow> True" and
+ "\<Theta> ; \<Phi> ; \<B> \<turnstile>\<^sub>w\<^sub>f ft \<Longrightarrow> True"
+proof(nominal_induct
+ b' and b and b and b and \<Phi> and \<Delta> and ftq and ft
+ avoiding: bv b B
+rule:wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.strong_induct)
+ case (wfE_valI \<Theta>' \<Phi>' \<B>' \<Gamma>' \<Delta>' v' b')
+ then show ?case unfolding subst_vb.simps subst_eb.simps using wf_b_subst1(1) Wellformed.wfE_valI by auto
+next
+ case (wfE_plusI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 v2)
+ then show ?case unfolding subst_eb.simps
+ using wf_b_subst_lemmas wf_b_subst1(1) Wellformed.wfE_plusI
+ proof -
+ have "\<forall>b ba v g f ts. (( ts ; f ; g[bv::=ba]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f v[bv::=ba]\<^sub>v\<^sub>b : b[bv::=ba]\<^sub>b\<^sub>b) \<or> \<not> ts ; \<B> ; g \<turnstile>\<^sub>w\<^sub>f v : b ) \<or> \<not> ts ; f \<turnstile>\<^sub>w\<^sub>f ba"
+ using wfE_plusI.prems(1) wf_b_subst1(1) by force (* 0.0 ms *)
+ then show "\<Theta> ; \<Phi> ; B ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b ; \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b \<turnstile>\<^sub>w\<^sub>f [ plus v1[bv::=b]\<^sub>v\<^sub>b v2[bv::=b]\<^sub>v\<^sub>b ]\<^sup>e : B_int[bv::=b]\<^sub>b\<^sub>b"
+
+ by (metis wfE_plusI.hyps(1) wfE_plusI.hyps(4) wfE_plusI.hyps(5) wfE_plusI.hyps(6) wfE_plusI.prems(1) wfE_plusI.prems(2) wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.wfE_plusI wf_b_subst_lemmas(86))
+ qed
+next
+ case (wfE_leqI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 v2)
+ then show ?case unfolding subst_eb.simps
+ using wf_b_subst_lemmas wf_b_subst1 Wellformed.wfE_leqI
+ proof -
+ have "\<And>ts f b ba g v. \<not> (ts ; f \<turnstile>\<^sub>w\<^sub>f b) \<or> \<not> (ts ; {|ba|} ; g \<turnstile>\<^sub>w\<^sub>f v : B_int) \<or> (ts ; f ; g[ba::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f v[ba::=b]\<^sub>v\<^sub>b : B_int)"
+ by (metis wf_b_subst1(1) wf_b_subst_lemmas(86)) (* 46 ms *)
+ then show "\<Theta> ; \<Phi> ; B ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b ; \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b \<turnstile>\<^sub>w\<^sub>f [ leq v1[bv::=b]\<^sub>v\<^sub>b v2[bv::=b]\<^sub>v\<^sub>b ]\<^sup>e : B_bool[bv::=b]\<^sub>b\<^sub>b"
+ by (metis (no_types) wfE_leqI.hyps(1) wfE_leqI.hyps(4) wfE_leqI.hyps(5) wfE_leqI.hyps(6) wfE_leqI.prems(1) wfE_leqI.prems(2) wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.wfE_leqI wf_b_subst_lemmas(87)) (* 46 ms *)
+ qed
+next
+ case (wfE_eqI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 bb v2)
+ show ?case unfolding subst_eb.simps subst_bb.simps proof
+ show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using wfX_wfY wfE_eqI by metis
+ show \<open> \<Theta> ; B ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b \<close> using wfX_wfY wfE_eqI by metis
+ show \<open> \<Theta> ; B ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f v1[bv::=b]\<^sub>v\<^sub>b : bb \<close> using subst_bb.simps wfE_eqI
+ by (metis (no_types, opaque_lifting) empty_iff insert_iff wf_b_subst1(1))
+ show \<open> \<Theta> ; B ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f v2[bv::=b]\<^sub>v\<^sub>b : bb \<close> using wfX_wfY wfE_eqI
+ by (metis insert_iff singleton_iff wf_b_subst1(1) wf_b_subst_lemmas(86) wf_b_subst_lemmas(87) wf_b_subst_lemmas(90))
+ show \<open>bb \<in> {B_bool, B_int, B_unit}\<close> using wfE_eqI by auto
+ qed
+next
+ case (wfE_fstI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 b1 b2)
+ then show ?case unfolding subst_eb.simps using wf_b_subst_lemmas(84) wf_b_subst1(1) Wellformed.wfE_fstI
+ by (metis wf_b_subst_lemmas(89))
+next
+ case (wfE_sndI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 b1 b2)
+ then show ?case unfolding subst_eb.simps using wf_b_subst_lemmas(86) wf_b_subst1(1) Wellformed.wfE_sndI
+ by (metis wf_b_subst_lemmas(89))
+next
+ case (wfE_concatI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 v2)
+then show ?case unfolding subst_eb.simps using wf_b_subst_lemmas(86) wf_b_subst1(1) Wellformed.wfE_concatI
+ by (metis wf_b_subst_lemmas(91))
+next
+ case (wfE_splitI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1 v2)
+ then show ?case unfolding subst_eb.simps using wf_b_subst_lemmas(86) wf_b_subst1(1) Wellformed.wfE_splitI
+ by (metis wf_b_subst_lemmas(89) wf_b_subst_lemmas(91))
+next
+ case (wfE_lenI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> v1)
+ then show ?case unfolding subst_eb.simps using wf_b_subst_lemmas(86) wf_b_subst1(1) Wellformed.wfE_lenI
+ by (metis wf_b_subst_lemmas(91) wf_b_subst_lemmas(89))
+next
+ case (wfE_appI \<Theta> \<Phi> \<B>' \<Gamma> \<Delta> f x b' c \<tau> s v)
+ hence bf: "atom bv \<sharp> b'" using wfPhi_f_simple_wfT wfT_supp bv_not_in_dom_g wfPhi_f_simple_supp_b fresh_def by fast
+ hence bseq: "b'[bv::=b]\<^sub>b\<^sub>b = b'" using subst_bb.simps wf_b_subst_lemmas by metis
+ have "\<Theta> ; \<Phi> ; B ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b ; \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b \<turnstile>\<^sub>w\<^sub>f (AE_app f (v[bv::=b]\<^sub>v\<^sub>b)) : (b_of (\<tau>[bv::=b]\<^sub>\<tau>\<^sub>b))"
+ proof
+ show "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi>" using wfE_appI by auto
+ show "\<Theta> ; B ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b " using wfE_appI by simp
+ have "atom bv \<sharp> \<tau>" using wfPhi_f_simple_wfT[OF wfE_appI(5) wfE_appI(1),THEN wfT_supp] bv_not_in_dom_g fresh_def by force
+ hence " \<tau>[bv::=b]\<^sub>\<tau>\<^sub>b = \<tau>" using forget_subst subst_b_\<tau>_def by metis
+ thus "Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x b' c \<tau>[bv::=b]\<^sub>\<tau>\<^sub>b s))) = lookup_fun \<Phi> f" using wfE_appI by simp
+ show "\<Theta> ; B ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f v[bv::=b]\<^sub>v\<^sub>b : b'" using wfE_appI bseq wf_b_subst1 by metis
+ qed
+ then show ?case using subst_eb.simps b_of_subst_bb_commute by simp
+next
+ case (wfE_appPI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> b' bv1 v1 \<tau>1 f x1 b1 c1 s1)
+ then have *: "atom bv \<sharp> b1" using wfPhi_f_supp(1) wfE_appPI(7,11)
+ by (metis fresh_def fresh_finsert singleton_iff subsetD fresh_def supp_at_base wfE_appPI.hyps(1))
+ have "\<Theta> ; \<Phi> ; B ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b ; \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b \<turnstile>\<^sub>w\<^sub>f AE_appP f b'[bv::=b]\<^sub>b\<^sub>b (v1[bv::=b]\<^sub>v\<^sub>b) : (b_of \<tau>1)[bv1::=b'[bv::=b]\<^sub>b\<^sub>b]\<^sub>b"
+ proof
+ show \<open> \<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> \<close> using wfE_appPI by auto
+ show \<open> \<Theta> ; B ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b \<close> using wfE_appPI by auto
+ show \<open> \<Theta> ; B \<turnstile>\<^sub>w\<^sub>f b'[bv::=b]\<^sub>b\<^sub>b \<close> using wfE_appPI wf_b_subst1 by auto
+ have "atom bv1 \<sharp> \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b" using fresh_subst_if subst_b_\<Gamma>_def wfE_appPI by metis
+ moreover have "atom bv1 \<sharp> b'[bv::=b]\<^sub>b\<^sub>b" using fresh_subst_if subst_b_b_def wfE_appPI by metis
+ moreover have "atom bv1 \<sharp> v1[bv::=b]\<^sub>v\<^sub>b" using fresh_subst_if subst_b_v_def wfE_appPI by metis
+ moreover have "atom bv1 \<sharp> \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b" using fresh_subst_if subst_b_\<Delta>_def wfE_appPI by metis
+ moreover have "atom bv1 \<sharp> (b_of \<tau>1)[bv1::=b'[bv::=b]\<^sub>b\<^sub>b]\<^sub>b\<^sub>b" using fresh_subst_if subst_b_b_def wfE_appPI by metis
+ ultimately show "atom bv1 \<sharp> (\<Phi>, \<Theta>, B, \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b, \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b, b'[bv::=b]\<^sub>b\<^sub>b, v1[bv::=b]\<^sub>v\<^sub>b, (b_of \<tau>1)[bv1::=b'[bv::=b]\<^sub>b\<^sub>b]\<^sub>b)"
+ using wfE_appPI using fresh_def fresh_prodN subst_b_b_def by metis
+ show \<open>Some (AF_fundef f (AF_fun_typ_some bv1 (AF_fun_typ x1 b1 c1 \<tau>1 s1))) = lookup_fun \<Phi> f\<close> using wfE_appPI by auto
+
+ have \<open> \<Theta> ; B ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f v1[bv::=b]\<^sub>v\<^sub>b : b1[bv1::=b']\<^sub>b[bv::=b]\<^sub>b\<^sub>b \<close>
+ using wfE_appPI subst_b_b_def * wf_b_subst1 by metis
+ thus \<open> \<Theta> ; B ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f v1[bv::=b]\<^sub>v\<^sub>b : b1[bv1::=b'[bv::=b]\<^sub>b\<^sub>b]\<^sub>b \<close>
+ using subst_bb_commute subst_b_b_def * by auto
+ qed
+ moreover have "atom bv \<sharp> b_of \<tau>1" proof -
+ have "supp (b_of \<tau>1) \<subseteq> { atom bv1 }" using wfPhi_f_poly_supp_b_of_t
+ using b_of.simps wfE_appPI wfPhi_f_supp(5) by simp
+ thus ?thesis using wfE_appPI
+ fresh_def fresh_finsert singleton_iff subsetD fresh_def supp_at_base wfE_appPI.hyps by metis
+ qed
+ ultimately show ?case using subst_eb.simps(3) subst_bb_commute subst_b_b_def * by simp
+next
+ case (wfE_mvarI \<Theta> \<Phi> \<B>' \<Gamma> \<Delta> u \<tau>)
+
+ have "\<Theta> ; \<Phi> ; B ; subst_gb \<Gamma> bv b ; subst_db \<Delta> bv b \<turnstile>\<^sub>w\<^sub>f (AE_mvar u)[bv::=b]\<^sub>e\<^sub>b : (b_of (\<tau>[bv::=b]\<^sub>\<tau>\<^sub>b))"
+ proof(subst subst_eb.simps,rule Wellformed.wfE_mvarI)
+ show "\<Theta> \<turnstile>\<^sub>w\<^sub>f \<Phi> " using wfE_mvarI by simp
+ show "\<Theta> ; B ; \<Gamma>[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b" using wfE_mvarI by metis
+ show "(u, \<tau>[bv::=b]\<^sub>\<tau>\<^sub>b) \<in> setD \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b"
+ using wfE_mvarI subst_db.simps set_insert subst_d_b_member by simp
+ qed
+ thus ?case using b_of_subst_bb_commute by auto
+
+next
+ case (wfS_seqI \<Theta> \<Phi> \<B> \<Gamma> \<Delta> s1 s2 b)
+ then show ?case using subst_bb.simps wf_intros wfX_wfY by metis
+next
+ case (wfD_emptyI \<Theta> \<B>' \<Gamma>)
+ then show ?case using subst_db.simps Wellformed.wfD_emptyI wf_b_subst1 by simp
+next
+ case (wfD_cons \<Theta> \<B>' \<Gamma>' \<Delta> \<tau> u)
+ show ?case proof(subst subst_db.simps, rule Wellformed.wfD_cons )
+ show "\<Theta> ; B ; \<Gamma>'[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b " using wfD_cons by auto
+ show "\<Theta> ; B ; \<Gamma>'[bv::=b]\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f \<tau>[bv::=b]\<^sub>\<tau>\<^sub>b " using wfD_cons wf_b_subst1 by auto
+ show "u \<notin> fst ` setD \<Delta>[bv::=b]\<^sub>\<Delta>\<^sub>b" using wfD_cons subst_b_lookup_d by metis
+ qed
+next
+ case (wfS_assertI \<Theta> \<Phi> \<B> x c \<Gamma> \<Delta> s b)
+ show ?case by auto
+qed(auto)
+
+lemmas wf_b_subst = wf_b_subst1 wf_b_subst2
+
+lemma wfT_subst_wfT:
+ fixes \<tau>::\<tau> and b'::b and bv::bv
+ assumes "\<Theta> ; {|bv|} ; (x,b,c) #\<^sub>\<Gamma>GNil \<turnstile>\<^sub>w\<^sub>f \<tau>" and "\<Theta> ; B \<turnstile>\<^sub>w\<^sub>f b'"
+ shows "\<Theta> ; B ; (x,b[bv::=b']\<^sub>b\<^sub>b,c[bv::=b']\<^sub>c\<^sub>b) #\<^sub>\<Gamma>GNil \<turnstile>\<^sub>w\<^sub>f (\<tau>[bv::=b']\<^sub>\<tau>\<^sub>b)"
+proof -
+ have "\<Theta> ; B ; ((x,b,c) #\<^sub>\<Gamma>GNil)[bv::=b']\<^sub>\<Gamma>\<^sub>b \<turnstile>\<^sub>w\<^sub>f (\<tau>[bv::=b']\<^sub>\<tau>\<^sub>b)"
+ using wf_b_subst assms by metis
+ thus ?thesis using subst_gb.simps wf_b_subst_lemmas wfCE_b_fresh by metis
+qed
+
+lemma wf_trans:
+ fixes \<Gamma>::\<Gamma> and \<Gamma>'::\<Gamma> and v::v and e::e and c::c and \<tau>::\<tau> and ts::"(string*\<tau>) list" and \<Delta>::\<Delta> and b::b and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def and s::s
+ and cs::branch_s and css::branch_list and \<Theta>::\<Theta>
+ shows "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f v : b' \<Longrightarrow> \<Gamma> = (x, b, c2) #\<^sub>\<Gamma> G \<Longrightarrow> \<Theta>; \<B>; (x, b, c1) #\<^sub>\<Gamma> G \<turnstile>\<^sub>w\<^sub>f c2 \<Longrightarrow> \<Theta>; \<B>; (x, b, c1) #\<^sub>\<Gamma> G \<turnstile>\<^sub>w\<^sub>f v : b'" and
+ "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f c \<Longrightarrow> \<Gamma> = (x, b, c2) #\<^sub>\<Gamma> G \<Longrightarrow> \<Theta>; \<B>; (x, b, c1) #\<^sub>\<Gamma> G \<turnstile>\<^sub>w\<^sub>f c2 \<Longrightarrow> \<Theta>; \<B>; (x, b, c1) #\<^sub>\<Gamma> G \<turnstile>\<^sub>w\<^sub>f c" and
+ "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f \<Gamma> \<Longrightarrow> True" and
+ "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<tau> \<Longrightarrow> True" and
+ "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ts \<Longrightarrow> True" and
+ "\<turnstile>\<^sub>w\<^sub>f \<Theta> \<Longrightarrow>True" and
+ "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b \<Longrightarrow> True " and
+ "\<Theta>; \<B>; \<Gamma> \<turnstile>\<^sub>w\<^sub>f ce : b' \<Longrightarrow> \<Gamma> = (x, b, c2) #\<^sub>\<Gamma> G \<Longrightarrow> \<Theta>; \<B>; (x, b, c1) #\<^sub>\<Gamma> G \<turnstile>\<^sub>w\<^sub>f c2 \<Longrightarrow> \<Theta>; \<B>; (x, b, c1) #\<^sub>\<Gamma> G \<turnstile>\<^sub>w\<^sub>f ce : b' " and
+ "\<Theta> \<turnstile>\<^sub>w\<^sub>f td \<Longrightarrow> True"
+proof(nominal_induct
+ b' and c and \<Gamma> and \<tau> and ts and \<Theta> and b and b' and td
+ avoiding: c1
+ arbitrary: \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1 and \<Gamma>\<^sub>1
+ rule:wfV_wfC_wfG_wfT_wfTs_wfTh_wfB_wfCE_wfTD.strong_induct)
+ case (wfV_varI \<Theta> \<B> \<Gamma> b' c' x')
+ have wbg: "\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f (x, b, c1) #\<^sub>\<Gamma> G " using wfC_wf wfV_varI by simp
+ show ?case proof(cases "x=x'")
+ case True
+ have "Some (b', c1) = lookup ((x, b, c1) #\<^sub>\<Gamma> G) x'" using lookup.simps wfV_varI using True by auto
+ then show ?thesis using Wellformed.wfV_varI wbg by simp
+ next
+ case False
+ then have "Some (b', c') = lookup ((x, b, c1) #\<^sub>\<Gamma> G) x'" using lookup.simps wfV_varI
+ by simp
+ then show ?thesis using Wellformed.wfV_varI wbg by simp
+ qed
+next
+ case (wfV_conspI s bv dclist \<Theta> dc x1 b' c \<B> b1 \<Gamma> v)
+ show ?case proof
+ show \<open>AF_typedef_poly s bv dclist \<in> set \<Theta>\<close> using wfV_conspI by auto
+ show \<open>(dc, \<lbrace> x1 : b' | c \<rbrace>) \<in> set dclist\<close> using wfV_conspI by auto
+ show \<open>\<Theta>; \<B> \<turnstile>\<^sub>w\<^sub>f b1 \<close> using wfV_conspI by auto
+ show \<open>atom bv \<sharp> (\<Theta>, \<B>, (x, b, c1) #\<^sub>\<Gamma> G, b1, v)\<close> unfolding fresh_prodN fresh_GCons using wfV_conspI fresh_prodN fresh_GCons by simp
+ show \<open>\<Theta>; \<B>; (x, b, c1) #\<^sub>\<Gamma> G \<turnstile>\<^sub>w\<^sub>f v : b'[bv::=b1]\<^sub>b\<^sub>b\<close> using wfV_conspI by auto
+ qed
+qed( (auto | metis wfC_wf wf_intros) +)
+
+
end
\ No newline at end of file
diff --git a/thys/MiniSail/document/root.tex b/thys/MiniSail/document/root.tex
--- a/thys/MiniSail/document/root.tex
+++ b/thys/MiniSail/document/root.tex
@@ -1,72 +1,72 @@
-\documentclass[11pt,a4paper]{report}
-\usepackage[T1]{fontenc}
-\usepackage{graphicx,isabelle,isabellesym}
-
-\usepackage[a4paper,includeheadfoot,margin=2.54cm]{geometry}
-
-% further packages required for unusual symbols (see also
-% isabellesym.sty), use only when needed
-
-\usepackage{amssymb}
- %for \<leadsto>, \<box>, \<diamond>, \<sqsupset>, \<mho>, \<Join>,
- %\<lhd>, \<lesssim>, \<greatersim>, \<lessapprox>, \<greaterapprox>,
- %\<triangleq>, \<yen>, \<lozenge>
-
-%\usepackage{eurosym}
- %for \<euro>
-
-%\usepackage[only,bigsqcap]{stmaryrd}
- %for \<Sqinter>
-
-%\usepackage{eufrak}
- %for \<AA> ... \<ZZ>, \<aa> ... \<zz> (also included in amssymb)
-
-%\usepackage{textcomp}
- %for \<onequarter>, \<onehalf>, \<threequarters>, \<degree>, \<cent>,
- %\<currency>
-
-% this should be the last package used
-\usepackage{pdfsetup}
-
-% urls in roman style, theory text in math-similar italics
-\urlstyle{rm}
-\isabellestyle{it}
-
-% for uniform font size
-%\renewcommand{\isastyle}{\isastyleminor}
-
-\begin{document}
-
-\title{MiniSail}
-\author{Mark P. Wassell}
-\maketitle
-
-\begin{abstract}
-MiniSail is a kernel language for Sail~\cite{Armstrong2019},
-an instruction set architecture (ISA) specification language.
-Sail is an imperative language with a light-weight dependent type system similar to refinement type systems such as~\cite{Vazou2014}.
-From an ISA specification, the Sail compiler can generate theorem prover code and C (or OCaml) to give an executable
-emulator for an architecture.
-The idea behind MiniSail is to capture the key and novel features of Sail in terms of their syntax,
-typing rules and operational semantics,
-and to confirm that they work together by proving progress and preservation lemmas.
-We use the Nominal2 library to handle binding.
-\end{abstract}
-
-\tableofcontents
-
-% sane default for proof documents
-\parindent 0pt\parskip 0.5ex
-
-% generated text of all theories
-\input{session}
-
-\begin{center}
- \includegraphics[width=\textwidth,height=\textheight,keepaspectratio]{session_graph}
-\end{center}
-
-\bibliographystyle{abbrv}
-\bibliography{root}
-
-\end{document}
-
+\documentclass[11pt,a4paper]{report}
+\usepackage[T1]{fontenc}
+\usepackage{graphicx,isabelle,isabellesym}
+
+\usepackage[a4paper,includeheadfoot,margin=2.54cm]{geometry}
+
+% further packages required for unusual symbols (see also
+% isabellesym.sty), use only when needed
+
+\usepackage{amssymb}
+ %for \<leadsto>, \<box>, \<diamond>, \<sqsupset>, \<mho>, \<Join>,
+ %\<lhd>, \<lesssim>, \<greatersim>, \<lessapprox>, \<greaterapprox>,
+ %\<triangleq>, \<yen>, \<lozenge>
+
+%\usepackage{eurosym}
+ %for \<euro>
+
+%\usepackage[only,bigsqcap]{stmaryrd}
+ %for \<Sqinter>
+
+%\usepackage{eufrak}
+ %for \<AA> ... \<ZZ>, \<aa> ... \<zz> (also included in amssymb)
+
+%\usepackage{textcomp}
+ %for \<onequarter>, \<onehalf>, \<threequarters>, \<degree>, \<cent>,
+ %\<currency>
+
+% this should be the last package used
+\usepackage{pdfsetup}
+
+% urls in roman style, theory text in math-similar italics
+\urlstyle{rm}
+\isabellestyle{it}
+
+% for uniform font size
+%\renewcommand{\isastyle}{\isastyleminor}
+
+\begin{document}
+
+\title{MiniSail}
+\author{Mark P. Wassell}
+\maketitle
+
+\begin{abstract}
+MiniSail is a kernel language for Sail~\cite{Armstrong2019},
+an instruction set architecture (ISA) specification language.
+Sail is an imperative language with a light-weight dependent type system similar to refinement type systems such as~\cite{Vazou2014}.
+From an ISA specification, the Sail compiler can generate theorem prover code and C (or OCaml) to give an executable
+emulator for an architecture.
+The idea behind MiniSail is to capture the key and novel features of Sail in terms of their syntax,
+typing rules and operational semantics,
+and to confirm that they work together by proving progress and preservation lemmas.
+We use the Nominal2 library to handle binding.
+\end{abstract}
+
+\tableofcontents
+
+% sane default for proof documents
+\parindent 0pt\parskip 0.5ex
+
+% generated text of all theories
+\input{session}
+
+\begin{center}
+ \includegraphics[width=\textwidth,height=\textheight,keepaspectratio]{session_graph}
+\end{center}
+
+\bibliographystyle{abbrv}
+\bibliography{root}
+
+\end{document}
+
diff --git a/thys/Modular_arithmetic_LLL_and_HNF_algorithms/Storjohann.thy b/thys/Modular_arithmetic_LLL_and_HNF_algorithms/Storjohann.thy
--- a/thys/Modular_arithmetic_LLL_and_HNF_algorithms/Storjohann.thy
+++ b/thys/Modular_arithmetic_LLL_and_HNF_algorithms/Storjohann.thy
@@ -1,2327 +1,2327 @@
-section \<open>Storjohann's basis reduction algorithm (abstract version)\<close>
-
-text \<open>This theory contains the soundness proofs of Storjohann's basis
- reduction algorithms, both for the normal and the improved-swap-order variant.
-
- The implementation of Storjohann's version of LLL uses modular operations throughout.
- It is an abstract implementation that is already quite close to what the actual implementation will be.
- In particular, the swap operation here is derived from the computation lemma for the swap
- operation in the old, integer-only formalization of LLL.\<close>
-
-theory Storjohann
- imports
- Storjohann_Mod_Operation
- LLL_Basis_Reduction.LLL_Number_Bounds
- Sqrt_Babylonian.NthRoot_Impl
-begin
-
-subsection \<open>Definition of algorithm\<close>
-
-text \<open>In the definition of the algorithm, the first-flag determines, whether only the first vector
- of the reduced basis should be computed, i.e., a short vector. Then the modulus can be slightly
- decreased in comparison to the required modulus for computing the whole reduced matrix.\<close>
-
-fun max_list_rats_with_index :: "(int * int * nat) list \<Rightarrow> (int * int * nat)" where
- "max_list_rats_with_index [x] = x" |
- "max_list_rats_with_index ((n1,d1,i1) # (n2,d2,i2) # xs)
- = max_list_rats_with_index ((if n1 * d2 \<le> n2 * d1 then (n2,d2,i2) else (n1,d1,i1)) # xs)"
-
-context LLL
-begin
-
-definition "log_base = (10 :: int)"
-
-definition bound_number :: "bool \<Rightarrow> nat" where
- "bound_number first = (if first \<and> m \<noteq> 0 then 1 else m)"
-
-definition compute_mod_of_max_gso_norm :: "bool \<Rightarrow> rat \<Rightarrow> int" where
- "compute_mod_of_max_gso_norm first mn = log_base ^ (log_ceiling log_base (max 2 (
- root_rat_ceiling 2 (mn * (rat_of_nat (bound_number first) + 3)) + 1)))"
-
-definition g_bnd_mode :: "bool \<Rightarrow> rat \<Rightarrow> int vec list \<Rightarrow> bool" where
- "g_bnd_mode first b fs = (if first \<and> m \<noteq> 0 then sq_norm (gso fs 0) \<le> b else g_bnd b fs)"
-
-definition d_of where "d_of dmu i = (if i = 0 then 1 :: int else dmu $$ (i - 1, i - 1))"
-
-definition compute_max_gso_norm :: "bool \<Rightarrow> int mat \<Rightarrow> rat \<times> nat" where
- "compute_max_gso_norm first dmu = (if m = 0 then (0,0) else
- case max_list_rats_with_index (map (\<lambda> i. (d_of dmu (Suc i), d_of dmu i, i)) [0 ..< (if first then 1 else m)])
- of (num, denom, i) \<Rightarrow> (of_int num / of_int denom, i))"
-
-
-context
- fixes p :: int \<comment> \<open>the modulus\<close>
- and first :: bool \<comment> \<open>only compute first vector of reduced basis\<close>
-begin
-
-definition basis_reduction_mod_add_row ::
- "int vec list \<Rightarrow> int mat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> (int vec list \<times> int mat)" where
- "basis_reduction_mod_add_row mfs dmu i j =
- (let c = round_num_denom (dmu $$ (i,j)) (d_of dmu (Suc j)) in
- (if c = 0 then (mfs, dmu)
- else (mfs[ i := (map_vec (\<lambda> x. x symmod p)) (mfs ! i - c \<cdot>\<^sub>v mfs ! j)],
- mat m m (\<lambda>(i',j'). (if (i' = i \<and> j' \<le> j)
- then (if j'=j then (dmu $$ (i,j') - c * dmu $$ (j,j'))
- else (dmu $$ (i,j') - c * dmu $$ (j,j'))
- symmod (p * d_of dmu j' * d_of dmu (Suc j')))
- else (dmu $$ (i',j')))))))"
-
-fun basis_reduction_mod_add_rows_loop where
- "basis_reduction_mod_add_rows_loop mfs dmu i 0 = (mfs, dmu)"
-| "basis_reduction_mod_add_rows_loop mfs dmu i (Suc j) = (
- let (mfs', dmu') = basis_reduction_mod_add_row mfs dmu i j
- in basis_reduction_mod_add_rows_loop mfs' dmu' i j)"
-
-definition basis_reduction_mod_swap_dmu_mod :: "int mat \<Rightarrow> nat \<Rightarrow> int mat" where
- "basis_reduction_mod_swap_dmu_mod dmu k = mat m m (\<lambda>(i, j). (
- if j < i \<and> (j = k \<or> j = k - 1) then
- dmu $$ (i, j) symmod (p * d_of dmu j * d_of dmu (Suc j))
- else dmu $$ (i, j)))"
-
-definition basis_reduction_mod_swap where
- "basis_reduction_mod_swap mfs dmu k =
- (mfs[k := mfs ! (k - 1), k - 1 := mfs ! k],
- basis_reduction_mod_swap_dmu_mod (mat m m (\<lambda>(i,j). (
- if j < i then
- if i = k - 1 then
- dmu $$ (k, j)
- else if i = k \<and> j \<noteq> k - 1 then
- dmu $$ (k - 1, j)
- else if i > k \<and> j = k then
- ((d_of dmu (Suc k)) * dmu $$ (i, k - 1) - dmu $$ (k, k - 1) * dmu $$ (i, j))
- div (d_of dmu k)
- else if i > k \<and> j = k - 1 then
- (dmu $$ (k, k - 1) * dmu $$ (i, j) + dmu $$ (i, k) * (d_of dmu (k-1)))
- div (d_of dmu k)
- else dmu $$ (i, j)
- else if i = j then
- if i = k - 1 then
- ((d_of dmu (Suc k)) * (d_of dmu (k-1)) + dmu $$ (k, k - 1) * dmu $$ (k, k - 1))
- div (d_of dmu k)
- else (d_of dmu (Suc i))
- else dmu $$ (i, j))
- )) k)"
-
-fun basis_reduction_adjust_mod where
- "basis_reduction_adjust_mod mfs dmu =
- (let (b,g_idx) = compute_max_gso_norm first dmu;
- p' = compute_mod_of_max_gso_norm first b
- in if p' < p then
- let mfs' = map (map_vec (\<lambda>x. x symmod p')) mfs;
- d_vec = vec (Suc m) (\<lambda> i. d_of dmu i);
- dmu' = mat m m (\<lambda> (i,j). if j < i then dmu $$ (i,j)
- symmod (p' * d_vec $ j * d_vec $ (Suc j)) else
- dmu $$ (i,j))
- in (p', mfs', dmu', g_idx)
- else (p, mfs, dmu, g_idx))"
-
-definition basis_reduction_adjust_swap_add_step where
- "basis_reduction_adjust_swap_add_step mfs dmu g_idx i = (
- let i1 = i - 1;
- (mfs1, dmu1) = basis_reduction_mod_add_row mfs dmu i i1;
- (mfs2, dmu2) = basis_reduction_mod_swap mfs1 dmu1 i
- in if i1 = g_idx then basis_reduction_adjust_mod mfs2 dmu2
- else (p, mfs2, dmu2, g_idx))"
-
-
-definition basis_reduction_mod_step where
- "basis_reduction_mod_step mfs dmu g_idx i (j :: int) = (if i = 0 then (p, mfs, dmu, g_idx, Suc i, j)
- else let di = d_of dmu i;
- (num, denom) = quotient_of \<alpha>
- in if di * di * denom \<le> num * d_of dmu (i - 1) * d_of dmu (Suc i) then
- (p, mfs, dmu, g_idx, Suc i, j)
- else let (p', mfs', dmu', g_idx') = basis_reduction_adjust_swap_add_step mfs dmu g_idx i
- in (p', mfs', dmu', g_idx', i - 1, j + 1))"
-
-primrec basis_reduction_mod_add_rows_outer_loop where
- "basis_reduction_mod_add_rows_outer_loop mfs dmu 0 = (mfs, dmu)" |
- "basis_reduction_mod_add_rows_outer_loop mfs dmu (Suc i) =
- (let (mfs', dmu') = basis_reduction_mod_add_rows_outer_loop mfs dmu i in
- basis_reduction_mod_add_rows_loop mfs' dmu' (Suc i) (Suc i))"
-end
-
-text \<open>the main loop of the normal Storjohann algorithm\<close>
-partial_function (tailrec) basis_reduction_mod_main where
- "basis_reduction_mod_main p first mfs dmu g_idx i (j :: int) = (
- (if i < m
- then
- case basis_reduction_mod_step p first mfs dmu g_idx i j
- of (p', mfs', dmu', g_idx', i', j') \<Rightarrow>
- basis_reduction_mod_main p' first mfs' dmu' g_idx' i' j'
- else
- (p, mfs, dmu)))"
-
-definition compute_max_gso_quot:: "int mat \<Rightarrow> (int * int * nat)" where
- "compute_max_gso_quot dmu = max_list_rats_with_index
- (map (\<lambda>i. ((d_of dmu (i+1)) * (d_of dmu (i+1)), (d_of dmu (i+2)) * (d_of dmu i), Suc i)) [0..<(m-1)])"
-
-text \<open>the main loop of Storjohann's algorithm with improved swap order\<close>
-partial_function (tailrec) basis_reduction_iso_main where
- "basis_reduction_iso_main p first mfs dmu g_idx (j :: int) = (
- (if m > 1 then
- (let (max_gso_num, max_gso_denum, indx) = compute_max_gso_quot dmu;
- (num, denum) = quotient_of \<alpha> in
- (if (max_gso_num * denum > num * max_gso_denum) then
- case basis_reduction_adjust_swap_add_step p first mfs dmu g_idx indx of
- (p', mfs', dmu', g_idx') \<Rightarrow>
- basis_reduction_iso_main p' first mfs' dmu' g_idx' (j + 1)
- else
- (p, mfs, dmu)))
- else (p, mfs, dmu)))"
-
-definition compute_initial_mfs where
- "compute_initial_mfs p = map (map_vec (\<lambda>x. x symmod p)) fs_init"
-
-definition compute_initial_dmu where
- "compute_initial_dmu p dmu = mat m m (\<lambda>(i',j'). if j' < i'
- then dmu $$ (i', j') symmod (p * d_of dmu j' * d_of dmu (Suc j'))
- else dmu $$ (i', j'))"
-
-definition "dmu_initial = (let dmu = d\<mu>_impl fs_init
- in mat m m (\<lambda> (i,j).
- if j \<le> i then d\<mu>_impl fs_init !! i !! j else 0))"
-
-definition "compute_initial_state first =
- (let dmu = dmu_initial;
- (b, g_idx) = compute_max_gso_norm first dmu;
- p = compute_mod_of_max_gso_norm first b
- in (p, compute_initial_mfs p, compute_initial_dmu p dmu, g_idx))"
-
-text \<open>Storjohann's algorithm\<close>
-definition reduce_basis_mod :: "int vec list" where
- "reduce_basis_mod = (
- let first = False;
- (p0, mfs0, dmu0, g_idx) = compute_initial_state first;
- (p', mfs', dmu') = basis_reduction_mod_main p0 first mfs0 dmu0 g_idx 0 0;
- (mfs'', dmu'') = basis_reduction_mod_add_rows_outer_loop p' mfs' dmu' (m-1)
- in mfs'')"
-
-text \<open>Storjohann's algorithm with improved swap order\<close>
-definition reduce_basis_iso :: "int vec list" where
- "reduce_basis_iso = (
- let first = False;
- (p0, mfs0, dmu0, g_idx) = compute_initial_state first;
- (p', mfs', dmu') = basis_reduction_iso_main p0 first mfs0 dmu0 g_idx 0;
- (mfs'', dmu'') = basis_reduction_mod_add_rows_outer_loop p' mfs' dmu' (m-1)
- in mfs'')"
-
-text \<open>Storjohann's algorithm for computing a short vector\<close>
-definition
- "short_vector_mod = (
- let first = True;
- (p0, mfs0, dmu0, g_idx) = compute_initial_state first;
- (p', mfs', dmu') = basis_reduction_mod_main p0 first mfs0 dmu0 g_idx 0 0
- in hd mfs')"
-
-text \<open>Storjohann's algorithm (iso-variant) for computing a short vector\<close>
-definition
- "short_vector_iso = (
- let first = True;
- (p0, mfs0, dmu0, g_idx) = compute_initial_state first;
- (p', mfs', dmu') = basis_reduction_iso_main p0 first mfs0 dmu0 g_idx 0
- in hd mfs')"
-end
-
-subsection \<open>Towards soundness of Storjohann's algorithm\<close>
-
-lemma max_list_rats_with_index_in_set:
- assumes max: "max_list_rats_with_index xs = (nm, dm, im)"
- and len: "length xs \<ge> 1"
-shows "(nm, dm, im) \<in> set xs"
- using assms
-proof (induct xs rule: max_list_rats_with_index.induct)
- case (2 n1 d1 i1 n2 d2 i2 xs)
- have "1 \<le> length ((if n1 * d2 \<le> n2 * d1 then (n2, d2, i2) else (n1, d1, i1)) # xs)" by simp
- moreover have "max_list_rats_with_index ((if n1 * d2 \<le> n2 * d1 then (n2, d2, i2) else (n1, d1, i1)) # xs)
- = (nm, dm, im)" using 2 by simp
- moreover have "(if n1 * d2 \<le> n2 * d1 then (n2, d2, i2) else (n1, d1, i1)) \<in>
- set ((n1, d1, i1) # (n2, d2, i2) # xs)" by simp
- moreover then have "set ((if n1 * d2 \<le> n2 * d1 then (n2, d2, i2) else (n1, d1, i1)) # xs) \<subseteq>
- set ((n1, d1, i1) # (n2, d2, i2) # xs)" by auto
- ultimately show ?case using 2(1) by auto
-qed auto
-
-lemma max_list_rats_with_index: assumes "\<And> n d i. (n,d,i) \<in> set xs \<Longrightarrow> d > 0"
- and max: "max_list_rats_with_index xs = (nm, dm, im)"
- and "(n,d,i) \<in> set xs"
-shows "rat_of_int n / of_int d \<le> of_int nm / of_int dm"
- using assms
-proof (induct xs arbitrary: n d i rule: max_list_rats_with_index.induct)
- case (2 n1 d1 i1 n2 d2 i2 xs n d i)
- let ?r = "rat_of_int"
- from 2(2) have "d1 > 0" "d2 > 0" by auto
- hence d: "?r d1 > 0" "?r d2 > 0" by auto
- have "(n1 * d2 \<le> n2 * d1) = (?r n1 * ?r d2 \<le> ?r n2 * ?r d1)"
- unfolding of_int_mult[symmetric] by presburger
- also have "\<dots> = (?r n1 / ?r d1 \<le> ?r n2 / ?r d2)" using d
- by (smt divide_strict_right_mono leD le_less_linear mult.commute nonzero_mult_div_cancel_left
- not_less_iff_gr_or_eq times_divide_eq_right)
- finally have id: "(n1 * d2 \<le> n2 * d1) = (?r n1 / ?r d1 \<le> ?r n2 / ?r d2)" .
- obtain n' d' i' where new: "(if n1 * d2 \<le> n2 * d1 then (n2, d2, i2) else (n1, d1, i1)) = (n',d',i')"
- by force
- have nd': "(n',d',i') \<in> {(n1,d1,i1), (n2, d2, i2)}" using new[symmetric] by auto
- from 2(3) have res: "max_list_rats_with_index ((n',d',i') # xs) = (nm, dm, im)" using new by auto
- note 2 = 2[unfolded new]
- show ?case
- proof (cases "(n,d,i) \<in> set xs")
- case True
- show ?thesis
- by (rule 2(1)[of n d, OF 2(2) res], insert True nd', force+)
- next
- case False
- with 2(4) have "n = n1 \<and> d = d1 \<or> n = n2 \<and> d = d2" by auto
- hence "?r n / ?r d \<le> ?r n' / ?r d'" using new[unfolded id]
- by (metis linear prod.inject)
- also have "?r n' / ?r d' \<le> ?r nm / ?r dm"
- by (rule 2(1)[of n' d', OF 2(2) res], insert nd', force+)
- finally show ?thesis .
- qed
-qed auto
-
-context LLL
-begin
-
-lemma log_base: "log_base \<ge> 2" unfolding log_base_def by auto
-
-definition LLL_invariant_weak' :: "nat \<Rightarrow> int vec list \<Rightarrow> bool" where
- "LLL_invariant_weak' i fs = (
- gs.lin_indpt_list (RAT fs) \<and>
- lattice_of fs = L \<and>
- weakly_reduced fs i \<and>
- i \<le> m \<and>
- length fs = m
- )"
-
-lemma LLL_invD_weak: assumes "LLL_invariant_weak' i fs"
- shows
- "lin_indep fs"
- "length (RAT fs) = m"
- "set fs \<subseteq> carrier_vec n"
- "\<And> i. i < m \<Longrightarrow> fs ! i \<in> carrier_vec n"
- "\<And> i. i < m \<Longrightarrow> gso fs i \<in> carrier_vec n"
- "length fs = m"
- "lattice_of fs = L"
- "weakly_reduced fs i"
- "i \<le> m"
-proof (atomize (full), goal_cases)
- case 1
- interpret gs': gram_schmidt_fs_lin_indpt n "RAT fs"
- by (standard) (use assms LLL_invariant_weak'_def gs.lin_indpt_list_def in auto)
- show ?case
- using assms gs'.fs_carrier gs'.f_carrier gs'.gso_carrier
- by (auto simp add: LLL_invariant_weak'_def gram_schmidt_fs.reduced_def)
-qed
-
-lemma LLL_invI_weak: assumes
- "set fs \<subseteq> carrier_vec n"
- "length fs = m"
- "lattice_of fs = L"
- "i \<le> m"
- "lin_indep fs"
- "weakly_reduced fs i"
-shows "LLL_invariant_weak' i fs"
- unfolding LLL_invariant_weak'_def Let_def using assms by auto
-
-lemma LLL_invw'_imp_w: "LLL_invariant_weak' i fs \<Longrightarrow> LLL_invariant_weak fs"
- unfolding LLL_invariant_weak'_def LLL_invariant_weak_def by auto
-
-lemma basis_reduction_add_row_weak:
- assumes Linvw: "LLL_invariant_weak' i fs"
- and i: "i < m" and j: "j < i"
- and fs': "fs' = fs[ i := fs ! i - c \<cdot>\<^sub>v fs ! j]"
-shows "LLL_invariant_weak' i fs'"
- "g_bnd B fs \<Longrightarrow> g_bnd B fs'"
-proof (atomize(full), goal_cases)
- case 1
- note Linv = LLL_invw'_imp_w[OF Linvw]
- note main = basis_reduction_add_row_main[OF Linv i j fs']
- have bnd: "g_bnd B fs \<Longrightarrow> g_bnd B fs'" using main(6) unfolding g_bnd_def by auto
- note new = LLL_inv_wD[OF main(1)]
- note old = LLL_invD_weak[OF Linvw]
- have red: "weakly_reduced fs' i" using \<open>weakly_reduced fs i\<close> main(6) \<open>i < m\<close>
- unfolding gram_schmidt_fs.weakly_reduced_def by auto
- have inv: "LLL_invariant_weak' i fs'" using LLL_inv_wD[OF main(1)] \<open>i < m\<close>
- by (intro LLL_invI_weak, auto intro: red)
- show ?case using inv red main bnd by auto
-qed
-
-lemma LLL_inv_weak_m_impl_i:
- assumes inv: "LLL_invariant_weak' m fs"
- and i: "i \<le> m"
-shows "LLL_invariant_weak' i fs"
-proof -
- have "weakly_reduced fs i" using LLL_invD_weak(8)[OF inv]
- by (meson assms(2) gram_schmidt_fs.weakly_reduced_def le_trans less_imp_le_nat linorder_not_less)
- then show ?thesis
- using LLL_invI_weak[of fs i, OF LLL_invD_weak(3,6,7)[OF inv] _ LLL_invD_weak(1)[OF inv]]
- LLL_invD_weak(2,4,5,8-)[OF inv] i by simp
-qed
-
-definition mod_invariant where
- "mod_invariant b p first = (b \<le> rat_of_int (p - 1)^2 / (rat_of_nat (bound_number first) + 3)
- \<and> (\<exists> e. p = log_base ^ e))"
-
-lemma compute_mod_of_max_gso_norm: assumes mn: "mn \<ge> 0"
- and m: "m = 0 \<Longrightarrow> mn = 0"
- and p: "p = compute_mod_of_max_gso_norm first mn"
-shows
- "p > 1"
- "mod_invariant mn p first"
-proof -
- let ?m = "bound_number first"
- define p' where "p' = root_rat_ceiling 2 (mn * (rat_of_nat ?m + 3)) + 1"
- define p'' where "p'' = max 2 p'"
- define q where "q = real_of_rat (mn * (rat_of_nat ?m + 3))"
- have *: "-1 < (0 :: real)" by simp
- also have "0 \<le> root 2 (real_of_rat (mn * (rat_of_nat ?m + 3)))" using mn by auto
- finally have "p' \<ge> 0 + 1" unfolding p'_def
- by (intro plus_left_mono, simp)
- hence p': "p' > 0" by auto
- have p'': "p'' > 1" unfolding p''_def by auto
- have pp'': "p \<ge> p''" unfolding compute_mod_of_max_gso_norm_def p p'_def[symmetric] p''_def[symmetric]
- using log_base p'' log_ceiling_sound by auto
- hence pp': "p \<ge> p'" unfolding p''_def by auto
- show "p > 1" using pp'' p'' by auto
-
- have q0: "q \<ge> 0" unfolding q_def using mn m by auto
- have "(mn \<le> rat_of_int (p' - 1)^2 / (rat_of_nat ?m + 3))
- = (real_of_rat mn \<le> real_of_rat (rat_of_int (p' - 1)^2 / (rat_of_nat ?m + 3)))" using of_rat_less_eq by blast
- also have "\<dots> = (real_of_rat mn \<le> real_of_rat (rat_of_int (p' - 1)^2) / real_of_rat (rat_of_nat ?m + 3))" by (simp add: of_rat_divide)
- also have "\<dots> = (real_of_rat mn \<le> ((real_of_int (p' - 1))^2) / real_of_rat (rat_of_nat ?m + 3))"
- by (metis of_rat_of_int_eq of_rat_power)
- also have "\<dots> = (real_of_rat mn \<le> (real_of_int \<lceil>sqrt q\<rceil>)^2 / real_of_rat (rat_of_nat ?m + 3))"
- unfolding p'_def sqrt_def q_def by simp
- also have "\<dots>"
- proof -
- have "real_of_rat mn \<le> q / real_of_rat (rat_of_nat ?m + 3)" unfolding q_def using m
- by (auto simp: of_rat_mult)
- also have "\<dots> \<le> (real_of_int \<lceil>sqrt q\<rceil>)^2 / real_of_rat (rat_of_nat ?m + 3)"
- proof (rule divide_right_mono)
- have "q = (sqrt q)^2" using q0 by simp
- also have "\<dots> \<le> (real_of_int \<lceil>sqrt q\<rceil>)^2"
- by (rule power_mono, auto simp: q0)
- finally show "q \<le> (real_of_int \<lceil>sqrt q\<rceil>)^2" .
- qed auto
- finally show ?thesis .
- qed
- finally have "mn \<le> rat_of_int (p' - 1)^2 / (rat_of_nat ?m + 3)" .
- also have "\<dots> \<le> rat_of_int (p - 1)^2 / (rat_of_nat ?m + 3)"
- unfolding power2_eq_square
- by (intro divide_right_mono mult_mono, insert p' pp', auto)
- finally have "mn \<le> rat_of_int (p - 1)^2 / (rat_of_nat ?m + 3)" .
- moreover have "\<exists> e. p = log_base ^ e" unfolding p compute_mod_of_max_gso_norm_def by auto
- ultimately show "mod_invariant mn p first" unfolding mod_invariant_def by auto
-qed
-
-lemma g_bnd_mode_cong: assumes "\<And> i. i < m \<Longrightarrow> gso fs i = gso fs' i"
- shows "g_bnd_mode first b fs = g_bnd_mode first b fs'"
- using assms unfolding g_bnd_mode_def g_bnd_def by auto
-
-definition LLL_invariant_mod :: "int vec list \<Rightarrow> int vec list \<Rightarrow> int mat \<Rightarrow> int \<Rightarrow> bool \<Rightarrow> rat \<Rightarrow> nat \<Rightarrow> bool" where
- "LLL_invariant_mod fs mfs dmu p first b i = (
- length fs = m \<and>
- length mfs = m \<and>
- i \<le> m \<and>
- lattice_of fs = L \<and>
- gs.lin_indpt_list (RAT fs) \<and>
- weakly_reduced fs i \<and>
- (map (map_vec (\<lambda>x. x symmod p)) fs = mfs) \<and>
- (\<forall>i' < m. \<forall> j' < i'. \<bar>d\<mu> fs i' j'\<bar> < p * d fs j' * d fs (Suc j')) \<and>
- (\<forall>i' < m. \<forall>j' < m. d\<mu> fs i' j' = dmu $$ (i',j')) \<and>
- p > 1 \<and>
- g_bnd_mode first b fs \<and>
- mod_invariant b p first
-)"
-
-lemma LLL_invD_mod: assumes "LLL_invariant_mod fs mfs dmu p first b i"
-shows
- "length mfs = m"
- "i \<le> m"
- "length fs = m"
- "lattice_of fs = L"
- "gs.lin_indpt_list (RAT fs)"
- "weakly_reduced fs i"
- "(map (map_vec (\<lambda>x. x symmod p)) fs = mfs)"
- "(\<forall>i' < m. \<forall>j' < i'. \<bar>d\<mu> fs i' j'\<bar> < p * d fs j' * d fs (Suc j'))"
- "(\<forall>i' < m. \<forall>j' < m. d\<mu> fs i' j' = dmu $$ (i',j'))"
- "\<And> i. i < m \<Longrightarrow> fs ! i \<in> carrier_vec n"
- "set fs \<subseteq> carrier_vec n"
- "\<And> i. i < m \<Longrightarrow> gso fs i \<in> carrier_vec n"
- "\<And> i. i < m \<Longrightarrow> mfs ! i \<in> carrier_vec n"
- "set mfs \<subseteq> carrier_vec n"
- "p > 1"
- "g_bnd_mode first b fs"
- "mod_invariant b p first"
-proof (atomize (full), goal_cases)
- case 1
- interpret gs': gram_schmidt_fs_lin_indpt n "RAT fs"
- using assms LLL_invariant_mod_def gs.lin_indpt_list_def
- by (meson gram_schmidt_fs_Rn.intro gram_schmidt_fs_lin_indpt.intro gram_schmidt_fs_lin_indpt_axioms.intro)
- have allfs: "\<forall>i < m. fs ! i \<in> carrier_vec n" using assms gs'.f_carrier
- by (simp add: LLL.LLL_invariant_mod_def)
- then have setfs: "set fs \<subseteq> carrier_vec n" by (metis LLL_invariant_mod_def assms in_set_conv_nth subsetI)
- have allgso: "(\<forall>i < m. gso fs i \<in> carrier_vec n)" using assms gs'.gso_carrier
- by (simp add: LLL.LLL_invariant_mod_def)
- show ?case
- using assms gs'.fs_carrier gs'.f_carrier gs'.gso_carrier allfs allgso
- LLL_invariant_mod_def gram_schmidt_fs.reduced_def in_set_conv_nth setfs by fastforce
-qed
-
-lemma LLL_invI_mod: assumes
- "length mfs = m"
- "i \<le> m"
- "length fs = m"
- "lattice_of fs = L"
- "gs.lin_indpt_list (RAT fs)"
- "weakly_reduced fs i"
- "map (map_vec (\<lambda>x. x symmod p)) fs = mfs"
- "(\<forall>i' < m. \<forall>j' < i'. \<bar>d\<mu> fs i' j'\<bar> < p * d fs j' * d fs (Suc j'))"
- "(\<forall>i' < m. \<forall>j' < m. d\<mu> fs i' j' = dmu $$ (i',j'))"
- "p > 1"
- "g_bnd_mode first b fs"
- "mod_invariant b p first"
-shows "LLL_invariant_mod fs mfs dmu p first b i"
- unfolding LLL_invariant_mod_def using assms by blast
-
-definition LLL_invariant_mod_weak :: "int vec list \<Rightarrow> int vec list \<Rightarrow> int mat \<Rightarrow> int \<Rightarrow> bool \<Rightarrow> rat \<Rightarrow> bool" where
- "LLL_invariant_mod_weak fs mfs dmu p first b = (
- length fs = m \<and>
- length mfs = m \<and>
- lattice_of fs = L \<and>
- gs.lin_indpt_list (RAT fs) \<and>
- (map (map_vec (\<lambda>x. x symmod p)) fs = mfs) \<and>
- (\<forall>i' < m. \<forall> j' < i'. \<bar>d\<mu> fs i' j'\<bar> < p * d fs j' * d fs (Suc j')) \<and>
- (\<forall>i' < m. \<forall>j' < m. d\<mu> fs i' j' = dmu $$ (i',j')) \<and>
- p > 1 \<and>
- g_bnd_mode first b fs \<and>
- mod_invariant b p first
-)"
-
-lemma LLL_invD_modw: assumes "LLL_invariant_mod_weak fs mfs dmu p first b"
-shows
- "length mfs = m"
- "length fs = m"
- "lattice_of fs = L"
- "gs.lin_indpt_list (RAT fs)"
- "(map (map_vec (\<lambda>x. x symmod p)) fs = mfs)"
- "(\<forall>i' < m. \<forall>j' < i'. \<bar>d\<mu> fs i' j'\<bar> < p * d fs j' * d fs (Suc j'))"
- "(\<forall>i' < m. \<forall>j' < m. d\<mu> fs i' j' = dmu $$ (i',j'))"
- "\<And> i. i < m \<Longrightarrow> fs ! i \<in> carrier_vec n"
- "set fs \<subseteq> carrier_vec n"
- "\<And> i. i < m \<Longrightarrow> gso fs i \<in> carrier_vec n"
- "\<And> i. i < m \<Longrightarrow> mfs ! i \<in> carrier_vec n"
- "set mfs \<subseteq> carrier_vec n"
- "p > 1"
- "g_bnd_mode first b fs"
- "mod_invariant b p first"
-proof (atomize (full), goal_cases)
- case 1
- interpret gs': gram_schmidt_fs_lin_indpt n "RAT fs"
- using assms LLL_invariant_mod_weak_def gs.lin_indpt_list_def
- by (meson gram_schmidt_fs_Rn.intro gram_schmidt_fs_lin_indpt.intro gram_schmidt_fs_lin_indpt_axioms.intro)
- have allfs: "\<forall>i < m. fs ! i \<in> carrier_vec n" using assms gs'.f_carrier
- by (simp add: LLL.LLL_invariant_mod_weak_def)
- then have setfs: "set fs \<subseteq> carrier_vec n" by (metis LLL_invariant_mod_weak_def assms in_set_conv_nth subsetI)
- have allgso: "(\<forall>i < m. gso fs i \<in> carrier_vec n)" using assms gs'.gso_carrier
- by (simp add: LLL.LLL_invariant_mod_weak_def)
- show ?case
- using assms gs'.fs_carrier gs'.f_carrier gs'.gso_carrier allfs allgso
- LLL_invariant_mod_weak_def gram_schmidt_fs.reduced_def in_set_conv_nth setfs by fastforce
-qed
-
-lemma LLL_invI_modw: assumes
- "length mfs = m"
- "length fs = m"
- "lattice_of fs = L"
- "gs.lin_indpt_list (RAT fs)"
- "map (map_vec (\<lambda>x. x symmod p)) fs = mfs"
- "(\<forall>i' < m. \<forall>j' < i'. \<bar>d\<mu> fs i' j'\<bar> < p * d fs j' * d fs (Suc j'))"
- "(\<forall>i' < m. \<forall>j' < m. d\<mu> fs i' j' = dmu $$ (i',j'))"
- "p > 1"
- "g_bnd_mode first b fs"
- "mod_invariant b p first"
-shows "LLL_invariant_mod_weak fs mfs dmu p first b"
- unfolding LLL_invariant_mod_weak_def using assms by blast
-
-lemma dd\<mu>:
- assumes i: "i < m"
- shows "d fs (Suc i) = d\<mu> fs i i"
-proof-
- have "\<mu> fs i i = 1" using i by (simp add: gram_schmidt_fs.\<mu>.simps)
- then show ?thesis using d\<mu>_def by simp
-qed
-
-lemma d_of_main: assumes "(\<forall>i' < m. d\<mu> fs i' i' = dmu $$ (i',i'))"
- and "i \<le> m"
-shows "d_of dmu i = d fs i"
-proof (cases "i = 0")
- case False
- with assms have "i - 1 < m" by auto
- from assms(1)[rule_format, OF this] dd\<mu>[OF this, of fs] False
- show ?thesis by (simp add: d_of_def)
-next
- case True
- thus ?thesis unfolding d_of_def True d_def by simp
-qed
-
-lemma d_of: assumes inv: "LLL_invariant_mod fs mfs dmu p b first j"
- and "i \<le> m"
-shows "d_of dmu i = d fs i"
- by (rule d_of_main[OF _ assms(2)], insert LLL_invD_mod(9)[OF inv], auto)
-
-lemma d_of_weak: assumes inv: "LLL_invariant_mod_weak fs mfs dmu p first b"
- and "i \<le> m"
-shows "d_of dmu i = d fs i"
- by (rule d_of_main[OF _ assms(2)], insert LLL_invD_modw(7)[OF inv], auto)
-
-lemma compute_max_gso_norm: assumes dmu: "(\<forall>i' < m. d\<mu> fs i' i' = dmu $$ (i',i'))"
- and Linv: "LLL_invariant_weak fs"
-shows "g_bnd_mode first (fst (compute_max_gso_norm first dmu)) fs"
- "fst (compute_max_gso_norm first dmu) \<ge> 0"
- "m = 0 \<Longrightarrow> fst (compute_max_gso_norm first dmu) = 0"
-proof -
- show gbnd: "g_bnd_mode first (fst (compute_max_gso_norm first dmu)) fs"
- proof (cases "first \<and> m \<noteq> 0")
- case False
- have "?thesis = (g_bnd (fst (compute_max_gso_norm first dmu)) fs)" unfolding g_bnd_mode_def using False by auto
- also have \<dots> unfolding g_bnd_def
- proof (intro allI impI)
- fix i
- assume i: "i < m"
- have id: "(if first then 1 else m) = m" using False i by auto
- define list where "list = map (\<lambda> i. (d_of dmu (Suc i), d_of dmu i, i)) [0 ..< m ]"
- obtain num denom j where ml: "max_list_rats_with_index list = (num, denom, j)"
- by (metis prod_cases3)
- have dpos: "d fs i > 0" using LLL_d_pos[OF Linv, of i] i by auto
- have pos: "(n, d, i) \<in> set list \<Longrightarrow> 0 < d" for n d i
- using LLL_d_pos[OF Linv] unfolding list_def using d_of_main[OF dmu] by auto
- from i have "list ! i \<in> set list" using i unfolding list_def by auto
- also have "list ! i = (d_of dmu (Suc i), d_of dmu i, i)" unfolding list_def using i by auto
- also have "\<dots> = (d fs (Suc i), d fs i, i)" using d_of_main[OF dmu] i by auto
- finally have "(d fs (Suc i), d fs i, i) \<in> set list" .
- from max_list_rats_with_index[OF pos ml this]
- have "of_int (d fs (Suc i)) / of_int (d fs i) \<le> fst (compute_max_gso_norm first dmu)"
- unfolding compute_max_gso_norm_def list_def[symmetric] ml id split using i by auto
- also have "of_int (d fs (Suc i)) / of_int (d fs i) = sq_norm (gso fs i)"
- using LLL_d_Suc[OF Linv i] dpos by auto
- finally show "sq_norm (gso fs i) \<le> fst (compute_max_gso_norm first dmu)" .
- qed
- finally show ?thesis .
- next
- case True
- thus ?thesis unfolding g_bnd_mode_def compute_max_gso_norm_def using d_of_main[OF dmu]
- LLL_d_Suc[OF Linv, of 0] LLL_d_pos[OF Linv, of 0] LLL_d_pos[OF Linv, of 1] by auto
- qed
- show "fst (compute_max_gso_norm first dmu) \<ge> 0"
- proof (cases "m = 0")
- case True
- thus ?thesis unfolding compute_max_gso_norm_def by simp
- next
- case False
- hence 0: "0 < m" by simp
- have "0 \<le> sq_norm (gso fs 0)" by blast
- also have "\<dots> \<le> fst (compute_max_gso_norm first dmu)"
- using gbnd[unfolded g_bnd_mode_def g_bnd_def] using 0 by metis
- finally show ?thesis .
- qed
-qed (auto simp: LLL.compute_max_gso_norm_def)
-
-
-lemma increase_i_mod:
- assumes Linv: "LLL_invariant_mod fs mfs dmu p first b i"
- and i: "i < m"
- and red_i: "i \<noteq> 0 \<Longrightarrow> sq_norm (gso fs (i - 1)) \<le> \<alpha> * sq_norm (gso fs i)"
-shows "LLL_invariant_mod fs mfs dmu p first b (Suc i)" "LLL_measure i fs > LLL_measure (Suc i) fs"
-proof -
- note inv = LLL_invD_mod[OF Linv]
- from inv have red: "weakly_reduced fs i" by (auto)
- from red red_i i have red: "weakly_reduced fs (Suc i)"
- unfolding gram_schmidt_fs.weakly_reduced_def
- by (intro allI impI, rename_tac ii, case_tac "Suc ii = i", auto)
- show "LLL_invariant_mod fs mfs dmu p first b (Suc i)"
- by (intro LLL_invI_mod, insert inv red i, auto)
- show "LLL_measure i fs > LLL_measure (Suc i) fs" unfolding LLL_measure_def using i by auto
-qed
-
-lemma basis_reduction_mod_add_row_main:
- assumes Linvmw: "LLL_invariant_mod_weak fs mfs dmu p first b"
- and i: "i < m" and j: "j < i"
- and c: "c = round (\<mu> fs i j)"
- and mfs': "mfs' = mfs[ i := (map_vec (\<lambda> x. x symmod p)) (mfs ! i - c \<cdot>\<^sub>v mfs ! j)]"
- and dmu': "dmu' = mat m m (\<lambda>(i',j'). (if (i' = i \<and> j' \<le> j)
- then (if j'=j then (dmu $$ (i,j') - c * dmu $$ (j,j'))
- else (dmu $$ (i,j') - c * dmu $$ (j,j'))
- symmod (p * (d_of dmu j') * (d_of dmu (Suc j'))))
- else (dmu $$ (i',j'))))"
-shows "(\<exists>fs'. LLL_invariant_mod_weak fs' mfs' dmu' p first b \<and>
- LLL_measure i fs' = LLL_measure i fs
- \<and> (\<mu>_small_row i fs (Suc j) \<longrightarrow> \<mu>_small_row i fs' j)
- \<and> (\<forall>k < m. gso fs' k = gso fs k)
- \<and> (\<forall>ii \<le> m. d fs' ii = d fs ii)
- \<and> \<bar>\<mu> fs' i j\<bar> \<le> 1 / 2
- \<and> (\<forall>i' j'. i' < i \<longrightarrow> j' \<le> i' \<longrightarrow> \<mu> fs' i' j' = \<mu> fs i' j')
- \<and> (LLL_invariant_mod fs mfs dmu p first b i \<longrightarrow> LLL_invariant_mod fs' mfs' dmu' p first b i))"
-proof -
- define fs' where "fs' = fs[ i := fs ! i - c \<cdot>\<^sub>v fs ! j]"
- from LLL_invD_modw[OF Linvmw] have gbnd: "g_bnd_mode first b fs" and p1: "p > 1" and pgtz: "p > 0" by auto
- have Linvww: "LLL_invariant_weak fs" using LLL_invD_modw[OF Linvmw] LLL_invariant_weak_def by simp
- have
- Linvw': "LLL_invariant_weak fs'" and
- 01: "c = round (\<mu> fs i j) \<Longrightarrow> \<mu>_small_row i fs (Suc j) \<Longrightarrow> \<mu>_small_row i fs' j" and
- 02: "LLL_measure i fs' = LLL_measure i fs" and
- 03: "\<And> i. i < m \<Longrightarrow> gso fs' i = gso fs i" and
- 04: "\<And> i' j'. i' < m \<Longrightarrow> j' < m \<Longrightarrow>
- \<mu> fs' i' j' = (if i' = i \<and> j' \<le> j then \<mu> fs i j' - of_int c * \<mu> fs j j' else \<mu> fs i' j')" and
- 05: "\<And> ii. ii \<le> m \<Longrightarrow> d fs' ii = d fs ii" and
- 06: "\<bar>\<mu> fs' i j\<bar> \<le> 1 / 2" and
- 061: "(\<forall>i' j'. i' < i \<longrightarrow> j' \<le> i' \<longrightarrow> \<mu> fs i' j' = \<mu> fs' i' j')"
- using basis_reduction_add_row_main[OF Linvww i j fs'_def] c i by auto
- have 07: "lin_indep fs'" and
- 08: "length fs' = m" and
- 09: "lattice_of fs' = L" using LLL_inv_wD Linvw' by auto
- have 091: "fs_int_indpt n fs'" using 07 using Gram_Schmidt_2.fs_int_indpt.intro by simp
- define I where "I = {(i',j'). i' = i \<and> j' < j}"
- have 10: "I \<subseteq> {(i',j'). i' < m \<and> j' < i'}" "(i,j)\<notin> I" "\<forall>j' \<ge> j. (i,j') \<notin> I" using I_def i j by auto
- obtain fs'' where
- 11: "lattice_of fs'' = L" and
- 12: "map (map_vec (\<lambda> x. x symmod p)) fs'' = map (map_vec (\<lambda> x. x symmod p)) fs'" and
- 13: "lin_indep fs''" and
- 14: "length fs'' = m" and
- 15: "(\<forall> k < m. gso fs'' k = gso fs' k)" and
- 16: "(\<forall> k \<le> m. d fs'' k = d fs' k)" and
- 17: "(\<forall> i' < m. \<forall> j' < m. d\<mu> fs'' i' j' =
- (if (i',j') \<in> I then d\<mu> fs' i' j' symmod (p * d fs' j' * d fs' (Suc j')) else d\<mu> fs' i' j'))"
- using mod_finite_set[OF 07 08 10(1) 09 pgtz] by blast
- have 171: "(\<forall>i' j'. i' < i \<longrightarrow> j' \<le> i' \<longrightarrow> \<mu> fs'' i' j' = \<mu> fs' i' j')"
- proof -
- {
- fix i' j'
- assume i'j': "i' < i" "j' \<le> i'"
- have "rat_of_int (d\<mu> fs'' i' j') = rat_of_int (d\<mu> fs' i' j')" using "17" I_def i i'j' by auto
- then have "rat_of_int (int_of_rat (rat_of_int (d fs'' (Suc j')) * \<mu> fs'' i' j')) =
- rat_of_int (int_of_rat (rat_of_int (d fs' (Suc j')) * \<mu> fs' i' j'))"
- using d\<mu>_def i'j' j by auto
- then have "rat_of_int (d fs'' (Suc j')) * \<mu> fs'' i' j' =
- rat_of_int (d fs' (Suc j')) * \<mu> fs' i' j'"
- by (smt "08" "091" "13" "14" d_def dual_order.strict_trans fs_int.d_def
- fs_int_indpt.fs_int_mu_d_Z fs_int_indpt.intro i i'j'(1) i'j'(2) int_of_rat(2))
- then have "\<mu> fs'' i' j' = \<mu> fs' i' j'" by (smt "16"
- LLL_d_pos[OF Linvw'] Suc_leI int_of_rat(1)
- dual_order.strict_trans fs'_def i i'j' j
- le_neq_implies_less nonzero_mult_div_cancel_left of_int_hom.hom_zero)
- }
- then show ?thesis by simp
- qed
- then have 172: "(\<forall>i' j'. i' < i \<longrightarrow> j' \<le> i' \<longrightarrow> \<mu> fs'' i' j' = \<mu> fs i' j')" using 061 by simp (* goal *)
- have 18: "LLL_measure i fs'' = LLL_measure i fs'" using 16 LLL_measure_def logD_def D_def by simp
- have 19: "(\<forall>k < m. gso fs'' k = gso fs k)" using 03 15 by simp
- have "\<forall>j' \<in> {j..(m-1)}. j' < m" using j i by auto
- then have 20: "\<forall>j' \<in> {j..(m-1)}. d\<mu> fs'' i j' = d\<mu> fs' i j'"
- using 10(3) 17 Suc_lessD less_trans_Suc by (meson atLeastAtMost_iff i)
- have 21: "\<forall>j' \<in> {j..(m-1)}. \<mu> fs'' i j' = \<mu> fs' i j'"
- proof -
- {
- fix j'
- assume j': "j' \<in> {j..(m-1)}"
- define \<mu>'' :: rat where "\<mu>'' = \<mu> fs'' i j'"
- define \<mu>' :: rat where "\<mu>' = \<mu> fs' i j'"
- have "rat_of_int (d\<mu> fs'' i j') = rat_of_int (d\<mu> fs' i j')" using 20 j' by simp
- moreover have "j' < length fs'" using i j' 08 by auto
- ultimately have "rat_of_int (d fs' (Suc j')) * gram_schmidt_fs.\<mu> n (map of_int_hom.vec_hom fs') i j'
- = rat_of_int (d fs'' (Suc j')) * gram_schmidt_fs.\<mu> n (map of_int_hom.vec_hom fs'') i j'"
- using 20 08 091 13 14 fs_int_indpt.d\<mu>_def fs_int.d_def fs_int_indpt.d\<mu> d\<mu>_def d_def i fs_int_indpt.intro j'
- by metis
- then have "rat_of_int (d fs' (Suc j')) * \<mu>'' = rat_of_int (d fs' (Suc j')) * \<mu>'"
- using 16 i j' \<mu>'_def \<mu>''_def unfolding d\<mu>_def by auto
- moreover have "0 < d fs' (Suc j')" using LLL_d_pos[OF Linvw', of "Suc j'"] i j' by auto
- ultimately have "\<mu> fs'' i j' = \<mu> fs' i j'" using \<mu>'_def \<mu>''_def by simp
- }
- then show ?thesis by simp
- qed
- then have 22: "\<mu> fs'' i j = \<mu> fs' i j" using i j by simp
- then have 23: "\<bar>\<mu> fs'' i j\<bar> \<le> 1 / 2" using 06 by simp (* goal *)
- have 24: "LLL_measure i fs'' = LLL_measure i fs" using 02 18 by simp (* goal *)
- have 25: "(\<forall> k \<le> m. d fs'' k = d fs k)" using 16 05 by simp (* goal *)
- have 26: "(\<forall> k < m. gso fs'' k = gso fs k)" using 15 03 by simp (* goal *)
- have 27: "\<mu>_small_row i fs (Suc j) \<Longrightarrow> \<mu>_small_row i fs'' j"
- using 21 01 \<mu>_small_row_def i j c by auto (* goal *)
- have 28: "length fs = m" "length mfs = m" using LLL_invD_modw[OF Linvmw] by auto
- have 29: "map (map_vec (\<lambda>x. x symmod p)) fs = mfs" using assms LLL_invD_modw by simp
- have 30: "\<And> i. i < m \<Longrightarrow> fs ! i \<in> carrier_vec n" "\<And> i. i < m \<Longrightarrow> mfs ! i \<in> carrier_vec n"
- using LLL_invD_modw[OF Linvmw] by auto
- have 31: "\<And> i. i < m \<Longrightarrow> fs' ! i \<in> carrier_vec n" using fs'_def 30(1)
- using "08" "091" fs_int_indpt.f_carrier by blast
- have 32: "\<And> i. i < m \<Longrightarrow> mfs' ! i \<in> carrier_vec n" unfolding mfs' using 30(2) 28(2)
- by (metis (no_types, lifting) Suc_lessD j less_trans_Suc map_carrier_vec minus_carrier_vec
- nth_list_update_eq nth_list_update_neq smult_closed)
- have 33: "length mfs' = m" using 28(2) mfs' by simp (* invariant goal *)
- then have 34: "map (map_vec (\<lambda>x. x symmod p)) fs' = mfs'"
- proof -
- {
- fix i' j'
- have j2: "j < m" using j i by auto
- assume i': "i' < m"
- assume j': "j' < n"
- then have fsij: "(fs ! i' $ j') symmod p = mfs ! i' $ j'" using 30 i' j' 28 29 by fastforce
- have "mfs' ! i $ j' = (mfs ! i $ j'- (c \<cdot>\<^sub>v mfs ! j) $ j') symmod p"
- unfolding mfs' using 30(2) j' 28 j2
- by (metis (no_types, lifting) carrier_vecD i index_map_vec(1) index_minus_vec(1)
- index_minus_vec(2) index_smult_vec(2) nth_list_update_eq)
- then have mfs'ij: "mfs' ! i $ j' = (mfs ! i $ j'- c * mfs ! j $ j') symmod p"
- unfolding mfs' using 30(2) i' j' 28 j2 by fastforce
- have "(fs' ! i' $ j') symmod p = mfs' ! i' $ j'"
- proof(cases "i' = i")
- case True
- show ?thesis using fs'_def mfs' True 28 fsij
- proof -
- have "fs' ! i' $ j' = (fs ! i' - c \<cdot>\<^sub>v fs ! j) $ j'" using fs'_def True i' j' 28(1) by simp
- also have "\<dots> = fs ! i' $ j' - (c \<cdot>\<^sub>v fs ! j) $ j'" using i' j' 30(1)
- by (metis Suc_lessD carrier_vecD i index_minus_vec(1) index_smult_vec(2) j less_trans_Suc)
- finally have "fs' ! i' $ j' = fs ! i' $ j' - (c \<cdot>\<^sub>v fs ! j) $ j'" by auto
- then have "(fs' ! i' $ j') symmod p = (fs ! i' $ j' - (c \<cdot>\<^sub>v fs ! j) $ j') symmod p" by auto
- also have "\<dots> = ((fs ! i' $ j') symmod p - ((c \<cdot>\<^sub>v fs ! j) $ j') symmod p) symmod p"
- by (simp add: sym_mod_diff_eq)
- also have "(c \<cdot>\<^sub>v fs ! j) $ j' = c * (fs ! j $ j')"
- using i' j' True 28 30(1) j
- by (metis Suc_lessD carrier_vecD index_smult_vec(1) less_trans_Suc)
- also have "((fs ! i' $ j') symmod p - (c * (fs ! j $ j')) symmod p) symmod p =
- ((fs ! i' $ j') symmod p - c * ((fs ! j $ j') symmod p)) symmod p"
- using i' j' True 28 30(1) j by (metis sym_mod_diff_right_eq sym_mod_mult_right_eq)
- also have "((fs ! j $ j') symmod p) = mfs ! j $ j'" using 30 i' j' 28 29 j2 by fastforce
- also have "((fs ! i' $ j') symmod p - c * mfs ! j $ j') symmod p =
- (mfs ! i' $ j' - c * mfs ! j $ j') symmod p" using fsij by simp
- finally show ?thesis using mfs'ij by (simp add: True)
- qed
- next
- case False
- show ?thesis using fs'_def mfs' False 28 fsij by simp
- qed
- }
- then have "\<forall>i' < m. (map_vec (\<lambda>x. x symmod p)) (fs' ! i') = mfs' ! i'"
- using 31 32 33 08 by fastforce
- then show ?thesis using 31 32 33 08 by (simp add: map_nth_eq_conv)
- qed
- then have 35: "map (map_vec (\<lambda>x. x symmod p)) fs'' = mfs'" using 12 by simp (* invariant req. *)
- have 36: "lin_indep fs''" using 13 by simp (* invariant req. *)
- have Linvw'': "LLL_invariant_weak fs''" using LLL_invariant_weak_def 11 13 14 by simp
- have 39: "(\<forall>i' < m. \<forall>j' < i'. \<bar>d\<mu> fs'' i' j'\<bar> < p * d fs'' j' * d fs'' (Suc j'))" (* invariant req. *)
- proof -
- {
- fix i' j'
- assume i': "i' < m"
- assume j': "j' < i'"
- define pdd where "pdd = (p * d fs'' j' * d fs'' (Suc j'))"
- then have pddgtz: "pdd > 0"
- using pgtz j' LLL_d_pos[OF Linvw', of "Suc j'"] LLL_d_pos[OF Linvw', of j'] j' i' 16 by simp
- have "\<bar>d\<mu> fs'' i' j'\<bar> < p * d fs'' j' * d fs'' (Suc j')"
- proof(cases "i' = i")
- case i'i: True
- then show ?thesis
- proof (cases "j' < j")
- case True
- then have eq'': "d\<mu> fs'' i' j' = d\<mu> fs' i' j' symmod (p * d fs'' j' * d fs'' (Suc j'))"
- using 16 17 10 I_def True i' j' i'i by simp
- have "0 < pdd" using pddgtz by simp
- then show ?thesis unfolding eq'' unfolding pdd_def[symmetric] using sym_mod_abs by blast
- next
- case fls: False
- then have "(i',j') \<notin> I" using I_def i'i by simp
- then have dmufs''fs': "d\<mu> fs'' i' j' = d\<mu> fs' i' j'" using 17 i' j' by simp
- show ?thesis
- proof (cases "j' = j")
- case True
- define \<mu>'' where "\<mu>'' = \<mu> fs'' i' j'"
- define d'' where "d'' = d fs'' (Suc j')"
- have pge1: "p \<ge> 1" using pgtz by simp
- have lh: "\<bar>\<mu>''\<bar> \<le> 1 / 2" using 23 True i'i \<mu>''_def by simp
- moreover have eq: "d\<mu> fs'' i' j' = \<mu>'' * d''" using d\<mu>_def i' j' \<mu>''_def d''_def
- by (smt "14" "36" LLL.d_def Suc_lessD fs_int.d_def fs_int_indpt.d\<mu> fs_int_indpt.intro
- int_of_rat(1) less_trans_Suc mult_of_int_commute of_rat_mult of_rat_of_int_eq)
- moreover have Sj': "Suc j' \<le> m" "j' \<le> m" using True j' i i' by auto
- moreover then have gtz: "0 < d''" using LLL_d_pos[OF Linvw''] d''_def by simp
- moreover have "rat_of_int \<bar>d\<mu> fs'' i' j'\<bar> = \<bar>\<mu>'' * (rat_of_int d'')\<bar>"
- using eq by (metis of_int_abs of_rat_hom.injectivity of_rat_mult of_rat_of_int_eq)
- moreover then have "\<bar>\<mu>'' * rat_of_int d'' \<bar> = \<bar>\<mu>''\<bar> * rat_of_int \<bar>d''\<bar>"
- by (metis (mono_tags, opaque_lifting) abs_mult of_int_abs)
- moreover have "\<dots> = \<bar>\<mu>''\<bar> * rat_of_int d'' " using gtz by simp
- moreover have "\<dots> < rat_of_int d''" using lh gtz by simp
- ultimately have "rat_of_int \<bar>d\<mu> fs'' i' j'\<bar> < rat_of_int d''" by simp
- then have "\<bar>d\<mu> fs'' i' j'\<bar> < d fs'' (Suc j')" using d''_def by simp
- then have "\<bar>d\<mu> fs'' i' j'\<bar> < p * d fs'' (Suc j')" using pge1
- by (smt mult_less_cancel_right2)
- then show ?thesis using pge1 LLL_d_pos[OF Linvw'' Sj'(2)] gtz unfolding d''_def
- by (smt mult_less_cancel_left2 mult_right_less_imp_less)
- next
- case False
- have "j' < m" using i' j' by simp
- moreover have "j' > j" using False fls by simp
- ultimately have "\<mu> fs' i' j' = \<mu> fs i' j'" using i' 04 i by simp
- then have "d\<mu> fs' i' j' = d\<mu> fs i' j'" using d\<mu>_def i' j' 05 by simp
- then have "d\<mu> fs'' i' j' = d\<mu> fs i' j'" using dmufs''fs' by simp
- then show ?thesis using LLL_invD_modw[OF Linvmw] i' j' 25 by simp
- qed
- qed
- next
- case False
- then have "(i',j') \<notin> I" using I_def by simp
- then have dmufs''fs': "d\<mu> fs'' i' j' = d\<mu> fs' i' j'" using 17 i' j' by simp
- have "\<mu> fs' i' j' = \<mu> fs i' j'" using i' 04 j' False by simp
- then have "d\<mu> fs' i' j' = d\<mu> fs i' j'" using d\<mu>_def i' j' 05 by simp
- moreover then have "d\<mu> fs'' i' j' = d\<mu> fs i' j'" using dmufs''fs' by simp
- then show ?thesis using LLL_invD_modw[OF Linvmw] i' j' 25 by simp
- qed
- }
- then show ?thesis by simp
- qed
- have 40: "(\<forall>i' < m. \<forall>j' < m. i' \<noteq> i \<or> j' > j \<longrightarrow> d\<mu> fs' i' j' = dmu $$ (i',j'))"
- proof -
- {
- fix i' j'
- assume i': "i' < m" and j': "j' < m"
- assume assm: "i' \<noteq> i \<or> j' > j"
- have "d\<mu> fs' i' j' = dmu $$ (i',j')"
- proof (cases "i' \<noteq> i")
- case True
- then show ?thesis using fs'_def LLL_invD_modw[OF Linvmw] d\<mu>_def i i' j j'
- 04 28(1) LLL_invI_weak basis_reduction_add_row_main(8)[OF Linvww] by auto
- next
- case False
- then show ?thesis
- using 05 LLL_invD_modw[OF Linvmw] d\<mu>_def i j j' 04 assm by simp
- qed
- }
- then show ?thesis by simp
- qed
- have 41: "\<forall>j' \<le> j. d\<mu> fs' i j' = dmu $$ (i,j') - c * dmu $$ (j,j')"
- proof -
- {
- let ?oi = "of_int :: _ \<Rightarrow> rat"
- fix j'
- assume j': "j' \<le> j"
- define dj' \<mu>i \<mu>j where "dj' = d fs (Suc j')" and "\<mu>i = \<mu> fs i j'" and "\<mu>j = \<mu> fs j j'"
- have "?oi (d\<mu> fs' i j') = ?oi (d fs (Suc j')) * (\<mu> fs i j' - ?oi c * \<mu> fs j j')"
- using j' 04 d\<mu>_def
- by (smt "05" "08" "091" Suc_leI d_def diff_diff_cancel fs_int.d_def
- fs_int_indpt.fs_int_mu_d_Z i int_of_rat(2) j less_imp_diff_less less_imp_le_nat)
- also have "\<dots> = (?oi dj') * (\<mu>i - of_int c * \<mu>j)"
- using dj'_def \<mu>i_def \<mu>j_def by (simp add: of_rat_mult)
- also have "\<dots> = (rat_of_int dj') * \<mu>i - of_int c * (rat_of_int dj') * \<mu>j" by algebra
- also have "\<dots> = rat_of_int (d\<mu> fs i j') - ?oi c * rat_of_int (d\<mu> fs j j')" unfolding dj'_def \<mu>i_def \<mu>j_def
- using i j j' d\<mu>_def
- using "28"(1) LLL.LLL_invD_modw(4) Linvmw d_def fs_int.d_def fs_int_indpt.fs_int_mu_d_Z fs_int_indpt.intro by auto
- also have "\<dots> = rat_of_int (dmu $$ (i,j')) - ?oi c * rat_of_int (dmu $$ (j,j'))"
- using LLL_invD_modw(7)[OF Linvmw] d\<mu>_def j' i j by auto
- finally have "?oi (d\<mu> fs' i j') = rat_of_int (dmu $$ (i,j')) - ?oi c * rat_of_int (dmu $$ (j,j'))" by simp
- then have "d\<mu> fs' i j' = dmu $$ (i,j') - c * dmu $$ (j,j')"
- using of_int_eq_iff by fastforce
- }
- then show ?thesis by simp
- qed
- have 42: "(\<forall>i' < m. \<forall>j' < m. d\<mu> fs'' i' j' = dmu' $$ (i',j'))"
- proof -
- {
- fix i' j'
- assume i': "i' < m" and j': "j' < m"
- have "d\<mu> fs'' i' j' = dmu' $$ (i',j')"
- proof (cases "i' = i")
- case i'i: True
- then show ?thesis
- proof (cases "j' > j")
- case True
- then have "(i',j')\<notin>I" using I_def by simp
- moreover then have "d\<mu> fs' i' j' = d\<mu> fs i' j'" using "04" "05" True Suc_leI d\<mu>_def i' j' by simp
- moreover have "dmu' $$ (i',j') = dmu $$ (i',j')" using dmu' True i' j' by simp
- ultimately show ?thesis using "17" "40" True i' j' by auto
- next
- case False
- then have j'lej: "j' \<le> j" by simp
- then have eq': "d\<mu> fs' i j' = dmu $$ (i,j') - c * dmu $$ (j,j')" using 41 by simp
- have id: "d_of dmu j' = d fs j'" "d_of dmu (Suc j') = d fs (Suc j')"
- using d_of_weak[OF Linvmw] \<open>j' < m\<close> by auto
- show ?thesis
- proof (cases "j' \<noteq> j")
- case True
- then have j'ltj: "j' < j" using True False by simp
- then have "(i',j') \<in> I" using I_def True i'i by simp
- then have "d\<mu> fs'' i' j' =
- (dmu $$ (i,j') - c * dmu $$ (j,j')) symmod (p * d fs' j' * d fs' (Suc j'))"
- using 17 i' 41 j'lej by (simp add: j' i'i)
- also have "\<dots> = (dmu $$ (i,j') - c * dmu $$ (j,j')) symmod (p * d fs j' * d fs (Suc j'))"
- using 05 i j'ltj j by simp
- also have "\<dots> = dmu' $$ (i,j')"
- unfolding dmu' index_mat(1)[OF \<open>i < m\<close> \<open>j' < m\<close>] split id using j'lej True by auto
- finally show ?thesis using i'i by simp
- next
- case False
- then have j'j: "j' = j" by simp
- then have "d\<mu> fs'' i j' = d\<mu> fs' i j'" using 20 j' by simp
- also have "\<dots> = dmu $$ (i,j') - c * dmu $$ (j,j')" using eq' by simp
- also have "\<dots> = dmu' $$ (i,j')" using dmu' j'j i j' by simp
- finally show ?thesis using i'i by simp
- qed
- qed
- next
- case False
- then have "(i',j')\<notin>I" using I_def by simp
- moreover then have "d\<mu> fs' i' j' = d\<mu> fs i' j'" by (simp add: "04" "05" False Suc_leI d\<mu>_def i' j')
- moreover then have "dmu' $$ (i',j') = dmu $$ (i',j')" using dmu' False i' j' by simp
- ultimately show ?thesis using "17" "40" False i' j' by auto
- qed
- }
- then show ?thesis by simp
- qed
- from gbnd 26 have gbnd: "g_bnd_mode first b fs''" using g_bnd_mode_cong[of fs'' fs] by simp
- {
- assume Linv: "LLL_invariant_mod fs mfs dmu p first b i"
- have Linvw: "LLL_invariant_weak' i fs" using Linv LLL_invD_mod LLL_invI_weak by simp
- note Linvww = LLL_invw'_imp_w[OF Linvw]
- have 00: "LLL_invariant_weak' i fs'" using Linvw basis_reduction_add_row_weak[OF Linvw i j fs'_def] by auto
- have 37: "weakly_reduced fs'' i" using 15 LLL_invD_weak(8)[OF 00] gram_schmidt_fs.weakly_reduced_def
- by (smt Suc_lessD i less_trans_Suc) (* invariant req. *)
- have 38: "LLL_invariant_weak' i fs''"
- using 00 11 14 36 37 i 31 12 LLL_invariant_weak'_def by blast
- have "LLL_invariant_mod fs'' mfs' dmu' p first b i"
- using LLL_invI_mod[OF 33 _ 14 11 13 37 35 39 42 p1 gbnd LLL_invD_mod(17)[OF Linv]] i by simp
- }
- moreover have "LLL_invariant_mod_weak fs'' mfs' dmu' p first b"
- using LLL_invI_modw[OF 33 14 11 13 35 39 42 p1 gbnd LLL_invD_modw(15)[OF Linvmw]] by simp
- ultimately show ?thesis using 27 23 24 25 26 172 by auto
-qed
-
-definition D_mod :: "int mat \<Rightarrow> nat" where "D_mod dmu = nat (\<Prod> i < m. d_of dmu i)"
-
-definition logD_mod :: "int mat \<Rightarrow> nat"
- where "logD_mod dmu = (if \<alpha> = 4/3 then (D_mod dmu) else nat (floor (log (1 / of_rat reduction) (D_mod dmu))))"
-end
-
-locale fs_int'_mod =
- fixes n m fs_init \<alpha> i fs mfs dmu p first b
- assumes LLL_inv_mod: "LLL.LLL_invariant_mod n m fs_init \<alpha> fs mfs dmu p first b i"
-
-context LLL_with_assms
-begin
-
-lemma basis_reduction_swap_weak': assumes Linvw: "LLL_invariant_weak' i fs"
- and i: "i < m"
- and i0: "i \<noteq> 0"
- and mu_F1_i: "\<bar>\<mu> fs i (i-1)\<bar> \<le> 1 / 2"
- and norm_ineq: "sq_norm (gso fs (i - 1)) > \<alpha> * sq_norm (gso fs i)"
- and fs'_def: "fs' = fs[i := fs ! (i - 1), i - 1 := fs ! i]"
-shows "LLL_invariant_weak' (i - 1) fs'"
-proof -
- note inv = LLL_invD_weak[OF Linvw]
- note invw = LLL_invw'_imp_w[OF Linvw]
- note main = basis_reduction_swap_main[OF invw disjI2[OF mu_F1_i] i i0 norm_ineq fs'_def]
- note inv' = LLL_inv_wD[OF main(1)]
- from \<open>weakly_reduced fs i\<close> have "weakly_reduced fs (i - 1)"
- unfolding gram_schmidt_fs.weakly_reduced_def by auto
- also have "weakly_reduced fs (i - 1) = weakly_reduced fs' (i - 1)"
- unfolding gram_schmidt_fs.weakly_reduced_def
- by (intro all_cong, insert i0 i main(5), auto)
- finally have red: "weakly_reduced fs' (i - 1)" .
- show "LLL_invariant_weak' (i - 1) fs'" using i
- by (intro LLL_invI_weak red inv', auto)
-qed
-
-lemma basis_reduction_add_row_done_weak:
- assumes Linv: "LLL_invariant_weak' i fs"
- and i: "i < m"
- and mu_small: "\<mu>_small_row i fs 0"
-shows "\<mu>_small fs i"
-proof -
- note inv = LLL_invD_weak[OF Linv]
- from mu_small
- have mu_small: "\<mu>_small fs i" unfolding \<mu>_small_row_def \<mu>_small_def by auto
- show ?thesis
- using i mu_small LLL_invI_weak[OF inv(3,6,7,9,1)] by auto
-qed
-
-lemma LLL_invariant_mod_to_weak_m_to_i: assumes
- inv: "LLL_invariant_mod fs mfs dmu p first b m"
- and i: "i \<le> m"
-shows "LLL_invariant_mod fs mfs dmu p first b i"
- "LLL_invariant_weak' m fs"
- "LLL_invariant_weak' i fs"
-proof -
- show "LLL_invariant_mod fs mfs dmu p first b i"
- proof -
- have "LLL_invariant_weak' m fs" using LLL_invD_mod[OF inv] LLL_invI_weak by simp
- then have "LLL_invariant_weak' i fs" using LLL_inv_weak_m_impl_i i by simp
- then have "weakly_reduced fs i" using i LLL_invD_weak(8) by simp
- then show ?thesis using LLL_invD_mod[OF inv] LLL_invI_mod i by simp
- qed
- then show fsinvwi: "LLL_invariant_weak' i fs" using LLL_invD_mod LLL_invI_weak by simp
- show "LLL_invariant_weak' m fs" using LLL_invD_mod[OF inv] LLL_invI_weak by simp
-qed
-
-lemma basis_reduction_mod_swap_main:
- assumes Linvmw: "LLL_invariant_mod_weak fs mfs dmu p first b"
- and k: "k < m"
- and k0: "k \<noteq> 0"
- and mu_F1_i: "\<bar>\<mu> fs k (k-1)\<bar> \<le> 1 / 2"
- and norm_ineq: "sq_norm (gso fs (k - 1)) > \<alpha> * sq_norm (gso fs k)"
- and mfs'_def: "mfs' = mfs[k := mfs ! (k - 1), k - 1 := mfs ! k]"
- and dmu'_def: "dmu' = (mat m m (\<lambda>(i,j). (
- if j < i then
- if i = k - 1 then
- dmu $$ (k, j)
- else if i = k \<and> j \<noteq> k - 1 then
- dmu $$ (k - 1, j)
- else if i > k \<and> j = k then
- ((d_of dmu (Suc k)) * dmu $$ (i, k - 1) - dmu $$ (k, k - 1) * dmu $$ (i, j))
- div (d_of dmu k)
- else if i > k \<and> j = k - 1 then
- (dmu $$ (k, k - 1) * dmu $$ (i, j) + dmu $$ (i, k) * (d_of dmu (k-1)))
- div (d_of dmu k)
- else dmu $$ (i, j)
- else if i = j then
- if i = k - 1 then
- ((d_of dmu (Suc k)) * (d_of dmu (k-1)) + dmu $$ (k, k - 1) * dmu $$ (k, k - 1))
- div (d_of dmu k)
- else (d_of dmu (Suc i))
- else dmu $$ (i, j))
- ))"
- and dmu'_mod_def: "dmu'_mod = mat m m (\<lambda>(i, j). (
- if j < i \<and> (j = k \<or> j = k - 1) then
- dmu' $$ (i, j) symmod (p * (d_of dmu' j) * (d_of dmu' (Suc j)))
- else dmu' $$ (i, j)))"
-shows "(\<exists>fs'. LLL_invariant_mod_weak fs' mfs' dmu'_mod p first b \<and>
- LLL_measure (k-1) fs' < LLL_measure k fs \<and>
- (LLL_invariant_mod fs mfs dmu p first b k \<longrightarrow> LLL_invariant_mod fs' mfs' dmu'_mod p first b (k-1)))"
-proof -
- define fs' where "fs' = fs[k := fs ! (k - 1), k - 1 := fs ! k]"
- have pgtz: "p > 0" and p1: "p > 1" using LLL_invD_modw[OF Linvmw] by auto
- have invw: "LLL_invariant_weak fs" using LLL_invD_modw[OF Linvmw] LLL_invariant_weak_def by simp
- note swap_main = basis_reduction_swap_main(3-)[OF invw disjI2[OF mu_F1_i] k k0 norm_ineq fs'_def]
- note dd\<mu>_swap = d_d\<mu>_swap[OF invw disjI2[OF mu_F1_i] k k0 norm_ineq fs'_def]
- have invw': "LLL_invariant_weak fs'" using fs'_def assms invw basis_reduction_swap_main(1) by simp
- have 02: "LLL_measure k fs > LLL_measure (k - 1) fs'" by fact
- have 03: "\<And> i j. i < m \<Longrightarrow> j < i \<Longrightarrow>
- d\<mu> fs' i j = (
- if i = k - 1 then
- d\<mu> fs k j
- else if i = k \<and> j \<noteq> k - 1 then
- d\<mu> fs (k - 1) j
- else if i > k \<and> j = k then
- (d fs (Suc k) * d\<mu> fs i (k - 1) - d\<mu> fs k (k - 1) * d\<mu> fs i j) div d fs k
- else if i > k \<and> j = k - 1 then
- (d\<mu> fs k (k - 1) * d\<mu> fs i j + d\<mu> fs i k * d fs (k - 1)) div d fs k
- else d\<mu> fs i j)"
- using dd\<mu>_swap by auto
- have 031: "\<And>i. i < k-1 \<Longrightarrow> gso fs' i = gso fs i"
- using swap_main(2) k k0 by auto
- have 032: "\<And> ii. ii \<le> m \<Longrightarrow> of_int (d fs' ii) = (if ii = k then
- sq_norm (gso fs' (k - 1)) / sq_norm (gso fs (k - 1)) * of_int (d fs k)
- else of_int (d fs ii))"
- by fact
- have gbnd: "g_bnd_mode first b fs'"
- proof (cases "first \<and> m \<noteq> 0")
- case True
- have "sq_norm (gso fs' 0) \<le> sq_norm (gso fs 0)"
- proof (cases "k - 1 = 0")
- case False
- thus ?thesis using 031[of 0] by simp
- next
- case *: True
- have k_1: "k - 1 < m" using k by auto
- from * k0 have k1: "k = 1" by simp
- (* this is a copy of what is done in LLL.swap-main, should be made accessible in swap-main *)
- have "sq_norm (gso fs' 0) \<le> abs (sq_norm (gso fs' 0))" by simp
- also have "\<dots> = abs (sq_norm (gso fs 1) + \<mu> fs 1 0 * \<mu> fs 1 0 * sq_norm (gso fs 0))"
- by (subst swap_main(3)[OF k_1, unfolded *], auto simp: k1)
- also have "\<dots> \<le> sq_norm (gso fs 1) + abs (\<mu> fs 1 0) * abs (\<mu> fs 1 0) * sq_norm (gso fs 0)"
- by (simp add: sq_norm_vec_ge_0)
- also have "\<dots> \<le> sq_norm (gso fs 1) + (1 / 2) * (1 / 2) * sq_norm (gso fs 0)"
- using mu_F1_i[unfolded k1]
- by (intro plus_right_mono mult_mono, auto)
- also have "\<dots> < 1 / \<alpha> * sq_norm (gso fs 0) + (1 / 2) * (1 / 2) * sq_norm (gso fs 0)"
- by (intro add_strict_right_mono, insert norm_ineq[unfolded mult.commute[of \<alpha>],
- THEN mult_imp_less_div_pos[OF \<alpha>0(1)]] k1, auto)
- also have "\<dots> = reduction * sq_norm (gso fs 0)" unfolding reduction_def
- using \<alpha>0 by (simp add: ring_distribs add_divide_distrib)
- also have "\<dots> \<le> 1 * sq_norm (gso fs 0)" using reduction(2)
- by (intro mult_right_mono, auto)
- finally show ?thesis by simp
- qed
- thus ?thesis using LLL_invD_modw(14)[OF Linvmw] True
- unfolding g_bnd_mode_def by auto
- next
- case False
- from LLL_invD_modw(14)[OF Linvmw] False have "g_bnd b fs" unfolding g_bnd_mode_def by auto
- hence "g_bnd b fs'" using g_bnd_swap[OF k k0 invw mu_F1_i norm_ineq fs'_def] by simp
- thus ?thesis using False unfolding g_bnd_mode_def by auto
- qed
- note d_of = d_of_weak[OF Linvmw]
- have 033: "\<And> i. i < m \<Longrightarrow> d\<mu> fs' i i = (
- if i = k - 1 then
- ((d_of dmu (Suc k)) * (d_of dmu (k-1)) + dmu $$ (k, k - 1) * dmu $$ (k, k - 1))
- div (d_of dmu k)
- else (d_of dmu (Suc i)))"
- proof -
- fix i
- assume i: "i < m"
- have "d\<mu> fs' i i = d fs' (Suc i)" using dd\<mu> i by simp
- also have "\<dots> = (if i = k - 1 then
- (d fs (Suc k) * d fs (k - 1) + d\<mu> fs k (k - 1) * d\<mu> fs k (k - 1)) div d fs k
- else d fs (Suc i))"
- by (subst dd\<mu>_swap, insert dd\<mu> k0 i, auto)
- also have "\<dots> = (if i = k - 1 then
- ((d_of dmu (Suc k)) * (d_of dmu (k-1)) + dmu $$ (k, k - 1) * dmu $$ (k, k - 1))
- div (d_of dmu k)
- else (d_of dmu (Suc i)))" (is "_ = ?r")
- using d_of i k LLL_invD_modw(7)[OF Linvmw] by auto
- finally show "d\<mu> fs' i i = ?r" .
- qed
- have 04: "lin_indep fs'" "length fs' = m" "lattice_of fs' = L" using LLL_inv_wD[OF invw'] by auto
- define I where "I = {(i, j). i < m \<and> j < i \<and> (j = k \<or> j = k - 1)}"
- then have Isubs: "I \<subseteq> {(i,j). i < m \<and> j < i}" using k k0 by auto
- obtain fs'' where
- 05: "lattice_of fs'' = L" and
- 06: "map (map_vec (\<lambda> x. x symmod p)) fs'' = map (map_vec (\<lambda> x. x symmod p)) fs'" and
- 07: "lin_indep fs''" and
- 08: "length fs'' = m" and
- 09: "(\<forall> k < m. gso fs'' k = gso fs' k)" and
- 10: "(\<forall> k \<le> m. d fs'' k = d fs' k)" and
- 11: "(\<forall> i' < m. \<forall> j' < m. d\<mu> fs'' i' j' =
- (if (i',j') \<in> I then d\<mu> fs' i' j' symmod (p * d fs' j' * d fs' (Suc j')) else d\<mu> fs' i' j'))"
- using mod_finite_set[OF 04(1) 04(2) Isubs 04(3) pgtz] by blast
- have 13: "length mfs' = m" using mfs'_def LLL_invD_modw(1)[OF Linvmw] by simp (* invariant requirement *)
- have 14: "map (map_vec (\<lambda> x. x symmod p)) fs'' = mfs'" (* invariant requirement *)
- using 06 fs'_def k k0 04(2) LLL_invD_modw(5)[OF Linvmw]
- by (metis (no_types, lifting) length_list_update less_imp_diff_less map_update mfs'_def nth_map)
- have "LLL_measure (k - 1) fs'' = LLL_measure (k - 1) fs'" using 10 LLL_measure_def logD_def D_def by simp
- then have 15: "LLL_measure (k - 1) fs'' < LLL_measure k fs" using 02 by simp (* goal *)
- {
- fix i' j'
- assume i'j': "i'<m" "j'<i'"
- and neq: "j' \<noteq> k" "j' \<noteq> k - 1"
- hence j'k: "j' \<noteq> k" "Suc j' \<noteq> k" using k0 by auto
- hence "d fs'' j' = d fs j'" "d fs'' (Suc j') = d fs (Suc j')"
- using \<open>k < m\<close> i'j' k0
- 10[rule_format, of j'] 032[rule_format, of j']
- 10[rule_format, of "Suc j'"] 032[rule_format, of "Suc j'"]
- by auto
- } note d_id = this
-
- have 16: "\<forall>i'<m. \<forall>j'<i'. \<bar>d\<mu> fs'' i' j'\<bar> < p * d fs'' j' * d fs'' (Suc j')" (* invariant requirement *)
- proof -
- {
- fix i' j'
- assume i'j': "i'<m" "j'<i'"
- have "\<bar>d\<mu> fs'' i' j'\<bar> < p * d fs'' j' * d fs'' (Suc j')"
- proof (cases "(i',j') \<in> I")
- case True
- define pdd where "pdd = (p * d fs' j' * d fs' (Suc j'))"
- have pdd_pos: "pdd > 0" using pgtz i'j' LLL_d_pos[OF invw'] pdd_def by simp
- have "d\<mu> fs'' i' j' = d\<mu> fs' i' j' symmod pdd" using True 11 i'j' pdd_def by simp
- then have "\<bar>d\<mu> fs'' i' j'\<bar> < pdd" using True 11 i'j' pdd_pos sym_mod_abs by simp
- then show ?thesis unfolding pdd_def using 10 i'j' by simp
- next
- case False
- from False[unfolded I_def] i'j' have neg: "j' \<noteq> k" "j' \<noteq> k - 1" by auto
-
- consider (1) "i' = k - 1 \<or> i' = k" | (2) "\<not> (i' = k - 1 \<or> i' = k)"
- using False i'j' unfolding I_def by linarith
- thus ?thesis
- proof cases
- case **: 1
- let ?i'' = "if i' = k - 1 then k else k -1"
- from ** neg i'j' have i'': "?i'' < m" "j' < ?i''" using k0 k by auto
- have "d\<mu> fs'' i' j' = d\<mu> fs' i' j'" using 11 False i'j' by simp
- also have "\<dots> = d\<mu> fs ?i'' j'" unfolding 03[OF \<open>i' < m\<close> \<open>j' < i'\<close>]
- using ** neg by auto
- finally show ?thesis using LLL_invD_modw(6)[OF Linvmw, rule_format, OF i''] unfolding d_id[OF i'j' neg] by auto
- next
- case **: 2
- hence neq: "j' \<noteq> k" "j' \<noteq> k - 1" using False k k0 i'j' unfolding I_def by auto
- have "d\<mu> fs'' i' j' = d\<mu> fs' i' j'" using 11 False i'j' by simp
- also have "\<dots> = d\<mu> fs i' j'" unfolding 03[OF \<open>i' < m\<close> \<open>j' < i'\<close>] using ** neq by auto
- finally show ?thesis using LLL_invD_modw(6)[OF Linvmw, rule_format, OF i'j'] using d_id[OF i'j' neq] by auto
- qed
- qed
- }
- then show ?thesis by simp
- qed
- have 17: "\<forall>i'<m. \<forall>j'<m. d\<mu> fs'' i' j' = dmu'_mod $$ (i', j')" (* invariant requirement *)
- proof -
- {
- fix i' j'
- assume i'j': "i'<m" "j'<i'"
- have d'dmu': "\<forall>j' < m. d fs' (Suc j') = dmu' $$ (j', j')" using dd\<mu> dmu'_def 033 by simp
- have eq': "d\<mu> fs' i' j' = dmu' $$ (i', j')"
- proof -
- have t00: "d\<mu> fs k j' = dmu $$ (k, j')" and
- t01: "d\<mu> fs (k - 1) j' = dmu $$ (k - 1, j')" and
- t04: "d\<mu> fs k (k - 1) = dmu $$ (k, k - 1)" and
- t05: "d\<mu> fs i' k = dmu $$ (i', k)"
- using LLL_invD_modw(7)[OF Linvmw] i'j' k dd\<mu> k0 by auto
- have t03: "d fs k = d\<mu> fs (k-1) (k-1)" using k0 k by (metis LLL.dd\<mu> Suc_diff_1 lessI not_gr_zero)
- have t06: "d fs (k - 1) = (d_of dmu (k-1))" using d_of k by auto
- have t07: "d fs k = (d_of dmu k)" using d_of k by auto
- have j': "j' < m" using i'j' by simp
- have "d\<mu> fs' i' j' = (if i' = k - 1 then
- dmu $$ (k, j')
- else if i' = k \<and> j' \<noteq> k - 1 then
- dmu $$ (k - 1, j')
- else if i' > k \<and> j' = k then
- (dmu $$ (k, k) * dmu $$ (i', k - 1) - dmu $$ (k, k - 1) * dmu $$ (i', j')) div (d_of dmu k)
- else if i' > k \<and> j' = k - 1 then
- (dmu $$ (k, k - 1) * dmu $$ (i', j') + dmu $$ (i', k) * d fs (k - 1)) div (d_of dmu k)
- else dmu $$ (i', j'))"
- using dd\<mu> k t00 t01 t03 LLL_invD_modw(7)[OF Linvmw] k i'j' j' 03 t07 by simp
- then show ?thesis using dmu'_def i'j' j' t06 t07 by (simp add: d_of_def)
- qed
- have "d\<mu> fs'' i' j' = dmu'_mod $$ (i', j')"
- proof (cases "(i',j') \<in> I")
- case i'j'I: True
- have j': "j' < m" using i'j' by simp
- show ?thesis
- proof -
- have "dmu'_mod $$ (i',j') = dmu' $$ (i',j')
- symmod (p * (d_of dmu' j') * (d_of dmu' (Suc j')))"
- using dmu'_mod_def i'j' i'j'I I_def by simp
- also have "d_of dmu' j' = d fs' j'"
- using j' d'dmu' d_def Suc_diff_1 less_imp_diff_less unfolding d_of_def
- by (cases j', auto)
- finally have "dmu'_mod $$ (i',j') = dmu' $$ (i',j') symmod (p * d fs' j' * d fs' (Suc j'))"
- using dd\<mu>[OF j'] d'dmu' j' by (auto simp: d_of_def)
- then show ?thesis using i'j'I 11 i'j' eq' by simp
- qed
- next
- case False
- have "d\<mu> fs'' i' j' = d\<mu> fs' i' j'" using False 11 i'j' by simp
- also have "\<dots> = dmu' $$ (i', j')" unfolding eq' ..
- finally show ?thesis unfolding dmu'_mod_def using False[unfolded I_def] i'j' by auto
- qed
- }
- moreover have "\<forall>i' j'. i' < m \<longrightarrow> j' < m \<longrightarrow> i' = j' \<longrightarrow> d\<mu> fs'' i' j' = dmu'_mod $$ (i', j')"
- using dd\<mu> dmu'_def 033 10 dmu'_mod_def 11 I_def by simp
- moreover {
- fix i' j'
- assume i'j'': "i' < m" "j' < m" "i' < j'"
- then have \<mu>z: "\<mu> fs'' i' j' = 0" by (simp add: gram_schmidt_fs.\<mu>.simps)
- have "dmu'_mod $$ (i',j') = dmu' $$ (i',j')" using dmu'_mod_def i'j'' by auto
- also have "\<dots> = d\<mu> fs i' j'" using LLL_invD_modw(7)[OF Linvmw] i'j'' dmu'_def by simp
- also have "\<dots> = 0" using d\<mu>_def i'j'' by (simp add: gram_schmidt_fs.\<mu>.simps)
- finally have "d\<mu> fs'' i' j' = dmu'_mod $$ (i',j')" using \<mu>z d_def i'j'' d\<mu>_def by simp
- }
- ultimately show ?thesis by (meson nat_neq_iff)
- qed
- from gbnd 09 have g_bnd: "g_bnd_mode first b fs''" using g_bnd_mode_cong[of fs' fs''] by auto
- {
- assume Linv: "LLL_invariant_mod fs mfs dmu p first b k"
- have 00: "LLL_invariant_weak' k fs" using LLL_invD_mod[OF Linv] LLL_invI_weak by simp
- note swap_weak' = basis_reduction_swap_weak'[OF 00 k k0 mu_F1_i norm_ineq fs'_def]
- have 01: "LLL_invariant_weak' (k - 1) fs'" by fact
- have 12: "weakly_reduced fs'' (k-1)" (* invariant requirement *)
- using 031 09 k LLL_invD_weak(8)[OF 00] unfolding gram_schmidt_fs.weakly_reduced_def by simp
- have "LLL_invariant_mod fs'' mfs' dmu'_mod p first b (k-1)"
- using LLL_invI_mod[OF 13 _ 08 05 07 12 14 16 17 p1 g_bnd LLL_invD_mod(17)[OF Linv]] k by simp
- }
- moreover have "LLL_invariant_mod_weak fs'' mfs' dmu'_mod p first b"
- using LLL_invI_modw[OF 13 08 05 07 14 16 17 p1 g_bnd LLL_invD_modw(15)[OF Linvmw]] by simp
- ultimately show ?thesis using 15 by auto
-qed
-
-lemma dmu_quot_is_round_of_\<mu>:
- assumes Linv: "LLL_invariant_mod fs mfs dmu p first b i'"
- and c: "c = round_num_denom (dmu $$ (i,j)) (d_of dmu (Suc j))"
- and i: "i < m"
- and j: "j < i"
- shows "c = round(\<mu> fs i j)"
-proof -
- have Linvw: "LLL_invariant_weak' i' fs" using LLL_invD_mod[OF Linv] LLL_invI_weak by simp
- have j2: "j < m" using i j by simp
- then have j3: "Suc j \<le> m" by simp
- have \<mu>1: "\<mu> fs j j = 1" using i j by (meson gram_schmidt_fs.\<mu>.elims less_irrefl_nat)
- have inZ: "rat_of_int (d fs (Suc j)) * \<mu> fs i j \<in> \<int>" using fs_int_indpt.fs_int_mu_d_Z_m_m i j
- LLL_invD_mod(5)[OF Linv] LLL_invD_weak(2) Linvw d_def fs_int.d_def fs_int_indpt.intro by auto
- have "c = round(rat_of_int (d\<mu> fs i j) / rat_of_int (d\<mu> fs j j))" using LLL_invD_mod(9) Linv i j c
- by (simp add: round_num_denom d_of_def)
- then show ?thesis using LLL_d_pos[OF LLL_invw'_imp_w[OF Linvw] j3] j i inZ d\<mu>_def \<mu>1 by simp
-qed
-
-lemma dmu_quot_is_round_of_\<mu>_weak:
- assumes Linv: "LLL_invariant_mod_weak fs mfs dmu p first b"
- and c: "c = round_num_denom (dmu $$ (i,j)) (d_of dmu (Suc j))"
- and i: "i < m"
- and j: "j < i"
- shows "c = round(\<mu> fs i j)"
-proof -
- have Linvww: "LLL_invariant_weak fs" using LLL_invD_modw[OF Linv] LLL_invariant_weak_def by simp
- have j2: "j < m" using i j by simp
- then have j3: "Suc j \<le> m" by simp
- have \<mu>1: "\<mu> fs j j = 1" using i j by (meson gram_schmidt_fs.\<mu>.elims less_irrefl_nat)
- have inZ: "rat_of_int (d fs (Suc j)) * \<mu> fs i j \<in> \<int>" using fs_int_indpt.fs_int_mu_d_Z_m_m i j
- LLL_invD_modw[OF Linv] d_def fs_int.d_def fs_int_indpt.intro by auto
- have "c = round(rat_of_int (d\<mu> fs i j) / rat_of_int (d\<mu> fs j j))" using LLL_invD_modw(7) Linv i j c
- by (simp add: round_num_denom d_of_def)
- then show ?thesis using LLL_d_pos[OF Linvww j3] j i inZ d\<mu>_def \<mu>1 by simp
-qed
-
-lemma basis_reduction_mod_add_row: assumes
- Linv: "LLL_invariant_mod_weak fs mfs dmu p first b"
- and res: "basis_reduction_mod_add_row p mfs dmu i j = (mfs', dmu')"
- and i: "i < m"
- and j: "j < i"
- and igtz: "i \<noteq> 0"
-shows "(\<exists>fs'. LLL_invariant_mod_weak fs' mfs' dmu' p first b \<and>
- LLL_measure i fs' = LLL_measure i fs \<and>
- (\<mu>_small_row i fs (Suc j) \<longrightarrow> \<mu>_small_row i fs' j) \<and>
- \<bar>\<mu> fs' i j\<bar> \<le> 1 / 2 \<and>
- (\<forall>i' j'. i' < i \<longrightarrow> j' \<le> i' \<longrightarrow> \<mu> fs' i' j' = \<mu> fs i' j') \<and>
- (LLL_invariant_mod fs mfs dmu p first b i \<longrightarrow> LLL_invariant_mod fs' mfs' dmu' p first b i) \<and>
- (\<forall>ii \<le> m. d fs' ii = d fs ii))"
-proof -
- define c where "c = round_num_denom (dmu $$ (i,j)) (d_of dmu (Suc j))"
- then have c: "c = round(\<mu> fs i j)" using dmu_quot_is_round_of_\<mu>_weak[OF Linv c_def i j] by simp
- show ?thesis
- proof (cases "c = 0")
- case True
- then have pair_id: "(mfs', dmu') = (mfs, dmu)"
- using res c_def unfolding basis_reduction_mod_add_row_def Let_def by auto
- moreover have "\<bar>\<mu> fs i j\<bar> \<le> inverse 2" using c[symmetric, unfolded True]
- by (simp add: round_def, linarith)
- moreover then have "(\<mu>_small_row i fs (Suc j) \<longrightarrow> \<mu>_small_row i fs j)"
- unfolding \<mu>_small_row_def using Suc_leI le_neq_implies_less by blast
- ultimately show ?thesis using Linv pair_id by auto
- next
- case False
- then have pair_id: "(mfs', dmu') = (mfs[i := map_vec (\<lambda>x. x symmod p) (mfs ! i - c \<cdot>\<^sub>v mfs ! j)],
- mat m m (\<lambda>(i', j'). if i' = i \<and> j' \<le> j
- then if j' = j then dmu $$ (i, j') - c * dmu $$ (j, j')
- else (dmu $$ (i,j') - c * dmu $$ (j,j'))
- symmod (p * (d_of dmu j') * (d_of dmu (Suc j')))
- else dmu $$ (i', j')))"
- using res c_def unfolding basis_reduction_mod_add_row_def Let_def by auto
- then have mfs': "mfs' = mfs[i := map_vec (\<lambda>x. x symmod p) (mfs ! i - c \<cdot>\<^sub>v mfs ! j)]"
- and dmu': "dmu' = mat m m (\<lambda>(i', j'). if i' = i \<and> j' \<le> j
- then if j' = j then dmu $$ (i, j') - c * dmu $$ (j, j')
- else (dmu $$ (i,j') - c * dmu $$ (j,j'))
- symmod (p * (d_of dmu j') * (d_of dmu (Suc j')))
- else dmu $$ (i', j'))" by auto
- show ?thesis using basis_reduction_mod_add_row_main[OF Linv i j c mfs' dmu'] by blast
- qed
-qed
-
-lemma basis_reduction_mod_swap: assumes
- Linv: "LLL_invariant_mod_weak fs mfs dmu p first b"
- and mu: "\<bar>\<mu> fs k (k-1)\<bar> \<le> 1 / 2"
- and res: "basis_reduction_mod_swap p mfs dmu k = (mfs', dmu'_mod)"
- and cond: "sq_norm (gso fs (k - 1)) > \<alpha> * sq_norm (gso fs k)"
- and i: "k < m" "k \<noteq> 0"
-shows "(\<exists>fs'. LLL_invariant_mod_weak fs' mfs' dmu'_mod p first b \<and>
- LLL_measure (k - 1) fs' < LLL_measure k fs \<and>
- (LLL_invariant_mod fs mfs dmu p first b k \<longrightarrow> LLL_invariant_mod fs' mfs' dmu'_mod p first b (k-1)))"
- using res[unfolded basis_reduction_mod_swap_def basis_reduction_mod_swap_dmu_mod_def]
- basis_reduction_mod_swap_main[OF Linv i mu cond] by blast
-
-lemma basis_reduction_adjust_mod: assumes
- Linv: "LLL_invariant_mod_weak fs mfs dmu p first b"
- and res: "basis_reduction_adjust_mod p first mfs dmu = (p', mfs', dmu', g_idx')"
-shows "(\<exists>fs' b'. (LLL_invariant_mod fs mfs dmu p first b i \<longrightarrow> LLL_invariant_mod fs' mfs' dmu' p' first b' i) \<and>
- LLL_invariant_mod_weak fs' mfs' dmu' p' first b' \<and>
- LLL_measure i fs' = LLL_measure i fs)"
-proof (cases "\<exists> g_idx. basis_reduction_adjust_mod p first mfs dmu = (p, mfs, dmu, g_idx)")
- case True
- thus ?thesis using res Linv by auto
-next
- case False
- obtain b' g_idx where norm: "compute_max_gso_norm first dmu = (b', g_idx)" by force
- define p'' where "p'' = compute_mod_of_max_gso_norm first b'"
- define d_vec where "d_vec = vec (Suc m) (\<lambda>i. d_of dmu i)"
- define mfs'' where "mfs'' = map (map_vec (\<lambda>x. x symmod p'')) mfs"
- define dmu'' where "dmu'' = mat m m (\<lambda>(i, j).
- if j < i then dmu $$ (i, j) symmod (p'' * d_vec $ j * d_vec $ Suc j)
- else dmu $$ (i, j))"
- note res = res False
- note res = res[unfolded basis_reduction_adjust_mod.simps Let_def norm split,
- folded p''_def, folded d_vec_def mfs''_def, folded dmu''_def]
- from res have pp': "p'' < p" and id: "dmu' = dmu''" "mfs' = mfs''" "p' = p''" "g_idx' = g_idx"
- by (auto split: if_splits)
- define I where "I = {(i',j'). i' < m \<and> j' < i'}"
- note inv = LLL_invD_modw[OF Linv]
- from inv(4) have lin: "gs.lin_indpt_list (RAT fs)" .
- from inv(3) have lat: "lattice_of fs = L" .
- from inv(2) have len: "length fs = m" .
- have weak: "LLL_invariant_weak fs" using Linv
- by (auto simp: LLL_invariant_mod_weak_def LLL_invariant_weak_def)
- from compute_max_gso_norm[OF _ weak, of dmu first, unfolded norm] inv(7)
- have bnd: "g_bnd_mode first b' fs" and b': "b' \<ge> 0" "m = 0 \<Longrightarrow> b' = 0" by auto
- from compute_mod_of_max_gso_norm[OF b' p''_def]
- have p'': "0 < p''" "1 < p''" "mod_invariant b' p'' first"
- by auto
- obtain fs' where
- 01: "lattice_of fs' = L" and
- 02: "map (map_vec (\<lambda> x. x symmod p'')) fs' = map (map_vec (\<lambda> x. x symmod p'')) fs" and
- 03: "lin_indep fs'" and
- 04: "length fs' = m" and
- 05: "(\<forall> k < m. gso fs' k = gso fs k)" and
- 06: "(\<forall> k \<le> m. d fs' k = d fs k)" and
- 07: "(\<forall> i' < m. \<forall> j' < m. d\<mu> fs' i' j' =
- (if (i',j') \<in> I then d\<mu> fs i' j' symmod (p'' * d fs j' * d fs (Suc j')) else d\<mu> fs i' j'))"
- using mod_finite_set[OF lin len _ lat, of I] I_def p'' by blast
- from bnd 05 have bnd: "g_bnd_mode first b' fs'" using g_bnd_mode_cong[of fs fs'] by auto
- have D: "D fs = D fs'" unfolding D_def using 06 by auto
-
-
- have Linv': "LLL_invariant_mod_weak fs' mfs'' dmu'' p'' first b'"
- proof (intro LLL_invI_modw p'' 04 03 01 bnd)
- {
- have "mfs'' = map (map_vec (\<lambda>x. x symmod p'')) mfs" by fact
- also have "\<dots> = map (map_vec (\<lambda>x. x symmod p'')) (map (map_vec (\<lambda>x. x symmod p)) fs)"
- using inv by simp
- also have "\<dots> = map (map_vec (\<lambda>x. x symmod p symmod p'')) fs" by auto
- also have "(\<lambda> x. x symmod p symmod p'') = (\<lambda> x. x symmod p'')"
- proof (intro ext)
- fix x
- from \<open>mod_invariant b p first\<close>[unfolded mod_invariant_def] obtain e where
- p: "p = log_base ^ e" by auto
- from p''[unfolded mod_invariant_def] obtain e' where
- p'': "p'' = log_base ^ e'" by auto
- from pp'[unfolded p p''] log_base have "e' \<le> e" by simp
- hence dvd: "p'' dvd p" unfolding p p'' using log_base by (metis le_imp_power_dvd)
- thus "x symmod p symmod p'' = x symmod p''"
- by (intro sym_mod_sym_mod_cancel)
- qed
- finally show "map (map_vec (\<lambda>x. x symmod p'')) fs' = mfs''" unfolding 02 ..
- }
- thus "length mfs'' = m" using 04 by auto
- show "\<forall>i'<m. \<forall>j'<i'. \<bar>d\<mu> fs' i' j'\<bar> < p'' * d fs' j' * d fs' (Suc j')"
- proof -
- {
- fix i' j'
- assume i'j': "i' < m" "j' < i'"
- then have "d\<mu> fs' i' j' = d\<mu> fs i' j' symmod (p'' * d fs' j' * d fs' (Suc j'))"
- using 07 06 unfolding I_def by simp
- then have "\<bar>d\<mu> fs' i' j'\<bar> < p'' * d fs' j' * d fs' (Suc j')"
- using sym_mod_abs p'' LLL_d_pos[OF weak] mult_pos_pos
- by (smt "06" i'j' less_imp_le_nat less_trans_Suc nat_SN.gt_trans)
- }
- then show ?thesis by simp
- qed
- from inv(7) have dmu: "i' < m \<Longrightarrow> j' < m \<Longrightarrow> dmu $$ (i', j') = d\<mu> fs i' j'" for i' j'
- by auto
- note d_of = d_of_weak[OF Linv]
- have dvec: "i \<le> m \<Longrightarrow> d_vec $ i = d fs i" for i unfolding d_vec_def using d_of by auto
- show "\<forall>i'<m. \<forall>j'<m. d\<mu> fs' i' j' = dmu'' $$ (i', j')"
- using 07 unfolding dmu''_def I_def
- by (auto simp: dmu dvec)
- qed
-
- moreover
- {
- assume linv: "LLL_invariant_mod fs mfs dmu p first b i"
- note inv = LLL_invD_mod[OF linv]
- hence i: "i \<le> m" by auto
- have norm: "j < m \<Longrightarrow> \<parallel>gso fs j\<parallel>\<^sup>2 = \<parallel>gso fs' j\<parallel>\<^sup>2" for j
- using 05 by auto
- have "weakly_reduced fs i = weakly_reduced fs' i"
- unfolding gram_schmidt_fs.weakly_reduced_def using i
- by (intro all_cong arg_cong2[where f = "(\<le>)"] arg_cong[where f = "\<lambda> x. _ * x"] norm, auto)
- with inv have "weakly_reduced fs' i" by auto
- hence "LLL_invariant_mod fs' mfs'' dmu'' p'' first b' i" using inv
- by (intro LLL_invI_mod LLL_invD_modw[OF Linv'])
- }
-
- moreover have "LLL_measure i fs' = LLL_measure i fs"
- unfolding LLL_measure_def logD_def D ..
- ultimately show ?thesis unfolding id by blast
-qed
-
-lemma alpha_comparison: assumes
- Linv: "LLL_invariant_mod_weak fs mfs dmu p first b"
- and alph: "quotient_of \<alpha> = (num, denom)"
- and i: "i < m"
- and i0: "i \<noteq> 0"
-shows "(d_of dmu i * d_of dmu i * denom \<le> num * d_of dmu (i - 1) * d_of dmu (Suc i))
- = (sq_norm (gso fs (i - 1)) \<le> \<alpha> * sq_norm (gso fs i))"
-proof -
- note inv = LLL_invD_modw[OF Linv]
- interpret fs_indep: fs_int_indpt n fs
- by (unfold_locales, insert inv, auto)
- from inv(2) i have ifs: "i < length fs" by auto
- note d_of_fs = d_of_weak[OF Linv]
- show ?thesis
- unfolding fs_indep.d_sq_norm_comparison[OF alph ifs i0, symmetric]
- by (subst (1 2 3 4) d_of_fs, use i d_def fs_indep.d_def in auto)
-qed
-
-lemma basis_reduction_adjust_swap_add_step: assumes
- Linv: "LLL_invariant_mod_weak fs mfs dmu p first b"
- and res: "basis_reduction_adjust_swap_add_step p first mfs dmu g_idx i = (p', mfs', dmu', g_idx')"
- and alph: "quotient_of \<alpha> = (num, denom)"
- and ineq: "\<not> (d_of dmu i * d_of dmu i * denom
- \<le> num * d_of dmu (i - 1) * d_of dmu (Suc i))"
- and i: "i < m"
- and i0: "i \<noteq> 0"
-shows "\<exists>fs' b'. LLL_invariant_mod_weak fs' mfs' dmu' p' first b' \<and>
- LLL_measure (i - 1) fs' < LLL_measure i fs \<and>
- LLL_measure (m - 1) fs' < LLL_measure (m - 1) fs \<and>
- (LLL_invariant_mod fs mfs dmu p first b i \<longrightarrow>
- LLL_invariant_mod fs' mfs' dmu' p' first b' (i - 1))"
-proof -
- obtain mfs0 dmu0 where add: "basis_reduction_mod_add_row p mfs dmu i (i-1) = (mfs0, dmu0)" by force
- obtain mfs1 dmu1 where swap: "basis_reduction_mod_swap p mfs0 dmu0 i = (mfs1, dmu1)" by force
- note res = res[unfolded basis_reduction_adjust_swap_add_step_def Let_def add split swap]
- from i0 have ii: "i - 1 < i" by auto
- from basis_reduction_mod_add_row[OF Linv add i ii i0]
- obtain fs0 where Linv0: "LLL_invariant_mod_weak fs0 mfs0 dmu0 p first b"
- and meas0: "LLL_measure i fs0 = LLL_measure i fs"
- and small: "\<bar>\<mu> fs0 i (i - 1)\<bar> \<le> 1 / 2"
- and Linv0': "LLL_invariant_mod fs mfs dmu p first b i \<Longrightarrow> LLL_invariant_mod fs0 mfs0 dmu0 p first b i"
- by blast
- {
- have id: "d_of dmu0 i = d_of dmu i" "d_of dmu0 (i - 1) = d_of dmu (i - 1)"
- "d_of dmu0 (Suc i) = d_of dmu (Suc i)"
- using i i0 add[unfolded basis_reduction_mod_add_row_def Let_def]
- by (auto split: if_splits simp: d_of_def)
- from ineq[folded id, unfolded alpha_comparison[OF Linv0 alph i i0]]
- have "\<parallel>gso fs0 (i - 1)\<parallel>\<^sup>2 > \<alpha> * \<parallel>gso fs0 i\<parallel>\<^sup>2" by simp
- } note ineq = this
- from Linv have "LLL_invariant_weak fs"
- by (auto simp: LLL_invariant_weak_def LLL_invariant_mod_weak_def)
- from basis_reduction_mod_swap[OF Linv0 small swap ineq i i0, unfolded meas0] Linv0'
- obtain fs1 where Linv1: "LLL_invariant_mod_weak fs1 mfs1 dmu1 p first b"
- and meas1: "LLL_measure (i - 1) fs1 < LLL_measure i fs"
- and Linv1': "LLL_invariant_mod fs mfs dmu p first b i \<Longrightarrow> LLL_invariant_mod fs1 mfs1 dmu1 p first b (i - 1)"
- by auto
- show ?thesis
- proof (cases "i - 1 = g_idx")
- case False
- with res have id: "p' = p" "mfs' = mfs1" "dmu' = dmu1" "g_idx' = g_idx" by auto
- show ?thesis unfolding id using Linv1' meas1 Linv1 by (intro exI[of _ fs1] exI[of _ b], auto simp: LLL_measure_def)
- next
- case True
- with res have adjust: "basis_reduction_adjust_mod p first mfs1 dmu1 = (p', mfs', dmu', g_idx')" by simp
- from basis_reduction_adjust_mod[OF Linv1 adjust, of "i - 1"] Linv1'
- obtain fs' b' where Linvw: "LLL_invariant_mod_weak fs' mfs' dmu' p' first b'"
- and Linv: "LLL_invariant_mod fs mfs dmu p first b i \<Longrightarrow> LLL_invariant_mod fs' mfs' dmu' p' first b' (i - 1)"
- and meas: "LLL_measure (i - 1) fs' = LLL_measure (i - 1) fs1"
- by blast
- note meas = meas1[folded meas]
- from meas have meas': "LLL_measure (m - 1) fs' < LLL_measure (m - 1) fs"
- unfolding LLL_measure_def using i by auto
- show ?thesis
- by (intro exI conjI impI, rule Linvw, rule meas, rule meas', rule Linv)
- qed
-qed
-
-
-lemma basis_reduction_mod_step: assumes
- Linv: "LLL_invariant_mod fs mfs dmu p first b i"
- and res: "basis_reduction_mod_step p first mfs dmu g_idx i j = (p', mfs', dmu', g_idx', i', j')"
- and i: "i < m"
-shows "\<exists>fs' b'. LLL_measure i' fs' < LLL_measure i fs \<and> LLL_invariant_mod fs' mfs' dmu' p' first b' i'"
-proof -
- note res = res[unfolded basis_reduction_mod_step_def Let_def]
- from Linv have Linvw: "LLL_invariant_mod_weak fs mfs dmu p first b"
- by (auto simp: LLL_invariant_mod_weak_def LLL_invariant_mod_def)
- show ?thesis
- proof (cases "i = 0")
- case True
- then have ids: "mfs' = mfs" "dmu' = dmu" "i' = Suc i" "p' = p" using res by auto
- have "LLL_measure i' fs < LLL_measure i fs \<and> LLL_invariant_mod fs mfs' dmu' p first b i'"
- using increase_i_mod[OF Linv i] True res ids inv by simp
- then show ?thesis using res ids inv by auto
- next
- case False
- hence id: "(i = 0) = False" by auto
- obtain num denom where alph: "quotient_of \<alpha> = (num, denom)" by force
- note res = res[unfolded id if_False alph split]
- let ?comp = "d_of dmu i * d_of dmu i * denom \<le> num * d_of dmu (i - 1) * d_of dmu (Suc i)"
- show ?thesis
- proof (cases ?comp)
- case False
- hence id: "?comp = False" by simp
- note res = res[unfolded id if_False]
- let ?step = "basis_reduction_adjust_swap_add_step p first mfs dmu g_idx i"
- from res have step: "?step = (p', mfs', dmu', g_idx')"
- and i': "i' = i - 1"
- by (cases ?step, auto)+
- from basis_reduction_adjust_swap_add_step[OF Linvw step alph False i \<open>i \<noteq> 0\<close>] Linv
- show ?thesis unfolding i' by blast
- next
- case True
- hence id: "?comp = True" by simp
- note res = res[unfolded id if_True]
- from res have ids: "p' = p" "mfs' = mfs" "dmu' = dmu" "i' = Suc i" by auto
- from True alpha_comparison[OF Linvw alph i False]
- have ineq: "sq_norm (gso fs (i - 1)) \<le> \<alpha> * sq_norm (gso fs i)" by simp
- from increase_i_mod[OF Linv i ineq]
- show ?thesis unfolding ids by auto
- qed
- qed
-qed
-
-lemma basis_reduction_mod_main: assumes "LLL_invariant_mod fs mfs dmu p first b i"
- and res: "basis_reduction_mod_main p first mfs dmu g_idx i j = (p', mfs', dmu')"
-shows "\<exists>fs' b'. LLL_invariant_mod fs' mfs' dmu' p' first b' m"
- using assms
-proof (induct "LLL_measure i fs" arbitrary: i mfs dmu j p b fs g_idx rule: less_induct)
- case (less i fs mfs dmu j p b g_idx)
- hence fsinv: "LLL_invariant_mod fs mfs dmu p first b i" by auto
- note res = less(3)[unfolded basis_reduction_mod_main.simps[of p first mfs dmu g_idx i j]]
- note inv = less(2)
- note IH = less(1)
- show ?case
- proof (cases "i < m")
- case i: True
- obtain p' mfs' dmu' g_idx' i' j' where step: "basis_reduction_mod_step p first mfs dmu g_idx i j = (p', mfs', dmu', g_idx', i', j')"
- (is "?step = _") by (cases ?step, auto)
- then obtain fs' b' where Linv: "LLL_invariant_mod fs' mfs' dmu' p' first b' i'"
- and decr: "LLL_measure i' fs' < LLL_measure i fs"
- using basis_reduction_mod_step[OF fsinv step i] i fsinv by blast
- note res = res[unfolded step split]
- from res i show ?thesis using IH[OF decr Linv] by auto
- next
- case False
- with LLL_invD_mod[OF fsinv] res have i: "i = m" "p' = p" by auto
- then obtain fs' b' where "LLL_invariant_mod fs' mfs' dmu' p first b' m" using False res fsinv by simp
- then show ?thesis using i by auto
- qed
-qed
-
-lemma compute_max_gso_quot_alpha:
- assumes inv: "LLL_invariant_mod_weak fs mfs dmu p first b"
- and max: "compute_max_gso_quot dmu = (msq_num, msq_denum, idx)"
- and alph: "quotient_of \<alpha> = (num, denum)"
- and cmp: "(msq_num * denum > num * msq_denum) = cmp"
- and m: "m > 1"
-shows "cmp \<Longrightarrow> idx \<noteq> 0 \<and> idx < m \<and> \<not> (d_of dmu idx * d_of dmu idx * denum
- \<le> num * d_of dmu (idx - 1) * d_of dmu (Suc idx))"
- and "\<not> cmp \<Longrightarrow> LLL_invariant_mod fs mfs dmu p first b m"
-proof -
- from inv
- have fsinv: "LLL_invariant_weak fs"
- by (simp add: LLL_invariant_mod_weak_def LLL_invariant_weak_def)
- define qt where "qt = (\<lambda>i. ((d_of dmu (i + 1)) * (d_of dmu (i + 1)),
- (d_of dmu (i + 2)) * (d_of dmu i), Suc i))"
- define lst where "lst = (map (\<lambda>i. qt i) [0..<(m-1)])"
- have msqlst: "(msq_num, msq_denum, idx) = max_list_rats_with_index lst"
- using max lst_def qt_def unfolding compute_max_gso_quot_def by simp
- have nz: "\<And>n d i. (n, d, i) \<in> set lst \<Longrightarrow> d > 0"
- unfolding lst_def qt_def using d_of_weak[OF inv] LLL_d_pos[OF fsinv] by auto
- have geq: "\<forall>(n, d, i) \<in> set lst. rat_of_int msq_num / of_int msq_denum \<ge> rat_of_int n / of_int d"
- using max_list_rats_with_index[of lst] nz msqlst by (metis (no_types, lifting) case_prodI2)
- have len: "length lst \<ge> 1" using m unfolding lst_def by simp
- have inset: "(msq_num, msq_denum, idx) \<in> set lst"
- using max_list_rats_with_index_in_set[OF msqlst[symmetric] len] nz by simp
- then have idxm: "idx \<in> {1..<m}" using lst_def[unfolded qt_def] by auto
- then have idx0: "idx \<noteq> 0" and idx: "idx < m" by auto
- have 00: "(msq_num, msq_denum, idx) = qt (idx - 1)" using lst_def inset qt_def by auto
- then have id_qt: "msq_num = d_of dmu idx * d_of dmu idx" "msq_denum = d_of dmu (Suc idx) * d_of dmu (idx - 1)"
- unfolding qt_def by auto
- have "msq_denum = (d_of dmu (idx + 1)) * (d_of dmu (idx - 1))"
- using 00 unfolding qt_def by simp
- then have dengt0: "msq_denum > 0" using d_of_weak[OF inv] idxm LLL_d_pos[OF fsinv] by auto
- have \<alpha>dengt0: "denum > 0" using alph by (metis quotient_of_denom_pos)
- from cmp[unfolded id_qt]
- have cmp: "cmp = (\<not> (d_of dmu idx * d_of dmu idx * denum \<le> num * d_of dmu (idx - 1) * d_of dmu (Suc idx)))"
- by (auto simp: ac_simps)
- {
- assume cmp
- from this[unfolded cmp]
- show "idx \<noteq> 0 \<and> idx < m \<and> \<not> (d_of dmu idx * d_of dmu idx * denum
- \<le> num * d_of dmu (idx - 1) * d_of dmu (Suc idx))" using idx0 idx by auto
- }
- {
- assume "\<not> cmp"
- from this[unfolded cmp] have small: "d_of dmu idx * d_of dmu idx * denum \<le> num * d_of dmu (idx - 1) * d_of dmu (Suc idx)" by auto
- note d_pos = LLL_d_pos[OF fsinv]
- have gso: "k < m \<Longrightarrow> sq_norm (gso fs k) = of_int (d fs (Suc k)) / of_int (d fs k)" for k using
- LLL_d_Suc[OF fsinv, of k] d_pos[of k] by simp
- have gso_pos: "k < m \<Longrightarrow> sq_norm (gso fs k) > 0" for k
- using gso[of k] d_pos[of k] d_pos[of "Suc k"] by auto
- from small[unfolded alpha_comparison[OF inv alph idx idx0]]
- have alph: "sq_norm (gso fs (idx - 1)) \<le> \<alpha> * sq_norm (gso fs idx)" .
- with gso_pos[OF idx] have alph: "sq_norm (gso fs (idx - 1)) / sq_norm (gso fs idx) \<le> \<alpha>"
- by (metis mult_imp_div_pos_le)
- have weak: "weakly_reduced fs m" unfolding gram_schmidt_fs.weakly_reduced_def
- proof (intro allI impI, goal_cases)
- case (1 i)
- from idx have idx1: "idx - 1 < m" by auto
- from geq[unfolded lst_def]
- have mem: "(d_of dmu (Suc i) * d_of dmu (Suc i),
- d_of dmu (Suc (Suc i)) * d_of dmu i, Suc i) \<in> set lst"
- unfolding lst_def qt_def using 1 by auto
- have "sq_norm (gso fs i) / sq_norm (gso fs (Suc i)) =
- of_int (d_of dmu (Suc i) * d_of dmu (Suc i)) / of_int (d_of dmu (Suc (Suc i)) * d_of dmu i)"
- using gso idx0 d_of_weak[OF inv] 1 by auto
- also have "\<dots> \<le> rat_of_int msq_num / rat_of_int msq_denum"
- using geq[rule_format, OF mem, unfolded split] by auto
- also have "\<dots> = sq_norm (gso fs (idx - 1)) / sq_norm (gso fs idx)"
- unfolding id_qt gso[OF idx] gso[OF idx1] using idx0 d_of_weak[OF inv] idx by auto
- also have "\<dots> \<le> \<alpha>" by fact
- finally show "sq_norm (gso fs i) \<le> \<alpha> * sq_norm (gso fs (Suc i))" using gso_pos[OF 1]
- using pos_divide_le_eq by blast
- qed
- with inv show "LLL_invariant_mod fs mfs dmu p first b m"
- by (auto simp: LLL_invariant_mod_weak_def LLL_invariant_mod_def)
- }
-qed
-
-
-lemma small_m:
- assumes inv: "LLL_invariant_mod_weak fs mfs dmu p first b"
- and m: "m \<le> 1"
-shows "LLL_invariant_mod fs mfs dmu p first b m"
-proof -
- have weak: "weakly_reduced fs m" unfolding gram_schmidt_fs.weakly_reduced_def using m
- by auto
- with inv show "LLL_invariant_mod fs mfs dmu p first b m"
- by (auto simp: LLL_invariant_mod_weak_def LLL_invariant_mod_def)
-qed
-
-lemma basis_reduction_iso_main: assumes "LLL_invariant_mod_weak fs mfs dmu p first b"
- and res: "basis_reduction_iso_main p first mfs dmu g_idx j = (p', mfs', dmu')"
-shows "\<exists>fs' b'. LLL_invariant_mod fs' mfs' dmu' p' first b' m"
- using assms
-proof (induct "LLL_measure (m-1) fs" arbitrary: fs mfs dmu j p b g_idx rule: less_induct)
- case (less fs mfs dmu j p b g_idx)
- have inv: "LLL_invariant_mod_weak fs mfs dmu p first b" using less by auto
- hence fsinv: "LLL_invariant_weak fs"
- by (simp add: LLL_invariant_mod_weak_def LLL_invariant_weak_def)
- note res = less(3)[unfolded basis_reduction_iso_main.simps[of p first mfs dmu g_idx j]]
- note IH = less(1)
- obtain msq_num msq_denum idx where max: "compute_max_gso_quot dmu = (msq_num, msq_denum, idx)"
- by (metis prod_cases3)
- obtain num denum where alph: "quotient_of \<alpha> = (num, denum)" by force
- note res = res[unfolded max alph Let_def split]
- consider (small) "m \<le> 1" | (final) "m > 1" "\<not> (num * msq_denum < msq_num * denum)" | (step) "m > 1" "num * msq_denum < msq_num * denum"
- by linarith
- thus ?case
- proof cases
- case *: step
- obtain p1 mfs1 dmu1 g_idx1 where step: "basis_reduction_adjust_swap_add_step p first mfs dmu g_idx idx = (p1, mfs1, dmu1, g_idx1)"
- by (metis prod_cases4)
- from res[unfolded step split] * have res: "basis_reduction_iso_main p1 first mfs1 dmu1 g_idx1 (j + 1) = (p', mfs', dmu')" by auto
- from compute_max_gso_quot_alpha(1)[OF inv max alph refl *]
- have idx0: "idx \<noteq> 0" and idx: "idx < m" and cmp: "\<not> d_of dmu idx * d_of dmu idx * denum \<le> num * d_of dmu (idx - 1) * d_of dmu (Suc idx)" by auto
- from basis_reduction_adjust_swap_add_step[OF inv step alph cmp idx idx0] obtain fs1 b1
- where inv1: "LLL_invariant_mod_weak fs1 mfs1 dmu1 p1 first b1" and meas: "LLL_measure (m - 1) fs1 < LLL_measure (m - 1) fs"
- by auto
- from IH[OF meas inv1 res] show ?thesis .
- next
- case small
- with res small_m[OF inv] show ?thesis by auto
- next
- case final
- from compute_max_gso_quot_alpha(2)[OF inv max alph refl final]
- final show ?thesis using res by auto
- qed
-qed
-
-lemma basis_reduction_mod_add_rows_loop_inv': assumes
- fsinv: "LLL_invariant_mod fs mfs dmu p first b m"
- and res: "basis_reduction_mod_add_rows_loop p mfs dmu i i = (mfs', dmu')"
- and i: "i < m"
-shows "\<exists>fs'. LLL_invariant_mod fs' mfs' dmu' p first b m \<and>
- (\<forall>i' j'. i' < i \<longrightarrow> j' \<le> i' \<longrightarrow> \<mu> fs i' j' = \<mu> fs' i' j') \<and>
- \<mu>_small fs' i"
-proof -
- {
- fix j
- assume j: "j \<le> i" and mu_small: "\<mu>_small_row i fs j"
- and resj: "basis_reduction_mod_add_rows_loop p mfs dmu i j = (mfs', dmu')"
- have "\<exists>fs'. LLL_invariant_mod fs' mfs' dmu' p first b m \<and>
- (\<forall>i' j'. i' < i \<longrightarrow> j' \<le> i' \<longrightarrow> \<mu> fs i' j' = \<mu> fs' i' j') \<and>
- (\<mu>_small fs' i)"
- proof (insert fsinv mu_small resj i j, induct j arbitrary: fs mfs dmu mfs' dmu')
- case (0 fs)
- then have "(mfs', dmu') = (mfs, dmu)" by simp
- then show ?case
- using LLL_invariant_mod_to_weak_m_to_i(3) basis_reduction_add_row_done_weak 0 by auto
- next
- case (Suc j)
- hence j: "j < i" by auto
- have in0: "i \<noteq> 0" using Suc(6) by simp
- define c where "c = round_num_denom (dmu $$ (i,j)) (d_of dmu (Suc j))"
- have c2: "c = round (\<mu> fs i j)" using dmu_quot_is_round_of_\<mu>[OF _ _ i j] c_def Suc by simp
- define mfs'' where "mfs'' = (if c=0 then mfs else mfs[ i := (map_vec (\<lambda> x. x symmod p)) (mfs ! i - c \<cdot>\<^sub>v mfs ! j)])"
- define dmu'' where "dmu'' = (if c=0 then dmu else mat m m (\<lambda>(i',j'). (if (i' = i \<and> j' \<le> j)
- then (if j'=j then (dmu $$ (i,j') - c * dmu $$ (j,j'))
- else (dmu $$ (i,j') - c * dmu $$ (j,j')) symmod (p * (d_of dmu j') * (d_of dmu (Suc j'))))
- else (dmu $$ (i',j')))))"
- have 00: "basis_reduction_mod_add_row p mfs dmu i j = (mfs'', dmu'')"
- using mfs''_def dmu''_def unfolding basis_reduction_mod_add_row_def c_def[symmetric] by simp
- then have 01: "basis_reduction_mod_add_rows_loop p mfs'' dmu'' i j = (mfs', dmu')"
- using basis_reduction_mod_add_rows_loop.simps(2)[of p mfs dmu i j] Suc by simp
- have fsinvi: "LLL_invariant_mod fs mfs dmu p first b i" using LLL_invariant_mod_to_weak_m_to_i[OF Suc(2)] i by simp
- then have fsinvmw: "LLL_invariant_mod_weak fs mfs dmu p first b" using LLL_invD_mod LLL_invI_modw by simp
- obtain fs'' where fs''invi: "LLL_invariant_mod fs'' mfs'' dmu'' p first b i" and
- \<mu>_small': "(\<mu>_small_row i fs (Suc j) \<longrightarrow> \<mu>_small_row i fs'' j)" and
- \<mu>s: "(\<forall>i' j'. i' < i \<longrightarrow> j' \<le> i' \<longrightarrow> \<mu> fs'' i' j' = \<mu> fs i' j')"
- using Suc basis_reduction_mod_add_row[OF fsinvmw 00 i j] fsinvi by auto
- moreover then have \<mu>sm: "\<mu>_small_row i fs'' j" using Suc by simp
- have fs''invwi: "LLL_invariant_weak' i fs''" using LLL_invD_mod[OF fs''invi] LLL_invI_weak by simp
- have fsinvwi: "LLL_invariant_weak' i fs" using LLL_invD_mod[OF fsinvi] LLL_invI_weak by simp
- note invw = LLL_invw'_imp_w[OF fsinvwi]
- note invw'' = LLL_invw'_imp_w[OF fs''invwi]
- have "LLL_invariant_mod fs'' mfs'' dmu'' p first b m"
- proof -
- have "(\<forall> l. Suc l < m \<longrightarrow> sq_norm (gso fs'' l) \<le> \<alpha> * sq_norm (gso fs'' (Suc l)))"
- proof -
- {
- fix l
- assume l: "Suc l < m"
- have "sq_norm (gso fs'' l) \<le> \<alpha> * sq_norm (gso fs'' (Suc l))"
- proof (cases "i \<le> Suc l")
- case True
- have deq: "\<And>k. k < m \<Longrightarrow> d fs (Suc k) = d fs'' (Suc k)"
- using dd\<mu> LLL_invD_mod(9)[OF fs''invi] LLL_invD_mod(9)[OF Suc(2)] dmu''_def j by simp
- {
- fix k
- assume k: "k < m"
- then have "d fs (Suc k) = d fs'' (Suc k)"
- using dd\<mu> LLL_invD_mod(9)[OF fs''invi] LLL_invD_mod(9)[OF Suc(2)] dmu''_def j by simp
- have "d fs 0 = 1" "d fs'' 0 = 1" using d_def by auto
- moreover have sqid: "sq_norm (gso fs'' k) = rat_of_int (d fs'' (Suc k)) / rat_of_int (d fs'' k)"
- using LLL_d_Suc[OF invw''] LLL_d_pos[OF invw''] k
- by (smt One_nat_def Suc_less_eq Suc_pred le_imp_less_Suc mult_eq_0_iff less_imp_le_nat
- nonzero_mult_div_cancel_right of_int_0_less_iff of_int_hom.hom_zero)
- moreover have "sq_norm (gso fs k) = rat_of_int (d fs (Suc k)) / rat_of_int (d fs k)"
- using LLL_d_Suc[OF invw] LLL_d_pos[OF invw] k
- by (smt One_nat_def Suc_less_eq Suc_pred le_imp_less_Suc mult_eq_0_iff less_imp_le_nat
- nonzero_mult_div_cancel_right of_int_0_less_iff of_int_hom.hom_zero)
- ultimately have "sq_norm (gso fs k) = sq_norm (gso fs'' k)" using k deq
- LLL_d_pos[OF invw] LLL_d_pos[OF invw'']
- by (metis (no_types, lifting) Nat.lessE Suc_lessD old.nat.inject zero_less_Suc)
- }
- then show ?thesis using LLL_invD_mod(6)[OF Suc(2)] by (simp add: gram_schmidt_fs.weakly_reduced_def l)
- next
- case False
- then show ?thesis using LLL_invD_mod(6)[OF fs''invi] gram_schmidt_fs.weakly_reduced_def
- by (metis less_or_eq_imp_le nat_neq_iff)
- qed
- }
- then show ?thesis by simp
- qed
- then have "weakly_reduced fs'' m" using gram_schmidt_fs.weakly_reduced_def by blast
- then show ?thesis using LLL_invD_mod[OF fs''invi] LLL_invI_mod by simp
- qed
- then show ?case using "01" Suc.hyps i j less_imp_le_nat \<mu>sm \<mu>s by metis
- qed
- }
- then show ?thesis using \<mu>_small_row_refl res by auto
-qed
-
-lemma basis_reduction_mod_add_rows_outer_loop_inv:
- assumes inv: "LLL_invariant_mod fs mfs dmu p first b m"
- and "(mfs', dmu') = basis_reduction_mod_add_rows_outer_loop p mfs dmu i"
- and i: "i < m"
-shows "(\<exists>fs'. LLL_invariant_mod fs' mfs' dmu' p first b m \<and>
- (\<forall>j. j \<le> i \<longrightarrow> \<mu>_small fs' j))"
-proof(insert assms, induct i arbitrary: fs mfs dmu mfs' dmu')
- case (0 fs)
- then show ?case using \<mu>_small_def by auto
-next
- case (Suc i fs mfs dmu mfs' dmu')
- obtain mfs'' dmu'' where mfs''dmu'': "(mfs'', dmu'')
- = basis_reduction_mod_add_rows_outer_loop p mfs dmu i" by (metis surj_pair)
- then obtain fs'' where fs'': "LLL_invariant_mod fs'' mfs'' dmu'' p first b m"
- and 00: "(\<forall>j. j \<le> i \<longrightarrow> \<mu>_small fs'' j)" using Suc by fastforce
- have "(mfs', dmu') = basis_reduction_mod_add_rows_loop p mfs'' dmu'' (Suc i) (Suc i)"
- using Suc(3,4) mfs''dmu'' by (smt basis_reduction_mod_add_rows_outer_loop.simps(2) case_prod_conv)
- then obtain fs' where 01: "LLL_invariant_mod fs' mfs' dmu' p first b m"
- and 02: "\<forall>i' j'. i' < (Suc i) \<longrightarrow> j' \<le> i' \<longrightarrow> \<mu> fs'' i' j' = \<mu> fs' i' j'" and 03: "\<mu>_small fs' (Suc i)"
- using fs'' basis_reduction_mod_add_rows_loop_inv' Suc by metis
- moreover have "\<forall>j. j \<le> (Suc i) \<longrightarrow> \<mu>_small fs' j" using 02 00 03 \<mu>_small_def by (simp add: le_Suc_eq)
- ultimately show ?case by blast
-qed
-
-lemma basis_reduction_mod_fs_bound:
- assumes Linv: "LLL_invariant_mod fs mfs dmu p first b k"
- and mu_small: "\<mu>_small fs i"
- and i: "i < m"
- and nFirst: "\<not> first"
-shows "fs ! i = mfs ! i"
-proof -
- from LLL_invD_mod(16-17)[OF Linv] nFirst g_bnd_mode_def
- have gbnd: "g_bnd b fs" and bp: "b \<le> (rat_of_int (p - 1))\<^sup>2 / (rat_of_nat m + 3)"
- by (auto simp: mod_invariant_def bound_number_def)
- have Linvw: "LLL_invariant_weak' k fs" using LLL_invD_mod[OF Linv] LLL_invI_weak by simp
- have "fs_int_indpt n fs" using LLL_invD_mod(5)[OF Linv] Gram_Schmidt_2.fs_int_indpt.intro by simp
- then interpret fs: fs_int_indpt n fs
- using fs_int_indpt.sq_norm_fs_via_sum_mu_gso by simp
- have "\<parallel>gso fs 0\<parallel>\<^sup>2 \<le> b" using gbnd i unfolding g_bnd_def by blast
- then have b0: "0 \<le> b" using sq_norm_vec_ge_0 dual_order.trans by auto
- have 00: "of_int \<parallel>fs ! i\<parallel>\<^sup>2 = (\<Sum>j\<leftarrow>[0..<Suc i]. (\<mu> fs i j)\<^sup>2 * \<parallel>gso fs j\<parallel>\<^sup>2)"
- using fs.sq_norm_fs_via_sum_mu_gso LLL_invD_mod[OF Linv] Gram_Schmidt_2.fs_int_indpt.intro i by simp
- have 01: "\<forall>j < i. (\<mu> fs i j)\<^sup>2 * \<parallel>gso fs j\<parallel>\<^sup>2 \<le> (1 / rat_of_int 4) * \<parallel>gso fs j\<parallel>\<^sup>2"
- proof -
- {
- fix j
- assume j: "j < i"
- then have "\<bar>fs.gs.\<mu> i j\<bar> \<le> 1 / (rat_of_int 2)"
- using mu_small Power.linordered_idom_class.abs_square_le_1 j unfolding \<mu>_small_def by simp
- moreover have "\<bar>\<mu> fs i j\<bar> \<ge> 0" by simp
- ultimately have "\<bar>\<mu> fs i j\<bar>\<^sup>2 \<le> (1 / rat_of_int 2)\<^sup>2"
- using Power.linordered_idom_class.abs_le_square_iff by fastforce
- also have "\<dots> = 1 / (rat_of_int 4)" by (simp add: field_simps)
- finally have "\<bar>\<mu> fs i j\<bar>\<^sup>2 \<le> 1 / rat_of_int 4" by simp
- }
- then show ?thesis using fs.gs.\<mu>.simps by (metis mult_right_mono power2_abs sq_norm_vec_ge_0)
- qed
- then have 0111: "\<And>j. j \<in> set [0..<i] \<Longrightarrow> (\<mu> fs i j)\<^sup>2 * \<parallel>gso fs j\<parallel>\<^sup>2 \<le> (1 / rat_of_int 4) * \<parallel>gso fs j\<parallel>\<^sup>2"
- by simp
- {
- fix j
- assume j: "j < n"
- have 011: "(\<mu> fs i i)\<^sup>2 * \<parallel>gso fs i\<parallel>\<^sup>2 = 1 * \<parallel>gso fs i\<parallel>\<^sup>2"
- using fs.gs.\<mu>.simps by simp
- have 02: "\<forall>j < Suc i. \<parallel>gso fs j\<parallel>\<^sup>2 \<le> b"
- using gbnd i unfolding g_bnd_def by simp
- have 03: "length [0..<Suc i] = (Suc i)" by simp
- have "of_int \<parallel>fs ! i\<parallel>\<^sup>2 = (\<Sum>j\<leftarrow>[0..<i]. (\<mu> fs i j)\<^sup>2 * \<parallel>gso fs j\<parallel>\<^sup>2) + \<parallel>gso fs i\<parallel>\<^sup>2"
- unfolding 00 using 011 by simp
- also have "(\<Sum>j\<leftarrow>[0..<i]. (\<mu> fs i j)\<^sup>2 * \<parallel>gso fs j\<parallel>\<^sup>2) \<le> (\<Sum>j\<leftarrow>[0..<i]. ((1 / rat_of_int 4) * \<parallel>gso fs j\<parallel>\<^sup>2))"
- using Groups_List.sum_list_mono[OF 0111] by fast
- finally have "of_int \<parallel>fs ! i\<parallel>\<^sup>2 \<le> (\<Sum>j\<leftarrow>[0..<i]. ((1 / rat_of_int 4) * \<parallel>gso fs j\<parallel>\<^sup>2)) + \<parallel>gso fs i\<parallel>\<^sup>2"
- by simp
- also have "(\<Sum>j\<leftarrow>[0..<i]. ((1 / rat_of_int 4) * \<parallel>gso fs j\<parallel>\<^sup>2)) \<le> (\<Sum>j\<leftarrow>[0..<i]. (1 / rat_of_int 4) * b)"
- by (intro sum_list_mono, insert 02, auto)
- also have " \<parallel>gso fs i\<parallel>\<^sup>2 \<le> b" using 02 by simp
- finally have "of_int \<parallel>fs ! i\<parallel>\<^sup>2 \<le> (\<Sum>j\<leftarrow>[0..<i]. (1 / rat_of_int 4) * b) + b" by simp
- also have "\<dots> = (rat_of_nat i) * ((1 / rat_of_int 4) * b) + b"
- using 03 sum_list_triv[of "(1 / rat_of_int 4) * b" "[0..<i]"] by simp
- also have "\<dots> = (rat_of_nat i) / 4 * b + b" by simp
- also have "\<dots> = ((rat_of_nat i) / 4 + 1)* b" by algebra
- also have "\<dots> = (rat_of_nat i + 4) / 4 * b" by simp
- finally have "of_int \<parallel>fs ! i\<parallel>\<^sup>2 \<le> (rat_of_nat i + 4) / 4 * b" by simp
- also have "\<dots> \<le> (rat_of_nat (m + 3)) / 4 * b" using i b0 times_left_mono by fastforce
- finally have "of_int \<parallel>fs ! i\<parallel>\<^sup>2 \<le> rat_of_nat (m+3) / 4 * b" by simp
- moreover have "\<bar>fs ! i $ j\<bar>\<^sup>2 \<le> \<parallel>fs ! i\<parallel>\<^sup>2" using vec_le_sq_norm LLL_invD_mod(10)[OF Linv] i j by blast
- ultimately have 04: "of_int (\<bar>fs ! i $ j\<bar>\<^sup>2) \<le> rat_of_nat (m+3) / 4 * b" using ge_trans i by linarith
- then have 05: "real_of_int (\<bar>fs ! i $ j\<bar>\<^sup>2) \<le> real_of_rat (rat_of_nat (m+3) / 4 * b)"
- proof -
- from j have "rat_of_int (\<bar>fs ! i $ j\<bar>\<^sup>2) \<le> rat_of_nat (m+3) / 4 * b" using 04 by simp
- then have "real_of_int (\<bar>fs ! i $ j\<bar>\<^sup>2) \<le> real_of_rat (rat_of_nat (m+3) / 4 * b)"
- using j of_rat_less_eq by (metis of_rat_of_int_eq)
- then show ?thesis by simp
- qed
- define rhs where "rhs = real_of_rat (rat_of_nat (m+3) / 4 * b)"
- have rhs0: "rhs \<ge> 0" using b0 i rhs_def by simp
- have fsij: "real_of_int \<bar>fs ! i $ j\<bar> \<ge> 0" by simp
- have "real_of_int (\<bar>fs ! i $ j\<bar>\<^sup>2) = (real_of_int \<bar>fs ! i $ j\<bar>)\<^sup>2" by simp
- then have "(real_of_int \<bar>fs ! i $ j\<bar>)\<^sup>2 \<le> rhs" using 05 j rhs_def by simp
- then have g1: "real_of_int \<bar>fs ! i $ j\<bar> \<le> sqrt rhs" using NthRoot.real_le_rsqrt by simp
- have pbnd: "2 * \<bar>fs ! i $ j\<bar> < p"
- proof -
- have "rat_of_nat (m+3) / 4 * b \<le> (rat_of_nat (m +3) / 4) * (rat_of_int (p - 1))\<^sup>2 / (rat_of_nat m+3)"
- using bp b0 i times_left_mono SN_Orders.of_nat_ge_zero gs.m_comm times_divide_eq_right
- by (smt gs.l_null le_divide_eq_numeral1(1))
- also have "\<dots> = (rat_of_int (p - 1))\<^sup>2 / 4 * (rat_of_nat (m + 3) / rat_of_nat (m + 3))"
- by (metis (no_types, lifting) gs.m_comm of_nat_add of_nat_numeral times_divide_eq_left)
- finally have "rat_of_nat (m+3) / 4 * b \<le> (rat_of_int (p - 1))\<^sup>2 / 4" by simp
- then have "sqrt rhs \<le> sqrt (real_of_rat ((rat_of_int (p - 1))\<^sup>2 / 4))"
- unfolding rhs_def using of_rat_less_eq by fastforce
- then have two_ineq:
- "2 * \<bar>fs ! i $ j\<bar> \<le> 2 * sqrt (real_of_rat ((rat_of_int (p - 1))\<^sup>2 / 4))"
- using g1 by linarith
- have "2 * sqrt (real_of_rat ((rat_of_int (p - 1))\<^sup>2 / 4)) =
- sqrt (real_of_rat (4 * ((rat_of_int (p - 1))\<^sup>2 / 4)))"
- by (metis (no_types, opaque_lifting) real_sqrt_mult of_int_numeral of_rat_hom.hom_mult
- of_rat_of_int_eq real_sqrt_four times_divide_eq_right)
- also have "\<dots> = sqrt (real_of_rat ((rat_of_int (p - 1))\<^sup>2))" using i by simp
- also have "(real_of_rat ((rat_of_int (p - 1))\<^sup>2)) = (real_of_rat (rat_of_int (p - 1)))\<^sup>2"
- using Rat.of_rat_power by blast
- also have "sqrt ((real_of_rat (rat_of_int (p - 1)))\<^sup>2) = real_of_rat (rat_of_int (p - 1))"
- using LLL_invD_mod(15)[OF Linv] by simp
- finally have "2 * sqrt (real_of_rat ((rat_of_int (p - 1))\<^sup>2 / 4)) =
- real_of_rat (rat_of_int (p - 1))" by simp
- then have "2 * \<bar>fs ! i $ j\<bar> \<le> real_of_rat (rat_of_int (p - 1))"
- using two_ineq by simp
- then show ?thesis by (metis of_int_le_iff of_rat_of_int_eq zle_diff1_eq)
- qed
- have p1: "p > 1" using LLL_invD_mod[OF Linv] by blast
- interpret pm: poly_mod_2 p
- by (unfold_locales, rule p1)
- from LLL_invD_mod[OF Linv] have len: "length fs = m" and fs: "set fs \<subseteq> carrier_vec n" by auto
- from pm.inv_M_rev[OF pbnd, unfolded pm.M_def] have "pm.inv_M (fs ! i $ j mod p) = fs ! i $ j" .
- also have "pm.inv_M (fs ! i $ j mod p) = mfs ! i $ j" unfolding LLL_invD_mod(7)[OF Linv, symmetric] sym_mod_def
- using i j len fs by auto
- finally have "fs ! i $ j = mfs ! i $ j" ..
- }
- thus "fs ! i = mfs ! i" using LLL_invD_mod(10,13)[OF Linv i] by auto
-qed
-
-lemma basis_reduction_mod_fs_bound_first:
- assumes Linv: "LLL_invariant_mod fs mfs dmu p first b k"
- and m0: "m > 0"
- and first: "first"
-shows "fs ! 0 = mfs ! 0"
-proof -
- from LLL_invD_mod(16-17)[OF Linv] first g_bnd_mode_def m0
- have gbnd: "sq_norm (gso fs 0) \<le> b" and bp: "b \<le> (rat_of_int (p - 1))\<^sup>2 / 4"
- by (auto simp: mod_invariant_def bound_number_def)
- from LLL_invD_mod[OF Linv] have p1: "p > 1" by blast
- have Linvw: "LLL_invariant_weak' k fs" using LLL_invD_mod[OF Linv] LLL_invI_weak by simp
- have "fs_int_indpt n fs" using LLL_invD_mod(5)[OF Linv] Gram_Schmidt_2.fs_int_indpt.intro by simp
- then interpret fs: fs_int_indpt n fs
- using fs_int_indpt.sq_norm_fs_via_sum_mu_gso by simp
- from gbnd have b0: "0 \<le> b" using sq_norm_vec_ge_0 dual_order.trans by auto
- have "of_int \<parallel>fs ! 0\<parallel>\<^sup>2 = (\<mu> fs 0 0)\<^sup>2 * \<parallel>gso fs 0\<parallel>\<^sup>2"
- using fs.sq_norm_fs_via_sum_mu_gso LLL_invD_mod[OF Linv] Gram_Schmidt_2.fs_int_indpt.intro m0 by simp
- also have "\<dots> = \<parallel>gso fs 0\<parallel>\<^sup>2" unfolding fs.gs.\<mu>.simps by (simp add: gs.\<mu>.simps)
- also have "\<dots> \<le> (rat_of_int (p - 1))\<^sup>2 / 4" using gbnd bp by auto
- finally have one: "of_int (sq_norm (fs ! 0)) \<le> (rat_of_int (p - 1))\<^sup>2 / 4" .
- {
- fix j
- assume j: "j < n"
- have leq: "\<bar>fs ! 0 $ j\<bar>\<^sup>2 \<le> \<parallel>fs ! 0\<parallel>\<^sup>2" using vec_le_sq_norm LLL_invD_mod(10)[OF Linv] m0 j by blast
- have "rat_of_int ((2 * \<bar>fs ! 0 $ j\<bar>)^2) = rat_of_int (4 * \<bar>fs ! 0 $ j\<bar>\<^sup>2)" by simp
- also have "\<dots> \<le> 4 * of_int \<parallel>fs ! 0\<parallel>\<^sup>2" using leq by simp
- also have "\<dots> \<le> 4 * (rat_of_int (p - 1))\<^sup>2 / 4" using one by simp
- also have "\<dots> = (rat_of_int (p - 1))\<^sup>2" by simp
- also have "\<dots> = rat_of_int ((p - 1)\<^sup>2)" by simp
- finally have "(2 * \<bar>fs ! 0 $ j\<bar>)^2 \<le> (p - 1)\<^sup>2" by linarith
- hence "2 * \<bar>fs ! 0 $ j\<bar> \<le> p - 1" using p1
- by (smt power_mono_iff zero_less_numeral)
- hence pbnd: "2 * \<bar>fs ! 0 $ j\<bar> < p" by simp
- interpret pm: poly_mod_2 p
- by (unfold_locales, rule p1)
- from LLL_invD_mod[OF Linv] m0 have len: "length fs = m" "length mfs = m"
- and fs: "fs ! 0 \<in> carrier_vec n" "mfs ! 0 \<in> carrier_vec n" by auto
- from pm.inv_M_rev[OF pbnd, unfolded pm.M_def] have "pm.inv_M (fs ! 0 $ j mod p) = fs ! 0 $ j" .
- also have "pm.inv_M (fs ! 0 $ j mod p) = mfs ! 0 $ j" unfolding LLL_invD_mod(7)[OF Linv, symmetric] sym_mod_def
- using m0 j len fs by auto
- finally have "mfs ! 0 $ j = fs ! 0 $ j" .
- }
- thus "fs ! 0 = mfs ! 0" using LLL_invD_mod(10,13)[OF Linv m0] by auto
-qed
-
-lemma dmu_initial: "dmu_initial = mat m m (\<lambda> (i,j). d\<mu> fs_init i j)"
-proof -
- interpret fs: fs_int_indpt n fs_init
- by (unfold_locales, intro lin_dep)
- show ?thesis unfolding dmu_initial_def Let_def
- proof (intro cong_mat refl refl, unfold split, goal_cases)
- case (1 i j)
- show ?case
- proof (cases "j \<le> i")
- case False
- thus ?thesis by (auto simp: d\<mu>_def gs.\<mu>.simps)
- next
- case True
- hence id: "d\<mu>_impl fs_init !! i !! j = fs.d\<mu> i j" unfolding fs.d\<mu>_impl
- by (subst of_fun_nth, use 1 len in force, subst of_fun_nth, insert True, auto)
- also have "\<dots> = d\<mu> fs_init i j" unfolding fs.d\<mu>_def d\<mu>_def fs.d_def d_def by simp
- finally show ?thesis using True by auto
- qed
- qed
-qed
-
-lemma LLL_initial_invariant_mod: assumes res: "compute_initial_state first = (p, mfs, dmu', g_idx)"
-shows "\<exists>fs b. LLL_invariant_mod fs mfs dmu' p first b 0"
-proof -
- from dmu_initial have dmu: "(\<forall>i' < m. \<forall>j' < m. d\<mu> fs_init i' j' = dmu_initial $$ (i',j'))" by auto
- obtain b g_idx where norm: "compute_max_gso_norm first dmu_initial = (b,g_idx)" by force
- note res = res[unfolded compute_initial_state_def Let_def norm split]
- from res have p: "p = compute_mod_of_max_gso_norm first b" by auto
- then have p0: "p > 0" unfolding compute_mod_of_max_gso_norm_def using log_base by simp
- then have p1: "p \<ge> 1" by simp
- note res = res[folded p]
- from res[unfolded compute_initial_mfs_def]
- have mfs: "mfs = map (map_vec (\<lambda>x. x symmod p)) fs_init" by auto
- from res[unfolded compute_initial_dmu_def]
- have dmu': "dmu' = mat m m (\<lambda>(i',j'). if j' < i'
- then dmu_initial $$ (i', j') symmod (p * d_of dmu_initial j' * d_of dmu_initial (Suc j'))
- else dmu_initial $$ (i',j'))" by auto
- have lat: "lattice_of fs_init = L" by (auto simp: L_def)
- define I where "I = {(i',j'). i' < m \<and> j' < i'}"
- obtain fs where
- 01: "lattice_of fs = L" and
- 02: "map (map_vec (\<lambda> x. x symmod p)) fs = map (map_vec (\<lambda> x. x symmod p)) fs_init" and
- 03: "lin_indep fs" and
- 04: "length fs = m" and
- 05: "(\<forall> k < m. gso fs k = gso fs_init k)" and
- 06: "(\<forall> k \<le> m. d fs k = d fs_init k)" and
- 07: "(\<forall> i' < m. \<forall> j' < m. d\<mu> fs i' j' =
- (if (i',j') \<in> I then d\<mu> fs_init i' j' symmod (p * d fs_init j' * d fs_init (Suc j')) else d\<mu> fs_init i' j'))"
- using mod_finite_set[OF lin_dep len _ lat p0, of I] I_def by blast
- have inv: "LLL_invariant_weak fs_init"
- by (intro LLL_inv_wI lat len lin_dep fs_init)
- have "\<forall>i'<m. d\<mu> fs_init i' i' = dmu_initial $$ (i', i')" unfolding dmu_initial by auto
- from compute_max_gso_norm[OF this inv, of first, unfolded norm] have gbnd: "g_bnd_mode first b fs_init"
- and b0: "0 \<le> b" and mb0: "m = 0 \<Longrightarrow> b = 0" by auto
- from gbnd 05 have gbnd: "g_bnd_mode first b fs" using g_bnd_mode_cong[of fs fs_init] by auto
- have d\<mu>dmu': "\<forall>i'<m. \<forall>j'<m. d\<mu> fs i' j' = dmu' $$ (i', j')" using 07 dmu d_of_main[of fs_init dmu_initial]
- unfolding I_def dmu' by simp
- have wred: "weakly_reduced fs 0" by (simp add: gram_schmidt_fs.weakly_reduced_def)
- have fs_carr: "set fs \<subseteq> carrier_vec n" using 03 unfolding gs.lin_indpt_list_def by force
- have m0: "m \<ge> 0" using len by auto
- have Linv: "LLL_invariant_weak' 0 fs"
- by (intro LLL_invI_weak 03 04 01 wred fs_carr m0)
- note Linvw = LLL_invw'_imp_w[OF Linv]
- from compute_mod_of_max_gso_norm[OF b0 mb0 p]
- have p: "mod_invariant b p first" "p > 1" by auto
- from len mfs have len': "length mfs = m" by auto
- have modbnd: "\<forall>i'<m. \<forall>j'<i'. \<bar>d\<mu> fs i' j'\<bar> < p * d fs j' * d fs (Suc j')"
- proof -
- have "\<forall> i' < m. \<forall> j' < i'. d\<mu> fs i' j' = d\<mu> fs i' j' symmod (p * d fs j' * d fs (Suc j'))"
- using I_def 07 06 by simp
- moreover have "\<forall>j' < m. p * d fs j' * d fs (Suc j') > 0" using p(2) LLL_d_pos[OF Linvw] by simp
- ultimately show ?thesis using sym_mod_abs
- by (smt Euclidean_Division.pos_mod_bound Euclidean_Division.pos_mod_sign less_trans)
- qed
- have "LLL_invariant_mod fs mfs dmu' p first b 0"
- using LLL_invI_mod[OF len' m0 04 01 03 wred _ modbnd d\<mu>dmu' p(2) gbnd p(1)] 02 mfs by simp
- then show ?thesis by auto
-qed
-
-subsection \<open>Soundness of Storjohann's algorithm\<close>
-
-text \<open>For all of these abstract algorithms, we actually formulate their soundness proofs by linking
- to the LLL-invariant (which implies that @{term fs} is reduced (@{term "LLL_invariant True m fs"})
- or that the first vector of @{term fs} is short (@{term "LLL_invariant_weak fs \<and> weakly_reduced fs m"}).\<close>
-
-text \<open>Soundness of Storjohann's algorithm\<close>
-lemma reduce_basis_mod_inv: assumes res: "reduce_basis_mod = fs"
- shows "LLL_invariant True m fs"
-proof (cases "m = 0")
- case True
- from True have *: "fs_init = []" using len by simp
- moreover have "fs = []" using res basis_reduction_mod_add_rows_outer_loop.simps(1)
- unfolding reduce_basis_mod_def Let_def basis_reduction_mod_main.simps[of _ _ _ _ _ 0]
- compute_initial_mfs_def compute_initial_state_def compute_initial_dmu_def
- unfolding True * by (auto split: prod.splits)
- ultimately show ?thesis using True LLL_inv_initial_state by blast
-next
- case False
- let ?first = False
- obtain p mfs0 dmu0 g_idx0 where init: "compute_initial_state ?first = (p, mfs0, dmu0, g_idx0)" by (metis prod_cases4)
- from LLL_initial_invariant_mod[OF init]
- obtain fs0 b where fs0: "LLL_invariant_mod fs0 mfs0 dmu0 p ?first b 0" by blast
- note res = res[unfolded reduce_basis_mod_def init Let_def split]
- obtain p1 mfs1 dmu1 where mfs1dmu1: "(p1, mfs1, dmu1) = basis_reduction_mod_main p ?first mfs0 dmu0 g_idx0 0 0"
- by (metis prod.exhaust)
- obtain fs1 b1 where Linv1: "LLL_invariant_mod fs1 mfs1 dmu1 p1 ?first b1 m"
- using basis_reduction_mod_main[OF fs0 mfs1dmu1[symmetric]] by auto
- obtain mfs2 dmu2 where mfs2dmu2:
- "(mfs2, dmu2) = basis_reduction_mod_add_rows_outer_loop p1 mfs1 dmu1 (m-1)" by (metis old.prod.exhaust)
- obtain fs2 where fs2: "LLL_invariant_mod fs2 mfs2 dmu2 p1 ?first b1 m"
- and \<mu>s: "((\<forall>j. j < m \<longrightarrow> \<mu>_small fs2 j))"
- using basis_reduction_mod_add_rows_outer_loop_inv[OF _ mfs2dmu2, of fs1 ?first b1] Linv1 False by auto
- have rbd: "LLL_invariant_weak' m fs2" "\<forall>j < m. \<mu>_small fs2 j"
- using LLL_invD_mod[OF fs2] LLL_invI_weak \<mu>s by auto
- have redfs2: "reduced fs2 m" using rbd LLL_invD_weak(8) gram_schmidt_fs.reduced_def \<mu>_small_def by blast
- have fs: "fs = mfs2"
- using res[folded mfs1dmu1, unfolded Let_def split, folded mfs2dmu2, unfolded split] ..
- have "\<forall>i < m. fs2 ! i = fs ! i"
- proof (intro allI impI)
- fix i
- assume i: "i < m"
- then have fs2i: "LLL_invariant_mod fs2 mfs2 dmu2 p1 ?first b1 i"
- using fs2 LLL_invariant_mod_to_weak_m_to_i by simp
- have \<mu>si: "\<mu>_small fs2 i" using \<mu>s i by simp
- show "fs2 ! i = fs ! i"
- using basis_reduction_mod_fs_bound(1)[OF fs2i \<mu>si i] fs by simp
- qed
- then have "fs2 = fs"
- using LLL_invD_mod(1,3,10,13)[OF fs2] fs by (metis nth_equalityI)
- then show ?thesis using redfs2 fs rbd(1) reduce_basis_def res LLL_invD_weak
- LLL_invariant_def by simp
-qed
-
-text \<open>Soundness of Storjohann's algorithm for computing a short vector.\<close>
-lemma short_vector_mod_inv: assumes res: "short_vector_mod = v"
- and m: "m > 0"
- shows "\<exists> fs. LLL_invariant_weak fs \<and> weakly_reduced fs m \<and> v = hd fs"
-proof -
- let ?first = True
- obtain p mfs0 dmu0 g_idx0 where init: "compute_initial_state ?first = (p, mfs0, dmu0, g_idx0)" by (metis prod_cases4)
- from LLL_initial_invariant_mod[OF init]
- obtain fs0 b where fs0: "LLL_invariant_mod fs0 mfs0 dmu0 p ?first b 0" by blast
- obtain p1 mfs1 dmu1 where main: "basis_reduction_mod_main p ?first mfs0 dmu0 g_idx0 0 0 = (p1, mfs1, dmu1)"
- by (metis prod.exhaust)
- obtain fs1 b1 where Linv1: "LLL_invariant_mod fs1 mfs1 dmu1 p1 ?first b1 m"
- using basis_reduction_mod_main[OF fs0 main] by auto
- have "v = hd mfs1" using res[unfolded short_vector_mod_def Let_def init split main] ..
- with basis_reduction_mod_fs_bound_first[OF Linv1 m] LLL_invD_mod(1,3)[OF Linv1] m
- have v: "v = hd fs1" by (cases fs1; cases mfs1; auto)
- from Linv1 have Linv1: "LLL_invariant_weak fs1" and red: "weakly_reduced fs1 m"
- unfolding LLL_invariant_mod_def LLL_invariant_weak_def by auto
- show ?thesis
- by (intro exI[of _ fs1] conjI Linv1 red v)
-qed
-
-text \<open>Soundness of Storjohann's algorithm with improved swap order\<close>
-lemma reduce_basis_iso_inv: assumes res: "reduce_basis_iso = fs"
- shows "LLL_invariant True m fs"
-proof (cases "m = 0")
- case True
- then have *: "fs_init = []" using len by simp
- moreover have "fs = []" using res basis_reduction_mod_add_rows_outer_loop.simps(1)
- unfolding reduce_basis_iso_def Let_def basis_reduction_iso_main.simps[of _ _ _ _ _ 0]
- compute_initial_mfs_def compute_initial_state_def compute_initial_dmu_def
- unfolding True * by (auto split: prod.splits)
- ultimately show ?thesis using True LLL_inv_initial_state by blast
-next
- case False
- let ?first = False
- obtain p mfs0 dmu0 g_idx0 where init: "compute_initial_state ?first = (p, mfs0, dmu0, g_idx0)" by (metis prod_cases4)
- from LLL_initial_invariant_mod[OF init]
- obtain fs0 b where fs0: "LLL_invariant_mod fs0 mfs0 dmu0 p ?first b 0" by blast
- have fs0w: "LLL_invariant_mod_weak fs0 mfs0 dmu0 p ?first b" using LLL_invD_mod[OF fs0] LLL_invI_modw by simp
- note res = res[unfolded reduce_basis_iso_def init Let_def split]
- obtain p1 mfs1 dmu1 where mfs1dmu1: "(p1, mfs1, dmu1) = basis_reduction_iso_main p ?first mfs0 dmu0 g_idx0 0"
- by (metis prod.exhaust)
- obtain fs1 b1 where Linv1: "LLL_invariant_mod fs1 mfs1 dmu1 p1 ?first b1 m"
- using basis_reduction_iso_main[OF fs0w mfs1dmu1[symmetric]] by auto
- obtain mfs2 dmu2 where mfs2dmu2:
- "(mfs2, dmu2) = basis_reduction_mod_add_rows_outer_loop p1 mfs1 dmu1 (m-1)" by (metis old.prod.exhaust)
- obtain fs2 where fs2: "LLL_invariant_mod fs2 mfs2 dmu2 p1 ?first b1 m"
- and \<mu>s: "((\<forall>j. j < m \<longrightarrow> \<mu>_small fs2 j))"
- using basis_reduction_mod_add_rows_outer_loop_inv[OF _ mfs2dmu2, of fs1 ?first b1] Linv1 False by auto
- have rbd: "LLL_invariant_weak' m fs2" "\<forall>j < m. \<mu>_small fs2 j"
- using LLL_invD_mod[OF fs2] LLL_invI_weak \<mu>s by auto
- have redfs2: "reduced fs2 m" using rbd LLL_invD_weak(8) gram_schmidt_fs.reduced_def \<mu>_small_def by blast
- have fs: "fs = mfs2"
- using res[folded mfs1dmu1, unfolded Let_def split, folded mfs2dmu2, unfolded split] ..
- have "\<forall>i < m. fs2 ! i = fs ! i"
- proof (intro allI impI)
- fix i
- assume i: "i < m"
- then have fs2i: "LLL_invariant_mod fs2 mfs2 dmu2 p1 ?first b1 i"
- using fs2 LLL_invariant_mod_to_weak_m_to_i by simp
- have \<mu>si: "\<mu>_small fs2 i" using \<mu>s i by simp
- show "fs2 ! i = fs ! i"
- using basis_reduction_mod_fs_bound(1)[OF fs2i \<mu>si i] fs by simp
- qed
- then have "fs2 = fs"
- using LLL_invD_mod(1,3,10,13)[OF fs2] fs by (metis nth_equalityI)
- then show ?thesis using redfs2 fs rbd(1) reduce_basis_def res LLL_invD_weak
- LLL_invariant_def by simp
-qed
-
-text \<open>Soundness of Storjohann's algorithm to compute short vectors with improved swap order\<close>
-lemma short_vector_iso_inv: assumes res: "short_vector_iso = v"
- and m: "m > 0"
- shows "\<exists> fs. LLL_invariant_weak fs \<and> weakly_reduced fs m \<and> v = hd fs"
-proof -
- let ?first = True
- obtain p mfs0 dmu0 g_idx0 where init: "compute_initial_state ?first = (p, mfs0, dmu0, g_idx0)" by (metis prod_cases4)
- from LLL_initial_invariant_mod[OF init]
- obtain fs0 b where fs0: "LLL_invariant_mod fs0 mfs0 dmu0 p ?first b 0" by blast
- have fs0w: "LLL_invariant_mod_weak fs0 mfs0 dmu0 p ?first b" using LLL_invD_mod[OF fs0] LLL_invI_modw by simp
- obtain p1 mfs1 dmu1 where main: "basis_reduction_iso_main p ?first mfs0 dmu0 g_idx0 0 = (p1, mfs1, dmu1)"
- by (metis prod.exhaust)
- obtain fs1 b1 where Linv1: "LLL_invariant_mod fs1 mfs1 dmu1 p1 ?first b1 m"
- using basis_reduction_iso_main[OF fs0w main] by auto
- have "v = hd mfs1" using res[unfolded short_vector_iso_def Let_def init split main] ..
- with basis_reduction_mod_fs_bound_first[OF Linv1 m] LLL_invD_mod(1,3)[OF Linv1] m
- have v: "v = hd fs1" by (cases fs1; cases mfs1; auto)
- from Linv1 have Linv1: "LLL_invariant_weak fs1" and red: "weakly_reduced fs1 m"
- unfolding LLL_invariant_mod_def LLL_invariant_weak_def by auto
- show ?thesis
- by (intro exI[of _ fs1] conjI Linv1 red v)
-qed
-
-end
-
-text \<open>From the soundness results of these abstract versions of the algorithms,
- one just needs to derive actual implementations that may integrate low-level
- optimizations.\<close>
-
-end
+section \<open>Storjohann's basis reduction algorithm (abstract version)\<close>
+
+text \<open>This theory contains the soundness proofs of Storjohann's basis
+ reduction algorithms, both for the normal and the improved-swap-order variant.
+
+ The implementation of Storjohann's version of LLL uses modular operations throughout.
+ It is an abstract implementation that is already quite close to what the actual implementation will be.
+ In particular, the swap operation here is derived from the computation lemma for the swap
+ operation in the old, integer-only formalization of LLL.\<close>
+
+theory Storjohann
+ imports
+ Storjohann_Mod_Operation
+ LLL_Basis_Reduction.LLL_Number_Bounds
+ Sqrt_Babylonian.NthRoot_Impl
+begin
+
+subsection \<open>Definition of algorithm\<close>
+
+text \<open>In the definition of the algorithm, the first-flag determines, whether only the first vector
+ of the reduced basis should be computed, i.e., a short vector. Then the modulus can be slightly
+ decreased in comparison to the required modulus for computing the whole reduced matrix.\<close>
+
+fun max_list_rats_with_index :: "(int * int * nat) list \<Rightarrow> (int * int * nat)" where
+ "max_list_rats_with_index [x] = x" |
+ "max_list_rats_with_index ((n1,d1,i1) # (n2,d2,i2) # xs)
+ = max_list_rats_with_index ((if n1 * d2 \<le> n2 * d1 then (n2,d2,i2) else (n1,d1,i1)) # xs)"
+
+context LLL
+begin
+
+definition "log_base = (10 :: int)"
+
+definition bound_number :: "bool \<Rightarrow> nat" where
+ "bound_number first = (if first \<and> m \<noteq> 0 then 1 else m)"
+
+definition compute_mod_of_max_gso_norm :: "bool \<Rightarrow> rat \<Rightarrow> int" where
+ "compute_mod_of_max_gso_norm first mn = log_base ^ (log_ceiling log_base (max 2 (
+ root_rat_ceiling 2 (mn * (rat_of_nat (bound_number first) + 3)) + 1)))"
+
+definition g_bnd_mode :: "bool \<Rightarrow> rat \<Rightarrow> int vec list \<Rightarrow> bool" where
+ "g_bnd_mode first b fs = (if first \<and> m \<noteq> 0 then sq_norm (gso fs 0) \<le> b else g_bnd b fs)"
+
+definition d_of where "d_of dmu i = (if i = 0 then 1 :: int else dmu $$ (i - 1, i - 1))"
+
+definition compute_max_gso_norm :: "bool \<Rightarrow> int mat \<Rightarrow> rat \<times> nat" where
+ "compute_max_gso_norm first dmu = (if m = 0 then (0,0) else
+ case max_list_rats_with_index (map (\<lambda> i. (d_of dmu (Suc i), d_of dmu i, i)) [0 ..< (if first then 1 else m)])
+ of (num, denom, i) \<Rightarrow> (of_int num / of_int denom, i))"
+
+
+context
+ fixes p :: int \<comment> \<open>the modulus\<close>
+ and first :: bool \<comment> \<open>only compute first vector of reduced basis\<close>
+begin
+
+definition basis_reduction_mod_add_row ::
+ "int vec list \<Rightarrow> int mat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> (int vec list \<times> int mat)" where
+ "basis_reduction_mod_add_row mfs dmu i j =
+ (let c = round_num_denom (dmu $$ (i,j)) (d_of dmu (Suc j)) in
+ (if c = 0 then (mfs, dmu)
+ else (mfs[ i := (map_vec (\<lambda> x. x symmod p)) (mfs ! i - c \<cdot>\<^sub>v mfs ! j)],
+ mat m m (\<lambda>(i',j'). (if (i' = i \<and> j' \<le> j)
+ then (if j'=j then (dmu $$ (i,j') - c * dmu $$ (j,j'))
+ else (dmu $$ (i,j') - c * dmu $$ (j,j'))
+ symmod (p * d_of dmu j' * d_of dmu (Suc j')))
+ else (dmu $$ (i',j')))))))"
+
+fun basis_reduction_mod_add_rows_loop where
+ "basis_reduction_mod_add_rows_loop mfs dmu i 0 = (mfs, dmu)"
+| "basis_reduction_mod_add_rows_loop mfs dmu i (Suc j) = (
+ let (mfs', dmu') = basis_reduction_mod_add_row mfs dmu i j
+ in basis_reduction_mod_add_rows_loop mfs' dmu' i j)"
+
+definition basis_reduction_mod_swap_dmu_mod :: "int mat \<Rightarrow> nat \<Rightarrow> int mat" where
+ "basis_reduction_mod_swap_dmu_mod dmu k = mat m m (\<lambda>(i, j). (
+ if j < i \<and> (j = k \<or> j = k - 1) then
+ dmu $$ (i, j) symmod (p * d_of dmu j * d_of dmu (Suc j))
+ else dmu $$ (i, j)))"
+
+definition basis_reduction_mod_swap where
+ "basis_reduction_mod_swap mfs dmu k =
+ (mfs[k := mfs ! (k - 1), k - 1 := mfs ! k],
+ basis_reduction_mod_swap_dmu_mod (mat m m (\<lambda>(i,j). (
+ if j < i then
+ if i = k - 1 then
+ dmu $$ (k, j)
+ else if i = k \<and> j \<noteq> k - 1 then
+ dmu $$ (k - 1, j)
+ else if i > k \<and> j = k then
+ ((d_of dmu (Suc k)) * dmu $$ (i, k - 1) - dmu $$ (k, k - 1) * dmu $$ (i, j))
+ div (d_of dmu k)
+ else if i > k \<and> j = k - 1 then
+ (dmu $$ (k, k - 1) * dmu $$ (i, j) + dmu $$ (i, k) * (d_of dmu (k-1)))
+ div (d_of dmu k)
+ else dmu $$ (i, j)
+ else if i = j then
+ if i = k - 1 then
+ ((d_of dmu (Suc k)) * (d_of dmu (k-1)) + dmu $$ (k, k - 1) * dmu $$ (k, k - 1))
+ div (d_of dmu k)
+ else (d_of dmu (Suc i))
+ else dmu $$ (i, j))
+ )) k)"
+
+fun basis_reduction_adjust_mod where
+ "basis_reduction_adjust_mod mfs dmu =
+ (let (b,g_idx) = compute_max_gso_norm first dmu;
+ p' = compute_mod_of_max_gso_norm first b
+ in if p' < p then
+ let mfs' = map (map_vec (\<lambda>x. x symmod p')) mfs;
+ d_vec = vec (Suc m) (\<lambda> i. d_of dmu i);
+ dmu' = mat m m (\<lambda> (i,j). if j < i then dmu $$ (i,j)
+ symmod (p' * d_vec $ j * d_vec $ (Suc j)) else
+ dmu $$ (i,j))
+ in (p', mfs', dmu', g_idx)
+ else (p, mfs, dmu, g_idx))"
+
+definition basis_reduction_adjust_swap_add_step where
+ "basis_reduction_adjust_swap_add_step mfs dmu g_idx i = (
+ let i1 = i - 1;
+ (mfs1, dmu1) = basis_reduction_mod_add_row mfs dmu i i1;
+ (mfs2, dmu2) = basis_reduction_mod_swap mfs1 dmu1 i
+ in if i1 = g_idx then basis_reduction_adjust_mod mfs2 dmu2
+ else (p, mfs2, dmu2, g_idx))"
+
+
+definition basis_reduction_mod_step where
+ "basis_reduction_mod_step mfs dmu g_idx i (j :: int) = (if i = 0 then (p, mfs, dmu, g_idx, Suc i, j)
+ else let di = d_of dmu i;
+ (num, denom) = quotient_of \<alpha>
+ in if di * di * denom \<le> num * d_of dmu (i - 1) * d_of dmu (Suc i) then
+ (p, mfs, dmu, g_idx, Suc i, j)
+ else let (p', mfs', dmu', g_idx') = basis_reduction_adjust_swap_add_step mfs dmu g_idx i
+ in (p', mfs', dmu', g_idx', i - 1, j + 1))"
+
+primrec basis_reduction_mod_add_rows_outer_loop where
+ "basis_reduction_mod_add_rows_outer_loop mfs dmu 0 = (mfs, dmu)" |
+ "basis_reduction_mod_add_rows_outer_loop mfs dmu (Suc i) =
+ (let (mfs', dmu') = basis_reduction_mod_add_rows_outer_loop mfs dmu i in
+ basis_reduction_mod_add_rows_loop mfs' dmu' (Suc i) (Suc i))"
+end
+
+text \<open>the main loop of the normal Storjohann algorithm\<close>
+partial_function (tailrec) basis_reduction_mod_main where
+ "basis_reduction_mod_main p first mfs dmu g_idx i (j :: int) = (
+ (if i < m
+ then
+ case basis_reduction_mod_step p first mfs dmu g_idx i j
+ of (p', mfs', dmu', g_idx', i', j') \<Rightarrow>
+ basis_reduction_mod_main p' first mfs' dmu' g_idx' i' j'
+ else
+ (p, mfs, dmu)))"
+
+definition compute_max_gso_quot:: "int mat \<Rightarrow> (int * int * nat)" where
+ "compute_max_gso_quot dmu = max_list_rats_with_index
+ (map (\<lambda>i. ((d_of dmu (i+1)) * (d_of dmu (i+1)), (d_of dmu (i+2)) * (d_of dmu i), Suc i)) [0..<(m-1)])"
+
+text \<open>the main loop of Storjohann's algorithm with improved swap order\<close>
+partial_function (tailrec) basis_reduction_iso_main where
+ "basis_reduction_iso_main p first mfs dmu g_idx (j :: int) = (
+ (if m > 1 then
+ (let (max_gso_num, max_gso_denum, indx) = compute_max_gso_quot dmu;
+ (num, denum) = quotient_of \<alpha> in
+ (if (max_gso_num * denum > num * max_gso_denum) then
+ case basis_reduction_adjust_swap_add_step p first mfs dmu g_idx indx of
+ (p', mfs', dmu', g_idx') \<Rightarrow>
+ basis_reduction_iso_main p' first mfs' dmu' g_idx' (j + 1)
+ else
+ (p, mfs, dmu)))
+ else (p, mfs, dmu)))"
+
+definition compute_initial_mfs where
+ "compute_initial_mfs p = map (map_vec (\<lambda>x. x symmod p)) fs_init"
+
+definition compute_initial_dmu where
+ "compute_initial_dmu p dmu = mat m m (\<lambda>(i',j'). if j' < i'
+ then dmu $$ (i', j') symmod (p * d_of dmu j' * d_of dmu (Suc j'))
+ else dmu $$ (i', j'))"
+
+definition "dmu_initial = (let dmu = d\<mu>_impl fs_init
+ in mat m m (\<lambda> (i,j).
+ if j \<le> i then d\<mu>_impl fs_init !! i !! j else 0))"
+
+definition "compute_initial_state first =
+ (let dmu = dmu_initial;
+ (b, g_idx) = compute_max_gso_norm first dmu;
+ p = compute_mod_of_max_gso_norm first b
+ in (p, compute_initial_mfs p, compute_initial_dmu p dmu, g_idx))"
+
+text \<open>Storjohann's algorithm\<close>
+definition reduce_basis_mod :: "int vec list" where
+ "reduce_basis_mod = (
+ let first = False;
+ (p0, mfs0, dmu0, g_idx) = compute_initial_state first;
+ (p', mfs', dmu') = basis_reduction_mod_main p0 first mfs0 dmu0 g_idx 0 0;
+ (mfs'', dmu'') = basis_reduction_mod_add_rows_outer_loop p' mfs' dmu' (m-1)
+ in mfs'')"
+
+text \<open>Storjohann's algorithm with improved swap order\<close>
+definition reduce_basis_iso :: "int vec list" where
+ "reduce_basis_iso = (
+ let first = False;
+ (p0, mfs0, dmu0, g_idx) = compute_initial_state first;
+ (p', mfs', dmu') = basis_reduction_iso_main p0 first mfs0 dmu0 g_idx 0;
+ (mfs'', dmu'') = basis_reduction_mod_add_rows_outer_loop p' mfs' dmu' (m-1)
+ in mfs'')"
+
+text \<open>Storjohann's algorithm for computing a short vector\<close>
+definition
+ "short_vector_mod = (
+ let first = True;
+ (p0, mfs0, dmu0, g_idx) = compute_initial_state first;
+ (p', mfs', dmu') = basis_reduction_mod_main p0 first mfs0 dmu0 g_idx 0 0
+ in hd mfs')"
+
+text \<open>Storjohann's algorithm (iso-variant) for computing a short vector\<close>
+definition
+ "short_vector_iso = (
+ let first = True;
+ (p0, mfs0, dmu0, g_idx) = compute_initial_state first;
+ (p', mfs', dmu') = basis_reduction_iso_main p0 first mfs0 dmu0 g_idx 0
+ in hd mfs')"
+end
+
+subsection \<open>Towards soundness of Storjohann's algorithm\<close>
+
+lemma max_list_rats_with_index_in_set:
+ assumes max: "max_list_rats_with_index xs = (nm, dm, im)"
+ and len: "length xs \<ge> 1"
+shows "(nm, dm, im) \<in> set xs"
+ using assms
+proof (induct xs rule: max_list_rats_with_index.induct)
+ case (2 n1 d1 i1 n2 d2 i2 xs)
+ have "1 \<le> length ((if n1 * d2 \<le> n2 * d1 then (n2, d2, i2) else (n1, d1, i1)) # xs)" by simp
+ moreover have "max_list_rats_with_index ((if n1 * d2 \<le> n2 * d1 then (n2, d2, i2) else (n1, d1, i1)) # xs)
+ = (nm, dm, im)" using 2 by simp
+ moreover have "(if n1 * d2 \<le> n2 * d1 then (n2, d2, i2) else (n1, d1, i1)) \<in>
+ set ((n1, d1, i1) # (n2, d2, i2) # xs)" by simp
+ moreover then have "set ((if n1 * d2 \<le> n2 * d1 then (n2, d2, i2) else (n1, d1, i1)) # xs) \<subseteq>
+ set ((n1, d1, i1) # (n2, d2, i2) # xs)" by auto
+ ultimately show ?case using 2(1) by auto
+qed auto
+
+lemma max_list_rats_with_index: assumes "\<And> n d i. (n,d,i) \<in> set xs \<Longrightarrow> d > 0"
+ and max: "max_list_rats_with_index xs = (nm, dm, im)"
+ and "(n,d,i) \<in> set xs"
+shows "rat_of_int n / of_int d \<le> of_int nm / of_int dm"
+ using assms
+proof (induct xs arbitrary: n d i rule: max_list_rats_with_index.induct)
+ case (2 n1 d1 i1 n2 d2 i2 xs n d i)
+ let ?r = "rat_of_int"
+ from 2(2) have "d1 > 0" "d2 > 0" by auto
+ hence d: "?r d1 > 0" "?r d2 > 0" by auto
+ have "(n1 * d2 \<le> n2 * d1) = (?r n1 * ?r d2 \<le> ?r n2 * ?r d1)"
+ unfolding of_int_mult[symmetric] by presburger
+ also have "\<dots> = (?r n1 / ?r d1 \<le> ?r n2 / ?r d2)" using d
+ by (smt divide_strict_right_mono leD le_less_linear mult.commute nonzero_mult_div_cancel_left
+ not_less_iff_gr_or_eq times_divide_eq_right)
+ finally have id: "(n1 * d2 \<le> n2 * d1) = (?r n1 / ?r d1 \<le> ?r n2 / ?r d2)" .
+ obtain n' d' i' where new: "(if n1 * d2 \<le> n2 * d1 then (n2, d2, i2) else (n1, d1, i1)) = (n',d',i')"
+ by force
+ have nd': "(n',d',i') \<in> {(n1,d1,i1), (n2, d2, i2)}" using new[symmetric] by auto
+ from 2(3) have res: "max_list_rats_with_index ((n',d',i') # xs) = (nm, dm, im)" using new by auto
+ note 2 = 2[unfolded new]
+ show ?case
+ proof (cases "(n,d,i) \<in> set xs")
+ case True
+ show ?thesis
+ by (rule 2(1)[of n d, OF 2(2) res], insert True nd', force+)
+ next
+ case False
+ with 2(4) have "n = n1 \<and> d = d1 \<or> n = n2 \<and> d = d2" by auto
+ hence "?r n / ?r d \<le> ?r n' / ?r d'" using new[unfolded id]
+ by (metis linear prod.inject)
+ also have "?r n' / ?r d' \<le> ?r nm / ?r dm"
+ by (rule 2(1)[of n' d', OF 2(2) res], insert nd', force+)
+ finally show ?thesis .
+ qed
+qed auto
+
+context LLL
+begin
+
+lemma log_base: "log_base \<ge> 2" unfolding log_base_def by auto
+
+definition LLL_invariant_weak' :: "nat \<Rightarrow> int vec list \<Rightarrow> bool" where
+ "LLL_invariant_weak' i fs = (
+ gs.lin_indpt_list (RAT fs) \<and>
+ lattice_of fs = L \<and>
+ weakly_reduced fs i \<and>
+ i \<le> m \<and>
+ length fs = m
+ )"
+
+lemma LLL_invD_weak: assumes "LLL_invariant_weak' i fs"
+ shows
+ "lin_indep fs"
+ "length (RAT fs) = m"
+ "set fs \<subseteq> carrier_vec n"
+ "\<And> i. i < m \<Longrightarrow> fs ! i \<in> carrier_vec n"
+ "\<And> i. i < m \<Longrightarrow> gso fs i \<in> carrier_vec n"
+ "length fs = m"
+ "lattice_of fs = L"
+ "weakly_reduced fs i"
+ "i \<le> m"
+proof (atomize (full), goal_cases)
+ case 1
+ interpret gs': gram_schmidt_fs_lin_indpt n "RAT fs"
+ by (standard) (use assms LLL_invariant_weak'_def gs.lin_indpt_list_def in auto)
+ show ?case
+ using assms gs'.fs_carrier gs'.f_carrier gs'.gso_carrier
+ by (auto simp add: LLL_invariant_weak'_def gram_schmidt_fs.reduced_def)
+qed
+
+lemma LLL_invI_weak: assumes
+ "set fs \<subseteq> carrier_vec n"
+ "length fs = m"
+ "lattice_of fs = L"
+ "i \<le> m"
+ "lin_indep fs"
+ "weakly_reduced fs i"
+shows "LLL_invariant_weak' i fs"
+ unfolding LLL_invariant_weak'_def Let_def using assms by auto
+
+lemma LLL_invw'_imp_w: "LLL_invariant_weak' i fs \<Longrightarrow> LLL_invariant_weak fs"
+ unfolding LLL_invariant_weak'_def LLL_invariant_weak_def by auto
+
+lemma basis_reduction_add_row_weak:
+ assumes Linvw: "LLL_invariant_weak' i fs"
+ and i: "i < m" and j: "j < i"
+ and fs': "fs' = fs[ i := fs ! i - c \<cdot>\<^sub>v fs ! j]"
+shows "LLL_invariant_weak' i fs'"
+ "g_bnd B fs \<Longrightarrow> g_bnd B fs'"
+proof (atomize(full), goal_cases)
+ case 1
+ note Linv = LLL_invw'_imp_w[OF Linvw]
+ note main = basis_reduction_add_row_main[OF Linv i j fs']
+ have bnd: "g_bnd B fs \<Longrightarrow> g_bnd B fs'" using main(6) unfolding g_bnd_def by auto
+ note new = LLL_inv_wD[OF main(1)]
+ note old = LLL_invD_weak[OF Linvw]
+ have red: "weakly_reduced fs' i" using \<open>weakly_reduced fs i\<close> main(6) \<open>i < m\<close>
+ unfolding gram_schmidt_fs.weakly_reduced_def by auto
+ have inv: "LLL_invariant_weak' i fs'" using LLL_inv_wD[OF main(1)] \<open>i < m\<close>
+ by (intro LLL_invI_weak, auto intro: red)
+ show ?case using inv red main bnd by auto
+qed
+
+lemma LLL_inv_weak_m_impl_i:
+ assumes inv: "LLL_invariant_weak' m fs"
+ and i: "i \<le> m"
+shows "LLL_invariant_weak' i fs"
+proof -
+ have "weakly_reduced fs i" using LLL_invD_weak(8)[OF inv]
+ by (meson assms(2) gram_schmidt_fs.weakly_reduced_def le_trans less_imp_le_nat linorder_not_less)
+ then show ?thesis
+ using LLL_invI_weak[of fs i, OF LLL_invD_weak(3,6,7)[OF inv] _ LLL_invD_weak(1)[OF inv]]
+ LLL_invD_weak(2,4,5,8-)[OF inv] i by simp
+qed
+
+definition mod_invariant where
+ "mod_invariant b p first = (b \<le> rat_of_int (p - 1)^2 / (rat_of_nat (bound_number first) + 3)
+ \<and> (\<exists> e. p = log_base ^ e))"
+
+lemma compute_mod_of_max_gso_norm: assumes mn: "mn \<ge> 0"
+ and m: "m = 0 \<Longrightarrow> mn = 0"
+ and p: "p = compute_mod_of_max_gso_norm first mn"
+shows
+ "p > 1"
+ "mod_invariant mn p first"
+proof -
+ let ?m = "bound_number first"
+ define p' where "p' = root_rat_ceiling 2 (mn * (rat_of_nat ?m + 3)) + 1"
+ define p'' where "p'' = max 2 p'"
+ define q where "q = real_of_rat (mn * (rat_of_nat ?m + 3))"
+ have *: "-1 < (0 :: real)" by simp
+ also have "0 \<le> root 2 (real_of_rat (mn * (rat_of_nat ?m + 3)))" using mn by auto
+ finally have "p' \<ge> 0 + 1" unfolding p'_def
+ by (intro plus_left_mono, simp)
+ hence p': "p' > 0" by auto
+ have p'': "p'' > 1" unfolding p''_def by auto
+ have pp'': "p \<ge> p''" unfolding compute_mod_of_max_gso_norm_def p p'_def[symmetric] p''_def[symmetric]
+ using log_base p'' log_ceiling_sound by auto
+ hence pp': "p \<ge> p'" unfolding p''_def by auto
+ show "p > 1" using pp'' p'' by auto
+
+ have q0: "q \<ge> 0" unfolding q_def using mn m by auto
+ have "(mn \<le> rat_of_int (p' - 1)^2 / (rat_of_nat ?m + 3))
+ = (real_of_rat mn \<le> real_of_rat (rat_of_int (p' - 1)^2 / (rat_of_nat ?m + 3)))" using of_rat_less_eq by blast
+ also have "\<dots> = (real_of_rat mn \<le> real_of_rat (rat_of_int (p' - 1)^2) / real_of_rat (rat_of_nat ?m + 3))" by (simp add: of_rat_divide)
+ also have "\<dots> = (real_of_rat mn \<le> ((real_of_int (p' - 1))^2) / real_of_rat (rat_of_nat ?m + 3))"
+ by (metis of_rat_of_int_eq of_rat_power)
+ also have "\<dots> = (real_of_rat mn \<le> (real_of_int \<lceil>sqrt q\<rceil>)^2 / real_of_rat (rat_of_nat ?m + 3))"
+ unfolding p'_def sqrt_def q_def by simp
+ also have "\<dots>"
+ proof -
+ have "real_of_rat mn \<le> q / real_of_rat (rat_of_nat ?m + 3)" unfolding q_def using m
+ by (auto simp: of_rat_mult)
+ also have "\<dots> \<le> (real_of_int \<lceil>sqrt q\<rceil>)^2 / real_of_rat (rat_of_nat ?m + 3)"
+ proof (rule divide_right_mono)
+ have "q = (sqrt q)^2" using q0 by simp
+ also have "\<dots> \<le> (real_of_int \<lceil>sqrt q\<rceil>)^2"
+ by (rule power_mono, auto simp: q0)
+ finally show "q \<le> (real_of_int \<lceil>sqrt q\<rceil>)^2" .
+ qed auto
+ finally show ?thesis .
+ qed
+ finally have "mn \<le> rat_of_int (p' - 1)^2 / (rat_of_nat ?m + 3)" .
+ also have "\<dots> \<le> rat_of_int (p - 1)^2 / (rat_of_nat ?m + 3)"
+ unfolding power2_eq_square
+ by (intro divide_right_mono mult_mono, insert p' pp', auto)
+ finally have "mn \<le> rat_of_int (p - 1)^2 / (rat_of_nat ?m + 3)" .
+ moreover have "\<exists> e. p = log_base ^ e" unfolding p compute_mod_of_max_gso_norm_def by auto
+ ultimately show "mod_invariant mn p first" unfolding mod_invariant_def by auto
+qed
+
+lemma g_bnd_mode_cong: assumes "\<And> i. i < m \<Longrightarrow> gso fs i = gso fs' i"
+ shows "g_bnd_mode first b fs = g_bnd_mode first b fs'"
+ using assms unfolding g_bnd_mode_def g_bnd_def by auto
+
+definition LLL_invariant_mod :: "int vec list \<Rightarrow> int vec list \<Rightarrow> int mat \<Rightarrow> int \<Rightarrow> bool \<Rightarrow> rat \<Rightarrow> nat \<Rightarrow> bool" where
+ "LLL_invariant_mod fs mfs dmu p first b i = (
+ length fs = m \<and>
+ length mfs = m \<and>
+ i \<le> m \<and>
+ lattice_of fs = L \<and>
+ gs.lin_indpt_list (RAT fs) \<and>
+ weakly_reduced fs i \<and>
+ (map (map_vec (\<lambda>x. x symmod p)) fs = mfs) \<and>
+ (\<forall>i' < m. \<forall> j' < i'. \<bar>d\<mu> fs i' j'\<bar> < p * d fs j' * d fs (Suc j')) \<and>
+ (\<forall>i' < m. \<forall>j' < m. d\<mu> fs i' j' = dmu $$ (i',j')) \<and>
+ p > 1 \<and>
+ g_bnd_mode first b fs \<and>
+ mod_invariant b p first
+)"
+
+lemma LLL_invD_mod: assumes "LLL_invariant_mod fs mfs dmu p first b i"
+shows
+ "length mfs = m"
+ "i \<le> m"
+ "length fs = m"
+ "lattice_of fs = L"
+ "gs.lin_indpt_list (RAT fs)"
+ "weakly_reduced fs i"
+ "(map (map_vec (\<lambda>x. x symmod p)) fs = mfs)"
+ "(\<forall>i' < m. \<forall>j' < i'. \<bar>d\<mu> fs i' j'\<bar> < p * d fs j' * d fs (Suc j'))"
+ "(\<forall>i' < m. \<forall>j' < m. d\<mu> fs i' j' = dmu $$ (i',j'))"
+ "\<And> i. i < m \<Longrightarrow> fs ! i \<in> carrier_vec n"
+ "set fs \<subseteq> carrier_vec n"
+ "\<And> i. i < m \<Longrightarrow> gso fs i \<in> carrier_vec n"
+ "\<And> i. i < m \<Longrightarrow> mfs ! i \<in> carrier_vec n"
+ "set mfs \<subseteq> carrier_vec n"
+ "p > 1"
+ "g_bnd_mode first b fs"
+ "mod_invariant b p first"
+proof (atomize (full), goal_cases)
+ case 1
+ interpret gs': gram_schmidt_fs_lin_indpt n "RAT fs"
+ using assms LLL_invariant_mod_def gs.lin_indpt_list_def
+ by (meson gram_schmidt_fs_Rn.intro gram_schmidt_fs_lin_indpt.intro gram_schmidt_fs_lin_indpt_axioms.intro)
+ have allfs: "\<forall>i < m. fs ! i \<in> carrier_vec n" using assms gs'.f_carrier
+ by (simp add: LLL.LLL_invariant_mod_def)
+ then have setfs: "set fs \<subseteq> carrier_vec n" by (metis LLL_invariant_mod_def assms in_set_conv_nth subsetI)
+ have allgso: "(\<forall>i < m. gso fs i \<in> carrier_vec n)" using assms gs'.gso_carrier
+ by (simp add: LLL.LLL_invariant_mod_def)
+ show ?case
+ using assms gs'.fs_carrier gs'.f_carrier gs'.gso_carrier allfs allgso
+ LLL_invariant_mod_def gram_schmidt_fs.reduced_def in_set_conv_nth setfs by fastforce
+qed
+
+lemma LLL_invI_mod: assumes
+ "length mfs = m"
+ "i \<le> m"
+ "length fs = m"
+ "lattice_of fs = L"
+ "gs.lin_indpt_list (RAT fs)"
+ "weakly_reduced fs i"
+ "map (map_vec (\<lambda>x. x symmod p)) fs = mfs"
+ "(\<forall>i' < m. \<forall>j' < i'. \<bar>d\<mu> fs i' j'\<bar> < p * d fs j' * d fs (Suc j'))"
+ "(\<forall>i' < m. \<forall>j' < m. d\<mu> fs i' j' = dmu $$ (i',j'))"
+ "p > 1"
+ "g_bnd_mode first b fs"
+ "mod_invariant b p first"
+shows "LLL_invariant_mod fs mfs dmu p first b i"
+ unfolding LLL_invariant_mod_def using assms by blast
+
+definition LLL_invariant_mod_weak :: "int vec list \<Rightarrow> int vec list \<Rightarrow> int mat \<Rightarrow> int \<Rightarrow> bool \<Rightarrow> rat \<Rightarrow> bool" where
+ "LLL_invariant_mod_weak fs mfs dmu p first b = (
+ length fs = m \<and>
+ length mfs = m \<and>
+ lattice_of fs = L \<and>
+ gs.lin_indpt_list (RAT fs) \<and>
+ (map (map_vec (\<lambda>x. x symmod p)) fs = mfs) \<and>
+ (\<forall>i' < m. \<forall> j' < i'. \<bar>d\<mu> fs i' j'\<bar> < p * d fs j' * d fs (Suc j')) \<and>
+ (\<forall>i' < m. \<forall>j' < m. d\<mu> fs i' j' = dmu $$ (i',j')) \<and>
+ p > 1 \<and>
+ g_bnd_mode first b fs \<and>
+ mod_invariant b p first
+)"
+
+lemma LLL_invD_modw: assumes "LLL_invariant_mod_weak fs mfs dmu p first b"
+shows
+ "length mfs = m"
+ "length fs = m"
+ "lattice_of fs = L"
+ "gs.lin_indpt_list (RAT fs)"
+ "(map (map_vec (\<lambda>x. x symmod p)) fs = mfs)"
+ "(\<forall>i' < m. \<forall>j' < i'. \<bar>d\<mu> fs i' j'\<bar> < p * d fs j' * d fs (Suc j'))"
+ "(\<forall>i' < m. \<forall>j' < m. d\<mu> fs i' j' = dmu $$ (i',j'))"
+ "\<And> i. i < m \<Longrightarrow> fs ! i \<in> carrier_vec n"
+ "set fs \<subseteq> carrier_vec n"
+ "\<And> i. i < m \<Longrightarrow> gso fs i \<in> carrier_vec n"
+ "\<And> i. i < m \<Longrightarrow> mfs ! i \<in> carrier_vec n"
+ "set mfs \<subseteq> carrier_vec n"
+ "p > 1"
+ "g_bnd_mode first b fs"
+ "mod_invariant b p first"
+proof (atomize (full), goal_cases)
+ case 1
+ interpret gs': gram_schmidt_fs_lin_indpt n "RAT fs"
+ using assms LLL_invariant_mod_weak_def gs.lin_indpt_list_def
+ by (meson gram_schmidt_fs_Rn.intro gram_schmidt_fs_lin_indpt.intro gram_schmidt_fs_lin_indpt_axioms.intro)
+ have allfs: "\<forall>i < m. fs ! i \<in> carrier_vec n" using assms gs'.f_carrier
+ by (simp add: LLL.LLL_invariant_mod_weak_def)
+ then have setfs: "set fs \<subseteq> carrier_vec n" by (metis LLL_invariant_mod_weak_def assms in_set_conv_nth subsetI)
+ have allgso: "(\<forall>i < m. gso fs i \<in> carrier_vec n)" using assms gs'.gso_carrier
+ by (simp add: LLL.LLL_invariant_mod_weak_def)
+ show ?case
+ using assms gs'.fs_carrier gs'.f_carrier gs'.gso_carrier allfs allgso
+ LLL_invariant_mod_weak_def gram_schmidt_fs.reduced_def in_set_conv_nth setfs by fastforce
+qed
+
+lemma LLL_invI_modw: assumes
+ "length mfs = m"
+ "length fs = m"
+ "lattice_of fs = L"
+ "gs.lin_indpt_list (RAT fs)"
+ "map (map_vec (\<lambda>x. x symmod p)) fs = mfs"
+ "(\<forall>i' < m. \<forall>j' < i'. \<bar>d\<mu> fs i' j'\<bar> < p * d fs j' * d fs (Suc j'))"
+ "(\<forall>i' < m. \<forall>j' < m. d\<mu> fs i' j' = dmu $$ (i',j'))"
+ "p > 1"
+ "g_bnd_mode first b fs"
+ "mod_invariant b p first"
+shows "LLL_invariant_mod_weak fs mfs dmu p first b"
+ unfolding LLL_invariant_mod_weak_def using assms by blast
+
+lemma dd\<mu>:
+ assumes i: "i < m"
+ shows "d fs (Suc i) = d\<mu> fs i i"
+proof-
+ have "\<mu> fs i i = 1" using i by (simp add: gram_schmidt_fs.\<mu>.simps)
+ then show ?thesis using d\<mu>_def by simp
+qed
+
+lemma d_of_main: assumes "(\<forall>i' < m. d\<mu> fs i' i' = dmu $$ (i',i'))"
+ and "i \<le> m"
+shows "d_of dmu i = d fs i"
+proof (cases "i = 0")
+ case False
+ with assms have "i - 1 < m" by auto
+ from assms(1)[rule_format, OF this] dd\<mu>[OF this, of fs] False
+ show ?thesis by (simp add: d_of_def)
+next
+ case True
+ thus ?thesis unfolding d_of_def True d_def by simp
+qed
+
+lemma d_of: assumes inv: "LLL_invariant_mod fs mfs dmu p b first j"
+ and "i \<le> m"
+shows "d_of dmu i = d fs i"
+ by (rule d_of_main[OF _ assms(2)], insert LLL_invD_mod(9)[OF inv], auto)
+
+lemma d_of_weak: assumes inv: "LLL_invariant_mod_weak fs mfs dmu p first b"
+ and "i \<le> m"
+shows "d_of dmu i = d fs i"
+ by (rule d_of_main[OF _ assms(2)], insert LLL_invD_modw(7)[OF inv], auto)
+
+lemma compute_max_gso_norm: assumes dmu: "(\<forall>i' < m. d\<mu> fs i' i' = dmu $$ (i',i'))"
+ and Linv: "LLL_invariant_weak fs"
+shows "g_bnd_mode first (fst (compute_max_gso_norm first dmu)) fs"
+ "fst (compute_max_gso_norm first dmu) \<ge> 0"
+ "m = 0 \<Longrightarrow> fst (compute_max_gso_norm first dmu) = 0"
+proof -
+ show gbnd: "g_bnd_mode first (fst (compute_max_gso_norm first dmu)) fs"
+ proof (cases "first \<and> m \<noteq> 0")
+ case False
+ have "?thesis = (g_bnd (fst (compute_max_gso_norm first dmu)) fs)" unfolding g_bnd_mode_def using False by auto
+ also have \<dots> unfolding g_bnd_def
+ proof (intro allI impI)
+ fix i
+ assume i: "i < m"
+ have id: "(if first then 1 else m) = m" using False i by auto
+ define list where "list = map (\<lambda> i. (d_of dmu (Suc i), d_of dmu i, i)) [0 ..< m ]"
+ obtain num denom j where ml: "max_list_rats_with_index list = (num, denom, j)"
+ by (metis prod_cases3)
+ have dpos: "d fs i > 0" using LLL_d_pos[OF Linv, of i] i by auto
+ have pos: "(n, d, i) \<in> set list \<Longrightarrow> 0 < d" for n d i
+ using LLL_d_pos[OF Linv] unfolding list_def using d_of_main[OF dmu] by auto
+ from i have "list ! i \<in> set list" using i unfolding list_def by auto
+ also have "list ! i = (d_of dmu (Suc i), d_of dmu i, i)" unfolding list_def using i by auto
+ also have "\<dots> = (d fs (Suc i), d fs i, i)" using d_of_main[OF dmu] i by auto
+ finally have "(d fs (Suc i), d fs i, i) \<in> set list" .
+ from max_list_rats_with_index[OF pos ml this]
+ have "of_int (d fs (Suc i)) / of_int (d fs i) \<le> fst (compute_max_gso_norm first dmu)"
+ unfolding compute_max_gso_norm_def list_def[symmetric] ml id split using i by auto
+ also have "of_int (d fs (Suc i)) / of_int (d fs i) = sq_norm (gso fs i)"
+ using LLL_d_Suc[OF Linv i] dpos by auto
+ finally show "sq_norm (gso fs i) \<le> fst (compute_max_gso_norm first dmu)" .
+ qed
+ finally show ?thesis .
+ next
+ case True
+ thus ?thesis unfolding g_bnd_mode_def compute_max_gso_norm_def using d_of_main[OF dmu]
+ LLL_d_Suc[OF Linv, of 0] LLL_d_pos[OF Linv, of 0] LLL_d_pos[OF Linv, of 1] by auto
+ qed
+ show "fst (compute_max_gso_norm first dmu) \<ge> 0"
+ proof (cases "m = 0")
+ case True
+ thus ?thesis unfolding compute_max_gso_norm_def by simp
+ next
+ case False
+ hence 0: "0 < m" by simp
+ have "0 \<le> sq_norm (gso fs 0)" by blast
+ also have "\<dots> \<le> fst (compute_max_gso_norm first dmu)"
+ using gbnd[unfolded g_bnd_mode_def g_bnd_def] using 0 by metis
+ finally show ?thesis .
+ qed
+qed (auto simp: LLL.compute_max_gso_norm_def)
+
+
+lemma increase_i_mod:
+ assumes Linv: "LLL_invariant_mod fs mfs dmu p first b i"
+ and i: "i < m"
+ and red_i: "i \<noteq> 0 \<Longrightarrow> sq_norm (gso fs (i - 1)) \<le> \<alpha> * sq_norm (gso fs i)"
+shows "LLL_invariant_mod fs mfs dmu p first b (Suc i)" "LLL_measure i fs > LLL_measure (Suc i) fs"
+proof -
+ note inv = LLL_invD_mod[OF Linv]
+ from inv have red: "weakly_reduced fs i" by (auto)
+ from red red_i i have red: "weakly_reduced fs (Suc i)"
+ unfolding gram_schmidt_fs.weakly_reduced_def
+ by (intro allI impI, rename_tac ii, case_tac "Suc ii = i", auto)
+ show "LLL_invariant_mod fs mfs dmu p first b (Suc i)"
+ by (intro LLL_invI_mod, insert inv red i, auto)
+ show "LLL_measure i fs > LLL_measure (Suc i) fs" unfolding LLL_measure_def using i by auto
+qed
+
+lemma basis_reduction_mod_add_row_main:
+ assumes Linvmw: "LLL_invariant_mod_weak fs mfs dmu p first b"
+ and i: "i < m" and j: "j < i"
+ and c: "c = round (\<mu> fs i j)"
+ and mfs': "mfs' = mfs[ i := (map_vec (\<lambda> x. x symmod p)) (mfs ! i - c \<cdot>\<^sub>v mfs ! j)]"
+ and dmu': "dmu' = mat m m (\<lambda>(i',j'). (if (i' = i \<and> j' \<le> j)
+ then (if j'=j then (dmu $$ (i,j') - c * dmu $$ (j,j'))
+ else (dmu $$ (i,j') - c * dmu $$ (j,j'))
+ symmod (p * (d_of dmu j') * (d_of dmu (Suc j'))))
+ else (dmu $$ (i',j'))))"
+shows "(\<exists>fs'. LLL_invariant_mod_weak fs' mfs' dmu' p first b \<and>
+ LLL_measure i fs' = LLL_measure i fs
+ \<and> (\<mu>_small_row i fs (Suc j) \<longrightarrow> \<mu>_small_row i fs' j)
+ \<and> (\<forall>k < m. gso fs' k = gso fs k)
+ \<and> (\<forall>ii \<le> m. d fs' ii = d fs ii)
+ \<and> \<bar>\<mu> fs' i j\<bar> \<le> 1 / 2
+ \<and> (\<forall>i' j'. i' < i \<longrightarrow> j' \<le> i' \<longrightarrow> \<mu> fs' i' j' = \<mu> fs i' j')
+ \<and> (LLL_invariant_mod fs mfs dmu p first b i \<longrightarrow> LLL_invariant_mod fs' mfs' dmu' p first b i))"
+proof -
+ define fs' where "fs' = fs[ i := fs ! i - c \<cdot>\<^sub>v fs ! j]"
+ from LLL_invD_modw[OF Linvmw] have gbnd: "g_bnd_mode first b fs" and p1: "p > 1" and pgtz: "p > 0" by auto
+ have Linvww: "LLL_invariant_weak fs" using LLL_invD_modw[OF Linvmw] LLL_invariant_weak_def by simp
+ have
+ Linvw': "LLL_invariant_weak fs'" and
+ 01: "c = round (\<mu> fs i j) \<Longrightarrow> \<mu>_small_row i fs (Suc j) \<Longrightarrow> \<mu>_small_row i fs' j" and
+ 02: "LLL_measure i fs' = LLL_measure i fs" and
+ 03: "\<And> i. i < m \<Longrightarrow> gso fs' i = gso fs i" and
+ 04: "\<And> i' j'. i' < m \<Longrightarrow> j' < m \<Longrightarrow>
+ \<mu> fs' i' j' = (if i' = i \<and> j' \<le> j then \<mu> fs i j' - of_int c * \<mu> fs j j' else \<mu> fs i' j')" and
+ 05: "\<And> ii. ii \<le> m \<Longrightarrow> d fs' ii = d fs ii" and
+ 06: "\<bar>\<mu> fs' i j\<bar> \<le> 1 / 2" and
+ 061: "(\<forall>i' j'. i' < i \<longrightarrow> j' \<le> i' \<longrightarrow> \<mu> fs i' j' = \<mu> fs' i' j')"
+ using basis_reduction_add_row_main[OF Linvww i j fs'_def] c i by auto
+ have 07: "lin_indep fs'" and
+ 08: "length fs' = m" and
+ 09: "lattice_of fs' = L" using LLL_inv_wD Linvw' by auto
+ have 091: "fs_int_indpt n fs'" using 07 using Gram_Schmidt_2.fs_int_indpt.intro by simp
+ define I where "I = {(i',j'). i' = i \<and> j' < j}"
+ have 10: "I \<subseteq> {(i',j'). i' < m \<and> j' < i'}" "(i,j)\<notin> I" "\<forall>j' \<ge> j. (i,j') \<notin> I" using I_def i j by auto
+ obtain fs'' where
+ 11: "lattice_of fs'' = L" and
+ 12: "map (map_vec (\<lambda> x. x symmod p)) fs'' = map (map_vec (\<lambda> x. x symmod p)) fs'" and
+ 13: "lin_indep fs''" and
+ 14: "length fs'' = m" and
+ 15: "(\<forall> k < m. gso fs'' k = gso fs' k)" and
+ 16: "(\<forall> k \<le> m. d fs'' k = d fs' k)" and
+ 17: "(\<forall> i' < m. \<forall> j' < m. d\<mu> fs'' i' j' =
+ (if (i',j') \<in> I then d\<mu> fs' i' j' symmod (p * d fs' j' * d fs' (Suc j')) else d\<mu> fs' i' j'))"
+ using mod_finite_set[OF 07 08 10(1) 09 pgtz] by blast
+ have 171: "(\<forall>i' j'. i' < i \<longrightarrow> j' \<le> i' \<longrightarrow> \<mu> fs'' i' j' = \<mu> fs' i' j')"
+ proof -
+ {
+ fix i' j'
+ assume i'j': "i' < i" "j' \<le> i'"
+ have "rat_of_int (d\<mu> fs'' i' j') = rat_of_int (d\<mu> fs' i' j')" using "17" I_def i i'j' by auto
+ then have "rat_of_int (int_of_rat (rat_of_int (d fs'' (Suc j')) * \<mu> fs'' i' j')) =
+ rat_of_int (int_of_rat (rat_of_int (d fs' (Suc j')) * \<mu> fs' i' j'))"
+ using d\<mu>_def i'j' j by auto
+ then have "rat_of_int (d fs'' (Suc j')) * \<mu> fs'' i' j' =
+ rat_of_int (d fs' (Suc j')) * \<mu> fs' i' j'"
+ by (smt "08" "091" "13" "14" d_def dual_order.strict_trans fs_int.d_def
+ fs_int_indpt.fs_int_mu_d_Z fs_int_indpt.intro i i'j'(1) i'j'(2) int_of_rat(2))
+ then have "\<mu> fs'' i' j' = \<mu> fs' i' j'" by (smt "16"
+ LLL_d_pos[OF Linvw'] Suc_leI int_of_rat(1)
+ dual_order.strict_trans fs'_def i i'j' j
+ le_neq_implies_less nonzero_mult_div_cancel_left of_int_hom.hom_zero)
+ }
+ then show ?thesis by simp
+ qed
+ then have 172: "(\<forall>i' j'. i' < i \<longrightarrow> j' \<le> i' \<longrightarrow> \<mu> fs'' i' j' = \<mu> fs i' j')" using 061 by simp (* goal *)
+ have 18: "LLL_measure i fs'' = LLL_measure i fs'" using 16 LLL_measure_def logD_def D_def by simp
+ have 19: "(\<forall>k < m. gso fs'' k = gso fs k)" using 03 15 by simp
+ have "\<forall>j' \<in> {j..(m-1)}. j' < m" using j i by auto
+ then have 20: "\<forall>j' \<in> {j..(m-1)}. d\<mu> fs'' i j' = d\<mu> fs' i j'"
+ using 10(3) 17 Suc_lessD less_trans_Suc by (meson atLeastAtMost_iff i)
+ have 21: "\<forall>j' \<in> {j..(m-1)}. \<mu> fs'' i j' = \<mu> fs' i j'"
+ proof -
+ {
+ fix j'
+ assume j': "j' \<in> {j..(m-1)}"
+ define \<mu>'' :: rat where "\<mu>'' = \<mu> fs'' i j'"
+ define \<mu>' :: rat where "\<mu>' = \<mu> fs' i j'"
+ have "rat_of_int (d\<mu> fs'' i j') = rat_of_int (d\<mu> fs' i j')" using 20 j' by simp
+ moreover have "j' < length fs'" using i j' 08 by auto
+ ultimately have "rat_of_int (d fs' (Suc j')) * gram_schmidt_fs.\<mu> n (map of_int_hom.vec_hom fs') i j'
+ = rat_of_int (d fs'' (Suc j')) * gram_schmidt_fs.\<mu> n (map of_int_hom.vec_hom fs'') i j'"
+ using 20 08 091 13 14 fs_int_indpt.d\<mu>_def fs_int.d_def fs_int_indpt.d\<mu> d\<mu>_def d_def i fs_int_indpt.intro j'
+ by metis
+ then have "rat_of_int (d fs' (Suc j')) * \<mu>'' = rat_of_int (d fs' (Suc j')) * \<mu>'"
+ using 16 i j' \<mu>'_def \<mu>''_def unfolding d\<mu>_def by auto
+ moreover have "0 < d fs' (Suc j')" using LLL_d_pos[OF Linvw', of "Suc j'"] i j' by auto
+ ultimately have "\<mu> fs'' i j' = \<mu> fs' i j'" using \<mu>'_def \<mu>''_def by simp
+ }
+ then show ?thesis by simp
+ qed
+ then have 22: "\<mu> fs'' i j = \<mu> fs' i j" using i j by simp
+ then have 23: "\<bar>\<mu> fs'' i j\<bar> \<le> 1 / 2" using 06 by simp (* goal *)
+ have 24: "LLL_measure i fs'' = LLL_measure i fs" using 02 18 by simp (* goal *)
+ have 25: "(\<forall> k \<le> m. d fs'' k = d fs k)" using 16 05 by simp (* goal *)
+ have 26: "(\<forall> k < m. gso fs'' k = gso fs k)" using 15 03 by simp (* goal *)
+ have 27: "\<mu>_small_row i fs (Suc j) \<Longrightarrow> \<mu>_small_row i fs'' j"
+ using 21 01 \<mu>_small_row_def i j c by auto (* goal *)
+ have 28: "length fs = m" "length mfs = m" using LLL_invD_modw[OF Linvmw] by auto
+ have 29: "map (map_vec (\<lambda>x. x symmod p)) fs = mfs" using assms LLL_invD_modw by simp
+ have 30: "\<And> i. i < m \<Longrightarrow> fs ! i \<in> carrier_vec n" "\<And> i. i < m \<Longrightarrow> mfs ! i \<in> carrier_vec n"
+ using LLL_invD_modw[OF Linvmw] by auto
+ have 31: "\<And> i. i < m \<Longrightarrow> fs' ! i \<in> carrier_vec n" using fs'_def 30(1)
+ using "08" "091" fs_int_indpt.f_carrier by blast
+ have 32: "\<And> i. i < m \<Longrightarrow> mfs' ! i \<in> carrier_vec n" unfolding mfs' using 30(2) 28(2)
+ by (metis (no_types, lifting) Suc_lessD j less_trans_Suc map_carrier_vec minus_carrier_vec
+ nth_list_update_eq nth_list_update_neq smult_closed)
+ have 33: "length mfs' = m" using 28(2) mfs' by simp (* invariant goal *)
+ then have 34: "map (map_vec (\<lambda>x. x symmod p)) fs' = mfs'"
+ proof -
+ {
+ fix i' j'
+ have j2: "j < m" using j i by auto
+ assume i': "i' < m"
+ assume j': "j' < n"
+ then have fsij: "(fs ! i' $ j') symmod p = mfs ! i' $ j'" using 30 i' j' 28 29 by fastforce
+ have "mfs' ! i $ j' = (mfs ! i $ j'- (c \<cdot>\<^sub>v mfs ! j) $ j') symmod p"
+ unfolding mfs' using 30(2) j' 28 j2
+ by (metis (no_types, lifting) carrier_vecD i index_map_vec(1) index_minus_vec(1)
+ index_minus_vec(2) index_smult_vec(2) nth_list_update_eq)
+ then have mfs'ij: "mfs' ! i $ j' = (mfs ! i $ j'- c * mfs ! j $ j') symmod p"
+ unfolding mfs' using 30(2) i' j' 28 j2 by fastforce
+ have "(fs' ! i' $ j') symmod p = mfs' ! i' $ j'"
+ proof(cases "i' = i")
+ case True
+ show ?thesis using fs'_def mfs' True 28 fsij
+ proof -
+ have "fs' ! i' $ j' = (fs ! i' - c \<cdot>\<^sub>v fs ! j) $ j'" using fs'_def True i' j' 28(1) by simp
+ also have "\<dots> = fs ! i' $ j' - (c \<cdot>\<^sub>v fs ! j) $ j'" using i' j' 30(1)
+ by (metis Suc_lessD carrier_vecD i index_minus_vec(1) index_smult_vec(2) j less_trans_Suc)
+ finally have "fs' ! i' $ j' = fs ! i' $ j' - (c \<cdot>\<^sub>v fs ! j) $ j'" by auto
+ then have "(fs' ! i' $ j') symmod p = (fs ! i' $ j' - (c \<cdot>\<^sub>v fs ! j) $ j') symmod p" by auto
+ also have "\<dots> = ((fs ! i' $ j') symmod p - ((c \<cdot>\<^sub>v fs ! j) $ j') symmod p) symmod p"
+ by (simp add: sym_mod_diff_eq)
+ also have "(c \<cdot>\<^sub>v fs ! j) $ j' = c * (fs ! j $ j')"
+ using i' j' True 28 30(1) j
+ by (metis Suc_lessD carrier_vecD index_smult_vec(1) less_trans_Suc)
+ also have "((fs ! i' $ j') symmod p - (c * (fs ! j $ j')) symmod p) symmod p =
+ ((fs ! i' $ j') symmod p - c * ((fs ! j $ j') symmod p)) symmod p"
+ using i' j' True 28 30(1) j by (metis sym_mod_diff_right_eq sym_mod_mult_right_eq)
+ also have "((fs ! j $ j') symmod p) = mfs ! j $ j'" using 30 i' j' 28 29 j2 by fastforce
+ also have "((fs ! i' $ j') symmod p - c * mfs ! j $ j') symmod p =
+ (mfs ! i' $ j' - c * mfs ! j $ j') symmod p" using fsij by simp
+ finally show ?thesis using mfs'ij by (simp add: True)
+ qed
+ next
+ case False
+ show ?thesis using fs'_def mfs' False 28 fsij by simp
+ qed
+ }
+ then have "\<forall>i' < m. (map_vec (\<lambda>x. x symmod p)) (fs' ! i') = mfs' ! i'"
+ using 31 32 33 08 by fastforce
+ then show ?thesis using 31 32 33 08 by (simp add: map_nth_eq_conv)
+ qed
+ then have 35: "map (map_vec (\<lambda>x. x symmod p)) fs'' = mfs'" using 12 by simp (* invariant req. *)
+ have 36: "lin_indep fs''" using 13 by simp (* invariant req. *)
+ have Linvw'': "LLL_invariant_weak fs''" using LLL_invariant_weak_def 11 13 14 by simp
+ have 39: "(\<forall>i' < m. \<forall>j' < i'. \<bar>d\<mu> fs'' i' j'\<bar> < p * d fs'' j' * d fs'' (Suc j'))" (* invariant req. *)
+ proof -
+ {
+ fix i' j'
+ assume i': "i' < m"
+ assume j': "j' < i'"
+ define pdd where "pdd = (p * d fs'' j' * d fs'' (Suc j'))"
+ then have pddgtz: "pdd > 0"
+ using pgtz j' LLL_d_pos[OF Linvw', of "Suc j'"] LLL_d_pos[OF Linvw', of j'] j' i' 16 by simp
+ have "\<bar>d\<mu> fs'' i' j'\<bar> < p * d fs'' j' * d fs'' (Suc j')"
+ proof(cases "i' = i")
+ case i'i: True
+ then show ?thesis
+ proof (cases "j' < j")
+ case True
+ then have eq'': "d\<mu> fs'' i' j' = d\<mu> fs' i' j' symmod (p * d fs'' j' * d fs'' (Suc j'))"
+ using 16 17 10 I_def True i' j' i'i by simp
+ have "0 < pdd" using pddgtz by simp
+ then show ?thesis unfolding eq'' unfolding pdd_def[symmetric] using sym_mod_abs by blast
+ next
+ case fls: False
+ then have "(i',j') \<notin> I" using I_def i'i by simp
+ then have dmufs''fs': "d\<mu> fs'' i' j' = d\<mu> fs' i' j'" using 17 i' j' by simp
+ show ?thesis
+ proof (cases "j' = j")
+ case True
+ define \<mu>'' where "\<mu>'' = \<mu> fs'' i' j'"
+ define d'' where "d'' = d fs'' (Suc j')"
+ have pge1: "p \<ge> 1" using pgtz by simp
+ have lh: "\<bar>\<mu>''\<bar> \<le> 1 / 2" using 23 True i'i \<mu>''_def by simp
+ moreover have eq: "d\<mu> fs'' i' j' = \<mu>'' * d''" using d\<mu>_def i' j' \<mu>''_def d''_def
+ by (smt "14" "36" LLL.d_def Suc_lessD fs_int.d_def fs_int_indpt.d\<mu> fs_int_indpt.intro
+ int_of_rat(1) less_trans_Suc mult_of_int_commute of_rat_mult of_rat_of_int_eq)
+ moreover have Sj': "Suc j' \<le> m" "j' \<le> m" using True j' i i' by auto
+ moreover then have gtz: "0 < d''" using LLL_d_pos[OF Linvw''] d''_def by simp
+ moreover have "rat_of_int \<bar>d\<mu> fs'' i' j'\<bar> = \<bar>\<mu>'' * (rat_of_int d'')\<bar>"
+ using eq by (metis of_int_abs of_rat_hom.injectivity of_rat_mult of_rat_of_int_eq)
+ moreover then have "\<bar>\<mu>'' * rat_of_int d'' \<bar> = \<bar>\<mu>''\<bar> * rat_of_int \<bar>d''\<bar>"
+ by (metis (mono_tags, opaque_lifting) abs_mult of_int_abs)
+ moreover have "\<dots> = \<bar>\<mu>''\<bar> * rat_of_int d'' " using gtz by simp
+ moreover have "\<dots> < rat_of_int d''" using lh gtz by simp
+ ultimately have "rat_of_int \<bar>d\<mu> fs'' i' j'\<bar> < rat_of_int d''" by simp
+ then have "\<bar>d\<mu> fs'' i' j'\<bar> < d fs'' (Suc j')" using d''_def by simp
+ then have "\<bar>d\<mu> fs'' i' j'\<bar> < p * d fs'' (Suc j')" using pge1
+ by (smt mult_less_cancel_right2)
+ then show ?thesis using pge1 LLL_d_pos[OF Linvw'' Sj'(2)] gtz unfolding d''_def
+ by (smt mult_less_cancel_left2 mult_right_less_imp_less)
+ next
+ case False
+ have "j' < m" using i' j' by simp
+ moreover have "j' > j" using False fls by simp
+ ultimately have "\<mu> fs' i' j' = \<mu> fs i' j'" using i' 04 i by simp
+ then have "d\<mu> fs' i' j' = d\<mu> fs i' j'" using d\<mu>_def i' j' 05 by simp
+ then have "d\<mu> fs'' i' j' = d\<mu> fs i' j'" using dmufs''fs' by simp
+ then show ?thesis using LLL_invD_modw[OF Linvmw] i' j' 25 by simp
+ qed
+ qed
+ next
+ case False
+ then have "(i',j') \<notin> I" using I_def by simp
+ then have dmufs''fs': "d\<mu> fs'' i' j' = d\<mu> fs' i' j'" using 17 i' j' by simp
+ have "\<mu> fs' i' j' = \<mu> fs i' j'" using i' 04 j' False by simp
+ then have "d\<mu> fs' i' j' = d\<mu> fs i' j'" using d\<mu>_def i' j' 05 by simp
+ moreover then have "d\<mu> fs'' i' j' = d\<mu> fs i' j'" using dmufs''fs' by simp
+ then show ?thesis using LLL_invD_modw[OF Linvmw] i' j' 25 by simp
+ qed
+ }
+ then show ?thesis by simp
+ qed
+ have 40: "(\<forall>i' < m. \<forall>j' < m. i' \<noteq> i \<or> j' > j \<longrightarrow> d\<mu> fs' i' j' = dmu $$ (i',j'))"
+ proof -
+ {
+ fix i' j'
+ assume i': "i' < m" and j': "j' < m"
+ assume assm: "i' \<noteq> i \<or> j' > j"
+ have "d\<mu> fs' i' j' = dmu $$ (i',j')"
+ proof (cases "i' \<noteq> i")
+ case True
+ then show ?thesis using fs'_def LLL_invD_modw[OF Linvmw] d\<mu>_def i i' j j'
+ 04 28(1) LLL_invI_weak basis_reduction_add_row_main(8)[OF Linvww] by auto
+ next
+ case False
+ then show ?thesis
+ using 05 LLL_invD_modw[OF Linvmw] d\<mu>_def i j j' 04 assm by simp
+ qed
+ }
+ then show ?thesis by simp
+ qed
+ have 41: "\<forall>j' \<le> j. d\<mu> fs' i j' = dmu $$ (i,j') - c * dmu $$ (j,j')"
+ proof -
+ {
+ let ?oi = "of_int :: _ \<Rightarrow> rat"
+ fix j'
+ assume j': "j' \<le> j"
+ define dj' \<mu>i \<mu>j where "dj' = d fs (Suc j')" and "\<mu>i = \<mu> fs i j'" and "\<mu>j = \<mu> fs j j'"
+ have "?oi (d\<mu> fs' i j') = ?oi (d fs (Suc j')) * (\<mu> fs i j' - ?oi c * \<mu> fs j j')"
+ using j' 04 d\<mu>_def
+ by (smt "05" "08" "091" Suc_leI d_def diff_diff_cancel fs_int.d_def
+ fs_int_indpt.fs_int_mu_d_Z i int_of_rat(2) j less_imp_diff_less less_imp_le_nat)
+ also have "\<dots> = (?oi dj') * (\<mu>i - of_int c * \<mu>j)"
+ using dj'_def \<mu>i_def \<mu>j_def by (simp add: of_rat_mult)
+ also have "\<dots> = (rat_of_int dj') * \<mu>i - of_int c * (rat_of_int dj') * \<mu>j" by algebra
+ also have "\<dots> = rat_of_int (d\<mu> fs i j') - ?oi c * rat_of_int (d\<mu> fs j j')" unfolding dj'_def \<mu>i_def \<mu>j_def
+ using i j j' d\<mu>_def
+ using "28"(1) LLL.LLL_invD_modw(4) Linvmw d_def fs_int.d_def fs_int_indpt.fs_int_mu_d_Z fs_int_indpt.intro by auto
+ also have "\<dots> = rat_of_int (dmu $$ (i,j')) - ?oi c * rat_of_int (dmu $$ (j,j'))"
+ using LLL_invD_modw(7)[OF Linvmw] d\<mu>_def j' i j by auto
+ finally have "?oi (d\<mu> fs' i j') = rat_of_int (dmu $$ (i,j')) - ?oi c * rat_of_int (dmu $$ (j,j'))" by simp
+ then have "d\<mu> fs' i j' = dmu $$ (i,j') - c * dmu $$ (j,j')"
+ using of_int_eq_iff by fastforce
+ }
+ then show ?thesis by simp
+ qed
+ have 42: "(\<forall>i' < m. \<forall>j' < m. d\<mu> fs'' i' j' = dmu' $$ (i',j'))"
+ proof -
+ {
+ fix i' j'
+ assume i': "i' < m" and j': "j' < m"
+ have "d\<mu> fs'' i' j' = dmu' $$ (i',j')"
+ proof (cases "i' = i")
+ case i'i: True
+ then show ?thesis
+ proof (cases "j' > j")
+ case True
+ then have "(i',j')\<notin>I" using I_def by simp
+ moreover then have "d\<mu> fs' i' j' = d\<mu> fs i' j'" using "04" "05" True Suc_leI d\<mu>_def i' j' by simp
+ moreover have "dmu' $$ (i',j') = dmu $$ (i',j')" using dmu' True i' j' by simp
+ ultimately show ?thesis using "17" "40" True i' j' by auto
+ next
+ case False
+ then have j'lej: "j' \<le> j" by simp
+ then have eq': "d\<mu> fs' i j' = dmu $$ (i,j') - c * dmu $$ (j,j')" using 41 by simp
+ have id: "d_of dmu j' = d fs j'" "d_of dmu (Suc j') = d fs (Suc j')"
+ using d_of_weak[OF Linvmw] \<open>j' < m\<close> by auto
+ show ?thesis
+ proof (cases "j' \<noteq> j")
+ case True
+ then have j'ltj: "j' < j" using True False by simp
+ then have "(i',j') \<in> I" using I_def True i'i by simp
+ then have "d\<mu> fs'' i' j' =
+ (dmu $$ (i,j') - c * dmu $$ (j,j')) symmod (p * d fs' j' * d fs' (Suc j'))"
+ using 17 i' 41 j'lej by (simp add: j' i'i)
+ also have "\<dots> = (dmu $$ (i,j') - c * dmu $$ (j,j')) symmod (p * d fs j' * d fs (Suc j'))"
+ using 05 i j'ltj j by simp
+ also have "\<dots> = dmu' $$ (i,j')"
+ unfolding dmu' index_mat(1)[OF \<open>i < m\<close> \<open>j' < m\<close>] split id using j'lej True by auto
+ finally show ?thesis using i'i by simp
+ next
+ case False
+ then have j'j: "j' = j" by simp
+ then have "d\<mu> fs'' i j' = d\<mu> fs' i j'" using 20 j' by simp
+ also have "\<dots> = dmu $$ (i,j') - c * dmu $$ (j,j')" using eq' by simp
+ also have "\<dots> = dmu' $$ (i,j')" using dmu' j'j i j' by simp
+ finally show ?thesis using i'i by simp
+ qed
+ qed
+ next
+ case False
+ then have "(i',j')\<notin>I" using I_def by simp
+ moreover then have "d\<mu> fs' i' j' = d\<mu> fs i' j'" by (simp add: "04" "05" False Suc_leI d\<mu>_def i' j')
+ moreover then have "dmu' $$ (i',j') = dmu $$ (i',j')" using dmu' False i' j' by simp
+ ultimately show ?thesis using "17" "40" False i' j' by auto
+ qed
+ }
+ then show ?thesis by simp
+ qed
+ from gbnd 26 have gbnd: "g_bnd_mode first b fs''" using g_bnd_mode_cong[of fs'' fs] by simp
+ {
+ assume Linv: "LLL_invariant_mod fs mfs dmu p first b i"
+ have Linvw: "LLL_invariant_weak' i fs" using Linv LLL_invD_mod LLL_invI_weak by simp
+ note Linvww = LLL_invw'_imp_w[OF Linvw]
+ have 00: "LLL_invariant_weak' i fs'" using Linvw basis_reduction_add_row_weak[OF Linvw i j fs'_def] by auto
+ have 37: "weakly_reduced fs'' i" using 15 LLL_invD_weak(8)[OF 00] gram_schmidt_fs.weakly_reduced_def
+ by (smt Suc_lessD i less_trans_Suc) (* invariant req. *)
+ have 38: "LLL_invariant_weak' i fs''"
+ using 00 11 14 36 37 i 31 12 LLL_invariant_weak'_def by blast
+ have "LLL_invariant_mod fs'' mfs' dmu' p first b i"
+ using LLL_invI_mod[OF 33 _ 14 11 13 37 35 39 42 p1 gbnd LLL_invD_mod(17)[OF Linv]] i by simp
+ }
+ moreover have "LLL_invariant_mod_weak fs'' mfs' dmu' p first b"
+ using LLL_invI_modw[OF 33 14 11 13 35 39 42 p1 gbnd LLL_invD_modw(15)[OF Linvmw]] by simp
+ ultimately show ?thesis using 27 23 24 25 26 172 by auto
+qed
+
+definition D_mod :: "int mat \<Rightarrow> nat" where "D_mod dmu = nat (\<Prod> i < m. d_of dmu i)"
+
+definition logD_mod :: "int mat \<Rightarrow> nat"
+ where "logD_mod dmu = (if \<alpha> = 4/3 then (D_mod dmu) else nat (floor (log (1 / of_rat reduction) (D_mod dmu))))"
+end
+
+locale fs_int'_mod =
+ fixes n m fs_init \<alpha> i fs mfs dmu p first b
+ assumes LLL_inv_mod: "LLL.LLL_invariant_mod n m fs_init \<alpha> fs mfs dmu p first b i"
+
+context LLL_with_assms
+begin
+
+lemma basis_reduction_swap_weak': assumes Linvw: "LLL_invariant_weak' i fs"
+ and i: "i < m"
+ and i0: "i \<noteq> 0"
+ and mu_F1_i: "\<bar>\<mu> fs i (i-1)\<bar> \<le> 1 / 2"
+ and norm_ineq: "sq_norm (gso fs (i - 1)) > \<alpha> * sq_norm (gso fs i)"
+ and fs'_def: "fs' = fs[i := fs ! (i - 1), i - 1 := fs ! i]"
+shows "LLL_invariant_weak' (i - 1) fs'"
+proof -
+ note inv = LLL_invD_weak[OF Linvw]
+ note invw = LLL_invw'_imp_w[OF Linvw]
+ note main = basis_reduction_swap_main[OF invw disjI2[OF mu_F1_i] i i0 norm_ineq fs'_def]
+ note inv' = LLL_inv_wD[OF main(1)]
+ from \<open>weakly_reduced fs i\<close> have "weakly_reduced fs (i - 1)"
+ unfolding gram_schmidt_fs.weakly_reduced_def by auto
+ also have "weakly_reduced fs (i - 1) = weakly_reduced fs' (i - 1)"
+ unfolding gram_schmidt_fs.weakly_reduced_def
+ by (intro all_cong, insert i0 i main(5), auto)
+ finally have red: "weakly_reduced fs' (i - 1)" .
+ show "LLL_invariant_weak' (i - 1) fs'" using i
+ by (intro LLL_invI_weak red inv', auto)
+qed
+
+lemma basis_reduction_add_row_done_weak:
+ assumes Linv: "LLL_invariant_weak' i fs"
+ and i: "i < m"
+ and mu_small: "\<mu>_small_row i fs 0"
+shows "\<mu>_small fs i"
+proof -
+ note inv = LLL_invD_weak[OF Linv]
+ from mu_small
+ have mu_small: "\<mu>_small fs i" unfolding \<mu>_small_row_def \<mu>_small_def by auto
+ show ?thesis
+ using i mu_small LLL_invI_weak[OF inv(3,6,7,9,1)] by auto
+qed
+
+lemma LLL_invariant_mod_to_weak_m_to_i: assumes
+ inv: "LLL_invariant_mod fs mfs dmu p first b m"
+ and i: "i \<le> m"
+shows "LLL_invariant_mod fs mfs dmu p first b i"
+ "LLL_invariant_weak' m fs"
+ "LLL_invariant_weak' i fs"
+proof -
+ show "LLL_invariant_mod fs mfs dmu p first b i"
+ proof -
+ have "LLL_invariant_weak' m fs" using LLL_invD_mod[OF inv] LLL_invI_weak by simp
+ then have "LLL_invariant_weak' i fs" using LLL_inv_weak_m_impl_i i by simp
+ then have "weakly_reduced fs i" using i LLL_invD_weak(8) by simp
+ then show ?thesis using LLL_invD_mod[OF inv] LLL_invI_mod i by simp
+ qed
+ then show fsinvwi: "LLL_invariant_weak' i fs" using LLL_invD_mod LLL_invI_weak by simp
+ show "LLL_invariant_weak' m fs" using LLL_invD_mod[OF inv] LLL_invI_weak by simp
+qed
+
+lemma basis_reduction_mod_swap_main:
+ assumes Linvmw: "LLL_invariant_mod_weak fs mfs dmu p first b"
+ and k: "k < m"
+ and k0: "k \<noteq> 0"
+ and mu_F1_i: "\<bar>\<mu> fs k (k-1)\<bar> \<le> 1 / 2"
+ and norm_ineq: "sq_norm (gso fs (k - 1)) > \<alpha> * sq_norm (gso fs k)"
+ and mfs'_def: "mfs' = mfs[k := mfs ! (k - 1), k - 1 := mfs ! k]"
+ and dmu'_def: "dmu' = (mat m m (\<lambda>(i,j). (
+ if j < i then
+ if i = k - 1 then
+ dmu $$ (k, j)
+ else if i = k \<and> j \<noteq> k - 1 then
+ dmu $$ (k - 1, j)
+ else if i > k \<and> j = k then
+ ((d_of dmu (Suc k)) * dmu $$ (i, k - 1) - dmu $$ (k, k - 1) * dmu $$ (i, j))
+ div (d_of dmu k)
+ else if i > k \<and> j = k - 1 then
+ (dmu $$ (k, k - 1) * dmu $$ (i, j) + dmu $$ (i, k) * (d_of dmu (k-1)))
+ div (d_of dmu k)
+ else dmu $$ (i, j)
+ else if i = j then
+ if i = k - 1 then
+ ((d_of dmu (Suc k)) * (d_of dmu (k-1)) + dmu $$ (k, k - 1) * dmu $$ (k, k - 1))
+ div (d_of dmu k)
+ else (d_of dmu (Suc i))
+ else dmu $$ (i, j))
+ ))"
+ and dmu'_mod_def: "dmu'_mod = mat m m (\<lambda>(i, j). (
+ if j < i \<and> (j = k \<or> j = k - 1) then
+ dmu' $$ (i, j) symmod (p * (d_of dmu' j) * (d_of dmu' (Suc j)))
+ else dmu' $$ (i, j)))"
+shows "(\<exists>fs'. LLL_invariant_mod_weak fs' mfs' dmu'_mod p first b \<and>
+ LLL_measure (k-1) fs' < LLL_measure k fs \<and>
+ (LLL_invariant_mod fs mfs dmu p first b k \<longrightarrow> LLL_invariant_mod fs' mfs' dmu'_mod p first b (k-1)))"
+proof -
+ define fs' where "fs' = fs[k := fs ! (k - 1), k - 1 := fs ! k]"
+ have pgtz: "p > 0" and p1: "p > 1" using LLL_invD_modw[OF Linvmw] by auto
+ have invw: "LLL_invariant_weak fs" using LLL_invD_modw[OF Linvmw] LLL_invariant_weak_def by simp
+ note swap_main = basis_reduction_swap_main(3-)[OF invw disjI2[OF mu_F1_i] k k0 norm_ineq fs'_def]
+ note dd\<mu>_swap = d_d\<mu>_swap[OF invw disjI2[OF mu_F1_i] k k0 norm_ineq fs'_def]
+ have invw': "LLL_invariant_weak fs'" using fs'_def assms invw basis_reduction_swap_main(1) by simp
+ have 02: "LLL_measure k fs > LLL_measure (k - 1) fs'" by fact
+ have 03: "\<And> i j. i < m \<Longrightarrow> j < i \<Longrightarrow>
+ d\<mu> fs' i j = (
+ if i = k - 1 then
+ d\<mu> fs k j
+ else if i = k \<and> j \<noteq> k - 1 then
+ d\<mu> fs (k - 1) j
+ else if i > k \<and> j = k then
+ (d fs (Suc k) * d\<mu> fs i (k - 1) - d\<mu> fs k (k - 1) * d\<mu> fs i j) div d fs k
+ else if i > k \<and> j = k - 1 then
+ (d\<mu> fs k (k - 1) * d\<mu> fs i j + d\<mu> fs i k * d fs (k - 1)) div d fs k
+ else d\<mu> fs i j)"
+ using dd\<mu>_swap by auto
+ have 031: "\<And>i. i < k-1 \<Longrightarrow> gso fs' i = gso fs i"
+ using swap_main(2) k k0 by auto
+ have 032: "\<And> ii. ii \<le> m \<Longrightarrow> of_int (d fs' ii) = (if ii = k then
+ sq_norm (gso fs' (k - 1)) / sq_norm (gso fs (k - 1)) * of_int (d fs k)
+ else of_int (d fs ii))"
+ by fact
+ have gbnd: "g_bnd_mode first b fs'"
+ proof (cases "first \<and> m \<noteq> 0")
+ case True
+ have "sq_norm (gso fs' 0) \<le> sq_norm (gso fs 0)"
+ proof (cases "k - 1 = 0")
+ case False
+ thus ?thesis using 031[of 0] by simp
+ next
+ case *: True
+ have k_1: "k - 1 < m" using k by auto
+ from * k0 have k1: "k = 1" by simp
+ (* this is a copy of what is done in LLL.swap-main, should be made accessible in swap-main *)
+ have "sq_norm (gso fs' 0) \<le> abs (sq_norm (gso fs' 0))" by simp
+ also have "\<dots> = abs (sq_norm (gso fs 1) + \<mu> fs 1 0 * \<mu> fs 1 0 * sq_norm (gso fs 0))"
+ by (subst swap_main(3)[OF k_1, unfolded *], auto simp: k1)
+ also have "\<dots> \<le> sq_norm (gso fs 1) + abs (\<mu> fs 1 0) * abs (\<mu> fs 1 0) * sq_norm (gso fs 0)"
+ by (simp add: sq_norm_vec_ge_0)
+ also have "\<dots> \<le> sq_norm (gso fs 1) + (1 / 2) * (1 / 2) * sq_norm (gso fs 0)"
+ using mu_F1_i[unfolded k1]
+ by (intro plus_right_mono mult_mono, auto)
+ also have "\<dots> < 1 / \<alpha> * sq_norm (gso fs 0) + (1 / 2) * (1 / 2) * sq_norm (gso fs 0)"
+ by (intro add_strict_right_mono, insert norm_ineq[unfolded mult.commute[of \<alpha>],
+ THEN mult_imp_less_div_pos[OF \<alpha>0(1)]] k1, auto)
+ also have "\<dots> = reduction * sq_norm (gso fs 0)" unfolding reduction_def
+ using \<alpha>0 by (simp add: ring_distribs add_divide_distrib)
+ also have "\<dots> \<le> 1 * sq_norm (gso fs 0)" using reduction(2)
+ by (intro mult_right_mono, auto)
+ finally show ?thesis by simp
+ qed
+ thus ?thesis using LLL_invD_modw(14)[OF Linvmw] True
+ unfolding g_bnd_mode_def by auto
+ next
+ case False
+ from LLL_invD_modw(14)[OF Linvmw] False have "g_bnd b fs" unfolding g_bnd_mode_def by auto
+ hence "g_bnd b fs'" using g_bnd_swap[OF k k0 invw mu_F1_i norm_ineq fs'_def] by simp
+ thus ?thesis using False unfolding g_bnd_mode_def by auto
+ qed
+ note d_of = d_of_weak[OF Linvmw]
+ have 033: "\<And> i. i < m \<Longrightarrow> d\<mu> fs' i i = (
+ if i = k - 1 then
+ ((d_of dmu (Suc k)) * (d_of dmu (k-1)) + dmu $$ (k, k - 1) * dmu $$ (k, k - 1))
+ div (d_of dmu k)
+ else (d_of dmu (Suc i)))"
+ proof -
+ fix i
+ assume i: "i < m"
+ have "d\<mu> fs' i i = d fs' (Suc i)" using dd\<mu> i by simp
+ also have "\<dots> = (if i = k - 1 then
+ (d fs (Suc k) * d fs (k - 1) + d\<mu> fs k (k - 1) * d\<mu> fs k (k - 1)) div d fs k
+ else d fs (Suc i))"
+ by (subst dd\<mu>_swap, insert dd\<mu> k0 i, auto)
+ also have "\<dots> = (if i = k - 1 then
+ ((d_of dmu (Suc k)) * (d_of dmu (k-1)) + dmu $$ (k, k - 1) * dmu $$ (k, k - 1))
+ div (d_of dmu k)
+ else (d_of dmu (Suc i)))" (is "_ = ?r")
+ using d_of i k LLL_invD_modw(7)[OF Linvmw] by auto
+ finally show "d\<mu> fs' i i = ?r" .
+ qed
+ have 04: "lin_indep fs'" "length fs' = m" "lattice_of fs' = L" using LLL_inv_wD[OF invw'] by auto
+ define I where "I = {(i, j). i < m \<and> j < i \<and> (j = k \<or> j = k - 1)}"
+ then have Isubs: "I \<subseteq> {(i,j). i < m \<and> j < i}" using k k0 by auto
+ obtain fs'' where
+ 05: "lattice_of fs'' = L" and
+ 06: "map (map_vec (\<lambda> x. x symmod p)) fs'' = map (map_vec (\<lambda> x. x symmod p)) fs'" and
+ 07: "lin_indep fs''" and
+ 08: "length fs'' = m" and
+ 09: "(\<forall> k < m. gso fs'' k = gso fs' k)" and
+ 10: "(\<forall> k \<le> m. d fs'' k = d fs' k)" and
+ 11: "(\<forall> i' < m. \<forall> j' < m. d\<mu> fs'' i' j' =
+ (if (i',j') \<in> I then d\<mu> fs' i' j' symmod (p * d fs' j' * d fs' (Suc j')) else d\<mu> fs' i' j'))"
+ using mod_finite_set[OF 04(1) 04(2) Isubs 04(3) pgtz] by blast
+ have 13: "length mfs' = m" using mfs'_def LLL_invD_modw(1)[OF Linvmw] by simp (* invariant requirement *)
+ have 14: "map (map_vec (\<lambda> x. x symmod p)) fs'' = mfs'" (* invariant requirement *)
+ using 06 fs'_def k k0 04(2) LLL_invD_modw(5)[OF Linvmw]
+ by (metis (no_types, lifting) length_list_update less_imp_diff_less map_update mfs'_def nth_map)
+ have "LLL_measure (k - 1) fs'' = LLL_measure (k - 1) fs'" using 10 LLL_measure_def logD_def D_def by simp
+ then have 15: "LLL_measure (k - 1) fs'' < LLL_measure k fs" using 02 by simp (* goal *)
+ {
+ fix i' j'
+ assume i'j': "i'<m" "j'<i'"
+ and neq: "j' \<noteq> k" "j' \<noteq> k - 1"
+ hence j'k: "j' \<noteq> k" "Suc j' \<noteq> k" using k0 by auto
+ hence "d fs'' j' = d fs j'" "d fs'' (Suc j') = d fs (Suc j')"
+ using \<open>k < m\<close> i'j' k0
+ 10[rule_format, of j'] 032[rule_format, of j']
+ 10[rule_format, of "Suc j'"] 032[rule_format, of "Suc j'"]
+ by auto
+ } note d_id = this
+
+ have 16: "\<forall>i'<m. \<forall>j'<i'. \<bar>d\<mu> fs'' i' j'\<bar> < p * d fs'' j' * d fs'' (Suc j')" (* invariant requirement *)
+ proof -
+ {
+ fix i' j'
+ assume i'j': "i'<m" "j'<i'"
+ have "\<bar>d\<mu> fs'' i' j'\<bar> < p * d fs'' j' * d fs'' (Suc j')"
+ proof (cases "(i',j') \<in> I")
+ case True
+ define pdd where "pdd = (p * d fs' j' * d fs' (Suc j'))"
+ have pdd_pos: "pdd > 0" using pgtz i'j' LLL_d_pos[OF invw'] pdd_def by simp
+ have "d\<mu> fs'' i' j' = d\<mu> fs' i' j' symmod pdd" using True 11 i'j' pdd_def by simp
+ then have "\<bar>d\<mu> fs'' i' j'\<bar> < pdd" using True 11 i'j' pdd_pos sym_mod_abs by simp
+ then show ?thesis unfolding pdd_def using 10 i'j' by simp
+ next
+ case False
+ from False[unfolded I_def] i'j' have neg: "j' \<noteq> k" "j' \<noteq> k - 1" by auto
+
+ consider (1) "i' = k - 1 \<or> i' = k" | (2) "\<not> (i' = k - 1 \<or> i' = k)"
+ using False i'j' unfolding I_def by linarith
+ thus ?thesis
+ proof cases
+ case **: 1
+ let ?i'' = "if i' = k - 1 then k else k -1"
+ from ** neg i'j' have i'': "?i'' < m" "j' < ?i''" using k0 k by auto
+ have "d\<mu> fs'' i' j' = d\<mu> fs' i' j'" using 11 False i'j' by simp
+ also have "\<dots> = d\<mu> fs ?i'' j'" unfolding 03[OF \<open>i' < m\<close> \<open>j' < i'\<close>]
+ using ** neg by auto
+ finally show ?thesis using LLL_invD_modw(6)[OF Linvmw, rule_format, OF i''] unfolding d_id[OF i'j' neg] by auto
+ next
+ case **: 2
+ hence neq: "j' \<noteq> k" "j' \<noteq> k - 1" using False k k0 i'j' unfolding I_def by auto
+ have "d\<mu> fs'' i' j' = d\<mu> fs' i' j'" using 11 False i'j' by simp
+ also have "\<dots> = d\<mu> fs i' j'" unfolding 03[OF \<open>i' < m\<close> \<open>j' < i'\<close>] using ** neq by auto
+ finally show ?thesis using LLL_invD_modw(6)[OF Linvmw, rule_format, OF i'j'] using d_id[OF i'j' neq] by auto
+ qed
+ qed
+ }
+ then show ?thesis by simp
+ qed
+ have 17: "\<forall>i'<m. \<forall>j'<m. d\<mu> fs'' i' j' = dmu'_mod $$ (i', j')" (* invariant requirement *)
+ proof -
+ {
+ fix i' j'
+ assume i'j': "i'<m" "j'<i'"
+ have d'dmu': "\<forall>j' < m. d fs' (Suc j') = dmu' $$ (j', j')" using dd\<mu> dmu'_def 033 by simp
+ have eq': "d\<mu> fs' i' j' = dmu' $$ (i', j')"
+ proof -
+ have t00: "d\<mu> fs k j' = dmu $$ (k, j')" and
+ t01: "d\<mu> fs (k - 1) j' = dmu $$ (k - 1, j')" and
+ t04: "d\<mu> fs k (k - 1) = dmu $$ (k, k - 1)" and
+ t05: "d\<mu> fs i' k = dmu $$ (i', k)"
+ using LLL_invD_modw(7)[OF Linvmw] i'j' k dd\<mu> k0 by auto
+ have t03: "d fs k = d\<mu> fs (k-1) (k-1)" using k0 k by (metis LLL.dd\<mu> Suc_diff_1 lessI not_gr_zero)
+ have t06: "d fs (k - 1) = (d_of dmu (k-1))" using d_of k by auto
+ have t07: "d fs k = (d_of dmu k)" using d_of k by auto
+ have j': "j' < m" using i'j' by simp
+ have "d\<mu> fs' i' j' = (if i' = k - 1 then
+ dmu $$ (k, j')
+ else if i' = k \<and> j' \<noteq> k - 1 then
+ dmu $$ (k - 1, j')
+ else if i' > k \<and> j' = k then
+ (dmu $$ (k, k) * dmu $$ (i', k - 1) - dmu $$ (k, k - 1) * dmu $$ (i', j')) div (d_of dmu k)
+ else if i' > k \<and> j' = k - 1 then
+ (dmu $$ (k, k - 1) * dmu $$ (i', j') + dmu $$ (i', k) * d fs (k - 1)) div (d_of dmu k)
+ else dmu $$ (i', j'))"
+ using dd\<mu> k t00 t01 t03 LLL_invD_modw(7)[OF Linvmw] k i'j' j' 03 t07 by simp
+ then show ?thesis using dmu'_def i'j' j' t06 t07 by (simp add: d_of_def)
+ qed
+ have "d\<mu> fs'' i' j' = dmu'_mod $$ (i', j')"
+ proof (cases "(i',j') \<in> I")
+ case i'j'I: True
+ have j': "j' < m" using i'j' by simp
+ show ?thesis
+ proof -
+ have "dmu'_mod $$ (i',j') = dmu' $$ (i',j')
+ symmod (p * (d_of dmu' j') * (d_of dmu' (Suc j')))"
+ using dmu'_mod_def i'j' i'j'I I_def by simp
+ also have "d_of dmu' j' = d fs' j'"
+ using j' d'dmu' d_def Suc_diff_1 less_imp_diff_less unfolding d_of_def
+ by (cases j', auto)
+ finally have "dmu'_mod $$ (i',j') = dmu' $$ (i',j') symmod (p * d fs' j' * d fs' (Suc j'))"
+ using dd\<mu>[OF j'] d'dmu' j' by (auto simp: d_of_def)
+ then show ?thesis using i'j'I 11 i'j' eq' by simp
+ qed
+ next
+ case False
+ have "d\<mu> fs'' i' j' = d\<mu> fs' i' j'" using False 11 i'j' by simp
+ also have "\<dots> = dmu' $$ (i', j')" unfolding eq' ..
+ finally show ?thesis unfolding dmu'_mod_def using False[unfolded I_def] i'j' by auto
+ qed
+ }
+ moreover have "\<forall>i' j'. i' < m \<longrightarrow> j' < m \<longrightarrow> i' = j' \<longrightarrow> d\<mu> fs'' i' j' = dmu'_mod $$ (i', j')"
+ using dd\<mu> dmu'_def 033 10 dmu'_mod_def 11 I_def by simp
+ moreover {
+ fix i' j'
+ assume i'j'': "i' < m" "j' < m" "i' < j'"
+ then have \<mu>z: "\<mu> fs'' i' j' = 0" by (simp add: gram_schmidt_fs.\<mu>.simps)
+ have "dmu'_mod $$ (i',j') = dmu' $$ (i',j')" using dmu'_mod_def i'j'' by auto
+ also have "\<dots> = d\<mu> fs i' j'" using LLL_invD_modw(7)[OF Linvmw] i'j'' dmu'_def by simp
+ also have "\<dots> = 0" using d\<mu>_def i'j'' by (simp add: gram_schmidt_fs.\<mu>.simps)
+ finally have "d\<mu> fs'' i' j' = dmu'_mod $$ (i',j')" using \<mu>z d_def i'j'' d\<mu>_def by simp
+ }
+ ultimately show ?thesis by (meson nat_neq_iff)
+ qed
+ from gbnd 09 have g_bnd: "g_bnd_mode first b fs''" using g_bnd_mode_cong[of fs' fs''] by auto
+ {
+ assume Linv: "LLL_invariant_mod fs mfs dmu p first b k"
+ have 00: "LLL_invariant_weak' k fs" using LLL_invD_mod[OF Linv] LLL_invI_weak by simp
+ note swap_weak' = basis_reduction_swap_weak'[OF 00 k k0 mu_F1_i norm_ineq fs'_def]
+ have 01: "LLL_invariant_weak' (k - 1) fs'" by fact
+ have 12: "weakly_reduced fs'' (k-1)" (* invariant requirement *)
+ using 031 09 k LLL_invD_weak(8)[OF 00] unfolding gram_schmidt_fs.weakly_reduced_def by simp
+ have "LLL_invariant_mod fs'' mfs' dmu'_mod p first b (k-1)"
+ using LLL_invI_mod[OF 13 _ 08 05 07 12 14 16 17 p1 g_bnd LLL_invD_mod(17)[OF Linv]] k by simp
+ }
+ moreover have "LLL_invariant_mod_weak fs'' mfs' dmu'_mod p first b"
+ using LLL_invI_modw[OF 13 08 05 07 14 16 17 p1 g_bnd LLL_invD_modw(15)[OF Linvmw]] by simp
+ ultimately show ?thesis using 15 by auto
+qed
+
+lemma dmu_quot_is_round_of_\<mu>:
+ assumes Linv: "LLL_invariant_mod fs mfs dmu p first b i'"
+ and c: "c = round_num_denom (dmu $$ (i,j)) (d_of dmu (Suc j))"
+ and i: "i < m"
+ and j: "j < i"
+ shows "c = round(\<mu> fs i j)"
+proof -
+ have Linvw: "LLL_invariant_weak' i' fs" using LLL_invD_mod[OF Linv] LLL_invI_weak by simp
+ have j2: "j < m" using i j by simp
+ then have j3: "Suc j \<le> m" by simp
+ have \<mu>1: "\<mu> fs j j = 1" using i j by (meson gram_schmidt_fs.\<mu>.elims less_irrefl_nat)
+ have inZ: "rat_of_int (d fs (Suc j)) * \<mu> fs i j \<in> \<int>" using fs_int_indpt.fs_int_mu_d_Z_m_m i j
+ LLL_invD_mod(5)[OF Linv] LLL_invD_weak(2) Linvw d_def fs_int.d_def fs_int_indpt.intro by auto
+ have "c = round(rat_of_int (d\<mu> fs i j) / rat_of_int (d\<mu> fs j j))" using LLL_invD_mod(9) Linv i j c
+ by (simp add: round_num_denom d_of_def)
+ then show ?thesis using LLL_d_pos[OF LLL_invw'_imp_w[OF Linvw] j3] j i inZ d\<mu>_def \<mu>1 by simp
+qed
+
+lemma dmu_quot_is_round_of_\<mu>_weak:
+ assumes Linv: "LLL_invariant_mod_weak fs mfs dmu p first b"
+ and c: "c = round_num_denom (dmu $$ (i,j)) (d_of dmu (Suc j))"
+ and i: "i < m"
+ and j: "j < i"
+ shows "c = round(\<mu> fs i j)"
+proof -
+ have Linvww: "LLL_invariant_weak fs" using LLL_invD_modw[OF Linv] LLL_invariant_weak_def by simp
+ have j2: "j < m" using i j by simp
+ then have j3: "Suc j \<le> m" by simp
+ have \<mu>1: "\<mu> fs j j = 1" using i j by (meson gram_schmidt_fs.\<mu>.elims less_irrefl_nat)
+ have inZ: "rat_of_int (d fs (Suc j)) * \<mu> fs i j \<in> \<int>" using fs_int_indpt.fs_int_mu_d_Z_m_m i j
+ LLL_invD_modw[OF Linv] d_def fs_int.d_def fs_int_indpt.intro by auto
+ have "c = round(rat_of_int (d\<mu> fs i j) / rat_of_int (d\<mu> fs j j))" using LLL_invD_modw(7) Linv i j c
+ by (simp add: round_num_denom d_of_def)
+ then show ?thesis using LLL_d_pos[OF Linvww j3] j i inZ d\<mu>_def \<mu>1 by simp
+qed
+
+lemma basis_reduction_mod_add_row: assumes
+ Linv: "LLL_invariant_mod_weak fs mfs dmu p first b"
+ and res: "basis_reduction_mod_add_row p mfs dmu i j = (mfs', dmu')"
+ and i: "i < m"
+ and j: "j < i"
+ and igtz: "i \<noteq> 0"
+shows "(\<exists>fs'. LLL_invariant_mod_weak fs' mfs' dmu' p first b \<and>
+ LLL_measure i fs' = LLL_measure i fs \<and>
+ (\<mu>_small_row i fs (Suc j) \<longrightarrow> \<mu>_small_row i fs' j) \<and>
+ \<bar>\<mu> fs' i j\<bar> \<le> 1 / 2 \<and>
+ (\<forall>i' j'. i' < i \<longrightarrow> j' \<le> i' \<longrightarrow> \<mu> fs' i' j' = \<mu> fs i' j') \<and>
+ (LLL_invariant_mod fs mfs dmu p first b i \<longrightarrow> LLL_invariant_mod fs' mfs' dmu' p first b i) \<and>
+ (\<forall>ii \<le> m. d fs' ii = d fs ii))"
+proof -
+ define c where "c = round_num_denom (dmu $$ (i,j)) (d_of dmu (Suc j))"
+ then have c: "c = round(\<mu> fs i j)" using dmu_quot_is_round_of_\<mu>_weak[OF Linv c_def i j] by simp
+ show ?thesis
+ proof (cases "c = 0")
+ case True
+ then have pair_id: "(mfs', dmu') = (mfs, dmu)"
+ using res c_def unfolding basis_reduction_mod_add_row_def Let_def by auto
+ moreover have "\<bar>\<mu> fs i j\<bar> \<le> inverse 2" using c[symmetric, unfolded True]
+ by (simp add: round_def, linarith)
+ moreover then have "(\<mu>_small_row i fs (Suc j) \<longrightarrow> \<mu>_small_row i fs j)"
+ unfolding \<mu>_small_row_def using Suc_leI le_neq_implies_less by blast
+ ultimately show ?thesis using Linv pair_id by auto
+ next
+ case False
+ then have pair_id: "(mfs', dmu') = (mfs[i := map_vec (\<lambda>x. x symmod p) (mfs ! i - c \<cdot>\<^sub>v mfs ! j)],
+ mat m m (\<lambda>(i', j'). if i' = i \<and> j' \<le> j
+ then if j' = j then dmu $$ (i, j') - c * dmu $$ (j, j')
+ else (dmu $$ (i,j') - c * dmu $$ (j,j'))
+ symmod (p * (d_of dmu j') * (d_of dmu (Suc j')))
+ else dmu $$ (i', j')))"
+ using res c_def unfolding basis_reduction_mod_add_row_def Let_def by auto
+ then have mfs': "mfs' = mfs[i := map_vec (\<lambda>x. x symmod p) (mfs ! i - c \<cdot>\<^sub>v mfs ! j)]"
+ and dmu': "dmu' = mat m m (\<lambda>(i', j'). if i' = i \<and> j' \<le> j
+ then if j' = j then dmu $$ (i, j') - c * dmu $$ (j, j')
+ else (dmu $$ (i,j') - c * dmu $$ (j,j'))
+ symmod (p * (d_of dmu j') * (d_of dmu (Suc j')))
+ else dmu $$ (i', j'))" by auto
+ show ?thesis using basis_reduction_mod_add_row_main[OF Linv i j c mfs' dmu'] by blast
+ qed
+qed
+
+lemma basis_reduction_mod_swap: assumes
+ Linv: "LLL_invariant_mod_weak fs mfs dmu p first b"
+ and mu: "\<bar>\<mu> fs k (k-1)\<bar> \<le> 1 / 2"
+ and res: "basis_reduction_mod_swap p mfs dmu k = (mfs', dmu'_mod)"
+ and cond: "sq_norm (gso fs (k - 1)) > \<alpha> * sq_norm (gso fs k)"
+ and i: "k < m" "k \<noteq> 0"
+shows "(\<exists>fs'. LLL_invariant_mod_weak fs' mfs' dmu'_mod p first b \<and>
+ LLL_measure (k - 1) fs' < LLL_measure k fs \<and>
+ (LLL_invariant_mod fs mfs dmu p first b k \<longrightarrow> LLL_invariant_mod fs' mfs' dmu'_mod p first b (k-1)))"
+ using res[unfolded basis_reduction_mod_swap_def basis_reduction_mod_swap_dmu_mod_def]
+ basis_reduction_mod_swap_main[OF Linv i mu cond] by blast
+
+lemma basis_reduction_adjust_mod: assumes
+ Linv: "LLL_invariant_mod_weak fs mfs dmu p first b"
+ and res: "basis_reduction_adjust_mod p first mfs dmu = (p', mfs', dmu', g_idx')"
+shows "(\<exists>fs' b'. (LLL_invariant_mod fs mfs dmu p first b i \<longrightarrow> LLL_invariant_mod fs' mfs' dmu' p' first b' i) \<and>
+ LLL_invariant_mod_weak fs' mfs' dmu' p' first b' \<and>
+ LLL_measure i fs' = LLL_measure i fs)"
+proof (cases "\<exists> g_idx. basis_reduction_adjust_mod p first mfs dmu = (p, mfs, dmu, g_idx)")
+ case True
+ thus ?thesis using res Linv by auto
+next
+ case False
+ obtain b' g_idx where norm: "compute_max_gso_norm first dmu = (b', g_idx)" by force
+ define p'' where "p'' = compute_mod_of_max_gso_norm first b'"
+ define d_vec where "d_vec = vec (Suc m) (\<lambda>i. d_of dmu i)"
+ define mfs'' where "mfs'' = map (map_vec (\<lambda>x. x symmod p'')) mfs"
+ define dmu'' where "dmu'' = mat m m (\<lambda>(i, j).
+ if j < i then dmu $$ (i, j) symmod (p'' * d_vec $ j * d_vec $ Suc j)
+ else dmu $$ (i, j))"
+ note res = res False
+ note res = res[unfolded basis_reduction_adjust_mod.simps Let_def norm split,
+ folded p''_def, folded d_vec_def mfs''_def, folded dmu''_def]
+ from res have pp': "p'' < p" and id: "dmu' = dmu''" "mfs' = mfs''" "p' = p''" "g_idx' = g_idx"
+ by (auto split: if_splits)
+ define I where "I = {(i',j'). i' < m \<and> j' < i'}"
+ note inv = LLL_invD_modw[OF Linv]
+ from inv(4) have lin: "gs.lin_indpt_list (RAT fs)" .
+ from inv(3) have lat: "lattice_of fs = L" .
+ from inv(2) have len: "length fs = m" .
+ have weak: "LLL_invariant_weak fs" using Linv
+ by (auto simp: LLL_invariant_mod_weak_def LLL_invariant_weak_def)
+ from compute_max_gso_norm[OF _ weak, of dmu first, unfolded norm] inv(7)
+ have bnd: "g_bnd_mode first b' fs" and b': "b' \<ge> 0" "m = 0 \<Longrightarrow> b' = 0" by auto
+ from compute_mod_of_max_gso_norm[OF b' p''_def]
+ have p'': "0 < p''" "1 < p''" "mod_invariant b' p'' first"
+ by auto
+ obtain fs' where
+ 01: "lattice_of fs' = L" and
+ 02: "map (map_vec (\<lambda> x. x symmod p'')) fs' = map (map_vec (\<lambda> x. x symmod p'')) fs" and
+ 03: "lin_indep fs'" and
+ 04: "length fs' = m" and
+ 05: "(\<forall> k < m. gso fs' k = gso fs k)" and
+ 06: "(\<forall> k \<le> m. d fs' k = d fs k)" and
+ 07: "(\<forall> i' < m. \<forall> j' < m. d\<mu> fs' i' j' =
+ (if (i',j') \<in> I then d\<mu> fs i' j' symmod (p'' * d fs j' * d fs (Suc j')) else d\<mu> fs i' j'))"
+ using mod_finite_set[OF lin len _ lat, of I] I_def p'' by blast
+ from bnd 05 have bnd: "g_bnd_mode first b' fs'" using g_bnd_mode_cong[of fs fs'] by auto
+ have D: "D fs = D fs'" unfolding D_def using 06 by auto
+
+
+ have Linv': "LLL_invariant_mod_weak fs' mfs'' dmu'' p'' first b'"
+ proof (intro LLL_invI_modw p'' 04 03 01 bnd)
+ {
+ have "mfs'' = map (map_vec (\<lambda>x. x symmod p'')) mfs" by fact
+ also have "\<dots> = map (map_vec (\<lambda>x. x symmod p'')) (map (map_vec (\<lambda>x. x symmod p)) fs)"
+ using inv by simp
+ also have "\<dots> = map (map_vec (\<lambda>x. x symmod p symmod p'')) fs" by auto
+ also have "(\<lambda> x. x symmod p symmod p'') = (\<lambda> x. x symmod p'')"
+ proof (intro ext)
+ fix x
+ from \<open>mod_invariant b p first\<close>[unfolded mod_invariant_def] obtain e where
+ p: "p = log_base ^ e" by auto
+ from p''[unfolded mod_invariant_def] obtain e' where
+ p'': "p'' = log_base ^ e'" by auto
+ from pp'[unfolded p p''] log_base have "e' \<le> e" by simp
+ hence dvd: "p'' dvd p" unfolding p p'' using log_base by (metis le_imp_power_dvd)
+ thus "x symmod p symmod p'' = x symmod p''"
+ by (intro sym_mod_sym_mod_cancel)
+ qed
+ finally show "map (map_vec (\<lambda>x. x symmod p'')) fs' = mfs''" unfolding 02 ..
+ }
+ thus "length mfs'' = m" using 04 by auto
+ show "\<forall>i'<m. \<forall>j'<i'. \<bar>d\<mu> fs' i' j'\<bar> < p'' * d fs' j' * d fs' (Suc j')"
+ proof -
+ {
+ fix i' j'
+ assume i'j': "i' < m" "j' < i'"
+ then have "d\<mu> fs' i' j' = d\<mu> fs i' j' symmod (p'' * d fs' j' * d fs' (Suc j'))"
+ using 07 06 unfolding I_def by simp
+ then have "\<bar>d\<mu> fs' i' j'\<bar> < p'' * d fs' j' * d fs' (Suc j')"
+ using sym_mod_abs p'' LLL_d_pos[OF weak] mult_pos_pos
+ by (smt "06" i'j' less_imp_le_nat less_trans_Suc nat_SN.gt_trans)
+ }
+ then show ?thesis by simp
+ qed
+ from inv(7) have dmu: "i' < m \<Longrightarrow> j' < m \<Longrightarrow> dmu $$ (i', j') = d\<mu> fs i' j'" for i' j'
+ by auto
+ note d_of = d_of_weak[OF Linv]
+ have dvec: "i \<le> m \<Longrightarrow> d_vec $ i = d fs i" for i unfolding d_vec_def using d_of by auto
+ show "\<forall>i'<m. \<forall>j'<m. d\<mu> fs' i' j' = dmu'' $$ (i', j')"
+ using 07 unfolding dmu''_def I_def
+ by (auto simp: dmu dvec)
+ qed
+
+ moreover
+ {
+ assume linv: "LLL_invariant_mod fs mfs dmu p first b i"
+ note inv = LLL_invD_mod[OF linv]
+ hence i: "i \<le> m" by auto
+ have norm: "j < m \<Longrightarrow> \<parallel>gso fs j\<parallel>\<^sup>2 = \<parallel>gso fs' j\<parallel>\<^sup>2" for j
+ using 05 by auto
+ have "weakly_reduced fs i = weakly_reduced fs' i"
+ unfolding gram_schmidt_fs.weakly_reduced_def using i
+ by (intro all_cong arg_cong2[where f = "(\<le>)"] arg_cong[where f = "\<lambda> x. _ * x"] norm, auto)
+ with inv have "weakly_reduced fs' i" by auto
+ hence "LLL_invariant_mod fs' mfs'' dmu'' p'' first b' i" using inv
+ by (intro LLL_invI_mod LLL_invD_modw[OF Linv'])
+ }
+
+ moreover have "LLL_measure i fs' = LLL_measure i fs"
+ unfolding LLL_measure_def logD_def D ..
+ ultimately show ?thesis unfolding id by blast
+qed
+
+lemma alpha_comparison: assumes
+ Linv: "LLL_invariant_mod_weak fs mfs dmu p first b"
+ and alph: "quotient_of \<alpha> = (num, denom)"
+ and i: "i < m"
+ and i0: "i \<noteq> 0"
+shows "(d_of dmu i * d_of dmu i * denom \<le> num * d_of dmu (i - 1) * d_of dmu (Suc i))
+ = (sq_norm (gso fs (i - 1)) \<le> \<alpha> * sq_norm (gso fs i))"
+proof -
+ note inv = LLL_invD_modw[OF Linv]
+ interpret fs_indep: fs_int_indpt n fs
+ by (unfold_locales, insert inv, auto)
+ from inv(2) i have ifs: "i < length fs" by auto
+ note d_of_fs = d_of_weak[OF Linv]
+ show ?thesis
+ unfolding fs_indep.d_sq_norm_comparison[OF alph ifs i0, symmetric]
+ by (subst (1 2 3 4) d_of_fs, use i d_def fs_indep.d_def in auto)
+qed
+
+lemma basis_reduction_adjust_swap_add_step: assumes
+ Linv: "LLL_invariant_mod_weak fs mfs dmu p first b"
+ and res: "basis_reduction_adjust_swap_add_step p first mfs dmu g_idx i = (p', mfs', dmu', g_idx')"
+ and alph: "quotient_of \<alpha> = (num, denom)"
+ and ineq: "\<not> (d_of dmu i * d_of dmu i * denom
+ \<le> num * d_of dmu (i - 1) * d_of dmu (Suc i))"
+ and i: "i < m"
+ and i0: "i \<noteq> 0"
+shows "\<exists>fs' b'. LLL_invariant_mod_weak fs' mfs' dmu' p' first b' \<and>
+ LLL_measure (i - 1) fs' < LLL_measure i fs \<and>
+ LLL_measure (m - 1) fs' < LLL_measure (m - 1) fs \<and>
+ (LLL_invariant_mod fs mfs dmu p first b i \<longrightarrow>
+ LLL_invariant_mod fs' mfs' dmu' p' first b' (i - 1))"
+proof -
+ obtain mfs0 dmu0 where add: "basis_reduction_mod_add_row p mfs dmu i (i-1) = (mfs0, dmu0)" by force
+ obtain mfs1 dmu1 where swap: "basis_reduction_mod_swap p mfs0 dmu0 i = (mfs1, dmu1)" by force
+ note res = res[unfolded basis_reduction_adjust_swap_add_step_def Let_def add split swap]
+ from i0 have ii: "i - 1 < i" by auto
+ from basis_reduction_mod_add_row[OF Linv add i ii i0]
+ obtain fs0 where Linv0: "LLL_invariant_mod_weak fs0 mfs0 dmu0 p first b"
+ and meas0: "LLL_measure i fs0 = LLL_measure i fs"
+ and small: "\<bar>\<mu> fs0 i (i - 1)\<bar> \<le> 1 / 2"
+ and Linv0': "LLL_invariant_mod fs mfs dmu p first b i \<Longrightarrow> LLL_invariant_mod fs0 mfs0 dmu0 p first b i"
+ by blast
+ {
+ have id: "d_of dmu0 i = d_of dmu i" "d_of dmu0 (i - 1) = d_of dmu (i - 1)"
+ "d_of dmu0 (Suc i) = d_of dmu (Suc i)"
+ using i i0 add[unfolded basis_reduction_mod_add_row_def Let_def]
+ by (auto split: if_splits simp: d_of_def)
+ from ineq[folded id, unfolded alpha_comparison[OF Linv0 alph i i0]]
+ have "\<parallel>gso fs0 (i - 1)\<parallel>\<^sup>2 > \<alpha> * \<parallel>gso fs0 i\<parallel>\<^sup>2" by simp
+ } note ineq = this
+ from Linv have "LLL_invariant_weak fs"
+ by (auto simp: LLL_invariant_weak_def LLL_invariant_mod_weak_def)
+ from basis_reduction_mod_swap[OF Linv0 small swap ineq i i0, unfolded meas0] Linv0'
+ obtain fs1 where Linv1: "LLL_invariant_mod_weak fs1 mfs1 dmu1 p first b"
+ and meas1: "LLL_measure (i - 1) fs1 < LLL_measure i fs"
+ and Linv1': "LLL_invariant_mod fs mfs dmu p first b i \<Longrightarrow> LLL_invariant_mod fs1 mfs1 dmu1 p first b (i - 1)"
+ by auto
+ show ?thesis
+ proof (cases "i - 1 = g_idx")
+ case False
+ with res have id: "p' = p" "mfs' = mfs1" "dmu' = dmu1" "g_idx' = g_idx" by auto
+ show ?thesis unfolding id using Linv1' meas1 Linv1 by (intro exI[of _ fs1] exI[of _ b], auto simp: LLL_measure_def)
+ next
+ case True
+ with res have adjust: "basis_reduction_adjust_mod p first mfs1 dmu1 = (p', mfs', dmu', g_idx')" by simp
+ from basis_reduction_adjust_mod[OF Linv1 adjust, of "i - 1"] Linv1'
+ obtain fs' b' where Linvw: "LLL_invariant_mod_weak fs' mfs' dmu' p' first b'"
+ and Linv: "LLL_invariant_mod fs mfs dmu p first b i \<Longrightarrow> LLL_invariant_mod fs' mfs' dmu' p' first b' (i - 1)"
+ and meas: "LLL_measure (i - 1) fs' = LLL_measure (i - 1) fs1"
+ by blast
+ note meas = meas1[folded meas]
+ from meas have meas': "LLL_measure (m - 1) fs' < LLL_measure (m - 1) fs"
+ unfolding LLL_measure_def using i by auto
+ show ?thesis
+ by (intro exI conjI impI, rule Linvw, rule meas, rule meas', rule Linv)
+ qed
+qed
+
+
+lemma basis_reduction_mod_step: assumes
+ Linv: "LLL_invariant_mod fs mfs dmu p first b i"
+ and res: "basis_reduction_mod_step p first mfs dmu g_idx i j = (p', mfs', dmu', g_idx', i', j')"
+ and i: "i < m"
+shows "\<exists>fs' b'. LLL_measure i' fs' < LLL_measure i fs \<and> LLL_invariant_mod fs' mfs' dmu' p' first b' i'"
+proof -
+ note res = res[unfolded basis_reduction_mod_step_def Let_def]
+ from Linv have Linvw: "LLL_invariant_mod_weak fs mfs dmu p first b"
+ by (auto simp: LLL_invariant_mod_weak_def LLL_invariant_mod_def)
+ show ?thesis
+ proof (cases "i = 0")
+ case True
+ then have ids: "mfs' = mfs" "dmu' = dmu" "i' = Suc i" "p' = p" using res by auto
+ have "LLL_measure i' fs < LLL_measure i fs \<and> LLL_invariant_mod fs mfs' dmu' p first b i'"
+ using increase_i_mod[OF Linv i] True res ids inv by simp
+ then show ?thesis using res ids inv by auto
+ next
+ case False
+ hence id: "(i = 0) = False" by auto
+ obtain num denom where alph: "quotient_of \<alpha> = (num, denom)" by force
+ note res = res[unfolded id if_False alph split]
+ let ?comp = "d_of dmu i * d_of dmu i * denom \<le> num * d_of dmu (i - 1) * d_of dmu (Suc i)"
+ show ?thesis
+ proof (cases ?comp)
+ case False
+ hence id: "?comp = False" by simp
+ note res = res[unfolded id if_False]
+ let ?step = "basis_reduction_adjust_swap_add_step p first mfs dmu g_idx i"
+ from res have step: "?step = (p', mfs', dmu', g_idx')"
+ and i': "i' = i - 1"
+ by (cases ?step, auto)+
+ from basis_reduction_adjust_swap_add_step[OF Linvw step alph False i \<open>i \<noteq> 0\<close>] Linv
+ show ?thesis unfolding i' by blast
+ next
+ case True
+ hence id: "?comp = True" by simp
+ note res = res[unfolded id if_True]
+ from res have ids: "p' = p" "mfs' = mfs" "dmu' = dmu" "i' = Suc i" by auto
+ from True alpha_comparison[OF Linvw alph i False]
+ have ineq: "sq_norm (gso fs (i - 1)) \<le> \<alpha> * sq_norm (gso fs i)" by simp
+ from increase_i_mod[OF Linv i ineq]
+ show ?thesis unfolding ids by auto
+ qed
+ qed
+qed
+
+lemma basis_reduction_mod_main: assumes "LLL_invariant_mod fs mfs dmu p first b i"
+ and res: "basis_reduction_mod_main p first mfs dmu g_idx i j = (p', mfs', dmu')"
+shows "\<exists>fs' b'. LLL_invariant_mod fs' mfs' dmu' p' first b' m"
+ using assms
+proof (induct "LLL_measure i fs" arbitrary: i mfs dmu j p b fs g_idx rule: less_induct)
+ case (less i fs mfs dmu j p b g_idx)
+ hence fsinv: "LLL_invariant_mod fs mfs dmu p first b i" by auto
+ note res = less(3)[unfolded basis_reduction_mod_main.simps[of p first mfs dmu g_idx i j]]
+ note inv = less(2)
+ note IH = less(1)
+ show ?case
+ proof (cases "i < m")
+ case i: True
+ obtain p' mfs' dmu' g_idx' i' j' where step: "basis_reduction_mod_step p first mfs dmu g_idx i j = (p', mfs', dmu', g_idx', i', j')"
+ (is "?step = _") by (cases ?step, auto)
+ then obtain fs' b' where Linv: "LLL_invariant_mod fs' mfs' dmu' p' first b' i'"
+ and decr: "LLL_measure i' fs' < LLL_measure i fs"
+ using basis_reduction_mod_step[OF fsinv step i] i fsinv by blast
+ note res = res[unfolded step split]
+ from res i show ?thesis using IH[OF decr Linv] by auto
+ next
+ case False
+ with LLL_invD_mod[OF fsinv] res have i: "i = m" "p' = p" by auto
+ then obtain fs' b' where "LLL_invariant_mod fs' mfs' dmu' p first b' m" using False res fsinv by simp
+ then show ?thesis using i by auto
+ qed
+qed
+
+lemma compute_max_gso_quot_alpha:
+ assumes inv: "LLL_invariant_mod_weak fs mfs dmu p first b"
+ and max: "compute_max_gso_quot dmu = (msq_num, msq_denum, idx)"
+ and alph: "quotient_of \<alpha> = (num, denum)"
+ and cmp: "(msq_num * denum > num * msq_denum) = cmp"
+ and m: "m > 1"
+shows "cmp \<Longrightarrow> idx \<noteq> 0 \<and> idx < m \<and> \<not> (d_of dmu idx * d_of dmu idx * denum
+ \<le> num * d_of dmu (idx - 1) * d_of dmu (Suc idx))"
+ and "\<not> cmp \<Longrightarrow> LLL_invariant_mod fs mfs dmu p first b m"
+proof -
+ from inv
+ have fsinv: "LLL_invariant_weak fs"
+ by (simp add: LLL_invariant_mod_weak_def LLL_invariant_weak_def)
+ define qt where "qt = (\<lambda>i. ((d_of dmu (i + 1)) * (d_of dmu (i + 1)),
+ (d_of dmu (i + 2)) * (d_of dmu i), Suc i))"
+ define lst where "lst = (map (\<lambda>i. qt i) [0..<(m-1)])"
+ have msqlst: "(msq_num, msq_denum, idx) = max_list_rats_with_index lst"
+ using max lst_def qt_def unfolding compute_max_gso_quot_def by simp
+ have nz: "\<And>n d i. (n, d, i) \<in> set lst \<Longrightarrow> d > 0"
+ unfolding lst_def qt_def using d_of_weak[OF inv] LLL_d_pos[OF fsinv] by auto
+ have geq: "\<forall>(n, d, i) \<in> set lst. rat_of_int msq_num / of_int msq_denum \<ge> rat_of_int n / of_int d"
+ using max_list_rats_with_index[of lst] nz msqlst by (metis (no_types, lifting) case_prodI2)
+ have len: "length lst \<ge> 1" using m unfolding lst_def by simp
+ have inset: "(msq_num, msq_denum, idx) \<in> set lst"
+ using max_list_rats_with_index_in_set[OF msqlst[symmetric] len] nz by simp
+ then have idxm: "idx \<in> {1..<m}" using lst_def[unfolded qt_def] by auto
+ then have idx0: "idx \<noteq> 0" and idx: "idx < m" by auto
+ have 00: "(msq_num, msq_denum, idx) = qt (idx - 1)" using lst_def inset qt_def by auto
+ then have id_qt: "msq_num = d_of dmu idx * d_of dmu idx" "msq_denum = d_of dmu (Suc idx) * d_of dmu (idx - 1)"
+ unfolding qt_def by auto
+ have "msq_denum = (d_of dmu (idx + 1)) * (d_of dmu (idx - 1))"
+ using 00 unfolding qt_def by simp
+ then have dengt0: "msq_denum > 0" using d_of_weak[OF inv] idxm LLL_d_pos[OF fsinv] by auto
+ have \<alpha>dengt0: "denum > 0" using alph by (metis quotient_of_denom_pos)
+ from cmp[unfolded id_qt]
+ have cmp: "cmp = (\<not> (d_of dmu idx * d_of dmu idx * denum \<le> num * d_of dmu (idx - 1) * d_of dmu (Suc idx)))"
+ by (auto simp: ac_simps)
+ {
+ assume cmp
+ from this[unfolded cmp]
+ show "idx \<noteq> 0 \<and> idx < m \<and> \<not> (d_of dmu idx * d_of dmu idx * denum
+ \<le> num * d_of dmu (idx - 1) * d_of dmu (Suc idx))" using idx0 idx by auto
+ }
+ {
+ assume "\<not> cmp"
+ from this[unfolded cmp] have small: "d_of dmu idx * d_of dmu idx * denum \<le> num * d_of dmu (idx - 1) * d_of dmu (Suc idx)" by auto
+ note d_pos = LLL_d_pos[OF fsinv]
+ have gso: "k < m \<Longrightarrow> sq_norm (gso fs k) = of_int (d fs (Suc k)) / of_int (d fs k)" for k using
+ LLL_d_Suc[OF fsinv, of k] d_pos[of k] by simp
+ have gso_pos: "k < m \<Longrightarrow> sq_norm (gso fs k) > 0" for k
+ using gso[of k] d_pos[of k] d_pos[of "Suc k"] by auto
+ from small[unfolded alpha_comparison[OF inv alph idx idx0]]
+ have alph: "sq_norm (gso fs (idx - 1)) \<le> \<alpha> * sq_norm (gso fs idx)" .
+ with gso_pos[OF idx] have alph: "sq_norm (gso fs (idx - 1)) / sq_norm (gso fs idx) \<le> \<alpha>"
+ by (metis mult_imp_div_pos_le)
+ have weak: "weakly_reduced fs m" unfolding gram_schmidt_fs.weakly_reduced_def
+ proof (intro allI impI, goal_cases)
+ case (1 i)
+ from idx have idx1: "idx - 1 < m" by auto
+ from geq[unfolded lst_def]
+ have mem: "(d_of dmu (Suc i) * d_of dmu (Suc i),
+ d_of dmu (Suc (Suc i)) * d_of dmu i, Suc i) \<in> set lst"
+ unfolding lst_def qt_def using 1 by auto
+ have "sq_norm (gso fs i) / sq_norm (gso fs (Suc i)) =
+ of_int (d_of dmu (Suc i) * d_of dmu (Suc i)) / of_int (d_of dmu (Suc (Suc i)) * d_of dmu i)"
+ using gso idx0 d_of_weak[OF inv] 1 by auto
+ also have "\<dots> \<le> rat_of_int msq_num / rat_of_int msq_denum"
+ using geq[rule_format, OF mem, unfolded split] by auto
+ also have "\<dots> = sq_norm (gso fs (idx - 1)) / sq_norm (gso fs idx)"
+ unfolding id_qt gso[OF idx] gso[OF idx1] using idx0 d_of_weak[OF inv] idx by auto
+ also have "\<dots> \<le> \<alpha>" by fact
+ finally show "sq_norm (gso fs i) \<le> \<alpha> * sq_norm (gso fs (Suc i))" using gso_pos[OF 1]
+ using pos_divide_le_eq by blast
+ qed
+ with inv show "LLL_invariant_mod fs mfs dmu p first b m"
+ by (auto simp: LLL_invariant_mod_weak_def LLL_invariant_mod_def)
+ }
+qed
+
+
+lemma small_m:
+ assumes inv: "LLL_invariant_mod_weak fs mfs dmu p first b"
+ and m: "m \<le> 1"
+shows "LLL_invariant_mod fs mfs dmu p first b m"
+proof -
+ have weak: "weakly_reduced fs m" unfolding gram_schmidt_fs.weakly_reduced_def using m
+ by auto
+ with inv show "LLL_invariant_mod fs mfs dmu p first b m"
+ by (auto simp: LLL_invariant_mod_weak_def LLL_invariant_mod_def)
+qed
+
+lemma basis_reduction_iso_main: assumes "LLL_invariant_mod_weak fs mfs dmu p first b"
+ and res: "basis_reduction_iso_main p first mfs dmu g_idx j = (p', mfs', dmu')"
+shows "\<exists>fs' b'. LLL_invariant_mod fs' mfs' dmu' p' first b' m"
+ using assms
+proof (induct "LLL_measure (m-1) fs" arbitrary: fs mfs dmu j p b g_idx rule: less_induct)
+ case (less fs mfs dmu j p b g_idx)
+ have inv: "LLL_invariant_mod_weak fs mfs dmu p first b" using less by auto
+ hence fsinv: "LLL_invariant_weak fs"
+ by (simp add: LLL_invariant_mod_weak_def LLL_invariant_weak_def)
+ note res = less(3)[unfolded basis_reduction_iso_main.simps[of p first mfs dmu g_idx j]]
+ note IH = less(1)
+ obtain msq_num msq_denum idx where max: "compute_max_gso_quot dmu = (msq_num, msq_denum, idx)"
+ by (metis prod_cases3)
+ obtain num denum where alph: "quotient_of \<alpha> = (num, denum)" by force
+ note res = res[unfolded max alph Let_def split]
+ consider (small) "m \<le> 1" | (final) "m > 1" "\<not> (num * msq_denum < msq_num * denum)" | (step) "m > 1" "num * msq_denum < msq_num * denum"
+ by linarith
+ thus ?case
+ proof cases
+ case *: step
+ obtain p1 mfs1 dmu1 g_idx1 where step: "basis_reduction_adjust_swap_add_step p first mfs dmu g_idx idx = (p1, mfs1, dmu1, g_idx1)"
+ by (metis prod_cases4)
+ from res[unfolded step split] * have res: "basis_reduction_iso_main p1 first mfs1 dmu1 g_idx1 (j + 1) = (p', mfs', dmu')" by auto
+ from compute_max_gso_quot_alpha(1)[OF inv max alph refl *]
+ have idx0: "idx \<noteq> 0" and idx: "idx < m" and cmp: "\<not> d_of dmu idx * d_of dmu idx * denum \<le> num * d_of dmu (idx - 1) * d_of dmu (Suc idx)" by auto
+ from basis_reduction_adjust_swap_add_step[OF inv step alph cmp idx idx0] obtain fs1 b1
+ where inv1: "LLL_invariant_mod_weak fs1 mfs1 dmu1 p1 first b1" and meas: "LLL_measure (m - 1) fs1 < LLL_measure (m - 1) fs"
+ by auto
+ from IH[OF meas inv1 res] show ?thesis .
+ next
+ case small
+ with res small_m[OF inv] show ?thesis by auto
+ next
+ case final
+ from compute_max_gso_quot_alpha(2)[OF inv max alph refl final]
+ final show ?thesis using res by auto
+ qed
+qed
+
+lemma basis_reduction_mod_add_rows_loop_inv': assumes
+ fsinv: "LLL_invariant_mod fs mfs dmu p first b m"
+ and res: "basis_reduction_mod_add_rows_loop p mfs dmu i i = (mfs', dmu')"
+ and i: "i < m"
+shows "\<exists>fs'. LLL_invariant_mod fs' mfs' dmu' p first b m \<and>
+ (\<forall>i' j'. i' < i \<longrightarrow> j' \<le> i' \<longrightarrow> \<mu> fs i' j' = \<mu> fs' i' j') \<and>
+ \<mu>_small fs' i"
+proof -
+ {
+ fix j
+ assume j: "j \<le> i" and mu_small: "\<mu>_small_row i fs j"
+ and resj: "basis_reduction_mod_add_rows_loop p mfs dmu i j = (mfs', dmu')"
+ have "\<exists>fs'. LLL_invariant_mod fs' mfs' dmu' p first b m \<and>
+ (\<forall>i' j'. i' < i \<longrightarrow> j' \<le> i' \<longrightarrow> \<mu> fs i' j' = \<mu> fs' i' j') \<and>
+ (\<mu>_small fs' i)"
+ proof (insert fsinv mu_small resj i j, induct j arbitrary: fs mfs dmu mfs' dmu')
+ case (0 fs)
+ then have "(mfs', dmu') = (mfs, dmu)" by simp
+ then show ?case
+ using LLL_invariant_mod_to_weak_m_to_i(3) basis_reduction_add_row_done_weak 0 by auto
+ next
+ case (Suc j)
+ hence j: "j < i" by auto
+ have in0: "i \<noteq> 0" using Suc(6) by simp
+ define c where "c = round_num_denom (dmu $$ (i,j)) (d_of dmu (Suc j))"
+ have c2: "c = round (\<mu> fs i j)" using dmu_quot_is_round_of_\<mu>[OF _ _ i j] c_def Suc by simp
+ define mfs'' where "mfs'' = (if c=0 then mfs else mfs[ i := (map_vec (\<lambda> x. x symmod p)) (mfs ! i - c \<cdot>\<^sub>v mfs ! j)])"
+ define dmu'' where "dmu'' = (if c=0 then dmu else mat m m (\<lambda>(i',j'). (if (i' = i \<and> j' \<le> j)
+ then (if j'=j then (dmu $$ (i,j') - c * dmu $$ (j,j'))
+ else (dmu $$ (i,j') - c * dmu $$ (j,j')) symmod (p * (d_of dmu j') * (d_of dmu (Suc j'))))
+ else (dmu $$ (i',j')))))"
+ have 00: "basis_reduction_mod_add_row p mfs dmu i j = (mfs'', dmu'')"
+ using mfs''_def dmu''_def unfolding basis_reduction_mod_add_row_def c_def[symmetric] by simp
+ then have 01: "basis_reduction_mod_add_rows_loop p mfs'' dmu'' i j = (mfs', dmu')"
+ using basis_reduction_mod_add_rows_loop.simps(2)[of p mfs dmu i j] Suc by simp
+ have fsinvi: "LLL_invariant_mod fs mfs dmu p first b i" using LLL_invariant_mod_to_weak_m_to_i[OF Suc(2)] i by simp
+ then have fsinvmw: "LLL_invariant_mod_weak fs mfs dmu p first b" using LLL_invD_mod LLL_invI_modw by simp
+ obtain fs'' where fs''invi: "LLL_invariant_mod fs'' mfs'' dmu'' p first b i" and
+ \<mu>_small': "(\<mu>_small_row i fs (Suc j) \<longrightarrow> \<mu>_small_row i fs'' j)" and
+ \<mu>s: "(\<forall>i' j'. i' < i \<longrightarrow> j' \<le> i' \<longrightarrow> \<mu> fs'' i' j' = \<mu> fs i' j')"
+ using Suc basis_reduction_mod_add_row[OF fsinvmw 00 i j] fsinvi by auto
+ moreover then have \<mu>sm: "\<mu>_small_row i fs'' j" using Suc by simp
+ have fs''invwi: "LLL_invariant_weak' i fs''" using LLL_invD_mod[OF fs''invi] LLL_invI_weak by simp
+ have fsinvwi: "LLL_invariant_weak' i fs" using LLL_invD_mod[OF fsinvi] LLL_invI_weak by simp
+ note invw = LLL_invw'_imp_w[OF fsinvwi]
+ note invw'' = LLL_invw'_imp_w[OF fs''invwi]
+ have "LLL_invariant_mod fs'' mfs'' dmu'' p first b m"
+ proof -
+ have "(\<forall> l. Suc l < m \<longrightarrow> sq_norm (gso fs'' l) \<le> \<alpha> * sq_norm (gso fs'' (Suc l)))"
+ proof -
+ {
+ fix l
+ assume l: "Suc l < m"
+ have "sq_norm (gso fs'' l) \<le> \<alpha> * sq_norm (gso fs'' (Suc l))"
+ proof (cases "i \<le> Suc l")
+ case True
+ have deq: "\<And>k. k < m \<Longrightarrow> d fs (Suc k) = d fs'' (Suc k)"
+ using dd\<mu> LLL_invD_mod(9)[OF fs''invi] LLL_invD_mod(9)[OF Suc(2)] dmu''_def j by simp
+ {
+ fix k
+ assume k: "k < m"
+ then have "d fs (Suc k) = d fs'' (Suc k)"
+ using dd\<mu> LLL_invD_mod(9)[OF fs''invi] LLL_invD_mod(9)[OF Suc(2)] dmu''_def j by simp
+ have "d fs 0 = 1" "d fs'' 0 = 1" using d_def by auto
+ moreover have sqid: "sq_norm (gso fs'' k) = rat_of_int (d fs'' (Suc k)) / rat_of_int (d fs'' k)"
+ using LLL_d_Suc[OF invw''] LLL_d_pos[OF invw''] k
+ by (smt One_nat_def Suc_less_eq Suc_pred le_imp_less_Suc mult_eq_0_iff less_imp_le_nat
+ nonzero_mult_div_cancel_right of_int_0_less_iff of_int_hom.hom_zero)
+ moreover have "sq_norm (gso fs k) = rat_of_int (d fs (Suc k)) / rat_of_int (d fs k)"
+ using LLL_d_Suc[OF invw] LLL_d_pos[OF invw] k
+ by (smt One_nat_def Suc_less_eq Suc_pred le_imp_less_Suc mult_eq_0_iff less_imp_le_nat
+ nonzero_mult_div_cancel_right of_int_0_less_iff of_int_hom.hom_zero)
+ ultimately have "sq_norm (gso fs k) = sq_norm (gso fs'' k)" using k deq
+ LLL_d_pos[OF invw] LLL_d_pos[OF invw'']
+ by (metis (no_types, lifting) Nat.lessE Suc_lessD old.nat.inject zero_less_Suc)
+ }
+ then show ?thesis using LLL_invD_mod(6)[OF Suc(2)] by (simp add: gram_schmidt_fs.weakly_reduced_def l)
+ next
+ case False
+ then show ?thesis using LLL_invD_mod(6)[OF fs''invi] gram_schmidt_fs.weakly_reduced_def
+ by (metis less_or_eq_imp_le nat_neq_iff)
+ qed
+ }
+ then show ?thesis by simp
+ qed
+ then have "weakly_reduced fs'' m" using gram_schmidt_fs.weakly_reduced_def by blast
+ then show ?thesis using LLL_invD_mod[OF fs''invi] LLL_invI_mod by simp
+ qed
+ then show ?case using "01" Suc.hyps i j less_imp_le_nat \<mu>sm \<mu>s by metis
+ qed
+ }
+ then show ?thesis using \<mu>_small_row_refl res by auto
+qed
+
+lemma basis_reduction_mod_add_rows_outer_loop_inv:
+ assumes inv: "LLL_invariant_mod fs mfs dmu p first b m"
+ and "(mfs', dmu') = basis_reduction_mod_add_rows_outer_loop p mfs dmu i"
+ and i: "i < m"
+shows "(\<exists>fs'. LLL_invariant_mod fs' mfs' dmu' p first b m \<and>
+ (\<forall>j. j \<le> i \<longrightarrow> \<mu>_small fs' j))"
+proof(insert assms, induct i arbitrary: fs mfs dmu mfs' dmu')
+ case (0 fs)
+ then show ?case using \<mu>_small_def by auto
+next
+ case (Suc i fs mfs dmu mfs' dmu')
+ obtain mfs'' dmu'' where mfs''dmu'': "(mfs'', dmu'')
+ = basis_reduction_mod_add_rows_outer_loop p mfs dmu i" by (metis surj_pair)
+ then obtain fs'' where fs'': "LLL_invariant_mod fs'' mfs'' dmu'' p first b m"
+ and 00: "(\<forall>j. j \<le> i \<longrightarrow> \<mu>_small fs'' j)" using Suc by fastforce
+ have "(mfs', dmu') = basis_reduction_mod_add_rows_loop p mfs'' dmu'' (Suc i) (Suc i)"
+ using Suc(3,4) mfs''dmu'' by (smt basis_reduction_mod_add_rows_outer_loop.simps(2) case_prod_conv)
+ then obtain fs' where 01: "LLL_invariant_mod fs' mfs' dmu' p first b m"
+ and 02: "\<forall>i' j'. i' < (Suc i) \<longrightarrow> j' \<le> i' \<longrightarrow> \<mu> fs'' i' j' = \<mu> fs' i' j'" and 03: "\<mu>_small fs' (Suc i)"
+ using fs'' basis_reduction_mod_add_rows_loop_inv' Suc by metis
+ moreover have "\<forall>j. j \<le> (Suc i) \<longrightarrow> \<mu>_small fs' j" using 02 00 03 \<mu>_small_def by (simp add: le_Suc_eq)
+ ultimately show ?case by blast
+qed
+
+lemma basis_reduction_mod_fs_bound:
+ assumes Linv: "LLL_invariant_mod fs mfs dmu p first b k"
+ and mu_small: "\<mu>_small fs i"
+ and i: "i < m"
+ and nFirst: "\<not> first"
+shows "fs ! i = mfs ! i"
+proof -
+ from LLL_invD_mod(16-17)[OF Linv] nFirst g_bnd_mode_def
+ have gbnd: "g_bnd b fs" and bp: "b \<le> (rat_of_int (p - 1))\<^sup>2 / (rat_of_nat m + 3)"
+ by (auto simp: mod_invariant_def bound_number_def)
+ have Linvw: "LLL_invariant_weak' k fs" using LLL_invD_mod[OF Linv] LLL_invI_weak by simp
+ have "fs_int_indpt n fs" using LLL_invD_mod(5)[OF Linv] Gram_Schmidt_2.fs_int_indpt.intro by simp
+ then interpret fs: fs_int_indpt n fs
+ using fs_int_indpt.sq_norm_fs_via_sum_mu_gso by simp
+ have "\<parallel>gso fs 0\<parallel>\<^sup>2 \<le> b" using gbnd i unfolding g_bnd_def by blast
+ then have b0: "0 \<le> b" using sq_norm_vec_ge_0 dual_order.trans by auto
+ have 00: "of_int \<parallel>fs ! i\<parallel>\<^sup>2 = (\<Sum>j\<leftarrow>[0..<Suc i]. (\<mu> fs i j)\<^sup>2 * \<parallel>gso fs j\<parallel>\<^sup>2)"
+ using fs.sq_norm_fs_via_sum_mu_gso LLL_invD_mod[OF Linv] Gram_Schmidt_2.fs_int_indpt.intro i by simp
+ have 01: "\<forall>j < i. (\<mu> fs i j)\<^sup>2 * \<parallel>gso fs j\<parallel>\<^sup>2 \<le> (1 / rat_of_int 4) * \<parallel>gso fs j\<parallel>\<^sup>2"
+ proof -
+ {
+ fix j
+ assume j: "j < i"
+ then have "\<bar>fs.gs.\<mu> i j\<bar> \<le> 1 / (rat_of_int 2)"
+ using mu_small Power.linordered_idom_class.abs_square_le_1 j unfolding \<mu>_small_def by simp
+ moreover have "\<bar>\<mu> fs i j\<bar> \<ge> 0" by simp
+ ultimately have "\<bar>\<mu> fs i j\<bar>\<^sup>2 \<le> (1 / rat_of_int 2)\<^sup>2"
+ using Power.linordered_idom_class.abs_le_square_iff by fastforce
+ also have "\<dots> = 1 / (rat_of_int 4)" by (simp add: field_simps)
+ finally have "\<bar>\<mu> fs i j\<bar>\<^sup>2 \<le> 1 / rat_of_int 4" by simp
+ }
+ then show ?thesis using fs.gs.\<mu>.simps by (metis mult_right_mono power2_abs sq_norm_vec_ge_0)
+ qed
+ then have 0111: "\<And>j. j \<in> set [0..<i] \<Longrightarrow> (\<mu> fs i j)\<^sup>2 * \<parallel>gso fs j\<parallel>\<^sup>2 \<le> (1 / rat_of_int 4) * \<parallel>gso fs j\<parallel>\<^sup>2"
+ by simp
+ {
+ fix j
+ assume j: "j < n"
+ have 011: "(\<mu> fs i i)\<^sup>2 * \<parallel>gso fs i\<parallel>\<^sup>2 = 1 * \<parallel>gso fs i\<parallel>\<^sup>2"
+ using fs.gs.\<mu>.simps by simp
+ have 02: "\<forall>j < Suc i. \<parallel>gso fs j\<parallel>\<^sup>2 \<le> b"
+ using gbnd i unfolding g_bnd_def by simp
+ have 03: "length [0..<Suc i] = (Suc i)" by simp
+ have "of_int \<parallel>fs ! i\<parallel>\<^sup>2 = (\<Sum>j\<leftarrow>[0..<i]. (\<mu> fs i j)\<^sup>2 * \<parallel>gso fs j\<parallel>\<^sup>2) + \<parallel>gso fs i\<parallel>\<^sup>2"
+ unfolding 00 using 011 by simp
+ also have "(\<Sum>j\<leftarrow>[0..<i]. (\<mu> fs i j)\<^sup>2 * \<parallel>gso fs j\<parallel>\<^sup>2) \<le> (\<Sum>j\<leftarrow>[0..<i]. ((1 / rat_of_int 4) * \<parallel>gso fs j\<parallel>\<^sup>2))"
+ using Groups_List.sum_list_mono[OF 0111] by fast
+ finally have "of_int \<parallel>fs ! i\<parallel>\<^sup>2 \<le> (\<Sum>j\<leftarrow>[0..<i]. ((1 / rat_of_int 4) * \<parallel>gso fs j\<parallel>\<^sup>2)) + \<parallel>gso fs i\<parallel>\<^sup>2"
+ by simp
+ also have "(\<Sum>j\<leftarrow>[0..<i]. ((1 / rat_of_int 4) * \<parallel>gso fs j\<parallel>\<^sup>2)) \<le> (\<Sum>j\<leftarrow>[0..<i]. (1 / rat_of_int 4) * b)"
+ by (intro sum_list_mono, insert 02, auto)
+ also have " \<parallel>gso fs i\<parallel>\<^sup>2 \<le> b" using 02 by simp
+ finally have "of_int \<parallel>fs ! i\<parallel>\<^sup>2 \<le> (\<Sum>j\<leftarrow>[0..<i]. (1 / rat_of_int 4) * b) + b" by simp
+ also have "\<dots> = (rat_of_nat i) * ((1 / rat_of_int 4) * b) + b"
+ using 03 sum_list_triv[of "(1 / rat_of_int 4) * b" "[0..<i]"] by simp
+ also have "\<dots> = (rat_of_nat i) / 4 * b + b" by simp
+ also have "\<dots> = ((rat_of_nat i) / 4 + 1)* b" by algebra
+ also have "\<dots> = (rat_of_nat i + 4) / 4 * b" by simp
+ finally have "of_int \<parallel>fs ! i\<parallel>\<^sup>2 \<le> (rat_of_nat i + 4) / 4 * b" by simp
+ also have "\<dots> \<le> (rat_of_nat (m + 3)) / 4 * b" using i b0 times_left_mono by fastforce
+ finally have "of_int \<parallel>fs ! i\<parallel>\<^sup>2 \<le> rat_of_nat (m+3) / 4 * b" by simp
+ moreover have "\<bar>fs ! i $ j\<bar>\<^sup>2 \<le> \<parallel>fs ! i\<parallel>\<^sup>2" using vec_le_sq_norm LLL_invD_mod(10)[OF Linv] i j by blast
+ ultimately have 04: "of_int (\<bar>fs ! i $ j\<bar>\<^sup>2) \<le> rat_of_nat (m+3) / 4 * b" using ge_trans i by linarith
+ then have 05: "real_of_int (\<bar>fs ! i $ j\<bar>\<^sup>2) \<le> real_of_rat (rat_of_nat (m+3) / 4 * b)"
+ proof -
+ from j have "rat_of_int (\<bar>fs ! i $ j\<bar>\<^sup>2) \<le> rat_of_nat (m+3) / 4 * b" using 04 by simp
+ then have "real_of_int (\<bar>fs ! i $ j\<bar>\<^sup>2) \<le> real_of_rat (rat_of_nat (m+3) / 4 * b)"
+ using j of_rat_less_eq by (metis of_rat_of_int_eq)
+ then show ?thesis by simp
+ qed
+ define rhs where "rhs = real_of_rat (rat_of_nat (m+3) / 4 * b)"
+ have rhs0: "rhs \<ge> 0" using b0 i rhs_def by simp
+ have fsij: "real_of_int \<bar>fs ! i $ j\<bar> \<ge> 0" by simp
+ have "real_of_int (\<bar>fs ! i $ j\<bar>\<^sup>2) = (real_of_int \<bar>fs ! i $ j\<bar>)\<^sup>2" by simp
+ then have "(real_of_int \<bar>fs ! i $ j\<bar>)\<^sup>2 \<le> rhs" using 05 j rhs_def by simp
+ then have g1: "real_of_int \<bar>fs ! i $ j\<bar> \<le> sqrt rhs" using NthRoot.real_le_rsqrt by simp
+ have pbnd: "2 * \<bar>fs ! i $ j\<bar> < p"
+ proof -
+ have "rat_of_nat (m+3) / 4 * b \<le> (rat_of_nat (m +3) / 4) * (rat_of_int (p - 1))\<^sup>2 / (rat_of_nat m+3)"
+ using bp b0 i times_left_mono SN_Orders.of_nat_ge_zero gs.m_comm times_divide_eq_right
+ by (smt gs.l_null le_divide_eq_numeral1(1))
+ also have "\<dots> = (rat_of_int (p - 1))\<^sup>2 / 4 * (rat_of_nat (m + 3) / rat_of_nat (m + 3))"
+ by (metis (no_types, lifting) gs.m_comm of_nat_add of_nat_numeral times_divide_eq_left)
+ finally have "rat_of_nat (m+3) / 4 * b \<le> (rat_of_int (p - 1))\<^sup>2 / 4" by simp
+ then have "sqrt rhs \<le> sqrt (real_of_rat ((rat_of_int (p - 1))\<^sup>2 / 4))"
+ unfolding rhs_def using of_rat_less_eq by fastforce
+ then have two_ineq:
+ "2 * \<bar>fs ! i $ j\<bar> \<le> 2 * sqrt (real_of_rat ((rat_of_int (p - 1))\<^sup>2 / 4))"
+ using g1 by linarith
+ have "2 * sqrt (real_of_rat ((rat_of_int (p - 1))\<^sup>2 / 4)) =
+ sqrt (real_of_rat (4 * ((rat_of_int (p - 1))\<^sup>2 / 4)))"
+ by (metis (no_types, opaque_lifting) real_sqrt_mult of_int_numeral of_rat_hom.hom_mult
+ of_rat_of_int_eq real_sqrt_four times_divide_eq_right)
+ also have "\<dots> = sqrt (real_of_rat ((rat_of_int (p - 1))\<^sup>2))" using i by simp
+ also have "(real_of_rat ((rat_of_int (p - 1))\<^sup>2)) = (real_of_rat (rat_of_int (p - 1)))\<^sup>2"
+ using Rat.of_rat_power by blast
+ also have "sqrt ((real_of_rat (rat_of_int (p - 1)))\<^sup>2) = real_of_rat (rat_of_int (p - 1))"
+ using LLL_invD_mod(15)[OF Linv] by simp
+ finally have "2 * sqrt (real_of_rat ((rat_of_int (p - 1))\<^sup>2 / 4)) =
+ real_of_rat (rat_of_int (p - 1))" by simp
+ then have "2 * \<bar>fs ! i $ j\<bar> \<le> real_of_rat (rat_of_int (p - 1))"
+ using two_ineq by simp
+ then show ?thesis by (metis of_int_le_iff of_rat_of_int_eq zle_diff1_eq)
+ qed
+ have p1: "p > 1" using LLL_invD_mod[OF Linv] by blast
+ interpret pm: poly_mod_2 p
+ by (unfold_locales, rule p1)
+ from LLL_invD_mod[OF Linv] have len: "length fs = m" and fs: "set fs \<subseteq> carrier_vec n" by auto
+ from pm.inv_M_rev[OF pbnd, unfolded pm.M_def] have "pm.inv_M (fs ! i $ j mod p) = fs ! i $ j" .
+ also have "pm.inv_M (fs ! i $ j mod p) = mfs ! i $ j" unfolding LLL_invD_mod(7)[OF Linv, symmetric] sym_mod_def
+ using i j len fs by auto
+ finally have "fs ! i $ j = mfs ! i $ j" ..
+ }
+ thus "fs ! i = mfs ! i" using LLL_invD_mod(10,13)[OF Linv i] by auto
+qed
+
+lemma basis_reduction_mod_fs_bound_first:
+ assumes Linv: "LLL_invariant_mod fs mfs dmu p first b k"
+ and m0: "m > 0"
+ and first: "first"
+shows "fs ! 0 = mfs ! 0"
+proof -
+ from LLL_invD_mod(16-17)[OF Linv] first g_bnd_mode_def m0
+ have gbnd: "sq_norm (gso fs 0) \<le> b" and bp: "b \<le> (rat_of_int (p - 1))\<^sup>2 / 4"
+ by (auto simp: mod_invariant_def bound_number_def)
+ from LLL_invD_mod[OF Linv] have p1: "p > 1" by blast
+ have Linvw: "LLL_invariant_weak' k fs" using LLL_invD_mod[OF Linv] LLL_invI_weak by simp
+ have "fs_int_indpt n fs" using LLL_invD_mod(5)[OF Linv] Gram_Schmidt_2.fs_int_indpt.intro by simp
+ then interpret fs: fs_int_indpt n fs
+ using fs_int_indpt.sq_norm_fs_via_sum_mu_gso by simp
+ from gbnd have b0: "0 \<le> b" using sq_norm_vec_ge_0 dual_order.trans by auto
+ have "of_int \<parallel>fs ! 0\<parallel>\<^sup>2 = (\<mu> fs 0 0)\<^sup>2 * \<parallel>gso fs 0\<parallel>\<^sup>2"
+ using fs.sq_norm_fs_via_sum_mu_gso LLL_invD_mod[OF Linv] Gram_Schmidt_2.fs_int_indpt.intro m0 by simp
+ also have "\<dots> = \<parallel>gso fs 0\<parallel>\<^sup>2" unfolding fs.gs.\<mu>.simps by (simp add: gs.\<mu>.simps)
+ also have "\<dots> \<le> (rat_of_int (p - 1))\<^sup>2 / 4" using gbnd bp by auto
+ finally have one: "of_int (sq_norm (fs ! 0)) \<le> (rat_of_int (p - 1))\<^sup>2 / 4" .
+ {
+ fix j
+ assume j: "j < n"
+ have leq: "\<bar>fs ! 0 $ j\<bar>\<^sup>2 \<le> \<parallel>fs ! 0\<parallel>\<^sup>2" using vec_le_sq_norm LLL_invD_mod(10)[OF Linv] m0 j by blast
+ have "rat_of_int ((2 * \<bar>fs ! 0 $ j\<bar>)^2) = rat_of_int (4 * \<bar>fs ! 0 $ j\<bar>\<^sup>2)" by simp
+ also have "\<dots> \<le> 4 * of_int \<parallel>fs ! 0\<parallel>\<^sup>2" using leq by simp
+ also have "\<dots> \<le> 4 * (rat_of_int (p - 1))\<^sup>2 / 4" using one by simp
+ also have "\<dots> = (rat_of_int (p - 1))\<^sup>2" by simp
+ also have "\<dots> = rat_of_int ((p - 1)\<^sup>2)" by simp
+ finally have "(2 * \<bar>fs ! 0 $ j\<bar>)^2 \<le> (p - 1)\<^sup>2" by linarith
+ hence "2 * \<bar>fs ! 0 $ j\<bar> \<le> p - 1" using p1
+ by (smt power_mono_iff zero_less_numeral)
+ hence pbnd: "2 * \<bar>fs ! 0 $ j\<bar> < p" by simp
+ interpret pm: poly_mod_2 p
+ by (unfold_locales, rule p1)
+ from LLL_invD_mod[OF Linv] m0 have len: "length fs = m" "length mfs = m"
+ and fs: "fs ! 0 \<in> carrier_vec n" "mfs ! 0 \<in> carrier_vec n" by auto
+ from pm.inv_M_rev[OF pbnd, unfolded pm.M_def] have "pm.inv_M (fs ! 0 $ j mod p) = fs ! 0 $ j" .
+ also have "pm.inv_M (fs ! 0 $ j mod p) = mfs ! 0 $ j" unfolding LLL_invD_mod(7)[OF Linv, symmetric] sym_mod_def
+ using m0 j len fs by auto
+ finally have "mfs ! 0 $ j = fs ! 0 $ j" .
+ }
+ thus "fs ! 0 = mfs ! 0" using LLL_invD_mod(10,13)[OF Linv m0] by auto
+qed
+
+lemma dmu_initial: "dmu_initial = mat m m (\<lambda> (i,j). d\<mu> fs_init i j)"
+proof -
+ interpret fs: fs_int_indpt n fs_init
+ by (unfold_locales, intro lin_dep)
+ show ?thesis unfolding dmu_initial_def Let_def
+ proof (intro cong_mat refl refl, unfold split, goal_cases)
+ case (1 i j)
+ show ?case
+ proof (cases "j \<le> i")
+ case False
+ thus ?thesis by (auto simp: d\<mu>_def gs.\<mu>.simps)
+ next
+ case True
+ hence id: "d\<mu>_impl fs_init !! i !! j = fs.d\<mu> i j" unfolding fs.d\<mu>_impl
+ by (subst of_fun_nth, use 1 len in force, subst of_fun_nth, insert True, auto)
+ also have "\<dots> = d\<mu> fs_init i j" unfolding fs.d\<mu>_def d\<mu>_def fs.d_def d_def by simp
+ finally show ?thesis using True by auto
+ qed
+ qed
+qed
+
+lemma LLL_initial_invariant_mod: assumes res: "compute_initial_state first = (p, mfs, dmu', g_idx)"
+shows "\<exists>fs b. LLL_invariant_mod fs mfs dmu' p first b 0"
+proof -
+ from dmu_initial have dmu: "(\<forall>i' < m. \<forall>j' < m. d\<mu> fs_init i' j' = dmu_initial $$ (i',j'))" by auto
+ obtain b g_idx where norm: "compute_max_gso_norm first dmu_initial = (b,g_idx)" by force
+ note res = res[unfolded compute_initial_state_def Let_def norm split]
+ from res have p: "p = compute_mod_of_max_gso_norm first b" by auto
+ then have p0: "p > 0" unfolding compute_mod_of_max_gso_norm_def using log_base by simp
+ then have p1: "p \<ge> 1" by simp
+ note res = res[folded p]
+ from res[unfolded compute_initial_mfs_def]
+ have mfs: "mfs = map (map_vec (\<lambda>x. x symmod p)) fs_init" by auto
+ from res[unfolded compute_initial_dmu_def]
+ have dmu': "dmu' = mat m m (\<lambda>(i',j'). if j' < i'
+ then dmu_initial $$ (i', j') symmod (p * d_of dmu_initial j' * d_of dmu_initial (Suc j'))
+ else dmu_initial $$ (i',j'))" by auto
+ have lat: "lattice_of fs_init = L" by (auto simp: L_def)
+ define I where "I = {(i',j'). i' < m \<and> j' < i'}"
+ obtain fs where
+ 01: "lattice_of fs = L" and
+ 02: "map (map_vec (\<lambda> x. x symmod p)) fs = map (map_vec (\<lambda> x. x symmod p)) fs_init" and
+ 03: "lin_indep fs" and
+ 04: "length fs = m" and
+ 05: "(\<forall> k < m. gso fs k = gso fs_init k)" and
+ 06: "(\<forall> k \<le> m. d fs k = d fs_init k)" and
+ 07: "(\<forall> i' < m. \<forall> j' < m. d\<mu> fs i' j' =
+ (if (i',j') \<in> I then d\<mu> fs_init i' j' symmod (p * d fs_init j' * d fs_init (Suc j')) else d\<mu> fs_init i' j'))"
+ using mod_finite_set[OF lin_dep len _ lat p0, of I] I_def by blast
+ have inv: "LLL_invariant_weak fs_init"
+ by (intro LLL_inv_wI lat len lin_dep fs_init)
+ have "\<forall>i'<m. d\<mu> fs_init i' i' = dmu_initial $$ (i', i')" unfolding dmu_initial by auto
+ from compute_max_gso_norm[OF this inv, of first, unfolded norm] have gbnd: "g_bnd_mode first b fs_init"
+ and b0: "0 \<le> b" and mb0: "m = 0 \<Longrightarrow> b = 0" by auto
+ from gbnd 05 have gbnd: "g_bnd_mode first b fs" using g_bnd_mode_cong[of fs fs_init] by auto
+ have d\<mu>dmu': "\<forall>i'<m. \<forall>j'<m. d\<mu> fs i' j' = dmu' $$ (i', j')" using 07 dmu d_of_main[of fs_init dmu_initial]
+ unfolding I_def dmu' by simp
+ have wred: "weakly_reduced fs 0" by (simp add: gram_schmidt_fs.weakly_reduced_def)
+ have fs_carr: "set fs \<subseteq> carrier_vec n" using 03 unfolding gs.lin_indpt_list_def by force
+ have m0: "m \<ge> 0" using len by auto
+ have Linv: "LLL_invariant_weak' 0 fs"
+ by (intro LLL_invI_weak 03 04 01 wred fs_carr m0)
+ note Linvw = LLL_invw'_imp_w[OF Linv]
+ from compute_mod_of_max_gso_norm[OF b0 mb0 p]
+ have p: "mod_invariant b p first" "p > 1" by auto
+ from len mfs have len': "length mfs = m" by auto
+ have modbnd: "\<forall>i'<m. \<forall>j'<i'. \<bar>d\<mu> fs i' j'\<bar> < p * d fs j' * d fs (Suc j')"
+ proof -
+ have "\<forall> i' < m. \<forall> j' < i'. d\<mu> fs i' j' = d\<mu> fs i' j' symmod (p * d fs j' * d fs (Suc j'))"
+ using I_def 07 06 by simp
+ moreover have "\<forall>j' < m. p * d fs j' * d fs (Suc j') > 0" using p(2) LLL_d_pos[OF Linvw] by simp
+ ultimately show ?thesis using sym_mod_abs
+ by (smt Euclidean_Division.pos_mod_bound Euclidean_Division.pos_mod_sign less_trans)
+ qed
+ have "LLL_invariant_mod fs mfs dmu' p first b 0"
+ using LLL_invI_mod[OF len' m0 04 01 03 wred _ modbnd d\<mu>dmu' p(2) gbnd p(1)] 02 mfs by simp
+ then show ?thesis by auto
+qed
+
+subsection \<open>Soundness of Storjohann's algorithm\<close>
+
+text \<open>For all of these abstract algorithms, we actually formulate their soundness proofs by linking
+ to the LLL-invariant (which implies that @{term fs} is reduced (@{term "LLL_invariant True m fs"})
+ or that the first vector of @{term fs} is short (@{term "LLL_invariant_weak fs \<and> weakly_reduced fs m"}).\<close>
+
+text \<open>Soundness of Storjohann's algorithm\<close>
+lemma reduce_basis_mod_inv: assumes res: "reduce_basis_mod = fs"
+ shows "LLL_invariant True m fs"
+proof (cases "m = 0")
+ case True
+ from True have *: "fs_init = []" using len by simp
+ moreover have "fs = []" using res basis_reduction_mod_add_rows_outer_loop.simps(1)
+ unfolding reduce_basis_mod_def Let_def basis_reduction_mod_main.simps[of _ _ _ _ _ 0]
+ compute_initial_mfs_def compute_initial_state_def compute_initial_dmu_def
+ unfolding True * by (auto split: prod.splits)
+ ultimately show ?thesis using True LLL_inv_initial_state by blast
+next
+ case False
+ let ?first = False
+ obtain p mfs0 dmu0 g_idx0 where init: "compute_initial_state ?first = (p, mfs0, dmu0, g_idx0)" by (metis prod_cases4)
+ from LLL_initial_invariant_mod[OF init]
+ obtain fs0 b where fs0: "LLL_invariant_mod fs0 mfs0 dmu0 p ?first b 0" by blast
+ note res = res[unfolded reduce_basis_mod_def init Let_def split]
+ obtain p1 mfs1 dmu1 where mfs1dmu1: "(p1, mfs1, dmu1) = basis_reduction_mod_main p ?first mfs0 dmu0 g_idx0 0 0"
+ by (metis prod.exhaust)
+ obtain fs1 b1 where Linv1: "LLL_invariant_mod fs1 mfs1 dmu1 p1 ?first b1 m"
+ using basis_reduction_mod_main[OF fs0 mfs1dmu1[symmetric]] by auto
+ obtain mfs2 dmu2 where mfs2dmu2:
+ "(mfs2, dmu2) = basis_reduction_mod_add_rows_outer_loop p1 mfs1 dmu1 (m-1)" by (metis old.prod.exhaust)
+ obtain fs2 where fs2: "LLL_invariant_mod fs2 mfs2 dmu2 p1 ?first b1 m"
+ and \<mu>s: "((\<forall>j. j < m \<longrightarrow> \<mu>_small fs2 j))"
+ using basis_reduction_mod_add_rows_outer_loop_inv[OF _ mfs2dmu2, of fs1 ?first b1] Linv1 False by auto
+ have rbd: "LLL_invariant_weak' m fs2" "\<forall>j < m. \<mu>_small fs2 j"
+ using LLL_invD_mod[OF fs2] LLL_invI_weak \<mu>s by auto
+ have redfs2: "reduced fs2 m" using rbd LLL_invD_weak(8) gram_schmidt_fs.reduced_def \<mu>_small_def by blast
+ have fs: "fs = mfs2"
+ using res[folded mfs1dmu1, unfolded Let_def split, folded mfs2dmu2, unfolded split] ..
+ have "\<forall>i < m. fs2 ! i = fs ! i"
+ proof (intro allI impI)
+ fix i
+ assume i: "i < m"
+ then have fs2i: "LLL_invariant_mod fs2 mfs2 dmu2 p1 ?first b1 i"
+ using fs2 LLL_invariant_mod_to_weak_m_to_i by simp
+ have \<mu>si: "\<mu>_small fs2 i" using \<mu>s i by simp
+ show "fs2 ! i = fs ! i"
+ using basis_reduction_mod_fs_bound(1)[OF fs2i \<mu>si i] fs by simp
+ qed
+ then have "fs2 = fs"
+ using LLL_invD_mod(1,3,10,13)[OF fs2] fs by (metis nth_equalityI)
+ then show ?thesis using redfs2 fs rbd(1) reduce_basis_def res LLL_invD_weak
+ LLL_invariant_def by simp
+qed
+
+text \<open>Soundness of Storjohann's algorithm for computing a short vector.\<close>
+lemma short_vector_mod_inv: assumes res: "short_vector_mod = v"
+ and m: "m > 0"
+ shows "\<exists> fs. LLL_invariant_weak fs \<and> weakly_reduced fs m \<and> v = hd fs"
+proof -
+ let ?first = True
+ obtain p mfs0 dmu0 g_idx0 where init: "compute_initial_state ?first = (p, mfs0, dmu0, g_idx0)" by (metis prod_cases4)
+ from LLL_initial_invariant_mod[OF init]
+ obtain fs0 b where fs0: "LLL_invariant_mod fs0 mfs0 dmu0 p ?first b 0" by blast
+ obtain p1 mfs1 dmu1 where main: "basis_reduction_mod_main p ?first mfs0 dmu0 g_idx0 0 0 = (p1, mfs1, dmu1)"
+ by (metis prod.exhaust)
+ obtain fs1 b1 where Linv1: "LLL_invariant_mod fs1 mfs1 dmu1 p1 ?first b1 m"
+ using basis_reduction_mod_main[OF fs0 main] by auto
+ have "v = hd mfs1" using res[unfolded short_vector_mod_def Let_def init split main] ..
+ with basis_reduction_mod_fs_bound_first[OF Linv1 m] LLL_invD_mod(1,3)[OF Linv1] m
+ have v: "v = hd fs1" by (cases fs1; cases mfs1; auto)
+ from Linv1 have Linv1: "LLL_invariant_weak fs1" and red: "weakly_reduced fs1 m"
+ unfolding LLL_invariant_mod_def LLL_invariant_weak_def by auto
+ show ?thesis
+ by (intro exI[of _ fs1] conjI Linv1 red v)
+qed
+
+text \<open>Soundness of Storjohann's algorithm with improved swap order\<close>
+lemma reduce_basis_iso_inv: assumes res: "reduce_basis_iso = fs"
+ shows "LLL_invariant True m fs"
+proof (cases "m = 0")
+ case True
+ then have *: "fs_init = []" using len by simp
+ moreover have "fs = []" using res basis_reduction_mod_add_rows_outer_loop.simps(1)
+ unfolding reduce_basis_iso_def Let_def basis_reduction_iso_main.simps[of _ _ _ _ _ 0]
+ compute_initial_mfs_def compute_initial_state_def compute_initial_dmu_def
+ unfolding True * by (auto split: prod.splits)
+ ultimately show ?thesis using True LLL_inv_initial_state by blast
+next
+ case False
+ let ?first = False
+ obtain p mfs0 dmu0 g_idx0 where init: "compute_initial_state ?first = (p, mfs0, dmu0, g_idx0)" by (metis prod_cases4)
+ from LLL_initial_invariant_mod[OF init]
+ obtain fs0 b where fs0: "LLL_invariant_mod fs0 mfs0 dmu0 p ?first b 0" by blast
+ have fs0w: "LLL_invariant_mod_weak fs0 mfs0 dmu0 p ?first b" using LLL_invD_mod[OF fs0] LLL_invI_modw by simp
+ note res = res[unfolded reduce_basis_iso_def init Let_def split]
+ obtain p1 mfs1 dmu1 where mfs1dmu1: "(p1, mfs1, dmu1) = basis_reduction_iso_main p ?first mfs0 dmu0 g_idx0 0"
+ by (metis prod.exhaust)
+ obtain fs1 b1 where Linv1: "LLL_invariant_mod fs1 mfs1 dmu1 p1 ?first b1 m"
+ using basis_reduction_iso_main[OF fs0w mfs1dmu1[symmetric]] by auto
+ obtain mfs2 dmu2 where mfs2dmu2:
+ "(mfs2, dmu2) = basis_reduction_mod_add_rows_outer_loop p1 mfs1 dmu1 (m-1)" by (metis old.prod.exhaust)
+ obtain fs2 where fs2: "LLL_invariant_mod fs2 mfs2 dmu2 p1 ?first b1 m"
+ and \<mu>s: "((\<forall>j. j < m \<longrightarrow> \<mu>_small fs2 j))"
+ using basis_reduction_mod_add_rows_outer_loop_inv[OF _ mfs2dmu2, of fs1 ?first b1] Linv1 False by auto
+ have rbd: "LLL_invariant_weak' m fs2" "\<forall>j < m. \<mu>_small fs2 j"
+ using LLL_invD_mod[OF fs2] LLL_invI_weak \<mu>s by auto
+ have redfs2: "reduced fs2 m" using rbd LLL_invD_weak(8) gram_schmidt_fs.reduced_def \<mu>_small_def by blast
+ have fs: "fs = mfs2"
+ using res[folded mfs1dmu1, unfolded Let_def split, folded mfs2dmu2, unfolded split] ..
+ have "\<forall>i < m. fs2 ! i = fs ! i"
+ proof (intro allI impI)
+ fix i
+ assume i: "i < m"
+ then have fs2i: "LLL_invariant_mod fs2 mfs2 dmu2 p1 ?first b1 i"
+ using fs2 LLL_invariant_mod_to_weak_m_to_i by simp
+ have \<mu>si: "\<mu>_small fs2 i" using \<mu>s i by simp
+ show "fs2 ! i = fs ! i"
+ using basis_reduction_mod_fs_bound(1)[OF fs2i \<mu>si i] fs by simp
+ qed
+ then have "fs2 = fs"
+ using LLL_invD_mod(1,3,10,13)[OF fs2] fs by (metis nth_equalityI)
+ then show ?thesis using redfs2 fs rbd(1) reduce_basis_def res LLL_invD_weak
+ LLL_invariant_def by simp
+qed
+
+text \<open>Soundness of Storjohann's algorithm to compute short vectors with improved swap order\<close>
+lemma short_vector_iso_inv: assumes res: "short_vector_iso = v"
+ and m: "m > 0"
+ shows "\<exists> fs. LLL_invariant_weak fs \<and> weakly_reduced fs m \<and> v = hd fs"
+proof -
+ let ?first = True
+ obtain p mfs0 dmu0 g_idx0 where init: "compute_initial_state ?first = (p, mfs0, dmu0, g_idx0)" by (metis prod_cases4)
+ from LLL_initial_invariant_mod[OF init]
+ obtain fs0 b where fs0: "LLL_invariant_mod fs0 mfs0 dmu0 p ?first b 0" by blast
+ have fs0w: "LLL_invariant_mod_weak fs0 mfs0 dmu0 p ?first b" using LLL_invD_mod[OF fs0] LLL_invI_modw by simp
+ obtain p1 mfs1 dmu1 where main: "basis_reduction_iso_main p ?first mfs0 dmu0 g_idx0 0 = (p1, mfs1, dmu1)"
+ by (metis prod.exhaust)
+ obtain fs1 b1 where Linv1: "LLL_invariant_mod fs1 mfs1 dmu1 p1 ?first b1 m"
+ using basis_reduction_iso_main[OF fs0w main] by auto
+ have "v = hd mfs1" using res[unfolded short_vector_iso_def Let_def init split main] ..
+ with basis_reduction_mod_fs_bound_first[OF Linv1 m] LLL_invD_mod(1,3)[OF Linv1] m
+ have v: "v = hd fs1" by (cases fs1; cases mfs1; auto)
+ from Linv1 have Linv1: "LLL_invariant_weak fs1" and red: "weakly_reduced fs1 m"
+ unfolding LLL_invariant_mod_def LLL_invariant_weak_def by auto
+ show ?thesis
+ by (intro exI[of _ fs1] conjI Linv1 red v)
+qed
+
+end
+
+text \<open>From the soundness results of these abstract versions of the algorithms,
+ one just needs to derive actual implementations that may integrate low-level
+ optimizations.\<close>
+
+end
diff --git a/thys/Projective_Measurements/CHSH_Inequality.thy b/thys/Projective_Measurements/CHSH_Inequality.thy
--- a/thys/Projective_Measurements/CHSH_Inequality.thy
+++ b/thys/Projective_Measurements/CHSH_Inequality.thy
@@ -1,1457 +1,1457 @@
-(*
-Author:
- Mnacho Echenim, Université Grenoble Alpes
-*)
-
-theory CHSH_Inequality imports
- Projective_Measurements
-
-
-begin
-
-
-section \<open>The CHSH inequality\<close>
-
-text \<open>The local hidden variable assumption for quantum mechanics was developed to make the case
-that quantum theory was incomplete. In this part we formalize the CHSH inequality, which provides
-an upper-bound of a quantity involving expectations in a probability space, and exploit this
-inequality to show that the local hidden variable assumption does not hold.\<close>
-
-subsection \<open>Inequality statement\<close>
-
-lemma chsh_real:
- fixes A0::real
- assumes "\<bar>A0 * B1\<bar> \<le> 1"
- and "\<bar>A0 * B0\<bar> \<le> 1"
- and "\<bar>A1 * B0\<bar> \<le> 1"
- and "\<bar>A1 * B1\<bar> \<le> 1"
- shows "\<bar>A0 * B1 - A0 * B0 + A1 * B0 + A1*B1\<bar> \<le> 2"
-proof -
- have "A0 * B1 - A0 * B0 = A0 * B1 - A0 * B0 + A0 * B1 * A1 * B0 - A0 * B1 * A1 * B0" by simp
- also have "... = A0 * B1 * (1 + A1*B0) - A0 * B0 * (1 + A1* B1)"
- by (metis (no_types, opaque_lifting) add_diff_cancel_right calculation diff_add_eq
- group_cancel.sub1 mult.commute mult.right_neutral
- vector_space_over_itself.scale_left_commute
- vector_space_over_itself.scale_right_diff_distrib
- vector_space_over_itself.scale_right_distrib
- vector_space_over_itself.scale_scale)
- finally have "A0 * B1 - A0 * B0 = A0 * B1 * (1 + A1*B0) - A0 * B0 * (1 + A1* B1)" .
- hence "\<bar>A0 * B1 - A0 * B0\<bar> \<le> \<bar>A0 * B1 * (1 + A1*B0)\<bar> + \<bar>A0 * B0 * (1 + A1* B1)\<bar>" by simp
- also have "... = \<bar>A0 * B1\<bar> * \<bar>(1 + A1*B0)\<bar> + \<bar>A0 * B0\<bar> * \<bar>(1 + A1* B1)\<bar>" by (simp add: abs_mult)
- also have "... \<le> \<bar>(1 + A1*B0)\<bar> + \<bar>(1 + A1* B1)\<bar>"
- proof-
- have "\<bar>A0 * B1\<bar> * \<bar>(1 + A1*B0)\<bar> \<le> \<bar>(1 + A1*B0)\<bar>"
- using assms(1) mult_left_le_one_le[of "\<bar>(1 + A1*B0)\<bar>"] by simp
- moreover have "\<bar>A0 * B0\<bar> *\<bar>(1 + A1* B1)\<bar> \<le> \<bar>(1 + A1* B1)\<bar>"
- using assms mult_left_le_one_le[of "\<bar>(1 + A1*B1)\<bar>"] by simp
- ultimately show ?thesis by simp
- qed
- also have "... = 1 + A1*B0 + 1 + A1* B1" using assms by (simp add: abs_le_iff)
- also have "... = 2 + A1 * B0 + A1 * B1" by simp
- finally have pls: "\<bar>A0 * B1 - A0 * B0\<bar> \<le> 2 + A1 * B0 + A1 * B1" .
- have "A0 * B1 - A0 * B0 = A0 * B1 - A0 * B0 + A0 * B1 * A1 * B0 - A0 * B1 * A1 * B0" by simp
- also have "... = A0 * B1 * (1 - A1*B0) - A0 * B0 * (1 - A1* B1)"
- proof -
- have "A0 * (B1 - (B0 - A1 * (B1 * B0)) - A1 * (B1 * B0)) = A0 * (B1 - B0)"
- by fastforce
- then show ?thesis
- by (metis (no_types) add.commute calculation diff_diff_add mult.right_neutral
- vector_space_over_itself.scale_left_commute
- vector_space_over_itself.scale_right_diff_distrib vector_space_over_itself.scale_scale)
- qed
- finally have "A0 * B1 - A0 * B0 = A0 * B1 * (1 - A1*B0) - A0 * B0 * (1 - A1* B1)" .
- hence "\<bar>A0 * B1 - A0 * B0\<bar> \<le> \<bar>A0 * B1 * (1 - A1*B0)\<bar> + \<bar>A0 * B0 * (1 - A1* B1)\<bar>" by simp
- also have "... = \<bar>A0 * B1\<bar> * \<bar>(1 - A1*B0)\<bar> + \<bar>A0 * B0\<bar> * \<bar>(1 - A1* B1)\<bar>" by (simp add: abs_mult)
- also have "... \<le> \<bar>(1 - A1*B0)\<bar> + \<bar>(1 - A1* B1)\<bar>"
- proof-
- have "\<bar>A0 * B1\<bar> * \<bar>(1 - A1*B0)\<bar> \<le> \<bar>(1 - A1*B0)\<bar>"
- using assms(1) mult_left_le_one_le[of "\<bar>(1 - A1*B0)\<bar>"] by simp
- moreover have "\<bar>A0 * B0\<bar> *\<bar>(1 - A1* B1)\<bar> \<le> \<bar>(1 - A1* B1)\<bar>"
- using assms mult_left_le_one_le[of "\<bar>(1 - A1*B1)\<bar>"] by simp
- ultimately show ?thesis by simp
- qed
- also have "... = 1 - A1*B0 + 1 - A1* B1" using assms by (simp add: abs_le_iff)
- also have "... = 2 - A1 * B0 - A1 * B1" by simp
- finally have mns: "\<bar>A0 * B1 - A0 * B0\<bar> \<le> 2 - (A1 * B0 + A1 * B1)" by simp
- have ls: "\<bar>A0 * B1 - A0 * B0\<bar> \<le> 2 - \<bar>A1 * B0 + A1 * B1\<bar>"
- proof (cases "0 \<le> A1 * B0 + A1 * B1")
- case True
- then show ?thesis using mns by simp
- next
- case False
- then show ?thesis using pls by simp
- qed
- have "\<bar>A0 * B1 - A0 * B0 + A1 * B0 + A1 * B1\<bar> \<le> \<bar>A0 * B1 - A0 * B0\<bar> + \<bar>A1 * B0 + A1 * B1\<bar>"
- by simp
- also have "... \<le> 2" using ls by simp
- finally show ?thesis .
-qed
-
-lemma (in prob_space) chsh_expect:
- fixes A0::"'a \<Rightarrow> real"
- assumes "AE w in M. \<bar>A0 w\<bar> \<le> 1"
- and "AE w in M. \<bar>A1 w\<bar> \<le> 1"
- and "AE w in M. \<bar>B0 w\<bar> \<le> 1"
- and "AE w in M. \<bar>B1 w\<bar> \<le> 1"
-and "integrable M (\<lambda>w. A0 w * B1 w)"
-and "integrable M (\<lambda>w. A1 w * B1 w)"
-and "integrable M (\<lambda>w. A1 w * B0 w)"
-and "integrable M (\<lambda>w. A0 w * B0 w)"
-shows "\<bar>expectation (\<lambda>w. A1 w * B0 w) + expectation (\<lambda>w. A0 w *B1 w) +
- expectation (\<lambda>w. A1 w * B1 w) - expectation (\<lambda>w. A0 w * B0 w)\<bar> \<le> 2"
-proof -
- have exeq: "expectation (\<lambda>w. A1 w * B0 w) + expectation (\<lambda>w. A0 w * B1 w) +
- expectation (\<lambda>w. A1 w * B1 w) - expectation (\<lambda>w. A0 w * B0 w) =
- expectation (\<lambda>w. A0 w * B1 w - A0 w * B0 w + A1 w * B0 w + A1 w * B1 w)"
- using assms by auto
- have "\<bar>expectation (\<lambda>w. A0 w * B1 w - A0 w * B0 w + A1 w * B0 w + A1 w * B1 w)\<bar> \<le>
- expectation (\<lambda>w. \<bar>A0 w * B1 w - A0 w * B0 w + A1 w * B0 w + A1 w * B1 w\<bar>)"
- using integral_abs_bound by blast
- also have "... \<le> 2"
- proof (rule integral_le_const)
- show "AE w in M. \<bar>A0 w * B1 w - A0 w * B0 w + A1 w * B0 w + A1 w * B1 w\<bar> \<le> (2::real)"
- proof (rule AE_mp)
- show "AE w in M. \<bar>A0 w\<bar> \<le> 1 \<and> \<bar>A1 w\<bar> \<le> 1 \<and> \<bar>B0 w\<bar> \<le> 1 \<and> \<bar>B1 w\<bar> \<le> 1"
- using assms by simp
- show "AE w in M. \<bar>A0 w\<bar> \<le> 1 \<and> \<bar>A1 w\<bar> \<le> 1 \<and> \<bar>B0 w\<bar> \<le> 1 \<and> \<bar>B1 w\<bar> \<le> 1 \<longrightarrow>
- \<bar>A0 w * B1 w - A0 w * B0 w + A1 w * B0 w + A1 w * B1 w\<bar> \<le> 2"
- proof
- fix w
- assume "w\<in> space M"
- show "\<bar>A0 w\<bar> \<le> 1 \<and> \<bar>A1 w\<bar> \<le> 1 \<and> \<bar>B0 w\<bar> \<le> 1 \<and> \<bar>B1 w\<bar> \<le> 1 \<longrightarrow>
- \<bar>A0 w * B1 w - A0 w * B0 w + A1 w * B0 w + A1 w * B1 w\<bar> \<le> 2"
- proof
- assume ineq: "\<bar>A0 w\<bar> \<le> 1 \<and> \<bar>A1 w\<bar> \<le> 1 \<and> \<bar>B0 w\<bar> \<le> 1 \<and> \<bar>B1 w\<bar> \<le> 1"
- show "\<bar>A0 w * B1 w - A0 w * B0 w + A1 w * B0 w + A1 w * B1 w\<bar> \<le> 2"
- proof (rule chsh_real)
- show "\<bar>A0 w * B1 w\<bar> \<le> 1" using ineq by (simp add: abs_mult mult_le_one)
- show "\<bar>A0 w * B0 w\<bar> \<le> 1" using ineq by (simp add: abs_mult mult_le_one)
- show "\<bar>A1 w * B1 w\<bar> \<le> 1" using ineq by (simp add: abs_mult mult_le_one)
- show "\<bar>A1 w * B0 w\<bar> \<le> 1" using ineq by (simp add: abs_mult mult_le_one)
- qed
- qed
- qed
- qed
- show "integrable M (\<lambda>x. \<bar>A0 x * B1 x - A0 x * B0 x + A1 x * B0 x + A1 x * B1 x\<bar>)"
- proof (rule Bochner_Integration.integrable_abs)
- show "integrable M (\<lambda>x. A0 x * B1 x - A0 x * B0 x + A1 x * B0 x + A1 x * B1 x)"
- using assms by auto
- qed
- qed
- finally show ?thesis using exeq by simp
-qed
-
-text \<open>The local hidden variable assumption states that separated quantum measurements are
-independent. It is standard for this assumption to be stated in a context where the hidden
-variable admits a density; it is stated here in a more general contest involving expectations,
-with no assumption on the existence of a density.\<close>
-
-definition pos_rv:: "'a measure \<Rightarrow> ('a \<Rightarrow> real) \<Rightarrow> bool" where
-"pos_rv M Xr \<equiv> Xr \<in> borel_measurable M \<and> (AE w in M. (0::real) \<le> Xr w)"
-
-definition prv_sum:: "'a measure \<Rightarrow> complex Matrix.mat \<Rightarrow> (complex \<Rightarrow> 'a \<Rightarrow> real) \<Rightarrow> bool" where
-"prv_sum M A Xr \<equiv> (AE w in M. (\<Sum> a\<in> spectrum A. Xr a w) = 1)"
-
-definition (in cpx_sq_mat) lhv where
-"lhv M A B R XA XB \<equiv>
- prob_space M \<and>
- (\<forall>a \<in>spectrum A. pos_rv M (XA a)) \<and>
- (prv_sum M A XA) \<and>
- (\<forall>b \<in>spectrum B. pos_rv M (XB b)) \<and>
- (prv_sum M B XB) \<and>
- (\<forall>a \<in>spectrum A . \<forall>b \<in> spectrum B.
- (integrable M (\<lambda>w. XA a w * XB b w)) \<and>
- integral\<^sup>L M (\<lambda>w. XA a w * XB b w) =
- Re (Complex_Matrix.trace(eigen_projector A a * (eigen_projector B b) * R)))"
-
-(*definition (in cpx_sq_mat) lhv where
-"lhv M A B R XA XB \<equiv>
- prob_space M \<and>
- (\<forall>a \<in>spectrum A. (XA a \<in> borel_measurable M) \<and>
- (AE w in M. ((0::real) \<le> XA a w))) \<and>
- (AE w in M. (\<Sum> a\<in> spectrum A. XA a w) = 1) \<and>
- (\<forall>b \<in>spectrum B. (XB b \<in> borel_measurable M) \<and>
- (AE w in M. (0 \<le> XB b w))) \<and>
- (AE w in M. (\<Sum> b\<in> spectrum B. XB b w) = 1) \<and>
- (\<forall>a \<in>spectrum A . \<forall>b \<in> spectrum B.
- (integrable M (\<lambda>w. XA a w * XB b w)) \<and>
- integral\<^sup>L M (\<lambda>w. XA a w * XB b w) =
- Re (Complex_Matrix.trace(eigen_projector A a * (eigen_projector B b) * R)))"*)
-
-lemma (in cpx_sq_mat) lhv_posl:
- assumes "lhv M A B R XA XB"
- shows "AE w in M. (\<forall> a \<in> spectrum A. 0 \<le> XA a w)"
-proof (rule AE_ball_countable[THEN iffD2])
- show "countable (spectrum A)" using spectrum_finite[of A]
- by (simp add: countable_finite)
- show "\<forall>a\<in>spectrum A. AE w in M. 0 \<le> XA a w" using assms unfolding lhv_def pos_rv_def by simp
-qed
-
-lemma (in cpx_sq_mat) lhv_lt1_l:
- assumes "lhv M A B R XA XB"
- shows "AE w in M. (\<forall> a \<in> spectrum A. XA a w \<le> 1)"
-proof (rule AE_mp)
- show "AE w in M. (\<forall> a \<in> spectrum A. 0 \<le> XA a w) \<and> (\<Sum> a\<in> spectrum A. XA a w) = 1"
- using lhv_posl assms unfolding lhv_def pos_rv_def prv_sum_def by simp
- show "AE w in M. (\<forall>a\<in>spectrum A. 0 \<le> XA a w) \<and> (\<Sum>a\<in>spectrum A. XA a w) = 1 \<longrightarrow>
- (\<forall>a\<in>spectrum A. XA a w \<le> 1)"
- proof
- fix w
- assume "w\<in> space M"
- show "(\<forall>a\<in>spectrum A. 0 \<le> XA a w) \<and> (\<Sum>a\<in>spectrum A. XA a w) = 1 \<longrightarrow>
- (\<forall>a\<in>spectrum A. XA a w \<le> 1)"
- proof (rule impI)
- assume pr: "(\<forall>a\<in>spectrum A. 0 \<le> XA a w) \<and> (\<Sum>a\<in>spectrum A. XA a w) = 1"
- show "\<forall>a\<in>spectrum A. XA a w \<le> 1"
- proof
- fix a
- assume "a\<in> spectrum A"
- show "XA a w \<le> 1"
- proof (rule pos_sum_1_le[of "spectrum A"])
- show "finite (spectrum A)" using spectrum_finite[of A] by simp
- show "a \<in> spectrum A" using \<open>a \<in> spectrum A\<close> .
- show " \<forall>i\<in>spectrum A. 0 \<le> XA i w" using pr by simp
- show "(\<Sum>i\<in>spectrum A. XA i w) = 1" using pr by simp
- qed
- qed
- qed
- qed
-qed
-
-lemma (in cpx_sq_mat) lhv_posr:
- assumes "lhv M A B R XA XB"
- shows "AE w in M. (\<forall> b \<in> spectrum B. 0 \<le> XB b w)"
-proof (rule AE_ball_countable[THEN iffD2])
- show "countable (spectrum B)" using spectrum_finite[of B]
- by (simp add: countable_finite)
- show "\<forall>b\<in>spectrum B. AE w in M. 0 \<le> XB b w" using assms unfolding lhv_def pos_rv_def by simp
-qed
-
-lemma (in cpx_sq_mat) lhv_lt1_r:
- assumes "lhv M A B R XA XB"
- shows "AE w in M. (\<forall> a \<in> spectrum B. XB a w \<le> 1)"
-proof (rule AE_mp)
- show "AE w in M. (\<forall> a \<in> spectrum B. 0 \<le> XB a w) \<and> (\<Sum> a\<in> spectrum B. XB a w) = 1"
- using lhv_posr assms unfolding lhv_def prv_sum_def pos_rv_def by simp
- show "AE w in M. (\<forall>a\<in>spectrum B. 0 \<le> XB a w) \<and> (\<Sum>a\<in>spectrum B. XB a w) = 1 \<longrightarrow>
- (\<forall>a\<in>spectrum B. XB a w \<le> 1)"
- proof
- fix w
- assume "w\<in> space M"
- show "(\<forall>a\<in>spectrum B. 0 \<le> XB a w) \<and> (\<Sum>a\<in>spectrum B. XB a w) = 1 \<longrightarrow>
- (\<forall>a\<in>spectrum B. XB a w \<le> 1)"
- proof (rule impI)
- assume pr: "(\<forall>a\<in>spectrum B. 0 \<le> XB a w) \<and> (\<Sum>a\<in>spectrum B. XB a w) = 1"
- show "\<forall>a\<in>spectrum B. XB a w \<le> 1"
- proof
- fix a
- assume "a\<in> spectrum B"
- show "XB a w \<le> 1"
- proof (rule pos_sum_1_le[of "spectrum B"])
- show "finite (spectrum B)" using spectrum_finite[of B] by simp
- show "a \<in> spectrum B" using \<open>a \<in> spectrum B\<close> .
- show " \<forall>i\<in>spectrum B. 0 \<le> XB i w" using pr by simp
- show "(\<Sum>i\<in>spectrum B. XB i w) = 1" using pr by simp
- qed
- qed
- qed
- qed
-qed
-
-lemma (in cpx_sq_mat) lhv_AE_propl:
- assumes "lhv M A B R XA XB"
- shows "AE w in M. (\<forall> a \<in> spectrum A. 0 \<le> XA a w \<and> XA a w \<le> 1) \<and> (\<Sum> a\<in> spectrum A. XA a w) = 1"
-proof (rule AE_conjI)
- show "AE w in M. \<forall>a\<in>spectrum A. 0 \<le> XA a w \<and> XA a w \<le> 1"
- proof (rule AE_mp)
- show "AE w in M. (\<forall> a \<in> spectrum A. 0 \<le> XA a w) \<and> (\<forall> a \<in> spectrum A. XA a w \<le> 1)"
- using assms lhv_posl[of M A] lhv_lt1_l[of M A] by simp
- show "AE w in M. (\<forall>a\<in>spectrum A. 0 \<le> XA a w) \<and> (\<forall>a\<in>spectrum A. XA a w \<le> 1) \<longrightarrow>
- (\<forall>a\<in>spectrum A. 0 \<le> XA a w \<and> XA a w \<le> 1)" by auto
- qed
- show "AE w in M. (\<Sum>a\<in>spectrum A. XA a w) = 1" using assms unfolding lhv_def prv_sum_def
- by simp
-qed
-
-lemma (in cpx_sq_mat) lhv_AE_propr:
- assumes "lhv M A B R XA XB"
- shows "AE w in M. (\<forall> a \<in> spectrum B. 0 \<le> XB a w \<and> XB a w \<le> 1) \<and> (\<Sum> a\<in> spectrum B. XB a w) = 1"
-proof (rule AE_conjI)
- show "AE w in M. \<forall>a\<in>spectrum B. 0 \<le> XB a w \<and> XB a w \<le> 1"
- proof (rule AE_mp)
- show "AE w in M. (\<forall> a \<in> spectrum B. 0 \<le> XB a w) \<and> (\<forall> a \<in> spectrum B. XB a w \<le> 1)"
- using assms lhv_posr[of M _ B] lhv_lt1_r[of M _ B] by simp
- show "AE w in M. (\<forall>a\<in>spectrum B. 0 \<le> XB a w) \<and> (\<forall>a\<in>spectrum B. XB a w \<le> 1) \<longrightarrow>
- (\<forall>a\<in>spectrum B. 0 \<le> XB a w \<and> XB a w \<le> 1)" by auto
- qed
- show "AE w in M. (\<Sum>a\<in>spectrum B. XB a w) = 1" using assms unfolding lhv_def prv_sum_def
- by simp
-qed
-
-lemma (in cpx_sq_mat) lhv_integral_eq:
- fixes c::real
- assumes "lhv M A B R XA XB"
- and "a\<in> spectrum A"
-and "b\<in> spectrum B"
-shows "integral\<^sup>L M (\<lambda>w. XA a w * XB b w) =
- Re (Complex_Matrix.trace(eigen_projector A a * (eigen_projector B b) * R))"
- using assms unfolding lhv_def by simp
-
-lemma (in cpx_sq_mat) lhv_integrable:
- fixes c::real
- assumes "lhv M A B R XA XB"
- and "a\<in> spectrum A"
-and "b\<in> spectrum B"
-shows "integrable M (\<lambda>w. XA a w * XB b w)" using assms unfolding lhv_def by simp
-
-lemma (in cpx_sq_mat) lhv_scal_integrable:
- fixes c::real
- assumes "lhv M A B R XA XB"
- and "a\<in> spectrum A"
-and "b\<in> spectrum B"
-shows "integrable M (\<lambda>w. c *XA a w * d * XB b w)"
-proof -
- {
- fix x
- assume "x\<in> space M"
- have "c * d * (XA a x * XB b x) = c * XA a x * d * XB b x" by simp
- } note eq = this
- have "integrable M (\<lambda>w. XA a w * XB b w)" using assms unfolding lhv_def by simp
- hence g:"integrable M (\<lambda>w. c * d * (XA a w * XB b w))" using integrable_real_mult_right by simp
- show ?thesis
- proof (rule Bochner_Integration.integrable_cong[THEN iffD2], simp)
- show "integrable M (\<lambda>w. c * d * (XA a w * XB b w))" using g .
- show "\<And>x. x \<in> space M \<Longrightarrow> c * XA a x * d * XB b x = c * d * (XA a x * XB b x)" using eq by simp
- qed
-qed
-
-lemma (in cpx_sq_mat) lhv_lsum_scal_integrable:
- assumes "lhv M A B R XA XB"
- and "a\<in> spectrum A"
-shows "integrable M (\<lambda>x. \<Sum>b\<in>spectrum B. c * XA a x * (f b) * XB b x)"
-proof (rule Bochner_Integration.integrable_sum)
- fix b
- assume "b\<in> spectrum B"
- thus "integrable M (\<lambda>x. c * XA a x * f b *XB b x)" using \<open>a\<in> spectrum A\<close> assms
- lhv_scal_integrable[of M] by simp
-qed
-
-lemma (in cpx_sq_mat) lhv_sum_integrable:
- assumes "lhv M A B R XA XB"
-shows "integrable M (\<lambda>w. (\<Sum> a \<in> spectrum A. (\<Sum> b \<in> spectrum B. f a * XA a w * g b * XB b w)))"
-proof (rule Bochner_Integration.integrable_sum)
- fix a
- assume "a\<in> spectrum A"
- thus "integrable M (\<lambda>x. \<Sum>b\<in>spectrum B. f a * XA a x * g b * XB b x)"
- using assms lhv_lsum_scal_integrable[of M]
- by simp
-qed
-
-lemma (in cpx_sq_mat) spectrum_abs_1_weighted_suml:
- assumes "lhv M A B R Va Vb"
-and "{Re x |x. x \<in> spectrum A} \<noteq> {}"
- and "{Re x |x. x \<in> spectrum A} \<subseteq> {- 1, 1}"
-and "hermitian A"
- and "A\<in> fc_mats"
-shows "AE w in M. \<bar>(\<Sum>a\<in>spectrum A. Re a * Va a w)\<bar> \<le> 1"
-proof (rule AE_mp)
- show "AE w in M. (\<forall>a\<in>spectrum A. 0 \<le> Va a w \<and> Va a w \<le> 1) \<and> (\<Sum>a\<in>spectrum A. Va a w) = 1"
- using assms lhv_AE_propl[of M A B _ Va] by simp
- show "AE w in M. (\<forall>a\<in>spectrum A. 0 \<le> Va a w \<and> Va a w \<le> 1) \<and> (\<Sum>a\<in>spectrum A. Va a w) = 1 \<longrightarrow>
- \<bar>\<Sum>a\<in>spectrum A. Re a * Va a w\<bar> \<le> 1"
- proof
- fix w
- assume "w\<in> space M"
- show "(\<forall>a\<in>spectrum A. 0 \<le> Va a w \<and> Va a w \<le> 1) \<and> (\<Sum>a\<in>spectrum A. Va a w) = 1 \<longrightarrow>
- \<bar>\<Sum>a\<in>spectrum A. Re a * Va a w\<bar> \<le> 1"
- proof (intro impI)
- assume pr: "(\<forall>a\<in>spectrum A. 0 \<le> Va a w \<and> Va a w \<le> 1) \<and> (\<Sum>a\<in>spectrum A. Va a w) = 1"
- show "\<bar>(\<Sum>a\<in>spectrum A. Re a * Va a w)\<bar> \<le> 1"
- proof (cases "{Re x |x. x \<in> spectrum A} = {- 1, 1}")
- case True
- hence sp: "spectrum A = {-1, 1}" using hermitian_Re_spectrum[of A] assms by simp
- hence scsum: "(\<Sum>a\<in>spectrum A. Re a * Va a w) = Va 1 w - Va (-1) w"
- using sum_2_elems by simp
- have sum: "(\<Sum>a\<in>spectrum A. Va a w) = Va (-1) w + Va 1 w"
- using sp sum_2_elems by simp
- have "\<bar>Va 1 w - Va (-1) w\<bar> \<le> 1"
- proof (rule fct_bound')
- have "1 \<in> spectrum A" using sp by simp
- thus "0 \<le> Va 1 w" using pr by simp
- have "-1 \<in> spectrum A" using sp by simp
- thus "0 \<le> Va (- 1) w" using pr by simp
- show "Va (- 1) w + Va 1 w = 1" using pr sum by simp
- qed
- thus ?thesis using scsum by simp
- next
- case False
- then show ?thesis
- proof (cases "{Re x |x. x \<in> spectrum A} = {- 1}")
- case True
- hence "spectrum A = {-1}" using hermitian_Re_spectrum[of A] assms by simp
- hence "(\<Sum>a\<in>spectrum A. Re a * Va a w) = - Va (-1) w" by simp
- moreover have "-1 \<in> spectrum A" using \<open>spectrum A = {-1}\<close> by simp
- ultimately show ?thesis using pr by simp
- next
- case False
- hence "{Re x |x. x \<in> spectrum A} = {1}" using assms \<open>{Re x |x. x \<in> spectrum A} \<noteq> {-1, 1}\<close>
- last_subset[of "{Re x |x. x \<in> spectrum A}"] by simp
- hence "spectrum A = {1}" using hermitian_Re_spectrum[of A] assms by simp
- hence "(\<Sum>a\<in>spectrum A. Re a * Va a w) = Va 1 w" by simp
- moreover have "1 \<in> spectrum A" using \<open>spectrum A = {1}\<close> by simp
- ultimately show ?thesis using pr by simp
- qed
- qed
- qed
- qed
-qed
-
-lemma (in cpx_sq_mat) spectrum_abs_1_weighted_sumr:
- assumes "lhv M B A R Vb Va"
-and "{Re x |x. x \<in> spectrum A} \<noteq> {}"
- and "{Re x |x. x \<in> spectrum A} \<subseteq> {- 1, 1}"
-and "hermitian A"
- and "A\<in> fc_mats"
-shows "AE w in M. \<bar>(\<Sum>a\<in>spectrum A. Re a * Va a w)\<bar> \<le> 1"
-proof (rule AE_mp)
- show "AE w in M. (\<forall>a\<in>spectrum A. 0 \<le> Va a w \<and> Va a w \<le> 1) \<and> (\<Sum>a\<in>spectrum A. Va a w) = 1"
- using assms lhv_AE_propr[of M B A _ Vb] by simp
- show "AE w in M. (\<forall>a\<in>spectrum A. 0 \<le> Va a w \<and> Va a w \<le> 1) \<and> (\<Sum>a\<in>spectrum A. Va a w) = 1 \<longrightarrow>
- \<bar>\<Sum>a\<in>spectrum A. Re a * Va a w\<bar> \<le> 1"
- proof
- fix w
- assume "w\<in> space M"
- show "(\<forall>a\<in>spectrum A. 0 \<le> Va a w \<and> Va a w \<le> 1) \<and> (\<Sum>a\<in>spectrum A. Va a w) = 1 \<longrightarrow>
- \<bar>\<Sum>a\<in>spectrum A. Re a * Va a w\<bar> \<le> 1"
- proof (intro impI)
- assume pr: "(\<forall>a\<in>spectrum A. 0 \<le> Va a w \<and> Va a w \<le> 1) \<and> (\<Sum>a\<in>spectrum A. Va a w) = 1"
- show "\<bar>(\<Sum>a\<in>spectrum A. Re a * Va a w)\<bar> \<le> 1"
- proof (cases "{Re x |x. x \<in> spectrum A} = {- 1, 1}")
- case True
- hence sp: "spectrum A = {-1, 1}" using hermitian_Re_spectrum[of A] assms by simp
- hence scsum: "(\<Sum>a\<in>spectrum A. Re a * Va a w) = Va 1 w - Va (-1) w"
- using sum_2_elems by simp
- have sum: "(\<Sum>a\<in>spectrum A. Va a w) = Va (-1) w + Va 1 w"
- using sp sum_2_elems by simp
- have "\<bar>Va 1 w - Va (-1) w\<bar> \<le> 1"
- proof (rule fct_bound')
- have "1 \<in> spectrum A" using sp by simp
- thus "0 \<le> Va 1 w" using pr by simp
- have "-1 \<in> spectrum A" using sp by simp
- thus "0 \<le> Va (- 1) w" using pr by simp
- show "Va (- 1) w + Va 1 w = 1" using pr sum by simp
- qed
- thus ?thesis using scsum by simp
- next
- case False
- then show ?thesis
- proof (cases "{Re x |x. x \<in> spectrum A} = {- 1}")
- case True
- hence "spectrum A = {-1}" using hermitian_Re_spectrum[of A] assms by simp
- hence "(\<Sum>a\<in>spectrum A. Re a * Va a w) = - Va (-1) w" by simp
- moreover have "-1 \<in> spectrum A" using \<open>spectrum A = {-1}\<close> by simp
- ultimately show ?thesis using pr by simp
- next
- case False
- hence "{Re x |x. x \<in> spectrum A} = {1}" using assms \<open>{Re x |x. x \<in> spectrum A} \<noteq> {-1, 1}\<close>
- last_subset[of "{Re x |x. x \<in> spectrum A}"] by simp
- hence "spectrum A = {1}" using hermitian_Re_spectrum[of A] assms by simp
- hence "(\<Sum>a\<in>spectrum A. Re a * Va a w) = Va 1 w" by simp
- moreover have "1 \<in> spectrum A" using \<open>spectrum A = {1}\<close> by simp
- ultimately show ?thesis using pr by simp
- qed
- qed
- qed
- qed
-qed
-
-definition qt_expect where
-"qt_expect A Va = (\<lambda>w. (\<Sum>a\<in>spectrum A. Re a * Va a w))"
-
-lemma (in cpx_sq_mat) spectr_sum_integrable:
-assumes "lhv M A B R Va Vb"
-shows "integrable M (\<lambda>w. qt_expect A Va w * qt_expect B Vb w)"
-proof (rule Bochner_Integration.integrable_cong[THEN iffD2])
- show "M = M" by simp
- show "\<And>w. w \<in> space M \<Longrightarrow> qt_expect A Va w * qt_expect B Vb w =
- (\<Sum>a\<in>spectrum A. (\<Sum>b\<in>spectrum B. Re a * Va a w * Re b * Vb b w))"
- proof -
- fix w
- assume "w\<in> space M"
- show "qt_expect A Va w * qt_expect B Vb w =
- (\<Sum>a\<in>spectrum A. (\<Sum>b\<in>spectrum B. Re a * Va a w * Re b * Vb b w))" unfolding qt_expect_def
- by (metis (mono_tags, lifting) Finite_Cartesian_Product.sum_cong_aux sum_product
- vector_space_over_itself.scale_scale)
- qed
- show "integrable M (\<lambda>w. \<Sum>a\<in>spectrum A. (\<Sum>b\<in>spectrum B. Re a * Va a w * Re b * Vb b w))"
- using lhv_sum_integrable[of M] assms by simp
-qed
-
-lemma (in cpx_sq_mat) lhv_sum_integral':
- assumes "lhv M A B R XA XB"
-shows "integral\<^sup>L M (\<lambda>w. (\<Sum> a \<in> spectrum A. f a * XA a w) * (\<Sum> b \<in> spectrum B. g b * XB b w)) =
- (\<Sum> a \<in> spectrum A. f a * (\<Sum> b \<in> spectrum B. g b * integral\<^sup>L M (\<lambda>w. XA a w * XB b w)))"
-proof -
- have "integral\<^sup>L M (\<lambda>w. (\<Sum> a \<in> spectrum A. f a * XA a w) * (\<Sum> b \<in> spectrum B. g b * XB b w)) =
- integral\<^sup>L M (\<lambda>w. (\<Sum> a \<in> spectrum A. (\<Sum> b \<in> spectrum B. f a * XA a w * g b * XB b w)))"
- proof (rule Bochner_Integration.integral_cong, simp)
- fix w
- assume "w\<in> space M"
- show "(\<Sum>a\<in>spectrum A. f a * XA a w) * (\<Sum>b\<in>spectrum B. g b * XB b w) =
- (\<Sum>a\<in>spectrum A. (\<Sum>b\<in>spectrum B. f a * XA a w * (g b) * XB b w))"
- by (simp add: sum_product vector_space_over_itself.scale_scale)
- qed
- also have "... = (\<Sum> a \<in> spectrum A.
- integral\<^sup>L M (\<lambda>w. (\<Sum> b \<in> spectrum B. f a * XA a w * g b * XB b w)))"
- proof (rule Bochner_Integration.integral_sum)
- fix a
- assume "a\<in> spectrum A"
- thus "integrable M (\<lambda>x. \<Sum>b\<in>spectrum B. f a * XA a x * g b * XB b x)"
- using lhv_lsum_scal_integrable[of M] assms by simp
- qed
- also have "... = (\<Sum> a \<in> spectrum A. f a *
- integral\<^sup>L M (\<lambda>w. (\<Sum> b \<in> spectrum B. XA a w * g b * XB b w)))"
- proof -
- have "\<forall> a\<in> spectrum A. integral\<^sup>L M (\<lambda>w. (\<Sum> b \<in> spectrum B. f a * XA a w * g b * XB b w)) =
- f a * integral\<^sup>L M (\<lambda>w. (\<Sum> b \<in> spectrum B. XA a w * g b * XB b w))"
- proof
- fix a
- assume "a\<in> spectrum A"
- have "(LINT w|M. (\<Sum>b\<in>spectrum B. f a * XA a w * g b * XB b w)) =
- (LINT w|M. f a* (\<Sum>b\<in>spectrum B. XA a w * g b * XB b w))"
- proof (rule Bochner_Integration.integral_cong, simp)
- fix x
- assume "x \<in> space M"
- show "(\<Sum>b\<in>spectrum B. f a * XA a x * g b * XB b x) =
- f a * (\<Sum>b\<in>spectrum B. XA a x * g b * XB b x)"
- by (metis (no_types, lifting) Finite_Cartesian_Product.sum_cong_aux
- vector_space_over_itself.scale_scale vector_space_over_itself.scale_sum_right)
- qed
- also have "... = f a * (LINT w|M. (\<Sum>b\<in>spectrum B. XA a w * g b * XB b w))" by simp
- finally show "(LINT w|M. (\<Sum>b\<in>spectrum B. f a * XA a w * g b * XB b w)) =
- f a * (LINT w|M. (\<Sum>b\<in>spectrum B. XA a w * g b * XB b w))" .
- qed
- thus ?thesis by simp
- qed
- also have "... = (\<Sum> a \<in> spectrum A. f a * (\<Sum> b \<in> spectrum B. g b *
- integral\<^sup>L M (\<lambda>w. XA a w * XB b w)))"
- proof (rule sum.cong, simp)
- fix a
- assume "a\<in> spectrum A"
- have "integral\<^sup>L M (\<lambda>w. (\<Sum> b \<in> spectrum B. XA a w * g b * XB b w)) = (\<Sum> b \<in> spectrum B.
- integral\<^sup>L M (\<lambda>w. XA a w * g b * XB b w))"
- proof (rule Bochner_Integration.integral_sum)
- show "\<And>b. b \<in> spectrum B \<Longrightarrow> integrable M (\<lambda>x. XA a x * g b * XB b x)"
- proof -
- fix b
- assume "b\<in> spectrum B"
- thus "integrable M (\<lambda>x. XA a x * g b * XB b x)"
- using assms lhv_scal_integrable[of M _ _ _ _ _ a b 1] \<open>a\<in> spectrum A\<close> by simp
- qed
- qed
- also have "... = (\<Sum> b \<in> spectrum B. g b * integral\<^sup>L M (\<lambda>w. XA a w * XB b w))"
- proof (rule sum.cong, simp)
- fix x
- assume "x\<in> spectrum B"
- have "LINT w|M. XA a w * g x * XB x w = LINT w|M. g x * (XA a w * XB x w)"
- by (rule Bochner_Integration.integral_cong, auto)
- also have "... = g x * (LINT w|M. XA a w * XB x w)"
- using Bochner_Integration.integral_mult_right_zero[of M "g x" "\<lambda>w. XA a w * XB x w"]
- by simp
- finally show "LINT w|M. XA a w * g x * XB x w = g x * (LINT w|M. XA a w * XB x w)" .
- qed
- finally have "integral\<^sup>L M (\<lambda>w. (\<Sum> b \<in> spectrum B. XA a w * g b * XB b w)) =
- (\<Sum> b \<in> spectrum B. g b * integral\<^sup>L M (\<lambda>w. XA a w * XB b w))" .
- thus "f a * (LINT w|M. (\<Sum>b\<in>spectrum B. XA a w * g b * XB b w)) =
- f a * (\<Sum> b \<in> spectrum B. g b * integral\<^sup>L M (\<lambda>w. XA a w * XB b w))" by simp
- qed
- finally show ?thesis .
-qed
-
-lemma (in cpx_sq_mat) sum_qt_expect_trace:
- assumes "lhv M A B R XA XB"
- shows "(\<Sum> a \<in> spectrum A. f a * (\<Sum> b \<in> spectrum B. g b * integral\<^sup>L M (\<lambda>w. XA a w * XB b w))) =
- (\<Sum> a \<in> spectrum A. f a * (\<Sum> b \<in> spectrum B. g b *
- Re (Complex_Matrix.trace(eigen_projector A a * (eigen_projector B b) * R))))"
-proof (rule sum.cong, simp)
- fix a
- assume "a\<in> spectrum A"
- have "(\<Sum>b\<in>spectrum B. g b * (LINT w|M. XA a w * XB b w)) =
- (\<Sum>b\<in>spectrum B. g b *
- Re (Complex_Matrix.trace (eigen_projector A a * eigen_projector B b * R)))"
- proof (rule sum.cong, simp)
- fix b
- assume "b\<in> spectrum B"
- show "g b * (LINT w|M. XA a w * XB b w) =
- g b * Re (Complex_Matrix.trace (eigen_projector A a * eigen_projector B b * R))"
- using lhv_integral_eq[of M] assms \<open>a \<in> spectrum A\<close> \<open>b \<in> spectrum B\<close> by simp
- qed
- thus "f a * (\<Sum>b\<in>spectrum B. g b * (LINT w|M. XA a w * XB b w)) =
- f a * (\<Sum>b\<in>spectrum B. g b *
- Re (Complex_Matrix.trace (eigen_projector A a * eigen_projector B b * R)))" by simp
-qed
-
-lemma (in cpx_sq_mat) sum_eigen_projector_trace_dist:
- assumes "hermitian B"
-and "A\<in> fc_mats"
-and "B\<in> fc_mats"
-and "R\<in> fc_mats"
- shows "(\<Sum> b \<in> spectrum B. (b *
- Complex_Matrix.trace(A * (eigen_projector B b) * R))) = Complex_Matrix.trace(A * B * R)"
-proof -
- have "(\<Sum>b\<in>spectrum B. b * Complex_Matrix.trace (A * eigen_projector B b * R)) =
- (\<Sum>b\<in>spectrum B. Complex_Matrix.trace (A * (b \<cdot>\<^sub>m (eigen_projector B b)) * R))"
- proof (rule sum.cong, simp)
- fix b
- assume "b\<in> spectrum B"
- have "b * Complex_Matrix.trace (A * eigen_projector B b * R) =
- Complex_Matrix.trace (b \<cdot>\<^sub>m (A * eigen_projector B b * R))"
- proof (rule trace_smult[symmetric])
- show "A * eigen_projector B b * R \<in> carrier_mat dimR dimR" using eigen_projector_carrier
- assms fc_mats_carrier dim_eq \<open>b \<in> spectrum B\<close> cpx_sq_mat_mult by meson
- qed
- also have "... = Complex_Matrix.trace (A * (b \<cdot>\<^sub>m eigen_projector B b) * R)"
- proof -
- have "b \<cdot>\<^sub>m (A * eigen_projector B b * R) = b \<cdot>\<^sub>m (A * (eigen_projector B b * R))"
- proof -
- have "A * eigen_projector B b * R = A * (eigen_projector B b * R)"
- by (metis \<open>b \<in> spectrum B\<close> assms(1) assms(2) assms(3) assms(4) assoc_mult_mat dim_eq
- fc_mats_carrier eigen_projector_carrier)
- thus ?thesis by simp
- qed
- also have "... = A * (b \<cdot>\<^sub>m (eigen_projector B b * R))"
- proof (rule mult_smult_distrib[symmetric])
- show "A \<in> carrier_mat dimR dimR" using eigen_projector_carrier assms
- fc_mats_carrier dim_eq by simp
- show "eigen_projector B b * R \<in> carrier_mat dimR dimR" using eigen_projector_carrier
- \<open>b \<in> spectrum B\<close> assms fc_mats_carrier dim_eq cpx_sq_mat_mult by blast
- qed
- also have "... = A * ((b \<cdot>\<^sub>m eigen_projector B b) * R)"
- proof -
- have "b \<cdot>\<^sub>m (eigen_projector B b * R) = (b \<cdot>\<^sub>m eigen_projector B b) * R"
- proof (rule mult_smult_assoc_mat[symmetric])
- show "eigen_projector B b \<in> carrier_mat dimR dimR" using eigen_projector_carrier
- \<open>b \<in> spectrum B\<close> assms fc_mats_carrier dim_eq by simp
- show "R \<in> carrier_mat dimR dimR" using assms fc_mats_carrier dim_eq by simp
- qed
- thus ?thesis by simp
- qed
- also have "... = A * (b \<cdot>\<^sub>m eigen_projector B b) * R"
- by (metis \<open>b \<in> spectrum B\<close> assms(1) assms(2) assms(3) assms(4) assoc_mult_mat
- cpx_sq_mat_smult dim_eq fc_mats_carrier eigen_projector_carrier)
- finally have "b \<cdot>\<^sub>m (A * eigen_projector B b * R) = A * (b \<cdot>\<^sub>m eigen_projector B b) * R" .
- then show ?thesis by simp
- qed
- finally show "b * Complex_Matrix.trace (A * eigen_projector B b * R) =
- Complex_Matrix.trace (A * (b \<cdot>\<^sub>m eigen_projector B b) * R)" .
- qed
- also have "... = Complex_Matrix.trace (A *
- (sum_mat (\<lambda>b. b \<cdot>\<^sub>m eigen_projector B b) (spectrum B)) * R)"
- proof (rule trace_sum_mat_mat_distrib, (auto simp add: assms))
- show "finite (spectrum B)" using spectrum_finite[of B] by simp
- fix b
- assume "b\<in> spectrum B"
- show "b \<cdot>\<^sub>m eigen_projector B b \<in> fc_mats"
- by (simp add: \<open>b \<in> spectrum B\<close> assms(1) assms(3) cpx_sq_mat_smult eigen_projector_carrier)
- qed
- also have "... = Complex_Matrix.trace (A * B * R)"
- proof -
- have "sum_mat (\<lambda>b. b \<cdot>\<^sub>m eigen_projector B b) (spectrum B) = B" using make_pm_sum' assms by simp
- thus ?thesis by simp
- qed
- finally show ?thesis .
-qed
-
-lemma (in cpx_sq_mat) sum_eigen_projector_trace_right:
- assumes "hermitian A"
-and "A\<in> fc_mats"
-and "B\<in> fc_mats"
-shows "(\<Sum> a \<in> spectrum A. Complex_Matrix.trace (a \<cdot>\<^sub>m eigen_projector A a * B)) =
- Complex_Matrix.trace (A * B)"
-proof -
- have "sum_mat (\<lambda>a. a \<cdot>\<^sub>m eigen_projector A a * B) (spectrum A) =
- sum_mat (\<lambda>a. a \<cdot>\<^sub>m eigen_projector A a) (spectrum A) * B"
- proof (rule mult_sum_mat_distrib_right)
- show "finite (spectrum A)" using spectrum_finite[of A] by simp
- show "\<And>a. a \<in> spectrum A \<Longrightarrow> a \<cdot>\<^sub>m eigen_projector A a \<in> fc_mats"
- using assms(1) assms(2) cpx_sq_mat_smult eigen_projector_carrier by blast
- show "B\<in> fc_mats" using assms by simp
- qed
- also have "... = A * B" using make_pm_sum' assms by simp
- finally have seq: "sum_mat (\<lambda>a. a \<cdot>\<^sub>m eigen_projector A a * B) (spectrum A) = A * B" .
- have "(\<Sum> a \<in> spectrum A. Complex_Matrix.trace (a \<cdot>\<^sub>m eigen_projector A a * B)) =
- Complex_Matrix.trace (sum_mat (\<lambda>a. a \<cdot>\<^sub>m eigen_projector A a * B) (spectrum A))"
- proof (rule trace_sum_mat[symmetric])
- show "finite (spectrum A)" using spectrum_finite[of A] by simp
- show "\<And>a. a \<in> spectrum A \<Longrightarrow> a \<cdot>\<^sub>m eigen_projector A a * B \<in> fc_mats"
- by (simp add: assms(1) assms(2) assms(3) cpx_sq_mat_mult cpx_sq_mat_smult
- eigen_projector_carrier)
- qed
- also have "... = Complex_Matrix.trace (A * B)" using seq by simp
- finally show ?thesis .
-qed
-
-lemma (in cpx_sq_mat) sum_eigen_projector_trace:
- assumes "hermitian A"
- and "hermitian B"
- and "A\<in> fc_mats"
- and "B\<in> fc_mats"
-and "R\<in> fc_mats"
- shows "(\<Sum> a \<in> spectrum A. a * (\<Sum> b \<in> spectrum B. (b *
- Complex_Matrix.trace(eigen_projector A a * (eigen_projector B b) * R)))) =
- Complex_Matrix.trace(A * B * R)"
-proof -
- have "(\<Sum> a \<in> spectrum A. a * (\<Sum> b \<in> spectrum B. (b *
- Complex_Matrix.trace(eigen_projector A a * (eigen_projector B b) * R)))) = (\<Sum> a \<in> spectrum A.
- Complex_Matrix.trace (a \<cdot>\<^sub>m eigen_projector A a * (B * R)))"
- proof (rule sum.cong, simp)
- fix a
- assume "a\<in> spectrum A"
- hence "(\<Sum>b\<in>spectrum B. b *
- Complex_Matrix.trace (eigen_projector A a * eigen_projector B b * R)) =
- Complex_Matrix.trace (eigen_projector A a * B * R)" using
- sum_eigen_projector_trace_dist[of B "eigen_projector A a" R] assms eigen_projector_carrier
- by blast
- hence "a * (\<Sum> b \<in> spectrum B. (b *
- Complex_Matrix.trace(eigen_projector A a * (eigen_projector B b) * R))) =
- a * Complex_Matrix.trace (eigen_projector A a * B * R)" by simp
- also have "... = Complex_Matrix.trace (a \<cdot>\<^sub>m (eigen_projector A a * B * R))"
- using trace_smult[symmetric, of "eigen_projector A a * B * R" dimR a] assms
- \<open>a \<in> spectrum A\<close> cpx_sq_mat_mult dim_eq fc_mats_carrier eigen_projector_carrier by force
- also have "... = Complex_Matrix.trace (a \<cdot>\<^sub>m eigen_projector A a * (B * R))"
- proof -
- have "a \<cdot>\<^sub>m (eigen_projector A a * B * R) = a \<cdot>\<^sub>m (eigen_projector A a * B) * R"
- proof (rule mult_smult_assoc_mat[symmetric])
- show "R\<in> carrier_mat dimR dimR" using assms fc_mats_carrier dim_eq by simp
- show "eigen_projector A a * B \<in> carrier_mat dimR dimR" using assms eigen_projector_carrier
- cpx_sq_mat_mult fc_mats_carrier dim_eq \<open>a \<in> spectrum A\<close> by blast
- qed
- also have "... = a \<cdot>\<^sub>m eigen_projector A a * B * R"
- proof -
- have "a \<cdot>\<^sub>m (eigen_projector A a * B) = a \<cdot>\<^sub>m eigen_projector A a * B"
- using mult_smult_assoc_mat[symmetric]
- proof -
- show ?thesis
- by (metis \<open>\<And>nr nc n k B A. \<lbrakk>A \<in> carrier_mat nr n; B \<in> carrier_mat n nc\<rbrakk> \<Longrightarrow>
- k \<cdot>\<^sub>m (A * B) = k \<cdot>\<^sub>m A * B\<close> \<open>a \<in> spectrum A\<close> assms(1) assms(3) assms(4) dim_eq
- fc_mats_carrier eigen_projector_carrier)
- qed
- thus ?thesis by simp
- qed
- also have "... = a \<cdot>\<^sub>m eigen_projector A a * (B * R)"
- by (metis \<open>a \<in> spectrum A\<close> assms(1) assms(3) assms(4) assms(5) assoc_mult_mat
- cpx_sq_mat_smult dim_eq fc_mats_carrier eigen_projector_carrier)
- finally show ?thesis by simp
- qed
- finally show "a * (\<Sum> b \<in> spectrum B. (b *
- Complex_Matrix.trace(eigen_projector A a * (eigen_projector B b) * R))) =
- Complex_Matrix.trace (a \<cdot>\<^sub>m eigen_projector A a * (B * R))" .
- qed
- also have "... = Complex_Matrix.trace (A * (B * R))"
- using sum_eigen_projector_trace_right[of A "B * R"] assms by (simp add: cpx_sq_mat_mult)
- also have "... = Complex_Matrix.trace (A * B * R)"
- proof -
- have "A * (B * R) = A * B * R"
- proof (rule assoc_mult_mat[symmetric])
- show "A\<in> carrier_mat dimR dimR" using assms fc_mats_carrier dim_eq by simp
- show "B\<in> carrier_mat dimR dimR" using assms fc_mats_carrier dim_eq by simp
- show "R\<in> carrier_mat dimR dimR" using assms fc_mats_carrier dim_eq by simp
- qed
- thus ?thesis by simp
- qed
- finally show ?thesis .
-qed
-
-text \<open>We obtain the main result of this part, which relates the quantum expectation value of a
-joint measurement with an expectation.\<close>
-
-lemma (in cpx_sq_mat) sum_qt_expect:
- assumes "lhv M A B R XA XB"
- and "A\<in> fc_mats"
- and "B\<in> fc_mats"
- and "R\<in> fc_mats"
- and "hermitian A"
- and "hermitian B"
- shows "integral\<^sup>L M (\<lambda>w. (qt_expect A XA w) * (qt_expect B XB w)) =
- Re (Complex_Matrix.trace(A * B * R))"
-proof -
- have br: "\<forall> b \<in> spectrum B. b\<in> Reals" using assms hermitian_spectrum_real[of B] by auto
- have ar: "\<forall>a \<in> spectrum A. a\<in> Reals" using hermitian_spectrum_real[of A] assms by auto
- have "integral\<^sup>L M (\<lambda>w. (\<Sum> a \<in> spectrum A. Re a* XA a w) * (\<Sum> b \<in> spectrum B. Re b *XB b w)) =
- (\<Sum> a \<in> spectrum A. Re a * (\<Sum> b \<in> spectrum B. Re b * integral\<^sup>L M (\<lambda>w. XA a w * XB b w)))"
- using lhv_sum_integral'[of M] assms by simp
- also have "... = (\<Sum> a \<in> spectrum A. Re a * (\<Sum> b \<in> spectrum B. Re b *
- Re (Complex_Matrix.trace(eigen_projector A a * (eigen_projector B b) * R))))"
- using assms sum_qt_expect_trace[of M] by simp
- also have "... = (\<Sum> a \<in> spectrum A. Re a * Re (\<Sum> b \<in> spectrum B. (b *
- Complex_Matrix.trace(eigen_projector A a * (eigen_projector B b) * R))))"
- proof (rule sum.cong, simp)
- fix a
- assume "a\<in> spectrum A"
- have "(\<Sum>b\<in>spectrum B. Re b *
- Re (Complex_Matrix.trace (eigen_projector A a * eigen_projector B b * R))) =
- (\<Sum> b \<in> spectrum B. Re (b *
- Complex_Matrix.trace(eigen_projector A a * (eigen_projector B b) * R)))"
- proof (rule sum.cong, simp)
- fix b
- assume "b\<in> spectrum B"
- hence "b\<in> Reals" using hermitian_spectrum_real[of B] assms by simp
- hence "Re b = b" by simp
- thus "Re b * Re (Complex_Matrix.trace (eigen_projector A a * eigen_projector B b * R)) =
- Re (b * Complex_Matrix.trace (eigen_projector A a * eigen_projector B b * R))"
- using hermitian_spectrum_real using \<open>b \<in> \<real>\<close> mult_real_cpx by auto
- qed
- also have "... =
- Re (\<Sum> b \<in> spectrum B. b *
- Complex_Matrix.trace(eigen_projector A a * (eigen_projector B b) * R))" by simp
- finally have "(\<Sum>b\<in>spectrum B. Re b *
- Re (Complex_Matrix.trace (eigen_projector A a * eigen_projector B b * R))) =
- Re (\<Sum> b \<in> spectrum B. b *
- Complex_Matrix.trace(eigen_projector A a * (eigen_projector B b) * R))" .
- thus "Re a * (\<Sum>b\<in>spectrum B. Re b *
- Re (Complex_Matrix.trace (eigen_projector A a * eigen_projector B b * R))) =
- Re a * Re (\<Sum>b\<in>spectrum B.
- (b * Complex_Matrix.trace (eigen_projector A a * eigen_projector B b * R)))"
- by simp
- qed
- also have "... = (\<Sum> a \<in> spectrum A. Re (a * (\<Sum> b \<in> spectrum B. (b *
- Complex_Matrix.trace(eigen_projector A a * (eigen_projector B b) * R)))))"
- proof (rule sum.cong, simp)
- fix x
- assume "x\<in> spectrum A"
- hence "Re x = x" using ar by simp
- thus "Re x * Re (\<Sum>b\<in>spectrum B. b *
- Complex_Matrix.trace (eigen_projector A x * eigen_projector B b * R)) =
- Re (x * (\<Sum>b\<in>spectrum B. b *
- Complex_Matrix.trace (eigen_projector A x * eigen_projector B b * R)))"
- using \<open>x \<in> spectrum A\<close> ar mult_real_cpx by auto
- qed
- also have "... = Re (\<Sum> a \<in> spectrum A. a * (\<Sum> b \<in> spectrum B. (b *
- Complex_Matrix.trace(eigen_projector A a * (eigen_projector B b) * R))))" by simp
- also have "... = Re (Complex_Matrix.trace(A *B * R))" using assms
- sum_eigen_projector_trace[of A B] by simp
- finally show ?thesis unfolding qt_expect_def .
-qed
-
-
-subsection \<open>Properties of specific observables\<close>
-
-text \<open>In this part we consider a specific density operator and specific observables corresponding
-to joint bipartite measurements. We will compute the quantum expectation value of this system and
-prove that it violates the CHSH inequality, thus proving that the local hidden variable assumption
-cannot hold.\<close>
-
-subsubsection \<open>Ket 0, Ket 1 and the corresponding projectors\<close>
-
-definition ket_0::"complex Matrix.vec" where
-"ket_0 = unit_vec 2 0"
-
-lemma ket_0_dim:
- shows "dim_vec ket_0 = 2" unfolding ket_0_def by simp
-
-lemma ket_0_norm:
- shows "\<parallel>ket_0\<parallel> = 1" using unit_cpx_vec_length unfolding ket_0_def by simp
-
-lemma ket_0_mat:
- shows "ket_vec ket_0 = Matrix.mat_of_cols_list 2 [[1, 0]]"
- by (auto simp add: ket_vec_def Matrix.mat_of_cols_list_def ket_0_def)
-
-definition ket_1::"complex Matrix.vec" where
-"ket_1 = unit_vec 2 1"
-
-lemma ket_1_dim:
- shows "dim_vec ket_1 = 2" unfolding ket_1_def by simp
-
-lemma ket_1_norm:
- shows "\<parallel>ket_1\<parallel> = 1" using unit_cpx_vec_length unfolding ket_1_def by simp
-
-definition ket_01
- where "ket_01 = tensor_vec ket_0 ket_1"
-
-lemma ket_01_dim:
- shows "dim_vec ket_01 = 4" unfolding ket_01_def
- by (simp add: ket_0_dim ket_1_dim)
-
-definition ket_10
- where "ket_10 = tensor_vec ket_1 ket_0"
-
-lemma ket_10_dim:
- shows "dim_vec ket_10 = 4" unfolding ket_10_def by (simp add: ket_0_dim ket_1_dim)
-
-text \<open>We define \verb+ket_psim+, one of the Bell states (or EPR pair).\<close>
-
-definition ket_psim where
-"ket_psim = 1/(sqrt 2) \<cdot>\<^sub>v (ket_01 - ket_10)"
-
-lemma ket_psim_dim:
- shows "dim_vec ket_psim = 4" using ket_01_dim ket_10_dim unfolding ket_psim_def by simp
-
-lemma ket_psim_norm:
- shows "\<parallel>ket_psim\<parallel> = 1"
-proof -
- have "dim_vec ket_psim = 2\<^sup>2" unfolding ket_psim_def ket_01_def ket_10_def ket_0_def ket_1_def
- by simp
- moreover have "(\<Sum>i<4. (cmod (vec_index ket_psim i))\<^sup>2) = 1"
- apply (auto simp add: ket_psim_def ket_01_def ket_10_def ket_0_def ket_1_def)
- apply (simp add: sum_4_elems)
- done
- ultimately show ?thesis by (simp add: cpx_vec_length_def)
-qed
-
-text \<open>\verb+rho_psim+ represents the density operator associated with the quantum
-state \verb+ket_psim+.\<close>
-
-definition rho_psim where
-"rho_psim = rank_1_proj ket_psim"
-
-lemma rho_psim_carrier:
- shows "rho_psim \<in> carrier_mat 4 4" using rank_1_proj_carrier[of ket_psim] ket_psim_dim
- rho_psim_def by simp
-
-lemma rho_psim_dim_row:
- shows "dim_row rho_psim = 4" using rho_psim_carrier by simp
-
-lemma rho_psim_density:
- shows "density_operator rho_psim" unfolding density_operator_def
-proof
- show "Complex_Matrix.positive rho_psim" using rank_1_proj_positive[of ket_psim] ket_psim_norm
- rho_psim_def by simp
- show "Complex_Matrix.trace rho_psim = 1" using rank_1_proj_trace[of ket_psim] ket_psim_norm
- rho_psim_def by simp
-qed
-
-
-subsubsection \<open>The X and Z matrices and two of their combinations\<close>
-
-text \<open>In this part we prove properties of two standard matrices in quantum theory, $X$ and $Z$,
-as well as two of their combinations: $\frac{X+Z}{\sqrt{2}}$ and $\frac{Z - X}{\sqrt{2}}$.
-Note that all of these matrices are observables, they will be used to violate the CHSH inequality.\<close>
-
-lemma Z_carrier: shows "Z \<in> carrier_mat 2 2" unfolding Z_def by simp
-
-lemma Z_hermitian:
- shows "hermitian Z" using dagger_adjoint dagger_of_Z unfolding hermitian_def by simp
-
-lemma unitary_Z:
- shows "Complex_Matrix.unitary Z"
-proof -
- have "Complex_Matrix.adjoint Z * Z = (1\<^sub>m 2)" using dagger_adjoint[of Z] by simp
- thus ?thesis unfolding unitary_def
- by (metis Complex_Matrix.adjoint_adjoint Complex_Matrix.unitary_def Z_carrier adjoint_dim
- carrier_matD(1) inverts_mat_def unitary_adjoint)
-qed
-
-lemma X_carrier: shows "X \<in> carrier_mat 2 2" unfolding X_def by simp
-
-lemma X_hermitian:
- shows "hermitian X" using dagger_adjoint dagger_of_X unfolding hermitian_def by simp
-
-lemma unitary_X:
- shows "Complex_Matrix.unitary X"
-proof -
- have "Complex_Matrix.adjoint X * X = (1\<^sub>m 2)" using dagger_adjoint[of X] by simp
- thus ?thesis unfolding unitary_def
- by (metis Complex_Matrix.adjoint_adjoint Complex_Matrix.unitary_def X_carrier adjoint_dim
- carrier_matD(1) inverts_mat_def unitary_adjoint)
-qed
-
-definition XpZ
- where "XpZ = -1/sqrt(2) \<cdot>\<^sub>m (X + Z)"
-
-lemma XpZ_carrier:
- shows "XpZ \<in> carrier_mat 2 2" unfolding XpZ_def X_def Z_def by simp
-
-lemma XpZ_hermitian:
- shows "hermitian XpZ"
-proof -
- have "X + Z \<in> carrier_mat 2 2" using Z_carrier X_carrier by simp
- moreover have "hermitian (X + Z)" using X_hermitian Z_hermitian hermitian_add Matrix.mat_carrier
- unfolding X_def Z_def by blast
- ultimately show ?thesis using hermitian_smult[of "X + Z" 2 "- 1 / sqrt 2"] unfolding XpZ_def
- by auto
-qed
-
-lemma XpZ_inv:
- "XpZ * XpZ = 1\<^sub>m 2" unfolding XpZ_def X_def Z_def times_mat_def one_mat_def
- apply (rule cong_mat, simp+)
- apply (auto simp add: Matrix.scalar_prod_def)
- apply (auto simp add: Gates.csqrt_2_sq)
- done
-
-lemma unitary_XpZ:
- shows "Complex_Matrix.unitary XpZ"
-proof -
- have "Complex_Matrix.adjoint XpZ * XpZ = (1\<^sub>m 2)" using XpZ_inv XpZ_hermitian
- unfolding hermitian_def by simp
- thus ?thesis unfolding unitary_def
- by (metis Complex_Matrix.adjoint_adjoint Complex_Matrix.unitary_def XpZ_carrier adjoint_dim
- carrier_matD(1) inverts_mat_def unitary_adjoint)
-qed
-
-definition ZmX
- where "ZmX = 1/sqrt(2) \<cdot>\<^sub>m (Z - X)"
-
-lemma ZmX_carrier:
- shows "ZmX \<in> carrier_mat 2 2" unfolding ZmX_def X_def Z_def
- by (simp add: minus_carrier_mat')
-
-lemma ZmX_hermitian:
- shows "hermitian ZmX"
-proof -
- have "Z - X \<in> carrier_mat 2 2" unfolding X_def Z_def
- by (simp add: minus_carrier_mat)
- moreover have "hermitian (Z - X)" using X_hermitian Z_hermitian hermitian_minus Matrix.mat_carrier
- unfolding X_def Z_def by blast
- ultimately show ?thesis using hermitian_smult[of "Z - X" 2 "1 / sqrt 2"] unfolding ZmX_def
- by auto
-qed
-
-lemma ZmX_inv:
- "ZmX * ZmX = 1\<^sub>m 2" unfolding ZmX_def X_def Z_def times_mat_def one_mat_def
- apply (rule cong_mat, simp+)
- apply (auto simp add: Matrix.scalar_prod_def)
- apply (auto simp add: Gates.csqrt_2_sq)
- done
-
-lemma unitary_ZmX:
- shows "Complex_Matrix.unitary ZmX"
-proof -
- have "Complex_Matrix.adjoint ZmX * ZmX = (1\<^sub>m 2)" using ZmX_inv ZmX_hermitian
- unfolding hermitian_def by simp
- thus ?thesis unfolding unitary_def
- by (metis Complex_Matrix.adjoint_adjoint Complex_Matrix.unitary_def ZmX_carrier adjoint_dim
- carrier_matD(1) inverts_mat_def unitary_adjoint)
-qed
-
-definition Z_XpZ
- where "Z_XpZ = tensor_mat Z XpZ"
-
-lemma Z_XpZ_carrier:
- shows "Z_XpZ \<in> carrier_mat 4 4" unfolding Z_XpZ_def using tensor_mat_carrier XpZ_carrier
- Z_carrier by auto
-
-definition X_XpZ
- where "X_XpZ = tensor_mat X XpZ"
-
-lemma X_XpZ_carrier:
- shows "X_XpZ \<in> carrier_mat 4 4" using tensor_mat_carrier XpZ_carrier X_carrier
- unfolding X_XpZ_def by auto
-
-definition Z_ZmX
- where "Z_ZmX = tensor_mat Z ZmX"
-
-lemma Z_ZmX_carrier:
- shows "Z_ZmX \<in> carrier_mat 4 4" using tensor_mat_carrier ZmX_carrier Z_carrier
- unfolding Z_ZmX_def by auto
-
-definition X_ZmX
- where "X_ZmX = tensor_mat X ZmX"
-
-lemma X_ZmX_carrier:
- shows "X_ZmX \<in> carrier_mat 4 4" using tensor_mat_carrier X_carrier ZmX_carrier
- unfolding X_ZmX_def by auto
-
-lemma X_ZmX_rho_psim[simp]:
- shows "Complex_Matrix.trace (rho_psim * X_ZmX) = 1/ (sqrt 2)"
- apply (auto simp add: ket_10_def ket_1_def X_ZmX_def ZmX_def X_def ket_01_def
- Z_def rho_psim_def ket_psim_def rank_1_proj_def outer_prod_def ket_0_def)
- apply (auto simp add: Complex_Matrix.trace_def)
- apply (simp add: sum_4_elems)
- apply (simp add: csqrt_2_sq)
- done
-
-lemma Z_ZmX_rho_psim[simp]:
- shows "Complex_Matrix.trace (rho_psim * Z_ZmX) = -1/ (sqrt 2)"
- apply (auto simp add: rho_psim_def ket_psim_def ket_10_def)
- apply (auto simp add: Z_ZmX_def Z_def ZmX_def X_def)
- apply (auto simp add: rank_1_proj_def outer_prod_def ket_01_def ket_1_def ket_0_def)
- apply (auto simp add: Complex_Matrix.trace_def sum_4_elems)
- apply (simp add: csqrt_2_sq)
- done
-
-lemma X_XpZ_rho_psim[simp]:
- shows "Complex_Matrix.trace (rho_psim * X_XpZ) =1/ (sqrt 2)"
- apply (auto simp add: rho_psim_def ket_psim_def ket_10_def)
- apply (auto simp add: X_XpZ_def Z_def XpZ_def X_def)
- apply (auto simp add: rank_1_proj_def outer_prod_def ket_01_def ket_1_def ket_0_def)
- apply (auto simp add: Complex_Matrix.trace_def sum_4_elems)
- apply (simp add: csqrt_2_sq)
- done
-
-lemma Z_XpZ_rho_psim[simp]:
- shows "Complex_Matrix.trace (rho_psim * Z_XpZ) =1/ (sqrt 2)"
- apply (auto simp add: rho_psim_def ket_psim_def ket_10_def)
- apply (auto simp add: Z_XpZ_def XpZ_def X_def Z_def)
- apply (auto simp add: rank_1_proj_def outer_prod_def ket_01_def ket_1_def ket_0_def)
- apply (auto simp add: Complex_Matrix.trace_def sum_4_elems)
- apply (simp add: csqrt_2_sq)
- done
-
-definition Z_I where
-"Z_I = tensor_mat Z (1\<^sub>m 2)"
-
-lemma Z_I_carrier:
- shows "Z_I \<in> carrier_mat 4 4" using tensor_mat_carrier Z_carrier unfolding Z_I_def by auto
-
-lemma Z_I_hermitian:
- shows "hermitian Z_I" unfolding Z_I_def using tensor_mat_hermitian[of Z 2 "1\<^sub>m 2" 2]
- by (simp add: Z_carrier Z_hermitian hermitian_one)
-
-lemma Z_I_unitary:
- shows "unitary Z_I" unfolding Z_I_def using tensor_mat_unitary[of Z "1\<^sub>m 2"] Z_carrier unitary_Z
- using unitary_one by auto
-
-lemma Z_I_spectrum:
- shows "{Re x |x. x \<in> spectrum Z_I} \<subseteq> {- 1, 1}" using unitary_hermitian_Re_spectrum Z_I_hermitian
- Z_I_unitary Z_I_carrier by simp
-
-definition X_I where
-"X_I = tensor_mat X (1\<^sub>m 2)"
-
-lemma X_I_carrier:
- shows "X_I \<in> carrier_mat 4 4" using tensor_mat_carrier X_carrier unfolding X_I_def by auto
-
-lemma X_I_hermitian:
- shows "hermitian X_I" unfolding X_I_def using tensor_mat_hermitian[of X 2 "1\<^sub>m 2" 2]
- by (simp add: X_carrier X_hermitian hermitian_one)
-
-lemma X_I_unitary:
- shows "unitary X_I" unfolding X_I_def using tensor_mat_unitary[of X "1\<^sub>m 2"] X_carrier unitary_X
- using unitary_one by auto
-
-lemma X_I_spectrum:
- shows "{Re x |x. x \<in> spectrum X_I} \<subseteq> {- 1, 1}" using unitary_hermitian_Re_spectrum X_I_hermitian
- X_I_unitary X_I_carrier by simp
-
-definition I_XpZ where
-"I_XpZ = tensor_mat (1\<^sub>m 2) XpZ"
-
-lemma I_XpZ_carrier:
- shows "I_XpZ \<in> carrier_mat 4 4" using tensor_mat_carrier XpZ_carrier unfolding I_XpZ_def by auto
-
-lemma I_XpZ_hermitian:
- shows "hermitian I_XpZ" unfolding I_XpZ_def using tensor_mat_hermitian[of "1\<^sub>m 2" 2 XpZ 2]
- by (simp add: XpZ_carrier XpZ_hermitian hermitian_one)
-
-lemma I_XpZ_unitary:
- shows "unitary I_XpZ" unfolding I_XpZ_def using tensor_mat_unitary[of "1\<^sub>m 2" XpZ] XpZ_carrier
- unitary_XpZ using unitary_one by auto
-
-lemma I_XpZ_spectrum:
- shows "{Re x |x. x \<in> spectrum I_XpZ} \<subseteq> {- 1, 1}" using unitary_hermitian_Re_spectrum
- I_XpZ_hermitian I_XpZ_unitary I_XpZ_carrier by simp
-
-definition I_ZmX where
-"I_ZmX = tensor_mat (1\<^sub>m 2) ZmX"
-
-lemma I_ZmX_carrier:
- shows "I_ZmX \<in> carrier_mat 4 4" using tensor_mat_carrier ZmX_carrier unfolding I_ZmX_def by auto
-
-lemma I_ZmX_hermitian:
- shows "hermitian I_ZmX" unfolding I_ZmX_def using tensor_mat_hermitian[of "1\<^sub>m 2" 2 ZmX 2]
- by (simp add: ZmX_carrier ZmX_hermitian hermitian_one)
-
-lemma I_ZmX_unitary:
- shows "unitary I_ZmX" unfolding I_ZmX_def using tensor_mat_unitary[of "1\<^sub>m 2" ZmX] ZmX_carrier
- unitary_ZmX using unitary_one by auto
-
-lemma I_ZmX_spectrum:
- shows "{Re x |x. x \<in> spectrum I_ZmX} \<subseteq> {- 1, 1}" using unitary_hermitian_Re_spectrum I_ZmX_hermitian
- I_ZmX_unitary I_ZmX_carrier by simp
-
-lemma X_I_ZmX_eq:
- shows "X_I * I_ZmX = X_ZmX" unfolding X_I_def I_ZmX_def X_ZmX_def using mult_distr_tensor
- by (metis (no_types, lifting) X_inv ZmX_inv index_mult_mat(2) index_mult_mat(3) index_one_mat(2)
- index_one_mat(3) left_mult_one_mat' pos2 right_mult_one_mat')
-
-lemma X_I_XpZ_eq:
- shows "X_I * I_XpZ = X_XpZ" unfolding X_I_def I_XpZ_def X_XpZ_def using mult_distr_tensor
- by (metis (no_types, lifting) X_inv XpZ_inv index_mult_mat(2) index_mult_mat(3) index_one_mat(2)
- index_one_mat(3) left_mult_one_mat' pos2 right_mult_one_mat')
-
-lemma Z_I_XpZ_eq:
- shows "Z_I * I_XpZ = Z_XpZ" unfolding Z_I_def I_XpZ_def Z_XpZ_def using mult_distr_tensor
- by (metis (no_types, lifting) Z_inv XpZ_inv index_mult_mat(2) index_mult_mat(3) index_one_mat(2)
- index_one_mat(3) left_mult_one_mat' pos2 right_mult_one_mat')
-
-lemma Z_I_ZmX_eq:
- shows "Z_I * I_ZmX = Z_ZmX" unfolding Z_I_def I_ZmX_def Z_ZmX_def using mult_distr_tensor
- by (metis (no_types, lifting) Z_inv ZmX_inv index_mult_mat(2) index_mult_mat(3) index_one_mat(2)
- index_one_mat(3) left_mult_one_mat' pos2 right_mult_one_mat')
-
-
-subsubsection \<open>No local hidden variable\<close>
-
-text \<open>We show that the local hidden variable hypothesis cannot hold by exhibiting a quantum
-expectation value that is greater than the upper-bound givven by the CHSH inequality.\<close>
-
-locale bin_cpx = cpx_sq_mat +
- assumes dim4: "dimR = 4"
-
-lemma (in bin_cpx) X_I_XpZ_trace:
- assumes "lhv M X_I I_XpZ R Vx Vp"
- and "R\<in> fc_mats"
- shows "LINT w|M. (qt_expect X_I Vx w) * (qt_expect I_XpZ Vp w) =
- Re (Complex_Matrix.trace (R * X_XpZ))"
-proof -
- have "LINT w|M. (qt_expect X_I Vx w) * (qt_expect I_XpZ Vp w) =
- Re (Complex_Matrix.trace (X_I * I_XpZ * R))"
- proof (rule sum_qt_expect, (simp add: assms))
- show "X_I \<in> fc_mats" using X_I_carrier dim4 fc_mats_carrier dim_eq by simp
- show "R\<in> fc_mats" using assms by simp
- show "I_XpZ \<in> fc_mats" using I_XpZ_carrier dim4 fc_mats_carrier dim_eq by simp
- show "hermitian X_I" using X_I_hermitian .
- show "hermitian I_XpZ" using I_XpZ_hermitian .
- qed
- also have "... = Re (Complex_Matrix.trace (X_XpZ * R))" using X_I_XpZ_eq by simp
- also have "... = Re (Complex_Matrix.trace (R * X_XpZ))"
- proof -
- have "Complex_Matrix.trace (X_XpZ * R) = Complex_Matrix.trace (R * X_XpZ)"
- using trace_comm[of X_XpZ 4 R] X_XpZ_carrier assms dim4 fc_mats_carrier dim_eq by simp
- thus ?thesis by simp
- qed
- finally show ?thesis .
-qed
-
-lemma (in bin_cpx) X_I_XpZ_chsh:
- assumes "lhv M X_I I_XpZ rho_psim Vx Vp"
- shows "LINT w|M. (qt_expect X_I Vx w) * (qt_expect I_XpZ Vp w) =
- 1/sqrt 2"
-proof -
- have "LINT w|M. (qt_expect X_I Vx w) * (qt_expect I_XpZ Vp w) =
- Re (Complex_Matrix.trace (rho_psim * X_XpZ))"
- proof (rule X_I_XpZ_trace, (simp add: assms))
- show "rho_psim \<in> fc_mats" using rho_psim_carrier fc_mats_carrier dim_eq dim4 by simp
- qed
- also have "... = 1/sqrt 2" using X_XpZ_rho_psim by simp
- finally show ?thesis .
-qed
-
-lemma (in bin_cpx) Z_I_XpZ_trace:
- assumes "lhv M Z_I I_XpZ R Vz Vp"
- and "R\<in> fc_mats"
- shows "LINT w|M. (qt_expect Z_I Vz w) * (qt_expect I_XpZ Vp w) =
- Re (Complex_Matrix.trace (R * Z_XpZ))"
-proof -
- have "LINT w|M. (qt_expect Z_I Vz w) * (qt_expect I_XpZ Vp w) =
- Re (Complex_Matrix.trace (Z_I * I_XpZ * R))"
- proof (rule sum_qt_expect, (simp add: assms))
- show "Z_I \<in> fc_mats" using Z_I_carrier dim4 fc_mats_carrier dim_eq by simp
- show "R\<in> fc_mats" using assms by simp
- show "I_XpZ \<in> fc_mats" using I_XpZ_carrier dim4 fc_mats_carrier dim_eq by simp
- show "hermitian Z_I" using Z_I_hermitian .
- show "hermitian I_XpZ" using I_XpZ_hermitian .
- qed
- also have "... = Re (Complex_Matrix.trace (Z_XpZ * R))" using Z_I_XpZ_eq by simp
- also have "... = Re (Complex_Matrix.trace (R * Z_XpZ))"
- proof -
- have "Complex_Matrix.trace (Z_XpZ * R) = Complex_Matrix.trace (R * Z_XpZ)"
- using trace_comm[of Z_XpZ 4 R] Z_XpZ_carrier assms dim4 fc_mats_carrier dim_eq by simp
- thus ?thesis by simp
- qed
- finally show ?thesis .
-qed
-
-lemma (in bin_cpx) Z_I_XpZ_chsh:
- assumes "lhv M Z_I I_XpZ rho_psim Vz Vp"
- shows "LINT w|M. (qt_expect Z_I Vz w) * (qt_expect I_XpZ Vp w) =
- 1/sqrt 2"
-proof -
- have "LINT w|M. (qt_expect Z_I Vz w) * (qt_expect I_XpZ Vp w) =
- Re (Complex_Matrix.trace (rho_psim * Z_XpZ))"
- proof (rule Z_I_XpZ_trace, (simp add: assms))
- show "rho_psim \<in> fc_mats" using rho_psim_carrier fc_mats_carrier dim_eq dim4 by simp
- qed
- also have "... = 1/sqrt 2" using Z_XpZ_rho_psim by simp
- finally show ?thesis unfolding qt_expect_def .
-qed
-
-lemma (in bin_cpx) X_I_ZmX_trace:
- assumes "lhv M X_I I_ZmX R Vx Vp"
- and "R\<in> fc_mats"
- shows "LINT w|M. (qt_expect X_I Vx w) * (qt_expect I_ZmX Vp w) =
- Re (Complex_Matrix.trace (R * X_ZmX))"
-proof -
- have "LINT w|M. (qt_expect X_I Vx w) * (qt_expect I_ZmX Vp w) =
- Re (Complex_Matrix.trace (X_I * I_ZmX * R))"
- proof (rule sum_qt_expect, (simp add: assms))
- show "X_I \<in> fc_mats" using X_I_carrier dim4 fc_mats_carrier dim_eq by simp
- show "R\<in> fc_mats" using assms by simp
- show "I_ZmX \<in> fc_mats" using I_ZmX_carrier dim4 fc_mats_carrier dim_eq by simp
- show "hermitian X_I" using X_I_hermitian .
- show "hermitian I_ZmX" using I_ZmX_hermitian .
- qed
- also have "... = Re (Complex_Matrix.trace (X_ZmX * R))" using X_I_ZmX_eq by simp
- also have "... = Re (Complex_Matrix.trace (R * X_ZmX))"
- proof -
- have "Complex_Matrix.trace (X_ZmX * R) = Complex_Matrix.trace (R * X_ZmX)"
- using trace_comm[of X_ZmX 4 R] X_ZmX_carrier assms dim4 fc_mats_carrier dim_eq by simp
- thus ?thesis by simp
- qed
- finally show ?thesis .
-qed
-
-lemma (in bin_cpx) X_I_ZmX_chsh:
- assumes "lhv M X_I I_ZmX rho_psim Vx Vp"
- shows "LINT w|M. (qt_expect X_I Vx w) * (qt_expect I_ZmX Vp w) =
- 1/sqrt 2"
-proof -
- have "LINT w|M. (qt_expect X_I Vx w) * (qt_expect I_ZmX Vp w) =
- Re (Complex_Matrix.trace (rho_psim * X_ZmX))"
- proof (rule X_I_ZmX_trace, (simp add: assms))
- show "rho_psim \<in> fc_mats" using rho_psim_carrier fc_mats_carrier dim_eq dim4 by simp
- qed
- also have "... = 1/sqrt 2" using X_ZmX_rho_psim by simp
- finally show ?thesis unfolding qt_expect_def .
-qed
-
-lemma (in bin_cpx) Z_I_ZmX_trace:
- assumes "lhv M Z_I I_ZmX R Vz Vp"
- and "R\<in> fc_mats"
- shows "LINT w|M. (qt_expect Z_I Vz w) * (qt_expect I_ZmX Vp w) =
- Re (Complex_Matrix.trace (R * Z_ZmX))"
-proof -
- have "LINT w|M. (qt_expect Z_I Vz w) * (qt_expect I_ZmX Vp w) =
- Re (Complex_Matrix.trace (Z_I * I_ZmX * R))"
- proof (rule sum_qt_expect, (simp add: assms))
- show "Z_I \<in> fc_mats" using Z_I_carrier dim4 fc_mats_carrier dim_eq by simp
- show "R\<in> fc_mats" using assms by simp
- show "I_ZmX \<in> fc_mats" using I_ZmX_carrier dim4 fc_mats_carrier dim_eq by simp
- show "hermitian Z_I" using Z_I_hermitian .
- show "hermitian I_ZmX" using I_ZmX_hermitian .
- qed
- also have "... = Re (Complex_Matrix.trace (Z_ZmX * R))" using Z_I_ZmX_eq by simp
- also have "... = Re (Complex_Matrix.trace (R * Z_ZmX))"
- proof -
- have "Complex_Matrix.trace (Z_ZmX * R) = Complex_Matrix.trace (R * Z_ZmX)"
- using trace_comm[of Z_ZmX 4 R] Z_ZmX_carrier assms dim4 fc_mats_carrier dim_eq by simp
- thus ?thesis by simp
- qed
- finally show ?thesis .
-qed
-
-lemma (in bin_cpx) Z_I_ZmX_chsh:
- assumes "lhv M Z_I I_ZmX rho_psim Vz Vp"
-shows "LINT w|M. (qt_expect Z_I Vz w) * (qt_expect I_ZmX Vp w) =
- -1/sqrt 2"
-proof -
- have "LINT w|M. (qt_expect Z_I Vz w) * (qt_expect I_ZmX Vp w) =
- Re (Complex_Matrix.trace (rho_psim * Z_ZmX))"
- proof (rule Z_I_ZmX_trace, (simp add: assms))
- show "rho_psim \<in> fc_mats" using rho_psim_carrier fc_mats_carrier dim_eq dim4 by simp
- qed
- also have "... = -1/sqrt 2" using Z_ZmX_rho_psim by simp
- finally show ?thesis unfolding qt_expect_def .
-qed
-
-lemma (in bin_cpx) chsh_upper_bound:
- assumes "prob_space M"
- and "lhv M X_I I_XpZ rho_psim Vx Vp"
- and "lhv M Z_I I_XpZ rho_psim Vz Vp"
- and "lhv M X_I I_ZmX rho_psim Vx Vm"
- and "lhv M Z_I I_ZmX rho_psim Vz Vm"
-shows "\<bar>(LINT w|M. qt_expect X_I Vx w * qt_expect I_ZmX Vm w) +
- (LINT w|M. qt_expect Z_I Vz w * qt_expect I_XpZ Vp w) +
- (LINT w|M. qt_expect X_I Vx w * qt_expect I_XpZ Vp w) -
- (LINT w|M. qt_expect Z_I Vz w * qt_expect I_ZmX Vm w)\<bar>
- \<le> 2"
-proof (rule prob_space.chsh_expect)
- show "AE w in M. \<bar>qt_expect X_I Vx w\<bar> \<le> 1" unfolding qt_expect_def
- proof (rule spectrum_abs_1_weighted_suml)
- show "X_I \<in> fc_mats" using X_I_carrier fc_mats_carrier dim_eq dim4 by simp
- show "hermitian X_I" using X_I_hermitian .
- show "lhv M X_I I_XpZ rho_psim Vx Vp" using assms by simp
- show "{Re x |x. x \<in> spectrum X_I} \<subseteq> {- 1, 1}" using X_I_spectrum by simp
- show "{Re x |x. x \<in> spectrum X_I} \<noteq> {}" using spectrum_ne X_I_hermitian \<open>X_I \<in> fc_mats\<close> by auto
- qed
- show "AE w in M. \<bar>qt_expect Z_I Vz w\<bar> \<le> 1" unfolding qt_expect_def
- proof (rule spectrum_abs_1_weighted_suml)
- show "Z_I \<in> fc_mats" using Z_I_carrier fc_mats_carrier dim_eq dim4 by simp
- show "hermitian Z_I" using Z_I_hermitian .
- show "lhv M Z_I I_XpZ rho_psim Vz Vp" using assms by simp
- show "{Re x |x. x \<in> spectrum Z_I} \<subseteq> {- 1, 1}" using Z_I_spectrum by simp
- show "{Re x |x. x \<in> spectrum Z_I} \<noteq> {}" using spectrum_ne Z_I_hermitian \<open>Z_I \<in> fc_mats\<close> by auto
- qed
- show "AE w in M. \<bar>qt_expect I_XpZ Vp w\<bar> \<le> 1" unfolding qt_expect_def
- proof (rule spectrum_abs_1_weighted_sumr)
- show "I_XpZ \<in> fc_mats" using I_XpZ_carrier fc_mats_carrier dim_eq dim4 by simp
- show "hermitian I_XpZ" using I_XpZ_hermitian .
- show "lhv M Z_I I_XpZ rho_psim Vz Vp" using assms by simp
- show "{Re x |x. x \<in> spectrum I_XpZ} \<subseteq> {- 1, 1}" using I_XpZ_spectrum by simp
- show "{Re x |x. x \<in> spectrum I_XpZ} \<noteq> {}" using spectrum_ne I_XpZ_hermitian \<open>I_XpZ \<in> fc_mats\<close>
- by auto
- qed
- show "AE w in M. \<bar>qt_expect I_ZmX Vm w\<bar> \<le> 1" unfolding qt_expect_def
- proof (rule spectrum_abs_1_weighted_sumr)
- show "I_ZmX \<in> fc_mats" using I_ZmX_carrier fc_mats_carrier dim_eq dim4 by simp
- show "hermitian I_ZmX" using I_ZmX_hermitian .
- show "lhv M Z_I I_ZmX rho_psim Vz Vm" using assms by simp
- show "{Re x |x. x \<in> spectrum I_ZmX} \<subseteq> {- 1, 1}" using I_ZmX_spectrum by simp
- show "{Re x |x. x \<in> spectrum I_ZmX} \<noteq> {}" using spectrum_ne I_ZmX_hermitian \<open>I_ZmX \<in> fc_mats\<close>
- by auto
- qed
- show "prob_space M" using assms by simp
- show "integrable M (\<lambda>w. qt_expect X_I Vx w * qt_expect I_ZmX Vm w)"
- using spectr_sum_integrable[of M] assms by simp
- show "integrable M (\<lambda>w. qt_expect Z_I Vz w * qt_expect I_ZmX Vm w)"
- using spectr_sum_integrable[of M] assms by simp
- show "integrable M (\<lambda>w. qt_expect X_I Vx w * qt_expect I_XpZ Vp w)"
- using spectr_sum_integrable[of M] assms by simp
- show "integrable M (\<lambda>w. qt_expect Z_I Vz w * qt_expect I_XpZ Vp w)"
- using spectr_sum_integrable[of M] assms by simp
-qed
-
-lemma (in bin_cpx) quantum_value:
- assumes "lhv M X_I I_XpZ rho_psim Vx Vp"
- and "lhv M Z_I I_XpZ rho_psim Vz Vp"
- and "lhv M X_I I_ZmX rho_psim Vx Vm"
- and "lhv M Z_I I_ZmX rho_psim Vz Vm"
-shows "\<bar>(LINT w|M. qt_expect X_I Vx w * qt_expect I_ZmX Vm w) +
- (LINT w|M. qt_expect Z_I Vz w * qt_expect I_XpZ Vp w) +
- (LINT w|M. qt_expect X_I Vx w * qt_expect I_XpZ Vp w) -
- (LINT w|M. qt_expect Z_I Vz w * qt_expect I_ZmX Vm w)\<bar>
- = 2* sqrt 2"
-proof -
- have "LINT w|M. qt_expect X_I Vx w * qt_expect I_ZmX Vm w = 1/sqrt 2"
- using X_I_ZmX_chsh[of M] assms unfolding qt_expect_def by simp
- moreover have b: "LINT w|M. qt_expect X_I Vx w * qt_expect I_XpZ Vp w = 1/sqrt 2"
- using X_I_XpZ_chsh[of M] assms unfolding qt_expect_def by simp
- moreover have c: "LINT w|M. qt_expect Z_I Vz w * qt_expect I_XpZ Vp w = 1/sqrt 2"
- using Z_I_XpZ_chsh[of M] assms unfolding qt_expect_def by simp
- moreover have "LINT w|M. qt_expect Z_I Vz w * qt_expect I_ZmX Vm w = -1/sqrt 2"
- using Z_I_ZmX_chsh[of M] assms unfolding qt_expect_def by simp
- ultimately have "(LINT w|M. qt_expect X_I Vx w * qt_expect I_ZmX Vm w) +
- (LINT w|M. qt_expect X_I Vx w * qt_expect I_XpZ Vp w) +
- (LINT w|M. qt_expect Z_I Vz w * qt_expect I_XpZ Vp w) -
- (LINT w|M. qt_expect Z_I Vz w * qt_expect I_ZmX Vm w) = 4/(sqrt 2)" by simp
- also have "... = (4 * (sqrt 2))/(sqrt 2 * (sqrt 2))"
- by (metis mult_numeral_1_right real_divide_square_eq times_divide_eq_right)
- also have "... = (4 * (sqrt 2)) / 2" by simp
- also have "... = 2 * (sqrt 2)" by simp
- finally have "(LINT w|M. qt_expect X_I Vx w * qt_expect I_ZmX Vm w) +
- (LINT w|M. qt_expect X_I Vx w * qt_expect I_XpZ Vp w) +
- (LINT w|M. qt_expect Z_I Vz w * qt_expect I_XpZ Vp w) -
- (LINT w|M. qt_expect Z_I Vz w * qt_expect I_ZmX Vm w) = 2 * sqrt 2" .
- thus ?thesis by (simp add: b c)
-qed
-
-lemma (in bin_cpx) no_lhv:
- assumes "lhv M X_I I_XpZ rho_psim Vx Vp"
- and "lhv M Z_I I_XpZ rho_psim Vz Vp"
- and "lhv M X_I I_ZmX rho_psim Vx Vm"
- and "lhv M Z_I I_ZmX rho_psim Vz Vm"
-shows False
-proof -
- have "prob_space M" using assms unfolding lhv_def by simp
- have "2 * sqrt 2 = \<bar>(LINT w|M. qt_expect X_I Vx w * qt_expect I_ZmX Vm w) +
- (LINT w|M. qt_expect Z_I Vz w * qt_expect I_XpZ Vp w) +
- (LINT w|M. qt_expect X_I Vx w * qt_expect I_XpZ Vp w) -
- (LINT w|M. qt_expect Z_I Vz w * qt_expect I_ZmX Vm w)\<bar>"
- using assms quantum_value[of M] by simp
- also have "... \<le> 2" using chsh_upper_bound[of M] assms \<open>prob_space M\<close> by simp
- finally have "2 * sqrt 2 \<le> 2" .
- thus False by simp
-qed
-
-
+(*
+Author:
+ Mnacho Echenim, Université Grenoble Alpes
+*)
+
+theory CHSH_Inequality imports
+ Projective_Measurements
+
+
+begin
+
+
+section \<open>The CHSH inequality\<close>
+
+text \<open>The local hidden variable assumption for quantum mechanics was developed to make the case
+that quantum theory was incomplete. In this part we formalize the CHSH inequality, which provides
+an upper-bound of a quantity involving expectations in a probability space, and exploit this
+inequality to show that the local hidden variable assumption does not hold.\<close>
+
+subsection \<open>Inequality statement\<close>
+
+lemma chsh_real:
+ fixes A0::real
+ assumes "\<bar>A0 * B1\<bar> \<le> 1"
+ and "\<bar>A0 * B0\<bar> \<le> 1"
+ and "\<bar>A1 * B0\<bar> \<le> 1"
+ and "\<bar>A1 * B1\<bar> \<le> 1"
+ shows "\<bar>A0 * B1 - A0 * B0 + A1 * B0 + A1*B1\<bar> \<le> 2"
+proof -
+ have "A0 * B1 - A0 * B0 = A0 * B1 - A0 * B0 + A0 * B1 * A1 * B0 - A0 * B1 * A1 * B0" by simp
+ also have "... = A0 * B1 * (1 + A1*B0) - A0 * B0 * (1 + A1* B1)"
+ by (metis (no_types, opaque_lifting) add_diff_cancel_right calculation diff_add_eq
+ group_cancel.sub1 mult.commute mult.right_neutral
+ vector_space_over_itself.scale_left_commute
+ vector_space_over_itself.scale_right_diff_distrib
+ vector_space_over_itself.scale_right_distrib
+ vector_space_over_itself.scale_scale)
+ finally have "A0 * B1 - A0 * B0 = A0 * B1 * (1 + A1*B0) - A0 * B0 * (1 + A1* B1)" .
+ hence "\<bar>A0 * B1 - A0 * B0\<bar> \<le> \<bar>A0 * B1 * (1 + A1*B0)\<bar> + \<bar>A0 * B0 * (1 + A1* B1)\<bar>" by simp
+ also have "... = \<bar>A0 * B1\<bar> * \<bar>(1 + A1*B0)\<bar> + \<bar>A0 * B0\<bar> * \<bar>(1 + A1* B1)\<bar>" by (simp add: abs_mult)
+ also have "... \<le> \<bar>(1 + A1*B0)\<bar> + \<bar>(1 + A1* B1)\<bar>"
+ proof-
+ have "\<bar>A0 * B1\<bar> * \<bar>(1 + A1*B0)\<bar> \<le> \<bar>(1 + A1*B0)\<bar>"
+ using assms(1) mult_left_le_one_le[of "\<bar>(1 + A1*B0)\<bar>"] by simp
+ moreover have "\<bar>A0 * B0\<bar> *\<bar>(1 + A1* B1)\<bar> \<le> \<bar>(1 + A1* B1)\<bar>"
+ using assms mult_left_le_one_le[of "\<bar>(1 + A1*B1)\<bar>"] by simp
+ ultimately show ?thesis by simp
+ qed
+ also have "... = 1 + A1*B0 + 1 + A1* B1" using assms by (simp add: abs_le_iff)
+ also have "... = 2 + A1 * B0 + A1 * B1" by simp
+ finally have pls: "\<bar>A0 * B1 - A0 * B0\<bar> \<le> 2 + A1 * B0 + A1 * B1" .
+ have "A0 * B1 - A0 * B0 = A0 * B1 - A0 * B0 + A0 * B1 * A1 * B0 - A0 * B1 * A1 * B0" by simp
+ also have "... = A0 * B1 * (1 - A1*B0) - A0 * B0 * (1 - A1* B1)"
+ proof -
+ have "A0 * (B1 - (B0 - A1 * (B1 * B0)) - A1 * (B1 * B0)) = A0 * (B1 - B0)"
+ by fastforce
+ then show ?thesis
+ by (metis (no_types) add.commute calculation diff_diff_add mult.right_neutral
+ vector_space_over_itself.scale_left_commute
+ vector_space_over_itself.scale_right_diff_distrib vector_space_over_itself.scale_scale)
+ qed
+ finally have "A0 * B1 - A0 * B0 = A0 * B1 * (1 - A1*B0) - A0 * B0 * (1 - A1* B1)" .
+ hence "\<bar>A0 * B1 - A0 * B0\<bar> \<le> \<bar>A0 * B1 * (1 - A1*B0)\<bar> + \<bar>A0 * B0 * (1 - A1* B1)\<bar>" by simp
+ also have "... = \<bar>A0 * B1\<bar> * \<bar>(1 - A1*B0)\<bar> + \<bar>A0 * B0\<bar> * \<bar>(1 - A1* B1)\<bar>" by (simp add: abs_mult)
+ also have "... \<le> \<bar>(1 - A1*B0)\<bar> + \<bar>(1 - A1* B1)\<bar>"
+ proof-
+ have "\<bar>A0 * B1\<bar> * \<bar>(1 - A1*B0)\<bar> \<le> \<bar>(1 - A1*B0)\<bar>"
+ using assms(1) mult_left_le_one_le[of "\<bar>(1 - A1*B0)\<bar>"] by simp
+ moreover have "\<bar>A0 * B0\<bar> *\<bar>(1 - A1* B1)\<bar> \<le> \<bar>(1 - A1* B1)\<bar>"
+ using assms mult_left_le_one_le[of "\<bar>(1 - A1*B1)\<bar>"] by simp
+ ultimately show ?thesis by simp
+ qed
+ also have "... = 1 - A1*B0 + 1 - A1* B1" using assms by (simp add: abs_le_iff)
+ also have "... = 2 - A1 * B0 - A1 * B1" by simp
+ finally have mns: "\<bar>A0 * B1 - A0 * B0\<bar> \<le> 2 - (A1 * B0 + A1 * B1)" by simp
+ have ls: "\<bar>A0 * B1 - A0 * B0\<bar> \<le> 2 - \<bar>A1 * B0 + A1 * B1\<bar>"
+ proof (cases "0 \<le> A1 * B0 + A1 * B1")
+ case True
+ then show ?thesis using mns by simp
+ next
+ case False
+ then show ?thesis using pls by simp
+ qed
+ have "\<bar>A0 * B1 - A0 * B0 + A1 * B0 + A1 * B1\<bar> \<le> \<bar>A0 * B1 - A0 * B0\<bar> + \<bar>A1 * B0 + A1 * B1\<bar>"
+ by simp
+ also have "... \<le> 2" using ls by simp
+ finally show ?thesis .
+qed
+
+lemma (in prob_space) chsh_expect:
+ fixes A0::"'a \<Rightarrow> real"
+ assumes "AE w in M. \<bar>A0 w\<bar> \<le> 1"
+ and "AE w in M. \<bar>A1 w\<bar> \<le> 1"
+ and "AE w in M. \<bar>B0 w\<bar> \<le> 1"
+ and "AE w in M. \<bar>B1 w\<bar> \<le> 1"
+and "integrable M (\<lambda>w. A0 w * B1 w)"
+and "integrable M (\<lambda>w. A1 w * B1 w)"
+and "integrable M (\<lambda>w. A1 w * B0 w)"
+and "integrable M (\<lambda>w. A0 w * B0 w)"
+shows "\<bar>expectation (\<lambda>w. A1 w * B0 w) + expectation (\<lambda>w. A0 w *B1 w) +
+ expectation (\<lambda>w. A1 w * B1 w) - expectation (\<lambda>w. A0 w * B0 w)\<bar> \<le> 2"
+proof -
+ have exeq: "expectation (\<lambda>w. A1 w * B0 w) + expectation (\<lambda>w. A0 w * B1 w) +
+ expectation (\<lambda>w. A1 w * B1 w) - expectation (\<lambda>w. A0 w * B0 w) =
+ expectation (\<lambda>w. A0 w * B1 w - A0 w * B0 w + A1 w * B0 w + A1 w * B1 w)"
+ using assms by auto
+ have "\<bar>expectation (\<lambda>w. A0 w * B1 w - A0 w * B0 w + A1 w * B0 w + A1 w * B1 w)\<bar> \<le>
+ expectation (\<lambda>w. \<bar>A0 w * B1 w - A0 w * B0 w + A1 w * B0 w + A1 w * B1 w\<bar>)"
+ using integral_abs_bound by blast
+ also have "... \<le> 2"
+ proof (rule integral_le_const)
+ show "AE w in M. \<bar>A0 w * B1 w - A0 w * B0 w + A1 w * B0 w + A1 w * B1 w\<bar> \<le> (2::real)"
+ proof (rule AE_mp)
+ show "AE w in M. \<bar>A0 w\<bar> \<le> 1 \<and> \<bar>A1 w\<bar> \<le> 1 \<and> \<bar>B0 w\<bar> \<le> 1 \<and> \<bar>B1 w\<bar> \<le> 1"
+ using assms by simp
+ show "AE w in M. \<bar>A0 w\<bar> \<le> 1 \<and> \<bar>A1 w\<bar> \<le> 1 \<and> \<bar>B0 w\<bar> \<le> 1 \<and> \<bar>B1 w\<bar> \<le> 1 \<longrightarrow>
+ \<bar>A0 w * B1 w - A0 w * B0 w + A1 w * B0 w + A1 w * B1 w\<bar> \<le> 2"
+ proof
+ fix w
+ assume "w\<in> space M"
+ show "\<bar>A0 w\<bar> \<le> 1 \<and> \<bar>A1 w\<bar> \<le> 1 \<and> \<bar>B0 w\<bar> \<le> 1 \<and> \<bar>B1 w\<bar> \<le> 1 \<longrightarrow>
+ \<bar>A0 w * B1 w - A0 w * B0 w + A1 w * B0 w + A1 w * B1 w\<bar> \<le> 2"
+ proof
+ assume ineq: "\<bar>A0 w\<bar> \<le> 1 \<and> \<bar>A1 w\<bar> \<le> 1 \<and> \<bar>B0 w\<bar> \<le> 1 \<and> \<bar>B1 w\<bar> \<le> 1"
+ show "\<bar>A0 w * B1 w - A0 w * B0 w + A1 w * B0 w + A1 w * B1 w\<bar> \<le> 2"
+ proof (rule chsh_real)
+ show "\<bar>A0 w * B1 w\<bar> \<le> 1" using ineq by (simp add: abs_mult mult_le_one)
+ show "\<bar>A0 w * B0 w\<bar> \<le> 1" using ineq by (simp add: abs_mult mult_le_one)
+ show "\<bar>A1 w * B1 w\<bar> \<le> 1" using ineq by (simp add: abs_mult mult_le_one)
+ show "\<bar>A1 w * B0 w\<bar> \<le> 1" using ineq by (simp add: abs_mult mult_le_one)
+ qed
+ qed
+ qed
+ qed
+ show "integrable M (\<lambda>x. \<bar>A0 x * B1 x - A0 x * B0 x + A1 x * B0 x + A1 x * B1 x\<bar>)"
+ proof (rule Bochner_Integration.integrable_abs)
+ show "integrable M (\<lambda>x. A0 x * B1 x - A0 x * B0 x + A1 x * B0 x + A1 x * B1 x)"
+ using assms by auto
+ qed
+ qed
+ finally show ?thesis using exeq by simp
+qed
+
+text \<open>The local hidden variable assumption states that separated quantum measurements are
+independent. It is standard for this assumption to be stated in a context where the hidden
+variable admits a density; it is stated here in a more general contest involving expectations,
+with no assumption on the existence of a density.\<close>
+
+definition pos_rv:: "'a measure \<Rightarrow> ('a \<Rightarrow> real) \<Rightarrow> bool" where
+"pos_rv M Xr \<equiv> Xr \<in> borel_measurable M \<and> (AE w in M. (0::real) \<le> Xr w)"
+
+definition prv_sum:: "'a measure \<Rightarrow> complex Matrix.mat \<Rightarrow> (complex \<Rightarrow> 'a \<Rightarrow> real) \<Rightarrow> bool" where
+"prv_sum M A Xr \<equiv> (AE w in M. (\<Sum> a\<in> spectrum A. Xr a w) = 1)"
+
+definition (in cpx_sq_mat) lhv where
+"lhv M A B R XA XB \<equiv>
+ prob_space M \<and>
+ (\<forall>a \<in>spectrum A. pos_rv M (XA a)) \<and>
+ (prv_sum M A XA) \<and>
+ (\<forall>b \<in>spectrum B. pos_rv M (XB b)) \<and>
+ (prv_sum M B XB) \<and>
+ (\<forall>a \<in>spectrum A . \<forall>b \<in> spectrum B.
+ (integrable M (\<lambda>w. XA a w * XB b w)) \<and>
+ integral\<^sup>L M (\<lambda>w. XA a w * XB b w) =
+ Re (Complex_Matrix.trace(eigen_projector A a * (eigen_projector B b) * R)))"
+
+(*definition (in cpx_sq_mat) lhv where
+"lhv M A B R XA XB \<equiv>
+ prob_space M \<and>
+ (\<forall>a \<in>spectrum A. (XA a \<in> borel_measurable M) \<and>
+ (AE w in M. ((0::real) \<le> XA a w))) \<and>
+ (AE w in M. (\<Sum> a\<in> spectrum A. XA a w) = 1) \<and>
+ (\<forall>b \<in>spectrum B. (XB b \<in> borel_measurable M) \<and>
+ (AE w in M. (0 \<le> XB b w))) \<and>
+ (AE w in M. (\<Sum> b\<in> spectrum B. XB b w) = 1) \<and>
+ (\<forall>a \<in>spectrum A . \<forall>b \<in> spectrum B.
+ (integrable M (\<lambda>w. XA a w * XB b w)) \<and>
+ integral\<^sup>L M (\<lambda>w. XA a w * XB b w) =
+ Re (Complex_Matrix.trace(eigen_projector A a * (eigen_projector B b) * R)))"*)
+
+lemma (in cpx_sq_mat) lhv_posl:
+ assumes "lhv M A B R XA XB"
+ shows "AE w in M. (\<forall> a \<in> spectrum A. 0 \<le> XA a w)"
+proof (rule AE_ball_countable[THEN iffD2])
+ show "countable (spectrum A)" using spectrum_finite[of A]
+ by (simp add: countable_finite)
+ show "\<forall>a\<in>spectrum A. AE w in M. 0 \<le> XA a w" using assms unfolding lhv_def pos_rv_def by simp
+qed
+
+lemma (in cpx_sq_mat) lhv_lt1_l:
+ assumes "lhv M A B R XA XB"
+ shows "AE w in M. (\<forall> a \<in> spectrum A. XA a w \<le> 1)"
+proof (rule AE_mp)
+ show "AE w in M. (\<forall> a \<in> spectrum A. 0 \<le> XA a w) \<and> (\<Sum> a\<in> spectrum A. XA a w) = 1"
+ using lhv_posl assms unfolding lhv_def pos_rv_def prv_sum_def by simp
+ show "AE w in M. (\<forall>a\<in>spectrum A. 0 \<le> XA a w) \<and> (\<Sum>a\<in>spectrum A. XA a w) = 1 \<longrightarrow>
+ (\<forall>a\<in>spectrum A. XA a w \<le> 1)"
+ proof
+ fix w
+ assume "w\<in> space M"
+ show "(\<forall>a\<in>spectrum A. 0 \<le> XA a w) \<and> (\<Sum>a\<in>spectrum A. XA a w) = 1 \<longrightarrow>
+ (\<forall>a\<in>spectrum A. XA a w \<le> 1)"
+ proof (rule impI)
+ assume pr: "(\<forall>a\<in>spectrum A. 0 \<le> XA a w) \<and> (\<Sum>a\<in>spectrum A. XA a w) = 1"
+ show "\<forall>a\<in>spectrum A. XA a w \<le> 1"
+ proof
+ fix a
+ assume "a\<in> spectrum A"
+ show "XA a w \<le> 1"
+ proof (rule pos_sum_1_le[of "spectrum A"])
+ show "finite (spectrum A)" using spectrum_finite[of A] by simp
+ show "a \<in> spectrum A" using \<open>a \<in> spectrum A\<close> .
+ show " \<forall>i\<in>spectrum A. 0 \<le> XA i w" using pr by simp
+ show "(\<Sum>i\<in>spectrum A. XA i w) = 1" using pr by simp
+ qed
+ qed
+ qed
+ qed
+qed
+
+lemma (in cpx_sq_mat) lhv_posr:
+ assumes "lhv M A B R XA XB"
+ shows "AE w in M. (\<forall> b \<in> spectrum B. 0 \<le> XB b w)"
+proof (rule AE_ball_countable[THEN iffD2])
+ show "countable (spectrum B)" using spectrum_finite[of B]
+ by (simp add: countable_finite)
+ show "\<forall>b\<in>spectrum B. AE w in M. 0 \<le> XB b w" using assms unfolding lhv_def pos_rv_def by simp
+qed
+
+lemma (in cpx_sq_mat) lhv_lt1_r:
+ assumes "lhv M A B R XA XB"
+ shows "AE w in M. (\<forall> a \<in> spectrum B. XB a w \<le> 1)"
+proof (rule AE_mp)
+ show "AE w in M. (\<forall> a \<in> spectrum B. 0 \<le> XB a w) \<and> (\<Sum> a\<in> spectrum B. XB a w) = 1"
+ using lhv_posr assms unfolding lhv_def prv_sum_def pos_rv_def by simp
+ show "AE w in M. (\<forall>a\<in>spectrum B. 0 \<le> XB a w) \<and> (\<Sum>a\<in>spectrum B. XB a w) = 1 \<longrightarrow>
+ (\<forall>a\<in>spectrum B. XB a w \<le> 1)"
+ proof
+ fix w
+ assume "w\<in> space M"
+ show "(\<forall>a\<in>spectrum B. 0 \<le> XB a w) \<and> (\<Sum>a\<in>spectrum B. XB a w) = 1 \<longrightarrow>
+ (\<forall>a\<in>spectrum B. XB a w \<le> 1)"
+ proof (rule impI)
+ assume pr: "(\<forall>a\<in>spectrum B. 0 \<le> XB a w) \<and> (\<Sum>a\<in>spectrum B. XB a w) = 1"
+ show "\<forall>a\<in>spectrum B. XB a w \<le> 1"
+ proof
+ fix a
+ assume "a\<in> spectrum B"
+ show "XB a w \<le> 1"
+ proof (rule pos_sum_1_le[of "spectrum B"])
+ show "finite (spectrum B)" using spectrum_finite[of B] by simp
+ show "a \<in> spectrum B" using \<open>a \<in> spectrum B\<close> .
+ show " \<forall>i\<in>spectrum B. 0 \<le> XB i w" using pr by simp
+ show "(\<Sum>i\<in>spectrum B. XB i w) = 1" using pr by simp
+ qed
+ qed
+ qed
+ qed
+qed
+
+lemma (in cpx_sq_mat) lhv_AE_propl:
+ assumes "lhv M A B R XA XB"
+ shows "AE w in M. (\<forall> a \<in> spectrum A. 0 \<le> XA a w \<and> XA a w \<le> 1) \<and> (\<Sum> a\<in> spectrum A. XA a w) = 1"
+proof (rule AE_conjI)
+ show "AE w in M. \<forall>a\<in>spectrum A. 0 \<le> XA a w \<and> XA a w \<le> 1"
+ proof (rule AE_mp)
+ show "AE w in M. (\<forall> a \<in> spectrum A. 0 \<le> XA a w) \<and> (\<forall> a \<in> spectrum A. XA a w \<le> 1)"
+ using assms lhv_posl[of M A] lhv_lt1_l[of M A] by simp
+ show "AE w in M. (\<forall>a\<in>spectrum A. 0 \<le> XA a w) \<and> (\<forall>a\<in>spectrum A. XA a w \<le> 1) \<longrightarrow>
+ (\<forall>a\<in>spectrum A. 0 \<le> XA a w \<and> XA a w \<le> 1)" by auto
+ qed
+ show "AE w in M. (\<Sum>a\<in>spectrum A. XA a w) = 1" using assms unfolding lhv_def prv_sum_def
+ by simp
+qed
+
+lemma (in cpx_sq_mat) lhv_AE_propr:
+ assumes "lhv M A B R XA XB"
+ shows "AE w in M. (\<forall> a \<in> spectrum B. 0 \<le> XB a w \<and> XB a w \<le> 1) \<and> (\<Sum> a\<in> spectrum B. XB a w) = 1"
+proof (rule AE_conjI)
+ show "AE w in M. \<forall>a\<in>spectrum B. 0 \<le> XB a w \<and> XB a w \<le> 1"
+ proof (rule AE_mp)
+ show "AE w in M. (\<forall> a \<in> spectrum B. 0 \<le> XB a w) \<and> (\<forall> a \<in> spectrum B. XB a w \<le> 1)"
+ using assms lhv_posr[of M _ B] lhv_lt1_r[of M _ B] by simp
+ show "AE w in M. (\<forall>a\<in>spectrum B. 0 \<le> XB a w) \<and> (\<forall>a\<in>spectrum B. XB a w \<le> 1) \<longrightarrow>
+ (\<forall>a\<in>spectrum B. 0 \<le> XB a w \<and> XB a w \<le> 1)" by auto
+ qed
+ show "AE w in M. (\<Sum>a\<in>spectrum B. XB a w) = 1" using assms unfolding lhv_def prv_sum_def
+ by simp
+qed
+
+lemma (in cpx_sq_mat) lhv_integral_eq:
+ fixes c::real
+ assumes "lhv M A B R XA XB"
+ and "a\<in> spectrum A"
+and "b\<in> spectrum B"
+shows "integral\<^sup>L M (\<lambda>w. XA a w * XB b w) =
+ Re (Complex_Matrix.trace(eigen_projector A a * (eigen_projector B b) * R))"
+ using assms unfolding lhv_def by simp
+
+lemma (in cpx_sq_mat) lhv_integrable:
+ fixes c::real
+ assumes "lhv M A B R XA XB"
+ and "a\<in> spectrum A"
+and "b\<in> spectrum B"
+shows "integrable M (\<lambda>w. XA a w * XB b w)" using assms unfolding lhv_def by simp
+
+lemma (in cpx_sq_mat) lhv_scal_integrable:
+ fixes c::real
+ assumes "lhv M A B R XA XB"
+ and "a\<in> spectrum A"
+and "b\<in> spectrum B"
+shows "integrable M (\<lambda>w. c *XA a w * d * XB b w)"
+proof -
+ {
+ fix x
+ assume "x\<in> space M"
+ have "c * d * (XA a x * XB b x) = c * XA a x * d * XB b x" by simp
+ } note eq = this
+ have "integrable M (\<lambda>w. XA a w * XB b w)" using assms unfolding lhv_def by simp
+ hence g:"integrable M (\<lambda>w. c * d * (XA a w * XB b w))" using integrable_real_mult_right by simp
+ show ?thesis
+ proof (rule Bochner_Integration.integrable_cong[THEN iffD2], simp)
+ show "integrable M (\<lambda>w. c * d * (XA a w * XB b w))" using g .
+ show "\<And>x. x \<in> space M \<Longrightarrow> c * XA a x * d * XB b x = c * d * (XA a x * XB b x)" using eq by simp
+ qed
+qed
+
+lemma (in cpx_sq_mat) lhv_lsum_scal_integrable:
+ assumes "lhv M A B R XA XB"
+ and "a\<in> spectrum A"
+shows "integrable M (\<lambda>x. \<Sum>b\<in>spectrum B. c * XA a x * (f b) * XB b x)"
+proof (rule Bochner_Integration.integrable_sum)
+ fix b
+ assume "b\<in> spectrum B"
+ thus "integrable M (\<lambda>x. c * XA a x * f b *XB b x)" using \<open>a\<in> spectrum A\<close> assms
+ lhv_scal_integrable[of M] by simp
+qed
+
+lemma (in cpx_sq_mat) lhv_sum_integrable:
+ assumes "lhv M A B R XA XB"
+shows "integrable M (\<lambda>w. (\<Sum> a \<in> spectrum A. (\<Sum> b \<in> spectrum B. f a * XA a w * g b * XB b w)))"
+proof (rule Bochner_Integration.integrable_sum)
+ fix a
+ assume "a\<in> spectrum A"
+ thus "integrable M (\<lambda>x. \<Sum>b\<in>spectrum B. f a * XA a x * g b * XB b x)"
+ using assms lhv_lsum_scal_integrable[of M]
+ by simp
+qed
+
+lemma (in cpx_sq_mat) spectrum_abs_1_weighted_suml:
+ assumes "lhv M A B R Va Vb"
+and "{Re x |x. x \<in> spectrum A} \<noteq> {}"
+ and "{Re x |x. x \<in> spectrum A} \<subseteq> {- 1, 1}"
+and "hermitian A"
+ and "A\<in> fc_mats"
+shows "AE w in M. \<bar>(\<Sum>a\<in>spectrum A. Re a * Va a w)\<bar> \<le> 1"
+proof (rule AE_mp)
+ show "AE w in M. (\<forall>a\<in>spectrum A. 0 \<le> Va a w \<and> Va a w \<le> 1) \<and> (\<Sum>a\<in>spectrum A. Va a w) = 1"
+ using assms lhv_AE_propl[of M A B _ Va] by simp
+ show "AE w in M. (\<forall>a\<in>spectrum A. 0 \<le> Va a w \<and> Va a w \<le> 1) \<and> (\<Sum>a\<in>spectrum A. Va a w) = 1 \<longrightarrow>
+ \<bar>\<Sum>a\<in>spectrum A. Re a * Va a w\<bar> \<le> 1"
+ proof
+ fix w
+ assume "w\<in> space M"
+ show "(\<forall>a\<in>spectrum A. 0 \<le> Va a w \<and> Va a w \<le> 1) \<and> (\<Sum>a\<in>spectrum A. Va a w) = 1 \<longrightarrow>
+ \<bar>\<Sum>a\<in>spectrum A. Re a * Va a w\<bar> \<le> 1"
+ proof (intro impI)
+ assume pr: "(\<forall>a\<in>spectrum A. 0 \<le> Va a w \<and> Va a w \<le> 1) \<and> (\<Sum>a\<in>spectrum A. Va a w) = 1"
+ show "\<bar>(\<Sum>a\<in>spectrum A. Re a * Va a w)\<bar> \<le> 1"
+ proof (cases "{Re x |x. x \<in> spectrum A} = {- 1, 1}")
+ case True
+ hence sp: "spectrum A = {-1, 1}" using hermitian_Re_spectrum[of A] assms by simp
+ hence scsum: "(\<Sum>a\<in>spectrum A. Re a * Va a w) = Va 1 w - Va (-1) w"
+ using sum_2_elems by simp
+ have sum: "(\<Sum>a\<in>spectrum A. Va a w) = Va (-1) w + Va 1 w"
+ using sp sum_2_elems by simp
+ have "\<bar>Va 1 w - Va (-1) w\<bar> \<le> 1"
+ proof (rule fct_bound')
+ have "1 \<in> spectrum A" using sp by simp
+ thus "0 \<le> Va 1 w" using pr by simp
+ have "-1 \<in> spectrum A" using sp by simp
+ thus "0 \<le> Va (- 1) w" using pr by simp
+ show "Va (- 1) w + Va 1 w = 1" using pr sum by simp
+ qed
+ thus ?thesis using scsum by simp
+ next
+ case False
+ then show ?thesis
+ proof (cases "{Re x |x. x \<in> spectrum A} = {- 1}")
+ case True
+ hence "spectrum A = {-1}" using hermitian_Re_spectrum[of A] assms by simp
+ hence "(\<Sum>a\<in>spectrum A. Re a * Va a w) = - Va (-1) w" by simp
+ moreover have "-1 \<in> spectrum A" using \<open>spectrum A = {-1}\<close> by simp
+ ultimately show ?thesis using pr by simp
+ next
+ case False
+ hence "{Re x |x. x \<in> spectrum A} = {1}" using assms \<open>{Re x |x. x \<in> spectrum A} \<noteq> {-1, 1}\<close>
+ last_subset[of "{Re x |x. x \<in> spectrum A}"] by simp
+ hence "spectrum A = {1}" using hermitian_Re_spectrum[of A] assms by simp
+ hence "(\<Sum>a\<in>spectrum A. Re a * Va a w) = Va 1 w" by simp
+ moreover have "1 \<in> spectrum A" using \<open>spectrum A = {1}\<close> by simp
+ ultimately show ?thesis using pr by simp
+ qed
+ qed
+ qed
+ qed
+qed
+
+lemma (in cpx_sq_mat) spectrum_abs_1_weighted_sumr:
+ assumes "lhv M B A R Vb Va"
+and "{Re x |x. x \<in> spectrum A} \<noteq> {}"
+ and "{Re x |x. x \<in> spectrum A} \<subseteq> {- 1, 1}"
+and "hermitian A"
+ and "A\<in> fc_mats"
+shows "AE w in M. \<bar>(\<Sum>a\<in>spectrum A. Re a * Va a w)\<bar> \<le> 1"
+proof (rule AE_mp)
+ show "AE w in M. (\<forall>a\<in>spectrum A. 0 \<le> Va a w \<and> Va a w \<le> 1) \<and> (\<Sum>a\<in>spectrum A. Va a w) = 1"
+ using assms lhv_AE_propr[of M B A _ Vb] by simp
+ show "AE w in M. (\<forall>a\<in>spectrum A. 0 \<le> Va a w \<and> Va a w \<le> 1) \<and> (\<Sum>a\<in>spectrum A. Va a w) = 1 \<longrightarrow>
+ \<bar>\<Sum>a\<in>spectrum A. Re a * Va a w\<bar> \<le> 1"
+ proof
+ fix w
+ assume "w\<in> space M"
+ show "(\<forall>a\<in>spectrum A. 0 \<le> Va a w \<and> Va a w \<le> 1) \<and> (\<Sum>a\<in>spectrum A. Va a w) = 1 \<longrightarrow>
+ \<bar>\<Sum>a\<in>spectrum A. Re a * Va a w\<bar> \<le> 1"
+ proof (intro impI)
+ assume pr: "(\<forall>a\<in>spectrum A. 0 \<le> Va a w \<and> Va a w \<le> 1) \<and> (\<Sum>a\<in>spectrum A. Va a w) = 1"
+ show "\<bar>(\<Sum>a\<in>spectrum A. Re a * Va a w)\<bar> \<le> 1"
+ proof (cases "{Re x |x. x \<in> spectrum A} = {- 1, 1}")
+ case True
+ hence sp: "spectrum A = {-1, 1}" using hermitian_Re_spectrum[of A] assms by simp
+ hence scsum: "(\<Sum>a\<in>spectrum A. Re a * Va a w) = Va 1 w - Va (-1) w"
+ using sum_2_elems by simp
+ have sum: "(\<Sum>a\<in>spectrum A. Va a w) = Va (-1) w + Va 1 w"
+ using sp sum_2_elems by simp
+ have "\<bar>Va 1 w - Va (-1) w\<bar> \<le> 1"
+ proof (rule fct_bound')
+ have "1 \<in> spectrum A" using sp by simp
+ thus "0 \<le> Va 1 w" using pr by simp
+ have "-1 \<in> spectrum A" using sp by simp
+ thus "0 \<le> Va (- 1) w" using pr by simp
+ show "Va (- 1) w + Va 1 w = 1" using pr sum by simp
+ qed
+ thus ?thesis using scsum by simp
+ next
+ case False
+ then show ?thesis
+ proof (cases "{Re x |x. x \<in> spectrum A} = {- 1}")
+ case True
+ hence "spectrum A = {-1}" using hermitian_Re_spectrum[of A] assms by simp
+ hence "(\<Sum>a\<in>spectrum A. Re a * Va a w) = - Va (-1) w" by simp
+ moreover have "-1 \<in> spectrum A" using \<open>spectrum A = {-1}\<close> by simp
+ ultimately show ?thesis using pr by simp
+ next
+ case False
+ hence "{Re x |x. x \<in> spectrum A} = {1}" using assms \<open>{Re x |x. x \<in> spectrum A} \<noteq> {-1, 1}\<close>
+ last_subset[of "{Re x |x. x \<in> spectrum A}"] by simp
+ hence "spectrum A = {1}" using hermitian_Re_spectrum[of A] assms by simp
+ hence "(\<Sum>a\<in>spectrum A. Re a * Va a w) = Va 1 w" by simp
+ moreover have "1 \<in> spectrum A" using \<open>spectrum A = {1}\<close> by simp
+ ultimately show ?thesis using pr by simp
+ qed
+ qed
+ qed
+ qed
+qed
+
+definition qt_expect where
+"qt_expect A Va = (\<lambda>w. (\<Sum>a\<in>spectrum A. Re a * Va a w))"
+
+lemma (in cpx_sq_mat) spectr_sum_integrable:
+assumes "lhv M A B R Va Vb"
+shows "integrable M (\<lambda>w. qt_expect A Va w * qt_expect B Vb w)"
+proof (rule Bochner_Integration.integrable_cong[THEN iffD2])
+ show "M = M" by simp
+ show "\<And>w. w \<in> space M \<Longrightarrow> qt_expect A Va w * qt_expect B Vb w =
+ (\<Sum>a\<in>spectrum A. (\<Sum>b\<in>spectrum B. Re a * Va a w * Re b * Vb b w))"
+ proof -
+ fix w
+ assume "w\<in> space M"
+ show "qt_expect A Va w * qt_expect B Vb w =
+ (\<Sum>a\<in>spectrum A. (\<Sum>b\<in>spectrum B. Re a * Va a w * Re b * Vb b w))" unfolding qt_expect_def
+ by (metis (mono_tags, lifting) Finite_Cartesian_Product.sum_cong_aux sum_product
+ vector_space_over_itself.scale_scale)
+ qed
+ show "integrable M (\<lambda>w. \<Sum>a\<in>spectrum A. (\<Sum>b\<in>spectrum B. Re a * Va a w * Re b * Vb b w))"
+ using lhv_sum_integrable[of M] assms by simp
+qed
+
+lemma (in cpx_sq_mat) lhv_sum_integral':
+ assumes "lhv M A B R XA XB"
+shows "integral\<^sup>L M (\<lambda>w. (\<Sum> a \<in> spectrum A. f a * XA a w) * (\<Sum> b \<in> spectrum B. g b * XB b w)) =
+ (\<Sum> a \<in> spectrum A. f a * (\<Sum> b \<in> spectrum B. g b * integral\<^sup>L M (\<lambda>w. XA a w * XB b w)))"
+proof -
+ have "integral\<^sup>L M (\<lambda>w. (\<Sum> a \<in> spectrum A. f a * XA a w) * (\<Sum> b \<in> spectrum B. g b * XB b w)) =
+ integral\<^sup>L M (\<lambda>w. (\<Sum> a \<in> spectrum A. (\<Sum> b \<in> spectrum B. f a * XA a w * g b * XB b w)))"
+ proof (rule Bochner_Integration.integral_cong, simp)
+ fix w
+ assume "w\<in> space M"
+ show "(\<Sum>a\<in>spectrum A. f a * XA a w) * (\<Sum>b\<in>spectrum B. g b * XB b w) =
+ (\<Sum>a\<in>spectrum A. (\<Sum>b\<in>spectrum B. f a * XA a w * (g b) * XB b w))"
+ by (simp add: sum_product vector_space_over_itself.scale_scale)
+ qed
+ also have "... = (\<Sum> a \<in> spectrum A.
+ integral\<^sup>L M (\<lambda>w. (\<Sum> b \<in> spectrum B. f a * XA a w * g b * XB b w)))"
+ proof (rule Bochner_Integration.integral_sum)
+ fix a
+ assume "a\<in> spectrum A"
+ thus "integrable M (\<lambda>x. \<Sum>b\<in>spectrum B. f a * XA a x * g b * XB b x)"
+ using lhv_lsum_scal_integrable[of M] assms by simp
+ qed
+ also have "... = (\<Sum> a \<in> spectrum A. f a *
+ integral\<^sup>L M (\<lambda>w. (\<Sum> b \<in> spectrum B. XA a w * g b * XB b w)))"
+ proof -
+ have "\<forall> a\<in> spectrum A. integral\<^sup>L M (\<lambda>w. (\<Sum> b \<in> spectrum B. f a * XA a w * g b * XB b w)) =
+ f a * integral\<^sup>L M (\<lambda>w. (\<Sum> b \<in> spectrum B. XA a w * g b * XB b w))"
+ proof
+ fix a
+ assume "a\<in> spectrum A"
+ have "(LINT w|M. (\<Sum>b\<in>spectrum B. f a * XA a w * g b * XB b w)) =
+ (LINT w|M. f a* (\<Sum>b\<in>spectrum B. XA a w * g b * XB b w))"
+ proof (rule Bochner_Integration.integral_cong, simp)
+ fix x
+ assume "x \<in> space M"
+ show "(\<Sum>b\<in>spectrum B. f a * XA a x * g b * XB b x) =
+ f a * (\<Sum>b\<in>spectrum B. XA a x * g b * XB b x)"
+ by (metis (no_types, lifting) Finite_Cartesian_Product.sum_cong_aux
+ vector_space_over_itself.scale_scale vector_space_over_itself.scale_sum_right)
+ qed
+ also have "... = f a * (LINT w|M. (\<Sum>b\<in>spectrum B. XA a w * g b * XB b w))" by simp
+ finally show "(LINT w|M. (\<Sum>b\<in>spectrum B. f a * XA a w * g b * XB b w)) =
+ f a * (LINT w|M. (\<Sum>b\<in>spectrum B. XA a w * g b * XB b w))" .
+ qed
+ thus ?thesis by simp
+ qed
+ also have "... = (\<Sum> a \<in> spectrum A. f a * (\<Sum> b \<in> spectrum B. g b *
+ integral\<^sup>L M (\<lambda>w. XA a w * XB b w)))"
+ proof (rule sum.cong, simp)
+ fix a
+ assume "a\<in> spectrum A"
+ have "integral\<^sup>L M (\<lambda>w. (\<Sum> b \<in> spectrum B. XA a w * g b * XB b w)) = (\<Sum> b \<in> spectrum B.
+ integral\<^sup>L M (\<lambda>w. XA a w * g b * XB b w))"
+ proof (rule Bochner_Integration.integral_sum)
+ show "\<And>b. b \<in> spectrum B \<Longrightarrow> integrable M (\<lambda>x. XA a x * g b * XB b x)"
+ proof -
+ fix b
+ assume "b\<in> spectrum B"
+ thus "integrable M (\<lambda>x. XA a x * g b * XB b x)"
+ using assms lhv_scal_integrable[of M _ _ _ _ _ a b 1] \<open>a\<in> spectrum A\<close> by simp
+ qed
+ qed
+ also have "... = (\<Sum> b \<in> spectrum B. g b * integral\<^sup>L M (\<lambda>w. XA a w * XB b w))"
+ proof (rule sum.cong, simp)
+ fix x
+ assume "x\<in> spectrum B"
+ have "LINT w|M. XA a w * g x * XB x w = LINT w|M. g x * (XA a w * XB x w)"
+ by (rule Bochner_Integration.integral_cong, auto)
+ also have "... = g x * (LINT w|M. XA a w * XB x w)"
+ using Bochner_Integration.integral_mult_right_zero[of M "g x" "\<lambda>w. XA a w * XB x w"]
+ by simp
+ finally show "LINT w|M. XA a w * g x * XB x w = g x * (LINT w|M. XA a w * XB x w)" .
+ qed
+ finally have "integral\<^sup>L M (\<lambda>w. (\<Sum> b \<in> spectrum B. XA a w * g b * XB b w)) =
+ (\<Sum> b \<in> spectrum B. g b * integral\<^sup>L M (\<lambda>w. XA a w * XB b w))" .
+ thus "f a * (LINT w|M. (\<Sum>b\<in>spectrum B. XA a w * g b * XB b w)) =
+ f a * (\<Sum> b \<in> spectrum B. g b * integral\<^sup>L M (\<lambda>w. XA a w * XB b w))" by simp
+ qed
+ finally show ?thesis .
+qed
+
+lemma (in cpx_sq_mat) sum_qt_expect_trace:
+ assumes "lhv M A B R XA XB"
+ shows "(\<Sum> a \<in> spectrum A. f a * (\<Sum> b \<in> spectrum B. g b * integral\<^sup>L M (\<lambda>w. XA a w * XB b w))) =
+ (\<Sum> a \<in> spectrum A. f a * (\<Sum> b \<in> spectrum B. g b *
+ Re (Complex_Matrix.trace(eigen_projector A a * (eigen_projector B b) * R))))"
+proof (rule sum.cong, simp)
+ fix a
+ assume "a\<in> spectrum A"
+ have "(\<Sum>b\<in>spectrum B. g b * (LINT w|M. XA a w * XB b w)) =
+ (\<Sum>b\<in>spectrum B. g b *
+ Re (Complex_Matrix.trace (eigen_projector A a * eigen_projector B b * R)))"
+ proof (rule sum.cong, simp)
+ fix b
+ assume "b\<in> spectrum B"
+ show "g b * (LINT w|M. XA a w * XB b w) =
+ g b * Re (Complex_Matrix.trace (eigen_projector A a * eigen_projector B b * R))"
+ using lhv_integral_eq[of M] assms \<open>a \<in> spectrum A\<close> \<open>b \<in> spectrum B\<close> by simp
+ qed
+ thus "f a * (\<Sum>b\<in>spectrum B. g b * (LINT w|M. XA a w * XB b w)) =
+ f a * (\<Sum>b\<in>spectrum B. g b *
+ Re (Complex_Matrix.trace (eigen_projector A a * eigen_projector B b * R)))" by simp
+qed
+
+lemma (in cpx_sq_mat) sum_eigen_projector_trace_dist:
+ assumes "hermitian B"
+and "A\<in> fc_mats"
+and "B\<in> fc_mats"
+and "R\<in> fc_mats"
+ shows "(\<Sum> b \<in> spectrum B. (b *
+ Complex_Matrix.trace(A * (eigen_projector B b) * R))) = Complex_Matrix.trace(A * B * R)"
+proof -
+ have "(\<Sum>b\<in>spectrum B. b * Complex_Matrix.trace (A * eigen_projector B b * R)) =
+ (\<Sum>b\<in>spectrum B. Complex_Matrix.trace (A * (b \<cdot>\<^sub>m (eigen_projector B b)) * R))"
+ proof (rule sum.cong, simp)
+ fix b
+ assume "b\<in> spectrum B"
+ have "b * Complex_Matrix.trace (A * eigen_projector B b * R) =
+ Complex_Matrix.trace (b \<cdot>\<^sub>m (A * eigen_projector B b * R))"
+ proof (rule trace_smult[symmetric])
+ show "A * eigen_projector B b * R \<in> carrier_mat dimR dimR" using eigen_projector_carrier
+ assms fc_mats_carrier dim_eq \<open>b \<in> spectrum B\<close> cpx_sq_mat_mult by meson
+ qed
+ also have "... = Complex_Matrix.trace (A * (b \<cdot>\<^sub>m eigen_projector B b) * R)"
+ proof -
+ have "b \<cdot>\<^sub>m (A * eigen_projector B b * R) = b \<cdot>\<^sub>m (A * (eigen_projector B b * R))"
+ proof -
+ have "A * eigen_projector B b * R = A * (eigen_projector B b * R)"
+ by (metis \<open>b \<in> spectrum B\<close> assms(1) assms(2) assms(3) assms(4) assoc_mult_mat dim_eq
+ fc_mats_carrier eigen_projector_carrier)
+ thus ?thesis by simp
+ qed
+ also have "... = A * (b \<cdot>\<^sub>m (eigen_projector B b * R))"
+ proof (rule mult_smult_distrib[symmetric])
+ show "A \<in> carrier_mat dimR dimR" using eigen_projector_carrier assms
+ fc_mats_carrier dim_eq by simp
+ show "eigen_projector B b * R \<in> carrier_mat dimR dimR" using eigen_projector_carrier
+ \<open>b \<in> spectrum B\<close> assms fc_mats_carrier dim_eq cpx_sq_mat_mult by blast
+ qed
+ also have "... = A * ((b \<cdot>\<^sub>m eigen_projector B b) * R)"
+ proof -
+ have "b \<cdot>\<^sub>m (eigen_projector B b * R) = (b \<cdot>\<^sub>m eigen_projector B b) * R"
+ proof (rule mult_smult_assoc_mat[symmetric])
+ show "eigen_projector B b \<in> carrier_mat dimR dimR" using eigen_projector_carrier
+ \<open>b \<in> spectrum B\<close> assms fc_mats_carrier dim_eq by simp
+ show "R \<in> carrier_mat dimR dimR" using assms fc_mats_carrier dim_eq by simp
+ qed
+ thus ?thesis by simp
+ qed
+ also have "... = A * (b \<cdot>\<^sub>m eigen_projector B b) * R"
+ by (metis \<open>b \<in> spectrum B\<close> assms(1) assms(2) assms(3) assms(4) assoc_mult_mat
+ cpx_sq_mat_smult dim_eq fc_mats_carrier eigen_projector_carrier)
+ finally have "b \<cdot>\<^sub>m (A * eigen_projector B b * R) = A * (b \<cdot>\<^sub>m eigen_projector B b) * R" .
+ then show ?thesis by simp
+ qed
+ finally show "b * Complex_Matrix.trace (A * eigen_projector B b * R) =
+ Complex_Matrix.trace (A * (b \<cdot>\<^sub>m eigen_projector B b) * R)" .
+ qed
+ also have "... = Complex_Matrix.trace (A *
+ (sum_mat (\<lambda>b. b \<cdot>\<^sub>m eigen_projector B b) (spectrum B)) * R)"
+ proof (rule trace_sum_mat_mat_distrib, (auto simp add: assms))
+ show "finite (spectrum B)" using spectrum_finite[of B] by simp
+ fix b
+ assume "b\<in> spectrum B"
+ show "b \<cdot>\<^sub>m eigen_projector B b \<in> fc_mats"
+ by (simp add: \<open>b \<in> spectrum B\<close> assms(1) assms(3) cpx_sq_mat_smult eigen_projector_carrier)
+ qed
+ also have "... = Complex_Matrix.trace (A * B * R)"
+ proof -
+ have "sum_mat (\<lambda>b. b \<cdot>\<^sub>m eigen_projector B b) (spectrum B) = B" using make_pm_sum' assms by simp
+ thus ?thesis by simp
+ qed
+ finally show ?thesis .
+qed
+
+lemma (in cpx_sq_mat) sum_eigen_projector_trace_right:
+ assumes "hermitian A"
+and "A\<in> fc_mats"
+and "B\<in> fc_mats"
+shows "(\<Sum> a \<in> spectrum A. Complex_Matrix.trace (a \<cdot>\<^sub>m eigen_projector A a * B)) =
+ Complex_Matrix.trace (A * B)"
+proof -
+ have "sum_mat (\<lambda>a. a \<cdot>\<^sub>m eigen_projector A a * B) (spectrum A) =
+ sum_mat (\<lambda>a. a \<cdot>\<^sub>m eigen_projector A a) (spectrum A) * B"
+ proof (rule mult_sum_mat_distrib_right)
+ show "finite (spectrum A)" using spectrum_finite[of A] by simp
+ show "\<And>a. a \<in> spectrum A \<Longrightarrow> a \<cdot>\<^sub>m eigen_projector A a \<in> fc_mats"
+ using assms(1) assms(2) cpx_sq_mat_smult eigen_projector_carrier by blast
+ show "B\<in> fc_mats" using assms by simp
+ qed
+ also have "... = A * B" using make_pm_sum' assms by simp
+ finally have seq: "sum_mat (\<lambda>a. a \<cdot>\<^sub>m eigen_projector A a * B) (spectrum A) = A * B" .
+ have "(\<Sum> a \<in> spectrum A. Complex_Matrix.trace (a \<cdot>\<^sub>m eigen_projector A a * B)) =
+ Complex_Matrix.trace (sum_mat (\<lambda>a. a \<cdot>\<^sub>m eigen_projector A a * B) (spectrum A))"
+ proof (rule trace_sum_mat[symmetric])
+ show "finite (spectrum A)" using spectrum_finite[of A] by simp
+ show "\<And>a. a \<in> spectrum A \<Longrightarrow> a \<cdot>\<^sub>m eigen_projector A a * B \<in> fc_mats"
+ by (simp add: assms(1) assms(2) assms(3) cpx_sq_mat_mult cpx_sq_mat_smult
+ eigen_projector_carrier)
+ qed
+ also have "... = Complex_Matrix.trace (A * B)" using seq by simp
+ finally show ?thesis .
+qed
+
+lemma (in cpx_sq_mat) sum_eigen_projector_trace:
+ assumes "hermitian A"
+ and "hermitian B"
+ and "A\<in> fc_mats"
+ and "B\<in> fc_mats"
+and "R\<in> fc_mats"
+ shows "(\<Sum> a \<in> spectrum A. a * (\<Sum> b \<in> spectrum B. (b *
+ Complex_Matrix.trace(eigen_projector A a * (eigen_projector B b) * R)))) =
+ Complex_Matrix.trace(A * B * R)"
+proof -
+ have "(\<Sum> a \<in> spectrum A. a * (\<Sum> b \<in> spectrum B. (b *
+ Complex_Matrix.trace(eigen_projector A a * (eigen_projector B b) * R)))) = (\<Sum> a \<in> spectrum A.
+ Complex_Matrix.trace (a \<cdot>\<^sub>m eigen_projector A a * (B * R)))"
+ proof (rule sum.cong, simp)
+ fix a
+ assume "a\<in> spectrum A"
+ hence "(\<Sum>b\<in>spectrum B. b *
+ Complex_Matrix.trace (eigen_projector A a * eigen_projector B b * R)) =
+ Complex_Matrix.trace (eigen_projector A a * B * R)" using
+ sum_eigen_projector_trace_dist[of B "eigen_projector A a" R] assms eigen_projector_carrier
+ by blast
+ hence "a * (\<Sum> b \<in> spectrum B. (b *
+ Complex_Matrix.trace(eigen_projector A a * (eigen_projector B b) * R))) =
+ a * Complex_Matrix.trace (eigen_projector A a * B * R)" by simp
+ also have "... = Complex_Matrix.trace (a \<cdot>\<^sub>m (eigen_projector A a * B * R))"
+ using trace_smult[symmetric, of "eigen_projector A a * B * R" dimR a] assms
+ \<open>a \<in> spectrum A\<close> cpx_sq_mat_mult dim_eq fc_mats_carrier eigen_projector_carrier by force
+ also have "... = Complex_Matrix.trace (a \<cdot>\<^sub>m eigen_projector A a * (B * R))"
+ proof -
+ have "a \<cdot>\<^sub>m (eigen_projector A a * B * R) = a \<cdot>\<^sub>m (eigen_projector A a * B) * R"
+ proof (rule mult_smult_assoc_mat[symmetric])
+ show "R\<in> carrier_mat dimR dimR" using assms fc_mats_carrier dim_eq by simp
+ show "eigen_projector A a * B \<in> carrier_mat dimR dimR" using assms eigen_projector_carrier
+ cpx_sq_mat_mult fc_mats_carrier dim_eq \<open>a \<in> spectrum A\<close> by blast
+ qed
+ also have "... = a \<cdot>\<^sub>m eigen_projector A a * B * R"
+ proof -
+ have "a \<cdot>\<^sub>m (eigen_projector A a * B) = a \<cdot>\<^sub>m eigen_projector A a * B"
+ using mult_smult_assoc_mat[symmetric]
+ proof -
+ show ?thesis
+ by (metis \<open>\<And>nr nc n k B A. \<lbrakk>A \<in> carrier_mat nr n; B \<in> carrier_mat n nc\<rbrakk> \<Longrightarrow>
+ k \<cdot>\<^sub>m (A * B) = k \<cdot>\<^sub>m A * B\<close> \<open>a \<in> spectrum A\<close> assms(1) assms(3) assms(4) dim_eq
+ fc_mats_carrier eigen_projector_carrier)
+ qed
+ thus ?thesis by simp
+ qed
+ also have "... = a \<cdot>\<^sub>m eigen_projector A a * (B * R)"
+ by (metis \<open>a \<in> spectrum A\<close> assms(1) assms(3) assms(4) assms(5) assoc_mult_mat
+ cpx_sq_mat_smult dim_eq fc_mats_carrier eigen_projector_carrier)
+ finally show ?thesis by simp
+ qed
+ finally show "a * (\<Sum> b \<in> spectrum B. (b *
+ Complex_Matrix.trace(eigen_projector A a * (eigen_projector B b) * R))) =
+ Complex_Matrix.trace (a \<cdot>\<^sub>m eigen_projector A a * (B * R))" .
+ qed
+ also have "... = Complex_Matrix.trace (A * (B * R))"
+ using sum_eigen_projector_trace_right[of A "B * R"] assms by (simp add: cpx_sq_mat_mult)
+ also have "... = Complex_Matrix.trace (A * B * R)"
+ proof -
+ have "A * (B * R) = A * B * R"
+ proof (rule assoc_mult_mat[symmetric])
+ show "A\<in> carrier_mat dimR dimR" using assms fc_mats_carrier dim_eq by simp
+ show "B\<in> carrier_mat dimR dimR" using assms fc_mats_carrier dim_eq by simp
+ show "R\<in> carrier_mat dimR dimR" using assms fc_mats_carrier dim_eq by simp
+ qed
+ thus ?thesis by simp
+ qed
+ finally show ?thesis .
+qed
+
+text \<open>We obtain the main result of this part, which relates the quantum expectation value of a
+joint measurement with an expectation.\<close>
+
+lemma (in cpx_sq_mat) sum_qt_expect:
+ assumes "lhv M A B R XA XB"
+ and "A\<in> fc_mats"
+ and "B\<in> fc_mats"
+ and "R\<in> fc_mats"
+ and "hermitian A"
+ and "hermitian B"
+ shows "integral\<^sup>L M (\<lambda>w. (qt_expect A XA w) * (qt_expect B XB w)) =
+ Re (Complex_Matrix.trace(A * B * R))"
+proof -
+ have br: "\<forall> b \<in> spectrum B. b\<in> Reals" using assms hermitian_spectrum_real[of B] by auto
+ have ar: "\<forall>a \<in> spectrum A. a\<in> Reals" using hermitian_spectrum_real[of A] assms by auto
+ have "integral\<^sup>L M (\<lambda>w. (\<Sum> a \<in> spectrum A. Re a* XA a w) * (\<Sum> b \<in> spectrum B. Re b *XB b w)) =
+ (\<Sum> a \<in> spectrum A. Re a * (\<Sum> b \<in> spectrum B. Re b * integral\<^sup>L M (\<lambda>w. XA a w * XB b w)))"
+ using lhv_sum_integral'[of M] assms by simp
+ also have "... = (\<Sum> a \<in> spectrum A. Re a * (\<Sum> b \<in> spectrum B. Re b *
+ Re (Complex_Matrix.trace(eigen_projector A a * (eigen_projector B b) * R))))"
+ using assms sum_qt_expect_trace[of M] by simp
+ also have "... = (\<Sum> a \<in> spectrum A. Re a * Re (\<Sum> b \<in> spectrum B. (b *
+ Complex_Matrix.trace(eigen_projector A a * (eigen_projector B b) * R))))"
+ proof (rule sum.cong, simp)
+ fix a
+ assume "a\<in> spectrum A"
+ have "(\<Sum>b\<in>spectrum B. Re b *
+ Re (Complex_Matrix.trace (eigen_projector A a * eigen_projector B b * R))) =
+ (\<Sum> b \<in> spectrum B. Re (b *
+ Complex_Matrix.trace(eigen_projector A a * (eigen_projector B b) * R)))"
+ proof (rule sum.cong, simp)
+ fix b
+ assume "b\<in> spectrum B"
+ hence "b\<in> Reals" using hermitian_spectrum_real[of B] assms by simp
+ hence "Re b = b" by simp
+ thus "Re b * Re (Complex_Matrix.trace (eigen_projector A a * eigen_projector B b * R)) =
+ Re (b * Complex_Matrix.trace (eigen_projector A a * eigen_projector B b * R))"
+ using hermitian_spectrum_real using \<open>b \<in> \<real>\<close> mult_real_cpx by auto
+ qed
+ also have "... =
+ Re (\<Sum> b \<in> spectrum B. b *
+ Complex_Matrix.trace(eigen_projector A a * (eigen_projector B b) * R))" by simp
+ finally have "(\<Sum>b\<in>spectrum B. Re b *
+ Re (Complex_Matrix.trace (eigen_projector A a * eigen_projector B b * R))) =
+ Re (\<Sum> b \<in> spectrum B. b *
+ Complex_Matrix.trace(eigen_projector A a * (eigen_projector B b) * R))" .
+ thus "Re a * (\<Sum>b\<in>spectrum B. Re b *
+ Re (Complex_Matrix.trace (eigen_projector A a * eigen_projector B b * R))) =
+ Re a * Re (\<Sum>b\<in>spectrum B.
+ (b * Complex_Matrix.trace (eigen_projector A a * eigen_projector B b * R)))"
+ by simp
+ qed
+ also have "... = (\<Sum> a \<in> spectrum A. Re (a * (\<Sum> b \<in> spectrum B. (b *
+ Complex_Matrix.trace(eigen_projector A a * (eigen_projector B b) * R)))))"
+ proof (rule sum.cong, simp)
+ fix x
+ assume "x\<in> spectrum A"
+ hence "Re x = x" using ar by simp
+ thus "Re x * Re (\<Sum>b\<in>spectrum B. b *
+ Complex_Matrix.trace (eigen_projector A x * eigen_projector B b * R)) =
+ Re (x * (\<Sum>b\<in>spectrum B. b *
+ Complex_Matrix.trace (eigen_projector A x * eigen_projector B b * R)))"
+ using \<open>x \<in> spectrum A\<close> ar mult_real_cpx by auto
+ qed
+ also have "... = Re (\<Sum> a \<in> spectrum A. a * (\<Sum> b \<in> spectrum B. (b *
+ Complex_Matrix.trace(eigen_projector A a * (eigen_projector B b) * R))))" by simp
+ also have "... = Re (Complex_Matrix.trace(A *B * R))" using assms
+ sum_eigen_projector_trace[of A B] by simp
+ finally show ?thesis unfolding qt_expect_def .
+qed
+
+
+subsection \<open>Properties of specific observables\<close>
+
+text \<open>In this part we consider a specific density operator and specific observables corresponding
+to joint bipartite measurements. We will compute the quantum expectation value of this system and
+prove that it violates the CHSH inequality, thus proving that the local hidden variable assumption
+cannot hold.\<close>
+
+subsubsection \<open>Ket 0, Ket 1 and the corresponding projectors\<close>
+
+definition ket_0::"complex Matrix.vec" where
+"ket_0 = unit_vec 2 0"
+
+lemma ket_0_dim:
+ shows "dim_vec ket_0 = 2" unfolding ket_0_def by simp
+
+lemma ket_0_norm:
+ shows "\<parallel>ket_0\<parallel> = 1" using unit_cpx_vec_length unfolding ket_0_def by simp
+
+lemma ket_0_mat:
+ shows "ket_vec ket_0 = Matrix.mat_of_cols_list 2 [[1, 0]]"
+ by (auto simp add: ket_vec_def Matrix.mat_of_cols_list_def ket_0_def)
+
+definition ket_1::"complex Matrix.vec" where
+"ket_1 = unit_vec 2 1"
+
+lemma ket_1_dim:
+ shows "dim_vec ket_1 = 2" unfolding ket_1_def by simp
+
+lemma ket_1_norm:
+ shows "\<parallel>ket_1\<parallel> = 1" using unit_cpx_vec_length unfolding ket_1_def by simp
+
+definition ket_01
+ where "ket_01 = tensor_vec ket_0 ket_1"
+
+lemma ket_01_dim:
+ shows "dim_vec ket_01 = 4" unfolding ket_01_def
+ by (simp add: ket_0_dim ket_1_dim)
+
+definition ket_10
+ where "ket_10 = tensor_vec ket_1 ket_0"
+
+lemma ket_10_dim:
+ shows "dim_vec ket_10 = 4" unfolding ket_10_def by (simp add: ket_0_dim ket_1_dim)
+
+text \<open>We define \verb+ket_psim+, one of the Bell states (or EPR pair).\<close>
+
+definition ket_psim where
+"ket_psim = 1/(sqrt 2) \<cdot>\<^sub>v (ket_01 - ket_10)"
+
+lemma ket_psim_dim:
+ shows "dim_vec ket_psim = 4" using ket_01_dim ket_10_dim unfolding ket_psim_def by simp
+
+lemma ket_psim_norm:
+ shows "\<parallel>ket_psim\<parallel> = 1"
+proof -
+ have "dim_vec ket_psim = 2\<^sup>2" unfolding ket_psim_def ket_01_def ket_10_def ket_0_def ket_1_def
+ by simp
+ moreover have "(\<Sum>i<4. (cmod (vec_index ket_psim i))\<^sup>2) = 1"
+ apply (auto simp add: ket_psim_def ket_01_def ket_10_def ket_0_def ket_1_def)
+ apply (simp add: sum_4_elems)
+ done
+ ultimately show ?thesis by (simp add: cpx_vec_length_def)
+qed
+
+text \<open>\verb+rho_psim+ represents the density operator associated with the quantum
+state \verb+ket_psim+.\<close>
+
+definition rho_psim where
+"rho_psim = rank_1_proj ket_psim"
+
+lemma rho_psim_carrier:
+ shows "rho_psim \<in> carrier_mat 4 4" using rank_1_proj_carrier[of ket_psim] ket_psim_dim
+ rho_psim_def by simp
+
+lemma rho_psim_dim_row:
+ shows "dim_row rho_psim = 4" using rho_psim_carrier by simp
+
+lemma rho_psim_density:
+ shows "density_operator rho_psim" unfolding density_operator_def
+proof
+ show "Complex_Matrix.positive rho_psim" using rank_1_proj_positive[of ket_psim] ket_psim_norm
+ rho_psim_def by simp
+ show "Complex_Matrix.trace rho_psim = 1" using rank_1_proj_trace[of ket_psim] ket_psim_norm
+ rho_psim_def by simp
+qed
+
+
+subsubsection \<open>The X and Z matrices and two of their combinations\<close>
+
+text \<open>In this part we prove properties of two standard matrices in quantum theory, $X$ and $Z$,
+as well as two of their combinations: $\frac{X+Z}{\sqrt{2}}$ and $\frac{Z - X}{\sqrt{2}}$.
+Note that all of these matrices are observables, they will be used to violate the CHSH inequality.\<close>
+
+lemma Z_carrier: shows "Z \<in> carrier_mat 2 2" unfolding Z_def by simp
+
+lemma Z_hermitian:
+ shows "hermitian Z" using dagger_adjoint dagger_of_Z unfolding hermitian_def by simp
+
+lemma unitary_Z:
+ shows "Complex_Matrix.unitary Z"
+proof -
+ have "Complex_Matrix.adjoint Z * Z = (1\<^sub>m 2)" using dagger_adjoint[of Z] by simp
+ thus ?thesis unfolding unitary_def
+ by (metis Complex_Matrix.adjoint_adjoint Complex_Matrix.unitary_def Z_carrier adjoint_dim
+ carrier_matD(1) inverts_mat_def unitary_adjoint)
+qed
+
+lemma X_carrier: shows "X \<in> carrier_mat 2 2" unfolding X_def by simp
+
+lemma X_hermitian:
+ shows "hermitian X" using dagger_adjoint dagger_of_X unfolding hermitian_def by simp
+
+lemma unitary_X:
+ shows "Complex_Matrix.unitary X"
+proof -
+ have "Complex_Matrix.adjoint X * X = (1\<^sub>m 2)" using dagger_adjoint[of X] by simp
+ thus ?thesis unfolding unitary_def
+ by (metis Complex_Matrix.adjoint_adjoint Complex_Matrix.unitary_def X_carrier adjoint_dim
+ carrier_matD(1) inverts_mat_def unitary_adjoint)
+qed
+
+definition XpZ
+ where "XpZ = -1/sqrt(2) \<cdot>\<^sub>m (X + Z)"
+
+lemma XpZ_carrier:
+ shows "XpZ \<in> carrier_mat 2 2" unfolding XpZ_def X_def Z_def by simp
+
+lemma XpZ_hermitian:
+ shows "hermitian XpZ"
+proof -
+ have "X + Z \<in> carrier_mat 2 2" using Z_carrier X_carrier by simp
+ moreover have "hermitian (X + Z)" using X_hermitian Z_hermitian hermitian_add Matrix.mat_carrier
+ unfolding X_def Z_def by blast
+ ultimately show ?thesis using hermitian_smult[of "X + Z" 2 "- 1 / sqrt 2"] unfolding XpZ_def
+ by auto
+qed
+
+lemma XpZ_inv:
+ "XpZ * XpZ = 1\<^sub>m 2" unfolding XpZ_def X_def Z_def times_mat_def one_mat_def
+ apply (rule cong_mat, simp+)
+ apply (auto simp add: Matrix.scalar_prod_def)
+ apply (auto simp add: Gates.csqrt_2_sq)
+ done
+
+lemma unitary_XpZ:
+ shows "Complex_Matrix.unitary XpZ"
+proof -
+ have "Complex_Matrix.adjoint XpZ * XpZ = (1\<^sub>m 2)" using XpZ_inv XpZ_hermitian
+ unfolding hermitian_def by simp
+ thus ?thesis unfolding unitary_def
+ by (metis Complex_Matrix.adjoint_adjoint Complex_Matrix.unitary_def XpZ_carrier adjoint_dim
+ carrier_matD(1) inverts_mat_def unitary_adjoint)
+qed
+
+definition ZmX
+ where "ZmX = 1/sqrt(2) \<cdot>\<^sub>m (Z - X)"
+
+lemma ZmX_carrier:
+ shows "ZmX \<in> carrier_mat 2 2" unfolding ZmX_def X_def Z_def
+ by (simp add: minus_carrier_mat')
+
+lemma ZmX_hermitian:
+ shows "hermitian ZmX"
+proof -
+ have "Z - X \<in> carrier_mat 2 2" unfolding X_def Z_def
+ by (simp add: minus_carrier_mat)
+ moreover have "hermitian (Z - X)" using X_hermitian Z_hermitian hermitian_minus Matrix.mat_carrier
+ unfolding X_def Z_def by blast
+ ultimately show ?thesis using hermitian_smult[of "Z - X" 2 "1 / sqrt 2"] unfolding ZmX_def
+ by auto
+qed
+
+lemma ZmX_inv:
+ "ZmX * ZmX = 1\<^sub>m 2" unfolding ZmX_def X_def Z_def times_mat_def one_mat_def
+ apply (rule cong_mat, simp+)
+ apply (auto simp add: Matrix.scalar_prod_def)
+ apply (auto simp add: Gates.csqrt_2_sq)
+ done
+
+lemma unitary_ZmX:
+ shows "Complex_Matrix.unitary ZmX"
+proof -
+ have "Complex_Matrix.adjoint ZmX * ZmX = (1\<^sub>m 2)" using ZmX_inv ZmX_hermitian
+ unfolding hermitian_def by simp
+ thus ?thesis unfolding unitary_def
+ by (metis Complex_Matrix.adjoint_adjoint Complex_Matrix.unitary_def ZmX_carrier adjoint_dim
+ carrier_matD(1) inverts_mat_def unitary_adjoint)
+qed
+
+definition Z_XpZ
+ where "Z_XpZ = tensor_mat Z XpZ"
+
+lemma Z_XpZ_carrier:
+ shows "Z_XpZ \<in> carrier_mat 4 4" unfolding Z_XpZ_def using tensor_mat_carrier XpZ_carrier
+ Z_carrier by auto
+
+definition X_XpZ
+ where "X_XpZ = tensor_mat X XpZ"
+
+lemma X_XpZ_carrier:
+ shows "X_XpZ \<in> carrier_mat 4 4" using tensor_mat_carrier XpZ_carrier X_carrier
+ unfolding X_XpZ_def by auto
+
+definition Z_ZmX
+ where "Z_ZmX = tensor_mat Z ZmX"
+
+lemma Z_ZmX_carrier:
+ shows "Z_ZmX \<in> carrier_mat 4 4" using tensor_mat_carrier ZmX_carrier Z_carrier
+ unfolding Z_ZmX_def by auto
+
+definition X_ZmX
+ where "X_ZmX = tensor_mat X ZmX"
+
+lemma X_ZmX_carrier:
+ shows "X_ZmX \<in> carrier_mat 4 4" using tensor_mat_carrier X_carrier ZmX_carrier
+ unfolding X_ZmX_def by auto
+
+lemma X_ZmX_rho_psim[simp]:
+ shows "Complex_Matrix.trace (rho_psim * X_ZmX) = 1/ (sqrt 2)"
+ apply (auto simp add: ket_10_def ket_1_def X_ZmX_def ZmX_def X_def ket_01_def
+ Z_def rho_psim_def ket_psim_def rank_1_proj_def outer_prod_def ket_0_def)
+ apply (auto simp add: Complex_Matrix.trace_def)
+ apply (simp add: sum_4_elems)
+ apply (simp add: csqrt_2_sq)
+ done
+
+lemma Z_ZmX_rho_psim[simp]:
+ shows "Complex_Matrix.trace (rho_psim * Z_ZmX) = -1/ (sqrt 2)"
+ apply (auto simp add: rho_psim_def ket_psim_def ket_10_def)
+ apply (auto simp add: Z_ZmX_def Z_def ZmX_def X_def)
+ apply (auto simp add: rank_1_proj_def outer_prod_def ket_01_def ket_1_def ket_0_def)
+ apply (auto simp add: Complex_Matrix.trace_def sum_4_elems)
+ apply (simp add: csqrt_2_sq)
+ done
+
+lemma X_XpZ_rho_psim[simp]:
+ shows "Complex_Matrix.trace (rho_psim * X_XpZ) =1/ (sqrt 2)"
+ apply (auto simp add: rho_psim_def ket_psim_def ket_10_def)
+ apply (auto simp add: X_XpZ_def Z_def XpZ_def X_def)
+ apply (auto simp add: rank_1_proj_def outer_prod_def ket_01_def ket_1_def ket_0_def)
+ apply (auto simp add: Complex_Matrix.trace_def sum_4_elems)
+ apply (simp add: csqrt_2_sq)
+ done
+
+lemma Z_XpZ_rho_psim[simp]:
+ shows "Complex_Matrix.trace (rho_psim * Z_XpZ) =1/ (sqrt 2)"
+ apply (auto simp add: rho_psim_def ket_psim_def ket_10_def)
+ apply (auto simp add: Z_XpZ_def XpZ_def X_def Z_def)
+ apply (auto simp add: rank_1_proj_def outer_prod_def ket_01_def ket_1_def ket_0_def)
+ apply (auto simp add: Complex_Matrix.trace_def sum_4_elems)
+ apply (simp add: csqrt_2_sq)
+ done
+
+definition Z_I where
+"Z_I = tensor_mat Z (1\<^sub>m 2)"
+
+lemma Z_I_carrier:
+ shows "Z_I \<in> carrier_mat 4 4" using tensor_mat_carrier Z_carrier unfolding Z_I_def by auto
+
+lemma Z_I_hermitian:
+ shows "hermitian Z_I" unfolding Z_I_def using tensor_mat_hermitian[of Z 2 "1\<^sub>m 2" 2]
+ by (simp add: Z_carrier Z_hermitian hermitian_one)
+
+lemma Z_I_unitary:
+ shows "unitary Z_I" unfolding Z_I_def using tensor_mat_unitary[of Z "1\<^sub>m 2"] Z_carrier unitary_Z
+ using unitary_one by auto
+
+lemma Z_I_spectrum:
+ shows "{Re x |x. x \<in> spectrum Z_I} \<subseteq> {- 1, 1}" using unitary_hermitian_Re_spectrum Z_I_hermitian
+ Z_I_unitary Z_I_carrier by simp
+
+definition X_I where
+"X_I = tensor_mat X (1\<^sub>m 2)"
+
+lemma X_I_carrier:
+ shows "X_I \<in> carrier_mat 4 4" using tensor_mat_carrier X_carrier unfolding X_I_def by auto
+
+lemma X_I_hermitian:
+ shows "hermitian X_I" unfolding X_I_def using tensor_mat_hermitian[of X 2 "1\<^sub>m 2" 2]
+ by (simp add: X_carrier X_hermitian hermitian_one)
+
+lemma X_I_unitary:
+ shows "unitary X_I" unfolding X_I_def using tensor_mat_unitary[of X "1\<^sub>m 2"] X_carrier unitary_X
+ using unitary_one by auto
+
+lemma X_I_spectrum:
+ shows "{Re x |x. x \<in> spectrum X_I} \<subseteq> {- 1, 1}" using unitary_hermitian_Re_spectrum X_I_hermitian
+ X_I_unitary X_I_carrier by simp
+
+definition I_XpZ where
+"I_XpZ = tensor_mat (1\<^sub>m 2) XpZ"
+
+lemma I_XpZ_carrier:
+ shows "I_XpZ \<in> carrier_mat 4 4" using tensor_mat_carrier XpZ_carrier unfolding I_XpZ_def by auto
+
+lemma I_XpZ_hermitian:
+ shows "hermitian I_XpZ" unfolding I_XpZ_def using tensor_mat_hermitian[of "1\<^sub>m 2" 2 XpZ 2]
+ by (simp add: XpZ_carrier XpZ_hermitian hermitian_one)
+
+lemma I_XpZ_unitary:
+ shows "unitary I_XpZ" unfolding I_XpZ_def using tensor_mat_unitary[of "1\<^sub>m 2" XpZ] XpZ_carrier
+ unitary_XpZ using unitary_one by auto
+
+lemma I_XpZ_spectrum:
+ shows "{Re x |x. x \<in> spectrum I_XpZ} \<subseteq> {- 1, 1}" using unitary_hermitian_Re_spectrum
+ I_XpZ_hermitian I_XpZ_unitary I_XpZ_carrier by simp
+
+definition I_ZmX where
+"I_ZmX = tensor_mat (1\<^sub>m 2) ZmX"
+
+lemma I_ZmX_carrier:
+ shows "I_ZmX \<in> carrier_mat 4 4" using tensor_mat_carrier ZmX_carrier unfolding I_ZmX_def by auto
+
+lemma I_ZmX_hermitian:
+ shows "hermitian I_ZmX" unfolding I_ZmX_def using tensor_mat_hermitian[of "1\<^sub>m 2" 2 ZmX 2]
+ by (simp add: ZmX_carrier ZmX_hermitian hermitian_one)
+
+lemma I_ZmX_unitary:
+ shows "unitary I_ZmX" unfolding I_ZmX_def using tensor_mat_unitary[of "1\<^sub>m 2" ZmX] ZmX_carrier
+ unitary_ZmX using unitary_one by auto
+
+lemma I_ZmX_spectrum:
+ shows "{Re x |x. x \<in> spectrum I_ZmX} \<subseteq> {- 1, 1}" using unitary_hermitian_Re_spectrum I_ZmX_hermitian
+ I_ZmX_unitary I_ZmX_carrier by simp
+
+lemma X_I_ZmX_eq:
+ shows "X_I * I_ZmX = X_ZmX" unfolding X_I_def I_ZmX_def X_ZmX_def using mult_distr_tensor
+ by (metis (no_types, lifting) X_inv ZmX_inv index_mult_mat(2) index_mult_mat(3) index_one_mat(2)
+ index_one_mat(3) left_mult_one_mat' pos2 right_mult_one_mat')
+
+lemma X_I_XpZ_eq:
+ shows "X_I * I_XpZ = X_XpZ" unfolding X_I_def I_XpZ_def X_XpZ_def using mult_distr_tensor
+ by (metis (no_types, lifting) X_inv XpZ_inv index_mult_mat(2) index_mult_mat(3) index_one_mat(2)
+ index_one_mat(3) left_mult_one_mat' pos2 right_mult_one_mat')
+
+lemma Z_I_XpZ_eq:
+ shows "Z_I * I_XpZ = Z_XpZ" unfolding Z_I_def I_XpZ_def Z_XpZ_def using mult_distr_tensor
+ by (metis (no_types, lifting) Z_inv XpZ_inv index_mult_mat(2) index_mult_mat(3) index_one_mat(2)
+ index_one_mat(3) left_mult_one_mat' pos2 right_mult_one_mat')
+
+lemma Z_I_ZmX_eq:
+ shows "Z_I * I_ZmX = Z_ZmX" unfolding Z_I_def I_ZmX_def Z_ZmX_def using mult_distr_tensor
+ by (metis (no_types, lifting) Z_inv ZmX_inv index_mult_mat(2) index_mult_mat(3) index_one_mat(2)
+ index_one_mat(3) left_mult_one_mat' pos2 right_mult_one_mat')
+
+
+subsubsection \<open>No local hidden variable\<close>
+
+text \<open>We show that the local hidden variable hypothesis cannot hold by exhibiting a quantum
+expectation value that is greater than the upper-bound givven by the CHSH inequality.\<close>
+
+locale bin_cpx = cpx_sq_mat +
+ assumes dim4: "dimR = 4"
+
+lemma (in bin_cpx) X_I_XpZ_trace:
+ assumes "lhv M X_I I_XpZ R Vx Vp"
+ and "R\<in> fc_mats"
+ shows "LINT w|M. (qt_expect X_I Vx w) * (qt_expect I_XpZ Vp w) =
+ Re (Complex_Matrix.trace (R * X_XpZ))"
+proof -
+ have "LINT w|M. (qt_expect X_I Vx w) * (qt_expect I_XpZ Vp w) =
+ Re (Complex_Matrix.trace (X_I * I_XpZ * R))"
+ proof (rule sum_qt_expect, (simp add: assms))
+ show "X_I \<in> fc_mats" using X_I_carrier dim4 fc_mats_carrier dim_eq by simp
+ show "R\<in> fc_mats" using assms by simp
+ show "I_XpZ \<in> fc_mats" using I_XpZ_carrier dim4 fc_mats_carrier dim_eq by simp
+ show "hermitian X_I" using X_I_hermitian .
+ show "hermitian I_XpZ" using I_XpZ_hermitian .
+ qed
+ also have "... = Re (Complex_Matrix.trace (X_XpZ * R))" using X_I_XpZ_eq by simp
+ also have "... = Re (Complex_Matrix.trace (R * X_XpZ))"
+ proof -
+ have "Complex_Matrix.trace (X_XpZ * R) = Complex_Matrix.trace (R * X_XpZ)"
+ using trace_comm[of X_XpZ 4 R] X_XpZ_carrier assms dim4 fc_mats_carrier dim_eq by simp
+ thus ?thesis by simp
+ qed
+ finally show ?thesis .
+qed
+
+lemma (in bin_cpx) X_I_XpZ_chsh:
+ assumes "lhv M X_I I_XpZ rho_psim Vx Vp"
+ shows "LINT w|M. (qt_expect X_I Vx w) * (qt_expect I_XpZ Vp w) =
+ 1/sqrt 2"
+proof -
+ have "LINT w|M. (qt_expect X_I Vx w) * (qt_expect I_XpZ Vp w) =
+ Re (Complex_Matrix.trace (rho_psim * X_XpZ))"
+ proof (rule X_I_XpZ_trace, (simp add: assms))
+ show "rho_psim \<in> fc_mats" using rho_psim_carrier fc_mats_carrier dim_eq dim4 by simp
+ qed
+ also have "... = 1/sqrt 2" using X_XpZ_rho_psim by simp
+ finally show ?thesis .
+qed
+
+lemma (in bin_cpx) Z_I_XpZ_trace:
+ assumes "lhv M Z_I I_XpZ R Vz Vp"
+ and "R\<in> fc_mats"
+ shows "LINT w|M. (qt_expect Z_I Vz w) * (qt_expect I_XpZ Vp w) =
+ Re (Complex_Matrix.trace (R * Z_XpZ))"
+proof -
+ have "LINT w|M. (qt_expect Z_I Vz w) * (qt_expect I_XpZ Vp w) =
+ Re (Complex_Matrix.trace (Z_I * I_XpZ * R))"
+ proof (rule sum_qt_expect, (simp add: assms))
+ show "Z_I \<in> fc_mats" using Z_I_carrier dim4 fc_mats_carrier dim_eq by simp
+ show "R\<in> fc_mats" using assms by simp
+ show "I_XpZ \<in> fc_mats" using I_XpZ_carrier dim4 fc_mats_carrier dim_eq by simp
+ show "hermitian Z_I" using Z_I_hermitian .
+ show "hermitian I_XpZ" using I_XpZ_hermitian .
+ qed
+ also have "... = Re (Complex_Matrix.trace (Z_XpZ * R))" using Z_I_XpZ_eq by simp
+ also have "... = Re (Complex_Matrix.trace (R * Z_XpZ))"
+ proof -
+ have "Complex_Matrix.trace (Z_XpZ * R) = Complex_Matrix.trace (R * Z_XpZ)"
+ using trace_comm[of Z_XpZ 4 R] Z_XpZ_carrier assms dim4 fc_mats_carrier dim_eq by simp
+ thus ?thesis by simp
+ qed
+ finally show ?thesis .
+qed
+
+lemma (in bin_cpx) Z_I_XpZ_chsh:
+ assumes "lhv M Z_I I_XpZ rho_psim Vz Vp"
+ shows "LINT w|M. (qt_expect Z_I Vz w) * (qt_expect I_XpZ Vp w) =
+ 1/sqrt 2"
+proof -
+ have "LINT w|M. (qt_expect Z_I Vz w) * (qt_expect I_XpZ Vp w) =
+ Re (Complex_Matrix.trace (rho_psim * Z_XpZ))"
+ proof (rule Z_I_XpZ_trace, (simp add: assms))
+ show "rho_psim \<in> fc_mats" using rho_psim_carrier fc_mats_carrier dim_eq dim4 by simp
+ qed
+ also have "... = 1/sqrt 2" using Z_XpZ_rho_psim by simp
+ finally show ?thesis unfolding qt_expect_def .
+qed
+
+lemma (in bin_cpx) X_I_ZmX_trace:
+ assumes "lhv M X_I I_ZmX R Vx Vp"
+ and "R\<in> fc_mats"
+ shows "LINT w|M. (qt_expect X_I Vx w) * (qt_expect I_ZmX Vp w) =
+ Re (Complex_Matrix.trace (R * X_ZmX))"
+proof -
+ have "LINT w|M. (qt_expect X_I Vx w) * (qt_expect I_ZmX Vp w) =
+ Re (Complex_Matrix.trace (X_I * I_ZmX * R))"
+ proof (rule sum_qt_expect, (simp add: assms))
+ show "X_I \<in> fc_mats" using X_I_carrier dim4 fc_mats_carrier dim_eq by simp
+ show "R\<in> fc_mats" using assms by simp
+ show "I_ZmX \<in> fc_mats" using I_ZmX_carrier dim4 fc_mats_carrier dim_eq by simp
+ show "hermitian X_I" using X_I_hermitian .
+ show "hermitian I_ZmX" using I_ZmX_hermitian .
+ qed
+ also have "... = Re (Complex_Matrix.trace (X_ZmX * R))" using X_I_ZmX_eq by simp
+ also have "... = Re (Complex_Matrix.trace (R * X_ZmX))"
+ proof -
+ have "Complex_Matrix.trace (X_ZmX * R) = Complex_Matrix.trace (R * X_ZmX)"
+ using trace_comm[of X_ZmX 4 R] X_ZmX_carrier assms dim4 fc_mats_carrier dim_eq by simp
+ thus ?thesis by simp
+ qed
+ finally show ?thesis .
+qed
+
+lemma (in bin_cpx) X_I_ZmX_chsh:
+ assumes "lhv M X_I I_ZmX rho_psim Vx Vp"
+ shows "LINT w|M. (qt_expect X_I Vx w) * (qt_expect I_ZmX Vp w) =
+ 1/sqrt 2"
+proof -
+ have "LINT w|M. (qt_expect X_I Vx w) * (qt_expect I_ZmX Vp w) =
+ Re (Complex_Matrix.trace (rho_psim * X_ZmX))"
+ proof (rule X_I_ZmX_trace, (simp add: assms))
+ show "rho_psim \<in> fc_mats" using rho_psim_carrier fc_mats_carrier dim_eq dim4 by simp
+ qed
+ also have "... = 1/sqrt 2" using X_ZmX_rho_psim by simp
+ finally show ?thesis unfolding qt_expect_def .
+qed
+
+lemma (in bin_cpx) Z_I_ZmX_trace:
+ assumes "lhv M Z_I I_ZmX R Vz Vp"
+ and "R\<in> fc_mats"
+ shows "LINT w|M. (qt_expect Z_I Vz w) * (qt_expect I_ZmX Vp w) =
+ Re (Complex_Matrix.trace (R * Z_ZmX))"
+proof -
+ have "LINT w|M. (qt_expect Z_I Vz w) * (qt_expect I_ZmX Vp w) =
+ Re (Complex_Matrix.trace (Z_I * I_ZmX * R))"
+ proof (rule sum_qt_expect, (simp add: assms))
+ show "Z_I \<in> fc_mats" using Z_I_carrier dim4 fc_mats_carrier dim_eq by simp
+ show "R\<in> fc_mats" using assms by simp
+ show "I_ZmX \<in> fc_mats" using I_ZmX_carrier dim4 fc_mats_carrier dim_eq by simp
+ show "hermitian Z_I" using Z_I_hermitian .
+ show "hermitian I_ZmX" using I_ZmX_hermitian .
+ qed
+ also have "... = Re (Complex_Matrix.trace (Z_ZmX * R))" using Z_I_ZmX_eq by simp
+ also have "... = Re (Complex_Matrix.trace (R * Z_ZmX))"
+ proof -
+ have "Complex_Matrix.trace (Z_ZmX * R) = Complex_Matrix.trace (R * Z_ZmX)"
+ using trace_comm[of Z_ZmX 4 R] Z_ZmX_carrier assms dim4 fc_mats_carrier dim_eq by simp
+ thus ?thesis by simp
+ qed
+ finally show ?thesis .
+qed
+
+lemma (in bin_cpx) Z_I_ZmX_chsh:
+ assumes "lhv M Z_I I_ZmX rho_psim Vz Vp"
+shows "LINT w|M. (qt_expect Z_I Vz w) * (qt_expect I_ZmX Vp w) =
+ -1/sqrt 2"
+proof -
+ have "LINT w|M. (qt_expect Z_I Vz w) * (qt_expect I_ZmX Vp w) =
+ Re (Complex_Matrix.trace (rho_psim * Z_ZmX))"
+ proof (rule Z_I_ZmX_trace, (simp add: assms))
+ show "rho_psim \<in> fc_mats" using rho_psim_carrier fc_mats_carrier dim_eq dim4 by simp
+ qed
+ also have "... = -1/sqrt 2" using Z_ZmX_rho_psim by simp
+ finally show ?thesis unfolding qt_expect_def .
+qed
+
+lemma (in bin_cpx) chsh_upper_bound:
+ assumes "prob_space M"
+ and "lhv M X_I I_XpZ rho_psim Vx Vp"
+ and "lhv M Z_I I_XpZ rho_psim Vz Vp"
+ and "lhv M X_I I_ZmX rho_psim Vx Vm"
+ and "lhv M Z_I I_ZmX rho_psim Vz Vm"
+shows "\<bar>(LINT w|M. qt_expect X_I Vx w * qt_expect I_ZmX Vm w) +
+ (LINT w|M. qt_expect Z_I Vz w * qt_expect I_XpZ Vp w) +
+ (LINT w|M. qt_expect X_I Vx w * qt_expect I_XpZ Vp w) -
+ (LINT w|M. qt_expect Z_I Vz w * qt_expect I_ZmX Vm w)\<bar>
+ \<le> 2"
+proof (rule prob_space.chsh_expect)
+ show "AE w in M. \<bar>qt_expect X_I Vx w\<bar> \<le> 1" unfolding qt_expect_def
+ proof (rule spectrum_abs_1_weighted_suml)
+ show "X_I \<in> fc_mats" using X_I_carrier fc_mats_carrier dim_eq dim4 by simp
+ show "hermitian X_I" using X_I_hermitian .
+ show "lhv M X_I I_XpZ rho_psim Vx Vp" using assms by simp
+ show "{Re x |x. x \<in> spectrum X_I} \<subseteq> {- 1, 1}" using X_I_spectrum by simp
+ show "{Re x |x. x \<in> spectrum X_I} \<noteq> {}" using spectrum_ne X_I_hermitian \<open>X_I \<in> fc_mats\<close> by auto
+ qed
+ show "AE w in M. \<bar>qt_expect Z_I Vz w\<bar> \<le> 1" unfolding qt_expect_def
+ proof (rule spectrum_abs_1_weighted_suml)
+ show "Z_I \<in> fc_mats" using Z_I_carrier fc_mats_carrier dim_eq dim4 by simp
+ show "hermitian Z_I" using Z_I_hermitian .
+ show "lhv M Z_I I_XpZ rho_psim Vz Vp" using assms by simp
+ show "{Re x |x. x \<in> spectrum Z_I} \<subseteq> {- 1, 1}" using Z_I_spectrum by simp
+ show "{Re x |x. x \<in> spectrum Z_I} \<noteq> {}" using spectrum_ne Z_I_hermitian \<open>Z_I \<in> fc_mats\<close> by auto
+ qed
+ show "AE w in M. \<bar>qt_expect I_XpZ Vp w\<bar> \<le> 1" unfolding qt_expect_def
+ proof (rule spectrum_abs_1_weighted_sumr)
+ show "I_XpZ \<in> fc_mats" using I_XpZ_carrier fc_mats_carrier dim_eq dim4 by simp
+ show "hermitian I_XpZ" using I_XpZ_hermitian .
+ show "lhv M Z_I I_XpZ rho_psim Vz Vp" using assms by simp
+ show "{Re x |x. x \<in> spectrum I_XpZ} \<subseteq> {- 1, 1}" using I_XpZ_spectrum by simp
+ show "{Re x |x. x \<in> spectrum I_XpZ} \<noteq> {}" using spectrum_ne I_XpZ_hermitian \<open>I_XpZ \<in> fc_mats\<close>
+ by auto
+ qed
+ show "AE w in M. \<bar>qt_expect I_ZmX Vm w\<bar> \<le> 1" unfolding qt_expect_def
+ proof (rule spectrum_abs_1_weighted_sumr)
+ show "I_ZmX \<in> fc_mats" using I_ZmX_carrier fc_mats_carrier dim_eq dim4 by simp
+ show "hermitian I_ZmX" using I_ZmX_hermitian .
+ show "lhv M Z_I I_ZmX rho_psim Vz Vm" using assms by simp
+ show "{Re x |x. x \<in> spectrum I_ZmX} \<subseteq> {- 1, 1}" using I_ZmX_spectrum by simp
+ show "{Re x |x. x \<in> spectrum I_ZmX} \<noteq> {}" using spectrum_ne I_ZmX_hermitian \<open>I_ZmX \<in> fc_mats\<close>
+ by auto
+ qed
+ show "prob_space M" using assms by simp
+ show "integrable M (\<lambda>w. qt_expect X_I Vx w * qt_expect I_ZmX Vm w)"
+ using spectr_sum_integrable[of M] assms by simp
+ show "integrable M (\<lambda>w. qt_expect Z_I Vz w * qt_expect I_ZmX Vm w)"
+ using spectr_sum_integrable[of M] assms by simp
+ show "integrable M (\<lambda>w. qt_expect X_I Vx w * qt_expect I_XpZ Vp w)"
+ using spectr_sum_integrable[of M] assms by simp
+ show "integrable M (\<lambda>w. qt_expect Z_I Vz w * qt_expect I_XpZ Vp w)"
+ using spectr_sum_integrable[of M] assms by simp
+qed
+
+lemma (in bin_cpx) quantum_value:
+ assumes "lhv M X_I I_XpZ rho_psim Vx Vp"
+ and "lhv M Z_I I_XpZ rho_psim Vz Vp"
+ and "lhv M X_I I_ZmX rho_psim Vx Vm"
+ and "lhv M Z_I I_ZmX rho_psim Vz Vm"
+shows "\<bar>(LINT w|M. qt_expect X_I Vx w * qt_expect I_ZmX Vm w) +
+ (LINT w|M. qt_expect Z_I Vz w * qt_expect I_XpZ Vp w) +
+ (LINT w|M. qt_expect X_I Vx w * qt_expect I_XpZ Vp w) -
+ (LINT w|M. qt_expect Z_I Vz w * qt_expect I_ZmX Vm w)\<bar>
+ = 2* sqrt 2"
+proof -
+ have "LINT w|M. qt_expect X_I Vx w * qt_expect I_ZmX Vm w = 1/sqrt 2"
+ using X_I_ZmX_chsh[of M] assms unfolding qt_expect_def by simp
+ moreover have b: "LINT w|M. qt_expect X_I Vx w * qt_expect I_XpZ Vp w = 1/sqrt 2"
+ using X_I_XpZ_chsh[of M] assms unfolding qt_expect_def by simp
+ moreover have c: "LINT w|M. qt_expect Z_I Vz w * qt_expect I_XpZ Vp w = 1/sqrt 2"
+ using Z_I_XpZ_chsh[of M] assms unfolding qt_expect_def by simp
+ moreover have "LINT w|M. qt_expect Z_I Vz w * qt_expect I_ZmX Vm w = -1/sqrt 2"
+ using Z_I_ZmX_chsh[of M] assms unfolding qt_expect_def by simp
+ ultimately have "(LINT w|M. qt_expect X_I Vx w * qt_expect I_ZmX Vm w) +
+ (LINT w|M. qt_expect X_I Vx w * qt_expect I_XpZ Vp w) +
+ (LINT w|M. qt_expect Z_I Vz w * qt_expect I_XpZ Vp w) -
+ (LINT w|M. qt_expect Z_I Vz w * qt_expect I_ZmX Vm w) = 4/(sqrt 2)" by simp
+ also have "... = (4 * (sqrt 2))/(sqrt 2 * (sqrt 2))"
+ by (metis mult_numeral_1_right real_divide_square_eq times_divide_eq_right)
+ also have "... = (4 * (sqrt 2)) / 2" by simp
+ also have "... = 2 * (sqrt 2)" by simp
+ finally have "(LINT w|M. qt_expect X_I Vx w * qt_expect I_ZmX Vm w) +
+ (LINT w|M. qt_expect X_I Vx w * qt_expect I_XpZ Vp w) +
+ (LINT w|M. qt_expect Z_I Vz w * qt_expect I_XpZ Vp w) -
+ (LINT w|M. qt_expect Z_I Vz w * qt_expect I_ZmX Vm w) = 2 * sqrt 2" .
+ thus ?thesis by (simp add: b c)
+qed
+
+lemma (in bin_cpx) no_lhv:
+ assumes "lhv M X_I I_XpZ rho_psim Vx Vp"
+ and "lhv M Z_I I_XpZ rho_psim Vz Vp"
+ and "lhv M X_I I_ZmX rho_psim Vx Vm"
+ and "lhv M Z_I I_ZmX rho_psim Vz Vm"
+shows False
+proof -
+ have "prob_space M" using assms unfolding lhv_def by simp
+ have "2 * sqrt 2 = \<bar>(LINT w|M. qt_expect X_I Vx w * qt_expect I_ZmX Vm w) +
+ (LINT w|M. qt_expect Z_I Vz w * qt_expect I_XpZ Vp w) +
+ (LINT w|M. qt_expect X_I Vx w * qt_expect I_XpZ Vp w) -
+ (LINT w|M. qt_expect Z_I Vz w * qt_expect I_ZmX Vm w)\<bar>"
+ using assms quantum_value[of M] by simp
+ also have "... \<le> 2" using chsh_upper_bound[of M] assms \<open>prob_space M\<close> by simp
+ finally have "2 * sqrt 2 \<le> 2" .
+ thus False by simp
+qed
+
+
end
\ No newline at end of file
diff --git a/thys/Projective_Measurements/Linear_Algebra_Complements.thy b/thys/Projective_Measurements/Linear_Algebra_Complements.thy
--- a/thys/Projective_Measurements/Linear_Algebra_Complements.thy
+++ b/thys/Projective_Measurements/Linear_Algebra_Complements.thy
@@ -1,2269 +1,2269 @@
-(*
-Author:
- Mnacho Echenim, Université Grenoble Alpes
-*)
-
-theory Linear_Algebra_Complements imports
- "Isabelle_Marries_Dirac.Tensor"
- "Isabelle_Marries_Dirac.More_Tensor"
- "QHLProver.Gates"
- "HOL-Types_To_Sets.Group_On_With"
- "HOL-Probability.Probability"
-
-
-begin
-hide_const(open) S
-section \<open>Preliminaries\<close>
-
-
-subsection \<open>Misc\<close>
-
-lemma mult_real_cpx:
- fixes a::complex
- fixes b::complex
- assumes "a\<in> Reals"
- shows "a* (Re b) = Re (a * b)" using assms
- by (metis Reals_cases complex.exhaust complex.sel(1) complex_of_real_mult_Complex of_real_mult)
-
-lemma fct_bound:
- fixes f::"complex\<Rightarrow> real"
- assumes "f(-1) + f 1 = 1"
-and "0 \<le> f 1"
-and "0 \<le> f (-1)"
-shows "-1 \<le> f 1 - f(-1) \<and> f 1 - f(-1) \<le> 1"
-proof
- have "f 1 - f(-1) = 1 - f(-1) - f(-1)" using assms by simp
- also have "...\<ge> -1" using assms by simp
- finally show "-1 \<le> f 1 - f(-1)" .
-next
- have "f(-1) - f 1 = 1 - f 1 - f 1 " using assms by simp
- also have "... \<ge> -1" using assms by simp
- finally have "-1 \<le> f(-1) - f 1" .
- thus "f 1 - f (-1) \<le> 1" by simp
-qed
-
-lemma fct_bound':
- fixes f::"complex\<Rightarrow> real"
- assumes "f(-1) + f 1 = 1"
-and "0 \<le> f 1"
-and "0 \<le> f (-1)"
-shows "\<bar>f 1 - f(-1)\<bar> \<le> 1" using assms fct_bound by auto
-
-lemma pos_sum_1_le:
- assumes "finite I"
-and "\<forall> i \<in> I. (0::real) \<le> f i"
-and "(\<Sum>i\<in> I. f i) = 1"
-and "j\<in> I"
-shows "f j \<le> 1"
-proof (rule ccontr)
- assume "\<not> f j \<le> 1"
- hence "1 < f j" by simp
- hence "1 < (\<Sum>i\<in> I. f i)" using assms by (metis \<open>\<not> f j \<le> 1\<close> sum_nonneg_leq_bound)
- thus False using assms by simp
-qed
-
-lemma last_subset:
- assumes "A \<subseteq> {a,b}"
- and "a\<noteq> b"
-and "A \<noteq> {a, b}"
-and "A\<noteq> {}"
-and "A \<noteq> {a}"
-shows "A = {b}" using assms by blast
-
-lemma disjoint_Un:
- assumes "disjoint_family_on A (insert x F)"
- and "x\<notin> F"
-shows "(A x) \<inter> (\<Union> a\<in> F. A a) = {}"
-proof -
- have "(A x) \<inter> (\<Union> a\<in> F. A a) = (\<Union>i\<in>F. (A x) \<inter> A i)" using Int_UN_distrib by simp
- also have "... = (\<Union>i\<in>F. {})" using assms disjoint_family_onD by fastforce
- also have "... = {}" using SUP_bot_conv(2) by simp
- finally show ?thesis .
-qed
-
-lemma sum_but_one:
- assumes "\<forall>j < (n::nat). j \<noteq>i \<longrightarrow> f j = (0::'a::ring)"
- and "i < n"
- shows "(\<Sum> j \<in> {0 ..< n}. f j * g j) = f i * g i"
-proof -
- have "sum (\<lambda>x. f x * g x) (({0 ..< n} - {i}) \<union> {i}) = sum (\<lambda>x. f x * g x) ({0 ..< n} - {i}) +
- sum (\<lambda>x. f x * g x) {i}" by (rule sum.union_disjoint, auto)
- also have "... = sum (\<lambda>x. f x * g x) {i}" using assms by auto
- also have "... = f i * g i" by simp
- finally have "sum (\<lambda>x. f x * g x) (({0 ..< n} - {i}) \<union> {i}) = f i * g i" .
- moreover have "{0 ..< n} = ({0 ..< n} - {i}) \<union> {i}" using assms by auto
- ultimately show ?thesis by simp
-qed
-
-lemma sum_2_elems:
- assumes "I = {a,b}"
- and "a\<noteq> b"
- shows "(\<Sum>a\<in>I. f a) = f a + f b"
-proof -
- have "(\<Sum>a\<in>I. f a) = (\<Sum>a\<in>(insert a {b}). f a)" using assms by simp
- also have "... = f a + (\<Sum>a\<in>{b}. f a)"
- proof (rule sum.insert)
- show "finite {b}" by simp
- show "a\<notin> {b}" using assms by simp
- qed
- also have "... = f a + f b" by simp
- finally show ?thesis .
-qed
-
-lemma sum_4_elems:
- shows "(\<Sum>i<(4::nat). f i) = f 0 + f 1 + f 2 + f 3"
-proof -
- have "(\<Sum>i<(4::nat). f i) = (\<Sum>i<(3::nat). f i) + f 3"
- by (metis Suc_numeral semiring_norm(2) semiring_norm(8) sum.lessThan_Suc)
- moreover have "(\<Sum>i<(3::nat). f i) = (\<Sum>i<(2::nat). f i) + f 2"
- by (metis Suc_1 add_2_eq_Suc' nat_1_add_1 numeral_code(3) numerals(1)
- one_plus_numeral_commute sum.lessThan_Suc)
- moreover have "(\<Sum>i<(2::nat). f i) = (\<Sum>i<(1::nat). f i) + f 1"
- by (metis Suc_1 sum.lessThan_Suc)
- ultimately show ?thesis by simp
-qed
-
-lemma disj_family_sum:
- shows "finite I \<Longrightarrow> disjoint_family_on A I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> finite (A i)) \<Longrightarrow>
- (\<Sum> i \<in> (\<Union>n \<in> I. A n). f i) = (\<Sum> n\<in> I. (\<Sum> i \<in> A n. f i))"
-proof (induct rule:finite_induct)
- case empty
- then show ?case by simp
-next
- case (insert x F)
- hence "disjoint_family_on A F"
- by (meson disjoint_family_on_mono subset_insertI)
- have "(\<Union>n \<in> (insert x F). A n) = A x \<union> (\<Union>n \<in> F. A n)" using insert by simp
- hence "(\<Sum> i \<in> (\<Union>n \<in> (insert x F). A n). f i) = (\<Sum> i \<in> (A x \<union> (\<Union>n \<in> F. A n)). f i)" by simp
- also have "... = (\<Sum> i \<in> A x. f i) + (\<Sum> i \<in> (\<Union>n \<in> F. A n). f i)"
- by (rule sum.union_disjoint, (simp add: insert disjoint_Un)+)
- also have "... = (\<Sum> i \<in> A x. f i) + (\<Sum>n\<in>F. sum f (A n))" using \<open>disjoint_family_on A F\<close>
- by (simp add: insert)
- also have "... = (\<Sum>n\<in>(insert x F). sum f (A n))" using insert by simp
- finally show ?case .
-qed
-
-lemma integrable_real_mult_right:
- fixes c::real
- assumes "integrable M f"
- shows "integrable M (\<lambda>w. c * f w)"
-proof (cases "c = 0")
- case True
- thus ?thesis by simp
-next
- case False
- thus ?thesis using integrable_mult_right[of c] assms by simp
-qed
-
-
-subsection \<open>Unifying notions between Isabelle Marries Dirac and QHLProver\<close>
-
-lemma mult_conj_cmod_square:
- fixes z::complex
- shows "z * conjugate z = (cmod z)\<^sup>2"
-proof -
- have "z * conjugate z = (Re z)\<^sup>2 + (Im z)\<^sup>2" using complex_mult_cnj by auto
- also have "... = (cmod z)\<^sup>2" unfolding cmod_def by simp
- finally show ?thesis .
-qed
-
-lemma vec_norm_sq_cpx_vec_length_sq:
- shows "(vec_norm v)\<^sup>2 = (cpx_vec_length v)\<^sup>2"
-proof -
- have "(vec_norm v)\<^sup>2 = inner_prod v v" unfolding vec_norm_def using power2_csqrt by blast
- also have "... = (\<Sum>i<dim_vec v. (cmod (Matrix.vec_index v i))\<^sup>2)" unfolding Matrix.scalar_prod_def
- proof -
- have "\<And>i. i < dim_vec v \<Longrightarrow> Matrix.vec_index v i * conjugate (Matrix.vec_index v i) =
- (cmod (Matrix.vec_index v i))\<^sup>2" using mult_conj_cmod_square by simp
- thus "(\<Sum>i = 0..<dim_vec (conjugate v). Matrix.vec_index v i *
- Matrix.vec_index (conjugate v) i) = (\<Sum>i<dim_vec v. (cmod (Matrix.vec_index v i))\<^sup>2)"
- by (simp add: lessThan_atLeast0)
- qed
- finally show "(vec_norm v)\<^sup>2 = (cpx_vec_length v)\<^sup>2" unfolding cpx_vec_length_def
- by (simp add: sum_nonneg)
-qed
-
-lemma vec_norm_eq_cpx_vec_length:
- shows "vec_norm v = cpx_vec_length v" using vec_norm_sq_cpx_vec_length_sq
-by (metis cpx_vec_length_inner_prod inner_prod_csqrt power2_csqrt vec_norm_def)
-
-lemma cpx_vec_length_square:
- shows "\<parallel>v\<parallel>\<^sup>2 = (\<Sum>i = 0..<dim_vec v. (cmod (Matrix.vec_index v i))\<^sup>2)" unfolding cpx_vec_length_def
- by (simp add: lessThan_atLeast0 sum_nonneg)
-
-lemma state_qbit_norm_sq:
- assumes "v\<in> state_qbit n"
- shows "(cpx_vec_length v)\<^sup>2 = 1"
-proof -
- have "cpx_vec_length v = 1" using assms unfolding state_qbit_def by simp
- thus ?thesis by simp
-qed
-
-lemma dagger_adjoint:
-shows "dagger M = Complex_Matrix.adjoint M" unfolding dagger_def Complex_Matrix.adjoint_def
- by (simp add: cong_mat)
-
-
-subsection \<open>Types to sets lemmata transfers\<close>
-
-context ab_group_add_on_with begin
-
-context includes lifting_syntax assumes ltd: "\<exists>(Rep::'s \<Rightarrow> 'a) (Abs::'a \<Rightarrow> 's).
- type_definition Rep Abs S" begin
-interpretation local_typedef_ab_group_add_on_with pls z mns um S "TYPE('s)" by unfold_locales fact
-
-lemmas lt_sum_union_disjoint = sum.union_disjoint
- [var_simplified explicit_ab_group_add,
- unoverload_type 'c,
- OF type.comm_monoid_add_axioms,
- untransferred]
-
-lemmas lt_disj_family_sum = disj_family_sum
- [var_simplified explicit_ab_group_add,
- unoverload_type 'd,
-OF type.comm_monoid_add_axioms,
- untransferred]
-
-lemmas lt_sum_reindex_cong = sum.reindex_cong
- [var_simplified explicit_ab_group_add,
- unoverload_type 'd,
-OF type.comm_monoid_add_axioms,
- untransferred]
-end
-
-lemmas sum_with_union_disjoint =
- lt_sum_union_disjoint
- [cancel_type_definition,
- OF carrier_ne,
- simplified pred_fun_def, simplified]
-
-lemmas disj_family_sum_with =
- lt_disj_family_sum
- [cancel_type_definition,
- OF carrier_ne,
- simplified pred_fun_def, simplified]
-
-lemmas sum_with_reindex_cong =
- lt_sum_reindex_cong
- [cancel_type_definition,
- OF carrier_ne,
- simplified pred_fun_def, simplified]
-
-end
-
-lemma (in comm_monoid_add_on_with) sum_with_cong':
- shows "finite I \<Longrightarrow> (\<And>i. i\<in> I \<Longrightarrow> A i = B i) \<Longrightarrow> (\<And>i. i\<in> I \<Longrightarrow> A i \<in> S) \<Longrightarrow>
- (\<And>i. i\<in> I \<Longrightarrow> B i \<in> S) \<Longrightarrow> sum_with pls z A I = sum_with pls z B I"
-proof (induct rule: finite_induct)
- case empty
- then show ?case by simp
-next
- case (insert x F)
- have "sum_with pls z A (insert x F) = pls (A x) (sum_with pls z A F)" using insert
- sum_with_insert[of A] by (simp add: image_subset_iff)
- also have "... = pls (B x) (sum_with pls z B F)" using insert by simp
- also have "... = sum_with pls z B (insert x F)" using insert sum_with_insert[of B]
- by (simp add: image_subset_iff)
- finally show ?case .
-qed
-
-
-section \<open>Linear algebra complements\<close>
-
-subsection \<open>Additional properties of matrices\<close>
-
-lemma smult_one:
- shows "(1::'a::monoid_mult) \<cdot>\<^sub>m A = A" by (simp add:eq_matI)
-
-lemma times_diag_index:
- fixes A::"'a::comm_ring Matrix.mat"
- assumes "A \<in> carrier_mat n n"
-and "B\<in> carrier_mat n n"
-and "diagonal_mat B"
-and "j < n"
-and "i < n"
-shows "Matrix.vec_index (Matrix.row (A*B) j) i = diag_mat B ! i *A $$ (j, i)"
-proof -
- have "Matrix.vec_index (Matrix.row (A*B) j) i = (A*B) $$ (j,i)"
- using Matrix.row_def[of "A*B" ] assms by simp
- also have "... = Matrix.scalar_prod (Matrix.row A j) (Matrix.col B i)" using assms
- times_mat_def[of A] by simp
- also have "... = Matrix.scalar_prod (Matrix.col B i) (Matrix.row A j)"
- using comm_scalar_prod[of "Matrix.row A j" n] assms by auto
- also have "... = (Matrix.vec_index (Matrix.col B i) i) * (Matrix.vec_index (Matrix.row A j) i)"
- unfolding Matrix.scalar_prod_def
- proof (rule sum_but_one)
- show "i < dim_vec (Matrix.row A j)" using assms by simp
- show "\<forall>ia<dim_vec (Matrix.row A j). ia \<noteq> i \<longrightarrow> Matrix.vec_index (Matrix.col B i) ia = 0" using assms
- by (metis carrier_matD(1) carrier_matD(2) diagonal_mat_def index_col index_row(2))
- qed
- also have "... = B $$(i,i) * (Matrix.vec_index (Matrix.row A j) i)" using assms by auto
- also have "... = diag_mat B ! i * (Matrix.vec_index (Matrix.row A j) i)" unfolding diag_mat_def
- using assms by simp
- also have "... = diag_mat B ! i * A $$ (j, i)" using assms by simp
- finally show ?thesis .
-qed
-
-lemma inner_prod_adjoint_comp:
- assumes "(U::'a::conjugatable_field Matrix.mat) \<in> carrier_mat n n"
-and "(V::'a::conjugatable_field Matrix.mat) \<in> carrier_mat n n"
-and "i < n"
-and "j < n"
-shows "Complex_Matrix.inner_prod (Matrix.col V i) (Matrix.col U j) =
- ((Complex_Matrix.adjoint V) * U) $$ (i, j)"
-proof -
- have "Complex_Matrix.inner_prod (Matrix.col V i) (Matrix.col U j) =
- Matrix.scalar_prod (Matrix.col U j) (Matrix.row (Complex_Matrix.adjoint V) i)"
- using adjoint_row[of i V] assms by simp
- also have "... = Matrix.scalar_prod (Matrix.row (Complex_Matrix.adjoint V) i) (Matrix.col U j)"
- by (metis adjoint_row assms(1) assms(2) assms(3) carrier_matD(1) carrier_matD(2) Matrix.col_dim
- conjugate_vec_sprod_comm)
- also have "... = ((Complex_Matrix.adjoint V) * U) $$ (i, j)" using assms
- by (simp add:times_mat_def)
- finally show ?thesis .
-qed
-
-lemma mat_unit_vec_col:
- assumes "(A::'a::conjugatable_field Matrix.mat) \<in> carrier_mat n n"
-and "i < n"
-shows "A *\<^sub>v (unit_vec n i) = Matrix.col A i"
-proof
- show "dim_vec (A *\<^sub>v unit_vec n i) = dim_vec (Matrix.col A i)" using assms by simp
- show "\<And>j. j < dim_vec (Matrix.col A i) \<Longrightarrow> Matrix.vec_index (A *\<^sub>v unit_vec n i) j =
- Matrix.vec_index (Matrix.col A i) j"
- proof -
- fix j
- assume "j < dim_vec (Matrix.col A i)"
- hence "Matrix.vec_index (A *\<^sub>v unit_vec n i) j =
- Matrix.scalar_prod (Matrix.row A j) (unit_vec n i)" unfolding mult_mat_vec_def by simp
- also have "... = Matrix.scalar_prod (unit_vec n i) (Matrix.row A j)" using comm_scalar_prod
- assms by auto
- also have "... = (Matrix.vec_index (unit_vec n i) i) * (Matrix.vec_index (Matrix.row A j) i)"
- unfolding Matrix.scalar_prod_def
- proof (rule sum_but_one)
- show "i < dim_vec (Matrix.row A j)" using assms by auto
- show "\<forall>ia<dim_vec (Matrix.row A j). ia \<noteq> i \<longrightarrow> Matrix.vec_index (unit_vec n i) ia = 0"
- using assms unfolding unit_vec_def by auto
- qed
- also have "... = (Matrix.vec_index (Matrix.row A j) i)" using assms by simp
- also have "... = A $$ (j, i)" using assms \<open>j < dim_vec (Matrix.col A i)\<close> by simp
- also have "... = Matrix.vec_index (Matrix.col A i) j" using assms \<open>j < dim_vec (Matrix.col A i)\<close> by simp
- finally show "Matrix.vec_index (A *\<^sub>v unit_vec n i) j =
- Matrix.vec_index (Matrix.col A i) j" .
- qed
-qed
-
-lemma mat_prod_unit_vec_cong:
- assumes "(A::'a::conjugatable_field Matrix.mat) \<in> carrier_mat n n"
-and "B\<in> carrier_mat n n"
-and "\<And>i. i < n \<Longrightarrow> A *\<^sub>v (unit_vec n i) = B *\<^sub>v (unit_vec n i)"
-shows "A = B"
-proof
- show "dim_row A = dim_row B" using assms by simp
- show "dim_col A = dim_col B" using assms by simp
- show "\<And>i j. i < dim_row B \<Longrightarrow> j < dim_col B \<Longrightarrow> A $$ (i, j) = B $$ (i, j)"
- proof -
- fix i j
- assume ij: "i < dim_row B" "j < dim_col B"
- hence "A $$ (i,j) = Matrix.vec_index (Matrix.col A j) i" using assms by simp
- also have "... = Matrix.vec_index (A *\<^sub>v (unit_vec n j)) i" using mat_unit_vec_col[of A] ij assms
- by simp
- also have "... = Matrix.vec_index (B *\<^sub>v (unit_vec n j)) i" using assms ij by simp
- also have "... = Matrix.vec_index (Matrix.col B j) i" using mat_unit_vec_col ij assms by simp
- also have "... = B $$ (i,j)" using assms ij by simp
- finally show "A $$ (i, j) = B $$ (i, j)" .
- qed
-qed
-
-lemma smult_smult_times:
- fixes a::"'a::semigroup_mult"
- shows "a\<cdot>\<^sub>m (k \<cdot>\<^sub>m A) = (a * k)\<cdot>\<^sub>m A"
-proof
- show r:"dim_row (a \<cdot>\<^sub>m (k \<cdot>\<^sub>m A)) = dim_row (a * k \<cdot>\<^sub>m A)" by simp
- show c:"dim_col (a \<cdot>\<^sub>m (k \<cdot>\<^sub>m A)) = dim_col (a * k \<cdot>\<^sub>m A)" by simp
- show "\<And>i j. i < dim_row (a * k \<cdot>\<^sub>m A) \<Longrightarrow>
- j < dim_col (a * k \<cdot>\<^sub>m A) \<Longrightarrow> (a \<cdot>\<^sub>m (k \<cdot>\<^sub>m A)) $$ (i, j) = (a * k \<cdot>\<^sub>m A) $$ (i, j)"
- proof -
- fix i j
- assume "i < dim_row (a * k \<cdot>\<^sub>m A)" and "j < dim_col (a * k \<cdot>\<^sub>m A)" note ij = this
- hence "(a \<cdot>\<^sub>m (k \<cdot>\<^sub>m A)) $$ (i, j) = a * (k \<cdot>\<^sub>m A) $$ (i, j)" by simp
- also have "... = a * (k * A $$ (i,j))" using ij by simp
- also have "... = (a * k) * A $$ (i,j)"
- by (simp add: semigroup_mult_class.mult.assoc)
- also have "... = (a * k \<cdot>\<^sub>m A) $$ (i, j)" using r c ij by simp
- finally show "(a \<cdot>\<^sub>m (k \<cdot>\<^sub>m A)) $$ (i, j) = (a * k \<cdot>\<^sub>m A) $$ (i, j)" .
- qed
-qed
-
-lemma mat_minus_minus:
- fixes A :: "'a :: ab_group_add Matrix.mat"
- assumes "A \<in> carrier_mat n m"
- and "B\<in> carrier_mat n m"
- and "C\<in> carrier_mat n m"
-shows "A - (B - C) = A - B + C"
-proof
- show "dim_row (A - (B - C)) = dim_row (A - B + C)" using assms by simp
- show "dim_col (A - (B - C)) = dim_col (A - B + C)" using assms by simp
- show "\<And>i j. i < dim_row (A - B + C) \<Longrightarrow> j < dim_col (A - B + C) \<Longrightarrow>
- (A - (B - C)) $$ (i, j) = (A - B + C) $$ (i, j)"
- proof -
- fix i j
- assume "i < dim_row (A - B + C)" and "j < dim_col (A - B + C)" note ij = this
- have "(A - (B - C)) $$ (i, j) = (A $$ (i,j) - B $$ (i,j) + C $$ (i,j))" using ij assms by simp
- also have "... = (A - B + C) $$ (i, j)" using assms ij by simp
- finally show "(A - (B - C)) $$ (i, j) = (A - B + C) $$ (i, j)" .
- qed
-qed
-
-
-subsection \<open>Complements on complex matrices\<close>
-
-lemma hermitian_square:
- assumes "hermitian M"
- shows "M \<in> carrier_mat (dim_row M) (dim_row M)"
-proof -
- have "dim_col M = dim_row M" using assms unfolding hermitian_def adjoint_def
- by (metis adjoint_dim_col)
- thus ?thesis by auto
-qed
-
-lemma hermitian_add:
- assumes "A\<in> carrier_mat n n"
- and "B\<in> carrier_mat n n"
-and "hermitian A"
-and "hermitian B"
-shows "hermitian (A + B)" unfolding hermitian_def
- by (metis adjoint_add assms hermitian_def)
-
-lemma hermitian_minus:
- assumes "A\<in> carrier_mat n n"
- and "B\<in> carrier_mat n n"
-and "hermitian A"
-and "hermitian B"
-shows "hermitian (A - B)" unfolding hermitian_def
- by (metis adjoint_minus assms hermitian_def)
-
-lemma hermitian_smult:
- fixes a::real
- fixes A::"complex Matrix.mat"
- assumes "A \<in> carrier_mat n n"
-and "hermitian A"
-shows "hermitian (a \<cdot>\<^sub>m A)"
-proof -
- have dim: "Complex_Matrix.adjoint A \<in> carrier_mat n n" using assms by (simp add: adjoint_dim)
- {
- fix i j
- assume "i < n" and "j < n"
- hence "Complex_Matrix.adjoint (a \<cdot>\<^sub>m A) $$ (i,j) = a * (Complex_Matrix.adjoint A $$ (i,j))"
- using adjoint_scale[of a A] assms by simp
- also have "... = a * (A $$ (i,j))" using assms unfolding hermitian_def by simp
- also have "... = (a \<cdot>\<^sub>m A) $$ (i,j)" using \<open>i < n\<close> \<open>j < n\<close> assms by simp
- finally have "Complex_Matrix.adjoint (a \<cdot>\<^sub>m A) $$ (i,j) = (a \<cdot>\<^sub>m A) $$ (i,j)" .
- }
- thus ?thesis using dim assms unfolding hermitian_def by auto
-qed
-
-lemma unitary_eigenvalues_norm_square:
- fixes U::"complex Matrix.mat"
- assumes "unitary U"
- and "U \<in> carrier_mat n n"
- and "eigenvalue U k"
-shows "conjugate k * k = 1"
-proof -
- have "\<exists>v. eigenvector U v k" using assms unfolding eigenvalue_def by simp
- from this obtain v where "eigenvector U v k" by auto
- define vn where "vn = vec_normalize v"
- have "eigenvector U vn k" using normalize_keep_eigenvector \<open>eigenvector U v k\<close>
- using assms(2) eigenvector_def vn_def by blast
- have "vn \<in> carrier_vec n"
- using \<open>eigenvector U v k\<close> assms(2) eigenvector_def normalized_vec_dim vn_def by blast
- have "Complex_Matrix.inner_prod vn vn = 1" using \<open>vn = vec_normalize v\<close> \<open>eigenvector U v k\<close>
- eigenvector_def normalized_vec_norm by blast
- hence "conjugate k * k = conjugate k * k * Complex_Matrix.inner_prod vn vn" by simp
- also have "... = conjugate k * Complex_Matrix.inner_prod vn (k \<cdot>\<^sub>v vn)"
- proof -
- have "k * Complex_Matrix.inner_prod vn vn = Complex_Matrix.inner_prod vn (k \<cdot>\<^sub>v vn)"
- using inner_prod_smult_left[of vn n vn k] \<open>vn \<in> carrier_vec n\<close> by simp
- thus ?thesis by simp
- qed
- also have "... = Complex_Matrix.inner_prod (k \<cdot>\<^sub>v vn) (k \<cdot>\<^sub>v vn)"
- using inner_prod_smult_right[of vn n _ k] by (simp add: \<open>vn \<in> carrier_vec n\<close>)
- also have "... = Complex_Matrix.inner_prod (U *\<^sub>v vn) (U *\<^sub>v vn)"
- using \<open>eigenvector U vn k\<close> unfolding eigenvector_def by simp
- also have "... =
- Complex_Matrix.inner_prod (Complex_Matrix.adjoint U *\<^sub>v (U *\<^sub>v vn)) vn"
- using adjoint_def_alter[of "U *\<^sub>v vn" n vn n U] assms
- by (metis \<open>eigenvector U vn k\<close> carrier_matD(1) carrier_vec_dim_vec dim_mult_mat_vec
- eigenvector_def)
- also have "... = Complex_Matrix.inner_prod vn vn"
- proof -
- have "Complex_Matrix.adjoint U *\<^sub>v (U *\<^sub>v vn) = (Complex_Matrix.adjoint U * U) *\<^sub>v vn"
- using assms
- by (metis \<open>eigenvector U vn k\<close> adjoint_dim assoc_mult_mat_vec carrier_matD(1) eigenvector_def)
- also have "... = vn" using assms unfolding unitary_def inverts_mat_def
- by (metis \<open>eigenvector U vn k\<close> assms(1) eigenvector_def one_mult_mat_vec unitary_simps(1))
- finally show ?thesis by simp
- qed
- also have "... = 1" using \<open>vn = vec_normalize v\<close> \<open>eigenvector U v k\<close> eigenvector_def
- normalized_vec_norm by blast
- finally show ?thesis .
-qed
-
-lemma outer_prod_smult_left:
- fixes v::"complex Matrix.vec"
- shows "outer_prod (a \<cdot>\<^sub>v v) w = a \<cdot>\<^sub>m outer_prod v w"
-proof -
- define paw where "paw = outer_prod (a \<cdot>\<^sub>v v) w"
- define apw where "apw = a \<cdot>\<^sub>m outer_prod v w"
- have "paw = apw"
- proof
- have "dim_row paw = dim_vec v" unfolding paw_def using outer_prod_dim
- by (metis carrier_matD(1) carrier_vec_dim_vec index_smult_vec(2))
- also have "... = dim_row apw" unfolding apw_def using outer_prod_dim
- by (metis carrier_matD(1) carrier_vec_dim_vec index_smult_mat(2))
- finally show dr: "dim_row paw = dim_row apw" .
- have "dim_col paw = dim_vec w" unfolding paw_def using outer_prod_dim
- using carrier_vec_dim_vec by blast
- also have "... = dim_col apw" unfolding apw_def using outer_prod_dim
- by (metis apw_def carrier_matD(2) carrier_vec_dim_vec smult_carrier_mat)
- finally show dc: "dim_col paw = dim_col apw" .
- show "\<And>i j. i < dim_row apw \<Longrightarrow> j < dim_col apw \<Longrightarrow> paw $$ (i, j) = apw $$ (i, j)"
- proof -
- fix i j
- assume "i < dim_row apw" and "j < dim_col apw" note ij = this
- hence "paw $$ (i,j) = a * (Matrix.vec_index v i) * cnj (Matrix.vec_index w j)"
- using dr dc unfolding paw_def outer_prod_def by simp
- also have "... = apw $$ (i,j)" using dr dc ij unfolding apw_def outer_prod_def by simp
- finally show "paw $$ (i, j) = apw $$ (i, j)" .
- qed
- qed
- thus ?thesis unfolding paw_def apw_def by simp
-qed
-
-lemma outer_prod_smult_right:
- fixes v::"complex Matrix.vec"
- shows "outer_prod v (a \<cdot>\<^sub>v w) = (conjugate a) \<cdot>\<^sub>m outer_prod v w"
-proof -
- define paw where "paw = outer_prod v (a \<cdot>\<^sub>v w)"
- define apw where "apw = (conjugate a) \<cdot>\<^sub>m outer_prod v w"
- have "paw = apw"
- proof
- have "dim_row paw = dim_vec v" unfolding paw_def using outer_prod_dim
- by (metis carrier_matD(1) carrier_vec_dim_vec)
- also have "... = dim_row apw" unfolding apw_def using outer_prod_dim
- by (metis carrier_matD(1) carrier_vec_dim_vec index_smult_mat(2))
- finally show dr: "dim_row paw = dim_row apw" .
- have "dim_col paw = dim_vec w" unfolding paw_def using outer_prod_dim
- using carrier_vec_dim_vec by (metis carrier_matD(2) index_smult_vec(2))
- also have "... = dim_col apw" unfolding apw_def using outer_prod_dim
- by (metis apw_def carrier_matD(2) carrier_vec_dim_vec smult_carrier_mat)
- finally show dc: "dim_col paw = dim_col apw" .
- show "\<And>i j. i < dim_row apw \<Longrightarrow> j < dim_col apw \<Longrightarrow> paw $$ (i, j) = apw $$ (i, j)"
- proof -
- fix i j
- assume "i < dim_row apw" and "j < dim_col apw" note ij = this
- hence "paw $$ (i,j) = (conjugate a) * (Matrix.vec_index v i) * cnj (Matrix.vec_index w j)"
- using dr dc unfolding paw_def outer_prod_def by simp
- also have "... = apw $$ (i,j)" using dr dc ij unfolding apw_def outer_prod_def by simp
- finally show "paw $$ (i, j) = apw $$ (i, j)" .
- qed
- qed
- thus ?thesis unfolding paw_def apw_def by simp
-qed
-
-lemma outer_prod_add_left:
- fixes v::"complex Matrix.vec"
- assumes "dim_vec v = dim_vec x"
- shows "outer_prod (v + x) w = outer_prod v w + (outer_prod x w)"
-proof -
- define paw where "paw = outer_prod (v+x) w"
- define apw where "apw = outer_prod v w + (outer_prod x w)"
- have "paw = apw"
- proof
- have rv: "dim_row paw = dim_vec v" unfolding paw_def using outer_prod_dim assms
- by (metis carrier_matD(1) carrier_vec_dim_vec index_add_vec(2) paw_def)
- also have "... = dim_row apw" unfolding apw_def using outer_prod_dim assms
- by (metis carrier_matD(1) carrier_vec_dim_vec index_add_mat(2))
- finally show dr: "dim_row paw = dim_row apw" .
- have cw: "dim_col paw = dim_vec w" unfolding paw_def using outer_prod_dim assms
- using carrier_vec_dim_vec by (metis carrier_matD(2))
- also have "... = dim_col apw" unfolding apw_def using outer_prod_dim
- by (metis apw_def carrier_matD(2) carrier_vec_dim_vec add_carrier_mat)
- finally show dc: "dim_col paw = dim_col apw" .
- show "\<And>i j. i < dim_row apw \<Longrightarrow> j < dim_col apw \<Longrightarrow> paw $$ (i, j) = apw $$ (i, j)"
- proof -
- fix i j
- assume "i < dim_row apw" and "j < dim_col apw" note ij = this
- hence "paw $$ (i,j) = (Matrix.vec_index v i + Matrix.vec_index x i) *
- cnj (Matrix.vec_index w j)"
- using dr dc unfolding paw_def outer_prod_def by simp
- also have "... = Matrix.vec_index v i * cnj (Matrix.vec_index w j) +
- Matrix.vec_index x i * cnj (Matrix.vec_index w j)"
- by (simp add: ring_class.ring_distribs(2))
- also have "... = (outer_prod v w) $$ (i,j) + (outer_prod x w) $$ (i,j)"
- using rv cw dr dc ij assms unfolding outer_prod_def by auto
- also have "... = apw $$ (i,j)" using dr dc ij unfolding apw_def outer_prod_def by simp
- finally show "paw $$ (i, j) = apw $$ (i, j)" .
- qed
- qed
- thus ?thesis unfolding paw_def apw_def by simp
-qed
-
-lemma outer_prod_add_right:
- fixes v::"complex Matrix.vec"
- assumes "dim_vec w = dim_vec x"
- shows "outer_prod v (w + x) = outer_prod v w + (outer_prod v x)"
-proof -
- define paw where "paw = outer_prod v (w+x)"
- define apw where "apw = outer_prod v w + (outer_prod v x)"
- have "paw = apw"
- proof
- have rv: "dim_row paw = dim_vec v" unfolding paw_def using outer_prod_dim assms
- by (metis carrier_matD(1) carrier_vec_dim_vec index_add_vec(2) paw_def)
- also have "... = dim_row apw" unfolding apw_def using outer_prod_dim assms
- by (metis carrier_matD(1) carrier_vec_dim_vec index_add_mat(2))
- finally show dr: "dim_row paw = dim_row apw" .
- have cw: "dim_col paw = dim_vec w" unfolding paw_def using outer_prod_dim assms
- using carrier_vec_dim_vec
- by (metis carrier_matD(2) index_add_vec(2) paw_def)
- also have "... = dim_col apw" unfolding apw_def using outer_prod_dim
- by (metis assms carrier_matD(2) carrier_vec_dim_vec index_add_mat(3))
- finally show dc: "dim_col paw = dim_col apw" .
- show "\<And>i j. i < dim_row apw \<Longrightarrow> j < dim_col apw \<Longrightarrow> paw $$ (i, j) = apw $$ (i, j)"
- proof -
- fix i j
- assume "i < dim_row apw" and "j < dim_col apw" note ij = this
- hence "paw $$ (i,j) = Matrix.vec_index v i *
- (cnj (Matrix.vec_index w j + (Matrix.vec_index x j)))"
- using dr dc unfolding paw_def outer_prod_def by simp
- also have "... = Matrix.vec_index v i * cnj (Matrix.vec_index w j) +
- Matrix.vec_index v i * cnj (Matrix.vec_index x j)"
- by (simp add: ring_class.ring_distribs(1))
- also have "... = (outer_prod v w) $$ (i,j) + (outer_prod v x) $$ (i,j)"
- using rv cw dr dc ij assms unfolding outer_prod_def by auto
- also have "... = apw $$ (i,j)" using dr dc ij unfolding apw_def outer_prod_def by simp
- finally show "paw $$ (i, j) = apw $$ (i, j)" .
- qed
- qed
- thus ?thesis unfolding paw_def apw_def by simp
-qed
-
-lemma outer_prod_minus_left:
- fixes v::"complex Matrix.vec"
- assumes "dim_vec v = dim_vec x"
- shows "outer_prod (v - x) w = outer_prod v w - (outer_prod x w)"
-proof -
- define paw where "paw = outer_prod (v-x) w"
- define apw where "apw = outer_prod v w - (outer_prod x w)"
- have "paw = apw"
- proof
- have rv: "dim_row paw = dim_vec v" unfolding paw_def using outer_prod_dim assms
- by (metis carrier_matD(1) carrier_vec_dim_vec index_minus_vec(2) paw_def)
- also have "... = dim_row apw" unfolding apw_def using outer_prod_dim assms
- by (metis carrier_matD(1) carrier_vec_dim_vec index_minus_mat(2))
- finally show dr: "dim_row paw = dim_row apw" .
- have cw: "dim_col paw = dim_vec w" unfolding paw_def using outer_prod_dim assms
- using carrier_vec_dim_vec by (metis carrier_matD(2))
- also have "... = dim_col apw" unfolding apw_def using outer_prod_dim
- by (metis apw_def carrier_matD(2) carrier_vec_dim_vec minus_carrier_mat)
- finally show dc: "dim_col paw = dim_col apw" .
- show "\<And>i j. i < dim_row apw \<Longrightarrow> j < dim_col apw \<Longrightarrow> paw $$ (i, j) = apw $$ (i, j)"
- proof -
- fix i j
- assume "i < dim_row apw" and "j < dim_col apw" note ij = this
- hence "paw $$ (i,j) = (Matrix.vec_index v i - Matrix.vec_index x i) *
- cnj (Matrix.vec_index w j)"
- using dr dc unfolding paw_def outer_prod_def by simp
- also have "... = Matrix.vec_index v i * cnj (Matrix.vec_index w j) -
- Matrix.vec_index x i * cnj (Matrix.vec_index w j)"
- by (simp add: ring_class.ring_distribs)
- also have "... = (outer_prod v w) $$ (i,j) - (outer_prod x w) $$ (i,j)"
- using rv cw dr dc ij assms unfolding outer_prod_def by auto
- also have "... = apw $$ (i,j)" using dr dc ij unfolding apw_def outer_prod_def by simp
- finally show "paw $$ (i, j) = apw $$ (i, j)" .
- qed
- qed
- thus ?thesis unfolding paw_def apw_def by simp
-qed
-
-lemma outer_prod_minus_right:
- fixes v::"complex Matrix.vec"
- assumes "dim_vec w = dim_vec x"
- shows "outer_prod v (w - x) = outer_prod v w - (outer_prod v x)"
-proof -
- define paw where "paw = outer_prod v (w-x)"
- define apw where "apw = outer_prod v w - (outer_prod v x)"
- have "paw = apw"
- proof
- have rv: "dim_row paw = dim_vec v" unfolding paw_def using outer_prod_dim assms
- by (metis carrier_matD(1) carrier_vec_dim_vec paw_def)
- also have "... = dim_row apw" unfolding apw_def using outer_prod_dim assms
- by (metis carrier_matD(1) carrier_vec_dim_vec index_minus_mat(2))
- finally show dr: "dim_row paw = dim_row apw" .
- have cw: "dim_col paw = dim_vec w" unfolding paw_def using outer_prod_dim assms
- using carrier_vec_dim_vec
- by (metis carrier_matD(2) index_minus_vec(2) paw_def)
- also have "... = dim_col apw" unfolding apw_def using outer_prod_dim
- by (metis assms carrier_matD(2) carrier_vec_dim_vec index_minus_mat(3))
- finally show dc: "dim_col paw = dim_col apw" .
- show "\<And>i j. i < dim_row apw \<Longrightarrow> j < dim_col apw \<Longrightarrow> paw $$ (i, j) = apw $$ (i, j)"
- proof -
- fix i j
- assume "i < dim_row apw" and "j < dim_col apw" note ij = this
- hence "paw $$ (i,j) = Matrix.vec_index v i *
- (cnj (Matrix.vec_index w j - (Matrix.vec_index x j)))"
- using dr dc unfolding paw_def outer_prod_def by simp
- also have "... = Matrix.vec_index v i * cnj (Matrix.vec_index w j) -
- Matrix.vec_index v i * cnj (Matrix.vec_index x j)"
- by (simp add: ring_class.ring_distribs)
- also have "... = (outer_prod v w) $$ (i,j) - (outer_prod v x) $$ (i,j)"
- using rv cw dr dc ij assms unfolding outer_prod_def by auto
- also have "... = apw $$ (i,j)" using dr dc ij unfolding apw_def outer_prod_def by simp
- finally show "paw $$ (i, j) = apw $$ (i, j)" .
- qed
- qed
- thus ?thesis unfolding paw_def apw_def by simp
-qed
-
-lemma outer_minus_minus:
- fixes a::"complex Matrix.vec"
- assumes "dim_vec a = dim_vec b"
- and "dim_vec u = dim_vec v"
- shows "outer_prod (a - b) (u - v) = outer_prod a u - outer_prod a v -
- outer_prod b u + outer_prod b v"
-proof -
- have "outer_prod (a - b) (u - v) = outer_prod a (u - v)
- - outer_prod b (u - v)" using outer_prod_minus_left assms by simp
- also have "... = outer_prod a u - outer_prod a v -
- outer_prod b (u - v)" using assms outer_prod_minus_right by simp
- also have "... = outer_prod a u - outer_prod a v -
- (outer_prod b u - outer_prod b v)" using assms outer_prod_minus_right by simp
- also have "... = outer_prod a u - outer_prod a v -
- outer_prod b u + outer_prod b v"
- proof (rule mat_minus_minus)
- show "outer_prod b u \<in> carrier_mat (dim_vec b) (dim_vec u)" by simp
- show "outer_prod b v \<in> carrier_mat (dim_vec b) (dim_vec u)" using assms by simp
- show "outer_prod a u - outer_prod a v \<in> carrier_mat (dim_vec b) (dim_vec u)" using assms
- by (metis carrier_vecI minus_carrier_mat outer_prod_dim)
- qed
- finally show ?thesis .
-qed
-
-lemma outer_trace_inner:
- assumes "A \<in> carrier_mat n n"
- and "dim_vec u = n"
-and "dim_vec v = n"
- shows "Complex_Matrix.trace (outer_prod u v * A) = Complex_Matrix.inner_prod v (A *\<^sub>v u)"
-proof -
- have "Complex_Matrix.trace (outer_prod u v * A) = Complex_Matrix.trace (A * outer_prod u v)"
- proof (rule trace_comm)
- show "A \<in> carrier_mat n n" using assms by simp
- show "outer_prod u v \<in> carrier_mat n n" using assms
- by (metis carrier_vec_dim_vec outer_prod_dim)
- qed
- also have "... = Complex_Matrix.inner_prod v (A *\<^sub>v u)" using trace_outer_prod_right[of A n u v]
- assms carrier_vec_dim_vec by metis
- finally show ?thesis .
-qed
-
-lemma zero_hermitian:
- shows "hermitian (0\<^sub>m n n)" unfolding hermitian_def
- by (metis adjoint_minus hermitian_def hermitian_one minus_r_inv_mat one_carrier_mat)
-
-lemma trace_1:
- shows "Complex_Matrix.trace ((1\<^sub>m n)::complex Matrix.mat) =(n::complex)" using one_mat_def
- by (simp add: Complex_Matrix.trace_def Matrix.mat_def)
-
-lemma trace_add:
- assumes "square_mat A"
- and "square_mat B"
- and "dim_row A = dim_row B"
- shows "Complex_Matrix.trace (A + B) = Complex_Matrix.trace A + Complex_Matrix.trace B"
- using assms by (simp add: Complex_Matrix.trace_def sum.distrib)
-
-lemma bra_vec_carrier:
- shows "bra_vec v \<in> carrier_mat 1 (dim_vec v)"
-proof -
- have "dim_row (ket_vec v) = dim_vec v" unfolding ket_vec_def by simp
- thus ?thesis using bra_bra_vec[of v] bra_def[of "ket_vec v"] by simp
-qed
-
-lemma mat_mult_ket_carrier:
- assumes "A\<in> carrier_mat n m"
-shows "A * |v\<rangle> \<in> carrier_mat n 1" using assms
- by (metis bra_bra_vec bra_vec_carrier carrier_matD(1) carrier_matI dagger_of_ket_is_bra
- dim_row_of_dagger index_mult_mat(2) index_mult_mat(3))
-
-lemma mat_mult_ket:
- assumes "A \<in> carrier_mat n m"
-and "dim_vec v = m"
-shows "A * |v\<rangle> = |A *\<^sub>v v\<rangle>"
-proof -
- have rn: "dim_row (A * |v\<rangle>) = n" unfolding times_mat_def using assms by simp
- have co: "dim_col |A *\<^sub>v v\<rangle> = 1" using assms unfolding ket_vec_def by simp
- have cov: "dim_col |v\<rangle> = 1" using assms unfolding ket_vec_def by simp
- have er: "dim_row (A * |v\<rangle>) = dim_row |A *\<^sub>v v\<rangle>" using assms
- by (metis bra_bra_vec bra_vec_carrier carrier_matD(2) dagger_of_ket_is_bra dim_col_of_dagger
- dim_mult_mat_vec index_mult_mat(2))
- have ec: "dim_col (A * |v\<rangle>) = dim_col |A *\<^sub>v v\<rangle>" using assms
- by (metis carrier_matD(2) index_mult_mat(3) mat_mult_ket_carrier)
- {
- fix i::nat
- fix j::nat
- assume "i < n"
- and "j < 1"
- hence "j = 0" by simp
- have "(A * |v\<rangle>) $$ (i,0) = Matrix.scalar_prod (Matrix.row A i) (Matrix.col |v\<rangle> 0)"
- using times_mat_def[of A] \<open>i < n\<close> rn cov by simp
- also have "... = Matrix.scalar_prod (Matrix.row A i) v" using ket_vec_col by simp
- also have "... = |A *\<^sub>v v\<rangle> $$ (i,j)" unfolding mult_mat_vec_def
- using \<open>i < n\<close> \<open>j = 0\<close> assms(1) by auto
- } note idx = this
- have "A * |v\<rangle> = Matrix.mat n 1 (\<lambda>(i, j). Matrix.scalar_prod (Matrix.row A i) (Matrix.col |v\<rangle> j))"
- using assms unfolding times_mat_def ket_vec_def by simp
- also have "... = |A *\<^sub>v v\<rangle>" using er ec idx rn co by auto
- finally show ?thesis .
-qed
-
-lemma unitary_density:
- assumes "density_operator R"
- and "unitary U"
- and "R\<in> carrier_mat n n"
- and "U\<in> carrier_mat n n"
-shows "density_operator (U * R * (Complex_Matrix.adjoint U))" unfolding density_operator_def
-proof (intro conjI)
- show "Complex_Matrix.positive (U * R * Complex_Matrix.adjoint U)"
- proof (rule positive_close_under_left_right_mult_adjoint)
- show "U \<in> carrier_mat n n" using assms by simp
- show "R \<in> carrier_mat n n" using assms by simp
- show "Complex_Matrix.positive R" using assms unfolding density_operator_def by simp
- qed
- have "Complex_Matrix.trace (U * R * Complex_Matrix.adjoint U) =
- Complex_Matrix.trace (Complex_Matrix.adjoint U * U * R)"
- using trace_comm[of "U * R" n "Complex_Matrix.adjoint U"] assms
- by (metis adjoint_dim mat_assoc_test(10))
- also have "... = Complex_Matrix.trace R" using assms by simp
- also have "... = 1" using assms unfolding density_operator_def by simp
- finally show "Complex_Matrix.trace (U * R * Complex_Matrix.adjoint U) = 1" .
-qed
-
-
-subsection \<open>Tensor product complements\<close>
-
-lemma tensor_vec_dim[simp]:
- shows "dim_vec (tensor_vec u v) = dim_vec u * (dim_vec v)"
-proof -
- have "length (mult.vec_vec_Tensor (*) (list_of_vec u) (list_of_vec v)) =
- length (list_of_vec u) * length (list_of_vec v)"
- using mult.vec_vec_Tensor_length[of "1::real" "(*)" "list_of_vec u" "list_of_vec v"]
- by (simp add: Matrix_Tensor.mult_def)
- thus ?thesis unfolding tensor_vec_def by simp
-qed
-
-lemma index_tensor_vec[simp]:
- assumes "0 < dim_vec v"
- and "i < dim_vec u * dim_vec v"
-shows "vec_index (tensor_vec u v) i =
- vec_index u (i div (dim_vec v)) * vec_index v (i mod dim_vec v)"
-proof -
- have m: "Matrix_Tensor.mult (1::complex) (*)" by (simp add: Matrix_Tensor.mult_def)
- have "length (list_of_vec v) = dim_vec v" using assms by simp
- hence "vec_index (tensor_vec u v) i = (*) (vec_index u (i div dim_vec v)) (vec_index v (i mod dim_vec v))"
- unfolding tensor_vec_def using mult.vec_vec_Tensor_elements assms m
- by (metis (mono_tags, lifting) length_greater_0_conv length_list_of_vec list_of_vec_index
- mult.vec_vec_Tensor_elements vec_of_list_index)
- thus ?thesis by simp
-qed
-
-lemma outer_prod_tensor_comm:
- fixes a::"complex Matrix.vec"
- fixes u::"complex Matrix.vec"
- assumes "0 < dim_vec a"
- and "0 < dim_vec b"
-shows "outer_prod (tensor_vec u v) (tensor_vec a b) = tensor_mat (outer_prod u a) (outer_prod v b)"
-proof -
- define ot where "ot = outer_prod (tensor_vec u v) (tensor_vec a b)"
- define to where "to = tensor_mat (outer_prod u a) (outer_prod v b)"
- define dv where "dv = dim_vec v"
- define db where "db = dim_vec b"
- have "ot = to"
- proof
- have ro: "dim_row ot = dim_vec u * dim_vec v" unfolding ot_def outer_prod_def by simp
- have "dim_row to = dim_row (outer_prod u a) * dim_row (outer_prod v b)"
- unfolding to_def by simp
- also have "... = dim_vec u * dim_vec v" using outer_prod_dim
- by (metis carrier_matD(1) carrier_vec_dim_vec)
- finally have rt: "dim_row to = dim_vec u * dim_vec v" .
- show "dim_row ot = dim_row to" using ro rt by simp
- have co: "dim_col ot = dim_vec a * dim_vec b" unfolding ot_def outer_prod_def by simp
- have "dim_col to = dim_col (outer_prod u a) * dim_col (outer_prod v b)" unfolding to_def by simp
- also have "... = dim_vec a * dim_vec b" using outer_prod_dim
- by (metis carrier_matD(2) carrier_vec_dim_vec)
- finally have ct: "dim_col to = dim_vec a * dim_vec b" .
- show "dim_col ot = dim_col to" using co ct by simp
- show "\<And>i j. i < dim_row to \<Longrightarrow> j < dim_col to \<Longrightarrow> ot $$ (i, j) = to $$ (i, j)"
- proof -
- fix i j
- assume "i < dim_row to" and "j < dim_col to" note ij = this
- have "ot $$ (i,j) = Matrix.vec_index (tensor_vec u v) i *
- (conjugate (Matrix.vec_index (tensor_vec a b) j))"
- unfolding ot_def outer_prod_def using ij rt ct by simp
- also have "... = vec_index u (i div dv) * vec_index v (i mod dv) *
- (conjugate (Matrix.vec_index (tensor_vec a b) j))" using ij rt assms
- unfolding dv_def
- by (metis index_tensor_vec less_nat_zero_code nat_0_less_mult_iff neq0_conv)
- also have "... = vec_index u (i div dv) * vec_index v (i mod dv) *
- (conjugate (vec_index a (j div db) * vec_index b (j mod db)))" using ij ct assms
- unfolding db_def by simp
- also have "... = vec_index u (i div dv) * vec_index v (i mod dv) *
- (conjugate (vec_index a (j div db))) * (conjugate (vec_index b (j mod db)))" by simp
- also have "... = vec_index u (i div dv) * (conjugate (vec_index a (j div db))) *
- vec_index v (i mod dv) * (conjugate (vec_index b (j mod db)))" by simp
- also have "... = (outer_prod u a) $$ (i div dv, j div db) *
- vec_index v (i mod dv) * (conjugate (vec_index b (j mod db)))"
- proof -
- have "i div dv < dim_vec u" using ij rt unfolding dv_def
- by (simp add: less_mult_imp_div_less)
- moreover have "j div db < dim_vec a" using ij ct assms unfolding db_def
- by (simp add: less_mult_imp_div_less)
- ultimately have "vec_index u (i div dv) * (conjugate (vec_index a (j div db))) =
- (outer_prod u a) $$ (i div dv, j div db)" unfolding outer_prod_def by simp
- thus ?thesis by simp
- qed
- also have "... = (outer_prod u a) $$ (i div dv, j div db) *
- (outer_prod v b) $$ (i mod dv, j mod db)"
- proof -
- have "i mod dv < dim_vec v" using ij rt unfolding dv_def
- using assms mod_less_divisor
- by (metis less_nat_zero_code mult.commute neq0_conv times_nat.simps(1))
- moreover have "j mod db < dim_vec b" using ij ct assms unfolding db_def
- by (simp add: less_mult_imp_div_less)
- ultimately have "vec_index v (i mod dv) * (conjugate (vec_index b (j mod db))) =
- (outer_prod v b) $$ (i mod dv, j mod db)" unfolding outer_prod_def by simp
- thus ?thesis by simp
- qed
- also have "... = tensor_mat (outer_prod u a) (outer_prod v b) $$ (i, j)"
- proof (rule index_tensor_mat[symmetric])
- show "dim_row (outer_prod u a) = dim_vec u" unfolding outer_prod_def by simp
- show "dim_row (outer_prod v b) = dv" unfolding outer_prod_def dv_def by simp
- show "dim_col (outer_prod v b) = db" unfolding db_def outer_prod_def by simp
- show "i < dim_vec u * dv" unfolding dv_def using ij rt by simp
- show "dim_col (outer_prod u a) = dim_vec a" unfolding outer_prod_def by simp
- show "j < dim_vec a * db" unfolding db_def using ij ct by simp
- show "0 < dim_vec a" using assms by simp
- show "0 < db" unfolding db_def using assms by simp
- qed
- finally show "ot $$ (i, j) = to $$ (i, j)" unfolding to_def .
- qed
- qed
- thus ?thesis unfolding ot_def to_def by simp
-qed
-
-lemma tensor_mat_adjoint:
- assumes "m1 \<in> carrier_mat r1 c1"
- and "m2 \<in> carrier_mat r2 c2"
- and "0 < c1"
- and "0 < c2"
-and "0 < r1"
-and "0 < r2"
- shows "Complex_Matrix.adjoint (tensor_mat m1 m2) =
- tensor_mat (Complex_Matrix.adjoint m1) (Complex_Matrix.adjoint m2)"
- apply (rule eq_matI, auto)
-proof -
- fix i j
- assume "i < dim_col m1 * dim_col m2" and "j < dim_row m1 * dim_row m2" note ij = this
- have c1: "dim_col m1 = c1" using assms by simp
- have r1: "dim_row m1 = r1" using assms by simp
- have c2: "dim_col m2 = c2" using assms by simp
- have r2: "dim_row m2 = r2" using assms by simp
- have "Complex_Matrix.adjoint (m1 \<Otimes> m2) $$ (i, j) = conjugate ((m1 \<Otimes> m2) $$ (j, i))"
- using ij by (simp add: adjoint_eval)
- also have "... = conjugate (m1 $$ (j div r2, i div c2) * m2 $$ (j mod r2, i mod c2))"
- proof -
- have "(m1 \<Otimes> m2) $$ (j, i) = m1 $$ (j div r2, i div c2) * m2 $$ (j mod r2, i mod c2)"
- proof (rule index_tensor_mat[of m1 r1 c1 m2 r2 c2 j i], (auto simp add: assms ij c1 c2 r1 r2))
- show "j < r1 * r2" using ij r1 r2 by simp
- show "i < c1 * c2" using ij c1 c2 by simp
- qed
- thus ?thesis by simp
- qed
- also have "... = conjugate (m1 $$ (j div r2, i div c2)) * conjugate ( m2 $$ (j mod r2, i mod c2))"
- by simp
- also have "... = (Complex_Matrix.adjoint m1) $$ (i div c2, j div r2) *
- conjugate ( m2 $$ (j mod r2, i mod c2))"
- by (metis adjoint_eval c2 ij less_mult_imp_div_less r2)
- also have "... = (Complex_Matrix.adjoint m1) $$ (i div c2, j div r2) *
- (Complex_Matrix.adjoint m2) $$ (i mod c2, j mod r2)"
- by (metis Euclidean_Division.div_eq_0_iff adjoint_eval assms(4) bits_mod_div_trivial c2
- gr_implies_not_zero ij(2) mult_not_zero r2)
- also have "... = (tensor_mat (Complex_Matrix.adjoint m1) (Complex_Matrix.adjoint m2)) $$ (i,j)"
- proof (rule index_tensor_mat[symmetric], (simp add: ij c1 c2 r1 r2) +)
- show "i < c1 * c2" using ij c1 c2 by simp
- show "j < r1 * r2" using ij r1 r2 by simp
- show "0 < r1" using assms by simp
- show "0 < r2" using assms by simp
- qed
- finally show "Complex_Matrix.adjoint (m1 \<Otimes> m2) $$ (i, j) =
- (Complex_Matrix.adjoint m1 \<Otimes> Complex_Matrix.adjoint m2) $$ (i, j)" .
-qed
-
-lemma index_tensor_mat':
- assumes "0 < dim_col A"
- and "0 < dim_col B"
- and "i < dim_row A * dim_row B"
- and "j < dim_col A * dim_col B"
- shows "(A \<Otimes> B) $$ (i, j) =
- A $$ (i div (dim_row B), j div (dim_col B)) * B $$ (i mod (dim_row B), j mod (dim_col B))"
- by (rule index_tensor_mat, (simp add: assms)+)
-
-lemma tensor_mat_carrier:
- shows "tensor_mat U V \<in> carrier_mat (dim_row U * dim_row V) (dim_col U * dim_col V)" by auto
-
-lemma tensor_mat_id:
- assumes "0 < d1"
- and "0 < d2"
-shows "tensor_mat (1\<^sub>m d1) (1\<^sub>m d2) = 1\<^sub>m (d1 * d2)"
-proof (rule eq_matI, auto)
- show "tensor_mat (1\<^sub>m d1) (1\<^sub>m d2) $$ (i, i) = 1" if "i < (d1 * d2)" for i
- using that index_tensor_mat'[of "1\<^sub>m d1" "1\<^sub>m d2"]
- by (simp add: assms less_mult_imp_div_less)
-next
- show "tensor_mat (1\<^sub>m d1) (1\<^sub>m d2) $$ (i, j) = 0" if "i < d1 * d2" "j < d1 * d2" "i \<noteq> j" for i j
- using that index_tensor_mat[of "1\<^sub>m d1" d1 d1 "1\<^sub>m d2" d2 d2 i j]
- by (metis assms(1) assms(2) index_one_mat(1) index_one_mat(2) index_one_mat(3)
- less_mult_imp_div_less mod_less_divisor mult_div_mod_eq mult_not_zero)
-qed
-
-lemma tensor_mat_hermitian:
- assumes "A \<in> carrier_mat n n"
- and "B \<in> carrier_mat n' n'"
- and "0 < n"
- and "0 < n'"
- and "hermitian A"
- and "hermitian B"
- shows "hermitian (tensor_mat A B)" using assms by (metis hermitian_def tensor_mat_adjoint)
-
-lemma tensor_mat_unitary:
- assumes "Complex_Matrix.unitary U"
- and "Complex_Matrix.unitary V"
-and "0 < dim_row U"
-and "0 < dim_row V"
-shows "Complex_Matrix.unitary (tensor_mat U V)"
-proof -
- define UI where "UI = tensor_mat U V"
- have "Complex_Matrix.adjoint UI =
- tensor_mat (Complex_Matrix.adjoint U) (Complex_Matrix.adjoint V)" unfolding UI_def
- proof (rule tensor_mat_adjoint)
- show "U \<in> carrier_mat (dim_row U) (dim_row U)" using assms unfolding Complex_Matrix.unitary_def
- by simp
- show "V \<in> carrier_mat (dim_row V) (dim_row V)" using assms unfolding Complex_Matrix.unitary_def
- by simp
- show "0 < dim_row V" using assms by simp
- show "0 < dim_row U" using assms by simp
- show "0 < dim_row V" using assms by simp
- show "0 < dim_row U" using assms by simp
- qed
- hence "UI * (Complex_Matrix.adjoint UI) =
- tensor_mat (U * Complex_Matrix.adjoint U) (V * Complex_Matrix.adjoint V)"
- using mult_distr_tensor[of U "Complex_Matrix.adjoint U" "V" "Complex_Matrix.adjoint V"]
- unfolding UI_def
- by (metis (no_types, lifting) Complex_Matrix.unitary_def adjoint_dim_col adjoint_dim_row
- assms carrier_matD(2) )
- also have "... = tensor_mat (1\<^sub>m (dim_row U)) (1\<^sub>m (dim_row V))" using assms unitary_simps(2)
- by (metis Complex_Matrix.unitary_def)
- also have "... = (1\<^sub>m (dim_row U * dim_row V))" using tensor_mat_id assms by simp
- finally have "UI * (Complex_Matrix.adjoint UI) = (1\<^sub>m (dim_row U * dim_row V))" .
- hence "inverts_mat UI (Complex_Matrix.adjoint UI)" unfolding inverts_mat_def UI_def by simp
- thus ?thesis using assms unfolding Complex_Matrix.unitary_def UI_def
- by (metis carrier_matD(2) carrier_matI dim_col_tensor_mat dim_row_tensor_mat)
-qed
-
-
-subsection \<open>Fixed carrier matrices locale\<close>
-
-text \<open>We define a locale for matrices with a fixed number of rows and columns, and define a
-finite sum operation on this locale. The \verb+Type_To_Sets+ transfer tools then permits to obtain
-lemmata on the locale for free. \<close>
-
-locale fixed_carrier_mat =
- fixes fc_mats::"'a::field Matrix.mat set"
- fixes dimR dimC
- assumes fc_mats_carrier: "fc_mats = carrier_mat dimR dimC"
-begin
-
-sublocale semigroup_add_on_with fc_mats "(+)"
-proof (unfold_locales)
- show "\<And>a b. a \<in> fc_mats \<Longrightarrow> b \<in> fc_mats \<Longrightarrow> a + b \<in> fc_mats" using fc_mats_carrier by simp
- show "\<And>a b c. a \<in> fc_mats \<Longrightarrow> b \<in> fc_mats \<Longrightarrow> c \<in> fc_mats \<Longrightarrow> a + b + c = a + (b + c)"
- using fc_mats_carrier by simp
-qed
-
-sublocale ab_semigroup_add_on_with fc_mats "(+)"
-proof (unfold_locales)
- show "\<And>a b. a \<in> fc_mats \<Longrightarrow> b \<in> fc_mats \<Longrightarrow> a + b = b + a" using fc_mats_carrier
- by (simp add: comm_add_mat)
-qed
-
-sublocale comm_monoid_add_on_with fc_mats "(+)" "0\<^sub>m dimR dimC"
-proof (unfold_locales)
- show "0\<^sub>m dimR dimC \<in> fc_mats" using fc_mats_carrier by simp
- show "\<And>a. a \<in> fc_mats \<Longrightarrow> 0\<^sub>m dimR dimC + a = a" using fc_mats_carrier by simp
-qed
-
-sublocale ab_group_add_on_with fc_mats "(+)" "0\<^sub>m dimR dimC" "(-)" "uminus"
-proof (unfold_locales)
- show "\<And>a. a \<in> fc_mats \<Longrightarrow> - a + a = 0\<^sub>m dimR dimC" using fc_mats_carrier by simp
- show "\<And>a b. a \<in> fc_mats \<Longrightarrow> b \<in> fc_mats \<Longrightarrow> a - b = a + - b" using fc_mats_carrier
- by (simp add: add_uminus_minus_mat)
- show "\<And>a. a \<in> fc_mats \<Longrightarrow> - a \<in> fc_mats" using fc_mats_carrier by simp
-qed
-end
-
-lemma (in fixed_carrier_mat) smult_mem:
- assumes "A \<in> fc_mats"
- shows "a \<cdot>\<^sub>m A \<in> fc_mats" using fc_mats_carrier assms by auto
-
-definition (in fixed_carrier_mat) sum_mat where
-"sum_mat A I = sum_with (+) (0\<^sub>m dimR dimC) A I"
-
-lemma (in fixed_carrier_mat) sum_mat_empty[simp]:
- shows "sum_mat A {} = 0\<^sub>m dimR dimC" unfolding sum_mat_def by simp
-
-lemma (in fixed_carrier_mat) sum_mat_carrier:
- shows "(\<And>i. i \<in> I \<Longrightarrow> (A i)\<in> fc_mats) \<Longrightarrow> sum_mat A I \<in> carrier_mat dimR dimC"
- unfolding sum_mat_def using sum_with_mem[of A I] fc_mats_carrier by auto
-
-lemma (in fixed_carrier_mat) sum_mat_insert:
- assumes "A x \<in> fc_mats" "A ` I \<subseteq> fc_mats"
- and A: "finite I" and x: "x \<notin> I"
- shows "sum_mat A (insert x I) = A x + sum_mat A I" unfolding sum_mat_def
- using assms sum_with_insert[of A x I] by simp
-
-
-subsection \<open>A locale for square matrices\<close>
-
-locale cpx_sq_mat = fixed_carrier_mat "fc_mats::complex Matrix.mat set" for fc_mats +
- assumes dim_eq: "dimR = dimC"
- and npos: "0 < dimR"
-
-lemma (in cpx_sq_mat) one_mem:
- shows "1\<^sub>m dimR \<in> fc_mats" using fc_mats_carrier dim_eq by simp
-
-lemma (in cpx_sq_mat) square_mats:
- assumes "A \<in> fc_mats"
- shows "square_mat A" using fc_mats_carrier dim_eq assms by simp
-
-lemma (in cpx_sq_mat) cpx_sq_mat_mult:
- assumes "A \<in> fc_mats"
- and "B \<in> fc_mats"
-shows "A * B \<in> fc_mats"
-proof -
- have "dim_row (A * B) = dimR" using assms fc_mats_carrier by simp
- moreover have "dim_col (A * B) = dimR" using assms fc_mats_carrier dim_eq by simp
- ultimately show ?thesis using fc_mats_carrier carrier_mat_def dim_eq by auto
-qed
-
-lemma (in cpx_sq_mat) sum_mat_distrib_left:
- shows "finite I \<Longrightarrow> R\<in> fc_mats \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> (A i)\<in> fc_mats) \<Longrightarrow>
- sum_mat (\<lambda>i. R * (A i)) I = R * (sum_mat A I)"
-proof (induct rule: finite_induct)
- case empty
- hence a: "sum_mat (\<lambda>i. R * (A i)) {} = 0\<^sub>m dimR dimC" unfolding sum_mat_def by simp
- have "sum_mat A {} = 0\<^sub>m dimR dimC" unfolding sum_mat_def by simp
- hence "R * (sum_mat A {}) = 0\<^sub>m dimR dimC" using fc_mats_carrier
- right_mult_zero_mat[of R dimR dimC dimC] empty dim_eq by simp
- thus ?case using a by simp
-next
- case (insert x F)
- hence "sum_mat (\<lambda>i. R * A i) (insert x F) = R * (A x) + sum_mat (\<lambda>i. R * A i) F"
- using sum_mat_insert[of "\<lambda>i. R * A i" x F] by (simp add: image_subsetI fc_mats_carrier dim_eq)
- also have "... = R * (A x) + R * (sum_mat A F)" using insert by simp
- also have "... = R * (A x + (sum_mat A F))"
- by (metis dim_eq fc_mats_carrier insert.prems(1) insert.prems(2) insertCI mult_add_distrib_mat
- sum_mat_carrier)
- also have "... = R * sum_mat A (insert x F)"
- proof -
- have "A x + (sum_mat A F) = sum_mat A (insert x F)"
- by (rule sum_mat_insert[symmetric], (auto simp add: insert))
- thus ?thesis by simp
- qed
- finally show ?case .
-qed
-
-lemma (in cpx_sq_mat) sum_mat_distrib_right:
- shows "finite I \<Longrightarrow> R\<in> fc_mats \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> (A i)\<in> fc_mats) \<Longrightarrow>
- sum_mat (\<lambda>i. (A i) * R) I = (sum_mat A I) * R"
-proof (induct rule: finite_induct)
- case empty
- hence a: "sum_mat (\<lambda>i. (A i) * R) {} = 0\<^sub>m dimR dimC" unfolding sum_mat_def by simp
- have "sum_mat A {} = 0\<^sub>m dimR dimC" unfolding sum_mat_def by simp
- hence "(sum_mat A {}) * R = 0\<^sub>m dimR dimC" using fc_mats_carrier right_mult_zero_mat[of R ]
- dim_eq empty by simp
- thus ?case using a by simp
-next
- case (insert x F)
- have a: "(\<lambda>i. A i * R) ` F \<subseteq> fc_mats" using insert cpx_sq_mat_mult
- by (simp add: image_subsetI)
- have "A x * R \<in> fc_mats" using insert
- by (metis insertI1 local.fc_mats_carrier mult_carrier_mat dim_eq)
- hence "sum_mat (\<lambda>i. A i * R) (insert x F) = (A x) * R + sum_mat (\<lambda>i. A i * R) F" using insert a
- using sum_mat_insert[of "\<lambda>i. A i * R" x F] by (simp add: image_subsetI local.fc_mats_carrier)
- also have "... = (A x) * R + (sum_mat A F) * R" using insert by simp
- also have "... = (A x + (sum_mat A F)) * R"
- proof (rule add_mult_distrib_mat[symmetric])
- show "A x \<in> carrier_mat dimR dimC" using insert fc_mats_carrier by simp
- show "sum_mat A F \<in> carrier_mat dimR dimC" using insert fc_mats_carrier sum_mat_carrier by blast
- show "R \<in> carrier_mat dimC dimC" using insert fc_mats_carrier dim_eq by simp
- qed
- also have "... = sum_mat A (insert x F) * R"
- proof -
- have "A x + (sum_mat A F) = sum_mat A (insert x F)"
- by (rule sum_mat_insert[symmetric], (auto simp add: insert))
- thus ?thesis by simp
- qed
- finally show ?case .
-qed
-
-lemma (in cpx_sq_mat) trace_sum_mat:
- fixes A::"'b \<Rightarrow> complex Matrix.mat"
- shows "finite I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> (A i)\<in> fc_mats) \<Longrightarrow>
- Complex_Matrix.trace (sum_mat A I) = (\<Sum> i\<in> I. Complex_Matrix.trace (A i))" unfolding sum_mat_def
-proof (induct rule: finite_induct)
- case empty
- then show ?case using trace_zero dim_eq by simp
-next
- case (insert x F)
- have "Complex_Matrix.trace (sum_with (+) (0\<^sub>m dimR dimC) A (insert x F)) =
- Complex_Matrix.trace (A x + sum_with (+) (0\<^sub>m dimR dimC) A F)"
- using sum_with_insert[of A x F] insert by (simp add: image_subset_iff dim_eq)
- also have "... = Complex_Matrix.trace (A x) +
- Complex_Matrix.trace (sum_with (+) (0\<^sub>m dimR dimC) A F)" using trace_add square_mats insert
- by (metis carrier_matD(1) fc_mats_carrier image_subset_iff insert_iff sum_with_mem)
- also have "... = Complex_Matrix.trace (A x) + (\<Sum> i\<in> F. Complex_Matrix.trace (A i))"
- using insert by simp
- also have "... = (\<Sum> i\<in> (insert x F). Complex_Matrix.trace (A i))"
- using sum_with_insert[of A x F] insert by (simp add: image_subset_iff)
- finally show ?case .
-qed
-
-lemma (in cpx_sq_mat) cpx_sq_mat_smult:
- assumes "A \<in> fc_mats"
- shows "x \<cdot>\<^sub>m A \<in> fc_mats"
- using assms fc_mats_carrier by auto
-
-lemma (in cpx_sq_mat) mult_add_distrib_right:
- assumes "A\<in> fc_mats" "B\<in> fc_mats" "C\<in> fc_mats"
- shows "A * (B + C) = A * B + A * C"
- using assms fc_mats_carrier mult_add_distrib_mat dim_eq by simp
-
-lemma (in cpx_sq_mat) mult_add_distrib_left:
- assumes "A\<in> fc_mats" "B\<in> fc_mats" "C\<in> fc_mats"
- shows "(B + C) * A = B * A + C * A"
- using assms fc_mats_carrier add_mult_distrib_mat dim_eq by simp
-
-lemma (in cpx_sq_mat) mult_sum_mat_distrib_left:
- shows "finite I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> (A i)\<in> fc_mats) \<Longrightarrow> B \<in> fc_mats \<Longrightarrow>
- (sum_mat (\<lambda>i. B * (A i)) I) = B * (sum_mat A I)"
-proof (induct rule: finite_induct)
- case empty
- hence "sum_mat A {} = 0\<^sub>m dimR dimC" using sum_mat_empty by simp
- hence "B * (sum_mat A {}) = 0\<^sub>m dimR dimC" using empty by (simp add: fc_mats_carrier dim_eq)
- moreover have "sum_mat (\<lambda>i. B * (A i)) {} = 0\<^sub>m dimR dimC" using sum_mat_empty[of "\<lambda>i. B * (A i)"]
- by simp
- ultimately show ?case by simp
-next
- case (insert x F)
- have "sum_mat (\<lambda>i. B * (A i)) (insert x F) = B * (A x) + sum_mat (\<lambda>i. B * (A i)) F"
- using sum_with_insert[of "\<lambda>i. B * (A i)" x F] insert
- by (simp add: image_subset_iff local.sum_mat_def cpx_sq_mat_mult)
- also have "... = B * (A x) + B * (sum_mat A F)" using insert by simp
- also have "... = B * (A x + (sum_mat A F))"
- proof (rule mult_add_distrib_right[symmetric])
- show "B\<in> fc_mats" using insert by simp
- show "A x \<in> fc_mats" using insert by simp
- show "sum_mat A F \<in> fc_mats" using insert by (simp add: fc_mats_carrier sum_mat_carrier)
- qed
- also have "... = B * (sum_mat A (insert x F))" using sum_with_insert[of A x F] insert
- by (simp add: image_subset_iff sum_mat_def)
- finally show ?case .
-qed
-
-lemma (in cpx_sq_mat) mult_sum_mat_distrib_right:
- shows "finite I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> (A i)\<in> fc_mats) \<Longrightarrow> B \<in> fc_mats \<Longrightarrow>
- (sum_mat (\<lambda>i. (A i) * B) I) = (sum_mat A I) * B"
-proof (induct rule: finite_induct)
- case empty
- hence "sum_mat A {} = 0\<^sub>m dimR dimC" using sum_mat_empty by simp
- hence "(sum_mat A {}) * B = 0\<^sub>m dimR dimC" using empty by (simp add: fc_mats_carrier dim_eq)
- moreover have "sum_mat (\<lambda>i. (A i) * B) {} = 0\<^sub>m dimR dimC" by simp
- ultimately show ?case by simp
-next
- case (insert x F)
- have "sum_mat (\<lambda>i. (A i) * B) (insert x F) = (A x) * B + sum_mat (\<lambda>i. (A i) * B) F"
- using sum_with_insert[of "\<lambda>i. (A i) * B" x F] insert
- by (simp add: image_subset_iff local.sum_mat_def cpx_sq_mat_mult)
- also have "... = (A x) * B + (sum_mat A F) * B" using insert by simp
- also have "... = (A x + (sum_mat A F)) * B"
- proof (rule mult_add_distrib_left[symmetric])
- show "B\<in> fc_mats" using insert by simp
- show "A x \<in> fc_mats" using insert by simp
- show "sum_mat A F \<in> fc_mats" using insert by (simp add: fc_mats_carrier sum_mat_carrier)
- qed
- also have "... = (sum_mat A (insert x F)) * B" using sum_with_insert[of A x F] insert
- by (simp add: image_subset_iff sum_mat_def)
- finally show ?case .
-qed
-
-lemma (in cpx_sq_mat) trace_sum_mat_mat_distrib:
- assumes "finite I"
-and "\<And>i. i\<in> I \<Longrightarrow> B i \<in> fc_mats"
-and "A\<in> fc_mats"
-and "C \<in> fc_mats"
-shows "(\<Sum> i\<in> I. Complex_Matrix.trace(A * (B i) * C)) =
- Complex_Matrix.trace (A * (sum_mat B I) * C)"
-proof -
- have seq: "sum_mat (\<lambda>i. A * (B i) * C) I = A * (sum_mat B I) * C"
- proof -
- have "sum_mat (\<lambda>i. A * (B i) * C) I = (sum_mat (\<lambda>i. A * (B i)) I) * C"
- proof (rule mult_sum_mat_distrib_right)
- show "finite I" using assms by simp
- show "C\<in> fc_mats" using assms by simp
- show "\<And>i. i \<in> I \<Longrightarrow> A * B i \<in> fc_mats" using assms cpx_sq_mat_mult by simp
- qed
- moreover have "sum_mat (\<lambda>i. A * (B i)) I = A * (sum_mat B I)"
- by (rule mult_sum_mat_distrib_left, (auto simp add: assms))
- ultimately show "sum_mat (\<lambda>i. A * (B i) * C) I = A * (sum_mat B I) * C" by simp
- qed
- have "(\<Sum> i\<in> I. Complex_Matrix.trace(A * (B i) * C)) =
- Complex_Matrix.trace (sum_mat (\<lambda>i. A * (B i) * C) I)"
- proof (rule trace_sum_mat[symmetric])
- show "finite I" using assms by simp
- fix i
- assume "i\<in> I"
- thus "A * B i * C \<in> fc_mats" using assms by (simp add: cpx_sq_mat_mult)
- qed
- also have "... = Complex_Matrix.trace (A * (sum_mat B I) * C)" using seq by simp
- finally show ?thesis .
-qed
-
-definition (in cpx_sq_mat) zero_col where
-"zero_col U = (\<lambda>i. if i < dimR then Matrix.col U i else 0\<^sub>v dimR)"
-
-lemma (in cpx_sq_mat) zero_col_dim:
- assumes "U \<in> fc_mats"
- shows "dim_vec (zero_col U i) = dimR" using assms fc_mats_carrier unfolding zero_col_def by simp
-
-lemma (in cpx_sq_mat) zero_col_col:
- assumes "i < dimR"
- shows "zero_col U i = Matrix.col U i" using assms unfolding zero_col_def by simp
-
-lemma (in cpx_sq_mat) sum_mat_index:
- shows "finite I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> (A i)\<in> fc_mats) \<Longrightarrow> i < dimR \<Longrightarrow> j < dimC \<Longrightarrow>
- (sum_mat (\<lambda>k. (A k)) I) $$ (i,j) = (\<Sum> k\<in>I. (A k) $$ (i,j))"
-proof (induct rule: finite_induct)
- case empty
- thus ?case unfolding sum_mat_def by simp
-next
- case (insert x F)
- hence "(sum_mat (\<lambda>k. (A k)) (insert x F)) $$ (i,j) =
- (A x + (sum_mat (\<lambda>k. (A k)) F)) $$ (i,j)" using insert sum_mat_insert[of A]
- by (simp add: image_subsetI local.fc_mats_carrier)
- also have "... = (A x) $$ (i,j) + (sum_mat (\<lambda>k. (A k)) F) $$ (i,j)" using insert
- sum_mat_carrier[of F A] fc_mats_carrier by simp
- also have "... = (A x) $$ (i,j) + (\<Sum> k\<in>F. (A k) $$ (i,j))" using insert by simp
- also have "... = (\<Sum> k\<in>(insert x F). (A k) $$ (i,j))" using insert by simp
- finally show ?case .
-qed
-
-lemma (in cpx_sq_mat) sum_mat_cong:
- shows "finite I \<Longrightarrow> (\<And>i. i\<in> I \<Longrightarrow> A i = B i) \<Longrightarrow> (\<And>i. i\<in> I \<Longrightarrow> A i \<in> fc_mats) \<Longrightarrow>
- (\<And>i. i\<in> I \<Longrightarrow> B i \<in> fc_mats) \<Longrightarrow> sum_mat A I = sum_mat B I"
-proof (induct rule: finite_induct)
- case empty
- then show ?case by simp
-next
- case (insert x F)
- have "sum_mat A (insert x F) = A x + sum_mat A F" using insert sum_mat_insert[of A]
- by (simp add: image_subset_iff)
- also have "... = B x + sum_mat B F" using insert by simp
- also have "... = sum_mat B (insert x F)" using insert sum_mat_insert[of B]
- by (simp add: image_subset_iff)
- finally show ?case .
-qed
-
-lemma (in cpx_sq_mat) smult_sum_mat:
- shows "finite I \<Longrightarrow> (\<And>i. i\<in> I \<Longrightarrow> A i \<in> fc_mats) \<Longrightarrow> a \<cdot>\<^sub>m sum_mat A I = sum_mat (\<lambda>i. a \<cdot>\<^sub>m (A i)) I"
-proof (induct rule: finite_induct)
- case empty
- then show ?case by simp
-next
- case (insert x F)
- have "a \<cdot>\<^sub>m sum_mat A (insert x F) = a \<cdot>\<^sub>m (A x + sum_mat A F)" using insert sum_mat_insert[of A]
- by (simp add: image_subset_iff)
- also have "... = a \<cdot>\<^sub>m A x + a \<cdot>\<^sub>m (sum_mat A F)" using insert
- by (metis add_smult_distrib_left_mat fc_mats_carrier insert_iff sum_mat_carrier)
- also have "... = a \<cdot>\<^sub>m A x + sum_mat (\<lambda>i. a \<cdot>\<^sub>m (A i)) F" using insert by simp
- also have "... = sum_mat (\<lambda>i. a \<cdot>\<^sub>m (A i)) (insert x F)" using insert
- sum_mat_insert[of "(\<lambda>i. a \<cdot>\<^sub>m (A i))"] by (simp add: image_subset_iff cpx_sq_mat_smult)
- finally show ?case .
-qed
-
-lemma (in cpx_sq_mat) zero_sum_mat:
- shows "finite I \<Longrightarrow> sum_mat (\<lambda>i. ((0\<^sub>m dimR dimR)::complex Matrix.mat)) I = ((0\<^sub>m dimR dimR)::complex Matrix.mat)"
-proof (induct rule:finite_induct)
- case empty
- then show ?case using dim_eq sum_mat_empty by auto
-next
- case (insert x F)
- have "sum_mat (\<lambda>i. ((0\<^sub>m dimR dimR)::complex Matrix.mat)) (insert x F) =
- 0\<^sub>m dimR dimR + sum_mat (\<lambda>i. 0\<^sub>m dimR dimR) F"
- using insert dim_eq zero_mem sum_mat_insert[of "\<lambda>i. ((0\<^sub>m dimR dimR)::complex Matrix.mat)"]
- by fastforce
- also have "... = ((0\<^sub>m dimR dimR)::complex Matrix.mat)" using insert by auto
- finally show ?case .
-qed
-
-lemma (in cpx_sq_mat) sum_mat_adjoint:
- shows "finite I \<Longrightarrow> (\<And>i. i\<in> I \<Longrightarrow> A i \<in> fc_mats) \<Longrightarrow>
- Complex_Matrix.adjoint (sum_mat A I) = sum_mat (\<lambda>i. Complex_Matrix.adjoint (A i)) I"
-proof (induct rule: finite_induct)
- case empty
- then show ?case using zero_hermitian[of dimR]
- by (metis (no_types) dim_eq hermitian_def sum_mat_empty)
-next
- case (insert x F)
- have "Complex_Matrix.adjoint (sum_mat A (insert x F)) =
- Complex_Matrix.adjoint (A x + sum_mat A F)" using insert sum_mat_insert[of A]
- by (simp add: image_subset_iff)
- also have "... = Complex_Matrix.adjoint (A x) + Complex_Matrix.adjoint (sum_mat A F)"
- proof (rule adjoint_add)
- show "A x \<in> carrier_mat dimR dimC" using insert fc_mats_carrier by simp
- show "sum_mat A F \<in> carrier_mat dimR dimC" using insert fc_mats_carrier sum_mat_carrier[of F]
- by simp
- qed
- also have "... = Complex_Matrix.adjoint (A x) + sum_mat (\<lambda>i. Complex_Matrix.adjoint (A i)) F"
- using insert by simp
- also have "... = sum_mat (\<lambda>i. Complex_Matrix.adjoint (A i)) (insert x F)"
- proof (rule sum_mat_insert[symmetric], (auto simp add: insert))
- show "Complex_Matrix.adjoint (A x) \<in> fc_mats" using insert fc_mats_carrier dim_eq
- by (simp add: adjoint_dim)
- show "\<And>i. i \<in> F \<Longrightarrow> Complex_Matrix.adjoint (A i) \<in> fc_mats" using insert fc_mats_carrier dim_eq
- by (simp add: adjoint_dim)
- qed
- finally show ?case .
-qed
-
-lemma (in cpx_sq_mat) sum_mat_hermitian:
- assumes "finite I"
-and "\<forall>i\<in> I. hermitian (A i)"
-and "\<forall>i\<in> I. A i\<in> fc_mats"
-shows "hermitian (sum_mat A I)"
-proof -
- have "Complex_Matrix.adjoint (sum_mat A I) = sum_mat (\<lambda>i. Complex_Matrix.adjoint (A i)) I"
- using assms sum_mat_adjoint[of I] by simp
- also have "... = sum_mat A I"
- proof (rule sum_mat_cong, (auto simp add: assms))
- show "\<And>i. i \<in> I \<Longrightarrow> Complex_Matrix.adjoint (A i) = A i" using assms
- unfolding hermitian_def by simp
- show "\<And>i. i \<in> I \<Longrightarrow> Complex_Matrix.adjoint (A i) \<in> fc_mats" using assms fc_mats_carrier dim_eq
- by (simp add: adjoint_dim)
- qed
- finally show ?thesis unfolding hermitian_def .
-qed
-
-lemma (in cpx_sq_mat) sum_mat_positive:
-shows "finite I \<Longrightarrow> (\<And>i. i\<in> I \<Longrightarrow> Complex_Matrix.positive (A i)) \<Longrightarrow>
- (\<And>i. i \<in> I \<Longrightarrow> A i \<in> fc_mats) \<Longrightarrow> Complex_Matrix.positive (sum_mat A I)"
-proof (induct rule: finite_induct)
- case empty
- then show ?case using positive_zero[of dimR] by (metis (no_types) dim_eq sum_mat_empty)
-next
- case (insert x F)
- hence "sum_mat A (insert x F) = A x + (sum_mat A F)" using sum_mat_insert[of A]
- by (simp add: image_subset_iff)
- moreover have "Complex_Matrix.positive (A x + (sum_mat A F))"
- proof (rule positive_add, (auto simp add: insert))
- show "A x \<in> carrier_mat dimR dimR" using insert fc_mats_carrier dim_eq by simp
- show "sum_mat A F \<in> carrier_mat dimR dimR" using insert sum_mat_carrier dim_eq
- by (metis insertCI)
- qed
- ultimately show "Complex_Matrix.positive (sum_mat A (insert x F))" by simp
-qed
-
-lemma (in cpx_sq_mat) sum_mat_left_ortho_zero:
- shows "finite I \<Longrightarrow>
- (\<And>i. i\<in> I \<Longrightarrow> A i \<in> fc_mats) \<Longrightarrow> (B \<in> fc_mats) \<Longrightarrow>
- (\<And> i. i\<in> I \<Longrightarrow> A i * B = (0\<^sub>m dimR dimR)) \<Longrightarrow>
- (sum_mat A I) * B = 0\<^sub>m dimR dimR"
-proof (induct rule:finite_induct)
- case empty
- then show ?case using dim_eq
- by (metis finite.intros(1) sum_mat_empty mult_sum_mat_distrib_right)
-next
- case (insert x F)
- have "(sum_mat A (insert x F)) * B =
- (A x + sum_mat A F) * B" using insert sum_mat_insert[of A]
- by (simp add: image_subset_iff)
- also have "... = A x * B + sum_mat A F * B"
- proof (rule add_mult_distrib_mat)
- show "A x \<in> carrier_mat dimR dimC" using insert fc_mats_carrier by simp
- show "sum_mat A F \<in> carrier_mat dimR dimC" using insert
- by (metis insert_iff local.fc_mats_carrier sum_mat_carrier)
- show "B \<in> carrier_mat dimC dimR" using insert fc_mats_carrier dim_eq by simp
- qed
- also have "... = A x * B + (0\<^sub>m dimR dimR)" using insert by simp
- also have "... = 0\<^sub>m dimR dimR" using insert by simp
- finally show ?case .
-qed
-
-lemma (in cpx_sq_mat) sum_mat_right_ortho_zero:
- shows "finite I \<Longrightarrow>
- (\<And>i. i\<in> I \<Longrightarrow> A i \<in> fc_mats) \<Longrightarrow> (B \<in> fc_mats) \<Longrightarrow>
- (\<And> i. i\<in> I \<Longrightarrow> B * A i = (0\<^sub>m dimR dimR)) \<Longrightarrow>
- B * (sum_mat A I) = 0\<^sub>m dimR dimR"
-proof (induct rule:finite_induct)
- case empty
- then show ?case using dim_eq
- by (metis finite.intros(1) sum_mat_empty mult_sum_mat_distrib_left)
-next
- case (insert x F)
- have "B * (sum_mat A (insert x F)) =
- B * (A x + sum_mat A F)" using insert sum_mat_insert[of A]
- by (simp add: image_subset_iff)
- also have "... = B * A x + B * sum_mat A F"
- proof (rule mult_add_distrib_mat)
- show "A x \<in> carrier_mat dimR dimC" using insert fc_mats_carrier by simp
- show "sum_mat A F \<in> carrier_mat dimR dimC" using insert
- by (metis insert_iff local.fc_mats_carrier sum_mat_carrier)
- show "B \<in> carrier_mat dimC dimR" using insert fc_mats_carrier dim_eq by simp
- qed
- also have "... = B * A x + (0\<^sub>m dimR dimR)" using insert by simp
- also have "... = 0\<^sub>m dimR dimR" using insert by simp
- finally show ?case .
-qed
-
-lemma (in cpx_sq_mat) sum_mat_ortho_square:
- shows "finite I \<Longrightarrow> (\<And>i. i\<in> I \<Longrightarrow> ((A i)::complex Matrix.mat) * (A i) = A i) \<Longrightarrow>
- (\<And>i. i\<in> I \<Longrightarrow> A i \<in> fc_mats) \<Longrightarrow>
- (\<And> i j. i\<in> I \<Longrightarrow> j\<in> I \<Longrightarrow> i\<noteq> j \<Longrightarrow> A i * (A j) = (0\<^sub>m dimR dimR)) \<Longrightarrow>
- (sum_mat A I) * (sum_mat A I) = (sum_mat A I)"
-proof (induct rule:finite_induct)
- case empty
- then show ?case using dim_eq
- by (metis fc_mats_carrier right_mult_zero_mat sum_mat_empty zero_mem)
-next
- case (insert x F)
- have "(sum_mat A (insert x F)) * (sum_mat A (insert x F)) =
- (A x + sum_mat A F) * (A x + sum_mat A F)" using insert sum_mat_insert[of A]
- by (simp add: \<open>\<And>i. i \<in> insert x F \<Longrightarrow> A i * A i = A i\<close> image_subset_iff)
- also have "... = A x * (A x + sum_mat A F) + sum_mat A F * (A x + sum_mat A F)"
- proof (rule add_mult_distrib_mat)
- show "A x \<in> carrier_mat dimR dimC" using insert fc_mats_carrier by simp
- show "sum_mat A F \<in> carrier_mat dimR dimC" using insert
- by (metis insert_iff local.fc_mats_carrier sum_mat_carrier)
- thus "A x + sum_mat A F \<in> carrier_mat dimC dimC" using insert dim_eq by simp
- qed
- also have "... = A x * A x + A x * (sum_mat A F) + sum_mat A F * (A x + sum_mat A F)"
- proof -
- have "A x * (A x + sum_mat A F) = A x * A x + A x * (sum_mat A F)"
- using dim_eq insert.prems(2) mult_add_distrib_right sum_mat_carrier
- by (metis fc_mats_carrier insertI1 subsetD subset_insertI)
- thus ?thesis by simp
- qed
- also have "... = A x * A x + A x * (sum_mat A F) + sum_mat A F * A x +
- sum_mat A F * (sum_mat A F)"
- proof -
- have "sum_mat A F * (A x + local.sum_mat A F) =
- sum_mat A F * A x + local.sum_mat A F * local.sum_mat A F"
- using insert dim_eq add_assoc add_mem mult_add_distrib_right cpx_sq_mat_mult sum_mat_carrier
- by (metis fc_mats_carrier insertI1 subsetD subset_insertI)
- hence "A x * A x + A x * sum_mat A F + sum_mat A F * (A x + sum_mat A F) =
- A x * A x + A x * sum_mat A F + (sum_mat A F * A x + sum_mat A F * sum_mat A F)" by simp
- also have "... = A x * A x + A x * sum_mat A F + sum_mat A F * A x + sum_mat A F * sum_mat A F"
- proof (rule assoc_add_mat[symmetric])
- show "A x * A x + A x * sum_mat A F \<in> carrier_mat dimR dimR" using sum_mat_carrier insert
- dim_eq fc_mats_carrier by (metis add_mem cpx_sq_mat_mult insertCI)
- show "sum_mat A F * A x \<in> carrier_mat dimR dimR" using sum_mat_carrier insert
- dim_eq fc_mats_carrier by (metis cpx_sq_mat_mult insertCI)
- show "sum_mat A F * sum_mat A F \<in> carrier_mat dimR dimR" using sum_mat_carrier insert
- dim_eq fc_mats_carrier by (metis cpx_sq_mat_mult insertCI)
- qed
- finally show ?thesis .
- qed
- also have "... = A x + sum_mat A F"
- proof -
- have "A x * A x = A x" using insert by simp
- moreover have "sum_mat A F * sum_mat A F = sum_mat A F" using insert by simp
- moreover have "A x * (sum_mat A F) = 0\<^sub>m dimR dimR"
- proof -
- have "A x * (sum_mat A F) = sum_mat (\<lambda>i. A x * (A i)) F"
- by (rule sum_mat_distrib_left[symmetric], (simp add: insert)+)
- also have "... = sum_mat (\<lambda>i. 0\<^sub>m dimR dimR) F"
- proof (rule sum_mat_cong, (auto simp add: insert zero_mem))
- show "\<And>i. i \<in> F \<Longrightarrow> A x * A i = 0\<^sub>m dimR dimR" using insert by auto
- show "\<And>i. i \<in> F \<Longrightarrow> A x * A i \<in> fc_mats" using insert cpx_sq_mat_mult by auto
- show "\<And>i. i \<in> F \<Longrightarrow> 0\<^sub>m dimR dimR \<in> fc_mats" using zero_mem dim_eq by simp
- qed
- also have "... = 0\<^sub>m dimR dimR" using zero_sum_mat insert by simp
- finally show ?thesis .
- qed
- moreover have "sum_mat A F * A x = 0\<^sub>m dimR dimR"
- proof -
- have "sum_mat A F * A x = sum_mat (\<lambda>i. A i * (A x)) F"
- by (rule sum_mat_distrib_right[symmetric], (simp add: insert)+)
- also have "... = sum_mat (\<lambda>i. 0\<^sub>m dimR dimR) F"
- proof (rule sum_mat_cong, (auto simp add: insert zero_mem))
- show "\<And>i. i \<in> F \<Longrightarrow> A i * A x = 0\<^sub>m dimR dimR" using insert by auto
- show "\<And>i. i \<in> F \<Longrightarrow> A i * A x \<in> fc_mats" using insert cpx_sq_mat_mult by auto
- show "\<And>i. i \<in> F \<Longrightarrow> 0\<^sub>m dimR dimR \<in> fc_mats" using zero_mem dim_eq by simp
- qed
- also have "... = 0\<^sub>m dimR dimR" using zero_sum_mat insert by simp
- finally show ?thesis .
- qed
- ultimately show ?thesis using add_commute add_zero insert.prems(2) zero_mem dim_eq by auto
- qed
- also have "... = sum_mat A (insert x F)" using insert sum_mat_insert[of A x F]
- by (simp add: \<open>\<And>i. i \<in> insert x F \<Longrightarrow> A i * A i = A i\<close> image_subsetI)
- finally show ?case .
-qed
-
-lemma diagonal_unit_vec:
- assumes "B \<in> carrier_mat n n"
-and "diagonal_mat (B::complex Matrix.mat)"
-shows "B *\<^sub>v (unit_vec n i) = B $$ (i,i) \<cdot>\<^sub>v (unit_vec n i)"
-proof -
- define v::"complex Matrix.vec" where "v = unit_vec n i"
- have "B *\<^sub>v v = Matrix.vec n (\<lambda> i. Matrix.scalar_prod (Matrix.row B i) v)"
- using assms unfolding mult_mat_vec_def by simp
- also have "... = Matrix.vec n (\<lambda> i. B $$(i,i) * Matrix.vec_index v i)"
- proof -
- have "\<forall>i < n. (Matrix.scalar_prod (Matrix.row B i) v = B $$(i,i) * Matrix.vec_index v i)"
- proof (intro allI impI)
- fix i
- assume "i < n"
- have "(Matrix.scalar_prod (Matrix.row B i) v) =
- (\<Sum> j \<in> {0 ..< n}. Matrix.vec_index (Matrix.row B i) j * Matrix.vec_index v j)" using assms
- unfolding Matrix.scalar_prod_def v_def by simp
- also have "... = Matrix.vec_index (Matrix.row B i) i * Matrix.vec_index v i"
- proof (rule sum_but_one)
- show "\<forall>j < n. j \<noteq> i \<longrightarrow> Matrix.vec_index (Matrix.row B i) j = 0"
- proof (intro allI impI)
- fix j
- assume "j < n" and "j \<noteq> i"
- hence "Matrix.vec_index (Matrix.row B i) j = B $$ (i,j)" using \<open>i < n\<close> \<open>j < n\<close> assms
- by auto
- also have "... = 0" using assms \<open>i < n\<close> \<open>j < n\<close> \<open>j\<noteq> i\<close> unfolding diagonal_mat_def by simp
- finally show "Matrix.vec_index (Matrix.row B i) j = 0" .
- qed
- show "i < n" using \<open>i < n\<close> .
- qed
- also have "... = B $$(i,i) * Matrix.vec_index v i" using assms \<open>i < n\<close> by auto
- finally show "(Matrix.scalar_prod (Matrix.row B i) v) = B $$(i,i) * Matrix.vec_index v i" .
- qed
- thus ?thesis by auto
- qed
- also have "... = B $$ (i,i) \<cdot>\<^sub>v v" unfolding v_def unit_vec_def by auto
- finally have "B *\<^sub>v v = B $$ (i,i) \<cdot>\<^sub>v v" .
- thus ?thesis unfolding v_def by simp
-qed
-
-lemma mat_vec_mult_assoc:
- assumes "A \<in> carrier_mat n p"
-and "B\<in> carrier_mat p q"
-and "dim_vec v = q"
-shows "A *\<^sub>v (B *\<^sub>v v) = (A * B) *\<^sub>v v" using assms by auto
-
-lemma (in cpx_sq_mat) similar_eigenvectors:
- assumes "A\<in> fc_mats"
- and "B\<in> fc_mats"
- and "P\<in> fc_mats"
- and "similar_mat_wit A B P (Complex_Matrix.adjoint P)"
- and "diagonal_mat B"
- and "i < n"
-shows "A *\<^sub>v (P *\<^sub>v (unit_vec dimR i)) = B $$ (i,i) \<cdot>\<^sub>v (P *\<^sub>v (unit_vec dimR i))"
-proof -
- have "A *\<^sub>v (P *\<^sub>v (unit_vec dimR i)) =
- (P * B * (Complex_Matrix.adjoint P)) *\<^sub>v (P *\<^sub>v (unit_vec dimR i))"
- using assms unfolding similar_mat_wit_def by metis
- also have "... = P * B * (Complex_Matrix.adjoint P) * P *\<^sub>v (unit_vec dimR i)"
- proof (rule mat_vec_mult_assoc[of _ dimR dimR], (auto simp add: assms fc_mats_carrier))
- show "P \<in> carrier_mat dimR dimR" using assms fc_mats_carrier dim_eq by simp
- show "P * B * Complex_Matrix.adjoint P \<in> carrier_mat dimR dimR"
- using assms fc_mats_carrier by auto
- qed
- also have "... = P * B * ((Complex_Matrix.adjoint P) * P) *\<^sub>v (unit_vec dimR i)" using assms dim_eq
- by (smt fc_mats_carrier mat_assoc_test(1) similar_mat_witD2(6) similar_mat_wit_sym)
- also have "... = P * B *\<^sub>v (unit_vec dimR i)"
- proof -
- have "(Complex_Matrix.adjoint P) * P = 1\<^sub>m dimR" using assms dim_eq unfolding similar_mat_wit_def
- by (simp add: fc_mats_carrier)
- thus ?thesis using assms(2) local.fc_mats_carrier dim_eq by auto
- qed
- also have "... = P *\<^sub>v (B *\<^sub>v (unit_vec dimR i))" using mat_vec_mult_assoc assms fc_mats_carrier
- dim_eq by simp
- also have "... = P *\<^sub>v (B $$ (i,i) \<cdot>\<^sub>v (unit_vec dimR i))" using assms diagonal_unit_vec
- fc_mats_carrier dim_eq by simp
- also have "... = B $$ (i,i) \<cdot>\<^sub>v (P *\<^sub>v (unit_vec dimR i))"
- proof (rule mult_mat_vec)
- show "P \<in> carrier_mat dimR dimC" using assms fc_mats_carrier by simp
- show "unit_vec dimR i \<in> carrier_vec dimC" using dim_eq by simp
- qed
- finally show ?thesis .
-qed
-
-
-subsection \<open>Projectors\<close>
-
-definition projector where
-"projector M \<longleftrightarrow> (hermitian M \<and> M * M = M)"
-
-lemma projector_hermitian:
- assumes "projector M"
- shows "hermitian M" using assms unfolding projector_def by simp
-
-lemma zero_projector[simp]:
- shows "projector (0\<^sub>m n n)" unfolding projector_def
-proof
- show "hermitian (0\<^sub>m n n)" using zero_hermitian[of n] by simp
- show "0\<^sub>m n n * 0\<^sub>m n n = 0\<^sub>m n n" by simp
-qed
-
-lemma projector_square_eq:
- assumes "projector M"
- shows "M * M = M" using assms unfolding projector_def by simp
-
-lemma projector_positive:
- assumes "projector M"
- shows "Complex_Matrix.positive M"
-proof (rule positive_if_decomp)
- show "M \<in> carrier_mat (dim_row M) (dim_row M)" using assms projector_hermitian hermitian_square
- by auto
-next
- have "M = Complex_Matrix.adjoint M" using assms projector_hermitian[of M]
- unfolding hermitian_def by simp
- hence "M * Complex_Matrix.adjoint M = M * M" by simp
- also have "... = M" using assms projector_square_eq by auto
- finally have "M * Complex_Matrix.adjoint M = M" .
- thus "\<exists>Ma. Ma * Complex_Matrix.adjoint Ma = M" by auto
-qed
-
-lemma projector_collapse_trace:
- assumes "projector (P::complex Matrix.mat)"
- and "P \<in> carrier_mat n n"
- and "R\<in> carrier_mat n n"
-shows "Complex_Matrix.trace (P * R * P) = Complex_Matrix.trace (R * P)"
-proof -
- have "Complex_Matrix.trace (R * P) = Complex_Matrix.trace (P * R)" using trace_comm assms by auto
- also have "... = Complex_Matrix.trace ((P * P) * R)" using assms projector_square_eq[of P] by simp
- also have "... = Complex_Matrix.trace (P * (P * R))" using assms by auto
- also have "... = Complex_Matrix.trace (P * R * P)" using trace_comm[of P n "P * R"] assms by auto
- finally have "Complex_Matrix.trace (R * P) = Complex_Matrix.trace (P * R * P)" .
- thus ?thesis by simp
-qed
-
-lemma positive_proj_trace:
- assumes "projector (P::complex Matrix.mat)"
- and "Complex_Matrix.positive R"
- and "P \<in> carrier_mat n n"
- and "R\<in> carrier_mat n n"
-shows "Complex_Matrix.trace (R * P) \<ge> 0"
-proof -
- have "Complex_Matrix.trace (R * P) = Complex_Matrix.trace ((P * R) * P)"
- using assms projector_collapse_trace by auto
- also have "... = Complex_Matrix.trace ((P * R) * (Complex_Matrix.adjoint P))"
- using assms projector_hermitian[of P]
- unfolding hermitian_def by simp
- also have "... \<ge> 0"
- proof (rule positive_trace)
- show " P * R * Complex_Matrix.adjoint P \<in> carrier_mat n n" using assms by auto
- show "Complex_Matrix.positive (P * R * Complex_Matrix.adjoint P)"
- by (rule positive_close_under_left_right_mult_adjoint[of _ n], (auto simp add: assms))
- qed
- finally show ?thesis .
-qed
-
-lemma trace_proj_pos_real:
- assumes "projector (P::complex Matrix.mat)"
- and "Complex_Matrix.positive R"
- and "P \<in> carrier_mat n n"
- and "R\<in> carrier_mat n n"
-shows "Re (Complex_Matrix.trace (R * P)) = Complex_Matrix.trace (R * P)"
-proof -
- have "Complex_Matrix.trace (R * P) \<ge> 0" using assms positive_proj_trace by simp
- thus ?thesis by (simp add: complex_eqI less_eq_complex_def)
-qed
-
-lemma (in cpx_sq_mat) trace_sum_mat_proj_pos_real:
- fixes f::"'a \<Rightarrow> real"
- assumes "finite I"
- and "\<forall> i\<in> I. projector (P i)"
- and "Complex_Matrix.positive R"
- and "\<forall>i\<in> I. P i \<in> fc_mats"
- and "R \<in> fc_mats"
-shows "Complex_Matrix.trace (R * (sum_mat (\<lambda>i. f i \<cdot>\<^sub>m (P i)) I)) =
- Re (Complex_Matrix.trace (R * (sum_mat (\<lambda>i. f i \<cdot>\<^sub>m (P i)) I)))"
-proof -
- have sm: "\<And>x. x \<in> I \<Longrightarrow> Complex_Matrix.trace (f x \<cdot>\<^sub>m (R * P x)) =
- f x * Complex_Matrix.trace (R * P x)"
- proof -
- fix i
- assume "i\<in> I"
- show "Complex_Matrix.trace (f i \<cdot>\<^sub>m (R * P i)) = f i * Complex_Matrix.trace (R * P i)"
- proof (rule trace_smult)
- show "R * P i \<in> carrier_mat dimR dimR" using assms cpx_sq_mat_mult fc_mats_carrier \<open>i\<in> I\<close>
- dim_eq by simp
- qed
- qed
- have sw: "\<And>x. x \<in> I \<Longrightarrow> R * (f x \<cdot>\<^sub>m P x) = f x \<cdot>\<^sub>m (R * P x)"
- proof -
- fix i
- assume "i \<in> I"
- show "R * (f i \<cdot>\<^sub>m P i) = f i \<cdot>\<^sub>m (R * P i)"
- proof (rule mult_smult_distrib)
- show "R\<in> carrier_mat dimR dimR" using assms fc_mats_carrier dim_eq by simp
- show "P i \<in> carrier_mat dimR dimR" using assms \<open>i\<in> I\<close> fc_mats_carrier dim_eq by simp
- qed
- qed
- have dr: "Complex_Matrix.trace (R * (sum_mat (\<lambda>i. f i \<cdot>\<^sub>m (P i)) I)) =
- Complex_Matrix.trace (sum_mat (\<lambda>i. (R * (f i \<cdot>\<^sub>m (P i)))) I)"
- using sum_mat_distrib_left[of I] assms by (simp add: cpx_sq_mat_smult)
- also have trs: "... = (\<Sum> i\<in> I. Complex_Matrix.trace (R * (f i \<cdot>\<^sub>m (P i))))"
- proof (rule trace_sum_mat, (simp add: assms))
- show "\<And>i. i \<in> I \<Longrightarrow> R * (f i \<cdot>\<^sub>m P i) \<in> fc_mats" using assms
- by (simp add: cpx_sq_mat_smult cpx_sq_mat_mult)
- qed
- also have "... = (\<Sum> i\<in> I. Complex_Matrix.trace (f i \<cdot>\<^sub>m (R * (P i))))"
- by (rule sum.cong, (simp add: sw)+)
- also have "... = (\<Sum> i\<in> I. f i * Complex_Matrix.trace (R * (P i)))"
- by (rule sum.cong, (simp add: sm)+)
- also have "... = (\<Sum> i\<in> I. complex_of_real (f i * Re (Complex_Matrix.trace (R * (P i)))))"
- proof (rule sum.cong, simp)
- show "\<And>x. x \<in> I \<Longrightarrow>
- complex_of_real (f x) * Complex_Matrix.trace (R * P x) =
- complex_of_real (f x * Re (Complex_Matrix.trace (R * P x)))"
- proof -
- fix x
- assume "x\<in> I"
- have "complex_of_real (f x) * Complex_Matrix.trace (R * P x) =
- complex_of_real (f x) * complex_of_real (Re (Complex_Matrix.trace (R * P x)))"
- using assms sum.cong[of I I] fc_mats_carrier trace_proj_pos_real \<open>x \<in> I\<close> dim_eq by auto
- also have "... = complex_of_real (f x * Re (Complex_Matrix.trace (R * P x)))" by simp
- finally show "complex_of_real (f x) * Complex_Matrix.trace (R * P x) =
- complex_of_real (f x * Re (Complex_Matrix.trace (R * P x)))" .
- qed
- qed
- also have "... = (\<Sum> i\<in> I. f i * Re (Complex_Matrix.trace (R * (P i))))" by simp
- also have "... = (\<Sum> i\<in> I. Re (Complex_Matrix.trace (f i \<cdot>\<^sub>m (R * (P i)))))"
- proof -
- have "(\<Sum> i\<in> I. f i * Re (Complex_Matrix.trace (R * (P i)))) =
- (\<Sum> i\<in> I. Re (Complex_Matrix.trace (f i \<cdot>\<^sub>m (R * (P i)))))"
- by (rule sum.cong, (simp add: sm)+)
- thus ?thesis by simp
- qed
- also have "... = (\<Sum> i\<in> I. Re (Complex_Matrix.trace (R * (f i \<cdot>\<^sub>m (P i)))))"
- proof -
- have "\<And>i. i \<in> I \<Longrightarrow> f i \<cdot>\<^sub>m (R * (P i)) = R * (f i \<cdot>\<^sub>m (P i))" using sw by simp
- thus ?thesis by simp
- qed
- also have "... = Re (\<Sum> i\<in> I. (Complex_Matrix.trace (R * (f i \<cdot>\<^sub>m (P i)))))" by simp
- also have "... = Re (Complex_Matrix.trace (sum_mat (\<lambda>i. R * (f i \<cdot>\<^sub>m (P i))) I))" using trs by simp
- also have "... = Re (Complex_Matrix.trace (R * (sum_mat (\<lambda>i. f i \<cdot>\<^sub>m (P i))) I))" using dr by simp
- finally show ?thesis .
-qed
-
-
-subsection \<open>Rank 1 projection\<close>
-
-definition rank_1_proj where
-"rank_1_proj v = outer_prod v v"
-
-lemma rank_1_proj_square_mat:
- shows "square_mat (rank_1_proj v)" using outer_prod_dim unfolding rank_1_proj_def
- by (metis carrier_matD(1) carrier_matD(2) carrier_vec_dim_vec square_mat.simps)
-
-lemma rank_1_proj_dim[simp]:
- shows "dim_row (rank_1_proj v) = dim_vec v" using outer_prod_dim unfolding rank_1_proj_def
- using carrier_vecI by blast
-
-lemma rank_1_proj_carrier[simp]:
- shows "rank_1_proj v \<in> carrier_mat (dim_vec v) (dim_vec v)" using outer_prod_dim
- unfolding rank_1_proj_def using carrier_vecI by blast
-
-lemma rank_1_proj_coord:
- assumes "i < dim_vec v"
- and "j < dim_vec v"
-shows "(rank_1_proj v) $$ (i, j) = Matrix.vec_index v i * (cnj (Matrix.vec_index v j))" using assms
- unfolding rank_1_proj_def outer_prod_def by auto
-
-lemma rank_1_proj_adjoint:
- shows "Complex_Matrix.adjoint (rank_1_proj (v::complex Matrix.vec)) = rank_1_proj v"
-proof
- show "dim_row (Complex_Matrix.adjoint (rank_1_proj v)) = dim_row (rank_1_proj v)"
- using rank_1_proj_square_mat by auto
- thus "dim_col (Complex_Matrix.adjoint (rank_1_proj v)) = dim_col (rank_1_proj v)" by auto
- fix i j
- assume "i < dim_row (rank_1_proj v)" and "j < dim_col (rank_1_proj v)" note ij = this
- have "Complex_Matrix.adjoint (rank_1_proj v) $$ (i, j) = conjugate ((rank_1_proj v) $$ (j, i))"
- using ij \<open>dim_col (Complex_Matrix.adjoint (rank_1_proj v)) = dim_col (rank_1_proj v)\<close>
- adjoint_eval by fastforce
- also have "... = conjugate (Matrix.vec_index v j * (cnj (Matrix.vec_index v i)))"
- using rank_1_proj_coord ij
- by (metis \<open>dim_col (Complex_Matrix.adjoint (rank_1_proj v)) = dim_col (rank_1_proj v)\<close>
- adjoint_dim_col rank_1_proj_dim)
- also have "... = Matrix.vec_index v i * (cnj (Matrix.vec_index v j))" by simp
- also have "... = rank_1_proj v $$ (i, j)" using ij rank_1_proj_coord
- by (metis \<open>dim_col (Complex_Matrix.adjoint (rank_1_proj v)) = dim_col (rank_1_proj v)\<close>
- adjoint_dim_col rank_1_proj_dim)
- finally show "Complex_Matrix.adjoint (rank_1_proj v) $$ (i, j) = rank_1_proj v $$ (i, j)".
-qed
-
-lemma rank_1_proj_hermitian:
- shows "hermitian (rank_1_proj (v::complex Matrix.vec))" using rank_1_proj_adjoint
- unfolding hermitian_def by simp
-
-lemma rank_1_proj_trace:
- assumes "\<parallel>v\<parallel> = 1"
- shows "Complex_Matrix.trace (rank_1_proj v) = 1"
-proof -
- have "Complex_Matrix.trace (rank_1_proj v) = inner_prod v v" using trace_outer_prod
- unfolding rank_1_proj_def using carrier_vecI by blast
- also have "... = (vec_norm v)\<^sup>2" unfolding vec_norm_def using power2_csqrt by presburger
- also have "... = \<parallel>v\<parallel>\<^sup>2" using vec_norm_sq_cpx_vec_length_sq by simp
- also have "... = 1" using assms by simp
- finally show ?thesis .
-qed
-
-lemma rank_1_proj_mat_col:
- assumes "A \<in> carrier_mat n n"
- and "i < n"
- and "j < n"
- and "k < n"
-shows "(rank_1_proj (Matrix.col A i)) $$ (j, k) = A $$ (j, i) * conjugate (A $$ (k,i))"
-proof -
- have "(rank_1_proj (Matrix.col A i)) $$ (j, k) = Matrix.vec_index (Matrix.col A i) j *
- conjugate (Matrix.vec_index (Matrix.col A i) k)" using index_outer_prod[of "Matrix.col A i" n "Matrix.col A i" n]
- assms unfolding rank_1_proj_def by simp
- also have "... = A $$ (j, i) * conjugate (A $$ (k,i))" using assms by simp
- finally show ?thesis .
-qed
-
-lemma (in cpx_sq_mat) weighted_sum_rank_1_proj_unitary_index:
- assumes "A \<in> fc_mats"
-and "B \<in> fc_mats"
-and "diagonal_mat B"
-and "Complex_Matrix.unitary A"
-and "j < dimR"
-and "k < dimR"
-shows "(sum_mat (\<lambda>i. (diag_mat B)!i \<cdot>\<^sub>m rank_1_proj (Matrix.col A i)) {..< dimR}) $$ (j,k) =
- (A * B * (Complex_Matrix.adjoint A)) $$ (j,k)"
-proof -
- have "(sum_mat (\<lambda>i. (diag_mat B)!i \<cdot>\<^sub>m rank_1_proj (Matrix.col A i)) {..< dimR}) $$ (j,k) =
- (\<Sum> i\<in> {..< dimR}. ((diag_mat B)!i \<cdot>\<^sub>m rank_1_proj (Matrix.col A i)) $$ (j,k))"
- proof (rule sum_mat_index, (auto simp add: fc_mats_carrier assms))
- show "\<And>i. i < dimR \<Longrightarrow> (diag_mat B)!i \<cdot>\<^sub>m rank_1_proj (Matrix.col A i) \<in> carrier_mat dimR dimC"
- using rank_1_proj_carrier assms fc_mats_carrier dim_eq
- by (metis smult_carrier_mat zero_col_col zero_col_dim)
- show "k < dimC" using assms dim_eq by simp
- qed
- also have "... = (\<Sum> i\<in> {..< dimR}. (diag_mat B)!i* A $$ (j, i) * conjugate (A $$ (k,i)))"
- proof (rule sum.cong, simp)
- have idx: "\<And>x. x \<in> {..<dimR} \<Longrightarrow> (rank_1_proj (Matrix.col A x)) $$ (j, k) =
- A $$ (j, x) * conjugate (A $$ (k, x))" using rank_1_proj_mat_col assms fc_mats_carrier dim_eq
- by blast
- show "\<And>x. x \<in> {..<dimR} \<Longrightarrow> ((diag_mat B)!x \<cdot>\<^sub>m rank_1_proj (Matrix.col A x)) $$ (j, k) =
- (diag_mat B)!x * A $$ (j, x) * conjugate (A $$ (k, x))"
- proof -
- fix x
- assume "x\<in> {..< dimR}"
- have "((diag_mat B)!x \<cdot>\<^sub>m rank_1_proj (Matrix.col A x)) $$ (j, k) =
- (diag_mat B)!x * (rank_1_proj (Matrix.col A x)) $$ (j, k)"
- proof (rule index_smult_mat)
- show "j < dim_row (rank_1_proj (Matrix.col A x))" using \<open>x\<in> {..< dimR}\<close> assms fc_mats_carrier by simp
- show "k < dim_col (rank_1_proj (Matrix.col A x))" using \<open>x\<in> {..< dimR}\<close> assms fc_mats_carrier
- rank_1_proj_carrier[of "Matrix.col A x"] by simp
- qed
- thus "((diag_mat B)!x \<cdot>\<^sub>m rank_1_proj (Matrix.col A x)) $$ (j, k) =
- (diag_mat B)!x * A $$ (j, x) * conjugate (A $$ (k, x))" using idx \<open>x\<in> {..< dimR}\<close> by simp
- qed
- qed
- also have "... = Matrix.scalar_prod (Matrix.col (Complex_Matrix.adjoint A) k) (Matrix.row (A*B) j) "
- unfolding Matrix.scalar_prod_def
- proof (rule sum.cong)
- show "{..<dimR} = {0..<dim_vec (Matrix.row (A*B) j)}"
- using assms fc_mats_carrier dim_eq by auto
- show "\<And>x. x \<in> {0..<dim_vec (Matrix.row (A*B) j)} \<Longrightarrow>
- diag_mat B ! x * A $$ (j, x) * conjugate (A $$ (k, x)) =
- Matrix.vec_index ((Matrix.col (Complex_Matrix.adjoint A) k)) x *
- Matrix.vec_index (Matrix.row (A*B) j) x"
- proof -
- fix x
- assume "x\<in> {0..<dim_vec (Matrix.row (A*B) j)}"
- hence "x\<in> {0..<dimR}" using assms fc_mats_carrier dim_eq by simp
- have "diag_mat B ! x *A $$ (j, x) = Matrix.vec_index (Matrix.row (A*B) j) x"
- proof (rule times_diag_index[symmetric, of _ dimR], (auto simp add: assms))
- show "A \<in> carrier_mat dimR dimR" using assms fc_mats_carrier dim_eq by simp
- show "B \<in> carrier_mat dimR dimR" using assms fc_mats_carrier dim_eq by simp
- show "x < dimR" using \<open>x\<in>{0..< dimR}\<close> by simp
- qed
- moreover have "conjugate (A $$ (k, x)) =
- Matrix.vec_index ((Matrix.col (Complex_Matrix.adjoint A) k)) x" using \<open>x\<in> {0..<dimR}\<close> assms
- fc_mats_carrier dim_eq by (simp add: adjoint_col)
- ultimately show "diag_mat B ! x * A $$ (j, x) * conjugate (A $$ (k, x)) =
- Matrix.vec_index ((Matrix.col (Complex_Matrix.adjoint A) k)) x *
- Matrix.vec_index (Matrix.row (A*B) j) x" by simp
- qed
- qed
- also have "... = (A*B * (Complex_Matrix.adjoint A)) $$ (j,k)"
- proof -
- have "Matrix.mat (dim_row (A*B)) (dim_col (Complex_Matrix.adjoint A))
- (\<lambda>(i, j). Matrix.scalar_prod (Matrix.row (A*B) i) (Matrix.col (Complex_Matrix.adjoint A) j)) $$
- (j, k) = Matrix.scalar_prod (Matrix.row (A*B) j) (Matrix.col (Complex_Matrix.adjoint A) k)"
- using assms fc_mats_carrier by simp
- also have "... = Matrix.scalar_prod (Matrix.col (Complex_Matrix.adjoint A) k) (Matrix.row (A*B) j)"
- using assms comm_scalar_prod[of "Matrix.row (A*B) j" dimR] fc_mats_carrier dim_eq
- by (metis adjoint_dim Matrix.col_carrier_vec row_carrier_vec cpx_sq_mat_mult)
- thus ?thesis using assms fc_mats_carrier unfolding times_mat_def by simp
- qed
- finally show ?thesis .
-qed
-
-lemma (in cpx_sq_mat) weighted_sum_rank_1_proj_unitary:
- assumes "A \<in> fc_mats"
-and "B \<in> fc_mats"
-and "diagonal_mat B"
-and "Complex_Matrix.unitary A"
-shows "(sum_mat (\<lambda>i. (diag_mat B)!i \<cdot>\<^sub>m rank_1_proj (Matrix.col A i)) {..< dimR}) =
- (A * B * (Complex_Matrix.adjoint A))"
-proof
- show "dim_row (sum_mat (\<lambda>i. diag_mat B ! i \<cdot>\<^sub>m rank_1_proj (Matrix.col A i)) {..<dimR}) =
- dim_row (A * B * Complex_Matrix.adjoint A)" using assms fc_mats_carrier dim_eq
- by (metis (no_types, lifting) carrier_matD(1) cpx_sq_mat_smult dim_col index_mult_mat(2)
- rank_1_proj_carrier sum_mat_carrier)
- show "dim_col (local.sum_mat (\<lambda>i. diag_mat B ! i \<cdot>\<^sub>m rank_1_proj (Matrix.col A i)) {..<dimR}) =
- dim_col (A * B * Complex_Matrix.adjoint A)" using assms fc_mats_carrier dim_eq
- by (metis (mono_tags, lifting) adjoint_dim_col carrier_matD(1) carrier_matI dim_col
- index_mult_mat(3) index_smult_mat(2) index_smult_mat(3) rank_1_proj_dim
- rank_1_proj_square_mat square_mat.elims(2) square_mats sum_mat_carrier)
- show "\<And>i j. i < dim_row (A * B * Complex_Matrix.adjoint A) \<Longrightarrow>
- j < dim_col (A * B * Complex_Matrix.adjoint A) \<Longrightarrow>
- local.sum_mat (\<lambda>i. diag_mat B ! i \<cdot>\<^sub>m rank_1_proj (Matrix.col A i)) {..<dimR} $$ (i, j) =
- (A * B * Complex_Matrix.adjoint A) $$ (i, j)"
- using weighted_sum_rank_1_proj_unitary_index[of A B] dim_eq fc_mats_carrier using assms by auto
-qed
-
-lemma rank_1_proj_projector:
- assumes "\<parallel>v\<parallel> = 1"
- shows "projector (rank_1_proj v)"
-proof -
- have "Complex_Matrix.adjoint (rank_1_proj v) = rank_1_proj v" using rank_1_proj_adjoint by simp
- hence a: "hermitian (rank_1_proj v)" unfolding hermitian_def by simp
- have "rank_1_proj v * rank_1_proj v = inner_prod v v \<cdot>\<^sub>m outer_prod v v" unfolding rank_1_proj_def
- using outer_prod_mult_outer_prod assms using carrier_vec_dim_vec by blast
- also have "... = (vec_norm v)\<^sup>2 \<cdot>\<^sub>m outer_prod v v" unfolding vec_norm_def using power2_csqrt
- by presburger
- also have "... = (cpx_vec_length v)\<^sup>2 \<cdot>\<^sub>m outer_prod v v " using vec_norm_sq_cpx_vec_length_sq
- by simp
- also have "... = outer_prod v v" using assms state_qbit_norm_sq smult_one[of "outer_prod v v"]
- by simp
- also have "... = rank_1_proj v" unfolding rank_1_proj_def by simp
- finally show ?thesis using a unfolding projector_def by simp
-qed
-
-lemma rank_1_proj_positive:
- assumes "\<parallel>v\<parallel> = 1"
- shows "Complex_Matrix.positive (rank_1_proj v)"
-proof -
- have "projector (rank_1_proj v)" using assms rank_1_proj_projector by simp
- thus ?thesis using projector_positive by simp
-qed
-
-lemma rank_1_proj_density:
- assumes "\<parallel>v\<parallel> = 1"
- shows "density_operator (rank_1_proj v)" unfolding density_operator_def using assms
- by (simp add: rank_1_proj_positive rank_1_proj_trace)
-
-lemma (in cpx_sq_mat) sum_rank_1_proj_unitary_index:
- assumes "A \<in> fc_mats"
-and "Complex_Matrix.unitary A"
-and "j < dimR"
-and "k < dimR"
-shows "(sum_mat (\<lambda>i. rank_1_proj (Matrix.col A i)) {..< dimR}) $$ (j,k) = (1\<^sub>m dimR) $$ (j,k)"
-proof -
- have "(sum_mat (\<lambda>i. rank_1_proj (Matrix.col A i)) {..< dimR}) $$ (j,k) =
- (\<Sum> i\<in> {..< dimR}. (rank_1_proj (Matrix.col A i)) $$ (j,k))"
- proof (rule sum_mat_index, (auto simp add: fc_mats_carrier assms))
- show "\<And>i. i < dimR \<Longrightarrow> rank_1_proj (Matrix.col A i) \<in> carrier_mat dimR dimC"
- proof -
- fix i
- assume "i < dimR"
- hence "dim_vec (Matrix.col A i) = dimR" using assms fc_mats_carrier by simp
- thus "rank_1_proj (Matrix.col A i) \<in> carrier_mat dimR dimC" using rank_1_proj_carrier assms
- fc_mats_carrier dim_eq fc_mats_carrier by (metis dim_col fc_mats_carrier)
- qed
- show "k < dimC" using assms dim_eq by simp
- qed
- also have "... = (\<Sum> i\<in> {..< dimR}. A $$ (j, i) * conjugate (A $$ (k,i)))"
- proof (rule sum.cong, simp)
- show "\<And>x. x \<in> {..<dimR} \<Longrightarrow> rank_1_proj (Matrix.col A x) $$ (j, k) =
- A $$ (j, x) * conjugate (A $$ (k, x))"
- using rank_1_proj_mat_col assms using local.fc_mats_carrier dim_eq by blast
- qed
- also have "... = Matrix.scalar_prod (Matrix.col (Complex_Matrix.adjoint A) k) (Matrix.row A j) "
- unfolding Matrix.scalar_prod_def
- proof (rule sum.cong)
- show "{..<dimR} = {0..<dim_vec (Matrix.row A j)}"
- using assms fc_mats_carrier dim_eq by auto
- show "\<And>x. x \<in> {0..<dim_vec (Matrix.row A j)} \<Longrightarrow>
- A $$ (j, x) * conjugate (A $$ (k, x)) =
- Matrix.vec_index ((Matrix.col (Complex_Matrix.adjoint A) k)) x * Matrix.vec_index (Matrix.row A j) x"
- proof -
- fix x
- assume "x\<in> {0..<dim_vec (Matrix.row A j)}"
- hence "x\<in> {0..<dimR}" using assms fc_mats_carrier dim_eq by simp
- hence "A $$ (j, x) = Matrix.vec_index (Matrix.row A j) x" using \<open>x\<in> {0..<dimR}\<close>
- assms fc_mats_carrier dim_eq by simp
- moreover have "conjugate (A $$ (k, x)) =
- Matrix.vec_index ((Matrix.col (Complex_Matrix.adjoint A) k)) x" using \<open>x\<in> {0..<dimR}\<close>
- assms fc_mats_carrier dim_eq by (simp add: adjoint_col)
- ultimately show "A $$ (j, x) * conjugate (A $$ (k, x)) =
- Matrix.vec_index ((Matrix.col (Complex_Matrix.adjoint A) k)) x *
- Matrix.vec_index (Matrix.row A j) x" by simp
- qed
- qed
- also have "... = (A * (Complex_Matrix.adjoint A)) $$ (j,k)"
- proof -
- have "Matrix.mat (dim_row A) (dim_col (Complex_Matrix.adjoint A))
- (\<lambda>(i, j). Matrix.scalar_prod (Matrix.row A i) (Matrix.col (Complex_Matrix.adjoint A) j)) $$
- (j, k) = Matrix.scalar_prod (Matrix.row A j) (Matrix.col (Complex_Matrix.adjoint A) k)"
- using assms fc_mats_carrier by simp
- also have "... = Matrix.scalar_prod (Matrix.col (Complex_Matrix.adjoint A) k) (Matrix.row A j)"
- using assms comm_scalar_prod[of "Matrix.row A j" dimR] fc_mats_carrier
- by (simp add: adjoint_dim dim_eq)
- thus ?thesis using assms fc_mats_carrier unfolding times_mat_def by simp
- qed
- also have "... = (1\<^sub>m dimR) $$ (j,k)" using assms dim_eq by (simp add: fc_mats_carrier)
- finally show ?thesis .
-qed
-
-lemma (in cpx_sq_mat) rank_1_proj_sum_density:
- assumes "finite I"
- and "\<forall>i\<in> I. \<parallel>u i\<parallel> = 1"
- and "\<forall>i\<in> I. dim_vec (u i) = dimR"
- and "\<forall>i\<in> I. 0 \<le> p i"
- and "(\<Sum>i\<in> I. p i) = 1"
-shows "density_operator (sum_mat (\<lambda>i. p i \<cdot>\<^sub>m (rank_1_proj (u i))) I)" unfolding density_operator_def
-proof
- have "Complex_Matrix.trace (sum_mat (\<lambda>i. p i \<cdot>\<^sub>m rank_1_proj (u i)) I) =
- (\<Sum>i\<in> I. Complex_Matrix.trace (p i \<cdot>\<^sub>m rank_1_proj (u i)))"
- proof (rule trace_sum_mat, (simp add: assms))
- show "\<And>i. i \<in> I \<Longrightarrow> p i \<cdot>\<^sub>m rank_1_proj (u i) \<in> fc_mats" using assms smult_mem fc_mats_carrier
- dim_eq by auto
- qed
- also have "... = (\<Sum>i\<in> I. p i * Complex_Matrix.trace (rank_1_proj (u i)))"
- proof -
- {
- fix i
- assume "i \<in> I"
- hence "Complex_Matrix.trace (p i \<cdot>\<^sub>m rank_1_proj (u i)) =
- p i * Complex_Matrix.trace (rank_1_proj (u i))"
- using trace_smult[of "rank_1_proj (u i)" dimR] assms fc_mats_carrier dim_eq by auto
- }
- thus ?thesis by simp
- qed
- also have "... = 1" using assms rank_1_proj_trace assms by auto
- finally show "Complex_Matrix.trace (sum_mat (\<lambda>i. p i \<cdot>\<^sub>m rank_1_proj (u i)) I) = 1" .
-next
- show "Complex_Matrix.positive (sum_mat (\<lambda>i. p i \<cdot>\<^sub>m rank_1_proj (u i)) I)"
- proof (rule sum_mat_positive, (simp add: assms))
- fix i
- assume "i\<in> I"
- thus "p i \<cdot>\<^sub>m rank_1_proj (u i) \<in> fc_mats" using assms smult_mem fc_mats_carrier dim_eq by auto
- show "Complex_Matrix.positive (p i \<cdot>\<^sub>m rank_1_proj (u i))" using assms \<open>i\<in> I\<close>
- rank_1_proj_positive positive_smult[of "rank_1_proj (u i)" dimR] dim_eq by auto
- qed
-qed
-
-
-
-lemma (in cpx_sq_mat) sum_rank_1_proj_unitary:
- assumes "A \<in> fc_mats"
-and "Complex_Matrix.unitary A"
-shows "(sum_mat (\<lambda>i. rank_1_proj (Matrix.col A i)) {..< dimR})= (1\<^sub>m dimR)"
-proof
- show "dim_row (sum_mat (\<lambda>i. rank_1_proj (Matrix.col A i)) {..<dimR}) = dim_row (1\<^sub>m dimR)"
- using assms fc_mats_carrier
- by (metis carrier_matD(1) dim_col dim_eq index_one_mat(2) rank_1_proj_carrier sum_mat_carrier)
- show "dim_col (sum_mat (\<lambda>i. rank_1_proj (Matrix.col A i)) {..<dimR}) = dim_col (1\<^sub>m dimR)"
- using assms fc_mats_carrier rank_1_proj_carrier sum_mat_carrier
- by (metis carrier_matD(2) dim_col dim_eq index_one_mat(3) square_mat.simps square_mats)
- show "\<And>i j. i < dim_row (1\<^sub>m dimR) \<Longrightarrow>
- j < dim_col (1\<^sub>m dimR) \<Longrightarrow> sum_mat (\<lambda>i. rank_1_proj (Matrix.col A i)) {..<dimR} $$ (i, j) =
- 1\<^sub>m dimR $$ (i, j)"
- using assms sum_rank_1_proj_unitary_index by simp
-qed
-
-lemma (in cpx_sq_mat) rank_1_proj_unitary:
- assumes "A \<in> fc_mats"
-and "Complex_Matrix.unitary A"
-and "j < dimR"
-and "k < dimR"
-shows "rank_1_proj (Matrix.col A j) * (rank_1_proj (Matrix.col A k)) =
- (1\<^sub>m dimR) $$ (j,k) \<cdot>\<^sub>m (outer_prod (Matrix.col A j) (Matrix.col A k))"
-proof -
- have "rank_1_proj (Matrix.col A j) * (rank_1_proj (Matrix.col A k)) =
- Complex_Matrix.inner_prod (Matrix.col A j) (Matrix.col A k) \<cdot>\<^sub>m outer_prod (Matrix.col A j) (Matrix.col A k)"
- using outer_prod_mult_outer_prod assms Matrix.col_dim local.fc_mats_carrier unfolding rank_1_proj_def
- by blast
- also have "... = (Complex_Matrix.adjoint A * A) $$ (j, k)\<cdot>\<^sub>m (outer_prod (Matrix.col A j) (Matrix.col A k))"
- using inner_prod_adjoint_comp[of A dimR A] assms fc_mats_carrier dim_eq by simp
- also have "... = (1\<^sub>m dimR) $$ (j,k) \<cdot>\<^sub>m (outer_prod (Matrix.col A j) (Matrix.col A k))" using assms
- unfolding Complex_Matrix.unitary_def
- by (metis add_commute assms(2) index_add_mat(2) index_one_mat(2) one_mem unitary_simps(1))
- finally show ?thesis .
-qed
-
-lemma (in cpx_sq_mat) rank_1_proj_unitary_ne:
- assumes "A \<in> fc_mats"
-and "Complex_Matrix.unitary A"
-and "j < dimR"
-and "k < dimR"
-and "j\<noteq> k"
-shows "rank_1_proj (Matrix.col A j) * (rank_1_proj (Matrix.col A k)) = (0\<^sub>m dimR dimR)"
-proof -
- have "dim_row (0 \<cdot>\<^sub>m outer_prod (Matrix.col A j) (Matrix.col A k)) = dim_row (outer_prod (Matrix.col A j) (Matrix.col A k))"
- by simp
- also have "... = dimR" using assms fc_mats_carrier dim_eq by auto
- finally have rw: "dim_row (0 \<cdot>\<^sub>m outer_prod (Matrix.col A j) (Matrix.col A k)) = dimR" .
- have "dim_col (0 \<cdot>\<^sub>m outer_prod (Matrix.col A j) (Matrix.col A k)) = dim_col (outer_prod (Matrix.col A j) (Matrix.col A k))"
- by simp
- also have "... = dimR" using assms fc_mats_carrier dim_eq by auto
- finally have cl: "dim_col (0 \<cdot>\<^sub>m outer_prod (Matrix.col A j) (Matrix.col A k)) = dimR" .
- have "rank_1_proj (Matrix.col A j) * (rank_1_proj (Matrix.col A k)) =
- (0::complex) \<cdot>\<^sub>m (outer_prod (Matrix.col A j) (Matrix.col A k))"
- using assms rank_1_proj_unitary by simp
- also have "... = (0\<^sub>m dimR dimR)"
- proof
- show "dim_row (0 \<cdot>\<^sub>m outer_prod (Matrix.col A j) (Matrix.col A k)) = dim_row (0\<^sub>m dimR dimR)" using rw by simp
- next
- show "dim_col (0 \<cdot>\<^sub>m outer_prod (Matrix.col A j) (Matrix.col A k)) = dim_col (0\<^sub>m dimR dimR)" using cl by simp
- next
- show "\<And>i p. i < dim_row (0\<^sub>m dimR dimR) \<Longrightarrow> p < dim_col (0\<^sub>m dimR dimR) \<Longrightarrow>
- (0 \<cdot>\<^sub>m outer_prod (Matrix.col A j) (Matrix.col A k)) $$ (i, p) = 0\<^sub>m dimR dimR $$ (i, p)" using rw cl by auto
- qed
- finally show ?thesis .
-qed
-
-lemma (in cpx_sq_mat) rank_1_proj_unitary_eq:
- assumes "A \<in> fc_mats"
-and "Complex_Matrix.unitary A"
-and "j < dimR"
-shows "rank_1_proj (Matrix.col A j) * (rank_1_proj (Matrix.col A j)) = rank_1_proj (Matrix.col A j)"
-proof -
- have "rank_1_proj (Matrix.col A j) * (rank_1_proj (Matrix.col A j)) = (1::complex) \<cdot>\<^sub>m (rank_1_proj (Matrix.col A j))"
- using assms rank_1_proj_unitary unfolding rank_1_proj_def by simp
- also have "... = (rank_1_proj (Matrix.col A j))" by (simp add: smult_one)
- finally show ?thesis .
-qed
-
-
+(*
+Author:
+ Mnacho Echenim, Université Grenoble Alpes
+*)
+
+theory Linear_Algebra_Complements imports
+ "Isabelle_Marries_Dirac.Tensor"
+ "Isabelle_Marries_Dirac.More_Tensor"
+ "QHLProver.Gates"
+ "HOL-Types_To_Sets.Group_On_With"
+ "HOL-Probability.Probability"
+
+
+begin
+hide_const(open) S
+section \<open>Preliminaries\<close>
+
+
+subsection \<open>Misc\<close>
+
+lemma mult_real_cpx:
+ fixes a::complex
+ fixes b::complex
+ assumes "a\<in> Reals"
+ shows "a* (Re b) = Re (a * b)" using assms
+ by (metis Reals_cases complex.exhaust complex.sel(1) complex_of_real_mult_Complex of_real_mult)
+
+lemma fct_bound:
+ fixes f::"complex\<Rightarrow> real"
+ assumes "f(-1) + f 1 = 1"
+and "0 \<le> f 1"
+and "0 \<le> f (-1)"
+shows "-1 \<le> f 1 - f(-1) \<and> f 1 - f(-1) \<le> 1"
+proof
+ have "f 1 - f(-1) = 1 - f(-1) - f(-1)" using assms by simp
+ also have "...\<ge> -1" using assms by simp
+ finally show "-1 \<le> f 1 - f(-1)" .
+next
+ have "f(-1) - f 1 = 1 - f 1 - f 1 " using assms by simp
+ also have "... \<ge> -1" using assms by simp
+ finally have "-1 \<le> f(-1) - f 1" .
+ thus "f 1 - f (-1) \<le> 1" by simp
+qed
+
+lemma fct_bound':
+ fixes f::"complex\<Rightarrow> real"
+ assumes "f(-1) + f 1 = 1"
+and "0 \<le> f 1"
+and "0 \<le> f (-1)"
+shows "\<bar>f 1 - f(-1)\<bar> \<le> 1" using assms fct_bound by auto
+
+lemma pos_sum_1_le:
+ assumes "finite I"
+and "\<forall> i \<in> I. (0::real) \<le> f i"
+and "(\<Sum>i\<in> I. f i) = 1"
+and "j\<in> I"
+shows "f j \<le> 1"
+proof (rule ccontr)
+ assume "\<not> f j \<le> 1"
+ hence "1 < f j" by simp
+ hence "1 < (\<Sum>i\<in> I. f i)" using assms by (metis \<open>\<not> f j \<le> 1\<close> sum_nonneg_leq_bound)
+ thus False using assms by simp
+qed
+
+lemma last_subset:
+ assumes "A \<subseteq> {a,b}"
+ and "a\<noteq> b"
+and "A \<noteq> {a, b}"
+and "A\<noteq> {}"
+and "A \<noteq> {a}"
+shows "A = {b}" using assms by blast
+
+lemma disjoint_Un:
+ assumes "disjoint_family_on A (insert x F)"
+ and "x\<notin> F"
+shows "(A x) \<inter> (\<Union> a\<in> F. A a) = {}"
+proof -
+ have "(A x) \<inter> (\<Union> a\<in> F. A a) = (\<Union>i\<in>F. (A x) \<inter> A i)" using Int_UN_distrib by simp
+ also have "... = (\<Union>i\<in>F. {})" using assms disjoint_family_onD by fastforce
+ also have "... = {}" using SUP_bot_conv(2) by simp
+ finally show ?thesis .
+qed
+
+lemma sum_but_one:
+ assumes "\<forall>j < (n::nat). j \<noteq>i \<longrightarrow> f j = (0::'a::ring)"
+ and "i < n"
+ shows "(\<Sum> j \<in> {0 ..< n}. f j * g j) = f i * g i"
+proof -
+ have "sum (\<lambda>x. f x * g x) (({0 ..< n} - {i}) \<union> {i}) = sum (\<lambda>x. f x * g x) ({0 ..< n} - {i}) +
+ sum (\<lambda>x. f x * g x) {i}" by (rule sum.union_disjoint, auto)
+ also have "... = sum (\<lambda>x. f x * g x) {i}" using assms by auto
+ also have "... = f i * g i" by simp
+ finally have "sum (\<lambda>x. f x * g x) (({0 ..< n} - {i}) \<union> {i}) = f i * g i" .
+ moreover have "{0 ..< n} = ({0 ..< n} - {i}) \<union> {i}" using assms by auto
+ ultimately show ?thesis by simp
+qed
+
+lemma sum_2_elems:
+ assumes "I = {a,b}"
+ and "a\<noteq> b"
+ shows "(\<Sum>a\<in>I. f a) = f a + f b"
+proof -
+ have "(\<Sum>a\<in>I. f a) = (\<Sum>a\<in>(insert a {b}). f a)" using assms by simp
+ also have "... = f a + (\<Sum>a\<in>{b}. f a)"
+ proof (rule sum.insert)
+ show "finite {b}" by simp
+ show "a\<notin> {b}" using assms by simp
+ qed
+ also have "... = f a + f b" by simp
+ finally show ?thesis .
+qed
+
+lemma sum_4_elems:
+ shows "(\<Sum>i<(4::nat). f i) = f 0 + f 1 + f 2 + f 3"
+proof -
+ have "(\<Sum>i<(4::nat). f i) = (\<Sum>i<(3::nat). f i) + f 3"
+ by (metis Suc_numeral semiring_norm(2) semiring_norm(8) sum.lessThan_Suc)
+ moreover have "(\<Sum>i<(3::nat). f i) = (\<Sum>i<(2::nat). f i) + f 2"
+ by (metis Suc_1 add_2_eq_Suc' nat_1_add_1 numeral_code(3) numerals(1)
+ one_plus_numeral_commute sum.lessThan_Suc)
+ moreover have "(\<Sum>i<(2::nat). f i) = (\<Sum>i<(1::nat). f i) + f 1"
+ by (metis Suc_1 sum.lessThan_Suc)
+ ultimately show ?thesis by simp
+qed
+
+lemma disj_family_sum:
+ shows "finite I \<Longrightarrow> disjoint_family_on A I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> finite (A i)) \<Longrightarrow>
+ (\<Sum> i \<in> (\<Union>n \<in> I. A n). f i) = (\<Sum> n\<in> I. (\<Sum> i \<in> A n. f i))"
+proof (induct rule:finite_induct)
+ case empty
+ then show ?case by simp
+next
+ case (insert x F)
+ hence "disjoint_family_on A F"
+ by (meson disjoint_family_on_mono subset_insertI)
+ have "(\<Union>n \<in> (insert x F). A n) = A x \<union> (\<Union>n \<in> F. A n)" using insert by simp
+ hence "(\<Sum> i \<in> (\<Union>n \<in> (insert x F). A n). f i) = (\<Sum> i \<in> (A x \<union> (\<Union>n \<in> F. A n)). f i)" by simp
+ also have "... = (\<Sum> i \<in> A x. f i) + (\<Sum> i \<in> (\<Union>n \<in> F. A n). f i)"
+ by (rule sum.union_disjoint, (simp add: insert disjoint_Un)+)
+ also have "... = (\<Sum> i \<in> A x. f i) + (\<Sum>n\<in>F. sum f (A n))" using \<open>disjoint_family_on A F\<close>
+ by (simp add: insert)
+ also have "... = (\<Sum>n\<in>(insert x F). sum f (A n))" using insert by simp
+ finally show ?case .
+qed
+
+lemma integrable_real_mult_right:
+ fixes c::real
+ assumes "integrable M f"
+ shows "integrable M (\<lambda>w. c * f w)"
+proof (cases "c = 0")
+ case True
+ thus ?thesis by simp
+next
+ case False
+ thus ?thesis using integrable_mult_right[of c] assms by simp
+qed
+
+
+subsection \<open>Unifying notions between Isabelle Marries Dirac and QHLProver\<close>
+
+lemma mult_conj_cmod_square:
+ fixes z::complex
+ shows "z * conjugate z = (cmod z)\<^sup>2"
+proof -
+ have "z * conjugate z = (Re z)\<^sup>2 + (Im z)\<^sup>2" using complex_mult_cnj by auto
+ also have "... = (cmod z)\<^sup>2" unfolding cmod_def by simp
+ finally show ?thesis .
+qed
+
+lemma vec_norm_sq_cpx_vec_length_sq:
+ shows "(vec_norm v)\<^sup>2 = (cpx_vec_length v)\<^sup>2"
+proof -
+ have "(vec_norm v)\<^sup>2 = inner_prod v v" unfolding vec_norm_def using power2_csqrt by blast
+ also have "... = (\<Sum>i<dim_vec v. (cmod (Matrix.vec_index v i))\<^sup>2)" unfolding Matrix.scalar_prod_def
+ proof -
+ have "\<And>i. i < dim_vec v \<Longrightarrow> Matrix.vec_index v i * conjugate (Matrix.vec_index v i) =
+ (cmod (Matrix.vec_index v i))\<^sup>2" using mult_conj_cmod_square by simp
+ thus "(\<Sum>i = 0..<dim_vec (conjugate v). Matrix.vec_index v i *
+ Matrix.vec_index (conjugate v) i) = (\<Sum>i<dim_vec v. (cmod (Matrix.vec_index v i))\<^sup>2)"
+ by (simp add: lessThan_atLeast0)
+ qed
+ finally show "(vec_norm v)\<^sup>2 = (cpx_vec_length v)\<^sup>2" unfolding cpx_vec_length_def
+ by (simp add: sum_nonneg)
+qed
+
+lemma vec_norm_eq_cpx_vec_length:
+ shows "vec_norm v = cpx_vec_length v" using vec_norm_sq_cpx_vec_length_sq
+by (metis cpx_vec_length_inner_prod inner_prod_csqrt power2_csqrt vec_norm_def)
+
+lemma cpx_vec_length_square:
+ shows "\<parallel>v\<parallel>\<^sup>2 = (\<Sum>i = 0..<dim_vec v. (cmod (Matrix.vec_index v i))\<^sup>2)" unfolding cpx_vec_length_def
+ by (simp add: lessThan_atLeast0 sum_nonneg)
+
+lemma state_qbit_norm_sq:
+ assumes "v\<in> state_qbit n"
+ shows "(cpx_vec_length v)\<^sup>2 = 1"
+proof -
+ have "cpx_vec_length v = 1" using assms unfolding state_qbit_def by simp
+ thus ?thesis by simp
+qed
+
+lemma dagger_adjoint:
+shows "dagger M = Complex_Matrix.adjoint M" unfolding dagger_def Complex_Matrix.adjoint_def
+ by (simp add: cong_mat)
+
+
+subsection \<open>Types to sets lemmata transfers\<close>
+
+context ab_group_add_on_with begin
+
+context includes lifting_syntax assumes ltd: "\<exists>(Rep::'s \<Rightarrow> 'a) (Abs::'a \<Rightarrow> 's).
+ type_definition Rep Abs S" begin
+interpretation local_typedef_ab_group_add_on_with pls z mns um S "TYPE('s)" by unfold_locales fact
+
+lemmas lt_sum_union_disjoint = sum.union_disjoint
+ [var_simplified explicit_ab_group_add,
+ unoverload_type 'c,
+ OF type.comm_monoid_add_axioms,
+ untransferred]
+
+lemmas lt_disj_family_sum = disj_family_sum
+ [var_simplified explicit_ab_group_add,
+ unoverload_type 'd,
+OF type.comm_monoid_add_axioms,
+ untransferred]
+
+lemmas lt_sum_reindex_cong = sum.reindex_cong
+ [var_simplified explicit_ab_group_add,
+ unoverload_type 'd,
+OF type.comm_monoid_add_axioms,
+ untransferred]
+end
+
+lemmas sum_with_union_disjoint =
+ lt_sum_union_disjoint
+ [cancel_type_definition,
+ OF carrier_ne,
+ simplified pred_fun_def, simplified]
+
+lemmas disj_family_sum_with =
+ lt_disj_family_sum
+ [cancel_type_definition,
+ OF carrier_ne,
+ simplified pred_fun_def, simplified]
+
+lemmas sum_with_reindex_cong =
+ lt_sum_reindex_cong
+ [cancel_type_definition,
+ OF carrier_ne,
+ simplified pred_fun_def, simplified]
+
+end
+
+lemma (in comm_monoid_add_on_with) sum_with_cong':
+ shows "finite I \<Longrightarrow> (\<And>i. i\<in> I \<Longrightarrow> A i = B i) \<Longrightarrow> (\<And>i. i\<in> I \<Longrightarrow> A i \<in> S) \<Longrightarrow>
+ (\<And>i. i\<in> I \<Longrightarrow> B i \<in> S) \<Longrightarrow> sum_with pls z A I = sum_with pls z B I"
+proof (induct rule: finite_induct)
+ case empty
+ then show ?case by simp
+next
+ case (insert x F)
+ have "sum_with pls z A (insert x F) = pls (A x) (sum_with pls z A F)" using insert
+ sum_with_insert[of A] by (simp add: image_subset_iff)
+ also have "... = pls (B x) (sum_with pls z B F)" using insert by simp
+ also have "... = sum_with pls z B (insert x F)" using insert sum_with_insert[of B]
+ by (simp add: image_subset_iff)
+ finally show ?case .
+qed
+
+
+section \<open>Linear algebra complements\<close>
+
+subsection \<open>Additional properties of matrices\<close>
+
+lemma smult_one:
+ shows "(1::'a::monoid_mult) \<cdot>\<^sub>m A = A" by (simp add:eq_matI)
+
+lemma times_diag_index:
+ fixes A::"'a::comm_ring Matrix.mat"
+ assumes "A \<in> carrier_mat n n"
+and "B\<in> carrier_mat n n"
+and "diagonal_mat B"
+and "j < n"
+and "i < n"
+shows "Matrix.vec_index (Matrix.row (A*B) j) i = diag_mat B ! i *A $$ (j, i)"
+proof -
+ have "Matrix.vec_index (Matrix.row (A*B) j) i = (A*B) $$ (j,i)"
+ using Matrix.row_def[of "A*B" ] assms by simp
+ also have "... = Matrix.scalar_prod (Matrix.row A j) (Matrix.col B i)" using assms
+ times_mat_def[of A] by simp
+ also have "... = Matrix.scalar_prod (Matrix.col B i) (Matrix.row A j)"
+ using comm_scalar_prod[of "Matrix.row A j" n] assms by auto
+ also have "... = (Matrix.vec_index (Matrix.col B i) i) * (Matrix.vec_index (Matrix.row A j) i)"
+ unfolding Matrix.scalar_prod_def
+ proof (rule sum_but_one)
+ show "i < dim_vec (Matrix.row A j)" using assms by simp
+ show "\<forall>ia<dim_vec (Matrix.row A j). ia \<noteq> i \<longrightarrow> Matrix.vec_index (Matrix.col B i) ia = 0" using assms
+ by (metis carrier_matD(1) carrier_matD(2) diagonal_mat_def index_col index_row(2))
+ qed
+ also have "... = B $$(i,i) * (Matrix.vec_index (Matrix.row A j) i)" using assms by auto
+ also have "... = diag_mat B ! i * (Matrix.vec_index (Matrix.row A j) i)" unfolding diag_mat_def
+ using assms by simp
+ also have "... = diag_mat B ! i * A $$ (j, i)" using assms by simp
+ finally show ?thesis .
+qed
+
+lemma inner_prod_adjoint_comp:
+ assumes "(U::'a::conjugatable_field Matrix.mat) \<in> carrier_mat n n"
+and "(V::'a::conjugatable_field Matrix.mat) \<in> carrier_mat n n"
+and "i < n"
+and "j < n"
+shows "Complex_Matrix.inner_prod (Matrix.col V i) (Matrix.col U j) =
+ ((Complex_Matrix.adjoint V) * U) $$ (i, j)"
+proof -
+ have "Complex_Matrix.inner_prod (Matrix.col V i) (Matrix.col U j) =
+ Matrix.scalar_prod (Matrix.col U j) (Matrix.row (Complex_Matrix.adjoint V) i)"
+ using adjoint_row[of i V] assms by simp
+ also have "... = Matrix.scalar_prod (Matrix.row (Complex_Matrix.adjoint V) i) (Matrix.col U j)"
+ by (metis adjoint_row assms(1) assms(2) assms(3) carrier_matD(1) carrier_matD(2) Matrix.col_dim
+ conjugate_vec_sprod_comm)
+ also have "... = ((Complex_Matrix.adjoint V) * U) $$ (i, j)" using assms
+ by (simp add:times_mat_def)
+ finally show ?thesis .
+qed
+
+lemma mat_unit_vec_col:
+ assumes "(A::'a::conjugatable_field Matrix.mat) \<in> carrier_mat n n"
+and "i < n"
+shows "A *\<^sub>v (unit_vec n i) = Matrix.col A i"
+proof
+ show "dim_vec (A *\<^sub>v unit_vec n i) = dim_vec (Matrix.col A i)" using assms by simp
+ show "\<And>j. j < dim_vec (Matrix.col A i) \<Longrightarrow> Matrix.vec_index (A *\<^sub>v unit_vec n i) j =
+ Matrix.vec_index (Matrix.col A i) j"
+ proof -
+ fix j
+ assume "j < dim_vec (Matrix.col A i)"
+ hence "Matrix.vec_index (A *\<^sub>v unit_vec n i) j =
+ Matrix.scalar_prod (Matrix.row A j) (unit_vec n i)" unfolding mult_mat_vec_def by simp
+ also have "... = Matrix.scalar_prod (unit_vec n i) (Matrix.row A j)" using comm_scalar_prod
+ assms by auto
+ also have "... = (Matrix.vec_index (unit_vec n i) i) * (Matrix.vec_index (Matrix.row A j) i)"
+ unfolding Matrix.scalar_prod_def
+ proof (rule sum_but_one)
+ show "i < dim_vec (Matrix.row A j)" using assms by auto
+ show "\<forall>ia<dim_vec (Matrix.row A j). ia \<noteq> i \<longrightarrow> Matrix.vec_index (unit_vec n i) ia = 0"
+ using assms unfolding unit_vec_def by auto
+ qed
+ also have "... = (Matrix.vec_index (Matrix.row A j) i)" using assms by simp
+ also have "... = A $$ (j, i)" using assms \<open>j < dim_vec (Matrix.col A i)\<close> by simp
+ also have "... = Matrix.vec_index (Matrix.col A i) j" using assms \<open>j < dim_vec (Matrix.col A i)\<close> by simp
+ finally show "Matrix.vec_index (A *\<^sub>v unit_vec n i) j =
+ Matrix.vec_index (Matrix.col A i) j" .
+ qed
+qed
+
+lemma mat_prod_unit_vec_cong:
+ assumes "(A::'a::conjugatable_field Matrix.mat) \<in> carrier_mat n n"
+and "B\<in> carrier_mat n n"
+and "\<And>i. i < n \<Longrightarrow> A *\<^sub>v (unit_vec n i) = B *\<^sub>v (unit_vec n i)"
+shows "A = B"
+proof
+ show "dim_row A = dim_row B" using assms by simp
+ show "dim_col A = dim_col B" using assms by simp
+ show "\<And>i j. i < dim_row B \<Longrightarrow> j < dim_col B \<Longrightarrow> A $$ (i, j) = B $$ (i, j)"
+ proof -
+ fix i j
+ assume ij: "i < dim_row B" "j < dim_col B"
+ hence "A $$ (i,j) = Matrix.vec_index (Matrix.col A j) i" using assms by simp
+ also have "... = Matrix.vec_index (A *\<^sub>v (unit_vec n j)) i" using mat_unit_vec_col[of A] ij assms
+ by simp
+ also have "... = Matrix.vec_index (B *\<^sub>v (unit_vec n j)) i" using assms ij by simp
+ also have "... = Matrix.vec_index (Matrix.col B j) i" using mat_unit_vec_col ij assms by simp
+ also have "... = B $$ (i,j)" using assms ij by simp
+ finally show "A $$ (i, j) = B $$ (i, j)" .
+ qed
+qed
+
+lemma smult_smult_times:
+ fixes a::"'a::semigroup_mult"
+ shows "a\<cdot>\<^sub>m (k \<cdot>\<^sub>m A) = (a * k)\<cdot>\<^sub>m A"
+proof
+ show r:"dim_row (a \<cdot>\<^sub>m (k \<cdot>\<^sub>m A)) = dim_row (a * k \<cdot>\<^sub>m A)" by simp
+ show c:"dim_col (a \<cdot>\<^sub>m (k \<cdot>\<^sub>m A)) = dim_col (a * k \<cdot>\<^sub>m A)" by simp
+ show "\<And>i j. i < dim_row (a * k \<cdot>\<^sub>m A) \<Longrightarrow>
+ j < dim_col (a * k \<cdot>\<^sub>m A) \<Longrightarrow> (a \<cdot>\<^sub>m (k \<cdot>\<^sub>m A)) $$ (i, j) = (a * k \<cdot>\<^sub>m A) $$ (i, j)"
+ proof -
+ fix i j
+ assume "i < dim_row (a * k \<cdot>\<^sub>m A)" and "j < dim_col (a * k \<cdot>\<^sub>m A)" note ij = this
+ hence "(a \<cdot>\<^sub>m (k \<cdot>\<^sub>m A)) $$ (i, j) = a * (k \<cdot>\<^sub>m A) $$ (i, j)" by simp
+ also have "... = a * (k * A $$ (i,j))" using ij by simp
+ also have "... = (a * k) * A $$ (i,j)"
+ by (simp add: semigroup_mult_class.mult.assoc)
+ also have "... = (a * k \<cdot>\<^sub>m A) $$ (i, j)" using r c ij by simp
+ finally show "(a \<cdot>\<^sub>m (k \<cdot>\<^sub>m A)) $$ (i, j) = (a * k \<cdot>\<^sub>m A) $$ (i, j)" .
+ qed
+qed
+
+lemma mat_minus_minus:
+ fixes A :: "'a :: ab_group_add Matrix.mat"
+ assumes "A \<in> carrier_mat n m"
+ and "B\<in> carrier_mat n m"
+ and "C\<in> carrier_mat n m"
+shows "A - (B - C) = A - B + C"
+proof
+ show "dim_row (A - (B - C)) = dim_row (A - B + C)" using assms by simp
+ show "dim_col (A - (B - C)) = dim_col (A - B + C)" using assms by simp
+ show "\<And>i j. i < dim_row (A - B + C) \<Longrightarrow> j < dim_col (A - B + C) \<Longrightarrow>
+ (A - (B - C)) $$ (i, j) = (A - B + C) $$ (i, j)"
+ proof -
+ fix i j
+ assume "i < dim_row (A - B + C)" and "j < dim_col (A - B + C)" note ij = this
+ have "(A - (B - C)) $$ (i, j) = (A $$ (i,j) - B $$ (i,j) + C $$ (i,j))" using ij assms by simp
+ also have "... = (A - B + C) $$ (i, j)" using assms ij by simp
+ finally show "(A - (B - C)) $$ (i, j) = (A - B + C) $$ (i, j)" .
+ qed
+qed
+
+
+subsection \<open>Complements on complex matrices\<close>
+
+lemma hermitian_square:
+ assumes "hermitian M"
+ shows "M \<in> carrier_mat (dim_row M) (dim_row M)"
+proof -
+ have "dim_col M = dim_row M" using assms unfolding hermitian_def adjoint_def
+ by (metis adjoint_dim_col)
+ thus ?thesis by auto
+qed
+
+lemma hermitian_add:
+ assumes "A\<in> carrier_mat n n"
+ and "B\<in> carrier_mat n n"
+and "hermitian A"
+and "hermitian B"
+shows "hermitian (A + B)" unfolding hermitian_def
+ by (metis adjoint_add assms hermitian_def)
+
+lemma hermitian_minus:
+ assumes "A\<in> carrier_mat n n"
+ and "B\<in> carrier_mat n n"
+and "hermitian A"
+and "hermitian B"
+shows "hermitian (A - B)" unfolding hermitian_def
+ by (metis adjoint_minus assms hermitian_def)
+
+lemma hermitian_smult:
+ fixes a::real
+ fixes A::"complex Matrix.mat"
+ assumes "A \<in> carrier_mat n n"
+and "hermitian A"
+shows "hermitian (a \<cdot>\<^sub>m A)"
+proof -
+ have dim: "Complex_Matrix.adjoint A \<in> carrier_mat n n" using assms by (simp add: adjoint_dim)
+ {
+ fix i j
+ assume "i < n" and "j < n"
+ hence "Complex_Matrix.adjoint (a \<cdot>\<^sub>m A) $$ (i,j) = a * (Complex_Matrix.adjoint A $$ (i,j))"
+ using adjoint_scale[of a A] assms by simp
+ also have "... = a * (A $$ (i,j))" using assms unfolding hermitian_def by simp
+ also have "... = (a \<cdot>\<^sub>m A) $$ (i,j)" using \<open>i < n\<close> \<open>j < n\<close> assms by simp
+ finally have "Complex_Matrix.adjoint (a \<cdot>\<^sub>m A) $$ (i,j) = (a \<cdot>\<^sub>m A) $$ (i,j)" .
+ }
+ thus ?thesis using dim assms unfolding hermitian_def by auto
+qed
+
+lemma unitary_eigenvalues_norm_square:
+ fixes U::"complex Matrix.mat"
+ assumes "unitary U"
+ and "U \<in> carrier_mat n n"
+ and "eigenvalue U k"
+shows "conjugate k * k = 1"
+proof -
+ have "\<exists>v. eigenvector U v k" using assms unfolding eigenvalue_def by simp
+ from this obtain v where "eigenvector U v k" by auto
+ define vn where "vn = vec_normalize v"
+ have "eigenvector U vn k" using normalize_keep_eigenvector \<open>eigenvector U v k\<close>
+ using assms(2) eigenvector_def vn_def by blast
+ have "vn \<in> carrier_vec n"
+ using \<open>eigenvector U v k\<close> assms(2) eigenvector_def normalized_vec_dim vn_def by blast
+ have "Complex_Matrix.inner_prod vn vn = 1" using \<open>vn = vec_normalize v\<close> \<open>eigenvector U v k\<close>
+ eigenvector_def normalized_vec_norm by blast
+ hence "conjugate k * k = conjugate k * k * Complex_Matrix.inner_prod vn vn" by simp
+ also have "... = conjugate k * Complex_Matrix.inner_prod vn (k \<cdot>\<^sub>v vn)"
+ proof -
+ have "k * Complex_Matrix.inner_prod vn vn = Complex_Matrix.inner_prod vn (k \<cdot>\<^sub>v vn)"
+ using inner_prod_smult_left[of vn n vn k] \<open>vn \<in> carrier_vec n\<close> by simp
+ thus ?thesis by simp
+ qed
+ also have "... = Complex_Matrix.inner_prod (k \<cdot>\<^sub>v vn) (k \<cdot>\<^sub>v vn)"
+ using inner_prod_smult_right[of vn n _ k] by (simp add: \<open>vn \<in> carrier_vec n\<close>)
+ also have "... = Complex_Matrix.inner_prod (U *\<^sub>v vn) (U *\<^sub>v vn)"
+ using \<open>eigenvector U vn k\<close> unfolding eigenvector_def by simp
+ also have "... =
+ Complex_Matrix.inner_prod (Complex_Matrix.adjoint U *\<^sub>v (U *\<^sub>v vn)) vn"
+ using adjoint_def_alter[of "U *\<^sub>v vn" n vn n U] assms
+ by (metis \<open>eigenvector U vn k\<close> carrier_matD(1) carrier_vec_dim_vec dim_mult_mat_vec
+ eigenvector_def)
+ also have "... = Complex_Matrix.inner_prod vn vn"
+ proof -
+ have "Complex_Matrix.adjoint U *\<^sub>v (U *\<^sub>v vn) = (Complex_Matrix.adjoint U * U) *\<^sub>v vn"
+ using assms
+ by (metis \<open>eigenvector U vn k\<close> adjoint_dim assoc_mult_mat_vec carrier_matD(1) eigenvector_def)
+ also have "... = vn" using assms unfolding unitary_def inverts_mat_def
+ by (metis \<open>eigenvector U vn k\<close> assms(1) eigenvector_def one_mult_mat_vec unitary_simps(1))
+ finally show ?thesis by simp
+ qed
+ also have "... = 1" using \<open>vn = vec_normalize v\<close> \<open>eigenvector U v k\<close> eigenvector_def
+ normalized_vec_norm by blast
+ finally show ?thesis .
+qed
+
+lemma outer_prod_smult_left:
+ fixes v::"complex Matrix.vec"
+ shows "outer_prod (a \<cdot>\<^sub>v v) w = a \<cdot>\<^sub>m outer_prod v w"
+proof -
+ define paw where "paw = outer_prod (a \<cdot>\<^sub>v v) w"
+ define apw where "apw = a \<cdot>\<^sub>m outer_prod v w"
+ have "paw = apw"
+ proof
+ have "dim_row paw = dim_vec v" unfolding paw_def using outer_prod_dim
+ by (metis carrier_matD(1) carrier_vec_dim_vec index_smult_vec(2))
+ also have "... = dim_row apw" unfolding apw_def using outer_prod_dim
+ by (metis carrier_matD(1) carrier_vec_dim_vec index_smult_mat(2))
+ finally show dr: "dim_row paw = dim_row apw" .
+ have "dim_col paw = dim_vec w" unfolding paw_def using outer_prod_dim
+ using carrier_vec_dim_vec by blast
+ also have "... = dim_col apw" unfolding apw_def using outer_prod_dim
+ by (metis apw_def carrier_matD(2) carrier_vec_dim_vec smult_carrier_mat)
+ finally show dc: "dim_col paw = dim_col apw" .
+ show "\<And>i j. i < dim_row apw \<Longrightarrow> j < dim_col apw \<Longrightarrow> paw $$ (i, j) = apw $$ (i, j)"
+ proof -
+ fix i j
+ assume "i < dim_row apw" and "j < dim_col apw" note ij = this
+ hence "paw $$ (i,j) = a * (Matrix.vec_index v i) * cnj (Matrix.vec_index w j)"
+ using dr dc unfolding paw_def outer_prod_def by simp
+ also have "... = apw $$ (i,j)" using dr dc ij unfolding apw_def outer_prod_def by simp
+ finally show "paw $$ (i, j) = apw $$ (i, j)" .
+ qed
+ qed
+ thus ?thesis unfolding paw_def apw_def by simp
+qed
+
+lemma outer_prod_smult_right:
+ fixes v::"complex Matrix.vec"
+ shows "outer_prod v (a \<cdot>\<^sub>v w) = (conjugate a) \<cdot>\<^sub>m outer_prod v w"
+proof -
+ define paw where "paw = outer_prod v (a \<cdot>\<^sub>v w)"
+ define apw where "apw = (conjugate a) \<cdot>\<^sub>m outer_prod v w"
+ have "paw = apw"
+ proof
+ have "dim_row paw = dim_vec v" unfolding paw_def using outer_prod_dim
+ by (metis carrier_matD(1) carrier_vec_dim_vec)
+ also have "... = dim_row apw" unfolding apw_def using outer_prod_dim
+ by (metis carrier_matD(1) carrier_vec_dim_vec index_smult_mat(2))
+ finally show dr: "dim_row paw = dim_row apw" .
+ have "dim_col paw = dim_vec w" unfolding paw_def using outer_prod_dim
+ using carrier_vec_dim_vec by (metis carrier_matD(2) index_smult_vec(2))
+ also have "... = dim_col apw" unfolding apw_def using outer_prod_dim
+ by (metis apw_def carrier_matD(2) carrier_vec_dim_vec smult_carrier_mat)
+ finally show dc: "dim_col paw = dim_col apw" .
+ show "\<And>i j. i < dim_row apw \<Longrightarrow> j < dim_col apw \<Longrightarrow> paw $$ (i, j) = apw $$ (i, j)"
+ proof -
+ fix i j
+ assume "i < dim_row apw" and "j < dim_col apw" note ij = this
+ hence "paw $$ (i,j) = (conjugate a) * (Matrix.vec_index v i) * cnj (Matrix.vec_index w j)"
+ using dr dc unfolding paw_def outer_prod_def by simp
+ also have "... = apw $$ (i,j)" using dr dc ij unfolding apw_def outer_prod_def by simp
+ finally show "paw $$ (i, j) = apw $$ (i, j)" .
+ qed
+ qed
+ thus ?thesis unfolding paw_def apw_def by simp
+qed
+
+lemma outer_prod_add_left:
+ fixes v::"complex Matrix.vec"
+ assumes "dim_vec v = dim_vec x"
+ shows "outer_prod (v + x) w = outer_prod v w + (outer_prod x w)"
+proof -
+ define paw where "paw = outer_prod (v+x) w"
+ define apw where "apw = outer_prod v w + (outer_prod x w)"
+ have "paw = apw"
+ proof
+ have rv: "dim_row paw = dim_vec v" unfolding paw_def using outer_prod_dim assms
+ by (metis carrier_matD(1) carrier_vec_dim_vec index_add_vec(2) paw_def)
+ also have "... = dim_row apw" unfolding apw_def using outer_prod_dim assms
+ by (metis carrier_matD(1) carrier_vec_dim_vec index_add_mat(2))
+ finally show dr: "dim_row paw = dim_row apw" .
+ have cw: "dim_col paw = dim_vec w" unfolding paw_def using outer_prod_dim assms
+ using carrier_vec_dim_vec by (metis carrier_matD(2))
+ also have "... = dim_col apw" unfolding apw_def using outer_prod_dim
+ by (metis apw_def carrier_matD(2) carrier_vec_dim_vec add_carrier_mat)
+ finally show dc: "dim_col paw = dim_col apw" .
+ show "\<And>i j. i < dim_row apw \<Longrightarrow> j < dim_col apw \<Longrightarrow> paw $$ (i, j) = apw $$ (i, j)"
+ proof -
+ fix i j
+ assume "i < dim_row apw" and "j < dim_col apw" note ij = this
+ hence "paw $$ (i,j) = (Matrix.vec_index v i + Matrix.vec_index x i) *
+ cnj (Matrix.vec_index w j)"
+ using dr dc unfolding paw_def outer_prod_def by simp
+ also have "... = Matrix.vec_index v i * cnj (Matrix.vec_index w j) +
+ Matrix.vec_index x i * cnj (Matrix.vec_index w j)"
+ by (simp add: ring_class.ring_distribs(2))
+ also have "... = (outer_prod v w) $$ (i,j) + (outer_prod x w) $$ (i,j)"
+ using rv cw dr dc ij assms unfolding outer_prod_def by auto
+ also have "... = apw $$ (i,j)" using dr dc ij unfolding apw_def outer_prod_def by simp
+ finally show "paw $$ (i, j) = apw $$ (i, j)" .
+ qed
+ qed
+ thus ?thesis unfolding paw_def apw_def by simp
+qed
+
+lemma outer_prod_add_right:
+ fixes v::"complex Matrix.vec"
+ assumes "dim_vec w = dim_vec x"
+ shows "outer_prod v (w + x) = outer_prod v w + (outer_prod v x)"
+proof -
+ define paw where "paw = outer_prod v (w+x)"
+ define apw where "apw = outer_prod v w + (outer_prod v x)"
+ have "paw = apw"
+ proof
+ have rv: "dim_row paw = dim_vec v" unfolding paw_def using outer_prod_dim assms
+ by (metis carrier_matD(1) carrier_vec_dim_vec index_add_vec(2) paw_def)
+ also have "... = dim_row apw" unfolding apw_def using outer_prod_dim assms
+ by (metis carrier_matD(1) carrier_vec_dim_vec index_add_mat(2))
+ finally show dr: "dim_row paw = dim_row apw" .
+ have cw: "dim_col paw = dim_vec w" unfolding paw_def using outer_prod_dim assms
+ using carrier_vec_dim_vec
+ by (metis carrier_matD(2) index_add_vec(2) paw_def)
+ also have "... = dim_col apw" unfolding apw_def using outer_prod_dim
+ by (metis assms carrier_matD(2) carrier_vec_dim_vec index_add_mat(3))
+ finally show dc: "dim_col paw = dim_col apw" .
+ show "\<And>i j. i < dim_row apw \<Longrightarrow> j < dim_col apw \<Longrightarrow> paw $$ (i, j) = apw $$ (i, j)"
+ proof -
+ fix i j
+ assume "i < dim_row apw" and "j < dim_col apw" note ij = this
+ hence "paw $$ (i,j) = Matrix.vec_index v i *
+ (cnj (Matrix.vec_index w j + (Matrix.vec_index x j)))"
+ using dr dc unfolding paw_def outer_prod_def by simp
+ also have "... = Matrix.vec_index v i * cnj (Matrix.vec_index w j) +
+ Matrix.vec_index v i * cnj (Matrix.vec_index x j)"
+ by (simp add: ring_class.ring_distribs(1))
+ also have "... = (outer_prod v w) $$ (i,j) + (outer_prod v x) $$ (i,j)"
+ using rv cw dr dc ij assms unfolding outer_prod_def by auto
+ also have "... = apw $$ (i,j)" using dr dc ij unfolding apw_def outer_prod_def by simp
+ finally show "paw $$ (i, j) = apw $$ (i, j)" .
+ qed
+ qed
+ thus ?thesis unfolding paw_def apw_def by simp
+qed
+
+lemma outer_prod_minus_left:
+ fixes v::"complex Matrix.vec"
+ assumes "dim_vec v = dim_vec x"
+ shows "outer_prod (v - x) w = outer_prod v w - (outer_prod x w)"
+proof -
+ define paw where "paw = outer_prod (v-x) w"
+ define apw where "apw = outer_prod v w - (outer_prod x w)"
+ have "paw = apw"
+ proof
+ have rv: "dim_row paw = dim_vec v" unfolding paw_def using outer_prod_dim assms
+ by (metis carrier_matD(1) carrier_vec_dim_vec index_minus_vec(2) paw_def)
+ also have "... = dim_row apw" unfolding apw_def using outer_prod_dim assms
+ by (metis carrier_matD(1) carrier_vec_dim_vec index_minus_mat(2))
+ finally show dr: "dim_row paw = dim_row apw" .
+ have cw: "dim_col paw = dim_vec w" unfolding paw_def using outer_prod_dim assms
+ using carrier_vec_dim_vec by (metis carrier_matD(2))
+ also have "... = dim_col apw" unfolding apw_def using outer_prod_dim
+ by (metis apw_def carrier_matD(2) carrier_vec_dim_vec minus_carrier_mat)
+ finally show dc: "dim_col paw = dim_col apw" .
+ show "\<And>i j. i < dim_row apw \<Longrightarrow> j < dim_col apw \<Longrightarrow> paw $$ (i, j) = apw $$ (i, j)"
+ proof -
+ fix i j
+ assume "i < dim_row apw" and "j < dim_col apw" note ij = this
+ hence "paw $$ (i,j) = (Matrix.vec_index v i - Matrix.vec_index x i) *
+ cnj (Matrix.vec_index w j)"
+ using dr dc unfolding paw_def outer_prod_def by simp
+ also have "... = Matrix.vec_index v i * cnj (Matrix.vec_index w j) -
+ Matrix.vec_index x i * cnj (Matrix.vec_index w j)"
+ by (simp add: ring_class.ring_distribs)
+ also have "... = (outer_prod v w) $$ (i,j) - (outer_prod x w) $$ (i,j)"
+ using rv cw dr dc ij assms unfolding outer_prod_def by auto
+ also have "... = apw $$ (i,j)" using dr dc ij unfolding apw_def outer_prod_def by simp
+ finally show "paw $$ (i, j) = apw $$ (i, j)" .
+ qed
+ qed
+ thus ?thesis unfolding paw_def apw_def by simp
+qed
+
+lemma outer_prod_minus_right:
+ fixes v::"complex Matrix.vec"
+ assumes "dim_vec w = dim_vec x"
+ shows "outer_prod v (w - x) = outer_prod v w - (outer_prod v x)"
+proof -
+ define paw where "paw = outer_prod v (w-x)"
+ define apw where "apw = outer_prod v w - (outer_prod v x)"
+ have "paw = apw"
+ proof
+ have rv: "dim_row paw = dim_vec v" unfolding paw_def using outer_prod_dim assms
+ by (metis carrier_matD(1) carrier_vec_dim_vec paw_def)
+ also have "... = dim_row apw" unfolding apw_def using outer_prod_dim assms
+ by (metis carrier_matD(1) carrier_vec_dim_vec index_minus_mat(2))
+ finally show dr: "dim_row paw = dim_row apw" .
+ have cw: "dim_col paw = dim_vec w" unfolding paw_def using outer_prod_dim assms
+ using carrier_vec_dim_vec
+ by (metis carrier_matD(2) index_minus_vec(2) paw_def)
+ also have "... = dim_col apw" unfolding apw_def using outer_prod_dim
+ by (metis assms carrier_matD(2) carrier_vec_dim_vec index_minus_mat(3))
+ finally show dc: "dim_col paw = dim_col apw" .
+ show "\<And>i j. i < dim_row apw \<Longrightarrow> j < dim_col apw \<Longrightarrow> paw $$ (i, j) = apw $$ (i, j)"
+ proof -
+ fix i j
+ assume "i < dim_row apw" and "j < dim_col apw" note ij = this
+ hence "paw $$ (i,j) = Matrix.vec_index v i *
+ (cnj (Matrix.vec_index w j - (Matrix.vec_index x j)))"
+ using dr dc unfolding paw_def outer_prod_def by simp
+ also have "... = Matrix.vec_index v i * cnj (Matrix.vec_index w j) -
+ Matrix.vec_index v i * cnj (Matrix.vec_index x j)"
+ by (simp add: ring_class.ring_distribs)
+ also have "... = (outer_prod v w) $$ (i,j) - (outer_prod v x) $$ (i,j)"
+ using rv cw dr dc ij assms unfolding outer_prod_def by auto
+ also have "... = apw $$ (i,j)" using dr dc ij unfolding apw_def outer_prod_def by simp
+ finally show "paw $$ (i, j) = apw $$ (i, j)" .
+ qed
+ qed
+ thus ?thesis unfolding paw_def apw_def by simp
+qed
+
+lemma outer_minus_minus:
+ fixes a::"complex Matrix.vec"
+ assumes "dim_vec a = dim_vec b"
+ and "dim_vec u = dim_vec v"
+ shows "outer_prod (a - b) (u - v) = outer_prod a u - outer_prod a v -
+ outer_prod b u + outer_prod b v"
+proof -
+ have "outer_prod (a - b) (u - v) = outer_prod a (u - v)
+ - outer_prod b (u - v)" using outer_prod_minus_left assms by simp
+ also have "... = outer_prod a u - outer_prod a v -
+ outer_prod b (u - v)" using assms outer_prod_minus_right by simp
+ also have "... = outer_prod a u - outer_prod a v -
+ (outer_prod b u - outer_prod b v)" using assms outer_prod_minus_right by simp
+ also have "... = outer_prod a u - outer_prod a v -
+ outer_prod b u + outer_prod b v"
+ proof (rule mat_minus_minus)
+ show "outer_prod b u \<in> carrier_mat (dim_vec b) (dim_vec u)" by simp
+ show "outer_prod b v \<in> carrier_mat (dim_vec b) (dim_vec u)" using assms by simp
+ show "outer_prod a u - outer_prod a v \<in> carrier_mat (dim_vec b) (dim_vec u)" using assms
+ by (metis carrier_vecI minus_carrier_mat outer_prod_dim)
+ qed
+ finally show ?thesis .
+qed
+
+lemma outer_trace_inner:
+ assumes "A \<in> carrier_mat n n"
+ and "dim_vec u = n"
+and "dim_vec v = n"
+ shows "Complex_Matrix.trace (outer_prod u v * A) = Complex_Matrix.inner_prod v (A *\<^sub>v u)"
+proof -
+ have "Complex_Matrix.trace (outer_prod u v * A) = Complex_Matrix.trace (A * outer_prod u v)"
+ proof (rule trace_comm)
+ show "A \<in> carrier_mat n n" using assms by simp
+ show "outer_prod u v \<in> carrier_mat n n" using assms
+ by (metis carrier_vec_dim_vec outer_prod_dim)
+ qed
+ also have "... = Complex_Matrix.inner_prod v (A *\<^sub>v u)" using trace_outer_prod_right[of A n u v]
+ assms carrier_vec_dim_vec by metis
+ finally show ?thesis .
+qed
+
+lemma zero_hermitian:
+ shows "hermitian (0\<^sub>m n n)" unfolding hermitian_def
+ by (metis adjoint_minus hermitian_def hermitian_one minus_r_inv_mat one_carrier_mat)
+
+lemma trace_1:
+ shows "Complex_Matrix.trace ((1\<^sub>m n)::complex Matrix.mat) =(n::complex)" using one_mat_def
+ by (simp add: Complex_Matrix.trace_def Matrix.mat_def)
+
+lemma trace_add:
+ assumes "square_mat A"
+ and "square_mat B"
+ and "dim_row A = dim_row B"
+ shows "Complex_Matrix.trace (A + B) = Complex_Matrix.trace A + Complex_Matrix.trace B"
+ using assms by (simp add: Complex_Matrix.trace_def sum.distrib)
+
+lemma bra_vec_carrier:
+ shows "bra_vec v \<in> carrier_mat 1 (dim_vec v)"
+proof -
+ have "dim_row (ket_vec v) = dim_vec v" unfolding ket_vec_def by simp
+ thus ?thesis using bra_bra_vec[of v] bra_def[of "ket_vec v"] by simp
+qed
+
+lemma mat_mult_ket_carrier:
+ assumes "A\<in> carrier_mat n m"
+shows "A * |v\<rangle> \<in> carrier_mat n 1" using assms
+ by (metis bra_bra_vec bra_vec_carrier carrier_matD(1) carrier_matI dagger_of_ket_is_bra
+ dim_row_of_dagger index_mult_mat(2) index_mult_mat(3))
+
+lemma mat_mult_ket:
+ assumes "A \<in> carrier_mat n m"
+and "dim_vec v = m"
+shows "A * |v\<rangle> = |A *\<^sub>v v\<rangle>"
+proof -
+ have rn: "dim_row (A * |v\<rangle>) = n" unfolding times_mat_def using assms by simp
+ have co: "dim_col |A *\<^sub>v v\<rangle> = 1" using assms unfolding ket_vec_def by simp
+ have cov: "dim_col |v\<rangle> = 1" using assms unfolding ket_vec_def by simp
+ have er: "dim_row (A * |v\<rangle>) = dim_row |A *\<^sub>v v\<rangle>" using assms
+ by (metis bra_bra_vec bra_vec_carrier carrier_matD(2) dagger_of_ket_is_bra dim_col_of_dagger
+ dim_mult_mat_vec index_mult_mat(2))
+ have ec: "dim_col (A * |v\<rangle>) = dim_col |A *\<^sub>v v\<rangle>" using assms
+ by (metis carrier_matD(2) index_mult_mat(3) mat_mult_ket_carrier)
+ {
+ fix i::nat
+ fix j::nat
+ assume "i < n"
+ and "j < 1"
+ hence "j = 0" by simp
+ have "(A * |v\<rangle>) $$ (i,0) = Matrix.scalar_prod (Matrix.row A i) (Matrix.col |v\<rangle> 0)"
+ using times_mat_def[of A] \<open>i < n\<close> rn cov by simp
+ also have "... = Matrix.scalar_prod (Matrix.row A i) v" using ket_vec_col by simp
+ also have "... = |A *\<^sub>v v\<rangle> $$ (i,j)" unfolding mult_mat_vec_def
+ using \<open>i < n\<close> \<open>j = 0\<close> assms(1) by auto
+ } note idx = this
+ have "A * |v\<rangle> = Matrix.mat n 1 (\<lambda>(i, j). Matrix.scalar_prod (Matrix.row A i) (Matrix.col |v\<rangle> j))"
+ using assms unfolding times_mat_def ket_vec_def by simp
+ also have "... = |A *\<^sub>v v\<rangle>" using er ec idx rn co by auto
+ finally show ?thesis .
+qed
+
+lemma unitary_density:
+ assumes "density_operator R"
+ and "unitary U"
+ and "R\<in> carrier_mat n n"
+ and "U\<in> carrier_mat n n"
+shows "density_operator (U * R * (Complex_Matrix.adjoint U))" unfolding density_operator_def
+proof (intro conjI)
+ show "Complex_Matrix.positive (U * R * Complex_Matrix.adjoint U)"
+ proof (rule positive_close_under_left_right_mult_adjoint)
+ show "U \<in> carrier_mat n n" using assms by simp
+ show "R \<in> carrier_mat n n" using assms by simp
+ show "Complex_Matrix.positive R" using assms unfolding density_operator_def by simp
+ qed
+ have "Complex_Matrix.trace (U * R * Complex_Matrix.adjoint U) =
+ Complex_Matrix.trace (Complex_Matrix.adjoint U * U * R)"
+ using trace_comm[of "U * R" n "Complex_Matrix.adjoint U"] assms
+ by (metis adjoint_dim mat_assoc_test(10))
+ also have "... = Complex_Matrix.trace R" using assms by simp
+ also have "... = 1" using assms unfolding density_operator_def by simp
+ finally show "Complex_Matrix.trace (U * R * Complex_Matrix.adjoint U) = 1" .
+qed
+
+
+subsection \<open>Tensor product complements\<close>
+
+lemma tensor_vec_dim[simp]:
+ shows "dim_vec (tensor_vec u v) = dim_vec u * (dim_vec v)"
+proof -
+ have "length (mult.vec_vec_Tensor (*) (list_of_vec u) (list_of_vec v)) =
+ length (list_of_vec u) * length (list_of_vec v)"
+ using mult.vec_vec_Tensor_length[of "1::real" "(*)" "list_of_vec u" "list_of_vec v"]
+ by (simp add: Matrix_Tensor.mult_def)
+ thus ?thesis unfolding tensor_vec_def by simp
+qed
+
+lemma index_tensor_vec[simp]:
+ assumes "0 < dim_vec v"
+ and "i < dim_vec u * dim_vec v"
+shows "vec_index (tensor_vec u v) i =
+ vec_index u (i div (dim_vec v)) * vec_index v (i mod dim_vec v)"
+proof -
+ have m: "Matrix_Tensor.mult (1::complex) (*)" by (simp add: Matrix_Tensor.mult_def)
+ have "length (list_of_vec v) = dim_vec v" using assms by simp
+ hence "vec_index (tensor_vec u v) i = (*) (vec_index u (i div dim_vec v)) (vec_index v (i mod dim_vec v))"
+ unfolding tensor_vec_def using mult.vec_vec_Tensor_elements assms m
+ by (metis (mono_tags, lifting) length_greater_0_conv length_list_of_vec list_of_vec_index
+ mult.vec_vec_Tensor_elements vec_of_list_index)
+ thus ?thesis by simp
+qed
+
+lemma outer_prod_tensor_comm:
+ fixes a::"complex Matrix.vec"
+ fixes u::"complex Matrix.vec"
+ assumes "0 < dim_vec a"
+ and "0 < dim_vec b"
+shows "outer_prod (tensor_vec u v) (tensor_vec a b) = tensor_mat (outer_prod u a) (outer_prod v b)"
+proof -
+ define ot where "ot = outer_prod (tensor_vec u v) (tensor_vec a b)"
+ define to where "to = tensor_mat (outer_prod u a) (outer_prod v b)"
+ define dv where "dv = dim_vec v"
+ define db where "db = dim_vec b"
+ have "ot = to"
+ proof
+ have ro: "dim_row ot = dim_vec u * dim_vec v" unfolding ot_def outer_prod_def by simp
+ have "dim_row to = dim_row (outer_prod u a) * dim_row (outer_prod v b)"
+ unfolding to_def by simp
+ also have "... = dim_vec u * dim_vec v" using outer_prod_dim
+ by (metis carrier_matD(1) carrier_vec_dim_vec)
+ finally have rt: "dim_row to = dim_vec u * dim_vec v" .
+ show "dim_row ot = dim_row to" using ro rt by simp
+ have co: "dim_col ot = dim_vec a * dim_vec b" unfolding ot_def outer_prod_def by simp
+ have "dim_col to = dim_col (outer_prod u a) * dim_col (outer_prod v b)" unfolding to_def by simp
+ also have "... = dim_vec a * dim_vec b" using outer_prod_dim
+ by (metis carrier_matD(2) carrier_vec_dim_vec)
+ finally have ct: "dim_col to = dim_vec a * dim_vec b" .
+ show "dim_col ot = dim_col to" using co ct by simp
+ show "\<And>i j. i < dim_row to \<Longrightarrow> j < dim_col to \<Longrightarrow> ot $$ (i, j) = to $$ (i, j)"
+ proof -
+ fix i j
+ assume "i < dim_row to" and "j < dim_col to" note ij = this
+ have "ot $$ (i,j) = Matrix.vec_index (tensor_vec u v) i *
+ (conjugate (Matrix.vec_index (tensor_vec a b) j))"
+ unfolding ot_def outer_prod_def using ij rt ct by simp
+ also have "... = vec_index u (i div dv) * vec_index v (i mod dv) *
+ (conjugate (Matrix.vec_index (tensor_vec a b) j))" using ij rt assms
+ unfolding dv_def
+ by (metis index_tensor_vec less_nat_zero_code nat_0_less_mult_iff neq0_conv)
+ also have "... = vec_index u (i div dv) * vec_index v (i mod dv) *
+ (conjugate (vec_index a (j div db) * vec_index b (j mod db)))" using ij ct assms
+ unfolding db_def by simp
+ also have "... = vec_index u (i div dv) * vec_index v (i mod dv) *
+ (conjugate (vec_index a (j div db))) * (conjugate (vec_index b (j mod db)))" by simp
+ also have "... = vec_index u (i div dv) * (conjugate (vec_index a (j div db))) *
+ vec_index v (i mod dv) * (conjugate (vec_index b (j mod db)))" by simp
+ also have "... = (outer_prod u a) $$ (i div dv, j div db) *
+ vec_index v (i mod dv) * (conjugate (vec_index b (j mod db)))"
+ proof -
+ have "i div dv < dim_vec u" using ij rt unfolding dv_def
+ by (simp add: less_mult_imp_div_less)
+ moreover have "j div db < dim_vec a" using ij ct assms unfolding db_def
+ by (simp add: less_mult_imp_div_less)
+ ultimately have "vec_index u (i div dv) * (conjugate (vec_index a (j div db))) =
+ (outer_prod u a) $$ (i div dv, j div db)" unfolding outer_prod_def by simp
+ thus ?thesis by simp
+ qed
+ also have "... = (outer_prod u a) $$ (i div dv, j div db) *
+ (outer_prod v b) $$ (i mod dv, j mod db)"
+ proof -
+ have "i mod dv < dim_vec v" using ij rt unfolding dv_def
+ using assms mod_less_divisor
+ by (metis less_nat_zero_code mult.commute neq0_conv times_nat.simps(1))
+ moreover have "j mod db < dim_vec b" using ij ct assms unfolding db_def
+ by (simp add: less_mult_imp_div_less)
+ ultimately have "vec_index v (i mod dv) * (conjugate (vec_index b (j mod db))) =
+ (outer_prod v b) $$ (i mod dv, j mod db)" unfolding outer_prod_def by simp
+ thus ?thesis by simp
+ qed
+ also have "... = tensor_mat (outer_prod u a) (outer_prod v b) $$ (i, j)"
+ proof (rule index_tensor_mat[symmetric])
+ show "dim_row (outer_prod u a) = dim_vec u" unfolding outer_prod_def by simp
+ show "dim_row (outer_prod v b) = dv" unfolding outer_prod_def dv_def by simp
+ show "dim_col (outer_prod v b) = db" unfolding db_def outer_prod_def by simp
+ show "i < dim_vec u * dv" unfolding dv_def using ij rt by simp
+ show "dim_col (outer_prod u a) = dim_vec a" unfolding outer_prod_def by simp
+ show "j < dim_vec a * db" unfolding db_def using ij ct by simp
+ show "0 < dim_vec a" using assms by simp
+ show "0 < db" unfolding db_def using assms by simp
+ qed
+ finally show "ot $$ (i, j) = to $$ (i, j)" unfolding to_def .
+ qed
+ qed
+ thus ?thesis unfolding ot_def to_def by simp
+qed
+
+lemma tensor_mat_adjoint:
+ assumes "m1 \<in> carrier_mat r1 c1"
+ and "m2 \<in> carrier_mat r2 c2"
+ and "0 < c1"
+ and "0 < c2"
+and "0 < r1"
+and "0 < r2"
+ shows "Complex_Matrix.adjoint (tensor_mat m1 m2) =
+ tensor_mat (Complex_Matrix.adjoint m1) (Complex_Matrix.adjoint m2)"
+ apply (rule eq_matI, auto)
+proof -
+ fix i j
+ assume "i < dim_col m1 * dim_col m2" and "j < dim_row m1 * dim_row m2" note ij = this
+ have c1: "dim_col m1 = c1" using assms by simp
+ have r1: "dim_row m1 = r1" using assms by simp
+ have c2: "dim_col m2 = c2" using assms by simp
+ have r2: "dim_row m2 = r2" using assms by simp
+ have "Complex_Matrix.adjoint (m1 \<Otimes> m2) $$ (i, j) = conjugate ((m1 \<Otimes> m2) $$ (j, i))"
+ using ij by (simp add: adjoint_eval)
+ also have "... = conjugate (m1 $$ (j div r2, i div c2) * m2 $$ (j mod r2, i mod c2))"
+ proof -
+ have "(m1 \<Otimes> m2) $$ (j, i) = m1 $$ (j div r2, i div c2) * m2 $$ (j mod r2, i mod c2)"
+ proof (rule index_tensor_mat[of m1 r1 c1 m2 r2 c2 j i], (auto simp add: assms ij c1 c2 r1 r2))
+ show "j < r1 * r2" using ij r1 r2 by simp
+ show "i < c1 * c2" using ij c1 c2 by simp
+ qed
+ thus ?thesis by simp
+ qed
+ also have "... = conjugate (m1 $$ (j div r2, i div c2)) * conjugate ( m2 $$ (j mod r2, i mod c2))"
+ by simp
+ also have "... = (Complex_Matrix.adjoint m1) $$ (i div c2, j div r2) *
+ conjugate ( m2 $$ (j mod r2, i mod c2))"
+ by (metis adjoint_eval c2 ij less_mult_imp_div_less r2)
+ also have "... = (Complex_Matrix.adjoint m1) $$ (i div c2, j div r2) *
+ (Complex_Matrix.adjoint m2) $$ (i mod c2, j mod r2)"
+ by (metis Euclidean_Division.div_eq_0_iff adjoint_eval assms(4) bits_mod_div_trivial c2
+ gr_implies_not_zero ij(2) mult_not_zero r2)
+ also have "... = (tensor_mat (Complex_Matrix.adjoint m1) (Complex_Matrix.adjoint m2)) $$ (i,j)"
+ proof (rule index_tensor_mat[symmetric], (simp add: ij c1 c2 r1 r2) +)
+ show "i < c1 * c2" using ij c1 c2 by simp
+ show "j < r1 * r2" using ij r1 r2 by simp
+ show "0 < r1" using assms by simp
+ show "0 < r2" using assms by simp
+ qed
+ finally show "Complex_Matrix.adjoint (m1 \<Otimes> m2) $$ (i, j) =
+ (Complex_Matrix.adjoint m1 \<Otimes> Complex_Matrix.adjoint m2) $$ (i, j)" .
+qed
+
+lemma index_tensor_mat':
+ assumes "0 < dim_col A"
+ and "0 < dim_col B"
+ and "i < dim_row A * dim_row B"
+ and "j < dim_col A * dim_col B"
+ shows "(A \<Otimes> B) $$ (i, j) =
+ A $$ (i div (dim_row B), j div (dim_col B)) * B $$ (i mod (dim_row B), j mod (dim_col B))"
+ by (rule index_tensor_mat, (simp add: assms)+)
+
+lemma tensor_mat_carrier:
+ shows "tensor_mat U V \<in> carrier_mat (dim_row U * dim_row V) (dim_col U * dim_col V)" by auto
+
+lemma tensor_mat_id:
+ assumes "0 < d1"
+ and "0 < d2"
+shows "tensor_mat (1\<^sub>m d1) (1\<^sub>m d2) = 1\<^sub>m (d1 * d2)"
+proof (rule eq_matI, auto)
+ show "tensor_mat (1\<^sub>m d1) (1\<^sub>m d2) $$ (i, i) = 1" if "i < (d1 * d2)" for i
+ using that index_tensor_mat'[of "1\<^sub>m d1" "1\<^sub>m d2"]
+ by (simp add: assms less_mult_imp_div_less)
+next
+ show "tensor_mat (1\<^sub>m d1) (1\<^sub>m d2) $$ (i, j) = 0" if "i < d1 * d2" "j < d1 * d2" "i \<noteq> j" for i j
+ using that index_tensor_mat[of "1\<^sub>m d1" d1 d1 "1\<^sub>m d2" d2 d2 i j]
+ by (metis assms(1) assms(2) index_one_mat(1) index_one_mat(2) index_one_mat(3)
+ less_mult_imp_div_less mod_less_divisor mult_div_mod_eq mult_not_zero)
+qed
+
+lemma tensor_mat_hermitian:
+ assumes "A \<in> carrier_mat n n"
+ and "B \<in> carrier_mat n' n'"
+ and "0 < n"
+ and "0 < n'"
+ and "hermitian A"
+ and "hermitian B"
+ shows "hermitian (tensor_mat A B)" using assms by (metis hermitian_def tensor_mat_adjoint)
+
+lemma tensor_mat_unitary:
+ assumes "Complex_Matrix.unitary U"
+ and "Complex_Matrix.unitary V"
+and "0 < dim_row U"
+and "0 < dim_row V"
+shows "Complex_Matrix.unitary (tensor_mat U V)"
+proof -
+ define UI where "UI = tensor_mat U V"
+ have "Complex_Matrix.adjoint UI =
+ tensor_mat (Complex_Matrix.adjoint U) (Complex_Matrix.adjoint V)" unfolding UI_def
+ proof (rule tensor_mat_adjoint)
+ show "U \<in> carrier_mat (dim_row U) (dim_row U)" using assms unfolding Complex_Matrix.unitary_def
+ by simp
+ show "V \<in> carrier_mat (dim_row V) (dim_row V)" using assms unfolding Complex_Matrix.unitary_def
+ by simp
+ show "0 < dim_row V" using assms by simp
+ show "0 < dim_row U" using assms by simp
+ show "0 < dim_row V" using assms by simp
+ show "0 < dim_row U" using assms by simp
+ qed
+ hence "UI * (Complex_Matrix.adjoint UI) =
+ tensor_mat (U * Complex_Matrix.adjoint U) (V * Complex_Matrix.adjoint V)"
+ using mult_distr_tensor[of U "Complex_Matrix.adjoint U" "V" "Complex_Matrix.adjoint V"]
+ unfolding UI_def
+ by (metis (no_types, lifting) Complex_Matrix.unitary_def adjoint_dim_col adjoint_dim_row
+ assms carrier_matD(2) )
+ also have "... = tensor_mat (1\<^sub>m (dim_row U)) (1\<^sub>m (dim_row V))" using assms unitary_simps(2)
+ by (metis Complex_Matrix.unitary_def)
+ also have "... = (1\<^sub>m (dim_row U * dim_row V))" using tensor_mat_id assms by simp
+ finally have "UI * (Complex_Matrix.adjoint UI) = (1\<^sub>m (dim_row U * dim_row V))" .
+ hence "inverts_mat UI (Complex_Matrix.adjoint UI)" unfolding inverts_mat_def UI_def by simp
+ thus ?thesis using assms unfolding Complex_Matrix.unitary_def UI_def
+ by (metis carrier_matD(2) carrier_matI dim_col_tensor_mat dim_row_tensor_mat)
+qed
+
+
+subsection \<open>Fixed carrier matrices locale\<close>
+
+text \<open>We define a locale for matrices with a fixed number of rows and columns, and define a
+finite sum operation on this locale. The \verb+Type_To_Sets+ transfer tools then permits to obtain
+lemmata on the locale for free. \<close>
+
+locale fixed_carrier_mat =
+ fixes fc_mats::"'a::field Matrix.mat set"
+ fixes dimR dimC
+ assumes fc_mats_carrier: "fc_mats = carrier_mat dimR dimC"
+begin
+
+sublocale semigroup_add_on_with fc_mats "(+)"
+proof (unfold_locales)
+ show "\<And>a b. a \<in> fc_mats \<Longrightarrow> b \<in> fc_mats \<Longrightarrow> a + b \<in> fc_mats" using fc_mats_carrier by simp
+ show "\<And>a b c. a \<in> fc_mats \<Longrightarrow> b \<in> fc_mats \<Longrightarrow> c \<in> fc_mats \<Longrightarrow> a + b + c = a + (b + c)"
+ using fc_mats_carrier by simp
+qed
+
+sublocale ab_semigroup_add_on_with fc_mats "(+)"
+proof (unfold_locales)
+ show "\<And>a b. a \<in> fc_mats \<Longrightarrow> b \<in> fc_mats \<Longrightarrow> a + b = b + a" using fc_mats_carrier
+ by (simp add: comm_add_mat)
+qed
+
+sublocale comm_monoid_add_on_with fc_mats "(+)" "0\<^sub>m dimR dimC"
+proof (unfold_locales)
+ show "0\<^sub>m dimR dimC \<in> fc_mats" using fc_mats_carrier by simp
+ show "\<And>a. a \<in> fc_mats \<Longrightarrow> 0\<^sub>m dimR dimC + a = a" using fc_mats_carrier by simp
+qed
+
+sublocale ab_group_add_on_with fc_mats "(+)" "0\<^sub>m dimR dimC" "(-)" "uminus"
+proof (unfold_locales)
+ show "\<And>a. a \<in> fc_mats \<Longrightarrow> - a + a = 0\<^sub>m dimR dimC" using fc_mats_carrier by simp
+ show "\<And>a b. a \<in> fc_mats \<Longrightarrow> b \<in> fc_mats \<Longrightarrow> a - b = a + - b" using fc_mats_carrier
+ by (simp add: add_uminus_minus_mat)
+ show "\<And>a. a \<in> fc_mats \<Longrightarrow> - a \<in> fc_mats" using fc_mats_carrier by simp
+qed
+end
+
+lemma (in fixed_carrier_mat) smult_mem:
+ assumes "A \<in> fc_mats"
+ shows "a \<cdot>\<^sub>m A \<in> fc_mats" using fc_mats_carrier assms by auto
+
+definition (in fixed_carrier_mat) sum_mat where
+"sum_mat A I = sum_with (+) (0\<^sub>m dimR dimC) A I"
+
+lemma (in fixed_carrier_mat) sum_mat_empty[simp]:
+ shows "sum_mat A {} = 0\<^sub>m dimR dimC" unfolding sum_mat_def by simp
+
+lemma (in fixed_carrier_mat) sum_mat_carrier:
+ shows "(\<And>i. i \<in> I \<Longrightarrow> (A i)\<in> fc_mats) \<Longrightarrow> sum_mat A I \<in> carrier_mat dimR dimC"
+ unfolding sum_mat_def using sum_with_mem[of A I] fc_mats_carrier by auto
+
+lemma (in fixed_carrier_mat) sum_mat_insert:
+ assumes "A x \<in> fc_mats" "A ` I \<subseteq> fc_mats"
+ and A: "finite I" and x: "x \<notin> I"
+ shows "sum_mat A (insert x I) = A x + sum_mat A I" unfolding sum_mat_def
+ using assms sum_with_insert[of A x I] by simp
+
+
+subsection \<open>A locale for square matrices\<close>
+
+locale cpx_sq_mat = fixed_carrier_mat "fc_mats::complex Matrix.mat set" for fc_mats +
+ assumes dim_eq: "dimR = dimC"
+ and npos: "0 < dimR"
+
+lemma (in cpx_sq_mat) one_mem:
+ shows "1\<^sub>m dimR \<in> fc_mats" using fc_mats_carrier dim_eq by simp
+
+lemma (in cpx_sq_mat) square_mats:
+ assumes "A \<in> fc_mats"
+ shows "square_mat A" using fc_mats_carrier dim_eq assms by simp
+
+lemma (in cpx_sq_mat) cpx_sq_mat_mult:
+ assumes "A \<in> fc_mats"
+ and "B \<in> fc_mats"
+shows "A * B \<in> fc_mats"
+proof -
+ have "dim_row (A * B) = dimR" using assms fc_mats_carrier by simp
+ moreover have "dim_col (A * B) = dimR" using assms fc_mats_carrier dim_eq by simp
+ ultimately show ?thesis using fc_mats_carrier carrier_mat_def dim_eq by auto
+qed
+
+lemma (in cpx_sq_mat) sum_mat_distrib_left:
+ shows "finite I \<Longrightarrow> R\<in> fc_mats \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> (A i)\<in> fc_mats) \<Longrightarrow>
+ sum_mat (\<lambda>i. R * (A i)) I = R * (sum_mat A I)"
+proof (induct rule: finite_induct)
+ case empty
+ hence a: "sum_mat (\<lambda>i. R * (A i)) {} = 0\<^sub>m dimR dimC" unfolding sum_mat_def by simp
+ have "sum_mat A {} = 0\<^sub>m dimR dimC" unfolding sum_mat_def by simp
+ hence "R * (sum_mat A {}) = 0\<^sub>m dimR dimC" using fc_mats_carrier
+ right_mult_zero_mat[of R dimR dimC dimC] empty dim_eq by simp
+ thus ?case using a by simp
+next
+ case (insert x F)
+ hence "sum_mat (\<lambda>i. R * A i) (insert x F) = R * (A x) + sum_mat (\<lambda>i. R * A i) F"
+ using sum_mat_insert[of "\<lambda>i. R * A i" x F] by (simp add: image_subsetI fc_mats_carrier dim_eq)
+ also have "... = R * (A x) + R * (sum_mat A F)" using insert by simp
+ also have "... = R * (A x + (sum_mat A F))"
+ by (metis dim_eq fc_mats_carrier insert.prems(1) insert.prems(2) insertCI mult_add_distrib_mat
+ sum_mat_carrier)
+ also have "... = R * sum_mat A (insert x F)"
+ proof -
+ have "A x + (sum_mat A F) = sum_mat A (insert x F)"
+ by (rule sum_mat_insert[symmetric], (auto simp add: insert))
+ thus ?thesis by simp
+ qed
+ finally show ?case .
+qed
+
+lemma (in cpx_sq_mat) sum_mat_distrib_right:
+ shows "finite I \<Longrightarrow> R\<in> fc_mats \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> (A i)\<in> fc_mats) \<Longrightarrow>
+ sum_mat (\<lambda>i. (A i) * R) I = (sum_mat A I) * R"
+proof (induct rule: finite_induct)
+ case empty
+ hence a: "sum_mat (\<lambda>i. (A i) * R) {} = 0\<^sub>m dimR dimC" unfolding sum_mat_def by simp
+ have "sum_mat A {} = 0\<^sub>m dimR dimC" unfolding sum_mat_def by simp
+ hence "(sum_mat A {}) * R = 0\<^sub>m dimR dimC" using fc_mats_carrier right_mult_zero_mat[of R ]
+ dim_eq empty by simp
+ thus ?case using a by simp
+next
+ case (insert x F)
+ have a: "(\<lambda>i. A i * R) ` F \<subseteq> fc_mats" using insert cpx_sq_mat_mult
+ by (simp add: image_subsetI)
+ have "A x * R \<in> fc_mats" using insert
+ by (metis insertI1 local.fc_mats_carrier mult_carrier_mat dim_eq)
+ hence "sum_mat (\<lambda>i. A i * R) (insert x F) = (A x) * R + sum_mat (\<lambda>i. A i * R) F" using insert a
+ using sum_mat_insert[of "\<lambda>i. A i * R" x F] by (simp add: image_subsetI local.fc_mats_carrier)
+ also have "... = (A x) * R + (sum_mat A F) * R" using insert by simp
+ also have "... = (A x + (sum_mat A F)) * R"
+ proof (rule add_mult_distrib_mat[symmetric])
+ show "A x \<in> carrier_mat dimR dimC" using insert fc_mats_carrier by simp
+ show "sum_mat A F \<in> carrier_mat dimR dimC" using insert fc_mats_carrier sum_mat_carrier by blast
+ show "R \<in> carrier_mat dimC dimC" using insert fc_mats_carrier dim_eq by simp
+ qed
+ also have "... = sum_mat A (insert x F) * R"
+ proof -
+ have "A x + (sum_mat A F) = sum_mat A (insert x F)"
+ by (rule sum_mat_insert[symmetric], (auto simp add: insert))
+ thus ?thesis by simp
+ qed
+ finally show ?case .
+qed
+
+lemma (in cpx_sq_mat) trace_sum_mat:
+ fixes A::"'b \<Rightarrow> complex Matrix.mat"
+ shows "finite I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> (A i)\<in> fc_mats) \<Longrightarrow>
+ Complex_Matrix.trace (sum_mat A I) = (\<Sum> i\<in> I. Complex_Matrix.trace (A i))" unfolding sum_mat_def
+proof (induct rule: finite_induct)
+ case empty
+ then show ?case using trace_zero dim_eq by simp
+next
+ case (insert x F)
+ have "Complex_Matrix.trace (sum_with (+) (0\<^sub>m dimR dimC) A (insert x F)) =
+ Complex_Matrix.trace (A x + sum_with (+) (0\<^sub>m dimR dimC) A F)"
+ using sum_with_insert[of A x F] insert by (simp add: image_subset_iff dim_eq)
+ also have "... = Complex_Matrix.trace (A x) +
+ Complex_Matrix.trace (sum_with (+) (0\<^sub>m dimR dimC) A F)" using trace_add square_mats insert
+ by (metis carrier_matD(1) fc_mats_carrier image_subset_iff insert_iff sum_with_mem)
+ also have "... = Complex_Matrix.trace (A x) + (\<Sum> i\<in> F. Complex_Matrix.trace (A i))"
+ using insert by simp
+ also have "... = (\<Sum> i\<in> (insert x F). Complex_Matrix.trace (A i))"
+ using sum_with_insert[of A x F] insert by (simp add: image_subset_iff)
+ finally show ?case .
+qed
+
+lemma (in cpx_sq_mat) cpx_sq_mat_smult:
+ assumes "A \<in> fc_mats"
+ shows "x \<cdot>\<^sub>m A \<in> fc_mats"
+ using assms fc_mats_carrier by auto
+
+lemma (in cpx_sq_mat) mult_add_distrib_right:
+ assumes "A\<in> fc_mats" "B\<in> fc_mats" "C\<in> fc_mats"
+ shows "A * (B + C) = A * B + A * C"
+ using assms fc_mats_carrier mult_add_distrib_mat dim_eq by simp
+
+lemma (in cpx_sq_mat) mult_add_distrib_left:
+ assumes "A\<in> fc_mats" "B\<in> fc_mats" "C\<in> fc_mats"
+ shows "(B + C) * A = B * A + C * A"
+ using assms fc_mats_carrier add_mult_distrib_mat dim_eq by simp
+
+lemma (in cpx_sq_mat) mult_sum_mat_distrib_left:
+ shows "finite I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> (A i)\<in> fc_mats) \<Longrightarrow> B \<in> fc_mats \<Longrightarrow>
+ (sum_mat (\<lambda>i. B * (A i)) I) = B * (sum_mat A I)"
+proof (induct rule: finite_induct)
+ case empty
+ hence "sum_mat A {} = 0\<^sub>m dimR dimC" using sum_mat_empty by simp
+ hence "B * (sum_mat A {}) = 0\<^sub>m dimR dimC" using empty by (simp add: fc_mats_carrier dim_eq)
+ moreover have "sum_mat (\<lambda>i. B * (A i)) {} = 0\<^sub>m dimR dimC" using sum_mat_empty[of "\<lambda>i. B * (A i)"]
+ by simp
+ ultimately show ?case by simp
+next
+ case (insert x F)
+ have "sum_mat (\<lambda>i. B * (A i)) (insert x F) = B * (A x) + sum_mat (\<lambda>i. B * (A i)) F"
+ using sum_with_insert[of "\<lambda>i. B * (A i)" x F] insert
+ by (simp add: image_subset_iff local.sum_mat_def cpx_sq_mat_mult)
+ also have "... = B * (A x) + B * (sum_mat A F)" using insert by simp
+ also have "... = B * (A x + (sum_mat A F))"
+ proof (rule mult_add_distrib_right[symmetric])
+ show "B\<in> fc_mats" using insert by simp
+ show "A x \<in> fc_mats" using insert by simp
+ show "sum_mat A F \<in> fc_mats" using insert by (simp add: fc_mats_carrier sum_mat_carrier)
+ qed
+ also have "... = B * (sum_mat A (insert x F))" using sum_with_insert[of A x F] insert
+ by (simp add: image_subset_iff sum_mat_def)
+ finally show ?case .
+qed
+
+lemma (in cpx_sq_mat) mult_sum_mat_distrib_right:
+ shows "finite I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> (A i)\<in> fc_mats) \<Longrightarrow> B \<in> fc_mats \<Longrightarrow>
+ (sum_mat (\<lambda>i. (A i) * B) I) = (sum_mat A I) * B"
+proof (induct rule: finite_induct)
+ case empty
+ hence "sum_mat A {} = 0\<^sub>m dimR dimC" using sum_mat_empty by simp
+ hence "(sum_mat A {}) * B = 0\<^sub>m dimR dimC" using empty by (simp add: fc_mats_carrier dim_eq)
+ moreover have "sum_mat (\<lambda>i. (A i) * B) {} = 0\<^sub>m dimR dimC" by simp
+ ultimately show ?case by simp
+next
+ case (insert x F)
+ have "sum_mat (\<lambda>i. (A i) * B) (insert x F) = (A x) * B + sum_mat (\<lambda>i. (A i) * B) F"
+ using sum_with_insert[of "\<lambda>i. (A i) * B" x F] insert
+ by (simp add: image_subset_iff local.sum_mat_def cpx_sq_mat_mult)
+ also have "... = (A x) * B + (sum_mat A F) * B" using insert by simp
+ also have "... = (A x + (sum_mat A F)) * B"
+ proof (rule mult_add_distrib_left[symmetric])
+ show "B\<in> fc_mats" using insert by simp
+ show "A x \<in> fc_mats" using insert by simp
+ show "sum_mat A F \<in> fc_mats" using insert by (simp add: fc_mats_carrier sum_mat_carrier)
+ qed
+ also have "... = (sum_mat A (insert x F)) * B" using sum_with_insert[of A x F] insert
+ by (simp add: image_subset_iff sum_mat_def)
+ finally show ?case .
+qed
+
+lemma (in cpx_sq_mat) trace_sum_mat_mat_distrib:
+ assumes "finite I"
+and "\<And>i. i\<in> I \<Longrightarrow> B i \<in> fc_mats"
+and "A\<in> fc_mats"
+and "C \<in> fc_mats"
+shows "(\<Sum> i\<in> I. Complex_Matrix.trace(A * (B i) * C)) =
+ Complex_Matrix.trace (A * (sum_mat B I) * C)"
+proof -
+ have seq: "sum_mat (\<lambda>i. A * (B i) * C) I = A * (sum_mat B I) * C"
+ proof -
+ have "sum_mat (\<lambda>i. A * (B i) * C) I = (sum_mat (\<lambda>i. A * (B i)) I) * C"
+ proof (rule mult_sum_mat_distrib_right)
+ show "finite I" using assms by simp
+ show "C\<in> fc_mats" using assms by simp
+ show "\<And>i. i \<in> I \<Longrightarrow> A * B i \<in> fc_mats" using assms cpx_sq_mat_mult by simp
+ qed
+ moreover have "sum_mat (\<lambda>i. A * (B i)) I = A * (sum_mat B I)"
+ by (rule mult_sum_mat_distrib_left, (auto simp add: assms))
+ ultimately show "sum_mat (\<lambda>i. A * (B i) * C) I = A * (sum_mat B I) * C" by simp
+ qed
+ have "(\<Sum> i\<in> I. Complex_Matrix.trace(A * (B i) * C)) =
+ Complex_Matrix.trace (sum_mat (\<lambda>i. A * (B i) * C) I)"
+ proof (rule trace_sum_mat[symmetric])
+ show "finite I" using assms by simp
+ fix i
+ assume "i\<in> I"
+ thus "A * B i * C \<in> fc_mats" using assms by (simp add: cpx_sq_mat_mult)
+ qed
+ also have "... = Complex_Matrix.trace (A * (sum_mat B I) * C)" using seq by simp
+ finally show ?thesis .
+qed
+
+definition (in cpx_sq_mat) zero_col where
+"zero_col U = (\<lambda>i. if i < dimR then Matrix.col U i else 0\<^sub>v dimR)"
+
+lemma (in cpx_sq_mat) zero_col_dim:
+ assumes "U \<in> fc_mats"
+ shows "dim_vec (zero_col U i) = dimR" using assms fc_mats_carrier unfolding zero_col_def by simp
+
+lemma (in cpx_sq_mat) zero_col_col:
+ assumes "i < dimR"
+ shows "zero_col U i = Matrix.col U i" using assms unfolding zero_col_def by simp
+
+lemma (in cpx_sq_mat) sum_mat_index:
+ shows "finite I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> (A i)\<in> fc_mats) \<Longrightarrow> i < dimR \<Longrightarrow> j < dimC \<Longrightarrow>
+ (sum_mat (\<lambda>k. (A k)) I) $$ (i,j) = (\<Sum> k\<in>I. (A k) $$ (i,j))"
+proof (induct rule: finite_induct)
+ case empty
+ thus ?case unfolding sum_mat_def by simp
+next
+ case (insert x F)
+ hence "(sum_mat (\<lambda>k. (A k)) (insert x F)) $$ (i,j) =
+ (A x + (sum_mat (\<lambda>k. (A k)) F)) $$ (i,j)" using insert sum_mat_insert[of A]
+ by (simp add: image_subsetI local.fc_mats_carrier)
+ also have "... = (A x) $$ (i,j) + (sum_mat (\<lambda>k. (A k)) F) $$ (i,j)" using insert
+ sum_mat_carrier[of F A] fc_mats_carrier by simp
+ also have "... = (A x) $$ (i,j) + (\<Sum> k\<in>F. (A k) $$ (i,j))" using insert by simp
+ also have "... = (\<Sum> k\<in>(insert x F). (A k) $$ (i,j))" using insert by simp
+ finally show ?case .
+qed
+
+lemma (in cpx_sq_mat) sum_mat_cong:
+ shows "finite I \<Longrightarrow> (\<And>i. i\<in> I \<Longrightarrow> A i = B i) \<Longrightarrow> (\<And>i. i\<in> I \<Longrightarrow> A i \<in> fc_mats) \<Longrightarrow>
+ (\<And>i. i\<in> I \<Longrightarrow> B i \<in> fc_mats) \<Longrightarrow> sum_mat A I = sum_mat B I"
+proof (induct rule: finite_induct)
+ case empty
+ then show ?case by simp
+next
+ case (insert x F)
+ have "sum_mat A (insert x F) = A x + sum_mat A F" using insert sum_mat_insert[of A]
+ by (simp add: image_subset_iff)
+ also have "... = B x + sum_mat B F" using insert by simp
+ also have "... = sum_mat B (insert x F)" using insert sum_mat_insert[of B]
+ by (simp add: image_subset_iff)
+ finally show ?case .
+qed
+
+lemma (in cpx_sq_mat) smult_sum_mat:
+ shows "finite I \<Longrightarrow> (\<And>i. i\<in> I \<Longrightarrow> A i \<in> fc_mats) \<Longrightarrow> a \<cdot>\<^sub>m sum_mat A I = sum_mat (\<lambda>i. a \<cdot>\<^sub>m (A i)) I"
+proof (induct rule: finite_induct)
+ case empty
+ then show ?case by simp
+next
+ case (insert x F)
+ have "a \<cdot>\<^sub>m sum_mat A (insert x F) = a \<cdot>\<^sub>m (A x + sum_mat A F)" using insert sum_mat_insert[of A]
+ by (simp add: image_subset_iff)
+ also have "... = a \<cdot>\<^sub>m A x + a \<cdot>\<^sub>m (sum_mat A F)" using insert
+ by (metis add_smult_distrib_left_mat fc_mats_carrier insert_iff sum_mat_carrier)
+ also have "... = a \<cdot>\<^sub>m A x + sum_mat (\<lambda>i. a \<cdot>\<^sub>m (A i)) F" using insert by simp
+ also have "... = sum_mat (\<lambda>i. a \<cdot>\<^sub>m (A i)) (insert x F)" using insert
+ sum_mat_insert[of "(\<lambda>i. a \<cdot>\<^sub>m (A i))"] by (simp add: image_subset_iff cpx_sq_mat_smult)
+ finally show ?case .
+qed
+
+lemma (in cpx_sq_mat) zero_sum_mat:
+ shows "finite I \<Longrightarrow> sum_mat (\<lambda>i. ((0\<^sub>m dimR dimR)::complex Matrix.mat)) I = ((0\<^sub>m dimR dimR)::complex Matrix.mat)"
+proof (induct rule:finite_induct)
+ case empty
+ then show ?case using dim_eq sum_mat_empty by auto
+next
+ case (insert x F)
+ have "sum_mat (\<lambda>i. ((0\<^sub>m dimR dimR)::complex Matrix.mat)) (insert x F) =
+ 0\<^sub>m dimR dimR + sum_mat (\<lambda>i. 0\<^sub>m dimR dimR) F"
+ using insert dim_eq zero_mem sum_mat_insert[of "\<lambda>i. ((0\<^sub>m dimR dimR)::complex Matrix.mat)"]
+ by fastforce
+ also have "... = ((0\<^sub>m dimR dimR)::complex Matrix.mat)" using insert by auto
+ finally show ?case .
+qed
+
+lemma (in cpx_sq_mat) sum_mat_adjoint:
+ shows "finite I \<Longrightarrow> (\<And>i. i\<in> I \<Longrightarrow> A i \<in> fc_mats) \<Longrightarrow>
+ Complex_Matrix.adjoint (sum_mat A I) = sum_mat (\<lambda>i. Complex_Matrix.adjoint (A i)) I"
+proof (induct rule: finite_induct)
+ case empty
+ then show ?case using zero_hermitian[of dimR]
+ by (metis (no_types) dim_eq hermitian_def sum_mat_empty)
+next
+ case (insert x F)
+ have "Complex_Matrix.adjoint (sum_mat A (insert x F)) =
+ Complex_Matrix.adjoint (A x + sum_mat A F)" using insert sum_mat_insert[of A]
+ by (simp add: image_subset_iff)
+ also have "... = Complex_Matrix.adjoint (A x) + Complex_Matrix.adjoint (sum_mat A F)"
+ proof (rule adjoint_add)
+ show "A x \<in> carrier_mat dimR dimC" using insert fc_mats_carrier by simp
+ show "sum_mat A F \<in> carrier_mat dimR dimC" using insert fc_mats_carrier sum_mat_carrier[of F]
+ by simp
+ qed
+ also have "... = Complex_Matrix.adjoint (A x) + sum_mat (\<lambda>i. Complex_Matrix.adjoint (A i)) F"
+ using insert by simp
+ also have "... = sum_mat (\<lambda>i. Complex_Matrix.adjoint (A i)) (insert x F)"
+ proof (rule sum_mat_insert[symmetric], (auto simp add: insert))
+ show "Complex_Matrix.adjoint (A x) \<in> fc_mats" using insert fc_mats_carrier dim_eq
+ by (simp add: adjoint_dim)
+ show "\<And>i. i \<in> F \<Longrightarrow> Complex_Matrix.adjoint (A i) \<in> fc_mats" using insert fc_mats_carrier dim_eq
+ by (simp add: adjoint_dim)
+ qed
+ finally show ?case .
+qed
+
+lemma (in cpx_sq_mat) sum_mat_hermitian:
+ assumes "finite I"
+and "\<forall>i\<in> I. hermitian (A i)"
+and "\<forall>i\<in> I. A i\<in> fc_mats"
+shows "hermitian (sum_mat A I)"
+proof -
+ have "Complex_Matrix.adjoint (sum_mat A I) = sum_mat (\<lambda>i. Complex_Matrix.adjoint (A i)) I"
+ using assms sum_mat_adjoint[of I] by simp
+ also have "... = sum_mat A I"
+ proof (rule sum_mat_cong, (auto simp add: assms))
+ show "\<And>i. i \<in> I \<Longrightarrow> Complex_Matrix.adjoint (A i) = A i" using assms
+ unfolding hermitian_def by simp
+ show "\<And>i. i \<in> I \<Longrightarrow> Complex_Matrix.adjoint (A i) \<in> fc_mats" using assms fc_mats_carrier dim_eq
+ by (simp add: adjoint_dim)
+ qed
+ finally show ?thesis unfolding hermitian_def .
+qed
+
+lemma (in cpx_sq_mat) sum_mat_positive:
+shows "finite I \<Longrightarrow> (\<And>i. i\<in> I \<Longrightarrow> Complex_Matrix.positive (A i)) \<Longrightarrow>
+ (\<And>i. i \<in> I \<Longrightarrow> A i \<in> fc_mats) \<Longrightarrow> Complex_Matrix.positive (sum_mat A I)"
+proof (induct rule: finite_induct)
+ case empty
+ then show ?case using positive_zero[of dimR] by (metis (no_types) dim_eq sum_mat_empty)
+next
+ case (insert x F)
+ hence "sum_mat A (insert x F) = A x + (sum_mat A F)" using sum_mat_insert[of A]
+ by (simp add: image_subset_iff)
+ moreover have "Complex_Matrix.positive (A x + (sum_mat A F))"
+ proof (rule positive_add, (auto simp add: insert))
+ show "A x \<in> carrier_mat dimR dimR" using insert fc_mats_carrier dim_eq by simp
+ show "sum_mat A F \<in> carrier_mat dimR dimR" using insert sum_mat_carrier dim_eq
+ by (metis insertCI)
+ qed
+ ultimately show "Complex_Matrix.positive (sum_mat A (insert x F))" by simp
+qed
+
+lemma (in cpx_sq_mat) sum_mat_left_ortho_zero:
+ shows "finite I \<Longrightarrow>
+ (\<And>i. i\<in> I \<Longrightarrow> A i \<in> fc_mats) \<Longrightarrow> (B \<in> fc_mats) \<Longrightarrow>
+ (\<And> i. i\<in> I \<Longrightarrow> A i * B = (0\<^sub>m dimR dimR)) \<Longrightarrow>
+ (sum_mat A I) * B = 0\<^sub>m dimR dimR"
+proof (induct rule:finite_induct)
+ case empty
+ then show ?case using dim_eq
+ by (metis finite.intros(1) sum_mat_empty mult_sum_mat_distrib_right)
+next
+ case (insert x F)
+ have "(sum_mat A (insert x F)) * B =
+ (A x + sum_mat A F) * B" using insert sum_mat_insert[of A]
+ by (simp add: image_subset_iff)
+ also have "... = A x * B + sum_mat A F * B"
+ proof (rule add_mult_distrib_mat)
+ show "A x \<in> carrier_mat dimR dimC" using insert fc_mats_carrier by simp
+ show "sum_mat A F \<in> carrier_mat dimR dimC" using insert
+ by (metis insert_iff local.fc_mats_carrier sum_mat_carrier)
+ show "B \<in> carrier_mat dimC dimR" using insert fc_mats_carrier dim_eq by simp
+ qed
+ also have "... = A x * B + (0\<^sub>m dimR dimR)" using insert by simp
+ also have "... = 0\<^sub>m dimR dimR" using insert by simp
+ finally show ?case .
+qed
+
+lemma (in cpx_sq_mat) sum_mat_right_ortho_zero:
+ shows "finite I \<Longrightarrow>
+ (\<And>i. i\<in> I \<Longrightarrow> A i \<in> fc_mats) \<Longrightarrow> (B \<in> fc_mats) \<Longrightarrow>
+ (\<And> i. i\<in> I \<Longrightarrow> B * A i = (0\<^sub>m dimR dimR)) \<Longrightarrow>
+ B * (sum_mat A I) = 0\<^sub>m dimR dimR"
+proof (induct rule:finite_induct)
+ case empty
+ then show ?case using dim_eq
+ by (metis finite.intros(1) sum_mat_empty mult_sum_mat_distrib_left)
+next
+ case (insert x F)
+ have "B * (sum_mat A (insert x F)) =
+ B * (A x + sum_mat A F)" using insert sum_mat_insert[of A]
+ by (simp add: image_subset_iff)
+ also have "... = B * A x + B * sum_mat A F"
+ proof (rule mult_add_distrib_mat)
+ show "A x \<in> carrier_mat dimR dimC" using insert fc_mats_carrier by simp
+ show "sum_mat A F \<in> carrier_mat dimR dimC" using insert
+ by (metis insert_iff local.fc_mats_carrier sum_mat_carrier)
+ show "B \<in> carrier_mat dimC dimR" using insert fc_mats_carrier dim_eq by simp
+ qed
+ also have "... = B * A x + (0\<^sub>m dimR dimR)" using insert by simp
+ also have "... = 0\<^sub>m dimR dimR" using insert by simp
+ finally show ?case .
+qed
+
+lemma (in cpx_sq_mat) sum_mat_ortho_square:
+ shows "finite I \<Longrightarrow> (\<And>i. i\<in> I \<Longrightarrow> ((A i)::complex Matrix.mat) * (A i) = A i) \<Longrightarrow>
+ (\<And>i. i\<in> I \<Longrightarrow> A i \<in> fc_mats) \<Longrightarrow>
+ (\<And> i j. i\<in> I \<Longrightarrow> j\<in> I \<Longrightarrow> i\<noteq> j \<Longrightarrow> A i * (A j) = (0\<^sub>m dimR dimR)) \<Longrightarrow>
+ (sum_mat A I) * (sum_mat A I) = (sum_mat A I)"
+proof (induct rule:finite_induct)
+ case empty
+ then show ?case using dim_eq
+ by (metis fc_mats_carrier right_mult_zero_mat sum_mat_empty zero_mem)
+next
+ case (insert x F)
+ have "(sum_mat A (insert x F)) * (sum_mat A (insert x F)) =
+ (A x + sum_mat A F) * (A x + sum_mat A F)" using insert sum_mat_insert[of A]
+ by (simp add: \<open>\<And>i. i \<in> insert x F \<Longrightarrow> A i * A i = A i\<close> image_subset_iff)
+ also have "... = A x * (A x + sum_mat A F) + sum_mat A F * (A x + sum_mat A F)"
+ proof (rule add_mult_distrib_mat)
+ show "A x \<in> carrier_mat dimR dimC" using insert fc_mats_carrier by simp
+ show "sum_mat A F \<in> carrier_mat dimR dimC" using insert
+ by (metis insert_iff local.fc_mats_carrier sum_mat_carrier)
+ thus "A x + sum_mat A F \<in> carrier_mat dimC dimC" using insert dim_eq by simp
+ qed
+ also have "... = A x * A x + A x * (sum_mat A F) + sum_mat A F * (A x + sum_mat A F)"
+ proof -
+ have "A x * (A x + sum_mat A F) = A x * A x + A x * (sum_mat A F)"
+ using dim_eq insert.prems(2) mult_add_distrib_right sum_mat_carrier
+ by (metis fc_mats_carrier insertI1 subsetD subset_insertI)
+ thus ?thesis by simp
+ qed
+ also have "... = A x * A x + A x * (sum_mat A F) + sum_mat A F * A x +
+ sum_mat A F * (sum_mat A F)"
+ proof -
+ have "sum_mat A F * (A x + local.sum_mat A F) =
+ sum_mat A F * A x + local.sum_mat A F * local.sum_mat A F"
+ using insert dim_eq add_assoc add_mem mult_add_distrib_right cpx_sq_mat_mult sum_mat_carrier
+ by (metis fc_mats_carrier insertI1 subsetD subset_insertI)
+ hence "A x * A x + A x * sum_mat A F + sum_mat A F * (A x + sum_mat A F) =
+ A x * A x + A x * sum_mat A F + (sum_mat A F * A x + sum_mat A F * sum_mat A F)" by simp
+ also have "... = A x * A x + A x * sum_mat A F + sum_mat A F * A x + sum_mat A F * sum_mat A F"
+ proof (rule assoc_add_mat[symmetric])
+ show "A x * A x + A x * sum_mat A F \<in> carrier_mat dimR dimR" using sum_mat_carrier insert
+ dim_eq fc_mats_carrier by (metis add_mem cpx_sq_mat_mult insertCI)
+ show "sum_mat A F * A x \<in> carrier_mat dimR dimR" using sum_mat_carrier insert
+ dim_eq fc_mats_carrier by (metis cpx_sq_mat_mult insertCI)
+ show "sum_mat A F * sum_mat A F \<in> carrier_mat dimR dimR" using sum_mat_carrier insert
+ dim_eq fc_mats_carrier by (metis cpx_sq_mat_mult insertCI)
+ qed
+ finally show ?thesis .
+ qed
+ also have "... = A x + sum_mat A F"
+ proof -
+ have "A x * A x = A x" using insert by simp
+ moreover have "sum_mat A F * sum_mat A F = sum_mat A F" using insert by simp
+ moreover have "A x * (sum_mat A F) = 0\<^sub>m dimR dimR"
+ proof -
+ have "A x * (sum_mat A F) = sum_mat (\<lambda>i. A x * (A i)) F"
+ by (rule sum_mat_distrib_left[symmetric], (simp add: insert)+)
+ also have "... = sum_mat (\<lambda>i. 0\<^sub>m dimR dimR) F"
+ proof (rule sum_mat_cong, (auto simp add: insert zero_mem))
+ show "\<And>i. i \<in> F \<Longrightarrow> A x * A i = 0\<^sub>m dimR dimR" using insert by auto
+ show "\<And>i. i \<in> F \<Longrightarrow> A x * A i \<in> fc_mats" using insert cpx_sq_mat_mult by auto
+ show "\<And>i. i \<in> F \<Longrightarrow> 0\<^sub>m dimR dimR \<in> fc_mats" using zero_mem dim_eq by simp
+ qed
+ also have "... = 0\<^sub>m dimR dimR" using zero_sum_mat insert by simp
+ finally show ?thesis .
+ qed
+ moreover have "sum_mat A F * A x = 0\<^sub>m dimR dimR"
+ proof -
+ have "sum_mat A F * A x = sum_mat (\<lambda>i. A i * (A x)) F"
+ by (rule sum_mat_distrib_right[symmetric], (simp add: insert)+)
+ also have "... = sum_mat (\<lambda>i. 0\<^sub>m dimR dimR) F"
+ proof (rule sum_mat_cong, (auto simp add: insert zero_mem))
+ show "\<And>i. i \<in> F \<Longrightarrow> A i * A x = 0\<^sub>m dimR dimR" using insert by auto
+ show "\<And>i. i \<in> F \<Longrightarrow> A i * A x \<in> fc_mats" using insert cpx_sq_mat_mult by auto
+ show "\<And>i. i \<in> F \<Longrightarrow> 0\<^sub>m dimR dimR \<in> fc_mats" using zero_mem dim_eq by simp
+ qed
+ also have "... = 0\<^sub>m dimR dimR" using zero_sum_mat insert by simp
+ finally show ?thesis .
+ qed
+ ultimately show ?thesis using add_commute add_zero insert.prems(2) zero_mem dim_eq by auto
+ qed
+ also have "... = sum_mat A (insert x F)" using insert sum_mat_insert[of A x F]
+ by (simp add: \<open>\<And>i. i \<in> insert x F \<Longrightarrow> A i * A i = A i\<close> image_subsetI)
+ finally show ?case .
+qed
+
+lemma diagonal_unit_vec:
+ assumes "B \<in> carrier_mat n n"
+and "diagonal_mat (B::complex Matrix.mat)"
+shows "B *\<^sub>v (unit_vec n i) = B $$ (i,i) \<cdot>\<^sub>v (unit_vec n i)"
+proof -
+ define v::"complex Matrix.vec" where "v = unit_vec n i"
+ have "B *\<^sub>v v = Matrix.vec n (\<lambda> i. Matrix.scalar_prod (Matrix.row B i) v)"
+ using assms unfolding mult_mat_vec_def by simp
+ also have "... = Matrix.vec n (\<lambda> i. B $$(i,i) * Matrix.vec_index v i)"
+ proof -
+ have "\<forall>i < n. (Matrix.scalar_prod (Matrix.row B i) v = B $$(i,i) * Matrix.vec_index v i)"
+ proof (intro allI impI)
+ fix i
+ assume "i < n"
+ have "(Matrix.scalar_prod (Matrix.row B i) v) =
+ (\<Sum> j \<in> {0 ..< n}. Matrix.vec_index (Matrix.row B i) j * Matrix.vec_index v j)" using assms
+ unfolding Matrix.scalar_prod_def v_def by simp
+ also have "... = Matrix.vec_index (Matrix.row B i) i * Matrix.vec_index v i"
+ proof (rule sum_but_one)
+ show "\<forall>j < n. j \<noteq> i \<longrightarrow> Matrix.vec_index (Matrix.row B i) j = 0"
+ proof (intro allI impI)
+ fix j
+ assume "j < n" and "j \<noteq> i"
+ hence "Matrix.vec_index (Matrix.row B i) j = B $$ (i,j)" using \<open>i < n\<close> \<open>j < n\<close> assms
+ by auto
+ also have "... = 0" using assms \<open>i < n\<close> \<open>j < n\<close> \<open>j\<noteq> i\<close> unfolding diagonal_mat_def by simp
+ finally show "Matrix.vec_index (Matrix.row B i) j = 0" .
+ qed
+ show "i < n" using \<open>i < n\<close> .
+ qed
+ also have "... = B $$(i,i) * Matrix.vec_index v i" using assms \<open>i < n\<close> by auto
+ finally show "(Matrix.scalar_prod (Matrix.row B i) v) = B $$(i,i) * Matrix.vec_index v i" .
+ qed
+ thus ?thesis by auto
+ qed
+ also have "... = B $$ (i,i) \<cdot>\<^sub>v v" unfolding v_def unit_vec_def by auto
+ finally have "B *\<^sub>v v = B $$ (i,i) \<cdot>\<^sub>v v" .
+ thus ?thesis unfolding v_def by simp
+qed
+
+lemma mat_vec_mult_assoc:
+ assumes "A \<in> carrier_mat n p"
+and "B\<in> carrier_mat p q"
+and "dim_vec v = q"
+shows "A *\<^sub>v (B *\<^sub>v v) = (A * B) *\<^sub>v v" using assms by auto
+
+lemma (in cpx_sq_mat) similar_eigenvectors:
+ assumes "A\<in> fc_mats"
+ and "B\<in> fc_mats"
+ and "P\<in> fc_mats"
+ and "similar_mat_wit A B P (Complex_Matrix.adjoint P)"
+ and "diagonal_mat B"
+ and "i < n"
+shows "A *\<^sub>v (P *\<^sub>v (unit_vec dimR i)) = B $$ (i,i) \<cdot>\<^sub>v (P *\<^sub>v (unit_vec dimR i))"
+proof -
+ have "A *\<^sub>v (P *\<^sub>v (unit_vec dimR i)) =
+ (P * B * (Complex_Matrix.adjoint P)) *\<^sub>v (P *\<^sub>v (unit_vec dimR i))"
+ using assms unfolding similar_mat_wit_def by metis
+ also have "... = P * B * (Complex_Matrix.adjoint P) * P *\<^sub>v (unit_vec dimR i)"
+ proof (rule mat_vec_mult_assoc[of _ dimR dimR], (auto simp add: assms fc_mats_carrier))
+ show "P \<in> carrier_mat dimR dimR" using assms fc_mats_carrier dim_eq by simp
+ show "P * B * Complex_Matrix.adjoint P \<in> carrier_mat dimR dimR"
+ using assms fc_mats_carrier by auto
+ qed
+ also have "... = P * B * ((Complex_Matrix.adjoint P) * P) *\<^sub>v (unit_vec dimR i)" using assms dim_eq
+ by (smt fc_mats_carrier mat_assoc_test(1) similar_mat_witD2(6) similar_mat_wit_sym)
+ also have "... = P * B *\<^sub>v (unit_vec dimR i)"
+ proof -
+ have "(Complex_Matrix.adjoint P) * P = 1\<^sub>m dimR" using assms dim_eq unfolding similar_mat_wit_def
+ by (simp add: fc_mats_carrier)
+ thus ?thesis using assms(2) local.fc_mats_carrier dim_eq by auto
+ qed
+ also have "... = P *\<^sub>v (B *\<^sub>v (unit_vec dimR i))" using mat_vec_mult_assoc assms fc_mats_carrier
+ dim_eq by simp
+ also have "... = P *\<^sub>v (B $$ (i,i) \<cdot>\<^sub>v (unit_vec dimR i))" using assms diagonal_unit_vec
+ fc_mats_carrier dim_eq by simp
+ also have "... = B $$ (i,i) \<cdot>\<^sub>v (P *\<^sub>v (unit_vec dimR i))"
+ proof (rule mult_mat_vec)
+ show "P \<in> carrier_mat dimR dimC" using assms fc_mats_carrier by simp
+ show "unit_vec dimR i \<in> carrier_vec dimC" using dim_eq by simp
+ qed
+ finally show ?thesis .
+qed
+
+
+subsection \<open>Projectors\<close>
+
+definition projector where
+"projector M \<longleftrightarrow> (hermitian M \<and> M * M = M)"
+
+lemma projector_hermitian:
+ assumes "projector M"
+ shows "hermitian M" using assms unfolding projector_def by simp
+
+lemma zero_projector[simp]:
+ shows "projector (0\<^sub>m n n)" unfolding projector_def
+proof
+ show "hermitian (0\<^sub>m n n)" using zero_hermitian[of n] by simp
+ show "0\<^sub>m n n * 0\<^sub>m n n = 0\<^sub>m n n" by simp
+qed
+
+lemma projector_square_eq:
+ assumes "projector M"
+ shows "M * M = M" using assms unfolding projector_def by simp
+
+lemma projector_positive:
+ assumes "projector M"
+ shows "Complex_Matrix.positive M"
+proof (rule positive_if_decomp)
+ show "M \<in> carrier_mat (dim_row M) (dim_row M)" using assms projector_hermitian hermitian_square
+ by auto
+next
+ have "M = Complex_Matrix.adjoint M" using assms projector_hermitian[of M]
+ unfolding hermitian_def by simp
+ hence "M * Complex_Matrix.adjoint M = M * M" by simp
+ also have "... = M" using assms projector_square_eq by auto
+ finally have "M * Complex_Matrix.adjoint M = M" .
+ thus "\<exists>Ma. Ma * Complex_Matrix.adjoint Ma = M" by auto
+qed
+
+lemma projector_collapse_trace:
+ assumes "projector (P::complex Matrix.mat)"
+ and "P \<in> carrier_mat n n"
+ and "R\<in> carrier_mat n n"
+shows "Complex_Matrix.trace (P * R * P) = Complex_Matrix.trace (R * P)"
+proof -
+ have "Complex_Matrix.trace (R * P) = Complex_Matrix.trace (P * R)" using trace_comm assms by auto
+ also have "... = Complex_Matrix.trace ((P * P) * R)" using assms projector_square_eq[of P] by simp
+ also have "... = Complex_Matrix.trace (P * (P * R))" using assms by auto
+ also have "... = Complex_Matrix.trace (P * R * P)" using trace_comm[of P n "P * R"] assms by auto
+ finally have "Complex_Matrix.trace (R * P) = Complex_Matrix.trace (P * R * P)" .
+ thus ?thesis by simp
+qed
+
+lemma positive_proj_trace:
+ assumes "projector (P::complex Matrix.mat)"
+ and "Complex_Matrix.positive R"
+ and "P \<in> carrier_mat n n"
+ and "R\<in> carrier_mat n n"
+shows "Complex_Matrix.trace (R * P) \<ge> 0"
+proof -
+ have "Complex_Matrix.trace (R * P) = Complex_Matrix.trace ((P * R) * P)"
+ using assms projector_collapse_trace by auto
+ also have "... = Complex_Matrix.trace ((P * R) * (Complex_Matrix.adjoint P))"
+ using assms projector_hermitian[of P]
+ unfolding hermitian_def by simp
+ also have "... \<ge> 0"
+ proof (rule positive_trace)
+ show " P * R * Complex_Matrix.adjoint P \<in> carrier_mat n n" using assms by auto
+ show "Complex_Matrix.positive (P * R * Complex_Matrix.adjoint P)"
+ by (rule positive_close_under_left_right_mult_adjoint[of _ n], (auto simp add: assms))
+ qed
+ finally show ?thesis .
+qed
+
+lemma trace_proj_pos_real:
+ assumes "projector (P::complex Matrix.mat)"
+ and "Complex_Matrix.positive R"
+ and "P \<in> carrier_mat n n"
+ and "R\<in> carrier_mat n n"
+shows "Re (Complex_Matrix.trace (R * P)) = Complex_Matrix.trace (R * P)"
+proof -
+ have "Complex_Matrix.trace (R * P) \<ge> 0" using assms positive_proj_trace by simp
+ thus ?thesis by (simp add: complex_eqI less_eq_complex_def)
+qed
+
+lemma (in cpx_sq_mat) trace_sum_mat_proj_pos_real:
+ fixes f::"'a \<Rightarrow> real"
+ assumes "finite I"
+ and "\<forall> i\<in> I. projector (P i)"
+ and "Complex_Matrix.positive R"
+ and "\<forall>i\<in> I. P i \<in> fc_mats"
+ and "R \<in> fc_mats"
+shows "Complex_Matrix.trace (R * (sum_mat (\<lambda>i. f i \<cdot>\<^sub>m (P i)) I)) =
+ Re (Complex_Matrix.trace (R * (sum_mat (\<lambda>i. f i \<cdot>\<^sub>m (P i)) I)))"
+proof -
+ have sm: "\<And>x. x \<in> I \<Longrightarrow> Complex_Matrix.trace (f x \<cdot>\<^sub>m (R * P x)) =
+ f x * Complex_Matrix.trace (R * P x)"
+ proof -
+ fix i
+ assume "i\<in> I"
+ show "Complex_Matrix.trace (f i \<cdot>\<^sub>m (R * P i)) = f i * Complex_Matrix.trace (R * P i)"
+ proof (rule trace_smult)
+ show "R * P i \<in> carrier_mat dimR dimR" using assms cpx_sq_mat_mult fc_mats_carrier \<open>i\<in> I\<close>
+ dim_eq by simp
+ qed
+ qed
+ have sw: "\<And>x. x \<in> I \<Longrightarrow> R * (f x \<cdot>\<^sub>m P x) = f x \<cdot>\<^sub>m (R * P x)"
+ proof -
+ fix i
+ assume "i \<in> I"
+ show "R * (f i \<cdot>\<^sub>m P i) = f i \<cdot>\<^sub>m (R * P i)"
+ proof (rule mult_smult_distrib)
+ show "R\<in> carrier_mat dimR dimR" using assms fc_mats_carrier dim_eq by simp
+ show "P i \<in> carrier_mat dimR dimR" using assms \<open>i\<in> I\<close> fc_mats_carrier dim_eq by simp
+ qed
+ qed
+ have dr: "Complex_Matrix.trace (R * (sum_mat (\<lambda>i. f i \<cdot>\<^sub>m (P i)) I)) =
+ Complex_Matrix.trace (sum_mat (\<lambda>i. (R * (f i \<cdot>\<^sub>m (P i)))) I)"
+ using sum_mat_distrib_left[of I] assms by (simp add: cpx_sq_mat_smult)
+ also have trs: "... = (\<Sum> i\<in> I. Complex_Matrix.trace (R * (f i \<cdot>\<^sub>m (P i))))"
+ proof (rule trace_sum_mat, (simp add: assms))
+ show "\<And>i. i \<in> I \<Longrightarrow> R * (f i \<cdot>\<^sub>m P i) \<in> fc_mats" using assms
+ by (simp add: cpx_sq_mat_smult cpx_sq_mat_mult)
+ qed
+ also have "... = (\<Sum> i\<in> I. Complex_Matrix.trace (f i \<cdot>\<^sub>m (R * (P i))))"
+ by (rule sum.cong, (simp add: sw)+)
+ also have "... = (\<Sum> i\<in> I. f i * Complex_Matrix.trace (R * (P i)))"
+ by (rule sum.cong, (simp add: sm)+)
+ also have "... = (\<Sum> i\<in> I. complex_of_real (f i * Re (Complex_Matrix.trace (R * (P i)))))"
+ proof (rule sum.cong, simp)
+ show "\<And>x. x \<in> I \<Longrightarrow>
+ complex_of_real (f x) * Complex_Matrix.trace (R * P x) =
+ complex_of_real (f x * Re (Complex_Matrix.trace (R * P x)))"
+ proof -
+ fix x
+ assume "x\<in> I"
+ have "complex_of_real (f x) * Complex_Matrix.trace (R * P x) =
+ complex_of_real (f x) * complex_of_real (Re (Complex_Matrix.trace (R * P x)))"
+ using assms sum.cong[of I I] fc_mats_carrier trace_proj_pos_real \<open>x \<in> I\<close> dim_eq by auto
+ also have "... = complex_of_real (f x * Re (Complex_Matrix.trace (R * P x)))" by simp
+ finally show "complex_of_real (f x) * Complex_Matrix.trace (R * P x) =
+ complex_of_real (f x * Re (Complex_Matrix.trace (R * P x)))" .
+ qed
+ qed
+ also have "... = (\<Sum> i\<in> I. f i * Re (Complex_Matrix.trace (R * (P i))))" by simp
+ also have "... = (\<Sum> i\<in> I. Re (Complex_Matrix.trace (f i \<cdot>\<^sub>m (R * (P i)))))"
+ proof -
+ have "(\<Sum> i\<in> I. f i * Re (Complex_Matrix.trace (R * (P i)))) =
+ (\<Sum> i\<in> I. Re (Complex_Matrix.trace (f i \<cdot>\<^sub>m (R * (P i)))))"
+ by (rule sum.cong, (simp add: sm)+)
+ thus ?thesis by simp
+ qed
+ also have "... = (\<Sum> i\<in> I. Re (Complex_Matrix.trace (R * (f i \<cdot>\<^sub>m (P i)))))"
+ proof -
+ have "\<And>i. i \<in> I \<Longrightarrow> f i \<cdot>\<^sub>m (R * (P i)) = R * (f i \<cdot>\<^sub>m (P i))" using sw by simp
+ thus ?thesis by simp
+ qed
+ also have "... = Re (\<Sum> i\<in> I. (Complex_Matrix.trace (R * (f i \<cdot>\<^sub>m (P i)))))" by simp
+ also have "... = Re (Complex_Matrix.trace (sum_mat (\<lambda>i. R * (f i \<cdot>\<^sub>m (P i))) I))" using trs by simp
+ also have "... = Re (Complex_Matrix.trace (R * (sum_mat (\<lambda>i. f i \<cdot>\<^sub>m (P i))) I))" using dr by simp
+ finally show ?thesis .
+qed
+
+
+subsection \<open>Rank 1 projection\<close>
+
+definition rank_1_proj where
+"rank_1_proj v = outer_prod v v"
+
+lemma rank_1_proj_square_mat:
+ shows "square_mat (rank_1_proj v)" using outer_prod_dim unfolding rank_1_proj_def
+ by (metis carrier_matD(1) carrier_matD(2) carrier_vec_dim_vec square_mat.simps)
+
+lemma rank_1_proj_dim[simp]:
+ shows "dim_row (rank_1_proj v) = dim_vec v" using outer_prod_dim unfolding rank_1_proj_def
+ using carrier_vecI by blast
+
+lemma rank_1_proj_carrier[simp]:
+ shows "rank_1_proj v \<in> carrier_mat (dim_vec v) (dim_vec v)" using outer_prod_dim
+ unfolding rank_1_proj_def using carrier_vecI by blast
+
+lemma rank_1_proj_coord:
+ assumes "i < dim_vec v"
+ and "j < dim_vec v"
+shows "(rank_1_proj v) $$ (i, j) = Matrix.vec_index v i * (cnj (Matrix.vec_index v j))" using assms
+ unfolding rank_1_proj_def outer_prod_def by auto
+
+lemma rank_1_proj_adjoint:
+ shows "Complex_Matrix.adjoint (rank_1_proj (v::complex Matrix.vec)) = rank_1_proj v"
+proof
+ show "dim_row (Complex_Matrix.adjoint (rank_1_proj v)) = dim_row (rank_1_proj v)"
+ using rank_1_proj_square_mat by auto
+ thus "dim_col (Complex_Matrix.adjoint (rank_1_proj v)) = dim_col (rank_1_proj v)" by auto
+ fix i j
+ assume "i < dim_row (rank_1_proj v)" and "j < dim_col (rank_1_proj v)" note ij = this
+ have "Complex_Matrix.adjoint (rank_1_proj v) $$ (i, j) = conjugate ((rank_1_proj v) $$ (j, i))"
+ using ij \<open>dim_col (Complex_Matrix.adjoint (rank_1_proj v)) = dim_col (rank_1_proj v)\<close>
+ adjoint_eval by fastforce
+ also have "... = conjugate (Matrix.vec_index v j * (cnj (Matrix.vec_index v i)))"
+ using rank_1_proj_coord ij
+ by (metis \<open>dim_col (Complex_Matrix.adjoint (rank_1_proj v)) = dim_col (rank_1_proj v)\<close>
+ adjoint_dim_col rank_1_proj_dim)
+ also have "... = Matrix.vec_index v i * (cnj (Matrix.vec_index v j))" by simp
+ also have "... = rank_1_proj v $$ (i, j)" using ij rank_1_proj_coord
+ by (metis \<open>dim_col (Complex_Matrix.adjoint (rank_1_proj v)) = dim_col (rank_1_proj v)\<close>
+ adjoint_dim_col rank_1_proj_dim)
+ finally show "Complex_Matrix.adjoint (rank_1_proj v) $$ (i, j) = rank_1_proj v $$ (i, j)".
+qed
+
+lemma rank_1_proj_hermitian:
+ shows "hermitian (rank_1_proj (v::complex Matrix.vec))" using rank_1_proj_adjoint
+ unfolding hermitian_def by simp
+
+lemma rank_1_proj_trace:
+ assumes "\<parallel>v\<parallel> = 1"
+ shows "Complex_Matrix.trace (rank_1_proj v) = 1"
+proof -
+ have "Complex_Matrix.trace (rank_1_proj v) = inner_prod v v" using trace_outer_prod
+ unfolding rank_1_proj_def using carrier_vecI by blast
+ also have "... = (vec_norm v)\<^sup>2" unfolding vec_norm_def using power2_csqrt by presburger
+ also have "... = \<parallel>v\<parallel>\<^sup>2" using vec_norm_sq_cpx_vec_length_sq by simp
+ also have "... = 1" using assms by simp
+ finally show ?thesis .
+qed
+
+lemma rank_1_proj_mat_col:
+ assumes "A \<in> carrier_mat n n"
+ and "i < n"
+ and "j < n"
+ and "k < n"
+shows "(rank_1_proj (Matrix.col A i)) $$ (j, k) = A $$ (j, i) * conjugate (A $$ (k,i))"
+proof -
+ have "(rank_1_proj (Matrix.col A i)) $$ (j, k) = Matrix.vec_index (Matrix.col A i) j *
+ conjugate (Matrix.vec_index (Matrix.col A i) k)" using index_outer_prod[of "Matrix.col A i" n "Matrix.col A i" n]
+ assms unfolding rank_1_proj_def by simp
+ also have "... = A $$ (j, i) * conjugate (A $$ (k,i))" using assms by simp
+ finally show ?thesis .
+qed
+
+lemma (in cpx_sq_mat) weighted_sum_rank_1_proj_unitary_index:
+ assumes "A \<in> fc_mats"
+and "B \<in> fc_mats"
+and "diagonal_mat B"
+and "Complex_Matrix.unitary A"
+and "j < dimR"
+and "k < dimR"
+shows "(sum_mat (\<lambda>i. (diag_mat B)!i \<cdot>\<^sub>m rank_1_proj (Matrix.col A i)) {..< dimR}) $$ (j,k) =
+ (A * B * (Complex_Matrix.adjoint A)) $$ (j,k)"
+proof -
+ have "(sum_mat (\<lambda>i. (diag_mat B)!i \<cdot>\<^sub>m rank_1_proj (Matrix.col A i)) {..< dimR}) $$ (j,k) =
+ (\<Sum> i\<in> {..< dimR}. ((diag_mat B)!i \<cdot>\<^sub>m rank_1_proj (Matrix.col A i)) $$ (j,k))"
+ proof (rule sum_mat_index, (auto simp add: fc_mats_carrier assms))
+ show "\<And>i. i < dimR \<Longrightarrow> (diag_mat B)!i \<cdot>\<^sub>m rank_1_proj (Matrix.col A i) \<in> carrier_mat dimR dimC"
+ using rank_1_proj_carrier assms fc_mats_carrier dim_eq
+ by (metis smult_carrier_mat zero_col_col zero_col_dim)
+ show "k < dimC" using assms dim_eq by simp
+ qed
+ also have "... = (\<Sum> i\<in> {..< dimR}. (diag_mat B)!i* A $$ (j, i) * conjugate (A $$ (k,i)))"
+ proof (rule sum.cong, simp)
+ have idx: "\<And>x. x \<in> {..<dimR} \<Longrightarrow> (rank_1_proj (Matrix.col A x)) $$ (j, k) =
+ A $$ (j, x) * conjugate (A $$ (k, x))" using rank_1_proj_mat_col assms fc_mats_carrier dim_eq
+ by blast
+ show "\<And>x. x \<in> {..<dimR} \<Longrightarrow> ((diag_mat B)!x \<cdot>\<^sub>m rank_1_proj (Matrix.col A x)) $$ (j, k) =
+ (diag_mat B)!x * A $$ (j, x) * conjugate (A $$ (k, x))"
+ proof -
+ fix x
+ assume "x\<in> {..< dimR}"
+ have "((diag_mat B)!x \<cdot>\<^sub>m rank_1_proj (Matrix.col A x)) $$ (j, k) =
+ (diag_mat B)!x * (rank_1_proj (Matrix.col A x)) $$ (j, k)"
+ proof (rule index_smult_mat)
+ show "j < dim_row (rank_1_proj (Matrix.col A x))" using \<open>x\<in> {..< dimR}\<close> assms fc_mats_carrier by simp
+ show "k < dim_col (rank_1_proj (Matrix.col A x))" using \<open>x\<in> {..< dimR}\<close> assms fc_mats_carrier
+ rank_1_proj_carrier[of "Matrix.col A x"] by simp
+ qed
+ thus "((diag_mat B)!x \<cdot>\<^sub>m rank_1_proj (Matrix.col A x)) $$ (j, k) =
+ (diag_mat B)!x * A $$ (j, x) * conjugate (A $$ (k, x))" using idx \<open>x\<in> {..< dimR}\<close> by simp
+ qed
+ qed
+ also have "... = Matrix.scalar_prod (Matrix.col (Complex_Matrix.adjoint A) k) (Matrix.row (A*B) j) "
+ unfolding Matrix.scalar_prod_def
+ proof (rule sum.cong)
+ show "{..<dimR} = {0..<dim_vec (Matrix.row (A*B) j)}"
+ using assms fc_mats_carrier dim_eq by auto
+ show "\<And>x. x \<in> {0..<dim_vec (Matrix.row (A*B) j)} \<Longrightarrow>
+ diag_mat B ! x * A $$ (j, x) * conjugate (A $$ (k, x)) =
+ Matrix.vec_index ((Matrix.col (Complex_Matrix.adjoint A) k)) x *
+ Matrix.vec_index (Matrix.row (A*B) j) x"
+ proof -
+ fix x
+ assume "x\<in> {0..<dim_vec (Matrix.row (A*B) j)}"
+ hence "x\<in> {0..<dimR}" using assms fc_mats_carrier dim_eq by simp
+ have "diag_mat B ! x *A $$ (j, x) = Matrix.vec_index (Matrix.row (A*B) j) x"
+ proof (rule times_diag_index[symmetric, of _ dimR], (auto simp add: assms))
+ show "A \<in> carrier_mat dimR dimR" using assms fc_mats_carrier dim_eq by simp
+ show "B \<in> carrier_mat dimR dimR" using assms fc_mats_carrier dim_eq by simp
+ show "x < dimR" using \<open>x\<in>{0..< dimR}\<close> by simp
+ qed
+ moreover have "conjugate (A $$ (k, x)) =
+ Matrix.vec_index ((Matrix.col (Complex_Matrix.adjoint A) k)) x" using \<open>x\<in> {0..<dimR}\<close> assms
+ fc_mats_carrier dim_eq by (simp add: adjoint_col)
+ ultimately show "diag_mat B ! x * A $$ (j, x) * conjugate (A $$ (k, x)) =
+ Matrix.vec_index ((Matrix.col (Complex_Matrix.adjoint A) k)) x *
+ Matrix.vec_index (Matrix.row (A*B) j) x" by simp
+ qed
+ qed
+ also have "... = (A*B * (Complex_Matrix.adjoint A)) $$ (j,k)"
+ proof -
+ have "Matrix.mat (dim_row (A*B)) (dim_col (Complex_Matrix.adjoint A))
+ (\<lambda>(i, j). Matrix.scalar_prod (Matrix.row (A*B) i) (Matrix.col (Complex_Matrix.adjoint A) j)) $$
+ (j, k) = Matrix.scalar_prod (Matrix.row (A*B) j) (Matrix.col (Complex_Matrix.adjoint A) k)"
+ using assms fc_mats_carrier by simp
+ also have "... = Matrix.scalar_prod (Matrix.col (Complex_Matrix.adjoint A) k) (Matrix.row (A*B) j)"
+ using assms comm_scalar_prod[of "Matrix.row (A*B) j" dimR] fc_mats_carrier dim_eq
+ by (metis adjoint_dim Matrix.col_carrier_vec row_carrier_vec cpx_sq_mat_mult)
+ thus ?thesis using assms fc_mats_carrier unfolding times_mat_def by simp
+ qed
+ finally show ?thesis .
+qed
+
+lemma (in cpx_sq_mat) weighted_sum_rank_1_proj_unitary:
+ assumes "A \<in> fc_mats"
+and "B \<in> fc_mats"
+and "diagonal_mat B"
+and "Complex_Matrix.unitary A"
+shows "(sum_mat (\<lambda>i. (diag_mat B)!i \<cdot>\<^sub>m rank_1_proj (Matrix.col A i)) {..< dimR}) =
+ (A * B * (Complex_Matrix.adjoint A))"
+proof
+ show "dim_row (sum_mat (\<lambda>i. diag_mat B ! i \<cdot>\<^sub>m rank_1_proj (Matrix.col A i)) {..<dimR}) =
+ dim_row (A * B * Complex_Matrix.adjoint A)" using assms fc_mats_carrier dim_eq
+ by (metis (no_types, lifting) carrier_matD(1) cpx_sq_mat_smult dim_col index_mult_mat(2)
+ rank_1_proj_carrier sum_mat_carrier)
+ show "dim_col (local.sum_mat (\<lambda>i. diag_mat B ! i \<cdot>\<^sub>m rank_1_proj (Matrix.col A i)) {..<dimR}) =
+ dim_col (A * B * Complex_Matrix.adjoint A)" using assms fc_mats_carrier dim_eq
+ by (metis (mono_tags, lifting) adjoint_dim_col carrier_matD(1) carrier_matI dim_col
+ index_mult_mat(3) index_smult_mat(2) index_smult_mat(3) rank_1_proj_dim
+ rank_1_proj_square_mat square_mat.elims(2) square_mats sum_mat_carrier)
+ show "\<And>i j. i < dim_row (A * B * Complex_Matrix.adjoint A) \<Longrightarrow>
+ j < dim_col (A * B * Complex_Matrix.adjoint A) \<Longrightarrow>
+ local.sum_mat (\<lambda>i. diag_mat B ! i \<cdot>\<^sub>m rank_1_proj (Matrix.col A i)) {..<dimR} $$ (i, j) =
+ (A * B * Complex_Matrix.adjoint A) $$ (i, j)"
+ using weighted_sum_rank_1_proj_unitary_index[of A B] dim_eq fc_mats_carrier using assms by auto
+qed
+
+lemma rank_1_proj_projector:
+ assumes "\<parallel>v\<parallel> = 1"
+ shows "projector (rank_1_proj v)"
+proof -
+ have "Complex_Matrix.adjoint (rank_1_proj v) = rank_1_proj v" using rank_1_proj_adjoint by simp
+ hence a: "hermitian (rank_1_proj v)" unfolding hermitian_def by simp
+ have "rank_1_proj v * rank_1_proj v = inner_prod v v \<cdot>\<^sub>m outer_prod v v" unfolding rank_1_proj_def
+ using outer_prod_mult_outer_prod assms using carrier_vec_dim_vec by blast
+ also have "... = (vec_norm v)\<^sup>2 \<cdot>\<^sub>m outer_prod v v" unfolding vec_norm_def using power2_csqrt
+ by presburger
+ also have "... = (cpx_vec_length v)\<^sup>2 \<cdot>\<^sub>m outer_prod v v " using vec_norm_sq_cpx_vec_length_sq
+ by simp
+ also have "... = outer_prod v v" using assms state_qbit_norm_sq smult_one[of "outer_prod v v"]
+ by simp
+ also have "... = rank_1_proj v" unfolding rank_1_proj_def by simp
+ finally show ?thesis using a unfolding projector_def by simp
+qed
+
+lemma rank_1_proj_positive:
+ assumes "\<parallel>v\<parallel> = 1"
+ shows "Complex_Matrix.positive (rank_1_proj v)"
+proof -
+ have "projector (rank_1_proj v)" using assms rank_1_proj_projector by simp
+ thus ?thesis using projector_positive by simp
+qed
+
+lemma rank_1_proj_density:
+ assumes "\<parallel>v\<parallel> = 1"
+ shows "density_operator (rank_1_proj v)" unfolding density_operator_def using assms
+ by (simp add: rank_1_proj_positive rank_1_proj_trace)
+
+lemma (in cpx_sq_mat) sum_rank_1_proj_unitary_index:
+ assumes "A \<in> fc_mats"
+and "Complex_Matrix.unitary A"
+and "j < dimR"
+and "k < dimR"
+shows "(sum_mat (\<lambda>i. rank_1_proj (Matrix.col A i)) {..< dimR}) $$ (j,k) = (1\<^sub>m dimR) $$ (j,k)"
+proof -
+ have "(sum_mat (\<lambda>i. rank_1_proj (Matrix.col A i)) {..< dimR}) $$ (j,k) =
+ (\<Sum> i\<in> {..< dimR}. (rank_1_proj (Matrix.col A i)) $$ (j,k))"
+ proof (rule sum_mat_index, (auto simp add: fc_mats_carrier assms))
+ show "\<And>i. i < dimR \<Longrightarrow> rank_1_proj (Matrix.col A i) \<in> carrier_mat dimR dimC"
+ proof -
+ fix i
+ assume "i < dimR"
+ hence "dim_vec (Matrix.col A i) = dimR" using assms fc_mats_carrier by simp
+ thus "rank_1_proj (Matrix.col A i) \<in> carrier_mat dimR dimC" using rank_1_proj_carrier assms
+ fc_mats_carrier dim_eq fc_mats_carrier by (metis dim_col fc_mats_carrier)
+ qed
+ show "k < dimC" using assms dim_eq by simp
+ qed
+ also have "... = (\<Sum> i\<in> {..< dimR}. A $$ (j, i) * conjugate (A $$ (k,i)))"
+ proof (rule sum.cong, simp)
+ show "\<And>x. x \<in> {..<dimR} \<Longrightarrow> rank_1_proj (Matrix.col A x) $$ (j, k) =
+ A $$ (j, x) * conjugate (A $$ (k, x))"
+ using rank_1_proj_mat_col assms using local.fc_mats_carrier dim_eq by blast
+ qed
+ also have "... = Matrix.scalar_prod (Matrix.col (Complex_Matrix.adjoint A) k) (Matrix.row A j) "
+ unfolding Matrix.scalar_prod_def
+ proof (rule sum.cong)
+ show "{..<dimR} = {0..<dim_vec (Matrix.row A j)}"
+ using assms fc_mats_carrier dim_eq by auto
+ show "\<And>x. x \<in> {0..<dim_vec (Matrix.row A j)} \<Longrightarrow>
+ A $$ (j, x) * conjugate (A $$ (k, x)) =
+ Matrix.vec_index ((Matrix.col (Complex_Matrix.adjoint A) k)) x * Matrix.vec_index (Matrix.row A j) x"
+ proof -
+ fix x
+ assume "x\<in> {0..<dim_vec (Matrix.row A j)}"
+ hence "x\<in> {0..<dimR}" using assms fc_mats_carrier dim_eq by simp
+ hence "A $$ (j, x) = Matrix.vec_index (Matrix.row A j) x" using \<open>x\<in> {0..<dimR}\<close>
+ assms fc_mats_carrier dim_eq by simp
+ moreover have "conjugate (A $$ (k, x)) =
+ Matrix.vec_index ((Matrix.col (Complex_Matrix.adjoint A) k)) x" using \<open>x\<in> {0..<dimR}\<close>
+ assms fc_mats_carrier dim_eq by (simp add: adjoint_col)
+ ultimately show "A $$ (j, x) * conjugate (A $$ (k, x)) =
+ Matrix.vec_index ((Matrix.col (Complex_Matrix.adjoint A) k)) x *
+ Matrix.vec_index (Matrix.row A j) x" by simp
+ qed
+ qed
+ also have "... = (A * (Complex_Matrix.adjoint A)) $$ (j,k)"
+ proof -
+ have "Matrix.mat (dim_row A) (dim_col (Complex_Matrix.adjoint A))
+ (\<lambda>(i, j). Matrix.scalar_prod (Matrix.row A i) (Matrix.col (Complex_Matrix.adjoint A) j)) $$
+ (j, k) = Matrix.scalar_prod (Matrix.row A j) (Matrix.col (Complex_Matrix.adjoint A) k)"
+ using assms fc_mats_carrier by simp
+ also have "... = Matrix.scalar_prod (Matrix.col (Complex_Matrix.adjoint A) k) (Matrix.row A j)"
+ using assms comm_scalar_prod[of "Matrix.row A j" dimR] fc_mats_carrier
+ by (simp add: adjoint_dim dim_eq)
+ thus ?thesis using assms fc_mats_carrier unfolding times_mat_def by simp
+ qed
+ also have "... = (1\<^sub>m dimR) $$ (j,k)" using assms dim_eq by (simp add: fc_mats_carrier)
+ finally show ?thesis .
+qed
+
+lemma (in cpx_sq_mat) rank_1_proj_sum_density:
+ assumes "finite I"
+ and "\<forall>i\<in> I. \<parallel>u i\<parallel> = 1"
+ and "\<forall>i\<in> I. dim_vec (u i) = dimR"
+ and "\<forall>i\<in> I. 0 \<le> p i"
+ and "(\<Sum>i\<in> I. p i) = 1"
+shows "density_operator (sum_mat (\<lambda>i. p i \<cdot>\<^sub>m (rank_1_proj (u i))) I)" unfolding density_operator_def
+proof
+ have "Complex_Matrix.trace (sum_mat (\<lambda>i. p i \<cdot>\<^sub>m rank_1_proj (u i)) I) =
+ (\<Sum>i\<in> I. Complex_Matrix.trace (p i \<cdot>\<^sub>m rank_1_proj (u i)))"
+ proof (rule trace_sum_mat, (simp add: assms))
+ show "\<And>i. i \<in> I \<Longrightarrow> p i \<cdot>\<^sub>m rank_1_proj (u i) \<in> fc_mats" using assms smult_mem fc_mats_carrier
+ dim_eq by auto
+ qed
+ also have "... = (\<Sum>i\<in> I. p i * Complex_Matrix.trace (rank_1_proj (u i)))"
+ proof -
+ {
+ fix i
+ assume "i \<in> I"
+ hence "Complex_Matrix.trace (p i \<cdot>\<^sub>m rank_1_proj (u i)) =
+ p i * Complex_Matrix.trace (rank_1_proj (u i))"
+ using trace_smult[of "rank_1_proj (u i)" dimR] assms fc_mats_carrier dim_eq by auto
+ }
+ thus ?thesis by simp
+ qed
+ also have "... = 1" using assms rank_1_proj_trace assms by auto
+ finally show "Complex_Matrix.trace (sum_mat (\<lambda>i. p i \<cdot>\<^sub>m rank_1_proj (u i)) I) = 1" .
+next
+ show "Complex_Matrix.positive (sum_mat (\<lambda>i. p i \<cdot>\<^sub>m rank_1_proj (u i)) I)"
+ proof (rule sum_mat_positive, (simp add: assms))
+ fix i
+ assume "i\<in> I"
+ thus "p i \<cdot>\<^sub>m rank_1_proj (u i) \<in> fc_mats" using assms smult_mem fc_mats_carrier dim_eq by auto
+ show "Complex_Matrix.positive (p i \<cdot>\<^sub>m rank_1_proj (u i))" using assms \<open>i\<in> I\<close>
+ rank_1_proj_positive positive_smult[of "rank_1_proj (u i)" dimR] dim_eq by auto
+ qed
+qed
+
+
+
+lemma (in cpx_sq_mat) sum_rank_1_proj_unitary:
+ assumes "A \<in> fc_mats"
+and "Complex_Matrix.unitary A"
+shows "(sum_mat (\<lambda>i. rank_1_proj (Matrix.col A i)) {..< dimR})= (1\<^sub>m dimR)"
+proof
+ show "dim_row (sum_mat (\<lambda>i. rank_1_proj (Matrix.col A i)) {..<dimR}) = dim_row (1\<^sub>m dimR)"
+ using assms fc_mats_carrier
+ by (metis carrier_matD(1) dim_col dim_eq index_one_mat(2) rank_1_proj_carrier sum_mat_carrier)
+ show "dim_col (sum_mat (\<lambda>i. rank_1_proj (Matrix.col A i)) {..<dimR}) = dim_col (1\<^sub>m dimR)"
+ using assms fc_mats_carrier rank_1_proj_carrier sum_mat_carrier
+ by (metis carrier_matD(2) dim_col dim_eq index_one_mat(3) square_mat.simps square_mats)
+ show "\<And>i j. i < dim_row (1\<^sub>m dimR) \<Longrightarrow>
+ j < dim_col (1\<^sub>m dimR) \<Longrightarrow> sum_mat (\<lambda>i. rank_1_proj (Matrix.col A i)) {..<dimR} $$ (i, j) =
+ 1\<^sub>m dimR $$ (i, j)"
+ using assms sum_rank_1_proj_unitary_index by simp
+qed
+
+lemma (in cpx_sq_mat) rank_1_proj_unitary:
+ assumes "A \<in> fc_mats"
+and "Complex_Matrix.unitary A"
+and "j < dimR"
+and "k < dimR"
+shows "rank_1_proj (Matrix.col A j) * (rank_1_proj (Matrix.col A k)) =
+ (1\<^sub>m dimR) $$ (j,k) \<cdot>\<^sub>m (outer_prod (Matrix.col A j) (Matrix.col A k))"
+proof -
+ have "rank_1_proj (Matrix.col A j) * (rank_1_proj (Matrix.col A k)) =
+ Complex_Matrix.inner_prod (Matrix.col A j) (Matrix.col A k) \<cdot>\<^sub>m outer_prod (Matrix.col A j) (Matrix.col A k)"
+ using outer_prod_mult_outer_prod assms Matrix.col_dim local.fc_mats_carrier unfolding rank_1_proj_def
+ by blast
+ also have "... = (Complex_Matrix.adjoint A * A) $$ (j, k)\<cdot>\<^sub>m (outer_prod (Matrix.col A j) (Matrix.col A k))"
+ using inner_prod_adjoint_comp[of A dimR A] assms fc_mats_carrier dim_eq by simp
+ also have "... = (1\<^sub>m dimR) $$ (j,k) \<cdot>\<^sub>m (outer_prod (Matrix.col A j) (Matrix.col A k))" using assms
+ unfolding Complex_Matrix.unitary_def
+ by (metis add_commute assms(2) index_add_mat(2) index_one_mat(2) one_mem unitary_simps(1))
+ finally show ?thesis .
+qed
+
+lemma (in cpx_sq_mat) rank_1_proj_unitary_ne:
+ assumes "A \<in> fc_mats"
+and "Complex_Matrix.unitary A"
+and "j < dimR"
+and "k < dimR"
+and "j\<noteq> k"
+shows "rank_1_proj (Matrix.col A j) * (rank_1_proj (Matrix.col A k)) = (0\<^sub>m dimR dimR)"
+proof -
+ have "dim_row (0 \<cdot>\<^sub>m outer_prod (Matrix.col A j) (Matrix.col A k)) = dim_row (outer_prod (Matrix.col A j) (Matrix.col A k))"
+ by simp
+ also have "... = dimR" using assms fc_mats_carrier dim_eq by auto
+ finally have rw: "dim_row (0 \<cdot>\<^sub>m outer_prod (Matrix.col A j) (Matrix.col A k)) = dimR" .
+ have "dim_col (0 \<cdot>\<^sub>m outer_prod (Matrix.col A j) (Matrix.col A k)) = dim_col (outer_prod (Matrix.col A j) (Matrix.col A k))"
+ by simp
+ also have "... = dimR" using assms fc_mats_carrier dim_eq by auto
+ finally have cl: "dim_col (0 \<cdot>\<^sub>m outer_prod (Matrix.col A j) (Matrix.col A k)) = dimR" .
+ have "rank_1_proj (Matrix.col A j) * (rank_1_proj (Matrix.col A k)) =
+ (0::complex) \<cdot>\<^sub>m (outer_prod (Matrix.col A j) (Matrix.col A k))"
+ using assms rank_1_proj_unitary by simp
+ also have "... = (0\<^sub>m dimR dimR)"
+ proof
+ show "dim_row (0 \<cdot>\<^sub>m outer_prod (Matrix.col A j) (Matrix.col A k)) = dim_row (0\<^sub>m dimR dimR)" using rw by simp
+ next
+ show "dim_col (0 \<cdot>\<^sub>m outer_prod (Matrix.col A j) (Matrix.col A k)) = dim_col (0\<^sub>m dimR dimR)" using cl by simp
+ next
+ show "\<And>i p. i < dim_row (0\<^sub>m dimR dimR) \<Longrightarrow> p < dim_col (0\<^sub>m dimR dimR) \<Longrightarrow>
+ (0 \<cdot>\<^sub>m outer_prod (Matrix.col A j) (Matrix.col A k)) $$ (i, p) = 0\<^sub>m dimR dimR $$ (i, p)" using rw cl by auto
+ qed
+ finally show ?thesis .
+qed
+
+lemma (in cpx_sq_mat) rank_1_proj_unitary_eq:
+ assumes "A \<in> fc_mats"
+and "Complex_Matrix.unitary A"
+and "j < dimR"
+shows "rank_1_proj (Matrix.col A j) * (rank_1_proj (Matrix.col A j)) = rank_1_proj (Matrix.col A j)"
+proof -
+ have "rank_1_proj (Matrix.col A j) * (rank_1_proj (Matrix.col A j)) = (1::complex) \<cdot>\<^sub>m (rank_1_proj (Matrix.col A j))"
+ using assms rank_1_proj_unitary unfolding rank_1_proj_def by simp
+ also have "... = (rank_1_proj (Matrix.col A j))" by (simp add: smult_one)
+ finally show ?thesis .
+qed
+
+
end
\ No newline at end of file
diff --git a/thys/Projective_Measurements/Projective_Measurements.thy b/thys/Projective_Measurements/Projective_Measurements.thy
--- a/thys/Projective_Measurements/Projective_Measurements.thy
+++ b/thys/Projective_Measurements/Projective_Measurements.thy
@@ -1,1670 +1,1670 @@
-(*
-Author:
- Mnacho Echenim, Université Grenoble Alpes
-*)
-
-theory Projective_Measurements imports
- Linear_Algebra_Complements
-
-
-begin
-
-
-section \<open>Projective measurements\<close>
-
-text \<open>In this part we define projective measurements, also refered to as von Neumann measurements.
- The latter are characterized by a set of orthogonal projectors, which are used to compute the
-probabilities of measure outcomes and to represent the state of the system after the measurement.\<close>
-
-text \<open>The state of the system (a density operator in this case) after a measurement is represented by
-the \verb+density_collapse+ function.\<close>
-
-subsection \<open>First definitions\<close>
-
-text \<open>We begin by defining a type synonym for couples of measurement values (reals) and the
-associated projectors (complex matrices).\<close>
-
-type_synonym measure_outcome = "real \<times> complex Matrix.mat"
-
-text \<open>The corresponding values and projectors are retrieved thanks to \verb+meas_outcome_val+
-and \verb+meas_outcome_prj+.\<close>
-
-definition meas_outcome_val::"measure_outcome \<Rightarrow> real" where
-"meas_outcome_val Mi = fst Mi"
-
-definition meas_outcome_prj::"measure_outcome \<Rightarrow> complex Matrix.mat" where
-"meas_outcome_prj Mi = snd Mi"
-
-text \<open>We define a predicate for projective measurements. A projective measurement is characterized
-by the number $p$ of possible measure outcomes, and a function $M$ mapping outcome $i$ to the
-corresponding \verb+measure_outcome+.\<close>
-
-definition (in cpx_sq_mat) proj_measurement::"nat \<Rightarrow> (nat \<Rightarrow> measure_outcome) \<Rightarrow> bool" where
- "proj_measurement n M \<longleftrightarrow>
- (inj_on (\<lambda>i. meas_outcome_val (M i)) {..< n}) \<and>
- (\<forall>j < n. meas_outcome_prj (M j) \<in> fc_mats \<and>
- projector (meas_outcome_prj (M j))) \<and>
- (\<forall> i < n. \<forall> j < n. i \<noteq> j \<longrightarrow>
- meas_outcome_prj (M i) * meas_outcome_prj (M j) = 0\<^sub>m dimR dimR) \<and>
- sum_mat (\<lambda>j. meas_outcome_prj (M j)) {..< n} = 1\<^sub>m dimR"
-
-lemma (in cpx_sq_mat) proj_measurement_inj:
- assumes "proj_measurement p M"
- shows "inj_on (\<lambda>i. meas_outcome_val (M i)) {..< p}" using assms
- unfolding proj_measurement_def by simp
-
-lemma (in cpx_sq_mat) proj_measurement_carrier:
- assumes "proj_measurement p M"
- and "i < p"
- shows "meas_outcome_prj (M i) \<in> fc_mats" using assms
- unfolding proj_measurement_def by simp
-
-lemma (in cpx_sq_mat) proj_measurement_ortho:
- assumes "proj_measurement p M"
-and "i < p"
-and "j < p"
-and "i\<noteq> j"
-shows "meas_outcome_prj (M i) * meas_outcome_prj (M j) = 0\<^sub>m dimR dimR" using assms
- unfolding proj_measurement_def by simp
-
-lemma (in cpx_sq_mat) proj_measurement_id:
- assumes "proj_measurement p M"
- shows "sum_mat (\<lambda>j. meas_outcome_prj (M j)) {..< p} = 1\<^sub>m dimR" using assms
- unfolding proj_measurement_def by simp
-
-lemma (in cpx_sq_mat) proj_measurement_square:
- assumes "proj_measurement p M"
-and "i < p"
-shows "meas_outcome_prj (M i) \<in> fc_mats" using assms unfolding proj_measurement_def by simp
-
-lemma (in cpx_sq_mat) proj_measurement_proj:
- assumes "proj_measurement p M"
-and "i < p"
-shows "projector (meas_outcome_prj (M i))" using assms unfolding proj_measurement_def by simp
-
-text \<open>We define the probability of obtaining a measurement outcome: this is a positive number and
-the sum over all the measurement outcomes is 1.\<close>
-
-definition (in cpx_sq_mat) meas_outcome_prob :: "complex Matrix.mat \<Rightarrow>
- (nat \<Rightarrow> real \<times> complex Matrix.mat) \<Rightarrow> nat \<Rightarrow> complex" where
-"meas_outcome_prob R M i = Complex_Matrix.trace (R* (meas_outcome_prj (M i)))"
-
-lemma (in cpx_sq_mat) meas_outcome_prob_real:
-assumes "R\<in> fc_mats"
- and "density_operator R"
- and "proj_measurement n M"
- and "i < n"
-shows "meas_outcome_prob R M i \<in> \<real>"
-proof -
- have "complex_of_real (Re (Complex_Matrix.trace (R * meas_outcome_prj (M i)))) =
- Complex_Matrix.trace (R * meas_outcome_prj (M i))"
- proof (rule trace_proj_pos_real)
- show "R \<in> carrier_mat dimR dimR" using assms fc_mats_carrier dim_eq by simp
- show "Complex_Matrix.positive R" using assms unfolding density_operator_def by simp
- show "projector (meas_outcome_prj (M i))" using assms proj_measurement_proj by simp
- show "meas_outcome_prj (M i) \<in> carrier_mat dimR dimR" using assms proj_measurement_carrier
- fc_mats_carrier dim_eq by simp
- qed
- thus ?thesis unfolding meas_outcome_prob_def by (metis Reals_of_real)
-qed
-
-lemma (in cpx_sq_mat) meas_outcome_prob_pos:
- assumes "R\<in> fc_mats"
- and "density_operator R"
- and "proj_measurement n M"
- and "i < n"
-shows "0 \<le> meas_outcome_prob R M i" unfolding meas_outcome_prob_def
-proof (rule positive_proj_trace)
- show "R \<in> carrier_mat dimR dimR" using assms fc_mats_carrier dim_eq by simp
- show "Complex_Matrix.positive R" using assms unfolding density_operator_def by simp
- show "projector (meas_outcome_prj (M i))" using assms proj_measurement_proj by simp
- show "meas_outcome_prj (M i) \<in> carrier_mat dimR dimR" using assms proj_measurement_carrier
- fc_mats_carrier dim_eq by simp
-qed
-
-lemma (in cpx_sq_mat) meas_outcome_prob_sum:
- assumes "density_operator R"
- and "R\<in> fc_mats"
- and" proj_measurement n M"
-shows "(\<Sum> j \<in> {..< n}. meas_outcome_prob R M j) = 1"
-proof -
- have "(\<Sum> j \<in> {..< n}. Complex_Matrix.trace (R* (meas_outcome_prj (M j)))) =
- Complex_Matrix.trace (sum_mat (\<lambda>j. R* (meas_outcome_prj (M j))) {..< n})"
- proof (rule trace_sum_mat[symmetric], auto)
- fix j
- assume "j < n"
- thus "R * meas_outcome_prj (M j) \<in> fc_mats" using cpx_sq_mat_mult assms
- unfolding proj_measurement_def by simp
- qed
- also have "... = Complex_Matrix.trace (R * (sum_mat (\<lambda>j. (meas_outcome_prj (M j))) {..< n}))"
- proof -
- have "sum_mat (\<lambda>j. R* (meas_outcome_prj (M j))) {..< n} =
- R * (sum_mat (\<lambda>j. (meas_outcome_prj (M j))) {..< n})"
- proof (rule mult_sum_mat_distrib_left, (auto simp add: assms))
- fix j
- assume "j < n"
- thus "meas_outcome_prj (M j) \<in> fc_mats" using assms unfolding proj_measurement_def by simp
- qed
- thus ?thesis by simp
- qed
- also have "... = Complex_Matrix.trace (R * 1\<^sub>m dimR)" using assms unfolding proj_measurement_def
- by simp
- also have "... = Complex_Matrix.trace R" using assms by (simp add: fc_mats_carrier dim_eq)
- also have "... = 1" using assms unfolding density_operator_def by simp
- finally show ?thesis unfolding meas_outcome_prob_def .
-qed
-
-text \<open>We introduce the maximally mixed density operator. Intuitively, this density operator
-corresponds to a uniform distribution of the states of an orthonormal basis.
-This operator will be used to define the density operator after a measurement for the measure
-outcome probabilities equal to zero.\<close>
-
-definition max_mix_density :: "nat \<Rightarrow> complex Matrix.mat" where
-"max_mix_density n = ((1::real)/ n) \<cdot>\<^sub>m (1\<^sub>m n)"
-
-lemma max_mix_density_carrier:
- shows "max_mix_density n \<in> carrier_mat n n" unfolding max_mix_density_def by simp
-
-lemma max_mix_is_density:
- assumes "0 < n"
- shows "density_operator (max_mix_density n)" unfolding density_operator_def max_mix_density_def
-proof
- have "Complex_Matrix.trace (complex_of_real ((1::real)/n) \<cdot>\<^sub>m 1\<^sub>m n) =
- (complex_of_real ((1::real)/n)) * (Complex_Matrix.trace ((1\<^sub>m n)::complex Matrix.mat))"
- using one_carrier_mat trace_smult[of "(1\<^sub>m n)::complex Matrix.mat"] by blast
- also have "... = (complex_of_real ((1::real)/n)) * (real n)" using trace_1[of n] by simp
- also have "... = 1" using assms by simp
- finally show "Complex_Matrix.trace (complex_of_real ((1::real)/n) \<cdot>\<^sub>m 1\<^sub>m n) = 1" .
-next
- show "Complex_Matrix.positive (complex_of_real (1 / real n) \<cdot>\<^sub>m 1\<^sub>m n)"
- by (rule positive_smult, (auto simp add: positive_one less_eq_complex_def))
-qed
-
-lemma (in cpx_sq_mat) max_mix_density_square:
- shows "max_mix_density dimR \<in> fc_mats" unfolding max_mix_density_def
- using fc_mats_carrier dim_eq by simp
-
-text \<open>Given a measurement outcome, \verb+density_collapse+ represents the resulting density
-operator. In practice only the measure outcomes with nonzero probabilities are of interest; we
-(arbitrarily) collapse the density operator for zero-probability outcomes to the maximally mixed
-density operator.\<close>
-
-definition density_collapse ::"complex Matrix.mat \<Rightarrow> complex Matrix.mat \<Rightarrow> complex Matrix.mat" where
-"density_collapse R P = (if ((Complex_Matrix.trace (R * P)) = 0) then (max_mix_density (dim_row R))
- else ((1::real)/ ((Complex_Matrix.trace (R * P)))) \<cdot>\<^sub>m (P * R * P))"
-
-lemma density_collapse_carrier:
- assumes "0 < dim_row R"
- and "P \<in> carrier_mat n n"
- and "R \<in> carrier_mat n n"
-shows "(density_collapse R P) \<in> carrier_mat n n"
-proof (cases "(Complex_Matrix.trace (R * P)) = 0")
- case True
- hence "density_collapse R P = max_mix_density (dim_row R)" unfolding density_collapse_def by simp
- then show ?thesis using max_mix_is_density assms max_mix_density_carrier by auto
-next
- case False
- hence "density_collapse R P = complex_of_real 1 / Complex_Matrix.trace (R * P) \<cdot>\<^sub>m (P * R * P)"
- unfolding density_collapse_def by simp
- thus ?thesis using assms by auto
-qed
-
-lemma density_collapse_operator:
- assumes "projector P"
- and "density_operator R"
- and "0 < dim_row R"
- and "P \<in> carrier_mat n n"
- and "R \<in> carrier_mat n n"
-shows "density_operator (density_collapse R P)"
-proof (cases "(Complex_Matrix.trace (R * P)) = 0")
- case True
- hence "density_collapse R P = max_mix_density (dim_row R)" unfolding density_collapse_def by simp
- then show ?thesis using max_mix_is_density assms by simp
-next
- case False
- show ?thesis unfolding density_operator_def
- proof (intro conjI)
- have "Complex_Matrix.positive ((1 / (Complex_Matrix.trace (R * P))) \<cdot>\<^sub>m (P * R * P))"
- proof (rule positive_smult)
- show "P * R * P \<in> carrier_mat n n" using assms by simp
- have "Complex_Matrix.positive R" using assms unfolding density_operator_def by simp
- hence "0 \<le> (Complex_Matrix.trace (R * P))" using positive_proj_trace[of P R n] assms
- False by auto
- hence "0 \<le> Re (Complex_Matrix.trace (R * P))" by (simp add: less_eq_complex_def)
- hence "0 \<le> 1/(Re (Complex_Matrix.trace (R * P)))" by simp
- have "Re (Complex_Matrix.trace (R * P)) = Complex_Matrix.trace (R * P)"
- using assms \<open>Complex_Matrix.positive R\<close> trace_proj_pos_real by simp
- hence inv: "1/ (Complex_Matrix.trace (R * P)) = 1/(Re (Complex_Matrix.trace (R * P)))" by simp
- thus "0 \<le> 1/ (Complex_Matrix.trace (R * P))"
- using \<open>0 \<le> 1/(Re (Complex_Matrix.trace (R * P)))\<close> by (simp add: inv less_eq_complex_def)
- show "Complex_Matrix.positive (P * R * P)" using assms
- positive_close_under_left_right_mult_adjoint[of P n R]
- by (simp add: \<open>Complex_Matrix.positive R\<close> hermitian_def projector_def)
- qed
- thus "Complex_Matrix.positive (density_collapse R P)" using False
- unfolding density_collapse_def by simp
- next
- have "Complex_Matrix.trace (density_collapse R P) =
- Complex_Matrix.trace ((1/ (Complex_Matrix.trace (R * P))) \<cdot>\<^sub>m (P * R * P))"
- using False unfolding density_collapse_def by simp
- also have "... = 1/ (Complex_Matrix.trace (R * P)) * Complex_Matrix.trace (P * R * P)"
- using trace_smult[of "P * R * P" n] assms by simp
- also have "... = 1/ (Complex_Matrix.trace (R * P)) * Complex_Matrix.trace (R * P)"
- using projector_collapse_trace assms by simp
- also have "... = 1" using False by simp
- finally show "Complex_Matrix.trace (density_collapse R P) = 1" .
- qed
-qed
-
-
-subsection \<open>Measurements with observables\<close>
-
-text \<open>It is standard in quantum mechanics to represent projective measurements with so-called
-\emph{observables}. These are Hermitian matrices which encode projective measurements as follows:
-the eigenvalues of an observable represent the possible projective measurement outcomes, and the
-associated projectors are the projectors onto the corresponding eigenspaces. The results in this
-part are based on the spectral theorem, which states that any Hermitian matrix admits an
-orthonormal basis consisting of eigenvectors of the matrix.\<close>
-
-
-subsubsection \<open>On the diagonal elements of a matrix\<close>
-
-text \<open>We begin by introducing definitions that will be used on the diagonalized version of a
-Hermitian matrix.\<close>
-
-definition diag_elems where
-"diag_elems B = {B$$(i,i) |i. i < dim_row B}"
-
-text \<open>Relationship between \verb+diag_elems+ and the list \verb+diag_mat+\<close>
-
-lemma diag_elems_set_diag_mat:
- shows "diag_elems B = set (diag_mat B)" unfolding diag_mat_def diag_elems_def
-proof
- show "{B $$ (i, i) |i. i < dim_row B} \<subseteq> set (map (\<lambda>i. B $$ (i, i)) [0..<dim_row B])"
- proof
- fix x
- assume "x \<in> {B $$ (i, i) |i. i < dim_row B}"
- hence "\<exists>i < dim_row B. x = B $$(i,i)" by auto
- from this obtain i where "i < dim_row B" and "x = B $$(i,i)" by auto
- thus "x \<in> set (map (\<lambda>i. B $$ (i, i)) [0..<dim_row B])" by auto
- qed
-next
- show "set (map (\<lambda>i. B $$ (i, i)) [0..<dim_row B]) \<subseteq> {B $$ (i, i) |i. i < dim_row B}"
- proof
- fix x
- assume "x \<in> set (map (\<lambda>i. B $$ (i, i)) [0..<dim_row B])"
- thus "x \<in> {B $$ (i, i) |i. i < dim_row B}" by auto
- qed
-qed
-
-lemma diag_elems_finite[simp]:
- shows "finite (diag_elems B)" unfolding diag_elems_def by simp
-
-lemma diag_elems_mem[simp]:
- assumes "i < dim_row B"
- shows "B $$(i,i) \<in> diag_elems B" using assms unfolding diag_elems_def by auto
-
-text \<open>When $x$ is a diagonal element of $B$, \verb+diag_elem_indices+ returns the set of diagonal
-indices of $B$ with value $x$.\<close>
-
-definition diag_elem_indices where
-"diag_elem_indices x B = {i|i. i < dim_row B \<and> B $$ (i,i) = x}"
-
-lemma diag_elem_indices_elem:
- assumes "a \<in> diag_elem_indices x B"
- shows "a < dim_row B \<and> B$$(a,a) = x" using assms unfolding diag_elem_indices_def by simp
-
-lemma diag_elem_indices_itself:
- assumes "i < dim_row B"
- shows "i \<in> diag_elem_indices (B $$(i,i)) B" using assms unfolding diag_elem_indices_def by simp
-
-lemma diag_elem_indices_finite:
- shows "finite (diag_elem_indices x B)" unfolding diag_elem_indices_def by simp
-
-text \<open>We can therefore partition the diagonal indices of a matrix $B$ depending on the value
-of the diagonal elements. If $B$ admits $p$ elements on its diagonal, then we define bijections
-between its set of diagonal elements and the initial segment $[0..p-1]$.\<close>
-
-definition dist_el_card where
-"dist_el_card B = card (diag_elems B)"
-
-definition diag_idx_to_el where
-"diag_idx_to_el B = (SOME h. bij_betw h {..< dist_el_card B} (diag_elems B))"
-
-definition diag_el_to_idx where
-"diag_el_to_idx B = inv_into {..< dist_el_card B} (diag_idx_to_el B)"
-
-lemma diag_idx_to_el_bij:
- shows "bij_betw (diag_idx_to_el B) {..< dist_el_card B} (diag_elems B)"
-proof -
- let ?V = "SOME h. bij_betw h {..< dist_el_card B} (diag_elems B)"
- have vprop: "bij_betw ?V {..< dist_el_card B} (diag_elems B)" using
- someI_ex[of "\<lambda>h. bij_betw h {..< dist_el_card B} (diag_elems B)"]
- diag_elems_finite unfolding dist_el_card_def using bij_betw_from_nat_into_finite by blast
- show ?thesis by (simp add:diag_idx_to_el_def vprop)
-qed
-
-lemma diag_el_to_idx_bij:
- shows "bij_betw (diag_el_to_idx B) (diag_elems B) {..< dist_el_card B}"
- unfolding diag_el_to_idx_def
-proof (rule bij_betw_inv_into_subset[of _ _ "diag_elems B"], (simp add: diag_idx_to_el_bij)+)
- show "diag_idx_to_el B ` {..<dist_el_card B} = diag_elems B"
- by (simp add: diag_idx_to_el_bij bij_betw_imp_surj_on)
-qed
-
-lemma diag_idx_to_el_less_inj:
- assumes "i < dist_el_card B"
-and "j < dist_el_card B"
- and "diag_idx_to_el B i = diag_idx_to_el B j"
-shows "i = j"
-proof -
- have "i \<in> {..< dist_el_card B}" using assms by simp
- moreover have "j\<in> {..< dist_el_card B}" using assms by simp
- moreover have "inj_on (diag_idx_to_el B) {..<dist_el_card B}"
- using diag_idx_to_el_bij[of B]
- unfolding bij_betw_def by simp
- ultimately show ?thesis by (meson assms(3) inj_on_contraD)
-qed
-
-lemma diag_idx_to_el_less_surj:
- assumes "x\<in> diag_elems B"
-shows "\<exists>k \<in> {..< dist_el_card B}. x = diag_idx_to_el B k"
-proof -
- have "diag_idx_to_el B ` {..<dist_el_card B} = diag_elems B"
- using diag_idx_to_el_bij[of B] unfolding bij_betw_def by simp
- thus ?thesis using assms by auto
-qed
-
-lemma diag_idx_to_el_img:
- assumes "k < dist_el_card B"
-shows "diag_idx_to_el B k \<in> diag_elems B"
-proof -
- have "diag_idx_to_el B ` {..<dist_el_card B} = diag_elems B"
- using diag_idx_to_el_bij[of B] unfolding bij_betw_def by simp
- thus ?thesis using assms by auto
-qed
-
-lemma diag_elems_real:
- fixes B::"complex Matrix.mat"
- assumes "\<forall>i < dim_row B. B$$(i,i) \<in> Reals"
- shows "diag_elems B \<subseteq> Reals"
-proof
- fix x
- assume "x\<in> diag_elems B"
- hence "\<exists>i < dim_row B. x = B $$(i,i)" using assms unfolding diag_elems_def by auto
- from this obtain i where "i < dim_row B" "x = B $$ (i,i)" by auto
- thus "x \<in> Reals" using assms by simp
-qed
-
-lemma diag_elems_Re:
- fixes B::"complex Matrix.mat"
- assumes "\<forall>i < (dim_row B). B$$(i,i) \<in> Reals"
- shows "{Re x|x. x\<in> diag_elems B} = diag_elems B"
-proof
- show "{complex_of_real (Re x) |x. x \<in> diag_elems B} \<subseteq> diag_elems B"
- proof
- fix x
- assume "x \<in> {complex_of_real (Re x) |x. x \<in> diag_elems B}"
- hence "\<exists>y \<in> diag_elems B. x = Re y" by auto
- from this obtain y where "y\<in> diag_elems B" and "x = Re y" by auto
- hence "y = x" using assms diag_elems_real[of B] by auto
- thus "x\<in> diag_elems B" using \<open>y\<in> diag_elems B\<close> by simp
- qed
- show "diag_elems B \<subseteq> {complex_of_real (Re x) |x. x \<in> diag_elems B}"
- proof
- fix x
- assume "x \<in> diag_elems B"
- hence "Re x = x" using assms diag_elems_real[of B] by auto
- thus "x\<in> {complex_of_real (Re x) |x. x \<in> diag_elems B}" using \<open>x\<in> diag_elems B\<close> by force
- qed
-qed
-
-lemma diag_idx_to_el_real:
- fixes B::"complex Matrix.mat"
- assumes "\<forall>i < dim_row B. B$$(i,i) \<in> Reals"
-and "i < dist_el_card B"
-shows "Re (diag_idx_to_el B i) = diag_idx_to_el B i"
-proof -
- have "diag_idx_to_el B i \<in> diag_elems B" using diag_idx_to_el_img[of i B] assms by simp
- hence "diag_idx_to_el B i \<in> Reals" using diag_elems_real[of B] assms by auto
- thus ?thesis by simp
-qed
-
-lemma diag_elem_indices_empty:
- assumes "B \<in> carrier_mat dimR dimC"
- and "i < (dist_el_card B)"
-and "j < (dist_el_card B)"
-and "i\<noteq> j"
-shows "diag_elem_indices (diag_idx_to_el B i) B \<inter>
- (diag_elem_indices (diag_idx_to_el B j) B) = {}"
-proof (rule ccontr)
- assume "diag_elem_indices (diag_idx_to_el B i) B \<inter>
- diag_elem_indices (diag_idx_to_el B j) B \<noteq> {}"
- hence "\<exists>x. x\<in> diag_elem_indices (diag_idx_to_el B i) B \<inter>
- diag_elem_indices (diag_idx_to_el B j) B" by auto
- from this obtain x where
- xprop: "x\<in> diag_elem_indices (diag_idx_to_el B i) B \<inter>
- diag_elem_indices (diag_idx_to_el B j) B" by auto
- hence "B $$ (x,x) = (diag_idx_to_el B i)"
- using diag_elem_indices_elem[of x "diag_idx_to_el B i"] by simp
- moreover have "B $$ (x,x) = (diag_idx_to_el B j)"
- using diag_elem_indices_elem[of x "diag_idx_to_el B j"] xprop by simp
- ultimately have "diag_idx_to_el B i = diag_idx_to_el B j" by simp
- hence "i = j" using diag_idx_to_el_less_inj assms by auto
- thus False using assms by simp
-qed
-
-lemma (in cpx_sq_mat) diag_elem_indices_disjoint:
- assumes "B\<in> carrier_mat dimR dimC"
- shows "disjoint_family_on (\<lambda>n. diag_elem_indices (diag_idx_to_el B n) B)
- {..<dist_el_card B}" unfolding disjoint_family_on_def
-proof (intro ballI impI)
- fix p m
- assume "m\<in> {..< dist_el_card B}" and "p\<in> {..< dist_el_card B}" and "m\<noteq> p"
- thus "diag_elem_indices (diag_idx_to_el B m) B \<inter>
- diag_elem_indices (diag_idx_to_el B p) B = {}"
- using diag_elem_indices_empty assms fc_mats_carrier by simp
-qed
-
-lemma diag_elem_indices_union:
- assumes "B\<in> carrier_mat dimR dimC"
- shows "(\<Union> i \<in> {..< dist_el_card B}. diag_elem_indices (diag_idx_to_el B i) B) =
- {..< dimR}"
-proof
- show "(\<Union>i<dist_el_card B. diag_elem_indices (diag_idx_to_el B i) B) \<subseteq> {..<dimR}"
- proof
- fix x
- assume "x \<in> (\<Union>i<dist_el_card B. diag_elem_indices (diag_idx_to_el B i) B)"
- hence "\<exists>i < dist_el_card B. x\<in> diag_elem_indices (diag_idx_to_el B i) B" by auto
- from this obtain i where "i < dist_el_card B"
- "x\<in> diag_elem_indices (diag_idx_to_el B i) B" by auto
- hence "x < dimR" using diag_elem_indices_elem[of x _ B] assms by simp
- thus "x \<in> {..< dimR}" by auto
- qed
-next
- show "{..<dimR} \<subseteq> (\<Union>i<dist_el_card B. diag_elem_indices (diag_idx_to_el B i) B)"
- proof
- fix j
- assume "j\<in> {..< dimR}"
- hence "j < dim_row B" using assms by simp
- hence "B$$(j,j) \<in> diag_elems B" by simp
- hence "\<exists>k \<in> {..< dist_el_card B}. B$$(j,j) = diag_idx_to_el B k"
- using diag_idx_to_el_less_surj[of "B $$(j,j)"] by simp
- from this obtain k where kprop: "k \<in> {..< dist_el_card B}"
- "B$$(j,j) = diag_idx_to_el B k" by auto
- hence "j \<in> diag_elem_indices (diag_idx_to_el B k) B" using \<open>j < dim_row B\<close>
- diag_elem_indices_itself by fastforce
- thus "j \<in> (\<Union>i<dist_el_card B. diag_elem_indices (diag_idx_to_el B i) B)"
- using kprop by auto
- qed
-qed
-
-
-subsubsection \<open>Construction of measurement outcomes\<close>
-
-text \<open>The construction of a projective measurement for a hermitian matrix $A$ is based on the Schur
-decomposition $A = U*B*U^\dagger$, where $B$ is diagonal and $U$ is unitary. The columns of $U$ are
-normalized and pairwise orthogonal; they are used to construct the projectors associated with
-a measurement value\<close>
-
-definition (in cpx_sq_mat) project_vecs where
-"project_vecs (f::nat \<Rightarrow> complex Matrix.vec) N = sum_mat (\<lambda>i. rank_1_proj (f i)) N"
-
-lemma (in cpx_sq_mat) project_vecs_dim:
- assumes "\<forall>i \<in> N. dim_vec (f i) = dimR"
- shows "project_vecs f N \<in> fc_mats"
-proof -
- have "project_vecs f N \<in> carrier_mat dimR dimC" unfolding project_vecs_def
- proof (rule sum_mat_carrier)
- show "\<And>i. i \<in> N \<Longrightarrow> rank_1_proj (f i) \<in> fc_mats" using assms fc_mats_carrier rank_1_proj_dim
- dim_eq rank_1_proj_carrier by fastforce
- qed
- thus ?thesis using fc_mats_carrier by simp
-qed
-
-definition (in cpx_sq_mat) mk_meas_outcome where
-"mk_meas_outcome B U = (\<lambda>i. (Re (diag_idx_to_el B i),
- project_vecs (\<lambda>i. zero_col U i) (diag_elem_indices (diag_idx_to_el B i) B)))"
-
-lemma (in cpx_sq_mat) mk_meas_outcome_carrier:
- assumes "Complex_Matrix.unitary U"
- and "U \<in> fc_mats"
- and "B\<in> fc_mats"
-shows "meas_outcome_prj ((mk_meas_outcome B U) j) \<in> fc_mats"
-proof -
- have "project_vecs (\<lambda>i. zero_col U i) (diag_elem_indices (diag_idx_to_el B j) B) \<in> fc_mats"
- using project_vecs_dim by (simp add: assms(2) zero_col_dim)
- thus ?thesis unfolding mk_meas_outcome_def meas_outcome_prj_def by simp
-qed
-
-lemma (in cpx_sq_mat) mk_meas_outcome_sum_id:
- assumes "Complex_Matrix.unitary U"
- and "U \<in> fc_mats"
- and "B\<in> fc_mats"
-shows "sum_mat (\<lambda>j. meas_outcome_prj ((mk_meas_outcome B U) j))
- {..<(dist_el_card B)} = 1\<^sub>m dimR"
-proof -
- have "sum_mat (\<lambda>j. meas_outcome_prj ((mk_meas_outcome B U) j)) {..<(dist_el_card B)} =
- sum_mat (\<lambda>j. project_vecs (\<lambda>i. zero_col U i) (diag_elem_indices (diag_idx_to_el B j) B))
- {..<(dist_el_card B)}"
- unfolding mk_meas_outcome_def meas_outcome_prj_def by simp
- also have "... = sum_mat (\<lambda>i. rank_1_proj (zero_col U i))
- (\<Union>j<dist_el_card B. diag_elem_indices (diag_idx_to_el B j) B)"
- unfolding project_vecs_def sum_mat_def
- proof (rule disj_family_sum_with[symmetric])
- show "\<forall>j. rank_1_proj (zero_col U j) \<in> fc_mats" using zero_col_dim fc_mats_carrier dim_eq
- by (metis assms(2) rank_1_proj_carrier)
- show "finite {..<dist_el_card B}" by simp
- show "\<And>i. i \<in> {..<dist_el_card B} \<Longrightarrow> finite (diag_elem_indices (diag_idx_to_el B i) B)"
- using diag_elem_indices_finite by simp
- show "disjoint_family_on (\<lambda>n. diag_elem_indices (diag_idx_to_el B n) B)
- {..<dist_el_card B}"
- unfolding disjoint_family_on_def
- proof (intro ballI impI)
- fix p m
- assume "m\<in> {..< dist_el_card B}" and "p\<in> {..< dist_el_card B}" and "m\<noteq> p"
- thus "diag_elem_indices (diag_idx_to_el B m) B \<inter>
- diag_elem_indices (diag_idx_to_el B p) B = {}"
- using diag_elem_indices_empty assms fc_mats_carrier by simp
- qed
- qed
- also have "... = sum_mat (\<lambda>i. rank_1_proj (zero_col U i)) {..< dimR}"
- using diag_elem_indices_union[of B] assms fc_mats_carrier by simp
- also have "... = sum_mat (\<lambda>i. rank_1_proj (Matrix.col U i)) {..< dimR}"
- proof (rule sum_mat_cong, simp)
- show "\<And>i. i \<in> {..<dimR} \<Longrightarrow> rank_1_proj (zero_col U i) \<in> fc_mats" using dim_eq
- by (metis assms(2) local.fc_mats_carrier rank_1_proj_carrier zero_col_dim)
- thus "\<And>i. i \<in> {..<dimR} \<Longrightarrow> rank_1_proj (Matrix.col U i) \<in> fc_mats" using dim_eq
- by (metis lessThan_iff zero_col_col)
- show "\<And>i. i \<in> {..<dimR} \<Longrightarrow> rank_1_proj (zero_col U i) = rank_1_proj (Matrix.col U i)"
- by (simp add: zero_col_col)
- qed
- also have "... = 1\<^sub>m dimR" using sum_rank_1_proj_unitary assms by simp
- finally show ?thesis .
-qed
-
-lemma (in cpx_sq_mat) make_meas_outcome_prj_ortho:
- assumes "Complex_Matrix.unitary U"
- and "U \<in> fc_mats"
- and "B\<in> fc_mats"
- and "i < dist_el_card B"
- and "j < dist_el_card B"
- and "i \<noteq> j"
-shows "meas_outcome_prj ((mk_meas_outcome B U) i) *
- meas_outcome_prj ((mk_meas_outcome B U) j) = 0\<^sub>m dimR dimR"
-proof -
- define Pi where "Pi = sum_mat (\<lambda>k. rank_1_proj (zero_col U k))
- (diag_elem_indices (diag_idx_to_el B i) B)"
- have sneqi: "meas_outcome_prj (mk_meas_outcome B U i) = Pi"
- unfolding mk_meas_outcome_def project_vecs_def Pi_def meas_outcome_prj_def by simp
- define Pj where "Pj = sum_mat (\<lambda>k. rank_1_proj (zero_col U k))
- (diag_elem_indices (diag_idx_to_el B j) B)"
- have sneqj: "meas_outcome_prj (mk_meas_outcome B U j) = Pj"
- unfolding mk_meas_outcome_def project_vecs_def Pj_def meas_outcome_prj_def by simp
- have "Pi * Pj = 0\<^sub>m dimR dimR" unfolding Pi_def
- proof (rule sum_mat_left_ortho_zero)
- show "finite (diag_elem_indices (diag_idx_to_el B i) B)"
- using diag_elem_indices_finite[of _ B] by simp
- show km: "\<And>k. k \<in> diag_elem_indices (diag_idx_to_el B i) B \<Longrightarrow>
- rank_1_proj (zero_col U k) \<in> fc_mats" using zero_col_dim assms fc_mats_carrier dim_eq
- by (metis rank_1_proj_carrier)
- show "Pj \<in> fc_mats" using sneqj assms mk_meas_outcome_carrier by auto
- show "\<And>k. k \<in> diag_elem_indices (diag_idx_to_el B i) B \<Longrightarrow>
- rank_1_proj (zero_col U k) * Pj = 0\<^sub>m dimR dimR"
- proof -
- fix k
- assume "k \<in> diag_elem_indices (diag_idx_to_el B i) B"
- show "rank_1_proj (zero_col U k) * Pj = 0\<^sub>m dimR dimR" unfolding Pj_def
- proof (rule sum_mat_right_ortho_zero)
- show "finite (diag_elem_indices (diag_idx_to_el B j) B)"
- using diag_elem_indices_finite[of _ B] by simp
- show "\<And>i. i \<in> diag_elem_indices (diag_idx_to_el B j) B \<Longrightarrow>
- rank_1_proj (zero_col U i) \<in> fc_mats" using zero_col_dim assms fc_mats_carrier dim_eq
- by (metis rank_1_proj_carrier)
- show "rank_1_proj (zero_col U k) \<in> fc_mats"
- by (simp add: km \<open>k \<in> diag_elem_indices (diag_idx_to_el B i) B\<close>)
- show "\<And>i. i \<in> diag_elem_indices (diag_idx_to_el B j) B \<Longrightarrow>
- rank_1_proj (zero_col U k) * rank_1_proj (zero_col U i) = 0\<^sub>m dimR dimR"
- proof -
- fix m
- assume "m \<in> diag_elem_indices (diag_idx_to_el B j) B"
- hence "m \<noteq> k" using \<open>k \<in> diag_elem_indices (diag_idx_to_el B i) B\<close>
- diag_elem_indices_disjoint[of B] fc_mats_carrier assms unfolding disjoint_family_on_def by auto
- have "\<And>i. i \<in> diag_elem_indices (diag_idx_to_el B j) B \<Longrightarrow> i < dimR"
- using diag_elem_indices_elem fc_mats_carrier assms dim_eq by (metis carrier_matD(1))
- hence "m < dimR" using \<open>m \<in> diag_elem_indices (diag_idx_to_el B j) B\<close> by simp
- have "\<And>k. k \<in> diag_elem_indices (diag_idx_to_el B i) B \<Longrightarrow> k < dimR"
- using diag_elem_indices_elem fc_mats_carrier assms dim_eq by (metis carrier_matD(1))
- hence "k < dimR" using \<open>k \<in> diag_elem_indices (diag_idx_to_el B i) B\<close> by simp
- show "rank_1_proj (zero_col U k) * rank_1_proj (zero_col U m) = 0\<^sub>m dimR dimR"
- using rank_1_proj_unitary_ne[of U k m] assms \<open>m < dimR\<close> \<open>k < dimR\<close>
- by (metis \<open>m \<noteq> k\<close> zero_col_col)
- qed
- qed
- qed
- qed
- thus ?thesis using sneqi sneqj by simp
-qed
-
-lemma (in cpx_sq_mat) make_meas_outcome_prjectors:
- assumes "Complex_Matrix.unitary U"
- and "U \<in> fc_mats"
- and "B\<in> fc_mats"
- and "j < dist_el_card B"
-shows "projector (meas_outcome_prj ((mk_meas_outcome B U) j))" unfolding projector_def
-proof
- define Pj where "Pj = sum_mat (\<lambda>i. rank_1_proj (zero_col U i))
- (diag_elem_indices (diag_idx_to_el B j) B)"
- have sneq: "meas_outcome_prj (mk_meas_outcome B U j) = Pj"
- unfolding mk_meas_outcome_def project_vecs_def Pj_def meas_outcome_prj_def by simp
- moreover have "hermitian Pj" unfolding Pj_def
- proof (rule sum_mat_hermitian)
- show "finite (diag_elem_indices (diag_idx_to_el B j) B)"
- using diag_elem_indices_finite[of _ B] by simp
- show "\<forall>i\<in>diag_elem_indices (diag_idx_to_el B j) B. hermitian (rank_1_proj (zero_col U i))"
- using rank_1_proj_hermitian by simp
- show "\<forall>i\<in>diag_elem_indices (diag_idx_to_el B j) B. rank_1_proj (zero_col U i) \<in> fc_mats"
- using zero_col_dim fc_mats_carrier dim_eq by (metis assms(2) rank_1_proj_carrier)
- qed
- ultimately show "hermitian (meas_outcome_prj (mk_meas_outcome B U j))" by simp
- have "Pj * Pj = Pj" unfolding Pj_def
- proof (rule sum_mat_ortho_square)
- show "finite (diag_elem_indices (diag_idx_to_el B j) B)"
- using diag_elem_indices_finite[of _ B] by simp
- show "\<And>i. i \<in> diag_elem_indices (diag_idx_to_el B j) B \<Longrightarrow>
- rank_1_proj (zero_col U i) * rank_1_proj (zero_col U i) = rank_1_proj (zero_col U i)"
- proof -
- fix i
- assume imem: "i\<in> diag_elem_indices (diag_idx_to_el B j) B"
- hence "i < dimR" using diag_elem_indices_elem fc_mats_carrier assms dim_eq
- by (metis carrier_matD(1))
- hence "zero_col U i = Matrix.col U i" using zero_col_col[of i] by simp
- hence "rank_1_proj (zero_col U i) = rank_1_proj (Matrix.col U i)" by simp
- moreover have "rank_1_proj (Matrix.col U i) * rank_1_proj (Matrix.col U i) =
- rank_1_proj (Matrix.col U i)" by (rule rank_1_proj_unitary_eq, (auto simp add: assms \<open>i < dimR\<close>))
- ultimately show "rank_1_proj (zero_col U i) * rank_1_proj (zero_col U i) =
- rank_1_proj (zero_col U i)" by simp
- qed
- show "\<And>i. i \<in> diag_elem_indices (diag_idx_to_el B j) B \<Longrightarrow>
- rank_1_proj (zero_col U i) \<in> fc_mats"
- using zero_col_dim assms fc_mats_carrier dim_eq by (metis rank_1_proj_carrier)
- have "\<And>i. i \<in> diag_elem_indices (diag_idx_to_el B j) B \<Longrightarrow> i < dimR"
- using diag_elem_indices_elem fc_mats_carrier assms dim_eq by (metis carrier_matD(1))
- thus "\<And>i ja.
- i \<in> diag_elem_indices (diag_idx_to_el B j) B \<Longrightarrow>
- ja \<in> diag_elem_indices (diag_idx_to_el B j) B \<Longrightarrow>
- i \<noteq> ja \<Longrightarrow> rank_1_proj (zero_col U i) * rank_1_proj (zero_col U ja) = 0\<^sub>m dimR dimR"
- using rank_1_proj_unitary_ne by (simp add: assms(1) assms(2) zero_col_col)
- qed
- thus "meas_outcome_prj (mk_meas_outcome B U j) *
- meas_outcome_prj (mk_meas_outcome B U j) =
- meas_outcome_prj (mk_meas_outcome B U j)"
- using sneq by simp
-qed
-
-lemma (in cpx_sq_mat) mk_meas_outcome_fst_inj:
- assumes "\<forall>i < (dim_row B). B$$(i,i) \<in> Reals"
- shows "inj_on (\<lambda>i. meas_outcome_val ((mk_meas_outcome B U) i)) {..<dist_el_card B}"
- unfolding inj_on_def
-proof (intro ballI impI)
- fix x y
- assume "x \<in> {..<dist_el_card B}" and "y \<in> {..<dist_el_card B}"
- and "meas_outcome_val (mk_meas_outcome B U x) =
- meas_outcome_val (mk_meas_outcome B U y)" note xy = this
- hence "diag_idx_to_el B x = Re (diag_idx_to_el B x)"
- using assms diag_idx_to_el_real by simp
- also have "... = Re (diag_idx_to_el B y)" using xy
- unfolding mk_meas_outcome_def meas_outcome_val_def by simp
- also have "... = diag_idx_to_el B y" using assms diag_idx_to_el_real xy by simp
- finally have "diag_idx_to_el B x = diag_idx_to_el B y" .
- thus "x = y " using diag_idx_to_el_less_inj xy by simp
-qed
-
-lemma (in cpx_sq_mat) mk_meas_outcome_fst_bij:
- assumes "\<forall>i < (dim_row B). B$$(i,i) \<in> Reals"
- shows "bij_betw (\<lambda>i. meas_outcome_val ((mk_meas_outcome B U) i)) {..< dist_el_card B}
- {Re x|x. x\<in> diag_elems B}"
- unfolding bij_betw_def
-proof
- have "inj_on (\<lambda>x. (meas_outcome_val (mk_meas_outcome B U x))) {..<dist_el_card B}"
- using assms mk_meas_outcome_fst_inj by simp
- moreover have "{Re x|x. x\<in> diag_elems B} = diag_elems B" using diag_elems_Re[of B] assms by simp
- ultimately show "inj_on (\<lambda>x. meas_outcome_val (mk_meas_outcome B U x))
- {..<dist_el_card B}" by simp
- show "(\<lambda>i. meas_outcome_val (mk_meas_outcome B U i)) ` {..<dist_el_card B} =
- {Re x |x. x \<in> diag_elems B}" unfolding meas_outcome_val_def mk_meas_outcome_def
- proof
- show "(\<lambda>i. fst (Re (diag_idx_to_el B i), project_vecs (zero_col U)
- (diag_elem_indices (diag_idx_to_el B i) B))) ` {..<dist_el_card B} \<subseteq>
- {Re x |x. x \<in> diag_elems B}"
- using diag_idx_to_el_bij[of B] bij_betw_apply by fastforce
- show "{Re x |x. x \<in> diag_elems B}
- \<subseteq> (\<lambda>i. fst (Re (diag_idx_to_el B i),
- project_vecs (zero_col U) (diag_elem_indices (diag_idx_to_el B i) B))) `
- {..<dist_el_card B}" using diag_idx_to_el_less_surj by fastforce
- qed
-qed
-
-
-subsubsection \<open>Projective measurement associated with an observable\<close>
-
-definition eigvals where
-"eigvals M = (SOME as. char_poly M = (\<Prod>a\<leftarrow>as. [:- a, 1:]) \<and> length as = dim_row M)"
-
-lemma eigvals_poly_length:
- assumes "(M::complex Matrix.mat) \<in> carrier_mat n n"
- shows "char_poly M = (\<Prod>a\<leftarrow>(eigvals M). [:- a, 1:]) \<and> length (eigvals M) = dim_row M"
-proof -
- let ?V = "SOME as. char_poly M = (\<Prod>a\<leftarrow>as. [:- a, 1:]) \<and> length as = dim_row M"
- have vprop: "char_poly M = (\<Prod>a\<leftarrow>?V. [:- a, 1:]) \<and> length ?V = dim_row M" using
- someI_ex[of "\<lambda>as. char_poly M = (\<Prod>a\<leftarrow>as. [:- a, 1:]) \<and> length as = dim_row M"]
- complex_mat_char_poly_factorizable assms by blast
- show ?thesis by (metis (no_types) eigvals_def vprop)
-qed
-
-text \<open>We define the spectrum of a matrix $A$: this is its set of eigenvalues; its elements are
-roots of the characteristic polynomial of $A$.\<close>
-
-definition spectrum where
-"spectrum M = set (eigvals M)"
-
-lemma spectrum_finite:
- shows "finite (spectrum M)" unfolding spectrum_def by simp
-
-lemma spectrum_char_poly_root:
- fixes A::"complex Matrix.mat"
- assumes "A\<in> carrier_mat n n"
-and "k \<in> spectrum A"
-shows "poly (char_poly A) k = 0" using eigvals_poly_length[of A n] assms
- unfolding spectrum_def eigenvalue_root_char_poly
- by (simp add: linear_poly_root)
-
-lemma spectrum_eigenvalues:
- fixes A::"complex Matrix.mat"
- assumes "A\<in> carrier_mat n n"
-and "k\<in> spectrum A"
-shows "eigenvalue A k" using eigenvalue_root_char_poly[of A n k]
- spectrum_char_poly_root[of A n k] assms by simp
-
-text \<open>The main result that is used to construct a projective measurement for a Hermitian matrix
-is that it is always possible to decompose it as $A = U*B*U^\dagger$, where $B$ is diagonal and
-only contains real elements, and $U$ is unitary.\<close>
-
-definition hermitian_decomp where
-"hermitian_decomp A B U \<equiv> similar_mat_wit A B U (Complex_Matrix.adjoint U) \<and> diagonal_mat B \<and>
- diag_mat B = (eigvals A) \<and> unitary U \<and> (\<forall>i < dim_row B. B$$(i, i) \<in> Reals)"
-
-lemma hermitian_decomp_sim:
- assumes "hermitian_decomp A B U"
- shows "similar_mat_wit A B U (Complex_Matrix.adjoint U)" using assms
- unfolding hermitian_decomp_def by simp
-
-lemma hermitian_decomp_diag_mat:
- assumes "hermitian_decomp A B U"
- shows "diagonal_mat B" using assms
- unfolding hermitian_decomp_def by simp
-
-lemma hermitian_decomp_eigenvalues:
- assumes "hermitian_decomp A B U"
- shows "diag_mat B = (eigvals A)" using assms
- unfolding hermitian_decomp_def by simp
-
-lemma hermitian_decomp_unitary:
- assumes "hermitian_decomp A B U"
- shows "unitary U" using assms
- unfolding hermitian_decomp_def by simp
-
-lemma hermitian_decomp_real_eigvals:
- assumes "hermitian_decomp A B U"
- shows "\<forall>i < dim_row B. B$$(i, i) \<in> Reals" using assms
- unfolding hermitian_decomp_def by simp
-
-lemma hermitian_decomp_dim_carrier:
- assumes "hermitian_decomp A B U"
- shows "B \<in> carrier_mat (dim_row A) (dim_col A)" using assms
- unfolding hermitian_decomp_def similar_mat_wit_def
- by (metis (full_types) index_mult_mat(3) index_one_mat(3) insert_subset)
-
-lemma similar_mat_wit_dim_row:
- assumes "similar_mat_wit A B Q R"
- shows "dim_row B = dim_row A" using assms Let_def unfolding similar_mat_wit_def
- by (meson carrier_matD(1) insert_subset)
-
-lemma (in cpx_sq_mat) hermitian_schur_decomp:
- assumes "hermitian A"
- and "A\<in> fc_mats"
-obtains B U where "hermitian_decomp A B U"
-proof -
- {
- have es: "char_poly A = (\<Prod> (e :: complex) \<leftarrow> (eigvals A). [:- e, 1:])"
- using assms fc_mats_carrier eigvals_poly_length dim_eq by auto
- obtain B U Q where us: "unitary_schur_decomposition A (eigvals A) = (B,U,Q)"
- by (cases "unitary_schur_decomposition A (eigvals A)")
- hence pr: "similar_mat_wit A B U (Complex_Matrix.adjoint U) \<and> diagonal_mat B \<and>
- diag_mat B = (eigvals A) \<and> unitary U \<and> (\<forall>i < dimR. B$$(i, i) \<in> Reals)"
- using hermitian_eigenvalue_real assms fc_mats_carrier es dim_eq by auto
- moreover have "dim_row B = dimR" using assms fc_mats_carrier dim_eq similar_mat_wit_dim_row[of A]
- pr by auto
- ultimately have "hermitian_decomp A B U" unfolding hermitian_decomp_def by simp
- hence "\<exists> B U. hermitian_decomp A B U" by auto
- }
- thus ?thesis using that by auto
-qed
-
-lemma (in cpx_sq_mat) hermitian_spectrum_real:
- assumes "A \<in> fc_mats"
- and "hermitian A"
- and "a \<in> spectrum A"
-shows "a \<in> Reals"
-proof -
- obtain B U where bu: "hermitian_decomp A B U" using assms hermitian_schur_decomp by auto
- hence dimB: "B \<in> carrier_mat dimR dimR" using assms fc_mats_carrier
- dim_eq hermitian_decomp_dim_carrier[of A] by simp
- hence Bii: "\<And>i. i < dimR \<Longrightarrow> B$$(i, i) \<in> Reals" using hermitian_decomp_real_eigvals[of A B U]
- bu assms fc_mats_carrier by simp
- have "diag_mat B = (eigvals A)" using bu hermitian_decomp_eigenvalues[of A B] by simp
- hence "a \<in> set (diag_mat B)" using assms unfolding spectrum_def by simp
- hence "\<exists>i < length (diag_mat B). a = diag_mat B ! i" by (metis in_set_conv_nth)
- from this obtain i where "i < length (diag_mat B)" and "a = diag_mat B ! i" by auto
- hence "a = B $$ (i,i)" unfolding diag_mat_def by simp
- moreover have "i < dimR" using dimB \<open>i < length (diag_mat B)\<close> unfolding diag_mat_def by simp
- ultimately show ?thesis using Bii by simp
-qed
-
-lemma (in cpx_sq_mat) spectrum_ne:
- assumes "A\<in> fc_mats"
-and "hermitian A"
-shows "spectrum A \<noteq> {}"
-proof -
- obtain B U where bu: "hermitian_decomp A B U" using assms hermitian_schur_decomp by auto
- hence dimB: "B \<in> carrier_mat dimR dimR" using assms fc_mats_carrier
- dim_eq hermitian_decomp_dim_carrier[of A] by simp
- have "diag_mat B = (eigvals A)" using bu hermitian_decomp_eigenvalues[of A B] by simp
- have "B$$(0,0) \<in> set (diag_mat B)" using dimB npos unfolding diag_mat_def by simp
- hence "set (diag_mat B) \<noteq> {}" by auto
- thus ?thesis unfolding spectrum_def using \<open>diag_mat B = eigvals A\<close> by auto
-qed
-
-lemma unitary_hermitian_eigenvalues:
- fixes U::"complex Matrix.mat"
- assumes "unitary U"
- and "hermitian U"
- and "U \<in> carrier_mat n n"
- and "0 < n"
- and "k \<in> spectrum U"
-shows "k \<in> {-1, 1}"
-proof -
- have "cpx_sq_mat n n (carrier_mat n n)" by (unfold_locales, (simp add: assms)+)
- have "eigenvalue U k" using spectrum_eigenvalues assms by simp
- have "k \<in> Reals" using assms \<open>cpx_sq_mat n n (carrier_mat n n)\<close>
- cpx_sq_mat.hermitian_spectrum_real by simp
- hence "conjugate k = k" by (simp add: Reals_cnj_iff)
- hence "k\<^sup>2 = 1" using unitary_eigenvalues_norm_square[of U n k] assms
- by (simp add: \<open>eigenvalue U k\<close> power2_eq_square)
- thus ?thesis using power2_eq_1_iff by auto
-qed
-
-lemma unitary_hermitian_Re_spectrum:
- fixes U::"complex Matrix.mat"
- assumes "unitary U"
- and "hermitian U"
- and "U \<in> carrier_mat n n"
- and "0 < n"
- shows "{Re x|x. x\<in> spectrum U} \<subseteq> {-1,1}"
-proof
- fix y
- assume "y\<in> {Re x|x. x\<in> spectrum U}"
- hence "\<exists>x \<in> spectrum U. y = Re x" by auto
- from this obtain x where "x\<in> spectrum U" and "y = Re x" by auto
- hence "x\<in> {-1,1}" using unitary_hermitian_eigenvalues assms by simp
- thus "y \<in> {-1, 1}" using \<open>y = Re x\<close> by auto
-qed
-
-text \<open>The projective measurement associated with matrix $M$ is obtained from its Schur
-decomposition, by considering the number of distinct elements on the resulting diagonal matrix
-and constructing the projectors associated with each of them.\<close>
-
-type_synonym proj_meas_rep = "nat \<times> (nat \<Rightarrow> measure_outcome)"
-
-definition proj_meas_size::"proj_meas_rep \<Rightarrow> nat" where
-"proj_meas_size P = fst P"
-
-definition proj_meas_outcomes::"proj_meas_rep \<Rightarrow> (nat \<Rightarrow> measure_outcome)" where
-"proj_meas_outcomes P = snd P"
-
-definition (in cpx_sq_mat) make_pm::"complex Matrix.mat \<Rightarrow> proj_meas_rep" where
-"make_pm A = (let (B,U,_) = unitary_schur_decomposition A (eigvals A) in
- (dist_el_card B, mk_meas_outcome B U))"
-
-lemma (in cpx_sq_mat) make_pm_decomp:
- shows "make_pm A = (proj_meas_size (make_pm A), proj_meas_outcomes (make_pm A))"
- unfolding proj_meas_size_def proj_meas_outcomes_def by simp
-
-lemma (in cpx_sq_mat) make_pm_proj_measurement:
- assumes "A \<in> fc_mats"
- and "hermitian A"
- and "make_pm A = (n, M)"
-shows "proj_measurement n M"
-proof -
- have es: "char_poly A = (\<Prod> (e :: complex) \<leftarrow> (eigvals A). [:- e, 1:])"
- using assms fc_mats_carrier eigvals_poly_length dim_eq by auto
- obtain B U Q where us: "unitary_schur_decomposition A (eigvals A) = (B,U,Q)"
- by (cases "unitary_schur_decomposition A (eigvals A)", auto)
- then have "similar_mat_wit A B U (Complex_Matrix.adjoint U) \<and> diagonal_mat B \<and>
- diag_mat B = (eigvals A) \<and> unitary U \<and> (\<forall>i < dimR. B$$(i, i) \<in> Reals)"
- using hermitian_eigenvalue_real assms fc_mats_carrier es dim_eq by auto
- hence A: "A = U * B * (Complex_Matrix.adjoint U)" and dB: "diagonal_mat B"
- and Bii: "\<And>i. i < dimR \<Longrightarrow> B$$(i, i) \<in> Reals"
- and dimB: "B \<in> carrier_mat dimR dimR" and dimP: "U \<in> carrier_mat dimR dimR" and
- dimaP: "Complex_Matrix.adjoint U \<in> carrier_mat dimR dimR" and unit: "unitary U"
- unfolding similar_mat_wit_def Let_def using assms fc_mats_carrier by auto
- hence mpeq: "make_pm A = (dist_el_card B, mk_meas_outcome B U)" using us Let_def
- unfolding make_pm_def by simp
- hence "n = dist_el_card B" using assms by simp
- have "M = mk_meas_outcome B U" using assms mpeq by simp
- show ?thesis unfolding proj_measurement_def
- proof (intro conjI)
- show "inj_on (\<lambda>i. meas_outcome_val (M i)) {..<n}"
- using mk_meas_outcome_fst_inj[of B U]
- \<open>n = dist_el_card B\<close> \<open>M = mk_meas_outcome B U\<close> Bii fc_mats_carrier dimB by auto
- show "\<forall>j<n. meas_outcome_prj (M j) \<in> fc_mats \<and> projector (meas_outcome_prj (M j))"
- proof (intro allI impI conjI)
- fix j
- assume "j < n"
- show "meas_outcome_prj (M j) \<in> fc_mats" using \<open>M = mk_meas_outcome B U\<close> assms
- fc_mats_carrier \<open>j < n\<close> \<open>n = dist_el_card B\<close> dim_eq mk_meas_outcome_carrier
- dimB dimP unit by blast
- show "projector (meas_outcome_prj (M j))" using make_meas_outcome_prjectors
- \<open>M = mk_meas_outcome B U\<close>
- fc_mats_carrier \<open>n = dist_el_card B\<close> unit \<open>j < n\<close> dimB dimP dim_eq by blast
- qed
- show "\<forall>i<n. \<forall>j<n. i \<noteq> j \<longrightarrow> meas_outcome_prj (M i) * meas_outcome_prj (M j) =
- 0\<^sub>m dimR dimR"
- proof (intro allI impI)
- fix i
- fix j
- assume "i < n" and "j < n" and "i\<noteq> j"
- thus "meas_outcome_prj (M i) * meas_outcome_prj (M j) = 0\<^sub>m dimR dimR"
- using make_meas_outcome_prj_ortho[of U B i j] assms dimB dimP fc_mats_carrier dim_eq
- by (simp add: \<open>M = mk_meas_outcome B U\<close> \<open>n = dist_el_card B\<close> unit)
- qed
- show "sum_mat (\<lambda>j. meas_outcome_prj (M j)) {..<n} = 1\<^sub>m dimR" using
- mk_meas_outcome_sum_id
- \<open>M = mk_meas_outcome B U\<close> fc_mats_carrier dim_eq \<open>n = dist_el_card B\<close> unit dimB dimP
- by blast
- qed
-qed
-
-lemma (in cpx_sq_mat) make_pm_proj_measurement':
- assumes "A\<in> fc_mats"
- and "hermitian A"
-shows "proj_measurement (proj_meas_size (make_pm A)) (proj_meas_outcomes (make_pm A))"
- unfolding proj_meas_size_def proj_meas_outcomes_def
- by (rule make_pm_proj_measurement[of A], (simp add: assms)+)
-
-lemma (in cpx_sq_mat) make_pm_projectors:
- assumes "A\<in> fc_mats"
- and "hermitian A"
-and "i < proj_meas_size (make_pm A)"
-shows "projector (meas_outcome_prj (proj_meas_outcomes (make_pm A) i))"
-proof -
- have "proj_measurement (proj_meas_size (make_pm A)) (proj_meas_outcomes (make_pm A))"
- using assms make_pm_proj_measurement' by simp
- thus ?thesis using proj_measurement_proj assms by simp
-qed
-
-lemma (in cpx_sq_mat) make_pm_square:
- assumes "A\<in> fc_mats"
- and "hermitian A"
-and "i < proj_meas_size (make_pm A)"
-shows "meas_outcome_prj (proj_meas_outcomes (make_pm A) i) \<in> fc_mats"
-proof -
- have "proj_measurement (proj_meas_size (make_pm A)) (proj_meas_outcomes (make_pm A))"
- using assms make_pm_proj_measurement' by simp
- thus ?thesis using proj_measurement_square assms by simp
-qed
-
-
-subsubsection \<open>Properties on the spectrum of a Hermitian matrix\<close>
-
-lemma (in cpx_sq_mat) hermitian_schur_decomp':
- assumes "hermitian A"
- and "A\<in> fc_mats"
-obtains B U where "hermitian_decomp A B U \<and>
- make_pm A = (dist_el_card B, mk_meas_outcome B U)"
-proof -
- {
- have es: "char_poly A = (\<Prod> (e :: complex) \<leftarrow> (eigvals A). [:- e, 1:])"
- using assms fc_mats_carrier eigvals_poly_length dim_eq by auto
- obtain B U Q where us: "unitary_schur_decomposition A (eigvals A) = (B,U,Q)"
- by (cases "unitary_schur_decomposition A (eigvals A)")
- hence pr: "similar_mat_wit A B U (Complex_Matrix.adjoint U) \<and> diagonal_mat B \<and>
- diag_mat B = (eigvals A) \<and> unitary U \<and> (\<forall>i < dimR. B$$(i, i) \<in> Reals)"
- using hermitian_eigenvalue_real assms fc_mats_carrier es dim_eq by auto
- moreover have "dim_row B = dimR" using assms fc_mats_carrier dim_eq similar_mat_wit_dim_row[of A]
- pr by auto
- ultimately have "hermitian_decomp A B U" unfolding hermitian_decomp_def by simp
- moreover have "make_pm A = (dist_el_card B, mk_meas_outcome B U)" using us Let_def
- unfolding make_pm_def hermitian_decomp_def by simp
- ultimately have "\<exists> B U. hermitian_decomp A B U \<and>
- make_pm A = (dist_el_card B, mk_meas_outcome B U)" by auto
- }
- thus ?thesis using that by auto
-qed
-
-lemma (in cpx_sq_mat) spectrum_meas_outcome_val_eq:
- assumes "hermitian A"
- and "A \<in> fc_mats"
-and "make_pm A = (p, M)"
-shows "spectrum A = (\<lambda>i. meas_outcome_val (M i)) `{..< p}"
-proof -
- obtain B U where bu: "hermitian_decomp A B U \<and>
- make_pm A = (dist_el_card B, mk_meas_outcome B U)"
- using assms hermitian_schur_decomp'[OF assms(1)] by auto
- hence "p = dist_el_card B" using assms by simp
- have dimB: "B \<in> carrier_mat dimR dimR" using hermitian_decomp_dim_carrier[of A] dim_eq bu assms
- fc_mats_carrier by auto
- have Bvals: "diag_mat B = eigvals A" using bu hermitian_decomp_eigenvalues[of A B U] by simp
- have Bii: "\<And>i. i < dimR \<Longrightarrow> B$$(i, i) \<in> Reals" using bu hermitian_decomp_real_eigvals[of A B U]
- dimB by simp
- have "diag_elems B = set (eigvals A)" using dimB Bvals diag_elems_set_diag_mat[of B] by simp
- have "M = mk_meas_outcome B U" using assms bu by simp
- {
- fix i
- assume "i < p"
- have "meas_outcome_val (M i) = Re (diag_idx_to_el B i)"
- using \<open>M = mk_meas_outcome B U\<close>
- unfolding meas_outcome_val_def mk_meas_outcome_def by simp
- also have "... = diag_idx_to_el B i" using diag_idx_to_el_real dimB \<open>i < p\<close>
- \<open>p = dist_el_card B\<close> Bii by simp
- finally have "meas_outcome_val (M i) = diag_idx_to_el B i" .
- } note eq = this
- have "bij_betw (diag_idx_to_el B) {..<dist_el_card B} (diag_elems B)"
- using diag_idx_to_el_bij[of B] by simp
- hence "diag_idx_to_el B ` {..< p} = diag_elems B" using \<open>p = dist_el_card B\<close>
- unfolding bij_betw_def by simp
- thus ?thesis using eq \<open>diag_elems B = set (eigvals A)\<close> unfolding spectrum_def by auto
-qed
-
-lemma (in cpx_sq_mat) spectrum_meas_outcome_val_eq':
- assumes "hermitian A"
- and "A \<in> fc_mats"
-and "make_pm A = (p, M)"
-shows "{Re x|x. x\<in> spectrum A} = (\<lambda>i. meas_outcome_val (M i)) `{..< p}"
-proof
- show "{Re x |x. x \<in> spectrum A} \<subseteq> (\<lambda>i. meas_outcome_val (M i)) ` {..<p}"
- using spectrum_meas_outcome_val_eq assms by force
- show "(\<lambda>i. meas_outcome_val (M i)) ` {..<p} \<subseteq> {Re x |x. x \<in> spectrum A}"
- using spectrum_meas_outcome_val_eq assms by force
-qed
-
-lemma (in cpx_sq_mat) make_pm_eigenvalues:
- assumes "A\<in> fc_mats"
- and "hermitian A"
-and "i < proj_meas_size (make_pm A)"
-shows "meas_outcome_val (proj_meas_outcomes (make_pm A) i) \<in> spectrum A"
- using spectrum_meas_outcome_val_eq[of A] assms make_pm_decomp by auto
-
-lemma (in cpx_sq_mat) make_pm_spectrum:
- assumes "A\<in> fc_mats"
- and "hermitian A"
- and "make_pm A = (p,M)"
-shows "(\<lambda>i. meas_outcome_val (proj_meas_outcomes (make_pm A) i)) ` {..< p} = spectrum A"
-proof
- show "(\<lambda>x. complex_of_real (meas_outcome_val (proj_meas_outcomes (make_pm A) x))) ` {..<p} \<subseteq>
- spectrum A"
- using assms make_pm_eigenvalues make_pm_decomp
- by (metis (mono_tags, lifting) Pair_inject image_subsetI lessThan_iff)
- show "spectrum A \<subseteq>
- (\<lambda>x. complex_of_real (meas_outcome_val (proj_meas_outcomes (make_pm A) x))) ` {..<p}"
- by (metis Pair_inject assms equalityE make_pm_decomp spectrum_meas_outcome_val_eq)
-qed
-
-lemma (in cpx_sq_mat) spectrum_size:
- assumes "hermitian A"
- and "A\<in> fc_mats"
-and "make_pm A = (p, M)"
-shows "p = card (spectrum A)"
-proof -
- obtain B U where bu: "hermitian_decomp A B U \<and>
- make_pm A = (dist_el_card B, mk_meas_outcome B U)"
- using assms hermitian_schur_decomp'[OF assms(1)] by auto
- hence "p = dist_el_card B" using assms by simp
- have "spectrum A = set (diag_mat B)" using bu hermitian_decomp_eigenvalues[of A B U]
- unfolding spectrum_def by simp
- also have "... = diag_elems B" using diag_elems_set_diag_mat[of B] by simp
- finally have "spectrum A = diag_elems B" .
- thus ?thesis using \<open>p = dist_el_card B\<close> unfolding dist_el_card_def by simp
-qed
-
-lemma (in cpx_sq_mat) spectrum_size':
- assumes "hermitian A"
-and "A\<in> fc_mats"
-shows "proj_meas_size (make_pm A) = card (spectrum A)" using spectrum_size
- unfolding proj_meas_size_def
- by (metis assms fst_conv surj_pair)
-
-lemma (in cpx_sq_mat) make_pm_projectors':
- assumes "hermitian A"
- and "A \<in> fc_mats"
- and "a<card (spectrum A)"
-shows "projector (meas_outcome_prj (proj_meas_outcomes (make_pm A) a))"
- by (rule make_pm_projectors, (simp add: assms spectrum_size')+)
-
-lemma (in cpx_sq_mat) meas_outcome_prj_carrier:
- assumes "hermitian A"
- and "A \<in> fc_mats"
- and "a<card (spectrum A)"
-shows "meas_outcome_prj (proj_meas_outcomes (make_pm A) a) \<in> fc_mats"
-proof (rule proj_measurement_square)
- show "proj_measurement (proj_meas_size (make_pm A)) (proj_meas_outcomes (make_pm A))"
- using make_pm_proj_measurement' assms by simp
- show "a < proj_meas_size (make_pm A)" using assms
- spectrum_size[of _ "proj_meas_size (make_pm A)"] make_pm_decomp[of A] by simp
-qed
-
-lemma (in cpx_sq_mat) meas_outcome_prj_trace_real:
- assumes "hermitian A"
- and "density_operator R"
- and "R \<in> fc_mats"
- and "A\<in> fc_mats"
- and "a<card (spectrum A)"
-shows "Re (Complex_Matrix.trace (R * meas_outcome_prj (proj_meas_outcomes (make_pm A) a))) =
- Complex_Matrix.trace (R * meas_outcome_prj (proj_meas_outcomes (make_pm A) a))"
-proof (rule trace_proj_pos_real)
- show "R \<in> carrier_mat dimR dimR" using fc_mats_carrier assms dim_eq by simp
- show "Complex_Matrix.positive R" using assms unfolding density_operator_def by simp
- show "projector (meas_outcome_prj (proj_meas_outcomes (make_pm A) a))" using assms
- make_pm_projectors' by simp
- show "meas_outcome_prj (proj_meas_outcomes (make_pm A) a) \<in> carrier_mat dimR dimR"
- using meas_outcome_prj_carrier assms
- dim_eq fc_mats_carrier by simp
-qed
-
-lemma (in cpx_sq_mat) sum_over_spectrum:
- assumes "A\<in> fc_mats"
-and "hermitian A"
-and "make_pm A = (p, M)"
-shows "(\<Sum> y \<in> spectrum A. f y) = (\<Sum> i < p. f (meas_outcome_val (M i)))"
-proof (rule sum.reindex_cong)
-show "spectrum A =(\<lambda>x. (meas_outcome_val (M x)))` {..< p}"
- using spectrum_meas_outcome_val_eq assms by simp
- show "inj_on (\<lambda>x. complex_of_real (meas_outcome_val (M x))) {..<p}"
- proof -
- have "inj_on (\<lambda>x. (meas_outcome_val (M x))) {..<p}"
- using make_pm_proj_measurement[of A p M] assms proj_measurement_inj by simp
- thus ?thesis by (simp add: inj_on_def)
- qed
-qed simp
-
-lemma (in cpx_sq_mat) sum_over_spectrum':
- assumes "A\<in> fc_mats"
-and "hermitian A"
-and "make_pm A = (p, M)"
-shows "(\<Sum> y \<in> {Re x|x. x \<in> spectrum A}. f y) = (\<Sum> i < p. f (meas_outcome_val (M i)))"
-proof (rule sum.reindex_cong)
- show "{Re x|x. x \<in> spectrum A} =(\<lambda>x. (meas_outcome_val (M x)))` {..< p}"
- using spectrum_meas_outcome_val_eq' assms by simp
- show "inj_on (\<lambda>x. (meas_outcome_val (M x))) {..<p}" using make_pm_proj_measurement[of A p M]
- assms proj_measurement_inj by simp
-qed simp
-
-text \<open>When a matrix $A$ is decomposed into a projective measurement $\{(\lambda_a, \pi_a)\}$, it
-can be recovered by the formula $A = \sum \lambda_a \pi_a$.\<close>
-
-lemma (in cpx_sq_mat) make_pm_sum:
- assumes "A \<in> fc_mats"
- and "hermitian A"
- and "make_pm A = (p, M)"
-shows "sum_mat (\<lambda>i. (meas_outcome_val (M i)) \<cdot>\<^sub>m meas_outcome_prj (M i)) {..< p} = A"
-proof -
- have es: "char_poly A = (\<Prod> (e :: complex) \<leftarrow> (eigvals A). [:- e, 1:])"
- using assms fc_mats_carrier eigvals_poly_length dim_eq by auto
- obtain B U Q where us: "unitary_schur_decomposition A (eigvals A) = (B,U,Q)"
- by (cases "unitary_schur_decomposition A (eigvals A)", auto)
- then have "similar_mat_wit A B U (Complex_Matrix.adjoint U) \<and> diagonal_mat B \<and>
- diag_mat B = (eigvals A)
- \<and> unitary U \<and> (\<forall>i < dimR. B$$(i, i) \<in> Reals)"
- using hermitian_eigenvalue_real assms fc_mats_carrier es dim_eq by auto
- hence A: "A = U * B * (Complex_Matrix.adjoint U)" and dB: "diagonal_mat B"
- and Bii: "\<And>i. i < dimR \<Longrightarrow> B$$(i, i) \<in> Reals"
- and dimB: "B \<in> carrier_mat dimR dimR" and dimP: "U \<in> carrier_mat dimR dimR" and
- dimaP: "Complex_Matrix.adjoint U \<in> carrier_mat dimR dimR" and unit: "unitary U"
- unfolding similar_mat_wit_def Let_def using assms fc_mats_carrier by auto
- hence mpeq: "make_pm A = (dist_el_card B, mk_meas_outcome B U)" using us Let_def
- unfolding make_pm_def by simp
- hence "p = dist_el_card B" using assms by simp
- have "M = mk_meas_outcome B U" using assms mpeq by simp
- have "sum_mat (\<lambda>i. meas_outcome_val (M i) \<cdot>\<^sub>m meas_outcome_prj (M i)) {..< p} =
- sum_mat (\<lambda>j. Re (diag_idx_to_el B j)\<cdot>\<^sub>m project_vecs (\<lambda>i. zero_col U i)
- (diag_elem_indices (diag_idx_to_el B j) B)) {..<(dist_el_card B)}"
- using \<open>p = dist_el_card B\<close>
- \<open>M = mk_meas_outcome B U\<close> unfolding meas_outcome_val_def meas_outcome_prj_def
- mk_meas_outcome_def by simp
- also have "... = sum_mat
- (\<lambda>j. sum_mat (\<lambda>i. (Re (diag_idx_to_el B j)) \<cdot>\<^sub>m rank_1_proj (zero_col U i))
- (diag_elem_indices (diag_idx_to_el B j) B)) {..<dist_el_card B}"
- unfolding project_vecs_def
- proof (rule sum_mat_cong)
- show "finite {..<dist_el_card B}" by simp
- show "\<And>i. i \<in> {..<dist_el_card B} \<Longrightarrow>
- complex_of_real (Re (diag_idx_to_el B i)) \<cdot>\<^sub>m
- sum_mat (\<lambda>i. rank_1_proj (zero_col U i)) (diag_elem_indices (diag_idx_to_el B i) B)
- \<in> fc_mats" using assms fc_mats_carrier dimP project_vecs_def project_vecs_dim zero_col_dim
- dim_eq by auto
- show "\<And>i. i \<in> {..<dist_el_card B} \<Longrightarrow>
- sum_mat (\<lambda>ia. complex_of_real (Re (diag_idx_to_el B i)) \<cdot>\<^sub>m rank_1_proj (zero_col U ia))
- (diag_elem_indices (diag_idx_to_el B i) B) \<in> fc_mats" using assms fc_mats_carrier dimP
- project_vecs_def project_vecs_dim zero_col_dim dim_eq
- by (metis (no_types, lifting) rank_1_proj_carrier cpx_sq_mat_smult sum_mat_carrier)
- show "\<And>j. j \<in> {..<dist_el_card B} \<Longrightarrow>
- (Re (diag_idx_to_el B j)) \<cdot>\<^sub>m sum_mat (\<lambda>i. rank_1_proj (zero_col U i))
- (diag_elem_indices (diag_idx_to_el B j) B) =
- sum_mat (\<lambda>i. complex_of_real (Re (diag_idx_to_el B j)) \<cdot>\<^sub>m rank_1_proj (zero_col U i))
- (diag_elem_indices (diag_idx_to_el B j) B)"
- proof -
- fix j
- assume "j \<in> {..<dist_el_card B}"
- show " (Re (diag_idx_to_el B j)) \<cdot>\<^sub>m sum_mat (\<lambda>i. rank_1_proj (zero_col U i))
- (diag_elem_indices (diag_idx_to_el B j) B) =
- sum_mat (\<lambda>i. (Re (diag_idx_to_el B j)) \<cdot>\<^sub>m rank_1_proj (zero_col U i))
- (diag_elem_indices (diag_idx_to_el B j) B)"
- proof (rule smult_sum_mat)
- show "finite (diag_elem_indices (diag_idx_to_el B j) B)"
- using diag_elem_indices_finite[of _ B] by simp
- show "\<And>i. i \<in> diag_elem_indices (diag_idx_to_el B j) B \<Longrightarrow>
- rank_1_proj (zero_col U i) \<in> fc_mats"
- using dim_eq by (metis dimP local.fc_mats_carrier rank_1_proj_carrier zero_col_dim)
- qed
- qed
- qed
- also have "... = sum_mat
- (\<lambda>j. sum_mat (\<lambda>ia. (diag_mat B ! ia) \<cdot>\<^sub>m rank_1_proj (zero_col U ia))
- (diag_elem_indices (diag_idx_to_el B j) B)) {..<dist_el_card B}"
- proof (rule sum_mat_cong)
- show "finite {..<dist_el_card B}" by simp
- show "\<And>i. i \<in> {..<dist_el_card B} \<Longrightarrow>
- sum_mat (\<lambda>ia. complex_of_real (Re (diag_idx_to_el B i)) \<cdot>\<^sub>m rank_1_proj (zero_col U ia))
- (diag_elem_indices (diag_idx_to_el B i) B) \<in> fc_mats" using assms fc_mats_carrier dimP
- project_vecs_def project_vecs_dim zero_col_dim dim_eq
- by (metis (no_types, lifting) rank_1_proj_carrier cpx_sq_mat_smult sum_mat_carrier)
- show "\<And>i. i \<in> {..<dist_el_card B} \<Longrightarrow>
- local.sum_mat (\<lambda>ia. (diag_mat B ! ia) \<cdot>\<^sub>m rank_1_proj (zero_col U ia))
- (diag_elem_indices (diag_idx_to_el B i) B) \<in> fc_mats" using assms fc_mats_carrier dimP
- project_vecs_def project_vecs_dim zero_col_dim dim_eq
- by (metis (no_types, lifting) rank_1_proj_carrier cpx_sq_mat_smult sum_mat_carrier)
- show "\<And>i. i \<in> {..<dist_el_card B} \<Longrightarrow>
- sum_mat (\<lambda>ia. (Re (diag_idx_to_el B i)) \<cdot>\<^sub>m rank_1_proj (zero_col U ia))
- (diag_elem_indices (diag_idx_to_el B i) B) =
- sum_mat (\<lambda>ia. (diag_mat B ! ia) \<cdot>\<^sub>m rank_1_proj (zero_col U ia))
- (diag_elem_indices (diag_idx_to_el B i) B)"
- proof -
- fix i
- assume "i\<in> {..< dist_el_card B}"
- show "sum_mat (\<lambda>ia. (Re (diag_idx_to_el B i)) \<cdot>\<^sub>m rank_1_proj (zero_col U ia))
- (diag_elem_indices (diag_idx_to_el B i) B) =
- sum_mat (\<lambda>ia. (diag_mat B ! ia) \<cdot>\<^sub>m rank_1_proj (zero_col U ia))
- (diag_elem_indices (diag_idx_to_el B i) B)"
- proof (rule sum_mat_cong)
- show "finite (diag_elem_indices (diag_idx_to_el B i) B)"
- using diag_elem_indices_finite[of _ B] by simp
- show "\<And>ia. ia \<in> diag_elem_indices (diag_idx_to_el B i) B \<Longrightarrow>
- (Re (diag_idx_to_el B i)) \<cdot>\<^sub>m rank_1_proj (zero_col U ia) \<in> fc_mats"
- using assms fc_mats_carrier dimP project_vecs_def project_vecs_dim zero_col_dim dim_eq
- by (metis (no_types, lifting) rank_1_proj_carrier cpx_sq_mat_smult)
- show "\<And>ia. ia \<in> diag_elem_indices (diag_idx_to_el B i) B \<Longrightarrow>
- (diag_mat B !ia) \<cdot>\<^sub>m rank_1_proj (zero_col U ia) \<in> fc_mats"
- using assms fc_mats_carrier dimP project_vecs_def project_vecs_dim zero_col_dim dim_eq
- by (metis (no_types, lifting) rank_1_proj_carrier cpx_sq_mat_smult)
- show "\<And>ia. ia \<in> diag_elem_indices (diag_idx_to_el B i) B \<Longrightarrow>
- (Re (diag_idx_to_el B i)) \<cdot>\<^sub>m rank_1_proj (zero_col U ia) =
- (diag_mat B ! ia) \<cdot>\<^sub>m rank_1_proj (zero_col U ia)"
- proof -
- fix ia
- assume ia: "ia \<in> diag_elem_indices (diag_idx_to_el B i) B"
- hence "ia < dim_row B" using diag_elem_indices_elem[of ia _ B] by simp
- have "Re (diag_idx_to_el B i) = Re (B $$ (ia, ia))"
- using diag_elem_indices_elem[of ia _ B] ia by simp
- also have "... = B $$ (ia, ia)" using Bii using \<open>ia < dim_row B\<close> dimB of_real_Re by blast
- also have "... = diag_mat B ! ia" using \<open>ia < dim_row B\<close> unfolding diag_mat_def by simp
- finally have "Re (diag_idx_to_el B i) = diag_mat B ! ia" .
- thus "(Re (diag_idx_to_el B i)) \<cdot>\<^sub>m rank_1_proj (zero_col U ia) =
- (diag_mat B ! ia) \<cdot>\<^sub>m rank_1_proj (zero_col U ia)" by simp
- qed
- qed
- qed
- qed
- also have "... = sum_mat
- (\<lambda>i. (diag_mat B ! i) \<cdot>\<^sub>m rank_1_proj (zero_col U i))
- (\<Union>j<dist_el_card B. diag_elem_indices (diag_idx_to_el B j) B)"
- unfolding project_vecs_def sum_mat_def
- proof (rule disj_family_sum_with[symmetric])
- show "finite {..<dist_el_card B}" by simp
- show "\<forall>j. (diag_mat B ! j) \<cdot>\<^sub>m rank_1_proj (zero_col U j) \<in> fc_mats" using assms fc_mats_carrier dimP
- project_vecs_def project_vecs_dim zero_col_dim dim_eq
- by (metis (no_types, lifting) rank_1_proj_carrier cpx_sq_mat_smult)
- show "\<And>i. i \<in> {..<dist_el_card B} \<Longrightarrow> finite (diag_elem_indices (diag_idx_to_el B i) B)"
- by (simp add: diag_elem_indices_finite)
- show "disjoint_family_on (\<lambda>n. diag_elem_indices (diag_idx_to_el B n) B)
- {..<dist_el_card B}"
- using diag_elem_indices_disjoint[of B] dimB dim_eq by simp
- qed
- also have "... = sum_mat (\<lambda>i. (diag_mat B ! i) \<cdot>\<^sub>m rank_1_proj (zero_col U i)) {..< dimR}"
- using diag_elem_indices_union[of B] dimB dim_eq by simp
- also have "... = sum_mat (\<lambda>i. (diag_mat B ! i) \<cdot>\<^sub>mrank_1_proj (Matrix.col U i)) {..< dimR}"
- proof (rule sum_mat_cong, simp)
- show res: "\<And>i. i \<in> {..<dimR} \<Longrightarrow> (diag_mat B ! i) \<cdot>\<^sub>m rank_1_proj (zero_col U i) \<in> fc_mats"
- using assms fc_mats_carrier dimP project_vecs_def project_vecs_dim zero_col_dim dim_eq
- by (metis (no_types, lifting) rank_1_proj_carrier cpx_sq_mat_smult)
- show "\<And>i. i \<in> {..<dimR} \<Longrightarrow> (diag_mat B ! i) \<cdot>\<^sub>m rank_1_proj (Matrix.col U i) \<in> fc_mats"
- using assms fc_mats_carrier dimP project_vecs_def project_vecs_dim zero_col_dim
- by (metis res lessThan_iff zero_col_col)
- show "\<And>i. i \<in> {..<dimR} \<Longrightarrow> (diag_mat B ! i) \<cdot>\<^sub>m rank_1_proj (zero_col U i) =
- (diag_mat B ! i) \<cdot>\<^sub>m rank_1_proj (Matrix.col U i)"
- by (simp add: zero_col_col)
- qed
- also have "... = U * B * Complex_Matrix.adjoint U"
- proof (rule weighted_sum_rank_1_proj_unitary)
- show "diagonal_mat B" using dB .
- show "Complex_Matrix.unitary U" using unit .
- show "U \<in> fc_mats" using fc_mats_carrier dim_eq dimP by simp
- show "B\<in> fc_mats" using fc_mats_carrier dim_eq dimB by simp
- qed
- also have "... = A" using A by simp
- finally show ?thesis .
-qed
-
-lemma (in cpx_sq_mat) trace_hermitian_pos_real:
- fixes f::"'a \<Rightarrow> real"
- assumes "hermitian A"
- and "Complex_Matrix.positive R"
- and "A \<in> fc_mats"
- and "R \<in> fc_mats"
-shows "Complex_Matrix.trace (R * A) =
- Re (Complex_Matrix.trace (R * A))"
-proof -
- define prj_mems where "prj_mems = make_pm A"
- define p where "p = proj_meas_size prj_mems"
- define meas where "meas = proj_meas_outcomes prj_mems"
- have tre: "Complex_Matrix.trace (R * A) =
- Complex_Matrix.trace (R *
- sum_mat (\<lambda>i. (meas_outcome_val (meas i)) \<cdot>\<^sub>m meas_outcome_prj (meas i)) {..< p})"
- using make_pm_sum assms meas_def p_def proj_meas_size_def proj_meas_outcomes_def prj_mems_def
- meas_outcome_val_def meas_outcome_prj_def by auto
- also have "... = Re (Complex_Matrix.trace (R *
- sum_mat (\<lambda>i. (meas_outcome_val (meas i)) \<cdot>\<^sub>m meas_outcome_prj (meas i)) {..< p}))"
- proof (rule trace_sum_mat_proj_pos_real, (auto simp add: assms))
- fix i
- assume "i < p"
- thus "projector (meas_outcome_prj (meas i))" using make_pm_projectors assms
- unfolding p_def meas_def prj_mems_def by simp
- show "meas_outcome_prj (meas i) \<in> fc_mats" using make_pm_square assms \<open>i < p\<close>
- unfolding p_def meas_def prj_mems_def by simp
- qed
- also have "... = Re (Complex_Matrix.trace (R * A))" using tre by simp
- finally show ?thesis .
-qed
-
-lemma (in cpx_sq_mat) hermitian_Re_spectrum:
- assumes "hermitian A"
-and "A\<in> fc_mats"
-and "{Re x |x. x \<in> spectrum A} = {a,b}"
-shows "spectrum A = {complex_of_real a, complex_of_real b}"
-proof
- have ar: "\<And>(a::complex). a \<in> spectrum A \<Longrightarrow> Re a = a" using hermitian_spectrum_real assms by simp
- show "spectrum A \<subseteq> {complex_of_real a, complex_of_real b}"
- proof
- fix x
- assume "x\<in> spectrum A"
- hence "Re x = x" using ar by simp
- have "Re x \<in> {a,b}" using \<open>x\<in> spectrum A\<close> assms by blast
- thus "x \<in> {complex_of_real a, complex_of_real b}" using \<open>Re x = x\<close> by auto
- qed
- show "{complex_of_real a, complex_of_real b} \<subseteq> spectrum A"
- proof
- fix x
- assume "x \<in> {complex_of_real a, complex_of_real b}"
- hence "x \<in> {complex_of_real (Re x) |x. x \<in> spectrum A}" using assms
- proof -
- have "\<And>r. r \<notin> {a, b} \<or> (\<exists>c. r = Re c \<and> c \<in> spectrum A)"
- using \<open>{Re x |x. x \<in> spectrum A} = {a, b}\<close> by blast
- then show ?thesis
- using \<open>x \<in> {complex_of_real a, complex_of_real b}\<close> by blast
- qed
- thus "x\<in> spectrum A" using ar by auto
- qed
-qed
-
-
-subsubsection \<open>Similar properties for eigenvalues rather than spectrum indices\<close>
-
-text \<open>In this part we go the other way round: given an eigenvalue of $A$, \verb+spectrum_to_pm_idx+
-permits to retrieve its index in the associated projective measurement\<close>
-
-definition (in cpx_sq_mat) spectrum_to_pm_idx
- where "spectrum_to_pm_idx A x = (let (B,U,_) = unitary_schur_decomposition A (eigvals A) in
- diag_el_to_idx B x)"
-
-lemma (in cpx_sq_mat) spectrum_to_pm_idx_bij:
-assumes "hermitian A"
- and "A\<in> fc_mats"
-shows "bij_betw (spectrum_to_pm_idx A) (spectrum A) {..< card (spectrum A)}"
-proof -
- define p where "p = proj_meas_size (make_pm A)"
- define M where "M = proj_meas_outcomes (make_pm A)"
- have es: "char_poly A = (\<Prod> (e :: complex) \<leftarrow> (eigvals A). [:- e, 1:])"
- using assms fc_mats_carrier eigvals_poly_length dim_eq by auto
- obtain B U Q where us: "unitary_schur_decomposition A (eigvals A) = (B,U,Q)"
- by (cases "unitary_schur_decomposition A (eigvals A)")
- hence pr: "similar_mat_wit A B U (Complex_Matrix.adjoint U) \<and>
- diag_mat B = (eigvals A)"
- using hermitian_eigenvalue_real assms fc_mats_carrier es dim_eq by auto
- have "(p,M) = make_pm A" unfolding p_def M_def using make_pm_decomp by simp
- hence "p = dist_el_card B" using assms us unfolding make_pm_def by simp
- have dimB: "B \<in> carrier_mat dimR dimR" using dim_eq pr assms
- fc_mats_carrier unfolding similar_mat_wit_def by auto
- have Bvals: "diag_mat B = eigvals A" using pr hermitian_decomp_eigenvalues[of A B U] by simp
- have "diag_elems B = spectrum A" unfolding spectrum_def using dimB Bvals
- diag_elems_set_diag_mat[of B] by simp
- moreover have "dist_el_card B = card (spectrum A)" using spectrum_size[of A p M] assms
- \<open>(p,M) = make_pm A\<close> \<open>p = dist_el_card B\<close> by simp
- ultimately show "bij_betw (spectrum_to_pm_idx A) (spectrum A) {..< card (spectrum A)}"
- using diag_el_to_idx_bij us unfolding spectrum_to_pm_idx_def Let_def
- by (metis (mono_tags, lifting) bij_betw_cong case_prod_conv)
-qed
-
-lemma (in cpx_sq_mat) spectrum_to_pm_idx_lt_card:
-assumes "A\<in> fc_mats"
- and "hermitian A"
-and "a\<in> spectrum A"
-shows "spectrum_to_pm_idx A a < card (spectrum A)" using spectrum_to_pm_idx_bij[of A] assms
- by (meson bij_betw_apply lessThan_iff)
-
-lemma (in cpx_sq_mat) spectrum_to_pm_idx_inj:
-assumes "hermitian A"
- and "A\<in> fc_mats"
-shows "inj_on (spectrum_to_pm_idx A) (spectrum A)" using assms spectrum_to_pm_idx_bij
- unfolding bij_betw_def by simp
-
-lemma (in cpx_sq_mat) spectrum_meas_outcome_val_inv:
-assumes "A\<in> fc_mats"
- and "hermitian A"
-and "make_pm A = (p,M)"
-and "i < p"
-shows "spectrum_to_pm_idx A (meas_outcome_val (M i)) = i"
-proof -
- have es: "char_poly A = (\<Prod> (e :: complex) \<leftarrow> (eigvals A). [:- e, 1:])"
- using assms fc_mats_carrier eigvals_poly_length dim_eq by auto
- obtain B U Q where us: "unitary_schur_decomposition A (eigvals A) = (B,U,Q)"
- by (cases "unitary_schur_decomposition A (eigvals A)")
- hence pr: "similar_mat_wit A B U (Complex_Matrix.adjoint U) \<and>
- diag_mat B = (eigvals A) \<and> (\<forall>i < dimR. B$$(i, i) \<in> Reals)"
- using hermitian_eigenvalue_real assms fc_mats_carrier es dim_eq by auto
- have "dim_row B = dimR" using assms fc_mats_carrier dim_eq similar_mat_wit_dim_row[of A]
- pr by auto
- hence "(p,M) = (dist_el_card B, mk_meas_outcome B U)" using assms us
- unfolding make_pm_def by simp
- hence "M = mk_meas_outcome B U" by simp
- have "meas_outcome_val (M i) = Re (diag_idx_to_el B i)"
- using \<open>M = mk_meas_outcome B U\<close> unfolding mk_meas_outcome_def
- meas_outcome_val_def by simp
- also have "... = diag_idx_to_el B i" using pr
- \<open>(p, M) = (dist_el_card B, mk_meas_outcome B U)\<close> \<open>dim_row B = dimR\<close> assms
- diag_idx_to_el_real by auto
- finally have "meas_outcome_val (M i) = diag_idx_to_el B i" .
- hence "spectrum_to_pm_idx A (meas_outcome_val (M i)) =
- spectrum_to_pm_idx A (diag_idx_to_el B i)" by simp
- also have "... = diag_el_to_idx B (diag_idx_to_el B i)" unfolding spectrum_to_pm_idx_def
- using us by simp
- also have "... = i" using assms unfolding diag_el_to_idx_def
- using \<open>(p, M) = (dist_el_card B, mk_meas_outcome B U)\<close> bij_betw_inv_into_left
- diag_idx_to_el_bij by fastforce
- finally show ?thesis .
-qed
-
-lemma (in cpx_sq_mat) meas_outcome_val_spectrum_inv:
- assumes "A\<in> fc_mats"
- and "hermitian A"
-and "x\<in> spectrum A"
-and "make_pm A = (p,M)"
-shows "meas_outcome_val (M (spectrum_to_pm_idx A x)) = x"
-proof -
- have es: "char_poly A = (\<Prod> (e :: complex) \<leftarrow> (eigvals A). [:- e, 1:])"
- using assms fc_mats_carrier eigvals_poly_length dim_eq by auto
- obtain B U Q where us: "unitary_schur_decomposition A (eigvals A) = (B,U,Q)"
- by (cases "unitary_schur_decomposition A (eigvals A)")
- hence pr: "similar_mat_wit A B U (Complex_Matrix.adjoint U) \<and> diagonal_mat B \<and>
- diag_mat B = (eigvals A) \<and> unitary U \<and> (\<forall>i < dimR. B$$(i, i) \<in> Reals)"
- using hermitian_eigenvalue_real assms fc_mats_carrier es dim_eq by auto
- have "dim_row B = dimR" using assms fc_mats_carrier dim_eq similar_mat_wit_dim_row[of A]
- pr by auto
- hence "(p,M) = (dist_el_card B, mk_meas_outcome B U)" using assms us
- unfolding make_pm_def by simp
- hence "M = mk_meas_outcome B U" by simp
- have "diag_mat B = (eigvals A)" using pr by simp
- hence "x \<in> set (diag_mat B)" using \<open>diag_mat B = eigvals A\<close> assms unfolding spectrum_def by simp
- hence "x \<in> diag_elems B" using diag_elems_set_diag_mat[of B] by simp
- hence "diag_idx_to_el B (diag_el_to_idx B x) = x" unfolding diag_el_to_idx_def
- by (meson bij_betw_inv_into_right diag_idx_to_el_bij)
- moreover have "spectrum_to_pm_idx A x = diag_el_to_idx B x" unfolding spectrum_to_pm_idx_def
- using us by simp
- moreover have "meas_outcome_val (M (spectrum_to_pm_idx A x)) =
- Re (diag_idx_to_el B (diag_el_to_idx B x))" using \<open>M = mk_meas_outcome B U\<close>
- unfolding mk_meas_outcome_def meas_outcome_val_def by (simp add: calculation(2))
- ultimately show ?thesis by simp
-qed
-
-definition (in cpx_sq_mat) eigen_projector where
-"eigen_projector A a =
- meas_outcome_prj ((proj_meas_outcomes (make_pm A)) (spectrum_to_pm_idx A a))"
-
-lemma (in cpx_sq_mat) eigen_projector_carrier:
- assumes "A\<in> fc_mats"
- and "a\<in> spectrum A"
- and "hermitian A"
-shows "eigen_projector A a \<in> fc_mats" unfolding eigen_projector_def
- using meas_outcome_prj_carrier[of A "spectrum_to_pm_idx A a"]
- spectrum_to_pm_idx_lt_card[of A a] assms by simp
-
-text \<open>We obtain the following result, which is similar to \verb+make_pm_sum+ but with a sum on
-the elements of the spectrum of Hermitian matrix $A$, which is a more standard statement of the
-spectral decomposition theorem.\<close>
-
-lemma (in cpx_sq_mat) make_pm_sum':
- assumes "A \<in> fc_mats"
- and "hermitian A"
-shows "sum_mat (\<lambda>a. a \<cdot>\<^sub>m (eigen_projector A a)) (spectrum A) = A"
-proof -
- define p where "p = proj_meas_size (make_pm A)"
- define M where "M = proj_meas_outcomes (make_pm A)"
- have "(p,M) = make_pm A" unfolding p_def M_def using make_pm_decomp by simp
- define g where
- "g = (\<lambda>i. (if i < p
- then complex_of_real (meas_outcome_val (M i)) \<cdot>\<^sub>m meas_outcome_prj (M i)
- else (0\<^sub>m dimR dimC)))"
- have g: "\<forall>x. g x \<in> fc_mats"
- proof
- fix x
- show "g x \<in> fc_mats"
- proof (cases "x < p")
- case True
- hence "(meas_outcome_val (M x)) \<cdot>\<^sub>m meas_outcome_prj (M x) \<in> fc_mats"
- using meas_outcome_prj_carrier
- spectrum_size assms cpx_sq_mat_smult make_pm_proj_measurement proj_measurement_square
- \<open>(p,M) = make_pm A\<close> by metis
- then show ?thesis unfolding g_def using True by simp
- next
- case False
- then show ?thesis unfolding g_def using zero_mem by simp
- qed
- qed
- define h where "h = (\<lambda>a. (if a \<in> (spectrum A) then a \<cdot>\<^sub>m eigen_projector A a else (0\<^sub>m dimR dimC)))"
- have h: "\<forall>x. h x \<in> fc_mats"
- proof
- fix x
- show "h x \<in> fc_mats"
- proof (cases "x\<in> spectrum A")
- case True
- then show ?thesis unfolding h_def using eigen_projector_carrier[of A x] assms True
- by (simp add: cpx_sq_mat_smult)
- next
- case False
- then show ?thesis unfolding h_def using zero_mem by simp
- qed
- qed
- have "inj_on (spectrum_to_pm_idx A) (spectrum A)" using assms spectrum_to_pm_idx_inj by simp
- moreover have "{..<p} = spectrum_to_pm_idx A ` spectrum A" using \<open>(p,M) = make_pm A\<close>
- spectrum_to_pm_idx_bij spectrum_size unfolding bij_betw_def
- by (metis assms(1) assms(2))
- moreover have "\<And>x. x \<in> spectrum A \<Longrightarrow> g (spectrum_to_pm_idx A x) = h x"
- proof -
- fix a
- assume "a \<in> spectrum A"
- hence "Re a = a" using hermitian_spectrum_real assms by simp
- have "spectrum_to_pm_idx A a < p" using spectrum_to_pm_idx_lt_card[of A] spectrum_size assms
- \<open>a \<in> spectrum A\<close> \<open>(p,M) = make_pm A\<close> by metis
- have "g (spectrum_to_pm_idx A a) =
- (meas_outcome_val (M (spectrum_to_pm_idx A a))) \<cdot>\<^sub>m
- meas_outcome_prj (M (spectrum_to_pm_idx A a))"
- using \<open>spectrum_to_pm_idx A a < p\<close> unfolding g_def by simp
- also have "... = a \<cdot>\<^sub>m meas_outcome_prj (M (spectrum_to_pm_idx A a))"
- using meas_outcome_val_spectrum_inv[of A "Re a"] \<open>Re a = a\<close> assms \<open>a \<in> spectrum A\<close>
- \<open>(p,M) = make_pm A\<close> by metis
- also have "... = h a" using \<open>a \<in> spectrum A\<close> unfolding eigen_projector_def M_def h_def by simp
- finally show "g (spectrum_to_pm_idx A a) = h a" .
- qed
- ultimately have "sum_mat h (spectrum A) =
- sum_mat g (spectrum_to_pm_idx A ` spectrum A)" unfolding sum_mat_def
- using sum_with_reindex_cong[symmetric, of g h "spectrum_to_pm_idx A" "spectrum A" "{..< p}"]
- g h by simp
- also have "... = sum_mat g {..< p}" using \<open>{..<p} = spectrum_to_pm_idx A ` spectrum A\<close> by simp
- also have "... = sum_mat (\<lambda>i. (meas_outcome_val (M i)) \<cdot>\<^sub>m meas_outcome_prj (M i)) {..<p}"
- proof (rule sum_mat_cong, simp)
- fix i
- assume "i \<in> {..< p}"
- hence "i < p" by simp
- show "g i \<in> fc_mats" using g by simp
- show "g i = (meas_outcome_val (M i)) \<cdot>\<^sub>m meas_outcome_prj (M i)" unfolding g_def
- using \<open>i < p\<close> by simp
- show "(meas_outcome_val (M i)) \<cdot>\<^sub>m meas_outcome_prj (M i) \<in> fc_mats"
- using meas_outcome_prj_carrier spectrum_size assms cpx_sq_mat_smult
- make_pm_proj_measurement proj_measurement_square \<open>i < p\<close> \<open>(p,M) = make_pm A\<close> by metis
- qed
- also have "... = A" using make_pm_sum assms \<open>(p,M) = make_pm A\<close> unfolding g_def by simp
- finally have "sum_mat h (spectrum A) = A" .
- moreover have "sum_mat h (spectrum A) = sum_mat (\<lambda>a. a \<cdot>\<^sub>m (eigen_projector A a)) (spectrum A)"
- proof (rule sum_mat_cong)
- show "finite (spectrum A)" using spectrum_finite[of A] by simp
- fix i
- assume "i\<in> spectrum A"
- thus "h i = i \<cdot>\<^sub>m eigen_projector A i" unfolding h_def by simp
- show "h i \<in> fc_mats" using h by simp
- show "i \<cdot>\<^sub>m eigen_projector A i \<in> fc_mats" using eigen_projector_carrier[of A i] assms
- \<open>i\<in> spectrum A\<close> by (metis \<open>h i = i \<cdot>\<^sub>m eigen_projector A i\<close> h)
- qed
- ultimately show ?thesis by simp
-qed
-
-
-
+(*
+Author:
+ Mnacho Echenim, Université Grenoble Alpes
+*)
+
+theory Projective_Measurements imports
+ Linear_Algebra_Complements
+
+
+begin
+
+
+section \<open>Projective measurements\<close>
+
+text \<open>In this part we define projective measurements, also refered to as von Neumann measurements.
+ The latter are characterized by a set of orthogonal projectors, which are used to compute the
+probabilities of measure outcomes and to represent the state of the system after the measurement.\<close>
+
+text \<open>The state of the system (a density operator in this case) after a measurement is represented by
+the \verb+density_collapse+ function.\<close>
+
+subsection \<open>First definitions\<close>
+
+text \<open>We begin by defining a type synonym for couples of measurement values (reals) and the
+associated projectors (complex matrices).\<close>
+
+type_synonym measure_outcome = "real \<times> complex Matrix.mat"
+
+text \<open>The corresponding values and projectors are retrieved thanks to \verb+meas_outcome_val+
+and \verb+meas_outcome_prj+.\<close>
+
+definition meas_outcome_val::"measure_outcome \<Rightarrow> real" where
+"meas_outcome_val Mi = fst Mi"
+
+definition meas_outcome_prj::"measure_outcome \<Rightarrow> complex Matrix.mat" where
+"meas_outcome_prj Mi = snd Mi"
+
+text \<open>We define a predicate for projective measurements. A projective measurement is characterized
+by the number $p$ of possible measure outcomes, and a function $M$ mapping outcome $i$ to the
+corresponding \verb+measure_outcome+.\<close>
+
+definition (in cpx_sq_mat) proj_measurement::"nat \<Rightarrow> (nat \<Rightarrow> measure_outcome) \<Rightarrow> bool" where
+ "proj_measurement n M \<longleftrightarrow>
+ (inj_on (\<lambda>i. meas_outcome_val (M i)) {..< n}) \<and>
+ (\<forall>j < n. meas_outcome_prj (M j) \<in> fc_mats \<and>
+ projector (meas_outcome_prj (M j))) \<and>
+ (\<forall> i < n. \<forall> j < n. i \<noteq> j \<longrightarrow>
+ meas_outcome_prj (M i) * meas_outcome_prj (M j) = 0\<^sub>m dimR dimR) \<and>
+ sum_mat (\<lambda>j. meas_outcome_prj (M j)) {..< n} = 1\<^sub>m dimR"
+
+lemma (in cpx_sq_mat) proj_measurement_inj:
+ assumes "proj_measurement p M"
+ shows "inj_on (\<lambda>i. meas_outcome_val (M i)) {..< p}" using assms
+ unfolding proj_measurement_def by simp
+
+lemma (in cpx_sq_mat) proj_measurement_carrier:
+ assumes "proj_measurement p M"
+ and "i < p"
+ shows "meas_outcome_prj (M i) \<in> fc_mats" using assms
+ unfolding proj_measurement_def by simp
+
+lemma (in cpx_sq_mat) proj_measurement_ortho:
+ assumes "proj_measurement p M"
+and "i < p"
+and "j < p"
+and "i\<noteq> j"
+shows "meas_outcome_prj (M i) * meas_outcome_prj (M j) = 0\<^sub>m dimR dimR" using assms
+ unfolding proj_measurement_def by simp
+
+lemma (in cpx_sq_mat) proj_measurement_id:
+ assumes "proj_measurement p M"
+ shows "sum_mat (\<lambda>j. meas_outcome_prj (M j)) {..< p} = 1\<^sub>m dimR" using assms
+ unfolding proj_measurement_def by simp
+
+lemma (in cpx_sq_mat) proj_measurement_square:
+ assumes "proj_measurement p M"
+and "i < p"
+shows "meas_outcome_prj (M i) \<in> fc_mats" using assms unfolding proj_measurement_def by simp
+
+lemma (in cpx_sq_mat) proj_measurement_proj:
+ assumes "proj_measurement p M"
+and "i < p"
+shows "projector (meas_outcome_prj (M i))" using assms unfolding proj_measurement_def by simp
+
+text \<open>We define the probability of obtaining a measurement outcome: this is a positive number and
+the sum over all the measurement outcomes is 1.\<close>
+
+definition (in cpx_sq_mat) meas_outcome_prob :: "complex Matrix.mat \<Rightarrow>
+ (nat \<Rightarrow> real \<times> complex Matrix.mat) \<Rightarrow> nat \<Rightarrow> complex" where
+"meas_outcome_prob R M i = Complex_Matrix.trace (R* (meas_outcome_prj (M i)))"
+
+lemma (in cpx_sq_mat) meas_outcome_prob_real:
+assumes "R\<in> fc_mats"
+ and "density_operator R"
+ and "proj_measurement n M"
+ and "i < n"
+shows "meas_outcome_prob R M i \<in> \<real>"
+proof -
+ have "complex_of_real (Re (Complex_Matrix.trace (R * meas_outcome_prj (M i)))) =
+ Complex_Matrix.trace (R * meas_outcome_prj (M i))"
+ proof (rule trace_proj_pos_real)
+ show "R \<in> carrier_mat dimR dimR" using assms fc_mats_carrier dim_eq by simp
+ show "Complex_Matrix.positive R" using assms unfolding density_operator_def by simp
+ show "projector (meas_outcome_prj (M i))" using assms proj_measurement_proj by simp
+ show "meas_outcome_prj (M i) \<in> carrier_mat dimR dimR" using assms proj_measurement_carrier
+ fc_mats_carrier dim_eq by simp
+ qed
+ thus ?thesis unfolding meas_outcome_prob_def by (metis Reals_of_real)
+qed
+
+lemma (in cpx_sq_mat) meas_outcome_prob_pos:
+ assumes "R\<in> fc_mats"
+ and "density_operator R"
+ and "proj_measurement n M"
+ and "i < n"
+shows "0 \<le> meas_outcome_prob R M i" unfolding meas_outcome_prob_def
+proof (rule positive_proj_trace)
+ show "R \<in> carrier_mat dimR dimR" using assms fc_mats_carrier dim_eq by simp
+ show "Complex_Matrix.positive R" using assms unfolding density_operator_def by simp
+ show "projector (meas_outcome_prj (M i))" using assms proj_measurement_proj by simp
+ show "meas_outcome_prj (M i) \<in> carrier_mat dimR dimR" using assms proj_measurement_carrier
+ fc_mats_carrier dim_eq by simp
+qed
+
+lemma (in cpx_sq_mat) meas_outcome_prob_sum:
+ assumes "density_operator R"
+ and "R\<in> fc_mats"
+ and" proj_measurement n M"
+shows "(\<Sum> j \<in> {..< n}. meas_outcome_prob R M j) = 1"
+proof -
+ have "(\<Sum> j \<in> {..< n}. Complex_Matrix.trace (R* (meas_outcome_prj (M j)))) =
+ Complex_Matrix.trace (sum_mat (\<lambda>j. R* (meas_outcome_prj (M j))) {..< n})"
+ proof (rule trace_sum_mat[symmetric], auto)
+ fix j
+ assume "j < n"
+ thus "R * meas_outcome_prj (M j) \<in> fc_mats" using cpx_sq_mat_mult assms
+ unfolding proj_measurement_def by simp
+ qed
+ also have "... = Complex_Matrix.trace (R * (sum_mat (\<lambda>j. (meas_outcome_prj (M j))) {..< n}))"
+ proof -
+ have "sum_mat (\<lambda>j. R* (meas_outcome_prj (M j))) {..< n} =
+ R * (sum_mat (\<lambda>j. (meas_outcome_prj (M j))) {..< n})"
+ proof (rule mult_sum_mat_distrib_left, (auto simp add: assms))
+ fix j
+ assume "j < n"
+ thus "meas_outcome_prj (M j) \<in> fc_mats" using assms unfolding proj_measurement_def by simp
+ qed
+ thus ?thesis by simp
+ qed
+ also have "... = Complex_Matrix.trace (R * 1\<^sub>m dimR)" using assms unfolding proj_measurement_def
+ by simp
+ also have "... = Complex_Matrix.trace R" using assms by (simp add: fc_mats_carrier dim_eq)
+ also have "... = 1" using assms unfolding density_operator_def by simp
+ finally show ?thesis unfolding meas_outcome_prob_def .
+qed
+
+text \<open>We introduce the maximally mixed density operator. Intuitively, this density operator
+corresponds to a uniform distribution of the states of an orthonormal basis.
+This operator will be used to define the density operator after a measurement for the measure
+outcome probabilities equal to zero.\<close>
+
+definition max_mix_density :: "nat \<Rightarrow> complex Matrix.mat" where
+"max_mix_density n = ((1::real)/ n) \<cdot>\<^sub>m (1\<^sub>m n)"
+
+lemma max_mix_density_carrier:
+ shows "max_mix_density n \<in> carrier_mat n n" unfolding max_mix_density_def by simp
+
+lemma max_mix_is_density:
+ assumes "0 < n"
+ shows "density_operator (max_mix_density n)" unfolding density_operator_def max_mix_density_def
+proof
+ have "Complex_Matrix.trace (complex_of_real ((1::real)/n) \<cdot>\<^sub>m 1\<^sub>m n) =
+ (complex_of_real ((1::real)/n)) * (Complex_Matrix.trace ((1\<^sub>m n)::complex Matrix.mat))"
+ using one_carrier_mat trace_smult[of "(1\<^sub>m n)::complex Matrix.mat"] by blast
+ also have "... = (complex_of_real ((1::real)/n)) * (real n)" using trace_1[of n] by simp
+ also have "... = 1" using assms by simp
+ finally show "Complex_Matrix.trace (complex_of_real ((1::real)/n) \<cdot>\<^sub>m 1\<^sub>m n) = 1" .
+next
+ show "Complex_Matrix.positive (complex_of_real (1 / real n) \<cdot>\<^sub>m 1\<^sub>m n)"
+ by (rule positive_smult, (auto simp add: positive_one less_eq_complex_def))
+qed
+
+lemma (in cpx_sq_mat) max_mix_density_square:
+ shows "max_mix_density dimR \<in> fc_mats" unfolding max_mix_density_def
+ using fc_mats_carrier dim_eq by simp
+
+text \<open>Given a measurement outcome, \verb+density_collapse+ represents the resulting density
+operator. In practice only the measure outcomes with nonzero probabilities are of interest; we
+(arbitrarily) collapse the density operator for zero-probability outcomes to the maximally mixed
+density operator.\<close>
+
+definition density_collapse ::"complex Matrix.mat \<Rightarrow> complex Matrix.mat \<Rightarrow> complex Matrix.mat" where
+"density_collapse R P = (if ((Complex_Matrix.trace (R * P)) = 0) then (max_mix_density (dim_row R))
+ else ((1::real)/ ((Complex_Matrix.trace (R * P)))) \<cdot>\<^sub>m (P * R * P))"
+
+lemma density_collapse_carrier:
+ assumes "0 < dim_row R"
+ and "P \<in> carrier_mat n n"
+ and "R \<in> carrier_mat n n"
+shows "(density_collapse R P) \<in> carrier_mat n n"
+proof (cases "(Complex_Matrix.trace (R * P)) = 0")
+ case True
+ hence "density_collapse R P = max_mix_density (dim_row R)" unfolding density_collapse_def by simp
+ then show ?thesis using max_mix_is_density assms max_mix_density_carrier by auto
+next
+ case False
+ hence "density_collapse R P = complex_of_real 1 / Complex_Matrix.trace (R * P) \<cdot>\<^sub>m (P * R * P)"
+ unfolding density_collapse_def by simp
+ thus ?thesis using assms by auto
+qed
+
+lemma density_collapse_operator:
+ assumes "projector P"
+ and "density_operator R"
+ and "0 < dim_row R"
+ and "P \<in> carrier_mat n n"
+ and "R \<in> carrier_mat n n"
+shows "density_operator (density_collapse R P)"
+proof (cases "(Complex_Matrix.trace (R * P)) = 0")
+ case True
+ hence "density_collapse R P = max_mix_density (dim_row R)" unfolding density_collapse_def by simp
+ then show ?thesis using max_mix_is_density assms by simp
+next
+ case False
+ show ?thesis unfolding density_operator_def
+ proof (intro conjI)
+ have "Complex_Matrix.positive ((1 / (Complex_Matrix.trace (R * P))) \<cdot>\<^sub>m (P * R * P))"
+ proof (rule positive_smult)
+ show "P * R * P \<in> carrier_mat n n" using assms by simp
+ have "Complex_Matrix.positive R" using assms unfolding density_operator_def by simp
+ hence "0 \<le> (Complex_Matrix.trace (R * P))" using positive_proj_trace[of P R n] assms
+ False by auto
+ hence "0 \<le> Re (Complex_Matrix.trace (R * P))" by (simp add: less_eq_complex_def)
+ hence "0 \<le> 1/(Re (Complex_Matrix.trace (R * P)))" by simp
+ have "Re (Complex_Matrix.trace (R * P)) = Complex_Matrix.trace (R * P)"
+ using assms \<open>Complex_Matrix.positive R\<close> trace_proj_pos_real by simp
+ hence inv: "1/ (Complex_Matrix.trace (R * P)) = 1/(Re (Complex_Matrix.trace (R * P)))" by simp
+ thus "0 \<le> 1/ (Complex_Matrix.trace (R * P))"
+ using \<open>0 \<le> 1/(Re (Complex_Matrix.trace (R * P)))\<close> by (simp add: inv less_eq_complex_def)
+ show "Complex_Matrix.positive (P * R * P)" using assms
+ positive_close_under_left_right_mult_adjoint[of P n R]
+ by (simp add: \<open>Complex_Matrix.positive R\<close> hermitian_def projector_def)
+ qed
+ thus "Complex_Matrix.positive (density_collapse R P)" using False
+ unfolding density_collapse_def by simp
+ next
+ have "Complex_Matrix.trace (density_collapse R P) =
+ Complex_Matrix.trace ((1/ (Complex_Matrix.trace (R * P))) \<cdot>\<^sub>m (P * R * P))"
+ using False unfolding density_collapse_def by simp
+ also have "... = 1/ (Complex_Matrix.trace (R * P)) * Complex_Matrix.trace (P * R * P)"
+ using trace_smult[of "P * R * P" n] assms by simp
+ also have "... = 1/ (Complex_Matrix.trace (R * P)) * Complex_Matrix.trace (R * P)"
+ using projector_collapse_trace assms by simp
+ also have "... = 1" using False by simp
+ finally show "Complex_Matrix.trace (density_collapse R P) = 1" .
+ qed
+qed
+
+
+subsection \<open>Measurements with observables\<close>
+
+text \<open>It is standard in quantum mechanics to represent projective measurements with so-called
+\emph{observables}. These are Hermitian matrices which encode projective measurements as follows:
+the eigenvalues of an observable represent the possible projective measurement outcomes, and the
+associated projectors are the projectors onto the corresponding eigenspaces. The results in this
+part are based on the spectral theorem, which states that any Hermitian matrix admits an
+orthonormal basis consisting of eigenvectors of the matrix.\<close>
+
+
+subsubsection \<open>On the diagonal elements of a matrix\<close>
+
+text \<open>We begin by introducing definitions that will be used on the diagonalized version of a
+Hermitian matrix.\<close>
+
+definition diag_elems where
+"diag_elems B = {B$$(i,i) |i. i < dim_row B}"
+
+text \<open>Relationship between \verb+diag_elems+ and the list \verb+diag_mat+\<close>
+
+lemma diag_elems_set_diag_mat:
+ shows "diag_elems B = set (diag_mat B)" unfolding diag_mat_def diag_elems_def
+proof
+ show "{B $$ (i, i) |i. i < dim_row B} \<subseteq> set (map (\<lambda>i. B $$ (i, i)) [0..<dim_row B])"
+ proof
+ fix x
+ assume "x \<in> {B $$ (i, i) |i. i < dim_row B}"
+ hence "\<exists>i < dim_row B. x = B $$(i,i)" by auto
+ from this obtain i where "i < dim_row B" and "x = B $$(i,i)" by auto
+ thus "x \<in> set (map (\<lambda>i. B $$ (i, i)) [0..<dim_row B])" by auto
+ qed
+next
+ show "set (map (\<lambda>i. B $$ (i, i)) [0..<dim_row B]) \<subseteq> {B $$ (i, i) |i. i < dim_row B}"
+ proof
+ fix x
+ assume "x \<in> set (map (\<lambda>i. B $$ (i, i)) [0..<dim_row B])"
+ thus "x \<in> {B $$ (i, i) |i. i < dim_row B}" by auto
+ qed
+qed
+
+lemma diag_elems_finite[simp]:
+ shows "finite (diag_elems B)" unfolding diag_elems_def by simp
+
+lemma diag_elems_mem[simp]:
+ assumes "i < dim_row B"
+ shows "B $$(i,i) \<in> diag_elems B" using assms unfolding diag_elems_def by auto
+
+text \<open>When $x$ is a diagonal element of $B$, \verb+diag_elem_indices+ returns the set of diagonal
+indices of $B$ with value $x$.\<close>
+
+definition diag_elem_indices where
+"diag_elem_indices x B = {i|i. i < dim_row B \<and> B $$ (i,i) = x}"
+
+lemma diag_elem_indices_elem:
+ assumes "a \<in> diag_elem_indices x B"
+ shows "a < dim_row B \<and> B$$(a,a) = x" using assms unfolding diag_elem_indices_def by simp
+
+lemma diag_elem_indices_itself:
+ assumes "i < dim_row B"
+ shows "i \<in> diag_elem_indices (B $$(i,i)) B" using assms unfolding diag_elem_indices_def by simp
+
+lemma diag_elem_indices_finite:
+ shows "finite (diag_elem_indices x B)" unfolding diag_elem_indices_def by simp
+
+text \<open>We can therefore partition the diagonal indices of a matrix $B$ depending on the value
+of the diagonal elements. If $B$ admits $p$ elements on its diagonal, then we define bijections
+between its set of diagonal elements and the initial segment $[0..p-1]$.\<close>
+
+definition dist_el_card where
+"dist_el_card B = card (diag_elems B)"
+
+definition diag_idx_to_el where
+"diag_idx_to_el B = (SOME h. bij_betw h {..< dist_el_card B} (diag_elems B))"
+
+definition diag_el_to_idx where
+"diag_el_to_idx B = inv_into {..< dist_el_card B} (diag_idx_to_el B)"
+
+lemma diag_idx_to_el_bij:
+ shows "bij_betw (diag_idx_to_el B) {..< dist_el_card B} (diag_elems B)"
+proof -
+ let ?V = "SOME h. bij_betw h {..< dist_el_card B} (diag_elems B)"
+ have vprop: "bij_betw ?V {..< dist_el_card B} (diag_elems B)" using
+ someI_ex[of "\<lambda>h. bij_betw h {..< dist_el_card B} (diag_elems B)"]
+ diag_elems_finite unfolding dist_el_card_def using bij_betw_from_nat_into_finite by blast
+ show ?thesis by (simp add:diag_idx_to_el_def vprop)
+qed
+
+lemma diag_el_to_idx_bij:
+ shows "bij_betw (diag_el_to_idx B) (diag_elems B) {..< dist_el_card B}"
+ unfolding diag_el_to_idx_def
+proof (rule bij_betw_inv_into_subset[of _ _ "diag_elems B"], (simp add: diag_idx_to_el_bij)+)
+ show "diag_idx_to_el B ` {..<dist_el_card B} = diag_elems B"
+ by (simp add: diag_idx_to_el_bij bij_betw_imp_surj_on)
+qed
+
+lemma diag_idx_to_el_less_inj:
+ assumes "i < dist_el_card B"
+and "j < dist_el_card B"
+ and "diag_idx_to_el B i = diag_idx_to_el B j"
+shows "i = j"
+proof -
+ have "i \<in> {..< dist_el_card B}" using assms by simp
+ moreover have "j\<in> {..< dist_el_card B}" using assms by simp
+ moreover have "inj_on (diag_idx_to_el B) {..<dist_el_card B}"
+ using diag_idx_to_el_bij[of B]
+ unfolding bij_betw_def by simp
+ ultimately show ?thesis by (meson assms(3) inj_on_contraD)
+qed
+
+lemma diag_idx_to_el_less_surj:
+ assumes "x\<in> diag_elems B"
+shows "\<exists>k \<in> {..< dist_el_card B}. x = diag_idx_to_el B k"
+proof -
+ have "diag_idx_to_el B ` {..<dist_el_card B} = diag_elems B"
+ using diag_idx_to_el_bij[of B] unfolding bij_betw_def by simp
+ thus ?thesis using assms by auto
+qed
+
+lemma diag_idx_to_el_img:
+ assumes "k < dist_el_card B"
+shows "diag_idx_to_el B k \<in> diag_elems B"
+proof -
+ have "diag_idx_to_el B ` {..<dist_el_card B} = diag_elems B"
+ using diag_idx_to_el_bij[of B] unfolding bij_betw_def by simp
+ thus ?thesis using assms by auto
+qed
+
+lemma diag_elems_real:
+ fixes B::"complex Matrix.mat"
+ assumes "\<forall>i < dim_row B. B$$(i,i) \<in> Reals"
+ shows "diag_elems B \<subseteq> Reals"
+proof
+ fix x
+ assume "x\<in> diag_elems B"
+ hence "\<exists>i < dim_row B. x = B $$(i,i)" using assms unfolding diag_elems_def by auto
+ from this obtain i where "i < dim_row B" "x = B $$ (i,i)" by auto
+ thus "x \<in> Reals" using assms by simp
+qed
+
+lemma diag_elems_Re:
+ fixes B::"complex Matrix.mat"
+ assumes "\<forall>i < (dim_row B). B$$(i,i) \<in> Reals"
+ shows "{Re x|x. x\<in> diag_elems B} = diag_elems B"
+proof
+ show "{complex_of_real (Re x) |x. x \<in> diag_elems B} \<subseteq> diag_elems B"
+ proof
+ fix x
+ assume "x \<in> {complex_of_real (Re x) |x. x \<in> diag_elems B}"
+ hence "\<exists>y \<in> diag_elems B. x = Re y" by auto
+ from this obtain y where "y\<in> diag_elems B" and "x = Re y" by auto
+ hence "y = x" using assms diag_elems_real[of B] by auto
+ thus "x\<in> diag_elems B" using \<open>y\<in> diag_elems B\<close> by simp
+ qed
+ show "diag_elems B \<subseteq> {complex_of_real (Re x) |x. x \<in> diag_elems B}"
+ proof
+ fix x
+ assume "x \<in> diag_elems B"
+ hence "Re x = x" using assms diag_elems_real[of B] by auto
+ thus "x\<in> {complex_of_real (Re x) |x. x \<in> diag_elems B}" using \<open>x\<in> diag_elems B\<close> by force
+ qed
+qed
+
+lemma diag_idx_to_el_real:
+ fixes B::"complex Matrix.mat"
+ assumes "\<forall>i < dim_row B. B$$(i,i) \<in> Reals"
+and "i < dist_el_card B"
+shows "Re (diag_idx_to_el B i) = diag_idx_to_el B i"
+proof -
+ have "diag_idx_to_el B i \<in> diag_elems B" using diag_idx_to_el_img[of i B] assms by simp
+ hence "diag_idx_to_el B i \<in> Reals" using diag_elems_real[of B] assms by auto
+ thus ?thesis by simp
+qed
+
+lemma diag_elem_indices_empty:
+ assumes "B \<in> carrier_mat dimR dimC"
+ and "i < (dist_el_card B)"
+and "j < (dist_el_card B)"
+and "i\<noteq> j"
+shows "diag_elem_indices (diag_idx_to_el B i) B \<inter>
+ (diag_elem_indices (diag_idx_to_el B j) B) = {}"
+proof (rule ccontr)
+ assume "diag_elem_indices (diag_idx_to_el B i) B \<inter>
+ diag_elem_indices (diag_idx_to_el B j) B \<noteq> {}"
+ hence "\<exists>x. x\<in> diag_elem_indices (diag_idx_to_el B i) B \<inter>
+ diag_elem_indices (diag_idx_to_el B j) B" by auto
+ from this obtain x where
+ xprop: "x\<in> diag_elem_indices (diag_idx_to_el B i) B \<inter>
+ diag_elem_indices (diag_idx_to_el B j) B" by auto
+ hence "B $$ (x,x) = (diag_idx_to_el B i)"
+ using diag_elem_indices_elem[of x "diag_idx_to_el B i"] by simp
+ moreover have "B $$ (x,x) = (diag_idx_to_el B j)"
+ using diag_elem_indices_elem[of x "diag_idx_to_el B j"] xprop by simp
+ ultimately have "diag_idx_to_el B i = diag_idx_to_el B j" by simp
+ hence "i = j" using diag_idx_to_el_less_inj assms by auto
+ thus False using assms by simp
+qed
+
+lemma (in cpx_sq_mat) diag_elem_indices_disjoint:
+ assumes "B\<in> carrier_mat dimR dimC"
+ shows "disjoint_family_on (\<lambda>n. diag_elem_indices (diag_idx_to_el B n) B)
+ {..<dist_el_card B}" unfolding disjoint_family_on_def
+proof (intro ballI impI)
+ fix p m
+ assume "m\<in> {..< dist_el_card B}" and "p\<in> {..< dist_el_card B}" and "m\<noteq> p"
+ thus "diag_elem_indices (diag_idx_to_el B m) B \<inter>
+ diag_elem_indices (diag_idx_to_el B p) B = {}"
+ using diag_elem_indices_empty assms fc_mats_carrier by simp
+qed
+
+lemma diag_elem_indices_union:
+ assumes "B\<in> carrier_mat dimR dimC"
+ shows "(\<Union> i \<in> {..< dist_el_card B}. diag_elem_indices (diag_idx_to_el B i) B) =
+ {..< dimR}"
+proof
+ show "(\<Union>i<dist_el_card B. diag_elem_indices (diag_idx_to_el B i) B) \<subseteq> {..<dimR}"
+ proof
+ fix x
+ assume "x \<in> (\<Union>i<dist_el_card B. diag_elem_indices (diag_idx_to_el B i) B)"
+ hence "\<exists>i < dist_el_card B. x\<in> diag_elem_indices (diag_idx_to_el B i) B" by auto
+ from this obtain i where "i < dist_el_card B"
+ "x\<in> diag_elem_indices (diag_idx_to_el B i) B" by auto
+ hence "x < dimR" using diag_elem_indices_elem[of x _ B] assms by simp
+ thus "x \<in> {..< dimR}" by auto
+ qed
+next
+ show "{..<dimR} \<subseteq> (\<Union>i<dist_el_card B. diag_elem_indices (diag_idx_to_el B i) B)"
+ proof
+ fix j
+ assume "j\<in> {..< dimR}"
+ hence "j < dim_row B" using assms by simp
+ hence "B$$(j,j) \<in> diag_elems B" by simp
+ hence "\<exists>k \<in> {..< dist_el_card B}. B$$(j,j) = diag_idx_to_el B k"
+ using diag_idx_to_el_less_surj[of "B $$(j,j)"] by simp
+ from this obtain k where kprop: "k \<in> {..< dist_el_card B}"
+ "B$$(j,j) = diag_idx_to_el B k" by auto
+ hence "j \<in> diag_elem_indices (diag_idx_to_el B k) B" using \<open>j < dim_row B\<close>
+ diag_elem_indices_itself by fastforce
+ thus "j \<in> (\<Union>i<dist_el_card B. diag_elem_indices (diag_idx_to_el B i) B)"
+ using kprop by auto
+ qed
+qed
+
+
+subsubsection \<open>Construction of measurement outcomes\<close>
+
+text \<open>The construction of a projective measurement for a hermitian matrix $A$ is based on the Schur
+decomposition $A = U*B*U^\dagger$, where $B$ is diagonal and $U$ is unitary. The columns of $U$ are
+normalized and pairwise orthogonal; they are used to construct the projectors associated with
+a measurement value\<close>
+
+definition (in cpx_sq_mat) project_vecs where
+"project_vecs (f::nat \<Rightarrow> complex Matrix.vec) N = sum_mat (\<lambda>i. rank_1_proj (f i)) N"
+
+lemma (in cpx_sq_mat) project_vecs_dim:
+ assumes "\<forall>i \<in> N. dim_vec (f i) = dimR"
+ shows "project_vecs f N \<in> fc_mats"
+proof -
+ have "project_vecs f N \<in> carrier_mat dimR dimC" unfolding project_vecs_def
+ proof (rule sum_mat_carrier)
+ show "\<And>i. i \<in> N \<Longrightarrow> rank_1_proj (f i) \<in> fc_mats" using assms fc_mats_carrier rank_1_proj_dim
+ dim_eq rank_1_proj_carrier by fastforce
+ qed
+ thus ?thesis using fc_mats_carrier by simp
+qed
+
+definition (in cpx_sq_mat) mk_meas_outcome where
+"mk_meas_outcome B U = (\<lambda>i. (Re (diag_idx_to_el B i),
+ project_vecs (\<lambda>i. zero_col U i) (diag_elem_indices (diag_idx_to_el B i) B)))"
+
+lemma (in cpx_sq_mat) mk_meas_outcome_carrier:
+ assumes "Complex_Matrix.unitary U"
+ and "U \<in> fc_mats"
+ and "B\<in> fc_mats"
+shows "meas_outcome_prj ((mk_meas_outcome B U) j) \<in> fc_mats"
+proof -
+ have "project_vecs (\<lambda>i. zero_col U i) (diag_elem_indices (diag_idx_to_el B j) B) \<in> fc_mats"
+ using project_vecs_dim by (simp add: assms(2) zero_col_dim)
+ thus ?thesis unfolding mk_meas_outcome_def meas_outcome_prj_def by simp
+qed
+
+lemma (in cpx_sq_mat) mk_meas_outcome_sum_id:
+ assumes "Complex_Matrix.unitary U"
+ and "U \<in> fc_mats"
+ and "B\<in> fc_mats"
+shows "sum_mat (\<lambda>j. meas_outcome_prj ((mk_meas_outcome B U) j))
+ {..<(dist_el_card B)} = 1\<^sub>m dimR"
+proof -
+ have "sum_mat (\<lambda>j. meas_outcome_prj ((mk_meas_outcome B U) j)) {..<(dist_el_card B)} =
+ sum_mat (\<lambda>j. project_vecs (\<lambda>i. zero_col U i) (diag_elem_indices (diag_idx_to_el B j) B))
+ {..<(dist_el_card B)}"
+ unfolding mk_meas_outcome_def meas_outcome_prj_def by simp
+ also have "... = sum_mat (\<lambda>i. rank_1_proj (zero_col U i))
+ (\<Union>j<dist_el_card B. diag_elem_indices (diag_idx_to_el B j) B)"
+ unfolding project_vecs_def sum_mat_def
+ proof (rule disj_family_sum_with[symmetric])
+ show "\<forall>j. rank_1_proj (zero_col U j) \<in> fc_mats" using zero_col_dim fc_mats_carrier dim_eq
+ by (metis assms(2) rank_1_proj_carrier)
+ show "finite {..<dist_el_card B}" by simp
+ show "\<And>i. i \<in> {..<dist_el_card B} \<Longrightarrow> finite (diag_elem_indices (diag_idx_to_el B i) B)"
+ using diag_elem_indices_finite by simp
+ show "disjoint_family_on (\<lambda>n. diag_elem_indices (diag_idx_to_el B n) B)
+ {..<dist_el_card B}"
+ unfolding disjoint_family_on_def
+ proof (intro ballI impI)
+ fix p m
+ assume "m\<in> {..< dist_el_card B}" and "p\<in> {..< dist_el_card B}" and "m\<noteq> p"
+ thus "diag_elem_indices (diag_idx_to_el B m) B \<inter>
+ diag_elem_indices (diag_idx_to_el B p) B = {}"
+ using diag_elem_indices_empty assms fc_mats_carrier by simp
+ qed
+ qed
+ also have "... = sum_mat (\<lambda>i. rank_1_proj (zero_col U i)) {..< dimR}"
+ using diag_elem_indices_union[of B] assms fc_mats_carrier by simp
+ also have "... = sum_mat (\<lambda>i. rank_1_proj (Matrix.col U i)) {..< dimR}"
+ proof (rule sum_mat_cong, simp)
+ show "\<And>i. i \<in> {..<dimR} \<Longrightarrow> rank_1_proj (zero_col U i) \<in> fc_mats" using dim_eq
+ by (metis assms(2) local.fc_mats_carrier rank_1_proj_carrier zero_col_dim)
+ thus "\<And>i. i \<in> {..<dimR} \<Longrightarrow> rank_1_proj (Matrix.col U i) \<in> fc_mats" using dim_eq
+ by (metis lessThan_iff zero_col_col)
+ show "\<And>i. i \<in> {..<dimR} \<Longrightarrow> rank_1_proj (zero_col U i) = rank_1_proj (Matrix.col U i)"
+ by (simp add: zero_col_col)
+ qed
+ also have "... = 1\<^sub>m dimR" using sum_rank_1_proj_unitary assms by simp
+ finally show ?thesis .
+qed
+
+lemma (in cpx_sq_mat) make_meas_outcome_prj_ortho:
+ assumes "Complex_Matrix.unitary U"
+ and "U \<in> fc_mats"
+ and "B\<in> fc_mats"
+ and "i < dist_el_card B"
+ and "j < dist_el_card B"
+ and "i \<noteq> j"
+shows "meas_outcome_prj ((mk_meas_outcome B U) i) *
+ meas_outcome_prj ((mk_meas_outcome B U) j) = 0\<^sub>m dimR dimR"
+proof -
+ define Pi where "Pi = sum_mat (\<lambda>k. rank_1_proj (zero_col U k))
+ (diag_elem_indices (diag_idx_to_el B i) B)"
+ have sneqi: "meas_outcome_prj (mk_meas_outcome B U i) = Pi"
+ unfolding mk_meas_outcome_def project_vecs_def Pi_def meas_outcome_prj_def by simp
+ define Pj where "Pj = sum_mat (\<lambda>k. rank_1_proj (zero_col U k))
+ (diag_elem_indices (diag_idx_to_el B j) B)"
+ have sneqj: "meas_outcome_prj (mk_meas_outcome B U j) = Pj"
+ unfolding mk_meas_outcome_def project_vecs_def Pj_def meas_outcome_prj_def by simp
+ have "Pi * Pj = 0\<^sub>m dimR dimR" unfolding Pi_def
+ proof (rule sum_mat_left_ortho_zero)
+ show "finite (diag_elem_indices (diag_idx_to_el B i) B)"
+ using diag_elem_indices_finite[of _ B] by simp
+ show km: "\<And>k. k \<in> diag_elem_indices (diag_idx_to_el B i) B \<Longrightarrow>
+ rank_1_proj (zero_col U k) \<in> fc_mats" using zero_col_dim assms fc_mats_carrier dim_eq
+ by (metis rank_1_proj_carrier)
+ show "Pj \<in> fc_mats" using sneqj assms mk_meas_outcome_carrier by auto
+ show "\<And>k. k \<in> diag_elem_indices (diag_idx_to_el B i) B \<Longrightarrow>
+ rank_1_proj (zero_col U k) * Pj = 0\<^sub>m dimR dimR"
+ proof -
+ fix k
+ assume "k \<in> diag_elem_indices (diag_idx_to_el B i) B"
+ show "rank_1_proj (zero_col U k) * Pj = 0\<^sub>m dimR dimR" unfolding Pj_def
+ proof (rule sum_mat_right_ortho_zero)
+ show "finite (diag_elem_indices (diag_idx_to_el B j) B)"
+ using diag_elem_indices_finite[of _ B] by simp
+ show "\<And>i. i \<in> diag_elem_indices (diag_idx_to_el B j) B \<Longrightarrow>
+ rank_1_proj (zero_col U i) \<in> fc_mats" using zero_col_dim assms fc_mats_carrier dim_eq
+ by (metis rank_1_proj_carrier)
+ show "rank_1_proj (zero_col U k) \<in> fc_mats"
+ by (simp add: km \<open>k \<in> diag_elem_indices (diag_idx_to_el B i) B\<close>)
+ show "\<And>i. i \<in> diag_elem_indices (diag_idx_to_el B j) B \<Longrightarrow>
+ rank_1_proj (zero_col U k) * rank_1_proj (zero_col U i) = 0\<^sub>m dimR dimR"
+ proof -
+ fix m
+ assume "m \<in> diag_elem_indices (diag_idx_to_el B j) B"
+ hence "m \<noteq> k" using \<open>k \<in> diag_elem_indices (diag_idx_to_el B i) B\<close>
+ diag_elem_indices_disjoint[of B] fc_mats_carrier assms unfolding disjoint_family_on_def by auto
+ have "\<And>i. i \<in> diag_elem_indices (diag_idx_to_el B j) B \<Longrightarrow> i < dimR"
+ using diag_elem_indices_elem fc_mats_carrier assms dim_eq by (metis carrier_matD(1))
+ hence "m < dimR" using \<open>m \<in> diag_elem_indices (diag_idx_to_el B j) B\<close> by simp
+ have "\<And>k. k \<in> diag_elem_indices (diag_idx_to_el B i) B \<Longrightarrow> k < dimR"
+ using diag_elem_indices_elem fc_mats_carrier assms dim_eq by (metis carrier_matD(1))
+ hence "k < dimR" using \<open>k \<in> diag_elem_indices (diag_idx_to_el B i) B\<close> by simp
+ show "rank_1_proj (zero_col U k) * rank_1_proj (zero_col U m) = 0\<^sub>m dimR dimR"
+ using rank_1_proj_unitary_ne[of U k m] assms \<open>m < dimR\<close> \<open>k < dimR\<close>
+ by (metis \<open>m \<noteq> k\<close> zero_col_col)
+ qed
+ qed
+ qed
+ qed
+ thus ?thesis using sneqi sneqj by simp
+qed
+
+lemma (in cpx_sq_mat) make_meas_outcome_prjectors:
+ assumes "Complex_Matrix.unitary U"
+ and "U \<in> fc_mats"
+ and "B\<in> fc_mats"
+ and "j < dist_el_card B"
+shows "projector (meas_outcome_prj ((mk_meas_outcome B U) j))" unfolding projector_def
+proof
+ define Pj where "Pj = sum_mat (\<lambda>i. rank_1_proj (zero_col U i))
+ (diag_elem_indices (diag_idx_to_el B j) B)"
+ have sneq: "meas_outcome_prj (mk_meas_outcome B U j) = Pj"
+ unfolding mk_meas_outcome_def project_vecs_def Pj_def meas_outcome_prj_def by simp
+ moreover have "hermitian Pj" unfolding Pj_def
+ proof (rule sum_mat_hermitian)
+ show "finite (diag_elem_indices (diag_idx_to_el B j) B)"
+ using diag_elem_indices_finite[of _ B] by simp
+ show "\<forall>i\<in>diag_elem_indices (diag_idx_to_el B j) B. hermitian (rank_1_proj (zero_col U i))"
+ using rank_1_proj_hermitian by simp
+ show "\<forall>i\<in>diag_elem_indices (diag_idx_to_el B j) B. rank_1_proj (zero_col U i) \<in> fc_mats"
+ using zero_col_dim fc_mats_carrier dim_eq by (metis assms(2) rank_1_proj_carrier)
+ qed
+ ultimately show "hermitian (meas_outcome_prj (mk_meas_outcome B U j))" by simp
+ have "Pj * Pj = Pj" unfolding Pj_def
+ proof (rule sum_mat_ortho_square)
+ show "finite (diag_elem_indices (diag_idx_to_el B j) B)"
+ using diag_elem_indices_finite[of _ B] by simp
+ show "\<And>i. i \<in> diag_elem_indices (diag_idx_to_el B j) B \<Longrightarrow>
+ rank_1_proj (zero_col U i) * rank_1_proj (zero_col U i) = rank_1_proj (zero_col U i)"
+ proof -
+ fix i
+ assume imem: "i\<in> diag_elem_indices (diag_idx_to_el B j) B"
+ hence "i < dimR" using diag_elem_indices_elem fc_mats_carrier assms dim_eq
+ by (metis carrier_matD(1))
+ hence "zero_col U i = Matrix.col U i" using zero_col_col[of i] by simp
+ hence "rank_1_proj (zero_col U i) = rank_1_proj (Matrix.col U i)" by simp
+ moreover have "rank_1_proj (Matrix.col U i) * rank_1_proj (Matrix.col U i) =
+ rank_1_proj (Matrix.col U i)" by (rule rank_1_proj_unitary_eq, (auto simp add: assms \<open>i < dimR\<close>))
+ ultimately show "rank_1_proj (zero_col U i) * rank_1_proj (zero_col U i) =
+ rank_1_proj (zero_col U i)" by simp
+ qed
+ show "\<And>i. i \<in> diag_elem_indices (diag_idx_to_el B j) B \<Longrightarrow>
+ rank_1_proj (zero_col U i) \<in> fc_mats"
+ using zero_col_dim assms fc_mats_carrier dim_eq by (metis rank_1_proj_carrier)
+ have "\<And>i. i \<in> diag_elem_indices (diag_idx_to_el B j) B \<Longrightarrow> i < dimR"
+ using diag_elem_indices_elem fc_mats_carrier assms dim_eq by (metis carrier_matD(1))
+ thus "\<And>i ja.
+ i \<in> diag_elem_indices (diag_idx_to_el B j) B \<Longrightarrow>
+ ja \<in> diag_elem_indices (diag_idx_to_el B j) B \<Longrightarrow>
+ i \<noteq> ja \<Longrightarrow> rank_1_proj (zero_col U i) * rank_1_proj (zero_col U ja) = 0\<^sub>m dimR dimR"
+ using rank_1_proj_unitary_ne by (simp add: assms(1) assms(2) zero_col_col)
+ qed
+ thus "meas_outcome_prj (mk_meas_outcome B U j) *
+ meas_outcome_prj (mk_meas_outcome B U j) =
+ meas_outcome_prj (mk_meas_outcome B U j)"
+ using sneq by simp
+qed
+
+lemma (in cpx_sq_mat) mk_meas_outcome_fst_inj:
+ assumes "\<forall>i < (dim_row B). B$$(i,i) \<in> Reals"
+ shows "inj_on (\<lambda>i. meas_outcome_val ((mk_meas_outcome B U) i)) {..<dist_el_card B}"
+ unfolding inj_on_def
+proof (intro ballI impI)
+ fix x y
+ assume "x \<in> {..<dist_el_card B}" and "y \<in> {..<dist_el_card B}"
+ and "meas_outcome_val (mk_meas_outcome B U x) =
+ meas_outcome_val (mk_meas_outcome B U y)" note xy = this
+ hence "diag_idx_to_el B x = Re (diag_idx_to_el B x)"
+ using assms diag_idx_to_el_real by simp
+ also have "... = Re (diag_idx_to_el B y)" using xy
+ unfolding mk_meas_outcome_def meas_outcome_val_def by simp
+ also have "... = diag_idx_to_el B y" using assms diag_idx_to_el_real xy by simp
+ finally have "diag_idx_to_el B x = diag_idx_to_el B y" .
+ thus "x = y " using diag_idx_to_el_less_inj xy by simp
+qed
+
+lemma (in cpx_sq_mat) mk_meas_outcome_fst_bij:
+ assumes "\<forall>i < (dim_row B). B$$(i,i) \<in> Reals"
+ shows "bij_betw (\<lambda>i. meas_outcome_val ((mk_meas_outcome B U) i)) {..< dist_el_card B}
+ {Re x|x. x\<in> diag_elems B}"
+ unfolding bij_betw_def
+proof
+ have "inj_on (\<lambda>x. (meas_outcome_val (mk_meas_outcome B U x))) {..<dist_el_card B}"
+ using assms mk_meas_outcome_fst_inj by simp
+ moreover have "{Re x|x. x\<in> diag_elems B} = diag_elems B" using diag_elems_Re[of B] assms by simp
+ ultimately show "inj_on (\<lambda>x. meas_outcome_val (mk_meas_outcome B U x))
+ {..<dist_el_card B}" by simp
+ show "(\<lambda>i. meas_outcome_val (mk_meas_outcome B U i)) ` {..<dist_el_card B} =
+ {Re x |x. x \<in> diag_elems B}" unfolding meas_outcome_val_def mk_meas_outcome_def
+ proof
+ show "(\<lambda>i. fst (Re (diag_idx_to_el B i), project_vecs (zero_col U)
+ (diag_elem_indices (diag_idx_to_el B i) B))) ` {..<dist_el_card B} \<subseteq>
+ {Re x |x. x \<in> diag_elems B}"
+ using diag_idx_to_el_bij[of B] bij_betw_apply by fastforce
+ show "{Re x |x. x \<in> diag_elems B}
+ \<subseteq> (\<lambda>i. fst (Re (diag_idx_to_el B i),
+ project_vecs (zero_col U) (diag_elem_indices (diag_idx_to_el B i) B))) `
+ {..<dist_el_card B}" using diag_idx_to_el_less_surj by fastforce
+ qed
+qed
+
+
+subsubsection \<open>Projective measurement associated with an observable\<close>
+
+definition eigvals where
+"eigvals M = (SOME as. char_poly M = (\<Prod>a\<leftarrow>as. [:- a, 1:]) \<and> length as = dim_row M)"
+
+lemma eigvals_poly_length:
+ assumes "(M::complex Matrix.mat) \<in> carrier_mat n n"
+ shows "char_poly M = (\<Prod>a\<leftarrow>(eigvals M). [:- a, 1:]) \<and> length (eigvals M) = dim_row M"
+proof -
+ let ?V = "SOME as. char_poly M = (\<Prod>a\<leftarrow>as. [:- a, 1:]) \<and> length as = dim_row M"
+ have vprop: "char_poly M = (\<Prod>a\<leftarrow>?V. [:- a, 1:]) \<and> length ?V = dim_row M" using
+ someI_ex[of "\<lambda>as. char_poly M = (\<Prod>a\<leftarrow>as. [:- a, 1:]) \<and> length as = dim_row M"]
+ complex_mat_char_poly_factorizable assms by blast
+ show ?thesis by (metis (no_types) eigvals_def vprop)
+qed
+
+text \<open>We define the spectrum of a matrix $A$: this is its set of eigenvalues; its elements are
+roots of the characteristic polynomial of $A$.\<close>
+
+definition spectrum where
+"spectrum M = set (eigvals M)"
+
+lemma spectrum_finite:
+ shows "finite (spectrum M)" unfolding spectrum_def by simp
+
+lemma spectrum_char_poly_root:
+ fixes A::"complex Matrix.mat"
+ assumes "A\<in> carrier_mat n n"
+and "k \<in> spectrum A"
+shows "poly (char_poly A) k = 0" using eigvals_poly_length[of A n] assms
+ unfolding spectrum_def eigenvalue_root_char_poly
+ by (simp add: linear_poly_root)
+
+lemma spectrum_eigenvalues:
+ fixes A::"complex Matrix.mat"
+ assumes "A\<in> carrier_mat n n"
+and "k\<in> spectrum A"
+shows "eigenvalue A k" using eigenvalue_root_char_poly[of A n k]
+ spectrum_char_poly_root[of A n k] assms by simp
+
+text \<open>The main result that is used to construct a projective measurement for a Hermitian matrix
+is that it is always possible to decompose it as $A = U*B*U^\dagger$, where $B$ is diagonal and
+only contains real elements, and $U$ is unitary.\<close>
+
+definition hermitian_decomp where
+"hermitian_decomp A B U \<equiv> similar_mat_wit A B U (Complex_Matrix.adjoint U) \<and> diagonal_mat B \<and>
+ diag_mat B = (eigvals A) \<and> unitary U \<and> (\<forall>i < dim_row B. B$$(i, i) \<in> Reals)"
+
+lemma hermitian_decomp_sim:
+ assumes "hermitian_decomp A B U"
+ shows "similar_mat_wit A B U (Complex_Matrix.adjoint U)" using assms
+ unfolding hermitian_decomp_def by simp
+
+lemma hermitian_decomp_diag_mat:
+ assumes "hermitian_decomp A B U"
+ shows "diagonal_mat B" using assms
+ unfolding hermitian_decomp_def by simp
+
+lemma hermitian_decomp_eigenvalues:
+ assumes "hermitian_decomp A B U"
+ shows "diag_mat B = (eigvals A)" using assms
+ unfolding hermitian_decomp_def by simp
+
+lemma hermitian_decomp_unitary:
+ assumes "hermitian_decomp A B U"
+ shows "unitary U" using assms
+ unfolding hermitian_decomp_def by simp
+
+lemma hermitian_decomp_real_eigvals:
+ assumes "hermitian_decomp A B U"
+ shows "\<forall>i < dim_row B. B$$(i, i) \<in> Reals" using assms
+ unfolding hermitian_decomp_def by simp
+
+lemma hermitian_decomp_dim_carrier:
+ assumes "hermitian_decomp A B U"
+ shows "B \<in> carrier_mat (dim_row A) (dim_col A)" using assms
+ unfolding hermitian_decomp_def similar_mat_wit_def
+ by (metis (full_types) index_mult_mat(3) index_one_mat(3) insert_subset)
+
+lemma similar_mat_wit_dim_row:
+ assumes "similar_mat_wit A B Q R"
+ shows "dim_row B = dim_row A" using assms Let_def unfolding similar_mat_wit_def
+ by (meson carrier_matD(1) insert_subset)
+
+lemma (in cpx_sq_mat) hermitian_schur_decomp:
+ assumes "hermitian A"
+ and "A\<in> fc_mats"
+obtains B U where "hermitian_decomp A B U"
+proof -
+ {
+ have es: "char_poly A = (\<Prod> (e :: complex) \<leftarrow> (eigvals A). [:- e, 1:])"
+ using assms fc_mats_carrier eigvals_poly_length dim_eq by auto
+ obtain B U Q where us: "unitary_schur_decomposition A (eigvals A) = (B,U,Q)"
+ by (cases "unitary_schur_decomposition A (eigvals A)")
+ hence pr: "similar_mat_wit A B U (Complex_Matrix.adjoint U) \<and> diagonal_mat B \<and>
+ diag_mat B = (eigvals A) \<and> unitary U \<and> (\<forall>i < dimR. B$$(i, i) \<in> Reals)"
+ using hermitian_eigenvalue_real assms fc_mats_carrier es dim_eq by auto
+ moreover have "dim_row B = dimR" using assms fc_mats_carrier dim_eq similar_mat_wit_dim_row[of A]
+ pr by auto
+ ultimately have "hermitian_decomp A B U" unfolding hermitian_decomp_def by simp
+ hence "\<exists> B U. hermitian_decomp A B U" by auto
+ }
+ thus ?thesis using that by auto
+qed
+
+lemma (in cpx_sq_mat) hermitian_spectrum_real:
+ assumes "A \<in> fc_mats"
+ and "hermitian A"
+ and "a \<in> spectrum A"
+shows "a \<in> Reals"
+proof -
+ obtain B U where bu: "hermitian_decomp A B U" using assms hermitian_schur_decomp by auto
+ hence dimB: "B \<in> carrier_mat dimR dimR" using assms fc_mats_carrier
+ dim_eq hermitian_decomp_dim_carrier[of A] by simp
+ hence Bii: "\<And>i. i < dimR \<Longrightarrow> B$$(i, i) \<in> Reals" using hermitian_decomp_real_eigvals[of A B U]
+ bu assms fc_mats_carrier by simp
+ have "diag_mat B = (eigvals A)" using bu hermitian_decomp_eigenvalues[of A B] by simp
+ hence "a \<in> set (diag_mat B)" using assms unfolding spectrum_def by simp
+ hence "\<exists>i < length (diag_mat B). a = diag_mat B ! i" by (metis in_set_conv_nth)
+ from this obtain i where "i < length (diag_mat B)" and "a = diag_mat B ! i" by auto
+ hence "a = B $$ (i,i)" unfolding diag_mat_def by simp
+ moreover have "i < dimR" using dimB \<open>i < length (diag_mat B)\<close> unfolding diag_mat_def by simp
+ ultimately show ?thesis using Bii by simp
+qed
+
+lemma (in cpx_sq_mat) spectrum_ne:
+ assumes "A\<in> fc_mats"
+and "hermitian A"
+shows "spectrum A \<noteq> {}"
+proof -
+ obtain B U where bu: "hermitian_decomp A B U" using assms hermitian_schur_decomp by auto
+ hence dimB: "B \<in> carrier_mat dimR dimR" using assms fc_mats_carrier
+ dim_eq hermitian_decomp_dim_carrier[of A] by simp
+ have "diag_mat B = (eigvals A)" using bu hermitian_decomp_eigenvalues[of A B] by simp
+ have "B$$(0,0) \<in> set (diag_mat B)" using dimB npos unfolding diag_mat_def by simp
+ hence "set (diag_mat B) \<noteq> {}" by auto
+ thus ?thesis unfolding spectrum_def using \<open>diag_mat B = eigvals A\<close> by auto
+qed
+
+lemma unitary_hermitian_eigenvalues:
+ fixes U::"complex Matrix.mat"
+ assumes "unitary U"
+ and "hermitian U"
+ and "U \<in> carrier_mat n n"
+ and "0 < n"
+ and "k \<in> spectrum U"
+shows "k \<in> {-1, 1}"
+proof -
+ have "cpx_sq_mat n n (carrier_mat n n)" by (unfold_locales, (simp add: assms)+)
+ have "eigenvalue U k" using spectrum_eigenvalues assms by simp
+ have "k \<in> Reals" using assms \<open>cpx_sq_mat n n (carrier_mat n n)\<close>
+ cpx_sq_mat.hermitian_spectrum_real by simp
+ hence "conjugate k = k" by (simp add: Reals_cnj_iff)
+ hence "k\<^sup>2 = 1" using unitary_eigenvalues_norm_square[of U n k] assms
+ by (simp add: \<open>eigenvalue U k\<close> power2_eq_square)
+ thus ?thesis using power2_eq_1_iff by auto
+qed
+
+lemma unitary_hermitian_Re_spectrum:
+ fixes U::"complex Matrix.mat"
+ assumes "unitary U"
+ and "hermitian U"
+ and "U \<in> carrier_mat n n"
+ and "0 < n"
+ shows "{Re x|x. x\<in> spectrum U} \<subseteq> {-1,1}"
+proof
+ fix y
+ assume "y\<in> {Re x|x. x\<in> spectrum U}"
+ hence "\<exists>x \<in> spectrum U. y = Re x" by auto
+ from this obtain x where "x\<in> spectrum U" and "y = Re x" by auto
+ hence "x\<in> {-1,1}" using unitary_hermitian_eigenvalues assms by simp
+ thus "y \<in> {-1, 1}" using \<open>y = Re x\<close> by auto
+qed
+
+text \<open>The projective measurement associated with matrix $M$ is obtained from its Schur
+decomposition, by considering the number of distinct elements on the resulting diagonal matrix
+and constructing the projectors associated with each of them.\<close>
+
+type_synonym proj_meas_rep = "nat \<times> (nat \<Rightarrow> measure_outcome)"
+
+definition proj_meas_size::"proj_meas_rep \<Rightarrow> nat" where
+"proj_meas_size P = fst P"
+
+definition proj_meas_outcomes::"proj_meas_rep \<Rightarrow> (nat \<Rightarrow> measure_outcome)" where
+"proj_meas_outcomes P = snd P"
+
+definition (in cpx_sq_mat) make_pm::"complex Matrix.mat \<Rightarrow> proj_meas_rep" where
+"make_pm A = (let (B,U,_) = unitary_schur_decomposition A (eigvals A) in
+ (dist_el_card B, mk_meas_outcome B U))"
+
+lemma (in cpx_sq_mat) make_pm_decomp:
+ shows "make_pm A = (proj_meas_size (make_pm A), proj_meas_outcomes (make_pm A))"
+ unfolding proj_meas_size_def proj_meas_outcomes_def by simp
+
+lemma (in cpx_sq_mat) make_pm_proj_measurement:
+ assumes "A \<in> fc_mats"
+ and "hermitian A"
+ and "make_pm A = (n, M)"
+shows "proj_measurement n M"
+proof -
+ have es: "char_poly A = (\<Prod> (e :: complex) \<leftarrow> (eigvals A). [:- e, 1:])"
+ using assms fc_mats_carrier eigvals_poly_length dim_eq by auto
+ obtain B U Q where us: "unitary_schur_decomposition A (eigvals A) = (B,U,Q)"
+ by (cases "unitary_schur_decomposition A (eigvals A)", auto)
+ then have "similar_mat_wit A B U (Complex_Matrix.adjoint U) \<and> diagonal_mat B \<and>
+ diag_mat B = (eigvals A) \<and> unitary U \<and> (\<forall>i < dimR. B$$(i, i) \<in> Reals)"
+ using hermitian_eigenvalue_real assms fc_mats_carrier es dim_eq by auto
+ hence A: "A = U * B * (Complex_Matrix.adjoint U)" and dB: "diagonal_mat B"
+ and Bii: "\<And>i. i < dimR \<Longrightarrow> B$$(i, i) \<in> Reals"
+ and dimB: "B \<in> carrier_mat dimR dimR" and dimP: "U \<in> carrier_mat dimR dimR" and
+ dimaP: "Complex_Matrix.adjoint U \<in> carrier_mat dimR dimR" and unit: "unitary U"
+ unfolding similar_mat_wit_def Let_def using assms fc_mats_carrier by auto
+ hence mpeq: "make_pm A = (dist_el_card B, mk_meas_outcome B U)" using us Let_def
+ unfolding make_pm_def by simp
+ hence "n = dist_el_card B" using assms by simp
+ have "M = mk_meas_outcome B U" using assms mpeq by simp
+ show ?thesis unfolding proj_measurement_def
+ proof (intro conjI)
+ show "inj_on (\<lambda>i. meas_outcome_val (M i)) {..<n}"
+ using mk_meas_outcome_fst_inj[of B U]
+ \<open>n = dist_el_card B\<close> \<open>M = mk_meas_outcome B U\<close> Bii fc_mats_carrier dimB by auto
+ show "\<forall>j<n. meas_outcome_prj (M j) \<in> fc_mats \<and> projector (meas_outcome_prj (M j))"
+ proof (intro allI impI conjI)
+ fix j
+ assume "j < n"
+ show "meas_outcome_prj (M j) \<in> fc_mats" using \<open>M = mk_meas_outcome B U\<close> assms
+ fc_mats_carrier \<open>j < n\<close> \<open>n = dist_el_card B\<close> dim_eq mk_meas_outcome_carrier
+ dimB dimP unit by blast
+ show "projector (meas_outcome_prj (M j))" using make_meas_outcome_prjectors
+ \<open>M = mk_meas_outcome B U\<close>
+ fc_mats_carrier \<open>n = dist_el_card B\<close> unit \<open>j < n\<close> dimB dimP dim_eq by blast
+ qed
+ show "\<forall>i<n. \<forall>j<n. i \<noteq> j \<longrightarrow> meas_outcome_prj (M i) * meas_outcome_prj (M j) =
+ 0\<^sub>m dimR dimR"
+ proof (intro allI impI)
+ fix i
+ fix j
+ assume "i < n" and "j < n" and "i\<noteq> j"
+ thus "meas_outcome_prj (M i) * meas_outcome_prj (M j) = 0\<^sub>m dimR dimR"
+ using make_meas_outcome_prj_ortho[of U B i j] assms dimB dimP fc_mats_carrier dim_eq
+ by (simp add: \<open>M = mk_meas_outcome B U\<close> \<open>n = dist_el_card B\<close> unit)
+ qed
+ show "sum_mat (\<lambda>j. meas_outcome_prj (M j)) {..<n} = 1\<^sub>m dimR" using
+ mk_meas_outcome_sum_id
+ \<open>M = mk_meas_outcome B U\<close> fc_mats_carrier dim_eq \<open>n = dist_el_card B\<close> unit dimB dimP
+ by blast
+ qed
+qed
+
+lemma (in cpx_sq_mat) make_pm_proj_measurement':
+ assumes "A\<in> fc_mats"
+ and "hermitian A"
+shows "proj_measurement (proj_meas_size (make_pm A)) (proj_meas_outcomes (make_pm A))"
+ unfolding proj_meas_size_def proj_meas_outcomes_def
+ by (rule make_pm_proj_measurement[of A], (simp add: assms)+)
+
+lemma (in cpx_sq_mat) make_pm_projectors:
+ assumes "A\<in> fc_mats"
+ and "hermitian A"
+and "i < proj_meas_size (make_pm A)"
+shows "projector (meas_outcome_prj (proj_meas_outcomes (make_pm A) i))"
+proof -
+ have "proj_measurement (proj_meas_size (make_pm A)) (proj_meas_outcomes (make_pm A))"
+ using assms make_pm_proj_measurement' by simp
+ thus ?thesis using proj_measurement_proj assms by simp
+qed
+
+lemma (in cpx_sq_mat) make_pm_square:
+ assumes "A\<in> fc_mats"
+ and "hermitian A"
+and "i < proj_meas_size (make_pm A)"
+shows "meas_outcome_prj (proj_meas_outcomes (make_pm A) i) \<in> fc_mats"
+proof -
+ have "proj_measurement (proj_meas_size (make_pm A)) (proj_meas_outcomes (make_pm A))"
+ using assms make_pm_proj_measurement' by simp
+ thus ?thesis using proj_measurement_square assms by simp
+qed
+
+
+subsubsection \<open>Properties on the spectrum of a Hermitian matrix\<close>
+
+lemma (in cpx_sq_mat) hermitian_schur_decomp':
+ assumes "hermitian A"
+ and "A\<in> fc_mats"
+obtains B U where "hermitian_decomp A B U \<and>
+ make_pm A = (dist_el_card B, mk_meas_outcome B U)"
+proof -
+ {
+ have es: "char_poly A = (\<Prod> (e :: complex) \<leftarrow> (eigvals A). [:- e, 1:])"
+ using assms fc_mats_carrier eigvals_poly_length dim_eq by auto
+ obtain B U Q where us: "unitary_schur_decomposition A (eigvals A) = (B,U,Q)"
+ by (cases "unitary_schur_decomposition A (eigvals A)")
+ hence pr: "similar_mat_wit A B U (Complex_Matrix.adjoint U) \<and> diagonal_mat B \<and>
+ diag_mat B = (eigvals A) \<and> unitary U \<and> (\<forall>i < dimR. B$$(i, i) \<in> Reals)"
+ using hermitian_eigenvalue_real assms fc_mats_carrier es dim_eq by auto
+ moreover have "dim_row B = dimR" using assms fc_mats_carrier dim_eq similar_mat_wit_dim_row[of A]
+ pr by auto
+ ultimately have "hermitian_decomp A B U" unfolding hermitian_decomp_def by simp
+ moreover have "make_pm A = (dist_el_card B, mk_meas_outcome B U)" using us Let_def
+ unfolding make_pm_def hermitian_decomp_def by simp
+ ultimately have "\<exists> B U. hermitian_decomp A B U \<and>
+ make_pm A = (dist_el_card B, mk_meas_outcome B U)" by auto
+ }
+ thus ?thesis using that by auto
+qed
+
+lemma (in cpx_sq_mat) spectrum_meas_outcome_val_eq:
+ assumes "hermitian A"
+ and "A \<in> fc_mats"
+and "make_pm A = (p, M)"
+shows "spectrum A = (\<lambda>i. meas_outcome_val (M i)) `{..< p}"
+proof -
+ obtain B U where bu: "hermitian_decomp A B U \<and>
+ make_pm A = (dist_el_card B, mk_meas_outcome B U)"
+ using assms hermitian_schur_decomp'[OF assms(1)] by auto
+ hence "p = dist_el_card B" using assms by simp
+ have dimB: "B \<in> carrier_mat dimR dimR" using hermitian_decomp_dim_carrier[of A] dim_eq bu assms
+ fc_mats_carrier by auto
+ have Bvals: "diag_mat B = eigvals A" using bu hermitian_decomp_eigenvalues[of A B U] by simp
+ have Bii: "\<And>i. i < dimR \<Longrightarrow> B$$(i, i) \<in> Reals" using bu hermitian_decomp_real_eigvals[of A B U]
+ dimB by simp
+ have "diag_elems B = set (eigvals A)" using dimB Bvals diag_elems_set_diag_mat[of B] by simp
+ have "M = mk_meas_outcome B U" using assms bu by simp
+ {
+ fix i
+ assume "i < p"
+ have "meas_outcome_val (M i) = Re (diag_idx_to_el B i)"
+ using \<open>M = mk_meas_outcome B U\<close>
+ unfolding meas_outcome_val_def mk_meas_outcome_def by simp
+ also have "... = diag_idx_to_el B i" using diag_idx_to_el_real dimB \<open>i < p\<close>
+ \<open>p = dist_el_card B\<close> Bii by simp
+ finally have "meas_outcome_val (M i) = diag_idx_to_el B i" .
+ } note eq = this
+ have "bij_betw (diag_idx_to_el B) {..<dist_el_card B} (diag_elems B)"
+ using diag_idx_to_el_bij[of B] by simp
+ hence "diag_idx_to_el B ` {..< p} = diag_elems B" using \<open>p = dist_el_card B\<close>
+ unfolding bij_betw_def by simp
+ thus ?thesis using eq \<open>diag_elems B = set (eigvals A)\<close> unfolding spectrum_def by auto
+qed
+
+lemma (in cpx_sq_mat) spectrum_meas_outcome_val_eq':
+ assumes "hermitian A"
+ and "A \<in> fc_mats"
+and "make_pm A = (p, M)"
+shows "{Re x|x. x\<in> spectrum A} = (\<lambda>i. meas_outcome_val (M i)) `{..< p}"
+proof
+ show "{Re x |x. x \<in> spectrum A} \<subseteq> (\<lambda>i. meas_outcome_val (M i)) ` {..<p}"
+ using spectrum_meas_outcome_val_eq assms by force
+ show "(\<lambda>i. meas_outcome_val (M i)) ` {..<p} \<subseteq> {Re x |x. x \<in> spectrum A}"
+ using spectrum_meas_outcome_val_eq assms by force
+qed
+
+lemma (in cpx_sq_mat) make_pm_eigenvalues:
+ assumes "A\<in> fc_mats"
+ and "hermitian A"
+and "i < proj_meas_size (make_pm A)"
+shows "meas_outcome_val (proj_meas_outcomes (make_pm A) i) \<in> spectrum A"
+ using spectrum_meas_outcome_val_eq[of A] assms make_pm_decomp by auto
+
+lemma (in cpx_sq_mat) make_pm_spectrum:
+ assumes "A\<in> fc_mats"
+ and "hermitian A"
+ and "make_pm A = (p,M)"
+shows "(\<lambda>i. meas_outcome_val (proj_meas_outcomes (make_pm A) i)) ` {..< p} = spectrum A"
+proof
+ show "(\<lambda>x. complex_of_real (meas_outcome_val (proj_meas_outcomes (make_pm A) x))) ` {..<p} \<subseteq>
+ spectrum A"
+ using assms make_pm_eigenvalues make_pm_decomp
+ by (metis (mono_tags, lifting) Pair_inject image_subsetI lessThan_iff)
+ show "spectrum A \<subseteq>
+ (\<lambda>x. complex_of_real (meas_outcome_val (proj_meas_outcomes (make_pm A) x))) ` {..<p}"
+ by (metis Pair_inject assms equalityE make_pm_decomp spectrum_meas_outcome_val_eq)
+qed
+
+lemma (in cpx_sq_mat) spectrum_size:
+ assumes "hermitian A"
+ and "A\<in> fc_mats"
+and "make_pm A = (p, M)"
+shows "p = card (spectrum A)"
+proof -
+ obtain B U where bu: "hermitian_decomp A B U \<and>
+ make_pm A = (dist_el_card B, mk_meas_outcome B U)"
+ using assms hermitian_schur_decomp'[OF assms(1)] by auto
+ hence "p = dist_el_card B" using assms by simp
+ have "spectrum A = set (diag_mat B)" using bu hermitian_decomp_eigenvalues[of A B U]
+ unfolding spectrum_def by simp
+ also have "... = diag_elems B" using diag_elems_set_diag_mat[of B] by simp
+ finally have "spectrum A = diag_elems B" .
+ thus ?thesis using \<open>p = dist_el_card B\<close> unfolding dist_el_card_def by simp
+qed
+
+lemma (in cpx_sq_mat) spectrum_size':
+ assumes "hermitian A"
+and "A\<in> fc_mats"
+shows "proj_meas_size (make_pm A) = card (spectrum A)" using spectrum_size
+ unfolding proj_meas_size_def
+ by (metis assms fst_conv surj_pair)
+
+lemma (in cpx_sq_mat) make_pm_projectors':
+ assumes "hermitian A"
+ and "A \<in> fc_mats"
+ and "a<card (spectrum A)"
+shows "projector (meas_outcome_prj (proj_meas_outcomes (make_pm A) a))"
+ by (rule make_pm_projectors, (simp add: assms spectrum_size')+)
+
+lemma (in cpx_sq_mat) meas_outcome_prj_carrier:
+ assumes "hermitian A"
+ and "A \<in> fc_mats"
+ and "a<card (spectrum A)"
+shows "meas_outcome_prj (proj_meas_outcomes (make_pm A) a) \<in> fc_mats"
+proof (rule proj_measurement_square)
+ show "proj_measurement (proj_meas_size (make_pm A)) (proj_meas_outcomes (make_pm A))"
+ using make_pm_proj_measurement' assms by simp
+ show "a < proj_meas_size (make_pm A)" using assms
+ spectrum_size[of _ "proj_meas_size (make_pm A)"] make_pm_decomp[of A] by simp
+qed
+
+lemma (in cpx_sq_mat) meas_outcome_prj_trace_real:
+ assumes "hermitian A"
+ and "density_operator R"
+ and "R \<in> fc_mats"
+ and "A\<in> fc_mats"
+ and "a<card (spectrum A)"
+shows "Re (Complex_Matrix.trace (R * meas_outcome_prj (proj_meas_outcomes (make_pm A) a))) =
+ Complex_Matrix.trace (R * meas_outcome_prj (proj_meas_outcomes (make_pm A) a))"
+proof (rule trace_proj_pos_real)
+ show "R \<in> carrier_mat dimR dimR" using fc_mats_carrier assms dim_eq by simp
+ show "Complex_Matrix.positive R" using assms unfolding density_operator_def by simp
+ show "projector (meas_outcome_prj (proj_meas_outcomes (make_pm A) a))" using assms
+ make_pm_projectors' by simp
+ show "meas_outcome_prj (proj_meas_outcomes (make_pm A) a) \<in> carrier_mat dimR dimR"
+ using meas_outcome_prj_carrier assms
+ dim_eq fc_mats_carrier by simp
+qed
+
+lemma (in cpx_sq_mat) sum_over_spectrum:
+ assumes "A\<in> fc_mats"
+and "hermitian A"
+and "make_pm A = (p, M)"
+shows "(\<Sum> y \<in> spectrum A. f y) = (\<Sum> i < p. f (meas_outcome_val (M i)))"
+proof (rule sum.reindex_cong)
+show "spectrum A =(\<lambda>x. (meas_outcome_val (M x)))` {..< p}"
+ using spectrum_meas_outcome_val_eq assms by simp
+ show "inj_on (\<lambda>x. complex_of_real (meas_outcome_val (M x))) {..<p}"
+ proof -
+ have "inj_on (\<lambda>x. (meas_outcome_val (M x))) {..<p}"
+ using make_pm_proj_measurement[of A p M] assms proj_measurement_inj by simp
+ thus ?thesis by (simp add: inj_on_def)
+ qed
+qed simp
+
+lemma (in cpx_sq_mat) sum_over_spectrum':
+ assumes "A\<in> fc_mats"
+and "hermitian A"
+and "make_pm A = (p, M)"
+shows "(\<Sum> y \<in> {Re x|x. x \<in> spectrum A}. f y) = (\<Sum> i < p. f (meas_outcome_val (M i)))"
+proof (rule sum.reindex_cong)
+ show "{Re x|x. x \<in> spectrum A} =(\<lambda>x. (meas_outcome_val (M x)))` {..< p}"
+ using spectrum_meas_outcome_val_eq' assms by simp
+ show "inj_on (\<lambda>x. (meas_outcome_val (M x))) {..<p}" using make_pm_proj_measurement[of A p M]
+ assms proj_measurement_inj by simp
+qed simp
+
+text \<open>When a matrix $A$ is decomposed into a projective measurement $\{(\lambda_a, \pi_a)\}$, it
+can be recovered by the formula $A = \sum \lambda_a \pi_a$.\<close>
+
+lemma (in cpx_sq_mat) make_pm_sum:
+ assumes "A \<in> fc_mats"
+ and "hermitian A"
+ and "make_pm A = (p, M)"
+shows "sum_mat (\<lambda>i. (meas_outcome_val (M i)) \<cdot>\<^sub>m meas_outcome_prj (M i)) {..< p} = A"
+proof -
+ have es: "char_poly A = (\<Prod> (e :: complex) \<leftarrow> (eigvals A). [:- e, 1:])"
+ using assms fc_mats_carrier eigvals_poly_length dim_eq by auto
+ obtain B U Q where us: "unitary_schur_decomposition A (eigvals A) = (B,U,Q)"
+ by (cases "unitary_schur_decomposition A (eigvals A)", auto)
+ then have "similar_mat_wit A B U (Complex_Matrix.adjoint U) \<and> diagonal_mat B \<and>
+ diag_mat B = (eigvals A)
+ \<and> unitary U \<and> (\<forall>i < dimR. B$$(i, i) \<in> Reals)"
+ using hermitian_eigenvalue_real assms fc_mats_carrier es dim_eq by auto
+ hence A: "A = U * B * (Complex_Matrix.adjoint U)" and dB: "diagonal_mat B"
+ and Bii: "\<And>i. i < dimR \<Longrightarrow> B$$(i, i) \<in> Reals"
+ and dimB: "B \<in> carrier_mat dimR dimR" and dimP: "U \<in> carrier_mat dimR dimR" and
+ dimaP: "Complex_Matrix.adjoint U \<in> carrier_mat dimR dimR" and unit: "unitary U"
+ unfolding similar_mat_wit_def Let_def using assms fc_mats_carrier by auto
+ hence mpeq: "make_pm A = (dist_el_card B, mk_meas_outcome B U)" using us Let_def
+ unfolding make_pm_def by simp
+ hence "p = dist_el_card B" using assms by simp
+ have "M = mk_meas_outcome B U" using assms mpeq by simp
+ have "sum_mat (\<lambda>i. meas_outcome_val (M i) \<cdot>\<^sub>m meas_outcome_prj (M i)) {..< p} =
+ sum_mat (\<lambda>j. Re (diag_idx_to_el B j)\<cdot>\<^sub>m project_vecs (\<lambda>i. zero_col U i)
+ (diag_elem_indices (diag_idx_to_el B j) B)) {..<(dist_el_card B)}"
+ using \<open>p = dist_el_card B\<close>
+ \<open>M = mk_meas_outcome B U\<close> unfolding meas_outcome_val_def meas_outcome_prj_def
+ mk_meas_outcome_def by simp
+ also have "... = sum_mat
+ (\<lambda>j. sum_mat (\<lambda>i. (Re (diag_idx_to_el B j)) \<cdot>\<^sub>m rank_1_proj (zero_col U i))
+ (diag_elem_indices (diag_idx_to_el B j) B)) {..<dist_el_card B}"
+ unfolding project_vecs_def
+ proof (rule sum_mat_cong)
+ show "finite {..<dist_el_card B}" by simp
+ show "\<And>i. i \<in> {..<dist_el_card B} \<Longrightarrow>
+ complex_of_real (Re (diag_idx_to_el B i)) \<cdot>\<^sub>m
+ sum_mat (\<lambda>i. rank_1_proj (zero_col U i)) (diag_elem_indices (diag_idx_to_el B i) B)
+ \<in> fc_mats" using assms fc_mats_carrier dimP project_vecs_def project_vecs_dim zero_col_dim
+ dim_eq by auto
+ show "\<And>i. i \<in> {..<dist_el_card B} \<Longrightarrow>
+ sum_mat (\<lambda>ia. complex_of_real (Re (diag_idx_to_el B i)) \<cdot>\<^sub>m rank_1_proj (zero_col U ia))
+ (diag_elem_indices (diag_idx_to_el B i) B) \<in> fc_mats" using assms fc_mats_carrier dimP
+ project_vecs_def project_vecs_dim zero_col_dim dim_eq
+ by (metis (no_types, lifting) rank_1_proj_carrier cpx_sq_mat_smult sum_mat_carrier)
+ show "\<And>j. j \<in> {..<dist_el_card B} \<Longrightarrow>
+ (Re (diag_idx_to_el B j)) \<cdot>\<^sub>m sum_mat (\<lambda>i. rank_1_proj (zero_col U i))
+ (diag_elem_indices (diag_idx_to_el B j) B) =
+ sum_mat (\<lambda>i. complex_of_real (Re (diag_idx_to_el B j)) \<cdot>\<^sub>m rank_1_proj (zero_col U i))
+ (diag_elem_indices (diag_idx_to_el B j) B)"
+ proof -
+ fix j
+ assume "j \<in> {..<dist_el_card B}"
+ show " (Re (diag_idx_to_el B j)) \<cdot>\<^sub>m sum_mat (\<lambda>i. rank_1_proj (zero_col U i))
+ (diag_elem_indices (diag_idx_to_el B j) B) =
+ sum_mat (\<lambda>i. (Re (diag_idx_to_el B j)) \<cdot>\<^sub>m rank_1_proj (zero_col U i))
+ (diag_elem_indices (diag_idx_to_el B j) B)"
+ proof (rule smult_sum_mat)
+ show "finite (diag_elem_indices (diag_idx_to_el B j) B)"
+ using diag_elem_indices_finite[of _ B] by simp
+ show "\<And>i. i \<in> diag_elem_indices (diag_idx_to_el B j) B \<Longrightarrow>
+ rank_1_proj (zero_col U i) \<in> fc_mats"
+ using dim_eq by (metis dimP local.fc_mats_carrier rank_1_proj_carrier zero_col_dim)
+ qed
+ qed
+ qed
+ also have "... = sum_mat
+ (\<lambda>j. sum_mat (\<lambda>ia. (diag_mat B ! ia) \<cdot>\<^sub>m rank_1_proj (zero_col U ia))
+ (diag_elem_indices (diag_idx_to_el B j) B)) {..<dist_el_card B}"
+ proof (rule sum_mat_cong)
+ show "finite {..<dist_el_card B}" by simp
+ show "\<And>i. i \<in> {..<dist_el_card B} \<Longrightarrow>
+ sum_mat (\<lambda>ia. complex_of_real (Re (diag_idx_to_el B i)) \<cdot>\<^sub>m rank_1_proj (zero_col U ia))
+ (diag_elem_indices (diag_idx_to_el B i) B) \<in> fc_mats" using assms fc_mats_carrier dimP
+ project_vecs_def project_vecs_dim zero_col_dim dim_eq
+ by (metis (no_types, lifting) rank_1_proj_carrier cpx_sq_mat_smult sum_mat_carrier)
+ show "\<And>i. i \<in> {..<dist_el_card B} \<Longrightarrow>
+ local.sum_mat (\<lambda>ia. (diag_mat B ! ia) \<cdot>\<^sub>m rank_1_proj (zero_col U ia))
+ (diag_elem_indices (diag_idx_to_el B i) B) \<in> fc_mats" using assms fc_mats_carrier dimP
+ project_vecs_def project_vecs_dim zero_col_dim dim_eq
+ by (metis (no_types, lifting) rank_1_proj_carrier cpx_sq_mat_smult sum_mat_carrier)
+ show "\<And>i. i \<in> {..<dist_el_card B} \<Longrightarrow>
+ sum_mat (\<lambda>ia. (Re (diag_idx_to_el B i)) \<cdot>\<^sub>m rank_1_proj (zero_col U ia))
+ (diag_elem_indices (diag_idx_to_el B i) B) =
+ sum_mat (\<lambda>ia. (diag_mat B ! ia) \<cdot>\<^sub>m rank_1_proj (zero_col U ia))
+ (diag_elem_indices (diag_idx_to_el B i) B)"
+ proof -
+ fix i
+ assume "i\<in> {..< dist_el_card B}"
+ show "sum_mat (\<lambda>ia. (Re (diag_idx_to_el B i)) \<cdot>\<^sub>m rank_1_proj (zero_col U ia))
+ (diag_elem_indices (diag_idx_to_el B i) B) =
+ sum_mat (\<lambda>ia. (diag_mat B ! ia) \<cdot>\<^sub>m rank_1_proj (zero_col U ia))
+ (diag_elem_indices (diag_idx_to_el B i) B)"
+ proof (rule sum_mat_cong)
+ show "finite (diag_elem_indices (diag_idx_to_el B i) B)"
+ using diag_elem_indices_finite[of _ B] by simp
+ show "\<And>ia. ia \<in> diag_elem_indices (diag_idx_to_el B i) B \<Longrightarrow>
+ (Re (diag_idx_to_el B i)) \<cdot>\<^sub>m rank_1_proj (zero_col U ia) \<in> fc_mats"
+ using assms fc_mats_carrier dimP project_vecs_def project_vecs_dim zero_col_dim dim_eq
+ by (metis (no_types, lifting) rank_1_proj_carrier cpx_sq_mat_smult)
+ show "\<And>ia. ia \<in> diag_elem_indices (diag_idx_to_el B i) B \<Longrightarrow>
+ (diag_mat B !ia) \<cdot>\<^sub>m rank_1_proj (zero_col U ia) \<in> fc_mats"
+ using assms fc_mats_carrier dimP project_vecs_def project_vecs_dim zero_col_dim dim_eq
+ by (metis (no_types, lifting) rank_1_proj_carrier cpx_sq_mat_smult)
+ show "\<And>ia. ia \<in> diag_elem_indices (diag_idx_to_el B i) B \<Longrightarrow>
+ (Re (diag_idx_to_el B i)) \<cdot>\<^sub>m rank_1_proj (zero_col U ia) =
+ (diag_mat B ! ia) \<cdot>\<^sub>m rank_1_proj (zero_col U ia)"
+ proof -
+ fix ia
+ assume ia: "ia \<in> diag_elem_indices (diag_idx_to_el B i) B"
+ hence "ia < dim_row B" using diag_elem_indices_elem[of ia _ B] by simp
+ have "Re (diag_idx_to_el B i) = Re (B $$ (ia, ia))"
+ using diag_elem_indices_elem[of ia _ B] ia by simp
+ also have "... = B $$ (ia, ia)" using Bii using \<open>ia < dim_row B\<close> dimB of_real_Re by blast
+ also have "... = diag_mat B ! ia" using \<open>ia < dim_row B\<close> unfolding diag_mat_def by simp
+ finally have "Re (diag_idx_to_el B i) = diag_mat B ! ia" .
+ thus "(Re (diag_idx_to_el B i)) \<cdot>\<^sub>m rank_1_proj (zero_col U ia) =
+ (diag_mat B ! ia) \<cdot>\<^sub>m rank_1_proj (zero_col U ia)" by simp
+ qed
+ qed
+ qed
+ qed
+ also have "... = sum_mat
+ (\<lambda>i. (diag_mat B ! i) \<cdot>\<^sub>m rank_1_proj (zero_col U i))
+ (\<Union>j<dist_el_card B. diag_elem_indices (diag_idx_to_el B j) B)"
+ unfolding project_vecs_def sum_mat_def
+ proof (rule disj_family_sum_with[symmetric])
+ show "finite {..<dist_el_card B}" by simp
+ show "\<forall>j. (diag_mat B ! j) \<cdot>\<^sub>m rank_1_proj (zero_col U j) \<in> fc_mats" using assms fc_mats_carrier dimP
+ project_vecs_def project_vecs_dim zero_col_dim dim_eq
+ by (metis (no_types, lifting) rank_1_proj_carrier cpx_sq_mat_smult)
+ show "\<And>i. i \<in> {..<dist_el_card B} \<Longrightarrow> finite (diag_elem_indices (diag_idx_to_el B i) B)"
+ by (simp add: diag_elem_indices_finite)
+ show "disjoint_family_on (\<lambda>n. diag_elem_indices (diag_idx_to_el B n) B)
+ {..<dist_el_card B}"
+ using diag_elem_indices_disjoint[of B] dimB dim_eq by simp
+ qed
+ also have "... = sum_mat (\<lambda>i. (diag_mat B ! i) \<cdot>\<^sub>m rank_1_proj (zero_col U i)) {..< dimR}"
+ using diag_elem_indices_union[of B] dimB dim_eq by simp
+ also have "... = sum_mat (\<lambda>i. (diag_mat B ! i) \<cdot>\<^sub>mrank_1_proj (Matrix.col U i)) {..< dimR}"
+ proof (rule sum_mat_cong, simp)
+ show res: "\<And>i. i \<in> {..<dimR} \<Longrightarrow> (diag_mat B ! i) \<cdot>\<^sub>m rank_1_proj (zero_col U i) \<in> fc_mats"
+ using assms fc_mats_carrier dimP project_vecs_def project_vecs_dim zero_col_dim dim_eq
+ by (metis (no_types, lifting) rank_1_proj_carrier cpx_sq_mat_smult)
+ show "\<And>i. i \<in> {..<dimR} \<Longrightarrow> (diag_mat B ! i) \<cdot>\<^sub>m rank_1_proj (Matrix.col U i) \<in> fc_mats"
+ using assms fc_mats_carrier dimP project_vecs_def project_vecs_dim zero_col_dim
+ by (metis res lessThan_iff zero_col_col)
+ show "\<And>i. i \<in> {..<dimR} \<Longrightarrow> (diag_mat B ! i) \<cdot>\<^sub>m rank_1_proj (zero_col U i) =
+ (diag_mat B ! i) \<cdot>\<^sub>m rank_1_proj (Matrix.col U i)"
+ by (simp add: zero_col_col)
+ qed
+ also have "... = U * B * Complex_Matrix.adjoint U"
+ proof (rule weighted_sum_rank_1_proj_unitary)
+ show "diagonal_mat B" using dB .
+ show "Complex_Matrix.unitary U" using unit .
+ show "U \<in> fc_mats" using fc_mats_carrier dim_eq dimP by simp
+ show "B\<in> fc_mats" using fc_mats_carrier dim_eq dimB by simp
+ qed
+ also have "... = A" using A by simp
+ finally show ?thesis .
+qed
+
+lemma (in cpx_sq_mat) trace_hermitian_pos_real:
+ fixes f::"'a \<Rightarrow> real"
+ assumes "hermitian A"
+ and "Complex_Matrix.positive R"
+ and "A \<in> fc_mats"
+ and "R \<in> fc_mats"
+shows "Complex_Matrix.trace (R * A) =
+ Re (Complex_Matrix.trace (R * A))"
+proof -
+ define prj_mems where "prj_mems = make_pm A"
+ define p where "p = proj_meas_size prj_mems"
+ define meas where "meas = proj_meas_outcomes prj_mems"
+ have tre: "Complex_Matrix.trace (R * A) =
+ Complex_Matrix.trace (R *
+ sum_mat (\<lambda>i. (meas_outcome_val (meas i)) \<cdot>\<^sub>m meas_outcome_prj (meas i)) {..< p})"
+ using make_pm_sum assms meas_def p_def proj_meas_size_def proj_meas_outcomes_def prj_mems_def
+ meas_outcome_val_def meas_outcome_prj_def by auto
+ also have "... = Re (Complex_Matrix.trace (R *
+ sum_mat (\<lambda>i. (meas_outcome_val (meas i)) \<cdot>\<^sub>m meas_outcome_prj (meas i)) {..< p}))"
+ proof (rule trace_sum_mat_proj_pos_real, (auto simp add: assms))
+ fix i
+ assume "i < p"
+ thus "projector (meas_outcome_prj (meas i))" using make_pm_projectors assms
+ unfolding p_def meas_def prj_mems_def by simp
+ show "meas_outcome_prj (meas i) \<in> fc_mats" using make_pm_square assms \<open>i < p\<close>
+ unfolding p_def meas_def prj_mems_def by simp
+ qed
+ also have "... = Re (Complex_Matrix.trace (R * A))" using tre by simp
+ finally show ?thesis .
+qed
+
+lemma (in cpx_sq_mat) hermitian_Re_spectrum:
+ assumes "hermitian A"
+and "A\<in> fc_mats"
+and "{Re x |x. x \<in> spectrum A} = {a,b}"
+shows "spectrum A = {complex_of_real a, complex_of_real b}"
+proof
+ have ar: "\<And>(a::complex). a \<in> spectrum A \<Longrightarrow> Re a = a" using hermitian_spectrum_real assms by simp
+ show "spectrum A \<subseteq> {complex_of_real a, complex_of_real b}"
+ proof
+ fix x
+ assume "x\<in> spectrum A"
+ hence "Re x = x" using ar by simp
+ have "Re x \<in> {a,b}" using \<open>x\<in> spectrum A\<close> assms by blast
+ thus "x \<in> {complex_of_real a, complex_of_real b}" using \<open>Re x = x\<close> by auto
+ qed
+ show "{complex_of_real a, complex_of_real b} \<subseteq> spectrum A"
+ proof
+ fix x
+ assume "x \<in> {complex_of_real a, complex_of_real b}"
+ hence "x \<in> {complex_of_real (Re x) |x. x \<in> spectrum A}" using assms
+ proof -
+ have "\<And>r. r \<notin> {a, b} \<or> (\<exists>c. r = Re c \<and> c \<in> spectrum A)"
+ using \<open>{Re x |x. x \<in> spectrum A} = {a, b}\<close> by blast
+ then show ?thesis
+ using \<open>x \<in> {complex_of_real a, complex_of_real b}\<close> by blast
+ qed
+ thus "x\<in> spectrum A" using ar by auto
+ qed
+qed
+
+
+subsubsection \<open>Similar properties for eigenvalues rather than spectrum indices\<close>
+
+text \<open>In this part we go the other way round: given an eigenvalue of $A$, \verb+spectrum_to_pm_idx+
+permits to retrieve its index in the associated projective measurement\<close>
+
+definition (in cpx_sq_mat) spectrum_to_pm_idx
+ where "spectrum_to_pm_idx A x = (let (B,U,_) = unitary_schur_decomposition A (eigvals A) in
+ diag_el_to_idx B x)"
+
+lemma (in cpx_sq_mat) spectrum_to_pm_idx_bij:
+assumes "hermitian A"
+ and "A\<in> fc_mats"
+shows "bij_betw (spectrum_to_pm_idx A) (spectrum A) {..< card (spectrum A)}"
+proof -
+ define p where "p = proj_meas_size (make_pm A)"
+ define M where "M = proj_meas_outcomes (make_pm A)"
+ have es: "char_poly A = (\<Prod> (e :: complex) \<leftarrow> (eigvals A). [:- e, 1:])"
+ using assms fc_mats_carrier eigvals_poly_length dim_eq by auto
+ obtain B U Q where us: "unitary_schur_decomposition A (eigvals A) = (B,U,Q)"
+ by (cases "unitary_schur_decomposition A (eigvals A)")
+ hence pr: "similar_mat_wit A B U (Complex_Matrix.adjoint U) \<and>
+ diag_mat B = (eigvals A)"
+ using hermitian_eigenvalue_real assms fc_mats_carrier es dim_eq by auto
+ have "(p,M) = make_pm A" unfolding p_def M_def using make_pm_decomp by simp
+ hence "p = dist_el_card B" using assms us unfolding make_pm_def by simp
+ have dimB: "B \<in> carrier_mat dimR dimR" using dim_eq pr assms
+ fc_mats_carrier unfolding similar_mat_wit_def by auto
+ have Bvals: "diag_mat B = eigvals A" using pr hermitian_decomp_eigenvalues[of A B U] by simp
+ have "diag_elems B = spectrum A" unfolding spectrum_def using dimB Bvals
+ diag_elems_set_diag_mat[of B] by simp
+ moreover have "dist_el_card B = card (spectrum A)" using spectrum_size[of A p M] assms
+ \<open>(p,M) = make_pm A\<close> \<open>p = dist_el_card B\<close> by simp
+ ultimately show "bij_betw (spectrum_to_pm_idx A) (spectrum A) {..< card (spectrum A)}"
+ using diag_el_to_idx_bij us unfolding spectrum_to_pm_idx_def Let_def
+ by (metis (mono_tags, lifting) bij_betw_cong case_prod_conv)
+qed
+
+lemma (in cpx_sq_mat) spectrum_to_pm_idx_lt_card:
+assumes "A\<in> fc_mats"
+ and "hermitian A"
+and "a\<in> spectrum A"
+shows "spectrum_to_pm_idx A a < card (spectrum A)" using spectrum_to_pm_idx_bij[of A] assms
+ by (meson bij_betw_apply lessThan_iff)
+
+lemma (in cpx_sq_mat) spectrum_to_pm_idx_inj:
+assumes "hermitian A"
+ and "A\<in> fc_mats"
+shows "inj_on (spectrum_to_pm_idx A) (spectrum A)" using assms spectrum_to_pm_idx_bij
+ unfolding bij_betw_def by simp
+
+lemma (in cpx_sq_mat) spectrum_meas_outcome_val_inv:
+assumes "A\<in> fc_mats"
+ and "hermitian A"
+and "make_pm A = (p,M)"
+and "i < p"
+shows "spectrum_to_pm_idx A (meas_outcome_val (M i)) = i"
+proof -
+ have es: "char_poly A = (\<Prod> (e :: complex) \<leftarrow> (eigvals A). [:- e, 1:])"
+ using assms fc_mats_carrier eigvals_poly_length dim_eq by auto
+ obtain B U Q where us: "unitary_schur_decomposition A (eigvals A) = (B,U,Q)"
+ by (cases "unitary_schur_decomposition A (eigvals A)")
+ hence pr: "similar_mat_wit A B U (Complex_Matrix.adjoint U) \<and>
+ diag_mat B = (eigvals A) \<and> (\<forall>i < dimR. B$$(i, i) \<in> Reals)"
+ using hermitian_eigenvalue_real assms fc_mats_carrier es dim_eq by auto
+ have "dim_row B = dimR" using assms fc_mats_carrier dim_eq similar_mat_wit_dim_row[of A]
+ pr by auto
+ hence "(p,M) = (dist_el_card B, mk_meas_outcome B U)" using assms us
+ unfolding make_pm_def by simp
+ hence "M = mk_meas_outcome B U" by simp
+ have "meas_outcome_val (M i) = Re (diag_idx_to_el B i)"
+ using \<open>M = mk_meas_outcome B U\<close> unfolding mk_meas_outcome_def
+ meas_outcome_val_def by simp
+ also have "... = diag_idx_to_el B i" using pr
+ \<open>(p, M) = (dist_el_card B, mk_meas_outcome B U)\<close> \<open>dim_row B = dimR\<close> assms
+ diag_idx_to_el_real by auto
+ finally have "meas_outcome_val (M i) = diag_idx_to_el B i" .
+ hence "spectrum_to_pm_idx A (meas_outcome_val (M i)) =
+ spectrum_to_pm_idx A (diag_idx_to_el B i)" by simp
+ also have "... = diag_el_to_idx B (diag_idx_to_el B i)" unfolding spectrum_to_pm_idx_def
+ using us by simp
+ also have "... = i" using assms unfolding diag_el_to_idx_def
+ using \<open>(p, M) = (dist_el_card B, mk_meas_outcome B U)\<close> bij_betw_inv_into_left
+ diag_idx_to_el_bij by fastforce
+ finally show ?thesis .
+qed
+
+lemma (in cpx_sq_mat) meas_outcome_val_spectrum_inv:
+ assumes "A\<in> fc_mats"
+ and "hermitian A"
+and "x\<in> spectrum A"
+and "make_pm A = (p,M)"
+shows "meas_outcome_val (M (spectrum_to_pm_idx A x)) = x"
+proof -
+ have es: "char_poly A = (\<Prod> (e :: complex) \<leftarrow> (eigvals A). [:- e, 1:])"
+ using assms fc_mats_carrier eigvals_poly_length dim_eq by auto
+ obtain B U Q where us: "unitary_schur_decomposition A (eigvals A) = (B,U,Q)"
+ by (cases "unitary_schur_decomposition A (eigvals A)")
+ hence pr: "similar_mat_wit A B U (Complex_Matrix.adjoint U) \<and> diagonal_mat B \<and>
+ diag_mat B = (eigvals A) \<and> unitary U \<and> (\<forall>i < dimR. B$$(i, i) \<in> Reals)"
+ using hermitian_eigenvalue_real assms fc_mats_carrier es dim_eq by auto
+ have "dim_row B = dimR" using assms fc_mats_carrier dim_eq similar_mat_wit_dim_row[of A]
+ pr by auto
+ hence "(p,M) = (dist_el_card B, mk_meas_outcome B U)" using assms us
+ unfolding make_pm_def by simp
+ hence "M = mk_meas_outcome B U" by simp
+ have "diag_mat B = (eigvals A)" using pr by simp
+ hence "x \<in> set (diag_mat B)" using \<open>diag_mat B = eigvals A\<close> assms unfolding spectrum_def by simp
+ hence "x \<in> diag_elems B" using diag_elems_set_diag_mat[of B] by simp
+ hence "diag_idx_to_el B (diag_el_to_idx B x) = x" unfolding diag_el_to_idx_def
+ by (meson bij_betw_inv_into_right diag_idx_to_el_bij)
+ moreover have "spectrum_to_pm_idx A x = diag_el_to_idx B x" unfolding spectrum_to_pm_idx_def
+ using us by simp
+ moreover have "meas_outcome_val (M (spectrum_to_pm_idx A x)) =
+ Re (diag_idx_to_el B (diag_el_to_idx B x))" using \<open>M = mk_meas_outcome B U\<close>
+ unfolding mk_meas_outcome_def meas_outcome_val_def by (simp add: calculation(2))
+ ultimately show ?thesis by simp
+qed
+
+definition (in cpx_sq_mat) eigen_projector where
+"eigen_projector A a =
+ meas_outcome_prj ((proj_meas_outcomes (make_pm A)) (spectrum_to_pm_idx A a))"
+
+lemma (in cpx_sq_mat) eigen_projector_carrier:
+ assumes "A\<in> fc_mats"
+ and "a\<in> spectrum A"
+ and "hermitian A"
+shows "eigen_projector A a \<in> fc_mats" unfolding eigen_projector_def
+ using meas_outcome_prj_carrier[of A "spectrum_to_pm_idx A a"]
+ spectrum_to_pm_idx_lt_card[of A a] assms by simp
+
+text \<open>We obtain the following result, which is similar to \verb+make_pm_sum+ but with a sum on
+the elements of the spectrum of Hermitian matrix $A$, which is a more standard statement of the
+spectral decomposition theorem.\<close>
+
+lemma (in cpx_sq_mat) make_pm_sum':
+ assumes "A \<in> fc_mats"
+ and "hermitian A"
+shows "sum_mat (\<lambda>a. a \<cdot>\<^sub>m (eigen_projector A a)) (spectrum A) = A"
+proof -
+ define p where "p = proj_meas_size (make_pm A)"
+ define M where "M = proj_meas_outcomes (make_pm A)"
+ have "(p,M) = make_pm A" unfolding p_def M_def using make_pm_decomp by simp
+ define g where
+ "g = (\<lambda>i. (if i < p
+ then complex_of_real (meas_outcome_val (M i)) \<cdot>\<^sub>m meas_outcome_prj (M i)
+ else (0\<^sub>m dimR dimC)))"
+ have g: "\<forall>x. g x \<in> fc_mats"
+ proof
+ fix x
+ show "g x \<in> fc_mats"
+ proof (cases "x < p")
+ case True
+ hence "(meas_outcome_val (M x)) \<cdot>\<^sub>m meas_outcome_prj (M x) \<in> fc_mats"
+ using meas_outcome_prj_carrier
+ spectrum_size assms cpx_sq_mat_smult make_pm_proj_measurement proj_measurement_square
+ \<open>(p,M) = make_pm A\<close> by metis
+ then show ?thesis unfolding g_def using True by simp
+ next
+ case False
+ then show ?thesis unfolding g_def using zero_mem by simp
+ qed
+ qed
+ define h where "h = (\<lambda>a. (if a \<in> (spectrum A) then a \<cdot>\<^sub>m eigen_projector A a else (0\<^sub>m dimR dimC)))"
+ have h: "\<forall>x. h x \<in> fc_mats"
+ proof
+ fix x
+ show "h x \<in> fc_mats"
+ proof (cases "x\<in> spectrum A")
+ case True
+ then show ?thesis unfolding h_def using eigen_projector_carrier[of A x] assms True
+ by (simp add: cpx_sq_mat_smult)
+ next
+ case False
+ then show ?thesis unfolding h_def using zero_mem by simp
+ qed
+ qed
+ have "inj_on (spectrum_to_pm_idx A) (spectrum A)" using assms spectrum_to_pm_idx_inj by simp
+ moreover have "{..<p} = spectrum_to_pm_idx A ` spectrum A" using \<open>(p,M) = make_pm A\<close>
+ spectrum_to_pm_idx_bij spectrum_size unfolding bij_betw_def
+ by (metis assms(1) assms(2))
+ moreover have "\<And>x. x \<in> spectrum A \<Longrightarrow> g (spectrum_to_pm_idx A x) = h x"
+ proof -
+ fix a
+ assume "a \<in> spectrum A"
+ hence "Re a = a" using hermitian_spectrum_real assms by simp
+ have "spectrum_to_pm_idx A a < p" using spectrum_to_pm_idx_lt_card[of A] spectrum_size assms
+ \<open>a \<in> spectrum A\<close> \<open>(p,M) = make_pm A\<close> by metis
+ have "g (spectrum_to_pm_idx A a) =
+ (meas_outcome_val (M (spectrum_to_pm_idx A a))) \<cdot>\<^sub>m
+ meas_outcome_prj (M (spectrum_to_pm_idx A a))"
+ using \<open>spectrum_to_pm_idx A a < p\<close> unfolding g_def by simp
+ also have "... = a \<cdot>\<^sub>m meas_outcome_prj (M (spectrum_to_pm_idx A a))"
+ using meas_outcome_val_spectrum_inv[of A "Re a"] \<open>Re a = a\<close> assms \<open>a \<in> spectrum A\<close>
+ \<open>(p,M) = make_pm A\<close> by metis
+ also have "... = h a" using \<open>a \<in> spectrum A\<close> unfolding eigen_projector_def M_def h_def by simp
+ finally show "g (spectrum_to_pm_idx A a) = h a" .
+ qed
+ ultimately have "sum_mat h (spectrum A) =
+ sum_mat g (spectrum_to_pm_idx A ` spectrum A)" unfolding sum_mat_def
+ using sum_with_reindex_cong[symmetric, of g h "spectrum_to_pm_idx A" "spectrum A" "{..< p}"]
+ g h by simp
+ also have "... = sum_mat g {..< p}" using \<open>{..<p} = spectrum_to_pm_idx A ` spectrum A\<close> by simp
+ also have "... = sum_mat (\<lambda>i. (meas_outcome_val (M i)) \<cdot>\<^sub>m meas_outcome_prj (M i)) {..<p}"
+ proof (rule sum_mat_cong, simp)
+ fix i
+ assume "i \<in> {..< p}"
+ hence "i < p" by simp
+ show "g i \<in> fc_mats" using g by simp
+ show "g i = (meas_outcome_val (M i)) \<cdot>\<^sub>m meas_outcome_prj (M i)" unfolding g_def
+ using \<open>i < p\<close> by simp
+ show "(meas_outcome_val (M i)) \<cdot>\<^sub>m meas_outcome_prj (M i) \<in> fc_mats"
+ using meas_outcome_prj_carrier spectrum_size assms cpx_sq_mat_smult
+ make_pm_proj_measurement proj_measurement_square \<open>i < p\<close> \<open>(p,M) = make_pm A\<close> by metis
+ qed
+ also have "... = A" using make_pm_sum assms \<open>(p,M) = make_pm A\<close> unfolding g_def by simp
+ finally have "sum_mat h (spectrum A) = A" .
+ moreover have "sum_mat h (spectrum A) = sum_mat (\<lambda>a. a \<cdot>\<^sub>m (eigen_projector A a)) (spectrum A)"
+ proof (rule sum_mat_cong)
+ show "finite (spectrum A)" using spectrum_finite[of A] by simp
+ fix i
+ assume "i\<in> spectrum A"
+ thus "h i = i \<cdot>\<^sub>m eigen_projector A i" unfolding h_def by simp
+ show "h i \<in> fc_mats" using h by simp
+ show "i \<cdot>\<^sub>m eigen_projector A i \<in> fc_mats" using eigen_projector_carrier[of A i] assms
+ \<open>i\<in> spectrum A\<close> by (metis \<open>h i = i \<cdot>\<^sub>m eigen_projector A i\<close> h)
+ qed
+ ultimately show ?thesis by simp
+qed
+
+
+
end
\ No newline at end of file
diff --git a/thys/Regression_Test_Selection/Common/CollectionBasedRTS.thy b/thys/Regression_Test_Selection/Common/CollectionBasedRTS.thy
--- a/thys/Regression_Test_Selection/Common/CollectionBasedRTS.thy
+++ b/thys/Regression_Test_Selection/Common/CollectionBasedRTS.thy
@@ -1,32 +1,32 @@
-(* Title: RTS/Common/CollectionBasedRTS.thy *)
-(* Author: Susannah Mansky, UIUC 2020 *)
-
-section "Collection-based RTS"
-
-theory CollectionBasedRTS
-imports RTS_safe CollectionSemantics
-begin
-
-locale CollectionBasedRTS_base = RTS_safe + CollectionSemantics
-
-text "General model for Regression Test Selection based on
- @{term CollectionSemantics}:"
-locale CollectionBasedRTS = CollectionBasedRTS_base where
- small = "small :: 'prog \<Rightarrow> 'state \<Rightarrow> 'state set" and
- collect = "collect :: 'prog \<Rightarrow> 'state \<Rightarrow> 'state \<Rightarrow> 'coll" and
- out = "out :: 'prog \<Rightarrow> 'test \<Rightarrow> ('state \<times> 'coll) set"
- for small collect out
-+
- fixes
- make_test_prog :: "'prog \<Rightarrow> 'test \<Rightarrow> 'prog" and
- collect_start :: "'prog \<Rightarrow> 'coll"
- assumes
- out_cbig:
- "\<exists>i. out P t = {(\<sigma>',coll'). \<exists>coll. (\<sigma>',coll) \<in> cbig (make_test_prog P t) i
- \<and> coll' = combine coll (collect_start P) }"
-
-context CollectionBasedRTS begin
-
-end \<comment> \<open> CollectionBasedRTS \<close>
-
-end
+(* Title: RTS/Common/CollectionBasedRTS.thy *)
+(* Author: Susannah Mansky, UIUC 2020 *)
+
+section "Collection-based RTS"
+
+theory CollectionBasedRTS
+imports RTS_safe CollectionSemantics
+begin
+
+locale CollectionBasedRTS_base = RTS_safe + CollectionSemantics
+
+text "General model for Regression Test Selection based on
+ @{term CollectionSemantics}:"
+locale CollectionBasedRTS = CollectionBasedRTS_base where
+ small = "small :: 'prog \<Rightarrow> 'state \<Rightarrow> 'state set" and
+ collect = "collect :: 'prog \<Rightarrow> 'state \<Rightarrow> 'state \<Rightarrow> 'coll" and
+ out = "out :: 'prog \<Rightarrow> 'test \<Rightarrow> ('state \<times> 'coll) set"
+ for small collect out
++
+ fixes
+ make_test_prog :: "'prog \<Rightarrow> 'test \<Rightarrow> 'prog" and
+ collect_start :: "'prog \<Rightarrow> 'coll"
+ assumes
+ out_cbig:
+ "\<exists>i. out P t = {(\<sigma>',coll'). \<exists>coll. (\<sigma>',coll) \<in> cbig (make_test_prog P t) i
+ \<and> coll' = combine coll (collect_start P) }"
+
+context CollectionBasedRTS begin
+
+end \<comment> \<open> CollectionBasedRTS \<close>
+
+end
diff --git a/thys/Regression_Test_Selection/Common/CollectionSemantics.thy b/thys/Regression_Test_Selection/Common/CollectionSemantics.thy
--- a/thys/Regression_Test_Selection/Common/CollectionSemantics.thy
+++ b/thys/Regression_Test_Selection/Common/CollectionSemantics.thy
@@ -1,258 +1,258 @@
-(* Title: RTS/Common/CollectionSemantics.thy *)
-(* Author: Susannah Mansky, UIUC 2020 *)
-
-section "Collection Semantics"
-
-theory CollectionSemantics
-imports Semantics
-begin
-
-text "General model for small step semantics instrumented
- with an information collection mechanism:"
-locale CollectionSemantics = Semantics +
- constrains
- small :: "'prog \<Rightarrow> 'state \<Rightarrow> 'state set" and
- endset :: "'state set"
- fixes
- collect :: "'prog \<Rightarrow> 'state \<Rightarrow> 'state \<Rightarrow> 'coll" and
- combine :: "'coll \<Rightarrow> 'coll \<Rightarrow> 'coll" and
- collect_id :: "'coll"
- assumes
- combine_assoc: "combine (combine c1 c2) c3 = combine c1 (combine c2 c3)" and
- collect_idl[simp]: "combine collect_id c = c" and
- collect_idr[simp]: "combine c collect_id = c"
-
-context CollectionSemantics begin
-
-subsection "Small-Step Collection Semantics"
-
-definition csmall :: "'prog \<Rightarrow> 'state \<Rightarrow> ('state \<times> 'coll) set" where
-"csmall P \<sigma> \<equiv> { (\<sigma>', coll). \<sigma>' \<in> small P \<sigma> \<and> collect P \<sigma> \<sigma>' = coll }"
-
-lemma small_det_csmall_det:
-assumes "\<forall>\<sigma>. small P \<sigma> = {} \<or> (\<exists>\<sigma>'. small P \<sigma> = {\<sigma>'})"
-shows "\<forall>\<sigma>. csmall P \<sigma> = {} \<or> (\<exists>o'. csmall P \<sigma> = {o'})"
-using assms by(fastforce simp: csmall_def)
-
-subsection "Extending @{term csmall} to multiple steps"
-
-primrec csmall_nstep :: "'prog \<Rightarrow> 'state \<Rightarrow> nat \<Rightarrow> ('state \<times> 'coll) set" where
-csmall_nstep_base:
- "csmall_nstep P \<sigma> 0 = {(\<sigma>, collect_id)}" |
-csmall_nstep_Rec:
- "csmall_nstep P \<sigma> (Suc n) =
- { (\<sigma>2, coll). \<exists>\<sigma>1 coll1 coll2. (\<sigma>1, coll1) \<in> csmall_nstep P \<sigma> n
- \<and> (\<sigma>2, coll2) \<in> csmall P \<sigma>1 \<and> combine coll1 coll2 = coll }"
-
-lemma small_nstep_csmall_nstep_equiv:
- "small_nstep P \<sigma> n
- = { \<sigma>'. \<exists>coll. (\<sigma>', coll) \<in> csmall_nstep P \<sigma> n }"
-proof (induct n) qed(simp_all add: csmall_def)
-
-lemma csmall_nstep_exists:
- "\<sigma>' \<in> big P \<sigma> \<Longrightarrow> \<exists>n coll. (\<sigma>', coll) \<in> csmall_nstep P \<sigma> n \<and> \<sigma>' \<in> endset"
-proof(drule bigD) qed(clarsimp simp: small_nstep_csmall_nstep_equiv)
-
-lemma csmall_det_csmall_nstep_det:
-assumes "\<forall>\<sigma>. csmall P \<sigma> = {} \<or> (\<exists>o'. csmall P \<sigma> = {o'})"
-shows "\<forall>\<sigma>. csmall_nstep P \<sigma> n = {} \<or> (\<exists>o'. csmall_nstep P \<sigma> n = {o'})"
-using assms
-proof(induct n)
- case (Suc n) then show ?case by fastforce
-qed(simp)
-
-lemma csmall_nstep_Rec2:
- "csmall_nstep P \<sigma> (Suc n) =
- { (\<sigma>2, coll). \<exists>\<sigma>1 coll1 coll2. (\<sigma>1, coll1) \<in> csmall P \<sigma>
- \<and> (\<sigma>2, coll2) \<in> csmall_nstep P \<sigma>1 n \<and> combine coll1 coll2 = coll }"
-proof(induct n arbitrary: \<sigma>)
- case (Suc n)
- have right: "\<And>\<sigma>' coll'. (\<sigma>', coll') \<in> csmall_nstep P \<sigma> (Suc(Suc n))
- \<Longrightarrow> \<exists>\<sigma>1 coll1 coll2. (\<sigma>1, coll1) \<in> csmall P \<sigma>
- \<and> (\<sigma>', coll2) \<in> csmall_nstep P \<sigma>1 (Suc n) \<and> combine coll1 coll2 = coll'"
- proof -
- fix \<sigma>' coll'
- assume "(\<sigma>', coll') \<in> csmall_nstep P \<sigma> (Suc(Suc n))"
- then obtain \<sigma>1 coll1 coll2 where Sucnstep: "(\<sigma>1, coll1) \<in> csmall_nstep P \<sigma> (Suc n)"
- "(\<sigma>', coll2) \<in> csmall P \<sigma>1" "combine coll1 coll2 = coll'" by fastforce
- obtain \<sigma>12 coll12 coll22 where nstep: "(\<sigma>12, coll12) \<in> csmall P \<sigma>
- \<and> (\<sigma>1, coll22) \<in> csmall_nstep P \<sigma>12 n \<and> combine coll12 coll22 = coll1"
- using Suc Sucnstep(1) by fastforce
- then show "\<exists>\<sigma>1 coll1 coll2. (\<sigma>1, coll1) \<in> csmall P \<sigma>
- \<and> (\<sigma>', coll2) \<in> csmall_nstep P \<sigma>1 (Suc n) \<and> combine coll1 coll2 = coll'"
- using combine_assoc[of coll12 coll22 coll2] Sucnstep by fastforce
- qed
- have left: "\<And>\<sigma>' coll'. \<exists>\<sigma>1 coll1 coll2. (\<sigma>1, coll1) \<in> csmall P \<sigma>
- \<and> (\<sigma>', coll2) \<in> csmall_nstep P \<sigma>1 (Suc n) \<and> combine coll1 coll2 = coll'
- \<Longrightarrow> (\<sigma>', coll') \<in> csmall_nstep P \<sigma> (Suc(Suc n))"
- proof -
- fix \<sigma>' coll'
- assume "\<exists>\<sigma>1 coll1 coll2. (\<sigma>1, coll1) \<in> csmall P \<sigma>
- \<and> (\<sigma>', coll2) \<in> csmall_nstep P \<sigma>1 (Suc n) \<and> combine coll1 coll2 = coll'"
- then obtain \<sigma>1 coll1 coll2 where Sucnstep: "(\<sigma>1, coll1) \<in> csmall P \<sigma>"
- "(\<sigma>', coll2) \<in> csmall_nstep P \<sigma>1 (Suc n)" "combine coll1 coll2 = coll'"
- by fastforce
- obtain \<sigma>12 coll12 coll22 where nstep: "(\<sigma>12, coll12) \<in> csmall_nstep P \<sigma>1 n
- \<and> (\<sigma>', coll22) \<in> csmall P \<sigma>12 \<and> combine coll12 coll22 = coll2"
- using Sucnstep(2) by auto
- then show "(\<sigma>', coll') \<in> csmall_nstep P \<sigma> (Suc(Suc n))"
- using combine_assoc[of coll1 coll12 coll22] Suc Sucnstep by fastforce
- qed
- show ?case using right left by fast
-qed(simp)
-
-lemma csmall_nstep_SucD:
-assumes "(\<sigma>',coll') \<in> csmall_nstep P \<sigma> (Suc n)"
-shows "\<exists>\<sigma>1 coll1. (\<sigma>1, coll1) \<in> csmall P \<sigma>
- \<and> (\<exists>coll. coll' = combine coll1 coll \<and> (\<sigma>',coll) \<in> csmall_nstep P \<sigma>1 n)"
- using csmall_nstep_Rec2 CollectionSemantics_axioms case_prodD assms by fastforce
-
-lemma csmall_nstep_Suc_nend: "o' \<in> csmall_nstep P \<sigma> (Suc n1) \<Longrightarrow> \<sigma> \<notin> endset"
- using endset_final csmall_nstep_SucD CollectionSemantics.csmall_def CollectionSemantics_axioms
- by fastforce
-
-lemma small_to_csmall_nstep_pres:
-assumes Qpres: "\<And>P \<sigma> \<sigma>'. Q P \<sigma> \<Longrightarrow> \<sigma>' \<in> small P \<sigma> \<Longrightarrow> Q P \<sigma>'"
-shows "Q P \<sigma> \<Longrightarrow> (\<sigma>', coll) \<in> csmall_nstep P \<sigma> n \<Longrightarrow> Q P \<sigma>'"
-proof(induct n arbitrary: \<sigma> \<sigma>' coll)
- case (Suc n)
- then obtain \<sigma>1 coll1 coll2 where nstep: "(\<sigma>1, coll1) \<in> csmall_nstep P \<sigma> n
- \<and> (\<sigma>', coll2) \<in> csmall P \<sigma>1 \<and> combine coll1 coll2 = coll" by clarsimp
- then show ?case using Suc Qpres[where P=P and \<sigma>=\<sigma>1 and \<sigma>'=\<sigma>'] by(auto simp: csmall_def)
-qed(simp)
-
-lemma csmall_to_csmall_nstep_prop:
-assumes cond: "\<And>P \<sigma> \<sigma>' coll. (\<sigma>', coll) \<in> csmall P \<sigma> \<Longrightarrow> R P coll \<Longrightarrow> Q P \<sigma> \<Longrightarrow> R' P \<sigma> \<sigma>' coll"
- and Rcomb: "\<And>P coll1 coll2. R P (combine coll1 coll2) = (R P coll1 \<and> R P coll2)"
- and Qpres: "\<And>P \<sigma> \<sigma>'. Q P \<sigma> \<Longrightarrow> \<sigma>' \<in> small P \<sigma> \<Longrightarrow> Q P \<sigma>'"
- and Rtrans': "\<And>P \<sigma> \<sigma>1 \<sigma>' coll1 coll2.
- R' P \<sigma> \<sigma>1 coll1 \<and> R' P \<sigma>1 \<sigma>' coll2 \<Longrightarrow> R' P \<sigma> \<sigma>' (combine coll1 coll2)"
- and base: "\<And>\<sigma>. R' P \<sigma> \<sigma> collect_id"
-shows "(\<sigma>', coll) \<in> csmall_nstep P \<sigma> n \<Longrightarrow> R P coll \<Longrightarrow> Q P \<sigma> \<Longrightarrow> R' P \<sigma> \<sigma>' coll"
-proof(induct n arbitrary: \<sigma> \<sigma>' coll)
- case (Suc n)
- then obtain \<sigma>1 coll1 coll2 where nstep: "(\<sigma>1, coll1) \<in> csmall_nstep P \<sigma> n
- \<and> (\<sigma>', coll2) \<in> csmall P \<sigma>1 \<and> combine coll1 coll2 = coll" by clarsimp
- then have "Q P \<sigma>1" using small_to_csmall_nstep_pres[where Q=Q] Qpres Suc by blast
- then show ?case using nstep assms Suc by auto blast
-qed(simp add: base)
-
-lemma csmall_to_csmall_nstep_prop2:
-assumes cond: "\<And>P P' \<sigma> \<sigma>' coll. (\<sigma>', coll) \<in> csmall P \<sigma>
- \<Longrightarrow> R P P' coll \<Longrightarrow> Q \<sigma> \<Longrightarrow> (\<sigma>', coll) \<in> csmall P' \<sigma>"
- and Rcomb: "\<And>P P' coll1 coll2. R P P' (combine coll1 coll2) = (R P P' coll1 \<and> R P P' coll2)"
- and Qpres: "\<And>P \<sigma> \<sigma>'. Q \<sigma> \<Longrightarrow> \<sigma>' \<in> small P \<sigma> \<Longrightarrow> Q \<sigma>'"
-shows "(\<sigma>', coll) \<in> csmall_nstep P \<sigma> n \<Longrightarrow> R P P' coll \<Longrightarrow> Q \<sigma> \<Longrightarrow> (\<sigma>', coll) \<in> csmall_nstep P' \<sigma> n"
-proof(induct n arbitrary: \<sigma> \<sigma>' coll)
- case (Suc n)
- then obtain \<sigma>1 coll1 coll2 where nstep: "(\<sigma>1, coll1) \<in> csmall_nstep P \<sigma> n
- \<and> (\<sigma>', coll2) \<in> csmall P \<sigma>1 \<and> combine coll1 coll2 = coll" by clarsimp
- then have "Q \<sigma>1" using small_to_csmall_nstep_pres[where Q="\<lambda>P. Q"] Qpres Suc by blast
- then show ?case using nstep assms Suc by auto blast
-qed(simp)
-
-subsection "Extending @{term csmall} to a big-step semantics"
-
-definition cbig :: "'prog \<Rightarrow> 'state \<Rightarrow> ('state \<times> 'coll) set" where
-"cbig P \<sigma> \<equiv>
- { (\<sigma>', coll). \<exists>n. (\<sigma>', coll) \<in> csmall_nstep P \<sigma> n \<and> \<sigma>' \<in> endset }"
-
-lemma cbigD:
- "\<lbrakk> (\<sigma>',coll') \<in> cbig P \<sigma> \<rbrakk> \<Longrightarrow> \<exists>n. (\<sigma>',coll') \<in> csmall_nstep P \<sigma> n \<and> \<sigma>' \<in> endset"
- by(simp add: cbig_def)
-
-lemma cbigD':
- "\<lbrakk> o' \<in> cbig P \<sigma> \<rbrakk> \<Longrightarrow> \<exists>n. o' \<in> csmall_nstep P \<sigma> n \<and> fst o' \<in> endset"
- by(cases o', simp add: cbig_def)
-
-lemma cbig_def2:
- "(\<sigma>', coll) \<in> cbig P \<sigma> \<longleftrightarrow> (\<exists>n. (\<sigma>', coll) \<in> csmall_nstep P \<sigma> n \<and> \<sigma>' \<in> endset)"
-proof(rule iffI)
- assume "(\<sigma>', coll) \<in> cbig P \<sigma>"
- then show "\<exists>n. (\<sigma>', coll) \<in> csmall_nstep P \<sigma> n \<and> \<sigma>' \<in> endset" using bigD cbig_def by auto
-next
- assume "\<exists>n. (\<sigma>', coll) \<in> csmall_nstep P \<sigma> n \<and> \<sigma>' \<in> endset"
- then show "(\<sigma>', coll) \<in> cbig P \<sigma>" using big_def cbig_def small_nstep_csmall_nstep_equiv by auto
-qed
-
-lemma cbig_big_equiv:
- "(\<exists>coll. (\<sigma>', coll) \<in> cbig P \<sigma>) \<longleftrightarrow> \<sigma>' \<in> big P \<sigma>"
-proof(rule iffI)
- assume "\<exists>coll. (\<sigma>', coll) \<in> cbig P \<sigma>"
- then show "\<sigma>' \<in> big P \<sigma>" by (auto simp: big_def cbig_def small_nstep_csmall_nstep_equiv)
-next
- assume "\<sigma>' \<in> big P \<sigma>"
- then show "\<exists>coll. (\<sigma>', coll) \<in> cbig P \<sigma>" by (fastforce simp: cbig_def dest: csmall_nstep_exists)
-qed
-
-lemma cbig_big_implies:
- "(\<sigma>', coll) \<in> cbig P \<sigma> \<Longrightarrow> \<sigma>' \<in> big P \<sigma>"
-using cbig_big_equiv by blast
-
-lemma csmall_to_cbig_prop:
-assumes "\<And>P \<sigma> \<sigma>' coll. (\<sigma>', coll) \<in> csmall P \<sigma> \<Longrightarrow> R P coll \<Longrightarrow> Q P \<sigma> \<Longrightarrow> R' P \<sigma> \<sigma>' coll"
- and "\<And>P coll1 coll2. R P (combine coll1 coll2) = (R P coll1 \<and> R P coll2)"
- and "\<And>P \<sigma> \<sigma>'. Q P \<sigma> \<Longrightarrow> \<sigma>' \<in> small P \<sigma> \<Longrightarrow> Q P \<sigma>'"
- and "\<And>P \<sigma> \<sigma>1 \<sigma>' coll1 coll2.
- R' P \<sigma> \<sigma>1 coll1 \<and> R' P \<sigma>1 \<sigma>' coll2 \<Longrightarrow> R' P \<sigma> \<sigma>' (combine coll1 coll2)"
- and "\<And>\<sigma>. R' P \<sigma> \<sigma> collect_id"
-shows "(\<sigma>', coll) \<in> cbig P \<sigma> \<Longrightarrow> R P coll \<Longrightarrow> Q P \<sigma> \<Longrightarrow> R' P \<sigma> \<sigma>' coll"
-using assms csmall_to_csmall_nstep_prop[where R=R and Q=Q and R'=R' and \<sigma>=\<sigma>]
- by(auto simp: cbig_def2)
-
-lemma csmall_to_cbig_prop2:
-assumes "\<And>P P' \<sigma> \<sigma>' coll. (\<sigma>', coll) \<in> csmall P \<sigma> \<Longrightarrow> R P P' coll \<Longrightarrow> Q \<sigma> \<Longrightarrow> (\<sigma>', coll) \<in> csmall P' \<sigma>"
- and "\<And>P P' coll1 coll2. R P P' (combine coll1 coll2) = (R P P' coll1 \<and> R P P' coll2)"
- and Qpres: "\<And>P \<sigma> \<sigma>'. Q \<sigma> \<Longrightarrow> \<sigma>' \<in> small P \<sigma> \<Longrightarrow> Q \<sigma>'"
-shows "(\<sigma>', coll) \<in> cbig P \<sigma> \<Longrightarrow> R P P' coll \<Longrightarrow> Q \<sigma> \<Longrightarrow> (\<sigma>', coll) \<in> cbig P' \<sigma>"
-using assms csmall_to_csmall_nstep_prop2[where R=R and Q=Q] by(auto simp: cbig_def2) blast
-
-lemma cbig_stepD:
-assumes cbig: "(\<sigma>',coll') \<in> cbig P \<sigma>" and nend: "\<sigma> \<notin> endset"
-shows "\<exists>\<sigma>1 coll1. (\<sigma>1, coll1) \<in> csmall P \<sigma>
- \<and> (\<exists>coll. coll' = combine coll1 coll \<and> (\<sigma>',coll) \<in> cbig P \<sigma>1)"
-proof -
- obtain n where n: "(\<sigma>', coll') \<in> csmall_nstep P \<sigma> n" "\<sigma>' \<in> endset"
- using cbig_def2 cbig by auto
- then show ?thesis using csmall_nstep_SucD nend cbig_def2 by(cases n, simp) blast
-qed
-
-(****)
-
-lemma csmall_nstep_det_last_eq:
-assumes det: "\<forall>\<sigma>. small P \<sigma> = {} \<or> (\<exists>\<sigma>'. small P \<sigma> = {\<sigma>'})"
-shows "\<lbrakk> (\<sigma>',coll') \<in> cbig P \<sigma>; (\<sigma>',coll') \<in> csmall_nstep P \<sigma> n; (\<sigma>',coll'') \<in> csmall_nstep P \<sigma> n' \<rbrakk>
- \<Longrightarrow> n = n'"
-proof(induct n arbitrary: n' \<sigma> \<sigma>' coll' coll'')
- case 0
- have "\<sigma>' = \<sigma>" using "0.prems"(2) csmall_nstep_base by blast
- then have endset: "\<sigma> \<in> endset" using "0.prems"(1) cbigD by blast
- show ?case
- proof(cases n')
- case Suc then show ?thesis using "0.prems"(3) csmall_nstep_Suc_nend endset by blast
- qed(simp)
-next
- case (Suc n1)
- then have endset: "\<sigma>' \<in> endset" using Suc.prems(1) cbigD by blast
- have nend: "\<sigma> \<notin> endset" using csmall_nstep_Suc_nend[OF Suc.prems(2)] by simp
- then have neq: "\<sigma>' \<noteq> \<sigma>" using endset by auto
- obtain \<sigma>1 coll coll1 where \<sigma>1: "(\<sigma>1,coll1) \<in> csmall P \<sigma>" "coll' = combine coll1 coll"
- "(\<sigma>',coll) \<in> csmall_nstep P \<sigma>1 n1"
- using csmall_nstep_SucD[OF Suc.prems(2)] by clarsimp
- then have cbig: "(\<sigma>',coll) \<in> cbig P \<sigma>1" using endset by(auto simp: cbig_def)
- show ?case
- proof(cases n')
- case 0 then show ?thesis using neq Suc.prems(3) using csmall_nstep_base by simp
- next
- case Suc': (Suc n1')
- then obtain \<sigma>1' coll2 coll1' where \<sigma>1': "(\<sigma>1',coll1') \<in> csmall P \<sigma>" "coll'' = combine coll1' coll2"
- "(\<sigma>',coll2) \<in> csmall_nstep P \<sigma>1' n1'"
- using csmall_nstep_SucD[where \<sigma>=\<sigma> and \<sigma>'=\<sigma>' and coll'=coll'' and n=n1'] Suc.prems(3) by blast
- then have "\<sigma>1=\<sigma>1'" using \<sigma>1 det csmall_def by auto
- then show ?thesis using Suc.hyps(1)[OF cbig \<sigma>1(3)] \<sigma>1'(3) Suc' by blast
- qed
-qed
-
-end \<comment> \<open> CollectionSemantics \<close>
-
+(* Title: RTS/Common/CollectionSemantics.thy *)
+(* Author: Susannah Mansky, UIUC 2020 *)
+
+section "Collection Semantics"
+
+theory CollectionSemantics
+imports Semantics
+begin
+
+text "General model for small step semantics instrumented
+ with an information collection mechanism:"
+locale CollectionSemantics = Semantics +
+ constrains
+ small :: "'prog \<Rightarrow> 'state \<Rightarrow> 'state set" and
+ endset :: "'state set"
+ fixes
+ collect :: "'prog \<Rightarrow> 'state \<Rightarrow> 'state \<Rightarrow> 'coll" and
+ combine :: "'coll \<Rightarrow> 'coll \<Rightarrow> 'coll" and
+ collect_id :: "'coll"
+ assumes
+ combine_assoc: "combine (combine c1 c2) c3 = combine c1 (combine c2 c3)" and
+ collect_idl[simp]: "combine collect_id c = c" and
+ collect_idr[simp]: "combine c collect_id = c"
+
+context CollectionSemantics begin
+
+subsection "Small-Step Collection Semantics"
+
+definition csmall :: "'prog \<Rightarrow> 'state \<Rightarrow> ('state \<times> 'coll) set" where
+"csmall P \<sigma> \<equiv> { (\<sigma>', coll). \<sigma>' \<in> small P \<sigma> \<and> collect P \<sigma> \<sigma>' = coll }"
+
+lemma small_det_csmall_det:
+assumes "\<forall>\<sigma>. small P \<sigma> = {} \<or> (\<exists>\<sigma>'. small P \<sigma> = {\<sigma>'})"
+shows "\<forall>\<sigma>. csmall P \<sigma> = {} \<or> (\<exists>o'. csmall P \<sigma> = {o'})"
+using assms by(fastforce simp: csmall_def)
+
+subsection "Extending @{term csmall} to multiple steps"
+
+primrec csmall_nstep :: "'prog \<Rightarrow> 'state \<Rightarrow> nat \<Rightarrow> ('state \<times> 'coll) set" where
+csmall_nstep_base:
+ "csmall_nstep P \<sigma> 0 = {(\<sigma>, collect_id)}" |
+csmall_nstep_Rec:
+ "csmall_nstep P \<sigma> (Suc n) =
+ { (\<sigma>2, coll). \<exists>\<sigma>1 coll1 coll2. (\<sigma>1, coll1) \<in> csmall_nstep P \<sigma> n
+ \<and> (\<sigma>2, coll2) \<in> csmall P \<sigma>1 \<and> combine coll1 coll2 = coll }"
+
+lemma small_nstep_csmall_nstep_equiv:
+ "small_nstep P \<sigma> n
+ = { \<sigma>'. \<exists>coll. (\<sigma>', coll) \<in> csmall_nstep P \<sigma> n }"
+proof (induct n) qed(simp_all add: csmall_def)
+
+lemma csmall_nstep_exists:
+ "\<sigma>' \<in> big P \<sigma> \<Longrightarrow> \<exists>n coll. (\<sigma>', coll) \<in> csmall_nstep P \<sigma> n \<and> \<sigma>' \<in> endset"
+proof(drule bigD) qed(clarsimp simp: small_nstep_csmall_nstep_equiv)
+
+lemma csmall_det_csmall_nstep_det:
+assumes "\<forall>\<sigma>. csmall P \<sigma> = {} \<or> (\<exists>o'. csmall P \<sigma> = {o'})"
+shows "\<forall>\<sigma>. csmall_nstep P \<sigma> n = {} \<or> (\<exists>o'. csmall_nstep P \<sigma> n = {o'})"
+using assms
+proof(induct n)
+ case (Suc n) then show ?case by fastforce
+qed(simp)
+
+lemma csmall_nstep_Rec2:
+ "csmall_nstep P \<sigma> (Suc n) =
+ { (\<sigma>2, coll). \<exists>\<sigma>1 coll1 coll2. (\<sigma>1, coll1) \<in> csmall P \<sigma>
+ \<and> (\<sigma>2, coll2) \<in> csmall_nstep P \<sigma>1 n \<and> combine coll1 coll2 = coll }"
+proof(induct n arbitrary: \<sigma>)
+ case (Suc n)
+ have right: "\<And>\<sigma>' coll'. (\<sigma>', coll') \<in> csmall_nstep P \<sigma> (Suc(Suc n))
+ \<Longrightarrow> \<exists>\<sigma>1 coll1 coll2. (\<sigma>1, coll1) \<in> csmall P \<sigma>
+ \<and> (\<sigma>', coll2) \<in> csmall_nstep P \<sigma>1 (Suc n) \<and> combine coll1 coll2 = coll'"
+ proof -
+ fix \<sigma>' coll'
+ assume "(\<sigma>', coll') \<in> csmall_nstep P \<sigma> (Suc(Suc n))"
+ then obtain \<sigma>1 coll1 coll2 where Sucnstep: "(\<sigma>1, coll1) \<in> csmall_nstep P \<sigma> (Suc n)"
+ "(\<sigma>', coll2) \<in> csmall P \<sigma>1" "combine coll1 coll2 = coll'" by fastforce
+ obtain \<sigma>12 coll12 coll22 where nstep: "(\<sigma>12, coll12) \<in> csmall P \<sigma>
+ \<and> (\<sigma>1, coll22) \<in> csmall_nstep P \<sigma>12 n \<and> combine coll12 coll22 = coll1"
+ using Suc Sucnstep(1) by fastforce
+ then show "\<exists>\<sigma>1 coll1 coll2. (\<sigma>1, coll1) \<in> csmall P \<sigma>
+ \<and> (\<sigma>', coll2) \<in> csmall_nstep P \<sigma>1 (Suc n) \<and> combine coll1 coll2 = coll'"
+ using combine_assoc[of coll12 coll22 coll2] Sucnstep by fastforce
+ qed
+ have left: "\<And>\<sigma>' coll'. \<exists>\<sigma>1 coll1 coll2. (\<sigma>1, coll1) \<in> csmall P \<sigma>
+ \<and> (\<sigma>', coll2) \<in> csmall_nstep P \<sigma>1 (Suc n) \<and> combine coll1 coll2 = coll'
+ \<Longrightarrow> (\<sigma>', coll') \<in> csmall_nstep P \<sigma> (Suc(Suc n))"
+ proof -
+ fix \<sigma>' coll'
+ assume "\<exists>\<sigma>1 coll1 coll2. (\<sigma>1, coll1) \<in> csmall P \<sigma>
+ \<and> (\<sigma>', coll2) \<in> csmall_nstep P \<sigma>1 (Suc n) \<and> combine coll1 coll2 = coll'"
+ then obtain \<sigma>1 coll1 coll2 where Sucnstep: "(\<sigma>1, coll1) \<in> csmall P \<sigma>"
+ "(\<sigma>', coll2) \<in> csmall_nstep P \<sigma>1 (Suc n)" "combine coll1 coll2 = coll'"
+ by fastforce
+ obtain \<sigma>12 coll12 coll22 where nstep: "(\<sigma>12, coll12) \<in> csmall_nstep P \<sigma>1 n
+ \<and> (\<sigma>', coll22) \<in> csmall P \<sigma>12 \<and> combine coll12 coll22 = coll2"
+ using Sucnstep(2) by auto
+ then show "(\<sigma>', coll') \<in> csmall_nstep P \<sigma> (Suc(Suc n))"
+ using combine_assoc[of coll1 coll12 coll22] Suc Sucnstep by fastforce
+ qed
+ show ?case using right left by fast
+qed(simp)
+
+lemma csmall_nstep_SucD:
+assumes "(\<sigma>',coll') \<in> csmall_nstep P \<sigma> (Suc n)"
+shows "\<exists>\<sigma>1 coll1. (\<sigma>1, coll1) \<in> csmall P \<sigma>
+ \<and> (\<exists>coll. coll' = combine coll1 coll \<and> (\<sigma>',coll) \<in> csmall_nstep P \<sigma>1 n)"
+ using csmall_nstep_Rec2 CollectionSemantics_axioms case_prodD assms by fastforce
+
+lemma csmall_nstep_Suc_nend: "o' \<in> csmall_nstep P \<sigma> (Suc n1) \<Longrightarrow> \<sigma> \<notin> endset"
+ using endset_final csmall_nstep_SucD CollectionSemantics.csmall_def CollectionSemantics_axioms
+ by fastforce
+
+lemma small_to_csmall_nstep_pres:
+assumes Qpres: "\<And>P \<sigma> \<sigma>'. Q P \<sigma> \<Longrightarrow> \<sigma>' \<in> small P \<sigma> \<Longrightarrow> Q P \<sigma>'"
+shows "Q P \<sigma> \<Longrightarrow> (\<sigma>', coll) \<in> csmall_nstep P \<sigma> n \<Longrightarrow> Q P \<sigma>'"
+proof(induct n arbitrary: \<sigma> \<sigma>' coll)
+ case (Suc n)
+ then obtain \<sigma>1 coll1 coll2 where nstep: "(\<sigma>1, coll1) \<in> csmall_nstep P \<sigma> n
+ \<and> (\<sigma>', coll2) \<in> csmall P \<sigma>1 \<and> combine coll1 coll2 = coll" by clarsimp
+ then show ?case using Suc Qpres[where P=P and \<sigma>=\<sigma>1 and \<sigma>'=\<sigma>'] by(auto simp: csmall_def)
+qed(simp)
+
+lemma csmall_to_csmall_nstep_prop:
+assumes cond: "\<And>P \<sigma> \<sigma>' coll. (\<sigma>', coll) \<in> csmall P \<sigma> \<Longrightarrow> R P coll \<Longrightarrow> Q P \<sigma> \<Longrightarrow> R' P \<sigma> \<sigma>' coll"
+ and Rcomb: "\<And>P coll1 coll2. R P (combine coll1 coll2) = (R P coll1 \<and> R P coll2)"
+ and Qpres: "\<And>P \<sigma> \<sigma>'. Q P \<sigma> \<Longrightarrow> \<sigma>' \<in> small P \<sigma> \<Longrightarrow> Q P \<sigma>'"
+ and Rtrans': "\<And>P \<sigma> \<sigma>1 \<sigma>' coll1 coll2.
+ R' P \<sigma> \<sigma>1 coll1 \<and> R' P \<sigma>1 \<sigma>' coll2 \<Longrightarrow> R' P \<sigma> \<sigma>' (combine coll1 coll2)"
+ and base: "\<And>\<sigma>. R' P \<sigma> \<sigma> collect_id"
+shows "(\<sigma>', coll) \<in> csmall_nstep P \<sigma> n \<Longrightarrow> R P coll \<Longrightarrow> Q P \<sigma> \<Longrightarrow> R' P \<sigma> \<sigma>' coll"
+proof(induct n arbitrary: \<sigma> \<sigma>' coll)
+ case (Suc n)
+ then obtain \<sigma>1 coll1 coll2 where nstep: "(\<sigma>1, coll1) \<in> csmall_nstep P \<sigma> n
+ \<and> (\<sigma>', coll2) \<in> csmall P \<sigma>1 \<and> combine coll1 coll2 = coll" by clarsimp
+ then have "Q P \<sigma>1" using small_to_csmall_nstep_pres[where Q=Q] Qpres Suc by blast
+ then show ?case using nstep assms Suc by auto blast
+qed(simp add: base)
+
+lemma csmall_to_csmall_nstep_prop2:
+assumes cond: "\<And>P P' \<sigma> \<sigma>' coll. (\<sigma>', coll) \<in> csmall P \<sigma>
+ \<Longrightarrow> R P P' coll \<Longrightarrow> Q \<sigma> \<Longrightarrow> (\<sigma>', coll) \<in> csmall P' \<sigma>"
+ and Rcomb: "\<And>P P' coll1 coll2. R P P' (combine coll1 coll2) = (R P P' coll1 \<and> R P P' coll2)"
+ and Qpres: "\<And>P \<sigma> \<sigma>'. Q \<sigma> \<Longrightarrow> \<sigma>' \<in> small P \<sigma> \<Longrightarrow> Q \<sigma>'"
+shows "(\<sigma>', coll) \<in> csmall_nstep P \<sigma> n \<Longrightarrow> R P P' coll \<Longrightarrow> Q \<sigma> \<Longrightarrow> (\<sigma>', coll) \<in> csmall_nstep P' \<sigma> n"
+proof(induct n arbitrary: \<sigma> \<sigma>' coll)
+ case (Suc n)
+ then obtain \<sigma>1 coll1 coll2 where nstep: "(\<sigma>1, coll1) \<in> csmall_nstep P \<sigma> n
+ \<and> (\<sigma>', coll2) \<in> csmall P \<sigma>1 \<and> combine coll1 coll2 = coll" by clarsimp
+ then have "Q \<sigma>1" using small_to_csmall_nstep_pres[where Q="\<lambda>P. Q"] Qpres Suc by blast
+ then show ?case using nstep assms Suc by auto blast
+qed(simp)
+
+subsection "Extending @{term csmall} to a big-step semantics"
+
+definition cbig :: "'prog \<Rightarrow> 'state \<Rightarrow> ('state \<times> 'coll) set" where
+"cbig P \<sigma> \<equiv>
+ { (\<sigma>', coll). \<exists>n. (\<sigma>', coll) \<in> csmall_nstep P \<sigma> n \<and> \<sigma>' \<in> endset }"
+
+lemma cbigD:
+ "\<lbrakk> (\<sigma>',coll') \<in> cbig P \<sigma> \<rbrakk> \<Longrightarrow> \<exists>n. (\<sigma>',coll') \<in> csmall_nstep P \<sigma> n \<and> \<sigma>' \<in> endset"
+ by(simp add: cbig_def)
+
+lemma cbigD':
+ "\<lbrakk> o' \<in> cbig P \<sigma> \<rbrakk> \<Longrightarrow> \<exists>n. o' \<in> csmall_nstep P \<sigma> n \<and> fst o' \<in> endset"
+ by(cases o', simp add: cbig_def)
+
+lemma cbig_def2:
+ "(\<sigma>', coll) \<in> cbig P \<sigma> \<longleftrightarrow> (\<exists>n. (\<sigma>', coll) \<in> csmall_nstep P \<sigma> n \<and> \<sigma>' \<in> endset)"
+proof(rule iffI)
+ assume "(\<sigma>', coll) \<in> cbig P \<sigma>"
+ then show "\<exists>n. (\<sigma>', coll) \<in> csmall_nstep P \<sigma> n \<and> \<sigma>' \<in> endset" using bigD cbig_def by auto
+next
+ assume "\<exists>n. (\<sigma>', coll) \<in> csmall_nstep P \<sigma> n \<and> \<sigma>' \<in> endset"
+ then show "(\<sigma>', coll) \<in> cbig P \<sigma>" using big_def cbig_def small_nstep_csmall_nstep_equiv by auto
+qed
+
+lemma cbig_big_equiv:
+ "(\<exists>coll. (\<sigma>', coll) \<in> cbig P \<sigma>) \<longleftrightarrow> \<sigma>' \<in> big P \<sigma>"
+proof(rule iffI)
+ assume "\<exists>coll. (\<sigma>', coll) \<in> cbig P \<sigma>"
+ then show "\<sigma>' \<in> big P \<sigma>" by (auto simp: big_def cbig_def small_nstep_csmall_nstep_equiv)
+next
+ assume "\<sigma>' \<in> big P \<sigma>"
+ then show "\<exists>coll. (\<sigma>', coll) \<in> cbig P \<sigma>" by (fastforce simp: cbig_def dest: csmall_nstep_exists)
+qed
+
+lemma cbig_big_implies:
+ "(\<sigma>', coll) \<in> cbig P \<sigma> \<Longrightarrow> \<sigma>' \<in> big P \<sigma>"
+using cbig_big_equiv by blast
+
+lemma csmall_to_cbig_prop:
+assumes "\<And>P \<sigma> \<sigma>' coll. (\<sigma>', coll) \<in> csmall P \<sigma> \<Longrightarrow> R P coll \<Longrightarrow> Q P \<sigma> \<Longrightarrow> R' P \<sigma> \<sigma>' coll"
+ and "\<And>P coll1 coll2. R P (combine coll1 coll2) = (R P coll1 \<and> R P coll2)"
+ and "\<And>P \<sigma> \<sigma>'. Q P \<sigma> \<Longrightarrow> \<sigma>' \<in> small P \<sigma> \<Longrightarrow> Q P \<sigma>'"
+ and "\<And>P \<sigma> \<sigma>1 \<sigma>' coll1 coll2.
+ R' P \<sigma> \<sigma>1 coll1 \<and> R' P \<sigma>1 \<sigma>' coll2 \<Longrightarrow> R' P \<sigma> \<sigma>' (combine coll1 coll2)"
+ and "\<And>\<sigma>. R' P \<sigma> \<sigma> collect_id"
+shows "(\<sigma>', coll) \<in> cbig P \<sigma> \<Longrightarrow> R P coll \<Longrightarrow> Q P \<sigma> \<Longrightarrow> R' P \<sigma> \<sigma>' coll"
+using assms csmall_to_csmall_nstep_prop[where R=R and Q=Q and R'=R' and \<sigma>=\<sigma>]
+ by(auto simp: cbig_def2)
+
+lemma csmall_to_cbig_prop2:
+assumes "\<And>P P' \<sigma> \<sigma>' coll. (\<sigma>', coll) \<in> csmall P \<sigma> \<Longrightarrow> R P P' coll \<Longrightarrow> Q \<sigma> \<Longrightarrow> (\<sigma>', coll) \<in> csmall P' \<sigma>"
+ and "\<And>P P' coll1 coll2. R P P' (combine coll1 coll2) = (R P P' coll1 \<and> R P P' coll2)"
+ and Qpres: "\<And>P \<sigma> \<sigma>'. Q \<sigma> \<Longrightarrow> \<sigma>' \<in> small P \<sigma> \<Longrightarrow> Q \<sigma>'"
+shows "(\<sigma>', coll) \<in> cbig P \<sigma> \<Longrightarrow> R P P' coll \<Longrightarrow> Q \<sigma> \<Longrightarrow> (\<sigma>', coll) \<in> cbig P' \<sigma>"
+using assms csmall_to_csmall_nstep_prop2[where R=R and Q=Q] by(auto simp: cbig_def2) blast
+
+lemma cbig_stepD:
+assumes cbig: "(\<sigma>',coll') \<in> cbig P \<sigma>" and nend: "\<sigma> \<notin> endset"
+shows "\<exists>\<sigma>1 coll1. (\<sigma>1, coll1) \<in> csmall P \<sigma>
+ \<and> (\<exists>coll. coll' = combine coll1 coll \<and> (\<sigma>',coll) \<in> cbig P \<sigma>1)"
+proof -
+ obtain n where n: "(\<sigma>', coll') \<in> csmall_nstep P \<sigma> n" "\<sigma>' \<in> endset"
+ using cbig_def2 cbig by auto
+ then show ?thesis using csmall_nstep_SucD nend cbig_def2 by(cases n, simp) blast
+qed
+
+(****)
+
+lemma csmall_nstep_det_last_eq:
+assumes det: "\<forall>\<sigma>. small P \<sigma> = {} \<or> (\<exists>\<sigma>'. small P \<sigma> = {\<sigma>'})"
+shows "\<lbrakk> (\<sigma>',coll') \<in> cbig P \<sigma>; (\<sigma>',coll') \<in> csmall_nstep P \<sigma> n; (\<sigma>',coll'') \<in> csmall_nstep P \<sigma> n' \<rbrakk>
+ \<Longrightarrow> n = n'"
+proof(induct n arbitrary: n' \<sigma> \<sigma>' coll' coll'')
+ case 0
+ have "\<sigma>' = \<sigma>" using "0.prems"(2) csmall_nstep_base by blast
+ then have endset: "\<sigma> \<in> endset" using "0.prems"(1) cbigD by blast
+ show ?case
+ proof(cases n')
+ case Suc then show ?thesis using "0.prems"(3) csmall_nstep_Suc_nend endset by blast
+ qed(simp)
+next
+ case (Suc n1)
+ then have endset: "\<sigma>' \<in> endset" using Suc.prems(1) cbigD by blast
+ have nend: "\<sigma> \<notin> endset" using csmall_nstep_Suc_nend[OF Suc.prems(2)] by simp
+ then have neq: "\<sigma>' \<noteq> \<sigma>" using endset by auto
+ obtain \<sigma>1 coll coll1 where \<sigma>1: "(\<sigma>1,coll1) \<in> csmall P \<sigma>" "coll' = combine coll1 coll"
+ "(\<sigma>',coll) \<in> csmall_nstep P \<sigma>1 n1"
+ using csmall_nstep_SucD[OF Suc.prems(2)] by clarsimp
+ then have cbig: "(\<sigma>',coll) \<in> cbig P \<sigma>1" using endset by(auto simp: cbig_def)
+ show ?case
+ proof(cases n')
+ case 0 then show ?thesis using neq Suc.prems(3) using csmall_nstep_base by simp
+ next
+ case Suc': (Suc n1')
+ then obtain \<sigma>1' coll2 coll1' where \<sigma>1': "(\<sigma>1',coll1') \<in> csmall P \<sigma>" "coll'' = combine coll1' coll2"
+ "(\<sigma>',coll2) \<in> csmall_nstep P \<sigma>1' n1'"
+ using csmall_nstep_SucD[where \<sigma>=\<sigma> and \<sigma>'=\<sigma>' and coll'=coll'' and n=n1'] Suc.prems(3) by blast
+ then have "\<sigma>1=\<sigma>1'" using \<sigma>1 det csmall_def by auto
+ then show ?thesis using Suc.hyps(1)[OF cbig \<sigma>1(3)] \<sigma>1'(3) Suc' by blast
+ qed
+qed
+
+end \<comment> \<open> CollectionSemantics \<close>
+
end
\ No newline at end of file
diff --git a/thys/Regression_Test_Selection/Common/RTS_safe.thy b/thys/Regression_Test_Selection/Common/RTS_safe.thy
--- a/thys/Regression_Test_Selection/Common/RTS_safe.thy
+++ b/thys/Regression_Test_Selection/Common/RTS_safe.thy
@@ -1,86 +1,86 @@
-(* Title: RTS/Common/RTS_safe.thy *)
-(* Author: Susannah Mansky, UIUC 2020 *)
-
-section "Regression Test Selection algorithm model"
-
-theory RTS_safe
-imports Main
-begin
-
-text \<open> This describes an \emph{existence safe} RTS algorithm: if a test
- is deselected based on an output, there is SOME equivalent output
- under the changed program. \<close>
-locale RTS_safe =
- fixes
- out :: "'prog \<Rightarrow> 'test \<Rightarrow> 'prog_out set" and
- equiv_out :: "'prog_out \<Rightarrow> 'prog_out \<Rightarrow> bool" and
- deselect :: "'prog \<Rightarrow> 'prog_out \<Rightarrow> 'prog \<Rightarrow> bool" and
- progs :: "'prog set" and
- tests :: "'test set"
- assumes
- existence_safe: "\<lbrakk> P \<in> progs; P' \<in> progs; t \<in> tests; o1 \<in> out P t; deselect P o1 P' \<rbrakk>
- \<Longrightarrow> (\<exists>o2 \<in> out P' t. equiv_out o1 o2)" and
- equiv_out_equiv: "equiv UNIV {(x,y). equiv_out x y}" and
- equiv_out_deselect: "\<lbrakk> equiv_out o1 o2; deselect P o1 P' \<rbrakk> \<Longrightarrow> deselect P o2 P'"
-
-context RTS_safe begin
-
-lemma equiv_out_refl: "equiv_out a a"
-using equiv_class_eq_iff equiv_out_equiv by fastforce
-
-lemma equiv_out_trans: "\<lbrakk> equiv_out a b; equiv_out b c \<rbrakk> \<Longrightarrow> equiv_out a c"
-using equiv_class_eq_iff equiv_out_equiv by fastforce
-
-text "This shows that it is safe to continue deselecting a test based
- on its output under a previous program, to an arbitrary number of
- program changes, as long as the test is continually deselected. This
- is useful because it means changed programs don't need to generate new
- outputs for deselected tests to ensure safety of future deselections."
-lemma existence_safe_trans:
-assumes Pst_in: "Ps \<noteq> []" "set Ps \<subseteq> progs" "t \<in> tests" and
- o0: "o\<^sub>0 \<in> out (Ps!0) t" and
- des: "\<forall>n < (length Ps) - 1. deselect (Ps!n) o\<^sub>0 (Ps!(Suc n))"
-shows "\<exists>o\<^sub>n \<in> out (last Ps) t. equiv_out o\<^sub>0 o\<^sub>n"
-using assms proof(induct "length Ps" arbitrary: Ps)
- case 0 with Pst_in show ?case by simp
-next
- case (Suc x) then show ?case
- proof(induct x)
- case z: 0
- from z.prems(2,3) have "Ps ! (length Ps - 2) = last Ps"
- by (simp add: last_conv_nth numeral_2_eq_2)
- with equiv_out_refl z.prems(2,6) show ?case by auto
- next
- case Suc':(Suc x')
- let ?Ps = "take (Suc x') Ps"
- have len': "Suc x' = length (take (Suc x') Ps)" using Suc'.prems(2) by auto
- moreover have nmt': "take (Suc x') Ps \<noteq> []" using len' by auto
- moreover have sub': "set (take (Suc x') Ps) \<subseteq> progs" using Suc.prems(2)
- by (meson order_trans set_take_subset)
- moreover have "t \<in> tests" using Pst_in(3) by simp
- moreover have "o\<^sub>0 \<in> out (take (Suc x') Ps ! 0) t" using Suc.prems(4) by simp
- moreover have "\<forall>n<length (take (Suc x') Ps) - 1.
- deselect (take (Suc x') Ps ! n) o\<^sub>0 (take (Suc x') Ps ! (Suc n))"
- using Suc.prems(5) len' by simp
- ultimately have "\<exists>o'\<in>out (last ?Ps) t. equiv_out o\<^sub>0 o'" by(rule Suc'.prems(1)[of ?Ps])
- then obtain o' where o': "o' \<in> out (last ?Ps) t" and eo: "equiv_out o\<^sub>0 o'" by clarify
- from Suc.prems(1) Suc'.prems(2) len' nmt'
- have "last (take (Suc x') Ps) = Ps!x'" "last Ps = Ps!(Suc x')"
- by (metis diff_Suc_1 last_conv_nth lessI nth_take)+
- moreover have "x' < length Ps - 1" using Suc'.prems(2) by linarith
- ultimately have des':"deselect (last (take (Suc x') Ps)) o\<^sub>0 (last Ps)"
- using Suc.prems(5) by simp
- from Suc.prems(1,2) sub' nmt' last_in_set
- have Ps_in: "last (take (Suc x') Ps) \<in> progs" "last Ps \<in> progs" by blast+
- have "\<exists>o\<^sub>n \<in> out (last Ps) t. equiv_out o' o\<^sub>n"
- by(rule existence_safe[where P="last (take (Suc x') Ps)" and P'="last Ps" and t=t,
- OF Ps_in Pst_in(3) o' equiv_out_deselect[OF eo des']])
- then obtain o\<^sub>n where oN: "o\<^sub>n \<in> out (last Ps) t" and eo': "equiv_out o' o\<^sub>n" by clarify
- moreover from eo eo' have "equiv_out o\<^sub>0 o\<^sub>n" by(rule equiv_out_trans)
- ultimately show ?case by auto
- qed
-qed
-
-end \<comment> \<open> @{text RTS_safe} \<close>
-
+(* Title: RTS/Common/RTS_safe.thy *)
+(* Author: Susannah Mansky, UIUC 2020 *)
+
+section "Regression Test Selection algorithm model"
+
+theory RTS_safe
+imports Main
+begin
+
+text \<open> This describes an \emph{existence safe} RTS algorithm: if a test
+ is deselected based on an output, there is SOME equivalent output
+ under the changed program. \<close>
+locale RTS_safe =
+ fixes
+ out :: "'prog \<Rightarrow> 'test \<Rightarrow> 'prog_out set" and
+ equiv_out :: "'prog_out \<Rightarrow> 'prog_out \<Rightarrow> bool" and
+ deselect :: "'prog \<Rightarrow> 'prog_out \<Rightarrow> 'prog \<Rightarrow> bool" and
+ progs :: "'prog set" and
+ tests :: "'test set"
+ assumes
+ existence_safe: "\<lbrakk> P \<in> progs; P' \<in> progs; t \<in> tests; o1 \<in> out P t; deselect P o1 P' \<rbrakk>
+ \<Longrightarrow> (\<exists>o2 \<in> out P' t. equiv_out o1 o2)" and
+ equiv_out_equiv: "equiv UNIV {(x,y). equiv_out x y}" and
+ equiv_out_deselect: "\<lbrakk> equiv_out o1 o2; deselect P o1 P' \<rbrakk> \<Longrightarrow> deselect P o2 P'"
+
+context RTS_safe begin
+
+lemma equiv_out_refl: "equiv_out a a"
+using equiv_class_eq_iff equiv_out_equiv by fastforce
+
+lemma equiv_out_trans: "\<lbrakk> equiv_out a b; equiv_out b c \<rbrakk> \<Longrightarrow> equiv_out a c"
+using equiv_class_eq_iff equiv_out_equiv by fastforce
+
+text "This shows that it is safe to continue deselecting a test based
+ on its output under a previous program, to an arbitrary number of
+ program changes, as long as the test is continually deselected. This
+ is useful because it means changed programs don't need to generate new
+ outputs for deselected tests to ensure safety of future deselections."
+lemma existence_safe_trans:
+assumes Pst_in: "Ps \<noteq> []" "set Ps \<subseteq> progs" "t \<in> tests" and
+ o0: "o\<^sub>0 \<in> out (Ps!0) t" and
+ des: "\<forall>n < (length Ps) - 1. deselect (Ps!n) o\<^sub>0 (Ps!(Suc n))"
+shows "\<exists>o\<^sub>n \<in> out (last Ps) t. equiv_out o\<^sub>0 o\<^sub>n"
+using assms proof(induct "length Ps" arbitrary: Ps)
+ case 0 with Pst_in show ?case by simp
+next
+ case (Suc x) then show ?case
+ proof(induct x)
+ case z: 0
+ from z.prems(2,3) have "Ps ! (length Ps - 2) = last Ps"
+ by (simp add: last_conv_nth numeral_2_eq_2)
+ with equiv_out_refl z.prems(2,6) show ?case by auto
+ next
+ case Suc':(Suc x')
+ let ?Ps = "take (Suc x') Ps"
+ have len': "Suc x' = length (take (Suc x') Ps)" using Suc'.prems(2) by auto
+ moreover have nmt': "take (Suc x') Ps \<noteq> []" using len' by auto
+ moreover have sub': "set (take (Suc x') Ps) \<subseteq> progs" using Suc.prems(2)
+ by (meson order_trans set_take_subset)
+ moreover have "t \<in> tests" using Pst_in(3) by simp
+ moreover have "o\<^sub>0 \<in> out (take (Suc x') Ps ! 0) t" using Suc.prems(4) by simp
+ moreover have "\<forall>n<length (take (Suc x') Ps) - 1.
+ deselect (take (Suc x') Ps ! n) o\<^sub>0 (take (Suc x') Ps ! (Suc n))"
+ using Suc.prems(5) len' by simp
+ ultimately have "\<exists>o'\<in>out (last ?Ps) t. equiv_out o\<^sub>0 o'" by(rule Suc'.prems(1)[of ?Ps])
+ then obtain o' where o': "o' \<in> out (last ?Ps) t" and eo: "equiv_out o\<^sub>0 o'" by clarify
+ from Suc.prems(1) Suc'.prems(2) len' nmt'
+ have "last (take (Suc x') Ps) = Ps!x'" "last Ps = Ps!(Suc x')"
+ by (metis diff_Suc_1 last_conv_nth lessI nth_take)+
+ moreover have "x' < length Ps - 1" using Suc'.prems(2) by linarith
+ ultimately have des':"deselect (last (take (Suc x') Ps)) o\<^sub>0 (last Ps)"
+ using Suc.prems(5) by simp
+ from Suc.prems(1,2) sub' nmt' last_in_set
+ have Ps_in: "last (take (Suc x') Ps) \<in> progs" "last Ps \<in> progs" by blast+
+ have "\<exists>o\<^sub>n \<in> out (last Ps) t. equiv_out o' o\<^sub>n"
+ by(rule existence_safe[where P="last (take (Suc x') Ps)" and P'="last Ps" and t=t,
+ OF Ps_in Pst_in(3) o' equiv_out_deselect[OF eo des']])
+ then obtain o\<^sub>n where oN: "o\<^sub>n \<in> out (last Ps) t" and eo': "equiv_out o' o\<^sub>n" by clarify
+ moreover from eo eo' have "equiv_out o\<^sub>0 o\<^sub>n" by(rule equiv_out_trans)
+ ultimately show ?case by auto
+ qed
+qed
+
+end \<comment> \<open> @{text RTS_safe} \<close>
+
end
\ No newline at end of file
diff --git a/thys/Regression_Test_Selection/Common/Semantics.thy b/thys/Regression_Test_Selection/Common/Semantics.thy
--- a/thys/Regression_Test_Selection/Common/Semantics.thy
+++ b/thys/Regression_Test_Selection/Common/Semantics.thy
@@ -1,135 +1,135 @@
-(* Title: RTS/Common/Semantics.thy *)
-(* Author: Susannah Mansky, UIUC 2020 *)
-
-section "Semantics model"
-
-theory Semantics
-imports Main
-begin
-
-text "General model for small-step semantics:"
-locale Semantics =
- fixes
- small :: "'prog \<Rightarrow> 'state \<Rightarrow> 'state set" and
- endset :: "'state set"
- assumes
- endset_final: "\<sigma> \<in> endset \<Longrightarrow> \<forall>P. small P \<sigma> = {}"
-
-context Semantics begin
-
-subsection "Extending @{term small} to multiple steps"
-
-primrec small_nstep :: "'prog \<Rightarrow> 'state \<Rightarrow> nat \<Rightarrow> 'state set" where
-small_nstep_base:
- "small_nstep P \<sigma> 0 = {\<sigma>}" |
-small_nstep_Rec:
- "small_nstep P \<sigma> (Suc n) =
- { \<sigma>2. \<exists>\<sigma>1. \<sigma>1 \<in> small_nstep P \<sigma> n \<and> \<sigma>2 \<in> small P \<sigma>1 }"
-
-lemma small_nstep_Rec2:
- "small_nstep P \<sigma> (Suc n) =
- { \<sigma>2. \<exists>\<sigma>1. \<sigma>1 \<in> small P \<sigma> \<and> \<sigma>2 \<in> small_nstep P \<sigma>1 n }"
-proof(induct n arbitrary: \<sigma>)
- case (Suc n)
- have right: "\<And>\<sigma>'. \<sigma>' \<in> small_nstep P \<sigma> (Suc(Suc n))
- \<Longrightarrow> \<exists>\<sigma>1. \<sigma>1 \<in> small P \<sigma> \<and> \<sigma>' \<in> small_nstep P \<sigma>1 (Suc n)"
- proof -
- fix \<sigma>'
- assume "\<sigma>' \<in> small_nstep P \<sigma> (Suc(Suc n))"
- then obtain \<sigma>1 where Sucnstep: "\<sigma>1 \<in> small_nstep P \<sigma> (Suc n)" "\<sigma>' \<in> small P \<sigma>1" by fastforce
- obtain \<sigma>12 where nstep: "\<sigma>12 \<in> small P \<sigma> \<and> \<sigma>1 \<in> small_nstep P \<sigma>12 n"
- using Suc Sucnstep(1) by fastforce
- then show "\<exists>\<sigma>1. \<sigma>1 \<in> small P \<sigma> \<and> \<sigma>' \<in> small_nstep P \<sigma>1 (Suc n)"
- using Sucnstep by fastforce
- qed
- have left: "\<And>\<sigma>' . \<exists>\<sigma>1. \<sigma>1 \<in> small P \<sigma> \<and> \<sigma>' \<in> small_nstep P \<sigma>1 (Suc n)
- \<Longrightarrow> \<sigma>' \<in> small_nstep P \<sigma> (Suc(Suc n))"
- proof -
- fix \<sigma>'
- assume "\<exists>\<sigma>1. \<sigma>1 \<in> small P \<sigma> \<and> \<sigma>' \<in> small_nstep P \<sigma>1 (Suc n)"
- then obtain \<sigma>1 where Sucnstep: "\<sigma>1 \<in> small P \<sigma>" "\<sigma>' \<in> small_nstep P \<sigma>1 (Suc n)"
- by fastforce
- obtain \<sigma>12 where nstep: "\<sigma>12 \<in> small_nstep P \<sigma>1 n \<and> \<sigma>' \<in> small P \<sigma>12"
- using Sucnstep(2) by auto
- then show "\<sigma>' \<in> small_nstep P \<sigma> (Suc(Suc n))" using Suc Sucnstep by fastforce
- qed
- show ?case using right left by fast
-qed(simp)
-
-lemma small_nstep_SucD:
-assumes "\<sigma>' \<in> small_nstep P \<sigma> (Suc n)"
-shows "\<exists>\<sigma>1. \<sigma>1 \<in> small P \<sigma> \<and> \<sigma>' \<in> small_nstep P \<sigma>1 n"
- using small_nstep_Rec2 case_prodD assms by fastforce
-
-lemma small_nstep_Suc_nend: "\<sigma>' \<in> small_nstep P \<sigma> (Suc n1) \<Longrightarrow> \<sigma> \<notin> endset"
- using endset_final small_nstep_SucD by fastforce
-
-subsection "Extending @{term small} to a big-step semantics"
-
-definition big :: "'prog \<Rightarrow> 'state \<Rightarrow> 'state set" where
-"big P \<sigma> \<equiv> { \<sigma>'. \<exists>n. \<sigma>' \<in> small_nstep P \<sigma> n \<and> \<sigma>' \<in> endset }"
-
-lemma bigI:
- "\<lbrakk> \<sigma>' \<in> small_nstep P \<sigma> n; \<sigma>' \<in> endset \<rbrakk> \<Longrightarrow> \<sigma>' \<in> big P \<sigma>"
-by (fastforce simp add: big_def)
-
-lemma bigD:
- "\<lbrakk> \<sigma>' \<in> big P \<sigma> \<rbrakk> \<Longrightarrow> \<exists>n. \<sigma>' \<in> small_nstep P \<sigma> n \<and> \<sigma>' \<in> endset"
-by (simp add: big_def)
-
-lemma big_def2:
- "\<sigma>' \<in> big P \<sigma> \<longleftrightarrow> (\<exists>n. \<sigma>' \<in> small_nstep P \<sigma> n \<and> \<sigma>' \<in> endset)"
-proof(rule iffI)
- assume "\<sigma>' \<in> big P \<sigma>"
- then show "\<exists>n. \<sigma>' \<in> small_nstep P \<sigma> n \<and> \<sigma>' \<in> endset" using bigD big_def by auto
-next
- assume "\<exists>n. \<sigma>' \<in> small_nstep P \<sigma> n \<and> \<sigma>' \<in> endset"
- then show "\<sigma>' \<in> big P \<sigma>" using big_def big_def by auto
-qed
-
-lemma big_stepD:
-assumes big: "\<sigma>' \<in> big P \<sigma>" and nend: "\<sigma> \<notin> endset"
-shows "\<exists>\<sigma>1. \<sigma>1 \<in> small P \<sigma> \<and> \<sigma>' \<in> big P \<sigma>1"
-proof -
- obtain n where n: "\<sigma>' \<in> small_nstep P \<sigma> n" "\<sigma>' \<in> endset"
- using big_def2 big by auto
- then show ?thesis using small_nstep_SucD nend big_def2 by(cases n, simp) blast
-qed
-
-(***)
-
-lemma small_nstep_det_last_eq:
-assumes det: "\<forall>\<sigma>. small P \<sigma> = {} \<or> (\<exists>\<sigma>'. small P \<sigma> = {\<sigma>'})"
-shows "\<lbrakk> \<sigma>' \<in> big P \<sigma>; \<sigma>' \<in> small_nstep P \<sigma> n; \<sigma>' \<in> small_nstep P \<sigma> n' \<rbrakk> \<Longrightarrow> n = n'"
-proof(induct n arbitrary: n' \<sigma> \<sigma>')
- case 0
- have "\<sigma>' = \<sigma>" using "0.prems"(2) small_nstep_base by blast
- then have endset: "\<sigma> \<in> endset" using "0.prems"(1) bigD by blast
- show ?case
- proof(cases n')
- case Suc then show ?thesis using "0.prems"(3) small_nstep_SucD endset_final[OF endset] by blast
- qed(simp)
-next
- case (Suc n1)
- then have endset: "\<sigma>' \<in> endset" using Suc.prems(1) bigD by blast
- have nend: "\<sigma> \<notin> endset" using small_nstep_Suc_nend[OF Suc.prems(2)] by simp
- then have neq: "\<sigma>' \<noteq> \<sigma>" using endset by auto
- obtain \<sigma>1 where \<sigma>1: "\<sigma>1 \<in> small P \<sigma>" "\<sigma>' \<in> small_nstep P \<sigma>1 n1"
- using small_nstep_SucD[OF Suc.prems(2)] by clarsimp
- then have big: "\<sigma>' \<in> big P \<sigma>1" using endset by(auto simp: big_def)
- show ?case
- proof(cases n')
- case 0 then show ?thesis using neq Suc.prems(3) using small_nstep_base by blast
- next
- case Suc': (Suc n1')
- then obtain \<sigma>1' where \<sigma>1': "\<sigma>1' \<in> small P \<sigma>" "\<sigma>' \<in> small_nstep P \<sigma>1' n1'"
- using small_nstep_SucD[where \<sigma>=\<sigma> and \<sigma>'=\<sigma>' and n=n1'] Suc.prems(3) by blast
- then have "\<sigma>1=\<sigma>1'" using \<sigma>1(1) det by auto
- then show ?thesis using Suc.hyps(1)[OF big \<sigma>1(2)] \<sigma>1'(2) Suc' by blast
- qed
-qed
-
-end \<comment> \<open> Semantics \<close>
-
-
+(* Title: RTS/Common/Semantics.thy *)
+(* Author: Susannah Mansky, UIUC 2020 *)
+
+section "Semantics model"
+
+theory Semantics
+imports Main
+begin
+
+text "General model for small-step semantics:"
+locale Semantics =
+ fixes
+ small :: "'prog \<Rightarrow> 'state \<Rightarrow> 'state set" and
+ endset :: "'state set"
+ assumes
+ endset_final: "\<sigma> \<in> endset \<Longrightarrow> \<forall>P. small P \<sigma> = {}"
+
+context Semantics begin
+
+subsection "Extending @{term small} to multiple steps"
+
+primrec small_nstep :: "'prog \<Rightarrow> 'state \<Rightarrow> nat \<Rightarrow> 'state set" where
+small_nstep_base:
+ "small_nstep P \<sigma> 0 = {\<sigma>}" |
+small_nstep_Rec:
+ "small_nstep P \<sigma> (Suc n) =
+ { \<sigma>2. \<exists>\<sigma>1. \<sigma>1 \<in> small_nstep P \<sigma> n \<and> \<sigma>2 \<in> small P \<sigma>1 }"
+
+lemma small_nstep_Rec2:
+ "small_nstep P \<sigma> (Suc n) =
+ { \<sigma>2. \<exists>\<sigma>1. \<sigma>1 \<in> small P \<sigma> \<and> \<sigma>2 \<in> small_nstep P \<sigma>1 n }"
+proof(induct n arbitrary: \<sigma>)
+ case (Suc n)
+ have right: "\<And>\<sigma>'. \<sigma>' \<in> small_nstep P \<sigma> (Suc(Suc n))
+ \<Longrightarrow> \<exists>\<sigma>1. \<sigma>1 \<in> small P \<sigma> \<and> \<sigma>' \<in> small_nstep P \<sigma>1 (Suc n)"
+ proof -
+ fix \<sigma>'
+ assume "\<sigma>' \<in> small_nstep P \<sigma> (Suc(Suc n))"
+ then obtain \<sigma>1 where Sucnstep: "\<sigma>1 \<in> small_nstep P \<sigma> (Suc n)" "\<sigma>' \<in> small P \<sigma>1" by fastforce
+ obtain \<sigma>12 where nstep: "\<sigma>12 \<in> small P \<sigma> \<and> \<sigma>1 \<in> small_nstep P \<sigma>12 n"
+ using Suc Sucnstep(1) by fastforce
+ then show "\<exists>\<sigma>1. \<sigma>1 \<in> small P \<sigma> \<and> \<sigma>' \<in> small_nstep P \<sigma>1 (Suc n)"
+ using Sucnstep by fastforce
+ qed
+ have left: "\<And>\<sigma>' . \<exists>\<sigma>1. \<sigma>1 \<in> small P \<sigma> \<and> \<sigma>' \<in> small_nstep P \<sigma>1 (Suc n)
+ \<Longrightarrow> \<sigma>' \<in> small_nstep P \<sigma> (Suc(Suc n))"
+ proof -
+ fix \<sigma>'
+ assume "\<exists>\<sigma>1. \<sigma>1 \<in> small P \<sigma> \<and> \<sigma>' \<in> small_nstep P \<sigma>1 (Suc n)"
+ then obtain \<sigma>1 where Sucnstep: "\<sigma>1 \<in> small P \<sigma>" "\<sigma>' \<in> small_nstep P \<sigma>1 (Suc n)"
+ by fastforce
+ obtain \<sigma>12 where nstep: "\<sigma>12 \<in> small_nstep P \<sigma>1 n \<and> \<sigma>' \<in> small P \<sigma>12"
+ using Sucnstep(2) by auto
+ then show "\<sigma>' \<in> small_nstep P \<sigma> (Suc(Suc n))" using Suc Sucnstep by fastforce
+ qed
+ show ?case using right left by fast
+qed(simp)
+
+lemma small_nstep_SucD:
+assumes "\<sigma>' \<in> small_nstep P \<sigma> (Suc n)"
+shows "\<exists>\<sigma>1. \<sigma>1 \<in> small P \<sigma> \<and> \<sigma>' \<in> small_nstep P \<sigma>1 n"
+ using small_nstep_Rec2 case_prodD assms by fastforce
+
+lemma small_nstep_Suc_nend: "\<sigma>' \<in> small_nstep P \<sigma> (Suc n1) \<Longrightarrow> \<sigma> \<notin> endset"
+ using endset_final small_nstep_SucD by fastforce
+
+subsection "Extending @{term small} to a big-step semantics"
+
+definition big :: "'prog \<Rightarrow> 'state \<Rightarrow> 'state set" where
+"big P \<sigma> \<equiv> { \<sigma>'. \<exists>n. \<sigma>' \<in> small_nstep P \<sigma> n \<and> \<sigma>' \<in> endset }"
+
+lemma bigI:
+ "\<lbrakk> \<sigma>' \<in> small_nstep P \<sigma> n; \<sigma>' \<in> endset \<rbrakk> \<Longrightarrow> \<sigma>' \<in> big P \<sigma>"
+by (fastforce simp add: big_def)
+
+lemma bigD:
+ "\<lbrakk> \<sigma>' \<in> big P \<sigma> \<rbrakk> \<Longrightarrow> \<exists>n. \<sigma>' \<in> small_nstep P \<sigma> n \<and> \<sigma>' \<in> endset"
+by (simp add: big_def)
+
+lemma big_def2:
+ "\<sigma>' \<in> big P \<sigma> \<longleftrightarrow> (\<exists>n. \<sigma>' \<in> small_nstep P \<sigma> n \<and> \<sigma>' \<in> endset)"
+proof(rule iffI)
+ assume "\<sigma>' \<in> big P \<sigma>"
+ then show "\<exists>n. \<sigma>' \<in> small_nstep P \<sigma> n \<and> \<sigma>' \<in> endset" using bigD big_def by auto
+next
+ assume "\<exists>n. \<sigma>' \<in> small_nstep P \<sigma> n \<and> \<sigma>' \<in> endset"
+ then show "\<sigma>' \<in> big P \<sigma>" using big_def big_def by auto
+qed
+
+lemma big_stepD:
+assumes big: "\<sigma>' \<in> big P \<sigma>" and nend: "\<sigma> \<notin> endset"
+shows "\<exists>\<sigma>1. \<sigma>1 \<in> small P \<sigma> \<and> \<sigma>' \<in> big P \<sigma>1"
+proof -
+ obtain n where n: "\<sigma>' \<in> small_nstep P \<sigma> n" "\<sigma>' \<in> endset"
+ using big_def2 big by auto
+ then show ?thesis using small_nstep_SucD nend big_def2 by(cases n, simp) blast
+qed
+
+(***)
+
+lemma small_nstep_det_last_eq:
+assumes det: "\<forall>\<sigma>. small P \<sigma> = {} \<or> (\<exists>\<sigma>'. small P \<sigma> = {\<sigma>'})"
+shows "\<lbrakk> \<sigma>' \<in> big P \<sigma>; \<sigma>' \<in> small_nstep P \<sigma> n; \<sigma>' \<in> small_nstep P \<sigma> n' \<rbrakk> \<Longrightarrow> n = n'"
+proof(induct n arbitrary: n' \<sigma> \<sigma>')
+ case 0
+ have "\<sigma>' = \<sigma>" using "0.prems"(2) small_nstep_base by blast
+ then have endset: "\<sigma> \<in> endset" using "0.prems"(1) bigD by blast
+ show ?case
+ proof(cases n')
+ case Suc then show ?thesis using "0.prems"(3) small_nstep_SucD endset_final[OF endset] by blast
+ qed(simp)
+next
+ case (Suc n1)
+ then have endset: "\<sigma>' \<in> endset" using Suc.prems(1) bigD by blast
+ have nend: "\<sigma> \<notin> endset" using small_nstep_Suc_nend[OF Suc.prems(2)] by simp
+ then have neq: "\<sigma>' \<noteq> \<sigma>" using endset by auto
+ obtain \<sigma>1 where \<sigma>1: "\<sigma>1 \<in> small P \<sigma>" "\<sigma>' \<in> small_nstep P \<sigma>1 n1"
+ using small_nstep_SucD[OF Suc.prems(2)] by clarsimp
+ then have big: "\<sigma>' \<in> big P \<sigma>1" using endset by(auto simp: big_def)
+ show ?case
+ proof(cases n')
+ case 0 then show ?thesis using neq Suc.prems(3) using small_nstep_base by blast
+ next
+ case Suc': (Suc n1')
+ then obtain \<sigma>1' where \<sigma>1': "\<sigma>1' \<in> small P \<sigma>" "\<sigma>' \<in> small_nstep P \<sigma>1' n1'"
+ using small_nstep_SucD[where \<sigma>=\<sigma> and \<sigma>'=\<sigma>' and n=n1'] Suc.prems(3) by blast
+ then have "\<sigma>1=\<sigma>1'" using \<sigma>1(1) det by auto
+ then show ?thesis using Suc.hyps(1)[OF big \<sigma>1(2)] \<sigma>1'(2) Suc' by blast
+ qed
+qed
+
+end \<comment> \<open> Semantics \<close>
+
+
end
\ No newline at end of file
diff --git a/thys/Regression_Test_Selection/JVM_RTS/JVMCollectionBasedRTS.thy b/thys/Regression_Test_Selection/JVM_RTS/JVMCollectionBasedRTS.thy
--- a/thys/Regression_Test_Selection/JVM_RTS/JVMCollectionBasedRTS.thy
+++ b/thys/Regression_Test_Selection/JVM_RTS/JVMCollectionBasedRTS.thy
@@ -1,2342 +1,2342 @@
-(* File: RTS/JVM_RTS/JVMCollectionBasedRTS.thy *)
-(* Author: Susannah Mansky, UIUC 2020 *)
-(* Proof of safety of certain collection based RTS algorithms for Jinja JVM *)
-
-section "Instantiating @{term CollectionBasedRTS} with Jinja JVM"
-
-theory JVMCollectionBasedRTS
-imports "../Common/CollectionBasedRTS" JVMCollectionSemantics
- JinjaDCI.BVSpecTypeSafe "../JinjaSuppl/JVMExecStepInductive"
-
-begin
-
-lemma eq_equiv[simp]: "equiv UNIV {(x, y). x = y}"
-by(simp add: equiv_def refl_on_def sym_def trans_def)
-
-(**********************************************)
-subsection \<open> Some @{text "classes_above"} lemmas \<close>
-(* here because they require ClassAdd/StartProg *)
-
-lemma start_prog_classes_above_Start:
- "classes_above (start_prog P C M) Start = {Object,Start}"
-using start_prog_Start_super[of C M P] subcls1_confluent by auto
-
-lemma class_add_classes_above:
-assumes ns: "\<not> is_class P C" and "\<not>P \<turnstile> D \<preceq>\<^sup>* C"
-shows "classes_above (class_add P (C, cdec)) D = classes_above P D"
-using assms by(auto intro: class_add_subcls class_add_subcls_rev)
-
-lemma class_add_classes_above_xcpts:
-assumes ns: "\<not> is_class P C"
- and ncp: "\<And>D. D \<in> sys_xcpts \<Longrightarrow> \<not>P \<turnstile> D \<preceq>\<^sup>* C"
-shows "classes_above_xcpts (class_add P (C, cdec)) = classes_above_xcpts P"
-using assms class_add_classes_above by simp
-
-(*********)
-subsection "JVM next-step lemmas for initialization calling"
-
-lemma JVM_New_next_step:
-assumes step: "\<sigma>' \<in> JVMsmall P \<sigma>"
- and nend: "\<sigma> \<notin> JVMendset"
- and curr: "curr_instr P (hd(frames_of \<sigma>)) = New C"
- and nDone: "\<not>(\<exists>sfs i. sheap \<sigma> C = Some(sfs,i) \<and> i = Done)"
- and ics: "ics_of(hd(frames_of \<sigma>)) = No_ics"
-shows "ics_of (hd(frames_of \<sigma>')) = Calling C [] \<and> sheap \<sigma> = sheap \<sigma>' \<and> \<sigma>' \<notin> JVMendset"
-proof -
- obtain xp h frs sh where \<sigma>: "\<sigma>=(xp,h,frs,sh)" by(cases \<sigma>)
- then obtain f1 frs1 where frs: "frs=f1#frs1" using nend by(cases frs, simp_all add: JVMendset_def)
- then obtain stk loc C' M' pc ics where f1:"f1=(stk,loc,C',M',pc,ics)" by(cases f1)
- have xp: "xp = None" using \<sigma> nend by(simp add: JVMendset_def)
- obtain xp' h' frs' sh' where \<sigma>': "\<sigma>'=(xp',h',frs',sh')" by(cases \<sigma>')
- have "ics_of (hd frs') = Calling C [] \<and> sh = sh' \<and> frs' \<noteq> [] \<and> xp' = None"
- proof(cases "sh C")
- case None then show ?thesis using \<sigma>' xp f1 frs \<sigma> assms by auto
- next
- case (Some a)
- then obtain sfs i where a: "a=(sfs,i)" by(cases a)
- then have nDone': "i \<noteq> Done" using nDone Some f1 frs \<sigma> by simp
- show ?thesis using a Some \<sigma>' xp f1 frs \<sigma> assms by(auto split: init_state.splits)
- qed
- then show ?thesis using ics \<sigma> \<sigma>' by(cases frs', auto simp: JVMendset_def)
-qed
-
-lemma JVM_Getstatic_next_step:
-assumes step: "\<sigma>' \<in> JVMsmall P \<sigma>"
- and nend: "\<sigma> \<notin> JVMendset"
- and curr: "curr_instr P (hd(frames_of \<sigma>)) = Getstatic C F D"
- and fC: "P \<turnstile> C has F,Static:t in D"
- and nDone: "\<not>(\<exists>sfs i. sheap \<sigma> D = Some(sfs,i) \<and> i = Done)"
- and ics: "ics_of(hd(frames_of \<sigma>)) = No_ics"
-shows "ics_of (hd(frames_of \<sigma>')) = Calling D [] \<and> sheap \<sigma> = sheap \<sigma>' \<and> \<sigma>' \<notin> JVMendset"
-proof -
- obtain xp h frs sh where \<sigma>: "\<sigma>=(xp,h,frs,sh)" by(cases \<sigma>)
- then obtain f1 frs1 where frs: "frs=f1#frs1" using nend by(cases frs, simp_all add: JVMendset_def)
- then obtain stk loc C' M' pc ics where f1:"f1=(stk,loc,C',M',pc,ics)" by(cases f1)
- have xp: "xp = None" using \<sigma> nend by(simp add: JVMendset_def)
- obtain xp' h' frs' sh' where \<sigma>': "\<sigma>'=(xp',h',frs',sh')" by(cases \<sigma>')
- have ex': "\<exists>t b. P \<turnstile> C has F,b:t in D" using fC by auto
- have field: "\<exists>t. field P D F = (D,Static,t)"
- using fC field_def2 has_field_idemp has_field_sees by blast
- have nCalled': "\<forall>Cs. ics \<noteq> Called Cs" using ics f1 frs \<sigma> by simp
- have "ics_of (hd frs') = Calling D [] \<and> sh = sh' \<and> frs' \<noteq> [] \<and> xp' = None"
- proof(cases "sh D")
- case None then show ?thesis using field ex' \<sigma>' xp f1 frs \<sigma> assms by auto
- next
- case (Some a)
- then obtain sfs i where a: "a=(sfs,i)" by(cases a)
- show ?thesis using field ex' a Some \<sigma>' xp f1 frs \<sigma> assms by(auto split: init_state.splits)
- qed
- then show ?thesis using ics \<sigma> \<sigma>' by(auto simp: JVMendset_def)
-qed
-
-lemma JVM_Putstatic_next_step:
-assumes step: "\<sigma>' \<in> JVMsmall P \<sigma>"
- and nend: "\<sigma> \<notin> JVMendset"
- and curr: "curr_instr P (hd(frames_of \<sigma>)) = Putstatic C F D"
- and fC: "P \<turnstile> C has F,Static:t in D"
- and nDone: "\<not>(\<exists>sfs i. sheap \<sigma> D = Some(sfs,i) \<and> i = Done)"
- and ics: "ics_of(hd(frames_of \<sigma>)) = No_ics"
-shows "ics_of (hd(frames_of \<sigma>')) = Calling D [] \<and> sheap \<sigma> = sheap \<sigma>' \<and> \<sigma>' \<notin> JVMendset"
-proof -
- obtain xp h frs sh where \<sigma>: "\<sigma>=(xp,h,frs,sh)" by(cases \<sigma>)
- then obtain f1 frs1 where frs: "frs=f1#frs1" using nend by(cases frs, simp_all add: JVMendset_def)
- then obtain stk loc C' M' pc ics where f1:"f1=(stk,loc,C',M',pc,ics)" by(cases f1)
- have xp: "xp = None" using \<sigma> nend by(simp add: JVMendset_def)
- obtain xp' h' frs' sh' where \<sigma>': "\<sigma>'=(xp',h',frs',sh')" by(cases \<sigma>')
- have ex': "\<exists>t b. P \<turnstile> C has F,b:t in D" using fC by auto
- have field: "field P D F = (D,Static,t)"
- using fC field_def2 has_field_idemp has_field_sees by blast
- have ics': "ics_of (hd frs') = Calling D [] \<and> sh = sh' \<and> frs' \<noteq> [] \<and> xp' = None"
- proof(cases "sh D")
- case None then show ?thesis using field ex' \<sigma>' xp f1 frs \<sigma> assms by auto
- next
- case (Some a)
- then obtain sfs i where a: "a=(sfs,i)" by(cases a)
- show ?thesis using field ex' a Some \<sigma>' xp f1 frs \<sigma> assms by(auto split: init_state.splits)
- qed
- then show ?thesis using ics \<sigma> \<sigma>' by(auto simp: JVMendset_def)
-qed
-
-lemma JVM_Invokestatic_next_step:
-assumes step: "\<sigma>' \<in> JVMsmall P \<sigma>"
- and nend: "\<sigma> \<notin> JVMendset"
- and curr: "curr_instr P (hd(frames_of \<sigma>)) = Invokestatic C M n"
- and mC: "P \<turnstile> C sees M,Static:Ts \<rightarrow> T = m in D"
- and nDone: "\<not>(\<exists>sfs i. sheap \<sigma> D = Some(sfs,i) \<and> i = Done)"
- and ics: "ics_of(hd(frames_of \<sigma>)) = No_ics"
-shows "ics_of (hd(frames_of \<sigma>')) = Calling D [] \<and> sheap \<sigma> = sheap \<sigma>' \<and> \<sigma>' \<notin> JVMendset"
-proof -
- obtain xp h frs sh where \<sigma>: "\<sigma>=(xp,h,frs,sh)" by(cases \<sigma>)
- then obtain f1 frs1 where frs: "frs=f1#frs1" using nend by(cases frs, simp_all add: JVMendset_def)
- then obtain stk loc C' M' pc ics where f1:"f1=(stk,loc,C',M',pc,ics)" by(cases f1)
- have xp: "xp = None" using \<sigma> nend by(simp add: JVMendset_def)
- obtain xp' h' frs' sh' where \<sigma>': "\<sigma>'=(xp',h',frs',sh')" by(cases \<sigma>')
- have ex': "\<exists>Ts T m D b. P \<turnstile> C sees M,b:Ts \<rightarrow> T = m in D" using mC by fastforce
- have method: "\<exists>m. method P C M = (D,Static,m)" using mC by(cases m, auto)
- have ics': "ics_of (hd frs') = Calling D [] \<and> sh = sh' \<and> frs' \<noteq> [] \<and> xp' = None"
- proof(cases "sh D")
- case None then show ?thesis using method ex' \<sigma>' xp f1 frs \<sigma> assms by auto
- next
- case (Some a)
- then obtain sfs i where a: "a=(sfs,i)" by(cases a)
- then have nDone': "i \<noteq> Done" using nDone Some f1 frs \<sigma> by simp
- show ?thesis using method ex' a Some \<sigma>' xp f1 frs \<sigma> assms by(auto split: init_state.splits)
- qed
- then show ?thesis using ics \<sigma> \<sigma>' by(auto simp: JVMendset_def)
-qed
-
-(***********************************************)
-subsection "Definitions"
-
-definition main :: "string" where "main = ''main''"
-definition Test :: "string" where "Test = ''Test''"
-definition test_oracle :: "string" where "test_oracle = ''oracle''"
-
-type_synonym jvm_class = "jvm_method cdecl"
-type_synonym jvm_prog_out = "jvm_state \<times> cname set"
-
-text "A deselection algorithm based on classes that have changed from
- @{text P1} to @{text P2}:"
-primrec jvm_deselect :: "jvm_prog \<Rightarrow> jvm_prog_out \<Rightarrow> jvm_prog \<Rightarrow> bool" where
-"jvm_deselect P1 (\<sigma>, cset) P2 = (cset \<inter> (classes_changed P1 P2) = {})"
-
-definition jvm_progs :: "jvm_prog set" where
-"jvm_progs \<equiv> {P. wf_jvm_prog P \<and> \<not>is_class P Start \<and> \<not>is_class P Test
- \<and> (\<forall>b' Ts' T' m' D'. P \<turnstile> Object sees start_m, b' : Ts'\<rightarrow>T' = m' in D'
- \<longrightarrow> b' = Static \<and> Ts' = [] \<and> T' = Void) }"
-
-definition jvm_tests :: "jvm_class set" where
-"jvm_tests = {t. fst t = Test
- \<and> (\<forall>P \<in> jvm_progs. wf_jvm_prog (t#P) \<and> (\<exists>m. t#P \<turnstile> Test sees main,Static: [] \<rightarrow> Void = m in Test)) }"
-
-abbreviation jvm_make_test_prog :: "jvm_prog \<Rightarrow> jvm_class \<Rightarrow> jvm_prog" where
-"jvm_make_test_prog P t \<equiv> start_prog (t#P) (fst t) main"
-
-declare jvm_progs_def [simp]
-declare jvm_tests_def [simp]
-
-(*****************************************************************************************)
-subsection "Definition lemmas"
-
-lemma jvm_progs_tests_nStart:
-assumes P: "P \<in> jvm_progs" and t: "t \<in> jvm_tests"
-shows "\<not>is_class (t#P) Start"
-using assms by(simp add: is_class_def class_def Start_def Test_def)
-
-lemma jvm_make_test_prog_classes_above_xcpts:
-assumes P: "P \<in> jvm_progs" and t: "t \<in> jvm_tests"
-shows "classes_above_xcpts (jvm_make_test_prog P t) = classes_above_xcpts P"
-proof -
- have nS: "\<not>is_class (t#P) Start" by(rule jvm_progs_tests_nStart[OF P t])
- from P have nT: "\<not>is_class P Test" by simp
- from P t have "wf_syscls (t#P) \<and> wf_syscls P"
- by(simp add: wf_jvm_prog_def wf_jvm_prog_phi_def wf_prog_def)
-
- then have [simp]: "\<And>D. D \<in> sys_xcpts \<Longrightarrow> is_class (t#P) D \<and> is_class P D"
- by(cases t, auto simp: wf_syscls_def is_class_def class_def dest!: weak_map_of_SomeI)
- from wf_nclass_nsub[OF _ _ nS] P t have nspS: "\<And>D. D \<in> sys_xcpts \<Longrightarrow> \<not>(t#P) \<turnstile> D \<preceq>\<^sup>* Start"
- by(auto simp: wf_jvm_prog_def wf_jvm_prog_phi_def)
- from wf_nclass_nsub[OF _ _ nT] P have nspT: "\<And>D. D \<in> sys_xcpts \<Longrightarrow> \<not>P \<turnstile> D \<preceq>\<^sup>* Test"
- by(auto simp: wf_jvm_prog_def wf_jvm_prog_phi_def)
-
- from class_add_classes_above_xcpts[where P="t#P" and C=Start, OF nS nspS]
- have "classes_above_xcpts (jvm_make_test_prog P t) = classes_above_xcpts (t#P)" by simp
- also from class_add_classes_above_xcpts[where P=P and C=Test, OF nT nspT] t
- have "\<dots> = classes_above_xcpts P" by(cases t, simp)
- finally show ?thesis by simp
-qed
-
-lemma jvm_make_test_prog_sees_Test_main:
-assumes P: "P \<in> jvm_progs" and t: "t \<in> jvm_tests"
-shows "\<exists>m. jvm_make_test_prog P t \<turnstile> Test sees main, Static : []\<rightarrow>Void = m in Test"
-proof -
- let ?P = "jvm_make_test_prog P t"
- from P t obtain m where
- meth: "t#P \<turnstile> Test sees main,Static:[] \<rightarrow> Void = m in Test" and
- nstart: "\<not> is_class (t # P) Start"
- by(auto simp: is_class_def class_def Start_def Test_def)
- from class_add_sees_method[OF meth nstart] show ?thesis by fastforce
-qed
-
-(************************************************)
-subsection "Naive RTS algorithm"
-
-subsubsection "Definitions"
-
-fun jvm_naive_out :: "jvm_prog \<Rightarrow> jvm_class \<Rightarrow> jvm_prog_out set" where
-"jvm_naive_out P t = JVMNaiveCollectionSemantics.cbig (jvm_make_test_prog P t) (start_state (t#P))"
-
-abbreviation jvm_naive_collect_start :: "jvm_prog \<Rightarrow> cname set" where
-"jvm_naive_collect_start P \<equiv> {}"
-
-lemma jvm_naive_out_xcpts_collected:
-assumes "o1 \<in> jvm_naive_out P t"
-shows "classes_above_xcpts (start_prog (t # P) (fst t) main) \<subseteq> snd o1"
-using assms
-proof -
- obtain \<sigma>' coll' where "o1 = (\<sigma>', coll')" and
- cbig: "(\<sigma>', coll') \<in> JVMNaiveCollectionSemantics.cbig (start_prog (t#P) (fst t) main) (start_state (t#P))"
- using assms by(cases o1, simp)
- with JVMNaiveCollectionSemantics.cbig_stepD[OF cbig start_state_nend]
- show ?thesis by(auto simp: JVMNaiveCollectionSemantics.csmall_def start_state_def)
-qed
-
-(***********************************************************)
-subsubsection "Naive algorithm correctness"
-
-text "We start with correctness over @{term exec_instr}, then all the
- functions/pieces that are used by naive @{term csmall} (that is, pieces
- used by @{term exec} - such as which frames are used based on @{term ics}
- - and all functions used by the collection function). We then prove that
- @{term csmall} is existence safe, extend this result to @{term cbig}, and
- finally prove the @{term existence_safe} statement over the locale pieces."
-
-\<comment> \<open> if collected classes unchanged, @{term exec_instr} unchanged \<close>
-lemma ncollect_exec_instr:
-assumes "JVMinstr_ncollect P i h stk \<inter> classes_changed P P' = {}"
- and above_C: "classes_above P C \<inter> classes_changed P P' = {}"
- and ics: "ics = Called [] \<or> ics = No_ics"
- and i: "i = instrs_of P C M ! pc"
-shows "exec_instr i P h stk loc C M pc ics frs sh = exec_instr i P' h stk loc C M pc ics frs sh"
-using assms proof(cases i)
- case (New C1) then show ?thesis using assms classes_above_blank[of C1 P P']
- by(auto split: init_state.splits option.splits)
-next
- case (Getfield F1 C1) show ?thesis
- proof(cases "hd stk = Null")
- case True then show ?thesis using Getfield assms by simp
- next
- case False
- let ?D = "(cname_of h (the_Addr (hd stk)))"
- have D: "classes_above P ?D \<inter> classes_changed P P' = {}"
- using False Getfield assms by simp
- show ?thesis
- proof(cases "\<exists>b t. P \<turnstile> ?D has F1,b:t in C1")
- case True
- then obtain b1 t1 where "P \<turnstile> ?D has F1,b1:t1 in C1" by auto
- then have has: "P' \<turnstile> ?D has F1,b1:t1 in C1"
- using Getfield assms classes_above_has_field[OF D] by auto
- have "P \<turnstile> ?D \<preceq>\<^sup>* C1" using has_field_decl_above True by auto
- then have "classes_above P C1 \<subseteq> classes_above P ?D" by(rule classes_above_subcls_subset)
- then have C1: "classes_above P C1 \<inter> classes_changed P P' = {}" using D by auto
- then show ?thesis using has True Getfield assms classes_above_field[of C1 P P' F1]
- by(cases "field P' C1 F1", cases "the (h (the_Addr (hd stk)))", auto)
- next
- case nex: False
- then have "\<nexists>b t. P' \<turnstile> ?D has F1,b:t in C1"
- using False Getfield assms
- classes_above_has_field2[where C="?D" and P=P and P'=P' and F=F1 and C'=C1]
- by auto
- then show ?thesis using nex Getfield assms classes_above_field[of C1 P P' F1]
- by(cases "field P' C1 F1", cases "the (h (the_Addr (hd stk)))", auto)
- qed
- qed
-next
- case (Getstatic C1 F1 D1)
- then have C1: "classes_above P C1 \<inter> classes_changed P P' = {}" using assms by auto
- show ?thesis
- proof(cases "\<exists>b t. P \<turnstile> C1 has F1,b:t in D1")
- case True
- then obtain b t where meth: "P \<turnstile> C1 has F1,b:t in D1" by auto
- then have "P \<turnstile> C1 \<preceq>\<^sup>* D1" by(rule has_field_decl_above)
- then have D1: "classes_above P D1 \<inter> classes_changed P P' = {}"
- using C1 rtrancl_trans by fastforce
- have "P' \<turnstile> C1 has F1,b:t in D1"
- using meth Getstatic assms classes_above_has_field[OF C1] by auto
- then show ?thesis using True D1 Getstatic assms classes_above_field[of D1 P P' F1]
- by(cases "field P' D1 F1", auto)
- next
- case False
- then have "\<nexists>b t. P' \<turnstile> C1 has F1,b:t in D1"
- using Getstatic assms
- classes_above_has_field2[where C=C1 and P=P and P'=P' and F=F1 and C'=D1]
- by auto
- then show ?thesis using False Getstatic assms
- by(cases "field P' D1 F1", auto)
- qed
-next
- case (Putfield F1 C1) show ?thesis
- proof(cases "hd(tl stk) = Null")
- case True then show ?thesis using Putfield assms by simp
- next
- case False
- let ?D = "(cname_of h (the_Addr (hd (tl stk))))"
- have D: "classes_above P ?D \<inter> classes_changed P P' = {}" using False Putfield assms by simp
- show ?thesis
- proof(cases "\<exists>b t. P \<turnstile> ?D has F1,b:t in C1")
- case True
- then obtain b1 t1 where "P \<turnstile> ?D has F1,b1:t1 in C1" by auto
- then have has: "P' \<turnstile> ?D has F1,b1:t1 in C1"
- using Putfield assms classes_above_has_field[OF D] by auto
- have "P \<turnstile> ?D \<preceq>\<^sup>* C1" using has_field_decl_above True by auto
- then have "classes_above P C1 \<subseteq> classes_above P ?D" by(rule classes_above_subcls_subset)
- then have C1: "classes_above P C1 \<inter> classes_changed P P' = {}" using D by auto
- then show ?thesis using has True Putfield assms classes_above_field[of C1 P P' F1]
- by(cases "field P' C1 F1", cases "the (h (the_Addr (hd (tl stk))))", auto)
- next
- case nex: False
- then have "\<nexists>b t. P' \<turnstile> ?D has F1,b:t in C1"
- using False Putfield assms
- classes_above_has_field2[where C="?D" and P=P and P'=P' and F=F1 and C'=C1]
- by auto
- then show ?thesis using nex Putfield assms classes_above_field[of C1 P P' F1]
- by(cases "field P' C1 F1", cases "the (h (the_Addr (hd (tl stk))))", auto)
- qed
- qed
-next
- case (Putstatic C1 F1 D1)
- then have C1: "classes_above P C1 \<inter> classes_changed P P' = {}" using Putstatic assms by auto
- show ?thesis
- proof(cases "\<exists>b t. P \<turnstile> C1 has F1,b:t in D1")
- case True
- then obtain b t where meth: "P \<turnstile> C1 has F1,b:t in D1" by auto
- then have "P \<turnstile> C1 \<preceq>\<^sup>* D1" by(rule has_field_decl_above)
- then have D1: "classes_above P D1 \<inter> classes_changed P P' = {}"
- using C1 rtrancl_trans by fastforce
- then have "P' \<turnstile> C1 has F1,b:t in D1"
- using meth Putstatic assms classes_above_has_field[OF C1] by auto
- then show ?thesis using True D1 Putstatic assms classes_above_field[of D1 P P' F1]
- by(cases "field P' D1 F1", auto)
- next
- case False
- then have "\<nexists>b t. P' \<turnstile> C1 has F1,b:t in D1"
- using Putstatic assms classes_above_has_field2[where C=C1 and P=P and P'=P' and F=F1 and C'=D1]
- by auto
- then show ?thesis using False Putstatic assms
- by(cases "field P' D1 F1", auto)
- qed
-next
- case (Checkcast C1)
- then show ?thesis using assms
- proof(cases "hd stk = Null")
- case False then show ?thesis
- using Checkcast assms classes_above_subcls classes_above_subcls2
- by(simp add: cast_ok_def) blast
- qed(simp add: cast_ok_def)
-next
- case (Invoke M n)
- let ?C = "cname_of h (the_Addr (stk ! n))"
- show ?thesis
- proof(cases "stk ! n = Null")
- case True then show ?thesis using Invoke assms by simp
- next
- case False
- then have above: "classes_above P ?C \<inter> classes_changed P P' = {}"
- using Invoke assms by simp
- obtain D b Ts T mxs mxl ins xt where meth: "method P' ?C M = (D,b,Ts,T,mxs,mxl,ins,xt)"
- by(cases "method P' ?C M", clarsimp)
- have meq: "method P ?C M = method P' ?C M"
- using classes_above_method[OF above] by simp
- then show ?thesis
- proof(cases "\<exists>Ts T m D b. P \<turnstile> ?C sees M,b:Ts \<rightarrow> T = m in D")
- case nex: False
- then have "\<not>(\<exists>Ts T m D b. P' \<turnstile> ?C sees M,b:Ts \<rightarrow> T = m in D)"
- using classes_above_sees_method2[OF above, of M] by clarsimp
- then show ?thesis using nex False Invoke by simp
- next
- case True
- then have "\<exists>Ts T m D b. P' \<turnstile> ?C sees M,b:Ts \<rightarrow> T = m in D"
- by(fastforce dest!: classes_above_sees_method[OF above, of M])
- then show ?thesis using meq meth True Invoke by simp
- qed
- qed
-next
- case (Invokestatic C1 M n)
- then have above: "classes_above P C1 \<inter> classes_changed P P' = {}"
- using assms by simp
- obtain D b Ts T mxs mxl ins xt where meth: "method P' C1 M = (D,b,Ts,T,mxs,mxl,ins,xt)"
- by(cases "method P' C1 M", clarsimp)
- have meq: "method P C1 M = method P' C1 M"
- using classes_above_method[OF above] by simp
- show ?thesis
- proof(cases "\<exists>Ts T m D b. P \<turnstile> C1 sees M,b:Ts \<rightarrow> T = m in D")
- case False
- then have "\<not>(\<exists>Ts T m D b. P' \<turnstile> C1 sees M,b:Ts \<rightarrow> T = m in D)"
- using classes_above_sees_method2[OF above, of M] by clarsimp
- then show ?thesis using False Invokestatic by simp
- next
- case True
- then have "\<exists>Ts T m D b. P' \<turnstile> C1 sees M,b:Ts \<rightarrow> T = m in D"
- by(fastforce dest!: classes_above_sees_method[OF above, of M])
- then show ?thesis using meq meth True Invokestatic by simp
- qed
-next
- case Return then show ?thesis using assms classes_above_method[OF above_C]
- by(cases frs, auto)
-next
- case (Load x1) then show ?thesis using assms by auto
-next
- case (Store x2) then show ?thesis using assms by auto
-next
- case (Push x3) then show ?thesis using assms by auto
-next
- case (Goto x15) then show ?thesis using assms by auto
-next
- case (IfFalse x17) then show ?thesis using assms by auto
-qed(auto)
-
-
-\<comment> \<open> if collected classes unchanged, instruction collection unchanged \<close>
-lemma ncollect_JVMinstr_ncollect:
-assumes "JVMinstr_ncollect P i h stk \<inter> classes_changed P P' = {}"
-shows "JVMinstr_ncollect P i h stk = JVMinstr_ncollect P' i h stk"
-proof(cases i)
- case (New C1)
- then show ?thesis using assms classes_above_set[of C1 P P'] by auto
-next
- case (Getfield F C1) show ?thesis
- proof(cases "hd stk = Null")
- case True then show ?thesis using Getfield assms by simp
- next
- case False
- let ?D = "cname_of h (the_Addr (hd stk))"
- have "classes_above P ?D \<inter> classes_changed P P' = {}" using False Getfield assms by auto
- then have "classes_above P ?D = classes_above P' ?D"
- using classes_above_set by blast
- then show ?thesis using assms Getfield by auto
- qed
-next
- case (Getstatic C1 P1 D1)
- then have "classes_above P C1 \<inter> classes_changed P P' = {}" using assms by auto
- then have "classes_above P C1 = classes_above P' C1"
- using classes_above_set assms Getstatic by blast
- then show ?thesis using assms Getstatic by auto
-next
- case (Putfield F C1) show ?thesis
- proof(cases "hd(tl stk) = Null")
- case True then show ?thesis using Putfield assms by simp
- next
- case False
- let ?D = "cname_of h (the_Addr (hd (tl stk)))"
- have "classes_above P ?D \<inter> classes_changed P P' = {}" using False Putfield assms by auto
- then have "classes_above P ?D = classes_above P' ?D"
- using classes_above_set by blast
- then show ?thesis using assms Putfield by auto
- qed
-next
- case (Putstatic C1 F D1)
- then have "classes_above P C1 \<inter> classes_changed P P' = {}" using assms by auto
- then have "classes_above P C1 = classes_above P' C1"
- using classes_above_set assms Putstatic by blast
- then show ?thesis using assms Putstatic by auto
-next
- case (Checkcast C1)
- then show ?thesis using assms
- classes_above_set[of "cname_of h (the_Addr (hd stk))" P P'] by auto
-next
- case (Invoke M n)
- then show ?thesis using assms
- classes_above_set[of "cname_of h (the_Addr (stk ! n))" P P'] by auto
-next
- case (Invokestatic C1 M n)
- then show ?thesis using assms classes_above_set[of C1 P P'] by auto
-next
- case Return
- then show ?thesis using assms classes_above_set[of _ P P'] by auto
-next
- case Throw
- then show ?thesis using assms
- classes_above_set[of "cname_of h (the_Addr (hd stk))" P P'] by auto
-qed(auto)
-
-\<comment> \<open> if collected classes unchanged, @{term exec_step} unchanged \<close>
-lemma ncollect_exec_step:
-assumes "JVMstep_ncollect P h stk C M pc ics \<inter> classes_changed P P' = {}"
- and above_C: "classes_above P C \<inter> classes_changed P P' = {}"
-shows "exec_step P h stk loc C M pc ics frs sh = exec_step P' h stk loc C M pc ics frs sh"
-proof(cases ics)
- case No_ics then show ?thesis
- using ncollect_exec_instr assms classes_above_method[OF above_C, THEN sym]
- by simp
-next
- case (Calling C1 Cs)
- then have above_C1: "classes_above P C1 \<inter> classes_changed P P' = {}"
- using assms(1) by auto
- show ?thesis
- proof(cases "sh C1")
- case None
- then show ?thesis using Calling assms classes_above_sblank[OF above_C1] by simp
- next
- case (Some a)
- then obtain sfs i where sfsi: "a = (sfs, i)" by(cases a)
- then show ?thesis using Calling Some assms
- proof(cases i)
- case Prepared then show ?thesis
- using above_C1 sfsi Calling Some assms classes_above_method[OF above_C1]
- by(simp add: split_beta classes_above_class classes_changed_class[where cn=C1])
- next
- case Error then show ?thesis
- using above_C1 sfsi Calling Some assms classes_above_method[of C1 P P']
- by(simp add: split_beta classes_above_class classes_changed_class[where cn=C1])
- qed(auto)
- qed
-next
- case (Called Cs) show ?thesis
- proof(cases Cs)
- case Nil then show ?thesis
- using ncollect_exec_instr assms classes_above_method[OF above_C, THEN sym] Called
- by simp
- next
- case (Cons C1 Cs1)
- then have above_C': "classes_above P C1 \<inter> classes_changed P P' = {}" using assms Called by auto
- show ?thesis using assms classes_above_method[OF above_C'] Cons Called by simp
- qed
-next
- case (Throwing Cs a) then show ?thesis using assms by(cases Cs; simp)
-qed
-
-\<comment> \<open> if collected classes unchanged, @{term exec_step} collection unchanged \<close>
-lemma ncollect_JVMstep_ncollect:
-assumes "JVMstep_ncollect P h stk C M pc ics \<inter> classes_changed P P' = {}"
- and above_C: "classes_above P C \<inter> classes_changed P P' = {}"
-shows "JVMstep_ncollect P h stk C M pc ics = JVMstep_ncollect P' h stk C M pc ics"
-proof(cases ics)
- case No_ics then show ?thesis
- using assms ncollect_JVMinstr_ncollect classes_above_method[OF above_C]
- by simp
-next
- case (Calling C1 Cs)
- then have above_C: "classes_above P C1 \<inter> classes_changed P P' = {}"
- using assms(1) by auto
- let ?C = "fst(method P C1 clinit)"
- show ?thesis using Calling assms classes_above_method[OF above_C]
- classes_above_set[OF above_C] by auto
-next
- case (Called Cs) show ?thesis
- proof(cases Cs)
- case Nil then show ?thesis
- using assms ncollect_JVMinstr_ncollect classes_above_method[OF above_C] Called
- by simp
- next
- case (Cons C1 Cs1)
- then have above_C1: "classes_above P C1 \<inter> classes_changed P P' = {}"
- and above_C': "classes_above P (fst (method P C1 clinit)) \<inter> classes_changed P P' = {}"
- using assms Called by auto
- show ?thesis using assms Cons Called classes_above_set[OF above_C1]
- classes_above_set[OF above_C'] classes_above_method[OF above_C1]
- by simp
- qed
-next
- case (Throwing Cs a) then show ?thesis
- using assms classes_above_set[of "cname_of h a" P P'] by simp
-qed
-
-\<comment> \<open> if collected classes unchanged, @{term classes_above_frames} unchanged \<close>
-lemma ncollect_classes_above_frames:
- "JVMexec_ncollect P (None, h, (stk,loc,C,M,pc,ics)#frs, sh) \<inter> classes_changed P P' = {}
- \<Longrightarrow> classes_above_frames P frs = classes_above_frames P' frs"
-proof(induct frs)
- case (Cons f frs')
- then obtain stk loc C M pc ics where f: "f = (stk,loc,C,M,pc,ics)" by(cases f)
- then have above_C: "classes_above P C \<inter> classes_changed P P' = {}" using Cons by auto
- show ?case using f Cons classes_above_subcls[OF above_C]
- classes_above_subcls2[OF above_C] by auto
-qed(auto)
-
-\<comment> \<open> if collected classes unchanged, @{term classes_above_xcpts} unchanged \<close>
-lemma ncollect_classes_above_xcpts:
-assumes "JVMexec_ncollect P (None, h, (stk,loc,C,M,pc,ics)#frs, sh) \<inter> classes_changed P P' = {}"
-shows "classes_above_xcpts P = classes_above_xcpts P'"
-proof -
- have left: "\<And>x x'. x' \<in> sys_xcpts \<Longrightarrow> P \<turnstile> x' \<preceq>\<^sup>* x \<Longrightarrow> \<exists>xa\<in>sys_xcpts. P' \<turnstile> xa \<preceq>\<^sup>* x"
- proof -
- fix x x'
- assume x': "x' \<in> sys_xcpts" and above: "P \<turnstile> x' \<preceq>\<^sup>* x"
- then show "\<exists>xa\<in>sys_xcpts. P' \<turnstile> xa \<preceq>\<^sup>* x" using assms classes_above_subcls[OF _ above]
- by(rule_tac x=x' in bexI) auto
- qed
- have right: "\<And>x x'. x' \<in> sys_xcpts \<Longrightarrow> P' \<turnstile> x' \<preceq>\<^sup>* x \<Longrightarrow> \<exists>xa\<in>sys_xcpts. P \<turnstile> xa \<preceq>\<^sup>* x"
- proof -
- fix x x'
- assume x': "x' \<in> sys_xcpts" and above: "P' \<turnstile> x' \<preceq>\<^sup>* x"
- then show "\<exists>xa\<in>sys_xcpts. P \<turnstile> xa \<preceq>\<^sup>* x" using assms classes_above_subcls2[OF _ above]
- by(rule_tac x=x' in bexI) auto
- qed
- show ?thesis using left right by auto
-qed
-
-\<comment> \<open> if collected classes unchanged, @{term exec} collection unchanged \<close>
-lemma ncollect_JVMexec_ncollect:
-assumes "JVMexec_ncollect P \<sigma> \<inter> classes_changed P P' = {}"
-shows "JVMexec_ncollect P \<sigma> = JVMexec_ncollect P' \<sigma>"
-proof -
- obtain xp h frs sh where \<sigma>: "\<sigma> = (xp,h,frs,sh)" by(cases \<sigma>)
- then show ?thesis using assms
- proof(cases "\<exists>x. xp = Some x \<or> frs = []")
- case False
- then obtain stk loc C M pc ics frs' where frs: "frs = (stk,loc,C,M,pc,ics)#frs'"
- by(cases frs, auto)
- have step: "JVMstep_ncollect P h stk C M pc ics \<inter> classes_changed P P' = {}"
- using False \<sigma> frs assms by(cases ics, auto simp: JVMNaiveCollectionSemantics.csmall_def)
- have above_C: "classes_above P C \<inter> classes_changed P P' = {}"
- using False \<sigma> frs assms by(auto simp: JVMNaiveCollectionSemantics.csmall_def)
- have frames: "classes_above_frames P frs' = classes_above_frames P' frs'"
- using ncollect_classes_above_frames frs \<sigma> False assms by simp
- have xcpts: "classes_above_xcpts P = classes_above_xcpts P'"
- using ncollect_classes_above_xcpts frs \<sigma> False assms by simp
- show ?thesis using False xcpts frames frs \<sigma> ncollect_JVMstep_ncollect[OF step above_C]
- classes_above_subcls[OF above_C] classes_above_subcls2[OF above_C]
- by auto
- qed(auto)
-qed
-
-\<comment> \<open> if collected classes unchanged, classes above an exception returned
- by @{term exec_instr} unchanged \<close>
-lemma ncollect_exec_instr_xcpts:
-assumes collect: "JVMinstr_ncollect P i h stk \<inter> classes_changed P P' = {}"
- and xpcollect: "classes_above_xcpts P \<inter> classes_changed P P' = {}"
- and prealloc: "preallocated h"
- and \<sigma>': "\<sigma>' = exec_instr i P h stk loc C M pc ics' frs sh"
- and xp: "fst \<sigma>' = Some a"
- and i: "i = instrs_of P C M ! pc"
-shows "classes_above P (cname_of h a) \<inter> classes_changed P P' = {}"
-using assms exec_instr_xcpts[OF \<sigma>' xp]
-proof(cases i)
- case Throw then show ?thesis using assms by(cases "hd stk", fastforce+)
-qed(fastforce+)
-
-\<comment> \<open> if collected classes unchanged, classes above an exception returned
- by @{term exec_step} unchanged \<close>
-lemma ncollect_exec_step_xcpts:
-assumes collect: "JVMstep_ncollect P h stk C M pc ics \<inter> classes_changed P P' = {}"
- and xpcollect: "classes_above_xcpts P \<inter> classes_changed P P' = {}"
- and prealloc: "preallocated h"
- and \<sigma>': "\<sigma>' = exec_step P h stk loc C M pc ics frs sh"
- and xp: "fst \<sigma>' = Some a"
-shows "classes_above P (cname_of h a) \<inter> classes_changed P P' = {}"
-proof(cases ics)
- case No_ics then show ?thesis using assms ncollect_exec_instr_xcpts by simp
-next
- case (Calling x21 x22)
- then show ?thesis using assms by(clarsimp split: option.splits init_state.splits if_split_asm)
-next
- case (Called Cs) then show ?thesis using assms ncollect_exec_instr_xcpts by(cases Cs; simp)
-next
- case (Throwing Cs a) then show ?thesis using assms by(cases Cs; simp)
-qed
-
-\<comment> \<open> if collected classes unchanged, if @{term csmall} returned a result
- under @{term P}, @{term P'} returns the same \<close>
-lemma ncollect_JVMsmall:
-assumes collect: "(\<sigma>', cset) \<in> JVMNaiveCollectionSemantics.csmall P \<sigma>"
- and intersect: "cset \<inter> classes_changed P P' = {}"
- and prealloc: "preallocated (fst(snd \<sigma>))"
-shows "(\<sigma>', cset) \<in> JVMNaiveCollectionSemantics.csmall P' \<sigma>"
-proof -
- obtain xp h frs sh where \<sigma>: "\<sigma> = (xp,h,frs,sh)" by(cases \<sigma>)
- then have prealloc': "preallocated h" using prealloc by simp
- show ?thesis using \<sigma> assms
- proof(cases "\<exists>x. xp = Some x \<or> frs = []")
- case False
- then obtain stk loc C M pc ics frs' where frs: "frs = (stk,loc,C,M,pc,ics)#frs'"
- by(cases frs, auto)
- have step: "JVMstep_ncollect P h stk C M pc ics \<inter> classes_changed P P' = {}"
- using False \<sigma> frs assms by(cases ics, auto simp: JVMNaiveCollectionSemantics.csmall_def)
- have above_C: "classes_above P C \<inter> classes_changed P P' = {}"
- using False \<sigma> frs assms by(auto simp: JVMNaiveCollectionSemantics.csmall_def)
- obtain xp1 h1 frs1 sh1 where exec: "exec_step P' h stk loc C M pc ics frs' sh = (xp1,h1,frs1,sh1)"
- by(cases "exec_step P' h stk loc C M pc ics frs' sh")
- have collect: "JVMexec_ncollect P \<sigma> = JVMexec_ncollect P' \<sigma>"
- using assms ncollect_JVMexec_ncollect by(simp add: JVMNaiveCollectionSemantics.csmall_def)
- have exec_instr: "exec_step P h stk loc C M pc ics frs' sh
- = exec_step P' h stk loc C M pc ics frs' sh"
- using ncollect_exec_step[OF step above_C] \<sigma> frs False by simp
- show ?thesis
- proof(cases xp1)
- case None then show ?thesis
- using None \<sigma> frs step False assms ncollect_exec_step[OF step above_C] collect exec
- by(auto simp: JVMNaiveCollectionSemantics.csmall_def)
- next
- case (Some a)
- then show ?thesis using \<sigma> assms
- proof(cases xp)
- case None
- have frames: "classes_above_frames P (frames_of \<sigma>) \<inter> classes_changed P P' = {}"
- using None Some frs \<sigma> assms by(auto simp: JVMNaiveCollectionSemantics.csmall_def)
- have frsi: "classes_above_frames P frs \<inter> classes_changed P P' = {}" using \<sigma> frames by simp
- have xpcollect: "classes_above_xcpts P \<inter> classes_changed P P' = {}"
- using None Some frs \<sigma> assms by(auto simp: JVMNaiveCollectionSemantics.csmall_def)
- have xp: "classes_above P (cname_of h a) \<inter> classes_changed P P' = {}"
- using ncollect_exec_step_xcpts[OF step xpcollect prealloc',
- where \<sigma>' = "(xp1,h1,frs1,sh1)" and frs=frs' and loc=loc and a=a and sh=sh]
- exec exec_instr Some assms by auto
- show ?thesis using Some exec \<sigma> frs False assms exec_instr collect
- classes_above_find_handler[where h=h and sh=sh, OF xp frsi]
- by(auto simp: JVMNaiveCollectionSemantics.csmall_def)
- qed(auto simp: JVMNaiveCollectionSemantics.csmall_def)
- qed
- qed(auto simp: JVMNaiveCollectionSemantics.csmall_def)
-qed
-
-\<comment> \<open> if collected classes unchanged, if @{term cbig} returned a result
- under @{term P}, @{term P'} returns the same \<close>
-lemma ncollect_JVMbig:
-assumes collect: "(\<sigma>', cset) \<in> JVMNaiveCollectionSemantics.cbig P \<sigma>"
- and intersect: "cset \<inter> classes_changed P P' = {}"
- and prealloc: "preallocated (fst(snd \<sigma>))"
-shows "(\<sigma>', cset) \<in> JVMNaiveCollectionSemantics.cbig P' \<sigma>"
-using JVMNaiveCollectionSemantics.csmall_to_cbig_prop2[where R="\<lambda>P P' cset. cset \<inter> classes_changed P P' = {}"
- and Q="\<lambda>\<sigma>. preallocated (fst(snd \<sigma>))" and P=P and P'=P' and \<sigma>=\<sigma> and \<sigma>'=\<sigma>' and coll=cset]
- ncollect_JVMsmall JVMsmall_prealloc_pres assms by auto
-
-\<comment> \<open> and finally, RTS algorithm based on @{term ncollect} is existence safe \<close>
-theorem jvm_naive_existence_safe:
-assumes p: "P \<in> jvm_progs" and "P' \<in> jvm_progs" and t: "t \<in> jvm_tests"
- and out: "o1 \<in> jvm_naive_out P t" and "jvm_deselect P o1 P'"
-shows "\<exists>o2 \<in> jvm_naive_out P' t. o1 = o2"
-using assms
-proof -
- let ?P = "start_prog (t#P) (fst t) main"
- let ?P' = "start_prog (t#P') (fst t) main"
- obtain wf_md where wf': "wf_prog wf_md (t#P)" using p t
- by(auto simp: wf_jvm_prog_def wf_jvm_prog_phi_def)
- have ns: "\<not>is_class (t#P) Start" using p t
- by(clarsimp simp: is_class_def class_def Start_def Test_def)
- obtain \<sigma>1 coll1 where o1: "o1 = (\<sigma>1, coll1)" by(cases o1)
- then have cbig: "(\<sigma>1, coll1) \<in> JVMNaiveCollectionSemantics.cbig ?P (start_state (t # P))"
- using assms by simp
- have "coll1 \<inter> classes_changed P P' = {}" using cbig o1 assms by auto
- then have changed: "coll1 \<inter> classes_changed (t#P) (t#P') = {}" by(rule classes_changed_int_Cons)
- then have changed': "coll1 \<inter> classes_changed ?P ?P' = {}" by(rule classes_changed_int_Cons)
- have "classes_above_xcpts ?P = classes_above_xcpts (t#P)"
- using class_add_classes_above[OF ns wf_sys_xcpt_nsub_Start[OF wf' ns]] by simp
- then have "classes_above_xcpts (t#P) \<inter> classes_changed (t#P) (t#P') = {}"
- using jvm_naive_out_xcpts_collected[OF out] o1 changed by auto
- then have ss_eq: "start_state (t#P) = start_state (t#P')"
- using classes_above_start_state by simp
- show ?thesis using ncollect_JVMbig[OF cbig changed']
- preallocated_start_state changed' ss_eq o1 assms by auto
-qed
-
-\<comment> \<open> ...thus @{term JVMNaiveCollection} is an instance of @{term CollectionBasedRTS} \<close>
-interpretation JVMNaiveCollectionRTS :
- CollectionBasedRTS "(=)" jvm_deselect jvm_progs jvm_tests
- JVMendset JVMcombine JVMcollect_id JVMsmall JVMNaiveCollect jvm_naive_out
- jvm_make_test_prog jvm_naive_collect_start
- by unfold_locales (rule jvm_naive_existence_safe, auto simp: start_state_def)
-
-(***********************************************************************************************)
-subsection "Smarter RTS algorithm"
-
-subsubsection "Definitions and helper lemmas"
-
-fun jvm_smart_out :: "jvm_prog \<Rightarrow> jvm_class \<Rightarrow> jvm_prog_out set" where
-"jvm_smart_out P t
- = {(\<sigma>',coll'). \<exists>coll. (\<sigma>',coll) \<in> JVMSmartCollectionSemantics.cbig
- (jvm_make_test_prog P t) (start_state (t#P))
- \<and> coll' = coll \<union> classes_above_xcpts P \<union> {Object,Start}}"
-
-abbreviation jvm_smart_collect_start :: "jvm_prog \<Rightarrow> cname set" where
-"jvm_smart_collect_start P \<equiv> classes_above_xcpts P \<union> {Object,Start}"
-
-
-lemma jvm_naive_iff_smart:
-"(\<exists>cset\<^sub>n. (\<sigma>',cset\<^sub>n) \<in> jvm_naive_out P t) \<longleftrightarrow> (\<exists>cset\<^sub>s. (\<sigma>',cset\<^sub>s) \<in> jvm_smart_out P t)"
- by(auto simp: JVMNaiveCollectionSemantics.cbig_big_equiv
- JVMSmartCollectionSemantics.cbig_big_equiv)
-
-(**************************************************)
-
-lemma jvm_smart_out_classes_above_xcpts:
-assumes s: "(\<sigma>',cset\<^sub>s) \<in> jvm_smart_out P t" and P: "P \<in> jvm_progs" and t: "t \<in> jvm_tests"
-shows "classes_above_xcpts (jvm_make_test_prog P t) \<subseteq> cset\<^sub>s"
- using jvm_make_test_prog_classes_above_xcpts[OF P t] s by clarsimp
-
-lemma jvm_smart_collect_start_make_test_prog:
- "\<lbrakk> P \<in> jvm_progs; t \<in> jvm_tests \<rbrakk>
- \<Longrightarrow> jvm_smart_collect_start (jvm_make_test_prog P t) = jvm_smart_collect_start P"
- using jvm_make_test_prog_classes_above_xcpts by simp
-
-lemma jvm_smart_out_classes_above_start_heap:
-assumes s: "(\<sigma>',cset\<^sub>s) \<in> jvm_smart_out P t" and h: "start_heap (t#P) a = Some(C,fs)"
- and P: "P \<in> jvm_progs" and t: "t \<in> jvm_tests"
-shows "classes_above (jvm_make_test_prog P t) C \<subseteq> cset\<^sub>s"
-using start_heap_classes[OF h] jvm_smart_out_classes_above_xcpts[OF s P t] by auto
-
-lemma jvm_smart_out_classes_above_start_sheap:
-assumes "(\<sigma>',cset\<^sub>s) \<in> jvm_smart_out P t" and "start_sheap C = Some(sfs,i)"
-shows "classes_above (jvm_make_test_prog P t) C \<subseteq> cset\<^sub>s"
-using assms start_prog_classes_above_Start by(clarsimp split: if_split_asm)
-
-lemma jvm_smart_out_classes_above_frames:
- "(\<sigma>',cset\<^sub>s) \<in> jvm_smart_out P t
- \<Longrightarrow> classes_above_frames (jvm_make_test_prog P t) (frames_of (start_state (t#P))) \<subseteq> cset\<^sub>s"
-using start_prog_classes_above_Start by(clarsimp split: if_split_asm simp: start_state_def)
-
-(**************************************************)
-subsubsection "Additional well-formedness conditions"
-
-\<comment> \<open> returns class to be initialized by given instruction, if applicable \<close>
-(* NOTE: similar to exp-taking init_class from J/EConform - but requires field existence checks *)
-fun coll_init_class :: "'m prog \<Rightarrow> instr \<Rightarrow> cname option" where
-"coll_init_class P (New C) = Some C" |
-"coll_init_class P (Getstatic C F D) = (if \<exists>t. P \<turnstile> C has F,Static:t in D
- then Some D else None)" |
-"coll_init_class P (Putstatic C F D) = (if \<exists>t. P \<turnstile> C has F,Static:t in D
- then Some D else None)" |
-"coll_init_class P (Invokestatic C M n) = seeing_class P C M" |
-"coll_init_class _ _ = None"
-
-\<comment> \<open> checks whether the given value is a pointer; if it's an address,
- checks whether it points to an object in the given heap \<close>
-fun is_ptr :: "heap \<Rightarrow> val \<Rightarrow> bool" where
-"is_ptr h Null = True" |
-"is_ptr h (Addr a) = (\<exists>Cfs. h a = Some Cfs)" |
-"is_ptr h _ = False"
-
-lemma is_ptrD: "is_ptr h v \<Longrightarrow> v = Null \<or> (\<exists>a. v = Addr a \<and> (\<exists>Cfs. h a = Some Cfs))"
- by(cases v, auto)
-
-\<comment> \<open> shorthand for: given stack has entries required by given instr,
- including pointer where necessary \<close>
-fun stack_safe :: "instr \<Rightarrow> heap \<Rightarrow> val list \<Rightarrow> bool" where
-"stack_safe (Getfield F C) h stk = (length stk > 0 \<and> is_ptr h (hd stk))" |
-"stack_safe (Putfield F C) h stk = (length stk > 1 \<and> is_ptr h (hd (tl stk)))" |
-"stack_safe (Checkcast C) h stk = (length stk > 0 \<and> is_ptr h (hd stk))" |
-"stack_safe (Invoke M n) h stk = (length stk > n \<and> is_ptr h (stk ! n))" |
-"stack_safe JVMInstructions.Throw h stk = (length stk > 0 \<and> is_ptr h (hd stk))" |
-"stack_safe i h stk = True"
-
-lemma well_formed_stack_safe:
-assumes wtp: "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P" and correct: "P,\<Phi> \<turnstile> (xp,h,(stk,loc,C,M,pc,ics)#frs,sh)\<surd>"
-shows "stack_safe (instrs_of P C M ! pc) h stk"
-proof -
- from correct obtain b Ts T mxs mxl\<^sub>0 ins xt where
- mC: "P \<turnstile> C sees M,b:Ts\<rightarrow>T=(mxs,mxl\<^sub>0,ins,xt) in C" and
- pc: "pc < length ins" by clarsimp
- from sees_wf_mdecl[OF _ mC] wtp have "wt_method P C b Ts T mxs mxl\<^sub>0 ins xt (\<Phi> C M)"
- by(auto simp: wf_jvm_prog_phi_def wf_mdecl_def)
- with pc have wt: "P,T,mxs,length ins,xt \<turnstile> ins ! pc,pc :: \<Phi> C M" by(simp add: wt_method_def)
- from mC correct obtain ST LT where
- \<Phi>: "\<Phi> C M ! pc = Some (ST,LT)" and
- stk: "P,h \<turnstile> stk [:\<le>] ST" by fastforce
- show ?thesis
- proof(cases "instrs_of P C M ! pc")
- case (Getfield F D)
- with mC \<Phi> wt stk obtain oT ST' where
- oT: "P \<turnstile> oT \<le> Class D" and
- ST: "ST = oT # ST'" by fastforce
- with stk obtain ref stk' where
- stk': "stk = ref#stk'" and
- ref: "P,h \<turnstile> ref :\<le> oT" by auto
- with ref oT have "ref = Null \<or> (ref \<noteq> Null \<and> P,h \<turnstile> ref :\<le> Class D)" by auto
- with Getfield mC have "is_ptr h ref" by(fastforce dest: non_npD)
- with stk' Getfield show ?thesis by auto
- next
- case (Putfield F D)
- with mC \<Phi> wt stk obtain vT oT ST' where
- oT: "P \<turnstile> oT \<le> Class D" and
- ST: "ST = vT # oT # ST'" by fastforce
- with stk obtain v ref stk' where
- stk': "stk = v#ref#stk'" and
- ref: "P,h \<turnstile> ref :\<le> oT" by auto
- with ref oT have "ref = Null \<or> (ref \<noteq> Null \<and> P,h \<turnstile> ref :\<le> Class D)" by auto
- with Putfield mC have "is_ptr h ref" by(fastforce dest: non_npD)
- with stk' Putfield show ?thesis by auto
- next
- case (Checkcast D)
- with mC \<Phi> wt stk obtain oT ST' where
- oT: "is_refT oT" and
- ST: "ST = oT # ST'" by fastforce
- with stk obtain ref stk' where
- stk': "stk = ref#stk'" and
- ref: "P,h \<turnstile> ref :\<le> oT" by auto
- with ref oT have "ref = Null \<or> (ref \<noteq> Null \<and> (\<exists>D'. P,h \<turnstile> ref :\<le> Class D'))"
- by(auto simp: is_refT_def)
- with Checkcast mC have "is_ptr h ref" by(fastforce dest: non_npD)
- with stk' Checkcast show ?thesis by auto
- next
- case (Invoke M1 n)
- with mC \<Phi> wt stk have
- ST: "n < size ST" and
- oT: "ST!n = NT \<or> (\<exists>D. ST!n = Class D)" by auto
- with stk have stk': "n < size stk" by (auto simp: list_all2_lengthD)
- with stk ST oT list_all2_nthD2
- have "stk!n = Null \<or> (stk!n \<noteq> Null \<and> (\<exists>D. P,h \<turnstile> stk!n :\<le> Class D))" by fastforce
- with Invoke mC have "is_ptr h (stk!n)" by(fastforce dest: non_npD)
- with stk' Invoke show ?thesis by auto
- next
- case Throw
- with mC \<Phi> wt stk obtain oT ST' where
- oT: "is_refT oT" and
- ST: "ST = oT # ST'" by fastforce
- with stk obtain ref stk' where
- stk': "stk = ref#stk'" and
- ref: "P,h \<turnstile> ref :\<le> oT" by auto
- with ref oT have "ref = Null \<or> (ref \<noteq> Null \<and> (\<exists>D'. P,h \<turnstile> ref :\<le> Class D'))"
- by(auto simp: is_refT_def)
- with Throw mC have "is_ptr h ref" by(fastforce dest: non_npD)
- with stk' Throw show ?thesis by auto
- qed(simp_all)
-qed
-
-(******************************************)
-subsubsection \<open> Proving naive @{text "\<subseteq>"} smart \<close>
-
-text \<open> We prove that, given well-formedness of the program and state, and "promises"
- about what has or will be collected in previous or future steps, @{term jvm_smart}
- collects everything @{term jvm_naive} does. We prove that promises about previously-
- collected classes ("backward promises") are maintained by execution, and promises
- about to-be-collected classes ("forward promises") are met by the end of execution.
- We then show that the required initial conditions (well-formedness and backward
- promises) are met by the defined start states, and thus that a run test will
- collect at least those classes collected by the naive algorithm. \<close>
-
-\<comment> \<open> Backward promises (classes that should already be collected) \<close>
- \<comment> \<open> - Classes of objects in the heap are collected \<close>
- \<comment> \<open> - Non-@{term None} classes on the static heap are collected \<close>
- \<comment> \<open> - Current classes from the frame stack are collected \<close>
- \<comment> \<open> - Classes of system exceptions are collected \<close>
-
-text "If backward promises have been kept, a single step preserves this property;
- i.e., any classes that have been added to this set (new heap objects, newly prepared
- sheap classes, new frames) are collected by the smart collection algorithm in that
- step or by forward promises:"
-lemma backward_coll_promises_kept:
-assumes
-\<comment> \<open> well-formedness \<close>
- wtp: "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P"
- and correct: "P,\<Phi> \<turnstile> (xp,h,frs,sh)\<surd>"
-\<comment> \<open> defs \<close>
- and f': "hd frs = (stk,loc,C',M',pc,ics)"
-\<comment> \<open> backward promises - will be collected prior \<close>
- and heap: "\<And>C fs. \<exists>a. h a = Some(C,fs) \<Longrightarrow> classes_above P C \<subseteq> cset"
- and sheap: "\<And>C sfs i. sh C = Some(sfs,i) \<Longrightarrow> classes_above P C \<subseteq> cset"
- and xcpts: "classes_above_xcpts P \<subseteq> cset"
- and frames: "classes_above_frames P frs \<subseteq> cset"
-\<comment> \<open> forward promises - will be collected after if not already \<close>
- and init_class_prom: "\<And>C. ics = Called [] \<or> ics = No_ics
- \<Longrightarrow> coll_init_class P (instrs_of P C' M' ! pc) = Some C \<Longrightarrow> classes_above P C \<subseteq> cset"
- and Calling_prom: "\<And>C' Cs'. ics = Calling C' Cs' \<Longrightarrow> classes_above P C' \<subseteq> cset"
-\<comment> \<open> collection and step \<close>
- and smart: "JVMexec_scollect P (xp,h,frs,sh) \<subseteq> cset"
- and small: "(xp',h',frs',sh') \<in> JVMsmall P (xp,h,frs,sh)"
-shows "(h' a = Some(C,fs) \<longrightarrow> classes_above P C \<subseteq> cset)
- \<and> (sh' C = Some(sfs',i') \<longrightarrow> classes_above P C \<subseteq> cset)
- \<and> (classes_above_frames P frs' \<subseteq> cset)"
-using assms
-proof(cases frs)
- case (Cons f1 frs1)
-(****)
- then have cr': "P,\<Phi> \<turnstile> (xp,h,(stk,loc,C',M',pc,ics)#frs1,sh)\<surd>" using correct f' by simp
- let ?i = "instrs_of P C' M' ! pc"
- from well_formed_stack_safe[OF wtp cr'] correct_state_Throwing_ex[OF cr'] obtain
- stack_safe: "stack_safe ?i h stk" and
- Throwing_ex: "\<And>Cs a. ics = Throwing Cs a \<Longrightarrow> \<exists>obj. h a = Some obj" by simp
- have confc: "conf_clinit P sh frs" using correct Cons by simp
- have Called_prom: "\<And>C' Cs'. ics = Called (C'#Cs')
- \<Longrightarrow> classes_above P C' \<subseteq> cset \<and> classes_above P (fst(method P C' clinit)) \<subseteq> cset"
- proof -
- fix C' Cs' assume [simp]: "ics = Called (C'#Cs')"
- then have "C' \<in> set(clinit_classes frs)" using f' Cons by simp
- then obtain sfs where shC': "sh C' = Some(sfs, Processing)" and "is_class P C'"
- using confc by(auto simp: conf_clinit_def)
- then have C'eq: "C' = fst(method P C' clinit)" using wf_sees_clinit wtp
- by(fastforce simp: is_class_def wf_jvm_prog_phi_def)
- then show "classes_above P C' \<subseteq> cset \<and> classes_above P (fst(method P C' clinit)) \<subseteq> cset"
- using sheap shC' by auto
- qed
- have Called_prom2: "\<And>Cs. ics = Called Cs \<Longrightarrow> \<exists>C1 sobj. Called_context P C1 ?i \<and> sh C1 = Some sobj"
- using cr' by(auto simp: conf_f_def2)
- have Throwing_prom: "\<And>C' Cs a. ics = Throwing (C'#Cs) a \<Longrightarrow> \<exists>sfs. sh C' = Some(sfs, Processing)"
- proof -
- fix C' Cs a assume [simp]: "ics = Throwing (C'#Cs) a"
- then have "C' \<in> set(clinit_classes frs)" using f' Cons by simp
- then show "\<exists>sfs. sh C' = Some(sfs, Processing)" using confc by(clarsimp simp: conf_clinit_def)
- qed
-(****)
- show ?thesis using Cons assms
- proof(cases xp)
- case None
- then have exec: "exec (P, None, h, (stk,loc,C',M',pc,ics)#frs1, sh) = Some (xp',h',frs',sh')"
- using small f' Cons by auto
- obtain si where si: "exec_step_input P C' M' pc ics = si" by simp
- obtain xp\<^sub>0 h\<^sub>0 frs\<^sub>0 sh\<^sub>0 where
- exec_step: "exec_step P h stk loc C' M' pc ics frs1 sh = (xp\<^sub>0, h\<^sub>0, frs\<^sub>0, sh\<^sub>0)"
- by(cases "exec_step P h stk loc C' M' pc ics frs1 sh")
- then have ind: "exec_step_ind si P h stk loc C' M' pc ics frs1 sh
- (xp\<^sub>0, h\<^sub>0, frs\<^sub>0, sh\<^sub>0)" using exec_step_ind_equiv si by auto
- then show ?thesis using heap sheap frames exec exec_step f' Cons
- si init_class_prom stack_safe Calling_prom Called_prom Called_prom2 Throwing_prom
- proof(induct rule: exec_step_ind.induct)
- case exec_step_ind_Load show ?case using exec_step_ind_Load.prems(1-7) by auto
- next
- case exec_step_ind_Store show ?case using exec_step_ind_Store.prems(1-7) by auto
- next
- case exec_step_ind_Push show ?case using exec_step_ind_Push.prems(1-7) by auto
- next
- case exec_step_ind_NewOOM_Called show ?case using exec_step_ind_NewOOM_Called.prems(1-7)
- by(auto simp del: find_handler.simps dest!: find_handler_pieces) blast+
- next
- case exec_step_ind_NewOOM_Done show ?case using exec_step_ind_NewOOM_Done.prems(1-7)
- by(auto simp del: find_handler.simps dest!: find_handler_pieces) blast+
- next
- case exec_step_ind_New_Called show ?case
- using exec_step_ind_New_Called.hyps exec_step_ind_New_Called.prems(1-9)
- by(auto split: if_split_asm simp: blank_def dest!: exec_step_input_StepID) blast+
- next
- case exec_step_ind_New_Done show ?case
- using exec_step_ind_New_Done.hyps exec_step_ind_New_Done.prems(1-9)
- by(auto split: if_split_asm simp: blank_def dest!: exec_step_input_StepID) blast+
- next
- case exec_step_ind_New_Init show ?case
- using exec_step_ind_New_Init.prems(1-7) by auto
- next
- case exec_step_ind_Getfield_Null show ?case using exec_step_ind_Getfield_Null.prems(1-7)
- by(auto simp del: find_handler.simps dest!: find_handler_pieces) blast+
- next
- case exec_step_ind_Getfield_NoField show ?case
- using exec_step_ind_Getfield_NoField.prems(1-7)
- by(auto simp del: find_handler.simps dest!: find_handler_pieces) blast+
- next
- case exec_step_ind_Getfield_Static show ?case
- using exec_step_ind_Getfield_Static.prems(1-7)
- by(auto simp del: find_handler.simps dest!: find_handler_pieces) blast+
- next
- case exec_step_ind_Getfield show ?case
- using exec_step_ind_Getfield.prems(1-7) by auto
- next
- case exec_step_ind_Getstatic_NoField show ?case
- using exec_step_ind_Getstatic_NoField.prems(1-7)
- by(auto simp del: find_handler.simps dest!: find_handler_pieces) blast+
- next
- case exec_step_ind_Getstatic_NonStatic show ?case
- using exec_step_ind_Getstatic_NonStatic.prems(1-7)
- by(auto simp del: find_handler.simps dest!: find_handler_pieces) blast+
- next
- case exec_step_ind_Getstatic_Called show ?case
- using exec_step_ind_Getstatic_Called.prems(1-7) by auto
- next
- case exec_step_ind_Getstatic_Done show ?case
- using exec_step_ind_Getstatic_Done.prems(1-7) by auto
- next
- case exec_step_ind_Getstatic_Init show ?case
- using exec_step_ind_Getstatic_Init.prems(1-7) by auto
- next
- case exec_step_ind_Putfield_Null show ?case
- using exec_step_ind_Putfield_Null.prems(1-7)
- by(auto simp del: find_handler.simps dest!: find_handler_pieces) blast+
- next
- case exec_step_ind_Putfield_NoField show ?case
- using exec_step_ind_Putfield_NoField.prems(1-7)
- by(auto simp del: find_handler.simps dest!: find_handler_pieces) blast+
- next
- case exec_step_ind_Putfield_Static show ?case
- using exec_step_ind_Putfield_Static.prems(1-7)
- by(auto simp del: find_handler.simps dest!: find_handler_pieces) blast+
- next
- case (exec_step_ind_Putfield v stk r a D fs h D' b t P C F loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
- obtain a C1 fs where addr: "hd (tl stk) = Null \<or> (hd (tl stk) = Addr a \<and> h a = Some(C1,fs))"
- using exec_step_ind_Putfield.prems(8,10) by(auto dest!: exec_step_input_StepID is_ptrD)
- then have "\<And>a. hd(tl stk) = Addr a \<Longrightarrow> classes_above P C1 \<subseteq> cset"
- using exec_step_ind_Putfield.prems(1) addr by auto
- then show ?case using exec_step_ind_Putfield.hyps exec_step_ind_Putfield.prems(1-7) addr
- by(auto split: if_split_asm) blast+
- next
- case exec_step_ind_Putstatic_NoField show ?case
- using exec_step_ind_Putstatic_NoField.prems(1-7)
- by(auto simp del: find_handler.simps dest!: find_handler_pieces) blast+
- next
- case exec_step_ind_Putstatic_NonStatic show ?case
- using exec_step_ind_Putstatic_NonStatic.prems(1-7)
- by(auto simp del: find_handler.simps dest!: find_handler_pieces) blast+
- next
- case (exec_step_ind_Putstatic_Called D' b t P D F C sh sfs i h stk loc C\<^sub>0 M\<^sub>0 pc Cs frs)
- then have "P \<turnstile> D sees F,Static:t in D" by(simp add: has_field_sees[OF has_field_idemp])
- then have D'eq: "D' = D" using exec_step_ind_Putstatic_Called.hyps(1) by simp
- obtain sobj where "sh D = Some sobj"
- using exec_step_ind_Putstatic_Called.hyps(2) exec_step_ind_Putstatic_Called.prems(8,13)
- by(fastforce dest!: exec_step_input_StepID)
- then show ?case using exec_step_ind_Putstatic_Called.hyps
- exec_step_ind_Putstatic_Called.prems(1-7) D'eq
- by(auto split: if_split_asm) blast+
- next
- case exec_step_ind_Putstatic_Done show ?case
- using exec_step_ind_Putstatic_Done.hyps exec_step_ind_Putstatic_Done.prems(1-7)
- by(auto split: if_split_asm) blast+
- next
- case exec_step_ind_Putstatic_Init show ?case
- using exec_step_ind_Putstatic_Init.hyps exec_step_ind_Putstatic_Init.prems(1-7)
- by(auto split: if_split_asm) blast+
- next
- case exec_step_ind_Checkcast show ?case
- using exec_step_ind_Checkcast.prems(1-7) by auto
- next
- case exec_step_ind_Checkcast_Error show ?case using exec_step_ind_Checkcast_Error.prems(1-7)
- by(auto simp del: find_handler.simps dest!: find_handler_pieces) blast+
- next
- case exec_step_ind_Invoke_Null show ?case using exec_step_ind_Invoke_Null.prems(1-7)
- by(auto simp del: find_handler.simps dest!: find_handler_pieces) blast+
- next
- case exec_step_ind_Invoke_NoMethod show ?case using exec_step_ind_Invoke_NoMethod.prems(1-7)
- by(auto simp del: find_handler.simps dest!: find_handler_pieces) blast+
- next
- case exec_step_ind_Invoke_Static show ?case using exec_step_ind_Invoke_Static.prems(1-7)
- by(auto simp del: find_handler.simps dest!: find_handler_pieces) blast+
- next
- case (exec_step_ind_Invoke ps n stk r C h D b Ts T mxs mxl\<^sub>0 ins xt P)
- have "classes_above P D \<subseteq> cset"
- using exec_step_ind_Invoke.hyps(2,3,5) exec_step_ind_Invoke.prems(1,8,10,13)
- rtrancl_trans[OF sees_method_decl_above[OF exec_step_ind_Invoke.hyps(6)]]
- by(auto dest!: exec_step_input_StepID is_ptrD) blast+
- then show ?case
- using exec_step_ind_Invoke.hyps(7) exec_step_ind_Invoke.prems(1-7) by auto
- next
- case exec_step_ind_Invokestatic_NoMethod
- show ?case using exec_step_ind_Invokestatic_NoMethod.prems(1-7)
- by(auto simp del: find_handler.simps dest!: find_handler_pieces) blast+
- next
- case exec_step_ind_Invokestatic_NonStatic
- show ?case using exec_step_ind_Invokestatic_NonStatic.prems(1-7)
- by(auto simp del: find_handler.simps dest!: find_handler_pieces) blast+
- next
- case (exec_step_ind_Invokestatic_Called ps n stk D b Ts T mxs mxl\<^sub>0 ins xt P C M)
- have "seeing_class P C M = Some D" using exec_step_ind_Invokestatic_Called.hyps(2,3)
- by(fastforce simp: seeing_class_def)
- then have "classes_above P D \<subseteq> cset" using exec_step_ind_Invokestatic_Called.prems(8-9)
- by(fastforce dest!: exec_step_input_StepID)
- then show ?case
- using exec_step_ind_Invokestatic_Called.hyps exec_step_ind_Invokestatic_Called.prems(1-7)
- by(auto simp: seeing_class_def)
- next
- case (exec_step_ind_Invokestatic_Done ps n stk D b Ts T mxs mxl\<^sub>0 ins xt P C M)
- have "seeing_class P C M = Some D" using exec_step_ind_Invokestatic_Done.hyps(2,3)
- by(fastforce simp: seeing_class_def)
- then have "classes_above P D \<subseteq> cset" using exec_step_ind_Invokestatic_Done.prems(8-9)
- by(fastforce dest!: exec_step_input_StepID)
- then show ?case
- using exec_step_ind_Invokestatic_Done.hyps exec_step_ind_Invokestatic_Done.prems(1-7)
- by auto blast+
- next
- case exec_step_ind_Invokestatic_Init show ?case
- using exec_step_ind_Invokestatic_Init.hyps exec_step_ind_Invokestatic_Init.prems(1-7)
- by auto blast+
- next
- case exec_step_ind_Return_Last_Init show ?case
- using exec_step_ind_Return_Last_Init.prems(1-7) by(auto split: if_split_asm) blast+
- next
- case exec_step_ind_Return_Last show ?case
- using exec_step_ind_Return_Last.prems(1-7) by auto
- next
- case exec_step_ind_Return_Init show ?case
- using exec_step_ind_Return_Init.prems(1-7) by(auto split: if_split_asm) blast+
- next
- case exec_step_ind_Return_NonStatic show ?case
- using exec_step_ind_Return_NonStatic.prems(1-7) by auto
- next
- case exec_step_ind_Return_Static show ?case
- using exec_step_ind_Return_Static.prems(1-7) by auto
- next
- case exec_step_ind_Pop show ?case using exec_step_ind_Pop.prems(1-7) by auto
- next
- case exec_step_ind_IAdd show ?case using exec_step_ind_IAdd.prems(1-7)by auto
- next
- case exec_step_ind_IfFalse_False show ?case
- using exec_step_ind_IfFalse_False.prems(1-7) by auto
- next
- case exec_step_ind_IfFalse_nFalse show ?case
- using exec_step_ind_IfFalse_nFalse.prems(1-7) by auto
- next
- case exec_step_ind_CmpEq show ?case using exec_step_ind_CmpEq.prems(1-7) by auto
- next
- case exec_step_ind_Goto show ?case using exec_step_ind_Goto.prems(1-7) by auto
- next
- case exec_step_ind_Throw show ?case using exec_step_ind_Throw.prems(1-7)
- by(auto simp del: find_handler.simps dest!: find_handler_pieces) blast+
- next
- case exec_step_ind_Throw_Null show ?case using exec_step_ind_Throw_Null.prems(1-7)
- by(auto simp del: find_handler.simps dest!: find_handler_pieces) blast+
- next
- case (exec_step_ind_Init_None_Called sh C Cs P)
- have "classes_above P C \<subseteq> cset" using exec_step_ind_Init_None_Called.prems(8,11)
- by(auto dest!: exec_step_input_StepCD)
- then show ?case using exec_step_ind_Init_None_Called.prems(1-7)
- by(auto split: if_split_asm) blast+
- next
- case exec_step_ind_Init_Done show ?case
- using exec_step_ind_Init_Done.prems(1-7) by auto
- next
- case exec_step_ind_Init_Processing show ?case
- using exec_step_ind_Init_Processing.prems(1-7) by auto
- next
- case exec_step_ind_Init_Error show ?case
- using exec_step_ind_Init_Error.prems(1-7) by auto
- next
- case exec_step_ind_Init_Prepared_Object show ?case
- using exec_step_ind_Init_Prepared_Object.hyps
- exec_step_ind_Init_Prepared_Object.prems(1-7,10)
- by(auto split: if_split_asm dest!: exec_step_input_StepCD) blast+
- next
- case exec_step_ind_Init_Prepared_nObject show ?case
- using exec_step_ind_Init_Prepared_nObject.hyps exec_step_ind_Init_Prepared_nObject.prems(1-7)
- by(auto split: if_split_asm) blast+
- next
- case exec_step_ind_Init show ?case
- using exec_step_ind_Init.prems(1-7,8,12)
- by(auto simp: split_beta dest!: exec_step_input_StepC2D)
- next
- case (exec_step_ind_InitThrow C Cs a P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
- obtain sfs where "sh C = Some(sfs,Processing)"
- using exec_step_ind_InitThrow.prems(8,14) by(fastforce dest!: exec_step_input_StepTD)
- then show ?case using exec_step_ind_InitThrow.prems(1-7)
- by(auto split: if_split_asm) blast+
- next
- case exec_step_ind_InitThrow_End show ?case using exec_step_ind_InitThrow_End.prems(1-7)
- by(auto simp del: find_handler.simps dest!: find_handler_pieces) blast+
- qed
- qed(simp)
-qed(simp)
-
-\<comment> \<open> Forward promises (classes that will be collected by the end of execution) \<close>
- \<comment> \<open> - Classes that the current instruction will check initialization for will be collected \<close>
- \<comment> \<open> - Class whose initialization is actively being called by the current frame will be collected \<close>
-
-text \<open> We prove that an @{term ics} of @{text "Calling C Cs"} (meaning @{text C}'s
- initialization procedure is actively being called) means that classes above @{text C}
- will be collected by @{term cbig} (i.e., by the end of execution) using proof by
- induction, proving the base and IH separately. \<close>
-
-\<comment> \<open> base case: @{term Object} \<close>
-lemma Calling_collects_base:
-assumes big: "(\<sigma>', cset') \<in> JVMSmartCollectionSemantics.cbig P \<sigma>"
- and nend: "\<sigma> \<notin> JVMendset"
- and ics: "ics_of (hd(frames_of \<sigma>)) = Calling Object Cs"
-shows "classes_above P Object \<subseteq> cset \<union> cset'"
-proof(cases "frames_of \<sigma>")
- case Nil then show ?thesis using nend by(clarsimp simp: JVMendset_def)
-next
- case (Cons f1 frs1)
- then obtain stk loc C M pc ics where "f1 = (stk,loc,C,M,pc,ics)" by(cases f1)
- then show ?thesis
- using JVMSmartCollectionSemantics.cbig_stepD[OF big nend] nend ics Cons
- by(clarsimp simp: JVMSmartCollectionSemantics.csmall_def JVMendset_def)
-qed
-
-\<comment> \<open> IH case where @{text C} has not been prepared yet \<close>
-lemma Calling_None_next_state:
-assumes ics: "ics_of (hd(frames_of \<sigma>)) = Calling C Cs"
- and none: "sheap \<sigma> C = None"
- and set: "\<forall>C'. P \<turnstile> C \<preceq>\<^sup>* C' \<longrightarrow> (\<exists>sfs i. sheap \<sigma> C' = Some(sfs,i))
- \<longrightarrow> classes_above P C' \<subseteq> cset"
- and \<sigma>': "(\<sigma>', cset') \<in> JVMSmartCollectionSemantics.csmall P \<sigma>"
-shows "\<sigma>' \<notin> JVMendset \<and> ics_of (hd(frames_of \<sigma>')) = Calling C Cs
- \<and> (\<exists>sfs. sheap \<sigma>' C = Some(sfs,Prepared))
- \<and> (\<forall>C'. P \<turnstile> C \<preceq>\<^sup>* C' \<longrightarrow> C \<noteq> C'
- \<longrightarrow> (\<exists>sfs i. sheap \<sigma>' C' = Some(sfs,i)) \<longrightarrow> classes_above P C' \<subseteq> cset)"
-proof(cases "frames_of \<sigma> = [] \<or> (\<exists>x. fst \<sigma> = Some x)")
- case True then show ?thesis using assms
- by(cases \<sigma>, auto simp: JVMSmartCollectionSemantics.csmall_def)
-next
- case False
- then obtain f1 frs1 where frs: "frames_of \<sigma> = f1#frs1" by(cases "frames_of \<sigma>", auto)
- obtain stk loc C' M pc ics where f1: "f1 = (stk,loc,C',M,pc,ics)" by(cases f1)
- show ?thesis using f1 frs False assms
- by(cases \<sigma>, cases "method P C clinit")
- (clarsimp simp: split_beta JVMSmartCollectionSemantics.csmall_def JVMendset_def)
-qed
-
-\<comment> \<open> IH case where @{text C} has been prepared (and has a direct superclass
- - i.e., is not @{term Object}) \<close>
-lemma Calling_Prepared_next_state:
-assumes sub: "P \<turnstile> C \<prec>\<^sup>1 D"
- and obj: "P \<turnstile> D \<preceq>\<^sup>* Object"
- and ics: "ics_of (hd(frames_of \<sigma>)) = Calling C Cs"
- and prep: "sheap \<sigma> C = Some(sfs,Prepared)"
- and set: "\<forall>C'. P \<turnstile> C \<preceq>\<^sup>* C' \<longrightarrow> C \<noteq> C' \<longrightarrow> (\<exists>sfs i. sheap \<sigma> C' = Some(sfs,i))
- \<longrightarrow> classes_above P C' \<subseteq> cset"
- and \<sigma>': "(\<sigma>', cset') \<in> JVMSmartCollectionSemantics.csmall P \<sigma>"
-shows "\<sigma>' \<notin> JVMendset \<and> ics_of (hd (frames_of \<sigma>')) = Calling D (C#Cs)
- \<and> (\<forall>C'. P \<turnstile> D \<preceq>\<^sup>* C' \<longrightarrow> (\<exists>sfs i. sheap \<sigma>' C' = Some(sfs,i))
- \<longrightarrow> classes_above P C' \<subseteq> cset)"
-using sub
-proof(cases "C=Object")
- case nobj:False show ?thesis
- proof(cases "frames_of \<sigma> = [] \<or> (\<exists>x. fst \<sigma> = Some x)")
- case True then show ?thesis using assms
- by(cases \<sigma>, auto simp: JVMSmartCollectionSemantics.csmall_def)
- next
- case False
- then obtain f1 frs1 where frs: "frames_of \<sigma> = f1#frs1" by(cases "frames_of \<sigma>", auto)
- obtain stk loc C' M pc ics where f1: "f1 = (stk,loc,C',M,pc,ics)" by(cases f1)
- have "C \<noteq> D" using sub obj subcls_self_superclass by auto
- then have dimp: "\<forall>C'. P \<turnstile> D \<preceq>\<^sup>* C' \<longrightarrow> P \<turnstile> C \<preceq>\<^sup>* C' \<and> C \<noteq> C'"
- using sub subcls_of_Obj_acyclic[OF obj] by fastforce
- have "\<forall>C'. P \<turnstile> C \<preceq>\<^sup>* C' \<longrightarrow> C \<noteq> C' \<longrightarrow> (\<exists>sfs i. sheap \<sigma>' C' = Some(sfs,i))
- \<longrightarrow> classes_above P C' \<subseteq> cset"
- using f1 frs False nobj assms
- by(cases \<sigma>, cases "method P C clinit")
- (auto simp: JVMSmartCollectionSemantics.csmall_def JVMendset_def)
- then have "\<forall>C'. P \<turnstile> D \<preceq>\<^sup>* C' \<longrightarrow> (\<exists>sfs i. sheap \<sigma>' C' = Some(sfs,i))
- \<longrightarrow> classes_above P C' \<subseteq> cset" using sub dimp by auto
- then show ?thesis using f1 frs False nobj assms
- by(cases \<sigma>, cases "method P C clinit")
- (auto dest:subcls1D simp: JVMSmartCollectionSemantics.csmall_def JVMendset_def)
- qed
-qed(simp)
-
-\<comment> \<open> completed IH case: non-@{term Object} (pulls together above IH cases) \<close>
-lemma Calling_collects_IH:
-assumes sub: "P \<turnstile> C \<prec>\<^sup>1 D"
- and obj: "P \<turnstile> D \<preceq>\<^sup>* Object"
- and step: "\<And>\<sigma> cset' Cs. (\<sigma>', cset') \<in> JVMSmartCollectionSemantics.cbig P \<sigma> \<Longrightarrow> \<sigma> \<notin> JVMendset
- \<Longrightarrow> ics_of (hd(frames_of \<sigma>)) = Calling D Cs
- \<Longrightarrow> \<forall>C'. P \<turnstile> D \<preceq>\<^sup>* C' \<longrightarrow> (\<exists>sfs i. sheap \<sigma> C' = Some(sfs,i))
- \<longrightarrow> classes_above P C' \<subseteq> cset
- \<Longrightarrow> classes_above P D \<subseteq> cset \<union> cset'"
- and big: "(\<sigma>', cset') \<in> JVMSmartCollectionSemantics.cbig P \<sigma>"
- and nend: "\<sigma> \<notin> JVMendset"
- and curr: "ics_of (hd(frames_of \<sigma>)) = Calling C Cs"
- and set: "\<forall>C'. P \<turnstile> C \<preceq>\<^sup>* C' \<longrightarrow> (\<exists>sfs i. sheap \<sigma> C' = Some(sfs,i))
- \<longrightarrow> classes_above P C' \<subseteq> cset"
-shows "classes_above P C \<subseteq> cset \<union> cset'"
-proof(cases "frames_of \<sigma>")
- case Nil then show ?thesis using nend by(clarsimp simp: JVMendset_def)
-next
- case (Cons f1 frs1)
- show ?thesis using sub
- proof(cases "\<exists>sfs i. sheap \<sigma> C = Some(sfs,i)")
- case True then show ?thesis using set by auto
- next
- case False
- obtain stk loc C' M pc ics where f1: "f1 = (stk,loc,C',M,pc,ics)" by(cases f1)
- then obtain \<sigma>1 coll1 coll where \<sigma>1: "(\<sigma>1, coll1) \<in> JVMSmartCollectionSemantics.csmall P \<sigma>"
- "cset' = coll1 \<union> coll" "(\<sigma>', coll) \<in> JVMSmartCollectionSemantics.cbig P \<sigma>1"
- using JVMSmartCollectionSemantics.cbig_stepD[OF big nend] by clarsimp
- show ?thesis
- proof(cases "\<exists>sfs. sheap \<sigma> C = Some(sfs,Prepared)")
- case True
- then obtain sfs where sfs: "sheap \<sigma> C = Some(sfs,Prepared)" by clarsimp
- have set': "\<forall>C'. P \<turnstile> C \<preceq>\<^sup>* C' \<longrightarrow> C\<noteq>C' \<longrightarrow> (\<exists>sfs i. sheap \<sigma> C' = Some(sfs,i))
- \<longrightarrow> classes_above P C' \<subseteq> cset" using set by auto
- then have "\<sigma>1 \<notin> JVMendset \<and> ics_of (hd (frames_of \<sigma>1)) = Calling D (C#Cs)"
- "\<forall>C'. P \<turnstile> D \<preceq>\<^sup>* C' \<longrightarrow> (\<exists>sfs i. sheap \<sigma>1 C' = Some(sfs,i))
- \<longrightarrow> classes_above P C' \<subseteq> cset"
- using Calling_Prepared_next_state[OF sub obj curr sfs set' \<sigma>1(1)]
- by(auto simp: JVMSmartCollectionSemantics.csmall_def)
- then show ?thesis using step[of coll \<sigma>1] classes_above_def2[OF sub] \<sigma>1 f1 Cons nend curr
- by(clarsimp simp: JVMSmartCollectionSemantics.csmall_def JVMendset_def)
- next
- case none: False \<comment> \<open> @{text "Calling C Cs"} is the next @{text ics}, but after that is @{text "Calling D (C#Cs)"} \<close>
- then have sNone: "sheap \<sigma> C = None" using False by(cases "sheap \<sigma> C", auto)
- then have nend1: "\<sigma>1 \<notin> JVMendset" and curr1: "ics_of (hd (frames_of \<sigma>1)) = Calling C Cs"
- and prep: "\<exists>sfs. sheap \<sigma>1 C = \<lfloor>(sfs, Prepared)\<rfloor>"
- and set1: "\<forall>C'. P \<turnstile> C \<preceq>\<^sup>* C' \<longrightarrow> C \<noteq> C' \<longrightarrow> (\<exists>sfs i. sheap \<sigma>1 C' = \<lfloor>(sfs, i)\<rfloor>)
- \<longrightarrow> classes_above P C' \<subseteq> cset"
- using Calling_None_next_state[OF curr sNone set \<sigma>1(1)] by simp+
- then obtain f2 frs2 where frs2: "frames_of \<sigma>1 = f2#frs2"
- by(cases \<sigma>1, cases "frames_of \<sigma>1", clarsimp simp: JVMendset_def)
- obtain sfs1 where sfs1: "sheap \<sigma>1 C = Some(sfs1,Prepared)" using prep by clarsimp
- obtain stk2 loc2 C2 M2 pc2 ics2 where f2: "f2 = (stk2,loc2,C2,M2,pc2,ics2)" by(cases f2)
- then obtain \<sigma>2 coll2 coll' where \<sigma>2: "(\<sigma>2, coll2) \<in> JVMSmartCollectionSemantics.csmall P \<sigma>1"
- "coll = coll2 \<union> coll'" "(\<sigma>', coll') \<in> JVMSmartCollectionSemantics.cbig P \<sigma>2"
- using JVMSmartCollectionSemantics.cbig_stepD[OF \<sigma>1(3) nend1] by clarsimp
- then have "\<sigma>2 \<notin> JVMendset \<and> ics_of (hd (frames_of \<sigma>2)) = Calling D (C#Cs)"
- "\<forall>C'. P \<turnstile> D \<preceq>\<^sup>* C' \<longrightarrow> (\<exists>sfs i. sheap \<sigma>2 C' = Some(sfs,i))
- \<longrightarrow> classes_above P C' \<subseteq> cset"
- using Calling_Prepared_next_state[OF sub obj curr1 sfs1 set1 \<sigma>2(1)]
- by(auto simp: JVMSmartCollectionSemantics.csmall_def)
- then show ?thesis using step[of coll' \<sigma>2] classes_above_def2[OF sub] \<sigma>2 \<sigma>1 f2 frs2 f1 Cons
- nend1 nend curr1 curr
- by(clarsimp simp: JVMSmartCollectionSemantics.csmall_def JVMendset_def)
- qed
- qed
-qed
-
-\<comment> \<open>pulls together above base and IH cases \<close>
-lemma Calling_collects:
-assumes sub: "P \<turnstile> C \<preceq>\<^sup>* Object"
- and "(\<sigma>', cset') \<in> JVMSmartCollectionSemantics.cbig P \<sigma>"
- and "\<sigma> \<notin> JVMendset"
- and "ics_of (hd(frames_of \<sigma>)) = Calling C Cs"
- and "\<forall>C'. P \<turnstile> C \<preceq>\<^sup>* C' \<longrightarrow> (\<exists>sfs i. sheap \<sigma> C' = Some(sfs,i))
- \<longrightarrow> classes_above P C' \<subseteq> cset"
- and "cset' \<subseteq> cset"
-shows "classes_above P C \<subseteq> cset"
-proof -
- have base: "\<forall>\<sigma> cset' Cs.
- (\<sigma>', cset') \<in> JVMSmartCollectionSemantics.cbig P \<sigma> \<longrightarrow> \<sigma> \<notin> JVMendset
- \<longrightarrow> ics_of (hd (frames_of \<sigma>)) = Calling Object Cs
- \<longrightarrow> (\<forall>C'. P \<turnstile> Object \<preceq>\<^sup>* C' \<longrightarrow> (\<exists>sfs i. sheap \<sigma> C' = \<lfloor>(sfs, i)\<rfloor>)
- \<longrightarrow> classes_above P C' \<subseteq> cset)
- \<longrightarrow> classes_above P Object \<subseteq> JVMcombine cset cset'" using Calling_collects_base by simp
- have IH: "\<And>y z. P \<turnstile> y \<prec>\<^sup>1 z \<Longrightarrow>
- P \<turnstile> z \<preceq>\<^sup>* Object \<Longrightarrow>
- \<forall>\<sigma> cset' Cs. (\<sigma>', cset') \<in> JVMSmartCollectionSemantics.cbig P \<sigma> \<longrightarrow> \<sigma> \<notin> JVMendset
- \<longrightarrow> ics_of (hd(frames_of \<sigma>)) = Calling z Cs
- \<longrightarrow> (\<forall>C'. P \<turnstile> z \<preceq>\<^sup>* C' \<longrightarrow> (\<exists>sfs i. sheap \<sigma> C' = Some(sfs,i))
- \<longrightarrow> classes_above P C' \<subseteq> cset)
- \<longrightarrow> classes_above P z \<subseteq> cset \<union> cset' \<Longrightarrow>
- \<forall>\<sigma> cset' Cs. (\<sigma>', cset') \<in> JVMSmartCollectionSemantics.cbig P \<sigma> \<longrightarrow> \<sigma> \<notin> JVMendset
- \<longrightarrow> ics_of (hd(frames_of \<sigma>)) = Calling y Cs
- \<longrightarrow> (\<forall>C'. P \<turnstile> y \<preceq>\<^sup>* C' \<longrightarrow> (\<exists>sfs i. sheap \<sigma> C' = Some(sfs,i))
- \<longrightarrow> classes_above P C' \<subseteq> cset)
- \<longrightarrow> classes_above P y \<subseteq> cset \<union> cset'"
- using Calling_collects_IH by blast
- have result: "\<forall>\<sigma> cset' Cs.
- (\<sigma>', cset') \<in> JVMSmartCollectionSemantics.cbig P \<sigma> \<longrightarrow> \<sigma> \<notin> JVMendset
- \<longrightarrow> ics_of (hd(frames_of \<sigma>)) = Calling C Cs
- \<longrightarrow> (\<forall>C'. P \<turnstile> C \<preceq>\<^sup>* C' \<longrightarrow> (\<exists>sfs i. sheap \<sigma> C' = Some(sfs,i))
- \<longrightarrow> classes_above P C' \<subseteq> cset)
- \<longrightarrow> classes_above P C \<subseteq> cset \<union> cset'"
- using converse_rtrancl_induct[OF sub,
- where P="\<lambda>C. \<forall>\<sigma> cset' Cs. (\<sigma>',cset') \<in> JVMSmartCollectionSemantics.cbig P \<sigma> \<longrightarrow> \<sigma> \<notin> JVMendset
- \<longrightarrow> ics_of (hd(frames_of \<sigma>)) = Calling C Cs
- \<longrightarrow> (\<forall>C'. P \<turnstile> C \<preceq>\<^sup>* C' \<longrightarrow> (\<exists>sfs i. sheap \<sigma> C' = Some(sfs,i))
- \<longrightarrow> classes_above P C' \<subseteq> cset)
- \<longrightarrow> classes_above P C \<subseteq> cset \<union> cset'"]
- using base IH by blast
- then show ?thesis using assms by blast
-qed
-
-(*******************)
-text "Instructions that call the initialization procedure will collect classes above
- the class initialized by the end of execution (using the above @{text Calling_collects})."
-
-lemma New_collects:
-assumes sub: "P \<turnstile> C \<preceq>\<^sup>* Object"
- and cbig: "(\<sigma>', cset') \<in> JVMSmartCollectionSemantics.cbig P \<sigma>"
- and nend: "\<sigma> \<notin> JVMendset"
- and curr: "curr_instr P (hd(frames_of \<sigma>)) = New C"
- and ics: "ics_of (hd(frames_of \<sigma>)) = No_ics"
- and sheap: "\<forall>C'. P \<turnstile> C \<preceq>\<^sup>* C' \<longrightarrow> (\<exists>sfs i. sheap \<sigma> C' = Some(sfs,i))
- \<longrightarrow> classes_above P C' \<subseteq> cset"
- and smart: "cset' \<subseteq> cset"
-shows "classes_above P C \<subseteq> cset"
-proof(cases "(\<exists>sfs i. sheap \<sigma> C = Some(sfs,i) \<and> i = Done)")
- case True then show ?thesis using sheap by auto
-next
- case False
- obtain n where nstep: "(\<sigma>', cset') \<in> JVMSmartCollectionSemantics.csmall_nstep P \<sigma> n"
- and "n \<noteq> 0" using nend cbig JVMSmartCollectionSemantics.cbig_def2
- JVMSmartCollectionSemantics.csmall_nstep_base by (metis empty_iff insert_iff)
- then show ?thesis
- proof(cases n)
- case (Suc n1)
- then obtain \<sigma>1 cset0 cset1 where \<sigma>1: "(\<sigma>1,cset1) \<in> JVMSmartCollectionSemantics.csmall P \<sigma>"
- "cset' = cset1 \<union> cset0" "(\<sigma>',cset0) \<in> JVMSmartCollectionSemantics.csmall_nstep P \<sigma>1 n1"
- using JVMSmartCollectionSemantics.csmall_nstep_SucD nstep by blast
- obtain xp h frs sh where "\<sigma>=(xp,h,frs,sh)" by(cases \<sigma>)
- then have ics1: "ics_of (hd(frames_of \<sigma>1)) = Calling C []"
- and sheap': "sheap \<sigma> = sheap \<sigma>1" and nend1: "\<sigma>1 \<notin> JVMendset"
- using JVM_New_next_step[OF _ nend curr] \<sigma>1(1) False ics
- by(simp add: JVMSmartCollectionSemantics.csmall_def)+
- have "\<sigma>' \<in> JVMendset" using cbig JVMSmartCollectionSemantics.cbig_def2 by blast
- then have cbig1: "(\<sigma>', cset0) \<in> JVMSmartCollectionSemantics.cbig P \<sigma>1"
- using JVMSmartCollectionSemantics.cbig_def2 \<sigma>1(3) by blast
- have sheap1: "\<forall>C'. P \<turnstile> C \<preceq>\<^sup>* C' \<longrightarrow> (\<exists>sfs i. sheap \<sigma>1 C' = \<lfloor>(sfs, i)\<rfloor>)
- \<longrightarrow> classes_above P C' \<subseteq> cset" using sheap' sheap by simp
- have "cset0 \<subseteq> cset" using \<sigma>1(2) smart by blast
- then have "classes_above P C \<subseteq> cset"
- using Calling_collects[OF sub cbig1 nend1 ics1 sheap1] by simp
- then show ?thesis using \<sigma>1(2) smart by auto
- qed(simp)
-qed
-
-lemma Getstatic_collects:
-assumes sub: "P \<turnstile> D \<preceq>\<^sup>* Object"
- and cbig: "(\<sigma>', cset') \<in> JVMSmartCollectionSemantics.cbig P \<sigma>"
- and nend: "\<sigma> \<notin> JVMendset"
- and curr: "curr_instr P (hd(frames_of \<sigma>)) = Getstatic C F D"
- and ics: "ics_of (hd(frames_of \<sigma>)) = No_ics"
- and fC: "P \<turnstile> C has F,Static:t in D"
- and sheap: "\<forall>C'. P \<turnstile> D \<preceq>\<^sup>* C' \<longrightarrow> (\<exists>sfs i. sheap \<sigma> C' = Some(sfs,i))
- \<longrightarrow> classes_above P C' \<subseteq> cset"
- and smart: "cset' \<subseteq> cset"
-shows "classes_above P D \<subseteq> cset"
-proof(cases "(\<exists>sfs i. sheap \<sigma> D = Some(sfs,i) \<and> i = Done)
- \<or> (ics_of(hd(frames_of \<sigma>)) = Called [])")
- case True then show ?thesis
- proof(cases "\<exists>sfs i. sheap \<sigma> D = Some(sfs,i) \<and> i = Done")
- case True then show ?thesis using sheap by auto
- next
- case False
- then have "ics_of(hd(frames_of \<sigma>)) = Called []" using True by clarsimp
- then show ?thesis using ics by auto
- qed
-next
- case False
- obtain n where nstep: "(\<sigma>', cset') \<in> JVMSmartCollectionSemantics.csmall_nstep P \<sigma> n"
- and "n \<noteq> 0" using nend cbig JVMSmartCollectionSemantics.cbig_def2
- JVMSmartCollectionSemantics.csmall_nstep_base by (metis empty_iff insert_iff)
- then show ?thesis
- proof(cases n)
- case (Suc n1)
- then obtain \<sigma>1 cset0 cset1 where \<sigma>1: "(\<sigma>1,cset1) \<in> JVMSmartCollectionSemantics.csmall P \<sigma>"
- "cset' = cset1 \<union> cset0" "(\<sigma>',cset0) \<in> JVMSmartCollectionSemantics.csmall_nstep P \<sigma>1 n1"
- using JVMSmartCollectionSemantics.csmall_nstep_SucD nstep by blast
- obtain xp h frs sh where "\<sigma>=(xp,h,frs,sh)" by(cases \<sigma>)
- then have curr1: "ics_of (hd(frames_of \<sigma>1)) = Calling D []"
- and sheap': "sheap \<sigma> = sheap \<sigma>1" and nend1: "\<sigma>1 \<notin> JVMendset"
- using JVM_Getstatic_next_step[OF _ nend curr fC] \<sigma>1(1) False ics
- by(simp add: JVMSmartCollectionSemantics.csmall_def)+
- have "\<sigma>' \<in> JVMendset" using cbig JVMSmartCollectionSemantics.cbig_def2 by blast
- then have cbig1: "(\<sigma>', cset0) \<in> JVMSmartCollectionSemantics.cbig P \<sigma>1"
- using JVMSmartCollectionSemantics.cbig_def2 \<sigma>1(3) by blast
- have sheap1: "\<forall>C'. P \<turnstile> D \<preceq>\<^sup>* C' \<longrightarrow> (\<exists>sfs i. sheap \<sigma>1 C' = \<lfloor>(sfs, i)\<rfloor>)
- \<longrightarrow> classes_above P C' \<subseteq> cset" using sheap' sheap by simp
- have "cset0 \<subseteq> cset" using \<sigma>1(2) smart by blast
- then have "classes_above P D \<subseteq> cset"
- using Calling_collects[OF sub cbig1 nend1 curr1 sheap1] by simp
- then show ?thesis using \<sigma>1(2) smart by auto
- qed(simp)
-qed
-
-lemma Putstatic_collects:
-assumes sub: "P \<turnstile> D \<preceq>\<^sup>* Object"
- and cbig: "(\<sigma>', cset') \<in> JVMSmartCollectionSemantics.cbig P \<sigma>"
- and nend: "\<sigma> \<notin> JVMendset"
- and curr: "curr_instr P (hd(frames_of \<sigma>)) = Putstatic C F D"
- and ics: "ics_of (hd(frames_of \<sigma>)) = No_ics"
- and fC: "P \<turnstile> C has F,Static:t in D"
- and sheap: "\<forall>C'. P \<turnstile> D \<preceq>\<^sup>* C' \<longrightarrow> (\<exists>sfs i. sheap \<sigma> C' = Some(sfs,i))
- \<longrightarrow> classes_above P C' \<subseteq> cset"
- and smart: "cset' \<subseteq> cset"
-shows "classes_above P D \<subseteq> cset"
-proof(cases "(\<exists>sfs i. sheap \<sigma> D = Some(sfs,i) \<and> i = Done)
- \<or> (ics_of(hd(frames_of \<sigma>)) = Called [])")
- case True then show ?thesis
- proof(cases "\<exists>sfs i. sheap \<sigma> D = Some(sfs,i) \<and> i = Done")
- case True then show ?thesis using sheap by auto
- next
- case False
- then have "ics_of(hd(frames_of \<sigma>)) = Called []" using True by clarsimp
- then show ?thesis using ics by auto
- qed
-next
- case False
- obtain n where nstep: "(\<sigma>', cset') \<in> JVMSmartCollectionSemantics.csmall_nstep P \<sigma> n"
- and "n \<noteq> 0" using nend cbig JVMSmartCollectionSemantics.cbig_def2
- JVMSmartCollectionSemantics.csmall_nstep_base by (metis empty_iff insert_iff)
- then show ?thesis
- proof(cases n)
- case (Suc n1)
- then obtain \<sigma>1 cset0 cset1 where \<sigma>1: "(\<sigma>1,cset1) \<in> JVMSmartCollectionSemantics.csmall P \<sigma>"
- "cset' = cset1 \<union> cset0" "(\<sigma>',cset0) \<in> JVMSmartCollectionSemantics.csmall_nstep P \<sigma>1 n1"
- using JVMSmartCollectionSemantics.csmall_nstep_SucD nstep by blast
- obtain xp h frs sh where "\<sigma>=(xp,h,frs,sh)" by(cases \<sigma>)
- then have curr1: "ics_of (hd(frames_of \<sigma>1)) = Calling D []"
- and sheap': "sheap \<sigma> = sheap \<sigma>1" and nend1: "\<sigma>1 \<notin> JVMendset"
- using JVM_Putstatic_next_step[OF _ nend curr fC] \<sigma>1(1) False ics
- by(simp add: JVMSmartCollectionSemantics.csmall_def)+
- have "\<sigma>' \<in> JVMendset" using cbig JVMSmartCollectionSemantics.cbig_def2 by blast
- then have cbig1: "(\<sigma>', cset0) \<in> JVMSmartCollectionSemantics.cbig P \<sigma>1"
- using JVMSmartCollectionSemantics.cbig_def2 \<sigma>1(3) by blast
- have sheap1: "\<forall>C'. P \<turnstile> D \<preceq>\<^sup>* C' \<longrightarrow> (\<exists>sfs i. sheap \<sigma>1 C' = \<lfloor>(sfs, i)\<rfloor>)
- \<longrightarrow> classes_above P C' \<subseteq> cset" using sheap' sheap by simp
- have "cset0 \<subseteq> cset" using \<sigma>1(2) smart by blast
- then have "classes_above P D \<subseteq> cset"
- using Calling_collects[OF sub cbig1 nend1 curr1 sheap1] by simp
- then show ?thesis using \<sigma>1(2) smart by auto
- qed(simp)
-qed
-
-lemma Invokestatic_collects:
-assumes sub: "P \<turnstile> D \<preceq>\<^sup>* Object"
- and cbig: "(\<sigma>', cset') \<in> JVMSmartCollectionSemantics.cbig P \<sigma>"
- and smart: "cset' \<subseteq> cset"
- and nend: "\<sigma> \<notin> JVMendset"
- and curr: "curr_instr P (hd(frames_of \<sigma>)) = Invokestatic C M n"
- and ics: "ics_of (hd(frames_of \<sigma>)) = No_ics"
- and mC: "P \<turnstile> C sees M,Static:Ts \<rightarrow> T = m in D"
- and sheap: "\<forall>C'. P \<turnstile> D \<preceq>\<^sup>* C' \<longrightarrow> (\<exists>sfs i. sheap \<sigma> C' = Some(sfs,i))
- \<longrightarrow> classes_above P C' \<subseteq> cset"
-shows "classes_above P D \<subseteq> cset"
-proof(cases "(\<exists>sfs i. sheap \<sigma> D = Some(sfs,i) \<and> i = Done)
- \<or> (ics_of(hd(frames_of \<sigma>)) = Called [])")
- case True then show ?thesis
- proof(cases "\<exists>sfs i. sheap \<sigma> D = Some(sfs,i) \<and> i = Done")
- case True then show ?thesis using sheap by auto
- next
- case False
- then have "ics_of(hd(frames_of \<sigma>)) = Called []" using True by clarsimp
- then show ?thesis using ics by auto
- qed
-next
- case False
- obtain n where nstep: "(\<sigma>', cset') \<in> JVMSmartCollectionSemantics.csmall_nstep P \<sigma> n"
- and "n \<noteq> 0" using nend cbig JVMSmartCollectionSemantics.cbig_def2
- JVMSmartCollectionSemantics.csmall_nstep_base by (metis empty_iff insert_iff)
- then show ?thesis
- proof(cases n)
- case (Suc n1)
- then obtain \<sigma>1 cset0 cset1 where \<sigma>1: "(\<sigma>1,cset1) \<in> JVMSmartCollectionSemantics.csmall P \<sigma>"
- "cset' = cset1 \<union> cset0" "(\<sigma>',cset0) \<in> JVMSmartCollectionSemantics.csmall_nstep P \<sigma>1 n1"
- using JVMSmartCollectionSemantics.csmall_nstep_SucD nstep by blast
- obtain xp h frs sh where "\<sigma>=(xp,h,frs,sh)" by(cases \<sigma>)
- then have curr1: "ics_of (hd(frames_of \<sigma>1)) = Calling D []"
- and sheap': "sheap \<sigma> = sheap \<sigma>1" and nend1: "\<sigma>1 \<notin> JVMendset"
- using JVM_Invokestatic_next_step[OF _ nend curr mC] \<sigma>1(1) False ics
- by(simp add: JVMSmartCollectionSemantics.csmall_def)+
- have "\<sigma>' \<in> JVMendset" using cbig JVMSmartCollectionSemantics.cbig_def2 by blast
- then have cbig1: "(\<sigma>', cset0) \<in> JVMSmartCollectionSemantics.cbig P \<sigma>1"
- using JVMSmartCollectionSemantics.cbig_def2 \<sigma>1(3) by blast
- have sheap1: "\<forall>C'. P \<turnstile> D \<preceq>\<^sup>* C' \<longrightarrow> (\<exists>sfs i. sheap \<sigma>1 C' = \<lfloor>(sfs, i)\<rfloor>)
- \<longrightarrow> classes_above P C' \<subseteq> cset" using sheap' sheap by simp
- have "cset0 \<subseteq> cset" using \<sigma>1(2) smart by blast
- then have "classes_above P D \<subseteq> cset"
- using Calling_collects[OF sub cbig1 nend1 curr1 sheap1] by simp
- then show ?thesis using \<sigma>1(2) smart by auto
- qed(simp)
-qed
-
-(*********)
-
-text "The @{text smart_out} execution function keeps the promise to
- collect above the initial class (@{term Test}):"
-lemma jvm_smart_out_classes_above_Test:
-assumes s: "(\<sigma>',cset\<^sub>s) \<in> jvm_smart_out P t" and P: "P \<in> jvm_progs" and t: "t \<in> jvm_tests"
-shows "classes_above (jvm_make_test_prog P t) Test \<subseteq> cset\<^sub>s"
- (is "classes_above ?P ?D \<subseteq> ?cset")
-proof -
- let ?\<sigma> = "start_state (t#P)" and ?M = main
- let ?ics = "ics_of (hd(frames_of ?\<sigma>))"
- have called: "?ics = Called [] \<Longrightarrow> classes_above ?P ?D \<subseteq> ?cset"
- by(simp add: start_state_def)
- then show ?thesis
- proof(cases "?ics = Called []")
- case True then show ?thesis using called by simp
- next
- case False
- from P t obtain wf_md where wf: "wf_prog wf_md (t#P)"
- by(auto simp: wf_jvm_prog_phi_def wf_jvm_prog_def)
- from jvm_make_test_prog_sees_Test_main[OF P t] obtain m where
- mC: "?P \<turnstile> ?D sees ?M,Static:[] \<rightarrow> Void = m in ?D" by fast
- (****)
- then have "?P \<turnstile> ?D \<preceq>\<^sup>* Object" by(rule sees_method_sub_Obj)
- moreover from s obtain cset' where
- cbig: "(\<sigma>', cset') \<in> JVMSmartCollectionSemantics.cbig ?P ?\<sigma>" and "cset' \<subseteq> ?cset" by clarsimp
- moreover have nend: "?\<sigma> \<notin> JVMendset" by(rule start_state_nend)
- moreover from start_prog_start_m_instrs[OF wf] t
- have curr: "curr_instr ?P (hd(frames_of ?\<sigma>)) = Invokestatic ?D ?M 0"
- by(simp add: start_state_def)
- moreover have ics: "?ics = No_ics"
- by(simp add: start_state_def)
- moreover note mC
- moreover from jvm_smart_out_classes_above_start_sheap[OF s]
- have sheap: "\<forall>C'. ?P \<turnstile> ?D \<preceq>\<^sup>* C' \<longrightarrow> (\<exists>sfs i. sheap ?\<sigma> C' = Some(sfs,i))
- \<longrightarrow> classes_above ?P C' \<subseteq> ?cset" by(simp add: start_state_def)
- ultimately show ?thesis by(rule Invokestatic_collects)
- qed
-qed
-
-(**********************************************)
-text "Using lemmas proving preservation of backward promises and keeping
- of forward promises, we prove that the smart algorithm collects at least
- the classes as the naive algorithm does."
-
-\<comment> \<open> first over a single execution step (assumes promises) \<close>
-lemma jvm_naive_to_smart_exec_collect:
-assumes
-\<comment> \<open> well-formedness \<close>
- wtp: "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P"
- and correct: "P,\<Phi> \<turnstile> (xp,h,frs,sh)\<surd>"
-\<comment> \<open> defs \<close>
- and f': "hd frs = (stk,loc,C',M',pc,ics)"
-\<comment> \<open> backward promises - will be collected prior \<close>
- and heap: "\<And>C fs. \<exists>a. h a = Some(C,fs) \<Longrightarrow> classes_above P C \<subseteq> cset"
- and sheap: "\<And>C sfs i. sh C = Some(sfs,i) \<Longrightarrow> classes_above P C \<subseteq> cset"
- and xcpts: "classes_above_xcpts P \<subseteq> cset"
- and frames: "classes_above_frames P frs \<subseteq> cset"
-\<comment> \<open> forward promises - will be collected after if not already \<close>
- and init_class_prom: "\<And>C. ics = Called [] \<or> ics = No_ics
- \<Longrightarrow> coll_init_class P (instrs_of P C' M' ! pc) = Some C \<Longrightarrow> classes_above P C \<subseteq> cset"
- and Calling_prom: "\<And>C' Cs'. ics = Calling C' Cs' \<Longrightarrow> classes_above P C' \<subseteq> cset"
-\<comment> \<open> collection \<close>
- and smart: "JVMexec_scollect P (xp,h,frs,sh) \<subseteq> cset"
-shows "JVMexec_ncollect P (xp,h,frs,sh) \<subseteq> cset"
-using assms
-proof(cases frs)
- case (Cons f' frs')
- then have [simp]: "classes_above P C' \<subseteq> cset" using frames f' by simp
- let ?i = "instrs_of P C' M' ! pc"
- have cr': "P,\<Phi> \<turnstile> (xp,h,(stk,loc,C',M',pc,ics)#frs',sh)\<surd>" using correct f' Cons by simp
- from well_formed_stack_safe[OF wtp cr'] correct_state_Throwing_ex[OF cr'] obtain
- stack_safe: "stack_safe ?i h stk" and
- Throwing_ex: "\<And>Cs a. ics = Throwing Cs a \<Longrightarrow> \<exists>obj. h a = Some obj" by simp
- have confc: "conf_clinit P sh frs" using correct Cons by simp
- have Called_prom: "\<And>C' Cs'. ics = Called (C'#Cs')
- \<Longrightarrow> classes_above P C' \<subseteq> cset \<and> classes_above P (fst(method P C' clinit)) \<subseteq> cset"
- proof -
- fix C' Cs' assume [simp]: "ics = Called (C'#Cs')"
- then have "C' \<in> set(clinit_classes frs)" using f' Cons by simp
- then obtain sfs where shC': "sh C' = Some(sfs, Processing)" and "is_class P C'"
- using confc by(auto simp: conf_clinit_def)
- then have C'eq: "C' = fst(method P C' clinit)" using wf_sees_clinit wtp
- by(fastforce simp: is_class_def wf_jvm_prog_phi_def)
- then show "classes_above P C' \<subseteq> cset \<and> classes_above P (fst(method P C' clinit)) \<subseteq> cset"
- using sheap shC' by auto
- qed
- show ?thesis using Cons assms
- proof(cases xp)
- case None
- { assume ics: "ics = Called [] \<or> ics = No_ics"
- then have [simp]: "JVMexec_ncollect P (xp,h,frs,sh)
- = JVMinstr_ncollect P ?i h stk \<union> classes_above P C'
- \<union> classes_above_frames P frs \<union> classes_above_xcpts P"
- and [simp]: "JVMexec_scollect P (xp,h,frs,sh) = JVMinstr_scollect P ?i"
- using f' None Cons by auto
- have ?thesis using assms
- proof(cases ?i)
- case (New C)
- then have "classes_above P C \<subseteq> cset" using ics New assms by simp
- then show ?thesis using New xcpts frames smart f' by auto
- next
- case (Getfield F C) show ?thesis
- proof(cases "hd stk = Null")
- case True then show ?thesis using Getfield assms by simp
- next
- case False
- let ?C = "cname_of h (the_Addr (hd stk))"
- have above_stk: "classes_above P ?C \<subseteq> cset"
- using stack_safe heap f' False Cons Getfield by(auto dest!: is_ptrD) blast
- then show ?thesis using Getfield assms by auto
- qed
- next
- case (Getstatic C F D)
- show ?thesis
- proof(cases "\<exists>t. P \<turnstile> C has F,Static:t in D")
- case True
- then have above_D: "classes_above P D \<subseteq> cset" using ics init_class_prom Getstatic by simp
- have sub: "P \<turnstile> C \<preceq>\<^sup>* D" using has_field_decl_above True by blast
- then have above_C: "classes_between P C D - {D} \<subseteq> cset"
- using True Getstatic above_D smart f' by simp
- then have "classes_above P C \<subseteq> cset"
- using classes_above_sub_classes_between_eq[OF sub] above_D above_C by auto
- then show ?thesis using Getstatic assms by auto
- next
- case False then show ?thesis using Getstatic assms by auto
- qed
- next
- case (Putfield F C) show ?thesis
- proof(cases "hd(tl stk) = Null")
- case True then show ?thesis using Putfield assms by simp
- next
- case False
- let ?C = "cname_of h (the_Addr (hd (tl stk)))"
- have above_stk: "classes_above P ?C \<subseteq> cset"
- using stack_safe heap f' False Cons Putfield by(auto dest!: is_ptrD) blast
- then show ?thesis using Putfield assms by auto
- qed
- next
- case (Putstatic C F D)
- show ?thesis
- proof(cases "\<exists>t. P \<turnstile> C has F,Static:t in D")
- case True
- then have above_D: "classes_above P D \<subseteq> cset" using ics init_class_prom Putstatic by simp
- have sub: "P \<turnstile> C \<preceq>\<^sup>* D" using has_field_decl_above True by blast
- then have above_C: "classes_between P C D - {D} \<subseteq> cset"
- using True Putstatic above_D smart f' by simp
- then have "classes_above P C \<subseteq> cset"
- using classes_above_sub_classes_between_eq[OF sub] above_D above_C by auto
- then show ?thesis using Putstatic assms by auto
- next
- case False then show ?thesis using Putstatic assms by auto
- qed
- next
- case (Checkcast C) show ?thesis
- proof(cases "hd stk = Null")
- case True then show ?thesis using Checkcast assms by simp
- next
- case False
- let ?C = "cname_of h (the_Addr (hd stk))"
- have above_stk: "classes_above P ?C \<subseteq> cset"
- using stack_safe heap False Cons f' Checkcast by(auto dest!: is_ptrD) blast
- then show ?thesis using above_stk Checkcast assms by(cases "hd stk = Null", auto)
- qed
- next
- case (Invoke M n) show ?thesis
- proof(cases "stk ! n = Null")
- case True then show ?thesis using Invoke assms by simp
- next
- case False
- let ?C = "cname_of h (the_Addr (stk ! n))"
- have above_stk: "classes_above P ?C \<subseteq> cset" using stack_safe heap False Cons f' Invoke
- by(auto dest!: is_ptrD) blast
- then show ?thesis using Invoke assms by auto
- qed
- next
- case (Invokestatic C M n)
- let ?D = "fst (method P C M)"
- show ?thesis
- proof(cases "\<exists>Ts T m D. P \<turnstile> C sees M,Static:Ts \<rightarrow> T = m in D")
- case True
- then have above_D: "classes_above P ?D \<subseteq> cset" using ics init_class_prom Invokestatic
- by(simp add: seeing_class_def)
- have sub: "P \<turnstile> C \<preceq>\<^sup>* ?D" using method_def2 sees_method_decl_above True by auto
- then show ?thesis
- proof(cases "C = ?D")
- case True then show ?thesis
- using Invokestatic above_D xcpts frames smart f' by auto
- next
- case False
- then have above_C: "classes_between P C ?D - {?D} \<subseteq> cset"
- using True Invokestatic above_D smart f' by simp
- then have "classes_above P C \<subseteq> cset"
- using classes_above_sub_classes_between_eq[OF sub] above_D above_C by auto
- then show ?thesis using Invokestatic assms by auto
- qed
- next
- case False then show ?thesis using Invokestatic assms by auto
- qed
- next
- case Throw show ?thesis
- proof(cases "hd stk = Null")
- case True then show ?thesis using Throw assms by simp
- next
- case False
- let ?C = "cname_of h (the_Addr (hd stk))"
- have above_stk: "classes_above P ?C \<subseteq> cset"
- using stack_safe heap False Cons f' Throw by(auto dest!: is_ptrD) blast
- then show ?thesis using above_stk Throw assms by auto
- qed
- next
- case Load then show ?thesis using assms by auto
- next
- case Store then show ?thesis using assms by auto
- next
- case Push then show ?thesis using assms by auto
- next
- case Goto then show ?thesis using assms by auto
- next
- case IfFalse then show ?thesis using assms by auto
- qed(auto)
- }
- moreover
- { fix C1 Cs1 assume ics: "ics = Called (C1#Cs1)"
- then have ?thesis using None Cons Called_prom[OF ics] xcpts frames f' by simp
- }
- moreover
- { fix Cs1 a assume ics: "ics = Throwing Cs1 a"
- then obtain C fs where "h a = Some(C,fs)" using Throwing_ex by fastforce
- then have above_stk: "classes_above P (cname_of h a) \<subseteq> cset" using heap by auto
- then have ?thesis using ics None Cons xcpts frames f' by simp
- }
- moreover
- { fix C1 Cs1 assume ics: "ics = Calling C1 Cs1"
- then have ?thesis using None Cons Calling_prom[OF ics] xcpts frames f' by simp
- }
- ultimately show ?thesis by (metis ics_classes.cases list.exhaust)
- qed(simp)
-qed(simp)
-
-\<comment> \<open> ... which is the same as @{term csmall} \<close>
-lemma jvm_naive_to_smart_csmall:
-assumes
-\<comment> \<open> well-formedness \<close>
- wtp: "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P"
- and correct: "P,\<Phi> \<turnstile> (xp,h,frs,sh)\<surd>"
-\<comment> \<open> defs \<close>
- and f': "hd frs = (stk,loc,C',M',pc,ics)"
-\<comment> \<open> backward promises - will be collected prior \<close>
- and heap: "\<And>C fs. \<exists>a. h a = Some(C,fs) \<Longrightarrow> classes_above P C \<subseteq> cset"
- and sheap: "\<And>C sfs i. sh C = Some(sfs,i) \<Longrightarrow> classes_above P C \<subseteq> cset"
- and xcpts: "classes_above_xcpts P \<subseteq> cset"
- and frames: "classes_above_frames P frs \<subseteq> cset"
-\<comment> \<open> forward promises - will be collected after if not already \<close>
- and init_class_prom: "\<And>C. ics = Called [] \<or> ics = No_ics
- \<Longrightarrow> coll_init_class P (instrs_of P C' M' ! pc) = Some C \<Longrightarrow> classes_above P C \<subseteq> cset"
- and Calling_prom: "\<And>C' Cs'. ics = Calling C' Cs' \<Longrightarrow> classes_above P C' \<subseteq> cset"
-\<comment> \<open> collections \<close>
- and smart_coll: "(\<sigma>', cset\<^sub>s) \<in> JVMSmartCollectionSemantics.csmall P (xp,h,frs,sh)"
- and naive_coll: "(\<sigma>', cset\<^sub>n) \<in> JVMNaiveCollectionSemantics.csmall P (xp,h,frs,sh)"
- and smart: "cset\<^sub>s \<subseteq> cset"
-shows "cset\<^sub>n \<subseteq> cset"
-using jvm_naive_to_smart_exec_collect[where h=h and sh=sh, OF assms(1-9)]
- smart smart_coll naive_coll
- by(fastforce simp: JVMNaiveCollectionSemantics.csmall_def
- JVMSmartCollectionSemantics.csmall_def)
-
-\<comment> \<open> ...which means over @{term csmall_nstep}, stepping from the end state
- (the point by which future promises will have been fulfilled) (uses backward
- and forward promise lemmas) \<close>
-lemma jvm_naive_to_smart_csmall_nstep:
-"\<lbrakk> wf_jvm_prog\<^bsub>\<Phi>\<^esub> P;
- P,\<Phi> \<turnstile> (xp,h,frs,sh)\<surd>;
- hd frs = (stk,loc,C',M',pc,ics);
- \<And>C fs. \<exists>a. h a = Some(C,fs) \<Longrightarrow> classes_above P C \<subseteq> cset;
- \<And>C sfs i. sh C = Some(sfs,i) \<Longrightarrow> classes_above P C \<subseteq> cset;
- classes_above_xcpts P \<subseteq> cset;
- classes_above_frames P frs \<subseteq> cset;
- \<And>C. ics = Called [] \<or> ics = No_ics
- \<Longrightarrow> coll_init_class P (instrs_of P C' M' ! pc) = Some C \<Longrightarrow> classes_above P C \<subseteq> cset;
- \<And>C' Cs'. ics = Calling C' Cs' \<Longrightarrow> classes_above P C' \<subseteq> cset;
- (\<sigma>', cset\<^sub>n) \<in> JVMNaiveCollectionSemantics.csmall_nstep P (xp,h,frs,sh) n;
- (\<sigma>', cset\<^sub>s) \<in> JVMSmartCollectionSemantics.csmall_nstep P (xp,h,frs,sh) n;
- cset\<^sub>s \<subseteq> cset;
- \<sigma>' \<in> JVMendset \<rbrakk>
- \<Longrightarrow> cset\<^sub>n \<subseteq> cset"
-proof(induct n arbitrary: xp h frs sh stk loc C' M' pc ics \<sigma>' cset\<^sub>n cset\<^sub>s cset)
- case 0 then show ?case
- using JVMNaiveCollectionSemantics.csmall_nstep_base subsetI old.prod.inject singletonD
- by (metis (no_types, lifting) equals0D)
-next
- case (Suc n1)
- let ?\<sigma> = "(xp,h,frs,sh)"
- obtain \<sigma>1 cset1 cset' where \<sigma>1: "(\<sigma>1, cset1) \<in> JVMNaiveCollectionSemantics.csmall P ?\<sigma>"
- "cset\<^sub>n = cset1 \<union> cset'" "(\<sigma>', cset') \<in> JVMNaiveCollectionSemantics.csmall_nstep P \<sigma>1 n1"
- using JVMNaiveCollectionSemantics.csmall_nstep_SucD[OF Suc.prems(10)] by clarsimp+
- obtain \<sigma>1' cset1' cset'' where \<sigma>1': "(\<sigma>1', cset1') \<in> JVMSmartCollectionSemantics.csmall P ?\<sigma>"
- "cset\<^sub>s = cset1' \<union> cset''" "(\<sigma>', cset'') \<in> JVMSmartCollectionSemantics.csmall_nstep P \<sigma>1' n1"
- using JVMSmartCollectionSemantics.csmall_nstep_SucD[OF Suc.prems(11)] by clarsimp+
- have \<sigma>_eq: "\<sigma>1 = \<sigma>1'" using \<sigma>1(1) \<sigma>1'(1) by(simp add: JVMNaiveCollectionSemantics.csmall_def
- JVMSmartCollectionSemantics.csmall_def)
- have sub1': "cset1' \<subseteq> cset" and sub'': "cset'' \<subseteq> cset" using Suc.prems(12) \<sigma>1'(2) by auto
- then have sub1: "cset1 \<subseteq> cset"
- using jvm_naive_to_smart_csmall[where h=h and sh=sh and \<sigma>'=\<sigma>1, OF Suc.prems(1-9) _ _ sub1']
- Suc.prems(11,12) \<sigma>1(1) \<sigma>1'(1) \<sigma>_eq by fastforce
- show ?case
- proof(cases n1)
- case 0 then show ?thesis using \<sigma>1(2,3) sub1 by auto
- next
- case Suc2: (Suc n2)
- then have nend1: "\<sigma>1 \<notin> JVMendset"
- using JVMNaiveCollectionSemantics.csmall_nstep_Suc_nend \<sigma>1(3) by blast
- obtain xp1 h1 frs1 sh1 where \<sigma>1_case [simp]: "\<sigma>1 = (xp1,h1,frs1,sh1)" by(cases \<sigma>1)
- obtain stk1 loc1 C1' M1' pc1 ics1 where f1': "hd frs1 = (stk1,loc1,C1',M1',pc1,ics1)"
- by(cases "hd frs1")
- then obtain frs1' where [simp]: "frs1 = (stk1,loc1,C1',M1',pc1,ics1)#frs1'"
- using JVMendset_def nend1 by(cases frs1, auto)
- have cbig1: "(\<sigma>', cset') \<in> JVMNaiveCollectionSemantics.cbig P \<sigma>1"
- "(\<sigma>', cset'') \<in> JVMSmartCollectionSemantics.cbig P \<sigma>1" using \<sigma>1(3) \<sigma>1'(3) Suc.prems(13) \<sigma>_eq
- using JVMNaiveCollectionSemantics.cbig_def2
- JVMSmartCollectionSemantics.cbig_def2 by blast+
- obtain \<sigma>2' cset2' cset2'' where \<sigma>2': "(\<sigma>2', cset2') \<in> JVMSmartCollectionSemantics.csmall P \<sigma>1"
- "cset'' = cset2' \<union> cset2''" "(\<sigma>', cset2'') \<in> JVMSmartCollectionSemantics.csmall_nstep P \<sigma>2' n2"
- using JVMSmartCollectionSemantics.csmall_nstep_SucD \<sigma>1'(3) Suc2 \<sigma>_eq by blast
-(*****)
- have wtp: "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P" by fact
- let ?i1 = "instrs_of P C1' M1' ! pc1"
- let ?ics1 = "ics_of (hd (frames_of \<sigma>1))"
- have step: "P \<turnstile> (xp,h,frs,sh) -jvm\<rightarrow> (xp1,h1,frs1,sh1)"
- proof -
- have "exec (P, ?\<sigma>) = \<lfloor>\<sigma>1'\<rfloor>" using JVMsmart_csmallD[OF \<sigma>1'(1)] by simp
- then have "P \<turnstile> ?\<sigma> -jvm\<rightarrow> \<sigma>1'" using jvm_one_step1[OF exec_1.exec_1I] by simp
- then show ?thesis using Suc.prems(12) \<sigma>_eq by fastforce
- qed
- have correct1: "P,\<Phi> \<turnstile> (xp1,h1,frs1,sh1)\<surd>" by(rule BV_correct[OF wtp step Suc.prems(2)])
-(**)
- have vics1: "P,h1,sh1 \<turnstile>\<^sub>i (C1', M1', pc1, ics1)"
- using correct1 Suc.prems(7) by(auto simp: conf_f_def2)
- from correct1 obtain b Ts T mxs mxl\<^sub>0 ins xt ST LT where
- meth1: "P \<turnstile> C1' sees M1',b:Ts\<rightarrow>T=(mxs,mxl\<^sub>0,ins,xt) in C1'" and
- pc1: "pc1 < length ins" and
- \<Phi>_pc1: "\<Phi> C1' M1'!pc1 = Some (ST,LT)" by(auto dest: sees_method_fun)
- then have wt1: "P,T,mxs,size ins,xt \<turnstile> ins!pc1,pc1 :: \<Phi> C1' M1'"
- using wt_jvm_prog_impl_wt_instr[OF wtp meth1] by fast
-(**)
- have "\<And>a C fs sfs' i'. (h1 a = \<lfloor>(C, fs)\<rfloor> \<longrightarrow> classes_above P C \<subseteq> cset) \<and>
- (sh1 C = \<lfloor>(sfs', i')\<rfloor> \<longrightarrow> classes_above P C \<subseteq> cset) \<and>
- classes_above_frames P frs1 \<subseteq> cset"
- proof -
- fix a C fs sfs' i'
- show "(h1 a = \<lfloor>(C, fs)\<rfloor> \<longrightarrow> classes_above P C \<subseteq> cset) \<and>
- (sh1 C = \<lfloor>(sfs', i')\<rfloor> \<longrightarrow> classes_above P C \<subseteq> cset) \<and>
- (classes_above_frames P frs1 \<subseteq> cset)"
- using Suc.prems(11-12) \<sigma>1' \<sigma>_eq[THEN sym] JVMsmart_csmallD[OF \<sigma>1'(1)]
- backward_coll_promises_kept[where h=h and xp=xp and sh=sh and frs=frs and frs'=frs1
- and xp'=xp1 and h'=h1 and sh'=sh1, OF Suc.prems(1-9)] by auto
- qed
- then have heap1: "\<And>C fs. \<exists>a. h1 a = Some(C,fs) \<Longrightarrow> classes_above P C \<subseteq> cset"
- and sheap1: "\<And>C sfs i. sh1 C = Some(sfs,i) \<Longrightarrow> classes_above P C \<subseteq> cset"
- and frames1: "classes_above_frames P frs1 \<subseteq> cset" by blast+
- have xcpts1: "classes_above_xcpts P \<subseteq> cset" using Suc.prems(6) by auto
-\<comment> \<open> @{text init_class} promise \<close>
- have sheap2: "\<And>C. coll_init_class P ?i1 = Some C
- \<Longrightarrow> \<forall>C'. P \<turnstile> C \<preceq>\<^sup>* C' \<longrightarrow> (\<exists>sfs i. sheap \<sigma>1 C' = \<lfloor>(sfs, i)\<rfloor>)
- \<longrightarrow> classes_above P C' \<subseteq> cset" using sheap1 by auto
- have called: "\<And>C. coll_init_class P ?i1 = Some C
- \<Longrightarrow> ics_of (hd (frames_of \<sigma>1)) = Called [] \<Longrightarrow> classes_above P C \<subseteq> cset"
- proof -
- fix C assume cic: "coll_init_class P ?i1 = Some C" and
- ics: "ics_of (hd (frames_of \<sigma>1)) = Called []"
- then obtain sobj where "sh1 C = Some sobj" using vics1 f1'
- by(cases ?i1, auto simp: seeing_class_def split: if_split_asm)
- then show "classes_above P C \<subseteq> cset" using sheap1 by(cases sobj, simp)
- qed
- have init_class_prom1: "\<And>C. ics1 = Called [] \<or> ics1 = No_ics
- \<Longrightarrow> coll_init_class P ?i1 = Some C \<Longrightarrow> classes_above P C \<subseteq> cset"
- proof -
- fix C assume "ics1 = Called [] \<or> ics1 = No_ics" and cic: "coll_init_class P ?i1 = Some C"
- then have ics: "?ics1 = Called [] \<or> ?ics1 = No_ics" using f1' by simp
- then show "classes_above P C \<subseteq> cset" using cic
- proof(cases "?ics1 = Called []")
- case True then show ?thesis using cic called by simp
- next
- case False
- then have ics': "?ics1 = No_ics" using ics by simp
- then show ?thesis using cic
- proof(cases ?i1)
- case (New C1)
- then have "is_class P C1" using \<Phi>_pc1 wt1 meth1 by auto
- then have "P \<turnstile> C1 \<preceq>\<^sup>* Object" using wtp is_class_is_subcls
- by(auto simp: wf_jvm_prog_phi_def)
- then show ?thesis using New_collects[OF _ cbig1(2) nend1 _ ics' sheap2 sub'']
- f1' ics cic New by auto
- next
- case (Getstatic C1 F1 D1)
- then obtain t where mC1: "P \<turnstile> C1 has F1,Static:t in D1" and eq: "C = D1"
- using cic by (metis coll_init_class.simps(2) option.inject option.simps(3))
- then have "is_class P C" using has_field_is_class'[OF mC1] by simp
- then have "P \<turnstile> C \<preceq>\<^sup>* Object" using wtp is_class_is_subcls
- by(auto simp: wf_jvm_prog_phi_def)
- then show ?thesis using Getstatic_collects[OF _ cbig1(2) nend1 _ ics' _ sheap2 sub'']
- eq f1' Getstatic ics cic by fastforce
- next
- case (Putstatic C1 F1 D1)
- then obtain t where mC1: "P \<turnstile> C1 has F1,Static:t in D1" and eq: "C = D1"
- using cic by (metis coll_init_class.simps(3) option.inject option.simps(3))
- then have "is_class P C" using has_field_is_class'[OF mC1] by simp
- then have "P \<turnstile> C \<preceq>\<^sup>* Object" using wtp is_class_is_subcls
- by(auto simp: wf_jvm_prog_phi_def)
- then show ?thesis using Putstatic_collects[OF _ cbig1(2) nend1 _ ics' _ sheap2 sub'']
- eq f1' Putstatic ics cic by fastforce
- next
- case (Invokestatic C1 M1 n')
- then obtain Ts T m where mC: "P \<turnstile> C1 sees M1, Static : Ts\<rightarrow>T = m in C"
- using cic by(fastforce simp: seeing_class_def split: if_split_asm)
- then have "is_class P C" by(rule sees_method_is_class')
- then have Obj: "P \<turnstile> C \<preceq>\<^sup>* Object" using wtp is_class_is_subcls
- by(auto simp: wf_jvm_prog_phi_def)
- show ?thesis using Invokestatic_collects[OF _ cbig1(2) sub'' nend1 _ ics' mC sheap2]
- Obj mC f1' Invokestatic ics cic by auto
- qed(simp+)
- qed
- qed
-\<comment> \<open> Calling promise \<close>
- have Calling_prom1: "\<And>C' Cs'. ics1 = Calling C' Cs' \<Longrightarrow> classes_above P C' \<subseteq> cset"
- proof -
- fix C' Cs' assume ics: "ics1 = Calling C' Cs'"
- then have "is_class P C'" using vics1 by simp
- then have obj: "P \<turnstile> C' \<preceq>\<^sup>* Object" using wtp is_class_is_subcls
- by(auto simp: wf_jvm_prog_phi_def)
- have sheap3: "\<forall>C1. P \<turnstile> C' \<preceq>\<^sup>* C1 \<longrightarrow> (\<exists>sfs i. sheap \<sigma>1 C1 = \<lfloor>(sfs, i)\<rfloor>)
- \<longrightarrow> classes_above P C1 \<subseteq> cset" using sheap1 by auto
- show "classes_above P C' \<subseteq> cset"
- using Calling_collects[OF obj cbig1(2) nend1 _ sheap3 sub''] ics f1' by simp
- qed
- have in_naive: "(\<sigma>', cset') \<in> JVMNaiveCollectionSemantics.csmall_nstep P (xp1, h1, frs1, sh1) n1"
- and in_smart: "(\<sigma>', cset'') \<in> JVMSmartCollectionSemantics.csmall_nstep P (xp1, h1, frs1, sh1) n1"
- using \<sigma>1(3) \<sigma>1'(3) \<sigma>_eq[THEN sym] by simp+
- have sub2: "cset' \<subseteq> cset"
- by(rule Suc.hyps[OF wtp correct1 f1' heap1 sheap1 xcpts1 frames1 init_class_prom1
- Calling_prom1 in_naive in_smart sub'' Suc.prems(13)]) simp_all
- then show ?thesis using \<sigma>1(2) \<sigma>1'(2) sub1 sub2 by fastforce
- qed
-qed
-
-\<comment> \<open> ...which means over @{term cbig} \<close>
-lemma jvm_naive_to_smart_cbig:
-assumes
-\<comment> \<open> well-formedness \<close>
- wtp: "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P"
- and correct: "P,\<Phi> \<turnstile> (xp,h,frs,sh)\<surd>"
-\<comment> \<open> defs \<close>
- and f': "hd frs = (stk,loc,C',M',pc,ics)"
-\<comment> \<open> backward promises - will be collected/maintained prior \<close>
- and heap: "\<And>C fs. \<exists>a. h a = Some(C,fs) \<Longrightarrow> classes_above P C \<subseteq> cset"
- and sheap: "\<And>C sfs i. sh C = Some(sfs,i) \<Longrightarrow> classes_above P C \<subseteq> cset"
- and xcpts: "classes_above_xcpts P \<subseteq> cset"
- and frames: "classes_above_frames P frs \<subseteq> cset"
-\<comment> \<open> forward promises - will be collected after if not already \<close>
- and init_class_prom: "\<And>C. ics = Called [] \<or> ics = No_ics
- \<Longrightarrow> coll_init_class P (instrs_of P C' M' ! pc) = Some C \<Longrightarrow> classes_above P C \<subseteq> cset"
- and Calling_prom: "\<And>C' Cs'. ics = Calling C' Cs' \<Longrightarrow> classes_above P C' \<subseteq> cset"
-\<comment> \<open> collections \<close>
- and n: "(\<sigma>', cset\<^sub>n) \<in> JVMNaiveCollectionSemantics.cbig P (xp,h,frs,sh)"
- and s: "(\<sigma>', cset\<^sub>s) \<in> JVMSmartCollectionSemantics.cbig P (xp,h,frs,sh)"
- and smart: "cset\<^sub>s \<subseteq> cset"
-shows "cset\<^sub>n \<subseteq> cset"
-proof -
- let ?\<sigma> = "(xp,h,frs,sh)"
- have nend: "\<sigma>' \<in> JVMendset" using n by(simp add: JVMNaiveCollectionSemantics.cbig_def)
- obtain n where n': "(\<sigma>', cset\<^sub>n) \<in> JVMNaiveCollectionSemantics.csmall_nstep P ?\<sigma> n" "\<sigma>' \<in> JVMendset"
- using JVMNaiveCollectionSemantics.cbig_def2 n by auto
- obtain s where s': "(\<sigma>', cset\<^sub>s) \<in> JVMSmartCollectionSemantics.csmall_nstep P ?\<sigma> s" "\<sigma>' \<in> JVMendset"
- using JVMSmartCollectionSemantics.cbig_def2 s by auto
- have "n=s" using jvm_naive_to_smart_csmall_nstep_last_eq[OF n n'(1) s'(1)] by simp
- then have sn: "(\<sigma>', cset\<^sub>s) \<in> JVMSmartCollectionSemantics.csmall_nstep P ?\<sigma> n" using s'(1) by simp
- then show ?thesis
- using jvm_naive_to_smart_csmall_nstep[OF assms(1-9) n'(1) sn assms(12) nend] by fast
-qed
-
-\<comment> \<open>...thus naive @{text "\<subseteq>"} smart over the out function, since all conditions will be met - and promises
- kept - by the defined starting point \<close>
-lemma jvm_naive_to_smart_collection:
-assumes naive: "(\<sigma>',cset\<^sub>n) \<in> jvm_naive_out P t" and smart: "(\<sigma>',cset\<^sub>s) \<in> jvm_smart_out P t"
- and P: "P \<in> jvm_progs" and t: "t \<in> jvm_tests"
-shows "cset\<^sub>n \<subseteq> cset\<^sub>s"
-proof -
- let ?P = "jvm_make_test_prog P t"
- let ?\<sigma> = "start_state (t#P)"
- let ?i = "instrs_of ?P Start start_m ! 0" and ?ics = No_ics
- obtain xp h frs sh where
- [simp]: "?\<sigma> = (xp,h,frs,sh)" and
- [simp]: "h = start_heap (t#P)" and
- [simp]: "frs = [([], [], Start, start_m, 0, No_ics)]" and
- [simp]: "sh = start_sheap"
- by(clarsimp simp: start_state_def)
-
- from P t have nS: "\<not> is_class (t # P) Start"
- by(simp add: is_class_def class_def Start_def Test_def)
- from P have nT: "\<not> is_class P Test" by simp
- from P t obtain m where tPm: "t # P \<turnstile> (fst t) sees main, Static : []\<rightarrow>Void = m in (fst t)"
- by auto
- have nclinit: "main \<noteq> clinit" by(simp add: main_def clinit_def)
- have Objp: "\<And>b' Ts' T' m' D'.
- t#P \<turnstile> Object sees start_m, b' : Ts'\<rightarrow>T' = m' in D' \<Longrightarrow> b' = Static \<and> Ts' = [] \<and> T' = Void"
- proof -
- fix b' Ts' T' m' D'
- assume mObj: "t#P \<turnstile> Object sees start_m, b' : Ts'\<rightarrow>T' = m' in D'"
- from P have ot_nsub: "\<not> P \<turnstile> Object \<preceq>\<^sup>* Test"
- by(clarsimp simp: wf_jvm_prog_def wf_jvm_prog_phi_def)
- from class_add_sees_method_rev[OF _ ot_nsub] mObj t
- have "P \<turnstile> Object sees start_m, b' : Ts'\<rightarrow>T' = m' in D'" by(cases t, auto)
- with P jvm_progs_def show "b' = Static \<and> Ts' = [] \<and> T' = Void" by blast
- qed
- from P t obtain \<Phi> where wtp0: "wf_jvm_prog\<^bsub>\<Phi>\<^esub> (t#P)" by(auto simp: wf_jvm_prog_def)
- let ?\<Phi>' = "\<lambda>C f. if C = Start \<and> (f = start_m \<or> f = clinit) then start_\<phi>\<^sub>m else \<Phi> C f"
- from wtp0 have wtp: "wf_jvm_prog\<^bsub>?\<Phi>'\<^esub> ?P"
- proof -
- note wtp0 nS tPm nclinit
- moreover obtain "\<And>C. C \<noteq> Start \<Longrightarrow> ?\<Phi>' C = \<Phi> C" "?\<Phi>' Start start_m = start_\<phi>\<^sub>m"
- "?\<Phi>' Start clinit = start_\<phi>\<^sub>m" by simp
- moreover note Objp
- ultimately show ?thesis by(rule start_prog_wf_jvm_prog_phi)
- qed
- have cic: "coll_init_class ?P ?i = Some Test"
- proof -
- from wtp0 obtain wf_md where wf: "wf_prog wf_md (t#P)"
- by(clarsimp dest!: wt_jvm_progD)
- with start_prog_start_m_instrs t have i: "?i = Invokestatic Test main 0" by simp
- from jvm_make_test_prog_sees_Test_main[OF P t] obtain m where
- "?P \<turnstile> Test sees main, Static : []\<rightarrow>Void = m in Test" by fast
- with t have "seeing_class (jvm_make_test_prog P t) Test main = \<lfloor>Test\<rfloor>"
- by(cases m, fastforce simp: seeing_class_def)
- with i show ?thesis by simp
- qed
-\<comment> \<open> well-formedness \<close>
- note wtp
- moreover have correct: "?P,?\<Phi>' \<turnstile> (xp,h,frs,sh)\<surd>"
- proof -
- note wtp0 nS tPm nclinit
- moreover have "?\<Phi>' Start start_m = start_\<phi>\<^sub>m" by simp
- ultimately have "?P,?\<Phi>' \<turnstile> ?\<sigma>\<surd>" by(rule BV_correct_initial)
- then show ?thesis by simp
- qed
-\<comment> \<open> defs \<close>
- moreover have "hd frs = ([], [], Start, start_m, 0, No_ics)" by simp
-\<comment> \<open> backward promises \<close>
- moreover from jvm_smart_out_classes_above_start_heap[OF smart _ P t]
- have heap: "\<And>C fs. \<exists>a. h a = Some(C,fs) \<Longrightarrow> classes_above ?P C \<subseteq> cset\<^sub>s" by auto
- moreover from jvm_smart_out_classes_above_start_sheap[OF smart]
- have sheap: "\<And>C sfs i. sh C = Some(sfs,i) \<Longrightarrow> classes_above ?P C \<subseteq> cset\<^sub>s" by simp
- moreover from jvm_smart_out_classes_above_xcpts[OF smart P t]
- have xcpts: "classes_above_xcpts ?P \<subseteq> cset\<^sub>s" by simp
- moreover from jvm_smart_out_classes_above_frames[OF smart]
- have frames: "classes_above_frames ?P frs \<subseteq> cset\<^sub>s" by simp
-\<comment> \<open> forward promises - will be collected after if not already \<close>
- moreover from jvm_smart_out_classes_above_Test[OF smart P t] cic
- have init_class_prom: "\<And>C. ?ics = Called [] \<or> ?ics = No_ics
- \<Longrightarrow> coll_init_class ?P ?i = Some C \<Longrightarrow> classes_above ?P C \<subseteq> cset\<^sub>s" by simp
- moreover have "\<And>C' Cs'. ?ics = Calling C' Cs' \<Longrightarrow> classes_above ?P C' \<subseteq> cset\<^sub>s" by simp
-\<comment> \<open> collections \<close>
- moreover from naive
- have n: "(\<sigma>', cset\<^sub>n) \<in> JVMNaiveCollectionSemantics.cbig ?P (xp,h,frs,sh)" by simp
- moreover from smart obtain cset\<^sub>s' where
- s: "(\<sigma>', cset\<^sub>s') \<in> JVMSmartCollectionSemantics.cbig ?P (xp,h,frs,sh)" and
- "cset\<^sub>s' \<subseteq> cset\<^sub>s"
- by clarsimp
- ultimately show "cset\<^sub>n \<subseteq> cset\<^sub>s" by(rule jvm_naive_to_smart_cbig; simp)
-qed
-
-
-(******************************************)
-subsubsection \<open> Proving smart @{text "\<subseteq>"} naive \<close>
-
-text "We prove that @{term jvm_naive} collects everything @{term jvm_smart}
- does. Combined with the other direction, this shows that the naive and smart
- algorithms collect the same set of classes."
-
-lemma jvm_smart_to_naive_exec_collect:
- "JVMexec_scollect P \<sigma> \<subseteq> JVMexec_ncollect P \<sigma>"
-proof -
- obtain xp h frs sh where \<sigma>: "\<sigma>=(xp,h,frs,sh)" by(cases \<sigma>)
- then show ?thesis
- proof(cases "\<exists>x. xp = Some x \<or> frs = []")
- case False
- then obtain stk loc C M pc ics frs'
- where none: "xp = None" and frs: "frs=(stk,loc,C,M,pc,ics)#frs'"
- by(cases xp, auto, cases frs, auto)
- have instr_case: "ics = Called [] \<or> ics = No_ics \<Longrightarrow> ?thesis"
- proof -
- assume ics: "ics = Called [] \<or> ics = No_ics"
- then show ?thesis using \<sigma> none frs
- proof(cases "curr_instr P (stk,loc,C,M,pc,ics)") qed(auto split: if_split_asm)
- qed
- then show ?thesis using \<sigma> none frs
- proof(cases ics)
- case(Called Cs) then show ?thesis using instr_case \<sigma> none frs by(cases Cs, auto)
- qed(auto)
- qed(auto)
-qed
-
-lemma jvm_smart_to_naive_csmall:
-assumes "(\<sigma>', cset\<^sub>n) \<in> JVMNaiveCollectionSemantics.csmall P \<sigma>"
- and "(\<sigma>', cset\<^sub>s) \<in> JVMSmartCollectionSemantics.csmall P \<sigma>"
-shows "cset\<^sub>s \<subseteq> cset\<^sub>n"
-using jvm_smart_to_naive_exec_collect assms
- by(auto simp: JVMNaiveCollectionSemantics.csmall_def
- JVMSmartCollectionSemantics.csmall_def)
-
-lemma jvm_smart_to_naive_csmall_nstep:
-"\<lbrakk> (\<sigma>', cset\<^sub>n) \<in> JVMNaiveCollectionSemantics.csmall_nstep P \<sigma> n;
- (\<sigma>', cset\<^sub>s) \<in> JVMSmartCollectionSemantics.csmall_nstep P \<sigma> n \<rbrakk>
- \<Longrightarrow> cset\<^sub>s \<subseteq> cset\<^sub>n"
-proof(induct n arbitrary: \<sigma> \<sigma>' cset\<^sub>n cset\<^sub>s)
- case (Suc n')
- obtain \<sigma>1 cset1 cset' where \<sigma>1: "(\<sigma>1, cset1) \<in> JVMNaiveCollectionSemantics.csmall P \<sigma>"
- "cset\<^sub>n = cset1 \<union> cset'" "(\<sigma>', cset') \<in> JVMNaiveCollectionSemantics.csmall_nstep P \<sigma>1 n'"
- using JVMNaiveCollectionSemantics.csmall_nstep_SucD [OF Suc.prems(1)] by clarsimp+
- obtain \<sigma>1' cset1' cset'' where \<sigma>1': "(\<sigma>1', cset1') \<in> JVMSmartCollectionSemantics.csmall P \<sigma>"
- "cset\<^sub>s = cset1' \<union> cset''" "(\<sigma>', cset'') \<in> JVMSmartCollectionSemantics.csmall_nstep P \<sigma>1' n'"
- using JVMSmartCollectionSemantics.csmall_nstep_SucD [OF Suc.prems(2)] by clarsimp+
- have \<sigma>_eq: "\<sigma>1 = \<sigma>1'" using \<sigma>1(1) \<sigma>1'(1) by(simp add: JVMNaiveCollectionSemantics.csmall_def
- JVMSmartCollectionSemantics.csmall_def)
- then have sub1: "cset1' \<subseteq> cset1" using \<sigma>1(1) \<sigma>1'(1) jvm_smart_to_naive_csmall by blast
- have sub2: "cset'' \<subseteq> cset'" using \<sigma>1(3) \<sigma>1'(3) \<sigma>_eq Suc.hyps by blast
- then show ?case using \<sigma>1(2) \<sigma>1'(2) sub1 sub2 by blast
-qed(simp)
-
-lemma jvm_smart_to_naive_cbig:
-assumes n: "(\<sigma>', cset\<^sub>n) \<in> JVMNaiveCollectionSemantics.cbig P \<sigma>"
- and s: "(\<sigma>', cset\<^sub>s) \<in> JVMSmartCollectionSemantics.cbig P \<sigma>"
-shows "cset\<^sub>s \<subseteq> cset\<^sub>n"
-proof -
- obtain n where n': "(\<sigma>', cset\<^sub>n) \<in> JVMNaiveCollectionSemantics.csmall_nstep P \<sigma> n" "\<sigma>' \<in> JVMendset"
- using JVMNaiveCollectionSemantics.cbig_def2 n by auto
- obtain s where s': "(\<sigma>', cset\<^sub>s) \<in> JVMSmartCollectionSemantics.csmall_nstep P \<sigma> s" "\<sigma>' \<in> JVMendset"
- using JVMSmartCollectionSemantics.cbig_def2 s by auto
- have "n=s" using jvm_naive_to_smart_csmall_nstep_last_eq[OF n n'(1) s'(1)] by simp
- then show ?thesis using jvm_smart_to_naive_csmall_nstep n'(1) s'(1) by blast
-qed
-
-lemma jvm_smart_to_naive_collection:
-assumes naive: "(\<sigma>',cset\<^sub>n) \<in> jvm_naive_out P t" and smart: "(\<sigma>',cset\<^sub>s) \<in> jvm_smart_out P t"
- and "P \<in> jvm_progs" and "t \<in> jvm_tests"
-shows "cset\<^sub>s \<subseteq> cset\<^sub>n"
-proof -
- have nend: "start_state (t#P) \<notin> JVMendset" by(simp add: JVMendset_def start_state_def)
- from naive obtain n where
- nstep: "(\<sigma>', cset\<^sub>n) \<in> JVMNaiveCollectionSemantics.csmall_nstep
- (jvm_make_test_prog P t) (start_state (t#P)) n"
- by(auto dest!: JVMNaiveCollectionSemantics.cbigD)
- with nend naive obtain n' where n': "n = Suc n'"
- by(cases n; simp add: JVMNaiveCollectionSemantics.cbig_def)
- from start_prog_classes_above_Start
- have "classes_above_frames (jvm_make_test_prog P t) (frames_of (start_state (t#P))) = {Object,Start}"
- by(simp add: start_state_def)
- with nstep n'
- have "jvm_smart_collect_start (jvm_make_test_prog P t) \<subseteq> cset\<^sub>n"
- by(auto simp: start_state_def JVMNaiveCollectionSemantics.csmall_def
- dest!: JVMNaiveCollectionSemantics.csmall_nstep_SucD
- simp del: JVMNaiveCollectionSemantics.csmall_nstep_Rec)
- with jvm_smart_to_naive_cbig[where P="jvm_make_test_prog P t" and \<sigma>="start_state (t#P)" and \<sigma>'=\<sigma>']
- jvm_smart_collect_start_make_test_prog assms show ?thesis by auto
-qed
-
-(**************************************************)
-subsubsection "Safety of the smart algorithm"
-
-text "Having proved containment in both directions, we get naive = smart:"
-lemma jvm_naive_eq_smart_collection:
-assumes naive: "(\<sigma>',cset\<^sub>n) \<in> jvm_naive_out P t" and smart: "(\<sigma>',cset\<^sub>s) \<in> jvm_smart_out P t"
- and "P \<in> jvm_progs" and "t \<in> jvm_tests"
-shows "cset\<^sub>n = cset\<^sub>s"
-using jvm_naive_to_smart_collection[OF assms] jvm_smart_to_naive_collection[OF assms] by simp
-
-text "Thus, since the RTS algorithm based on @{term ncollect} is existence safe,
- the algorithm based on @{term scollect} is as well."
-theorem jvm_smart_existence_safe:
-assumes P: "P \<in> jvm_progs" and P': "P' \<in> jvm_progs" and t: "t \<in> jvm_tests"
- and out: "o1 \<in> jvm_smart_out P t" and dss: "jvm_deselect P o1 P'"
-shows "\<exists>o2 \<in> jvm_smart_out P' t. o1 = o2"
-proof -
- obtain \<sigma>' cset\<^sub>s where o1[simp]: "o1=(\<sigma>',cset\<^sub>s)" by(cases o1)
- with jvm_naive_iff_smart out obtain cset\<^sub>n where n: "(\<sigma>',cset\<^sub>n) \<in> jvm_naive_out P t" by blast
-
- from jvm_naive_eq_smart_collection[OF n _ P t] out have eq: "cset\<^sub>n = cset\<^sub>s" by simp
- with jvm_naive_existence_safe[OF P P' t n] dss have n': "(\<sigma>',cset\<^sub>n) \<in> jvm_naive_out P' t" by simp
- with jvm_naive_iff_smart obtain cset\<^sub>s' where s': "(\<sigma>', cset\<^sub>s') \<in> jvm_smart_out P' t" by blast
-
- from jvm_naive_eq_smart_collection[OF n' s' P' t] eq have "cset\<^sub>s = cset\<^sub>s'" by simp
- then show ?thesis using s' by simp
-qed
-
-text "...thus @{term JVMSmartCollection} is an instance of @{term CollectionBasedRTS}:"
-interpretation JVMSmartCollectionRTS :
- CollectionBasedRTS "(=)" jvm_deselect jvm_progs jvm_tests
- JVMendset JVMcombine JVMcollect_id JVMsmall JVMSmartCollect jvm_smart_out
- jvm_make_test_prog jvm_smart_collect_start
- by unfold_locales (rule jvm_smart_existence_safe, auto simp: start_state_def)
-
+(* File: RTS/JVM_RTS/JVMCollectionBasedRTS.thy *)
+(* Author: Susannah Mansky, UIUC 2020 *)
+(* Proof of safety of certain collection based RTS algorithms for Jinja JVM *)
+
+section "Instantiating @{term CollectionBasedRTS} with Jinja JVM"
+
+theory JVMCollectionBasedRTS
+imports "../Common/CollectionBasedRTS" JVMCollectionSemantics
+ JinjaDCI.BVSpecTypeSafe "../JinjaSuppl/JVMExecStepInductive"
+
+begin
+
+lemma eq_equiv[simp]: "equiv UNIV {(x, y). x = y}"
+by(simp add: equiv_def refl_on_def sym_def trans_def)
+
+(**********************************************)
+subsection \<open> Some @{text "classes_above"} lemmas \<close>
+(* here because they require ClassAdd/StartProg *)
+
+lemma start_prog_classes_above_Start:
+ "classes_above (start_prog P C M) Start = {Object,Start}"
+using start_prog_Start_super[of C M P] subcls1_confluent by auto
+
+lemma class_add_classes_above:
+assumes ns: "\<not> is_class P C" and "\<not>P \<turnstile> D \<preceq>\<^sup>* C"
+shows "classes_above (class_add P (C, cdec)) D = classes_above P D"
+using assms by(auto intro: class_add_subcls class_add_subcls_rev)
+
+lemma class_add_classes_above_xcpts:
+assumes ns: "\<not> is_class P C"
+ and ncp: "\<And>D. D \<in> sys_xcpts \<Longrightarrow> \<not>P \<turnstile> D \<preceq>\<^sup>* C"
+shows "classes_above_xcpts (class_add P (C, cdec)) = classes_above_xcpts P"
+using assms class_add_classes_above by simp
+
+(*********)
+subsection "JVM next-step lemmas for initialization calling"
+
+lemma JVM_New_next_step:
+assumes step: "\<sigma>' \<in> JVMsmall P \<sigma>"
+ and nend: "\<sigma> \<notin> JVMendset"
+ and curr: "curr_instr P (hd(frames_of \<sigma>)) = New C"
+ and nDone: "\<not>(\<exists>sfs i. sheap \<sigma> C = Some(sfs,i) \<and> i = Done)"
+ and ics: "ics_of(hd(frames_of \<sigma>)) = No_ics"
+shows "ics_of (hd(frames_of \<sigma>')) = Calling C [] \<and> sheap \<sigma> = sheap \<sigma>' \<and> \<sigma>' \<notin> JVMendset"
+proof -
+ obtain xp h frs sh where \<sigma>: "\<sigma>=(xp,h,frs,sh)" by(cases \<sigma>)
+ then obtain f1 frs1 where frs: "frs=f1#frs1" using nend by(cases frs, simp_all add: JVMendset_def)
+ then obtain stk loc C' M' pc ics where f1:"f1=(stk,loc,C',M',pc,ics)" by(cases f1)
+ have xp: "xp = None" using \<sigma> nend by(simp add: JVMendset_def)
+ obtain xp' h' frs' sh' where \<sigma>': "\<sigma>'=(xp',h',frs',sh')" by(cases \<sigma>')
+ have "ics_of (hd frs') = Calling C [] \<and> sh = sh' \<and> frs' \<noteq> [] \<and> xp' = None"
+ proof(cases "sh C")
+ case None then show ?thesis using \<sigma>' xp f1 frs \<sigma> assms by auto
+ next
+ case (Some a)
+ then obtain sfs i where a: "a=(sfs,i)" by(cases a)
+ then have nDone': "i \<noteq> Done" using nDone Some f1 frs \<sigma> by simp
+ show ?thesis using a Some \<sigma>' xp f1 frs \<sigma> assms by(auto split: init_state.splits)
+ qed
+ then show ?thesis using ics \<sigma> \<sigma>' by(cases frs', auto simp: JVMendset_def)
+qed
+
+lemma JVM_Getstatic_next_step:
+assumes step: "\<sigma>' \<in> JVMsmall P \<sigma>"
+ and nend: "\<sigma> \<notin> JVMendset"
+ and curr: "curr_instr P (hd(frames_of \<sigma>)) = Getstatic C F D"
+ and fC: "P \<turnstile> C has F,Static:t in D"
+ and nDone: "\<not>(\<exists>sfs i. sheap \<sigma> D = Some(sfs,i) \<and> i = Done)"
+ and ics: "ics_of(hd(frames_of \<sigma>)) = No_ics"
+shows "ics_of (hd(frames_of \<sigma>')) = Calling D [] \<and> sheap \<sigma> = sheap \<sigma>' \<and> \<sigma>' \<notin> JVMendset"
+proof -
+ obtain xp h frs sh where \<sigma>: "\<sigma>=(xp,h,frs,sh)" by(cases \<sigma>)
+ then obtain f1 frs1 where frs: "frs=f1#frs1" using nend by(cases frs, simp_all add: JVMendset_def)
+ then obtain stk loc C' M' pc ics where f1:"f1=(stk,loc,C',M',pc,ics)" by(cases f1)
+ have xp: "xp = None" using \<sigma> nend by(simp add: JVMendset_def)
+ obtain xp' h' frs' sh' where \<sigma>': "\<sigma>'=(xp',h',frs',sh')" by(cases \<sigma>')
+ have ex': "\<exists>t b. P \<turnstile> C has F,b:t in D" using fC by auto
+ have field: "\<exists>t. field P D F = (D,Static,t)"
+ using fC field_def2 has_field_idemp has_field_sees by blast
+ have nCalled': "\<forall>Cs. ics \<noteq> Called Cs" using ics f1 frs \<sigma> by simp
+ have "ics_of (hd frs') = Calling D [] \<and> sh = sh' \<and> frs' \<noteq> [] \<and> xp' = None"
+ proof(cases "sh D")
+ case None then show ?thesis using field ex' \<sigma>' xp f1 frs \<sigma> assms by auto
+ next
+ case (Some a)
+ then obtain sfs i where a: "a=(sfs,i)" by(cases a)
+ show ?thesis using field ex' a Some \<sigma>' xp f1 frs \<sigma> assms by(auto split: init_state.splits)
+ qed
+ then show ?thesis using ics \<sigma> \<sigma>' by(auto simp: JVMendset_def)
+qed
+
+lemma JVM_Putstatic_next_step:
+assumes step: "\<sigma>' \<in> JVMsmall P \<sigma>"
+ and nend: "\<sigma> \<notin> JVMendset"
+ and curr: "curr_instr P (hd(frames_of \<sigma>)) = Putstatic C F D"
+ and fC: "P \<turnstile> C has F,Static:t in D"
+ and nDone: "\<not>(\<exists>sfs i. sheap \<sigma> D = Some(sfs,i) \<and> i = Done)"
+ and ics: "ics_of(hd(frames_of \<sigma>)) = No_ics"
+shows "ics_of (hd(frames_of \<sigma>')) = Calling D [] \<and> sheap \<sigma> = sheap \<sigma>' \<and> \<sigma>' \<notin> JVMendset"
+proof -
+ obtain xp h frs sh where \<sigma>: "\<sigma>=(xp,h,frs,sh)" by(cases \<sigma>)
+ then obtain f1 frs1 where frs: "frs=f1#frs1" using nend by(cases frs, simp_all add: JVMendset_def)
+ then obtain stk loc C' M' pc ics where f1:"f1=(stk,loc,C',M',pc,ics)" by(cases f1)
+ have xp: "xp = None" using \<sigma> nend by(simp add: JVMendset_def)
+ obtain xp' h' frs' sh' where \<sigma>': "\<sigma>'=(xp',h',frs',sh')" by(cases \<sigma>')
+ have ex': "\<exists>t b. P \<turnstile> C has F,b:t in D" using fC by auto
+ have field: "field P D F = (D,Static,t)"
+ using fC field_def2 has_field_idemp has_field_sees by blast
+ have ics': "ics_of (hd frs') = Calling D [] \<and> sh = sh' \<and> frs' \<noteq> [] \<and> xp' = None"
+ proof(cases "sh D")
+ case None then show ?thesis using field ex' \<sigma>' xp f1 frs \<sigma> assms by auto
+ next
+ case (Some a)
+ then obtain sfs i where a: "a=(sfs,i)" by(cases a)
+ show ?thesis using field ex' a Some \<sigma>' xp f1 frs \<sigma> assms by(auto split: init_state.splits)
+ qed
+ then show ?thesis using ics \<sigma> \<sigma>' by(auto simp: JVMendset_def)
+qed
+
+lemma JVM_Invokestatic_next_step:
+assumes step: "\<sigma>' \<in> JVMsmall P \<sigma>"
+ and nend: "\<sigma> \<notin> JVMendset"
+ and curr: "curr_instr P (hd(frames_of \<sigma>)) = Invokestatic C M n"
+ and mC: "P \<turnstile> C sees M,Static:Ts \<rightarrow> T = m in D"
+ and nDone: "\<not>(\<exists>sfs i. sheap \<sigma> D = Some(sfs,i) \<and> i = Done)"
+ and ics: "ics_of(hd(frames_of \<sigma>)) = No_ics"
+shows "ics_of (hd(frames_of \<sigma>')) = Calling D [] \<and> sheap \<sigma> = sheap \<sigma>' \<and> \<sigma>' \<notin> JVMendset"
+proof -
+ obtain xp h frs sh where \<sigma>: "\<sigma>=(xp,h,frs,sh)" by(cases \<sigma>)
+ then obtain f1 frs1 where frs: "frs=f1#frs1" using nend by(cases frs, simp_all add: JVMendset_def)
+ then obtain stk loc C' M' pc ics where f1:"f1=(stk,loc,C',M',pc,ics)" by(cases f1)
+ have xp: "xp = None" using \<sigma> nend by(simp add: JVMendset_def)
+ obtain xp' h' frs' sh' where \<sigma>': "\<sigma>'=(xp',h',frs',sh')" by(cases \<sigma>')
+ have ex': "\<exists>Ts T m D b. P \<turnstile> C sees M,b:Ts \<rightarrow> T = m in D" using mC by fastforce
+ have method: "\<exists>m. method P C M = (D,Static,m)" using mC by(cases m, auto)
+ have ics': "ics_of (hd frs') = Calling D [] \<and> sh = sh' \<and> frs' \<noteq> [] \<and> xp' = None"
+ proof(cases "sh D")
+ case None then show ?thesis using method ex' \<sigma>' xp f1 frs \<sigma> assms by auto
+ next
+ case (Some a)
+ then obtain sfs i where a: "a=(sfs,i)" by(cases a)
+ then have nDone': "i \<noteq> Done" using nDone Some f1 frs \<sigma> by simp
+ show ?thesis using method ex' a Some \<sigma>' xp f1 frs \<sigma> assms by(auto split: init_state.splits)
+ qed
+ then show ?thesis using ics \<sigma> \<sigma>' by(auto simp: JVMendset_def)
+qed
+
+(***********************************************)
+subsection "Definitions"
+
+definition main :: "string" where "main = ''main''"
+definition Test :: "string" where "Test = ''Test''"
+definition test_oracle :: "string" where "test_oracle = ''oracle''"
+
+type_synonym jvm_class = "jvm_method cdecl"
+type_synonym jvm_prog_out = "jvm_state \<times> cname set"
+
+text "A deselection algorithm based on classes that have changed from
+ @{text P1} to @{text P2}:"
+primrec jvm_deselect :: "jvm_prog \<Rightarrow> jvm_prog_out \<Rightarrow> jvm_prog \<Rightarrow> bool" where
+"jvm_deselect P1 (\<sigma>, cset) P2 = (cset \<inter> (classes_changed P1 P2) = {})"
+
+definition jvm_progs :: "jvm_prog set" where
+"jvm_progs \<equiv> {P. wf_jvm_prog P \<and> \<not>is_class P Start \<and> \<not>is_class P Test
+ \<and> (\<forall>b' Ts' T' m' D'. P \<turnstile> Object sees start_m, b' : Ts'\<rightarrow>T' = m' in D'
+ \<longrightarrow> b' = Static \<and> Ts' = [] \<and> T' = Void) }"
+
+definition jvm_tests :: "jvm_class set" where
+"jvm_tests = {t. fst t = Test
+ \<and> (\<forall>P \<in> jvm_progs. wf_jvm_prog (t#P) \<and> (\<exists>m. t#P \<turnstile> Test sees main,Static: [] \<rightarrow> Void = m in Test)) }"
+
+abbreviation jvm_make_test_prog :: "jvm_prog \<Rightarrow> jvm_class \<Rightarrow> jvm_prog" where
+"jvm_make_test_prog P t \<equiv> start_prog (t#P) (fst t) main"
+
+declare jvm_progs_def [simp]
+declare jvm_tests_def [simp]
+
+(*****************************************************************************************)
+subsection "Definition lemmas"
+
+lemma jvm_progs_tests_nStart:
+assumes P: "P \<in> jvm_progs" and t: "t \<in> jvm_tests"
+shows "\<not>is_class (t#P) Start"
+using assms by(simp add: is_class_def class_def Start_def Test_def)
+
+lemma jvm_make_test_prog_classes_above_xcpts:
+assumes P: "P \<in> jvm_progs" and t: "t \<in> jvm_tests"
+shows "classes_above_xcpts (jvm_make_test_prog P t) = classes_above_xcpts P"
+proof -
+ have nS: "\<not>is_class (t#P) Start" by(rule jvm_progs_tests_nStart[OF P t])
+ from P have nT: "\<not>is_class P Test" by simp
+ from P t have "wf_syscls (t#P) \<and> wf_syscls P"
+ by(simp add: wf_jvm_prog_def wf_jvm_prog_phi_def wf_prog_def)
+
+ then have [simp]: "\<And>D. D \<in> sys_xcpts \<Longrightarrow> is_class (t#P) D \<and> is_class P D"
+ by(cases t, auto simp: wf_syscls_def is_class_def class_def dest!: weak_map_of_SomeI)
+ from wf_nclass_nsub[OF _ _ nS] P t have nspS: "\<And>D. D \<in> sys_xcpts \<Longrightarrow> \<not>(t#P) \<turnstile> D \<preceq>\<^sup>* Start"
+ by(auto simp: wf_jvm_prog_def wf_jvm_prog_phi_def)
+ from wf_nclass_nsub[OF _ _ nT] P have nspT: "\<And>D. D \<in> sys_xcpts \<Longrightarrow> \<not>P \<turnstile> D \<preceq>\<^sup>* Test"
+ by(auto simp: wf_jvm_prog_def wf_jvm_prog_phi_def)
+
+ from class_add_classes_above_xcpts[where P="t#P" and C=Start, OF nS nspS]
+ have "classes_above_xcpts (jvm_make_test_prog P t) = classes_above_xcpts (t#P)" by simp
+ also from class_add_classes_above_xcpts[where P=P and C=Test, OF nT nspT] t
+ have "\<dots> = classes_above_xcpts P" by(cases t, simp)
+ finally show ?thesis by simp
+qed
+
+lemma jvm_make_test_prog_sees_Test_main:
+assumes P: "P \<in> jvm_progs" and t: "t \<in> jvm_tests"
+shows "\<exists>m. jvm_make_test_prog P t \<turnstile> Test sees main, Static : []\<rightarrow>Void = m in Test"
+proof -
+ let ?P = "jvm_make_test_prog P t"
+ from P t obtain m where
+ meth: "t#P \<turnstile> Test sees main,Static:[] \<rightarrow> Void = m in Test" and
+ nstart: "\<not> is_class (t # P) Start"
+ by(auto simp: is_class_def class_def Start_def Test_def)
+ from class_add_sees_method[OF meth nstart] show ?thesis by fastforce
+qed
+
+(************************************************)
+subsection "Naive RTS algorithm"
+
+subsubsection "Definitions"
+
+fun jvm_naive_out :: "jvm_prog \<Rightarrow> jvm_class \<Rightarrow> jvm_prog_out set" where
+"jvm_naive_out P t = JVMNaiveCollectionSemantics.cbig (jvm_make_test_prog P t) (start_state (t#P))"
+
+abbreviation jvm_naive_collect_start :: "jvm_prog \<Rightarrow> cname set" where
+"jvm_naive_collect_start P \<equiv> {}"
+
+lemma jvm_naive_out_xcpts_collected:
+assumes "o1 \<in> jvm_naive_out P t"
+shows "classes_above_xcpts (start_prog (t # P) (fst t) main) \<subseteq> snd o1"
+using assms
+proof -
+ obtain \<sigma>' coll' where "o1 = (\<sigma>', coll')" and
+ cbig: "(\<sigma>', coll') \<in> JVMNaiveCollectionSemantics.cbig (start_prog (t#P) (fst t) main) (start_state (t#P))"
+ using assms by(cases o1, simp)
+ with JVMNaiveCollectionSemantics.cbig_stepD[OF cbig start_state_nend]
+ show ?thesis by(auto simp: JVMNaiveCollectionSemantics.csmall_def start_state_def)
+qed
+
+(***********************************************************)
+subsubsection "Naive algorithm correctness"
+
+text "We start with correctness over @{term exec_instr}, then all the
+ functions/pieces that are used by naive @{term csmall} (that is, pieces
+ used by @{term exec} - such as which frames are used based on @{term ics}
+ - and all functions used by the collection function). We then prove that
+ @{term csmall} is existence safe, extend this result to @{term cbig}, and
+ finally prove the @{term existence_safe} statement over the locale pieces."
+
+\<comment> \<open> if collected classes unchanged, @{term exec_instr} unchanged \<close>
+lemma ncollect_exec_instr:
+assumes "JVMinstr_ncollect P i h stk \<inter> classes_changed P P' = {}"
+ and above_C: "classes_above P C \<inter> classes_changed P P' = {}"
+ and ics: "ics = Called [] \<or> ics = No_ics"
+ and i: "i = instrs_of P C M ! pc"
+shows "exec_instr i P h stk loc C M pc ics frs sh = exec_instr i P' h stk loc C M pc ics frs sh"
+using assms proof(cases i)
+ case (New C1) then show ?thesis using assms classes_above_blank[of C1 P P']
+ by(auto split: init_state.splits option.splits)
+next
+ case (Getfield F1 C1) show ?thesis
+ proof(cases "hd stk = Null")
+ case True then show ?thesis using Getfield assms by simp
+ next
+ case False
+ let ?D = "(cname_of h (the_Addr (hd stk)))"
+ have D: "classes_above P ?D \<inter> classes_changed P P' = {}"
+ using False Getfield assms by simp
+ show ?thesis
+ proof(cases "\<exists>b t. P \<turnstile> ?D has F1,b:t in C1")
+ case True
+ then obtain b1 t1 where "P \<turnstile> ?D has F1,b1:t1 in C1" by auto
+ then have has: "P' \<turnstile> ?D has F1,b1:t1 in C1"
+ using Getfield assms classes_above_has_field[OF D] by auto
+ have "P \<turnstile> ?D \<preceq>\<^sup>* C1" using has_field_decl_above True by auto
+ then have "classes_above P C1 \<subseteq> classes_above P ?D" by(rule classes_above_subcls_subset)
+ then have C1: "classes_above P C1 \<inter> classes_changed P P' = {}" using D by auto
+ then show ?thesis using has True Getfield assms classes_above_field[of C1 P P' F1]
+ by(cases "field P' C1 F1", cases "the (h (the_Addr (hd stk)))", auto)
+ next
+ case nex: False
+ then have "\<nexists>b t. P' \<turnstile> ?D has F1,b:t in C1"
+ using False Getfield assms
+ classes_above_has_field2[where C="?D" and P=P and P'=P' and F=F1 and C'=C1]
+ by auto
+ then show ?thesis using nex Getfield assms classes_above_field[of C1 P P' F1]
+ by(cases "field P' C1 F1", cases "the (h (the_Addr (hd stk)))", auto)
+ qed
+ qed
+next
+ case (Getstatic C1 F1 D1)
+ then have C1: "classes_above P C1 \<inter> classes_changed P P' = {}" using assms by auto
+ show ?thesis
+ proof(cases "\<exists>b t. P \<turnstile> C1 has F1,b:t in D1")
+ case True
+ then obtain b t where meth: "P \<turnstile> C1 has F1,b:t in D1" by auto
+ then have "P \<turnstile> C1 \<preceq>\<^sup>* D1" by(rule has_field_decl_above)
+ then have D1: "classes_above P D1 \<inter> classes_changed P P' = {}"
+ using C1 rtrancl_trans by fastforce
+ have "P' \<turnstile> C1 has F1,b:t in D1"
+ using meth Getstatic assms classes_above_has_field[OF C1] by auto
+ then show ?thesis using True D1 Getstatic assms classes_above_field[of D1 P P' F1]
+ by(cases "field P' D1 F1", auto)
+ next
+ case False
+ then have "\<nexists>b t. P' \<turnstile> C1 has F1,b:t in D1"
+ using Getstatic assms
+ classes_above_has_field2[where C=C1 and P=P and P'=P' and F=F1 and C'=D1]
+ by auto
+ then show ?thesis using False Getstatic assms
+ by(cases "field P' D1 F1", auto)
+ qed
+next
+ case (Putfield F1 C1) show ?thesis
+ proof(cases "hd(tl stk) = Null")
+ case True then show ?thesis using Putfield assms by simp
+ next
+ case False
+ let ?D = "(cname_of h (the_Addr (hd (tl stk))))"
+ have D: "classes_above P ?D \<inter> classes_changed P P' = {}" using False Putfield assms by simp
+ show ?thesis
+ proof(cases "\<exists>b t. P \<turnstile> ?D has F1,b:t in C1")
+ case True
+ then obtain b1 t1 where "P \<turnstile> ?D has F1,b1:t1 in C1" by auto
+ then have has: "P' \<turnstile> ?D has F1,b1:t1 in C1"
+ using Putfield assms classes_above_has_field[OF D] by auto
+ have "P \<turnstile> ?D \<preceq>\<^sup>* C1" using has_field_decl_above True by auto
+ then have "classes_above P C1 \<subseteq> classes_above P ?D" by(rule classes_above_subcls_subset)
+ then have C1: "classes_above P C1 \<inter> classes_changed P P' = {}" using D by auto
+ then show ?thesis using has True Putfield assms classes_above_field[of C1 P P' F1]
+ by(cases "field P' C1 F1", cases "the (h (the_Addr (hd (tl stk))))", auto)
+ next
+ case nex: False
+ then have "\<nexists>b t. P' \<turnstile> ?D has F1,b:t in C1"
+ using False Putfield assms
+ classes_above_has_field2[where C="?D" and P=P and P'=P' and F=F1 and C'=C1]
+ by auto
+ then show ?thesis using nex Putfield assms classes_above_field[of C1 P P' F1]
+ by(cases "field P' C1 F1", cases "the (h (the_Addr (hd (tl stk))))", auto)
+ qed
+ qed
+next
+ case (Putstatic C1 F1 D1)
+ then have C1: "classes_above P C1 \<inter> classes_changed P P' = {}" using Putstatic assms by auto
+ show ?thesis
+ proof(cases "\<exists>b t. P \<turnstile> C1 has F1,b:t in D1")
+ case True
+ then obtain b t where meth: "P \<turnstile> C1 has F1,b:t in D1" by auto
+ then have "P \<turnstile> C1 \<preceq>\<^sup>* D1" by(rule has_field_decl_above)
+ then have D1: "classes_above P D1 \<inter> classes_changed P P' = {}"
+ using C1 rtrancl_trans by fastforce
+ then have "P' \<turnstile> C1 has F1,b:t in D1"
+ using meth Putstatic assms classes_above_has_field[OF C1] by auto
+ then show ?thesis using True D1 Putstatic assms classes_above_field[of D1 P P' F1]
+ by(cases "field P' D1 F1", auto)
+ next
+ case False
+ then have "\<nexists>b t. P' \<turnstile> C1 has F1,b:t in D1"
+ using Putstatic assms classes_above_has_field2[where C=C1 and P=P and P'=P' and F=F1 and C'=D1]
+ by auto
+ then show ?thesis using False Putstatic assms
+ by(cases "field P' D1 F1", auto)
+ qed
+next
+ case (Checkcast C1)
+ then show ?thesis using assms
+ proof(cases "hd stk = Null")
+ case False then show ?thesis
+ using Checkcast assms classes_above_subcls classes_above_subcls2
+ by(simp add: cast_ok_def) blast
+ qed(simp add: cast_ok_def)
+next
+ case (Invoke M n)
+ let ?C = "cname_of h (the_Addr (stk ! n))"
+ show ?thesis
+ proof(cases "stk ! n = Null")
+ case True then show ?thesis using Invoke assms by simp
+ next
+ case False
+ then have above: "classes_above P ?C \<inter> classes_changed P P' = {}"
+ using Invoke assms by simp
+ obtain D b Ts T mxs mxl ins xt where meth: "method P' ?C M = (D,b,Ts,T,mxs,mxl,ins,xt)"
+ by(cases "method P' ?C M", clarsimp)
+ have meq: "method P ?C M = method P' ?C M"
+ using classes_above_method[OF above] by simp
+ then show ?thesis
+ proof(cases "\<exists>Ts T m D b. P \<turnstile> ?C sees M,b:Ts \<rightarrow> T = m in D")
+ case nex: False
+ then have "\<not>(\<exists>Ts T m D b. P' \<turnstile> ?C sees M,b:Ts \<rightarrow> T = m in D)"
+ using classes_above_sees_method2[OF above, of M] by clarsimp
+ then show ?thesis using nex False Invoke by simp
+ next
+ case True
+ then have "\<exists>Ts T m D b. P' \<turnstile> ?C sees M,b:Ts \<rightarrow> T = m in D"
+ by(fastforce dest!: classes_above_sees_method[OF above, of M])
+ then show ?thesis using meq meth True Invoke by simp
+ qed
+ qed
+next
+ case (Invokestatic C1 M n)
+ then have above: "classes_above P C1 \<inter> classes_changed P P' = {}"
+ using assms by simp
+ obtain D b Ts T mxs mxl ins xt where meth: "method P' C1 M = (D,b,Ts,T,mxs,mxl,ins,xt)"
+ by(cases "method P' C1 M", clarsimp)
+ have meq: "method P C1 M = method P' C1 M"
+ using classes_above_method[OF above] by simp
+ show ?thesis
+ proof(cases "\<exists>Ts T m D b. P \<turnstile> C1 sees M,b:Ts \<rightarrow> T = m in D")
+ case False
+ then have "\<not>(\<exists>Ts T m D b. P' \<turnstile> C1 sees M,b:Ts \<rightarrow> T = m in D)"
+ using classes_above_sees_method2[OF above, of M] by clarsimp
+ then show ?thesis using False Invokestatic by simp
+ next
+ case True
+ then have "\<exists>Ts T m D b. P' \<turnstile> C1 sees M,b:Ts \<rightarrow> T = m in D"
+ by(fastforce dest!: classes_above_sees_method[OF above, of M])
+ then show ?thesis using meq meth True Invokestatic by simp
+ qed
+next
+ case Return then show ?thesis using assms classes_above_method[OF above_C]
+ by(cases frs, auto)
+next
+ case (Load x1) then show ?thesis using assms by auto
+next
+ case (Store x2) then show ?thesis using assms by auto
+next
+ case (Push x3) then show ?thesis using assms by auto
+next
+ case (Goto x15) then show ?thesis using assms by auto
+next
+ case (IfFalse x17) then show ?thesis using assms by auto
+qed(auto)
+
+
+\<comment> \<open> if collected classes unchanged, instruction collection unchanged \<close>
+lemma ncollect_JVMinstr_ncollect:
+assumes "JVMinstr_ncollect P i h stk \<inter> classes_changed P P' = {}"
+shows "JVMinstr_ncollect P i h stk = JVMinstr_ncollect P' i h stk"
+proof(cases i)
+ case (New C1)
+ then show ?thesis using assms classes_above_set[of C1 P P'] by auto
+next
+ case (Getfield F C1) show ?thesis
+ proof(cases "hd stk = Null")
+ case True then show ?thesis using Getfield assms by simp
+ next
+ case False
+ let ?D = "cname_of h (the_Addr (hd stk))"
+ have "classes_above P ?D \<inter> classes_changed P P' = {}" using False Getfield assms by auto
+ then have "classes_above P ?D = classes_above P' ?D"
+ using classes_above_set by blast
+ then show ?thesis using assms Getfield by auto
+ qed
+next
+ case (Getstatic C1 P1 D1)
+ then have "classes_above P C1 \<inter> classes_changed P P' = {}" using assms by auto
+ then have "classes_above P C1 = classes_above P' C1"
+ using classes_above_set assms Getstatic by blast
+ then show ?thesis using assms Getstatic by auto
+next
+ case (Putfield F C1) show ?thesis
+ proof(cases "hd(tl stk) = Null")
+ case True then show ?thesis using Putfield assms by simp
+ next
+ case False
+ let ?D = "cname_of h (the_Addr (hd (tl stk)))"
+ have "classes_above P ?D \<inter> classes_changed P P' = {}" using False Putfield assms by auto
+ then have "classes_above P ?D = classes_above P' ?D"
+ using classes_above_set by blast
+ then show ?thesis using assms Putfield by auto
+ qed
+next
+ case (Putstatic C1 F D1)
+ then have "classes_above P C1 \<inter> classes_changed P P' = {}" using assms by auto
+ then have "classes_above P C1 = classes_above P' C1"
+ using classes_above_set assms Putstatic by blast
+ then show ?thesis using assms Putstatic by auto
+next
+ case (Checkcast C1)
+ then show ?thesis using assms
+ classes_above_set[of "cname_of h (the_Addr (hd stk))" P P'] by auto
+next
+ case (Invoke M n)
+ then show ?thesis using assms
+ classes_above_set[of "cname_of h (the_Addr (stk ! n))" P P'] by auto
+next
+ case (Invokestatic C1 M n)
+ then show ?thesis using assms classes_above_set[of C1 P P'] by auto
+next
+ case Return
+ then show ?thesis using assms classes_above_set[of _ P P'] by auto
+next
+ case Throw
+ then show ?thesis using assms
+ classes_above_set[of "cname_of h (the_Addr (hd stk))" P P'] by auto
+qed(auto)
+
+\<comment> \<open> if collected classes unchanged, @{term exec_step} unchanged \<close>
+lemma ncollect_exec_step:
+assumes "JVMstep_ncollect P h stk C M pc ics \<inter> classes_changed P P' = {}"
+ and above_C: "classes_above P C \<inter> classes_changed P P' = {}"
+shows "exec_step P h stk loc C M pc ics frs sh = exec_step P' h stk loc C M pc ics frs sh"
+proof(cases ics)
+ case No_ics then show ?thesis
+ using ncollect_exec_instr assms classes_above_method[OF above_C, THEN sym]
+ by simp
+next
+ case (Calling C1 Cs)
+ then have above_C1: "classes_above P C1 \<inter> classes_changed P P' = {}"
+ using assms(1) by auto
+ show ?thesis
+ proof(cases "sh C1")
+ case None
+ then show ?thesis using Calling assms classes_above_sblank[OF above_C1] by simp
+ next
+ case (Some a)
+ then obtain sfs i where sfsi: "a = (sfs, i)" by(cases a)
+ then show ?thesis using Calling Some assms
+ proof(cases i)
+ case Prepared then show ?thesis
+ using above_C1 sfsi Calling Some assms classes_above_method[OF above_C1]
+ by(simp add: split_beta classes_above_class classes_changed_class[where cn=C1])
+ next
+ case Error then show ?thesis
+ using above_C1 sfsi Calling Some assms classes_above_method[of C1 P P']
+ by(simp add: split_beta classes_above_class classes_changed_class[where cn=C1])
+ qed(auto)
+ qed
+next
+ case (Called Cs) show ?thesis
+ proof(cases Cs)
+ case Nil then show ?thesis
+ using ncollect_exec_instr assms classes_above_method[OF above_C, THEN sym] Called
+ by simp
+ next
+ case (Cons C1 Cs1)
+ then have above_C': "classes_above P C1 \<inter> classes_changed P P' = {}" using assms Called by auto
+ show ?thesis using assms classes_above_method[OF above_C'] Cons Called by simp
+ qed
+next
+ case (Throwing Cs a) then show ?thesis using assms by(cases Cs; simp)
+qed
+
+\<comment> \<open> if collected classes unchanged, @{term exec_step} collection unchanged \<close>
+lemma ncollect_JVMstep_ncollect:
+assumes "JVMstep_ncollect P h stk C M pc ics \<inter> classes_changed P P' = {}"
+ and above_C: "classes_above P C \<inter> classes_changed P P' = {}"
+shows "JVMstep_ncollect P h stk C M pc ics = JVMstep_ncollect P' h stk C M pc ics"
+proof(cases ics)
+ case No_ics then show ?thesis
+ using assms ncollect_JVMinstr_ncollect classes_above_method[OF above_C]
+ by simp
+next
+ case (Calling C1 Cs)
+ then have above_C: "classes_above P C1 \<inter> classes_changed P P' = {}"
+ using assms(1) by auto
+ let ?C = "fst(method P C1 clinit)"
+ show ?thesis using Calling assms classes_above_method[OF above_C]
+ classes_above_set[OF above_C] by auto
+next
+ case (Called Cs) show ?thesis
+ proof(cases Cs)
+ case Nil then show ?thesis
+ using assms ncollect_JVMinstr_ncollect classes_above_method[OF above_C] Called
+ by simp
+ next
+ case (Cons C1 Cs1)
+ then have above_C1: "classes_above P C1 \<inter> classes_changed P P' = {}"
+ and above_C': "classes_above P (fst (method P C1 clinit)) \<inter> classes_changed P P' = {}"
+ using assms Called by auto
+ show ?thesis using assms Cons Called classes_above_set[OF above_C1]
+ classes_above_set[OF above_C'] classes_above_method[OF above_C1]
+ by simp
+ qed
+next
+ case (Throwing Cs a) then show ?thesis
+ using assms classes_above_set[of "cname_of h a" P P'] by simp
+qed
+
+\<comment> \<open> if collected classes unchanged, @{term classes_above_frames} unchanged \<close>
+lemma ncollect_classes_above_frames:
+ "JVMexec_ncollect P (None, h, (stk,loc,C,M,pc,ics)#frs, sh) \<inter> classes_changed P P' = {}
+ \<Longrightarrow> classes_above_frames P frs = classes_above_frames P' frs"
+proof(induct frs)
+ case (Cons f frs')
+ then obtain stk loc C M pc ics where f: "f = (stk,loc,C,M,pc,ics)" by(cases f)
+ then have above_C: "classes_above P C \<inter> classes_changed P P' = {}" using Cons by auto
+ show ?case using f Cons classes_above_subcls[OF above_C]
+ classes_above_subcls2[OF above_C] by auto
+qed(auto)
+
+\<comment> \<open> if collected classes unchanged, @{term classes_above_xcpts} unchanged \<close>
+lemma ncollect_classes_above_xcpts:
+assumes "JVMexec_ncollect P (None, h, (stk,loc,C,M,pc,ics)#frs, sh) \<inter> classes_changed P P' = {}"
+shows "classes_above_xcpts P = classes_above_xcpts P'"
+proof -
+ have left: "\<And>x x'. x' \<in> sys_xcpts \<Longrightarrow> P \<turnstile> x' \<preceq>\<^sup>* x \<Longrightarrow> \<exists>xa\<in>sys_xcpts. P' \<turnstile> xa \<preceq>\<^sup>* x"
+ proof -
+ fix x x'
+ assume x': "x' \<in> sys_xcpts" and above: "P \<turnstile> x' \<preceq>\<^sup>* x"
+ then show "\<exists>xa\<in>sys_xcpts. P' \<turnstile> xa \<preceq>\<^sup>* x" using assms classes_above_subcls[OF _ above]
+ by(rule_tac x=x' in bexI) auto
+ qed
+ have right: "\<And>x x'. x' \<in> sys_xcpts \<Longrightarrow> P' \<turnstile> x' \<preceq>\<^sup>* x \<Longrightarrow> \<exists>xa\<in>sys_xcpts. P \<turnstile> xa \<preceq>\<^sup>* x"
+ proof -
+ fix x x'
+ assume x': "x' \<in> sys_xcpts" and above: "P' \<turnstile> x' \<preceq>\<^sup>* x"
+ then show "\<exists>xa\<in>sys_xcpts. P \<turnstile> xa \<preceq>\<^sup>* x" using assms classes_above_subcls2[OF _ above]
+ by(rule_tac x=x' in bexI) auto
+ qed
+ show ?thesis using left right by auto
+qed
+
+\<comment> \<open> if collected classes unchanged, @{term exec} collection unchanged \<close>
+lemma ncollect_JVMexec_ncollect:
+assumes "JVMexec_ncollect P \<sigma> \<inter> classes_changed P P' = {}"
+shows "JVMexec_ncollect P \<sigma> = JVMexec_ncollect P' \<sigma>"
+proof -
+ obtain xp h frs sh where \<sigma>: "\<sigma> = (xp,h,frs,sh)" by(cases \<sigma>)
+ then show ?thesis using assms
+ proof(cases "\<exists>x. xp = Some x \<or> frs = []")
+ case False
+ then obtain stk loc C M pc ics frs' where frs: "frs = (stk,loc,C,M,pc,ics)#frs'"
+ by(cases frs, auto)
+ have step: "JVMstep_ncollect P h stk C M pc ics \<inter> classes_changed P P' = {}"
+ using False \<sigma> frs assms by(cases ics, auto simp: JVMNaiveCollectionSemantics.csmall_def)
+ have above_C: "classes_above P C \<inter> classes_changed P P' = {}"
+ using False \<sigma> frs assms by(auto simp: JVMNaiveCollectionSemantics.csmall_def)
+ have frames: "classes_above_frames P frs' = classes_above_frames P' frs'"
+ using ncollect_classes_above_frames frs \<sigma> False assms by simp
+ have xcpts: "classes_above_xcpts P = classes_above_xcpts P'"
+ using ncollect_classes_above_xcpts frs \<sigma> False assms by simp
+ show ?thesis using False xcpts frames frs \<sigma> ncollect_JVMstep_ncollect[OF step above_C]
+ classes_above_subcls[OF above_C] classes_above_subcls2[OF above_C]
+ by auto
+ qed(auto)
+qed
+
+\<comment> \<open> if collected classes unchanged, classes above an exception returned
+ by @{term exec_instr} unchanged \<close>
+lemma ncollect_exec_instr_xcpts:
+assumes collect: "JVMinstr_ncollect P i h stk \<inter> classes_changed P P' = {}"
+ and xpcollect: "classes_above_xcpts P \<inter> classes_changed P P' = {}"
+ and prealloc: "preallocated h"
+ and \<sigma>': "\<sigma>' = exec_instr i P h stk loc C M pc ics' frs sh"
+ and xp: "fst \<sigma>' = Some a"
+ and i: "i = instrs_of P C M ! pc"
+shows "classes_above P (cname_of h a) \<inter> classes_changed P P' = {}"
+using assms exec_instr_xcpts[OF \<sigma>' xp]
+proof(cases i)
+ case Throw then show ?thesis using assms by(cases "hd stk", fastforce+)
+qed(fastforce+)
+
+\<comment> \<open> if collected classes unchanged, classes above an exception returned
+ by @{term exec_step} unchanged \<close>
+lemma ncollect_exec_step_xcpts:
+assumes collect: "JVMstep_ncollect P h stk C M pc ics \<inter> classes_changed P P' = {}"
+ and xpcollect: "classes_above_xcpts P \<inter> classes_changed P P' = {}"
+ and prealloc: "preallocated h"
+ and \<sigma>': "\<sigma>' = exec_step P h stk loc C M pc ics frs sh"
+ and xp: "fst \<sigma>' = Some a"
+shows "classes_above P (cname_of h a) \<inter> classes_changed P P' = {}"
+proof(cases ics)
+ case No_ics then show ?thesis using assms ncollect_exec_instr_xcpts by simp
+next
+ case (Calling x21 x22)
+ then show ?thesis using assms by(clarsimp split: option.splits init_state.splits if_split_asm)
+next
+ case (Called Cs) then show ?thesis using assms ncollect_exec_instr_xcpts by(cases Cs; simp)
+next
+ case (Throwing Cs a) then show ?thesis using assms by(cases Cs; simp)
+qed
+
+\<comment> \<open> if collected classes unchanged, if @{term csmall} returned a result
+ under @{term P}, @{term P'} returns the same \<close>
+lemma ncollect_JVMsmall:
+assumes collect: "(\<sigma>', cset) \<in> JVMNaiveCollectionSemantics.csmall P \<sigma>"
+ and intersect: "cset \<inter> classes_changed P P' = {}"
+ and prealloc: "preallocated (fst(snd \<sigma>))"
+shows "(\<sigma>', cset) \<in> JVMNaiveCollectionSemantics.csmall P' \<sigma>"
+proof -
+ obtain xp h frs sh where \<sigma>: "\<sigma> = (xp,h,frs,sh)" by(cases \<sigma>)
+ then have prealloc': "preallocated h" using prealloc by simp
+ show ?thesis using \<sigma> assms
+ proof(cases "\<exists>x. xp = Some x \<or> frs = []")
+ case False
+ then obtain stk loc C M pc ics frs' where frs: "frs = (stk,loc,C,M,pc,ics)#frs'"
+ by(cases frs, auto)
+ have step: "JVMstep_ncollect P h stk C M pc ics \<inter> classes_changed P P' = {}"
+ using False \<sigma> frs assms by(cases ics, auto simp: JVMNaiveCollectionSemantics.csmall_def)
+ have above_C: "classes_above P C \<inter> classes_changed P P' = {}"
+ using False \<sigma> frs assms by(auto simp: JVMNaiveCollectionSemantics.csmall_def)
+ obtain xp1 h1 frs1 sh1 where exec: "exec_step P' h stk loc C M pc ics frs' sh = (xp1,h1,frs1,sh1)"
+ by(cases "exec_step P' h stk loc C M pc ics frs' sh")
+ have collect: "JVMexec_ncollect P \<sigma> = JVMexec_ncollect P' \<sigma>"
+ using assms ncollect_JVMexec_ncollect by(simp add: JVMNaiveCollectionSemantics.csmall_def)
+ have exec_instr: "exec_step P h stk loc C M pc ics frs' sh
+ = exec_step P' h stk loc C M pc ics frs' sh"
+ using ncollect_exec_step[OF step above_C] \<sigma> frs False by simp
+ show ?thesis
+ proof(cases xp1)
+ case None then show ?thesis
+ using None \<sigma> frs step False assms ncollect_exec_step[OF step above_C] collect exec
+ by(auto simp: JVMNaiveCollectionSemantics.csmall_def)
+ next
+ case (Some a)
+ then show ?thesis using \<sigma> assms
+ proof(cases xp)
+ case None
+ have frames: "classes_above_frames P (frames_of \<sigma>) \<inter> classes_changed P P' = {}"
+ using None Some frs \<sigma> assms by(auto simp: JVMNaiveCollectionSemantics.csmall_def)
+ have frsi: "classes_above_frames P frs \<inter> classes_changed P P' = {}" using \<sigma> frames by simp
+ have xpcollect: "classes_above_xcpts P \<inter> classes_changed P P' = {}"
+ using None Some frs \<sigma> assms by(auto simp: JVMNaiveCollectionSemantics.csmall_def)
+ have xp: "classes_above P (cname_of h a) \<inter> classes_changed P P' = {}"
+ using ncollect_exec_step_xcpts[OF step xpcollect prealloc',
+ where \<sigma>' = "(xp1,h1,frs1,sh1)" and frs=frs' and loc=loc and a=a and sh=sh]
+ exec exec_instr Some assms by auto
+ show ?thesis using Some exec \<sigma> frs False assms exec_instr collect
+ classes_above_find_handler[where h=h and sh=sh, OF xp frsi]
+ by(auto simp: JVMNaiveCollectionSemantics.csmall_def)
+ qed(auto simp: JVMNaiveCollectionSemantics.csmall_def)
+ qed
+ qed(auto simp: JVMNaiveCollectionSemantics.csmall_def)
+qed
+
+\<comment> \<open> if collected classes unchanged, if @{term cbig} returned a result
+ under @{term P}, @{term P'} returns the same \<close>
+lemma ncollect_JVMbig:
+assumes collect: "(\<sigma>', cset) \<in> JVMNaiveCollectionSemantics.cbig P \<sigma>"
+ and intersect: "cset \<inter> classes_changed P P' = {}"
+ and prealloc: "preallocated (fst(snd \<sigma>))"
+shows "(\<sigma>', cset) \<in> JVMNaiveCollectionSemantics.cbig P' \<sigma>"
+using JVMNaiveCollectionSemantics.csmall_to_cbig_prop2[where R="\<lambda>P P' cset. cset \<inter> classes_changed P P' = {}"
+ and Q="\<lambda>\<sigma>. preallocated (fst(snd \<sigma>))" and P=P and P'=P' and \<sigma>=\<sigma> and \<sigma>'=\<sigma>' and coll=cset]
+ ncollect_JVMsmall JVMsmall_prealloc_pres assms by auto
+
+\<comment> \<open> and finally, RTS algorithm based on @{term ncollect} is existence safe \<close>
+theorem jvm_naive_existence_safe:
+assumes p: "P \<in> jvm_progs" and "P' \<in> jvm_progs" and t: "t \<in> jvm_tests"
+ and out: "o1 \<in> jvm_naive_out P t" and "jvm_deselect P o1 P'"
+shows "\<exists>o2 \<in> jvm_naive_out P' t. o1 = o2"
+using assms
+proof -
+ let ?P = "start_prog (t#P) (fst t) main"
+ let ?P' = "start_prog (t#P') (fst t) main"
+ obtain wf_md where wf': "wf_prog wf_md (t#P)" using p t
+ by(auto simp: wf_jvm_prog_def wf_jvm_prog_phi_def)
+ have ns: "\<not>is_class (t#P) Start" using p t
+ by(clarsimp simp: is_class_def class_def Start_def Test_def)
+ obtain \<sigma>1 coll1 where o1: "o1 = (\<sigma>1, coll1)" by(cases o1)
+ then have cbig: "(\<sigma>1, coll1) \<in> JVMNaiveCollectionSemantics.cbig ?P (start_state (t # P))"
+ using assms by simp
+ have "coll1 \<inter> classes_changed P P' = {}" using cbig o1 assms by auto
+ then have changed: "coll1 \<inter> classes_changed (t#P) (t#P') = {}" by(rule classes_changed_int_Cons)
+ then have changed': "coll1 \<inter> classes_changed ?P ?P' = {}" by(rule classes_changed_int_Cons)
+ have "classes_above_xcpts ?P = classes_above_xcpts (t#P)"
+ using class_add_classes_above[OF ns wf_sys_xcpt_nsub_Start[OF wf' ns]] by simp
+ then have "classes_above_xcpts (t#P) \<inter> classes_changed (t#P) (t#P') = {}"
+ using jvm_naive_out_xcpts_collected[OF out] o1 changed by auto
+ then have ss_eq: "start_state (t#P) = start_state (t#P')"
+ using classes_above_start_state by simp
+ show ?thesis using ncollect_JVMbig[OF cbig changed']
+ preallocated_start_state changed' ss_eq o1 assms by auto
+qed
+
+\<comment> \<open> ...thus @{term JVMNaiveCollection} is an instance of @{term CollectionBasedRTS} \<close>
+interpretation JVMNaiveCollectionRTS :
+ CollectionBasedRTS "(=)" jvm_deselect jvm_progs jvm_tests
+ JVMendset JVMcombine JVMcollect_id JVMsmall JVMNaiveCollect jvm_naive_out
+ jvm_make_test_prog jvm_naive_collect_start
+ by unfold_locales (rule jvm_naive_existence_safe, auto simp: start_state_def)
+
+(***********************************************************************************************)
+subsection "Smarter RTS algorithm"
+
+subsubsection "Definitions and helper lemmas"
+
+fun jvm_smart_out :: "jvm_prog \<Rightarrow> jvm_class \<Rightarrow> jvm_prog_out set" where
+"jvm_smart_out P t
+ = {(\<sigma>',coll'). \<exists>coll. (\<sigma>',coll) \<in> JVMSmartCollectionSemantics.cbig
+ (jvm_make_test_prog P t) (start_state (t#P))
+ \<and> coll' = coll \<union> classes_above_xcpts P \<union> {Object,Start}}"
+
+abbreviation jvm_smart_collect_start :: "jvm_prog \<Rightarrow> cname set" where
+"jvm_smart_collect_start P \<equiv> classes_above_xcpts P \<union> {Object,Start}"
+
+
+lemma jvm_naive_iff_smart:
+"(\<exists>cset\<^sub>n. (\<sigma>',cset\<^sub>n) \<in> jvm_naive_out P t) \<longleftrightarrow> (\<exists>cset\<^sub>s. (\<sigma>',cset\<^sub>s) \<in> jvm_smart_out P t)"
+ by(auto simp: JVMNaiveCollectionSemantics.cbig_big_equiv
+ JVMSmartCollectionSemantics.cbig_big_equiv)
+
+(**************************************************)
+
+lemma jvm_smart_out_classes_above_xcpts:
+assumes s: "(\<sigma>',cset\<^sub>s) \<in> jvm_smart_out P t" and P: "P \<in> jvm_progs" and t: "t \<in> jvm_tests"
+shows "classes_above_xcpts (jvm_make_test_prog P t) \<subseteq> cset\<^sub>s"
+ using jvm_make_test_prog_classes_above_xcpts[OF P t] s by clarsimp
+
+lemma jvm_smart_collect_start_make_test_prog:
+ "\<lbrakk> P \<in> jvm_progs; t \<in> jvm_tests \<rbrakk>
+ \<Longrightarrow> jvm_smart_collect_start (jvm_make_test_prog P t) = jvm_smart_collect_start P"
+ using jvm_make_test_prog_classes_above_xcpts by simp
+
+lemma jvm_smart_out_classes_above_start_heap:
+assumes s: "(\<sigma>',cset\<^sub>s) \<in> jvm_smart_out P t" and h: "start_heap (t#P) a = Some(C,fs)"
+ and P: "P \<in> jvm_progs" and t: "t \<in> jvm_tests"
+shows "classes_above (jvm_make_test_prog P t) C \<subseteq> cset\<^sub>s"
+using start_heap_classes[OF h] jvm_smart_out_classes_above_xcpts[OF s P t] by auto
+
+lemma jvm_smart_out_classes_above_start_sheap:
+assumes "(\<sigma>',cset\<^sub>s) \<in> jvm_smart_out P t" and "start_sheap C = Some(sfs,i)"
+shows "classes_above (jvm_make_test_prog P t) C \<subseteq> cset\<^sub>s"
+using assms start_prog_classes_above_Start by(clarsimp split: if_split_asm)
+
+lemma jvm_smart_out_classes_above_frames:
+ "(\<sigma>',cset\<^sub>s) \<in> jvm_smart_out P t
+ \<Longrightarrow> classes_above_frames (jvm_make_test_prog P t) (frames_of (start_state (t#P))) \<subseteq> cset\<^sub>s"
+using start_prog_classes_above_Start by(clarsimp split: if_split_asm simp: start_state_def)
+
+(**************************************************)
+subsubsection "Additional well-formedness conditions"
+
+\<comment> \<open> returns class to be initialized by given instruction, if applicable \<close>
+(* NOTE: similar to exp-taking init_class from J/EConform - but requires field existence checks *)
+fun coll_init_class :: "'m prog \<Rightarrow> instr \<Rightarrow> cname option" where
+"coll_init_class P (New C) = Some C" |
+"coll_init_class P (Getstatic C F D) = (if \<exists>t. P \<turnstile> C has F,Static:t in D
+ then Some D else None)" |
+"coll_init_class P (Putstatic C F D) = (if \<exists>t. P \<turnstile> C has F,Static:t in D
+ then Some D else None)" |
+"coll_init_class P (Invokestatic C M n) = seeing_class P C M" |
+"coll_init_class _ _ = None"
+
+\<comment> \<open> checks whether the given value is a pointer; if it's an address,
+ checks whether it points to an object in the given heap \<close>
+fun is_ptr :: "heap \<Rightarrow> val \<Rightarrow> bool" where
+"is_ptr h Null = True" |
+"is_ptr h (Addr a) = (\<exists>Cfs. h a = Some Cfs)" |
+"is_ptr h _ = False"
+
+lemma is_ptrD: "is_ptr h v \<Longrightarrow> v = Null \<or> (\<exists>a. v = Addr a \<and> (\<exists>Cfs. h a = Some Cfs))"
+ by(cases v, auto)
+
+\<comment> \<open> shorthand for: given stack has entries required by given instr,
+ including pointer where necessary \<close>
+fun stack_safe :: "instr \<Rightarrow> heap \<Rightarrow> val list \<Rightarrow> bool" where
+"stack_safe (Getfield F C) h stk = (length stk > 0 \<and> is_ptr h (hd stk))" |
+"stack_safe (Putfield F C) h stk = (length stk > 1 \<and> is_ptr h (hd (tl stk)))" |
+"stack_safe (Checkcast C) h stk = (length stk > 0 \<and> is_ptr h (hd stk))" |
+"stack_safe (Invoke M n) h stk = (length stk > n \<and> is_ptr h (stk ! n))" |
+"stack_safe JVMInstructions.Throw h stk = (length stk > 0 \<and> is_ptr h (hd stk))" |
+"stack_safe i h stk = True"
+
+lemma well_formed_stack_safe:
+assumes wtp: "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P" and correct: "P,\<Phi> \<turnstile> (xp,h,(stk,loc,C,M,pc,ics)#frs,sh)\<surd>"
+shows "stack_safe (instrs_of P C M ! pc) h stk"
+proof -
+ from correct obtain b Ts T mxs mxl\<^sub>0 ins xt where
+ mC: "P \<turnstile> C sees M,b:Ts\<rightarrow>T=(mxs,mxl\<^sub>0,ins,xt) in C" and
+ pc: "pc < length ins" by clarsimp
+ from sees_wf_mdecl[OF _ mC] wtp have "wt_method P C b Ts T mxs mxl\<^sub>0 ins xt (\<Phi> C M)"
+ by(auto simp: wf_jvm_prog_phi_def wf_mdecl_def)
+ with pc have wt: "P,T,mxs,length ins,xt \<turnstile> ins ! pc,pc :: \<Phi> C M" by(simp add: wt_method_def)
+ from mC correct obtain ST LT where
+ \<Phi>: "\<Phi> C M ! pc = Some (ST,LT)" and
+ stk: "P,h \<turnstile> stk [:\<le>] ST" by fastforce
+ show ?thesis
+ proof(cases "instrs_of P C M ! pc")
+ case (Getfield F D)
+ with mC \<Phi> wt stk obtain oT ST' where
+ oT: "P \<turnstile> oT \<le> Class D" and
+ ST: "ST = oT # ST'" by fastforce
+ with stk obtain ref stk' where
+ stk': "stk = ref#stk'" and
+ ref: "P,h \<turnstile> ref :\<le> oT" by auto
+ with ref oT have "ref = Null \<or> (ref \<noteq> Null \<and> P,h \<turnstile> ref :\<le> Class D)" by auto
+ with Getfield mC have "is_ptr h ref" by(fastforce dest: non_npD)
+ with stk' Getfield show ?thesis by auto
+ next
+ case (Putfield F D)
+ with mC \<Phi> wt stk obtain vT oT ST' where
+ oT: "P \<turnstile> oT \<le> Class D" and
+ ST: "ST = vT # oT # ST'" by fastforce
+ with stk obtain v ref stk' where
+ stk': "stk = v#ref#stk'" and
+ ref: "P,h \<turnstile> ref :\<le> oT" by auto
+ with ref oT have "ref = Null \<or> (ref \<noteq> Null \<and> P,h \<turnstile> ref :\<le> Class D)" by auto
+ with Putfield mC have "is_ptr h ref" by(fastforce dest: non_npD)
+ with stk' Putfield show ?thesis by auto
+ next
+ case (Checkcast D)
+ with mC \<Phi> wt stk obtain oT ST' where
+ oT: "is_refT oT" and
+ ST: "ST = oT # ST'" by fastforce
+ with stk obtain ref stk' where
+ stk': "stk = ref#stk'" and
+ ref: "P,h \<turnstile> ref :\<le> oT" by auto
+ with ref oT have "ref = Null \<or> (ref \<noteq> Null \<and> (\<exists>D'. P,h \<turnstile> ref :\<le> Class D'))"
+ by(auto simp: is_refT_def)
+ with Checkcast mC have "is_ptr h ref" by(fastforce dest: non_npD)
+ with stk' Checkcast show ?thesis by auto
+ next
+ case (Invoke M1 n)
+ with mC \<Phi> wt stk have
+ ST: "n < size ST" and
+ oT: "ST!n = NT \<or> (\<exists>D. ST!n = Class D)" by auto
+ with stk have stk': "n < size stk" by (auto simp: list_all2_lengthD)
+ with stk ST oT list_all2_nthD2
+ have "stk!n = Null \<or> (stk!n \<noteq> Null \<and> (\<exists>D. P,h \<turnstile> stk!n :\<le> Class D))" by fastforce
+ with Invoke mC have "is_ptr h (stk!n)" by(fastforce dest: non_npD)
+ with stk' Invoke show ?thesis by auto
+ next
+ case Throw
+ with mC \<Phi> wt stk obtain oT ST' where
+ oT: "is_refT oT" and
+ ST: "ST = oT # ST'" by fastforce
+ with stk obtain ref stk' where
+ stk': "stk = ref#stk'" and
+ ref: "P,h \<turnstile> ref :\<le> oT" by auto
+ with ref oT have "ref = Null \<or> (ref \<noteq> Null \<and> (\<exists>D'. P,h \<turnstile> ref :\<le> Class D'))"
+ by(auto simp: is_refT_def)
+ with Throw mC have "is_ptr h ref" by(fastforce dest: non_npD)
+ with stk' Throw show ?thesis by auto
+ qed(simp_all)
+qed
+
+(******************************************)
+subsubsection \<open> Proving naive @{text "\<subseteq>"} smart \<close>
+
+text \<open> We prove that, given well-formedness of the program and state, and "promises"
+ about what has or will be collected in previous or future steps, @{term jvm_smart}
+ collects everything @{term jvm_naive} does. We prove that promises about previously-
+ collected classes ("backward promises") are maintained by execution, and promises
+ about to-be-collected classes ("forward promises") are met by the end of execution.
+ We then show that the required initial conditions (well-formedness and backward
+ promises) are met by the defined start states, and thus that a run test will
+ collect at least those classes collected by the naive algorithm. \<close>
+
+\<comment> \<open> Backward promises (classes that should already be collected) \<close>
+ \<comment> \<open> - Classes of objects in the heap are collected \<close>
+ \<comment> \<open> - Non-@{term None} classes on the static heap are collected \<close>
+ \<comment> \<open> - Current classes from the frame stack are collected \<close>
+ \<comment> \<open> - Classes of system exceptions are collected \<close>
+
+text "If backward promises have been kept, a single step preserves this property;
+ i.e., any classes that have been added to this set (new heap objects, newly prepared
+ sheap classes, new frames) are collected by the smart collection algorithm in that
+ step or by forward promises:"
+lemma backward_coll_promises_kept:
+assumes
+\<comment> \<open> well-formedness \<close>
+ wtp: "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P"
+ and correct: "P,\<Phi> \<turnstile> (xp,h,frs,sh)\<surd>"
+\<comment> \<open> defs \<close>
+ and f': "hd frs = (stk,loc,C',M',pc,ics)"
+\<comment> \<open> backward promises - will be collected prior \<close>
+ and heap: "\<And>C fs. \<exists>a. h a = Some(C,fs) \<Longrightarrow> classes_above P C \<subseteq> cset"
+ and sheap: "\<And>C sfs i. sh C = Some(sfs,i) \<Longrightarrow> classes_above P C \<subseteq> cset"
+ and xcpts: "classes_above_xcpts P \<subseteq> cset"
+ and frames: "classes_above_frames P frs \<subseteq> cset"
+\<comment> \<open> forward promises - will be collected after if not already \<close>
+ and init_class_prom: "\<And>C. ics = Called [] \<or> ics = No_ics
+ \<Longrightarrow> coll_init_class P (instrs_of P C' M' ! pc) = Some C \<Longrightarrow> classes_above P C \<subseteq> cset"
+ and Calling_prom: "\<And>C' Cs'. ics = Calling C' Cs' \<Longrightarrow> classes_above P C' \<subseteq> cset"
+\<comment> \<open> collection and step \<close>
+ and smart: "JVMexec_scollect P (xp,h,frs,sh) \<subseteq> cset"
+ and small: "(xp',h',frs',sh') \<in> JVMsmall P (xp,h,frs,sh)"
+shows "(h' a = Some(C,fs) \<longrightarrow> classes_above P C \<subseteq> cset)
+ \<and> (sh' C = Some(sfs',i') \<longrightarrow> classes_above P C \<subseteq> cset)
+ \<and> (classes_above_frames P frs' \<subseteq> cset)"
+using assms
+proof(cases frs)
+ case (Cons f1 frs1)
+(****)
+ then have cr': "P,\<Phi> \<turnstile> (xp,h,(stk,loc,C',M',pc,ics)#frs1,sh)\<surd>" using correct f' by simp
+ let ?i = "instrs_of P C' M' ! pc"
+ from well_formed_stack_safe[OF wtp cr'] correct_state_Throwing_ex[OF cr'] obtain
+ stack_safe: "stack_safe ?i h stk" and
+ Throwing_ex: "\<And>Cs a. ics = Throwing Cs a \<Longrightarrow> \<exists>obj. h a = Some obj" by simp
+ have confc: "conf_clinit P sh frs" using correct Cons by simp
+ have Called_prom: "\<And>C' Cs'. ics = Called (C'#Cs')
+ \<Longrightarrow> classes_above P C' \<subseteq> cset \<and> classes_above P (fst(method P C' clinit)) \<subseteq> cset"
+ proof -
+ fix C' Cs' assume [simp]: "ics = Called (C'#Cs')"
+ then have "C' \<in> set(clinit_classes frs)" using f' Cons by simp
+ then obtain sfs where shC': "sh C' = Some(sfs, Processing)" and "is_class P C'"
+ using confc by(auto simp: conf_clinit_def)
+ then have C'eq: "C' = fst(method P C' clinit)" using wf_sees_clinit wtp
+ by(fastforce simp: is_class_def wf_jvm_prog_phi_def)
+ then show "classes_above P C' \<subseteq> cset \<and> classes_above P (fst(method P C' clinit)) \<subseteq> cset"
+ using sheap shC' by auto
+ qed
+ have Called_prom2: "\<And>Cs. ics = Called Cs \<Longrightarrow> \<exists>C1 sobj. Called_context P C1 ?i \<and> sh C1 = Some sobj"
+ using cr' by(auto simp: conf_f_def2)
+ have Throwing_prom: "\<And>C' Cs a. ics = Throwing (C'#Cs) a \<Longrightarrow> \<exists>sfs. sh C' = Some(sfs, Processing)"
+ proof -
+ fix C' Cs a assume [simp]: "ics = Throwing (C'#Cs) a"
+ then have "C' \<in> set(clinit_classes frs)" using f' Cons by simp
+ then show "\<exists>sfs. sh C' = Some(sfs, Processing)" using confc by(clarsimp simp: conf_clinit_def)
+ qed
+(****)
+ show ?thesis using Cons assms
+ proof(cases xp)
+ case None
+ then have exec: "exec (P, None, h, (stk,loc,C',M',pc,ics)#frs1, sh) = Some (xp',h',frs',sh')"
+ using small f' Cons by auto
+ obtain si where si: "exec_step_input P C' M' pc ics = si" by simp
+ obtain xp\<^sub>0 h\<^sub>0 frs\<^sub>0 sh\<^sub>0 where
+ exec_step: "exec_step P h stk loc C' M' pc ics frs1 sh = (xp\<^sub>0, h\<^sub>0, frs\<^sub>0, sh\<^sub>0)"
+ by(cases "exec_step P h stk loc C' M' pc ics frs1 sh")
+ then have ind: "exec_step_ind si P h stk loc C' M' pc ics frs1 sh
+ (xp\<^sub>0, h\<^sub>0, frs\<^sub>0, sh\<^sub>0)" using exec_step_ind_equiv si by auto
+ then show ?thesis using heap sheap frames exec exec_step f' Cons
+ si init_class_prom stack_safe Calling_prom Called_prom Called_prom2 Throwing_prom
+ proof(induct rule: exec_step_ind.induct)
+ case exec_step_ind_Load show ?case using exec_step_ind_Load.prems(1-7) by auto
+ next
+ case exec_step_ind_Store show ?case using exec_step_ind_Store.prems(1-7) by auto
+ next
+ case exec_step_ind_Push show ?case using exec_step_ind_Push.prems(1-7) by auto
+ next
+ case exec_step_ind_NewOOM_Called show ?case using exec_step_ind_NewOOM_Called.prems(1-7)
+ by(auto simp del: find_handler.simps dest!: find_handler_pieces) blast+
+ next
+ case exec_step_ind_NewOOM_Done show ?case using exec_step_ind_NewOOM_Done.prems(1-7)
+ by(auto simp del: find_handler.simps dest!: find_handler_pieces) blast+
+ next
+ case exec_step_ind_New_Called show ?case
+ using exec_step_ind_New_Called.hyps exec_step_ind_New_Called.prems(1-9)
+ by(auto split: if_split_asm simp: blank_def dest!: exec_step_input_StepID) blast+
+ next
+ case exec_step_ind_New_Done show ?case
+ using exec_step_ind_New_Done.hyps exec_step_ind_New_Done.prems(1-9)
+ by(auto split: if_split_asm simp: blank_def dest!: exec_step_input_StepID) blast+
+ next
+ case exec_step_ind_New_Init show ?case
+ using exec_step_ind_New_Init.prems(1-7) by auto
+ next
+ case exec_step_ind_Getfield_Null show ?case using exec_step_ind_Getfield_Null.prems(1-7)
+ by(auto simp del: find_handler.simps dest!: find_handler_pieces) blast+
+ next
+ case exec_step_ind_Getfield_NoField show ?case
+ using exec_step_ind_Getfield_NoField.prems(1-7)
+ by(auto simp del: find_handler.simps dest!: find_handler_pieces) blast+
+ next
+ case exec_step_ind_Getfield_Static show ?case
+ using exec_step_ind_Getfield_Static.prems(1-7)
+ by(auto simp del: find_handler.simps dest!: find_handler_pieces) blast+
+ next
+ case exec_step_ind_Getfield show ?case
+ using exec_step_ind_Getfield.prems(1-7) by auto
+ next
+ case exec_step_ind_Getstatic_NoField show ?case
+ using exec_step_ind_Getstatic_NoField.prems(1-7)
+ by(auto simp del: find_handler.simps dest!: find_handler_pieces) blast+
+ next
+ case exec_step_ind_Getstatic_NonStatic show ?case
+ using exec_step_ind_Getstatic_NonStatic.prems(1-7)
+ by(auto simp del: find_handler.simps dest!: find_handler_pieces) blast+
+ next
+ case exec_step_ind_Getstatic_Called show ?case
+ using exec_step_ind_Getstatic_Called.prems(1-7) by auto
+ next
+ case exec_step_ind_Getstatic_Done show ?case
+ using exec_step_ind_Getstatic_Done.prems(1-7) by auto
+ next
+ case exec_step_ind_Getstatic_Init show ?case
+ using exec_step_ind_Getstatic_Init.prems(1-7) by auto
+ next
+ case exec_step_ind_Putfield_Null show ?case
+ using exec_step_ind_Putfield_Null.prems(1-7)
+ by(auto simp del: find_handler.simps dest!: find_handler_pieces) blast+
+ next
+ case exec_step_ind_Putfield_NoField show ?case
+ using exec_step_ind_Putfield_NoField.prems(1-7)
+ by(auto simp del: find_handler.simps dest!: find_handler_pieces) blast+
+ next
+ case exec_step_ind_Putfield_Static show ?case
+ using exec_step_ind_Putfield_Static.prems(1-7)
+ by(auto simp del: find_handler.simps dest!: find_handler_pieces) blast+
+ next
+ case (exec_step_ind_Putfield v stk r a D fs h D' b t P C F loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
+ obtain a C1 fs where addr: "hd (tl stk) = Null \<or> (hd (tl stk) = Addr a \<and> h a = Some(C1,fs))"
+ using exec_step_ind_Putfield.prems(8,10) by(auto dest!: exec_step_input_StepID is_ptrD)
+ then have "\<And>a. hd(tl stk) = Addr a \<Longrightarrow> classes_above P C1 \<subseteq> cset"
+ using exec_step_ind_Putfield.prems(1) addr by auto
+ then show ?case using exec_step_ind_Putfield.hyps exec_step_ind_Putfield.prems(1-7) addr
+ by(auto split: if_split_asm) blast+
+ next
+ case exec_step_ind_Putstatic_NoField show ?case
+ using exec_step_ind_Putstatic_NoField.prems(1-7)
+ by(auto simp del: find_handler.simps dest!: find_handler_pieces) blast+
+ next
+ case exec_step_ind_Putstatic_NonStatic show ?case
+ using exec_step_ind_Putstatic_NonStatic.prems(1-7)
+ by(auto simp del: find_handler.simps dest!: find_handler_pieces) blast+
+ next
+ case (exec_step_ind_Putstatic_Called D' b t P D F C sh sfs i h stk loc C\<^sub>0 M\<^sub>0 pc Cs frs)
+ then have "P \<turnstile> D sees F,Static:t in D" by(simp add: has_field_sees[OF has_field_idemp])
+ then have D'eq: "D' = D" using exec_step_ind_Putstatic_Called.hyps(1) by simp
+ obtain sobj where "sh D = Some sobj"
+ using exec_step_ind_Putstatic_Called.hyps(2) exec_step_ind_Putstatic_Called.prems(8,13)
+ by(fastforce dest!: exec_step_input_StepID)
+ then show ?case using exec_step_ind_Putstatic_Called.hyps
+ exec_step_ind_Putstatic_Called.prems(1-7) D'eq
+ by(auto split: if_split_asm) blast+
+ next
+ case exec_step_ind_Putstatic_Done show ?case
+ using exec_step_ind_Putstatic_Done.hyps exec_step_ind_Putstatic_Done.prems(1-7)
+ by(auto split: if_split_asm) blast+
+ next
+ case exec_step_ind_Putstatic_Init show ?case
+ using exec_step_ind_Putstatic_Init.hyps exec_step_ind_Putstatic_Init.prems(1-7)
+ by(auto split: if_split_asm) blast+
+ next
+ case exec_step_ind_Checkcast show ?case
+ using exec_step_ind_Checkcast.prems(1-7) by auto
+ next
+ case exec_step_ind_Checkcast_Error show ?case using exec_step_ind_Checkcast_Error.prems(1-7)
+ by(auto simp del: find_handler.simps dest!: find_handler_pieces) blast+
+ next
+ case exec_step_ind_Invoke_Null show ?case using exec_step_ind_Invoke_Null.prems(1-7)
+ by(auto simp del: find_handler.simps dest!: find_handler_pieces) blast+
+ next
+ case exec_step_ind_Invoke_NoMethod show ?case using exec_step_ind_Invoke_NoMethod.prems(1-7)
+ by(auto simp del: find_handler.simps dest!: find_handler_pieces) blast+
+ next
+ case exec_step_ind_Invoke_Static show ?case using exec_step_ind_Invoke_Static.prems(1-7)
+ by(auto simp del: find_handler.simps dest!: find_handler_pieces) blast+
+ next
+ case (exec_step_ind_Invoke ps n stk r C h D b Ts T mxs mxl\<^sub>0 ins xt P)
+ have "classes_above P D \<subseteq> cset"
+ using exec_step_ind_Invoke.hyps(2,3,5) exec_step_ind_Invoke.prems(1,8,10,13)
+ rtrancl_trans[OF sees_method_decl_above[OF exec_step_ind_Invoke.hyps(6)]]
+ by(auto dest!: exec_step_input_StepID is_ptrD) blast+
+ then show ?case
+ using exec_step_ind_Invoke.hyps(7) exec_step_ind_Invoke.prems(1-7) by auto
+ next
+ case exec_step_ind_Invokestatic_NoMethod
+ show ?case using exec_step_ind_Invokestatic_NoMethod.prems(1-7)
+ by(auto simp del: find_handler.simps dest!: find_handler_pieces) blast+
+ next
+ case exec_step_ind_Invokestatic_NonStatic
+ show ?case using exec_step_ind_Invokestatic_NonStatic.prems(1-7)
+ by(auto simp del: find_handler.simps dest!: find_handler_pieces) blast+
+ next
+ case (exec_step_ind_Invokestatic_Called ps n stk D b Ts T mxs mxl\<^sub>0 ins xt P C M)
+ have "seeing_class P C M = Some D" using exec_step_ind_Invokestatic_Called.hyps(2,3)
+ by(fastforce simp: seeing_class_def)
+ then have "classes_above P D \<subseteq> cset" using exec_step_ind_Invokestatic_Called.prems(8-9)
+ by(fastforce dest!: exec_step_input_StepID)
+ then show ?case
+ using exec_step_ind_Invokestatic_Called.hyps exec_step_ind_Invokestatic_Called.prems(1-7)
+ by(auto simp: seeing_class_def)
+ next
+ case (exec_step_ind_Invokestatic_Done ps n stk D b Ts T mxs mxl\<^sub>0 ins xt P C M)
+ have "seeing_class P C M = Some D" using exec_step_ind_Invokestatic_Done.hyps(2,3)
+ by(fastforce simp: seeing_class_def)
+ then have "classes_above P D \<subseteq> cset" using exec_step_ind_Invokestatic_Done.prems(8-9)
+ by(fastforce dest!: exec_step_input_StepID)
+ then show ?case
+ using exec_step_ind_Invokestatic_Done.hyps exec_step_ind_Invokestatic_Done.prems(1-7)
+ by auto blast+
+ next
+ case exec_step_ind_Invokestatic_Init show ?case
+ using exec_step_ind_Invokestatic_Init.hyps exec_step_ind_Invokestatic_Init.prems(1-7)
+ by auto blast+
+ next
+ case exec_step_ind_Return_Last_Init show ?case
+ using exec_step_ind_Return_Last_Init.prems(1-7) by(auto split: if_split_asm) blast+
+ next
+ case exec_step_ind_Return_Last show ?case
+ using exec_step_ind_Return_Last.prems(1-7) by auto
+ next
+ case exec_step_ind_Return_Init show ?case
+ using exec_step_ind_Return_Init.prems(1-7) by(auto split: if_split_asm) blast+
+ next
+ case exec_step_ind_Return_NonStatic show ?case
+ using exec_step_ind_Return_NonStatic.prems(1-7) by auto
+ next
+ case exec_step_ind_Return_Static show ?case
+ using exec_step_ind_Return_Static.prems(1-7) by auto
+ next
+ case exec_step_ind_Pop show ?case using exec_step_ind_Pop.prems(1-7) by auto
+ next
+ case exec_step_ind_IAdd show ?case using exec_step_ind_IAdd.prems(1-7)by auto
+ next
+ case exec_step_ind_IfFalse_False show ?case
+ using exec_step_ind_IfFalse_False.prems(1-7) by auto
+ next
+ case exec_step_ind_IfFalse_nFalse show ?case
+ using exec_step_ind_IfFalse_nFalse.prems(1-7) by auto
+ next
+ case exec_step_ind_CmpEq show ?case using exec_step_ind_CmpEq.prems(1-7) by auto
+ next
+ case exec_step_ind_Goto show ?case using exec_step_ind_Goto.prems(1-7) by auto
+ next
+ case exec_step_ind_Throw show ?case using exec_step_ind_Throw.prems(1-7)
+ by(auto simp del: find_handler.simps dest!: find_handler_pieces) blast+
+ next
+ case exec_step_ind_Throw_Null show ?case using exec_step_ind_Throw_Null.prems(1-7)
+ by(auto simp del: find_handler.simps dest!: find_handler_pieces) blast+
+ next
+ case (exec_step_ind_Init_None_Called sh C Cs P)
+ have "classes_above P C \<subseteq> cset" using exec_step_ind_Init_None_Called.prems(8,11)
+ by(auto dest!: exec_step_input_StepCD)
+ then show ?case using exec_step_ind_Init_None_Called.prems(1-7)
+ by(auto split: if_split_asm) blast+
+ next
+ case exec_step_ind_Init_Done show ?case
+ using exec_step_ind_Init_Done.prems(1-7) by auto
+ next
+ case exec_step_ind_Init_Processing show ?case
+ using exec_step_ind_Init_Processing.prems(1-7) by auto
+ next
+ case exec_step_ind_Init_Error show ?case
+ using exec_step_ind_Init_Error.prems(1-7) by auto
+ next
+ case exec_step_ind_Init_Prepared_Object show ?case
+ using exec_step_ind_Init_Prepared_Object.hyps
+ exec_step_ind_Init_Prepared_Object.prems(1-7,10)
+ by(auto split: if_split_asm dest!: exec_step_input_StepCD) blast+
+ next
+ case exec_step_ind_Init_Prepared_nObject show ?case
+ using exec_step_ind_Init_Prepared_nObject.hyps exec_step_ind_Init_Prepared_nObject.prems(1-7)
+ by(auto split: if_split_asm) blast+
+ next
+ case exec_step_ind_Init show ?case
+ using exec_step_ind_Init.prems(1-7,8,12)
+ by(auto simp: split_beta dest!: exec_step_input_StepC2D)
+ next
+ case (exec_step_ind_InitThrow C Cs a P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
+ obtain sfs where "sh C = Some(sfs,Processing)"
+ using exec_step_ind_InitThrow.prems(8,14) by(fastforce dest!: exec_step_input_StepTD)
+ then show ?case using exec_step_ind_InitThrow.prems(1-7)
+ by(auto split: if_split_asm) blast+
+ next
+ case exec_step_ind_InitThrow_End show ?case using exec_step_ind_InitThrow_End.prems(1-7)
+ by(auto simp del: find_handler.simps dest!: find_handler_pieces) blast+
+ qed
+ qed(simp)
+qed(simp)
+
+\<comment> \<open> Forward promises (classes that will be collected by the end of execution) \<close>
+ \<comment> \<open> - Classes that the current instruction will check initialization for will be collected \<close>
+ \<comment> \<open> - Class whose initialization is actively being called by the current frame will be collected \<close>
+
+text \<open> We prove that an @{term ics} of @{text "Calling C Cs"} (meaning @{text C}'s
+ initialization procedure is actively being called) means that classes above @{text C}
+ will be collected by @{term cbig} (i.e., by the end of execution) using proof by
+ induction, proving the base and IH separately. \<close>
+
+\<comment> \<open> base case: @{term Object} \<close>
+lemma Calling_collects_base:
+assumes big: "(\<sigma>', cset') \<in> JVMSmartCollectionSemantics.cbig P \<sigma>"
+ and nend: "\<sigma> \<notin> JVMendset"
+ and ics: "ics_of (hd(frames_of \<sigma>)) = Calling Object Cs"
+shows "classes_above P Object \<subseteq> cset \<union> cset'"
+proof(cases "frames_of \<sigma>")
+ case Nil then show ?thesis using nend by(clarsimp simp: JVMendset_def)
+next
+ case (Cons f1 frs1)
+ then obtain stk loc C M pc ics where "f1 = (stk,loc,C,M,pc,ics)" by(cases f1)
+ then show ?thesis
+ using JVMSmartCollectionSemantics.cbig_stepD[OF big nend] nend ics Cons
+ by(clarsimp simp: JVMSmartCollectionSemantics.csmall_def JVMendset_def)
+qed
+
+\<comment> \<open> IH case where @{text C} has not been prepared yet \<close>
+lemma Calling_None_next_state:
+assumes ics: "ics_of (hd(frames_of \<sigma>)) = Calling C Cs"
+ and none: "sheap \<sigma> C = None"
+ and set: "\<forall>C'. P \<turnstile> C \<preceq>\<^sup>* C' \<longrightarrow> (\<exists>sfs i. sheap \<sigma> C' = Some(sfs,i))
+ \<longrightarrow> classes_above P C' \<subseteq> cset"
+ and \<sigma>': "(\<sigma>', cset') \<in> JVMSmartCollectionSemantics.csmall P \<sigma>"
+shows "\<sigma>' \<notin> JVMendset \<and> ics_of (hd(frames_of \<sigma>')) = Calling C Cs
+ \<and> (\<exists>sfs. sheap \<sigma>' C = Some(sfs,Prepared))
+ \<and> (\<forall>C'. P \<turnstile> C \<preceq>\<^sup>* C' \<longrightarrow> C \<noteq> C'
+ \<longrightarrow> (\<exists>sfs i. sheap \<sigma>' C' = Some(sfs,i)) \<longrightarrow> classes_above P C' \<subseteq> cset)"
+proof(cases "frames_of \<sigma> = [] \<or> (\<exists>x. fst \<sigma> = Some x)")
+ case True then show ?thesis using assms
+ by(cases \<sigma>, auto simp: JVMSmartCollectionSemantics.csmall_def)
+next
+ case False
+ then obtain f1 frs1 where frs: "frames_of \<sigma> = f1#frs1" by(cases "frames_of \<sigma>", auto)
+ obtain stk loc C' M pc ics where f1: "f1 = (stk,loc,C',M,pc,ics)" by(cases f1)
+ show ?thesis using f1 frs False assms
+ by(cases \<sigma>, cases "method P C clinit")
+ (clarsimp simp: split_beta JVMSmartCollectionSemantics.csmall_def JVMendset_def)
+qed
+
+\<comment> \<open> IH case where @{text C} has been prepared (and has a direct superclass
+ - i.e., is not @{term Object}) \<close>
+lemma Calling_Prepared_next_state:
+assumes sub: "P \<turnstile> C \<prec>\<^sup>1 D"
+ and obj: "P \<turnstile> D \<preceq>\<^sup>* Object"
+ and ics: "ics_of (hd(frames_of \<sigma>)) = Calling C Cs"
+ and prep: "sheap \<sigma> C = Some(sfs,Prepared)"
+ and set: "\<forall>C'. P \<turnstile> C \<preceq>\<^sup>* C' \<longrightarrow> C \<noteq> C' \<longrightarrow> (\<exists>sfs i. sheap \<sigma> C' = Some(sfs,i))
+ \<longrightarrow> classes_above P C' \<subseteq> cset"
+ and \<sigma>': "(\<sigma>', cset') \<in> JVMSmartCollectionSemantics.csmall P \<sigma>"
+shows "\<sigma>' \<notin> JVMendset \<and> ics_of (hd (frames_of \<sigma>')) = Calling D (C#Cs)
+ \<and> (\<forall>C'. P \<turnstile> D \<preceq>\<^sup>* C' \<longrightarrow> (\<exists>sfs i. sheap \<sigma>' C' = Some(sfs,i))
+ \<longrightarrow> classes_above P C' \<subseteq> cset)"
+using sub
+proof(cases "C=Object")
+ case nobj:False show ?thesis
+ proof(cases "frames_of \<sigma> = [] \<or> (\<exists>x. fst \<sigma> = Some x)")
+ case True then show ?thesis using assms
+ by(cases \<sigma>, auto simp: JVMSmartCollectionSemantics.csmall_def)
+ next
+ case False
+ then obtain f1 frs1 where frs: "frames_of \<sigma> = f1#frs1" by(cases "frames_of \<sigma>", auto)
+ obtain stk loc C' M pc ics where f1: "f1 = (stk,loc,C',M,pc,ics)" by(cases f1)
+ have "C \<noteq> D" using sub obj subcls_self_superclass by auto
+ then have dimp: "\<forall>C'. P \<turnstile> D \<preceq>\<^sup>* C' \<longrightarrow> P \<turnstile> C \<preceq>\<^sup>* C' \<and> C \<noteq> C'"
+ using sub subcls_of_Obj_acyclic[OF obj] by fastforce
+ have "\<forall>C'. P \<turnstile> C \<preceq>\<^sup>* C' \<longrightarrow> C \<noteq> C' \<longrightarrow> (\<exists>sfs i. sheap \<sigma>' C' = Some(sfs,i))
+ \<longrightarrow> classes_above P C' \<subseteq> cset"
+ using f1 frs False nobj assms
+ by(cases \<sigma>, cases "method P C clinit")
+ (auto simp: JVMSmartCollectionSemantics.csmall_def JVMendset_def)
+ then have "\<forall>C'. P \<turnstile> D \<preceq>\<^sup>* C' \<longrightarrow> (\<exists>sfs i. sheap \<sigma>' C' = Some(sfs,i))
+ \<longrightarrow> classes_above P C' \<subseteq> cset" using sub dimp by auto
+ then show ?thesis using f1 frs False nobj assms
+ by(cases \<sigma>, cases "method P C clinit")
+ (auto dest:subcls1D simp: JVMSmartCollectionSemantics.csmall_def JVMendset_def)
+ qed
+qed(simp)
+
+\<comment> \<open> completed IH case: non-@{term Object} (pulls together above IH cases) \<close>
+lemma Calling_collects_IH:
+assumes sub: "P \<turnstile> C \<prec>\<^sup>1 D"
+ and obj: "P \<turnstile> D \<preceq>\<^sup>* Object"
+ and step: "\<And>\<sigma> cset' Cs. (\<sigma>', cset') \<in> JVMSmartCollectionSemantics.cbig P \<sigma> \<Longrightarrow> \<sigma> \<notin> JVMendset
+ \<Longrightarrow> ics_of (hd(frames_of \<sigma>)) = Calling D Cs
+ \<Longrightarrow> \<forall>C'. P \<turnstile> D \<preceq>\<^sup>* C' \<longrightarrow> (\<exists>sfs i. sheap \<sigma> C' = Some(sfs,i))
+ \<longrightarrow> classes_above P C' \<subseteq> cset
+ \<Longrightarrow> classes_above P D \<subseteq> cset \<union> cset'"
+ and big: "(\<sigma>', cset') \<in> JVMSmartCollectionSemantics.cbig P \<sigma>"
+ and nend: "\<sigma> \<notin> JVMendset"
+ and curr: "ics_of (hd(frames_of \<sigma>)) = Calling C Cs"
+ and set: "\<forall>C'. P \<turnstile> C \<preceq>\<^sup>* C' \<longrightarrow> (\<exists>sfs i. sheap \<sigma> C' = Some(sfs,i))
+ \<longrightarrow> classes_above P C' \<subseteq> cset"
+shows "classes_above P C \<subseteq> cset \<union> cset'"
+proof(cases "frames_of \<sigma>")
+ case Nil then show ?thesis using nend by(clarsimp simp: JVMendset_def)
+next
+ case (Cons f1 frs1)
+ show ?thesis using sub
+ proof(cases "\<exists>sfs i. sheap \<sigma> C = Some(sfs,i)")
+ case True then show ?thesis using set by auto
+ next
+ case False
+ obtain stk loc C' M pc ics where f1: "f1 = (stk,loc,C',M,pc,ics)" by(cases f1)
+ then obtain \<sigma>1 coll1 coll where \<sigma>1: "(\<sigma>1, coll1) \<in> JVMSmartCollectionSemantics.csmall P \<sigma>"
+ "cset' = coll1 \<union> coll" "(\<sigma>', coll) \<in> JVMSmartCollectionSemantics.cbig P \<sigma>1"
+ using JVMSmartCollectionSemantics.cbig_stepD[OF big nend] by clarsimp
+ show ?thesis
+ proof(cases "\<exists>sfs. sheap \<sigma> C = Some(sfs,Prepared)")
+ case True
+ then obtain sfs where sfs: "sheap \<sigma> C = Some(sfs,Prepared)" by clarsimp
+ have set': "\<forall>C'. P \<turnstile> C \<preceq>\<^sup>* C' \<longrightarrow> C\<noteq>C' \<longrightarrow> (\<exists>sfs i. sheap \<sigma> C' = Some(sfs,i))
+ \<longrightarrow> classes_above P C' \<subseteq> cset" using set by auto
+ then have "\<sigma>1 \<notin> JVMendset \<and> ics_of (hd (frames_of \<sigma>1)) = Calling D (C#Cs)"
+ "\<forall>C'. P \<turnstile> D \<preceq>\<^sup>* C' \<longrightarrow> (\<exists>sfs i. sheap \<sigma>1 C' = Some(sfs,i))
+ \<longrightarrow> classes_above P C' \<subseteq> cset"
+ using Calling_Prepared_next_state[OF sub obj curr sfs set' \<sigma>1(1)]
+ by(auto simp: JVMSmartCollectionSemantics.csmall_def)
+ then show ?thesis using step[of coll \<sigma>1] classes_above_def2[OF sub] \<sigma>1 f1 Cons nend curr
+ by(clarsimp simp: JVMSmartCollectionSemantics.csmall_def JVMendset_def)
+ next
+ case none: False \<comment> \<open> @{text "Calling C Cs"} is the next @{text ics}, but after that is @{text "Calling D (C#Cs)"} \<close>
+ then have sNone: "sheap \<sigma> C = None" using False by(cases "sheap \<sigma> C", auto)
+ then have nend1: "\<sigma>1 \<notin> JVMendset" and curr1: "ics_of (hd (frames_of \<sigma>1)) = Calling C Cs"
+ and prep: "\<exists>sfs. sheap \<sigma>1 C = \<lfloor>(sfs, Prepared)\<rfloor>"
+ and set1: "\<forall>C'. P \<turnstile> C \<preceq>\<^sup>* C' \<longrightarrow> C \<noteq> C' \<longrightarrow> (\<exists>sfs i. sheap \<sigma>1 C' = \<lfloor>(sfs, i)\<rfloor>)
+ \<longrightarrow> classes_above P C' \<subseteq> cset"
+ using Calling_None_next_state[OF curr sNone set \<sigma>1(1)] by simp+
+ then obtain f2 frs2 where frs2: "frames_of \<sigma>1 = f2#frs2"
+ by(cases \<sigma>1, cases "frames_of \<sigma>1", clarsimp simp: JVMendset_def)
+ obtain sfs1 where sfs1: "sheap \<sigma>1 C = Some(sfs1,Prepared)" using prep by clarsimp
+ obtain stk2 loc2 C2 M2 pc2 ics2 where f2: "f2 = (stk2,loc2,C2,M2,pc2,ics2)" by(cases f2)
+ then obtain \<sigma>2 coll2 coll' where \<sigma>2: "(\<sigma>2, coll2) \<in> JVMSmartCollectionSemantics.csmall P \<sigma>1"
+ "coll = coll2 \<union> coll'" "(\<sigma>', coll') \<in> JVMSmartCollectionSemantics.cbig P \<sigma>2"
+ using JVMSmartCollectionSemantics.cbig_stepD[OF \<sigma>1(3) nend1] by clarsimp
+ then have "\<sigma>2 \<notin> JVMendset \<and> ics_of (hd (frames_of \<sigma>2)) = Calling D (C#Cs)"
+ "\<forall>C'. P \<turnstile> D \<preceq>\<^sup>* C' \<longrightarrow> (\<exists>sfs i. sheap \<sigma>2 C' = Some(sfs,i))
+ \<longrightarrow> classes_above P C' \<subseteq> cset"
+ using Calling_Prepared_next_state[OF sub obj curr1 sfs1 set1 \<sigma>2(1)]
+ by(auto simp: JVMSmartCollectionSemantics.csmall_def)
+ then show ?thesis using step[of coll' \<sigma>2] classes_above_def2[OF sub] \<sigma>2 \<sigma>1 f2 frs2 f1 Cons
+ nend1 nend curr1 curr
+ by(clarsimp simp: JVMSmartCollectionSemantics.csmall_def JVMendset_def)
+ qed
+ qed
+qed
+
+\<comment> \<open>pulls together above base and IH cases \<close>
+lemma Calling_collects:
+assumes sub: "P \<turnstile> C \<preceq>\<^sup>* Object"
+ and "(\<sigma>', cset') \<in> JVMSmartCollectionSemantics.cbig P \<sigma>"
+ and "\<sigma> \<notin> JVMendset"
+ and "ics_of (hd(frames_of \<sigma>)) = Calling C Cs"
+ and "\<forall>C'. P \<turnstile> C \<preceq>\<^sup>* C' \<longrightarrow> (\<exists>sfs i. sheap \<sigma> C' = Some(sfs,i))
+ \<longrightarrow> classes_above P C' \<subseteq> cset"
+ and "cset' \<subseteq> cset"
+shows "classes_above P C \<subseteq> cset"
+proof -
+ have base: "\<forall>\<sigma> cset' Cs.
+ (\<sigma>', cset') \<in> JVMSmartCollectionSemantics.cbig P \<sigma> \<longrightarrow> \<sigma> \<notin> JVMendset
+ \<longrightarrow> ics_of (hd (frames_of \<sigma>)) = Calling Object Cs
+ \<longrightarrow> (\<forall>C'. P \<turnstile> Object \<preceq>\<^sup>* C' \<longrightarrow> (\<exists>sfs i. sheap \<sigma> C' = \<lfloor>(sfs, i)\<rfloor>)
+ \<longrightarrow> classes_above P C' \<subseteq> cset)
+ \<longrightarrow> classes_above P Object \<subseteq> JVMcombine cset cset'" using Calling_collects_base by simp
+ have IH: "\<And>y z. P \<turnstile> y \<prec>\<^sup>1 z \<Longrightarrow>
+ P \<turnstile> z \<preceq>\<^sup>* Object \<Longrightarrow>
+ \<forall>\<sigma> cset' Cs. (\<sigma>', cset') \<in> JVMSmartCollectionSemantics.cbig P \<sigma> \<longrightarrow> \<sigma> \<notin> JVMendset
+ \<longrightarrow> ics_of (hd(frames_of \<sigma>)) = Calling z Cs
+ \<longrightarrow> (\<forall>C'. P \<turnstile> z \<preceq>\<^sup>* C' \<longrightarrow> (\<exists>sfs i. sheap \<sigma> C' = Some(sfs,i))
+ \<longrightarrow> classes_above P C' \<subseteq> cset)
+ \<longrightarrow> classes_above P z \<subseteq> cset \<union> cset' \<Longrightarrow>
+ \<forall>\<sigma> cset' Cs. (\<sigma>', cset') \<in> JVMSmartCollectionSemantics.cbig P \<sigma> \<longrightarrow> \<sigma> \<notin> JVMendset
+ \<longrightarrow> ics_of (hd(frames_of \<sigma>)) = Calling y Cs
+ \<longrightarrow> (\<forall>C'. P \<turnstile> y \<preceq>\<^sup>* C' \<longrightarrow> (\<exists>sfs i. sheap \<sigma> C' = Some(sfs,i))
+ \<longrightarrow> classes_above P C' \<subseteq> cset)
+ \<longrightarrow> classes_above P y \<subseteq> cset \<union> cset'"
+ using Calling_collects_IH by blast
+ have result: "\<forall>\<sigma> cset' Cs.
+ (\<sigma>', cset') \<in> JVMSmartCollectionSemantics.cbig P \<sigma> \<longrightarrow> \<sigma> \<notin> JVMendset
+ \<longrightarrow> ics_of (hd(frames_of \<sigma>)) = Calling C Cs
+ \<longrightarrow> (\<forall>C'. P \<turnstile> C \<preceq>\<^sup>* C' \<longrightarrow> (\<exists>sfs i. sheap \<sigma> C' = Some(sfs,i))
+ \<longrightarrow> classes_above P C' \<subseteq> cset)
+ \<longrightarrow> classes_above P C \<subseteq> cset \<union> cset'"
+ using converse_rtrancl_induct[OF sub,
+ where P="\<lambda>C. \<forall>\<sigma> cset' Cs. (\<sigma>',cset') \<in> JVMSmartCollectionSemantics.cbig P \<sigma> \<longrightarrow> \<sigma> \<notin> JVMendset
+ \<longrightarrow> ics_of (hd(frames_of \<sigma>)) = Calling C Cs
+ \<longrightarrow> (\<forall>C'. P \<turnstile> C \<preceq>\<^sup>* C' \<longrightarrow> (\<exists>sfs i. sheap \<sigma> C' = Some(sfs,i))
+ \<longrightarrow> classes_above P C' \<subseteq> cset)
+ \<longrightarrow> classes_above P C \<subseteq> cset \<union> cset'"]
+ using base IH by blast
+ then show ?thesis using assms by blast
+qed
+
+(*******************)
+text "Instructions that call the initialization procedure will collect classes above
+ the class initialized by the end of execution (using the above @{text Calling_collects})."
+
+lemma New_collects:
+assumes sub: "P \<turnstile> C \<preceq>\<^sup>* Object"
+ and cbig: "(\<sigma>', cset') \<in> JVMSmartCollectionSemantics.cbig P \<sigma>"
+ and nend: "\<sigma> \<notin> JVMendset"
+ and curr: "curr_instr P (hd(frames_of \<sigma>)) = New C"
+ and ics: "ics_of (hd(frames_of \<sigma>)) = No_ics"
+ and sheap: "\<forall>C'. P \<turnstile> C \<preceq>\<^sup>* C' \<longrightarrow> (\<exists>sfs i. sheap \<sigma> C' = Some(sfs,i))
+ \<longrightarrow> classes_above P C' \<subseteq> cset"
+ and smart: "cset' \<subseteq> cset"
+shows "classes_above P C \<subseteq> cset"
+proof(cases "(\<exists>sfs i. sheap \<sigma> C = Some(sfs,i) \<and> i = Done)")
+ case True then show ?thesis using sheap by auto
+next
+ case False
+ obtain n where nstep: "(\<sigma>', cset') \<in> JVMSmartCollectionSemantics.csmall_nstep P \<sigma> n"
+ and "n \<noteq> 0" using nend cbig JVMSmartCollectionSemantics.cbig_def2
+ JVMSmartCollectionSemantics.csmall_nstep_base by (metis empty_iff insert_iff)
+ then show ?thesis
+ proof(cases n)
+ case (Suc n1)
+ then obtain \<sigma>1 cset0 cset1 where \<sigma>1: "(\<sigma>1,cset1) \<in> JVMSmartCollectionSemantics.csmall P \<sigma>"
+ "cset' = cset1 \<union> cset0" "(\<sigma>',cset0) \<in> JVMSmartCollectionSemantics.csmall_nstep P \<sigma>1 n1"
+ using JVMSmartCollectionSemantics.csmall_nstep_SucD nstep by blast
+ obtain xp h frs sh where "\<sigma>=(xp,h,frs,sh)" by(cases \<sigma>)
+ then have ics1: "ics_of (hd(frames_of \<sigma>1)) = Calling C []"
+ and sheap': "sheap \<sigma> = sheap \<sigma>1" and nend1: "\<sigma>1 \<notin> JVMendset"
+ using JVM_New_next_step[OF _ nend curr] \<sigma>1(1) False ics
+ by(simp add: JVMSmartCollectionSemantics.csmall_def)+
+ have "\<sigma>' \<in> JVMendset" using cbig JVMSmartCollectionSemantics.cbig_def2 by blast
+ then have cbig1: "(\<sigma>', cset0) \<in> JVMSmartCollectionSemantics.cbig P \<sigma>1"
+ using JVMSmartCollectionSemantics.cbig_def2 \<sigma>1(3) by blast
+ have sheap1: "\<forall>C'. P \<turnstile> C \<preceq>\<^sup>* C' \<longrightarrow> (\<exists>sfs i. sheap \<sigma>1 C' = \<lfloor>(sfs, i)\<rfloor>)
+ \<longrightarrow> classes_above P C' \<subseteq> cset" using sheap' sheap by simp
+ have "cset0 \<subseteq> cset" using \<sigma>1(2) smart by blast
+ then have "classes_above P C \<subseteq> cset"
+ using Calling_collects[OF sub cbig1 nend1 ics1 sheap1] by simp
+ then show ?thesis using \<sigma>1(2) smart by auto
+ qed(simp)
+qed
+
+lemma Getstatic_collects:
+assumes sub: "P \<turnstile> D \<preceq>\<^sup>* Object"
+ and cbig: "(\<sigma>', cset') \<in> JVMSmartCollectionSemantics.cbig P \<sigma>"
+ and nend: "\<sigma> \<notin> JVMendset"
+ and curr: "curr_instr P (hd(frames_of \<sigma>)) = Getstatic C F D"
+ and ics: "ics_of (hd(frames_of \<sigma>)) = No_ics"
+ and fC: "P \<turnstile> C has F,Static:t in D"
+ and sheap: "\<forall>C'. P \<turnstile> D \<preceq>\<^sup>* C' \<longrightarrow> (\<exists>sfs i. sheap \<sigma> C' = Some(sfs,i))
+ \<longrightarrow> classes_above P C' \<subseteq> cset"
+ and smart: "cset' \<subseteq> cset"
+shows "classes_above P D \<subseteq> cset"
+proof(cases "(\<exists>sfs i. sheap \<sigma> D = Some(sfs,i) \<and> i = Done)
+ \<or> (ics_of(hd(frames_of \<sigma>)) = Called [])")
+ case True then show ?thesis
+ proof(cases "\<exists>sfs i. sheap \<sigma> D = Some(sfs,i) \<and> i = Done")
+ case True then show ?thesis using sheap by auto
+ next
+ case False
+ then have "ics_of(hd(frames_of \<sigma>)) = Called []" using True by clarsimp
+ then show ?thesis using ics by auto
+ qed
+next
+ case False
+ obtain n where nstep: "(\<sigma>', cset') \<in> JVMSmartCollectionSemantics.csmall_nstep P \<sigma> n"
+ and "n \<noteq> 0" using nend cbig JVMSmartCollectionSemantics.cbig_def2
+ JVMSmartCollectionSemantics.csmall_nstep_base by (metis empty_iff insert_iff)
+ then show ?thesis
+ proof(cases n)
+ case (Suc n1)
+ then obtain \<sigma>1 cset0 cset1 where \<sigma>1: "(\<sigma>1,cset1) \<in> JVMSmartCollectionSemantics.csmall P \<sigma>"
+ "cset' = cset1 \<union> cset0" "(\<sigma>',cset0) \<in> JVMSmartCollectionSemantics.csmall_nstep P \<sigma>1 n1"
+ using JVMSmartCollectionSemantics.csmall_nstep_SucD nstep by blast
+ obtain xp h frs sh where "\<sigma>=(xp,h,frs,sh)" by(cases \<sigma>)
+ then have curr1: "ics_of (hd(frames_of \<sigma>1)) = Calling D []"
+ and sheap': "sheap \<sigma> = sheap \<sigma>1" and nend1: "\<sigma>1 \<notin> JVMendset"
+ using JVM_Getstatic_next_step[OF _ nend curr fC] \<sigma>1(1) False ics
+ by(simp add: JVMSmartCollectionSemantics.csmall_def)+
+ have "\<sigma>' \<in> JVMendset" using cbig JVMSmartCollectionSemantics.cbig_def2 by blast
+ then have cbig1: "(\<sigma>', cset0) \<in> JVMSmartCollectionSemantics.cbig P \<sigma>1"
+ using JVMSmartCollectionSemantics.cbig_def2 \<sigma>1(3) by blast
+ have sheap1: "\<forall>C'. P \<turnstile> D \<preceq>\<^sup>* C' \<longrightarrow> (\<exists>sfs i. sheap \<sigma>1 C' = \<lfloor>(sfs, i)\<rfloor>)
+ \<longrightarrow> classes_above P C' \<subseteq> cset" using sheap' sheap by simp
+ have "cset0 \<subseteq> cset" using \<sigma>1(2) smart by blast
+ then have "classes_above P D \<subseteq> cset"
+ using Calling_collects[OF sub cbig1 nend1 curr1 sheap1] by simp
+ then show ?thesis using \<sigma>1(2) smart by auto
+ qed(simp)
+qed
+
+lemma Putstatic_collects:
+assumes sub: "P \<turnstile> D \<preceq>\<^sup>* Object"
+ and cbig: "(\<sigma>', cset') \<in> JVMSmartCollectionSemantics.cbig P \<sigma>"
+ and nend: "\<sigma> \<notin> JVMendset"
+ and curr: "curr_instr P (hd(frames_of \<sigma>)) = Putstatic C F D"
+ and ics: "ics_of (hd(frames_of \<sigma>)) = No_ics"
+ and fC: "P \<turnstile> C has F,Static:t in D"
+ and sheap: "\<forall>C'. P \<turnstile> D \<preceq>\<^sup>* C' \<longrightarrow> (\<exists>sfs i. sheap \<sigma> C' = Some(sfs,i))
+ \<longrightarrow> classes_above P C' \<subseteq> cset"
+ and smart: "cset' \<subseteq> cset"
+shows "classes_above P D \<subseteq> cset"
+proof(cases "(\<exists>sfs i. sheap \<sigma> D = Some(sfs,i) \<and> i = Done)
+ \<or> (ics_of(hd(frames_of \<sigma>)) = Called [])")
+ case True then show ?thesis
+ proof(cases "\<exists>sfs i. sheap \<sigma> D = Some(sfs,i) \<and> i = Done")
+ case True then show ?thesis using sheap by auto
+ next
+ case False
+ then have "ics_of(hd(frames_of \<sigma>)) = Called []" using True by clarsimp
+ then show ?thesis using ics by auto
+ qed
+next
+ case False
+ obtain n where nstep: "(\<sigma>', cset') \<in> JVMSmartCollectionSemantics.csmall_nstep P \<sigma> n"
+ and "n \<noteq> 0" using nend cbig JVMSmartCollectionSemantics.cbig_def2
+ JVMSmartCollectionSemantics.csmall_nstep_base by (metis empty_iff insert_iff)
+ then show ?thesis
+ proof(cases n)
+ case (Suc n1)
+ then obtain \<sigma>1 cset0 cset1 where \<sigma>1: "(\<sigma>1,cset1) \<in> JVMSmartCollectionSemantics.csmall P \<sigma>"
+ "cset' = cset1 \<union> cset0" "(\<sigma>',cset0) \<in> JVMSmartCollectionSemantics.csmall_nstep P \<sigma>1 n1"
+ using JVMSmartCollectionSemantics.csmall_nstep_SucD nstep by blast
+ obtain xp h frs sh where "\<sigma>=(xp,h,frs,sh)" by(cases \<sigma>)
+ then have curr1: "ics_of (hd(frames_of \<sigma>1)) = Calling D []"
+ and sheap': "sheap \<sigma> = sheap \<sigma>1" and nend1: "\<sigma>1 \<notin> JVMendset"
+ using JVM_Putstatic_next_step[OF _ nend curr fC] \<sigma>1(1) False ics
+ by(simp add: JVMSmartCollectionSemantics.csmall_def)+
+ have "\<sigma>' \<in> JVMendset" using cbig JVMSmartCollectionSemantics.cbig_def2 by blast
+ then have cbig1: "(\<sigma>', cset0) \<in> JVMSmartCollectionSemantics.cbig P \<sigma>1"
+ using JVMSmartCollectionSemantics.cbig_def2 \<sigma>1(3) by blast
+ have sheap1: "\<forall>C'. P \<turnstile> D \<preceq>\<^sup>* C' \<longrightarrow> (\<exists>sfs i. sheap \<sigma>1 C' = \<lfloor>(sfs, i)\<rfloor>)
+ \<longrightarrow> classes_above P C' \<subseteq> cset" using sheap' sheap by simp
+ have "cset0 \<subseteq> cset" using \<sigma>1(2) smart by blast
+ then have "classes_above P D \<subseteq> cset"
+ using Calling_collects[OF sub cbig1 nend1 curr1 sheap1] by simp
+ then show ?thesis using \<sigma>1(2) smart by auto
+ qed(simp)
+qed
+
+lemma Invokestatic_collects:
+assumes sub: "P \<turnstile> D \<preceq>\<^sup>* Object"
+ and cbig: "(\<sigma>', cset') \<in> JVMSmartCollectionSemantics.cbig P \<sigma>"
+ and smart: "cset' \<subseteq> cset"
+ and nend: "\<sigma> \<notin> JVMendset"
+ and curr: "curr_instr P (hd(frames_of \<sigma>)) = Invokestatic C M n"
+ and ics: "ics_of (hd(frames_of \<sigma>)) = No_ics"
+ and mC: "P \<turnstile> C sees M,Static:Ts \<rightarrow> T = m in D"
+ and sheap: "\<forall>C'. P \<turnstile> D \<preceq>\<^sup>* C' \<longrightarrow> (\<exists>sfs i. sheap \<sigma> C' = Some(sfs,i))
+ \<longrightarrow> classes_above P C' \<subseteq> cset"
+shows "classes_above P D \<subseteq> cset"
+proof(cases "(\<exists>sfs i. sheap \<sigma> D = Some(sfs,i) \<and> i = Done)
+ \<or> (ics_of(hd(frames_of \<sigma>)) = Called [])")
+ case True then show ?thesis
+ proof(cases "\<exists>sfs i. sheap \<sigma> D = Some(sfs,i) \<and> i = Done")
+ case True then show ?thesis using sheap by auto
+ next
+ case False
+ then have "ics_of(hd(frames_of \<sigma>)) = Called []" using True by clarsimp
+ then show ?thesis using ics by auto
+ qed
+next
+ case False
+ obtain n where nstep: "(\<sigma>', cset') \<in> JVMSmartCollectionSemantics.csmall_nstep P \<sigma> n"
+ and "n \<noteq> 0" using nend cbig JVMSmartCollectionSemantics.cbig_def2
+ JVMSmartCollectionSemantics.csmall_nstep_base by (metis empty_iff insert_iff)
+ then show ?thesis
+ proof(cases n)
+ case (Suc n1)
+ then obtain \<sigma>1 cset0 cset1 where \<sigma>1: "(\<sigma>1,cset1) \<in> JVMSmartCollectionSemantics.csmall P \<sigma>"
+ "cset' = cset1 \<union> cset0" "(\<sigma>',cset0) \<in> JVMSmartCollectionSemantics.csmall_nstep P \<sigma>1 n1"
+ using JVMSmartCollectionSemantics.csmall_nstep_SucD nstep by blast
+ obtain xp h frs sh where "\<sigma>=(xp,h,frs,sh)" by(cases \<sigma>)
+ then have curr1: "ics_of (hd(frames_of \<sigma>1)) = Calling D []"
+ and sheap': "sheap \<sigma> = sheap \<sigma>1" and nend1: "\<sigma>1 \<notin> JVMendset"
+ using JVM_Invokestatic_next_step[OF _ nend curr mC] \<sigma>1(1) False ics
+ by(simp add: JVMSmartCollectionSemantics.csmall_def)+
+ have "\<sigma>' \<in> JVMendset" using cbig JVMSmartCollectionSemantics.cbig_def2 by blast
+ then have cbig1: "(\<sigma>', cset0) \<in> JVMSmartCollectionSemantics.cbig P \<sigma>1"
+ using JVMSmartCollectionSemantics.cbig_def2 \<sigma>1(3) by blast
+ have sheap1: "\<forall>C'. P \<turnstile> D \<preceq>\<^sup>* C' \<longrightarrow> (\<exists>sfs i. sheap \<sigma>1 C' = \<lfloor>(sfs, i)\<rfloor>)
+ \<longrightarrow> classes_above P C' \<subseteq> cset" using sheap' sheap by simp
+ have "cset0 \<subseteq> cset" using \<sigma>1(2) smart by blast
+ then have "classes_above P D \<subseteq> cset"
+ using Calling_collects[OF sub cbig1 nend1 curr1 sheap1] by simp
+ then show ?thesis using \<sigma>1(2) smart by auto
+ qed(simp)
+qed
+
+(*********)
+
+text "The @{text smart_out} execution function keeps the promise to
+ collect above the initial class (@{term Test}):"
+lemma jvm_smart_out_classes_above_Test:
+assumes s: "(\<sigma>',cset\<^sub>s) \<in> jvm_smart_out P t" and P: "P \<in> jvm_progs" and t: "t \<in> jvm_tests"
+shows "classes_above (jvm_make_test_prog P t) Test \<subseteq> cset\<^sub>s"
+ (is "classes_above ?P ?D \<subseteq> ?cset")
+proof -
+ let ?\<sigma> = "start_state (t#P)" and ?M = main
+ let ?ics = "ics_of (hd(frames_of ?\<sigma>))"
+ have called: "?ics = Called [] \<Longrightarrow> classes_above ?P ?D \<subseteq> ?cset"
+ by(simp add: start_state_def)
+ then show ?thesis
+ proof(cases "?ics = Called []")
+ case True then show ?thesis using called by simp
+ next
+ case False
+ from P t obtain wf_md where wf: "wf_prog wf_md (t#P)"
+ by(auto simp: wf_jvm_prog_phi_def wf_jvm_prog_def)
+ from jvm_make_test_prog_sees_Test_main[OF P t] obtain m where
+ mC: "?P \<turnstile> ?D sees ?M,Static:[] \<rightarrow> Void = m in ?D" by fast
+ (****)
+ then have "?P \<turnstile> ?D \<preceq>\<^sup>* Object" by(rule sees_method_sub_Obj)
+ moreover from s obtain cset' where
+ cbig: "(\<sigma>', cset') \<in> JVMSmartCollectionSemantics.cbig ?P ?\<sigma>" and "cset' \<subseteq> ?cset" by clarsimp
+ moreover have nend: "?\<sigma> \<notin> JVMendset" by(rule start_state_nend)
+ moreover from start_prog_start_m_instrs[OF wf] t
+ have curr: "curr_instr ?P (hd(frames_of ?\<sigma>)) = Invokestatic ?D ?M 0"
+ by(simp add: start_state_def)
+ moreover have ics: "?ics = No_ics"
+ by(simp add: start_state_def)
+ moreover note mC
+ moreover from jvm_smart_out_classes_above_start_sheap[OF s]
+ have sheap: "\<forall>C'. ?P \<turnstile> ?D \<preceq>\<^sup>* C' \<longrightarrow> (\<exists>sfs i. sheap ?\<sigma> C' = Some(sfs,i))
+ \<longrightarrow> classes_above ?P C' \<subseteq> ?cset" by(simp add: start_state_def)
+ ultimately show ?thesis by(rule Invokestatic_collects)
+ qed
+qed
+
+(**********************************************)
+text "Using lemmas proving preservation of backward promises and keeping
+ of forward promises, we prove that the smart algorithm collects at least
+ the classes as the naive algorithm does."
+
+\<comment> \<open> first over a single execution step (assumes promises) \<close>
+lemma jvm_naive_to_smart_exec_collect:
+assumes
+\<comment> \<open> well-formedness \<close>
+ wtp: "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P"
+ and correct: "P,\<Phi> \<turnstile> (xp,h,frs,sh)\<surd>"
+\<comment> \<open> defs \<close>
+ and f': "hd frs = (stk,loc,C',M',pc,ics)"
+\<comment> \<open> backward promises - will be collected prior \<close>
+ and heap: "\<And>C fs. \<exists>a. h a = Some(C,fs) \<Longrightarrow> classes_above P C \<subseteq> cset"
+ and sheap: "\<And>C sfs i. sh C = Some(sfs,i) \<Longrightarrow> classes_above P C \<subseteq> cset"
+ and xcpts: "classes_above_xcpts P \<subseteq> cset"
+ and frames: "classes_above_frames P frs \<subseteq> cset"
+\<comment> \<open> forward promises - will be collected after if not already \<close>
+ and init_class_prom: "\<And>C. ics = Called [] \<or> ics = No_ics
+ \<Longrightarrow> coll_init_class P (instrs_of P C' M' ! pc) = Some C \<Longrightarrow> classes_above P C \<subseteq> cset"
+ and Calling_prom: "\<And>C' Cs'. ics = Calling C' Cs' \<Longrightarrow> classes_above P C' \<subseteq> cset"
+\<comment> \<open> collection \<close>
+ and smart: "JVMexec_scollect P (xp,h,frs,sh) \<subseteq> cset"
+shows "JVMexec_ncollect P (xp,h,frs,sh) \<subseteq> cset"
+using assms
+proof(cases frs)
+ case (Cons f' frs')
+ then have [simp]: "classes_above P C' \<subseteq> cset" using frames f' by simp
+ let ?i = "instrs_of P C' M' ! pc"
+ have cr': "P,\<Phi> \<turnstile> (xp,h,(stk,loc,C',M',pc,ics)#frs',sh)\<surd>" using correct f' Cons by simp
+ from well_formed_stack_safe[OF wtp cr'] correct_state_Throwing_ex[OF cr'] obtain
+ stack_safe: "stack_safe ?i h stk" and
+ Throwing_ex: "\<And>Cs a. ics = Throwing Cs a \<Longrightarrow> \<exists>obj. h a = Some obj" by simp
+ have confc: "conf_clinit P sh frs" using correct Cons by simp
+ have Called_prom: "\<And>C' Cs'. ics = Called (C'#Cs')
+ \<Longrightarrow> classes_above P C' \<subseteq> cset \<and> classes_above P (fst(method P C' clinit)) \<subseteq> cset"
+ proof -
+ fix C' Cs' assume [simp]: "ics = Called (C'#Cs')"
+ then have "C' \<in> set(clinit_classes frs)" using f' Cons by simp
+ then obtain sfs where shC': "sh C' = Some(sfs, Processing)" and "is_class P C'"
+ using confc by(auto simp: conf_clinit_def)
+ then have C'eq: "C' = fst(method P C' clinit)" using wf_sees_clinit wtp
+ by(fastforce simp: is_class_def wf_jvm_prog_phi_def)
+ then show "classes_above P C' \<subseteq> cset \<and> classes_above P (fst(method P C' clinit)) \<subseteq> cset"
+ using sheap shC' by auto
+ qed
+ show ?thesis using Cons assms
+ proof(cases xp)
+ case None
+ { assume ics: "ics = Called [] \<or> ics = No_ics"
+ then have [simp]: "JVMexec_ncollect P (xp,h,frs,sh)
+ = JVMinstr_ncollect P ?i h stk \<union> classes_above P C'
+ \<union> classes_above_frames P frs \<union> classes_above_xcpts P"
+ and [simp]: "JVMexec_scollect P (xp,h,frs,sh) = JVMinstr_scollect P ?i"
+ using f' None Cons by auto
+ have ?thesis using assms
+ proof(cases ?i)
+ case (New C)
+ then have "classes_above P C \<subseteq> cset" using ics New assms by simp
+ then show ?thesis using New xcpts frames smart f' by auto
+ next
+ case (Getfield F C) show ?thesis
+ proof(cases "hd stk = Null")
+ case True then show ?thesis using Getfield assms by simp
+ next
+ case False
+ let ?C = "cname_of h (the_Addr (hd stk))"
+ have above_stk: "classes_above P ?C \<subseteq> cset"
+ using stack_safe heap f' False Cons Getfield by(auto dest!: is_ptrD) blast
+ then show ?thesis using Getfield assms by auto
+ qed
+ next
+ case (Getstatic C F D)
+ show ?thesis
+ proof(cases "\<exists>t. P \<turnstile> C has F,Static:t in D")
+ case True
+ then have above_D: "classes_above P D \<subseteq> cset" using ics init_class_prom Getstatic by simp
+ have sub: "P \<turnstile> C \<preceq>\<^sup>* D" using has_field_decl_above True by blast
+ then have above_C: "classes_between P C D - {D} \<subseteq> cset"
+ using True Getstatic above_D smart f' by simp
+ then have "classes_above P C \<subseteq> cset"
+ using classes_above_sub_classes_between_eq[OF sub] above_D above_C by auto
+ then show ?thesis using Getstatic assms by auto
+ next
+ case False then show ?thesis using Getstatic assms by auto
+ qed
+ next
+ case (Putfield F C) show ?thesis
+ proof(cases "hd(tl stk) = Null")
+ case True then show ?thesis using Putfield assms by simp
+ next
+ case False
+ let ?C = "cname_of h (the_Addr (hd (tl stk)))"
+ have above_stk: "classes_above P ?C \<subseteq> cset"
+ using stack_safe heap f' False Cons Putfield by(auto dest!: is_ptrD) blast
+ then show ?thesis using Putfield assms by auto
+ qed
+ next
+ case (Putstatic C F D)
+ show ?thesis
+ proof(cases "\<exists>t. P \<turnstile> C has F,Static:t in D")
+ case True
+ then have above_D: "classes_above P D \<subseteq> cset" using ics init_class_prom Putstatic by simp
+ have sub: "P \<turnstile> C \<preceq>\<^sup>* D" using has_field_decl_above True by blast
+ then have above_C: "classes_between P C D - {D} \<subseteq> cset"
+ using True Putstatic above_D smart f' by simp
+ then have "classes_above P C \<subseteq> cset"
+ using classes_above_sub_classes_between_eq[OF sub] above_D above_C by auto
+ then show ?thesis using Putstatic assms by auto
+ next
+ case False then show ?thesis using Putstatic assms by auto
+ qed
+ next
+ case (Checkcast C) show ?thesis
+ proof(cases "hd stk = Null")
+ case True then show ?thesis using Checkcast assms by simp
+ next
+ case False
+ let ?C = "cname_of h (the_Addr (hd stk))"
+ have above_stk: "classes_above P ?C \<subseteq> cset"
+ using stack_safe heap False Cons f' Checkcast by(auto dest!: is_ptrD) blast
+ then show ?thesis using above_stk Checkcast assms by(cases "hd stk = Null", auto)
+ qed
+ next
+ case (Invoke M n) show ?thesis
+ proof(cases "stk ! n = Null")
+ case True then show ?thesis using Invoke assms by simp
+ next
+ case False
+ let ?C = "cname_of h (the_Addr (stk ! n))"
+ have above_stk: "classes_above P ?C \<subseteq> cset" using stack_safe heap False Cons f' Invoke
+ by(auto dest!: is_ptrD) blast
+ then show ?thesis using Invoke assms by auto
+ qed
+ next
+ case (Invokestatic C M n)
+ let ?D = "fst (method P C M)"
+ show ?thesis
+ proof(cases "\<exists>Ts T m D. P \<turnstile> C sees M,Static:Ts \<rightarrow> T = m in D")
+ case True
+ then have above_D: "classes_above P ?D \<subseteq> cset" using ics init_class_prom Invokestatic
+ by(simp add: seeing_class_def)
+ have sub: "P \<turnstile> C \<preceq>\<^sup>* ?D" using method_def2 sees_method_decl_above True by auto
+ then show ?thesis
+ proof(cases "C = ?D")
+ case True then show ?thesis
+ using Invokestatic above_D xcpts frames smart f' by auto
+ next
+ case False
+ then have above_C: "classes_between P C ?D - {?D} \<subseteq> cset"
+ using True Invokestatic above_D smart f' by simp
+ then have "classes_above P C \<subseteq> cset"
+ using classes_above_sub_classes_between_eq[OF sub] above_D above_C by auto
+ then show ?thesis using Invokestatic assms by auto
+ qed
+ next
+ case False then show ?thesis using Invokestatic assms by auto
+ qed
+ next
+ case Throw show ?thesis
+ proof(cases "hd stk = Null")
+ case True then show ?thesis using Throw assms by simp
+ next
+ case False
+ let ?C = "cname_of h (the_Addr (hd stk))"
+ have above_stk: "classes_above P ?C \<subseteq> cset"
+ using stack_safe heap False Cons f' Throw by(auto dest!: is_ptrD) blast
+ then show ?thesis using above_stk Throw assms by auto
+ qed
+ next
+ case Load then show ?thesis using assms by auto
+ next
+ case Store then show ?thesis using assms by auto
+ next
+ case Push then show ?thesis using assms by auto
+ next
+ case Goto then show ?thesis using assms by auto
+ next
+ case IfFalse then show ?thesis using assms by auto
+ qed(auto)
+ }
+ moreover
+ { fix C1 Cs1 assume ics: "ics = Called (C1#Cs1)"
+ then have ?thesis using None Cons Called_prom[OF ics] xcpts frames f' by simp
+ }
+ moreover
+ { fix Cs1 a assume ics: "ics = Throwing Cs1 a"
+ then obtain C fs where "h a = Some(C,fs)" using Throwing_ex by fastforce
+ then have above_stk: "classes_above P (cname_of h a) \<subseteq> cset" using heap by auto
+ then have ?thesis using ics None Cons xcpts frames f' by simp
+ }
+ moreover
+ { fix C1 Cs1 assume ics: "ics = Calling C1 Cs1"
+ then have ?thesis using None Cons Calling_prom[OF ics] xcpts frames f' by simp
+ }
+ ultimately show ?thesis by (metis ics_classes.cases list.exhaust)
+ qed(simp)
+qed(simp)
+
+\<comment> \<open> ... which is the same as @{term csmall} \<close>
+lemma jvm_naive_to_smart_csmall:
+assumes
+\<comment> \<open> well-formedness \<close>
+ wtp: "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P"
+ and correct: "P,\<Phi> \<turnstile> (xp,h,frs,sh)\<surd>"
+\<comment> \<open> defs \<close>
+ and f': "hd frs = (stk,loc,C',M',pc,ics)"
+\<comment> \<open> backward promises - will be collected prior \<close>
+ and heap: "\<And>C fs. \<exists>a. h a = Some(C,fs) \<Longrightarrow> classes_above P C \<subseteq> cset"
+ and sheap: "\<And>C sfs i. sh C = Some(sfs,i) \<Longrightarrow> classes_above P C \<subseteq> cset"
+ and xcpts: "classes_above_xcpts P \<subseteq> cset"
+ and frames: "classes_above_frames P frs \<subseteq> cset"
+\<comment> \<open> forward promises - will be collected after if not already \<close>
+ and init_class_prom: "\<And>C. ics = Called [] \<or> ics = No_ics
+ \<Longrightarrow> coll_init_class P (instrs_of P C' M' ! pc) = Some C \<Longrightarrow> classes_above P C \<subseteq> cset"
+ and Calling_prom: "\<And>C' Cs'. ics = Calling C' Cs' \<Longrightarrow> classes_above P C' \<subseteq> cset"
+\<comment> \<open> collections \<close>
+ and smart_coll: "(\<sigma>', cset\<^sub>s) \<in> JVMSmartCollectionSemantics.csmall P (xp,h,frs,sh)"
+ and naive_coll: "(\<sigma>', cset\<^sub>n) \<in> JVMNaiveCollectionSemantics.csmall P (xp,h,frs,sh)"
+ and smart: "cset\<^sub>s \<subseteq> cset"
+shows "cset\<^sub>n \<subseteq> cset"
+using jvm_naive_to_smart_exec_collect[where h=h and sh=sh, OF assms(1-9)]
+ smart smart_coll naive_coll
+ by(fastforce simp: JVMNaiveCollectionSemantics.csmall_def
+ JVMSmartCollectionSemantics.csmall_def)
+
+\<comment> \<open> ...which means over @{term csmall_nstep}, stepping from the end state
+ (the point by which future promises will have been fulfilled) (uses backward
+ and forward promise lemmas) \<close>
+lemma jvm_naive_to_smart_csmall_nstep:
+"\<lbrakk> wf_jvm_prog\<^bsub>\<Phi>\<^esub> P;
+ P,\<Phi> \<turnstile> (xp,h,frs,sh)\<surd>;
+ hd frs = (stk,loc,C',M',pc,ics);
+ \<And>C fs. \<exists>a. h a = Some(C,fs) \<Longrightarrow> classes_above P C \<subseteq> cset;
+ \<And>C sfs i. sh C = Some(sfs,i) \<Longrightarrow> classes_above P C \<subseteq> cset;
+ classes_above_xcpts P \<subseteq> cset;
+ classes_above_frames P frs \<subseteq> cset;
+ \<And>C. ics = Called [] \<or> ics = No_ics
+ \<Longrightarrow> coll_init_class P (instrs_of P C' M' ! pc) = Some C \<Longrightarrow> classes_above P C \<subseteq> cset;
+ \<And>C' Cs'. ics = Calling C' Cs' \<Longrightarrow> classes_above P C' \<subseteq> cset;
+ (\<sigma>', cset\<^sub>n) \<in> JVMNaiveCollectionSemantics.csmall_nstep P (xp,h,frs,sh) n;
+ (\<sigma>', cset\<^sub>s) \<in> JVMSmartCollectionSemantics.csmall_nstep P (xp,h,frs,sh) n;
+ cset\<^sub>s \<subseteq> cset;
+ \<sigma>' \<in> JVMendset \<rbrakk>
+ \<Longrightarrow> cset\<^sub>n \<subseteq> cset"
+proof(induct n arbitrary: xp h frs sh stk loc C' M' pc ics \<sigma>' cset\<^sub>n cset\<^sub>s cset)
+ case 0 then show ?case
+ using JVMNaiveCollectionSemantics.csmall_nstep_base subsetI old.prod.inject singletonD
+ by (metis (no_types, lifting) equals0D)
+next
+ case (Suc n1)
+ let ?\<sigma> = "(xp,h,frs,sh)"
+ obtain \<sigma>1 cset1 cset' where \<sigma>1: "(\<sigma>1, cset1) \<in> JVMNaiveCollectionSemantics.csmall P ?\<sigma>"
+ "cset\<^sub>n = cset1 \<union> cset'" "(\<sigma>', cset') \<in> JVMNaiveCollectionSemantics.csmall_nstep P \<sigma>1 n1"
+ using JVMNaiveCollectionSemantics.csmall_nstep_SucD[OF Suc.prems(10)] by clarsimp+
+ obtain \<sigma>1' cset1' cset'' where \<sigma>1': "(\<sigma>1', cset1') \<in> JVMSmartCollectionSemantics.csmall P ?\<sigma>"
+ "cset\<^sub>s = cset1' \<union> cset''" "(\<sigma>', cset'') \<in> JVMSmartCollectionSemantics.csmall_nstep P \<sigma>1' n1"
+ using JVMSmartCollectionSemantics.csmall_nstep_SucD[OF Suc.prems(11)] by clarsimp+
+ have \<sigma>_eq: "\<sigma>1 = \<sigma>1'" using \<sigma>1(1) \<sigma>1'(1) by(simp add: JVMNaiveCollectionSemantics.csmall_def
+ JVMSmartCollectionSemantics.csmall_def)
+ have sub1': "cset1' \<subseteq> cset" and sub'': "cset'' \<subseteq> cset" using Suc.prems(12) \<sigma>1'(2) by auto
+ then have sub1: "cset1 \<subseteq> cset"
+ using jvm_naive_to_smart_csmall[where h=h and sh=sh and \<sigma>'=\<sigma>1, OF Suc.prems(1-9) _ _ sub1']
+ Suc.prems(11,12) \<sigma>1(1) \<sigma>1'(1) \<sigma>_eq by fastforce
+ show ?case
+ proof(cases n1)
+ case 0 then show ?thesis using \<sigma>1(2,3) sub1 by auto
+ next
+ case Suc2: (Suc n2)
+ then have nend1: "\<sigma>1 \<notin> JVMendset"
+ using JVMNaiveCollectionSemantics.csmall_nstep_Suc_nend \<sigma>1(3) by blast
+ obtain xp1 h1 frs1 sh1 where \<sigma>1_case [simp]: "\<sigma>1 = (xp1,h1,frs1,sh1)" by(cases \<sigma>1)
+ obtain stk1 loc1 C1' M1' pc1 ics1 where f1': "hd frs1 = (stk1,loc1,C1',M1',pc1,ics1)"
+ by(cases "hd frs1")
+ then obtain frs1' where [simp]: "frs1 = (stk1,loc1,C1',M1',pc1,ics1)#frs1'"
+ using JVMendset_def nend1 by(cases frs1, auto)
+ have cbig1: "(\<sigma>', cset') \<in> JVMNaiveCollectionSemantics.cbig P \<sigma>1"
+ "(\<sigma>', cset'') \<in> JVMSmartCollectionSemantics.cbig P \<sigma>1" using \<sigma>1(3) \<sigma>1'(3) Suc.prems(13) \<sigma>_eq
+ using JVMNaiveCollectionSemantics.cbig_def2
+ JVMSmartCollectionSemantics.cbig_def2 by blast+
+ obtain \<sigma>2' cset2' cset2'' where \<sigma>2': "(\<sigma>2', cset2') \<in> JVMSmartCollectionSemantics.csmall P \<sigma>1"
+ "cset'' = cset2' \<union> cset2''" "(\<sigma>', cset2'') \<in> JVMSmartCollectionSemantics.csmall_nstep P \<sigma>2' n2"
+ using JVMSmartCollectionSemantics.csmall_nstep_SucD \<sigma>1'(3) Suc2 \<sigma>_eq by blast
+(*****)
+ have wtp: "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P" by fact
+ let ?i1 = "instrs_of P C1' M1' ! pc1"
+ let ?ics1 = "ics_of (hd (frames_of \<sigma>1))"
+ have step: "P \<turnstile> (xp,h,frs,sh) -jvm\<rightarrow> (xp1,h1,frs1,sh1)"
+ proof -
+ have "exec (P, ?\<sigma>) = \<lfloor>\<sigma>1'\<rfloor>" using JVMsmart_csmallD[OF \<sigma>1'(1)] by simp
+ then have "P \<turnstile> ?\<sigma> -jvm\<rightarrow> \<sigma>1'" using jvm_one_step1[OF exec_1.exec_1I] by simp
+ then show ?thesis using Suc.prems(12) \<sigma>_eq by fastforce
+ qed
+ have correct1: "P,\<Phi> \<turnstile> (xp1,h1,frs1,sh1)\<surd>" by(rule BV_correct[OF wtp step Suc.prems(2)])
+(**)
+ have vics1: "P,h1,sh1 \<turnstile>\<^sub>i (C1', M1', pc1, ics1)"
+ using correct1 Suc.prems(7) by(auto simp: conf_f_def2)
+ from correct1 obtain b Ts T mxs mxl\<^sub>0 ins xt ST LT where
+ meth1: "P \<turnstile> C1' sees M1',b:Ts\<rightarrow>T=(mxs,mxl\<^sub>0,ins,xt) in C1'" and
+ pc1: "pc1 < length ins" and
+ \<Phi>_pc1: "\<Phi> C1' M1'!pc1 = Some (ST,LT)" by(auto dest: sees_method_fun)
+ then have wt1: "P,T,mxs,size ins,xt \<turnstile> ins!pc1,pc1 :: \<Phi> C1' M1'"
+ using wt_jvm_prog_impl_wt_instr[OF wtp meth1] by fast
+(**)
+ have "\<And>a C fs sfs' i'. (h1 a = \<lfloor>(C, fs)\<rfloor> \<longrightarrow> classes_above P C \<subseteq> cset) \<and>
+ (sh1 C = \<lfloor>(sfs', i')\<rfloor> \<longrightarrow> classes_above P C \<subseteq> cset) \<and>
+ classes_above_frames P frs1 \<subseteq> cset"
+ proof -
+ fix a C fs sfs' i'
+ show "(h1 a = \<lfloor>(C, fs)\<rfloor> \<longrightarrow> classes_above P C \<subseteq> cset) \<and>
+ (sh1 C = \<lfloor>(sfs', i')\<rfloor> \<longrightarrow> classes_above P C \<subseteq> cset) \<and>
+ (classes_above_frames P frs1 \<subseteq> cset)"
+ using Suc.prems(11-12) \<sigma>1' \<sigma>_eq[THEN sym] JVMsmart_csmallD[OF \<sigma>1'(1)]
+ backward_coll_promises_kept[where h=h and xp=xp and sh=sh and frs=frs and frs'=frs1
+ and xp'=xp1 and h'=h1 and sh'=sh1, OF Suc.prems(1-9)] by auto
+ qed
+ then have heap1: "\<And>C fs. \<exists>a. h1 a = Some(C,fs) \<Longrightarrow> classes_above P C \<subseteq> cset"
+ and sheap1: "\<And>C sfs i. sh1 C = Some(sfs,i) \<Longrightarrow> classes_above P C \<subseteq> cset"
+ and frames1: "classes_above_frames P frs1 \<subseteq> cset" by blast+
+ have xcpts1: "classes_above_xcpts P \<subseteq> cset" using Suc.prems(6) by auto
+\<comment> \<open> @{text init_class} promise \<close>
+ have sheap2: "\<And>C. coll_init_class P ?i1 = Some C
+ \<Longrightarrow> \<forall>C'. P \<turnstile> C \<preceq>\<^sup>* C' \<longrightarrow> (\<exists>sfs i. sheap \<sigma>1 C' = \<lfloor>(sfs, i)\<rfloor>)
+ \<longrightarrow> classes_above P C' \<subseteq> cset" using sheap1 by auto
+ have called: "\<And>C. coll_init_class P ?i1 = Some C
+ \<Longrightarrow> ics_of (hd (frames_of \<sigma>1)) = Called [] \<Longrightarrow> classes_above P C \<subseteq> cset"
+ proof -
+ fix C assume cic: "coll_init_class P ?i1 = Some C" and
+ ics: "ics_of (hd (frames_of \<sigma>1)) = Called []"
+ then obtain sobj where "sh1 C = Some sobj" using vics1 f1'
+ by(cases ?i1, auto simp: seeing_class_def split: if_split_asm)
+ then show "classes_above P C \<subseteq> cset" using sheap1 by(cases sobj, simp)
+ qed
+ have init_class_prom1: "\<And>C. ics1 = Called [] \<or> ics1 = No_ics
+ \<Longrightarrow> coll_init_class P ?i1 = Some C \<Longrightarrow> classes_above P C \<subseteq> cset"
+ proof -
+ fix C assume "ics1 = Called [] \<or> ics1 = No_ics" and cic: "coll_init_class P ?i1 = Some C"
+ then have ics: "?ics1 = Called [] \<or> ?ics1 = No_ics" using f1' by simp
+ then show "classes_above P C \<subseteq> cset" using cic
+ proof(cases "?ics1 = Called []")
+ case True then show ?thesis using cic called by simp
+ next
+ case False
+ then have ics': "?ics1 = No_ics" using ics by simp
+ then show ?thesis using cic
+ proof(cases ?i1)
+ case (New C1)
+ then have "is_class P C1" using \<Phi>_pc1 wt1 meth1 by auto
+ then have "P \<turnstile> C1 \<preceq>\<^sup>* Object" using wtp is_class_is_subcls
+ by(auto simp: wf_jvm_prog_phi_def)
+ then show ?thesis using New_collects[OF _ cbig1(2) nend1 _ ics' sheap2 sub'']
+ f1' ics cic New by auto
+ next
+ case (Getstatic C1 F1 D1)
+ then obtain t where mC1: "P \<turnstile> C1 has F1,Static:t in D1" and eq: "C = D1"
+ using cic by (metis coll_init_class.simps(2) option.inject option.simps(3))
+ then have "is_class P C" using has_field_is_class'[OF mC1] by simp
+ then have "P \<turnstile> C \<preceq>\<^sup>* Object" using wtp is_class_is_subcls
+ by(auto simp: wf_jvm_prog_phi_def)
+ then show ?thesis using Getstatic_collects[OF _ cbig1(2) nend1 _ ics' _ sheap2 sub'']
+ eq f1' Getstatic ics cic by fastforce
+ next
+ case (Putstatic C1 F1 D1)
+ then obtain t where mC1: "P \<turnstile> C1 has F1,Static:t in D1" and eq: "C = D1"
+ using cic by (metis coll_init_class.simps(3) option.inject option.simps(3))
+ then have "is_class P C" using has_field_is_class'[OF mC1] by simp
+ then have "P \<turnstile> C \<preceq>\<^sup>* Object" using wtp is_class_is_subcls
+ by(auto simp: wf_jvm_prog_phi_def)
+ then show ?thesis using Putstatic_collects[OF _ cbig1(2) nend1 _ ics' _ sheap2 sub'']
+ eq f1' Putstatic ics cic by fastforce
+ next
+ case (Invokestatic C1 M1 n')
+ then obtain Ts T m where mC: "P \<turnstile> C1 sees M1, Static : Ts\<rightarrow>T = m in C"
+ using cic by(fastforce simp: seeing_class_def split: if_split_asm)
+ then have "is_class P C" by(rule sees_method_is_class')
+ then have Obj: "P \<turnstile> C \<preceq>\<^sup>* Object" using wtp is_class_is_subcls
+ by(auto simp: wf_jvm_prog_phi_def)
+ show ?thesis using Invokestatic_collects[OF _ cbig1(2) sub'' nend1 _ ics' mC sheap2]
+ Obj mC f1' Invokestatic ics cic by auto
+ qed(simp+)
+ qed
+ qed
+\<comment> \<open> Calling promise \<close>
+ have Calling_prom1: "\<And>C' Cs'. ics1 = Calling C' Cs' \<Longrightarrow> classes_above P C' \<subseteq> cset"
+ proof -
+ fix C' Cs' assume ics: "ics1 = Calling C' Cs'"
+ then have "is_class P C'" using vics1 by simp
+ then have obj: "P \<turnstile> C' \<preceq>\<^sup>* Object" using wtp is_class_is_subcls
+ by(auto simp: wf_jvm_prog_phi_def)
+ have sheap3: "\<forall>C1. P \<turnstile> C' \<preceq>\<^sup>* C1 \<longrightarrow> (\<exists>sfs i. sheap \<sigma>1 C1 = \<lfloor>(sfs, i)\<rfloor>)
+ \<longrightarrow> classes_above P C1 \<subseteq> cset" using sheap1 by auto
+ show "classes_above P C' \<subseteq> cset"
+ using Calling_collects[OF obj cbig1(2) nend1 _ sheap3 sub''] ics f1' by simp
+ qed
+ have in_naive: "(\<sigma>', cset') \<in> JVMNaiveCollectionSemantics.csmall_nstep P (xp1, h1, frs1, sh1) n1"
+ and in_smart: "(\<sigma>', cset'') \<in> JVMSmartCollectionSemantics.csmall_nstep P (xp1, h1, frs1, sh1) n1"
+ using \<sigma>1(3) \<sigma>1'(3) \<sigma>_eq[THEN sym] by simp+
+ have sub2: "cset' \<subseteq> cset"
+ by(rule Suc.hyps[OF wtp correct1 f1' heap1 sheap1 xcpts1 frames1 init_class_prom1
+ Calling_prom1 in_naive in_smart sub'' Suc.prems(13)]) simp_all
+ then show ?thesis using \<sigma>1(2) \<sigma>1'(2) sub1 sub2 by fastforce
+ qed
+qed
+
+\<comment> \<open> ...which means over @{term cbig} \<close>
+lemma jvm_naive_to_smart_cbig:
+assumes
+\<comment> \<open> well-formedness \<close>
+ wtp: "wf_jvm_prog\<^bsub>\<Phi>\<^esub> P"
+ and correct: "P,\<Phi> \<turnstile> (xp,h,frs,sh)\<surd>"
+\<comment> \<open> defs \<close>
+ and f': "hd frs = (stk,loc,C',M',pc,ics)"
+\<comment> \<open> backward promises - will be collected/maintained prior \<close>
+ and heap: "\<And>C fs. \<exists>a. h a = Some(C,fs) \<Longrightarrow> classes_above P C \<subseteq> cset"
+ and sheap: "\<And>C sfs i. sh C = Some(sfs,i) \<Longrightarrow> classes_above P C \<subseteq> cset"
+ and xcpts: "classes_above_xcpts P \<subseteq> cset"
+ and frames: "classes_above_frames P frs \<subseteq> cset"
+\<comment> \<open> forward promises - will be collected after if not already \<close>
+ and init_class_prom: "\<And>C. ics = Called [] \<or> ics = No_ics
+ \<Longrightarrow> coll_init_class P (instrs_of P C' M' ! pc) = Some C \<Longrightarrow> classes_above P C \<subseteq> cset"
+ and Calling_prom: "\<And>C' Cs'. ics = Calling C' Cs' \<Longrightarrow> classes_above P C' \<subseteq> cset"
+\<comment> \<open> collections \<close>
+ and n: "(\<sigma>', cset\<^sub>n) \<in> JVMNaiveCollectionSemantics.cbig P (xp,h,frs,sh)"
+ and s: "(\<sigma>', cset\<^sub>s) \<in> JVMSmartCollectionSemantics.cbig P (xp,h,frs,sh)"
+ and smart: "cset\<^sub>s \<subseteq> cset"
+shows "cset\<^sub>n \<subseteq> cset"
+proof -
+ let ?\<sigma> = "(xp,h,frs,sh)"
+ have nend: "\<sigma>' \<in> JVMendset" using n by(simp add: JVMNaiveCollectionSemantics.cbig_def)
+ obtain n where n': "(\<sigma>', cset\<^sub>n) \<in> JVMNaiveCollectionSemantics.csmall_nstep P ?\<sigma> n" "\<sigma>' \<in> JVMendset"
+ using JVMNaiveCollectionSemantics.cbig_def2 n by auto
+ obtain s where s': "(\<sigma>', cset\<^sub>s) \<in> JVMSmartCollectionSemantics.csmall_nstep P ?\<sigma> s" "\<sigma>' \<in> JVMendset"
+ using JVMSmartCollectionSemantics.cbig_def2 s by auto
+ have "n=s" using jvm_naive_to_smart_csmall_nstep_last_eq[OF n n'(1) s'(1)] by simp
+ then have sn: "(\<sigma>', cset\<^sub>s) \<in> JVMSmartCollectionSemantics.csmall_nstep P ?\<sigma> n" using s'(1) by simp
+ then show ?thesis
+ using jvm_naive_to_smart_csmall_nstep[OF assms(1-9) n'(1) sn assms(12) nend] by fast
+qed
+
+\<comment> \<open>...thus naive @{text "\<subseteq>"} smart over the out function, since all conditions will be met - and promises
+ kept - by the defined starting point \<close>
+lemma jvm_naive_to_smart_collection:
+assumes naive: "(\<sigma>',cset\<^sub>n) \<in> jvm_naive_out P t" and smart: "(\<sigma>',cset\<^sub>s) \<in> jvm_smart_out P t"
+ and P: "P \<in> jvm_progs" and t: "t \<in> jvm_tests"
+shows "cset\<^sub>n \<subseteq> cset\<^sub>s"
+proof -
+ let ?P = "jvm_make_test_prog P t"
+ let ?\<sigma> = "start_state (t#P)"
+ let ?i = "instrs_of ?P Start start_m ! 0" and ?ics = No_ics
+ obtain xp h frs sh where
+ [simp]: "?\<sigma> = (xp,h,frs,sh)" and
+ [simp]: "h = start_heap (t#P)" and
+ [simp]: "frs = [([], [], Start, start_m, 0, No_ics)]" and
+ [simp]: "sh = start_sheap"
+ by(clarsimp simp: start_state_def)
+
+ from P t have nS: "\<not> is_class (t # P) Start"
+ by(simp add: is_class_def class_def Start_def Test_def)
+ from P have nT: "\<not> is_class P Test" by simp
+ from P t obtain m where tPm: "t # P \<turnstile> (fst t) sees main, Static : []\<rightarrow>Void = m in (fst t)"
+ by auto
+ have nclinit: "main \<noteq> clinit" by(simp add: main_def clinit_def)
+ have Objp: "\<And>b' Ts' T' m' D'.
+ t#P \<turnstile> Object sees start_m, b' : Ts'\<rightarrow>T' = m' in D' \<Longrightarrow> b' = Static \<and> Ts' = [] \<and> T' = Void"
+ proof -
+ fix b' Ts' T' m' D'
+ assume mObj: "t#P \<turnstile> Object sees start_m, b' : Ts'\<rightarrow>T' = m' in D'"
+ from P have ot_nsub: "\<not> P \<turnstile> Object \<preceq>\<^sup>* Test"
+ by(clarsimp simp: wf_jvm_prog_def wf_jvm_prog_phi_def)
+ from class_add_sees_method_rev[OF _ ot_nsub] mObj t
+ have "P \<turnstile> Object sees start_m, b' : Ts'\<rightarrow>T' = m' in D'" by(cases t, auto)
+ with P jvm_progs_def show "b' = Static \<and> Ts' = [] \<and> T' = Void" by blast
+ qed
+ from P t obtain \<Phi> where wtp0: "wf_jvm_prog\<^bsub>\<Phi>\<^esub> (t#P)" by(auto simp: wf_jvm_prog_def)
+ let ?\<Phi>' = "\<lambda>C f. if C = Start \<and> (f = start_m \<or> f = clinit) then start_\<phi>\<^sub>m else \<Phi> C f"
+ from wtp0 have wtp: "wf_jvm_prog\<^bsub>?\<Phi>'\<^esub> ?P"
+ proof -
+ note wtp0 nS tPm nclinit
+ moreover obtain "\<And>C. C \<noteq> Start \<Longrightarrow> ?\<Phi>' C = \<Phi> C" "?\<Phi>' Start start_m = start_\<phi>\<^sub>m"
+ "?\<Phi>' Start clinit = start_\<phi>\<^sub>m" by simp
+ moreover note Objp
+ ultimately show ?thesis by(rule start_prog_wf_jvm_prog_phi)
+ qed
+ have cic: "coll_init_class ?P ?i = Some Test"
+ proof -
+ from wtp0 obtain wf_md where wf: "wf_prog wf_md (t#P)"
+ by(clarsimp dest!: wt_jvm_progD)
+ with start_prog_start_m_instrs t have i: "?i = Invokestatic Test main 0" by simp
+ from jvm_make_test_prog_sees_Test_main[OF P t] obtain m where
+ "?P \<turnstile> Test sees main, Static : []\<rightarrow>Void = m in Test" by fast
+ with t have "seeing_class (jvm_make_test_prog P t) Test main = \<lfloor>Test\<rfloor>"
+ by(cases m, fastforce simp: seeing_class_def)
+ with i show ?thesis by simp
+ qed
+\<comment> \<open> well-formedness \<close>
+ note wtp
+ moreover have correct: "?P,?\<Phi>' \<turnstile> (xp,h,frs,sh)\<surd>"
+ proof -
+ note wtp0 nS tPm nclinit
+ moreover have "?\<Phi>' Start start_m = start_\<phi>\<^sub>m" by simp
+ ultimately have "?P,?\<Phi>' \<turnstile> ?\<sigma>\<surd>" by(rule BV_correct_initial)
+ then show ?thesis by simp
+ qed
+\<comment> \<open> defs \<close>
+ moreover have "hd frs = ([], [], Start, start_m, 0, No_ics)" by simp
+\<comment> \<open> backward promises \<close>
+ moreover from jvm_smart_out_classes_above_start_heap[OF smart _ P t]
+ have heap: "\<And>C fs. \<exists>a. h a = Some(C,fs) \<Longrightarrow> classes_above ?P C \<subseteq> cset\<^sub>s" by auto
+ moreover from jvm_smart_out_classes_above_start_sheap[OF smart]
+ have sheap: "\<And>C sfs i. sh C = Some(sfs,i) \<Longrightarrow> classes_above ?P C \<subseteq> cset\<^sub>s" by simp
+ moreover from jvm_smart_out_classes_above_xcpts[OF smart P t]
+ have xcpts: "classes_above_xcpts ?P \<subseteq> cset\<^sub>s" by simp
+ moreover from jvm_smart_out_classes_above_frames[OF smart]
+ have frames: "classes_above_frames ?P frs \<subseteq> cset\<^sub>s" by simp
+\<comment> \<open> forward promises - will be collected after if not already \<close>
+ moreover from jvm_smart_out_classes_above_Test[OF smart P t] cic
+ have init_class_prom: "\<And>C. ?ics = Called [] \<or> ?ics = No_ics
+ \<Longrightarrow> coll_init_class ?P ?i = Some C \<Longrightarrow> classes_above ?P C \<subseteq> cset\<^sub>s" by simp
+ moreover have "\<And>C' Cs'. ?ics = Calling C' Cs' \<Longrightarrow> classes_above ?P C' \<subseteq> cset\<^sub>s" by simp
+\<comment> \<open> collections \<close>
+ moreover from naive
+ have n: "(\<sigma>', cset\<^sub>n) \<in> JVMNaiveCollectionSemantics.cbig ?P (xp,h,frs,sh)" by simp
+ moreover from smart obtain cset\<^sub>s' where
+ s: "(\<sigma>', cset\<^sub>s') \<in> JVMSmartCollectionSemantics.cbig ?P (xp,h,frs,sh)" and
+ "cset\<^sub>s' \<subseteq> cset\<^sub>s"
+ by clarsimp
+ ultimately show "cset\<^sub>n \<subseteq> cset\<^sub>s" by(rule jvm_naive_to_smart_cbig; simp)
+qed
+
+
+(******************************************)
+subsubsection \<open> Proving smart @{text "\<subseteq>"} naive \<close>
+
+text "We prove that @{term jvm_naive} collects everything @{term jvm_smart}
+ does. Combined with the other direction, this shows that the naive and smart
+ algorithms collect the same set of classes."
+
+lemma jvm_smart_to_naive_exec_collect:
+ "JVMexec_scollect P \<sigma> \<subseteq> JVMexec_ncollect P \<sigma>"
+proof -
+ obtain xp h frs sh where \<sigma>: "\<sigma>=(xp,h,frs,sh)" by(cases \<sigma>)
+ then show ?thesis
+ proof(cases "\<exists>x. xp = Some x \<or> frs = []")
+ case False
+ then obtain stk loc C M pc ics frs'
+ where none: "xp = None" and frs: "frs=(stk,loc,C,M,pc,ics)#frs'"
+ by(cases xp, auto, cases frs, auto)
+ have instr_case: "ics = Called [] \<or> ics = No_ics \<Longrightarrow> ?thesis"
+ proof -
+ assume ics: "ics = Called [] \<or> ics = No_ics"
+ then show ?thesis using \<sigma> none frs
+ proof(cases "curr_instr P (stk,loc,C,M,pc,ics)") qed(auto split: if_split_asm)
+ qed
+ then show ?thesis using \<sigma> none frs
+ proof(cases ics)
+ case(Called Cs) then show ?thesis using instr_case \<sigma> none frs by(cases Cs, auto)
+ qed(auto)
+ qed(auto)
+qed
+
+lemma jvm_smart_to_naive_csmall:
+assumes "(\<sigma>', cset\<^sub>n) \<in> JVMNaiveCollectionSemantics.csmall P \<sigma>"
+ and "(\<sigma>', cset\<^sub>s) \<in> JVMSmartCollectionSemantics.csmall P \<sigma>"
+shows "cset\<^sub>s \<subseteq> cset\<^sub>n"
+using jvm_smart_to_naive_exec_collect assms
+ by(auto simp: JVMNaiveCollectionSemantics.csmall_def
+ JVMSmartCollectionSemantics.csmall_def)
+
+lemma jvm_smart_to_naive_csmall_nstep:
+"\<lbrakk> (\<sigma>', cset\<^sub>n) \<in> JVMNaiveCollectionSemantics.csmall_nstep P \<sigma> n;
+ (\<sigma>', cset\<^sub>s) \<in> JVMSmartCollectionSemantics.csmall_nstep P \<sigma> n \<rbrakk>
+ \<Longrightarrow> cset\<^sub>s \<subseteq> cset\<^sub>n"
+proof(induct n arbitrary: \<sigma> \<sigma>' cset\<^sub>n cset\<^sub>s)
+ case (Suc n')
+ obtain \<sigma>1 cset1 cset' where \<sigma>1: "(\<sigma>1, cset1) \<in> JVMNaiveCollectionSemantics.csmall P \<sigma>"
+ "cset\<^sub>n = cset1 \<union> cset'" "(\<sigma>', cset') \<in> JVMNaiveCollectionSemantics.csmall_nstep P \<sigma>1 n'"
+ using JVMNaiveCollectionSemantics.csmall_nstep_SucD [OF Suc.prems(1)] by clarsimp+
+ obtain \<sigma>1' cset1' cset'' where \<sigma>1': "(\<sigma>1', cset1') \<in> JVMSmartCollectionSemantics.csmall P \<sigma>"
+ "cset\<^sub>s = cset1' \<union> cset''" "(\<sigma>', cset'') \<in> JVMSmartCollectionSemantics.csmall_nstep P \<sigma>1' n'"
+ using JVMSmartCollectionSemantics.csmall_nstep_SucD [OF Suc.prems(2)] by clarsimp+
+ have \<sigma>_eq: "\<sigma>1 = \<sigma>1'" using \<sigma>1(1) \<sigma>1'(1) by(simp add: JVMNaiveCollectionSemantics.csmall_def
+ JVMSmartCollectionSemantics.csmall_def)
+ then have sub1: "cset1' \<subseteq> cset1" using \<sigma>1(1) \<sigma>1'(1) jvm_smart_to_naive_csmall by blast
+ have sub2: "cset'' \<subseteq> cset'" using \<sigma>1(3) \<sigma>1'(3) \<sigma>_eq Suc.hyps by blast
+ then show ?case using \<sigma>1(2) \<sigma>1'(2) sub1 sub2 by blast
+qed(simp)
+
+lemma jvm_smart_to_naive_cbig:
+assumes n: "(\<sigma>', cset\<^sub>n) \<in> JVMNaiveCollectionSemantics.cbig P \<sigma>"
+ and s: "(\<sigma>', cset\<^sub>s) \<in> JVMSmartCollectionSemantics.cbig P \<sigma>"
+shows "cset\<^sub>s \<subseteq> cset\<^sub>n"
+proof -
+ obtain n where n': "(\<sigma>', cset\<^sub>n) \<in> JVMNaiveCollectionSemantics.csmall_nstep P \<sigma> n" "\<sigma>' \<in> JVMendset"
+ using JVMNaiveCollectionSemantics.cbig_def2 n by auto
+ obtain s where s': "(\<sigma>', cset\<^sub>s) \<in> JVMSmartCollectionSemantics.csmall_nstep P \<sigma> s" "\<sigma>' \<in> JVMendset"
+ using JVMSmartCollectionSemantics.cbig_def2 s by auto
+ have "n=s" using jvm_naive_to_smart_csmall_nstep_last_eq[OF n n'(1) s'(1)] by simp
+ then show ?thesis using jvm_smart_to_naive_csmall_nstep n'(1) s'(1) by blast
+qed
+
+lemma jvm_smart_to_naive_collection:
+assumes naive: "(\<sigma>',cset\<^sub>n) \<in> jvm_naive_out P t" and smart: "(\<sigma>',cset\<^sub>s) \<in> jvm_smart_out P t"
+ and "P \<in> jvm_progs" and "t \<in> jvm_tests"
+shows "cset\<^sub>s \<subseteq> cset\<^sub>n"
+proof -
+ have nend: "start_state (t#P) \<notin> JVMendset" by(simp add: JVMendset_def start_state_def)
+ from naive obtain n where
+ nstep: "(\<sigma>', cset\<^sub>n) \<in> JVMNaiveCollectionSemantics.csmall_nstep
+ (jvm_make_test_prog P t) (start_state (t#P)) n"
+ by(auto dest!: JVMNaiveCollectionSemantics.cbigD)
+ with nend naive obtain n' where n': "n = Suc n'"
+ by(cases n; simp add: JVMNaiveCollectionSemantics.cbig_def)
+ from start_prog_classes_above_Start
+ have "classes_above_frames (jvm_make_test_prog P t) (frames_of (start_state (t#P))) = {Object,Start}"
+ by(simp add: start_state_def)
+ with nstep n'
+ have "jvm_smart_collect_start (jvm_make_test_prog P t) \<subseteq> cset\<^sub>n"
+ by(auto simp: start_state_def JVMNaiveCollectionSemantics.csmall_def
+ dest!: JVMNaiveCollectionSemantics.csmall_nstep_SucD
+ simp del: JVMNaiveCollectionSemantics.csmall_nstep_Rec)
+ with jvm_smart_to_naive_cbig[where P="jvm_make_test_prog P t" and \<sigma>="start_state (t#P)" and \<sigma>'=\<sigma>']
+ jvm_smart_collect_start_make_test_prog assms show ?thesis by auto
+qed
+
+(**************************************************)
+subsubsection "Safety of the smart algorithm"
+
+text "Having proved containment in both directions, we get naive = smart:"
+lemma jvm_naive_eq_smart_collection:
+assumes naive: "(\<sigma>',cset\<^sub>n) \<in> jvm_naive_out P t" and smart: "(\<sigma>',cset\<^sub>s) \<in> jvm_smart_out P t"
+ and "P \<in> jvm_progs" and "t \<in> jvm_tests"
+shows "cset\<^sub>n = cset\<^sub>s"
+using jvm_naive_to_smart_collection[OF assms] jvm_smart_to_naive_collection[OF assms] by simp
+
+text "Thus, since the RTS algorithm based on @{term ncollect} is existence safe,
+ the algorithm based on @{term scollect} is as well."
+theorem jvm_smart_existence_safe:
+assumes P: "P \<in> jvm_progs" and P': "P' \<in> jvm_progs" and t: "t \<in> jvm_tests"
+ and out: "o1 \<in> jvm_smart_out P t" and dss: "jvm_deselect P o1 P'"
+shows "\<exists>o2 \<in> jvm_smart_out P' t. o1 = o2"
+proof -
+ obtain \<sigma>' cset\<^sub>s where o1[simp]: "o1=(\<sigma>',cset\<^sub>s)" by(cases o1)
+ with jvm_naive_iff_smart out obtain cset\<^sub>n where n: "(\<sigma>',cset\<^sub>n) \<in> jvm_naive_out P t" by blast
+
+ from jvm_naive_eq_smart_collection[OF n _ P t] out have eq: "cset\<^sub>n = cset\<^sub>s" by simp
+ with jvm_naive_existence_safe[OF P P' t n] dss have n': "(\<sigma>',cset\<^sub>n) \<in> jvm_naive_out P' t" by simp
+ with jvm_naive_iff_smart obtain cset\<^sub>s' where s': "(\<sigma>', cset\<^sub>s') \<in> jvm_smart_out P' t" by blast
+
+ from jvm_naive_eq_smart_collection[OF n' s' P' t] eq have "cset\<^sub>s = cset\<^sub>s'" by simp
+ then show ?thesis using s' by simp
+qed
+
+text "...thus @{term JVMSmartCollection} is an instance of @{term CollectionBasedRTS}:"
+interpretation JVMSmartCollectionRTS :
+ CollectionBasedRTS "(=)" jvm_deselect jvm_progs jvm_tests
+ JVMendset JVMcombine JVMcollect_id JVMsmall JVMSmartCollect jvm_smart_out
+ jvm_make_test_prog jvm_smart_collect_start
+ by unfold_locales (rule jvm_smart_existence_safe, auto simp: start_state_def)
+
end
\ No newline at end of file
diff --git a/thys/Regression_Test_Selection/JVM_RTS/JVMCollectionSemantics.thy b/thys/Regression_Test_Selection/JVM_RTS/JVMCollectionSemantics.thy
--- a/thys/Regression_Test_Selection/JVM_RTS/JVMCollectionSemantics.thy
+++ b/thys/Regression_Test_Selection/JVM_RTS/JVMCollectionSemantics.thy
@@ -1,221 +1,221 @@
-(* File: RTS/JVM_RTS/JVMCollectionSemantics.thy *)
-(* Author: Susannah Mansky, UIUC 2020 *)
-
-section "Instantiating @{term CollectionSemantics} with Jinja JVM"
-
-theory JVMCollectionSemantics
-imports "../Common/CollectionSemantics" JVMSemantics "../JinjaSuppl/ClassesAbove"
-
-begin
-
-abbreviation JVMcombine :: "cname set \<Rightarrow> cname set \<Rightarrow> cname set" where
-"JVMcombine C C' \<equiv> C \<union> C'"
-
-abbreviation JVMcollect_id :: "cname set" where
-"JVMcollect_id \<equiv> {}"
-
-(*******************************************)
-subsection \<open> JVM-specific @{text "classes_above"} theory \<close>
-
-fun classes_above_frames :: "'m prog \<Rightarrow> frame list \<Rightarrow> cname set" where
-"classes_above_frames P ((stk,loc,C,M,pc,ics)#frs) = classes_above P C \<union> classes_above_frames P frs" |
-"classes_above_frames P [] = {}"
-
-lemma classes_above_start_state:
-assumes above_xcpts: "classes_above_xcpts P \<inter> classes_changed P P' = {}"
-shows "start_state P = start_state P'"
-using assms classes_above_start_heap by(simp add: start_state_def)
-
-lemma classes_above_matches_ex_entry:
- "classes_above P C \<inter> classes_changed P P' = {}
- \<Longrightarrow> matches_ex_entry P C pc xcp = matches_ex_entry P' C pc xcp"
-using classes_above_subcls classes_above_subcls2
- by(auto simp: matches_ex_entry_def)
-
-lemma classes_above_match_ex_table:
-assumes "classes_above P C \<inter> classes_changed P P' = {}"
-shows "match_ex_table P C pc es = match_ex_table P' C pc es"
-using classes_above_matches_ex_entry[OF assms] proof(induct es) qed(auto)
-
-lemma classes_above_find_handler:
-assumes "classes_above P (cname_of h a) \<inter> classes_changed P P' = {}"
-shows "classes_above_frames P frs \<inter> classes_changed P P' = {}
- \<Longrightarrow> find_handler P a h frs sh = find_handler P' a h frs sh"
-proof(induct frs)
- case (Cons fr' frs')
- obtain stk loc C M pc ics where fr': "fr' = (stk,loc,C,M,pc,ics)" by(cases fr')
- with Cons have
- intC: "classes_above P C \<inter> classes_changed P P' = {}"
- and int: "classes_above_frames P frs' \<inter> classes_changed P P' = {}" by auto
- show ?case using Cons fr' int classes_above_method[OF intC]
- classes_above_match_ex_table[OF assms(1)] by(auto split: bool.splits)
-qed(simp)
-
-lemma find_handler_classes_above_frames:
- "find_handler P a h frs sh = (xp',h',frs',sh')
- \<Longrightarrow> classes_above_frames P frs' \<subseteq> classes_above_frames P frs"
-proof(induct frs)
- case (Cons f1 frs1)
- then obtain stk loc C M pc ics where f1: "f1 = (stk,loc,C,M,pc,ics)" by(cases f1)
- show ?case
- proof(cases "match_ex_table P (cname_of h a) pc (ex_table_of P C M)")
- case None then show ?thesis using f1 None Cons
- by(auto split: bool.splits list.splits init_call_status.splits)
- next
- case (Some a) then show ?thesis using f1 Some Cons by auto
- qed
-qed(simp)
-
-lemma find_handler_pieces:
- "find_handler P a h frs sh = (xp',h',frs',sh')
- \<Longrightarrow> h = h' \<and> sh = sh' \<and> classes_above_frames P frs' \<subseteq> classes_above_frames P frs"
-using find_handler_classes_above_frames by(auto dest: find_handler_heap find_handler_sheap)
-
-(**************************************)
-
-subsection "Naive RTS algorithm"
-
-fun JVMinstr_ncollect ::
- "[jvm_prog, instr, heap, val list] \<Rightarrow> cname set" where
-"JVMinstr_ncollect P (New C) h stk = classes_above P C" |
-"JVMinstr_ncollect P (Getfield F C) h stk =
- (if (hd stk) = Null then {}
- else classes_above P (cname_of h (the_Addr (hd stk))))" |
-"JVMinstr_ncollect P (Getstatic C F D) h stk = classes_above P C" |
-"JVMinstr_ncollect P (Putfield F C) h stk =
- (if (hd (tl stk)) = Null then {}
- else classes_above P (cname_of h (the_Addr (hd (tl stk)))))" |
-"JVMinstr_ncollect P (Putstatic C F D) h stk = classes_above P C" |
-"JVMinstr_ncollect P (Checkcast C) h stk =
- (if (hd stk) = Null then {}
- else classes_above P (cname_of h (the_Addr (hd stk))))" |
-"JVMinstr_ncollect P (Invoke M n) h stk =
- (if (stk ! n) = Null then {}
- else classes_above P (cname_of h (the_Addr (stk ! n))))" |
-"JVMinstr_ncollect P (Invokestatic C M n) h stk = classes_above P C" |
-"JVMinstr_ncollect P Throw h stk =
- (if (hd stk) = Null then {}
- else classes_above P (cname_of h (the_Addr (hd stk))))" |
-"JVMinstr_ncollect P _ h stk = {}"
-
-fun JVMstep_ncollect ::
- "[jvm_prog, heap, val list, cname, mname, pc, init_call_status] \<Rightarrow> cname set" where
-"JVMstep_ncollect P h stk C M pc (Calling C' Cs) = classes_above P C'" |
-"JVMstep_ncollect P h stk C M pc (Called (C'#Cs))
- = classes_above P C' \<union> classes_above P (fst(method P C' clinit))" |
-"JVMstep_ncollect P h stk C M pc (Throwing Cs a) = classes_above P (cname_of h a)" |
-"JVMstep_ncollect P h stk C M pc ics = JVMinstr_ncollect P (instrs_of P C M ! pc) h stk"
-
-\<comment> \<open> naive collection function \<close>
-fun JVMexec_ncollect :: "jvm_prog \<Rightarrow> jvm_state \<Rightarrow> cname set" where
-"JVMexec_ncollect P (None, h, (stk,loc,C,M,pc,ics)#frs, sh) =
- (JVMstep_ncollect P h stk C M pc ics
- \<union> classes_above P C \<union> classes_above_frames P frs \<union> classes_above_xcpts P
- )"
-| "JVMexec_ncollect P _ = {}"
-
-(****)
-
-fun JVMNaiveCollect :: "jvm_prog \<Rightarrow> jvm_state \<Rightarrow> jvm_state \<Rightarrow> cname set" where
-"JVMNaiveCollect P \<sigma> \<sigma>' = JVMexec_ncollect P \<sigma>"
-
-interpretation JVMNaiveCollectionSemantics:
- CollectionSemantics JVMsmall JVMendset JVMNaiveCollect JVMcombine JVMcollect_id
- by unfold_locales auto
-
-(**************************************)
-
-subsection "Smarter RTS algorithm"
-
-fun JVMinstr_scollect ::
- "[jvm_prog, instr] \<Rightarrow> cname set" where
-"JVMinstr_scollect P (Getstatic C F D)
- = (if \<not>(\<exists>t. P \<turnstile> C has F,Static:t in D) then classes_above P C
- else classes_between P C D - {D})" |
-"JVMinstr_scollect P (Putstatic C F D)
- = (if \<not>(\<exists>t. P \<turnstile> C has F,Static:t in D) then classes_above P C
- else classes_between P C D - {D})" |
-"JVMinstr_scollect P (Invokestatic C M n)
- = (if \<not>(\<exists>Ts T m D. P \<turnstile> C sees M,Static:Ts \<rightarrow> T = m in D) then classes_above P C
- else classes_between P C (fst(method P C M)) - {fst(method P C M)})" |
-"JVMinstr_scollect P _ = {}"
-
-fun JVMstep_scollect ::
- "[jvm_prog, instr, init_call_status] \<Rightarrow> cname set" where
-"JVMstep_scollect P i (Calling C' Cs) = {C'}" |
-"JVMstep_scollect P i (Called (C'#Cs)) = {}" |
-"JVMstep_scollect P i (Throwing Cs a) = {}" |
-"JVMstep_scollect P i ics = JVMinstr_scollect P i"
-
-\<comment> \<open> smarter collection function \<close>
-fun JVMexec_scollect :: "jvm_prog \<Rightarrow> jvm_state \<Rightarrow> cname set" where
-"JVMexec_scollect P (None, h, (stk,loc,C,M,pc,ics)#frs, sh) =
- JVMstep_scollect P (instrs_of P C M ! pc) ics"
-| "JVMexec_scollect P _ = {}"
-
-(****)
-
-fun JVMSmartCollect :: "jvm_prog \<Rightarrow> jvm_state \<Rightarrow> jvm_state \<Rightarrow> cname set" where
-"JVMSmartCollect P \<sigma> \<sigma>' = JVMexec_scollect P \<sigma>"
-
-interpretation JVMSmartCollectionSemantics:
- CollectionSemantics JVMsmall JVMendset JVMSmartCollect JVMcombine JVMcollect_id
- by unfold_locales
-
-(***********************************************)
-
-subsection "A few lemmas using the instantiations"
-
-lemma JVMnaive_csmallD:
-"(\<sigma>', cset) \<in> JVMNaiveCollectionSemantics.csmall P \<sigma>
- \<Longrightarrow> JVMexec_ncollect P \<sigma> = cset \<and> \<sigma>' \<in> JVMsmall P \<sigma>"
- by(simp add: JVMNaiveCollectionSemantics.csmall_def)
-
-lemma JVMsmart_csmallD:
-"(\<sigma>', cset) \<in> JVMSmartCollectionSemantics.csmall P \<sigma>
- \<Longrightarrow> JVMexec_scollect P \<sigma> = cset \<and> \<sigma>' \<in> JVMsmall P \<sigma>"
- by(simp add: JVMSmartCollectionSemantics.csmall_def)
-
-
-lemma jvm_naive_to_smart_csmall_nstep_last_eq:
- "\<lbrakk> (\<sigma>',cset\<^sub>n) \<in> JVMNaiveCollectionSemantics.cbig P \<sigma>;
- (\<sigma>',cset\<^sub>n) \<in> JVMNaiveCollectionSemantics.csmall_nstep P \<sigma> n;
- (\<sigma>',cset\<^sub>s) \<in> JVMSmartCollectionSemantics.csmall_nstep P \<sigma> n' \<rbrakk>
- \<Longrightarrow> n = n'"
-proof(induct n arbitrary: n' \<sigma> \<sigma>' cset\<^sub>n cset\<^sub>s)
- case 0
- have "\<sigma>' = \<sigma>" using "0.prems"(2) JVMNaiveCollectionSemantics.csmall_nstep_base by blast
- then have endset: "\<sigma> \<in> JVMendset" using "0.prems"(1) JVMNaiveCollectionSemantics.cbigD by blast
- show ?case
- proof(cases n')
- case Suc then show ?thesis using "0.prems"(3) JVMSmartCollectionSemantics.csmall_nstep_Suc_nend
- endset by blast
- qed(simp)
-next
- case (Suc n1)
- then have endset: "\<sigma>' \<in> JVMendset" using Suc.prems(1) JVMNaiveCollectionSemantics.cbigD by blast
- have nend: "\<sigma> \<notin> JVMendset"
- using JVMNaiveCollectionSemantics.csmall_nstep_Suc_nend[OF Suc.prems(2)] by simp
- then have neq: "\<sigma>' \<noteq> \<sigma>" using endset by auto
- obtain \<sigma>1 cset cset1 where \<sigma>1: "(\<sigma>1,cset1) \<in> JVMNaiveCollectionSemantics.csmall P \<sigma>"
- "cset\<^sub>n = cset1 \<union> cset" "(\<sigma>',cset) \<in> JVMNaiveCollectionSemantics.csmall_nstep P \<sigma>1 n1"
- using JVMNaiveCollectionSemantics.csmall_nstep_SucD[OF Suc.prems(2)] by clarsimp
- then have cbig: "(\<sigma>',cset) \<in> JVMNaiveCollectionSemantics.cbig P \<sigma>1"
- using endset by(auto simp: JVMNaiveCollectionSemantics.cbig_def)
- show ?case
- proof(cases n')
- case 0 then show ?thesis
- using neq Suc.prems(3) JVMSmartCollectionSemantics.csmall_nstep_base by blast
- next
- case Suc': (Suc n1')
- then obtain \<sigma>1' cset2 cset1' where \<sigma>1': "(\<sigma>1',cset1') \<in> JVMSmartCollectionSemantics.csmall P \<sigma>"
- "cset\<^sub>s = cset1' \<union> cset2" "(\<sigma>',cset2) \<in> JVMSmartCollectionSemantics.csmall_nstep P \<sigma>1' n1'"
- using JVMSmartCollectionSemantics.csmall_nstep_SucD[where \<sigma>=\<sigma> and \<sigma>'=\<sigma>' and coll'=cset\<^sub>s
- and n=n1'] Suc.prems(3) by blast
- then have "\<sigma>1=\<sigma>1'" using \<sigma>1 JVMNaiveCollectionSemantics.csmall_def
- JVMSmartCollectionSemantics.csmall_def by auto
- then show ?thesis using Suc.hyps(1)[OF cbig \<sigma>1(3)] \<sigma>1'(3) Suc' by blast
- qed
-qed
-
+(* File: RTS/JVM_RTS/JVMCollectionSemantics.thy *)
+(* Author: Susannah Mansky, UIUC 2020 *)
+
+section "Instantiating @{term CollectionSemantics} with Jinja JVM"
+
+theory JVMCollectionSemantics
+imports "../Common/CollectionSemantics" JVMSemantics "../JinjaSuppl/ClassesAbove"
+
+begin
+
+abbreviation JVMcombine :: "cname set \<Rightarrow> cname set \<Rightarrow> cname set" where
+"JVMcombine C C' \<equiv> C \<union> C'"
+
+abbreviation JVMcollect_id :: "cname set" where
+"JVMcollect_id \<equiv> {}"
+
+(*******************************************)
+subsection \<open> JVM-specific @{text "classes_above"} theory \<close>
+
+fun classes_above_frames :: "'m prog \<Rightarrow> frame list \<Rightarrow> cname set" where
+"classes_above_frames P ((stk,loc,C,M,pc,ics)#frs) = classes_above P C \<union> classes_above_frames P frs" |
+"classes_above_frames P [] = {}"
+
+lemma classes_above_start_state:
+assumes above_xcpts: "classes_above_xcpts P \<inter> classes_changed P P' = {}"
+shows "start_state P = start_state P'"
+using assms classes_above_start_heap by(simp add: start_state_def)
+
+lemma classes_above_matches_ex_entry:
+ "classes_above P C \<inter> classes_changed P P' = {}
+ \<Longrightarrow> matches_ex_entry P C pc xcp = matches_ex_entry P' C pc xcp"
+using classes_above_subcls classes_above_subcls2
+ by(auto simp: matches_ex_entry_def)
+
+lemma classes_above_match_ex_table:
+assumes "classes_above P C \<inter> classes_changed P P' = {}"
+shows "match_ex_table P C pc es = match_ex_table P' C pc es"
+using classes_above_matches_ex_entry[OF assms] proof(induct es) qed(auto)
+
+lemma classes_above_find_handler:
+assumes "classes_above P (cname_of h a) \<inter> classes_changed P P' = {}"
+shows "classes_above_frames P frs \<inter> classes_changed P P' = {}
+ \<Longrightarrow> find_handler P a h frs sh = find_handler P' a h frs sh"
+proof(induct frs)
+ case (Cons fr' frs')
+ obtain stk loc C M pc ics where fr': "fr' = (stk,loc,C,M,pc,ics)" by(cases fr')
+ with Cons have
+ intC: "classes_above P C \<inter> classes_changed P P' = {}"
+ and int: "classes_above_frames P frs' \<inter> classes_changed P P' = {}" by auto
+ show ?case using Cons fr' int classes_above_method[OF intC]
+ classes_above_match_ex_table[OF assms(1)] by(auto split: bool.splits)
+qed(simp)
+
+lemma find_handler_classes_above_frames:
+ "find_handler P a h frs sh = (xp',h',frs',sh')
+ \<Longrightarrow> classes_above_frames P frs' \<subseteq> classes_above_frames P frs"
+proof(induct frs)
+ case (Cons f1 frs1)
+ then obtain stk loc C M pc ics where f1: "f1 = (stk,loc,C,M,pc,ics)" by(cases f1)
+ show ?case
+ proof(cases "match_ex_table P (cname_of h a) pc (ex_table_of P C M)")
+ case None then show ?thesis using f1 None Cons
+ by(auto split: bool.splits list.splits init_call_status.splits)
+ next
+ case (Some a) then show ?thesis using f1 Some Cons by auto
+ qed
+qed(simp)
+
+lemma find_handler_pieces:
+ "find_handler P a h frs sh = (xp',h',frs',sh')
+ \<Longrightarrow> h = h' \<and> sh = sh' \<and> classes_above_frames P frs' \<subseteq> classes_above_frames P frs"
+using find_handler_classes_above_frames by(auto dest: find_handler_heap find_handler_sheap)
+
+(**************************************)
+
+subsection "Naive RTS algorithm"
+
+fun JVMinstr_ncollect ::
+ "[jvm_prog, instr, heap, val list] \<Rightarrow> cname set" where
+"JVMinstr_ncollect P (New C) h stk = classes_above P C" |
+"JVMinstr_ncollect P (Getfield F C) h stk =
+ (if (hd stk) = Null then {}
+ else classes_above P (cname_of h (the_Addr (hd stk))))" |
+"JVMinstr_ncollect P (Getstatic C F D) h stk = classes_above P C" |
+"JVMinstr_ncollect P (Putfield F C) h stk =
+ (if (hd (tl stk)) = Null then {}
+ else classes_above P (cname_of h (the_Addr (hd (tl stk)))))" |
+"JVMinstr_ncollect P (Putstatic C F D) h stk = classes_above P C" |
+"JVMinstr_ncollect P (Checkcast C) h stk =
+ (if (hd stk) = Null then {}
+ else classes_above P (cname_of h (the_Addr (hd stk))))" |
+"JVMinstr_ncollect P (Invoke M n) h stk =
+ (if (stk ! n) = Null then {}
+ else classes_above P (cname_of h (the_Addr (stk ! n))))" |
+"JVMinstr_ncollect P (Invokestatic C M n) h stk = classes_above P C" |
+"JVMinstr_ncollect P Throw h stk =
+ (if (hd stk) = Null then {}
+ else classes_above P (cname_of h (the_Addr (hd stk))))" |
+"JVMinstr_ncollect P _ h stk = {}"
+
+fun JVMstep_ncollect ::
+ "[jvm_prog, heap, val list, cname, mname, pc, init_call_status] \<Rightarrow> cname set" where
+"JVMstep_ncollect P h stk C M pc (Calling C' Cs) = classes_above P C'" |
+"JVMstep_ncollect P h stk C M pc (Called (C'#Cs))
+ = classes_above P C' \<union> classes_above P (fst(method P C' clinit))" |
+"JVMstep_ncollect P h stk C M pc (Throwing Cs a) = classes_above P (cname_of h a)" |
+"JVMstep_ncollect P h stk C M pc ics = JVMinstr_ncollect P (instrs_of P C M ! pc) h stk"
+
+\<comment> \<open> naive collection function \<close>
+fun JVMexec_ncollect :: "jvm_prog \<Rightarrow> jvm_state \<Rightarrow> cname set" where
+"JVMexec_ncollect P (None, h, (stk,loc,C,M,pc,ics)#frs, sh) =
+ (JVMstep_ncollect P h stk C M pc ics
+ \<union> classes_above P C \<union> classes_above_frames P frs \<union> classes_above_xcpts P
+ )"
+| "JVMexec_ncollect P _ = {}"
+
+(****)
+
+fun JVMNaiveCollect :: "jvm_prog \<Rightarrow> jvm_state \<Rightarrow> jvm_state \<Rightarrow> cname set" where
+"JVMNaiveCollect P \<sigma> \<sigma>' = JVMexec_ncollect P \<sigma>"
+
+interpretation JVMNaiveCollectionSemantics:
+ CollectionSemantics JVMsmall JVMendset JVMNaiveCollect JVMcombine JVMcollect_id
+ by unfold_locales auto
+
+(**************************************)
+
+subsection "Smarter RTS algorithm"
+
+fun JVMinstr_scollect ::
+ "[jvm_prog, instr] \<Rightarrow> cname set" where
+"JVMinstr_scollect P (Getstatic C F D)
+ = (if \<not>(\<exists>t. P \<turnstile> C has F,Static:t in D) then classes_above P C
+ else classes_between P C D - {D})" |
+"JVMinstr_scollect P (Putstatic C F D)
+ = (if \<not>(\<exists>t. P \<turnstile> C has F,Static:t in D) then classes_above P C
+ else classes_between P C D - {D})" |
+"JVMinstr_scollect P (Invokestatic C M n)
+ = (if \<not>(\<exists>Ts T m D. P \<turnstile> C sees M,Static:Ts \<rightarrow> T = m in D) then classes_above P C
+ else classes_between P C (fst(method P C M)) - {fst(method P C M)})" |
+"JVMinstr_scollect P _ = {}"
+
+fun JVMstep_scollect ::
+ "[jvm_prog, instr, init_call_status] \<Rightarrow> cname set" where
+"JVMstep_scollect P i (Calling C' Cs) = {C'}" |
+"JVMstep_scollect P i (Called (C'#Cs)) = {}" |
+"JVMstep_scollect P i (Throwing Cs a) = {}" |
+"JVMstep_scollect P i ics = JVMinstr_scollect P i"
+
+\<comment> \<open> smarter collection function \<close>
+fun JVMexec_scollect :: "jvm_prog \<Rightarrow> jvm_state \<Rightarrow> cname set" where
+"JVMexec_scollect P (None, h, (stk,loc,C,M,pc,ics)#frs, sh) =
+ JVMstep_scollect P (instrs_of P C M ! pc) ics"
+| "JVMexec_scollect P _ = {}"
+
+(****)
+
+fun JVMSmartCollect :: "jvm_prog \<Rightarrow> jvm_state \<Rightarrow> jvm_state \<Rightarrow> cname set" where
+"JVMSmartCollect P \<sigma> \<sigma>' = JVMexec_scollect P \<sigma>"
+
+interpretation JVMSmartCollectionSemantics:
+ CollectionSemantics JVMsmall JVMendset JVMSmartCollect JVMcombine JVMcollect_id
+ by unfold_locales
+
+(***********************************************)
+
+subsection "A few lemmas using the instantiations"
+
+lemma JVMnaive_csmallD:
+"(\<sigma>', cset) \<in> JVMNaiveCollectionSemantics.csmall P \<sigma>
+ \<Longrightarrow> JVMexec_ncollect P \<sigma> = cset \<and> \<sigma>' \<in> JVMsmall P \<sigma>"
+ by(simp add: JVMNaiveCollectionSemantics.csmall_def)
+
+lemma JVMsmart_csmallD:
+"(\<sigma>', cset) \<in> JVMSmartCollectionSemantics.csmall P \<sigma>
+ \<Longrightarrow> JVMexec_scollect P \<sigma> = cset \<and> \<sigma>' \<in> JVMsmall P \<sigma>"
+ by(simp add: JVMSmartCollectionSemantics.csmall_def)
+
+
+lemma jvm_naive_to_smart_csmall_nstep_last_eq:
+ "\<lbrakk> (\<sigma>',cset\<^sub>n) \<in> JVMNaiveCollectionSemantics.cbig P \<sigma>;
+ (\<sigma>',cset\<^sub>n) \<in> JVMNaiveCollectionSemantics.csmall_nstep P \<sigma> n;
+ (\<sigma>',cset\<^sub>s) \<in> JVMSmartCollectionSemantics.csmall_nstep P \<sigma> n' \<rbrakk>
+ \<Longrightarrow> n = n'"
+proof(induct n arbitrary: n' \<sigma> \<sigma>' cset\<^sub>n cset\<^sub>s)
+ case 0
+ have "\<sigma>' = \<sigma>" using "0.prems"(2) JVMNaiveCollectionSemantics.csmall_nstep_base by blast
+ then have endset: "\<sigma> \<in> JVMendset" using "0.prems"(1) JVMNaiveCollectionSemantics.cbigD by blast
+ show ?case
+ proof(cases n')
+ case Suc then show ?thesis using "0.prems"(3) JVMSmartCollectionSemantics.csmall_nstep_Suc_nend
+ endset by blast
+ qed(simp)
+next
+ case (Suc n1)
+ then have endset: "\<sigma>' \<in> JVMendset" using Suc.prems(1) JVMNaiveCollectionSemantics.cbigD by blast
+ have nend: "\<sigma> \<notin> JVMendset"
+ using JVMNaiveCollectionSemantics.csmall_nstep_Suc_nend[OF Suc.prems(2)] by simp
+ then have neq: "\<sigma>' \<noteq> \<sigma>" using endset by auto
+ obtain \<sigma>1 cset cset1 where \<sigma>1: "(\<sigma>1,cset1) \<in> JVMNaiveCollectionSemantics.csmall P \<sigma>"
+ "cset\<^sub>n = cset1 \<union> cset" "(\<sigma>',cset) \<in> JVMNaiveCollectionSemantics.csmall_nstep P \<sigma>1 n1"
+ using JVMNaiveCollectionSemantics.csmall_nstep_SucD[OF Suc.prems(2)] by clarsimp
+ then have cbig: "(\<sigma>',cset) \<in> JVMNaiveCollectionSemantics.cbig P \<sigma>1"
+ using endset by(auto simp: JVMNaiveCollectionSemantics.cbig_def)
+ show ?case
+ proof(cases n')
+ case 0 then show ?thesis
+ using neq Suc.prems(3) JVMSmartCollectionSemantics.csmall_nstep_base by blast
+ next
+ case Suc': (Suc n1')
+ then obtain \<sigma>1' cset2 cset1' where \<sigma>1': "(\<sigma>1',cset1') \<in> JVMSmartCollectionSemantics.csmall P \<sigma>"
+ "cset\<^sub>s = cset1' \<union> cset2" "(\<sigma>',cset2) \<in> JVMSmartCollectionSemantics.csmall_nstep P \<sigma>1' n1'"
+ using JVMSmartCollectionSemantics.csmall_nstep_SucD[where \<sigma>=\<sigma> and \<sigma>'=\<sigma>' and coll'=cset\<^sub>s
+ and n=n1'] Suc.prems(3) by blast
+ then have "\<sigma>1=\<sigma>1'" using \<sigma>1 JVMNaiveCollectionSemantics.csmall_def
+ JVMSmartCollectionSemantics.csmall_def by auto
+ then show ?thesis using Suc.hyps(1)[OF cbig \<sigma>1(3)] \<sigma>1'(3) Suc' by blast
+ qed
+qed
+
end
\ No newline at end of file
diff --git a/thys/Regression_Test_Selection/JVM_RTS/JVMSemantics.thy b/thys/Regression_Test_Selection/JVM_RTS/JVMSemantics.thy
--- a/thys/Regression_Test_Selection/JVM_RTS/JVMSemantics.thy
+++ b/thys/Regression_Test_Selection/JVM_RTS/JVMSemantics.thy
@@ -1,35 +1,35 @@
-(* Title: RTS/JVM_RTS/JVMSemantics.thy *)
-(* Author: Susannah Mansky, UIUC 2020 *)
-
-section "Instantiating @{term Semantics} with Jinja JVM"
-
-theory JVMSemantics
-imports "../Common/Semantics" JinjaDCI.JVMExec
-begin
-
-fun JVMsmall :: "jvm_prog \<Rightarrow> jvm_state \<Rightarrow> jvm_state set" where
-"JVMsmall P \<sigma> = { \<sigma>'. exec (P, \<sigma>) = Some \<sigma>' }"
-
-lemma JVMsmall_prealloc_pres:
-assumes pre: "preallocated (fst(snd \<sigma>))"
- and "\<sigma>' \<in> JVMsmall P \<sigma>"
-shows "preallocated (fst(snd \<sigma>'))"
-using exec_prealloc_pres[OF pre] assms by(cases \<sigma>, cases \<sigma>', auto)
-
-lemma JVMsmall_det: "JVMsmall P \<sigma> = {} \<or> (\<exists>\<sigma>'. JVMsmall P \<sigma> = {\<sigma>'})"
-by auto
-
-definition JVMendset :: "jvm_state set" where
-"JVMendset \<equiv> { (xp,h,frs,sh). frs = [] \<or> (\<exists>x. xp = Some x) }"
-
-lemma JVMendset_final: "\<sigma> \<in> JVMendset \<Longrightarrow> \<forall>P. JVMsmall P \<sigma> = {}"
- by(auto simp: JVMendset_def)
-
-lemma start_state_nend:
- "start_state P \<notin> JVMendset"
- by(simp add: start_state_def JVMendset_def)
-
-interpretation JVMSemantics: Semantics JVMsmall JVMendset
- by unfold_locales (auto dest: JVMendset_final)
-
+(* Title: RTS/JVM_RTS/JVMSemantics.thy *)
+(* Author: Susannah Mansky, UIUC 2020 *)
+
+section "Instantiating @{term Semantics} with Jinja JVM"
+
+theory JVMSemantics
+imports "../Common/Semantics" JinjaDCI.JVMExec
+begin
+
+fun JVMsmall :: "jvm_prog \<Rightarrow> jvm_state \<Rightarrow> jvm_state set" where
+"JVMsmall P \<sigma> = { \<sigma>'. exec (P, \<sigma>) = Some \<sigma>' }"
+
+lemma JVMsmall_prealloc_pres:
+assumes pre: "preallocated (fst(snd \<sigma>))"
+ and "\<sigma>' \<in> JVMsmall P \<sigma>"
+shows "preallocated (fst(snd \<sigma>'))"
+using exec_prealloc_pres[OF pre] assms by(cases \<sigma>, cases \<sigma>', auto)
+
+lemma JVMsmall_det: "JVMsmall P \<sigma> = {} \<or> (\<exists>\<sigma>'. JVMsmall P \<sigma> = {\<sigma>'})"
+by auto
+
+definition JVMendset :: "jvm_state set" where
+"JVMendset \<equiv> { (xp,h,frs,sh). frs = [] \<or> (\<exists>x. xp = Some x) }"
+
+lemma JVMendset_final: "\<sigma> \<in> JVMendset \<Longrightarrow> \<forall>P. JVMsmall P \<sigma> = {}"
+ by(auto simp: JVMendset_def)
+
+lemma start_state_nend:
+ "start_state P \<notin> JVMendset"
+ by(simp add: start_state_def JVMendset_def)
+
+interpretation JVMSemantics: Semantics JVMsmall JVMendset
+ by unfold_locales (auto dest: JVMendset_final)
+
end
\ No newline at end of file
diff --git a/thys/Regression_Test_Selection/JinjaSuppl/ClassesAbove.thy b/thys/Regression_Test_Selection/JinjaSuppl/ClassesAbove.thy
--- a/thys/Regression_Test_Selection/JinjaSuppl/ClassesAbove.thy
+++ b/thys/Regression_Test_Selection/JinjaSuppl/ClassesAbove.thy
@@ -1,286 +1,286 @@
-(* Title: RTS/JinjaSuppl/ClassesAbove.thy
- Author: Susannah Mansky, UIUC 2020
-*)
-
-section "@{term classes_above} theory"
-
-text "This section contains theory around the classes above
- (superclasses of) a class in the class structure, in particular
- noting that if their contents have not changed, then much of what
- that class sees (methods, fields) stays the same."
-
-theory ClassesAbove
-imports ClassesChanged Subcls JinjaDCI.Exceptions
-begin
-
-abbreviation classes_above :: "'m prog \<Rightarrow> cname \<Rightarrow> cname set" where
-"classes_above P c \<equiv> { cn. P \<turnstile> c \<preceq>\<^sup>* cn }"
-
-abbreviation classes_between :: "'m prog \<Rightarrow> cname \<Rightarrow> cname \<Rightarrow> cname set" where
-"classes_between P c d \<equiv> { cn. (P \<turnstile> c \<preceq>\<^sup>* cn \<and> P \<turnstile> cn \<preceq>\<^sup>* d) }"
-
-abbreviation classes_above_xcpts :: "'m prog \<Rightarrow> cname set" where
-"classes_above_xcpts P \<equiv> \<Union>x\<in>sys_xcpts. classes_above P x"
-
-(******************************************************************************)
-
-lemma classes_above_def2:
- "P \<turnstile> C \<prec>\<^sup>1 D \<Longrightarrow> classes_above P C = {C} \<union> classes_above P D"
-using subcls1_confluent by auto
-
-lemma classes_above_class:
- "\<lbrakk> classes_above P C \<inter> classes_changed P P' = {}; P \<turnstile> C \<preceq>\<^sup>* C' \<rbrakk>
- \<Longrightarrow> class P C' = class P' C'"
- by (drule classes_changed_class_set, simp)
-
-lemma classes_above_subset:
-assumes "classes_above P C \<inter> classes_changed P P' = {}"
-shows "classes_above P C \<subseteq> classes_above P' C"
-proof -
- have ind: "\<And>x. P \<turnstile> C \<preceq>\<^sup>* x \<Longrightarrow> P' \<turnstile> C \<preceq>\<^sup>* x"
- proof -
- fix x assume sub: "P \<turnstile> C \<preceq>\<^sup>* x"
- then show "P' \<turnstile> C \<preceq>\<^sup>* x"
- proof(induct rule: rtrancl_induct)
- case base then show ?case by simp
- next
- case (step y z)
- have "P' \<turnstile> y \<prec>\<^sup>1 z" by(rule class_subcls1[OF classes_above_class[OF assms step(1)] step(2)])
- then show ?case using step(3) by simp
- qed
- qed
- with classes_changed_class_set[OF assms] show ?thesis by clarsimp
-qed
-
-lemma classes_above_subcls:
- "\<lbrakk> classes_above P C \<inter> classes_changed P P' = {}; P \<turnstile> C \<preceq>\<^sup>* C' \<rbrakk>
- \<Longrightarrow> P' \<turnstile> C \<preceq>\<^sup>* C'"
- by (fastforce dest: classes_above_subset)
-
-lemma classes_above_subset2:
-assumes "classes_above P C \<inter> classes_changed P P' = {}"
-shows "classes_above P' C \<subseteq> classes_above P C"
-proof -
- have ind: "\<And>x. P' \<turnstile> C \<preceq>\<^sup>* x \<Longrightarrow> P \<turnstile> C \<preceq>\<^sup>* x"
- proof -
- fix x assume sub: "P' \<turnstile> C \<preceq>\<^sup>* x"
- then show "P \<turnstile> C \<preceq>\<^sup>* x"
- proof(induct rule: rtrancl_induct)
- case base then show ?case by simp
- next
- case (step y z)
- with class_subcls1 classes_above_class[OF assms] rtrancl_into_rtrancl show ?case by metis
- qed
- qed
- with classes_changed_class_set[OF assms] show ?thesis by clarsimp
-qed
-
-lemma classes_above_subcls2:
- "\<lbrakk> classes_above P C \<inter> classes_changed P P' = {}; P' \<turnstile> C \<preceq>\<^sup>* C' \<rbrakk>
- \<Longrightarrow> P \<turnstile> C \<preceq>\<^sup>* C'"
- by (fastforce dest: classes_above_subset2)
-
-lemma classes_above_set:
- "\<lbrakk> classes_above P C \<inter> classes_changed P P' = {} \<rbrakk>
- \<Longrightarrow> classes_above P C = classes_above P' C"
- by(fastforce dest: classes_above_subset classes_above_subset2)
-
-lemma classes_above_classes_changed_sym:
-assumes "classes_above P C \<inter> classes_changed P P' = {}"
-shows "classes_above P' C \<inter> classes_changed P' P = {}"
-proof -
- have "classes_above P C = classes_above P' C" by(rule classes_above_set[OF assms])
- with classes_changed_sym[where P=P] assms show ?thesis by simp
-qed
-
-lemma classes_above_sub_classes_between_eq:
- "P \<turnstile> C \<preceq>\<^sup>* D \<Longrightarrow> classes_above P C = (classes_between P C D - {D}) \<union> classes_above P D"
-using subcls_confluent by auto
-
-lemma classes_above_subcls_subset:
- "\<lbrakk> P \<turnstile> C \<preceq>\<^sup>* C' \<rbrakk> \<Longrightarrow> classes_above P C' \<subseteq> classes_above P C"
- by auto
-
-(************************************************************)
-subsection "Methods"
-
-lemma classes_above_sees_methods:
-assumes int: "classes_above P C \<inter> classes_changed P P' = {}" and ms: "P \<turnstile> C sees_methods Mm"
-shows "P' \<turnstile> C sees_methods Mm"
-proof -
- have cls: "\<forall>C'\<in>classes_above P C. class P C' = class P' C'"
- by(rule classes_changed_class_set[OF int])
-
- have "\<And>C Mm. P \<turnstile> C sees_methods Mm \<Longrightarrow>
- \<forall>C'\<in>classes_above P C. class P C' = class P' C' \<Longrightarrow> P' \<turnstile> C sees_methods Mm"
- proof -
- fix C Mm assume "P \<turnstile> C sees_methods Mm" and "\<forall>C'\<in>classes_above P C. class P C' = class P' C'"
- then show "P' \<turnstile> C sees_methods Mm"
- proof(induct rule: Methods.induct)
- case Obj: (sees_methods_Object D fs ms Mm)
- with cls have "class P' Object = \<lfloor>(D, fs, ms)\<rfloor>" by simp
- with Obj show ?case by(auto intro!: sees_methods_Object)
- next
- case rec: (sees_methods_rec C D fs ms Mm Mm')
- then have "P \<turnstile> C \<preceq>\<^sup>* D" by (simp add: r_into_rtrancl[OF subcls1I])
- with converse_rtrancl_into_rtrancl have "\<And>x. P \<turnstile> D \<preceq>\<^sup>* x \<Longrightarrow> P \<turnstile> C \<preceq>\<^sup>* x" by simp
- with rec.prems(1) have "\<forall>C'\<in>classes_above P D. class P C' = class P' C'" by simp
- with rec show ?case by(auto intro: sees_methods_rec)
- qed
- qed
- with ms cls show ?thesis by simp
-qed
-
-lemma classes_above_sees_method:
- "\<lbrakk> classes_above P C \<inter> classes_changed P P' = {};
- P \<turnstile> C sees M,b: Ts\<rightarrow>T = m in C' \<rbrakk>
- \<Longrightarrow> P' \<turnstile> C sees M,b: Ts\<rightarrow>T = m in C'"
- by (auto dest: classes_above_sees_methods simp: Method_def)
-
-lemma classes_above_sees_method2:
- "\<lbrakk> classes_above P C \<inter> classes_changed P P' = {};
- P' \<turnstile> C sees M,b: Ts\<rightarrow>T = m in C' \<rbrakk>
- \<Longrightarrow> P \<turnstile> C sees M,b: Ts\<rightarrow>T = m in C'"
- by (auto dest: classes_above_classes_changed_sym intro: classes_above_sees_method)
-
-lemma classes_above_method:
-assumes "classes_above P C \<inter> classes_changed P P' = {}"
-shows "method P C M = method P' C M"
-proof(cases "\<exists>Ts T m D b. P \<turnstile> C sees M,b : Ts\<rightarrow>T = m in D")
- case True
- with assms show ?thesis by (auto dest: classes_above_sees_method)
-next
- case False
- with assms have "\<not>(\<exists>Ts T m D b. P' \<turnstile> C sees M,b : Ts\<rightarrow>T = m in D)"
- by (auto dest: classes_above_sees_method2)
- with False show ?thesis by(simp add: method_def)
-qed
-
-(*********************************************)
-subsection "Fields"
-
-lemma classes_above_has_fields:
-assumes int: "classes_above P C \<inter> classes_changed P P' = {}" and fs: "P \<turnstile> C has_fields FDTs"
-shows "P' \<turnstile> C has_fields FDTs"
-proof -
- have cls: "\<forall>C'\<in>classes_above P C. class P C' = class P' C'"
- by(rule classes_changed_class_set[OF int])
-
- have "\<And>C Mm. P \<turnstile> C has_fields FDTs \<Longrightarrow>
- \<forall>C'\<in>classes_above P C. class P C' = class P' C' \<Longrightarrow> P' \<turnstile> C has_fields FDTs"
- proof -
- fix C Mm assume "P \<turnstile> C has_fields FDTs" and "\<forall>C'\<in>classes_above P C. class P C' = class P' C'"
- then show "P' \<turnstile> C has_fields FDTs"
- proof(induct rule: Fields.induct)
- case Obj: (has_fields_Object D fs ms FDTs)
- with cls have "class P' Object = \<lfloor>(D, fs, ms)\<rfloor>" by simp
- with Obj show ?case by(auto intro!: has_fields_Object)
- next
- case rec: (has_fields_rec C D fs ms FDTs FDTs')
- then have "P \<turnstile> C \<preceq>\<^sup>* D" by (simp add: r_into_rtrancl[OF subcls1I])
- with converse_rtrancl_into_rtrancl have "\<And>x. P \<turnstile> D \<preceq>\<^sup>* x \<Longrightarrow> P \<turnstile> C \<preceq>\<^sup>* x" by simp
- with rec.prems(1) have "\<forall>x. P \<turnstile> D \<preceq>\<^sup>* x \<longrightarrow> class P x = class P' x" by simp
- with rec show ?case by(auto intro: has_fields_rec)
- qed
- qed
- with fs cls show ?thesis by simp
-qed
-
-lemma classes_above_has_fields_dne:
-assumes "classes_above P C \<inter> classes_changed P P' = {}"
-shows "(\<forall>FDTs. \<not> P \<turnstile> C has_fields FDTs) = (\<forall>FDTs. \<not> P' \<turnstile> C has_fields FDTs)"
-proof(rule iffI)
- assume asm: "\<forall>FDTs. \<not> P \<turnstile> C has_fields FDTs"
- from assms classes_changed_sym[where P=P] classes_above_set[OF assms]
- have int': "classes_above P' C \<inter> classes_changed P' P = {}" by simp
- from asm classes_above_has_fields[OF int'] show "\<forall>FDTs. \<not> P' \<turnstile> C has_fields FDTs" by auto
-next
- assume "\<forall>FDTs. \<not> P' \<turnstile> C has_fields FDTs"
- with assms show "\<forall>FDTs. \<not> P \<turnstile> C has_fields FDTs" by(auto dest: classes_above_has_fields)
-qed
-
-lemma classes_above_has_field:
- "\<lbrakk> classes_above P C \<inter> classes_changed P P' = {};
- P \<turnstile> C has F,b:t in C' \<rbrakk>
- \<Longrightarrow> P' \<turnstile> C has F,b:t in C'"
- by(auto dest: classes_above_has_fields simp: has_field_def)
-
-lemma classes_above_has_field2:
- "\<lbrakk> classes_above P C \<inter> classes_changed P P' = {};
- P' \<turnstile> C has F,b:t in C' \<rbrakk>
- \<Longrightarrow> P \<turnstile> C has F,b:t in C'"
- by(auto intro: classes_above_has_field dest: classes_above_classes_changed_sym)
-
-lemma classes_above_sees_field:
- "\<lbrakk> classes_above P C \<inter> classes_changed P P' = {};
- P \<turnstile> C sees F,b:t in C' \<rbrakk>
- \<Longrightarrow> P' \<turnstile> C sees F,b:t in C'"
- by(auto dest: classes_above_has_fields simp: sees_field_def)
-
-lemma classes_above_sees_field2:
- "\<lbrakk> classes_above P C \<inter> classes_changed P P' = {};
- P' \<turnstile> C sees F,b:t in C' \<rbrakk>
- \<Longrightarrow> P \<turnstile> C sees F,b:t in C'"
- by (auto intro: classes_above_sees_field dest: classes_above_classes_changed_sym)
-
-lemma classes_above_field:
-assumes "classes_above P C \<inter> classes_changed P P' = {}"
-shows "field P C F = field P' C F"
-proof(cases "\<exists>T D b. P \<turnstile> C sees F,b : T in D")
- case True
- with assms show ?thesis by (auto dest: classes_above_sees_field)
-next
- case False
- with assms have "\<not>(\<exists>T D b. P' \<turnstile> C sees F,b : T in D)"
- by (auto dest: classes_above_sees_field2)
- with False show ?thesis by(simp add: field_def)
-qed
-
-lemma classes_above_fields:
-assumes "classes_above P C \<inter> classes_changed P P' = {}"
-shows "fields P C = fields P' C"
-proof(cases "\<exists>FDTs. P \<turnstile> C has_fields FDTs")
- case True
- with assms show ?thesis by(auto dest: classes_above_has_fields)
-next
- case False
- with assms show ?thesis by (auto dest: classes_above_has_fields_dne simp: fields_def)
-qed
-
-lemma classes_above_ifields:
- "\<lbrakk> classes_above P C \<inter> classes_changed P P' = {} \<rbrakk>
- \<Longrightarrow>
- ifields P C = ifields P' C"
- by (simp add: ifields_def classes_above_fields)
-
-
-lemma classes_above_blank:
- "\<lbrakk> classes_above P C \<inter> classes_changed P P' = {} \<rbrakk>
- \<Longrightarrow>
- blank P C = blank P' C"
- by (simp add: blank_def classes_above_ifields)
-
-lemma classes_above_isfields:
- "\<lbrakk> classes_above P C \<inter> classes_changed P P' = {} \<rbrakk>
- \<Longrightarrow>
- isfields P C = isfields P' C"
- by (simp add: isfields_def classes_above_fields)
-
-lemma classes_above_sblank:
- "\<lbrakk> classes_above P C \<inter> classes_changed P P' = {} \<rbrakk>
- \<Longrightarrow>
- sblank P C = sblank P' C"
- by (simp add: sblank_def classes_above_isfields)
-
-(******************************************)
-subsection "Other"
-
-lemma classes_above_start_heap:
-assumes "classes_above_xcpts P \<inter> classes_changed P P' = {}"
-shows "start_heap P = start_heap P'"
-proof -
- from assms have "\<forall>C \<in> sys_xcpts. blank P C = blank P' C" by (auto intro: classes_above_blank)
- then show ?thesis by(simp add: start_heap_def)
-qed
-
+(* Title: RTS/JinjaSuppl/ClassesAbove.thy
+ Author: Susannah Mansky, UIUC 2020
+*)
+
+section "@{term classes_above} theory"
+
+text "This section contains theory around the classes above
+ (superclasses of) a class in the class structure, in particular
+ noting that if their contents have not changed, then much of what
+ that class sees (methods, fields) stays the same."
+
+theory ClassesAbove
+imports ClassesChanged Subcls JinjaDCI.Exceptions
+begin
+
+abbreviation classes_above :: "'m prog \<Rightarrow> cname \<Rightarrow> cname set" where
+"classes_above P c \<equiv> { cn. P \<turnstile> c \<preceq>\<^sup>* cn }"
+
+abbreviation classes_between :: "'m prog \<Rightarrow> cname \<Rightarrow> cname \<Rightarrow> cname set" where
+"classes_between P c d \<equiv> { cn. (P \<turnstile> c \<preceq>\<^sup>* cn \<and> P \<turnstile> cn \<preceq>\<^sup>* d) }"
+
+abbreviation classes_above_xcpts :: "'m prog \<Rightarrow> cname set" where
+"classes_above_xcpts P \<equiv> \<Union>x\<in>sys_xcpts. classes_above P x"
+
+(******************************************************************************)
+
+lemma classes_above_def2:
+ "P \<turnstile> C \<prec>\<^sup>1 D \<Longrightarrow> classes_above P C = {C} \<union> classes_above P D"
+using subcls1_confluent by auto
+
+lemma classes_above_class:
+ "\<lbrakk> classes_above P C \<inter> classes_changed P P' = {}; P \<turnstile> C \<preceq>\<^sup>* C' \<rbrakk>
+ \<Longrightarrow> class P C' = class P' C'"
+ by (drule classes_changed_class_set, simp)
+
+lemma classes_above_subset:
+assumes "classes_above P C \<inter> classes_changed P P' = {}"
+shows "classes_above P C \<subseteq> classes_above P' C"
+proof -
+ have ind: "\<And>x. P \<turnstile> C \<preceq>\<^sup>* x \<Longrightarrow> P' \<turnstile> C \<preceq>\<^sup>* x"
+ proof -
+ fix x assume sub: "P \<turnstile> C \<preceq>\<^sup>* x"
+ then show "P' \<turnstile> C \<preceq>\<^sup>* x"
+ proof(induct rule: rtrancl_induct)
+ case base then show ?case by simp
+ next
+ case (step y z)
+ have "P' \<turnstile> y \<prec>\<^sup>1 z" by(rule class_subcls1[OF classes_above_class[OF assms step(1)] step(2)])
+ then show ?case using step(3) by simp
+ qed
+ qed
+ with classes_changed_class_set[OF assms] show ?thesis by clarsimp
+qed
+
+lemma classes_above_subcls:
+ "\<lbrakk> classes_above P C \<inter> classes_changed P P' = {}; P \<turnstile> C \<preceq>\<^sup>* C' \<rbrakk>
+ \<Longrightarrow> P' \<turnstile> C \<preceq>\<^sup>* C'"
+ by (fastforce dest: classes_above_subset)
+
+lemma classes_above_subset2:
+assumes "classes_above P C \<inter> classes_changed P P' = {}"
+shows "classes_above P' C \<subseteq> classes_above P C"
+proof -
+ have ind: "\<And>x. P' \<turnstile> C \<preceq>\<^sup>* x \<Longrightarrow> P \<turnstile> C \<preceq>\<^sup>* x"
+ proof -
+ fix x assume sub: "P' \<turnstile> C \<preceq>\<^sup>* x"
+ then show "P \<turnstile> C \<preceq>\<^sup>* x"
+ proof(induct rule: rtrancl_induct)
+ case base then show ?case by simp
+ next
+ case (step y z)
+ with class_subcls1 classes_above_class[OF assms] rtrancl_into_rtrancl show ?case by metis
+ qed
+ qed
+ with classes_changed_class_set[OF assms] show ?thesis by clarsimp
+qed
+
+lemma classes_above_subcls2:
+ "\<lbrakk> classes_above P C \<inter> classes_changed P P' = {}; P' \<turnstile> C \<preceq>\<^sup>* C' \<rbrakk>
+ \<Longrightarrow> P \<turnstile> C \<preceq>\<^sup>* C'"
+ by (fastforce dest: classes_above_subset2)
+
+lemma classes_above_set:
+ "\<lbrakk> classes_above P C \<inter> classes_changed P P' = {} \<rbrakk>
+ \<Longrightarrow> classes_above P C = classes_above P' C"
+ by(fastforce dest: classes_above_subset classes_above_subset2)
+
+lemma classes_above_classes_changed_sym:
+assumes "classes_above P C \<inter> classes_changed P P' = {}"
+shows "classes_above P' C \<inter> classes_changed P' P = {}"
+proof -
+ have "classes_above P C = classes_above P' C" by(rule classes_above_set[OF assms])
+ with classes_changed_sym[where P=P] assms show ?thesis by simp
+qed
+
+lemma classes_above_sub_classes_between_eq:
+ "P \<turnstile> C \<preceq>\<^sup>* D \<Longrightarrow> classes_above P C = (classes_between P C D - {D}) \<union> classes_above P D"
+using subcls_confluent by auto
+
+lemma classes_above_subcls_subset:
+ "\<lbrakk> P \<turnstile> C \<preceq>\<^sup>* C' \<rbrakk> \<Longrightarrow> classes_above P C' \<subseteq> classes_above P C"
+ by auto
+
+(************************************************************)
+subsection "Methods"
+
+lemma classes_above_sees_methods:
+assumes int: "classes_above P C \<inter> classes_changed P P' = {}" and ms: "P \<turnstile> C sees_methods Mm"
+shows "P' \<turnstile> C sees_methods Mm"
+proof -
+ have cls: "\<forall>C'\<in>classes_above P C. class P C' = class P' C'"
+ by(rule classes_changed_class_set[OF int])
+
+ have "\<And>C Mm. P \<turnstile> C sees_methods Mm \<Longrightarrow>
+ \<forall>C'\<in>classes_above P C. class P C' = class P' C' \<Longrightarrow> P' \<turnstile> C sees_methods Mm"
+ proof -
+ fix C Mm assume "P \<turnstile> C sees_methods Mm" and "\<forall>C'\<in>classes_above P C. class P C' = class P' C'"
+ then show "P' \<turnstile> C sees_methods Mm"
+ proof(induct rule: Methods.induct)
+ case Obj: (sees_methods_Object D fs ms Mm)
+ with cls have "class P' Object = \<lfloor>(D, fs, ms)\<rfloor>" by simp
+ with Obj show ?case by(auto intro!: sees_methods_Object)
+ next
+ case rec: (sees_methods_rec C D fs ms Mm Mm')
+ then have "P \<turnstile> C \<preceq>\<^sup>* D" by (simp add: r_into_rtrancl[OF subcls1I])
+ with converse_rtrancl_into_rtrancl have "\<And>x. P \<turnstile> D \<preceq>\<^sup>* x \<Longrightarrow> P \<turnstile> C \<preceq>\<^sup>* x" by simp
+ with rec.prems(1) have "\<forall>C'\<in>classes_above P D. class P C' = class P' C'" by simp
+ with rec show ?case by(auto intro: sees_methods_rec)
+ qed
+ qed
+ with ms cls show ?thesis by simp
+qed
+
+lemma classes_above_sees_method:
+ "\<lbrakk> classes_above P C \<inter> classes_changed P P' = {};
+ P \<turnstile> C sees M,b: Ts\<rightarrow>T = m in C' \<rbrakk>
+ \<Longrightarrow> P' \<turnstile> C sees M,b: Ts\<rightarrow>T = m in C'"
+ by (auto dest: classes_above_sees_methods simp: Method_def)
+
+lemma classes_above_sees_method2:
+ "\<lbrakk> classes_above P C \<inter> classes_changed P P' = {};
+ P' \<turnstile> C sees M,b: Ts\<rightarrow>T = m in C' \<rbrakk>
+ \<Longrightarrow> P \<turnstile> C sees M,b: Ts\<rightarrow>T = m in C'"
+ by (auto dest: classes_above_classes_changed_sym intro: classes_above_sees_method)
+
+lemma classes_above_method:
+assumes "classes_above P C \<inter> classes_changed P P' = {}"
+shows "method P C M = method P' C M"
+proof(cases "\<exists>Ts T m D b. P \<turnstile> C sees M,b : Ts\<rightarrow>T = m in D")
+ case True
+ with assms show ?thesis by (auto dest: classes_above_sees_method)
+next
+ case False
+ with assms have "\<not>(\<exists>Ts T m D b. P' \<turnstile> C sees M,b : Ts\<rightarrow>T = m in D)"
+ by (auto dest: classes_above_sees_method2)
+ with False show ?thesis by(simp add: method_def)
+qed
+
+(*********************************************)
+subsection "Fields"
+
+lemma classes_above_has_fields:
+assumes int: "classes_above P C \<inter> classes_changed P P' = {}" and fs: "P \<turnstile> C has_fields FDTs"
+shows "P' \<turnstile> C has_fields FDTs"
+proof -
+ have cls: "\<forall>C'\<in>classes_above P C. class P C' = class P' C'"
+ by(rule classes_changed_class_set[OF int])
+
+ have "\<And>C Mm. P \<turnstile> C has_fields FDTs \<Longrightarrow>
+ \<forall>C'\<in>classes_above P C. class P C' = class P' C' \<Longrightarrow> P' \<turnstile> C has_fields FDTs"
+ proof -
+ fix C Mm assume "P \<turnstile> C has_fields FDTs" and "\<forall>C'\<in>classes_above P C. class P C' = class P' C'"
+ then show "P' \<turnstile> C has_fields FDTs"
+ proof(induct rule: Fields.induct)
+ case Obj: (has_fields_Object D fs ms FDTs)
+ with cls have "class P' Object = \<lfloor>(D, fs, ms)\<rfloor>" by simp
+ with Obj show ?case by(auto intro!: has_fields_Object)
+ next
+ case rec: (has_fields_rec C D fs ms FDTs FDTs')
+ then have "P \<turnstile> C \<preceq>\<^sup>* D" by (simp add: r_into_rtrancl[OF subcls1I])
+ with converse_rtrancl_into_rtrancl have "\<And>x. P \<turnstile> D \<preceq>\<^sup>* x \<Longrightarrow> P \<turnstile> C \<preceq>\<^sup>* x" by simp
+ with rec.prems(1) have "\<forall>x. P \<turnstile> D \<preceq>\<^sup>* x \<longrightarrow> class P x = class P' x" by simp
+ with rec show ?case by(auto intro: has_fields_rec)
+ qed
+ qed
+ with fs cls show ?thesis by simp
+qed
+
+lemma classes_above_has_fields_dne:
+assumes "classes_above P C \<inter> classes_changed P P' = {}"
+shows "(\<forall>FDTs. \<not> P \<turnstile> C has_fields FDTs) = (\<forall>FDTs. \<not> P' \<turnstile> C has_fields FDTs)"
+proof(rule iffI)
+ assume asm: "\<forall>FDTs. \<not> P \<turnstile> C has_fields FDTs"
+ from assms classes_changed_sym[where P=P] classes_above_set[OF assms]
+ have int': "classes_above P' C \<inter> classes_changed P' P = {}" by simp
+ from asm classes_above_has_fields[OF int'] show "\<forall>FDTs. \<not> P' \<turnstile> C has_fields FDTs" by auto
+next
+ assume "\<forall>FDTs. \<not> P' \<turnstile> C has_fields FDTs"
+ with assms show "\<forall>FDTs. \<not> P \<turnstile> C has_fields FDTs" by(auto dest: classes_above_has_fields)
+qed
+
+lemma classes_above_has_field:
+ "\<lbrakk> classes_above P C \<inter> classes_changed P P' = {};
+ P \<turnstile> C has F,b:t in C' \<rbrakk>
+ \<Longrightarrow> P' \<turnstile> C has F,b:t in C'"
+ by(auto dest: classes_above_has_fields simp: has_field_def)
+
+lemma classes_above_has_field2:
+ "\<lbrakk> classes_above P C \<inter> classes_changed P P' = {};
+ P' \<turnstile> C has F,b:t in C' \<rbrakk>
+ \<Longrightarrow> P \<turnstile> C has F,b:t in C'"
+ by(auto intro: classes_above_has_field dest: classes_above_classes_changed_sym)
+
+lemma classes_above_sees_field:
+ "\<lbrakk> classes_above P C \<inter> classes_changed P P' = {};
+ P \<turnstile> C sees F,b:t in C' \<rbrakk>
+ \<Longrightarrow> P' \<turnstile> C sees F,b:t in C'"
+ by(auto dest: classes_above_has_fields simp: sees_field_def)
+
+lemma classes_above_sees_field2:
+ "\<lbrakk> classes_above P C \<inter> classes_changed P P' = {};
+ P' \<turnstile> C sees F,b:t in C' \<rbrakk>
+ \<Longrightarrow> P \<turnstile> C sees F,b:t in C'"
+ by (auto intro: classes_above_sees_field dest: classes_above_classes_changed_sym)
+
+lemma classes_above_field:
+assumes "classes_above P C \<inter> classes_changed P P' = {}"
+shows "field P C F = field P' C F"
+proof(cases "\<exists>T D b. P \<turnstile> C sees F,b : T in D")
+ case True
+ with assms show ?thesis by (auto dest: classes_above_sees_field)
+next
+ case False
+ with assms have "\<not>(\<exists>T D b. P' \<turnstile> C sees F,b : T in D)"
+ by (auto dest: classes_above_sees_field2)
+ with False show ?thesis by(simp add: field_def)
+qed
+
+lemma classes_above_fields:
+assumes "classes_above P C \<inter> classes_changed P P' = {}"
+shows "fields P C = fields P' C"
+proof(cases "\<exists>FDTs. P \<turnstile> C has_fields FDTs")
+ case True
+ with assms show ?thesis by(auto dest: classes_above_has_fields)
+next
+ case False
+ with assms show ?thesis by (auto dest: classes_above_has_fields_dne simp: fields_def)
+qed
+
+lemma classes_above_ifields:
+ "\<lbrakk> classes_above P C \<inter> classes_changed P P' = {} \<rbrakk>
+ \<Longrightarrow>
+ ifields P C = ifields P' C"
+ by (simp add: ifields_def classes_above_fields)
+
+
+lemma classes_above_blank:
+ "\<lbrakk> classes_above P C \<inter> classes_changed P P' = {} \<rbrakk>
+ \<Longrightarrow>
+ blank P C = blank P' C"
+ by (simp add: blank_def classes_above_ifields)
+
+lemma classes_above_isfields:
+ "\<lbrakk> classes_above P C \<inter> classes_changed P P' = {} \<rbrakk>
+ \<Longrightarrow>
+ isfields P C = isfields P' C"
+ by (simp add: isfields_def classes_above_fields)
+
+lemma classes_above_sblank:
+ "\<lbrakk> classes_above P C \<inter> classes_changed P P' = {} \<rbrakk>
+ \<Longrightarrow>
+ sblank P C = sblank P' C"
+ by (simp add: sblank_def classes_above_isfields)
+
+(******************************************)
+subsection "Other"
+
+lemma classes_above_start_heap:
+assumes "classes_above_xcpts P \<inter> classes_changed P P' = {}"
+shows "start_heap P = start_heap P'"
+proof -
+ from assms have "\<forall>C \<in> sys_xcpts. blank P C = blank P' C" by (auto intro: classes_above_blank)
+ then show ?thesis by(simp add: start_heap_def)
+qed
+
end
\ No newline at end of file
diff --git a/thys/Regression_Test_Selection/JinjaSuppl/ClassesChanged.thy b/thys/Regression_Test_Selection/JinjaSuppl/ClassesChanged.thy
--- a/thys/Regression_Test_Selection/JinjaSuppl/ClassesChanged.thy
+++ b/thys/Regression_Test_Selection/JinjaSuppl/ClassesChanged.thy
@@ -1,75 +1,75 @@
-(* Title: RTS/JinjaSuppl/ClassesChanged.thy
- Author: Susannah Mansky, UIUC 2020
- Description: Theory around the classes changed from one program to another
-*)
-
-section "@{term classes_changed} theory"
-
-theory ClassesChanged
-imports JinjaDCI.Decl
-begin
-
-text "A class is considered changed if it exists only in one program or the other,
- or exists in both but is different."
-definition classes_changed :: "'m prog \<Rightarrow> 'm prog \<Rightarrow> cname set" where
-"classes_changed P1 P2 = {cn. class P1 cn \<noteq> class P2 cn}"
-
-definition class_changed :: "'m prog \<Rightarrow> 'm prog \<Rightarrow> cname \<Rightarrow> bool" where
-"class_changed P1 P2 cn = (class P1 cn \<noteq> class P2 cn)"
-
-lemma classes_changed_class_changed[simp]: "cn \<in> classes_changed P1 P2 = class_changed P1 P2 cn"
- by (simp add: classes_changed_def class_changed_def)
-
-lemma classes_changed_self[simp]: "classes_changed P P = {}"
- by (auto simp: class_changed_def)
-
-lemma classes_changed_sym: "classes_changed P P' = classes_changed P' P"
- by (auto simp: class_changed_def)
-
-lemma classes_changed_class: "\<lbrakk> cn \<notin> classes_changed P P'\<rbrakk> \<Longrightarrow> class P cn = class P' cn"
- by (clarsimp simp: class_changed_def)
-
-lemma classes_changed_class_set: "\<lbrakk> S \<inter> classes_changed P P' = {} \<rbrakk>
- \<Longrightarrow> \<forall>C \<in> S. class P C = class P' C"
- by (fastforce simp: disjoint_iff_not_equal dest: classes_changed_class)
-
-text "We now relate @{term classes_changed} over two programs to those
- over programs with an added class (such as a test class)."
-
-lemma classes_changed_cons_eq:
- "classes_changed (t # P) P' = (classes_changed P P' - {fst t})
- \<union> (if class_changed [t] P' (fst t) then {fst t} else {})"
- by (auto simp: classes_changed_def class_changed_def class_def)
-
-lemma class_changed_cons:
- "fst t \<notin> classes_changed (t#P) (t#P')"
- by (simp add: class_changed_def class_def)
-
-lemma classes_changed_cons:
- "classes_changed (t # P) (t # P') = classes_changed P P' - {fst t}"
-proof(cases "fst t \<in> classes_changed P P'")
- case True
- then show ?thesis using class_changed_cons[where t=t and P=P and P'=P']
- classes_changed_cons_eq[where t=t] by (auto simp: class_changed_def class_cons)
-next
- case False
- then show ?thesis using class_changed_cons[where t=t and P=P and P'=P']
- by (auto simp: class_changed_def) (metis (no_types, lifting) class_cons)+
-qed
-
-lemma classes_changed_int_Cons:
-assumes "coll \<inter> classes_changed P P' = {}"
-shows "coll \<inter> classes_changed (t # P) (t # P') = {}"
-proof(cases "fst t \<in> classes_changed P P'")
- case True
- then have "classes_changed P P' = classes_changed (t # P) (t # P') \<union> {fst t}"
- using classes_changed_cons[where t=t and P=P and P'=P'] by fastforce
- then show ?thesis using assms by simp
-next
- case False
- then have "classes_changed P P' = classes_changed (t # P) (t # P')"
- using classes_changed_cons[where t=t and P=P and P'=P'] by fastforce
- then show ?thesis using assms by simp
-qed
-
+(* Title: RTS/JinjaSuppl/ClassesChanged.thy
+ Author: Susannah Mansky, UIUC 2020
+ Description: Theory around the classes changed from one program to another
+*)
+
+section "@{term classes_changed} theory"
+
+theory ClassesChanged
+imports JinjaDCI.Decl
+begin
+
+text "A class is considered changed if it exists only in one program or the other,
+ or exists in both but is different."
+definition classes_changed :: "'m prog \<Rightarrow> 'm prog \<Rightarrow> cname set" where
+"classes_changed P1 P2 = {cn. class P1 cn \<noteq> class P2 cn}"
+
+definition class_changed :: "'m prog \<Rightarrow> 'm prog \<Rightarrow> cname \<Rightarrow> bool" where
+"class_changed P1 P2 cn = (class P1 cn \<noteq> class P2 cn)"
+
+lemma classes_changed_class_changed[simp]: "cn \<in> classes_changed P1 P2 = class_changed P1 P2 cn"
+ by (simp add: classes_changed_def class_changed_def)
+
+lemma classes_changed_self[simp]: "classes_changed P P = {}"
+ by (auto simp: class_changed_def)
+
+lemma classes_changed_sym: "classes_changed P P' = classes_changed P' P"
+ by (auto simp: class_changed_def)
+
+lemma classes_changed_class: "\<lbrakk> cn \<notin> classes_changed P P'\<rbrakk> \<Longrightarrow> class P cn = class P' cn"
+ by (clarsimp simp: class_changed_def)
+
+lemma classes_changed_class_set: "\<lbrakk> S \<inter> classes_changed P P' = {} \<rbrakk>
+ \<Longrightarrow> \<forall>C \<in> S. class P C = class P' C"
+ by (fastforce simp: disjoint_iff_not_equal dest: classes_changed_class)
+
+text "We now relate @{term classes_changed} over two programs to those
+ over programs with an added class (such as a test class)."
+
+lemma classes_changed_cons_eq:
+ "classes_changed (t # P) P' = (classes_changed P P' - {fst t})
+ \<union> (if class_changed [t] P' (fst t) then {fst t} else {})"
+ by (auto simp: classes_changed_def class_changed_def class_def)
+
+lemma class_changed_cons:
+ "fst t \<notin> classes_changed (t#P) (t#P')"
+ by (simp add: class_changed_def class_def)
+
+lemma classes_changed_cons:
+ "classes_changed (t # P) (t # P') = classes_changed P P' - {fst t}"
+proof(cases "fst t \<in> classes_changed P P'")
+ case True
+ then show ?thesis using class_changed_cons[where t=t and P=P and P'=P']
+ classes_changed_cons_eq[where t=t] by (auto simp: class_changed_def class_cons)
+next
+ case False
+ then show ?thesis using class_changed_cons[where t=t and P=P and P'=P']
+ by (auto simp: class_changed_def) (metis (no_types, lifting) class_cons)+
+qed
+
+lemma classes_changed_int_Cons:
+assumes "coll \<inter> classes_changed P P' = {}"
+shows "coll \<inter> classes_changed (t # P) (t # P') = {}"
+proof(cases "fst t \<in> classes_changed P P'")
+ case True
+ then have "classes_changed P P' = classes_changed (t # P) (t # P') \<union> {fst t}"
+ using classes_changed_cons[where t=t and P=P and P'=P'] by fastforce
+ then show ?thesis using assms by simp
+next
+ case False
+ then have "classes_changed P P' = classes_changed (t # P) (t # P')"
+ using classes_changed_cons[where t=t and P=P and P'=P'] by fastforce
+ then show ?thesis using assms by simp
+qed
+
end
\ No newline at end of file
diff --git a/thys/Regression_Test_Selection/JinjaSuppl/JVMExecStepInductive.thy b/thys/Regression_Test_Selection/JinjaSuppl/JVMExecStepInductive.thy
--- a/thys/Regression_Test_Selection/JinjaSuppl/JVMExecStepInductive.thy
+++ b/thys/Regression_Test_Selection/JinjaSuppl/JVMExecStepInductive.thy
@@ -1,919 +1,919 @@
-(* Title: RTS/JinjaSuppl/JVMExecStepInductive.thy
- Author: Susannah Mansky
- 2020, UIUC
-
- Program Execution in the JVM as an inductive
-*)
-
-section "Inductive JVM execution"
-
-theory JVMExecStepInductive
-imports JinjaDCI.JVMExec
-begin
-
-datatype step_input = StepI instr |
- StepC cname "cname list" | StepC2 cname "cname list" |
- StepT "cname list" addr
-
-
-inductive exec_step_ind :: "[step_input, jvm_prog, heap, val list, val list,
- cname, mname, pc, init_call_status, frame list, sheap,jvm_state] \<Rightarrow> bool"
-where
- exec_step_ind_Load:
-"exec_step_ind (StepI (Load n)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
- (None, h, ((loc ! n) # stk, loc, C\<^sub>0, M\<^sub>0, Suc pc, ics)#frs, sh)"
-
-| exec_step_ind_Store:
-"exec_step_ind (StepI (Store n)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
- (None, h, (tl stk, loc[n:=hd stk], C\<^sub>0, M\<^sub>0, Suc pc, ics)#frs, sh)"
-
-| exec_step_ind_Push:
-"exec_step_ind (StepI (Push v)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
- (None, h, (v # stk, loc, C\<^sub>0, M\<^sub>0, Suc pc, ics)#frs, sh)"
-
-| exec_step_ind_NewOOM_Called:
-"new_Addr h = None
- \<Longrightarrow> exec_step_ind (StepI (New C)) P h stk loc C\<^sub>0 M\<^sub>0 pc (Called Cs) frs sh
- (\<lfloor>addr_of_sys_xcpt OutOfMemory\<rfloor>, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, No_ics)#frs, sh)"
-
-| exec_step_ind_NewOOM_Done:
-"\<lbrakk> sh C = Some(obj, Done); new_Addr h = None; \<forall>Cs. ics \<noteq> Called Cs \<rbrakk>
- \<Longrightarrow> exec_step_ind (StepI (New C)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
- (\<lfloor>addr_of_sys_xcpt OutOfMemory\<rfloor>, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, ics)#frs, sh)"
-
-| exec_step_ind_New_Called:
-"new_Addr h = Some a
- \<Longrightarrow> exec_step_ind (StepI (New C)) P h stk loc C\<^sub>0 M\<^sub>0 pc (Called Cs) frs sh
- (None, h(a\<mapsto>blank P C), (Addr a#stk, loc, C\<^sub>0, M\<^sub>0, Suc pc, No_ics)#frs, sh)"
-
-| exec_step_ind_New_Done:
-"\<lbrakk> sh C = Some(obj, Done); new_Addr h = Some a; \<forall>Cs. ics \<noteq> Called Cs \<rbrakk>
- \<Longrightarrow> exec_step_ind (StepI (New C)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
- (None, h(a\<mapsto>blank P C), (Addr a#stk, loc, C\<^sub>0, M\<^sub>0, Suc pc, ics)#frs, sh)"
-
-| exec_step_ind_New_Init:
-"\<lbrakk> \<forall>obj. sh C \<noteq> Some(obj, Done); \<forall>Cs. ics \<noteq> Called Cs \<rbrakk>
- \<Longrightarrow> exec_step_ind (StepI (New C)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
- (None, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, Calling C [])#frs, sh)"
-
-| exec_step_ind_Getfield_Null:
-"hd stk = Null
- \<Longrightarrow> exec_step_ind (StepI (Getfield F C)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
- (\<lfloor>addr_of_sys_xcpt NullPointer\<rfloor>, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, ics)#frs, sh)"
-
-| exec_step_ind_Getfield_NoField:
-"\<lbrakk> v = hd stk; (D,fs) = the(h(the_Addr v)); v \<noteq> Null; \<not>(\<exists>t b. P \<turnstile> D has F,b:t in C) \<rbrakk>
- \<Longrightarrow> exec_step_ind (StepI (Getfield F C)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
- (\<lfloor>addr_of_sys_xcpt NoSuchFieldError\<rfloor>, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, ics)#frs, sh)"
-
-| exec_step_ind_Getfield_Static:
-"\<lbrakk> v = hd stk; (D,fs) = the(h(the_Addr v)); v \<noteq> Null; P \<turnstile> D has F,Static:t in C \<rbrakk>
- \<Longrightarrow> exec_step_ind (StepI (Getfield F C)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
- (\<lfloor>addr_of_sys_xcpt IncompatibleClassChangeError\<rfloor>, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, ics)#frs, sh)"
-
-| exec_step_ind_Getfield:
-"\<lbrakk> v = hd stk; (D,fs) = the(h(the_Addr v)); (D',b,t) = field P C F; v \<noteq> Null;
- P \<turnstile> D has F,NonStatic:t in C \<rbrakk>
- \<Longrightarrow> exec_step_ind (StepI (Getfield F C)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
- (None, h, (the(fs(F,C))#(tl stk), loc, C\<^sub>0, M\<^sub>0, pc+1, ics)#frs, sh)"
-
-| exec_step_ind_Getstatic_NoField:
-"\<not>(\<exists>t b. P \<turnstile> C has F,b:t in D)
- \<Longrightarrow> exec_step_ind (StepI (Getstatic C F D)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
- (\<lfloor>addr_of_sys_xcpt NoSuchFieldError\<rfloor>, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, ics)#frs, sh)"
-
-| exec_step_ind_Getstatic_NonStatic:
-"P \<turnstile> C has F,NonStatic:t in D
- \<Longrightarrow> exec_step_ind (StepI (Getstatic C F D)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
- (\<lfloor>addr_of_sys_xcpt IncompatibleClassChangeError\<rfloor>, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, ics)#frs, sh)"
-
-| exec_step_ind_Getstatic_Called:
-"\<lbrakk> (D',b,t) = field P D F; P \<turnstile> C has F,Static:t in D;
- v = the ((fst(the(sh D'))) F) \<rbrakk>
- \<Longrightarrow> exec_step_ind (StepI (Getstatic C F D)) P h stk loc C\<^sub>0 M\<^sub>0 pc (Called Cs) frs sh
- (None, h, (v#stk, loc, C\<^sub>0, M\<^sub>0, Suc pc, No_ics)#frs, sh)"
-
-| exec_step_ind_Getstatic_Done:
-"\<lbrakk> (D',b,t) = field P D F; P \<turnstile> C has F,Static:t in D;
- \<forall>Cs. ics \<noteq> Called Cs; sh D' = Some(sfs,Done);
- v = the (sfs F) \<rbrakk>
- \<Longrightarrow> exec_step_ind (StepI (Getstatic C F D)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
- (None, h, (v#stk, loc, C\<^sub>0, M\<^sub>0, Suc pc, ics)#frs, sh)"
-
-| exec_step_ind_Getstatic_Init:
-"\<lbrakk> (D',b,t) = field P D F; P \<turnstile> C has F,Static:t in D;
- \<forall>sfs. sh D' \<noteq> Some(sfs,Done); \<forall>Cs. ics \<noteq> Called Cs \<rbrakk>
- \<Longrightarrow> exec_step_ind (StepI (Getstatic C F D)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
- (None, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, Calling D' [])#frs, sh)"
-
-| exec_step_ind_Putfield_Null:
-"hd(tl stk) = Null
- \<Longrightarrow> exec_step_ind (StepI (Putfield F C)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
- (\<lfloor>addr_of_sys_xcpt NullPointer\<rfloor>, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, ics)#frs, sh)"
-
-| exec_step_ind_Putfield_NoField:
-"\<lbrakk> r = hd(tl stk); a = the_Addr r; (D,fs) = the (h a); r \<noteq> Null; \<not>(\<exists>t b. P \<turnstile> D has F,b:t in C) \<rbrakk>
- \<Longrightarrow> exec_step_ind (StepI (Putfield F C)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
- (\<lfloor>addr_of_sys_xcpt NoSuchFieldError\<rfloor>, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, ics)#frs, sh)"
-
-| exec_step_ind_Putfield_Static:
-"\<lbrakk> r = hd(tl stk); a = the_Addr r; (D,fs) = the (h a); r \<noteq> Null; P \<turnstile> D has F,Static:t in C \<rbrakk>
- \<Longrightarrow> exec_step_ind (StepI (Putfield F C)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
- (\<lfloor>addr_of_sys_xcpt IncompatibleClassChangeError\<rfloor>, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, ics)#frs, sh)"
-
-| exec_step_ind_Putfield:
-"\<lbrakk> v = hd stk; r = hd(tl stk); a = the_Addr r; (D,fs) = the (h a); (D',b,t) = field P C F;
- r \<noteq> Null; P \<turnstile> D has F,NonStatic:t in C \<rbrakk>
- \<Longrightarrow> exec_step_ind (StepI (Putfield F C)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
- (None, h(a \<mapsto> (D, fs((F,C) \<mapsto> v))), (tl (tl stk), loc, C\<^sub>0, M\<^sub>0, pc+1, ics)#frs, sh)"
-
-| exec_step_ind_Putstatic_NoField:
-"\<not>(\<exists>t b. P \<turnstile> C has F,b:t in D)
- \<Longrightarrow> exec_step_ind (StepI (Putstatic C F D)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
- (\<lfloor>addr_of_sys_xcpt NoSuchFieldError\<rfloor>, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, ics)#frs, sh)"
-
-| exec_step_ind_Putstatic_NonStatic:
-"P \<turnstile> C has F,NonStatic:t in D
- \<Longrightarrow> exec_step_ind (StepI (Putstatic C F D)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
- (\<lfloor>addr_of_sys_xcpt IncompatibleClassChangeError\<rfloor>, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, ics)#frs, sh)"
-
-| exec_step_ind_Putstatic_Called:
-"\<lbrakk> (D',b,t) = field P D F; P \<turnstile> C has F,Static:t in D; the(sh D') = (sfs,i) \<rbrakk>
- \<Longrightarrow> exec_step_ind (StepI (Putstatic C F D)) P h stk loc C\<^sub>0 M\<^sub>0 pc (Called Cs) frs sh
- (None, h, (tl stk, loc, C\<^sub>0, M\<^sub>0, Suc pc, No_ics)#frs, sh(D':=Some ((sfs(F \<mapsto> hd stk)), i)))"
-
-| exec_step_ind_Putstatic_Done:
-"\<lbrakk> (D',b,t) = field P D F; P \<turnstile> C has F,Static:t in D;
- \<forall>Cs. ics \<noteq> Called Cs; sh D' = Some (sfs, Done) \<rbrakk>
- \<Longrightarrow> exec_step_ind (StepI (Putstatic C F D)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
- (None, h, (tl stk, loc, C\<^sub>0, M\<^sub>0, Suc pc, ics)#frs, sh(D':=Some ((sfs(F \<mapsto> hd stk)), Done)))"
-
-| exec_step_ind_Putstatic_Init:
-"\<lbrakk> (D',b,t) = field P D F; P \<turnstile> C has F,Static:t in D;
- \<forall>sfs. sh D' \<noteq> Some (sfs, Done); \<forall>Cs. ics \<noteq> Called Cs \<rbrakk>
- \<Longrightarrow> exec_step_ind (StepI (Putstatic C F D)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
- (None, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, Calling D' [])#frs, sh)"
-
-| exec_step_ind_Checkcast:
-"cast_ok P C h (hd stk)
- \<Longrightarrow> exec_step_ind (StepI (Checkcast C)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
- (None, h, (stk, loc, C\<^sub>0, M\<^sub>0, Suc pc, ics)#frs, sh)"
-
-| exec_step_ind_Checkcast_Error:
-"\<not>cast_ok P C h (hd stk)
- \<Longrightarrow> exec_step_ind (StepI (Checkcast C)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
- (\<lfloor>addr_of_sys_xcpt ClassCast\<rfloor>, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, ics)#frs, sh)"
-
-| exec_step_ind_Invoke_Null:
-"stk!n = Null
- \<Longrightarrow> exec_step_ind (StepI (Invoke M n)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
- (\<lfloor>addr_of_sys_xcpt NullPointer\<rfloor>, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, ics)#frs, sh)"
-
-| exec_step_ind_Invoke_NoMethod:
-"\<lbrakk> r = stk!n; C = fst(the(h(the_Addr r))); r \<noteq> Null;
- \<not>(\<exists>Ts T m D b. P \<turnstile> C sees M,b:Ts \<rightarrow> T = m in D) \<rbrakk>
- \<Longrightarrow> exec_step_ind (StepI (Invoke M n)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
- (\<lfloor>addr_of_sys_xcpt NoSuchMethodError\<rfloor>, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, ics)#frs, sh)"
-
-| exec_step_ind_Invoke_Static:
-"\<lbrakk> r = stk!n; C = fst(the(h(the_Addr r)));
- (D,b,Ts,T,mxs,mxl\<^sub>0,ins,xt)= method P C M; r \<noteq> Null;
- P \<turnstile> C sees M,Static:Ts \<rightarrow> T = m in D \<rbrakk>
- \<Longrightarrow> exec_step_ind (StepI (Invoke M n)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
- (\<lfloor>addr_of_sys_xcpt IncompatibleClassChangeError\<rfloor>, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, ics)#frs, sh)"
-
-| exec_step_ind_Invoke:
-"\<lbrakk> ps = take n stk; r = stk!n; C = fst(the(h(the_Addr r)));
- (D,b,Ts,T,mxs,mxl\<^sub>0,ins,xt)= method P C M; r \<noteq> Null;
- P \<turnstile> C sees M,NonStatic:Ts \<rightarrow> T = m in D;
- f' = ([],[r]@(rev ps)@(replicate mxl\<^sub>0 undefined),D,M,0,No_ics) \<rbrakk>
- \<Longrightarrow> exec_step_ind (StepI (Invoke M n)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
- (None, h, f'#(stk, loc, C\<^sub>0, M\<^sub>0, pc, ics)#frs, sh)"
-
-| exec_step_ind_Invokestatic_NoMethod:
-"\<lbrakk> (D,b,Ts,T,mxs,mxl\<^sub>0,ins,xt)= method P C M; \<not>(\<exists>Ts T m D b. P \<turnstile> C sees M,b:Ts \<rightarrow> T = m in D) \<rbrakk>
- \<Longrightarrow> exec_step_ind (StepI (Invokestatic C M n)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
- (\<lfloor>addr_of_sys_xcpt NoSuchMethodError\<rfloor>, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, ics)#frs, sh)"
-
-| exec_step_ind_Invokestatic_NonStatic:
-"\<lbrakk> (D,b,Ts,T,mxs,mxl\<^sub>0,ins,xt)= method P C M; P \<turnstile> C sees M,NonStatic:Ts \<rightarrow> T = m in D \<rbrakk>
- \<Longrightarrow> exec_step_ind (StepI (Invokestatic C M n)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
- (\<lfloor>addr_of_sys_xcpt IncompatibleClassChangeError\<rfloor>, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, ics)#frs, sh)"
-
-| exec_step_ind_Invokestatic_Called:
-"\<lbrakk> ps = take n stk; (D,b,Ts,T,mxs,mxl\<^sub>0,ins,xt) = method P C M;
- P \<turnstile> C sees M,Static:Ts \<rightarrow> T = m in D;
- f' = ([],(rev ps)@(replicate mxl\<^sub>0 undefined),D,M,0,No_ics) \<rbrakk>
- \<Longrightarrow> exec_step_ind (StepI (Invokestatic C M n)) P h stk loc C\<^sub>0 M\<^sub>0 pc (Called Cs) frs sh
- (None, h, f'#(stk, loc, C\<^sub>0, M\<^sub>0, pc, No_ics)#frs, sh)"
-
-| exec_step_ind_Invokestatic_Done:
-"\<lbrakk> ps = take n stk; (D,b,Ts,T,mxs,mxl\<^sub>0,ins,xt) = method P C M;
- P \<turnstile> C sees M,Static:Ts \<rightarrow> T = m in D;
- \<forall>Cs. ics \<noteq> Called Cs; sh D = Some (sfs, Done);
- f' = ([],(rev ps)@(replicate mxl\<^sub>0 undefined),D,M,0,No_ics) \<rbrakk>
- \<Longrightarrow> exec_step_ind (StepI (Invokestatic C M n)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
- (None, h, f'#(stk, loc, C\<^sub>0, M\<^sub>0, pc, ics)#frs, sh)"
-
-| exec_step_ind_Invokestatic_Init:
-"\<lbrakk> (D,b,Ts,T,mxs,mxl\<^sub>0,ins,xt) = method P C M;
- P \<turnstile> C sees M,Static:Ts \<rightarrow> T = m in D;
- \<forall>sfs. sh D \<noteq> Some (sfs, Done); \<forall>Cs. ics \<noteq> Called Cs \<rbrakk>
- \<Longrightarrow> exec_step_ind (StepI (Invokestatic C M n)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
- (None, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, Calling D [])#frs, sh)"
-
-| exec_step_ind_Return_Last_Init:
- "exec_step_ind (StepI Return) P h stk\<^sub>0 loc\<^sub>0 C\<^sub>0 clinit pc ics [] sh
- (None, h, [], sh(C\<^sub>0:=Some(fst(the(sh C\<^sub>0)), Done)))"
-
-| exec_step_ind_Return_Last:
- "M\<^sub>0 \<noteq> clinit
- \<Longrightarrow> exec_step_ind (StepI Return) P h stk\<^sub>0 loc\<^sub>0 C\<^sub>0 M\<^sub>0 pc ics [] sh (None, h, [], sh)"
-
-| exec_step_ind_Return_Init:
- "\<lbrakk> (D,b,Ts,T,m) = method P C\<^sub>0 clinit \<rbrakk>
- \<Longrightarrow> exec_step_ind (StepI Return) P h stk\<^sub>0 loc\<^sub>0 C\<^sub>0 clinit pc ics ((stk',loc',C',m',pc',ics')#frs') sh
- (None, h, (stk',loc',C',m',pc',ics')#frs', sh(C\<^sub>0:=Some(fst(the(sh C\<^sub>0)), Done)))"
-
-| exec_step_ind_Return_NonStatic:
- "\<lbrakk> (D,NonStatic,Ts,T,m) = method P C\<^sub>0 M\<^sub>0; M\<^sub>0 \<noteq> clinit \<rbrakk>
- \<Longrightarrow> exec_step_ind (StepI Return) P h stk\<^sub>0 loc\<^sub>0 C\<^sub>0 M\<^sub>0 pc ics ((stk',loc',C',m',pc',ics')#frs') sh
- (None, h, ((hd stk\<^sub>0)#(drop (length Ts + 1) stk'),loc',C',m',Suc pc',ics')#frs', sh)"
-
-| exec_step_ind_Return_Static:
- "\<lbrakk> (D,Static,Ts,T,m) = method P C\<^sub>0 M\<^sub>0; M\<^sub>0 \<noteq> clinit \<rbrakk>
- \<Longrightarrow> exec_step_ind (StepI Return) P h stk\<^sub>0 loc\<^sub>0 C\<^sub>0 M\<^sub>0 pc ics ((stk',loc',C',m',pc',ics')#frs') sh
- (None, h, ((hd stk\<^sub>0)#(drop (length Ts) stk'),loc',C',m',Suc pc',ics')#frs', sh)"
-
-| exec_step_ind_Pop:
-"exec_step_ind (StepI Pop) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
- (None, h, (tl stk, loc, C\<^sub>0, M\<^sub>0, Suc pc, ics)#frs, sh)"
-
-| exec_step_ind_IAdd:
-"exec_step_ind (StepI IAdd) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
- (None, h, (Intg (the_Intg (hd (tl stk)) + the_Intg (hd stk))#(tl (tl stk)), loc, C\<^sub>0, M\<^sub>0, Suc pc, ics)#frs, sh)"
-
-| exec_step_ind_IfFalse_False:
-"hd stk = Bool False
- \<Longrightarrow> exec_step_ind (StepI (IfFalse i)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
- (None, h, (tl stk, loc, C\<^sub>0, M\<^sub>0, nat(int pc+i), ics)#frs, sh)"
-
-| exec_step_ind_IfFalse_nFalse:
-"hd stk \<noteq> Bool False
- \<Longrightarrow> exec_step_ind (StepI (IfFalse i)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
- (None, h, (tl stk, loc, C\<^sub>0, M\<^sub>0, Suc pc, ics)#frs, sh)"
-
-| exec_step_ind_CmpEq:
-"exec_step_ind (StepI CmpEq) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
- (None, h, (Bool (hd (tl stk) = hd stk) # tl (tl stk), loc, C\<^sub>0, M\<^sub>0, Suc pc, ics)#frs, sh)"
-
-| exec_step_ind_Goto:
-"exec_step_ind (StepI (Goto i)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
- (None, h, (stk, loc, C\<^sub>0, M\<^sub>0, nat(int pc+i), ics)#frs, sh)"
-
-| exec_step_ind_Throw:
-"hd stk \<noteq> Null
- \<Longrightarrow> exec_step_ind (StepI Throw) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
- (\<lfloor>the_Addr (hd stk)\<rfloor>, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, ics)#frs, sh)"
-
-| exec_step_ind_Throw_Null:
-"hd stk = Null
- \<Longrightarrow> exec_step_ind (StepI Throw) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
- (\<lfloor>addr_of_sys_xcpt NullPointer\<rfloor>, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, ics)#frs, sh)"
-
-| exec_step_ind_Init_None_Called:
-"\<lbrakk> sh C = None \<rbrakk>
- \<Longrightarrow> exec_step_ind (StepC C Cs) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
- (None, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, Calling C Cs)#frs, sh(C := Some (sblank P C, Prepared)))"
-
-| exec_step_ind_Init_Done:
-"sh C = Some (sfs, Done)
- \<Longrightarrow> exec_step_ind (StepC C Cs) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
- (None, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, Called Cs)#frs, sh)"
-
-| exec_step_ind_Init_Processing:
-"sh C = Some (sfs, Processing)
- \<Longrightarrow> exec_step_ind (StepC C Cs) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
- (None, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, Called Cs)#frs, sh)"
-
-| exec_step_ind_Init_Error:
-"\<lbrakk> sh C = Some (sfs, Error) \<rbrakk>
- \<Longrightarrow> exec_step_ind (StepC C Cs) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
- (None, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, Throwing Cs (addr_of_sys_xcpt NoClassDefFoundError))#frs, sh)"
-
-| exec_step_ind_Init_Prepared_Object:
-"\<lbrakk> sh C = Some (sfs, Prepared);
- sh' = sh(C:=Some(fst(the(sh C)), Processing));
- C = Object \<rbrakk>
- \<Longrightarrow> exec_step_ind (StepC C Cs) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
- (None, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, Called (C#Cs))#frs, sh')"
-
-| exec_step_ind_Init_Prepared_nObject:
-"\<lbrakk> sh C = Some (sfs, Prepared);
- sh' = sh(C:=Some(fst(the(sh C)), Processing));
- C \<noteq> Object; D = fst(the(class P C)) \<rbrakk>
- \<Longrightarrow> exec_step_ind (StepC C Cs) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
- (None, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, Calling D (C#Cs))#frs, sh')"
-
-| exec_step_ind_Init:
-"exec_step_ind (StepC2 C Cs) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
- (None, h, create_init_frame P C#(stk, loc, C\<^sub>0, M\<^sub>0, pc, Called Cs)#frs, sh)"
-
-| exec_step_ind_InitThrow:
-"exec_step_ind (StepT (C#Cs) a) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
- (None, h, (stk,loc,C\<^sub>0,M\<^sub>0,pc,Throwing Cs a)#frs, (sh(C \<mapsto> (fst(the(sh C)), Error))))"
-
-| exec_step_ind_InitThrow_End:
-"exec_step_ind (StepT [] a) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
- (\<lfloor>a\<rfloor>, h, (stk,loc,C\<^sub>0,M\<^sub>0,pc,No_ics)#frs, sh)"
-
-(** ******* **)
-
-inductive_cases exec_step_ind_cases [cases set]:
- "exec_step_ind (StepI (Load n)) P h stk loc C M pc ics frs sh \<sigma>"
- "exec_step_ind (StepI (Store n)) P h stk loc C M pc ics frs sh \<sigma>"
- "exec_step_ind (StepI (Push v)) P h stk loc C M pc ics frs sh \<sigma>"
- "exec_step_ind (StepI (New C)) P h stk loc C M pc ics frs sh \<sigma>"
- "exec_step_ind (StepI (Getfield F C)) P h stk loc C M pc ics frs sh \<sigma>"
- "exec_step_ind (StepI (Getstatic C F D)) P h stk loc C M pc ics frs sh \<sigma>"
- "exec_step_ind (StepI (Putfield F C)) P h stk loc C M pc ics frs sh \<sigma>"
- "exec_step_ind (StepI (Putstatic C F D)) P h stk loc C M pc ics frs sh \<sigma>"
- "exec_step_ind (StepI (Checkcast C)) P h stk loc C M pc ics frs sh \<sigma>"
- "exec_step_ind (StepI (Invoke M n)) P h stk loc C M pc ics frs sh \<sigma>"
- "exec_step_ind (StepI (Invokestatic C M n)) P h stk loc C M pc ics frs sh \<sigma>"
- "exec_step_ind (StepI Return) P h stk loc C M pc ics frs sh \<sigma>"
- "exec_step_ind (StepI Pop) P h stk loc C M pc ics frs sh \<sigma>"
- "exec_step_ind (StepI IAdd) P h stk loc C M pc ics frs sh \<sigma>"
- "exec_step_ind (StepI (IfFalse i)) P h stk loc C M pc ics frs sh \<sigma>"
- "exec_step_ind (StepI CmpEq) P h stk loc C M pc ics frs sh \<sigma>"
- "exec_step_ind (StepI (Goto i)) P h stk loc C M pc ics frs sh \<sigma>"
- "exec_step_ind (StepI Throw) P h stk loc C M pc ics frs sh \<sigma>"
- "exec_step_ind (StepC C' Cs) P h stk loc C M pc ics frs sh \<sigma>"
- "exec_step_ind (StepC2 C' Cs) P h stk loc C M pc ics frs sh \<sigma>"
- "exec_step_ind (StepT Cs a) P h stk loc C M pc ics frs sh \<sigma>"
-
-
-\<comment> \<open> Deriving @{term step_input} for @{term exec_step_ind} from @{term exec_step} arguments \<close>
-fun exec_step_input :: "[jvm_prog, cname, mname, pc, init_call_status] \<Rightarrow> step_input" where
-"exec_step_input P C M pc (Calling C' Cs) = StepC C' Cs" |
-"exec_step_input P C M pc (Called (C'#Cs)) = StepC2 C' Cs" |
-"exec_step_input P C M pc (Throwing Cs a) = StepT Cs a" |
-"exec_step_input P C M pc ics = StepI (instrs_of P C M ! pc)"
-
-lemma exec_step_input_StepTD[simp]:
-assumes "exec_step_input P C M pc ics = StepT Cs a" shows "ics = Throwing Cs a"
-using assms proof(cases ics)
- case (Called Cs) with assms show ?thesis by(cases Cs; simp)
-qed(auto)
-
-lemma exec_step_input_StepCD[simp]:
-assumes "exec_step_input P C M pc ics = StepC C' Cs" shows "ics = Calling C' Cs"
-using assms proof(cases ics)
- case (Called Cs) with assms show ?thesis by(cases Cs; simp)
-qed(auto)
-
-lemma exec_step_input_StepC2D[simp]:
-assumes "exec_step_input P C M pc ics = StepC2 C' Cs" shows "ics = Called (C'#Cs)"
-using assms proof(cases ics)
- case (Called Cs) with assms show ?thesis by(cases Cs; simp)
-qed(auto)
-
-lemma exec_step_input_StepID:
-assumes "exec_step_input P C M pc ics = StepI i"
-shows "(ics = Called [] \<or> ics = No_ics) \<and> instrs_of P C M ! pc = i"
-using assms proof(cases ics)
- case (Called Cs) with assms show ?thesis by(cases Cs; simp)
-qed(auto)
-
-subsection "Equivalence of @{term exec_step} and @{term exec_step_input}"
-
-lemma exec_step_imp_exec_step_ind:
-assumes es: "exec_step P h stk loc C M pc ics frs sh = (xp', h', frs', sh')"
-shows "exec_step_ind (exec_step_input P C M pc ics) P h stk loc C M pc ics frs sh (xp', h', frs', sh')"
-proof(cases "exec_step_input P C M pc ics")
- case (StepT Cs a)
- then have "ics = Throwing Cs a" by simp
- then show ?thesis using exec_step_ind_InitThrow exec_step_ind_InitThrow_End StepT es
- by(cases Cs, auto)
-next
- case (StepC C1 Cs)
- then have ics: "ics = Calling C1 Cs" by simp
- obtain D b Ts T m where lets: "method P C1 clinit = (D,b,Ts,T,m)" by(cases "method P C1 clinit")
- then obtain mxs mxl\<^sub>0 ins xt where m: "m = (mxs,mxl\<^sub>0,ins,xt)" by(cases m)
- show ?thesis
- proof(cases "sh C1")
- case None then show ?thesis
- using exec_step_ind_Init_None_Called ics assms by auto
- next
- case (Some a)
- then obtain sfs i where sfsi: "a = (sfs,i)" by(cases a)
- then show ?thesis using exec_step_ind_Init_Done exec_step_ind_Init_Processing
- exec_step_ind_Init_Error m lets Some ics assms
- proof(cases i)
- case Prepared
- show ?thesis
- using exec_step_ind_Init_Prepared_Object[where P=P] exec_step_ind_Init_Prepared_nObject
- sfsi m lets Prepared Some ics assms by(auto split: if_split_asm)
- qed(auto)
- qed
-next
- case (StepC2 C1 Cs)
- then have ics: "ics = Called (C1#Cs)" by simp
- then show ?thesis using exec_step_ind_Init assms by auto
-next
- case (StepI i)
- then have
- ics: "ics = Called [] \<or> ics = No_ics" and
- exec_instr: "exec_instr i P h stk loc C M pc ics frs sh = (xp', h', frs', sh')"
- using assms by(auto dest!: exec_step_input_StepID)
- show ?thesis
- proof(cases i)
- case (Load x1) then show ?thesis using exec_instr exec_step_ind_Load StepI by auto
- next
- case (Store x2) then show ?thesis using exec_instr exec_step_ind_Store StepI by auto
- next
- case (Push x3) then show ?thesis using exec_instr exec_step_ind_Push StepI by auto
- next
- case (New C1)
- then obtain sfs i where sfsi: "the(sh C1) = (sfs,i)" by(cases "the(sh C1)")
- then show ?thesis using exec_step_ind_New_Called exec_step_ind_NewOOM_Called
- exec_step_ind_New_Done exec_step_ind_NewOOM_Done
- exec_step_ind_New_Init sfsi New StepI exec_instr ics by(auto split: init_state.splits)
- next
- case (Getfield F1 C1)
- then obtain D fs D' b t where lets: "the(h(the_Addr (hd stk))) = (D,fs)"
- "field P C1 F1 = (D',b,t)" by(cases "the(h(the_Addr (hd stk)))", cases "field P C1 F1")
- then have "\<And>b' t'. P \<turnstile> D has F1,b':t' in C1 \<Longrightarrow> (D', b, t) = (C1, b', t')"
- using field_def2 has_field_idemp has_field_sees by fastforce
- then show ?thesis using exec_step_ind_Getfield_Null exec_step_ind_Getfield_NoField
- exec_step_ind_Getfield_Static exec_step_ind_Getfield lets Getfield StepI exec_instr
- by(auto split: if_split_asm staticb.splits) metis+
- next
- case (Getstatic C1 F1 D1)
- then obtain D' b t where lets: "field P D1 F1 = (D',b,t)" by(cases "field P D1 F1")
- then have field: "\<And>b' t'. P \<turnstile> C1 has F1,b':t' in D1 \<Longrightarrow> (D', b, t) = (D1, b', t')"
- using field_def2 has_field_idemp has_field_sees by fastforce
- show ?thesis
- proof(cases b)
- case NonStatic then show ?thesis
- using exec_step_ind_Getstatic_NoField exec_step_ind_Getstatic_NonStatic
- field lets Getstatic exec_instr StepI by(auto split: if_split_asm) fastforce
- next
- case Static show ?thesis
- proof(cases "ics = Called []")
- case True then show ?thesis using exec_step_ind_Getstatic_NoField
- exec_step_ind_Getstatic_Called exec_step_ind_Getstatic_Init
- Static field lets Getstatic exec_instr StepI ics
- by(auto simp: split_beta split: if_split_asm) metis
- next
- case False
- then have nCalled: "\<forall>Cs. ics \<noteq> Called Cs" using ics by simp
- show ?thesis
- proof(cases "sh D1")
- case None
- then have nDone: "\<forall>sfs. sh D1 \<noteq> Some(sfs, Done)" by simp
- then show ?thesis using exec_step_ind_Getstatic_NoField
- exec_step_ind_Getstatic_Init[where sh=sh, OF _ _ nDone nCalled]
- field lets None False Static Getstatic exec_instr StepI ics
- by(auto split: if_split_asm) metis
- next
- case (Some a)
- then obtain sfs i where sfsi: "a=(sfs,i)" by(cases a)
- show ?thesis using exec_step_ind_Getstatic_NoField
- exec_step_ind_Getstatic_Init sfsi False Static Some field lets Getstatic exec_instr
- proof(cases "i = Done")
- case True then show ?thesis using exec_step_ind_Getstatic_NoField
- exec_step_ind_Getstatic_Done[OF _ _ nCalled] exec_step_ind_Getstatic_Init
- sfsi False Static Some field lets Getstatic exec_instr StepI ics
- by(auto split: if_split_asm) metis
- next
- case nD: False
- then have nDone: "\<forall>sfs. sh D1 \<noteq> Some(sfs, Done)" using sfsi Some by simp
- show ?thesis using nD
- proof(cases i)
- case Processing then show ?thesis using exec_step_ind_Getstatic_NoField
- exec_step_ind_Getstatic_Init[where sh=sh, OF _ _ nDone nCalled]
- sfsi False Static Some field lets Getstatic exec_instr StepI ics
- by(auto split: if_split_asm) metis
- next
- case Prepared then show ?thesis using exec_step_ind_Getstatic_NoField
- exec_step_ind_Getstatic_Init[where sh=sh, OF _ _ nDone nCalled]
- sfsi False Static Some field lets Getstatic exec_instr StepI ics
- by(auto split: if_split_asm) metis
- next
- case Error then show ?thesis using exec_step_ind_Getstatic_NoField
- exec_step_ind_Getstatic_Init[where sh=sh, OF _ _ nDone nCalled]
- sfsi False Static Some field lets Getstatic exec_instr StepI ics
- by(auto split: if_split_asm) metis
- qed(simp)
- qed
- qed
- qed
- qed
- next
- case (Putfield F1 C1)
- then obtain D fs D' b t where lets: "the(h(the_Addr (hd(tl stk)))) = (D,fs)"
- "field P C1 F1 = (D',b,t)" by(cases "the(h(the_Addr (hd(tl stk))))", cases "field P C1 F1")
- then have "\<And>b' t'. P \<turnstile> D has F1,b':t' in C1 \<Longrightarrow> (D', b, t) = (C1, b', t')"
- using field_def2 has_field_idemp has_field_sees by fastforce
- then show ?thesis using exec_step_ind_Putfield_Null exec_step_ind_Putfield_NoField
- exec_step_ind_Putfield_Static exec_step_ind_Putfield lets Putfield exec_instr StepI
- by(auto split: if_split_asm staticb.splits) metis+
- next
- case (Putstatic C1 F1 D1)
- then obtain D' b t where lets: "field P D1 F1 = (D',b,t)" by(cases "field P D1 F1")
- then have field: "\<And>b' t'. P \<turnstile> C1 has F1,b':t' in D1 \<Longrightarrow> (D', b, t) = (D1, b', t')"
- using field_def2 has_field_idemp has_field_sees by fastforce
- show ?thesis
- proof(cases b)
- case NonStatic then show ?thesis
- using exec_step_ind_Putstatic_NoField exec_step_ind_Putstatic_NonStatic
- field lets Putstatic exec_instr StepI by(auto split: if_split_asm) fastforce
- next
- case Static show ?thesis
- proof(cases "ics = Called []")
- case True then show ?thesis using exec_step_ind_Putstatic_NoField
- exec_step_ind_Putstatic_Called exec_step_ind_Putstatic_Init
- Static field lets Putstatic exec_instr StepI ics
- by(cases "the(sh D1)", auto split: if_split_asm) metis
- next
- case False
- then have nCalled: "\<forall>Cs. ics \<noteq> Called Cs" using ics by simp
- show ?thesis
- proof(cases "sh D1")
- case None
- then have nDone: "\<forall>sfs. sh D1 \<noteq> Some(sfs, Done)" by simp
- then show ?thesis using exec_step_ind_Putstatic_NoField
- exec_step_ind_Putstatic_Init[where sh=sh, OF _ _ nDone nCalled]
- field lets None False Static Putstatic exec_instr StepI ics
- by(auto split: if_split_asm) metis
- next
- case (Some a)
- then obtain sfs i where sfsi: "a=(sfs,i)" by(cases a)
- show ?thesis using exec_step_ind_Putstatic_NoField
- exec_step_ind_Putstatic_Init sfsi False Static Some field lets Putstatic exec_instr
- proof(cases "i = Done")
- case True then show ?thesis using exec_step_ind_Putstatic_NoField
- exec_step_ind_Putstatic_Done[OF _ _ nCalled] exec_step_ind_Putstatic_Init
- sfsi False Static Some field lets Putstatic exec_instr StepI ics
- by(auto split: if_split_asm) metis
- next
- case nD: False
- then have nDone: "\<forall>sfs. sh D1 \<noteq> Some(sfs, Done)" using sfsi Some by simp
- show ?thesis using nD
- proof(cases i)
- case Processing then show ?thesis using exec_step_ind_Putstatic_NoField
- exec_step_ind_Putstatic_Init[where sh=sh, OF _ _ nDone nCalled]
- sfsi False Static Some field lets Putstatic exec_instr StepI ics
- by(auto split: if_split_asm) metis
- next
- case Prepared then show ?thesis using exec_step_ind_Putstatic_NoField
- exec_step_ind_Putstatic_Init[where sh=sh, OF _ _ nDone nCalled]
- sfsi False Static Some field lets Putstatic exec_instr StepI ics
- by(auto split: if_split_asm) metis
- next
- case Error then show ?thesis using exec_step_ind_Putstatic_NoField
- exec_step_ind_Putstatic_Init[where sh=sh, OF _ _ nDone nCalled]
- sfsi False Static Some field lets Putstatic exec_instr StepI ics
- by(auto split: if_split_asm) metis
- qed(simp)
- qed
- qed
- qed
- qed
- next
- case Checkcast then show ?thesis
- using exec_step_ind_Checkcast exec_step_ind_Checkcast_Error exec_instr StepI
- by(auto split: if_split_asm)
- next
- case (Invoke M1 n) show ?thesis
- proof(cases "stk!n = Null")
- case True then show ?thesis using exec_step_ind_Invoke_Null Invoke exec_instr StepI
- by clarsimp
- next
- case False
- let ?C = "cname_of h (the_Addr (stk ! n))"
- obtain D b Ts T m where method: "method P ?C M1 = (D,b,Ts,T,m)" by(cases "method P ?C M1")
- then obtain mxs mxl\<^sub>0 ins xt where "m = (mxs,mxl\<^sub>0,ins,xt)" by(cases m)
- then show ?thesis using exec_step_ind_Invoke_NoMethod
- exec_step_ind_Invoke_Static exec_step_ind_Invoke method False Invoke exec_instr StepI
- by(auto split: if_split_asm staticb.splits)
- qed
- next
- case (Invokestatic C1 M1 n)
- obtain D b Ts T m where lets: "method P C1 M1 = (D,b,Ts,T,m)" by(cases "method P C1 M1")
- then obtain mxs mxl\<^sub>0 ins xt where m: "m = (mxs,mxl\<^sub>0,ins,xt)" by(cases m)
- have method: "\<And>b' Ts' t' m' D'. P \<turnstile> C1 sees M1,b':Ts' \<rightarrow> t' = m' in D'
- \<Longrightarrow> (D,b,Ts,T,m) = (D',b',Ts',t',m')" using lets by auto
- show ?thesis
- proof(cases b)
- case NonStatic then show ?thesis
- using exec_step_ind_Invokestatic_NoMethod exec_step_ind_Invokestatic_NonStatic
- m method lets Invokestatic exec_instr StepI by(auto split: if_split_asm)
- next
- case Static show ?thesis
- proof(cases "ics = Called []")
- case True then show ?thesis using exec_step_ind_Invokestatic_NoMethod
- exec_step_ind_Invokestatic_Called exec_step_ind_Invokestatic_Init
- Static m method lets Invokestatic exec_instr StepI ics
- by(auto split: if_split_asm)
- next
- case False
- then have nCalled: "\<forall>Cs. ics \<noteq> Called Cs" using ics by simp
- show ?thesis
- proof(cases "sh D")
- case None
- then have nDone: "\<forall>sfs. sh D \<noteq> Some(sfs, Done)" by simp
- show ?thesis using exec_step_ind_Invokestatic_NoMethod
- exec_step_ind_Invokestatic_Init[where sh=sh, OF _ _ nDone nCalled]
- method m lets None False Static Invokestatic exec_instr StepI ics
- by(auto split: if_split_asm)
- next
- case (Some a)
- then obtain sfs i where sfsi: "a=(sfs,i)" by(cases a)
- show ?thesis using exec_step_ind_Invokestatic_NoMethod
- exec_step_ind_Invokestatic_Init sfsi False Static Some method lets Invokestatic exec_instr
- proof(cases "i = Done")
- case True then show ?thesis using exec_step_ind_Invokestatic_NoMethod
- exec_step_ind_Invokestatic_Done[OF _ _ _ nCalled] exec_step_ind_Invokestatic_Init
- sfsi False Static Some m method lets Invokestatic exec_instr StepI ics
- by(auto split: if_split_asm)
- next
- case nD: False
- then have nDone: "\<forall>sfs. sh D \<noteq> Some(sfs, Done)" using sfsi Some by simp
- show ?thesis using nD
- proof(cases i)
- case Processing then show ?thesis using exec_step_ind_Invokestatic_NoMethod
- exec_step_ind_Invokestatic_Init[where sh=sh, OF _ _ nDone nCalled]
- sfsi False Static Some m method lets Invokestatic exec_instr StepI ics
- by(auto split: if_split_asm)
- next
- case Prepared then show ?thesis using exec_step_ind_Invokestatic_NoMethod
- exec_step_ind_Invokestatic_Init[where sh=sh, OF _ _ nDone nCalled]
- sfsi False Static Some m method lets Invokestatic exec_instr StepI ics
- by(auto split: if_split_asm)
- next
- case Error then show ?thesis using exec_step_ind_Invokestatic_NoMethod
- exec_step_ind_Invokestatic_Init[where sh=sh, OF _ _ nDone nCalled]
- sfsi False Static Some m method lets Invokestatic exec_instr StepI ics
- by(auto split: if_split_asm)
- qed(simp)
- qed
- qed
- qed
- qed
- next
- case Return
- obtain D b Ts T m where method: "method P C M = (D,b,Ts,T,m)" by(cases "method P C M")
- then obtain mxs mxl\<^sub>0 ins xt where "m = (mxs,mxl\<^sub>0,ins,xt)" by(cases m)
- then show ?thesis using exec_step_ind_Return_Last_Init exec_step_ind_Return_Last
- exec_step_ind_Return_Init exec_step_ind_Return_NonStatic exec_step_ind_Return_Static
- method Return exec_instr StepI ics
- by(auto split: if_split_asm staticb.splits bool.splits list.splits)
- next
- case Pop then show ?thesis using exec_instr StepI exec_step_ind_Pop by auto
- next
- case IAdd then show ?thesis using exec_instr StepI exec_step_ind_IAdd by auto
- next
- case Goto then show ?thesis using exec_instr StepI exec_step_ind_Goto by auto
- next
- case CmpEq then show ?thesis using exec_instr StepI exec_step_ind_CmpEq by auto
- next
- case (IfFalse x17) then show ?thesis
- using exec_instr StepI exec_step_ind_IfFalse_nFalse exec_step_ind_IfFalse_False
- exec_instr StepI by(auto split: val.splits staticb.splits)
- next
- case Throw then show ?thesis
- using exec_instr StepI exec_step_ind_Throw exec_step_ind_Throw_Null
- by(auto split: val.splits)
- qed
-qed
-
-lemma exec_step_ind_imp_exec_step:
-assumes esi: "exec_step_ind si P h stk loc C M pc ics frs sh (xp', h', frs', sh')"
- and si: "exec_step_input P C M pc ics = si"
-shows "exec_step P h stk loc C M pc ics frs sh = (xp', h', frs', sh')"
-proof -
- have StepI:
- "\<And>P C M pc Cs i . exec_step_input P C M pc (Called Cs) = StepI i
- \<Longrightarrow> instrs_of P C M ! pc = i \<and> Cs = []"
- proof -
- fix P C M pc Cs i show "exec_step_input P C M pc (Called Cs) = StepI i
- \<Longrightarrow> instrs_of P C M ! pc = i \<and> Cs = []" by(cases Cs; simp)
- qed
- have StepC:
- "\<And>P C M pc ics C' Cs. exec_step_input P C M pc ics = StepC C' Cs \<Longrightarrow> ics = Calling C' Cs"
- by simp
- have StepT:
- "\<And>P C M pc ics Cs a. exec_step_input P C M pc ics = StepT Cs a \<Longrightarrow> ics = Throwing Cs a"
- by simp
- show ?thesis using assms
- proof(induct rule: exec_step_ind.induct)
- case (exec_step_ind_NewOOM_Done sh C obj h ics P stk loc C\<^sub>0 M\<^sub>0 pc frs)
- then show ?case by(cases ics, auto)
- next
- case (exec_step_ind_New_Done sh C obj h a ics P stk loc C\<^sub>0 M\<^sub>0 pc frs)
- then show ?case by(cases ics, auto)
- next
- case (exec_step_ind_New_Init sh C ics P h stk loc C\<^sub>0 M\<^sub>0 pc frs)
- then show ?case by(cases ics, auto split: init_state.splits)
- next
- case (exec_step_ind_Getfield_NoField v stk D fs h P F C loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
- then show ?case by(cases "the (h (the_Addr (hd stk)))", cases ics, auto dest!: StepI)
- next
- case (exec_step_ind_Getfield_Static v stk D fs h P F t C loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
- then show ?case
- by(cases "the (h (the_Addr (hd stk)))", cases "fst(snd(field P C F))",
- cases ics, auto simp: split_beta dest: has_field_sees[OF has_field_idemp] dest!: StepI)
- next
- case (exec_step_ind_Getfield v stk D fs h D' b t P C F loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
- then show ?case
- by(cases "the (h (the_Addr (hd stk)))",
- cases ics; fastforce simp: split_beta dest: has_field_sees[OF has_field_idemp] dest!: StepI)
- next
- case (exec_step_ind_Getstatic_NonStatic P C F t D h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
- then show ?case
- by(cases ics; fastforce simp: split_beta split: staticb.splits
- dest: has_field_sees[OF has_field_idemp] dest!: StepI)
- next
- case exec_step_ind_Getstatic_Called
- then show ?case by(fastforce simp: split_beta split: staticb.splits dest!: StepI
- dest: has_field_sees[OF has_field_idemp])
- next
- case (exec_step_ind_Getstatic_Done D' b t P D F C ics sh sfs v h stk loc C\<^sub>0 M\<^sub>0 pc frs)
- then show ?case by(cases ics, auto simp: split_beta split: staticb.splits
- dest: has_field_sees[OF has_field_idemp])
- next
- case (exec_step_ind_Getstatic_Init D' b t P D F C sh ics h stk loc C\<^sub>0 M\<^sub>0 pc frs)
- then show ?case
- by(cases ics, auto simp: split_beta split: init_state.splits staticb.splits
- dest: has_field_sees[OF has_field_idemp])
- next
- case (exec_step_ind_Putfield_NoField r stk a D fs h P F C loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
- then show ?case by(cases "the (h (the_Addr (hd(tl stk))))", cases ics, auto dest!: StepI)
- next
- case (exec_step_ind_Putfield_Static r stk a D fs h P F t C loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
- then show ?case
- by(cases "the (h (the_Addr (hd(tl stk))))", cases "fst(snd(field P C F))",
- cases ics, auto simp: split_beta dest: has_field_sees[OF has_field_idemp] dest!: StepI)
- next
- case (exec_step_ind_Putfield v stk r a D fs h D' b t P C F loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
- then show ?case
- by(cases "the (h (the_Addr (hd(tl stk))))",
- cases ics; fastforce simp: split_beta dest: has_field_sees[OF has_field_idemp] dest!: StepI)
- next
- case (exec_step_ind_Putstatic_NonStatic P C F t D h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
- then show ?case
- by(cases ics; fastforce simp: split_beta split: staticb.splits
- dest: has_field_sees[OF has_field_idemp] dest!: StepI)
- next
- case exec_step_ind_Putstatic_Called
- then show ?case by(fastforce simp: split_beta split: staticb.splits dest!: StepI
- dest: has_field_sees[OF has_field_idemp])
- next
- case (exec_step_ind_Putstatic_Done D' b t P D F C ics sh sfs h stk loc C\<^sub>0 M\<^sub>0 pc frs)
- then show ?case by(cases ics, auto simp: split_beta split: staticb.splits
- dest: has_field_sees[OF has_field_idemp])
- next
- case (exec_step_ind_Putstatic_Init D' b t P D F C sh ics h stk loc C\<^sub>0 M\<^sub>0 pc frs)
- then show ?case
- by(cases ics, auto simp: split_beta split: staticb.splits init_state.splits
- dest: has_field_sees[OF has_field_idemp])
- next
- case (exec_step_ind_Invoke ps n stk r C h D b Ts T mxs mxl\<^sub>0 ins xt P M m f' loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
- then show ?case by(cases ics; fastforce dest!: StepI)
- next
- case (exec_step_ind_Invokestatic_Called ps n stk D b Ts T mxs mxl\<^sub>0 ins xt P C M m ics ics' sh)
- then show ?case by(cases ics; fastforce dest!: StepI)
- next
- case (exec_step_ind_Invokestatic_Done ps n stk D b Ts T mxs mxl\<^sub>0 ins xt P C M m ics sh sfs f')
- then show ?case by(cases ics; fastforce)
- next
- case (exec_step_ind_Invokestatic_Init D b Ts T mxs mxl\<^sub>0 ins xt P C M m sh ics n h stk loc C\<^sub>0 M\<^sub>0 pc frs)
- then show ?case by(cases ics; fastforce split: init_state.splits)
- next
- case (exec_step_ind_Return_NonStatic D Ts T m P C\<^sub>0 M\<^sub>0 h stk\<^sub>0 loc\<^sub>0 pc ics stk' loc' C' m' pc' ics' frs' sh)
- then show ?case by(cases "method P C\<^sub>0 M\<^sub>0", cases ics, auto dest!: StepI)
- next
- case (exec_step_ind_Return_Static D Ts T m P C\<^sub>0 M\<^sub>0 h stk\<^sub>0 loc\<^sub>0 pc ics stk' loc' C' m' pc' ics' frs' sh)
- then show ?case by(cases "method P C\<^sub>0 M\<^sub>0", cases ics, auto dest!: StepI)
- next
- case (exec_step_ind_IfFalse_nFalse stk i P h loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
- then show ?case by(cases "hd stk"; cases ics, auto dest!: StepI)
- next
- case (exec_step_ind_Throw_Null stk P h loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
- then show ?case by(cases "hd stk"; cases ics, auto dest!: StepI)
- next
- case (exec_step_ind_Init C Cs P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
- then have "ics = Called (C#Cs)" by simp
- then show ?case by auto
- (***)
- next
- case (exec_step_ind_Load n P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
- then show ?case by(cases ics, auto dest!: StepI)
- next
- case (exec_step_ind_Store n P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
- then show ?case by(cases ics, auto dest!: StepI)
- next
- case (exec_step_ind_Push v P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
- then show ?case by(cases ics, auto dest!: StepI)
- next
- case (exec_step_ind_NewOOM_Called h C P stk loc C\<^sub>0 M\<^sub>0 pc frs sh ics')
- then show ?case by(auto dest!: StepI)
- next
- case (exec_step_ind_New_Called h a C P stk loc C\<^sub>0 M\<^sub>0 pc frs sh ics')
- then show ?case by(auto dest!: StepI)
- next
- case (exec_step_ind_Getfield_Null stk F C P h loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
- then show ?case by(cases ics, auto dest!: StepI)
- next
- case (exec_step_ind_Getstatic_NoField P C F D h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
- then show ?case by(cases ics, auto dest!: StepI)
- next
- case (exec_step_ind_Putfield_Null stk F C P h loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
- then show ?case by(cases ics, auto dest!: StepI)
- next
- case (exec_step_ind_Putstatic_NoField P C F D h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
- then show ?case by(cases ics, auto dest!: StepI)
- next
- case (exec_step_ind_Checkcast P C h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
- then show ?case by(cases ics, auto dest!: StepI)
- next
- case (exec_step_ind_Checkcast_Error P C h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
- then show ?case by(cases ics, auto dest!: StepI)
- next
- case (exec_step_ind_Invoke_Null stk n M P h loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
- then show ?case by(cases ics, auto dest!: StepI)
- next
- case (exec_step_ind_Invoke_NoMethod r stk n C h P M loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
- then show ?case by(cases ics, auto dest!: StepI)
- next
- case (exec_step_ind_Invoke_Static r stk n C h D b Ts T mxs mxl\<^sub>0 ins xt P M m loc C\<^sub>0 M\<^sub>0 pc ics)
- then show ?case by(cases ics, auto dest!: StepI)
- next
- case (exec_step_ind_Invokestatic_NoMethod D b Ts T mxs mxl\<^sub>0 ins xt P C M n h stk loc C\<^sub>0 M\<^sub>0 pc ics)
- then show ?case by(cases ics, auto dest!: StepI)
- next
- case (exec_step_ind_Invokestatic_NonStatic D b Ts T mxs mxl\<^sub>0 ins xt P C M m n h stk loc C\<^sub>0 M\<^sub>0 pc ics)
- then show ?case by(cases ics, auto dest!: StepI)
- next
- case (exec_step_ind_Return_Last_Init P h stk\<^sub>0 loc\<^sub>0 C\<^sub>0 pc ics sh)
- then show ?case by(cases ics, auto dest!: StepI)
- next
- case (exec_step_ind_Return_Last M\<^sub>0 P h stk\<^sub>0 loc\<^sub>0 C\<^sub>0 pc ics sh)
- then show ?case by(cases ics, auto dest!: StepI)
- next
- case (exec_step_ind_Return_Init D b Ts T m P C\<^sub>0 h stk\<^sub>0 loc\<^sub>0 pc ics stk' loc' C' m' pc' ics')
- then show ?case by(cases ics, auto dest!: StepI)
- next
- case (exec_step_ind_Pop P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
- then show ?case by(cases ics, auto dest!: StepI)
- next
- case (exec_step_ind_IAdd P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
- then show ?case by(cases ics, auto dest!: StepI)
- next
- case (exec_step_ind_IfFalse_False stk i P h loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
- then show ?case by(cases ics, auto dest!: StepI)
- next
- case (exec_step_ind_CmpEq P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
- then show ?case by(cases ics, auto dest!: StepI)
- next
- case (exec_step_ind_Goto i P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
- then show ?case by(cases ics, auto dest!: StepI)
- next
- case (exec_step_ind_Throw stk P h loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
- then show ?case by(cases ics, auto dest!: StepI)
- next
- case (exec_step_ind_Init_None_Called sh C Cs P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs)
- then show ?case by(auto dest!: StepC)
- next
- case (exec_step_ind_Init_Done sh C sfs Cs P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs)
- then show ?case by(auto dest!: StepC)
- next
- case (exec_step_ind_Init_Processing sh C sfs Cs P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs)
- then show ?case by(auto dest!: StepC)
- next
- case (exec_step_ind_Init_Error sh C sfs Cs P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs)
- then show ?case by(auto dest!: StepC)
- next
- case (exec_step_ind_Init_Prepared_Object sh C sfs sh' Cs P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs)
- then show ?case by(auto dest!: StepC)
- next
- case (exec_step_ind_Init_Prepared_nObject sh C sfs sh' D P Cs h stk loc C\<^sub>0 M\<^sub>0 pc ics frs)
- then show ?case by(auto dest!: StepC)
- next
- case (exec_step_ind_InitThrow C Cs a P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
- then show ?case by(auto dest!: StepT)
- next
- case (exec_step_ind_InitThrow_End a P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
- then show ?case by(auto dest!: StepT)
- qed
-qed
-
-\<comment> \<open> @{term exec_step} and @{term exec_step_ind} reach the same result given equivalent input \<close>
-lemma exec_step_ind_equiv:
- "exec_step P h stk loc C M pc ics frs sh = (xp', h', frs', sh')
- = exec_step_ind (exec_step_input P C M pc ics) P h stk loc C M pc ics frs sh (xp', h', frs', sh')"
- using exec_step_imp_exec_step_ind exec_step_ind_imp_exec_step by auto
-
-end
+(* Title: RTS/JinjaSuppl/JVMExecStepInductive.thy
+ Author: Susannah Mansky
+ 2020, UIUC
+
+ Program Execution in the JVM as an inductive
+*)
+
+section "Inductive JVM execution"
+
+theory JVMExecStepInductive
+imports JinjaDCI.JVMExec
+begin
+
+datatype step_input = StepI instr |
+ StepC cname "cname list" | StepC2 cname "cname list" |
+ StepT "cname list" addr
+
+
+inductive exec_step_ind :: "[step_input, jvm_prog, heap, val list, val list,
+ cname, mname, pc, init_call_status, frame list, sheap,jvm_state] \<Rightarrow> bool"
+where
+ exec_step_ind_Load:
+"exec_step_ind (StepI (Load n)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
+ (None, h, ((loc ! n) # stk, loc, C\<^sub>0, M\<^sub>0, Suc pc, ics)#frs, sh)"
+
+| exec_step_ind_Store:
+"exec_step_ind (StepI (Store n)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
+ (None, h, (tl stk, loc[n:=hd stk], C\<^sub>0, M\<^sub>0, Suc pc, ics)#frs, sh)"
+
+| exec_step_ind_Push:
+"exec_step_ind (StepI (Push v)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
+ (None, h, (v # stk, loc, C\<^sub>0, M\<^sub>0, Suc pc, ics)#frs, sh)"
+
+| exec_step_ind_NewOOM_Called:
+"new_Addr h = None
+ \<Longrightarrow> exec_step_ind (StepI (New C)) P h stk loc C\<^sub>0 M\<^sub>0 pc (Called Cs) frs sh
+ (\<lfloor>addr_of_sys_xcpt OutOfMemory\<rfloor>, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, No_ics)#frs, sh)"
+
+| exec_step_ind_NewOOM_Done:
+"\<lbrakk> sh C = Some(obj, Done); new_Addr h = None; \<forall>Cs. ics \<noteq> Called Cs \<rbrakk>
+ \<Longrightarrow> exec_step_ind (StepI (New C)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
+ (\<lfloor>addr_of_sys_xcpt OutOfMemory\<rfloor>, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, ics)#frs, sh)"
+
+| exec_step_ind_New_Called:
+"new_Addr h = Some a
+ \<Longrightarrow> exec_step_ind (StepI (New C)) P h stk loc C\<^sub>0 M\<^sub>0 pc (Called Cs) frs sh
+ (None, h(a\<mapsto>blank P C), (Addr a#stk, loc, C\<^sub>0, M\<^sub>0, Suc pc, No_ics)#frs, sh)"
+
+| exec_step_ind_New_Done:
+"\<lbrakk> sh C = Some(obj, Done); new_Addr h = Some a; \<forall>Cs. ics \<noteq> Called Cs \<rbrakk>
+ \<Longrightarrow> exec_step_ind (StepI (New C)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
+ (None, h(a\<mapsto>blank P C), (Addr a#stk, loc, C\<^sub>0, M\<^sub>0, Suc pc, ics)#frs, sh)"
+
+| exec_step_ind_New_Init:
+"\<lbrakk> \<forall>obj. sh C \<noteq> Some(obj, Done); \<forall>Cs. ics \<noteq> Called Cs \<rbrakk>
+ \<Longrightarrow> exec_step_ind (StepI (New C)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
+ (None, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, Calling C [])#frs, sh)"
+
+| exec_step_ind_Getfield_Null:
+"hd stk = Null
+ \<Longrightarrow> exec_step_ind (StepI (Getfield F C)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
+ (\<lfloor>addr_of_sys_xcpt NullPointer\<rfloor>, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, ics)#frs, sh)"
+
+| exec_step_ind_Getfield_NoField:
+"\<lbrakk> v = hd stk; (D,fs) = the(h(the_Addr v)); v \<noteq> Null; \<not>(\<exists>t b. P \<turnstile> D has F,b:t in C) \<rbrakk>
+ \<Longrightarrow> exec_step_ind (StepI (Getfield F C)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
+ (\<lfloor>addr_of_sys_xcpt NoSuchFieldError\<rfloor>, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, ics)#frs, sh)"
+
+| exec_step_ind_Getfield_Static:
+"\<lbrakk> v = hd stk; (D,fs) = the(h(the_Addr v)); v \<noteq> Null; P \<turnstile> D has F,Static:t in C \<rbrakk>
+ \<Longrightarrow> exec_step_ind (StepI (Getfield F C)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
+ (\<lfloor>addr_of_sys_xcpt IncompatibleClassChangeError\<rfloor>, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, ics)#frs, sh)"
+
+| exec_step_ind_Getfield:
+"\<lbrakk> v = hd stk; (D,fs) = the(h(the_Addr v)); (D',b,t) = field P C F; v \<noteq> Null;
+ P \<turnstile> D has F,NonStatic:t in C \<rbrakk>
+ \<Longrightarrow> exec_step_ind (StepI (Getfield F C)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
+ (None, h, (the(fs(F,C))#(tl stk), loc, C\<^sub>0, M\<^sub>0, pc+1, ics)#frs, sh)"
+
+| exec_step_ind_Getstatic_NoField:
+"\<not>(\<exists>t b. P \<turnstile> C has F,b:t in D)
+ \<Longrightarrow> exec_step_ind (StepI (Getstatic C F D)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
+ (\<lfloor>addr_of_sys_xcpt NoSuchFieldError\<rfloor>, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, ics)#frs, sh)"
+
+| exec_step_ind_Getstatic_NonStatic:
+"P \<turnstile> C has F,NonStatic:t in D
+ \<Longrightarrow> exec_step_ind (StepI (Getstatic C F D)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
+ (\<lfloor>addr_of_sys_xcpt IncompatibleClassChangeError\<rfloor>, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, ics)#frs, sh)"
+
+| exec_step_ind_Getstatic_Called:
+"\<lbrakk> (D',b,t) = field P D F; P \<turnstile> C has F,Static:t in D;
+ v = the ((fst(the(sh D'))) F) \<rbrakk>
+ \<Longrightarrow> exec_step_ind (StepI (Getstatic C F D)) P h stk loc C\<^sub>0 M\<^sub>0 pc (Called Cs) frs sh
+ (None, h, (v#stk, loc, C\<^sub>0, M\<^sub>0, Suc pc, No_ics)#frs, sh)"
+
+| exec_step_ind_Getstatic_Done:
+"\<lbrakk> (D',b,t) = field P D F; P \<turnstile> C has F,Static:t in D;
+ \<forall>Cs. ics \<noteq> Called Cs; sh D' = Some(sfs,Done);
+ v = the (sfs F) \<rbrakk>
+ \<Longrightarrow> exec_step_ind (StepI (Getstatic C F D)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
+ (None, h, (v#stk, loc, C\<^sub>0, M\<^sub>0, Suc pc, ics)#frs, sh)"
+
+| exec_step_ind_Getstatic_Init:
+"\<lbrakk> (D',b,t) = field P D F; P \<turnstile> C has F,Static:t in D;
+ \<forall>sfs. sh D' \<noteq> Some(sfs,Done); \<forall>Cs. ics \<noteq> Called Cs \<rbrakk>
+ \<Longrightarrow> exec_step_ind (StepI (Getstatic C F D)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
+ (None, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, Calling D' [])#frs, sh)"
+
+| exec_step_ind_Putfield_Null:
+"hd(tl stk) = Null
+ \<Longrightarrow> exec_step_ind (StepI (Putfield F C)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
+ (\<lfloor>addr_of_sys_xcpt NullPointer\<rfloor>, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, ics)#frs, sh)"
+
+| exec_step_ind_Putfield_NoField:
+"\<lbrakk> r = hd(tl stk); a = the_Addr r; (D,fs) = the (h a); r \<noteq> Null; \<not>(\<exists>t b. P \<turnstile> D has F,b:t in C) \<rbrakk>
+ \<Longrightarrow> exec_step_ind (StepI (Putfield F C)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
+ (\<lfloor>addr_of_sys_xcpt NoSuchFieldError\<rfloor>, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, ics)#frs, sh)"
+
+| exec_step_ind_Putfield_Static:
+"\<lbrakk> r = hd(tl stk); a = the_Addr r; (D,fs) = the (h a); r \<noteq> Null; P \<turnstile> D has F,Static:t in C \<rbrakk>
+ \<Longrightarrow> exec_step_ind (StepI (Putfield F C)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
+ (\<lfloor>addr_of_sys_xcpt IncompatibleClassChangeError\<rfloor>, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, ics)#frs, sh)"
+
+| exec_step_ind_Putfield:
+"\<lbrakk> v = hd stk; r = hd(tl stk); a = the_Addr r; (D,fs) = the (h a); (D',b,t) = field P C F;
+ r \<noteq> Null; P \<turnstile> D has F,NonStatic:t in C \<rbrakk>
+ \<Longrightarrow> exec_step_ind (StepI (Putfield F C)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
+ (None, h(a \<mapsto> (D, fs((F,C) \<mapsto> v))), (tl (tl stk), loc, C\<^sub>0, M\<^sub>0, pc+1, ics)#frs, sh)"
+
+| exec_step_ind_Putstatic_NoField:
+"\<not>(\<exists>t b. P \<turnstile> C has F,b:t in D)
+ \<Longrightarrow> exec_step_ind (StepI (Putstatic C F D)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
+ (\<lfloor>addr_of_sys_xcpt NoSuchFieldError\<rfloor>, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, ics)#frs, sh)"
+
+| exec_step_ind_Putstatic_NonStatic:
+"P \<turnstile> C has F,NonStatic:t in D
+ \<Longrightarrow> exec_step_ind (StepI (Putstatic C F D)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
+ (\<lfloor>addr_of_sys_xcpt IncompatibleClassChangeError\<rfloor>, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, ics)#frs, sh)"
+
+| exec_step_ind_Putstatic_Called:
+"\<lbrakk> (D',b,t) = field P D F; P \<turnstile> C has F,Static:t in D; the(sh D') = (sfs,i) \<rbrakk>
+ \<Longrightarrow> exec_step_ind (StepI (Putstatic C F D)) P h stk loc C\<^sub>0 M\<^sub>0 pc (Called Cs) frs sh
+ (None, h, (tl stk, loc, C\<^sub>0, M\<^sub>0, Suc pc, No_ics)#frs, sh(D':=Some ((sfs(F \<mapsto> hd stk)), i)))"
+
+| exec_step_ind_Putstatic_Done:
+"\<lbrakk> (D',b,t) = field P D F; P \<turnstile> C has F,Static:t in D;
+ \<forall>Cs. ics \<noteq> Called Cs; sh D' = Some (sfs, Done) \<rbrakk>
+ \<Longrightarrow> exec_step_ind (StepI (Putstatic C F D)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
+ (None, h, (tl stk, loc, C\<^sub>0, M\<^sub>0, Suc pc, ics)#frs, sh(D':=Some ((sfs(F \<mapsto> hd stk)), Done)))"
+
+| exec_step_ind_Putstatic_Init:
+"\<lbrakk> (D',b,t) = field P D F; P \<turnstile> C has F,Static:t in D;
+ \<forall>sfs. sh D' \<noteq> Some (sfs, Done); \<forall>Cs. ics \<noteq> Called Cs \<rbrakk>
+ \<Longrightarrow> exec_step_ind (StepI (Putstatic C F D)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
+ (None, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, Calling D' [])#frs, sh)"
+
+| exec_step_ind_Checkcast:
+"cast_ok P C h (hd stk)
+ \<Longrightarrow> exec_step_ind (StepI (Checkcast C)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
+ (None, h, (stk, loc, C\<^sub>0, M\<^sub>0, Suc pc, ics)#frs, sh)"
+
+| exec_step_ind_Checkcast_Error:
+"\<not>cast_ok P C h (hd stk)
+ \<Longrightarrow> exec_step_ind (StepI (Checkcast C)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
+ (\<lfloor>addr_of_sys_xcpt ClassCast\<rfloor>, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, ics)#frs, sh)"
+
+| exec_step_ind_Invoke_Null:
+"stk!n = Null
+ \<Longrightarrow> exec_step_ind (StepI (Invoke M n)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
+ (\<lfloor>addr_of_sys_xcpt NullPointer\<rfloor>, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, ics)#frs, sh)"
+
+| exec_step_ind_Invoke_NoMethod:
+"\<lbrakk> r = stk!n; C = fst(the(h(the_Addr r))); r \<noteq> Null;
+ \<not>(\<exists>Ts T m D b. P \<turnstile> C sees M,b:Ts \<rightarrow> T = m in D) \<rbrakk>
+ \<Longrightarrow> exec_step_ind (StepI (Invoke M n)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
+ (\<lfloor>addr_of_sys_xcpt NoSuchMethodError\<rfloor>, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, ics)#frs, sh)"
+
+| exec_step_ind_Invoke_Static:
+"\<lbrakk> r = stk!n; C = fst(the(h(the_Addr r)));
+ (D,b,Ts,T,mxs,mxl\<^sub>0,ins,xt)= method P C M; r \<noteq> Null;
+ P \<turnstile> C sees M,Static:Ts \<rightarrow> T = m in D \<rbrakk>
+ \<Longrightarrow> exec_step_ind (StepI (Invoke M n)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
+ (\<lfloor>addr_of_sys_xcpt IncompatibleClassChangeError\<rfloor>, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, ics)#frs, sh)"
+
+| exec_step_ind_Invoke:
+"\<lbrakk> ps = take n stk; r = stk!n; C = fst(the(h(the_Addr r)));
+ (D,b,Ts,T,mxs,mxl\<^sub>0,ins,xt)= method P C M; r \<noteq> Null;
+ P \<turnstile> C sees M,NonStatic:Ts \<rightarrow> T = m in D;
+ f' = ([],[r]@(rev ps)@(replicate mxl\<^sub>0 undefined),D,M,0,No_ics) \<rbrakk>
+ \<Longrightarrow> exec_step_ind (StepI (Invoke M n)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
+ (None, h, f'#(stk, loc, C\<^sub>0, M\<^sub>0, pc, ics)#frs, sh)"
+
+| exec_step_ind_Invokestatic_NoMethod:
+"\<lbrakk> (D,b,Ts,T,mxs,mxl\<^sub>0,ins,xt)= method P C M; \<not>(\<exists>Ts T m D b. P \<turnstile> C sees M,b:Ts \<rightarrow> T = m in D) \<rbrakk>
+ \<Longrightarrow> exec_step_ind (StepI (Invokestatic C M n)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
+ (\<lfloor>addr_of_sys_xcpt NoSuchMethodError\<rfloor>, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, ics)#frs, sh)"
+
+| exec_step_ind_Invokestatic_NonStatic:
+"\<lbrakk> (D,b,Ts,T,mxs,mxl\<^sub>0,ins,xt)= method P C M; P \<turnstile> C sees M,NonStatic:Ts \<rightarrow> T = m in D \<rbrakk>
+ \<Longrightarrow> exec_step_ind (StepI (Invokestatic C M n)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
+ (\<lfloor>addr_of_sys_xcpt IncompatibleClassChangeError\<rfloor>, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, ics)#frs, sh)"
+
+| exec_step_ind_Invokestatic_Called:
+"\<lbrakk> ps = take n stk; (D,b,Ts,T,mxs,mxl\<^sub>0,ins,xt) = method P C M;
+ P \<turnstile> C sees M,Static:Ts \<rightarrow> T = m in D;
+ f' = ([],(rev ps)@(replicate mxl\<^sub>0 undefined),D,M,0,No_ics) \<rbrakk>
+ \<Longrightarrow> exec_step_ind (StepI (Invokestatic C M n)) P h stk loc C\<^sub>0 M\<^sub>0 pc (Called Cs) frs sh
+ (None, h, f'#(stk, loc, C\<^sub>0, M\<^sub>0, pc, No_ics)#frs, sh)"
+
+| exec_step_ind_Invokestatic_Done:
+"\<lbrakk> ps = take n stk; (D,b,Ts,T,mxs,mxl\<^sub>0,ins,xt) = method P C M;
+ P \<turnstile> C sees M,Static:Ts \<rightarrow> T = m in D;
+ \<forall>Cs. ics \<noteq> Called Cs; sh D = Some (sfs, Done);
+ f' = ([],(rev ps)@(replicate mxl\<^sub>0 undefined),D,M,0,No_ics) \<rbrakk>
+ \<Longrightarrow> exec_step_ind (StepI (Invokestatic C M n)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
+ (None, h, f'#(stk, loc, C\<^sub>0, M\<^sub>0, pc, ics)#frs, sh)"
+
+| exec_step_ind_Invokestatic_Init:
+"\<lbrakk> (D,b,Ts,T,mxs,mxl\<^sub>0,ins,xt) = method P C M;
+ P \<turnstile> C sees M,Static:Ts \<rightarrow> T = m in D;
+ \<forall>sfs. sh D \<noteq> Some (sfs, Done); \<forall>Cs. ics \<noteq> Called Cs \<rbrakk>
+ \<Longrightarrow> exec_step_ind (StepI (Invokestatic C M n)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
+ (None, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, Calling D [])#frs, sh)"
+
+| exec_step_ind_Return_Last_Init:
+ "exec_step_ind (StepI Return) P h stk\<^sub>0 loc\<^sub>0 C\<^sub>0 clinit pc ics [] sh
+ (None, h, [], sh(C\<^sub>0:=Some(fst(the(sh C\<^sub>0)), Done)))"
+
+| exec_step_ind_Return_Last:
+ "M\<^sub>0 \<noteq> clinit
+ \<Longrightarrow> exec_step_ind (StepI Return) P h stk\<^sub>0 loc\<^sub>0 C\<^sub>0 M\<^sub>0 pc ics [] sh (None, h, [], sh)"
+
+| exec_step_ind_Return_Init:
+ "\<lbrakk> (D,b,Ts,T,m) = method P C\<^sub>0 clinit \<rbrakk>
+ \<Longrightarrow> exec_step_ind (StepI Return) P h stk\<^sub>0 loc\<^sub>0 C\<^sub>0 clinit pc ics ((stk',loc',C',m',pc',ics')#frs') sh
+ (None, h, (stk',loc',C',m',pc',ics')#frs', sh(C\<^sub>0:=Some(fst(the(sh C\<^sub>0)), Done)))"
+
+| exec_step_ind_Return_NonStatic:
+ "\<lbrakk> (D,NonStatic,Ts,T,m) = method P C\<^sub>0 M\<^sub>0; M\<^sub>0 \<noteq> clinit \<rbrakk>
+ \<Longrightarrow> exec_step_ind (StepI Return) P h stk\<^sub>0 loc\<^sub>0 C\<^sub>0 M\<^sub>0 pc ics ((stk',loc',C',m',pc',ics')#frs') sh
+ (None, h, ((hd stk\<^sub>0)#(drop (length Ts + 1) stk'),loc',C',m',Suc pc',ics')#frs', sh)"
+
+| exec_step_ind_Return_Static:
+ "\<lbrakk> (D,Static,Ts,T,m) = method P C\<^sub>0 M\<^sub>0; M\<^sub>0 \<noteq> clinit \<rbrakk>
+ \<Longrightarrow> exec_step_ind (StepI Return) P h stk\<^sub>0 loc\<^sub>0 C\<^sub>0 M\<^sub>0 pc ics ((stk',loc',C',m',pc',ics')#frs') sh
+ (None, h, ((hd stk\<^sub>0)#(drop (length Ts) stk'),loc',C',m',Suc pc',ics')#frs', sh)"
+
+| exec_step_ind_Pop:
+"exec_step_ind (StepI Pop) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
+ (None, h, (tl stk, loc, C\<^sub>0, M\<^sub>0, Suc pc, ics)#frs, sh)"
+
+| exec_step_ind_IAdd:
+"exec_step_ind (StepI IAdd) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
+ (None, h, (Intg (the_Intg (hd (tl stk)) + the_Intg (hd stk))#(tl (tl stk)), loc, C\<^sub>0, M\<^sub>0, Suc pc, ics)#frs, sh)"
+
+| exec_step_ind_IfFalse_False:
+"hd stk = Bool False
+ \<Longrightarrow> exec_step_ind (StepI (IfFalse i)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
+ (None, h, (tl stk, loc, C\<^sub>0, M\<^sub>0, nat(int pc+i), ics)#frs, sh)"
+
+| exec_step_ind_IfFalse_nFalse:
+"hd stk \<noteq> Bool False
+ \<Longrightarrow> exec_step_ind (StepI (IfFalse i)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
+ (None, h, (tl stk, loc, C\<^sub>0, M\<^sub>0, Suc pc, ics)#frs, sh)"
+
+| exec_step_ind_CmpEq:
+"exec_step_ind (StepI CmpEq) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
+ (None, h, (Bool (hd (tl stk) = hd stk) # tl (tl stk), loc, C\<^sub>0, M\<^sub>0, Suc pc, ics)#frs, sh)"
+
+| exec_step_ind_Goto:
+"exec_step_ind (StepI (Goto i)) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
+ (None, h, (stk, loc, C\<^sub>0, M\<^sub>0, nat(int pc+i), ics)#frs, sh)"
+
+| exec_step_ind_Throw:
+"hd stk \<noteq> Null
+ \<Longrightarrow> exec_step_ind (StepI Throw) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
+ (\<lfloor>the_Addr (hd stk)\<rfloor>, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, ics)#frs, sh)"
+
+| exec_step_ind_Throw_Null:
+"hd stk = Null
+ \<Longrightarrow> exec_step_ind (StepI Throw) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
+ (\<lfloor>addr_of_sys_xcpt NullPointer\<rfloor>, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, ics)#frs, sh)"
+
+| exec_step_ind_Init_None_Called:
+"\<lbrakk> sh C = None \<rbrakk>
+ \<Longrightarrow> exec_step_ind (StepC C Cs) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
+ (None, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, Calling C Cs)#frs, sh(C := Some (sblank P C, Prepared)))"
+
+| exec_step_ind_Init_Done:
+"sh C = Some (sfs, Done)
+ \<Longrightarrow> exec_step_ind (StepC C Cs) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
+ (None, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, Called Cs)#frs, sh)"
+
+| exec_step_ind_Init_Processing:
+"sh C = Some (sfs, Processing)
+ \<Longrightarrow> exec_step_ind (StepC C Cs) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
+ (None, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, Called Cs)#frs, sh)"
+
+| exec_step_ind_Init_Error:
+"\<lbrakk> sh C = Some (sfs, Error) \<rbrakk>
+ \<Longrightarrow> exec_step_ind (StepC C Cs) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
+ (None, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, Throwing Cs (addr_of_sys_xcpt NoClassDefFoundError))#frs, sh)"
+
+| exec_step_ind_Init_Prepared_Object:
+"\<lbrakk> sh C = Some (sfs, Prepared);
+ sh' = sh(C:=Some(fst(the(sh C)), Processing));
+ C = Object \<rbrakk>
+ \<Longrightarrow> exec_step_ind (StepC C Cs) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
+ (None, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, Called (C#Cs))#frs, sh')"
+
+| exec_step_ind_Init_Prepared_nObject:
+"\<lbrakk> sh C = Some (sfs, Prepared);
+ sh' = sh(C:=Some(fst(the(sh C)), Processing));
+ C \<noteq> Object; D = fst(the(class P C)) \<rbrakk>
+ \<Longrightarrow> exec_step_ind (StepC C Cs) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
+ (None, h, (stk, loc, C\<^sub>0, M\<^sub>0, pc, Calling D (C#Cs))#frs, sh')"
+
+| exec_step_ind_Init:
+"exec_step_ind (StepC2 C Cs) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
+ (None, h, create_init_frame P C#(stk, loc, C\<^sub>0, M\<^sub>0, pc, Called Cs)#frs, sh)"
+
+| exec_step_ind_InitThrow:
+"exec_step_ind (StepT (C#Cs) a) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
+ (None, h, (stk,loc,C\<^sub>0,M\<^sub>0,pc,Throwing Cs a)#frs, (sh(C \<mapsto> (fst(the(sh C)), Error))))"
+
+| exec_step_ind_InitThrow_End:
+"exec_step_ind (StepT [] a) P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh
+ (\<lfloor>a\<rfloor>, h, (stk,loc,C\<^sub>0,M\<^sub>0,pc,No_ics)#frs, sh)"
+
+(** ******* **)
+
+inductive_cases exec_step_ind_cases [cases set]:
+ "exec_step_ind (StepI (Load n)) P h stk loc C M pc ics frs sh \<sigma>"
+ "exec_step_ind (StepI (Store n)) P h stk loc C M pc ics frs sh \<sigma>"
+ "exec_step_ind (StepI (Push v)) P h stk loc C M pc ics frs sh \<sigma>"
+ "exec_step_ind (StepI (New C)) P h stk loc C M pc ics frs sh \<sigma>"
+ "exec_step_ind (StepI (Getfield F C)) P h stk loc C M pc ics frs sh \<sigma>"
+ "exec_step_ind (StepI (Getstatic C F D)) P h stk loc C M pc ics frs sh \<sigma>"
+ "exec_step_ind (StepI (Putfield F C)) P h stk loc C M pc ics frs sh \<sigma>"
+ "exec_step_ind (StepI (Putstatic C F D)) P h stk loc C M pc ics frs sh \<sigma>"
+ "exec_step_ind (StepI (Checkcast C)) P h stk loc C M pc ics frs sh \<sigma>"
+ "exec_step_ind (StepI (Invoke M n)) P h stk loc C M pc ics frs sh \<sigma>"
+ "exec_step_ind (StepI (Invokestatic C M n)) P h stk loc C M pc ics frs sh \<sigma>"
+ "exec_step_ind (StepI Return) P h stk loc C M pc ics frs sh \<sigma>"
+ "exec_step_ind (StepI Pop) P h stk loc C M pc ics frs sh \<sigma>"
+ "exec_step_ind (StepI IAdd) P h stk loc C M pc ics frs sh \<sigma>"
+ "exec_step_ind (StepI (IfFalse i)) P h stk loc C M pc ics frs sh \<sigma>"
+ "exec_step_ind (StepI CmpEq) P h stk loc C M pc ics frs sh \<sigma>"
+ "exec_step_ind (StepI (Goto i)) P h stk loc C M pc ics frs sh \<sigma>"
+ "exec_step_ind (StepI Throw) P h stk loc C M pc ics frs sh \<sigma>"
+ "exec_step_ind (StepC C' Cs) P h stk loc C M pc ics frs sh \<sigma>"
+ "exec_step_ind (StepC2 C' Cs) P h stk loc C M pc ics frs sh \<sigma>"
+ "exec_step_ind (StepT Cs a) P h stk loc C M pc ics frs sh \<sigma>"
+
+
+\<comment> \<open> Deriving @{term step_input} for @{term exec_step_ind} from @{term exec_step} arguments \<close>
+fun exec_step_input :: "[jvm_prog, cname, mname, pc, init_call_status] \<Rightarrow> step_input" where
+"exec_step_input P C M pc (Calling C' Cs) = StepC C' Cs" |
+"exec_step_input P C M pc (Called (C'#Cs)) = StepC2 C' Cs" |
+"exec_step_input P C M pc (Throwing Cs a) = StepT Cs a" |
+"exec_step_input P C M pc ics = StepI (instrs_of P C M ! pc)"
+
+lemma exec_step_input_StepTD[simp]:
+assumes "exec_step_input P C M pc ics = StepT Cs a" shows "ics = Throwing Cs a"
+using assms proof(cases ics)
+ case (Called Cs) with assms show ?thesis by(cases Cs; simp)
+qed(auto)
+
+lemma exec_step_input_StepCD[simp]:
+assumes "exec_step_input P C M pc ics = StepC C' Cs" shows "ics = Calling C' Cs"
+using assms proof(cases ics)
+ case (Called Cs) with assms show ?thesis by(cases Cs; simp)
+qed(auto)
+
+lemma exec_step_input_StepC2D[simp]:
+assumes "exec_step_input P C M pc ics = StepC2 C' Cs" shows "ics = Called (C'#Cs)"
+using assms proof(cases ics)
+ case (Called Cs) with assms show ?thesis by(cases Cs; simp)
+qed(auto)
+
+lemma exec_step_input_StepID:
+assumes "exec_step_input P C M pc ics = StepI i"
+shows "(ics = Called [] \<or> ics = No_ics) \<and> instrs_of P C M ! pc = i"
+using assms proof(cases ics)
+ case (Called Cs) with assms show ?thesis by(cases Cs; simp)
+qed(auto)
+
+subsection "Equivalence of @{term exec_step} and @{term exec_step_input}"
+
+lemma exec_step_imp_exec_step_ind:
+assumes es: "exec_step P h stk loc C M pc ics frs sh = (xp', h', frs', sh')"
+shows "exec_step_ind (exec_step_input P C M pc ics) P h stk loc C M pc ics frs sh (xp', h', frs', sh')"
+proof(cases "exec_step_input P C M pc ics")
+ case (StepT Cs a)
+ then have "ics = Throwing Cs a" by simp
+ then show ?thesis using exec_step_ind_InitThrow exec_step_ind_InitThrow_End StepT es
+ by(cases Cs, auto)
+next
+ case (StepC C1 Cs)
+ then have ics: "ics = Calling C1 Cs" by simp
+ obtain D b Ts T m where lets: "method P C1 clinit = (D,b,Ts,T,m)" by(cases "method P C1 clinit")
+ then obtain mxs mxl\<^sub>0 ins xt where m: "m = (mxs,mxl\<^sub>0,ins,xt)" by(cases m)
+ show ?thesis
+ proof(cases "sh C1")
+ case None then show ?thesis
+ using exec_step_ind_Init_None_Called ics assms by auto
+ next
+ case (Some a)
+ then obtain sfs i where sfsi: "a = (sfs,i)" by(cases a)
+ then show ?thesis using exec_step_ind_Init_Done exec_step_ind_Init_Processing
+ exec_step_ind_Init_Error m lets Some ics assms
+ proof(cases i)
+ case Prepared
+ show ?thesis
+ using exec_step_ind_Init_Prepared_Object[where P=P] exec_step_ind_Init_Prepared_nObject
+ sfsi m lets Prepared Some ics assms by(auto split: if_split_asm)
+ qed(auto)
+ qed
+next
+ case (StepC2 C1 Cs)
+ then have ics: "ics = Called (C1#Cs)" by simp
+ then show ?thesis using exec_step_ind_Init assms by auto
+next
+ case (StepI i)
+ then have
+ ics: "ics = Called [] \<or> ics = No_ics" and
+ exec_instr: "exec_instr i P h stk loc C M pc ics frs sh = (xp', h', frs', sh')"
+ using assms by(auto dest!: exec_step_input_StepID)
+ show ?thesis
+ proof(cases i)
+ case (Load x1) then show ?thesis using exec_instr exec_step_ind_Load StepI by auto
+ next
+ case (Store x2) then show ?thesis using exec_instr exec_step_ind_Store StepI by auto
+ next
+ case (Push x3) then show ?thesis using exec_instr exec_step_ind_Push StepI by auto
+ next
+ case (New C1)
+ then obtain sfs i where sfsi: "the(sh C1) = (sfs,i)" by(cases "the(sh C1)")
+ then show ?thesis using exec_step_ind_New_Called exec_step_ind_NewOOM_Called
+ exec_step_ind_New_Done exec_step_ind_NewOOM_Done
+ exec_step_ind_New_Init sfsi New StepI exec_instr ics by(auto split: init_state.splits)
+ next
+ case (Getfield F1 C1)
+ then obtain D fs D' b t where lets: "the(h(the_Addr (hd stk))) = (D,fs)"
+ "field P C1 F1 = (D',b,t)" by(cases "the(h(the_Addr (hd stk)))", cases "field P C1 F1")
+ then have "\<And>b' t'. P \<turnstile> D has F1,b':t' in C1 \<Longrightarrow> (D', b, t) = (C1, b', t')"
+ using field_def2 has_field_idemp has_field_sees by fastforce
+ then show ?thesis using exec_step_ind_Getfield_Null exec_step_ind_Getfield_NoField
+ exec_step_ind_Getfield_Static exec_step_ind_Getfield lets Getfield StepI exec_instr
+ by(auto split: if_split_asm staticb.splits) metis+
+ next
+ case (Getstatic C1 F1 D1)
+ then obtain D' b t where lets: "field P D1 F1 = (D',b,t)" by(cases "field P D1 F1")
+ then have field: "\<And>b' t'. P \<turnstile> C1 has F1,b':t' in D1 \<Longrightarrow> (D', b, t) = (D1, b', t')"
+ using field_def2 has_field_idemp has_field_sees by fastforce
+ show ?thesis
+ proof(cases b)
+ case NonStatic then show ?thesis
+ using exec_step_ind_Getstatic_NoField exec_step_ind_Getstatic_NonStatic
+ field lets Getstatic exec_instr StepI by(auto split: if_split_asm) fastforce
+ next
+ case Static show ?thesis
+ proof(cases "ics = Called []")
+ case True then show ?thesis using exec_step_ind_Getstatic_NoField
+ exec_step_ind_Getstatic_Called exec_step_ind_Getstatic_Init
+ Static field lets Getstatic exec_instr StepI ics
+ by(auto simp: split_beta split: if_split_asm) metis
+ next
+ case False
+ then have nCalled: "\<forall>Cs. ics \<noteq> Called Cs" using ics by simp
+ show ?thesis
+ proof(cases "sh D1")
+ case None
+ then have nDone: "\<forall>sfs. sh D1 \<noteq> Some(sfs, Done)" by simp
+ then show ?thesis using exec_step_ind_Getstatic_NoField
+ exec_step_ind_Getstatic_Init[where sh=sh, OF _ _ nDone nCalled]
+ field lets None False Static Getstatic exec_instr StepI ics
+ by(auto split: if_split_asm) metis
+ next
+ case (Some a)
+ then obtain sfs i where sfsi: "a=(sfs,i)" by(cases a)
+ show ?thesis using exec_step_ind_Getstatic_NoField
+ exec_step_ind_Getstatic_Init sfsi False Static Some field lets Getstatic exec_instr
+ proof(cases "i = Done")
+ case True then show ?thesis using exec_step_ind_Getstatic_NoField
+ exec_step_ind_Getstatic_Done[OF _ _ nCalled] exec_step_ind_Getstatic_Init
+ sfsi False Static Some field lets Getstatic exec_instr StepI ics
+ by(auto split: if_split_asm) metis
+ next
+ case nD: False
+ then have nDone: "\<forall>sfs. sh D1 \<noteq> Some(sfs, Done)" using sfsi Some by simp
+ show ?thesis using nD
+ proof(cases i)
+ case Processing then show ?thesis using exec_step_ind_Getstatic_NoField
+ exec_step_ind_Getstatic_Init[where sh=sh, OF _ _ nDone nCalled]
+ sfsi False Static Some field lets Getstatic exec_instr StepI ics
+ by(auto split: if_split_asm) metis
+ next
+ case Prepared then show ?thesis using exec_step_ind_Getstatic_NoField
+ exec_step_ind_Getstatic_Init[where sh=sh, OF _ _ nDone nCalled]
+ sfsi False Static Some field lets Getstatic exec_instr StepI ics
+ by(auto split: if_split_asm) metis
+ next
+ case Error then show ?thesis using exec_step_ind_Getstatic_NoField
+ exec_step_ind_Getstatic_Init[where sh=sh, OF _ _ nDone nCalled]
+ sfsi False Static Some field lets Getstatic exec_instr StepI ics
+ by(auto split: if_split_asm) metis
+ qed(simp)
+ qed
+ qed
+ qed
+ qed
+ next
+ case (Putfield F1 C1)
+ then obtain D fs D' b t where lets: "the(h(the_Addr (hd(tl stk)))) = (D,fs)"
+ "field P C1 F1 = (D',b,t)" by(cases "the(h(the_Addr (hd(tl stk))))", cases "field P C1 F1")
+ then have "\<And>b' t'. P \<turnstile> D has F1,b':t' in C1 \<Longrightarrow> (D', b, t) = (C1, b', t')"
+ using field_def2 has_field_idemp has_field_sees by fastforce
+ then show ?thesis using exec_step_ind_Putfield_Null exec_step_ind_Putfield_NoField
+ exec_step_ind_Putfield_Static exec_step_ind_Putfield lets Putfield exec_instr StepI
+ by(auto split: if_split_asm staticb.splits) metis+
+ next
+ case (Putstatic C1 F1 D1)
+ then obtain D' b t where lets: "field P D1 F1 = (D',b,t)" by(cases "field P D1 F1")
+ then have field: "\<And>b' t'. P \<turnstile> C1 has F1,b':t' in D1 \<Longrightarrow> (D', b, t) = (D1, b', t')"
+ using field_def2 has_field_idemp has_field_sees by fastforce
+ show ?thesis
+ proof(cases b)
+ case NonStatic then show ?thesis
+ using exec_step_ind_Putstatic_NoField exec_step_ind_Putstatic_NonStatic
+ field lets Putstatic exec_instr StepI by(auto split: if_split_asm) fastforce
+ next
+ case Static show ?thesis
+ proof(cases "ics = Called []")
+ case True then show ?thesis using exec_step_ind_Putstatic_NoField
+ exec_step_ind_Putstatic_Called exec_step_ind_Putstatic_Init
+ Static field lets Putstatic exec_instr StepI ics
+ by(cases "the(sh D1)", auto split: if_split_asm) metis
+ next
+ case False
+ then have nCalled: "\<forall>Cs. ics \<noteq> Called Cs" using ics by simp
+ show ?thesis
+ proof(cases "sh D1")
+ case None
+ then have nDone: "\<forall>sfs. sh D1 \<noteq> Some(sfs, Done)" by simp
+ then show ?thesis using exec_step_ind_Putstatic_NoField
+ exec_step_ind_Putstatic_Init[where sh=sh, OF _ _ nDone nCalled]
+ field lets None False Static Putstatic exec_instr StepI ics
+ by(auto split: if_split_asm) metis
+ next
+ case (Some a)
+ then obtain sfs i where sfsi: "a=(sfs,i)" by(cases a)
+ show ?thesis using exec_step_ind_Putstatic_NoField
+ exec_step_ind_Putstatic_Init sfsi False Static Some field lets Putstatic exec_instr
+ proof(cases "i = Done")
+ case True then show ?thesis using exec_step_ind_Putstatic_NoField
+ exec_step_ind_Putstatic_Done[OF _ _ nCalled] exec_step_ind_Putstatic_Init
+ sfsi False Static Some field lets Putstatic exec_instr StepI ics
+ by(auto split: if_split_asm) metis
+ next
+ case nD: False
+ then have nDone: "\<forall>sfs. sh D1 \<noteq> Some(sfs, Done)" using sfsi Some by simp
+ show ?thesis using nD
+ proof(cases i)
+ case Processing then show ?thesis using exec_step_ind_Putstatic_NoField
+ exec_step_ind_Putstatic_Init[where sh=sh, OF _ _ nDone nCalled]
+ sfsi False Static Some field lets Putstatic exec_instr StepI ics
+ by(auto split: if_split_asm) metis
+ next
+ case Prepared then show ?thesis using exec_step_ind_Putstatic_NoField
+ exec_step_ind_Putstatic_Init[where sh=sh, OF _ _ nDone nCalled]
+ sfsi False Static Some field lets Putstatic exec_instr StepI ics
+ by(auto split: if_split_asm) metis
+ next
+ case Error then show ?thesis using exec_step_ind_Putstatic_NoField
+ exec_step_ind_Putstatic_Init[where sh=sh, OF _ _ nDone nCalled]
+ sfsi False Static Some field lets Putstatic exec_instr StepI ics
+ by(auto split: if_split_asm) metis
+ qed(simp)
+ qed
+ qed
+ qed
+ qed
+ next
+ case Checkcast then show ?thesis
+ using exec_step_ind_Checkcast exec_step_ind_Checkcast_Error exec_instr StepI
+ by(auto split: if_split_asm)
+ next
+ case (Invoke M1 n) show ?thesis
+ proof(cases "stk!n = Null")
+ case True then show ?thesis using exec_step_ind_Invoke_Null Invoke exec_instr StepI
+ by clarsimp
+ next
+ case False
+ let ?C = "cname_of h (the_Addr (stk ! n))"
+ obtain D b Ts T m where method: "method P ?C M1 = (D,b,Ts,T,m)" by(cases "method P ?C M1")
+ then obtain mxs mxl\<^sub>0 ins xt where "m = (mxs,mxl\<^sub>0,ins,xt)" by(cases m)
+ then show ?thesis using exec_step_ind_Invoke_NoMethod
+ exec_step_ind_Invoke_Static exec_step_ind_Invoke method False Invoke exec_instr StepI
+ by(auto split: if_split_asm staticb.splits)
+ qed
+ next
+ case (Invokestatic C1 M1 n)
+ obtain D b Ts T m where lets: "method P C1 M1 = (D,b,Ts,T,m)" by(cases "method P C1 M1")
+ then obtain mxs mxl\<^sub>0 ins xt where m: "m = (mxs,mxl\<^sub>0,ins,xt)" by(cases m)
+ have method: "\<And>b' Ts' t' m' D'. P \<turnstile> C1 sees M1,b':Ts' \<rightarrow> t' = m' in D'
+ \<Longrightarrow> (D,b,Ts,T,m) = (D',b',Ts',t',m')" using lets by auto
+ show ?thesis
+ proof(cases b)
+ case NonStatic then show ?thesis
+ using exec_step_ind_Invokestatic_NoMethod exec_step_ind_Invokestatic_NonStatic
+ m method lets Invokestatic exec_instr StepI by(auto split: if_split_asm)
+ next
+ case Static show ?thesis
+ proof(cases "ics = Called []")
+ case True then show ?thesis using exec_step_ind_Invokestatic_NoMethod
+ exec_step_ind_Invokestatic_Called exec_step_ind_Invokestatic_Init
+ Static m method lets Invokestatic exec_instr StepI ics
+ by(auto split: if_split_asm)
+ next
+ case False
+ then have nCalled: "\<forall>Cs. ics \<noteq> Called Cs" using ics by simp
+ show ?thesis
+ proof(cases "sh D")
+ case None
+ then have nDone: "\<forall>sfs. sh D \<noteq> Some(sfs, Done)" by simp
+ show ?thesis using exec_step_ind_Invokestatic_NoMethod
+ exec_step_ind_Invokestatic_Init[where sh=sh, OF _ _ nDone nCalled]
+ method m lets None False Static Invokestatic exec_instr StepI ics
+ by(auto split: if_split_asm)
+ next
+ case (Some a)
+ then obtain sfs i where sfsi: "a=(sfs,i)" by(cases a)
+ show ?thesis using exec_step_ind_Invokestatic_NoMethod
+ exec_step_ind_Invokestatic_Init sfsi False Static Some method lets Invokestatic exec_instr
+ proof(cases "i = Done")
+ case True then show ?thesis using exec_step_ind_Invokestatic_NoMethod
+ exec_step_ind_Invokestatic_Done[OF _ _ _ nCalled] exec_step_ind_Invokestatic_Init
+ sfsi False Static Some m method lets Invokestatic exec_instr StepI ics
+ by(auto split: if_split_asm)
+ next
+ case nD: False
+ then have nDone: "\<forall>sfs. sh D \<noteq> Some(sfs, Done)" using sfsi Some by simp
+ show ?thesis using nD
+ proof(cases i)
+ case Processing then show ?thesis using exec_step_ind_Invokestatic_NoMethod
+ exec_step_ind_Invokestatic_Init[where sh=sh, OF _ _ nDone nCalled]
+ sfsi False Static Some m method lets Invokestatic exec_instr StepI ics
+ by(auto split: if_split_asm)
+ next
+ case Prepared then show ?thesis using exec_step_ind_Invokestatic_NoMethod
+ exec_step_ind_Invokestatic_Init[where sh=sh, OF _ _ nDone nCalled]
+ sfsi False Static Some m method lets Invokestatic exec_instr StepI ics
+ by(auto split: if_split_asm)
+ next
+ case Error then show ?thesis using exec_step_ind_Invokestatic_NoMethod
+ exec_step_ind_Invokestatic_Init[where sh=sh, OF _ _ nDone nCalled]
+ sfsi False Static Some m method lets Invokestatic exec_instr StepI ics
+ by(auto split: if_split_asm)
+ qed(simp)
+ qed
+ qed
+ qed
+ qed
+ next
+ case Return
+ obtain D b Ts T m where method: "method P C M = (D,b,Ts,T,m)" by(cases "method P C M")
+ then obtain mxs mxl\<^sub>0 ins xt where "m = (mxs,mxl\<^sub>0,ins,xt)" by(cases m)
+ then show ?thesis using exec_step_ind_Return_Last_Init exec_step_ind_Return_Last
+ exec_step_ind_Return_Init exec_step_ind_Return_NonStatic exec_step_ind_Return_Static
+ method Return exec_instr StepI ics
+ by(auto split: if_split_asm staticb.splits bool.splits list.splits)
+ next
+ case Pop then show ?thesis using exec_instr StepI exec_step_ind_Pop by auto
+ next
+ case IAdd then show ?thesis using exec_instr StepI exec_step_ind_IAdd by auto
+ next
+ case Goto then show ?thesis using exec_instr StepI exec_step_ind_Goto by auto
+ next
+ case CmpEq then show ?thesis using exec_instr StepI exec_step_ind_CmpEq by auto
+ next
+ case (IfFalse x17) then show ?thesis
+ using exec_instr StepI exec_step_ind_IfFalse_nFalse exec_step_ind_IfFalse_False
+ exec_instr StepI by(auto split: val.splits staticb.splits)
+ next
+ case Throw then show ?thesis
+ using exec_instr StepI exec_step_ind_Throw exec_step_ind_Throw_Null
+ by(auto split: val.splits)
+ qed
+qed
+
+lemma exec_step_ind_imp_exec_step:
+assumes esi: "exec_step_ind si P h stk loc C M pc ics frs sh (xp', h', frs', sh')"
+ and si: "exec_step_input P C M pc ics = si"
+shows "exec_step P h stk loc C M pc ics frs sh = (xp', h', frs', sh')"
+proof -
+ have StepI:
+ "\<And>P C M pc Cs i . exec_step_input P C M pc (Called Cs) = StepI i
+ \<Longrightarrow> instrs_of P C M ! pc = i \<and> Cs = []"
+ proof -
+ fix P C M pc Cs i show "exec_step_input P C M pc (Called Cs) = StepI i
+ \<Longrightarrow> instrs_of P C M ! pc = i \<and> Cs = []" by(cases Cs; simp)
+ qed
+ have StepC:
+ "\<And>P C M pc ics C' Cs. exec_step_input P C M pc ics = StepC C' Cs \<Longrightarrow> ics = Calling C' Cs"
+ by simp
+ have StepT:
+ "\<And>P C M pc ics Cs a. exec_step_input P C M pc ics = StepT Cs a \<Longrightarrow> ics = Throwing Cs a"
+ by simp
+ show ?thesis using assms
+ proof(induct rule: exec_step_ind.induct)
+ case (exec_step_ind_NewOOM_Done sh C obj h ics P stk loc C\<^sub>0 M\<^sub>0 pc frs)
+ then show ?case by(cases ics, auto)
+ next
+ case (exec_step_ind_New_Done sh C obj h a ics P stk loc C\<^sub>0 M\<^sub>0 pc frs)
+ then show ?case by(cases ics, auto)
+ next
+ case (exec_step_ind_New_Init sh C ics P h stk loc C\<^sub>0 M\<^sub>0 pc frs)
+ then show ?case by(cases ics, auto split: init_state.splits)
+ next
+ case (exec_step_ind_Getfield_NoField v stk D fs h P F C loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
+ then show ?case by(cases "the (h (the_Addr (hd stk)))", cases ics, auto dest!: StepI)
+ next
+ case (exec_step_ind_Getfield_Static v stk D fs h P F t C loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
+ then show ?case
+ by(cases "the (h (the_Addr (hd stk)))", cases "fst(snd(field P C F))",
+ cases ics, auto simp: split_beta dest: has_field_sees[OF has_field_idemp] dest!: StepI)
+ next
+ case (exec_step_ind_Getfield v stk D fs h D' b t P C F loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
+ then show ?case
+ by(cases "the (h (the_Addr (hd stk)))",
+ cases ics; fastforce simp: split_beta dest: has_field_sees[OF has_field_idemp] dest!: StepI)
+ next
+ case (exec_step_ind_Getstatic_NonStatic P C F t D h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
+ then show ?case
+ by(cases ics; fastforce simp: split_beta split: staticb.splits
+ dest: has_field_sees[OF has_field_idemp] dest!: StepI)
+ next
+ case exec_step_ind_Getstatic_Called
+ then show ?case by(fastforce simp: split_beta split: staticb.splits dest!: StepI
+ dest: has_field_sees[OF has_field_idemp])
+ next
+ case (exec_step_ind_Getstatic_Done D' b t P D F C ics sh sfs v h stk loc C\<^sub>0 M\<^sub>0 pc frs)
+ then show ?case by(cases ics, auto simp: split_beta split: staticb.splits
+ dest: has_field_sees[OF has_field_idemp])
+ next
+ case (exec_step_ind_Getstatic_Init D' b t P D F C sh ics h stk loc C\<^sub>0 M\<^sub>0 pc frs)
+ then show ?case
+ by(cases ics, auto simp: split_beta split: init_state.splits staticb.splits
+ dest: has_field_sees[OF has_field_idemp])
+ next
+ case (exec_step_ind_Putfield_NoField r stk a D fs h P F C loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
+ then show ?case by(cases "the (h (the_Addr (hd(tl stk))))", cases ics, auto dest!: StepI)
+ next
+ case (exec_step_ind_Putfield_Static r stk a D fs h P F t C loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
+ then show ?case
+ by(cases "the (h (the_Addr (hd(tl stk))))", cases "fst(snd(field P C F))",
+ cases ics, auto simp: split_beta dest: has_field_sees[OF has_field_idemp] dest!: StepI)
+ next
+ case (exec_step_ind_Putfield v stk r a D fs h D' b t P C F loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
+ then show ?case
+ by(cases "the (h (the_Addr (hd(tl stk))))",
+ cases ics; fastforce simp: split_beta dest: has_field_sees[OF has_field_idemp] dest!: StepI)
+ next
+ case (exec_step_ind_Putstatic_NonStatic P C F t D h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
+ then show ?case
+ by(cases ics; fastforce simp: split_beta split: staticb.splits
+ dest: has_field_sees[OF has_field_idemp] dest!: StepI)
+ next
+ case exec_step_ind_Putstatic_Called
+ then show ?case by(fastforce simp: split_beta split: staticb.splits dest!: StepI
+ dest: has_field_sees[OF has_field_idemp])
+ next
+ case (exec_step_ind_Putstatic_Done D' b t P D F C ics sh sfs h stk loc C\<^sub>0 M\<^sub>0 pc frs)
+ then show ?case by(cases ics, auto simp: split_beta split: staticb.splits
+ dest: has_field_sees[OF has_field_idemp])
+ next
+ case (exec_step_ind_Putstatic_Init D' b t P D F C sh ics h stk loc C\<^sub>0 M\<^sub>0 pc frs)
+ then show ?case
+ by(cases ics, auto simp: split_beta split: staticb.splits init_state.splits
+ dest: has_field_sees[OF has_field_idemp])
+ next
+ case (exec_step_ind_Invoke ps n stk r C h D b Ts T mxs mxl\<^sub>0 ins xt P M m f' loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
+ then show ?case by(cases ics; fastforce dest!: StepI)
+ next
+ case (exec_step_ind_Invokestatic_Called ps n stk D b Ts T mxs mxl\<^sub>0 ins xt P C M m ics ics' sh)
+ then show ?case by(cases ics; fastforce dest!: StepI)
+ next
+ case (exec_step_ind_Invokestatic_Done ps n stk D b Ts T mxs mxl\<^sub>0 ins xt P C M m ics sh sfs f')
+ then show ?case by(cases ics; fastforce)
+ next
+ case (exec_step_ind_Invokestatic_Init D b Ts T mxs mxl\<^sub>0 ins xt P C M m sh ics n h stk loc C\<^sub>0 M\<^sub>0 pc frs)
+ then show ?case by(cases ics; fastforce split: init_state.splits)
+ next
+ case (exec_step_ind_Return_NonStatic D Ts T m P C\<^sub>0 M\<^sub>0 h stk\<^sub>0 loc\<^sub>0 pc ics stk' loc' C' m' pc' ics' frs' sh)
+ then show ?case by(cases "method P C\<^sub>0 M\<^sub>0", cases ics, auto dest!: StepI)
+ next
+ case (exec_step_ind_Return_Static D Ts T m P C\<^sub>0 M\<^sub>0 h stk\<^sub>0 loc\<^sub>0 pc ics stk' loc' C' m' pc' ics' frs' sh)
+ then show ?case by(cases "method P C\<^sub>0 M\<^sub>0", cases ics, auto dest!: StepI)
+ next
+ case (exec_step_ind_IfFalse_nFalse stk i P h loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
+ then show ?case by(cases "hd stk"; cases ics, auto dest!: StepI)
+ next
+ case (exec_step_ind_Throw_Null stk P h loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
+ then show ?case by(cases "hd stk"; cases ics, auto dest!: StepI)
+ next
+ case (exec_step_ind_Init C Cs P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
+ then have "ics = Called (C#Cs)" by simp
+ then show ?case by auto
+ (***)
+ next
+ case (exec_step_ind_Load n P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
+ then show ?case by(cases ics, auto dest!: StepI)
+ next
+ case (exec_step_ind_Store n P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
+ then show ?case by(cases ics, auto dest!: StepI)
+ next
+ case (exec_step_ind_Push v P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
+ then show ?case by(cases ics, auto dest!: StepI)
+ next
+ case (exec_step_ind_NewOOM_Called h C P stk loc C\<^sub>0 M\<^sub>0 pc frs sh ics')
+ then show ?case by(auto dest!: StepI)
+ next
+ case (exec_step_ind_New_Called h a C P stk loc C\<^sub>0 M\<^sub>0 pc frs sh ics')
+ then show ?case by(auto dest!: StepI)
+ next
+ case (exec_step_ind_Getfield_Null stk F C P h loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
+ then show ?case by(cases ics, auto dest!: StepI)
+ next
+ case (exec_step_ind_Getstatic_NoField P C F D h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
+ then show ?case by(cases ics, auto dest!: StepI)
+ next
+ case (exec_step_ind_Putfield_Null stk F C P h loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
+ then show ?case by(cases ics, auto dest!: StepI)
+ next
+ case (exec_step_ind_Putstatic_NoField P C F D h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
+ then show ?case by(cases ics, auto dest!: StepI)
+ next
+ case (exec_step_ind_Checkcast P C h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
+ then show ?case by(cases ics, auto dest!: StepI)
+ next
+ case (exec_step_ind_Checkcast_Error P C h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
+ then show ?case by(cases ics, auto dest!: StepI)
+ next
+ case (exec_step_ind_Invoke_Null stk n M P h loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
+ then show ?case by(cases ics, auto dest!: StepI)
+ next
+ case (exec_step_ind_Invoke_NoMethod r stk n C h P M loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
+ then show ?case by(cases ics, auto dest!: StepI)
+ next
+ case (exec_step_ind_Invoke_Static r stk n C h D b Ts T mxs mxl\<^sub>0 ins xt P M m loc C\<^sub>0 M\<^sub>0 pc ics)
+ then show ?case by(cases ics, auto dest!: StepI)
+ next
+ case (exec_step_ind_Invokestatic_NoMethod D b Ts T mxs mxl\<^sub>0 ins xt P C M n h stk loc C\<^sub>0 M\<^sub>0 pc ics)
+ then show ?case by(cases ics, auto dest!: StepI)
+ next
+ case (exec_step_ind_Invokestatic_NonStatic D b Ts T mxs mxl\<^sub>0 ins xt P C M m n h stk loc C\<^sub>0 M\<^sub>0 pc ics)
+ then show ?case by(cases ics, auto dest!: StepI)
+ next
+ case (exec_step_ind_Return_Last_Init P h stk\<^sub>0 loc\<^sub>0 C\<^sub>0 pc ics sh)
+ then show ?case by(cases ics, auto dest!: StepI)
+ next
+ case (exec_step_ind_Return_Last M\<^sub>0 P h stk\<^sub>0 loc\<^sub>0 C\<^sub>0 pc ics sh)
+ then show ?case by(cases ics, auto dest!: StepI)
+ next
+ case (exec_step_ind_Return_Init D b Ts T m P C\<^sub>0 h stk\<^sub>0 loc\<^sub>0 pc ics stk' loc' C' m' pc' ics')
+ then show ?case by(cases ics, auto dest!: StepI)
+ next
+ case (exec_step_ind_Pop P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
+ then show ?case by(cases ics, auto dest!: StepI)
+ next
+ case (exec_step_ind_IAdd P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
+ then show ?case by(cases ics, auto dest!: StepI)
+ next
+ case (exec_step_ind_IfFalse_False stk i P h loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
+ then show ?case by(cases ics, auto dest!: StepI)
+ next
+ case (exec_step_ind_CmpEq P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
+ then show ?case by(cases ics, auto dest!: StepI)
+ next
+ case (exec_step_ind_Goto i P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
+ then show ?case by(cases ics, auto dest!: StepI)
+ next
+ case (exec_step_ind_Throw stk P h loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
+ then show ?case by(cases ics, auto dest!: StepI)
+ next
+ case (exec_step_ind_Init_None_Called sh C Cs P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs)
+ then show ?case by(auto dest!: StepC)
+ next
+ case (exec_step_ind_Init_Done sh C sfs Cs P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs)
+ then show ?case by(auto dest!: StepC)
+ next
+ case (exec_step_ind_Init_Processing sh C sfs Cs P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs)
+ then show ?case by(auto dest!: StepC)
+ next
+ case (exec_step_ind_Init_Error sh C sfs Cs P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs)
+ then show ?case by(auto dest!: StepC)
+ next
+ case (exec_step_ind_Init_Prepared_Object sh C sfs sh' Cs P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs)
+ then show ?case by(auto dest!: StepC)
+ next
+ case (exec_step_ind_Init_Prepared_nObject sh C sfs sh' D P Cs h stk loc C\<^sub>0 M\<^sub>0 pc ics frs)
+ then show ?case by(auto dest!: StepC)
+ next
+ case (exec_step_ind_InitThrow C Cs a P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
+ then show ?case by(auto dest!: StepT)
+ next
+ case (exec_step_ind_InitThrow_End a P h stk loc C\<^sub>0 M\<^sub>0 pc ics frs sh)
+ then show ?case by(auto dest!: StepT)
+ qed
+qed
+
+\<comment> \<open> @{term exec_step} and @{term exec_step_ind} reach the same result given equivalent input \<close>
+lemma exec_step_ind_equiv:
+ "exec_step P h stk loc C M pc ics frs sh = (xp', h', frs', sh')
+ = exec_step_ind (exec_step_input P C M pc ics) P h stk loc C M pc ics frs sh (xp', h', frs', sh')"
+ using exec_step_imp_exec_step_ind exec_step_ind_imp_exec_step by auto
+
+end
diff --git a/thys/Regression_Test_Selection/JinjaSuppl/Subcls.thy b/thys/Regression_Test_Selection/JinjaSuppl/Subcls.thy
--- a/thys/Regression_Test_Selection/JinjaSuppl/Subcls.thy
+++ b/thys/Regression_Test_Selection/JinjaSuppl/Subcls.thy
@@ -1,66 +1,66 @@
-(* Title: RTS/JinjaSuppl/Subcls.thy
- Author: Susannah Mansky, UIUC 2020
- Description: Theory for the subcls relation
-*)
-
-section "@{term subcls} theory"
-
-theory Subcls
-imports JinjaDCI.TypeRel
-begin
-
-lemma subcls_class_ex: "\<lbrakk> P \<turnstile> C \<preceq>\<^sup>* C'; C \<noteq> C' \<rbrakk>
- \<Longrightarrow> \<exists>D' fs ms. class P C = \<lfloor>(D', fs, ms)\<rfloor>"
-proof(induct rule: converse_rtrancl_induct)
- case (step y z) then show ?case by(auto dest: subcls1D)
-qed(simp)
-
-lemma class_subcls1:
- "\<lbrakk> class P y = class P' y; P \<turnstile> y \<prec>\<^sup>1 z \<rbrakk> \<Longrightarrow> P' \<turnstile> y \<prec>\<^sup>1 z"
- by(auto dest!: subcls1D intro!: subcls1I intro: sym)
-
-
-lemma subcls1_single_valued: "single_valued (subcls1 P)"
-by (clarsimp simp: single_valued_def subcls1.simps)
-
-lemma subcls_confluent:
- "\<lbrakk> P \<turnstile> C \<preceq>\<^sup>* C'; P \<turnstile> C \<preceq>\<^sup>* C'' \<rbrakk> \<Longrightarrow> P \<turnstile> C' \<preceq>\<^sup>* C'' \<or> P \<turnstile> C'' \<preceq>\<^sup>* C'"
- by (simp add: single_valued_confluent subcls1_single_valued)
-
-lemma subcls1_confluent: "\<lbrakk> P \<turnstile> a \<prec>\<^sup>1 b; P \<turnstile> a \<preceq>\<^sup>* c; a \<noteq> c \<rbrakk> \<Longrightarrow> P \<turnstile> b \<preceq>\<^sup>* c"
-using subcls1_single_valued
- by (auto elim!: converse_rtranclE[where x=a] simp: single_valued_def)
-
-
-lemma subcls_self_superclass: "\<lbrakk> P \<turnstile> C \<prec>\<^sup>1 C; P \<turnstile> C \<preceq>\<^sup>* D \<rbrakk> \<Longrightarrow> D = C"
-using subcls1_single_valued
- by (auto elim!: rtrancl_induct[where b=D] simp: single_valued_def)
-
-lemma subcls_of_Obj_acyclic:
- "\<lbrakk> P \<turnstile> C \<preceq>\<^sup>* Object; C \<noteq> D\<rbrakk> \<Longrightarrow> \<not>(P \<turnstile> C \<preceq>\<^sup>* D \<and> P \<turnstile> D \<preceq>\<^sup>* C)"
-proof(induct arbitrary: D rule: converse_rtrancl_induct)
- case base then show ?case by simp
-next
- case (step y z) show ?case
- proof(cases "y=z")
- case True with step show ?thesis by simp
- next
- case False show ?thesis
- proof(cases "z = D")
- case True with False step.hyps(3)[of y] show ?thesis by clarsimp
- next
- case neq: False
- with step.hyps(3) have "\<not>(P \<turnstile> z \<preceq>\<^sup>* D \<and> P \<turnstile> D \<preceq>\<^sup>* z)" by simp
- moreover from step.hyps(1)
- have "P \<turnstile> D \<preceq>\<^sup>* y \<Longrightarrow> P \<turnstile> D \<preceq>\<^sup>* z" by(simp add: rtrancl_into_rtrancl)
- moreover from step.hyps(1) step.prems(1)
- have "P \<turnstile> y \<preceq>\<^sup>* D \<Longrightarrow> P \<turnstile> z \<preceq>\<^sup>* D" by(simp add: subcls1_confluent)
- ultimately show ?thesis by clarsimp
- qed
- qed
-qed
-
-lemma subcls_of_Obj: "\<lbrakk> P \<turnstile> C \<preceq>\<^sup>* Object; P \<turnstile> C \<preceq>\<^sup>* D \<rbrakk> \<Longrightarrow> P \<turnstile> D \<preceq>\<^sup>* Object"
- by(auto dest: subcls_confluent)
-
+(* Title: RTS/JinjaSuppl/Subcls.thy
+ Author: Susannah Mansky, UIUC 2020
+ Description: Theory for the subcls relation
+*)
+
+section "@{term subcls} theory"
+
+theory Subcls
+imports JinjaDCI.TypeRel
+begin
+
+lemma subcls_class_ex: "\<lbrakk> P \<turnstile> C \<preceq>\<^sup>* C'; C \<noteq> C' \<rbrakk>
+ \<Longrightarrow> \<exists>D' fs ms. class P C = \<lfloor>(D', fs, ms)\<rfloor>"
+proof(induct rule: converse_rtrancl_induct)
+ case (step y z) then show ?case by(auto dest: subcls1D)
+qed(simp)
+
+lemma class_subcls1:
+ "\<lbrakk> class P y = class P' y; P \<turnstile> y \<prec>\<^sup>1 z \<rbrakk> \<Longrightarrow> P' \<turnstile> y \<prec>\<^sup>1 z"
+ by(auto dest!: subcls1D intro!: subcls1I intro: sym)
+
+
+lemma subcls1_single_valued: "single_valued (subcls1 P)"
+by (clarsimp simp: single_valued_def subcls1.simps)
+
+lemma subcls_confluent:
+ "\<lbrakk> P \<turnstile> C \<preceq>\<^sup>* C'; P \<turnstile> C \<preceq>\<^sup>* C'' \<rbrakk> \<Longrightarrow> P \<turnstile> C' \<preceq>\<^sup>* C'' \<or> P \<turnstile> C'' \<preceq>\<^sup>* C'"
+ by (simp add: single_valued_confluent subcls1_single_valued)
+
+lemma subcls1_confluent: "\<lbrakk> P \<turnstile> a \<prec>\<^sup>1 b; P \<turnstile> a \<preceq>\<^sup>* c; a \<noteq> c \<rbrakk> \<Longrightarrow> P \<turnstile> b \<preceq>\<^sup>* c"
+using subcls1_single_valued
+ by (auto elim!: converse_rtranclE[where x=a] simp: single_valued_def)
+
+
+lemma subcls_self_superclass: "\<lbrakk> P \<turnstile> C \<prec>\<^sup>1 C; P \<turnstile> C \<preceq>\<^sup>* D \<rbrakk> \<Longrightarrow> D = C"
+using subcls1_single_valued
+ by (auto elim!: rtrancl_induct[where b=D] simp: single_valued_def)
+
+lemma subcls_of_Obj_acyclic:
+ "\<lbrakk> P \<turnstile> C \<preceq>\<^sup>* Object; C \<noteq> D\<rbrakk> \<Longrightarrow> \<not>(P \<turnstile> C \<preceq>\<^sup>* D \<and> P \<turnstile> D \<preceq>\<^sup>* C)"
+proof(induct arbitrary: D rule: converse_rtrancl_induct)
+ case base then show ?case by simp
+next
+ case (step y z) show ?case
+ proof(cases "y=z")
+ case True with step show ?thesis by simp
+ next
+ case False show ?thesis
+ proof(cases "z = D")
+ case True with False step.hyps(3)[of y] show ?thesis by clarsimp
+ next
+ case neq: False
+ with step.hyps(3) have "\<not>(P \<turnstile> z \<preceq>\<^sup>* D \<and> P \<turnstile> D \<preceq>\<^sup>* z)" by simp
+ moreover from step.hyps(1)
+ have "P \<turnstile> D \<preceq>\<^sup>* y \<Longrightarrow> P \<turnstile> D \<preceq>\<^sup>* z" by(simp add: rtrancl_into_rtrancl)
+ moreover from step.hyps(1) step.prems(1)
+ have "P \<turnstile> y \<preceq>\<^sup>* D \<Longrightarrow> P \<turnstile> z \<preceq>\<^sup>* D" by(simp add: subcls1_confluent)
+ ultimately show ?thesis by clarsimp
+ qed
+ qed
+qed
+
+lemma subcls_of_Obj: "\<lbrakk> P \<turnstile> C \<preceq>\<^sup>* Object; P \<turnstile> C \<preceq>\<^sup>* D \<rbrakk> \<Longrightarrow> P \<turnstile> D \<preceq>\<^sup>* Object"
+ by(auto dest: subcls_confluent)
+
end
\ No newline at end of file
diff --git a/thys/Three_Circles/Bernstein.thy b/thys/Three_Circles/Bernstein.thy
--- a/thys/Three_Circles/Bernstein.thy
+++ b/thys/Three_Circles/Bernstein.thy
@@ -1,406 +1,406 @@
-section \<open>Bernstein Polynomials over any finite interval\<close>
-
-theory Bernstein
- imports "Bernstein_01"
-begin
-
-subsection \<open>Definition and relation to Bernstein Polynomials over [0, 1]\<close>
-
-definition Bernstein_Poly :: "nat \<Rightarrow> nat \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real poly" where
- "Bernstein_Poly j p c d = smult ((p choose j)/(d - c)^p)
- (((monom 1 j) \<circ>\<^sub>p [:-c, 1:]) * (monom 1 (p-j) \<circ>\<^sub>p [:d, -1:]))"
-
-lemma Bernstein_Poly_altdef:
- assumes "c \<noteq> d" and "j \<le> p"
- shows "Bernstein_Poly j p c d = smult (p choose j)
- ([:-c/(d-c), 1/(d-c):]^j * [:d/(d-c), -1/(d-c):]^(p-j))"
- (is "?L = ?R")
-proof -
- have "?L = smult (p choose j) (smult ((1/(d - c))^j)
- (smult ((1/(d - c))^(p-j)) ([:-c, 1:]^j * [:d, -1:]^(p-j))))"
- using assms by (auto simp: Bernstein_Poly_def monom_altdef hom_distribs
- pcompose_pCons smult_eq_iff field_simps power_add[symmetric])
- also have "... = ?R"
- apply (subst mult_smult_right[symmetric])
- apply (subst mult_smult_left[symmetric])
- apply (subst smult_power)
- apply (subst smult_power)
- by auto
- finally show ?thesis .
-qed
-
-lemma Bernstein_Poly_nonneg:
- assumes "c \<le> x" and "x \<le> d"
- shows "poly (Bernstein_Poly j p c d) x \<ge> 0"
- using assms by (auto simp: Bernstein_Poly_def poly_pcompose poly_monom)
-
-lemma Bernstein_Poly_01: "Bernstein_Poly j p 0 1 = Bernstein_Poly_01 j p"
- by (auto simp: Bernstein_Poly_def Bernstein_Poly_01_def monom_altdef)
-
-lemma Bernstein_Poly_rescale:
- assumes "a \<noteq> b"
- shows "Bernstein_Poly j p c d \<circ>\<^sub>p [:a, 1:] \<circ>\<^sub>p [:0, b-a:]
- = Bernstein_Poly j p ((c-a)/(b-a)) ((d-a)/(b-a))"
- (is "?L = ?R")
-proof -
- have "?R = smult (real (p choose j)
- / ((d - a) / (b - a) - (c - a) / (b - a)) ^ p)
- ([:- ((c - a) / (b - a)), 1:] ^ j
- * [:(d - a) / (b - a), - 1:] ^ (p - j))"
- by (auto simp: Bernstein_Poly_def monom_altdef hom_distribs
- pcompose_pCons)
- also have "... = smult (real (p choose j) / ((d - c) / (b - a)) ^ p)
- ([:- ((c - a) / (b - a)), 1:] ^ j * [:(d - a) / (b - a), - 1:]
- ^ (p - j))"
- by argo
- also have "... = smult (real (p choose j) / (d - c) ^ p)
- (smult ((b - a) ^ (p - j)) (smult ((b - a) ^ j)
- ([:- ((c - a) / (b - a)), 1:] ^ j * [:(d - a) / (b - a), - 1:]
- ^ (p - j))))"
- by (auto simp: power_add[symmetric] power_divide)
- also have "... = smult (real (p choose j) / (d - c) ^ p)
- ([:- (c - a), b - a:] ^ j * [:d - a, -(b - a):] ^ (p - j))"
- apply (subst mult_smult_left[symmetric])
- apply (subst mult_smult_right[symmetric])
- using assms by (auto simp: smult_power)
- also have "... = ?L"
- using assms
- by (auto simp: Bernstein_Poly_def monom_altdef pcompose_mult
- pcompose_smult hom_distribs pcompose_pCons)
- finally show ?thesis by presburger
-qed
-
-lemma Bernstein_Poly_rescale_01:
- assumes "c \<noteq> d"
- shows "Bernstein_Poly j p c d \<circ>\<^sub>p [:c, 1:] \<circ>\<^sub>p [:0, d-c:]
- = Bernstein_Poly_01 j p"
- apply (subst Bernstein_Poly_rescale)
- using assms by (auto simp: Bernstein_Poly_01)
-
-lemma Bernstein_Poly_eq_rescale_01:
- assumes "c \<noteq> d"
- shows "Bernstein_Poly j p c d = Bernstein_Poly_01 j p
- \<circ>\<^sub>p [:0, 1/(d-c):] \<circ>\<^sub>p [:-c, 1:]"
- apply (subst Bernstein_Poly_rescale_01[symmetric])
- using assms by (auto simp: pcompose_pCons pcompose_assoc[symmetric])
-
-lemma coeff_Bernstein_sum:
- fixes b::"nat \<Rightarrow> real" and p::nat and c d::real
- defines "P \<equiv> (\<Sum>j = 0..p. (smult (b j) (Bernstein_Poly j p c d)))"
- assumes "i \<le> p" and "c \<noteq> d"
- shows "coeff ((reciprocal_poly p (P \<circ>\<^sub>p [:c, 1:]
- \<circ>\<^sub>p [:0, d-c:])) \<circ>\<^sub>p [:1, 1:]) (p - i) = (p choose i) * (b i)"
-proof -
- have h: "P \<circ>\<^sub>p [:c, 1:] \<circ>\<^sub>p [:0, d-c:]
- = (\<Sum>j = 0..p. (smult (b j) (Bernstein_Poly_01 j p)))"
- using assms
- by (auto simp: P_def pcompose_sum pcompose_smult
- pcompose_add Bernstein_Poly_rescale_01)
- then show ?thesis
- using coeff_Bernstein_sum_01 assms by simp
-qed
-
-lemma Bernstein_sum:
- assumes "c \<noteq> d" and "degree P \<le> p"
- shows "P = (\<Sum>j = 0..p. smult (inverse (real (p choose j))
- * coeff (reciprocal_poly p (P \<circ>\<^sub>p [:c, 1:] \<circ>\<^sub>p [:0, d-c:])
- \<circ>\<^sub>p [:1, 1:]) (p-j)) (Bernstein_Poly j p c d))"
- apply (subst Bernstein_Poly_eq_rescale_01)
- subgoal using assms by blast
- subgoal
- apply (subst pcompose_smult[symmetric])
- apply (subst pcompose_sum[symmetric])
- apply (subst pcompose_smult[symmetric])
- apply (subst pcompose_sum[symmetric])
- apply (subst Bernstein_sum_01[symmetric])
- using assms by (auto simp: degree_pcompose pcompose_assoc[symmetric]
- pcompose_pCons)
- done
-
-lemma Bernstein_Poly_span1:
- assumes "c \<noteq> d" and "degree P \<le> p"
- shows "P \<in> poly_vs.span {Bernstein_Poly x p c d | x. x \<le> p}"
-proof (subst Bernstein_sum[OF assms], rule poly_vs.span_sum)
- fix x :: nat
- assume "x \<in> {0..p}"
- then have "\<exists>n. Bernstein_Poly x p c d = Bernstein_Poly n p c d \<and> n \<le> p"
- by auto
- then have
- "Bernstein_Poly x p c d \<in> poly_vs.span {Bernstein_Poly n p c d |n. n \<le> p}"
- by (simp add: poly_vs.span_base)
- thus "smult (inverse (real (p choose x)) *
- coeff (reciprocal_poly p (P \<circ>\<^sub>p [:c, 1:] \<circ>\<^sub>p [:0, d - c:]) \<circ>\<^sub>p [:1, 1:])
- (p - x)) (Bernstein_Poly x p c d)
- \<in> poly_vs.span {Bernstein_Poly x p c d |x. x \<le> p}"
- by (rule poly_vs.span_scale)
-qed
-
-lemma Bernstein_Poly_span:
- assumes "c \<noteq> d"
- shows "poly_vs.span {Bernstein_Poly x p c d | x. x \<le> p} = {x. degree x \<le> p}"
-proof (subst Bernstein_Poly_01_span[symmetric], subst poly_vs.span_eq, rule conjI)
- show "{Bernstein_Poly x p c d |x. x \<le> p}
- \<subseteq> poly_vs.span {Bernstein_Poly_01 x p |x. x \<le> p}"
- apply (subst Setcompr_subset)
- apply (rule allI, rule impI)
- apply (rule Bernstein_Poly_01_span1)
- using assms by (auto simp: degree_Bernstein_le Bernstein_Poly_eq_rescale_01
- degree_pcompose)
-
- show "{Bernstein_Poly_01 x p |x. x \<le> p}
- \<subseteq> poly_vs.span {Bernstein_Poly x p c d |x. x \<le> p}"
- apply (subst Setcompr_subset)
- apply (rule allI, rule impI)
- apply (rule Bernstein_Poly_span1)
- using assms by (auto simp: degree_Bernstein_le)
-qed
-
-lemma Bernstein_Poly_independent: assumes "c \<noteq> d"
- shows "poly_vs.independent {Bernstein_Poly x p c d | x. x \<in> {..p}}"
-proof (rule poly_vs.card_le_dim_spanning)
- show "{Bernstein_Poly x p c d |x. x \<in> {.. p}} \<subseteq> {x. degree x \<le> p}"
- using assms
- by (auto simp: degree_Bernstein Bernstein_Poly_eq_rescale_01 degree_pcompose)
- show "{x. degree x \<le> p} \<subseteq> poly_vs.span {Bernstein_Poly x p c d |x. x \<in> {..p}}"
- using assms by (auto simp: Bernstein_Poly_span1)
- show "finite {Bernstein_Poly x p c d |x. x \<in> {..p}}" by fastforce
- show "card {Bernstein_Poly x p c d |x. x \<in> {..p}} \<le> poly_vs.dim {x. degree x \<le> p}"
- apply (rule le_trans)
- apply (subst image_Collect[symmetric], rule card_image_le, force)
- by (force simp: dim_degree)
-qed
-
-subsection \<open>Bernstein coefficients and changes over any interval\<close>
-
-definition Bernstein_coeffs ::
- "nat \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real poly \<Rightarrow> real list" where
- "Bernstein_coeffs p c d P =
- [(inverse (real (p choose j)) *
- coeff (reciprocal_poly p (P \<circ>\<^sub>p [:c, 1:] \<circ>\<^sub>p [:0, d-c:]) \<circ>\<^sub>p [:1, 1:]) (p-j)).
- j \<leftarrow> [0..<(p+1)]]"
-
-lemma Bernstein_coeffs_eq_rescale: assumes "c \<noteq> d"
- shows "Bernstein_coeffs p c d P = Bernstein_coeffs_01 p (P \<circ>\<^sub>p [:c, 1:] \<circ>\<^sub>p [:0, d-c:])"
- using assms by (auto simp: pcompose_pCons pcompose_assoc[symmetric]
- Bernstein_coeffs_def Bernstein_coeffs_01_def)
-
-lemma nth_default_Bernstein_coeffs: assumes "degree P \<le> p"
- shows "nth_default 0 (Bernstein_coeffs p c d P) i =
- inverse (p choose i) * coeff
- (reciprocal_poly p (P \<circ>\<^sub>p [:c, 1:] \<circ>\<^sub>p [:0, d-c:]) \<circ>\<^sub>p [:1, 1:]) (p-i)"
- apply (cases "p = i")
- using assms by (auto simp: Bernstein_coeffs_def nth_default_append
- nth_default_Cons Nitpick.case_nat_unfold binomial_eq_0)
-
-lemma Bernstein_coeffs_sum: assumes "c \<noteq> d" and hP: "degree P \<le> p"
- shows "P = (\<Sum>j = 0..p. smult (nth_default 0 (Bernstein_coeffs p c d P) j)
- (Bernstein_Poly j p c d))"
- apply (subst nth_default_Bernstein_coeffs[OF hP])
- apply (subst Bernstein_sum[OF assms])
- by argo
-
-definition Bernstein_changes :: "nat \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real poly \<Rightarrow> int" where
- "Bernstein_changes p c d P = nat (changes (Bernstein_coeffs p c d P))"
-
-lemma Bernstein_changes_eq_rescale: assumes "c \<noteq> d" and "degree P \<le> p"
- shows "Bernstein_changes p c d P =
- Bernstein_changes_01 p (P \<circ>\<^sub>p [:c, 1:] \<circ>\<^sub>p [:0, d-c:])"
- using assms by (auto simp: Bernstein_coeffs_eq_rescale Bernstein_changes_def
- Bernstein_changes_01_def)
-
-text \<open>This is related and mostly equivalent to previous Descartes test \cite{li2019counting}\<close>
-lemma Bernstein_changes_test:
- fixes P::"real poly"
- assumes "degree P \<le> p" and "P \<noteq> 0" and "c < d"
- shows "proots_count P {x. c < x \<and> x < d} \<le> Bernstein_changes p c d P \<and>
- even (Bernstein_changes p c d P - proots_count P {x. c < x \<and> x < d})"
-proof -
- define Q where "Q=P \<circ>\<^sub>p [:c, 1:] \<circ>\<^sub>p [:0, d - c:]"
-
- have "int (proots_count Q {x. 0 < x \<and> x < 1})
- \<le> Bernstein_changes_01 p Q \<and>
- even (Bernstein_changes_01 p Q -
- int (proots_count Q {x. 0 < x \<and> x < 1}))"
- unfolding Q_def
- apply (rule Bernstein_changes_01_test)
- subgoal using assms by fastforce
- subgoal using assms by (auto simp: pcompose_eq_0)
- done
- moreover have "proots_count P {x. c < x \<and> x < d} =
- proots_count Q {x. 0 < x \<and> x < 1}"
- unfolding Q_def
- proof (subst proots_pcompose)
- have "poly [:c, 1:] ` poly [:0, d - c:] ` {x. 0 < x \<and> x < 1} =
- {x. c < x \<and> x < d}" (is "?L = ?R")
- proof
- have "c + x * (d - c) < d" if "x < 1" for x
- proof -
- have "x * (d - c) < 1 * (d - c)"
- using \<open>c < d\<close> that by force
- then show ?thesis by fastforce
- qed
- then show "?L \<subseteq> ?R"
- using assms by auto
- next
- show "?R \<subseteq> ?L"
- proof
- fix x::real assume "x \<in> ?R"
- hence "c < x" and "x < d" by auto
- thus "x \<in> ?L"
- proof (subst image_eqI)
- show "x = poly [:c, 1:] (x - c)" by force
- assume "c < x" and "x < d"
- thus "x - c \<in> poly [:0, d - c:] ` {x. 0 < x \<and> x < 1}"
- proof (subst image_eqI)
- show "x - c = poly [:0, d - c:] ((x - c)/(d - c))"
- using assms by fastforce
- assume "c < x" and "x < d"
- thus "(x - c) / (d - c) \<in> {x. 0 < x \<and> x < 1}"
- by auto
- qed fast
- qed fast
- qed
- qed
- then show "proots_count P {x. c < x \<and> x < d} =
- proots_count (P \<circ>\<^sub>p [:c, 1:])
- (poly [:0, d - c:] ` {x. 0 < x \<and> x < 1})"
- using assms by (auto simp:proots_pcompose)
- show "P \<circ>\<^sub>p [:c, 1:] \<noteq> 0"
- by (simp add: pcompose_eq_0 assms(2))
- show "degree [:0, d - c:] = 1"
- using assms by auto
- qed
- moreover have " Bernstein_changes p c d P = Bernstein_changes_01 p Q"
- unfolding Q_def
- apply (rule Bernstein_changes_eq_rescale)
- using assms by auto
- ultimately show ?thesis by auto
-qed
-
-subsection \<open>The control polygon of a polynomial\<close>
-
-definition control_points ::
- "nat \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real poly \<Rightarrow> (real \<times> real) list"
-where
- "control_points p c d P =
- [(((real i)*d + (real (p - i))*c)/p,
- nth_default 0 (Bernstein_coeffs p c d P) i).
- i \<leftarrow> [0..<(p+1)]]"
-
-lemma line_above:
- fixes a b c d :: real and p :: nat and P :: "real poly"
- assumes hline: "\<And>i. i \<le> p \<Longrightarrow> a * (((real i)*d + (real (p - i))*c)/p) + b \<ge>
- nth_default 0 (Bernstein_coeffs p c d P) i"
- and hp: "p \<noteq> 0" and hcd: "c \<noteq> d" and hP: "degree P \<le> p"
- shows "\<And>x. c \<le> x \<Longrightarrow> x \<le> d \<Longrightarrow> a*x + b \<ge> poly P x"
-proof -
- fix x
- assume hc: "c \<le> x" and hd: "x \<le> d"
-
- have bern_eq:"Bernstein_coeffs p c d [:b, a:] =
- [a*(real i * d + real (p - i) * c)/p + b. i \<leftarrow> [0..<(p+1)]]"
- proof -
- have "Bernstein_coeffs p c d [:b, a:] = map (nth_default 0
- (Bernstein_coeffs_01 p ([:b, a:] \<circ>\<^sub>p [:c, 1:] \<circ>\<^sub>p [:0, d - c:])))
- [0..<p+1]"
- apply (subst Bernstein_coeffs_eq_rescale["OF" hcd])
- apply (subst map_nth_default[symmetric])
- apply (subst length_Bernstein_coeffs_01)
- by blast
- also have
- "... = map (\<lambda>i. a * (real i * d + real (p - i) * c) / real p + b) [0..<p + 1]"
- proof (rule map_cong)
- fix x assume hx: "x \<in> set [0..<p + 1]"
- have "nth_default 0 (Bernstein_coeffs_01 p
- ([:b, a:] \<circ>\<^sub>p [:c, 1:] \<circ>\<^sub>p [:0, d - c:])) x =
- nth_default 0 (Bernstein_coeffs_01 p
- (smult (b + a*c) 1 + smult (a*(d - c)) (monom 1 1))) x"
- proof-
- have "[:b, a:] \<circ>\<^sub>p [:c, 1:] \<circ>\<^sub>p [:0, d - c:] =
- smult (b + a*c) 1 + smult (a*(d - c)) (monom 1 1)"
- by (simp add: monom_altdef pcompose_pCons)
- then show ?thesis by auto
- qed
- also have "... =
- nth_default 0 (Bernstein_coeffs_01 p (smult (b + a * c) 1)) x +
- nth_default 0 (Bernstein_coeffs_01 p (smult (a * (d - c)) (monom 1 1))) x"
- apply (subst Bernstein_coeffs_01_add)
- using hp by (auto simp: degree_monom_eq)
- also have "... =
- (b + a*c) * nth_default 0 (Bernstein_coeffs_01 p 1) x +
- (a*(d - c)) * nth_default 0 (Bernstein_coeffs_01 p (monom 1 1)) x"
- apply (subst Bernstein_coeffs_01_smult)
- using hp by (auto simp: Bernstein_coeffs_01_smult degree_monom_eq)
- also have "... =
- (b + a * c) * (if x < p + 1 then 1 else 0) +
- a * (d - c) * (real (nth_default 0 [0..<p + 1] x) / real p)"
- apply (subst Bernstein_coeffs_01_1, subst Bernstein_coeffs_01_x[OF hp])
- apply (subst nth_default_replicate_eq, subst nth_default_map_eq[of _ 0])
- by auto
- also have "... =
- (b + a * c) * (if x < p + 1 then 1 else 0) +
- a * (d - c) * (real ([0..<p + 1] ! x) / real p)"
- apply (subst nth_default_nth)
- using hx by auto
- also have "... = (b + a * c) * (if x < p + 1 then 1 else 0) +
- a * (d - c) * (real (0 + x) / real p)"
- apply (subst nth_upt)
- using hx by auto
- also have "... = a * (real x * d + real (p - x) * c) / real p + b"
- apply (subst of_nat_diff)
- using hx hp by (auto simp: field_simps)
- finally show "nth_default 0 (Bernstein_coeffs_01 p
- ([:b, a:] \<circ>\<^sub>p [:c, 1:] \<circ>\<^sub>p [:0, d - c:])) x =
- a * (real x * d + real (p - x) * c) / real p + b" .
- qed blast
- finally show ?thesis .
- qed
-
- have nth_default_geq:"nth_default 0 (Bernstein_coeffs p c d [:b, a:]) i \<ge>
- nth_default 0 (Bernstein_coeffs p c d P) i" for i
- proof -
- show "nth_default 0 (Bernstein_coeffs p c d [:b, a:]) i \<ge>
- nth_default 0 (Bernstein_coeffs p c d P) i"
- proof cases
- define p1 where "p1 \<equiv> p+1"
- assume h: "i \<le> p"
- hence "nth_default 0 (Bernstein_coeffs p c d P) i \<le>
- a * (((real i)*d + (real (p - i))*c)/p) + b"
- by (rule hline)
- also have "... = nth_default 0 (map (\<lambda>i. a * (real i * d
- + real (p - i) * c) / real p + b) [0..<p + 1]) i"
- apply (subst p1_def[symmetric])
- using h apply (auto simp: nth_default_def)
- by (auto simp: p1_def)
- also have "... = nth_default 0 (Bernstein_coeffs p c d [:b, a:]) i"
- using bern_eq by simp
- finally show ?thesis .
- next
- assume h: "\<not>i \<le> p"
- thus ?thesis
- using assms
- by (auto simp: nth_default_def Bernstein_coeffs_eq_rescale
- length_Bernstein_coeffs_01)
- qed
- qed
-
- have "poly P x = (\<Sum>k = 0..p.
- poly (smult (nth_default 0 (Bernstein_coeffs p c d P) k)
- (Bernstein_Poly k p c d)) x)"
- apply (subst Bernstein_coeffs_sum[OF hcd hP])
- by (rule poly_sum)
- also have "... \<le> (\<Sum>k = 0..p.
- poly (smult (nth_default 0 (Bernstein_coeffs p c d [:b, a:]) k)
- (Bernstein_Poly k p c d)) x)"
- apply (rule sum_mono)
- using mult_right_mono[OF nth_default_geq] Bernstein_Poly_nonneg[OF hc hd]
- by auto
- also have "... = poly [:b, a:] x"
- apply (subst(2) Bernstein_coeffs_sum[of c d "[:b, a:]" p])
- using assms apply auto[2]
- by (rule poly_sum[symmetric])
- also have "... = a*x + b" by force
- finally show "poly P x \<le> a*x + b" .
-qed
-
+section \<open>Bernstein Polynomials over any finite interval\<close>
+
+theory Bernstein
+ imports "Bernstein_01"
+begin
+
+subsection \<open>Definition and relation to Bernstein Polynomials over [0, 1]\<close>
+
+definition Bernstein_Poly :: "nat \<Rightarrow> nat \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real poly" where
+ "Bernstein_Poly j p c d = smult ((p choose j)/(d - c)^p)
+ (((monom 1 j) \<circ>\<^sub>p [:-c, 1:]) * (monom 1 (p-j) \<circ>\<^sub>p [:d, -1:]))"
+
+lemma Bernstein_Poly_altdef:
+ assumes "c \<noteq> d" and "j \<le> p"
+ shows "Bernstein_Poly j p c d = smult (p choose j)
+ ([:-c/(d-c), 1/(d-c):]^j * [:d/(d-c), -1/(d-c):]^(p-j))"
+ (is "?L = ?R")
+proof -
+ have "?L = smult (p choose j) (smult ((1/(d - c))^j)
+ (smult ((1/(d - c))^(p-j)) ([:-c, 1:]^j * [:d, -1:]^(p-j))))"
+ using assms by (auto simp: Bernstein_Poly_def monom_altdef hom_distribs
+ pcompose_pCons smult_eq_iff field_simps power_add[symmetric])
+ also have "... = ?R"
+ apply (subst mult_smult_right[symmetric])
+ apply (subst mult_smult_left[symmetric])
+ apply (subst smult_power)
+ apply (subst smult_power)
+ by auto
+ finally show ?thesis .
+qed
+
+lemma Bernstein_Poly_nonneg:
+ assumes "c \<le> x" and "x \<le> d"
+ shows "poly (Bernstein_Poly j p c d) x \<ge> 0"
+ using assms by (auto simp: Bernstein_Poly_def poly_pcompose poly_monom)
+
+lemma Bernstein_Poly_01: "Bernstein_Poly j p 0 1 = Bernstein_Poly_01 j p"
+ by (auto simp: Bernstein_Poly_def Bernstein_Poly_01_def monom_altdef)
+
+lemma Bernstein_Poly_rescale:
+ assumes "a \<noteq> b"
+ shows "Bernstein_Poly j p c d \<circ>\<^sub>p [:a, 1:] \<circ>\<^sub>p [:0, b-a:]
+ = Bernstein_Poly j p ((c-a)/(b-a)) ((d-a)/(b-a))"
+ (is "?L = ?R")
+proof -
+ have "?R = smult (real (p choose j)
+ / ((d - a) / (b - a) - (c - a) / (b - a)) ^ p)
+ ([:- ((c - a) / (b - a)), 1:] ^ j
+ * [:(d - a) / (b - a), - 1:] ^ (p - j))"
+ by (auto simp: Bernstein_Poly_def monom_altdef hom_distribs
+ pcompose_pCons)
+ also have "... = smult (real (p choose j) / ((d - c) / (b - a)) ^ p)
+ ([:- ((c - a) / (b - a)), 1:] ^ j * [:(d - a) / (b - a), - 1:]
+ ^ (p - j))"
+ by argo
+ also have "... = smult (real (p choose j) / (d - c) ^ p)
+ (smult ((b - a) ^ (p - j)) (smult ((b - a) ^ j)
+ ([:- ((c - a) / (b - a)), 1:] ^ j * [:(d - a) / (b - a), - 1:]
+ ^ (p - j))))"
+ by (auto simp: power_add[symmetric] power_divide)
+ also have "... = smult (real (p choose j) / (d - c) ^ p)
+ ([:- (c - a), b - a:] ^ j * [:d - a, -(b - a):] ^ (p - j))"
+ apply (subst mult_smult_left[symmetric])
+ apply (subst mult_smult_right[symmetric])
+ using assms by (auto simp: smult_power)
+ also have "... = ?L"
+ using assms
+ by (auto simp: Bernstein_Poly_def monom_altdef pcompose_mult
+ pcompose_smult hom_distribs pcompose_pCons)
+ finally show ?thesis by presburger
+qed
+
+lemma Bernstein_Poly_rescale_01:
+ assumes "c \<noteq> d"
+ shows "Bernstein_Poly j p c d \<circ>\<^sub>p [:c, 1:] \<circ>\<^sub>p [:0, d-c:]
+ = Bernstein_Poly_01 j p"
+ apply (subst Bernstein_Poly_rescale)
+ using assms by (auto simp: Bernstein_Poly_01)
+
+lemma Bernstein_Poly_eq_rescale_01:
+ assumes "c \<noteq> d"
+ shows "Bernstein_Poly j p c d = Bernstein_Poly_01 j p
+ \<circ>\<^sub>p [:0, 1/(d-c):] \<circ>\<^sub>p [:-c, 1:]"
+ apply (subst Bernstein_Poly_rescale_01[symmetric])
+ using assms by (auto simp: pcompose_pCons pcompose_assoc[symmetric])
+
+lemma coeff_Bernstein_sum:
+ fixes b::"nat \<Rightarrow> real" and p::nat and c d::real
+ defines "P \<equiv> (\<Sum>j = 0..p. (smult (b j) (Bernstein_Poly j p c d)))"
+ assumes "i \<le> p" and "c \<noteq> d"
+ shows "coeff ((reciprocal_poly p (P \<circ>\<^sub>p [:c, 1:]
+ \<circ>\<^sub>p [:0, d-c:])) \<circ>\<^sub>p [:1, 1:]) (p - i) = (p choose i) * (b i)"
+proof -
+ have h: "P \<circ>\<^sub>p [:c, 1:] \<circ>\<^sub>p [:0, d-c:]
+ = (\<Sum>j = 0..p. (smult (b j) (Bernstein_Poly_01 j p)))"
+ using assms
+ by (auto simp: P_def pcompose_sum pcompose_smult
+ pcompose_add Bernstein_Poly_rescale_01)
+ then show ?thesis
+ using coeff_Bernstein_sum_01 assms by simp
+qed
+
+lemma Bernstein_sum:
+ assumes "c \<noteq> d" and "degree P \<le> p"
+ shows "P = (\<Sum>j = 0..p. smult (inverse (real (p choose j))
+ * coeff (reciprocal_poly p (P \<circ>\<^sub>p [:c, 1:] \<circ>\<^sub>p [:0, d-c:])
+ \<circ>\<^sub>p [:1, 1:]) (p-j)) (Bernstein_Poly j p c d))"
+ apply (subst Bernstein_Poly_eq_rescale_01)
+ subgoal using assms by blast
+ subgoal
+ apply (subst pcompose_smult[symmetric])
+ apply (subst pcompose_sum[symmetric])
+ apply (subst pcompose_smult[symmetric])
+ apply (subst pcompose_sum[symmetric])
+ apply (subst Bernstein_sum_01[symmetric])
+ using assms by (auto simp: degree_pcompose pcompose_assoc[symmetric]
+ pcompose_pCons)
+ done
+
+lemma Bernstein_Poly_span1:
+ assumes "c \<noteq> d" and "degree P \<le> p"
+ shows "P \<in> poly_vs.span {Bernstein_Poly x p c d | x. x \<le> p}"
+proof (subst Bernstein_sum[OF assms], rule poly_vs.span_sum)
+ fix x :: nat
+ assume "x \<in> {0..p}"
+ then have "\<exists>n. Bernstein_Poly x p c d = Bernstein_Poly n p c d \<and> n \<le> p"
+ by auto
+ then have
+ "Bernstein_Poly x p c d \<in> poly_vs.span {Bernstein_Poly n p c d |n. n \<le> p}"
+ by (simp add: poly_vs.span_base)
+ thus "smult (inverse (real (p choose x)) *
+ coeff (reciprocal_poly p (P \<circ>\<^sub>p [:c, 1:] \<circ>\<^sub>p [:0, d - c:]) \<circ>\<^sub>p [:1, 1:])
+ (p - x)) (Bernstein_Poly x p c d)
+ \<in> poly_vs.span {Bernstein_Poly x p c d |x. x \<le> p}"
+ by (rule poly_vs.span_scale)
+qed
+
+lemma Bernstein_Poly_span:
+ assumes "c \<noteq> d"
+ shows "poly_vs.span {Bernstein_Poly x p c d | x. x \<le> p} = {x. degree x \<le> p}"
+proof (subst Bernstein_Poly_01_span[symmetric], subst poly_vs.span_eq, rule conjI)
+ show "{Bernstein_Poly x p c d |x. x \<le> p}
+ \<subseteq> poly_vs.span {Bernstein_Poly_01 x p |x. x \<le> p}"
+ apply (subst Setcompr_subset)
+ apply (rule allI, rule impI)
+ apply (rule Bernstein_Poly_01_span1)
+ using assms by (auto simp: degree_Bernstein_le Bernstein_Poly_eq_rescale_01
+ degree_pcompose)
+
+ show "{Bernstein_Poly_01 x p |x. x \<le> p}
+ \<subseteq> poly_vs.span {Bernstein_Poly x p c d |x. x \<le> p}"
+ apply (subst Setcompr_subset)
+ apply (rule allI, rule impI)
+ apply (rule Bernstein_Poly_span1)
+ using assms by (auto simp: degree_Bernstein_le)
+qed
+
+lemma Bernstein_Poly_independent: assumes "c \<noteq> d"
+ shows "poly_vs.independent {Bernstein_Poly x p c d | x. x \<in> {..p}}"
+proof (rule poly_vs.card_le_dim_spanning)
+ show "{Bernstein_Poly x p c d |x. x \<in> {.. p}} \<subseteq> {x. degree x \<le> p}"
+ using assms
+ by (auto simp: degree_Bernstein Bernstein_Poly_eq_rescale_01 degree_pcompose)
+ show "{x. degree x \<le> p} \<subseteq> poly_vs.span {Bernstein_Poly x p c d |x. x \<in> {..p}}"
+ using assms by (auto simp: Bernstein_Poly_span1)
+ show "finite {Bernstein_Poly x p c d |x. x \<in> {..p}}" by fastforce
+ show "card {Bernstein_Poly x p c d |x. x \<in> {..p}} \<le> poly_vs.dim {x. degree x \<le> p}"
+ apply (rule le_trans)
+ apply (subst image_Collect[symmetric], rule card_image_le, force)
+ by (force simp: dim_degree)
+qed
+
+subsection \<open>Bernstein coefficients and changes over any interval\<close>
+
+definition Bernstein_coeffs ::
+ "nat \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real poly \<Rightarrow> real list" where
+ "Bernstein_coeffs p c d P =
+ [(inverse (real (p choose j)) *
+ coeff (reciprocal_poly p (P \<circ>\<^sub>p [:c, 1:] \<circ>\<^sub>p [:0, d-c:]) \<circ>\<^sub>p [:1, 1:]) (p-j)).
+ j \<leftarrow> [0..<(p+1)]]"
+
+lemma Bernstein_coeffs_eq_rescale: assumes "c \<noteq> d"
+ shows "Bernstein_coeffs p c d P = Bernstein_coeffs_01 p (P \<circ>\<^sub>p [:c, 1:] \<circ>\<^sub>p [:0, d-c:])"
+ using assms by (auto simp: pcompose_pCons pcompose_assoc[symmetric]
+ Bernstein_coeffs_def Bernstein_coeffs_01_def)
+
+lemma nth_default_Bernstein_coeffs: assumes "degree P \<le> p"
+ shows "nth_default 0 (Bernstein_coeffs p c d P) i =
+ inverse (p choose i) * coeff
+ (reciprocal_poly p (P \<circ>\<^sub>p [:c, 1:] \<circ>\<^sub>p [:0, d-c:]) \<circ>\<^sub>p [:1, 1:]) (p-i)"
+ apply (cases "p = i")
+ using assms by (auto simp: Bernstein_coeffs_def nth_default_append
+ nth_default_Cons Nitpick.case_nat_unfold binomial_eq_0)
+
+lemma Bernstein_coeffs_sum: assumes "c \<noteq> d" and hP: "degree P \<le> p"
+ shows "P = (\<Sum>j = 0..p. smult (nth_default 0 (Bernstein_coeffs p c d P) j)
+ (Bernstein_Poly j p c d))"
+ apply (subst nth_default_Bernstein_coeffs[OF hP])
+ apply (subst Bernstein_sum[OF assms])
+ by argo
+
+definition Bernstein_changes :: "nat \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real poly \<Rightarrow> int" where
+ "Bernstein_changes p c d P = nat (changes (Bernstein_coeffs p c d P))"
+
+lemma Bernstein_changes_eq_rescale: assumes "c \<noteq> d" and "degree P \<le> p"
+ shows "Bernstein_changes p c d P =
+ Bernstein_changes_01 p (P \<circ>\<^sub>p [:c, 1:] \<circ>\<^sub>p [:0, d-c:])"
+ using assms by (auto simp: Bernstein_coeffs_eq_rescale Bernstein_changes_def
+ Bernstein_changes_01_def)
+
+text \<open>This is related and mostly equivalent to previous Descartes test \cite{li2019counting}\<close>
+lemma Bernstein_changes_test:
+ fixes P::"real poly"
+ assumes "degree P \<le> p" and "P \<noteq> 0" and "c < d"
+ shows "proots_count P {x. c < x \<and> x < d} \<le> Bernstein_changes p c d P \<and>
+ even (Bernstein_changes p c d P - proots_count P {x. c < x \<and> x < d})"
+proof -
+ define Q where "Q=P \<circ>\<^sub>p [:c, 1:] \<circ>\<^sub>p [:0, d - c:]"
+
+ have "int (proots_count Q {x. 0 < x \<and> x < 1})
+ \<le> Bernstein_changes_01 p Q \<and>
+ even (Bernstein_changes_01 p Q -
+ int (proots_count Q {x. 0 < x \<and> x < 1}))"
+ unfolding Q_def
+ apply (rule Bernstein_changes_01_test)
+ subgoal using assms by fastforce
+ subgoal using assms by (auto simp: pcompose_eq_0)
+ done
+ moreover have "proots_count P {x. c < x \<and> x < d} =
+ proots_count Q {x. 0 < x \<and> x < 1}"
+ unfolding Q_def
+ proof (subst proots_pcompose)
+ have "poly [:c, 1:] ` poly [:0, d - c:] ` {x. 0 < x \<and> x < 1} =
+ {x. c < x \<and> x < d}" (is "?L = ?R")
+ proof
+ have "c + x * (d - c) < d" if "x < 1" for x
+ proof -
+ have "x * (d - c) < 1 * (d - c)"
+ using \<open>c < d\<close> that by force
+ then show ?thesis by fastforce
+ qed
+ then show "?L \<subseteq> ?R"
+ using assms by auto
+ next
+ show "?R \<subseteq> ?L"
+ proof
+ fix x::real assume "x \<in> ?R"
+ hence "c < x" and "x < d" by auto
+ thus "x \<in> ?L"
+ proof (subst image_eqI)
+ show "x = poly [:c, 1:] (x - c)" by force
+ assume "c < x" and "x < d"
+ thus "x - c \<in> poly [:0, d - c:] ` {x. 0 < x \<and> x < 1}"
+ proof (subst image_eqI)
+ show "x - c = poly [:0, d - c:] ((x - c)/(d - c))"
+ using assms by fastforce
+ assume "c < x" and "x < d"
+ thus "(x - c) / (d - c) \<in> {x. 0 < x \<and> x < 1}"
+ by auto
+ qed fast
+ qed fast
+ qed
+ qed
+ then show "proots_count P {x. c < x \<and> x < d} =
+ proots_count (P \<circ>\<^sub>p [:c, 1:])
+ (poly [:0, d - c:] ` {x. 0 < x \<and> x < 1})"
+ using assms by (auto simp:proots_pcompose)
+ show "P \<circ>\<^sub>p [:c, 1:] \<noteq> 0"
+ by (simp add: pcompose_eq_0 assms(2))
+ show "degree [:0, d - c:] = 1"
+ using assms by auto
+ qed
+ moreover have " Bernstein_changes p c d P = Bernstein_changes_01 p Q"
+ unfolding Q_def
+ apply (rule Bernstein_changes_eq_rescale)
+ using assms by auto
+ ultimately show ?thesis by auto
+qed
+
+subsection \<open>The control polygon of a polynomial\<close>
+
+definition control_points ::
+ "nat \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real poly \<Rightarrow> (real \<times> real) list"
+where
+ "control_points p c d P =
+ [(((real i)*d + (real (p - i))*c)/p,
+ nth_default 0 (Bernstein_coeffs p c d P) i).
+ i \<leftarrow> [0..<(p+1)]]"
+
+lemma line_above:
+ fixes a b c d :: real and p :: nat and P :: "real poly"
+ assumes hline: "\<And>i. i \<le> p \<Longrightarrow> a * (((real i)*d + (real (p - i))*c)/p) + b \<ge>
+ nth_default 0 (Bernstein_coeffs p c d P) i"
+ and hp: "p \<noteq> 0" and hcd: "c \<noteq> d" and hP: "degree P \<le> p"
+ shows "\<And>x. c \<le> x \<Longrightarrow> x \<le> d \<Longrightarrow> a*x + b \<ge> poly P x"
+proof -
+ fix x
+ assume hc: "c \<le> x" and hd: "x \<le> d"
+
+ have bern_eq:"Bernstein_coeffs p c d [:b, a:] =
+ [a*(real i * d + real (p - i) * c)/p + b. i \<leftarrow> [0..<(p+1)]]"
+ proof -
+ have "Bernstein_coeffs p c d [:b, a:] = map (nth_default 0
+ (Bernstein_coeffs_01 p ([:b, a:] \<circ>\<^sub>p [:c, 1:] \<circ>\<^sub>p [:0, d - c:])))
+ [0..<p+1]"
+ apply (subst Bernstein_coeffs_eq_rescale["OF" hcd])
+ apply (subst map_nth_default[symmetric])
+ apply (subst length_Bernstein_coeffs_01)
+ by blast
+ also have
+ "... = map (\<lambda>i. a * (real i * d + real (p - i) * c) / real p + b) [0..<p + 1]"
+ proof (rule map_cong)
+ fix x assume hx: "x \<in> set [0..<p + 1]"
+ have "nth_default 0 (Bernstein_coeffs_01 p
+ ([:b, a:] \<circ>\<^sub>p [:c, 1:] \<circ>\<^sub>p [:0, d - c:])) x =
+ nth_default 0 (Bernstein_coeffs_01 p
+ (smult (b + a*c) 1 + smult (a*(d - c)) (monom 1 1))) x"
+ proof-
+ have "[:b, a:] \<circ>\<^sub>p [:c, 1:] \<circ>\<^sub>p [:0, d - c:] =
+ smult (b + a*c) 1 + smult (a*(d - c)) (monom 1 1)"
+ by (simp add: monom_altdef pcompose_pCons)
+ then show ?thesis by auto
+ qed
+ also have "... =
+ nth_default 0 (Bernstein_coeffs_01 p (smult (b + a * c) 1)) x +
+ nth_default 0 (Bernstein_coeffs_01 p (smult (a * (d - c)) (monom 1 1))) x"
+ apply (subst Bernstein_coeffs_01_add)
+ using hp by (auto simp: degree_monom_eq)
+ also have "... =
+ (b + a*c) * nth_default 0 (Bernstein_coeffs_01 p 1) x +
+ (a*(d - c)) * nth_default 0 (Bernstein_coeffs_01 p (monom 1 1)) x"
+ apply (subst Bernstein_coeffs_01_smult)
+ using hp by (auto simp: Bernstein_coeffs_01_smult degree_monom_eq)
+ also have "... =
+ (b + a * c) * (if x < p + 1 then 1 else 0) +
+ a * (d - c) * (real (nth_default 0 [0..<p + 1] x) / real p)"
+ apply (subst Bernstein_coeffs_01_1, subst Bernstein_coeffs_01_x[OF hp])
+ apply (subst nth_default_replicate_eq, subst nth_default_map_eq[of _ 0])
+ by auto
+ also have "... =
+ (b + a * c) * (if x < p + 1 then 1 else 0) +
+ a * (d - c) * (real ([0..<p + 1] ! x) / real p)"
+ apply (subst nth_default_nth)
+ using hx by auto
+ also have "... = (b + a * c) * (if x < p + 1 then 1 else 0) +
+ a * (d - c) * (real (0 + x) / real p)"
+ apply (subst nth_upt)
+ using hx by auto
+ also have "... = a * (real x * d + real (p - x) * c) / real p + b"
+ apply (subst of_nat_diff)
+ using hx hp by (auto simp: field_simps)
+ finally show "nth_default 0 (Bernstein_coeffs_01 p
+ ([:b, a:] \<circ>\<^sub>p [:c, 1:] \<circ>\<^sub>p [:0, d - c:])) x =
+ a * (real x * d + real (p - x) * c) / real p + b" .
+ qed blast
+ finally show ?thesis .
+ qed
+
+ have nth_default_geq:"nth_default 0 (Bernstein_coeffs p c d [:b, a:]) i \<ge>
+ nth_default 0 (Bernstein_coeffs p c d P) i" for i
+ proof -
+ show "nth_default 0 (Bernstein_coeffs p c d [:b, a:]) i \<ge>
+ nth_default 0 (Bernstein_coeffs p c d P) i"
+ proof cases
+ define p1 where "p1 \<equiv> p+1"
+ assume h: "i \<le> p"
+ hence "nth_default 0 (Bernstein_coeffs p c d P) i \<le>
+ a * (((real i)*d + (real (p - i))*c)/p) + b"
+ by (rule hline)
+ also have "... = nth_default 0 (map (\<lambda>i. a * (real i * d
+ + real (p - i) * c) / real p + b) [0..<p + 1]) i"
+ apply (subst p1_def[symmetric])
+ using h apply (auto simp: nth_default_def)
+ by (auto simp: p1_def)
+ also have "... = nth_default 0 (Bernstein_coeffs p c d [:b, a:]) i"
+ using bern_eq by simp
+ finally show ?thesis .
+ next
+ assume h: "\<not>i \<le> p"
+ thus ?thesis
+ using assms
+ by (auto simp: nth_default_def Bernstein_coeffs_eq_rescale
+ length_Bernstein_coeffs_01)
+ qed
+ qed
+
+ have "poly P x = (\<Sum>k = 0..p.
+ poly (smult (nth_default 0 (Bernstein_coeffs p c d P) k)
+ (Bernstein_Poly k p c d)) x)"
+ apply (subst Bernstein_coeffs_sum[OF hcd hP])
+ by (rule poly_sum)
+ also have "... \<le> (\<Sum>k = 0..p.
+ poly (smult (nth_default 0 (Bernstein_coeffs p c d [:b, a:]) k)
+ (Bernstein_Poly k p c d)) x)"
+ apply (rule sum_mono)
+ using mult_right_mono[OF nth_default_geq] Bernstein_Poly_nonneg[OF hc hd]
+ by auto
+ also have "... = poly [:b, a:] x"
+ apply (subst(2) Bernstein_coeffs_sum[of c d "[:b, a:]" p])
+ using assms apply auto[2]
+ by (rule poly_sum[symmetric])
+ also have "... = a*x + b" by force
+ finally show "poly P x \<le> a*x + b" .
+qed
+
end
\ No newline at end of file
diff --git a/thys/Three_Circles/Bernstein_01.thy b/thys/Three_Circles/Bernstein_01.thy
--- a/thys/Three_Circles/Bernstein_01.thy
+++ b/thys/Three_Circles/Bernstein_01.thy
@@ -1,619 +1,619 @@
-section \<open>Bernstein Polynomials over the interval [0, 1]\<close>
-
-theory Bernstein_01
- imports "HOL-Computational_Algebra.Computational_Algebra"
- "Budan_Fourier.Budan_Fourier"
- "RRI_Misc"
-begin
-
-text \<open>
-The theorem of three circles is a statement about the Bernstein coefficients of a polynomial, the
-coefficients when a polynomial is expressed as a sum of Bernstein polynomials. These coefficients
-behave nicely under translations and rescaling and are the coefficients of a particular polynomial
-in the [0, 1] case. We shall define the [0, 1] case now and consider the general case later,
-deriving all the results by rescaling.
-\<close>
-
-subsection \<open>Definition and basic results\<close>
-
-definition Bernstein_Poly_01 :: "nat \<Rightarrow> nat \<Rightarrow> real poly" where
- "Bernstein_Poly_01 j p = (monom (p choose j) j)
- * (monom 1 (p-j) \<circ>\<^sub>p [:1, -1:])"
-
-lemma degree_Bernstein:
- assumes hb: "j \<le> p"
- shows "degree (Bernstein_Poly_01 j p) = p"
-proof -
- have ha: "monom (p choose j) j \<noteq> (0::real poly)" using hb by force
- have hb: "monom 1 (p-j) \<circ>\<^sub>p [:1, -1:] \<noteq> (0::real poly)"
- proof
- assume "monom 1 (p-j) \<circ>\<^sub>p [:1, -1:] = (0::real poly)"
- hence "lead_coeff (monom 1 (p - j) \<circ>\<^sub>p [:1, -1:]) = (0::real)"
- apply (subst leading_coeff_0_iff)
- by simp
- moreover have "lead_coeff (monom (1::real) (p - j)
- \<circ>\<^sub>p [:1, -1:]) = (((- 1) ^ (p - j))::real)"
- by (subst lead_coeff_comp, auto simp: degree_monom_eq)
- ultimately show "False" by auto
- qed
- from ha hb show ?thesis
- by (auto simp add: Bernstein_Poly_01_def degree_mult_eq
- degree_monom_eq degree_pcompose)
-qed
-
-lemma coeff_gt:
- assumes hb: "j > p"
- shows "Bernstein_Poly_01 j p = 0"
- by (simp add: hb Bernstein_Poly_01_def)
-
-lemma degree_Bernstein_le: "degree (Bernstein_Poly_01 j p) \<le> p"
- apply (cases "j \<le> p")
- by (simp_all add: degree_Bernstein coeff_gt)
-
-lemma poly_Bernstein_nonneg:
- assumes "x \<ge> 0" and "1 \<ge> x"
- shows "poly (Bernstein_Poly_01 j p) x \<ge> 0"
- using assms by (simp add: poly_monom poly_pcompose Bernstein_Poly_01_def)
-
-lemma Bernstein_symmetry:
- assumes "j \<le> p"
- shows "(Bernstein_Poly_01 j p) \<circ>\<^sub>p [:1, -1:] = Bernstein_Poly_01 (p-j) p"
-proof -
- have "(Bernstein_Poly_01 j p) \<circ>\<^sub>p [:1, -1:]
- = ((monom (p choose j) j) * (monom 1 (p-j) \<circ>\<^sub>p [:1, -1:])) \<circ>\<^sub>p [:1, -1:]"
- by (simp add: Bernstein_Poly_01_def)
- also have "... = (monom (p choose (p-j)) j *
- (monom 1 (p-j) \<circ>\<^sub>p [:1, -1:])) \<circ>\<^sub>p [:1, -1:]"
- by (fastforce simp: binomial_symmetric[OF assms])
- also have "... = monom (p choose (p-j)) j \<circ>\<^sub>p [:1, -1:] *
- (monom 1 (p-j)) \<circ>\<^sub>p ([:1, -1:] \<circ>\<^sub>p [:1, -1:])"
- by (force simp: pcompose_mult pcompose_assoc)
- also have "... = (monom (p choose (p-j)) j \<circ>\<^sub>p [:1, -1:]) * monom 1 (p-j)"
- by (force simp: pcompose_pCons)
- also have "... = smult (p choose (p-j)) (monom 1 j \<circ>\<^sub>p [:1, -1:])
- * monom 1 (p-j)"
- by (simp add: assms smult_monom pcompose_smult[symmetric])
- also have "... = (monom 1 j \<circ>\<^sub>p [:1, -1:]) * monom (p choose (p-j)) (p-j)"
- apply (subst mult_smult_left)
- apply (subst mult_smult_right[symmetric])
- apply (subst smult_monom)
- by force
- also have "... = Bernstein_Poly_01 (p-j) p" using assms
- by (auto simp: Bernstein_Poly_01_def)
- finally show ?thesis .
-qed
-
-subsection \<open>@{term Bernstein_Poly_01} and @{term reciprocal_poly}\<close>
-
-lemma Bernstein_reciprocal:
- "reciprocal_poly p (Bernstein_Poly_01 i p)
- = smult (p choose i) ([:-1, 1:]^(p-i))"
-proof cases
- assume "i \<le> p"
- hence "reciprocal_poly p (Bernstein_Poly_01 i p) =
- reciprocal_poly (degree (Bernstein_Poly_01 i p)) (Bernstein_Poly_01 i p)"
- by (auto simp: degree_Bernstein)
- also have "... = reflect_poly (Bernstein_Poly_01 i p)"
- by (rule reciprocal_degree)
- also have "... = smult (p choose i) ([:-1, 1:]^(p-i))"
- by (auto simp: Bernstein_Poly_01_def reflect_poly_simps monom_altdef
- pcompose_pCons reflect_poly_pCons' hom_distribs)
- finally show ?thesis .
-next
- assume h:"\<not> i \<le> p"
- hence "reciprocal_poly p (Bernstein_Poly_01 i p) = (0::real poly)"
- by (auto simp: coeff_gt reciprocal_poly_def)
- also have "... = smult (p choose i) ([:-1, 1:]^(p - i))" using h
- by fastforce
- finally show ?thesis .
-qed
-
-lemma Bernstein_reciprocal_translate:
- "reciprocal_poly p (Bernstein_Poly_01 i p) \<circ>\<^sub>p [:1, 1:] =
- monom (p choose i) (p - i)"
- by (auto simp: Bernstein_reciprocal pcompose_smult pcompose_pCons monom_altdef hom_distribs)
-
-lemma coeff_Bernstein_sum_01: fixes b::"nat \<Rightarrow> real" assumes hi: "p \<ge> i"
- shows
- "coeff (reciprocal_poly p
- (\<Sum>x = 0..p. smult (b x) (Bernstein_Poly_01 x p)) \<circ>\<^sub>p [:1, 1:])
- (p - i) = (p choose i) * (b i)" (is "?L = ?R")
-proof -
- define P where "P \<equiv> (\<Sum>x = 0..p. (smult (b x) (Bernstein_Poly_01 x p)))"
-
- have "\<And>x. degree (smult (b x) (Bernstein_Poly_01 x p)) \<le> p"
- proof -
- fix x
- show "degree (smult (b x) (Bernstein_Poly_01 x p)) \<le> p"
- apply (cases "x \<le> p")
- by (auto simp: degree_Bernstein coeff_gt)
- qed
- hence "reciprocal_poly p P =
- (\<Sum>x = 0..p. reciprocal_poly p (smult (b x) (Bernstein_Poly_01 x p)))"
- apply (subst P_def)
- apply (rule reciprocal_sum)
- by presburger
- also have
- "... = (\<Sum>x = 0..p. (smult (b x * (p choose x)) ([:-1, 1:]^(p-x))))"
- proof (rule sum.cong)
- fix x assume "x \<in> {0..p}"
- hence "x \<le> p" by simp
- thus "reciprocal_poly p (smult (b x) (Bernstein_Poly_01 x p)) =
- smult ((b x) * (p choose x)) ([:-1, 1:]^(p-x))"
- by (auto simp add: reciprocal_smult degree_Bernstein Bernstein_reciprocal)
- qed (simp)
- finally have
- "reciprocal_poly p P =
- (\<Sum>x = 0..p. (smult ((b x) * (p choose x)) ([:-1, 1:]^(p-x))))" .
- hence
- "(reciprocal_poly p P) \<circ>\<^sub>p [:1, 1:] =
- (\<Sum>x = 0..p. (smult ((b x) * (p choose x)) ([:-1, 1:]^(p-x))) \<circ>\<^sub>p [:1, 1:])"
- by (simp add: pcompose_sum pcompose_add)
- also have "... = (\<Sum>x = 0..p. (monom ((b x) * (p choose x)) (p - x)))"
- proof (rule sum.cong)
- fix x assume "x \<in> {0..p}"
- hence "x \<le> p" by simp
- thus "smult (b x * (p choose x)) ([:- 1, 1:] ^ (p - x)) \<circ>\<^sub>p [:1, 1:] =
- monom (b x * (p choose x)) (p - x)"
- by (simp add: hom_distribs pcompose_smult pcompose_pCons monom_altdef)
- qed (simp)
- finally have "(reciprocal_poly p P) \<circ>\<^sub>p [:1, 1:] =
- (\<Sum>x = 0..p. (monom ((b x) * (p choose x)) (p - x)))" .
- hence "?L = (\<Sum>x = 0..p. if p - x = p - i then b x * real (p choose x) else 0)"
- by (auto simp add: P_def coeff_sum)
- also have "... = (\<Sum>x = 0..p. if x = i then b x * real (p choose x) else 0)"
- proof (rule sum.cong)
- fix x assume "x \<in> {0..p}"
- hence "x \<le> p" by simp
- thus "(if p - x = p - i then b x * real (p choose x) else 0) =
- (if x = i then b x * real (p choose x) else 0)" using hi
- by (auto simp add: leI)
- qed (simp)
- also have "... = ?R" by simp
- finally show ?thesis .
-qed
-
-lemma Bernstein_sum_01: assumes hP: "degree P \<le> p"
- shows
- "P = (\<Sum>j = 0..p. smult
- (inverse (real (p choose j)) *
- coeff (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:]) (p-j))
- (Bernstein_Poly_01 j p))"
-proof -
- define Q where "Q \<equiv> reciprocal_poly p P \<circ>\<^sub>p [:1, 1:]"
- from hP Q_def have hQ: "degree Q \<le> p"
- by (auto simp: degree_reciprocal degree_pcompose)
- have "reciprocal_poly p (\<Sum>j = 0..p.
- smult (inverse (real (p choose j)) * coeff Q (p-j))
- (Bernstein_Poly_01 j p)) \<circ>\<^sub>p [:1, 1:] = Q"
- proof (rule poly_eqI)
- fix n
- show "coeff (reciprocal_poly p (\<Sum>j = 0..p.
- smult (inverse (real (p choose j)) * coeff Q (p-j))
- (Bernstein_Poly_01 j p)) \<circ>\<^sub>p [:1, 1:]) n = coeff Q n"
- (is "?L = ?R")
- proof cases
- assume hn: "n \<le> p"
- hence "?L = coeff (reciprocal_poly p (\<Sum>j = 0..p.
- smult (inverse (real (p choose j)) * coeff Q (p-j))
- (Bernstein_Poly_01 j p)) \<circ>\<^sub>p [:1, 1:]) (p - (p - n))"
- by force
- also have "... = (p choose (p-n)) *
- (inverse (real (p choose (p-n))) *
- coeff Q (p-(p-n)))"
- apply (subst coeff_Bernstein_sum_01)
- by auto
- also have "... = ?R" using hn
- by fastforce
- finally show "?L = ?R" .
- next
- assume hn: "\<not> n \<le> p"
- have "degree (\<Sum>j = 0..p.
- smult (inverse (real (p choose j)) * coeff Q (p - j))
- (Bernstein_Poly_01 j p)) \<le> p"
- proof (rule degree_sum_le)
- fix q assume "q \<in> {0..p}"
- hence "q \<le> p" by fastforce
- thus "degree (smult (inverse (real (p choose q)) *
- coeff Q (p - q)) (Bernstein_Poly_01 q p)) \<le> p"
- by (auto simp add: degree_Bernstein degree_smult_le)
- qed simp
- hence "degree (reciprocal_poly p (\<Sum>j = 0..p.
- smult (inverse (real (p choose j)) * coeff Q (p - j))
- (Bernstein_Poly_01 j p)) \<circ>\<^sub>p [:1, 1:]) \<le> p"
- by (auto simp add: degree_pcompose degree_reciprocal)
- hence "?L = 0" using hn by (auto simp add: coeff_eq_0)
- thus "?L = ?R" using hQ hn by (simp add: coeff_eq_0)
- qed
- qed
- hence "reciprocal_poly p P \<circ>\<^sub>p [:1, 1:] =
- reciprocal_poly p (\<Sum>j = 0..p.
- smult (inverse (real (p choose j)) *
- coeff (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:]) (p-j))
- (Bernstein_Poly_01 j p)) \<circ>\<^sub>p [:1, 1:]"
- by (auto simp: degree_reciprocal degree_pcompose Q_def)
- hence "reciprocal_poly p P \<circ>\<^sub>p ([:1, 1:] \<circ>\<^sub>p [:-1, 1:]) =
- reciprocal_poly p (\<Sum>j = 0..p. smult (inverse (real (p choose j)) *
- coeff (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:]) (p-j))
- (Bernstein_Poly_01 j p)) \<circ>\<^sub>p ([:1, 1:] \<circ>\<^sub>p [:-1, 1:])"
- by (auto simp: pcompose_assoc)
- hence "reciprocal_poly p P = reciprocal_poly p (\<Sum>j = 0..p.
- smult (inverse (real (p choose j)) *
- coeff (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:]) (p-j)) (Bernstein_Poly_01 j p))"
- by (auto simp: pcompose_pCons)
- hence "reciprocal_poly p (reciprocal_poly p P) =
- reciprocal_poly p (reciprocal_poly p (\<Sum>j = 0..p.
- smult (inverse (real (p choose j)) *
- coeff (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:]) (p-j)) (Bernstein_Poly_01 j p)))"
- by argo
- thus "P = (\<Sum>j = 0..p. smult (inverse (real (p choose j)) *
- coeff (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:]) (p-j)) (Bernstein_Poly_01 j p))"
- using hP by (auto simp: reciprocal_reciprocal degree_sum_le degree_smult_le
- degree_Bernstein degree_add_le)
-qed
-
-lemma Bernstein_Poly_01_span1:
- assumes hP: "degree P \<le> p"
- shows "P \<in> poly_vs.span {Bernstein_Poly_01 x p | x. x \<le> p}"
-proof -
- have "Bernstein_Poly_01 x p
- \<in> poly_vs.span {Bernstein_Poly_01 x p |x. x \<le> p}"
- if "x \<in> {0..p}" for x
- proof -
- have "\<exists>n. Bernstein_Poly_01 x p = Bernstein_Poly_01 n p \<and> n \<le> p"
- using that by force
- then show
- "Bernstein_Poly_01 x p \<in> poly_vs.span {Bernstein_Poly_01 n p |n. n \<le> p}"
- by (simp add: poly_vs.span_base)
- qed
- thus ?thesis
- apply (subst Bernstein_sum_01[OF hP])
- apply (rule poly_vs.span_sum)
- apply (rule poly_vs.span_scale)
- by blast
-qed
-
-lemma Bernstein_Poly_01_span:
- "poly_vs.span {Bernstein_Poly_01 x p | x. x \<le> p}
- = {x. degree x \<le> p}"
- apply (subst monom_span[symmetric])
- apply (subst poly_vs.span_eq)
- by (auto simp: monom_span degree_Bernstein_le
- Bernstein_Poly_01_span1 degree_monom_eq)
-
-subsection \<open>Bernstein coefficients and changes\<close>
-
-definition Bernstein_coeffs_01 :: "nat \<Rightarrow> real poly \<Rightarrow> real list" where
- "Bernstein_coeffs_01 p P =
- [(inverse (real (p choose j)) *
- coeff (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:]) (p-j)). j \<leftarrow> [0..<(p+1)]]"
-
-lemma length_Bernstein_coeffs_01: "length (Bernstein_coeffs_01 p P) = p + 1"
- by (auto simp: Bernstein_coeffs_01_def)
-
-lemma nth_default_Bernstein_coeffs_01: assumes "degree P \<le> p"
- shows "nth_default 0 (Bernstein_coeffs_01 p P) i =
- inverse (p choose i) * coeff (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:]) (p-i)"
- apply (cases "p = i")
- using assms by (auto simp: Bernstein_coeffs_01_def nth_default_append
- nth_default_Cons Nitpick.case_nat_unfold binomial_eq_0)
-
-lemma Bernstein_coeffs_01_sum: assumes "degree P \<le> p"
- shows "P = (\<Sum>j = 0..p. smult (nth_default 0 (Bernstein_coeffs_01 p P) j)
- (Bernstein_Poly_01 j p))"
- apply (subst nth_default_Bernstein_coeffs_01[OF assms])
- apply (subst Bernstein_sum_01[OF assms])
- by argo
-
-definition Bernstein_changes_01 :: "nat \<Rightarrow> real poly \<Rightarrow> int" where
- "Bernstein_changes_01 p P = nat (changes (Bernstein_coeffs_01 p P))"
-
-lemma Bernstein_changes_01_def':
- "Bernstein_changes_01 p P = nat (changes [(inverse (real (p choose j)) *
- coeff (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:]) (p-j)). j \<leftarrow> [0..<p + 1]])"
- by (simp add: Bernstein_changes_01_def Bernstein_coeffs_01_def)
-
-lemma Bernstein_changes_01_eq_changes:
- assumes hP: "degree P \<le> p"
- shows "Bernstein_changes_01 p P =
- changes (coeffs ((reciprocal_poly p P) \<circ>\<^sub>p [:1, 1:]))"
-proof (subst Bernstein_changes_01_def')
- have h:
- "map (\<lambda>j. inverse (real (p choose j)) *
- coeff (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:]) (p - j)) [0..<p + 1] =
- map (\<lambda>j. inverse (real (p choose j)) *
- nth_default 0 [nth_default 0 (coeffs (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:]))
- (p - j). j \<leftarrow> [0..<p + 1]] j) [0..<p + 1]"
- proof (rule map_cong)
- fix x
- assume "x \<in> set [0..<p+1]"
- hence hx: "x \<le> p" by fastforce
- moreover have 1:
- "length (map (\<lambda>j. nth_default 0
- (coeffs (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:])) (p - j)) [0..<p + 1]) \<le> Suc p"
- by force
- moreover have "length (coeffs (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:])) \<le> Suc p"
- proof (cases "P=0")
- case False
- then have "reciprocal_poly p P \<circ>\<^sub>p [:1, 1:] \<noteq> 0"
- using hP by (simp add: Missing_Polynomial.pcompose_eq_0 reciprocal_0_iff)
- moreover have "Suc (degree (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:])) \<le> Suc p"
- using hP by (auto simp: degree_pcompose degree_reciprocal)
- ultimately show ?thesis
- using length_coeffs_degree by force
- qed (auto simp: reciprocal_0)
- ultimately have h:
- "nth_default 0 (map (\<lambda>j. nth_default 0 (coeffs
- (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:])) (p - j)) [0..<p + 1]) x =
- nth_default 0 (coeffs (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:])) (p - x)"
- (is "?L = ?R")
- proof -
- have "?L = (map (\<lambda>j. nth_default 0 (coeffs
- (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:])) (p - j)) [0..<p + 1]) ! x"
- using hx by (auto simp: nth_default_nth)
- also have "... = nth_default 0
- (coeffs (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:])) (p - [0..<p + 1] ! x)"
- apply (subst nth_map)
- using hx by auto
- also have "... = ?R"
- apply (subst nth_upt)
- using hx by auto
- finally show ?thesis .
- qed
- show "inverse (real (p choose x)) *
- coeff (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:]) (p - x) =
- inverse (real (p choose x)) *
- nth_default 0 (map (\<lambda>j. nth_default 0
- (coeffs (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:])) (p - j)) [0..<p + 1]) x"
- apply (subst h)
- apply (subst nth_default_coeffs_eq)
- by blast
- qed auto
-
- have 1:
- "rev (map (\<lambda>j. nth_default 0 (coeffs (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:]))
- (p - j)) [0..<p + 1]) = map (\<lambda>j. nth_default 0 (coeffs
- (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:])) j) [0..<p + 1]"
- proof (subst rev_map, rule map_cong')
- have "\<And>q. (q \<ge> p \<longrightarrow> rev [q-p..<q+1] = map ((-) q) [0..<p+1])"
- proof (induction p)
- case 0
- then show ?case by simp
- next
- case (Suc p)
- have IH: "\<And>q. (q \<ge> p \<longrightarrow> rev [q-p..<q+1] = map ((-) q) [0..<p+1])"
- using Suc.IH by blast
- show ?case
- proof
- assume hq: "Suc p \<le> q"
- then have h: "rev [q - p..<q + 1] = map ((-) (q)) [0..<p + 1]"
- apply (subst IH)
- using hq by auto
- have "[q - Suc p..<q + 1] = (q - Suc p) # [q - p..<q + 1]"
- by (simp add: Suc_diff_Suc Suc_le_lessD hq upt_conv_Cons)
- hence "rev [q - Suc p..<q + 1] = rev [q - p..<q + 1] @ [q - Suc p]"
- by force
- also have "... = map ((-) (q)) [0..<p + 1] @ [q - Suc p]"
- using h by blast
- also have "... = map ((-) q) [0..<Suc p + 1]"
- by force
- finally show "rev [q - Suc p..<q + 1] = map ((-) q) [0..<Suc p + 1]" .
- qed
- qed
- thus "rev [0..<p + 1] = map ((-) p) [0..<p + 1]"
- by force
- next
- fix y
- assume "y \<in> set [0..<p + 1]"
- hence "y \<le> p" by fastforce
- thus "nth_default 0 (coeffs (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:])) (p - (p - y)) =
- nth_default 0 (coeffs (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:])) y"
- by fastforce
- qed
-
- have 2: "\<And> f. f \<noteq> 0 \<longrightarrow> degree f \<le> p \<longrightarrow>
- map (nth_default 0 (coeffs f)) [0..<p + 1] =
- coeffs f @ replicate (p - degree f) 0"
- proof (induction p)
- case 0
- then show ?case by (auto simp: degree_0_iff)
- next
- fix f
- case (Suc p)
- hence IH: "(f \<noteq> 0 \<longrightarrow>
- degree f \<le> p \<longrightarrow>
- map (nth_default 0 (coeffs f)) [0..<p + 1] =
- coeffs f @ replicate (p - degree f) 0)" by blast
- then show ?case
- proof (cases)
- assume h': "Suc p = degree f"
- hence h: "[0..<Suc p + 1] = [0..<length (coeffs f)]"
- by (metis add_is_0 degree_0 length_coeffs plus_1_eq_Suc zero_neq_one)
- thus ?thesis
- apply (subst h)
- apply (subst map_nth_default)
- using h' by fastforce
- next
- assume h': "Suc p \<noteq> degree f"
- show ?thesis
- proof
- assume hf: "f \<noteq> 0"
- show "degree f \<le> Suc p \<longrightarrow>
- map (nth_default 0 (coeffs f)) [0..<Suc p + 1] =
- coeffs f @ replicate (Suc p - degree f) 0"
- proof
- assume "degree f \<le> Suc p"
- hence 1: "degree f \<le> p" using h' by fastforce
- hence 2: "map (nth_default 0 (coeffs f)) [0..<p + 1] =
- coeffs f @ replicate (p - degree f) 0" using IH hf by blast
- have "map (nth_default 0 (coeffs f)) [0..<Suc p + 1] =
- map (nth_default 0 (coeffs f)) [0..<p + 1] @
- [nth_default 0 (coeffs f) (Suc p)]"
- by fastforce
- also have
- "... = coeffs f @ replicate (p - degree f) 0 @ [coeff f (Suc p)]"
- using 2
- by (auto simp: nth_default_coeffs_eq)
- also have "... = coeffs f @ replicate (p - degree f) 0 @ [0]"
- using \<open>degree f \<le> Suc p\<close> h' le_antisym le_degree by blast
- also have "... = coeffs f @ replicate (Suc p - degree f) 0" using 1
- by (simp add: Suc_diff_le replicate_app_Cons_same)
- finally show "map (nth_default 0 (coeffs f)) [0..<Suc p + 1] =
- coeffs f @ replicate (Suc p - degree f) 0" .
- qed
- qed
- qed
- qed
-
- thus "int (nat (changes (map (\<lambda>j. inverse (real (p choose j)) *
- coeff (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:]) (p - j)) [0..<p + 1]))) =
- changes (coeffs (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:]))"
- proof cases
- assume hP: "P = 0"
- show "int (nat (changes (map (\<lambda>j. inverse (real (p choose j)) *
- coeff (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:]) (p - j)) [0..<p + 1]))) =
- changes (coeffs (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:]))" (is "?L = ?R")
- proof -
- have "?L = int (nat (changes (map (\<lambda>j. 0::real) [0..<p+1])))"
- using hP by (auto simp: reciprocal_0 changes_nonneg)
- also have "... = 0"
- apply (induction p)
- by (auto simp: map_replicate_trivial changes_nonneg
- replicate_app_Cons_same)
- also have "0 = changes ([]::real list)" by simp
- also have "... = ?R" using hP by (auto simp: reciprocal_0)
- finally show ?thesis .
- qed
- next
- assume hP': "P \<noteq> 0"
- thus ?thesis
- apply (subst h)
- apply (subst changes_scale)
- apply auto[2]
- apply (subst changes_rev[symmetric])
- apply (subst 1)
- apply (subst 2)
- apply (simp add: pcompose_eq_0 hP reciprocal_0_iff)
- using assms apply (auto simp: degree_reciprocal)[1]
- by (auto simp: changes_append_replicate_0 changes_nonneg)
- qed
-qed
-
-lemma Bernstein_changes_01_test: fixes P::"real poly"
- assumes hP: "degree P \<le> p" and h0: "P \<noteq> 0"
- shows "proots_count P {x. 0 < x \<and> x < 1} \<le> Bernstein_changes_01 p P \<and>
- even (Bernstein_changes_01 p P - proots_count P {x. 0 < x \<and> x < 1})"
-proof -
- let ?Q = "(reciprocal_poly p P) \<circ>\<^sub>p [:1, 1:]"
-
- have 1: "changes (coeffs ?Q) \<ge> proots_count ?Q {x. 0 < x} \<and>
- even (changes (coeffs ?Q) - proots_count ?Q {x. 0 < x})"
- apply (rule descartes_sign)
- by (simp add: Missing_Polynomial.pcompose_eq_0 h0 hP reciprocal_0_iff)
-
- have "((+) (1::real) ` Collect ((<) (0::real))) = {x. (1::real)<x}"
- proof
- show "{x::real. 1 < x} \<subseteq> (+) 1 ` Collect ((<) 0)"
- proof
- fix x::real assume "x \<in> {x. 1 < x}"
- hence "1 < x" by simp
- hence "-1 + x \<in> Collect ((<) 0)" by auto
- hence "1 + (-1 + x) \<in> (+) 1 ` Collect ((<) 0)" by blast
- thus "x \<in> (+) 1 ` Collect ((<) 0)" by argo
- qed
- qed auto
- hence 2: "proots_count P {x. 0 < x \<and> x < 1} = proots_count ?Q {x. 0 < x}"
- using assms
- by (auto simp: proots_pcompose reciprocal_0_iff proots_count_reciprocal')
-
- show ?thesis
- apply (subst Bernstein_changes_01_eq_changes[OF hP])
- apply (subst Bernstein_changes_01_eq_changes[OF hP])
- apply (subst 2)
- apply (subst 2)
- by (rule 1)
-qed
-
-subsection \<open>Expression as a Bernstein sum\<close>
-
-lemma Bernstein_coeffs_01_0: "Bernstein_coeffs_01 p 0 = replicate (p+1) 0"
- by (auto simp: Bernstein_coeffs_01_def reciprocal_0 map_replicate_trivial
- replicate_append_same)
-
-lemma Bernstein_coeffs_01_1: "Bernstein_coeffs_01 p 1 = replicate (p+1) 1"
-proof -
- have "Bernstein_coeffs_01 p 1 =
- map (\<lambda>j. inverse (real (p choose j)) *
- coeff (\<Sum>k\<le>p. smult (real (p choose k)) ([:0, 1:] ^ k)) (p - j)) [0..<(p+1)]"
- by (auto simp: Bernstein_coeffs_01_def reciprocal_1 monom_altdef
- hom_distribs pcompose_pCons poly_0_coeff_0[symmetric] poly_binomial)
- also have "... = map (\<lambda>j. inverse (real (p choose j)) *
- real (p choose (p - j))) [0..<(p+1)]"
- by (auto simp: monom_altdef[symmetric] coeff_sum binomial)
- also have "... = map (\<lambda>j. 1) [0..<(p+1)]"
- apply (rule map_cong)
- subgoal by argo
- subgoal apply (subst binomial_symmetric)
- by auto
- done
- also have "... = replicate (p+1) 1"
- by (auto simp: map_replicate_trivial replicate_append_same)
- finally show ?thesis .
-qed
-
-lemma Bernstein_coeffs_01_x: assumes "p \<noteq> 0"
- shows "Bernstein_coeffs_01 p (monom 1 1) = [i/p. i \<leftarrow> [0..<(p+1)]]"
-proof -
- have
- "Bernstein_coeffs_01 p (monom 1 1) = map (\<lambda>j. inverse (real (p choose j)) *
- coeff (monom 1 (p - Suc 0) \<circ>\<^sub>p [:1, 1:]) (p - j)) [0..<(p+1)]"
- using assms by (auto simp: Bernstein_coeffs_01_def reciprocal_monom)
- also have
- "... = map (\<lambda>j. inverse (real (p choose j)) *
- (\<Sum>k\<le>p - Suc 0. coeff (monom (real (p - 1 choose k)) k) (p - j))) [0..<(p+1)]"
- by (auto simp: monom_altdef hom_distribs pcompose_pCons poly_binomial coeff_sum)
- also have"... = map (\<lambda>j. inverse (real (p choose j)) *
- real (p - 1 choose (p - j))) [0..<(p+1)]"
- by auto
- also have "... = map (\<lambda>j. j/p) [0..<(p+1)]"
- proof (rule map_cong)
- fix x assume "x \<in> set [0..<(p+1)]"
- hence "x \<le> p" by force
- thus "inverse (real (p choose x)) * real (p - 1 choose (p - x)) =
- real x / real p"
- proof (cases "x = 0")
- show "x = 0 \<Longrightarrow> ?thesis"
- using assms by fastforce
- assume 1: "x \<le> p" and 2: "x \<noteq> 0"
- hence "p - x \<le> p - 1" by force
- hence "(p - 1 choose (p - x)) = (p - 1 choose (x - 1))"
- apply (subst binomial_symmetric)
- using 1 2 by auto
- hence "x * (p choose x) = p * (p - 1 choose (p - x))"
- using 2 times_binomial_minus1_eq by simp
- hence "real x * real (p choose x) = real p * real (p - 1 choose (p - x))"
- by (metis of_nat_mult)
- thus ?thesis using 1 2
- by (auto simp: divide_simps)
- qed
- qed blast
- finally show ?thesis .
-qed
-
-lemma Bernstein_coeffs_01_add:
- assumes "degree P \<le> p" and "degree Q \<le> p"
- shows "nth_default 0 (Bernstein_coeffs_01 p (P + Q)) i =
- nth_default 0 (Bernstein_coeffs_01 p P) i +
- nth_default 0 (Bernstein_coeffs_01 p Q) i"
- using assms by (auto simp: nth_default_Bernstein_coeffs_01 degree_add_le
- reciprocal_add pcompose_add algebra_simps)
-
-lemma Bernstein_coeffs_01_smult:
- assumes "degree P \<le> p"
- shows "nth_default 0 (Bernstein_coeffs_01 p (smult a P)) i =
- a * nth_default 0 (Bernstein_coeffs_01 p P) i"
- using assms
- by (auto simp: nth_default_Bernstein_coeffs_01 reciprocal_smult
- pcompose_smult)
-
-end
+section \<open>Bernstein Polynomials over the interval [0, 1]\<close>
+
+theory Bernstein_01
+ imports "HOL-Computational_Algebra.Computational_Algebra"
+ "Budan_Fourier.Budan_Fourier"
+ "RRI_Misc"
+begin
+
+text \<open>
+The theorem of three circles is a statement about the Bernstein coefficients of a polynomial, the
+coefficients when a polynomial is expressed as a sum of Bernstein polynomials. These coefficients
+behave nicely under translations and rescaling and are the coefficients of a particular polynomial
+in the [0, 1] case. We shall define the [0, 1] case now and consider the general case later,
+deriving all the results by rescaling.
+\<close>
+
+subsection \<open>Definition and basic results\<close>
+
+definition Bernstein_Poly_01 :: "nat \<Rightarrow> nat \<Rightarrow> real poly" where
+ "Bernstein_Poly_01 j p = (monom (p choose j) j)
+ * (monom 1 (p-j) \<circ>\<^sub>p [:1, -1:])"
+
+lemma degree_Bernstein:
+ assumes hb: "j \<le> p"
+ shows "degree (Bernstein_Poly_01 j p) = p"
+proof -
+ have ha: "monom (p choose j) j \<noteq> (0::real poly)" using hb by force
+ have hb: "monom 1 (p-j) \<circ>\<^sub>p [:1, -1:] \<noteq> (0::real poly)"
+ proof
+ assume "monom 1 (p-j) \<circ>\<^sub>p [:1, -1:] = (0::real poly)"
+ hence "lead_coeff (monom 1 (p - j) \<circ>\<^sub>p [:1, -1:]) = (0::real)"
+ apply (subst leading_coeff_0_iff)
+ by simp
+ moreover have "lead_coeff (monom (1::real) (p - j)
+ \<circ>\<^sub>p [:1, -1:]) = (((- 1) ^ (p - j))::real)"
+ by (subst lead_coeff_comp, auto simp: degree_monom_eq)
+ ultimately show "False" by auto
+ qed
+ from ha hb show ?thesis
+ by (auto simp add: Bernstein_Poly_01_def degree_mult_eq
+ degree_monom_eq degree_pcompose)
+qed
+
+lemma coeff_gt:
+ assumes hb: "j > p"
+ shows "Bernstein_Poly_01 j p = 0"
+ by (simp add: hb Bernstein_Poly_01_def)
+
+lemma degree_Bernstein_le: "degree (Bernstein_Poly_01 j p) \<le> p"
+ apply (cases "j \<le> p")
+ by (simp_all add: degree_Bernstein coeff_gt)
+
+lemma poly_Bernstein_nonneg:
+ assumes "x \<ge> 0" and "1 \<ge> x"
+ shows "poly (Bernstein_Poly_01 j p) x \<ge> 0"
+ using assms by (simp add: poly_monom poly_pcompose Bernstein_Poly_01_def)
+
+lemma Bernstein_symmetry:
+ assumes "j \<le> p"
+ shows "(Bernstein_Poly_01 j p) \<circ>\<^sub>p [:1, -1:] = Bernstein_Poly_01 (p-j) p"
+proof -
+ have "(Bernstein_Poly_01 j p) \<circ>\<^sub>p [:1, -1:]
+ = ((monom (p choose j) j) * (monom 1 (p-j) \<circ>\<^sub>p [:1, -1:])) \<circ>\<^sub>p [:1, -1:]"
+ by (simp add: Bernstein_Poly_01_def)
+ also have "... = (monom (p choose (p-j)) j *
+ (monom 1 (p-j) \<circ>\<^sub>p [:1, -1:])) \<circ>\<^sub>p [:1, -1:]"
+ by (fastforce simp: binomial_symmetric[OF assms])
+ also have "... = monom (p choose (p-j)) j \<circ>\<^sub>p [:1, -1:] *
+ (monom 1 (p-j)) \<circ>\<^sub>p ([:1, -1:] \<circ>\<^sub>p [:1, -1:])"
+ by (force simp: pcompose_mult pcompose_assoc)
+ also have "... = (monom (p choose (p-j)) j \<circ>\<^sub>p [:1, -1:]) * monom 1 (p-j)"
+ by (force simp: pcompose_pCons)
+ also have "... = smult (p choose (p-j)) (monom 1 j \<circ>\<^sub>p [:1, -1:])
+ * monom 1 (p-j)"
+ by (simp add: assms smult_monom pcompose_smult[symmetric])
+ also have "... = (monom 1 j \<circ>\<^sub>p [:1, -1:]) * monom (p choose (p-j)) (p-j)"
+ apply (subst mult_smult_left)
+ apply (subst mult_smult_right[symmetric])
+ apply (subst smult_monom)
+ by force
+ also have "... = Bernstein_Poly_01 (p-j) p" using assms
+ by (auto simp: Bernstein_Poly_01_def)
+ finally show ?thesis .
+qed
+
+subsection \<open>@{term Bernstein_Poly_01} and @{term reciprocal_poly}\<close>
+
+lemma Bernstein_reciprocal:
+ "reciprocal_poly p (Bernstein_Poly_01 i p)
+ = smult (p choose i) ([:-1, 1:]^(p-i))"
+proof cases
+ assume "i \<le> p"
+ hence "reciprocal_poly p (Bernstein_Poly_01 i p) =
+ reciprocal_poly (degree (Bernstein_Poly_01 i p)) (Bernstein_Poly_01 i p)"
+ by (auto simp: degree_Bernstein)
+ also have "... = reflect_poly (Bernstein_Poly_01 i p)"
+ by (rule reciprocal_degree)
+ also have "... = smult (p choose i) ([:-1, 1:]^(p-i))"
+ by (auto simp: Bernstein_Poly_01_def reflect_poly_simps monom_altdef
+ pcompose_pCons reflect_poly_pCons' hom_distribs)
+ finally show ?thesis .
+next
+ assume h:"\<not> i \<le> p"
+ hence "reciprocal_poly p (Bernstein_Poly_01 i p) = (0::real poly)"
+ by (auto simp: coeff_gt reciprocal_poly_def)
+ also have "... = smult (p choose i) ([:-1, 1:]^(p - i))" using h
+ by fastforce
+ finally show ?thesis .
+qed
+
+lemma Bernstein_reciprocal_translate:
+ "reciprocal_poly p (Bernstein_Poly_01 i p) \<circ>\<^sub>p [:1, 1:] =
+ monom (p choose i) (p - i)"
+ by (auto simp: Bernstein_reciprocal pcompose_smult pcompose_pCons monom_altdef hom_distribs)
+
+lemma coeff_Bernstein_sum_01: fixes b::"nat \<Rightarrow> real" assumes hi: "p \<ge> i"
+ shows
+ "coeff (reciprocal_poly p
+ (\<Sum>x = 0..p. smult (b x) (Bernstein_Poly_01 x p)) \<circ>\<^sub>p [:1, 1:])
+ (p - i) = (p choose i) * (b i)" (is "?L = ?R")
+proof -
+ define P where "P \<equiv> (\<Sum>x = 0..p. (smult (b x) (Bernstein_Poly_01 x p)))"
+
+ have "\<And>x. degree (smult (b x) (Bernstein_Poly_01 x p)) \<le> p"
+ proof -
+ fix x
+ show "degree (smult (b x) (Bernstein_Poly_01 x p)) \<le> p"
+ apply (cases "x \<le> p")
+ by (auto simp: degree_Bernstein coeff_gt)
+ qed
+ hence "reciprocal_poly p P =
+ (\<Sum>x = 0..p. reciprocal_poly p (smult (b x) (Bernstein_Poly_01 x p)))"
+ apply (subst P_def)
+ apply (rule reciprocal_sum)
+ by presburger
+ also have
+ "... = (\<Sum>x = 0..p. (smult (b x * (p choose x)) ([:-1, 1:]^(p-x))))"
+ proof (rule sum.cong)
+ fix x assume "x \<in> {0..p}"
+ hence "x \<le> p" by simp
+ thus "reciprocal_poly p (smult (b x) (Bernstein_Poly_01 x p)) =
+ smult ((b x) * (p choose x)) ([:-1, 1:]^(p-x))"
+ by (auto simp add: reciprocal_smult degree_Bernstein Bernstein_reciprocal)
+ qed (simp)
+ finally have
+ "reciprocal_poly p P =
+ (\<Sum>x = 0..p. (smult ((b x) * (p choose x)) ([:-1, 1:]^(p-x))))" .
+ hence
+ "(reciprocal_poly p P) \<circ>\<^sub>p [:1, 1:] =
+ (\<Sum>x = 0..p. (smult ((b x) * (p choose x)) ([:-1, 1:]^(p-x))) \<circ>\<^sub>p [:1, 1:])"
+ by (simp add: pcompose_sum pcompose_add)
+ also have "... = (\<Sum>x = 0..p. (monom ((b x) * (p choose x)) (p - x)))"
+ proof (rule sum.cong)
+ fix x assume "x \<in> {0..p}"
+ hence "x \<le> p" by simp
+ thus "smult (b x * (p choose x)) ([:- 1, 1:] ^ (p - x)) \<circ>\<^sub>p [:1, 1:] =
+ monom (b x * (p choose x)) (p - x)"
+ by (simp add: hom_distribs pcompose_smult pcompose_pCons monom_altdef)
+ qed (simp)
+ finally have "(reciprocal_poly p P) \<circ>\<^sub>p [:1, 1:] =
+ (\<Sum>x = 0..p. (monom ((b x) * (p choose x)) (p - x)))" .
+ hence "?L = (\<Sum>x = 0..p. if p - x = p - i then b x * real (p choose x) else 0)"
+ by (auto simp add: P_def coeff_sum)
+ also have "... = (\<Sum>x = 0..p. if x = i then b x * real (p choose x) else 0)"
+ proof (rule sum.cong)
+ fix x assume "x \<in> {0..p}"
+ hence "x \<le> p" by simp
+ thus "(if p - x = p - i then b x * real (p choose x) else 0) =
+ (if x = i then b x * real (p choose x) else 0)" using hi
+ by (auto simp add: leI)
+ qed (simp)
+ also have "... = ?R" by simp
+ finally show ?thesis .
+qed
+
+lemma Bernstein_sum_01: assumes hP: "degree P \<le> p"
+ shows
+ "P = (\<Sum>j = 0..p. smult
+ (inverse (real (p choose j)) *
+ coeff (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:]) (p-j))
+ (Bernstein_Poly_01 j p))"
+proof -
+ define Q where "Q \<equiv> reciprocal_poly p P \<circ>\<^sub>p [:1, 1:]"
+ from hP Q_def have hQ: "degree Q \<le> p"
+ by (auto simp: degree_reciprocal degree_pcompose)
+ have "reciprocal_poly p (\<Sum>j = 0..p.
+ smult (inverse (real (p choose j)) * coeff Q (p-j))
+ (Bernstein_Poly_01 j p)) \<circ>\<^sub>p [:1, 1:] = Q"
+ proof (rule poly_eqI)
+ fix n
+ show "coeff (reciprocal_poly p (\<Sum>j = 0..p.
+ smult (inverse (real (p choose j)) * coeff Q (p-j))
+ (Bernstein_Poly_01 j p)) \<circ>\<^sub>p [:1, 1:]) n = coeff Q n"
+ (is "?L = ?R")
+ proof cases
+ assume hn: "n \<le> p"
+ hence "?L = coeff (reciprocal_poly p (\<Sum>j = 0..p.
+ smult (inverse (real (p choose j)) * coeff Q (p-j))
+ (Bernstein_Poly_01 j p)) \<circ>\<^sub>p [:1, 1:]) (p - (p - n))"
+ by force
+ also have "... = (p choose (p-n)) *
+ (inverse (real (p choose (p-n))) *
+ coeff Q (p-(p-n)))"
+ apply (subst coeff_Bernstein_sum_01)
+ by auto
+ also have "... = ?R" using hn
+ by fastforce
+ finally show "?L = ?R" .
+ next
+ assume hn: "\<not> n \<le> p"
+ have "degree (\<Sum>j = 0..p.
+ smult (inverse (real (p choose j)) * coeff Q (p - j))
+ (Bernstein_Poly_01 j p)) \<le> p"
+ proof (rule degree_sum_le)
+ fix q assume "q \<in> {0..p}"
+ hence "q \<le> p" by fastforce
+ thus "degree (smult (inverse (real (p choose q)) *
+ coeff Q (p - q)) (Bernstein_Poly_01 q p)) \<le> p"
+ by (auto simp add: degree_Bernstein degree_smult_le)
+ qed simp
+ hence "degree (reciprocal_poly p (\<Sum>j = 0..p.
+ smult (inverse (real (p choose j)) * coeff Q (p - j))
+ (Bernstein_Poly_01 j p)) \<circ>\<^sub>p [:1, 1:]) \<le> p"
+ by (auto simp add: degree_pcompose degree_reciprocal)
+ hence "?L = 0" using hn by (auto simp add: coeff_eq_0)
+ thus "?L = ?R" using hQ hn by (simp add: coeff_eq_0)
+ qed
+ qed
+ hence "reciprocal_poly p P \<circ>\<^sub>p [:1, 1:] =
+ reciprocal_poly p (\<Sum>j = 0..p.
+ smult (inverse (real (p choose j)) *
+ coeff (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:]) (p-j))
+ (Bernstein_Poly_01 j p)) \<circ>\<^sub>p [:1, 1:]"
+ by (auto simp: degree_reciprocal degree_pcompose Q_def)
+ hence "reciprocal_poly p P \<circ>\<^sub>p ([:1, 1:] \<circ>\<^sub>p [:-1, 1:]) =
+ reciprocal_poly p (\<Sum>j = 0..p. smult (inverse (real (p choose j)) *
+ coeff (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:]) (p-j))
+ (Bernstein_Poly_01 j p)) \<circ>\<^sub>p ([:1, 1:] \<circ>\<^sub>p [:-1, 1:])"
+ by (auto simp: pcompose_assoc)
+ hence "reciprocal_poly p P = reciprocal_poly p (\<Sum>j = 0..p.
+ smult (inverse (real (p choose j)) *
+ coeff (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:]) (p-j)) (Bernstein_Poly_01 j p))"
+ by (auto simp: pcompose_pCons)
+ hence "reciprocal_poly p (reciprocal_poly p P) =
+ reciprocal_poly p (reciprocal_poly p (\<Sum>j = 0..p.
+ smult (inverse (real (p choose j)) *
+ coeff (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:]) (p-j)) (Bernstein_Poly_01 j p)))"
+ by argo
+ thus "P = (\<Sum>j = 0..p. smult (inverse (real (p choose j)) *
+ coeff (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:]) (p-j)) (Bernstein_Poly_01 j p))"
+ using hP by (auto simp: reciprocal_reciprocal degree_sum_le degree_smult_le
+ degree_Bernstein degree_add_le)
+qed
+
+lemma Bernstein_Poly_01_span1:
+ assumes hP: "degree P \<le> p"
+ shows "P \<in> poly_vs.span {Bernstein_Poly_01 x p | x. x \<le> p}"
+proof -
+ have "Bernstein_Poly_01 x p
+ \<in> poly_vs.span {Bernstein_Poly_01 x p |x. x \<le> p}"
+ if "x \<in> {0..p}" for x
+ proof -
+ have "\<exists>n. Bernstein_Poly_01 x p = Bernstein_Poly_01 n p \<and> n \<le> p"
+ using that by force
+ then show
+ "Bernstein_Poly_01 x p \<in> poly_vs.span {Bernstein_Poly_01 n p |n. n \<le> p}"
+ by (simp add: poly_vs.span_base)
+ qed
+ thus ?thesis
+ apply (subst Bernstein_sum_01[OF hP])
+ apply (rule poly_vs.span_sum)
+ apply (rule poly_vs.span_scale)
+ by blast
+qed
+
+lemma Bernstein_Poly_01_span:
+ "poly_vs.span {Bernstein_Poly_01 x p | x. x \<le> p}
+ = {x. degree x \<le> p}"
+ apply (subst monom_span[symmetric])
+ apply (subst poly_vs.span_eq)
+ by (auto simp: monom_span degree_Bernstein_le
+ Bernstein_Poly_01_span1 degree_monom_eq)
+
+subsection \<open>Bernstein coefficients and changes\<close>
+
+definition Bernstein_coeffs_01 :: "nat \<Rightarrow> real poly \<Rightarrow> real list" where
+ "Bernstein_coeffs_01 p P =
+ [(inverse (real (p choose j)) *
+ coeff (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:]) (p-j)). j \<leftarrow> [0..<(p+1)]]"
+
+lemma length_Bernstein_coeffs_01: "length (Bernstein_coeffs_01 p P) = p + 1"
+ by (auto simp: Bernstein_coeffs_01_def)
+
+lemma nth_default_Bernstein_coeffs_01: assumes "degree P \<le> p"
+ shows "nth_default 0 (Bernstein_coeffs_01 p P) i =
+ inverse (p choose i) * coeff (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:]) (p-i)"
+ apply (cases "p = i")
+ using assms by (auto simp: Bernstein_coeffs_01_def nth_default_append
+ nth_default_Cons Nitpick.case_nat_unfold binomial_eq_0)
+
+lemma Bernstein_coeffs_01_sum: assumes "degree P \<le> p"
+ shows "P = (\<Sum>j = 0..p. smult (nth_default 0 (Bernstein_coeffs_01 p P) j)
+ (Bernstein_Poly_01 j p))"
+ apply (subst nth_default_Bernstein_coeffs_01[OF assms])
+ apply (subst Bernstein_sum_01[OF assms])
+ by argo
+
+definition Bernstein_changes_01 :: "nat \<Rightarrow> real poly \<Rightarrow> int" where
+ "Bernstein_changes_01 p P = nat (changes (Bernstein_coeffs_01 p P))"
+
+lemma Bernstein_changes_01_def':
+ "Bernstein_changes_01 p P = nat (changes [(inverse (real (p choose j)) *
+ coeff (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:]) (p-j)). j \<leftarrow> [0..<p + 1]])"
+ by (simp add: Bernstein_changes_01_def Bernstein_coeffs_01_def)
+
+lemma Bernstein_changes_01_eq_changes:
+ assumes hP: "degree P \<le> p"
+ shows "Bernstein_changes_01 p P =
+ changes (coeffs ((reciprocal_poly p P) \<circ>\<^sub>p [:1, 1:]))"
+proof (subst Bernstein_changes_01_def')
+ have h:
+ "map (\<lambda>j. inverse (real (p choose j)) *
+ coeff (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:]) (p - j)) [0..<p + 1] =
+ map (\<lambda>j. inverse (real (p choose j)) *
+ nth_default 0 [nth_default 0 (coeffs (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:]))
+ (p - j). j \<leftarrow> [0..<p + 1]] j) [0..<p + 1]"
+ proof (rule map_cong)
+ fix x
+ assume "x \<in> set [0..<p+1]"
+ hence hx: "x \<le> p" by fastforce
+ moreover have 1:
+ "length (map (\<lambda>j. nth_default 0
+ (coeffs (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:])) (p - j)) [0..<p + 1]) \<le> Suc p"
+ by force
+ moreover have "length (coeffs (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:])) \<le> Suc p"
+ proof (cases "P=0")
+ case False
+ then have "reciprocal_poly p P \<circ>\<^sub>p [:1, 1:] \<noteq> 0"
+ using hP by (simp add: Missing_Polynomial.pcompose_eq_0 reciprocal_0_iff)
+ moreover have "Suc (degree (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:])) \<le> Suc p"
+ using hP by (auto simp: degree_pcompose degree_reciprocal)
+ ultimately show ?thesis
+ using length_coeffs_degree by force
+ qed (auto simp: reciprocal_0)
+ ultimately have h:
+ "nth_default 0 (map (\<lambda>j. nth_default 0 (coeffs
+ (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:])) (p - j)) [0..<p + 1]) x =
+ nth_default 0 (coeffs (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:])) (p - x)"
+ (is "?L = ?R")
+ proof -
+ have "?L = (map (\<lambda>j. nth_default 0 (coeffs
+ (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:])) (p - j)) [0..<p + 1]) ! x"
+ using hx by (auto simp: nth_default_nth)
+ also have "... = nth_default 0
+ (coeffs (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:])) (p - [0..<p + 1] ! x)"
+ apply (subst nth_map)
+ using hx by auto
+ also have "... = ?R"
+ apply (subst nth_upt)
+ using hx by auto
+ finally show ?thesis .
+ qed
+ show "inverse (real (p choose x)) *
+ coeff (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:]) (p - x) =
+ inverse (real (p choose x)) *
+ nth_default 0 (map (\<lambda>j. nth_default 0
+ (coeffs (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:])) (p - j)) [0..<p + 1]) x"
+ apply (subst h)
+ apply (subst nth_default_coeffs_eq)
+ by blast
+ qed auto
+
+ have 1:
+ "rev (map (\<lambda>j. nth_default 0 (coeffs (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:]))
+ (p - j)) [0..<p + 1]) = map (\<lambda>j. nth_default 0 (coeffs
+ (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:])) j) [0..<p + 1]"
+ proof (subst rev_map, rule map_cong')
+ have "\<And>q. (q \<ge> p \<longrightarrow> rev [q-p..<q+1] = map ((-) q) [0..<p+1])"
+ proof (induction p)
+ case 0
+ then show ?case by simp
+ next
+ case (Suc p)
+ have IH: "\<And>q. (q \<ge> p \<longrightarrow> rev [q-p..<q+1] = map ((-) q) [0..<p+1])"
+ using Suc.IH by blast
+ show ?case
+ proof
+ assume hq: "Suc p \<le> q"
+ then have h: "rev [q - p..<q + 1] = map ((-) (q)) [0..<p + 1]"
+ apply (subst IH)
+ using hq by auto
+ have "[q - Suc p..<q + 1] = (q - Suc p) # [q - p..<q + 1]"
+ by (simp add: Suc_diff_Suc Suc_le_lessD hq upt_conv_Cons)
+ hence "rev [q - Suc p..<q + 1] = rev [q - p..<q + 1] @ [q - Suc p]"
+ by force
+ also have "... = map ((-) (q)) [0..<p + 1] @ [q - Suc p]"
+ using h by blast
+ also have "... = map ((-) q) [0..<Suc p + 1]"
+ by force
+ finally show "rev [q - Suc p..<q + 1] = map ((-) q) [0..<Suc p + 1]" .
+ qed
+ qed
+ thus "rev [0..<p + 1] = map ((-) p) [0..<p + 1]"
+ by force
+ next
+ fix y
+ assume "y \<in> set [0..<p + 1]"
+ hence "y \<le> p" by fastforce
+ thus "nth_default 0 (coeffs (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:])) (p - (p - y)) =
+ nth_default 0 (coeffs (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:])) y"
+ by fastforce
+ qed
+
+ have 2: "\<And> f. f \<noteq> 0 \<longrightarrow> degree f \<le> p \<longrightarrow>
+ map (nth_default 0 (coeffs f)) [0..<p + 1] =
+ coeffs f @ replicate (p - degree f) 0"
+ proof (induction p)
+ case 0
+ then show ?case by (auto simp: degree_0_iff)
+ next
+ fix f
+ case (Suc p)
+ hence IH: "(f \<noteq> 0 \<longrightarrow>
+ degree f \<le> p \<longrightarrow>
+ map (nth_default 0 (coeffs f)) [0..<p + 1] =
+ coeffs f @ replicate (p - degree f) 0)" by blast
+ then show ?case
+ proof (cases)
+ assume h': "Suc p = degree f"
+ hence h: "[0..<Suc p + 1] = [0..<length (coeffs f)]"
+ by (metis add_is_0 degree_0 length_coeffs plus_1_eq_Suc zero_neq_one)
+ thus ?thesis
+ apply (subst h)
+ apply (subst map_nth_default)
+ using h' by fastforce
+ next
+ assume h': "Suc p \<noteq> degree f"
+ show ?thesis
+ proof
+ assume hf: "f \<noteq> 0"
+ show "degree f \<le> Suc p \<longrightarrow>
+ map (nth_default 0 (coeffs f)) [0..<Suc p + 1] =
+ coeffs f @ replicate (Suc p - degree f) 0"
+ proof
+ assume "degree f \<le> Suc p"
+ hence 1: "degree f \<le> p" using h' by fastforce
+ hence 2: "map (nth_default 0 (coeffs f)) [0..<p + 1] =
+ coeffs f @ replicate (p - degree f) 0" using IH hf by blast
+ have "map (nth_default 0 (coeffs f)) [0..<Suc p + 1] =
+ map (nth_default 0 (coeffs f)) [0..<p + 1] @
+ [nth_default 0 (coeffs f) (Suc p)]"
+ by fastforce
+ also have
+ "... = coeffs f @ replicate (p - degree f) 0 @ [coeff f (Suc p)]"
+ using 2
+ by (auto simp: nth_default_coeffs_eq)
+ also have "... = coeffs f @ replicate (p - degree f) 0 @ [0]"
+ using \<open>degree f \<le> Suc p\<close> h' le_antisym le_degree by blast
+ also have "... = coeffs f @ replicate (Suc p - degree f) 0" using 1
+ by (simp add: Suc_diff_le replicate_app_Cons_same)
+ finally show "map (nth_default 0 (coeffs f)) [0..<Suc p + 1] =
+ coeffs f @ replicate (Suc p - degree f) 0" .
+ qed
+ qed
+ qed
+ qed
+
+ thus "int (nat (changes (map (\<lambda>j. inverse (real (p choose j)) *
+ coeff (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:]) (p - j)) [0..<p + 1]))) =
+ changes (coeffs (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:]))"
+ proof cases
+ assume hP: "P = 0"
+ show "int (nat (changes (map (\<lambda>j. inverse (real (p choose j)) *
+ coeff (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:]) (p - j)) [0..<p + 1]))) =
+ changes (coeffs (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:]))" (is "?L = ?R")
+ proof -
+ have "?L = int (nat (changes (map (\<lambda>j. 0::real) [0..<p+1])))"
+ using hP by (auto simp: reciprocal_0 changes_nonneg)
+ also have "... = 0"
+ apply (induction p)
+ by (auto simp: map_replicate_trivial changes_nonneg
+ replicate_app_Cons_same)
+ also have "0 = changes ([]::real list)" by simp
+ also have "... = ?R" using hP by (auto simp: reciprocal_0)
+ finally show ?thesis .
+ qed
+ next
+ assume hP': "P \<noteq> 0"
+ thus ?thesis
+ apply (subst h)
+ apply (subst changes_scale)
+ apply auto[2]
+ apply (subst changes_rev[symmetric])
+ apply (subst 1)
+ apply (subst 2)
+ apply (simp add: pcompose_eq_0 hP reciprocal_0_iff)
+ using assms apply (auto simp: degree_reciprocal)[1]
+ by (auto simp: changes_append_replicate_0 changes_nonneg)
+ qed
+qed
+
+lemma Bernstein_changes_01_test: fixes P::"real poly"
+ assumes hP: "degree P \<le> p" and h0: "P \<noteq> 0"
+ shows "proots_count P {x. 0 < x \<and> x < 1} \<le> Bernstein_changes_01 p P \<and>
+ even (Bernstein_changes_01 p P - proots_count P {x. 0 < x \<and> x < 1})"
+proof -
+ let ?Q = "(reciprocal_poly p P) \<circ>\<^sub>p [:1, 1:]"
+
+ have 1: "changes (coeffs ?Q) \<ge> proots_count ?Q {x. 0 < x} \<and>
+ even (changes (coeffs ?Q) - proots_count ?Q {x. 0 < x})"
+ apply (rule descartes_sign)
+ by (simp add: Missing_Polynomial.pcompose_eq_0 h0 hP reciprocal_0_iff)
+
+ have "((+) (1::real) ` Collect ((<) (0::real))) = {x. (1::real)<x}"
+ proof
+ show "{x::real. 1 < x} \<subseteq> (+) 1 ` Collect ((<) 0)"
+ proof
+ fix x::real assume "x \<in> {x. 1 < x}"
+ hence "1 < x" by simp
+ hence "-1 + x \<in> Collect ((<) 0)" by auto
+ hence "1 + (-1 + x) \<in> (+) 1 ` Collect ((<) 0)" by blast
+ thus "x \<in> (+) 1 ` Collect ((<) 0)" by argo
+ qed
+ qed auto
+ hence 2: "proots_count P {x. 0 < x \<and> x < 1} = proots_count ?Q {x. 0 < x}"
+ using assms
+ by (auto simp: proots_pcompose reciprocal_0_iff proots_count_reciprocal')
+
+ show ?thesis
+ apply (subst Bernstein_changes_01_eq_changes[OF hP])
+ apply (subst Bernstein_changes_01_eq_changes[OF hP])
+ apply (subst 2)
+ apply (subst 2)
+ by (rule 1)
+qed
+
+subsection \<open>Expression as a Bernstein sum\<close>
+
+lemma Bernstein_coeffs_01_0: "Bernstein_coeffs_01 p 0 = replicate (p+1) 0"
+ by (auto simp: Bernstein_coeffs_01_def reciprocal_0 map_replicate_trivial
+ replicate_append_same)
+
+lemma Bernstein_coeffs_01_1: "Bernstein_coeffs_01 p 1 = replicate (p+1) 1"
+proof -
+ have "Bernstein_coeffs_01 p 1 =
+ map (\<lambda>j. inverse (real (p choose j)) *
+ coeff (\<Sum>k\<le>p. smult (real (p choose k)) ([:0, 1:] ^ k)) (p - j)) [0..<(p+1)]"
+ by (auto simp: Bernstein_coeffs_01_def reciprocal_1 monom_altdef
+ hom_distribs pcompose_pCons poly_0_coeff_0[symmetric] poly_binomial)
+ also have "... = map (\<lambda>j. inverse (real (p choose j)) *
+ real (p choose (p - j))) [0..<(p+1)]"
+ by (auto simp: monom_altdef[symmetric] coeff_sum binomial)
+ also have "... = map (\<lambda>j. 1) [0..<(p+1)]"
+ apply (rule map_cong)
+ subgoal by argo
+ subgoal apply (subst binomial_symmetric)
+ by auto
+ done
+ also have "... = replicate (p+1) 1"
+ by (auto simp: map_replicate_trivial replicate_append_same)
+ finally show ?thesis .
+qed
+
+lemma Bernstein_coeffs_01_x: assumes "p \<noteq> 0"
+ shows "Bernstein_coeffs_01 p (monom 1 1) = [i/p. i \<leftarrow> [0..<(p+1)]]"
+proof -
+ have
+ "Bernstein_coeffs_01 p (monom 1 1) = map (\<lambda>j. inverse (real (p choose j)) *
+ coeff (monom 1 (p - Suc 0) \<circ>\<^sub>p [:1, 1:]) (p - j)) [0..<(p+1)]"
+ using assms by (auto simp: Bernstein_coeffs_01_def reciprocal_monom)
+ also have
+ "... = map (\<lambda>j. inverse (real (p choose j)) *
+ (\<Sum>k\<le>p - Suc 0. coeff (monom (real (p - 1 choose k)) k) (p - j))) [0..<(p+1)]"
+ by (auto simp: monom_altdef hom_distribs pcompose_pCons poly_binomial coeff_sum)
+ also have"... = map (\<lambda>j. inverse (real (p choose j)) *
+ real (p - 1 choose (p - j))) [0..<(p+1)]"
+ by auto
+ also have "... = map (\<lambda>j. j/p) [0..<(p+1)]"
+ proof (rule map_cong)
+ fix x assume "x \<in> set [0..<(p+1)]"
+ hence "x \<le> p" by force
+ thus "inverse (real (p choose x)) * real (p - 1 choose (p - x)) =
+ real x / real p"
+ proof (cases "x = 0")
+ show "x = 0 \<Longrightarrow> ?thesis"
+ using assms by fastforce
+ assume 1: "x \<le> p" and 2: "x \<noteq> 0"
+ hence "p - x \<le> p - 1" by force
+ hence "(p - 1 choose (p - x)) = (p - 1 choose (x - 1))"
+ apply (subst binomial_symmetric)
+ using 1 2 by auto
+ hence "x * (p choose x) = p * (p - 1 choose (p - x))"
+ using 2 times_binomial_minus1_eq by simp
+ hence "real x * real (p choose x) = real p * real (p - 1 choose (p - x))"
+ by (metis of_nat_mult)
+ thus ?thesis using 1 2
+ by (auto simp: divide_simps)
+ qed
+ qed blast
+ finally show ?thesis .
+qed
+
+lemma Bernstein_coeffs_01_add:
+ assumes "degree P \<le> p" and "degree Q \<le> p"
+ shows "nth_default 0 (Bernstein_coeffs_01 p (P + Q)) i =
+ nth_default 0 (Bernstein_coeffs_01 p P) i +
+ nth_default 0 (Bernstein_coeffs_01 p Q) i"
+ using assms by (auto simp: nth_default_Bernstein_coeffs_01 degree_add_le
+ reciprocal_add pcompose_add algebra_simps)
+
+lemma Bernstein_coeffs_01_smult:
+ assumes "degree P \<le> p"
+ shows "nth_default 0 (Bernstein_coeffs_01 p (smult a P)) i =
+ a * nth_default 0 (Bernstein_coeffs_01 p P) i"
+ using assms
+ by (auto simp: nth_default_Bernstein_coeffs_01 reciprocal_smult
+ pcompose_smult)
+
+end
diff --git a/thys/Three_Circles/Normal_Poly.thy b/thys/Three_Circles/Normal_Poly.thy
--- a/thys/Three_Circles/Normal_Poly.thy
+++ b/thys/Three_Circles/Normal_Poly.thy
@@ -1,1106 +1,1106 @@
-section \<open>Normal Polynomials\<close>
-
-theory Normal_Poly
- imports "RRI_Misc"
-begin
-
-text \<open>
-Here we define normal polynomials as defined in
- Basu, S., Pollack, R., Roy, M.-F.: Algorithms in Real Algebraic Geometry.
- Springer Berlin Heidelberg, Berlin, Heidelberg (2016).
-\<close>
-
-definition normal_poly :: "('a::{comm_ring_1,ord}) poly \<Rightarrow> bool" where
-"normal_poly p \<equiv>
- (p \<noteq> 0) \<and>
- (\<forall> i. 0 \<le> coeff p i) \<and>
- (\<forall> i. coeff p i * coeff p (i+2) \<le> (coeff p (i+1))^2) \<and>
- (\<forall> i j k. i \<le> j \<longrightarrow> j \<le> k \<longrightarrow> 0 < coeff p i
- \<longrightarrow> 0 < coeff p k \<longrightarrow> 0 < coeff p j)"
-
-lemma normal_non_zero: "normal_poly p \<Longrightarrow> p \<noteq> 0"
- using normal_poly_def by blast
-
-lemma normal_coeff_nonneg: "normal_poly p \<Longrightarrow> 0 \<le> coeff p i"
- using normal_poly_def by metis
-
-lemma normal_poly_coeff_mult:
- "normal_poly p \<Longrightarrow> coeff p i * coeff p (i+2) \<le> (coeff p (i+1))^2"
- using normal_poly_def by blast
-
-lemma normal_poly_pos_interval:
- "normal_poly p \<Longrightarrow> i \<le> j \<Longrightarrow> j \<le> k \<Longrightarrow> 0 < coeff p i \<Longrightarrow> 0 < coeff p k
- \<Longrightarrow> 0 < coeff p j"
- using normal_poly_def by blast
-
-lemma normal_polyI:
- assumes "(p \<noteq> 0)"
- and "(\<And> i. 0 \<le> coeff p i)"
- and "(\<And> i. coeff p i * coeff p (i+2) \<le> (coeff p (i+1))^2)"
- and "(\<And> i j k. i \<le> j \<Longrightarrow> j \<le> k \<Longrightarrow> 0 < coeff p i \<Longrightarrow> 0 < coeff p k \<Longrightarrow> 0 < coeff p j)"
- shows "normal_poly p"
- using assms by (force simp: normal_poly_def)
-
-lemma linear_normal_iff:
- fixes x::real
- shows "normal_poly [:-x, 1:] \<longleftrightarrow> x \<le> 0"
-proof
- assume "normal_poly [:-x, 1:]"
- thus "x \<le> 0" using normal_coeff_nonneg[of "[:-x, 1:]" 0] by auto
-next
- assume "x \<le> 0"
- then have "0 \<le> coeff [:- x, 1:] i" for i
- by (cases i) (simp_all add: pCons_one)
- moreover have "0 < coeff [:- x, 1:] j"
- if "i \<le> j" "j \<le> k" "0 < coeff [:- x, 1:] i"
- "0 < coeff [:- x, 1:] k" for i j k
- apply (cases "k=0 \<or> i=0")
- subgoal using that
- by (smt (z3) bot_nat_0.extremum_uniqueI degree_pCons_eq_if
- le_antisym le_degree not_less_eq_eq)
- subgoal using that
- by (smt (z3) One_nat_def degree_pCons_eq_if le_degree less_one
- not_le one_neq_zero pCons_one verit_la_disequality)
- done
- ultimately show "normal_poly [:-x, 1:]"
- unfolding normal_poly_def by auto
-qed
-
-lemma quadratic_normal_iff:
- fixes z::complex
- shows "normal_poly [:(cmod z)\<^sup>2, -2*Re z, 1:]
- \<longleftrightarrow> Re z \<le> 0 \<and> 4*(Re z)^2 \<ge> (cmod z)^2"
-proof
- assume "normal_poly [:(cmod z)\<^sup>2, - 2 * Re z, 1:]"
- hence "-2*Re z \<ge> 0 \<and> (cmod z)^2 \<ge> 0 \<and> (-2*Re z)^2 \<ge> (cmod z)^2"
- using normal_coeff_nonneg[of _ 1] normal_poly_coeff_mult[of _ 0]
- by fastforce
- thus "Re z \<le> 0 \<and> 4*(Re z)^2 \<ge> (cmod z)^2"
- by auto
-next
- assume asm:"Re z \<le> 0 \<and> 4*(Re z)^2 \<ge> (cmod z)^2"
- define P where "P=[:(cmod z)\<^sup>2, - 2 * Re z, 1:]"
-
- have "0 \<le> coeff P i" for i
- unfolding P_def using asm
- apply (cases "i=0\<or>i=1\<or>i=2")
- by (auto simp:numeral_2_eq_2 coeff_eq_0)
- moreover have "coeff P i * coeff P (i + 2) \<le> (coeff P (i + 1))\<^sup>2" for i
- apply (cases "i=0\<or>i=1\<or>i=2")
- using asm
- unfolding P_def by (auto simp:coeff_eq_0)
- moreover have "0 < coeff P j"
- if "0 < coeff P k" "0 < coeff P i" "j \<le> k" "i \<le> j"
- for i j k
- using that unfolding P_def
- apply (cases "k=0 \<or> k=1 \<or> k=2")
- subgoal using asm
- by (smt (z3) One_nat_def Suc_1 bot_nat_0.extremum_uniqueI
- coeff_pCons_0 coeff_pCons_Suc le_Suc_eq
- zero_less_power2)
- subgoal by (auto simp:coeff_eq_0)
- done
- moreover have "P\<noteq>0" unfolding P_def by auto
- ultimately show "normal_poly P"
- unfolding normal_poly_def by blast
-qed
-
-lemma normal_of_no_zero_root:
- fixes f::"real poly"
- assumes hzero: "poly f 0 \<noteq> 0" and hdeg: "i \<le> degree f"
- and hnorm: "normal_poly f"
- shows "0 < coeff f i"
-proof -
- have "coeff f 0 > 0" using hzero normal_coeff_nonneg[OF hnorm]
- by (metis eq_iff not_le_imp_less poly_0_coeff_0)
- moreover have "coeff f (degree f) > 0" using normal_coeff_nonneg[OF hnorm] normal_non_zero[OF hnorm]
- by (meson dual_order.irrefl eq_iff eq_zero_or_degree_less not_le_imp_less)
- moreover have "0 \<le> i" by simp
- ultimately show "0 < coeff f i" using hdeg normal_poly_pos_interval[OF hnorm] by blast
-qed
-
-lemma normal_divide_x:
- fixes f::"real poly"
- assumes hnorm: "normal_poly (f*[:0,1:])"
- shows "normal_poly f"
-proof (rule normal_polyI)
- show "f \<noteq> 0"
- using normal_non_zero[OF hnorm] by auto
-next
- fix i
- show "0 \<le> coeff f i"
- using normal_coeff_nonneg[OF hnorm, of "Suc i"] by (simp add: coeff_pCons)
-next
- fix i
- show "coeff f i * coeff f (i + 2) \<le> (coeff f (i + 1))\<^sup>2"
- using normal_poly_coeff_mult[OF hnorm, of "Suc i"] by (simp add: coeff_pCons)
-next
- fix i j k
- show "i \<le> j \<Longrightarrow> j \<le> k \<Longrightarrow> 0 < coeff f i \<Longrightarrow> 0 < coeff f k \<Longrightarrow> 0 < coeff f j"
- using normal_poly_pos_interval[of _ "Suc i" "Suc j" "Suc k", OF hnorm]
- by (simp add: coeff_pCons)
-qed
-
-lemma normal_mult_x:
- fixes f::"real poly"
- assumes hnorm: "normal_poly f"
- shows "normal_poly (f * [:0, 1:])"
-proof (rule normal_polyI)
- show "f * [:0, 1:] \<noteq> 0"
- using normal_non_zero[OF hnorm] by auto
-next
- fix i
- show "0 \<le> coeff (f * [:0, 1:]) i"
- using normal_coeff_nonneg[OF hnorm, of "i-1"] by (cases i, auto simp: coeff_pCons)
-next
- fix i
- show "coeff (f * [:0, 1:]) i * coeff (f * [:0, 1:]) (i + 2) \<le> (coeff (f * [:0, 1:]) (i + 1))\<^sup>2"
- using normal_poly_coeff_mult[OF hnorm, of "i-1"] by (cases i, auto simp: coeff_pCons)
-next
- fix i j k
- show "i \<le> j \<Longrightarrow> j \<le> k \<Longrightarrow> 0 < coeff (f * [:0, 1:]) i \<Longrightarrow> 0 < coeff (f * [:0, 1:]) k \<Longrightarrow> 0 < coeff (f * [:0, 1:]) j"
- using normal_poly_pos_interval[of _ "i-1" "j-1" "k-1", OF hnorm]
- apply (cases i, force)
- apply (cases j, force)
- apply (cases k, force)
- by (auto simp: coeff_pCons)
-qed
-
-lemma normal_poly_general_coeff_mult:
- fixes f::"real poly"
- assumes "normal_poly f" and "h \<le> j"
- shows "coeff f (h+1) * coeff f (j+1) \<ge> coeff f h * coeff f (j+2)"
-using assms proof (induction j)
- case 0
- then show ?case
- using normal_poly_coeff_mult by (auto simp: power2_eq_square)[1]
-next
- case (Suc j)
- then show ?case
- proof (cases "h = Suc j")
- assume "h = Suc j" "normal_poly f"
- thus ?thesis
- using normal_poly_coeff_mult by (auto simp: power2_eq_square)
- next
- assume "(normal_poly f \<Longrightarrow>
- h \<le> j \<Longrightarrow> coeff f h * coeff f (j + 2) \<le> coeff f (h + 1) * coeff f (j + 1))"
- "normal_poly f" and h: "h \<le> Suc j" "h \<noteq> Suc j"
- hence IH: "coeff f h * coeff f (j + 2) \<le> coeff f (h + 1) * coeff f (j + 1)"
- by linarith
- show ?thesis
- proof (cases "coeff f (Suc j + 1) = 0", cases "coeff f (Suc j + 2) = 0")
- show "coeff f (Suc j + 1) = 0 \<Longrightarrow> coeff f (Suc j + 2) = 0 \<Longrightarrow>
- coeff f h * coeff f (Suc j + 2) \<le> coeff f (h + 1) * coeff f (Suc j + 1)"
- by (metis assms(1) mult_zero_right normal_coeff_nonneg)
- next
- assume 1: "coeff f (Suc j + 1) = 0" "coeff f (Suc j + 2) \<noteq> 0"
- hence "coeff f (Suc j + 2) > 0" "\<not>coeff f (Suc j + 1) > 0"
- using normal_coeff_nonneg[of f "Suc j + 2"] assms(1) by auto
- hence "coeff f h > 0 \<Longrightarrow> False"
- using normal_poly_pos_interval[of f h "Suc j + 1" "Suc j + 2"] assms(1) h by force
- hence "coeff f h = 0"
- using normal_coeff_nonneg[OF assms(1)] less_eq_real_def by auto
- thus "coeff f h * coeff f (Suc j + 2) \<le> coeff f (h + 1) * coeff f (Suc j + 1)"
- using 1 by fastforce
- next
- assume 1: "coeff f (Suc j + 1) \<noteq> 0"
- show "coeff f h * coeff f (Suc j + 2) \<le> coeff f (h + 1) * coeff f (Suc j + 1)"
- proof (cases "coeff f (Suc j) = 0")
- assume 2: "coeff f (Suc j) = 0"
- hence "coeff f (Suc j + 1) > 0" "\<not>coeff f (Suc j) > 0"
- using normal_coeff_nonneg[of f "Suc j + 1"] assms(1) 1 by auto
- hence "coeff f h > 0 \<Longrightarrow> False"
- using normal_poly_pos_interval[of f h "Suc j" "Suc j + 1"] assms(1) h by force
- hence "coeff f h = 0"
- using normal_coeff_nonneg[OF assms(1)] less_eq_real_def by auto
- thus "coeff f h * coeff f (Suc j + 2) \<le> coeff f (h + 1) * coeff f (Suc j + 1)"
- by (simp add: assms(1) normal_coeff_nonneg)
- next
- assume 2: "coeff f (Suc j) \<noteq> 0"
- from normal_poly_coeff_mult[OF assms(1), of "Suc j"] normal_coeff_nonneg[OF assms(1), of "Suc j"]
- normal_coeff_nonneg[OF assms(1), of "Suc (Suc j)"] 1 2
- have 3: "coeff f (Suc j + 1) / coeff f (Suc j) \<ge> coeff f (Suc j + 2) / coeff f (Suc j + 1)"
- by (auto simp: power2_eq_square divide_simps algebra_simps)
- have "(coeff f h * coeff f (j + 2)) * (coeff f (Suc j + 2) / coeff f (Suc j + 1)) \<le> (coeff f (h + 1) * coeff f (j + 1)) * (coeff f (Suc j + 1) / coeff f (Suc j))"
- apply (rule mult_mono[OF IH])
- using 3 by (simp_all add: assms(1) normal_coeff_nonneg)
- thus "coeff f h * coeff f (Suc j + 2) \<le> coeff f (h + 1) * coeff f (Suc j + 1)"
- using 1 2 by fastforce
- qed
- qed
- qed
-qed
-
-lemma normal_mult:
- fixes f g::"real poly"
- assumes hf: "normal_poly f" and hg: "normal_poly g"
- defines "df \<equiv> degree f" and "dg \<equiv> degree g"
- shows "normal_poly (f*g)"
-using df_def hf proof (induction df arbitrary: f)
-text \<open>We shall first show that without loss of generality we may assume \<open>poly f 0 \<noteq> 0\<close>,
- this is done by induction on the degree, if 0 is a root then we derive the result from \<open>f/[:0,1:]\<close>.\<close>
- fix f::"real poly" fix i::nat
- assume "0 = degree f" and hf: "normal_poly f"
- then obtain a where "f = [:a:]" using degree_0_iff by auto
- then show "normal_poly (f*g)"
- apply (subst normal_polyI)
- subgoal using normal_non_zero[OF hf] normal_non_zero[OF hg] by auto
- subgoal
- using normal_coeff_nonneg[of _ 0, OF hf] normal_coeff_nonneg[OF hg]
- by simp
- subgoal
- using normal_coeff_nonneg[of _ 0, OF hf] normal_poly_coeff_mult[OF hg]
- by (auto simp: algebra_simps power2_eq_square mult_left_mono)[1]
- subgoal
- using normal_non_zero[OF hf] normal_coeff_nonneg[of _ 0, OF hf] normal_poly_pos_interval[OF hg]
- by (simp add: zero_less_mult_iff)
- subgoal by simp
- done
-next
- case (Suc df)
- then show ?case
- proof (cases "poly f 0 = 0")
- assume "poly f 0 = 0" and hf:"normal_poly f"
- moreover then obtain f' where hdiv: "f = f'*[:0,1:]"
- by (smt (verit) dvdE mult.commute poly_eq_0_iff_dvd)
- ultimately have hf': "normal_poly f'" using normal_divide_x by blast
- assume "Suc df = degree f"
- hence "degree f' = df" using hdiv normal_non_zero[OF hf'] by (auto simp: degree_mult_eq)
- moreover assume "\<And>f. df = degree f \<Longrightarrow> normal_poly f \<Longrightarrow> normal_poly (f * g)"
- ultimately have "normal_poly (f'*g)" using hf' by blast
- thus "normal_poly (f*g)" using hdiv normal_mult_x by fastforce
- next
- assume hf: "normal_poly f" and hf0: "poly f 0 \<noteq> 0"
- define dg where "dg \<equiv> degree g"
- show "normal_poly (f * g)"
- using dg_def hg proof (induction dg arbitrary: g)
- text \<open>Similarly we may assume \<open>poly g 0 \<noteq> 0\<close>.\<close>
- fix g::"real poly" fix i::nat
- assume "0 = degree g" and hg: "normal_poly g"
- then obtain a where "g = [:a:]" using degree_0_iff by auto
- then show "normal_poly (f*g)"
- apply (subst normal_polyI)
- subgoal
- using normal_non_zero[OF hg] normal_non_zero[OF hf] by auto
- subgoal
- using normal_coeff_nonneg[of _ 0, OF hg] normal_coeff_nonneg[OF hf]
- by simp
- subgoal
- using normal_coeff_nonneg[of _ 0, OF hg] normal_poly_coeff_mult[OF hf]
- by (auto simp: algebra_simps power2_eq_square mult_left_mono)
- subgoal
- using normal_non_zero[OF hf] normal_coeff_nonneg[of _ 0, OF hg]
- normal_poly_pos_interval[OF hf]
- by (simp add: zero_less_mult_iff)
- by simp
- next
- case (Suc dg)
- then show ?case
- proof (cases "poly g 0 = 0")
- assume "poly g 0 = 0" and hg:"normal_poly g"
- moreover then obtain g' where hdiv: "g = g'*[:0,1:]"
- by (smt (verit) dvdE mult.commute poly_eq_0_iff_dvd)
- ultimately have hg': "normal_poly g'" using normal_divide_x by blast
- assume "Suc dg = degree g"
- hence "degree g' = dg" using hdiv normal_non_zero[OF hg'] by (auto simp: degree_mult_eq)
- moreover assume "\<And>g. dg = degree g \<Longrightarrow> normal_poly g \<Longrightarrow> normal_poly (f * g)"
- ultimately have "normal_poly (f*g')" using hg' by blast
- thus "normal_poly (f*g)" using hdiv normal_mult_x by fastforce
- next
- text \<open>It now remains to show that $(fg)_i \geq 0$. This follows by decomposing $\{(h, j) \in
- \mathbf{Z}^2 | h > j\} = \{(h, j) \in \mathbf{Z}^2 | h \leq j\} \cup \{(h, h - 1) \in
- \mathbf{Z}^2 | h \in \mathbf{Z}\}$.
- Note in order to avoid working with infinite sums over integers all these sets are
- bounded, which adds some complexity compared to the proof of lemma 2.55 in
- Basu, S., Pollack, R., Roy, M.-F.: Algorithms in Real Algebraic Geometry.
- Springer Berlin Heidelberg, Berlin, Heidelberg (2016).\<close>
- assume hg0: "poly g 0 \<noteq> 0" and hg: "normal_poly g"
- have "f * g \<noteq> 0" using hf hg by (simp add: normal_non_zero Suc.prems)
- moreover have "\<And>i. coeff (f*g) i \<ge> 0"
- apply (subst coeff_mult, rule sum_nonneg, rule mult_nonneg_nonneg)
- using normal_coeff_nonneg[OF hf] normal_coeff_nonneg[OF hg] by auto
- moreover have "
- coeff (f*g) i * coeff (f*g) (i+2) \<le> (coeff (f*g) (i+1))^2" for i
- proof -
-
- text \<open>$(fg)_{i+1}^2 - (fg)_i(fg)_{i+2} = \left(\sum_x f_xg_{i+1-x}\right)^2 -
- \left(\sum_x f_xg_{i+2-x}\right)\left(\sum_x f_xg_{i-x}\right)$\<close>
- have "(coeff (f*g) (i+1))^2 - coeff (f*g) i * coeff (f*g) (i+2) =
- (\<Sum>x\<le>i+1. coeff f x * coeff g (i + 1 - x)) *
- (\<Sum>x\<le>i+1. coeff f x * coeff g (i + 1 - x)) -
- (\<Sum>x\<le>i+2. coeff f x * coeff g (i + 2 - x)) *
- (\<Sum>x\<le>i. coeff f x * coeff g (i - x))"
- by (auto simp: coeff_mult power2_eq_square algebra_simps)
-
- text \<open>$\dots = \sum_{x, y} f_xg_{i+1-x}f_yg_{i+1-y} - \sum_{x, y} f_xg_{i+2-x}f_yg_{i-y}$\<close>
- also have "... =
- (\<Sum>x\<le>i+1. \<Sum>y\<le>i+1. coeff f x * coeff g (i + 1 - x) *
- coeff f y * coeff g (i + 1 - y)) -
- (\<Sum>x\<le>i+2. \<Sum>y\<le>i. coeff f x * coeff g (i + 2 - x) *
- coeff f y * coeff g (i - y))"
- by (subst sum_product, subst sum_product, auto simp: algebra_simps)
-
- text \<open>$\dots = \sum_{h \leq j} f_hg_{i+1-h}f_jg_{i+1-j} + \sum_{h>j} f_hg_{i+1-h}f_jg_{i+1-j}
- - \sum_{h \leq j} f_hg_{i+2-h}f_jg_{i-j} - \sum_{h>j} f_hg_{i+2-h}f_jg_{i-j}$\<close>
- also have "... =
- (\<Sum>(h, j)\<in>{(h, j). i+1 \<ge> j \<and> j \<ge> h}.
- coeff f h * coeff g (i + 1 - h) * coeff f j * coeff g (i + 1 - j)) +
- (\<Sum>(h, j)\<in>{(h, j). i+1 \<ge> h \<and> h > j}.
- coeff f h * coeff g (i + 1 - h) * coeff f j * coeff g (i + 1 - j)) -
- ((\<Sum>(h, j)\<in>{(h, j). i \<ge> j \<and> j \<ge> h}.
- coeff f h * coeff g (i + 2 - h) * coeff f j * coeff g (i - j)) +
- (\<Sum>(h, j)\<in>{(h, j). i + 2 \<ge> h \<and> h > j \<and> i \<ge> j}.
- coeff f h * coeff g (i + 2 - h) * coeff f j * coeff g (i - j)))"
- proof -
- have "(\<Sum>x\<le>i + 1. \<Sum>y\<le>i + 1. coeff f x * coeff g (i + 1 - x) * coeff f y * coeff g (i + 1 - y)) =
- (\<Sum>(h, j)\<in>{(h, j). j \<le> i + 1 \<and> h \<le> j}.
- coeff f h * coeff g (i + 1 - h) * coeff f j * coeff g (i + 1 - j)) +
- (\<Sum>(h, j)\<in>{(h, j). h \<le> i + 1 \<and> j < h}.
- coeff f h * coeff g (i + 1 - h) * coeff f j * coeff g (i + 1 - j))"
- proof (subst sum.union_disjoint[symmetric])
- have H:"{(h, j). j \<le> i + 1 \<and> h \<le> j} \<subseteq> {..i+1} \<times> {..i+1}"
- "{(h, j). h \<le> i + 1 \<and> j < h} \<subseteq> {..i+1} \<times> {..i+1}"
- "finite ({..i+1} \<times> {..i+1})"
- by (fastforce, fastforce, fastforce)
- show "finite {(h, j). j \<le> i + 1 \<and> h \<le> j}"
- apply (rule finite_subset) using H by (blast, blast)
- show "finite {(h, j). h \<le> i + 1 \<and> j < h}"
- apply (rule finite_subset) using H by (blast, blast)
- show "{(h, j). j \<le> i + 1 \<and> h \<le> j} \<inter> {(h, j). h \<le> i + 1 \<and> j < h} = {}"
- by fastforce
- show "(\<Sum>x\<le>i + 1. \<Sum>y\<le>i + 1. coeff f x * coeff g (i + 1 - x) * coeff f y * coeff g (i + 1 - y)) =
- (\<Sum>(h, j)\<in>{(h, j). j \<le> i + 1 \<and> h \<le> j} \<union> {(h, j). h \<le> i + 1 \<and> j < h}.
- coeff f h * coeff g (i + 1 - h) * coeff f j * coeff g (i + 1 - j))"
- apply (subst sum.cartesian_product, rule sum.cong)
- apply force by blast
- qed
- moreover have "(\<Sum>x\<le>i + 2. \<Sum>y\<le>i. coeff f x * coeff g (i + 2 - x) * coeff f y * coeff g (i - y)) =
- (\<Sum>(h, j)\<in>{(h, j). j \<le> i \<and> h \<le> j}.
- coeff f h * coeff g (i + 2 - h) * coeff f j * coeff g (i - j)) +
- (\<Sum>(h, j)\<in>{(h, j). i + 2 \<ge> h \<and> h > j \<and> i \<ge> j}.
- coeff f h * coeff g (i + 2 - h) * coeff f j * coeff g (i - j))"
- proof (subst sum.union_disjoint[symmetric])
- have H:"{(h, j). j \<le> i \<and> h \<le> j} \<subseteq> {..i+2} \<times> {..i}"
- "{(h, j). i + 2 \<ge> h \<and> h > j \<and> i \<ge> j} \<subseteq> {..i+2} \<times> {..i}"
- "finite ({..i+2} \<times> {..i})"
- by (fastforce, fastforce, fastforce)
- show "finite {(h, j). j \<le> i \<and> h \<le> j}"
- apply (rule finite_subset) using H by (blast, blast)
- show "finite {(h, j). i + 2 \<ge> h \<and> h > j \<and> i \<ge> j}"
- apply (rule finite_subset) using H by (blast, blast)
- show "{(h, j). j \<le> i \<and> h \<le> j} \<inter> {(h, j). i + 2 \<ge> h \<and> h > j \<and> i \<ge> j} = {}"
- by fastforce
- show "(\<Sum>x\<le>i + 2. \<Sum>y\<le>i. coeff f x * coeff g (i + 2 - x) * coeff f y * coeff g (i - y)) =
- (\<Sum>(h, j)\<in>{(h, j). j \<le> i \<and> h \<le> j} \<union> {(h, j). i + 2 \<ge> h \<and> h > j \<and> i \<ge> j}.
- coeff f h * coeff g (i + 2 - h) * coeff f j * coeff g (i - j))"
- apply (subst sum.cartesian_product, rule sum.cong)
- apply force by blast
- qed
- ultimately show ?thesis by presburger
- qed
-
- text \<open>$\dots = \sum_{h \leq j} f_hg_{i+1-h}f_jg_{i+1-j} + \sum_{h \leq j} f_{j+1}g_{i-j}f_{h-2}g_{i+2-h}
- + \sum_h f_hg_{i+1-h}f_{h-1}g_{i+2-h} - \sum_{h \leq j} f_hg_{i+2-h}f_jg_{i-j}
- - \sum_{h \leq j} f_{j+1}g_{i+1-j}f_{h-2}g_{i+1-h} - \sum_h f_hg_{i+2-h}f_{h-1}g_{i+1-h}$\<close>
- also have "... =
- (\<Sum>(h, j)\<in>{(h, j). j \<le> i + 1 \<and> h \<le> j}.
- coeff f h * coeff g (i + 1 - h) * coeff f j * coeff g (i + 1 - j)) +
- (\<Sum>(h, j)\<in>{(h, j). j \<le> i \<and> h \<le> j \<and> 0 < h}.
- coeff f (j+1) * coeff g (i - j) * coeff f (h-1) * coeff g (i + 2 - h)) +
- (\<Sum>h\<in>{1..i+1}.
- coeff f h * coeff g (i + 1 - h) * coeff f (h-1) * coeff g (i + 2 - h)) -
- ((\<Sum>(h, j)\<in>{(h, j). j \<le> i \<and> h \<le> j}.
- coeff f h * coeff g (i + 2 - h) * coeff f j * coeff g (i - j)) +
- (\<Sum>(h, j)\<in>{(h, j). j \<le> i + 1 \<and> h \<le> j \<and> 0 < h}.
- coeff f (j+1) * coeff g (i + 1 - j) * coeff f (h-1) * coeff g (i + 1 - h)) +
- (\<Sum>h\<in>{1..i+1}.
- coeff f h * coeff g (i + 2 - h) * coeff f (h-1) * coeff g (i + 1 - h)))"
- proof -
- have "(\<Sum>(h, j)\<in>{(h, j). h \<le> i + 1 \<and> j < h}.
- coeff f h * coeff g (i + 1 - h) * coeff f j * coeff g (i + 1 - j)) =
- (\<Sum>(h, j)\<in>{(h, j). j \<le> i \<and> h \<le> j \<and> 0 < h}.
- coeff f (j + 1) * coeff g (i - j) * coeff f (h - 1) * coeff g (i + 2 - h)) +
- (\<Sum>h = 1..i + 1. coeff f h * coeff g (i + 1 - h) * coeff f (h - 1) * coeff g (i + 2 - h))"
- proof -
- have 1: "(\<Sum>(h, j)\<in>{(h, j). j \<le> i \<and> h \<le> j \<and> 0 < h}.
- coeff f (j + 1) * coeff g (i - j) * coeff f (h - 1) * coeff g (i + 2 - h)) =
- (\<Sum>(h, j)\<in>{(h, j). h \<le> i + 1 \<and> j < h \<and> h \<noteq> j + 1}.
- coeff f h * coeff g (i + 1 - h) * coeff f j * coeff g (i + 1 - j))"
- proof (rule sum.reindex_cong)
- show "{(h, j). j \<le> i \<and> h \<le> j \<and> 0 < h} = (\<lambda>(h, j). (j+1, h-1)) ` {(h, j). h \<le> i + 1 \<and> j < h \<and> h \<noteq> j + 1}"
- proof
- show "(\<lambda>(h, j). (j + 1, h - 1)) ` {(h, j). h \<le> i + 1 \<and> j < h \<and> h \<noteq> j + 1} \<subseteq> {(h, j). j \<le> i \<and> h \<le> j \<and> 0 < h}"
- by fastforce
- show "{(h, j). j \<le> i \<and> h \<le> j \<and> 0 < h} \<subseteq> (\<lambda>(h, j). (j + 1, h - 1)) ` {(h, j). h \<le> i + 1 \<and> j < h \<and> h \<noteq> j + 1}"
- proof
- fix x
- assume "x \<in> {(h, j). j \<le> i \<and> h \<le> j \<and> 0 < h}"
- then obtain h j where "x = (h, j)" "j \<le> i" "h \<le> j" "0 < h" by blast
- hence "j + 1 \<le> i + 1 \<and> h - 1 < j + 1 \<and> j + 1 \<noteq> h - 1 + 1 \<and> x = ((h - 1) + 1, (j + 1) - 1)"
- by auto
- thus "x \<in> (\<lambda>(h, j). (j + 1, h - 1)) ` {(h, j). h \<le> i + 1 \<and> j < h \<and> h \<noteq> j + 1}"
- by (auto simp: image_iff)
- qed
- qed
- show "inj_on (\<lambda>(h, j). (j + 1, h - 1)) {(h, j). h \<le> i + 1 \<and> j < h \<and> h \<noteq> j + 1}"
- proof
- fix x y::"nat\<times>nat"
- assume "x \<in> {(h, j). h \<le> i + 1 \<and> j < h \<and> h \<noteq> j + 1}" "y \<in> {(h, j). h \<le> i + 1 \<and> j < h \<and> h \<noteq> j + 1}"
- thus "(case x of (h, j) \<Rightarrow> (j + 1, h - 1)) = (case y of (h, j) \<Rightarrow> (j + 1, h - 1)) \<Longrightarrow> x = y"
- by auto
- qed
- show "\<And>x. x \<in> {(h, j). h \<le> i + 1 \<and> j < h \<and> h \<noteq> j + 1} \<Longrightarrow>
- (case case x of (h, j) \<Rightarrow> (j + 1, h - 1) of
- (h, j) \<Rightarrow> coeff f (j + 1) * coeff g (i - j) * coeff f (h - 1) * coeff g (i + 2 - h)) =
- (case x of (h, j) \<Rightarrow> coeff f h * coeff g (i + 1 - h) * coeff f j * coeff g (i + 1 - j))"
- by fastforce
- qed
- have 2: "(\<Sum>h = 1..i + 1. coeff f h * coeff g (i + 1 - h) * coeff f (h - 1) * coeff g (i + 2 - h)) =
- (\<Sum>(h, j)\<in>{(h, j). h \<le> i + 1 \<and> j < h \<and> h = j + 1}.
- coeff f h * coeff g (i + 1 - h) * coeff f j * coeff g (i + 1 - j))"
- proof (rule sum.reindex_cong)
- show "{1..i + 1} = fst ` {(h, j). h \<le> i + 1 \<and> j < h \<and> h = j + 1}"
- proof
- show "{1..i + 1} \<subseteq> fst ` {(h, j). h \<le> i + 1 \<and> j < h \<and> h = j + 1}"
- proof
- fix x
- assume "x \<in> {1..i + 1}"
- hence "x \<le> i + 1 \<and> x - 1 < x \<and> x = x - 1 + 1 \<and> x = fst (x, x-1)"
- by auto
- thus "x \<in> fst ` {(h, j). h \<le> i + 1 \<and> j < h \<and> h = j + 1}"
- by blast
- qed
- show "fst ` {(h, j). h \<le> i + 1 \<and> j < h \<and> h = j + 1} \<subseteq> {1..i + 1}"
- by force
- qed
- show "inj_on fst {(h, j). h \<le> i + 1 \<and> j < h \<and> h = j + 1}"
- proof
- fix x y
- assume "x \<in> {(h, j). h \<le> i + 1 \<and> j < h \<and> h = j + 1}"
- "y \<in> {(h, j). h \<le> i + 1 \<and> j < h \<and> h = j + 1}"
- hence "x = (fst x, fst x - 1)" "y = (fst y, fst y - 1)" "fst x > 0" "fst y > 0"
- by auto
- thus "fst x = fst y \<Longrightarrow> x = y" by presburger
- qed
- show "\<And>x. x \<in> {(h, j). h \<le> i + 1 \<and> j < h \<and> h = j + 1} \<Longrightarrow>
- coeff f (fst x) * coeff g (i + 1 - fst x) * coeff f (fst x - 1) * coeff g (i + 2 - fst x) =
- (case x of (h, j) \<Rightarrow> coeff f h * coeff g (i + 1 - h) * coeff f j * coeff g (i + 1 - j))"
- by fastforce
- qed
- have H: "{(h, j). h \<le> i + 1 \<and> j < h \<and> h \<noteq> j + 1} \<subseteq> {0..i+1}\<times>{0..i+1}"
- "{(h, j). h \<le> i + 1 \<and> j < h \<and> h = j + 1} \<subseteq> {0..i+1}\<times>{0..i+1}"
- "finite ({0..i+1}\<times>{0..i+1})"
- by (fastforce, fastforce, fastforce)
- have "finite {(h, j). h \<le> i + 1 \<and> j < h \<and> h \<noteq> j + 1}"
- "finite {(h, j). h \<le> i + 1 \<and> j < h \<and> h = j + 1}"
- apply (rule finite_subset) using H apply (simp, simp)
- apply (rule finite_subset) using H apply (simp, simp)
- done
- thus ?thesis
- apply (subst 1, subst 2, subst sum.union_disjoint[symmetric])
- apply auto[3]
- apply (rule sum.cong)
- by auto
- qed
- moreover have "(\<Sum>(h, j)\<in>{(h, j). h \<le> i + 2 \<and> j < h \<and> j \<le> i}.
- coeff f h * coeff g (i + 2 - h) * coeff f j * coeff g (i - j)) =
- (\<Sum>(h, j)\<in>{(h, j). j \<le> i + 1 \<and> h \<le> j \<and> 0 < h}.
- coeff f (j + 1) * coeff g (i + 1 - j) * coeff f (h - 1) * coeff g (i + 1 - h)) +
- (\<Sum>h = 1..i + 1. coeff f h * coeff g (i + 2 - h) * coeff f (h - 1) * coeff g (i + 1 - h))"
- proof -
- have 1: "(\<Sum>(h, j)\<in>{(h, j). j \<le> i + 1 \<and> h \<le> j \<and> 0 < h}.
- coeff f (j + 1) * coeff g (i + 1 - j) * coeff f (h - 1) * coeff g (i + 1 - h)) =
- (\<Sum>(h, j)\<in>{(h, j). h \<le> i + 2 \<and> j < h \<and> j \<le> i \<and> h \<noteq> j + 1}.
- coeff f h * coeff g (i + 2 - h) * coeff f j * coeff g (i - j))"
- proof (rule sum.reindex_cong)
- show "{(h, j). j \<le> i + 1 \<and> h \<le> j \<and> 0 < h} = (\<lambda>(h, j). (j+1, h-1)) ` {(h, j). h \<le> i + 2 \<and> j < h \<and> j \<le> i \<and> h \<noteq> j + 1}"
- proof
- show "(\<lambda>(h, j). (j + 1, h - 1)) ` {(h, j). h \<le> i + 2 \<and> j < h \<and> j \<le> i \<and> h \<noteq> j + 1} \<subseteq> {(h, j). j \<le> i + 1 \<and> h \<le> j \<and> 0 < h}"
- by fastforce
- show "{(h, j). j \<le> i + 1 \<and> h \<le> j \<and> 0 < h} \<subseteq> (\<lambda>(h, j). (j + 1, h - 1)) ` {(h, j). h \<le> i + 2 \<and> j < h \<and> j \<le> i \<and> h \<noteq> j + 1}"
- proof
- fix x
- assume "x \<in> {(h, j). j \<le> i + 1 \<and> h \<le> j \<and> 0 < h}"
- then obtain h j where "x = (h, j)" "j \<le> i + 1" "h \<le> j" "0 < h" by blast
- hence "j + 1 \<le> i + 2 \<and> h - 1 < j + 1 \<and> h - 1 \<le> i \<and> j + 1 \<noteq> h - 1 + 1 \<and> x = ((h - 1) + 1, (j + 1) - 1)"
- by auto
- thus "x \<in> (\<lambda>(h, j). (j + 1, h - 1)) ` {(h, j). h \<le> i + 2 \<and> j < h \<and> j \<le> i \<and> h \<noteq> j + 1}"
- by (auto simp: image_iff)
- qed
- qed
- show "inj_on (\<lambda>(h, j). (j + 1, h - 1)) {(h, j). h \<le> i + 2 \<and> j < h \<and> j \<le> i \<and> h \<noteq> j + 1}"
- proof
- fix x y::"nat\<times>nat"
- assume "x \<in> {(h, j). h \<le> i + 2 \<and> j < h \<and> j \<le> i \<and> h \<noteq> j + 1}" "y \<in> {(h, j). h \<le> i + 2 \<and> j < h \<and> j \<le> i \<and> h \<noteq> j + 1}"
- thus "(case x of (h, j) \<Rightarrow> (j + 1, h - 1)) = (case y of (h, j) \<Rightarrow> (j + 1, h - 1)) \<Longrightarrow> x = y"
- by auto
- qed
- show "\<And>x. x \<in> {(h, j). h \<le> i + 2 \<and> j < h \<and> j \<le> i \<and> h \<noteq> j + 1} \<Longrightarrow>
- (case case x of (h, j) \<Rightarrow> (j + 1, h - 1) of
- (h, j) \<Rightarrow> coeff f (j + 1) * coeff g (i + 1 - j) * coeff f (h - 1) * coeff g (i + 1 - h)) =
- (case x of (h, j) \<Rightarrow> coeff f h * coeff g (i + 2 - h) * coeff f j * coeff g (i - j))"
- by fastforce
- qed
- have 2: "(\<Sum>h = 1..i + 1. coeff f h * coeff g (i + 2 - h) * coeff f (h - 1) * coeff g (i + 1 - h)) =
- (\<Sum>(h, j)\<in>{(h, j). h \<le> i + 2 \<and> j < h \<and> j \<le> i \<and> h = j + 1}.
- coeff f h * coeff g (i + 2 - h) * coeff f j * coeff g (i - j))"
- proof (rule sum.reindex_cong)
- show "{1..i + 1} = fst ` {(h, j). h \<le> i + 2 \<and> j < h \<and> j \<le> i \<and> h = j + 1}"
- proof
- show "{1..i + 1} \<subseteq> fst ` {(h, j). h \<le> i + 2 \<and> j < h \<and> j \<le> i \<and> h = j + 1}"
- proof
- fix x
- assume "x \<in> {1..i + 1}"
- hence "x \<le> i + 2 \<and> x - 1 < x \<and> x - 1 \<le> i \<and> x = x - 1 + 1 \<and> x = fst (x, x-1)"
- by auto
- thus "x \<in> fst ` {(h, j). h \<le> i + 2 \<and> j < h \<and> j \<le> i \<and> h = j + 1}"
- by blast
- qed
- show "fst ` {(h, j). h \<le> i + 2 \<and> j < h \<and> j \<le> i \<and> h = j + 1} \<subseteq> {1..i + 1}"
- by force
- qed
- show "inj_on fst {(h, j). h \<le> i + 2 \<and> j < h \<and> j \<le> i \<and> h = j + 1}"
- proof
- fix x y
- assume "x \<in> {(h, j). h \<le> i + 2 \<and> j < h \<and> j \<le> i \<and> h = j + 1}"
- "y \<in> {(h, j). h \<le> i + 2 \<and> j < h \<and> j \<le> i \<and> h = j + 1}"
- hence "x = (fst x, fst x - 1)" "y = (fst y, fst y - 1)" "fst x > 0" "fst y > 0"
- by auto
- thus "fst x = fst y \<Longrightarrow> x = y" by presburger
- qed
- show "\<And>x. x \<in> {(h, j). h \<le> i + 2 \<and> j < h \<and> j \<le> i \<and> h = j + 1} \<Longrightarrow>
- coeff f (fst x) * coeff g (i + 2 - fst x) * coeff f (fst x - 1) * coeff g (i + 1 - fst x) =
- (case x of (h, j) \<Rightarrow> coeff f h * coeff g (i + 2 - h) * coeff f j * coeff g (i - j))"
- by fastforce
- qed
- have H: "{(h, j). h \<le> i + 2 \<and> j < h \<and> j \<le> i \<and> h \<noteq> j + 1} \<subseteq> {0..i+2}\<times>{0..i}"
- "{(h, j). h \<le> i + 2 \<and> j < h \<and> j \<le> i \<and> h = j + 1} \<subseteq> {0..i+2}\<times>{0..i}"
- "finite ({0..i+2}\<times>{0..i})"
- by (fastforce, fastforce, fastforce)
- have "finite {(h, j). h \<le> i + 2 \<and> j < h \<and> j \<le> i \<and> h \<noteq> j + 1}"
- "finite {(h, j). h \<le> i + 2 \<and> j < h \<and> j \<le> i \<and> h = j + 1}"
- apply (rule finite_subset) using H apply (simp, simp)
- apply (rule finite_subset) using H apply (simp, simp)
- done
- thus ?thesis
- apply (subst 1, subst 2, subst sum.union_disjoint[symmetric])
- apply auto[3]
- apply (rule sum.cong)
- by auto
- qed
- ultimately show ?thesis
- by algebra
- qed
-
- text \<open>$\dots = \sum_{h \leq j} f_hf_j\left(g_{i+1-h}g_{i+1-j} - g_{i+2-h}g_{i-j}\right) +
- \sum_{h \leq j} f_{j+1}f_{h-1}\left(g_{i-j}g_{i+2-h} - g_{i+1-j}f_jg_{i+1-h}\right)$
-
- Note we have to also consider the edge cases caused by making these sums finite.\<close>
- also have "... =
- (\<Sum>(h, j)\<in>{(h, j). j = i + 1 \<and> h \<le> j}.
- coeff f h * coeff f j * (coeff g (i + 1 - h) * coeff g (i + 1 - j))) +
- (\<Sum>(h, j)\<in>{(h, j). j \<le> i \<and> h \<le> j}.
- coeff f h * coeff f j * (coeff g (i + 1 - h) * coeff g (i + 1 - j) - coeff g (i + 2 - h) * coeff g (i - j))) +
- (\<Sum>(h, j)\<in>{(h, j). j \<le> i \<and> h \<le> j \<and> 0 < h}.
- coeff f (j+1) * coeff f (h-1) * (coeff g (i - j) * coeff g (i + 2 - h) - coeff g (i + 1 - j) * coeff g (i + 1 - h))) -
- ((\<Sum>(h, j)\<in>{(h, j). j = i + 1 \<and> h \<le> j \<and> 0 < h}.
- coeff f (j+1) * coeff g (i + 1 - j) * coeff f (h-1) * coeff g (i + 1 - h)))" (is "?L = ?R")
- proof -
- have "?R =
- (\<Sum>(h, j)\<in>{(h, j). j = i + 1 \<and> h \<le> j}.
- coeff f h * coeff f j * (coeff g (i + 1 - h) * coeff g (i + 1 - j))) +
- ((\<Sum>(h, j)\<in>{(h, j). j \<le> i \<and> h \<le> j}.
- coeff f h * coeff f j * coeff g (i + 1 - h) * coeff g (i + 1 - j)) -
- (\<Sum>(h, j)\<in>{(h, j). j \<le> i \<and> h \<le> j}.
- coeff f h * coeff f j * coeff g (i + 2 - h) * coeff g (i - j))) +
- ((\<Sum>(h, j)\<in>{(h, j). j \<le> i \<and> h \<le> j \<and> 0 < h}.
- coeff f (j+1) * coeff f (h-1) * coeff g (i - j) * coeff g (i + 2 - h)) -
- (\<Sum>(h, j)\<in>{(h, j). j \<le> i \<and> h \<le> j \<and> 0 < h}.
- coeff f (j+1) * coeff f (h-1) * coeff g (i + 1 - j) * coeff g (i + 1 - h))) -
- ((\<Sum>(h, j)\<in>{(h, j). j = i + 1 \<and> h \<le> j \<and> 0 < h}.
- coeff f (j+1) * coeff g (i + 1 - j) * coeff f (h-1) * coeff g (i + 1 - h)))"
- apply (subst sum_subtractf[symmetric], subst sum_subtractf[symmetric])
- by (auto simp: algebra_simps split_beta)
- also have "... =
- ((\<Sum>(h, j)\<in>{(h, j). j = i + 1 \<and> h \<le> j}.
- coeff f h * coeff f j * (coeff g (i + 1 - h) * coeff g (i + 1 - j))) +
- (\<Sum>(h, j)\<in>{(h, j). j \<le> i \<and> h \<le> j}.
- coeff f h * coeff f j * (coeff g (i + 1 - h) * coeff g (i + 1 - j)))) -
- (\<Sum>(h, j)\<in>{(h, j). j \<le> i \<and> h \<le> j}.
- coeff f h * coeff f j * coeff g (i + 2 - h) * coeff g (i - j)) +
- (\<Sum>(h, j)\<in>{(h, j). j \<le> i \<and> h \<le> j \<and> 0 < h}.
- coeff f (j + 1) * coeff f (h - 1) * coeff g (i - j) * coeff g (i + 2 - h)) -
- ((\<Sum>(h, j)\<in>{(h, j). j \<le> i \<and> h \<le> j \<and> 0 < h}.
- coeff f (j + 1) * coeff g (i + 1 - j) * coeff f (h - 1) * coeff g (i + 1 - h)) +
- (\<Sum>(h, j)\<in>{(h, j). j = i + 1 \<and> h \<le> j \<and> 0 < h}.
- coeff f (j + 1) * coeff g (i + 1 - j) * coeff f (h - 1) * coeff g (i + 1 - h)))"
- by (auto simp: algebra_simps)
- also have "... = ?L"
- proof -
- have "(\<Sum>(h, j)\<in>{(h, j). j = i + 1 \<and> h \<le> j}.
- coeff f h * coeff f j * (coeff g (i + 1 - h) * coeff g (i + 1 - j))) +
- (\<Sum>(h, j)\<in>{(h, j). j \<le> i \<and> h \<le> j}.
- coeff f h * coeff f j * (coeff g (i + 1 - h) * coeff g (i + 1 - j))) =
- (\<Sum>(h, j)\<in>{(h, j). j \<le> i + 1 \<and> h \<le> j}.
- coeff f h * coeff f j * (coeff g (i + 1 - h) * coeff g (i + 1 - j)))"
- proof (subst sum.union_disjoint[symmetric])
- have "{(h, j). j = i + 1 \<and> h \<le> j} \<subseteq> {..i + 1} \<times> {..i + 1}"
- "{(h, j). j \<le> i \<and> h \<le> j} \<subseteq> {..i + 1} \<times> {..i + 1}"
- by (fastforce, fastforce)
- thus "finite {(h, j). j = i + 1 \<and> h \<le> j}" "finite {(h, j). j \<le> i \<and> h \<le> j}"
- by (auto simp: finite_subset)
- show "{(h, j). j = i + 1 \<and> h \<le> j} \<inter> {(h, j). j \<le> i \<and> h \<le> j} = {}"
- by fastforce
- qed (rule sum.cong, auto)
- moreover have "(\<Sum>(h, j)\<in>{(h, j). j \<le> i \<and> h \<le> j \<and> 0 < h}.
- coeff f (j + 1) * coeff g (i + 1 - j) * coeff f (h - 1) * coeff g (i + 1 - h)) +
- (\<Sum>(h, j)\<in>{(h, j). j = i + 1 \<and> h \<le> j \<and> 0 < h}.
- coeff f (j + 1) * coeff g (i + 1 - j) * coeff f (h - 1) * coeff g (i + 1 - h)) =
- (\<Sum>(h, j)\<in>{(h, j). j \<le> i + 1 \<and> h \<le> j \<and> 0 < h}.
- coeff f (j + 1) * coeff g (i + 1 - j) * coeff f (h - 1) * coeff g (i + 1 - h))"
- proof (subst sum.union_disjoint[symmetric])
- have "{(h, j). j \<le> i \<and> h \<le> j \<and> 0 < h} \<subseteq> {..i + 1} \<times> {..i + 1}"
- "{(h, j). j = i + 1 \<and> h \<le> j \<and> 0 < h} \<subseteq> {..i + 1} \<times> {..i + 1}"
- by (fastforce, fastforce)
- thus "finite {(h, j). j \<le> i \<and> h \<le> j \<and> 0 < h}" "finite {(h, j). j = i + 1 \<and> h \<le> j \<and> 0 < h}"
- by (auto simp: finite_subset)
- show "{(h, j). j \<le> i \<and> h \<le> j \<and> 0 < h} \<inter> {(h, j). j = i + 1 \<and> h \<le> j \<and> 0 < h} = {}"
- by fastforce
- qed (rule sum.cong, auto)
- ultimately show ?thesis
- by (auto simp: algebra_simps)
- qed
- finally show ?thesis by presburger
- qed
-
- text \<open>$\dots = \sum_{h \leq j} \left(f_hf_j - f_{j+1}f_{h-1}\right)
- \left(g_{i+1-h}g_{i+1-j} - g_{i+2-h}g_{i-j}\right)$\<close>
- also have "... =
- (\<Sum>(h, j)\<in>{(h, j). j \<le> i \<and> h \<le> j \<and> 0 < h}.
- -(coeff f h * coeff f j - coeff f (j+1) * coeff f (h-1)) * (coeff g (i - j) * coeff g (i + 2 - h) - coeff g (i + 1 - j) * coeff g (i + 1 - h))) +
- (\<Sum>(h, j)\<in>{(h, j). j \<le> i \<and> h \<le> j \<and> h = 0}.
- coeff f h * coeff f j * (coeff g (i + 1 - h) * coeff g (i + 1 - j) - coeff g (i + 2 - h) * coeff g (i - j))) +
- (\<Sum>(h, j)\<in>{(h, j). j = i + 1 \<and> h \<le> j}.
- coeff f h * coeff f j * (coeff g (i + 1 - h) * coeff g (i + 1 - j))) -
- ((\<Sum>(h, j)\<in>{(h, j). j = i + 1 \<and> h \<le> j \<and> 0 < h}.
- coeff f (j+1) * coeff g (i + 1 - j) * coeff f (h-1) * coeff g (i + 1 - h)))" (is "?L = ?R")
- proof -
- have "(\<Sum>(h, j)\<in>{(h, j). j \<le> i \<and> h \<le> j}.
- coeff f h * coeff f j *
- (coeff g (i + 1 - h) * coeff g (i + 1 - j) - coeff g (i + 2 - h) * coeff g (i - j))) =
- (\<Sum>(h, j)\<in>{(h, j). j \<le> i \<and> h \<le> j \<and> 0 < h}.
- coeff f h * coeff f j *
- (coeff g (i + 1 - h) * coeff g (i + 1 - j) - coeff g (i + 2 - h) * coeff g (i - j))) +
- (\<Sum>(h, j)\<in>{(h, j). j \<le> i \<and> h \<le> j \<and> 0 = h}.
- coeff f h * coeff f j *
- (coeff g (i + 1 - h) * coeff g (i + 1 - j) - coeff g (i + 2 - h) * coeff g (i - j)))"
- proof (subst sum.union_disjoint[symmetric])
- have "{(h, j). j \<le> i \<and> h \<le> j \<and> 0 < h} \<subseteq> {..i}\<times>{..i}" "{(h, j). j \<le> i \<and> h \<le> j \<and> 0 = h} \<subseteq> {..i}\<times>{..i}"
- by (force, force)
- thus "finite {(h, j). j \<le> i \<and> h \<le> j \<and> 0 < h}" "finite {(h, j). j \<le> i \<and> h \<le> j \<and> 0 = h}"
- by (auto simp: finite_subset)
- show "{(h, j). j \<le> i \<and> h \<le> j \<and> 0 < h} \<inter> {(h, j). j \<le> i \<and> h \<le> j \<and> 0 = h} = {}"
- by fast
- qed (rule sum.cong, auto)
-
- moreover have "(\<Sum>(h, j)\<in>{(h, j). j \<le> i \<and> h \<le> j \<and> 0 < h}.
- (-coeff f h * coeff f j + coeff f (j + 1) * coeff f (h - 1)) *
- (coeff g (i - j) * coeff g (i + 2 - h) - coeff g (i + 1 - j) * coeff g (i + 1 - h))) =
- (\<Sum>(h, j)\<in>{(h, j). j \<le> i \<and> h \<le> j \<and> 0 < h}.
- coeff f h * coeff f j * (coeff g (i + 1 - h) * coeff g (i + 1 - j) - coeff g (i + 2 - h) * coeff g (i - j))) +
- (\<Sum>(h, j)\<in>{(h, j). j \<le> i \<and> h \<le> j \<and> 0 < h}.
- coeff f (j + 1) * coeff f (h - 1) *
- (coeff g (i - j) * coeff g (i + 2 - h) - coeff g (i + 1 - j) * coeff g (i + 1 - h)))"
- by (subst sum.distrib[symmetric], rule sum.cong, fast, auto simp: algebra_simps)
-
- ultimately show ?thesis
- by (auto simp: algebra_simps)
- qed
-
- text \<open>$\dots \geq 0$ by \<open>normal_poly_general_coeff_mult\<close>\<close>
- also have "... \<ge> 0"
- proof -
- have "0 \<le> (\<Sum>(h, j)\<in>{(h, j). j \<le> i \<and> h \<le> j \<and> 0 < h}.
- - (coeff f h * coeff f j - coeff f (j + 1) * coeff f (h - 1)) *
- (coeff g (i - j) * coeff g (i + 2 - h) - coeff g (i + 1 - j) * coeff g (i + 1 - h)))"
- proof (rule sum_nonneg)
- fix x assume "x \<in> {(h, j). j \<le> i \<and> h \<le> j \<and> 0 < h}"
- then obtain h j where H: "x = (h, j)" "j \<le> i" "h \<le> j" "0 < h" by fast
- hence "h - 1 \<le> j - 1" by force
- hence 1: "coeff f h * coeff f j - coeff f (j + 1) * coeff f (h - 1) \<ge> 0"
- using normal_poly_general_coeff_mult[OF hf, of "h-1" "j-1"] H
- by (auto simp: algebra_simps)
- from H have "i - j \<le> i - h" by force
- hence 2: "coeff g (i - j) * coeff g (i + 2 - h) - coeff g (i + 1 - j) * coeff g (i + 1 - h) \<le> 0"
- using normal_poly_general_coeff_mult[OF hg, of "i - j" "i - h"] H
- by (smt (verit, del_insts) Nat.add_diff_assoc2 le_trans)
- show "0 \<le> (case x of
- (h, j) \<Rightarrow>
- - (coeff f h * coeff f j - coeff f (j + 1) * coeff f (h - 1)) *
- (coeff g (i - j) * coeff g (i + 2 - h) -
- coeff g (i + 1 - j) * coeff g (i + 1 - h)))"
- apply (subst H(1), subst split, rule mult_nonpos_nonpos, subst neg_le_0_iff_le)
- subgoal using 1 by blast
- subgoal using 2 by blast
- done
- qed
- moreover have "0 \<le> (\<Sum>(h, j)\<in>{(h, j). j \<le> i \<and> h \<le> j \<and> h = 0}.
- coeff f h * coeff f j *
- (coeff g (i + 1 - h) * coeff g (i + 1 - j) - coeff g (i + 2 - h) * coeff g (i - j)))"
- proof (rule sum_nonneg)
- fix x assume "x \<in> {(h, j). j \<le> i \<and> h \<le> j \<and> h = 0}"
- then obtain h j where H: "x = (h, j)" "j \<le> i" "h \<le> j" "h = 0" by fast
- have 1: "coeff f h * coeff f j \<ge> 0"
- by (simp add: hf normal_coeff_nonneg)
- from H have "i - j \<le> i - h" by force
- hence 2: "coeff g (i - j) * coeff g (i + 2 - h) - coeff g (i + 1 - j) * coeff g (i + 1 - h) \<le> 0"
- using normal_poly_general_coeff_mult[OF hg, of "i - j" "i - h"] H
- by (smt (verit, del_insts) Nat.add_diff_assoc2 le_trans)
- show "0 \<le> (case x of
- (h, j) \<Rightarrow>
- coeff f h * coeff f j *
- (coeff g (i + 1 - h) * coeff g (i + 1 - j) -
- coeff g (i + 2 - h) * coeff g (i - j)))"
- apply (subst H(1), subst split, rule mult_nonneg_nonneg)
- subgoal using 1 by blast
- subgoal using 2 by argo
- done
- qed
- moreover have "0 \<le> (\<Sum>(h, j)\<in>{(h, j). j = i + 1 \<and> h \<le> j}. coeff f h * coeff f j * (coeff g (i + 1 - h) * coeff g (i + 1 - j))) -
- (\<Sum>(h, j)\<in>{(h, j). j = i + 1 \<and> h \<le> j \<and> 0 < h}.
- coeff f (j + 1) * coeff g (i + 1 - j) * coeff f (h - 1) * coeff g (i + 1 - h))"
- proof -
- have "(\<Sum>(h, j)\<in>{(h, j). j = i + 1 \<and> h \<le> j}. coeff f h * coeff f j * (coeff g (i + 1 - h) * coeff g (i + 1 - j))) -
- (\<Sum>(h, j)\<in>{(h, j). j = i + 1 \<and> h \<le> j \<and> 0 < h}.
- coeff f (j + 1) * coeff g (i + 1 - j) * coeff f (h - 1) * coeff g (i + 1 - h)) =
- (\<Sum>(h, j)\<in>{(h, j). j = i + 1 \<and> h \<le> j \<and> h = 0}. coeff f h * coeff f j * (coeff g (i + 1 - h) * coeff g (i + 1 - j))) +
- (\<Sum>(h, j)\<in>{(h, j). j = i + 1 \<and> h \<le> j \<and> 0 < h}. coeff f h * coeff f j * (coeff g (i + 1 - h) * coeff g (i + 1 - j))) -
- (\<Sum>(h, j)\<in>{(h, j). j = i + 1 \<and> h \<le> j \<and> 0 < h}.
- coeff f (j + 1) * coeff g (i + 1 - j) * coeff f (h - 1) * coeff g (i + 1 - h))"
- proof (subst sum.union_disjoint[symmetric])
- have "{(h, j). j = i + 1 \<and> h \<le> j \<and> h = 0} = {(0, i + 1)}"
- "{(h, j). j = i + 1 \<and> h \<le> j \<and> 0 < h} = {1..i+1} \<times> {i + 1}"
- by (fastforce, force)
- thus "finite {(h, j). j = i + 1 \<and> h \<le> j \<and> h = 0}"
- "finite {(h, j). j = i + 1 \<and> h \<le> j \<and> 0 < h}"
- by auto
- show "{(h, j). j = i + 1 \<and> h \<le> j \<and> h = 0} \<inter> {(h, j). j = i + 1 \<and> h \<le> j \<and> 0 < h} = {}"
- by fastforce
- have "{(h, j). j = i + 1 \<and> h \<le> j \<and> h = 0} \<union> {(h, j). j = i + 1 \<and> h \<le> j \<and> 0 < h} = {(h, j). j = i + 1 \<and> h \<le> j}"
- by fastforce
- thus "(\<Sum>(h, j)\<in>{(h, j). j = i + 1 \<and> h \<le> j}. coeff f h * coeff f j * (coeff g (i + 1 - h) * coeff g (i + 1 - j))) -
- (\<Sum>(h, j)\<in>{(h, j). j = i + 1 \<and> h \<le> j \<and> 0 < h}.
- coeff f (j + 1) * coeff g (i + 1 - j) * coeff f (h - 1) * coeff g (i + 1 - h)) =
- (\<Sum>(h, j)\<in>{(h, j). j = i + 1 \<and> h \<le> j \<and> h = 0} \<union> {(h, j). j = i + 1 \<and> h \<le> j \<and> 0 < h}.
- coeff f h * coeff f j * (coeff g (i + 1 - h) * coeff g (i + 1 - j))) -
- (\<Sum>(h, j)\<in>{(h, j). j = i + 1 \<and> h \<le> j \<and> 0 < h}.
- coeff f (j + 1) * coeff g (i + 1 - j) * coeff f (h - 1) * coeff g (i + 1 - h))"
- by presburger
- qed
- also have "... =
- (\<Sum>(h, j)\<in>{(h, j). j = i + 1 \<and> h \<le> j \<and> h = 0}. coeff f h * coeff f j * (coeff g (i + 1 - h) * coeff g (i + 1 - j))) +
- (\<Sum>(h, j)\<in>{(h, j). j = i + 1 \<and> h \<le> j \<and> 0 < h}.
- (coeff f h * coeff f j - coeff f (j + 1) * coeff f (h - 1)) * (coeff g (i + 1 - h) * coeff g (i + 1 - j)))"
- by (subst add_diff_eq[symmetric], subst sum_subtractf[symmetric], subst add_left_cancel, rule sum.cong, auto simp: algebra_simps)
- also have "... \<ge> 0"
- proof (rule add_nonneg_nonneg)
- show "0 \<le> (\<Sum>(h, j)\<in>{(h, j). j = i + 1 \<and> h \<le> j \<and> 0 < h}.
- (coeff f h * coeff f j - coeff f (j + 1) * coeff f (h - 1)) *
- (coeff g (i + 1 - h) * coeff g (i + 1 - j)))"
- proof (rule sum_nonneg)
- fix x assume "x \<in> {(h, j). j = i + 1 \<and> h \<le> j \<and> 0 < h}"
- then obtain h j where H: "x = (h, j)" "j = i + 1" "h \<le> j" "0 < h" by fast
- hence "h - 1 \<le> j - 1" by force
- hence 1: "coeff f h * coeff f j - coeff f (j + 1) * coeff f (h - 1) \<ge> 0"
- using normal_poly_general_coeff_mult[OF hf, of "h-1" "j-1"] H
- by (auto simp: algebra_simps)
- hence 2: "0 \<le> coeff g (i + 1 - h) * coeff g (i + 1 - j)"
- by (meson hg mult_nonneg_nonneg normal_coeff_nonneg)
- show "0 \<le> (case x of
- (h, j) \<Rightarrow>
- (coeff f h * coeff f j - coeff f (j + 1) * coeff f (h - 1)) *
- (coeff g (i + 1 - h) * coeff g (i + 1 - j)))"
- apply (subst H(1), subst split, rule mult_nonneg_nonneg)
- subgoal using 1 by blast
- subgoal using 2 by blast
- done
- qed
- qed (rule sum_nonneg, auto simp: hf hg normal_coeff_nonneg)[1]
- finally show ?thesis .
- qed
- ultimately show ?thesis by auto
- qed
- finally show "coeff (f * g) i * coeff (f * g) (i + 2) \<le> (coeff (f * g) (i + 1))\<^sup>2" by (auto simp: power2_eq_square)
- qed
- moreover have "\<And>i j k. i \<le> j \<Longrightarrow> j \<le> k \<Longrightarrow> 0 < coeff (f*g) i \<Longrightarrow> 0 < coeff (f*g) k \<Longrightarrow> 0 < coeff (f*g) j"
- proof -
- fix j k
- assume "0 < coeff (f * g) k"
- hence "k \<le> degree (f * g)" using le_degree by force
- moreover assume "j \<le> k"
- ultimately have "j \<le> degree (f * g)" by auto
- hence 1: "j \<le> degree f + degree g"
- by (simp add: degree_mult_eq hf hg normal_non_zero)
- show "0 < coeff (f * g) j"
- apply (subst coeff_mult, rule sum_pos2[of _ "min j (degree f)"], simp, simp)
- apply (rule mult_pos_pos, rule normal_of_no_zero_root, simp add: hf0, simp)
- using hf apply auto[1]
- apply (rule normal_of_no_zero_root)
- apply (simp add: hg0)
- using 1 apply force
- using hg apply auto[1]
- by (simp add: hf hg normal_coeff_nonneg)
- qed
- ultimately show "normal_poly (f*g)"
- by (rule normal_polyI)
- qed
- qed
- qed
-qed
-
-lemma normal_poly_of_roots:
- fixes p::"real poly"
- assumes "\<And>z. poly (map_poly complex_of_real p) z = 0
- \<Longrightarrow> Re z \<le> 0 \<and> 4*(Re z)^2 \<ge> (cmod z)^2"
- and "lead_coeff p = 1"
- shows "normal_poly p"
- using assms
-proof (induction p rule: real_poly_roots_induct)
- fix p::"real poly" and x::real
- assume "lead_coeff (p * [:- x, 1:]) = 1"
- hence 1: "lead_coeff p = 1"
- by (metis coeff_degree_mult lead_coeff_pCons(1) mult_cancel_left1 pCons_one zero_neq_one)
- assume h: "(\<And>z. poly (map_poly complex_of_real (p * [:- x, 1:])) z = 0 \<Longrightarrow>
- Re z \<le> 0 \<and> (cmod z)\<^sup>2 \<le> 4 * (Re z)\<^sup>2)"
- hence 2: "(\<And>z. poly (map_poly complex_of_real p) z = 0 \<Longrightarrow>
- Re z \<le> 0 \<and> (cmod z)\<^sup>2 \<le> 4 * (Re z)\<^sup>2)"
- by (metis four_x_squared mult_zero_left of_real_poly_map_mult poly_mult)
- have 3: "normal_poly [:-x, 1:]"
- apply (subst linear_normal_iff,
- subst Re_complex_of_real[symmetric], rule conjunct1)
- by (rule h[of x], subst of_real_poly_map_poly[symmetric], force)
- assume "(\<And>z. poly (map_poly complex_of_real p) z = 0
- \<Longrightarrow> Re z \<le> 0 \<and> (cmod z)\<^sup>2 \<le> 4 * (Re z)\<^sup>2) \<Longrightarrow>
- lead_coeff p = 1 \<Longrightarrow> normal_poly p"
- hence "normal_poly p" using 1 2 by fast
- then show "normal_poly (p * [:-x, 1:])"
- using 3 by (rule normal_mult)
-next
- fix p::"real poly" and a b::real
- assume "lead_coeff (p * [:a * a + b * b, - 2 * a, 1:]) = 1"
- hence 1: "lead_coeff p = 1"
- by (smt (verit) coeff_degree_mult lead_coeff_pCons(1) mult_cancel_left1 pCons_eq_0_iff pCons_one)
- assume h: "(\<And>z. poly (map_poly complex_of_real (p * [:a * a + b * b, - 2 * a, 1:])) z = 0 \<Longrightarrow>
- Re z \<le> 0 \<and> (cmod z)\<^sup>2 \<le> 4 * (Re z)\<^sup>2)"
- hence 2: "(\<And>z. poly (map_poly complex_of_real p) z = 0 \<Longrightarrow>
- Re z \<le> 0 \<and> (cmod z)\<^sup>2 \<le> 4 * (Re z)\<^sup>2)"
- proof -
- fix z :: complex
- assume "poly (map_poly complex_of_real p) z = 0"
- then have "\<forall>q. 0 = poly (map_poly complex_of_real (p * q)) z"
- by simp
- then show "Re z \<le> 0 \<and> (cmod z)\<^sup>2 \<le> 4 * (Re z)\<^sup>2"
- using h by presburger
- qed
- have 3: "[:a * a + b * b, - 2 * a, 1:] = [:cmod (a + \<i>*b) ^ 2, -2 * Re (a + \<i>*b), 1:]"
- by (force simp: cmod_def power2_eq_square)
- interpret map_poly_idom_hom complex_of_real ..
- have 4: "normal_poly [:a * a + b * b, - 2 * a, 1:]"
- apply (subst 3, subst quadratic_normal_iff)
- apply (rule h, unfold hom_mult poly_mult)
- by (auto simp: algebra_simps)
- assume "(\<And>z. poly (map_poly complex_of_real p) z = 0 \<Longrightarrow> Re z \<le> 0 \<and> (cmod z)\<^sup>2 \<le> 4 * (Re z)\<^sup>2) \<Longrightarrow>
- lead_coeff p = 1 \<Longrightarrow> normal_poly p"
- hence "normal_poly p" using 1 2 by fast
- then show "normal_poly (p * [:a * a + b * b, - 2 * a, 1:])"
- using 4 by (rule normal_mult)
-next
- fix a::real
- assume "lead_coeff [:a:] = 1"
- moreover have "\<And>i j k.
- lead_coeff [:a:] = 1 \<Longrightarrow>
- i \<le> j \<Longrightarrow>
- j \<le> k \<Longrightarrow> 0 < coeff [:a:] i \<Longrightarrow> 0 < coeff [:a:] k \<Longrightarrow> 0 < coeff [:a:] j"
- by (metis bot_nat_0.extremum_uniqueI coeff_eq_0 degree_pCons_0 leI
- less_numeral_extra(3))
- ultimately show "normal_poly [:a:]"
- apply (subst normal_polyI)
- by (auto simp:pCons_one)
-qed
-
-lemma normal_changes:
- fixes f::"real poly"
- assumes hf: "normal_poly f" and hx: "x > 0"
- defines "df \<equiv> degree f"
- shows "changes (coeffs (f*[:-x,1:])) = 1"
- using df_def hf
-proof (induction df arbitrary: f)
- case 0
- then obtain a where "f = [:a:]" using degree_0_iff by auto
- thus "changes (coeffs (f*[:-x, 1:])) = 1"
- using normal_non_zero[OF \<open>normal_poly f\<close>] hx
- by (auto simp: algebra_simps zero_less_mult_iff mult_less_0_iff)
-next
- case (Suc df)
- then show ?case
- proof (cases "poly f 0 = 0")
- assume "poly f 0 = 0" and hf:"normal_poly f"
- moreover then obtain f' where hdiv: "f = f'*[:0, 1:]"
- by (smt (verit) dvdE mult.commute poly_eq_0_iff_dvd)
- ultimately have hf': "normal_poly f'" using normal_divide_x by blast
- assume "Suc df = degree f"
- hence "degree f' = df" using hdiv normal_non_zero[OF hf'] by (auto simp: degree_mult_eq)
- moreover assume "\<And>f::real poly. df = degree f \<Longrightarrow> normal_poly f \<Longrightarrow> changes (coeffs (f * [:- x, 1:])) = 1"
- ultimately have "changes (coeffs (f' * [:- x, 1:])) = 1" using hf' by fast
- thus "changes (coeffs (f * [:- x, 1:])) = 1"
- apply (subst hdiv, subst mult_pCons_right, subst smult_0_left, subst add_0)
- apply (subst mult_pCons_left, subst smult_0_left, subst add_0)
- by (subst changes_pCons, auto)
- next
- assume hf:"normal_poly f" and "poly f 0 \<noteq> 0"
- hence h': "\<And>i. i \<le> degree f \<Longrightarrow> coeff f i > 0"
- by (auto simp: normal_of_no_zero_root)
- hence "\<And>i. i < degree f - 1 \<Longrightarrow> (coeff f i)/(coeff f (i+1)) \<le> (coeff f (i+1))/(coeff f (i+2))"
- using normal_poly_coeff_mult[OF hf] normal_coeff_nonneg[OF hf]
- by (auto simp: divide_simps power2_eq_square)
- hence h'': "\<And>i. i < degree f - 1 \<Longrightarrow> (coeff f i)/(coeff f (i+1)) - x \<le> (coeff f (i+1))/(coeff f (i+2)) - x"
- by fastforce
- have hdeg: "degree (pCons 0 f - smult x f) = degree f + 1"
- apply (subst diff_conv_add_uminus)
- apply (subst degree_add_eq_left)
- by (auto simp: hf normal_non_zero)
-
- let ?f = "\<lambda> z w. \<lambda>i. if i=0 then z/(x * coeff f 0) else (if i = degree (pCons 0 f - smult x f) then w/(lead_coeff f) else inverse (coeff f i))"
-
- have 1: "\<And>z w. 0 < z \<Longrightarrow> 0 < w \<Longrightarrow> changes (coeffs (f * [:-x, 1:])) =
- changes (-z # map (\<lambda>i. (coeff f (i-1))/(coeff f i) - x) [1..<degree (pCons 0 f - smult x f)] @ [w])"
- proof -
- fix z w :: real
- assume hz: "0 < z" and hw: "0 < w"
-
- have "-z # map (\<lambda>i. (coeff f (i-1))/(coeff f i) - x) [1..<degree (pCons 0 f - smult x f)] @ [w] =
- map (\<lambda>i. if i = 0 then -z else if i = degree (pCons 0 f - smult x f) then w else
- (coeff f (i-1))/(coeff f i) - x) [0..<degree (pCons 0 f - smult x f) + 1]"
- proof (rule nth_equalityI)
- fix i assume "i < length (- z # map (\<lambda>i. coeff f (i - 1) / coeff f i - x) [1..<degree (pCons 0 f - smult x f)] @ [w])"
- hence "i \<le> degree (pCons 0 f - smult x f)"
- using hdeg Suc.hyps(2) by auto
- then consider (a)"i = 0" | (b)"(0 < i \<and> i < degree (pCons 0 f - smult x f))" |
- (c)"i = degree (pCons 0 f - smult x f)"
- by fastforce
- then show "(- z #
- map (\<lambda>i. coeff f (i - 1) / coeff f i - x)
- [1..<degree (pCons 0 f - smult x f)] @
- [w]) ! i =
- map (\<lambda>i. if i = 0 then - z
- else if i = degree (pCons 0 f - smult x f) then w
- else coeff f (i - 1) / coeff f i - x)
- [0..<degree (pCons 0 f - smult x f) + 1] ! i"
- apply (cases)
- by (auto simp: nth_append)
- qed (force simp: hdeg)
-
- also have "... = [?f z w i * (nth_default 0 (coeffs (f * [:-x, 1:])) i).
- i \<leftarrow> [0..<Suc (degree (pCons 0 f - smult x f))]]"
- proof (rule map_cong)
- fix i assume "i \<in> set [0..<Suc (degree (pCons 0 f - smult x f))]"
- then consider (a)"i = 0" | (b)"(0 \<noteq> i \<and> i < degree (pCons 0 f - smult x f))" |
- (c)"i = degree (pCons 0 f - smult x f)"
- by fastforce
- then show "(if i = 0 then - z
- else if i = degree (pCons 0 f - smult x f) then w
- else coeff f (i - 1) / coeff f i - x) =
- (if i = 0 then z / (x * coeff f 0)
- else if i = degree (pCons 0 f - smult x f) then w / lead_coeff f
- else inverse (coeff f i)) *
- nth_default 0 (coeffs (f * [:- x, 1:])) i"
- proof (cases)
- case (a)
- thus ?thesis using hx \<open>poly f 0 \<noteq> 0\<close> by (auto simp: nth_default_coeffs_eq poly_0_coeff_0)
- next
- case (b)
- thus ?thesis using hx h'[of i] hdeg
- by (auto simp: field_simps nth_default_coeffs_eq coeff_pCons nat.split poly_0_coeff_0)
- next
- case (c)
- thus ?thesis using hdeg by (auto simp: nth_default_coeffs_eq coeff_eq_0)
- qed
- qed force
-
- finally have 1: " - z #
- map (\<lambda>i. coeff f (i - 1) / coeff f i - x) [1..<degree (pCons 0 f - smult x f)] @ [w] =
- map (\<lambda>i. (if i = 0 then z / (x * coeff f 0)
- else if i = degree (pCons 0 f - smult x f) then w / lead_coeff f
- else inverse (coeff f i)) *
- nth_default 0 (coeffs (f * [:- x, 1:])) i)
- [0..<Suc (degree (pCons 0 f - smult x f))]" .
-
- have "f * [:-x, 1:] \<noteq> 0" using hdeg by force
-
- show "changes (coeffs (f * [:- x, 1:])) =
- changes
- (- z #
- map (\<lambda>i. coeff f (i - 1) / coeff f i - x)
- [1..<degree (pCons 0 f - smult x f)] @
- [w])"
- apply (subst 1)
- apply (rule changes_scale[symmetric])
- subgoal using hz hw hx h' hdeg by auto
- subgoal using hdeg \<open>f * [:-x, 1:] \<noteq> 0\<close>
- by (auto simp: length_coeffs)
- done
- qed
-
- hence "changes (coeffs (f * [:- x, 1:])) =
- changes
- (- (max 1 (-(coeff f 0 / coeff f 1 - x))) #
- map (\<lambda>i. coeff f (i - 1) / coeff f i - x)
- [1..<degree (pCons 0 f - smult x f)] @
- [max 1 (coeff f (degree f - 1) / lead_coeff f - x)])"
- by force
-
- also have "... = 1"
- proof (rule changes_increasing)
- fix i
- assume "i < length
- (- max 1 (- (coeff f 0 / coeff f 1 - x)) #
- map (\<lambda>i. coeff f (i - 1) / coeff f i - x) [1..<degree (pCons 0 f - smult x f)] @
- [max 1 (coeff f (degree f - 1) / lead_coeff f - x)]) - 1"
- hence "i < degree (pCons 0 f - smult x f)"
- using hdeg Suc.hyps(2) by fastforce
- then consider (a)"i = 0" | (b)"0 \<noteq> i \<and> i < degree (pCons 0 f - smult x f) - 1" |
- (c)"i = degree (pCons 0 f - smult x f) - 1"
- by fastforce
- then show "(- max 1 (- (coeff f 0 / coeff f 1 - x)) #
- map (\<lambda>i. coeff f (i - 1) / coeff f i - x)
- [1..<degree (pCons 0 f - smult x f)] @
- [max 1 (coeff f (degree f - 1) / lead_coeff f - x)]) !
- i
- \<le> (- max 1 (- (coeff f 0 / coeff f 1 - x)) #
- map (\<lambda>i. coeff f (i - 1) / coeff f i - x)
- [1..<degree (pCons 0 f - smult x f)] @
- [max 1 (coeff f (degree f - 1) / lead_coeff f - x)]) !
- (i + 1)"
- proof (cases)
- case a
- then show ?thesis by (auto simp: nth_append)
- next
- case b
- have "coeff f (i - 1) * coeff f (i - 1 + 2) \<le> (coeff f (i - 1 + 1))\<^sup>2"
- by (rule normal_poly_coeff_mult[OF hf, of "i - 1"])
- hence "coeff f (i - 1) / coeff f i \<le> coeff f i / coeff f (i + 1)"
- using h'[of i] h'[of "i+1"] h'[of "i-1"] h' b hdeg
- by (auto simp: power2_eq_square divide_simps)
- then show ?thesis
- using b by (auto simp: nth_append)
- next
- case c
- then show ?thesis using hdeg by (auto simp: nth_append not_le)
- qed
- qed auto
-
- finally show "changes (coeffs (f * [:-x, 1:])) = 1" .
- qed
-qed
-
+section \<open>Normal Polynomials\<close>
+
+theory Normal_Poly
+ imports "RRI_Misc"
+begin
+
+text \<open>
+Here we define normal polynomials as defined in
+ Basu, S., Pollack, R., Roy, M.-F.: Algorithms in Real Algebraic Geometry.
+ Springer Berlin Heidelberg, Berlin, Heidelberg (2016).
+\<close>
+
+definition normal_poly :: "('a::{comm_ring_1,ord}) poly \<Rightarrow> bool" where
+"normal_poly p \<equiv>
+ (p \<noteq> 0) \<and>
+ (\<forall> i. 0 \<le> coeff p i) \<and>
+ (\<forall> i. coeff p i * coeff p (i+2) \<le> (coeff p (i+1))^2) \<and>
+ (\<forall> i j k. i \<le> j \<longrightarrow> j \<le> k \<longrightarrow> 0 < coeff p i
+ \<longrightarrow> 0 < coeff p k \<longrightarrow> 0 < coeff p j)"
+
+lemma normal_non_zero: "normal_poly p \<Longrightarrow> p \<noteq> 0"
+ using normal_poly_def by blast
+
+lemma normal_coeff_nonneg: "normal_poly p \<Longrightarrow> 0 \<le> coeff p i"
+ using normal_poly_def by metis
+
+lemma normal_poly_coeff_mult:
+ "normal_poly p \<Longrightarrow> coeff p i * coeff p (i+2) \<le> (coeff p (i+1))^2"
+ using normal_poly_def by blast
+
+lemma normal_poly_pos_interval:
+ "normal_poly p \<Longrightarrow> i \<le> j \<Longrightarrow> j \<le> k \<Longrightarrow> 0 < coeff p i \<Longrightarrow> 0 < coeff p k
+ \<Longrightarrow> 0 < coeff p j"
+ using normal_poly_def by blast
+
+lemma normal_polyI:
+ assumes "(p \<noteq> 0)"
+ and "(\<And> i. 0 \<le> coeff p i)"
+ and "(\<And> i. coeff p i * coeff p (i+2) \<le> (coeff p (i+1))^2)"
+ and "(\<And> i j k. i \<le> j \<Longrightarrow> j \<le> k \<Longrightarrow> 0 < coeff p i \<Longrightarrow> 0 < coeff p k \<Longrightarrow> 0 < coeff p j)"
+ shows "normal_poly p"
+ using assms by (force simp: normal_poly_def)
+
+lemma linear_normal_iff:
+ fixes x::real
+ shows "normal_poly [:-x, 1:] \<longleftrightarrow> x \<le> 0"
+proof
+ assume "normal_poly [:-x, 1:]"
+ thus "x \<le> 0" using normal_coeff_nonneg[of "[:-x, 1:]" 0] by auto
+next
+ assume "x \<le> 0"
+ then have "0 \<le> coeff [:- x, 1:] i" for i
+ by (cases i) (simp_all add: pCons_one)
+ moreover have "0 < coeff [:- x, 1:] j"
+ if "i \<le> j" "j \<le> k" "0 < coeff [:- x, 1:] i"
+ "0 < coeff [:- x, 1:] k" for i j k
+ apply (cases "k=0 \<or> i=0")
+ subgoal using that
+ by (smt (z3) bot_nat_0.extremum_uniqueI degree_pCons_eq_if
+ le_antisym le_degree not_less_eq_eq)
+ subgoal using that
+ by (smt (z3) One_nat_def degree_pCons_eq_if le_degree less_one
+ not_le one_neq_zero pCons_one verit_la_disequality)
+ done
+ ultimately show "normal_poly [:-x, 1:]"
+ unfolding normal_poly_def by auto
+qed
+
+lemma quadratic_normal_iff:
+ fixes z::complex
+ shows "normal_poly [:(cmod z)\<^sup>2, -2*Re z, 1:]
+ \<longleftrightarrow> Re z \<le> 0 \<and> 4*(Re z)^2 \<ge> (cmod z)^2"
+proof
+ assume "normal_poly [:(cmod z)\<^sup>2, - 2 * Re z, 1:]"
+ hence "-2*Re z \<ge> 0 \<and> (cmod z)^2 \<ge> 0 \<and> (-2*Re z)^2 \<ge> (cmod z)^2"
+ using normal_coeff_nonneg[of _ 1] normal_poly_coeff_mult[of _ 0]
+ by fastforce
+ thus "Re z \<le> 0 \<and> 4*(Re z)^2 \<ge> (cmod z)^2"
+ by auto
+next
+ assume asm:"Re z \<le> 0 \<and> 4*(Re z)^2 \<ge> (cmod z)^2"
+ define P where "P=[:(cmod z)\<^sup>2, - 2 * Re z, 1:]"
+
+ have "0 \<le> coeff P i" for i
+ unfolding P_def using asm
+ apply (cases "i=0\<or>i=1\<or>i=2")
+ by (auto simp:numeral_2_eq_2 coeff_eq_0)
+ moreover have "coeff P i * coeff P (i + 2) \<le> (coeff P (i + 1))\<^sup>2" for i
+ apply (cases "i=0\<or>i=1\<or>i=2")
+ using asm
+ unfolding P_def by (auto simp:coeff_eq_0)
+ moreover have "0 < coeff P j"
+ if "0 < coeff P k" "0 < coeff P i" "j \<le> k" "i \<le> j"
+ for i j k
+ using that unfolding P_def
+ apply (cases "k=0 \<or> k=1 \<or> k=2")
+ subgoal using asm
+ by (smt (z3) One_nat_def Suc_1 bot_nat_0.extremum_uniqueI
+ coeff_pCons_0 coeff_pCons_Suc le_Suc_eq
+ zero_less_power2)
+ subgoal by (auto simp:coeff_eq_0)
+ done
+ moreover have "P\<noteq>0" unfolding P_def by auto
+ ultimately show "normal_poly P"
+ unfolding normal_poly_def by blast
+qed
+
+lemma normal_of_no_zero_root:
+ fixes f::"real poly"
+ assumes hzero: "poly f 0 \<noteq> 0" and hdeg: "i \<le> degree f"
+ and hnorm: "normal_poly f"
+ shows "0 < coeff f i"
+proof -
+ have "coeff f 0 > 0" using hzero normal_coeff_nonneg[OF hnorm]
+ by (metis eq_iff not_le_imp_less poly_0_coeff_0)
+ moreover have "coeff f (degree f) > 0" using normal_coeff_nonneg[OF hnorm] normal_non_zero[OF hnorm]
+ by (meson dual_order.irrefl eq_iff eq_zero_or_degree_less not_le_imp_less)
+ moreover have "0 \<le> i" by simp
+ ultimately show "0 < coeff f i" using hdeg normal_poly_pos_interval[OF hnorm] by blast
+qed
+
+lemma normal_divide_x:
+ fixes f::"real poly"
+ assumes hnorm: "normal_poly (f*[:0,1:])"
+ shows "normal_poly f"
+proof (rule normal_polyI)
+ show "f \<noteq> 0"
+ using normal_non_zero[OF hnorm] by auto
+next
+ fix i
+ show "0 \<le> coeff f i"
+ using normal_coeff_nonneg[OF hnorm, of "Suc i"] by (simp add: coeff_pCons)
+next
+ fix i
+ show "coeff f i * coeff f (i + 2) \<le> (coeff f (i + 1))\<^sup>2"
+ using normal_poly_coeff_mult[OF hnorm, of "Suc i"] by (simp add: coeff_pCons)
+next
+ fix i j k
+ show "i \<le> j \<Longrightarrow> j \<le> k \<Longrightarrow> 0 < coeff f i \<Longrightarrow> 0 < coeff f k \<Longrightarrow> 0 < coeff f j"
+ using normal_poly_pos_interval[of _ "Suc i" "Suc j" "Suc k", OF hnorm]
+ by (simp add: coeff_pCons)
+qed
+
+lemma normal_mult_x:
+ fixes f::"real poly"
+ assumes hnorm: "normal_poly f"
+ shows "normal_poly (f * [:0, 1:])"
+proof (rule normal_polyI)
+ show "f * [:0, 1:] \<noteq> 0"
+ using normal_non_zero[OF hnorm] by auto
+next
+ fix i
+ show "0 \<le> coeff (f * [:0, 1:]) i"
+ using normal_coeff_nonneg[OF hnorm, of "i-1"] by (cases i, auto simp: coeff_pCons)
+next
+ fix i
+ show "coeff (f * [:0, 1:]) i * coeff (f * [:0, 1:]) (i + 2) \<le> (coeff (f * [:0, 1:]) (i + 1))\<^sup>2"
+ using normal_poly_coeff_mult[OF hnorm, of "i-1"] by (cases i, auto simp: coeff_pCons)
+next
+ fix i j k
+ show "i \<le> j \<Longrightarrow> j \<le> k \<Longrightarrow> 0 < coeff (f * [:0, 1:]) i \<Longrightarrow> 0 < coeff (f * [:0, 1:]) k \<Longrightarrow> 0 < coeff (f * [:0, 1:]) j"
+ using normal_poly_pos_interval[of _ "i-1" "j-1" "k-1", OF hnorm]
+ apply (cases i, force)
+ apply (cases j, force)
+ apply (cases k, force)
+ by (auto simp: coeff_pCons)
+qed
+
+lemma normal_poly_general_coeff_mult:
+ fixes f::"real poly"
+ assumes "normal_poly f" and "h \<le> j"
+ shows "coeff f (h+1) * coeff f (j+1) \<ge> coeff f h * coeff f (j+2)"
+using assms proof (induction j)
+ case 0
+ then show ?case
+ using normal_poly_coeff_mult by (auto simp: power2_eq_square)[1]
+next
+ case (Suc j)
+ then show ?case
+ proof (cases "h = Suc j")
+ assume "h = Suc j" "normal_poly f"
+ thus ?thesis
+ using normal_poly_coeff_mult by (auto simp: power2_eq_square)
+ next
+ assume "(normal_poly f \<Longrightarrow>
+ h \<le> j \<Longrightarrow> coeff f h * coeff f (j + 2) \<le> coeff f (h + 1) * coeff f (j + 1))"
+ "normal_poly f" and h: "h \<le> Suc j" "h \<noteq> Suc j"
+ hence IH: "coeff f h * coeff f (j + 2) \<le> coeff f (h + 1) * coeff f (j + 1)"
+ by linarith
+ show ?thesis
+ proof (cases "coeff f (Suc j + 1) = 0", cases "coeff f (Suc j + 2) = 0")
+ show "coeff f (Suc j + 1) = 0 \<Longrightarrow> coeff f (Suc j + 2) = 0 \<Longrightarrow>
+ coeff f h * coeff f (Suc j + 2) \<le> coeff f (h + 1) * coeff f (Suc j + 1)"
+ by (metis assms(1) mult_zero_right normal_coeff_nonneg)
+ next
+ assume 1: "coeff f (Suc j + 1) = 0" "coeff f (Suc j + 2) \<noteq> 0"
+ hence "coeff f (Suc j + 2) > 0" "\<not>coeff f (Suc j + 1) > 0"
+ using normal_coeff_nonneg[of f "Suc j + 2"] assms(1) by auto
+ hence "coeff f h > 0 \<Longrightarrow> False"
+ using normal_poly_pos_interval[of f h "Suc j + 1" "Suc j + 2"] assms(1) h by force
+ hence "coeff f h = 0"
+ using normal_coeff_nonneg[OF assms(1)] less_eq_real_def by auto
+ thus "coeff f h * coeff f (Suc j + 2) \<le> coeff f (h + 1) * coeff f (Suc j + 1)"
+ using 1 by fastforce
+ next
+ assume 1: "coeff f (Suc j + 1) \<noteq> 0"
+ show "coeff f h * coeff f (Suc j + 2) \<le> coeff f (h + 1) * coeff f (Suc j + 1)"
+ proof (cases "coeff f (Suc j) = 0")
+ assume 2: "coeff f (Suc j) = 0"
+ hence "coeff f (Suc j + 1) > 0" "\<not>coeff f (Suc j) > 0"
+ using normal_coeff_nonneg[of f "Suc j + 1"] assms(1) 1 by auto
+ hence "coeff f h > 0 \<Longrightarrow> False"
+ using normal_poly_pos_interval[of f h "Suc j" "Suc j + 1"] assms(1) h by force
+ hence "coeff f h = 0"
+ using normal_coeff_nonneg[OF assms(1)] less_eq_real_def by auto
+ thus "coeff f h * coeff f (Suc j + 2) \<le> coeff f (h + 1) * coeff f (Suc j + 1)"
+ by (simp add: assms(1) normal_coeff_nonneg)
+ next
+ assume 2: "coeff f (Suc j) \<noteq> 0"
+ from normal_poly_coeff_mult[OF assms(1), of "Suc j"] normal_coeff_nonneg[OF assms(1), of "Suc j"]
+ normal_coeff_nonneg[OF assms(1), of "Suc (Suc j)"] 1 2
+ have 3: "coeff f (Suc j + 1) / coeff f (Suc j) \<ge> coeff f (Suc j + 2) / coeff f (Suc j + 1)"
+ by (auto simp: power2_eq_square divide_simps algebra_simps)
+ have "(coeff f h * coeff f (j + 2)) * (coeff f (Suc j + 2) / coeff f (Suc j + 1)) \<le> (coeff f (h + 1) * coeff f (j + 1)) * (coeff f (Suc j + 1) / coeff f (Suc j))"
+ apply (rule mult_mono[OF IH])
+ using 3 by (simp_all add: assms(1) normal_coeff_nonneg)
+ thus "coeff f h * coeff f (Suc j + 2) \<le> coeff f (h + 1) * coeff f (Suc j + 1)"
+ using 1 2 by fastforce
+ qed
+ qed
+ qed
+qed
+
+lemma normal_mult:
+ fixes f g::"real poly"
+ assumes hf: "normal_poly f" and hg: "normal_poly g"
+ defines "df \<equiv> degree f" and "dg \<equiv> degree g"
+ shows "normal_poly (f*g)"
+using df_def hf proof (induction df arbitrary: f)
+text \<open>We shall first show that without loss of generality we may assume \<open>poly f 0 \<noteq> 0\<close>,
+ this is done by induction on the degree, if 0 is a root then we derive the result from \<open>f/[:0,1:]\<close>.\<close>
+ fix f::"real poly" fix i::nat
+ assume "0 = degree f" and hf: "normal_poly f"
+ then obtain a where "f = [:a:]" using degree_0_iff by auto
+ then show "normal_poly (f*g)"
+ apply (subst normal_polyI)
+ subgoal using normal_non_zero[OF hf] normal_non_zero[OF hg] by auto
+ subgoal
+ using normal_coeff_nonneg[of _ 0, OF hf] normal_coeff_nonneg[OF hg]
+ by simp
+ subgoal
+ using normal_coeff_nonneg[of _ 0, OF hf] normal_poly_coeff_mult[OF hg]
+ by (auto simp: algebra_simps power2_eq_square mult_left_mono)[1]
+ subgoal
+ using normal_non_zero[OF hf] normal_coeff_nonneg[of _ 0, OF hf] normal_poly_pos_interval[OF hg]
+ by (simp add: zero_less_mult_iff)
+ subgoal by simp
+ done
+next
+ case (Suc df)
+ then show ?case
+ proof (cases "poly f 0 = 0")
+ assume "poly f 0 = 0" and hf:"normal_poly f"
+ moreover then obtain f' where hdiv: "f = f'*[:0,1:]"
+ by (smt (verit) dvdE mult.commute poly_eq_0_iff_dvd)
+ ultimately have hf': "normal_poly f'" using normal_divide_x by blast
+ assume "Suc df = degree f"
+ hence "degree f' = df" using hdiv normal_non_zero[OF hf'] by (auto simp: degree_mult_eq)
+ moreover assume "\<And>f. df = degree f \<Longrightarrow> normal_poly f \<Longrightarrow> normal_poly (f * g)"
+ ultimately have "normal_poly (f'*g)" using hf' by blast
+ thus "normal_poly (f*g)" using hdiv normal_mult_x by fastforce
+ next
+ assume hf: "normal_poly f" and hf0: "poly f 0 \<noteq> 0"
+ define dg where "dg \<equiv> degree g"
+ show "normal_poly (f * g)"
+ using dg_def hg proof (induction dg arbitrary: g)
+ text \<open>Similarly we may assume \<open>poly g 0 \<noteq> 0\<close>.\<close>
+ fix g::"real poly" fix i::nat
+ assume "0 = degree g" and hg: "normal_poly g"
+ then obtain a where "g = [:a:]" using degree_0_iff by auto
+ then show "normal_poly (f*g)"
+ apply (subst normal_polyI)
+ subgoal
+ using normal_non_zero[OF hg] normal_non_zero[OF hf] by auto
+ subgoal
+ using normal_coeff_nonneg[of _ 0, OF hg] normal_coeff_nonneg[OF hf]
+ by simp
+ subgoal
+ using normal_coeff_nonneg[of _ 0, OF hg] normal_poly_coeff_mult[OF hf]
+ by (auto simp: algebra_simps power2_eq_square mult_left_mono)
+ subgoal
+ using normal_non_zero[OF hf] normal_coeff_nonneg[of _ 0, OF hg]
+ normal_poly_pos_interval[OF hf]
+ by (simp add: zero_less_mult_iff)
+ by simp
+ next
+ case (Suc dg)
+ then show ?case
+ proof (cases "poly g 0 = 0")
+ assume "poly g 0 = 0" and hg:"normal_poly g"
+ moreover then obtain g' where hdiv: "g = g'*[:0,1:]"
+ by (smt (verit) dvdE mult.commute poly_eq_0_iff_dvd)
+ ultimately have hg': "normal_poly g'" using normal_divide_x by blast
+ assume "Suc dg = degree g"
+ hence "degree g' = dg" using hdiv normal_non_zero[OF hg'] by (auto simp: degree_mult_eq)
+ moreover assume "\<And>g. dg = degree g \<Longrightarrow> normal_poly g \<Longrightarrow> normal_poly (f * g)"
+ ultimately have "normal_poly (f*g')" using hg' by blast
+ thus "normal_poly (f*g)" using hdiv normal_mult_x by fastforce
+ next
+ text \<open>It now remains to show that $(fg)_i \geq 0$. This follows by decomposing $\{(h, j) \in
+ \mathbf{Z}^2 | h > j\} = \{(h, j) \in \mathbf{Z}^2 | h \leq j\} \cup \{(h, h - 1) \in
+ \mathbf{Z}^2 | h \in \mathbf{Z}\}$.
+ Note in order to avoid working with infinite sums over integers all these sets are
+ bounded, which adds some complexity compared to the proof of lemma 2.55 in
+ Basu, S., Pollack, R., Roy, M.-F.: Algorithms in Real Algebraic Geometry.
+ Springer Berlin Heidelberg, Berlin, Heidelberg (2016).\<close>
+ assume hg0: "poly g 0 \<noteq> 0" and hg: "normal_poly g"
+ have "f * g \<noteq> 0" using hf hg by (simp add: normal_non_zero Suc.prems)
+ moreover have "\<And>i. coeff (f*g) i \<ge> 0"
+ apply (subst coeff_mult, rule sum_nonneg, rule mult_nonneg_nonneg)
+ using normal_coeff_nonneg[OF hf] normal_coeff_nonneg[OF hg] by auto
+ moreover have "
+ coeff (f*g) i * coeff (f*g) (i+2) \<le> (coeff (f*g) (i+1))^2" for i
+ proof -
+
+ text \<open>$(fg)_{i+1}^2 - (fg)_i(fg)_{i+2} = \left(\sum_x f_xg_{i+1-x}\right)^2 -
+ \left(\sum_x f_xg_{i+2-x}\right)\left(\sum_x f_xg_{i-x}\right)$\<close>
+ have "(coeff (f*g) (i+1))^2 - coeff (f*g) i * coeff (f*g) (i+2) =
+ (\<Sum>x\<le>i+1. coeff f x * coeff g (i + 1 - x)) *
+ (\<Sum>x\<le>i+1. coeff f x * coeff g (i + 1 - x)) -
+ (\<Sum>x\<le>i+2. coeff f x * coeff g (i + 2 - x)) *
+ (\<Sum>x\<le>i. coeff f x * coeff g (i - x))"
+ by (auto simp: coeff_mult power2_eq_square algebra_simps)
+
+ text \<open>$\dots = \sum_{x, y} f_xg_{i+1-x}f_yg_{i+1-y} - \sum_{x, y} f_xg_{i+2-x}f_yg_{i-y}$\<close>
+ also have "... =
+ (\<Sum>x\<le>i+1. \<Sum>y\<le>i+1. coeff f x * coeff g (i + 1 - x) *
+ coeff f y * coeff g (i + 1 - y)) -
+ (\<Sum>x\<le>i+2. \<Sum>y\<le>i. coeff f x * coeff g (i + 2 - x) *
+ coeff f y * coeff g (i - y))"
+ by (subst sum_product, subst sum_product, auto simp: algebra_simps)
+
+ text \<open>$\dots = \sum_{h \leq j} f_hg_{i+1-h}f_jg_{i+1-j} + \sum_{h>j} f_hg_{i+1-h}f_jg_{i+1-j}
+ - \sum_{h \leq j} f_hg_{i+2-h}f_jg_{i-j} - \sum_{h>j} f_hg_{i+2-h}f_jg_{i-j}$\<close>
+ also have "... =
+ (\<Sum>(h, j)\<in>{(h, j). i+1 \<ge> j \<and> j \<ge> h}.
+ coeff f h * coeff g (i + 1 - h) * coeff f j * coeff g (i + 1 - j)) +
+ (\<Sum>(h, j)\<in>{(h, j). i+1 \<ge> h \<and> h > j}.
+ coeff f h * coeff g (i + 1 - h) * coeff f j * coeff g (i + 1 - j)) -
+ ((\<Sum>(h, j)\<in>{(h, j). i \<ge> j \<and> j \<ge> h}.
+ coeff f h * coeff g (i + 2 - h) * coeff f j * coeff g (i - j)) +
+ (\<Sum>(h, j)\<in>{(h, j). i + 2 \<ge> h \<and> h > j \<and> i \<ge> j}.
+ coeff f h * coeff g (i + 2 - h) * coeff f j * coeff g (i - j)))"
+ proof -
+ have "(\<Sum>x\<le>i + 1. \<Sum>y\<le>i + 1. coeff f x * coeff g (i + 1 - x) * coeff f y * coeff g (i + 1 - y)) =
+ (\<Sum>(h, j)\<in>{(h, j). j \<le> i + 1 \<and> h \<le> j}.
+ coeff f h * coeff g (i + 1 - h) * coeff f j * coeff g (i + 1 - j)) +
+ (\<Sum>(h, j)\<in>{(h, j). h \<le> i + 1 \<and> j < h}.
+ coeff f h * coeff g (i + 1 - h) * coeff f j * coeff g (i + 1 - j))"
+ proof (subst sum.union_disjoint[symmetric])
+ have H:"{(h, j). j \<le> i + 1 \<and> h \<le> j} \<subseteq> {..i+1} \<times> {..i+1}"
+ "{(h, j). h \<le> i + 1 \<and> j < h} \<subseteq> {..i+1} \<times> {..i+1}"
+ "finite ({..i+1} \<times> {..i+1})"
+ by (fastforce, fastforce, fastforce)
+ show "finite {(h, j). j \<le> i + 1 \<and> h \<le> j}"
+ apply (rule finite_subset) using H by (blast, blast)
+ show "finite {(h, j). h \<le> i + 1 \<and> j < h}"
+ apply (rule finite_subset) using H by (blast, blast)
+ show "{(h, j). j \<le> i + 1 \<and> h \<le> j} \<inter> {(h, j). h \<le> i + 1 \<and> j < h} = {}"
+ by fastforce
+ show "(\<Sum>x\<le>i + 1. \<Sum>y\<le>i + 1. coeff f x * coeff g (i + 1 - x) * coeff f y * coeff g (i + 1 - y)) =
+ (\<Sum>(h, j)\<in>{(h, j). j \<le> i + 1 \<and> h \<le> j} \<union> {(h, j). h \<le> i + 1 \<and> j < h}.
+ coeff f h * coeff g (i + 1 - h) * coeff f j * coeff g (i + 1 - j))"
+ apply (subst sum.cartesian_product, rule sum.cong)
+ apply force by blast
+ qed
+ moreover have "(\<Sum>x\<le>i + 2. \<Sum>y\<le>i. coeff f x * coeff g (i + 2 - x) * coeff f y * coeff g (i - y)) =
+ (\<Sum>(h, j)\<in>{(h, j). j \<le> i \<and> h \<le> j}.
+ coeff f h * coeff g (i + 2 - h) * coeff f j * coeff g (i - j)) +
+ (\<Sum>(h, j)\<in>{(h, j). i + 2 \<ge> h \<and> h > j \<and> i \<ge> j}.
+ coeff f h * coeff g (i + 2 - h) * coeff f j * coeff g (i - j))"
+ proof (subst sum.union_disjoint[symmetric])
+ have H:"{(h, j). j \<le> i \<and> h \<le> j} \<subseteq> {..i+2} \<times> {..i}"
+ "{(h, j). i + 2 \<ge> h \<and> h > j \<and> i \<ge> j} \<subseteq> {..i+2} \<times> {..i}"
+ "finite ({..i+2} \<times> {..i})"
+ by (fastforce, fastforce, fastforce)
+ show "finite {(h, j). j \<le> i \<and> h \<le> j}"
+ apply (rule finite_subset) using H by (blast, blast)
+ show "finite {(h, j). i + 2 \<ge> h \<and> h > j \<and> i \<ge> j}"
+ apply (rule finite_subset) using H by (blast, blast)
+ show "{(h, j). j \<le> i \<and> h \<le> j} \<inter> {(h, j). i + 2 \<ge> h \<and> h > j \<and> i \<ge> j} = {}"
+ by fastforce
+ show "(\<Sum>x\<le>i + 2. \<Sum>y\<le>i. coeff f x * coeff g (i + 2 - x) * coeff f y * coeff g (i - y)) =
+ (\<Sum>(h, j)\<in>{(h, j). j \<le> i \<and> h \<le> j} \<union> {(h, j). i + 2 \<ge> h \<and> h > j \<and> i \<ge> j}.
+ coeff f h * coeff g (i + 2 - h) * coeff f j * coeff g (i - j))"
+ apply (subst sum.cartesian_product, rule sum.cong)
+ apply force by blast
+ qed
+ ultimately show ?thesis by presburger
+ qed
+
+ text \<open>$\dots = \sum_{h \leq j} f_hg_{i+1-h}f_jg_{i+1-j} + \sum_{h \leq j} f_{j+1}g_{i-j}f_{h-2}g_{i+2-h}
+ + \sum_h f_hg_{i+1-h}f_{h-1}g_{i+2-h} - \sum_{h \leq j} f_hg_{i+2-h}f_jg_{i-j}
+ - \sum_{h \leq j} f_{j+1}g_{i+1-j}f_{h-2}g_{i+1-h} - \sum_h f_hg_{i+2-h}f_{h-1}g_{i+1-h}$\<close>
+ also have "... =
+ (\<Sum>(h, j)\<in>{(h, j). j \<le> i + 1 \<and> h \<le> j}.
+ coeff f h * coeff g (i + 1 - h) * coeff f j * coeff g (i + 1 - j)) +
+ (\<Sum>(h, j)\<in>{(h, j). j \<le> i \<and> h \<le> j \<and> 0 < h}.
+ coeff f (j+1) * coeff g (i - j) * coeff f (h-1) * coeff g (i + 2 - h)) +
+ (\<Sum>h\<in>{1..i+1}.
+ coeff f h * coeff g (i + 1 - h) * coeff f (h-1) * coeff g (i + 2 - h)) -
+ ((\<Sum>(h, j)\<in>{(h, j). j \<le> i \<and> h \<le> j}.
+ coeff f h * coeff g (i + 2 - h) * coeff f j * coeff g (i - j)) +
+ (\<Sum>(h, j)\<in>{(h, j). j \<le> i + 1 \<and> h \<le> j \<and> 0 < h}.
+ coeff f (j+1) * coeff g (i + 1 - j) * coeff f (h-1) * coeff g (i + 1 - h)) +
+ (\<Sum>h\<in>{1..i+1}.
+ coeff f h * coeff g (i + 2 - h) * coeff f (h-1) * coeff g (i + 1 - h)))"
+ proof -
+ have "(\<Sum>(h, j)\<in>{(h, j). h \<le> i + 1 \<and> j < h}.
+ coeff f h * coeff g (i + 1 - h) * coeff f j * coeff g (i + 1 - j)) =
+ (\<Sum>(h, j)\<in>{(h, j). j \<le> i \<and> h \<le> j \<and> 0 < h}.
+ coeff f (j + 1) * coeff g (i - j) * coeff f (h - 1) * coeff g (i + 2 - h)) +
+ (\<Sum>h = 1..i + 1. coeff f h * coeff g (i + 1 - h) * coeff f (h - 1) * coeff g (i + 2 - h))"
+ proof -
+ have 1: "(\<Sum>(h, j)\<in>{(h, j). j \<le> i \<and> h \<le> j \<and> 0 < h}.
+ coeff f (j + 1) * coeff g (i - j) * coeff f (h - 1) * coeff g (i + 2 - h)) =
+ (\<Sum>(h, j)\<in>{(h, j). h \<le> i + 1 \<and> j < h \<and> h \<noteq> j + 1}.
+ coeff f h * coeff g (i + 1 - h) * coeff f j * coeff g (i + 1 - j))"
+ proof (rule sum.reindex_cong)
+ show "{(h, j). j \<le> i \<and> h \<le> j \<and> 0 < h} = (\<lambda>(h, j). (j+1, h-1)) ` {(h, j). h \<le> i + 1 \<and> j < h \<and> h \<noteq> j + 1}"
+ proof
+ show "(\<lambda>(h, j). (j + 1, h - 1)) ` {(h, j). h \<le> i + 1 \<and> j < h \<and> h \<noteq> j + 1} \<subseteq> {(h, j). j \<le> i \<and> h \<le> j \<and> 0 < h}"
+ by fastforce
+ show "{(h, j). j \<le> i \<and> h \<le> j \<and> 0 < h} \<subseteq> (\<lambda>(h, j). (j + 1, h - 1)) ` {(h, j). h \<le> i + 1 \<and> j < h \<and> h \<noteq> j + 1}"
+ proof
+ fix x
+ assume "x \<in> {(h, j). j \<le> i \<and> h \<le> j \<and> 0 < h}"
+ then obtain h j where "x = (h, j)" "j \<le> i" "h \<le> j" "0 < h" by blast
+ hence "j + 1 \<le> i + 1 \<and> h - 1 < j + 1 \<and> j + 1 \<noteq> h - 1 + 1 \<and> x = ((h - 1) + 1, (j + 1) - 1)"
+ by auto
+ thus "x \<in> (\<lambda>(h, j). (j + 1, h - 1)) ` {(h, j). h \<le> i + 1 \<and> j < h \<and> h \<noteq> j + 1}"
+ by (auto simp: image_iff)
+ qed
+ qed
+ show "inj_on (\<lambda>(h, j). (j + 1, h - 1)) {(h, j). h \<le> i + 1 \<and> j < h \<and> h \<noteq> j + 1}"
+ proof
+ fix x y::"nat\<times>nat"
+ assume "x \<in> {(h, j). h \<le> i + 1 \<and> j < h \<and> h \<noteq> j + 1}" "y \<in> {(h, j). h \<le> i + 1 \<and> j < h \<and> h \<noteq> j + 1}"
+ thus "(case x of (h, j) \<Rightarrow> (j + 1, h - 1)) = (case y of (h, j) \<Rightarrow> (j + 1, h - 1)) \<Longrightarrow> x = y"
+ by auto
+ qed
+ show "\<And>x. x \<in> {(h, j). h \<le> i + 1 \<and> j < h \<and> h \<noteq> j + 1} \<Longrightarrow>
+ (case case x of (h, j) \<Rightarrow> (j + 1, h - 1) of
+ (h, j) \<Rightarrow> coeff f (j + 1) * coeff g (i - j) * coeff f (h - 1) * coeff g (i + 2 - h)) =
+ (case x of (h, j) \<Rightarrow> coeff f h * coeff g (i + 1 - h) * coeff f j * coeff g (i + 1 - j))"
+ by fastforce
+ qed
+ have 2: "(\<Sum>h = 1..i + 1. coeff f h * coeff g (i + 1 - h) * coeff f (h - 1) * coeff g (i + 2 - h)) =
+ (\<Sum>(h, j)\<in>{(h, j). h \<le> i + 1 \<and> j < h \<and> h = j + 1}.
+ coeff f h * coeff g (i + 1 - h) * coeff f j * coeff g (i + 1 - j))"
+ proof (rule sum.reindex_cong)
+ show "{1..i + 1} = fst ` {(h, j). h \<le> i + 1 \<and> j < h \<and> h = j + 1}"
+ proof
+ show "{1..i + 1} \<subseteq> fst ` {(h, j). h \<le> i + 1 \<and> j < h \<and> h = j + 1}"
+ proof
+ fix x
+ assume "x \<in> {1..i + 1}"
+ hence "x \<le> i + 1 \<and> x - 1 < x \<and> x = x - 1 + 1 \<and> x = fst (x, x-1)"
+ by auto
+ thus "x \<in> fst ` {(h, j). h \<le> i + 1 \<and> j < h \<and> h = j + 1}"
+ by blast
+ qed
+ show "fst ` {(h, j). h \<le> i + 1 \<and> j < h \<and> h = j + 1} \<subseteq> {1..i + 1}"
+ by force
+ qed
+ show "inj_on fst {(h, j). h \<le> i + 1 \<and> j < h \<and> h = j + 1}"
+ proof
+ fix x y
+ assume "x \<in> {(h, j). h \<le> i + 1 \<and> j < h \<and> h = j + 1}"
+ "y \<in> {(h, j). h \<le> i + 1 \<and> j < h \<and> h = j + 1}"
+ hence "x = (fst x, fst x - 1)" "y = (fst y, fst y - 1)" "fst x > 0" "fst y > 0"
+ by auto
+ thus "fst x = fst y \<Longrightarrow> x = y" by presburger
+ qed
+ show "\<And>x. x \<in> {(h, j). h \<le> i + 1 \<and> j < h \<and> h = j + 1} \<Longrightarrow>
+ coeff f (fst x) * coeff g (i + 1 - fst x) * coeff f (fst x - 1) * coeff g (i + 2 - fst x) =
+ (case x of (h, j) \<Rightarrow> coeff f h * coeff g (i + 1 - h) * coeff f j * coeff g (i + 1 - j))"
+ by fastforce
+ qed
+ have H: "{(h, j). h \<le> i + 1 \<and> j < h \<and> h \<noteq> j + 1} \<subseteq> {0..i+1}\<times>{0..i+1}"
+ "{(h, j). h \<le> i + 1 \<and> j < h \<and> h = j + 1} \<subseteq> {0..i+1}\<times>{0..i+1}"
+ "finite ({0..i+1}\<times>{0..i+1})"
+ by (fastforce, fastforce, fastforce)
+ have "finite {(h, j). h \<le> i + 1 \<and> j < h \<and> h \<noteq> j + 1}"
+ "finite {(h, j). h \<le> i + 1 \<and> j < h \<and> h = j + 1}"
+ apply (rule finite_subset) using H apply (simp, simp)
+ apply (rule finite_subset) using H apply (simp, simp)
+ done
+ thus ?thesis
+ apply (subst 1, subst 2, subst sum.union_disjoint[symmetric])
+ apply auto[3]
+ apply (rule sum.cong)
+ by auto
+ qed
+ moreover have "(\<Sum>(h, j)\<in>{(h, j). h \<le> i + 2 \<and> j < h \<and> j \<le> i}.
+ coeff f h * coeff g (i + 2 - h) * coeff f j * coeff g (i - j)) =
+ (\<Sum>(h, j)\<in>{(h, j). j \<le> i + 1 \<and> h \<le> j \<and> 0 < h}.
+ coeff f (j + 1) * coeff g (i + 1 - j) * coeff f (h - 1) * coeff g (i + 1 - h)) +
+ (\<Sum>h = 1..i + 1. coeff f h * coeff g (i + 2 - h) * coeff f (h - 1) * coeff g (i + 1 - h))"
+ proof -
+ have 1: "(\<Sum>(h, j)\<in>{(h, j). j \<le> i + 1 \<and> h \<le> j \<and> 0 < h}.
+ coeff f (j + 1) * coeff g (i + 1 - j) * coeff f (h - 1) * coeff g (i + 1 - h)) =
+ (\<Sum>(h, j)\<in>{(h, j). h \<le> i + 2 \<and> j < h \<and> j \<le> i \<and> h \<noteq> j + 1}.
+ coeff f h * coeff g (i + 2 - h) * coeff f j * coeff g (i - j))"
+ proof (rule sum.reindex_cong)
+ show "{(h, j). j \<le> i + 1 \<and> h \<le> j \<and> 0 < h} = (\<lambda>(h, j). (j+1, h-1)) ` {(h, j). h \<le> i + 2 \<and> j < h \<and> j \<le> i \<and> h \<noteq> j + 1}"
+ proof
+ show "(\<lambda>(h, j). (j + 1, h - 1)) ` {(h, j). h \<le> i + 2 \<and> j < h \<and> j \<le> i \<and> h \<noteq> j + 1} \<subseteq> {(h, j). j \<le> i + 1 \<and> h \<le> j \<and> 0 < h}"
+ by fastforce
+ show "{(h, j). j \<le> i + 1 \<and> h \<le> j \<and> 0 < h} \<subseteq> (\<lambda>(h, j). (j + 1, h - 1)) ` {(h, j). h \<le> i + 2 \<and> j < h \<and> j \<le> i \<and> h \<noteq> j + 1}"
+ proof
+ fix x
+ assume "x \<in> {(h, j). j \<le> i + 1 \<and> h \<le> j \<and> 0 < h}"
+ then obtain h j where "x = (h, j)" "j \<le> i + 1" "h \<le> j" "0 < h" by blast
+ hence "j + 1 \<le> i + 2 \<and> h - 1 < j + 1 \<and> h - 1 \<le> i \<and> j + 1 \<noteq> h - 1 + 1 \<and> x = ((h - 1) + 1, (j + 1) - 1)"
+ by auto
+ thus "x \<in> (\<lambda>(h, j). (j + 1, h - 1)) ` {(h, j). h \<le> i + 2 \<and> j < h \<and> j \<le> i \<and> h \<noteq> j + 1}"
+ by (auto simp: image_iff)
+ qed
+ qed
+ show "inj_on (\<lambda>(h, j). (j + 1, h - 1)) {(h, j). h \<le> i + 2 \<and> j < h \<and> j \<le> i \<and> h \<noteq> j + 1}"
+ proof
+ fix x y::"nat\<times>nat"
+ assume "x \<in> {(h, j). h \<le> i + 2 \<and> j < h \<and> j \<le> i \<and> h \<noteq> j + 1}" "y \<in> {(h, j). h \<le> i + 2 \<and> j < h \<and> j \<le> i \<and> h \<noteq> j + 1}"
+ thus "(case x of (h, j) \<Rightarrow> (j + 1, h - 1)) = (case y of (h, j) \<Rightarrow> (j + 1, h - 1)) \<Longrightarrow> x = y"
+ by auto
+ qed
+ show "\<And>x. x \<in> {(h, j). h \<le> i + 2 \<and> j < h \<and> j \<le> i \<and> h \<noteq> j + 1} \<Longrightarrow>
+ (case case x of (h, j) \<Rightarrow> (j + 1, h - 1) of
+ (h, j) \<Rightarrow> coeff f (j + 1) * coeff g (i + 1 - j) * coeff f (h - 1) * coeff g (i + 1 - h)) =
+ (case x of (h, j) \<Rightarrow> coeff f h * coeff g (i + 2 - h) * coeff f j * coeff g (i - j))"
+ by fastforce
+ qed
+ have 2: "(\<Sum>h = 1..i + 1. coeff f h * coeff g (i + 2 - h) * coeff f (h - 1) * coeff g (i + 1 - h)) =
+ (\<Sum>(h, j)\<in>{(h, j). h \<le> i + 2 \<and> j < h \<and> j \<le> i \<and> h = j + 1}.
+ coeff f h * coeff g (i + 2 - h) * coeff f j * coeff g (i - j))"
+ proof (rule sum.reindex_cong)
+ show "{1..i + 1} = fst ` {(h, j). h \<le> i + 2 \<and> j < h \<and> j \<le> i \<and> h = j + 1}"
+ proof
+ show "{1..i + 1} \<subseteq> fst ` {(h, j). h \<le> i + 2 \<and> j < h \<and> j \<le> i \<and> h = j + 1}"
+ proof
+ fix x
+ assume "x \<in> {1..i + 1}"
+ hence "x \<le> i + 2 \<and> x - 1 < x \<and> x - 1 \<le> i \<and> x = x - 1 + 1 \<and> x = fst (x, x-1)"
+ by auto
+ thus "x \<in> fst ` {(h, j). h \<le> i + 2 \<and> j < h \<and> j \<le> i \<and> h = j + 1}"
+ by blast
+ qed
+ show "fst ` {(h, j). h \<le> i + 2 \<and> j < h \<and> j \<le> i \<and> h = j + 1} \<subseteq> {1..i + 1}"
+ by force
+ qed
+ show "inj_on fst {(h, j). h \<le> i + 2 \<and> j < h \<and> j \<le> i \<and> h = j + 1}"
+ proof
+ fix x y
+ assume "x \<in> {(h, j). h \<le> i + 2 \<and> j < h \<and> j \<le> i \<and> h = j + 1}"
+ "y \<in> {(h, j). h \<le> i + 2 \<and> j < h \<and> j \<le> i \<and> h = j + 1}"
+ hence "x = (fst x, fst x - 1)" "y = (fst y, fst y - 1)" "fst x > 0" "fst y > 0"
+ by auto
+ thus "fst x = fst y \<Longrightarrow> x = y" by presburger
+ qed
+ show "\<And>x. x \<in> {(h, j). h \<le> i + 2 \<and> j < h \<and> j \<le> i \<and> h = j + 1} \<Longrightarrow>
+ coeff f (fst x) * coeff g (i + 2 - fst x) * coeff f (fst x - 1) * coeff g (i + 1 - fst x) =
+ (case x of (h, j) \<Rightarrow> coeff f h * coeff g (i + 2 - h) * coeff f j * coeff g (i - j))"
+ by fastforce
+ qed
+ have H: "{(h, j). h \<le> i + 2 \<and> j < h \<and> j \<le> i \<and> h \<noteq> j + 1} \<subseteq> {0..i+2}\<times>{0..i}"
+ "{(h, j). h \<le> i + 2 \<and> j < h \<and> j \<le> i \<and> h = j + 1} \<subseteq> {0..i+2}\<times>{0..i}"
+ "finite ({0..i+2}\<times>{0..i})"
+ by (fastforce, fastforce, fastforce)
+ have "finite {(h, j). h \<le> i + 2 \<and> j < h \<and> j \<le> i \<and> h \<noteq> j + 1}"
+ "finite {(h, j). h \<le> i + 2 \<and> j < h \<and> j \<le> i \<and> h = j + 1}"
+ apply (rule finite_subset) using H apply (simp, simp)
+ apply (rule finite_subset) using H apply (simp, simp)
+ done
+ thus ?thesis
+ apply (subst 1, subst 2, subst sum.union_disjoint[symmetric])
+ apply auto[3]
+ apply (rule sum.cong)
+ by auto
+ qed
+ ultimately show ?thesis
+ by algebra
+ qed
+
+ text \<open>$\dots = \sum_{h \leq j} f_hf_j\left(g_{i+1-h}g_{i+1-j} - g_{i+2-h}g_{i-j}\right) +
+ \sum_{h \leq j} f_{j+1}f_{h-1}\left(g_{i-j}g_{i+2-h} - g_{i+1-j}f_jg_{i+1-h}\right)$
+
+ Note we have to also consider the edge cases caused by making these sums finite.\<close>
+ also have "... =
+ (\<Sum>(h, j)\<in>{(h, j). j = i + 1 \<and> h \<le> j}.
+ coeff f h * coeff f j * (coeff g (i + 1 - h) * coeff g (i + 1 - j))) +
+ (\<Sum>(h, j)\<in>{(h, j). j \<le> i \<and> h \<le> j}.
+ coeff f h * coeff f j * (coeff g (i + 1 - h) * coeff g (i + 1 - j) - coeff g (i + 2 - h) * coeff g (i - j))) +
+ (\<Sum>(h, j)\<in>{(h, j). j \<le> i \<and> h \<le> j \<and> 0 < h}.
+ coeff f (j+1) * coeff f (h-1) * (coeff g (i - j) * coeff g (i + 2 - h) - coeff g (i + 1 - j) * coeff g (i + 1 - h))) -
+ ((\<Sum>(h, j)\<in>{(h, j). j = i + 1 \<and> h \<le> j \<and> 0 < h}.
+ coeff f (j+1) * coeff g (i + 1 - j) * coeff f (h-1) * coeff g (i + 1 - h)))" (is "?L = ?R")
+ proof -
+ have "?R =
+ (\<Sum>(h, j)\<in>{(h, j). j = i + 1 \<and> h \<le> j}.
+ coeff f h * coeff f j * (coeff g (i + 1 - h) * coeff g (i + 1 - j))) +
+ ((\<Sum>(h, j)\<in>{(h, j). j \<le> i \<and> h \<le> j}.
+ coeff f h * coeff f j * coeff g (i + 1 - h) * coeff g (i + 1 - j)) -
+ (\<Sum>(h, j)\<in>{(h, j). j \<le> i \<and> h \<le> j}.
+ coeff f h * coeff f j * coeff g (i + 2 - h) * coeff g (i - j))) +
+ ((\<Sum>(h, j)\<in>{(h, j). j \<le> i \<and> h \<le> j \<and> 0 < h}.
+ coeff f (j+1) * coeff f (h-1) * coeff g (i - j) * coeff g (i + 2 - h)) -
+ (\<Sum>(h, j)\<in>{(h, j). j \<le> i \<and> h \<le> j \<and> 0 < h}.
+ coeff f (j+1) * coeff f (h-1) * coeff g (i + 1 - j) * coeff g (i + 1 - h))) -
+ ((\<Sum>(h, j)\<in>{(h, j). j = i + 1 \<and> h \<le> j \<and> 0 < h}.
+ coeff f (j+1) * coeff g (i + 1 - j) * coeff f (h-1) * coeff g (i + 1 - h)))"
+ apply (subst sum_subtractf[symmetric], subst sum_subtractf[symmetric])
+ by (auto simp: algebra_simps split_beta)
+ also have "... =
+ ((\<Sum>(h, j)\<in>{(h, j). j = i + 1 \<and> h \<le> j}.
+ coeff f h * coeff f j * (coeff g (i + 1 - h) * coeff g (i + 1 - j))) +
+ (\<Sum>(h, j)\<in>{(h, j). j \<le> i \<and> h \<le> j}.
+ coeff f h * coeff f j * (coeff g (i + 1 - h) * coeff g (i + 1 - j)))) -
+ (\<Sum>(h, j)\<in>{(h, j). j \<le> i \<and> h \<le> j}.
+ coeff f h * coeff f j * coeff g (i + 2 - h) * coeff g (i - j)) +
+ (\<Sum>(h, j)\<in>{(h, j). j \<le> i \<and> h \<le> j \<and> 0 < h}.
+ coeff f (j + 1) * coeff f (h - 1) * coeff g (i - j) * coeff g (i + 2 - h)) -
+ ((\<Sum>(h, j)\<in>{(h, j). j \<le> i \<and> h \<le> j \<and> 0 < h}.
+ coeff f (j + 1) * coeff g (i + 1 - j) * coeff f (h - 1) * coeff g (i + 1 - h)) +
+ (\<Sum>(h, j)\<in>{(h, j). j = i + 1 \<and> h \<le> j \<and> 0 < h}.
+ coeff f (j + 1) * coeff g (i + 1 - j) * coeff f (h - 1) * coeff g (i + 1 - h)))"
+ by (auto simp: algebra_simps)
+ also have "... = ?L"
+ proof -
+ have "(\<Sum>(h, j)\<in>{(h, j). j = i + 1 \<and> h \<le> j}.
+ coeff f h * coeff f j * (coeff g (i + 1 - h) * coeff g (i + 1 - j))) +
+ (\<Sum>(h, j)\<in>{(h, j). j \<le> i \<and> h \<le> j}.
+ coeff f h * coeff f j * (coeff g (i + 1 - h) * coeff g (i + 1 - j))) =
+ (\<Sum>(h, j)\<in>{(h, j). j \<le> i + 1 \<and> h \<le> j}.
+ coeff f h * coeff f j * (coeff g (i + 1 - h) * coeff g (i + 1 - j)))"
+ proof (subst sum.union_disjoint[symmetric])
+ have "{(h, j). j = i + 1 \<and> h \<le> j} \<subseteq> {..i + 1} \<times> {..i + 1}"
+ "{(h, j). j \<le> i \<and> h \<le> j} \<subseteq> {..i + 1} \<times> {..i + 1}"
+ by (fastforce, fastforce)
+ thus "finite {(h, j). j = i + 1 \<and> h \<le> j}" "finite {(h, j). j \<le> i \<and> h \<le> j}"
+ by (auto simp: finite_subset)
+ show "{(h, j). j = i + 1 \<and> h \<le> j} \<inter> {(h, j). j \<le> i \<and> h \<le> j} = {}"
+ by fastforce
+ qed (rule sum.cong, auto)
+ moreover have "(\<Sum>(h, j)\<in>{(h, j). j \<le> i \<and> h \<le> j \<and> 0 < h}.
+ coeff f (j + 1) * coeff g (i + 1 - j) * coeff f (h - 1) * coeff g (i + 1 - h)) +
+ (\<Sum>(h, j)\<in>{(h, j). j = i + 1 \<and> h \<le> j \<and> 0 < h}.
+ coeff f (j + 1) * coeff g (i + 1 - j) * coeff f (h - 1) * coeff g (i + 1 - h)) =
+ (\<Sum>(h, j)\<in>{(h, j). j \<le> i + 1 \<and> h \<le> j \<and> 0 < h}.
+ coeff f (j + 1) * coeff g (i + 1 - j) * coeff f (h - 1) * coeff g (i + 1 - h))"
+ proof (subst sum.union_disjoint[symmetric])
+ have "{(h, j). j \<le> i \<and> h \<le> j \<and> 0 < h} \<subseteq> {..i + 1} \<times> {..i + 1}"
+ "{(h, j). j = i + 1 \<and> h \<le> j \<and> 0 < h} \<subseteq> {..i + 1} \<times> {..i + 1}"
+ by (fastforce, fastforce)
+ thus "finite {(h, j). j \<le> i \<and> h \<le> j \<and> 0 < h}" "finite {(h, j). j = i + 1 \<and> h \<le> j \<and> 0 < h}"
+ by (auto simp: finite_subset)
+ show "{(h, j). j \<le> i \<and> h \<le> j \<and> 0 < h} \<inter> {(h, j). j = i + 1 \<and> h \<le> j \<and> 0 < h} = {}"
+ by fastforce
+ qed (rule sum.cong, auto)
+ ultimately show ?thesis
+ by (auto simp: algebra_simps)
+ qed
+ finally show ?thesis by presburger
+ qed
+
+ text \<open>$\dots = \sum_{h \leq j} \left(f_hf_j - f_{j+1}f_{h-1}\right)
+ \left(g_{i+1-h}g_{i+1-j} - g_{i+2-h}g_{i-j}\right)$\<close>
+ also have "... =
+ (\<Sum>(h, j)\<in>{(h, j). j \<le> i \<and> h \<le> j \<and> 0 < h}.
+ -(coeff f h * coeff f j - coeff f (j+1) * coeff f (h-1)) * (coeff g (i - j) * coeff g (i + 2 - h) - coeff g (i + 1 - j) * coeff g (i + 1 - h))) +
+ (\<Sum>(h, j)\<in>{(h, j). j \<le> i \<and> h \<le> j \<and> h = 0}.
+ coeff f h * coeff f j * (coeff g (i + 1 - h) * coeff g (i + 1 - j) - coeff g (i + 2 - h) * coeff g (i - j))) +
+ (\<Sum>(h, j)\<in>{(h, j). j = i + 1 \<and> h \<le> j}.
+ coeff f h * coeff f j * (coeff g (i + 1 - h) * coeff g (i + 1 - j))) -
+ ((\<Sum>(h, j)\<in>{(h, j). j = i + 1 \<and> h \<le> j \<and> 0 < h}.
+ coeff f (j+1) * coeff g (i + 1 - j) * coeff f (h-1) * coeff g (i + 1 - h)))" (is "?L = ?R")
+ proof -
+ have "(\<Sum>(h, j)\<in>{(h, j). j \<le> i \<and> h \<le> j}.
+ coeff f h * coeff f j *
+ (coeff g (i + 1 - h) * coeff g (i + 1 - j) - coeff g (i + 2 - h) * coeff g (i - j))) =
+ (\<Sum>(h, j)\<in>{(h, j). j \<le> i \<and> h \<le> j \<and> 0 < h}.
+ coeff f h * coeff f j *
+ (coeff g (i + 1 - h) * coeff g (i + 1 - j) - coeff g (i + 2 - h) * coeff g (i - j))) +
+ (\<Sum>(h, j)\<in>{(h, j). j \<le> i \<and> h \<le> j \<and> 0 = h}.
+ coeff f h * coeff f j *
+ (coeff g (i + 1 - h) * coeff g (i + 1 - j) - coeff g (i + 2 - h) * coeff g (i - j)))"
+ proof (subst sum.union_disjoint[symmetric])
+ have "{(h, j). j \<le> i \<and> h \<le> j \<and> 0 < h} \<subseteq> {..i}\<times>{..i}" "{(h, j). j \<le> i \<and> h \<le> j \<and> 0 = h} \<subseteq> {..i}\<times>{..i}"
+ by (force, force)
+ thus "finite {(h, j). j \<le> i \<and> h \<le> j \<and> 0 < h}" "finite {(h, j). j \<le> i \<and> h \<le> j \<and> 0 = h}"
+ by (auto simp: finite_subset)
+ show "{(h, j). j \<le> i \<and> h \<le> j \<and> 0 < h} \<inter> {(h, j). j \<le> i \<and> h \<le> j \<and> 0 = h} = {}"
+ by fast
+ qed (rule sum.cong, auto)
+
+ moreover have "(\<Sum>(h, j)\<in>{(h, j). j \<le> i \<and> h \<le> j \<and> 0 < h}.
+ (-coeff f h * coeff f j + coeff f (j + 1) * coeff f (h - 1)) *
+ (coeff g (i - j) * coeff g (i + 2 - h) - coeff g (i + 1 - j) * coeff g (i + 1 - h))) =
+ (\<Sum>(h, j)\<in>{(h, j). j \<le> i \<and> h \<le> j \<and> 0 < h}.
+ coeff f h * coeff f j * (coeff g (i + 1 - h) * coeff g (i + 1 - j) - coeff g (i + 2 - h) * coeff g (i - j))) +
+ (\<Sum>(h, j)\<in>{(h, j). j \<le> i \<and> h \<le> j \<and> 0 < h}.
+ coeff f (j + 1) * coeff f (h - 1) *
+ (coeff g (i - j) * coeff g (i + 2 - h) - coeff g (i + 1 - j) * coeff g (i + 1 - h)))"
+ by (subst sum.distrib[symmetric], rule sum.cong, fast, auto simp: algebra_simps)
+
+ ultimately show ?thesis
+ by (auto simp: algebra_simps)
+ qed
+
+ text \<open>$\dots \geq 0$ by \<open>normal_poly_general_coeff_mult\<close>\<close>
+ also have "... \<ge> 0"
+ proof -
+ have "0 \<le> (\<Sum>(h, j)\<in>{(h, j). j \<le> i \<and> h \<le> j \<and> 0 < h}.
+ - (coeff f h * coeff f j - coeff f (j + 1) * coeff f (h - 1)) *
+ (coeff g (i - j) * coeff g (i + 2 - h) - coeff g (i + 1 - j) * coeff g (i + 1 - h)))"
+ proof (rule sum_nonneg)
+ fix x assume "x \<in> {(h, j). j \<le> i \<and> h \<le> j \<and> 0 < h}"
+ then obtain h j where H: "x = (h, j)" "j \<le> i" "h \<le> j" "0 < h" by fast
+ hence "h - 1 \<le> j - 1" by force
+ hence 1: "coeff f h * coeff f j - coeff f (j + 1) * coeff f (h - 1) \<ge> 0"
+ using normal_poly_general_coeff_mult[OF hf, of "h-1" "j-1"] H
+ by (auto simp: algebra_simps)
+ from H have "i - j \<le> i - h" by force
+ hence 2: "coeff g (i - j) * coeff g (i + 2 - h) - coeff g (i + 1 - j) * coeff g (i + 1 - h) \<le> 0"
+ using normal_poly_general_coeff_mult[OF hg, of "i - j" "i - h"] H
+ by (smt (verit, del_insts) Nat.add_diff_assoc2 le_trans)
+ show "0 \<le> (case x of
+ (h, j) \<Rightarrow>
+ - (coeff f h * coeff f j - coeff f (j + 1) * coeff f (h - 1)) *
+ (coeff g (i - j) * coeff g (i + 2 - h) -
+ coeff g (i + 1 - j) * coeff g (i + 1 - h)))"
+ apply (subst H(1), subst split, rule mult_nonpos_nonpos, subst neg_le_0_iff_le)
+ subgoal using 1 by blast
+ subgoal using 2 by blast
+ done
+ qed
+ moreover have "0 \<le> (\<Sum>(h, j)\<in>{(h, j). j \<le> i \<and> h \<le> j \<and> h = 0}.
+ coeff f h * coeff f j *
+ (coeff g (i + 1 - h) * coeff g (i + 1 - j) - coeff g (i + 2 - h) * coeff g (i - j)))"
+ proof (rule sum_nonneg)
+ fix x assume "x \<in> {(h, j). j \<le> i \<and> h \<le> j \<and> h = 0}"
+ then obtain h j where H: "x = (h, j)" "j \<le> i" "h \<le> j" "h = 0" by fast
+ have 1: "coeff f h * coeff f j \<ge> 0"
+ by (simp add: hf normal_coeff_nonneg)
+ from H have "i - j \<le> i - h" by force
+ hence 2: "coeff g (i - j) * coeff g (i + 2 - h) - coeff g (i + 1 - j) * coeff g (i + 1 - h) \<le> 0"
+ using normal_poly_general_coeff_mult[OF hg, of "i - j" "i - h"] H
+ by (smt (verit, del_insts) Nat.add_diff_assoc2 le_trans)
+ show "0 \<le> (case x of
+ (h, j) \<Rightarrow>
+ coeff f h * coeff f j *
+ (coeff g (i + 1 - h) * coeff g (i + 1 - j) -
+ coeff g (i + 2 - h) * coeff g (i - j)))"
+ apply (subst H(1), subst split, rule mult_nonneg_nonneg)
+ subgoal using 1 by blast
+ subgoal using 2 by argo
+ done
+ qed
+ moreover have "0 \<le> (\<Sum>(h, j)\<in>{(h, j). j = i + 1 \<and> h \<le> j}. coeff f h * coeff f j * (coeff g (i + 1 - h) * coeff g (i + 1 - j))) -
+ (\<Sum>(h, j)\<in>{(h, j). j = i + 1 \<and> h \<le> j \<and> 0 < h}.
+ coeff f (j + 1) * coeff g (i + 1 - j) * coeff f (h - 1) * coeff g (i + 1 - h))"
+ proof -
+ have "(\<Sum>(h, j)\<in>{(h, j). j = i + 1 \<and> h \<le> j}. coeff f h * coeff f j * (coeff g (i + 1 - h) * coeff g (i + 1 - j))) -
+ (\<Sum>(h, j)\<in>{(h, j). j = i + 1 \<and> h \<le> j \<and> 0 < h}.
+ coeff f (j + 1) * coeff g (i + 1 - j) * coeff f (h - 1) * coeff g (i + 1 - h)) =
+ (\<Sum>(h, j)\<in>{(h, j). j = i + 1 \<and> h \<le> j \<and> h = 0}. coeff f h * coeff f j * (coeff g (i + 1 - h) * coeff g (i + 1 - j))) +
+ (\<Sum>(h, j)\<in>{(h, j). j = i + 1 \<and> h \<le> j \<and> 0 < h}. coeff f h * coeff f j * (coeff g (i + 1 - h) * coeff g (i + 1 - j))) -
+ (\<Sum>(h, j)\<in>{(h, j). j = i + 1 \<and> h \<le> j \<and> 0 < h}.
+ coeff f (j + 1) * coeff g (i + 1 - j) * coeff f (h - 1) * coeff g (i + 1 - h))"
+ proof (subst sum.union_disjoint[symmetric])
+ have "{(h, j). j = i + 1 \<and> h \<le> j \<and> h = 0} = {(0, i + 1)}"
+ "{(h, j). j = i + 1 \<and> h \<le> j \<and> 0 < h} = {1..i+1} \<times> {i + 1}"
+ by (fastforce, force)
+ thus "finite {(h, j). j = i + 1 \<and> h \<le> j \<and> h = 0}"
+ "finite {(h, j). j = i + 1 \<and> h \<le> j \<and> 0 < h}"
+ by auto
+ show "{(h, j). j = i + 1 \<and> h \<le> j \<and> h = 0} \<inter> {(h, j). j = i + 1 \<and> h \<le> j \<and> 0 < h} = {}"
+ by fastforce
+ have "{(h, j). j = i + 1 \<and> h \<le> j \<and> h = 0} \<union> {(h, j). j = i + 1 \<and> h \<le> j \<and> 0 < h} = {(h, j). j = i + 1 \<and> h \<le> j}"
+ by fastforce
+ thus "(\<Sum>(h, j)\<in>{(h, j). j = i + 1 \<and> h \<le> j}. coeff f h * coeff f j * (coeff g (i + 1 - h) * coeff g (i + 1 - j))) -
+ (\<Sum>(h, j)\<in>{(h, j). j = i + 1 \<and> h \<le> j \<and> 0 < h}.
+ coeff f (j + 1) * coeff g (i + 1 - j) * coeff f (h - 1) * coeff g (i + 1 - h)) =
+ (\<Sum>(h, j)\<in>{(h, j). j = i + 1 \<and> h \<le> j \<and> h = 0} \<union> {(h, j). j = i + 1 \<and> h \<le> j \<and> 0 < h}.
+ coeff f h * coeff f j * (coeff g (i + 1 - h) * coeff g (i + 1 - j))) -
+ (\<Sum>(h, j)\<in>{(h, j). j = i + 1 \<and> h \<le> j \<and> 0 < h}.
+ coeff f (j + 1) * coeff g (i + 1 - j) * coeff f (h - 1) * coeff g (i + 1 - h))"
+ by presburger
+ qed
+ also have "... =
+ (\<Sum>(h, j)\<in>{(h, j). j = i + 1 \<and> h \<le> j \<and> h = 0}. coeff f h * coeff f j * (coeff g (i + 1 - h) * coeff g (i + 1 - j))) +
+ (\<Sum>(h, j)\<in>{(h, j). j = i + 1 \<and> h \<le> j \<and> 0 < h}.
+ (coeff f h * coeff f j - coeff f (j + 1) * coeff f (h - 1)) * (coeff g (i + 1 - h) * coeff g (i + 1 - j)))"
+ by (subst add_diff_eq[symmetric], subst sum_subtractf[symmetric], subst add_left_cancel, rule sum.cong, auto simp: algebra_simps)
+ also have "... \<ge> 0"
+ proof (rule add_nonneg_nonneg)
+ show "0 \<le> (\<Sum>(h, j)\<in>{(h, j). j = i + 1 \<and> h \<le> j \<and> 0 < h}.
+ (coeff f h * coeff f j - coeff f (j + 1) * coeff f (h - 1)) *
+ (coeff g (i + 1 - h) * coeff g (i + 1 - j)))"
+ proof (rule sum_nonneg)
+ fix x assume "x \<in> {(h, j). j = i + 1 \<and> h \<le> j \<and> 0 < h}"
+ then obtain h j where H: "x = (h, j)" "j = i + 1" "h \<le> j" "0 < h" by fast
+ hence "h - 1 \<le> j - 1" by force
+ hence 1: "coeff f h * coeff f j - coeff f (j + 1) * coeff f (h - 1) \<ge> 0"
+ using normal_poly_general_coeff_mult[OF hf, of "h-1" "j-1"] H
+ by (auto simp: algebra_simps)
+ hence 2: "0 \<le> coeff g (i + 1 - h) * coeff g (i + 1 - j)"
+ by (meson hg mult_nonneg_nonneg normal_coeff_nonneg)
+ show "0 \<le> (case x of
+ (h, j) \<Rightarrow>
+ (coeff f h * coeff f j - coeff f (j + 1) * coeff f (h - 1)) *
+ (coeff g (i + 1 - h) * coeff g (i + 1 - j)))"
+ apply (subst H(1), subst split, rule mult_nonneg_nonneg)
+ subgoal using 1 by blast
+ subgoal using 2 by blast
+ done
+ qed
+ qed (rule sum_nonneg, auto simp: hf hg normal_coeff_nonneg)[1]
+ finally show ?thesis .
+ qed
+ ultimately show ?thesis by auto
+ qed
+ finally show "coeff (f * g) i * coeff (f * g) (i + 2) \<le> (coeff (f * g) (i + 1))\<^sup>2" by (auto simp: power2_eq_square)
+ qed
+ moreover have "\<And>i j k. i \<le> j \<Longrightarrow> j \<le> k \<Longrightarrow> 0 < coeff (f*g) i \<Longrightarrow> 0 < coeff (f*g) k \<Longrightarrow> 0 < coeff (f*g) j"
+ proof -
+ fix j k
+ assume "0 < coeff (f * g) k"
+ hence "k \<le> degree (f * g)" using le_degree by force
+ moreover assume "j \<le> k"
+ ultimately have "j \<le> degree (f * g)" by auto
+ hence 1: "j \<le> degree f + degree g"
+ by (simp add: degree_mult_eq hf hg normal_non_zero)
+ show "0 < coeff (f * g) j"
+ apply (subst coeff_mult, rule sum_pos2[of _ "min j (degree f)"], simp, simp)
+ apply (rule mult_pos_pos, rule normal_of_no_zero_root, simp add: hf0, simp)
+ using hf apply auto[1]
+ apply (rule normal_of_no_zero_root)
+ apply (simp add: hg0)
+ using 1 apply force
+ using hg apply auto[1]
+ by (simp add: hf hg normal_coeff_nonneg)
+ qed
+ ultimately show "normal_poly (f*g)"
+ by (rule normal_polyI)
+ qed
+ qed
+ qed
+qed
+
+lemma normal_poly_of_roots:
+ fixes p::"real poly"
+ assumes "\<And>z. poly (map_poly complex_of_real p) z = 0
+ \<Longrightarrow> Re z \<le> 0 \<and> 4*(Re z)^2 \<ge> (cmod z)^2"
+ and "lead_coeff p = 1"
+ shows "normal_poly p"
+ using assms
+proof (induction p rule: real_poly_roots_induct)
+ fix p::"real poly" and x::real
+ assume "lead_coeff (p * [:- x, 1:]) = 1"
+ hence 1: "lead_coeff p = 1"
+ by (metis coeff_degree_mult lead_coeff_pCons(1) mult_cancel_left1 pCons_one zero_neq_one)
+ assume h: "(\<And>z. poly (map_poly complex_of_real (p * [:- x, 1:])) z = 0 \<Longrightarrow>
+ Re z \<le> 0 \<and> (cmod z)\<^sup>2 \<le> 4 * (Re z)\<^sup>2)"
+ hence 2: "(\<And>z. poly (map_poly complex_of_real p) z = 0 \<Longrightarrow>
+ Re z \<le> 0 \<and> (cmod z)\<^sup>2 \<le> 4 * (Re z)\<^sup>2)"
+ by (metis four_x_squared mult_zero_left of_real_poly_map_mult poly_mult)
+ have 3: "normal_poly [:-x, 1:]"
+ apply (subst linear_normal_iff,
+ subst Re_complex_of_real[symmetric], rule conjunct1)
+ by (rule h[of x], subst of_real_poly_map_poly[symmetric], force)
+ assume "(\<And>z. poly (map_poly complex_of_real p) z = 0
+ \<Longrightarrow> Re z \<le> 0 \<and> (cmod z)\<^sup>2 \<le> 4 * (Re z)\<^sup>2) \<Longrightarrow>
+ lead_coeff p = 1 \<Longrightarrow> normal_poly p"
+ hence "normal_poly p" using 1 2 by fast
+ then show "normal_poly (p * [:-x, 1:])"
+ using 3 by (rule normal_mult)
+next
+ fix p::"real poly" and a b::real
+ assume "lead_coeff (p * [:a * a + b * b, - 2 * a, 1:]) = 1"
+ hence 1: "lead_coeff p = 1"
+ by (smt (verit) coeff_degree_mult lead_coeff_pCons(1) mult_cancel_left1 pCons_eq_0_iff pCons_one)
+ assume h: "(\<And>z. poly (map_poly complex_of_real (p * [:a * a + b * b, - 2 * a, 1:])) z = 0 \<Longrightarrow>
+ Re z \<le> 0 \<and> (cmod z)\<^sup>2 \<le> 4 * (Re z)\<^sup>2)"
+ hence 2: "(\<And>z. poly (map_poly complex_of_real p) z = 0 \<Longrightarrow>
+ Re z \<le> 0 \<and> (cmod z)\<^sup>2 \<le> 4 * (Re z)\<^sup>2)"
+ proof -
+ fix z :: complex
+ assume "poly (map_poly complex_of_real p) z = 0"
+ then have "\<forall>q. 0 = poly (map_poly complex_of_real (p * q)) z"
+ by simp
+ then show "Re z \<le> 0 \<and> (cmod z)\<^sup>2 \<le> 4 * (Re z)\<^sup>2"
+ using h by presburger
+ qed
+ have 3: "[:a * a + b * b, - 2 * a, 1:] = [:cmod (a + \<i>*b) ^ 2, -2 * Re (a + \<i>*b), 1:]"
+ by (force simp: cmod_def power2_eq_square)
+ interpret map_poly_idom_hom complex_of_real ..
+ have 4: "normal_poly [:a * a + b * b, - 2 * a, 1:]"
+ apply (subst 3, subst quadratic_normal_iff)
+ apply (rule h, unfold hom_mult poly_mult)
+ by (auto simp: algebra_simps)
+ assume "(\<And>z. poly (map_poly complex_of_real p) z = 0 \<Longrightarrow> Re z \<le> 0 \<and> (cmod z)\<^sup>2 \<le> 4 * (Re z)\<^sup>2) \<Longrightarrow>
+ lead_coeff p = 1 \<Longrightarrow> normal_poly p"
+ hence "normal_poly p" using 1 2 by fast
+ then show "normal_poly (p * [:a * a + b * b, - 2 * a, 1:])"
+ using 4 by (rule normal_mult)
+next
+ fix a::real
+ assume "lead_coeff [:a:] = 1"
+ moreover have "\<And>i j k.
+ lead_coeff [:a:] = 1 \<Longrightarrow>
+ i \<le> j \<Longrightarrow>
+ j \<le> k \<Longrightarrow> 0 < coeff [:a:] i \<Longrightarrow> 0 < coeff [:a:] k \<Longrightarrow> 0 < coeff [:a:] j"
+ by (metis bot_nat_0.extremum_uniqueI coeff_eq_0 degree_pCons_0 leI
+ less_numeral_extra(3))
+ ultimately show "normal_poly [:a:]"
+ apply (subst normal_polyI)
+ by (auto simp:pCons_one)
+qed
+
+lemma normal_changes:
+ fixes f::"real poly"
+ assumes hf: "normal_poly f" and hx: "x > 0"
+ defines "df \<equiv> degree f"
+ shows "changes (coeffs (f*[:-x,1:])) = 1"
+ using df_def hf
+proof (induction df arbitrary: f)
+ case 0
+ then obtain a where "f = [:a:]" using degree_0_iff by auto
+ thus "changes (coeffs (f*[:-x, 1:])) = 1"
+ using normal_non_zero[OF \<open>normal_poly f\<close>] hx
+ by (auto simp: algebra_simps zero_less_mult_iff mult_less_0_iff)
+next
+ case (Suc df)
+ then show ?case
+ proof (cases "poly f 0 = 0")
+ assume "poly f 0 = 0" and hf:"normal_poly f"
+ moreover then obtain f' where hdiv: "f = f'*[:0, 1:]"
+ by (smt (verit) dvdE mult.commute poly_eq_0_iff_dvd)
+ ultimately have hf': "normal_poly f'" using normal_divide_x by blast
+ assume "Suc df = degree f"
+ hence "degree f' = df" using hdiv normal_non_zero[OF hf'] by (auto simp: degree_mult_eq)
+ moreover assume "\<And>f::real poly. df = degree f \<Longrightarrow> normal_poly f \<Longrightarrow> changes (coeffs (f * [:- x, 1:])) = 1"
+ ultimately have "changes (coeffs (f' * [:- x, 1:])) = 1" using hf' by fast
+ thus "changes (coeffs (f * [:- x, 1:])) = 1"
+ apply (subst hdiv, subst mult_pCons_right, subst smult_0_left, subst add_0)
+ apply (subst mult_pCons_left, subst smult_0_left, subst add_0)
+ by (subst changes_pCons, auto)
+ next
+ assume hf:"normal_poly f" and "poly f 0 \<noteq> 0"
+ hence h': "\<And>i. i \<le> degree f \<Longrightarrow> coeff f i > 0"
+ by (auto simp: normal_of_no_zero_root)
+ hence "\<And>i. i < degree f - 1 \<Longrightarrow> (coeff f i)/(coeff f (i+1)) \<le> (coeff f (i+1))/(coeff f (i+2))"
+ using normal_poly_coeff_mult[OF hf] normal_coeff_nonneg[OF hf]
+ by (auto simp: divide_simps power2_eq_square)
+ hence h'': "\<And>i. i < degree f - 1 \<Longrightarrow> (coeff f i)/(coeff f (i+1)) - x \<le> (coeff f (i+1))/(coeff f (i+2)) - x"
+ by fastforce
+ have hdeg: "degree (pCons 0 f - smult x f) = degree f + 1"
+ apply (subst diff_conv_add_uminus)
+ apply (subst degree_add_eq_left)
+ by (auto simp: hf normal_non_zero)
+
+ let ?f = "\<lambda> z w. \<lambda>i. if i=0 then z/(x * coeff f 0) else (if i = degree (pCons 0 f - smult x f) then w/(lead_coeff f) else inverse (coeff f i))"
+
+ have 1: "\<And>z w. 0 < z \<Longrightarrow> 0 < w \<Longrightarrow> changes (coeffs (f * [:-x, 1:])) =
+ changes (-z # map (\<lambda>i. (coeff f (i-1))/(coeff f i) - x) [1..<degree (pCons 0 f - smult x f)] @ [w])"
+ proof -
+ fix z w :: real
+ assume hz: "0 < z" and hw: "0 < w"
+
+ have "-z # map (\<lambda>i. (coeff f (i-1))/(coeff f i) - x) [1..<degree (pCons 0 f - smult x f)] @ [w] =
+ map (\<lambda>i. if i = 0 then -z else if i = degree (pCons 0 f - smult x f) then w else
+ (coeff f (i-1))/(coeff f i) - x) [0..<degree (pCons 0 f - smult x f) + 1]"
+ proof (rule nth_equalityI)
+ fix i assume "i < length (- z # map (\<lambda>i. coeff f (i - 1) / coeff f i - x) [1..<degree (pCons 0 f - smult x f)] @ [w])"
+ hence "i \<le> degree (pCons 0 f - smult x f)"
+ using hdeg Suc.hyps(2) by auto
+ then consider (a)"i = 0" | (b)"(0 < i \<and> i < degree (pCons 0 f - smult x f))" |
+ (c)"i = degree (pCons 0 f - smult x f)"
+ by fastforce
+ then show "(- z #
+ map (\<lambda>i. coeff f (i - 1) / coeff f i - x)
+ [1..<degree (pCons 0 f - smult x f)] @
+ [w]) ! i =
+ map (\<lambda>i. if i = 0 then - z
+ else if i = degree (pCons 0 f - smult x f) then w
+ else coeff f (i - 1) / coeff f i - x)
+ [0..<degree (pCons 0 f - smult x f) + 1] ! i"
+ apply (cases)
+ by (auto simp: nth_append)
+ qed (force simp: hdeg)
+
+ also have "... = [?f z w i * (nth_default 0 (coeffs (f * [:-x, 1:])) i).
+ i \<leftarrow> [0..<Suc (degree (pCons 0 f - smult x f))]]"
+ proof (rule map_cong)
+ fix i assume "i \<in> set [0..<Suc (degree (pCons 0 f - smult x f))]"
+ then consider (a)"i = 0" | (b)"(0 \<noteq> i \<and> i < degree (pCons 0 f - smult x f))" |
+ (c)"i = degree (pCons 0 f - smult x f)"
+ by fastforce
+ then show "(if i = 0 then - z
+ else if i = degree (pCons 0 f - smult x f) then w
+ else coeff f (i - 1) / coeff f i - x) =
+ (if i = 0 then z / (x * coeff f 0)
+ else if i = degree (pCons 0 f - smult x f) then w / lead_coeff f
+ else inverse (coeff f i)) *
+ nth_default 0 (coeffs (f * [:- x, 1:])) i"
+ proof (cases)
+ case (a)
+ thus ?thesis using hx \<open>poly f 0 \<noteq> 0\<close> by (auto simp: nth_default_coeffs_eq poly_0_coeff_0)
+ next
+ case (b)
+ thus ?thesis using hx h'[of i] hdeg
+ by (auto simp: field_simps nth_default_coeffs_eq coeff_pCons nat.split poly_0_coeff_0)
+ next
+ case (c)
+ thus ?thesis using hdeg by (auto simp: nth_default_coeffs_eq coeff_eq_0)
+ qed
+ qed force
+
+ finally have 1: " - z #
+ map (\<lambda>i. coeff f (i - 1) / coeff f i - x) [1..<degree (pCons 0 f - smult x f)] @ [w] =
+ map (\<lambda>i. (if i = 0 then z / (x * coeff f 0)
+ else if i = degree (pCons 0 f - smult x f) then w / lead_coeff f
+ else inverse (coeff f i)) *
+ nth_default 0 (coeffs (f * [:- x, 1:])) i)
+ [0..<Suc (degree (pCons 0 f - smult x f))]" .
+
+ have "f * [:-x, 1:] \<noteq> 0" using hdeg by force
+
+ show "changes (coeffs (f * [:- x, 1:])) =
+ changes
+ (- z #
+ map (\<lambda>i. coeff f (i - 1) / coeff f i - x)
+ [1..<degree (pCons 0 f - smult x f)] @
+ [w])"
+ apply (subst 1)
+ apply (rule changes_scale[symmetric])
+ subgoal using hz hw hx h' hdeg by auto
+ subgoal using hdeg \<open>f * [:-x, 1:] \<noteq> 0\<close>
+ by (auto simp: length_coeffs)
+ done
+ qed
+
+ hence "changes (coeffs (f * [:- x, 1:])) =
+ changes
+ (- (max 1 (-(coeff f 0 / coeff f 1 - x))) #
+ map (\<lambda>i. coeff f (i - 1) / coeff f i - x)
+ [1..<degree (pCons 0 f - smult x f)] @
+ [max 1 (coeff f (degree f - 1) / lead_coeff f - x)])"
+ by force
+
+ also have "... = 1"
+ proof (rule changes_increasing)
+ fix i
+ assume "i < length
+ (- max 1 (- (coeff f 0 / coeff f 1 - x)) #
+ map (\<lambda>i. coeff f (i - 1) / coeff f i - x) [1..<degree (pCons 0 f - smult x f)] @
+ [max 1 (coeff f (degree f - 1) / lead_coeff f - x)]) - 1"
+ hence "i < degree (pCons 0 f - smult x f)"
+ using hdeg Suc.hyps(2) by fastforce
+ then consider (a)"i = 0" | (b)"0 \<noteq> i \<and> i < degree (pCons 0 f - smult x f) - 1" |
+ (c)"i = degree (pCons 0 f - smult x f) - 1"
+ by fastforce
+ then show "(- max 1 (- (coeff f 0 / coeff f 1 - x)) #
+ map (\<lambda>i. coeff f (i - 1) / coeff f i - x)
+ [1..<degree (pCons 0 f - smult x f)] @
+ [max 1 (coeff f (degree f - 1) / lead_coeff f - x)]) !
+ i
+ \<le> (- max 1 (- (coeff f 0 / coeff f 1 - x)) #
+ map (\<lambda>i. coeff f (i - 1) / coeff f i - x)
+ [1..<degree (pCons 0 f - smult x f)] @
+ [max 1 (coeff f (degree f - 1) / lead_coeff f - x)]) !
+ (i + 1)"
+ proof (cases)
+ case a
+ then show ?thesis by (auto simp: nth_append)
+ next
+ case b
+ have "coeff f (i - 1) * coeff f (i - 1 + 2) \<le> (coeff f (i - 1 + 1))\<^sup>2"
+ by (rule normal_poly_coeff_mult[OF hf, of "i - 1"])
+ hence "coeff f (i - 1) / coeff f i \<le> coeff f i / coeff f (i + 1)"
+ using h'[of i] h'[of "i+1"] h'[of "i-1"] h' b hdeg
+ by (auto simp: power2_eq_square divide_simps)
+ then show ?thesis
+ using b by (auto simp: nth_append)
+ next
+ case c
+ then show ?thesis using hdeg by (auto simp: nth_append not_le)
+ qed
+ qed auto
+
+ finally show "changes (coeffs (f * [:-x, 1:])) = 1" .
+ qed
+qed
+
end
\ No newline at end of file
diff --git a/thys/Three_Circles/ROOT b/thys/Three_Circles/ROOT
--- a/thys/Three_Circles/ROOT
+++ b/thys/Three_Circles/ROOT
@@ -1,12 +1,12 @@
-chapter AFP
-
-session Three_Circles (AFP) = "HOL-Computational_Algebra" +
- options [timeout = 600]
- sessions
- Budan_Fourier
- Polynomial_Interpolation
- theories
- Three_Circles
- document_files
- "root.tex"
- "root.bib"
+chapter AFP
+
+session Three_Circles (AFP) = "HOL-Computational_Algebra" +
+ options [timeout = 600]
+ sessions
+ Budan_Fourier
+ Polynomial_Interpolation
+ theories
+ Three_Circles
+ document_files
+ "root.tex"
+ "root.bib"
diff --git a/thys/Three_Circles/RRI_Misc.thy b/thys/Three_Circles/RRI_Misc.thy
--- a/thys/Three_Circles/RRI_Misc.thy
+++ b/thys/Three_Circles/RRI_Misc.thy
@@ -1,1099 +1,1099 @@
-section \<open>Misc results about polynomials\<close>
-
-theory RRI_Misc imports
- "HOL-Computational_Algebra.Computational_Algebra"
- "Budan_Fourier.BF_Misc"
- "Polynomial_Interpolation.Ring_Hom_Poly"
-begin
-
-subsection \<open>Misc\<close>
-
-declare pcompose_pCons[simp del]
-
-lemma Setcompr_subset: "\<And>f P S. {f x | x. P x} \<subseteq> S = (\<forall> x. P x \<longrightarrow> f x \<in> S)"
- by blast
-
-lemma map_cong':
- assumes "xs = map h ys" and "\<And>y. y \<in> set ys \<Longrightarrow> f (h y) = g y"
- shows "map f xs = map g ys"
- using assms map_replicate_trivial by simp
-
-lemma nth_default_replicate_eq:
- "nth_default dflt (replicate n x) i = (if i < n then x else dflt)"
- by (auto simp: nth_default_def)
-
-lemma square_bounded_less:
- fixes a b::"'a :: linordered_ring_strict"
- shows "-a < b \<and> b < a \<Longrightarrow> b*b < a*a"
- by (metis (no_types, lifting) leD leI minus_less_iff minus_mult_minus mult_strict_mono'
- neg_less_eq_nonneg neg_less_pos verit_minus_simplify(4) zero_le_mult_iff zero_le_square)
-
-lemma square_bounded_le:
- fixes a b::"'a :: linordered_ring_strict"
- shows "-a \<le> b \<and> b \<le> a \<Longrightarrow> b*b \<le> a*a"
- by (metis le_less minus_mult_minus square_bounded_less)
-
-context vector_space
-begin
-
-lemma card_le_dim_spanning:
- assumes BV: "B \<subseteq> V"
- and VB: "V \<subseteq> span B"
- and fB: "finite B"
- and dVB: "dim V \<ge> card B"
- shows "independent B"
-proof -
- {
- fix a
- assume a: "a \<in> B" "a \<in> span (B - {a})"
- from a fB have c0: "card B \<noteq> 0"
- by auto
- from a fB have cb: "card (B - {a}) = card B - 1"
- by auto
- {
- fix x
- assume x: "x \<in> V"
- from a have eq: "insert a (B - {a}) = B"
- by blast
- from x VB have x': "x \<in> span B"
- by blast
- from span_trans[OF a(2), unfolded eq, OF x']
- have "x \<in> span (B - {a})" .
- }
- then have th1: "V \<subseteq> span (B - {a})"
- by blast
- have th2: "finite (B - {a})"
- using fB by auto
- from dim_le_card[OF th1 th2]
- have c: "dim V \<le> card (B - {a})" .
- from c c0 dVB cb have False by simp
- }
- then show ?thesis
- unfolding dependent_def by blast
-qed
-
-end
-
-subsection \<open>Misc results about polynomials\<close>
-
-lemma smult_power: "smult (x^n) (p^n) = smult x p ^ n"
- apply (induction n)
- subgoal by fastforce
- by (metis smult_power)
-
-lemma reflect_poly_monom: "reflect_poly (monom n i) = monom n 0"
- apply (induction i)
- by (auto simp: coeffs_eq_iff coeffs_monom coeffs_reflect_poly)
-
-lemma poly_eq_by_eval:
- fixes P Q :: "'a::{comm_ring_1,ring_no_zero_divisors,ring_char_0} poly"
- assumes h: "\<And>x. poly P x = poly Q x" shows "P = Q"
-proof -
- have "poly P = poly Q" using h by fast
- thus ?thesis by (auto simp: poly_eq_poly_eq_iff)
-qed
-
-lemma poly_binomial:
- "[:(1::'a::comm_ring_1), 1:]^n = (\<Sum>k\<le>n. monom (of_nat (n choose k)) k)"
-proof -
- have "[:(1::'a::comm_ring_1), 1:]^n = (monom 1 1 + 1)^n"
- by (metis (no_types, lifting) add.left_neutral add.right_neutral add_pCons
- monom_altdef pCons_one power_one_right smult_1_left)
- also have "... = (\<Sum>k\<le>n. of_nat (n choose k) * monom 1 1 ^ k)"
- apply (subst binomial_ring)
- by force
- also have "... = (\<Sum>k\<le>n. monom (of_nat (n choose k)) k)"
- by (auto simp: monom_altdef of_nat_poly)
- finally show ?thesis .
-qed
-
-lemma degree_0_iff: "degree P = 0 \<longleftrightarrow> (\<exists>a. P = [:a:])"
- by (meson degree_eq_zeroE degree_pCons_0)
-
-interpretation poly_vs: vector_space smult
- by (simp add: vector_space_def smult_add_right smult_add_left)
-
-lemma degree_subspace: "poly_vs.subspace {x. degree x \<le> n}"
- by (auto simp: poly_vs.subspace_def degree_add_le)
-
-lemma monom_span:
- "poly_vs.span {monom 1 x | x. x \<le> p} = {(x::'a::field poly). degree x \<le> p}"
-(is "?L = ?R")
-proof
- show "?L \<subseteq> ?R"
- proof
- fix x assume "x \<in> ?L"
- moreover have hfin: "finite {P. \<exists>x \<in> {..p}. P = monom 1 x}"
- by auto
- ultimately have
- "x \<in> range (\<lambda>u. \<Sum>v\<in>{monom 1 x | x. x \<in> {..p}}. smult (u v) v)"
- by (simp add: poly_vs.span_finite)
- hence "\<exists> u. x = (\<Sum>v\<in>{monom 1 x | x. x \<in> {..p}}. smult (u v) v)"
- by (auto simp: image_iff)
- then obtain u
- where p': "x = (\<Sum>v\<in>{monom 1 x | x. x \<in> {..p}}. smult (u v) v)"
- by blast
- have "\<And>v. v \<in> {monom 1 x | x. x \<in> {..p}} \<Longrightarrow> degree (smult (u v) v) \<le> p"
- by (auto simp add: degree_monom_eq)
- hence "degree x \<le> p" using hfin
- apply (subst p')
- apply (rule degree_sum_le)
- by auto
- thus "x \<in> {x. degree x \<le> p}" by force
- qed
-next
- show "?R \<subseteq> ?L"
- proof
- fix x assume "x \<in> ?R"
- hence "degree x \<le> p" by force
- hence "x = (\<Sum>i\<le>p. monom (coeff x i) i)"
- by (simp add: poly_as_sum_of_monoms')
- also have
- "... = (\<Sum>i\<le>p. smult (coeff x (degree (monom (1::'a) i))) (monom 1 i))"
- by (auto simp add: smult_monom degree_monom_eq)
- also have
- "... = (\<Sum>v\<in>{monom 1 x | x. x \<in> {..p}}. smult ((\<lambda>v. coeff x (degree v)) v) v)"
- proof (rule sum.reindex_cong)
- show "inj_on degree {monom (1::'a) x |x. x \<in> {..p}}"
- proof
- fix x
- assume "x \<in> {monom (1::'a) x |x. x \<in> {..p}}"
- hence "\<exists> a. x = monom 1 a" by blast
- then obtain a where hx: "x = monom 1 a" by blast
- fix y
- assume "y \<in> {monom (1::'a) x |x. x \<in> {..p}}"
- hence "\<exists> b. y = monom 1 b" by blast
- then obtain b where hy: "y = monom 1 b" by blast
- assume "degree x = degree y"
- thus "x = y" using hx hy by (simp add: degree_monom_eq)
- qed
- show "{..p} = degree ` {monom (1::'a) x |x. x \<in> {..p}}"
- proof
- show "{..p} \<subseteq> degree ` {monom (1::'a) x |x. x \<in> {..p}}"
- proof
- fix x assume "x \<in> {..p}"
- hence "monom (1::'a) x \<in> {monom 1 x |x. x \<in> {..p}}" by force
- moreover have "degree (monom (1::'a) x) = x"
- by (simp add: degree_monom_eq)
- ultimately show "x \<in> degree ` {monom (1::'a) x |x. x \<in> {..p}}" by auto
- qed
- show "degree ` {monom (1::'a) x |x. x \<in> {..p}} \<subseteq> {..p}"
- by (auto simp add: degree_monom_eq)
- qed
- next
- fix y assume "y \<in> {monom (1::'a) x |x. x \<in> {..p}}"
- hence "\<exists>z \<in> {..p}. y = monom (1::'a) z" by blast
- then obtain z where "y = monom (1::'a) z" by blast
- thus
- "smult (coeff x (degree (monom (1::'a) (degree y)))) (monom (1::'a) (degree y)) =
- smult (coeff x (degree y)) y"
- by (simp add: smult_monom degree_monom_eq)
- qed
- finally have "x = (\<Sum>v\<in>{monom 1 x | x. x \<in> {..p}}.
- smult ((\<lambda>v. coeff x (degree v)) v) v)" .
- thus "x \<in> ?L" by (auto simp add: poly_vs.span_finite)
- qed
-qed
-
-lemma monom_independent:
- "poly_vs.independent {monom (1::'a::field) x | x. x \<le> p}"
-proof (rule poly_vs.independent_if_scalars_zero)
- fix f::"'a poly \<Rightarrow> 'a"
- assume h: "(\<Sum>x\<in>{monom 1 x |x. x \<le> p}. smult (f x) x) = 0"
- have h': "(\<Sum>i\<le>p. monom (f (monom (1::'a) i)) i) =
- (\<Sum>x\<in>{monom (1::'a) x |x. x \<le> p}. smult (f x) x)"
- proof (rule sum.reindex_cong)
- show "inj_on degree {monom (1::'a) x |x. x \<le> p}"
- by (smt (verit) degree_monom_eq inj_on_def mem_Collect_eq zero_neq_one)
- show "{..p} = degree ` {monom (1::'a) x |x. x \<le> p}"
- proof
- show "{..p} \<subseteq> degree ` {monom (1::'a) x |x. x \<le> p}"
- proof
- fix x assume "x \<in> {..p}"
- then have "x = degree (monom (1::'a) x) \<and> x \<le> p"
- by (auto simp: degree_monom_eq)
- thus "x \<in> degree ` {monom (1::'a) x |x. x \<le> p}"
- by blast
- qed
- show "degree ` {monom (1::'a) x |x. x \<le> p} \<subseteq> {..p}"
- by (force simp: degree_monom_eq)
- qed
- qed (auto simp: degree_monom_eq smult_monom)
-
- fix x::"'a poly"
- assume "x \<in> {monom 1 x |x. x \<le> p}"
- then obtain y where "y \<le> p" and "x = monom 1 y" by blast
- hence "f x = coeff (\<Sum>x\<in>{monom 1 x |x. x \<le> p}. smult (f x) x) y"
- by (auto simp: coeff_sum h'[symmetric])
- thus "f x = 0"
- using h by auto
-qed force
-
-lemma dim_degree: "poly_vs.dim {x. degree x \<le> n} = n + 1"
- using poly_vs.dim_eq_card_independent[OF monom_independent]
- by (auto simp: monom_span[symmetric] card_image image_Collect[symmetric]
- inj_on_def monom_eq_iff')
-
-lemma degree_div:
- fixes p q::"('a::idom_divide) poly"
- assumes "q dvd p"
- shows "degree (p div q) = degree p - degree q" using assms
- by (metis (no_types, lifting) add_diff_cancel_left' degree_0 degree_mult_eq
- diff_add_zero diff_zero div_by_0 dvd_div_eq_0_iff dvd_mult_div_cancel)
-
-lemma lead_coeff_div:
- fixes p q::"('a::{idom_divide, inverse}) poly"
- assumes "q dvd p"
- shows "lead_coeff (p div q) = lead_coeff p / lead_coeff q" using assms
- by (smt (z3) div_by_0 dvd_div_mult_self lead_coeff_mult leading_coeff_0_iff
- nonzero_mult_div_cancel_right)
-
-lemma complex_poly_eq:
- "r = map_poly of_real (map_poly Re r) + smult \<i> (map_poly of_real (map_poly Im r))"
- by (auto simp: poly_eq_iff coeff_map_poly complex_eq)
-
-lemma complex_poly_cong:
- "(map_poly Re p = map_poly Re q \<and> map_poly Im p = map_poly Im q) = (p = q)"
- by (metis complex_poly_eq)
-
-lemma map_poly_Im_of_real: "map_poly Im (map_poly of_real p) = 0"
- by (auto simp: poly_eq_iff coeff_map_poly)
-
-lemma mult_map_poly_imp_map_poly:
- assumes "map_poly complex_of_real q = r * map_poly complex_of_real p"
- "p \<noteq> 0"
- shows "r = map_poly complex_of_real (map_poly Re r)"
-proof -
- have h: "Im \<circ> (*) \<i> \<circ> complex_of_real = id" by fastforce
- have "map_poly complex_of_real q = r * map_poly complex_of_real p"
- using assms by blast
- also have "... = (map_poly of_real (map_poly Re r) +
- smult \<i> (map_poly of_real (map_poly Im r))) *
- map_poly complex_of_real p"
- using complex_poly_eq by fastforce
- also have "... = map_poly of_real (map_poly Re r * p) +
- smult \<i> (map_poly of_real (map_poly Im r * p))"
- by (simp add: mult_poly_add_left)
- finally have "map_poly complex_of_real q =
- map_poly of_real (map_poly Re r * p) +
- smult \<i> (map_poly of_real (map_poly Im r * p))" .
- hence "0 = map_poly Im (map_poly of_real (map_poly Re r * p) +
- smult \<i> (map_poly of_real (map_poly Im r * p)))"
- by (auto simp: complex_poly_cong[symmetric] map_poly_Im_of_real)
- also have "... = map_poly of_real (map_poly Im r * p)"
- apply (rule poly_eqI)
- by (auto simp: coeff_map_poly coeff_mult)
- finally have "0 = map_poly complex_of_real (map_poly Im r) *
- map_poly complex_of_real p"
- by auto
- hence "map_poly complex_of_real (map_poly Im r) = 0" using assms by fastforce
- thus ?thesis apply (subst complex_poly_eq) by auto
-qed
-
-lemma map_poly_dvd:
- fixes p q::"real poly"
- assumes hdvd: "map_poly complex_of_real p dvd
- map_poly complex_of_real q" "q \<noteq> 0"
- shows "p dvd q"
-proof -
- from hdvd obtain r
- where h:"map_poly complex_of_real q = r * map_poly complex_of_real p"
- by fastforce
- hence "r = map_poly complex_of_real (map_poly Re r)"
- using mult_map_poly_imp_map_poly assms by force
- hence "map_poly complex_of_real q = map_poly complex_of_real (p * map_poly Re r)"
- using h by auto
- hence "q = p * map_poly Re r" using of_real_poly_eq_iff by blast
- thus "p dvd q" by force
-qed
-
-lemma div_poly_eq_0:
- fixes p q::"('a::idom_divide) poly"
- assumes "q dvd p" "poly (p div q) x = 0" "q \<noteq> 0"
- shows "poly p x = 0"
- using assms by fastforce
-
-lemma poly_map_poly_of_real_cnj:
- "poly (map_poly of_real p) (cnj z) = cnj (poly (map_poly of_real p) z)"
- apply (induction p)
- by auto
-
-text \<open>
-An induction rule on real polynomials, if $P \neq 0$ then either $(X-x)|P$ or
-$(X-z)(X-cnj z)|P$, we induct by dividing by these polynomials.
-\<close>
-lemma real_poly_roots_induct:
- fixes P::"real poly \<Rightarrow> bool" and p::"real poly"
- assumes IH_real: "\<And>p x. P p \<Longrightarrow> P (p * [:-x, 1:])"
- and IH_complex: "\<And>p a b. b \<noteq> 0 \<Longrightarrow> P p
- \<Longrightarrow> P (p * [: a*a + b*b, -2*a, 1 :])"
- and H0: "\<And>a. P [:a:]"
- defines "d \<equiv> degree p"
- shows "P p"
- using d_def
-proof (induction d arbitrary: p rule: less_induct)
- fix p::"real poly"
- assume IH: "(\<And>q. degree q < degree p \<Longrightarrow> P q)"
- show "P p"
- proof (cases "0 = degree p")
- fix p::"real poly" assume "0 = degree p"
- hence "\<exists> a. p = [:a:]"
- by (simp add: degree_0_iff)
- thus "P p" using H0 by blast
- next
- assume hdeg: "0 \<noteq> degree p"
- hence "\<not> constant (poly (map_poly of_real p))"
- by (metis (no_types, opaque_lifting) constant_def constant_degree of_real_eq_iff of_real_poly_map_poly)
- then obtain z::complex where h: "poly (map_poly of_real p) z = 0"
- using fundamental_theorem_of_algebra by blast
- show "P p"
- proof cases
- assume "Im z = 0"
- hence "z = Re z" by (simp add: complex_is_Real_iff)
- moreover have "[:-z, 1:] dvd map_poly of_real p"
- using h poly_eq_0_iff_dvd by blast
- ultimately have "[:-(Re z), 1:] dvd p"
- by (smt (z3) dvd_iff_poly_eq_0 h of_real_0 of_real_eq_iff of_real_poly_map_poly)
- hence 2:"P (p div [:-Re z, 1:])"
- apply (subst IH)
- using hdeg by (auto simp: degree_div)
- moreover have 1:"p = (p div [:- Re z, 1:]) * [:-Re z, 1:]"
- by (metis \<open>[:- Re z, 1:] dvd p\<close> dvd_div_mult_self)
- ultimately show "P p"
- apply (subst 1)
- by (rule IH_real[OF 2])
- next
- assume "Im z \<noteq> 0"
- hence hcnj: "cnj z \<noteq> z" by (metis cnj.simps(2) neg_equal_zero)
- have h2: "poly (map_poly of_real p) (cnj z) = 0"
- using h poly_map_poly_of_real_cnj by force
- have "[:-z, 1:] * [:-cnj z, 1:] dvd map_poly of_real p"
- proof (rule divides_mult)
- have "\<And>c. c dvd [:-z, 1:] \<Longrightarrow> c dvd [:- cnj z, 1:] \<Longrightarrow> is_unit c"
- proof -
- fix c
- assume h:"c dvd [:-z, 1:]" hence "degree c \<le> 1" using divides_degree by fastforce
- hence "degree c = 0 \<or> degree c = 1" by linarith
- thus "c dvd [:- cnj z, 1:] \<Longrightarrow> is_unit c"
- proof
- assume "degree c = 0"
- moreover have "c \<noteq> 0" using h by fastforce
- ultimately show "is_unit c" by (simp add: is_unit_iff_degree)
- next
- assume hdeg: "degree c = 1"
- then obtain x where 1:"[:-z, 1:] = x*c" using h by fastforce
- hence "degree [:-z, 1:] = degree x + degree c"
- by (metis add.inverse_neutral degree_mult_eq mult_cancel_right
- mult_poly_0_left pCons_eq_0_iff zero_neq_neg_one)
- hence "degree x = 0" using hdeg by auto
- then obtain x' where 2: "x = [:x':]" using degree_0_iff by auto
- assume "c dvd [:-cnj z, 1:]"
- then obtain y where 3: "[:-cnj z, 1:] = y*c" by fastforce
- hence "degree [:-cnj z, 1:] = degree y + degree c"
- by (metis add.inverse_neutral degree_mult_eq mult_cancel_right
- mult_poly_0_left pCons_eq_0_iff zero_neq_neg_one)
- hence "degree y = 0" using hdeg by auto
- then obtain y' where 4: "y = [:y':]" using degree_0_iff by auto
- moreover from hdeg obtain a b where 5:"c = [:a, b:]" and 6: "b \<noteq> 0"
- by (meson degree_eq_oneE)
- from 1 2 5 6 have "x' = inverse b" by (auto simp: field_simps)
- moreover from 3 4 5 6 have "y' = inverse b" by (auto simp: field_simps)
- ultimately have "x = y" using 2 4 by presburger
- then have "z = cnj z" using 1 3 by (metis neg_equal_iff_equal pCons_eq_iff)
- thus "is_unit c" using hcnj by argo
- qed
- qed
- thus "coprime [:- z, 1:] [:- cnj z, 1:]" by (meson not_coprimeE)
- show "[:- z, 1:] dvd map_poly complex_of_real p"
- using h poly_eq_0_iff_dvd by auto
- show "[:- cnj z, 1:] dvd map_poly complex_of_real p"
- using h2 poly_eq_0_iff_dvd by blast
- qed
- moreover have "[:- z, 1:] * [:- cnj z, 1:] =
- map_poly of_real [:Re z*Re z + Im z*Im z, -2*Re z, 1:]"
- by (auto simp: complex_eqI)
- ultimately have hdvd:
- "map_poly complex_of_real [:Re z*Re z + Im z*Im z, -2*Re z, 1:] dvd
- map_poly complex_of_real p"
- by force
- hence "[:Re z*Re z + Im z*Im z, -2*Re z, 1:] dvd p" using map_poly_dvd
- by blast
- hence 2:"P (p div [:Re z*Re z + Im z*Im z, -2*Re z, 1:])"
- apply (subst IH)
- using hdeg by (auto simp: degree_div)
- moreover have 1:
- "p = (p div [:Re z*Re z + Im z*Im z, -2*Re z, 1:]) *
- [:Re z*Re z + Im z*Im z, -2*Re z, 1:]"
- apply (subst dvd_div_mult_self)
- using \<open>[:Re z*Re z + Im z*Im z, -2*Re z, 1:] dvd p\<close> by auto
- ultimately show "P p"
- apply (subst 1)
- apply (rule IH_complex[of "Im z" _ "Re z"])
- apply (meson \<open>Im z \<noteq> 0\<close>)
- by blast
- qed
- qed
-qed
-
-subsection \<open>The reciprocal polynomial\<close>
-
-definition reciprocal_poly :: "nat \<Rightarrow> 'a::zero poly \<Rightarrow> 'a poly"
- where "reciprocal_poly p P =
- Poly (rev ((coeffs P) @ (replicate (p - degree P) 0)))"
-
-lemma reciprocal_0: "reciprocal_poly p 0 = 0" by (simp add: reciprocal_poly_def)
-
-lemma reciprocal_1: "reciprocal_poly p 1 = monom 1 p"
- by (simp add: reciprocal_poly_def monom_altdef Poly_append)
-
-lemma coeff_reciprocal:
- assumes hi: "i \<le> p" and hP: "degree P \<le> p"
- shows "coeff (reciprocal_poly p P) i = coeff P (p - i)"
-proof cases
- assume "i < p - degree P"
- hence "degree P < p - i" using hP by linarith
- thus "coeff (reciprocal_poly p P) i = coeff P (p - i)"
- by (auto simp: reciprocal_poly_def nth_default_append coeff_eq_0)
-next
- assume h: "\<not>i < p - degree P"
- show "coeff (reciprocal_poly p P) i = coeff P (p - i)"
- proof cases
- assume "P = 0"
- thus "coeff (reciprocal_poly p P) i = coeff P (p - i)"
- by (simp add: reciprocal_0)
- next
- assume hP': "P \<noteq> 0"
- have "degree P \<ge> p - i" using h hP by linarith
- moreover hence "(i - (p - degree P)) < length (rev (coeffs P))"
- using hP' hP hi by (auto simp: length_coeffs)
- thus "coeff (reciprocal_poly p P) i = coeff P (p - i)"
- by (auto simp: reciprocal_poly_def nth_default_append coeff_eq_0 hP hP'
- nth_default_nth rev_nth calculation coeffs_nth length_coeffs_degree)
- qed
-qed
-
-lemma coeff_reciprocal_less:
- assumes hn: "p < i" and hP: "degree P \<le> p"
- shows "coeff (reciprocal_poly p P) i = 0"
-proof cases
- assume "P = 0"
- thus ?thesis by (auto simp: reciprocal_0)
-next
- assume "P \<noteq> 0"
- thus ?thesis
- using hn
- by (auto simp: reciprocal_poly_def nth_default_append
- nth_default_eq_dflt_iff hP length_coeffs)
-qed
-
-lemma reciprocal_monom:
- assumes "n \<le> p"
- shows "reciprocal_poly p (monom a n) = monom a (p-n)"
-proof (cases "a=0")
- case True
- then show ?thesis by (simp add: reciprocal_0)
-next
- case False
- with \<open>n\<le>p\<close> show ?thesis
- apply (rule_tac poly_eqI)
- by (metis coeff_monom coeff_reciprocal coeff_reciprocal_less
- diff_diff_cancel diff_le_self lead_coeff_monom not_le_imp_less)
-qed
-
-lemma reciprocal_degree: "reciprocal_poly (degree P) P = reflect_poly P"
- by (auto simp add: reciprocal_poly_def reflect_poly_def)
-
-lemma degree_reciprocal:
- fixes P :: "('a::zero) poly"
- assumes hP: "degree P \<le> p"
- shows "degree (reciprocal_poly p P) \<le> p"
-proof (auto simp add: reciprocal_poly_def)
- have "degree (reciprocal_poly p P) \<le>
- length (replicate (p - degree P) (0::'a) @ rev (coeffs P))"
- by (metis degree_Poly reciprocal_poly_def rev_append rev_replicate)
- thus "degree (Poly (replicate (p - degree P) 0 @ rev (coeffs P))) \<le> p"
- by (smt Suc_le_mono add_Suc_right coeffs_Poly degree_0 hP le_SucE le_SucI
- le_add_diff_inverse2 le_zero_eq length_append length_coeffs_degree
- length_replicate length_rev length_strip_while_le reciprocal_0
- reciprocal_poly_def rev_append rev_replicate)
-qed
-
-lemma reciprocal_0_iff:
- assumes hP: "degree P \<le> p"
- shows "(reciprocal_poly p P = 0) = (P = 0)"
-proof
- assume h: "reciprocal_poly p P = 0"
- show "P = 0"
- proof (rule poly_eqI)
- fix n
- show "coeff P n = coeff 0 n"
- proof cases
- assume hn: "n \<le> p"
- hence "p - n \<le> p" by auto
- hence "coeff (reciprocal_poly p P) (p - n) = coeff P n"
- using hP hn by (auto simp: coeff_reciprocal)
- thus ?thesis using h by auto
- next
- assume hn: "\<not> n \<le> p"
- thus ?thesis using hP by (metis coeff_0 dual_order.trans le_degree)
- qed
- qed
-next
- assume "P = 0"
- thus "reciprocal_poly p P = 0" using reciprocal_0 by fast
-qed
-
-lemma poly_reciprocal:
- fixes P::"'a::field poly"
- assumes hp: "degree P \<le> p" and hx: "x \<noteq> 0"
- shows "poly (reciprocal_poly p P) x = x^p * (poly P (inverse x))"
-proof -
- have "poly (reciprocal_poly p P) x
- = poly ((Poly ((replicate (p - degree P) 0) @ rev (coeffs P)))) x"
- by (auto simp add: hx reflect_poly_def reciprocal_poly_def)
- also have "... = poly ((monom 1 (p - degree P)) * (reflect_poly P)) x"
- by (auto simp add: reflect_poly_def Poly_append)
- also have "... = x^(p - degree P) * x ^ degree P * poly P (inverse x)"
- by (auto simp add: poly_reflect_poly_nz poly_monom hx)
- also have "... = x^p * poly P (inverse x)"
- by (auto simp add: hp power_add[symmetric])
- finally show ?thesis .
-qed
-
-lemma reciprocal_fcompose:
- fixes P::"('a::{ring_char_0,field}) poly"
- assumes hP: "degree P \<le> p"
- shows "reciprocal_poly p P = monom 1 (p - degree P) * fcompose P 1 [:0, 1:]"
-proof (rule poly_eq_by_eval, cases)
- fix x::'a
- assume hx: "x = 0"
- hence "poly (reciprocal_poly p P) x = coeff P p"
- using hP by (auto simp: poly_0_coeff_0 coeff_reciprocal)
- moreover have "poly (monom 1 (p - degree P)
- * fcompose P 1 [:0, 1:]) x = coeff P p"
- proof cases
- assume "degree P = p"
- thus ?thesis
- apply (induction P arbitrary: p)
- using hx by (auto simp: poly_monom degree_0_iff fcompose_pCons)
- next
- assume "degree P \<noteq> p"
- hence "degree P < p" using hP by auto
- thus ?thesis
- using hx by (auto simp: poly_monom coeff_eq_0)
- qed
- ultimately show "poly (reciprocal_poly p P) x = poly (monom 1 (p - degree P) * fcompose P 1 [:0, 1:]) x"
- by presburger
-next
- fix x::'a assume "x \<noteq> 0"
- thus "poly (reciprocal_poly p P) x =
- poly (monom 1 (p - degree P) * fcompose P 1 [:0, 1:]) x"
- using hP
- by (auto simp: poly_reciprocal poly_fcompose inverse_eq_divide
- poly_monom power_diff)
-qed
-
-lemma reciprocal_reciprocal:
- fixes P :: "'a::{field,ring_char_0} poly"
- assumes hP: "degree P \<le> p"
- shows "reciprocal_poly p (reciprocal_poly p P) = P"
-proof (rule poly_eq_by_eval)
- fix x
- show "poly (reciprocal_poly p (reciprocal_poly p P)) x = poly P x"
- proof cases
- assume "x = 0"
- thus "poly (reciprocal_poly p (reciprocal_poly p P)) x = poly P x"
- using hP
- by (auto simp: poly_0_coeff_0 coeff_reciprocal degree_reciprocal)
- next
- assume hx: "x \<noteq> 0"
- hence "poly (reciprocal_poly p (reciprocal_poly p P)) x
- = x ^ p * (inverse x ^ p * poly P x)" using hP
- by (auto simp: poly_reciprocal degree_reciprocal)
- thus "poly (reciprocal_poly p (reciprocal_poly p P)) x = poly P x"
- using hP hx left_right_inverse_power right_inverse by auto
- qed
-qed
-
-lemma reciprocal_smult:
- fixes P :: "'a::idom poly"
- assumes h: "degree P \<le> p"
- shows "reciprocal_poly p (smult n P) = smult n (reciprocal_poly p P)"
-proof cases
- assume "n = 0"
- thus ?thesis by (auto simp add: reciprocal_poly_def)
-next
- assume "n \<noteq> 0"
- thus ?thesis
- by (auto simp add: reciprocal_poly_def smult_Poly coeffs_smult
- rev_map[symmetric])
-qed
-
-lemma reciprocal_add:
- fixes P Q :: "'a::comm_semiring_0 poly"
- assumes "degree P \<le> p" and "degree Q \<le> p"
- shows "reciprocal_poly p (P + Q) = reciprocal_poly p P + reciprocal_poly p Q"
-(is "?L = ?R")
-proof (rule poly_eqI, cases)
- fix n
- assume "n \<le> p"
- then show "coeff ?L n = coeff ?R n"
- using assms by (auto simp: degree_add_le coeff_reciprocal)
-next
- fix n assume "\<not>n \<le> p"
- then show "coeff ?L n = coeff ?R n"
- using assms by (auto simp: degree_add_le coeff_reciprocal_less)
-qed
-
-lemma reciprocal_diff:
- fixes P Q :: "'a::comm_ring poly"
- assumes "degree P \<le> p" and "degree Q \<le> p"
- shows "reciprocal_poly p (P - Q) = reciprocal_poly p P - reciprocal_poly p Q"
- by (metis (no_types, lifting) ab_group_add_class.ab_diff_conv_add_uminus assms
- add_diff_cancel degree_add_le degree_minus diff_add_cancel reciprocal_add)
-
-lemma reciprocal_sum:
- fixes P :: "'a \<Rightarrow> 'b::comm_semiring_0 poly"
- assumes hP: "\<And>k. degree (P k) \<le> p"
- shows "reciprocal_poly p (\<Sum>k\<in>A. P k) = (\<Sum>k\<in>A. reciprocal_poly p (P k))"
-proof (induct A rule: infinite_finite_induct)
- case (infinite A)
- then show ?case by (simp add: reciprocal_0)
-next
- case empty
- then show ?case by (simp add: reciprocal_0)
-next
- case (insert x F)
- assume "x \<notin> F"
- and "reciprocal_poly p (sum P F) = (\<Sum>k\<in>F. reciprocal_poly p (P k))"
- and "finite F"
- moreover hence "reciprocal_poly p (sum P (insert x F))
- = reciprocal_poly p (P x) + reciprocal_poly p (sum P F)"
- by (auto simp add: reciprocal_add hP degree_sum_le)
- ultimately show "reciprocal_poly p (sum P (insert x F))
- = (\<Sum>k\<in>insert x F. reciprocal_poly p (P k))"
- by (auto simp: Groups_Big.comm_monoid_add_class.sum.insert_if)
-qed
-
-lemma reciprocal_mult:
- fixes P Q::"'a::{ring_char_0,field} poly"
- assumes "degree (P * Q) \<le> p"
- and "degree P \<le> p" and "degree Q \<le> p"
- shows "monom 1 p * reciprocal_poly p (P * Q) =
- reciprocal_poly p P * reciprocal_poly p Q"
-proof (cases "P=0 \<or> Q=0")
- case True
- then show ?thesis using assms(1)
- by (auto simp: reciprocal_fcompose fcompose_mult)
-next
- case False
- then show ?thesis
- using assms
- by (auto simp: degree_mult_eq mult_monom reciprocal_fcompose fcompose_mult)
-qed
-
-lemma reciprocal_reflect_poly:
- fixes P::"'a::{ring_char_0,field} poly"
- assumes hP: "degree P \<le> p"
- shows "reciprocal_poly p P = monom 1 (p - degree P) * reflect_poly P"
-proof (rule poly_eqI)
- fix n
- show "coeff (reciprocal_poly p P) n =
- coeff (monom 1 (p - degree P) * reflect_poly P) n"
- proof cases
- assume "n \<le> p"
- thus ?thesis using hP
- by (auto simp: coeff_reciprocal coeff_monom_mult coeff_reflect_poly coeff_eq_0)
- next
- assume "\<not> n \<le> p"
- thus ?thesis using hP
- by (auto simp: coeff_reciprocal_less coeff_monom_mult coeff_reflect_poly)
- qed
-qed
-
-lemma map_poly_reciprocal:
- assumes "degree P \<le> p" and "f 0 = 0"
- shows "map_poly f (reciprocal_poly p P) = reciprocal_poly p (map_poly f P)"
-proof (rule poly_eqI)
- fix n
- show "coeff (map_poly f (reciprocal_poly p P)) n =
- coeff (reciprocal_poly p (map_poly f P)) n"
- proof (cases "n\<le>p")
- case True
- then show ?thesis
- apply (subst coeff_reciprocal[OF True])
- subgoal by (meson assms(1) assms(2) degree_map_poly_le le_trans)
- by (simp add: assms(1) assms(2) coeff_map_poly coeff_reciprocal)
- next
- case False
- then show ?thesis
- by (metis assms(1) assms(2) coeff_map_poly coeff_reciprocal_less
- degree_map_poly_le dual_order.trans leI)
- qed
-qed
-
-subsection \<open>More about @{term proots_count}\<close>
-
-lemma proots_count_monom:
- assumes "0 \<notin> A"
- shows "proots_count (monom 1 d) A = 0"
- using assms by (auto simp: proots_count_def poly_monom)
-
-lemma proots_count_reciprocal:
- fixes P::"'a::{ring_char_0,field} poly"
- assumes hP: "degree P \<le> p" and h0: "P \<noteq> 0" and h0': "0 \<notin> A"
- shows "proots_count (reciprocal_poly p P) A = proots_count P {x. inverse x \<in> A}"
-proof -
- have "proots_count (reciprocal_poly p P) A =
- proots_count (fcompose P 1 [:0, 1:]) A"
- apply (subst reciprocal_fcompose[OF hP], subst proots_count_times)
- subgoal using h0 by (metis hP reciprocal_0_iff reciprocal_fcompose)
- subgoal using h0' by (auto simp: proots_count_monom)
- done
-
- also have "... = proots_count P {x. inverse x \<in> A}"
- proof (rule proots_fcompose_bij_eq[symmetric])
- show "bij_betw (\<lambda>x. poly 1 x / poly [:0, 1:] x) A {x. inverse x \<in> A}"
- proof (rule bij_betw_imageI)
- show "inj_on (\<lambda>x. poly 1 x / poly [:0, 1:] x) A"
- by (simp add: inj_on_def)
-
- show "(\<lambda>x. poly 1 x / poly [:0, 1:] x) ` A = {x. inverse x \<in> A}"
- proof
- show "(\<lambda>x. poly 1 x / poly [:0, 1:] x) ` A \<subseteq> {x. inverse x \<in> A}"
- by force
- show "{x. inverse x \<in> A} \<subseteq> (\<lambda>x. poly 1 x / poly [:0, 1:] x) ` A"
- proof
- fix x assume "x \<in> {x::'a. inverse x \<in> A}"
- hence "x = poly 1 (inverse x) / poly [:0, 1:] (inverse x) \<and> inverse x \<in> A"
- by (auto simp: inverse_eq_divide)
- thus "x \<in> (\<lambda>x. poly 1 x / poly [:0, 1:] x) ` A" by blast
- qed
- qed
- qed
- next
- show "\<forall>c. 1 \<noteq> smult c [:0, 1:]"
- by (metis coeff_pCons_0 degree_1 lead_coeff_1 pCons_0_0 pcompose_0'
- pcompose_smult smult_0_right zero_neq_one)
- qed (auto simp: assms infinite_UNIV_char_0)
- finally show ?thesis by linarith
-qed
-
-lemma proots_count_reciprocal':
- fixes P::"real poly"
- assumes hP: "degree P \<le> p" and h0: "P \<noteq> 0"
- shows "proots_count P {x. 0 < x \<and> x < 1} =
- proots_count (reciprocal_poly p P) {x. 1 < x}"
-proof (subst proots_count_reciprocal)
- show "proots_count P {x. 0 < x \<and> x < 1} =
- proots_count P {x. inverse x \<in> Collect ((<) 1)}"
- apply (rule arg_cong[of _ _ "proots_count P"])
- using one_less_inverse_iff by fastforce
-qed (use assms in auto)
-
-lemma proots_count_pos:
- assumes "proots_count P S > 0"
- shows "\<exists>x \<in> S. poly P x = 0"
-proof (rule ccontr)
- assume "\<not> (\<exists>x\<in>S. poly P x = 0)"
- hence "\<And>x. x \<in> S \<Longrightarrow> poly P x \<noteq> 0" by blast
- hence "\<And>x. x \<in> S \<Longrightarrow> order x P = 0" using order_0I by blast
- hence "proots_count P S = 0" by (force simp: proots_count_def)
- thus False using assms by presburger
-qed
-
-lemma proots_count_of_root_set:
- assumes "P \<noteq> 0" "R \<subseteq> S" and "\<And>x. x\<in>R \<Longrightarrow> poly P x = 0"
- shows "proots_count P S \<ge> card R"
-proof -
- have "card R \<le> card (proots_within P S)"
- apply (rule card_mono)
- subgoal using assms by auto
- subgoal using assms(2) assms(3) by (auto simp: proots_within_def)
- done
- also have "... \<le> proots_count P S"
- by (rule card_proots_within_leq[OF assms(1)])
- finally show ?thesis .
-qed
-
-lemma proots_count_of_root: assumes "P \<noteq> 0" "x\<in>S" "poly P x = 0"
- shows "proots_count P S > 0"
- using proots_count_of_root_set[of P "{x}" S] assms by force
-
-subsection \<open>More about @{term changes}\<close>
-
-lemma changes_nonneg: "0 \<le> changes xs"
- apply (induction xs rule: changes.induct)
- by simp_all
-
-lemma changes_replicate_0: shows "changes (replicate n 0) = 0"
- apply (induction n)
- by auto
-
-lemma changes_append_replicate_0: "changes (xs @ replicate n 0) = changes xs"
-proof (induction xs rule: changes.induct)
- case (2 uu)
- then show ?case
- apply (induction n)
- by auto
-qed (auto simp: changes_replicate_0)
-
-lemma changes_scale_Cons:
- fixes xs:: "real list" assumes hs: "s > 0"
- shows "changes (s * x # xs) = changes (x # xs)"
- apply (induction xs rule: changes.induct)
- using assms by (auto simp: mult_less_0_iff zero_less_mult_iff)
-
-lemma changes_scale:
- fixes xs::"('a::linordered_idom) list"
- assumes hs: "\<And>i. i < n \<Longrightarrow> s i > 0" and hn: "length xs \<le> n"
- shows "changes [s i * (nth_default 0 xs i). i \<leftarrow> [0..<n]] = changes xs"
-using assms
-proof (induction xs arbitrary: s n rule: changes.induct)
- case 1
- show ?case by (auto simp: map_replicate_const changes_replicate_0)
-next
- case (2 uu)
- show ?case
- proof (cases n)
- case 0
- then show ?thesis by force
- next
- case (Suc m)
- hence "map (\<lambda>i. s i * nth_default 0 [uu] i) [0..<n] = [s 0 * uu] @ replicate m 0"
- proof (induction m arbitrary: n)
- case (Suc m n)
- from Suc have "map (\<lambda>i. s i * nth_default 0 [uu] i) [0..<Suc m] =
- [s 0 * uu] @ replicate m 0"
- by meson
- hence "map (\<lambda>i. s i * nth_default 0 [uu] i) [0..<n] =
- [s 0 * uu] @ replicate m 0 @ [0]"
- using Suc by auto
- also have "... = [s 0 * uu] @ replicate (Suc m) 0"
- by (simp add: replicate_append_same)
- finally show ?case .
- qed fastforce
- then show ?thesis
- by (metis changes.simps(2) changes_append_replicate_0)
- qed
-next
- case (3 a b xs s n)
- obtain m where hn: "n = m + 2"
- using 3(5)
- by (metis add_2_eq_Suc' diff_diff_cancel diff_le_self length_Suc_conv
- nat_arith.suc1 ordered_cancel_comm_monoid_diff_class.add_diff_inverse)
- hence h:
- "map (\<lambda>i. s i * nth_default 0 (a # b # xs) i) [0..<n] =
- s 0 * a # s 1 * b # map (\<lambda>i.
- (\<lambda> i. s (i+2)) i * nth_default 0 (xs) i) [0..<m]"
- apply (induction m arbitrary: n)
- by auto
- consider (neg)"a*b<0" | (nil)"b=0" | (pos)"\<not>a*b<0 \<and> \<not>b=0" by linarith
- then show ?case
- proof (cases)
- case neg
- hence
- "changes (map (\<lambda>i. s i * nth_default 0 (a # b # xs) i) [0..<n]) =
- 1 + changes (s 1 * b # map (\<lambda>i. (\<lambda> i. s (i+2)) i
- * nth_default 0 (xs) i) [0..<m])"
- apply (subst h)
- using 3(4)[of 0] 3(4)[of 1] hn
- by (metis (no_types, lifting) changes.simps(3) mult_less_0_iff pos2
- mult_pos_pos one_less_numeral_iff semiring_norm(76) trans_less_add2)
- also have
- "... = 1 + changes (map (\<lambda>i. s (Suc i) * nth_default 0 (b # xs) i) [0..<Suc m])"
- apply (rule arg_cong[of _ _ "\<lambda> x. 1 + changes x"])
- apply (induction m)
- by auto
- also have "... = changes (a # b # xs)"
- apply (subst 3(1)[OF neg])
- using 3 neg hn by auto
- finally show ?thesis .
- next
- case nil
- hence "changes (map (\<lambda>i. s i * nth_default 0 (a # b # xs) i) [0..<n]) =
- changes (s 0 * a # map (\<lambda>i. (\<lambda> i. s (i+2)) i * nth_default 0 (xs) i) [0..<m])"
- apply (subst h)
- using 3(4)[of 0] 3(4)[of 1] hn
- by auto
- also have
- "... = changes (map
- (\<lambda>i. s (if i = 0 then 0 else Suc i) * nth_default 0 (a # xs) i)
- [0..<Suc m])"
- apply (rule arg_cong[of _ _ "\<lambda> x. changes x"])
- apply (induction m)
- by auto
- also have "... = changes (a # b # xs)"
- apply (subst 3(2))
- using 3 nil hn by auto
- finally show ?thesis .
- next
- case pos
- hence "changes (map (\<lambda>i. s i * nth_default 0 (a # b # xs) i) [0..<n]) =
- changes (s 1 * b # map (\<lambda>i. (\<lambda> i. s (i+2)) i * nth_default 0 (xs) i) [0..<m])"
- apply (subst h)
- using 3(4)[of 0] 3(4)[of 1] hn
- by (metis (no_types, lifting) changes.simps(3) divisors_zero
- mult_less_0_iff nat_1_add_1 not_square_less_zero one_less_numeral_iff
- semiring_norm(76) trans_less_add2 zero_less_mult_pos zero_less_two)
- also have
- "... = changes (map (\<lambda>i. s (Suc i) * nth_default 0 (b # xs) i) [0..<Suc m])"
- apply (rule arg_cong[of _ _ "\<lambda> x. changes x"])
- apply (induction m)
- by auto
- also have "... = changes (a # b # xs)"
- apply (subst 3(3))
- using 3 pos hn by auto
- finally show ?thesis .
- qed
-qed
-
-lemma changes_scale_const: fixes xs::"'a::linordered_idom list"
- assumes hs: "s \<noteq> 0"
- shows "changes (map ((*) s) xs) = changes xs"
- apply (induction xs rule: changes.induct)
- apply (simp, force)
- using hs by (auto simp: mult_less_0_iff zero_less_mult_iff)
-
-lemma changes_snoc: fixes xs::"'a::linordered_idom list"
- shows "changes (xs @ [b, a]) = (if a * b < 0 then 1 + changes (xs @ [b])
- else if b = 0 then changes (xs @ [a]) else changes (xs @ [b]))"
- apply (induction xs rule: changes.induct)
- subgoal by (force simp: mult_less_0_iff)
- subgoal by (force simp: mult_less_0_iff)
- subgoal by force
- done
-
-lemma changes_rev: fixes xs:: "'a::linordered_idom list"
- shows "changes (rev xs) = changes xs"
- apply (induction xs rule: changes.induct)
- by (auto simp: changes_snoc)
-
-lemma changes_rev_about: fixes xs:: "'a::linordered_idom list"
- shows "changes (replicate (p - length xs) 0 @ rev xs) = changes xs"
-proof (induction p)
- case (Suc p)
- then show ?case
- proof cases
- assume "\<not>Suc p \<le> length xs"
- hence "Suc p - length xs = Suc (p - length xs)" by linarith
- thus ?case using Suc.IH changes_rev by auto
- qed (auto simp: changes_rev)
-qed (auto simp: changes_rev)
-
-lemma changes_add_between:
- assumes "a \<le> x" and "x \<le> b"
- shows "changes (as @ [a, b] @ bs) = changes (as @ [a, x, b] @ bs)"
-proof (induction as rule: changes.induct)
- case 1
- then show ?case using assms
- apply (induction bs)
- by (auto simp: mult_less_0_iff)
-next
- case (2 c)
- then show ?case
- apply (induction bs)
- using assms by (auto simp: mult_less_0_iff)
-next
- case (3 y z as)
- then show ?case
- using assms by (auto simp: mult_less_0_iff)
-qed
-
-lemma changes_all_nonneg: assumes "\<And>i. nth_default 0 xs i \<ge> 0" shows "changes xs = 0"
- using assms
-proof (induction xs rule: changes.induct)
- case (3 x1 x2 xs)
- moreover assume "(\<And>i. 0 \<le> nth_default 0 (x1 # x2 # xs) i)"
- moreover hence "(\<And>i. 0 \<le> nth_default 0 (x1 # xs) i)"
- and "(\<And>i. 0 \<le> nth_default 0 (x2 # xs) i)"
- and "x1 * x2 \<ge> 0"
- proof -
- fix i
- assume h:"(\<And>i. 0 \<le> nth_default 0 (x1 # x2 # xs) i)"
- show "0 \<le> nth_default 0 (x1 # xs) i"
- proof (cases i)
- case 0
- then show ?thesis using h[of 0] by force
- next
- case (Suc nat)
- then show ?thesis using h[of "Suc (Suc nat)"] by force
- qed
- show "0 \<le> nth_default 0 (x2 # xs) i" using h[of "Suc i"] by simp
- show "x1*x2 \<ge> 0" using h[of 0] h[of 1] by simp
- qed
- ultimately show ?case by auto
-qed auto
-
-lemma changes_pCons: "changes (coeffs (pCons 0 f)) = changes (coeffs f)"
- by (auto simp: cCons_def)
-
-lemma changes_increasing:
- assumes "\<And>i. i < length xs - 1 \<Longrightarrow> xs ! (i + 1) \<ge> xs ! i"
- and "length xs > 1"
- and "hd xs < 0"
- and "last xs > 0"
- shows "changes xs = 1"
- using assms
-proof (induction xs rule:changes.induct)
- case (3 x y xs)
- consider (neg)"x*y<0" | (nil)"y=0" | (pos)"\<not>x*y<0 \<and> \<not>y=0" by linarith
- then show ?case
- proof cases
- case neg
- have "changes (y # xs) = 0"
- proof (rule changes_all_nonneg)
- fix i
- show "0 \<le> nth_default 0 (y # xs) i"
- proof (cases "i < length (y # xs)")
- case True
- then show ?thesis using 3(4)[of i]
- apply (induction i)
- subgoal using 3(6) neg by (fastforce simp: mult_less_0_iff)
- subgoal using 3(4) by (auto simp: nth_default_def)
- done
- next
- case False
- then show ?thesis by (simp add: nth_default_def)
- qed
- qed
- thus "changes (x # y # xs) = 1"
- using neg by force
- next
- case nil
- hence "xs \<noteq> []" using 3(7) by force
- have h: "\<And>i. i < length (x # xs) - 1 \<Longrightarrow> (x # xs) ! i \<le> (x # xs) ! (i + 1)"
- proof -
- fix i assume "i < length (x # xs) - 1"
- thus "(x # xs) ! i \<le> (x # xs) ! (i + 1)"
- apply (cases "i = 0")
- subgoal using 3(4)[of 0] 3(4)[of 1] \<open>xs \<noteq> []\<close> by force
- using 3(4)[of "i+1"] by simp
- qed
- have "changes (x # xs) = 1"
- apply (rule 3(2))
- using nil h \<open>xs \<noteq> []\<close> 3(6) 3(7) by auto
- thus ?thesis
- using nil by force
- next
- case pos
- hence "xs \<noteq> []" using 3(6) 3(7) by (fastforce simp: mult_less_0_iff)
- have "changes (y # xs) = 1"
- proof (rule 3(3))
- show "\<not> x * y < 0" "y \<noteq> 0"
- using pos by auto
- show "\<And>i. i < length (y # xs) - 1
- \<Longrightarrow> (y # xs) ! i \<le> (y # xs) ! (i + 1)"
- using 3(4) by force
- show "1 < length (y # xs)"
- using \<open>xs \<noteq> []\<close> by force
- show "hd (y # xs) < 0"
- using 3(6) pos by (force simp: mult_less_0_iff)
- show "0 < last (y # xs)"
- using 3(7) by force
- qed
- thus ?thesis using pos by auto
- qed
-qed auto
-
+section \<open>Misc results about polynomials\<close>
+
+theory RRI_Misc imports
+ "HOL-Computational_Algebra.Computational_Algebra"
+ "Budan_Fourier.BF_Misc"
+ "Polynomial_Interpolation.Ring_Hom_Poly"
+begin
+
+subsection \<open>Misc\<close>
+
+declare pcompose_pCons[simp del]
+
+lemma Setcompr_subset: "\<And>f P S. {f x | x. P x} \<subseteq> S = (\<forall> x. P x \<longrightarrow> f x \<in> S)"
+ by blast
+
+lemma map_cong':
+ assumes "xs = map h ys" and "\<And>y. y \<in> set ys \<Longrightarrow> f (h y) = g y"
+ shows "map f xs = map g ys"
+ using assms map_replicate_trivial by simp
+
+lemma nth_default_replicate_eq:
+ "nth_default dflt (replicate n x) i = (if i < n then x else dflt)"
+ by (auto simp: nth_default_def)
+
+lemma square_bounded_less:
+ fixes a b::"'a :: linordered_ring_strict"
+ shows "-a < b \<and> b < a \<Longrightarrow> b*b < a*a"
+ by (metis (no_types, lifting) leD leI minus_less_iff minus_mult_minus mult_strict_mono'
+ neg_less_eq_nonneg neg_less_pos verit_minus_simplify(4) zero_le_mult_iff zero_le_square)
+
+lemma square_bounded_le:
+ fixes a b::"'a :: linordered_ring_strict"
+ shows "-a \<le> b \<and> b \<le> a \<Longrightarrow> b*b \<le> a*a"
+ by (metis le_less minus_mult_minus square_bounded_less)
+
+context vector_space
+begin
+
+lemma card_le_dim_spanning:
+ assumes BV: "B \<subseteq> V"
+ and VB: "V \<subseteq> span B"
+ and fB: "finite B"
+ and dVB: "dim V \<ge> card B"
+ shows "independent B"
+proof -
+ {
+ fix a
+ assume a: "a \<in> B" "a \<in> span (B - {a})"
+ from a fB have c0: "card B \<noteq> 0"
+ by auto
+ from a fB have cb: "card (B - {a}) = card B - 1"
+ by auto
+ {
+ fix x
+ assume x: "x \<in> V"
+ from a have eq: "insert a (B - {a}) = B"
+ by blast
+ from x VB have x': "x \<in> span B"
+ by blast
+ from span_trans[OF a(2), unfolded eq, OF x']
+ have "x \<in> span (B - {a})" .
+ }
+ then have th1: "V \<subseteq> span (B - {a})"
+ by blast
+ have th2: "finite (B - {a})"
+ using fB by auto
+ from dim_le_card[OF th1 th2]
+ have c: "dim V \<le> card (B - {a})" .
+ from c c0 dVB cb have False by simp
+ }
+ then show ?thesis
+ unfolding dependent_def by blast
+qed
+
+end
+
+subsection \<open>Misc results about polynomials\<close>
+
+lemma smult_power: "smult (x^n) (p^n) = smult x p ^ n"
+ apply (induction n)
+ subgoal by fastforce
+ by (metis smult_power)
+
+lemma reflect_poly_monom: "reflect_poly (monom n i) = monom n 0"
+ apply (induction i)
+ by (auto simp: coeffs_eq_iff coeffs_monom coeffs_reflect_poly)
+
+lemma poly_eq_by_eval:
+ fixes P Q :: "'a::{comm_ring_1,ring_no_zero_divisors,ring_char_0} poly"
+ assumes h: "\<And>x. poly P x = poly Q x" shows "P = Q"
+proof -
+ have "poly P = poly Q" using h by fast
+ thus ?thesis by (auto simp: poly_eq_poly_eq_iff)
+qed
+
+lemma poly_binomial:
+ "[:(1::'a::comm_ring_1), 1:]^n = (\<Sum>k\<le>n. monom (of_nat (n choose k)) k)"
+proof -
+ have "[:(1::'a::comm_ring_1), 1:]^n = (monom 1 1 + 1)^n"
+ by (metis (no_types, lifting) add.left_neutral add.right_neutral add_pCons
+ monom_altdef pCons_one power_one_right smult_1_left)
+ also have "... = (\<Sum>k\<le>n. of_nat (n choose k) * monom 1 1 ^ k)"
+ apply (subst binomial_ring)
+ by force
+ also have "... = (\<Sum>k\<le>n. monom (of_nat (n choose k)) k)"
+ by (auto simp: monom_altdef of_nat_poly)
+ finally show ?thesis .
+qed
+
+lemma degree_0_iff: "degree P = 0 \<longleftrightarrow> (\<exists>a. P = [:a:])"
+ by (meson degree_eq_zeroE degree_pCons_0)
+
+interpretation poly_vs: vector_space smult
+ by (simp add: vector_space_def smult_add_right smult_add_left)
+
+lemma degree_subspace: "poly_vs.subspace {x. degree x \<le> n}"
+ by (auto simp: poly_vs.subspace_def degree_add_le)
+
+lemma monom_span:
+ "poly_vs.span {monom 1 x | x. x \<le> p} = {(x::'a::field poly). degree x \<le> p}"
+(is "?L = ?R")
+proof
+ show "?L \<subseteq> ?R"
+ proof
+ fix x assume "x \<in> ?L"
+ moreover have hfin: "finite {P. \<exists>x \<in> {..p}. P = monom 1 x}"
+ by auto
+ ultimately have
+ "x \<in> range (\<lambda>u. \<Sum>v\<in>{monom 1 x | x. x \<in> {..p}}. smult (u v) v)"
+ by (simp add: poly_vs.span_finite)
+ hence "\<exists> u. x = (\<Sum>v\<in>{monom 1 x | x. x \<in> {..p}}. smult (u v) v)"
+ by (auto simp: image_iff)
+ then obtain u
+ where p': "x = (\<Sum>v\<in>{monom 1 x | x. x \<in> {..p}}. smult (u v) v)"
+ by blast
+ have "\<And>v. v \<in> {monom 1 x | x. x \<in> {..p}} \<Longrightarrow> degree (smult (u v) v) \<le> p"
+ by (auto simp add: degree_monom_eq)
+ hence "degree x \<le> p" using hfin
+ apply (subst p')
+ apply (rule degree_sum_le)
+ by auto
+ thus "x \<in> {x. degree x \<le> p}" by force
+ qed
+next
+ show "?R \<subseteq> ?L"
+ proof
+ fix x assume "x \<in> ?R"
+ hence "degree x \<le> p" by force
+ hence "x = (\<Sum>i\<le>p. monom (coeff x i) i)"
+ by (simp add: poly_as_sum_of_monoms')
+ also have
+ "... = (\<Sum>i\<le>p. smult (coeff x (degree (monom (1::'a) i))) (monom 1 i))"
+ by (auto simp add: smult_monom degree_monom_eq)
+ also have
+ "... = (\<Sum>v\<in>{monom 1 x | x. x \<in> {..p}}. smult ((\<lambda>v. coeff x (degree v)) v) v)"
+ proof (rule sum.reindex_cong)
+ show "inj_on degree {monom (1::'a) x |x. x \<in> {..p}}"
+ proof
+ fix x
+ assume "x \<in> {monom (1::'a) x |x. x \<in> {..p}}"
+ hence "\<exists> a. x = monom 1 a" by blast
+ then obtain a where hx: "x = monom 1 a" by blast
+ fix y
+ assume "y \<in> {monom (1::'a) x |x. x \<in> {..p}}"
+ hence "\<exists> b. y = monom 1 b" by blast
+ then obtain b where hy: "y = monom 1 b" by blast
+ assume "degree x = degree y"
+ thus "x = y" using hx hy by (simp add: degree_monom_eq)
+ qed
+ show "{..p} = degree ` {monom (1::'a) x |x. x \<in> {..p}}"
+ proof
+ show "{..p} \<subseteq> degree ` {monom (1::'a) x |x. x \<in> {..p}}"
+ proof
+ fix x assume "x \<in> {..p}"
+ hence "monom (1::'a) x \<in> {monom 1 x |x. x \<in> {..p}}" by force
+ moreover have "degree (monom (1::'a) x) = x"
+ by (simp add: degree_monom_eq)
+ ultimately show "x \<in> degree ` {monom (1::'a) x |x. x \<in> {..p}}" by auto
+ qed
+ show "degree ` {monom (1::'a) x |x. x \<in> {..p}} \<subseteq> {..p}"
+ by (auto simp add: degree_monom_eq)
+ qed
+ next
+ fix y assume "y \<in> {monom (1::'a) x |x. x \<in> {..p}}"
+ hence "\<exists>z \<in> {..p}. y = monom (1::'a) z" by blast
+ then obtain z where "y = monom (1::'a) z" by blast
+ thus
+ "smult (coeff x (degree (monom (1::'a) (degree y)))) (monom (1::'a) (degree y)) =
+ smult (coeff x (degree y)) y"
+ by (simp add: smult_monom degree_monom_eq)
+ qed
+ finally have "x = (\<Sum>v\<in>{monom 1 x | x. x \<in> {..p}}.
+ smult ((\<lambda>v. coeff x (degree v)) v) v)" .
+ thus "x \<in> ?L" by (auto simp add: poly_vs.span_finite)
+ qed
+qed
+
+lemma monom_independent:
+ "poly_vs.independent {monom (1::'a::field) x | x. x \<le> p}"
+proof (rule poly_vs.independent_if_scalars_zero)
+ fix f::"'a poly \<Rightarrow> 'a"
+ assume h: "(\<Sum>x\<in>{monom 1 x |x. x \<le> p}. smult (f x) x) = 0"
+ have h': "(\<Sum>i\<le>p. monom (f (monom (1::'a) i)) i) =
+ (\<Sum>x\<in>{monom (1::'a) x |x. x \<le> p}. smult (f x) x)"
+ proof (rule sum.reindex_cong)
+ show "inj_on degree {monom (1::'a) x |x. x \<le> p}"
+ by (smt (verit) degree_monom_eq inj_on_def mem_Collect_eq zero_neq_one)
+ show "{..p} = degree ` {monom (1::'a) x |x. x \<le> p}"
+ proof
+ show "{..p} \<subseteq> degree ` {monom (1::'a) x |x. x \<le> p}"
+ proof
+ fix x assume "x \<in> {..p}"
+ then have "x = degree (monom (1::'a) x) \<and> x \<le> p"
+ by (auto simp: degree_monom_eq)
+ thus "x \<in> degree ` {monom (1::'a) x |x. x \<le> p}"
+ by blast
+ qed
+ show "degree ` {monom (1::'a) x |x. x \<le> p} \<subseteq> {..p}"
+ by (force simp: degree_monom_eq)
+ qed
+ qed (auto simp: degree_monom_eq smult_monom)
+
+ fix x::"'a poly"
+ assume "x \<in> {monom 1 x |x. x \<le> p}"
+ then obtain y where "y \<le> p" and "x = monom 1 y" by blast
+ hence "f x = coeff (\<Sum>x\<in>{monom 1 x |x. x \<le> p}. smult (f x) x) y"
+ by (auto simp: coeff_sum h'[symmetric])
+ thus "f x = 0"
+ using h by auto
+qed force
+
+lemma dim_degree: "poly_vs.dim {x. degree x \<le> n} = n + 1"
+ using poly_vs.dim_eq_card_independent[OF monom_independent]
+ by (auto simp: monom_span[symmetric] card_image image_Collect[symmetric]
+ inj_on_def monom_eq_iff')
+
+lemma degree_div:
+ fixes p q::"('a::idom_divide) poly"
+ assumes "q dvd p"
+ shows "degree (p div q) = degree p - degree q" using assms
+ by (metis (no_types, lifting) add_diff_cancel_left' degree_0 degree_mult_eq
+ diff_add_zero diff_zero div_by_0 dvd_div_eq_0_iff dvd_mult_div_cancel)
+
+lemma lead_coeff_div:
+ fixes p q::"('a::{idom_divide, inverse}) poly"
+ assumes "q dvd p"
+ shows "lead_coeff (p div q) = lead_coeff p / lead_coeff q" using assms
+ by (smt (z3) div_by_0 dvd_div_mult_self lead_coeff_mult leading_coeff_0_iff
+ nonzero_mult_div_cancel_right)
+
+lemma complex_poly_eq:
+ "r = map_poly of_real (map_poly Re r) + smult \<i> (map_poly of_real (map_poly Im r))"
+ by (auto simp: poly_eq_iff coeff_map_poly complex_eq)
+
+lemma complex_poly_cong:
+ "(map_poly Re p = map_poly Re q \<and> map_poly Im p = map_poly Im q) = (p = q)"
+ by (metis complex_poly_eq)
+
+lemma map_poly_Im_of_real: "map_poly Im (map_poly of_real p) = 0"
+ by (auto simp: poly_eq_iff coeff_map_poly)
+
+lemma mult_map_poly_imp_map_poly:
+ assumes "map_poly complex_of_real q = r * map_poly complex_of_real p"
+ "p \<noteq> 0"
+ shows "r = map_poly complex_of_real (map_poly Re r)"
+proof -
+ have h: "Im \<circ> (*) \<i> \<circ> complex_of_real = id" by fastforce
+ have "map_poly complex_of_real q = r * map_poly complex_of_real p"
+ using assms by blast
+ also have "... = (map_poly of_real (map_poly Re r) +
+ smult \<i> (map_poly of_real (map_poly Im r))) *
+ map_poly complex_of_real p"
+ using complex_poly_eq by fastforce
+ also have "... = map_poly of_real (map_poly Re r * p) +
+ smult \<i> (map_poly of_real (map_poly Im r * p))"
+ by (simp add: mult_poly_add_left)
+ finally have "map_poly complex_of_real q =
+ map_poly of_real (map_poly Re r * p) +
+ smult \<i> (map_poly of_real (map_poly Im r * p))" .
+ hence "0 = map_poly Im (map_poly of_real (map_poly Re r * p) +
+ smult \<i> (map_poly of_real (map_poly Im r * p)))"
+ by (auto simp: complex_poly_cong[symmetric] map_poly_Im_of_real)
+ also have "... = map_poly of_real (map_poly Im r * p)"
+ apply (rule poly_eqI)
+ by (auto simp: coeff_map_poly coeff_mult)
+ finally have "0 = map_poly complex_of_real (map_poly Im r) *
+ map_poly complex_of_real p"
+ by auto
+ hence "map_poly complex_of_real (map_poly Im r) = 0" using assms by fastforce
+ thus ?thesis apply (subst complex_poly_eq) by auto
+qed
+
+lemma map_poly_dvd:
+ fixes p q::"real poly"
+ assumes hdvd: "map_poly complex_of_real p dvd
+ map_poly complex_of_real q" "q \<noteq> 0"
+ shows "p dvd q"
+proof -
+ from hdvd obtain r
+ where h:"map_poly complex_of_real q = r * map_poly complex_of_real p"
+ by fastforce
+ hence "r = map_poly complex_of_real (map_poly Re r)"
+ using mult_map_poly_imp_map_poly assms by force
+ hence "map_poly complex_of_real q = map_poly complex_of_real (p * map_poly Re r)"
+ using h by auto
+ hence "q = p * map_poly Re r" using of_real_poly_eq_iff by blast
+ thus "p dvd q" by force
+qed
+
+lemma div_poly_eq_0:
+ fixes p q::"('a::idom_divide) poly"
+ assumes "q dvd p" "poly (p div q) x = 0" "q \<noteq> 0"
+ shows "poly p x = 0"
+ using assms by fastforce
+
+lemma poly_map_poly_of_real_cnj:
+ "poly (map_poly of_real p) (cnj z) = cnj (poly (map_poly of_real p) z)"
+ apply (induction p)
+ by auto
+
+text \<open>
+An induction rule on real polynomials, if $P \neq 0$ then either $(X-x)|P$ or
+$(X-z)(X-cnj z)|P$, we induct by dividing by these polynomials.
+\<close>
+lemma real_poly_roots_induct:
+ fixes P::"real poly \<Rightarrow> bool" and p::"real poly"
+ assumes IH_real: "\<And>p x. P p \<Longrightarrow> P (p * [:-x, 1:])"
+ and IH_complex: "\<And>p a b. b \<noteq> 0 \<Longrightarrow> P p
+ \<Longrightarrow> P (p * [: a*a + b*b, -2*a, 1 :])"
+ and H0: "\<And>a. P [:a:]"
+ defines "d \<equiv> degree p"
+ shows "P p"
+ using d_def
+proof (induction d arbitrary: p rule: less_induct)
+ fix p::"real poly"
+ assume IH: "(\<And>q. degree q < degree p \<Longrightarrow> P q)"
+ show "P p"
+ proof (cases "0 = degree p")
+ fix p::"real poly" assume "0 = degree p"
+ hence "\<exists> a. p = [:a:]"
+ by (simp add: degree_0_iff)
+ thus "P p" using H0 by blast
+ next
+ assume hdeg: "0 \<noteq> degree p"
+ hence "\<not> constant (poly (map_poly of_real p))"
+ by (metis (no_types, opaque_lifting) constant_def constant_degree of_real_eq_iff of_real_poly_map_poly)
+ then obtain z::complex where h: "poly (map_poly of_real p) z = 0"
+ using fundamental_theorem_of_algebra by blast
+ show "P p"
+ proof cases
+ assume "Im z = 0"
+ hence "z = Re z" by (simp add: complex_is_Real_iff)
+ moreover have "[:-z, 1:] dvd map_poly of_real p"
+ using h poly_eq_0_iff_dvd by blast
+ ultimately have "[:-(Re z), 1:] dvd p"
+ by (smt (z3) dvd_iff_poly_eq_0 h of_real_0 of_real_eq_iff of_real_poly_map_poly)
+ hence 2:"P (p div [:-Re z, 1:])"
+ apply (subst IH)
+ using hdeg by (auto simp: degree_div)
+ moreover have 1:"p = (p div [:- Re z, 1:]) * [:-Re z, 1:]"
+ by (metis \<open>[:- Re z, 1:] dvd p\<close> dvd_div_mult_self)
+ ultimately show "P p"
+ apply (subst 1)
+ by (rule IH_real[OF 2])
+ next
+ assume "Im z \<noteq> 0"
+ hence hcnj: "cnj z \<noteq> z" by (metis cnj.simps(2) neg_equal_zero)
+ have h2: "poly (map_poly of_real p) (cnj z) = 0"
+ using h poly_map_poly_of_real_cnj by force
+ have "[:-z, 1:] * [:-cnj z, 1:] dvd map_poly of_real p"
+ proof (rule divides_mult)
+ have "\<And>c. c dvd [:-z, 1:] \<Longrightarrow> c dvd [:- cnj z, 1:] \<Longrightarrow> is_unit c"
+ proof -
+ fix c
+ assume h:"c dvd [:-z, 1:]" hence "degree c \<le> 1" using divides_degree by fastforce
+ hence "degree c = 0 \<or> degree c = 1" by linarith
+ thus "c dvd [:- cnj z, 1:] \<Longrightarrow> is_unit c"
+ proof
+ assume "degree c = 0"
+ moreover have "c \<noteq> 0" using h by fastforce
+ ultimately show "is_unit c" by (simp add: is_unit_iff_degree)
+ next
+ assume hdeg: "degree c = 1"
+ then obtain x where 1:"[:-z, 1:] = x*c" using h by fastforce
+ hence "degree [:-z, 1:] = degree x + degree c"
+ by (metis add.inverse_neutral degree_mult_eq mult_cancel_right
+ mult_poly_0_left pCons_eq_0_iff zero_neq_neg_one)
+ hence "degree x = 0" using hdeg by auto
+ then obtain x' where 2: "x = [:x':]" using degree_0_iff by auto
+ assume "c dvd [:-cnj z, 1:]"
+ then obtain y where 3: "[:-cnj z, 1:] = y*c" by fastforce
+ hence "degree [:-cnj z, 1:] = degree y + degree c"
+ by (metis add.inverse_neutral degree_mult_eq mult_cancel_right
+ mult_poly_0_left pCons_eq_0_iff zero_neq_neg_one)
+ hence "degree y = 0" using hdeg by auto
+ then obtain y' where 4: "y = [:y':]" using degree_0_iff by auto
+ moreover from hdeg obtain a b where 5:"c = [:a, b:]" and 6: "b \<noteq> 0"
+ by (meson degree_eq_oneE)
+ from 1 2 5 6 have "x' = inverse b" by (auto simp: field_simps)
+ moreover from 3 4 5 6 have "y' = inverse b" by (auto simp: field_simps)
+ ultimately have "x = y" using 2 4 by presburger
+ then have "z = cnj z" using 1 3 by (metis neg_equal_iff_equal pCons_eq_iff)
+ thus "is_unit c" using hcnj by argo
+ qed
+ qed
+ thus "coprime [:- z, 1:] [:- cnj z, 1:]" by (meson not_coprimeE)
+ show "[:- z, 1:] dvd map_poly complex_of_real p"
+ using h poly_eq_0_iff_dvd by auto
+ show "[:- cnj z, 1:] dvd map_poly complex_of_real p"
+ using h2 poly_eq_0_iff_dvd by blast
+ qed
+ moreover have "[:- z, 1:] * [:- cnj z, 1:] =
+ map_poly of_real [:Re z*Re z + Im z*Im z, -2*Re z, 1:]"
+ by (auto simp: complex_eqI)
+ ultimately have hdvd:
+ "map_poly complex_of_real [:Re z*Re z + Im z*Im z, -2*Re z, 1:] dvd
+ map_poly complex_of_real p"
+ by force
+ hence "[:Re z*Re z + Im z*Im z, -2*Re z, 1:] dvd p" using map_poly_dvd
+ by blast
+ hence 2:"P (p div [:Re z*Re z + Im z*Im z, -2*Re z, 1:])"
+ apply (subst IH)
+ using hdeg by (auto simp: degree_div)
+ moreover have 1:
+ "p = (p div [:Re z*Re z + Im z*Im z, -2*Re z, 1:]) *
+ [:Re z*Re z + Im z*Im z, -2*Re z, 1:]"
+ apply (subst dvd_div_mult_self)
+ using \<open>[:Re z*Re z + Im z*Im z, -2*Re z, 1:] dvd p\<close> by auto
+ ultimately show "P p"
+ apply (subst 1)
+ apply (rule IH_complex[of "Im z" _ "Re z"])
+ apply (meson \<open>Im z \<noteq> 0\<close>)
+ by blast
+ qed
+ qed
+qed
+
+subsection \<open>The reciprocal polynomial\<close>
+
+definition reciprocal_poly :: "nat \<Rightarrow> 'a::zero poly \<Rightarrow> 'a poly"
+ where "reciprocal_poly p P =
+ Poly (rev ((coeffs P) @ (replicate (p - degree P) 0)))"
+
+lemma reciprocal_0: "reciprocal_poly p 0 = 0" by (simp add: reciprocal_poly_def)
+
+lemma reciprocal_1: "reciprocal_poly p 1 = monom 1 p"
+ by (simp add: reciprocal_poly_def monom_altdef Poly_append)
+
+lemma coeff_reciprocal:
+ assumes hi: "i \<le> p" and hP: "degree P \<le> p"
+ shows "coeff (reciprocal_poly p P) i = coeff P (p - i)"
+proof cases
+ assume "i < p - degree P"
+ hence "degree P < p - i" using hP by linarith
+ thus "coeff (reciprocal_poly p P) i = coeff P (p - i)"
+ by (auto simp: reciprocal_poly_def nth_default_append coeff_eq_0)
+next
+ assume h: "\<not>i < p - degree P"
+ show "coeff (reciprocal_poly p P) i = coeff P (p - i)"
+ proof cases
+ assume "P = 0"
+ thus "coeff (reciprocal_poly p P) i = coeff P (p - i)"
+ by (simp add: reciprocal_0)
+ next
+ assume hP': "P \<noteq> 0"
+ have "degree P \<ge> p - i" using h hP by linarith
+ moreover hence "(i - (p - degree P)) < length (rev (coeffs P))"
+ using hP' hP hi by (auto simp: length_coeffs)
+ thus "coeff (reciprocal_poly p P) i = coeff P (p - i)"
+ by (auto simp: reciprocal_poly_def nth_default_append coeff_eq_0 hP hP'
+ nth_default_nth rev_nth calculation coeffs_nth length_coeffs_degree)
+ qed
+qed
+
+lemma coeff_reciprocal_less:
+ assumes hn: "p < i" and hP: "degree P \<le> p"
+ shows "coeff (reciprocal_poly p P) i = 0"
+proof cases
+ assume "P = 0"
+ thus ?thesis by (auto simp: reciprocal_0)
+next
+ assume "P \<noteq> 0"
+ thus ?thesis
+ using hn
+ by (auto simp: reciprocal_poly_def nth_default_append
+ nth_default_eq_dflt_iff hP length_coeffs)
+qed
+
+lemma reciprocal_monom:
+ assumes "n \<le> p"
+ shows "reciprocal_poly p (monom a n) = monom a (p-n)"
+proof (cases "a=0")
+ case True
+ then show ?thesis by (simp add: reciprocal_0)
+next
+ case False
+ with \<open>n\<le>p\<close> show ?thesis
+ apply (rule_tac poly_eqI)
+ by (metis coeff_monom coeff_reciprocal coeff_reciprocal_less
+ diff_diff_cancel diff_le_self lead_coeff_monom not_le_imp_less)
+qed
+
+lemma reciprocal_degree: "reciprocal_poly (degree P) P = reflect_poly P"
+ by (auto simp add: reciprocal_poly_def reflect_poly_def)
+
+lemma degree_reciprocal:
+ fixes P :: "('a::zero) poly"
+ assumes hP: "degree P \<le> p"
+ shows "degree (reciprocal_poly p P) \<le> p"
+proof (auto simp add: reciprocal_poly_def)
+ have "degree (reciprocal_poly p P) \<le>
+ length (replicate (p - degree P) (0::'a) @ rev (coeffs P))"
+ by (metis degree_Poly reciprocal_poly_def rev_append rev_replicate)
+ thus "degree (Poly (replicate (p - degree P) 0 @ rev (coeffs P))) \<le> p"
+ by (smt Suc_le_mono add_Suc_right coeffs_Poly degree_0 hP le_SucE le_SucI
+ le_add_diff_inverse2 le_zero_eq length_append length_coeffs_degree
+ length_replicate length_rev length_strip_while_le reciprocal_0
+ reciprocal_poly_def rev_append rev_replicate)
+qed
+
+lemma reciprocal_0_iff:
+ assumes hP: "degree P \<le> p"
+ shows "(reciprocal_poly p P = 0) = (P = 0)"
+proof
+ assume h: "reciprocal_poly p P = 0"
+ show "P = 0"
+ proof (rule poly_eqI)
+ fix n
+ show "coeff P n = coeff 0 n"
+ proof cases
+ assume hn: "n \<le> p"
+ hence "p - n \<le> p" by auto
+ hence "coeff (reciprocal_poly p P) (p - n) = coeff P n"
+ using hP hn by (auto simp: coeff_reciprocal)
+ thus ?thesis using h by auto
+ next
+ assume hn: "\<not> n \<le> p"
+ thus ?thesis using hP by (metis coeff_0 dual_order.trans le_degree)
+ qed
+ qed
+next
+ assume "P = 0"
+ thus "reciprocal_poly p P = 0" using reciprocal_0 by fast
+qed
+
+lemma poly_reciprocal:
+ fixes P::"'a::field poly"
+ assumes hp: "degree P \<le> p" and hx: "x \<noteq> 0"
+ shows "poly (reciprocal_poly p P) x = x^p * (poly P (inverse x))"
+proof -
+ have "poly (reciprocal_poly p P) x
+ = poly ((Poly ((replicate (p - degree P) 0) @ rev (coeffs P)))) x"
+ by (auto simp add: hx reflect_poly_def reciprocal_poly_def)
+ also have "... = poly ((monom 1 (p - degree P)) * (reflect_poly P)) x"
+ by (auto simp add: reflect_poly_def Poly_append)
+ also have "... = x^(p - degree P) * x ^ degree P * poly P (inverse x)"
+ by (auto simp add: poly_reflect_poly_nz poly_monom hx)
+ also have "... = x^p * poly P (inverse x)"
+ by (auto simp add: hp power_add[symmetric])
+ finally show ?thesis .
+qed
+
+lemma reciprocal_fcompose:
+ fixes P::"('a::{ring_char_0,field}) poly"
+ assumes hP: "degree P \<le> p"
+ shows "reciprocal_poly p P = monom 1 (p - degree P) * fcompose P 1 [:0, 1:]"
+proof (rule poly_eq_by_eval, cases)
+ fix x::'a
+ assume hx: "x = 0"
+ hence "poly (reciprocal_poly p P) x = coeff P p"
+ using hP by (auto simp: poly_0_coeff_0 coeff_reciprocal)
+ moreover have "poly (monom 1 (p - degree P)
+ * fcompose P 1 [:0, 1:]) x = coeff P p"
+ proof cases
+ assume "degree P = p"
+ thus ?thesis
+ apply (induction P arbitrary: p)
+ using hx by (auto simp: poly_monom degree_0_iff fcompose_pCons)
+ next
+ assume "degree P \<noteq> p"
+ hence "degree P < p" using hP by auto
+ thus ?thesis
+ using hx by (auto simp: poly_monom coeff_eq_0)
+ qed
+ ultimately show "poly (reciprocal_poly p P) x = poly (monom 1 (p - degree P) * fcompose P 1 [:0, 1:]) x"
+ by presburger
+next
+ fix x::'a assume "x \<noteq> 0"
+ thus "poly (reciprocal_poly p P) x =
+ poly (monom 1 (p - degree P) * fcompose P 1 [:0, 1:]) x"
+ using hP
+ by (auto simp: poly_reciprocal poly_fcompose inverse_eq_divide
+ poly_monom power_diff)
+qed
+
+lemma reciprocal_reciprocal:
+ fixes P :: "'a::{field,ring_char_0} poly"
+ assumes hP: "degree P \<le> p"
+ shows "reciprocal_poly p (reciprocal_poly p P) = P"
+proof (rule poly_eq_by_eval)
+ fix x
+ show "poly (reciprocal_poly p (reciprocal_poly p P)) x = poly P x"
+ proof cases
+ assume "x = 0"
+ thus "poly (reciprocal_poly p (reciprocal_poly p P)) x = poly P x"
+ using hP
+ by (auto simp: poly_0_coeff_0 coeff_reciprocal degree_reciprocal)
+ next
+ assume hx: "x \<noteq> 0"
+ hence "poly (reciprocal_poly p (reciprocal_poly p P)) x
+ = x ^ p * (inverse x ^ p * poly P x)" using hP
+ by (auto simp: poly_reciprocal degree_reciprocal)
+ thus "poly (reciprocal_poly p (reciprocal_poly p P)) x = poly P x"
+ using hP hx left_right_inverse_power right_inverse by auto
+ qed
+qed
+
+lemma reciprocal_smult:
+ fixes P :: "'a::idom poly"
+ assumes h: "degree P \<le> p"
+ shows "reciprocal_poly p (smult n P) = smult n (reciprocal_poly p P)"
+proof cases
+ assume "n = 0"
+ thus ?thesis by (auto simp add: reciprocal_poly_def)
+next
+ assume "n \<noteq> 0"
+ thus ?thesis
+ by (auto simp add: reciprocal_poly_def smult_Poly coeffs_smult
+ rev_map[symmetric])
+qed
+
+lemma reciprocal_add:
+ fixes P Q :: "'a::comm_semiring_0 poly"
+ assumes "degree P \<le> p" and "degree Q \<le> p"
+ shows "reciprocal_poly p (P + Q) = reciprocal_poly p P + reciprocal_poly p Q"
+(is "?L = ?R")
+proof (rule poly_eqI, cases)
+ fix n
+ assume "n \<le> p"
+ then show "coeff ?L n = coeff ?R n"
+ using assms by (auto simp: degree_add_le coeff_reciprocal)
+next
+ fix n assume "\<not>n \<le> p"
+ then show "coeff ?L n = coeff ?R n"
+ using assms by (auto simp: degree_add_le coeff_reciprocal_less)
+qed
+
+lemma reciprocal_diff:
+ fixes P Q :: "'a::comm_ring poly"
+ assumes "degree P \<le> p" and "degree Q \<le> p"
+ shows "reciprocal_poly p (P - Q) = reciprocal_poly p P - reciprocal_poly p Q"
+ by (metis (no_types, lifting) ab_group_add_class.ab_diff_conv_add_uminus assms
+ add_diff_cancel degree_add_le degree_minus diff_add_cancel reciprocal_add)
+
+lemma reciprocal_sum:
+ fixes P :: "'a \<Rightarrow> 'b::comm_semiring_0 poly"
+ assumes hP: "\<And>k. degree (P k) \<le> p"
+ shows "reciprocal_poly p (\<Sum>k\<in>A. P k) = (\<Sum>k\<in>A. reciprocal_poly p (P k))"
+proof (induct A rule: infinite_finite_induct)
+ case (infinite A)
+ then show ?case by (simp add: reciprocal_0)
+next
+ case empty
+ then show ?case by (simp add: reciprocal_0)
+next
+ case (insert x F)
+ assume "x \<notin> F"
+ and "reciprocal_poly p (sum P F) = (\<Sum>k\<in>F. reciprocal_poly p (P k))"
+ and "finite F"
+ moreover hence "reciprocal_poly p (sum P (insert x F))
+ = reciprocal_poly p (P x) + reciprocal_poly p (sum P F)"
+ by (auto simp add: reciprocal_add hP degree_sum_le)
+ ultimately show "reciprocal_poly p (sum P (insert x F))
+ = (\<Sum>k\<in>insert x F. reciprocal_poly p (P k))"
+ by (auto simp: Groups_Big.comm_monoid_add_class.sum.insert_if)
+qed
+
+lemma reciprocal_mult:
+ fixes P Q::"'a::{ring_char_0,field} poly"
+ assumes "degree (P * Q) \<le> p"
+ and "degree P \<le> p" and "degree Q \<le> p"
+ shows "monom 1 p * reciprocal_poly p (P * Q) =
+ reciprocal_poly p P * reciprocal_poly p Q"
+proof (cases "P=0 \<or> Q=0")
+ case True
+ then show ?thesis using assms(1)
+ by (auto simp: reciprocal_fcompose fcompose_mult)
+next
+ case False
+ then show ?thesis
+ using assms
+ by (auto simp: degree_mult_eq mult_monom reciprocal_fcompose fcompose_mult)
+qed
+
+lemma reciprocal_reflect_poly:
+ fixes P::"'a::{ring_char_0,field} poly"
+ assumes hP: "degree P \<le> p"
+ shows "reciprocal_poly p P = monom 1 (p - degree P) * reflect_poly P"
+proof (rule poly_eqI)
+ fix n
+ show "coeff (reciprocal_poly p P) n =
+ coeff (monom 1 (p - degree P) * reflect_poly P) n"
+ proof cases
+ assume "n \<le> p"
+ thus ?thesis using hP
+ by (auto simp: coeff_reciprocal coeff_monom_mult coeff_reflect_poly coeff_eq_0)
+ next
+ assume "\<not> n \<le> p"
+ thus ?thesis using hP
+ by (auto simp: coeff_reciprocal_less coeff_monom_mult coeff_reflect_poly)
+ qed
+qed
+
+lemma map_poly_reciprocal:
+ assumes "degree P \<le> p" and "f 0 = 0"
+ shows "map_poly f (reciprocal_poly p P) = reciprocal_poly p (map_poly f P)"
+proof (rule poly_eqI)
+ fix n
+ show "coeff (map_poly f (reciprocal_poly p P)) n =
+ coeff (reciprocal_poly p (map_poly f P)) n"
+ proof (cases "n\<le>p")
+ case True
+ then show ?thesis
+ apply (subst coeff_reciprocal[OF True])
+ subgoal by (meson assms(1) assms(2) degree_map_poly_le le_trans)
+ by (simp add: assms(1) assms(2) coeff_map_poly coeff_reciprocal)
+ next
+ case False
+ then show ?thesis
+ by (metis assms(1) assms(2) coeff_map_poly coeff_reciprocal_less
+ degree_map_poly_le dual_order.trans leI)
+ qed
+qed
+
+subsection \<open>More about @{term proots_count}\<close>
+
+lemma proots_count_monom:
+ assumes "0 \<notin> A"
+ shows "proots_count (monom 1 d) A = 0"
+ using assms by (auto simp: proots_count_def poly_monom)
+
+lemma proots_count_reciprocal:
+ fixes P::"'a::{ring_char_0,field} poly"
+ assumes hP: "degree P \<le> p" and h0: "P \<noteq> 0" and h0': "0 \<notin> A"
+ shows "proots_count (reciprocal_poly p P) A = proots_count P {x. inverse x \<in> A}"
+proof -
+ have "proots_count (reciprocal_poly p P) A =
+ proots_count (fcompose P 1 [:0, 1:]) A"
+ apply (subst reciprocal_fcompose[OF hP], subst proots_count_times)
+ subgoal using h0 by (metis hP reciprocal_0_iff reciprocal_fcompose)
+ subgoal using h0' by (auto simp: proots_count_monom)
+ done
+
+ also have "... = proots_count P {x. inverse x \<in> A}"
+ proof (rule proots_fcompose_bij_eq[symmetric])
+ show "bij_betw (\<lambda>x. poly 1 x / poly [:0, 1:] x) A {x. inverse x \<in> A}"
+ proof (rule bij_betw_imageI)
+ show "inj_on (\<lambda>x. poly 1 x / poly [:0, 1:] x) A"
+ by (simp add: inj_on_def)
+
+ show "(\<lambda>x. poly 1 x / poly [:0, 1:] x) ` A = {x. inverse x \<in> A}"
+ proof
+ show "(\<lambda>x. poly 1 x / poly [:0, 1:] x) ` A \<subseteq> {x. inverse x \<in> A}"
+ by force
+ show "{x. inverse x \<in> A} \<subseteq> (\<lambda>x. poly 1 x / poly [:0, 1:] x) ` A"
+ proof
+ fix x assume "x \<in> {x::'a. inverse x \<in> A}"
+ hence "x = poly 1 (inverse x) / poly [:0, 1:] (inverse x) \<and> inverse x \<in> A"
+ by (auto simp: inverse_eq_divide)
+ thus "x \<in> (\<lambda>x. poly 1 x / poly [:0, 1:] x) ` A" by blast
+ qed
+ qed
+ qed
+ next
+ show "\<forall>c. 1 \<noteq> smult c [:0, 1:]"
+ by (metis coeff_pCons_0 degree_1 lead_coeff_1 pCons_0_0 pcompose_0'
+ pcompose_smult smult_0_right zero_neq_one)
+ qed (auto simp: assms infinite_UNIV_char_0)
+ finally show ?thesis by linarith
+qed
+
+lemma proots_count_reciprocal':
+ fixes P::"real poly"
+ assumes hP: "degree P \<le> p" and h0: "P \<noteq> 0"
+ shows "proots_count P {x. 0 < x \<and> x < 1} =
+ proots_count (reciprocal_poly p P) {x. 1 < x}"
+proof (subst proots_count_reciprocal)
+ show "proots_count P {x. 0 < x \<and> x < 1} =
+ proots_count P {x. inverse x \<in> Collect ((<) 1)}"
+ apply (rule arg_cong[of _ _ "proots_count P"])
+ using one_less_inverse_iff by fastforce
+qed (use assms in auto)
+
+lemma proots_count_pos:
+ assumes "proots_count P S > 0"
+ shows "\<exists>x \<in> S. poly P x = 0"
+proof (rule ccontr)
+ assume "\<not> (\<exists>x\<in>S. poly P x = 0)"
+ hence "\<And>x. x \<in> S \<Longrightarrow> poly P x \<noteq> 0" by blast
+ hence "\<And>x. x \<in> S \<Longrightarrow> order x P = 0" using order_0I by blast
+ hence "proots_count P S = 0" by (force simp: proots_count_def)
+ thus False using assms by presburger
+qed
+
+lemma proots_count_of_root_set:
+ assumes "P \<noteq> 0" "R \<subseteq> S" and "\<And>x. x\<in>R \<Longrightarrow> poly P x = 0"
+ shows "proots_count P S \<ge> card R"
+proof -
+ have "card R \<le> card (proots_within P S)"
+ apply (rule card_mono)
+ subgoal using assms by auto
+ subgoal using assms(2) assms(3) by (auto simp: proots_within_def)
+ done
+ also have "... \<le> proots_count P S"
+ by (rule card_proots_within_leq[OF assms(1)])
+ finally show ?thesis .
+qed
+
+lemma proots_count_of_root: assumes "P \<noteq> 0" "x\<in>S" "poly P x = 0"
+ shows "proots_count P S > 0"
+ using proots_count_of_root_set[of P "{x}" S] assms by force
+
+subsection \<open>More about @{term changes}\<close>
+
+lemma changes_nonneg: "0 \<le> changes xs"
+ apply (induction xs rule: changes.induct)
+ by simp_all
+
+lemma changes_replicate_0: shows "changes (replicate n 0) = 0"
+ apply (induction n)
+ by auto
+
+lemma changes_append_replicate_0: "changes (xs @ replicate n 0) = changes xs"
+proof (induction xs rule: changes.induct)
+ case (2 uu)
+ then show ?case
+ apply (induction n)
+ by auto
+qed (auto simp: changes_replicate_0)
+
+lemma changes_scale_Cons:
+ fixes xs:: "real list" assumes hs: "s > 0"
+ shows "changes (s * x # xs) = changes (x # xs)"
+ apply (induction xs rule: changes.induct)
+ using assms by (auto simp: mult_less_0_iff zero_less_mult_iff)
+
+lemma changes_scale:
+ fixes xs::"('a::linordered_idom) list"
+ assumes hs: "\<And>i. i < n \<Longrightarrow> s i > 0" and hn: "length xs \<le> n"
+ shows "changes [s i * (nth_default 0 xs i). i \<leftarrow> [0..<n]] = changes xs"
+using assms
+proof (induction xs arbitrary: s n rule: changes.induct)
+ case 1
+ show ?case by (auto simp: map_replicate_const changes_replicate_0)
+next
+ case (2 uu)
+ show ?case
+ proof (cases n)
+ case 0
+ then show ?thesis by force
+ next
+ case (Suc m)
+ hence "map (\<lambda>i. s i * nth_default 0 [uu] i) [0..<n] = [s 0 * uu] @ replicate m 0"
+ proof (induction m arbitrary: n)
+ case (Suc m n)
+ from Suc have "map (\<lambda>i. s i * nth_default 0 [uu] i) [0..<Suc m] =
+ [s 0 * uu] @ replicate m 0"
+ by meson
+ hence "map (\<lambda>i. s i * nth_default 0 [uu] i) [0..<n] =
+ [s 0 * uu] @ replicate m 0 @ [0]"
+ using Suc by auto
+ also have "... = [s 0 * uu] @ replicate (Suc m) 0"
+ by (simp add: replicate_append_same)
+ finally show ?case .
+ qed fastforce
+ then show ?thesis
+ by (metis changes.simps(2) changes_append_replicate_0)
+ qed
+next
+ case (3 a b xs s n)
+ obtain m where hn: "n = m + 2"
+ using 3(5)
+ by (metis add_2_eq_Suc' diff_diff_cancel diff_le_self length_Suc_conv
+ nat_arith.suc1 ordered_cancel_comm_monoid_diff_class.add_diff_inverse)
+ hence h:
+ "map (\<lambda>i. s i * nth_default 0 (a # b # xs) i) [0..<n] =
+ s 0 * a # s 1 * b # map (\<lambda>i.
+ (\<lambda> i. s (i+2)) i * nth_default 0 (xs) i) [0..<m]"
+ apply (induction m arbitrary: n)
+ by auto
+ consider (neg)"a*b<0" | (nil)"b=0" | (pos)"\<not>a*b<0 \<and> \<not>b=0" by linarith
+ then show ?case
+ proof (cases)
+ case neg
+ hence
+ "changes (map (\<lambda>i. s i * nth_default 0 (a # b # xs) i) [0..<n]) =
+ 1 + changes (s 1 * b # map (\<lambda>i. (\<lambda> i. s (i+2)) i
+ * nth_default 0 (xs) i) [0..<m])"
+ apply (subst h)
+ using 3(4)[of 0] 3(4)[of 1] hn
+ by (metis (no_types, lifting) changes.simps(3) mult_less_0_iff pos2
+ mult_pos_pos one_less_numeral_iff semiring_norm(76) trans_less_add2)
+ also have
+ "... = 1 + changes (map (\<lambda>i. s (Suc i) * nth_default 0 (b # xs) i) [0..<Suc m])"
+ apply (rule arg_cong[of _ _ "\<lambda> x. 1 + changes x"])
+ apply (induction m)
+ by auto
+ also have "... = changes (a # b # xs)"
+ apply (subst 3(1)[OF neg])
+ using 3 neg hn by auto
+ finally show ?thesis .
+ next
+ case nil
+ hence "changes (map (\<lambda>i. s i * nth_default 0 (a # b # xs) i) [0..<n]) =
+ changes (s 0 * a # map (\<lambda>i. (\<lambda> i. s (i+2)) i * nth_default 0 (xs) i) [0..<m])"
+ apply (subst h)
+ using 3(4)[of 0] 3(4)[of 1] hn
+ by auto
+ also have
+ "... = changes (map
+ (\<lambda>i. s (if i = 0 then 0 else Suc i) * nth_default 0 (a # xs) i)
+ [0..<Suc m])"
+ apply (rule arg_cong[of _ _ "\<lambda> x. changes x"])
+ apply (induction m)
+ by auto
+ also have "... = changes (a # b # xs)"
+ apply (subst 3(2))
+ using 3 nil hn by auto
+ finally show ?thesis .
+ next
+ case pos
+ hence "changes (map (\<lambda>i. s i * nth_default 0 (a # b # xs) i) [0..<n]) =
+ changes (s 1 * b # map (\<lambda>i. (\<lambda> i. s (i+2)) i * nth_default 0 (xs) i) [0..<m])"
+ apply (subst h)
+ using 3(4)[of 0] 3(4)[of 1] hn
+ by (metis (no_types, lifting) changes.simps(3) divisors_zero
+ mult_less_0_iff nat_1_add_1 not_square_less_zero one_less_numeral_iff
+ semiring_norm(76) trans_less_add2 zero_less_mult_pos zero_less_two)
+ also have
+ "... = changes (map (\<lambda>i. s (Suc i) * nth_default 0 (b # xs) i) [0..<Suc m])"
+ apply (rule arg_cong[of _ _ "\<lambda> x. changes x"])
+ apply (induction m)
+ by auto
+ also have "... = changes (a # b # xs)"
+ apply (subst 3(3))
+ using 3 pos hn by auto
+ finally show ?thesis .
+ qed
+qed
+
+lemma changes_scale_const: fixes xs::"'a::linordered_idom list"
+ assumes hs: "s \<noteq> 0"
+ shows "changes (map ((*) s) xs) = changes xs"
+ apply (induction xs rule: changes.induct)
+ apply (simp, force)
+ using hs by (auto simp: mult_less_0_iff zero_less_mult_iff)
+
+lemma changes_snoc: fixes xs::"'a::linordered_idom list"
+ shows "changes (xs @ [b, a]) = (if a * b < 0 then 1 + changes (xs @ [b])
+ else if b = 0 then changes (xs @ [a]) else changes (xs @ [b]))"
+ apply (induction xs rule: changes.induct)
+ subgoal by (force simp: mult_less_0_iff)
+ subgoal by (force simp: mult_less_0_iff)
+ subgoal by force
+ done
+
+lemma changes_rev: fixes xs:: "'a::linordered_idom list"
+ shows "changes (rev xs) = changes xs"
+ apply (induction xs rule: changes.induct)
+ by (auto simp: changes_snoc)
+
+lemma changes_rev_about: fixes xs:: "'a::linordered_idom list"
+ shows "changes (replicate (p - length xs) 0 @ rev xs) = changes xs"
+proof (induction p)
+ case (Suc p)
+ then show ?case
+ proof cases
+ assume "\<not>Suc p \<le> length xs"
+ hence "Suc p - length xs = Suc (p - length xs)" by linarith
+ thus ?case using Suc.IH changes_rev by auto
+ qed (auto simp: changes_rev)
+qed (auto simp: changes_rev)
+
+lemma changes_add_between:
+ assumes "a \<le> x" and "x \<le> b"
+ shows "changes (as @ [a, b] @ bs) = changes (as @ [a, x, b] @ bs)"
+proof (induction as rule: changes.induct)
+ case 1
+ then show ?case using assms
+ apply (induction bs)
+ by (auto simp: mult_less_0_iff)
+next
+ case (2 c)
+ then show ?case
+ apply (induction bs)
+ using assms by (auto simp: mult_less_0_iff)
+next
+ case (3 y z as)
+ then show ?case
+ using assms by (auto simp: mult_less_0_iff)
+qed
+
+lemma changes_all_nonneg: assumes "\<And>i. nth_default 0 xs i \<ge> 0" shows "changes xs = 0"
+ using assms
+proof (induction xs rule: changes.induct)
+ case (3 x1 x2 xs)
+ moreover assume "(\<And>i. 0 \<le> nth_default 0 (x1 # x2 # xs) i)"
+ moreover hence "(\<And>i. 0 \<le> nth_default 0 (x1 # xs) i)"
+ and "(\<And>i. 0 \<le> nth_default 0 (x2 # xs) i)"
+ and "x1 * x2 \<ge> 0"
+ proof -
+ fix i
+ assume h:"(\<And>i. 0 \<le> nth_default 0 (x1 # x2 # xs) i)"
+ show "0 \<le> nth_default 0 (x1 # xs) i"
+ proof (cases i)
+ case 0
+ then show ?thesis using h[of 0] by force
+ next
+ case (Suc nat)
+ then show ?thesis using h[of "Suc (Suc nat)"] by force
+ qed
+ show "0 \<le> nth_default 0 (x2 # xs) i" using h[of "Suc i"] by simp
+ show "x1*x2 \<ge> 0" using h[of 0] h[of 1] by simp
+ qed
+ ultimately show ?case by auto
+qed auto
+
+lemma changes_pCons: "changes (coeffs (pCons 0 f)) = changes (coeffs f)"
+ by (auto simp: cCons_def)
+
+lemma changes_increasing:
+ assumes "\<And>i. i < length xs - 1 \<Longrightarrow> xs ! (i + 1) \<ge> xs ! i"
+ and "length xs > 1"
+ and "hd xs < 0"
+ and "last xs > 0"
+ shows "changes xs = 1"
+ using assms
+proof (induction xs rule:changes.induct)
+ case (3 x y xs)
+ consider (neg)"x*y<0" | (nil)"y=0" | (pos)"\<not>x*y<0 \<and> \<not>y=0" by linarith
+ then show ?case
+ proof cases
+ case neg
+ have "changes (y # xs) = 0"
+ proof (rule changes_all_nonneg)
+ fix i
+ show "0 \<le> nth_default 0 (y # xs) i"
+ proof (cases "i < length (y # xs)")
+ case True
+ then show ?thesis using 3(4)[of i]
+ apply (induction i)
+ subgoal using 3(6) neg by (fastforce simp: mult_less_0_iff)
+ subgoal using 3(4) by (auto simp: nth_default_def)
+ done
+ next
+ case False
+ then show ?thesis by (simp add: nth_default_def)
+ qed
+ qed
+ thus "changes (x # y # xs) = 1"
+ using neg by force
+ next
+ case nil
+ hence "xs \<noteq> []" using 3(7) by force
+ have h: "\<And>i. i < length (x # xs) - 1 \<Longrightarrow> (x # xs) ! i \<le> (x # xs) ! (i + 1)"
+ proof -
+ fix i assume "i < length (x # xs) - 1"
+ thus "(x # xs) ! i \<le> (x # xs) ! (i + 1)"
+ apply (cases "i = 0")
+ subgoal using 3(4)[of 0] 3(4)[of 1] \<open>xs \<noteq> []\<close> by force
+ using 3(4)[of "i+1"] by simp
+ qed
+ have "changes (x # xs) = 1"
+ apply (rule 3(2))
+ using nil h \<open>xs \<noteq> []\<close> 3(6) 3(7) by auto
+ thus ?thesis
+ using nil by force
+ next
+ case pos
+ hence "xs \<noteq> []" using 3(6) 3(7) by (fastforce simp: mult_less_0_iff)
+ have "changes (y # xs) = 1"
+ proof (rule 3(3))
+ show "\<not> x * y < 0" "y \<noteq> 0"
+ using pos by auto
+ show "\<And>i. i < length (y # xs) - 1
+ \<Longrightarrow> (y # xs) ! i \<le> (y # xs) ! (i + 1)"
+ using 3(4) by force
+ show "1 < length (y # xs)"
+ using \<open>xs \<noteq> []\<close> by force
+ show "hd (y # xs) < 0"
+ using 3(6) pos by (force simp: mult_less_0_iff)
+ show "0 < last (y # xs)"
+ using 3(7) by force
+ qed
+ thus ?thesis using pos by auto
+ qed
+qed auto
+
end
\ No newline at end of file
diff --git a/thys/Three_Circles/Three_Circles.thy b/thys/Three_Circles/Three_Circles.thy
--- a/thys/Three_Circles/Three_Circles.thy
+++ b/thys/Three_Circles/Three_Circles.thy
@@ -1,769 +1,769 @@
-section \<open>Proof of the theorem of three circles\<close>
-
-theory Three_Circles
- imports "Bernstein" "Normal_Poly"
-begin
-
-text \<open>
-The theorem of three circles is a result in real algebraic geometry about the number of real roots
-in an interval. It says if the number of roots in certain circles in the complex plane are zero or
-one then the number of roots in the circles is equal to the sign changes of the Bernstein
-coefficients on that interval for which the circles intersect the real line. This can then be used
-to determine if an interval has a real root in the bisection procedure, which is more efficient than
-Descartes' rule of signs.
-
-The proof here follows Theorem 10.50 in
- Basu, S., Pollack, R., Roy, M.-F.: Algorithms in Real Algebraic Geometry.
- Springer Berlin Heidelberg, Berlin, Heidelberg (2016).
-
-This theorem has also been fomalised in Coq \cite{zsido2014theorem}. The relationship
-between this theorem and root isolation has been elaborated in Eigenwillig's PhD
-thesis \cite{eigenwillig2008real}.
-\<close>
-
-subsection \<open>No sign changes case\<close>
-
-declare degree_pcompose[simp del]
-
-corollary descartes_sign_zero:
- fixes p::"real poly"
- assumes "\<And>x::complex. poly (map_poly of_real p) x = 0 \<Longrightarrow> Re x \<le> 0"
- and "lead_coeff p = 1"
- shows "coeff p i \<ge> 0"
- using assms
-proof (induction p arbitrary: i rule: real_poly_roots_induct)
- case (1 p x)
- interpret map_poly_idom_hom complex_of_real ..
- have h: "\<And> i. 0 \<le> coeff p i"
- apply (rule 1(1))
- using 1(2) apply (metis lambda_zero of_real_poly_map_mult poly_mult)
- using 1(3) apply (metis lead_coeff_1 lead_coeff_mult lead_coeff_pCons(1)
- mult_cancel_right2 pCons_one zero_neq_one)
- done
- have "x \<le> 0"
- apply (subst Re_complex_of_real[symmetric])
- apply (rule 1(2))
- apply (subst hom_mult)
- by (auto)
- thus ?case
- apply (cases i)
- subgoal using h[of i] h[of "i-1"]
- by (fastforce simp: coeff_pCons mult_nonneg_nonpos2)
- subgoal using h[of i] h[of "i-1"] mult_left_mono_neg
- by (fastforce simp: coeff_pCons)
- done
-next
- case (2 p a b)
- interpret map_poly_idom_hom complex_of_real ..
- have h: "\<And> i. 0 \<le> coeff p i"
- apply (rule 2(2))
- using 2(3) apply (metis lambda_zero of_real_poly_map_mult poly_mult)
- using 2(4) apply (metis lead_coeff_1 lead_coeff_mult lead_coeff_pCons(1)
- mult_cancel_right2 pCons_one zero_neq_one)
- done
- have "Re (a + b * \<i>) \<le> 0"
- apply (rule 2(3))
- apply (subst hom_mult)
- by (auto simp: algebra_simps)
- hence 1: "0 \<le> - 2 * a" by fastforce
- have 2: "0 \<le> a * a + b * b" by fastforce
- have "\<And> x. 0 \<le> coeff [:a * a + b * b, - 2 * a, 1:] x"
- proof -
- fix x
- show "0 \<le> coeff [:a * a + b * b, - 2 * a, 1:] x"
- using 2 apply (cases "x = 0", fastforce)
- using 1 apply (cases "x = 1", fastforce)
- apply (cases "x = 2", fastforce simp: coeff_pCons)
- by (auto simp: coeff_eq_0)
- qed
- thus ?case
- apply (subst mult.commute, subst coeff_mult)
- apply (rule sum_nonneg, rule mult_nonneg_nonneg[OF _ h])
- by auto
-next
- case (3 a)
- then show ?case
- by (smt (z3) bot_nat_0.extremum_uniqueI degree_1 le_degree
- lead_coeff_pCons(2) pCons_one)
-qed
-
-definition circle_01_diam :: "complex set" where
-"circle_01_diam =
- {x. cmod (x - (of_nat 1 :: complex)/(of_nat 2)) < (real 1)/(real 2)}"
-
-lemma pos_real_map:
- "{x::complex. 1 / x \<in> (\<lambda>x. x + 1) ` {x. 0 < Re x}} = circle_01_diam"
-proof
- show "{x. 1 / x \<in> (\<lambda>x. x + 1) ` {x. 0 < Re x}} \<subseteq> circle_01_diam"
- proof
- fix x assume "x \<in> {x. 1 / x \<in> (\<lambda>x. x + 1) ` {x. 0 < Re x}}"
- then obtain y where h: "1 / x = y + 1" and hy: "0 < Re y" by blast
- hence hy': "y = 1 / x - 1" by fastforce
- hence hy'': "y + 1 \<noteq> 0" using h hy by fastforce
- hence "x = 1 / (y + 1)" using h
- by (metis div_by_1 divide_divide_eq_right mult.left_neutral)
- have "\<bar>Re y - 1\<bar> < \<bar>Re y + 1\<bar>" using hy by simp
- hence "cmod (y - 1) < cmod (y + 1)"
- by (smt (z3) cmod_Re_le_iff minus_complex.simps(1) minus_complex.simps(2)
- one_complex.simps plus_complex.simps(1) plus_complex.simps(2))
- hence "cmod ((y - 1)/(y + 1)) < 1"
- by (smt (verit) divide_less_eq_1_pos nonzero_norm_divide zero_less_norm_iff)
- thus "x \<in> circle_01_diam" using hy' hy''
- by (auto simp: field_simps norm_minus_commute circle_01_diam_def)
- qed
- show "circle_01_diam \<subseteq> {x. 1 / x \<in> (\<lambda>x. x + 1) ` {x. 0 < Re x}}"
- proof
- fix x assume "x \<in> circle_01_diam"
- hence "cmod (x - 1 / 2) * 2 < 1" by (auto simp: circle_01_diam_def)
- hence h: "x \<noteq> 0" and "cmod (x - 1 / 2) * cmod 2 < 1" by auto
- hence "cmod (2*x - 1) < 1"
- by (smt (verit) dbl_simps(3) dbl_simps(5) div_self times_divide_eq_left
- left_diff_distrib_numeral mult.commute mult_numeral_1
- norm_eq_zero norm_mult norm_numeral norm_one numeral_One)
- hence "cmod (((1/x - 1) - 1)/((1/x - 1) + 1)) < 1"
- by (auto simp: divide_simps norm_minus_commute)
- hence "cmod (((1/x - 1) - 1)/ cmod ((1/x - 1) + 1)) < 1"
- by (metis (no_types, lifting) abs_norm_cancel norm_divide norm_of_real)
- hence "cmod ((1/x - 1) - 1) < cmod ((1/x - 1) + 1)" using h
- by (smt (verit) diff_add_cancel divide_eq_0_iff divide_less_eq_1_pos
- norm_divide norm_of_real zero_less_norm_iff zero_neq_one)
- hence "\<bar>Re (1/x - 1) - 1\<bar> < \<bar>Re (1/x - 1) + 1\<bar>"
- by (smt (z3) cmod_Re_le_iff minus_complex.simps(1) minus_complex.simps(2)
- one_complex.simps plus_complex.simps(1) plus_complex.simps(2))
- hence "0 < Re (1/x - 1)" by linarith
- moreover have "1 / x = (1/x - 1) + 1" by simp
- ultimately have "0 < Re (1/x - 1) \<and> 1 / x = (1/x - 1) + 1" by blast
- hence "\<exists>xa. 0 < Re xa \<and> 1 / x = xa + 1" by blast
- thus "x \<in> {x. 1 / x \<in> (\<lambda>x. x + 1) ` {x. 0 < Re x}}" by blast
- qed
-qed
-
-lemma one_circle_01: fixes P::"real poly" assumes hP: "degree P \<le> p" and "P \<noteq> 0"
- and "proots_count (map_poly of_real P) circle_01_diam = 0"
-shows "Bernstein_changes_01 p P = 0"
-proof -
- let ?Q = "(reciprocal_poly p P) \<circ>\<^sub>p [:1, 1:]"
- have hQ: "?Q \<noteq> 0"
- using assms
- by (simp add: Missing_Polynomial.pcompose_eq_0 reciprocal_0_iff)
-
- hence 1: "changes (coeffs ?Q) \<ge> proots_count ?Q {x. 0 < x} \<and>
- even (changes (coeffs ?Q) - proots_count ?Q {x. 0 < x})"
- by (rule descartes_sign)
-
- have hdeg: "degree (map_poly complex_of_real P) \<le> p"
- by (rule le_trans, rule degree_map_poly_le, auto simp: assms)
- have hx: "\<And>x. 1 + x = 0 \<Longrightarrow> 0 < Re x \<Longrightarrow> False"
- proof -
- fix x::complex assume "1 + x = 0"
- hence "x = -1" by algebra
- thus "0 < Re x \<Longrightarrow> False" by auto
- qed
-
- have 2: "proots_count (map_poly of_real P) circle_01_diam =
- proots_count (map_poly of_real ?Q) {x. 0 < Re x}"
- apply (subst pos_real_map[symmetric])
- apply (subst of_real_hom.map_poly_pcompose)
- apply (subst map_poly_reciprocal) using assms apply auto[2]
- apply (subst proots_pcompose)
- using assms apply (auto simp: reciprocal_0_iff degree_map_poly)[2]
- apply (subst proots_count_reciprocal)
- using assms apply (auto simp: degree_map_poly inverse_eq_divide)[2]
- using hx apply fastforce
- by (auto simp: inverse_eq_divide algebra_simps)
-
- hence 3:"proots_count (map_poly of_real ?Q) {x. 0 < Re x} = 0"
- using assms(3) by presburger
-
- hence "\<And>x::complex.
- poly (map_poly of_real (smult (inverse (lead_coeff ?Q)) ?Q)) x = 0 \<Longrightarrow>
- Re x \<le> 0"
- proof cases
- fix x::complex show "Re x \<le> 0 \<Longrightarrow> Re x \<le> 0" by fast
- assume "\<not>Re x \<le> 0" hence h:"0 < Re x" by simp
- assume "poly (map_poly of_real (smult (inverse (lead_coeff ?Q)) ?Q)) x = 0"
- hence h2:"poly (map_poly of_real ?Q) x = 0" by fastforce
- hence "order x (map_poly complex_of_real (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:])) > 0"
- using assms by (fastforce simp: order_root pcompose_eq_0 reciprocal_0_iff)
- hence "proots_count (map_poly of_real ?Q) {x. 0 < Re x} \<noteq> 0"
- proof -
- have h3: "finite {x. poly (map_poly complex_of_real
- (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:])) x = 0}"
- apply (rule poly_roots_finite)
- using assms by (fastforce simp: order_root pcompose_eq_0 reciprocal_0_iff)
- have "0 < order x (map_poly complex_of_real (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:]))"
- using h2 assms by (fastforce simp: order_root pcompose_eq_0 reciprocal_0_iff)
- also have "... \<le> (\<Sum>r\<in>{x. 0 < Re x \<and>
- poly (map_poly complex_of_real (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:])) x =
- 0}.
- order r (map_poly complex_of_real (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:])))"
- apply (rule member_le_sum) using h h2 h3 by auto
- finally have
- "0 < (\<Sum>r\<in>{x. 0 < Re x \<and>
- poly (map_poly complex_of_real (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:])) x = 0}.
- order r (map_poly complex_of_real (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:])))" .
- thus
- "0 < order x (map_poly complex_of_real (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:])) \<Longrightarrow>
- proots_count (map_poly complex_of_real (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:]))
- {x. 0 < Re x} \<noteq> 0 "
- by (auto simp: proots_count_def proots_within_def)
- qed
- thus "Re x \<le> 0" using 3 by blast
- qed
- hence "\<And>i. coeff (smult (inverse (lead_coeff ?Q)) ?Q) i \<ge> 0"
- apply (frule descartes_sign_zero)
- using assms by (fastforce simp: pcompose_eq_0 reciprocal_0_iff)
- hence "changes (coeffs (smult (inverse (lead_coeff ?Q)) ?Q)) = 0"
- by (subst changes_all_nonneg, auto simp: nth_default_coeffs_eq)
- hence "changes (coeffs ?Q) = 0"
- using hQ by (auto simp: coeffs_smult changes_scale_const)
-
- thus ?thesis
- apply (subst Bernstein_changes_01_eq_changes["OF" hP])
- by blast
-qed
-
-definition circle_diam :: "real \<Rightarrow> real \<Rightarrow> complex set" where
-"circle_diam l r = {x. cmod ((x - l) - (r - l)/2) < (r - l)/2}"
-
-lemma circle_diam_rescale: assumes "l < r"
- shows "circle_diam l r = (\<lambda> x . (x*(r - l) + l)) ` circle_01_diam"
-proof
- show "circle_diam l r \<subseteq> (\<lambda>x. x * (complex_of_real r - complex_of_real l) +
- complex_of_real l) ` circle_01_diam"
- proof
- fix x assume "x \<in> circle_diam l r"
- hence "cmod ((x - l) - (r - l)/2) < (r - l)/2" by (auto simp: circle_diam_def)
- hence "cmod ((r - l) * ((x - l)/(r - l) - 1/2)) < (r - l)/2" using assms
- by (subst right_diff_distrib, fastforce)
- hence "(r - l) * cmod ((x - l)/(r - l) - 1/2) < (r - l) * 1/2"
- apply (subst(2) abs_of_pos[symmetric])
- subgoal using assms by argo
- subgoal
- apply (subst norm_scaleR[symmetric])
- by (simp add: scaleR_conv_of_real)
- done
- hence "cmod ((x - l)/(r - l) - 1/2) < 1/2"
- apply (subst mult_less_cancel_left_pos[of "r-l",symmetric])
- using assms by auto
- hence
- "cmod ((x-l)/(r-l) - 1 / 2) * 2 < 1 \<and>
- x = (x-l)/(r-l) * (complex_of_real r - complex_of_real l) + complex_of_real l"
- by force
- thus "x \<in> (\<lambda>x. x * (complex_of_real r - complex_of_real l) + complex_of_real l) `
- circle_01_diam"
- by (force simp: circle_01_diam_def)
- qed
- show "(\<lambda>x. x * (complex_of_real r - complex_of_real l) + complex_of_real l) `
- circle_01_diam \<subseteq> circle_diam l r"
- proof
- fix x::complex
- assume
- "x \<in> (\<lambda>x. x * (complex_of_real r - complex_of_real l) + complex_of_real l) `
- circle_01_diam"
- then obtain y::complex where "x = y * (r - l) + l" "cmod (y - 1/2) < 1/2"
- by (fastforce simp: circle_01_diam_def)
- moreover hence "y = (x - l) / (r - l)" using assms by force
- ultimately have "cmod ((x - l) / (r - l) - 1/2) < 1/2" by presburger
- hence "(r - l) * (cmod ((x - l) / (r - l) - 1/2)) < (r - l) * (1/2)"
- apply (subst mult_less_cancel_left_pos)
- using assms by auto
- hence "cmod ((x - l) - (r - l)/2) < (r - l)/2"
- apply (subst(asm) (2) abs_of_pos[symmetric])
- using assms apply argo
- apply (subst(asm) norm_scaleR[symmetric])
- by (smt (verit, del_insts)
- \<open>x = y * complex_of_real (r - l) + complex_of_real l\<close>
- \<open>y = (x - complex_of_real l) / complex_of_real (r - l)\<close>
- add_diff_cancel divide_divide_eq_right divide_numeral_1 mult.commute
- of_real_1 of_real_add of_real_divide one_add_one scaleR_conv_of_real
- scale_right_diff_distrib times_divide_eq_right)
- thus "x \<in> circle_diam l r"
- by (force simp: circle_diam_def)
- qed
-qed
-
-lemma one_circle: fixes P::"real poly" assumes "l < r"
- and "proots_count (map_poly of_real P) (circle_diam l r) = 0"
- and "P \<noteq> 0"
- and "degree P \<le> p"
-shows "Bernstein_changes p l r P = 0"
-proof (subst Bernstein_changes_eq_rescale)
- show "l \<noteq> r" using assms(1) by force
- show "degree P \<le> p" using assms(4) by blast
- show "Bernstein_changes_01 p (P \<circ>\<^sub>p [:l, 1:] \<circ>\<^sub>p [:0, r - l:]) = 0"
- proof (rule one_circle_01)
- show "degree (P \<circ>\<^sub>p [:l, 1:] \<circ>\<^sub>p [:0, r - l:]) \<le> p"
- using assms(4) by (force simp: degree_pcompose)
- show "P \<circ>\<^sub>p [:l, 1:] \<circ>\<^sub>p [:0, r - l:] \<noteq> 0"
- using assms by (smt (z3) degree_0_iff gr_zeroI pCons_eq_0_iff pCons_eq_iff
- pcompose_eq_0)
-
- have "proots_count (map_poly of_real P) (circle_diam l r) =
- proots_count (map_poly complex_of_real (P \<circ>\<^sub>p [:l, 1:] \<circ>\<^sub>p [:0, r - l:]))
- circle_01_diam"
- apply (subst of_real_hom.map_poly_pcompose)
- apply (subst proots_pcompose)
- apply (metis assms(3) degree_eq_zeroE of_real_poly_eq_0_iff
- pCons_eq_iff pCons_one pcompose_eq_0 zero_neq_one)
- using assms(1) apply fastforce
- apply (subst of_real_hom.map_poly_pcompose)
- apply (subst proots_pcompose)
- apply (auto simp: assms(3))[2]
- apply (subst circle_diam_rescale[OF assms(1)])
- apply (rule arg_cong[of _ _ "proots_count (map_poly complex_of_real P)"])
- by fastforce
-
- thus "proots_count (map_poly complex_of_real (P \<circ>\<^sub>p [:l, 1:] \<circ>\<^sub>p [:0, r - l:]))
- circle_01_diam = 0"
- using assms(2) by presburger
- qed
-qed
-
-subsection \<open>One sign change case\<close>
-
-definition upper_circle_01 :: "complex set" where
-"upper_circle_01 = {x. cmod (x - (1/2 + sqrt(3)/6 * \<i>)) < sqrt 3 / 3}"
-
-lemma upper_circle_map:
- "{x::complex. 1 / x \<in> (\<lambda>x. x + 1) ` {x. Im x < sqrt 3 * Re x}} = upper_circle_01"
-proof
- show "{x::complex. 1 / x \<in> (\<lambda>x. x + 1) ` {x. Im x < sqrt 3 * Re x}} \<subseteq> upper_circle_01"
- proof
- fix x
- assume "x \<in> {x. 1 / x \<in> (\<lambda>x. x + 1) ` {x. Im x < sqrt 3 * Re x}}"
- then obtain y where "1 / x = y + 1" and h: "Im y < sqrt 3 * Re y" by fastforce
- hence hy: "y = 1/x - 1" by simp
- hence hx: "x = 1/(y+1)" by auto
- from h have hy1: "y \<noteq> -1" by fastforce
- hence hx0: "x \<noteq> 0" using hy by fastforce
- from h have "0 < Re ((\<i> + sqrt 3) * y)" by fastforce
- hence "cmod ((\<i> + sqrt 3) * y - 1) < cmod ((\<i> + sqrt 3) * y + 1)"
- by (auto simp: cmod_def power2_eq_square algebra_simps)
- hence 1: "cmod (((\<i> + sqrt 3) * y - 1)/((\<i> + sqrt 3) * y + 1)) < 1"
- by (auto simp: norm_divide divide_simps)
- also have "cmod (((\<i> + sqrt 3) * y - 1)/((\<i> + sqrt 3) * y + 1)) =
- cmod (((\<i> + sqrt 3) * y * x - x)/((\<i> + sqrt 3) * y * x + x))"
- apply (subst mult_divide_mult_cancel_right[symmetric, OF hx0])
- apply (subst ring_distribs(2)[of _ _ x])
- apply (subst left_diff_distrib[of _ _ x])
- by simp
- also have "... = cmod
- (((-1 - complex_of_real (sqrt 3) - \<i>) * x + (complex_of_real (sqrt 3) + \<i>)) /
- (( 1 - complex_of_real (sqrt 3) - \<i>) * x + (complex_of_real (sqrt 3) + \<i>)))"
- by (auto simp: hy algebra_simps hx0)
-
- also have
- "... = cmod ((-1 - complex_of_real (sqrt 3) - \<i>) * x +
- (complex_of_real (sqrt 3) + \<i>)) /
- cmod (( 1 - complex_of_real (sqrt 3) - \<i>) * x +
- (complex_of_real (sqrt 3) + \<i>))"
- by (auto simp: norm_divide)
-
- finally have
- "cmod ((-1 - complex_of_real (sqrt 3) - \<i>) * x + (complex_of_real (sqrt 3) + \<i>))
- < cmod ((1 - complex_of_real (sqrt 3) - \<i>) * x + (complex_of_real (sqrt 3) + \<i>))"
- proof (subst divide_less_eq_1_pos[symmetric], subst zero_less_norm_iff)
- show "(1 - complex_of_real (sqrt 3) - \<i>) * x + (complex_of_real (sqrt 3) + \<i>) \<noteq> 0"
- proof
- have "-\<i> + 1 \<noteq> complex_of_real (sqrt 3)" by (auto simp: complex_eq_iff)
- moreover assume
- "(1 - complex_of_real (sqrt 3) - \<i>) * x + (complex_of_real (sqrt 3) + \<i>) = 0"
- ultimately have
- "x = (-complex_of_real (sqrt 3) - \<i>)/(1 - complex_of_real (sqrt 3) - \<i>)"
- by (auto simp: divide_simps algebra_simps)
- thus False
- using h by (auto simp: hy field_simps Im_divide Re_divide)
- qed
- qed
-
- hence "cmod (x - (1/2 + sqrt(3)/6 * \<i>)) < sqrt 3 / 3"
- apply (subst(3) abs_of_pos[symmetric, of 3]) apply auto[1]
- apply (subst real_sqrt_abs2[symmetric], subst real_sqrt_divide[symmetric])
- apply (subst cmod_def, subst real_sqrt_less_iff)
- apply (rule mult_right_less_imp_less[of _ "sqrt 3 /3"])
- by (auto simp: cmod_def power2_eq_square algebra_simps)
-
- thus "x \<in> upper_circle_01"
- by (auto simp: upper_circle_01_def)
- qed
-
- show "upper_circle_01 \<subseteq> {x. 1 / x \<in> (\<lambda>x. x + 1) ` {x. sqrt 3 * Re x > Im x}}"
- proof
- fix x assume "x \<in> upper_circle_01"
- hence "cmod (x - (1/2 + sqrt(3)/6 * \<i>)) < sqrt 3 / 3"
- by (force simp: upper_circle_01_def)
- hence "sqrt ((Re x - 1/2)^2 + (Im x - sqrt(3)/6)^2) < sqrt (1/3)"
- by (auto simp: cmod_def sqrt_divide_self_eq real_sqrt_inverse[symmetric])
- hence 1: "- Im x * sqrt 3 + (Im x * (Im x * 3) + Re x * (Re x * 3)) < Re x * 3"
- by (auto simp: power2_eq_square algebra_simps)
- have 2: "- Im x + (Im x * (Im x * sqrt 3) + Re x * (Re x * sqrt 3)) < Re x * sqrt 3"
- apply (rule mult_right_less_imp_less[of _ "sqrt 3"])
- apply (subst mult.assoc[of _ "sqrt 3"], subst real_sqrt_mult_self)
- using 1 by (auto simp: algebra_simps)
- have "sqrt 3 + (-Im x) / (Im x * Im x + Re x * Re x) <
- Re x * sqrt 3 / (Im x * Im x + Re x * Re x)"
- apply (rule mult_right_less_imp_less[of _ "(Im x * Im x + Re x * Re x)"])
- apply (rule subst, rule arg_cong2[of _ _ _ _ "(<)"])
- prefer 3 apply (rule 2)
- apply (subst distrib_right)
- using 2 apply auto
- by (auto simp: algebra_simps)
-
- hence "0 < - Im (1/x-1) + sqrt 3 * Re (1/x-1)"
- by (auto simp: power2_eq_square algebra_simps Re_divide Im_divide)
- hence "sqrt 3 * Re (1/x-1) > Im (1/x-1)"
- by argo
- hence "(1/x-1) \<in> {x. sqrt 3 * Re x > Im x}" by fast
- moreover have "1/x = (1/x-1) + 1" by simp
- ultimately show "x \<in> {x. 1 / x \<in> (\<lambda>x. x + 1) ` {x. sqrt 3 * Re x > Im x}}"
- by blast
- qed
-qed
-
-definition lower_circle_01 :: "complex set" where
-"lower_circle_01 = {x. cmod (x - (1/2 - sqrt(3)/6 * \<i>)) < sqrt 3 / 3}"
-
-lemma cnj_upper_circle_01: "cnj ` upper_circle_01 = lower_circle_01"
-proof
- show "cnj ` upper_circle_01 \<subseteq> lower_circle_01"
- proof
- fix x assume "x \<in> cnj ` upper_circle_01"
- then obtain y where "y \<in> upper_circle_01" and "x = cnj y" by blast
- thus "x \<in> lower_circle_01"
- apply (subst lower_circle_01_def, subst complex_mod_cnj[symmetric])
- by (auto simp add: upper_circle_01_def del: complex_mod_cnj)
- qed
- show "lower_circle_01 \<subseteq> cnj ` upper_circle_01"
- proof
- fix x assume "x \<in> lower_circle_01"
- hence "cnj x \<in> upper_circle_01" and "x = cnj (cnj x)"
- apply (subst upper_circle_01_def, subst complex_mod_cnj[symmetric])
- by (auto simp add: lower_circle_01_def del: complex_mod_cnj)
- thus "x \<in> cnj ` upper_circle_01"
- by blast
- qed
-qed
-
-lemma lower_circle_map:
- "{x::complex. 1 / x \<in> (\<lambda>x. x + 1) ` {x. Im x > - sqrt 3 * Re x}} = lower_circle_01"
-proof (subst cnj_upper_circle_01[symmetric], subst upper_circle_map[symmetric])
- have "{x. 1 / x \<in> (\<lambda>x. x + 1) ` {x. - sqrt 3 * Re x < Im x}} =
- {x. 1 / x \<in> (\<lambda>x. x + 1) ` {x. sqrt 3 * Re (cnj x) > Im (cnj x)}}"
- by auto
- also have "... = {x. 1 / x \<in> (\<lambda>x. x + 1) ` cnj ` {x. sqrt 3 * Re x > Im x}}"
- apply (subst(2) bij_image_Collect_eq)
- apply (metis bijI' complex_cnj_cnj)
- by (auto simp: inj_def inj_imp_inv_eq[of _ cnj])
- also have "... = {x. 1 / x \<in> cnj ` (\<lambda>x. x + 1) ` {x. sqrt 3 * Re x > Im x}}"
- by (auto simp: image_image)
- also have "... = {x. cnj (1 / x) \<in> (\<lambda>x. x + 1) ` {x. sqrt 3 * Re x > Im x}}"
- by (metis (no_types, lifting) complex_cnj_cnj image_iff)
- also have "... = cnj ` {x. 1 / x \<in> (\<lambda>x. x + 1) ` {x. sqrt 3 * Re x > Im x}}"
- apply (subst(2) bij_image_Collect_eq)
- apply (metis bijI' complex_cnj_cnj)
- by (auto simp: inj_def inj_imp_inv_eq[of _ cnj])
- finally show "{x. 1 / x \<in> (\<lambda>x. x + 1) ` {x. - sqrt 3 * Re x < Im x}} =
- cnj ` {x. 1 / x \<in> (\<lambda>x. x + 1) ` {x. Im x < sqrt 3 * Re x}}" .
-qed
-
-lemma two_circles_01:
- fixes P::"real poly"
- assumes hP: "degree P \<le> p" and hP0: "P \<noteq> 0" and hp0: "p \<noteq> 0"
- and h: "proots_count (map_poly of_real P)
- (upper_circle_01 \<union> lower_circle_01) = 1"
-shows "Bernstein_changes_01 p P = 1"
-proof (subst Bernstein_changes_01_eq_changes[OF hP])
- let ?Q = "reciprocal_poly p P \<circ>\<^sub>p [:1, 1:]"
- have hQ0: "?Q \<noteq> 0" using hP0
- by (simp add: pcompose_eq_0 hP reciprocal_0_iff)
-
- from h obtain x' where hroot': "poly (map_poly of_real P) x' = 0"
- and hx':"x' \<in> upper_circle_01 \<union> lower_circle_01"
- using proots_count_pos by (metis less_numeral_extra(1))
-
- obtain x where hxx': "x' = complex_of_real x"
- proof (cases "Im x' = 0")
- assume "Im x' = 0" and h: "\<And>x. x' = complex_of_real x \<Longrightarrow> thesis"
- then show thesis using h[of "Re x'"] by (simp add: complex_is_Real_iff)
- next
- assume hx'': "Im x' \<noteq> 0"
- have 1: "card {x', cnj x'} = 2"
- proof (subst card_2_iff)
- from hx'' have "x' \<noteq> cnj x'" and "{x', cnj x'} = {x', cnj x'}"
- by (metis cnj.simps(2) neg_equal_zero, argo)
- thus "\<exists>x y. {x', cnj x'} = {x, y} \<and> x \<noteq> y" by blast
- qed
- moreover have "{x', cnj x'} \<subseteq> upper_circle_01 \<union> lower_circle_01" using hx'
- apply (rule UnE)
- by (auto simp: cnj_upper_circle_01[symmetric])
- moreover have "\<And>x. x \<in> {x', cnj x'} \<Longrightarrow> poly (map_poly of_real P) x = 0"
- using hroot' poly_map_poly_of_real_cnj by auto
- ultimately have
- "proots_count (map_poly of_real P) (upper_circle_01 \<union> lower_circle_01) \<ge> 2"
- apply (subst 1[symmetric])
- apply (rule proots_count_of_root_set)
- using assms(2) of_real_poly_eq_0_iff by (blast, blast, blast)
- thus thesis using assms(4) by linarith
- qed
- hence hroot: "poly P x = 0"
- using hroot' by (metis of_real_0 of_real_eq_iff of_real_poly_map_poly)
- have that: "3 * sqrt (x * x + 1 / 3 - x) < sqrt 3" using hx'
- apply (rule UnE)
- by (auto simp: cmod_def power2_eq_square algebra_simps upper_circle_01_def
- lower_circle_01_def hxx')
- have hx: "0 < x \<and> x < 1"
- proof -
- have "3 * sqrt (x * x + 1 / 3 - x) = sqrt (9 * (x * x + 1 / 3 - x))"
- by (subst real_sqrt_mult, simp)
- hence "9 * (x * x + 1 / 3 - x) < 3" using that real_sqrt_less_iff by metis
- hence "x*x - x < 0" by auto
- thus "0 < x \<and> x < 1"
- using mult_eq_0_iff mult_less_cancel_right_disj by fastforce
- qed
-
- let ?y = "1/x - 1"
- from hroot hx assms have "poly ?Q ?y = 0"
- by (auto simp: poly_pcompose poly_reciprocal)
- hence "[:-?y, 1:] dvd ?Q" using poly_eq_0_iff_dvd by blast
- then obtain R where hR: "?Q = R * [:-?y, 1:]" by auto
- hence hR0: "R \<noteq> 0" using hQ0 by force
- interpret map_poly_idom_hom complex_of_real ..
-
-
- have "normal_poly (smult (inverse (lead_coeff (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:]))) R)"
- proof (rule normal_poly_of_roots)
- show "\<And>z. poly (map_poly complex_of_real
- (smult (inverse (lead_coeff (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:]))) R)) z = 0 \<Longrightarrow>
- Re z \<le> 0 \<and> (cmod z)\<^sup>2 \<le> 4 * (Re z)\<^sup>2"
- proof -
- fix z
- assume
- "poly (map_poly complex_of_real
- (smult (inverse (lead_coeff (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:]))) R)) z = 0"
- hence hroot2: "poly (map_poly complex_of_real R) z = 0"
- by (auto simp: map_poly_smult hQ0)
- show "Re z \<le> 0 \<and> (cmod z)\<^sup>2 \<le> 4 * (Re z)\<^sup>2"
- proof (rule ccontr)
- assume "\<not> (Re z \<le> 0 \<and> (cmod z)\<^sup>2 \<le> 4 * (Re z)\<^sup>2)"
- hence 1: "0 < Re z \<or> 4 * (Re z)\<^sup>2 < (cmod z)\<^sup>2" by linarith
- hence hz: "z \<noteq> -1" by force
- have "Im z > - sqrt 3 * Re z \<or> sqrt 3 * Re z > Im z"
- proof (cases "Im z \<ge> sqrt 3 * Re z", cases "- sqrt 3 * Re z \<ge> Im z")
- assume 2: "sqrt 3 * Re z \<le> Im z" "Im z \<le> - sqrt 3 * Re z"
- hence "sqrt 3 * Re z \<le> sqrt 3 * - Re z" by force
- hence "Re z \<le> - Re z"
- apply (rule mult_left_le_imp_le)
- by fastforce
- hence "Re z \<le> 0" by simp
- moreover have "(Im z)^2 \<le> (-sqrt 3 * Re z)^2"
- apply (subst power2_eq_square, subst power2_eq_square)
- apply (rule square_bounded_le)
- using 2 by auto
- ultimately have False using 1
- by (auto simp: power2_eq_square cmod_def algebra_simps)
- thus ?thesis by fast
- qed auto
-
- hence "z \<in> {z. - sqrt 3 * Re z < Im z} \<union> {z. Im z < sqrt 3 * Re z}"
- by blast
-
- hence 1: "inverse (1 + z) \<in> upper_circle_01 \<union> lower_circle_01"
- by (force simp: inverse_eq_divide upper_circle_map[symmetric]
- lower_circle_map[symmetric])
-
- have hRdeg': "degree R < p"
- apply (rule less_le_trans[of "degree R" "degree ?Q"])
- apply (subst hR, subst degree_mult_eq[OF hR0], fastforce, fastforce)
- by (auto simp: degree_pcompose degree_reciprocal hP)
- hence hRdeg: "degree R \<le> p" by fastforce
- have 2: "map_poly complex_of_real (reciprocal_poly p (R \<circ>\<^sub>p [:-1, 1:])) \<noteq> 0"
- apply (subst of_real_poly_eq_0_iff, subst reciprocal_0_iff)
- apply (force simp: hRdeg degree_pcompose)
- using hR0 pcompose_eq_0
- by (metis degree_eq_zeroE gr_zeroI pCons_eq_iff pCons_one zero_neq_one)
- have 3:
- "poly (map_poly complex_of_real (reciprocal_poly p (R \<circ>\<^sub>p [:-1, 1:])))
- (inverse (1 + z)) = 0"
- apply (subst map_poly_reciprocal)
- using hRdeg apply (force simp: degree_pcompose)
- apply auto[1]
- apply (subst poly_reciprocal)
- using hRdeg apply (force simp: degree_map_poly degree_pcompose)
- using hz apply (metis inverse_nonzero_iff_nonzero neg_eq_iff_add_eq_0)
- by (auto simp: of_real_hom.map_poly_pcompose poly_pcompose hroot2)
-
- have "proots_count (map_poly of_real (reciprocal_poly p (R \<circ>\<^sub>p [:-1, 1:])))
- (upper_circle_01 \<union> lower_circle_01) > 0"
- by (rule proots_count_of_root[OF 2 1 3])
- moreover have "proots_count
- (map_poly complex_of_real
- (reciprocal_poly p ([:1 - 1 / x, 1:] \<circ>\<^sub>p [:- 1, 1:])))
- (upper_circle_01 \<union> lower_circle_01) > 0"
- apply (subst map_poly_reciprocal)
- using hp0 less_eq_Suc_le apply (simp add: degree_pcompose)
- apply simp
- apply (subst proots_count_reciprocal)
- using hp0 less_eq_Suc_le
- apply (simp add: degree_pcompose degree_map_poly)
- apply (auto simp: pcompose_pCons)[1]
- apply (auto simp: cmod_def power2_eq_square real_sqrt_divide
- real_div_sqrt upper_circle_01_def lower_circle_01_def)[1]
- apply (subst of_real_hom.map_poly_pcompose)
- apply (subst proots_pcompose, fastforce, force)
- apply (subst lower_circle_map[symmetric])
- apply (subst upper_circle_map[symmetric])
- apply (rule proots_count_of_root[of _ "of_real (1/x - 1)"])
- apply simp
- apply (auto simp: bij_image_Collect_eq bij_def inj_def image_iff
- inverse_eq_divide inj_imp_inv_eq[of _ "\<lambda> x. x+1"] hx)[1]
- by force
-
- ultimately have "proots_count
- (map_poly complex_of_real (reciprocal_poly p (R \<circ>\<^sub>p [:- 1, 1:])))
- (upper_circle_01 \<union> lower_circle_01) +
- proots_count
- (map_poly complex_of_real
- (reciprocal_poly p ([:1 - 1 / x, 1:] \<circ>\<^sub>p [:- 1, 1:])))
- (upper_circle_01 \<union> lower_circle_01) > 1"
- by fastforce
- also have "... = proots_count (map_poly complex_of_real
- (monom 1 p * reciprocal_poly p (?Q \<circ>\<^sub>p [:- 1, 1:])))
- (upper_circle_01 \<union> lower_circle_01)"
- apply (subst eq_commute, subst hR, subst pcompose_mult)
- apply (subst reciprocal_mult, subst degree_mult_eq)
- using hR0 apply (fastforce simp: pcompose_eq_0)
- apply (fastforce simp: pcompose_pCons)
- using hRdeg' apply (simp add: degree_pcompose)
- using hRdeg apply (simp add: degree_pcompose)
- using hp0 apply (auto simp: degree_pcompose)[1]
- apply (subst hom_mult)
- apply (subst proots_count_times)
- using hp0 hRdeg' hR0
- apply (fastforce simp: reciprocal_0_iff degree_pcompose pcompose_eq_0
- pcompose_pCons)
- by simp
- also have "... = proots_count
- (map_poly complex_of_real
- (reciprocal_poly p (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:] \<circ>\<^sub>p [:- 1, 1:])))
- (upper_circle_01 \<union> lower_circle_01)"
- apply (subst hom_mult)
- apply (subst proots_count_times)
- using hp0 hP hP0
- apply (auto simp: map_poly_reciprocal degree_pcompose
- degree_reciprocal of_real_hom.map_poly_pcompose
- reciprocal_0_iff degree_map_poly pcompose_eq_0)[1]
- apply (subst map_poly_monom, fastforce)
- apply (subst of_real_1, subst proots_count_monom)
- apply (auto simp: cmod_def power2_eq_square real_sqrt_divide
- real_div_sqrt upper_circle_01_def lower_circle_01_def)[1]
- by presburger
- also have "... = 1"
- by (auto simp: pcompose_assoc["symmetric"] pcompose_pCons
- reciprocal_reciprocal hP h)
- finally show False by blast
- qed
- qed
- show "lead_coeff
- (smult (inverse (lead_coeff (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:]))) R) = 1"
- by (auto simp: hR degree_add_eq_right hR0 coeff_eq_0)
- qed
-
- hence "changes (coeffs (smult (inverse (lead_coeff ?Q)) ?Q)) = 1"
- apply (subst hR, subst mult_smult_left[symmetric], rule normal_changes)
- by (auto simp: hx)
-
- moreover have "changes (coeffs (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:])) =
- changes (coeffs (smult (inverse (lead_coeff (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:])))
- (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:])))"
- by (auto simp: pcompose_eq_0 reciprocal_0_iff hP hP0 coeffs_smult
- changes_scale_const[symmetric])
-
- ultimately show "changes (coeffs (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:])) = 1" by argo
-qed
-
-definition upper_circle :: "real \<Rightarrow> real \<Rightarrow> complex set" where
-"upper_circle l r = {x::complex.
- cmod ((x-of_real l)/(of_real (r-l)) - (1/2 + of_real (sqrt(3))/6 * \<i>)) < sqrt 3 / 3}"
-
-lemma upper_circle_rescale: assumes "l < r"
- shows "upper_circle l r = (\<lambda> x . (x*(r - l) + l)) ` upper_circle_01"
-proof
- show "upper_circle l r \<subseteq>
- (\<lambda>x. x * (of_real r - of_real l) + of_real l) ` upper_circle_01"
- apply (rule subsetI)
- apply (rule image_eqI[of _ _ "(_ - of_real l)/(of_real r - of_real l)"])
- using assms apply (auto simp: divide_simps)[1]
- by (auto simp: upper_circle_01_def upper_circle_def)
- show "(\<lambda>x. x * (of_real r - of_real l) + of_real l) ` upper_circle_01 \<subseteq>
- upper_circle l r"
- apply (rule subsetI, subst(asm) image_iff)
- using assms by (auto simp: upper_circle_01_def upper_circle_def)
-qed
-
-definition lower_circle :: "real \<Rightarrow> real \<Rightarrow> complex set" where
-"lower_circle l r = {x::complex.
- cmod ((x-of_real l)/(of_real (r-l)) - (1/2 - of_real (sqrt(3))/6 * \<i>)) < sqrt 3 / 3}"
-
-lemma lower_circle_rescale:
- assumes "l < r"
- shows "lower_circle l r = (\<lambda> x . (x*(r - l) + l)) ` lower_circle_01"
-proof
- show "lower_circle l r \<subseteq> (\<lambda>x. x * (of_real r - of_real l) + of_real l) `
- lower_circle_01"
- apply (rule subsetI)
- apply (rule image_eqI[of _ _ "(_ - of_real l)/(of_real r - of_real l)"])
- using assms apply (auto simp: divide_simps)[1]
- by (auto simp: lower_circle_01_def lower_circle_def)
- show "(\<lambda>x. x * (of_real r - of_real l) + of_real l) ` lower_circle_01 \<subseteq>
- lower_circle l r"
- apply (rule subsetI, subst(asm) image_iff)
- using assms by (auto simp: lower_circle_01_def lower_circle_def)
-qed
-
-lemma two_circles:
- fixes P::"real poly" and l r::real
- assumes hlr: "l < r"
- and hP: "degree P \<le> p"
- and hP0: "P \<noteq> 0"
- and hp0: "p \<noteq> 0"
- and h: "proots_count (map_poly of_real P)
- (upper_circle l r \<union> lower_circle l r) = 1"
-shows "Bernstein_changes p l r P = 1"
-proof -
- have 1: "Bernstein_changes p l r P =
- Bernstein_changes_01 p (P \<circ>\<^sub>p [:l, 1:] \<circ>\<^sub>p [:0, r - l:])"
- using assms by (force simp: Bernstein_changes_eq_rescale)
- have "proots_count (map_poly complex_of_real (P \<circ>\<^sub>p [:l, 1:] \<circ>\<^sub>p [:0, r - l:]))
- (upper_circle_01 \<union> lower_circle_01) = 1"
- using assms
- by (auto simp: upper_circle_rescale lower_circle_rescale proots_pcompose image_Un
- of_real_hom.map_poly_pcompose pcompose_eq_0 image_image algebra_simps)
- thus ?thesis
- apply (subst 1)
- apply (rule two_circles_01)
- using hP apply (force simp: degree_pcompose)
- using hP0 hlr apply (fastforce simp: pcompose_eq_0)
- using hp0 apply blast
- by blast
-qed
-
-subsection \<open>The theorem of three circles\<close>
-
-theorem three_circles:
- fixes P::"real poly" and l r::real
- assumes "l < r"
- and hP: "degree P \<le> p"
- and hP0: "P \<noteq> 0"
- and hp0: "p \<noteq> 0"
-shows "proots_count (map_poly of_real P) (circle_diam l r) = 0 \<Longrightarrow>
- Bernstein_changes p l r P = 0"
- and "proots_count (map_poly of_real P)
- (upper_circle l r \<union> lower_circle l r) = 1 \<Longrightarrow>
- Bernstein_changes p l r P = 1"
- apply (rule one_circle)
- using assms apply auto[4]
- apply (rule two_circles)
- using assms by auto
-
+section \<open>Proof of the theorem of three circles\<close>
+
+theory Three_Circles
+ imports "Bernstein" "Normal_Poly"
+begin
+
+text \<open>
+The theorem of three circles is a result in real algebraic geometry about the number of real roots
+in an interval. It says if the number of roots in certain circles in the complex plane are zero or
+one then the number of roots in the circles is equal to the sign changes of the Bernstein
+coefficients on that interval for which the circles intersect the real line. This can then be used
+to determine if an interval has a real root in the bisection procedure, which is more efficient than
+Descartes' rule of signs.
+
+The proof here follows Theorem 10.50 in
+ Basu, S., Pollack, R., Roy, M.-F.: Algorithms in Real Algebraic Geometry.
+ Springer Berlin Heidelberg, Berlin, Heidelberg (2016).
+
+This theorem has also been fomalised in Coq \cite{zsido2014theorem}. The relationship
+between this theorem and root isolation has been elaborated in Eigenwillig's PhD
+thesis \cite{eigenwillig2008real}.
+\<close>
+
+subsection \<open>No sign changes case\<close>
+
+declare degree_pcompose[simp del]
+
+corollary descartes_sign_zero:
+ fixes p::"real poly"
+ assumes "\<And>x::complex. poly (map_poly of_real p) x = 0 \<Longrightarrow> Re x \<le> 0"
+ and "lead_coeff p = 1"
+ shows "coeff p i \<ge> 0"
+ using assms
+proof (induction p arbitrary: i rule: real_poly_roots_induct)
+ case (1 p x)
+ interpret map_poly_idom_hom complex_of_real ..
+ have h: "\<And> i. 0 \<le> coeff p i"
+ apply (rule 1(1))
+ using 1(2) apply (metis lambda_zero of_real_poly_map_mult poly_mult)
+ using 1(3) apply (metis lead_coeff_1 lead_coeff_mult lead_coeff_pCons(1)
+ mult_cancel_right2 pCons_one zero_neq_one)
+ done
+ have "x \<le> 0"
+ apply (subst Re_complex_of_real[symmetric])
+ apply (rule 1(2))
+ apply (subst hom_mult)
+ by (auto)
+ thus ?case
+ apply (cases i)
+ subgoal using h[of i] h[of "i-1"]
+ by (fastforce simp: coeff_pCons mult_nonneg_nonpos2)
+ subgoal using h[of i] h[of "i-1"] mult_left_mono_neg
+ by (fastforce simp: coeff_pCons)
+ done
+next
+ case (2 p a b)
+ interpret map_poly_idom_hom complex_of_real ..
+ have h: "\<And> i. 0 \<le> coeff p i"
+ apply (rule 2(2))
+ using 2(3) apply (metis lambda_zero of_real_poly_map_mult poly_mult)
+ using 2(4) apply (metis lead_coeff_1 lead_coeff_mult lead_coeff_pCons(1)
+ mult_cancel_right2 pCons_one zero_neq_one)
+ done
+ have "Re (a + b * \<i>) \<le> 0"
+ apply (rule 2(3))
+ apply (subst hom_mult)
+ by (auto simp: algebra_simps)
+ hence 1: "0 \<le> - 2 * a" by fastforce
+ have 2: "0 \<le> a * a + b * b" by fastforce
+ have "\<And> x. 0 \<le> coeff [:a * a + b * b, - 2 * a, 1:] x"
+ proof -
+ fix x
+ show "0 \<le> coeff [:a * a + b * b, - 2 * a, 1:] x"
+ using 2 apply (cases "x = 0", fastforce)
+ using 1 apply (cases "x = 1", fastforce)
+ apply (cases "x = 2", fastforce simp: coeff_pCons)
+ by (auto simp: coeff_eq_0)
+ qed
+ thus ?case
+ apply (subst mult.commute, subst coeff_mult)
+ apply (rule sum_nonneg, rule mult_nonneg_nonneg[OF _ h])
+ by auto
+next
+ case (3 a)
+ then show ?case
+ by (smt (z3) bot_nat_0.extremum_uniqueI degree_1 le_degree
+ lead_coeff_pCons(2) pCons_one)
+qed
+
+definition circle_01_diam :: "complex set" where
+"circle_01_diam =
+ {x. cmod (x - (of_nat 1 :: complex)/(of_nat 2)) < (real 1)/(real 2)}"
+
+lemma pos_real_map:
+ "{x::complex. 1 / x \<in> (\<lambda>x. x + 1) ` {x. 0 < Re x}} = circle_01_diam"
+proof
+ show "{x. 1 / x \<in> (\<lambda>x. x + 1) ` {x. 0 < Re x}} \<subseteq> circle_01_diam"
+ proof
+ fix x assume "x \<in> {x. 1 / x \<in> (\<lambda>x. x + 1) ` {x. 0 < Re x}}"
+ then obtain y where h: "1 / x = y + 1" and hy: "0 < Re y" by blast
+ hence hy': "y = 1 / x - 1" by fastforce
+ hence hy'': "y + 1 \<noteq> 0" using h hy by fastforce
+ hence "x = 1 / (y + 1)" using h
+ by (metis div_by_1 divide_divide_eq_right mult.left_neutral)
+ have "\<bar>Re y - 1\<bar> < \<bar>Re y + 1\<bar>" using hy by simp
+ hence "cmod (y - 1) < cmod (y + 1)"
+ by (smt (z3) cmod_Re_le_iff minus_complex.simps(1) minus_complex.simps(2)
+ one_complex.simps plus_complex.simps(1) plus_complex.simps(2))
+ hence "cmod ((y - 1)/(y + 1)) < 1"
+ by (smt (verit) divide_less_eq_1_pos nonzero_norm_divide zero_less_norm_iff)
+ thus "x \<in> circle_01_diam" using hy' hy''
+ by (auto simp: field_simps norm_minus_commute circle_01_diam_def)
+ qed
+ show "circle_01_diam \<subseteq> {x. 1 / x \<in> (\<lambda>x. x + 1) ` {x. 0 < Re x}}"
+ proof
+ fix x assume "x \<in> circle_01_diam"
+ hence "cmod (x - 1 / 2) * 2 < 1" by (auto simp: circle_01_diam_def)
+ hence h: "x \<noteq> 0" and "cmod (x - 1 / 2) * cmod 2 < 1" by auto
+ hence "cmod (2*x - 1) < 1"
+ by (smt (verit) dbl_simps(3) dbl_simps(5) div_self times_divide_eq_left
+ left_diff_distrib_numeral mult.commute mult_numeral_1
+ norm_eq_zero norm_mult norm_numeral norm_one numeral_One)
+ hence "cmod (((1/x - 1) - 1)/((1/x - 1) + 1)) < 1"
+ by (auto simp: divide_simps norm_minus_commute)
+ hence "cmod (((1/x - 1) - 1)/ cmod ((1/x - 1) + 1)) < 1"
+ by (metis (no_types, lifting) abs_norm_cancel norm_divide norm_of_real)
+ hence "cmod ((1/x - 1) - 1) < cmod ((1/x - 1) + 1)" using h
+ by (smt (verit) diff_add_cancel divide_eq_0_iff divide_less_eq_1_pos
+ norm_divide norm_of_real zero_less_norm_iff zero_neq_one)
+ hence "\<bar>Re (1/x - 1) - 1\<bar> < \<bar>Re (1/x - 1) + 1\<bar>"
+ by (smt (z3) cmod_Re_le_iff minus_complex.simps(1) minus_complex.simps(2)
+ one_complex.simps plus_complex.simps(1) plus_complex.simps(2))
+ hence "0 < Re (1/x - 1)" by linarith
+ moreover have "1 / x = (1/x - 1) + 1" by simp
+ ultimately have "0 < Re (1/x - 1) \<and> 1 / x = (1/x - 1) + 1" by blast
+ hence "\<exists>xa. 0 < Re xa \<and> 1 / x = xa + 1" by blast
+ thus "x \<in> {x. 1 / x \<in> (\<lambda>x. x + 1) ` {x. 0 < Re x}}" by blast
+ qed
+qed
+
+lemma one_circle_01: fixes P::"real poly" assumes hP: "degree P \<le> p" and "P \<noteq> 0"
+ and "proots_count (map_poly of_real P) circle_01_diam = 0"
+shows "Bernstein_changes_01 p P = 0"
+proof -
+ let ?Q = "(reciprocal_poly p P) \<circ>\<^sub>p [:1, 1:]"
+ have hQ: "?Q \<noteq> 0"
+ using assms
+ by (simp add: Missing_Polynomial.pcompose_eq_0 reciprocal_0_iff)
+
+ hence 1: "changes (coeffs ?Q) \<ge> proots_count ?Q {x. 0 < x} \<and>
+ even (changes (coeffs ?Q) - proots_count ?Q {x. 0 < x})"
+ by (rule descartes_sign)
+
+ have hdeg: "degree (map_poly complex_of_real P) \<le> p"
+ by (rule le_trans, rule degree_map_poly_le, auto simp: assms)
+ have hx: "\<And>x. 1 + x = 0 \<Longrightarrow> 0 < Re x \<Longrightarrow> False"
+ proof -
+ fix x::complex assume "1 + x = 0"
+ hence "x = -1" by algebra
+ thus "0 < Re x \<Longrightarrow> False" by auto
+ qed
+
+ have 2: "proots_count (map_poly of_real P) circle_01_diam =
+ proots_count (map_poly of_real ?Q) {x. 0 < Re x}"
+ apply (subst pos_real_map[symmetric])
+ apply (subst of_real_hom.map_poly_pcompose)
+ apply (subst map_poly_reciprocal) using assms apply auto[2]
+ apply (subst proots_pcompose)
+ using assms apply (auto simp: reciprocal_0_iff degree_map_poly)[2]
+ apply (subst proots_count_reciprocal)
+ using assms apply (auto simp: degree_map_poly inverse_eq_divide)[2]
+ using hx apply fastforce
+ by (auto simp: inverse_eq_divide algebra_simps)
+
+ hence 3:"proots_count (map_poly of_real ?Q) {x. 0 < Re x} = 0"
+ using assms(3) by presburger
+
+ hence "\<And>x::complex.
+ poly (map_poly of_real (smult (inverse (lead_coeff ?Q)) ?Q)) x = 0 \<Longrightarrow>
+ Re x \<le> 0"
+ proof cases
+ fix x::complex show "Re x \<le> 0 \<Longrightarrow> Re x \<le> 0" by fast
+ assume "\<not>Re x \<le> 0" hence h:"0 < Re x" by simp
+ assume "poly (map_poly of_real (smult (inverse (lead_coeff ?Q)) ?Q)) x = 0"
+ hence h2:"poly (map_poly of_real ?Q) x = 0" by fastforce
+ hence "order x (map_poly complex_of_real (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:])) > 0"
+ using assms by (fastforce simp: order_root pcompose_eq_0 reciprocal_0_iff)
+ hence "proots_count (map_poly of_real ?Q) {x. 0 < Re x} \<noteq> 0"
+ proof -
+ have h3: "finite {x. poly (map_poly complex_of_real
+ (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:])) x = 0}"
+ apply (rule poly_roots_finite)
+ using assms by (fastforce simp: order_root pcompose_eq_0 reciprocal_0_iff)
+ have "0 < order x (map_poly complex_of_real (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:]))"
+ using h2 assms by (fastforce simp: order_root pcompose_eq_0 reciprocal_0_iff)
+ also have "... \<le> (\<Sum>r\<in>{x. 0 < Re x \<and>
+ poly (map_poly complex_of_real (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:])) x =
+ 0}.
+ order r (map_poly complex_of_real (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:])))"
+ apply (rule member_le_sum) using h h2 h3 by auto
+ finally have
+ "0 < (\<Sum>r\<in>{x. 0 < Re x \<and>
+ poly (map_poly complex_of_real (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:])) x = 0}.
+ order r (map_poly complex_of_real (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:])))" .
+ thus
+ "0 < order x (map_poly complex_of_real (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:])) \<Longrightarrow>
+ proots_count (map_poly complex_of_real (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:]))
+ {x. 0 < Re x} \<noteq> 0 "
+ by (auto simp: proots_count_def proots_within_def)
+ qed
+ thus "Re x \<le> 0" using 3 by blast
+ qed
+ hence "\<And>i. coeff (smult (inverse (lead_coeff ?Q)) ?Q) i \<ge> 0"
+ apply (frule descartes_sign_zero)
+ using assms by (fastforce simp: pcompose_eq_0 reciprocal_0_iff)
+ hence "changes (coeffs (smult (inverse (lead_coeff ?Q)) ?Q)) = 0"
+ by (subst changes_all_nonneg, auto simp: nth_default_coeffs_eq)
+ hence "changes (coeffs ?Q) = 0"
+ using hQ by (auto simp: coeffs_smult changes_scale_const)
+
+ thus ?thesis
+ apply (subst Bernstein_changes_01_eq_changes["OF" hP])
+ by blast
+qed
+
+definition circle_diam :: "real \<Rightarrow> real \<Rightarrow> complex set" where
+"circle_diam l r = {x. cmod ((x - l) - (r - l)/2) < (r - l)/2}"
+
+lemma circle_diam_rescale: assumes "l < r"
+ shows "circle_diam l r = (\<lambda> x . (x*(r - l) + l)) ` circle_01_diam"
+proof
+ show "circle_diam l r \<subseteq> (\<lambda>x. x * (complex_of_real r - complex_of_real l) +
+ complex_of_real l) ` circle_01_diam"
+ proof
+ fix x assume "x \<in> circle_diam l r"
+ hence "cmod ((x - l) - (r - l)/2) < (r - l)/2" by (auto simp: circle_diam_def)
+ hence "cmod ((r - l) * ((x - l)/(r - l) - 1/2)) < (r - l)/2" using assms
+ by (subst right_diff_distrib, fastforce)
+ hence "(r - l) * cmod ((x - l)/(r - l) - 1/2) < (r - l) * 1/2"
+ apply (subst(2) abs_of_pos[symmetric])
+ subgoal using assms by argo
+ subgoal
+ apply (subst norm_scaleR[symmetric])
+ by (simp add: scaleR_conv_of_real)
+ done
+ hence "cmod ((x - l)/(r - l) - 1/2) < 1/2"
+ apply (subst mult_less_cancel_left_pos[of "r-l",symmetric])
+ using assms by auto
+ hence
+ "cmod ((x-l)/(r-l) - 1 / 2) * 2 < 1 \<and>
+ x = (x-l)/(r-l) * (complex_of_real r - complex_of_real l) + complex_of_real l"
+ by force
+ thus "x \<in> (\<lambda>x. x * (complex_of_real r - complex_of_real l) + complex_of_real l) `
+ circle_01_diam"
+ by (force simp: circle_01_diam_def)
+ qed
+ show "(\<lambda>x. x * (complex_of_real r - complex_of_real l) + complex_of_real l) `
+ circle_01_diam \<subseteq> circle_diam l r"
+ proof
+ fix x::complex
+ assume
+ "x \<in> (\<lambda>x. x * (complex_of_real r - complex_of_real l) + complex_of_real l) `
+ circle_01_diam"
+ then obtain y::complex where "x = y * (r - l) + l" "cmod (y - 1/2) < 1/2"
+ by (fastforce simp: circle_01_diam_def)
+ moreover hence "y = (x - l) / (r - l)" using assms by force
+ ultimately have "cmod ((x - l) / (r - l) - 1/2) < 1/2" by presburger
+ hence "(r - l) * (cmod ((x - l) / (r - l) - 1/2)) < (r - l) * (1/2)"
+ apply (subst mult_less_cancel_left_pos)
+ using assms by auto
+ hence "cmod ((x - l) - (r - l)/2) < (r - l)/2"
+ apply (subst(asm) (2) abs_of_pos[symmetric])
+ using assms apply argo
+ apply (subst(asm) norm_scaleR[symmetric])
+ by (smt (verit, del_insts)
+ \<open>x = y * complex_of_real (r - l) + complex_of_real l\<close>
+ \<open>y = (x - complex_of_real l) / complex_of_real (r - l)\<close>
+ add_diff_cancel divide_divide_eq_right divide_numeral_1 mult.commute
+ of_real_1 of_real_add of_real_divide one_add_one scaleR_conv_of_real
+ scale_right_diff_distrib times_divide_eq_right)
+ thus "x \<in> circle_diam l r"
+ by (force simp: circle_diam_def)
+ qed
+qed
+
+lemma one_circle: fixes P::"real poly" assumes "l < r"
+ and "proots_count (map_poly of_real P) (circle_diam l r) = 0"
+ and "P \<noteq> 0"
+ and "degree P \<le> p"
+shows "Bernstein_changes p l r P = 0"
+proof (subst Bernstein_changes_eq_rescale)
+ show "l \<noteq> r" using assms(1) by force
+ show "degree P \<le> p" using assms(4) by blast
+ show "Bernstein_changes_01 p (P \<circ>\<^sub>p [:l, 1:] \<circ>\<^sub>p [:0, r - l:]) = 0"
+ proof (rule one_circle_01)
+ show "degree (P \<circ>\<^sub>p [:l, 1:] \<circ>\<^sub>p [:0, r - l:]) \<le> p"
+ using assms(4) by (force simp: degree_pcompose)
+ show "P \<circ>\<^sub>p [:l, 1:] \<circ>\<^sub>p [:0, r - l:] \<noteq> 0"
+ using assms by (smt (z3) degree_0_iff gr_zeroI pCons_eq_0_iff pCons_eq_iff
+ pcompose_eq_0)
+
+ have "proots_count (map_poly of_real P) (circle_diam l r) =
+ proots_count (map_poly complex_of_real (P \<circ>\<^sub>p [:l, 1:] \<circ>\<^sub>p [:0, r - l:]))
+ circle_01_diam"
+ apply (subst of_real_hom.map_poly_pcompose)
+ apply (subst proots_pcompose)
+ apply (metis assms(3) degree_eq_zeroE of_real_poly_eq_0_iff
+ pCons_eq_iff pCons_one pcompose_eq_0 zero_neq_one)
+ using assms(1) apply fastforce
+ apply (subst of_real_hom.map_poly_pcompose)
+ apply (subst proots_pcompose)
+ apply (auto simp: assms(3))[2]
+ apply (subst circle_diam_rescale[OF assms(1)])
+ apply (rule arg_cong[of _ _ "proots_count (map_poly complex_of_real P)"])
+ by fastforce
+
+ thus "proots_count (map_poly complex_of_real (P \<circ>\<^sub>p [:l, 1:] \<circ>\<^sub>p [:0, r - l:]))
+ circle_01_diam = 0"
+ using assms(2) by presburger
+ qed
+qed
+
+subsection \<open>One sign change case\<close>
+
+definition upper_circle_01 :: "complex set" where
+"upper_circle_01 = {x. cmod (x - (1/2 + sqrt(3)/6 * \<i>)) < sqrt 3 / 3}"
+
+lemma upper_circle_map:
+ "{x::complex. 1 / x \<in> (\<lambda>x. x + 1) ` {x. Im x < sqrt 3 * Re x}} = upper_circle_01"
+proof
+ show "{x::complex. 1 / x \<in> (\<lambda>x. x + 1) ` {x. Im x < sqrt 3 * Re x}} \<subseteq> upper_circle_01"
+ proof
+ fix x
+ assume "x \<in> {x. 1 / x \<in> (\<lambda>x. x + 1) ` {x. Im x < sqrt 3 * Re x}}"
+ then obtain y where "1 / x = y + 1" and h: "Im y < sqrt 3 * Re y" by fastforce
+ hence hy: "y = 1/x - 1" by simp
+ hence hx: "x = 1/(y+1)" by auto
+ from h have hy1: "y \<noteq> -1" by fastforce
+ hence hx0: "x \<noteq> 0" using hy by fastforce
+ from h have "0 < Re ((\<i> + sqrt 3) * y)" by fastforce
+ hence "cmod ((\<i> + sqrt 3) * y - 1) < cmod ((\<i> + sqrt 3) * y + 1)"
+ by (auto simp: cmod_def power2_eq_square algebra_simps)
+ hence 1: "cmod (((\<i> + sqrt 3) * y - 1)/((\<i> + sqrt 3) * y + 1)) < 1"
+ by (auto simp: norm_divide divide_simps)
+ also have "cmod (((\<i> + sqrt 3) * y - 1)/((\<i> + sqrt 3) * y + 1)) =
+ cmod (((\<i> + sqrt 3) * y * x - x)/((\<i> + sqrt 3) * y * x + x))"
+ apply (subst mult_divide_mult_cancel_right[symmetric, OF hx0])
+ apply (subst ring_distribs(2)[of _ _ x])
+ apply (subst left_diff_distrib[of _ _ x])
+ by simp
+ also have "... = cmod
+ (((-1 - complex_of_real (sqrt 3) - \<i>) * x + (complex_of_real (sqrt 3) + \<i>)) /
+ (( 1 - complex_of_real (sqrt 3) - \<i>) * x + (complex_of_real (sqrt 3) + \<i>)))"
+ by (auto simp: hy algebra_simps hx0)
+
+ also have
+ "... = cmod ((-1 - complex_of_real (sqrt 3) - \<i>) * x +
+ (complex_of_real (sqrt 3) + \<i>)) /
+ cmod (( 1 - complex_of_real (sqrt 3) - \<i>) * x +
+ (complex_of_real (sqrt 3) + \<i>))"
+ by (auto simp: norm_divide)
+
+ finally have
+ "cmod ((-1 - complex_of_real (sqrt 3) - \<i>) * x + (complex_of_real (sqrt 3) + \<i>))
+ < cmod ((1 - complex_of_real (sqrt 3) - \<i>) * x + (complex_of_real (sqrt 3) + \<i>))"
+ proof (subst divide_less_eq_1_pos[symmetric], subst zero_less_norm_iff)
+ show "(1 - complex_of_real (sqrt 3) - \<i>) * x + (complex_of_real (sqrt 3) + \<i>) \<noteq> 0"
+ proof
+ have "-\<i> + 1 \<noteq> complex_of_real (sqrt 3)" by (auto simp: complex_eq_iff)
+ moreover assume
+ "(1 - complex_of_real (sqrt 3) - \<i>) * x + (complex_of_real (sqrt 3) + \<i>) = 0"
+ ultimately have
+ "x = (-complex_of_real (sqrt 3) - \<i>)/(1 - complex_of_real (sqrt 3) - \<i>)"
+ by (auto simp: divide_simps algebra_simps)
+ thus False
+ using h by (auto simp: hy field_simps Im_divide Re_divide)
+ qed
+ qed
+
+ hence "cmod (x - (1/2 + sqrt(3)/6 * \<i>)) < sqrt 3 / 3"
+ apply (subst(3) abs_of_pos[symmetric, of 3]) apply auto[1]
+ apply (subst real_sqrt_abs2[symmetric], subst real_sqrt_divide[symmetric])
+ apply (subst cmod_def, subst real_sqrt_less_iff)
+ apply (rule mult_right_less_imp_less[of _ "sqrt 3 /3"])
+ by (auto simp: cmod_def power2_eq_square algebra_simps)
+
+ thus "x \<in> upper_circle_01"
+ by (auto simp: upper_circle_01_def)
+ qed
+
+ show "upper_circle_01 \<subseteq> {x. 1 / x \<in> (\<lambda>x. x + 1) ` {x. sqrt 3 * Re x > Im x}}"
+ proof
+ fix x assume "x \<in> upper_circle_01"
+ hence "cmod (x - (1/2 + sqrt(3)/6 * \<i>)) < sqrt 3 / 3"
+ by (force simp: upper_circle_01_def)
+ hence "sqrt ((Re x - 1/2)^2 + (Im x - sqrt(3)/6)^2) < sqrt (1/3)"
+ by (auto simp: cmod_def sqrt_divide_self_eq real_sqrt_inverse[symmetric])
+ hence 1: "- Im x * sqrt 3 + (Im x * (Im x * 3) + Re x * (Re x * 3)) < Re x * 3"
+ by (auto simp: power2_eq_square algebra_simps)
+ have 2: "- Im x + (Im x * (Im x * sqrt 3) + Re x * (Re x * sqrt 3)) < Re x * sqrt 3"
+ apply (rule mult_right_less_imp_less[of _ "sqrt 3"])
+ apply (subst mult.assoc[of _ "sqrt 3"], subst real_sqrt_mult_self)
+ using 1 by (auto simp: algebra_simps)
+ have "sqrt 3 + (-Im x) / (Im x * Im x + Re x * Re x) <
+ Re x * sqrt 3 / (Im x * Im x + Re x * Re x)"
+ apply (rule mult_right_less_imp_less[of _ "(Im x * Im x + Re x * Re x)"])
+ apply (rule subst, rule arg_cong2[of _ _ _ _ "(<)"])
+ prefer 3 apply (rule 2)
+ apply (subst distrib_right)
+ using 2 apply auto
+ by (auto simp: algebra_simps)
+
+ hence "0 < - Im (1/x-1) + sqrt 3 * Re (1/x-1)"
+ by (auto simp: power2_eq_square algebra_simps Re_divide Im_divide)
+ hence "sqrt 3 * Re (1/x-1) > Im (1/x-1)"
+ by argo
+ hence "(1/x-1) \<in> {x. sqrt 3 * Re x > Im x}" by fast
+ moreover have "1/x = (1/x-1) + 1" by simp
+ ultimately show "x \<in> {x. 1 / x \<in> (\<lambda>x. x + 1) ` {x. sqrt 3 * Re x > Im x}}"
+ by blast
+ qed
+qed
+
+definition lower_circle_01 :: "complex set" where
+"lower_circle_01 = {x. cmod (x - (1/2 - sqrt(3)/6 * \<i>)) < sqrt 3 / 3}"
+
+lemma cnj_upper_circle_01: "cnj ` upper_circle_01 = lower_circle_01"
+proof
+ show "cnj ` upper_circle_01 \<subseteq> lower_circle_01"
+ proof
+ fix x assume "x \<in> cnj ` upper_circle_01"
+ then obtain y where "y \<in> upper_circle_01" and "x = cnj y" by blast
+ thus "x \<in> lower_circle_01"
+ apply (subst lower_circle_01_def, subst complex_mod_cnj[symmetric])
+ by (auto simp add: upper_circle_01_def del: complex_mod_cnj)
+ qed
+ show "lower_circle_01 \<subseteq> cnj ` upper_circle_01"
+ proof
+ fix x assume "x \<in> lower_circle_01"
+ hence "cnj x \<in> upper_circle_01" and "x = cnj (cnj x)"
+ apply (subst upper_circle_01_def, subst complex_mod_cnj[symmetric])
+ by (auto simp add: lower_circle_01_def del: complex_mod_cnj)
+ thus "x \<in> cnj ` upper_circle_01"
+ by blast
+ qed
+qed
+
+lemma lower_circle_map:
+ "{x::complex. 1 / x \<in> (\<lambda>x. x + 1) ` {x. Im x > - sqrt 3 * Re x}} = lower_circle_01"
+proof (subst cnj_upper_circle_01[symmetric], subst upper_circle_map[symmetric])
+ have "{x. 1 / x \<in> (\<lambda>x. x + 1) ` {x. - sqrt 3 * Re x < Im x}} =
+ {x. 1 / x \<in> (\<lambda>x. x + 1) ` {x. sqrt 3 * Re (cnj x) > Im (cnj x)}}"
+ by auto
+ also have "... = {x. 1 / x \<in> (\<lambda>x. x + 1) ` cnj ` {x. sqrt 3 * Re x > Im x}}"
+ apply (subst(2) bij_image_Collect_eq)
+ apply (metis bijI' complex_cnj_cnj)
+ by (auto simp: inj_def inj_imp_inv_eq[of _ cnj])
+ also have "... = {x. 1 / x \<in> cnj ` (\<lambda>x. x + 1) ` {x. sqrt 3 * Re x > Im x}}"
+ by (auto simp: image_image)
+ also have "... = {x. cnj (1 / x) \<in> (\<lambda>x. x + 1) ` {x. sqrt 3 * Re x > Im x}}"
+ by (metis (no_types, lifting) complex_cnj_cnj image_iff)
+ also have "... = cnj ` {x. 1 / x \<in> (\<lambda>x. x + 1) ` {x. sqrt 3 * Re x > Im x}}"
+ apply (subst(2) bij_image_Collect_eq)
+ apply (metis bijI' complex_cnj_cnj)
+ by (auto simp: inj_def inj_imp_inv_eq[of _ cnj])
+ finally show "{x. 1 / x \<in> (\<lambda>x. x + 1) ` {x. - sqrt 3 * Re x < Im x}} =
+ cnj ` {x. 1 / x \<in> (\<lambda>x. x + 1) ` {x. Im x < sqrt 3 * Re x}}" .
+qed
+
+lemma two_circles_01:
+ fixes P::"real poly"
+ assumes hP: "degree P \<le> p" and hP0: "P \<noteq> 0" and hp0: "p \<noteq> 0"
+ and h: "proots_count (map_poly of_real P)
+ (upper_circle_01 \<union> lower_circle_01) = 1"
+shows "Bernstein_changes_01 p P = 1"
+proof (subst Bernstein_changes_01_eq_changes[OF hP])
+ let ?Q = "reciprocal_poly p P \<circ>\<^sub>p [:1, 1:]"
+ have hQ0: "?Q \<noteq> 0" using hP0
+ by (simp add: pcompose_eq_0 hP reciprocal_0_iff)
+
+ from h obtain x' where hroot': "poly (map_poly of_real P) x' = 0"
+ and hx':"x' \<in> upper_circle_01 \<union> lower_circle_01"
+ using proots_count_pos by (metis less_numeral_extra(1))
+
+ obtain x where hxx': "x' = complex_of_real x"
+ proof (cases "Im x' = 0")
+ assume "Im x' = 0" and h: "\<And>x. x' = complex_of_real x \<Longrightarrow> thesis"
+ then show thesis using h[of "Re x'"] by (simp add: complex_is_Real_iff)
+ next
+ assume hx'': "Im x' \<noteq> 0"
+ have 1: "card {x', cnj x'} = 2"
+ proof (subst card_2_iff)
+ from hx'' have "x' \<noteq> cnj x'" and "{x', cnj x'} = {x', cnj x'}"
+ by (metis cnj.simps(2) neg_equal_zero, argo)
+ thus "\<exists>x y. {x', cnj x'} = {x, y} \<and> x \<noteq> y" by blast
+ qed
+ moreover have "{x', cnj x'} \<subseteq> upper_circle_01 \<union> lower_circle_01" using hx'
+ apply (rule UnE)
+ by (auto simp: cnj_upper_circle_01[symmetric])
+ moreover have "\<And>x. x \<in> {x', cnj x'} \<Longrightarrow> poly (map_poly of_real P) x = 0"
+ using hroot' poly_map_poly_of_real_cnj by auto
+ ultimately have
+ "proots_count (map_poly of_real P) (upper_circle_01 \<union> lower_circle_01) \<ge> 2"
+ apply (subst 1[symmetric])
+ apply (rule proots_count_of_root_set)
+ using assms(2) of_real_poly_eq_0_iff by (blast, blast, blast)
+ thus thesis using assms(4) by linarith
+ qed
+ hence hroot: "poly P x = 0"
+ using hroot' by (metis of_real_0 of_real_eq_iff of_real_poly_map_poly)
+ have that: "3 * sqrt (x * x + 1 / 3 - x) < sqrt 3" using hx'
+ apply (rule UnE)
+ by (auto simp: cmod_def power2_eq_square algebra_simps upper_circle_01_def
+ lower_circle_01_def hxx')
+ have hx: "0 < x \<and> x < 1"
+ proof -
+ have "3 * sqrt (x * x + 1 / 3 - x) = sqrt (9 * (x * x + 1 / 3 - x))"
+ by (subst real_sqrt_mult, simp)
+ hence "9 * (x * x + 1 / 3 - x) < 3" using that real_sqrt_less_iff by metis
+ hence "x*x - x < 0" by auto
+ thus "0 < x \<and> x < 1"
+ using mult_eq_0_iff mult_less_cancel_right_disj by fastforce
+ qed
+
+ let ?y = "1/x - 1"
+ from hroot hx assms have "poly ?Q ?y = 0"
+ by (auto simp: poly_pcompose poly_reciprocal)
+ hence "[:-?y, 1:] dvd ?Q" using poly_eq_0_iff_dvd by blast
+ then obtain R where hR: "?Q = R * [:-?y, 1:]" by auto
+ hence hR0: "R \<noteq> 0" using hQ0 by force
+ interpret map_poly_idom_hom complex_of_real ..
+
+
+ have "normal_poly (smult (inverse (lead_coeff (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:]))) R)"
+ proof (rule normal_poly_of_roots)
+ show "\<And>z. poly (map_poly complex_of_real
+ (smult (inverse (lead_coeff (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:]))) R)) z = 0 \<Longrightarrow>
+ Re z \<le> 0 \<and> (cmod z)\<^sup>2 \<le> 4 * (Re z)\<^sup>2"
+ proof -
+ fix z
+ assume
+ "poly (map_poly complex_of_real
+ (smult (inverse (lead_coeff (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:]))) R)) z = 0"
+ hence hroot2: "poly (map_poly complex_of_real R) z = 0"
+ by (auto simp: map_poly_smult hQ0)
+ show "Re z \<le> 0 \<and> (cmod z)\<^sup>2 \<le> 4 * (Re z)\<^sup>2"
+ proof (rule ccontr)
+ assume "\<not> (Re z \<le> 0 \<and> (cmod z)\<^sup>2 \<le> 4 * (Re z)\<^sup>2)"
+ hence 1: "0 < Re z \<or> 4 * (Re z)\<^sup>2 < (cmod z)\<^sup>2" by linarith
+ hence hz: "z \<noteq> -1" by force
+ have "Im z > - sqrt 3 * Re z \<or> sqrt 3 * Re z > Im z"
+ proof (cases "Im z \<ge> sqrt 3 * Re z", cases "- sqrt 3 * Re z \<ge> Im z")
+ assume 2: "sqrt 3 * Re z \<le> Im z" "Im z \<le> - sqrt 3 * Re z"
+ hence "sqrt 3 * Re z \<le> sqrt 3 * - Re z" by force
+ hence "Re z \<le> - Re z"
+ apply (rule mult_left_le_imp_le)
+ by fastforce
+ hence "Re z \<le> 0" by simp
+ moreover have "(Im z)^2 \<le> (-sqrt 3 * Re z)^2"
+ apply (subst power2_eq_square, subst power2_eq_square)
+ apply (rule square_bounded_le)
+ using 2 by auto
+ ultimately have False using 1
+ by (auto simp: power2_eq_square cmod_def algebra_simps)
+ thus ?thesis by fast
+ qed auto
+
+ hence "z \<in> {z. - sqrt 3 * Re z < Im z} \<union> {z. Im z < sqrt 3 * Re z}"
+ by blast
+
+ hence 1: "inverse (1 + z) \<in> upper_circle_01 \<union> lower_circle_01"
+ by (force simp: inverse_eq_divide upper_circle_map[symmetric]
+ lower_circle_map[symmetric])
+
+ have hRdeg': "degree R < p"
+ apply (rule less_le_trans[of "degree R" "degree ?Q"])
+ apply (subst hR, subst degree_mult_eq[OF hR0], fastforce, fastforce)
+ by (auto simp: degree_pcompose degree_reciprocal hP)
+ hence hRdeg: "degree R \<le> p" by fastforce
+ have 2: "map_poly complex_of_real (reciprocal_poly p (R \<circ>\<^sub>p [:-1, 1:])) \<noteq> 0"
+ apply (subst of_real_poly_eq_0_iff, subst reciprocal_0_iff)
+ apply (force simp: hRdeg degree_pcompose)
+ using hR0 pcompose_eq_0
+ by (metis degree_eq_zeroE gr_zeroI pCons_eq_iff pCons_one zero_neq_one)
+ have 3:
+ "poly (map_poly complex_of_real (reciprocal_poly p (R \<circ>\<^sub>p [:-1, 1:])))
+ (inverse (1 + z)) = 0"
+ apply (subst map_poly_reciprocal)
+ using hRdeg apply (force simp: degree_pcompose)
+ apply auto[1]
+ apply (subst poly_reciprocal)
+ using hRdeg apply (force simp: degree_map_poly degree_pcompose)
+ using hz apply (metis inverse_nonzero_iff_nonzero neg_eq_iff_add_eq_0)
+ by (auto simp: of_real_hom.map_poly_pcompose poly_pcompose hroot2)
+
+ have "proots_count (map_poly of_real (reciprocal_poly p (R \<circ>\<^sub>p [:-1, 1:])))
+ (upper_circle_01 \<union> lower_circle_01) > 0"
+ by (rule proots_count_of_root[OF 2 1 3])
+ moreover have "proots_count
+ (map_poly complex_of_real
+ (reciprocal_poly p ([:1 - 1 / x, 1:] \<circ>\<^sub>p [:- 1, 1:])))
+ (upper_circle_01 \<union> lower_circle_01) > 0"
+ apply (subst map_poly_reciprocal)
+ using hp0 less_eq_Suc_le apply (simp add: degree_pcompose)
+ apply simp
+ apply (subst proots_count_reciprocal)
+ using hp0 less_eq_Suc_le
+ apply (simp add: degree_pcompose degree_map_poly)
+ apply (auto simp: pcompose_pCons)[1]
+ apply (auto simp: cmod_def power2_eq_square real_sqrt_divide
+ real_div_sqrt upper_circle_01_def lower_circle_01_def)[1]
+ apply (subst of_real_hom.map_poly_pcompose)
+ apply (subst proots_pcompose, fastforce, force)
+ apply (subst lower_circle_map[symmetric])
+ apply (subst upper_circle_map[symmetric])
+ apply (rule proots_count_of_root[of _ "of_real (1/x - 1)"])
+ apply simp
+ apply (auto simp: bij_image_Collect_eq bij_def inj_def image_iff
+ inverse_eq_divide inj_imp_inv_eq[of _ "\<lambda> x. x+1"] hx)[1]
+ by force
+
+ ultimately have "proots_count
+ (map_poly complex_of_real (reciprocal_poly p (R \<circ>\<^sub>p [:- 1, 1:])))
+ (upper_circle_01 \<union> lower_circle_01) +
+ proots_count
+ (map_poly complex_of_real
+ (reciprocal_poly p ([:1 - 1 / x, 1:] \<circ>\<^sub>p [:- 1, 1:])))
+ (upper_circle_01 \<union> lower_circle_01) > 1"
+ by fastforce
+ also have "... = proots_count (map_poly complex_of_real
+ (monom 1 p * reciprocal_poly p (?Q \<circ>\<^sub>p [:- 1, 1:])))
+ (upper_circle_01 \<union> lower_circle_01)"
+ apply (subst eq_commute, subst hR, subst pcompose_mult)
+ apply (subst reciprocal_mult, subst degree_mult_eq)
+ using hR0 apply (fastforce simp: pcompose_eq_0)
+ apply (fastforce simp: pcompose_pCons)
+ using hRdeg' apply (simp add: degree_pcompose)
+ using hRdeg apply (simp add: degree_pcompose)
+ using hp0 apply (auto simp: degree_pcompose)[1]
+ apply (subst hom_mult)
+ apply (subst proots_count_times)
+ using hp0 hRdeg' hR0
+ apply (fastforce simp: reciprocal_0_iff degree_pcompose pcompose_eq_0
+ pcompose_pCons)
+ by simp
+ also have "... = proots_count
+ (map_poly complex_of_real
+ (reciprocal_poly p (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:] \<circ>\<^sub>p [:- 1, 1:])))
+ (upper_circle_01 \<union> lower_circle_01)"
+ apply (subst hom_mult)
+ apply (subst proots_count_times)
+ using hp0 hP hP0
+ apply (auto simp: map_poly_reciprocal degree_pcompose
+ degree_reciprocal of_real_hom.map_poly_pcompose
+ reciprocal_0_iff degree_map_poly pcompose_eq_0)[1]
+ apply (subst map_poly_monom, fastforce)
+ apply (subst of_real_1, subst proots_count_monom)
+ apply (auto simp: cmod_def power2_eq_square real_sqrt_divide
+ real_div_sqrt upper_circle_01_def lower_circle_01_def)[1]
+ by presburger
+ also have "... = 1"
+ by (auto simp: pcompose_assoc["symmetric"] pcompose_pCons
+ reciprocal_reciprocal hP h)
+ finally show False by blast
+ qed
+ qed
+ show "lead_coeff
+ (smult (inverse (lead_coeff (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:]))) R) = 1"
+ by (auto simp: hR degree_add_eq_right hR0 coeff_eq_0)
+ qed
+
+ hence "changes (coeffs (smult (inverse (lead_coeff ?Q)) ?Q)) = 1"
+ apply (subst hR, subst mult_smult_left[symmetric], rule normal_changes)
+ by (auto simp: hx)
+
+ moreover have "changes (coeffs (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:])) =
+ changes (coeffs (smult (inverse (lead_coeff (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:])))
+ (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:])))"
+ by (auto simp: pcompose_eq_0 reciprocal_0_iff hP hP0 coeffs_smult
+ changes_scale_const[symmetric])
+
+ ultimately show "changes (coeffs (reciprocal_poly p P \<circ>\<^sub>p [:1, 1:])) = 1" by argo
+qed
+
+definition upper_circle :: "real \<Rightarrow> real \<Rightarrow> complex set" where
+"upper_circle l r = {x::complex.
+ cmod ((x-of_real l)/(of_real (r-l)) - (1/2 + of_real (sqrt(3))/6 * \<i>)) < sqrt 3 / 3}"
+
+lemma upper_circle_rescale: assumes "l < r"
+ shows "upper_circle l r = (\<lambda> x . (x*(r - l) + l)) ` upper_circle_01"
+proof
+ show "upper_circle l r \<subseteq>
+ (\<lambda>x. x * (of_real r - of_real l) + of_real l) ` upper_circle_01"
+ apply (rule subsetI)
+ apply (rule image_eqI[of _ _ "(_ - of_real l)/(of_real r - of_real l)"])
+ using assms apply (auto simp: divide_simps)[1]
+ by (auto simp: upper_circle_01_def upper_circle_def)
+ show "(\<lambda>x. x * (of_real r - of_real l) + of_real l) ` upper_circle_01 \<subseteq>
+ upper_circle l r"
+ apply (rule subsetI, subst(asm) image_iff)
+ using assms by (auto simp: upper_circle_01_def upper_circle_def)
+qed
+
+definition lower_circle :: "real \<Rightarrow> real \<Rightarrow> complex set" where
+"lower_circle l r = {x::complex.
+ cmod ((x-of_real l)/(of_real (r-l)) - (1/2 - of_real (sqrt(3))/6 * \<i>)) < sqrt 3 / 3}"
+
+lemma lower_circle_rescale:
+ assumes "l < r"
+ shows "lower_circle l r = (\<lambda> x . (x*(r - l) + l)) ` lower_circle_01"
+proof
+ show "lower_circle l r \<subseteq> (\<lambda>x. x * (of_real r - of_real l) + of_real l) `
+ lower_circle_01"
+ apply (rule subsetI)
+ apply (rule image_eqI[of _ _ "(_ - of_real l)/(of_real r - of_real l)"])
+ using assms apply (auto simp: divide_simps)[1]
+ by (auto simp: lower_circle_01_def lower_circle_def)
+ show "(\<lambda>x. x * (of_real r - of_real l) + of_real l) ` lower_circle_01 \<subseteq>
+ lower_circle l r"
+ apply (rule subsetI, subst(asm) image_iff)
+ using assms by (auto simp: lower_circle_01_def lower_circle_def)
+qed
+
+lemma two_circles:
+ fixes P::"real poly" and l r::real
+ assumes hlr: "l < r"
+ and hP: "degree P \<le> p"
+ and hP0: "P \<noteq> 0"
+ and hp0: "p \<noteq> 0"
+ and h: "proots_count (map_poly of_real P)
+ (upper_circle l r \<union> lower_circle l r) = 1"
+shows "Bernstein_changes p l r P = 1"
+proof -
+ have 1: "Bernstein_changes p l r P =
+ Bernstein_changes_01 p (P \<circ>\<^sub>p [:l, 1:] \<circ>\<^sub>p [:0, r - l:])"
+ using assms by (force simp: Bernstein_changes_eq_rescale)
+ have "proots_count (map_poly complex_of_real (P \<circ>\<^sub>p [:l, 1:] \<circ>\<^sub>p [:0, r - l:]))
+ (upper_circle_01 \<union> lower_circle_01) = 1"
+ using assms
+ by (auto simp: upper_circle_rescale lower_circle_rescale proots_pcompose image_Un
+ of_real_hom.map_poly_pcompose pcompose_eq_0 image_image algebra_simps)
+ thus ?thesis
+ apply (subst 1)
+ apply (rule two_circles_01)
+ using hP apply (force simp: degree_pcompose)
+ using hP0 hlr apply (fastforce simp: pcompose_eq_0)
+ using hp0 apply blast
+ by blast
+qed
+
+subsection \<open>The theorem of three circles\<close>
+
+theorem three_circles:
+ fixes P::"real poly" and l r::real
+ assumes "l < r"
+ and hP: "degree P \<le> p"
+ and hP0: "P \<noteq> 0"
+ and hp0: "p \<noteq> 0"
+shows "proots_count (map_poly of_real P) (circle_diam l r) = 0 \<Longrightarrow>
+ Bernstein_changes p l r P = 0"
+ and "proots_count (map_poly of_real P)
+ (upper_circle l r \<union> lower_circle l r) = 1 \<Longrightarrow>
+ Bernstein_changes p l r P = 1"
+ apply (rule one_circle)
+ using assms apply auto[4]
+ apply (rule two_circles)
+ using assms by auto
+
end
\ No newline at end of file
diff --git a/thys/Three_Circles/document/root.bib b/thys/Three_Circles/document/root.bib
--- a/thys/Three_Circles/document/root.bib
+++ b/thys/Three_Circles/document/root.bib
@@ -1,42 +1,42 @@
-@article{eigenwillig2008real,
- title={Real root isolation for exact and approximate polynomials using Descartes' rule of signs},
- author={Eigenwillig, Arno},
- year={2008}
-}
-
-@inproceedings{li2019counting,
- title={Counting polynomial roots in Isabelle/HOL: a formal proof of the Budan-Fourier theorem},
- author={Li, Wenda and Paulson, Lawrence C},
- booktitle={Proceedings of the 8th ACM SIGPLAN International Conference on Certified Programs and Proofs},
- pages={52--64},
- year={2019}
-}
-
-@article{zsido2014theorem,
- title={Theorem of three circles in Coq},
- author={Zsid{\'o}, Julianna},
- journal={Journal of automated reasoning},
- volume={53},
- number={2},
- pages={105--127},
- year={2014},
- publisher={Springer}
-}
-
-@book{Basu:2016bo,
- author = {Basu, Saugata and Pollack, Richard and Roy, Marie-Franco̧ise},
- title = {{Algorithms in Real Algebraic Geometry}},
- publisher = {Springer Berlin Heidelberg},
- year = {2016},
- volume = {10},
- series = {Algorithms and Computation in Mathematics},
- address = {Berlin, Heidelberg}
-}
-
-
-
-
-
-
-
-
+@article{eigenwillig2008real,
+ title={Real root isolation for exact and approximate polynomials using Descartes' rule of signs},
+ author={Eigenwillig, Arno},
+ year={2008}
+}
+
+@inproceedings{li2019counting,
+ title={Counting polynomial roots in Isabelle/HOL: a formal proof of the Budan-Fourier theorem},
+ author={Li, Wenda and Paulson, Lawrence C},
+ booktitle={Proceedings of the 8th ACM SIGPLAN International Conference on Certified Programs and Proofs},
+ pages={52--64},
+ year={2019}
+}
+
+@article{zsido2014theorem,
+ title={Theorem of three circles in Coq},
+ author={Zsid{\'o}, Julianna},
+ journal={Journal of automated reasoning},
+ volume={53},
+ number={2},
+ pages={105--127},
+ year={2014},
+ publisher={Springer}
+}
+
+@book{Basu:2016bo,
+ author = {Basu, Saugata and Pollack, Richard and Roy, Marie-Franco̧ise},
+ title = {{Algorithms in Real Algebraic Geometry}},
+ publisher = {Springer Berlin Heidelberg},
+ year = {2016},
+ volume = {10},
+ series = {Algorithms and Computation in Mathematics},
+ address = {Berlin, Heidelberg}
+}
+
+
+
+
+
+
+
+
diff --git a/thys/Three_Circles/document/root.tex b/thys/Three_Circles/document/root.tex
--- a/thys/Three_Circles/document/root.tex
+++ b/thys/Three_Circles/document/root.tex
@@ -1,75 +1,75 @@
-\documentclass[11pt,a4paper]{article}
-\usepackage{isabelle,isabellesym}
-
-% further packages required for unusual symbols (see also
-% isabellesym.sty), use only when needed
-
-%\usepackage{amssymb}
- %for \<leadsto>, \<box>, \<diamond>, \<sqsupset>, \<mho>, \<Join>,
- %\<lhd>, \<lesssim>, \<greatersim>, \<lessapprox>, \<greaterapprox>,
- %\<triangleq>, \<yen>, \<lozenge>
-
-%\usepackage{eurosym}
- %for \<euro>
-
-%\usepackage[only,bigsqcap]{stmaryrd}
- %for \<Sqinter>
-
-%\usepackage{eufrak}
- %for \<AA> ... \<ZZ>, \<aa> ... \<zz> (also included in amssymb)
-
-%\usepackage{textcomp}
- %for \<onequarter>, \<onehalf>, \<threequarters>, \<degree>, \<cent>,
- %\<currency>
-
-% this should be the last package used
-\usepackage{pdfsetup}
-
-% urls in roman style, theory text in math-similar italics
-\urlstyle{rm}
-\isabellestyle{it}
-
-% for uniform font size
-%\renewcommand{\isastyle}{\isastyleminor}
-
-
-\begin{document}
-
-\title{The Theorem of Three Circles}
-\author{Fox Thomson, Wenda Li}
-\maketitle
-
-
-\begin{abstract}
- The Descartes test based on Bernstein coefficients and Descartes' rule of signs
- effectively (over-)approximates the number of real roots of a univariate polynomial over
- an interval. In this entry we formalise the theorem of three circles (Theorem 10.50 in \cite{Basu:2016bo}),
- which gives sufficient conditions for when the Descartes test
- returns 0 or 1. This is the first step for efficient root isolation.
-\end{abstract}
-
-\tableofcontents
-
-
-% sane default for proof documents
-\parindent 0pt\parskip 0.5ex
-
-% generated text of all theories
-\input{session}
-
-
-\section{Acknowledgements}
-
-The work has been jointly supported by the Cambridge Mathematics Placements (CMP)
-Programme and the ERC Advanced Grant ALEXANDRIA (Project GA 742178).
-
-% optional bibliography
-\bibliographystyle{abbrv}
-\bibliography{root}
-
-\end{document}
-
-%%% Local Variables:
-%%% mode: latex
-%%% TeX-master: t
-%%% End:
+\documentclass[11pt,a4paper]{article}
+\usepackage{isabelle,isabellesym}
+
+% further packages required for unusual symbols (see also
+% isabellesym.sty), use only when needed
+
+%\usepackage{amssymb}
+ %for \<leadsto>, \<box>, \<diamond>, \<sqsupset>, \<mho>, \<Join>,
+ %\<lhd>, \<lesssim>, \<greatersim>, \<lessapprox>, \<greaterapprox>,
+ %\<triangleq>, \<yen>, \<lozenge>
+
+%\usepackage{eurosym}
+ %for \<euro>
+
+%\usepackage[only,bigsqcap]{stmaryrd}
+ %for \<Sqinter>
+
+%\usepackage{eufrak}
+ %for \<AA> ... \<ZZ>, \<aa> ... \<zz> (also included in amssymb)
+
+%\usepackage{textcomp}
+ %for \<onequarter>, \<onehalf>, \<threequarters>, \<degree>, \<cent>,
+ %\<currency>
+
+% this should be the last package used
+\usepackage{pdfsetup}
+
+% urls in roman style, theory text in math-similar italics
+\urlstyle{rm}
+\isabellestyle{it}
+
+% for uniform font size
+%\renewcommand{\isastyle}{\isastyleminor}
+
+
+\begin{document}
+
+\title{The Theorem of Three Circles}
+\author{Fox Thomson, Wenda Li}
+\maketitle
+
+
+\begin{abstract}
+ The Descartes test based on Bernstein coefficients and Descartes' rule of signs
+ effectively (over-)approximates the number of real roots of a univariate polynomial over
+ an interval. In this entry we formalise the theorem of three circles (Theorem 10.50 in \cite{Basu:2016bo}),
+ which gives sufficient conditions for when the Descartes test
+ returns 0 or 1. This is the first step for efficient root isolation.
+\end{abstract}
+
+\tableofcontents
+
+
+% sane default for proof documents
+\parindent 0pt\parskip 0.5ex
+
+% generated text of all theories
+\input{session}
+
+
+\section{Acknowledgements}
+
+The work has been jointly supported by the Cambridge Mathematics Placements (CMP)
+Programme and the ERC Advanced Grant ALEXANDRIA (Project GA 742178).
+
+% optional bibliography
+\bibliographystyle{abbrv}
+\bibliography{root}
+
+\end{document}
+
+%%% Local Variables:
+%%% mode: latex
+%%% TeX-master: t
+%%% End:
diff --git a/thys/Weighted_Path_Order/KBO_as_WPO.thy b/thys/Weighted_Path_Order/KBO_as_WPO.thy
--- a/thys/Weighted_Path_Order/KBO_as_WPO.thy
+++ b/thys/Weighted_Path_Order/KBO_as_WPO.thy
@@ -1,421 +1,421 @@
-subsection \<open>A restricted equality between KBO and WPO\<close>
-
-text \<open>The remaining difficulty to make KBO an instance of WPO is the
- different treatment of lexicographic comparisons, which is unrestricted in KBO,
- but there is a length-restriction in WPO.
- Therefore we will only show that KBO is an instance of WPO if we compare terms with
- bounded arity.\<close>
-
-text \<open>This restriction does however not prohibit us from lifting properties of WPO to KBO.
- For instance, for several properties one can choose a large-enough bound restriction of WPO,
- since there are only finitely many arities occurring in a property.\<close>
-
-theory KBO_as_WPO
- imports
- WPO
- KBO_Transformation
-begin
-
-definition bounded_arity :: "nat \<Rightarrow> ('f \<times> nat)set \<Rightarrow> bool" where
- "bounded_arity b F = (\<forall> (f,n) \<in> F. n \<le> b)"
-
-lemma finite_funas_term[simp,intro]: "finite (funas_term t)"
- by (induct t, auto)
-
-context weight_fun begin
-
-definition "weight_le s t \<equiv>
- (vars_term_ms (SCF s) \<subseteq># vars_term_ms (SCF t) \<and> weight s \<le> weight t)"
-
-definition "weight_less s t \<equiv>
- (vars_term_ms (SCF s) \<subseteq># vars_term_ms (SCF t) \<and> weight s < weight t)"
-
-lemma weight_le_less_iff: "weight_le s t \<Longrightarrow> weight_less s t \<longleftrightarrow> weight s < weight t"
- by (auto simp: weight_le_def weight_less_def)
-
-lemma weight_less_iff: "weight_less s t \<Longrightarrow> weight_le s t \<and> weight s < weight t"
- by (auto simp: weight_le_def weight_less_def)
-
-
-abbreviation "weight_NS \<equiv> {(t,s). weight_le s t}"
-
-abbreviation "weight_S \<equiv> {(t,s). weight_less s t}"
-
-lemma weight_le_mono_one:
- assumes S: "weight_le s t"
- shows "weight_le (Fun f (ss1 @ s # ss2)) (Fun f (ss1 @ t # ss2))" (is "weight_le ?s ?t")
-proof -
- from S have w: "weight s \<le> weight t" and v: "vars_term_ms (SCF s) \<subseteq># vars_term_ms (SCF t)"
- by (auto simp: weight_le_def)
- have v': "vars_term_ms (SCF ?s) \<subseteq># vars_term_ms (SCF ?t)"
- using mset_replicate_mono[OF v] by simp
- have w': "weight ?s \<le> weight ?t" using sum_list_replicate_mono[OF w] by simp
- from v' w' show ?thesis by (auto simp: weight_le_def)
-qed
-
-lemma weight_le_ctxt: "weight_le s t \<Longrightarrow> weight_le (C\<langle>s\<rangle>) (C\<langle>t\<rangle>)"
- by (induct C, auto intro: weight_le_mono_one)
-
-lemma SCF_stable:
- assumes "vars_term_ms (SCF s) \<subseteq># vars_term_ms (SCF t)"
- shows "vars_term_ms (SCF (s \<cdot> \<sigma>)) \<subseteq># vars_term_ms (SCF (t \<cdot> \<sigma>))"
- unfolding scf_term_subst
- using vars_term_ms_subst_mono[OF assms].
-
-lemma SN_weight_S: "SN weight_S"
-proof-
- from wf_inv_image[OF wf_less]
- have wf: "wf {(s,t). weight s < weight t}" by (auto simp: inv_image_def)
- show ?thesis
- by (unfold SN_iff_wf, rule wf_subset[OF wf], auto simp: weight_less_def)
-qed
-
-lemma weight_less_imp_le: "weight_less s t \<Longrightarrow> weight_le s t" by (simp add: weight_less_def weight_le_def)
-
-lemma weight_le_Var_Var: "weight_le (Var x) (Var y) \<longleftrightarrow> x = y"
- by (auto simp: weight_le_def)
-end
-
-context kbo begin
-
-lemma kbo_altdef:
- "kbo s t = (if weight_le t s
- then if weight_less t s
- then (True, True)
- else (case s of
- Var y \<Rightarrow> (False, (case t of Var x \<Rightarrow> x = y | Fun g ts \<Rightarrow> ts = [] \<and> least g))
- | Fun f ss \<Rightarrow> (case t of
- Var x \<Rightarrow> (True, True)
- | Fun g ts \<Rightarrow> if pr_strict (f, length ss) (g, length ts)
- then (True, True)
- else if pr_weak (f, length ss) (g, length ts)
- then lex_ext_unbounded kbo ss ts
- else (False, False)))
- else (False, False))"
- by (simp add: weight_le_less_iff weight_le_def)
-
-end
-
-context admissible_kbo begin
-
-lemma weight_le_stable:
- assumes "weight_le s t"
- shows "weight_le (s \<cdot> \<sigma>) (t \<cdot> \<sigma>)"
- using assms weight_stable_le SCF_stable by (auto simp: weight_le_def)
-
-lemma weight_less_stable:
- assumes "weight_less s t"
- shows "weight_less (s \<cdot> \<sigma>) (t \<cdot> \<sigma>)"
- using assms weight_stable_lt SCF_stable by (auto simp: weight_less_def)
-
-lemma simple_arg_pos_weight: "simple_arg_pos weight_NS (f,n) i"
- unfolding simple_arg_pos_def
-proof (intro allI impI, unfold snd_conv fst_conv)
- fix ts :: "('f,'a)term list"
- assume i: "i < n" and len: "length ts = n"
- from id_take_nth_drop[OF i[folded len]] i[folded len]
- obtain us vs where id: "Fun f ts = Fun f (us @ ts ! i # vs)"
- and us: "us = take i ts"
- and len: "length us = i" by auto
- have "length us < Suc (length us + length vs)" by auto
- from scf[OF this, of f] obtain j where [simp]: "scf (f, Suc (length us + length vs)) (length us) = Suc j"
- by (rule lessE)
- show "(Fun f ts, ts ! i) \<in> weight_NS"
- unfolding weight_le_def id by (auto simp: o_def)
-qed
-
-lemma weight_lemmas:
- shows "refl weight_NS" and "trans weight_NS" and "trans weight_S"
- and "weight_NS O weight_S \<subseteq> weight_S" and "weight_S O weight_NS \<subseteq> weight_S"
- by (auto intro!: refl_onI transI simp: weight_le_def weight_less_def)
-
-interpretation kbo': admissible_kbo w w0 pr_strict' pr_weak' least scf
- by (rule admissible_kbo')
-
-context
- assumes least_global: "\<And> f g. least f \<Longrightarrow> pr_weak g (f,0)"
- and least_trans: "\<And> f g. least f \<Longrightarrow> pr_weak (f,0) g \<Longrightarrow> least (fst g) \<and> snd g = 0"
- fixes n :: nat
-begin
-
-lemma kbo_instance_of_wpo_with_assms: "wpo_with_assms
- weight_S weight_NS (\<lambda>f g. (pr_strict f g, pr_weak f g))
- (\<lambda>(f, n). n = 0 \<and> least f) full_status False (\<lambda>f. False)"
- apply (unfold_locales)
- apply (auto simp: weight_lemmas SN_weight_S pr_SN pr_strict_irrefl
- weight_less_stable weight_le_stable weight_le_mono_one weight_less_imp_le
- simple_arg_pos_weight)
- apply (force dest: least_global least_trans simp: pr_strict)+
- apply (auto simp: pr_strict least dest:pr_weak_trans)
- done
-
-interpretation wpo: wpo_with_assms
- where S = weight_S and NS = weight_NS
- and prc = "\<lambda>f g. (pr_strict f g, pr_weak f g)" and prl = "\<lambda>(f,n). n = 0 \<and> least f"
- and c = "\<lambda>_. Lex"
- and ssimple = False and large = "\<lambda>f. False" and \<sigma>\<sigma> = full_status
- and n = n
- by (rule kbo_instance_of_wpo_with_assms)
-
-lemma kbo_as_wpo_with_assms: assumes "bounded_arity n (funas_term t)"
- shows "kbo s t = wpo.wpo s t"
-proof -
- define m where "m = size s + size t"
- from m_def assms show ?thesis
- proof (induct m arbitrary: s t rule: less_induct)
- case (less m s t)
- hence IH: "size si + size ti < size s + size t \<Longrightarrow> bounded_arity n (funas_term ti) \<Longrightarrow> kbo si ti = wpo.wpo si ti" for si ti :: "('f,'a)term" by auto
- note wpo_sI = arg_cong[OF wpo.wpo.simps, of fst, THEN iffD2]
- note wpo_nsI = arg_cong[OF wpo.wpo.simps, of snd, THEN iffD2]
- note bounded = less(3)
- show ?case
- proof (cases s)
- case s: (Var x)
- have "\<not> weight_less t (Var x)"
- by (metis leD weight.simps(1) weight_le_less_iff weight_less_imp_le weight_w0)
- thus ?thesis
- by (cases t, auto simp add: s kbo_altdef wpo.wpo.simps, simp add: weight_le_def)
- next
- case s: (Fun f ss)
- show ?thesis
- proof (cases t)
- case t: (Var y)
- { assume "weight_le t s"
- then have "\<exists>s' \<in> set ss. weight_le t s'"
- apply (auto simp: s t weight_le_def)
- by (metis scf set_scf_list weight_w0)
- then obtain s' where s': "s' \<in> set ss" and "weight_le t s'" by auto
- from this(2) have "wpo.wpo_ns s' t"
- proof (induct s')
- case (Var x)
- then show ?case by (auto intro!:wpo_nsI simp: t weight_le_Var_Var)
- next
- case (Fun f' ss')
- from this(2) have "\<exists>s'' \<in> set ss'. weight_le t s''"
- apply (auto simp: t weight_le_def)
- by (metis scf set_scf_list weight_w0)
- then obtain s'' where "s'' \<in> set ss'" and "weight_le t s''" by auto
- with Fun(1)[OF this] Fun(2)
- show ?case by (auto intro!: wpo_nsI simp: t in_set_conv_nth)
- qed
- with s' have "\<exists>s' \<in> set ss. wpo.wpo_ns s' t" by auto
- }
- then
- show ?thesis unfolding wpo.wpo.simps[of s t] kbo_altdef[of s t]
- by (auto simp add: s t weight_less_iff set_conv_nth, auto)
- next
- case t: (Fun g ts)
- {
- fix j
- assume "j < length ts"
- hence "ts ! j \<in> set ts" by auto
- hence "funas_term (ts ! j) \<subseteq> funas_term t" unfolding t by auto
- with bounded have "bounded_arity n (funas_term (ts ! j))" unfolding bounded_arity_def by auto
- } note bounded_tj = this
- note IH_tj = IH[OF _ this]
- show ?thesis
- proof (cases "\<not> weight_le t s \<or> weight_less t s")
- case True
- thus ?thesis unfolding wpo.wpo.simps[of s t] kbo_altdef[of s t]
- unfolding s t by (auto simp: weight_less_iff)
- next
- case False
- let ?f = "(f,length ss)"
- let ?g = "(g,length ts)"
- from False have wle: "weight_le t s = True" "weight_less t s = False"
- "(s, t) \<in> weight_NS \<longleftrightarrow> True" "(s, t) \<in> weight_S \<longleftrightarrow> False" by auto
- have lex: "(Lex = Lex \<and> Lex = Lex) = True" by simp
- have sig: "set (wpo.\<sigma> ?f) = {..<length ss}"
- "set (wpo.\<sigma> ?g) = {..<length ts}" by auto
- have map: "map ((!) ss) (wpo.\<sigma> ?f) = ss"
- "map ((!) ts) (wpo.\<sigma> ?g) = ts"
- by (auto simp: map_nth)
- have sizes: "i < length ss \<Longrightarrow> size (ss ! i) < size s" for i unfolding s
- by (simp add: size_simp1)
- have sizet: "i < length ts \<Longrightarrow> size (ts ! i) < size t" for i unfolding t
- by (simp add: size_simp1)
- have wpo: "wpo.wpo s t =
- (if \<exists>i\<in>{..<length ss}. wpo.wpo_ns (ss ! i) t then (True, True)
- else if pr_weak ?f ?g \<and> (\<forall>j\<in>{..<length ts}. wpo.wpo_s s (ts ! j))
- then if pr_strict ?f ?g then (True, True) else lex_ext wpo.wpo n ss ts
- else (False, False))"
- unfolding wpo.wpo.simps[of s t]
- unfolding s t term.simps split Let_def lex if_True sig map
- unfolding s[symmetric] t[symmetric] wle if_True weight_less_iff if_False False snd_conv by auto
- have "kbo s t = (if pr_strict ?f ?g then (True, True)
- else if pr_weak ?f ?g then lex_ext_unbounded kbo ss ts
- else (False, False))"
- unfolding kbo_altdef[of s t]
- unfolding s t term.simps split Let_def if_True
- unfolding s[symmetric] t[symmetric] wle if_True weight_less_iff if_False by auto
- also have "lex_ext_unbounded kbo ss ts = lex_ext kbo n ss ts"
- using bounded[unfolded t] unfolding bounded_arity_def lex_ext_def by auto
- also have "\<dots> = lex_ext wpo.wpo n ss ts"
- by (rule lex_ext_cong[OF refl refl refl], rule IH_tj, auto dest!: sizes sizet)
- finally have kbo: "kbo s t =
- (if pr_strict ?f ?g then (True, True)
- else if pr_weak ?f ?g then lex_ext wpo.wpo n ss ts
- else (False, False))" .
- show ?thesis
- proof (cases "\<exists>i\<in>{..<length ss}. wpo.wpo_ns (ss ! i) t")
- case True
- then obtain i where i: "i < length ss" and "wpo.wpo_ns (ss ! i) t" by auto
- then obtain b where "wpo.wpo (ss ! i) t = (b, True)" by (cases "wpo.wpo (ss ! i) t", auto)
- also have "wpo.wpo (ss ! i) t = kbo (ss ! i) t" using i by (intro IH[symmetric, OF _ bounded], auto dest: sizes)
- finally have "NS (ss ! i) t" by simp
- from kbo_supt_one[OF this]
- have "S (Fun f (take i ss @ ss ! i # drop (Suc i) ss)) t" .
- also have "(take i ss @ ss ! i # drop (Suc i) ss) = ss" using i by (metis id_take_nth_drop)
- also have "Fun f ss = s" unfolding s by simp
- finally have "S s t" .
- with S_imp_NS[OF this]
- have "kbo s t = (True,True)" by (cases "kbo s t", auto)
- with True show ?thesis unfolding wpo by auto
- next
- case False
- hence False: "(\<exists>i\<in>{..<length ss}. wpo.wpo_ns (ss ! i) t) = False" by simp
- {
- fix j
- assume NS: "NS s t"
- assume j: "j < length ts"
- (* here we make use of proven properties of KBO: subterm-property and transitivity,
- perhaps there is a simple proof without already using these properties *)
- from kbo_supt_one[OF NS_refl, of g "take j ts" "ts ! j" "drop (Suc j) ts"]
- have S: "S t (ts ! j)" using id_take_nth_drop[OF j] unfolding t by auto
- from kbo_trans[of s t "ts ! j"] NS S have "S s (ts ! j)" by auto
- with S S_imp_NS[OF this]
- have "kbo s (ts ! j) = (True, True)" by (cases "kbo s (ts ! j)", auto)
- hence "wpo.wpo_s s (ts ! j)"
- by (subst IH_tj[symmetric], insert sizet[OF j] j, auto)
- }
- thus ?thesis unfolding wpo kbo False if_False using lex_ext_stri_imp_nstri[of wpo.wpo n ss ts]
- by (cases "lex_ext wpo.wpo n ss ts", auto simp: pr_strict split: if_splits)
- qed
- qed
- qed
- qed
- qed
-qed
-end
-
-text \<open>This is the main theorem. It tells us that KBO can be seen as an instance of WPO, under mild preconditions:
- the parameter $n$ for the lexicographic extension has to be chosen high enough to cover the arities of all
- terms that should be compared.\<close>
-lemma defines "prec \<equiv> ((\<lambda>f g. (pr_strict' f g, pr_weak' f g)))"
- and "prl \<equiv> (\<lambda>(f, n). n = 0 \<and> least f)"
- shows
- kbo_encoding_is_valid_wpo: "wpo_with_assms weight_S weight_NS prec prl full_status False (\<lambda>f. False)"
- and
- kbo_as_wpo: "bounded_arity n (funas_term t) \<Longrightarrow> kbo s t = wpo.wpo n weight_S weight_NS prec prl full_status (\<lambda>_. Lex) False (\<lambda>f. False) s t"
- unfolding prec_def prl_def
- subgoal by (intro admissible_kbo.kbo_instance_of_wpo_with_assms[OF admissible_kbo']
- least_pr_weak' least_pr_weak'_trans)
- apply (subst kbo'_eq_kbo[symmetric])
- apply (subst admissible_kbo.kbo_as_wpo_with_assms[OF admissible_kbo' least_pr_weak' least_pr_weak'_trans, symmetric], (auto)[3])
- by auto
-
-text \<open>As a proof-of-concept we show that now properties of WPO can be used to prove these properties for KBO.
- Here, as example we consider closure under substitutions and strong normalization,
- but the following idea can be applied for several more properties:
- if the property involves only terms where the arities are bounded, then just choose the parameter $n$ large enough.
- This even works for strong normalization, since in an infinite chain of KBO-decreases $t_1 > t_2 > t_3 > ...$ all terms have
- a weight of at most the weight of $t_1$, and this weight is also a bound on the arities.\<close>
-
-lemma KBO_stable_via_WPO: "S s t \<Longrightarrow> S (s \<cdot> (\<sigma> :: ('f,'a) subst)) (t \<cdot> \<sigma>)"
-proof -
- let ?terms = "{t, t \<cdot> \<sigma>}" (* collect all rhss of comparisons *)
- let ?prec = "((\<lambda>f g. (pr_strict' f g, pr_weak' f g)))"
- let ?prl = "(\<lambda>(f, n). n = 0 \<and> least f)"
- have "finite (\<Union> (funas_term ` ?terms))"
- by auto
- from finite_list[OF this] obtain F where F: "set F = \<Union> (funas_term ` ?terms)" by auto
- (* since there only finitely many symbols, we can take n as the maximal arity *)
- define n where "n = max_list (map snd F)"
-
- (* now get a WPO for this choice of n *)
- interpret wpo: wpo_with_assms
- where S = weight_S and NS = weight_NS
- and prc = ?prec and prl = ?prl
- and c = "\<lambda>_. Lex"
- and ssimple = False and large = "\<lambda>f. False" and \<sigma>\<sigma> = full_status
- and n = n
- by (rule kbo_encoding_is_valid_wpo)
-
- {
- fix t
- assume "t \<in> ?terms"
- hence "funas_term t \<subseteq> set F" unfolding F by auto
- hence "bounded_arity n (funas_term t)" unfolding bounded_arity_def
- using max_list[of _ "map snd F", folded n_def] by fastforce
- }
- (* for all the terms we have that KBO = WPO *)
- note kbo_as_wpo = kbo_as_wpo[OF this]
-
- (* and finally transfer the existing property of WPO to KBO *)
- from wpo.WPO_S_subst[of s t \<sigma>]
- show "S s t \<Longrightarrow> S (s \<cdot> \<sigma>) (t \<cdot> \<sigma>)"
- using kbo_as_wpo by auto
-qed
-
-lemma weight_is_arity_bound: "weight t \<le> b \<Longrightarrow> bounded_arity b (funas_term t)"
-proof (induct t)
- case (Fun f ts)
- have "sum_list (map weight ts) \<le> weight (Fun f ts)"
- using sum_list_scf_list[of ts "scf (f,length ts)", OF scf] by auto
- also have "\<dots> \<le> b" using Fun by auto
- finally have sum_b: "sum_list (map weight ts) \<le> b" .
- {
- fix t
- assume t: "t \<in> set ts"
- from split_list[OF this] have "weight t \<le> sum_list (map weight ts)" by auto
- with sum_b have "bounded_arity b (funas_term t)" using t Fun by auto
- } note IH = this
- have "length ts = sum_list (map (\<lambda> _. 1) ts)" by (induct ts, auto)
- also have "\<dots> \<le> sum_list (map weight ts)"
- apply (rule sum_list_mono)
- subgoal for t using weight_gt_0[of t] by auto
- done
- also have "\<dots> \<le> b" by fact
- finally have len: "length ts \<le> b" by auto
- from IH len show ?case unfolding bounded_arity_def by auto
-qed (auto simp: bounded_arity_def)
-
-lemma KBO_SN_via_WPO: "SN {(s,t). S s t}"
-proof
- fix f :: "nat \<Rightarrow> ('f,'a)term"
- assume "\<forall>i. (f i, f (Suc i)) \<in> {(s, t). S s t}"
- hence steps: "S (f i) (f (Suc i))" for i by auto
- define n where "n = weight (f 0)"
-
- have w_bound: "weight (f i) \<le> n" for i
- proof (induct i)
- case (Suc i)
- from steps[of i] have "weight (f (Suc i)) \<le> weight (f i)"
- unfolding kbo.simps[of "f i"] by (auto split: if_splits)
- with Suc show ?case by simp
- qed (auto simp: n_def)
-
- let ?prec = "((\<lambda>f g. (pr_strict' f g, pr_weak' f g)))"
- let ?prl = "(\<lambda>(f, n). n = 0 \<and> least f)"
-
- (* now get a WPO for this choice of n *)
- interpret wpo: wpo_with_assms
- where S = weight_S and NS = weight_NS
- and prc = ?prec and prl = ?prl
- and c = "\<lambda>_. Lex"
- and ssimple = False and large = "\<lambda>f. False" and \<sigma>\<sigma> = full_status
- and n = n
- by (rule kbo_encoding_is_valid_wpo)
-
- have "kbo (f i) (f (Suc i)) = wpo.wpo (f i) (f (Suc i))" for i
- by (rule kbo_as_wpo[OF weight_is_arity_bound[OF w_bound]])
- (* for all the terms in the infinite sequence f 0, f 1, ...
- we have that KBO = WPO *)
-
- (* and finally derive contradiction to SN-property of WPO *)
- from steps[unfolded this] wpo.WPO_S_SN show False by auto
-qed
-
-end
-
+subsection \<open>A restricted equality between KBO and WPO\<close>
+
+text \<open>The remaining difficulty to make KBO an instance of WPO is the
+ different treatment of lexicographic comparisons, which is unrestricted in KBO,
+ but there is a length-restriction in WPO.
+ Therefore we will only show that KBO is an instance of WPO if we compare terms with
+ bounded arity.\<close>
+
+text \<open>This restriction does however not prohibit us from lifting properties of WPO to KBO.
+ For instance, for several properties one can choose a large-enough bound restriction of WPO,
+ since there are only finitely many arities occurring in a property.\<close>
+
+theory KBO_as_WPO
+ imports
+ WPO
+ KBO_Transformation
+begin
+
+definition bounded_arity :: "nat \<Rightarrow> ('f \<times> nat)set \<Rightarrow> bool" where
+ "bounded_arity b F = (\<forall> (f,n) \<in> F. n \<le> b)"
+
+lemma finite_funas_term[simp,intro]: "finite (funas_term t)"
+ by (induct t, auto)
+
+context weight_fun begin
+
+definition "weight_le s t \<equiv>
+ (vars_term_ms (SCF s) \<subseteq># vars_term_ms (SCF t) \<and> weight s \<le> weight t)"
+
+definition "weight_less s t \<equiv>
+ (vars_term_ms (SCF s) \<subseteq># vars_term_ms (SCF t) \<and> weight s < weight t)"
+
+lemma weight_le_less_iff: "weight_le s t \<Longrightarrow> weight_less s t \<longleftrightarrow> weight s < weight t"
+ by (auto simp: weight_le_def weight_less_def)
+
+lemma weight_less_iff: "weight_less s t \<Longrightarrow> weight_le s t \<and> weight s < weight t"
+ by (auto simp: weight_le_def weight_less_def)
+
+
+abbreviation "weight_NS \<equiv> {(t,s). weight_le s t}"
+
+abbreviation "weight_S \<equiv> {(t,s). weight_less s t}"
+
+lemma weight_le_mono_one:
+ assumes S: "weight_le s t"
+ shows "weight_le (Fun f (ss1 @ s # ss2)) (Fun f (ss1 @ t # ss2))" (is "weight_le ?s ?t")
+proof -
+ from S have w: "weight s \<le> weight t" and v: "vars_term_ms (SCF s) \<subseteq># vars_term_ms (SCF t)"
+ by (auto simp: weight_le_def)
+ have v': "vars_term_ms (SCF ?s) \<subseteq># vars_term_ms (SCF ?t)"
+ using mset_replicate_mono[OF v] by simp
+ have w': "weight ?s \<le> weight ?t" using sum_list_replicate_mono[OF w] by simp
+ from v' w' show ?thesis by (auto simp: weight_le_def)
+qed
+
+lemma weight_le_ctxt: "weight_le s t \<Longrightarrow> weight_le (C\<langle>s\<rangle>) (C\<langle>t\<rangle>)"
+ by (induct C, auto intro: weight_le_mono_one)
+
+lemma SCF_stable:
+ assumes "vars_term_ms (SCF s) \<subseteq># vars_term_ms (SCF t)"
+ shows "vars_term_ms (SCF (s \<cdot> \<sigma>)) \<subseteq># vars_term_ms (SCF (t \<cdot> \<sigma>))"
+ unfolding scf_term_subst
+ using vars_term_ms_subst_mono[OF assms].
+
+lemma SN_weight_S: "SN weight_S"
+proof-
+ from wf_inv_image[OF wf_less]
+ have wf: "wf {(s,t). weight s < weight t}" by (auto simp: inv_image_def)
+ show ?thesis
+ by (unfold SN_iff_wf, rule wf_subset[OF wf], auto simp: weight_less_def)
+qed
+
+lemma weight_less_imp_le: "weight_less s t \<Longrightarrow> weight_le s t" by (simp add: weight_less_def weight_le_def)
+
+lemma weight_le_Var_Var: "weight_le (Var x) (Var y) \<longleftrightarrow> x = y"
+ by (auto simp: weight_le_def)
+end
+
+context kbo begin
+
+lemma kbo_altdef:
+ "kbo s t = (if weight_le t s
+ then if weight_less t s
+ then (True, True)
+ else (case s of
+ Var y \<Rightarrow> (False, (case t of Var x \<Rightarrow> x = y | Fun g ts \<Rightarrow> ts = [] \<and> least g))
+ | Fun f ss \<Rightarrow> (case t of
+ Var x \<Rightarrow> (True, True)
+ | Fun g ts \<Rightarrow> if pr_strict (f, length ss) (g, length ts)
+ then (True, True)
+ else if pr_weak (f, length ss) (g, length ts)
+ then lex_ext_unbounded kbo ss ts
+ else (False, False)))
+ else (False, False))"
+ by (simp add: weight_le_less_iff weight_le_def)
+
+end
+
+context admissible_kbo begin
+
+lemma weight_le_stable:
+ assumes "weight_le s t"
+ shows "weight_le (s \<cdot> \<sigma>) (t \<cdot> \<sigma>)"
+ using assms weight_stable_le SCF_stable by (auto simp: weight_le_def)
+
+lemma weight_less_stable:
+ assumes "weight_less s t"
+ shows "weight_less (s \<cdot> \<sigma>) (t \<cdot> \<sigma>)"
+ using assms weight_stable_lt SCF_stable by (auto simp: weight_less_def)
+
+lemma simple_arg_pos_weight: "simple_arg_pos weight_NS (f,n) i"
+ unfolding simple_arg_pos_def
+proof (intro allI impI, unfold snd_conv fst_conv)
+ fix ts :: "('f,'a)term list"
+ assume i: "i < n" and len: "length ts = n"
+ from id_take_nth_drop[OF i[folded len]] i[folded len]
+ obtain us vs where id: "Fun f ts = Fun f (us @ ts ! i # vs)"
+ and us: "us = take i ts"
+ and len: "length us = i" by auto
+ have "length us < Suc (length us + length vs)" by auto
+ from scf[OF this, of f] obtain j where [simp]: "scf (f, Suc (length us + length vs)) (length us) = Suc j"
+ by (rule lessE)
+ show "(Fun f ts, ts ! i) \<in> weight_NS"
+ unfolding weight_le_def id by (auto simp: o_def)
+qed
+
+lemma weight_lemmas:
+ shows "refl weight_NS" and "trans weight_NS" and "trans weight_S"
+ and "weight_NS O weight_S \<subseteq> weight_S" and "weight_S O weight_NS \<subseteq> weight_S"
+ by (auto intro!: refl_onI transI simp: weight_le_def weight_less_def)
+
+interpretation kbo': admissible_kbo w w0 pr_strict' pr_weak' least scf
+ by (rule admissible_kbo')
+
+context
+ assumes least_global: "\<And> f g. least f \<Longrightarrow> pr_weak g (f,0)"
+ and least_trans: "\<And> f g. least f \<Longrightarrow> pr_weak (f,0) g \<Longrightarrow> least (fst g) \<and> snd g = 0"
+ fixes n :: nat
+begin
+
+lemma kbo_instance_of_wpo_with_assms: "wpo_with_assms
+ weight_S weight_NS (\<lambda>f g. (pr_strict f g, pr_weak f g))
+ (\<lambda>(f, n). n = 0 \<and> least f) full_status False (\<lambda>f. False)"
+ apply (unfold_locales)
+ apply (auto simp: weight_lemmas SN_weight_S pr_SN pr_strict_irrefl
+ weight_less_stable weight_le_stable weight_le_mono_one weight_less_imp_le
+ simple_arg_pos_weight)
+ apply (force dest: least_global least_trans simp: pr_strict)+
+ apply (auto simp: pr_strict least dest:pr_weak_trans)
+ done
+
+interpretation wpo: wpo_with_assms
+ where S = weight_S and NS = weight_NS
+ and prc = "\<lambda>f g. (pr_strict f g, pr_weak f g)" and prl = "\<lambda>(f,n). n = 0 \<and> least f"
+ and c = "\<lambda>_. Lex"
+ and ssimple = False and large = "\<lambda>f. False" and \<sigma>\<sigma> = full_status
+ and n = n
+ by (rule kbo_instance_of_wpo_with_assms)
+
+lemma kbo_as_wpo_with_assms: assumes "bounded_arity n (funas_term t)"
+ shows "kbo s t = wpo.wpo s t"
+proof -
+ define m where "m = size s + size t"
+ from m_def assms show ?thesis
+ proof (induct m arbitrary: s t rule: less_induct)
+ case (less m s t)
+ hence IH: "size si + size ti < size s + size t \<Longrightarrow> bounded_arity n (funas_term ti) \<Longrightarrow> kbo si ti = wpo.wpo si ti" for si ti :: "('f,'a)term" by auto
+ note wpo_sI = arg_cong[OF wpo.wpo.simps, of fst, THEN iffD2]
+ note wpo_nsI = arg_cong[OF wpo.wpo.simps, of snd, THEN iffD2]
+ note bounded = less(3)
+ show ?case
+ proof (cases s)
+ case s: (Var x)
+ have "\<not> weight_less t (Var x)"
+ by (metis leD weight.simps(1) weight_le_less_iff weight_less_imp_le weight_w0)
+ thus ?thesis
+ by (cases t, auto simp add: s kbo_altdef wpo.wpo.simps, simp add: weight_le_def)
+ next
+ case s: (Fun f ss)
+ show ?thesis
+ proof (cases t)
+ case t: (Var y)
+ { assume "weight_le t s"
+ then have "\<exists>s' \<in> set ss. weight_le t s'"
+ apply (auto simp: s t weight_le_def)
+ by (metis scf set_scf_list weight_w0)
+ then obtain s' where s': "s' \<in> set ss" and "weight_le t s'" by auto
+ from this(2) have "wpo.wpo_ns s' t"
+ proof (induct s')
+ case (Var x)
+ then show ?case by (auto intro!:wpo_nsI simp: t weight_le_Var_Var)
+ next
+ case (Fun f' ss')
+ from this(2) have "\<exists>s'' \<in> set ss'. weight_le t s''"
+ apply (auto simp: t weight_le_def)
+ by (metis scf set_scf_list weight_w0)
+ then obtain s'' where "s'' \<in> set ss'" and "weight_le t s''" by auto
+ with Fun(1)[OF this] Fun(2)
+ show ?case by (auto intro!: wpo_nsI simp: t in_set_conv_nth)
+ qed
+ with s' have "\<exists>s' \<in> set ss. wpo.wpo_ns s' t" by auto
+ }
+ then
+ show ?thesis unfolding wpo.wpo.simps[of s t] kbo_altdef[of s t]
+ by (auto simp add: s t weight_less_iff set_conv_nth, auto)
+ next
+ case t: (Fun g ts)
+ {
+ fix j
+ assume "j < length ts"
+ hence "ts ! j \<in> set ts" by auto
+ hence "funas_term (ts ! j) \<subseteq> funas_term t" unfolding t by auto
+ with bounded have "bounded_arity n (funas_term (ts ! j))" unfolding bounded_arity_def by auto
+ } note bounded_tj = this
+ note IH_tj = IH[OF _ this]
+ show ?thesis
+ proof (cases "\<not> weight_le t s \<or> weight_less t s")
+ case True
+ thus ?thesis unfolding wpo.wpo.simps[of s t] kbo_altdef[of s t]
+ unfolding s t by (auto simp: weight_less_iff)
+ next
+ case False
+ let ?f = "(f,length ss)"
+ let ?g = "(g,length ts)"
+ from False have wle: "weight_le t s = True" "weight_less t s = False"
+ "(s, t) \<in> weight_NS \<longleftrightarrow> True" "(s, t) \<in> weight_S \<longleftrightarrow> False" by auto
+ have lex: "(Lex = Lex \<and> Lex = Lex) = True" by simp
+ have sig: "set (wpo.\<sigma> ?f) = {..<length ss}"
+ "set (wpo.\<sigma> ?g) = {..<length ts}" by auto
+ have map: "map ((!) ss) (wpo.\<sigma> ?f) = ss"
+ "map ((!) ts) (wpo.\<sigma> ?g) = ts"
+ by (auto simp: map_nth)
+ have sizes: "i < length ss \<Longrightarrow> size (ss ! i) < size s" for i unfolding s
+ by (simp add: size_simp1)
+ have sizet: "i < length ts \<Longrightarrow> size (ts ! i) < size t" for i unfolding t
+ by (simp add: size_simp1)
+ have wpo: "wpo.wpo s t =
+ (if \<exists>i\<in>{..<length ss}. wpo.wpo_ns (ss ! i) t then (True, True)
+ else if pr_weak ?f ?g \<and> (\<forall>j\<in>{..<length ts}. wpo.wpo_s s (ts ! j))
+ then if pr_strict ?f ?g then (True, True) else lex_ext wpo.wpo n ss ts
+ else (False, False))"
+ unfolding wpo.wpo.simps[of s t]
+ unfolding s t term.simps split Let_def lex if_True sig map
+ unfolding s[symmetric] t[symmetric] wle if_True weight_less_iff if_False False snd_conv by auto
+ have "kbo s t = (if pr_strict ?f ?g then (True, True)
+ else if pr_weak ?f ?g then lex_ext_unbounded kbo ss ts
+ else (False, False))"
+ unfolding kbo_altdef[of s t]
+ unfolding s t term.simps split Let_def if_True
+ unfolding s[symmetric] t[symmetric] wle if_True weight_less_iff if_False by auto
+ also have "lex_ext_unbounded kbo ss ts = lex_ext kbo n ss ts"
+ using bounded[unfolded t] unfolding bounded_arity_def lex_ext_def by auto
+ also have "\<dots> = lex_ext wpo.wpo n ss ts"
+ by (rule lex_ext_cong[OF refl refl refl], rule IH_tj, auto dest!: sizes sizet)
+ finally have kbo: "kbo s t =
+ (if pr_strict ?f ?g then (True, True)
+ else if pr_weak ?f ?g then lex_ext wpo.wpo n ss ts
+ else (False, False))" .
+ show ?thesis
+ proof (cases "\<exists>i\<in>{..<length ss}. wpo.wpo_ns (ss ! i) t")
+ case True
+ then obtain i where i: "i < length ss" and "wpo.wpo_ns (ss ! i) t" by auto
+ then obtain b where "wpo.wpo (ss ! i) t = (b, True)" by (cases "wpo.wpo (ss ! i) t", auto)
+ also have "wpo.wpo (ss ! i) t = kbo (ss ! i) t" using i by (intro IH[symmetric, OF _ bounded], auto dest: sizes)
+ finally have "NS (ss ! i) t" by simp
+ from kbo_supt_one[OF this]
+ have "S (Fun f (take i ss @ ss ! i # drop (Suc i) ss)) t" .
+ also have "(take i ss @ ss ! i # drop (Suc i) ss) = ss" using i by (metis id_take_nth_drop)
+ also have "Fun f ss = s" unfolding s by simp
+ finally have "S s t" .
+ with S_imp_NS[OF this]
+ have "kbo s t = (True,True)" by (cases "kbo s t", auto)
+ with True show ?thesis unfolding wpo by auto
+ next
+ case False
+ hence False: "(\<exists>i\<in>{..<length ss}. wpo.wpo_ns (ss ! i) t) = False" by simp
+ {
+ fix j
+ assume NS: "NS s t"
+ assume j: "j < length ts"
+ (* here we make use of proven properties of KBO: subterm-property and transitivity,
+ perhaps there is a simple proof without already using these properties *)
+ from kbo_supt_one[OF NS_refl, of g "take j ts" "ts ! j" "drop (Suc j) ts"]
+ have S: "S t (ts ! j)" using id_take_nth_drop[OF j] unfolding t by auto
+ from kbo_trans[of s t "ts ! j"] NS S have "S s (ts ! j)" by auto
+ with S S_imp_NS[OF this]
+ have "kbo s (ts ! j) = (True, True)" by (cases "kbo s (ts ! j)", auto)
+ hence "wpo.wpo_s s (ts ! j)"
+ by (subst IH_tj[symmetric], insert sizet[OF j] j, auto)
+ }
+ thus ?thesis unfolding wpo kbo False if_False using lex_ext_stri_imp_nstri[of wpo.wpo n ss ts]
+ by (cases "lex_ext wpo.wpo n ss ts", auto simp: pr_strict split: if_splits)
+ qed
+ qed
+ qed
+ qed
+ qed
+qed
+end
+
+text \<open>This is the main theorem. It tells us that KBO can be seen as an instance of WPO, under mild preconditions:
+ the parameter $n$ for the lexicographic extension has to be chosen high enough to cover the arities of all
+ terms that should be compared.\<close>
+lemma defines "prec \<equiv> ((\<lambda>f g. (pr_strict' f g, pr_weak' f g)))"
+ and "prl \<equiv> (\<lambda>(f, n). n = 0 \<and> least f)"
+ shows
+ kbo_encoding_is_valid_wpo: "wpo_with_assms weight_S weight_NS prec prl full_status False (\<lambda>f. False)"
+ and
+ kbo_as_wpo: "bounded_arity n (funas_term t) \<Longrightarrow> kbo s t = wpo.wpo n weight_S weight_NS prec prl full_status (\<lambda>_. Lex) False (\<lambda>f. False) s t"
+ unfolding prec_def prl_def
+ subgoal by (intro admissible_kbo.kbo_instance_of_wpo_with_assms[OF admissible_kbo']
+ least_pr_weak' least_pr_weak'_trans)
+ apply (subst kbo'_eq_kbo[symmetric])
+ apply (subst admissible_kbo.kbo_as_wpo_with_assms[OF admissible_kbo' least_pr_weak' least_pr_weak'_trans, symmetric], (auto)[3])
+ by auto
+
+text \<open>As a proof-of-concept we show that now properties of WPO can be used to prove these properties for KBO.
+ Here, as example we consider closure under substitutions and strong normalization,
+ but the following idea can be applied for several more properties:
+ if the property involves only terms where the arities are bounded, then just choose the parameter $n$ large enough.
+ This even works for strong normalization, since in an infinite chain of KBO-decreases $t_1 > t_2 > t_3 > ...$ all terms have
+ a weight of at most the weight of $t_1$, and this weight is also a bound on the arities.\<close>
+
+lemma KBO_stable_via_WPO: "S s t \<Longrightarrow> S (s \<cdot> (\<sigma> :: ('f,'a) subst)) (t \<cdot> \<sigma>)"
+proof -
+ let ?terms = "{t, t \<cdot> \<sigma>}" (* collect all rhss of comparisons *)
+ let ?prec = "((\<lambda>f g. (pr_strict' f g, pr_weak' f g)))"
+ let ?prl = "(\<lambda>(f, n). n = 0 \<and> least f)"
+ have "finite (\<Union> (funas_term ` ?terms))"
+ by auto
+ from finite_list[OF this] obtain F where F: "set F = \<Union> (funas_term ` ?terms)" by auto
+ (* since there only finitely many symbols, we can take n as the maximal arity *)
+ define n where "n = max_list (map snd F)"
+
+ (* now get a WPO for this choice of n *)
+ interpret wpo: wpo_with_assms
+ where S = weight_S and NS = weight_NS
+ and prc = ?prec and prl = ?prl
+ and c = "\<lambda>_. Lex"
+ and ssimple = False and large = "\<lambda>f. False" and \<sigma>\<sigma> = full_status
+ and n = n
+ by (rule kbo_encoding_is_valid_wpo)
+
+ {
+ fix t
+ assume "t \<in> ?terms"
+ hence "funas_term t \<subseteq> set F" unfolding F by auto
+ hence "bounded_arity n (funas_term t)" unfolding bounded_arity_def
+ using max_list[of _ "map snd F", folded n_def] by fastforce
+ }
+ (* for all the terms we have that KBO = WPO *)
+ note kbo_as_wpo = kbo_as_wpo[OF this]
+
+ (* and finally transfer the existing property of WPO to KBO *)
+ from wpo.WPO_S_subst[of s t \<sigma>]
+ show "S s t \<Longrightarrow> S (s \<cdot> \<sigma>) (t \<cdot> \<sigma>)"
+ using kbo_as_wpo by auto
+qed
+
+lemma weight_is_arity_bound: "weight t \<le> b \<Longrightarrow> bounded_arity b (funas_term t)"
+proof (induct t)
+ case (Fun f ts)
+ have "sum_list (map weight ts) \<le> weight (Fun f ts)"
+ using sum_list_scf_list[of ts "scf (f,length ts)", OF scf] by auto
+ also have "\<dots> \<le> b" using Fun by auto
+ finally have sum_b: "sum_list (map weight ts) \<le> b" .
+ {
+ fix t
+ assume t: "t \<in> set ts"
+ from split_list[OF this] have "weight t \<le> sum_list (map weight ts)" by auto
+ with sum_b have "bounded_arity b (funas_term t)" using t Fun by auto
+ } note IH = this
+ have "length ts = sum_list (map (\<lambda> _. 1) ts)" by (induct ts, auto)
+ also have "\<dots> \<le> sum_list (map weight ts)"
+ apply (rule sum_list_mono)
+ subgoal for t using weight_gt_0[of t] by auto
+ done
+ also have "\<dots> \<le> b" by fact
+ finally have len: "length ts \<le> b" by auto
+ from IH len show ?case unfolding bounded_arity_def by auto
+qed (auto simp: bounded_arity_def)
+
+lemma KBO_SN_via_WPO: "SN {(s,t). S s t}"
+proof
+ fix f :: "nat \<Rightarrow> ('f,'a)term"
+ assume "\<forall>i. (f i, f (Suc i)) \<in> {(s, t). S s t}"
+ hence steps: "S (f i) (f (Suc i))" for i by auto
+ define n where "n = weight (f 0)"
+
+ have w_bound: "weight (f i) \<le> n" for i
+ proof (induct i)
+ case (Suc i)
+ from steps[of i] have "weight (f (Suc i)) \<le> weight (f i)"
+ unfolding kbo.simps[of "f i"] by (auto split: if_splits)
+ with Suc show ?case by simp
+ qed (auto simp: n_def)
+
+ let ?prec = "((\<lambda>f g. (pr_strict' f g, pr_weak' f g)))"
+ let ?prl = "(\<lambda>(f, n). n = 0 \<and> least f)"
+
+ (* now get a WPO for this choice of n *)
+ interpret wpo: wpo_with_assms
+ where S = weight_S and NS = weight_NS
+ and prc = ?prec and prl = ?prl
+ and c = "\<lambda>_. Lex"
+ and ssimple = False and large = "\<lambda>f. False" and \<sigma>\<sigma> = full_status
+ and n = n
+ by (rule kbo_encoding_is_valid_wpo)
+
+ have "kbo (f i) (f (Suc i)) = wpo.wpo (f i) (f (Suc i))" for i
+ by (rule kbo_as_wpo[OF weight_is_arity_bound[OF w_bound]])
+ (* for all the terms in the infinite sequence f 0, f 1, ...
+ we have that KBO = WPO *)
+
+ (* and finally derive contradiction to SN-property of WPO *)
+ from steps[unfolded this] wpo.WPO_S_SN show False by auto
+qed
+
+end
+
end
\ No newline at end of file

File Metadata

Mime Type
application/octet-stream
Expires
Sun, May 5, 5:06 AM (1 d, 23 h)
Storage Engine
chunks
Storage Format
Chunks
Storage Handle
seLtXmSmRIBo
Default Alt Text
(7 MB)

Event Timeline