diff --git a/thys/Conditional_Transfer_Rule/CTR/Tests/CTR_Tests.thy b/thys/Conditional_Transfer_Rule/CTR/Tests/CTR_Tests.thy --- a/thys/Conditional_Transfer_Rule/CTR/Tests/CTR_Tests.thy +++ b/thys/Conditional_Transfer_Rule/CTR/Tests/CTR_Tests.thy @@ -1,250 +1,252 @@ (* Title: CTR/Tests/CTR_Tests.thy Author: Mihails Milehins Copyright 2021 (C) Mihails Milehins A test suite for the sub-framework CTR. *) section\A test suite for CTR\ theory CTR_Tests imports "../CTR" "../../IML_UT/IML_UT" Complex_Main keywords "ctr_test" :: thy_defn begin subsection\Background\ ML\ type ctr_test_data = { ctr_type : string, synthesis : (string * thm list option) option, elems: (string, string, Facts.ref) Element.ctxt list, type_specs : (string * string) list, thm_specs : ((binding option * thm) * mixfix) list }; structure CTRTestData = Generic_Data ( type T = ctr_test_data Symtab.table val empty = Symtab.empty val merge = Symtab.merge (K true) ); val get_ctr_test_data_generic = CTRTestData.get; val get_ctr_test_data_proof = Context.Proof #> get_ctr_test_data_generic; val get_ctr_test_data_global = Context.Theory #> get_ctr_test_data_generic; fun test_data_of_generic context = context |> get_ctr_test_data_generic |> Symtab.lookup; val ctr_test_data_of_proof = Context.Proof #> test_data_of_generic; (*oversimplified: to be used with care*) fun update_ctr_test_data k ctr_test_data = Local_Theory.declaration {pervasive=true, syntax=false} (fn _ => (k, ctr_test_data) |> Symtab.update |> CTRTestData.map); fun process_ctr_test_data (k, args) (lthy : local_theory) = let fun preprocess_thm_specs lthy = map (apfst (apsnd (singleton (Attrib.eval_thms lthy)))) fun process_ctrs_impl (CTR.ALG_PP _) (lthy : local_theory) = lthy | process_ctrs_impl (CTR.ALG_RP (((synthesis, elems), type_specs), thm_specs)) (lthy : local_theory) = let val thm_specs' = preprocess_thm_specs lthy thm_specs val synthesis' = Option.map (apsnd (Option.map ((single #> Attrib.eval_thms lthy)))) synthesis val data : ctr_test_data = { ctr_type = "relativization", synthesis = synthesis', elems = elems, type_specs = type_specs, thm_specs = thm_specs' } in update_ctr_test_data k data lthy end in process_ctrs_impl args lthy end; val ctr_test_parser = Parse.string -- CTR.ctr_parser; val _ = Outer_Syntax.local_theory \<^command_keyword>\ctr_test\ "test setup for the command ctr" (ctr_test_parser >> process_ctr_test_data); \ -ud \order.mono\ -ud mono' \mono\ +definition mono where + "mono f \ (\x y. x \ y \ f x \ f y)" + +ud \mono\ definition mono_ow :: "'a set \ ('b \ 'b \ bool) \ ('a \ 'a \ bool) \ ('a \ 'b) \ bool" where "mono_ow UB leb lea f \ \x\UB. \y\UB. lea x y \ leb (f x) (f y)" typedef 'a K = \{xs::'a list. length xs = 2}\ by (simp add: Ex_list_of_length) definition KK :: "'a K \ 'a K \ bool" where "KK k1 k2 \ k1 = k2" typedef 'a L = \{xs::'a list. length xs = 2}\ by (simp add: Ex_list_of_length) definition LL :: "'a L \ 'a L \ bool" where "LL k1 k2 \ k1 = k2" definition rel_L :: "('a::group_add \ 'b::group_add \ bool) \ 'a::group_add L \ 'b::group_add L \ bool" where "rel_L A b c = True" ctr_relator rel_L definition not_binders_binrelT :: "('a \ 'b \ bool) \ ('c \ bool) \ 'a \ 'b \ bool" where "not_binders_binrelT R1 R2 a b = True" definition no_dup_binrelT :: "('a \ 'b \ bool) \ ('c \ 'a \ bool) \ 'a \ 'b \ bool" where "no_dup_binrelT R1 R2 a b = True" definition not_binders_binrelT_ftv_stv :: "('a \ 'b \ bool) \ (nat \ 'c \ bool) \ 'a \ 'b \ bool" where "not_binders_binrelT_ftv_stv R1 R2 a b = True" definition not_type_constructor_lhs :: "('a \ 'b \ bool) \ ('c \ 'd \ bool) \ 'a \ 'a K \ bool" where "not_type_constructor_lhs R1 R2 a b = True" definition not_type_constructor_rhs :: "('a \ 'b \ bool) \ ('c \ 'd \ bool) \ 'a K \ 'e \ bool" where "not_type_constructor_rhs R1 R2 a b = True" definition not_identical_type_constructors :: "('a \ 'b \ bool) \ ('c \ 'd \ bool) \ 'a K \ 'e L \ bool" where "not_identical_type_constructors R1 R2 a b = True" definition not_identical_type_constructors_lhs :: "('a \ 'b \ bool) \ ('c \ 'd \ bool) \ 'a K \ 'b K \ bool" where "not_identical_type_constructors_lhs R1 R2 a b = True" definition not_identical_type_constructors_rhs :: "('a \ 'b \ bool) \ 'a K \ 'c K \ bool" where "not_identical_type_constructors_rhs R1 a b = True" subsection\Test data\ lemma mono_ow_transfer': includes lifting_syntax assumes [transfer_domain_rule, transfer_rule]: "Domainp B = (\x. x \ UB)" and [transfer_rule]: "right_total B" shows "((A ===> A ===> (=)) ===> (B ===> B ===> (=)) ===> (B ===> A) ===> (=)) (mono_ow UB) mono.with" unfolding mono_ow_def mono.with_def by (transfer_prover_start, transfer_step+) simp ctr_test "mono_with" relativization synthesis ctr_simps_Collect_mem_eq assumes [transfer_domain_rule, transfer_rule]: "Domainp (B::'c\'d\bool) = (\x. x \ UB)" and [transfer_rule]: "right_total B" trp (?'b \A::'a\'b\bool\) and (?'a B) in mono_ow': mono.with_def ctr_test "exI" relativization in mono_ow'': exI ctr_test "binrel" relativization synthesis ctr_simps_Collect_mem_eq assumes [transfer_domain_rule, transfer_rule]: "Domainp (B::'c\'d\bool) = (\x. x \ UB)" and [transfer_rule]: "right_total B" trp (?'b A) and (?'a B) in mono_ow': mono.with_def ctr_test "binrel_ftv" relativization synthesis ctr_simps_Collect_mem_eq assumes [transfer_domain_rule, transfer_rule]: "Domainp (B::'c\'d\bool) = (\x. x \ UB)" and [transfer_rule]: "right_total B" trp (?'b \A::nat\'b\bool\) and (?'a B) in mono_ow': mono.with_def ctr_test "dup_stvs" relativization synthesis ctr_simps_Collect_mem_eq assumes [transfer_domain_rule, transfer_rule]: "Domainp (B::'c\'d\bool) = (\x. x \ UB)" and [transfer_rule]: "right_total B" trp (?'b \A::'a\'b\bool\) and (?'b B) in mono_ow': mono.with_def ctr_test "dup_binrel_ftvs" relativization synthesis ctr_simps_Collect_mem_eq assumes [transfer_domain_rule, transfer_rule]: "Domainp (B::'c\'d\bool) = (\x. x \ UB)" and [transfer_rule]: "right_total B" trp (?'b \A::'a\'d\bool\) and (?'a B) in mono_ow': mono.with_def ctr_test "no_relator" relativization synthesis ctr_simps_Collect_mem_eq assumes [transfer_domain_rule, transfer_rule]: "Domainp (B::'c\'d\bool) = (\x. x \ UB)" and [transfer_rule]: "right_total B" trp (?'b \A::'a\'b\bool\) and (?'a B) in KK_def ctr_test "invalid_relator" relativization synthesis ctr_simps_Collect_mem_eq assumes [transfer_domain_rule, transfer_rule]: "Domainp (B::'c\'d\bool) = (\x. x \ UB)" and [transfer_rule]: "right_total B" trp (?'b \A::'a\'b\bool\) and (?'a B) in LL_def subsection\Tests\ subsubsection\\process_relativization\\ ML_file\CTR_TEST_PROCESS_RELATIVIZATION.ML\ context includes lifting_syntax begin ML\ Lecker.test_group @{context} () [ctr_test_process_relativization.test_suite] \ end subsubsection\\process_ctr_relator\\ ML_file\CTR_TEST_PROCESS_CTR_RELATOR.ML\ context includes lifting_syntax begin ML\ Lecker.test_group @{context} () [ctr_test_process_ctr_relator.test_suite] \ end end \ No newline at end of file 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,348 +1,340 @@ (* Title: UD/UD_Reference.thy Author: Mihails Milehins Copyright 2021 (C) Mihails Milehins Reference manual for the UD. *) section\UD\ theory UD_Reference imports UD "../Reference_Prerequisites" begin subsection\Introduction\ subsubsection\Background\ text\ 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}. \ subsubsection\Purpose and scope\ text\ 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. \ subsubsection\Related and previous work\ text\ 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. Lastly, it should be mentioned that the framework SpecCheck \cite{kappelmann_speccheck_2021} was used for unit testing the framework UD. \ subsection\Theory\label{sec:ud_theory}\ text\ 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. \ subsection\Syntax\ text\ 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. \ text\ \begin{matharray}{rcl} @{command_def "ud"} & : & \theory \ theory\\\ \end{matharray} \<^medskip> \<^rail>\@@{command ud} binding? const mixfix?\ \<^descr> \<^theory_text>\ud\ (\b\) \const\ (\mixfix\) 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 \b\ is used for the specification of the names of the entities added by the command to the theory and the optional argument \mixfix\ 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 \b\ or \mixfix\ 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}. \ subsection\Examples\label{sec:ud_ex}\ text\ 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. \ subsubsection\Type classes\ +definition mono where + "mono f \ (\x y. x \ y \ f x \ f y)" + text\ 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 +It is defined as \begin{center} @{thm [names_short = true] mono_def[no_vars]} \end{center} for any @{term [show_sorts] "f::'a::order\'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}: +The constants is unoverloaded using the command @{command ud}: \ -ud \order.mono\ -ud mono' \mono\ +ud \mono\ text\ -The invocation of the commands above declares the constant @{const mono.with} +The invocation of the command 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]}. +@{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}. \ subsubsection\Low-level overloading\ text\ The following example closely follows Example 5 in section 5.2. in \cite{kaufmann_mechanized_2010}. \ consts pls :: "'a \ 'a \ 'a" overloading pls_nat \ "pls::nat \ nat \ nat" pls_times \ "pls::'a \ 'b \ 'a \ 'b \ 'a \ 'b" begin definition pls_nat :: "nat \ nat \ nat" where "pls_nat a b = a + b" definition pls_times :: "'a \ 'b \ 'a \ 'b \ 'a \ 'b" where "pls_times \ \x y. (pls (fst x) (fst y), pls (snd x) (snd y))" end ud pls_nat \pls::nat \ nat \ nat\ ud pls_times \pls::'a \ 'b \ 'a \ 'b \ 'a \ 'b\ text\ As expected, two new unoverloaded constants are produced via the invocations of the command @{command ud} above. The first constant, \<^const>\pls_nat.with\, corresponds to \pls::nat \ nat \ nat\ and is given by \begin{center} @{thm pls_nat.with_def[no_vars]}, \end{center} the second constant, \<^const>\pls_times.with\, corresponds to \begin{center} \pls::'a \ 'b \ 'a \ 'b \ 'a \ 'b\ \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>\pls_nat.with\ and \<^const>\pls_times.with\ 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}. \ text\\newpage\ end diff --git a/thys/CryptHOL/Partial_Function_Set.thy b/thys/CryptHOL/Partial_Function_Set.thy --- a/thys/CryptHOL/Partial_Function_Set.thy +++ b/thys/CryptHOL/Partial_Function_Set.thy @@ -1,196 +1,193 @@ (* Title: Partial_Function_Set.thy Author: Andreas Lochbihler, ETH Zurich *) theory Partial_Function_Set imports Main begin subsection \Setup for \partial_function\ for sets\ lemma (in complete_lattice) lattice_partial_function_definition: "partial_function_definitions (\) Sup" by(unfold_locales)(auto intro: Sup_upper Sup_least) interpretation set: partial_function_definitions "(\)" Union by(rule lattice_partial_function_definition) lemma fun_lub_Sup: "fun_lub Sup = (Sup :: _ \ _ :: complete_lattice)" by(fastforce simp add: fun_lub_def fun_eq_iff Sup_fun_def intro: Sup_eqI SUP_upper SUP_least) lemma set_admissible: "set.admissible (\f :: 'a \ 'b set. \x y. y \ f x \ P x y)" by(rule ccpo.admissibleI)(auto simp add: fun_lub_Sup) abbreviation "mono_set \ monotone (fun_ord (\)) (\)" lemma fixp_induct_set_scott: fixes F :: "'c \ 'c" and U :: "'c \ 'b \ 'a set" and C :: "('b \ 'a set) \ 'c" and P :: "'b \ 'a \ bool" and x and y assumes mono: "\x. mono_set (\f. U (F (C f)) x)" and eq: "f \ C (ccpo.fixp (fun_lub Sup) (fun_ord (\)) (\f. U (F (C f))))" and inverse2: "\f. U (C f) = f" and step: "\f x y. \ \x y. y \ U f x \ P x y; y \ U (F f) x \ \ P x y" and enforce_variable_ordering: "x = x" and elem: "y \ U f x" shows "P x y" using step elem set.fixp_induct_uc[of U F C, OF mono eq inverse2 set_admissible, of P] by blast lemma fixp_Sup_le: defines "le \ ((\) :: _ :: complete_lattice \ _)" shows "ccpo.fixp Sup le = ccpo_class.fixp" proof - have "class.ccpo Sup le (<)" unfolding le_def by unfold_locales thus ?thesis by(simp add: ccpo.fixp_def fixp_def ccpo.iterates_def iterates_def ccpo.iteratesp_def iteratesp_def fun_eq_iff le_def) qed lemma fun_ord_le: "fun_ord (\) = (\)" by(auto simp add: fun_ord_def fun_eq_iff le_fun_def) -lemma monotone_le_le: "monotone (\) (\) = mono" -by(simp add: monotone_def[abs_def] mono_def[abs_def]) - lemma fixp_induct_set: fixes F :: "'c \ 'c" and U :: "'c \ 'b \ 'a set" and C :: "('b \ 'a set) \ 'c" and P :: "'b \ 'a \ bool" and x and y assumes mono: "\x. mono_set (\f. U (F (C f)) x)" and eq: "f \ C (ccpo.fixp (fun_lub Sup) (fun_ord (\)) (\f. U (F (C f))))" and inverse2: "\f. U (C f) = f" and step: "\f' x y. \ \x. U f' x = U f' x; y \ U (F (C (inf (U f) (\x. {y. P x y})))) x \ \ P x y" \ \partial\_function requires a quantifier over f', so let's have a fake one\ and elem: "y \ U f x" shows "P x y" proof - from mono have mono': "mono (\f. U (F (C f)))" - by(simp add: fun_ord_le monotone_le_le mono_def le_fun_def) + by(simp add: fun_ord_le mono_def le_fun_def) hence eq': "f \ C (lfp (\f. U (F (C f))))" using eq unfolding fun_ord_le fun_lub_Sup fixp_Sup_le by(simp add: lfp_eq_fixp) let ?f = "C (lfp (\f. U (F (C f))))" have step': "\x y. \ y \ U (F (C (inf (U ?f) (\x. {y. P x y})))) x \ \ P x y" unfolding eq'[symmetric] by(rule step[OF refl]) let ?P = "\x. {y. P x y}" from mono' have "lfp (\f. U (F (C f))) \ ?P" by(rule lfp_induct)(auto intro!: le_funI step' simp add: inverse2) with elem show ?thesis by(subst (asm) eq')(auto simp add: inverse2 le_fun_def) qed declaration \Partial_Function.init "set" @{term set.fixp_fun} @{term set.mono_body} @{thm set.fixp_rule_uc} @{thm set.fixp_induct_uc} (SOME @{thm fixp_induct_set})\ lemma [partial_function_mono]: shows insert_mono: "mono_set A \ mono_set (\f. insert x (A f))" and UNION_mono: "\mono_set B; \y. mono_set (\f. C y f)\ \ mono_set (\f. \y\B f. C y f)" and set_bind_mono: "\mono_set B; \y. mono_set (\f. C y f)\ \ mono_set (\f. Set.bind (B f) (\y. C y f))" and Un_mono: "\ mono_set A; mono_set B \ \ mono_set (\f. A f \ B f)" and Int_mono: "\ mono_set A; mono_set B \ \ mono_set (\f. A f \ B f)" and Diff_mono1: "mono_set A \ mono_set (\f. A f - X)" and image_mono: "mono_set A \ mono_set (\f. g ` A f)" and vimage_mono: "mono_set A \ mono_set (\f. g -` A f)" unfolding bind_UNION by(fast intro!: monotoneI dest: monotoneD)+ partial_function (set) test :: "'a list \ nat \ bool \ int set" where "test xs i j = insert 4 (test [] 0 j \ test [] 1 True \ test [] 2 False - {5} \ uminus ` test [undefined] 0 True \ uminus -` test [] 1 False)" interpretation coset: partial_function_definitions "(\)" Inter by(rule complete_lattice.lattice_partial_function_definition[OF dual_complete_lattice]) lemma fun_lub_Inf: "fun_lub Inf = (Inf :: _ \ _ :: complete_lattice)" by(auto simp add: fun_lub_def fun_eq_iff Inf_fun_def intro: Inf_eqI INF_lower INF_greatest) lemma fun_ord_ge: "fun_ord (\) = (\)" by(auto simp add: fun_ord_def fun_eq_iff le_fun_def) lemma coset_admissible: "coset.admissible (\f :: 'a \ 'b set. \x y. P x y \ y \ f x)" by(rule ccpo.admissibleI)(auto simp add: fun_lub_Inf) abbreviation "mono_coset \ monotone (fun_ord (\)) (\)" lemma gfp_eq_fixp: fixes f :: "'a :: complete_lattice \ 'a" assumes f: "monotone (\) (\) f" shows "gfp f = ccpo.fixp Inf (\) f" proof (rule antisym) from f have f': "mono f" by(simp add: mono_def monotone_def) interpret ccpo Inf "(\)" "mk_less (\) :: 'a \ _" by(rule ccpo)(rule complete_lattice.lattice_partial_function_definition[OF dual_complete_lattice]) show "ccpo.fixp Inf (\) f \ gfp f" by(rule gfp_upperbound)(subst fixp_unfold[OF f], rule order_refl) show "gfp f \ ccpo.fixp Inf (\) f" by(rule fixp_lowerbound[OF f])(subst gfp_unfold[OF f'], rule order_refl) qed lemma fixp_coinduct_set: fixes F :: "'c \ 'c" and U :: "'c \ 'b \ 'a set" and C :: "('b \ 'a set) \ 'c" and P :: "'b \ 'a \ bool" and x and y assumes mono: "\x. mono_coset (\f. U (F (C f)) x)" and eq: "f \ C (ccpo.fixp (fun_lub Inter) (fun_ord (\)) (\f. U (F (C f))))" and inverse2: "\f. U (C f) = f" and step: "\f' x y. \ \x. U f' x = U f' x; \ P x y \ \ y \ U (F (C (sup (\x. {y. \ P x y}) (U f)))) x" \ \partial\_function requires a quantifier over f', so let's have a fake one\ and elem: "y \ U f x" shows "P x y" using elem proof(rule contrapos_np) have mono': "monotone (\) (\) (\f. U (F (C f)))" and mono'': "mono (\f. U (F (C f)))" using mono by(simp_all add: monotone_def fun_ord_def le_fun_def mono_def) hence eq': "U f = gfp (\f. U (F (C f)))" by(subst eq)(simp add: fun_lub_Inf fun_ord_ge gfp_eq_fixp inverse2) let ?P = "\x. {y. \ P x y}" have "?P \ gfp (\f. U (F (C f)))" using mono'' by(rule coinduct)(auto intro!: le_funI dest: step[OF refl] simp add: eq') moreover assume "\ P x y" ultimately show "y \ U f x" by(auto simp add: le_fun_def eq') qed declaration \Partial_Function.init "coset" @{term coset.fixp_fun} @{term coset.mono_body} @{thm coset.fixp_rule_uc} @{thm coset.fixp_induct_uc} (SOME @{thm fixp_coinduct_set})\ abbreviation "mono_set' \ monotone (fun_ord (\)) (\)" lemma [partial_function_mono]: shows insert_mono': "mono_set' A \ mono_set' (\f. insert x (A f))" and UNION_mono': "\mono_set' B; \y. mono_set' (\f. C y f)\ \ mono_set' (\f. \y\B f. C y f)" and set_bind_mono': "\mono_set' B; \y. mono_set' (\f. C y f)\ \ mono_set' (\f. Set.bind (B f) (\y. C y f))" and Un_mono': "\ mono_set' A; mono_set' B \ \ mono_set' (\f. A f \ B f)" and Int_mono': "\ mono_set' A; mono_set' B \ \ mono_set' (\f. A f \ B f)" unfolding bind_UNION by(fast intro!: monotoneI dest: monotoneD)+ context begin private partial_function (coset) test2 :: "nat \ nat set" where "test2 x = insert x (test2 (Suc x))" private lemma test2_coinduct: assumes "P x y" and *: "\x y. P x y \ y = x \ (P (Suc x) y \ y \ test2 (Suc x))" shows "y \ test2 x" using \P x y\ apply(rule contrapos_pp) apply(erule test2.raw_induct[rotated]) apply(simp add: *) done end end diff --git a/thys/Lp/Lp.thy b/thys/Lp/Lp.thy --- a/thys/Lp/Lp.thy +++ b/thys/Lp/Lp.thy @@ -1,2246 +1,2248 @@ (* Author: Sébastien Gouëzel sebastien.gouezel@univ-rennes1.fr License: BSD *) theory Lp imports Functional_Spaces begin text \The material in this file is essentially of analytic nature. However, one of the central proofs (the proof of Holder inequality below) uses a probability space, and Jensen's inequality there. Hence, we need to import \verb+Probability+. Moreover, we use several lemmas from \verb+SG_Library_Complement+.\ section \Conjugate exponents\ text \Two numbers $p$ and $q$ are \emph{conjugate} if $1/p + 1/q = 1$. This relation keeps appearing in the theory of $L^p$ spaces, as the dual of $L^p$ is $L^q$ where $q$ is the conjugate of $p$. This relation makes sense for real numbers, but also for ennreals (where the case $p=1$ and $q=\infty$ is most important). Unfortunately, manipulating the previous relation with ennreals is tedious as there is no good simproc involving addition and division there. To mitigate this difficulty, we prove once and for all most useful properties of such conjugates exponents in this paragraph.\ lemma Lp_cases_1_PInf: assumes "p \ (1::ennreal)" obtains (gr) p2 where "p = ennreal p2" "p2 > 1" "p > 1" | (one) "p = 1" | (PInf) "p = \" using assms by (metis (full_types) antisym_conv ennreal_cases ennreal_le_1 infinity_ennreal_def not_le) lemma Lp_cases: obtains (real_pos) p2 where "p = ennreal p2" "p2 > 0" "p > 0" | (zero) "p = 0" | (PInf) "p = \" by (metis enn2real_positive_iff ennreal_enn2real_if infinity_ennreal_def not_gr_zero top.not_eq_extremum) definition "conjugate_exponent p = 1 + 1/(p-1)" lemma conjugate_exponent_real: assumes "p > (1::real)" shows "1/p + 1/(conjugate_exponent p) = 1" "conjugate_exponent p > 1" "conjugate_exponent(conjugate_exponent p) = p" "(p-1) * conjugate_exponent p = p" "p - p / conjugate_exponent p = 1" unfolding conjugate_exponent_def using assms by (auto simp add: algebra_simps divide_simps) lemma conjugate_exponent_real_iff: assumes "p > (1::real)" shows "q = conjugate_exponent p \ (1/p + 1/q = 1)" unfolding conjugate_exponent_def using assms by (auto simp add: algebra_simps divide_simps) lemma conjugate_exponent_real_2 [simp]: "conjugate_exponent (2::real) = 2" unfolding conjugate_exponent_def by (auto simp add: algebra_simps divide_simps) lemma conjugate_exponent_realI: assumes "p > (0::real)" "q > 0" "1/p + 1/q = 1" shows "p > 1" "q = conjugate_exponent p" "q > 1" "p = conjugate_exponent q" unfolding conjugate_exponent_def using assms apply (auto simp add: algebra_simps divide_simps) apply (metis assms(3) divide_less_eq_1_pos less_add_same_cancel1 zero_less_divide_1_iff) using mult_less_cancel_left_pos by fastforce lemma conjugate_exponent_real_ennreal: assumes "p> (1::real)" shows "conjugate_exponent(ennreal p) = ennreal(conjugate_exponent p)" unfolding conjugate_exponent_def using assms by (auto, metis diff_gt_0_iff_gt divide_ennreal ennreal_1 ennreal_minus zero_le_one) lemma conjugate_exponent_ennreal_1_2_PInf [simp]: "conjugate_exponent (1::ennreal) = \" "conjugate_exponent (\::ennreal) = 1" "conjugate_exponent (\::ennreal) = 1" "conjugate_exponent (2::ennreal) = 2" using conjugate_exponent_real_ennreal[of 2] by (auto simp add: conjugate_exponent_def) lemma conjugate_exponent_ennreal: assumes "p \ (1::ennreal)" shows "1/p + 1/(conjugate_exponent p) = 1" "conjugate_exponent p \ 1" "conjugate_exponent(conjugate_exponent p) = p" proof - have "(1/p + 1/(conjugate_exponent p) = 1) \ (conjugate_exponent p \ 1) \ conjugate_exponent(conjugate_exponent p) = p" using \p \ 1\ proof (cases rule: Lp_cases_1_PInf) case (gr p2) then have *: "conjugate_exponent p = ennreal (conjugate_exponent p2)" using conjugate_exponent_real_ennreal[OF \p2 > 1\] by auto have a: "conjugate_exponent p \ 1" using * conjugate_exponent_real[OF \p2 > 1\] by auto have b: "conjugate_exponent(conjugate_exponent p) = p" using conjugate_exponent_real(3)[OF \p2 > 1\] conjugate_exponent_real_ennreal[OF \p2 > 1\] conjugate_exponent_real_ennreal[OF conjugate_exponent_real(2)[OF \p2 > 1\]] unfolding * \p = ennreal p2\ by auto have "1 / p + 1 / conjugate_exponent p = ennreal(1/p2 + 1/(conjugate_exponent p2))" unfolding * unfolding \p = ennreal p2\ using conjugate_exponent_real(2)[OF \p2 > 1\] \p2 > 1\ apply (subst ennreal_plus, auto) apply (subst divide_ennreal[symmetric], auto) using divide_ennreal_def inverse_ennreal inverse_eq_divide by auto then have c: "1 / p + 1 / conjugate_exponent p = 1" using conjugate_exponent_real[OF \p2 > 1\] by auto show ?thesis using a b c by simp qed (auto) then show "1/p + 1/(conjugate_exponent p) = 1" "conjugate_exponent p \ 1" "conjugate_exponent(conjugate_exponent p) = p" by auto qed lemma conjugate_exponent_ennreal_iff: assumes "p \ (1::ennreal)" shows "q = conjugate_exponent p \ (1/p + 1/q = 1)" using conjugate_exponent_ennreal[OF assms] by (auto, metis ennreal_add_diff_cancel_left ennreal_add_eq_top ennreal_top_neq_one one_divide_one_divide_ennreal) lemma conjugate_exponent_ennrealI: assumes "1/p + 1/q = (1::ennreal)" shows "p \ 1" "q \ 1" "p = conjugate_exponent q" "q = conjugate_exponent p" proof - have "1/p \ 1" using assms using le_iff_add by fastforce then show "p \ 1" by (metis assms divide_ennreal_def ennreal_add_eq_top ennreal_divide_self ennreal_divide_zero ennreal_le_epsilon ennreal_one_neq_top mult.left_neutral mult_left_le zero_le) then show "q = conjugate_exponent p" using conjugate_exponent_ennreal_iff assms by auto then show "q \ 1" using conjugate_exponent_ennreal[OF \p \ 1\] by auto show "p = conjugate_exponent q" using conjugate_exponent_ennreal_iff[OF \q\1\, of p] assms by (simp add: add.commute) qed section \Convexity inequalities and integration\ text \In this paragraph, we describe the basic inequalities relating the integral of a function and of its $p$-th power, for $p > 0$. These inequalities imply in particular that the $L^p$ norm satisfies the triangular inequality, a feature we will need when defining the $L^p$ spaces below. In particular, we prove the Hölder and Minkowski inequalities. The Hölder inequality, especially, is the basis of all further inequalities for $L^p$ spaces. \ lemma (in prob_space) bound_L1_Lp: assumes "p \ (1::real)" "f \ borel_measurable M" "integrable M (\x. \f x\ powr p)" shows "integrable M f" "abs(\x. f x \M) powr p \ (\x. \f x\ powr p \M)" "abs(\x. f x \M) \ (\x. \f x\ powr p \M) powr (1/p)" proof - have *: "norm x \ 1 + (norm x) powr p" for x::real apply (cases "norm x \ 1") apply (meson le_add_same_cancel1 order.trans powr_ge_pzero) apply (metis add_le_same_cancel2 assms(1) less_le_trans linear not_less not_one_le_zero powr_le_cancel_iff powr_one_gt_zero_iff) done show *: "integrable M f" apply (rule Bochner_Integration.integrable_bound[of _ "\x. 1 + \f x\ powr p"], auto simp add: assms) using * by auto show "abs(\x. f x \M) powr p \ (\x. \f x\ powr p \M)" by (rule jensens_inequality[OF * _ _ assms(3) convex_abs_powr[OF \p \ 1\]], auto) then have "(abs(\x. f x \M) powr p) powr (1/p) \ (\x. \f x\ powr p \M) powr (1/p)" using assms(1) powr_mono2 by auto then show "abs(\x. f x \M) \ (\x. \f x\ powr p \M) powr (1/p)" using \p \ 1\ by (auto simp add: powr_powr) qed theorem Holder_inequality: assumes "p > (0::real)" "q > 0" "1/p + 1/q = 1" and [measurable]: "f \ borel_measurable M" "g \ borel_measurable M" "integrable M (\x. \f x\ powr p)" "integrable M (\x. \g x\ powr q)" shows "integrable M (\x. f x * g x)" "(\x. \f x * g x\ \M) \ (\x. \f x\ powr p \M) powr (1/p) * (\x. \g x\ powr q \M) powr (1/q)" "abs(\x. f x * g x \M) \ (\x. \f x\ powr p \M) powr (1/p) * (\x. \g x\ powr q \M) powr (1/q)" proof - have "p > 1" using conjugate_exponent_realI(1)[OF \p>0\ \q>0\ \1/p+1/q=1\]. have *: "x * y \ x powr p + y powr q" if "x \ 0" "y \ 0" for x y proof - have "x * y = (x powr p) powr (1/p) * (y powr q) powr (1/q)" using \p > 0\ \q > 0\ powr_powr that(1) that(2) by auto also have "... \ (max (x powr p) (y powr q)) powr (1/p) * (max (x powr p) (y powr q)) powr (1/q)" apply (rule mult_mono, auto) using assms(1) assms(2) powr_mono2 by auto also have "... = max (x powr p) (y powr q)" by (metis max_def mult.right_neutral powr_add powr_powr assms(3)) also have "... \ x powr p + y powr q" by auto finally show ?thesis by simp qed show [simp]: "integrable M (\x. f x * g x)" apply (rule Bochner_Integration.integrable_bound[of _ "\x. \f x\ powr p + \g x\ powr q"], auto) by (rule Bochner_Integration.integrable_add, auto simp add: assms * abs_mult) text \The proof of the main inequality is done by applying the inequality $(\int |h| d\mu \leq \int |h|^p d\mu)^{1/p}$ to the right function $h$ in the right probability space. One should take $h = f \cdot |g|^{1-q}$, and $d\mu = |g|^q dM / I$, where $I = \int |g|^q$. This readily gives the result.\ show *: "(\x. \f x * g x\ \M) \ (\x. \f x\ powr p \M) powr (1/p) * (\x. \g x\ powr q \M) powr (1/q)" proof (cases "(\x. \g x\ powr q \M) = 0") case True then have "AE x in M. \g x\ powr q = 0" by (subst integral_nonneg_eq_0_iff_AE[symmetric], auto simp add: assms) then have *: "AE x in M. f x * g x = 0" using \q > 0\ by auto have "(\x. \f x * g x\ \M) = (\x. 0 \M)" apply (rule integral_cong_AE) using * by auto then show ?thesis by auto next case False moreover have "(\x. \g x\ powr q \M) \ (\x. 0 \M)" by (rule integral_mono, auto simp add: assms) ultimately have *: "(\x. \g x\ powr q \M) > 0" by (simp add: le_less) define I where "I = (\x. \g x\ powr q \M)" have [simp]: "I > 0" unfolding I_def using * by auto define M2 where "M2 = density M (\x. \g x\ powr q / I)" interpret prob_space M2 apply (standard, unfold M2_def, auto, subst emeasure_density, auto) apply (subst divide_ennreal[symmetric], auto, subst nn_integral_divide, auto) apply (subst nn_integral_eq_integral, auto simp add: assms, unfold I_def) using * by auto have [simp]: "p \ 1" "p \ 0" using \p > 1\ by auto have A: "q + (1 - q) * p = 0" using assms by (auto simp add: divide_simps algebra_simps) have B: "1 - 1/p = 1/q" using \1/p + 1/q = 1\ by auto define f2 where "f2 = (\x. f x * indicator {y\ space M. g y \ 0} x)" have [measurable]: "f2 \ borel_measurable M" unfolding f2_def by auto define h where "h = (\x. \f2 x\ * \g x\ powr (1-q))" have [measurable]: "h \ borel_measurable M" unfolding h_def by auto have [measurable]: "h \ borel_measurable M2" unfolding M2_def by auto have Eq: "(\g x\ powr q / I) *\<^sub>R \h x\ powr p = \f2 x\ powr p / I" for x apply (insert \I>0\, auto simp add: divide_simps, unfold h_def) apply (auto simp add: divide_nonneg_pos divide_simps powr_mult powr_powr powr_add[symmetric] A) unfolding f2_def by auto have "integrable M2 (\x. \h x\ powr p)" unfolding M2_def apply (subst integrable_density, simp, simp, simp add: divide_simps) apply (subst Eq, rule integrable_divide, rule Bochner_Integration.integrable_bound[of _ "\x. \f x\ powr p"], unfold f2_def) by (unfold indicator_def, auto simp add: \integrable M (\x. \f x\ powr p)\) then have "integrable M2 (\x. \h x\)" by (metis bound_L1_Lp(1) \random_variable borel h\ \p > 1\ integrable_abs le_less) have "(\x. \h x\ powr p \M2) = (\x. (\g x\ powr q / I) *\<^sub>R (\h x\ powr p) \M)" unfolding M2_def by (rule integral_density[of "\x. \h x\ powr p" M "\x. \g x\ powr q / I"], auto simp add: divide_simps) also have "... = (\x. \f2 x\ powr p / I \M)" apply (rule Bochner_Integration.integral_cong) using Eq by auto also have "... \ (\x. \f x\ powr p / I \M)" apply (rule integral_mono', rule integrable_divide[OF \integrable M (\x. \f x\ powr p)\]) unfolding f2_def indicator_def using \I > 0\ by (auto simp add: divide_simps) finally have C: "(\x. \h x\ powr p \M2) \ (\x. \f x\ powr p / I \M)" by simp have "(\x. \f x * g x\ \M) / I = (\x. \f x * g x\ / I \M)" by auto also have "... = (\x. \f2 x * g x\ / I \M)" by (auto simp add: divide_simps, rule Bochner_Integration.integral_cong, unfold f2_def indicator_def, auto) also have "... = (\x. \h x\ \M2)" apply (unfold M2_def, subst integral_density, simp, simp, simp add: divide_simps) by (rule Bochner_Integration.integral_cong, unfold h_def, auto simp add: divide_simps algebra_simps powr_add[symmetric] abs_mult) also have "... \ abs (\x. \h x\ \M2)" by auto also have "... \ (\x. abs(\h x\) powr p \M2) powr (1/p)" apply (rule bound_L1_Lp(3)[of p "\x. \h x\"]) by (auto simp add: \integrable M2 (\x. \h x\ powr p)\) also have "... \ (\x. \f x\ powr p / I \M) powr (1/p)" by (rule powr_mono2, insert C, auto) also have "... \ ((\x. \f x\ powr p \M) / I) powr (1/p)" apply (rule powr_mono2, auto simp add: divide_simps) using \p \ 0\ by auto also have "... = (\x. \f x\ powr p \M) powr (1/p) * I powr(-1/p)" by (auto simp add: less_imp_le powr_divide powr_minus_divide) finally have "(\x. \f x * g x\ \M) \ (\x. \f x\ powr p \M) powr (1/p) * I * I powr(-1/p)" by (auto simp add: divide_simps algebra_simps) also have "... = (\x. \f x\ powr p \M) powr (1/p) * I powr (1-1/p)" by (auto simp add: powr_mult_base less_imp_le) also have "... = (\x. \f x\ powr p \M) powr (1/p) * (\x. \g x\ powr q \M) powr (1/q)" unfolding I_def using B by auto finally show ?thesis by simp qed have "abs(\x. f x * g x \M) \ (\x. \f x * g x\ \M)" by auto then show "abs(\x. f x * g x \M) \ (\x. \f x\ powr p \M) powr (1/p) * (\x. \g x\ powr q \M) powr (1/q)" using * by linarith qed theorem Minkowski_inequality: assumes "p \ (1::real)" and [measurable, simp]: "f \ borel_measurable M" "g \ borel_measurable M" "integrable M (\x. \f x\ powr p)" "integrable M (\x. \g x\ powr p)" shows "integrable M (\x. \f x + g x\ powr p)" "(\x. \f x + g x\ powr p \M) powr (1/p) \ (\x. \f x\ powr p \M) powr (1/p) + (\x. \g x\ powr p \M) powr (1/p)" proof - have *: "\x + y\ powr p \ 2 powr p * (\x\ powr p + \y\ powr p)" for x y::real proof - have "\x + y\ \ \x\ + \y\" by auto also have "... \ (max \x\ \y\) + max \x\ \y\" by auto also have "... = 2 * max \x\ \y\" by auto finally have "\x + y\ powr p \ (2 * max \x\ \y\) powr p" using powr_mono2 \p \ 1\ by auto also have "... = 2 powr p * (max \x\ \y\) powr p" using powr_mult by auto also have "... \ 2 powr p * (\x\ powr p + \y\ powr p)" unfolding max_def by auto finally show ?thesis by simp qed show [simp]: "integrable M (\x. \f x + g x\ powr p)" by (rule Bochner_Integration.integrable_bound[of _ "\x. 2 powr p * (\f x\ powr p + \g x\ powr p)"], auto simp add: *) show "(\x. \f x + g x\ powr p \M) powr (1/p) \ (\x. \f x\ powr p \M) powr (1/p) + (\x. \g x\ powr p \M) powr (1/p)" proof (cases "p=1") case True then show ?thesis apply (auto, subst Bochner_Integration.integral_add[symmetric], insert assms(4) assms(5), simp, simp) by (rule integral_mono', auto) next case False then have [simp]: "p > 1" "p \ 1" "p > 0" "p \ 0" using assms(1) by auto define q where "q = conjugate_exponent p" have [simp]: "q > 1" "q > 0" "1/p + 1/q = 1" "(p-1) * q = p" unfolding q_def using conjugate_exponent_real[OF \p>1\] by auto then have [simp]: "(z powr (p-1)) powr q = z powr p" for z by (simp add: powr_powr) have "(\x. \f x + g x\ powr p \M) = (\x. \f x + g x\ * \f x + g x\ powr (p-1) \M)" by (subst powr_mult_base, auto) also have "... \ (\x. \f x\ * \f x + g x\ powr (p-1) + \g x\ * \f x + g x\ powr (p-1) \M)" apply (rule integral_mono', rule Bochner_Integration.integrable_add) apply (rule Holder_inequality(1)[of p q], auto) apply (rule Holder_inequality(1)[of p q], auto) by (metis abs_ge_zero abs_triangle_ineq comm_semiring_class.distrib le_less mult_mono' powr_ge_pzero) also have "... = (\x. \f x\ * \f x + g x\ powr (p-1) \M) + (\x. \g x\ * \f x + g x\ powr (p-1) \M)" apply (rule Bochner_Integration.integral_add) by (rule Holder_inequality(1)[of p q], auto)+ also have "... \ abs (\x. \f x\ * \f x + g x\ powr (p-1) \M) + abs (\x. \g x\ * \f x + g x\ powr (p-1) \M)" by auto also have "... \ (\x. abs(\f x\) powr p \M) powr (1/p) * (\x. abs(\f x + g x\ powr (p-1)) powr q \M) powr (1/q) + (\x. abs(\g x\) powr p \M) powr (1/p) * (\x. abs(\f x + g x\ powr (p-1)) powr q \M) powr (1/q)" apply (rule add_mono) apply (rule Holder_inequality(3)[of p q], simp, simp, simp, simp, simp, simp, simp) apply (rule Holder_inequality(3)[of p q], simp, simp, simp, simp, simp, simp, simp) done also have "... = (\x. \f x + g x\ powr p \M) powr (1/q) * ((\x. abs(\f x\) powr p \M) powr (1/p) + (\x. abs(\g x\) powr p \M) powr (1/p))" by (auto simp add: algebra_simps) finally have *: "(\x. \f x + g x\ powr p \M) \ (\x. \f x + g x\ powr p \M) powr (1/q) * ((\x. abs(\f x\) powr p \M) powr (1/p) + (\x. abs(\g x\) powr p \M) powr (1/p))" by simp show ?thesis proof (cases "(\x. \f x + g x\ powr p \M) = 0") case True then show ?thesis by auto next case False then have **: "(\x. \f x + g x\ powr p \M) powr (1/q) > 0" by auto have "(\x. \f x + g x\ powr p \M) powr (1/q) * (\x. \f x + g x\ powr p \M) powr (1/p) = (\x. \f x + g x\ powr p \M)" by (auto simp add: powr_add[symmetric] add.commute) then have "(\x. \f x + g x\ powr p \M) powr (1/q) * (\x. \f x + g x\ powr p \M) powr (1/p) \ (\x. \f x + g x\ powr p \M) powr (1/q) * ((\x. abs(\f x\) powr p \M) powr (1/p) + (\x. abs(\g x\) powr p \M) powr (1/p))" using * by auto then show ?thesis using ** by auto qed qed qed text \When $p<1$, the function $x \mapsto |x|^p$ is not convex any more. Hence, the $L^p$ ``norm'' is not a norm any more, but a quasinorm. This is proved using a different convexity argument, as follows.\ theorem Minkowski_inequality_le_1: assumes "p > (0::real)" "p \ 1" and [measurable, simp]: "f \ borel_measurable M" "g \ borel_measurable M" "integrable M (\x. \f x\ powr p)" "integrable M (\x. \g x\ powr p)" shows "integrable M (\x. \f x + g x\ powr p)" "(\x. \f x + g x\ powr p \M) powr (1/p) \ 2 powr (1/p-1) * (\x. \f x\ powr p \M) powr (1/p) + 2 powr (1/p-1) * (\x. \g x\ powr p \M) powr (1/p)" proof - have *: "\a + b\ powr p \ \a\ powr p + \b\ powr p" for a b using x_plus_y_p_le_xp_plus_yp[OF \p > 0\ \p \ 1\, of "\a\" "\b\"] by (auto, meson abs_ge_zero abs_triangle_ineq assms(1) le_less order.trans powr_mono2) show "integrable M (\x. \f x + g x\ powr p)" by (rule Bochner_Integration.integrable_bound[of _ "\x. \f x\ powr p + \g x\ powr p"], auto simp add: *) have "(\x. \f x + g x\ powr p \M) powr (1/p) \ (\x. \f x\ powr p + \g x\ powr p \M) powr (1/p)" by (rule powr_mono2, simp add: \p > 0\ less_imp_le, simp, rule integral_mono', auto simp add: *) also have "... = 2 powr (1/p) * (((\x. \f x\ powr p \M) + (\x. \g x\ powr p \M)) / 2) powr (1/p)" by (auto simp add: powr_mult[symmetric] add_divide_distrib) also have "... \ 2 powr (1/p) * (((\x. \f x\ powr p \M) powr (1/p) + (\x. \g x\ powr p \M) powr (1/p)) / 2)" apply (rule mult_mono, simp, rule convex_on_mean_ineq[OF convex_powr[of "1/p"]]) using \p \ 1\ \p > 0\ by auto also have "... = 2 powr (1/p - 1) * ((\x. \f x\ powr p \M) powr (1/p) + (\x. \g x\ powr p \M) powr (1/p))" by (simp add: powr_diff) finally show "(\x. \f x + g x\ powr p \M) powr (1/p) \ 2 powr (1/p-1) * (\x. \f x\ powr p \M) powr (1/p) + 2 powr (1/p-1) * (\x. \g x\ powr p \M) powr (1/p)" by (auto simp add: algebra_simps) qed section \$L^p$ spaces\ text \We define $L^p$ spaces by giving their defining quasinorm. It is a norm for $p\in [1, \infty]$, and a quasinorm for $p \in (0,1)$. The construction of a quasinorm from a formula only makes sense if this formula is indeed a quasinorm, i.e., it is homogeneous and satisfies the triangular inequality with the given multiplicative defect. Thus, we have to show that this is indeed the case to be able to use the definition.\ definition Lp_space::"ennreal \ 'a measure \ ('a \ real) quasinorm" where "Lp_space p M = ( if p = 0 then quasinorm_of (1, (\f. if (f \ borel_measurable M) then 0 else \)) else if p < \ then quasinorm_of ( if p < 1 then 2 powr (1/enn2real p - 1) else 1, (\f. if (f \ borel_measurable M \ integrable M (\x. \f x\ powr (enn2real p))) then (\x. \f x\ powr (enn2real p) \M) powr (1/(enn2real p)) else (\::ennreal))) else quasinorm_of (1, (\f. if f \ borel_measurable M then esssup M (\x. ereal \f x\) else (\::ennreal))))" abbreviation "\ == Lp_space" subsection \$L^\infty$\ text \Let us check that, for $L^\infty$, the above definition makes sense.\ lemma L_infinity: "eNorm (\ \ M) f = (if f \ borel_measurable M then esssup M (\x. ereal \f x\) else (\::ennreal))" "defect (\ \ M) = 1" proof - have T: "esssup M (\x. ereal \(f + g) x\) \ e2ennreal (esssup M (\x. ereal \f x\)) + esssup M (\x. ereal \g x\)" if [measurable]: "f \ borel_measurable M" "g \ borel_measurable M" for f g proof (cases "emeasure M (space M) = 0") case True then have "e2ennreal (esssup M (\x. ereal \(f + g) x\)) = 0" using esssup_zero_space[OF True] by (simp add: e2ennreal_neg) then show ?thesis by simp next case False have *: "esssup M (\x. \h x\) \ 0" for h::"'a \ real" proof - have "esssup M (\x. 0) \ esssup M (\x. \h x\)" by (rule esssup_mono, auto) then show ?thesis using esssup_const[OF False, of "0::ereal"] by simp qed have "esssup M (\x. ereal \(f + g) x\) \ esssup M (\x. ereal \f x\ + ereal \g x\)" by (rule esssup_mono, auto simp add: plus_fun_def) also have "... \ esssup M (\x. ereal \f x\) + esssup M (\x. ereal \g x\)" by (rule esssup_add) finally show ?thesis using * by (simp add: e2ennreal_mono eq_onp_def plus_ennreal.abs_eq) qed have H: "esssup M (\x. ereal \(c *\<^sub>R f) x\) \ ennreal \c\ * esssup M (\x. ereal \f x\)" if "c \ 0" for f c proof - have "abs c > 0" "ereal \c\ \ 0" using that by auto have *: "esssup M (\x. abs(c *\<^sub>R f x)) = abs c * esssup M (\x. \f x\)" apply (subst esssup_cmult[OF \abs c > 0\, of M "\x. ereal \f x\", symmetric]) using times_ereal.simps(1) by (auto simp add: abs_mult) show ?thesis unfolding e2ennreal_mult[OF \ereal \c\ \ 0\] * scaleR_fun_def by simp qed have "esssup M (\x. ereal 0) \ 0" using esssup_I by auto then have Z: "e2ennreal (esssup M (\x. ereal 0)) = 0" using e2ennreal_neg by auto have *: "quasinorm_on (borel_measurable M) 1 (\(f::'a\real). e2ennreal(esssup M (\x. ereal \f x\)))" apply (rule quasinorm_onI) using T H Z by auto have **: "quasinorm_on UNIV 1 (\(f::'a\real). if f \ borel_measurable M then e2ennreal(esssup M (\x. ereal \f x\)) else \)" by (rule extend_quasinorm[OF *]) show "eNorm (\ \ M) f = (if f \ borel_measurable M then e2ennreal(esssup M (\x. \f x\)) else \)" "defect (\ \ M) = 1" unfolding Lp_space_def using quasinorm_of[OF **] by auto qed lemma L_infinity_space: "space\<^sub>N (\ \ M) = {f \ borel_measurable M. \C. AE x in M. \f x\ \ C}" proof (auto simp del: infinity_ennreal_def) fix f assume H: "f \ space\<^sub>N (\ \ M)" then show "f \ borel_measurable M" unfolding space\<^sub>N_def using L_infinity(1)[of M] top.not_eq_extremum by fastforce then have *: "esssup M (\x. \f x\) < \" using H unfolding space\<^sub>N_def L_infinity(1)[of M] by (auto simp add: e2ennreal_infty) define C where "C = real_of_ereal(esssup M (\x. \f x\))" have "AE x in M. ereal \f x\ \ ereal C" proof (cases "emeasure M (space M) = 0") case True then show ?thesis using emeasure_0_AE by simp next case False then have "esssup M (\x. \f x\) \ 0" using esssup_mono[of "\x. 0" M "(\x. \f x\)"] esssup_const[OF False, of "0::ereal"] by auto then have "esssup M (\x. \f x\) = ereal C" unfolding C_def using * ereal_real by auto then show ?thesis using esssup_AE[of "(\x. ereal \f x\)" M] by simp qed then have "AE x in M. \f x\ \ C" by auto then show "\C. AE x in M. \f x\ \ C" by blast next fix f::"'a \ real" and C::real assume H: "f \ borel_measurable M" "AE x in M. \f x\ \ C" then have "esssup M (\x. \f x\) \ C" using esssup_I by auto then have "eNorm (\ \ M) f \ C" unfolding L_infinity(1) using H(1) by auto (metis e2ennreal_ereal e2ennreal_mono) then show "f \ space\<^sub>N (\ \ M)" using spaceN_iff le_less_trans by fastforce qed lemma L_infinity_zero_space: "zero_space\<^sub>N (\ \ M) = {f \ borel_measurable M. AE x in M. f x = 0}" proof (auto simp del: infinity_ennreal_def) fix f assume H: "f \ zero_space\<^sub>N (\ \ M)" then show "f \ borel_measurable M" unfolding zero_space\<^sub>N_def using L_infinity(1)[of M] top.not_eq_extremum by fastforce then have *: "e2ennreal(esssup M (\x. \f x\)) = 0" using H unfolding zero_space\<^sub>N_def using L_infinity(1)[of M] e2ennreal_infty by auto then have "esssup M (\x. \f x\) \ 0" by (metis e2ennreal_infty e2ennreal_mult ennreal_top_neq_zero ereal_mult_infty leI linear mult_zero_left) then have "f x = 0" if "ereal \f x\ \ esssup M (\x. \f x\)" for x using that order.trans by fastforce then show "AE x in M. f x = 0" using esssup_AE[of "\x. ereal \f x\" M] by auto next fix f::"'a \ real" assume H: "f \ borel_measurable M" "AE x in M. f x = 0" then have "esssup M (\x. \f x\) \ 0" using esssup_I by auto then have "eNorm (\ \ M) f = 0" unfolding L_infinity(1) using H(1) by (simp add: e2ennreal_neg) then show "f \ zero_space\<^sub>N (\ \ M)" using zero_spaceN_iff by auto qed lemma L_infinity_AE_ebound: "AE x in M. ennreal \f x\ \ eNorm (\ \ M) f" proof (cases "f \ borel_measurable M") case False then have "eNorm (\ \ M) f = \" unfolding L_infinity(1) by auto then show ?thesis by simp next case True then have "ennreal \f x\ \ eNorm (\ \ M) f" if "\f x\ \ esssup M (\x. \f x\)" for x unfolding L_infinity(1) using that e2ennreal_mono by fastforce then show ?thesis using esssup_AE[of "\x. ereal \f x\"] by force qed lemma L_infinity_AE_bound: assumes "f \ space\<^sub>N (\ \ M)" shows "AE x in M. \f x\ \ Norm (\ \ M) f" using L_infinity_AE_ebound[of f M] unfolding eNorm_Norm[OF assms] by (simp) text \In the next lemma, the assumption $C \geq 0$ that might seem useless is in fact necessary for the second statement when the space has zero measure. Indeed, any function is then almost surely bounded by any constant!\ lemma L_infinity_I: assumes "f \ borel_measurable M" "AE x in M. \f x\ \ C" "C \ 0" shows "f \ space\<^sub>N (\ \ M)" "Norm (\ \ M) f \ C" proof - show "f \ space\<^sub>N (\ \ M)" using L_infinity_space assms(1) assms(2) by force have "esssup M (\x. \f x\) \ C" using assms(1) assms(2) esssup_I by auto then have "eNorm (\ \ M) f \ ereal C" unfolding L_infinity(1) using assms(1) e2ennreal_mono by force then have "ennreal (Norm (\ \ M) f) \ ennreal C" using eNorm_Norm[OF \f \ space\<^sub>N (\ \ M)\] assms(3) by auto then show "Norm (\ \ M) f \ C" using assms(3) by auto qed lemma L_infinity_I': assumes [measurable]: "f \ borel_measurable M" and "AE x in M. ennreal \f x\ \ C" shows "eNorm (\ \ M) f \ C" proof - have "esssup M (\x. \f x\) \ enn2ereal C" apply (rule esssup_I, auto) using assms(2) less_eq_ennreal.rep_eq by auto then show ?thesis unfolding L_infinity using assms apply auto using e2ennreal_mono by fastforce qed lemma L_infinity_pos_measure: assumes [measurable]: "f \ borel_measurable M" and "eNorm (\ \ M) f > (C::real)" shows "emeasure M {x \ space M. \f x\ > C} > 0" proof - have *: "esssup M (\x. ereal(\f x\)) > ereal C" using \eNorm (\ \ M) f > C\ unfolding L_infinity proof (auto) assume a1: "ennreal C < e2ennreal (esssup M (\x. ereal \f x\))" have "\ e2ennreal (esssup M (\a. ereal \f a\)) \ e2ennreal (ereal C)" if "\ C < 0" using a1 that by (metis (no_types) e2ennreal_enn2ereal enn2ereal_ennreal leD leI) then have "e2ennreal (esssup M (\a. ereal \f a\)) \ e2ennreal (ereal C) \ (\e\esssup M (\a. ereal \f a\). ereal C < e)" using a1 e2ennreal_neg by fastforce then show ?thesis by (meson e2ennreal_mono leI less_le_trans) qed have "emeasure M {x \ space M. ereal(\f x\) > C} > 0" by (rule esssup_pos_measure[OF _ *], auto) then show ?thesis by auto qed lemma L_infinity_tendsto_AE: assumes "tendsto_in\<^sub>N (\ \ M) f g" "\n. f n \ space\<^sub>N (\ \ M)" "g \ space\<^sub>N (\ \ M)" shows "AE x in M. (\n. f n x) \ g x" proof - have *: "AE x in M. \(f n - g) x\ \ Norm (\ \ M) (f n - g)" for n apply (rule L_infinity_AE_bound) using assms spaceN_diff by blast have "AE x in M. \n. \(f n - g) x\ \ Norm (\ \ M) (f n - g)" apply (subst AE_all_countable) using * by auto moreover have "(\n. f n x) \ g x" if "\n. \(f n - g) x\ \ Norm (\ \ M) (f n - g)" for x proof - have "(\n. \(f n - g) x\) \ 0" apply (rule tendsto_sandwich[of "\n. 0" _ _ "\n. Norm (\ \ M) (f n - g)"]) using that \tendsto_in\<^sub>N (\ \ M) f g\ unfolding tendsto_in\<^sub>N_def by auto then have "(\n. \f n x - g x\) \ 0" by auto then show ?thesis by (simp add: \(\n. \f n x - g x\) \ 0\ LIM_zero_cancel tendsto_rabs_zero_cancel) qed ultimately show ?thesis by auto qed text \As an illustration of the mechanism of spaces inclusion, let us show that bounded continuous functions belong to $L^\infty$.\ lemma bcontfun_subset_L_infinity: assumes "sets M = sets borel" shows "space\<^sub>N bcontfun\<^sub>N \ space\<^sub>N (\ \ M)" "\f. f \ space\<^sub>N bcontfun\<^sub>N \ Norm (\ \ M) f \ Norm bcontfun\<^sub>N f" "\f. eNorm (\ \ M) f \ eNorm bcontfun\<^sub>N f" "bcontfun\<^sub>N \\<^sub>N \ \ M" proof - have *: "f \ space\<^sub>N (\ \ M) \ Norm (\ \ M) f \ Norm bcontfun\<^sub>N f" if "f \ space\<^sub>N bcontfun\<^sub>N" for f proof - have H: "continuous_on UNIV f" "\x. abs(f x) \ Norm bcontfun\<^sub>N f" using bcontfun\<^sub>ND[OF \f \ space\<^sub>N bcontfun\<^sub>N\] by auto then have "f \ borel_measurable borel" using borel_measurable_continuous_onI by simp then have "f \ borel_measurable M" using assms by auto have *: "AE x in M. \f x\ \ Norm bcontfun\<^sub>N f" using H(2) by auto show ?thesis using L_infinity_I[OF \f \ borel_measurable M\ * Norm_nonneg] by auto qed show "space\<^sub>N bcontfun\<^sub>N \ space\<^sub>N (\ \ M)" "\f. f \ space\<^sub>N bcontfun\<^sub>N \ Norm (\ \ M) f \ Norm bcontfun\<^sub>N f" using * by auto show **: "bcontfun\<^sub>N \\<^sub>N \ \ M" apply (rule quasinorm_subsetI'[of _ _ 1]) using * by auto have "eNorm (\ \ M) f \ ennreal 1 * eNorm bcontfun\<^sub>N f" for f apply (rule quasinorm_subset_Norm_eNorm) using * ** by auto then show "eNorm (\ \ M) f \ eNorm bcontfun\<^sub>N f" for f by simp qed subsection \$L^p$ for $0 < p < \infty$\ lemma Lp: assumes "p \ (1::real)" shows "eNorm (\ p M) f = (if (f \ borel_measurable M \ integrable M (\x. \f x\ powr p)) then (\x. \f x\ powr p \M) powr (1/p) else (\::ennreal))" "defect (\ p M) = 1" proof - define F where "F = {f \ borel_measurable M. integrable M (\x. \f x\ powr p)}" have *: "quasinorm_on F 1 (\(f::'a\real). (\x. \f x\ powr p \M) powr (1/p))" proof (rule quasinorm_onI) fix f g assume "f \ F" "g \ F" then show "f + g \ F" unfolding F_def plus_fun_def apply (auto) by (rule Minkowski_inequality(1), auto simp add: \p \ 1\) show "ennreal ((\x. \(f + g) x\ powr p \M) powr (1/p)) \ ennreal 1 * (\x. \f x\ powr p \M) powr (1/p) + ennreal 1 * (\x. \g x\ powr p \M) powr (1/p)" apply (auto, subst ennreal_plus[symmetric], simp, simp, rule ennreal_leI) unfolding plus_fun_def apply (rule Minkowski_inequality(2)[of p f M g], auto simp add: \p \ 1\) using \f \ F\ \g \ F\ unfolding F_def by auto next fix f and c::real assume "f \ F" show "c *\<^sub>R f \ F" using \f \ F\ unfolding scaleR_fun_def F_def by (auto simp add: abs_mult powr_mult) show "(\x. \(c *\<^sub>R f) x\ powr p \M) powr (1/p) \ ennreal(abs(c)) * (\x. \f x\ powr p \M) powr (1/p)" apply (rule eq_refl, subst ennreal_mult[symmetric], simp, simp, rule ennreal_cong) apply (unfold scaleR_fun_def, simp add: abs_mult powr_mult powr_powr) using \p \ 1\ by auto next show "0 \ F" unfolding zero_fun_def F_def by auto qed (auto) have "p \ 0" using \p \ 1\ by auto have **: "\ p M = quasinorm_of (1, (\f. if (f \ borel_measurable M \ integrable M (\x. \f x\ powr p)) then (\x. \f x\ powr p \M) powr (1/p) else (\::ennreal)))" unfolding Lp_space_def using enn2real_ennreal[OF \p \ 0\] \p \ 1\ apply auto using enn2real_ennreal[OF \p \ 0\] by presburger show "eNorm (\ p M) f = (if (f \ borel_measurable M \ integrable M (\x. \f x\ powr p)) then (\x. \f x\ powr p \M) powr (1/p) else (\::ennreal))" "defect (\ p M) = 1" unfolding ** using quasinorm_of[OF extend_quasinorm[OF *]] unfolding F_def by auto qed lemma Lp_le_1: assumes "p > 0" "p \ (1::real)" shows "eNorm (\ p M) f = (if (f \ borel_measurable M \ integrable M (\x. \f x\ powr p)) then (\x. \f x\ powr p \M) powr (1/p) else (\::ennreal))" "defect (\ p M) = 2 powr (1/p - 1)" proof - define F where "F = {f \ borel_measurable M. integrable M (\x. \f x\ powr p)}" have *: "quasinorm_on F (2 powr (1/p-1)) (\(f::'a\real). (\x. \f x\ powr p \M) powr (1/p))" proof (rule quasinorm_onI) fix f g assume "f \ F" "g \ F" then show "f + g \ F" unfolding F_def plus_fun_def apply (auto) by (rule Minkowski_inequality_le_1(1), auto simp add: \p > 0\ \p \ 1\) show "ennreal ((\x. \(f + g) x\ powr p \M) powr (1/p)) \ ennreal (2 powr (1/p-1)) * (\x. \f x\ powr p \M) powr (1/p) + ennreal (2 powr (1/p-1)) * (\x. \g x\ powr p \M) powr (1/p)" apply (subst ennreal_mult[symmetric], auto)+ apply (subst ennreal_plus[symmetric], simp, simp) apply (rule ennreal_leI) unfolding plus_fun_def apply (rule Minkowski_inequality_le_1(2)[of p f M g], auto simp add: \p > 0\ \p \ 1\) using \f \ F\ \g \ F\ unfolding F_def by auto next fix f and c::real assume "f \ F" show "c *\<^sub>R f \ F" using \f \ F\ unfolding scaleR_fun_def F_def by (auto simp add: abs_mult powr_mult) show "(\x. \(c *\<^sub>R f) x\ powr p \M) powr (1/p) \ ennreal(abs(c)) * (\x. \f x\ powr p \M) powr (1/p)" apply (rule eq_refl, subst ennreal_mult[symmetric], simp, simp, rule ennreal_cong) apply (unfold scaleR_fun_def, simp add: abs_mult powr_mult powr_powr) using \p > 0\ by auto next show "0 \ F" unfolding zero_fun_def F_def by auto show "1 \ 2 powr (1 / p - 1)" using \p > 0\ \p \ 1\ by (auto simp add: ge_one_powr_ge_zero) qed (auto) have "p \ 0" using \p > 0\ by auto have **: "\ p M = quasinorm_of (2 powr (1/p-1), (\f. if (f \ borel_measurable M \ integrable M (\x. \f x\ powr p)) then (\x. \f x\ powr p \M) powr (1/p) else (\::ennreal)))" unfolding Lp_space_def using \p > 0\ \p \ 1\ using enn2real_ennreal[OF \p \ 0\] apply auto by (insert enn2real_ennreal[OF \p \ 0\], presburger)+ show "eNorm (\ p M) f = (if (f \ borel_measurable M \ integrable M (\x. \f x\ powr p)) then (\x. \f x\ powr p \M) powr (1/p) else (\::ennreal))" "defect (\ p M) = 2 powr (1/p-1)" unfolding ** using quasinorm_of[OF extend_quasinorm[OF *]] unfolding F_def by auto qed lemma Lp_space: assumes "p > (0::real)" shows "space\<^sub>N (\ p M) = {f \ borel_measurable M. integrable M (\x. \f x\ powr p)}" apply (auto simp add: spaceN_iff) using Lp(1) Lp_le_1(1) \p > 0\ apply (metis infinity_ennreal_def less_le not_less) using Lp(1) Lp_le_1(1) \p > 0\ apply (metis infinity_ennreal_def less_le not_less) using Lp(1) Lp_le_1(1) \p > 0\ by (metis ennreal_neq_top linear top.not_eq_extremum) lemma Lp_I: assumes "p > (0::real)" "f \ borel_measurable M" "integrable M (\x. \f x\ powr p)" shows "f \ space\<^sub>N (\ p M)" "Norm (\ p M) f = (\x. \f x\ powr p \M) powr (1/p)" "eNorm (\ p M) f = (\x. \f x\ powr p \M) powr (1/p)" proof - have *: "eNorm (\ p M) f = (\x. \f x\ powr p \M) powr (1/p)" by (cases "p \ 1", insert assms, auto simp add: Lp_le_1(1) Lp(1)) then show **: "f \ space\<^sub>N (\ p M)" unfolding space\<^sub>N_def by auto show "Norm (\ p M) f = (\x. \f x\ powr p \M) powr (1/p)" using * unfolding Norm_def by auto then show "eNorm (\ p M) f = (\x. \f x\ powr p \M) powr (1/p)" using eNorm_Norm[OF **] by auto qed lemma Lp_D: assumes "p>0" "f \ space\<^sub>N (\ p M)" shows "f \ borel_measurable M" "integrable M (\x. \f x\ powr p)" "Norm (\ p M) f = (\x. \f x\ powr p \M) powr (1/p)" "eNorm (\ p M) f = (\x. \f x\ powr p \M) powr (1/p)" proof - show *: "f \ borel_measurable M" "integrable M (\x. \f x\ powr p)" using Lp_space[OF \p > 0\] assms(2) by auto then show "Norm (\ p M) f = (\x. \f x\ powr p \M) powr (1/p)" "eNorm (\ p M) f = (\x. \f x\ powr p \M) powr (1/p)" using Lp_I[OF \p > 0\] by auto qed lemma Lp_Norm: assumes "p > (0::real)" "f \ borel_measurable M" shows "Norm (\ p M) f = (\x. \f x\ powr p \M) powr (1/p)" "(Norm (\ p M) f) powr p = (\x. \f x\ powr p \M)" proof - show *: "Norm (\ p M) f = (\x. \f x\ powr p \M) powr (1/p)" proof (cases "integrable M (\x. \f x\ powr p)") case True then show ?thesis using Lp_I[OF assms True] by auto next case False then have "f \ space\<^sub>N (\ p M)" using Lp_space[OF \p > 0\, of M] by auto then have *: "Norm (\ p M) f = 0" using eNorm_Norm' by auto have "(\x. \f x\ powr p \M) = 0" using False by (simp add: not_integrable_integral_eq) then have "(\x. \f x\ powr p \M) powr (1/p) = 0" by auto then show ?thesis using * by auto qed then show "(Norm (\ p M) f) powr p = (\x. \f x\ powr p \M)" unfolding * using powr_powr \p > 0\ by auto qed lemma Lp_zero_space: assumes "p > (0::real)" shows "zero_space\<^sub>N (\ p M) = {f \ borel_measurable M. AE x in M. f x = 0}" proof (auto) fix f assume H: "f \ zero_space\<^sub>N (\ p M)" then have *: "f \ {f \ borel_measurable M. integrable M (\x. \f x\ powr p)}" using Lp_space[OF assms] zero_spaceN_subset_spaceN by auto then show "f \ borel_measurable M" by auto have "eNorm (\ p M) f = (\x. \f x\ powr p \M) powr (1/p)" by (cases "p \ 1", insert * \p > 0\, auto simp add: Lp_le_1(1) Lp(1)) then have "(\x. \f x\ powr p \M) = 0" using H unfolding zero_space\<^sub>N_def by auto then have "AE x in M. \f x\ powr p = 0" by (subst integral_nonneg_eq_0_iff_AE[symmetric], insert *, auto) then show "AE x in M. f x = 0" by auto next fix f::"'a \ real" assume H [measurable]: "f \ borel_measurable M" "AE x in M. f x = 0" then have *: "AE x in M. \f x\ powr p = 0" by auto have "integrable M (\x. \f x\ powr p)" using integrable_cong_AE[OF _ _ *] by auto have **: "(\x. \f x\ powr p \M) = 0" using integral_cong_AE[OF _ _ *] by auto have "eNorm (\ p M) f = (\x. \f x\ powr p \M) powr (1/p)" by (cases "p \ 1", insert H(1) \integrable M (\x. \f x\ powr p)\ \p > 0\, auto simp add: Lp_le_1(1) Lp(1)) then have "eNorm (\ p M) f = 0" using ** by simp then show "f \ zero_space\<^sub>N (\ p M)" using zero_spaceN_iff by auto qed lemma Lp_tendsto_AE_subseq: assumes "p>(0::real)" "tendsto_in\<^sub>N (\ p M) f g" "\n. f n \ space\<^sub>N (\ p M)" "g \ space\<^sub>N (\ p M)" shows "\r. strict_mono r \ (AE x in M. (\n. f (r n) x) \ g x)" proof - have "f n - g \ space\<^sub>N (\ p M)" for n using spaceN_diff[OF \\n. f n \ space\<^sub>N (\ p M)\ \g \ space\<^sub>N (\ p M)\] by simp have int: "integrable M (\x. \f n x - g x\ powr p)" for n using Lp_D(2)[OF \p > 0\ \f n - g \ space\<^sub>N (\ p M)\] by auto have "(\n. Norm (\ p M) (f n - g)) \ 0" using \tendsto_in\<^sub>N (\ p M) f g\ unfolding tendsto_in\<^sub>N_def by auto then have *: "(\n. (\x. \f n x - g x\ powr p \M) powr (1/p)) \ 0" using Lp_D(3)[OF \p > 0\ \\n. f n - g \ space\<^sub>N (\ p M)\] by auto have "(\n. ((\x. \f n x - g x\ powr p \M) powr (1/p)) powr p) \ 0" apply (rule tendsto_zero_powrI[of _ _ _ p]) using \p > 0\ * by auto then have **: "(\n. (\x. \f n x - g x\ powr p \M)) \ 0" using powr_powr \p > 0\ by auto have "\r. strict_mono r \ (AE x in M. (\n. \f (r n) x - g x\ powr p) \ 0)" apply (rule tendsto_L1_AE_subseq) using int ** by auto then obtain r where "strict_mono r" "AE x in M. (\n. \f (r n) x - g x\ powr p) \ 0" by blast moreover have "(\n. f (r n) x) \ g x" if "(\n. \f (r n) x - g x\ powr p) \ 0" for x proof - have "(\n. (\f (r n) x - g x\ powr p) powr (1/p)) \ 0" apply (rule tendsto_zero_powrI[of _ _ _ "1/p"]) using \p > 0\ that by auto then have "(\n. \f (r n) x - g x\) \ 0" using powr_powr \p > 0\ by auto show ?thesis by (simp add: \(\n. \f (r n) x - g x\) \ 0\ Limits.LIM_zero_cancel tendsto_rabs_zero_cancel) qed ultimately have "AE x in M. (\n. f (r n) x) \ g x" by auto then show ?thesis using \strict_mono r\ by auto qed subsection \Specialization to $L^1$\ lemma L1_space: "space\<^sub>N (\ 1 M) = {f. integrable M f}" unfolding one_ereal_def using Lp_space[of 1 M] integrable_abs_iff by auto lemma L1_I: assumes "integrable M f" shows "f \ space\<^sub>N (\ 1 M)" "Norm (\ 1 M) f = (\x. \f x\ \M)" "eNorm (\ 1 M) f = (\x. \f x\ \M)" unfolding one_ereal_def using Lp_I[of 1, OF _ borel_measurable_integrable[OF assms]] assms powr_to_1 by auto lemma L1_D: assumes "f \ space\<^sub>N (\ 1 M)" shows "f \ borel_measurable M" "integrable M f" "Norm (\ 1 M) f = (\x. \f x\ \M)" "eNorm (\ 1 M) f = (\x. \f x\ \M)" using assms by (auto simp add: L1_space L1_I) lemma L1_int_ineq: "abs(\x. f x \M) \ Norm (\ 1 M) f" proof (cases "integrable M f") case True then show ?thesis using L1_I(2)[OF True] by auto next case False then have "(\x. f x \M) = 0" by (simp add: not_integrable_integral_eq) then show ?thesis using Norm_nonneg by auto qed text \In $L^1$, one can give a direct formula for the eNorm of a measurable function, using a nonnegative integral. The same formula holds in $L^p$ for $p > 0$, with additional powers $p$ and $1/p$, but one can not write it down since \verb+powr+ is not defined on \verb+ennreal+.\ lemma L1_Norm: assumes [measurable]: "f \ borel_measurable M" shows "Norm (\ 1 M) f = (\x. \f x\ \M)" "eNorm (\ 1 M) f = (\\<^sup>+x. \f x\ \M)" proof - show *: "Norm (\ 1 M) f = (\x. \f x\ \M)" using Lp_Norm[of 1, OF _ assms] unfolding one_ereal_def by auto show "eNorm (\ 1 M) f = (\\<^sup>+x. \f x\ \M)" proof (cases "integrable M f") case True then have "f \ space\<^sub>N (\ 1 M)" using L1_space by auto then have "eNorm (\ 1 M) f = ennreal (Norm (\ 1 M) f)" using eNorm_Norm by auto then show ?thesis by (metis (mono_tags) * AE_I2 True abs_ge_zero integrable_abs nn_integral_eq_integral) next case False then have "eNorm (\ 1 M) f = \" using L1_space space\<^sub>N_def by (metis ennreal_add_eq_top infinity_ennreal_def le_iff_add le_less_linear mem_Collect_eq) moreover have "(\\<^sup>+x. \f x\ \M) = \" apply (rule nn_integral_nonneg_infinite) using False by (auto simp add: integrable_abs_iff) ultimately show ?thesis by simp qed qed lemma L1_indicator: assumes [measurable]: "A \ sets M" shows "eNorm (\ 1 M) (indicator A) = emeasure M A" by (subst L1_Norm, auto, metis assms ennreal_indicator nn_integral_cong nn_integral_indicator) lemma L1_indicator': assumes [measurable]: "A \ sets M" and "emeasure M A \ \" shows "indicator A \ space\<^sub>N (\ 1 M)" "Norm (\ 1 M) (indicator A) = measure M A" unfolding space\<^sub>N_def Norm_def using L1_indicator[OF \A \ sets M\] \emeasure M A \ \\ by (auto simp add: top.not_eq_extremum Sigma_Algebra.measure_def) subsection \$L^0$\ text \We have defined $L^p$ for all exponents $p$, although it does not really make sense for $p = 0$. We have chosen a definition in this case (the space of all measurable functions) so that many statements are true for all exponents. In this paragraph, we show the consistency of this definition.\ lemma L_zero: "eNorm (\ 0 M) f = (if f \ borel_measurable M then 0 else \)" "defect (\ 0 M) = 1" proof - have *: "quasinorm_on UNIV 1 (\(f::'a\real). (if f \ borel_measurable M then 0 else \))" by (rule extend_quasinorm, rule quasinorm_onI, auto) show "eNorm (\ 0 M) f = (if f \ borel_measurable M then 0 else \)" "defect (\ 0 M) = 1" using quasinorm_of[OF *] unfolding Lp_space_def by auto qed lemma L_zero_space [simp]: "space\<^sub>N (\ 0 M) = borel_measurable M" "zero_space\<^sub>N (\ 0 M) = borel_measurable M" apply (auto simp add: spaceN_iff zero_spaceN_iff L_zero(1)) using top.not_eq_extremum by force+ subsection \Basic results on $L^p$ for general $p$\ lemma Lp_measurable_subset: "space\<^sub>N (\ p M) \ borel_measurable M" proof (cases rule: Lp_cases[of p]) case zero then show ?thesis using L_zero_space by auto next case (real_pos p2) then show ?thesis using Lp_space[OF \p2 > 0\] by auto next case PInf then show ?thesis using L_infinity_space by auto qed lemma Lp_measurable: assumes "f \ space\<^sub>N (\ p M)" shows "f \ borel_measurable M" using assms Lp_measurable_subset by auto lemma Lp_infinity_zero_space: assumes "p > (0::ennreal)" shows "zero_space\<^sub>N (\ p M) = {f \ borel_measurable M. AE x in M. f x = 0}" proof (cases rule: Lp_cases[of p]) case PInf then show ?thesis using L_infinity_zero_space by auto next case (real_pos p2) then show ?thesis using Lp_zero_space[OF \p2 > 0\] unfolding \p = ennreal p2\ by auto next case zero then have False using assms by auto then show ?thesis by simp qed lemma (in prob_space) Lp_subset_Lq: assumes "p \ q" shows "\f. eNorm (\ p M) f \ eNorm (\ q M) f" "\ q M \\<^sub>N \ p M" "space\<^sub>N (\ q M) \ space\<^sub>N (\ p M)" "\f. f \ space\<^sub>N (\ q M) \ Norm (\ p M) f \ Norm (\ q M) f" proof - show "eNorm (\ p M) f \ eNorm (\ q M) f" for f proof (cases "eNorm (\ q M) f < \") case True then have "f \ space\<^sub>N (\ q M)" using spaceN_iff by auto then have f_meas [measurable]: "f \ borel_measurable M" using Lp_measurable by auto consider "p = 0" | "p = q" | "p > 0 \ p < \ \ q = \" | "p > 0 \ p < q \ q < \" using \p \ q\ apply (simp add: top.not_eq_extremum) using not_less_iff_gr_or_eq order.order_iff_strict by fastforce then show ?thesis proof (cases) case 1 then show ?thesis by (simp add: L_zero(1)) next case 2 then show ?thesis by auto next case 3 then have "q = \" by simp obtain p2 where "p = ennreal p2" "p2 > 0" using 3 enn2real_positive_iff[of p] by (cases p) auto have *: "AE x in M. \f x\ \ Norm (\ \ M) f" using L_infinity_AE_bound \f \ space\<^sub>N (\ q M)\ \q = \\ by auto have **: "integrable M (\x. \f x\ powr p2)" apply (rule Bochner_Integration.integrable_bound[of _ "\x. (Norm (\ \ M) f) powr p2"], auto) using * powr_mono2 \p2 > 0\ by force then have "eNorm (\ p2 M) f = (\x. \f x\ powr p2 \M) powr (1/p2)" using Lp_I(3)[OF \p2 > 0\ f_meas] by simp also have "... \ (\x. (Norm (\ \ M) f) powr p2 \M) powr (1/p2)" apply (rule ennreal_leI, rule powr_mono2, simp add: \p2 > 0\ less_imp_le, simp) apply (rule integral_mono_AE, auto simp add: **) using * powr_mono2 \p2 > 0\ by force also have "... = Norm (\ \ M) f" using \p2 > 0\ by (auto simp add: prob_space powr_powr) finally show ?thesis using \p = ennreal p2\ \q = \\ eNorm_Norm[OF \f \ space\<^sub>N (\ q M)\] by auto next case 4 then have "0 < p" "p < \" by auto then obtain p2 where "p = ennreal p2" "p2 > 0" using enn2real_positive_iff[of p] by (cases p) auto have "0 < q" "q < \" using 4 by auto then obtain q2 where "q = ennreal q2" "q2 > 0" using enn2real_positive_iff[of q] by (cases q) auto have "p2 < q2" using 4 \p = ennreal p2\ \q = ennreal q2\ using ennreal_less_iff by auto define r2 where "r2 = q2 / p2" have "r2 \ 1" unfolding r2_def using \p2 < q2\ \p2 > 0\ by auto have *: "abs (\z\ powr p2) powr r2 = \z\ powr q2" for z::real unfolding r2_def using \p2 > 0\ by (simp add: powr_powr) have I: "integrable M (\x. abs(\f x\ powr p2) powr r2)" unfolding * using \f \ space\<^sub>N (\ q M)\ \q = ennreal q2\ Lp_D(2)[OF \q2 > 0\] by auto have J: "integrable M (\x. \f x\ powr p2)" by (rule bound_L1_Lp(1)[OF \r2 \ 1\ _ I], auto) have "f \ space\<^sub>N (\ p2 M)" by (rule Lp_I(1)[OF \p2 > 0\ _ J], simp) have "(\x. \f x\ powr p2 \M) powr (1/p2) = abs(\x. \f x\ powr p2 \M) powr (1/p2)" by auto also have "... \ ((\x. abs (\f x\ powr p2) powr r2 \M) powr (1/r2)) powr (1/p2)" apply (subst powr_mono2, simp add: \p2 > 0\ less_imp_le, simp) apply (rule bound_L1_Lp, simp add: \r2 \ 1\, simp) unfolding * using \f \ space\<^sub>N (\ q M)\ \q = ennreal q2\ Lp_D(2)[OF \q2 > 0\] by auto also have "... = (\x. \f x\ powr q2 \M) powr (1/q2)" unfolding * using \p2 > 0\ by (simp add: powr_powr r2_def) finally show ?thesis using \f \ space\<^sub>N (\ q M)\ Lp_D(4)[OF \q2 > 0\] ennreal_leI unfolding \p = ennreal p2\ \q = ennreal q2\ Lp_D(4)[OF \p2 > 0\ \f \ space\<^sub>N (\ p2 M)\] by force qed next case False then have "eNorm (\ q M) f = \" using top.not_eq_extremum by fastforce then show ?thesis by auto qed then show "\ q M \\<^sub>N \ p M" using quasinorm_subsetI[of _ _ 1] by auto then show "space\<^sub>N (\ q M) \ space\<^sub>N (\ p M)" using quasinorm_subset_space by auto then show "Norm (\ p M) f \ Norm (\ q M) f" if "f \ space\<^sub>N (\ q M)" for f using eNorm_Norm that \eNorm (\ p M) f \ eNorm (\ q M) f\ ennreal_le_iff Norm_nonneg by (metis rev_subsetD) qed proposition Lp_domination: assumes [measurable]: "g \ borel_measurable M" and "f \ space\<^sub>N (\ p M)" "AE x in M. \g x\ \ \f x\" shows "g \ space\<^sub>N (\ p M)" "Norm (\ p M) g \ Norm (\ p M) f" proof - have [measurable]: "f \ borel_measurable M" using Lp_measurable[OF assms(2)] by simp have "g \ space\<^sub>N (\ p M) \ Norm (\ p M) g \ Norm (\ p M) f" proof (cases rule: Lp_cases[of p]) case zero then have "Norm (\ p M) g = 0" unfolding Norm_def using L_zero(1)[of M] by auto then have "Norm (\ p M) g \ Norm (\ p M) f" using Norm_nonneg by auto then show ?thesis unfolding \p = 0\ L_zero_space by auto next case (real_pos p2) have *: "integrable M (\x. \f x\ powr p2)" using \f \ space\<^sub>N (\ p M)\ unfolding \p = ennreal p2\ using Lp_D(2) \p2 > 0\ by auto have **: "integrable M (\x. \g x\ powr p2)" apply (rule Bochner_Integration.integrable_bound[of _ "\x. \f x\ powr p2"]) using * apply auto using assms(3) powr_mono2 \p2 > 0\ by (auto simp add: less_imp_le) then have "g \ space\<^sub>N (\ p M)" unfolding \p = ennreal p2\ using Lp_space[OF \p2 > 0\, of M] by auto have "Norm (\ p M) g = (\x. \g x\ powr p2 \M) powr (1/p2)" unfolding \p = ennreal p2\ by (rule Lp_I(2)[OF \p2 > 0\ _ **], simp) also have "... \ (\x. \f x\ powr p2 \M) powr (1/p2)" apply (rule powr_mono2, simp add: \p2 > 0\ less_imp_le, simp) apply (rule integral_mono_AE, auto simp add: * **) using \p2 > 0\ less_imp_le powr_mono2 assms(3) by auto also have "... = Norm (\ p M) f" unfolding \p = ennreal p2\ by (rule Lp_I(2)[OF \p2 > 0\ _ *, symmetric], simp) finally show ?thesis using \g \ space\<^sub>N (\ p M)\ by auto next case PInf have "AE x in M. \f x\ \ Norm (\ p M) f" using \f \ space\<^sub>N (\ p M)\ L_infinity_AE_bound unfolding \p = \\ by auto then have *: "AE x in M. \g x\ \ Norm (\ p M) f" using assms(3) by auto show ?thesis using L_infinity_I[OF assms(1) *] Norm_nonneg \p = \\ by auto qed then show "g \ space\<^sub>N (\ p M)" "Norm (\ p M) g \ Norm (\ p M) f" by auto qed lemma Lp_Banach_lattice: assumes "f \ space\<^sub>N (\ p M)" shows "(\x. \f x\) \ space\<^sub>N (\ p M)" "Norm (\ p M) (\x. \f x\) = Norm (\ p M) f" proof - have [measurable]: "f \ borel_measurable M" using Lp_measurable[OF assms] by simp show "(\x. \f x\) \ space\<^sub>N (\ p M)" by (rule Lp_domination(1)[OF _ assms], auto) have "Norm (\ p M) (\x. \f x\) \ Norm (\ p M) f" by (rule Lp_domination[OF _ assms], auto) moreover have "Norm (\ p M) f \ Norm (\ p M) (\x. \f x\)" by (rule Lp_domination[OF _ \(\x. \f x\) \ space\<^sub>N (\ p M)\], auto) finally show "Norm (\ p M) (\x. \f x\) = Norm (\ p M) f" by auto qed lemma Lp_bounded_bounded_support: assumes [measurable]: "f \ borel_measurable M" and "AE x in M. \f x\ \ C" "emeasure M {x \ space M. f x \ 0} \ \" shows "f \ space\<^sub>N(\ p M)" proof (cases rule: Lp_cases[of p]) case zero then show ?thesis using L_zero_space assms by blast next case PInf then show ?thesis using L_infinity_space assms by blast next case (real_pos p2) have *: "integrable M (\x. \f x\ powr p2)" apply (rule integrableI_bounded_set[of "{x \ space M. f x \ 0}" _ _ "C powr p2"]) using assms powr_mono2[OF less_imp_le[OF \p2 > 0\]] by (auto simp add: top.not_eq_extremum) show ?thesis unfolding \p = ennreal p2\ apply (rule Lp_I[OF \p2 > 0\]) using * by auto qed subsection \$L^p$ versions of the main theorems in integration theory\ text \The space $L^p$ is stable under almost sure convergence, for sequence with bounded norm. This is a version of Fatou's lemma (and it indeed follows from this lemma in the only nontrivial situation where $p \in (0, +\infty)$.\ proposition Lp_AE_limit: assumes [measurable]: "g \ borel_measurable M" and "AE x in M. (\n. f n x) \ g x" shows "eNorm (\ p M) g \ liminf (\n. eNorm (\ p M) (f n))" proof (cases "liminf (\n. eNorm (\ p M) (f n)) = \") case True then show ?thesis by auto next case False define le where "le = liminf (\n. eNorm (\ p M) (f n))" then have "le < \" using False by (simp add: top.not_eq_extremum) obtain r0 where r0: "strict_mono r0" "(\n. eNorm (\ p M) (f (r0 n))) \ le" - using liminf_subseq_lim unfolding comp_def le_def by force + using liminf_subseq_lim[of "\n. eNorm (\ p M) (f n)"] + unfolding comp_def le_def + by blast then have "eventually (\n. eNorm (\ p M) (f (r0 n)) < \) sequentially" using False unfolding order_tendsto_iff le_def by (simp add: top.not_eq_extremum) then obtain N where N: "\n. n \ N \ eNorm (\ p M) (f (r0 n)) < \" unfolding eventually_sequentially by blast define r where "r = (\n. r0 (n + N))" have "strict_mono r" unfolding r_def using \strict_mono r0\ by (simp add: strict_mono_Suc_iff) have *: "(\n. eNorm (\ p M) (f (r n))) \ le" unfolding r_def using LIMSEQ_ignore_initial_segment[OF r0(2), of N]. have "f (r n) \ space\<^sub>N (\ p M)" for n using N spaceN_iff unfolding r_def by force then have [measurable]: "f (r n) \ borel_measurable M" for n using Lp_measurable by auto define l where "l = enn2real le" have "l \ 0" unfolding l_def by auto have "le = ennreal l" using \le < \\ unfolding l_def by auto have [tendsto_intros]: "(\n. Norm (\ p M) (f (r n))) \ l" apply (rule tendsto_ennrealD) using * \le < \\ unfolding eNorm_Norm[OF \\n. f (r n) \ space\<^sub>N (\ p M)\] l_def by auto show ?thesis proof (cases rule: Lp_cases[of p]) case zero then have "eNorm (\ p M) g = 0" using assms(1) by (simp add: L_zero(1)) then show ?thesis by auto next case (real_pos p2) then have "f (r n) \ space\<^sub>N (\ p2 M)" for n using \\n. f (r n) \ space\<^sub>N(\ p M)\ by auto have "liminf (\n. ennreal(\f (r n) x\ powr p2)) = \g x\ powr p2" if "(\n. f n x) \ g x" for x apply (rule lim_imp_Liminf, auto intro!: tendsto_intros simp add: \p2 > 0\) using LIMSEQ_subseq_LIMSEQ[OF that \strict_mono r\] unfolding comp_def by auto then have *: "AE x in M. liminf (\n. ennreal(\f (r n) x\ powr p2)) = \g x\ powr p2" using \AE x in M. (\n. f n x) \ g x\ by auto have "(\\<^sup>+x. ennreal(\f (r n) x\ powr p2) \M) = ennreal((Norm (\ p M) (f (r n))) powr p2)" for n proof - have "(\\<^sup>+x. ennreal(\f (r n) x\ powr p2) \M) = ennreal (\x. \f (r n) x\ powr p2 \M)" by (rule nn_integral_eq_integral, auto simp add: Lp_D(2)[OF \p2 > 0\ \f (r n) \ space\<^sub>N (\ p2 M)\]) also have "... = ennreal((Norm (\ p2 M) (f (r n))) powr p2)" unfolding Lp_D(3)[OF \p2 > 0\ \f (r n) \ space\<^sub>N (\ p2 M)\] using powr_powr \p2 > 0\ by auto finally show ?thesis using \p = ennreal p2\ by simp qed moreover have "(\n. ennreal((Norm (\ p M) (f (r n))) powr p2)) \ ennreal(l powr p2)" by (auto intro!:tendsto_intros simp add: \p2 > 0\) ultimately have **: "liminf (\n. (\\<^sup>+x. ennreal(\f (r n) x\ powr p2) \M)) = ennreal(l powr p2)" using lim_imp_Liminf by force have "(\\<^sup>+x. \g x\ powr p2 \M) = (\\<^sup>+x. liminf (\n. ennreal(\f (r n) x\ powr p2)) \M)" apply (rule nn_integral_cong_AE) using * by auto also have "... \ liminf (\n. \\<^sup>+x. ennreal(\f (r n) x\ powr p2) \M)" by (rule nn_integral_liminf, auto) finally have "(\\<^sup>+x. \g x\ powr p2 \M) \ ennreal(l powr p2)" using ** by auto then have "(\\<^sup>+x. \g x\ powr p2 \M) < \" using le_less_trans by fastforce then have intg: "integrable M (\x. \g x\ powr p2)" apply (intro integrableI_nonneg) by auto then have "g \ space\<^sub>N (\ p2 M)" using Lp_I(1)[OF \p2 > 0\, of _ M] by fastforce have "ennreal((Norm (\ p2 M) g) powr p2) = ennreal(\x. \g x\ powr p2 \M)" unfolding Lp_D(3)[OF \p2 > 0\ \g \ space\<^sub>N (\ p2 M)\] using powr_powr \p2 > 0\ by auto also have "... = (\\<^sup>+x. \g x\ powr p2 \M)" by (rule nn_integral_eq_integral[symmetric], auto simp add: intg) finally have "ennreal((Norm (\ p2 M) g) powr p2) \ ennreal(l powr p2)" using \(\\<^sup>+x. \g x\ powr p2 \M) \ ennreal(l powr p2)\ by auto then have "((Norm (\ p2 M) g) powr p2) powr (1/p2) \ (l powr p2) powr (1/p2)" using ennreal_le_iff \l \ 0\ \p2 > 0\ powr_mono2 by auto then have "Norm (\ p2 M) g \ l" using \p2 > 0\ \l \ 0\ by (auto simp add: powr_powr) then have "eNorm (\ p2 M) g \ le" unfolding eNorm_Norm[OF \g \ space\<^sub>N (\ p2 M)\] \le = ennreal l\ using ennreal_leI by auto then show ?thesis unfolding le_def \p = ennreal p2\ by simp next case PInf then have "AE x in M. \n. \f (r n) x\ \ Norm (\ \ M) (f (r n))" apply (subst AE_all_countable) using L_infinity_AE_bound \\n. f (r n) \ space\<^sub>N (\ p M)\ by blast moreover have "\g x\ \ l" if "\n. \f (r n) x\ \ Norm (\ \ M) (f (r n))" "(\n. f n x) \ g x" for x proof - have "(\n. f (r n) x) \ g x" using that LIMSEQ_subseq_LIMSEQ[OF _ \strict_mono r\] unfolding comp_def by auto then have *: "(\n. \f (r n) x\) \ \g x\" by (auto intro!:tendsto_intros) show ?thesis apply (rule LIMSEQ_le[OF *]) using that(1) \(\n. Norm (\ p M) (f (r n))) \ l\ unfolding PInf by auto qed ultimately have "AE x in M. \g x\ \ l" using \AE x in M. (\n. f n x) \ g x\ by auto then have "g \ space\<^sub>N (\ \ M)" "Norm (\ \ M) g \ l" using L_infinity_I[OF \g \ borel_measurable M\ _ \l \ 0\] by auto then have "eNorm (\ \ M) g \ le" unfolding eNorm_Norm[OF \g \ space\<^sub>N (\ \ M)\] \le = ennreal l\ using ennreal_leI by auto then show ?thesis unfolding le_def \p = \\ by simp qed qed lemma Lp_AE_limit': assumes "g \ borel_measurable M" "\n. f n \ space\<^sub>N (\ p M)" "AE x in M. (\n. f n x) \ g x" "(\n. Norm (\ p M) (f n)) \ l" shows "g \ space\<^sub>N (\ p M)" "Norm (\ p M) g \ l" proof - have "l \ 0" by (rule LIMSEQ_le_const[OF \(\n. Norm (\ p M) (f n)) \ l\], auto) have "(\n. eNorm (\ p M) (f n)) \ ennreal l" unfolding eNorm_Norm[OF \\n. f n \ space\<^sub>N (\ p M)\] using \(\n. Norm (\ p M) (f n)) \ l\ by auto then have *: "ennreal l = liminf (\n. eNorm (\ p M) (f n))" using lim_imp_Liminf[symmetric] trivial_limit_sequentially by blast have "eNorm (\ p M) g \ ennreal l" unfolding * apply (rule Lp_AE_limit) using assms by auto then have "eNorm (\ p M) g < \" using le_less_trans by fastforce then show "g \ space\<^sub>N (\ p M)" using spaceN_iff by auto show "Norm (\ p M) g \ l" using \eNorm (\ p M) g \ ennreal l\ ennreal_le_iff[OF \l \ 0\] unfolding eNorm_Norm[OF \g \ space\<^sub>N (\ p M)\] by auto qed lemma Lp_AE_limit'': assumes "g \ borel_measurable M" "\n. f n \ space\<^sub>N (\ p M)" "AE x in M. (\n. f n x) \ g x" "\n. Norm (\ p M) (f n) \ C" shows "g \ space\<^sub>N (\ p M)" "Norm (\ p M) g \ C" proof - have "C \ 0" by (rule order_trans[OF Norm_nonneg[of "\ p M" "f 0"] \Norm (\ p M) (f 0) \ C\]) have *: "liminf (\n. ennreal C) = ennreal C" using Liminf_const trivial_limit_at_top_linorder by blast have "eNorm (\ p M) (f n) \ ennreal C" for n unfolding eNorm_Norm[OF \f n \ space\<^sub>N (\ p M)\] using \Norm (\ p M) (f n) \ C\ by (auto simp add: ennreal_leI) then have "liminf (\n. eNorm (\ p M) (f n)) \ ennreal C" using Liminf_mono[of "(\n. eNorm (\ p M) (f n))" "\_. C" sequentially] * by auto then have "eNorm (\ p M) g \ ennreal C" using Lp_AE_limit[OF \g \ borel_measurable M\ \AE x in M. (\n. f n x) \ g x\, of p] by auto then have "eNorm (\ p M) g < \" using le_less_trans by fastforce then show "g \ space\<^sub>N (\ p M)" using spaceN_iff by auto show "Norm (\ p M) g \ C" using \eNorm (\ p M) g \ ennreal C\ ennreal_le_iff[OF \C \ 0\] unfolding eNorm_Norm[OF \g \ space\<^sub>N (\ p M)\] by auto qed text \We give the version of Lebesgue dominated convergence theorem in the setting of $L^p$ spaces.\ proposition Lp_domination_limit: fixes p::real assumes [measurable]: "g \ borel_measurable M" "\n. f n \ borel_measurable M" and "m \ space\<^sub>N (\ p M)" "AE x in M. (\n. f n x) \ g x" "\n. AE x in M. \f n x\ \ m x" shows "g \ space\<^sub>N (\ p M)" "tendsto_in\<^sub>N (\ p M) f g" proof - have [measurable]: "m \ borel_measurable M" using Lp_measurable[OF \m \ space\<^sub>N (\ p M)\] by auto have "f n \ space\<^sub>N(\ p M)" for n apply (rule Lp_domination[OF _ \m \ space\<^sub>N (\ p M)\]) using \AE x in M. \f n x\ \ m x\ by auto have "AE x in M. \n. \f n x\ \ m x" apply (subst AE_all_countable) using \\n. AE x in M. \f n x\ \ m x\ by auto moreover have "\g x\ \ m x" if "\n. \f n x\ \ m x" "(\n. f n x) \ g x" for x apply (rule LIMSEQ_le_const2[of "\n. \f n x\"]) using that by (auto intro!:tendsto_intros) ultimately have *: "AE x in M. \g x\ \ m x" using \AE x in M. (\n. f n x) \ g x\ by auto show "g \ space\<^sub>N(\ p M)" apply (rule Lp_domination[OF _ \m \ space\<^sub>N (\ p M)\]) using * by auto have "(\n. Norm (\ p M) (f n - g)) \ 0" proof (cases "p \ 0") case True then have "ennreal p = 0" by (simp add: ennreal_eq_0_iff) then show ?thesis unfolding Norm_def by (auto simp add: L_zero(1)) next case False then have "p > 0" by auto have "(\n. (\x. \f n x - g x\ powr p \M)) \ (\x. \0\ powr p \M)" proof (rule integral_dominated_convergence[of _ _ _ "(\x. \2 * m x\ powr p)"], auto) show "integrable M (\x. \2 * m x\ powr p)" unfolding abs_mult apply (subst powr_mult) using Lp_D(2)[OF \p > 0\ \m \ space\<^sub>N (\ p M)\] by auto have "(\n. \f n x - g x\ powr p) \ \0\ powr p" if "(\n. f n x) \ g x" for x apply (rule tendsto_powr') using \p > 0\ that apply (auto) using Lim_null tendsto_rabs_zero_iff by fastforce then show "AE x in M. (\n. \f n x - g x\ powr p) \ 0" using \AE x in M. (\n. f n x) \ g x\ by auto have "\f n x - g x\ powr p \ \2 * m x\ powr p" if "\f n x\ \ m x" "\g x\ \ m x" for n x using powr_mono2 \p > 0\ that by auto then show "AE x in M. \f n x - g x\ powr p \ \2 * m x\ powr p" for n using \AE x in M. \f n x\ \ m x\ \AE x in M. \g x\ \ m x\ by auto qed then have "(\n. (Norm (\ p M) (f n - g)) powr p) \ (Norm (\ p M) 0) powr p" unfolding Lp_D[OF \p > 0\ spaceN_diff[OF \\n. f n \ space\<^sub>N(\ p M)\ \g \ space\<^sub>N(\ p M)\]] using \p > 0\ by (auto simp add: powr_powr) then have "(\n. ((Norm (\ p M) (f n - g)) powr p) powr (1/p)) \ ((Norm (\ p M) 0) powr p) powr (1/p)" by (rule tendsto_powr', auto simp add: \p > 0\) then show ?thesis using powr_powr \p > 0\ by auto qed then show "tendsto_in\<^sub>N (\ p M) f g" unfolding tendsto_in\<^sub>N_def by auto qed text \We give the version of the monotone convergence theorem in the setting of $L^p$ spaces.\ proposition Lp_monotone_limit: fixes f::"nat \ 'a \ real" assumes "p > (0::ennreal)" "AE x in M. incseq (\n. f n x)" "\n. Norm (\ p M) (f n) \ C" "\n. f n \ space\<^sub>N (\ p M)" shows "AE x in M. convergent (\n. f n x)" "(\x. lim (\n. f n x)) \ space\<^sub>N (\ p M)" "Norm (\ p M) (\x. lim (\n. f n x)) \ C" proof - have [measurable]: "f n \ borel_measurable M" for n using Lp_measurable[OF assms(4)]. show "AE x in M. convergent (\n. f n x)" proof (cases rule: Lp_cases[of p]) case PInf have "AE x in M. \f n x\ \ C" for n using L_infinity_AE_bound[of "f n" M] \Norm (\ p M) (f n) \ C\ \f n \ space\<^sub>N (\ p M)\ unfolding \p=\\ by auto then have *: "AE x in M. \n. \f n x\ \ C" by (subst AE_all_countable, auto) have "(\n. f n x) \ (SUP n. f n x)" if "incseq (\n. f n x)" "\n. \f n x\ \ C" for x apply (rule LIMSEQ_incseq_SUP[OF _ \incseq (\n. f n x)\]) using that(2) abs_le_D1 by fastforce then have "convergent (\n. f n x)" if "incseq (\n. f n x)" "\n. \f n x\ \ C" for x unfolding convergent_def using that by auto then show ?thesis using \AE x in M. incseq (\n. f n x)\ * by auto next case (real_pos p2) define g where "g = (\n. f n - f 0)" have "AE x in M. incseq (\n. g n x)" unfolding g_def using \AE x in M. incseq (\n. f n x)\ by (simp add: incseq_def) have "g n \ space\<^sub>N (\ p2 M)" for n unfolding g_def using \\n. f n \ space\<^sub>N (\ p M)\ unfolding \p = ennreal p2\ by auto then have [measurable]: "g n \ borel_measurable M" for n using Lp_measurable by auto define D where "D = defect (\ p2 M) * C + defect (\ p2 M) * C" have "Norm (\ p2 M) (g n) \ D" for n proof - have "f n \ space\<^sub>N (\ p2 M)" using \f n \ space\<^sub>N (\ p M)\ unfolding \p = ennreal p2\ by auto have "Norm (\ p2 M) (g n) \ defect (\ p2 M) * Norm (\ p2 M) (f n) + defect (\ p2 M) * Norm (\ p2 M) (f 0)" unfolding g_def using Norm_triangular_ineq_diff[OF \f n \ space\<^sub>N (\ p2 M)\] by auto also have "... \ D" unfolding D_def apply(rule add_mono) using mult_left_mono defect_ge_1[of "\ p2 M"] \\n. Norm (\ p M) (f n) \ C\ unfolding \p = ennreal p2\ by auto finally show ?thesis by simp qed have g_bound: "(\\<^sup>+x. \g n x\ powr p2 \M) \ ennreal(D powr p2)" for n proof - have "(\\<^sup>+x. \g n x\ powr p2 \M) = ennreal(\x. \g n x\ powr p2 \M)" apply (rule nn_integral_eq_integral) using Lp_D(2)[OF \p2 > 0\ \g n \ space\<^sub>N (\ p2 M)\] by auto also have "... = ennreal((Norm (\ p2 M) (g n)) powr p2)" apply (subst Lp_Norm(2)[OF \p2 > 0\, of "g n", symmetric]) by auto also have "... \ ennreal(D powr p2)" by (auto intro!: powr_mono2 simp add: less_imp_le[OF \p2 > 0\] \Norm (\ p2 M) (g n) \ D\) finally show ?thesis by simp qed have "\n. g n x \ 0" if "incseq (\n. f n x)" for x unfolding g_def using that by (auto simp add: incseq_def) then have "AE x in M. \n. g n x \ 0" using \AE x in M. incseq (\n. f n x)\ by auto define h where "h = (\n x. ennreal(\g n x\ powr p2))" have [measurable]: "h n \ borel_measurable M" for n unfolding h_def by auto define H where "H = (\x. (SUP n. h n x))" have [measurable]: "H \ borel_measurable M" unfolding H_def by auto have "\n. h n x \ h (Suc n) x" if "\n. g n x \ 0" "incseq (\n. g n x)" for x unfolding h_def apply (auto intro!:powr_mono2) apply (auto simp add: less_imp_le[OF \p2 > 0\]) using that incseq_SucD by auto then have *: "AE x in M. h n x \ h (Suc n) x" for n using \AE x in M. \n. g n x \ 0\ \AE x in M. incseq (\n. g n x)\ by auto have "(\\<^sup>+x. H x \M) = (SUP n. \\<^sup>+x. h n x \M)" unfolding H_def by (rule nn_integral_monotone_convergence_SUP_AE, auto simp add: *) also have "... \ ennreal(D powr p2)" unfolding H_def h_def using g_bound by (simp add: SUP_least) finally have "(\\<^sup>+x. H x \M) < \" by (simp add: le_less_trans) then have "AE x in M. H x \ \" by (metis (mono_tags, lifting) \H \ borel_measurable M\ infinity_ennreal_def nn_integral_noteq_infinite top.not_eq_extremum) have "convergent (\n. f n x)" if "H x \ \" "incseq (\n. f n x)" for x proof - define A where "A = enn2real(H x)" then have "H x = ennreal A" using \H x \ \\ by (simp add: ennreal_enn2real_if) have "f n x \ f 0 x + A powr (1/p2)" for n proof - have "ennreal(\g n x\ powr p2) \ ennreal A" unfolding \H x = ennreal A\[symmetric] H_def h_def by (meson SUP_upper2 UNIV_I order_refl) then have "\g n x\ powr p2 \ A" by (subst ennreal_le_iff[symmetric], auto simp add: A_def) have "\g n x\ = (\g n x\ powr p2) powr (1/p2)" using \p2 > 0\ by (simp add: powr_powr) also have "... \ A powr (1/p2)" apply (rule powr_mono2) using \p2 > 0\ \\g n x\ powr p2 \ A\ by auto finally have "\g n x\ \ A powr (1/p2)" by simp then show ?thesis unfolding g_def by auto qed then show "convergent (\n. f n x)" using LIMSEQ_incseq_SUP[OF _ \incseq (\n. f n x)\] convergent_def by (metis bdd_aboveI2) qed then show "AE x in M. convergent (\n. f n x)" using \AE x in M. H x \ \\ \AE x in M. incseq (\n. f n x)\ by auto qed (insert \p>0\, simp) then have lim: "AE x in M. (\n. f n x) \ lim (\n. f n x)" using convergent_LIMSEQ_iff by auto show "(\x. lim (\n. f n x)) \ space\<^sub>N (\ p M)" apply (rule Lp_AE_limit''[of _ _ f, OF _ \\n. f n \ space\<^sub>N (\ p M)\ lim \\n. Norm (\ p M) (f n) \ C\]) by auto show "Norm (\ p M) (\x. lim (\n. f n x)) \ C" apply (rule Lp_AE_limit''[of _ _ f, OF _ \\n. f n \ space\<^sub>N (\ p M)\ lim \\n. Norm (\ p M) (f n) \ C\]) by auto qed subsection \Completeness of $L^p$\ text \We prove the completeness of $L^p$.\ theorem Lp_complete: "complete\<^sub>N (\ p M)" proof (cases rule: Lp_cases[of p]) case zero show ?thesis proof (rule complete\<^sub>N_I) fix u assume "\(n::nat). u n \ space\<^sub>N (\ p M)" then have "tendsto_in\<^sub>N (\ p M) u 0" unfolding tendsto_in\<^sub>N_def Norm_def \p = 0\ L_zero(1) L_zero_space by auto then show "\x\space\<^sub>N (\ p M). tendsto_in\<^sub>N (\ p M) u x" by auto qed next case (real_pos p2) show ?thesis proof (rule complete\<^sub>N_I'[of "\n. (1/2)^n * (1/(defect (\ p M))^(Suc n))"], unfold \p = ennreal p2\) show "0 < (1/2) ^ n * (1 / defect (\ (ennreal p2) M) ^ Suc n)" for n using defect_ge_1[of "\ (ennreal p2) M"] by (auto simp add: divide_simps) fix u assume "\(n::nat). u n \ space\<^sub>N (\ p2 M)" "\n. Norm (\ p2 M) (u n) \ (1/2)^n * (1/(defect (\ p2 M))^(Suc n))" then have H: "\n. u n \ space\<^sub>N (\ p2 M)" "\n. Norm (\ p2 M) (u n) \ (1/2) ^ n * (1/(defect (\ p2 M))^(Suc n))" unfolding \p = ennreal p2\ by auto have [measurable]: "u n \ borel_measurable M" for n using Lp_measurable[OF H(1)]. define w where "w = (\N x. (\n\{..u n x\))" have w2: "w = (\N. sum (\n x. \u n x\) {..N. w N x)" for x unfolding w2 by (rule incseq_SucI, auto) then have wN_inc: "AE x in M. incseq (\N. w N x)" by simp have abs_u_space: "(\x. \u n x\) \ space\<^sub>N (\ p2 M)" for n by (rule Lp_Banach_lattice[OF \u n \ space\<^sub>N (\ p2 M)\]) then have wN_space: "w N \ space\<^sub>N (\ p2 M)" for N unfolding w2 using H(1) by auto have abs_u_Norm: "Norm (\ p2 M) (\x. \u n x\) \ (1/2) ^ n * (1/(defect (\ p2 M))^(Suc n))" for n using Lp_Banach_lattice(2)[OF \u n \ space\<^sub>N (\ p2 M)\] H(2) by auto have wN_Norm: "Norm (\ p2 M) (w N) \ 2" for N proof - have *: "(defect (\ p2 M))^(Suc n) \ 0" "(defect (\ p2 M))^(Suc n) > 0" for n using defect_ge_1[of "\ p2 M"] by auto have "Norm (\ p2 M) (w N) \ (\n p2 M))^(Suc n) * Norm (\ p2 M) (\x. \u n x\))" unfolding w2 lessThan_Suc_atMost[symmetric] by (rule Norm_sum, simp add: abs_u_space) also have "... \ (\n p2 M))^(Suc n) * ((1/2) ^ n * (1/(defect (\ p2 M))^(Suc n))))" apply (rule sum_mono, rule mult_left_mono) using abs_u_Norm * by auto also have "... = (\n p2 M"] by (auto simp add: algebra_simps) also have "... \ (\n. (1/2) ^ n)" unfolding lessThan_Suc_atMost[symmetric] by (rule sum_le_suminf, rule summable_geometric[of "1/2"], auto) also have "... = 2" using suminf_geometric[of "1/2"] by auto finally show ?thesis by simp qed have "AE x in M. convergent (\N. w N x)" apply (rule Lp_monotone_limit[OF \p > 0\, of _ _ 2], unfold \p = ennreal p2\) using wN_inc wN_Norm wN_space by auto define m where "m = (\x. lim (\N. w N x))" have m_space: "m \ space\<^sub>N (\ p2 M)" unfolding m_def \p = ennreal p2\[symmetric] apply (rule Lp_monotone_limit[OF \p > 0\, of _ _ 2], unfold \p = ennreal p2\) using wN_inc wN_Norm wN_space by auto define v where "v = (\x. (\n. u n x))" have v_meas: "v \ borel_measurable M" unfolding v_def by auto have u_meas: "\n. (sum u {0.. borel_measurable M" by auto { fix x assume "convergent (\N. w N x)" then have S: "summable (\n. \u n x\)" unfolding w_def using summable_iff_convergent by auto then have "m x = (\n. \u n x\)" unfolding m_def w_def by (metis suminf_eq_lim) have "summable (\n. u n x)" using S by (rule summable_rabs_cancel) then have *: "(\n. (sum u {.. v x" unfolding v_def fun_sum_apply by (metis convergent_LIMSEQ_iff suminf_eq_lim summable_iff_convergent) have "\(sum u {.. \ m x" for n proof - have "\(sum u {.. \ (\i\{..u i x\)" unfolding fun_sum_apply by auto also have "... \ (\i. \u i x\)" apply (rule sum_le_suminf) using S by auto finally show ?thesis using \m x = (\n. \u n x\)\ by simp qed then have "(\n. \(sum u {0.. \ m x) \ (\n. (sum u {0.. v x" unfolding atLeast0LessThan using * by auto } then have m_bound: "\n. AE x in M. \(sum u {0.. \ m x" and u_conv: "AE x in M. (\n. (sum u {0.. v x" using \AE x in M. convergent (\N. w N x)\ by auto have "tendsto_in\<^sub>N (\ p2 M) (\n. sum u {0.. space\<^sub>N (\ p2 M)" by (rule Lp_domination_limit[OF v_meas u_meas m_space u_conv m_bound]) ultimately show "\v \ space\<^sub>N (\ p2 M). tendsto_in\<^sub>N (\ p2 M) (\n. sum u {0..N_I'[of "\n. (1/2)^n"]) fix u assume "\(n::nat). u n \ space\<^sub>N (\ p M)" "\n. Norm (\ p M) (u n) \ (1/2) ^ n" then have H: "\n. u n \ space\<^sub>N (\ \ M)" "\n. Norm (\ \ M) (u n) \ (1/2) ^ n" using PInf by auto have [measurable]: "u n \ borel_measurable M" for n using Lp_measurable[OF H(1)] by auto define v where "v = (\x. \n. u n x)" have [measurable]: "v \ borel_measurable M" unfolding v_def by auto define w where "w = (\N x. (\n\{0.. borel_measurable M" for N unfolding w_def by auto have "AE x in M. \u n x\ \ (1/2)^n" for n using L_infinity_AE_bound[OF H(1), of n] H(2)[of n] by auto then have "AE x in M. \n. \u n x\ \ (1/2)^n" by (subst AE_all_countable, auto) moreover have "\w N x - v x\ \ (1/2)^N * 2" if "\n. \u n x\ \ (1/2)^n" for N x proof - have *: "\n. \u n x\ \ (1/2)^n" using that by auto have **: "summable (\n. \u n x\)" apply (rule summable_norm_cancel, rule summable_comparison_test'[OF summable_geometric[of "1/2"]]) using * by auto have "\w N x - v x\ = \(\n. u (n + N) x)\" unfolding v_def w_def apply (subst suminf_split_initial_segment[OF summable_rabs_cancel[OF \summable (\n. \u n x\)\], of "N"]) by (simp add: lessThan_atLeast0) also have "... \ (\n. \u (n + N) x\)" apply (rule summable_rabs, subst summable_iff_shift) using ** by auto also have "... \ (\n. (1/2)^(n + N))" proof (rule suminf_le) show "\n. \u (n + N) x\ \ (1/2) ^ (n + N)" using *[of "_ + N"] by simp show "summable (\n. \u (n + N) x\)" using ** by (subst summable_iff_shift) simp show "summable (\n. (1/2::real) ^ (n + N))" using summable_geometric [of "1/2"] by (subst summable_iff_shift) simp qed also have "... = (1/2)^N * (\n. (1/2)^n)" by (subst power_add, subst suminf_mult2[symmetric], auto simp add: summable_geometric[of "1/2"]) also have "... = (1/2)^N * 2" by (subst suminf_geometric, auto) finally show ?thesis by simp qed ultimately have *: "AE x in M. \w N x - v x\ \ (1/2)^N * 2" for N by auto have **: "w N - v \ space\<^sub>N (\ \ M)" "Norm (\ \ M) (w N - v) \ (1/2)^N * 2" for N unfolding fun_diff_def using L_infinity_I[OF _ *] by auto have l: "(\N. ((1/2)^N) * (2::real)) \ 0 * 2" by (rule tendsto_mult, auto simp add: LIMSEQ_realpow_zero[of "1/2"]) have "tendsto_in\<^sub>N (\ \ M) w v" unfolding tendsto_in\<^sub>N_def apply (rule tendsto_sandwich[of "\_. 0" _ _ "\n. (1/2)^n * 2"]) using l **(2) by auto have "v = - (w 0 - v)" unfolding w_def by auto then have "v \ space\<^sub>N (\ \ M)" using **(1)[of 0] spaceN_add spaceN_diff by fastforce then show "\v \ space\<^sub>N (\ p M). tendsto_in\<^sub>N (\ p M) (\n. sum u {0..tendsto_in\<^sub>N (\ \ M) w v\ unfolding \p = \\ w_def fun_sum_apply[symmetric] by auto qed (simp) qed subsection \Multiplication of functions, duality\ text \The next theorem asserts that the multiplication of two functions in $L^p$ and $L^q$ belongs to $L^r$, where $r$ is determined by the equality $1/r = 1/p + 1/q$. This is essentially a case by case analysis, depending on the kind of $L^p$ space we are considering. The only nontrivial case is when $p$, $q$ (and $r$) are finite and nonzero. In this case, it reduces to H\"older inequality.\ theorem Lp_Lq_mult: fixes p q r::ennreal assumes "1/p + 1/q = 1/r" and "f \ space\<^sub>N (\ p M)" "g \ space\<^sub>N (\ q M)" shows "(\x. f x * g x) \ space\<^sub>N (\ r M)" "Norm (\ r M) (\x. f x * g x) \ Norm (\ p M) f * Norm (\ q M) g" proof - have [measurable]: "f \ borel_measurable M" "g \ borel_measurable M" using Lp_measurable assms by auto have "(\x. f x * g x) \ space\<^sub>N (\ r M) \ Norm (\ r M) (\x. f x * g x) \ Norm (\ p M) f * Norm (\ q M) g" proof (cases rule: Lp_cases[of r]) case zero have *: "(\x. f x * g x) \ borel_measurable M" by auto then have "Norm (\ r M) (\x. f x * g x) = 0" using L_zero[of M] unfolding Norm_def zero by auto then have "Norm (\ r M) (\x. f x * g x) \ Norm (\ p M) f * Norm (\ q M) g" using Norm_nonneg by auto then show ?thesis unfolding zero using * L_zero_space[of M] by auto next case (real_pos r2) have "p > 0" "q > 0" using \1/p + 1/q = 1/r\ \r > 0\ by (metis ennreal_add_eq_top ennreal_divide_eq_top_iff ennreal_top_neq_one gr_zeroI zero_neq_one)+ consider "p = \" | "q = \" | "p < \ \ q < \" using top.not_eq_extremum by force then show ?thesis proof (cases) case 1 then have "q = r" using \1/p + 1/q = 1/r\ by (metis ennreal_divide_top infinity_ennreal_def one_divide_one_divide_ennreal semiring_normalization_rules(5)) have "AE x in M. \f x\ \ Norm (\ p M) f" using \f \ space\<^sub>N (\ p M)\ L_infinity_AE_bound unfolding \p = \\ by auto then have *: "AE x in M. \f x * g x\ \ \Norm (\ p M) f * g x\" unfolding abs_mult using Norm_nonneg[of "\ p M" f] mult_right_mono by fastforce have **: "(\x. Norm (\ p M) f * g x) \ space\<^sub>N (\ r M)" using spaceN_cmult[OF \g \ space\<^sub>N (\ q M)\] unfolding \q = r\ scaleR_fun_def by simp have ***: "Norm (\ r M) (\x. Norm (\ p M) f * g x) = Norm (\ p M) f * Norm (\ q M) g" using Norm_cmult[of "\ r M"] unfolding \q = r\ scaleR_fun_def by auto then show ?thesis using Lp_domination[of "\x. f x * g x" M "\x. Norm (\ p M) f * g x" r] unfolding \q = r\ using * ** *** by auto next case 2 then have "p = r" using \1/p + 1/q = 1/r\ by (metis add.right_neutral ennreal_divide_top infinity_ennreal_def one_divide_one_divide_ennreal) have "AE x in M. \g x\ \ Norm (\ q M) g" using \g \ space\<^sub>N (\ q M)\ L_infinity_AE_bound unfolding \q = \\ by auto then have *: "AE x in M. \f x * g x\ \ \Norm (\ q M) g * f x\" apply (simp only: mult.commute[of "Norm (\ q M) g" _]) unfolding abs_mult using mult_left_mono Norm_nonneg[of "\ q M" g] by fastforce have **: "(\x. Norm (\ q M) g * f x) \ space\<^sub>N (\ r M)" using spaceN_cmult[OF \f \ space\<^sub>N (\ p M)\] unfolding \p = r\ scaleR_fun_def by simp have ***: "Norm (\ r M) (\x. Norm (\ q M) g * f x) = Norm (\ p M) f * Norm (\ q M) g" using Norm_cmult[of "\ r M"] unfolding \p = r\ scaleR_fun_def by auto then show ?thesis using Lp_domination[of "\x. f x * g x" M "\x. Norm (\ q M) g * f x" r] unfolding \p = r\ using * ** *** by auto next case 3 obtain p2 where "p = ennreal p2" "p2 > 0" using enn2real_positive_iff[of p] 3 \p > 0\ by (cases p) auto obtain q2 where "q = ennreal q2" "q2 > 0" using enn2real_positive_iff[of q] 3 \q > 0\ by (cases q) auto have "ennreal(1/r2) = 1/r" using \r = ennreal r2\ \r2 > 0\ divide_ennreal zero_le_one by fastforce also have "... = 1/p + 1/q" using assms by auto also have "... = ennreal(1/p2 + 1/q2)" using \p = ennreal p2\ \p2 > 0\ \q = ennreal q2\ \q2 > 0\ apply (simp only: divide_ennreal ennreal_1[symmetric]) using ennreal_plus[of "1/p2" "1/q2", symmetric] by auto finally have *: "1/r2 = 1/p2 + 1/q2" using ennreal_inj \p2 > 0\ \q2 > 0\ \r2 > 0\ by (metis divide_pos_pos ennreal_less_zero_iff le_less zero_less_one) define P where "P = p2 / r2" define Q where "Q = q2 / r2" have [simp]: "P > 0" "Q > 0" and "1/P + 1/Q = 1" using \p2 > 0\ \q2 > 0\ \r2 > 0\ * unfolding P_def Q_def by (auto simp add: divide_simps algebra_simps) have Pa: "(\z\ powr r2) powr P = \z\ powr p2" for z unfolding P_def powr_powr using \r2 > 0\ by auto have Qa: "(\z\ powr r2) powr Q = \z\ powr q2" for z unfolding Q_def powr_powr using \r2 > 0\ by auto have *: "integrable M (\x. \f x\ powr r2 * \g x\ powr r2)" apply (rule Holder_inequality[OF \P>0\ \Q>0\ \1/P + 1/Q = 1\], auto simp add: Pa Qa) using \f \ space\<^sub>N (\ p M)\ unfolding \p = ennreal p2\ using Lp_space[OF \p2 > 0\] apply auto using \g \ space\<^sub>N (\ q M)\ unfolding \q = ennreal q2\ using Lp_space[OF \q2 > 0\] by auto have "(\x. f x * g x) \ space\<^sub>N (\ r M)" unfolding \r = ennreal r2\ using Lp_space[OF \r2 > 0\, of M] by (auto simp add: * abs_mult powr_mult) have "Norm (\ r M) (\x. f x * g x) = (\x. \f x * g x\ powr r2 \M) powr (1/r2)" unfolding \r = ennreal r2\ using Lp_Norm[OF \r2 > 0\, of _ M] by auto also have "... = abs (\x. \f x\ powr r2 * \g x\ powr r2 \M) powr (1/r2)" by (auto simp add: powr_mult abs_mult) also have "... \ ((\x. \ \f x\ powr r2 \ powr P \M) powr (1/P) * (\x. \ \g x\ powr r2 \ powr Q \M) powr (1/Q)) powr (1/r2)" apply (rule powr_mono2, simp add: \r2 > 0\ less_imp_le, simp) apply (rule Holder_inequality[OF \P>0\ \Q>0\ \1/P + 1/Q = 1\], auto simp add: Pa Qa) using \f \ space\<^sub>N (\ p M)\ unfolding \p = ennreal p2\ using Lp_space[OF \p2 > 0\] apply auto using \g \ space\<^sub>N (\ q M)\ unfolding \q = ennreal q2\ using Lp_space[OF \q2 > 0\] by auto also have "... = (\x. \f x\ powr p2 \M) powr (1/p2) * (\x. \g x\ powr q2 \M) powr (1/q2)" apply (auto simp add: powr_mult powr_powr) unfolding P_def Q_def using \r2 > 0\ by auto also have "... = Norm (\ p M) f * Norm (\ q M) g" unfolding \p = ennreal p2\ \q = ennreal q2\ using Lp_Norm[OF \p2 > 0\, of _ M] Lp_Norm[OF \q2 > 0\, of _ M] by auto finally show ?thesis using \(\x. f x * g x) \ space\<^sub>N (\ r M)\ by auto qed next case PInf then have "p = \" "q = r" using \1/p + 1/q = 1/r\ by (metis add_eq_0_iff_both_eq_0 ennreal_divide_eq_0_iff infinity_ennreal_def not_one_le_zero order.order_iff_strict)+ have "AE x in M. \f x\ \ Norm (\ p M) f" using \f \ space\<^sub>N (\ p M)\ L_infinity_AE_bound unfolding \p = \\ by auto then have *: "AE x in M. \f x * g x\ \ \Norm (\ p M) f * g x\" unfolding abs_mult using Norm_nonneg[of "\ p M" f] mult_right_mono by fastforce have **: "(\x. Norm (\ p M) f * g x) \ space\<^sub>N (\ r M)" using spaceN_cmult[OF \g \ space\<^sub>N (\ q M)\] unfolding \q = r\ scaleR_fun_def by simp have ***: "Norm (\ r M) (\x. Norm (\ p M) f * g x) = Norm (\ p M) f * Norm (\ q M) g" using Norm_cmult[of "\ r M"] unfolding \q = r\ scaleR_fun_def by auto then show ?thesis using Lp_domination[of "\x. f x * g x" M "\x. Norm (\ p M) f * g x" r] unfolding \q = r\ using * ** *** by auto qed then show "(\x. f x * g x) \ space\<^sub>N (\ r M)" "Norm (\ r M) (\x. f x * g x) \ Norm (\ p M) f * Norm (\ q M) g" by auto qed text \The previous theorem admits an eNorm version in which one does not assume a priori that the functions under consideration belong to $L^p$ or $L^q$.\ theorem Lp_Lq_emult: fixes p q r::ennreal assumes "1/p + 1/q = 1/r" "f \ borel_measurable M" "g \ borel_measurable M" shows "eNorm (\ r M) (\x. f x * g x) \ eNorm (\ p M) f * eNorm (\ q M) g" proof (cases "r = 0") case True then have "eNorm (\ r M) (\x. f x * g x) = 0" using assms by (simp add: L_zero(1)) then show ?thesis by auto next case False then have "r > 0" using not_gr_zero by blast then have "p > 0" "q > 0" using \1/p + 1/q = 1/r\ by (metis ennreal_add_eq_top ennreal_divide_eq_top_iff ennreal_top_neq_one gr_zeroI zero_neq_one)+ then have Z: "zero_space\<^sub>N (\ p M) = {f \ borel_measurable M. AE x in M. f x = 0}" "zero_space\<^sub>N (\ q M) = {f \ borel_measurable M. AE x in M. f x = 0}" "zero_space\<^sub>N (\ r M) = {f \ borel_measurable M. AE x in M. f x = 0}" using \r > 0\ Lp_infinity_zero_space by auto have [measurable]: "(\x. f x * g x) \ borel_measurable M" using assms by auto consider "eNorm (\ p M) f = 0 \ eNorm (\ q M) g = 0" | "(eNorm (\ p M) f > 0 \ eNorm (\ q M) g = \) \ (eNorm (\ p M) f = \ \ eNorm (\ q M) g > 0)" | "eNorm (\ p M) f < \ \ eNorm (\ q M) g < \" using less_top by fastforce then show ?thesis proof (cases) case 1 then have "(AE x in M. f x = 0) \ (AE x in M. g x = 0)" using Z unfolding zero_space\<^sub>N_def by auto then have "AE x in M. f x * g x = 0" by auto then have "eNorm (\ r M) (\x. f x * g x) = 0" using Z unfolding zero_space\<^sub>N_def by auto then show ?thesis by simp next case 2 then have "eNorm (\ p M) f * eNorm (\ q M) g = \" using ennreal_mult_eq_top_iff by force then show ?thesis by auto next case 3 then have *: "f \ space\<^sub>N (\ p M)" "g \ space\<^sub>N (\ q M)" unfolding space\<^sub>N_def by auto then have "(\x. f x * g x) \ space\<^sub>N (\ r M)" using Lp_Lq_mult(1)[OF assms(1)] by auto then show ?thesis using Lp_Lq_mult(2)[OF assms(1) *] by (simp add: eNorm_Norm * ennreal_mult'[symmetric]) qed qed lemma Lp_Lq_duality_bound: fixes p q::ennreal assumes "1/p + 1/q = 1" "f \ space\<^sub>N (\ p M)" "g \ space\<^sub>N (\ q M)" shows "integrable M (\x. f x * g x)" "abs(\x. f x * g x \M) \ Norm (\ p M) f * Norm (\ q M) g" proof - have "(\x. f x * g x) \ space\<^sub>N (\ 1 M)" apply (rule Lp_Lq_mult[OF _ \f \ space\<^sub>N (\ p M)\ \g \ space\<^sub>N (\ q M)\]) using \1/p + 1/q = 1\ by auto then show "integrable M (\x. f x * g x)" using L1_space by auto have "abs(\x. f x * g x \M) \ Norm (\ 1 M) (\x. f x * g x)" using L1_int_ineq by auto also have "... \ Norm (\ p M) f * Norm (\ q M) g" apply (rule Lp_Lq_mult[OF _ \f \ space\<^sub>N (\ p M)\ \g \ space\<^sub>N (\ q M)\]) using \1/p + 1/q = 1\ by auto finally show "abs(\x. f x * g x \M) \ Norm (\ p M) f * Norm (\ q M) g" by simp qed text \The next theorem asserts that the norm of an $L^p$ function $f$ can be obtained by estimating the integrals of $fg$ over all $L^q$ functions $g$, where $1/p + 1/q = 1$. When $p = \infty$, it is necessary to assume that the space is sigma-finite: for instance, if the space is one single atom of infinite mass, then there is no nonzero $L^1$ function, so taking for $f$ the constant function equal to $1$, it has $L^\infty$ norm equal to $1$, but $\int fg = 0$ for all $L^1$ function $g$.\ theorem Lp_Lq_duality: fixes p q::ennreal assumes "f \ space\<^sub>N (\ p M)" "1/p + 1/q = 1" "p = \ \ sigma_finite_measure M" shows "bdd_above ((\g. (\x. f x * g x \M))`{g \ space\<^sub>N (\ q M). Norm (\ q M) g \ 1})" "Norm (\ p M) f = (SUP g\{g \ space\<^sub>N (\ q M). Norm (\ q M) g \ 1}. (\x. f x * g x \M))" proof - have [measurable]: "f \ borel_measurable M" using Lp_measurable[OF assms(1)] by auto have B: "(\x. f x * g x \M) \ Norm (\ p M) f" if "g \ {g \ space\<^sub>N (\ q M). Norm (\ q M) g \ 1}" for g proof - have g: "g \ space\<^sub>N (\ q M)" "Norm (\ q M) g \ 1" using that by auto have "(\x. f x * g x \M) \ abs(\x. f x * g x \M)" by auto also have "... \ Norm (\ p M) f * Norm (\ q M) g" using Lp_Lq_duality_bound(2)[OF \1/p + 1/q = 1\ \f \ space\<^sub>N (\ p M)\ g(1)] by auto also have "... \ Norm (\ p M) f" using g(2) Norm_nonneg[of "\ p M" f] mult_left_le by blast finally show "(\x. f x * g x \M) \ Norm (\ p M) f" by simp qed then show "bdd_above ((\g. (\x. f x * g x \M))`{g \ space\<^sub>N (\ q M). Norm (\ q M) g \ 1})" by (meson bdd_aboveI2) show "Norm (\ p M) f = (SUP g\{g \ space\<^sub>N (\ q M). Norm (\ q M) g \ 1}. (\x. f x * g x \M))" proof (rule antisym) show "(SUP g\{g \ space\<^sub>N (\ q M). Norm (\ q M) g \ 1}. \x. f x * g x \M) \ Norm (\ p M) f" by (rule cSUP_least, auto, rule exI[of _ 0], auto simp add: B) have "p \ 1" using conjugate_exponent_ennrealI(1)[OF \1/p + 1/q = 1\] by simp show "Norm (\ p M) f \ (SUP g\{g \ space\<^sub>N (\ q M). Norm (\ q M) g \ 1}. (\x. f x * g x \M))" using \p \ 1\ proof (cases rule: Lp_cases_1_PInf) case PInf then have "f \ space\<^sub>N (\ \ M)" using \f \ space\<^sub>N(\ p M)\ by simp have "q = 1" using \1/p + 1/q = 1\ \p = \\ by (simp add: divide_eq_1_ennreal) have "c \ (SUP g\{g \ space\<^sub>N (\ q M). Norm (\ q M) g \ 1}. (\x. f x * g x \M))" if "c < Norm (\ p M) f" for c proof (cases "c < 0") case True then have "c \ (\x. f x * 0 x \M)" by auto also have "... \ (SUP g\{g \ space\<^sub>N (\ q M). Norm (\ q M) g \ 1}. (\x. f x * g x \M))" apply (rule cSUP_upper, auto simp add: zero_fun_def[symmetric]) using B by (meson bdd_aboveI2) finally show ?thesis by simp next case False then have "ennreal c < eNorm (\ \ M) f" using eNorm_Norm[OF \f \ space\<^sub>N (\ p M)\] that ennreal_less_iff unfolding \p = \\ by auto then have *: "emeasure M {x \ space M. \f x\ > c} > 0" using L_infinity_pos_measure[of f M c] by auto obtain A where [measurable]: "\(n::nat). A n \ sets M" and "(\i. A i) = space M" "\i. emeasure M (A i) \ \" using sigma_finite_measure.sigma_finite[OF \p = \ \ sigma_finite_measure M\[OF \p = \\]] by (metis UNIV_I sets_range) define Y where "Y = (\n::nat. {x \ A n. \f x\ > c})" have [measurable]: "Y n \ sets M" for n unfolding Y_def by auto have "{x \ space M. \f x\ > c} = (\n. Y n)" unfolding Y_def using \(\i. A i) = space M\ by auto then have "emeasure M (\n. Y n) > 0" using * by auto then obtain n where "emeasure M (Y n) > 0" using emeasure_pos_unionE[of Y, OF \\n. Y n \ sets M\] by auto have "emeasure M (Y n) \ emeasure M (A n)" apply (rule emeasure_mono) unfolding Y_def by auto then have "emeasure M (Y n) \ \" using \emeasure M (A n) \ \\ by (metis infinity_ennreal_def neq_top_trans) then have "measure M (Y n) > 0" using \emeasure M (Y n) > 0\ unfolding measure_def by (simp add: enn2real_positive_iff top.not_eq_extremum) have "\f x\ \ c" if "x \ Y n" for x using that less_imp_le unfolding Y_def by auto define g where "g = (\x. indicator (Y n) x * sgn(f x)) /\<^sub>R measure M (Y n)" have "g \ space\<^sub>N (\ 1 M)" apply (rule Lp_domination[of _ _ "indicator (Y n) /\<^sub>R measure M (Y n)"]) unfolding g_def using L1_indicator'[OF \Y n \ sets M\ \emeasure M (Y n) \ \\] by (auto simp add: abs_mult indicator_def abs_sgn_eq) have "Norm (\ 1 M) g = Norm (\ 1 M) (\x. indicator (Y n) x * sgn(f x)) / abs(measure M (Y n))" unfolding g_def Norm_cmult by (simp add: divide_inverse) also have "... \ Norm (\ 1 M) (indicator (Y n)) / abs(measure M (Y n))" using \measure M (Y n) > 0\ apply (auto simp add: divide_simps) apply (rule Lp_domination) using L1_indicator'[OF \Y n \ sets M\ \emeasure M (Y n) \ \\] by (auto simp add: abs_mult indicator_def abs_sgn_eq) also have "... = measure M (Y n) / abs(measure M (Y n))" using L1_indicator'[OF \Y n \ sets M\ \emeasure M (Y n) \ \\] by (auto simp add: abs_mult indicator_def abs_sgn_eq) also have "... = 1" using \measure M (Y n) > 0\ by auto finally have "Norm (\ 1 M) g \ 1" by simp have "c * measure M (Y n) = (\x. c * indicator (Y n) x \M)" using \measure M (Y n) > 0\ \emeasure M (Y n) \ \\ by auto also have "... \ (\x. \f x\ * indicator (Y n) x \M)" apply (rule integral_mono) using \emeasure M (Y n) \ \\ \0 < Sigma_Algebra.measure M (Y n)\ not_integrable_integral_eq apply fastforce apply (rule Bochner_Integration.integrable_bound[of _ "\x. Norm (\ \ M) f * indicator (Y n) x"]) using \emeasure M (Y n) \ \\ \0 < Sigma_Algebra.measure M (Y n)\ not_integrable_integral_eq apply fastforce using L_infinity_AE_bound[OF \f \ space\<^sub>N (\ \ M)\] by (auto simp add: indicator_def Y_def) finally have "c \ (\x. \f x\ * indicator (Y n) x \M) / measure M (Y n)" using \measure M (Y n) > 0\ by (auto simp add: divide_simps) also have "... = (\x. f x * indicator (Y n) x * sgn(f x) / measure M (Y n) \M)" using \measure M (Y n) > 0\ by (simp add: abs_sgn mult.commute mult.left_commute) also have "... = (\x. f x * g x \M)" unfolding divide_inverse g_def divideR_apply by (auto simp add: algebra_simps) also have "... \ (SUP g\{g \ space\<^sub>N (\ q M). Norm (\ q M) g \ 1}. (\x. f x * g x \M))" unfolding \q = 1\ apply (rule cSUP_upper, auto) using \g \ space\<^sub>N (\ 1 M)\ \Norm (\ 1 M) g \ 1\ apply auto using B \p = \\ \q = 1\ by (meson bdd_aboveI2) finally show ?thesis by simp qed then show ?thesis using dense_le by auto next case one then have "q = \" using \1/p + 1/q = 1\ by simp define g where "g = (\x. sgn (f x))" have [measurable]: "g \ space\<^sub>N (\ \ M)" apply (rule L_infinity_I[of g M 1]) unfolding g_def by (auto simp add: abs_sgn_eq) have "Norm (\ \ M) g \ 1" apply (rule L_infinity_I[of g M 1]) unfolding g_def by (auto simp add: abs_sgn_eq) have "Norm (\ p M) f = (\x. \f x\ \M)" unfolding \p = 1\ apply (rule L1_D(3)) using \f \ space\<^sub>N (\ p M)\ unfolding \p = 1\ by auto also have "... = (\x. f x * g x \M)" unfolding g_def by (simp add: abs_sgn) also have "... \ (SUP g\{g \ space\<^sub>N (\ q M). Norm (\ q M) g \ 1}. (\x. f x * g x \M))" unfolding \q = \\ apply (rule cSUP_upper, auto) using \g \ space\<^sub>N (\ \ M)\ \Norm (\ \ M) g \ 1\ apply auto using B \q = \\ by fastforce finally show ?thesis by simp next case (gr p2) then have "p2 > 0" by simp have "f \ space\<^sub>N (\ p2 M)" using \f \ space\<^sub>N (\ p M)\ \p = ennreal p2\ by auto define q2 where "q2 = conjugate_exponent p2" have "q2 > 1" "q2 > 0" using conjugate_exponent_real(2)[OF \p2 > 1\] unfolding q2_def by auto have "q = ennreal q2" unfolding q2_def conjugate_exponent_real_ennreal[OF \p2 > 1\, symmetric] \p = ennreal p2\[symmetric] using conjugate_exponent_ennreal_iff[OF \p \ 1\] \1/p + 1/q = 1\ by auto show ?thesis proof (cases "Norm (\ p M) f = 0") case True then have "Norm (\ p M) f \ (\x. f x * 0 x \M)" by auto also have "... \ (SUP g\{g \ space\<^sub>N (\ q M). Norm (\ q M) g \ 1}. (\x. f x * g x \M))" apply (rule cSUP_upper, auto simp add: zero_fun_def[symmetric]) using B by (meson bdd_aboveI2) finally show ?thesis by simp next case False then have "Norm (\ p2 M) f > 0" unfolding \p = ennreal p2\ using Norm_nonneg[of "\ p2 M" f] by linarith define h where "h = (\x. sgn(f x) * \f x\ powr (p2 - 1))" have [measurable]: "h \ borel_measurable M" unfolding h_def by auto have "(\\<^sup>+x. \h x\ powr q2 \M) = (\\<^sup>+x. (\f x\ powr (p2 - 1)) powr q2 \M)" unfolding h_def by (rule nn_integral_cong, auto simp add: abs_mult abs_sgn_eq) also have "... = (\\<^sup>+x. \f x\ powr p2 \M)" unfolding powr_powr q2_def using conjugate_exponent_real(4)[OF \p2 > 1\] by auto also have "... = (Norm (\ p2 M) f) powr p2" apply (subst Lp_Norm(2), auto simp add: \p2 > 0\) by (rule nn_integral_eq_integral, auto simp add: Lp_D(2)[OF \p2 > 0\ \f \ space\<^sub>N (\ p2 M)\]) finally have *: "(\\<^sup>+x. \h x\ powr q2 \M) = (Norm (\ p2 M) f) powr p2" by simp have "integrable M (\x. \h x\ powr q2)" apply (rule integrableI_bounded, auto) using * by auto then have "(\x. \h x\ powr q2 \M) = (\\<^sup>+x. \h x\ powr q2 \M)" by (rule nn_integral_eq_integral[symmetric], auto) then have **: "(\x. \h x\ powr q2 \M) = (Norm (\ p2 M) f) powr p2" using * by auto define g where "g = (\x. h x / (Norm (\ p2 M) f) powr (p2 / q2))" have [measurable]: "g \ borel_measurable M" unfolding g_def by auto have intg: "integrable M (\x. \g x\ powr q2)" unfolding g_def using \Norm (\ p2 M) f > 0\ \q2 > 1\ apply (simp add: abs_mult powr_divide powr_powr) using \integrable M (\x. \h x\ powr q2)\ integrable_divide_zero by blast have "g \ space\<^sub>N (\ q2 M)" by (rule Lp_I(1)[OF \q2 > 0\ _ intg], auto) have "(\x. \g x\ powr q2 \M) = 1" unfolding g_def using \Norm (\ p2 M) f > 0\ \q2 > 1\ by (simp add: abs_mult powr_divide powr_powr **) then have "Norm (\ q2 M) g = 1" apply (subst Lp_D[OF \q2 > 0\]) using \g \ space\<^sub>N (\ q2 M)\ by auto have "(\x. f x * g x \M) = (\x. f x * sgn(f x) * \f x\ powr (p2 - 1) / (Norm (\ p2 M) f) powr (p2 / q2) \M)" unfolding g_def h_def by (simp add: mult.assoc) also have "... = (\x. \f x\ * \f x\ powr (p2-1) \M) / (Norm (\ p2 M) f) powr (p2 / q2)" by (auto simp add: abs_sgn) also have "... = (\x. \f x\ powr p2 \M) / (Norm (\ p2 M) f) powr (p2 / q2)" by (subst powr_mult_base, auto) also have "... = (Norm (\ p2 M) f) powr p2 / (Norm (\ p2 M) f) powr (p2 / q2)" by (subst Lp_Norm(2)[OF \p2 > 0\], auto) also have "... = (Norm (\ p2 M) f) powr (p2 - p2/q2)" by (simp add: powr_diff [symmetric] ) also have "... = Norm (\ p2 M) f" unfolding q2_def using conjugate_exponent_real(5)[OF \p2 > 1\] by auto finally have "Norm (\ p M) f = (\x. f x * g x \M)" unfolding \p = ennreal p2\ by simp also have "... \ (SUP g\{g \ space\<^sub>N (\ q M). Norm (\ q M) g \ 1}. (\x. f x * g x \M))" unfolding \q = ennreal q2\ apply (rule cSUP_upper, auto) using \g \ space\<^sub>N (\ q2 M)\ \Norm (\ q2 M) g = 1\ apply auto using B \q = ennreal q2\ by fastforce finally show ?thesis by simp qed qed qed qed text \The previous theorem admits a version in which one does not assume a priori that the function under consideration belongs to $L^p$. This gives an efficient criterion to check if a function is indeed in $L^p$. In this case, it is always necessary to assume that the measure is sigma-finite. Note that, in the statement, the Bochner integral $\int fg$ vanishes by definition if $fg$ is not integrable. Hence, the statement really says that the eNorm can be estimated using functions $g$ for which $fg$ is integrable. It is precisely the construction of such functions $g$ that requires the space to be sigma-finite.\ theorem Lp_Lq_duality': fixes p q::ennreal assumes "1/p + 1/q = 1" "sigma_finite_measure M" and [measurable]: "f \ borel_measurable M" shows "eNorm (\ p M) f = (SUP g\{g \ space\<^sub>N (\ q M). Norm (\ q M) g \ 1}. ennreal(\x. f x * g x \M))" proof (cases "eNorm (\ p M) f \ \") case True then have "f \ space\<^sub>N (\ p M)" unfolding space\<^sub>N_def by (simp add: top.not_eq_extremum) show ?thesis unfolding eNorm_Norm[OF \f \ space\<^sub>N (\ p M)\] Lp_Lq_duality[OF \f \ space\<^sub>N (\ p M)\ \1/p + 1/q = 1\ \sigma_finite_measure M\] apply (rule SUP_real_ennreal[symmetric], auto, rule exI[of _ 0], auto) by (rule Lp_Lq_duality[OF \f \ space\<^sub>N (\ p M)\ \1/p + 1/q = 1\ \sigma_finite_measure M\]) next case False have B: "\g \ {g \ space\<^sub>N (\ q M). Norm (\ q M) g \ 1}. (\x. f x * g x \M) \ C" if "C < \" for C::ennreal proof - obtain Cr where "C = ennreal Cr" "Cr \ 0" using \C < \\ ennreal_cases less_irrefl by auto obtain A where A: "\n::nat. A n \ sets M" "incseq A" "(\n. A n) = space M" "\n. emeasure M (A n) \ \" using sigma_finite_measure.sigma_finite_incseq[OF \sigma_finite_measure M\] by (metis range_subsetD) define Y where "Y = (\n. {x \ A n. \f x\ \ n})" have [measurable]: "\n. Y n \ sets M" unfolding Y_def using \\n::nat. A n \ sets M\ by auto have "incseq Y" apply (rule incseq_SucI) unfolding Y_def using incseq_SucD[OF \incseq A\] by auto have *: "\N. \n \ N. f x * indicator (Y n) x = f x" if "x \ space M" for x proof - obtain n0 where n0: "x \ A n0" using \x \ space M\ \(\n. A n) = space M\ by auto obtain n1::nat where n1: "\f x\ \ n1" using real_arch_simple by blast have "x \ Y (max n0 n1)" unfolding Y_def using n1 apply auto using n0 \incseq A\ incseq_def max.cobounded1 by blast then have *: "x \ Y n" if "n \ max n0 n1" for n using \incseq Y\ that incseq_def by blast show ?thesis by (rule exI[of _ "max n0 n1"], auto simp add: *) qed have *: "(\n. f x * indicator (Y n) x) \ f x" if "x \ space M" for x using *[OF that] unfolding eventually_sequentially[symmetric] by (simp add: tendsto_eventually) have "liminf (\n. eNorm (\ p M) (\x. f x * indicator (Y n) x)) \ eNorm (\ p M) f" apply (rule Lp_AE_limit) using * by auto then have "liminf (\n. eNorm (\ p M) (\x. f x * indicator (Y n) x)) > Cr" using False neq_top_trans by force then have "limsup (\n. eNorm (\ p M) (\x. f x * indicator (Y n) x)) > Cr" using Liminf_le_Limsup less_le_trans trivial_limit_sequentially by blast then obtain n where n: "eNorm (\ p M) (\x. f x * indicator (Y n) x) > Cr" using Limsup_obtain by blast have "(\x. f x * indicator (Y n) x) \ space\<^sub>N (\ p M)" apply (rule Lp_bounded_bounded_support[of _ _ n], auto) unfolding Y_def indicator_def apply auto by (metis (mono_tags, lifting) A(1) A(4) emeasure_mono infinity_ennreal_def mem_Collect_eq neq_top_trans subsetI) have "Norm (\ p M) (\x. f x * indicator (Y n) x) > Cr" using n unfolding eNorm_Norm[OF \(\x. f x * indicator (Y n) x) \ space\<^sub>N (\ p M)\] by (meson ennreal_leI not_le) then have "(SUP g\{g \ space\<^sub>N (\ q M). Norm (\ q M) g \ 1}. (\x. f x * indicator (Y n) x * g x \M)) > Cr" using Lp_Lq_duality(2)[OF \(\x. f x * indicator (Y n) x) \ space\<^sub>N (\ p M)\ \1/p + 1/q = 1\ \sigma_finite_measure M\] by auto then have "\g \ {g \ space\<^sub>N (\ q M). Norm (\ q M) g \ 1}. (\x. f x * indicator (Y n) x * g x \M) > Cr" apply (subst less_cSUP_iff[symmetric]) using Lp_Lq_duality(1)[OF \(\x. f x * indicator (Y n) x) \ space\<^sub>N (\ p M)\ \1/p + 1/q = 1\ \sigma_finite_measure M\] apply auto by (rule exI[of _ 0], auto) then obtain g where g: "g \ space\<^sub>N (\ q M)" "Norm (\ q M) g \ 1" "(\x. f x * indicator (Y n) x * g x \M) > Cr" by auto then have [measurable]: "g \ borel_measurable M" using Lp_measurable by auto define h where "h = (\x. indicator (Y n) x * g x)" have "Norm (\ q M) h \ Norm (\ q M) g" apply (rule Lp_domination[of _ _ g]) unfolding h_def indicator_def using \g \ space\<^sub>N (\ q M)\ by auto then have a: "Norm (\ q M) h \ 1" using \Norm (\ q M) g \ 1\ by auto have b: "h \ space\<^sub>N (\ q M)" apply (rule Lp_domination[of _ _ g]) unfolding h_def indicator_def using \g \ space\<^sub>N (\ q M)\ by auto have "(\x. f x * h x \M) > Cr" unfolding h_def using g(3) by (auto simp add: mult.assoc) then have "(\x. f x * h x \M) > C" unfolding \C = ennreal Cr\ using \Cr \ 0\ by (simp add: ennreal_less_iff) then show ?thesis using a b by auto qed have "(SUP g\{g \ space\<^sub>N (\ q M). Norm (\ q M) g \ 1}. ennreal(\x. f x * g x \M)) \ \" apply (rule dense_le) using B by (meson SUP_upper2) then show ?thesis using False neq_top_trans by force qed subsection \Conditional expectations and $L^p$\ text \The $L^p$ space with respect to a subalgebra is included in the whole $L^p$ space.\ lemma Lp_subalgebra: assumes "subalgebra M F" shows "\f. eNorm (\ p M) f \ eNorm (\ p (restr_to_subalg M F)) f" "(\ p (restr_to_subalg M F)) \\<^sub>N \ p M" "space\<^sub>N ((\ p (restr_to_subalg M F))) \ space\<^sub>N (\ p M)" "\f. f \ space\<^sub>N ((\ p (restr_to_subalg M F))) \ Norm (\ p M) f = Norm (\ p (restr_to_subalg M F)) f" proof - have *: "f \ space\<^sub>N (\ p M) \ Norm (\ p M) f = Norm (\ p (restr_to_subalg M F)) f" if "f \ space\<^sub>N (\ p (restr_to_subalg M F))" for f proof - have [measurable]: "f \ borel_measurable (restr_to_subalg M F)" using that Lp_measurable by auto then have [measurable]: "f \ borel_measurable M" using assms measurable_from_subalg measurable_in_subalg' by blast show ?thesis proof (cases rule: Lp_cases[of p]) case zero then show ?thesis using that unfolding \p = 0\ L_zero_space Norm_def L_zero by auto next case PInf have [measurable]: "f \ borel_measurable (restr_to_subalg M F)" using that Lp_measurable by auto then have [measurable]: "f \ borel_measurable F" using assms measurable_in_subalg' by blast then have [measurable]: "f \ borel_measurable M" using assms measurable_from_subalg by blast have "AE x in (restr_to_subalg M F). \f x\ \ Norm (\ \ (restr_to_subalg M F)) f" using L_infinity_AE_bound that unfolding \p = \\ by auto then have a: "AE x in M. \f x\ \ Norm (\ \ (restr_to_subalg M F)) f" using assms AE_restr_to_subalg by blast have *: "f \ space\<^sub>N (\ \ M)" "Norm (\ \ M) f \ Norm (\ \ (restr_to_subalg M F)) f" using L_infinity_I[OF \f \ borel_measurable M\ a] by auto then have b: "AE x in M. \f x\ \ Norm (\ \ M) f" using L_infinity_AE_bound by auto have c: "AE x in (restr_to_subalg M F). \f x\ \ Norm (\ \ M) f" apply (rule AE_restr_to_subalg2[OF assms]) using b by auto have "Norm (\ \ (restr_to_subalg M F)) f \ Norm (\ \ M) f" using L_infinity_I[OF \f \ borel_measurable (restr_to_subalg M F)\ c] by auto then show ?thesis using * unfolding \p = \\ by auto next case (real_pos p2) then have a [measurable]: "f \ space\<^sub>N (\ p2 (restr_to_subalg M F))" using that unfolding \p = ennreal p2\ by auto then have b [measurable]: "f \ space\<^sub>N (\ p2 M)" unfolding Lp_space[OF \p2 > 0\] using integrable_from_subalg[OF assms] by auto show ?thesis unfolding \p = ennreal p2\ Lp_D[OF \p2 > 0\ a] Lp_D[OF \p2 > 0\ b] using integral_subalgebra2[OF assms, symmetric, of f] apply (auto simp add: b) by (metis (mono_tags, lifting) \integrable (restr_to_subalg M F) (\x. \f x\ powr p2)\ assms integrableD(1) integral_subalgebra2 measurable_in_subalg') qed qed show "space\<^sub>N ((\ p (restr_to_subalg M F))) \ space\<^sub>N (\ p M)" using * by auto show "Norm (\ p M) f = Norm (\ p (restr_to_subalg M F)) f" if "f \ space\<^sub>N ((\ p (restr_to_subalg M F)))" for f using * that by auto show "eNorm (\ p M) f \ eNorm (\ p (restr_to_subalg M F)) f" for f by (metis "*" eNorm_Norm eq_iff infinity_ennreal_def less_imp_le spaceN_iff top.not_eq_extremum) then show "(\ p (restr_to_subalg M F)) \\<^sub>N \ p M" by (metis ennreal_1 mult.left_neutral quasinorm_subsetI) qed text \For $p \geq 1$, the conditional expectation of an $L^p$ function still belongs to $L^p$, with an $L^p$ norm which is bounded by the norm of the original function. This is wrong for $p < 1$. One can prove this separating the cases and using the conditional version of Jensen's inequality, but it is much more efficient to do it with duality arguments, as follows.\ proposition Lp_real_cond_exp: assumes [simp]: "subalgebra M F" and "p \ (1::ennreal)" "sigma_finite_measure (restr_to_subalg M F)" "f \ space\<^sub>N (\ p M)" shows "real_cond_exp M F f \ space\<^sub>N (\ p (restr_to_subalg M F))" "Norm (\ p (restr_to_subalg M F)) (real_cond_exp M F f) \ Norm (\ p M) f" proof - have [measurable]: "f \ borel_measurable M" using Lp_measurable assms by auto define q where "q = conjugate_exponent p" have "1/p + 1/q = 1" unfolding q_def using conjugate_exponent_ennreal[OF \p \ 1\] by simp have "eNorm (\ p (restr_to_subalg M F)) (real_cond_exp M F f) = (SUP g\{g \ space\<^sub>N (\ q (restr_to_subalg M F)). Norm (\ q (restr_to_subalg M F)) g \ 1}. ennreal(\x. (real_cond_exp M F f) x * g x \(restr_to_subalg M F)))" by (rule Lp_Lq_duality'[OF \1/p + 1/q = 1\ \sigma_finite_measure (restr_to_subalg M F)\], simp) also have "... \ (SUP g\{g \ space\<^sub>N (\ q M). Norm (\ q M) g \ 1}. ennreal(\x. f x * g x \M))" proof (rule SUP_mono, auto) fix g assume H: "g \ space\<^sub>N (\ q (restr_to_subalg M F))" "Norm (\ q (restr_to_subalg M F)) g \ 1" then have H2: "g \ space\<^sub>N (\ q M)" "Norm (\ q M) g \ 1" using Lp_subalgebra[OF \subalgebra M F\] by (auto simp add: subset_iff) have [measurable]: "g \ borel_measurable M" "g \ borel_measurable F" using Lp_measurable[OF H(1)] Lp_measurable[OF H2(1)] by auto have int: "integrable M (\x. f x * g x)" using Lp_Lq_duality_bound(1)[OF \1/p + 1/q = 1\ \f \ space\<^sub>N (\ p M)\ H2(1)]. have "(\x. (real_cond_exp M F f) x * g x \(restr_to_subalg M F)) = (\x. g x * (real_cond_exp M F f) x \M)" by (subst mult.commute, rule integral_subalgebra2[OF \subalgebra M F\], auto) also have "... = (\x. g x * f x \M)" apply (rule sigma_finite_subalgebra.real_cond_exp_intg, auto simp add: int mult.commute) unfolding sigma_finite_subalgebra_def using assms by auto finally have "ennreal (\x. (real_cond_exp M F f) x * g x \(restr_to_subalg M F)) \ ennreal (\x. f x * g x \M)" by (auto intro!: ennreal_leI simp add: mult.commute) then show "\m. m \ space\<^sub>N (\ q M) \ Norm (\ q M) m \ 1 \ ennreal (LINT x|restr_to_subalg M F. real_cond_exp M F f x * g x) \ ennreal (LINT x|M. f x * m x)" using H2 by blast qed also have "... = eNorm (\ p M) f" apply (rule Lp_Lq_duality'[OF \1/p + 1/q = 1\, symmetric], auto intro!: sigma_finite_subalgebra_is_sigma_finite[of _ F]) unfolding sigma_finite_subalgebra_def using assms by auto finally have *: "eNorm (\ p (restr_to_subalg M F)) (real_cond_exp M F f) \ eNorm (\ p M) f" by simp then show a: "real_cond_exp M F f \ space\<^sub>N (\ p (restr_to_subalg M F))" apply (subst spaceN_iff) using \f \ space\<^sub>N (\ p M)\ by (simp add: space\<^sub>N_def) show "Norm (\ p (restr_to_subalg M F)) (real_cond_exp M F f) \ Norm (\ p M) f" using * unfolding eNorm_Norm[OF \f \ space\<^sub>N (\ p M)\] eNorm_Norm[OF a] by simp qed lemma Lp_real_cond_exp_eNorm: assumes [simp]: "subalgebra M F" and "p \ (1::ennreal)" "sigma_finite_measure (restr_to_subalg M F)" shows "eNorm (\ p (restr_to_subalg M F)) (real_cond_exp M F f) \ eNorm (\ p M) f" proof (cases "eNorm (\ p M) f = \") case False then have *: "f \ space\<^sub>N (\ p M)" unfolding spaceN_iff by (simp add: top.not_eq_extremum) show ?thesis using Lp_real_cond_exp[OF assms \f \ space\<^sub>N (\ p M)\] by (subst eNorm_Norm, auto simp: \f \ space\<^sub>N (\ p M)\)+ qed (simp) end diff --git a/thys/Markov_Models/ex/PGCL.thy b/thys/Markov_Models/ex/PGCL.thy --- a/thys/Markov_Models/ex/PGCL.thy +++ b/thys/Markov_Models/ex/PGCL.thy @@ -1,262 +1,262 @@ (* Author: Johannes Hölzl *) section \Probabilistic Guarded Command Language (pGCL)\ theory PGCL imports "../Markov_Decision_Process" begin subsection \Syntax\ datatype 's pgcl = Skip | Abort | Assign "'s \ 's" | Seq "'s pgcl" "'s pgcl" | Par "'s pgcl" "'s pgcl" | If "'s \ bool" "'s pgcl" "'s pgcl" | Prob "bool pmf" "'s pgcl" "'s pgcl" | While "'s \ bool" "'s pgcl" subsection \Denotational Semantics\ primrec wp :: "'s pgcl \ ('s \ ennreal) \ ('s \ ennreal)" where "wp Skip f = f" | "wp Abort f = (\_. 0)" | "wp (Assign u) f = f \ u" | "wp (Seq c\<^sub>1 c\<^sub>2) f = wp c\<^sub>1 (wp c\<^sub>2 f)" | "wp (If b c\<^sub>1 c\<^sub>2) f = (\s. if b s then wp c\<^sub>1 f s else wp c\<^sub>2 f s)" | "wp (Par c\<^sub>1 c\<^sub>2) f = wp c\<^sub>1 f \ wp c\<^sub>2 f" | "wp (Prob p c\<^sub>1 c\<^sub>2) f = (\s. pmf p True * wp c\<^sub>1 f s + pmf p False * wp c\<^sub>2 f s)" | "wp (While b c) f = lfp (\X s. if b s then wp c X s else f s)" lemma wp_mono: "mono (wp c)" by (induction c) - (auto simp: mono_def le_fun_def intro: order_trans le_infI1 le_infI2 + (auto simp: monotone_def le_fun_def intro: order_trans le_infI1 le_infI2 intro!: add_mono mult_left_mono lfp_mono[THEN le_funD]) abbreviation det :: "'s pgcl \ 's \ ('s pgcl \ 's) pmf set" ("\ _, _ \") where "det c s \ {return_pmf (c, s)}" subsection \Operational Semantics\ fun step :: "('s pgcl \ 's) \ ('s pgcl \ 's) pmf set" where "step (Skip, s) = \Skip, s\" | "step (Abort, s) = \Abort, s\" | "step (Assign u, s) = \Skip, u s\" | "step (Seq c\<^sub>1 c\<^sub>2, s) = (map_pmf (\(p1', s'). (if p1' = Skip then c\<^sub>2 else Seq p1' c\<^sub>2, s'))) ` step (c\<^sub>1, s)" | "step (If b c\<^sub>1 c\<^sub>2, s) = (if b s then \c\<^sub>1, s\ else \c\<^sub>2, s\)" | "step (Par c\<^sub>1 c\<^sub>2, s) = \c\<^sub>1, s\ \ \c\<^sub>2, s\" | "step (Prob p c\<^sub>1 c\<^sub>2, s) = {map_pmf (\b. if b then (c\<^sub>1, s) else (c\<^sub>2, s)) p}" | "step (While b c, s) = (if b s then \Seq c (While b c), s\ else \Skip, s\)" lemma step_finite: "finite (step x)" by (induction x rule: step.induct) simp_all lemma step_non_empty: "step x \ {}" by (induction x rule: step.induct) simp_all interpretation step: Markov_Decision_Process step proof qed (rule step_non_empty) definition rF :: "('s \ ennreal) \ (('s pgcl \ 's) stream \ ennreal) \ ('s pgcl \ 's) stream \ ennreal" where "rF f F \ = (if fst (shd \) = Skip then f (snd (shd \)) else F (stl \))" abbreviation r :: "('s \ ennreal) \ ('s pgcl \ 's) stream \ ennreal" where "r f \ lfp (rF f)" lemma continuous_rF: "sup_continuous (rF f)" unfolding rF_def[abs_def] by (auto simp: sup_continuous_def fun_eq_iff SUP_sup_distrib [symmetric] image_comp split: prod.splits pgcl.splits) lemma mono_rF: "mono (rF f)" using continuous_rF by (rule sup_continuous_mono) lemma r_unfold: "r f \ = (if fst (shd \) = Skip then f (snd (shd \)) else r f (stl \))" by (subst lfp_unfold[OF mono_rF]) (simp add: rF_def) lemma mono_r: "F \ G \ r F \ \ r G \" by (rule le_funD[of _ _ \], rule lfp_mono) (auto intro!: lfp_mono simp: rF_def le_fun_def max.coboundedI2) lemma measurable_rF: assumes F[measurable]: "F \ borel_measurable step.St" shows "rF f F \ borel_measurable step.St" unfolding rF_def[abs_def] apply measurable apply (rule measurable_compose[OF measurable_shd]) apply measurable [] apply (rule measurable_compose[OF measurable_stl]) apply measurable [] apply (rule predE) apply (rule measurable_compose[OF measurable_shd]) apply measurable done lemma measurable_r[measurable]: "r f \ borel_measurable step.St" using continuous_rF measurable_rF by (rule borel_measurable_lfp) lemma mono_r': "mono (\F s. \D\step s. \\<^sup>+ t. (if fst t = Skip then f (snd t) else F t) \measure_pmf D)" by (auto intro!: monoI le_funI INF_mono[OF bexI] nn_integral_mono simp: le_fun_def) lemma E_inf_r: "step.E_inf s (r f) = lfp (\F s. \D\step s. \\<^sup>+ t. (if fst t = Skip then f (snd t) else F t) \measure_pmf D) s" proof - have "step.E_inf s (r f) = lfp (\F s. \D\step s. \\<^sup>+ t. (if fst t = Skip then f (snd t) else F t) \measure_pmf D) s" unfolding rF_def[abs_def] proof (rule step.E_inf_lfp[THEN fun_cong]) let ?F = "\t x. (if fst t = Skip then f (snd t) else x)" show "(\(s, x). ?F s x) \ borel_measurable (count_space UNIV \\<^sub>M borel)" apply (simp add: measurable_split_conv split_beta') apply (intro borel_measurable_max borel_measurable_const measurable_If predE measurable_compose[OF measurable_snd] measurable_compose[OF measurable_fst]) apply measurable done show "\s. sup_continuous (?F s)" by (auto simp: sup_continuous_def SUP_sup_distrib[symmetric] split: prod.split pgcl.split) show "\F cfg. (\\<^sup>+\. ?F (state cfg) (F \) \step.T cfg) = ?F (state cfg) (nn_integral (step.T cfg) F)" by (auto simp: split: pgcl.split prod.split) qed (rule step_finite) then show ?thesis by simp qed lemma E_inf_r_unfold: "step.E_inf s (r f) = (\D\step s. \\<^sup>+ t. (if fst t = Skip then f (snd t) else step.E_inf t (r f)) \measure_pmf D)" unfolding E_inf_r by (simp add: lfp_unfold[OF mono_r']) lemma E_inf_r_induct[consumes 1, case_names step]: assumes "P s y" assumes *: "\F s y. P s y \ (\s y. P s y \ F s \ y) \ (\s. F s \ step.E_inf s (r f)) \ (\D\step s. \\<^sup>+ t. (if fst t = Skip then f (snd t) else F t) \measure_pmf D) \ y" shows "step.E_inf s (r f) \ y" using \P s y\ unfolding E_inf_r proof (induction arbitrary: s y rule: lfp_ordinal_induct[OF mono_r'[where f=f]]) case (1 F) with *[of s y F] show ?case unfolding le_fun_def E_inf_r[where f=f, symmetric] by simp qed (auto intro: SUP_least) lemma E_inf_Skip: "step.E_inf (Skip, s) (r f) = f s" by (subst E_inf_r_unfold) simp lemma E_inf_Seq: assumes [simp]: "\x. 0 \ f x" shows "step.E_inf (Seq a b, s) (r f) = step.E_inf (a, s) (r (\s. step.E_inf (b, s) (r f)))" proof (rule antisym) show "step.E_inf (Seq a b, s) (r f) \ step.E_inf (a, s) (r (\s. step.E_inf (b, s) (r f)))" proof (coinduction arbitrary: a s rule: E_inf_r_induct) case step then show ?case by (rewrite in "_ \ \" E_inf_r_unfold) (force intro!: INF_mono[OF bexI] nn_integral_mono intro: le_infI2 simp: E_inf_Skip image_comp) qed show "step.E_inf (a, s) (r (\s. step.E_inf (b, s) (r f))) \ step.E_inf (Seq a b, s) (r f)" proof (coinduction arbitrary: a s rule: E_inf_r_induct) case step then show ?case by (rewrite in "_ \ \" E_inf_r_unfold) (force intro!: INF_mono[OF bexI] nn_integral_mono intro: le_infI2 simp: E_inf_Skip image_comp) qed qed lemma E_inf_While: "step.E_inf (While g c, s) (r f) = lfp (\F s. if g s then step.E_inf (c, s) (r F) else f s) s" proof (rule antisym) have E_inf_While_step: "step.E_inf (While g c, s) (r f) = (if g s then step.E_inf (c, s) (r (\s. step.E_inf (While g c, s) (r f))) else f s)" for f s by (rewrite E_inf_r_unfold) (simp add: min_absorb1 E_inf_Seq) have "mono (\F s. if g s then step.E_inf (c, s) (r F) else f s)" (is "mono ?F") by (auto intro!: mono_r step.E_inf_mono simp: mono_def le_fun_def max.coboundedI2) then show "lfp ?F s \ step.E_inf (While g c, s) (r f)" proof (induction arbitrary: s rule: lfp_ordinal_induct[consumes 1]) case mono then show ?case by (rewrite E_inf_While_step) (auto intro!: step.E_inf_mono mono_r le_funI) qed (auto intro: SUP_least) define w where "w F s = (\D\step s. \\<^sup>+ t. (if fst t = Skip then if g (snd t) then F (c, snd t) else f (snd t) else F t) \measure_pmf D)" for F s have "mono w" by (auto simp: w_def mono_def le_fun_def intro!: INF_mono[OF bexI] nn_integral_mono) [] define d where "d = c" define t where "t = Seq d (While g c)" then have "(t = While g c \ d = c \ g s) \ t = Seq d (While g c)" by auto then have "step.E_inf (t, s) (r f) \ lfp w (d, s)" proof (coinduction arbitrary: t d s rule: E_inf_r_induct) case (step F t d s) from step(1) show ?case proof (elim conjE disjE) { fix s have "\ g s \ F (While g c, s) \ f s" using step(3)[of "(While g c, s)"] by (simp add: E_inf_While_step) } note [simp] = this assume "t = Seq d (While g c)" then show ?thesis by (rewrite lfp_unfold[OF \mono w\]) (auto simp: max.absorb2 w_def intro!: INF_mono[OF bexI] nn_integral_mono step) qed (auto intro!: step) qed also have "lfp w = lfp (\F s. step.E_inf s (r (\s. if g s then F (c, s) else f s)))" unfolding E_inf_r w_def by (rule lfp_lfp[symmetric]) (auto simp: le_fun_def intro!: INF_mono[OF bexI] nn_integral_mono) finally have "step.E_inf (While g c, s) (r f) \ (if g s then \ (c, s) else f s)" unfolding t_def d_def by (rewrite E_inf_r_unfold) simp also have "\ = lfp ?F s" by (rewrite lfp_rolling[symmetric, of "\F s. if g s then F (c, s) else f s" "\F s. step.E_inf s (r F)"]) (auto simp: mono_def le_fun_def sup_apply[abs_def] if_distrib[of "max 0"] max.coboundedI2 max.absorb2 intro!: step.E_inf_mono mono_r cong del: if_weak_cong) finally show "step.E_inf (While g c, s) (r f) \ \" . qed subsection \Equate Both Semantics\ lemma E_inf_r_eq_wp: "step.E_inf (c, s) (r f) = wp c f s" proof (induction c arbitrary: f s) case Skip then show ?case by (simp add: E_inf_Skip) next case Abort then show ?case proof (intro antisym) have "lfp (\F s. \D\step s. \\<^sup>+ t. (if fst t = Skip then f (snd t) else F t) \measure_pmf D) \ (\s. if \t. s = (Abort, t) then 0 else \)" by (intro lfp_lowerbound) (auto simp: le_fun_def) then show "step.E_inf (Abort, s) (r f) \ wp Abort f s" by (auto simp: E_inf_r le_fun_def split: if_split_asm) qed simp next case Assign then show ?case by (rewrite E_inf_r_unfold) (simp add: min_absorb1) next case (If b c1 c2) then show ?case by (rewrite E_inf_r_unfold) auto next case (Prob p c1 c2) then show ?case apply (rewrite E_inf_r_unfold) apply auto apply (rewrite nn_integral_measure_pmf_support[of "UNIV::bool set"]) apply (auto simp: UNIV_bool ac_simps) done next case (Par c1 c2) then show ?case by (rewrite E_inf_r_unfold) (auto intro: inf.commute) next case (Seq c1 c2) then show ?case by (simp add: E_inf_Seq) next case (While g c) then show ?case apply (simp add: E_inf_While) apply (rewrite While) apply auto done qed end diff --git a/thys/MonoBoolTranAlgebra/Mono_Bool_Tran_Algebra.thy b/thys/MonoBoolTranAlgebra/Mono_Bool_Tran_Algebra.thy --- a/thys/MonoBoolTranAlgebra/Mono_Bool_Tran_Algebra.thy +++ b/thys/MonoBoolTranAlgebra/Mono_Bool_Tran_Algebra.thy @@ -1,961 +1,962 @@ section \Algebra of Monotonic Boolean Transformers\ theory Mono_Bool_Tran_Algebra imports Mono_Bool_Tran begin text\ In this section we introduce the {\em algebra of monotonic boolean transformers}. This is a bounded distributive lattice with a monoid operation, a dual operator and an iteration operator. The standard model for this algebra is the set of monotonic boolean transformers introduced in the previous section. \ class dual = fixes dual::"'a \ 'a" ("_ ^ o" [81] 80) class omega = fixes omega::"'a \ 'a" ("_ ^ \" [81] 80) class star = fixes star::"'a \ 'a" ("(_ ^ *)" [81] 80) class dual_star = fixes dual_star::"'a \ 'a" ("(_ ^ \)" [81] 80) class mbt_algebra = monoid_mult + dual + omega + distrib_lattice + order_top + order_bot + star + dual_star + assumes dual_le: "(x \ y) = (y ^ o \ x ^ o)" and dual_dual [simp]: "(x ^ o) ^ o = x" and dual_comp: "(x * y) ^ o = x ^ o * y ^ o" and dual_one [simp]: "1 ^ o = 1" and top_comp [simp]: "\ * x = \" and inf_comp: "(x \ y) * z = (x * z) \ (y * z)" and le_comp: "x \ y \ z * x \ z * y" and dual_neg: "(x * \) \ (x ^ o * \) = \" and omega_fix: "x ^ \ = (x * (x ^ \)) \ 1" and omega_least: "(x * z) \ y \ z \ (x ^ \) * y \ z" and star_fix: "x ^ * = (x * (x ^ *)) \ 1" and star_greatest: "z \ (x * z) \ y \ z \ (x ^ *) * y" and dual_star_def: "(x ^ \) = (((x ^ o) ^ *) ^ o)" begin lemma le_comp_right: "x \ y \ x * z \ y * z" apply (cut_tac x = x and y = y and z = z in inf_comp) apply (simp add: inf_absorb1) apply (subgoal_tac "x * z \ (y * z) \ y * z") apply simp by (rule inf_le2) subclass bounded_lattice proof qed end instantiation MonoTran :: (complete_boolean_algebra) mbt_algebra begin lift_definition dual_MonoTran :: "'a MonoTran \ 'a MonoTran" is dual_fun by (fact mono_dual_fun) lift_definition omega_MonoTran :: "'a MonoTran \ 'a MonoTran" is omega_fun by (fact mono_omega_fun) lift_definition star_MonoTran :: "'a MonoTran \ 'a MonoTran" is star_fun by (fact mono_star_fun) definition dual_star_MonoTran :: "'a MonoTran \ 'a MonoTran" where "(x::('a MonoTran)) ^ \ = ((x ^ o) ^ *) ^ o" instance proof fix x y :: "'a MonoTran" show "(x \ y) = (y ^ o \ x ^ o)" apply transfer apply (auto simp add: fun_eq_iff le_fun_def) apply (drule_tac x = "-xa" in spec) apply simp done next fix x :: "'a MonoTran" show "(x ^ o) ^ o = x" apply transfer apply (simp add: fun_eq_iff) done next fix x y :: "'a MonoTran" show "(x * y) ^ o = x ^ o * y ^ o" apply transfer apply (simp add: fun_eq_iff) done next show "(1::'a MonoTran) ^ o = 1" apply transfer apply (simp add: fun_eq_iff) done next fix x :: "'a MonoTran" show "\ * x = \" apply transfer apply (simp add: fun_eq_iff) done next fix x y z :: "'a MonoTran" show "(x \ y) * z = (x * z) \ (y * z)" apply transfer apply (simp add: fun_eq_iff) done next fix x y z :: "'a MonoTran" assume A: "x \ y" from A show " z * x \ z * y" apply transfer apply (auto simp add: le_fun_def elim: monoE) done next fix x :: "'a MonoTran" show "x * \ \ (x ^ o * \) = \" apply transfer apply (simp add: fun_eq_iff) done next fix x :: "'a MonoTran" show "x ^ \ = x * x ^ \ \ 1" apply transfer apply (simp add: fun_eq_iff) apply (simp add: omega_fun_def Omega_fun_def) apply (subst lfp_unfold, simp_all add: ac_simps) apply (auto intro!: mono_comp mono_comp_fun) done next fix x y z :: "'a MonoTran" assume A: "x * z \ y \ z" from A show "x ^ \ * y \ z" apply transfer apply (auto simp add: lfp_omega lfp_def) apply (rule Inf_lower) apply (auto simp add: Omega_fun_def ac_simps) done next fix x :: "'a MonoTran" show "x ^ * = x * x ^ * \ 1" apply transfer apply (auto simp add: star_fun_def Omega_fun_def) apply (subst gfp_unfold, simp_all add: ac_simps) apply (auto intro!: mono_comp mono_comp_fun) done next fix x y z :: "'a MonoTran" assume A: "z \ x * z \ y" from A show "z \ x ^ * * y" apply transfer apply (auto simp add: gfp_star gfp_def) apply (rule Sup_upper) apply (auto simp add: Omega_fun_def) done next fix x :: "'a MonoTran" show "x ^ \ = ((x ^ o) ^ *) ^ o" by (simp add: dual_star_MonoTran_def) qed end context mbt_algebra begin lemma dual_top [simp]: "\ ^ o = \" apply (rule order.antisym, simp_all) by (subst dual_le, simp) lemma dual_bot [simp]: "\ ^ o = \" apply (rule order.antisym, simp_all) by (subst dual_le, simp) lemma dual_inf: "(x \ y) ^ o = (x ^ o) \ (y ^ o)" apply (rule order.antisym, simp_all, safe) apply (subst dual_le, simp, safe) apply (subst dual_le, simp) apply (subst dual_le, simp) apply (subst dual_le, simp) by (subst dual_le, simp) lemma dual_sup: "(x \ y) ^ o = (x ^ o) \ (y ^ o)" apply (rule order.antisym, simp_all, safe) apply (subst dual_le, simp) apply (subst dual_le, simp) apply (subst dual_le, simp, safe) apply (subst dual_le, simp) by (subst dual_le, simp) lemma sup_comp: "(x \ y) * z = (x * z) \ (y * z)" apply (subgoal_tac "((x ^ o \ y ^ o) * z ^ o) ^ o = ((x ^ o * z ^ o) \ (y ^ o * z ^ o)) ^ o") apply (simp add: dual_inf dual_comp) by (simp add: inf_comp) lemma dual_eq: "x ^ o = y ^ o \ x = y" apply (subgoal_tac "(x ^ o) ^ o = (y ^ o) ^ o") apply (subst (asm) dual_dual) apply (subst (asm) dual_dual) by simp_all lemma dual_neg_top [simp]: "(x ^ o * \) \ (x * \) = \" apply (rule dual_eq) by(simp add: dual_sup dual_comp dual_neg) lemma bot_comp [simp]: "\ * x = \" by (rule dual_eq, simp add: dual_comp) lemma [simp]: "(x * \) * y = x * \" by (simp add: mult.assoc) lemma [simp]: "(x * \) * y = x * \" by (simp add: mult.assoc) lemma gt_one_comp: "1 \ x \ y \ x * y" by (cut_tac x = 1 and y = x and z = y in le_comp_right, simp_all) theorem omega_comp_fix: "x ^ \ * y = (x * (x ^ \) * y) \ y" apply (subst omega_fix) by (simp add: inf_comp) theorem dual_star_fix: "x^\ = (x * (x^\)) \ 1" by (metis dual_comp dual_dual dual_inf dual_one dual_star_def star_fix) theorem star_comp_fix: "x ^ * * y = (x * (x ^ *) * y) \ y" apply (subst star_fix) by (simp add: inf_comp) theorem dual_star_comp_fix: "x^\ * y = (x * (x^\) * y) \ y" apply (subst dual_star_fix) by (simp add: sup_comp) theorem dual_star_least: "(x * z) \ y \ z \ (x^\) * y \ z" apply (subst dual_le) apply (simp add: dual_star_def dual_comp) apply (rule star_greatest) apply (subst dual_le) by (simp add: dual_inf dual_comp) lemma omega_one [simp]: "1 ^ \ = \" apply (rule order.antisym, simp_all) by (cut_tac x = "1::'a" and y = 1 and z = \ in omega_least, simp_all) lemma omega_mono: "x \ y \ x ^ \ \ y ^ \" apply (cut_tac x = x and y = 1 and z = "y ^ \" in omega_least, simp_all) apply (subst (2) omega_fix, simp_all) apply (rule_tac y = "x * y ^ \" in order_trans, simp) by (rule le_comp_right, simp) end sublocale mbt_algebra < conjunctive "inf" "inf" "times" done sublocale mbt_algebra < disjunctive "sup" "sup" "times" done context mbt_algebra begin lemma dual_conjunctive: "x \ conjunctive \ x ^ o \ disjunctive" apply (simp add: conjunctive_def disjunctive_def) apply safe apply (rule dual_eq) by (simp add: dual_comp dual_sup) lemma dual_disjunctive: "x \ disjunctive \ x ^ o \ conjunctive" apply (simp add: conjunctive_def disjunctive_def) apply safe apply (rule dual_eq) by (simp add: dual_comp dual_inf) lemma comp_pres_conj: "x \ conjunctive \ y \ conjunctive \ x * y \ conjunctive" apply (subst conjunctive_def, safe) by (simp add: mult.assoc conjunctiveD) lemma comp_pres_disj: "x \ disjunctive \ y \ disjunctive \ x * y \ disjunctive" apply (subst disjunctive_def, safe) by (simp add: mult.assoc disjunctiveD) lemma start_pres_conj: "x \ conjunctive \ (x ^ *) \ conjunctive" apply (subst conjunctive_def, safe) apply (rule order.antisym, simp_all) apply (metis inf_le1 inf_le2 le_comp) apply (rule star_greatest) apply (subst conjunctiveD, simp) apply (subst star_comp_fix) apply (subst star_comp_fix) by (metis inf.assoc inf_left_commute mult.assoc order_refl) lemma dual_star_pres_disj: "x \ disjunctive \ x^\ \ disjunctive" apply (simp add: dual_star_def) apply (rule dual_conjunctive) apply (rule start_pres_conj) by (rule dual_disjunctive, simp) subsection\Assertions\ text\ Usually, in Kleene algebra with tests or in other progrm algebras, tests or assertions or assumptions are defined using an existential quantifier. An element of the algebra is a test if it has a complement with respect to $\bot$ and $1$. In this formalization assertions can be defined much simpler using the dual operator. \ definition "assertion = {x . x \ 1 \ (x * \) \ (x ^ o) = x}" lemma assertion_prop: "x \ assertion \ (x * \) \ 1 = x" apply (simp add: assertion_def) apply safe apply (rule order.antisym) apply simp_all proof - assume [simp]: "x \ 1" assume A: "x * \ \ x ^ o = x" have "x * \ \ 1 \ x * \ \ x ^ o" apply simp apply (rule_tac y = 1 in order_trans) apply simp apply (subst dual_le) by simp also have "\ = x" by (cut_tac A, simp) finally show "x * \ \ 1 \ x" . next assume A: "x * \ \ x ^ o = x" have "x = x * \ \ x ^ o" by (simp add: A) also have "\ \ x * \" by simp finally show "x \ x * \" . qed lemma dual_assertion_prop: "x \ assertion \ ((x ^ o) * \) \ 1 = x ^ o" apply (rule dual_eq) by (simp add: dual_sup dual_comp assertion_prop) lemma assertion_disjunctive: "x \ assertion \ x \ disjunctive" apply (simp add: disjunctive_def, safe) apply (drule assertion_prop) proof - assume A: "x * \ \ 1 = x" fix y z::"'a" have "x * (y \ z) = (x * \ \ 1) * (y \ z)" by (cut_tac A, simp) also have "\ = (x * \) \ (y \ z)" by (simp add: inf_comp) also have "\ = ((x * \) \ y) \ ((x * \) \ z)" by (simp add: inf_sup_distrib) also have "\ = (((x * \) \ 1) * y) \ (((x * \) \ 1) * z)" by (simp add: inf_comp) also have "\ = x * y \ x * z" by (cut_tac A, simp) finally show "x * (y \ z) = x * y \ x * z" . qed lemma Abs_MonoTran_injective: "mono x \ mono y \ Abs_MonoTran x = Abs_MonoTran y \ x = y" apply (subgoal_tac "Rep_MonoTran (Abs_MonoTran x) = Rep_MonoTran (Abs_MonoTran y)") apply (subst (asm) Abs_MonoTran_inverse, simp) by (subst (asm) Abs_MonoTran_inverse, simp_all) end lemma mbta_MonoTran_disjunctive: "Rep_MonoTran ` disjunctive = Apply.disjunctive" apply (simp add: disjunctive_def Apply.disjunctive_def) apply transfer apply auto proof - fix f :: "'a \ 'a" and a b assume prem: "\y. mono y \ (\z. mono z \ f \ y \ z = (f \ y) \ (f \ z))" { fix g h :: "'b \ 'a" assume "mono g" and "mono h" then have "f \ g \ h = (f \ g) \ (f \ h)" using prem by blast } note * = this assume "mono f" show "f (a \ b) = f a \ f b" (is "?P = ?Q") proof (rule order_antisym) show "?P \ ?Q" using * [of "\_. a" "\_. b"] by (simp add: comp_def fun_eq_iff) next - from \mono f\ show "?Q \ ?P" by (rule Lattices.semilattice_sup_class.mono_sup) + from \mono f\ show "?Q \ ?P" + using Fun.semilattice_sup_class.mono_sup by blast qed next fix f :: "'a \ 'a" assume "\y z. f (y \ z) = f y \ f z" then have *: "\y z. f (y \ z) = f y \ f z" by blast show "mono f" proof fix a b :: 'a assume "a \ b" then show "f a \ f b" unfolding sup.order_iff * [symmetric] by simp qed qed lemma assertion_MonoTran: "assertion = Abs_MonoTran ` assertion_fun" apply (safe) apply (subst assertion_fun_disj_less_one) apply (simp add: image_def) apply (rule_tac x = "Rep_MonoTran x" in bexI) apply (simp add: Rep_MonoTran_inverse) apply safe apply (drule assertion_disjunctive) apply (unfold mbta_MonoTran_disjunctive [THEN sym], simp) apply (simp add: assertion_def less_eq_MonoTran_def one_MonoTran_def Abs_MonoTran_inverse) apply (simp add: assertion_def) by (simp_all add: inf_MonoTran_def less_eq_MonoTran_def times_MonoTran_def dual_MonoTran_def top_MonoTran_def Abs_MonoTran_inverse one_MonoTran_def assertion_fun_dual) context mbt_algebra begin lemma assertion_conjunctive: "x \ assertion \ x \ conjunctive" apply (simp add: conjunctive_def, safe) apply (drule assertion_prop) proof - assume A: "x * \ \ 1 = x" fix y z::"'a" have "x * (y \ z) = (x * \ \ 1) * (y \ z)" by (cut_tac A, simp) also have "\ = (x * \) \ (y \ z)" by (simp add: inf_comp) also have "\ = ((x * \) \ y) \ ((x * \) \ z)" apply (rule order.antisym, simp_all, safe) apply (rule_tac y = "y \ z" in order_trans) apply (rule inf_le2) apply simp apply (rule_tac y = "y \ z" in order_trans) apply (rule inf_le2) apply simp_all apply (simp add: inf_assoc) apply (rule_tac y = " x * \ \ y" in order_trans) apply (rule inf_le1) apply simp apply (rule_tac y = " x * \ \ z" in order_trans) apply (rule inf_le2) by simp also have "\ = (((x * \) \ 1) * y) \ (((x * \) \ 1) * z)" by (simp add: inf_comp) also have "\ = (x * y) \ (x * z)" by (cut_tac A, simp) finally show "x * (y \ z) = (x * y) \ (x * z)" . qed lemma dual_assertion_conjunctive: "x \ assertion \ x ^ o \ conjunctive" apply (drule assertion_disjunctive) by (rule dual_disjunctive, simp) lemma dual_assertion_disjunct: "x \ assertion \ x ^ o \ disjunctive" apply (drule assertion_conjunctive) by (rule dual_conjunctive, simp) lemma [simp]: "x \ assertion \ y \ assertion \ x \ y \ x * y" apply (simp add: assertion_def, safe) proof - assume A: "x \ 1" assume B: "x * \ \ x ^ o = x" assume C: "y \ 1" assume D: "y * \ \ y ^ o = y" have "x \ y = (x * \ \ x ^ o) \ (y * \ \ y ^ o)" by (cut_tac B D, simp) also have "\ \ (x * \) \ (((x^o) * (y * \)) \ ((x^o) * (y^o)))" apply (simp, safe) apply (rule_tac y = "x * \ \ x ^ o" in order_trans) apply (rule inf_le1) apply simp apply (rule_tac y = "y * \" in order_trans) apply (rule_tac y = "y * \ \ y ^ o" in order_trans) apply (rule inf_le2) apply simp apply (rule gt_one_comp) apply (subst dual_le, simp add: A) apply (rule_tac y = "y ^ o" in order_trans) apply (rule_tac y = "y * \ \ y ^ o" in order_trans) apply (rule inf_le2) apply simp apply (rule gt_one_comp) by (subst dual_le, simp add: A) also have "... = ((x * \) \ (x ^ o)) * ((y * \) \ (y ^ o))" apply (cut_tac x = x in dual_assertion_conjunctive) apply (cut_tac A, cut_tac B, simp add: assertion_def) by (simp add: inf_comp conjunctiveD) also have "... = x * y" by (cut_tac B, cut_tac D, simp) finally show "x \ y \ x * y" . qed lemma [simp]: "x \ assertion \ x * y \ y" by (unfold assertion_def, cut_tac x = x and y = 1 and z = y in le_comp_right, simp_all) lemma [simp]: "x \ assertion \ y \ assertion \ x * y \ x" apply (subgoal_tac "x * y \ (x * \) \ (x ^ o)") apply (simp add: assertion_def) apply (simp, safe) apply (rule le_comp, simp) apply (rule_tac y = 1 in order_trans) apply (rule_tac y = y in order_trans) apply simp apply (simp add: assertion_def) by (subst dual_le, simp add: assertion_def) lemma assertion_inf_comp_eq: "x \ assertion \ y \ assertion \ x \ y = x * y" by (rule order.antisym, simp_all) lemma one_right_assertion [simp]: "x \ assertion \ x * 1 = x" apply (drule assertion_prop) proof - assume A: "x * \ \ 1 = x" have "x * 1 = (x * \ \ 1) * 1" by (simp add: A) also have "\ = x * \ \ 1" by (simp add: inf_comp) also have "\ = x" by (simp add: A) finally show ?thesis . qed lemma [simp]: "x \ assertion \ x \ 1 = 1" by (rule order.antisym, simp_all add: assertion_def) lemma [simp]: "x \ assertion \ 1 \ x = 1" by (rule order.antisym, simp_all add: assertion_def) lemma [simp]: "x \ assertion \ x \ 1 = x" by (rule order.antisym, simp_all add: assertion_def) lemma [simp]: "x \ assertion \ 1 \ x = x" by (rule order.antisym, simp_all add: assertion_def) lemma [simp]: "x \ assertion \ x \ x * \" by (cut_tac x = 1 and y = \ and z = x in le_comp, simp_all) lemma [simp]: "x \ assertion \ x \ 1" by (simp add: assertion_def) definition "neg_assert (x::'a) = (x ^ o * \) \ 1" lemma sup_uminus[simp]: "x \ assertion \ x \ neg_assert x = 1" apply (simp add: neg_assert_def) apply (simp add: sup_inf_distrib) apply (rule order.antisym, simp_all) apply (unfold assertion_def) apply safe apply (subst dual_le) apply (simp add: dual_sup dual_comp) apply (subst inf_commute) by simp lemma inf_uminus[simp]: "x \ assertion \ x \ neg_assert x = \" apply (simp add: neg_assert_def) apply (rule order.antisym, simp_all) apply (rule_tac y = "x \ (x ^ o * \)" in order_trans) apply simp apply (rule_tac y = "x ^ o * \ \ 1" in order_trans) apply (rule inf_le2) apply simp apply (rule_tac y = "(x * \) \ (x ^ o * \)" in order_trans) apply simp apply (rule_tac y = x in order_trans) apply simp_all by (simp add: dual_neg) lemma uminus_assertion[simp]: "x \ assertion \ neg_assert x \ assertion" apply (subst assertion_def) apply (simp add: neg_assert_def) apply (simp add: inf_comp dual_inf dual_comp inf_sup_distrib) apply (subst inf_commute) by (simp add: dual_neg) lemma uminus_uminus [simp]: "x \ assertion \ neg_assert (neg_assert x) = x" apply (simp add: neg_assert_def) by (simp add: dual_inf dual_comp sup_comp assertion_prop) lemma dual_comp_neg [simp]: "x ^ o * y \ (neg_assert x) * \ = x ^ o * y" apply (simp add: neg_assert_def inf_comp) apply (rule order.antisym, simp_all) by (rule le_comp, simp) lemma [simp]: "(neg_assert x) ^ o * y \ x * \ = (neg_assert x) ^ o * y" apply (simp add: neg_assert_def inf_comp dual_inf dual_comp sup_comp) by (rule order.antisym, simp_all) lemma [simp]: " x * \ \ (neg_assert x) ^ o * y= (neg_assert x) ^ o * y" by (simp add: neg_assert_def inf_comp dual_inf dual_comp sup_comp) lemma inf_assertion [simp]: "x \ assertion \ y \ assertion \ x \ y \ assertion" apply (subst assertion_def) apply safe apply (rule_tac y = x in order_trans) apply simp_all apply (simp add: assertion_inf_comp_eq) proof - assume A: "x \ assertion" assume B: "y \ assertion" have C: "(x * \) \ (x ^ o) = x" by (cut_tac A, unfold assertion_def, simp) have D: "(y * \) \ (y ^ o) = y" by (cut_tac B, unfold assertion_def, simp) have "x * y = ((x * \) \ (x ^ o)) * ((y * \) \ (y ^ o))" by (simp add: C D) also have "\ = x * \ \ ((x ^ o) * ((y * \) \ (y ^ o)))" by (simp add: inf_comp) also have "\ = x * \ \ ((x ^ o) * (y * \)) \ ((x ^ o) *(y ^ o))" by (cut_tac A, cut_tac x = x in dual_assertion_conjunctive, simp_all add: conjunctiveD inf_assoc) also have "\ = (((x * \) \ (x ^ o)) * (y * \)) \ ((x ^ o) *(y ^ o))" by (simp add: inf_comp) also have "\ = (x * y * \) \ ((x * y) ^ o)" by (simp add: C mult.assoc dual_comp) finally show "(x * y * \) \ ((x * y) ^ o) = x * y" by simp qed lemma comp_assertion [simp]: "x \ assertion \ y \ assertion \ x * y \ assertion" by (subst assertion_inf_comp_eq [THEN sym], simp_all) lemma sup_assertion [simp]: "x \ assertion \ y \ assertion \ x \ y \ assertion" apply (subst assertion_def) apply safe apply (unfold assertion_def) apply simp apply safe proof - assume [simp]: "x \ 1" assume [simp]: "y \ 1" assume A: "x * \ \ x ^ o = x" assume B: "y * \ \ y ^ o = y" have "(y * \) \ (x ^ o) \ (y ^ o) = (x ^ o) \ (y * \) \ (y ^ o)" by (simp add: inf_commute) also have "\ = (x ^ o) \ ((y * \) \ (y ^ o))" by (simp add: inf_assoc) also have "\ = (x ^ o) \ y" by (simp add: B) also have "\ = y" apply (rule order.antisym, simp_all) apply (rule_tac y = 1 in order_trans) apply simp by (subst dual_le, simp) finally have [simp]: "(y * \) \ (x ^ o) \ (y ^ o) = y" . have "x * \ \ (x ^ o) \ (y ^ o) = x \ (y ^ o)" by (simp add: A) also have "\ = x" apply (rule order.antisym, simp_all) apply (rule_tac y = 1 in order_trans) apply simp by (subst dual_le, simp) finally have [simp]: "x * \ \ (x ^ o) \ (y ^ o) = x" . have "(x \ y) * \ \ (x \ y) ^ o = (x * \ \ y * \) \ ((x ^ o) \ (y ^ o))" by (simp add: sup_comp dual_sup) also have "\ = x \ y" by (simp add: inf_sup_distrib inf_assoc [THEN sym]) finally show "(x \ y) * \ \ (x \ y) ^ o = x \ y" . qed lemma [simp]: "x \ assertion \ x * x = x" by (simp add: assertion_inf_comp_eq [THEN sym]) lemma [simp]: "x \ assertion \ (x ^ o) * (x ^ o) = x ^ o" apply (rule dual_eq) by (simp add: dual_comp assertion_inf_comp_eq [THEN sym]) lemma [simp]: "x \ assertion \ x * (x ^ o) = x" proof - assume A: "x \ assertion" have B: "x * \ \ (x ^ o) = x" by (cut_tac A, unfold assertion_def, simp) have "x * x ^ o = (x * \ \ (x ^ o)) * x ^ o" by (simp add: B) also have "\ = x * \ \ (x ^ o)" by (cut_tac A, simp add: inf_comp) also have "\ = x" by (simp add: B) finally show ?thesis . qed lemma [simp]: "x \ assertion \ (x ^ o) * x = x ^ o" apply (rule dual_eq) by (simp add: dual_comp) lemma [simp]: "\ \ assertion" by (unfold assertion_def, simp) lemma [simp]: "1 \ assertion" by (unfold assertion_def, simp) subsection \Weakest precondition of true\ definition "wpt x = (x * \) \ 1" lemma wpt_is_assertion [simp]: "wpt x \ assertion" apply (unfold wpt_def assertion_def, safe) apply simp apply (simp add: inf_comp dual_inf dual_comp inf_sup_distrib) apply (rule order.antisym) by (simp_all add: dual_neg) lemma wpt_comp: "(wpt x) * x = x" apply (simp add: wpt_def inf_comp) apply (rule order.antisym, simp_all) by (cut_tac x = 1 and y = \ and z = x in le_comp, simp_all) lemma wpt_comp_2: "wpt (x * y) = wpt (x * (wpt y))" by (simp add: wpt_def inf_comp mult.assoc) lemma wpt_assertion [simp]: "x \ assertion \ wpt x = x" by (simp add: wpt_def assertion_prop) lemma wpt_le_assertion: "x \ assertion \ x * y = y \ wpt y \ x" apply (simp add: wpt_def) proof - assume A: "x \ assertion" assume B: "x * y = y" have "y * \ \ 1 = x * (y * \) \ 1" by (simp add: B mult.assoc [THEN sym]) also have "\ \ x * \ \ 1" apply simp apply (rule_tac y = "x * (y * \)" in order_trans) apply simp_all by (rule le_comp, simp) also have "\ = x" by (cut_tac A, simp add: assertion_prop) finally show "y * \ \ 1 \ x" . qed lemma wpt_choice: "wpt (x \ y) = wpt x \ wpt y" apply (simp add: wpt_def inf_comp) proof - have "x * \ \ 1 \ (y * \ \ 1) = x * \ \ ((y * \ \ 1) \ 1)" apply (subst inf_assoc) by (simp add: inf_commute) also have "... = x * \ \ (y * \ \ 1)" by (subst inf_assoc, simp) also have "... = (x * \) \ (y * \) \ 1" by (subst inf_assoc, simp) finally show "x * \ \ (y * \) \ 1 = x * \ \ 1 \ (y * \ \ 1)" by simp qed end context lattice begin lemma [simp]: "x \ y \ x \ y = x" by (simp add: inf_absorb1) end context mbt_algebra begin lemma wpt_dual_assertion_comp: "x \ assertion \ y \ assertion \ wpt ((x ^ o) * y) = (neg_assert x) \ y" apply (simp add: wpt_def neg_assert_def) proof - assume A: "x \ assertion" assume B: "y \ assertion" have C: "((x ^ o) * \) \ 1 = x ^ o" by (rule dual_assertion_prop, rule A) have "x ^ o * y * \ \ 1 = (((x ^ o) * \) \ 1) * y * \ \ 1" by (simp add: C) also have "\ = ((x ^ o) * \ \ (y * \)) \ 1" by (simp add: sup_comp) also have "\ = (((x ^ o) * \) \ 1) \ ((y * \) \ 1)" by (simp add: inf_sup_distrib2) also have "\ = (((x ^ o) * \) \ 1) \ y" by (cut_tac B, drule assertion_prop, simp) finally show "x ^ o * y * \ \ 1 = (((x ^ o) * \) \ 1) \ y" . qed lemma le_comp_left_right: "x \ y \ u \ v \ x * u \ y * v" apply (rule_tac y = "x * v" in order_trans) apply (rule le_comp, simp) by (rule le_comp_right, simp) lemma wpt_dual_assertion: "x \ assertion \ wpt (x ^ o) = 1" apply (simp add: wpt_def) apply (rule order.antisym) apply simp_all apply (cut_tac x = 1 and y = "x ^ o" and u = 1 and v = \ in le_comp_left_right) apply simp_all apply (subst dual_le) by simp lemma assertion_commute: "x \ assertion \ y \ conjunctive \ y * x = wpt(y * x) * y" apply (simp add: wpt_def) apply (simp add: inf_comp) apply (drule_tac x = y and y = "x * \" and z = 1 in conjunctiveD) by (simp add: mult.assoc [THEN sym] assertion_prop) lemma wpt_mono: "x \ y \ wpt x \ wpt y" apply (simp add: wpt_def) apply (rule_tac y = "x * \" in order_trans, simp_all) by (rule le_comp_right, simp) lemma "a \ conjunctive \ x * a \ a * y \ (x ^ \) * a \ a * (y ^ \)" apply (rule omega_least) apply (simp add: mult.assoc [THEN sym]) apply (rule_tac y = "a * y * y ^ \ \ a" in order_trans) apply (simp) apply (rule_tac y = "x * a * y ^ \" in order_trans, simp_all) apply (rule le_comp_right, simp) apply (simp add: mult.assoc) apply (subst (2) omega_fix) by (simp add: conjunctiveD) lemma [simp]: "x \ 1 \ y * x \ y" by (cut_tac x = x and y = 1 and z = y in le_comp, simp_all) lemma [simp]: "x \ x * \" by (cut_tac x = 1 and y = \ and z = x in le_comp, simp_all) lemma [simp]: "x * \ \ x" by (cut_tac x = \ and y = 1 and z = x in le_comp, simp_all) end subsection\Monotonic Boolean trasformers algebra with post condition statement\ definition "post_fun (p::'a::order) q = (if p \ q then (\::'b::{order_bot,order_top}) else \)" lemma mono_post_fun [simp]: "mono (post_fun (p::_::{order_bot,order_top}))" apply (simp add: post_fun_def mono_def, safe) apply (subgoal_tac "p \ y", simp) apply (rule_tac y = x in order_trans) apply simp_all done lemma post_top [simp]: "post_fun p p = \" by (simp add: post_fun_def) lemma post_refin [simp]: "mono S \ ((S p)::'a::bounded_lattice) \ (post_fun p) x \ S x" apply (simp add: le_fun_def assert_fun_def post_fun_def, safe) by (rule_tac f = S in monoD, simp_all) class post_mbt_algebra = mbt_algebra + fixes post :: "'a \ 'a" assumes post_1: "(post x) * x * \ = \" and post_2: "y * x * \ \ (post x) \ y" instantiation MonoTran :: (complete_boolean_algebra) post_mbt_algebra begin lift_definition post_MonoTran :: "'a::complete_boolean_algebra MonoTran \ 'a::complete_boolean_algebra MonoTran" is "\x. post_fun (x \)" by (rule mono_post_fun) instance proof fix x :: "'a MonoTran" show "post x * x * \ = \" apply transfer apply (simp add: fun_eq_iff) done fix x y :: "'a MonoTran" show "y * x * \ \ post x \ y" apply transfer apply (simp add: le_fun_def) done qed end subsection\Complete monotonic Boolean transformers algebra\ class complete_mbt_algebra = post_mbt_algebra + complete_distrib_lattice + assumes Inf_comp: "(Inf X) * z = (INF x \ X . (x * z))" instance MonoTran :: (complete_boolean_algebra) complete_mbt_algebra apply intro_classes apply transfer apply (simp add: Inf_comp_fun) done context complete_mbt_algebra begin lemma dual_Inf: "(Inf X) ^ o = (SUP x\ X . x ^ o)" apply (rule order.antisym) apply (subst dual_le, simp) apply (rule Inf_greatest) apply (subst dual_le, simp) apply (rule SUP_upper, simp) apply (rule SUP_least) apply (subst dual_le, simp) by (rule Inf_lower, simp) lemma dual_Sup: "(Sup X) ^ o = (INF x\ X . x ^ o)" apply (rule order.antisym) apply (rule INF_greatest) apply (subst dual_le, simp) apply (rule Sup_upper, simp) apply (subst dual_le, simp) apply (rule Sup_least) apply (subst dual_le, simp) by (rule INF_lower, simp) lemma INF_comp: "(\(f ` A)) * z = (INF a \ A . (f a) * z)" unfolding Inf_comp apply (subgoal_tac "((\x::'a. x * z) ` f ` A) = ((\a::'b. f a * z) ` A)") by auto lemma dual_INF: "(\(f ` A)) ^ o = (SUP a \ A . (f a) ^ o)" unfolding Inf_comp dual_Inf apply (subgoal_tac "(dual ` f ` A) = ((\a::'b. f a ^ o) ` A)") by auto lemma dual_SUP: "(\(f ` A)) ^ o = (INF a \ A . (f a) ^ o)" unfolding dual_Sup apply (subgoal_tac "(dual ` f ` A) = ((\a::'b. f a ^ o) ` A)") by auto lemma Sup_comp: "(Sup X) * z = (SUP x \ X . (x * z))" apply (rule dual_eq) by (simp add: dual_comp dual_Sup dual_SUP INF_comp image_comp) lemma SUP_comp: "(\(f ` A)) * z = (SUP a \ A . (f a) * z)" unfolding Sup_comp apply (subgoal_tac "((\x::'a. x * z) ` f ` A) = ((\a::'b. f a * z) ` A)") by auto lemma Sup_assertion [simp]: "X \ assertion \ Sup X \ assertion" apply (unfold assertion_def) apply safe apply (rule Sup_least) apply blast apply (simp add: Sup_comp dual_Sup Sup_inf) apply (subgoal_tac "((\y . y \ \(dual ` X)) ` (\x . x * \) ` X) = X") apply simp proof - assume A: "X \ {x. x \ 1 \ x * \ \ x ^ o = x}" have B [simp]: "!! x . x \ X \ x * \ \ (\(dual ` X)) = x" proof - fix x assume C: "x \ X" have "x * \ \ \(dual ` X) = x * \ \ (x ^ o \ \(dual ` X))" apply (subgoal_tac "\(dual ` X) = (x ^ o \ \(dual ` X))", simp) apply (rule order.antisym, simp_all) apply (rule Inf_lower, cut_tac C, simp) done also have "\ = x \ \(dual ` X)" by (unfold inf_assoc [THEN sym], cut_tac A, cut_tac C, auto) also have "\ = x" apply (rule order.antisym, simp_all) apply (rule INF_greatest) apply (cut_tac A C) apply (rule_tac y = 1 in order_trans) apply auto[1] apply (subst dual_le, auto) done finally show "x * \ \ \(dual ` X) = x" . qed show "(\y. y \ \(dual ` X)) ` (\x . x * \) ` X = X" by (simp add: image_comp) qed lemma Sup_range_assertion [simp]: "(!!w . p w \ assertion) \ Sup (range p) \ assertion" by (rule Sup_assertion, auto) lemma Sup_less_assertion [simp]: "(!!w . p w \ assertion) \ Sup_less p w \ assertion" by (unfold Sup_less_def, rule Sup_assertion, auto) theorem omega_lfp: "x ^ \ * y = lfp (\ z . (x * z) \ y)" apply (rule order.antisym) apply (rule lfp_greatest) apply (drule omega_least, simp) apply (rule lfp_lowerbound) apply (subst (2) omega_fix) by (simp add: inf_comp mult.assoc) end lemma [simp]: "mono (\ (t::'a::mbt_algebra) . x * t \ y)" apply (simp add: mono_def, safe) apply (rule_tac y = "x * xa" in order_trans, simp) by (rule le_comp, simp) class mbt_algebra_fusion = mbt_algebra + assumes fusion: "(\ t . x * t \ y \ z \ u * (t \ z) \ v) \ (x ^ \) * y \ z \ (u ^ \) * v " lemma "class.mbt_algebra_fusion (1::'a::complete_mbt_algebra) ((*)) (\) (\) (<) (\) dual dual_star omega star \ \" apply unfold_locales apply (cut_tac h = "\ t . t \ z" and f = "\ t . x * t \ y" and g = "\ t . u * t \ v" in weak_fusion) apply (rule inf_Disj) apply simp_all apply (simp add: le_fun_def) by (simp add: omega_lfp) context mbt_algebra_fusion begin lemma omega_star: "x \ conjunctive \ x ^ \ = wpt (x ^ \) * (x ^ *)" apply (simp add: wpt_def inf_comp) apply (rule order.antisym) apply (cut_tac x = x and y = 1 and z = "x ^ \ * \ \ x ^ *" in omega_least) apply (simp_all add: conjunctiveD,safe) apply (subst (2) omega_fix) apply (simp add: inf_comp inf_assoc mult.assoc) apply (metis inf.commute inf_assoc inf_le1 star_fix) apply (cut_tac x = x and y = \ and z = "x ^ *" and u = x and v = 1 in fusion) apply (simp add: conjunctiveD) apply (metis inf_commute inf_le1 le_infE star_fix) by (metis mult.right_neutral) lemma omega_pres_conj: "x \ conjunctive \ x ^ \ \ conjunctive" apply (subst omega_star, simp) apply (rule comp_pres_conj) apply (rule assertion_conjunctive, simp) by (rule start_pres_conj, simp) end end diff --git a/thys/Refine_Monadic/Refine_Mono_Prover.thy b/thys/Refine_Monadic/Refine_Mono_Prover.thy --- a/thys/Refine_Monadic/Refine_Mono_Prover.thy +++ b/thys/Refine_Monadic/Refine_Mono_Prover.thy @@ -1,43 +1,43 @@ theory Refine_Mono_Prover imports Main Automatic_Refinement.Refine_Lib begin ML_file \refine_mono_prover.ML\ setup Refine_Mono_Prover.setup declaration Refine_Mono_Prover.decl_setup method_setup refine_mono = \Scan.succeed (fn ctxt => SIMPLE_METHOD' ( Refine_Mono_Prover.untriggered_mono_tac ctxt ))\ "Refinement framework: Monotonicity prover" locale mono_setup_loc = \ \Locale to set up monotonicity prover for given ordering operator\ fixes le :: "'a \ 'a \ bool" assumes refl: "le x x" begin lemma monoI: "(\f g x. (\x. le (f x) (g x)) \ le (B f x) (B g x)) \ monotone (fun_ord le) (fun_ord le) B" unfolding monotone_def fun_ord_def by blast lemma mono_if: "\le t t'; le e e'\ \ le (If b t e) (If b t' e')" by auto lemma mono_let: "(\x. le (f x) (f' x)) \ le (Let x f) (Let x f')" by auto lemmas mono_thms[refine_mono] = monoI mono_if mono_let refl declaration \Refine_Mono_Prover.declare_mono_triggers @{thms monoI}\ end interpretation order_mono_setup: mono_setup_loc "(\) :: 'a::preorder \ _" by standard auto - declaration \Refine_Mono_Prover.declare_mono_triggers @{thms Orderings.monoI}\ + declaration \Refine_Mono_Prover.declare_mono_triggers @{thms monoI}\ lemmas [refine_mono] = lfp_mono[OF le_funI, THEN le_funD] gfp_mono[OF le_funI, THEN le_funD] end diff --git a/thys/Transformer_Semantics/Isotone_Transformers.thy b/thys/Transformer_Semantics/Isotone_Transformers.thy --- a/thys/Transformer_Semantics/Isotone_Transformers.thy +++ b/thys/Transformer_Semantics/Isotone_Transformers.thy @@ -1,344 +1,344 @@ (* Title: Isotone Transformers Between Complete Lattices Author: Georg Struth Maintainer: Georg Struth *) section \Isotone Transformers Between Complete Lattices\ theory Isotone_Transformers imports "Order_Lattice_Props.Fixpoint_Fusion" "Quantales.Quantale_Star" begin text \A transformer is a function between lattices; an isotone transformer preserves the order (or is monotone). In this component, statements are developed in a type-driven way. Statements are developed in more general contexts or even the most general one.\ subsection \Basic Properties\ text \First I show that some basic transformers are isotone...\ lemma iso_id: "mono id" by (simp add: monoI) lemma iso_botf: "mono \" by (simp add: monoI) lemma iso_topf: "mono \" by (simp add: monoI) text \... and that compositions, Infs and Sups preserve isotonicity.\ lemma iso_fcomp: "mono f \ mono g \ mono (f \ g)" by (simp add: mono_def) lemma iso_fSup: fixes F :: "('a::order \ 'b::complete_lattice) set" shows "(\f \ F. mono f) \ mono (\F)" by (simp add: mono_def SUP_subset_mono) lemma iso_fsup: "mono f \ mono g \ mono (f \ g)" unfolding mono_def using sup_mono by fastforce lemma iso_fInf: fixes F :: "('a::order \ 'b::complete_lattice) set" shows "\f \ F. mono f \ mono (\F)" by (simp add: mono_def, safe, rule Inf_greatest, auto simp: INF_lower2) lemma iso_finf: "mono f \ mono g \ mono (f \ g)" unfolding mono_def using inf_mono by fastforce lemma fun_isol: "mono f \ g \ h \ (f \ g) \ (f \ h)" by (simp add: le_fun_def monoD) lemma fun_isor: "mono f \ g \ h \ (g \ f) \ (h \ f)" by (simp add: le_fun_def monoD) subsection \Pre-Quantale of Isotone Transformers\ text \It is well known, and has been formalised within Isabelle, that functions into complete lattices form complete lattices. In the following proof, this needs to be replayed because isotone functions are considered and closure conditions need to be respected. Functions must now be restricted to a single type.\ instantiation iso :: (complete_lattice) unital_pre_quantale begin lift_definition one_iso :: "'a::complete_lattice iso" is id by (simp add: iso_id) lift_definition times_iso :: "'a::complete_lattice iso \ 'a iso \ 'a iso" is "(\)" by (simp add: iso_fcomp) instance by (intro_classes; transfer, simp_all add: comp_assoc fInf_distr_var fInf_subdistl_var) end text \I have previously worked in (pre)quantales with many types or quantaloids. Formally, these are categories enriched over the category of Sup-lattices (complete lattices with Sup-preserving functions). An advantage of the single-typed approach is that the definition of the Kleene star for (pre)quantales is available in this setting.\ subsection \Propositional Hoare Logic for Transformers without Star\ text \The rules of an abstract Propositional Hoare logic are derivable.\ lemma H_iso_cond1: "(x::'a::preorder) \ y \ y \ f z \ x \ f z" using order_trans by auto lemma H_iso_cond2: "mono f \ y \ z \ x \ f y \ x \ f z" by (meson mono_def order_subst1) lemma H_iso_seq: "mono f \ x \ f y \ y \ g z \ x \ f (g z)" using H_iso_cond2 by force lemma H_iso_seq_var: "mono f \ x \ f y \ y \ g z \ x \ (f \ g) z" by (simp add: H_iso_cond2) lemma H_iso_fInf: fixes F :: "('a \ 'b::complete_lattice) set" shows "(\f \ F. x \ f y) \ x \ (\F) y" by (simp add: le_INF_iff) lemma H_iso_fSup: fixes F :: "('a \ 'b::complete_lattice) set" shows "F \ {} \ (\f \ F. x \ f y) \ x \ (\F) y" using SUP_upper2 by fastforce text \These rules are suitable for weakest liberal preconditions. Order-dual ones, in which the order relation is swapped, are consistent with other kinds of transformers. In the context of dynamic logic, the first set corresponds to box modalities whereas the second one would correspond to diamonds.\ subsection \Kleene Star of Isotone Transformers\ text \The Hoare rule for loops requires some preparation. On the way I verify some Kleene-algebra-style axioms for iteration.\ text \First I show that functions form monoids.\ interpretation fun_mon: monoid_mult "id::'a \ 'a" "(\)" by unfold_locales auto definition fiter_fun :: "('a \ 'c::semilattice_inf) \ ('b \ 'c) \ ('a \ 'b) \ 'a \ 'c" where "fiter_fun f g = (\) f \ (\) g" definition fiter :: "('a \ 'b::complete_lattice) \ ('b \ 'b) \ 'a \ 'b" where "fiter f g = gfp (fiter_fun f g)" definition fiter_id :: "('a::complete_lattice \ 'a) \ 'a \ 'a" where "fiter_id = fiter id" abbreviation "fpower \ fun_mon.power" definition fstar :: "('a::complete_lattice \ 'a) \ 'a \ 'a" where "fstar f = (\i. fpower f i)" text \The types in the following statements are often more general than those in the prequantale setting. I develop them generally, instead of inheriting (most of them) with more restrictive types from the quantale components.\ lemma fiter_fun_exp: "fiter_fun f g h = f \ (g \ h)" unfolding fiter_fun_def by simp text \The two lemmas that follow set up the relationship between the star for transformers and those in quantales.\ lemma fiter_qiter1: "Abs_iso (fiter_fun (Rep_iso f) (Rep_iso g) (Rep_iso h)) = qiter_fun f g h" unfolding fiter_fun_def qiter_fun_def by (metis Rep_iso_inverse comp_def sup_iso.rep_eq times_iso.rep_eq) lemma fiter_qiter4: "mono f \ mono g \ mono h \ Rep_iso (qiter_fun (Abs_iso f) (Abs_iso g) (Abs_iso h)) = fiter_fun f g h" - by (metis Abs_iso_inverse fiter_fun_exp fiter_qiter1 iso_fcomp iso_finf mem_Collect_eq) + by (simp add: Abs_iso_inverse fiter_fun_exp qiter_fun_exp sup_iso.rep_eq times_iso.rep_eq) text \The type coercions are needed to deal with isotone (monotone) functions, which had to be redefined to one single type above, in order to cooperate with the type classes for quantales. Having to deal with these coercions would be another drawback of using the quantale-based setting for the development.\ lemma iso_fiter_fun: "mono f \ mono (fiter_fun f)" by (simp add: fiter_fun_exp le_fun_def mono_def inf.coboundedI2) lemma iso_fiter_fun2: "mono f \ mono g \ mono (fiter_fun f g)" by (simp add: fiter_fun_exp le_fun_def mono_def inf.coboundedI2) lemma fiter_unfoldl: fixes f :: "'a::complete_lattice \ 'a" shows "mono f \ mono g \ f \ (g \ fiter f g) = fiter f g" by (metis fiter_def fiter_fun_exp gfp_unfold iso_fiter_fun2) lemma fiter_inductl: fixes f :: "'a::complete_lattice \ 'a" shows "mono f \ mono g \ h \ f \ (g \ h) \ h \ fiter f g" by (simp add: fiter_def fiter_fun_def gfp_upperbound) lemma fiter_fusion: fixes f :: "'a::complete_lattice \ 'a" assumes "mono f" and "mono g" shows "fiter f g = fiter_id g \ f" proof- have h1: "mono (fiter_fun id g)" by (simp add: assms(2) iso_fiter_fun2 iso_id) have h2: "mono (fiter_fun f g)" by (simp add: assms(1) assms(2) iso_fiter_fun2) have h3: "Inf \ image (\x. x \ f) = (\x. x \ f) \ Inf" by (simp add: fun_eq_iff image_comp) have "(\x. x \ f) \ (fiter_fun id g) = (fiter_fun f g) \ (\x. x \ f)" by (simp add: fun_eq_iff fiter_fun_def) thus ?thesis using gfp_fusion_inf_pres by (metis fiter_def fiter_id_def h1 h2 h3) qed lemma fpower_supdistl: fixes f :: "'a::complete_lattice \ 'b::complete_lattice" shows "mono f \ f \ fstar g \ (\i. f \ fpower g i)" by (simp add: Isotone_Transformers.fun_isol fstar_def mono_INF mono_def) lemma fpower_distr: "fstar f \ g = (\i. fpower f i \ g)" by (auto simp: fstar_def image_comp) lemma fpower_Sup_subcomm: "mono f \ f \ fstar f \ fstar f \ f" - by (metis (mono_tags, lifting) fun_mon.power_commutes le_INF_iff fpower_distr fpower_supdistl) + unfolding fpower_distr fun_mon.power_commutes by (rule fpower_supdistl) lemma fpower_inductl: fixes f :: "'a::complete_lattice \ 'a" shows "mono f \ mono g \ h \ g \ (f \ h) \ h \ fpower f i \ g" apply (induct i, simp_all) by (metis (no_types, opaque_lifting) fun.map_comp fun_isol order_trans) lemma fpower_inductr: fixes f :: "'a::complete_lattice \ 'a" shows "mono f \ mono g \ h \ g \ (h \ f) \ h \ g \ fpower f i" by (induct i, simp_all add: le_fun_def, metis comp_eq_elim fun_mon.power_commutes order_trans) lemma fiter_fstar: "mono f \ fiter_id f \ fstar f" by (metis (no_types, lifting) fiter_id_def fiter_unfoldl fpower_inductl fstar_def iso_id le_INF_iff o_id order_refl) lemma iso_fiter_ext: fixes f :: "'a::order \ 'b::complete_lattice" shows "mono f \ mono (\x. y \ f x)" by (simp add: le_infI2 mono_def) lemma fstar_pred_char: fixes f :: "'a::complete_lattice \ 'a" shows "mono f \ fiter_id f x = gfp (\y. x \ f y)" proof - assume hyp: "mono f" have "\g. (id \ (f \ g)) x = x \ f (g x)" by simp hence "\g. fiter_fun id f g x = (\y. x \ f y) (g x)" unfolding fiter_fun_def by simp thus ?thesis by (simp add: fiter_id_def fiter_def gfp_fusion_var hyp iso_fiter_fun2 iso_id iso_fiter_ext) qed subsection \Propositional Hoare Logic Completed\ lemma H_weak_loop: "mono f \ x \ f x \ x \ fiter_id f x" by (force simp: fstar_pred_char gfp_def intro: Sup_upper) lemma iso_fiter: "mono f \ mono (fiter_id f)" unfolding mono_def by (subst fstar_pred_char, simp add: mono_def)+ (auto intro: gfp_mono inf_mono) text \As already mentioned, a dual Hoare logic can be built for the dual lattice. In this case, weak iteration is defined with respect to Sup.\ text \The following standard construction lifts elements of (meet semi)lattices to transformers. I allow a more general type.\ definition fqtran :: "'a::inf \ 'a \ 'a" where "fqtran x \ \y. x \ y" text \The following standard construction lifts elements of boolean algebras to transformers.\ definition bqtran :: "'a::boolean_algebra \ 'a \ 'a" ("\_\") where "\x\ y = -x \ y" text \The conditional and while rule of Hoare logic are now derivable.\ lemma bqtran_iso: "mono \x\" by (metis bqtran_def monoI order_refl sup.mono) lemma cond_iso: "mono f \ mono g \ mono (\x\ \ f \ \y\ \ g)" by (simp add: bqtran_iso iso_fcomp iso_finf) lemma loop_iso: "mono f \ mono (fiter_id (\x\ \ f) \ \y\)" by (simp add: bqtran_iso iso_fcomp iso_fiter) lemma H_iso_cond: "mono f \ mono g \ p \ x \ f y \ q \ x \ g y \ x \ (inf (\p\ \ f) (\q\ \ g)) y" by (metis (full_types) bqtran_def comp_apply inf_apply inf_commute le_inf_iff shunt1) lemma H_iso_loop: "mono f \ p \ x \ f x \ x \ ((fiter_id (\p\ \ f)) \ \q\) (x \ q)" proof- assume a: "mono f" and "p \ x \ f x" hence "x \ (\p\ \ f) x" using H_iso_cond by fastforce hence "x \ (fiter_id (\p\ \ f)) x" by (simp add: H_weak_loop a bqtran_iso iso_fcomp) also have "... \ (fiter_id (\p\ \ f)) (-q \ (x \ q))" by (meson a bqtran_iso dual_order.refl iso_fcomp iso_fiter monoD shunt1) finally show "x \ ((fiter_id (\p\ \ f)) \ \q\) (x \ q)" by (simp add: bqtran_def) qed lemma btran_spec: "x \ \y\ (x \ y)" by (simp add: bqtran_def sup_inf_distrib1) lemma btran_neg_spec: "x \ \-y\ (x - y)" by (simp add: btran_spec diff_eq) subsection \A Propositional Refinement Calculus\ text \Next I derive the laws of an abstract Propositional Refinement Calculus, Morgan-style. These are given without the co-called frames, which capture information about local and global variables in variants of this calculus.\ definition "Ri x y z = \{f z |f. x \ f y \ mono (f::'a::order \ 'b::complete_lattice)}" lemma Ri_least: "mono f \ x \ f y \ Ri x y z \ f z" unfolding Ri_def by (metis (mono_tags, lifting) Inf_lower mem_Collect_eq) lemma Ri_spec: "x \ Ri x y y" unfolding Ri_def by (rule Inf_greatest, safe) lemma Ri_spec_var: "(\z. Ri x y z \ f z) \ x \ f y" using Ri_spec dual_order.trans by blast lemma Ri_prop: "mono f \ x \ f y \ (\z. Ri x y z \ f z)" using Ri_least Ri_spec_var by blast lemma iso_Ri: "mono (Ri x y)" unfolding mono_def Ri_def by (auto intro!: Inf_mono) lemma Ri_weaken: "x \ x' \ y' \ y \ Ri x y z \ Ri x' y' z" by (meson H_iso_cond2 Ri_least Ri_spec iso_Ri order.trans) lemma Ri_seq: "Ri x y z \ Ri x w (Ri w y z)" by (metis (no_types, opaque_lifting) H_iso_cond2 Ri_prop Ri_spec iso_Ri iso_fcomp o_apply) lemma Ri_seq_var: "Ri x y z \ ((Ri x w) \ (Ri w y)) z" by (simp add: Ri_seq) lemma Ri_Inf: " Ri (\ X) y z \ \{Ri x y z |x. x \ X}" by (safe intro!: Inf_greatest, simp add: Ri_weaken Inf_lower) lemma Ri_weak_iter: "Ri x x y \ fiter_id (Ri x x) y" by (simp add: H_weak_loop Ri_least Ri_spec iso_Ri iso_fiter) lemma Ri_cond: "Ri x y z \ (inf (\p\ \ (Ri (p \ x) y)) ((\q\ \ (Ri (q \ x) y)))) z" by (meson H_iso_cond Ri_least Ri_spec bqtran_iso iso_Ri iso_fcomp iso_finf) lemma Ri_loop: "Ri x (q \ x) y \ ((fiter_id (\p\ \ (Ri (x \ p) x))) \ \q\) (q \ y)" proof- have "(p \ x) \ Ri (p \ x) x x" by (simp add: Ri_spec) hence "x \ ((fiter_id (\p\ \ (Ri (x \ p) x))) \ \q\) (q \ x)" by (metis H_iso_loop inf_commute iso_Ri) thus ?thesis apply (subst Ri_least, safe, simp_all add: mono_def) by (metis bqtran_iso inf_mono iso_Ri iso_fcomp iso_fiter mono_def order_refl) qed end diff --git a/thys/Types_To_Sets_Extension/Examples/SML_Relativization/Lattices/SML_Complete_Lattices.thy b/thys/Types_To_Sets_Extension/Examples/SML_Relativization/Lattices/SML_Complete_Lattices.thy --- a/thys/Types_To_Sets_Extension/Examples/SML_Relativization/Lattices/SML_Complete_Lattices.thy +++ b/thys/Types_To_Sets_Extension/Examples/SML_Relativization/Lattices/SML_Complete_Lattices.thy @@ -1,1240 +1,1152 @@ (* Title: Examples/SML_Relativization/Lattices/SML_Complete_Lattices.thy Author: Mihails Milehins Copyright 2021 (C) Mihails Milehins *) section\Relativization of the results about complete lattices\ theory SML_Complete_Lattices imports SML_Lattices begin subsection\Simple complete lattices\ subsubsection\Definitions and common properties\ locale Inf_ow = fixes U :: "'al set" and Inf (\\\<^sub>o\<^sub>w\) assumes Inf_closed[simp]: "\\<^sub>o\<^sub>w ` Pow U \ U" begin notation Inf (\\\<^sub>o\<^sub>w\) lemma Inf_closed'[simp]: "\x\U. \\<^sub>o\<^sub>w x \ U" using Inf_closed by blast end locale Inf_pair_ow = Inf\<^sub>1: Inf_ow U\<^sub>1 Inf\<^sub>1 + Inf\<^sub>2: Inf_ow U\<^sub>2 Inf\<^sub>2 for U\<^sub>1 :: "'al set" and Inf\<^sub>1 and U\<^sub>2 :: "'bl set" and Inf\<^sub>2 begin notation Inf\<^sub>1 (\\\<^sub>o\<^sub>w\<^sub>.\<^sub>1\) notation Inf\<^sub>2 (\\\<^sub>o\<^sub>w\<^sub>.\<^sub>2\) end locale Sup_ow = fixes U :: "'al set" and Sup (\\\<^sub>o\<^sub>w\) assumes Sup_closed[simp]: "\\<^sub>o\<^sub>w ` Pow U \ U" begin notation Sup (\\\<^sub>o\<^sub>w\) lemma Sup_closed'[simp]: "\x\U. \\<^sub>o\<^sub>w x \ U" using Sup_closed by blast end locale Sup_pair_ow = Sup\<^sub>1: Sup_ow U\<^sub>1 Sup\<^sub>1 + Sup\<^sub>2: Sup_ow U\<^sub>2 Sup\<^sub>2 for U\<^sub>1 :: "'al set" and Sup\<^sub>1 and U\<^sub>2 :: "'bl set" and Sup\<^sub>2 begin notation Sup\<^sub>1 (\\\<^sub>o\<^sub>w\<^sub>.\<^sub>1\) notation Sup\<^sub>2 (\\\<^sub>o\<^sub>w\<^sub>.\<^sub>2\) end locale complete_lattice_ow = lattice_ow U inf le ls sup + Inf_ow U Inf + Sup_ow U Sup + bot_ow U bot + top_ow U top for U :: "'al set" and Inf Sup inf le ls sup bot top + assumes Inf_lower: "\ A \ U; x \ A \ \ \\<^sub>o\<^sub>w A \\<^sub>o\<^sub>w x" and Inf_greatest: "\ A \ U; z \ U; (\x. x \ A \ z \\<^sub>o\<^sub>w x) \ \ z \\<^sub>o\<^sub>w \\<^sub>o\<^sub>w A" and Sup_upper: "\ x \ A; A \ U \ \ x \\<^sub>o\<^sub>w \\<^sub>o\<^sub>w A" and Sup_least: "\ A \ U; z \ U; (\x. x \ A \ x \\<^sub>o\<^sub>w z) \ \ \\<^sub>o\<^sub>w A \\<^sub>o\<^sub>w z" and Inf_empty[simp]: "\\<^sub>o\<^sub>w {} = \\<^sub>o\<^sub>w" and Sup_empty[simp]: "\\<^sub>o\<^sub>w {} = \\<^sub>o\<^sub>w" begin sublocale bounded_lattice_ow U \(\\<^sub>o\<^sub>w)\ \(\\<^sub>o\<^sub>w)\ \(<\<^sub>o\<^sub>w)\ \(\\<^sub>o\<^sub>w)\ \\\<^sub>o\<^sub>w\ \\\<^sub>o\<^sub>w\ apply standard subgoal using Sup_least by force subgoal using Inf_greatest by fastforce done tts_register_sbts \\\<^sub>o\<^sub>w\ | U proof- assume ALA: "Domainp ALA = (\x. x \ U)" have "Domainp ALA \\<^sub>o\<^sub>w" unfolding ALA by simp then obtain bot' where "ALA \\<^sub>o\<^sub>w bot'" by blast then show "\bot'. ALA \\<^sub>o\<^sub>w bot'" by auto qed tts_register_sbts \\\<^sub>o\<^sub>w\ | U proof- assume ALA: "Domainp ALA = (\x. x \ U)" have "Domainp ALA \\<^sub>o\<^sub>w" unfolding ALA by simp then obtain top' where "ALA \\<^sub>o\<^sub>w top'" by blast then show "\top'. ALA \\<^sub>o\<^sub>w top'" by auto qed end locale complete_lattice_pair_ow = cl\<^sub>1: complete_lattice_ow U\<^sub>1 Inf\<^sub>1 Sup\<^sub>1 inf\<^sub>1 le\<^sub>1 ls\<^sub>1 sup\<^sub>1 bot\<^sub>1 top\<^sub>1 + cl\<^sub>2: complete_lattice_ow U\<^sub>2 Inf\<^sub>2 Sup\<^sub>2 inf\<^sub>2 le\<^sub>2 ls\<^sub>2 sup\<^sub>2 bot\<^sub>2 top\<^sub>2 for U\<^sub>1 :: "'al set" and Inf\<^sub>1 Sup\<^sub>1 inf\<^sub>1 le\<^sub>1 ls\<^sub>1 sup\<^sub>1 bot\<^sub>1 top\<^sub>1 and U\<^sub>2 :: "'bl set" and Inf\<^sub>2 Sup\<^sub>2 inf\<^sub>2 le\<^sub>2 ls\<^sub>2 sup\<^sub>2 bot\<^sub>2 top\<^sub>2 begin sublocale bounded_lattice_pair_ow .. sublocale Inf_pair_ow .. sublocale Sup_pair_ow .. end subsubsection\Transfer rules\ context includes lifting_syntax begin lemma complete_lattice_transfer[transfer_rule]: assumes [transfer_rule]: "bi_unique A" "right_total A" shows "( (rel_set A ===> A) ===> (rel_set A ===> A) ===> (A ===> A ===> A) ===> (A ===> A ===> (=)) ===> (A ===> A ===> (=)) ===> (A ===> A ===> A) ===> A ===> A ===> (=) ) (complete_lattice_ow (Collect (Domainp A))) class.complete_lattice" proof- let ?P = "( (rel_set A ===> A) ===> (rel_set A ===> A) ===> (A ===> A ===> A) ===> (A ===> A ===> (=)) ===> (A ===> A ===> (=)) ===> (A ===> A ===> A) ===> A ===> A ===> (=) )" let ?complete_lattice_ow = "(complete_lattice_ow (Collect (Domainp A)))" let ?Inf_Sup_UNIV = "(\F'::'b set \ 'b. (F' ` Pow UNIV \ UNIV))" have rrng: "((A2 \ A3 \ A4 \ A5 \ A1 \ A6 \ A7 \ A8 \ A9 \ A10 \ A11) = F') \ ((A1 \ A2 \ A3 \ A4 \ A5 \ A6 \ A7 \ A8 \ A9 \ A10 \ A11) = F')" for A1 A2 A3 A4 A5 A6 A7 A8 A9 A10 A11 F' by auto have "?P ?complete_lattice_ow ( \Inf Sup inf le ls sup bot top. ?Inf_Sup_UNIV Inf \ ?Inf_Sup_UNIV Sup \ bot \ UNIV \ top \ UNIV \ class.complete_lattice Inf Sup inf le ls sup bot top )" unfolding complete_lattice_ow_def class.complete_lattice_def complete_lattice_ow_axioms_def class.complete_lattice_axioms_def Inf_ow_def Sup_ow_def bot_ow_def top_ow_def apply transfer_prover_start apply transfer_step+ apply(simp, intro ext, rule rrng, intro arg_cong2[where f="(\)"]) subgoal unfolding Ball_Collect by simp subgoal unfolding Ball_Collect by simp subgoal by simp subgoal by simp subgoal by simp subgoal unfolding Ball_Collect[symmetric] by auto subgoal unfolding Ball_Collect[symmetric] by metis subgoal unfolding Ball_Collect[symmetric] by auto subgoal unfolding Ball_Collect[symmetric] by metis subgoal by simp subgoal by simp done then show ?thesis by simp qed end subsubsection\Relativization\ context complete_lattice_ow begin tts_context tts: (?'a to U) rewriting ctr_simps substituting complete_lattice_ow_axioms and sup_bot.sl_neut.not_empty eliminating \?a \ ?A\ and \?A \ ?B\ through auto applying [OF Inf_closed' Sup_closed' inf_closed' sup_closed'] begin tts_lemma Inf_atLeast: assumes "x \ U" shows "\\<^sub>o\<^sub>w {x..\<^sub>o\<^sub>w} = x" is complete_lattice_class.Inf_atLeast. tts_lemma Inf_atMost: assumes "x \ U" shows "\\<^sub>o\<^sub>w {..\<^sub>o\<^sub>wx} = \\<^sub>o\<^sub>w" is complete_lattice_class.Inf_atMost. tts_lemma Sup_atLeast: assumes "x \ U" shows "\\<^sub>o\<^sub>w {x..\<^sub>o\<^sub>w} = \\<^sub>o\<^sub>w" is complete_lattice_class.Sup_atLeast. tts_lemma Sup_atMost: assumes "y \ U" shows "\\<^sub>o\<^sub>w {..\<^sub>o\<^sub>wy} = y" is complete_lattice_class.Sup_atMost. tts_lemma Inf_insert: assumes "a \ U" and "A \ U" shows "\\<^sub>o\<^sub>w (insert a A) = a \\<^sub>o\<^sub>w \\<^sub>o\<^sub>w A" is complete_lattice_class.Inf_insert. tts_lemma Sup_insert: assumes "a \ U" and "A \ U" shows "\\<^sub>o\<^sub>w (insert a A) = a \\<^sub>o\<^sub>w \\<^sub>o\<^sub>w A" is complete_lattice_class.Sup_insert. tts_lemma Inf_atMostLessThan: assumes "x \ U" and "\\<^sub>o\<^sub>w <\<^sub>o\<^sub>w x" shows "\\<^sub>o\<^sub>w {..<\<^sub>o\<^sub>wx} = \\<^sub>o\<^sub>w" is complete_lattice_class.Inf_atMostLessThan. tts_lemma Sup_greaterThanAtLeast: assumes "x \ U" and "x <\<^sub>o\<^sub>w \\<^sub>o\<^sub>w" shows "\\<^sub>o\<^sub>w {x<\<^sub>o\<^sub>w..} = \\<^sub>o\<^sub>w" is complete_lattice_class.Sup_greaterThanAtLeast. tts_lemma le_Inf_iff: assumes "b \ U" and "A \ U" shows "(b \\<^sub>o\<^sub>w \\<^sub>o\<^sub>w A) = Ball A ((\\<^sub>o\<^sub>w) b)" is complete_lattice_class.le_Inf_iff. tts_lemma Sup_le_iff: assumes "A \ U" and "b \ U" shows "(\\<^sub>o\<^sub>w A \\<^sub>o\<^sub>w b) = (\x\A. x \\<^sub>o\<^sub>w b)" is complete_lattice_class.Sup_le_iff. tts_lemma Inf_atLeastLessThan: assumes "x \ U" and "y \ U" and "x <\<^sub>o\<^sub>w y" shows "\\<^sub>o\<^sub>w (on U with (\\<^sub>o\<^sub>w) (<\<^sub>o\<^sub>w) : {x.. U" and "y \ U" and "x <\<^sub>o\<^sub>w y" shows "\\<^sub>o\<^sub>w {x<\<^sub>o\<^sub>w..y} = y" is complete_lattice_class.Sup_greaterThanAtMost. tts_lemma Inf_atLeastAtMost: assumes "x \ U" and "y \ U" and "x \\<^sub>o\<^sub>w y" shows "\\<^sub>o\<^sub>w {x..\<^sub>o\<^sub>wy} = x" is complete_lattice_class.Inf_atLeastAtMost. tts_lemma Sup_atLeastAtMost: assumes "x \ U" and "y \ U" and "x \\<^sub>o\<^sub>w y" shows "\\<^sub>o\<^sub>w {x..\<^sub>o\<^sub>wy} = y" is complete_lattice_class.Sup_atLeastAtMost. tts_lemma Inf_lower2: assumes "A \ U" and "v \ U" and "u \ A" and "u \\<^sub>o\<^sub>w v" shows "\\<^sub>o\<^sub>w A \\<^sub>o\<^sub>w v" is complete_lattice_class.Inf_lower2. tts_lemma Sup_upper2: assumes "A \ U" and "v \ U" and "u \ A" and "v \\<^sub>o\<^sub>w u" shows "v \\<^sub>o\<^sub>w \\<^sub>o\<^sub>w A" is complete_lattice_class.Sup_upper2. tts_lemma Inf_less_eq: assumes "A \ U" and "u \ U" and "\v. v \ A \ v \\<^sub>o\<^sub>w u" and "A \ {}" shows "\\<^sub>o\<^sub>w A \\<^sub>o\<^sub>w u" given complete_lattice_class.Inf_less_eq by auto tts_lemma less_eq_Sup: assumes "A \ U" and "u \ U" and "\v. v \ A \ u \\<^sub>o\<^sub>w v" and "A \ {}" shows "u \\<^sub>o\<^sub>w \\<^sub>o\<^sub>w A" given complete_lattice_class.less_eq_Sup by auto tts_lemma Sup_eqI: assumes "A \ U" and "x \ U" and "\y. y \ A \ y \\<^sub>o\<^sub>w x" and "\y. \y \ U; \z. z \ A \ z \\<^sub>o\<^sub>w y\ \ x \\<^sub>o\<^sub>w y" shows "\\<^sub>o\<^sub>w A = x" given complete_lattice_class.Sup_eqI by (simp add: Sup_least Sup_upper order.antisym) tts_lemma Inf_eqI: assumes "A \ U" and "x \ U" and "\i. i \ A \ x \\<^sub>o\<^sub>w i" and "\y. \y \ U; \i. i \ A \ y \\<^sub>o\<^sub>w i\ \ y \\<^sub>o\<^sub>w x" shows "\\<^sub>o\<^sub>w A = x" given complete_lattice_class.Inf_eqI by (simp add: subset_eq) tts_lemma Inf_union_distrib: assumes "A \ U" and "B \ U" shows "\\<^sub>o\<^sub>w (A \ B) = \\<^sub>o\<^sub>w A \\<^sub>o\<^sub>w \\<^sub>o\<^sub>w B" is complete_lattice_class.Inf_union_distrib. tts_lemma Sup_union_distrib: assumes "A \ U" and "B \ U" shows "\\<^sub>o\<^sub>w (A \ B) = \\<^sub>o\<^sub>w A \\<^sub>o\<^sub>w \\<^sub>o\<^sub>w B" is complete_lattice_class.Sup_union_distrib. tts_lemma Sup_inter_less_eq: assumes "A \ U" and "B \ U" shows "\\<^sub>o\<^sub>w (A \ B) \\<^sub>o\<^sub>w \\<^sub>o\<^sub>w A \\<^sub>o\<^sub>w \\<^sub>o\<^sub>w B" is complete_lattice_class.Sup_inter_less_eq. tts_lemma less_eq_Inf_inter: assumes "A \ U" and "B \ U" shows "\\<^sub>o\<^sub>w A \\<^sub>o\<^sub>w \\<^sub>o\<^sub>w B \\<^sub>o\<^sub>w \\<^sub>o\<^sub>w (A \ B)" is complete_lattice_class.less_eq_Inf_inter. tts_lemma Sup_subset_mono: assumes "B \ U" and "A \ B" shows "\\<^sub>o\<^sub>w A \\<^sub>o\<^sub>w \\<^sub>o\<^sub>w B" is complete_lattice_class.Sup_subset_mono. tts_lemma Inf_superset_mono: assumes "A \ U" and "B \ A" shows "\\<^sub>o\<^sub>w A \\<^sub>o\<^sub>w \\<^sub>o\<^sub>w B" is complete_lattice_class.Inf_superset_mono. tts_lemma Sup_bot_conv: assumes "A \ U" shows "(\\<^sub>o\<^sub>w A = \\<^sub>o\<^sub>w) = (\x\A. x = \\<^sub>o\<^sub>w)" "(\\<^sub>o\<^sub>w = \\<^sub>o\<^sub>w A) = (\x\A. x = \\<^sub>o\<^sub>w)" is complete_lattice_class.Sup_bot_conv. tts_lemma Inf_top_conv: assumes "A \ U" shows "(\\<^sub>o\<^sub>w A = \\<^sub>o\<^sub>w) = (\x\A. x = \\<^sub>o\<^sub>w)" "(\\<^sub>o\<^sub>w = \\<^sub>o\<^sub>w A) = (\x\A. x = \\<^sub>o\<^sub>w)" is complete_lattice_class.Inf_top_conv. tts_lemma Inf_le_Sup: assumes "A \ U" and "A \ {}" shows "\\<^sub>o\<^sub>w A \\<^sub>o\<^sub>w \\<^sub>o\<^sub>w A" is complete_lattice_class.Inf_le_Sup. tts_lemma INF_UNIV_bool_expand: assumes "range A \ U" shows "\\<^sub>o\<^sub>w (range A) = A True \\<^sub>o\<^sub>w A False" is complete_lattice_class.INF_UNIV_bool_expand. tts_lemma SUP_UNIV_bool_expand: assumes "range A \ U" shows "\\<^sub>o\<^sub>w (range A) = A True \\<^sub>o\<^sub>w A False" is complete_lattice_class.SUP_UNIV_bool_expand. tts_lemma Sup_mono: assumes "A \ U" and "B \ U" and "\a. a \ A \ Bex B ((\\<^sub>o\<^sub>w) a)" shows "\\<^sub>o\<^sub>w A \\<^sub>o\<^sub>w \\<^sub>o\<^sub>w B" given complete_lattice_class.Sup_mono by simp tts_lemma Inf_mono: assumes "B \ U" and "A \ U" and "\b. b \ B \ \x\A. x \\<^sub>o\<^sub>w b" shows "\\<^sub>o\<^sub>w A \\<^sub>o\<^sub>w \\<^sub>o\<^sub>w B" given complete_lattice_class.Inf_mono by simp tts_lemma Sup_Inf_le: assumes "A \ Pow U" shows "\\<^sub>o\<^sub>w ( \\<^sub>o\<^sub>w ` {x. x \ U \ (\f\{f. f ` Pow U \ U}. x = f ` A \ (\Y\A. f Y \ Y))} ) \\<^sub>o\<^sub>w \\<^sub>o\<^sub>w (\\<^sub>o\<^sub>w ` A)" is complete_lattice_class.Sup_Inf_le. end tts_context tts: (?'a to U) rewriting ctr_simps substituting complete_lattice_ow_axioms and sup_bot.sl_neut.not_empty applying [ OF Inf_closed' Sup_closed' inf_closed' sup_closed' bot_closed top_closed ] begin tts_lemma Inf_UNIV: "\\<^sub>o\<^sub>w U = \\<^sub>o\<^sub>w" is complete_lattice_class.Inf_UNIV. tts_lemma Sup_UNIV: "\\<^sub>o\<^sub>w U = \\<^sub>o\<^sub>w" is complete_lattice_class.Sup_UNIV. end context fixes U\<^sub>2 :: "'b set" begin lemmas [transfer_rule] = image_transfer[where ?'a='b] tts_context tts: (?'a to U) and (?'b to \U\<^sub>2::'b set\) rewriting ctr_simps substituting complete_lattice_ow_axioms and sup_bot.sl_neut.not_empty eliminating through (insert Sup_least, auto) begin tts_lemma SUP_upper2: assumes "A \ U\<^sub>2" and "u \ U" and "\x\U\<^sub>2. f x \ U" and "i \ A" and "u \\<^sub>o\<^sub>w f i" shows "u \\<^sub>o\<^sub>w \\<^sub>o\<^sub>w (f ` A)" is complete_lattice_class.SUP_upper2. tts_lemma INF_lower2: assumes "A \ U\<^sub>2" and "\x\U\<^sub>2. f x \ U" and "u \ U" and "i \ A" and "f i \\<^sub>o\<^sub>w u" shows "\\<^sub>o\<^sub>w (f ` A) \\<^sub>o\<^sub>w u" is complete_lattice_class.INF_lower2. tts_lemma less_INF_D: assumes "y \ U" and "\x\U\<^sub>2. f x \ U" and "A \ U\<^sub>2" and "y <\<^sub>o\<^sub>w \\<^sub>o\<^sub>w (f ` A)" and "i \ A" shows "y <\<^sub>o\<^sub>w f i" is complete_lattice_class.less_INF_D. tts_lemma SUP_lessD: assumes "\x\U\<^sub>2. f x \ U" and "A \ U\<^sub>2" and "y \ U" and "\\<^sub>o\<^sub>w (f ` A) <\<^sub>o\<^sub>w y" and "i \ A" shows "f i <\<^sub>o\<^sub>w y" is complete_lattice_class.SUP_lessD. tts_lemma SUP_le_iff: assumes "\x\U\<^sub>2. f x \ U" and "A \ U\<^sub>2" and "u \ U" shows "(\\<^sub>o\<^sub>w (f ` A) \\<^sub>o\<^sub>w u) = (\x\A. f x \\<^sub>o\<^sub>w u)" is complete_lattice_class.SUP_le_iff. end end tts_context tts: (?'a to U) and (?'b to \U\<^sub>2::'b set\) rewriting ctr_simps substituting complete_lattice_ow_axioms and sup_bot.sl_neut.not_empty eliminating through (auto dest: top_le) begin tts_lemma INF_eqI: assumes "A \ U\<^sub>2" and "x \ U" and "\x\U\<^sub>2. f x \ U" and "\i. \i \ U\<^sub>2; i \ A\ \ x \\<^sub>o\<^sub>w f i" and "\y. \y \ U; \i. \i \ U\<^sub>2; i \ A\ \ y \\<^sub>o\<^sub>w f i\ \ y \\<^sub>o\<^sub>w x" shows "\\<^sub>o\<^sub>w (f ` A) = x" is complete_lattice_class.INF_eqI. end tts_context tts: (?'a to U) and (?'b to \U\<^sub>2::'b set\) rewriting ctr_simps substituting complete_lattice_ow_axioms and sup_bot.sl_neut.not_empty eliminating through (auto simp: order.eq_iff) begin tts_lemma SUP_eqI: assumes "A \ U\<^sub>2" and "\x\U\<^sub>2. f x \ U" and "x \ U" and "\i. \i \ U\<^sub>2; i \ A\ \ f i \\<^sub>o\<^sub>w x" and "\y. \y \ U; \i. \i \ U\<^sub>2; i \ A\ \ f i \\<^sub>o\<^sub>w y\ \ x \\<^sub>o\<^sub>w y" shows "\\<^sub>o\<^sub>w (f ` A) = x" is complete_lattice_class.SUP_eqI. end tts_context tts: (?'a to U) and (?'b to \U\<^sub>2::'b set\) rewriting ctr_simps substituting complete_lattice_ow_axioms and sup_bot.sl_neut.not_empty eliminating through simp begin tts_lemma INF_le_SUP: assumes "A \ U\<^sub>2" and "\x\U\<^sub>2. f x \ U" and "A \ {}" shows "\\<^sub>o\<^sub>w (f ` A) \\<^sub>o\<^sub>w \\<^sub>o\<^sub>w (f ` A)" is complete_lattice_class.INF_le_SUP. tts_lemma INF_inf_const1: assumes "I \ U\<^sub>2" and "x \ U" and "\x\U\<^sub>2. f x \ U" and "I \ {}" shows "\\<^sub>o\<^sub>w ((\i. x \\<^sub>o\<^sub>w f i) ` I) = x \\<^sub>o\<^sub>w \\<^sub>o\<^sub>w (f ` I)" is complete_lattice_class.INF_inf_const1. tts_lemma INF_inf_const2: assumes "I \ U\<^sub>2" and "\x\U\<^sub>2. f x \ U" and "x \ U" and "I \ {}" shows "\\<^sub>o\<^sub>w ((\i. f i \\<^sub>o\<^sub>w x) ` I) = \\<^sub>o\<^sub>w (f ` I) \\<^sub>o\<^sub>w x" is complete_lattice_class.INF_inf_const2. end tts_context tts: (?'a to U) and (?'b to \U\<^sub>2::'b set\) rewriting ctr_simps substituting complete_lattice_ow_axioms and sup_bot.sl_neut.not_empty eliminating through auto begin tts_lemma INF_eq_const: assumes "I \ U\<^sub>2" and "\x\U\<^sub>2. f x \ U" and "I \ {}" and "\i. i \ I \ f i = x" shows "\\<^sub>o\<^sub>w (f ` I) = x" given complete_lattice_class.INF_eq_const proof- assume "\I \ U\<^sub>2; \x\U\<^sub>2. f x \ U; I \ {}; \i. \i \ U\<^sub>2; i \ I\ \ f i = x\ \ \\<^sub>o\<^sub>w (f ` I) = x" for I :: "'a set" and U\<^sub>2 f x then have "\I \ U\<^sub>2; \x\U\<^sub>2. f x \ U; I \ {}; \i. i \ I \ f i = x\ \ \\<^sub>o\<^sub>w (f ` I) = x" by presburger then show "\\<^sub>o\<^sub>w (f ` I) = x" using assms by simp qed tts_lemma SUP_eq_const: assumes "I \ U\<^sub>2" and "\x\U\<^sub>2. f x \ U" and "I \ {}" and "\i. i \ I \ f i = x" shows "\\<^sub>o\<^sub>w (f ` I) = x" given complete_lattice_class.SUP_eq_const proof- assume "\I \ U\<^sub>2; \x\U\<^sub>2. f x \ U; I \ {}; \i. \i \ U\<^sub>2; i \ I\ \ f i = x\ \ \\<^sub>o\<^sub>w (f ` I) = x" for I :: "'a set" and U\<^sub>2 f x then have "\I \ U\<^sub>2; \x\U\<^sub>2. f x \ U; I \ {}; \i. i \ I \ f i = x\ \ \\<^sub>o\<^sub>w (f ` I) = x" by presburger then show "\\<^sub>o\<^sub>w (f ` I) = x" using assms by simp qed tts_lemma SUP_eq_iff: assumes "I \ U\<^sub>2" and "c \ U" and "\x\U\<^sub>2. f x \ U" and "I \ {}" and "\i. i \ I \ c \\<^sub>o\<^sub>w f i" shows "(\\<^sub>o\<^sub>w (f ` I) = c) = (\x\I. f x = c)" given complete_lattice_class.SUP_eq_iff proof- assume "\ I \ U\<^sub>2; c \ U; \x\U\<^sub>2. f x \ U; I \ {}; \i. \i \ U\<^sub>2; i \ I\ \ c \\<^sub>o\<^sub>w f i \ \ (\\<^sub>o\<^sub>w (f ` I) = c) = (\x\I. f x = c)" for I :: "'a set" and U\<^sub>2 c f then have "\ I \ U\<^sub>2; c \ U; \x\U\<^sub>2. f x \ U; I \ {}; \i. i \ I \ c \\<^sub>o\<^sub>w f i \ \ (\\<^sub>o\<^sub>w (f ` I) = c) = (\x\I. f x = c)" by presburger then show "(\\<^sub>o\<^sub>w (f ` I) = c) = (\x\I. f x = c)" using assms by auto qed tts_lemma INF_eq_iff: assumes "I \ U\<^sub>2" and "\x\U\<^sub>2. f x \ U" and "c \ U" and "I \ {}" and "\i. i \ I \ f i \\<^sub>o\<^sub>w c" shows "(\\<^sub>o\<^sub>w (f ` I) = c) = (\x\I. f x = c)" given complete_lattice_class.INF_eq_iff proof- assume "\ I \ U\<^sub>2; \x\U\<^sub>2. f x \ U; c \ U; I \ {}; \i. \i \ U\<^sub>2; i \ I\ \ f i \\<^sub>o\<^sub>w c \ \ (\\<^sub>o\<^sub>w (f ` I) = c) = (\x\I. f x = c)" for I :: "'a set" and U\<^sub>2 f c then have "\ I \ U\<^sub>2; \x\U\<^sub>2. f x \ U; c \ U; I \ {}; \i. i \ I \ f i \\<^sub>o\<^sub>w c \ \ (\\<^sub>o\<^sub>w (f ` I) = c) = (\x\I. f x = c)" by presburger then show "(\\<^sub>o\<^sub>w (f ` I) = c) = (\x\I. f x = c)" using assms by auto qed end context fixes U\<^sub>2 :: "'b set" begin lemmas [transfer_rule] = image_transfer[where ?'a='b] tts_context tts: (?'a to U) and (?'b to \U\<^sub>2::'b set\) rewriting ctr_simps substituting complete_lattice_ow_axioms and sup_bot.sl_neut.not_empty eliminating through (auto simp: sup_bot.top_unique top_le intro: Sup_least) begin tts_lemma INF_insert: assumes "\x\U\<^sub>2. f x \ U" and "a \ U\<^sub>2" and "A \ U\<^sub>2" shows "\\<^sub>o\<^sub>w (f ` insert a A) = f a \\<^sub>o\<^sub>w \\<^sub>o\<^sub>w (f ` A)" is complete_lattice_class.INF_insert. tts_lemma SUP_insert: assumes "\x\U\<^sub>2. f x \ U" and "a \ U\<^sub>2" and "A \ U\<^sub>2" shows "\\<^sub>o\<^sub>w (f ` insert a A) = f a \\<^sub>o\<^sub>w \\<^sub>o\<^sub>w (f ` A)" is complete_lattice_class.SUP_insert. tts_lemma le_INF_iff: assumes "u \ U" and "\x\U\<^sub>2. f x \ U" and "A \ U\<^sub>2" shows "(u \\<^sub>o\<^sub>w \\<^sub>o\<^sub>w (f ` A)) = (\x\A. u \\<^sub>o\<^sub>w f x)" is complete_lattice_class.le_INF_iff. tts_lemma INF_union: assumes "\x\U\<^sub>2. M x \ U" and "A \ U\<^sub>2" and "B \ U\<^sub>2" shows "\\<^sub>o\<^sub>w (M ` (A \ B)) = \\<^sub>o\<^sub>w (M ` A) \\<^sub>o\<^sub>w \\<^sub>o\<^sub>w (M ` B)" is complete_lattice_class.INF_union. tts_lemma SUP_union: assumes "\x\U\<^sub>2. M x \ U" and "A \ U\<^sub>2" and "B \ U\<^sub>2" shows "\\<^sub>o\<^sub>w (M ` (A \ B)) = \\<^sub>o\<^sub>w (M ` A) \\<^sub>o\<^sub>w \\<^sub>o\<^sub>w (M ` B)" is complete_lattice_class.SUP_union. tts_lemma SUP_bot_conv: assumes "\x\U\<^sub>2. B x \ U" and "A \ U\<^sub>2" shows "(\\<^sub>o\<^sub>w (B ` A) = \\<^sub>o\<^sub>w) = (\x\A. B x = \\<^sub>o\<^sub>w)" "(\\<^sub>o\<^sub>w = \\<^sub>o\<^sub>w (B ` A)) = (\x\A. B x = \\<^sub>o\<^sub>w)" is complete_lattice_class.SUP_bot_conv. tts_lemma INF_top_conv: assumes "\x\U\<^sub>2. B x \ U" and "A \ U\<^sub>2" shows "(\\<^sub>o\<^sub>w (B ` A) = \\<^sub>o\<^sub>w) = (\x\A. B x = \\<^sub>o\<^sub>w)" "(\\<^sub>o\<^sub>w = \\<^sub>o\<^sub>w (B ` A)) = (\x\A. B x = \\<^sub>o\<^sub>w)" is complete_lattice_class.INF_top_conv. tts_lemma SUP_upper: assumes "A \ U\<^sub>2" and "\x\U\<^sub>2. f x \ U" and "i \ A" shows "f i \\<^sub>o\<^sub>w \\<^sub>o\<^sub>w (f ` A)" is complete_lattice_class.SUP_upper. tts_lemma INF_lower: assumes "A \ U\<^sub>2" and "\x\U\<^sub>2. f x \ U" and "i \ A" shows "\\<^sub>o\<^sub>w (f ` A) \\<^sub>o\<^sub>w f i" is complete_lattice_class.INF_lower. tts_lemma INF_inf_distrib: assumes "\x\U\<^sub>2. f x \ U" and "A \ U\<^sub>2" and "\x\U\<^sub>2. g x \ U" shows "\\<^sub>o\<^sub>w (f ` A) \\<^sub>o\<^sub>w \\<^sub>o\<^sub>w (g ` A) = \\<^sub>o\<^sub>w ((\a. f a \\<^sub>o\<^sub>w g a) ` A)" is complete_lattice_class.INF_inf_distrib. tts_lemma SUP_sup_distrib: assumes "\x\U\<^sub>2. f x \ U" and "A \ U\<^sub>2" and "\x\U\<^sub>2. g x \ U" shows "\\<^sub>o\<^sub>w (f ` A) \\<^sub>o\<^sub>w \\<^sub>o\<^sub>w (g ` A) = \\<^sub>o\<^sub>w ((\a. f a \\<^sub>o\<^sub>w g a) ` A)" is complete_lattice_class.SUP_sup_distrib. tts_lemma SUP_subset_mono: assumes "B \ U\<^sub>2" and "\x\U\<^sub>2. f x \ U" and "\x\U\<^sub>2. g x \ U" and "A \ B" and "\x. x \ A \ f x \\<^sub>o\<^sub>w g x" shows "\\<^sub>o\<^sub>w (f ` A) \\<^sub>o\<^sub>w \\<^sub>o\<^sub>w (g ` B)" given complete_lattice_class.SUP_subset_mono proof- assume "\ B \ U\<^sub>2; \x\U\<^sub>2. f x \ U; \x\U\<^sub>2. g x \ U; A \ B; \x. \x \ U\<^sub>2; x \ A\ \ f x \\<^sub>o\<^sub>w g x \ \ \\<^sub>o\<^sub>w (f ` A) \\<^sub>o\<^sub>w \\<^sub>o\<^sub>w (g ` B)" for B :: "'b set" and f g A then have "\ B \ U\<^sub>2; \x\U\<^sub>2. f x \ U; \x\U\<^sub>2. g x \ U; A \ B; \x. x \ A \ f x \\<^sub>o\<^sub>w g x \ \ \\<^sub>o\<^sub>w (f ` A) \\<^sub>o\<^sub>w \\<^sub>o\<^sub>w (g ` B)" by auto then show "\\<^sub>o\<^sub>w (f ` A) \\<^sub>o\<^sub>w \\<^sub>o\<^sub>w (g ` B)" using assms by blast qed tts_lemma INF_superset_mono: assumes "A \ U\<^sub>2" and "\x\U\<^sub>2. f x \ U" and "\x\U\<^sub>2. g x \ U" and "B \ A" and "\x. \x \ U\<^sub>2; x \ B\ \ f x \\<^sub>o\<^sub>w g x" shows "\\<^sub>o\<^sub>w (f ` A) \\<^sub>o\<^sub>w \\<^sub>o\<^sub>w (g ` B)" given complete_lattice_class.INF_superset_mono proof- assume "\ A \ U\<^sub>2; \x\U\<^sub>2. f x \ U; \x\U\<^sub>2. g x \ U; B \ A; \x. \x \ U\<^sub>2; x \ B\ \ f x \\<^sub>o\<^sub>w g x \ \ \\<^sub>o\<^sub>w (f ` A) \\<^sub>o\<^sub>w \\<^sub>o\<^sub>w (g ` B)" for A :: "'b set" and f g B then have "\ A \ U\<^sub>2; \x\U\<^sub>2. f x \ U; \x\U\<^sub>2. g x \ U; B \ A; \x. x \ B \ f x \\<^sub>o\<^sub>w g x \ \ \\<^sub>o\<^sub>w (f ` A) \\<^sub>o\<^sub>w \\<^sub>o\<^sub>w (g ` B)" by auto then show "\\<^sub>o\<^sub>w (f ` A) \\<^sub>o\<^sub>w \\<^sub>o\<^sub>w (g ` B)" using assms by blast qed tts_lemma INF_absorb: assumes "I \ U\<^sub>2" and "\x\U\<^sub>2. A x \ U" and "k \ I" shows "A k \\<^sub>o\<^sub>w \\<^sub>o\<^sub>w (A ` I) = \\<^sub>o\<^sub>w (A ` I)" is complete_lattice_class.INF_absorb. tts_lemma SUP_absorb: assumes "I \ U\<^sub>2" and "\x\U\<^sub>2. A x \ U" and "k \ I" shows "A k \\<^sub>o\<^sub>w \\<^sub>o\<^sub>w (A ` I) = \\<^sub>o\<^sub>w (A ` I)" is complete_lattice_class.SUP_absorb. end end context fixes f :: "'b \ 'al" and U\<^sub>2 :: "'b set" assumes f[transfer_rule]: "\x \ U\<^sub>2. f x = \\<^sub>o\<^sub>w" begin tts_context tts: (?'a to U) and (?'b to \U\<^sub>2::'b set\) sbterms: (\Orderings.bot::?'a::complete_lattice\ to \\\<^sub>o\<^sub>w\) rewriting ctr_simps substituting complete_lattice_ow_axioms and sup_bot.sl_neut.not_empty eliminating through simp begin tts_lemma SUP_bot: assumes "A \ U\<^sub>2" shows "\\<^sub>o\<^sub>w (f ` A) = \\<^sub>o\<^sub>w" is complete_lattice_class.SUP_bot[folded const_fun_def]. end end context fixes f :: "'b \ 'al" and U\<^sub>2 :: "'b set" assumes f[transfer_rule]: "\x \ U\<^sub>2. f x = \\<^sub>o\<^sub>w" begin tts_context tts: (?'a to U) and (?'b to \U\<^sub>2::'b set\) sbterms: (\Orderings.top::?'a::complete_lattice\ to \\\<^sub>o\<^sub>w\) rewriting ctr_simps substituting complete_lattice_ow_axioms and sup_bot.sl_neut.not_empty eliminating through simp begin tts_lemma INF_top: assumes "A \ U\<^sub>2" shows "\\<^sub>o\<^sub>w (f ` A) = \\<^sub>o\<^sub>w" is complete_lattice_class.INF_top[folded const_fun_def]. end end context fixes f :: "'b \ 'al" and c and F and U\<^sub>2 :: "'b set" assumes c_closed[transfer_rule]: "c \ U" assumes f[transfer_rule]: "\x \ U\<^sub>2. f x = c" begin tts_register_sbts \c\ | U proof- assume ALC: "Domainp ALC = (\x. x \ U)" have "Domainp ALC c" unfolding ALC by (simp add: c_closed) then obtain c' where "ALC c c'" by blast then show "\c'. ALC c c'" by auto qed tts_context tts: (?'a to U) and (?'b to \U\<^sub>2::'b set\) sbterms: (\?c::?'a::complete_lattice\ to \c\) rewriting ctr_simps substituting complete_lattice_ow_axioms and sup_bot.sl_neut.not_empty eliminating through simp begin tts_lemma INF_constant: assumes "A \ U\<^sub>2" shows "\\<^sub>o\<^sub>w (f ` A) = (if A = {} then \\<^sub>o\<^sub>w else c)" is complete_lattice_class.INF_constant[folded const_fun_def]. tts_lemma SUP_constant: assumes "A \ U\<^sub>2" shows "\\<^sub>o\<^sub>w (f ` A) = (if A = {} then \\<^sub>o\<^sub>w else c)" is complete_lattice_class.SUP_constant[folded const_fun_def]. end tts_context tts: (?'a to U) and (?'b to \U\<^sub>2::'b set\) sbterms: (\?f::?'a::complete_lattice\ to \c\) rewriting ctr_simps substituting complete_lattice_ow_axioms and sup_bot.sl_neut.not_empty eliminating through simp begin tts_lemma INF_const: assumes "A \ U\<^sub>2" and "A \ {}" shows "\\<^sub>o\<^sub>w ((\i. c) ` A) = c" is complete_lattice_class.INF_const. tts_lemma SUP_const: assumes "A \ U\<^sub>2" and "A \ {}" shows "\\<^sub>o\<^sub>w ((\i. c) ` A) = c" is complete_lattice_class.SUP_const. end end end context complete_lattice_ow begin tts_context tts: (?'a to U) and (?'b to \U\<^sub>2::'b set\) and (?'c to \U\<^sub>3::'c set\) rewriting ctr_simps substituting complete_lattice_ow_axioms and sup_bot.sl_neut.not_empty eliminating \?U \ {}\ through (force simp: subset_iff INF_top equals0D) applying [ OF Inf_closed' Sup_closed' inf_closed' sup_closed' bot_closed top_closed ] begin tts_lemma SUP_eq: assumes "A \ U\<^sub>2" and "B \ U\<^sub>3" and "\x\U\<^sub>2. f (x::'a) \ U" and "\x\U\<^sub>3. g x \ U" and "\i. i \ A \ \x\B. f i \\<^sub>o\<^sub>w g x" and "\j. j \ B \ \x\A. g j \\<^sub>o\<^sub>w f x" shows "\\<^sub>o\<^sub>w (f ` A) = \\<^sub>o\<^sub>w (g ` B)" given complete_lattice_class.SUP_eq proof- assume "\ A \ U\<^sub>2; B \ U\<^sub>3; \x\U\<^sub>2. f x \ U; \x\U\<^sub>3. g x \ U; \i. \i \ U\<^sub>2; i \ A\ \ \x\B. f i \\<^sub>o\<^sub>w g x; \j. \j \ U\<^sub>3; j \ B\ \ \x\A. g j \\<^sub>o\<^sub>w f x \ \ \\<^sub>o\<^sub>w (f ` A) = \\<^sub>o\<^sub>w (g ` B)" for A :: "'a set" and U\<^sub>2 and B :: "'b set" and U\<^sub>3 f g then have "\ A \ U\<^sub>2; B \ U\<^sub>3; \x\U\<^sub>2. f x \ U; \x\U\<^sub>3. g x \ U; \i. i \ A \ \x\B. f i \\<^sub>o\<^sub>w g x; \j. j \ B \ \x\A. g j \\<^sub>o\<^sub>w f x \ \ \\<^sub>o\<^sub>w (f ` A) = \\<^sub>o\<^sub>w (g ` B)" by simp then show "\\<^sub>o\<^sub>w (f ` A) = \\<^sub>o\<^sub>w (g ` B)" using assms by simp qed tts_lemma INF_eq: assumes "A \ U\<^sub>2" and "B \ U\<^sub>3" and "\x\U\<^sub>3. g x \ U" and "\x\U\<^sub>2. f (x::'a) \ U" and "\i. i \ A \ \x\B. g x \\<^sub>o\<^sub>w f i" and "\j. j \ B \ \x\A. f x \\<^sub>o\<^sub>w g j" shows "\\<^sub>o\<^sub>w (f ` A) = \\<^sub>o\<^sub>w (g ` B)" given complete_lattice_class.INF_eq proof- assume "\ A \ U\<^sub>2; B \ U\<^sub>3; \x\U\<^sub>3. g x \ U; \x\U\<^sub>2. f x \ U; \i. \i \ U\<^sub>2; i \ A\ \ \x\B. g x \\<^sub>o\<^sub>w f i; \j. \j \ U\<^sub>3; j \ B\ \ \x\A. f x \\<^sub>o\<^sub>w g j \ \ \\<^sub>o\<^sub>w (f ` A) = \\<^sub>o\<^sub>w (g ` B)" for A :: "'a set" and U\<^sub>2 and B :: "'b set" and U\<^sub>3 g f then have "\ A \ U\<^sub>2; B \ U\<^sub>3; \x\U\<^sub>3. g x \ U; \x\U\<^sub>2. f x \ U; \i. i \ A \ \x\B. g x \\<^sub>o\<^sub>w f i; \j. j \ B \ \x\A. f x \\<^sub>o\<^sub>w g j \ \ \\<^sub>o\<^sub>w (f ` A) = \\<^sub>o\<^sub>w (g ` B)" by simp then show "\\<^sub>o\<^sub>w (f ` A) = \\<^sub>o\<^sub>w (g ` B)" using assms by simp qed end end context complete_lattice_ow begin context fixes U\<^sub>2 :: "'b set" and U\<^sub>3 :: "'c set" begin lemmas [transfer_rule] = image_transfer[where ?'a='b] image_transfer[where ?'a='c] tts_context tts: (?'a to U) and (?'b to \U\<^sub>2::'b set\) and (?'c to \U\<^sub>3::'c set\) rewriting ctr_simps substituting complete_lattice_ow_axioms and sup_bot.sl_neut.not_empty applying [ OF _ _ Inf_closed' Sup_closed' inf_closed' sup_closed' bot_closed top_closed ] begin tts_lemma ne_INF_commute: assumes "U\<^sub>2 \ {}" and "U\<^sub>3 \ {}" and "\x\U\<^sub>2. \y\U\<^sub>3. f (x::'b) y \ U" and "B \ U\<^sub>3" and "A \ U\<^sub>2" shows "\\<^sub>o\<^sub>w ((\i. \\<^sub>o\<^sub>w (f i ` B)) ` A) = \\<^sub>o\<^sub>w ((\j. \\<^sub>o\<^sub>w ((\i. f i j) ` A)) ` B)" is complete_lattice_class.INF_commute. tts_lemma ne_SUP_commute: assumes "U\<^sub>2 \ {}" and "U\<^sub>3 \ {}" and "\x\U\<^sub>2. \y\U\<^sub>3. f (x::'b) y \ U" and "B \ U\<^sub>3" and "A \ U\<^sub>2" shows "\\<^sub>o\<^sub>w ((\i. \\<^sub>o\<^sub>w (f i ` B)) ` A) = \\<^sub>o\<^sub>w ((\j. \\<^sub>o\<^sub>w ((\i. f i j) ` A)) ` B)" is complete_lattice_class.SUP_commute. tts_lemma ne_SUP_mono: assumes "U\<^sub>2 \ {}" and "U\<^sub>3 \ {}" and "A \ U\<^sub>2" and "B \ U\<^sub>3" and "\x\U\<^sub>2. f (x::'b) \ U" and "\x\U\<^sub>3. g x \ U" and "\n. \n \ U\<^sub>2; n \ A\ \ \x\B. f n \\<^sub>o\<^sub>w g x" shows "\\<^sub>o\<^sub>w (f ` A) \\<^sub>o\<^sub>w \\<^sub>o\<^sub>w (g ` B)" is complete_lattice_class.SUP_mono. tts_lemma ne_INF_mono: assumes "U\<^sub>2 \ {}" and "U\<^sub>3 \ {}" and "B \ U\<^sub>2" and "A \ U\<^sub>3" and "\x\U\<^sub>3. f x \ U" and "\x\U\<^sub>2. g (x::'b) \ U" and "\m. \m \ U\<^sub>2; m \ B\ \ \x\A. f x \\<^sub>o\<^sub>w g m" shows "\\<^sub>o\<^sub>w (f ` A) \\<^sub>o\<^sub>w \\<^sub>o\<^sub>w (g ` B)" is complete_lattice_class.INF_mono. end end lemma INF_commute: assumes "\x\U\<^sub>2. \y\U\<^sub>3. f x y \ U" and "B \ U\<^sub>3" and "A \ U\<^sub>2" shows "\\<^sub>o\<^sub>w ((\x. \\<^sub>o\<^sub>w (f x ` B)) ` A) = \\<^sub>o\<^sub>w ((\j. \\<^sub>o\<^sub>w ((\i. f i j) ` A)) ` B)" proof(cases \U\<^sub>2 = {}\) case True then show ?thesis using assms by (simp add: inf_top.sl_neut.neutral_map Inf_top_conv(2)) next case ne_U\<^sub>2: False show ?thesis proof(cases \U\<^sub>3 = {}\) case True show ?thesis proof- from assms(2) have "B = {}" unfolding True by simp from assms(1) show ?thesis unfolding \B = {}\ by (force intro: INF_top) qed next case False from ne_U\<^sub>2 False assms show ?thesis by (rule ne_INF_commute) qed qed lemma SUP_commute: assumes "\x\U\<^sub>2. \y\U\<^sub>3. f x y \ U" and "B \ U\<^sub>3" and "A \ U\<^sub>2" shows "\\<^sub>o\<^sub>w ((\x. \\<^sub>o\<^sub>w (f x ` B)) ` A) = \\<^sub>o\<^sub>w ((\j. \\<^sub>o\<^sub>w ((\i. f i j) ` A)) ` B)" proof(cases \U\<^sub>2 = {}\) case True show ?thesis proof- from assms(3) have "A = {}" unfolding True by simp from assms(2) show ?thesis unfolding \A = {}\ by (simp add: sup_bot.sl_neut.neutral_map inf_absorb1 SUP_bot) qed next case ne_U\<^sub>2: False show ?thesis proof(cases \U\<^sub>3 = {}\) case True show ?thesis proof- from assms(2) have "B = {}" unfolding True by simp from assms(1) show ?thesis unfolding \B = {}\ by (simp add: sup_bot.sl_neut.neutral_map Sup_bot_conv(1)) qed next case False from ne_U\<^sub>2 False assms show ?thesis by (rule ne_SUP_commute) qed qed lemma SUP_mono: assumes "A \ U\<^sub>2" and "B \ U\<^sub>3" and "\x\U\<^sub>2. f x \ U" and "\x\U\<^sub>3. g x \ U" and "\n. n \ A \ \m\B. f n \\<^sub>o\<^sub>w g m" shows "\\<^sub>o\<^sub>w (f ` A) \\<^sub>o\<^sub>w \\<^sub>o\<^sub>w (g ` B)" proof(cases \U\<^sub>2 = {}\) case True show ?thesis proof- from assms(1) have "A = {}" unfolding True by simp have "f ` A = {}" unfolding \A = {}\ by simp from assms(2,4) show ?thesis unfolding \f ` A = {}\ by (simp add: image_subset_iff in_mono) qed next case ne_U\<^sub>2: False show ?thesis proof(cases \U\<^sub>3 = {}\) case True show ?thesis proof- from assms(2) have "B = {}" unfolding True by simp have "g ` B = {}" unfolding \B = {}\ by simp from assms(1,3,5) show ?thesis unfolding \g ` B = {}\ \B = {}\ by (force simp: subset_iff) qed next case False from ne_U\<^sub>2 False assms show ?thesis by (rule ne_SUP_mono) qed qed lemma INF_mono: assumes "B \ U\<^sub>2" and "A \ U\<^sub>3" and "\x\U\<^sub>3. f x \ U" and "\x\U\<^sub>2. g x \ U" and "\m. m \ B \ \n\A. f n \\<^sub>o\<^sub>w g m" shows "\\<^sub>o\<^sub>w (f ` A) \\<^sub>o\<^sub>w \\<^sub>o\<^sub>w (g ` B)" proof(cases \U\<^sub>2 = {}\) case True show ?thesis proof- from assms(1) have "B = {}" unfolding True by simp have "g ` B = {}" unfolding \B = {}\ by simp from assms(2,3) show ?thesis unfolding \g ` B = {}\ by (simp add: image_subset_iff in_mono) qed next case ne_U\<^sub>2: False show ?thesis proof(cases \U\<^sub>3 = {}\) case True show ?thesis proof- from assms(2) have "A = {}" unfolding True by simp have "f ` A = {}" unfolding \A = {}\ by simp from assms show ?thesis unfolding \f ` A = {}\ \A = {}\ by (auto simp: subset_iff) fastforce qed next case False from ne_U\<^sub>2 False assms show ?thesis by (rule ne_INF_mono) qed qed end context complete_lattice_pair_ow begin -tts_context - tts: (?'a to U\<^sub>1) and (?'b to U\<^sub>2) - rewriting ctr_simps - substituting cl\<^sub>1.complete_lattice_ow_axioms - and cl\<^sub>2.complete_lattice_ow_axioms - and cl\<^sub>1.inf_top.sl_neut.not_empty - and cl\<^sub>2.inf_top.sl_neut.not_empty - applying - [ - OF - cl\<^sub>1.Inf_closed' - cl\<^sub>1.Sup_closed' - cl\<^sub>1.inf_closed' - cl\<^sub>1.sup_closed' - cl\<^sub>1.bot_closed - cl\<^sub>1.top_closed - cl\<^sub>2.Inf_closed' - cl\<^sub>2.Sup_closed' - cl\<^sub>2.inf_closed' - cl\<^sub>2.sup_closed' - cl\<^sub>2.bot_closed - cl\<^sub>2.top_closed - ] -begin - -tts_lemma mono_Inf: - assumes "\x\U\<^sub>1. f x \ U\<^sub>2" - and "A \ U\<^sub>1" - and "on U\<^sub>1 with (\\<^sub>o\<^sub>w\<^sub>.\<^sub>2) (\\<^sub>o\<^sub>w\<^sub>.\<^sub>1) : \mono\ f" - shows "f (\\<^sub>o\<^sub>w\<^sub>.\<^sub>1 A) \\<^sub>o\<^sub>w\<^sub>.\<^sub>2 \\<^sub>o\<^sub>w\<^sub>.\<^sub>2 (f ` A)" - is complete_lattice_class.mono_Inf. - -tts_lemma mono_Sup: - assumes "\x\U\<^sub>1. f x \ U\<^sub>2" - and "A \ U\<^sub>1" - and "on U\<^sub>1 with (\\<^sub>o\<^sub>w\<^sub>.\<^sub>2) (\\<^sub>o\<^sub>w\<^sub>.\<^sub>1) : \mono\ f" - shows "\\<^sub>o\<^sub>w\<^sub>.\<^sub>2 (f ` A) \\<^sub>o\<^sub>w\<^sub>.\<^sub>2 f (\\<^sub>o\<^sub>w\<^sub>.\<^sub>1 A)" - is complete_lattice_class.mono_Sup. - -end - context fixes U\<^sub>3 :: "'c set" begin lemmas [transfer_rule] = image_transfer[where ?'a='c] -tts_context - tts: (?'a to U\<^sub>1) and (?'b to U\<^sub>2) and (?'c to \U\<^sub>3::'c set\) - rewriting ctr_simps - substituting cl\<^sub>1.complete_lattice_ow_axioms - and cl\<^sub>2.complete_lattice_ow_axioms - and cl\<^sub>1.inf_top.sl_neut.not_empty - and cl\<^sub>2.inf_top.sl_neut.not_empty - eliminating through (simp add: image_subset_iff image_subset_iff') - applying - [ - OF - _ - cl\<^sub>1.Inf_closed' - cl\<^sub>1.Sup_closed' - cl\<^sub>1.inf_closed' - cl\<^sub>1.sup_closed' - cl\<^sub>1.bot_closed - cl\<^sub>1.top_closed - cl\<^sub>2.Inf_closed' - cl\<^sub>2.Sup_closed' - cl\<^sub>2.inf_closed' - cl\<^sub>2.sup_closed' - cl\<^sub>2.bot_closed - cl\<^sub>2.top_closed - ] -begin - -tts_lemma mono_SUP: - assumes "U\<^sub>3 \ {}" - and "\x\U\<^sub>1. f x \ U\<^sub>2" - and "\x\U\<^sub>3. A x \ U\<^sub>1" - and "I \ U\<^sub>3" - and "on U\<^sub>1 with (\\<^sub>o\<^sub>w\<^sub>.\<^sub>2) (\\<^sub>o\<^sub>w\<^sub>.\<^sub>1) : \mono\ f" - shows "\\<^sub>o\<^sub>w\<^sub>.\<^sub>2 ((\x. f (A x)) ` I) \\<^sub>o\<^sub>w\<^sub>.\<^sub>2 f (\\<^sub>o\<^sub>w\<^sub>.\<^sub>1 (A ` I))" - is complete_lattice_class.mono_SUP. - -tts_lemma mono_INF: - assumes "U\<^sub>3 \ {}" - and "\x\U\<^sub>1. f x \ U\<^sub>2" - and "\x\U\<^sub>3. A x \ U\<^sub>1" - and "I \ U\<^sub>3" - and "on U\<^sub>1 with (\\<^sub>o\<^sub>w\<^sub>.\<^sub>2) (\\<^sub>o\<^sub>w\<^sub>.\<^sub>1) : \mono\ f" - shows "f (\\<^sub>o\<^sub>w\<^sub>.\<^sub>1 (A ` I)) \\<^sub>o\<^sub>w\<^sub>.\<^sub>2 \\<^sub>o\<^sub>w\<^sub>.\<^sub>2 ((\x. f (A x)) ` I)" - is complete_lattice_class.mono_INF. - -end - end end text\\newpage\ end \ No newline at end of file diff --git a/thys/Types_To_Sets_Extension/Examples/SML_Relativization/Lattices/SML_Linorders.thy b/thys/Types_To_Sets_Extension/Examples/SML_Relativization/Lattices/SML_Linorders.thy --- a/thys/Types_To_Sets_Extension/Examples/SML_Relativization/Lattices/SML_Linorders.thy +++ b/thys/Types_To_Sets_Extension/Examples/SML_Relativization/Lattices/SML_Linorders.thy @@ -1,569 +1,472 @@ (* Title: Examples/SML_Relativization/Lattices/SML_Linorders.thy Author: Mihails Milehins Copyright 2021 (C) Mihails Milehins *) section\Relativization of the results about linear orders\ theory SML_Linorders imports SML_Semilattices begin subsection\Linear orders\ subsubsection\Definitions and further properties\ locale linorder_ow = order_ow + assumes linear: "\ x \ U; y \ U \ \ x \\<^sub>o\<^sub>w y \ y \\<^sub>o\<^sub>w x" begin sublocale min: semilattice_order_ow U \(\x y. (with (\\<^sub>o\<^sub>w) : \min\ x y))\ \(\\<^sub>o\<^sub>w)\ \(<\<^sub>o\<^sub>w)\ apply unfold_locales subgoal unfolding min.with_def by simp subgoal by (metis linear order_trans min.with_def) subgoal unfolding min.with_def by (auto dest: linear simp: antisym) subgoal unfolding min.with_def by simp subgoal unfolding min.with_def by (simp add: eq_iff) subgoal unfolding min.with_def by (simp add: less_le) done sublocale max: semilattice_order_ow U \(\x y. (with (\\<^sub>o\<^sub>w) : \max\ x y))\ \(\\<^sub>o\<^sub>w)\ \(>\<^sub>o\<^sub>w)\ apply unfold_locales subgoal unfolding max.with_def by simp subgoal by (metis linear order_trans max.with_def) subgoal unfolding max.with_def by (auto dest: linear simp: antisym) subgoal unfolding max.with_def by simp subgoal unfolding max.with_def by (auto dest: linear simp: antisym) subgoal unfolding max.with_def by (auto dest: linear simp: less_le_not_le) done end locale ord_linorder_ow = ord?: ord_ow U\<^sub>1 le\<^sub>1 ls\<^sub>1 + lo?: linorder_ow U\<^sub>2 le\<^sub>2 ls\<^sub>2 for U\<^sub>1 :: "'ao set" and le\<^sub>1 ls\<^sub>1 and U\<^sub>2 :: "'bo set" and le\<^sub>2 ls\<^sub>2 begin sublocale ord_order_ow .. end locale preorder_linorder_ow = po?: preorder_ow U\<^sub>1 le\<^sub>1 ls\<^sub>1 + lo?: linorder_ow U\<^sub>2 le\<^sub>2 ls\<^sub>2 for U\<^sub>1 :: "'ao set" and le\<^sub>1 ls\<^sub>1 and U\<^sub>2 :: "'bo set" and le\<^sub>2 ls\<^sub>2 begin sublocale preorder_order_ow .. end locale order_linorder_ow = order?: order_ow U\<^sub>1 le\<^sub>1 ls\<^sub>1 + lo?: linorder_ow U\<^sub>2 le\<^sub>2 ls\<^sub>2 for U\<^sub>1 :: "'ao set" and le\<^sub>1 ls\<^sub>1 and U\<^sub>2 :: "'bo set" and le\<^sub>2 ls\<^sub>2 begin sublocale order_pair_ow .. end locale linorder_pair_ow = lo\<^sub>1?: linorder_ow U\<^sub>1 le\<^sub>1 ls\<^sub>1 + lo\<^sub>2?: linorder_ow U\<^sub>2 le\<^sub>2 ls\<^sub>2 for U\<^sub>1 :: "'ao set" and le\<^sub>1 ls\<^sub>1 and U\<^sub>2 :: "'bo set" and le\<^sub>2 ls\<^sub>2 begin sublocale order_linorder_ow .. end subsubsection\Transfer rules\ context includes lifting_syntax begin lemma linorder_transfer[transfer_rule]: assumes [transfer_rule]: "bi_unique A" "right_total A" shows "((A ===> A ===> (=)) ===> (A ===> A ===> (=)) ===> (=)) (linorder_ow (Collect (Domainp A))) class.linorder" unfolding linorder_ow_def class.linorder_def linorder_ow_axioms_def class.linorder_axioms_def apply transfer_prover_start apply transfer_step+ by simp end subsubsection\Relativization\ context linorder_ow begin tts_context tts: (?'a to U) rewriting ctr_simps substituting linorder_ow_axioms eliminating through simp begin tts_lemma le_less_linear: assumes "x \ U" and "y \ U" shows "x \\<^sub>o\<^sub>w y \ y <\<^sub>o\<^sub>w x" is linorder_class.le_less_linear. tts_lemma not_less: assumes "x \ U" and "y \ U" shows "(\ x <\<^sub>o\<^sub>w y) = (y \\<^sub>o\<^sub>w x)" is linorder_class.not_less. tts_lemma not_le: assumes "x \ U" and "y \ U" shows "(\ x \\<^sub>o\<^sub>w y) = (y <\<^sub>o\<^sub>w x)" is linorder_class.not_le. tts_lemma lessThan_minus_lessThan: assumes "n \ U" and "m \ U" shows "{..<\<^sub>o\<^sub>wn} - {..<\<^sub>o\<^sub>wm} = (on U with (\\<^sub>o\<^sub>w) (<\<^sub>o\<^sub>w) : {m.. U" and "b \ U" shows "({a..\<^sub>o\<^sub>w} \ {b<\<^sub>o\<^sub>w..}) = (b <\<^sub>o\<^sub>w a)" is linorder_class.Ici_subset_Ioi_iff. tts_lemma Iic_subset_Iio_iff: assumes "a \ U" and "b \ U" shows "({..\<^sub>o\<^sub>wa} \ {..<\<^sub>o\<^sub>wb}) = (a <\<^sub>o\<^sub>w b)" is linorder_class.Iic_subset_Iio_iff. tts_lemma leI: assumes "x \ U" and "y \ U" and "\ x <\<^sub>o\<^sub>w y" shows "y \\<^sub>o\<^sub>w x" is linorder_class.leI. tts_lemma not_le_imp_less: assumes "y \ U" and "x \ U" and "\ y \\<^sub>o\<^sub>w x" shows "x <\<^sub>o\<^sub>w y" is linorder_class.not_le_imp_less. tts_lemma Int_atMost: assumes "a \ U" and "b \ U" shows "{..\<^sub>o\<^sub>wa} \ {..\<^sub>o\<^sub>wb} = {..\<^sub>o\<^sub>wmin a b}" is linorder_class.Int_atMost. tts_lemma lessThan_Int_lessThan: assumes "a \ U" and "b \ U" shows "{a<\<^sub>o\<^sub>w..} \ {b<\<^sub>o\<^sub>w..} = {max a b<\<^sub>o\<^sub>w..}" is linorder_class.lessThan_Int_lessThan. tts_lemma greaterThan_Int_greaterThan: assumes "a \ U" and "b \ U" shows "{..<\<^sub>o\<^sub>wa} \ {..<\<^sub>o\<^sub>wb} = {..<\<^sub>o\<^sub>wmin a b}" is linorder_class.greaterThan_Int_greaterThan. tts_lemma less_linear: assumes "x \ U" and "y \ U" shows "x <\<^sub>o\<^sub>w y \ x = y \ y <\<^sub>o\<^sub>w x" is linorder_class.less_linear. tts_lemma Int_atLeastAtMostR2: assumes "a \ U" and "c \ U" and "d \ U" shows "{a..\<^sub>o\<^sub>w} \ {c..\<^sub>o\<^sub>wd} = {max a c..\<^sub>o\<^sub>wd}" is linorder_class.Int_atLeastAtMostR2. tts_lemma Int_atLeastAtMostR1: assumes "b \ U" and "c \ U" and "d \ U" shows "{..\<^sub>o\<^sub>wb} \ {c..\<^sub>o\<^sub>wd} = {c..\<^sub>o\<^sub>wmin b d}" is linorder_class.Int_atLeastAtMostR1. tts_lemma Int_atLeastAtMostL2: assumes "a \ U" and "b \ U" and "c \ U" shows "{a..\<^sub>o\<^sub>wb} \ {c..\<^sub>o\<^sub>w} = {max a c..\<^sub>o\<^sub>wb}" is linorder_class.Int_atLeastAtMostL2. tts_lemma Int_atLeastAtMostL1: assumes "a \ U" and "b \ U" and "d \ U" shows "{a..\<^sub>o\<^sub>wb} \ {..\<^sub>o\<^sub>wd} = {a..\<^sub>o\<^sub>wmin b d}" is linorder_class.Int_atLeastAtMostL1. tts_lemma neq_iff: assumes "x \ U" and "y \ U" shows "(x \ y) = (x <\<^sub>o\<^sub>w y \ y <\<^sub>o\<^sub>w x)" is linorder_class.neq_iff. tts_lemma not_less_iff_gr_or_eq: assumes "x \ U" and "y \ U" shows "(\ x <\<^sub>o\<^sub>w y) = (y <\<^sub>o\<^sub>w x \ x = y)" is linorder_class.not_less_iff_gr_or_eq. tts_lemma max_min_distrib2: assumes "a \ U" and "b \ U" and "c \ U" shows "max a (min b c) = min (max a b) (max a c)" is linorder_class.max_min_distrib2. tts_lemma max_min_distrib1: assumes "b \ U" and "c \ U" and "a \ U" shows "max (min b c) a = min (max b a) (max c a)" is linorder_class.max_min_distrib1. tts_lemma min_max_distrib2: assumes "a \ U" and "b \ U" and "c \ U" shows "min a (max b c) = max (min a b) (min a c)" is linorder_class.min_max_distrib2. tts_lemma min_max_distrib1: assumes "b \ U" and "c \ U" and "a \ U" shows "min (max b c) a = max (min b a) (min c a)" is linorder_class.min_max_distrib1. tts_lemma atLeastAtMost_diff_ends: assumes "a \ U" and "b \ U" shows "{a..\<^sub>o\<^sub>wb} - {a, b} = {a<\<^sub>o\<^sub>w..<\<^sub>o\<^sub>wb}" is linorder_class.atLeastAtMost_diff_ends. tts_lemma less_max_iff_disj: assumes "z \ U" and "x \ U" and "y \ U" shows "(z <\<^sub>o\<^sub>w max x y) = (z <\<^sub>o\<^sub>w x \ z <\<^sub>o\<^sub>w y)" is linorder_class.less_max_iff_disj. tts_lemma min_less_iff_conj: assumes "z \ U" and "x \ U" and "y \ U" shows "(z <\<^sub>o\<^sub>w min x y) = (z <\<^sub>o\<^sub>w x \ z <\<^sub>o\<^sub>w y)" is linorder_class.min_less_iff_conj. tts_lemma max_less_iff_conj: assumes "x \ U" and "y \ U" and "z \ U" shows "(max x y <\<^sub>o\<^sub>w z) = (x <\<^sub>o\<^sub>w z \ y <\<^sub>o\<^sub>w z)" is linorder_class.max_less_iff_conj. tts_lemma min_less_iff_disj: assumes "x \ U" and "y \ U" and "z \ U" shows "(min x y <\<^sub>o\<^sub>w z) = (x <\<^sub>o\<^sub>w z \ y <\<^sub>o\<^sub>w z)" is linorder_class.min_less_iff_disj. tts_lemma le_max_iff_disj: assumes "z \ U" and "x \ U" and "y \ U" shows "(z \\<^sub>o\<^sub>w max x y) = (z \\<^sub>o\<^sub>w x \ z \\<^sub>o\<^sub>w y)" is linorder_class.le_max_iff_disj. tts_lemma min_le_iff_disj: assumes "x \ U" and "y \ U" and "z \ U" shows "(min x y \\<^sub>o\<^sub>w z) = (x \\<^sub>o\<^sub>w z \ y \\<^sub>o\<^sub>w z)" is linorder_class.min_le_iff_disj. tts_lemma antisym_conv3: assumes "y \ U" and "x \ U" and "\ y <\<^sub>o\<^sub>w x" shows "(\ x <\<^sub>o\<^sub>w y) = (x = y)" is linorder_class.antisym_conv3. tts_lemma Int_atLeastAtMost: assumes "a \ U" and "b \ U" and "c \ U" and "d \ U" shows "{a..\<^sub>o\<^sub>wb} \ {c..\<^sub>o\<^sub>wd} = {max a c..\<^sub>o\<^sub>wmin b d}" is linorder_class.Int_atLeastAtMost. tts_lemma Int_atLeastLessThan: assumes "a \ U" and "b \ U" and "c \ U" and "d \ U" shows "(on U with (\\<^sub>o\<^sub>w) (<\<^sub>o\<^sub>w) : {a.. (on U with (\\<^sub>o\<^sub>w) (<\<^sub>o\<^sub>w) : {c..\<^sub>o\<^sub>w) (<\<^sub>o\<^sub>w) : {(max a c)..<(min b d)})" is linorder_class.Int_atLeastLessThan. tts_lemma Int_greaterThanAtMost: assumes "a \ U" and "b \ U" and "c \ U" and "d \ U" shows "{a<\<^sub>o\<^sub>w..b} \ {c<\<^sub>o\<^sub>w..d} = {max a c<\<^sub>o\<^sub>w..min b d}" is linorder_class.Int_greaterThanAtMost. tts_lemma Int_greaterThanLessThan: assumes "a \ U" and "b \ U" and "c \ U" and "d \ U" shows "{a<\<^sub>o\<^sub>w..<\<^sub>o\<^sub>wb} \ {c<\<^sub>o\<^sub>w..<\<^sub>o\<^sub>wd} = {max a c<\<^sub>o\<^sub>w..<\<^sub>o\<^sub>wmin b d}" is linorder_class.Int_greaterThanLessThan. tts_lemma le_cases: assumes "x \ U" and "y \ U" and "x \\<^sub>o\<^sub>w y \ P" and "y \\<^sub>o\<^sub>w x \ P" shows P is linorder_class.le_cases. tts_lemma split_max: assumes "i \ U" and "j \ U" shows "P (max i j) = ((i \\<^sub>o\<^sub>w j \ P j) \ (\ i \\<^sub>o\<^sub>w j \ P i))" is linorder_class.split_max. tts_lemma split_min: assumes "i \ U" and "j \ U" shows "P (min i j) = ((i \\<^sub>o\<^sub>w j \ P i) \ (\ i \\<^sub>o\<^sub>w j \ P j))" is linorder_class.split_min. tts_lemma Ioc_subset_iff: assumes "a \ U" and "b \ U" and "c \ U" and "d \ U" shows "({a<\<^sub>o\<^sub>w..b} \ {c<\<^sub>o\<^sub>w..d}) = (b \\<^sub>o\<^sub>w a \ b \\<^sub>o\<^sub>w d \ c \\<^sub>o\<^sub>w a)" is linorder_class.Ioc_subset_iff. tts_lemma atLeastLessThan_subset_iff: assumes "a \ U" and "b \ U" and "c \ U" and "d \ U" and "(on U with (\\<^sub>o\<^sub>w) (<\<^sub>o\<^sub>w) : {a.. (on U with (\\<^sub>o\<^sub>w) (<\<^sub>o\<^sub>w) : {c..\<^sub>o\<^sub>w a \ b \\<^sub>o\<^sub>w d \ c \\<^sub>o\<^sub>w a" is linorder_class.atLeastLessThan_subset_iff. tts_lemma Ioc_inj: assumes "a \ U" and "b \ U" and "c \ U" and "d \ U" shows "({a<\<^sub>o\<^sub>w..b} = {c<\<^sub>o\<^sub>w..d}) = (b \\<^sub>o\<^sub>w a \ d \\<^sub>o\<^sub>w c \ a = c \ b = d)" is linorder_class.Ioc_inj. tts_lemma neqE: assumes "x \ U" and "y \ U" and "x \ y" and "x <\<^sub>o\<^sub>w y \ R" and "y <\<^sub>o\<^sub>w x \ R" shows R is linorder_class.neqE. tts_lemma Ioc_disjoint: assumes "a \ U" and "b \ U" and "c \ U" and "d \ U" shows "({a<\<^sub>o\<^sub>w..b} \ {c<\<^sub>o\<^sub>w..d} = {}) = (b \\<^sub>o\<^sub>w a \ d \\<^sub>o\<^sub>w c \ b \\<^sub>o\<^sub>w c \ d \\<^sub>o\<^sub>w a)" is linorder_class.Ioc_disjoint. tts_lemma linorder_cases: assumes "x \ U" and "y \ U" and "x <\<^sub>o\<^sub>w y \ P" and "x = y \ P" and "y <\<^sub>o\<^sub>w x \ P" shows P is linorder_class.linorder_cases. tts_lemma le_cases3: assumes "x \ U" and "y \ U" and "z \ U" and "\x \\<^sub>o\<^sub>w y; y \\<^sub>o\<^sub>w z\ \ P" and "\y \\<^sub>o\<^sub>w x; x \\<^sub>o\<^sub>w z\ \ P" and "\x \\<^sub>o\<^sub>w z; z \\<^sub>o\<^sub>w y\ \ P" and "\z \\<^sub>o\<^sub>w y; y \\<^sub>o\<^sub>w x\ \ P" and "\y \\<^sub>o\<^sub>w z; z \\<^sub>o\<^sub>w x\ \ P" and "\z \\<^sub>o\<^sub>w x; x \\<^sub>o\<^sub>w y\ \ P" shows P is linorder_class.le_cases3. end end -context order_linorder_ow -begin - -tts_context - tts: (?'a to U\<^sub>2) and (?'b to U\<^sub>1) - rewriting ctr_simps - substituting order.order_ow_axioms and lo.linorder_ow_axioms - eliminating through simp -begin - - -tts_lemma strict_mono_imp_inj_on: - assumes "\x\U\<^sub>2. f x \ U\<^sub>1" - and "A \ U\<^sub>2" - and "on U\<^sub>2 with (<\<^sub>o\<^sub>w\<^sub>.\<^sub>1) (<\<^sub>o\<^sub>w\<^sub>.\<^sub>2) : \strict_mono\ f" - shows "inj_on f A" - is linorder_class.strict_mono_imp_inj_on. - -tts_lemma strict_mono_eq: - assumes "\x\U\<^sub>2. f x \ U\<^sub>1" - and "x \ U\<^sub>2" - and "y \ U\<^sub>2" - and "on U\<^sub>2 with (<\<^sub>o\<^sub>w\<^sub>.\<^sub>1) (<\<^sub>o\<^sub>w\<^sub>.\<^sub>2) : \strict_mono\ f" - shows "(f x = f y) = (x = y)" - is linorder_class.strict_mono_eq. - -tts_lemma strict_mono_less: - assumes "\x\U\<^sub>2. f x \ U\<^sub>1" - and "x \ U\<^sub>2" - and "y \ U\<^sub>2" - and "on U\<^sub>2 with (<\<^sub>o\<^sub>w\<^sub>.\<^sub>1) (<\<^sub>o\<^sub>w\<^sub>.\<^sub>2) : \strict_mono\ f" - shows "(f x <\<^sub>o\<^sub>w\<^sub>.\<^sub>1 f y) = (x <\<^sub>o\<^sub>w\<^sub>.\<^sub>2 y)" - is linorder_class.strict_mono_less. - -tts_lemma strict_mono_less_eq: - assumes "\x\U\<^sub>2. f x \ U\<^sub>1" - and "x \ U\<^sub>2" - and "y \ U\<^sub>2" - and "on U\<^sub>2 with (<\<^sub>o\<^sub>w\<^sub>.\<^sub>1) (<\<^sub>o\<^sub>w\<^sub>.\<^sub>2) : \strict_mono\ f" - shows "(f x \\<^sub>o\<^sub>w\<^sub>.\<^sub>1 f y) = (x \\<^sub>o\<^sub>w\<^sub>.\<^sub>2 y)" - is linorder_class.strict_mono_less_eq. - -tts_lemma mono_strict_invE: - assumes "\x\U\<^sub>2. f x \ U\<^sub>1" - and "x \ U\<^sub>2" - and "y \ U\<^sub>2" - and "on U\<^sub>2 with (\\<^sub>o\<^sub>w\<^sub>.\<^sub>1) (\\<^sub>o\<^sub>w\<^sub>.\<^sub>2) : \mono\ f" - and "f x <\<^sub>o\<^sub>w\<^sub>.\<^sub>1 f y" - and "x <\<^sub>o\<^sub>w\<^sub>.\<^sub>2 y \ thesis" - shows thesis - is linorder_class.mono_strict_invE. - -tts_lemma mono_invE: - assumes "\x\U\<^sub>2. f x \ U\<^sub>1" - and "x \ U\<^sub>2" - and "y \ U\<^sub>2" - and "on U\<^sub>2 with (\\<^sub>o\<^sub>w\<^sub>.\<^sub>1) (\\<^sub>o\<^sub>w\<^sub>.\<^sub>2) : \mono\ f" - and "f x <\<^sub>o\<^sub>w\<^sub>.\<^sub>1 f y" - and "x \\<^sub>o\<^sub>w\<^sub>.\<^sub>2 y \ thesis" - shows thesis - is linorder_class.mono_invE. - -end - -end - -context linorder_pair_ow -begin - -tts_context - tts: (?'a to U\<^sub>1) and (?'b to U\<^sub>2) - rewriting ctr_simps - substituting lo\<^sub>1.linorder_ow_axioms and lo\<^sub>2.linorder_ow_axioms - eliminating through simp -begin - -tts_lemma max_of_mono: - assumes "\x\U\<^sub>1. f x \ U\<^sub>2" - and "m \ U\<^sub>1" - and "n \ U\<^sub>1" - and "on U\<^sub>1 with (\\<^sub>o\<^sub>w\<^sub>.\<^sub>2) (\\<^sub>o\<^sub>w\<^sub>.\<^sub>1) : \mono\ f" - shows "lo\<^sub>2.max (f m) (f n) = f (lo\<^sub>1.max m n)" - is linorder_class.max_of_mono. - -tts_lemma min_of_mono: - assumes "\x\U\<^sub>1. f x \ U\<^sub>2" - and "m \ U\<^sub>1" - and "n \ U\<^sub>1" - and "on U\<^sub>1 with (\\<^sub>o\<^sub>w\<^sub>.\<^sub>2) (\\<^sub>o\<^sub>w\<^sub>.\<^sub>1) : \mono\ f" - shows "lo\<^sub>2.min (f m) (f n) = f (lo\<^sub>1.min m n)" - is linorder_class.min_of_mono. - -end - -end - - subsection\Dense linear orders\ subsubsection\Definitions and further properties\ locale dense_linorder_ow = linorder_ow U le ls + dense_order_ow U le ls for U :: "'ao set" and le (infix \\\<^sub>o\<^sub>w\ 50) and ls (infix \<\<^sub>o\<^sub>w\ 50) subsubsection\Transfer rules\ context includes lifting_syntax begin lemma dense_linorder_transfer[transfer_rule]: assumes [transfer_rule]: "bi_unique A" "right_total A" shows "((A ===> A ===> (=)) ===> (A ===> A ===> (=)) ===> (=)) (dense_linorder_ow (Collect (Domainp A))) class.dense_linorder" unfolding dense_linorder_ow_def class.dense_linorder_def apply transfer_prover_start apply transfer_step+ by auto end subsubsection\Relativization\ context dense_linorder_ow begin tts_context tts: (?'a to U) rewriting ctr_simps substituting dense_linorder_ow_axioms eliminating through simp begin tts_lemma infinite_Icc: assumes "a \ U" and "b \ U" and "a <\<^sub>o\<^sub>w b" shows "infinite {a..\<^sub>o\<^sub>wb}" is dense_linorder_class.infinite_Icc. tts_lemma infinite_Ico: assumes "a \ U" and "b \ U" and "a <\<^sub>o\<^sub>w b" shows "infinite (on U with (\\<^sub>o\<^sub>w) (<\<^sub>o\<^sub>w) : {a.. U" and "b \ U" and "a <\<^sub>o\<^sub>w b" shows "infinite {a<\<^sub>o\<^sub>w..b}" is dense_linorder_class.infinite_Ioc. tts_lemma infinite_Ioo: assumes "a \ U" and "b \ U" and "a <\<^sub>o\<^sub>w b" shows "infinite {a<\<^sub>o\<^sub>w..<\<^sub>o\<^sub>wb}" is dense_linorder_class.infinite_Ioo. tts_lemma atLeastLessThan_subseteq_atLeastAtMost_iff: assumes "a \ U" and "b \ U" and "c \ U" and "d \ U" shows "((on U with (\\<^sub>o\<^sub>w) (<\<^sub>o\<^sub>w) : {a.. {c..\<^sub>o\<^sub>wd}) = (a <\<^sub>o\<^sub>w b \ b \\<^sub>o\<^sub>w d \ c \\<^sub>o\<^sub>w a)" is dense_linorder_class.atLeastLessThan_subseteq_atLeastAtMost_iff. tts_lemma greaterThanAtMost_subseteq_atLeastAtMost_iff: assumes "a \ U" and "b \ U" and "c \ U" and "d \ U" shows "({a<\<^sub>o\<^sub>w..b} \ {c..\<^sub>o\<^sub>wd}) = (a <\<^sub>o\<^sub>w b \ b \\<^sub>o\<^sub>w d \ c \\<^sub>o\<^sub>w a)" is dense_linorder_class.greaterThanAtMost_subseteq_atLeastAtMost_iff. tts_lemma greaterThanAtMost_subseteq_atLeastLessThan_iff: assumes "a \ U" and "b \ U" and "c \ U" and "d \ U" shows "({a<\<^sub>o\<^sub>w..b} \ (on U with (\\<^sub>o\<^sub>w) (<\<^sub>o\<^sub>w) : {c..o\<^sub>w b \ b <\<^sub>o\<^sub>w d \ c \\<^sub>o\<^sub>w a)" is dense_linorder_class.greaterThanAtMost_subseteq_atLeastLessThan_iff. tts_lemma greaterThanLessThan_subseteq_atLeastAtMost_iff: assumes "a \ U" and "b \ U" and "c \ U" and "d \ U" shows "({a<\<^sub>o\<^sub>w..<\<^sub>o\<^sub>wb} \ {c..\<^sub>o\<^sub>wd}) = (a <\<^sub>o\<^sub>w b \ b \\<^sub>o\<^sub>w d \ c \\<^sub>o\<^sub>w a)" is dense_linorder_class.greaterThanLessThan_subseteq_atLeastAtMost_iff. tts_lemma greaterThanLessThan_subseteq_atLeastLessThan_iff: assumes "a \ U" and "b \ U" and "c \ U" and "d \ U" shows "({a<\<^sub>o\<^sub>w..<\<^sub>o\<^sub>wb} \ (on U with (\\<^sub>o\<^sub>w) (<\<^sub>o\<^sub>w) : {c..o\<^sub>w b \ b \\<^sub>o\<^sub>w d \ c \\<^sub>o\<^sub>w a)" is dense_linorder_class.greaterThanLessThan_subseteq_atLeastLessThan_iff. tts_lemma greaterThanLessThan_subseteq_greaterThanAtMost_iff: assumes "a \ U" and "b \ U" and "c \ U" and "d \ U" shows "({a<\<^sub>o\<^sub>w..<\<^sub>o\<^sub>wb} \ {c<\<^sub>o\<^sub>w..d}) = (a <\<^sub>o\<^sub>w b \ b \\<^sub>o\<^sub>w d \ c \\<^sub>o\<^sub>w a)" is dense_linorder_class.greaterThanLessThan_subseteq_greaterThanAtMost_iff. tts_lemma greaterThanLessThan_subseteq_greaterThanLessThan: assumes "a \ U" and "b \ U" and "c \ U" and "d \ U" shows "({a<\<^sub>o\<^sub>w..<\<^sub>o\<^sub>wb} \ {c<\<^sub>o\<^sub>w..<\<^sub>o\<^sub>wd}) = (a <\<^sub>o\<^sub>w b \ b \\<^sub>o\<^sub>w d \ c \\<^sub>o\<^sub>w a)" is dense_linorder_class.greaterThanLessThan_subseteq_greaterThanLessThan. end end text\\newpage\ end \ No newline at end of file diff --git a/thys/Types_To_Sets_Extension/Examples/SML_Relativization/Simple_Orders/SML_Simple_Orders.thy b/thys/Types_To_Sets_Extension/Examples/SML_Relativization/Simple_Orders/SML_Simple_Orders.thy --- a/thys/Types_To_Sets_Extension/Examples/SML_Relativization/Simple_Orders/SML_Simple_Orders.thy +++ b/thys/Types_To_Sets_Extension/Examples/SML_Relativization/Simple_Orders/SML_Simple_Orders.thy @@ -1,1985 +1,1821 @@ (* Title: Examples/SML_Relativization/Simple_Orders/SML_Simple_Orders.thy Author: Mihails Milehins Copyright 2021 (C) Mihails Milehins *) section\Relativization of the results about orders\ theory SML_Simple_Orders imports "../../Introduction" "../Foundations/SML_Relations" Complex_Main begin subsection\Class \<^class>\ord\\ subsubsection\Definitions and common properties\ locale ord_ow = fixes U :: "'ao set" and le :: "['ao, 'ao] \ bool" (infix \\\<^sub>o\<^sub>w\ 50) and ls :: "['ao, 'ao] \ bool" (infix \<\<^sub>o\<^sub>w\ 50) begin notation le (\'(\\<^sub>o\<^sub>w')\) and le (infix \\\<^sub>o\<^sub>w\ 50) and ls (\'(<\<^sub>o\<^sub>w')\) and ls (infix \<\<^sub>o\<^sub>w\ 50) abbreviation (input) ge (infix \\\<^sub>o\<^sub>w\ 50) where "x \\<^sub>o\<^sub>w y \ y \\<^sub>o\<^sub>w x" abbreviation (input) gt (infix \>\<^sub>o\<^sub>w\ 50) where "x >\<^sub>o\<^sub>w y \ y <\<^sub>o\<^sub>w x" notation ge (\'(\\<^sub>o\<^sub>w')\) and ge (infix \\\<^sub>o\<^sub>w\ 50) and gt (\'(>\<^sub>o\<^sub>w')\) and gt (infix \>\<^sub>o\<^sub>w\ 50) tts_register_sbts \(\\<^sub>o\<^sub>w)\ | U proof- assume "Domainp AOA = (\x. x \ U)" "bi_unique AOA" "right_total AOA" from tts_AA_eq_transfer[OF this] show ?thesis by auto qed tts_register_sbts \(<\<^sub>o\<^sub>w)\ | U proof- assume "Domainp AOA = (\x. x \ U)" "bi_unique AOA" "right_total AOA" from tts_AA_eq_transfer[OF this] show ?thesis by auto qed end locale ord_pair_ow = ord\<^sub>1: ord_ow U\<^sub>1 le\<^sub>1 ls\<^sub>1 + ord\<^sub>2: ord_ow U\<^sub>2 le\<^sub>2 ls\<^sub>2 for U\<^sub>1 :: "'ao set" and le\<^sub>1 ls\<^sub>1 and U\<^sub>2 :: "'bo set" and le\<^sub>2 ls\<^sub>2 begin notation le\<^sub>1 (\'(\\<^sub>o\<^sub>w\<^sub>.\<^sub>1')\) and le\<^sub>1 (infix \\\<^sub>o\<^sub>w\<^sub>.\<^sub>1\ 50) and ls\<^sub>1 (\'(<\<^sub>o\<^sub>w\<^sub>.\<^sub>1')\) and ls\<^sub>1 (infix \<\<^sub>o\<^sub>w\<^sub>.\<^sub>1\ 50) and le\<^sub>2 (\'(\\<^sub>o\<^sub>w\<^sub>.\<^sub>2')\) and le\<^sub>2 (infix \\\<^sub>o\<^sub>w\<^sub>.\<^sub>2\ 50) and ls\<^sub>2 (\'(<\<^sub>o\<^sub>w\<^sub>.\<^sub>2')\) and ls\<^sub>2 (infix \<\<^sub>o\<^sub>w\<^sub>.\<^sub>2\ 50) notation ord\<^sub>1.ge (\'(\\<^sub>o\<^sub>w\<^sub>.\<^sub>1')\) and ord\<^sub>1.ge (infix \\\<^sub>o\<^sub>w\<^sub>.\<^sub>1\ 50) and ord\<^sub>1.gt (\'(>\<^sub>o\<^sub>w\<^sub>.\<^sub>1')\) and ord\<^sub>1.gt (infix \>\<^sub>o\<^sub>w\<^sub>.\<^sub>1\ 50) and ord\<^sub>2.ge (\'(\\<^sub>o\<^sub>w\<^sub>.\<^sub>2')\) and ord\<^sub>2.ge (infix \\\<^sub>o\<^sub>w\<^sub>.\<^sub>2\ 50) and ord\<^sub>2.gt (\'(>\<^sub>o\<^sub>w\<^sub>.\<^sub>2')\) and ord\<^sub>2.gt (infix \>\<^sub>o\<^sub>w\<^sub>.\<^sub>2\ 50) end ud \ord.lessThan\ (\(with _ : ({..<_}))\ [1000] 10) ud lessThan' \lessThan\ ud \ord.atMost\ (\(with _ : ({.._}))\ [1000] 10) ud atMost' \atMost\ ud \ord.greaterThan\ (\(with _ : ({_<..}))\ [1000] 10) ud greaterThan' \greaterThan\ ud \ord.atLeast\ (\(with _ : ({_..}))\ [1000] 10) ud atLeast' \atLeast\ ud \ord.greaterThanLessThan\ (\(with _ : ({_<..<_}))\ [1000, 999, 1000] 10) ud greaterThanLessThan' \greaterThanLessThan\ ud \ord.atLeastLessThan\ (\(with _ _ : ({_..<_}))\ [1000, 999, 1000, 1000] 10) ud atLeastLessThan' \atLeastLessThan\ ud \ord.greaterThanAtMost\ (\(with _ _ : ({_<.._}))\ [1000, 999, 1000, 999] 10) ud greaterThanAtMost' \greaterThanAtMost\ ud \ord.atLeastAtMost\ (\(with _ : ({_.._}))\ [1000, 1000, 1000] 10) ud atLeastAtMost' \atLeastAtMost\ ud \ord.min\ (\(with _ : \min\ _ _)\ [1000, 1000, 999] 10) ud min' \min\ ud \ord.max\ (\(with _ : \max\ _ _)\ [1000, 1000, 999] 10) ud max' \max\ ctr relativization synthesis ctr_simps assumes [transfer_domain_rule, transfer_rule]: "Domainp A = (\x. x \ U)" and [transfer_rule]: "right_total A" trp (?'a A) in lessThan_ow: lessThan.with_def (\(on _ with _ : ({..<_}))\ [1000, 1000, 1000] 10) and atMost_ow: atMost.with_def (\(on _ with _ : ({.._}))\ [1000, 1000, 1000] 10) and greaterThan_ow: greaterThan.with_def (\(on _ with _: ({_<..}))\ [1000, 1000, 1000] 10) and atLeast_ow: atLeast.with_def (\(on _ with _ : ({_..}))\ [1000, 1000, 1000] 10) ctr relativization synthesis ctr_simps assumes [transfer_domain_rule, transfer_rule]: "Domainp A = (\x. x \ U)" and [transfer_rule]: "bi_unique A" "right_total A" trp (?'a A) in greaterThanLessThan_ow: greaterThanLessThan.with_def (\(on _ with _ : ({_<..<_}))\ [1000, 1000, 1000, 1000] 10) and atLeastLessThan_ow: atLeastLessThan.with_def (\(on _ with _ _ : ({_..<_}))\ [1000, 1000, 999, 1000, 1000] 10) and greaterThanAtMost_ow: greaterThanAtMost.with_def (\(on _ with _ _ : ({_<.._}))\ [1000, 1000, 999, 1000, 1000] 10) and atLeastAtMost_ow: atLeastAtMost.with_def (\(on _ with _ : ({_.._}))\ [1000, 1000, 1000, 1000] 10) ctr parametricity in min_ow: min.with_def and max_ow: max.with_def context ord_ow begin abbreviation lessThan :: "'ao \ 'ao set" ("(1{..<\<^sub>o\<^sub>w_})") where "{..<\<^sub>o\<^sub>w u} \ on U with (<\<^sub>o\<^sub>w) : {.. 'ao set" ("(1{..\<^sub>o\<^sub>w_})") where "{..\<^sub>o\<^sub>w u} \ on U with (\\<^sub>o\<^sub>w) : {..u}" abbreviation greaterThan :: "'ao \ 'ao set" ("(1{_<\<^sub>o\<^sub>w..})") where "{l<\<^sub>o\<^sub>w..} \ on U with (<\<^sub>o\<^sub>w) : {l<..}" abbreviation atLeast :: "'ao \ 'ao set" ("(1{_..\<^sub>o\<^sub>w})") where "atLeast l \ on U with (\\<^sub>o\<^sub>w) : {l..}" abbreviation greaterThanLessThan :: "'ao \ 'ao \ 'ao set" ("(1{_<\<^sub>o\<^sub>w..<\<^sub>o\<^sub>w_})") where "{l<\<^sub>o\<^sub>w..<\<^sub>o\<^sub>wu} \ on U with (<\<^sub>o\<^sub>w) : {l<.. 'ao \ 'ao set" ("(1{_..<\<^sub>o\<^sub>w_})") where "{l..<\<^sub>o\<^sub>w u} \ on U with (\\<^sub>o\<^sub>w) (<\<^sub>o\<^sub>w) : {l<..u}" abbreviation greaterThanAtMost :: "'ao \ 'ao \ 'ao set" ("(1{_<\<^sub>o\<^sub>w.._})") where "{l<\<^sub>o\<^sub>w..u} \ on U with (\\<^sub>o\<^sub>w) (<\<^sub>o\<^sub>w) : {l<..u}" abbreviation atLeastAtMost :: "'ao \ 'ao \ 'ao set" ("(1{_..\<^sub>o\<^sub>w_})") where "{l..\<^sub>o\<^sub>wu} \ on U with (\\<^sub>o\<^sub>w) : {l..u}" abbreviation min :: "'ao \ 'ao \ 'ao" where "min \ min.with (\\<^sub>o\<^sub>w)" abbreviation max :: "'ao \ 'ao \ 'ao" where "max \ max.with (\\<^sub>o\<^sub>w)" end context ord_pair_ow begin notation ord\<^sub>1.lessThan ("(1{..<\<^sub>o\<^sub>w\<^sub>.\<^sub>1_})") notation ord\<^sub>1.atMost ("(1{..\<^sub>o\<^sub>w\<^sub>.\<^sub>1_})") notation ord\<^sub>1.greaterThan ("(1{_<\<^sub>o\<^sub>w\<^sub>.\<^sub>1..})") notation ord\<^sub>1.atLeast ("(1{_..\<^sub>o\<^sub>w\<^sub>.\<^sub>1})") notation ord\<^sub>1.greaterThanLessThan ("(1{_<\<^sub>o\<^sub>w\<^sub>.\<^sub>1..<\<^sub>o\<^sub>w\<^sub>.\<^sub>1_})") notation ord\<^sub>1.atLeastLessThan ("(1{_..<\<^sub>o\<^sub>w\<^sub>.\<^sub>1_})") notation ord\<^sub>1.greaterThanAtMost ("(1{_<\<^sub>o\<^sub>w\<^sub>.\<^sub>1.._})") notation ord\<^sub>1.atLeastAtMost ("(1{_..\<^sub>o\<^sub>w\<^sub>.\<^sub>1_})") notation ord\<^sub>2.lessThan ("(1{..<\<^sub>o\<^sub>w\<^sub>.\<^sub>2_})") notation ord\<^sub>2.atMost ("(1{..\<^sub>o\<^sub>w\<^sub>.\<^sub>2_})") notation ord\<^sub>2.greaterThan ("(1{_<\<^sub>o\<^sub>w\<^sub>.\<^sub>2..})") notation ord\<^sub>2.atLeast ("(1{_..\<^sub>o\<^sub>w\<^sub>.\<^sub>2})") notation ord\<^sub>2.greaterThanLessThan ("(1{_<\<^sub>o\<^sub>w\<^sub>.\<^sub>2..<\<^sub>o\<^sub>w\<^sub>.\<^sub>2_})") notation ord\<^sub>2.atLeastLessThan ("(1{_..<\<^sub>o\<^sub>w\<^sub>.\<^sub>2_})") notation ord\<^sub>2.greaterThanAtMost ("(1{_<\<^sub>o\<^sub>w\<^sub>.\<^sub>2.._})") notation ord\<^sub>2.atLeastAtMost ("(1{_..\<^sub>o\<^sub>w\<^sub>.\<^sub>2_})") end subsection\Preorders\ subsubsection\Definitions and common properties\ locale partial_preordering_ow = fixes U :: "'ao set" and le :: "'ao \ 'ao \ bool" (infix "\<^bold>\\<^sub>o\<^sub>w" 50) assumes refl: "a \ U \ a \<^bold>\\<^sub>o\<^sub>w a" and trans: "\ a \ U; b \ U; c \ U; a \<^bold>\\<^sub>o\<^sub>w b; b \<^bold>\\<^sub>o\<^sub>w c \ \ a \<^bold>\\<^sub>o\<^sub>w c" begin notation le (infix "\<^bold>\\<^sub>o\<^sub>w" 50) end locale preordering_ow = partial_preordering_ow U le for U :: "'ao set" and le :: "'ao \ 'ao \ bool" (infix "\<^bold>\\<^sub>o\<^sub>w" 50) + fixes ls :: \'ao \ 'ao \ bool\ (infix "\<^bold><\<^sub>o\<^sub>w" 50) assumes strict_iff_not: "\ a \ U; b \ U \ \ a \<^bold><\<^sub>o\<^sub>w b \ a \<^bold>\\<^sub>o\<^sub>w b \ \ b \<^bold>\\<^sub>o\<^sub>w a" locale preorder_ow = ord_ow U le ls for U :: "'ao set" and le ls + assumes less_le_not_le: "\ x \ U; y \ U \ \ x <\<^sub>o\<^sub>w y \ x \\<^sub>o\<^sub>w y \ \ (y \\<^sub>o\<^sub>w x)" and order_refl[iff]: "x \ U \ x \\<^sub>o\<^sub>w x" and order_trans: "\ x \ U; y \ U; z \ U; x \\<^sub>o\<^sub>w y; y \\<^sub>o\<^sub>w z \ \ x \\<^sub>o\<^sub>w z" begin sublocale order: preordering_ow U \(\\<^sub>o\<^sub>w)\ \(<\<^sub>o\<^sub>w)\ + dual_order: preordering_ow U \(\\<^sub>o\<^sub>w)\ \(>\<^sub>o\<^sub>w)\ apply unfold_locales subgoal by auto subgoal by (meson order_trans) subgoal using less_le_not_le by simp subgoal by (meson order_trans) subgoal by (metis less_le_not_le) done end locale ord_preorder_ow = ord\<^sub>1: ord_ow U\<^sub>1 le\<^sub>1 ls\<^sub>1 + ord\<^sub>2: preorder_ow U\<^sub>2 le\<^sub>2 ls\<^sub>2 for U\<^sub>1 :: "'ao set" and le\<^sub>1 ls\<^sub>1 and U\<^sub>2 :: "'bo set" and le\<^sub>2 ls\<^sub>2 begin sublocale ord_pair_ow . end locale preorder_pair_ow = ord\<^sub>1: preorder_ow U\<^sub>1 le\<^sub>1 ls\<^sub>1 + ord\<^sub>2: preorder_ow U\<^sub>2 le\<^sub>2 ls\<^sub>2 for U\<^sub>1 :: "'ao set" and le\<^sub>1 and ls\<^sub>1 and U\<^sub>2 :: "'bo set" and le\<^sub>2 and ls\<^sub>2 begin sublocale ord_preorder_ow .. end ud \preordering_bdd.bdd\ ud \preorder.bdd_above\ ud bdd_above' \bdd_above\ ud \preorder.bdd_below\ ud bdd_below' \bdd_below\ ctr relativization synthesis ctr_simps assumes [transfer_domain_rule, transfer_rule]: "Domainp A = (\x. x \ U)" and [transfer_rule]: "right_total A" trp (?'a A) in bdd_ow: bdd.with_def (\(on _ with _ : \bdd\ _)\ [1000, 1000, 1000] 10) and bdd_above_ow: bdd_above.with_def (\(on _ with _ : \bdd'_above\ _)\ [1000, 1000, 1000] 10) and bdd_below_ow: bdd_below.with_def (\(on _ with _ : \bdd'_below\ _)\ [1000, 1000, 1000] 10) declare bdd.with[ud_with del] context preorder_ow begin abbreviation bdd_above :: "'ao set \ bool" where "bdd_above \ bdd_above_ow U (\\<^sub>o\<^sub>w)" abbreviation bdd_below :: "'ao set \ bool" where "bdd_below \ bdd_below_ow U (\\<^sub>o\<^sub>w)" end subsubsection\Transfer rules\ context includes lifting_syntax begin lemma partial_preordering_transfer[transfer_rule]: assumes [transfer_rule]: "right_total A" shows "((A ===> A ===> (=)) ===> (=)) (partial_preordering_ow (Collect (Domainp A))) partial_preordering" unfolding partial_preordering_ow_def partial_preordering_def apply transfer_prover_start apply transfer_step+ by blast lemma preordering_transfer[transfer_rule]: assumes [transfer_rule]: "right_total A" shows "((A ===> A ===> (=)) ===> (A ===> A ===> (=)) ===> (=)) (preordering_ow (Collect (Domainp A))) preordering" unfolding preordering_ow_def preordering_ow_axioms_def preordering_def preordering_axioms_def apply transfer_prover_start apply transfer_step+ by blast lemma preorder_transfer[transfer_rule]: assumes [transfer_rule]: "right_total A" shows "((A ===> A ===> (=)) ===> (A ===> A ===> (=)) ===> (=)) (preorder_ow (Collect (Domainp A))) class.preorder" unfolding preorder_ow_def class.preorder_def apply transfer_prover_start apply transfer_step+ by blast end subsubsection\Relativization\ context preordering_ow begin tts_context tts: (?'a to U) rewriting ctr_simps substituting preordering_ow_axioms eliminating through auto begin tts_lemma strict_implies_order: assumes "a \ U" and "b \ U" and "a \<^bold><\<^sub>o\<^sub>w b" shows "a \<^bold>\\<^sub>o\<^sub>w b" is preordering.strict_implies_order. tts_lemma irrefl: assumes "a \ U" shows "\a \<^bold><\<^sub>o\<^sub>w a" is preordering.irrefl. tts_lemma asym: assumes "a \ U" and "b \ U" and "a \<^bold><\<^sub>o\<^sub>w b" and "b \<^bold><\<^sub>o\<^sub>w a" shows False is preordering.asym. tts_lemma strict_trans1: assumes "a \ U" and "b \ U" and "c \ U" and "a \<^bold>\\<^sub>o\<^sub>w b" and "b \<^bold><\<^sub>o\<^sub>w c" shows "a \<^bold><\<^sub>o\<^sub>w c" is preordering.strict_trans1. tts_lemma strict_trans2: assumes "a \ U" and "b \ U" and "c \ U" and "a \<^bold><\<^sub>o\<^sub>w b" and "b \<^bold>\\<^sub>o\<^sub>w c" shows "a \<^bold><\<^sub>o\<^sub>w c" is preordering.strict_trans2. tts_lemma strict_trans: assumes "a \ U" and "b \ U" and "c \ U" and "a \<^bold><\<^sub>o\<^sub>w b" and "b \<^bold><\<^sub>o\<^sub>w c" shows "a \<^bold><\<^sub>o\<^sub>w c" is preordering.strict_trans. end end context preorder_ow begin tts_context tts: (?'a to U) rewriting ctr_simps substituting preorder_ow_axioms eliminating through auto begin tts_lemma less_irrefl: assumes "x \ U" shows "\ x <\<^sub>o\<^sub>w x" is preorder_class.less_irrefl. tts_lemma bdd_below_Ioc: assumes "a \ U" and "b \ U" shows "bdd_below {a<\<^sub>o\<^sub>w..b}" is preorder_class.bdd_below_Ioc. tts_lemma bdd_above_Ioc: assumes "a \ U" and "b \ U" shows "bdd_above {a<\<^sub>o\<^sub>w..b}" is preorder_class.bdd_above_Ioc. tts_lemma bdd_above_Iic: assumes "b \ U" shows "bdd_above {..\<^sub>o\<^sub>wb}" is preorder_class.bdd_above_Iic. tts_lemma bdd_above_Iio: assumes "b \ U" shows "bdd_above {..<\<^sub>o\<^sub>wb}" is preorder_class.bdd_above_Iio. tts_lemma bdd_below_Ici: assumes "a \ U" shows "bdd_below {a..\<^sub>o\<^sub>w}" is preorder_class.bdd_below_Ici. tts_lemma bdd_below_Ioi: assumes "a \ U" shows "bdd_below {a<\<^sub>o\<^sub>w..}" is preorder_class.bdd_below_Ioi. tts_lemma bdd_above_Icc: assumes "a \ U" and "b \ U" shows "bdd_above {a..\<^sub>o\<^sub>wb}" is preorder_class.bdd_above_Icc. tts_lemma bdd_above_Ioo: assumes "a \ U" and "b \ U" shows "bdd_above {a<\<^sub>o\<^sub>w..<\<^sub>o\<^sub>wb}" is preorder_class.bdd_above_Ioo. tts_lemma bdd_below_Icc: assumes "a \ U" and "b \ U" shows "bdd_below {a..\<^sub>o\<^sub>wb}" is preorder_class.bdd_below_Icc. tts_lemma bdd_below_Ioo: assumes "a \ U" and "b \ U" shows "bdd_below {a<\<^sub>o\<^sub>w..<\<^sub>o\<^sub>wb}" is preorder_class.bdd_below_Ioo. tts_lemma bdd_above_Ico: assumes "a \ U" and "b \ U" shows "bdd_above (on U with (\\<^sub>o\<^sub>w) (<\<^sub>o\<^sub>w) : {a.. U" and "b \ U" shows "bdd_below (on U with (\\<^sub>o\<^sub>w) (<\<^sub>o\<^sub>w) : {a.. U" shows "{a<\<^sub>o\<^sub>w..} \ {a..\<^sub>o\<^sub>w}" is preorder_class.Ioi_le_Ico. tts_lemma eq_refl: assumes "y \ U" and "x = y" shows "x \\<^sub>o\<^sub>w y" is preorder_class.eq_refl. tts_lemma less_imp_le: assumes "x \ U" and "y \ U" and "x <\<^sub>o\<^sub>w y" shows "x \\<^sub>o\<^sub>w y" is preorder_class.less_imp_le. tts_lemma less_not_sym: assumes "x \ U" and "y \ U" and "x <\<^sub>o\<^sub>w y" shows "\ y <\<^sub>o\<^sub>w x" is preorder_class.less_not_sym. tts_lemma less_imp_not_less: assumes "x \ U" and "y \ U" and "x <\<^sub>o\<^sub>w y" shows "(\ y <\<^sub>o\<^sub>w x) = True" is preorder_class.less_imp_not_less. tts_lemma less_asym': assumes "a \ U" and "b \ U" and "a <\<^sub>o\<^sub>w b" and "b <\<^sub>o\<^sub>w a" shows P is preorder_class.less_asym'. tts_lemma less_imp_triv: assumes "x \ U" and "y \ U" and "x <\<^sub>o\<^sub>w y" shows "(y <\<^sub>o\<^sub>w x \ P) = True" is preorder_class.less_imp_triv. tts_lemma less_trans: assumes "x \ U" and "y \ U" and "z \ U" and "x <\<^sub>o\<^sub>w y" and "y <\<^sub>o\<^sub>w z" shows "x <\<^sub>o\<^sub>w z" is preorder_class.less_trans. tts_lemma less_le_trans: assumes "x \ U" and "y \ U" and "z \ U" and "x <\<^sub>o\<^sub>w y" and "y \\<^sub>o\<^sub>w z" shows "x <\<^sub>o\<^sub>w z" is preorder_class.less_le_trans. tts_lemma le_less_trans: assumes "x \ U" and "y \ U" and "z \ U" and "x \\<^sub>o\<^sub>w y" and "y <\<^sub>o\<^sub>w z" shows "x <\<^sub>o\<^sub>w z" is preorder_class.le_less_trans. tts_lemma bdd_aboveI: assumes "A \ U" and "M \ U" and "\x. \x \ U; x \ A\ \ x \\<^sub>o\<^sub>w M" shows "bdd_above A" is preorder_class.bdd_aboveI. tts_lemma bdd_belowI: assumes "A \ U" and "m \ U" and "\x. \x \ U; x \ A\ \ m \\<^sub>o\<^sub>w x" shows "bdd_below A" is preorder_class.bdd_belowI. tts_lemma less_asym: assumes "x \ U" and "y \ U" and "x <\<^sub>o\<^sub>w y" and "\ P \ y <\<^sub>o\<^sub>w x" shows P is preorder_class.less_asym. tts_lemma transp_less: "transp_on U (<\<^sub>o\<^sub>w)" is preorder_class.transp_less. tts_lemma transp_le: "transp_on U (\\<^sub>o\<^sub>w)" is preorder_class.transp_le. tts_lemma transp_gr: "transp_on U (\x y. y <\<^sub>o\<^sub>w x)" is preorder_class.transp_gr. tts_lemma transp_ge: "transp_on U (\x y. y \\<^sub>o\<^sub>w x)" is preorder_class.transp_ge. tts_lemma bdd_above_Int1: assumes "A \ U" and "B \ U" and "bdd_above A" shows "bdd_above (A \ B)" is preorder_class.bdd_above_Int1. tts_lemma bdd_above_Int2: assumes "B \ U" and "A \ U" and "bdd_above B" shows "bdd_above (A \ B)" is preorder_class.bdd_above_Int2. tts_lemma bdd_below_Int1: assumes "A \ U" and "B \ U" and "bdd_below A" shows "bdd_below (A \ B)" is preorder_class.bdd_below_Int1. tts_lemma bdd_below_Int2: assumes "B \ U" and "A \ U" and "bdd_below B" shows "bdd_below (A \ B)" is preorder_class.bdd_below_Int2. tts_lemma bdd_above_mono: assumes "B \ U" and "bdd_above B" and "A \ B" shows "bdd_above A" is preorder_class.bdd_above_mono. tts_lemma bdd_below_mono: assumes "B \ U" and "bdd_below B" and "A \ B" shows "bdd_below A" is preorder_class.bdd_below_mono. tts_lemma atLeastAtMost_subseteq_atLeastLessThan_iff: assumes "a \ U" and "b \ U" and "c \ U" and "d \ U" shows "({a..\<^sub>o\<^sub>wb} \ (on U with (\\<^sub>o\<^sub>w) (<\<^sub>o\<^sub>w) : {c..\<^sub>o\<^sub>w b \ b <\<^sub>o\<^sub>w d \ c \\<^sub>o\<^sub>w a)" is preorder_class.atLeastAtMost_subseteq_atLeastLessThan_iff. tts_lemma atMost_subset_iff: assumes "x \ U" and "y \ U" shows "({..\<^sub>o\<^sub>wx} \ {..\<^sub>o\<^sub>wy}) = (x \\<^sub>o\<^sub>w y)" is Set_Interval.atMost_subset_iff. tts_lemma single_Diff_lessThan: assumes "k \ U" shows "{k} - {..<\<^sub>o\<^sub>wk} = {k}" is Set_Interval.single_Diff_lessThan. tts_lemma atLeast_subset_iff: assumes "x \ U" and "y \ U" shows "({x..\<^sub>o\<^sub>w} \ {y..\<^sub>o\<^sub>w}) = (y \\<^sub>o\<^sub>w x)" is Set_Interval.atLeast_subset_iff. tts_lemma atLeastatMost_psubset_iff: assumes "a \ U" and "b \ U" and "c \ U" and "d \ U" shows "({a..\<^sub>o\<^sub>wb} \ {c..\<^sub>o\<^sub>wd}) = (c \\<^sub>o\<^sub>w d \ (\ a \\<^sub>o\<^sub>w b \ c \\<^sub>o\<^sub>w a \ b \\<^sub>o\<^sub>w d \ (c <\<^sub>o\<^sub>w a \ b <\<^sub>o\<^sub>w d)))" is preorder_class.atLeastatMost_psubset_iff. tts_lemma not_empty_eq_Iic_eq_empty: assumes "h \ U" shows "{} \ {..\<^sub>o\<^sub>wh}" is preorder_class.not_empty_eq_Iic_eq_empty. tts_lemma atLeastatMost_subset_iff: assumes "a \ U" and "b \ U" and "c \ U" and "d \ U" shows "({a..\<^sub>o\<^sub>wb} \ {c..\<^sub>o\<^sub>wd}) = (\ a \\<^sub>o\<^sub>w b \ b \\<^sub>o\<^sub>w d \ c \\<^sub>o\<^sub>w a)" is preorder_class.atLeastatMost_subset_iff. tts_lemma Icc_subset_Ici_iff: assumes "l \ U" and "h \ U" and "l' \ U" shows "({l..\<^sub>o\<^sub>wh} \ {l'..\<^sub>o\<^sub>w}) = (\ l \\<^sub>o\<^sub>w h \ l' \\<^sub>o\<^sub>w l)" is preorder_class.Icc_subset_Ici_iff. tts_lemma Icc_subset_Iic_iff: assumes "l \ U" and "h \ U" and "h' \ U" shows "({l..\<^sub>o\<^sub>wh} \ {..\<^sub>o\<^sub>wh'}) = (\ l \\<^sub>o\<^sub>w h \ h \\<^sub>o\<^sub>w h')" is preorder_class.Icc_subset_Iic_iff. tts_lemma not_empty_eq_Ici_eq_empty: assumes "l \ U" shows "{} \ {l..\<^sub>o\<^sub>w}" is preorder_class.not_empty_eq_Ici_eq_empty. tts_lemma not_Ici_eq_empty: assumes "l \ U" shows "{l..\<^sub>o\<^sub>w} \ {}" is preorder_class.not_Ici_eq_empty. tts_lemma not_Iic_eq_empty: assumes "h \ U" shows "{..\<^sub>o\<^sub>wh} \ {}" is preorder_class.not_Iic_eq_empty. tts_lemma atLeastatMost_empty_iff2: assumes "a \ U" and "b \ U" shows "({} = {a..\<^sub>o\<^sub>wb}) = (\ a \\<^sub>o\<^sub>w b)" is preorder_class.atLeastatMost_empty_iff2. tts_lemma atLeastLessThan_empty_iff2: assumes "a \ U" and "b \ U" shows "({} = (on U with (\\<^sub>o\<^sub>w) (<\<^sub>o\<^sub>w) : {a.. a <\<^sub>o\<^sub>w b)" is preorder_class.atLeastLessThan_empty_iff2. tts_lemma greaterThanAtMost_empty_iff2: assumes "k \ U" and "l \ U" shows "({} = {k<\<^sub>o\<^sub>w..l}) = (\ k <\<^sub>o\<^sub>w l)" is preorder_class.greaterThanAtMost_empty_iff2. tts_lemma atLeastatMost_empty_iff: assumes "a \ U" and "b \ U" shows "({a..\<^sub>o\<^sub>wb} = {}) = (\ a \\<^sub>o\<^sub>w b)" is preorder_class.atLeastatMost_empty_iff. tts_lemma atLeastLessThan_empty_iff: assumes "a \ U" and "b \ U" shows "((on U with (\\<^sub>o\<^sub>w) (<\<^sub>o\<^sub>w) : {a.. a <\<^sub>o\<^sub>w b)" is preorder_class.atLeastLessThan_empty_iff. tts_lemma greaterThanAtMost_empty_iff: assumes "k \ U" and "l \ U" shows "({k<\<^sub>o\<^sub>w..l} = {}) = (\ k <\<^sub>o\<^sub>w l)" is preorder_class.greaterThanAtMost_empty_iff. end tts_context tts: (?'a to U) substituting preorder_ow_axioms begin tts_lemma bdd_above_empty: assumes "U \ {}" shows "bdd_above {}" is preorder_class.bdd_above_empty. tts_lemma bdd_below_empty: assumes "U \ {}" shows "bdd_below {}" is preorder_class.bdd_below_empty. end tts_context tts: (?'a to U) and (?'b to \U\<^sub>2::'a set\) rewriting ctr_simps substituting preorder_ow_axioms eliminating through (auto intro: bdd_above_empty bdd_below_empty) begin tts_lemma bdd_belowI2: assumes "A \ U\<^sub>2" and "m \ U" and "\x\U\<^sub>2. f x \ U" and "\x. x \ A \ m \\<^sub>o\<^sub>w f x" shows "bdd_below (f ` A)" given preorder_class.bdd_belowI2 by blast tts_lemma bdd_aboveI2: assumes "A \ U\<^sub>2" and "\x\U\<^sub>2. f x \ U" and "M \ U" and "\x. x \ A \ f x \\<^sub>o\<^sub>w M" shows "bdd_above (f ` A)" given preorder_class.bdd_aboveI2 by blast end end subsection\Partial orders\ subsubsection\Definitions and common properties\ locale ordering_ow = partial_preordering_ow U le for U :: "'ao set" and le :: "'ao \ 'ao \ bool" (infix "\<^bold>\\<^sub>o\<^sub>w" 50) + fixes ls :: "'ao \ 'ao \ bool" (infix "\<^bold><\<^sub>o\<^sub>w" 50) assumes strict_iff_order: "\ a \ U; b \ U \ \ a \<^bold><\<^sub>o\<^sub>w b \ a \<^bold>\\<^sub>o\<^sub>w b \ a \ b" and antisym: "\ a \ U; b \ U; a \<^bold>\\<^sub>o\<^sub>w b; b \<^bold>\\<^sub>o\<^sub>w a \ \ a = b" begin notation le (infix "\<^bold>\\<^sub>o\<^sub>w" 50) and ls (infix "\<^bold><\<^sub>o\<^sub>w" 50) sublocale preordering_ow U \(\<^bold>\\<^sub>o\<^sub>w)\ \(\<^bold><\<^sub>o\<^sub>w)\ using local.antisym strict_iff_order by unfold_locales blast end locale order_ow = preorder_ow U le ls for U :: "'ao set" and le ls + assumes antisym: "\ x \ U; y \ U; x \\<^sub>o\<^sub>w y; y \\<^sub>o\<^sub>w x \ \ x = y" begin sublocale order: ordering_ow U \(\\<^sub>o\<^sub>w)\ \(<\<^sub>o\<^sub>w)\ + dual_order: ordering_ow U \(\\<^sub>o\<^sub>w)\ \(>\<^sub>o\<^sub>w)\ apply unfold_locales subgoal by (force simp: less_le_not_le antisym) subgoal by (simp add: antisym) subgoal by (force simp: less_le_not_le antisym) subgoal by (simp add: antisym) done no_notation le (infix "\<^bold>\\<^sub>o\<^sub>w" 50) and ls (infix "\<^bold><\<^sub>o\<^sub>w" 50) end locale ord_order_ow = ord\<^sub>1: ord_ow U\<^sub>1 le\<^sub>1 ls\<^sub>1 + ord\<^sub>2: order_ow U\<^sub>2 le\<^sub>2 ls\<^sub>2 for U\<^sub>1 :: "'ao set" and le\<^sub>1 ls\<^sub>1 and U\<^sub>2 :: "'bo set" and le\<^sub>2 ls\<^sub>2 sublocale ord_order_ow \ ord_preorder_ow .. locale preorder_order_ow = ord\<^sub>1: preorder_ow U\<^sub>1 le\<^sub>1 ls\<^sub>1 + ord\<^sub>2: order_ow U\<^sub>2 le\<^sub>2 ls\<^sub>2 for U\<^sub>1 :: "'ao set" and le\<^sub>1 ls\<^sub>1 and U\<^sub>2 :: "'bo set" and le\<^sub>2 ls\<^sub>2 sublocale preorder_order_ow \ preorder_pair_ow .. locale order_pair_ow = ord\<^sub>1: order_ow U\<^sub>1 le\<^sub>1 ls\<^sub>1 + ord\<^sub>2: order_ow U\<^sub>2 le\<^sub>2 ls\<^sub>2 for U\<^sub>1 :: "'ao set" and le\<^sub>1 ls\<^sub>1 and U\<^sub>2 :: "'bo set" and le\<^sub>2 ls\<^sub>2 sublocale order_pair_ow \ preorder_order_ow .. -ud \order.mono\ (\(with _ _ : \mono\ _)\ [1000, 999, 1000] 10) -ud mono' \mono\ -ud \order.strict_mono\ (\(with _ _ : \strict'_mono\ _)\ [1000, 999, 1000] 10) -ud strict_mono' \strict_mono\ +definition mono where + "mono f \ (\x y. x \ y \ f x \ f y)" + +definition strict_mono where + "strict_mono f \ (\x y. x < y \ f x < f y)" + +lemma mono_iff_mono: "mono f \ Fun.mono f" + by (simp add: mono_def Fun.mono_def) + +ud \mono\ (\(with _ _ : \mono\ _)\ [1000, 999, 1000] 10) +ud \strict_mono\ (\(with _ _ : \strict'_mono\ _)\ [1000, 999, 1000] 10) ud \order.antimono\ (\(with _ _ : \strict'_mono\ _)\ [1000, 999, 1000] 10) ud antimono' \antimono\ ud \monoseq\ (\(with _ : \monoseq\ _)\ [1000, 1000] 10) ctr relativization synthesis ctr_simps assumes [transfer_domain_rule, transfer_rule]: "Domainp (B::'c\'d\bool) = (\x. x \ U\<^sub>2)" and [transfer_rule]: "right_total B" trp (?'b \A::'a\'b\bool\) and (?'a B) in mono_ow: mono.with_def (\(on _ with _ _ : \mono\ _)\ [1000, 1000, 999, 1000] 10) and strict_mono_ow: strict_mono.with_def (\(on _ with _ _ : \strict'_mono\ _)\ [1000, 1000, 999, 1000] 10) and antimono_ow: antimono.with_def (\(on _ with _ _ : \antimono\ _)\ [1000, 1000, 999, 1000] 10) and monoseq_ow: monoseq.with_def subsubsection\Transfer rules\ context includes lifting_syntax begin lemma ordering_transfer[transfer_rule]: assumes [transfer_rule]: "bi_unique A" "right_total A" shows "((A ===> A ===> (=)) ===> (A ===> A ===> (=)) ===> (=)) (ordering_ow (Collect (Domainp A))) ordering" unfolding ordering_ow_def ordering_ow_axioms_def ordering_def ordering_axioms_def apply transfer_prover_start apply transfer_step+ unfolding Ball_Collect[symmetric] by (intro ext HOL.arg_cong2[where f="(\)"]) auto lemma order_transfer[transfer_rule]: assumes [transfer_rule]: "bi_unique A" "right_total A" shows "((A ===> A ===> (=)) ===> (A ===> A ===> (=)) ===> (=)) (order_ow (Collect (Domainp A))) class.order" unfolding order_ow_def class.order_def order_ow_axioms_def class.order_axioms_def apply transfer_prover_start apply transfer_step+ by simp end subsubsection\Relativization\ context ordering_ow begin tts_context tts: (?'a to U) rewriting ctr_simps substituting ordering_ow_axioms eliminating through simp begin tts_lemma strict_implies_not_eq: assumes "a \ U" and "b \ U" and "a \<^bold><\<^sub>o\<^sub>w b" shows "a \ b" is ordering.strict_implies_not_eq. tts_lemma order_iff_strict: assumes "a \ U" and "b \ U" shows "(a \<^bold>\\<^sub>o\<^sub>w b) = (a \<^bold><\<^sub>o\<^sub>w b \ a = b)" is ordering.order_iff_strict. tts_lemma not_eq_order_implies_strict: assumes "a \ U" and "b \ U" and "a \ b" and "a \<^bold>\\<^sub>o\<^sub>w b" shows "a \<^bold><\<^sub>o\<^sub>w b" is ordering.not_eq_order_implies_strict. tts_lemma eq_iff: assumes "a \ U" and "b \ U" shows "(a = b) = (a \<^bold>\\<^sub>o\<^sub>w b \ b \<^bold>\\<^sub>o\<^sub>w a)" is ordering.eq_iff. end end context order_ow begin tts_context tts: (?'a to U) rewriting ctr_simps substituting order_ow_axioms eliminating through clarsimp begin tts_lemma atLeastAtMost_singleton: assumes "a \ U" shows "{a..\<^sub>o\<^sub>wa} = {a}" is order_class.atLeastAtMost_singleton. tts_lemma less_imp_neq: assumes "y \ U" and "x <\<^sub>o\<^sub>w y" shows "x \ y" is order_class.less_imp_neq. tts_lemma atLeastatMost_empty: assumes "b \ U" and "a \ U" and "b <\<^sub>o\<^sub>w a" shows "{a..\<^sub>o\<^sub>wb} = {}" is order_class.atLeastatMost_empty. tts_lemma less_imp_not_eq: assumes "y \ U" and "x <\<^sub>o\<^sub>w y" shows "(x = y) = False" is order_class.less_imp_not_eq. tts_lemma less_imp_not_eq2: assumes "y \ U" and "x <\<^sub>o\<^sub>w y" shows "(y = x) = False" is order_class.less_imp_not_eq2. tts_lemma atLeastLessThan_empty: assumes "b \ U" and "a \ U" and "b \\<^sub>o\<^sub>w a" shows "(on U with (\\<^sub>o\<^sub>w) (<\<^sub>o\<^sub>w) : {a.. U" and "k \ U" and "l \\<^sub>o\<^sub>w k" shows "{k<\<^sub>o\<^sub>w..l} = {}" is order_class.greaterThanAtMost_empty. tts_lemma antisym_conv1: assumes "x \ U" and "y \ U" and "\ x <\<^sub>o\<^sub>w y" shows "(x \\<^sub>o\<^sub>w y) = (x = y)" is order_class.antisym_conv1. tts_lemma antisym_conv2: assumes "x \ U" and "y \ U" and "x \\<^sub>o\<^sub>w y" shows "(\ x <\<^sub>o\<^sub>w y) = (x = y)" is order_class.antisym_conv2. tts_lemma greaterThanLessThan_empty: assumes "l \ U" and "k \ U" and "l \\<^sub>o\<^sub>w k" shows "{k<\<^sub>o\<^sub>w..<\<^sub>o\<^sub>wl} = {}" is order_class.greaterThanLessThan_empty. tts_lemma atLeastLessThan_eq_atLeastAtMost_diff: assumes "a \ U" and "b \ U" shows "(on U with (\\<^sub>o\<^sub>w) (<\<^sub>o\<^sub>w) : {a..o\<^sub>wb} - {b}" is order_class.atLeastLessThan_eq_atLeastAtMost_diff. tts_lemma greaterThanAtMost_eq_atLeastAtMost_diff: assumes "a \ U" and "b \ U" shows "{a<\<^sub>o\<^sub>w..b} = {a..\<^sub>o\<^sub>wb} - {a}" is order_class.greaterThanAtMost_eq_atLeastAtMost_diff. tts_lemma less_separate: assumes "x \ U" and "y \ U" and "x <\<^sub>o\<^sub>w y" shows "\x'\U. \y'\U. x \ {..<\<^sub>o\<^sub>wx'} \ y \ {y'<\<^sub>o\<^sub>w..} \ {..<\<^sub>o\<^sub>wx'} \ {y'<\<^sub>o\<^sub>w..} = {}" is order_class.less_separate. tts_lemma order_iff_strict: assumes "a \ U" and "b \ U" shows "(a \\<^sub>o\<^sub>w b) = (a <\<^sub>o\<^sub>w b \ a = b)" is order_class.order.order_iff_strict. tts_lemma le_less: assumes "x \ U" and "y \ U" shows "(x \\<^sub>o\<^sub>w y) = (x <\<^sub>o\<^sub>w y \ x = y)" is order_class.le_less. tts_lemma strict_iff_order: assumes "a \ U" and "b \ U" shows "(a <\<^sub>o\<^sub>w b) = (a \\<^sub>o\<^sub>w b \ a \ b)" is order_class.order.strict_iff_order. tts_lemma less_le: assumes "x \ U" and "y \ U" shows "(x <\<^sub>o\<^sub>w y) = (x \\<^sub>o\<^sub>w y \ x \ y)" is order_class.less_le. tts_lemma atLeastAtMost_singleton': assumes "b \ U" and "a = b" shows "{a..\<^sub>o\<^sub>wb} = {a}" is order_class.atLeastAtMost_singleton'. tts_lemma le_imp_less_or_eq: assumes "x \ U" and "y \ U" and "x \\<^sub>o\<^sub>w y" shows "x <\<^sub>o\<^sub>w y \ x = y" is order_class.le_imp_less_or_eq. tts_lemma antisym_conv: assumes "y \ U" and "x \ U" and "y \\<^sub>o\<^sub>w x" shows "(x \\<^sub>o\<^sub>w y) = (x = y)" is order_class.antisym_conv. tts_lemma le_neq_trans: assumes "a \ U" and "b \ U" and "a \\<^sub>o\<^sub>w b" and "a \ b" shows "a <\<^sub>o\<^sub>w b" is order_class.le_neq_trans. tts_lemma neq_le_trans: assumes "a \ U" and "b \ U" and "a \ b" and "a \\<^sub>o\<^sub>w b" shows "a <\<^sub>o\<^sub>w b" is order_class.neq_le_trans. tts_lemma Iio_Int_singleton: assumes "k \ U" and "x \ U" shows "{..<\<^sub>o\<^sub>wk} \ {x} = (if x <\<^sub>o\<^sub>w k then {x} else {})" is order_class.Iio_Int_singleton. tts_lemma atLeastAtMost_singleton_iff: assumes "a \ U" and "b \ U" and "c \ U" shows "({a..\<^sub>o\<^sub>wb} = {c}) = (a = b \ b = c)" is order_class.atLeastAtMost_singleton_iff. tts_lemma Icc_eq_Icc: assumes "l \ U" and "h \ U" and "l' \ U" and "h' \ U" shows "({l..\<^sub>o\<^sub>wh} = {l'..\<^sub>o\<^sub>wh'}) = (h = h' \ l = l' \ \ l' \\<^sub>o\<^sub>w h' \ \ l \\<^sub>o\<^sub>w h)" is order_class.Icc_eq_Icc. tts_lemma lift_Suc_mono_less_iff: assumes "range f \ U" and "\n. f n <\<^sub>o\<^sub>w f (Suc n)" shows "(f n <\<^sub>o\<^sub>w f m) = (n < m)" is order_class.lift_Suc_mono_less_iff. tts_lemma lift_Suc_mono_less: assumes "range f \ U" and "\n. f n <\<^sub>o\<^sub>w f (Suc n)" and "n < n'" shows "f n <\<^sub>o\<^sub>w f n'" is order_class.lift_Suc_mono_less. tts_lemma lift_Suc_mono_le: assumes "range f \ U" and "\n. f n \\<^sub>o\<^sub>w f (Suc n)" and "n \ n'" shows "f n \\<^sub>o\<^sub>w f n'" is order_class.lift_Suc_mono_le. tts_lemma lift_Suc_antimono_le: assumes "range f \ U" and "\n. f (Suc n) \\<^sub>o\<^sub>w f n" and "n \ n'" shows "f n' \\<^sub>o\<^sub>w f n" is order_class.lift_Suc_antimono_le. tts_lemma ivl_disj_int_two: assumes "l \ U" and "m \ U" and "u \ U" shows "{l<\<^sub>o\<^sub>w..<\<^sub>o\<^sub>wm} \ (on U with (\\<^sub>o\<^sub>w) (<\<^sub>o\<^sub>w) : {m..o\<^sub>w..m} \ {m<\<^sub>o\<^sub>w..<\<^sub>o\<^sub>wu} = {}" "(on U with (\\<^sub>o\<^sub>w) (<\<^sub>o\<^sub>w) : {l.. (on U with (\\<^sub>o\<^sub>w) (<\<^sub>o\<^sub>w) : {m..o\<^sub>wm} \ {m<\<^sub>o\<^sub>w..<\<^sub>o\<^sub>wu} = {}" "{l<\<^sub>o\<^sub>w..<\<^sub>o\<^sub>wm} \ {m..\<^sub>o\<^sub>wu} = {}" "{l<\<^sub>o\<^sub>w..m} \ {m<\<^sub>o\<^sub>w..u} = {}" "(on U with (\\<^sub>o\<^sub>w) (<\<^sub>o\<^sub>w) : {l.. {m..\<^sub>o\<^sub>wu} = {}" "{l..\<^sub>o\<^sub>wm} \ {m<\<^sub>o\<^sub>w..u} = {}" is Set_Interval.ivl_disj_int_two. tts_lemma ivl_disj_int_one: assumes "l \ U" and "u \ U" shows "{..\<^sub>o\<^sub>wl} \ {l<\<^sub>o\<^sub>w..<\<^sub>o\<^sub>wu} = {}" "{..<\<^sub>o\<^sub>wl} \ (on U with (\\<^sub>o\<^sub>w) (<\<^sub>o\<^sub>w) : {l..o\<^sub>wl} \ {l<\<^sub>o\<^sub>w..u} = {}" "{..<\<^sub>o\<^sub>wl} \ {l..\<^sub>o\<^sub>wu} = {}" "{l<\<^sub>o\<^sub>w..u} \ {u<\<^sub>o\<^sub>w..} = {}" "{l<\<^sub>o\<^sub>w..<\<^sub>o\<^sub>wu} \ {u..\<^sub>o\<^sub>w} = {}" "{l..\<^sub>o\<^sub>wu} \ {u<\<^sub>o\<^sub>w..} = {}" "(on U with (\\<^sub>o\<^sub>w) (<\<^sub>o\<^sub>w) : {l.. {u..\<^sub>o\<^sub>w} = {}" is Set_Interval.ivl_disj_int_one. tts_lemma min_absorb2: assumes "y \ U" and "x \ U" and "y \\<^sub>o\<^sub>w x" shows "local.min x y = y" is Orderings.min_absorb2. tts_lemma max_absorb1: assumes "y \ U" and "x \ U" and "y \\<^sub>o\<^sub>w x" shows "local.max x y = x" is Orderings.max_absorb1. - -tts_lemma finite_mono_remains_stable_implies_strict_prefix: - assumes "range f \ U" - and "finite (range f)" - and "on UNIV with (\\<^sub>o\<^sub>w) (\) : \mono\ f" - and "\n. f n = f (Suc n) \ f (Suc n) = f (Suc (Suc n))" - shows "\N. (\n\N. f N = f n) \ (\n\N. \m\N. m < n \ f m <\<^sub>o\<^sub>w f n)" - is Hilbert_Choice.finite_mono_remains_stable_implies_strict_prefix. - -tts_lemma incseq_imp_monoseq: - assumes "range X \ U" and "on UNIV with (\\<^sub>o\<^sub>w) (\) : \mono\ X" - shows "with (\\<^sub>o\<^sub>w) : \monoseq\ X" - is Topological_Spaces.incseq_imp_monoseq. tts_lemma decseq_imp_monoseq: assumes "range X \ U" and "on UNIV with (\\<^sub>o\<^sub>w) (\) : \antimono\ X" shows "with (\\<^sub>o\<^sub>w) : \monoseq\ X" is Topological_Spaces.decseq_imp_monoseq. -tts_lemma incseq_Suc_iff: - assumes "range f \ U" - shows "(on UNIV with (\\<^sub>o\<^sub>w) (\) : \mono\ f) = (\x. f x \\<^sub>o\<^sub>w f (Suc x))" - is Topological_Spaces.incseq_Suc_iff. - tts_lemma decseq_Suc_iff: assumes "range f \ U" shows "(on UNIV with (\\<^sub>o\<^sub>w) (\) : \antimono\ f) = (\x. f (Suc x) \\<^sub>o\<^sub>w f x)" is Topological_Spaces.decseq_Suc_iff. -tts_lemma incseq_const: - assumes "k \ U" - shows "on (UNIV::nat set) with (\\<^sub>o\<^sub>w) (\) : \mono\ (\x. k)" - is Topological_Spaces.incseq_const. - tts_lemma decseq_const: assumes "k \ U" shows "on (UNIV::nat set) with (\\<^sub>o\<^sub>w) (\) : \antimono\ (\x. k)" is Topological_Spaces.decseq_const. tts_lemma atMost_Int_atLeast: assumes "n \ U" shows "{..\<^sub>o\<^sub>wn} \ {n..\<^sub>o\<^sub>w} = {n}" is Set_Interval.atMost_Int_atLeast. -tts_lemma monoseq_iff: - assumes "range X \ U" - shows - "(with (\\<^sub>o\<^sub>w) : \monoseq\ X) = - ( - (on UNIV with (\\<^sub>o\<^sub>w) (\) : \mono\ X) \ - (on UNIV with (\\<^sub>o\<^sub>w) (\) : \antimono\ X) - )" - is Topological_Spaces.monoseq_iff. - tts_lemma monoseq_Suc: assumes "range X \ U" shows "(with (\\<^sub>o\<^sub>w) : \monoseq\ X) = ((\x. X x \\<^sub>o\<^sub>w X (Suc x)) \ (\x. X (Suc x) \\<^sub>o\<^sub>w X x))" is Topological_Spaces.monoseq_Suc. -tts_lemma incseq_SucI: - assumes "range X \ U" and "\n. X n \\<^sub>o\<^sub>w X (Suc n)" - shows "on UNIV with (\\<^sub>o\<^sub>w) (\) : \mono\ X" - is Topological_Spaces.incseq_SucI. - -tts_lemma incseq_SucD: - assumes "range A \ U" and "on UNIV with (\\<^sub>o\<^sub>w) (\) : \mono\ A" - shows "A i \\<^sub>o\<^sub>w A (Suc i)" - is Topological_Spaces.incseq_SucD. - tts_lemma decseq_SucI: assumes "range X \ U" and "\n. X (Suc n) \\<^sub>o\<^sub>w X n" shows "on UNIV with (\\<^sub>o\<^sub>w) (\) : \antimono\ X" is Topological_Spaces.decseq_SucI. tts_lemma decseq_SucD: assumes "range A \ U" and "on UNIV with (\\<^sub>o\<^sub>w) (\) : \antimono\ A" shows "A (Suc i) \\<^sub>o\<^sub>w A i" is Topological_Spaces.decseq_SucD. tts_lemma mono_SucI2: assumes "range X \ U" and "\x. X (Suc x) \\<^sub>o\<^sub>w X x" shows "with (\\<^sub>o\<^sub>w) : \monoseq\ X" is Topological_Spaces.mono_SucI2. tts_lemma mono_SucI1: assumes "range X \ U" and "\x. X x \\<^sub>o\<^sub>w X (Suc x)" shows "with (\\<^sub>o\<^sub>w) : \monoseq\ X" is Topological_Spaces.mono_SucI1. -tts_lemma incseqD: - assumes "range f \ U" - and "on UNIV with (\\<^sub>o\<^sub>w) (\) : \mono\ f" - and "(i::nat) \ j" - shows "f i \\<^sub>o\<^sub>w f j" - is Topological_Spaces.incseqD. - tts_lemma decseqD: assumes "range f \ U" and "on UNIV with (\\<^sub>o\<^sub>w) (\) : \antimono\ f" and "(i::nat) \ j" shows "f j \\<^sub>o\<^sub>w f i" is Topological_Spaces.decseqD. tts_lemma monoI2: assumes "range X \ U" and "\x y. x \ y \ X y \\<^sub>o\<^sub>w X x" shows "with (\\<^sub>o\<^sub>w) : \monoseq\ X" is Topological_Spaces.monoI2. tts_lemma monoI1: assumes "range X \ U" and "\x y. x \ y \ X x \\<^sub>o\<^sub>w X y" shows "with (\\<^sub>o\<^sub>w) : \monoseq\ X" is Topological_Spaces.monoI1. tts_lemma antimono_iff_le_Suc: assumes "range f \ U" shows "(on UNIV with (\\<^sub>o\<^sub>w) (\) : \antimono\ f) = (\x. f (Suc x) \\<^sub>o\<^sub>w f x)" is Nat.antimono_iff_le_Suc. -tts_lemma mono_iff_le_Suc: - assumes "range f \ U" - shows "(on UNIV with (\\<^sub>o\<^sub>w) (\) : \mono\ f) = (\x. f x \\<^sub>o\<^sub>w f (Suc x))" - is Nat.mono_iff_le_Suc. - -tts_lemma funpow_mono2: - assumes "\x\U. f x \ U" - and "x \ U" - and "y \ U" - and "on U with (\\<^sub>o\<^sub>w) (\\<^sub>o\<^sub>w) : \mono\ f" - and "i \ j" - and "x \\<^sub>o\<^sub>w y" - and "x \\<^sub>o\<^sub>w f x" - shows "(f ^^ i) x \\<^sub>o\<^sub>w (f ^^ j) y" - is Nat.funpow_mono2. - -tts_lemma funpow_mono: - assumes "\x\U. f x \ U" - and "A \ U" - and "B \ U" - and "on U with (\\<^sub>o\<^sub>w) (\\<^sub>o\<^sub>w) : \mono\ f" - and "A \\<^sub>o\<^sub>w B" - shows "(f ^^ n) A \\<^sub>o\<^sub>w (f ^^ n) B" - is Nat.funpow_mono. - end tts_context tts: (?'a to U) rewriting ctr_simps substituting order_ow_axioms eliminating through clarsimp begin tts_lemma ex_min_if_finite: assumes "S \ U" and "finite S" and "S \ {}" shows "\x\S. \ (\y\S. y <\<^sub>o\<^sub>w x)" is Lattices_Big.ex_min_if_finite. end tts_context tts: (?'a to U) sbterms: (\(\)::['a::order, 'a::order] \ bool\ to \(\\<^sub>o\<^sub>w)\) and (\(<)::['a::order, 'a::order] \ bool\ to \(<\<^sub>o\<^sub>w)\) substituting order_ow_axioms eliminating through clarsimp begin tts_lemma xt1: shows "\a = b; c <\<^sub>o\<^sub>w b\ \ c <\<^sub>o\<^sub>w a" "\b <\<^sub>o\<^sub>w a; b = c\ \ c <\<^sub>o\<^sub>w a" "\a = b; c \\<^sub>o\<^sub>w b\ \ c \\<^sub>o\<^sub>w a" "\b \\<^sub>o\<^sub>w a; b = c\ \ c \\<^sub>o\<^sub>w a" "\y \ U; x \ U; y \\<^sub>o\<^sub>w x; x \\<^sub>o\<^sub>w y\ \ x = y" "\y \ U; x \ U; z \ U; y \\<^sub>o\<^sub>w x; z \\<^sub>o\<^sub>w y\ \ z \\<^sub>o\<^sub>w x" "\y \ U; x \ U; z \ U; y <\<^sub>o\<^sub>w x; z \\<^sub>o\<^sub>w y\ \ z <\<^sub>o\<^sub>w x" "\y \ U; x \ U; z \ U; y \\<^sub>o\<^sub>w x; z <\<^sub>o\<^sub>w y\ \ z <\<^sub>o\<^sub>w x" "\b \ U; a \ U; b <\<^sub>o\<^sub>w a; a <\<^sub>o\<^sub>w b\ \ P" "\y \ U; x \ U; z \ U; y <\<^sub>o\<^sub>w x; z <\<^sub>o\<^sub>w y\ \ z <\<^sub>o\<^sub>w x" "\b \ U; a \ U; b \\<^sub>o\<^sub>w a; a \ b\ \ b <\<^sub>o\<^sub>w a" "\a \ U; b \ U; a \ b; b \\<^sub>o\<^sub>w a\ \ b <\<^sub>o\<^sub>w a" "\ b \ U; c \ U; a = f b; c <\<^sub>o\<^sub>w b; \x y. \x \ U; y \ U; y <\<^sub>o\<^sub>w x\ \ f y <\<^sub>o\<^sub>w f x \ \ f c <\<^sub>o\<^sub>w a" "\ b \ U; a \ U; b <\<^sub>o\<^sub>w a; f b = c; \x y. \x \ U; y \ U; y <\<^sub>o\<^sub>w x\ \ f y <\<^sub>o\<^sub>w f x \ \ c <\<^sub>o\<^sub>w f a" "\ b \ U; c \ U; a = f b; c \\<^sub>o\<^sub>w b; \x y. \x \ U; y \ U; y \\<^sub>o\<^sub>w x\ \ f y \\<^sub>o\<^sub>w f x \ \ f c \\<^sub>o\<^sub>w a" "\ b \ U; a \ U; b \\<^sub>o\<^sub>w a; f b = c; \x y. \x \ U; y \ U; y \\<^sub>o\<^sub>w x\ \ f y \\<^sub>o\<^sub>w f x \ \ c \\<^sub>o\<^sub>w f a" is Orderings.xt1. end -tts_context - tts: (?'a to U) and (?'b to \U\<^sub>2::'b set\) - rewriting ctr_simps - substituting order_ow_axioms - eliminating through (simp add: mono_ow_def) -begin - -tts_lemma coinduct3_mono_lemma: - assumes "\x\U. f x \ U\<^sub>2" - and "X \ U\<^sub>2" - and "B \ U\<^sub>2" - and "on U with (\) (\\<^sub>o\<^sub>w) : \mono\ f" - shows "on U with (\) (\\<^sub>o\<^sub>w) : \mono\ (\x. f x \ (X \ B))" - is Inductive.coinduct3_mono_lemma. - -end - end context ord_order_ow begin tts_context tts: (?'a to U\<^sub>2) and (?'b to U\<^sub>1) sbterms: (\(\)::[?'a::order, ?'a::order] \ bool\ to \(\\<^sub>o\<^sub>w\<^sub>.\<^sub>2)\) and (\(<)::[?'a::order, ?'a::order] \ bool\ to \(<\<^sub>o\<^sub>w\<^sub>.\<^sub>2)\) and (\(\)::[?'b::ord, ?'b::ord] \ bool\ to \(\\<^sub>o\<^sub>w\<^sub>.\<^sub>1)\) and (\(<)::[?'b::ord, ?'b::ord] \ bool\ to \(<\<^sub>o\<^sub>w\<^sub>.\<^sub>1)\) rewriting ctr_simps substituting ord\<^sub>2.order_ow_axioms eliminating through clarsimp begin tts_lemma xt2: assumes "\x\U\<^sub>1. f x \ U\<^sub>2" and "b \ U\<^sub>1" and "a \ U\<^sub>2" and "c \ U\<^sub>1" and "f b \\<^sub>o\<^sub>w\<^sub>.\<^sub>2 a" and "c \\<^sub>o\<^sub>w\<^sub>.\<^sub>1 b" and "\x y. \x \ U\<^sub>1; y \ U\<^sub>1; y \\<^sub>o\<^sub>w\<^sub>.\<^sub>1 x\ \ f y \\<^sub>o\<^sub>w\<^sub>.\<^sub>2 f x" shows "f c \\<^sub>o\<^sub>w\<^sub>.\<^sub>2 a" is Orderings.xt2. tts_lemma xt6: assumes "\x\U\<^sub>1. f x \ U\<^sub>2" and "b \ U\<^sub>1" and "a \ U\<^sub>2" and "c \ U\<^sub>1" and "f b \\<^sub>o\<^sub>w\<^sub>.\<^sub>2 a" and "c <\<^sub>o\<^sub>w\<^sub>.\<^sub>1 b" and "\x y. \x \ U\<^sub>1; y \ U\<^sub>1; y <\<^sub>o\<^sub>w\<^sub>.\<^sub>1 x\ \ f y <\<^sub>o\<^sub>w\<^sub>.\<^sub>2 f x" shows "f c <\<^sub>o\<^sub>w\<^sub>.\<^sub>2 a" is Orderings.xt6. end end context order_pair_ow begin tts_context tts: (?'a to U\<^sub>1) and (?'b to U\<^sub>2) rewriting ctr_simps substituting ord\<^sub>1.order_ow_axioms and ord\<^sub>2.order_ow_axioms eliminating through ( unfold strict_mono_ow_def mono_ow_def antimono_ow_def bdd_above_ow_def bdd_below_ow_def bdd_ow_def, clarsimp ) begin tts_lemma antimonoD: assumes "x \ U\<^sub>1" and "y \ U\<^sub>1" and "on U\<^sub>1 with (\\<^sub>o\<^sub>w\<^sub>.\<^sub>2) (\\<^sub>o\<^sub>w\<^sub>.\<^sub>1) : \antimono\ f" and "x \\<^sub>o\<^sub>w\<^sub>.\<^sub>1 y" shows "f y \\<^sub>o\<^sub>w\<^sub>.\<^sub>2 f x" is Orderings.antimonoD. -tts_lemma monoD: - assumes "x \ U\<^sub>1" - and "y \ U\<^sub>1" - and "on U\<^sub>1 with (\\<^sub>o\<^sub>w\<^sub>.\<^sub>2) (\\<^sub>o\<^sub>w\<^sub>.\<^sub>1) : \mono\ f" - and "x \\<^sub>o\<^sub>w\<^sub>.\<^sub>1 y" - shows "f x \\<^sub>o\<^sub>w\<^sub>.\<^sub>2 f y" - is Orderings.monoD. - -tts_lemma strict_monoD: - assumes "x \ U\<^sub>1" - and "y \ U\<^sub>1" - and "on U\<^sub>1 with (<\<^sub>o\<^sub>w\<^sub>.\<^sub>2) (<\<^sub>o\<^sub>w\<^sub>.\<^sub>1) : \strict_mono\ f" - and "x <\<^sub>o\<^sub>w\<^sub>.\<^sub>1 y" - shows "f x <\<^sub>o\<^sub>w\<^sub>.\<^sub>2 f y" - is Orderings.strict_monoD. - -tts_lemma strict_monoI: - assumes "\x y. \x \ U\<^sub>1; y \ U\<^sub>1; x <\<^sub>o\<^sub>w\<^sub>.\<^sub>1 y\ \ f x <\<^sub>o\<^sub>w\<^sub>.\<^sub>2 f y" - shows "on U\<^sub>1 with (<\<^sub>o\<^sub>w\<^sub>.\<^sub>2) (<\<^sub>o\<^sub>w\<^sub>.\<^sub>1) : \strict_mono\ f" - is Orderings.strict_monoI. - tts_lemma antimonoI: assumes "\x y. \x \ U\<^sub>1; y \ U\<^sub>1; x \\<^sub>o\<^sub>w\<^sub>.\<^sub>1 y\ \ f y \\<^sub>o\<^sub>w\<^sub>.\<^sub>2 f x" shows "on U\<^sub>1 with (\\<^sub>o\<^sub>w\<^sub>.\<^sub>2) (\\<^sub>o\<^sub>w\<^sub>.\<^sub>1) : \antimono\ f" is Orderings.antimonoI. -tts_lemma monoI: - assumes "\x y. \x \ U\<^sub>1; y \ U\<^sub>1; x \\<^sub>o\<^sub>w\<^sub>.\<^sub>1 y\ \ f x \\<^sub>o\<^sub>w\<^sub>.\<^sub>2 f y" - shows "on U\<^sub>1 with (\\<^sub>o\<^sub>w\<^sub>.\<^sub>2) (\\<^sub>o\<^sub>w\<^sub>.\<^sub>1) : \mono\ f" - is Orderings.monoI. - tts_lemma antimonoE: assumes "x \ U\<^sub>1" and "y \ U\<^sub>1" and "on U\<^sub>1 with (\\<^sub>o\<^sub>w\<^sub>.\<^sub>2) (\\<^sub>o\<^sub>w\<^sub>.\<^sub>1) : \antimono\ f" and "x \\<^sub>o\<^sub>w\<^sub>.\<^sub>1 y" and "f y \\<^sub>o\<^sub>w\<^sub>.\<^sub>2 f x \ thesis" shows thesis is Orderings.antimonoE. -tts_lemma monoE: - assumes "x \ U\<^sub>1" - and "y \ U\<^sub>1" - and "on U\<^sub>1 with (\\<^sub>o\<^sub>w\<^sub>.\<^sub>2) (\\<^sub>o\<^sub>w\<^sub>.\<^sub>1) : \mono\ f" - and "x \\<^sub>o\<^sub>w\<^sub>.\<^sub>1 y" - and "f x \\<^sub>o\<^sub>w\<^sub>.\<^sub>2 f y \ thesis" - shows thesis - is Orderings.monoE. - -tts_lemma strict_mono_mono: - assumes "\x\U\<^sub>1. f x \ U\<^sub>2" - and "on U\<^sub>1 with (<\<^sub>o\<^sub>w\<^sub>.\<^sub>2) (<\<^sub>o\<^sub>w\<^sub>.\<^sub>1) : \strict_mono\ f" - shows "on U\<^sub>1 with (\\<^sub>o\<^sub>w\<^sub>.\<^sub>2) (\\<^sub>o\<^sub>w\<^sub>.\<^sub>1) : \mono\ f" - is Orderings.strict_mono_mono. - tts_lemma bdd_below_image_antimono: assumes "\x\U\<^sub>1. f x \ U\<^sub>2" and "A \ U\<^sub>1" and "on U\<^sub>1 with (\\<^sub>o\<^sub>w\<^sub>.\<^sub>2) (\\<^sub>o\<^sub>w\<^sub>.\<^sub>1) : \antimono\ f" and "ord\<^sub>1.bdd_above A" shows "ord\<^sub>2.bdd_below (f ` A)" is Conditionally_Complete_Lattices.bdd_below_image_antimono. tts_lemma bdd_above_image_antimono: assumes "\x\U\<^sub>1. f x \ U\<^sub>2" and "A \ U\<^sub>1" and "on U\<^sub>1 with (\\<^sub>o\<^sub>w\<^sub>.\<^sub>2) (\\<^sub>o\<^sub>w\<^sub>.\<^sub>1) : \antimono\ f" and "ord\<^sub>1.bdd_below A" shows "ord\<^sub>2.bdd_above (f ` A)" is Conditionally_Complete_Lattices.bdd_above_image_antimono. -tts_lemma bdd_below_image_mono: - assumes "\x\U\<^sub>1. f x \ U\<^sub>2" - and "A \ U\<^sub>1" - and "on U\<^sub>1 with (\\<^sub>o\<^sub>w\<^sub>.\<^sub>2) (\\<^sub>o\<^sub>w\<^sub>.\<^sub>1) : \mono\ f" - and "ord\<^sub>1.bdd_below A" - shows "ord\<^sub>2.bdd_below (f ` A)" - is Conditionally_Complete_Lattices.bdd_below_image_mono. - -tts_lemma bdd_above_image_mono: - assumes "\x\U\<^sub>1. f x \ U\<^sub>2" - and "A \ U\<^sub>1" - and "on U\<^sub>1 with (\\<^sub>o\<^sub>w\<^sub>.\<^sub>2) (\\<^sub>o\<^sub>w\<^sub>.\<^sub>1) : \mono\ f" - and "ord\<^sub>1.bdd_above A" - shows "ord\<^sub>2.bdd_above (f ` A)" - is Conditionally_Complete_Lattices.bdd_above_image_mono. - -tts_lemma strict_mono_leD: - assumes "\x\U\<^sub>1. r x \ U\<^sub>2" - and "m \ U\<^sub>1" - and "n \ U\<^sub>1" - and "on U\<^sub>1 with (<\<^sub>o\<^sub>w\<^sub>.\<^sub>2) (<\<^sub>o\<^sub>w\<^sub>.\<^sub>1) : \strict_mono\ r" - and "m \\<^sub>o\<^sub>w\<^sub>.\<^sub>1 n" - shows "r m \\<^sub>o\<^sub>w\<^sub>.\<^sub>2 r n" - is Topological_Spaces.strict_mono_leD. - -tts_lemma mono_image_least: - assumes "\x\U\<^sub>1. f x \ U\<^sub>2" - and "m \ U\<^sub>1" - and "n \ U\<^sub>1" - and "m' \ U\<^sub>2" - and "n' \ U\<^sub>2" - and "on U\<^sub>1 with (\\<^sub>o\<^sub>w\<^sub>.\<^sub>2) (\\<^sub>o\<^sub>w\<^sub>.\<^sub>1) : \mono\ f" - and "f ` (on U\<^sub>1 with (\\<^sub>o\<^sub>w\<^sub>.\<^sub>1) (<\<^sub>o\<^sub>w\<^sub>.\<^sub>1) : {m..2 with (\\<^sub>o\<^sub>w\<^sub>.\<^sub>2) (<\<^sub>o\<^sub>w\<^sub>.\<^sub>2) : {m'..o\<^sub>w\<^sub>.\<^sub>1 n" - shows "f m = m'" - is Set_Interval.mono_image_least. - end tts_context tts: (?'a to U\<^sub>1) and (?'b to U\<^sub>2) sbterms: (\(\)::[?'a::order, ?'a::order] \ bool\ to \(\\<^sub>o\<^sub>w\<^sub>.\<^sub>1)\) and (\(<)::[?'a::order, ?'a::order] \ bool\ to \(<\<^sub>o\<^sub>w\<^sub>.\<^sub>1)\) and (\(\)::[?'b::order, ?'b::order] \ bool\ to \(\\<^sub>o\<^sub>w\<^sub>.\<^sub>2)\) and (\(<)::[?'b::order, ?'b::order] \ bool\ to \(<\<^sub>o\<^sub>w\<^sub>.\<^sub>2)\) rewriting ctr_simps substituting ord\<^sub>1.order_ow_axioms and ord\<^sub>2.order_ow_axioms eliminating through clarsimp begin tts_lemma xt3: assumes "b \ U\<^sub>1" and "a \ U\<^sub>1" and "c \ U\<^sub>2" and "\x\U\<^sub>1. f x \ U\<^sub>2" and "b \\<^sub>o\<^sub>w\<^sub>.\<^sub>1 a" and "c \\<^sub>o\<^sub>w\<^sub>.\<^sub>2 f b" and "\x y. \x \ U\<^sub>1; y \ U\<^sub>1; y \\<^sub>o\<^sub>w\<^sub>.\<^sub>1 x\ \ f y \\<^sub>o\<^sub>w\<^sub>.\<^sub>2 f x" shows "c \\<^sub>o\<^sub>w\<^sub>.\<^sub>2 f a" is Orderings.xt3. tts_lemma xt4: assumes "\x\U\<^sub>2. f x \ U\<^sub>1" and "b \ U\<^sub>2" and "a \ U\<^sub>1" and "c \ U\<^sub>2" and "f b <\<^sub>o\<^sub>w\<^sub>.\<^sub>1 a" and "c \\<^sub>o\<^sub>w\<^sub>.\<^sub>2 b" and "\x y. \x \ U\<^sub>2; y \ U\<^sub>2; y \\<^sub>o\<^sub>w\<^sub>.\<^sub>2 x\ \ f y \\<^sub>o\<^sub>w\<^sub>.\<^sub>1 f x" shows "f c <\<^sub>o\<^sub>w\<^sub>.\<^sub>1 a" is Orderings.xt4. tts_lemma xt5: assumes "b \ U\<^sub>1" and "a \ U\<^sub>1" and "c \ U\<^sub>2" and "\x\U\<^sub>1. f x \ U\<^sub>2" and "b <\<^sub>o\<^sub>w\<^sub>.\<^sub>1 a" and "c \\<^sub>o\<^sub>w\<^sub>.\<^sub>2 f b" and "\x y. \x \ U\<^sub>1; y \ U\<^sub>1; y <\<^sub>o\<^sub>w\<^sub>.\<^sub>1 x\ \ f y <\<^sub>o\<^sub>w\<^sub>.\<^sub>2 f x" shows "c <\<^sub>o\<^sub>w\<^sub>.\<^sub>2 f a" is Orderings.xt5. tts_lemma xt7: assumes "b \ U\<^sub>1" and "a \ U\<^sub>1" and "c \ U\<^sub>2" and "\x\U\<^sub>1. f x \ U\<^sub>2" and "b \\<^sub>o\<^sub>w\<^sub>.\<^sub>1 a" and "c <\<^sub>o\<^sub>w\<^sub>.\<^sub>2 f b" and "\x y. \x \ U\<^sub>1; y \ U\<^sub>1; y \\<^sub>o\<^sub>w\<^sub>.\<^sub>1 x\ \ f y \\<^sub>o\<^sub>w\<^sub>.\<^sub>2 f x" shows "c <\<^sub>o\<^sub>w\<^sub>.\<^sub>2 f a" is Orderings.xt7. tts_lemma xt8: assumes "\x\U\<^sub>2. f x \ U\<^sub>1" and "b \ U\<^sub>2" and "a \ U\<^sub>1" and "c \ U\<^sub>2" and "f b <\<^sub>o\<^sub>w\<^sub>.\<^sub>1 a" and "c <\<^sub>o\<^sub>w\<^sub>.\<^sub>2 b" and "\x y. \x \ U\<^sub>2; y \ U\<^sub>2; y <\<^sub>o\<^sub>w\<^sub>.\<^sub>2 x\ \ f y <\<^sub>o\<^sub>w\<^sub>.\<^sub>1 f x" shows "f c <\<^sub>o\<^sub>w\<^sub>.\<^sub>1 a" is Orderings.xt8. tts_lemma xt9: assumes "b \ U\<^sub>1" and "a \ U\<^sub>1" and "c \ U\<^sub>2" and "\x\U\<^sub>1. f x \ U\<^sub>2" and "b <\<^sub>o\<^sub>w\<^sub>.\<^sub>1 a" and "c <\<^sub>o\<^sub>w\<^sub>.\<^sub>2 f b" and "\x y. \x \ U\<^sub>1; y \ U\<^sub>1; y <\<^sub>o\<^sub>w\<^sub>.\<^sub>1 x\ \ f y <\<^sub>o\<^sub>w\<^sub>.\<^sub>2 f x" shows "c <\<^sub>o\<^sub>w\<^sub>.\<^sub>2 f a" is Orderings.xt9. end tts_context tts: (?'a to U\<^sub>1) and (?'b to U\<^sub>2) sbterms: (\(\)::[?'a::order, ?'a::order] \ bool\ to \(\\<^sub>o\<^sub>w\<^sub>.\<^sub>1)\) and (\(<)::[?'a::order, ?'a::order] \ bool\ to \(<\<^sub>o\<^sub>w\<^sub>.\<^sub>1)\) and (\(\)::[?'b::order, ?'b::order] \ bool\ to \(\\<^sub>o\<^sub>w\<^sub>.\<^sub>2)\) and (\(<)::[?'b::order, ?'b::order] \ bool\ to \(<\<^sub>o\<^sub>w\<^sub>.\<^sub>2)\) rewriting ctr_simps substituting ord\<^sub>1.order_ow_axioms and ord\<^sub>2.order_ow_axioms eliminating through simp begin tts_lemma order_less_subst1: assumes "a \ U\<^sub>1" and "\x\U\<^sub>2. f x \ U\<^sub>1" and "b \ U\<^sub>2" and "c \ U\<^sub>2" and "a <\<^sub>o\<^sub>w\<^sub>.\<^sub>1 f b" and "b <\<^sub>o\<^sub>w\<^sub>.\<^sub>2 c" and "\x y. \x \ U\<^sub>2; y \ U\<^sub>2; x <\<^sub>o\<^sub>w\<^sub>.\<^sub>2 y\ \ f x <\<^sub>o\<^sub>w\<^sub>.\<^sub>1 f y" shows "a <\<^sub>o\<^sub>w\<^sub>.\<^sub>1 f c" is Orderings.order_less_subst1. tts_lemma order_subst1: assumes "a \ U\<^sub>1" and "\x\U\<^sub>2. f x \ U\<^sub>1" and "b \ U\<^sub>2" and "c \ U\<^sub>2" and "a \\<^sub>o\<^sub>w\<^sub>.\<^sub>1 f b" and "b \\<^sub>o\<^sub>w\<^sub>.\<^sub>2 c" and "\x y. \x \ U\<^sub>2; y \ U\<^sub>2; x \\<^sub>o\<^sub>w\<^sub>.\<^sub>2 y\ \ f x \\<^sub>o\<^sub>w\<^sub>.\<^sub>1 f y" shows "a \\<^sub>o\<^sub>w\<^sub>.\<^sub>1 f c" is Orderings.order_subst1. end tts_context tts: (?'a to U\<^sub>1) and (?'c to U\<^sub>2) sbterms: (\(\)::[?'a::order, ?'a::order] \ bool\ to \(\\<^sub>o\<^sub>w\<^sub>.\<^sub>1)\) and (\(<)::[?'a::order, ?'a::order] \ bool\ to \(<\<^sub>o\<^sub>w\<^sub>.\<^sub>1)\) and (\(\)::[?'c::order, ?'c::order] \ bool\ to \(\\<^sub>o\<^sub>w\<^sub>.\<^sub>2)\) and (\(<)::[?'c::order, ?'c::order] \ bool\ to \(<\<^sub>o\<^sub>w\<^sub>.\<^sub>2)\) rewriting ctr_simps substituting ord\<^sub>1.order_ow_axioms and ord\<^sub>2.order_ow_axioms eliminating through simp begin tts_lemma order_less_le_subst2: assumes "a \ U\<^sub>1" and "b \ U\<^sub>1" and "\x\U\<^sub>1. f x \ U\<^sub>2" and "c \ U\<^sub>2" and "a <\<^sub>o\<^sub>w\<^sub>.\<^sub>1 b" and "f b \\<^sub>o\<^sub>w\<^sub>.\<^sub>2 c" and "\x y. \x \ U\<^sub>1; y \ U\<^sub>1; x <\<^sub>o\<^sub>w\<^sub>.\<^sub>1 y\ \ f x <\<^sub>o\<^sub>w\<^sub>.\<^sub>2 f y" shows "f a <\<^sub>o\<^sub>w\<^sub>.\<^sub>2 c" is Orderings.order_less_le_subst2. tts_lemma order_le_less_subst2: assumes "a \ U\<^sub>1" and "b \ U\<^sub>1" and "\x\U\<^sub>1. f x \ U\<^sub>2" and "c \ U\<^sub>2" and "a \\<^sub>o\<^sub>w\<^sub>.\<^sub>1 b" and "f b <\<^sub>o\<^sub>w\<^sub>.\<^sub>2 c" and "\x y. \x \ U\<^sub>1; y \ U\<^sub>1; x \\<^sub>o\<^sub>w\<^sub>.\<^sub>1 y\ \ f x \\<^sub>o\<^sub>w\<^sub>.\<^sub>2 f y" shows "f a <\<^sub>o\<^sub>w\<^sub>.\<^sub>2 c" is Orderings.order_le_less_subst2. tts_lemma order_less_subst2: assumes "a \ U\<^sub>1" and "b \ U\<^sub>1" and "\x\U\<^sub>1. f x \ U\<^sub>2" and "c \ U\<^sub>2" and "a <\<^sub>o\<^sub>w\<^sub>.\<^sub>1 b" and "f b <\<^sub>o\<^sub>w\<^sub>.\<^sub>2 c" and "\x y. \x \ U\<^sub>1; y \ U\<^sub>1; x <\<^sub>o\<^sub>w\<^sub>.\<^sub>1 y\ \ f x <\<^sub>o\<^sub>w\<^sub>.\<^sub>2 f y" shows "f a <\<^sub>o\<^sub>w\<^sub>.\<^sub>2 c" is Orderings.order_less_subst2. tts_lemma order_subst2: assumes "a \ U\<^sub>1" and "b \ U\<^sub>1" and "\x\U\<^sub>1. f x \ U\<^sub>2" and "c \ U\<^sub>2" and "a \\<^sub>o\<^sub>w\<^sub>.\<^sub>1 b" and "f b \\<^sub>o\<^sub>w\<^sub>.\<^sub>2 c" and "\x y. \x \ U\<^sub>1; y \ U\<^sub>1; x \\<^sub>o\<^sub>w\<^sub>.\<^sub>1 y\ \ f x \\<^sub>o\<^sub>w\<^sub>.\<^sub>2 f y" shows "f a \\<^sub>o\<^sub>w\<^sub>.\<^sub>2 c" is Orderings.order_subst2. end end subsection\Dense orders\ subsubsection\Definitions and common properties\ locale dense_order_ow = order_ow U le ls for U :: "'ao set" and le ls + assumes dense: "\ x \ U; y \ U; x <\<^sub>o\<^sub>w y \ \ (\z\U. x <\<^sub>o\<^sub>w z \ z <\<^sub>o\<^sub>w y)" subsubsection\Transfer rules\ context includes lifting_syntax begin lemma dense_order_transfer[transfer_rule]: assumes [transfer_rule]: "bi_unique A" "right_total A" shows "((A ===> A ===> (=)) ===> (A ===> A ===> (=)) ===> (=)) (dense_order_ow (Collect (Domainp A))) class.dense_order" unfolding dense_order_ow_def class.dense_order_def dense_order_ow_axioms_def class.dense_order_axioms_def apply transfer_prover_start apply transfer_step+ by simp end subsection\ Partial orders with the greatest element and partial orders with the least elements \ subsubsection\Definitions and common properties\ locale ordering_top_ow = ordering_ow U le ls for U :: "'ao set" and le ls + fixes top :: "'ao" ("\<^bold>\\<^sub>o\<^sub>w") assumes top_closed[simp]: "\<^bold>\\<^sub>o\<^sub>w \ U" assumes extremum[simp]: "a \ U \ a \<^bold>\\<^sub>o\<^sub>w \<^bold>\\<^sub>o\<^sub>w" begin notation top ("\<^bold>\\<^sub>o\<^sub>w") end locale bot_ow = fixes U :: "'ao set" and bot (\\\<^sub>o\<^sub>w\) assumes bot_closed[simp]: "\\<^sub>o\<^sub>w \ U" begin notation bot (\\\<^sub>o\<^sub>w\) end locale bot_pair_ow = ord\<^sub>1: bot_ow U\<^sub>1 bot\<^sub>1 + ord\<^sub>2: bot_ow U\<^sub>2 bot\<^sub>2 for U\<^sub>1 :: "'ao set" and bot\<^sub>1 and U\<^sub>2 :: "'bo set" and bot\<^sub>2 begin notation bot\<^sub>1 (\\\<^sub>o\<^sub>w\<^sub>.\<^sub>1\) notation bot\<^sub>2 (\\\<^sub>o\<^sub>w\<^sub>.\<^sub>2\) end locale order_bot_ow = order_ow U le ls + bot_ow U bot for U :: "'ao set" and bot le ls + assumes bot_least: "a \ U \ \\<^sub>o\<^sub>w \\<^sub>o\<^sub>w a" begin sublocale bot: ordering_top_ow U \(\\<^sub>o\<^sub>w)\ \(>\<^sub>o\<^sub>w)\ \\\<^sub>o\<^sub>w\ apply unfold_locales subgoal by simp subgoal by (simp add: bot_least) done no_notation le (infix "\<^bold>\\<^sub>o\<^sub>w" 50) and ls (infix "\<^bold><\<^sub>o\<^sub>w" 50) and top ("\<^bold>\\<^sub>o\<^sub>w") end locale order_bot_pair_ow = ord\<^sub>1: order_bot_ow U\<^sub>1 bot\<^sub>1 le\<^sub>1 ls\<^sub>1 + ord\<^sub>2: order_bot_ow U\<^sub>2 bot\<^sub>2 le\<^sub>2 ls\<^sub>2 for U\<^sub>1 :: "'ao set" and bot\<^sub>1 le\<^sub>1 ls\<^sub>1 and U\<^sub>2 :: "'bo set" and bot\<^sub>2 le\<^sub>2 ls\<^sub>2 begin sublocale bot_pair_ow .. sublocale order_pair_ow .. end locale top_ow = fixes U :: "'ao set" and top (\\\<^sub>o\<^sub>w\) assumes top_closed[simp]: "\\<^sub>o\<^sub>w \ U" begin notation top (\\\<^sub>o\<^sub>w\) end locale top_pair_ow = ord\<^sub>1: top_ow U\<^sub>1 top\<^sub>1 + ord\<^sub>2: top_ow U\<^sub>2 top\<^sub>2 for U\<^sub>1 :: "'ao set" and top\<^sub>1 and U\<^sub>2 :: "'bo set" and top\<^sub>2 begin notation top\<^sub>1 (\\\<^sub>o\<^sub>w\<^sub>.\<^sub>1\) notation top\<^sub>2 (\\\<^sub>o\<^sub>w\<^sub>.\<^sub>2\) end locale order_top_ow = order_ow U le ls + top_ow U top for U :: "'ao set" and le ls top + assumes top_greatest: "a \ U \ a \\<^sub>o\<^sub>w \\<^sub>o\<^sub>w" begin sublocale top: ordering_top_ow U \(\\<^sub>o\<^sub>w)\ \(<\<^sub>o\<^sub>w)\ \\\<^sub>o\<^sub>w\ apply unfold_locales subgoal by simp subgoal by (simp add: top_greatest) done no_notation le (infix "\<^bold>\\<^sub>o\<^sub>w" 50) and ls (infix "\<^bold><\<^sub>o\<^sub>w" 50) and top ("\<^bold>\\<^sub>o\<^sub>w") end locale order_top_pair_ow = ord\<^sub>1: order_top_ow U\<^sub>1 le\<^sub>1 ls\<^sub>1 top\<^sub>1 + ord\<^sub>2: order_top_ow U\<^sub>2 le\<^sub>2 ls\<^sub>2 top\<^sub>2 for U\<^sub>1 :: "'ao set" and le\<^sub>1 ls\<^sub>1 top\<^sub>1 and U\<^sub>2 :: "'bo set" and le\<^sub>2 ls\<^sub>2 top\<^sub>2 begin sublocale top_pair_ow .. sublocale order_pair_ow .. end subsubsection\Transfer rules\ context includes lifting_syntax begin lemma ordering_top_transfer[transfer_rule]: assumes [transfer_rule]: "bi_unique A" "right_total A" shows "((A ===> A ===> (=)) ===> (A ===> A ===> (=)) ===> A ===> (=)) (ordering_top_ow (Collect (Domainp A))) ordering_top" proof- let ?P = "((A ===> A ===> (=)) ===> (A ===> A ===> (=)) ===> A ===> (=))" let ?ordering_top_ow = "ordering_top_ow (Collect (Domainp A))" have "?P ?ordering_top_ow (\le ls ext. ext \ UNIV \ ordering_top le ls ext)" unfolding ordering_top_ow_def ordering_top_def ordering_top_ow_axioms_def ordering_top_axioms_def apply transfer_prover_start apply transfer_step+ by blast thus ?thesis by simp qed lemma order_bot_transfer[transfer_rule]: assumes [transfer_rule]: "bi_unique A" "right_total A" shows "(A ===> (A ===> A ===> (=)) ===> (A ===> A ===> (=)) ===> (=)) (order_bot_ow (Collect (Domainp A))) class.order_bot" proof- let ?P = "(A ===> (A ===> A ===> (=)) ===> (A ===> A ===> (=)) ===> (=))" let ?order_bot_ow = "order_bot_ow (Collect (Domainp A))" have "?P ?order_bot_ow (\bot le ls. bot \ UNIV \ class.order_bot bot le ls)" unfolding class.order_bot_def order_bot_ow_def class.order_bot_axioms_def order_bot_ow_axioms_def bot_ow_def apply transfer_prover_start apply transfer_step+ by blast thus ?thesis by simp qed lemma order_top_transfer[transfer_rule]: assumes [transfer_rule]: "bi_unique A" "right_total A" shows "((A ===> A ===> (=)) ===> (A ===> A ===> (=)) ===> A ===> (=)) (order_top_ow (Collect (Domainp A))) class.order_top" proof- let ?P = "((A ===> A ===> (=)) ===> (A ===> A ===> (=)) ===> A ===> (=))" let ?order_top_ow = "order_top_ow (Collect (Domainp A))" have "?P ?order_top_ow (\le ls top. top \ UNIV \ class.order_top le ls top)" unfolding class.order_top_def order_top_ow_def class.order_top_axioms_def order_top_ow_axioms_def top_ow_def apply transfer_prover_start apply transfer_step+ by blast thus ?thesis by simp qed end subsubsection\Relativization\ context ordering_top_ow begin tts_context tts: (?'a to U) rewriting ctr_simps substituting ordering_top_ow_axioms eliminating through simp applying [OF top_closed] begin tts_lemma extremum_uniqueI: assumes "\<^bold>\\<^sub>o\<^sub>w \<^bold>\\<^sub>o\<^sub>w \<^bold>\\<^sub>o\<^sub>w" shows "\<^bold>\\<^sub>o\<^sub>w = \<^bold>\\<^sub>o\<^sub>w" is ordering_top.extremum_uniqueI. tts_lemma extremum_unique: shows "(\<^bold>\\<^sub>o\<^sub>w \<^bold>\\<^sub>o\<^sub>w \<^bold>\\<^sub>o\<^sub>w) = (\<^bold>\\<^sub>o\<^sub>w = \<^bold>\\<^sub>o\<^sub>w)" is ordering_top.extremum_unique. tts_lemma extremum_strict: shows "\ \<^bold>\\<^sub>o\<^sub>w \<^bold><\<^sub>o\<^sub>w \<^bold>\\<^sub>o\<^sub>w" is ordering_top.extremum_strict. tts_lemma not_eq_extremum: shows "(\<^bold>\\<^sub>o\<^sub>w \ \<^bold>\\<^sub>o\<^sub>w) = (\<^bold>\\<^sub>o\<^sub>w \<^bold><\<^sub>o\<^sub>w \<^bold>\\<^sub>o\<^sub>w)" is ordering_top.not_eq_extremum. end end context order_bot_ow begin tts_context tts: (?'a to U) rewriting ctr_simps substituting order_bot_ow_axioms eliminating through simp begin tts_lemma bdd_above_bot: assumes "A \ U" shows "bdd_below A" is order_bot_class.bdd_below_bot. tts_lemma not_less_bot: assumes "a \ U" shows "\ a <\<^sub>o\<^sub>w \\<^sub>o\<^sub>w" is order_bot_class.not_less_bot. tts_lemma max_bot: assumes "x \ U" shows "max \\<^sub>o\<^sub>w x = x" is order_bot_class.max_bot. tts_lemma max_bot2: assumes "x \ U" shows "max x \\<^sub>o\<^sub>w = x" is order_bot_class.max_bot2. tts_lemma min_bot: assumes "x \ U" shows "min \\<^sub>o\<^sub>w x = \\<^sub>o\<^sub>w" is order_bot_class.min_bot. tts_lemma min_bot2: assumes "x \ U" shows "min x \\<^sub>o\<^sub>w = \\<^sub>o\<^sub>w" is order_bot_class.min_bot2. tts_lemma bot_unique: assumes "a \ U" shows "(a \\<^sub>o\<^sub>w \\<^sub>o\<^sub>w) = (a = \\<^sub>o\<^sub>w)" is order_bot_class.bot_unique. tts_lemma bot_less: assumes "a \ U" shows "(a \ \\<^sub>o\<^sub>w) = (\\<^sub>o\<^sub>w <\<^sub>o\<^sub>w a)" is order_bot_class.bot_less. tts_lemma atLeast_eq_UNIV_iff: assumes "x \ U" shows "({x..\<^sub>o\<^sub>w} = U) = (x = \\<^sub>o\<^sub>w)" is order_bot_class.atLeast_eq_UNIV_iff. tts_lemma le_bot: assumes "a \ U" and "a \\<^sub>o\<^sub>w \\<^sub>o\<^sub>w" shows "a = \\<^sub>o\<^sub>w" is order_bot_class.le_bot. end end context order_top_ow begin tts_context tts: (?'a to U) rewriting ctr_simps substituting order_top_ow_axioms eliminating through simp begin tts_lemma bdd_above_top: assumes "A \ U" shows "bdd_above A" is order_top_class.bdd_above_top. tts_lemma not_top_less: assumes "a \ U" shows "\ \\<^sub>o\<^sub>w <\<^sub>o\<^sub>w a" is order_top_class.not_top_less. tts_lemma max_top: assumes "x \ U" shows "max \\<^sub>o\<^sub>w x = \\<^sub>o\<^sub>w" is order_top_class.max_top. tts_lemma max_top2: assumes "x \ U" shows "max x \\<^sub>o\<^sub>w = \\<^sub>o\<^sub>w" is order_top_class.max_top2. tts_lemma min_top: assumes "x \ U" shows "min \\<^sub>o\<^sub>w x = x" is order_top_class.min_top. tts_lemma min_top2: assumes "x \ U" shows "min x \\<^sub>o\<^sub>w = x" is order_top_class.min_top2. tts_lemma top_unique: assumes "a \ U" shows "(\\<^sub>o\<^sub>w \\<^sub>o\<^sub>w a) = (a = \\<^sub>o\<^sub>w)" is order_top_class.top_unique. tts_lemma less_top: assumes "a \ U" shows "(a \ \\<^sub>o\<^sub>w) = (a <\<^sub>o\<^sub>w \\<^sub>o\<^sub>w)" is order_top_class.less_top. tts_lemma atMost_eq_UNIV_iff: assumes "x \ U" shows "({..\<^sub>o\<^sub>wx} = U) = (x = \\<^sub>o\<^sub>w)" is order_top_class.atMost_eq_UNIV_iff. tts_lemma top_le: assumes "a \ U" and "\\<^sub>o\<^sub>w \\<^sub>o\<^sub>w a" shows "a = \\<^sub>o\<^sub>w" is order_top_class.top_le. end end text\\newpage\ end \ No newline at end of file