diff --git a/thys/Collections/ICF/gen_algo/Algos.thy b/thys/Collections/ICF/gen_algo/Algos.thy --- a/thys/Collections/ICF/gen_algo/Algos.thy +++ b/thys/Collections/ICF/gen_algo/Algos.thy @@ -1,181 +1,181 @@ (* Title: Isabelle Collections Framework Author: Peter Lammich Maintainer: Peter Lammich *) section \\isaheader{More Generic Algorithms}\ theory Algos imports "../spec/SetSpec" "../spec/MapSpec" "../spec/ListSpec" begin text_raw \\label{thy:Algos}\ subsection "Injective Map to Naturals" text "Whether a set is an initial segment of the natural numbers" definition inatseg :: "nat set \ bool" where "inatseg s == \k. s = {i::nat. i 'm" where "map_to_nat s == snd (s.iterate s (\x (c,m). (c+1,m.update x c m)) (0,m.empty ()))" lemma map_to_nat_correct: assumes INV[simp]: "s.invar s" shows \ \All elements have got a number\ "dom (m.\ (map_to_nat s)) = s.\ s" (is ?T1) and \ \No two elements got the same number\ [rule_format]: "inj_on (m.\ (map_to_nat s)) (s.\ s)" (is ?T2) and \ \Numbering is inatseg\ [rule_format]: "inatseg (ran (m.\ (map_to_nat s)))" (is ?T3) and \ \The result satisfies the map invariant\ "m.invar (map_to_nat s)" (is ?T4) proof - have i_aux: "!!m S S' k v. \inj_on m S; S' = insert k S; v\ran m\ \ inj_on (m(k\v)) S'" apply (rule inj_onI) apply (simp split: if_split_asm) apply (simp add: ran_def) apply (simp add: ran_def) apply blast apply (blast dest: inj_onD) done have "?T1 \ ?T2 \ ?T3 \ ?T4" apply (unfold map_to_nat_def) apply (rule_tac I="\it (c,m). m.invar m \ dom (m.\ m) = s.\ s - it \ inj_on (m.\ m) (s.\ s - it) \ (ran (m.\ m) = {i. i) apply (simp add: m.empty_correct m.update_correct) apply (intro conjI) apply blast apply clarify - apply (rule_tac m="m.\ ba" and - k=x and v=aa and - S'="(s.\ s - (it - {x}))" and - S="(s.\ s - it)" + apply (rule_tac m2="m.\ ba" and + k2=x and v2=aa and + S'2="(s.\ s - (it - {x}))" and + S2="(s.\ s - it)" in i_aux) apply auto [3] apply auto [1] apply (case_tac \) apply (auto simp add: inatseg_def) done thus ?T1 ?T2 ?T3 ?T4 by auto qed end subsection "Map from Set" text "Build a map using a set of keys and a function to compute the values." locale it_dom_fun_to_map_loc = s: StdSet s_ops + m: StdMap m_ops for s_ops :: "('k,'s,'more1) set_ops_scheme" and m_ops :: "('k,'v,'m,'more2) map_ops_scheme" begin definition it_dom_fun_to_map :: "'s \ ('k \ 'v) \ 'm" where "it_dom_fun_to_map s f == s.iterate s (\k m. m.update_dj k (f k) m) (m.empty ())" lemma it_dom_fun_to_map_correct: assumes INV: "s.invar s" shows "m.\ (it_dom_fun_to_map s f) k = (if k \ s.\ s then Some (f k) else None)" (is ?G1) and "m.invar (it_dom_fun_to_map s f)" (is ?G2) proof - have "m.\ (it_dom_fun_to_map s f) k = (if k \ s.\ s then Some (f k) else None) \ m.invar (it_dom_fun_to_map s f)" unfolding it_dom_fun_to_map_def apply (rule s.iterate_rule_P[where I = "\it res. m.invar res \ (\k. m.\ res k = (if (k \ (s.\ s) - it) then Some (f k) else None))" ]) apply (simp add: INV) apply (simp add: m.empty_correct) apply (subgoal_tac "x\dom (m.\ \)") apply (auto simp: INV m.empty_correct m.update_dj_correct) [] apply auto [2] done thus ?G1 and ?G2 by auto qed end locale set_to_list_defs_loc = s: StdSetDefs s_ops + l: StdListDefs l_ops for s_ops :: "('x,'s,'more1) set_ops_scheme" and l_ops :: "('x,'l,'more2) list_ops_scheme" begin definition "g_set_to_listl s \ s.iterate s l.appendl (l.empty ())" definition "g_set_to_listr s \ s.iterate s l.appendr (l.empty ())" end locale set_to_list_loc = set_to_list_defs_loc s_ops l_ops + s: StdSet s_ops + l: StdList l_ops for s_ops :: "('x,'s,'more1) set_ops_scheme" and l_ops :: "('x,'l,'more2) list_ops_scheme" begin lemma g_set_to_listl_correct: assumes I: "s.invar s" shows "List.set (l.\ (g_set_to_listl s)) = s.\ s" and "l.invar (g_set_to_listl s)" and "distinct (l.\ (g_set_to_listl s))" using I apply - unfolding g_set_to_listl_def apply (rule_tac I="\it \. l.invar \ \ List.set (l.\ \) = it \ distinct (l.\ \)" in s.iterate_rule_insert_P, auto simp: l.correct)+ done lemma g_set_to_listr_correct: assumes I: "s.invar s" shows "List.set (l.\ (g_set_to_listr s)) = s.\ s" and "l.invar (g_set_to_listr s)" and "distinct (l.\ (g_set_to_listr s))" using I apply - unfolding g_set_to_listr_def apply (rule_tac I="\it \. l.invar \ \ List.set (l.\ \) = it \ distinct (l.\ \)" in s.iterate_rule_insert_P, auto simp: l.correct)+ done end end diff --git a/thys/Featherweight_OCL/UML_State.thy b/thys/Featherweight_OCL/UML_State.thy --- a/thys/Featherweight_OCL/UML_State.thy +++ b/thys/Featherweight_OCL/UML_State.thy @@ -1,1265 +1,1265 @@ (***************************************************************************** * Featherweight-OCL --- A Formal Semantics for UML-OCL Version OCL 2.5 * for the OMG Standard. * http://www.brucker.ch/projects/hol-testgen/ * * UML_State.thy --- State Operations and Objects. * This file is part of HOL-TestGen. * * Copyright (c) 2012-2015 Université Paris-Saclay, Univ. Paris-Sud, France * 2013-2015 IRT SystemX, France * * All rights reserved. * * Redistribution and use in source and binary forms, with or without * modification, are permitted provided that the following conditions are * met: * * * Redistributions of source code must retain the above copyright * notice, this list of conditions and the following disclaimer. * * * Redistributions in binary form must reproduce the above * copyright notice, this list of conditions and the following * disclaimer in the documentation and/or other materials provided * with the distribution. * * * Neither the name of the copyright holders nor the names of its * contributors may be used to endorse or promote products derived * from this software without specific prior written permission. * * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS * "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT * LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR * A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT * OWNER OR CONTRIBUTORS BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, * SPECIAL, EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT * LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, * DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY * THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT * (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE. ******************************************************************************) chapter\Formalization III: UML/OCL constructs: State Operations and Objects\ theory UML_State imports UML_Library begin no_notation None ("\") section\Introduction: States over Typed Object Universes\ text\\label{sec:universe} In the following, we will refine the concepts of a user-defined data-model (implied by a class-diagram) as well as the notion of $\state{}$ used in the previous section to much more detail. Surprisingly, even without a concrete notion of an objects and a universe of object representation, the generic infrastructure of state-related operations is fairly rich. \ subsection\Fundamental Properties on Objects: Core Referential Equality\ subsubsection\Definition\ text\Generic referential equality - to be used for instantiations with concrete object types ...\ definition StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t :: "('\,'a::{object,null})val \ ('\,'a)val \ ('\)Boolean" where "StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t x y \ \ \. if (\ x) \ = true \ \ (\ y) \ = true \ then if x \ = null \ y \ = null then \\x \ = null \ y \ = null\\ else \\(oid_of (x \)) = (oid_of (y \)) \\ else invalid \" subsubsection\Strictness and context passing\ lemma StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_strict1[simp,code_unfold] : "(StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t x invalid) = invalid" by(rule ext, simp add: StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def true_def false_def) lemma StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_strict2[simp,code_unfold] : "(StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t invalid x) = invalid" by(rule ext, simp add: StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def true_def false_def) lemma cp_StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t: "(StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t x y \) = (StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t (\_. x \) (\_. y \)) \" by(auto simp: StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def cp_valid[symmetric]) text_raw\\isatagafp\ lemmas cp0_StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t= cp_StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t[THEN allI[THEN allI[THEN allI[THEN cpI2]], of "StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t"]] lemmas cp_intro''[intro!,simp,code_unfold] = cp_intro'' cp_StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t[THEN allI[THEN allI[THEN allI[THEN cpI2]], of "StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t"]] text_raw\\endisatagafp\ subsection\Logic and Algebraic Layer on Object\ subsubsection\Validity and Definedness Properties\ text\We derive the usual laws on definedness for (generic) object equality:\ lemma StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_defargs: "\ \ (StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t x (y::('\,'a::{null,object})val))\ (\ \(\ x)) \ (\ \(\ y))" by(simp add: StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def OclValid_def true_def invalid_def bot_option_def split: bool.split_asm HOL.if_split_asm) lemma defined_StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_I: assumes val_x : "\ \ \ x" assumes val_x : "\ \ \ y" shows "\ \ \ (StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t x y)" apply(insert assms, simp add: StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def OclValid_def) by(subst cp_defined, simp add: true_def) lemma StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def_homo : "\(StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t x (y::('\,'a::{null,object})val)) = ((\ x) and (\ y))" oops (* sorry *) subsubsection\Symmetry\ lemma StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_sym : assumes x_val : "\ \ \ x" shows "\ \ StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t x x" by(simp add: StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def true_def OclValid_def x_val[simplified OclValid_def]) subsubsection\Behavior vs StrongEq\ text\It remains to clarify the role of the state invariant $\inv_\sigma(\sigma)$ mentioned above that states the condition that there is a ``one-to-one'' correspondence between object representations and $\oid$'s: $\forall \mathit{oid} \in \dom\ap \sigma\spot oid = \HolOclOidOf\ap \drop{\sigma(\mathit{oid})}$. This condition is also mentioned in~\<^cite>\\Annex A\ in "omg:ocl:2012"\ and goes back to \<^citet>\"richters:precise:2002"\; however, we state this condition as an invariant on states rather than a global axiom. It can, therefore, not be taken for granted that an $\oid$ makes sense both in pre- and post-states of OCL expressions. \ text\We capture this invariant in the predicate WFF :\ definition WFF :: "('\::object)st \ bool" where "WFF \ = ((\ x \ ran(heap(fst \)). \heap(fst \) (oid_of x)\ = x) \ (\ x \ ran(heap(snd \)). \heap(snd \) (oid_of x)\ = x))" text\It turns out that WFF is a key-concept for linking strict referential equality to logical equality: in well-formed states (i.e. those states where the self (oid-of) field contains the pointer to which the object is associated to in the state), referential equality coincides with logical equality.\ text\We turn now to the generic definition of referential equality on objects: Equality on objects in a state is reduced to equality on the references to these objects. As in HOL-OCL~\<^cite>\"brucker.ea:hol-ocl:2008" and "brucker.ea:hol-ocl-book:2006"\, we will store the reference of an object inside the object in a (ghost) field. By establishing certain invariants (``consistent state''), it can be assured that there is a ``one-to-one-correspondence'' of objects to their references---and therefore the definition below behaves as we expect.\ text\Generic Referential Equality enjoys the usual properties: (quasi) reflexivity, symmetry, transitivity, substitutivity for defined values. For type-technical reasons, for each concrete object type, the equality \\\ is defined by generic referential equality.\ theorem StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_vs_StrongEq: assumes WFF: "WFF \" and valid_x: "\ \(\ x)" and valid_y: "\ \(\ y)" and x_present_pre: "x \ \ ran (heap(fst \))" and y_present_pre: "y \ \ ran (heap(fst \))" and x_present_post:"x \ \ ran (heap(snd \))" and y_present_post:"y \ \ ran (heap(snd \))" (* x and y must be object representations that exist in either the pre or post state *) shows "(\ \ (StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t x y)) = (\ \ (x \ y))" apply(insert WFF valid_x valid_y x_present_pre y_present_pre x_present_post y_present_post) apply(auto simp: StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def OclValid_def WFF_def StrongEq_def true_def Ball_def) apply(erule_tac x="x \" in allE', simp_all) done theorem StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_vs_StrongEq': assumes WFF: "WFF \" and valid_x: "\ \(\ (x :: ('\::object,'\::{null,object})val))" and valid_y: "\ \(\ y)" and oid_preserve: "\x. x \ ran (heap(fst \)) \ x \ ran (heap(snd \)) \ H x \ \ \ oid_of (H x) = oid_of x" and xy_together: "x \ \ H ` ran (heap(fst \)) \ y \ \ H ` ran (heap(fst \)) \ x \ \ H ` ran (heap(snd \)) \ y \ \ H ` ran (heap(snd \))" (* x and y must be object representations that exist in either the pre or post state *) shows "(\ \ (StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t x y)) = (\ \ (x \ y))" apply(insert WFF valid_x valid_y xy_together) apply(simp add: WFF_def) apply(auto simp: StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def OclValid_def WFF_def StrongEq_def true_def Ball_def) by (metis foundation18' oid_preserve valid_x valid_y)+ text\So, if two object descriptions live in the same state (both pre or post), the referential equality on objects implies in a WFF state the logical equality.\ section\Operations on Object\ subsection\Initial States (for testing and code generation)\ definition \\<^sub>0 :: "('\)st" where "\\<^sub>0 \ (\heap=Map.empty, assocs = Map.empty\, \heap=Map.empty, assocs = Map.empty\)" subsection\OclAllInstances\ text\To denote OCL types occurring in OCL expressions syntactically---as, for example, as ``argument'' of \inlineocl{oclAllInstances()}---we use the inverses of the injection functions into the object universes; we show that this is a sufficient ``characterization.''\ definition OclAllInstances_generic :: "(('\::object) st \ '\ state) \ ('\::object \ '\) \ ('\, '\ option option) Set" where "OclAllInstances_generic fst_snd H = (\\. Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\ Some ` ((H ` ran (heap (fst_snd \))) - { None }) \\)" lemma OclAllInstances_generic_defined: "\ \ \ (OclAllInstances_generic pre_post H)" apply(simp add: defined_def OclValid_def OclAllInstances_generic_def false_def true_def bot_fun_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_fun_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def) apply(rule conjI) apply(rule notI, subst (asm) Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject, simp, (rule disjI2)+, metis bot_option_def option.distinct(1), (simp add: bot_option_def null_option_def)+)+ done lemma OclAllInstances_generic_init_empty: assumes [simp]: "\x. pre_post (x, x) = x" shows "\\<^sub>0 \ OclAllInstances_generic pre_post H \ Set{}" by(simp add: StrongEq_def OclAllInstances_generic_def OclValid_def \\<^sub>0_def mtSet_def) lemma represented_generic_objects_nonnull: assumes A: "\ \ ((OclAllInstances_generic pre_post (H::('\::object \ '\))) ->includes\<^sub>S\<^sub>e\<^sub>t(x))" shows "\ \ not(x \ null)" proof - have B: "\ \ \ (OclAllInstances_generic pre_post H)" by (simp add: OclAllInstances_generic_defined) have C: "\ \ \ x" by (metis OclIncludes.def_valid_then_def OclIncludes_valid_args_valid A foundation6) show ?thesis apply(insert A) apply(simp add: StrongEq_def OclValid_def OclNot_def null_def true_def OclIncludes_def B[simplified OclValid_def] C[simplified OclValid_def]) apply(simp add:OclAllInstances_generic_def) apply(erule contrapos_pn) apply(subst Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e.Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse, auto simp: null_fun_def null_option_def bot_option_def) done qed lemma represented_generic_objects_defined: assumes A: "\ \ ((OclAllInstances_generic pre_post (H::('\::object \ '\))) ->includes\<^sub>S\<^sub>e\<^sub>t(x))" shows "\ \ \ (OclAllInstances_generic pre_post H) \ \ \ \ x" by (metis OclAllInstances_generic_defined A[THEN represented_generic_objects_nonnull] OclIncludes.defined_args_valid A foundation16' foundation18 foundation24 foundation6) text\One way to establish the actual presence of an object representation in a state is:\ definition "is_represented_in_state fst_snd x H \ = (x \ \ (Some o H) ` ran (heap (fst_snd \)))" lemma represented_generic_objects_in_state: assumes A: "\ \ (OclAllInstances_generic pre_post H)->includes\<^sub>S\<^sub>e\<^sub>t(x)" shows "is_represented_in_state pre_post x H \" proof - have B: "(\ (OclAllInstances_generic pre_post H)) \ = true \" by(simp add: OclValid_def[symmetric] OclAllInstances_generic_defined) have C: "(\ x) \ = true \" by (metis OclValid_def UML_Set.OclIncludes_def assms bot_option_def option.distinct(1) true_def) have F: "Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\Some ` (H ` ran (heap (pre_post \)) - {None})\\) = \\Some ` (H ` ran (heap (pre_post \)) - {None})\\" by(subst Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e.Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse,simp_all add: bot_option_def) show ?thesis apply(insert A) apply(simp add: is_represented_in_state_def OclIncludes_def OclValid_def ran_def B C image_def true_def) apply(simp add: OclAllInstances_generic_def) apply(simp add: F) apply(simp add: ran_def) by(fastforce) qed lemma state_update_vs_allInstances_generic_empty: assumes [simp]: "\a. pre_post (mk a) = a" shows "(mk \heap=Map.empty, assocs=A\) \ OclAllInstances_generic pre_post Type \ Set{}" proof - have state_update_vs_allInstances_empty: "(OclAllInstances_generic pre_post Type) (mk \heap=Map.empty, assocs=A\) = Set{} (mk \heap=Map.empty, assocs=A\)" by(simp add: OclAllInstances_generic_def mtSet_def) show ?thesis apply(simp only: OclValid_def, subst StrictRefEq\<^sub>S\<^sub>e\<^sub>t.cp0, simp only: state_update_vs_allInstances_empty StrictRefEq\<^sub>S\<^sub>e\<^sub>t.refl_ext) apply(simp add: OclIf_def valid_def mtSet_def defined_def bot_fun_def null_fun_def null_option_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def) by(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject, (simp add: bot_option_def true_def)+) qed text\Here comes a couple of operational rules that allow to infer the value of \inlineisar+oclAllInstances+ from the context $\tau$. These rules are a special-case in the sense that they are the only rules that relate statements with \emph{different} $\tau$'s. For that reason, new concepts like ``constant contexts P'' are necessary (for which we do not elaborate an own theory for reasons of space limitations; in examples, we will prove resulting constraints straight forward by hand).\ lemma state_update_vs_allInstances_generic_including': assumes [simp]: "\a. pre_post (mk a) = a" assumes "\x. \' oid = Some x \ x = Object" and "Type Object \ None" shows "(OclAllInstances_generic pre_post Type) (mk \heap=\'(oid\Object), assocs=A\) = ((OclAllInstances_generic pre_post Type)->including\<^sub>S\<^sub>e\<^sub>t(\ _. \\ drop (Type Object) \\)) (mk \heap=\',assocs=A\)" proof - have drop_none : "\x. x \ None \ \\x\\ = x" by(case_tac x, simp+) have insert_diff : "\x S. insert \x\ (S - {None}) = (insert \x\ S) - {None}" by (metis insert_Diff_if option.distinct(1) singletonE) show ?thesis apply(simp add: UML_Set.OclIncluding_def OclAllInstances_generic_defined[simplified OclValid_def], simp add: OclAllInstances_generic_def) apply(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse, simp add: bot_option_def, simp add: comp_def, subst image_insert[symmetric], subst drop_none, simp add: assms) apply(case_tac "Type Object", simp add: assms, simp only:, subst insert_diff, drule sym, simp) apply(subgoal_tac "ran (\'(oid \ Object)) = insert Object (ran \')", simp) apply(case_tac "\ (\x. \' oid = Some x)") apply(rule ran_map_upd, simp) apply(simp, erule exE, frule assms, simp) apply(subgoal_tac "Object \ ran \'") prefer 2 apply(rule ranI, simp) by(subst insert_absorb, simp, metis fun_upd_apply) qed lemma state_update_vs_allInstances_generic_including: assumes [simp]: "\a. pre_post (mk a) = a" assumes "\x. \' oid = Some x \ x = Object" and "Type Object \ None" shows "(OclAllInstances_generic pre_post Type) (mk \heap=\'(oid\Object), assocs=A\) = ((\_. (OclAllInstances_generic pre_post Type) (mk \heap=\', assocs=A\))->including\<^sub>S\<^sub>e\<^sub>t(\ _. \\ drop (Type Object) \\)) (mk \heap=\'(oid\Object), assocs=A\)" apply(subst state_update_vs_allInstances_generic_including', (simp add: assms)+, subst UML_Set.OclIncluding.cp0, simp add: UML_Set.OclIncluding_def) apply(subst (1 3) cp_defined[symmetric], simp add: OclAllInstances_generic_defined[simplified OclValid_def]) apply(simp add: defined_def OclValid_def OclAllInstances_generic_def invalid_def bot_fun_def null_fun_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def) apply(subst (1 3) Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject) by(simp add: bot_option_def null_option_def)+ lemma state_update_vs_allInstances_generic_noincluding': assumes [simp]: "\a. pre_post (mk a) = a" assumes "\x. \' oid = Some x \ x = Object" and "Type Object = None" shows "(OclAllInstances_generic pre_post Type) (mk \heap=\'(oid\Object), assocs=A\) = (OclAllInstances_generic pre_post Type) (mk \heap=\', assocs=A\)" proof - have drop_none : "\x. x \ None \ \\x\\ = x" by(case_tac x, simp+) have insert_diff : "\x S. insert \x\ (S - {None}) = (insert \x\ S) - {None}" by (metis insert_Diff_if option.distinct(1) singletonE) show ?thesis apply(simp add: OclIncluding_def OclAllInstances_generic_defined[simplified OclValid_def] OclAllInstances_generic_def) apply(subgoal_tac "ran (\'(oid \ Object)) = insert Object (ran \')", simp add: assms) apply(case_tac "\ (\x. \' oid = Some x)") apply(rule ran_map_upd, simp) apply(simp, erule exE, frule assms, simp) apply(subgoal_tac "Object \ ran \'") prefer 2 apply(rule ranI, simp) apply(subst insert_absorb, simp) by (metis fun_upd_apply) qed theorem state_update_vs_allInstances_generic_ntc: assumes [simp]: "\a. pre_post (mk a) = a" assumes oid_def: "oid\dom \'" and non_type_conform: "Type Object = None " and cp_ctxt: "cp P" and const_ctxt: "\X. const X \ const (P X)" shows "(mk \heap=\'(oid\Object),assocs=A\ \ P (OclAllInstances_generic pre_post Type)) = (mk \heap=\', assocs=A\ \ P (OclAllInstances_generic pre_post Type))" (is "(?\ \ P ?\) = (?\' \ P ?\)") proof - have P_cp : "\x \. P x \ = P (\_. x \) \" by (metis (full_types) cp_ctxt cp_def) have A : "const (P (\_. ?\ ?\))" by(simp add: const_ctxt const_ss) have "(?\ \ P ?\) = (?\ \ \_. P ?\ ?\)" by(subst foundation23, rule refl) also have "... = (?\ \ \_. P (\_. ?\ ?\) ?\)" by(subst P_cp, rule refl) also have "... = (?\' \ \_. P (\_. ?\ ?\) ?\')" apply(simp add: OclValid_def) by(subst A[simplified const_def], subst const_true[simplified const_def], simp) finally have X: "(?\ \ P ?\) = (?\' \ \_. P (\_. ?\ ?\) ?\')" by simp show ?thesis apply(subst X) apply(subst foundation23[symmetric]) apply(rule StrongEq_L_subst3[OF cp_ctxt]) apply(simp add: OclValid_def StrongEq_def true_def) apply(rule state_update_vs_allInstances_generic_noincluding') by(insert oid_def, auto simp: non_type_conform) qed theorem state_update_vs_allInstances_generic_tc: assumes [simp]: "\a. pre_post (mk a) = a" assumes oid_def: "oid\dom \'" and type_conform: "Type Object \ None " and cp_ctxt: "cp P" and const_ctxt: "\X. const X \ const (P X)" shows "(mk \heap=\'(oid\Object),assocs=A\ \ P (OclAllInstances_generic pre_post Type)) = (mk \heap=\', assocs=A\ \ P ((OclAllInstances_generic pre_post Type) ->including\<^sub>S\<^sub>e\<^sub>t(\ _. \(Type Object)\)))" (is "(?\ \ P ?\) = (?\' \ P ?\')") proof - have P_cp : "\x \. P x \ = P (\_. x \) \" by (metis (full_types) cp_ctxt cp_def) have A : "const (P (\_. ?\ ?\))" by(simp add: const_ctxt const_ss) have "(?\ \ P ?\) = (?\ \ \_. P ?\ ?\)" by(subst foundation23, rule refl) also have "... = (?\ \ \_. P (\_. ?\ ?\) ?\)" by(subst P_cp, rule refl) also have "... = (?\' \ \_. P (\_. ?\ ?\) ?\')" apply(simp add: OclValid_def) by(subst A[simplified const_def], subst const_true[simplified const_def], simp) finally have X: "(?\ \ P ?\) = (?\' \ \_. P (\_. ?\ ?\) ?\')" by simp let ?allInstances = "OclAllInstances_generic pre_post Type" have "?allInstances ?\ = \_. ?allInstances ?\'->including\<^sub>S\<^sub>e\<^sub>t(\_.\\\Type Object\\\) ?\" apply(rule state_update_vs_allInstances_generic_including) by(insert oid_def, auto simp: type_conform) also have "... = ((\_. ?allInstances ?\')->including\<^sub>S\<^sub>e\<^sub>t(\_. (\_.\\\Type Object\\\) ?\') ?\')" by(subst const_OclIncluding[simplified const_def], simp+) also have "... = (?allInstances->including\<^sub>S\<^sub>e\<^sub>t(\ _. \Type Object\) ?\')" apply(subst UML_Set.OclIncluding.cp0[symmetric]) by(insert type_conform, auto) finally have Y : "?allInstances ?\ = (?allInstances->including\<^sub>S\<^sub>e\<^sub>t(\ _. \Type Object\) ?\')" by auto show ?thesis apply(subst X) apply(subst foundation23[symmetric]) apply(rule StrongEq_L_subst3[OF cp_ctxt]) apply(simp add: OclValid_def StrongEq_def Y true_def) done qed declare OclAllInstances_generic_def [simp] subsubsection\OclAllInstances (@post)\ definition OclAllInstances_at_post :: "('\ :: object \ '\) \ ('\, '\ option option) Set" ("_ .allInstances'(')") where "OclAllInstances_at_post = OclAllInstances_generic snd" lemma OclAllInstances_at_post_defined: "\ \ \ (H .allInstances())" unfolding OclAllInstances_at_post_def by(rule OclAllInstances_generic_defined) lemma "\\<^sub>0 \ H .allInstances() \ Set{}" unfolding OclAllInstances_at_post_def by(rule OclAllInstances_generic_init_empty, simp) lemma represented_at_post_objects_nonnull: assumes A: "\ \ (((H::('\::object \ '\)).allInstances()) ->includes\<^sub>S\<^sub>e\<^sub>t(x))" shows "\ \ not(x \ null)" by(rule represented_generic_objects_nonnull[OF A[simplified OclAllInstances_at_post_def]]) lemma represented_at_post_objects_defined: assumes A: "\ \ (((H::('\::object \ '\)).allInstances()) ->includes\<^sub>S\<^sub>e\<^sub>t(x))" shows "\ \ \ (H .allInstances()) \ \ \ \ x" unfolding OclAllInstances_at_post_def by(rule represented_generic_objects_defined[OF A[simplified OclAllInstances_at_post_def]]) text\One way to establish the actual presence of an object representation in a state is:\ lemma assumes A: "\ \ H .allInstances()->includes\<^sub>S\<^sub>e\<^sub>t(x)" shows "is_represented_in_state snd x H \" by(rule represented_generic_objects_in_state[OF A[simplified OclAllInstances_at_post_def]]) lemma state_update_vs_allInstances_at_post_empty: shows "(\, \heap=Map.empty, assocs=A\) \ Type .allInstances() \ Set{}" unfolding OclAllInstances_at_post_def by(rule state_update_vs_allInstances_generic_empty[OF snd_conv]) text\Here comes a couple of operational rules that allow to infer the value of \inlineisar+oclAllInstances+ from the context $\tau$. These rules are a special-case in the sense that they are the only rules that relate statements with \emph{different} $\tau$'s. For that reason, new concepts like ``constant contexts P'' are necessary (for which we do not elaborate an own theory for reasons of space limitations; in examples, we will prove resulting constraints straight forward by hand).\ lemma state_update_vs_allInstances_at_post_including': assumes "\x. \' oid = Some x \ x = Object" and "Type Object \ None" shows "(Type .allInstances()) (\, \heap=\'(oid\Object), assocs=A\) = ((Type .allInstances())->including\<^sub>S\<^sub>e\<^sub>t(\ _. \\ drop (Type Object) \\)) (\, \heap=\',assocs=A\)" unfolding OclAllInstances_at_post_def by(rule state_update_vs_allInstances_generic_including'[OF snd_conv], insert assms) lemma state_update_vs_allInstances_at_post_including: assumes "\x. \' oid = Some x \ x = Object" and "Type Object \ None" shows "(Type .allInstances()) (\, \heap=\'(oid\Object), assocs=A\) = ((\_. (Type .allInstances()) (\, \heap=\', assocs=A\))->including\<^sub>S\<^sub>e\<^sub>t(\ _. \\ drop (Type Object) \\)) (\, \heap=\'(oid\Object), assocs=A\)" unfolding OclAllInstances_at_post_def by(rule state_update_vs_allInstances_generic_including[OF snd_conv], insert assms) lemma state_update_vs_allInstances_at_post_noincluding': assumes "\x. \' oid = Some x \ x = Object" and "Type Object = None" shows "(Type .allInstances()) (\, \heap=\'(oid\Object), assocs=A\) = (Type .allInstances()) (\, \heap=\', assocs=A\)" unfolding OclAllInstances_at_post_def by(rule state_update_vs_allInstances_generic_noincluding'[OF snd_conv], insert assms) theorem state_update_vs_allInstances_at_post_ntc: assumes oid_def: "oid\dom \'" and non_type_conform: "Type Object = None " and cp_ctxt: "cp P" and const_ctxt: "\X. const X \ const (P X)" shows "((\, \heap=\'(oid\Object),assocs=A\) \ (P(Type .allInstances()))) = ((\, \heap=\', assocs=A\) \ (P(Type .allInstances())))" unfolding OclAllInstances_at_post_def by(rule state_update_vs_allInstances_generic_ntc[OF snd_conv], insert assms) theorem state_update_vs_allInstances_at_post_tc: assumes oid_def: "oid\dom \'" and type_conform: "Type Object \ None " and cp_ctxt: "cp P" and const_ctxt: "\X. const X \ const (P X)" shows "((\, \heap=\'(oid\Object),assocs=A\) \ (P(Type .allInstances()))) = ((\, \heap=\', assocs=A\) \ (P((Type .allInstances()) ->including\<^sub>S\<^sub>e\<^sub>t(\ _. \(Type Object)\))))" unfolding OclAllInstances_at_post_def by(rule state_update_vs_allInstances_generic_tc[OF snd_conv], insert assms) subsubsection\OclAllInstances (@pre)\ definition OclAllInstances_at_pre :: "('\ :: object \ '\) \ ('\, '\ option option) Set" ("_ .allInstances@pre'(')") where "OclAllInstances_at_pre = OclAllInstances_generic fst" lemma OclAllInstances_at_pre_defined: "\ \ \ (H .allInstances@pre())" unfolding OclAllInstances_at_pre_def by(rule OclAllInstances_generic_defined) lemma "\\<^sub>0 \ H .allInstances@pre() \ Set{}" unfolding OclAllInstances_at_pre_def by(rule OclAllInstances_generic_init_empty, simp) lemma represented_at_pre_objects_nonnull: assumes A: "\ \ (((H::('\::object \ '\)).allInstances@pre()) ->includes\<^sub>S\<^sub>e\<^sub>t(x))" shows "\ \ not(x \ null)" by(rule represented_generic_objects_nonnull[OF A[simplified OclAllInstances_at_pre_def]]) lemma represented_at_pre_objects_defined: assumes A: "\ \ (((H::('\::object \ '\)).allInstances@pre()) ->includes\<^sub>S\<^sub>e\<^sub>t(x))" shows "\ \ \ (H .allInstances@pre()) \ \ \ \ x" unfolding OclAllInstances_at_pre_def by(rule represented_generic_objects_defined[OF A[simplified OclAllInstances_at_pre_def]]) text\One way to establish the actual presence of an object representation in a state is:\ lemma assumes A: "\ \ H .allInstances@pre()->includes\<^sub>S\<^sub>e\<^sub>t(x)" shows "is_represented_in_state fst x H \" by(rule represented_generic_objects_in_state[OF A[simplified OclAllInstances_at_pre_def]]) lemma state_update_vs_allInstances_at_pre_empty: shows "(\heap=Map.empty, assocs=A\, \) \ Type .allInstances@pre() \ Set{}" unfolding OclAllInstances_at_pre_def by(rule state_update_vs_allInstances_generic_empty[OF fst_conv]) text\Here comes a couple of operational rules that allow to infer the value of \inlineisar+oclAllInstances@pre+ from the context $\tau$. These rules are a special-case in the sense that they are the only rules that relate statements with \emph{different} $\tau$'s. For that reason, new concepts like ``constant contexts P'' are necessary (for which we do not elaborate an own theory for reasons of space limitations; in examples, we will prove resulting constraints straight forward by hand).\ lemma state_update_vs_allInstances_at_pre_including': assumes "\x. \' oid = Some x \ x = Object" and "Type Object \ None" shows "(Type .allInstances@pre()) (\heap=\'(oid\Object), assocs=A\, \) = ((Type .allInstances@pre())->including\<^sub>S\<^sub>e\<^sub>t(\ _. \\ drop (Type Object) \\)) (\heap=\',assocs=A\, \)" unfolding OclAllInstances_at_pre_def by(rule state_update_vs_allInstances_generic_including'[OF fst_conv], insert assms) lemma state_update_vs_allInstances_at_pre_including: assumes "\x. \' oid = Some x \ x = Object" and "Type Object \ None" shows "(Type .allInstances@pre()) (\heap=\'(oid\Object), assocs=A\, \) = ((\_. (Type .allInstances@pre()) (\heap=\', assocs=A\, \))->including\<^sub>S\<^sub>e\<^sub>t(\ _. \\ drop (Type Object) \\)) (\heap=\'(oid\Object), assocs=A\, \)" unfolding OclAllInstances_at_pre_def by(rule state_update_vs_allInstances_generic_including[OF fst_conv], insert assms) lemma state_update_vs_allInstances_at_pre_noincluding': assumes "\x. \' oid = Some x \ x = Object" and "Type Object = None" shows "(Type .allInstances@pre()) (\heap=\'(oid\Object), assocs=A\, \) = (Type .allInstances@pre()) (\heap=\', assocs=A\, \)" unfolding OclAllInstances_at_pre_def by(rule state_update_vs_allInstances_generic_noincluding'[OF fst_conv], insert assms) theorem state_update_vs_allInstances_at_pre_ntc: assumes oid_def: "oid\dom \'" and non_type_conform: "Type Object = None " and cp_ctxt: "cp P" and const_ctxt: "\X. const X \ const (P X)" shows "((\heap=\'(oid\Object),assocs=A\, \) \ (P(Type .allInstances@pre()))) = ((\heap=\', assocs=A\, \) \ (P(Type .allInstances@pre())))" unfolding OclAllInstances_at_pre_def by(rule state_update_vs_allInstances_generic_ntc[OF fst_conv], insert assms) theorem state_update_vs_allInstances_at_pre_tc: assumes oid_def: "oid\dom \'" and type_conform: "Type Object \ None " and cp_ctxt: "cp P" and const_ctxt: "\X. const X \ const (P X)" shows "((\heap=\'(oid\Object),assocs=A\, \) \ (P(Type .allInstances@pre()))) = ((\heap=\', assocs=A\, \) \ (P((Type .allInstances@pre()) ->including\<^sub>S\<^sub>e\<^sub>t(\ _. \(Type Object)\))))" unfolding OclAllInstances_at_pre_def by(rule state_update_vs_allInstances_generic_tc[OF fst_conv], insert assms) subsubsection\@post or @pre\ theorem StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_vs_StrongEq'': assumes WFF: "WFF \" and valid_x: "\ \(\ (x :: ('\::object,'\::object option option)val))" and valid_y: "\ \(\ y)" and oid_preserve: "\x. x \ ran (heap(fst \)) \ x \ ran (heap(snd \)) \ oid_of (H x) = oid_of x" and xy_together: "\ \ ((H .allInstances()->includes\<^sub>S\<^sub>e\<^sub>t(x) and H .allInstances()->includes\<^sub>S\<^sub>e\<^sub>t(y)) or (H .allInstances@pre()->includes\<^sub>S\<^sub>e\<^sub>t(x) and H .allInstances@pre()->includes\<^sub>S\<^sub>e\<^sub>t(y)))" (* x and y must be object representations that exist in either the pre or post state *) shows "(\ \ (StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t x y)) = (\ \ (x \ y))" proof - have at_post_def : "\x. \ \ \ x \ \ \ \ (H .allInstances()->includes\<^sub>S\<^sub>e\<^sub>t(x))" apply(simp add: OclIncludes_def OclValid_def OclAllInstances_at_post_defined[simplified OclValid_def]) by(subst cp_defined, simp) have at_pre_def : "\x. \ \ \ x \ \ \ \ (H .allInstances@pre()->includes\<^sub>S\<^sub>e\<^sub>t(x))" apply(simp add: OclIncludes_def OclValid_def OclAllInstances_at_pre_defined[simplified OclValid_def]) by(subst cp_defined, simp) have F: "Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\Some ` (H ` ran (heap (fst \)) - {None})\\) = \\Some ` (H ` ran (heap (fst \)) - {None})\\" by(subst Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e.Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse,simp_all add: bot_option_def) have F': "Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\Some ` (H ` ran (heap (snd \)) - {None})\\) = \\Some ` (H ` ran (heap (snd \)) - {None})\\" by(subst Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e.Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse,simp_all add: bot_option_def) show ?thesis apply(rule StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_vs_StrongEq'[OF WFF valid_x valid_y, where H = "Some o H"]) apply(subst oid_preserve[symmetric], simp, simp add: oid_of_option_def) apply(insert xy_together, subst (asm) foundation11, metis at_post_def defined_and_I valid_x valid_y, metis at_pre_def defined_and_I valid_x valid_y) apply(erule disjE) by(drule foundation5, simp add: OclAllInstances_at_pre_def OclAllInstances_at_post_def OclValid_def OclIncludes_def true_def F F' valid_x[simplified OclValid_def] valid_y[simplified OclValid_def] bot_option_def split: if_split_asm, simp add: comp_def image_def, fastforce)+ qed subsection\OclIsNew, OclIsDeleted, OclIsMaintained, OclIsAbsent\ definition OclIsNew:: "('\, '\::{null,object})val \ ('\)Boolean" ("(_).oclIsNew'(')") where "X .oclIsNew() \ (\\ . if (\ X) \ = true \ then \\oid_of (X \) \ dom(heap(fst \)) \ oid_of (X \) \ dom(heap(snd \))\\ else invalid \)" text\The following predicates --- which are not part of the OCL standard descriptions --- complete the goal of \inlineisar+oclIsNew+ by describing where an object belongs. \ definition OclIsDeleted:: "('\, '\::{null,object})val \ ('\)Boolean" ("(_).oclIsDeleted'(')") where "X .oclIsDeleted() \ (\\ . if (\ X) \ = true \ then \\oid_of (X \) \ dom(heap(fst \)) \ oid_of (X \) \ dom(heap(snd \))\\ else invalid \)" definition OclIsMaintained:: "('\, '\::{null,object})val \ ('\)Boolean"("(_).oclIsMaintained'(')") where "X .oclIsMaintained() \ (\\ . if (\ X) \ = true \ then \\oid_of (X \) \ dom(heap(fst \)) \ oid_of (X \) \ dom(heap(snd \))\\ else invalid \)" definition OclIsAbsent:: "('\, '\::{null,object})val \ ('\)Boolean" ("(_).oclIsAbsent'(')") where "X .oclIsAbsent() \ (\\ . if (\ X) \ = true \ then \\oid_of (X \) \ dom(heap(fst \)) \ oid_of (X \) \ dom(heap(snd \))\\ else invalid \)" lemma state_split : "\ \ \ X \ \ \ (X .oclIsNew()) \ \ \ (X .oclIsDeleted()) \ \ \ (X .oclIsMaintained()) \ \ \ (X .oclIsAbsent())" by(simp add: OclIsDeleted_def OclIsNew_def OclIsMaintained_def OclIsAbsent_def OclValid_def true_def, blast) lemma notNew_vs_others : "\ \ \ X \ (\ \ \ (X .oclIsNew())) = (\ \ (X .oclIsDeleted()) \ \ \ (X .oclIsMaintained()) \ \ \ (X .oclIsAbsent()))" by(simp add: OclIsDeleted_def OclIsNew_def OclIsMaintained_def OclIsAbsent_def OclNot_def OclValid_def true_def, blast) subsection\OclIsModifiedOnly\ subsubsection\Definition\ text\The following predicate---which is not part of the OCL standard---provides a simple, but powerful means to describe framing conditions. For any formal approach, be it animation of OCL contracts, test-case generation or die-hard theorem proving, the specification of the part of a system transition that \emph{does not change} is of primordial importance. The following operator establishes the equality between old and new objects in the state (provided that they exist in both states), with the exception of those objects.\ definition OclIsModifiedOnly ::"('\::object,'\::{null,object})Set \ '\ Boolean" ("_->oclIsModifiedOnly'(')") where "X->oclIsModifiedOnly() \ (\(\,\'). let X' = (oid_of ` \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e(X(\,\'))\\); S = ((dom (heap \) \ dom (heap \')) - X') in if (\ X) (\,\') = true (\,\') \ (\x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e(X(\,\'))\\. x \ null) then \\\ x \ S. (heap \) x = (heap \') x\\ else invalid (\,\'))" subsubsection\Execution with Invalid or Null or Null Element as Argument\ lemma "invalid->oclIsModifiedOnly() = invalid" by(simp add: OclIsModifiedOnly_def) lemma "null->oclIsModifiedOnly() = invalid" by(simp add: OclIsModifiedOnly_def) lemma assumes X_null : "\ \ X->includes\<^sub>S\<^sub>e\<^sub>t(null)" shows "\ \ X->oclIsModifiedOnly() \ invalid" apply(insert X_null, simp add : OclIncludes_def OclIsModifiedOnly_def StrongEq_def OclValid_def true_def) apply(case_tac \, simp) apply(simp split: if_split_asm) by(simp add: null_fun_def, blast) subsubsection\Context Passing\ lemma cp_OclIsModifiedOnly : "X->oclIsModifiedOnly() \ = (\_. X \)->oclIsModifiedOnly() \" by(simp only: OclIsModifiedOnly_def, case_tac \, simp only:, subst cp_defined, simp) subsection\OclSelf\ text\The following predicate---which is not part of the OCL standard---explicitly retrieves in the pre or post state the original OCL expression given as argument.\ definition [simp]: "OclSelf x H fst_snd = (\\ . if (\ x) \ = true \ then if oid_of (x \) \ dom(heap(fst \)) \ oid_of (x \) \ dom(heap (snd \)) then H \(heap(fst_snd \))(oid_of (x \))\ else invalid \ else invalid \)" definition OclSelf_at_pre :: "('\::object,'\::{null,object})val \ ('\ \ '\) \ ('\::object,'\::{null,object})val" ("(_)@pre(_)") where "x @pre H = OclSelf x H fst" definition OclSelf_at_post :: "('\::object,'\::{null,object})val \ ('\ \ '\) \ ('\::object,'\::{null,object})val" ("(_)@post(_)") where "x @post H = OclSelf x H snd" subsection\Framing Theorem\ lemma all_oid_diff: assumes def_x : "\ \ \ x" assumes def_X : "\ \ \ X" assumes def_X' : "\x. x \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\ \ x \ null" defines "P \ (\a. not (StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t x a))" shows "(\ \ X->forAll\<^sub>S\<^sub>e\<^sub>t(a| P a)) = (oid_of (x \) \ oid_of ` \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\)" proof - have P_null_bot: "\b. b = null \ b = \ \ \ (\x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\. P (\(_:: 'a state \ 'a state). x) \ = b \)" apply(erule disjE) apply(simp, rule ballI, simp add: P_def StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def, rename_tac x', subst cp_OclNot, simp, subgoal_tac "x \ \ null \ x' \ null", simp, simp add: OclNot_def null_fun_def null_option_def bot_option_def bot_fun_def invalid_def, ( metis def_X' def_x foundation16[THEN iffD1] | (metis bot_fun_def OclValid_def Set_inv_lemma def_X def_x defined_def valid_def, metis def_X' def_x foundation16[THEN iffD1])))+ done have not_inj : "\x y. ((not x) \ = (not y) \) = (x \ = y \)" by (metis foundation21 foundation22) have P_false : "\x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\. P (\_. x) \ = false \ \ oid_of (x \) \ oid_of ` \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\" apply(erule bexE, rename_tac x') apply(simp add: P_def) apply(simp only: OclNot3[symmetric], simp only: not_inj) apply(simp add: StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def split: if_split_asm) apply(subgoal_tac "x \ \ null \ x' \ null", simp) apply (metis (mono_tags) drop.simps def_x foundation16[THEN iffD1] true_def) by(simp add: invalid_def bot_option_def true_def)+ have P_true : "\x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\. P (\_. x) \ = true \ \ oid_of (x \) \ oid_of ` \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\" apply(subgoal_tac "\x'\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\. oid_of x' \ oid_of (x \)") apply (metis imageE) apply(rule ballI, drule_tac x = "x'" in ballE) prefer 3 apply assumption apply(simp add: P_def) apply(simp only: OclNot4[symmetric], simp only: not_inj) apply(simp add: StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def false_def split: if_split_asm) apply(subgoal_tac "x \ \ null \ x' \ null", simp) apply (metis def_X' def_x foundation16[THEN iffD1]) by(simp add: invalid_def bot_option_def false_def)+ have bool_split : "\x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\. P (\_. x) \ \ null \ \ \x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\. P (\_. x) \ \ \ \ \ \x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\. P (\_. x) \ \ false \ \ \x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\. P (\_. x) \ = true \" apply(rule ballI) apply(drule_tac x = x in ballE) prefer 3 apply assumption apply(drule_tac x = x in ballE) prefer 3 apply assumption apply(drule_tac x = x in ballE) prefer 3 apply assumption apply (metis (full_types) bot_fun_def OclNot4 OclValid_def foundation16 foundation9 not_inj null_fun_def) by(fast+) show ?thesis apply(subst OclForall_rep_set_true[OF def_X], simp add: OclValid_def) apply(rule iffI, simp add: P_true) by (metis P_false P_null_bot bool_split) qed theorem framing: assumes modifiesclause:"\ \ (X->excluding\<^sub>S\<^sub>e\<^sub>t(x))->oclIsModifiedOnly()" and oid_is_typerepr : "\ \ X->forAll\<^sub>S\<^sub>e\<^sub>t(a| not (StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t x a))" shows "\ \ (x @pre P \ (x @post P))" apply(case_tac "\ \ \ x") proof - show "\ \ \ x \ ?thesis" proof - assume def_x : "\ \ \ x" show ?thesis proof - have def_X : "\ \ \ X" apply(insert oid_is_typerepr, simp add: UML_Set.OclForall_def OclValid_def split: if_split_asm) by(simp add: bot_option_def true_def) have def_X' : "\x. x \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\ \ x \ null" apply(insert modifiesclause, simp add: OclIsModifiedOnly_def OclValid_def split: if_split_asm) apply(case_tac \, simp split: if_split_asm) apply(simp add: UML_Set.OclExcluding_def split: if_split_asm) apply(subst (asm) (2) Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse) apply(simp, (rule disjI2)+) apply (metis (opaque_lifting, mono_tags) Diff_iff Set_inv_lemma def_X) apply(simp) apply(erule ballE[where P = "\x. x \ null"]) apply(assumption) apply(simp) apply (metis (opaque_lifting, no_types) def_x foundation16[THEN iffD1]) apply (metis (opaque_lifting, no_types) OclValid_def def_X def_x foundation20 OclExcluding_valid_args_valid OclExcluding_valid_args_valid'') by(simp add: invalid_def bot_option_def) have oid_is_typerepr : "oid_of (x \) \ oid_of ` \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\" by(rule all_oid_diff[THEN iffD1, OF def_x def_X def_X' oid_is_typerepr]) show ?thesis apply(simp add: StrongEq_def OclValid_def true_def OclSelf_at_pre_def OclSelf_at_post_def def_x[simplified OclValid_def]) apply(rule conjI, rule impI) apply(rule_tac f = "\x. P \x\" in arg_cong) apply(insert modifiesclause[simplified OclIsModifiedOnly_def OclValid_def]) apply(case_tac \, rename_tac \ \', simp split: if_split_asm) apply(subst (asm) (2) UML_Set.OclExcluding_def) apply(drule foundation5[simplified OclValid_def true_def], simp) apply(subst (asm) Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse, simp) apply(rule disjI2)+ apply (metis (opaque_lifting, no_types) DiffD1 OclValid_def Set_inv_lemma def_x foundation16 foundation18') apply(simp) apply(erule_tac x = "oid_of (x (\, \'))" in ballE) apply simp+ apply (metis (opaque_lifting, no_types) DiffD1 image_iff image_insert insert_Diff_single insert_absorb oid_is_typerepr) apply(simp add: invalid_def bot_option_def)+ by blast qed qed qed(simp add: OclSelf_at_post_def OclSelf_at_pre_def OclValid_def StrongEq_def true_def)+ text\As corollary, the framing property can be expressed with only the strong equality as comparison operator.\ theorem framing': assumes wff : "WFF \" assumes modifiesclause:"\ \ (X->excluding\<^sub>S\<^sub>e\<^sub>t(x))->oclIsModifiedOnly()" and oid_is_typerepr : "\ \ X->forAll\<^sub>S\<^sub>e\<^sub>t(a| not (x \ a))" and oid_preserve: "\x. x \ ran (heap(fst \)) \ x \ ran (heap(snd \)) \ oid_of (H x) = oid_of x" and xy_together: "\ \ X->forAll\<^sub>S\<^sub>e\<^sub>t(y | (H .allInstances()->includes\<^sub>S\<^sub>e\<^sub>t(x) and H .allInstances()->includes\<^sub>S\<^sub>e\<^sub>t(y)) or (H .allInstances@pre()->includes\<^sub>S\<^sub>e\<^sub>t(x) and H .allInstances@pre()->includes\<^sub>S\<^sub>e\<^sub>t(y)))" shows "\ \ (x @pre P \ (x @post P))" proof - have def_X : "\ \ \ X" apply(insert oid_is_typerepr, simp add: UML_Set.OclForall_def OclValid_def split: if_split_asm) by(simp add: bot_option_def true_def) show ?thesis apply(case_tac "\ \ \ x", drule foundation20) apply(rule framing[OF modifiesclause]) apply(rule OclForall_cong'[OF _ oid_is_typerepr xy_together], rename_tac y) apply(cut_tac Set_inv_lemma'[OF def_X]) prefer 2 apply assumption apply(rule OclNot_contrapos_nn, simp add: StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def) apply(simp add: OclValid_def, subst cp_defined, simp, assumption) apply(rule StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_vs_StrongEq''[THEN iffD1, OF wff _ _ oid_preserve], assumption+) by(simp add: OclSelf_at_post_def OclSelf_at_pre_def OclValid_def StrongEq_def true_def)+ qed subsection\Miscellaneous\ lemma pre_post_new: "\ \ (x .oclIsNew()) \ \ (\ \ \(x @pre H1)) \ \ (\ \ \(x @post H2))" by(simp add: OclIsNew_def OclSelf_at_pre_def OclSelf_at_post_def OclValid_def StrongEq_def true_def false_def bot_option_def invalid_def bot_fun_def valid_def split: if_split_asm) lemma pre_post_old: "\ \ (x .oclIsDeleted()) \ \ (\ \ \(x @pre H1)) \ \ (\ \ \(x @post H2))" by(simp add: OclIsDeleted_def OclSelf_at_pre_def OclSelf_at_post_def OclValid_def StrongEq_def true_def false_def bot_option_def invalid_def bot_fun_def valid_def split: if_split_asm) lemma pre_post_absent: "\ \ (x .oclIsAbsent()) \ \ (\ \ \(x @pre H1)) \ \ (\ \ \(x @post H2))" by(simp add: OclIsAbsent_def OclSelf_at_pre_def OclSelf_at_post_def OclValid_def StrongEq_def true_def false_def bot_option_def invalid_def bot_fun_def valid_def split: if_split_asm) lemma pre_post_maintained: "(\ \ \(x @pre H1) \ \ \ \(x @post H2)) \ \ \ (x .oclIsMaintained())" by(simp add: OclIsMaintained_def OclSelf_at_pre_def OclSelf_at_post_def OclValid_def StrongEq_def true_def false_def bot_option_def invalid_def bot_fun_def valid_def split: if_split_asm) lemma pre_post_maintained': "\ \ (x .oclIsMaintained()) \ (\ \ \(x @pre (Some o H1)) \ \ \ \(x @post (Some o H2)))" by(simp add: OclIsMaintained_def OclSelf_at_pre_def OclSelf_at_post_def OclValid_def StrongEq_def true_def false_def bot_option_def invalid_def bot_fun_def valid_def split: if_split_asm) lemma framing_same_state: "(\, \) \ (x @pre H \ (x @post H))" by(simp add: OclSelf_at_pre_def OclSelf_at_post_def OclValid_def StrongEq_def) section\Accessors on Object\ subsection\Definition\ definition "select_object mt incl smash deref l = smash (foldl incl mt (map deref l)) \ \smash returns null with \mt\ in input (in this case, object contains null pointer)\" text\The continuation \f\ is usually instantiated with a smashing function which is either the identity @{term id} or, for \inlineocl{0..1} cardinalities of associations, the @{term OclANY}-selector which also handles the @{term null}-cases appropriately. A standard use-case for this combinator is for example:\ term "(select_object mtSet UML_Set.OclIncluding UML_Set.OclANY f l oid )::('\, 'a::null)val" definition "select_object\<^sub>S\<^sub>e\<^sub>t = select_object mtSet UML_Set.OclIncluding id" definition "select_object_any0\<^sub>S\<^sub>e\<^sub>t f s_set = UML_Set.OclANY (select_object\<^sub>S\<^sub>e\<^sub>t f s_set)" definition "select_object_any\<^sub>S\<^sub>e\<^sub>t f s_set = (let s = select_object\<^sub>S\<^sub>e\<^sub>t f s_set in if s->size\<^sub>S\<^sub>e\<^sub>t() \ \ then s->any\<^sub>S\<^sub>e\<^sub>t() else \ endif)" definition "select_object\<^sub>S\<^sub>e\<^sub>q = select_object mtSequence UML_Sequence.OclIncluding id" definition "select_object_any\<^sub>S\<^sub>e\<^sub>q f s_set = UML_Sequence.OclANY (select_object\<^sub>S\<^sub>e\<^sub>q f s_set)" definition "select_object\<^sub>P\<^sub>a\<^sub>i\<^sub>r f1 f2 = (\(a,b). OclPair (f1 a) (f2 b))" subsection\Validity and Definedness Properties\ lemma select_fold_exec\<^sub>S\<^sub>e\<^sub>q: assumes "list_all (\f. (\ \ \ f)) l" shows "\\Rep_Sequence\<^sub>b\<^sub>a\<^sub>s\<^sub>e (foldl UML_Sequence.OclIncluding Sequence{} l \)\\ = List.map (\f. f \) l" proof - have def_fold: "\l. list_all (\f. \ \ \ f) l \ \ \ (\ foldl UML_Sequence.OclIncluding Sequence{} l)" apply(rule rev_induct[where P = "\l. list_all (\f. (\ \ \ f)) l \ \ \ (\ foldl UML_Sequence.OclIncluding Sequence{} l)", THEN mp], simp) by(simp add: foundation10') show ?thesis apply(rule rev_induct[where P = "\l. list_all (\f. (\ \ \ f)) l \ \\Rep_Sequence\<^sub>b\<^sub>a\<^sub>s\<^sub>e (foldl UML_Sequence.OclIncluding Sequence{} l \)\\ = List.map (\f. f \) l", THEN mp], simp) apply(simp add: mtSequence_def) apply(subst Abs_Sequence\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse, (simp | intro impI)+) apply(simp add: UML_Sequence.OclIncluding_def, intro conjI impI) apply(subst Abs_Sequence\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse, simp, (rule disjI2)+) apply(simp add: list_all_iff foundation18', simp) apply(subst (asm) def_fold[simplified (no_asm) OclValid_def], simp, simp add: OclValid_def) by (rule assms) qed lemma select_fold_exec\<^sub>S\<^sub>e\<^sub>t: assumes "list_all (\f. (\ \ \ f)) l" shows "\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (foldl UML_Set.OclIncluding Set{} l \)\\ = set (List.map (\f. f \) l)" proof - have def_fold: "\l. list_all (\f. \ \ \ f) l \ \ \ (\ foldl UML_Set.OclIncluding Set{} l)" apply(rule rev_induct[where P = "\l. list_all (\f. (\ \ \ f)) l \ \ \ (\ foldl UML_Set.OclIncluding Set{} l)", THEN mp], simp) by(simp add: foundation10') show ?thesis apply(rule rev_induct[where P = "\l. list_all (\f. (\ \ \ f)) l \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (foldl UML_Set.OclIncluding Set{} l \)\\ = set (List.map (\f. f \) l)", THEN mp], simp) apply(simp add: mtSet_def) apply(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse, (simp | intro impI)+) apply(simp add: UML_Set.OclIncluding_def, intro conjI impI) apply(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse, simp, (rule disjI2)+) apply(simp add: list_all_iff foundation18', simp) apply(subst (asm) def_fold[simplified (no_asm) OclValid_def], simp, simp add: OclValid_def) by (rule assms) qed lemma fold_val_elem\<^sub>S\<^sub>e\<^sub>q: assumes "\ \ \ (foldl UML_Sequence.OclIncluding Sequence{} (List.map (f p) s_set))" shows "list_all (\x. (\ \ \ (f p x))) s_set" apply(rule rev_induct[where P = "\s_set. \ \ \ foldl UML_Sequence.OclIncluding Sequence{} (List.map (f p) s_set) \ list_all (\x. \ \ \ f p x) s_set", THEN mp]) apply(simp, auto) apply (metis (opaque_lifting, mono_tags) UML_Sequence.OclIncluding.def_valid_then_def UML_Sequence.OclIncluding.defined_args_valid foundation20)+ by(simp add: assms) lemma fold_val_elem\<^sub>S\<^sub>e\<^sub>t: assumes "\ \ \ (foldl UML_Set.OclIncluding Set{} (List.map (f p) s_set))" shows "list_all (\x. (\ \ \ (f p x))) s_set" apply(rule rev_induct[where P = "\s_set. \ \ \ foldl UML_Set.OclIncluding Set{} (List.map (f p) s_set) \ list_all (\x. \ \ \ f p x) s_set", THEN mp]) apply(simp, auto) apply (metis (opaque_lifting, mono_tags) foundation10' foundation20)+ by(simp add: assms) lemma select_object_any_defined\<^sub>S\<^sub>e\<^sub>q: assumes def_sel: "\ \ \ (select_object_any\<^sub>S\<^sub>e\<^sub>q f s_set)" shows "s_set \ []" apply(insert def_sel, case_tac s_set) apply(simp add: select_object_any\<^sub>S\<^sub>e\<^sub>q_def UML_Sequence.OclANY_def select_object\<^sub>S\<^sub>e\<^sub>q_def select_object_def defined_def OclValid_def false_def true_def bot_fun_def bot_option_def split: if_split_asm) apply(simp add: mtSequence_def, subst (asm) Abs_Sequence\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse, simp, simp) by(simp) lemma (*select_object_any_defined\<^sub>S\<^sub>e\<^sub>t:*) assumes def_sel: "\ \ \ (select_object_any0\<^sub>S\<^sub>e\<^sub>t f s_set)" shows "s_set \ []" apply(insert def_sel, case_tac s_set) apply(simp add: select_object_any0\<^sub>S\<^sub>e\<^sub>t_def UML_Sequence.OclANY_def select_object\<^sub>S\<^sub>e\<^sub>t_def select_object_def defined_def OclValid_def false_def true_def bot_fun_def bot_option_def split: if_split_asm) by(simp) lemma select_object_any_defined\<^sub>S\<^sub>e\<^sub>t: assumes def_sel: "\ \ \ (select_object_any\<^sub>S\<^sub>e\<^sub>t f s_set)" shows "s_set \ []" apply(insert def_sel, case_tac s_set) apply(simp add: select_object_any\<^sub>S\<^sub>e\<^sub>t_def UML_Sequence.OclANY_def select_object\<^sub>S\<^sub>e\<^sub>t_def select_object_def defined_def OclValid_def false_def true_def bot_fun_def bot_option_def OclInt0_def OclInt1_def StrongEq_def OclIf_def null_fun_def null_option_def split: if_split_asm) by(simp) lemma select_object_any_exec0\<^sub>S\<^sub>e\<^sub>q: assumes def_sel: "\ \ \ (select_object_any\<^sub>S\<^sub>e\<^sub>q f s_set)" shows "\ \ (select_object_any\<^sub>S\<^sub>e\<^sub>q f s_set \ f (hd s_set))" apply(insert def_sel[simplified foundation16], simp add: select_object_any\<^sub>S\<^sub>e\<^sub>q_def foundation22 UML_Sequence.OclANY_def split: if_split_asm) apply(case_tac "\\Rep_Sequence\<^sub>b\<^sub>a\<^sub>s\<^sub>e (select_object\<^sub>S\<^sub>e\<^sub>q f s_set \)\\", simp add: bot_option_def, simp) apply(simp add: select_object\<^sub>S\<^sub>e\<^sub>q_def select_object_def) apply(subst (asm) select_fold_exec\<^sub>S\<^sub>e\<^sub>q) apply(rule fold_val_elem\<^sub>S\<^sub>e\<^sub>q, simp add: foundation18' invalid_def) apply(simp) by(drule arg_cong[where f = hd], subst (asm) hd_map, simp add: select_object_any_defined\<^sub>S\<^sub>e\<^sub>q[OF def_sel], simp) lemma select_object_any_exec\<^sub>S\<^sub>e\<^sub>q: assumes def_sel: "\ \ \ (select_object_any\<^sub>S\<^sub>e\<^sub>q f s_set)" shows "\e. List.member s_set e \ (\ \ (select_object_any\<^sub>S\<^sub>e\<^sub>q f s_set \ f e))" apply(insert select_object_any_exec0\<^sub>S\<^sub>e\<^sub>q[OF def_sel]) apply(rule exI[where x = "hd s_set"], simp) apply(case_tac s_set, simp add: select_object_any_defined\<^sub>S\<^sub>e\<^sub>q[OF def_sel]) by (metis list.sel member_rec(1)) lemma (*select_object_any_exec\<^sub>S\<^sub>e\<^sub>t:*) assumes def_sel: "\ \ \ (select_object_any0\<^sub>S\<^sub>e\<^sub>t f s_set)" shows "\ e. List.member s_set e \ (\ \ (select_object_any0\<^sub>S\<^sub>e\<^sub>t f s_set \ f e))" proof - have list_all_map: "\P f l. list_all P (List.map f l) = list_all (P o f) l" by(induct_tac l, simp_all) fix z show ?thesis when "\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (select_object\<^sub>S\<^sub>e\<^sub>t f s_set \)\\ = z" apply(insert that def_sel[simplified foundation16], simp add: select_object_any0\<^sub>S\<^sub>e\<^sub>t_def foundation22 UML_Set.OclANY_def null_fun_def split: if_split_asm) apply(simp add: select_object\<^sub>S\<^sub>e\<^sub>t_def select_object_def) apply(subst (asm) select_fold_exec\<^sub>S\<^sub>e\<^sub>t) apply(rule fold_val_elem\<^sub>S\<^sub>e\<^sub>t, simp add: OclValid_def) apply(simp add: comp_def) apply(case_tac s_set, simp, simp add: false_def true_def, simp) proof - fix a l show "insert (f a \) ((\x. f x \) ` set l) = z \ \e. List.member (a # l) e \ (SOME y. y \ z) = f e \" apply(rule list.induct[where P = "\l. \z a. insert (f a \) ((\x. f x \) ` set l) = z \ (\e. List.member (a # l) e \ ((SOME y. y \ z) = f e \))", THEN spec, THEN spec, THEN mp], intro allI impI) proof - fix x xa show "insert (f xa \) ((\x. f x \) ` set []) = x \ \e. List.member [xa] e \ (SOME y. y \ x) = f e \" apply(rule exI[where x = xa], simp add: List.member_def) apply(subst some_equality[where a = "f xa \"]) apply (metis (mono_tags) insertI1) apply (metis (mono_tags) empty_iff insert_iff) by(simp) apply_end(intro allI impI, simp) fix x list xa xb show " \x. \e. List.member (x # list) e \ (SOME y. y = f x \ \ y \ (\x. f x \) ` set list) = f e \ \ insert (f xb \) (insert (f x \) ((\x. f x \) ` set list)) = xa \ \e. List.member (xb # x # list) e \ (SOME y. y \ xa) = f e \" apply(case_tac "x = xb", simp) apply(erule allE[where x = xb]) apply(erule exE) proof - fix e show "insert (f xb \) ((\x. f x \) ` set list) = xa \ x = xb \ List.member (xb # list) e \ (SOME y. y = f xb \ \ y \ (\x. f x \) ` set list) = f e \ \ \e. List.member (xb # xb # list) e \ (SOME y. y \ xa) = f e \" apply(rule exI[where x = e], auto) by (metis member_rec(1)) apply_end(case_tac "List.member list x") apply_end(erule allE[where x = xb]) apply_end(erule exE) fix e let ?P = "\y. y = f xb \ \ y \ (\x. f x \) ` set list" show "insert (f xb \) (insert (f x \) ((\x. f x \) ` set list)) = xa \ x \ xb \ List.member list x \ List.member (xb # list) e \ (SOME y. y = f xb \ \ y \ (\x. f x \) ` set list) = f e \ \ \e. List.member (xb # x # list) e \ (SOME y. y \ xa) = f e \" apply(rule exI[where x = e], auto) apply (metis member_rec(1)) apply(subgoal_tac "?P (f e \)") prefer 2 apply(case_tac "xb = e", simp) apply (metis (mono_tags) image_eqI in_set_member member_rec(1)) apply(rule someI2[where a = "f e \"]) apply(erule disjE, simp)+ apply(rule disjI2)+ apply(simp) oops lemma select_object_any_exec\<^sub>S\<^sub>e\<^sub>t: assumes def_sel: "\ \ \ (select_object_any\<^sub>S\<^sub>e\<^sub>t f s_set)" shows "\ e. List.member s_set e \ (\ \ (select_object_any\<^sub>S\<^sub>e\<^sub>t f s_set \ f e))" proof - have card_singl: "\A a. finite A \ card (insert a A) = 1 \ A \ {a}" by (auto, metis Suc_inject card_Suc_eq card_eq_0_iff insert_absorb insert_not_empty singleton_iff) have list_same: "\f s_set z' x. f ` set s_set = {z'} \ List.member s_set x \ f x = z'" by (metis (full_types) empty_iff imageI in_set_member insert_iff) fix z show ?thesis when "\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (select_object\<^sub>S\<^sub>e\<^sub>t f s_set \)\\ = z" apply(insert that def_sel[simplified foundation16], simp add: select_object_any\<^sub>S\<^sub>e\<^sub>t_def foundation22 Let_def null_fun_def bot_fun_def OclIf_def split: if_split_asm) apply(simp add: StrongEq_def OclInt1_def true_def UML_Set.OclSize_def bot_option_def UML_Set.OclANY_def null_fun_def split: if_split_asm) apply(subgoal_tac "\z'. z = {z'}") prefer 2 apply(rule finite.cases[where a = z], simp, simp, simp) apply(rule card_singl, simp, simp) apply(erule exE, clarsimp) apply(simp add: select_object\<^sub>S\<^sub>e\<^sub>t_def select_object_def) apply(subst (asm) select_fold_exec\<^sub>S\<^sub>e\<^sub>t) apply(rule fold_val_elem\<^sub>S\<^sub>e\<^sub>t, simp add: OclValid_def true_def) apply(simp add: comp_def) apply(case_tac s_set, simp) proof - fix z' a list show "(\x. f x \) ` set s_set = {z'} \ s_set = a # list \ \e. List.member s_set e \ z' = f e \" - apply(drule list_same[where x = a]) + apply(drule list_same[where x1 = a]) apply (metis member_rec(1)) by (metis (opaque_lifting, mono_tags) ListMem_iff elem in_set_member) qed qed blast+ end diff --git a/thys/Featherweight_OCL/collection_types/UML_Set.thy b/thys/Featherweight_OCL/collection_types/UML_Set.thy --- a/thys/Featherweight_OCL/collection_types/UML_Set.thy +++ b/thys/Featherweight_OCL/collection_types/UML_Set.thy @@ -1,3172 +1,3172 @@ (***************************************************************************** * Featherweight-OCL --- A Formal Semantics for UML-OCL Version OCL 2.5 * for the OMG Standard. * http://www.brucker.ch/projects/hol-testgen/ * * UML_Set.thy --- Library definitions. * This file is part of HOL-TestGen. * * Copyright (c) 2012-2015 Université Paris-Saclay, Univ. Paris-Sud, France * 2013-2015 IRT SystemX, France * * All rights reserved. * * Redistribution and use in source and binary forms, with or without * modification, are permitted provided that the following conditions are * met: * * * Redistributions of source code must retain the above copyright * notice, this list of conditions and the following disclaimer. * * * Redistributions in binary form must reproduce the above * copyright notice, this list of conditions and the following * disclaimer in the documentation and/or other materials provided * with the distribution. * * * Neither the name of the copyright holders nor the names of its * contributors may be used to endorse or promote products derived * from this software without specific prior written permission. * * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS * "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT * LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR * A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT * OWNER OR CONTRIBUTORS BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, * SPECIAL, EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT * LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, * DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY * THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT * (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE. ******************************************************************************) theory UML_Set imports "../basic_types/UML_Void" "../basic_types/UML_Boolean" "../basic_types/UML_Integer" "../basic_types/UML_String" "../basic_types/UML_Real" begin no_notation None ("\") section\Collection Type Set: Operations \label{formal-set}\ subsection\As a Motivation for the (infinite) Type Construction: Type-Extensions as Sets \label{sec:type-extensions}\ text\Our notion of typed set goes beyond the usual notion of a finite executable set and is powerful enough to capture \emph{the extension of a type} in UML and OCL. This means we can have in Featherweight OCL Sets containing all possible elements of a type, not only those (finite) ones representable in a state. This holds for base types as well as class types, although the notion for class-types --- involving object id's not occurring in a state --- requires some care. In a world with @{term invalid} and @{term null}, there are two notions extensions possible: \begin{enumerate} \item the set of all \emph{defined} values of a type @{term T} (for which we will introduce the constant @{term T}) \item the set of all \emph{valid} values of a type @{term T}, so including @{term null} (for which we will introduce the constant @{term T\<^sub>n\<^sub>u\<^sub>l\<^sub>l}). \end{enumerate} \ text\We define the set extensions for the base type @{term Integer} as follows:\ definition Integer :: "('\,Integer\<^sub>b\<^sub>a\<^sub>s\<^sub>e) Set" where "Integer \ (\ \. (Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e o Some o Some) ((Some o Some) ` (UNIV::int set)))" definition Integer\<^sub>n\<^sub>u\<^sub>l\<^sub>l :: "('\,Integer\<^sub>b\<^sub>a\<^sub>s\<^sub>e) Set" where "Integer\<^sub>n\<^sub>u\<^sub>l\<^sub>l \ (\ \. (Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e o Some o Some) (Some ` (UNIV::int option set)))" lemma Integer_defined : "\ Integer = true" apply(rule ext, auto simp: Integer_def defined_def false_def true_def bot_fun_def null_fun_def null_option_def) by(simp_all add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject bot_option_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_option_def) lemma Integer\<^sub>n\<^sub>u\<^sub>l\<^sub>l_defined : "\ Integer\<^sub>n\<^sub>u\<^sub>l\<^sub>l = true" apply(rule ext, auto simp: Integer\<^sub>n\<^sub>u\<^sub>l\<^sub>l_def defined_def false_def true_def bot_fun_def null_fun_def null_option_def) by(simp_all add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject bot_option_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_option_def) text\This allows the theorems: \\ \ \ x \ \ \ (Integer->includes\<^sub>S\<^sub>e\<^sub>t(x))\ \\ \ \ x \ \ \ Integer \ (Integer->including\<^sub>S\<^sub>e\<^sub>t(x))\ and \\ \ \ x \ \ \ (Integer\<^sub>n\<^sub>u\<^sub>l\<^sub>l->includes\<^sub>S\<^sub>e\<^sub>t(x))\ \\ \ \ x \ \ \ Integer\<^sub>n\<^sub>u\<^sub>l\<^sub>l \ (Integer\<^sub>n\<^sub>u\<^sub>l\<^sub>l->including\<^sub>S\<^sub>e\<^sub>t(x))\ which characterize the infiniteness of these sets by a recursive property on these sets. \ text\In the same spirit, we proceed similarly for the remaining base types:\ definition Void\<^sub>n\<^sub>u\<^sub>l\<^sub>l :: "('\,Void\<^sub>b\<^sub>a\<^sub>s\<^sub>e) Set" where "Void\<^sub>n\<^sub>u\<^sub>l\<^sub>l \ (\ \. (Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e o Some o Some) {Abs_Void\<^sub>b\<^sub>a\<^sub>s\<^sub>e (Some None)})" definition Void\<^sub>e\<^sub>m\<^sub>p\<^sub>t\<^sub>y :: "('\,Void\<^sub>b\<^sub>a\<^sub>s\<^sub>e) Set" where "Void\<^sub>e\<^sub>m\<^sub>p\<^sub>t\<^sub>y \ (\ \. (Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e o Some o Some) {})" lemma Void\<^sub>n\<^sub>u\<^sub>l\<^sub>l_defined : "\ Void\<^sub>n\<^sub>u\<^sub>l\<^sub>l = true" apply(rule ext, auto simp: Void\<^sub>n\<^sub>u\<^sub>l\<^sub>l_def defined_def false_def true_def bot_fun_def null_fun_def null_option_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def) by((subst (asm) Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject, auto simp add: bot_option_def null_option_def bot_Void_def), (subst (asm) Abs_Void\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject, auto simp add: bot_option_def null_option_def))+ lemma Void\<^sub>e\<^sub>m\<^sub>p\<^sub>t\<^sub>y_defined : "\ Void\<^sub>e\<^sub>m\<^sub>p\<^sub>t\<^sub>y = true" apply(rule ext, auto simp: Void\<^sub>e\<^sub>m\<^sub>p\<^sub>t\<^sub>y_def defined_def false_def true_def bot_fun_def null_fun_def null_option_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def) by((subst (asm) Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject, auto simp add: bot_option_def null_option_def bot_Void_def))+ lemma assumes "\ \ \ (V :: ('\,Void\<^sub>b\<^sub>a\<^sub>s\<^sub>e) Set)" shows "\ \ V \ Void\<^sub>n\<^sub>u\<^sub>l\<^sub>l \ \ \ V \ Void\<^sub>e\<^sub>m\<^sub>p\<^sub>t\<^sub>y" proof - have A:"\x y. x \ {} \ \y. y\ x" by (metis all_not_in_conv) show "?thesis" apply(case_tac "V \") proof - fix y show "V \ = Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e y \ y \ {X. X = \ \ X = null \ (\x\\\X\\. x \ \)} \ \ \ V \ Void\<^sub>n\<^sub>u\<^sub>l\<^sub>l \ \ \ V \ Void\<^sub>e\<^sub>m\<^sub>p\<^sub>t\<^sub>y" apply(insert assms, case_tac y, simp add: bot_option_def, simp add: bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def foundation16) apply(simp add: bot_option_def null_option_def) apply(erule disjE, metis OclValid_def defined_def foundation2 null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_fun_def true_def) proof - fix a show "V \ = Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \a\ \ \x\\a\. x \ \ \ \ \ V \ Void\<^sub>n\<^sub>u\<^sub>l\<^sub>l \ \ \ V \ Void\<^sub>e\<^sub>m\<^sub>p\<^sub>t\<^sub>y" apply(case_tac a, simp, insert assms, metis OclValid_def foundation16 null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def true_def) apply(simp) proof - fix aa show " V \ = Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\aa\\ \ \x\aa. x \ \ \ \ \ V \ Void\<^sub>n\<^sub>u\<^sub>l\<^sub>l \ \ \ V \ Void\<^sub>e\<^sub>m\<^sub>p\<^sub>t\<^sub>y" apply(case_tac "aa = {}", rule disjI2, insert assms, simp add: Void\<^sub>e\<^sub>m\<^sub>p\<^sub>t\<^sub>y_def OclValid_def StrongEq_def true_def, rule disjI1) apply(subgoal_tac "aa = {Abs_Void\<^sub>b\<^sub>a\<^sub>s\<^sub>e \None\}", simp add: StrongEq_def OclValid_def true_def Void\<^sub>n\<^sub>u\<^sub>l\<^sub>l_def) apply(drule A, erule exE) proof - fix y show "V \ = Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\aa\\ \ \x\aa. x \ \ \ \ \ \ V \ y \ aa \ aa = {Abs_Void\<^sub>b\<^sub>a\<^sub>s\<^sub>e \None\}" apply(rule equalityI, rule subsetI, simp) proof - fix x show " V \ = Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\aa\\ \ \x\aa. x \ \ \ \ \ \ V \ y \ aa \ x \ aa \ x = Abs_Void\<^sub>b\<^sub>a\<^sub>s\<^sub>e \None\" apply(case_tac x, simp) by (metis bot_Void_def bot_option_def null_option_def) apply_end(simp_all) apply_end(erule ballE[where x = y], simp_all) apply_end(case_tac y, simp add: bot_option_def null_option_def OclValid_def defined_def split: if_split_asm, simp add: false_def true_def) qed (erule disjE, simp add: bot_Void_def, simp) qed qed qed qed qed definition Boolean :: "('\,Boolean\<^sub>b\<^sub>a\<^sub>s\<^sub>e) Set" where "Boolean \ (\ \. (Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e o Some o Some) ((Some o Some) ` (UNIV::bool set)))" definition Boolean\<^sub>n\<^sub>u\<^sub>l\<^sub>l :: "('\,Boolean\<^sub>b\<^sub>a\<^sub>s\<^sub>e) Set" where "Boolean\<^sub>n\<^sub>u\<^sub>l\<^sub>l \ (\ \. (Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e o Some o Some) (Some ` (UNIV::bool option set)))" lemma Boolean_defined : "\ Boolean = true" apply(rule ext, auto simp: Boolean_def defined_def false_def true_def bot_fun_def null_fun_def null_option_def) by(simp_all add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject bot_option_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_option_def) lemma Boolean\<^sub>n\<^sub>u\<^sub>l\<^sub>l_defined : "\ Boolean\<^sub>n\<^sub>u\<^sub>l\<^sub>l = true" apply(rule ext, auto simp: Boolean\<^sub>n\<^sub>u\<^sub>l\<^sub>l_def defined_def false_def true_def bot_fun_def null_fun_def null_option_def) by(simp_all add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject bot_option_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_option_def) definition String :: "('\,String\<^sub>b\<^sub>a\<^sub>s\<^sub>e) Set" where "String \ (\ \. (Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e o Some o Some) ((Some o Some) ` (UNIV::string set)))" definition String\<^sub>n\<^sub>u\<^sub>l\<^sub>l :: "('\,String\<^sub>b\<^sub>a\<^sub>s\<^sub>e) Set" where "String\<^sub>n\<^sub>u\<^sub>l\<^sub>l \ (\ \. (Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e o Some o Some) (Some ` (UNIV::string option set)))" lemma String_defined : "\ String = true" apply(rule ext, auto simp: String_def defined_def false_def true_def bot_fun_def null_fun_def null_option_def) by(simp_all add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject bot_option_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_option_def) lemma String\<^sub>n\<^sub>u\<^sub>l\<^sub>l_defined : "\ String\<^sub>n\<^sub>u\<^sub>l\<^sub>l = true" apply(rule ext, auto simp: String\<^sub>n\<^sub>u\<^sub>l\<^sub>l_def defined_def false_def true_def bot_fun_def null_fun_def null_option_def) by(simp_all add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject bot_option_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_option_def) definition Real :: "('\,Real\<^sub>b\<^sub>a\<^sub>s\<^sub>e) Set" where "Real \ (\ \. (Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e o Some o Some) ((Some o Some) ` (UNIV::real set)))" definition Real\<^sub>n\<^sub>u\<^sub>l\<^sub>l :: "('\,Real\<^sub>b\<^sub>a\<^sub>s\<^sub>e) Set" where "Real\<^sub>n\<^sub>u\<^sub>l\<^sub>l \ (\ \. (Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e o Some o Some) (Some ` (UNIV::real option set)))" lemma Real_defined : "\ Real = true" apply(rule ext, auto simp: Real_def defined_def false_def true_def bot_fun_def null_fun_def null_option_def) by(simp_all add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject bot_option_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_option_def) lemma Real\<^sub>n\<^sub>u\<^sub>l\<^sub>l_defined : "\ Real\<^sub>n\<^sub>u\<^sub>l\<^sub>l = true" apply(rule ext, auto simp: Real\<^sub>n\<^sub>u\<^sub>l\<^sub>l_def defined_def false_def true_def bot_fun_def null_fun_def null_option_def) by(simp_all add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject bot_option_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_option_def) subsection\Basic Properties of the Set Type\ text\Every element in a defined set is valid.\ lemma Set_inv_lemma: "\ \ (\ X) \ \x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\. x \ bot" apply(insert Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e [of "X \"], simp) apply(auto simp: OclValid_def defined_def false_def true_def cp_def bot_fun_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_fun_def split:if_split_asm) apply(erule contrapos_pp [of "Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \) = bot"]) apply(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject[symmetric], rule Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e, simp) apply(simp add: Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_option_def) apply(erule contrapos_pp [of "Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \) = null"]) apply(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject[symmetric], rule Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e, simp) apply(simp add: Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse null_option_def) by (simp add: bot_option_def) lemma Set_inv_lemma' : assumes x_def : "\ \ \ X" and e_mem : "e \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\" shows "\ \ \ (\_. e)" apply(rule Set_inv_lemma[OF x_def, THEN ballE[where x = e]]) apply(simp add: foundation18') by(simp add: e_mem) lemma abs_rep_simp' : assumes S_all_def : "\ \ \ S" shows "Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\\\ = S \" proof - have discr_eq_false_true : "\\. (false \ = true \) = False" by(simp add: false_def true_def) show ?thesis apply(insert S_all_def, simp add: OclValid_def defined_def) apply(rule mp[OF Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_induct[where P = "\S. (if S = \ \ \ S = null \ then false \ else true \) = true \ \ Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e S\\\\ = S"]], rename_tac S') apply(simp add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse discr_eq_false_true) apply(case_tac S') apply(simp add: bot_fun_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def)+ apply(rename_tac S'', case_tac S'') apply(simp add: null_fun_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def)+ done qed lemma S_lift' : assumes S_all_def : "(\ :: '\ st) \ \ S" shows "\S'. (\a (_::'\ st). a) ` \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\ = (\a (_::'\ st). \a\) ` S'" apply(rule_tac x = "(\a. \a\) ` \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\" in exI) apply(simp only: image_comp) apply(simp add: comp_def) apply(rule image_cong, fast) (* *) apply(drule Set_inv_lemma'[OF S_all_def]) by(case_tac x, (simp add: bot_option_def foundation18')+) lemma invalid_set_OclNot_defined [simp,code_unfold]:"\(invalid::('\,'\::null) Set) = false" by simp lemma null_set_OclNot_defined [simp,code_unfold]:"\(null::('\,'\::null) Set) = false" by(simp add: defined_def null_fun_def) lemma invalid_set_valid [simp,code_unfold]:"\(invalid::('\,'\::null) Set) = false" by simp lemma null_set_valid [simp,code_unfold]:"\(null::('\,'\::null) Set) = true" apply(simp add: valid_def null_fun_def bot_fun_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def) apply(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject,simp_all add: null_option_def bot_option_def) done text\... which means that we can have a type \('\,('\,('\) Integer) Set) Set\ corresponding exactly to Set(Set(Integer)) in OCL notation. Note that the parameter \'\\ still refers to the object universe; making the OCL semantics entirely parametric in the object universe makes it possible to study (and prove) its properties independently from a concrete class diagram.\ subsection\Definition: Strict Equality \label{sec:set-strict-equality}\ text\After the part of foundational operations on sets, we detail here equality on sets. Strong equality is inherited from the OCL core, but we have to consider the case of the strict equality. We decide to overload strict equality in the same way we do for other value's in OCL:\ overloading StrictRefEq \ "StrictRefEq :: [('\,'\::null)Set,('\,'\::null)Set] \ ('\)Boolean" begin definition StrictRefEq\<^sub>S\<^sub>e\<^sub>t : "(x::('\,'\::null)Set) \ y \ \ \. if (\ x) \ = true \ \ (\ y) \ = true \ then (x \ y)\ else invalid \" end text\One might object here that for the case of objects, this is an empty definition. The answer is no, we will restrain later on states and objects such that any object has its oid stored inside the object (so the ref, under which an object can be referenced in the store will represented in the object itself). For such well-formed stores that satisfy this invariant (the WFF-invariant), the referential equality and the strong equality---and therefore the strict equality on sets in the sense above---coincides.\ text\Property proof in terms of @{term "profile_bin\<^sub>S\<^sub>t\<^sub>r\<^sub>o\<^sub>n\<^sub>g\<^sub>E\<^sub>q_\<^sub>v_\<^sub>v"}\ interpretation StrictRefEq\<^sub>S\<^sub>e\<^sub>t : profile_bin\<^sub>S\<^sub>t\<^sub>r\<^sub>o\<^sub>n\<^sub>g\<^sub>E\<^sub>q_\<^sub>v_\<^sub>v "\ x y. (x::('\,'\::null)Set) \ y" by unfold_locales (auto simp: StrictRefEq\<^sub>S\<^sub>e\<^sub>t) subsection\Constants: mtSet\ definition mtSet::"('\,'\::null) Set" ("Set{}") where "Set{} \ (\ \. Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\{}::'\ set\\ )" lemma mtSet_defined[simp,code_unfold]:"\(Set{}) = true" apply(rule ext, auto simp: mtSet_def defined_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_fun_def null_fun_def) by(simp_all add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject bot_option_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_option_def) lemma mtSet_valid[simp,code_unfold]:"\(Set{}) = true" apply(rule ext,auto simp: mtSet_def valid_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_fun_def null_fun_def) by(simp_all add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject bot_option_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_option_def) lemma mtSet_rep_set: "\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (Set{} \)\\ = {}" apply(simp add: mtSet_def, subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse) by(simp add: bot_option_def)+ lemma [simp,code_unfold]: "const Set{}" by(simp add: const_def mtSet_def) text\Note that the collection types in OCL allow for null to be included; however, there is the null-collection into which inclusion yields invalid.\ subsection\Definition: Including\ definition OclIncluding :: "[('\,'\::null) Set,('\,'\) val] \ ('\,'\) Set" where "OclIncluding x y = (\ \. if (\ x) \ = true \ \ (\ y) \ = true \ then Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (x \)\\ \ {y \} \\ else invalid \ )" notation OclIncluding ("_->including\<^sub>S\<^sub>e\<^sub>t'(_')") interpretation OclIncluding : profile_bin\<^sub>d_\<^sub>v OclIncluding "\x y. Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e\\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e x\\ \ {y}\\" proof - have A : "None \ {X. X = bot \ X = null \ (\x\\\X\\. x \ bot)}" by(simp add: bot_option_def) have B : "\None\ \ {X. X = bot \ X = null \ (\x\\\X\\. x \ bot)}" by(simp add: null_option_def bot_option_def) have C : "\x y. x \ \ \ x \ null \ y \ \ \ \\insert y \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e x\\\\ \ {X. X = bot \ X = null \ (\x\\\X\\. x \ bot)}" by(auto intro!:Set_inv_lemma[simplified OclValid_def defined_def false_def true_def null_fun_def bot_fun_def]) show "profile_bin\<^sub>d_\<^sub>v OclIncluding (\x y. Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e\\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e x\\ \ {y}\\)" apply unfold_locales apply(auto simp:OclIncluding_def bot_option_def null_option_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def) apply(erule_tac Q="Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e\\insert y \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e x\\\\ = Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e None" in contrapos_pp) apply(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject[OF C A]) apply(simp_all add: null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_option_def) apply(erule_tac Q="Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e\\insert y \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e x\\\\ = Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \None\" in contrapos_pp) apply(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject[OF C B]) apply(simp_all add: null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_option_def) done qed syntax "_OclFinset" :: "args => ('\,'a::null) Set" ("Set{(_)}") translations "Set{x, xs}" == "CONST OclIncluding (Set{xs}) x" "Set{x}" == "CONST OclIncluding (Set{}) x " subsection\Definition: Excluding\ definition OclExcluding :: "[('\,'\::null) Set,('\,'\) val] \ ('\,'\) Set" where "OclExcluding x y = (\ \. if (\ x) \ = true \ \ (\ y) \ = true \ then Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (x \)\\ - {y \} \\ else \ )" notation OclExcluding ("_->excluding\<^sub>S\<^sub>e\<^sub>t'(_')") lemma OclExcluding_inv: "(x:: Set('b::{null})) \ \ \ x \ null \ y \ \ \ \\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e x\\ - {y}\\ \ {X. X = bot \ X = null \ (\x\\\X\\. x \ bot)}" proof - fix X :: "'a state \ 'a state \ Set('b)" fix \ show "x \ \ \ x \ null \ y \ \ \ ?thesis" when "x = X \" by(simp add: that Set_inv_lemma[simplified OclValid_def defined_def null_fun_def bot_fun_def, of X \]) qed simp_all interpretation OclExcluding : profile_bin\<^sub>d_\<^sub>v OclExcluding "\x y. Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e\\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e x\\ - {y}\\" proof - have A : "None \ {X. X = bot \ X = null \ (\x\\\X\\. x \ bot)}" by(simp add: bot_option_def) have B : "\None\ \ {X. X = bot \ X = null \ (\x\\\X\\. x \ bot)}" by(simp add: null_option_def bot_option_def) show "profile_bin\<^sub>d_\<^sub>v OclExcluding (\x y. Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e\\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (x:: Set('b))\\ - {y}\\)" apply unfold_locales apply(auto simp:OclExcluding_def bot_option_def null_option_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def invalid_def) apply(erule_tac Q="Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e\\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e x\\ - {y}\\ = Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e None" in contrapos_pp) apply(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject[OF OclExcluding_inv A]) apply(simp_all add: null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_option_def) apply(erule_tac Q="Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e\\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e x\\ - {y}\\ = Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \None\" in contrapos_pp) apply(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject[OF OclExcluding_inv B]) apply(simp_all add: null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_option_def) done qed subsection\Definition: Includes\ definition OclIncludes :: "[('\,'\::null) Set,('\,'\) val] \ '\ Boolean" where "OclIncludes x y = (\ \. if (\ x) \ = true \ \ (\ y) \ = true \ then \\(y \) \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (x \)\\ \\ else \ )" notation OclIncludes ("_->includes\<^sub>S\<^sub>e\<^sub>t'(_')" (*[66,65]65*)) interpretation OclIncludes : profile_bin\<^sub>d_\<^sub>v OclIncludes "\x y. \\y \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e x\\\\" by(unfold_locales, auto simp:OclIncludes_def bot_option_def null_option_def invalid_def) subsection\Definition: Excludes\ definition OclExcludes :: "[('\,'\::null) Set,('\,'\) val] \ '\ Boolean" where "OclExcludes x y = (not(OclIncludes x y))" notation OclExcludes ("_->excludes\<^sub>S\<^sub>e\<^sub>t'(_')" (*[66,65]65*)) text\The case of the size definition is somewhat special, we admit explicitly in Featherweight OCL the possibility of infinite sets. For the size definition, this requires an extra condition that assures that the cardinality of the set is actually a defined integer.\ interpretation OclExcludes : profile_bin\<^sub>d_\<^sub>v OclExcludes "\x y. \\y \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e x\\\\" by(unfold_locales, auto simp:OclExcludes_def OclIncludes_def OclNot_def bot_option_def null_option_def invalid_def) subsection\Definition: Size\ definition OclSize :: "('\,'\::null)Set \ '\ Integer" where "OclSize x = (\ \. if (\ x) \ = true \ \ finite(\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (x \)\\) then \\ int(card \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (x \)\\) \\ else \ )" notation (* standard ascii syntax *) OclSize ("_->size\<^sub>S\<^sub>e\<^sub>t'(')" (*[66]*)) text\The following definition follows the requirement of the standard to treat null as neutral element of sets. It is a well-documented exception from the general strictness rule and the rule that the distinguished argument self should be non-null.\ (*TODO Locale - Equivalent*) subsection\Definition: IsEmpty\ definition OclIsEmpty :: "('\,'\::null) Set \ '\ Boolean" where "OclIsEmpty x = ((\ x and not (\ x)) or ((OclSize x) \ \))" notation OclIsEmpty ("_->isEmpty\<^sub>S\<^sub>e\<^sub>t'(')" (*[66]*)) (*TODO Locale - Equivalent*) subsection\Definition: NotEmpty\ definition OclNotEmpty :: "('\,'\::null) Set \ '\ Boolean" where "OclNotEmpty x = not(OclIsEmpty x)" notation OclNotEmpty ("_->notEmpty\<^sub>S\<^sub>e\<^sub>t'(')" (*[66]*)) (*TODO Locale - Equivalent*) subsection\Definition: Any\ (* Slight breach of naming convention in order to avoid naming conflict on constant.*) definition OclANY :: "[('\,'\::null) Set] \ ('\,'\) val" where "OclANY x = (\ \. if (\ x) \ = true \ then if (\ x and OclNotEmpty x) \ = true \ then SOME y. y \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (x \)\\ else null \ else \ )" notation OclANY ("_->any\<^sub>S\<^sub>e\<^sub>t'(')") (*TODO Locale - Equivalent*) (* actually, this definition covers only: X->any\<^sub>S\<^sub>e\<^sub>t(true) of the standard, which foresees a (totally correct) high-level definition source->any\<^sub>S\<^sub>e\<^sub>t(iterator | body) = source->select(iterator | body)->asSequence()->first(). Since we don't have sequences, we have to go for a direct---restricted---definition. *) subsection\Definition: Forall\ text\The definition of OclForall mimics the one of @{term "OclAnd"}: OclForall is not a strict operation.\ definition OclForall :: "[('\,'\::null)Set,('\,'\)val\('\)Boolean] \ '\ Boolean" where "OclForall S P = (\ \. if (\ S) \ = true \ then if (\x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\. P(\ _. x) \ = false \) then false \ else if (\x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\. P(\ _. x) \ = invalid \) then invalid \ else if (\x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\. P(\ _. x) \ = null \) then null \ else true \ else \)" syntax "_OclForallSet" :: "[('\,'\::null) Set,id,('\)Boolean] \ '\ Boolean" ("(_)->forAll\<^sub>S\<^sub>e\<^sub>t'(_|_')") translations "X->forAll\<^sub>S\<^sub>e\<^sub>t(x | P)" == "CONST UML_Set.OclForall X (%x. P)" (*TODO Locale - Equivalent*) subsection\Definition: Exists\ text\Like OclForall, OclExists is also not strict.\ definition OclExists :: "[('\,'\::null) Set,('\,'\)val\('\)Boolean] \ '\ Boolean" where "OclExists S P = not(UML_Set.OclForall S (\ X. not (P X)))" syntax "_OclExistSet" :: "[('\,'\::null) Set,id,('\)Boolean] \ '\ Boolean" ("(_)->exists\<^sub>S\<^sub>e\<^sub>t'(_|_')") translations "X->exists\<^sub>S\<^sub>e\<^sub>t(x | P)" == "CONST UML_Set.OclExists X (%x. P)" (*TODO Locale - Equivalent*) subsection\Definition: Iterate\ definition OclIterate :: "[('\,'\::null) Set,('\,'\::null)val, ('\,'\)val\('\,'\)val\('\,'\)val] \ ('\,'\)val" where "OclIterate S A F = (\ \. if (\ S) \ = true \ \ (\ A) \ = true \ \ finite\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\ then (Finite_Set.fold (F) (A) ((\a \. a) ` \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\))\ else \)" syntax "_OclIterateSet" :: "[('\,'\::null) Set, idt, idt, '\, '\] => ('\,'\)val" ("_ ->iterate\<^sub>S\<^sub>e\<^sub>t'(_;_=_ | _')" (*[71,100,70]50*)) translations "X->iterate\<^sub>S\<^sub>e\<^sub>t(a; x = A | P)" == "CONST OclIterate X A (%a. (% x. P))" (*TODO Locale - Equivalent*) subsection\Definition: Select\ definition OclSelect :: "[('\,'\::null)Set,('\,'\)val\('\)Boolean] \ ('\,'\)Set" where "OclSelect S P = (\\. if (\ S) \ = true \ then if (\x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\. P(\ _. x) \ = invalid \) then invalid \ else Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\{x\\\ Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\. P (\_. x) \ \ false \}\\ else invalid \)" syntax "_OclSelectSet" :: "[('\,'\::null) Set,id,('\)Boolean] \ '\ Boolean" ("(_)->select\<^sub>S\<^sub>e\<^sub>t'(_|_')") translations "X->select\<^sub>S\<^sub>e\<^sub>t(x | P)" == "CONST OclSelect X (% x. P)" (*TODO Locale - Equivalent*) subsection\Definition: Reject\ definition OclReject :: "[('\,'\::null)Set,('\,'\)val\('\)Boolean] \ ('\,'\::null)Set" where "OclReject S P = OclSelect S (not o P)" syntax "_OclRejectSet" :: "[('\,'\::null) Set,id,('\)Boolean] \ '\ Boolean" ("(_)->reject\<^sub>S\<^sub>e\<^sub>t'(_|_')") translations "X->reject\<^sub>S\<^sub>e\<^sub>t(x | P)" == "CONST OclReject X (% x. P)" (*TODO Locale - Equivalent*) subsection\Definition: IncludesAll\ definition OclIncludesAll :: "[('\,'\::null) Set,('\,'\) Set] \ '\ Boolean" where "OclIncludesAll x y = (\ \. if (\ x) \ = true \ \ (\ y) \ = true \ then \\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (y \)\\ \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (x \)\\ \\ else \ )" notation OclIncludesAll ("_->includesAll\<^sub>S\<^sub>e\<^sub>t'(_')" (*[66,65]65*)) interpretation OclIncludesAll : profile_bin\<^sub>d_\<^sub>d OclIncludesAll "\x y. \\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e y\\ \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e x\\\\" by(unfold_locales, auto simp:OclIncludesAll_def bot_option_def null_option_def invalid_def) subsection\Definition: ExcludesAll\ definition OclExcludesAll :: "[('\,'\::null) Set,('\,'\) Set] \ '\ Boolean" where "OclExcludesAll x y = (\ \. if (\ x) \ = true \ \ (\ y) \ = true \ then \\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (y \)\\ \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (x \)\\ = {} \\ else \ )" notation OclExcludesAll ("_->excludesAll\<^sub>S\<^sub>e\<^sub>t'(_')" (*[66,65]65*)) interpretation OclExcludesAll : profile_bin\<^sub>d_\<^sub>d OclExcludesAll "\x y. \\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e y\\ \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e x\\ = {}\\" by(unfold_locales, auto simp:OclExcludesAll_def bot_option_def null_option_def invalid_def) subsection\Definition: Union\ definition OclUnion :: "[('\,'\::null) Set,('\,'\) Set] \ ('\,'\) Set" where "OclUnion x y = (\ \. if (\ x) \ = true \ \ (\ y) \ = true \ then Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e\\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (y \)\\ \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (x \)\\ \\ else \ )" notation OclUnion ("_->union\<^sub>S\<^sub>e\<^sub>t'(_')" (*[66,65]65*)) lemma OclUnion_inv: "(x:: Set('b::{null})) \ \ \ x \ null \ y \ \ \ y \ null \ \\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e y\\ \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e x\\\\ \ {X. X = bot \ X = null \ (\x\\\X\\. x \ bot)}" proof - fix X Y :: "'a state \ 'a state \ Set('b)" fix \ show "x \ \ \ x \ null \ y \ \ \ y \ null \ ?thesis" when "x = X \" "y = Y \" by(auto simp: that, insert Set_inv_lemma[simplified OclValid_def defined_def null_fun_def bot_fun_def, of Y \] Set_inv_lemma[simplified OclValid_def defined_def null_fun_def bot_fun_def, of X \], auto) qed simp_all interpretation OclUnion : profile_bin\<^sub>d_\<^sub>d OclUnion "\x y. Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e\\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e y\\ \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e x\\\\" proof - have A : "None \ {X. X = bot \ X = null \ (\x\\\X\\. x \ bot)}" by(simp add: bot_option_def) have B : "\None\ \ {X. X = bot \ X = null \ (\x\\\X\\. x \ bot)}" by(simp add: null_option_def bot_option_def) show "profile_bin\<^sub>d_\<^sub>d OclUnion (\x y. Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e\\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e y\\ \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e x\\\\)" apply unfold_locales apply(auto simp:OclUnion_def bot_option_def null_option_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def invalid_def) apply(erule_tac Q="Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e\\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e y\\ \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e x\\\\ = Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e None" in contrapos_pp) apply(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject[OF OclUnion_inv A]) apply(simp_all add: null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_option_def) apply(erule_tac Q="Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e\\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e y\\ \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e x\\\\ = Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \None\" in contrapos_pp) apply(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject[OF OclUnion_inv B]) apply(simp_all add: null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_option_def) done qed subsection\Definition: Intersection\ definition OclIntersection :: "[('\,'\::null) Set,('\,'\) Set] \ ('\,'\) Set" where "OclIntersection x y = (\ \. if (\ x) \ = true \ \ (\ y) \ = true \ then Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e\\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (y \)\\ \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (x \)\\\\ else \ )" notation OclIntersection("_->intersection\<^sub>S\<^sub>e\<^sub>t'(_')" (*[71,70]70*)) lemma OclIntersection_inv: "(x:: Set('b::{null})) \ \ \ x \ null \ y \ \ \ y \ null \ \\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e y\\ \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e x\\\\ \ {X. X = bot \ X = null \ (\x\\\X\\. x \ bot)}" proof - fix X Y :: "'a state \ 'a state \ Set('b)" fix \ show "x \ \ \ x \ null \ y \ \ \ y \ null \ ?thesis" when "x = X \" "y = Y \" by(auto simp: that, insert Set_inv_lemma[simplified OclValid_def defined_def null_fun_def bot_fun_def, of Y \] Set_inv_lemma[simplified OclValid_def defined_def null_fun_def bot_fun_def, of X \], auto) qed simp_all interpretation OclIntersection : profile_bin\<^sub>d_\<^sub>d OclIntersection "\x y. Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e\\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e y\\ \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e x\\\\" proof - have A : "None \ {X. X = bot \ X = null \ (\x\\\X\\. x \ bot)}" by(simp add: bot_option_def) have B : "\None\ \ {X. X = bot \ X = null \ (\x\\\X\\. x \ bot)}" by(simp add: null_option_def bot_option_def) show "profile_bin\<^sub>d_\<^sub>d OclIntersection (\x y. Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e\\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e y\\ \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e x\\\\)" apply unfold_locales apply(auto simp:OclIntersection_def bot_option_def null_option_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def invalid_def) apply(erule_tac Q="Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e\\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e y\\ \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e x\\\\ = Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e None" in contrapos_pp) apply(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject[OF OclIntersection_inv A]) apply(simp_all add: null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_option_def) apply(erule_tac Q="Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e\\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e y\\ \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e x\\\\ = Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \None\" in contrapos_pp) apply(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject[OF OclIntersection_inv B]) apply(simp_all add: null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_option_def) done qed subsection\Definition (future operators)\ consts (* abstract set collection operations *) OclCount :: "[('\,'\::null) Set,('\,'\) Set] \ '\ Integer" OclSum :: " ('\,'\::null) Set \ '\ Integer" notation OclCount ("_->count\<^sub>S\<^sub>e\<^sub>t'(_')" (*[66,65]65*)) notation OclSum ("_->sum\<^sub>S\<^sub>e\<^sub>t'(')" (*[66]*)) subsection\Logical Properties\ text\OclIncluding\ lemma OclIncluding_valid_args_valid: "(\ \ \(X->including\<^sub>S\<^sub>e\<^sub>t(x))) = ((\ \(\ X)) \ (\ \(\ x)))" by (metis (opaque_lifting, no_types) OclIncluding.def_valid_then_def OclIncluding.defined_args_valid) lemma OclIncluding_valid_args_valid''[simp,code_unfold]: "\(X->including\<^sub>S\<^sub>e\<^sub>t(x)) = ((\ X) and (\ x))" by (simp add: OclIncluding.def_valid_then_def) text\etc. etc.\ text_raw\\isatagafp\ text\OclExcluding\ lemma OclExcluding_valid_args_valid: "(\ \ \(X->excluding\<^sub>S\<^sub>e\<^sub>t(x))) = ((\ \(\ X)) \ (\ \(\ x)))" by (metis OclExcluding.def_valid_then_def OclExcluding.defined_args_valid) lemma OclExcluding_valid_args_valid''[simp,code_unfold]: "\(X->excluding\<^sub>S\<^sub>e\<^sub>t(x)) = ((\ X) and (\ x))" by (simp add: OclExcluding.def_valid_then_def) text\OclIncludes\ lemma OclIncludes_valid_args_valid: "(\ \ \(X->includes\<^sub>S\<^sub>e\<^sub>t(x))) = ((\ \(\ X)) \ (\ \(\ x)))" by (simp add: OclIncludes.def_valid_then_def foundation10') lemma OclIncludes_valid_args_valid''[simp,code_unfold]: "\(X->includes\<^sub>S\<^sub>e\<^sub>t(x)) = ((\ X) and (\ x))" by (simp add: OclIncludes.def_valid_then_def) text\OclExcludes\ lemma OclExcludes_valid_args_valid: "(\ \ \(X->excludes\<^sub>S\<^sub>e\<^sub>t(x))) = ((\ \(\ X)) \ (\ \(\ x)))" by (simp add: OclExcludes.def_valid_then_def foundation10') lemma OclExcludes_valid_args_valid''[simp,code_unfold]: "\(X->excludes\<^sub>S\<^sub>e\<^sub>t(x)) = ((\ X) and (\ x))" by (simp add: OclExcludes.def_valid_then_def) text\OclSize\ lemma OclSize_defined_args_valid: "\ \ \ (X->size\<^sub>S\<^sub>e\<^sub>t()) \ \ \ \ X" by(auto simp: OclSize_def OclValid_def true_def valid_def false_def StrongEq_def defined_def invalid_def bot_fun_def null_fun_def split: bool.split_asm HOL.if_split_asm option.split) lemma OclSize_infinite: assumes non_finite:"\ \ not(\(S->size\<^sub>S\<^sub>e\<^sub>t()))" shows "(\ \ not(\(S))) \ \ finite \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\" apply(insert non_finite, simp) apply(rule impI) apply(simp add: OclSize_def OclValid_def defined_def) apply(case_tac "finite \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\", simp_all add:null_fun_def null_option_def bot_fun_def bot_option_def) done lemma "\ \ \ X \ \ finite \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\ \ \ \ \ \ (X->size\<^sub>S\<^sub>e\<^sub>t())" by(simp add: OclSize_def OclValid_def defined_def bot_fun_def false_def true_def) lemma size_defined: assumes X_finite: "\\. finite \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\" shows "\ (X->size\<^sub>S\<^sub>e\<^sub>t()) = \ X" apply(rule ext, simp add: cp_defined[of "X->size\<^sub>S\<^sub>e\<^sub>t()"] OclSize_def) apply(simp add: defined_def bot_option_def bot_fun_def null_option_def null_fun_def X_finite) done lemma size_defined': assumes X_finite: "finite \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\" shows "(\ \ \ (X->size\<^sub>S\<^sub>e\<^sub>t())) = (\ \ \ X)" apply(simp add: cp_defined[of "X->size\<^sub>S\<^sub>e\<^sub>t()"] OclSize_def OclValid_def) apply(simp add: defined_def bot_option_def bot_fun_def null_option_def null_fun_def X_finite) done text\OclIsEmpty\ lemma OclIsEmpty_defined_args_valid:"\ \ \ (X->isEmpty\<^sub>S\<^sub>e\<^sub>t()) \ \ \ \ X" apply(auto simp: OclIsEmpty_def OclValid_def defined_def valid_def false_def true_def bot_fun_def null_fun_def OclAnd_def OclOr_def OclNot_def split: if_split_asm) apply(case_tac "(X->size\<^sub>S\<^sub>e\<^sub>t() \ \) \", simp add: bot_option_def, simp, rename_tac x) apply(case_tac x, simp add: null_option_def bot_option_def, simp) apply(simp add: OclSize_def StrictRefEq\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r valid_def) by (metis (opaque_lifting, no_types) bot_fun_def OclValid_def defined_def foundation2 invalid_def) lemma "\ \ \ (null->isEmpty\<^sub>S\<^sub>e\<^sub>t())" by(auto simp: OclIsEmpty_def OclValid_def defined_def valid_def false_def true_def bot_fun_def null_fun_def OclAnd_def OclOr_def OclNot_def null_is_valid split: if_split_asm) lemma OclIsEmpty_infinite: "\ \ \ X \ \ finite \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\ \ \ \ \ \ (X->isEmpty\<^sub>S\<^sub>e\<^sub>t())" apply(auto simp: OclIsEmpty_def OclValid_def defined_def valid_def false_def true_def bot_fun_def null_fun_def OclAnd_def OclOr_def OclNot_def split: if_split_asm) apply(case_tac "(X->size\<^sub>S\<^sub>e\<^sub>t() \ \) \", simp add: bot_option_def, simp, rename_tac x) apply(case_tac x, simp add: null_option_def bot_option_def, simp) by(simp add: OclSize_def StrictRefEq\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r valid_def bot_fun_def false_def true_def invalid_def) text\OclNotEmpty\ lemma OclNotEmpty_defined_args_valid:"\ \ \ (X->notEmpty\<^sub>S\<^sub>e\<^sub>t()) \ \ \ \ X" by (metis (opaque_lifting, no_types) OclNotEmpty_def OclNot_defargs OclNot_not foundation6 foundation9 OclIsEmpty_defined_args_valid) lemma "\ \ \ (null->notEmpty\<^sub>S\<^sub>e\<^sub>t())" by (metis (opaque_lifting, no_types) OclNotEmpty_def OclAnd_false1 OclAnd_idem OclIsEmpty_def OclNot3 OclNot4 OclOr_def defined2 defined4 transform1 valid2) lemma OclNotEmpty_infinite: "\ \ \ X \ \ finite \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\ \ \ \ \ \ (X->notEmpty\<^sub>S\<^sub>e\<^sub>t())" apply(simp add: OclNotEmpty_def) apply(drule OclIsEmpty_infinite, simp) by (metis OclNot_defargs OclNot_not foundation6 foundation9) lemma OclNotEmpty_has_elt : "\ \ \ X \ \ \ X->notEmpty\<^sub>S\<^sub>e\<^sub>t() \ \e. e \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\" apply(simp add: OclNotEmpty_def OclIsEmpty_def deMorgan1 deMorgan2, drule foundation5) apply(subst (asm) (2) OclNot_def, simp add: OclValid_def StrictRefEq\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r StrongEq_def split: if_split_asm) prefer 2 apply(simp add: invalid_def bot_option_def true_def) apply(simp add: OclSize_def valid_def split: if_split_asm, simp_all add: false_def true_def bot_option_def bot_fun_def OclInt0_def) by (metis equals0I) text\OclANY\ lemma OclANY_defined_args_valid: "\ \ \ (X->any\<^sub>S\<^sub>e\<^sub>t()) \ \ \ \ X" by(auto simp: OclANY_def OclValid_def true_def valid_def false_def StrongEq_def defined_def invalid_def bot_fun_def null_fun_def OclAnd_def split: bool.split_asm HOL.if_split_asm option.split) lemma "\ \ \ X \ \ \ X->isEmpty\<^sub>S\<^sub>e\<^sub>t() \ \ \ \ \ (X->any\<^sub>S\<^sub>e\<^sub>t())" apply(simp add: OclANY_def OclValid_def) apply(subst cp_defined, subst cp_OclAnd, simp add: OclNotEmpty_def, subst (1 2) cp_OclNot, simp add: cp_OclNot[symmetric] cp_OclAnd[symmetric] cp_defined[symmetric], simp add: false_def true_def) by(drule foundation20[simplified OclValid_def true_def], simp) lemma OclANY_valid_args_valid: "(\ \ \(X->any\<^sub>S\<^sub>e\<^sub>t())) = (\ \ \ X)" proof - have A: "(\ \ \(X->any\<^sub>S\<^sub>e\<^sub>t())) \ ((\ \(\ X)))" by(auto simp: OclANY_def OclValid_def true_def valid_def false_def StrongEq_def defined_def invalid_def bot_fun_def null_fun_def split: bool.split_asm HOL.if_split_asm option.split) have B: "(\ \(\ X)) \ (\ \ \(X->any\<^sub>S\<^sub>e\<^sub>t()))" apply(auto simp: OclANY_def OclValid_def true_def false_def StrongEq_def defined_def invalid_def valid_def bot_fun_def null_fun_def bot_option_def null_option_def null_is_valid OclAnd_def split: bool.split_asm HOL.if_split_asm option.split) apply(frule Set_inv_lemma[OF foundation16[THEN iffD2], OF conjI], simp) apply(subgoal_tac "(\ X) \ = true \") prefer 2 apply (metis (opaque_lifting, no_types) OclValid_def foundation16) apply(simp add: true_def, drule OclNotEmpty_has_elt[simplified OclValid_def true_def], simp) by(erule exE, insert someI2[where Q = "\x. x \ \" and P = "\y. y \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\"], simp) show ?thesis by(auto dest:A intro:B) qed lemma OclANY_valid_args_valid''[simp,code_unfold]: "\(X->any\<^sub>S\<^sub>e\<^sub>t()) = (\ X)" by(auto intro!: OclANY_valid_args_valid transform2_rev) (* and higher order ones : forall, exists, iterate, select, reject... *) text_raw\\endisatagafp\ subsection\Execution Laws with Invalid or Null or Infinite Set as Argument\ text\OclIncluding\ (* properties already generated by the corresponding locale *) text\OclExcluding\ (* properties already generated by the corresponding locale *) text\OclIncludes\ (* properties already generated by the corresponding locale *) text\OclExcludes\ (* properties already generated by the corresponding locale *) text\OclSize\ lemma OclSize_invalid[simp,code_unfold]:"(invalid->size\<^sub>S\<^sub>e\<^sub>t()) = invalid" by(simp add: bot_fun_def OclSize_def invalid_def defined_def valid_def false_def true_def) lemma OclSize_null[simp,code_unfold]:"(null->size\<^sub>S\<^sub>e\<^sub>t()) = invalid" by(rule ext, simp add: bot_fun_def null_fun_def null_is_valid OclSize_def invalid_def defined_def valid_def false_def true_def) text\OclIsEmpty\ lemma OclIsEmpty_invalid[simp,code_unfold]:"(invalid->isEmpty\<^sub>S\<^sub>e\<^sub>t()) = invalid" by(simp add: OclIsEmpty_def) lemma OclIsEmpty_null[simp,code_unfold]:"(null->isEmpty\<^sub>S\<^sub>e\<^sub>t()) = true" by(simp add: OclIsEmpty_def) text\OclNotEmpty\ lemma OclNotEmpty_invalid[simp,code_unfold]:"(invalid->notEmpty\<^sub>S\<^sub>e\<^sub>t()) = invalid" by(simp add: OclNotEmpty_def) lemma OclNotEmpty_null[simp,code_unfold]:"(null->notEmpty\<^sub>S\<^sub>e\<^sub>t()) = false" by(simp add: OclNotEmpty_def) text\OclANY\ lemma OclANY_invalid[simp,code_unfold]:"(invalid->any\<^sub>S\<^sub>e\<^sub>t()) = invalid" by(simp add: bot_fun_def OclANY_def invalid_def defined_def valid_def false_def true_def) lemma OclANY_null[simp,code_unfold]:"(null->any\<^sub>S\<^sub>e\<^sub>t()) = null" by(simp add: OclANY_def false_def true_def) text\OclForall\ lemma OclForall_invalid[simp,code_unfold]:"invalid->forAll\<^sub>S\<^sub>e\<^sub>t(a| P a) = invalid" by(simp add: bot_fun_def invalid_def OclForall_def defined_def valid_def false_def true_def) lemma OclForall_null[simp,code_unfold]:"null->forAll\<^sub>S\<^sub>e\<^sub>t(a | P a) = invalid" by(simp add: bot_fun_def invalid_def OclForall_def defined_def valid_def false_def true_def) text\OclExists\ lemma OclExists_invalid[simp,code_unfold]:"invalid->exists\<^sub>S\<^sub>e\<^sub>t(a| P a) = invalid" by(simp add: OclExists_def) lemma OclExists_null[simp,code_unfold]:"null->exists\<^sub>S\<^sub>e\<^sub>t(a | P a) = invalid" by(simp add: OclExists_def) text\OclIterate\ lemma OclIterate_invalid[simp,code_unfold]:"invalid->iterate\<^sub>S\<^sub>e\<^sub>t(a; x = A | P a x) = invalid" by(simp add: bot_fun_def invalid_def OclIterate_def defined_def valid_def false_def true_def) lemma OclIterate_null[simp,code_unfold]:"null->iterate\<^sub>S\<^sub>e\<^sub>t(a; x = A | P a x) = invalid" by(simp add: bot_fun_def invalid_def OclIterate_def defined_def valid_def false_def true_def) lemma OclIterate_invalid_args[simp,code_unfold]:"S->iterate\<^sub>S\<^sub>e\<^sub>t(a; x = invalid | P a x) = invalid" by(simp add: bot_fun_def invalid_def OclIterate_def defined_def valid_def false_def true_def) text\An open question is this ...\ lemma (*OclIterate_null_args[simp,code_unfold]:*) "S->iterate\<^sub>S\<^sub>e\<^sub>t(a; x = null | P a x) = invalid" oops (* In the definition above, this does not hold in general. And I believe, this is how it should be ... *) lemma OclIterate_infinite: assumes non_finite: "\ \ not(\(S->size\<^sub>S\<^sub>e\<^sub>t()))" shows "(OclIterate S A F) \ = invalid \" apply(insert non_finite [THEN OclSize_infinite]) apply(subst (asm) foundation9, simp) by(metis OclIterate_def OclValid_def invalid_def) text\OclSelect\ lemma OclSelect_invalid[simp,code_unfold]:"invalid->select\<^sub>S\<^sub>e\<^sub>t(a | P a) = invalid" by(simp add: bot_fun_def invalid_def OclSelect_def defined_def valid_def false_def true_def) lemma OclSelect_null[simp,code_unfold]:"null->select\<^sub>S\<^sub>e\<^sub>t(a | P a) = invalid" by(simp add: bot_fun_def invalid_def OclSelect_def defined_def valid_def false_def true_def) text\OclReject\ lemma OclReject_invalid[simp,code_unfold]:"invalid->reject\<^sub>S\<^sub>e\<^sub>t(a | P a) = invalid" by(simp add: OclReject_def) lemma OclReject_null[simp,code_unfold]:"null->reject\<^sub>S\<^sub>e\<^sub>t(a | P a) = invalid" by(simp add: OclReject_def) text_raw\\isatagafp\ subsubsection\Context Passing\ lemma cp_OclIncludes1: "(X->includes\<^sub>S\<^sub>e\<^sub>t(x)) \ = (X->includes\<^sub>S\<^sub>e\<^sub>t(\ _. x \)) \" by(auto simp: OclIncludes_def StrongEq_def invalid_def cp_defined[symmetric] cp_valid[symmetric]) lemma cp_OclSize: "X->size\<^sub>S\<^sub>e\<^sub>t() \ = ((\_. X \)->size\<^sub>S\<^sub>e\<^sub>t()) \" by(simp add: OclSize_def cp_defined[symmetric]) lemma cp_OclIsEmpty: "X->isEmpty\<^sub>S\<^sub>e\<^sub>t() \ = ((\_. X \)->isEmpty\<^sub>S\<^sub>e\<^sub>t()) \" apply(simp only: OclIsEmpty_def) apply(subst (2) cp_OclOr, subst cp_OclAnd, subst cp_OclNot, subst StrictRefEq\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r.cp0) by(simp add: cp_defined[symmetric] cp_valid[symmetric] StrictRefEq\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r.cp0[symmetric] cp_OclSize[symmetric] cp_OclNot[symmetric] cp_OclAnd[symmetric] cp_OclOr[symmetric]) lemma cp_OclNotEmpty: "X->notEmpty\<^sub>S\<^sub>e\<^sub>t() \ = ((\_. X \)->notEmpty\<^sub>S\<^sub>e\<^sub>t()) \" apply(simp only: OclNotEmpty_def) apply(subst (2) cp_OclNot) by(simp add: cp_OclNot[symmetric] cp_OclIsEmpty[symmetric]) lemma cp_OclANY: "X->any\<^sub>S\<^sub>e\<^sub>t() \ = ((\_. X \)->any\<^sub>S\<^sub>e\<^sub>t()) \" apply(simp only: OclANY_def) apply(subst (2) cp_OclAnd) by(simp only: cp_OclAnd[symmetric] cp_defined[symmetric] cp_valid[symmetric] cp_OclNotEmpty[symmetric]) lemma cp_OclForall: "(S->forAll\<^sub>S\<^sub>e\<^sub>t(x | P x)) \ = ((\ _. S \)->forAll\<^sub>S\<^sub>e\<^sub>t(x | P (\ _. x \))) \" by(simp add: OclForall_def cp_defined[symmetric]) (* first-order version !*) lemma cp_OclForall1 [simp,intro!]: "cp S \ cp (\X. ((S X)->forAll\<^sub>S\<^sub>e\<^sub>t(x | P x)))" apply(simp add: cp_def) apply(erule exE, rule exI, intro allI) apply(erule_tac x=X in allE) by(subst cp_OclForall, simp) lemma (*cp_OclForall2 [simp,intro!]:*) "cp (\X St x. P (\\. x) X St) \ cp S \ cp (\X. (S X)->forAll\<^sub>S\<^sub>e\<^sub>t(x|P x X)) " apply(simp only: cp_def) oops lemma (*cp_OclForall:*) "cp S \ (\ x. cp(P x)) \ cp(\X. ((S X)->forAll\<^sub>S\<^sub>e\<^sub>t(x | P x X)))" oops lemma cp_OclExists: "(S->exists\<^sub>S\<^sub>e\<^sub>t(x | P x)) \ = ((\ _. S \)->exists\<^sub>S\<^sub>e\<^sub>t(x | P (\ _. x \))) \" by(simp add: OclExists_def OclNot_def, subst cp_OclForall, simp) (* first-order version !*) lemma cp_OclExists1 [simp,intro!]: "cp S \ cp (\X. ((S X)->exists\<^sub>S\<^sub>e\<^sub>t(x | P x)))" apply(simp add: cp_def) apply(erule exE, rule exI, intro allI) apply(erule_tac x=X in allE) by(subst cp_OclExists,simp) lemma cp_OclIterate: "(X->iterate\<^sub>S\<^sub>e\<^sub>t(a; x = A | P a x)) \ = ((\ _. X \)->iterate\<^sub>S\<^sub>e\<^sub>t(a; x = A | P a x)) \" by(simp add: OclIterate_def cp_defined[symmetric]) lemma cp_OclSelect: "(X->select\<^sub>S\<^sub>e\<^sub>t(a | P a)) \ = ((\ _. X \)->select\<^sub>S\<^sub>e\<^sub>t(a | P a)) \" by(simp add: OclSelect_def cp_defined[symmetric]) lemma cp_OclReject: "(X->reject\<^sub>S\<^sub>e\<^sub>t(a | P a)) \ = ((\ _. X \)->reject\<^sub>S\<^sub>e\<^sub>t(a | P a)) \" by(simp add: OclReject_def, subst cp_OclSelect, simp) lemmas cp_intro''\<^sub>S\<^sub>e\<^sub>t[intro!,simp,code_unfold] = cp_OclSize [THEN allI[THEN allI[THEN cpI1], of "OclSize"]] cp_OclIsEmpty [THEN allI[THEN allI[THEN cpI1], of "OclIsEmpty"]] cp_OclNotEmpty [THEN allI[THEN allI[THEN cpI1], of "OclNotEmpty"]] cp_OclANY [THEN allI[THEN allI[THEN cpI1], of "OclANY"]] subsubsection\Const\ lemma const_OclIncluding[simp,code_unfold] : assumes const_x : "const x" and const_S : "const S" shows "const (S->including\<^sub>S\<^sub>e\<^sub>t(x))" proof - have A:"\\ \'. \ (\ \ \ x) \ (S->including\<^sub>S\<^sub>e\<^sub>t(x) \) = (S->including\<^sub>S\<^sub>e\<^sub>t(x) \')" apply(simp add: foundation18) apply(erule const_subst[OF const_x const_invalid],simp_all) by(rule const_charn[OF const_invalid]) have B: "\ \ \'. \ (\ \ \ S) \ (S->including\<^sub>S\<^sub>e\<^sub>t(x) \) = (S->including\<^sub>S\<^sub>e\<^sub>t(x) \')" apply(simp add: foundation16', elim disjE) apply(erule const_subst[OF const_S const_invalid],simp_all) apply(rule const_charn[OF const_invalid]) apply(erule const_subst[OF const_S const_null],simp_all) by(rule const_charn[OF const_invalid]) show ?thesis apply(simp only: const_def,intro allI, rename_tac \ \') apply(case_tac "\ (\ \ \ x)", simp add: A) apply(case_tac "\ (\ \ \ S)", simp_all add: B) apply(frule_tac \'1= \' in const_OclValid2[OF const_x, THEN iffD1]) apply(frule_tac \'1= \' in const_OclValid1[OF const_S, THEN iffD1]) apply(simp add: OclIncluding_def OclValid_def) apply(subst const_charn[OF const_x]) apply(subst const_charn[OF const_S]) by simp qed text_raw\\endisatagafp\ subsection\General Algebraic Execution Rules\ subsubsection\Execution Rules on Including\ lemma OclIncluding_finite_rep_set : assumes X_def : "\ \ \ X" and x_val : "\ \ \ x" shows "finite \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X->including\<^sub>S\<^sub>e\<^sub>t(x) \)\\ = finite \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\" proof - have C : "\\insert (x \) \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\\\ \ {X. X = bot \ X = null \ (\x\\\X\\. x \ bot)}" by(insert X_def x_val, frule Set_inv_lemma, simp add: foundation18 invalid_def) show "?thesis" by(insert X_def x_val, auto simp: OclIncluding_def Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse[OF C] dest: foundation13[THEN iffD2, THEN foundation22[THEN iffD1]]) qed lemma OclIncluding_rep_set: assumes S_def: "\ \ \ S" shows "\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S->including\<^sub>S\<^sub>e\<^sub>t(\_. \\x\\) \)\\ = insert \\x\\ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\" apply(simp add: OclIncluding_def S_def[simplified OclValid_def]) apply(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse, simp add: bot_option_def null_option_def) apply(insert Set_inv_lemma[OF S_def], metis bot_option_def not_Some_eq) by(simp) lemma OclIncluding_notempty_rep_set: assumes X_def: "\ \ \ X" and a_val: "\ \ \ a" shows "\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X->including\<^sub>S\<^sub>e\<^sub>t(a) \)\\ \ {}" apply(simp add: OclIncluding_def X_def[simplified OclValid_def] a_val[simplified OclValid_def]) apply(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse, simp add: bot_option_def null_option_def) apply(insert Set_inv_lemma[OF X_def], metis a_val foundation18') by(simp) lemma OclIncluding_includes0: assumes "\ \ X->includes\<^sub>S\<^sub>e\<^sub>t(x)" shows "X->including\<^sub>S\<^sub>e\<^sub>t(x) \ = X \" proof - have includes_def: "\ \ X->includes\<^sub>S\<^sub>e\<^sub>t(x) \ \ \ \ X" by (metis bot_fun_def OclIncludes_def OclValid_def defined3 foundation16) have includes_val: "\ \ X->includes\<^sub>S\<^sub>e\<^sub>t(x) \ \ \ \ x" using foundation5 foundation6 by fastforce show ?thesis apply(insert includes_def[OF assms] includes_val[OF assms] assms, simp add: OclIncluding_def OclIncludes_def OclValid_def true_def) apply(drule insert_absorb, simp, subst abs_rep_simp') by(simp_all add: OclValid_def true_def) qed lemma OclIncluding_includes: assumes "\ \ X->includes\<^sub>S\<^sub>e\<^sub>t(x)" shows "\ \ X->including\<^sub>S\<^sub>e\<^sub>t(x) \ X" by(simp add: StrongEq_def OclValid_def true_def OclIncluding_includes0[OF assms]) lemma OclIncluding_commute0 : assumes S_def : "\ \ \ S" and i_val : "\ \ \ i" and j_val : "\ \ \ j" shows "\ \ ((S :: ('\, 'a::null) Set)->including\<^sub>S\<^sub>e\<^sub>t(i)->including\<^sub>S\<^sub>e\<^sub>t(j) \ (S->including\<^sub>S\<^sub>e\<^sub>t(j)->including\<^sub>S\<^sub>e\<^sub>t(i)))" proof - have A : "\\insert (i \) \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\\\ \ {X. X = bot \ X = null \ (\x\\\X\\. x \ bot)}" by(insert S_def i_val, frule Set_inv_lemma, simp add: foundation18 invalid_def) have B : "\\insert (j \) \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\\\ \ {X. X = bot \ X = null \ (\x\\\X\\. x \ bot)}" by(insert S_def j_val, frule Set_inv_lemma, simp add: foundation18 invalid_def) have G1 : "Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\insert (i \) \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\\\ \ Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e None" by(insert A, simp add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject bot_option_def null_option_def) have G2 : "Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\insert (i \) \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\\\ \ Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \None\" by(insert A, simp add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject bot_option_def null_option_def) have G3 : "Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\insert (j \) \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\\\ \ Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e None" by(insert B, simp add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject bot_option_def null_option_def) have G4 : "Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\insert (j \) \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\\\ \ Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \None\" by(insert B, simp add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject bot_option_def null_option_def) have * : "(\ (\_. Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\insert (i \) \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\\\)) \ = \\True\\" by(auto simp: OclValid_def false_def defined_def null_fun_def true_def bot_fun_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def S_def i_val G1 G2) have ** : "(\ (\_. Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\insert (j \) \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\\\)) \ = \\True\\" by(auto simp: OclValid_def false_def defined_def null_fun_def true_def bot_fun_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def S_def i_val G3 G4) have *** : "Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\insert(j \)\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e(Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e\\insert(i \)\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e(S \)\\\\)\\\\ = Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\insert(i \)\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e(Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e\\insert(j \)\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e(S \)\\\\)\\\\" by(simp add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse[OF A] Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse[OF B] Set.insert_commute) show ?thesis apply(simp add: OclIncluding_def S_def[simplified OclValid_def] i_val[simplified OclValid_def] j_val[simplified OclValid_def] true_def OclValid_def StrongEq_def) apply(subst cp_defined, simp add: S_def[simplified OclValid_def] i_val[simplified OclValid_def] j_val[simplified OclValid_def] true_def *) apply(subst cp_defined, simp add: S_def[simplified OclValid_def] i_val[simplified OclValid_def] j_val[simplified OclValid_def] true_def ** ***) apply(subst cp_defined, simp add: S_def[simplified OclValid_def] i_val[simplified OclValid_def] j_val[simplified OclValid_def] true_def *) apply(subst cp_defined, simp add: S_def[simplified OclValid_def] i_val[simplified OclValid_def] j_val[simplified OclValid_def] true_def * ) apply(subst cp_defined, simp add: S_def[simplified OclValid_def] i_val[simplified OclValid_def] j_val[simplified OclValid_def] true_def * **) done qed lemma OclIncluding_commute[simp,code_unfold]: "((S :: ('\, 'a::null) Set)->including\<^sub>S\<^sub>e\<^sub>t(i)->including\<^sub>S\<^sub>e\<^sub>t(j) = (S->including\<^sub>S\<^sub>e\<^sub>t(j)->including\<^sub>S\<^sub>e\<^sub>t(i)))" proof - have A: "\ \. \ \ (i \ invalid) \ (S->including\<^sub>S\<^sub>e\<^sub>t(i)->including\<^sub>S\<^sub>e\<^sub>t(j)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) have A': "\ \. \ \ (i \ invalid) \ (S->including\<^sub>S\<^sub>e\<^sub>t(j)->including\<^sub>S\<^sub>e\<^sub>t(i)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) have B:"\ \. \ \ (j \ invalid) \ (S->including\<^sub>S\<^sub>e\<^sub>t(i)->including\<^sub>S\<^sub>e\<^sub>t(j)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) have B':"\ \. \ \ (j \ invalid) \ (S->including\<^sub>S\<^sub>e\<^sub>t(j)->including\<^sub>S\<^sub>e\<^sub>t(i)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) have C: "\ \. \ \ (S \ invalid) \ (S->including\<^sub>S\<^sub>e\<^sub>t(i)->including\<^sub>S\<^sub>e\<^sub>t(j)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) have C': "\ \. \ \ (S \ invalid) \ (S->including\<^sub>S\<^sub>e\<^sub>t(j)->including\<^sub>S\<^sub>e\<^sub>t(i)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) have D: "\ \. \ \ (S \ null) \ (S->including\<^sub>S\<^sub>e\<^sub>t(i)->including\<^sub>S\<^sub>e\<^sub>t(j)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) have D': "\ \. \ \ (S \ null) \ (S->including\<^sub>S\<^sub>e\<^sub>t(j)->including\<^sub>S\<^sub>e\<^sub>t(i)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) show ?thesis apply(rule ext, rename_tac \) apply(case_tac "\ \ (\ i)") apply(case_tac "\ \ (\ j)") apply(case_tac "\ \ (\ S)") apply(simp only: OclIncluding_commute0[THEN foundation22[THEN iffD1]]) apply(simp add: foundation16', elim disjE) apply(simp add: C[OF foundation22[THEN iffD2]] C'[OF foundation22[THEN iffD2]]) apply(simp add: D[OF foundation22[THEN iffD2]] D'[OF foundation22[THEN iffD2]]) apply(simp add:foundation18 B[OF foundation22[THEN iffD2]] B'[OF foundation22[THEN iffD2]]) apply(simp add:foundation18 A[OF foundation22[THEN iffD2]] A'[OF foundation22[THEN iffD2]]) done qed subsubsection\Execution Rules on Excluding\ lemma OclExcluding_finite_rep_set : assumes X_def : "\ \ \ X" and x_val : "\ \ \ x" shows "finite \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X->excluding\<^sub>S\<^sub>e\<^sub>t(x) \)\\ = finite \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\" proof - have C : "\\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\ - {x \}\\ \ {X. X = bot \ X = null \ (\x\\\X\\. x \ bot)}" apply(insert X_def x_val, frule Set_inv_lemma) apply(simp add: foundation18 invalid_def) done show "?thesis" by(insert X_def x_val, auto simp: OclExcluding_def Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse[OF C] dest: foundation13[THEN iffD2, THEN foundation22[THEN iffD1]]) qed lemma OclExcluding_rep_set: assumes S_def: "\ \ \ S" shows "\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S->excluding\<^sub>S\<^sub>e\<^sub>t(\_. \\x\\) \)\\ = \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\ - {\\x\\}" apply(simp add: OclExcluding_def S_def[simplified OclValid_def]) apply(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse, simp add: bot_option_def null_option_def) apply(insert Set_inv_lemma[OF S_def], metis Diff_iff bot_option_def not_None_eq) by(simp) lemma OclExcluding_excludes0: assumes "\ \ X->excludes\<^sub>S\<^sub>e\<^sub>t(x)" shows "X->excluding\<^sub>S\<^sub>e\<^sub>t(x) \ = X \" proof - have excludes_def: "\ \ X->excludes\<^sub>S\<^sub>e\<^sub>t(x) \ \ \ \ X" by (metis OclExcludes.def_valid_then_def OclExcludes_valid_args_valid'' foundation10' foundation6) have excludes_val: "\ \ X->excludes\<^sub>S\<^sub>e\<^sub>t(x) \ \ \ \ x" by (metis OclExcludes.def_valid_then_def OclExcludes_valid_args_valid'' foundation10' foundation6) show ?thesis apply(insert excludes_def[OF assms] excludes_val[OF assms] assms, simp add: OclExcluding_def OclExcludes_def OclIncludes_def OclNot_def OclValid_def true_def) by (metis (opaque_lifting, no_types) abs_rep_simp' assms excludes_def) qed lemma OclExcluding_excludes: assumes "\ \ X->excludes\<^sub>S\<^sub>e\<^sub>t(x)" shows "\ \ X->excluding\<^sub>S\<^sub>e\<^sub>t(x) \ X" by(simp add: StrongEq_def OclValid_def true_def OclExcluding_excludes0[OF assms]) lemma OclExcluding_charn0[simp]: assumes val_x:"\ \ (\ x)" shows "\ \ ((Set{}->excluding\<^sub>S\<^sub>e\<^sub>t(x)) \ Set{})" proof - have A : "\None\ \ {X. X = bot \ X = null \ (\x\\\X\\. x \ bot)}" by(simp add: null_option_def bot_option_def) have B : "\\{}\\ \ {X. X = bot \ X = null \ (\x\\\X\\. x \ bot)}" by(simp add: mtSet_def) show ?thesis using val_x apply(auto simp: OclValid_def OclIncludes_def OclNot_def false_def true_def StrongEq_def OclExcluding_def mtSet_def defined_def bot_fun_def null_fun_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def) apply(auto simp: mtSet_def Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e.Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e.Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject[OF B A]) done qed lemma OclExcluding_commute0 : assumes S_def : "\ \ \ S" and i_val : "\ \ \ i" and j_val : "\ \ \ j" shows "\ \ ((S :: ('\, 'a::null) Set)->excluding\<^sub>S\<^sub>e\<^sub>t(i)->excluding\<^sub>S\<^sub>e\<^sub>t(j) \ (S->excluding\<^sub>S\<^sub>e\<^sub>t(j)->excluding\<^sub>S\<^sub>e\<^sub>t(i)))" proof - have A : "\\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\ - {i \}\\ \ {X. X = bot \ X = null \ (\x\\\X\\. x \ bot)}" by(insert S_def i_val, frule Set_inv_lemma, simp add: foundation18 invalid_def) have B : "\\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\ - {j \}\\ \ {X. X = bot \ X = null \ (\x\\\X\\. x \ bot)}" by(insert S_def j_val, frule Set_inv_lemma, simp add: foundation18 invalid_def) have G1 : "Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\ - {i \}\\ \ Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e None" by(insert A, simp add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject bot_option_def null_option_def) have G2 : "Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\ - {i \}\\ \ Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \None\" by(insert A, simp add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject bot_option_def null_option_def) have G3 : "Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\ - {j \}\\ \ Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e None" by(insert B, simp add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject bot_option_def null_option_def) have G4 : "Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\ - {j \}\\ \ Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \None\" by(insert B, simp add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject bot_option_def null_option_def) have * : "(\ (\_. Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\ - {i \}\\)) \ = \\True\\" by(auto simp: OclValid_def false_def defined_def null_fun_def true_def bot_fun_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def S_def i_val G1 G2) have ** : "(\ (\_. Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\ - {j \}\\)) \ = \\True\\" by(auto simp: OclValid_def false_def defined_def null_fun_def true_def bot_fun_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def S_def i_val G3 G4) have *** : "Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e(Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e\\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e(S \)\\-{i \}\\)\\-{j \}\\ = Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e(Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e\\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e(S \)\\-{j \}\\)\\-{i \}\\" apply(simp add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse[OF A] Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse[OF B]) by (metis Diff_insert2 insert_commute) show ?thesis apply(simp add: OclExcluding_def S_def[simplified OclValid_def] i_val[simplified OclValid_def] j_val[simplified OclValid_def] true_def OclValid_def StrongEq_def) apply(subst cp_defined, simp add: S_def[simplified OclValid_def] i_val[simplified OclValid_def] j_val[simplified OclValid_def] true_def *) apply(subst cp_defined, simp add: S_def[simplified OclValid_def] i_val[simplified OclValid_def] j_val[simplified OclValid_def] true_def ** ***) apply(subst cp_defined, simp add: S_def[simplified OclValid_def] i_val[simplified OclValid_def] j_val[simplified OclValid_def] true_def *) apply(subst cp_defined, simp add: S_def[simplified OclValid_def] i_val[simplified OclValid_def] j_val[simplified OclValid_def] true_def * ) apply(subst cp_defined, simp add: S_def[simplified OclValid_def] i_val[simplified OclValid_def] j_val[simplified OclValid_def] true_def * **) done qed lemma OclExcluding_commute[simp,code_unfold]: "((S :: ('\, 'a::null) Set)->excluding\<^sub>S\<^sub>e\<^sub>t(i)->excluding\<^sub>S\<^sub>e\<^sub>t(j) = (S->excluding\<^sub>S\<^sub>e\<^sub>t(j)->excluding\<^sub>S\<^sub>e\<^sub>t(i)))" proof - have A: "\ \. \ \ i \ invalid \ (S->excluding\<^sub>S\<^sub>e\<^sub>t(i)->excluding\<^sub>S\<^sub>e\<^sub>t(j)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) have A': "\ \. \ \ i \ invalid \ (S->excluding\<^sub>S\<^sub>e\<^sub>t(j)->excluding\<^sub>S\<^sub>e\<^sub>t(i)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) have B:"\ \. \ \ j \ invalid \ (S->excluding\<^sub>S\<^sub>e\<^sub>t(i)->excluding\<^sub>S\<^sub>e\<^sub>t(j)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) have B':"\ \. \ \ j \ invalid \ (S->excluding\<^sub>S\<^sub>e\<^sub>t(j)->excluding\<^sub>S\<^sub>e\<^sub>t(i)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) have C: "\ \. \ \ S \ invalid \ (S->excluding\<^sub>S\<^sub>e\<^sub>t(i)->excluding\<^sub>S\<^sub>e\<^sub>t(j)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) have C': "\ \. \ \ S \ invalid \ (S->excluding\<^sub>S\<^sub>e\<^sub>t(j)->excluding\<^sub>S\<^sub>e\<^sub>t(i)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) have D: "\ \. \ \ S \ null \ (S->excluding\<^sub>S\<^sub>e\<^sub>t(i)->excluding\<^sub>S\<^sub>e\<^sub>t(j)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) have D': "\ \. \ \ S \ null \ (S->excluding\<^sub>S\<^sub>e\<^sub>t(j)->excluding\<^sub>S\<^sub>e\<^sub>t(i)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) show ?thesis apply(rule ext, rename_tac \) apply(case_tac "\ \ (\ i)") apply(case_tac "\ \ (\ j)") apply(case_tac "\ \ (\ S)") apply(simp only: OclExcluding_commute0[THEN foundation22[THEN iffD1]]) apply(simp add: foundation16', elim disjE) apply(simp add: C[OF foundation22[THEN iffD2]] C'[OF foundation22[THEN iffD2]]) apply(simp add: D[OF foundation22[THEN iffD2]] D'[OF foundation22[THEN iffD2]]) apply(simp add:foundation18 B[OF foundation22[THEN iffD2]] B'[OF foundation22[THEN iffD2]]) apply(simp add:foundation18 A[OF foundation22[THEN iffD2]] A'[OF foundation22[THEN iffD2]]) done qed lemma OclExcluding_charn0_exec[simp,code_unfold]: "(Set{}->excluding\<^sub>S\<^sub>e\<^sub>t(x)) = (if (\ x) then Set{} else invalid endif)" proof - have A: "\ \. (Set{}->excluding\<^sub>S\<^sub>e\<^sub>t(invalid)) \ = (if (\ invalid) then Set{} else invalid endif) \" by simp have B: "\ \ x. \ \ (\ x) \ (Set{}->excluding\<^sub>S\<^sub>e\<^sub>t(x)) \ = (if (\ x) then Set{} else invalid endif) \" by(simp add: OclExcluding_charn0[THEN foundation22[THEN iffD1]]) show ?thesis apply(rule ext, rename_tac \) apply(case_tac "\ \ (\ x)") apply(simp add: B) apply(simp add: foundation18) apply(subst OclExcluding.cp0, simp) apply(simp add: cp_OclIf[symmetric] OclExcluding.cp0[symmetric] cp_valid[symmetric] A) done qed lemma OclExcluding_charn1: assumes def_X:"\ \ (\ X)" and val_x:"\ \ (\ x)" and val_y:"\ \ (\ y)" and neq :"\ \ not(x \ y)" shows "\ \ ((X->including\<^sub>S\<^sub>e\<^sub>t(x))->excluding\<^sub>S\<^sub>e\<^sub>t(y)) \ ((X->excluding\<^sub>S\<^sub>e\<^sub>t(y))->including\<^sub>S\<^sub>e\<^sub>t(x))" proof - have C : "\\insert (x \) \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\\\ \ {X. X = bot \ X = null \ (\x\\\X\\. x \ bot)}" by(insert def_X val_x, frule Set_inv_lemma, simp add: foundation18 invalid_def) have D : "\\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\ - {y \}\\ \ {X. X = bot \ X = null \ (\x\\\X\\. x \ bot)}" by(insert def_X val_x, frule Set_inv_lemma, simp add: foundation18 invalid_def) have E : "x \ \ y \" by(insert neq, auto simp: OclValid_def bot_fun_def OclIncluding_def OclIncludes_def false_def true_def defined_def valid_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_fun_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def StrongEq_def OclNot_def) have G1 : "Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\insert (x \) \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\\\ \ Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e None" by(insert C, simp add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject bot_option_def null_option_def) have G2 : "Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\insert (x \) \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\\\ \ Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \None\" by(insert C, simp add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject bot_option_def null_option_def) have G : "(\ (\_. Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\insert (x \) \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\\\)) \ = true \" by(auto simp: OclValid_def false_def true_def defined_def bot_fun_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_fun_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def G1 G2) have H1 : "Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\ - {y \}\\ \ Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e None" by(insert D, simp add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject bot_option_def null_option_def) have H2 : "Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\ - {y \}\\ \ Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \None\" by(insert D, simp add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject bot_option_def null_option_def) have H : "(\ (\_. Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\ - {y \}\\)) \ = true \" by(auto simp: OclValid_def false_def true_def defined_def bot_fun_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_fun_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def H1 H2) have Z : "insert (x \) \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\ - {y \} = insert (x \) (\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\ - {y \})" by(auto simp: E) show ?thesis apply(insert def_X[THEN foundation13[THEN iffD2]] val_x[THEN foundation13[THEN iffD2]] val_y[THEN foundation13[THEN iffD2]]) apply(simp add: foundation22 OclIncluding_def OclExcluding_def def_X[THEN foundation16[THEN iffD1]]) apply(subst cp_defined, simp)+ apply(simp add: G H Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse[OF C] Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse[OF D] Z) done qed lemma OclExcluding_charn2: assumes def_X:"\ \ (\ X)" and val_x:"\ \ (\ x)" shows "\ \ (((X->including\<^sub>S\<^sub>e\<^sub>t(x))->excluding\<^sub>S\<^sub>e\<^sub>t(x)) \ (X->excluding\<^sub>S\<^sub>e\<^sub>t(x)))" proof - have C : "\\insert (x \) \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\\\ \ {X. X = bot \ X = null \ (\x\\\X\\. x \ bot)}" by(insert def_X val_x, frule Set_inv_lemma, simp add: foundation18 invalid_def) have G1 : "Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\insert (x \) \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\\\ \ Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e None" by(insert C, simp add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject bot_option_def null_option_def) have G2 : "Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\insert (x \) \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\\\ \ Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \None\" by(insert C, simp add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject bot_option_def null_option_def) show ?thesis apply(insert def_X[THEN foundation16[THEN iffD1]] val_x[THEN foundation18[THEN iffD1]]) apply(auto simp: OclValid_def bot_fun_def OclIncluding_def OclIncludes_def false_def true_def invalid_def defined_def valid_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_fun_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def StrongEq_def) apply(subst OclExcluding.cp0) apply(auto simp:OclExcluding_def) apply(simp add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse[OF C]) apply(simp_all add: false_def true_def defined_def valid_def null_fun_def bot_fun_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def split: bool.split_asm HOL.if_split_asm option.split) apply(auto simp: G1 G2) done qed theorem OclExcluding_charn3: "((X->including\<^sub>S\<^sub>e\<^sub>t(x))->excluding\<^sub>S\<^sub>e\<^sub>t(x)) = (X->excluding\<^sub>S\<^sub>e\<^sub>t(x))" proof - have A1 : "\\. \ \ (X \ invalid) \ (X->including\<^sub>S\<^sub>e\<^sub>t(x)->excluding\<^sub>S\<^sub>e\<^sub>t(x)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) have A1': "\\. \ \ (X \ invalid) \ (X->excluding\<^sub>S\<^sub>e\<^sub>t(x)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) have A2 : "\\. \ \ (X \ null) \ (X->including\<^sub>S\<^sub>e\<^sub>t(x)->excluding\<^sub>S\<^sub>e\<^sub>t(x)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) have A2': "\\. \ \ (X \ null) \ (X->excluding\<^sub>S\<^sub>e\<^sub>t(x)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) have A3 : "\\. \ \ (x \ invalid) \ (X->including\<^sub>S\<^sub>e\<^sub>t(x)->excluding\<^sub>S\<^sub>e\<^sub>t(x)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) have A3': "\\. \ \ (x \ invalid) \ (X->excluding\<^sub>S\<^sub>e\<^sub>t(x)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) show ?thesis apply(rule ext, rename_tac "\") apply(case_tac "\ \ (\ x)") apply(case_tac "\ \ (\ X)") apply(simp only: OclExcluding_charn2[THEN foundation22[THEN iffD1]]) apply(simp add: foundation16', elim disjE) apply(simp add: A1[OF foundation22[THEN iffD2]] A1'[OF foundation22[THEN iffD2]]) apply(simp add: A2[OF foundation22[THEN iffD2]] A2'[OF foundation22[THEN iffD2]]) apply(simp add:foundation18 A3[OF foundation22[THEN iffD2]] A3'[OF foundation22[THEN iffD2]]) done qed text\One would like a generic theorem of the form: \begin{isar}[mathescape] lemma OclExcluding_charn_exec: "(X->including$_{Set}$(x::('$\mathfrak{A}$,'a::null)val)->excluding$_{Set}$(y)) = (if \ X then if x \ y then X->excluding$_{Set}$(y) else X->excluding$_{Set}$(y)->including$_{Set}$(x) endif else invalid endif)" \end{isar} Unfortunately, this does not hold in general, since referential equality is an overloaded concept and has to be defined for each type individually. Consequently, it is only valid for concrete type instances for Boolean, Integer, and Sets thereof... \ text\The computational law \emph{OclExcluding-charn-exec} becomes generic since it uses strict equality which in itself is generic. It is possible to prove the following generic theorem and instantiate it later (using properties that link the polymorphic logical strong equality with the concrete instance of strict quality).\ lemma OclExcluding_charn_exec: assumes strict1: "(invalid \ y) = invalid" and strict2: "(x \ invalid) = invalid" and StrictRefEq_valid_args_valid: "\ (x::('\,'a::null)val) y \. (\ \ \ (x \ y)) = ((\ \ (\ x)) \ (\ \ \ y))" and cp_StrictRefEq: "\ (X::('\,'a::null)val) Y \. (X \ Y) \ = ((\_. X \) \ (\_. Y \)) \" and StrictRefEq_vs_StrongEq: "\ (x::('\,'a::null)val) y \. \ \ \ x \ \ \ \ y \ (\ \ ((x \ y) \ (x \ y)))" shows "(X->including\<^sub>S\<^sub>e\<^sub>t(x::('\,'a::null)val)->excluding\<^sub>S\<^sub>e\<^sub>t(y)) = (if \ X then if x \ y then X->excluding\<^sub>S\<^sub>e\<^sub>t(y) else X->excluding\<^sub>S\<^sub>e\<^sub>t(y)->including\<^sub>S\<^sub>e\<^sub>t(x) endif else invalid endif)" proof - (* Lifting theorems, largely analogous OclIncludes_execute_generic, with the same problems wrt. strict equality. *) have A1: "\\. \ \ (X \ invalid) \ (X->including\<^sub>S\<^sub>e\<^sub>t(x)->includes\<^sub>S\<^sub>e\<^sub>t(y)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) have B1: "\\. \ \ (X \ null) \ (X->including\<^sub>S\<^sub>e\<^sub>t(x)->includes\<^sub>S\<^sub>e\<^sub>t(y)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) have A2: "\\. \ \ (X \ invalid) \ X->including\<^sub>S\<^sub>e\<^sub>t(x)->excluding\<^sub>S\<^sub>e\<^sub>t(y) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) have B2: "\\. \ \ (X \ null) \ X->including\<^sub>S\<^sub>e\<^sub>t(x)->excluding\<^sub>S\<^sub>e\<^sub>t(y) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) note [simp] = cp_StrictRefEq [THEN allI[THEN allI[THEN allI[THEN cpI2]], of "StrictRefEq"]] have C: "\\. \ \ (x \ invalid) \ (X->including\<^sub>S\<^sub>e\<^sub>t(x)->excluding\<^sub>S\<^sub>e\<^sub>t(y)) \ = (if x \ y then X->excluding\<^sub>S\<^sub>e\<^sub>t(y) else X->excluding\<^sub>S\<^sub>e\<^sub>t(y)->including\<^sub>S\<^sub>e\<^sub>t(x) endif) \" apply(rule foundation22[THEN iffD1]) apply(erule StrongEq_L_subst2_rev,simp,simp) by(simp add: strict1) have D: "\\. \ \ (y \ invalid) \ (X->including\<^sub>S\<^sub>e\<^sub>t(x)->excluding\<^sub>S\<^sub>e\<^sub>t(y)) \ = (if x \ y then X->excluding\<^sub>S\<^sub>e\<^sub>t(y) else X->excluding\<^sub>S\<^sub>e\<^sub>t(y)->including\<^sub>S\<^sub>e\<^sub>t(x) endif) \" apply(rule foundation22[THEN iffD1]) apply(erule StrongEq_L_subst2_rev,simp,simp) by (simp add: strict2) have E: "\\. \ \ \ x \ \ \ \ y \ (if x \ y then X->excluding\<^sub>S\<^sub>e\<^sub>t(y) else X->excluding\<^sub>S\<^sub>e\<^sub>t(y)->including\<^sub>S\<^sub>e\<^sub>t(x) endif) \ = (if x \ y then X->excluding\<^sub>S\<^sub>e\<^sub>t(y) else X->excluding\<^sub>S\<^sub>e\<^sub>t(y)->including\<^sub>S\<^sub>e\<^sub>t(x) endif) \" apply(subst cp_OclIf) apply(subst StrictRefEq_vs_StrongEq[THEN foundation22[THEN iffD1]]) by(simp_all add: cp_OclIf[symmetric]) have F: "\\. \ \ \ X \ \ \ \ x \ \ \ (x \ y) \ (X->including\<^sub>S\<^sub>e\<^sub>t(x)->excluding\<^sub>S\<^sub>e\<^sub>t(y) \) = (X->excluding\<^sub>S\<^sub>e\<^sub>t(y) \)" apply(drule StrongEq_L_sym) apply(rule foundation22[THEN iffD1]) apply(erule StrongEq_L_subst2_rev,simp) by(simp add: OclExcluding_charn2) show ?thesis apply(rule ext, rename_tac "\") apply(case_tac "\ (\ \ (\ X))", simp add:defined_split,elim disjE A1 B1 A2 B2) apply(case_tac "\ (\ \ (\ x))", simp add:foundation18 foundation22[symmetric], drule StrongEq_L_sym) apply(simp add: foundation22 C) apply(case_tac "\ (\ \ (\ y))", simp add:foundation18 foundation22[symmetric], drule StrongEq_L_sym, simp add: foundation22 D, simp) apply(subst E,simp_all) apply(case_tac "\ \ not (x \ y)") apply(simp add: OclExcluding_charn1[simplified foundation22] OclExcluding_charn2[simplified foundation22]) apply(simp add: foundation9 F) done qed (* Hack to work around OF-Bug *) schematic_goal OclExcluding_charn_exec\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r[simp,code_unfold]: "?X" by(rule OclExcluding_charn_exec[OF StrictRefEq\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r.strict1 StrictRefEq\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r.strict2 StrictRefEq\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r.defined_args_valid StrictRefEq\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r.cp0 StrictRefEq\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r.StrictRefEq_vs_StrongEq], simp_all) schematic_goal OclExcluding_charn_exec\<^sub>B\<^sub>o\<^sub>o\<^sub>l\<^sub>e\<^sub>a\<^sub>n[simp,code_unfold]: "?X" by(rule OclExcluding_charn_exec[OF StrictRefEq\<^sub>B\<^sub>o\<^sub>o\<^sub>l\<^sub>e\<^sub>a\<^sub>n.strict1 StrictRefEq\<^sub>B\<^sub>o\<^sub>o\<^sub>l\<^sub>e\<^sub>a\<^sub>n.strict2 StrictRefEq\<^sub>B\<^sub>o\<^sub>o\<^sub>l\<^sub>e\<^sub>a\<^sub>n.defined_args_valid StrictRefEq\<^sub>B\<^sub>o\<^sub>o\<^sub>l\<^sub>e\<^sub>a\<^sub>n.cp0 StrictRefEq\<^sub>B\<^sub>o\<^sub>o\<^sub>l\<^sub>e\<^sub>a\<^sub>n.StrictRefEq_vs_StrongEq], simp_all) schematic_goal OclExcluding_charn_exec\<^sub>S\<^sub>e\<^sub>t[simp,code_unfold]: "?X" by(rule OclExcluding_charn_exec[OF StrictRefEq\<^sub>S\<^sub>e\<^sub>t.strict1 StrictRefEq\<^sub>S\<^sub>e\<^sub>t.strict2 StrictRefEq\<^sub>S\<^sub>e\<^sub>t.defined_args_valid StrictRefEq\<^sub>S\<^sub>e\<^sub>t.cp0 StrictRefEq\<^sub>S\<^sub>e\<^sub>t.StrictRefEq_vs_StrongEq], simp_all) subsubsection\Execution Rules on Includes\ lemma OclIncludes_charn0[simp]: assumes val_x:"\ \ (\ x)" shows "\ \ not(Set{}->includes\<^sub>S\<^sub>e\<^sub>t(x))" using val_x apply(auto simp: OclValid_def OclIncludes_def OclNot_def false_def true_def) apply(auto simp: mtSet_def Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e.Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse) done lemma OclIncludes_charn0'[simp,code_unfold]: "Set{}->includes\<^sub>S\<^sub>e\<^sub>t(x) = (if \ x then false else invalid endif)" proof - have A: "\ \. (Set{}->includes\<^sub>S\<^sub>e\<^sub>t(invalid)) \ = (if (\ invalid) then false else invalid endif) \" by simp have B: "\ \ x. \ \ (\ x) \ (Set{}->includes\<^sub>S\<^sub>e\<^sub>t(x)) \ = (if \ x then false else invalid endif) \" apply(frule OclIncludes_charn0, simp add: OclValid_def) apply(rule foundation21[THEN fun_cong, simplified StrongEq_def,simplified, THEN iffD1, of _ _ "false"]) by simp show ?thesis apply(rule ext, rename_tac \) apply(case_tac "\ \ (\ x)") apply(simp_all add: B foundation18) apply(subst OclIncludes.cp0, simp add: OclIncludes.cp0[symmetric] A) done qed lemma OclIncludes_charn1: assumes def_X:"\ \ (\ X)" assumes val_x:"\ \ (\ x)" shows "\ \ (X->including\<^sub>S\<^sub>e\<^sub>t(x)->includes\<^sub>S\<^sub>e\<^sub>t(x))" proof - have C : "\\insert (x \) \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\\\ \ {X. X = bot \ X = null \ (\x\\\X\\. x \ bot)}" by(insert def_X val_x, frule Set_inv_lemma, simp add: foundation18 invalid_def) show ?thesis apply(subst OclIncludes_def, simp add: foundation10[simplified OclValid_def] OclValid_def def_X[simplified OclValid_def] val_x[simplified OclValid_def]) apply(simp add: OclIncluding_def def_X[simplified OclValid_def] val_x[simplified OclValid_def] Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse[OF C] true_def) done qed lemma OclIncludes_charn2: assumes def_X:"\ \ (\ X)" and val_x:"\ \ (\ x)" and val_y:"\ \ (\ y)" and neq :"\ \ not(x \ y)" shows "\ \ (X->including\<^sub>S\<^sub>e\<^sub>t(x)->includes\<^sub>S\<^sub>e\<^sub>t(y)) \ (X->includes\<^sub>S\<^sub>e\<^sub>t(y))" proof - have C : "\\insert (x \) \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\\\ \ {X. X = bot \ X = null \ (\x\\\X\\. x \ bot)}" by(insert def_X val_x, frule Set_inv_lemma, simp add: foundation18 invalid_def) show ?thesis apply(subst OclIncludes_def, simp add: def_X[simplified OclValid_def] val_x[simplified OclValid_def] val_y[simplified OclValid_def] foundation10[simplified OclValid_def] OclValid_def StrongEq_def) apply(simp add: OclIncluding_def OclIncludes_def def_X[simplified OclValid_def] val_x[simplified OclValid_def] val_y[simplified OclValid_def] Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse[OF C] true_def) by(metis foundation22 foundation6 foundation9 neq) qed text\Here is again a generic theorem similar as above.\ lemma OclIncludes_execute_generic: assumes strict1: "(invalid \ y) = invalid" and strict2: "(x \ invalid) = invalid" and cp_StrictRefEq: "\ (X::('\,'a::null)val) Y \. (X \ Y) \ = ((\_. X \) \ (\_. Y \)) \" and StrictRefEq_vs_StrongEq: "\ (x::('\,'a::null)val) y \. \ \ \ x \ \ \ \ y \ (\ \ ((x \ y) \ (x \ y)))" shows "(X->including\<^sub>S\<^sub>e\<^sub>t(x::('\,'a::null)val)->includes\<^sub>S\<^sub>e\<^sub>t(y)) = (if \ X then if x \ y then true else X->includes\<^sub>S\<^sub>e\<^sub>t(y) endif else invalid endif)" proof - have A: "\\. \ \ (X \ invalid) \ (X->including\<^sub>S\<^sub>e\<^sub>t(x)->includes\<^sub>S\<^sub>e\<^sub>t(y)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev,simp,simp) have B: "\\. \ \ (X \ null) \ (X->including\<^sub>S\<^sub>e\<^sub>t(x)->includes\<^sub>S\<^sub>e\<^sub>t(y)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev,simp,simp) note [simp] = cp_StrictRefEq [THEN allI[THEN allI[THEN allI[THEN cpI2]], of "StrictRefEq"]] have C: "\\. \ \ (x \ invalid) \ (X->including\<^sub>S\<^sub>e\<^sub>t(x)->includes\<^sub>S\<^sub>e\<^sub>t(y)) \ = (if x \ y then true else X->includes\<^sub>S\<^sub>e\<^sub>t(y) endif) \" apply(rule foundation22[THEN iffD1]) apply(erule StrongEq_L_subst2_rev,simp,simp) by (simp add: strict1) have D:"\\. \ \ (y \ invalid) \ (X->including\<^sub>S\<^sub>e\<^sub>t(x)->includes\<^sub>S\<^sub>e\<^sub>t(y)) \ = (if x \ y then true else X->includes\<^sub>S\<^sub>e\<^sub>t(y) endif) \" apply(rule foundation22[THEN iffD1]) apply(erule StrongEq_L_subst2_rev,simp,simp) by (simp add: strict2) have E: "\\. \ \ \ x \ \ \ \ y \ (if x \ y then true else X->includes\<^sub>S\<^sub>e\<^sub>t(y) endif) \ = (if x \ y then true else X->includes\<^sub>S\<^sub>e\<^sub>t(y) endif) \" apply(subst cp_OclIf) apply(subst StrictRefEq_vs_StrongEq[THEN foundation22[THEN iffD1]]) by(simp_all add: cp_OclIf[symmetric]) have F: "\\. \ \ (x \ y) \ (X->including\<^sub>S\<^sub>e\<^sub>t(x)->includes\<^sub>S\<^sub>e\<^sub>t(y)) \ = (X->including\<^sub>S\<^sub>e\<^sub>t(x)->includes\<^sub>S\<^sub>e\<^sub>t(x)) \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev,simp, simp) show ?thesis apply(rule ext, rename_tac "\") apply(case_tac "\ (\ \ (\ X))", simp add:defined_split,elim disjE A B) apply(case_tac "\ (\ \ (\ x))", simp add:foundation18 foundation22[symmetric], drule StrongEq_L_sym) apply(simp add: foundation22 C) apply(case_tac "\ (\ \ (\ y))", simp add:foundation18 foundation22[symmetric], drule StrongEq_L_sym, simp add: foundation22 D, simp) apply(subst E,simp_all) apply(case_tac "\ \ not(x \ y)") apply(simp add: OclIncludes_charn2[simplified foundation22]) apply(simp add: foundation9 F OclIncludes_charn1[THEN foundation13[THEN iffD2], THEN foundation22[THEN iffD1]]) done qed (* Hack to work around OF-Bug *) schematic_goal OclIncludes_execute\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r[simp,code_unfold]: "?X" by(rule OclIncludes_execute_generic[OF StrictRefEq\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r.strict1 StrictRefEq\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r.strict2 StrictRefEq\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r.cp0 StrictRefEq\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r.StrictRefEq_vs_StrongEq], simp_all) schematic_goal OclIncludes_execute\<^sub>B\<^sub>o\<^sub>o\<^sub>l\<^sub>e\<^sub>a\<^sub>n[simp,code_unfold]: "?X" by(rule OclIncludes_execute_generic[OF StrictRefEq\<^sub>B\<^sub>o\<^sub>o\<^sub>l\<^sub>e\<^sub>a\<^sub>n.strict1 StrictRefEq\<^sub>B\<^sub>o\<^sub>o\<^sub>l\<^sub>e\<^sub>a\<^sub>n.strict2 StrictRefEq\<^sub>B\<^sub>o\<^sub>o\<^sub>l\<^sub>e\<^sub>a\<^sub>n.cp0 StrictRefEq\<^sub>B\<^sub>o\<^sub>o\<^sub>l\<^sub>e\<^sub>a\<^sub>n.StrictRefEq_vs_StrongEq], simp_all) schematic_goal OclIncludes_execute\<^sub>S\<^sub>e\<^sub>t[simp,code_unfold]: "?X" by(rule OclIncludes_execute_generic[OF StrictRefEq\<^sub>S\<^sub>e\<^sub>t.strict1 StrictRefEq\<^sub>S\<^sub>e\<^sub>t.strict2 StrictRefEq\<^sub>S\<^sub>e\<^sub>t.cp0 StrictRefEq\<^sub>S\<^sub>e\<^sub>t.StrictRefEq_vs_StrongEq], simp_all) lemma OclIncludes_including_generic : assumes OclIncludes_execute_generic [simp] : "\X x y. (X->including\<^sub>S\<^sub>e\<^sub>t(x::('\,'a::null)val)->includes\<^sub>S\<^sub>e\<^sub>t(y)) = (if \ X then if x \ y then true else X->includes\<^sub>S\<^sub>e\<^sub>t(y) endif else invalid endif)" and StrictRefEq_strict'' : "\x y. \ ((x::('\,'a::null)val) \ y) = (\(x) and \(y))" and a_val : "\ \ \ a" and x_val : "\ \ \ x" and S_incl : "\ \ (S)->includes\<^sub>S\<^sub>e\<^sub>t((x::('\,'a::null)val))" shows "\ \ S->including\<^sub>S\<^sub>e\<^sub>t((a::('\,'a::null)val))->includes\<^sub>S\<^sub>e\<^sub>t(x)" proof - have discr_eq_bot1_true : "\\. (\ \ = true \) = False" by (metis bot_fun_def foundation1 foundation18' valid3) have discr_eq_bot2_true : "\\. (\ = true \) = False" by (metis bot_fun_def discr_eq_bot1_true) have discr_neq_invalid_true : "\\. (invalid \ \ true \) = True" by (metis discr_eq_bot2_true invalid_def) have discr_eq_invalid_true : "\\. (invalid \ = true \) = False" by (metis bot_option_def invalid_def option.simps(2) true_def) show ?thesis apply(simp) apply(subgoal_tac "\ \ \ S") prefer 2 apply(insert S_incl[simplified OclIncludes_def], simp add: OclValid_def) apply(metis discr_eq_bot2_true) apply(simp add: cp_OclIf[of "\ S"] OclValid_def OclIf_def x_val[simplified OclValid_def] discr_neq_invalid_true discr_eq_invalid_true) by (metis OclValid_def S_incl StrictRefEq_strict'' a_val foundation10 foundation6 x_val) qed lemmas OclIncludes_including\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r = OclIncludes_including_generic[OF OclIncludes_execute\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r StrictRefEq\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r.def_homo] subsubsection\Execution Rules on Excludes\ lemma OclExcludes_charn1: assumes def_X:"\ \ (\ X)" assumes val_x:"\ \ (\ x)" shows "\ \ (X->excluding\<^sub>S\<^sub>e\<^sub>t(x)->excludes\<^sub>S\<^sub>e\<^sub>t(x))" proof - let ?OclSet = "\S. \\S\\ \ {X. X = \ \ X = null \ (\x\\\X\\. x \ \)}" have diff_in_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e : "?OclSet (\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\ - {x \})" apply(simp, (rule disjI2)+) by (metis (opaque_lifting, no_types) Diff_iff Set_inv_lemma def_X) show ?thesis apply(subst OclExcludes_def, simp add: foundation10[simplified OclValid_def] OclValid_def def_X[simplified OclValid_def] val_x[simplified OclValid_def]) apply(subst OclIncludes_def, simp add: OclNot_def) apply(simp add: OclExcluding_def def_X[simplified OclValid_def] val_x[simplified OclValid_def] Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse[OF diff_in_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e] true_def) by(simp add: OclAnd_def def_X[simplified OclValid_def] val_x[simplified OclValid_def] true_def) qed subsubsection\Execution Rules on Size\ lemma [simp,code_unfold]: "Set{} ->size\<^sub>S\<^sub>e\<^sub>t() = \" apply(rule ext) apply(simp add: defined_def mtSet_def OclSize_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_fun_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_fun_def) apply(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject, simp_all add: bot_option_def null_option_def) + by(simp add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse bot_option_def null_option_def OclInt0_def) lemma OclSize_including_exec[simp,code_unfold]: "((X ->including\<^sub>S\<^sub>e\<^sub>t(x)) ->size\<^sub>S\<^sub>e\<^sub>t()) = (if \ X and \ x then X ->size\<^sub>S\<^sub>e\<^sub>t() +\<^sub>i\<^sub>n\<^sub>t if X ->includes\<^sub>S\<^sub>e\<^sub>t(x) then \ else \ endif else invalid endif)" proof - have valid_inject_true : "\\ P. (\ P) \ \ true \ \ (\ P) \ = false \" apply(simp add: valid_def true_def false_def bot_fun_def bot_option_def null_fun_def null_option_def) by (case_tac "P \ = \", simp_all add: true_def) have defined_inject_true : "\\ P. (\ P) \ \ true \ \ (\ P) \ = false \" apply(simp add: defined_def true_def false_def bot_fun_def bot_option_def null_fun_def null_option_def) by (case_tac " P \ = \ \ P \ = null", simp_all add: true_def) show ?thesis apply(rule ext, rename_tac \) proof - fix \ have includes_notin: "\ \ \ X->includes\<^sub>S\<^sub>e\<^sub>t(x) \ (\ X) \ = true \ \ (\ x) \ = true \ \ x \ \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\" by(simp add: OclIncludes_def OclValid_def true_def) have includes_def: "\ \ X->includes\<^sub>S\<^sub>e\<^sub>t(x) \ \ \ \ X" by (metis bot_fun_def OclIncludes_def OclValid_def defined3 foundation16) have includes_val: "\ \ X->includes\<^sub>S\<^sub>e\<^sub>t(x) \ \ \ \ x" using foundation5 foundation6 by fastforce have ins_in_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e: "\ \ \ X \ \ \ \ x \ \\insert (x \) \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\\\ \ {X. X = \ \ X = null \ (\x\\\X\\. x \ \)}" apply(simp add: bot_option_def null_option_def) by (metis (opaque_lifting, no_types) Set_inv_lemma foundation18' foundation5) have m : "\\. (\_. \) = (\_. invalid \)" by(rule ext, simp add:invalid_def) show "X->including\<^sub>S\<^sub>e\<^sub>t(x)->size\<^sub>S\<^sub>e\<^sub>t() \ = (if \ X and \ x then X->size\<^sub>S\<^sub>e\<^sub>t() +\<^sub>i\<^sub>n\<^sub>t if X->includes\<^sub>S\<^sub>e\<^sub>t(x) then \ else \ endif else invalid endif) \" apply(case_tac "\ \ \ X and \ x", simp) apply(subst OclAdd\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r.cp0) apply(case_tac "\ \ X->includes\<^sub>S\<^sub>e\<^sub>t(x)", simp add: OclAdd\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r.cp0[symmetric]) apply(case_tac "\ \ ((\ (X->size\<^sub>S\<^sub>e\<^sub>t())) and not (\ (X->size\<^sub>S\<^sub>e\<^sub>t())))", simp) apply(drule foundation5[where P = "\ X->size\<^sub>S\<^sub>e\<^sub>t()"], erule conjE) apply(drule OclSize_infinite) apply(frule includes_def, drule includes_val, simp) apply(subst OclSize_def, subst OclIncluding_finite_rep_set, assumption+) apply (metis (opaque_lifting, no_types) invalid_def) apply(subst OclIf_false', metis (opaque_lifting, no_types) defined5 defined6 defined_and_I defined_not_I foundation1 foundation9) apply(subst cp_OclSize, simp add: OclIncluding_includes0 cp_OclSize[symmetric]) (* *) apply(subst OclIf_false', subst foundation9, auto, simp add: OclSize_def) apply(drule foundation5) apply(subst (1 2) OclIncluding_finite_rep_set, fast+) apply(subst (1 2) cp_OclAnd, subst (1 2) OclAdd\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r.cp0, simp) apply(rule conjI) apply(simp add: OclIncluding_def) apply(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse[OF ins_in_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e], fast+) apply(subst (asm) (2 3) OclValid_def, simp add: OclAdd\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r_def OclInt1_def) apply(rule impI) apply(drule Finite_Set.card.insert[where x = "x \"]) apply(rule includes_notin, simp, simp) apply (metis Suc_eq_plus1 of_nat_1 of_nat_add) apply(subst (1 2) m[of \], simp only: OclAdd\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r.cp0[symmetric],simp, simp add:invalid_def) apply(subst OclIncluding_finite_rep_set, fast+, simp add: OclValid_def) (* *) apply(subst OclIf_false', metis (opaque_lifting, no_types) defined6 foundation1 foundation9 OclExcluding_valid_args_valid'') by (metis cp_OclSize foundation18' OclIncluding_valid_args_valid'' invalid_def OclSize_invalid) qed qed subsubsection\Execution Rules on IsEmpty\ lemma [simp,code_unfold]: "Set{}->isEmpty\<^sub>S\<^sub>e\<^sub>t() = true" by(simp add: OclIsEmpty_def) lemma OclIsEmpty_including [simp]: assumes X_def: "\ \ \ X" and X_finite: "finite \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\" and a_val: "\ \ \ a" shows "X->including\<^sub>S\<^sub>e\<^sub>t(a)->isEmpty\<^sub>S\<^sub>e\<^sub>t() \ = false \" proof - have A1 : "\\ X. X \ = true \ \ X \ = false \ \ (X and not X) \ = false \" by (metis (no_types) OclAnd_false1 OclAnd_idem OclImplies_def OclNot3 OclNot_not OclOr_false1 cp_OclAnd cp_OclNot deMorgan1 deMorgan2) have defined_inject_true : "\\ P. (\ P) \ \ true \ \ (\ P) \ = false \" apply(simp add: defined_def true_def false_def bot_fun_def bot_option_def null_fun_def null_option_def) by (case_tac " P \ = \ \ P \ = null", simp_all add: true_def) have B : "\X \. \ \ \ X \ X \ \ \ \ \ (X \ \) \ = false \" apply(simp add: foundation22[symmetric] foundation14 foundation9) apply(erule StrongEq_L_subst4_rev[THEN iffD2, OF StrictRefEq\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r.StrictRefEq_vs_StrongEq]) by(simp_all) show ?thesis apply(simp add: OclIsEmpty_def del: OclSize_including_exec) apply(subst cp_OclOr, subst A1) apply (metis OclExcludes.def_homo defined_inject_true) apply(simp add: cp_OclOr[symmetric] del: OclSize_including_exec) apply(rule B, rule foundation20, metis OclIncluding.def_homo OclIncluding_finite_rep_set X_def X_finite a_val foundation10' size_defined') apply(simp add: OclSize_def OclIncluding_finite_rep_set[OF X_def a_val] X_finite OclInt0_def) by (metis OclValid_def X_def a_val foundation10 foundation6 OclIncluding_notempty_rep_set[OF X_def a_val]) qed subsubsection\Execution Rules on NotEmpty\ lemma [simp,code_unfold]: "Set{}->notEmpty\<^sub>S\<^sub>e\<^sub>t() = false" by(simp add: OclNotEmpty_def) lemma OclNotEmpty_including [simp,code_unfold]: assumes X_def: "\ \ \ X" and X_finite: "finite \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\" and a_val: "\ \ \ a" shows "X->including\<^sub>S\<^sub>e\<^sub>t(a)->notEmpty\<^sub>S\<^sub>e\<^sub>t() \ = true \" apply(simp add: OclNotEmpty_def) apply(subst cp_OclNot, subst OclIsEmpty_including, simp_all add: assms) by (metis OclNot4 cp_OclNot) subsubsection\Execution Rules on Any\ lemma [simp,code_unfold]: "Set{}->any\<^sub>S\<^sub>e\<^sub>t() = null" by(rule ext, simp add: OclANY_def, simp add: false_def true_def) lemma OclANY_singleton_exec[simp,code_unfold]: "(Set{}->including\<^sub>S\<^sub>e\<^sub>t(a))->any\<^sub>S\<^sub>e\<^sub>t() = a" apply(rule ext, rename_tac \, simp add: mtSet_def OclANY_def) apply(case_tac "\ \ \ a") apply(simp add: OclValid_def mtSet_defined[simplified mtSet_def] mtSet_valid[simplified mtSet_def] mtSet_rep_set[simplified mtSet_def]) apply(subst (1 2) cp_OclAnd, subst (1 2) OclNotEmpty_including[where X = "Set{}", simplified mtSet_def]) apply(simp add: mtSet_defined[simplified mtSet_def]) apply(metis (opaque_lifting, no_types) finite.emptyI mtSet_def mtSet_rep_set) apply(simp add: OclValid_def) apply(simp add: OclIncluding_def) apply(rule conjI) apply(subst (1 2) Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse, simp add: bot_option_def null_option_def) apply(simp, metis OclValid_def foundation18') apply(simp) apply(simp add: mtSet_defined[simplified mtSet_def]) (* *) apply(subgoal_tac "a \ = \") prefer 2 apply(simp add: OclValid_def valid_def bot_fun_def split: if_split_asm) apply(simp) apply(subst (1 2 3 4) cp_OclAnd, simp add: mtSet_defined[simplified mtSet_def] valid_def bot_fun_def) by(simp add: cp_OclAnd[symmetric], rule impI, simp add: false_def true_def) subsubsection\Execution Rules on Forall\ lemma OclForall_mtSet_exec[simp,code_unfold] :"((Set{})->forAll\<^sub>S\<^sub>e\<^sub>t(z| P(z))) = true" apply(simp add: OclForall_def) apply(subst mtSet_def)+ apply(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse, simp_all add: true_def)+ done text\The following rule is a main theorem of our approach: From a denotational definition that assures consistency, but may be --- as in the case of the @{term "X->forAll\<^sub>S\<^sub>e\<^sub>t(x | P x)"} --- dauntingly complex, we derive operational rules that can serve as a gold-standard for operational execution, since they may be evaluated in whatever situation and according to whatever strategy. In the case of @{term "X->forAll\<^sub>S\<^sub>e\<^sub>t(x | P x)"}, the operational rule gives immediately a way to evaluation in any finite (in terms of conventional OCL: denotable) set, although the rule also holds for the infinite case: @{term "Integer\<^sub>n\<^sub>u\<^sub>l\<^sub>l ->forAll\<^sub>S\<^sub>e\<^sub>t(x | (Integer\<^sub>n\<^sub>u\<^sub>l\<^sub>l ->forAll\<^sub>S\<^sub>e\<^sub>t(y | x +\<^sub>i\<^sub>n\<^sub>t y \ y +\<^sub>i\<^sub>n\<^sub>t x)))"} or even: @{term "Integer ->forAll\<^sub>S\<^sub>e\<^sub>t(x | (Integer ->forAll\<^sub>S\<^sub>e\<^sub>t(y | x +\<^sub>i\<^sub>n\<^sub>t y \ y +\<^sub>i\<^sub>n\<^sub>t x)))"} are valid OCL statements in any context $\tau$. \ theorem OclForall_including_exec[simp,code_unfold] : assumes cp0 : "cp P" shows "((S->including\<^sub>S\<^sub>e\<^sub>t(x))->forAll\<^sub>S\<^sub>e\<^sub>t(z | P(z))) = (if \ S and \ x then P x and (S->forAll\<^sub>S\<^sub>e\<^sub>t(z | P(z))) else invalid endif)" proof - have cp: "\\. P x \ = P (\_. x \) \" by(insert cp0, auto simp: cp_def) have cp_eq : "\\ v. (P x \ = v) = (P (\_. x \) \ = v)" by(subst cp, simp) have cp_OclNot_eq : "\\ v. (P x \ \ v) = (P (\_. x \) \ \ v)" by(subst cp, simp) have insert_in_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e : "\\. (\ \(\ S)) \ (\ \(\ x)) \ \\insert (x \) \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\\\ \ {X. X = bot \ X = null \ (\x\\\X\\. x \ bot)}" by(frule Set_inv_lemma, simp add: foundation18 invalid_def) have forall_including_invert : "\\ f. (f x \ = f (\ _. x \) \) \ \ \ (\ S and \ x) \ (\x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S->including\<^sub>S\<^sub>e\<^sub>t(x) \)\\. f (\_. x) \) = (f x \ \ (\x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\. f (\_. x) \))" apply(drule foundation5, simp add: OclIncluding_def) apply(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse) apply(rule insert_in_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e, fast+) by(simp add: OclValid_def) have exists_including_invert : "\\ f. (f x \ = f (\ _. x \) \) \ \ \ (\ S and \ x) \ (\x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S->including\<^sub>S\<^sub>e\<^sub>t(x) \)\\. f (\_. x) \) = (f x \ \ (\x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\. f (\_. x) \))" apply(subst arg_cong[where f = "\x. \x", OF forall_including_invert[where f = "\x \. \ (f x \)"], simplified]) by simp_all have contradict_Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e: "\\ S f. \x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e S\\. f (\_. x) \ \ (\x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e S\\. \ (f (\_. x) \)) = False" by(case_tac "(\x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e S\\. \ (f (\_. x) \)) = True", simp_all) have bot_invalid : "\ = invalid" by(rule ext, simp add: invalid_def bot_fun_def) have bot_invalid2 : "\\. \ = invalid \" by(simp add: invalid_def) have C1 : "\\. P x \ = false \ \ (\x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\. P (\_. x) \ = false \) \ \ \ (\ S and \ x) \ false \ = (P x and OclForall S P) \" apply(simp add: cp_OclAnd[of "P x"]) apply(elim disjE, simp) apply(simp only: cp_OclAnd[symmetric], simp) apply(subgoal_tac "OclForall S P \ = false \") apply(simp only: cp_OclAnd[symmetric], simp) apply(simp add: OclForall_def) apply(fold OclValid_def, simp add: foundation10') done have C2 : "\\. \ \ (\ S and \ x) \ P x \ = null \ \ (\x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\. P (\_. x) \ = null \) \ P x \ = invalid \ \ (\x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\. P (\_. x) \ = invalid \) \ \x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S->including\<^sub>S\<^sub>e\<^sub>t(x) \)\\. P (\_. x) \ \ false \ \ invalid \ = (P x and OclForall S P) \" apply(subgoal_tac "(\ S)\ = true \") prefer 2 apply(simp add: foundation10', simp add: OclValid_def) apply(drule forall_including_invert[of "\ x \. P x \ \ false \", OF cp_OclNot_eq, THEN iffD1]) apply(assumption) apply(simp add: cp_OclAnd[of "P x"],elim disjE, simp_all) apply(simp add: invalid_def null_fun_def null_option_def bot_fun_def bot_option_def) apply(subgoal_tac "OclForall S P \ = invalid \") apply(simp only:cp_OclAnd[symmetric],simp,simp add:invalid_def bot_fun_def) apply(unfold OclForall_def, simp add: invalid_def false_def bot_fun_def,simp) apply(simp add:cp_OclAnd[symmetric],simp) apply(erule conjE) apply(subgoal_tac "(P x \ = invalid \) \ (P x \ = null \) \ (P x \ = true \) \ (P x \ = false \)") prefer 2 apply(rule bool_split_0) apply(elim disjE, simp_all) apply(simp only:cp_OclAnd[symmetric],simp)+ done have A : "\\. \ \ (\ S and \ x) \ OclForall (S->including\<^sub>S\<^sub>e\<^sub>t(x)) P \ = (P x and OclForall S P) \" proof - fix \ assume 0 : "\ \ (\ S and \ x)" let ?S = "\ocl. P x \ \ ocl \ \ (\x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\. P (\_. x) \ \ ocl \)" let ?S' = "\ocl. \x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S->including\<^sub>S\<^sub>e\<^sub>t(x) \)\\. P (\_. x) \ \ ocl \" let ?assms_1 = "?S' null" let ?assms_2 = "?S' invalid" let ?assms_3 = "?S' false" have 4 : "?assms_3 \ ?S false" apply(subst forall_including_invert[of "\ x \. P x \ \ false \",symmetric]) by(simp_all add: cp_OclNot_eq 0) have 5 : "?assms_2 \ ?S invalid" apply(subst forall_including_invert[of "\ x \. P x \ \ invalid \",symmetric]) by(simp_all add: cp_OclNot_eq 0) have 6 : "?assms_1 \ ?S null" apply(subst forall_including_invert[of "\ x \. P x \ \ null \",symmetric]) by(simp_all add: cp_OclNot_eq 0) have 7 : "(\ S) \ = true \" by(insert 0, simp add: foundation10', simp add: OclValid_def) show "?thesis \" apply(subst OclForall_def) apply(simp add: cp_OclAnd[THEN sym] OclValid_def contradict_Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e) apply(intro conjI impI,fold OclValid_def) apply(simp_all add: exists_including_invert[where f = "\ x \. P x \ = null \", OF cp_eq]) apply(simp_all add: exists_including_invert[where f = "\ x \. P x \ = invalid \", OF cp_eq]) apply(simp_all add: exists_including_invert[where f = "\ x \. P x \ = false \", OF cp_eq]) proof - assume 1 : "P x \ = null \ \ (\x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\. P (\_. x) \ = null \)" and 2 : ?assms_2 and 3 : ?assms_3 show "null \ = (P x and OclForall S P) \" proof - note 4 = 4[OF 3] note 5 = 5[OF 2] have 6 : "P x \ = null \ \ P x \ = true \" by(metis 4 5 bool_split_0) show ?thesis apply(insert 6, elim disjE) apply(subst cp_OclAnd) apply(simp add: OclForall_def 7 4[THEN conjunct2] 5[THEN conjunct2]) apply(simp_all add:cp_OclAnd[symmetric]) apply(subst cp_OclAnd, simp_all add:cp_OclAnd[symmetric] OclForall_def) apply(simp add:4[THEN conjunct2] 5[THEN conjunct2] 0[simplified OclValid_def] 7) apply(insert 1, elim disjE, auto) done qed next assume 1 : ?assms_1 and 2 : "P x \ = invalid \ \ (\x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\. P (\_. x) \ = invalid \)" and 3 : ?assms_3 show "invalid \ = (P x and OclForall S P) \" proof - note 4 = 4[OF 3] note 6 = 6[OF 1] have 5 : "P x \ = invalid \ \ P x \ = true \" by(metis 4 6 bool_split_0) show ?thesis apply(insert 5, elim disjE) apply(subst cp_OclAnd) apply(simp add: OclForall_def 4[THEN conjunct2] 6[THEN conjunct2] 7) apply(simp_all add:cp_OclAnd[symmetric]) apply(subst cp_OclAnd, simp_all add:cp_OclAnd[symmetric] OclForall_def) apply(insert 2, elim disjE, simp add: invalid_def true_def bot_option_def) apply(simp add: 0[simplified OclValid_def] 4[THEN conjunct2] 6[THEN conjunct2] 7) by(auto) qed next assume 1 : ?assms_1 and 2 : ?assms_2 and 3 : ?assms_3 show "true \ = (P x and OclForall S P) \" proof - note 4 = 4[OF 3] note 5 = 5[OF 2] note 6 = 6[OF 1] have 8 : "P x \ = true \" by(metis 4 5 6 bool_split_0) show ?thesis apply(subst cp_OclAnd, simp add: 8 cp_OclAnd[symmetric]) by(simp add: OclForall_def 4 5 6 7) qed qed ( simp add: 0 | rule C1, simp+ | rule C2, simp add: 0 )+ qed have B : "\\. \ (\ \ (\ S and \ x)) \ OclForall (S->including\<^sub>S\<^sub>e\<^sub>t(x)) P \ = invalid \" apply(rule foundation22[THEN iffD1]) apply(simp only: foundation10' de_Morgan_conj foundation18'', elim disjE) apply(simp add: defined_split, elim disjE) apply(erule StrongEq_L_subst2_rev, simp+)+ done show ?thesis apply(rule ext, rename_tac \) apply(simp add: OclIf_def) apply(simp add: cp_defined[of "\ S and \ x"] cp_defined[THEN sym]) apply(intro conjI impI) by(auto intro!: A B simp: OclValid_def) qed subsubsection\Execution Rules on Exists\ lemma OclExists_mtSet_exec[simp,code_unfold] : "((Set{})->exists\<^sub>S\<^sub>e\<^sub>t(z | P(z))) = false" by(simp add: OclExists_def) lemma OclExists_including_exec[simp,code_unfold] : assumes cp: "cp P" shows "((S->including\<^sub>S\<^sub>e\<^sub>t(x))->exists\<^sub>S\<^sub>e\<^sub>t(z | P(z))) = (if \ S and \ x then P x or (S->exists\<^sub>S\<^sub>e\<^sub>t(z | P(z))) else invalid endif)" by(simp add: OclExists_def OclOr_def cp OclNot_inject) subsubsection\Execution Rules on Iterate\ lemma OclIterate_empty[simp,code_unfold]: "((Set{})->iterate\<^sub>S\<^sub>e\<^sub>t(a; x = A | P a x)) = A" proof - have C : "\ \. (\ (\\. Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\{}\\)) \ = true \" by (metis (no_types) defined_def mtSet_def mtSet_defined null_fun_def) show ?thesis apply(simp add: OclIterate_def mtSet_def Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse valid_def C) apply(rule ext, rename_tac \) apply(case_tac "A \ = \ \", simp_all, simp add:true_def false_def bot_fun_def) apply(simp add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse) done qed text\In particular, this does hold for A = null.\ lemma OclIterate_including: assumes S_finite: "\ \ \(S->size\<^sub>S\<^sub>e\<^sub>t())" and F_valid_arg: "(\ A) \ = (\ (F a A)) \" and F_commute: "comp_fun_commute F" and F_cp: "\ x y \. F x y \ = F (\ _. x \) y \" shows "((S->including\<^sub>S\<^sub>e\<^sub>t(a))->iterate\<^sub>S\<^sub>e\<^sub>t(a; x = A | F a x)) \ = ((S->excluding\<^sub>S\<^sub>e\<^sub>t(a))->iterate\<^sub>S\<^sub>e\<^sub>t(a; x = F a A | F a x)) \" proof - have insert_in_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e : "\\. (\ \(\ S)) \ (\ \(\ a)) \ \\insert (a \) \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\\\ \ {X. X = bot \ X = null \ (\x\\\X\\. x \ bot)}" by(frule Set_inv_lemma, simp add: foundation18 invalid_def) have insert_defined : "\\. (\ \(\ S)) \ (\ \(\ a)) \ (\ (\_. Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\insert (a \) \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\\\)) \ = true \" apply(subst defined_def) apply(simp add: bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_fun_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_fun_def) by(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject, rule insert_in_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e, simp_all add: null_option_def bot_option_def)+ have remove_finite : "finite \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\ \ finite ((\a \. a) ` (\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\ - {a \}))" by(simp) have remove_in_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e : "\\. (\ \(\ S)) \ (\ \(\ a)) \ \\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\ - {a \}\\ \ {X. X = bot \ X = null \ (\x\\\X\\. x \ bot)}" by(frule Set_inv_lemma, simp add: foundation18 invalid_def) have remove_defined : "\\. (\ \(\ S)) \ (\ \(\ a)) \ (\ (\_. Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\ - {a \}\\)) \ = true \" apply(subst defined_def) apply(simp add: bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_fun_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_fun_def) by(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject, rule remove_in_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e, simp_all add: null_option_def bot_option_def)+ have abs_rep: "\x. \\x\\ \ {X. X = bot \ X = null \ (\x\\\X\\. x \ bot)} \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\x\\)\\ = x" by(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse, simp_all) have inject : "inj (\a \. a)" by(rule inj_fun, simp) interpret F_commute: comp_fun_commute "F" by (fact F_commute) show ?thesis apply(subst (1 2) cp_OclIterate, subst OclIncluding_def, subst OclExcluding_def) apply(case_tac "\ ((\ S) \ = true \ \ (\ a) \ = true \)", simp add: invalid_def) apply(subgoal_tac "OclIterate (\_. \) A F \ = OclIterate (\_. \) (F a A) F \", simp) apply(rule conjI, blast+) apply(simp add: OclIterate_def defined_def bot_option_def bot_fun_def false_def true_def) apply(simp add: OclIterate_def) apply((subst abs_rep[OF insert_in_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e[simplified OclValid_def], of \], simp_all)+, (subst abs_rep[OF remove_in_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e[simplified OclValid_def], of \], simp_all)+, (subst insert_defined, simp_all add: OclValid_def)+, (subst remove_defined, simp_all add: OclValid_def)+) apply(case_tac "\ ((\ A) \ = true \)", (simp add: F_valid_arg)+) apply(rule impI, subst F_commute.fold_fun_left_comm[symmetric], rule remove_finite, simp) apply(subst image_set_diff[OF inject], simp) apply(subgoal_tac "Finite_Set.fold F A (insert (\\'. a \) ((\a \. a) ` \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\)) \ = F (\\'. a \) (Finite_Set.fold F A ((\a \. a) ` \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\ - {\\'. a \})) \") apply(subst F_cp, simp) by(subst F_commute.fold_insert_remove, simp+) qed subsubsection\Execution Rules on Select\ lemma OclSelect_mtSet_exec[simp,code_unfold]: "OclSelect mtSet P = mtSet" apply(rule ext, rename_tac \) apply(simp add: OclSelect_def mtSet_def defined_def false_def true_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_fun_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_fun_def) by(( subst (1 2 3 4 5) Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse | subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject), (simp add: null_option_def bot_option_def)+)+ definition "OclSelect_body :: _ \ _ \ _ \ ('\, 'a option option) Set \ (\P x acc. if P x \ false then acc else acc->including\<^sub>S\<^sub>e\<^sub>t(x) endif)" theorem OclSelect_including_exec[simp,code_unfold]: assumes P_cp : "cp P" shows "OclSelect (X->including\<^sub>S\<^sub>e\<^sub>t(y)) P = OclSelect_body P y (OclSelect (X->excluding\<^sub>S\<^sub>e\<^sub>t(y)) P)" (is "_ = ?select") proof - have P_cp: "\x \. P x \ = P (\_. x \) \" by(insert P_cp, auto simp: cp_def) have ex_including : "\f X y \. \ \ \ X \ \ \ \ y \ (\x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X->including\<^sub>S\<^sub>e\<^sub>t(y) \)\\. f (P (\_. x)) \) = (f (P (\_. y \)) \ \ (\x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\. f (P (\_. x)) \))" apply(simp add: OclIncluding_def OclValid_def) apply(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse, simp, (rule disjI2)+) by (metis (opaque_lifting, no_types) OclValid_def Set_inv_lemma foundation18',simp) have al_including : "\f X y \. \ \ \ X \ \ \ \ y \ (\x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X->including\<^sub>S\<^sub>e\<^sub>t(y) \)\\. f (P (\_. x)) \) = (f (P (\_. y \)) \ \ (\x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\. f (P (\_. x)) \))" apply(simp add: OclIncluding_def OclValid_def) apply(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse, simp, (rule disjI2)+) by (metis (opaque_lifting, no_types) OclValid_def Set_inv_lemma foundation18', simp) have ex_excluding1 : "\f X y \. \ \ \ X \ \ \ \ y \ \ (f (P (\_. y \)) \) \ (\x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\. f (P (\_. x)) \) = (\x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X->excluding\<^sub>S\<^sub>e\<^sub>t(y) \)\\. f (P (\_. x)) \)" apply(simp add: OclExcluding_def OclValid_def) apply(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse, simp, (rule disjI2)+) by (metis (no_types) Diff_iff OclValid_def Set_inv_lemma) auto have al_excluding1 : "\f X y \. \ \ \ X \ \ \ \ y \ f (P (\_. y \)) \ \ (\x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\. f (P (\_. x)) \) = (\x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X->excluding\<^sub>S\<^sub>e\<^sub>t(y) \)\\. f (P (\_. x)) \)" apply(simp add: OclExcluding_def OclValid_def) apply(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse, simp, (rule disjI2)+) by (metis (no_types) Diff_iff OclValid_def Set_inv_lemma) auto have in_including : "\f X y \. \ \ \ X \ \ \ \ y \ {x \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X->including\<^sub>S\<^sub>e\<^sub>t(y) \)\\. f (P (\_. x) \)} = (let s = {x \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\. f (P (\_. x) \)} in if f (P (\_. y \) \) then insert (y \) s else s)" apply(simp add: OclIncluding_def OclValid_def) apply(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse, simp, (rule disjI2)+) apply (metis (opaque_lifting, no_types) OclValid_def Set_inv_lemma foundation18') by(simp add: Let_def, auto) let ?OclSet = "\S. \\S\\ \ {X. X = \ \ X = null \ (\x\\\X\\. x \ \)}" have diff_in_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e : "\\. (\ X) \ = true \ \ ?OclSet (\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\ - {y \})" apply(simp, (rule disjI2)+) by (metis (mono_tags) Diff_iff OclValid_def Set_inv_lemma) have ins_in_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e : "\\. (\ X) \ = true \ \ (\ y) \ = true \ \ ?OclSet (insert (y \) {x \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\. P (\_. x) \ \ false \})" apply(simp, (rule disjI2)+) by (metis (opaque_lifting, no_types) OclValid_def Set_inv_lemma foundation18') have ins_in_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e' : "\\. (\ X) \ = true \ \ (\ y) \ = true \ \ ?OclSet (insert (y \) {x \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\. x \ y \ \ P (\_. x) \ \ false \})" apply(simp, (rule disjI2)+) by (metis (opaque_lifting, no_types) OclValid_def Set_inv_lemma foundation18') have ins_in_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e'' : "\\. (\ X) \ = true \ \ ?OclSet {x \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\. P (\_. x) \ \ false \}" apply(simp, (rule disjI2)+) by (metis (opaque_lifting, no_types) OclValid_def Set_inv_lemma) have ins_in_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e''' : "\\. (\ X) \ = true \ \ ?OclSet {x \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\. x \ y \ \ P (\_. x) \ \ false \}" apply(simp, (rule disjI2)+) by(metis (opaque_lifting, no_types) OclValid_def Set_inv_lemma) have if_same : "\a b c d \. \ \ \ a \ b \ = d \ \ c \ = d \ \ (if a then b else c endif) \ = d \" by(simp add: OclIf_def OclValid_def) have invert_including : "\P y \. P \ = \ \ P->including\<^sub>S\<^sub>e\<^sub>t(y) \ = \" by (metis (opaque_lifting, no_types) foundation16[THEN iffD1] foundation18' OclIncluding_valid_args_valid) have exclude_defined : "\\. \ \ \ X \ (\(\_. Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\{x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\. x \ y \ \ P (\_. x) \\false \}\\)) \ = true \" apply(subst defined_def, simp add: false_def true_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_fun_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_fun_def) by(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject[OF ins_in_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e'''[simplified false_def]], (simp add: OclValid_def bot_option_def null_option_def)+)+ have if_eq : "\x A B \. \ \ \ x \ \ \ ((if x \ false then A else B endif) \ (if x \ false then A else B endif))" apply(simp add: StrictRefEq\<^sub>B\<^sub>o\<^sub>o\<^sub>l\<^sub>e\<^sub>a\<^sub>n OclValid_def) apply(subst (2) StrongEq_def) by(subst cp_OclIf, simp add: cp_OclIf[symmetric] true_def) have OclSelect_body_bot: "\\. \ \ \ X \ \ \ \ y \ P y \ \ \ \ (\x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\. P (\_. x) \ = \) \ \ = ?select \" - apply(drule ex_excluding1[where X = X and y = y and f = "\x \. x \ = \"], + apply(drule ex_excluding1[where X2 = X and y2 = y and f2 = "\x \. x \ = \"], (simp add: P_cp[symmetric])+) apply(subgoal_tac "\ \ (\ \ ?select)", simp add: OclValid_def StrongEq_def true_def bot_fun_def) apply(simp add: OclSelect_body_def) apply(subst StrongEq_L_subst3[OF _ if_eq], simp, metis foundation18') apply(simp add: OclValid_def, subst StrongEq_def, subst true_def, simp) apply(subgoal_tac "\x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X->excluding\<^sub>S\<^sub>e\<^sub>t(y) \)\\. P (\_. x) \ = \ \") prefer 2 apply (metis bot_fun_def ) - apply(subst if_same[where d = "\"]) + apply(subst if_same[where d5 = "\"]) apply (metis defined7 transform1) apply(simp add: OclSelect_def bot_option_def bot_fun_def invalid_def) apply(subst invert_including) by(simp add: OclSelect_def bot_option_def bot_fun_def invalid_def)+ have d_and_v_inject : "\\ X y. (\ X and \ y) \ \ true \ \ (\ X and \ y) \ = false \" apply(fold OclValid_def, subst foundation22[symmetric]) apply(auto simp:foundation10' defined_split) apply(erule StrongEq_L_subst2_rev,simp,simp) apply(erule StrongEq_L_subst2_rev,simp,simp) by(erule foundation7'[THEN iffD2, THEN foundation15[THEN iffD2, THEN StrongEq_L_subst2_rev]],simp,simp) have OclSelect_body_bot': "\\. (\ X and \ y) \ \ true \ \ \ = ?select \" apply(drule d_and_v_inject) apply(simp add: OclSelect_def OclSelect_body_def) apply(subst cp_OclIf, subst OclIncluding.cp0, simp add: false_def true_def) apply(subst cp_OclIf[symmetric], subst OclIncluding.cp0[symmetric]) by (metis (lifting, no_types) OclIf_def foundation18 foundation18' invert_including) have conj_split2 : "\a b c \. ((a \ false) \ = false \ \ b) \ ((a \ false) \ = true \ \ c) \ (a \ \ false \ \ b) \ (a \ = false \ \ c)" by (metis OclValid_def defined7 foundation14 foundation22 foundation9) have defined_inject_true : "\\ P. (\ P) \ \ true \ \ (\ P) \ = false \" apply(simp add: defined_def true_def false_def bot_fun_def bot_option_def null_fun_def null_option_def) by (case_tac " P \ = \ \ P \ = null", simp_all add: true_def) have cp_OclSelect_body : "\\. ?select \ = OclSelect_body P y (\_.(OclSelect (X->excluding\<^sub>S\<^sub>e\<^sub>t(y))P)\)\" apply(simp add: OclSelect_body_def) by(subst (1 2) cp_OclIf, subst (1 2) OclIncluding.cp0, blast) have OclSelect_body_strict1 : "OclSelect_body P y invalid = invalid" by(rule ext, simp add: OclSelect_body_def OclIf_def) have bool_invalid: "\(x::('\)Boolean) y \. \ (\ \ \ x) \ \ \ ((x \ y) \ invalid)" by(simp add: StrictRefEq\<^sub>B\<^sub>o\<^sub>o\<^sub>l\<^sub>e\<^sub>a\<^sub>n OclValid_def StrongEq_def true_def) have conj_comm : "\p q r. (p \ q \ r) = ((p \ q) \ r)" by blast have inv_bot : "\\. invalid \ = \ \" by (metis bot_fun_def invalid_def) have inv_bot' : "\\. invalid \ = \" by (simp add: invalid_def) show ?thesis apply(rule ext, rename_tac \) apply(subst OclSelect_def) apply(case_tac "(\ (X->including\<^sub>S\<^sub>e\<^sub>t(y))) \ = true \", simp) apply(( subst ex_including | subst in_including), metis OclValid_def foundation5, metis OclValid_def foundation5)+ apply(simp add: Let_def inv_bot) apply(subst (2 4 7 9) bot_fun_def) apply(subst (4) false_def, subst (4) bot_fun_def, simp add: bot_option_def P_cp[symmetric]) (* *) apply(case_tac "\ (\ \ (\ P y))") apply(subgoal_tac "P y \ \ false \") prefer 2 apply (metis (opaque_lifting, no_types) foundation1 foundation18' valid4) apply(simp) (* *) apply(subst conj_comm, rule conjI) - apply(drule_tac y = false in bool_invalid) + apply(drule_tac y11 = false in bool_invalid) apply(simp only: OclSelect_body_def, metis OclIf_def OclValid_def defined_def foundation2 foundation22 bot_fun_def invalid_def) (* *) apply(drule foundation5[simplified OclValid_def], subst al_including[simplified OclValid_def], simp, simp) apply(simp add: P_cp[symmetric]) apply (metis bot_fun_def foundation18') apply(simp add: foundation18' bot_fun_def OclSelect_body_bot OclSelect_body_bot') (* *) apply(subst (1 2) al_including, metis OclValid_def foundation5, metis OclValid_def foundation5) apply(simp add: P_cp[symmetric], subst (4) false_def, subst (4) bot_option_def, simp) apply(simp add: OclSelect_def[simplified inv_bot'] OclSelect_body_def StrictRefEq\<^sub>B\<^sub>o\<^sub>o\<^sub>l\<^sub>e\<^sub>a\<^sub>n) apply(subst (1 2 3 4) cp_OclIf, subst (1 2 3 4) foundation18'[THEN iffD2, simplified OclValid_def], simp, simp only: cp_OclIf[symmetric] refl if_True) apply(subst (1 2) OclIncluding.cp0, rule conj_split2, simp add: cp_OclIf[symmetric]) apply(subst (1 2 3 4 5 6 7 8) cp_OclIf[symmetric], simp) apply(( subst ex_excluding1[symmetric] | subst al_excluding1[symmetric] ), metis OclValid_def foundation5, metis OclValid_def foundation5, simp add: P_cp[symmetric] bot_fun_def)+ apply(simp add: bot_fun_def) apply(subst (1 2) invert_including, simp+) (* *) apply(rule conjI, blast) apply(intro impI conjI) apply(subst OclExcluding_def) apply(drule foundation5[simplified OclValid_def], simp) apply(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse[OF diff_in_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e], fast) apply(simp add: OclIncluding_def cp_valid[symmetric]) apply((erule conjE)+, frule exclude_defined[simplified OclValid_def], simp) apply(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse[OF ins_in_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e'''], simp+) apply(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject[OF ins_in_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e ins_in_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e'], fast+) (* *) apply(simp add: OclExcluding_def) apply(simp add: foundation10[simplified OclValid_def]) apply(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse[OF diff_in_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e], simp+) apply(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject[OF ins_in_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e'' ins_in_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e'''], simp+) apply(subgoal_tac "P (\_. y \) \ = false \") prefer 2 apply(subst P_cp[symmetric], metis OclValid_def foundation22) apply(rule equalityI) apply(rule subsetI, simp, metis) apply(rule subsetI, simp) (* *) apply(drule defined_inject_true) apply(subgoal_tac "\ (\ \ \ X) \ \ (\ \ \ y)") prefer 2 apply (metis OclIncluding.def_homo OclIncluding_valid_args_valid OclIncluding_valid_args_valid'' OclValid_def foundation18 valid1) apply(subst cp_OclSelect_body, subst cp_OclSelect, subst OclExcluding_def) apply(simp add: OclValid_def false_def true_def, rule conjI, blast) apply(simp add: OclSelect_invalid[simplified invalid_def] OclSelect_body_strict1[simplified invalid_def] inv_bot') done qed subsubsection\Execution Rules on Reject\ lemma OclReject_mtSet_exec[simp,code_unfold]: "OclReject mtSet P = mtSet" by(simp add: OclReject_def) lemma OclReject_including_exec[simp,code_unfold]: assumes P_cp : "cp P" shows "OclReject (X->including\<^sub>S\<^sub>e\<^sub>t(y)) P = OclSelect_body (not o P) y (OclReject (X->excluding\<^sub>S\<^sub>e\<^sub>t(y)) P)" apply(simp add: OclReject_def comp_def, rule OclSelect_including_exec) by (metis assms cp_intro'(5)) subsubsection\Execution Rules Combining Previous Operators\ text\OclIncluding\ (* logical level : *) lemma OclIncluding_idem0 : assumes "\ \ \ S" and "\ \ \ i" shows "\ \ (S->including\<^sub>S\<^sub>e\<^sub>t(i)->including\<^sub>S\<^sub>e\<^sub>t(i) \ (S->including\<^sub>S\<^sub>e\<^sub>t(i)))" by(simp add: OclIncluding_includes OclIncludes_charn1 assms) (* Pure algebraic level *) theorem OclIncluding_idem[simp,code_unfold]: "((S :: ('\,'a::null)Set)->including\<^sub>S\<^sub>e\<^sub>t(i)->including\<^sub>S\<^sub>e\<^sub>t(i) = (S->including\<^sub>S\<^sub>e\<^sub>t(i)))" proof - have A: "\ \. \ \ (i \ invalid) \ (S->including\<^sub>S\<^sub>e\<^sub>t(i)->including\<^sub>S\<^sub>e\<^sub>t(i)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) have A':"\ \. \ \ (i \ invalid) \ (S->including\<^sub>S\<^sub>e\<^sub>t(i)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) have C: "\ \. \ \ (S \ invalid) \ (S->including\<^sub>S\<^sub>e\<^sub>t(i)->including\<^sub>S\<^sub>e\<^sub>t(i)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) have C': "\ \. \ \ (S \ invalid) \ (S->including\<^sub>S\<^sub>e\<^sub>t(i)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) have D: "\ \. \ \ (S \ null) \ (S->including\<^sub>S\<^sub>e\<^sub>t(i)->including\<^sub>S\<^sub>e\<^sub>t(i)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) have D': "\ \. \ \ (S \ null) \ (S->including\<^sub>S\<^sub>e\<^sub>t(i)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) show ?thesis apply(rule ext, rename_tac \) apply(case_tac "\ \ (\ i)") apply(case_tac "\ \ (\ S)") apply(simp only: OclIncluding_idem0[THEN foundation22[THEN iffD1]]) apply(simp add: foundation16', elim disjE) apply(simp add: C[OF foundation22[THEN iffD2]] C'[OF foundation22[THEN iffD2]]) apply(simp add: D[OF foundation22[THEN iffD2]] D'[OF foundation22[THEN iffD2]]) apply(simp add:foundation18 A[OF foundation22[THEN iffD2]] A'[OF foundation22[THEN iffD2]]) done qed text\OclExcluding\ (* logical level : *) lemma OclExcluding_idem0 : assumes "\ \ \ S" and "\ \ \ i" shows "\ \ (S->excluding\<^sub>S\<^sub>e\<^sub>t(i)->excluding\<^sub>S\<^sub>e\<^sub>t(i) \ (S->excluding\<^sub>S\<^sub>e\<^sub>t(i)))" by(simp add: OclExcluding_excludes OclExcludes_charn1 assms) (* Pure algebraic level *) theorem OclExcluding_idem[simp,code_unfold]: "((S->excluding\<^sub>S\<^sub>e\<^sub>t(i))->excluding\<^sub>S\<^sub>e\<^sub>t(i)) = (S->excluding\<^sub>S\<^sub>e\<^sub>t(i))" proof - have A: "\ \. \ \ (i \ invalid) \ (S->excluding\<^sub>S\<^sub>e\<^sub>t(i)->excluding\<^sub>S\<^sub>e\<^sub>t(i)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) have A':"\ \. \ \ (i \ invalid) \ (S->excluding\<^sub>S\<^sub>e\<^sub>t(i)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) have C: "\ \. \ \ (S \ invalid) \ (S->excluding\<^sub>S\<^sub>e\<^sub>t(i)->excluding\<^sub>S\<^sub>e\<^sub>t(i)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) have C': "\ \. \ \ (S \ invalid) \ (S->excluding\<^sub>S\<^sub>e\<^sub>t(i)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) have D: "\ \. \ \ (S \ null) \ (S->excluding\<^sub>S\<^sub>e\<^sub>t(i)->excluding\<^sub>S\<^sub>e\<^sub>t(i)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) have D': "\ \. \ \ (S \ null) \ (S->excluding\<^sub>S\<^sub>e\<^sub>t(i)) \ = invalid \" apply(rule foundation22[THEN iffD1]) by(erule StrongEq_L_subst2_rev, simp,simp) show ?thesis apply(rule ext, rename_tac \) apply(case_tac "\ \ (\ i)") apply(case_tac "\ \ (\ S)") apply(simp only: OclExcluding_idem0[THEN foundation22[THEN iffD1]]) apply(simp add: foundation16', elim disjE) apply(simp add: C[OF foundation22[THEN iffD2]] C'[OF foundation22[THEN iffD2]]) apply(simp add: D[OF foundation22[THEN iffD2]] D'[OF foundation22[THEN iffD2]]) apply(simp add:foundation18 A[OF foundation22[THEN iffD2]] A'[OF foundation22[THEN iffD2]]) done qed text\OclIncludes\ lemma OclIncludes_any[simp,code_unfold]: "X->includes\<^sub>S\<^sub>e\<^sub>t(X->any\<^sub>S\<^sub>e\<^sub>t()) = (if \ X then if \ (X->size\<^sub>S\<^sub>e\<^sub>t()) then not(X->isEmpty\<^sub>S\<^sub>e\<^sub>t()) else X->includes\<^sub>S\<^sub>e\<^sub>t(null) endif else invalid endif)" proof - have defined_inject_true : "\\ P. (\ P) \ \ true \ \ (\ P) \ = false \" apply(simp add: defined_def true_def false_def bot_fun_def bot_option_def null_fun_def null_option_def) by (case_tac " P \ = \ \ P \ = null", simp_all add: true_def) have valid_inject_true : "\\ P. (\ P) \ \ true \ \ (\ P) \ = false \" apply(simp add: valid_def true_def false_def bot_fun_def bot_option_def null_fun_def null_option_def) by (case_tac "P \ = \", simp_all add: true_def) have notempty': "\\ X. \ \ \ X \ finite \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\ \ not (X->isEmpty\<^sub>S\<^sub>e\<^sub>t()) \ \ true \ \ X \ = Set{} \" apply(case_tac "X \", rename_tac X', simp add: mtSet_def Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject) apply(erule disjE, metis (opaque_lifting, mono_tags) bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_option_def foundation16) apply(erule disjE, metis (opaque_lifting, no_types) bot_option_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_option_def foundation16[THEN iffD1]) apply(case_tac X', simp, metis (opaque_lifting, no_types) bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def foundation16[THEN iffD1]) apply(rename_tac X'', case_tac X'', simp) apply (metis (opaque_lifting, no_types) foundation16[THEN iffD1] null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def) apply(simp add: OclIsEmpty_def OclSize_def) apply(subst (asm) cp_OclNot, subst (asm) cp_OclOr, subst (asm) StrictRefEq\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r.cp0, subst (asm) cp_OclAnd, subst (asm) cp_OclNot) apply(simp only: OclValid_def foundation20[simplified OclValid_def] cp_OclNot[symmetric] cp_OclAnd[symmetric] cp_OclOr[symmetric]) apply(simp add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse split: if_split_asm) by(simp add: true_def OclInt0_def OclNot_def StrictRefEq\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r StrongEq_def) have B: "\X \. \ finite \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\ \ (\ (X->size\<^sub>S\<^sub>e\<^sub>t())) \ = false \" apply(subst cp_defined) apply(simp add: OclSize_def) by (metis bot_fun_def defined_def) show ?thesis apply(rule ext, rename_tac \, simp only: OclIncludes_def OclANY_def) apply(subst cp_OclIf, subst (2) cp_valid) apply(case_tac "(\ X) \ = true \", simp only: foundation20[simplified OclValid_def] cp_OclIf[symmetric], simp, subst (1 2) cp_OclAnd, simp add: cp_OclAnd[symmetric]) apply(case_tac "finite \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\") apply(frule size_defined'[THEN iffD2, simplified OclValid_def], assumption) apply(subst (1 2 3 4) cp_OclIf, simp) apply(subst (1 2 3 4) cp_OclIf[symmetric], simp) apply(case_tac "(X->notEmpty\<^sub>S\<^sub>e\<^sub>t()) \ = true \", simp) apply(frule OclNotEmpty_has_elt[simplified OclValid_def], simp) apply(simp add: OclNotEmpty_def cp_OclIf[symmetric]) apply(subgoal_tac "(SOME y. y \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\) \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\", simp add: true_def) apply(metis OclValid_def Set_inv_lemma foundation18' null_option_def true_def) apply(rule someI_ex, simp) apply(simp add: OclNotEmpty_def cp_valid[symmetric]) apply(subgoal_tac "\ (null \ \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\)", simp) apply(subst OclIsEmpty_def, simp add: OclSize_def) apply(subst cp_OclNot, subst cp_OclOr, subst StrictRefEq\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r.cp0, subst cp_OclAnd, subst cp_OclNot, simp add: OclValid_def foundation20[simplified OclValid_def] cp_OclNot[symmetric] cp_OclAnd[symmetric] cp_OclOr[symmetric]) apply(frule notempty'[simplified OclValid_def], (simp add: mtSet_def Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse OclInt0_def false_def)+) apply(drule notempty'[simplified OclValid_def], simp, simp) apply (metis (opaque_lifting, no_types) empty_iff mtSet_rep_set) (* *) apply(frule B) apply(subst (1 2 3 4) cp_OclIf, simp) apply(subst (1 2 3 4) cp_OclIf[symmetric], simp) apply(case_tac "(X->notEmpty\<^sub>S\<^sub>e\<^sub>t()) \ = true \", simp) apply(frule OclNotEmpty_has_elt[simplified OclValid_def], simp) apply(simp add: OclNotEmpty_def OclIsEmpty_def) apply(subgoal_tac "X->size\<^sub>S\<^sub>e\<^sub>t() \ = \") prefer 2 apply (metis (opaque_lifting, no_types) OclSize_def) apply(subst (asm) cp_OclNot, subst (asm) cp_OclOr, subst (asm) StrictRefEq\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r.cp0, subst (asm) cp_OclAnd, subst (asm) cp_OclNot) apply(simp add: OclValid_def foundation20[simplified OclValid_def] cp_OclNot[symmetric] cp_OclAnd[symmetric] cp_OclOr[symmetric]) apply(simp add: OclNot_def StrongEq_def StrictRefEq\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r valid_def false_def true_def bot_option_def bot_fun_def invalid_def) apply (metis bot_fun_def null_fun_def null_is_valid valid_def) by(drule defined_inject_true, simp add: false_def true_def OclIf_false[simplified false_def] invalid_def) qed text\OclSize\ lemma [simp,code_unfold]: "\ (Set{} ->size\<^sub>S\<^sub>e\<^sub>t()) = true" by simp lemma [simp,code_unfold]: "\ ((X ->including\<^sub>S\<^sub>e\<^sub>t(x)) ->size\<^sub>S\<^sub>e\<^sub>t()) = (\(X->size\<^sub>S\<^sub>e\<^sub>t()) and \(x))" proof - have defined_inject_true : "\\ P. (\ P) \ \ true \ \ (\ P) \ = false \" apply(simp add: defined_def true_def false_def bot_fun_def bot_option_def null_fun_def null_option_def) by (case_tac " P \ = \ \ P \ = null", simp_all add: true_def) have valid_inject_true : "\\ P. (\ P) \ \ true \ \ (\ P) \ = false \" apply(simp add: valid_def true_def false_def bot_fun_def bot_option_def null_fun_def null_option_def) by (case_tac "P \ = \", simp_all add: true_def) have OclIncluding_finite_rep_set : "\\. (\ X and \ x) \ = true \ \ finite \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X->including\<^sub>S\<^sub>e\<^sub>t(x) \)\\ = finite \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\" apply(rule OclIncluding_finite_rep_set) by(metis OclValid_def foundation5)+ have card_including_exec : "\\. (\ (\_. \\int (card \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X->including\<^sub>S\<^sub>e\<^sub>t(x) \)\\)\\)) \ = (\ (\_. \\int (card \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\)\\)) \" by(simp add: defined_def bot_fun_def bot_option_def null_fun_def null_option_def) show ?thesis apply(rule ext, rename_tac \) apply(case_tac "(\ (X->including\<^sub>S\<^sub>e\<^sub>t(x)->size\<^sub>S\<^sub>e\<^sub>t())) \ = true \", simp del: OclSize_including_exec) apply(subst cp_OclAnd, subst cp_defined, simp only: cp_defined[of "X->including\<^sub>S\<^sub>e\<^sub>t(x)->size\<^sub>S\<^sub>e\<^sub>t()"], simp add: OclSize_def) apply(case_tac "((\ X and \ x) \ = true \ \ finite \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X->including\<^sub>S\<^sub>e\<^sub>t(x) \)\\)", simp) apply(erule conjE, simp add: OclIncluding_finite_rep_set[simplified OclValid_def] card_including_exec cp_OclAnd[of "\ X" "\ x"] cp_OclAnd[of "true", THEN sym]) apply(subgoal_tac "(\ X) \ = true \ \ (\ x) \ = true \", simp) apply(rule foundation5[of _ "\ X" "\ x", simplified OclValid_def], simp only: cp_OclAnd[THEN sym]) apply(simp, simp add: defined_def true_def false_def bot_fun_def bot_option_def) apply(drule defined_inject_true[of "X->including\<^sub>S\<^sub>e\<^sub>t(x)->size\<^sub>S\<^sub>e\<^sub>t()"], simp del: OclSize_including_exec, simp only: cp_OclAnd[of "\ (X->size\<^sub>S\<^sub>e\<^sub>t())" "\ x"], simp add: cp_defined[of "X->including\<^sub>S\<^sub>e\<^sub>t(x)->size\<^sub>S\<^sub>e\<^sub>t()" ] cp_defined[of "X->size\<^sub>S\<^sub>e\<^sub>t()" ] del: OclSize_including_exec, simp add: OclSize_def card_including_exec del: OclSize_including_exec) apply(case_tac "(\ X and \ x) \ = true \ \ finite \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\", simp add: OclIncluding_finite_rep_set[simplified OclValid_def] card_including_exec, simp only: cp_OclAnd[THEN sym], simp add: defined_def bot_fun_def) apply(split if_split_asm) apply(simp add: OclIncluding_finite_rep_set[simplified OclValid_def] card_including_exec)+ apply(simp only: cp_OclAnd[THEN sym], simp, rule impI, erule conjE) apply(case_tac "(\ x) \ = true \", simp add: cp_OclAnd[of "\ X" "\ x"]) by(drule valid_inject_true[of "x"], simp add: cp_OclAnd[of _ "\ x"]) qed lemma [simp,code_unfold]: "\ ((X ->excluding\<^sub>S\<^sub>e\<^sub>t(x)) ->size\<^sub>S\<^sub>e\<^sub>t()) = (\(X->size\<^sub>S\<^sub>e\<^sub>t()) and \(x))" proof - have defined_inject_true : "\\ P. (\ P) \ \ true \ \ (\ P) \ = false \" apply(simp add: defined_def true_def false_def bot_fun_def bot_option_def null_fun_def null_option_def) by (case_tac " P \ = \ \ P \ = null", simp_all add: true_def) have valid_inject_true : "\\ P. (\ P) \ \ true \ \ (\ P) \ = false \" apply(simp add: valid_def true_def false_def bot_fun_def bot_option_def null_fun_def null_option_def) by (case_tac "P \ = \", simp_all add: true_def) have OclExcluding_finite_rep_set : "\\. (\ X and \ x) \ = true \ \ finite \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X->excluding\<^sub>S\<^sub>e\<^sub>t(x) \)\\ = finite \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\" apply(rule OclExcluding_finite_rep_set) by(metis OclValid_def foundation5)+ have card_excluding_exec : "\\. (\ (\_. \\int (card \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X->excluding\<^sub>S\<^sub>e\<^sub>t(x) \)\\)\\)) \ = (\ (\_. \\int (card \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\)\\)) \" by(simp add: defined_def bot_fun_def bot_option_def null_fun_def null_option_def) show ?thesis apply(rule ext, rename_tac \) apply(case_tac "(\ (X->excluding\<^sub>S\<^sub>e\<^sub>t(x)->size\<^sub>S\<^sub>e\<^sub>t())) \ = true \", simp) apply(subst cp_OclAnd, subst cp_defined, simp only: cp_defined[of "X->excluding\<^sub>S\<^sub>e\<^sub>t(x)->size\<^sub>S\<^sub>e\<^sub>t()"], simp add: OclSize_def) apply(case_tac "((\ X and \ x) \ = true \ \ finite \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X->excluding\<^sub>S\<^sub>e\<^sub>t(x) \)\\)", simp) apply(erule conjE, simp add: OclExcluding_finite_rep_set[simplified OclValid_def] card_excluding_exec cp_OclAnd[of "\ X" "\ x"] cp_OclAnd[of "true", THEN sym]) apply(subgoal_tac "(\ X) \ = true \ \ (\ x) \ = true \", simp) apply(rule foundation5[of _ "\ X" "\ x", simplified OclValid_def], simp only: cp_OclAnd[THEN sym]) apply(simp, simp add: defined_def true_def false_def bot_fun_def bot_option_def) apply(drule defined_inject_true[of "X->excluding\<^sub>S\<^sub>e\<^sub>t(x)->size\<^sub>S\<^sub>e\<^sub>t()"], simp, simp only: cp_OclAnd[of "\ (X->size\<^sub>S\<^sub>e\<^sub>t())" "\ x"], simp add: cp_defined[of "X->excluding\<^sub>S\<^sub>e\<^sub>t(x)->size\<^sub>S\<^sub>e\<^sub>t()" ] cp_defined[of "X->size\<^sub>S\<^sub>e\<^sub>t()" ], simp add: OclSize_def card_excluding_exec) apply(case_tac "(\ X and \ x) \ = true \ \ finite \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\", simp add: OclExcluding_finite_rep_set[simplified OclValid_def] card_excluding_exec, simp only: cp_OclAnd[THEN sym], simp add: defined_def bot_fun_def) apply(split if_split_asm) apply(simp add: OclExcluding_finite_rep_set[simplified OclValid_def] card_excluding_exec)+ apply(simp only: cp_OclAnd[THEN sym], simp, rule impI, erule conjE) apply(case_tac "(\ x) \ = true \", simp add: cp_OclAnd[of "\ X" "\ x"]) by(drule valid_inject_true[of "x"], simp add: cp_OclAnd[of _ "\ x"]) qed lemma [simp]: assumes X_finite: "\\. finite \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\" shows "\ ((X ->including\<^sub>S\<^sub>e\<^sub>t(x)) ->size\<^sub>S\<^sub>e\<^sub>t()) = (\(X) and \(x))" by(simp add: size_defined[OF X_finite] del: OclSize_including_exec) text\OclForall\ lemma OclForall_rep_set_false: assumes "\ \ \ X" shows "(OclForall X P \ = false \) = (\x \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\. P (\\. x) \ = false \)" by(insert assms, simp add: OclForall_def OclValid_def false_def true_def invalid_def bot_fun_def bot_option_def null_fun_def null_option_def) lemma OclForall_rep_set_true: assumes "\ \ \ X" shows "(\ \ OclForall X P) = (\x \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\. \ \ P (\\. x))" proof - have destruct_ocl : "\x \. x = true \ \ x = false \ \ x = null \ \ x = \ \" apply(case_tac x) apply (metis bot_Boolean_def) apply(rename_tac x', case_tac x') apply (metis null_Boolean_def) apply(rename_tac x'', case_tac x'') apply (metis (full_types) true_def) by (metis (full_types) false_def) have disjE4 : "\ P1 P2 P3 P4 R. (P1 \ P2 \ P3 \ P4) \ (P1 \ R) \ (P2 \ R) \ (P3 \ R) \ (P4 \ R) \ R" by metis show ?thesis apply(simp add: OclForall_def OclValid_def true_def false_def invalid_def bot_fun_def bot_option_def null_fun_def null_option_def split: if_split_asm) apply(rule conjI, rule impI) apply (metis drop.simps option.distinct(1) invalid_def) apply(rule impI, rule conjI, rule impI) apply (metis option.distinct(1)) apply(rule impI, rule conjI, rule impI) apply (metis drop.simps) apply(intro conjI impI ballI) proof - fix x show "\x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\. P (\_. x) \ \ \None\ \ \x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\. \y. P (\_. x) \ = \y\ \ \x\\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\. P (\_. x) \ \ \\False\\ \ x \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\ \ P (\\. x) \ = \\True\\" apply(erule_tac x = x in ballE)+ by(rule disjE4[OF destruct_ocl[of "P (\\. x) \"]], (simp add: true_def false_def null_fun_def null_option_def bot_fun_def bot_option_def)+) qed(simp add: assms[simplified OclValid_def true_def])+ qed lemma OclForall_includes : assumes x_def : "\ \ \ x" and y_def : "\ \ \ y" shows "(\ \ OclForall x (OclIncludes y)) = (\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (x \)\\ \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (y \)\\)" apply(simp add: OclForall_rep_set_true[OF x_def], simp add: OclIncludes_def OclValid_def y_def[simplified OclValid_def]) apply(insert Set_inv_lemma[OF x_def], simp add: valid_def false_def true_def bot_fun_def) by(rule iffI, simp add: subsetI, simp add: subsetD) lemma OclForall_not_includes : assumes x_def : "\ \ \ x" and y_def : "\ \ \ y" shows "(OclForall x (OclIncludes y) \ = false \) = (\ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (x \)\\ \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (y \)\\)" apply(simp add: OclForall_rep_set_false[OF x_def], simp add: OclIncludes_def OclValid_def y_def[simplified OclValid_def]) apply(insert Set_inv_lemma[OF x_def], simp add: valid_def false_def true_def bot_fun_def) by(rule iffI, metis rev_subsetD, metis subsetI) lemma OclForall_iterate: assumes S_finite: "finite \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\" shows "S->forAll\<^sub>S\<^sub>e\<^sub>t(x | P x) \ = (S->iterate\<^sub>S\<^sub>e\<^sub>t(x; acc = true | acc and P x)) \" proof - interpret and_comm: comp_fun_commute "(\x acc. acc and P x)" apply(simp add: comp_fun_commute_def comp_def) by (metis OclAnd_assoc OclAnd_commute) have ex_insert : "\x F P. (\x\insert x F. P x) = (P x \ (\x\F. P x))" by (metis insert_iff) have destruct_ocl : "\x \. x = true \ \ x = false \ \ x = null \ \ x = \ \" apply(case_tac x) apply (metis bot_Boolean_def) apply(rename_tac x', case_tac x') apply (metis null_Boolean_def) apply(rename_tac x'', case_tac x'') apply (metis (full_types) true_def) by (metis (full_types) false_def) have disjE4 : "\ P1 P2 P3 P4 R. (P1 \ P2 \ P3 \ P4) \ (P1 \ R) \ (P2 \ R) \ (P3 \ R) \ (P4 \ R) \ R" by metis let ?P_eq = "\x b \. P (\_. x) \ = b \" let ?P = "\set b \. \x\set. ?P_eq x b \" let ?if = "\f b c. if f b \ then b \ else c" let ?forall = "\P. ?if P false (?if P invalid (?if P null (true \)))" show ?thesis apply(simp only: OclForall_def OclIterate_def) apply(case_tac "\ \ \ S", simp only: OclValid_def) apply(subgoal_tac "let set = \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\ in ?forall (?P set) = Finite_Set.fold (\x acc. acc and P x) true ((\a \. a) ` set) \", simp only: Let_def, simp add: S_finite, simp only: Let_def) apply(case_tac "\\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (S \)\\ = {}", simp) apply(rule finite_ne_induct[OF S_finite], simp) (* *) apply(simp only: image_insert) apply(subst and_comm.fold_insert, simp) apply (metis empty_iff image_empty) apply(simp add: invalid_def) apply (metis bot_fun_def destruct_ocl null_fun_def) (* *) apply(simp only: image_insert) apply(subst and_comm.fold_insert, simp) apply (metis (mono_tags) imageE) (* *) apply(subst cp_OclAnd) apply(drule sym, drule sym, simp only:, drule sym, simp only:) apply(simp only: ex_insert) apply(subgoal_tac "\x. x\F") prefer 2 apply(metis all_not_in_conv) proof - fix x F show "(\ S) \ = true \ \ \x. x \ F \ ?forall (\b \. ?P_eq x b \ \ ?P F b \) = ((\_. ?forall (?P F)) and (\_. P (\\. x) \)) \" - apply(rule disjE4[OF destruct_ocl[where x = "P (\\. x) \"]]) + apply(rule disjE4[OF destruct_ocl[where x1 = "P (\\. x) \"]]) apply(simp_all add: true_def false_def invalid_def OclAnd_def null_fun_def null_option_def bot_fun_def bot_option_def) by (metis (lifting) option.distinct(1))+ qed(simp add: OclValid_def)+ qed lemma OclForall_cong: assumes "\x. x \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\ \ \ \ P (\\. x) \ \ \ Q (\\. x)" assumes P: "\ \ OclForall X P" shows "\ \ OclForall X Q" proof - have def_X: "\ \ \ X" by(insert P, simp add: OclForall_def OclValid_def bot_option_def true_def split: if_split_asm) show ?thesis apply(insert P) apply(subst (asm) OclForall_rep_set_true[OF def_X], subst OclForall_rep_set_true[OF def_X]) by (simp add: assms) qed lemma OclForall_cong': assumes "\x. x \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (X \)\\ \ \ \ P (\\. x) \ \ \ Q (\\. x) \ \ \ R (\\. x)" assumes P: "\ \ OclForall X P" assumes Q: "\ \ OclForall X Q" shows "\ \ OclForall X R" proof - have def_X: "\ \ \ X" by(insert P, simp add: OclForall_def OclValid_def bot_option_def true_def split: if_split_asm) show ?thesis apply(insert P Q) apply(subst (asm) (1 2) OclForall_rep_set_true[OF def_X], subst OclForall_rep_set_true[OF def_X]) by (simp add: assms) qed text\Strict Equality\ lemma StrictRefEq\<^sub>S\<^sub>e\<^sub>t_defined : assumes x_def: "\ \ \ x" assumes y_def: "\ \ \ y" shows "((x::('\,'\::null)Set) \ y) \ = (x->forAll\<^sub>S\<^sub>e\<^sub>t(z| y->includes\<^sub>S\<^sub>e\<^sub>t(z)) and (y->forAll\<^sub>S\<^sub>e\<^sub>t(z| x->includes\<^sub>S\<^sub>e\<^sub>t(z)))) \" proof - have rep_set_inj : "\\. (\ x) \ = true \ \ (\ y) \ = true \ \ x \ \ y \ \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (y \)\\ \ \\Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (x \)\\" apply(simp add: defined_def) apply(split if_split_asm, simp add: false_def true_def)+ apply(simp add: null_fun_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def bot_fun_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def) apply(case_tac "x \", rename_tac x') apply(case_tac x', simp_all, rename_tac x'') apply(case_tac x'', simp_all) apply(case_tac "y \", rename_tac y') apply(case_tac y', simp_all, rename_tac y'') apply(case_tac y'', simp_all) apply(simp add: Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse) by(blast) show ?thesis apply(simp add: StrictRefEq\<^sub>S\<^sub>e\<^sub>t StrongEq_def foundation20[OF x_def, simplified OclValid_def] foundation20[OF y_def, simplified OclValid_def]) apply(subgoal_tac "\\x \ = y \\\ = true \ \ \\x \ = y \\\ = false \") prefer 2 apply(simp add: false_def true_def) (* *) apply(erule disjE) apply(simp add: true_def) apply(subgoal_tac "(\ \ OclForall x (OclIncludes y)) \ (\ \ OclForall y (OclIncludes x))") apply(subst cp_OclAnd, simp add: true_def OclValid_def) apply(simp add: OclForall_includes[OF x_def y_def] OclForall_includes[OF y_def x_def]) (* *) apply(simp) apply(subgoal_tac "OclForall x (OclIncludes y) \ = false \ \ OclForall y (OclIncludes x) \ = false \") apply(subst cp_OclAnd, metis OclAnd_false1 OclAnd_false2 cp_OclAnd) apply(simp only: OclForall_not_includes[OF x_def y_def, simplified OclValid_def] OclForall_not_includes[OF y_def x_def, simplified OclValid_def], simp add: false_def) by (metis OclValid_def rep_set_inj subset_antisym x_def y_def) qed lemma StrictRefEq\<^sub>S\<^sub>e\<^sub>t_exec[simp,code_unfold] : "((x::('\,'\::null)Set) \ y) = (if \ x then (if \ y then ((x->forAll\<^sub>S\<^sub>e\<^sub>t(z| y->includes\<^sub>S\<^sub>e\<^sub>t(z)) and (y->forAll\<^sub>S\<^sub>e\<^sub>t(z| x->includes\<^sub>S\<^sub>e\<^sub>t(z))))) else if \ y then false \ \\x'->includes = null\\ else invalid endif endif) else if \ x \ \\null = ???\\ then if \ y then not(\ y) else invalid endif else invalid endif endif)" proof - have defined_inject_true : "\\ P. (\ (\ \ \ P)) = ((\ P) \ = false \)" by (metis bot_fun_def OclValid_def defined_def foundation16 null_fun_def) have valid_inject_true : "\\ P. (\ (\ \ \ P)) = ((\ P) \ = false \)" by (metis bot_fun_def OclIf_true' OclIncludes_charn0 OclIncludes_charn0' OclValid_def valid_def foundation6 foundation9) show ?thesis apply(rule ext, rename_tac \) (* *) apply(simp add: OclIf_def defined_inject_true[simplified OclValid_def] valid_inject_true[simplified OclValid_def], subst false_def, subst true_def, simp) apply(subst (1 2) cp_OclNot, simp, simp add: cp_OclNot[symmetric]) apply(simp add: StrictRefEq\<^sub>S\<^sub>e\<^sub>t_defined[simplified OclValid_def]) by(simp add: StrictRefEq\<^sub>S\<^sub>e\<^sub>t StrongEq_def false_def true_def valid_def defined_def) qed lemma StrictRefEq\<^sub>S\<^sub>e\<^sub>t_L_subst1 : "cp P \ \ \ \ x \ \ \ \ y \ \ \ \ P x \ \ \ \ P y \ \ \ (x::('\,'\::null)Set) \ y \ \ \ (P x ::('\,'\::null)Set) \ P y" apply(simp only: StrictRefEq\<^sub>S\<^sub>e\<^sub>t OclValid_def) apply(split if_split_asm) apply(simp add: StrongEq_L_subst1[simplified OclValid_def]) by (simp add: invalid_def bot_option_def true_def) lemma OclIncluding_cong' : shows "\ \ \ s \ \ \ \ t \ \ \ \ x \ \ \ ((s::('\,'a::null)Set) \ t) \ \ \ (s->including\<^sub>S\<^sub>e\<^sub>t(x) \ (t->including\<^sub>S\<^sub>e\<^sub>t(x)))" proof - have cp: "cp (\s. (s->including\<^sub>S\<^sub>e\<^sub>t(x)))" apply(simp add: cp_def, subst OclIncluding.cp0) by (rule_tac x = "(\xab ab. ((\_. xab)->including\<^sub>S\<^sub>e\<^sub>t(\_. x ab)) ab)" in exI, simp) show "\ \ \ s \ \ \ \ t \ \ \ \ x \ \ \ (s \ t) \ ?thesis" apply(rule_tac P = "\s. (s->including\<^sub>S\<^sub>e\<^sub>t(x))" in StrictRefEq\<^sub>S\<^sub>e\<^sub>t_L_subst1) apply(rule cp) apply(simp add: foundation20) apply(simp add: foundation20) apply (simp add: foundation10 foundation6)+ done qed lemma OclIncluding_cong : "\(s::('\,'a::null)Set) t x y \. \ \ \ t \ \ \ \ y \ \ \ s \ t \ x = y \ \ \ s->including\<^sub>S\<^sub>e\<^sub>t(x) \ (t->including\<^sub>S\<^sub>e\<^sub>t(y))" apply(simp only:) apply(rule OclIncluding_cong', simp_all only:) by(auto simp: OclValid_def OclIf_def invalid_def bot_option_def OclNot_def split : if_split_asm) (* < *) lemma const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_empty : "const X \ const (X \ Set{})" apply(rule StrictRefEq\<^sub>S\<^sub>e\<^sub>t.const, assumption) by(simp) lemma const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including : "const a \ const S \ const X \ const (X \ S->including\<^sub>S\<^sub>e\<^sub>t(a))" apply(rule StrictRefEq\<^sub>S\<^sub>e\<^sub>t.const, assumption) by(rule const_OclIncluding) subsection\Test Statements\ Assert "(\ \ (Set{\_. \\x\\} \ Set{\_. \\x\\}))" Assert "(\ \ (Set{\_. \x\} \ Set{\_. \x\}))" instantiation Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e :: (equal)equal begin definition "HOL.equal k l \ (k::('a::equal)Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e) = l" instance by standard (rule equal_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def) end lemma equal_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_code [code]: "HOL.equal k (l::('a::{equal,null})Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e) \ Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e k = Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e l" by (auto simp add: equal Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e.Rep_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject) Assert "\ \ (Set{} \ Set{})" Assert "\ \ (Set{\,\} \ Set{}->including\<^sub>S\<^sub>e\<^sub>t(\)->including\<^sub>S\<^sub>e\<^sub>t(\))" Assert "\ \ (Set{\,invalid,\} \ invalid)" Assert "\ \ (Set{\,\}->including\<^sub>S\<^sub>e\<^sub>t(null) \ Set{null,\,\})" Assert "\ \ (Set{\,\}->including\<^sub>S\<^sub>e\<^sub>t(null) \ Set{\,\,null})" (* Assert "\ (\ \ (Set{\,\,\} \ Set{\,\}))" Assert "\ (\ \ (Set{\,\} \ Set{\,\}))" *) (* > *) end diff --git a/thys/Iptables_Semantics/Primitive_Matchers/Transform.thy b/thys/Iptables_Semantics/Primitive_Matchers/Transform.thy --- a/thys/Iptables_Semantics/Primitive_Matchers/Transform.thy +++ b/thys/Iptables_Semantics/Primitive_Matchers/Transform.thy @@ -1,1875 +1,1875 @@ section\Optimizing and Normalizing Primitives\ theory Transform imports Common_Primitive_Lemmas "../Semantics_Ternary/Semantics_Ternary" "../Semantics_Ternary/Negation_Type_Matching" Ports_Normalize IpAddresses_Normalize Interfaces_Normalize Protocols_Normalize "../Common/Remdups_Rev" Interface_Replace "../Semantics_Ternary/Optimizing" begin text\This transform theory plugs a lot of stuff together. We perform several normalization and optimization steps on complete firewall rulesets. We show that it preserves the semantics and also, that structural properties are preserved. For example, if you normalize interfaces and afterwards normalize protocols, the interfaces remain normalized and no new interfaces are added when doing the protocol normalization.\ (*Maintainer note: we plug a lot of lemmata together to show that structural properties are preserved. Yes, there is a huge set of apply style in there but there is no magic happening, it is just pushing through invariants about structural properties.*) definition compress_normalize_besteffort :: "'i::len common_primitive match_expr \ 'i common_primitive match_expr option" where "compress_normalize_besteffort m \ compress_normalize_primitive_monad [compress_normalize_protocols, compress_normalize_input_interfaces, compress_normalize_output_interfaces] m" context begin private lemma compress_normalize_besteffort_normalized: "f \ set [compress_normalize_protocols, compress_normalize_input_interfaces, compress_normalize_output_interfaces] \ normalized_nnf_match m \ f m = Some m' \ normalized_nnf_match m'" apply(simp) apply(elim disjE) using compress_normalize_protocols_nnf apply blast using compress_normalize_input_interfaces_nnf apply blast using compress_normalize_output_interfaces_nnf apply blast done private lemma compress_normalize_besteffort_matches: assumes generic: "primitive_matcher_generic \" shows "f \ set [compress_normalize_protocols, compress_normalize_input_interfaces, compress_normalize_output_interfaces] \ normalized_nnf_match m \ f m = Some m' \ matches (\, \) m' a p = matches (\, \) m a p" apply(simp) apply(elim disjE) using primitive_matcher_generic.compress_normalize_protocols_Some[OF generic] apply blast using compress_normalize_input_interfaces_Some[OF generic] apply blast using compress_normalize_output_interfaces_Some[OF generic] apply blast done lemma compress_normalize_besteffort_Some: assumes generic: "primitive_matcher_generic \" shows "normalized_nnf_match m \ compress_normalize_besteffort m = Some m' \ matches (\, \) m' a p = matches (\, \) m a p" unfolding compress_normalize_besteffort_def apply(rule compress_normalize_primitive_monad) using compress_normalize_besteffort_normalized compress_normalize_besteffort_matches[OF generic] by blast+ lemma compress_normalize_besteffort_None: assumes generic: "primitive_matcher_generic \" shows "normalized_nnf_match m \ compress_normalize_besteffort m = None \ \ matches (\, \) m a p" proof - have notmatches: "f \ set [compress_normalize_protocols, compress_normalize_input_interfaces, compress_normalize_output_interfaces] \ normalized_nnf_match m \ f m = None \ \ matches (\, \) m a p" for f m apply(simp) using primitive_matcher_generic.compress_normalize_protocols_None[OF generic] compress_normalize_input_interfaces_None[OF generic] compress_normalize_output_interfaces_None[OF generic] by blast show "normalized_nnf_match m \ compress_normalize_besteffort m = None \ \ matches (\, \) m a p" unfolding compress_normalize_besteffort_def apply(rule compress_normalize_primitive_monad_None) using compress_normalize_besteffort_normalized compress_normalize_besteffort_matches[OF generic] notmatches by blast+ qed lemma compress_normalize_besteffort_nnf: "normalized_nnf_match m \ compress_normalize_besteffort m = Some m' \ normalized_nnf_match m'" unfolding compress_normalize_besteffort_def apply(rule compress_normalize_primitive_monad) using compress_normalize_besteffort_normalized compress_normalize_besteffort_matches[OF primitive_matcher_generic_common_matcher] by blast+ lemma compress_normalize_besteffort_not_introduces_Iiface: "\ has_disc is_Iiface m \ normalized_nnf_match m \ compress_normalize_besteffort m = Some m' \ \ has_disc is_Iiface m'" unfolding compress_normalize_besteffort_def apply(rule compress_normalize_primitive_monad_preserves[THEN conjunct2]) apply(drule(3) compress_normalize_besteffort_normalized) apply(auto dest: compress_normalize_input_interfaces_not_introduces_Iiface compress_normalize_protocols_hasdisc compress_normalize_output_interfaces_hasdisc) done lemma compress_normalize_besteffort_not_introduces_Oiface: "\ has_disc is_Oiface m \ normalized_nnf_match m \ compress_normalize_besteffort m = Some m' \ \ has_disc is_Oiface m'" unfolding compress_normalize_besteffort_def apply(rule compress_normalize_primitive_monad_preserves[THEN conjunct2]) apply(drule(3) compress_normalize_besteffort_normalized) apply(auto dest: compress_normalize_output_interfaces_hasdisc compress_normalize_output_interfaces_not_introduces_Oiface compress_normalize_protocols_hasdisc compress_normalize_input_interfaces_hasdisc) done lemma compress_normalize_besteffort_not_introduces_Iiface_negated: "\ has_disc_negated is_Iiface False m \ normalized_nnf_match m \ compress_normalize_besteffort m = Some m' \ \ has_disc_negated is_Iiface False m'" unfolding compress_normalize_besteffort_def apply(rule compress_normalize_primitive_monad_preserves[THEN conjunct2]) apply(drule(3) compress_normalize_besteffort_normalized) apply(auto dest: compress_normalize_besteffort_normalized compress_normalize_input_interfaces_not_introduces_Iiface_negated compress_normalize_protocols_hasdisc_negated compress_normalize_output_interfaces_hasdisc_negated) done lemma compress_normalize_besteffort_not_introduces_Oiface_negated: "\ has_disc_negated is_Oiface False m \ normalized_nnf_match m \ compress_normalize_besteffort m = Some m' \ \ has_disc_negated is_Oiface False m'" unfolding compress_normalize_besteffort_def apply(rule compress_normalize_primitive_monad_preserves[THEN conjunct2]) apply(drule(3) compress_normalize_besteffort_normalized) apply(auto dest: compress_normalize_output_interfaces_not_introduces_Oiface_negated compress_normalize_input_interfaces_hasdisc_negated compress_normalize_protocols_hasdisc_negated) done lemma compress_normalize_besteffort_not_introduces_Prot_negated: "\ has_disc_negated is_Prot False m \ normalized_nnf_match m \ compress_normalize_besteffort m = Some m' \ \ has_disc_negated is_Prot False m'" unfolding compress_normalize_besteffort_def apply(rule compress_normalize_primitive_monad_preserves[THEN conjunct2]) apply(drule(3) compress_normalize_besteffort_normalized) apply(auto dest: compress_normalize_input_interfaces_hasdisc_negated compress_normalize_protocols_not_introduces_Prot_negated compress_normalize_output_interfaces_hasdisc_negated) done lemma compress_normalize_besteffort_hasdisc: "\ has_disc disc m \ (\a. \ disc (IIface a)) \ (\a. \ disc (OIface a)) \ (\a. \ disc (Prot a)) \ normalized_nnf_match m \ compress_normalize_besteffort m = Some m' \ normalized_nnf_match m' \ \ has_disc disc m'" unfolding compress_normalize_besteffort_def apply(rule compress_normalize_primitive_monad_preserves) apply(drule(3) compress_normalize_besteffort_normalized) apply(auto dest: compress_normalize_input_interfaces_hasdisc compress_normalize_output_interfaces_hasdisc compress_normalize_protocols_hasdisc) done lemma compress_normalize_besteffort_hasdisc_negated: "\ has_disc_negated disc False m \ (\a. \ disc (IIface a)) \ (\a. \ disc (OIface a)) \ (\a. \ disc (Prot a)) \ normalized_nnf_match m \ compress_normalize_besteffort m = Some m' \ normalized_nnf_match m' \ \ has_disc_negated disc False m'" (*due to protocols, we can only show for neg := False*) unfolding compress_normalize_besteffort_def apply(rule compress_normalize_primitive_monad_preserves) apply(drule(3) compress_normalize_besteffort_normalized) apply(simp split: option.split_asm) using compress_normalize_input_interfaces_hasdisc_negated compress_normalize_output_interfaces_hasdisc_negated compress_normalize_protocols_hasdisc_negated apply blast apply simp_all done lemma compress_normalize_besteffort_preserves_normalized_n_primitive: "normalized_n_primitive (disc, sel) P m \ (\a. \ disc (IIface a)) \ (\a. \ disc (OIface a)) \ (\a. \ disc (Prot a)) \ normalized_nnf_match m \ compress_normalize_besteffort m = Some m' \ normalized_nnf_match m' \ normalized_n_primitive (disc, sel) P m'" unfolding compress_normalize_besteffort_def apply(rule compress_normalize_primitive_monad_preserves) apply(drule(3) compress_normalize_besteffort_normalized) apply(auto dest: compress_normalize_input_interfaces_preserves_normalized_n_primitive compress_normalize_output_interfaces_preserves_normalized_n_primitive compress_normalize_protocols_preserves_normalized_n_primitive) done end section\Transforming rulesets\ subsection\Optimizations\ lemma approximating_bigstep_fun_remdups_rev: "approximating_bigstep_fun \ p (remdups_rev rs) s = approximating_bigstep_fun \ p rs s" proof(induction \ p rs s rule: approximating_bigstep_fun.induct) case 1 thus ?case by(simp add: remdups_rev_def) next case 2 thus ?case by (simp add: Decision_approximating_bigstep_fun) next case (3 \ p m a rs) thus ?case proof(cases "matches \ m a p") case False with 3 show ?thesis by(simp add: remdups_rev_fst remdups_rev_removeAll not_matches_removeAll) next case True { fix rs s have "approximating_bigstep_fun \ p (filter ((\) (Rule m Log)) rs) s = approximating_bigstep_fun \ p rs s" proof(induction \ p rs s rule: approximating_bigstep_fun_induct) qed(auto simp add: Decision_approximating_bigstep_fun split: action.split) } note helper_Log=this { fix rs s have "approximating_bigstep_fun \ p (filter ((\) (Rule m Empty)) rs) s = approximating_bigstep_fun \ p rs s" proof(induction \ p rs s rule: approximating_bigstep_fun_induct) qed(auto simp add: Decision_approximating_bigstep_fun split: action.split) } note helper_Empty=this from True 3 show ?thesis apply(simp add: remdups_rev_fst split: action.split) apply(safe) apply(simp_all) apply(simp_all add: remdups_rev_removeAll) apply(simp_all add: removeAll_filter_not_eq helper_Empty helper_Log) done qed qed lemma remdups_rev_simplers: "simple_ruleset rs \ simple_ruleset (remdups_rev rs)" by(induction rs) (simp_all add: remdups_rev_def simple_ruleset_def) lemma remdups_rev_preserve_matches: "\r\set rs. P (get_match r) \ \r\set (remdups_rev rs). P (get_match r)" by(induction rs) (simp_all add: remdups_rev_def simple_ruleset_def) subsection\Optimize and Normalize to NNF form\ (*without normalize_rules_dnf, the result cannot be normalized as optimize_primitive_univ can contain MatchNot MatchAny*) definition transform_optimize_dnf_strict :: "'i::len common_primitive rule list \ 'i common_primitive rule list" where "transform_optimize_dnf_strict = cut_off_after_match_any \ (optimize_matches opt_MatchAny_match_expr \ normalize_rules_dnf \ (optimize_matches (opt_MatchAny_match_expr \ optimize_primitive_univ)))" theorem transform_optimize_dnf_strict_structure: assumes simplers: "simple_ruleset rs" and wf\: "wf_unknown_match_tac \" shows "simple_ruleset (transform_optimize_dnf_strict rs)" and "\ r \ set rs. \ has_disc disc (get_match r) \ \ r \ set (transform_optimize_dnf_strict rs). \ has_disc disc (get_match r)" and "\ r \ set (transform_optimize_dnf_strict rs). normalized_nnf_match (get_match r)" and "\ r \ set rs. normalized_n_primitive disc_sel f (get_match r) \ \ r \ set (transform_optimize_dnf_strict rs). normalized_n_primitive disc_sel f (get_match r)" and "\ r \ set rs. \ has_disc_negated disc neg (get_match r) \ \ r \ set (transform_optimize_dnf_strict rs). \ has_disc_negated disc neg (get_match r)" proof - show simplers_transform: "simple_ruleset (transform_optimize_dnf_strict rs)" unfolding transform_optimize_dnf_strict_def using simplers by (simp add: cut_off_after_match_any_simplers optimize_matches_simple_ruleset simple_ruleset_normalize_rules_dnf) define transform_optimize_dnf_strict_inner where "transform_optimize_dnf_strict_inner = optimize_matches (opt_MatchAny_match_expr :: 'a common_primitive match_expr \ 'a common_primitive match_expr) \ normalize_rules_dnf \ (optimize_matches (opt_MatchAny_match_expr \ optimize_primitive_univ))" have inner_outer: "transform_optimize_dnf_strict = (cut_off_after_match_any \ transform_optimize_dnf_strict_inner)" by(auto simp add: transform_optimize_dnf_strict_def transform_optimize_dnf_strict_inner_def) have tf1: "transform_optimize_dnf_strict_inner (r#rs) = (optimize_matches opt_MatchAny_match_expr (normalize_rules_dnf (optimize_matches (opt_MatchAny_match_expr \ optimize_primitive_univ) [r])))@ transform_optimize_dnf_strict_inner rs" for r rs unfolding transform_optimize_dnf_strict_inner_def apply(simp) apply(subst optimize_matches_fst) apply(simp add: normalize_rules_dnf_append optimize_matches_append) done \ \if the individual optimization functions preserve a property, then the whole thing does\ { fix P :: "'a::len common_primitive match_expr \ bool" assume p1: "\m. P m \ P (optimize_primitive_univ m)" assume p2: "\m. P m \ P (opt_MatchAny_match_expr m)" assume p3: "\m. P m \ (\m' \ set (normalize_match m). P m')" { fix rs have "\ r \ set rs. P (get_match r) \ \ r \ set (optimize_matches (opt_MatchAny_match_expr \ optimize_primitive_univ) rs). P (get_match r)" apply(rule optimize_matches_preserves) using p1 p2 by simp } note opt1=this { fix rs have "\ r \ set rs. P (get_match r) \ \ r \ set (normalize_rules_dnf rs). P (get_match r)" apply(induction rs rule: normalize_rules_dnf.induct) apply(simp; fail) apply(simp) apply(safe) apply(simp_all) using p3 by(simp) } note opt2=this { fix rs have "\ r \ set rs. P (get_match r) \ \ r \ set (optimize_matches opt_MatchAny_match_expr rs). P (get_match r)" apply(rule optimize_matches_preserves) using p2 by simp } note opt3=this have "\ r \ set rs. P (get_match r) \ \ r \ set (transform_optimize_dnf_strict rs). P (get_match r)" unfolding transform_optimize_dnf_strict_def apply(drule opt1) apply(drule opt2) apply(drule opt3) using cut_off_after_match_any_preserve_matches by auto } note matchpred_rule=this { fix m have "\ has_disc disc m \ \ has_disc disc (optimize_primitive_univ m)" by(induction m rule: optimize_primitive_univ.induct) (simp_all) } moreover { fix m have "\ has_disc disc m \ (\m' \ set (normalize_match m). \ has_disc disc m')" using normalize_match_preserves_nodisc by fast } ultimately show "\ r \ set rs. \ has_disc disc (get_match r) \ \ r \ set (transform_optimize_dnf_strict rs). \ has_disc disc (get_match r)" using not_has_disc_opt_MatchAny_match_expr matchpred_rule[of "\m. \ has_disc disc m"] by fast { fix m have "\ has_disc_negated disc neg m \ \ has_disc_negated disc neg (optimize_primitive_univ m)" apply(induction disc neg m rule: has_disc_negated.induct) apply(simp_all) apply(rename_tac a) apply(subgoal_tac "optimize_primitive_univ (Match a) = Match a \ optimize_primitive_univ (Match a) = MatchAny") apply safe apply simp_all using optimize_primitive_univ_unchanged_primitives by blast } with not_has_disc_negated_opt_MatchAny_match_expr not_has_disc_normalize_match show "\ r \ set rs. \ has_disc_negated disc neg (get_match r) \ \ r \ set (transform_optimize_dnf_strict rs). \ has_disc_negated disc neg (get_match r)" using matchpred_rule[of "\m. \ has_disc_negated disc neg m"] by fast { fix P and a::"'a common_primitive" have "(optimize_primitive_univ (Match a)) = (Match a) \ (optimize_primitive_univ (Match a)) = MatchAny" by(induction "(Match a)" rule: optimize_primitive_univ.induct) (auto) hence "((optimize_primitive_univ (Match a)) = Match a \ P a) \ (optimize_primitive_univ (Match a) = MatchAny \ P a) \ P a" by blast } note optimize_primitive_univ_match_cases=this { fix m have "normalized_n_primitive disc_sel f m \ normalized_n_primitive disc_sel f (optimize_primitive_univ m)" apply(induction disc_sel f m rule: normalized_n_primitive.induct) apply(simp_all split: if_split_asm) apply(rule optimize_primitive_univ_match_cases, simp_all)+ done } moreover { fix m have "normalized_n_primitive disc_sel f m \ (\m' \ set (normalize_match m). normalized_n_primitive disc_sel f m')" using normalize_match_preserves_normalized_n_primitive by blast } ultimately show "\ r \ set rs. normalized_n_primitive disc_sel f (get_match r) \ \ r \ set (transform_optimize_dnf_strict rs). normalized_n_primitive disc_sel f (get_match r)" using matchpred_rule[of "\m. normalized_n_primitive disc_sel f m"] normalized_n_primitive_opt_MatchAny_match_expr by fast { fix rs::"'a::len common_primitive rule list" from normalize_rules_dnf_normalized_nnf_match[of "rs"] have "\x \ set (normalize_rules_dnf rs). normalized_nnf_match (get_match x)" . hence "\r \ set (optimize_matches opt_MatchAny_match_expr (normalize_rules_dnf rs)). normalized_nnf_match (get_match r)" apply - apply(rule optimize_matches_preserves) using normalized_nnf_match_opt_MatchAny_match_expr by blast } from this have "\ r \ set (transform_optimize_dnf_strict_inner rs). normalized_nnf_match (get_match r)" unfolding transform_optimize_dnf_strict_inner_def by simp thus "\ r \ set (transform_optimize_dnf_strict rs). normalized_nnf_match (get_match r)" unfolding inner_outer apply simp apply(rule cut_off_after_match_any_preserve_matches) . qed theorem transform_optimize_dnf_strict: assumes simplers: "simple_ruleset rs" and wf\: "wf_unknown_match_tac \" shows "(common_matcher, \),p\ \transform_optimize_dnf_strict rs, s\ \\<^sub>\ t \ (common_matcher, \),p\ \rs, s\ \\<^sub>\ t" proof - let ?\="(common_matcher, \)" let ?fw="\rs. approximating_bigstep_fun ?\ p rs s" have simplers_transform: "simple_ruleset (transform_optimize_dnf_strict rs)" unfolding transform_optimize_dnf_strict_def using simplers by (simp add: cut_off_after_match_any_simplers optimize_matches_simple_ruleset simple_ruleset_normalize_rules_dnf) have simplers1: "simple_ruleset (optimize_matches (opt_MatchAny_match_expr \ optimize_primitive_univ) rs)" using simplers optimize_matches_simple_ruleset by (metis) have 1: "?\,p\ \rs, s\ \\<^sub>\ t \ ?fw rs = t" using approximating_semantics_iff_fun_good_ruleset[OF simple_imp_good_ruleset[OF simplers]] by fast have "?fw rs = ?fw (optimize_matches (opt_MatchAny_match_expr \ optimize_primitive_univ) rs)" apply(rule optimize_matches[symmetric]) using optimize_primitive_univ_correct_matchexpr opt_MatchAny_match_expr_correct by (metis comp_apply) also have "\ = ?fw (normalize_rules_dnf (optimize_matches (opt_MatchAny_match_expr \ optimize_primitive_univ) rs))" apply(rule normalize_rules_dnf_correct[symmetric]) using simplers1 by (metis good_imp_wf_ruleset simple_imp_good_ruleset) also have "\ = ?fw (optimize_matches opt_MatchAny_match_expr (normalize_rules_dnf (optimize_matches (opt_MatchAny_match_expr \ optimize_primitive_univ) rs)))" apply(rule optimize_matches[symmetric]) using opt_MatchAny_match_expr_correct by (metis) finally have rs: "?fw rs = ?fw (transform_optimize_dnf_strict rs)" unfolding transform_optimize_dnf_strict_def by(simp add: cut_off_after_match_any) have 2: "?fw (transform_optimize_dnf_strict rs) = t \ ?\,p\ \transform_optimize_dnf_strict rs, s\ \\<^sub>\ t " using approximating_semantics_iff_fun_good_ruleset[OF simple_imp_good_ruleset[OF simplers_transform], symmetric] by fast from 1 2 rs show "?\,p\ \transform_optimize_dnf_strict rs, s\ \\<^sub>\ t \ ?\,p\ \rs, s\ \\<^sub>\ t" by simp qed subsection\Abstracting over unknowns\ definition transform_remove_unknowns_generic :: "('a, 'packet) match_tac \ 'a rule list \ 'a rule list" where "transform_remove_unknowns_generic \ = optimize_matches_a (remove_unknowns_generic \) " theorem transform_remove_unknowns_generic: assumes simplers: "simple_ruleset rs" and wf\: "wf_unknown_match_tac \" and packet_independent_\: "packet_independent_\ \" and wf\: "packet_independent_\_unknown \" shows "(\, \),p\ \transform_remove_unknowns_generic (\, \) rs, s\ \\<^sub>\ t \ (\, \),p\ \rs, s\ \\<^sub>\ t" and "simple_ruleset (transform_remove_unknowns_generic (\, \) rs)" and "\ r \ set rs. \ has_disc disc (get_match r) \ \ r \ set (transform_remove_unknowns_generic (\, \) rs). \ has_disc disc (get_match r)" and "\ r \ set (transform_remove_unknowns_generic (\, \) rs). \ has_unknowns \ (get_match r)" (*may return MatchNot MatchAny*) and "\ r \ set rs. normalized_n_primitive disc_sel f (get_match r) \ \ r \ set (transform_remove_unknowns_generic (\, \) rs). normalized_n_primitive disc_sel f (get_match r)" and "\ r \ set rs. \ has_disc_negated disc neg (get_match r) \ \ r \ set (transform_remove_unknowns_generic (\, \) rs). \ has_disc_negated disc neg (get_match r)" proof - let ?\="(\, \)" let ?fw="\rs. approximating_bigstep_fun ?\ p rs s" show simplers1: "simple_ruleset (transform_remove_unknowns_generic ?\ rs)" unfolding transform_remove_unknowns_generic_def using simplers optimize_matches_a_simple_ruleset by blast show "?\,p\ \transform_remove_unknowns_generic ?\ rs, s\ \\<^sub>\ t \ ?\,p\ \rs, s\ \\<^sub>\ t" unfolding approximating_semantics_iff_fun_good_ruleset[OF simple_imp_good_ruleset[OF simplers1]] unfolding approximating_semantics_iff_fun_good_ruleset[OF simple_imp_good_ruleset[OF simplers]] unfolding transform_remove_unknowns_generic_def using optimize_matches_a_simplers[OF simplers] remove_unknowns_generic by metis from remove_unknowns_generic_not_has_disc show "\ r \ set rs. \ has_disc disc (get_match r) \ \ r \ set (transform_remove_unknowns_generic ?\ rs). \ has_disc disc (get_match r)" unfolding transform_remove_unknowns_generic_def by(intro optimize_matches_a_preserves) blast from remove_unknowns_generic_normalized_n_primitive show "\ r \ set rs. normalized_n_primitive disc_sel f (get_match r) \ \ r \ set (transform_remove_unknowns_generic ?\ rs). normalized_n_primitive disc_sel f (get_match r)" unfolding transform_remove_unknowns_generic_def by(intro optimize_matches_a_preserves) blast show "\ r \ set (transform_remove_unknowns_generic ?\ rs). \ has_unknowns \ (get_match r)" unfolding transform_remove_unknowns_generic_def apply(rule optimize_matches_a_preserves) apply(rule remove_unknowns_generic_specification[OF _ packet_independent_\ wf\]) using simplers by(simp add: simple_ruleset_def) from remove_unknowns_generic_not_has_disc_negated show "\ r \ set rs. \ has_disc_negated disc neg (get_match r) \ \ r \ set (transform_remove_unknowns_generic ?\ rs). \ has_disc_negated disc neg (get_match r)" unfolding transform_remove_unknowns_generic_def by(rule optimize_matches_a_preserves) blast qed thm transform_remove_unknowns_generic[OF _ _ _ packet_independent_\_unknown_common_matcher] corollary transform_remove_unknowns_upper: defines "upper \ optimize_matches_a upper_closure_matchexpr" assumes simplers: "simple_ruleset rs" shows "(common_matcher, in_doubt_allow),p\ \upper rs, s\ \\<^sub>\ t \ (common_matcher, in_doubt_allow),p\ \rs, s\ \\<^sub>\ t" and "simple_ruleset (upper rs)" and "\ r \ set rs. \ has_disc disc (get_match r) \ \ r \ set (upper rs). \ has_disc disc (get_match r)" and "\ r \ set (upper rs). \ has_disc is_Extra (get_match r)" and "\ r \ set rs. normalized_n_primitive disc_sel f (get_match r) \ \ r \ set (upper rs). normalized_n_primitive disc_sel f (get_match r)" and "\ r \ set rs. \ has_disc_negated disc neg (get_match r) \ \ r \ set (upper rs). \ has_disc_negated disc neg (get_match r)" proof - from simplers have upper: "upper rs = transform_remove_unknowns_generic (common_matcher, in_doubt_allow) rs" apply(simp add: transform_remove_unknowns_generic_def upper_def) apply(erule optimize_matches_a_simple_ruleset_eq) by (simp add: upper_closure_matchexpr_generic) note * = transform_remove_unknowns_generic[OF simplers wf_in_doubt_allow packet_independent_unknown_match_tacs(1) packet_independent_\_unknown_common_matcher, simplified upper_closure_matchexpr_generic] from *(1)[where p = p] show "(common_matcher, in_doubt_allow),p\ \upper rs, s\ \\<^sub>\ t \ (common_matcher, in_doubt_allow),p\ \rs, s\ \\<^sub>\ t" by(subst upper) from *(2) show "simple_ruleset (upper rs)" by(subst upper) from *(3) show "\ r \ set rs. \ has_disc disc (get_match r) \ \ r \ set (upper rs). \ has_disc disc (get_match r)" by(subst upper) fast from *(4) show "\ r \ set (upper rs). \ has_disc is_Extra (get_match r)" apply(subst upper) using has_unknowns_common_matcher by auto from *(5) show "\ r \ set rs. normalized_n_primitive disc_sel f (get_match r) \ \ r \ set (upper rs). normalized_n_primitive disc_sel f (get_match r)" apply(subst upper) using packet_independent_unknown_match_tacs(1) simplers transform_remove_unknowns_generic(5)[OF _ _ _ packet_independent_\_unknown_common_matcher] wf_in_doubt_allow by blast from *(6) show "\ r \ set rs. \ has_disc_negated disc neg (get_match r) \ \ r \ set (upper rs). \ has_disc_negated disc neg (get_match r)" by(subst upper) fast qed (*copy&paste from transform_remove_unknowns_upper*) corollary transform_remove_unknowns_lower: defines "lower \ optimize_matches_a lower_closure_matchexpr" assumes simplers: "simple_ruleset rs" shows "(common_matcher, in_doubt_deny),p\ \lower rs, s\ \\<^sub>\ t \ (common_matcher, in_doubt_deny),p\ \rs, s\ \\<^sub>\ t" and "simple_ruleset (lower rs)" and "\ r \ set rs. \ has_disc disc (get_match r) \ \ r \ set (lower rs). \ has_disc disc (get_match r)" and "\ r \ set (lower rs). \ has_disc is_Extra (get_match r)" and "\ r \ set rs. normalized_n_primitive disc_sel f (get_match r) \ \ r \ set (lower rs). normalized_n_primitive disc_sel f (get_match r)" and "\ r \ set rs. \ has_disc_negated disc neg (get_match r) \ \ r \ set (lower rs). \ has_disc_negated disc neg (get_match r)" proof - from simplers have lower: "lower rs = transform_remove_unknowns_generic (common_matcher, in_doubt_deny) rs" apply(simp add: transform_remove_unknowns_generic_def lower_def) apply(erule optimize_matches_a_simple_ruleset_eq) by (simp add: lower_closure_matchexpr_generic) note * = transform_remove_unknowns_generic[OF simplers wf_in_doubt_deny packet_independent_unknown_match_tacs(2) packet_independent_\_unknown_common_matcher, simplified lower_closure_matchexpr_generic] from *(1)[where p = p] show "(common_matcher, in_doubt_deny),p\ \lower rs, s\ \\<^sub>\ t \ (common_matcher, in_doubt_deny),p\ \rs, s\ \\<^sub>\ t" by(subst lower) from *(2) show "simple_ruleset (lower rs)" by(subst lower) from *(3) show "\ r \ set rs. \ has_disc disc (get_match r) \ \ r \ set (lower rs). \ has_disc disc (get_match r)" by(subst lower) fast from *(4) show "\ r \ set (lower rs). \ has_disc is_Extra (get_match r)" apply(subst lower) using has_unknowns_common_matcher by auto from *(5) show "\ r \ set rs. normalized_n_primitive disc_sel f (get_match r) \ \ r \ set (lower rs). normalized_n_primitive disc_sel f (get_match r)" apply(subst lower) using packet_independent_unknown_match_tacs(1) simplers transform_remove_unknowns_generic(5)[OF _ _ _ packet_independent_\_unknown_common_matcher] wf_in_doubt_deny by blast from *(6) show "\ r \ set rs. \ has_disc_negated disc neg (get_match r) \ \ r \ set (lower rs). \ has_disc_negated disc neg (get_match r)" by(subst lower) fast qed subsection\Normalizing and Transforming Primitives\ text\Rewrite the primitives IPs and Ports such that can be used by the simple firewall.\ definition transform_normalize_primitives :: "'i::len common_primitive rule list \ 'i common_primitive rule list" where "transform_normalize_primitives = optimize_matches_option compress_normalize_besteffort \ \ \normalizes protocols, needs to go last\ normalize_rules normalize_dst_ips \ normalize_rules normalize_src_ips \ normalize_rules normalize_dst_ports \ \ \may introduce new matches on protocols\ normalize_rules normalize_src_ports \ \ \may introduce new matches in protocols\ normalize_rules rewrite_MultiportPorts \ \introduces \Src_Ports\ and \Dst_Ports\ matches\" thm normalize_primitive_extract_preserves_unrelated_normalized_n_primitive lemma normalize_rules_preserves_unrelated_normalized_n_primitive: assumes "\ r \ set rs. normalized_nnf_match (get_match r) \ normalized_n_primitive (disc2, sel2) P (get_match r)" and "wf_disc_sel (disc1, sel1) C" and "\a. \ disc2 (C a)" shows "\r \ set (normalize_rules (normalize_primitive_extract (disc1, sel1) C f) rs). normalized_nnf_match (get_match r) \ normalized_n_primitive (disc2, sel2) P (get_match r)" thm normalize_rules_preserves[where P="\m. normalized_nnf_match m \ normalized_n_primitive (disc2, sel2) P m" and f="normalize_primitive_extract (disc1, sel1) C f"] apply(rule normalize_rules_preserves[where P="\m. normalized_nnf_match m \ normalized_n_primitive (disc2, sel2) P m" and f="normalize_primitive_extract (disc1, sel1) C f"]) using assms(1) apply(simp) apply(safe) using normalize_primitive_extract_preserves_nnf_normalized[OF _ assms(2)] apply fast using normalize_primitive_extract_preserves_unrelated_normalized_n_primitive[OF _ _ assms(2) assms(3)] by blast lemma normalize_rules_normalized_n_primitive: assumes "\ r \ set rs. normalized_nnf_match (get_match r)" and "\m. normalized_nnf_match m \ (\m' \ set (normalize_primitive_extract (disc, sel) C f m). normalized_n_primitive (disc, sel) P m')" shows "\r \ set (normalize_rules (normalize_primitive_extract (disc, sel) C f) rs). normalized_n_primitive (disc, sel) P (get_match r)" apply(rule normalize_rules_property[where P=normalized_nnf_match and f="normalize_primitive_extract (disc, sel) C f"]) using assms(1) apply simp using assms(2) by simp lemma optimize_matches_option_compress_normalize_besteffort_preserves_unrelated_normalized_n_primitive: assumes "\ r \ set rs. normalized_nnf_match (get_match r) \ normalized_n_primitive (disc2, sel2) P (get_match r)" and "\a. \ disc2 (IIface a)" and "\a. \ disc2 (OIface a)" and "\a. \ disc2 (Prot a)" shows "\r \ set (optimize_matches_option compress_normalize_besteffort rs). normalized_nnf_match (get_match r) \ normalized_n_primitive (disc2, sel2) P (get_match r)" thm optimize_matches_option_preserves apply(rule optimize_matches_option_preserves[where P="\m. normalized_nnf_match m \ normalized_n_primitive (disc2, sel2) P m" and f="compress_normalize_besteffort"]) apply(rule compress_normalize_besteffort_preserves_normalized_n_primitive) apply(simp_all add: assms) done (*We write (\a. \ disc (Src_Ports a)) to say that, basically, disc is not the function is_Src_Ports. But hey, equality on functions, ....*) theorem transform_normalize_primitives: \ \all discriminators which will not be normalized remain unchanged\ defines "unchanged disc \ (\a. \ disc (Src_Ports a)) \ (\a. \ disc (Dst_Ports a)) \ (\a. \ disc (Src a)) \ (\a. \ disc (Dst a))" \ \also holds for these discriminators, but not for @{const Prot}, which might be changed\ and "changeddisc disc \ ((\a. \ disc (IIface a)) \ disc = is_Iiface) \ ((\a. \ disc (OIface a)) \ disc = is_Oiface)" (*port normalization may introduce new matches on protocols*) assumes simplers: "simple_ruleset (rs :: 'i::len common_primitive rule list)" and wf\: "wf_unknown_match_tac \" and normalized: "\ r \ set rs. normalized_nnf_match (get_match r)" shows "(common_matcher, \),p\ \transform_normalize_primitives rs, s\ \\<^sub>\ t \ (common_matcher, \),p\ \rs, s\ \\<^sub>\ t" and "simple_ruleset (transform_normalize_primitives rs)" and "unchanged disc1 \ changeddisc disc1 \ \a. \ disc1 (Prot a) \ \ r \ set rs. \ has_disc disc1 (get_match r) \ \ r \ set (transform_normalize_primitives rs). \ has_disc disc1 (get_match r)" and "\ r \ set (transform_normalize_primitives rs). normalized_nnf_match (get_match r)" and "\ r \ set (transform_normalize_primitives rs). normalized_src_ports (get_match r) \ normalized_dst_ports (get_match r) \ normalized_src_ips (get_match r) \ normalized_dst_ips (get_match r) \ \ has_disc is_MultiportPorts (get_match r)" and "unchanged disc2 \ (\a. \ disc2 (IIface a)) \ (\a. \ disc2 (OIface a)) \ (\a. \ disc2 (Prot a)) \ \ r \ set rs. normalized_n_primitive (disc2, sel2) f (get_match r) \ \ r \ set (transform_normalize_primitives rs). normalized_n_primitive (disc2, sel2) f (get_match r)" \ \For disc3, we do not allow ports and ips, because these are changed. Here is the complicated part: (It is only complicated if, basically disc3 is @{const is_Prot}) In addition, either it must not be protocol or (complicated case) there must be no negated port matches in the ruleset. Note that negated @{const Src_Ports} or @{const Dst_Ports} can also be introduced by rewriting @{const MultiportPorts}\ and "unchanged disc3 \ changeddisc disc3 \ (\a. \ disc3 (Prot a)) \ (disc3 = is_Prot \ (\ r \ set rs. \ has_disc_negated is_Src_Ports False (get_match r) \ \ has_disc_negated is_Dst_Ports False (get_match r) \ \ has_disc is_MultiportPorts (get_match r))) \ \ r \ set rs. \ has_disc_negated disc3 False (get_match r) \ \ r \ set (transform_normalize_primitives rs). \ has_disc_negated disc3 False (get_match r)" proof - let ?\="(common_matcher, \)" let ?fw="\rs. approximating_bigstep_fun ?\ p rs s" show simplers_t: "simple_ruleset (transform_normalize_primitives rs)" unfolding transform_normalize_primitives_def by(simp add: simple_ruleset_normalize_rules simplers optimize_matches_option_simple_ruleset) let ?rs0="normalize_rules rewrite_MultiportPorts rs" let ?rs1="normalize_rules normalize_src_ports ?rs0" let ?rs2="normalize_rules normalize_dst_ports ?rs1" let ?rs3="normalize_rules normalize_src_ips ?rs2" let ?rs4="normalize_rules normalize_dst_ips ?rs3" let ?rs5="optimize_matches_option compress_normalize_besteffort ?rs4" have normalized_rs0: "\r \ set ?rs0. normalized_nnf_match (get_match r)" apply(intro normalize_rules_preserves[OF normalized]) apply(simp add: rewrite_MultiportPorts_def) using normalized_nnf_match_normalize_match by blast from normalize_src_ports_nnf have normalized_rs1: "\r \ set ?rs1. normalized_nnf_match (get_match r)" apply(intro normalize_rules_preserves[OF normalized_rs0]) using normalize_dst_ports_nnf by blast from normalize_dst_ports_nnf have normalized_rs2: "\r \ set ?rs2. normalized_nnf_match (get_match r)" apply(intro normalize_rules_preserves[OF normalized_rs1]) by blast from normalize_rules_primitive_extract_preserves_nnf_normalized[OF this wf_disc_sel_common_primitive(3)] normalize_src_ips_def have normalized_rs3: "\r \ set ?rs3. normalized_nnf_match (get_match r)" by metis from normalize_rules_primitive_extract_preserves_nnf_normalized[OF this wf_disc_sel_common_primitive(4)] normalize_dst_ips_def have normalized_rs4: "\r \ set ?rs4. normalized_nnf_match (get_match r)" by metis have normalized_rs5: "\r \ set ?rs5. normalized_nnf_match (get_match r)" apply(intro optimize_matches_option_preserves) apply(erule compress_normalize_besteffort_nnf[rotated]) by(simp add: normalized_rs4) thus "\ r \ set (transform_normalize_primitives rs). normalized_nnf_match (get_match r)" unfolding transform_normalize_primitives_def by simp (*add this as generic simp rule somewhere? But simplifier loops? what to do? cong_rule?*) have local_simp: "\rs1 rs2. approximating_bigstep_fun ?\ p rs1 s = approximating_bigstep_fun ?\ p rs2 s \ (approximating_bigstep_fun ?\ p rs1 s = t) = (approximating_bigstep_fun ?\ p rs2 s = t)" by simp have opt_compress_rule: "approximating_bigstep_fun (common_matcher, \) p (optimize_matches_option compress_normalize_besteffort rs1) s = approximating_bigstep_fun (common_matcher, \) p rs2 s" if rs1_n: "\r \ set rs1. normalized_nnf_match (get_match r)" and rs1rs2: "approximating_bigstep_fun (common_matcher, \) p rs1 s = approximating_bigstep_fun (common_matcher, \) p rs2 s" for rs1 rs2 apply(subst optimize_matches_option_generic[where P="\ m a. normalized_nnf_match m"]) apply(simp_all add: normalized compress_normalize_besteffort_Some[OF primitive_matcher_generic_common_matcher] compress_normalize_besteffort_None[OF primitive_matcher_generic_common_matcher] rs1_n) using rs1rs2 by simp show "?\,p\ \transform_normalize_primitives rs, s\ \\<^sub>\ t \ ?\,p\ \rs, s\ \\<^sub>\ t" unfolding approximating_semantics_iff_fun_good_ruleset[OF simple_imp_good_ruleset[OF simplers_t]] unfolding approximating_semantics_iff_fun_good_ruleset[OF simple_imp_good_ruleset[OF simplers]] unfolding transform_normalize_primitives_def apply(simp) apply(subst local_simp, simp_all) apply(rule opt_compress_rule[OF normalized_rs4]) apply(subst normalize_rules_match_list_semantics_3[of normalized_nnf_match]) using normalize_dst_ips[where p = p] apply(simp; fail) using simplers simple_ruleset_normalize_rules apply blast using normalized_rs3 apply(simp; fail) apply(subst normalize_rules_match_list_semantics_3[of normalized_nnf_match]) using normalize_src_ips[where p = p] apply(simp; fail) using simplers simple_ruleset_normalize_rules apply blast using normalized_rs2 apply(simp; fail) apply(subst normalize_rules_match_list_semantics_3[of normalized_nnf_match]) using normalize_dst_ports[OF primitive_matcher_generic_common_matcher,where p = p] apply(simp; fail) using simplers simple_ruleset_normalize_rules apply blast using normalized_rs1 apply(simp; fail) apply(subst normalize_rules_match_list_semantics_3[of normalized_nnf_match]) using normalize_src_ports[OF primitive_matcher_generic_common_matcher, where p = p] apply(simp; fail) using simplers simple_ruleset_normalize_rules apply blast using normalized_rs0 apply(simp; fail) apply(subst normalize_rules_match_list_semantics_3[of normalized_nnf_match]) using rewrite_MultiportPorts[OF primitive_matcher_generic_common_matcher, where p = p] apply(simp; fail) using simplers apply blast using normalized apply(simp; fail) .. (*naming: does not "normalize" but eliminate all multiportPorts!*) from rewrite_MultiportPorts_removes_MultiportsPorts normalize_rules_property[OF normalized, where f=rewrite_MultiportPorts and Q="\m. \ has_disc is_MultiportPorts m"] have rewrite_MultiportPorts_normalizes_Multiports: "\r \ set ?rs0. \ has_disc is_MultiportPorts (get_match r)" by blast from normalize_src_ports_normalized_n_primitive have normalized_src_ports: "\r \ set ?rs1. normalized_src_ports (get_match r)" apply(intro normalize_rules_property[OF normalized_rs0, where f=normalize_src_ports and Q=normalized_src_ports]) by blast from normalize_dst_ports_normalized_n_primitive normalize_rules_property[OF normalized_rs1, where f=normalize_dst_ports and Q=normalized_dst_ports] have normalized_dst_ports: "\r \ set ?rs2. normalized_dst_ports (get_match r)" by fast from normalize_src_ips_normalized_n_primitive normalize_rules_property[OF normalized_rs2, where f=normalize_src_ips and Q=normalized_src_ips] have normalized_src_ips: "\r \ set ?rs3. normalized_src_ips (get_match r)" by fast from normalize_dst_ips_normalized_n_primitive normalize_rules_property[OF normalized_rs3, where f=normalize_dst_ips and Q=normalized_dst_ips] normalized_rs4 have normalized_dst_ips_rs4: "\r \ set ?rs4. normalized_nnf_match (get_match r) \ normalized_dst_ips (get_match r)" by fast with optimize_matches_option_compress_normalize_besteffort_preserves_unrelated_normalized_n_primitive[ of _ is_Dst dst_sel normalized_cidr_ip , folded normalized_dst_ips_def2] have normalized_dst_rs5: "\r \ set ?rs5. normalized_dst_ips (get_match r)" by fastforce have normalize_dst_ports_preserves_normalized_src_ports: "m' \ set (normalize_dst_ports m) \ normalized_nnf_match m \ normalized_src_ports m \ normalized_src_ports m'" for m m' :: " 'i common_primitive match_expr" unfolding normalized_src_ports_def2 apply(rule normalize_ports_generic_preserves_normalized_n_primitive[OF _ wf_disc_sel_common_primitive(2)]) apply(simp_all) by (simp add: normalize_dst_ports_def normalize_ports_generic_def normalize_positive_dst_ports_def rewrite_negated_dst_ports_def) from normalize_rules_preserves_unrelated_normalized_n_primitive[of _ is_MultiportPorts multiportports_sel "\_. False"] have preserve_normalized_multiport_ports: " \r\ set rs. normalized_nnf_match (get_match r) \ \r\ set rs. \ has_disc is_MultiportPorts (get_match r) \ wf_disc_sel (disc, sel) C \ \a. \ is_MultiportPorts (C a) \ \r\ set (normalize_rules (normalize_primitive_extract (disc, sel) C f) rs). \ has_disc is_MultiportPorts (get_match r)" for f :: "'c negation_type list \ 'c list" and rs disc sel and C :: "'c \ 'i::len common_primitive" using normalized_n_primitive_false_eq_notdisc by blast (*TODO: push through*) have normalized_multiportports_rs1: "\r \ set ?rs1. \ has_disc is_MultiportPorts (get_match r)" apply(rule normalize_rules_property[where P="\m. normalized_nnf_match m \ \ has_disc is_MultiportPorts m"]) using normalized_rs0 rewrite_MultiportPorts_normalizes_Multiports apply blast apply(intro allI impI ballI) apply(rule normalize_src_ports_preserves_normalized_not_has_disc) by(simp_all) have normalized_multiportports_rs2: "\r \ set ?rs2. \ has_disc is_MultiportPorts (get_match r)" apply(rule normalize_rules_property[where P="\m. normalized_nnf_match m \ \ has_disc is_MultiportPorts m"]) using normalized_rs1 normalized_multiportports_rs1 apply blast apply(intro allI impI ballI) apply(rule normalize_dst_ports_preserves_normalized_not_has_disc) by(simp_all) from preserve_normalized_multiport_ports[OF normalized_rs2 normalized_multiportports_rs2 wf_disc_sel_common_primitive(3), where f2=ipt_iprange_compress, folded normalize_src_ips_def] have normalized_multiportports_rs3: "\r \ set ?rs3. \ has_disc is_MultiportPorts (get_match r)" by simp from preserve_normalized_multiport_ports[OF normalized_rs3 normalized_multiportports_rs3 wf_disc_sel_common_primitive(4), where f2=ipt_iprange_compress, folded normalize_dst_ips_def] normalized_rs4 have normalized_multiportports_rs4: "\r \ set ?rs4. normalized_nnf_match (get_match r) \ \ has_disc is_MultiportPorts (get_match r)" by simp with optimize_matches_option_compress_normalize_besteffort_preserves_unrelated_normalized_n_primitive[ of _ is_MultiportPorts multiportports_sel "\_. False" , simplified] have normalized_multiportports_rs5: "\r \ set ?rs5. \ has_disc is_MultiportPorts (get_match r)" using normalized_n_primitive_false_eq_notdisc by fastforce from normalize_rules_preserves_unrelated_normalized_n_primitive[of _ is_Src_Ports src_ports_sel "(\ps. case ps of L4Ports _ pts \ length pts \ 1)", folded normalized_src_ports_def2] have preserve_normalized_src_ports: " \r\ set rs. normalized_nnf_match (get_match r) \ \r\ set rs. normalized_src_ports (get_match r) \ wf_disc_sel (disc, sel) C \ \a. \ is_Src_Ports (C a) \ \r\ set (normalize_rules (normalize_primitive_extract (disc, sel) C f) rs). normalized_src_ports (get_match r)" for f :: "'c negation_type list \ 'c list" and rs disc sel and C :: "'c \ 'i::len common_primitive" by blast have normalized_src_ports_rs2: "\r \ set ?rs2. normalized_src_ports (get_match r)" apply(rule normalize_rules_property[where P="\m. normalized_nnf_match m \ normalized_src_ports m"]) using normalized_rs1 normalized_src_ports apply blast apply(clarify) using normalize_dst_ports_preserves_normalized_src_ports by blast from preserve_normalized_src_ports[OF normalized_rs2 normalized_src_ports_rs2 wf_disc_sel_common_primitive(3), where f3=ipt_iprange_compress, folded normalize_src_ips_def] have normalized_src_ports_rs3: "\r \ set ?rs3. normalized_src_ports (get_match r)" by simp from preserve_normalized_src_ports[OF normalized_rs3 normalized_src_ports_rs3 wf_disc_sel_common_primitive(4), where f3=ipt_iprange_compress, folded normalize_dst_ips_def] normalized_rs4 have normalized_src_ports_rs4: "\r \ set ?rs4. normalized_nnf_match (get_match r) \ normalized_src_ports (get_match r)" by simp with optimize_matches_option_compress_normalize_besteffort_preserves_unrelated_normalized_n_primitive[ of _ is_Src_Ports src_ports_sel "(\ps. case ps of L4Ports _ pts \ length pts \ 1)" , folded normalized_src_ports_def2] have normalized_src_ports_rs5: "\r \ set ?rs5. normalized_src_ports (get_match r)" by fastforce from normalize_rules_preserves_unrelated_normalized_n_primitive[of _ is_Dst_Ports dst_ports_sel "(\ps. case ps of L4Ports _ pts \ length pts \ 1)", folded normalized_dst_ports_def2] have preserve_normalized_dst_ports: "\rs disc sel C f. \r\set rs. normalized_nnf_match (get_match r) \ \r\set rs. normalized_dst_ports (get_match r) \ wf_disc_sel (disc, sel) C \ \a. \ is_Dst_Ports (C a) \ \r\ set (normalize_rules (normalize_primitive_extract (disc, sel) C f) rs). normalized_dst_ports (get_match r)" by blast from preserve_normalized_dst_ports[OF normalized_rs2 normalized_dst_ports wf_disc_sel_common_primitive(3), - where f=ipt_iprange_compress, folded normalize_src_ips_def] + where f3=ipt_iprange_compress, folded normalize_src_ips_def] have normalized_dst_ports_rs3: "\r \ set ?rs3. normalized_dst_ports (get_match r)" by force from preserve_normalized_dst_ports[OF normalized_rs3 normalized_dst_ports_rs3 wf_disc_sel_common_primitive(4), - where f=ipt_iprange_compress, folded normalize_dst_ips_def] + where f3=ipt_iprange_compress, folded normalize_dst_ips_def] normalized_rs4 have normalized_dst_ports_rs4: "\r \ set ?rs4. normalized_nnf_match (get_match r) \ normalized_dst_ports (get_match r)" by force with optimize_matches_option_compress_normalize_besteffort_preserves_unrelated_normalized_n_primitive[ of _ is_Dst_Ports dst_ports_sel "(\ps. case ps of L4Ports _ pts \ length pts \ 1)" , folded normalized_dst_ports_def2] have normalized_dst_ports_rs5: "\r \ set ?rs5. normalized_dst_ports (get_match r)" by fastforce from normalize_rules_preserves_unrelated_normalized_n_primitive[of ?rs3 is_Src src_sel normalized_cidr_ip, OF _ wf_disc_sel_common_primitive(4), where f=ipt_iprange_compress, folded normalize_dst_ips_def normalized_src_ips_def2] normalized_rs3 normalized_src_ips have normalized_src_rs4: "\r \ set ?rs4. normalized_nnf_match (get_match r) \ normalized_src_ips (get_match r)" by force with optimize_matches_option_compress_normalize_besteffort_preserves_unrelated_normalized_n_primitive[ of _ is_Src src_sel normalized_cidr_ip , folded normalized_src_ips_def2] have normalized_src_rs5: "\r \ set ?rs5. normalized_src_ips (get_match r)" by fastforce from normalized_multiportports_rs5 normalized_src_ports_rs5 normalized_dst_ports_rs5 normalized_src_rs5 normalized_dst_rs5 show "\ r \ set (transform_normalize_primitives rs). normalized_src_ports (get_match r) \ normalized_dst_ports (get_match r) \ normalized_src_ips (get_match r) \ normalized_dst_ips (get_match r) \ \ has_disc is_MultiportPorts (get_match r)" unfolding transform_normalize_primitives_def by simp show "unchanged disc2 \ (\a. \ disc2 (IIface a)) \ (\a. \ disc2 (OIface a)) \ (\a. \ disc2 (Prot a)) \ \ r \ set rs. normalized_n_primitive (disc2, sel2) f (get_match r) \ \ r \ set (transform_normalize_primitives rs). normalized_n_primitive (disc2, sel2) f (get_match r)" unfolding unchanged_def proof(elim conjE) assume "\r\ set rs. normalized_n_primitive (disc2, sel2) f (get_match r)" with normalized have a': "\r\ set rs. normalized_nnf_match (get_match r) \ normalized_n_primitive (disc2, sel2) f (get_match r)" by blast assume a_Src_Ports: "\a. \ disc2 (Src_Ports a)" assume a_Dst_Ports: "\a. \ disc2 (Dst_Ports a)" assume a_Src: "\a. \ disc2 (Src a)" assume a_Dst: "\a. \ disc2 (Dst a)" assume a_IIface: "(\a. \ disc2 (IIface a))" assume a_OIface: "(\a. \ disc2 (OIface a))" assume a_Prot: "(\a. \ disc2 (Prot a))" have normalized_n_primitive_rs0: "\r\set ?rs0. normalized_n_primitive (disc2, sel2) f (get_match r)" apply(intro normalize_rules_property[where P="\m. normalized_nnf_match m \ normalized_n_primitive (disc2, sel2) f m"]) using a' apply blast using rewrite_MultiportPorts_preserves_normalized_n_primitive[OF _ a_Src_Ports a_Dst_Ports] by blast have normalized_n_primitive_rs1: "\r\set ?rs1. normalized_n_primitive (disc2, sel2) f (get_match r)" (*by blast*) apply(rule normalize_rules_property[where P="\m. normalized_nnf_match m \ normalized_n_primitive (disc2, sel2) f m"]) using normalized_n_primitive_rs0 normalized_rs0 apply blast using normalize_src_ports_preserves_normalized_n_primitive[OF _ a_Src_Ports] a_Prot by blast have "\r\set ?rs2. normalized_n_primitive (disc2, sel2) f (get_match r)" apply(rule normalize_rules_property[where P="\m. normalized_nnf_match m \ normalized_n_primitive (disc2, sel2) f m"]) using normalized_n_primitive_rs1 normalized_rs1 apply blast using normalize_dst_ports_preserves_normalized_n_primitive[OF _ a_Dst_Ports] a_Prot by blast with normalized_rs2 normalize_rules_preserves_unrelated_normalized_n_primitive[OF _ wf_disc_sel_common_primitive(3) a_Src, of ?rs2 sel2 f ipt_iprange_compress, folded normalize_src_ips_def] have "\r\set ?rs3. normalized_n_primitive (disc2, sel2) f (get_match r)" by blast with normalized_rs3 normalize_rules_preserves_unrelated_normalized_n_primitive[OF _ wf_disc_sel_common_primitive(4) a_Dst, of ?rs3 sel2 f ipt_iprange_compress, folded normalize_dst_ips_def] have "\r\set ?rs4. normalized_nnf_match (get_match r) \ normalized_n_primitive (disc2, sel2) f (get_match r)" by blast hence "\r\set ?rs5. normalized_nnf_match (get_match r) \ normalized_n_primitive (disc2, sel2) f (get_match r)" apply(intro optimize_matches_option_compress_normalize_besteffort_preserves_unrelated_normalized_n_primitive) using a_IIface a_OIface a_Prot by simp_all thus ?thesis unfolding transform_normalize_primitives_def by simp qed \ \Pushing through properties through the ip normalizer\ { fix m and m' and disc::"('i::len common_primitive \ bool)" and sel::"('i::len common_primitive \ 'x)" and C'::" ('x \ 'i::len common_primitive)" and f'::"('x negation_type list \ 'x list)" assume am: "\ has_disc disc1 m" and nm: "normalized_nnf_match m" and am': "m' \ set (normalize_primitive_extract (disc, sel) C' f' m)" and wfdiscsel: "wf_disc_sel (disc,sel) C'" and disc_different: "\a. \ disc1 (C' a)" (*from wfdiscsel disc_different have "\a. \ disc1 (C' a)" apply(simp add: wf_disc_sel.simps)*) from disc_different have af: "\spts. (\a \ Match ` C' ` set (f' spts). \ has_disc disc1 a)" by(simp) obtain as ms where asms: "primitive_extractor (disc, sel) m = (as, ms)" by fastforce from am' asms have "m' \ (\spt. MatchAnd (Match (C' spt)) ms) ` set (f' as)" unfolding normalize_primitive_extract_def by(simp) hence goalrule:"\spt \ set (f' as). \ has_disc disc1 (Match (C' spt)) \ \ has_disc disc1 ms \ \ has_disc disc1 m'" by fastforce from am primitive_extractor_correct(4)[OF nm wfdiscsel asms] have 1: "\ has_disc disc1 ms" by simp from af have 2: "\spt \ set (f' as). \ has_disc disc1 (Match (C' spt))" by simp from goalrule[OF 2 1] have "\ has_disc disc1 m'" . moreover from nm have "normalized_nnf_match m'" by (metis am' normalize_primitive_extract_preserves_nnf_normalized wfdiscsel) ultimately have "\ has_disc disc1 m' \ normalized_nnf_match m'" by simp } hence x: "\disc sel C' f'. wf_disc_sel (disc, sel) C' \ \a. \ disc1 (C' a) \ \m. normalized_nnf_match m \ \ has_disc disc1 m \ (\m'\set (normalize_primitive_extract (disc, sel) C' f' m). normalized_nnf_match m' \ \ has_disc disc1 m')" by blast \ \Pushing through properties through the ports normalizer\ from normalize_src_ports_preserves_normalized_not_has_disc normalize_src_ports_nnf have x_src_ports: "\a. \ disc (Src_Ports a) \ \a. \ disc (Prot a) \ m' \ set (normalize_src_ports m) \ normalized_nnf_match m \ \ has_disc disc m \ \ has_disc disc m' \ normalized_nnf_match m'" for m m' and disc :: "'i common_primitive \ bool" by blast from normalize_dst_ports_preserves_normalized_not_has_disc normalize_dst_ports_nnf have x_dst_ports: "\a. \ disc (Dst_Ports a) \ \a. \ disc (Prot a) \ m'\ set (normalize_dst_ports m) \ normalized_nnf_match m \ \ has_disc disc m \ \ has_disc disc m' \ normalized_nnf_match m'" for m m' and disc :: "'i common_primitive \ bool" by blast { fix rs assume "(\a. \ disc1 (IIface a)) \ disc1 = is_Iiface" and "((\a. \ disc1 (OIface a)) \ disc1 = is_Oiface)" and "(\a. \ disc1 (Prot a))" hence "\m\set rs. \ has_disc disc1 (get_match m) \ normalized_nnf_match (get_match m) \ \m\set (optimize_matches_option compress_normalize_besteffort rs). normalized_nnf_match (get_match m) \ \ has_disc disc1 (get_match m)" apply - apply(rule optimize_matches_option_preserves) apply(elim disjE) using compress_normalize_besteffort_hasdisc apply blast using compress_normalize_besteffort_nnf compress_normalize_besteffort_not_introduces_Iiface compress_normalize_besteffort_not_introduces_Oiface by blast+ } note y=this have "\a. \ disc1 (Src_Ports a) \ \a. \ disc1 (Dst_Ports a) \ \a. \ disc1 (Src a) \ \a. \ disc1 (Dst a) \ (\a. \ disc1 (IIface a)) \ disc1 = is_Iiface \ (\a. \ disc1 (OIface a)) \ disc1 = is_Oiface \ (\a. \ disc1 (Prot a)) \ \ r\set rs. \ has_disc disc1 (get_match r) \ normalized_nnf_match (get_match r) \ \ r \ set (transform_normalize_primitives rs). normalized_nnf_match (get_match r) \ \ has_disc disc1 (get_match r)" unfolding transform_normalize_primitives_def apply(simp) apply(rule y) apply(simp; fail) apply(simp; fail) apply(simp; fail) apply(rule normalize_rules_preserves)+ apply(simp; fail) subgoal apply(intro allI impI conjI ballI) apply(rule rewrite_MultiportPorts_preserves_normalized_not_has_disc, simp_all) by(simp add: rewrite_MultiportPorts_normalized_nnf_match) subgoal apply clarify apply(rule x_src_ports) by simp+ subgoal apply clarify by(rule x_dst_ports) simp+ using x[OF wf_disc_sel_common_primitive(3), of ipt_iprange_compress,folded normalize_src_ips_def] apply blast using x[OF wf_disc_sel_common_primitive(4), of ipt_iprange_compress,folded normalize_dst_ips_def] apply blast done thus "unchanged disc1 \ changeddisc disc1 \ \a. \ disc1 (Prot a) \ \ r \ set rs. \ has_disc disc1 (get_match r) \ \ r \ set (transform_normalize_primitives rs). \ has_disc disc1 (get_match r)" unfolding unchanged_def changeddisc_def using normalized by blast (*copy pasta*) { fix m and m' and disc::"('i::len common_primitive \ bool)" and sel::"('i::len common_primitive \ 'x)" and C'::" ('x \ 'i::len common_primitive)" and f'::"('x negation_type list \ 'x list)" and neg and disc3 assume am: "\ has_disc_negated disc3 neg m" and nm: "normalized_nnf_match m" and am': "m' \ set (normalize_primitive_extract (disc, sel) C' f' m)" and wfdiscsel: "wf_disc_sel (disc,sel) C'" and disc_different: "\a. \ disc3 (C' a)" from disc_different have af: "\spts. (\a \ Match ` C' ` set (f' spts). \ has_disc disc3 a)" by(simp) obtain as ms where asms: "primitive_extractor (disc, sel) m = (as, ms)" by fastforce from am' asms have "m' \ (\spt. MatchAnd (Match (C' spt)) ms) ` set (f' as)" unfolding normalize_primitive_extract_def by(simp) hence goalrule:"\spt \ set (f' as). \ has_disc_negated disc3 neg (Match (C' spt)) \ \ has_disc_negated disc3 neg ms \ \ has_disc_negated disc3 neg m'" by fastforce from am primitive_extractor_correct(6)[OF nm wfdiscsel asms] have 1: "\ has_disc_negated disc3 neg ms" by simp from af have 2: "\spt \ set (f' as). \ has_disc_negated disc3 neg (Match (C' spt))" by simp from goalrule[OF 2 1] have "\ has_disc_negated disc3 neg m'" . moreover from nm have "normalized_nnf_match m'" by (metis am' normalize_primitive_extract_preserves_nnf_normalized wfdiscsel) ultimately have "\ has_disc_negated disc3 neg m' \ normalized_nnf_match m'" by simp } note x_generic=this hence x: "wf_disc_sel (disc, sel) C' \ \a. \ disc3 (C' a) \ \m. normalized_nnf_match m \ \ has_disc_negated disc3 False m \ (\m' \ set (normalize_primitive_extract (disc, sel) C' f' m). normalized_nnf_match m' \ \ has_disc_negated disc3 False m')" for disc :: "'i common_primitive \ bool" and sel and C' :: "'c \ 'i common_primitive" and f' and disc3 by blast \ \Pushing through properties through the ports normalizer\ from normalize_src_ports_preserves_normalized_not_has_disc_negated normalize_src_ports_nnf have x_src_ports: "\a. \ disc (Src_Ports a) \ (\a. \ disc (Prot a)) \ \ has_disc_negated is_Src_Ports False m \ m' \ set (normalize_src_ports m) \ normalized_nnf_match m \ \ has_disc_negated disc False m \ \ has_disc_negated disc False m' \ normalized_nnf_match m'" for m m' and disc :: "'i common_primitive \ bool" by blast from normalize_dst_ports_preserves_normalized_not_has_disc_negated normalize_dst_ports_nnf have x_dst_ports: "\a. \ disc (Src_Ports a) \ (\a. \ disc (Prot a)) \ \ has_disc_negated is_Dst_Ports False m \ m' \ set (normalize_dst_ports m) \ normalized_nnf_match m \ \ has_disc_negated disc False m \ \ has_disc_negated disc False m' \ normalized_nnf_match m'" for m m' and disc :: "'i common_primitive \ bool" by blast (*push arbitrary P m through too. y is then \_. True [simplified]*) { fix rs fix P :: "'i common_primitive match_expr \ bool" assume "(\a. \ disc3 (IIface a)) \ disc3 = is_Iiface" and "(\a. \ disc3 (OIface a)) \ disc3 = is_Oiface" and "(\a. \ disc3 (Prot a)) \ disc3 = is_Prot" and P_prop: "\m m'. normalized_nnf_match m \ P m \ compress_normalize_besteffort m = Some m' \ P m'" hence "\r\set rs. \ has_disc_negated disc3 False (get_match r) \ normalized_nnf_match (get_match r) \ P (get_match r) \ \r\set (optimize_matches_option compress_normalize_besteffort rs). normalized_nnf_match (get_match r) \ \ has_disc_negated disc3 False (get_match r) \ P (get_match r)" apply - thm optimize_matches_option_preserves[where P= "\m. normalized_nnf_match m \ \ has_disc_negated disc3 False m \ P m"] apply(rule optimize_matches_option_preserves[where P= "\m. normalized_nnf_match m \ \ has_disc_negated disc3 False m \ P m"]) apply(elim disjE) using compress_normalize_besteffort_nnf compress_normalize_besteffort_hasdisc_negated apply blast using compress_normalize_besteffort_nnf compress_normalize_besteffort_not_introduces_Iiface_negated compress_normalize_besteffort_not_introduces_Oiface_negated compress_normalize_besteffort_not_introduces_Prot_negated by blast+ } note y_generic=this note y=y_generic[of "\_. True", simplified] have case_disc3_is_not_prot: "\a. \ disc3 (Src_Ports a) \ \a. \ disc3 (Dst_Ports a) \ \a. \ disc3 (Src a) \ \a. \ disc3 (Dst a) \ (\a. \ disc3 (IIface a)) \ disc3 = is_Iiface \ (\a. \ disc3 (OIface a)) \ disc3 = is_Oiface \ (\a. \ disc3 (Prot a)) \ \ r \ set rs. \ has_disc_negated disc3 False (get_match r) \ normalized_nnf_match (get_match r) \ \ r \ set (transform_normalize_primitives rs). normalized_nnf_match (get_match r) \ \ has_disc_negated disc3 False (get_match r)" unfolding transform_normalize_primitives_def apply(simp) apply(rule y) apply(simp; fail) apply(simp; fail) apply(blast) apply(rule normalize_rules_preserves)+ apply(simp; fail) subgoal apply(intro allI impI conjI ballI) apply(rule rewrite_MultiportPorts_preserves_normalized_not_has_disc_negated, simp_all) by(simp add: rewrite_MultiportPorts_normalized_nnf_match) subgoal apply(clarify) apply(rule_tac m6=m in x_src_ports) by(simp)+ subgoal apply(clarify) by(rule x_dst_ports) simp+ using x[OF wf_disc_sel_common_primitive(3), of disc3 ipt_iprange_compress, folded normalize_src_ips_def] apply blast using x[OF wf_disc_sel_common_primitive(4), of disc3 ipt_iprange_compress, folded normalize_dst_ips_def] apply blast done have case_disc3_is_prot_optimize_matches_option:"\r\set rs. \ has_disc_negated is_Prot False (get_match r) \ normalized_nnf_match (get_match r) \ \ has_disc_negated is_Src_Ports False (get_match r) \ \ has_disc_negated is_Dst_Ports False (get_match r) \ \r\set (optimize_matches_option compress_normalize_besteffort rs). normalized_nnf_match (get_match r) \ \ has_disc_negated is_Prot False (get_match r) \ \ has_disc_negated is_Src_Ports False (get_match r) \ \ has_disc_negated is_Dst_Ports False (get_match r)" if isprot: "disc3 = is_Prot" for rs :: "'i common_primitive rule list" apply(rule y_generic[where P8="\m. \ has_disc_negated is_Src_Ports False m \ \ has_disc_negated is_Dst_Ports False m", simplified isprot]) apply simp+ apply(clarify) apply(intro conjI) using compress_normalize_besteffort_hasdisc_negated[of is_Src_Ports] apply fastforce using compress_normalize_besteffort_hasdisc_negated[of is_Dst_Ports] apply fastforce by simp (*copy from above, specific version for is_Prot*) (*Push through things, but now more complicated because several things could introduce Prots*) have case_disc3_is_prot: "disc3 = is_Prot \ \ r \ set rs. \ has_disc_negated disc3 False (get_match r) \ normalized_nnf_match (get_match r) \ \ has_disc_negated is_Src_Ports False (get_match r) \ \ has_disc_negated is_Dst_Ports False (get_match r) & \ has_disc is_MultiportPorts (get_match r) \ \MultiportPorts could be rewritten to negated \Src\/\Dst\ Ports\ \ \ r \ set (transform_normalize_primitives rs). normalized_nnf_match (get_match r) \ \ has_disc_negated disc3 False (get_match r) \ \ has_disc_negated is_Src_Ports False (get_match r) \ \ has_disc_negated is_Dst_Ports False (get_match r)" unfolding transform_normalize_primitives_def apply(simp) apply(rule case_disc3_is_prot_optimize_matches_option) apply(simp; fail) thm normalize_rules_property[ where P="\m. normalized_nnf_match m \ \ has_disc_negated disc3 False m"] apply(rule normalize_rules_property[ where P="\m. normalized_nnf_match m \ \ has_disc_negated disc3 False m \ \ has_disc_negated is_Src_Ports False m \ \ has_disc_negated is_Dst_Ports False m"]) (*dst ips*) apply(rule normalize_rules_property[ where P="\m. normalized_nnf_match m \ \ has_disc_negated disc3 False m \ \ has_disc_negated is_Src_Ports False m \ \ has_disc_negated is_Dst_Ports False m"])(*src ips*) apply(rule normalize_rules_property[ where P="\m. normalized_nnf_match m \ \ has_disc_negated disc3 False m \ \ has_disc_negated is_Src_Ports False m \ \ has_disc_negated is_Dst_Ports False m"])(*dst ports*) apply(rule normalize_rules_property[ where P="\m. normalized_nnf_match m \ \ has_disc_negated disc3 False m \ \ has_disc_negated is_Src_Ports False m \ \ has_disc_negated is_Dst_Ports False m"])(*src ports*) apply(rule normalize_rules_property[ where P="\m. normalized_nnf_match m \ \ has_disc_negated disc3 False m \ \ has_disc_negated is_Src_Ports False m \ \ has_disc_negated is_Dst_Ports False m \ \ has_disc is_MultiportPorts m"]) (*multiports, needs \ has_disc is_MultiportPorts m*) apply(simp; fail) subgoal apply(intro allI impI conjI ballI) apply(simp add: rewrite_MultiportPorts_normalized_nnf_match; fail) apply(rule rewrite_MultiportPorts_preserves_normalized_not_has_disc_negated, simp_all) \ \Here we need @{term "\ has_disc is_MultiportPorts m"}\ using rewrite_MultiportPorts_unchanged_if_not_has_disc by fastforce+ subgoal (*yeah, just need to consider the other cases*) apply(clarify) apply(frule_tac m6=m in x_src_ports[rotated 2]) apply(simp_all) apply simp using normalize_src_ports_preserves_normalized_not_has_disc_negated by blast subgoal apply(clarify) apply(frule_tac m6=m in x_dst_ports[rotated 2]) apply(simp_all) apply simp using normalize_dst_ports_preserves_normalized_not_has_disc_negated by blast using x[OF wf_disc_sel_common_primitive(3), of disc3 ipt_iprange_compress, folded normalize_src_ips_def] x[OF wf_disc_sel_common_primitive(3), of is_Dst_Ports ipt_iprange_compress, folded normalize_src_ips_def] x_generic[OF _ _ _ wf_disc_sel_common_primitive(3), of is_Src_Ports False _ _ ipt_iprange_compress, folded normalize_src_ips_def] apply (meson common_primitive.disc(45) common_primitive.disc(56) common_primitive.disc(67); fail) using x[OF wf_disc_sel_common_primitive(4), of disc3 ipt_iprange_compress, folded normalize_dst_ips_def] x[OF wf_disc_sel_common_primitive(4), of is_Src_Ports ipt_iprange_compress, folded normalize_dst_ips_def] x_generic[OF _ _ _ wf_disc_sel_common_primitive(4), of is_Dst_Ports False _ _ ipt_iprange_compress, folded normalize_dst_ips_def] apply (meson common_primitive.disc(46) common_primitive.disc(57) common_primitive.disc(68); fail) done show "unchanged disc3 \ changeddisc disc3 \ (\a. \ disc3 (Prot a)) \ (disc3 = is_Prot \ (\ r \ set rs. \ has_disc_negated is_Src_Ports False (get_match r) \ \ has_disc_negated is_Dst_Ports False (get_match r) \ \ has_disc is_MultiportPorts (get_match r))) \ \ r \ set rs. \ has_disc_negated disc3 False (get_match r) \ \ r \ set (transform_normalize_primitives rs). \ has_disc_negated disc3 False (get_match r)" unfolding unchanged_def changeddisc_def apply(elim disjE) using normalized case_disc3_is_not_prot apply blast using normalized case_disc3_is_prot by blast qed theorem iiface_constrain: assumes simplers: "simple_ruleset rs" and normalized: "\ r \ set rs. normalized_nnf_match (get_match r)" and wf_ipassmt: "ipassmt_sanity_nowildcards ipassmt" and nospoofing: "\ips. ipassmt (Iface (p_iiface p)) = Some ips \ p_src p \ ipcidr_union_set (set ips)" shows "(common_matcher, \),p\ \optimize_matches (iiface_constrain ipassmt) rs, s\ \\<^sub>\ t \ (common_matcher, \),p\ \rs, s\ \\<^sub>\ t" and "simple_ruleset (optimize_matches (iiface_constrain ipassmt) rs)" proof - show simplers_t: "simple_ruleset (optimize_matches (iiface_constrain ipassmt) rs)" by (simp add: optimize_matches_simple_ruleset simplers) have my_arg_cong: "\P Q. P s = Q s \ (P s = t) \ (Q s = t)" by simp show "(common_matcher, \),p\ \optimize_matches (iiface_constrain ipassmt) rs, s\ \\<^sub>\ t \ (common_matcher, \),p\ \rs, s\ \\<^sub>\ t" unfolding approximating_semantics_iff_fun_good_ruleset[OF simple_imp_good_ruleset[OF simplers_t]] unfolding approximating_semantics_iff_fun_good_ruleset[OF simple_imp_good_ruleset[OF simplers]] apply(rule my_arg_cong) apply(rule optimize_matches_generic[where P="\ m _. normalized_nnf_match m"]) apply(simp add: normalized) apply(rule matches_iiface_constrain) apply(simp_all add: wf_ipassmt nospoofing) done qed text\In contrast to @{thm iiface_constrain}, this requires @{const ipassmt_sanity_disjoint} and as much stronger nospoof assumption: This assumption requires that the packet is actually in ipassmt!\ theorem iiface_rewrite: assumes simplers: "simple_ruleset rs" and normalized: "\ r \ set rs. normalized_nnf_match (get_match r)" and wf_ipassmt: "ipassmt_sanity_nowildcards ipassmt" and disjoint_ipassmt: "ipassmt_sanity_disjoint ipassmt" and nospoofing: "\ips. ipassmt (Iface (p_iiface p)) = Some ips \ p_src p \ ipcidr_union_set (set ips)" shows "(common_matcher, \),p\ \optimize_matches (iiface_rewrite ipassmt) rs, s\ \\<^sub>\ t \ (common_matcher, \),p\ \rs, s\ \\<^sub>\ t" and "simple_ruleset (optimize_matches (iiface_rewrite ipassmt) rs)" proof - show simplers_t: "simple_ruleset (optimize_matches (iiface_rewrite ipassmt) rs)" by(simp add: simplers optimize_matches_simple_ruleset) \ \packet must come from a defined interface!\ from nospoofing have "Iface (p_iiface p) \ dom ipassmt" by blast have my_arg_cong: "\P Q. P s = Q s \ (P s = t) \ (Q s = t)" by simp show "(common_matcher, \),p\ \optimize_matches (iiface_rewrite ipassmt) rs, s\ \\<^sub>\ t \ (common_matcher, \),p\ \rs, s\ \\<^sub>\ t" unfolding approximating_semantics_iff_fun_good_ruleset[OF simple_imp_good_ruleset[OF simplers_t]] unfolding approximating_semantics_iff_fun_good_ruleset[OF simple_imp_good_ruleset[OF simplers]] apply(rule my_arg_cong) apply(rule optimize_matches_generic[where P="\ m _. normalized_nnf_match m"]) apply(simp add: normalized) apply(rule matches_iiface_rewrite) apply(simp_all add: wf_ipassmt nospoofing disjoint_ipassmt) done qed (* Copy of iiface_rewrite *) theorem oiface_rewrite: assumes simplers: "simple_ruleset rs" and normalized: "\ r \ set rs. normalized_nnf_match (get_match r)" and wf_ipassmt: "ipassmt_sanity_nowildcards ipassmt" and ipassmt_from_rt: "ipassmt = map_of (routing_ipassmt rt)" and correct_routing: "correct_routing rt" and rtbl_decided: "output_iface (routing_table_semantics rt (p_dst p)) = p_oiface p" shows "(common_matcher, \),p\ \optimize_matches (oiface_rewrite ipassmt) rs, s\ \\<^sub>\ t \ (common_matcher, \),p\ \rs, s\ \\<^sub>\ t" and "simple_ruleset (optimize_matches (oiface_rewrite ipassmt) rs)" proof - show simplers_t: "simple_ruleset (optimize_matches (oiface_rewrite ipassmt) rs)" using simplers by(fact optimize_matches_simple_ruleset) show "(common_matcher, \),p\ \optimize_matches (oiface_rewrite ipassmt) rs, s\ \\<^sub>\ t \ (common_matcher, \),p\ \rs, s\ \\<^sub>\ t" unfolding approximating_semantics_iff_fun_good_ruleset[OF simple_imp_good_ruleset[OF simplers_t]] unfolding approximating_semantics_iff_fun_good_ruleset[OF simple_imp_good_ruleset[OF simplers]] apply(rule arg_cong[where f="\x. x = t"]) apply(rule optimize_matches_generic[where P="\ m _. normalized_nnf_match m"]) apply(simp add: normalized ;fail) apply(rule matches_oiface_rewrite[OF _ _ _ ipassmt_from_rt]; assumption?) apply(simp_all add: wf_ipassmt correct_routing rtbl_decided) done qed definition upper_closure :: "'i::len common_primitive rule list \ 'i common_primitive rule list" where "upper_closure rs == remdups_rev (transform_optimize_dnf_strict (transform_normalize_primitives (transform_optimize_dnf_strict (optimize_matches_a upper_closure_matchexpr rs))))" definition lower_closure :: "'i::len common_primitive rule list \ 'i common_primitive rule list" where "lower_closure rs == remdups_rev (transform_optimize_dnf_strict (transform_normalize_primitives (transform_optimize_dnf_strict (optimize_matches_a lower_closure_matchexpr rs))))" text\putting it all together\ lemma transform_upper_closure: assumes simplers: "simple_ruleset rs" \ \semantics are preserved\ shows "(common_matcher, in_doubt_allow),p\ \upper_closure rs, s\ \\<^sub>\ t \ (common_matcher, in_doubt_allow),p\ \rs, s\ \\<^sub>\ t" and "simple_ruleset (upper_closure rs)" \ \simple, normalized rules without unknowns\ and "\ r \ set (upper_closure rs). normalized_nnf_match (get_match r) \ normalized_src_ports (get_match r) \ normalized_dst_ports (get_match r) \ normalized_src_ips (get_match r) \ normalized_dst_ips (get_match r) \ \ has_disc is_MultiportPorts (get_match r) \ \ has_disc is_Extra (get_match r)" \ \no new primitives are introduced\ and "\a. \ disc (Src_Ports a) \ \a. \ disc (Dst_Ports a) \ \a. \ disc (Src a) \ \a. \ disc (Dst a) \ \a. \ disc (IIface a) \ disc = is_Iiface \ \a. \ disc (OIface a) \ disc = is_Oiface \ \a. \ disc (Prot a) \ \ r \ set rs. \ has_disc disc (get_match r) \ \ r \ set (upper_closure rs). \ has_disc disc (get_match r)" and "\a. \ disc (Src_Ports a) \ \a. \ disc (Dst_Ports a) \ \a. \ disc (Src a) \ \a. \ disc (Dst a) \ \a. \ disc (IIface a) \ disc = is_Iiface \ \a. \ disc (OIface a) \ disc = is_Oiface \ (\a. \ disc (Prot a)) \ disc = is_Prot \ \ \if it is prot, there must not be negated matches on ports\ (\ r \ set rs. \ has_disc_negated is_Src_Ports False (get_match r) \ \ has_disc_negated is_Dst_Ports False (get_match r) \ \ has_disc is_MultiportPorts (get_match r)) \ \ r \ set rs. \ has_disc_negated disc False (get_match r) \ \ r \ set (upper_closure rs). \ has_disc_negated disc False (get_match r)" proof - let ?rs1="optimize_matches_a upper_closure_matchexpr rs" let ?rs2="transform_optimize_dnf_strict ?rs1" let ?rs3="transform_normalize_primitives ?rs2" let ?rs4="transform_optimize_dnf_strict ?rs3" { fix m a have "Rule m a \ set (upper_closure rs) \ (a = action.Accept \ a = action.Drop) \ normalized_nnf_match m \ normalized_src_ports m \ normalized_dst_ports m \ normalized_src_ips m \ normalized_dst_ips m \ \ has_disc is_MultiportPorts m \ \ has_disc is_Extra m" using simplers unfolding upper_closure_def apply(simp add: remdups_rev_set) apply(frule transform_remove_unknowns_upper(4)) apply(drule transform_remove_unknowns_upper(2)) thm transform_optimize_dnf_strict[OF _ wf_in_doubt_allow] apply(frule(1) transform_optimize_dnf_strict_structure(2)[OF _ wf_in_doubt_allow, where disc=is_Extra]) apply(thin_tac "\r\ set (optimize_matches_a upper_closure_matchexpr rs). \ has_disc is_Extra (get_match r)") apply(frule transform_optimize_dnf_strict_structure(3)[OF _ wf_in_doubt_allow]) apply(drule transform_optimize_dnf_strict_structure(1)[OF _ wf_in_doubt_allow]) thm transform_normalize_primitives[OF _ wf_in_doubt_allow] apply(frule(1) transform_normalize_primitives(3)[OF _ wf_in_doubt_allow, of _ is_Extra]) apply(simp;fail) apply(simp;fail) apply(simp;fail) apply blast apply(thin_tac "\r\ set ?rs2. \ has_disc is_Extra (get_match r)") apply(frule(1) transform_normalize_primitives(5)[OF _ wf_in_doubt_allow]) apply(drule transform_normalize_primitives(2)[OF _ wf_in_doubt_allow], simp) thm transform_optimize_dnf_strict[OF _ wf_in_doubt_allow] apply(frule(1) transform_optimize_dnf_strict_structure(2)[OF _ wf_in_doubt_allow, where disc=is_Extra]) apply(frule transform_optimize_dnf_strict_structure(2)[OF _ wf_in_doubt_allow, where disc=is_MultiportPorts]) apply blast apply(frule transform_optimize_dnf_strict_structure(3)[OF _ wf_in_doubt_allow]) apply(frule transform_optimize_dnf_strict_structure(4)[OF _ wf_in_doubt_allow, of _ "(is_Src_Ports, src_ports_sel)" "(\ps. case ps of L4Ports _ pts \ length pts \ 1)"]) apply(simp add: normalized_src_ports_def2; fail) apply(frule transform_optimize_dnf_strict_structure(4)[OF _ wf_in_doubt_allow, of _ "(is_Dst_Ports, dst_ports_sel)" "(\ps. case ps of L4Ports _ pts \ length pts \ 1)"]) apply(simp add: normalized_dst_ports_def2; fail) apply(frule transform_optimize_dnf_strict_structure(4)[OF _ wf_in_doubt_allow, of _ "(is_Src, src_sel)" normalized_cidr_ip]) apply(simp add: normalized_src_ips_def2; fail) apply(frule transform_optimize_dnf_strict_structure(4)[OF _ wf_in_doubt_allow, of _ "(is_Dst, dst_sel)" normalized_cidr_ip]) apply(simp add: normalized_dst_ips_def2; fail) apply(thin_tac "\r\set ?rs2. _ r")+ apply(thin_tac "\r\set ?rs3. _ r")+ apply(drule transform_optimize_dnf_strict_structure(1)[OF _ wf_in_doubt_allow]) apply(subgoal_tac "(a = action.Accept \ a = action.Drop)") prefer 2 apply(simp_all add: simple_ruleset_def) apply fastforce apply(simp add: normalized_src_ports_def2 normalized_dst_ports_def2 normalized_src_ips_def2 normalized_dst_ips_def2) apply(intro conjI) by fastforce+ (*1s*) } note 1=this from 1 show "simple_ruleset (upper_closure rs)" apply(simp add: simple_ruleset_def) apply(clarify) apply(rename_tac r) apply(case_tac r) apply(simp) by blast from 1 show "\ r \ set (upper_closure rs). normalized_nnf_match (get_match r) \ normalized_src_ports (get_match r) \ normalized_dst_ports (get_match r) \ normalized_src_ips (get_match r) \ normalized_dst_ips (get_match r) \ \ has_disc is_MultiportPorts (get_match r) \ \ has_disc is_Extra (get_match r)" apply(clarify) apply(rename_tac r) apply(case_tac r) apply(simp) done show "\a. \ disc (Src_Ports a) \ \a. \ disc (Dst_Ports a) \ \a. \ disc (Src a) \ \a. \ disc (Dst a) \ \a. \ disc (IIface a) \ disc = is_Iiface \ \a. \ disc (OIface a) \ disc = is_Oiface \ \a. \ disc (Prot a) \ \ r \ set rs. \ has_disc disc (get_match r) \ \ r \ set (upper_closure rs). \ has_disc disc (get_match r)" using simplers unfolding upper_closure_def apply - apply(frule(1) transform_remove_unknowns_upper(3)[where disc=disc]) apply(drule transform_remove_unknowns_upper(2)) apply(frule(1) transform_optimize_dnf_strict_structure(2)[OF _ wf_in_doubt_allow, where disc=disc]) apply(frule transform_optimize_dnf_strict_structure(3)[OF _ wf_in_doubt_allow]) apply(drule transform_optimize_dnf_strict_structure(1)[OF _ wf_in_doubt_allow]) apply(frule(1) transform_normalize_primitives(3)[OF _ wf_in_doubt_allow, of _ disc]) apply(simp;fail) apply blast apply(simp;fail) apply(simp;fail) apply(drule transform_normalize_primitives(2)[OF _ wf_in_doubt_allow], simp) apply(frule(1) transform_optimize_dnf_strict_structure(2)[OF _ wf_in_doubt_allow, where disc=disc]) apply(simp add: remdups_rev_set) done have no_ports_1: "\ has_disc_negated is_Src_Ports False (get_match m) \ \ has_disc_negated is_Dst_Ports False (get_match m) \ \ has_disc is_MultiportPorts (get_match m)" if no_ports: "\r\set rs. \ has_disc_negated is_Src_Ports False (get_match r) \ \ has_disc_negated is_Dst_Ports False (get_match r) \ \ has_disc is_MultiportPorts (get_match r)" and m: "m \ set (transform_optimize_dnf_strict (optimize_matches_a upper_closure_matchexpr rs))" for m proof - from no_ports transform_remove_unknowns_upper(3,6)[OF simplers] have "\r\ set (optimize_matches_a upper_closure_matchexpr rs). \ has_disc_negated is_Src_Ports False (get_match r) \ \ has_disc_negated is_Dst_Ports False (get_match r) \ \ has_disc is_MultiportPorts (get_match r)" by blast with m transform_optimize_dnf_strict_structure(2,5)[OF optimize_matches_a_simple_ruleset[OF simplers] wf_in_doubt_allow, of upper_closure_matchexpr] show ?thesis by blast qed show"\a. \ disc (Src_Ports a) \ \a. \ disc (Dst_Ports a) \ \a. \ disc (Src a) \ \a. \ disc (Dst a) \ \a. \ disc (IIface a) \ disc = is_Iiface \ \a. \ disc (OIface a) \ disc = is_Oiface \ (\a. \ disc (Prot a)) \ disc = is_Prot \ (\ r \ set rs. \ has_disc_negated is_Src_Ports False (get_match r) \ \ has_disc_negated is_Dst_Ports False (get_match r) \ \ has_disc is_MultiportPorts (get_match r)) \ \ r \ set rs. \ has_disc_negated disc False (get_match r) \ \ r \ set (upper_closure rs). \ has_disc_negated disc False (get_match r)" using simplers unfolding upper_closure_def apply - apply(frule(1) transform_remove_unknowns_upper(6)[where disc=disc]) apply(drule transform_remove_unknowns_upper(2)) apply(frule(1) transform_optimize_dnf_strict_structure(5)[OF _ wf_in_doubt_allow, where disc=disc]) apply(frule transform_optimize_dnf_strict_structure(3)[OF _ wf_in_doubt_allow]) apply(drule transform_optimize_dnf_strict_structure(1)[OF _ wf_in_doubt_allow]) apply(frule(1) transform_normalize_primitives(7)[OF _ wf_in_doubt_allow, of _ disc]) apply(simp;fail) apply blast apply(elim disjE) apply blast apply(simp) using no_ports_1 apply fast apply(simp;fail) apply(drule transform_normalize_primitives(2)[OF _ wf_in_doubt_allow], simp) apply(frule(1) transform_optimize_dnf_strict_structure(5)[OF _ wf_in_doubt_allow, where disc=disc]) apply(simp add: remdups_rev_set) done show "(common_matcher, in_doubt_allow),p\ \upper_closure rs, s\ \\<^sub>\ t \ (common_matcher, in_doubt_allow),p\ \rs, s\ \\<^sub>\ t" using simplers unfolding upper_closure_def apply - apply(frule transform_remove_unknowns_upper(1)[where p=p and s=s and t=t]) apply(drule transform_remove_unknowns_upper(2)) apply(frule transform_optimize_dnf_strict[OF _ wf_in_doubt_allow, where p=p and s=s and t=t]) apply(frule transform_optimize_dnf_strict_structure(3)[OF _ wf_in_doubt_allow]) apply(drule transform_optimize_dnf_strict_structure(1)[OF _ wf_in_doubt_allow]) apply(frule(1) transform_normalize_primitives(1)[OF _ wf_in_doubt_allow, where p=p and s=s and t=t]) apply(drule transform_normalize_primitives(2)[OF _ wf_in_doubt_allow], simp) apply(frule transform_optimize_dnf_strict[OF _ wf_in_doubt_allow, where p=p and s=s and t=t]) apply(drule transform_optimize_dnf_strict_structure(1)[OF _ wf_in_doubt_allow]) apply(simp) using approximating_bigstep_fun_remdups_rev by (simp add: approximating_bigstep_fun_remdups_rev approximating_semantics_iff_fun_good_ruleset remdups_rev_simplers simple_imp_good_ruleset) qed (*copy&paste from transform_upper_closure*) lemma transform_lower_closure: assumes simplers: "simple_ruleset rs" \ \semantics are preserved\ shows "(common_matcher, in_doubt_deny),p\ \lower_closure rs, s\ \\<^sub>\ t \ (common_matcher, in_doubt_deny),p\ \rs, s\ \\<^sub>\ t" and "simple_ruleset (lower_closure rs)" \ \simple, normalized rules without unknowns\ and "\ r \ set (lower_closure rs). normalized_nnf_match (get_match r) \ normalized_src_ports (get_match r) \ normalized_dst_ports (get_match r) \ normalized_src_ips (get_match r) \ normalized_dst_ips (get_match r) \ \ has_disc is_MultiportPorts (get_match r) \ \ has_disc is_Extra (get_match r)" \ \no new primitives are introduced\ and "\a. \ disc (Src_Ports a) \ \a. \ disc (Dst_Ports a) \ \a. \ disc (Src a) \ \a. \ disc (Dst a) \ \a. \ disc (IIface a) \ disc = is_Iiface \ \a. \ disc (OIface a) \ disc = is_Oiface \ \a. \ disc (Prot a) \ \ r \ set rs. \ has_disc disc (get_match r) \ \ r \ set (lower_closure rs). \ has_disc disc (get_match r)" and "\a. \ disc (Src_Ports a) \ \a. \ disc (Dst_Ports a) \ \a. \ disc (Src a) \ \a. \ disc (Dst a) \ \a. \ disc (IIface a) \ disc = is_Iiface \ \a. \ disc (OIface a) \ disc = is_Oiface \ (\a. \ disc (Prot a)) \ disc = is_Prot \ (\ r \ set rs. \ has_disc_negated is_Src_Ports False (get_match r) \ \ has_disc_negated is_Dst_Ports False (get_match r) \ \ has_disc is_MultiportPorts (get_match r)) \ \ r \ set rs. \ has_disc_negated disc False (get_match r) \ \ r \ set (lower_closure rs). \ has_disc_negated disc False (get_match r)" proof - let ?rs1="optimize_matches_a lower_closure_matchexpr rs" let ?rs2="transform_optimize_dnf_strict ?rs1" let ?rs3="transform_normalize_primitives ?rs2" let ?rs4="transform_optimize_dnf_strict ?rs3" { fix m a have "Rule m a \ set (lower_closure rs) \ (a = action.Accept \ a = action.Drop) \ normalized_nnf_match m \ normalized_src_ports m \ normalized_dst_ports m \ normalized_src_ips m \ normalized_dst_ips m \ \ has_disc is_MultiportPorts m \ \ has_disc is_Extra m" using simplers unfolding lower_closure_def apply(simp add: remdups_rev_set) apply(frule transform_remove_unknowns_lower(4)) apply(drule transform_remove_unknowns_lower(2)) thm transform_optimize_dnf_strict[OF _ wf_in_doubt_deny] apply(frule(1) transform_optimize_dnf_strict_structure(2)[OF _ wf_in_doubt_deny, where disc=is_Extra]) apply(thin_tac "\r\ set (optimize_matches_a lower_closure_matchexpr rs). \ has_disc is_Extra (get_match r)") apply(frule transform_optimize_dnf_strict_structure(3)[OF _ wf_in_doubt_deny]) apply(drule transform_optimize_dnf_strict_structure(1)[OF _ wf_in_doubt_deny]) thm transform_normalize_primitives[OF _ wf_in_doubt_deny] apply(frule(1) transform_normalize_primitives(3)[OF _ wf_in_doubt_deny, of _ is_Extra]) apply(simp;fail) apply(simp;fail) apply(simp;fail) apply blast apply(thin_tac "\r\ set ?rs2. \ has_disc is_Extra (get_match r)") apply(frule(1) transform_normalize_primitives(5)[OF _ wf_in_doubt_deny]) apply(drule transform_normalize_primitives(2)[OF _ wf_in_doubt_deny], simp) thm transform_optimize_dnf_strict[OF _ wf_in_doubt_deny] apply(frule(1) transform_optimize_dnf_strict_structure(2)[OF _ wf_in_doubt_deny, where disc=is_Extra]) apply(frule transform_optimize_dnf_strict_structure(2)[OF _ wf_in_doubt_deny, where disc=is_MultiportPorts]) apply blast apply(frule transform_optimize_dnf_strict_structure(3)[OF _ wf_in_doubt_deny]) apply(frule transform_optimize_dnf_strict_structure(4)[OF _ wf_in_doubt_deny, of _ "(is_Src_Ports, src_ports_sel)" "(\ps. case ps of L4Ports _ pts \ length pts \ 1)"]) apply(simp add: normalized_src_ports_def2; fail) apply(frule transform_optimize_dnf_strict_structure(4)[OF _ wf_in_doubt_deny, of _ "(is_Dst_Ports, dst_ports_sel)" "(\ps. case ps of L4Ports _ pts \ length pts \ 1)"]) apply(simp add: normalized_dst_ports_def2; fail) apply(frule transform_optimize_dnf_strict_structure(4)[OF _ wf_in_doubt_deny, of _ "(is_Src, src_sel)" normalized_cidr_ip]) apply(simp add: normalized_src_ips_def2; fail) apply(frule transform_optimize_dnf_strict_structure(4)[OF _ wf_in_doubt_deny, of _ "(is_Dst, dst_sel)" normalized_cidr_ip]) apply(simp add: normalized_dst_ips_def2; fail) apply(thin_tac "\r\set ?rs2. _ r")+ apply(thin_tac "\r\set ?rs3. _ r")+ apply(drule transform_optimize_dnf_strict_structure(1)[OF _ wf_in_doubt_deny]) apply(subgoal_tac "(a = action.Accept \ a = action.Drop)") prefer 2 apply(simp_all add: simple_ruleset_def) apply fastforce apply(simp add: normalized_src_ports_def2 normalized_dst_ports_def2 normalized_src_ips_def2 normalized_dst_ips_def2) apply(intro conjI) by fastforce+ (*1s*) } note 1=this from 1 show "simple_ruleset (lower_closure rs)" apply(simp add: simple_ruleset_def) apply(clarify) apply(rename_tac r) apply(case_tac r) apply(simp) by blast from 1 show "\ r \ set (lower_closure rs). normalized_nnf_match (get_match r) \ normalized_src_ports (get_match r) \ normalized_dst_ports (get_match r) \ normalized_src_ips (get_match r) \ normalized_dst_ips (get_match r) \ \ has_disc is_MultiportPorts (get_match r) \ \ has_disc is_Extra (get_match r)" apply(clarify) apply(rename_tac r) apply(case_tac r) apply(simp) done show "\a. \ disc (Src_Ports a) \ \a. \ disc (Dst_Ports a) \ \a. \ disc (Src a) \ \a. \ disc (Dst a) \ \a. \ disc (IIface a) \ disc = is_Iiface \ \a. \ disc (OIface a) \ disc = is_Oiface \ \a. \ disc (Prot a) \ \ r \ set rs. \ has_disc disc (get_match r) \ \ r \ set (lower_closure rs). \ has_disc disc (get_match r)" using simplers unfolding lower_closure_def apply - apply(frule(1) transform_remove_unknowns_lower(3)[where disc=disc]) apply(drule transform_remove_unknowns_lower(2)) apply(frule(1) transform_optimize_dnf_strict_structure(2)[OF _ wf_in_doubt_deny, where disc=disc]) apply(frule transform_optimize_dnf_strict_structure(3)[OF _ wf_in_doubt_deny]) apply(drule transform_optimize_dnf_strict_structure(1)[OF _ wf_in_doubt_deny]) apply(frule(1) transform_normalize_primitives(3)[OF _ wf_in_doubt_deny, of _ disc]) apply(simp;fail) apply blast apply(simp;fail) apply(simp;fail) apply(drule transform_normalize_primitives(2)[OF _ wf_in_doubt_deny], simp) apply(frule(1) transform_optimize_dnf_strict_structure(2)[OF _ wf_in_doubt_deny, where disc=disc]) apply(simp add: remdups_rev_set) done have no_ports_1: "\ has_disc_negated is_Src_Ports False (get_match m) \ \ has_disc_negated is_Dst_Ports False (get_match m) \ \ has_disc is_MultiportPorts (get_match m)" if no_ports: "\r\set rs. \ has_disc_negated is_Src_Ports False (get_match r) \ \ has_disc_negated is_Dst_Ports False (get_match r) \ \ has_disc is_MultiportPorts (get_match r)" and m: "m \ set (transform_optimize_dnf_strict (optimize_matches_a lower_closure_matchexpr rs))" for m proof - from no_ports transform_remove_unknowns_lower(3,6)[OF simplers] have "\r\ set (optimize_matches_a lower_closure_matchexpr rs). \ has_disc_negated is_Src_Ports False (get_match r) \ \ has_disc_negated is_Dst_Ports False (get_match r) \ \ has_disc is_MultiportPorts (get_match r)" by blast from m this transform_optimize_dnf_strict_structure(2,5)[OF optimize_matches_a_simple_ruleset[OF simplers] wf_in_doubt_deny, of lower_closure_matchexpr] show ?thesis by blast qed show"\a. \ disc (Src_Ports a) \ \a. \ disc (Dst_Ports a) \ \a. \ disc (Src a) \ \a. \ disc (Dst a) \ \a. \ disc (IIface a) \ disc = is_Iiface \ \a. \ disc (OIface a) \ disc = is_Oiface \ (\a. \ disc (Prot a)) \ disc = is_Prot \ (\ r \ set rs. \ has_disc_negated is_Src_Ports False (get_match r) \ \ has_disc_negated is_Dst_Ports False (get_match r) \ \ has_disc is_MultiportPorts (get_match r)) \ \ r \ set rs. \ has_disc_negated disc False (get_match r) \ \ r \ set (lower_closure rs). \ has_disc_negated disc False (get_match r)" using simplers unfolding lower_closure_def apply - apply(frule(1) transform_remove_unknowns_lower(6)[where disc=disc]) apply(drule transform_remove_unknowns_lower(2)) apply(frule(1) transform_optimize_dnf_strict_structure(5)[OF _ wf_in_doubt_deny, where disc=disc]) apply(frule transform_optimize_dnf_strict_structure(3)[OF _ wf_in_doubt_deny]) apply(drule transform_optimize_dnf_strict_structure(1)[OF _ wf_in_doubt_deny]) apply(frule(1) transform_normalize_primitives(7)[OF _ wf_in_doubt_deny, of _ disc]) apply(simp;fail) apply blast apply(elim disjE) apply blast apply(simp) using no_ports_1 apply fast apply(simp;fail) apply(drule transform_normalize_primitives(2)[OF _ wf_in_doubt_deny], simp) apply(frule(1) transform_optimize_dnf_strict_structure(5)[OF _ wf_in_doubt_deny, where disc=disc]) apply(simp add: remdups_rev_set) done show "(common_matcher, in_doubt_deny),p\ \lower_closure rs, s\ \\<^sub>\ t \ (common_matcher, in_doubt_deny),p\ \rs, s\ \\<^sub>\ t" using simplers unfolding lower_closure_def apply - apply(frule transform_remove_unknowns_lower(1)[where p=p and s=s and t=t]) apply(drule transform_remove_unknowns_lower(2)) apply(frule transform_optimize_dnf_strict[OF _ wf_in_doubt_deny, where p=p and s=s and t=t]) apply(frule transform_optimize_dnf_strict_structure(3)[OF _ wf_in_doubt_deny]) apply(drule transform_optimize_dnf_strict_structure(1)[OF _ wf_in_doubt_deny]) apply(frule(1) transform_normalize_primitives(1)[OF _ wf_in_doubt_deny, where p=p and s=s and t=t]) apply(drule transform_normalize_primitives(2)[OF _ wf_in_doubt_deny], simp) apply(frule transform_optimize_dnf_strict[OF _ wf_in_doubt_deny, where p=p and s=s and t=t]) apply(drule transform_optimize_dnf_strict_structure(1)[OF _ wf_in_doubt_deny]) apply(simp) using approximating_bigstep_fun_remdups_rev by (simp add: approximating_bigstep_fun_remdups_rev approximating_semantics_iff_fun_good_ruleset remdups_rev_simplers simple_imp_good_ruleset) qed definition iface_try_rewrite :: "(iface \ ('i::len word \ nat) list) list \ 'i prefix_routing option \ 'i common_primitive rule list \ 'i common_primitive rule list" where "iface_try_rewrite ipassmt rtblo rs \ let o_rewrite = (case rtblo of None \ id | Some rtbl \ transform_optimize_dnf_strict \ optimize_matches (oiface_rewrite (map_of_ipassmt (routing_ipassmt rtbl)))) in if ipassmt_sanity_disjoint (map_of ipassmt) \ ipassmt_sanity_defined rs (map_of ipassmt) then optimize_matches (iiface_rewrite (map_of_ipassmt ipassmt)) (o_rewrite rs) else optimize_matches (iiface_constrain (map_of_ipassmt ipassmt)) (o_rewrite rs)" text\Where @{typ "(iface \ ('i::len word \ nat) list) list"} is @{const map_of}@{typ "'i::len ipassignment"}. The sanity checkers need to iterate over the interfaces, hence we don't pass a map but a list of tuples.\ text\In @{file \Transform.thy\} there should be the final correctness theorem for \iface_try_rewrite\. Here are some structural properties.\ lemma iface_try_rewrite_simplers: "simple_ruleset rs \ simple_ruleset (iface_try_rewrite ipassmt rtblo rs)" by(simp add: iface_try_rewrite_def optimize_matches_simple_ruleset transform_optimize_dnf_strict_structure(1)[OF _ wf_in_doubt_allow (* The wf_unknown_match_tac is only required for some other parts of that lemma group, so any wellfounded tactic will do. *)] Let_def split: option.splits) lemma iiface_rewrite_preserves_nodisc: "\a. \ disc (Src a) \ \ has_disc disc m \ \ has_disc disc (iiface_rewrite ipassmt m)" proof(induction ipassmt m rule: iiface_rewrite.induct) case 2 have "\a. \ disc (Src a) \ \ disc (IIface ifce) \ \ has_disc disc (ipassmt_iface_replace_srcip_mexpr ipassmt ifce)" for ifce ipassmt apply(simp add: ipassmt_iface_replace_srcip_mexpr_def split: option.split) apply(intro allI impI, rename_tac ips) apply(drule_tac X=Src and ls="map (uncurry IpAddrNetmask) ips" in match_list_to_match_expr_not_has_disc) apply(simp) done with 2 show ?case by simp qed(simp_all) lemma iiface_constrain_preserves_nodisc: "\a. \ disc (Src a) \ \ has_disc disc m \ \ has_disc disc (iiface_constrain ipassmt m)" proof(induction ipassmt m rule: iiface_rewrite.induct) case 2 have "\a. \ disc (Src a) \ \ disc (IIface ifce) \ \ has_disc disc (ipassmt_iface_constrain_srcip_mexpr ipassmt ifce)" for ifce ipassmt apply(simp add: ipassmt_iface_constrain_srcip_mexpr_def split: option.split) apply(intro allI impI, rename_tac ips) apply(drule_tac X=Src and ls="map (uncurry IpAddrNetmask) ips" in match_list_to_match_expr_not_has_disc) apply(simp) done with 2 show ?case by simp qed(simp_all) lemma iface_try_rewrite_preserves_nodisc: " simple_ruleset rs \ \a. \ disc (Src a) \ \a. \ disc (Dst a) \ \r\ set rs. \ has_disc disc (get_match r) \ \r\ set (iface_try_rewrite ipassmt rtblo rs). \ has_disc disc (get_match r)" apply(insert wf_in_doubt_deny) (* to appease transform_optimize_dnf_strict_structure *) apply(simp add: iface_try_rewrite_def Let_def) apply(intro conjI impI optimize_matches_preserves) apply(case_tac[!] rtblo) apply(simp_all add: oiface_rewrite_preserves_nodisc iiface_rewrite_preserves_nodisc iiface_constrain_preserves_nodisc) (* solves the two None-cases *) apply(rule iiface_rewrite_preserves_nodisc; assumption?) apply(rule transform_optimize_dnf_strict_structure(2)[THEN bspec]; (assumption|simp add: optimize_matches_simple_ruleset; fail)?) apply(rule optimize_matches_preserves) apply(rule oiface_rewrite_preserves_nodisc; simp; fail) apply(rule iiface_constrain_preserves_nodisc; assumption?) apply(rule transform_optimize_dnf_strict_structure(2)[THEN bspec]; (assumption|simp add: optimize_matches_simple_ruleset; fail)?) apply(rule optimize_matches_preserves) apply(rule oiface_rewrite_preserves_nodisc; simp; fail) done theorem iface_try_rewrite_no_rtbl: assumes simplers: "simple_ruleset rs" and normalized: "\ r \ set rs. normalized_nnf_match (get_match r)" and wf_ipassmt1: "ipassmt_sanity_nowildcards (map_of ipassmt)" and wf_ipassmt2: "distinct (map fst ipassmt)" and nospoofing: "\ips. (map_of ipassmt) (Iface (p_iiface p)) = Some ips \ p_src p \ ipcidr_union_set (set ips)" shows "(common_matcher, \),p\ \iface_try_rewrite ipassmt None rs, s\ \\<^sub>\ t \ (common_matcher, \),p\ \rs, s\ \\<^sub>\ t" proof - show "(common_matcher, \),p\ \iface_try_rewrite ipassmt None rs, s\ \\<^sub>\ t \ (common_matcher, \),p\ \rs, s\ \\<^sub>\ t" apply(simp add: iface_try_rewrite_def Let_def comp_def) apply(simp add: map_of_ipassmt_def wf_ipassmt1 wf_ipassmt2) apply(intro conjI impI) apply(elim conjE) using iiface_rewrite(1)[OF simplers normalized wf_ipassmt1 _ nospoofing] apply blast using iiface_constrain(1)[OF simplers normalized wf_ipassmt1, where p = p] nospoofing apply force done qed lemma optimize_matches_comp: assumes mono: "\m. matcheq_matchNone m \ matcheq_matchNone (g m)" shows "optimize_matches (g \ f) rs = optimize_matches g ((optimize_matches f) rs)" unfolding optimize_matches_def proof(induction rs) case (Cons r rs) obtain m a where [simp]: "r = Rule m a" by(cases r) show ?case proof(cases "matcheq_matchNone (f m)") case True hence mn: "matcheq_matchNone (g (f m))" by(fact mono) show ?thesis by(unfold comp_def (* occasionally, the simplifier is weird *); simp add: mn Cons.IH[unfolded comp_def]) next case False show ?thesis by(unfold comp_def; simp add: False Cons.IH[unfolded comp_def]) qed qed simp (* optimize_matches_comp is a really nice lemma. The problem is that it is useless because I cannot execute the two rewrites after each other without going back to nnf. *) context begin private lemma iiface_rewrite_monoNone: "matcheq_matchNone m \ matcheq_matchNone (iiface_rewrite ipassmt m)" by(induction m rule: matcheq_matchNone.induct) auto private lemma iiface_constrain_monoNone: "matcheq_matchNone m \ matcheq_matchNone (iiface_constrain ipassmt m)" by(induction m rule: matcheq_matchNone.induct) auto private lemmas optimize_matches_iiface_comp = optimize_matches_comp[OF iiface_rewrite_monoNone] optimize_matches_comp[OF iiface_constrain_monoNone] end theorem iface_try_rewrite_rtbl: assumes simplers: "simple_ruleset rs" and normalized: "\ r \ set rs. normalized_nnf_match (get_match r)" and wf_ipassmt: "ipassmt_sanity_nowildcards (map_of ipassmt)" "distinct (map fst ipassmt)" and nospoofing: "\ips. (map_of ipassmt) (Iface (p_iiface p)) = Some ips \ p_src p \ ipcidr_union_set (set ips)" and routing_decided: "output_iface (routing_table_semantics rtbl (p_dst p)) = p_oiface p" and correct_routing: "correct_routing rtbl" and wf_ipassmt_o: "ipassmt_sanity_nowildcards (map_of (routing_ipassmt rtbl))" and wf_match_tac: "wf_unknown_match_tac \" shows "(common_matcher, \),p\ \iface_try_rewrite ipassmt (Some rtbl) rs, s\ \\<^sub>\ t \ (common_matcher, \),p\ \rs, s\ \\<^sub>\ t" proof - note oiface_rewrite = oiface_rewrite[OF simplers normalized wf_ipassmt_o refl correct_routing routing_decided] let ?ors = "optimize_matches (oiface_rewrite (map_of (routing_ipassmt rtbl))) rs" let ?nrs = "transform_optimize_dnf_strict ?ors" have osimplers: "simple_ruleset ?ors" using oiface_rewrite(2) . have nsimplers: "simple_ruleset ?nrs" using transform_optimize_dnf_strict_structure(1)[OF osimplers wf_match_tac] . have nnormalized: "\ r \ set ?nrs. normalized_nnf_match (get_match r)" using transform_optimize_dnf_strict_structure(3)[OF osimplers wf_match_tac] . note nnf = transform_optimize_dnf_strict[OF osimplers wf_match_tac] have nospoofing_alt: "\ips. map_of ipassmt (Iface (p_iiface p)) = Some ips \ p_src p \ ipcidr_union_set (set ips)" using nospoofing by simp show "(common_matcher, \),p\ \iface_try_rewrite ipassmt (Some rtbl) rs, s\ \\<^sub>\ t \ (common_matcher, \),p\ \rs, s\ \\<^sub>\ t" apply(simp add: iface_try_rewrite_def Let_def) apply(simp add: map_of_ipassmt_def wf_ipassmt routing_ipassmt_distinct wf_ipassmt_o) apply(intro conjI impI; (elim conjE)?) subgoal using iiface_rewrite(1)[OF nsimplers nnormalized wf_ipassmt(1) _ nospoofing] oiface_rewrite(1) nnf by simp subgoal using iiface_constrain(1)[OF nsimplers nnormalized wf_ipassmt(1), where p = p] nospoofing_alt oiface_rewrite(1) nnf by simp done qed end diff --git a/thys/Psi_Calculi/Bisim_Struct_Cong.thy b/thys/Psi_Calculi/Bisim_Struct_Cong.thy --- a/thys/Psi_Calculi/Bisim_Struct_Cong.thy +++ b/thys/Psi_Calculi/Bisim_Struct_Cong.thy @@ -1,1227 +1,1227 @@ (* Title: Psi-calculi Author/Maintainer: Jesper Bengtson (jebe@itu.dk), 2012 *) theory Bisim_Struct_Cong imports Bisim_Pres Sim_Struct_Cong Structural_Congruence begin context env begin lemma bisimParComm: fixes \ :: 'b and P :: "('a, 'b, 'c) psi" and Q :: "('a, 'b, 'c) psi" shows "\ \ P \ Q \ Q \ P" proof - let ?X = "{((\::'b), \\*xvec\((P::('a, 'b, 'c) psi) \ Q), \\*xvec\(Q \ P)) | xvec \ P Q. xvec \* \}" have "eqvt ?X" by(force simp add: eqvt_def pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst] eqvts) have "(\, P \ Q, Q \ P) \ ?X" apply auto by(rule_tac x="[]" in exI) auto thus ?thesis proof(coinduct rule: bisimWeakCoinduct) case(cStatEq \ PQ QP) from \(\, PQ, QP) \ ?X\ obtain xvec P Q where PFrQ: "PQ = \\*xvec\(P \ Q)" and QFrP: "QP = \\*xvec\(Q \ P)" and "xvec \* \" by auto obtain A\<^sub>P \\<^sub>P where FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" and "A\<^sub>P \* \" and "A\<^sub>P \* Q" by(rule_tac C="(\, Q)" in freshFrame) auto obtain A\<^sub>Q \\<^sub>Q where FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" and "A\<^sub>Q \* \" and "A\<^sub>Q \* A\<^sub>P" and "A\<^sub>Q \* \\<^sub>P" by(rule_tac C="(\, A\<^sub>P, \\<^sub>P)" in freshFrame) auto from FrQ \A\<^sub>Q \* A\<^sub>P\ \A\<^sub>P \* Q\ have "A\<^sub>P \* \\<^sub>Q" by(force dest: extractFrameFreshChain) have "\(xvec@A\<^sub>P@A\<^sub>Q), \ \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \(xvec@A\<^sub>Q@A\<^sub>P), \ \ \\<^sub>Q \ \\<^sub>P\" by(simp add: frameChainAppend) (metis frameResChainPres frameResChainComm frameNilStatEq compositionSym Associativity Commutativity FrameStatEqTrans) with FrP FrQ PFrQ QFrP \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>Q \* \\<^sub>P\ \A\<^sub>Q \* A\<^sub>P\ \xvec \* \\ \A\<^sub>P \* \\ \A\<^sub>Q \* \\ show ?case by(auto simp add: frameChainAppend) next case(cSim \ PQ QP) from \(\, PQ, QP) \ ?X\ obtain xvec P Q where PFrQ: "PQ = \\*xvec\(P \ Q)" and QFrP: "QP = \\*xvec\(Q \ P)" and "xvec \* \" by auto moreover have "\ \ \\*xvec\(P \ Q) \[?X] \\*xvec\(Q \ P)" proof - have "\ \ P \ Q \[?X] Q \ P" proof - note \eqvt ?X\ moreover have "\\ P Q. (\, P \ Q, Q \ P) \ ?X" apply auto by(rule_tac x="[]" in exI) auto moreover have "\\ P Q xvec. \(\, P, Q) \ ?X; xvec \* \\ \ (\, \\*xvec\P, \\*xvec\Q) \ ?X" apply(induct xvec, auto) by(rule_tac x="xvec@xveca" in exI) (auto simp add: resChainAppend) ultimately show ?thesis by(rule simParComm) qed moreover note \eqvt ?X\ \xvec \* \\ moreover have "\\ P Q x. \(\, P, Q) \ ?X; x \ \\ \ (\, \\x\P, \\x\Q) \ ?X" apply auto by(rule_tac x="x#xvec" in exI) auto ultimately show ?thesis by(rule resChainPres) qed ultimately show ?case by simp next case(cExt \ PQ QP \') from \(\, PQ, QP) \ ?X\ obtain xvec P Q where PFrQ: "PQ = \\*xvec\(P \ Q)" and QFrP: "QP = \\*xvec\(Q \ P)" and "xvec \* \" by auto obtain p where "(p \ xvec) \* \" and "(p \ xvec) \* P" and "(p \ xvec) \* Q" and "(p \ xvec) \* \'" and S: "(set p) \ (set xvec) \ (set(p \ xvec))" and "distinctPerm p" by(rule_tac c="(\, P, Q, \')" in name_list_avoiding) auto from \(p \ xvec) \* P\ \(p \ xvec) \* Q\ S have "\\*xvec\(P \ Q) = \\*(p \ xvec)\(p \ (P \ Q))" by(subst resChainAlpha) auto hence PQAlpha: "\\*xvec\(P \ Q) = \\*(p \ xvec)\((p \ P) \ (p \ Q))" by(simp add: eqvts) from \(p \ xvec) \* P\ \(p \ xvec) \* Q\ S have "\\*xvec\(Q \ P) = \\*(p \ xvec)\(p \ (Q \ P))" by(subst resChainAlpha) auto hence QPAlpha: "\\*xvec\(Q \ P) = \\*(p \ xvec)\((p \ Q) \ (p \ P))" by(simp add: eqvts) from \(p \ xvec) \* \\ \(p \ xvec) \* \'\ have "(\ \ \', \\*(p \ xvec)\((p \ P) \ (p \ Q)), \\*(p \ xvec)\((p \ Q) \ (p \ P))) \ ?X" by auto with PFrQ QFrP PQAlpha QPAlpha show ?case by simp next case(cSym \ PR QR) thus ?case by blast qed qed lemma bisimResComm: fixes x :: name and \ :: 'b and y :: name and P :: "('a, 'b, 'c) psi" shows "\ \ \\x\(\\y\P) \ \\y\(\\x\P)" proof(cases "x=y") case True thus ?thesis by(blast intro: bisimReflexive) next case False { fix x::name and y::name and P::"('a, 'b, 'c) psi" assume "x \ \" and "y \ \" let ?X = "{((\::'b), \\x\(\\y\(P::('a, 'b, 'c) psi)), \\y\(\\x\P)) | \ x y P. x \ \ \ y \ \}" from \x \ \\ \y \ \\ have "(\, \\x\(\\y\P), \\y\(\\x\P)) \ ?X" by auto hence "\ \ \\x\(\\y\P) \ \\y\(\\x\P)" proof(coinduct rule: bisimCoinduct) case(cStatEq \ xyP yxP) from \(\, xyP, yxP) \ ?X\ obtain x y P where "x \ \" and "y \ \" and "xyP = \\x\(\\y\P)" and "yxP = \\y\(\\x\P)" by auto moreover obtain A\<^sub>P \\<^sub>P where "extractFrame P = \A\<^sub>P, \\<^sub>P\" and "A\<^sub>P \* \" and "x \ A\<^sub>P" and "y \ A\<^sub>P" by(rule_tac C="(x, y, \)" in freshFrame) auto ultimately show ?case by(force intro: frameResComm FrameStatEqTrans) next case(cSim \ xyP yxP) from \(\, xyP, yxP) \ ?X\ obtain x y P where "x \ \" and "y \ \" and "xyP = \\x\(\\y\P)" and "yxP = \\y\(\\x\P)" by auto note \x \ \\ \y \ \\ moreover have "eqvt ?X" by(force simp add: eqvt_def pt_fresh_bij[OF pt_name_inst, OF at_name_inst]) hence "eqvt(?X \ bisim)" by auto moreover have "\\ P. (\, P, P) \ ?X \ bisim" by(blast intro: bisimReflexive) moreover have "\\ x y P. \x \ \; y \ \\ \ (\, \\x\(\\y\P), \\y\(\\x\P)) \ ?X \ bisim" by auto ultimately have "\ \ \\x\(\\y\P) \[(?X \ bisim)] \\y\(\\x\P)" by(rule resComm) with \xyP = \\x\(\\y\P)\ \yxP = \\y\(\\x\P)\ show ?case by simp next case(cExt \ xyP yxP \') from \(\, xyP, yxP) \ ?X\ obtain x y P where "x \ \" and "y \ \" and xyPeq: "xyP = \\x\(\\y\P)" and yxPeq: "yxP = \\y\(\\x\P)" by auto show ?case proof(case_tac "x=y") assume "x = y" with xyPeq yxPeq show ?case by(blast intro: bisimReflexive) next assume "x \ y" obtain x' where "x' \ \" and "x' \ \'" and "x' \ x" and "x' \ y" and "x' \ P" by(generate_fresh "name") (auto simp add: fresh_prod) obtain y' where "y' \ \" and "y' \ \'" and "y' \ x" and "x' \ y'" and "y' \ y" and "y' \ P" by(generate_fresh "name") (auto simp add: fresh_prod) with xyPeq \y' \ P\ \x' \ P\ \x \ y\ \x' \ y\ \y' \ x\ have "\\x\(\\y\P) = \\x'\(\\y'\([(x, x')] \ [(y, y')] \ P))" apply(subst alphaRes[of x']) apply(simp add: abs_fresh) by(subst alphaRes[of y' _ y]) (auto simp add: eqvts calc_atm) moreover with yxPeq \y' \ P\ \x' \ P\ \x \ y\ \x' \ y\ \y' \ x\ \x' \ y'\ have "\\y\(\\x\P) = \\y'\(\\x'\([(y, y')] \ [(x, x')] \ P))" apply(subst alphaRes[of y']) apply(simp add: abs_fresh) by(subst alphaRes[of x' _ x]) (auto simp add: eqvts calc_atm) with \x \ y\ \x' \ y\ \y' \ y\ \x' \ x\ \y' \ x\ \x' \ y'\ have "\\y\(\\x\P) = \\y'\(\\x'\([(x, x')] \ [(y, y')] \ P))" by(subst perm_compose) (simp add: calc_atm) moreover from \x' \ \\ \x' \ \'\ \y' \ \\ \y' \ \'\ have "(\ \ \', \\x'\(\\y'\([(x, x')] \ [(y, y')] \ P)), \\y'\(\\x'\([(x, x')] \ [(y, y')] \ P))) \ ?X" by auto ultimately show ?case using xyPeq yxPeq by simp qed next case(cSym \ xyP yxP) thus ?case by auto qed } moreover obtain x'::name where "x' \ \" and "x' \ P" and "x' \ x" and "x' \ y" by(generate_fresh "name") auto moreover obtain y'::name where "y' \ \" and "y' \ P" and "y' \ x" and "y' \ y" and "y' \ x'" by(generate_fresh "name") auto ultimately have "\ \ \\x'\(\\y'\([(y, y'), (x, x')] \ P)) \ \\y'\(\\x'\([(y, y'), (x, x')] \ P))" by auto thus ?thesis using \x' \ P\ \x' \ x\ \x' \ y\ \y' \ P\ \y' \ x\ \y' \ y\ \y' \ x'\ \x \ y\ apply(subst alphaRes[where x=x and y=x' and P=P], auto) apply(subst alphaRes[where x=y and y=y' and P=P], auto) apply(subst alphaRes[where x=x and y=x' and P="\\y'\([(y, y')] \ P)"], auto simp add: abs_fresh fresh_left) apply(subst alphaRes[where x=y and y=y' and P="\\x'\([(x, x')] \ P)"], auto simp add: abs_fresh fresh_left) by(subst perm_compose) (simp add: eqvts calc_atm) qed lemma bisimResComm': fixes x :: name and \ :: 'b and xvec :: "name list" and P :: "('a, 'b, 'c) psi" assumes "x \ \" and "xvec \* \" shows "\ \ \\x\(\\*xvec\P) \ \\*xvec\(\\x\P)" using assms by(induct xvec) (auto intro: bisimResComm bisimReflexive bisimResPres bisimTransitive) lemma bisimScopeExt: fixes x :: name and \ :: 'b and P :: "('a, 'b, 'c) psi" and Q :: "('a, 'b, 'c) psi" assumes "x \ P" shows "\ \ \\x\(P \ Q) \ P \ \\x\Q" proof - { fix x::name and Q :: "('a, 'b, 'c) psi" assume "x \ \" and "x \ P" let ?X1 = "{((\::'b), \\*xvec\(\\x\((P::('a, 'b, 'c) psi) \ Q)), \\*xvec\(P \ \\x\Q)) | \ xvec x P Q. x \ \ \ x \ P \ xvec \* \}" let ?X2 = "{((\::'b), \\*xvec\((P::('a, 'b, 'c) psi) \ \\x\Q), \\*xvec\(\\x\(P \ Q))) | \ xvec x P Q. x \ \ \ x \ P \ xvec \* \}" let ?X = "?X1 \ ?X2" from \x \ \\ \x \ P\ have "(\, \\x\(P \ Q), P \ \\x\Q) \ ?X" by(auto, rule_tac x="[]" in exI) (auto simp add: fresh_list_nil) moreover have "eqvt ?X" by(rule eqvtUnion) (fastforce simp add: eqvt_def eqvts pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst] pt_fresh_bij[OF pt_name_inst, OF at_name_inst])+ ultimately have "\ \ \\x\(P \ Q) \ P \ \\x\Q" proof(coinduct rule: transitiveCoinduct) case(cStatEq \ R T) show ?case proof(case_tac "(\, R, T) \ ?X1") assume "(\, R, T) \ ?X1" then obtain xvec x P Q where "R = \\*xvec\(\\x\(P \ Q))" and "T = \\*xvec\(P \ \\x\Q)" and "xvec \* \" and "x \ P" and "x \ \" by auto moreover obtain A\<^sub>P \\<^sub>P where FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" and "A\<^sub>P \* \" and "x \ A\<^sub>P" and "A\<^sub>P \* Q" by(rule_tac C="(\, x, Q)" in freshFrame) auto moreover obtain A\<^sub>Q \\<^sub>Q where FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" and "A\<^sub>Q \* \" and "x \ A\<^sub>Q" and "A\<^sub>Q \* A\<^sub>P" and "A\<^sub>Q \* \\<^sub>P" by(rule_tac C="(\, x, A\<^sub>P, \\<^sub>P)" in freshFrame) auto moreover from FrQ \A\<^sub>P \* Q\ \A\<^sub>Q \* A\<^sub>P\ have "A\<^sub>P \* \\<^sub>Q" by(drule_tac extractFrameFreshChain) auto moreover from \x \ P\ \x \ A\<^sub>P\ FrP have "x \ \\<^sub>P" by(drule_tac extractFrameFresh) auto ultimately show ?case by(force simp add: frameChainAppend intro: frameResComm' FrameStatEqTrans frameResChainPres) next assume "(\, R, T) \ ?X1" with \(\, R, T) \ ?X\ have "(\, R, T) \ ?X2" by blast then obtain xvec x P Q where "T = \\*xvec\(\\x\(P \ Q))" and "R = \\*xvec\(P \ \\x\Q)" and "xvec \* \" and "x \ P" and "x \ \" by auto moreover obtain A\<^sub>P \\<^sub>P where FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" and "A\<^sub>P \* \" and "x \ A\<^sub>P" and "A\<^sub>P \* Q" by(rule_tac C="(\, x, Q)" in freshFrame) auto moreover obtain A\<^sub>Q \\<^sub>Q where FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" and "A\<^sub>Q \* \" and "x \ A\<^sub>Q" and "A\<^sub>Q \* A\<^sub>P" and "A\<^sub>Q \* \\<^sub>P" by(rule_tac C="(\, x, A\<^sub>P, \\<^sub>P)" in freshFrame) auto moreover from FrQ \A\<^sub>P \* Q\ \A\<^sub>Q \* A\<^sub>P\ have "A\<^sub>P \* \\<^sub>Q" by(drule_tac extractFrameFreshChain) auto moreover from \x \ P\ \x \ A\<^sub>P\ FrP have "x \ \\<^sub>P" by(drule_tac extractFrameFresh) auto ultimately show ?case apply auto by(force simp add: frameChainAppend intro: frameResComm' FrameStatEqTrans frameResChainPres FrameStatEqSym) qed next case(cSim \ R T) let ?Y = "{(\, P, Q) | \ P P' Q' Q. \ \ P \ P' \ ((\, P', Q') \ ?X \ \ \ P' \ Q') \ \ \ Q' \ Q}" from \eqvt ?X\ have "eqvt ?Y" by blast have C1: "\\ R T y. \(\, R, T) \ ?Y; (y::name) \ \\ \ (\, \\y\R, \\y\T) \ ?Y" proof - fix \ R T y assume "(\, R, T) \ ?Y" then obtain R' T' where "\ \ R \ R'" and "(\, R', T') \ (?X \ bisim)" and "\ \ T' \ T" by fastforce assume "(y::name) \ \" show "(\, \\y\R, \\y\T) \ ?Y" proof(case_tac "(\, R', T') \ ?X") assume "(\, R', T') \ ?X" show ?thesis proof(case_tac "(\, R', T') \ ?X1") assume "(\, R', T') \ ?X1" then obtain xvec x P Q where R'eq: "R' = \\*xvec\(\\x\(P \ Q))" and T'eq: "T' = \\*xvec\(P \ \\x\Q)" and "xvec \* \" and "x \ P" and "x \ \" by auto from \\ \ R \ R'\ \y \ \\ have "\ \ \\y\R \ \\y\R'" by(rule bisimResPres) moreover from \xvec \* \\ \y \ \\ \x \ P\ \x \ \\ have "(\, \\*(y#xvec)\\\x\(P \ Q), \\*(y#xvec)\(P \ \\x\Q)) \ ?X1" by(force simp del: resChain.simps) with R'eq T'eq have "(\, \\y\R', \\y\T') \ ?X \ bisim" by simp moreover from \\ \ T' \ T\ \y \ \\ have "\ \ \\y\T' \ \\y\T" by(rule bisimResPres) ultimately show ?thesis by blast next assume "(\, R', T') \ ?X1" with \(\, R', T') \ ?X\ have "(\, R', T') \ ?X2" by blast then obtain xvec x P Q where T'eq: "T' = \\*xvec\(\\x\(P \ Q))" and R'eq: "R' = \\*xvec\(P \ \\x\Q)" and "xvec \* \" and "x \ P" and "x \ \" by auto from \\ \ R \ R'\ \y \ \\ have "\ \ \\y\R \ \\y\R'" by(rule bisimResPres) moreover from \xvec \* \\ \y \ \\ \x \ P\ \x \ \\ have "(\, \\*(y#xvec)\(P \ \\x\Q), \\*(y#xvec)\(\\x\(P \ Q))) \ ?X2" by(force simp del: resChain.simps) with R'eq T'eq have "(\, \\y\R', \\y\T') \ ?X \ bisim" by simp moreover from \\ \ T' \ T\ \y \ \\ have "\ \ \\y\T' \ \\y\T" by(rule bisimResPres) ultimately show ?thesis by blast qed next assume "(\, R', T') \ ?X" with \(\, R', T') \ ?X \ bisim\ have "\ \ R' \ T'" by blast with \\ \ R \ R'\ \\ \ T' \ T\ \y \ \\ show ?thesis by(blast dest: bisimResPres) qed qed show ?case proof(case_tac "(\, R, T) \ ?X1") assume "(\, R, T) \ ?X1" then obtain xvec x P Q where Req: "R = \\*xvec\(\\x\(P \ Q))" and Teq: "T = \\*xvec\(P \ \\x\Q)" and "xvec \* \" and "x \ P" and "x \ \" by auto have "\ \ \\*xvec\(\\x\(P \ Q)) \[?Y] \\*xvec\(P \ \\x\Q)" proof - have "\ \ \\x\(P \ Q) \[?Y] P \ \\x\Q" proof - note \x \ P\ \x \ \\ \eqvt ?Y\ moreover have "\\ P. (\, P, P) \ ?Y" by(blast intro: bisimReflexive) moreover have "\x \ P Q xvec. \x \ \; x \ P; xvec \* \\ \ (\, \\x\(\\*xvec\(P \ Q)), \\*xvec\(P \ \\x\Q)) \ ?Y" proof - fix x \ P Q xvec assume "(x::name) \ (\::'b)" and "x \ (P::('a, 'b, 'c) psi)" and "(xvec::name list) \* \" from \x \ \\ \xvec \* \\ have "\ \ \\x\(\\*xvec\(P \ Q)) \ \\*xvec\(\\x\(P \ Q))" by(rule bisimResComm') moreover from \xvec \* \\ \x \ \\ \x \ P\ have "(\, \\*xvec\(\\x\(P \ Q)), \\*xvec\(P \ \\x\Q)) \ ?X \ bisim" by blast ultimately show "(\, \\x\(\\*xvec\(P \ Q)), \\*xvec\(P \ \\x\Q)) \ ?Y" by(blast intro: bisimReflexive) qed moreover have "\\ xvec P x. \x \ \; xvec \* \\ \ (\, \\x\(\\*xvec\P), \\*xvec\(\\x\P)) \ ?Y" by(blast intro: bisimResComm' bisimReflexive) ultimately show ?thesis by(rule scopeExtLeft) qed thus ?thesis using \eqvt ?Y\ \xvec \* \\ C1 by(rule resChainPres) qed with Req Teq show ?case by simp next assume "(\, R, T) \ ?X1" with \(\, R, T) \ ?X\ have "(\, R, T) \ ?X2" by blast then obtain xvec x P Q where Teq: "T = \\*xvec\(\\x\(P \ Q))" and Req: "R = \\*xvec\(P \ \\x\Q)" and "xvec \* \" and "x \ P" and "x \ \" by auto have "\ \ \\*xvec\(P \ \\x\Q) \[?Y] \\*xvec\(\\x\(P \ Q))" proof - have "\ \ P \ \\x\Q \[?Y] \\x\(P \ Q)" proof - note \x \ P\ \x \ \\ \eqvt ?Y\ moreover have "\\ P. (\, P, P) \ ?Y" by(blast intro: bisimReflexive) moreover have "\x \ P Q xvec. \x \ \; x \ P; xvec \* \\ \ (\, \\*xvec\(P \ \\x\Q), \\x\(\\*xvec\(P \ Q))) \ ?Y" proof - fix x \ P Q xvec assume "(x::name) \ (\::'b)" and "x \ (P::('a, 'b, 'c) psi)" and "(xvec::name list) \* \" from \xvec \* \\ \x \ \\ \x \ P\ have "(\, \\*xvec\(P \ \\x\Q), \\*xvec\(\\x\(P \ Q))) \ ?X \ bisim" by blast moreover from \x \ \\ \xvec \* \\ have "\ \ \\*xvec\(\\x\(P \ Q)) \ \\x\(\\*xvec\(P \ Q))" by(blast intro: bisimResComm' bisimE) ultimately show "(\, \\*xvec\(P \ \\x\Q), \\x\(\\*xvec\(P \ Q))) \ ?Y" by(blast intro: bisimReflexive) qed ultimately show ?thesis by(rule scopeExtRight) qed thus ?thesis using \eqvt ?Y\ \xvec \* \\ C1 by(rule resChainPres) qed with Req Teq show ?case by simp qed next case(cExt \ R T \') show ?case proof(case_tac "(\, R, T) \ ?X1") assume "(\, R, T) \ ?X1" then obtain xvec x P Q where Req: "R = \\*xvec\(\\x\(P \ Q))" and Teq: "T = \\*xvec\(P \ \\x\Q)" and "xvec \* \" and "x \ P" and "x \ \" by auto obtain y::name where "y \ P" and "y \ Q" and "y \ xvec" and "y \ \" and "y \ \'" by(generate_fresh "name", auto simp add: fresh_prod) obtain p where "(p \ xvec) \* \" and "(p \ xvec) \* P" and "(p \ xvec) \* Q" and "(p \ xvec) \* \'" and "x \ (p \ xvec)" and "y \ (p \ xvec)" and S: "(set p) \ (set xvec) \ (set(p \ xvec))" and "distinctPerm p" by(rule_tac c="(\, P, Q, x, y, \')" in name_list_avoiding) auto from \y \ P\ have "(p \ y) \ (p \ P)" by(simp add: pt_fresh_bij[OF pt_name_inst, OF at_name_inst]) with S \y \ xvec\ \y \ (p \ xvec)\ have "y \ (p \ P)" by simp with \(p \ xvec) \* \\ \y \ \\ \(p \ xvec) \* \'\ \y \ \'\ have "(\ \ \', \\*(p \ xvec)\(\\y\((p \ P) \ (p \ [(x, y)] \ Q))), \\*(p \ xvec)\((p \ P) \ (\\y\(p \ [(x, y)] \ Q)))) \ ?X" by auto moreover from Req \(p \ xvec) \* P\ \(p \ xvec) \* Q\ \y \ xvec\ \y \ (p \ xvec)\ \x \ (p \ xvec)\ \y \ P\ \y \ Q\ \x \ P\ S have "R = \\*(p \ xvec)\(\\y\((p \ P) \ (p \ [(x, y)] \ Q)))" apply(erule_tac rev_mp) apply(subst alphaRes[of y]) apply(clarsimp simp add: eqvts) apply(subst resChainAlpha[of p]) by(auto simp add: eqvts) moreover from Teq \(p \ xvec) \* P\ \(p \ xvec) \* Q\ \y \ xvec\ \y \ (p \ xvec)\ \x \ (p \ xvec)\ \y \ P\ \y \ Q\ \x \ P\ S have "T = \\*(p \ xvec)\((p \ P) \ \\y\(p \ [(x, y)] \ Q))" apply(erule_tac rev_mp) apply(subst alphaRes[of y]) apply(clarsimp simp add: eqvts) apply(subst resChainAlpha[of p]) by(auto simp add: eqvts) ultimately show ?case by blast next assume "(\, R, T) \ ?X1" with \(\, R, T) \ ?X\ have "(\, R, T) \ ?X2" by blast then obtain xvec x P Q where Teq: "T = \\*xvec\(\\x\(P \ Q))" and Req: "R = \\*xvec\(P \ \\x\Q)" and "xvec \* \" and "x \ P" and "x \ \" by auto obtain y::name where "y \ P" and "y \ Q" and "y \ xvec" and "y \ \" and "y \ \'" by(generate_fresh "name", auto simp add: fresh_prod) obtain p where "(p \ xvec) \* \" and "(p \ xvec) \* P" and "(p \ xvec) \* Q" and "(p \ xvec) \* \'" and "x \ (p \ xvec)" and "y \ (p \ xvec)" and S: "(set p) \ (set xvec) \ (set(p \ xvec))" and "distinctPerm p" by(rule_tac c="(\, P, Q, x, y, \')" in name_list_avoiding) auto from \y \ P\ have "(p \ y) \ (p \ P)" by(simp add: pt_fresh_bij[OF pt_name_inst, OF at_name_inst]) with S \y \ xvec\ \y \ (p \ xvec)\ have "y \ (p \ P)" by simp with \(p \ xvec) \* \\ \y \ \\ \(p \ xvec) \* \'\ \y \ \'\ have "(\ \ \', \\*(p \ xvec)\((p \ P) \ \\y\(p \ [(x, y)] \ Q)), \\*(p \ xvec)\(\\y\((p \ P) \ (p \ [(x, y)] \ Q)))) \ ?X2" by auto moreover from Teq \(p \ xvec) \* P\ \(p \ xvec) \* Q\ \y \ xvec\ \y \ (p \ xvec)\ \x \ (p \ xvec)\ \y \ P\ \y \ Q\ \x \ P\ S have "T = \\*(p \ xvec)\(\\y\((p \ P) \ (p \ [(x, y)] \ Q)))" apply(erule_tac rev_mp) apply(subst alphaRes[of y]) apply(clarsimp simp add: eqvts) apply(subst resChainAlpha[of p]) by(auto simp add: eqvts) moreover from Req \(p \ xvec) \* P\ \(p \ xvec) \* Q\ \y \ xvec\ \y \ (p \ xvec)\ \x \ (p \ xvec)\ \y \ P\ \y \ Q\ \x \ P\ S have "R = \\*(p \ xvec)\((p \ P) \ \\y\(p \ [(x, y)] \ Q))" apply(erule_tac rev_mp) apply(subst alphaRes[of y]) apply(clarsimp simp add: eqvts) apply(subst resChainAlpha[of p]) by(auto simp add: eqvts) ultimately show ?case by blast qed next case(cSym \ P Q) thus ?case by(blast dest: bisimE) qed } moreover obtain y::name where "y \ \" and "y \ P" "y \ Q" by(generate_fresh "name") auto ultimately have "\ \ \\y\(P \ ([(x, y)] \ Q)) \ P \ \\y\([(x, y)] \ Q)" by auto thus ?thesis using assms \y \ P\ \y \ Q\ apply(subst alphaRes[where x=x and y=y and P=Q], auto) by(subst alphaRes[where x=x and y=y and P="P \ Q"]) auto qed lemma bisimScopeExtChain: fixes xvec :: "name list" and \ :: 'b and P :: "('a, 'b, 'c) psi" and Q :: "('a, 'b, 'c) psi" assumes "xvec \* \" and "xvec \* P" shows "\ \ \\*xvec\(P \ Q) \ P \ (\\*xvec\Q)" using assms by(induct xvec) (auto intro: bisimScopeExt bisimReflexive bisimTransitive bisimResPres) lemma bisimParAssoc: fixes \ :: 'b and P :: "('a, 'b, 'c) psi" and Q :: "('a, 'b, 'c) psi" and R :: "('a, 'b, 'c) psi" shows "\ \ (P \ Q) \ R \ P \ (Q \ R)" proof - let ?X = "{(\, \\*xvec\((P \ Q) \ R), \\*xvec\(P \ (Q \ R))) | \ xvec P Q R. xvec \* \}" let ?Y = "{(\, P, Q) | \ P P' Q' Q. \ \ P \ P' \ (\, P', Q') \ ?X \ \ \ Q' \ Q}" have "(\, (P \ Q) \ R, P \ (Q \ R)) \ ?X" by(auto, rule_tac x="[]" in exI) auto moreover have "eqvt ?X" by(force simp add: eqvt_def simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst] eqvts) ultimately show ?thesis proof(coinduct rule: weakTransitiveCoinduct') case(cStatEq \ PQR PQR') from \(\, PQR, PQR') \ ?X\ obtain xvec P Q R where "xvec \* \" and "PQR = \\*xvec\((P \ Q) \ R)" and "PQR' = \\*xvec\(P \ (Q \ R))" by auto moreover obtain A\<^sub>P \\<^sub>P where FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" and "A\<^sub>P \* \" and "A\<^sub>P \* Q" and "A\<^sub>P \* R" by(rule_tac C="(\, Q, R)" in freshFrame) auto moreover obtain A\<^sub>Q \\<^sub>Q where FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" and "A\<^sub>Q \* \" and "A\<^sub>Q \* A\<^sub>P" and "A\<^sub>Q \* \\<^sub>P" and "A\<^sub>Q \* R" by(rule_tac C="(\, A\<^sub>P, \\<^sub>P, R)" in freshFrame) auto moreover obtain A\<^sub>R \\<^sub>R where FrR: "extractFrame R = \A\<^sub>R, \\<^sub>R\" and "A\<^sub>R \* \" and "A\<^sub>R \* A\<^sub>P" and "A\<^sub>R \* \\<^sub>P" and "A\<^sub>R \* A\<^sub>Q" and "A\<^sub>R \* \\<^sub>Q" by(rule_tac C="(\, A\<^sub>P, \\<^sub>P, A\<^sub>Q, \\<^sub>Q)" in freshFrame) auto moreover from FrQ \A\<^sub>P \* Q\ \A\<^sub>Q \* A\<^sub>P\ have "A\<^sub>P \* \\<^sub>Q" by(drule_tac extractFrameFreshChain) auto moreover from FrR \A\<^sub>P \* R\ \A\<^sub>R \* A\<^sub>P\ have "A\<^sub>P \* \\<^sub>R" by(drule_tac extractFrameFreshChain) auto moreover from FrR \A\<^sub>Q \* R\ \A\<^sub>R \* A\<^sub>Q\ have "A\<^sub>Q \* \\<^sub>R" by(drule_tac extractFrameFreshChain) auto ultimately show ?case using freshCompChain by auto (metis frameChainAppend compositionSym Associativity frameNilStatEq frameResChainPres) next case(cSim \ T S) from \(\, T, S) \ ?X\ obtain xvec P Q R where "xvec \* \" and TEq: "T = \\*xvec\((P \ Q) \ R)" and SEq: "S = \\*xvec\(P \ (Q \ R))" by auto from \eqvt ?X\have "eqvt ?Y" by blast have C1: "\\ T S yvec. \(\, T, S) \ ?Y; yvec \* \\ \ (\, \\*yvec\T, \\*yvec\S) \ ?Y" proof - fix \ T S yvec assume "(\, T, S) \ ?Y" then obtain T' S' where "\ \ T \ T'" and "(\, T', S') \ ?X" and "\ \ S' \ S" by fastforce assume "(yvec::name list) \* \" from \(\, T', S') \ ?X\ obtain xvec P Q R where T'eq: "T' = \\*xvec\((P \ Q) \ R)" and S'eq: "S' = \\*xvec\(P \ (Q \ R))" and "xvec \* \" by auto from \\ \ T \ T'\ \yvec \* \\ have "\ \ \\*yvec\T \ \\*yvec\T'" by(rule bisimResChainPres) moreover from \xvec \* \\ \yvec \* \\ have "(\, \\*(yvec@xvec)\((P \ Q) \ R), \\*(yvec@xvec)\(P \ (Q \ R))) \ ?X" by force with T'eq S'eq have "(\, \\*yvec\T', \\*yvec\S') \ ?X" by(simp add: resChainAppend) moreover from \\ \ S' \ S\ \yvec \* \\ have "\ \ \\*yvec\S' \ \\*yvec\S" by(rule bisimResChainPres) ultimately show "(\, \\*yvec\T, \\*yvec\S) \ ?Y" by blast qed have C2: "\\ T S y. \(\, T, S) \ ?Y; y \ \\ \ (\, \\y\T, \\y\S) \ ?Y" - by(drule_tac yvec="[y]" in C1) auto + by(drule_tac yvec2="[y]" in C1) auto have "\ \ \\*xvec\((P \ Q) \ R) \[?Y] \\*xvec\(P \ (Q \ R))" proof - have "\ \ (P \ Q) \ R \[?Y] P \ (Q \ R)" proof - note \eqvt ?Y\ moreover have "\\ P Q R. (\, (P \ Q) \ R, P \ (Q \ R)) \ ?Y" proof - fix \ P Q R have "(\::'b, ((P::('a, 'b, 'c) psi) \ Q) \ R, P \ (Q \ R)) \ ?X" by(auto, rule_tac x="[]" in exI) auto thus "(\, (P \ Q) \ R, P \ (Q \ R)) \ ?Y" by(blast intro: bisimReflexive) qed moreover have "\xvec \ P Q R. \xvec \* \; xvec \* P\ \ (\, \\*xvec\((P \ Q) \ R), P \ (\\*xvec\(Q \ R))) \ ?Y" proof - fix xvec \ P Q R assume "(xvec::name list) \* (\::'b)" and "xvec \* (P::('a, 'b, 'c) psi)" from \xvec \* \\ have "(\, \\*xvec\((P \ Q) \ R), \\*xvec\(P \ (Q \ R))) \ ?X" by blast moreover from \xvec \* \\ \xvec \* P\ have "\ \ \\*xvec\(P \ (Q \ R)) \ P \ (\\*xvec\(Q \ R))" by(rule bisimScopeExtChain) ultimately show "(\, \\*xvec\((P \ Q) \ R), P \ (\\*xvec\(Q \ R))) \ ?Y" by(blast intro: bisimReflexive) qed moreover have "\xvec \ P Q R. \xvec \* \; xvec \* R\ \ (\, (\\*xvec\(P \ Q)) \ R, \\*xvec\(P \ (Q \ R))) \ ?Y" proof - fix xvec \ P Q R assume "(xvec::name list) \* (\::'b)" and "xvec \* (R::('a, 'b, 'c) psi)" have "\ \ (\\*xvec\(P \ Q)) \ R \ R \ (\\*xvec\(P \ Q))" by(rule bisimParComm) moreover from \xvec \* \\ \xvec \* R\ have "\ \ \\*xvec\(R \ (P \ Q)) \ R \ (\\*xvec\(P \ Q))" by(rule bisimScopeExtChain) hence "\ \ R \ (\\*xvec\(P \ Q)) \ \\*xvec\(R \ (P \ Q))" by(rule bisimE) moreover from \xvec \* \\ have "\ \ \\*xvec\(R \ (P \ Q)) \ \\*xvec\((P \ Q) \ R)" by(metis bisimResChainPres bisimParComm) moreover from \xvec \* \\ have "(\, \\*xvec\((P \ Q) \ R), \\*xvec\(P \ (Q \ R))) \ ?X" by blast ultimately show "(\, (\\*xvec\(P \ Q)) \ R, \\*xvec\(P \ (Q \ R))) \ ?Y" by(blast dest: bisimTransitive intro: bisimReflexive) qed ultimately show ?thesis using C1 by(rule parAssocLeft) qed thus ?thesis using \eqvt ?Y\ \xvec \* \\ C2 by(rule resChainPres) qed with TEq SEq show ?case by simp next case(cExt \ T S \') from \(\, T, S) \ ?X\ obtain xvec P Q R where "xvec \* \" and TEq: "T = \\*xvec\((P \ Q) \ R)" and SEq: "S = \\*xvec\(P \ (Q \ R))" by auto obtain p where "(p \ xvec) \* \" and "(p \ xvec) \* P" and "(p \ xvec) \* Q" and "(p \ xvec) \* R" and "(p \ xvec) \* \'" and S: "(set p) \ (set xvec) \ (set(p \ xvec))" and "distinctPerm p" by(rule_tac c="(\, P, Q, R, \')" in name_list_avoiding) auto from \(p \ xvec) \* \\ \(p \ xvec) \* \'\ have "(\ \ \', \\*(p \ xvec)\(((p \ P) \ (p \ Q)) \ (p \ R)), \\*(p \ xvec)\((p \ P) \ ((p \ Q) \ (p \ R)))) \ ?X" by auto moreover from TEq \(p \ xvec) \* P\ \(p \ xvec) \* Q\ \(p \ xvec) \* R\ S have "T = \\*(p \ xvec)\(((p \ P) \ (p \ Q)) \ (p \ R))" apply auto by(subst resChainAlpha[of p]) auto moreover from SEq \(p \ xvec) \* P\ \(p \ xvec) \* Q\ \(p \ xvec) \* R\ S have "S = \\*(p \ xvec)\((p \ P) \ ((p \ Q) \ (p \ R)))" apply auto by(subst resChainAlpha[of p]) auto ultimately show ?case by simp next case(cSym \ T S) from \(\, T, S) \ ?X\ obtain xvec P Q R where "xvec \* \" and TEq: "T = \\*xvec\((P \ Q) \ R)" and SEq: "\\*xvec\(P \ (Q \ R)) = S" by auto from \xvec \* \\ have "\ \ \\*xvec\(P \ (Q \ R)) \ \\*xvec\((R \ Q) \ P)" by(metis bisimParComm bisimParPres bisimTransitive bisimResChainPres) moreover from \xvec \* \\ have "(\, \\*xvec\((R \ Q) \ P), \\*xvec\(R \ (Q \ P))) \ ?X" by blast moreover from \xvec \* \\ have "\ \ \\*xvec\(R \ (Q \ P)) \ \\*xvec\((P \ Q) \ R)" by(metis bisimParComm bisimParPres bisimTransitive bisimResChainPres) ultimately show ?case using TEq SEq by(blast dest: bisimTransitive) qed qed lemma bisimParNil: fixes P :: "('a, 'b, 'c) psi" shows "\ \ P \ \ \ P" proof - let ?X1 = "{(\, P \ \, P) | \ P. True}" let ?X2 = "{(\, P, P \ \) | \ P. True}" let ?X = "?X1 \ ?X2" have "eqvt ?X" by(auto simp add: eqvt_def) have "(\, P \ \, P) \ ?X" by simp thus ?thesis proof(coinduct rule: bisimWeakCoinduct) case(cStatEq \ Q R) show ?case proof(case_tac "(\, Q, R) \ ?X1") assume "(\, Q, R) \ ?X1" then obtain P where "Q = P \ \" and "R = P" by auto moreover obtain A\<^sub>P \\<^sub>P where "extractFrame P = \A\<^sub>P, \\<^sub>P\" and "A\<^sub>P \* \" by(rule freshFrame) ultimately show ?case apply auto by(metis frameResChainPres frameNilStatEq Identity Associativity AssertionStatEqTrans Commutativity) next assume "(\, Q, R) \ ?X1" with \(\, Q, R) \ ?X\ have "(\, Q, R) \ ?X2" by blast then obtain P where "Q = P" and "R = P \ \" by auto moreover obtain A\<^sub>P \\<^sub>P where "extractFrame P = \A\<^sub>P, \\<^sub>P\" and "A\<^sub>P \* \" by(rule freshFrame) ultimately show ?case apply auto by(metis frameResChainPres frameNilStatEq Identity Associativity AssertionStatEqTrans AssertionStatEqSym Commutativity) qed next case(cSim \ Q R) thus ?case using \eqvt ?X\ by(auto intro: parNilLeft parNilRight) next case(cExt \ Q R \') thus ?case by auto next case(cSym \ Q R) thus ?case by auto qed qed lemma bisimResNil: fixes x :: name and \ :: 'b shows "\ \ \\x\\ \ \" proof - { fix x::name assume "x \ \" have "\ \ \\x\\ \ \" proof - let ?X1 = "{(\, \\x\\, \) | \ x. x \ \}" let ?X2 = "{(\, \, \\x\\) | \ x. x \ \}" let ?X = "?X1 \ ?X2" from \x \ \\ have "(\, \\x\\, \) \ ?X" by auto thus ?thesis proof(coinduct rule: bisimWeakCoinduct) case(cStatEq \ P Q) thus ?case using freshComp by(force intro: frameResFresh FrameStatEqSym) next case(cSim \ P Q) thus ?case by(force intro: resNilLeft resNilRight) next case(cExt \ P Q \') obtain y where "y \ \" and "y \ \'" and "y \ x" by(generate_fresh "name") (auto simp add: fresh_prod) show ?case proof(case_tac "(\, P, Q) \ ?X1") assume "(\, P, Q) \ ?X1" then obtain x where "P = \\x\\" and "Q = \" by auto moreover have "\\x\\ = \\y\ \" by(subst alphaRes) auto ultimately show ?case using \y \ \\ \y \ \'\ by auto next assume "(\, P, Q) \ ?X1" with \(\, P, Q) \ ?X\ have "(\, P, Q) \ ?X2" by auto then obtain x where "Q = \\x\\" and "P = \" by auto moreover have "\\x\\ = \\y\ \" by(subst alphaRes) auto ultimately show ?case using \y \ \\ \y \ \'\ by auto qed next case(cSym \ P Q) thus ?case by auto qed qed } moreover obtain y::name where "y \ \" by(generate_fresh "name") auto ultimately have "\ \ \\y\\ \ \" by auto thus ?thesis by(subst alphaRes[where x=x and y=y]) auto qed lemma bisimOutputPushRes: fixes x :: name and \ :: 'b and M :: 'a and N :: 'a and P :: "('a, 'b, 'c) psi" assumes "x \ M" and "x \ N" shows "\ \ \\x\(M\N\.P) \ M\N\.\\x\P" proof - { fix x::name and P::"('a, 'b, 'c) psi" assume "x \ \" and "x \ M" and "x \ N" let ?X1 = "{(\, \\x\(M\N\.P), M\N\.\\x\P) | \ x M N P. x \ \ \ x \ M \ x \ N}" let ?X2 = "{(\, M\N\.\\x\P, \\x\(M\N\.P)) | \ x M N P. x \ \ \ x \ M \ x \ N}" let ?X = "?X1 \ ?X2" have "eqvt ?X" by(rule_tac eqvtUnion) (force simp add: eqvt_def pt_fresh_bij[OF pt_name_inst, OF at_name_inst] eqvts)+ from \x \ \\ \x \ M\ \x \ N\ have "(\, \\x\(M\N\.P), M\N\.\\x\P) \ ?X" by auto hence "\ \ \\x\(M\N\.P) \ M\N\.\\x\P" proof(coinduct rule: bisimCoinduct) case(cStatEq \ Q R) thus ?case using freshComp by(force intro: frameResFresh FrameStatEqSym) next case(cSim \ Q R) thus ?case using \eqvt ?X\ by(fastforce intro: outputPushResLeft outputPushResRight bisimReflexive) next case(cExt \ Q R \') show ?case proof(case_tac "(\, Q, R) \ ?X1") assume "(\, Q, R) \ ?X1" then obtain x M N P where Qeq: "Q = \\x\(M\N\.P)" and Req: "R = M\N\.\\x\P" and "x \ \" and "x \ M" and "x \ N" by auto obtain y::name where "y \ \" and "y \ \'" and "y \ M" and "y \ N" and "y \ P" by(generate_fresh "name") (auto simp add: fresh_prod) moreover hence "(\ \ \', \\y\(M\N\.([(x, y)] \ P)), M\N\.\\y\([(x, y)] \ P)) \ ?X" by auto moreover from Qeq \x \ M\ \y \ M\ \x \ N\ \y \ N\ \y \ P\ have "Q = \\y\(M\N\.([(x, y)] \ P))" apply auto by(subst alphaRes[of y]) (auto simp add: eqvts) moreover from Req \y \ P\ have "R = M\N\.\\y\([(x, y)] \ P)" apply auto by(subst alphaRes[of y]) (auto simp add: eqvts) ultimately show ?case by blast next assume "(\, Q, R) \ ?X1" with \(\, Q, R) \ ?X\ have "(\, Q, R) \ ?X2" by blast then obtain x M N P where Req: "R = \\x\(M\N\.P)" and Qeq: "Q = M\N\.\\x\P" and "x \ \" and "x \ M" and "x \ N" by auto obtain y::name where "y \ \" and "y \ \'" and "y \ M" and "y \ N" and "y \ P" by(generate_fresh "name") (auto simp add: fresh_prod) moreover hence "(\ \ \', \\y\(M\N\.([(x, y)] \ P)), M\N\.\\y\([(x, y)] \ P)) \ ?X" by auto moreover from Req \x \ M\ \y \ M\ \x \ N\ \y \ N\ \y \ P\ have "R = \\y\(M\N\.([(x, y)] \ P))" apply auto by(subst alphaRes[of y]) (auto simp add: eqvts) moreover from Qeq \y \ P\ have "Q = M\N\.\\y\([(x, y)] \ P)" apply auto by(subst alphaRes[of y]) (auto simp add: eqvts) ultimately show ?case by blast qed next case(cSym \ R Q) thus ?case by blast qed } moreover obtain y::name where "y \ \" and "y \ M" and "y \ N" "y \ P" by(generate_fresh "name") auto ultimately have "\ \ \\y\(M\N\.([(x, y)] \ P)) \ M\N\.\\y\([(x, y)] \ P)" by auto thus ?thesis using assms \y \ P\ \y \ M\ \y \ N\ apply(subst alphaRes[where x=x and y=y and P=P], auto) by(subst alphaRes[where x=x and y=y and P="M\N\.P"]) auto qed lemma bisimInputPushRes: fixes x :: name and \ :: 'b and M :: 'a and xvec :: "name list" and N :: 'a and P :: "('a, 'b, 'c) psi" assumes "x \ M" and "x \ xvec" and "x \ N" shows "\ \ \\x\(M\\*xvec N\.P) \ M\\*xvec N\.\\x\P" proof - { fix x::name and P::"('a, 'b, 'c) psi" assume "x \ \" and "x \ M" and "x \ N" and "x \ xvec" let ?X1 = "{(\, \\x\(M\\*xvec N\.P), M\\*xvec N\.\\x\P) | \ x M xvec N P. x \ \ \ x \ M \ x \ xvec \ x \ N}" let ?X2 = "{(\, M\\*xvec N\.\\x\P, \\x\(M\\*xvec N\.P)) | \ x M xvec N P. x \ \ \ x \ M \ x \ xvec \ x \ N}" let ?X = "?X1 \ ?X2" have "eqvt ?X" by(rule_tac eqvtUnion) (force simp add: eqvt_def pt_fresh_bij[OF pt_name_inst, OF at_name_inst] eqvts)+ from \x \ \\ \x \ M\ \x \ xvec\ \x \ N\ have "(\, \\x\(M\\*xvec N\.P), M\\*xvec N\.\\x\P) \ ?X" by blast hence "\ \ \\x\(M\\*xvec N\.P) \ M\\*xvec N\.\\x\P" proof(coinduct rule: bisimCoinduct) case(cStatEq \ Q R) thus ?case using freshComp by(force intro: frameResFresh FrameStatEqSym) next case(cSim \ Q R) thus ?case using \eqvt ?X\ by(fastforce intro: inputPushResLeft inputPushResRight bisimReflexive) next case(cExt \ Q R \') show ?case proof(case_tac "(\, Q, R) \ ?X1") assume "(\, Q, R) \ ?X1" then obtain x M xvec N P where Qeq: "Q = \\x\(M\\*xvec N\.P)" and Req: "R = M\\*xvec N\.\\x\P" and "x \ \" and "x \ M" and "x \ xvec" and "x \ N" by auto obtain y::name where "y \ \" and "y \ \'" and "y \ M" and "y \ N" and "y \ P" and "y \ xvec" by(generate_fresh "name") (auto simp add: fresh_prod) moreover hence "(\ \ \', \\y\(M\\*xvec N\.([(x, y)] \ P)), M\\*xvec N\.\\y\([(x, y)] \ P)) \ ?X" by fastforce moreover from Qeq \x \ M\ \y \ M\ \x \ xvec\ \y \ xvec\ \x \ N\ \y \ N\ \y \ P\ have "Q = \\y\(M\\*xvec N\.([(x, y)] \ P))" apply auto by(subst alphaRes[of y]) (auto simp add: eqvts inputChainFresh) moreover from Req \y \ P\ have "R = M\\*xvec N \.\\y\([(x, y)] \ P)" apply auto by(subst alphaRes[of y]) (auto simp add: eqvts) ultimately show ?case by blast next assume "(\, Q, R) \ ?X1" with \(\, Q, R) \ ?X\ have "(\, Q, R) \ ?X2" by blast then obtain x M xvec N P where Req: "R = \\x\(M\\*xvec N\.P)" and Qeq: "Q = M\\*xvec N\.\\x\P" and "x \ \" and "x \ M" and "x \ xvec" and "x \ N" by auto obtain y::name where "y \ \" and "y \ \'" and "y \ M" and "y \ N" and "y \ P" and "y \ xvec" by(generate_fresh "name") (auto simp add: fresh_prod) moreover hence "(\ \ \', \\y\(M\\*xvec N\.([(x, y)] \ P)), M\\*xvec N\.\\y\([(x, y)] \ P)) \ ?X" by fastforce moreover from Req \x \ M\ \y \ M\ \x \ xvec\ \y \ xvec\ \x \ N\ \y \ N\ \y \ P\ have "R = \\y\(M\\*xvec N\.([(x, y)] \ P))" apply auto by(subst alphaRes[of y]) (auto simp add: eqvts inputChainFresh) moreover from Qeq \y \ P\ have "Q = M\\*xvec N \.\\y\([(x, y)] \ P)" apply auto by(subst alphaRes[of y]) (auto simp add: eqvts) ultimately show ?case by blast qed next case(cSym \ R Q) thus ?case by blast qed } moreover obtain y::name where "y \ \" and "y \ M" and "y \ N" and "y \ P" and "y \ xvec" by(generate_fresh "name") auto ultimately have "\ \ \\y\(M\\*xvec N\.([(x, y)] \ P)) \ M\\*xvec N\.\\y\([(x, y)] \ P)" by auto thus ?thesis using assms \y \ P\ \y \ M\ \y \ N\ \y \ xvec\ apply(subst alphaRes[where x=x and y=y and P=P], auto) by(subst alphaRes[where x=x and y=y and P="M\\*xvec N\.P"]) (auto simp add: inputChainFresh eqvts) qed lemma bisimCasePushRes: fixes x :: name and \ :: 'b and Cs :: "('c \ ('a, 'b, 'c) psi) list" assumes "x \ (map fst Cs)" shows "\ \ \\x\(Cases Cs) \ Cases(map (\(\, P). (\, \\x\P)) Cs)" proof - { fix x::name and Cs::"('c \ ('a, 'b, 'c) psi) list" assume "x \ \" and "x \ (map fst Cs)" let ?X1 = "{(\, \\x\(Cases Cs), Cases(map (\(\, P). (\, \\x\P)) Cs)) | \ x Cs. x \ \ \ x \ (map fst Cs)}" let ?X2 = "{(\, Cases(map (\(\, P). (\, \\x\P)) Cs), \\x\(Cases Cs)) | \ x Cs. x \ \ \ x \ (map fst Cs)}" let ?X = "?X1 \ ?X2" have "eqvt ?X" apply(rule_tac eqvtUnion) apply(auto simp add: eqvt_def eqvts) apply(rule_tac x="p \ x" in exI) apply(rule_tac x="p \ Cs" in exI) apply(perm_extend_simp) apply(auto simp add: eqvts) apply(simp add: pt_fresh_bij[OF pt_name_inst, OF at_name_inst]) apply(drule_tac pi=p in pt_fresh_bij1[OF pt_name_inst, OF at_name_inst]) apply(drule_tac pi=p in pt_fresh_bij1[OF pt_name_inst, OF at_name_inst]) apply(simp add: eqvts) apply(perm_extend_simp) apply(simp add: eqvts) apply(rule_tac x="p \ x" in exI) apply(rule_tac x="p \ Cs" in exI) apply auto apply(perm_extend_simp) apply(simp add: pt_fresh_bij[OF pt_name_inst, OF at_name_inst]) apply(drule_tac pi=p in pt_fresh_bij1[OF pt_name_inst, OF at_name_inst]) apply(drule_tac pi=p in pt_fresh_bij1[OF pt_name_inst, OF at_name_inst]) apply(simp add: eqvts) apply(perm_extend_simp) by(simp add: eqvts) from \x \ \\ \x \ map fst Cs\ have "(\, \\x\(Cases Cs), Cases(map (\(\, P). (\, \\x\P)) Cs)) \ ?X" by auto hence "\ \ \\x\(Cases Cs) \ Cases(map (\(\, P). (\, \\x\P)) Cs)" proof(coinduct rule: bisimCoinduct) case(cStatEq \ Q R) thus ?case using freshComp by(force intro: frameResFresh FrameStatEqSym) next case(cSim \ Q R) thus ?case using \eqvt ?X\ by(fastforce intro: casePushResLeft casePushResRight bisimReflexive) next case(cExt \ Q R \') show ?case proof(case_tac "(\, Q, R) \ ?X1") assume "(\, Q, R) \ ?X1" then obtain x Cs where Qeq: "Q = \\x\(Cases Cs)" and Req: "R = Cases(map (\(\, P). (\, \\x\P)) Cs)" and "x \ \" and "x \ (map fst Cs)" by blast obtain y::name where "y \ \" and "y \ \'" and "y \ Cs" by(generate_fresh "name") (auto simp add: fresh_prod) from \y \ Cs\ \x \ (map fst Cs)\ have "y \ map fst ([(x, y)] \ Cs)" by(induct Cs) (auto simp add: fresh_list_cons fresh_list_nil) moreover with \y \ \\ \y \ \'\ have "(\ \ \', \\y\(Cases ([(x, y)] \ Cs)), Cases(map (\(\, P). (\, \\y\P)) ([(x, y)] \ Cs))) \ ?X" by auto moreover from Qeq \y \ Cs\ have "Q = \\y\(Cases([(x, y)] \ Cs))" apply auto by(subst alphaRes[of y]) (auto simp add: eqvts) moreover from Req \y \ Cs\ \x \ (map fst Cs)\ have "R = Cases(map (\(\, P). (\, \\y\P)) ([(x, y)] \ Cs))" by(induct Cs arbitrary: R) (auto simp add: fresh_list_cons fresh_prod alphaRes) ultimately show ?case by blast next assume "(\, Q, R) \ ?X1" with \(\, Q, R) \ ?X\ have "(\, Q, R) \ ?X2" by blast then obtain x Cs where Req: "R = \\x\(Cases Cs)" and Qeq: "Q = Cases(map (\(\, P). (\, \\x\P)) Cs)" and "x \ \" and "x \ (map fst Cs)" by blast obtain y::name where "y \ \" and "y \ \'" and "y \ Cs" by(generate_fresh "name") (auto simp add: fresh_prod) from \y \ Cs\ \x \ (map fst Cs)\ have "y \ map fst ([(x, y)] \ Cs)" by(induct Cs) (auto simp add: fresh_list_cons fresh_list_nil) moreover with \y \ \\ \y \ \'\ have "(\ \ \', \\y\(Cases ([(x, y)] \ Cs)), Cases(map (\(\, P). (\, \\y\P)) ([(x, y)] \ Cs))) \ ?X" by auto moreover from Req \y \ Cs\ have "R = \\y\(Cases([(x, y)] \ Cs))" apply auto by(subst alphaRes[of y]) (auto simp add: eqvts) moreover from Qeq \y \ Cs\ \x \ (map fst Cs)\ have "Q = Cases(map (\(\, P). (\, \\y\P)) ([(x, y)] \ Cs))" by(induct Cs arbitrary: Q) (auto simp add: fresh_list_cons fresh_prod alphaRes) ultimately show ?case by blast qed next case(cSym \ R Q) thus ?case by blast qed } moreover obtain y::name where "y \ \" and "y \ Cs" by(generate_fresh "name") auto moreover from \x \ map fst Cs\ have "y \ map fst([(x, y)] \ Cs)" by(induct Cs) (auto simp add: fresh_left calc_atm) ultimately have "\ \ \\y\(Cases ([(x, y)] \ Cs)) \ Cases(map (\(\, P). (\, \\y\P)) ([(x, y)] \ Cs))" by auto moreover from \y \ Cs\ have "\\y\(Cases ([(x, y)] \ Cs)) = \\x\(Cases Cs)" by(simp add: alphaRes eqvts) moreover from \x \ map fst Cs\ \y \ Cs\ have "Cases(map (\(\, P). (\, \\y\P)) ([(x, y)] \ Cs)) = Cases(map (\(\, P). (\, \\x\P)) Cs)" by(induct Cs) (auto simp add: alphaRes) ultimately show ?thesis by auto qed lemma bangExt: fixes \ :: 'b and P :: "('a, 'b, 'c) psi" assumes "guarded P" shows "\ \ !P \ P \ !P" proof - let ?X = "{(\, !P, P \ !P) | \ P. guarded P} \ {(\, P \ !P, !P) | \ P. guarded P}" from \guarded P\ have "(\, !P, P \ !P) \ ?X" by auto thus ?thesis proof(coinduct rule: bisimCoinduct) case(cStatEq \ Q R) from \(\, Q, R) \ ?X\ obtain P where Eq: "(Q = !P \ R = P \ !P) \ (Q = P \ !P \ R = !P)" and "guarded P" by auto obtain A\<^sub>P \\<^sub>P where FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" and "A\<^sub>P \* \" by(rule freshFrame) from FrP \guarded P\ have "\\<^sub>P \ SBottom'" by(blast dest: guardedStatEq) from \\\<^sub>P \ SBottom'\ have "\ \ SBottom' \ \ \ \\<^sub>P \ SBottom'" by(metis Identity Composition AssertionStatEqTrans Commutativity AssertionStatEqSym) hence "\A\<^sub>P, \ \ SBottom'\ \\<^sub>F \A\<^sub>P, \ \ \\<^sub>P \ SBottom'\" by(force intro: frameResChainPres) moreover from \A\<^sub>P \* \\ have "\\, \ \ SBottom'\ \\<^sub>F \A\<^sub>P, \ \ SBottom'\" by(rule_tac FrameStatEqSym) (fastforce intro: frameResFreshChain) ultimately show ?case using Eq \A\<^sub>P \* \\ FrP by auto (blast dest: FrameStatEqTrans FrameStatEqSym)+ next case(cSim \ Q R) thus ?case by(auto intro: bangExtLeft bangExtRight bisimReflexive) next case(cExt \ Q R) thus ?case by auto next case(cSym \ Q R) thus ?case by auto qed qed lemma bisimParPresSym: fixes \ :: 'b and P :: "('a, 'b, 'c) psi" and Q :: "('a, 'b, 'c) psi" and R :: "('a, 'b, 'c) psi" assumes "\ \ P \ Q" shows "\ \ R \ P \ R \ Q" using assms by(metis bisimParComm bisimParPres bisimTransitive) lemma bisimScopeExtSym: fixes x :: name and Q :: "('a, 'b, 'c) psi" and P :: "('a, 'b, 'c) psi" assumes "x \ \" and "x \ Q" shows "\ \ \\x\(P \ Q) \ (\\x\P) \ Q" using assms by(metis bisimScopeExt bisimTransitive bisimParComm bisimSymmetric bisimResPres) lemma bisimScopeExtChainSym: fixes xvec :: "name list" and Q :: "('a, 'b, 'c) psi" and P :: "('a, 'b, 'c) psi" assumes "xvec \* \" and "xvec \* Q" shows "\ \ \\*xvec\(P \ Q) \ (\\*xvec\P) \ Q" using assms by(induct xvec) (auto intro: bisimScopeExtSym bisimReflexive bisimTransitive bisimResPres) lemma bisimParPresAuxSym: fixes \ :: 'b and P :: "('a, 'b, 'c) psi" and Q :: "('a, 'b, 'c) psi" and R :: "('a, 'b, 'c) psi" assumes "\ \ \\<^sub>R \ P \ Q" and "extractFrame R = \A\<^sub>R, \\<^sub>R\" and "A\<^sub>R \* \" and "A\<^sub>R \* P" and "A\<^sub>R \* Q" shows "\ \ R \ P \ R \ Q" using assms by(metis bisimParComm bisimParPresAux bisimTransitive) lemma bangDerivative: fixes \ :: 'b and P :: "('a, 'b, 'c) psi" and \ :: "'a action" and P' :: "('a, 'b, 'c) psi" assumes "\ \ !P \\ \ P'" and "\ \ P \ Q" and "bn \ \* \" and "bn \ \* P" and "bn \ \* Q" and "bn \ \* subject \" and "guarded Q" obtains Q' R T where "\ \ !Q \\ \ Q'" and "\ \ P' \ R \ !P" and "\ \ Q' \ T \ !Q" and "\ \ R \ T" and "((supp R)::name set) \ supp P'" and "((supp T)::name set) \ supp Q'" proof - from \\ \ !P \\ \ P'\ have "guarded P" apply - by(ind_cases "\ \ !P \\ \ P'") (auto simp add: psi.inject) assume "\Q' R T. \\ \ !Q \\ \ Q'; \ \ P' \ R \ !P; \ \ Q' \ T \ !Q; \ \ R \ T; ((supp R)::name set) \ supp P'; ((supp T)::name set) \ supp Q'\ \ thesis" moreover from \\ \ !P \\ \ P'\ \bn \ \* subject \\ \bn \ \* \\ \bn \ \* P\ \bn \ \* Q\ \\ \ P \ Q\ \guarded Q\ have "\Q' T R . \ \ !Q \\ \ Q' \ \ \ P' \ R \ !P \ \ \ Q' \ T \ !Q \ \ \ R \ T \ ((supp R)::name set) \ supp P' \ ((supp T)::name set) \ supp Q'" proof(nominal_induct avoiding: Q rule: bangInduct') case(cAlpha \ P' p Q) then obtain Q' T R where QTrans: "\ \ !Q \\ \ Q'" and "\ \ P' \ R \ (P \ !P)" and "\ \ Q' \ T \ !Q" and "\ \ R \ T" and suppR: "((supp R)::name set) \ supp P'" and suppT: "((supp T)::name set) \ supp Q'" by blast from QTrans have "distinct(bn \)" by(rule boundOutputDistinct) have S: "set p \ set(bn \) \ set(bn(p \ \))" by fact from QTrans \bn(p \ \) \* Q\ \bn(p \ \) \* \\ \bn \ \* subject \\ \distinct(bn \)\ have "bn(p \ \) \* Q'" by(drule_tac freeFreshChainDerivative) simp+ with QTrans \bn(p \ \) \* \\ S \bn \ \* subject \\ have "\ \ !Q \(p \ \) \ (p \ Q')" by(force simp add: residualAlpha) moreover from \\ \ P' \ R \ (P \ !P)\ have "(p \ \) \ (p \ P') \ (p \ (R \ (P \ !P)))" by(rule bisimClosed) with \bn \ \* \\ \bn \ \* P\ \bn(p \ \) \* \\ \bn(p \ \) \* P\ S have "\ \ (p \ P') \ (p \ R) \ (P \ !P)" by(simp add: eqvts) moreover from \\ \ Q' \ T \ !Q\ have "(p \ \) \ (p \ Q') \ (p \ (T \ !Q))" by(rule bisimClosed) with \bn \ \* \\ \bn \ \* Q\ \bn(p \ \) \* \\ \bn(p \ \) \* Q\ S have "\ \ (p \ Q') \ (p \ T) \ !Q" by(simp add: eqvts) moreover from \\ \ R \ T\ have "(p \ \) \ (p \ R) \ (p \ T)" by(rule bisimClosed) with \bn \ \* \\ \bn(p \ \) \* \\ S have "\ \ (p \ R) \ (p \ T)" by(simp add: eqvts) moreover from suppR have "((supp(p \ R))::name set) \ supp(p \ P')" apply(erule_tac rev_mp) by(subst subsetClosed[of p, symmetric]) (simp add: eqvts) moreover from suppT have "((supp(p \ T))::name set) \ supp(p \ Q')" apply(erule_tac rev_mp) by(subst subsetClosed[of p, symmetric]) (simp add: eqvts) ultimately show ?case by blast next case(cPar1 \ P' Q) from \\ \ P \ Q\ \\ \ P \\ \ P'\ \bn \ \* \\ \bn \ \* Q\ obtain Q' where QTrans: "\ \ Q \\ \ Q'" and "\ \ P' \ Q'" by(blast dest: bisimE simE) from QTrans have "\ \ SBottom' \ Q \\ \ Q'" by(metis statEqTransition Identity AssertionStatEqSym) hence "\ \ Q \ !Q \\ \ (Q' \ !Q)" using \bn \ \* Q\ by(rule_tac Par1) (assumption | simp)+ hence "\ \ !Q \\ \ (Q' \ !Q)" using \guarded Q\ by(rule Bang) moreover from \guarded P\ have "\ \ P' \ !P \ P' \ (P \ !P)" by(metis bangExt bisimParPresSym) moreover have "\ \ Q' \ !Q \ Q' \ !Q" by(rule bisimReflexive) ultimately show ?case using \\ \ P' \ Q'\ by(force simp add: psi.supp) next case(cPar2 \ P' Q) then obtain Q' T R where QTrans: "\ \ !Q \\ \ Q'" and "\ \ P' \ R \ !P" and "\ \ Q' \ T \ !Q" and "\ \ R \ T" and suppR: "((supp R)::name set) \ supp P'" and suppT: "((supp T)::name set) \ supp Q'" by blast note QTrans from \\ \ P' \ R \ !P\ have "\ \ P \ P' \ R \ (P \ !P)" by(metis bisimParPresSym bisimParComm bisimTransitive bisimParAssoc) with QTrans show ?case using \\ \ Q' \ T \ !Q\ \\ \ R \ T\ suppR suppT by(force simp add: psi.supp) next case(cComm1 M N P' K xvec P'' Q) from \\ \ P \ Q\ have "\ \ Q \[bisim] P" by(metis bisimE) with \\ \ P \M\N\ \ P'\ obtain Q' where QTrans: "\ \ Q \M\N\ \ Q'" and "\ \ Q' \ P'" by(force dest: simE) from QTrans have "\ \ SBottom' \ Q \M\N\ \ Q'" by(metis statEqTransition Identity AssertionStatEqSym) moreover obtain A\<^sub>Q \\<^sub>Q where FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" and "A\<^sub>Q \* \" and "A\<^sub>Q \* Q" and "A\<^sub>Q \* M" by(rule_tac C="(\, Q, M)" in freshFrame) auto note FrQ moreover from FrQ \guarded Q\ have "\\<^sub>Q \ SBottom'" by(blast dest: guardedStatEq) from \\ \ !P \K\\*xvec\\N\ \ P''\ \xvec \* K\ \\ \ P \ Q\ \xvec \* \\ \xvec \* P\ \xvec \* Q\ \guarded Q\ obtain Q'' T R where QTrans': "\ \ !Q \K\\*xvec\\N\ \ Q''" and "\ \ P'' \ R \ !P" and "\ \ Q'' \ T \ !Q" and "\ \ R \ T" and suppR: "((supp R)::name set) \ supp P''" and suppT: "((supp T)::name set) \ supp Q''" using cComm1 by fastforce from QTrans' \\\<^sub>Q \ SBottom'\ have "\ \ \\<^sub>Q \ !Q \K\\*xvec\\N\ \ Q''" by(metis statEqTransition Identity compositionSym AssertionStatEqSym) moreover from \\ \ M \ K\ \\\<^sub>Q \ SBottom'\ have "\ \ \\<^sub>Q \ SBottom' \ M \ K" by(metis statEqEnt Identity compositionSym AssertionStatEqSym) ultimately have "\ \ Q \ !Q \\ \ (\\*xvec\(Q' \ Q''))" using \A\<^sub>Q \* \\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* M\ \xvec \* Q\ by(rule_tac Comm1) (assumption | simp)+ hence "\ \ !Q \\ \ (\\*xvec\(Q' \ Q''))" using \guarded Q\ by(rule Bang) moreover from \\ \ P'' \ R \ !P\ \guarded P\ \xvec \* \\ have "\ \ \\*xvec\(P' \ P'') \ \\*xvec\((P' \ R) \ (P \ !P))" by(metis bisimParPresSym bangExt bisimTransitive bisimParAssoc bisimSymmetric bisimResChainPres) with \xvec \* \\ \xvec \* P\ have "\ \ \\*xvec\(P' \ P'') \ (\\*xvec\(P' \ R)) \ (P \ !P)" by(metis bisimScopeExtChainSym bisimTransitive psiFreshVec) moreover from \\ \ Q'' \ T \ !Q\ \xvec \* \\ \xvec \* Q\ have "\ \ \\*xvec\(Q' \ Q'') \ (\\*xvec\(Q' \ T)) \ !Q" by(metis bisimParPresSym bisimTransitive bisimParAssoc bisimSymmetric bisimResChainPres bisimScopeExtChainSym psiFreshVec) moreover from \\ \ R \ T\ \\ \ Q' \ P'\ \xvec \* \\ have "\ \ \\*xvec\(P' \ R) \ \\*xvec\(Q' \ T)" by(metis bisimParPresSym bisimTransitive bisimResChainPres bisimParComm bisimE(4)) moreover from suppR have "((supp(\\*xvec\(P' \ R)))::name set) \ supp((\\*xvec\(P' \ P'')))" by(auto simp add: psi.supp resChainSupp) moreover from suppT have "((supp(\\*xvec\(Q' \ T)))::name set) \ supp((\\*xvec\(Q' \ Q'')))" by(auto simp add: psi.supp resChainSupp) ultimately show ?case by blast next case(cComm2 M xvec N P' K P'' Q) from \\ \ P \ Q\ \\ \ P \M\\*xvec\\N\ \ P'\ \xvec \* \\ \xvec \* Q\ obtain Q' where QTrans: "\ \ Q \M\\*xvec\\N\ \ Q'" and "\ \ P' \ Q'" by(metis bisimE simE bn.simps) from QTrans have "\ \ SBottom' \ Q \M\\*xvec\\N\ \ Q'" by(metis statEqTransition Identity AssertionStatEqSym) moreover obtain A\<^sub>Q \\<^sub>Q where FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" and "A\<^sub>Q \* \" and "A\<^sub>Q \* Q" and "A\<^sub>Q \* M" by(rule_tac C="(\, Q, M)" in freshFrame) auto note FrQ moreover from FrQ \guarded Q\ have "\\<^sub>Q \ SBottom'" by(blast dest: guardedStatEq) from \\ \ !P \K\N\ \ P''\ \\ \ P \ Q\ \guarded Q\ obtain Q'' T R where QTrans': "\ \ !Q \K\N\ \ Q''" and "\ \ P'' \ R \ !P" and "\ \ Q'' \ T \ !Q" and "\ \ R \ T" and suppR: "((supp R)::name set) \ supp P''" and suppT: "((supp T)::name set) \ supp Q''" using cComm2 by fastforce from QTrans' \\\<^sub>Q \ SBottom'\ have "\ \ \\<^sub>Q \ !Q \K\N\ \ Q''" by(metis statEqTransition Identity compositionSym AssertionStatEqSym) moreover from \\ \ M \ K\ \\\<^sub>Q \ SBottom'\ have "\ \ \\<^sub>Q \ SBottom' \ M \ K" by(metis statEqEnt Identity compositionSym AssertionStatEqSym) ultimately have "\ \ Q \ !Q \\ \ (\\*xvec\(Q' \ Q''))" using \A\<^sub>Q \* \\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* M\ \xvec \* Q\ by(rule_tac Comm2) (assumption | simp)+ hence "\ \ !Q \\ \ (\\*xvec\(Q' \ Q''))" using \guarded Q\ by(rule Bang) moreover from \\ \ P'' \ R \ !P\ \guarded P\ \xvec \* \\ have "\ \ \\*xvec\(P' \ P'') \ \\*xvec\((P' \ R) \ (P \ !P))" by(metis bisimParPresSym bangExt bisimTransitive bisimParAssoc bisimSymmetric bisimResChainPres) with \xvec \* \\ \xvec \* P\ have "\ \ \\*xvec\(P' \ P'') \ (\\*xvec\(P' \ R)) \ (P \ !P)" by(metis bisimScopeExtChainSym bisimTransitive psiFreshVec) moreover from \\ \ Q'' \ T \ !Q\ \xvec \* \\ \xvec \* Q\ have "\ \ \\*xvec\(Q' \ Q'') \ (\\*xvec\(Q' \ T)) \ !Q" by(metis bisimParPresSym bisimTransitive bisimParAssoc bisimSymmetric bisimResChainPres bisimScopeExtChainSym psiFreshVec) moreover from \\ \ R \ T\ \\ \ P' \ Q'\ \xvec \* \\ have "\ \ \\*xvec\(P' \ R) \ \\*xvec\(Q' \ T)" by(metis bisimParPresSym bisimTransitive bisimResChainPres bisimParComm) moreover from suppR have "((supp(\\*xvec\(P' \ R)))::name set) \ supp((\\*xvec\(P' \ P'')))" by(auto simp add: psi.supp resChainSupp) moreover from suppT have "((supp(\\*xvec\(Q' \ T)))::name set) \ supp((\\*xvec\(Q' \ Q'')))" by(auto simp add: psi.supp resChainSupp) ultimately show ?case by blast next case(cBang \ P' Q) then obtain Q' T R where QTrans: "\ \ !Q \\ \ Q'" and "\ \ P' \ R \ (P \ !P)" and "\ \ Q' \ T \ !Q" and "\ \ R \ T" and suppR: "((supp R)::name set) \ supp P'" and suppT: "((supp T)::name set) \ supp Q'" by blast from \\ \ P' \ R \ (P \ !P)\ \guarded P\ have "\ \ P' \ R \ !P" by(metis bangExt bisimParPresSym bisimTransitive bisimSymmetric) with QTrans show ?case using \\ \ Q' \ T \ !Q\ \\ \ R \ T\ suppR suppT by blast qed ultimately show ?thesis by blast qed lemma structCongBisim: fixes P :: "('a, 'b, 'c) psi" and Q :: "('a, 'b, 'c) psi" assumes "P \\<^sub>s Q" shows "P \ Q" using assms by(induct rule: structCong.induct) (auto intro: bisimReflexive bisimSymmetric bisimTransitive bisimParComm bisimParAssoc bisimParNil bisimResNil bisimResComm bisimScopeExt bisimCasePushRes bisimInputPushRes bisimOutputPushRes bangExt) lemma bisimBangPres: fixes \ :: 'b and P :: "('a, 'b, 'c) psi" and Q :: "('a, 'b, 'c) psi" assumes "\ \ P \ Q" and "guarded P" and "guarded Q" shows "\ \ !P \ !Q" proof - let ?X = "{(\, R \ !P, R \ !Q) | \ P Q R. \ \ P \ Q \ guarded P \ guarded Q}" let ?Y = "{(\, P, Q) | \ P P' Q' Q. \ \ P \ P' \ (\, P', Q') \ ?X \ \ \ Q' \ Q}" from assms have "(\, \ \ !P, \ \ !Q) \ ?X" by(blast intro: bisimReflexive) moreover have "eqvt ?X" apply(auto simp add: eqvt_def) apply(drule_tac p=p in bisimClosed) by fastforce ultimately have "\ \ \ \ !P \ \ \ !Q" proof(coinduct rule: weakTransitiveCoinduct) case(cStatEq \ P Q) thus ?case by auto next case(cSim \ RP RQ) from \(\, RP, RQ) \ ?X\ obtain P Q R where "\ \ P \ Q" and "guarded P" and "guarded Q" and "RP = R \ !P" and "RQ = R \ !Q" by auto note \\ \ P \ Q\ moreover from \eqvt ?X\ have "eqvt ?Y" by blast moreover note \guarded P\ \guarded Q\ bisimE(2) bisimE(3) bisimE(4) statEqBisim bisimClosed bisimParAssoc[THEN bisimSymmetric] bisimParPres bisimParPresAuxSym bisimResChainPres bisimScopeExtChainSym bisimTransitive moreover have "\\ P Q R T. \\ \ P \ Q; (\, Q, R) \ ?Y; \ \ R \ T\ \ (\, P, T) \ ?Y" by auto (metis bisimTransitive) moreover have "\\ P Q R. \\ \ P \ Q; guarded P; guarded Q\ \ (\, R \ !P, R \ !Q) \ ?Y" by(blast intro: bisimReflexive) moreover have "\\ P \ P' Q. \\ \ !P \\ \ P'; \ \ P \ Q; bn \ \* \; bn \ \* P; bn \ \* Q; guarded Q; bn \ \* subject \\ \ \Q' R T. \ \ !Q \\ \ Q' \ \ \ P' \ R \ !P \ \ \ Q' \ T \ !Q \ \ \ R \ T \ ((supp R)::name set) \ supp P' \ ((supp T)::name set) \ supp Q'" by(blast elim: bangDerivative) ultimately have "\ \ R \ !P \[?Y] R \ !Q" by(rule bangPres) with \RP = R \ !P\ \RQ = R \ !Q\ show ?case by blast next case(cExt \ RP RQ \') thus ?case by(blast dest: bisimE) next case(cSym \ RP RQ) thus ?case by(blast dest: bisimE) qed thus ?thesis by(metis bisimTransitive bisimParNil bisimSymmetric bisimParComm) qed end end diff --git a/thys/QHLProver/Complex_Matrix.thy b/thys/QHLProver/Complex_Matrix.thy --- a/thys/QHLProver/Complex_Matrix.thy +++ b/thys/QHLProver/Complex_Matrix.thy @@ -1,2352 +1,2352 @@ section \Complex matrices\ theory Complex_Matrix imports "Jordan_Normal_Form.Matrix" "Jordan_Normal_Form.Conjugate" "Jordan_Normal_Form.Jordan_Normal_Form_Existence" begin subsection \Trace of a matrix\ definition trace :: "'a::ring mat \ 'a" where "trace A = (\ i \ {0 ..< dim_row A}. A $$ (i,i))" lemma trace_zero [simp]: "trace (0\<^sub>m n n) = 0" by (simp add: trace_def) lemma trace_id [simp]: "trace (1\<^sub>m n) = n" by (simp add: trace_def) lemma trace_comm: fixes A B :: "'a::comm_ring mat" assumes A: "A \ carrier_mat n n" and B: "B \ carrier_mat n n" shows "trace (A * B) = trace (B * A)" proof (simp add: trace_def) have "(\i = 0..i = 0..j = 0.. = (\j = 0..i = 0.. = (\j = 0.. row B j)" by (metis (no_types, lifting) A B atLeastLessThan_iff carrier_matD index_col index_row scalar_prod_def sum.cong) also have "\ = (\j = 0.. col A j)" apply (rule sum.cong) apply auto apply (subst comm_scalar_prod[where n=n]) apply auto using assms by auto also have "\ = (\j = 0..i = 0..i = 0.. carrier_mat n n" and B: "B \ carrier_mat n n" shows "trace (A + B) = trace A + trace B" (is "?lhs = ?rhs") proof - have "?lhs = (\i=0.. = (\i=0..i=0..i=0..i=0..i=0..i=0.. carrier_mat n n" and B: "B \ carrier_mat n n" shows "trace (A - B) = trace A - trace B" (is "?lhs = ?rhs") proof - have "?lhs = (\i=0.. = (\i=0..i=0..i=0..i=0..i=0..i=0.. carrier_mat n n" shows "trace (c \\<^sub>m A) = c * trace A" proof - have "trace (c \\<^sub>m A) = (\i = 0.. = c * (\i = 0.. = c * trace A" unfolding trace_def by auto ultimately show ?thesis by auto qed subsection \Conjugate of a vector\ lemma conjugate_scalar_prod: fixes v w :: "'a::conjugatable_ring vec" assumes "dim_vec v = dim_vec w" shows "conjugate (v \ w) = conjugate v \ conjugate w" using assms by (simp add: scalar_prod_def sum_conjugate conjugate_dist_mul) subsection \Inner product\ abbreviation inner_prod :: "'a vec \ 'a vec \ 'a :: conjugatable_ring" where "inner_prod v w \ w \c v" lemma conjugate_scalar_prod_Im [simp]: "Im (v \c v) = 0" by (simp add: scalar_prod_def conjugate_vec_def sum.neutral) lemma conjugate_scalar_prod_Re [simp]: "Re (v \c v) \ 0" by (simp add: scalar_prod_def conjugate_vec_def sum_nonneg) lemma self_cscalar_prod_geq_0: fixes v :: "'a::conjugatable_ordered_field vec" shows "v \c v \ 0" by (auto simp add: scalar_prod_def, rule sum_nonneg, rule conjugate_square_positive) lemma inner_prod_distrib_left: fixes u v w :: "('a::conjugatable_field) vec" assumes dimu: "u \ carrier_vec n" and dimv:"v \ carrier_vec n" and dimw: "w \ carrier_vec n" shows "inner_prod (v + w) u = inner_prod v u + inner_prod w u" (is "?lhs = ?rhs") proof - have dimcv: "conjugate v \ carrier_vec n" and dimcw: "conjugate w \ carrier_vec n" using assms by auto have dimvw: "conjugate (v + w) \ carrier_vec n" using assms by auto have "u \ (conjugate (v + w)) = u \ conjugate v + u \ conjugate w" using dimv dimw dimu dimcv dimcw by (metis conjugate_add_vec scalar_prod_add_distrib) then show ?thesis by auto qed lemma inner_prod_distrib_right: fixes u v w :: "('a::conjugatable_field) vec" assumes dimu: "u \ carrier_vec n" and dimv:"v \ carrier_vec n" and dimw: "w \ carrier_vec n" shows "inner_prod u (v + w) = inner_prod u v + inner_prod u w" (is "?lhs = ?rhs") proof - have dimvw: "v + w \ carrier_vec n" using assms by auto have dimcu: "conjugate u \ carrier_vec n" using assms by auto have "(v + w) \ (conjugate u) = v \ conjugate u + w \ conjugate u" apply (simp add: comm_scalar_prod[OF dimvw dimcu]) apply (simp add: scalar_prod_add_distrib[OF dimcu dimv dimw]) apply (insert dimv dimw dimcu, simp add: comm_scalar_prod[of _ n]) done then show ?thesis by auto qed lemma inner_prod_minus_distrib_right: fixes u v w :: "('a::conjugatable_field) vec" assumes dimu: "u \ carrier_vec n" and dimv:"v \ carrier_vec n" and dimw: "w \ carrier_vec n" shows "inner_prod u (v - w) = inner_prod u v - inner_prod u w" (is "?lhs = ?rhs") proof - have dimvw: "v - w \ carrier_vec n" using assms by auto have dimcu: "conjugate u \ carrier_vec n" using assms by auto have "(v - w) \ (conjugate u) = v \ conjugate u - w \ conjugate u" apply (simp add: comm_scalar_prod[OF dimvw dimcu]) apply (simp add: scalar_prod_minus_distrib[OF dimcu dimv dimw]) apply (insert dimv dimw dimcu, simp add: comm_scalar_prod[of _ n]) done then show ?thesis by auto qed lemma inner_prod_smult_right: fixes u v :: "complex vec" assumes dimu: "u \ carrier_vec n" and dimv:"v \ carrier_vec n" shows "inner_prod (a \\<^sub>v u) v = conjugate a * inner_prod u v" (is "?lhs = ?rhs") using assms apply (simp add: scalar_prod_def conjugate_dist_mul) apply (subst sum_distrib_left) by (rule sum.cong, auto) lemma inner_prod_smult_left: fixes u v :: "complex vec" assumes dimu: "u \ carrier_vec n" and dimv: "v \ carrier_vec n" shows "inner_prod u (a \\<^sub>v v) = a * inner_prod u v" (is "?lhs = ?rhs") using assms apply (simp add: scalar_prod_def) apply (subst sum_distrib_left) by (rule sum.cong, auto) lemma inner_prod_smult_left_right: fixes u v :: "complex vec" assumes dimu: "u \ carrier_vec n" and dimv: "v \ carrier_vec n" shows "inner_prod (a \\<^sub>v u) (b \\<^sub>v v) = conjugate a * b * inner_prod u v" (is "?lhs = ?rhs") using assms apply (simp add: scalar_prod_def) apply (subst sum_distrib_left) by (rule sum.cong, auto) lemma inner_prod_swap: fixes x y :: "complex vec" assumes "y \ carrier_vec n" and "x \ carrier_vec n" shows "inner_prod y x = conjugate (inner_prod x y)" apply (simp add: scalar_prod_def) apply (rule sum.cong) using assms by auto text \Cauchy-Schwarz theorem for complex vectors. This is analogous to aux\_Cauchy and Cauchy\_Schwarz\_ineq in Generalizations2.thy in QR\_Decomposition. Consider merging and moving to Isabelle library.\ lemma aux_Cauchy: fixes x y :: "complex vec" assumes "x \ carrier_vec n" and "y \ carrier_vec n" shows "0 \ inner_prod x x + a * (inner_prod x y) + (cnj a) * ((cnj (inner_prod x y)) + a * (inner_prod y y))" proof - have "(inner_prod (x+ a \\<^sub>v y) (x+a \\<^sub>v y)) = (inner_prod (x+a \\<^sub>v y) x) + (inner_prod (x+a \\<^sub>v y) (a \\<^sub>v y))" apply (subst inner_prod_distrib_right) using assms by auto also have "\ = inner_prod x x + (a) * (inner_prod x y) + cnj a * ((cnj (inner_prod x y)) + (a) * (inner_prod y y))" apply (subst (1 2) inner_prod_distrib_left[of _ n]) apply (auto simp add: assms) apply (subst (1 2) inner_prod_smult_right[of _ n]) apply (auto simp add: assms) apply (subst inner_prod_smult_left[of _ n]) apply (auto simp add: assms) apply (subst inner_prod_swap[of y n x]) apply (auto simp add: assms) unfolding distrib_left by auto finally show ?thesis by (metis self_cscalar_prod_geq_0) qed lemma Cauchy_Schwarz_complex_vec: fixes x y :: "complex vec" assumes "x \ carrier_vec n" and "y \ carrier_vec n" shows "inner_prod x y * inner_prod y x \ inner_prod x x * inner_prod y y" proof - define cnj_a where "cnj_a = - (inner_prod x y)/ cnj (inner_prod y y)" define a where "a = cnj (cnj_a)" have cnj_rw: "(cnj a) = cnj_a" unfolding a_def by (simp) have rw_0: "cnj (inner_prod x y) + a * (inner_prod y y) = 0" unfolding a_def cnj_a_def using assms(1) assms(2) conjugate_square_eq_0_vec by fastforce have "0 \ (inner_prod x x + a * (inner_prod x y) + (cnj a) * ((cnj (inner_prod x y)) + a * (inner_prod y y)))" using aux_Cauchy assms by auto also have "\ = (inner_prod x x + a * (inner_prod x y))" unfolding rw_0 by auto also have "\ = (inner_prod x x - (inner_prod x y) * cnj (inner_prod x y) / (inner_prod y y))" unfolding a_def cnj_a_def by simp finally have " 0 \ (inner_prod x x - (inner_prod x y) * cnj (inner_prod x y) / (inner_prod y y)) " . hence "0 \ (inner_prod x x - (inner_prod x y) * cnj (inner_prod x y) / (inner_prod y y)) * (inner_prod y y)" by (auto simp: less_eq_complex_def) also have "\ = ((inner_prod x x)*(inner_prod y y) - (inner_prod x y) * cnj (inner_prod x y))" by (smt add.inverse_neutral add_diff_cancel diff_0 diff_divide_eq_iff divide_cancel_right mult_eq_0_iff nonzero_mult_div_cancel_right rw_0) finally have "(inner_prod x y) * cnj (inner_prod x y) \ (inner_prod x x)*(inner_prod y y)" by auto then show ?thesis apply (subst inner_prod_swap[of y n x]) by (auto simp add: assms) qed subsection \Hermitian adjoint of a matrix\ abbreviation adjoint where "adjoint \ mat_adjoint" lemma adjoint_dim_row [simp]: "dim_row (adjoint A) = dim_col A" by (simp add: mat_adjoint_def) lemma adjoint_dim_col [simp]: "dim_col (adjoint A) = dim_row A" by (simp add: mat_adjoint_def) lemma adjoint_dim: "A \ carrier_mat n n \ adjoint A \ carrier_mat n n" using adjoint_dim_col adjoint_dim_row by blast lemma adjoint_def: "adjoint A = mat (dim_col A) (dim_row A) (\(i,j). conjugate (A $$ (j,i)))" unfolding mat_adjoint_def mat_of_rows_def by auto lemma adjoint_eval: assumes "i < dim_col A" "j < dim_row A" shows "(adjoint A) $$ (i,j) = conjugate (A $$ (j,i))" using assms by (simp add: adjoint_def) lemma adjoint_row: assumes "i < dim_col A" shows "row (adjoint A) i = conjugate (col A i)" apply (rule eq_vecI) using assms by (auto simp add: adjoint_eval) lemma adjoint_col: assumes "i < dim_row A" shows "col (adjoint A) i = conjugate (row A i)" apply (rule eq_vecI) using assms by (auto simp add: adjoint_eval) text \The identity = \ lemma adjoint_def_alter: fixes v w :: "'a::conjugatable_field vec" and A :: "'a::conjugatable_field mat" assumes dims: "v \ carrier_vec n" "w \ carrier_vec m" "A \ carrier_mat n m" shows "inner_prod v (A *\<^sub>v w) = inner_prod (adjoint A *\<^sub>v v) w" (is "?lhs = ?rhs") proof - from dims have "?lhs = (\i=0..j=0..i=0..j=0..m n) = (1\<^sub>m n::complex mat)" apply (rule eq_matI) by (auto simp add: adjoint_eval) lemma adjoint_scale: fixes A :: "'a::conjugatable_field mat" shows "adjoint (a \\<^sub>m A) = (conjugate a) \\<^sub>m adjoint A" apply (rule eq_matI) using conjugatable_ring_class.conjugate_dist_mul by (auto simp add: adjoint_eval) lemma adjoint_add: fixes A B :: "'a::conjugatable_field mat" assumes "A \ carrier_mat n m" "B \ carrier_mat n m" shows "adjoint (A + B) = adjoint A + adjoint B" apply (rule eq_matI) using assms conjugatable_ring_class.conjugate_dist_add by( auto simp add: adjoint_eval) lemma adjoint_minus: fixes A B :: "'a::conjugatable_field mat" assumes "A \ carrier_mat n m" "B \ carrier_mat n m" shows "adjoint (A - B) = adjoint A - adjoint B" apply (rule eq_matI) using assms apply(auto simp add: adjoint_eval) by (metis add_uminus_conv_diff conjugate_dist_add conjugate_neg) lemma adjoint_mult: fixes A B :: "'a::conjugatable_field mat" assumes "A \ carrier_mat n m" "B \ carrier_mat m l" shows "adjoint (A * B) = adjoint B * adjoint A" proof (rule eq_matI, auto simp add: adjoint_eval adjoint_row adjoint_col) fix i j assume "i < dim_col B" "j < dim_row A" show "conjugate (row A j \ col B i) = conjugate (col B i) \ conjugate (row A j)" using assms apply (simp add: conjugate_scalar_prod) apply (subst comm_scalar_prod[where n="dim_row B"]) by (auto simp add: carrier_vecI) qed lemma adjoint_adjoint: fixes A :: "'a::conjugatable_field mat" shows "adjoint (adjoint A) = A" by (rule eq_matI, auto simp add: adjoint_eval) lemma trace_adjoint_positive: fixes A :: "complex mat" shows "trace (A * adjoint A) \ 0" apply (auto simp add: trace_def adjoint_col) apply (rule sum_nonneg) by auto subsection \Algebraic manipulations on matrices\ lemma right_add_zero_mat[simp]: "(A :: 'a :: monoid_add mat) \ carrier_mat nr nc \ A + 0\<^sub>m nr nc = A" by (intro eq_matI, auto) lemma add_carrier_mat': "A \ carrier_mat nr nc \ B \ carrier_mat nr nc \ A + B \ carrier_mat nr nc" by simp lemma minus_carrier_mat': "A \ carrier_mat nr nc \ B \ carrier_mat nr nc \ A - B \ carrier_mat nr nc" by auto lemma swap_plus_mat: fixes A B C :: "'a::semiring_1 mat" assumes "A \ carrier_mat n n" "B \ carrier_mat n n" "C \ carrier_mat n n" shows "A + B + C = A + C + B" by (metis assms assoc_add_mat comm_add_mat) lemma uminus_mat: fixes A :: "complex mat" assumes "A \ carrier_mat n n" shows "-A = (-1) \\<^sub>m A" by auto ML_file "mat_alg.ML" method_setup mat_assoc = \mat_assoc_method\ "Normalization of expressions on matrices" lemma mat_assoc_test: fixes A B C D :: "complex mat" assumes "A \ carrier_mat n n" "B \ carrier_mat n n" "C \ carrier_mat n n" "D \ carrier_mat n n" shows "(A * B) * (C * D) = A * B * C * D" "adjoint (A * adjoint B) * C = B * (adjoint A * C)" "A * 1\<^sub>m n * 1\<^sub>m n * B * 1\<^sub>m n = A * B" "(A - B) + (B - C) = A + (-B) + B + (-C)" "A + (B - C) = A + B - C" "A - (B + C + D) = A - B - C - D" "(A + B) * (B + C) = A * B + B * B + A * C + B * C" "A - B = A + (-1) \\<^sub>m B" "A * (B - C) * D = A * B * D - A * C * D" "trace (A * B * C) = trace (B * C * A)" "trace (A * B * C * D) = trace (C * D * A * B)" "trace (A + B * C) = trace A + trace (C * B)" "A + B = B + A" "A + B + C = C + B + A" "A + B + (C + D) = A + C + (B + D)" using assms by (mat_assoc n)+ subsection \Hermitian matrices\ text \A Hermitian matrix is a matrix that is equal to its Hermitian adjoint.\ definition hermitian :: "'a::conjugatable_field mat \ bool" where "hermitian A \ (adjoint A = A)" lemma hermitian_one: shows "hermitian ((1\<^sub>m n)::('a::conjugatable_field mat))" unfolding hermitian_def proof- have "conjugate (1::'a) = 1" apply (subst mult_1_right[symmetric, of "conjugate 1"]) apply (subst conjugate_id[symmetric, of "conjugate 1 * 1"]) apply (subst conjugate_dist_mul) apply auto done then show "adjoint ((1\<^sub>m n)::('a::conjugatable_field mat)) = (1\<^sub>m n)" by (auto simp add: adjoint_eval) qed subsection \Inverse matrices\ lemma inverts_mat_symm: fixes A B :: "'a::field mat" assumes dim: "A \ carrier_mat n n" "B \ carrier_mat n n" and AB: "inverts_mat A B" shows "inverts_mat B A" proof - have "A * B = 1\<^sub>m n" using dim AB unfolding inverts_mat_def by auto with dim have "B * A = 1\<^sub>m n" by (rule mat_mult_left_right_inverse) then show "inverts_mat B A" using dim inverts_mat_def by auto qed lemma inverts_mat_unique: fixes A B C :: "'a::field mat" assumes dim: "A \ carrier_mat n n" "B \ carrier_mat n n" "C \ carrier_mat n n" and AB: "inverts_mat A B" and AC: "inverts_mat A C" shows "B = C" proof - have AB1: "A * B = 1\<^sub>m n" using AB dim unfolding inverts_mat_def by auto have "A * C = 1\<^sub>m n" using AC dim unfolding inverts_mat_def by auto then have CA1: "C * A = 1\<^sub>m n" using mat_mult_left_right_inverse[of A n C] dim by auto then have "C = C * 1\<^sub>m n" using dim by auto also have "\ = C * (A * B)" using AB1 by auto also have "\ = (C * A) * B" using dim by auto also have "\ = 1\<^sub>m n * B" using CA1 by auto also have "\ = B" using dim by auto finally show "B = C" .. qed subsection \Unitary matrices\ text \A unitary matrix is a matrix whose Hermitian adjoint is also its inverse.\ definition unitary :: "'a::conjugatable_field mat \ bool" where "unitary A \ A \ carrier_mat (dim_row A) (dim_row A) \ inverts_mat A (adjoint A)" lemma unitaryD2: assumes "A \ carrier_mat n n" shows "unitary A \ inverts_mat (adjoint A) A" using assms adjoint_dim inverts_mat_symm unitary_def by blast lemma unitary_simps [simp]: "A \ carrier_mat n n \ unitary A \ adjoint A * A = 1\<^sub>m n" "A \ carrier_mat n n \ unitary A \ A * adjoint A = 1\<^sub>m n" apply (metis adjoint_dim_row carrier_matD(2) inverts_mat_def unitaryD2) by (simp add: inverts_mat_def unitary_def) lemma unitary_adjoint [simp]: assumes "A \ carrier_mat n n" "unitary A" shows "unitary (adjoint A)" unfolding unitary_def using adjoint_dim[OF assms(1)] assms by (auto simp add: unitaryD2[OF assms] adjoint_adjoint) lemma unitary_one: shows "unitary ((1\<^sub>m n)::('a::conjugatable_field mat))" unfolding unitary_def proof - define I where I_def[simp]: "I \ ((1\<^sub>m n)::('a::conjugatable_field mat))" have dim: "I \ carrier_mat n n" by auto have "hermitian I" using hermitian_one by auto hence "adjoint I = I" using hermitian_def by auto with dim show "I \ carrier_mat (dim_row I) (dim_row I) \ inverts_mat I (adjoint I)" unfolding inverts_mat_def using dim by auto qed lemma unitary_zero: fixes A :: "'a::conjugatable_field mat" assumes "A \ carrier_mat 0 0" shows "unitary A" unfolding unitary_def inverts_mat_def Let_def using assms by auto lemma unitary_elim: assumes dims: "A \ carrier_mat n n" "B \ carrier_mat n n" "P \ carrier_mat n n" and uP: "unitary P" and eq: "P * A * adjoint P = P * B * adjoint P" shows "A = B" proof - have dimaP: "adjoint P \ carrier_mat n n" using dims by auto have iv: "inverts_mat P (adjoint P)" using uP unitary_def by auto then have "P * (adjoint P) = 1\<^sub>m n" using inverts_mat_def dims by auto then have aPP: "adjoint P * P = 1\<^sub>m n" using mat_mult_left_right_inverse[OF dims(3) dimaP] by auto have "adjoint P * (P * A * adjoint P) * P = (adjoint P * P) * A * (adjoint P * P)" using dims dimaP by (mat_assoc n) also have "\ = 1\<^sub>m n * A * 1\<^sub>m n" using aPP by auto also have "\ = A" using dims by auto finally have eqA: "A = adjoint P * (P * A * adjoint P) * P" .. have "adjoint P * (P * B * adjoint P) * P = (adjoint P * P) * B * (adjoint P * P)" using dims dimaP by (mat_assoc n) also have "\ = 1\<^sub>m n * B * 1\<^sub>m n" using aPP by auto also have "\ = B" using dims by auto finally have eqB: "B = adjoint P * (P * B * adjoint P) * P" .. then show ?thesis using eqA eqB eq by auto qed lemma unitary_is_corthogonal: fixes U :: "'a::conjugatable_field mat" assumes dim: "U \ carrier_mat n n" and U: "unitary U" shows "corthogonal_mat U" unfolding corthogonal_mat_def Let_def proof (rule conjI) have dima: "adjoint U \ carrier_mat n n" using dim by auto have aUU: "mat_adjoint U * U = (1\<^sub>m n)" apply (insert U[unfolded unitary_def] dim dima, drule conjunct2) apply (drule inverts_mat_symm[of "U", OF dim dima], unfold inverts_mat_def, auto) done then show "diagonal_mat (mat_adjoint U * U)" by (simp add: diagonal_mat_def) show "\i 0" using dim by (simp add: aUU) qed lemma unitary_times_unitary: fixes P Q :: "'a:: conjugatable_field mat" assumes dim: "P \ carrier_mat n n" "Q \ carrier_mat n n" and uP: "unitary P" and uQ: "unitary Q" shows "unitary (P * Q)" proof - have dim_pq: "P * Q \ carrier_mat n n" using dim by auto have "(P * Q) * adjoint (P * Q) = P * (Q * adjoint Q) * adjoint P" using dim by (mat_assoc n) also have "\ = P * (1\<^sub>m n) * adjoint P" using uQ dim by auto also have "\ = P * adjoint P" using dim by (mat_assoc n) also have "\ = 1\<^sub>m n" using uP dim by simp finally have "(P * Q) * adjoint (P * Q) = 1\<^sub>m n" by auto hence "inverts_mat (P * Q) (adjoint (P * Q))" using inverts_mat_def dim_pq by auto thus "unitary (P*Q)" using unitary_def dim_pq by auto qed lemma unitary_operator_keep_trace: fixes U A :: "complex mat" assumes dU: "U \ carrier_mat n n" and dA: "A \ carrier_mat n n" and u: "unitary U" shows "trace A = trace (adjoint U * A * U)" proof - have u': "U * adjoint U = 1\<^sub>m n" using u unfolding unitary_def inverts_mat_def using dU by auto have "trace (adjoint U * A * U) = trace (U * adjoint U * A)" using dU dA by (mat_assoc n) also have "\ = trace A" using u' dA by auto finally show ?thesis by auto qed subsection \Normalization of vectors\ definition vec_norm :: "complex vec \ complex" where "vec_norm v \ csqrt (v \c v)" lemma vec_norm_geq_0: fixes v :: "complex vec" shows "vec_norm v \ 0" unfolding vec_norm_def by (insert self_cscalar_prod_geq_0[of v], simp add: less_eq_complex_def) lemma vec_norm_zero: fixes v :: "complex vec" assumes dim: "v \ carrier_vec n" shows "vec_norm v = 0 \ v = 0\<^sub>v n" unfolding vec_norm_def by (subst conjugate_square_eq_0_vec[OF dim, symmetric], rule csqrt_eq_0) lemma vec_norm_ge_0: fixes v :: "complex vec" assumes dim_v: "v \ carrier_vec n" and neq0: "v \ 0\<^sub>v n" shows "vec_norm v > 0" proof - have geq: "vec_norm v \ 0" using vec_norm_geq_0 by auto have neq: "vec_norm v \ 0" apply (insert dim_v neq0) apply (drule vec_norm_zero, auto) done show ?thesis using neq geq by (rule dual_order.not_eq_order_implies_strict) qed definition vec_normalize :: "complex vec \ complex vec" where "vec_normalize v = (if (v = 0\<^sub>v (dim_vec v)) then v else 1 / (vec_norm v) \\<^sub>v v)" lemma normalized_vec_dim[simp]: assumes "(v::complex vec) \ carrier_vec n" shows "vec_normalize v \ carrier_vec n" unfolding vec_normalize_def using assms by auto lemma vec_eq_norm_smult_normalized: shows "v = vec_norm v \\<^sub>v vec_normalize v" proof (cases "v = 0\<^sub>v (dim_vec v)") define n where "n = dim_vec v" then have dimv: "v \ carrier_vec n" by auto then have dimnv: "vec_normalize v \ carrier_vec n" by auto { case True then have v0: "v = 0\<^sub>v n" using n_def by auto then have n0: "vec_norm v = 0" using vec_norm_def by auto have "vec_norm v \\<^sub>v vec_normalize v = 0\<^sub>v n" unfolding smult_vec_def by (auto simp add: n0 carrier_vecD[OF dimnv]) then show ?thesis using v0 by auto next case False then have v: "v \ 0\<^sub>v n" using n_def by auto then have ge0: "vec_norm v > 0" using vec_norm_ge_0 dimv by auto have "vec_normalize v = (1 / vec_norm v) \\<^sub>v v" using False vec_normalize_def by auto then have "vec_norm v \\<^sub>v vec_normalize v = (vec_norm v * (1 / vec_norm v)) \\<^sub>v v" using smult_smult_assoc by auto also have "\ = v" using ge0 by auto finally have "v = vec_norm v \\<^sub>v vec_normalize v".. then show "v = vec_norm v \\<^sub>v vec_normalize v" using v by auto } qed lemma normalized_cscalar_prod: fixes v w :: "complex vec" assumes dim_v: "v \ carrier_vec n" and dim_w: "w \ carrier_vec n" shows "v \c w = (vec_norm v * vec_norm w) * (vec_normalize v \c vec_normalize w)" unfolding vec_normalize_def apply (split if_split, split if_split) proof (intro conjI impI) note dim0 = dim_v dim_w have dim: "dim_vec v = n" "dim_vec w = n" using dim0 by auto { assume "w = 0\<^sub>v n" "v = 0\<^sub>v n" then have lhs: "v \c w = 0" by auto then moreover have rhs: "vec_norm v * vec_norm w * (v \c w) = 0" by auto ultimately have "v \c w = vec_norm v * vec_norm w * (v \c w)" by auto } with dim show "w = 0\<^sub>v (dim_vec w) \ v = 0\<^sub>v (dim_vec v) \ v \c w = vec_norm v * vec_norm w * (v \c w)" by auto { assume asm: "w = 0\<^sub>v n" "v \ 0\<^sub>v n" then have w0: "conjugate w = 0\<^sub>v n" by auto with dim0 have "(1 / vec_norm v \\<^sub>v v) \c w = 0" by auto then moreover have rhs: "vec_norm v * vec_norm w * ((1 / vec_norm v \\<^sub>v v) \c w) = 0" by auto moreover have "v \c w = 0" using w0 dim0 by auto ultimately have "v \c w = vec_norm v * vec_norm w * ((1 / vec_norm v \\<^sub>v v) \c w)" by auto } with dim show "w = 0\<^sub>v (dim_vec w) \ v \ 0\<^sub>v (dim_vec v) \ v \c w = vec_norm v * vec_norm w * ((1 / vec_norm v \\<^sub>v v) \c w)" by auto { assume asm: "w \ 0\<^sub>v n" "v = 0\<^sub>v n" with dim0 have "v \c (1 / vec_norm w \\<^sub>v w) = 0" by auto then moreover have rhs: "vec_norm v * vec_norm w * (v \c (1 / vec_norm w \\<^sub>v w)) = 0" by auto moreover have "v \c w = 0" using asm dim0 by auto ultimately have "v \c w = vec_norm v * vec_norm w * (v \c (1 / vec_norm w \\<^sub>v w))" by auto } with dim show "w \ 0\<^sub>v (dim_vec w) \ v = 0\<^sub>v (dim_vec v) \ v \c w = vec_norm v * vec_norm w * (v \c (1 / vec_norm w \\<^sub>v w))" by auto { assume asmw: "w \ 0\<^sub>v n" and asmv: "v \ 0\<^sub>v n" have "vec_norm w > 0" by (insert asmw dim0, rule vec_norm_ge_0, auto) then have cw: "conjugate (1 / vec_norm w) = 1 / vec_norm w" by (simp add: complex_eq_iff complex_is_Real_iff less_complex_def) from dim0 have "((1 / vec_norm v \\<^sub>v v) \c (1 / vec_norm w \\<^sub>v w)) = 1 / vec_norm v * (v \c (1 / vec_norm w \\<^sub>v w))" by auto also have "\ = 1 / vec_norm v * (v \ (conjugate (1 / vec_norm w) \\<^sub>v conjugate w))" by (subst conjugate_smult_vec, auto) also have "\ = 1 / vec_norm v * conjugate (1 / vec_norm w) * (v \ conjugate w)" using dim by auto also have "\ = 1 / vec_norm v * (1 / vec_norm w) * (v \c w)" using vec_norm_ge_0 cw by auto finally have eq1: "(1 / vec_norm v \\<^sub>v v) \c (1 / vec_norm w \\<^sub>v w) = 1 / vec_norm v * (1 / vec_norm w) * (v \c w)" . then have "vec_norm v * vec_norm w * ((1 / vec_norm v \\<^sub>v v) \c (1 / vec_norm w \\<^sub>v w)) = (v \c w)" by (subst eq1, insert vec_norm_ge_0[of v n, OF dim_v asmv] vec_norm_ge_0[of w n, OF dim_w asmw], auto) } with dim show " w \ 0\<^sub>v (dim_vec w) \ v \ 0\<^sub>v (dim_vec v) \ v \c w = vec_norm v * vec_norm w * ((1 / vec_norm v \\<^sub>v v) \c (1 / vec_norm w \\<^sub>v w))" by auto qed lemma normalized_vec_norm : fixes v :: "complex vec" assumes dim_v: "v \ carrier_vec n" and neq0: "v \ 0\<^sub>v n" shows "vec_normalize v \c vec_normalize v = 1" unfolding vec_normalize_def proof (simp, rule conjI) show "v = 0\<^sub>v (dim_vec v) \ v \c v = 1" using neq0 dim_v by auto have dim_a: "(vec_normalize v) \ carrier_vec n" "conjugate (vec_normalize v) \ carrier_vec n" using dim_v vec_normalize_def by auto note dim = dim_v dim_a have nvge0: "vec_norm v > 0" using vec_norm_ge_0 neq0 dim_v by auto then have vvvv: "v \c v = (vec_norm v) * (vec_norm v)" unfolding vec_norm_def by (metis power2_csqrt power2_eq_square) from nvge0 have "conjugate (vec_norm v) = vec_norm v" by (simp add: complex_eq_iff complex_is_Real_iff less_complex_def) then have "v \c (1 / vec_norm v \\<^sub>v v) = 1 / vec_norm v * (v \c v)" by (subst conjugate_smult_vec, auto) also have "\ = 1 / vec_norm v * vec_norm v * vec_norm v" using vvvv by auto also have "\ = vec_norm v" by auto finally have "v \c (1 / vec_norm v \\<^sub>v v) = vec_norm v". then show "v \ 0\<^sub>v (dim_vec v) \ vec_norm v \ 0 \ v \c (1 / vec_norm v \\<^sub>v v) = vec_norm v" using neq0 nvge0 by auto qed lemma normalize_zero: assumes "v \ carrier_vec n" shows "vec_normalize v = 0\<^sub>v n \ v = 0\<^sub>v n" proof show "v = 0\<^sub>v n \ vec_normalize v = 0\<^sub>v n" unfolding vec_normalize_def by auto next have "v \ 0\<^sub>v n \ vec_normalize v \ 0\<^sub>v n" unfolding vec_normalize_def proof (simp, rule impI) assume asm: "v \ 0\<^sub>v n" then have "vec_norm v > 0" using vec_norm_ge_0 assms by auto then have nvge0: "1 / vec_norm v > 0" by (simp add: complex_is_Real_iff less_complex_def) have "\k < n. v $ k \ 0" using asm assms by auto then obtain k where kn: "k < n" and vkneq0: "v $ k \ 0" by auto then have "(1 / vec_norm v \\<^sub>v v) $ k = (1 / vec_norm v) * (v $ k)" using assms carrier_vecD index_smult_vec(1) by blast with nvge0 vkneq0 have "(1 / vec_norm v \\<^sub>v v) $ k \ 0" by auto then show "1 / vec_norm v \\<^sub>v v \ 0\<^sub>v n" using assms kn by fastforce qed then show "vec_normalize v = 0\<^sub>v n \ v = 0\<^sub>v n" by auto qed lemma normalize_normalize[simp]: "vec_normalize (vec_normalize v) = vec_normalize v" proof (rule disjE[of "v = 0\<^sub>v (dim_vec v)" "v \ 0\<^sub>v (dim_vec v)"], auto) let ?n = "dim_vec v" { assume "v = 0\<^sub>v ?n" then have "vec_normalize v = v" unfolding vec_normalize_def by auto then show "vec_normalize (vec_normalize v) = vec_normalize v" by auto } assume neq0: "v \ 0\<^sub>v ?n" have dim: "v \ carrier_vec ?n" by auto have "vec_norm (vec_normalize v) = 1" unfolding vec_norm_def using normalized_vec_norm[OF dim neq0] by auto then show "vec_normalize (vec_normalize v) = vec_normalize v" by (subst (1) vec_normalize_def, simp) qed subsection \Spectral decomposition of normal complex matrices\ lemma normalize_keep_corthogonal: fixes vs :: "complex vec list" assumes cor: "corthogonal vs" and dims: "set vs \ carrier_vec n" shows "corthogonal (map vec_normalize vs)" unfolding corthogonal_def proof (rule allI, rule impI, rule allI, rule impI, goal_cases) case c: (1 i j) let ?m = "length vs" have len: "length (map vec_normalize vs) = ?m" by auto have dim: "\k. k < ?m \ (vs ! k) \ carrier_vec n" using dims by auto have map: "\k. k < ?m \ map vec_normalize vs ! k = vec_normalize (vs ! k)" by auto have eq1: "\j k. j < ?m \ k < ?m \ ((vs ! j) \c (vs ! k) = 0) = (j \ k)" using assms unfolding corthogonal_def by auto then have "\k. k < ?m \ (vs ! k) \c (vs ! k) \ 0 " by auto then have "\k. k < ?m \ (vs ! k) \ (0\<^sub>v n)" using dim by (auto simp add: conjugate_square_eq_0_vec[of _ n, OF dim]) then have vnneq0: "\k. k < ?m \ vec_norm (vs ! k) \ 0" using vec_norm_zero[OF dim] by auto then have i0: "vec_norm (vs ! i) \ 0" and j0: "vec_norm (vs ! j) \ 0" using c by auto have "(vs ! i) \c (vs ! j) = vec_norm (vs ! i) * vec_norm (vs ! j) * (vec_normalize (vs ! i) \c vec_normalize (vs ! j))" by (subst normalized_cscalar_prod[of "vs ! i" n "vs ! j"], auto, insert dim c, auto) with i0 j0 have "(vec_normalize (vs ! i) \c vec_normalize (vs ! j) = 0) = ((vs ! i) \c (vs ! j) = 0)" by auto with eq1 c have "(vec_normalize (vs ! i) \c vec_normalize (vs ! j) = 0) = (i \ j)" by auto with map c show "(map vec_normalize vs ! i \c map vec_normalize vs ! j = 0) = (i \ j)" by auto qed lemma normalized_corthogonal_mat_is_unitary: assumes W: "set ws \ carrier_vec n" and orth: "corthogonal ws" and len: "length ws = n" shows "unitary (mat_of_cols n (map vec_normalize ws))" (is "unitary ?W") proof - define vs where "vs = map vec_normalize ws" define W where "W = mat_of_cols n vs" have W': "set vs \ carrier_vec n" using assms vs_def by auto then have W'': "\k. k < length vs \ vs ! k \ carrier_vec n" by auto have orth': "corthogonal vs" using assms normalize_keep_corthogonal vs_def by auto have len'[simp]: "length vs = n" using assms vs_def by auto have dimW: "W \ carrier_mat n n" using W_def len by auto have "adjoint W \ carrier_mat n n" using dimW by auto then have dimaW: "mat_adjoint W \ carrier_mat n n" by auto { fix i j assume i: "i < n" and j: "j < n" have dimws: "(ws ! i) \ carrier_vec n" "(ws ! j) \ carrier_vec n" using W len i j by auto have "(ws ! i) \c (ws ! i) \ 0" "(ws ! j) \c (ws ! j) \ 0" using orth corthogonal_def[of ws] len i j by auto then have neq0: "(ws ! i) \ 0\<^sub>v n" "(ws ! j) \ 0\<^sub>v n" by (auto simp add: conjugate_square_eq_0_vec[of "ws ! i" n]) then have "vec_norm (ws ! i) > 0" "vec_norm (ws ! j) > 0" using vec_norm_ge_0 dimws by auto then have ge0: "vec_norm (ws ! i) * vec_norm (ws ! j) > 0" by (auto simp: less_complex_def) have ws': "vs ! i = vec_normalize (ws ! i)" "vs ! j = vec_normalize (ws ! j)" using len i j vs_def by auto have ii1: "(vs ! i) \c (vs ! i) = 1" apply (simp add: ws') apply (rule normalized_vec_norm[of "ws ! i"], rule dimws, rule neq0) done have ij0: "i \ j \ (ws ! i) \c (ws ! j) = 0" using i j by (insert orth, auto simp add: corthogonal_def[of ws] len) have "i \ j \ (ws ! i) \c (ws ! j) = (vec_norm (ws ! i) * vec_norm (ws ! j)) * ((vs ! i) \c (vs ! j))" apply (auto simp add: ws') apply (rule normalized_cscalar_prod) apply (rule dimws, rule dimws) done with ij0 have ij0': "i \ j \ (vs ! i) \c (vs ! j) = 0" using ge0 by auto have cWk: "\k. k < n \ col W k = vs ! k" unfolding W_def apply (subst col_mat_of_cols) apply (auto simp add: W'') done have "(mat_adjoint W * W) $$ (j, i) = row (mat_adjoint W) j \ col W i" by (insert dimW i j dimaW, auto) also have "\ = conjugate (col W j) \ col W i" by (insert dimW i j dimaW, auto simp add: mat_adjoint_def) also have "\ = col W i \ conjugate (col W j)" using comm_scalar_prod[of "col W i" n] dimW by auto also have "\ = (vs ! i) \c (vs ! j)" using W_def col_mat_of_cols i j len cWk by auto finally have "(mat_adjoint W * W) $$ (j, i) = (vs ! i) \c (vs ! j)". then have "(mat_adjoint W * W) $$ (j, i) = (if (j = i) then 1 else 0)" by (auto simp add: ii1 ij0') } note maWW = this then have "mat_adjoint W * W = 1\<^sub>m n" unfolding one_mat_def using dimW dimaW by (auto simp add: maWW adjoint_def) then have iv0: "adjoint W * W = 1\<^sub>m n" by auto have dimaW: "adjoint W \ carrier_mat n n" using dimaW by auto then have iv1: "W * adjoint W = 1\<^sub>m n" using mat_mult_left_right_inverse dimW iv0 by auto then show "unitary W" unfolding unitary_def inverts_mat_def using dimW dimaW iv0 iv1 by auto qed lemma normalize_keep_eigenvector: assumes ev: "eigenvector A v e" and dim: "A \ carrier_mat n n" "v \ carrier_vec n" shows "eigenvector A (vec_normalize v) e" unfolding eigenvector_def proof show "vec_normalize v \ carrier_vec (dim_row A)" using dim by auto have eg: "A *\<^sub>v v = e \\<^sub>v v" using ev dim eigenvector_def by auto have vneq0: "v \ 0\<^sub>v n" using ev dim unfolding eigenvector_def by auto then have s0: "vec_normalize v \ 0\<^sub>v n" by (insert dim, subst normalize_zero[of v], auto) from vneq0 have vvge0: "vec_norm v > 0" using vec_norm_ge_0 dim by auto have s1: "A *\<^sub>v vec_normalize v = e \\<^sub>v vec_normalize v" unfolding vec_normalize_def using vneq0 dim apply (auto, simp add: mult_mat_vec) apply (subst eg, auto) done with s0 dim show "vec_normalize v \ 0\<^sub>v (dim_row A) \ A *\<^sub>v vec_normalize v = e \\<^sub>v vec_normalize v" by auto qed lemma four_block_mat_adjoint: fixes A B C D :: "'a::conjugatable_field mat" assumes dim: "A \ carrier_mat nr1 nc1" "B \ carrier_mat nr1 nc2" "C \ carrier_mat nr2 nc1" "D \ carrier_mat nr2 nc2" shows "adjoint (four_block_mat A B C D) = four_block_mat (adjoint A) (adjoint C) (adjoint B) (adjoint D)" by (rule eq_matI, insert dim, auto simp add: adjoint_eval) fun unitary_schur_decomposition :: "complex mat \ complex list \ complex mat \ complex mat \ complex mat" where "unitary_schur_decomposition A [] = (A, 1\<^sub>m (dim_row A), 1\<^sub>m (dim_row A))" | "unitary_schur_decomposition A (e # es) = (let n = dim_row A; n1 = n - 1; v' = find_eigenvector A e; v = vec_normalize v'; ws0 = gram_schmidt n (basis_completion v); ws = map vec_normalize ws0; W = mat_of_cols n ws; W' = corthogonal_inv W; A' = W' * A * W; (A1,A2,A0,A3) = split_block A' 1 1; (B,P,Q) = unitary_schur_decomposition A3 es; z_row = (0\<^sub>m 1 n1); z_col = (0\<^sub>m n1 1); one_1 = 1\<^sub>m 1 in (four_block_mat A1 (A2 * P) A0 B, W * four_block_mat one_1 z_row z_col P, four_block_mat one_1 z_row z_col Q * W'))" theorem unitary_schur_decomposition: assumes A: "(A::complex mat) \ carrier_mat n n" and c: "char_poly A = (\ (e :: complex) \ es. [:- e, 1:])" and B: "unitary_schur_decomposition A es = (B,P,Q)" shows "similar_mat_wit A B P Q \ upper_triangular B \ diag_mat B = es \ unitary P \ (Q = adjoint P)" using assms proof (induct es arbitrary: n A B P Q) case Nil with degree_monic_char_poly[of A n] show ?case by (auto intro: similar_mat_wit_refl simp: diag_mat_def unitary_zero) next case (Cons e es n A C P Q) let ?n1 = "n - 1" from Cons have A: "A \ carrier_mat n n" and dim: "dim_row A = n" by auto let ?cp = "char_poly A" from Cons(3) have cp: "?cp = [: -e, 1 :] * (\e \ es. [:- e, 1:])" by auto have mon: "monic (\e\ es. [:- e, 1:])" by (rule monic_prod_list, auto) have deg: "degree ?cp = Suc (degree (\e\ es. [:- e, 1:]))" unfolding cp by (subst degree_mult_eq, insert mon, auto) with degree_monic_char_poly[OF A] have n: "n \ 0" by auto define v' where "v' = find_eigenvector A e" define v where "v = vec_normalize v'" define b where "b = basis_completion v" define ws0 where "ws0 = gram_schmidt n b" define ws where "ws = map vec_normalize ws0" define W where "W = mat_of_cols n ws" define W' where "W' = corthogonal_inv W" define A' where "A' = W' * A * W" obtain A1 A2 A0 A3 where splitA': "split_block A' 1 1 = (A1,A2,A0,A3)" by (cases "split_block A' 1 1", auto) obtain B P' Q' where schur: "unitary_schur_decomposition A3 es = (B,P',Q')" by (cases "unitary_schur_decomposition A3 es", auto) let ?P' = "four_block_mat (1\<^sub>m 1) (0\<^sub>m 1 ?n1) (0\<^sub>m ?n1 1) P'" let ?Q' = "four_block_mat (1\<^sub>m 1) (0\<^sub>m 1 ?n1) (0\<^sub>m ?n1 1) Q'" have C: "C = four_block_mat A1 (A2 * P') A0 B" and P: "P = W * ?P'" and Q: "Q = ?Q' * W'" using Cons(4) unfolding unitary_schur_decomposition.simps Let_def list.sel dim v'_def[symmetric] v_def[symmetric] b_def[symmetric] ws0_def[symmetric] ws_def[symmetric] W'_def[symmetric] W_def[symmetric] A'_def[symmetric] split splitA' schur by auto have e: "eigenvalue A e" unfolding eigenvalue_root_char_poly[OF A] cp by simp from find_eigenvector[OF A e] have ev': "eigenvector A v' e" unfolding v'_def . then have "v' \ carrier_vec n" unfolding eigenvector_def using A by auto with ev' have ev: "eigenvector A v e" unfolding v_def using A dim normalize_keep_eigenvector by auto from this[unfolded eigenvector_def] have v[simp]: "v \ carrier_vec n" and v0: "v \ 0\<^sub>v n" using A by auto interpret cof_vec_space n "TYPE(complex)" . from basis_completion[OF v v0, folded b_def] have span_b: "span (set b) = carrier_vec n" and dist_b: "distinct b" and indep: "\ lin_dep (set b)" and b: "set b \ carrier_vec n" and hdb: "hd b = v" and len_b: "length b = n" by auto from hdb len_b n obtain vs where bv: "b = v # vs" by (cases b, auto) from gram_schmidt_result[OF b dist_b indep refl, folded ws0_def] have ws0: "set ws0 \ carrier_vec n" "corthogonal ws0" "length ws0 = n" by (auto simp: len_b) then have ws: "set ws \ carrier_vec n" "corthogonal ws" "length ws = n" unfolding ws_def using normalize_keep_corthogonal by auto have ws0ne: "ws0 \ []" using \length ws0 = n\ n by auto from gram_schmidt_hd[OF v, of vs, folded bv] have hdws0: "hd ws0 = (vec_normalize v')" unfolding ws0_def v_def . have "hd ws = vec_normalize (hd ws0)" unfolding ws_def using hd_map[OF ws0ne] by auto then have hdws: "hd ws = v" unfolding v_def using normalize_normalize[of v'] hdws0 by auto have orth_W: "corthogonal_mat W" using orthogonal_mat_of_cols ws unfolding W_def. have W: "W \ carrier_mat n n" using ws unfolding W_def using mat_of_cols_carrier(1)[of n ws] by auto have W': "W' \ carrier_mat n n" unfolding W'_def corthogonal_inv_def using W by (auto simp: mat_of_rows_def) from corthogonal_inv_result[OF orth_W] have W'W: "inverts_mat W' W" unfolding W'_def . hence WW': "inverts_mat W W'" using mat_mult_left_right_inverse[OF W' W] W' W unfolding inverts_mat_def by auto have A': "A' \ carrier_mat n n" using W W' A unfolding A'_def by auto have A'A_wit: "similar_mat_wit A' A W' W" by (rule similar_mat_witI[of _ _ n], insert W W' A A' W'W WW', auto simp: A'_def inverts_mat_def) hence A'A: "similar_mat A' A" unfolding similar_mat_def by blast from similar_mat_wit_sym[OF A'A_wit] have simAA': "similar_mat_wit A A' W W'" by auto have eigen[simp]: "A *\<^sub>v v = e \\<^sub>v v" and v0: "v \ 0\<^sub>v n" using v_def v'_def find_eigenvector[OF A e] A normalize_keep_eigenvector unfolding eigenvector_def by auto let ?f = "(\ i. if i = 0 then e else 0)" have col0: "col A' 0 = vec n ?f" unfolding A'_def W'_def W_def using corthogonal_col_ev_0[OF A v v0 eigen n hdws ws]. from A' n have "dim_row A' = 1 + ?n1" "dim_col A' = 1 + ?n1" by auto from split_block[OF splitA' this] have A2: "A2 \ carrier_mat 1 ?n1" and A3: "A3 \ carrier_mat ?n1 ?n1" and A'block: "A' = four_block_mat A1 A2 A0 A3" by auto have A1id: "A1 = mat 1 1 (\ _. e)" using splitA'[unfolded split_block_def Let_def] arg_cong[OF col0, of "\ v. v $ 0"] A' n by (auto simp: col_def) have A1: "A1 \ carrier_mat 1 1" unfolding A1id by auto { fix i assume "i < ?n1" with arg_cong[OF col0, of "\ v. v $ Suc i"] A' have "A' $$ (Suc i, 0) = 0" by auto } note A'0 = this have A0id: "A0 = 0\<^sub>m ?n1 1" using splitA'[unfolded split_block_def Let_def] A'0 A' by auto have A0: "A0 \ carrier_mat ?n1 1" unfolding A0id by auto from cp char_poly_similar[OF A'A] have cp: "char_poly A' = [: -e,1 :] * (\ e \ es. [:- e, 1:])" by simp also have "char_poly A' = char_poly A1 * char_poly A3" unfolding A'block A0id by (rule char_poly_four_block_zeros_col[OF A1 A2 A3]) also have "char_poly A1 = [: -e,1 :]" by (simp add: A1id char_poly_defs det_def) finally have cp: "char_poly A3 = (\ e \ es. [:- e, 1:])" by (metis mult_cancel_left pCons_eq_0_iff zero_neq_one) from Cons(1)[OF A3 cp schur] have simIH: "similar_mat_wit A3 B P' Q'" and ut: "upper_triangular B" and diag: "diag_mat B = es" and uP': "unitary P'" and Q'P': "Q' = adjoint P'" by auto from similar_mat_witD2[OF A3 simIH] have B: "B \ carrier_mat ?n1 ?n1" and P': "P' \ carrier_mat ?n1 ?n1" and Q': "Q' \ carrier_mat ?n1 ?n1" and PQ': "P' * Q' = 1\<^sub>m ?n1" by auto have A0_eq: "A0 = P' * A0 * 1\<^sub>m 1" unfolding A0id using P' by auto have simA'C: "similar_mat_wit A' C ?P' ?Q'" unfolding A'block C by (rule similar_mat_wit_four_block[OF similar_mat_wit_refl[OF A1] simIH _ A0_eq A1 A3 A0], insert PQ' A2 P' Q', auto) have ut1: "upper_triangular A1" unfolding A1id by auto have ut: "upper_triangular C" unfolding C A0id by (intro upper_triangular_four_block[OF _ B ut1 ut], auto simp: A1id) from A1id have diagA1: "diag_mat A1 = [e]" unfolding diag_mat_def by auto from diag_four_block_mat[OF A1 B] have diag: "diag_mat C = e # es" unfolding diag diagA1 C by simp have aW: "adjoint W \ carrier_mat n n" using W by auto have aW': "adjoint W' \ carrier_mat n n" using W' by auto have "unitary W" using W_def ws_def ws0 normalized_corthogonal_mat_is_unitary by auto then have ivWaW: "inverts_mat W (adjoint W)" using unitary_def W aW by auto with WW' have W'aW: "W' = (adjoint W)" using inverts_mat_unique W W' aW by auto then have "adjoint W' = W" using adjoint_adjoint by auto with ivWaW have "inverts_mat W' (adjoint W')" using inverts_mat_symm W aW W'aW by auto then have "unitary W'" using unitary_def W' by auto have newP': "P' \ carrier_mat (n - Suc 0) (n - Suc 0)" using P' by auto have rl: "\ x1 x2 x3 x4 y1 y2 y3 y4 f. x1 = y1 \ x2 = y2 \ x3 = y3 \ x4 = y4 \ f x1 x2 x3 x4 = f y1 y2 y3 y4" by simp have Q'aP': "?Q' = adjoint ?P'" apply (subst four_block_mat_adjoint, auto simp add: newP') - apply (rule rl[where f = four_block_mat]) + apply (rule rl[where f2 = four_block_mat]) apply (auto simp add: eq_matI adjoint_eval Q'P') done have "adjoint P = adjoint ?P' * adjoint W" using W newP' n apply (simp add: P) apply (subst adjoint_mult[of W, symmetric]) apply (auto simp add: W P' carrier_matD[of W n n]) done also have "\ = ?Q' * W'" using Q'aP' W'aW by auto also have "\ = Q" using Q by auto finally have QaP: "Q = adjoint P" .. from similar_mat_wit_trans[OF simAA' simA'C, folded P Q] have smw: "similar_mat_wit A C P Q" by blast then have dimP: "P \ carrier_mat n n" and dimQ: "Q \ carrier_mat n n" unfolding similar_mat_wit_def using A by auto from smw have "P * Q = 1\<^sub>m n" unfolding similar_mat_wit_def using A by auto then have "inverts_mat P Q" using inverts_mat_def dimP by auto then have uP: "unitary P" using QaP unitary_def dimP by auto from ut similar_mat_wit_trans[OF simAA' simA'C, folded P Q] diag uP QaP show ?case by blast qed lemma complex_mat_char_poly_factorizable: fixes A :: "complex mat" assumes "A \ carrier_mat n n" shows "\as. char_poly A = (\ a \ as. [:- a, 1:]) \ length as = n" proof - let ?ca = "char_poly A" have ex0: "\bs. Polynomial.smult (lead_coeff ?ca) (\b\bs. [:- b, 1:]) = ?ca \ length bs = degree ?ca" by (simp add: fundamental_theorem_algebra_factorized) then obtain bs where " Polynomial.smult (lead_coeff ?ca) (\b\bs. [:- b, 1:]) = ?ca \ length bs = degree ?ca" by auto moreover have "lead_coeff ?ca = (1::complex)" using assms degree_monic_char_poly by blast ultimately have ex1: "?ca = (\b\bs. [:- b, 1:]) \ length bs = degree ?ca" by auto moreover have "degree ?ca = n" by (simp add: assms degree_monic_char_poly) ultimately show ?thesis by auto qed lemma complex_mat_has_unitary_schur_decomposition: fixes A :: "complex mat" assumes "A \ carrier_mat n n" shows "\B P es. similar_mat_wit A B P (adjoint P) \ unitary P \ char_poly A = (\ (e :: complex) \ es. [:- e, 1:]) \ diag_mat B = es" proof - have "\es. char_poly A = (\ e \ es. [:- e, 1:]) \ length es = n" using assms by (simp add: complex_mat_char_poly_factorizable) then obtain es where es: "char_poly A = (\ e \ es. [:- e, 1:]) \ length es = n" by auto obtain B P Q where B: "unitary_schur_decomposition A es = (B,P,Q)" by (cases "unitary_schur_decomposition A es", auto) have "similar_mat_wit A B P Q \ upper_triangular B \ unitary P \ (Q = adjoint P) \ char_poly A = (\ (e :: complex) \ es. [:- e, 1:]) \ diag_mat B = es" using assms es B by (auto simp add: unitary_schur_decomposition) then show ?thesis by auto qed lemma normal_upper_triangular_matrix_is_diagonal: fixes A :: "'a::conjugatable_ordered_field mat" assumes "A \ carrier_mat n n" and tri: "upper_triangular A" and norm: "A * adjoint A = adjoint A * A" shows "diagonal_mat A" proof (rule disjE[of "n = 0" "n > 0"], blast) have dim: "dim_row A = n" "dim_col A = n" using assms by auto from norm have eq0: "\i j. (A * adjoint A)$$(i,j) = (adjoint A * A)$$(i,j)" by auto have nat_induct_strong: "\P. (P::nat\bool) 0 \ (\i. i < n \ (\k. k < i \ P k) \ P i) \ (\i. i < n \ P i)" by (metis dual_order.strict_trans infinite_descent0 linorder_neqE_nat) show "n = 0 \ ?thesis" using dim unfolding diagonal_mat_def by auto show "n > 0 \ ?thesis" unfolding diagonal_mat_def dim apply (rule allI, rule impI) apply (rule nat_induct_strong) proof (rule allI, rule impI, rule impI) assume asm: "n > 0" from tri upper_triangularD[of A 0 j] dim have z0: "\j. 0< j \ j < n \ A$$(j, 0) = 0" by auto then have ada00: "(adjoint A * A)$$(0,0) = conjugate (A$$(0,0)) * A$$(0,0)" using asm dim by (auto simp add: scalar_prod_def adjoint_eval sum.atLeast_Suc_lessThan) have aad00: "(A * adjoint A)$$(0,0) = (\k=0.. = A$$(0,0) * conjugate (A$$(0,0)) + (\k=1..k. A$$(0, k) * conjugate (A$$(0, k))"], auto) ultimately have f1tneq0: "(\k=(Suc 0)..k. k < n \ A$$(0, k) * conjugate (A$$(0, k)) \ 0" using conjugate_square_positive by auto have "\k. 1 \ k \ k < n \ A$$(0, k) * conjugate (A$$(0, k)) = 0" by (rule sum_nonneg_0[of "{1..j. 0 < n \ j < n \ 0 \ j \ A $$ (0, j) = 0" by auto { fix i assume asm: "n > 0" "i < n" "i > 0" and ih: "\k. k < i \ \j j \ A $$ (k, j) = 0" then have "\j. j i \ j \ A $$ (i, j) = 0" proof - have inter_part: "\b m e. (b::nat) < e \ b < m \ m < e \ {b.. {m..b m e f. (b::nat) < e \ b < m \ m < e \ (\k=b..k\{b..{m..b m e f. (b::nat) < e \ b < m \ m < e \ (\k=b..k=b..k=m..j. j < i \ A$$(i, j) = 0" by auto from tri upper_triangularD[of A j i] asm dim have zsi1: "\k. i < k \ k < n \ A$$(k, i) = 0" by auto have "(A * adjoint A)$$(i, i) = (\k=0.. = (\k=0..k=i.. = (\k=i.. = conjugate (A$$(i, i)) * A$$(i, i) + (\k=(Suc i)..k=(Suc i)..k=0.. = (\k=0..k=i.. = (\k=i.. = conjugate (A$$(i, i)) * A$$(i, i)" using asm zsi1 by (auto simp add: sum.atLeast_Suc_lessThan) finally have "(adjoint A * A)$$(i, i) = conjugate (A$$(i, i)) * A$$(i, i)" . with adaii eq0 have fsitoneq0: "(\k=(Suc i)..k. k i < k \ conjugate (A$$(i, k)) * A$$(i, k) = 0" by (rule sum_nonneg_0[of "{(Suc i)..k. k i A $$ (i, k) = 0" by auto with zsi0 show "\j. j i \ j \ A $$ (i, j) = 0" by (metis linorder_neqE_nat) qed } with case0 show "\i ia. 0 < n \ i < n \ ia < n \ (\k. k < ia \ \j j \ A $$ (k, j) = 0) \ \j j \ A $$ (ia, j) = 0" by auto qed qed lemma normal_complex_mat_has_spectral_decomposition: assumes A: "(A::complex mat) \ carrier_mat n n" and normal: "A * adjoint A = adjoint A * A" and c: "char_poly A = (\ (e :: complex) \ es. [:- e, 1:])" and B: "unitary_schur_decomposition A es = (B,P,Q)" shows "similar_mat_wit A B P (adjoint P) \ diagonal_mat B \ diag_mat B = es \ unitary P" proof - have smw: "similar_mat_wit A B P (adjoint P)" and ut: "upper_triangular B" and uP: "unitary P" and dB: "diag_mat B = es" and "(Q = adjoint P)" using assms by (auto simp add: unitary_schur_decomposition) from smw have dimP: "P \ carrier_mat n n" and dimB: "B \ carrier_mat n n" and dimaP: "adjoint P \ carrier_mat n n" unfolding similar_mat_wit_def using A by auto have dimaB: "adjoint B \ carrier_mat n n" using dimB by auto note dims = dimP dimB dimaP dimaB have "inverts_mat P (adjoint P)" using unitary_def uP dims by auto then have iaPP: "inverts_mat (adjoint P) P" using inverts_mat_symm using dims by auto have aPP: "adjoint P * P = 1\<^sub>m n" using dims iaPP unfolding inverts_mat_def by auto from smw have A: "A = P * B * (adjoint P)" unfolding similar_mat_wit_def Let_def by auto then have aA: "adjoint A = P * adjoint B * adjoint P" by (insert A dimP dimB dimaP, auto simp add: adjoint_mult[of _ n n _ n] adjoint_adjoint) have "A * adjoint A = (P * B * adjoint P) * (P * adjoint B * adjoint P)" using A aA by auto also have "\ = P * B * (adjoint P * P) * (adjoint B * adjoint P)" using dims by (mat_assoc n) also have "\ = P * B * 1\<^sub>m n * (adjoint B * adjoint P)" using dims aPP by (auto) also have "\ = P * B * adjoint B * adjoint P" using dims by (mat_assoc n) finally have "A * adjoint A = P * B * adjoint B * adjoint P". then have "adjoint P * (A * adjoint A) * P = (adjoint P * P) * B * adjoint B * (adjoint P * P)" using dims by (simp add: assoc_mult_mat[of _ n n _ n _ n]) also have "\ = 1\<^sub>m n * B * adjoint B * 1\<^sub>m n" using aPP by auto also have "\ = B * adjoint B" using dims by auto finally have eq0: "adjoint P * (A * adjoint A) * P = B * adjoint B". have "adjoint A * A = (P * adjoint B * adjoint P) * (P * B * adjoint P)" using A aA by auto also have "\ = P * adjoint B * (adjoint P * P) * (B * adjoint P)" using dims by (mat_assoc n) also have "\ = P * adjoint B * 1\<^sub>m n * (B * adjoint P)" using dims aPP by (auto) also have "\ = P * adjoint B * B * adjoint P" using dims by (mat_assoc n) finally have "adjoint A * A = P * adjoint B * B * adjoint P" by auto then have "adjoint P * (adjoint A * A) * P = (adjoint P * P) * adjoint B * B * (adjoint P * P)" using dims by (simp add: assoc_mult_mat[of _ n n _ n _ n]) also have "\ = 1\<^sub>m n * adjoint B * B * 1\<^sub>m n" using aPP by auto also have "\ = adjoint B * B" using dims by auto finally have eq1: "adjoint P * (adjoint A * A) * P = adjoint B * B". from normal have "adjoint P * (adjoint A * A) * P = adjoint P * (A * adjoint A) * P" by auto with eq0 eq1 have "B * adjoint B = adjoint B * B" by auto with ut dims have "diagonal_mat B" using normal_upper_triangular_matrix_is_diagonal by auto with smw uP dB show "similar_mat_wit A B P (adjoint P) \ diagonal_mat B \ diag_mat B = es \ unitary P" by auto qed lemma complex_mat_has_jordan_nf: fixes A :: "complex mat" assumes "A \ carrier_mat n n" shows "\n_as. jordan_nf A n_as" proof - have "\as. char_poly A = (\ a \ as. [:- a, 1:]) \ length as = n" using assms by (simp add: complex_mat_char_poly_factorizable) then show ?thesis using assms by (auto simp add: jordan_nf_iff_linear_factorization) qed lemma hermitian_is_normal: assumes "hermitian A" shows "A * adjoint A = adjoint A * A" using assms by (auto simp add: hermitian_def) lemma hermitian_eigenvalue_real: assumes dim: "(A::complex mat) \ carrier_mat n n" and hA: "hermitian A" and c: "char_poly A = (\ (e :: complex) \ es. [:- e, 1:])" and B: "unitary_schur_decomposition A es = (B,P,Q)" shows "similar_mat_wit A B P (adjoint P) \ diagonal_mat B \ diag_mat B = es \ unitary P \ (\i < n. B$$(i, i) \ Reals)" proof - have normal: "A * adjoint A = adjoint A * A" using hA hermitian_is_normal by auto then have schur: "similar_mat_wit A B P (adjoint P) \ diagonal_mat B \ diag_mat B = es \ unitary P" using normal_complex_mat_has_spectral_decomposition[OF dim normal c B] by (simp) then have "similar_mat_wit A B P (adjoint P)" and uP: "unitary P" and dB: "diag_mat B = es" using assms by auto then have A: "A = P * B * (adjoint P)" and dimB: "B \ carrier_mat n n" and dimP: "P \ carrier_mat n n" unfolding similar_mat_wit_def Let_def using dim by auto then have dimaB: "adjoint B \ carrier_mat n n" by auto have "adjoint A = adjoint (adjoint P) * adjoint (P * B)" apply (subst A) apply (subst adjoint_mult[of "P * B" n n "adjoint P" n]) apply (insert dimB dimP, auto) done also have "\ = P * adjoint (P * B)" by (auto simp add: adjoint_adjoint) also have "\ = P * (adjoint B * adjoint P)" using dimB dimP by (auto simp add: adjoint_mult) also have "\ = P * adjoint B * adjoint P" using dimB dimP by (subst assoc_mult_mat[symmetric, of P n n "adjoint B" n "adjoint P" n], auto) finally have aA: "adjoint A = P * adjoint B * adjoint P" . have "A = adjoint A" using hA hermitian_def[of A] by auto then have "P * B * adjoint P = P * adjoint B * adjoint P" using A aA by auto then have BaB: "B = adjoint B" using unitary_elim[OF dimB dimaB dimP] uP by auto { fix i assume "i < n" then have "B$$(i, i) = conjugate (B$$(i, i))" apply (subst BaB) by (insert dimB, simp add: adjoint_eval) then have "B$$(i, i) \ Reals" unfolding conjugate_complex_def using Reals_cnj_iff by auto } then have "\i Reals" by auto with schur show ?thesis by auto qed lemma hermitian_inner_prod_real: assumes dimA: "(A::complex mat) \ carrier_mat n n" and dimv: "v \ carrier_vec n" and hA: "hermitian A" shows "inner_prod v (A *\<^sub>v v) \ Reals" proof - obtain es where es: "char_poly A = (\ (e :: complex) \ es. [:- e, 1:])" using complex_mat_char_poly_factorizable dimA by auto obtain B P Q where "unitary_schur_decomposition A es = (B,P,Q)" by (cases "unitary_schur_decomposition A es", auto) then have "similar_mat_wit A B P (adjoint P) \ diagonal_mat B \ diag_mat B = es \ unitary P \ (\i < n. B$$(i, i) \ Reals)" using hermitian_eigenvalue_real dimA es hA by auto then have A: "A = P * B * (adjoint P)" and dB: "diagonal_mat B" and Bii: "\i. i < n \ B$$(i, i) \ Reals" and dimB: "B \ carrier_mat n n" and dimP: "P \ carrier_mat n n" and dimaP: "adjoint P \ carrier_mat n n" unfolding similar_mat_wit_def Let_def using dimA by auto define w where "w = (adjoint P) *\<^sub>v v" then have dimw: "w \ carrier_vec n" using dimaP dimv by auto from A have "inner_prod v (A *\<^sub>v v) = inner_prod v ((P * B * (adjoint P)) *\<^sub>v v)" by auto also have "\ = inner_prod v ((P * B) *\<^sub>v ((adjoint P) *\<^sub>v v))" using dimP dimB dimv by (subst assoc_mult_mat_vec[of _ n n "adjoint P" n], auto) also have "\ = inner_prod v (P *\<^sub>v (B *\<^sub>v ((adjoint P) *\<^sub>v v)))" using dimP dimB dimv dimaP by (subst assoc_mult_mat_vec[of _ n n "B" n], auto) also have "\ = inner_prod w (B *\<^sub>v w)" unfolding w_def apply (rule adjoint_def_alter[OF _ _ dimP]) apply (insert mult_mat_vec_carrier[OF dimB mult_mat_vec_carrier[OF dimaP dimv]], auto simp add: dimv) done also have "\ = (\i=0..j=0.. = (\i=0..v v) = (\i=0..i. i < n \ B$$(i, i) * conjugate (w$i) * w$i \ Reals" using Bii by (simp add: Reals_cnj_iff) then have "(\i=0.. Reals" by auto then show ?thesis using sum by auto qed lemma unit_vec_bracket: fixes A :: "complex mat" assumes dimA: "A \ carrier_mat n n" and i: "i < n" shows "inner_prod (unit_vec n i) (A *\<^sub>v (unit_vec n i)) = A$$(i, i)" proof - define w where "(w::complex vec) = unit_vec n i" have "A *\<^sub>v w = col A i" using i dimA w_def by auto then have 1: "inner_prod w (A *\<^sub>v w) = inner_prod w (col A i)" using w_def by auto have "conjugate w = w" unfolding w_def unit_vec_def conjugate_vec_def using i by auto then have 2: "inner_prod w (col A i) = A$$(i, i)" using i dimA w_def by auto from 1 2 show "inner_prod w (A *\<^sub>v w) = A$$(i, i)" by auto qed lemma spectral_decomposition_extract_diag: fixes P B :: "complex mat" assumes dimP: "P \ carrier_mat n n" and dimB: "B \ carrier_mat n n" and uP: "unitary P" and dB: "diagonal_mat B" and i: "i < n" shows "inner_prod (col P i) (P * B * (adjoint P) *\<^sub>v (col P i)) = B$$(i, i)" proof - have dimaP: "adjoint P\ carrier_mat n n" using dimP by auto have uaP: "unitary (adjoint P)" using unitary_adjoint uP dimP by auto then have "inverts_mat (adjoint P) P" by (simp add: unitary_def adjoint_adjoint) then have iv: "(adjoint P) * P = 1\<^sub>m n" using dimaP inverts_mat_def by auto define v where "v = col P i" then have dimv: "v \ carrier_vec n" using dimP by auto define w where "(w::complex vec) = unit_vec n i" then have dimw: "w \ carrier_vec n" by auto have BaPv: "B *\<^sub>v (adjoint P *\<^sub>v v) \ carrier_vec n" using dimB dimaP dimv by auto have "(adjoint P) *\<^sub>v v = (col (adjoint P * P) i)" by (simp add: col_mult2[OF dimaP dimP i, symmetric] v_def) then have aPv: "(adjoint P) *\<^sub>v v = w" by (auto simp add: iv i w_def) have "inner_prod v (P * B * (adjoint P) *\<^sub>v v) = inner_prod v ((P * B) *\<^sub>v ((adjoint P) *\<^sub>v v))" using dimP dimB dimv by (subst assoc_mult_mat_vec[of _ n n "adjoint P" n], auto) also have "\ = inner_prod v (P *\<^sub>v (B *\<^sub>v ((adjoint P) *\<^sub>v v)))" using dimP dimB dimv dimaP by (subst assoc_mult_mat_vec[of _ n n "B" n], auto) also have "\ = inner_prod (adjoint P *\<^sub>v v) (B *\<^sub>v (adjoint P *\<^sub>v v))" by (simp add: adjoint_def_alter[OF dimv BaPv dimP]) also have "\ = inner_prod w (B *\<^sub>v w)" using aPv by auto also have "\ = B$$(i, i)" using w_def unit_vec_bracket dimB i by auto finally show "inner_prod v (P * B * (adjoint P) *\<^sub>v v) = B$$(i, i)". qed lemma hermitian_inner_prod_zero: fixes A :: "complex mat" assumes dimA: "A \ carrier_mat n n" and hA: "hermitian A" and zero: "\v\carrier_vec n. inner_prod v (A *\<^sub>v v) = 0" shows "A = 0\<^sub>m n n" proof - obtain es where es: "char_poly A = (\ (e :: complex) \ es. [:- e, 1:])" using complex_mat_char_poly_factorizable dimA by auto obtain B P Q where "unitary_schur_decomposition A es = (B,P,Q)" by (cases "unitary_schur_decomposition A es", auto) then have "similar_mat_wit A B P (adjoint P) \ diagonal_mat B \ diag_mat B = es \ unitary P \ (\i < n. B$$(i, i) \ Reals)" using hermitian_eigenvalue_real dimA es hA by auto then have A: "A = P * B * (adjoint P)" and dB: "diagonal_mat B" and Bii: "\i. i < n \ B$$(i, i) \ Reals" and dimB: "B \ carrier_mat n n" and dimP: "P \ carrier_mat n n" and dimaP: "adjoint P \ carrier_mat n n" and uP: "unitary P" unfolding similar_mat_wit_def Let_def unitary_def using dimA by auto then have uaP: "unitary (adjoint P)" using unitary_adjoint by auto then have "inverts_mat (adjoint P) P" by (simp add: unitary_def adjoint_adjoint) then have iv: "adjoint P * P = 1\<^sub>m n" using dimaP inverts_mat_def by auto have "B = 0\<^sub>m n n" proof- { fix i assume i: "i < n" define v where "v = col P i" then have dimv: "v \ carrier_vec n" using v_def dimP by auto have "inner_prod v (A *\<^sub>v v) = B$$(i, i)" unfolding A v_def using spectral_decomposition_extract_diag[OF dimP dimB uP dB i] by auto moreover have "inner_prod v (A *\<^sub>v v) = 0" using dimv zero by auto ultimately have "B$$(i, i) = 0" by auto } note zB = this show "B = 0\<^sub>m n n" by (insert zB dB dimB, rule eq_matI, auto simp add: diagonal_mat_def) qed then show "A = 0\<^sub>m n n" using A dimB dimP dimaP by auto qed lemma complex_mat_decomposition_to_hermitian: fixes A :: "complex mat" assumes dim: "A \ carrier_mat n n" shows "\B C. hermitian B \ hermitian C \ A = B + \ \\<^sub>m C \ B \ carrier_mat n n \ C \ carrier_mat n n" proof - obtain B C where B: "B = (1 / 2) \\<^sub>m (A + adjoint A)" and C: "C = (-\ / 2) \\<^sub>m (A - adjoint A)" by auto then have dimB: "B \ carrier_mat n n" and dimC: "C \ carrier_mat n n" using dim by auto have "hermitian B" unfolding B hermitian_def using dim by (auto simp add: adjoint_eval) moreover have "hermitian C" unfolding C hermitian_def using dim apply (subst eq_matI) apply (auto simp add: adjoint_eval algebra_simps) done moreover have "A = B + \ \\<^sub>m C" using dim B C apply (subst eq_matI) apply (auto simp add: adjoint_eval algebra_simps) done ultimately show ?thesis using dimB dimC by auto qed subsection \Outer product\ definition outer_prod :: "'a::conjugatable_field vec \ 'a vec \ 'a mat" where "outer_prod v w = mat (dim_vec v) 1 (\(i, j). v $ i) * mat 1 (dim_vec w) (\(i, j). (conjugate w) $ j)" lemma outer_prod_dim[simp]: fixes v w :: "'a::conjugatable_field vec" assumes v: "v \ carrier_vec n" and w: "w \ carrier_vec m" shows "outer_prod v w \ carrier_mat n m" unfolding outer_prod_def using assms mat_of_cols_carrier mat_of_rows_carrier by auto lemma mat_of_vec_mult_eq_scalar_prod: fixes v w :: "'a::conjugatable_field vec" assumes "v \ carrier_vec n" and "w \ carrier_vec n" shows "mat 1 (dim_vec v) (\(i, j). (conjugate v) $ j) * mat (dim_vec w) 1 (\(i, j). w $ i) = mat 1 1 (\k. inner_prod v w)" apply (rule eq_matI) using assms apply (simp add: scalar_prod_def) apply (rule sum.cong) by auto lemma one_dim_mat_mult_is_scale: fixes A B :: "('a::conjugatable_field mat)" assumes "B \ carrier_mat 1 n" shows "(mat 1 1 (\k. a)) * B = a \\<^sub>m B" apply (rule eq_matI) using assms by (auto simp add: scalar_prod_def) lemma outer_prod_mult_outer_prod: fixes a b c d :: "'a::conjugatable_field vec" assumes a: "a \ carrier_vec d1" and b: "b \ carrier_vec d2" and c: "c \ carrier_vec d2" and d: "d \ carrier_vec d3" shows "outer_prod a b * outer_prod c d = inner_prod b c \\<^sub>m outer_prod a d" proof - let ?ma = "mat (dim_vec a) 1 (\(i, j). a $ i)" let ?mb = "mat 1 (dim_vec b) (\(i, j). (conjugate b) $ j)" let ?mc = "mat (dim_vec c) 1 (\(i, j). c $ i)" let ?md = "mat 1 (dim_vec d) (\(i, j). (conjugate d) $ j)" have "(?ma * ?mb) * (?mc * ?md) = ?ma * (?mb * (?mc * ?md))" apply (subst assoc_mult_mat[of "?ma" d1 1 "?mb" d2 "?mc * ?md" d3] ) using assms by auto also have "\ = ?ma * ((?mb * ?mc) * ?md)" apply (subst assoc_mult_mat[symmetric, of "?mb" 1 d2 "?mc" 1 "?md" d3]) using assms by auto also have "\ = ?ma * ((mat 1 1 (\k. inner_prod b c)) * ?md)" apply (subst mat_of_vec_mult_eq_scalar_prod[of b d2 c]) using assms by auto also have "\ = ?ma * (inner_prod b c \\<^sub>m ?md)" apply (subst one_dim_mat_mult_is_scale) using assms by auto also have "\ = (inner_prod b c) \\<^sub>m (?ma * ?md)" using assms by auto finally show ?thesis unfolding outer_prod_def by auto qed lemma index_outer_prod: fixes v w :: "'a::conjugatable_field vec" assumes v: "v \ carrier_vec n" and w: "w \ carrier_vec m" and ij: "i < n" "j < m" shows "(outer_prod v w)$$(i, j) = v $ i * conjugate (w $ j)" unfolding outer_prod_def using assms by (simp add: scalar_prod_def) lemma mat_of_vec_mult_vec: fixes a b c :: "'a::conjugatable_field vec" assumes a: "a \ carrier_vec d" and b: "b \ carrier_vec d" shows "mat 1 d (\(i, j). (conjugate a) $ j) *\<^sub>v b = vec 1 (\k. inner_prod a b)" apply (rule eq_vecI) apply (simp add: scalar_prod_def carrier_vecD[OF a] carrier_vecD[OF b]) apply (rule sum.cong) by auto lemma mat_of_vec_mult_one_dim_vec: fixes a b :: "'a::conjugatable_field vec" assumes a: "a \ carrier_vec d" shows "mat d 1 (\(i, j). a $ i) *\<^sub>v vec 1 (\k. c) = c \\<^sub>v a" apply (rule eq_vecI) by (auto simp add: scalar_prod_def carrier_vecD[OF a]) lemma outer_prod_mult_vec: fixes a b c :: "'a::conjugatable_field vec" assumes a: "a \ carrier_vec d1" and b: "b \ carrier_vec d2" and c: "c \ carrier_vec d2" shows "outer_prod a b *\<^sub>v c = inner_prod b c \\<^sub>v a" proof - have "outer_prod a b *\<^sub>v c = mat d1 1 (\(i, j). a $ i) * mat 1 d2 (\(i, j). (conjugate b) $ j) *\<^sub>v c" unfolding outer_prod_def using assms by auto also have "\ = mat d1 1 (\(i, j). a $ i) *\<^sub>v (mat 1 d2 (\(i, j). (conjugate b) $ j) *\<^sub>v c)" apply (subst assoc_mult_mat_vec) using assms by auto also have "\ = mat d1 1 (\(i, j). a $ i) *\<^sub>v vec 1 (\k. inner_prod b c)" using mat_of_vec_mult_vec[of b] assms by auto also have "\ = inner_prod b c \\<^sub>v a" using mat_of_vec_mult_one_dim_vec assms by auto finally show ?thesis by auto qed lemma trace_outer_prod_right: fixes A :: "'a::conjugatable_field mat" and v w :: "'a vec" assumes A: "A \ carrier_mat n n" and v: "v \ carrier_vec n" and w: "w \ carrier_vec n" shows "trace (A * outer_prod v w) = inner_prod w (A *\<^sub>v v)" (is "?lhs = ?rhs") proof - define B where "B = outer_prod v w" then have B: "B \ carrier_mat n n" using assms by auto have "trace(A * B) = (\i = 0..j = 0.. = (\i = 0..j = 0..i = 0..j = 0..i = 0..j = 0.. carrier_vec n" and w: "w \ carrier_vec n" shows "trace (outer_prod v w) = inner_prod w v" (is "?lhs = ?rhs") proof - have "(1\<^sub>m n) * (outer_prod v w) = outer_prod v w" apply (subst left_mult_one_mat) using outer_prod_dim assms by auto moreover have "1\<^sub>m n *\<^sub>v v = v" using assms by auto ultimately show ?thesis using trace_outer_prod_right[of "1\<^sub>m n" n v w] assms by auto qed lemma inner_prod_outer_prod: fixes a b c d :: "'a::conjugatable_field vec" assumes a: "a \ carrier_vec n" and b: "b \ carrier_vec n" and c: "c \ carrier_vec m" and d: "d \ carrier_vec m" shows "inner_prod a (outer_prod b c *\<^sub>v d) = inner_prod a b * inner_prod c d" (is "?lhs = ?rhs") proof - define P where "P = outer_prod b c" then have dimP: "P \ carrier_mat n m" using assms by auto have "inner_prod a (P *\<^sub>v d) = (\i=0..j=0.. = (\i=0..j=0..i=0..j=0..i=0..j=0.. = (\i=0..j=0..Semi-definite matrices\ definition positive :: "complex mat \ bool" where "positive A \ A \ carrier_mat (dim_col A) (dim_col A) \ (\v. dim_vec v = dim_col A \ inner_prod v (A *\<^sub>v v) \ 0)" lemma positive_iff_normalized_vec: "positive A \ A \ carrier_mat (dim_col A) (dim_col A) \ (\v. (dim_vec v = dim_col A \ vec_norm v = 1) \ inner_prod v (A *\<^sub>v v) \ 0)" proof (rule) assume "positive A" then show "A \ carrier_mat (dim_col A) (dim_col A) \ (\v. dim_vec v = dim_col A \ vec_norm v = 1 \ 0 \ inner_prod v (A *\<^sub>v v))" unfolding positive_def by auto next define n where "n = dim_col A" assume "A \ carrier_mat (dim_col A) (dim_col A) \ (\v. dim_vec v = dim_col A \ vec_norm v = 1 \ 0 \ inner_prod v (A *\<^sub>v v))" then have A: "A \ carrier_mat (dim_col A) (dim_col A)" and geq0: "\v. dim_vec v = dim_col A \ vec_norm v = 1 \ 0 \ inner_prod v (A *\<^sub>v v)" by auto then have dimA: "A \ carrier_mat n n" using n_def[symmetric] by auto { fix v assume dimv: "(v::complex vec) \ carrier_vec n" have "0 \ inner_prod v (A *\<^sub>v v)" proof (cases "v = 0\<^sub>v n") case True then show "0 \ inner_prod v (A *\<^sub>v v)" using dimA by auto next case False then have 1: "vec_norm v > 0" using vec_norm_ge_0 dimv by auto then have cnv: "cnj (vec_norm v) = vec_norm v" using Reals_cnj_iff complex_is_Real_iff less_complex_def by auto define w where "w = vec_normalize v" then have dimw: "w \ carrier_vec n" using dimv by auto have nvw: "v = vec_norm v \\<^sub>v w" using w_def vec_eq_norm_smult_normalized by auto have "vec_norm w = 1" using normalized_vec_norm[OF dimv False] vec_norm_def w_def by auto then have 2: "0 \ inner_prod w (A *\<^sub>v w)" using geq0 dimw dimA by auto have "inner_prod v (A *\<^sub>v v) = vec_norm v * vec_norm v * inner_prod w (A *\<^sub>v w)" using dimA dimv dimw apply (subst (1 2) nvw) apply (subst mult_mat_vec, simp, simp) apply (subst scalar_prod_smult_left[of "(A *\<^sub>v w)" "conjugate (vec_norm v \\<^sub>v w)" "vec_norm v"], simp) apply (simp add: conjugate_smult_vec cnv) done also have "\ \ 0" using 1 2 by auto finally show "0 \ inner_prod v (A *\<^sub>v v)" by auto qed } then have geq: "\v. dim_vec v = dim_col A \ 0 \ inner_prod v (A *\<^sub>v v)" using dimA by auto show "positive A" unfolding positive_def by (rule, simp add: A, rule geq) qed lemma positive_is_hermitian: fixes A :: "complex mat" assumes pA: "positive A" shows "hermitian A" proof - define n where "n = dim_col A" then have dimA: "A \ carrier_mat n n" using positive_def pA by auto obtain B C where B: "hermitian B" and C: "hermitian C" and A: "A = B + \ \\<^sub>m C" and dimB: "B \ carrier_mat n n" and dimC: "C \ carrier_mat n n" and dimiC: "\ \\<^sub>m C \ carrier_mat n n" using complex_mat_decomposition_to_hermitian[OF dimA] by auto { fix v :: "complex vec" assume dimv: "v \ carrier_vec n" have dimvA: "dim_vec v = dim_col A" using dimv dimA by auto have "inner_prod v (A *\<^sub>v v) = inner_prod v (B *\<^sub>v v) + inner_prod v ((\ \\<^sub>m C) *\<^sub>v v)" unfolding A using dimB dimiC dimv by (simp add: add_mult_distrib_mat_vec inner_prod_distrib_right) moreover have "inner_prod v ((\ \\<^sub>m C) *\<^sub>v v) = \ * inner_prod v (C *\<^sub>v v)" using dimv dimC apply (simp add: scalar_prod_def sum_distrib_left cong: sum.cong) apply (rule sum.cong, auto) done ultimately have ABC: "inner_prod v (A *\<^sub>v v) = inner_prod v (B *\<^sub>v v) + \ * inner_prod v (C *\<^sub>v v)" by auto moreover have "inner_prod v (B *\<^sub>v v) \ Reals" using B dimB dimv hermitian_inner_prod_real by auto moreover have "inner_prod v (C *\<^sub>v v) \ Reals" using C dimC dimv hermitian_inner_prod_real by auto moreover have "inner_prod v (A *\<^sub>v v) \ Reals" using pA unfolding positive_def apply (rule) apply (fold n_def) apply (simp add: complex_is_Real_iff[of "inner_prod v (A *\<^sub>v v)"]) apply (auto simp add: dimvA less_complex_def less_eq_complex_def) done ultimately have "inner_prod v (C *\<^sub>v v) = 0" using of_real_Re by fastforce } then have "C = 0\<^sub>m n n" using hermitian_inner_prod_zero dimC C by auto then have "A = B" using A dimC dimB by auto then show "hermitian A" using B by auto qed lemma positive_eigenvalue_positive: assumes dimA: "(A::complex mat) \ carrier_mat n n" and pA: "positive A" and c: "char_poly A = (\ (e :: complex) \ es. [:- e, 1:])" and B: "unitary_schur_decomposition A es = (B,P,Q)" shows "\i. i < n \ B$$(i, i) \ 0" proof - have hA: "hermitian A" using positive_is_hermitian pA by auto have "similar_mat_wit A B P (adjoint P) \ diagonal_mat B \ diag_mat B = es \ unitary P \ (\i < n. B$$(i, i) \ Reals)" using hermitian_eigenvalue_real dimA hA B c by auto then have A: "A = P * B * (adjoint P)" and dB: "diagonal_mat B" and Bii: "\i. i < n \ B$$(i, i) \ Reals" and dimB: "B \ carrier_mat n n" and dimP: "P \ carrier_mat n n" and dimaP: "adjoint P \ carrier_mat n n" and uP: "unitary P" unfolding similar_mat_wit_def Let_def unitary_def using dimA by auto { fix i assume i: "i < n" define v where "v = col P i" then have dimv: "v \ carrier_vec n" using v_def dimP by auto have "inner_prod v (A *\<^sub>v v) = B$$(i, i)" unfolding A v_def using spectral_decomposition_extract_diag[OF dimP dimB uP dB i] by auto moreover have "inner_prod v (A *\<^sub>v v) \ 0" using dimv pA dimA positive_def by auto ultimately show "B$$(i, i) \ 0" by auto } qed lemma diag_mat_mult_diag_mat: fixes B D :: "'a::semiring_0 mat" assumes dimB: "B \ carrier_mat n n" and dimD: "D \ carrier_mat n n" and dB: "diagonal_mat B" and dD: "diagonal_mat D" shows "B * D = mat n n (\(i,j). (if i = j then (B$$(i, i)) * (D$$(i, i)) else 0))" proof(rule eq_matI, auto) have Bij: "\x y. x < n \ y < n \ x \ y \ B$$(x, y) = 0" using dB diagonal_mat_def dimB by auto have Dij: "\x y. x < n \ y < n \ x \ y \ D$$(x, y) = 0" using dD diagonal_mat_def dimD by auto { fix i j assume ij: "i < n" "j < n" have "(B * D) $$ (i, j) = (\k=0.. = B$$(i, i) * D$$(i, j)" apply (simp add: sum.remove[of _i] ij) apply (simp add: Bij Dij ij) done finally have "(B * D) $$ (i, j) = B$$(i, i) * D$$(i, j)". } note BDij = this from BDij show "\j. j < n \ (B * D) $$ (j, j) = B $$ (j, j) * D $$ (j, j)" by auto from BDij show "\i j. i < n \ j < n \ i \ j \ (B * D) $$ (i, j) = 0" using Bij Dij by auto from assms show "dim_row B = n" "dim_col D = n" by auto qed lemma positive_only_if_decomp: assumes dimA: "A \ carrier_mat n n" and pA: "positive A" shows "\M \ carrier_mat n n. M * adjoint M = A" proof - from pA have hA: "hermitian A" using positive_is_hermitian by auto obtain es where es: "char_poly A = (\ (e :: complex) \ es. [:- e, 1:])" using complex_mat_char_poly_factorizable dimA by auto obtain B P Q where schur: "unitary_schur_decomposition A es = (B,P,Q)" by (cases "unitary_schur_decomposition A es", auto) then have "similar_mat_wit A B P (adjoint P) \ diagonal_mat B \ diag_mat B = es \ unitary P \ (\i < n. B$$(i, i) \ Reals)" using hermitian_eigenvalue_real dimA es hA by auto then have A: "A = P * B * (adjoint P)" and dB: "diagonal_mat B" and Bii: "\i. i < n \ B$$(i, i) \ Reals" and dimB: "B \ carrier_mat n n" and dimP: "P \ carrier_mat n n" and dimaP: "adjoint P \ carrier_mat n n" unfolding similar_mat_wit_def Let_def using dimA by auto have Bii: "\i. i < n \ B$$(i, i) \ 0" using pA dimA es schur positive_eigenvalue_positive by auto define D where "D = mat n n (\(i, j). (if (i = j) then csqrt (B$$(i, i)) else 0))" then have dimD: "D \ carrier_mat n n" and dimaD: "adjoint D \ carrier_mat n n" using dimB by auto have dD: "diagonal_mat D" using dB D_def unfolding diagonal_mat_def by auto then have daD: "diagonal_mat (adjoint D)" by (simp add: adjoint_eval diagonal_mat_def) have Dii: "\i. i < n \ D$$(i, i) = csqrt (B$$(i, i))" using dimD D_def by auto { fix i assume i: "i < n" define c where "c = csqrt (B$$(i, i))" have c: "c \ 0" using Bii i c_def by (auto simp: less_complex_def less_eq_complex_def) then have "conjugate c = c" using Reals_cnj_iff complex_is_Real_iff unfolding less_complex_def less_eq_complex_def by auto then have "c * cnj c = B$$(i, i)" using c_def c unfolding conjugate_complex_def by (metis power2_csqrt power2_eq_square) } note cBii = this have "D * adjoint D = mat n n (\(i,j). (if (i = j) then B$$(i, i) else 0))" apply (simp add: diag_mat_mult_diag_mat[OF dimD dimaD dD daD]) apply (rule eq_matI, auto simp add: D_def adjoint_eval cBii) done also have "\ = B" using dimB dB[unfolded diagonal_mat_def] by auto finally have DaDB: "D * adjoint D = B". define M where "M = P * D" then have dimM: "M \ carrier_mat n n" using dimP dimD by auto have "M * adjoint M = (P * D) * (adjoint D * adjoint P)" using M_def adjoint_mult[OF dimP dimD] by auto also have "\ = P * (D * adjoint D) * (adjoint P)" using dimP dimD by (mat_assoc n) also have "\ = P * B * (adjoint P)" using DaDB by auto finally have "M * adjoint M = A" using A by auto with dimM show "\M \ carrier_mat n n. M * adjoint M = A" by auto qed lemma positive_if_decomp: assumes dimA: "A \ carrier_mat n n" and "\M. M * adjoint M = A" shows "positive A" proof - from assms obtain M where M: "M * adjoint M = A" by auto define m where "m = dim_col M" have dimM: "M \ carrier_mat n m" using M dimA m_def by auto { fix v assume dimv: "(v::complex vec) \ carrier_vec n" have dimaM: "adjoint M \ carrier_mat m n" using dimM by auto have dimaMv: "(adjoint M) *\<^sub>v v \ carrier_vec m" using dimaM dimv by auto have "inner_prod v (A *\<^sub>v v) = inner_prod v (M * adjoint M *\<^sub>v v)" using M by auto also have "\ = inner_prod v (M *\<^sub>v (adjoint M *\<^sub>v v))" using assoc_mult_mat_vec dimM dimaM dimv by auto also have "\ = inner_prod (adjoint M *\<^sub>v v) (adjoint M *\<^sub>v v)" using adjoint_def_alter[OF dimv dimaMv dimM] by auto also have "\ \ 0" using self_cscalar_prod_geq_0 by auto finally have "inner_prod v (A *\<^sub>v v) \ 0". } note geq0 = this from dimA geq0 show "positive A" using positive_def by auto qed lemma positive_iff_decomp: assumes dimA: "A \ carrier_mat n n" shows "positive A \ (\M\carrier_mat n n. M * adjoint M = A)" proof assume pA: "positive A" then show "\M\carrier_mat n n. M * adjoint M = A" using positive_only_if_decomp assms by auto next assume "\M\carrier_mat n n. M * adjoint M = A" then obtain M where M: "M * adjoint M = A" by auto then show "positive A" using M positive_if_decomp assms by auto qed lemma positive_dim_eq: assumes "positive A" shows "dim_row A = dim_col A" using carrier_matD(1)[of A "dim_col A" "dim_col A"] assms[unfolded positive_def] by simp lemma positive_zero: "positive (0\<^sub>m n n)" by (simp add: positive_def zero_mat_def mult_mat_vec_def scalar_prod_def) lemma positive_one: "positive (1\<^sub>m n)" proof (rule positive_if_decomp) show "1\<^sub>m n \ carrier_mat n n" by auto have "adjoint (1\<^sub>m n) = 1\<^sub>m n" using hermitian_one hermitian_def by auto then have "1\<^sub>m n * adjoint (1\<^sub>m n) = 1\<^sub>m n" by auto then show "\M. M * adjoint M = 1\<^sub>m n" by fastforce qed lemma positive_antisym: assumes pA: "positive A" and pnA: "positive (-A)" shows "A = 0\<^sub>m (dim_col A) (dim_col A)" proof - define n where "n = dim_col A" from pA have dimA: "A \ carrier_mat n n" and dimnA: "-A \ carrier_mat n n" using positive_def n_def by auto from pA have hA: "hermitian A" using positive_is_hermitian by auto obtain es where es: "char_poly A = (\ (e :: complex) \ es. [:- e, 1:])" using complex_mat_char_poly_factorizable dimA by auto obtain B P Q where schur: "unitary_schur_decomposition A es = (B,P,Q)" by (cases "unitary_schur_decomposition A es", auto) then have "similar_mat_wit A B P (adjoint P) \ diagonal_mat B \ unitary P" using hermitian_eigenvalue_real dimA es hA by auto then have A: "A = P * B * (adjoint P)" and dB: "diagonal_mat B" and uP: "unitary P" and dimB: "B \ carrier_mat n n" and dimnB: "-B \ carrier_mat n n" and dimP: "P \ carrier_mat n n" and dimaP: "adjoint P \ carrier_mat n n" unfolding similar_mat_wit_def Let_def using dimA by auto from es schur have geq0: "\i. i < n \ B$$(i, i) \ 0" using positive_eigenvalue_positive dimA pA by auto from A have nA: "-A = P * (-B) * (adjoint P)" using mult_smult_assoc_mat dimB dimP dimaP by auto from dB have dnB: "diagonal_mat (-B)" by (simp add: diagonal_mat_def) { fix i assume i: "i < n" define v where "v = col P i" then have dimv: "v \ carrier_vec n" using v_def dimP by auto have "inner_prod v ((-A) *\<^sub>v v) = (-B)$$(i, i)" unfolding nA v_def using spectral_decomposition_extract_diag[OF dimP dimnB uP dnB i] by auto moreover have "inner_prod v ((-A) *\<^sub>v v) \ 0" using dimv pnA dimnA positive_def by auto ultimately have "B$$(i, i) \ 0" using dimB i by auto moreover have "B$$(i, i) \ 0" using i geq0 by auto ultimately have "B$$(i, i) = 0" by (metis antisym) } then have "B = 0\<^sub>m n n" using dimB dB[unfolded diagonal_mat_def] by (subst eq_matI, auto) then show "A = 0\<^sub>m n n" using A dimB dimP dimaP by auto qed lemma positive_add: assumes pA: "positive A" and pB: "positive B" and dimA: "A \ carrier_mat n n" and dimB: "B \ carrier_mat n n" shows "positive (A + B)" unfolding positive_def proof have dimApB: "A + B \ carrier_mat n n" using dimA dimB by auto then show "A + B \ carrier_mat (dim_col (A + B)) (dim_col (A + B))" using carrier_matD[of "A+B"] by auto { fix v assume dimv: "(v::complex vec) \ carrier_vec n" have 1: "inner_prod v (A *\<^sub>v v) \ 0" using dimv pA[unfolded positive_def] dimA by auto have 2: "inner_prod v (B *\<^sub>v v) \ 0" using dimv pB[unfolded positive_def] dimB by auto have "inner_prod v ((A + B) *\<^sub>v v) = inner_prod v (A *\<^sub>v v) + inner_prod v (B *\<^sub>v v)" using dimA dimB dimv by (simp add: add_mult_distrib_mat_vec inner_prod_distrib_right) also have "\ \ 0" using 1 2 by auto finally have "inner_prod v ((A + B) *\<^sub>v v) \ 0". } note geq0 = this then have "\v. dim_vec v = n \ 0 \ inner_prod v ((A + B) *\<^sub>v v)" by auto then show "\v. dim_vec v = dim_col (A + B) \ 0 \ inner_prod v ((A + B) *\<^sub>v v)" using dimApB by auto qed lemma positive_trace: assumes "A \ carrier_mat n n" and "positive A" shows "trace A \ 0" using assms positive_iff_decomp trace_adjoint_positive by auto lemma positive_close_under_left_right_mult_adjoint: fixes M A :: "complex mat" assumes dM: "M \ carrier_mat n n" and dA: "A \ carrier_mat n n" and pA: "positive A" shows "positive (M * A * adjoint M)" unfolding positive_def proof (rule, simp add: mult_carrier_mat[OF mult_carrier_mat[OF dM dA] adjoint_dim[OF dM]] carrier_matD[OF dM], rule, rule) have daM: "adjoint M \ carrier_mat n n" using dM by auto fix v::"complex vec" assume "dim_vec v = dim_col (M * A * adjoint M)" then have dv: "v \ carrier_vec n" using assms by auto then have "adjoint M *\<^sub>v v \ carrier_vec n" using daM by auto have assoc: "M * A * adjoint M *\<^sub>v v = M *\<^sub>v (A *\<^sub>v (adjoint M *\<^sub>v v))" using dA dM daM dv by (auto simp add: assoc_mult_mat_vec[of _ n n _ n]) have "inner_prod v (M * A * adjoint M *\<^sub>v v) = inner_prod (adjoint M *\<^sub>v v) (A *\<^sub>v (adjoint M *\<^sub>v v))" apply (subst assoc) apply (subst adjoint_def_alter[where ?A = "M"]) by (auto simp add: dv dA daM dM carrier_matD[OF dM] mult_mat_vec_carrier[of _ n n]) also have "\ \ 0" using dA dv daM pA positive_def by auto finally show "inner_prod v (M * A * adjoint M *\<^sub>v v) \ 0" by auto qed lemma positive_same_outer_prod: fixes v w :: "complex vec" assumes v: "v \ carrier_vec n" shows "positive (outer_prod v v)" proof - have d1: "adjoint (mat (dim_vec v) 1 (\(i, j). v $ i)) \ carrier_mat 1 n" using assms by auto have d2: "mat 1 (dim_vec v) (\(i, y). conjugate v $ y) \ carrier_mat 1 n" using assms by auto have dv: "dim_vec v = n" using assms by auto have "mat 1 (dim_vec v) (\(i, y). conjugate v $ y) = adjoint (mat (dim_vec v) 1 (\(i, j). v $ i))" (is "?r = adjoint ?l") apply (rule eq_matI) subgoal for i j by (simp add: dv adjoint_eval) using d1 d2 by auto then have "outer_prod v v = ?l * adjoint ?l" unfolding outer_prod_def by auto then have "\M. M * adjoint M = outer_prod v v" by auto then show "positive (outer_prod v v)" using positive_if_decomp[OF outer_prod_dim[OF v v]] by auto qed lemma smult_smult_mat: fixes k :: complex and l :: complex assumes "A \ carrier_mat nr n" shows "k \\<^sub>m (l \\<^sub>m A) = (k * l) \\<^sub>m A" by auto lemma positive_smult: assumes "A \ carrier_mat n n" and "positive A" and "c \ 0" shows "positive (c \\<^sub>m A)" proof - have sc: "csqrt c \ 0" using assms(3) by (fastforce simp: less_eq_complex_def) obtain M where dimM: "M \ carrier_mat n n" and A: "M * adjoint M = A" using assms(1-2) positive_iff_decomp by auto have "c \\<^sub>m A = c \\<^sub>m (M * adjoint M)" using A by auto have ccsq: "conjugate (csqrt c) = (csqrt c)" using sc Reals_cnj_iff[of "csqrt c"] complex_is_Real_iff by (auto simp: less_eq_complex_def) have MM: "(M * adjoint M) \ carrier_mat n n" using A assms by fastforce have leftd: "c \\<^sub>m (M * adjoint M) \ carrier_mat n n" using A assms by fastforce have rightd: "(csqrt c \\<^sub>m M) * (adjoint (csqrt c \\<^sub>m M))\ carrier_mat n n" using A assms by fastforce have "(csqrt c \\<^sub>m M) * (adjoint (csqrt c \\<^sub>m M)) = (csqrt c \\<^sub>m M) * ((conjugate (csqrt c)) \\<^sub>m adjoint M)" using adjoint_scale assms(1) by (metis adjoint_scale) also have "\ = (csqrt c \\<^sub>m M) * (csqrt c \\<^sub>m adjoint M)" using sc ccsq by fastforce also have "\ = csqrt c \\<^sub>m (M * (csqrt c \\<^sub>m adjoint M))" using mult_smult_assoc_mat index_smult_mat(2,3) by fastforce also have "\ = csqrt c \\<^sub>m ((csqrt c) \\<^sub>m (M * adjoint M))" using mult_smult_distrib by fastforce also have "\ = c \\<^sub>m (M * adjoint M)" using smult_smult_mat[of "M * adjoint M" n n "(csqrt c)" "(csqrt c)"] MM sc by (metis power2_csqrt power2_eq_square ) also have "\ = c \\<^sub>m A" using A by auto finally have "(csqrt c \\<^sub>m M) * (adjoint (csqrt c \\<^sub>m M)) = c \\<^sub>m A" by auto moreover have "c \\<^sub>m A \ carrier_mat n n" using assms(1) by auto moreover have "csqrt c \\<^sub>m M \ carrier_mat n n" using dimM by auto ultimately show ?thesis using positive_iff_decomp by auto qed text \Version of previous theorem for real numbers\ lemma positive_scale: fixes c :: real assumes "A \ carrier_mat n n" and "positive A" and "c \ 0" shows "positive (c \\<^sub>m A)" apply (rule positive_smult) using assms by (auto simp: less_eq_complex_def) subsection \L\"{o}wner partial order\ definition lowner_le :: "complex mat \ complex mat \ bool" (infix "\\<^sub>L" 50) where "A \\<^sub>L B \ dim_row A = dim_row B \ dim_col A = dim_col B \ positive (B - A)" lemma lowner_le_refl: assumes "A \ carrier_mat n n" shows "A \\<^sub>L A" unfolding lowner_le_def apply (simp add: minus_r_inv_mat[OF assms]) by (rule positive_zero) lemma lowner_le_antisym: assumes A: "A \ carrier_mat n n" and B: "B \ carrier_mat n n" and L1: "A \\<^sub>L B" and L2: "B \\<^sub>L A" shows "A = B" proof - from L1 have P1: "positive (B - A)" by (simp add: lowner_le_def) from L2 have P2: "positive (A - B)" by (simp add: lowner_le_def) have "A - B = - (B - A)" using A B by auto then have P3: "positive (- (B - A))" using P2 by auto have BA: "B - A \ carrier_mat n n" using A B by auto have "B - A = 0\<^sub>m n n" using BA by (subst positive_antisym[OF P1 P3], auto) then have "B + (-A) + A = 0\<^sub>m n n + A" using A B minus_add_uminus_mat[OF B A] by auto then have "B + (-A + A) = 0\<^sub>m n n + A" using A B by auto then show "A = B" using A B BA uminus_l_inv_mat[OF A] by auto qed lemma lowner_le_inner_prod_le: fixes A B :: "complex mat" and v :: "complex vec" assumes A: "A \ carrier_mat n n" and B: "B \ carrier_mat n n" and v: "v \ carrier_vec n" and "A \\<^sub>L B" shows "inner_prod v (A *\<^sub>v v) \ inner_prod v (B *\<^sub>v v)" proof - from assms have "positive (B-A)" by (auto simp add: lowner_le_def) with assms have geq: "inner_prod v ((B-A) *\<^sub>v v) \ 0" unfolding positive_def by auto have "inner_prod v ((B-A) *\<^sub>v v) = inner_prod v (B *\<^sub>v v) - inner_prod v (A *\<^sub>v v)" unfolding minus_add_uminus_mat[OF B A] by (subst add_mult_distrib_mat_vec[OF B _ v], insert A B v, auto simp add: inner_prod_distrib_right[OF v]) then show ?thesis using geq by auto qed lemma lowner_le_trans: fixes A B C :: "complex mat" assumes A: "A \ carrier_mat n n" and B: "B \ carrier_mat n n" and C: "C \ carrier_mat n n" and L1: "A \\<^sub>L B" and L2: "B \\<^sub>L C" shows "A \\<^sub>L C" unfolding lowner_le_def proof (auto simp add: carrier_matD[OF A] carrier_matD[OF C]) have dim: "C - A \ carrier_mat n n" using A C by auto { fix v assume v: "(v::complex vec) \ carrier_vec n" from L1 have "inner_prod v (A *\<^sub>v v) \ inner_prod v (B *\<^sub>v v)" using lowner_le_inner_prod_le A B v by auto also from L2 have "\ \ inner_prod v (C *\<^sub>v v)" using lowner_le_inner_prod_le B C v by auto finally have "inner_prod v (A *\<^sub>v v) \ inner_prod v (C *\<^sub>v v)". then have "inner_prod v (C *\<^sub>v v) - inner_prod v (A *\<^sub>v v) \ 0" by auto then have "inner_prod v ((C - A) *\<^sub>v v) \ 0" using A C v apply (subst minus_add_uminus_mat[OF C A]) apply (subst add_mult_distrib_mat_vec[OF C _ v], simp) apply (simp add: inner_prod_distrib_right[OF v]) done } note leq = this show "positive (C - A)" unfolding positive_def apply (rule, simp add: carrier_matD[OF A] dim) apply (subst carrier_matD[OF dim], insert leq, auto) done qed lemma lowner_le_imp_trace_le: assumes "A \ carrier_mat n n" and "B \ carrier_mat n n" and "A \\<^sub>L B" shows "trace A \ trace B" proof - have "positive (B - A)" using assms lowner_le_def by auto moreover have "B - A \ carrier_mat n n" using assms by auto ultimately have "trace (B - A) \ 0" using positive_trace by auto moreover have "trace (B - A) = trace B - trace A" using trace_minus_linear assms by auto ultimately have "trace B - trace A \ 0" by auto then show "trace A \ trace B" by auto qed lemma lowner_le_add: assumes "A \ carrier_mat n n" "B \ carrier_mat n n" "C \ carrier_mat n n" "D \ carrier_mat n n" and "A \\<^sub>L B" "C \\<^sub>L D" shows "A + C \\<^sub>L B + D" proof - have "B + D - (A + C) = B - A + (D - C) " using assms by auto then have "positive (B + D - (A + C))" using assms unfolding lowner_le_def using positive_add by (metis minus_carrier_mat) then show "A + C \\<^sub>L B + D" unfolding lowner_le_def using assms by fastforce qed lemma lowner_le_swap: assumes "A \ carrier_mat n n" "B \ carrier_mat n n" and "A \\<^sub>L B" shows "-B \\<^sub>L -A" proof - have "positive (B - A)" using assms lowner_le_def by fastforce moreover have "B - A = (-A) - (-B)" using assms by fastforce ultimately have "positive ((-A) - (-B))" by auto then show ?thesis using lowner_le_def assms by fastforce qed lemma lowner_le_minus: assumes "A \ carrier_mat n n" "B \ carrier_mat n n" "C \ carrier_mat n n" "D \ carrier_mat n n" and "A \\<^sub>L B" "C \\<^sub>L D" shows "A - D \\<^sub>L B - C" proof - have "positive (D - C)" using assms lowner_le_def by auto then have "-D \\<^sub>L -C" using lowner_le_swap assms by auto then have "A + (-D) \\<^sub>L B + (-C)" using lowner_le_add[of "A" n "B"] assms by auto moreover have "A + (-D) = A - D" and "B + (-C) = B - C" by auto ultimately show ?thesis by auto qed lemma outer_prod_le_one: assumes "v \ carrier_vec n" and "inner_prod v v \ 1" shows "outer_prod v v \\<^sub>L 1\<^sub>m n" proof - let ?o = "outer_prod v v" have do: "?o \ carrier_mat n n" using assms by auto { fix u :: "complex vec" assume "dim_vec u = n" then have du: "u \ carrier_vec n" by auto have r: "inner_prod u u \ Reals" apply (simp add: scalar_prod_def carrier_vecD[OF du]) using complex_In_mult_cnj_zero complex_is_Real_iff by blast have geq0: "inner_prod u u \ 0" using self_cscalar_prod_geq_0 by auto have "inner_prod u (?o *\<^sub>v u) = inner_prod u v * inner_prod v u" apply (subst inner_prod_outer_prod) using du assms by auto also have "\ \ inner_prod u u * inner_prod v v" using Cauchy_Schwarz_complex_vec du assms by auto also have "\ \ inner_prod u u" using assms(2) r geq0 by (simp add: mult_right_le_one_le less_eq_complex_def) finally have le: "inner_prod u (?o *\<^sub>v u) \ inner_prod u u". have "inner_prod u ((1\<^sub>m n - ?o) *\<^sub>v u) = inner_prod u ((1\<^sub>m n *\<^sub>v u) - ?o *\<^sub>v u)" apply (subst minus_mult_distrib_mat_vec) using do du by auto also have "\ = inner_prod u u - inner_prod u (?o *\<^sub>v u)" apply (subst inner_prod_minus_distrib_right) using du do by auto also have "\ \ 0" using le by auto finally have "inner_prod u ((1\<^sub>m n - ?o) *\<^sub>v u) \ 0" by auto } then have "positive (1\<^sub>m n - outer_prod v v)" unfolding positive_def using do by auto then show ?thesis unfolding lowner_le_def using do by auto qed lemma zero_lowner_le_positiveD: fixes A :: "complex mat" assumes dA: "A \ carrier_mat n n" and le: "0\<^sub>m n n \\<^sub>L A" shows "positive A" using assms unfolding lowner_le_def by (subgoal_tac "A - 0\<^sub>m n n = A", auto) lemma zero_lowner_le_positiveI: fixes A :: "complex mat" assumes dA: "A \ carrier_mat n n" and le: "positive A" shows "0\<^sub>m n n \\<^sub>L A" using assms unfolding lowner_le_def by (subgoal_tac "A - 0\<^sub>m n n = A", auto) lemma lowner_le_trans_positiveI: fixes A B :: "complex mat" assumes dA: "A \ carrier_mat n n" and pA: "positive A" and le: "A \\<^sub>L B" shows "positive B" proof - have dB: "B \ carrier_mat n n" using le dA lowner_le_def by auto have "0\<^sub>m n n \\<^sub>L A" using zero_lowner_le_positiveI dA pA by auto then have "0\<^sub>m n n \\<^sub>L B" using dA dB le by (simp add: lowner_le_trans[of _ n A B]) then show ?thesis using dB zero_lowner_le_positiveD by auto qed lemma lowner_le_keep_under_measurement: fixes M A B :: "complex mat" assumes dM: "M \ carrier_mat n n" and dA: "A \ carrier_mat n n" and dB: "B \ carrier_mat n n" and le: "A \\<^sub>L B" shows "adjoint M * A * M \\<^sub>L adjoint M * B * M" unfolding lowner_le_def proof (rule conjI, fastforce)+ have daM: "adjoint M \ carrier_mat n n" using dM by auto have dBmA: "B - A \ carrier_mat n n" using dB dA by fastforce have "positive (B - A)" using le lowner_le_def by auto then have p: "positive (adjoint M * (B - A) * M)" using positive_close_under_left_right_mult_adjoint[OF daM dBmA] adjoint_adjoint[of M] by auto moreover have e: "adjoint M * (B - A) * M = adjoint M * B * M - adjoint M * A * M" using dM dB dA by (mat_assoc n) ultimately show "positive (adjoint M * B * M - adjoint M * A * M)" by auto qed lemma smult_distrib_left_minus_mat: fixes A B :: "'a::comm_ring_1 mat" assumes "A \ carrier_mat n n" "B \ carrier_mat n n" shows "c \\<^sub>m (B - A) = c \\<^sub>m B - c \\<^sub>m A" using assms by (auto simp add: minus_add_uminus_mat add_smult_distrib_left_mat) lemma lowner_le_smultc: fixes c :: complex assumes "c \ 0" "A \\<^sub>L B" "A \ carrier_mat n n" "B \ carrier_mat n n" shows "c \\<^sub>m A \\<^sub>L c \\<^sub>m B" proof - have eqBA: "c \\<^sub>m (B - A) = c \\<^sub>m B - c \\<^sub>m A" using assms by (auto simp add: smult_distrib_left_minus_mat) have "positive (B - A)" using assms(2) unfolding lowner_le_def by auto then have "positive (c \\<^sub>m (B - A))" using positive_smult[of "B-A" n c] assms by fastforce moreover have "c \\<^sub>m A \ carrier_mat n n" using index_smult_mat(2,3) assms(3) by auto moreover have "c \\<^sub>m B \ carrier_mat n n" using index_smult_mat(2,3) assms(4) by auto ultimately show ?thesis unfolding lowner_le_def using eqBA by fastforce qed lemma lowner_le_smult: fixes c :: real assumes "c \ 0" "A \\<^sub>L B" "A \ carrier_mat n n" "B \ carrier_mat n n" shows "c \\<^sub>m A \\<^sub>L c \\<^sub>m B" apply (rule lowner_le_smultc) using assms by (auto simp: less_eq_complex_def) lemma minus_smult_vec_distrib: fixes w :: "'a::comm_ring_1 vec" shows "(a - b) \\<^sub>v w = a \\<^sub>v w - b \\<^sub>v w" apply (rule eq_vecI) by (auto simp add: scalar_prod_def algebra_simps) lemma smult_mat_mult_mat_vec_assoc: fixes A :: "'a::comm_ring_1 mat" assumes A: "A \ carrier_mat n m" and w: "w \ carrier_vec m" shows "a \\<^sub>m A *\<^sub>v w = a \\<^sub>v (A *\<^sub>v w)" apply (rule eq_vecI) apply (simp add: scalar_prod_def carrier_matD[OF A] carrier_vecD[OF w]) apply (subst sum_distrib_left) apply (rule sum.cong, simp) by auto lemma mult_mat_vec_smult_vec_assoc: fixes A :: "'a::comm_ring_1 mat" assumes A: "A \ carrier_mat n m" and w: "w \ carrier_vec m" shows "A *\<^sub>v (a \\<^sub>v w) = a \\<^sub>v (A *\<^sub>v w)" apply (rule eq_vecI) apply (simp add: scalar_prod_def carrier_matD[OF A] carrier_vecD[OF w]) apply (subst sum_distrib_left) apply (rule sum.cong, simp) by auto lemma outer_prod_left_right_mat: fixes A B :: "complex mat" assumes du: "u \ carrier_vec d2" and dv: "v \ carrier_vec d3" and dA: "A \ carrier_mat d1 d2" and dB: "B \ carrier_mat d3 d4" shows "A * (outer_prod u v) * B = (outer_prod (A *\<^sub>v u) (adjoint B *\<^sub>v v))" unfolding outer_prod_def proof - have eq1: "A * (mat (dim_vec u) 1 (\(i, j). u $ i)) = mat (dim_vec (A *\<^sub>v u)) 1 (\(i, j). (A *\<^sub>v u) $ i)" apply (rule eq_matI) by (auto simp add: dA du scalar_prod_def) have conj: "conjugate a * b = conjugate ((a::complex) * conjugate b) " for a b by auto have eq2: "mat 1 (dim_vec v) (\(i, y). conjugate v $ y) * B = mat 1 (dim_vec (adjoint B *\<^sub>v v)) (\(i, y). conjugate (adjoint B *\<^sub>v v) $ y)" apply (rule eq_matI) apply (auto simp add: carrier_matD[OF dB] carrier_vecD[OF dv] scalar_prod_def adjoint_def conjugate_vec_def sum_conjugate ) apply (rule sum.cong) by (auto simp add: conj) have "A * (mat (dim_vec u) 1 (\(i, j). u $ i) * mat 1 (dim_vec v) (\(i, y). conjugate v $ y)) * B = (A * (mat (dim_vec u) 1 (\(i, j). u $ i))) *(mat 1 (dim_vec v) (\(i, y). conjugate v $ y)) * B" using dA du dv dB assoc_mult_mat[OF dA, of "mat (dim_vec u) 1 (\(i, j). u $ i)" 1 "mat 1 (dim_vec v) (\(i, y). conjugate v $ y)"] by fastforce also have "\ = (A * (mat (dim_vec u) 1 (\(i, j). u $ i))) *((mat 1 (dim_vec v) (\(i, y). conjugate v $ y)) * B)" using dA du dv dB assoc_mult_mat[OF _ _ dB, of "(A * (mat (dim_vec u) 1 (\(i, j). u $ i)))" d1 1] by fastforce finally show "A * (mat (dim_vec u) 1 (\(i, j). u $ i) * mat 1 (dim_vec v) (\(i, y). conjugate v $ y)) * B = mat (dim_vec (A *\<^sub>v u)) 1 (\(i, j). (A *\<^sub>v u) $ i) * mat 1 (dim_vec (adjoint B *\<^sub>v v)) (\(i, y). conjugate (adjoint B *\<^sub>v v) $ y)" using eq1 eq2 by auto qed subsection \Density operators\ definition density_operator :: "complex mat \ bool" where "density_operator A \ positive A \ trace A = 1" definition partial_density_operator :: "complex mat \ bool" where "partial_density_operator A \ positive A \ trace A \ 1" lemma pure_state_self_outer_prod_is_partial_density_operator: fixes v :: "complex vec" assumes dimv: "v \ carrier_vec n" and nv: "vec_norm v = 1" shows "partial_density_operator (outer_prod v v)" unfolding partial_density_operator_def proof have dimov: "outer_prod v v \ carrier_mat n n" using dimv by auto show "positive (outer_prod v v)" unfolding positive_def proof (rule, simp add: carrier_matD(2)[OF dimov] dimov, rule allI, rule impI) fix w assume "dim_vec (w::complex vec) = dim_col (outer_prod v v)" then have dimw: "w \ carrier_vec n" using dimov carrier_vecI by auto then have "inner_prod w ((outer_prod v v) *\<^sub>v w) = inner_prod w v * inner_prod v w" using inner_prod_outer_prod dimw dimv by auto also have "\ = inner_prod w v * conjugate (inner_prod w v)" using dimw dimv apply (subst conjugate_scalar_prod[of v "conjugate w"], simp) apply (subst conjugate_vec_sprod_comm[of "conjugate v" _ "conjugate w"], auto) apply (rule carrier_vec_conjugate[OF dimv]) apply (rule carrier_vec_conjugate[OF dimw]) done also have "\ \ 0" by (auto simp: less_eq_complex_def) finally show "inner_prod w ((outer_prod v v) *\<^sub>v w) \ 0". qed have eq: "trace (outer_prod v v) = (\i=0..i=0..i=0.. 1" by auto qed (* Lemma 2.1 *) lemma lowner_le_trace: assumes A: "A \ carrier_mat n n" and B: "B \ carrier_mat n n" shows "A \\<^sub>L B \ (\\\carrier_mat n n. partial_density_operator \ \ trace (A * \) \ trace (B * \))" proof (rule iffI) have dimBmA: "B - A \ carrier_mat n n" using A B by auto { assume "A \\<^sub>L B" then have pBmA: "positive (B - A)" using lowner_le_def by auto moreover have "B - A \ carrier_mat n n" using assms by auto ultimately have "\M\carrier_mat n n. M * adjoint M = B - A" using positive_iff_decomp[of "B - A"] by auto then obtain M where dimM: "M \ carrier_mat n n" and M: "M * adjoint M = B - A" by auto { fix \ assume dimr: "\ \ carrier_mat n n" and pdr: "partial_density_operator \" have eq: "trace(B * \) - trace(A * \) = trace((B - A) * \)" using A B dimr apply (subst minus_mult_distrib_mat, auto) apply (subst trace_minus_linear, auto) done have pr: "positive \" using pdr partial_density_operator_def by auto then have "\P\carrier_mat n n. \ = P * adjoint P" using positive_iff_decomp dimr by auto then obtain P where dimP: "P \ carrier_mat n n" and P: "\ = P * adjoint P" by auto have "trace((B - A) * \) = trace(M * adjoint M * (P * adjoint P))" using P M by auto also have "\ = trace((adjoint P * M) * adjoint (adjoint P * M))" using dimM dimP by (mat_assoc n) also have "\ \ 0" using trace_adjoint_positive by auto finally have "trace((B - A) * \) \ 0". with eq have " trace (B * \) - trace (A * \) \ 0" by auto } then show "\\\carrier_mat n n. partial_density_operator \ \ trace (A * \) \ trace (B * \)" by auto } { assume asm: "\\\carrier_mat n n. partial_density_operator \ \ trace (A * \) \ trace (B * \)" have "positive (B - A)" proof - { fix v assume "dim_vec (v::complex vec) = dim_col (B - A) \ vec_norm v = 1" then have dimv: "v \ carrier_vec n" and nv: "vec_norm v = 1" using carrier_matD[OF dimBmA] by (auto intro: carrier_vecI) have dimov: "outer_prod v v \ carrier_mat n n" using dimv by auto then have "partial_density_operator (outer_prod v v)" using dimv nv pure_state_self_outer_prod_is_partial_density_operator by auto then have leq: "trace(A * (outer_prod v v)) \ trace(B * (outer_prod v v))" using asm dimov by auto have "trace((B - A) * (outer_prod v v)) = trace(B * (outer_prod v v)) - trace(A * (outer_prod v v))" using A B dimov apply (subst minus_mult_distrib_mat, auto) apply (subst trace_minus_linear, auto) done then have "trace((B - A) * (outer_prod v v)) \ 0" using leq by auto then have "inner_prod v ((B - A) *\<^sub>v v) \ 0" using trace_outer_prod_right[OF dimBmA dimv dimv] by auto } then show "positive (B - A)" using positive_iff_normalized_vec[of "B - A"] dimBmA A by simp qed then show "A \\<^sub>L B" using lowner_le_def A B by auto } qed lemma lowner_le_traceI: assumes "A \ carrier_mat n n" and "B \ carrier_mat n n" and "\\. \ \ carrier_mat n n \ partial_density_operator \ \ trace (A * \) \ trace (B * \)" shows "A \\<^sub>L B" using lowner_le_trace assms by auto lemma trace_pdo_eq_imp_eq: assumes A: "A \ carrier_mat n n" and B: "B \ carrier_mat n n" and teq: "\\. \ \ carrier_mat n n \ partial_density_operator \ \ trace (A * \) = trace (B * \)" shows "A = B" proof - from teq have "A \\<^sub>L B" using lowner_le_trace[OF A B] teq by auto moreover from teq have "B \\<^sub>L A" using lowner_le_trace[OF B A] teq by auto ultimately show "A = B" using lowner_le_antisym A B by auto qed lemma lowner_le_traceD: assumes "A \ carrier_mat n n" "B \ carrier_mat n n" "\ \ carrier_mat n n" and "A \\<^sub>L B" and "partial_density_operator \" shows "trace (A * \) \ trace (B * \)" using lowner_le_trace assms by blast lemma sum_only_one_neq_0: assumes "finite A" and "j \ A" and "\i. i \ A \ i \ j \ g i = 0" shows "sum g A = g j" proof - have "{j} \ A" using assms by auto moreover have "\i\A - {j}. g i = 0" using assms by simp ultimately have "sum g A = sum g {j}" using assms by (auto simp add: comm_monoid_add_class.sum.mono_neutral_right[of A "{j}" g]) moreover have "sum g {j} = g j" by simp ultimately show ?thesis by auto qed end diff --git a/thys/VectorSpace/VectorSpace.thy b/thys/VectorSpace/VectorSpace.thy --- a/thys/VectorSpace/VectorSpace.thy +++ b/thys/VectorSpace/VectorSpace.thy @@ -1,1210 +1,1210 @@ section \Basic theory of vector spaces, using locales\ theory VectorSpace imports Main "HOL-Algebra.Module" "HOL-Algebra.Coset" RingModuleFacts MonoidSums LinearCombinations SumSpaces begin subsection \Basic definitions and facts carried over from modules\ text \A \vectorspace\ is a module where the ring is a field. Note that we switch notation from $(R, M)$ to $(K, V)$.\ locale vectorspace = module?: module K V + field?: field K for K and V (* Use sets for bases, and functions from the sets to carrier K represent the coefficients. *) text \A \subspace\ of a vectorspace is a nonempty subset that is closed under addition and scalar multiplication. These properties have already been defined in submodule. Caution: W is a set, while V is a module record. To get W as a vectorspace, write vs W.\ locale subspace = fixes K and W and V (structure) assumes vs: "vectorspace K V" and submod: "submodule K W V" lemma (in vectorspace) is_module[simp]: "subspace K W V\submodule K W V" by (unfold subspace_def, auto) text \We introduce some basic facts and definitions copied from module. We introduce some abbreviations, to match convention.\ abbreviation (in vectorspace) vs::"'c set \ ('a, 'c, 'd) module_scheme" where "vs W \ V\carrier :=W\" lemma (in vectorspace) carrier_vs_is_self [simp]: "carrier (vs W) = W" by auto lemma (in vectorspace) subspace_is_vs: fixes W::"'c set" assumes 0: "subspace K W V" shows "vectorspace K (vs W)" proof - from 0 show ?thesis apply (unfold vectorspace_def subspace_def, auto) by (intro submodule_is_module, auto) qed abbreviation (in module) subspace_sum:: "['c set, 'c set] \ 'c set" where "subspace_sum W1 W2 \submodule_sum W1 W2" lemma (in vectorspace) vs_zero_lin_dep: assumes h2: "S\carrier V" and h3: "lin_indpt S" shows "\\<^bsub>V\<^esub> \ S" proof - have vs: "vectorspace K V".. from vs have nonzero: "carrier K \{\\<^bsub>K\<^esub>}" by (metis one_zeroI zero_not_one) from h2 h3 nonzero show ?thesis by (rule zero_nin_lin_indpt) qed text \A \linear_map\ is a module homomorphism between 2 vectorspaces over the same field.\ locale linear_map = V?: vectorspace K V + W?: vectorspace K W + mod_hom?: mod_hom K V W T for K and V and W and T context linear_map begin lemmas T_hom = f_hom lemmas T_add = f_add lemmas T_smult = f_smult lemmas T_im = f_im lemmas T_neg = f_neg lemmas T_minus = f_minus lemmas T_ker = f_ker abbreviation imT:: "'e set" where "imT \ mod_hom.im" abbreviation kerT:: "'c set" where "kerT \ mod_hom.ker" lemmas T0_is_0[simp] = f0_is_0 lemma kerT_is_subspace: "subspace K ker V" proof - have vs: "vectorspace K V".. from vs show ?thesis apply (unfold subspace_def, auto) by (rule ker_is_submodule) qed lemma imT_is_subspace: "subspace K imT W" proof - have vs: "vectorspace K W".. from vs show ?thesis apply (unfold subspace_def, auto) by (rule im_is_submodule) qed end lemma vs_criteria: fixes K and V assumes field: "field K" and zero: "\\<^bsub>V\<^esub>\ carrier V" and add: "\v w. v\carrier V \ w\carrier V\ v\\<^bsub>V\<^esub> w\ carrier V" and neg: "\v\carrier V. (\neg_v\carrier V. v\\<^bsub>V\<^esub>neg_v=\\<^bsub>V\<^esub>)" and smult: "\c v. c\ carrier K \ v\carrier V\ c\\<^bsub>V\<^esub> v \ carrier V" and comm: "\v w. v\carrier V \ w\carrier V\ v\\<^bsub>V\<^esub> w=w\\<^bsub>V\<^esub> v" and assoc: "\v w x. v\carrier V \ w\carrier V \ x\carrier V\ (v\\<^bsub>V\<^esub> w)\\<^bsub>V\<^esub> x= v\\<^bsub>V\<^esub> (w\\<^bsub>V\<^esub> x)" and add_id: "\v\carrier V. (v\\<^bsub>V\<^esub> \\<^bsub>V\<^esub> =v)" and compat: "\a b v. a\ carrier K \ b\ carrier K \ v\carrier V\ (a\\<^bsub>K\<^esub> b)\\<^bsub>V\<^esub> v =a\\<^bsub>V\<^esub> (b\\<^bsub>V\<^esub> v)" and smult_id: "\v\carrier V. (\\<^bsub>K\<^esub> \\<^bsub>V\<^esub> v =v)" and dist_f: "\a b v. a\ carrier K \ b\ carrier K \ v\carrier V\ (a\\<^bsub>K\<^esub> b)\\<^bsub>V\<^esub> v =(a\\<^bsub>V\<^esub> v) \\<^bsub>V\<^esub> (b\\<^bsub>V\<^esub> v)" and dist_add: "\a v w. a\ carrier K \ v\carrier V \ w\carrier V\ a\\<^bsub>V\<^esub> (v\\<^bsub>V\<^esub> w) =(a\\<^bsub>V\<^esub> v) \\<^bsub>V\<^esub> (a\\<^bsub>V\<^esub> w)" shows "vectorspace K V" proof - from field have 1: "cring K" by (unfold field_def domain_def, auto) from assms 1 have 2: "module K V" by (intro module_criteria, auto) from field 2 show ?thesis by (unfold vectorspace_def module_def, auto) qed text \For any set $S$, the space of functions $S\to K$ forms a vector space.\ lemma (in vectorspace) func_space_is_vs: fixes S shows "vectorspace K (func_space S)" proof - have 0: "field K".. have 1: "module K (func_space S)" by (rule func_space_is_module) from 0 1 show ?thesis by (unfold vectorspace_def module_def, auto) qed lemma direct_sum_is_vs: fixes K V1 V2 assumes h1: "vectorspace K V1" and h2: "vectorspace K V2" shows "vectorspace K (direct_sum V1 V2)" proof - from h1 h2 have mod: "module K (direct_sum V1 V2)" by (unfold vectorspace_def, intro direct_sum_is_module, auto) from mod h1 show ?thesis by (unfold vectorspace_def, auto) qed lemma inj1_linear: fixes K V1 V2 assumes h1: "vectorspace K V1" and h2: "vectorspace K V2" shows "linear_map K V1 (direct_sum V1 V2) (inj1 V1 V2)" proof - from h1 h2 have mod: "mod_hom K V1 (direct_sum V1 V2) (inj1 V1 V2)" by (unfold vectorspace_def, intro inj1_hom, auto) from mod h1 h2 show ?thesis by (unfold linear_map_def vectorspace_def , auto, intro direct_sum_is_module, auto) qed lemma inj2_linear: fixes K V1 V2 assumes h1: "vectorspace K V1" and h2: "vectorspace K V2" shows "linear_map K V2 (direct_sum V1 V2) (inj2 V1 V2)" proof - from h1 h2 have mod: "mod_hom K V2 (direct_sum V1 V2) (inj2 V1 V2)" by (unfold vectorspace_def, intro inj2_hom, auto) from mod h1 h2 show ?thesis by (unfold linear_map_def vectorspace_def , auto, intro direct_sum_is_module, auto) qed text \For subspaces $V_1,V_2\subseteq V$, the map $V_1\oplus V_2\to V$ given by $(v_1,v_2)\mapsto v_1+v_2$ is linear.\ lemma (in vectorspace) sum_map_linear: fixes V1 V2 assumes h1: "subspace K V1 V" and h2: "subspace K V2 V" shows "linear_map K (direct_sum (vs V1) (vs V2)) V (\ v. (fst v) \\<^bsub>V\<^esub> (snd v))" proof - from h1 h2 have mod: "mod_hom K (direct_sum (vs V1) (vs V2)) V (\ v. (fst v) \\<^bsub>V\<^esub> (snd v))" by ( intro sum_map_hom, unfold subspace_def, auto) from mod h1 h2 show ?thesis apply (unfold linear_map_def, auto) apply (intro direct_sum_is_vs subspace_is_vs, auto).. qed lemma (in vectorspace) sum_is_subspace: fixes W1 W2 assumes h1: "subspace K W1 V" and h2: "subspace K W2 V" shows "subspace K (subspace_sum W1 W2) V" proof - from h1 h2 have mod: "submodule K (submodule_sum W1 W2) V" by ( intro sum_is_submodule, unfold subspace_def, auto) from mod h1 h2 show ?thesis by (unfold subspace_def, auto) qed text \If $W_1,W_2\subseteq V$ are subspaces, $W_1\subseteq W_1 + W_2$\ lemma (in vectorspace) in_sum_vs: fixes W1 W2 assumes h1: "subspace K W1 V" and h2: "subspace K W2 V" shows "W1 \ subspace_sum W1 W2" proof - from h1 h2 show ?thesis by (intro in_sum, unfold subspace_def, auto) qed lemma (in vectorspace) vsum_comm: fixes W1 W2 assumes h1: "subspace K W1 V" and h2: "subspace K W2 V" shows "(subspace_sum W1 W2) = (subspace_sum W2 W1)" proof - from h1 h2 show ?thesis by (intro msum_comm, unfold subspace_def, auto) qed text \If $W_1,W_2\subseteq V$ are subspaces, then $W_1+W_2$ is the minimal subspace such that both $W_1\subseteq W$ and $W_2\subseteq W$.\ lemma (in vectorspace) vsum_is_minimal: fixes W W1 W2 assumes h1: "subspace K W1 V" and h2: "subspace K W2 V" and h3: "subspace K W V" shows "(subspace_sum W1 W2) \ W \ W1 \ W \ W2 \ W" proof - from h1 h2 h3 show ?thesis by (intro sum_is_minimal, unfold subspace_def, auto) qed lemma (in vectorspace) span_is_subspace: fixes S assumes h2: "S\carrier V" shows "subspace K (span S) V" proof - have 0: "vectorspace K V".. from h2 have 1: "submodule K (span S) V" by (rule span_is_submodule) from 0 1 show ?thesis by (unfold subspace_def mod_hom_def linear_map_def, auto) qed subsubsection \Facts specific to vector spaces\ text \If $av = w$ and $a\neq 0$, $v=a^{-1}w$.\ lemma (in vectorspace) mult_inverse: assumes h1: "a\carrier K" and h2: "v\carrier V" and h3: "a \\<^bsub>V\<^esub> v = w" and h4: "a\\\<^bsub>K\<^esub>" shows "v = (inv\<^bsub>K\<^esub> a )\\<^bsub>V\<^esub> w" proof - from h1 h2 h3 have 1: "w\carrier V" by auto from h3 1 have 2: "(inv\<^bsub>K\<^esub> a )\\<^bsub>V\<^esub>(a \\<^bsub>V\<^esub> v) =(inv\<^bsub>K\<^esub> a )\\<^bsub>V\<^esub>w" by auto from h1 h4 have 3: "inv\<^bsub>K\<^esub> a\carrier K" by auto interpret g: group "(units_group K)" by (rule units_form_group) have f: "field K".. from f h1 h4 have 4: "a\Units K" by (unfold field_def field_axioms_def, simp) from 4 h1 h4 have 5: "((inv\<^bsub>K\<^esub> a) \\<^bsub>K\<^esub>a) = \\<^bsub>K\<^esub>" by (intro Units_l_inv, auto) from 5 have 6: "(inv\<^bsub>K\<^esub> a )\\<^bsub>V\<^esub>(a \\<^bsub>V\<^esub> v) = v" proof - from h1 h2 h4 have 7: "(inv\<^bsub>K\<^esub> a )\\<^bsub>V\<^esub>(a \\<^bsub>V\<^esub> v) =(inv\<^bsub>K\<^esub> a \\<^bsub>K\<^esub>a) \\<^bsub>V\<^esub> v" by (auto simp add: smult_assoc1) from 5 h2 have 8: "(inv\<^bsub>K\<^esub> a \\<^bsub>K\<^esub>a) \\<^bsub>V\<^esub> v = v" by auto from 7 8 show ?thesis by auto qed from 2 6 show ?thesis by auto qed text \If $w\in S$ and $\sum_{w\in S} a_ww=0$, we have $v=\sum_{w\not\in S}a_v^{-1}a_ww$\ lemma (in vectorspace) lincomb_isolate: fixes A v assumes h1: "finite A" and h2: "A\carrier V" and h3: "a\A\carrier K" and h4: "v\A" and h5: "a v \ \\<^bsub>K\<^esub>" and h6: "lincomb a A=\\<^bsub>V\<^esub>" shows "v=lincomb (\w. \\<^bsub>K\<^esub>(inv\<^bsub>K\<^esub> (a v)) \\<^bsub>K\<^esub> a w) (A-{v})" and "v\ span (A-{v})" proof - from h1 h2 h3 h4 have 1: "lincomb a A = ((a v) \\<^bsub>V\<^esub> v) \\<^bsub>V\<^esub> lincomb a (A-{v})" by (rule lincomb_del2) from 1 have 2: "\\<^bsub>V\<^esub>= ((a v) \\<^bsub>V\<^esub> v) \\<^bsub>V\<^esub> lincomb a (A-{v})" by (simp add: h6) from h1 h2 h3 have 5: "lincomb a (A-{v}) \carrier V" by auto (*intro lincomb_closed*) from 2 h1 h2 h3 h4 have 3: " \\<^bsub>V\<^esub> lincomb a (A-{v}) = ((a v) \\<^bsub>V\<^esub> v)" by (auto intro!: M.minus_equality) have 6: "v = (\\<^bsub>K\<^esub> (inv\<^bsub>K\<^esub> (a v))) \\<^bsub>V\<^esub> lincomb a (A-{v})" proof - from h2 h3 h4 h5 3 have 7: "v = inv\<^bsub>K\<^esub> (a v) \\<^bsub>V\<^esub> (\\<^bsub>V\<^esub> lincomb a (A-{v}))" by (intro mult_inverse, auto) from assms have 8: "inv\<^bsub>K\<^esub> (a v)\carrier K" by auto from assms 5 8 have 9: "inv\<^bsub>K\<^esub> (a v) \\<^bsub>V\<^esub> (\\<^bsub>V\<^esub> lincomb a (A-{v})) = (\\<^bsub>K\<^esub> (inv\<^bsub>K\<^esub> (a v))) \\<^bsub>V\<^esub> lincomb a (A-{v})" by (simp add: smult_assoc_simp smult_minus_1_back r_minus) from 7 9 show ?thesis by auto qed from h1 have 10: "finite (A-{v})" by auto from assms have 11 : "(\\<^bsub>K\<^esub> (inv\<^bsub>K\<^esub> (a v)))\ carrier K" by auto from assms have 12: "lincomb (\w. \\<^bsub>K\<^esub>(inv\<^bsub>K\<^esub> (a v)) \\<^bsub>K\<^esub> a w) (A-{v}) = (\\<^bsub>K\<^esub> (inv\<^bsub>K\<^esub> (a v))) \\<^bsub>V\<^esub> lincomb a (A-{v})" by (intro lincomb_smult, auto) from 6 12 show "v=lincomb (\w. \\<^bsub>K\<^esub>(inv\<^bsub>K\<^esub> (a v)) \\<^bsub>K\<^esub> a w) (A-{v})" by auto with assms show "v\ span (A-{v})" unfolding span_def by (force simp add: 11 ring_subset_carrier) qed text \The map $(S\to K)\mapsto V$ given by $(a_v)_{v\in S}\mapsto \sum_{v\in S} a_v v$ is linear.\ lemma (in vectorspace) lincomb_is_linear: fixes S assumes h: "finite S" and h2: "S\carrier V" shows "linear_map K (func_space S) V (\a. lincomb a S)" proof - have 0: "vectorspace K V".. from h h2 have 1: "mod_hom K (func_space S) V (\a. lincomb a S)" by (rule lincomb_is_mod_hom) from 0 1 show ?thesis by (unfold vectorspace_def mod_hom_def linear_map_def, auto) qed subsection \Basic facts about span and linear independence\ text \If $S$ is linearly independent, then $v\in \text{span}S$ iff $S\cup \{v\}$ is linearly dependent.\ theorem (in vectorspace) lin_dep_iff_in_span: fixes A v S assumes h1: "S \ carrier V" and h2: "lin_indpt S" and h3: "v\ carrier V" and h4: "v\S" shows "v\ span S \ lin_dep (S \ {v})" proof - let ?T = "S \ {v}" have 0: "v\?T " by auto from h1 h3 have h1_1: "?T \ carrier V" by auto have a1:"lin_dep ?T \ v\ span S" proof - assume a11: "lin_dep ?T" from a11 obtain a w A where a: "(finite A \ A\?T \ (a\ (A\carrier K)) \ (lincomb a A = \\<^bsub>V\<^esub>) \ (w\A) \ (a w\ \\<^bsub>K\<^esub>))" by (metis lin_dep_def) from assms a have nz2: "\v\A-S. a v\\\<^bsub>K\<^esub>" by (intro lincomb_must_include[where ?v="w" and ?T="S\{v}"], auto) from a nz2 have singleton: "{v}=A-S" by auto from singleton nz2 have nz3: "a v\\\<^bsub>K\<^esub>" by auto (*Can modularize this whole section out. "solving for one variable"*) let ?b="(\w. \\<^bsub>K\<^esub> (inv\<^bsub>K\<^esub> (a v)) \\<^bsub>K\<^esub> (a w))" from singleton have Ains: "(A\S) = A-{v}" by auto from assms a singleton nz3 have a31: "v= lincomb ?b (A\S)" apply (subst Ains) by (intro lincomb_isolate(1), auto) from a a31 nz3 singleton show ?thesis apply (unfold span_def, auto) apply (rule_tac x="?b" in exI) apply (rule_tac x="A\S" in exI) by (auto intro!: m_closed) qed have a2: "v\ (span S) \ lin_dep ?T" proof - assume inspan: "v\ (span S)" from inspan obtain a A where a: "A\S \ finite A \ (v = lincomb a A)\ a\A\carrier K" by (simp add: span_def, auto) let ?b = "\ w. if (w=v) then (\\<^bsub>K\<^esub> \\<^bsub>K\<^esub>) else a w" (*consider -v + \sum a_w w*) have lc0: " lincomb ?b (A\{v})=\\<^bsub>V\<^esub>" proof - from assms a have lc_ins: "lincomb ?b (A\{v}) = ((?b v) \\<^bsub>V\<^esub> v) \\<^bsub>V\<^esub> lincomb ?b A" by (intro lincomb_insert, auto) from assms a have lc_elim: "lincomb ?b A=lincomb a A" by (intro lincomb_elim_if, auto) from assms lc_ins lc_elim a show ?thesis by (simp add: M.l_neg smult_minus_1) qed from a lc0 show ?thesis apply (unfold lin_dep_def) apply (rule_tac x="A\{v}" in exI) apply (rule_tac x="?b" in exI) apply (rule_tac x="v" in exI) by auto qed from a1 a2 show ?thesis by auto qed text \If $v\in \text{span} A$ then $\text{span}A =\text{span}(A\cup \{v\})$\ lemma (in vectorspace) already_in_span: fixes v A assumes inC: "A\carrier V" and inspan: "v\span A" shows "span A= span (A\{v})" proof - from inC inspan have dir1: "span A \ span (A\{v})" by (intro span_is_monotone, auto) from inC have inown: "A\span A" by (rule in_own_span) from inC have subm: "submodule K (span A) V" by (rule span_is_submodule) from inown inspan subm have dir2: "span (A \ {v}) \ span A" by (intro span_is_subset, auto) from dir1 dir2 show ?thesis by auto qed subsection \The Replacement Theorem\ text \If $A,B\subseteq V$ are finite, $A$ is linearly independent, $B$ generates $W$, and $A\subseteq W$, then there exists $C\subseteq V$ disjoint from $A$ such that $\text{span}(A\cup C) = W$ and $|C|\le |B|-|A|$. In other words, we can complete any linearly independent set to a generating set of $W$ by adding at most $|B|-|A|$ more elements.\ theorem (in vectorspace) replacement: fixes A B (*A B are lists of vectors (colloquially we refer to them as sets)*) assumes h1: "finite A" and h2: "finite B" and h3: "B\carrier V" and h4: "lin_indpt A" (*A is linearly independent*) and h5: "A\span B" (*All entries of A are in K*) shows "\C. finite C \ C\carrier V \ C\span B \ C\A={} \ int (card C) \ (int (card B)) - (int (card A)) \ (span (A \ C) = span B)" (is "\C. ?P A B C") (*There is a set C of cardinality |B| - |A| such that A\C generates V*) using h1 h2 h3 h4 h5 proof (induct "card A" arbitrary: A B) case 0 from "0.prems"(1) "0.hyps" have a0: "A={}" by auto from "0.prems"(3) have a3: "B\span B" by (rule in_own_span) from a0 a3 "0.prems" show ?case by (rule_tac x="B" in exI, auto) next case (Suc m) let ?W="span B" from Suc.prems(3) have BinC: "span B\carrier V" by (rule span_is_subset2) (*everything you want to know about A*) from Suc.prems Suc.hyps BinC have A: "finite A" "lin_indpt A" "A\span B" "Suc m = card A" "A\carrier V" by auto (*everything you want to know about B*) from Suc.prems have B: "finite B" "B\carrier V" by auto (*A B are lists of vectors (colloquially we refer to them as sets)*) from Suc.hyps(2) obtain v where v: "v\A" by fastforce let ?A'="A-{v}" (*?A' is linearly independent because it is the subset of a linearly independent set, A.*) from A(2) have liA': "lin_indpt ?A'" apply (intro subset_li_is_li[of "A" "?A'"]) by auto from v liA' Suc.prems Suc.hyps(2) have "\C'. ?P ?A' B C'" apply (intro Suc.hyps(1)) by auto from this obtain C' where C': "?P ?A' B C'" by auto show ?case proof (cases "v\ C'") case True have vinC': "v\C'" by fact from vinC' v have seteq: "A - {v} \ C' = A \ (C' - {v})" by auto from C' seteq have spaneq: "span (A \ (C' - {v})) = span (B)" by algebra from Suc.prems Suc.hyps C' vinC' v spaneq show ?thesis apply (rule_tac x="C'-{v}" in exI) apply (subgoal_tac "card C' >0") by auto next case False have f: "v\C'" by fact from A v C' have "\a. a\(?A'\C')\carrier K \ lincomb a (?A' \ C') =v" by (intro finite_in_span, auto) from this obtain a where a: "a\(?A'\C')\carrier K \ v= lincomb a (?A' \ C')" by metis let ?b="(\ w. if (w=v) then \\<^bsub>K\<^esub>\\<^bsub>K\<^esub> else a w)" from a have b: "?b\A\C'\carrier K" by auto from v have rewrite_ins: "A\C'=(?A'\C')\{v}" by auto from f have "v\?A'\C'" by auto from this A C' v a f have lcb: "lincomb ?b (A \ C') = \\<^bsub>V\<^esub>" apply (subst rewrite_ins) apply (subst lincomb_insert) apply (simp_all add: ring_subset_carrier coeff_in_ring) apply (auto split: if_split_asm) apply (subst lincomb_elim_if) by (auto simp add: smult_minus_1 l_neg ring_subset_carrier) (*NOTE: l_neg got deleted from the simp rules, but it is very useful.*) from C' f have rewrite_minus: "C'=(A\C')-A" by auto from A C' b lcb v have exw: "\w\ C'. ?b w\\\<^bsub>K\<^esub>" apply (subst rewrite_minus) apply (intro lincomb_must_include[where ?T="A\C'" and ?v="v"]) by auto from exw obtain w where w: "w\ C'" "?b w\\\<^bsub>K\<^esub>" by auto from A C' w f b lcb have w_in: "w\span ((A\ C') -{w})" apply (intro lincomb_isolate[where a="?b"]) by auto have spaneq2: "span (A\(C'-{w})) = span B" proof - have 1: "span (?A'\C') = span (A\ C')" (*adding v doesn't change the span*) proof - from A C' v have m1: "span (?A'\C') = span ((?A'\ C')\{v})" apply (intro already_in_span) by auto from f m1 show ?thesis by (metis rewrite_ins) qed have 2: "span (A\ (C'-{w})) = span (A\ C')" (*removing w doesn't change the span*) proof - from C' w(1) f have b60: "A\ (C'-{w}) = (A\ C') -{w}" by auto from w(1) have b61: "A\ C'= (A\ C' -{w})\{w}" by auto from A C' w_in show ?thesis apply (subst b61) apply (subst b60) apply (intro already_in_span) by auto qed from C' 1 2 show ?thesis by auto qed(* "span (A\(C'-{w})) = span B"*) from A C' w f v spaneq2 show ?thesis apply (rule_tac x="C'-{w}" in exI) apply (subgoal_tac "card C' >0") by auto qed qed subsection \Defining dimension and bases.\ text \Finite dimensional is defined as having a finite generating set.\ definition (in vectorspace) fin_dim:: "bool" where "fin_dim = (\ A. ((finite A) \ (A \ carrier V) \ (gen_set A)))" text \The dimension is the size of the smallest generating set. For equivalent characterizations see below.\ definition (in vectorspace) dim:: "nat" where "dim = (LEAST n. (\ A. ((finite A) \ (card A = n) \ (A \ carrier V) \ (gen_set A))))" text \A \basis\ is a linearly independent generating set.\ definition (in vectorspace) basis:: "'c set \ bool" where "basis A = ((lin_indpt A) \ (gen_set A)\ (A\carrier V))" text \From the replacement theorem, any linearly independent set is smaller than any generating set.\ lemma (in vectorspace) li_smaller_than_gen: fixes A B assumes h1: "finite A" and h2: "finite B" and h3: "A\carrier V" and h4: "B\carrier V" and h5: "lin_indpt A" and h6: "gen_set B" shows "card A \ card B" proof - from h3 h6 have 1: "A\span B" by auto from h1 h2 h4 h5 1 obtain C where 2: "finite C \ C\carrier V \ C\span B \ C\A={} \ int (card C) \ int (card B) - int (card A) \ (span (A \ C) = span B)" by (metis replacement) from 2 show ?thesis by arith qed text \The dimension is the cardinality of any basis. (In particular, all bases are the same size.)\ lemma (in vectorspace) dim_basis: fixes A assumes fin: "finite A" and h2: "basis A" shows "dim = card A" proof - have 0: "\B m. ((finite B) \ (card B = m) \ (B \ carrier V) \ (gen_set B)) \ card A \ m" proof - fix B m assume 1: "((finite B) \ (card B = m) \ (B \ carrier V) \ (gen_set B))" from 1 fin h2 have 2: "card A \ card B" apply (unfold basis_def) apply (intro li_smaller_than_gen) by auto from 1 2 show "?thesis B m" by auto qed from fin h2 0 show ?thesis apply (unfold dim_def basis_def) apply (intro Least_equality) apply (rule_tac x="A" in exI) by auto qed (*can define more generally with posets*) text \A \maximal\ set with respect to $P$ is such that if $B\supseteq A$ and $P$ is also satisfied for $B$, then $B=A$.\ definition maximal::"'a set \ ('a set \ bool) \ bool" where "maximal A P = ((P A) \ (\B. B\A \ P B \ B = A))" text \A \minimal\ set with respect to $P$ is such that if $B\subseteq A$ and $P$ is also satisfied for $B$, then $B=A$.\ definition minimal::"'a set \ ('a set \ bool) \ bool" where "minimal A P = ((P A) \ (\B. B\A \ P B \ B = A))" text \A maximal linearly independent set is a generating set.\ lemma (in vectorspace) max_li_is_gen: fixes A assumes h1: "maximal A (\S. S\carrier V \ lin_indpt S)" shows "gen_set A" proof (rule ccontr) assume 0: "\(gen_set A)" from h1 have 1: " A\ carrier V \ lin_indpt A" by (unfold maximal_def, auto) from 1 have 2: "span A \ carrier V" by (intro span_is_subset2, auto) from 0 1 2 have 3: "\v. v\carrier V \ v \ (span A)" by auto from 3 obtain v where 4: "v\carrier V \ v \ (span A)" by auto have 5: "v\A" proof - from h1 1 have 51: "A\span A" apply (intro in_own_span) by auto from 4 51 show ?thesis by auto qed from lin_dep_iff_in_span have 6: "\S v. S \ carrier V\ lin_indpt S \ v\ carrier V \ v\S \ v\ span S \ (lin_indpt (S \ {v}))" by auto from 1 4 5 have 7: "lin_indpt (A \ {v})" apply (intro 6) by auto (* assumes h0:"finite S" and h1: "S \ carrier V" and h2: "lin_indpt S" and h3: "v\ carrier V" and h4: "v\S" shows "v\ span S \ \ (lin_indpt (S \ {v}))"*) have 9: "\(maximal A (\S. S\carrier V \ lin_indpt S))" proof - from 1 4 5 7 have 8: "(\B. A \ B \ B \ carrier V \ lin_indpt B \ B \ A)" apply (rule_tac x="A\{v}" in exI) by auto from 8 show ?thesis apply (unfold maximal_def) by simp qed from h1 9 show False by auto qed text \A minimal generating set is linearly independent.\ lemma (in vectorspace) min_gen_is_li: fixes A assumes h1: "minimal A (\S. S\carrier V \ gen_set S)" shows "lin_indpt A" proof (rule ccontr) assume 0: "\lin_indpt A" from h1 have 1: " A\ carrier V \ gen_set A" by (unfold minimal_def, auto) from 1 have 2: "span A = carrier V" by auto from 0 1 obtain a v A' where 3: "finite A' \ A'\A \ a \ A' \ carrier K \ LinearCombinations.module.lincomb V a A' = \\<^bsub>V\<^esub> \ v \ A' \ a v \ \\<^bsub>K\<^esub>" by (unfold lin_dep_def, auto) have 4: "gen_set (A-{v})" proof - from 1 3 have 5: "v\span (A'-{v})" apply (intro lincomb_isolate[where a="a" and v="v"]) by auto from 3 5 have 51: "v\span (A-{v})" apply (intro subsetD[where ?A="span (A'-{v})" and ?B="span (A-{v})" and ?c="v"]) by (intro span_is_monotone, auto) from 1 have 6: "A\span A" apply (intro in_own_span) by auto from 1 51 have 7: "span (A-{v}) = span ((A-{v})\{v})" apply (intro already_in_span) by auto from 3 have 8: "A = ((A-{v})\{v})" by auto from 2 7 8 have 9:"span (A-{v}) = carrier V" by auto (*can't use 3 directly :( *) from 9 show ?thesis by auto qed have 10: "\(minimal A (\S. S\carrier V \ gen_set S))" proof - from 1 3 4 have 11: "(\B. A \ B \ B \ carrier V \ gen_set B \ B \ A)" apply (rule_tac x="A-{v}" in exI) by auto from 11 show ?thesis apply (unfold minimal_def) by auto qed from h1 10 show False by auto qed text \Given that some finite set satisfies $P$, there is a minimal set that satisfies $P$.\ lemma minimal_exists: fixes A P assumes h1: "finite A" and h2: "P A" shows "\B. B\A \ minimal B P" using h1 h2 proof (induct "card A" arbitrary: A rule: less_induct) case (less A) show ?case proof (cases "card A = 0") case True from True less.hyps less.prems show ?thesis apply (rule_tac x="{}" in exI) apply (unfold minimal_def) by auto next case False show ?thesis proof (cases "minimal A P") case True then show ?thesis apply (rule_tac x="A" in exI) by auto next case False have 2: "\minimal A P" by fact from less.prems 2 have 3: "\B. P B \ B \ A \ B\A" apply (unfold minimal_def) by auto from 3 obtain B where 4: "P B \ B \ A \ B\A" by auto from 4 have 5: "card B < card A" by (metis less.prems(1) psubset_card_mono) from less.hyps less.prems 3 4 5 have 6: "\C\B. minimal C P" apply (intro less.hyps) apply auto by (metis rev_finite_subset) from 6 obtain C where 7: "C\B \ minimal C P" by auto from 4 7 show ?thesis apply (rule_tac x="C" in exI) apply (unfold minimal_def) by auto qed qed qed text \If $V$ is finite-dimensional, then any linearly independent set is finite.\ lemma (in vectorspace) fin_dim_li_fin: assumes fd: "fin_dim" and li: "lin_indpt A" and inC: "A\carrier V" shows fin: "finite A" proof (rule ccontr) assume A: "\finite A" from fd obtain C where C: "finite C \ C\carrier V \ gen_set C" by (unfold fin_dim_def, auto) from A obtain B where B: "B\A\ finite B \ card B = card C + 1" by (metis infinite_arbitrarily_large) from B li have liB: "lin_indpt B" by (intro subset_li_is_li[where ?A="A" and ?B="B"], auto) from B C liB inC have "card B \ card C" by (intro li_smaller_than_gen, auto) from this B show False by auto qed text \If $V$ is finite-dimensional (has a finite generating set), then a finite basis exists.\ lemma (in vectorspace) finite_basis_exists: assumes h1: "fin_dim" shows "\\. finite \ \ basis \" proof - from h1 obtain A where 1: "finite A \ A\carrier V \ gen_set A" by (metis fin_dim_def) hence 2: "\\. \\A \ minimal \ (\S. S\carrier V \ gen_set S)" apply (intro minimal_exists) by auto then obtain \ where 3: "\\A \ minimal \ (\S. S\carrier V \ gen_set S)" by auto hence 4: "lin_indpt \" apply (intro min_gen_is_li) by auto moreover from 3 have 5: "gen_set \ \ \\carrier V" apply (unfold minimal_def) by auto moreover from 1 3 have 6: "finite \" by (auto simp add: finite_subset) ultimately show ?thesis apply (unfold basis_def) by auto qed text\ The proof is as follows. \begin{enumerate} \item Because $V$ is finite-dimensional, there is a finite generating set (we took this as our definition of finite-dimensional). \item Hence, there is a minimal $\beta \subseteq A$ such that $\beta$ generates $V$. \item $\beta$ is linearly independent because a minimal generating set is linearly independent. \end{enumerate} Finally, $\beta$ is a basis because it is both generating and linearly independent. \ text \Any linearly independent set has cardinality at most equal to the dimension.\ lemma (in vectorspace) li_le_dim: fixes A assumes fd: "fin_dim" and c: "A\carrier V" and l: "lin_indpt A" shows "finite A" "card A \ dim" proof - from fd c l show fa: "finite A" by (intro fin_dim_li_fin, auto) from fd obtain \ where 1: "finite \ \ basis \" by (metis finite_basis_exists) from assms fa 1 have 2: "card A \ card \" apply (intro li_smaller_than_gen, auto) by (unfold basis_def, auto) from assms 1 have 3: "dim = card \" by (intro dim_basis, auto) from 2 3 show "card A \ dim" by auto qed text \Any generating set has cardinality at least equal to the dimension.\ lemma (in vectorspace) gen_ge_dim: fixes A assumes fa: "finite A" and c: "A\carrier V" and l: "gen_set A" shows "card A \ dim" proof - from assms have fd: "fin_dim" by (unfold fin_dim_def, auto) from fd obtain \ where 1: "finite \ \ basis \" by (metis finite_basis_exists) from assms 1 have 2: "card A \ card \" apply (intro li_smaller_than_gen, auto) by (unfold basis_def, auto) from assms 1 have 3: "dim = card \" by (intro dim_basis, auto) from 2 3 show ?thesis by auto qed text \If there is an upper bound on the cardinality of sets satisfying $P$, then there is a maximal set satisfying $P$.\ lemma maximal_exists: fixes P B N assumes maxc: "\A. P A \ finite A \ (card A \N)" and b: "P B" shows "\A. finite A \ maximal A P" proof - (*take the maximal*) let ?S="{card A| A. P A}" let ?n="Max ?S" from maxc have 1:"finite ?S" apply (simp add: finite_nat_set_iff_bounded_le) by auto from 1 have 2: "?n\?S" by (metis (mono_tags, lifting) Collect_empty_eq Max_in b) from assms 2 have 3: "\A. P A \ finite A \ card A = ?n" by auto from 3 obtain A where 4: "P A \ finite A \ card A = ?n" by auto from 1 maxc have 5: "\A. P A \ finite A \ (card A \?n)" by (metis (mono_tags, lifting) Max.coboundedI mem_Collect_eq) from 4 5 have 6: "maximal A P" apply (unfold maximal_def) by (metis card_seteq) from 4 6 show ?thesis by auto qed text \Any maximal linearly independent set is a basis.\ lemma (in vectorspace) max_li_is_basis: fixes A assumes h1: "maximal A (\S. S\carrier V \ lin_indpt S)" shows "basis A" proof - from h1 have 1: "gen_set A" by (rule max_li_is_gen) from assms 1 show ?thesis by (unfold basis_def maximal_def, auto) qed text \Any minimal linearly independent set is a generating set.\ lemma (in vectorspace) min_gen_is_basis: fixes A assumes h1: "minimal A (\S. S\carrier V \ gen_set S)" shows "basis A" proof - from h1 have 1: "lin_indpt A" by (rule min_gen_is_li) from assms 1 show ?thesis by (unfold basis_def minimal_def, auto) qed text \Any linearly independent set with cardinality at least the dimension is a basis.\ lemma (in vectorspace) dim_li_is_basis: fixes A assumes fd: "fin_dim" and fa: "finite A" and ca: "A\carrier V" and li: "lin_indpt A" and d: "card A \ dim" (*\*) shows "basis A" proof - from fd have 0: "\S. S\carrier V \ lin_indpt S \ finite S \ card S \dim" by (auto intro: li_le_dim) (*fin_dim_li_fin*) from 0 assms have h1: "finite A \ maximal A (\S. S\carrier V \ lin_indpt S)" apply (unfold maximal_def) apply auto by (metis card_seteq eq_iff) from h1 show ?thesis by (auto intro: max_li_is_basis) qed text \Any generating set with cardinality at most the dimension is a basis.\ lemma (in vectorspace) dim_gen_is_basis: fixes A assumes fa: "finite A" and ca: "A\carrier V" and li: "gen_set A" and d: "card A \ dim" shows "basis A" proof - have 0: "\S. finite S\ S\carrier V \ gen_set S \ card S \dim" by (intro gen_ge_dim, auto) (*li_le_dim*) from 0 assms have h1: "minimal A (\S. finite S \ S\carrier V \ gen_set S)" apply (unfold minimal_def) apply auto by (metis card_seteq eq_iff) (*slightly annoying: we have to get rid of "finite S" inside.*) from h1 have h: "\B. B \ A \ B \ carrier V \ LinearCombinations.module.gen_set K V B \ B = A" proof - fix B assume asm: "B \ A \ B \ carrier V \ LinearCombinations.module.gen_set K V B" from asm h1 have "finite B" apply (unfold minimal_def) apply (intro finite_subset[where ?A="B" and ?B="A"]) by auto from h1 asm this show "?thesis B" apply (unfold minimal_def) by simp qed from h1 h have h2: "minimal A (\S. S\carrier V \ gen_set S)" apply (unfold minimal_def) by presburger from h2 show ?thesis by (rule min_gen_is_basis) qed text \$\beta$ is a basis iff for all $v\in V$, there exists a unique $(a_v)_{v\in S}$ such that $\sum_{v\in S} a_v v=v$.\ lemma (in vectorspace) basis_criterion: assumes A_fin: "finite A" and AinC: "A\carrier V" shows "basis A \ (\v. v\ carrier V \(\! a. a\A \\<^sub>E carrier K \ lincomb a A = v))" proof - have 1: "\(\v. v\ carrier V \(\! a. a\A \\<^sub>E carrier K \ lincomb a A = v)) \ \basis A" proof - assume "\(\v. v\ carrier V \(\! a. a\A \\<^sub>E carrier K \ lincomb a A = v))" then obtain v where v: "v\ carrier V \ \(\! a. a\A \\<^sub>E carrier K \ lincomb a A = v)" by metis (*either there is more than 1 rep, or no reps*) from v have vinC: "v\carrier V" by auto from v have "\(\ a. a\A \\<^sub>E carrier K \ lincomb a A = v) \ (\ a b. a\A \\<^sub>E carrier K \ lincomb a A = v \ b\A \\<^sub>E carrier K \ lincomb b A = v \ a\b)" by metis then show ?thesis proof assume a: "\(\ a. a\A \\<^sub>E carrier K \ lincomb a A = v)" from A_fin AinC have "\a. a\A \ carrier K \ lincomb a A = lincomb (restrict a A) A" unfolding lincomb_def restrict_def by (simp cong: finsum_cong add: ring_subset_carrier coeff_in_ring) with a have "\(\ a. a\A \ carrier K \ lincomb a A = v)" by auto with A_fin AinC have "v\span A" using finite_in_span by blast with AinC v show "\(basis A)" by (unfold basis_def, auto) next assume a2: "(\ a b. a\A \\<^sub>E carrier K \ lincomb a A = v \ b\A \\<^sub>E carrier K \ lincomb b A = v \ a\b)" then obtain a b where ab: "a\A \\<^sub>E carrier K \ lincomb a A = v \ b\A \\<^sub>E carrier K \ lincomb b A = v \ a\b" by metis from ab obtain w where w: "w\A \ a w \ b w" apply (unfold PiE_def, auto) by (metis extensionalityI) let ?c="\ x. (if x\A then ((a x) \\<^bsub>K\<^esub> (b x)) else undefined)" from ab have a_fun: "a\A \ carrier K" and b_fun: "b\A \ carrier K" by (unfold PiE_def, auto) from w a_fun b_fun have abinC: "a w \carrier K" "b w \carrier K" by auto from abinC w have nz: "a w \\<^bsub>K\<^esub> b w \ \\<^bsub>K\<^esub>" by auto (*uses M.minus_other_side*) from A_fin AinC a_fun b_fun ab vinC have a_b: "LinearCombinations.module.lincomb V (\x. if x \ A then a x \\<^bsub>K\<^esub> b x else undefined) A = \\<^bsub>V\<^esub>" by (simp cong: lincomb_cong add: coeff_in_ring lincomb_diff) from A_fin AinC ab w v nz a_b have "lin_dep A" apply (intro lin_dep_crit[where ?A="A" and ?a="?c" and ?v="w"]) apply (auto simp add: PiE_def) by auto thus "\basis A" by (unfold basis_def, auto) qed qed have 2: "(\v. v\ carrier V \(\! a. a\A \\<^sub>E carrier K \ lincomb a A = v)) \ basis A" proof - assume b1: "(\v. v\ carrier V \(\! a. a\A \\<^sub>E carrier K \ lincomb a A = v))" (is "(\v. v\ carrier V \(\! a. ?Q a v))") from b1 have b2: "(\v. v\ carrier V \(\ a. a\A \ carrier K \ lincomb a A = v))" apply (unfold PiE_def) by blast from A_fin AinC b2 have "gen_set A" apply (unfold span_def) by blast from b1 have A_li: "lin_indpt A" proof - let ?z="\ x. (if (x\A) then \\<^bsub>K\<^esub> else undefined)" from A_fin AinC have zero: "?Q ?z \\<^bsub>V\<^esub>" by (unfold PiE_def extensional_def lincomb_def, auto simp add: ring_subset_carrier) (*uses finsum_all0*) from A_fin AinC show ?thesis proof (rule finite_lin_indpt2) fix a assume a_fun: "a \ A \ carrier K" and lc_a: "LinearCombinations.module.lincomb V a A = \\<^bsub>V\<^esub>" from a_fun have a_res: "restrict a A \ A \\<^sub>E carrier K" by auto from a_fun A_fin AinC lc_a have lc_a_res: "LinearCombinations.module.lincomb V (restrict a A) A = \\<^bsub>V\<^esub>" apply (unfold lincomb_def restrict_def) by (simp cong: finsum_cong2 add: coeff_in_ring ring_subset_carrier) from a_fun a_res lc_a_res zero b1 have "restrict a A = ?z" by auto from this show "\v\A. a v = \\<^bsub>K\<^esub>" apply (unfold restrict_def) by meson qed qed have A_gen: "gen_set A" proof - from AinC have dir1: "span A\carrier V" by (rule span_is_subset2) have dir2: "carrier V\span A" proof (auto) fix v assume v: "v\carrier V" from v b2 obtain a where "a\A \ carrier K \ lincomb a A = v" by auto from this A_fin AinC show "v\span A" by (subst finite_span, auto) qed from dir1 dir2 show ?thesis by auto qed from A_li A_gen AinC show "basis A" by (unfold basis_def, auto) qed from 1 2 show ?thesis by satx qed lemma (in linear_map) surj_imp_imT_carrier: assumes surj: "T` (carrier V) = carrier W" shows "(imT) = carrier W" by (simp add: surj im_def) subsection \The rank-nullity (dimension) theorem\ text \If $V$ is finite-dimensional and $T:V\to W$ is a linear map, then $\text{dim}(\text{im}(T))+ \text{dim}(\text{ker}(T)) = \text{dim} V$. Moreover, we prove that if $T$ is surjective linear-map between $V$ and $W$, where $V$ is finite-dimensional, then also $W$ is finite-dimensional.\ theorem (in linear_map) rank_nullity_main: assumes fd: "V.fin_dim" shows "(vectorspace.dim K (W.vs imT)) + (vectorspace.dim K (V.vs kerT)) = V.dim" "T ` (carrier V) = carrier W \ W.fin_dim" proof - \ \First interpret kerT, imT as vectorspaces\ have subs_ker: "subspace K kerT V" by (intro kerT_is_subspace) from subs_ker have vs_ker: "vectorspace K (V.vs kerT)" by (rule V.subspace_is_vs) from vs_ker interpret ker: vectorspace K "(V.vs kerT)" by auto have kerInC: "kerT\carrier V" by (unfold ker_def, auto) have subs_im: "subspace K imT W" by (intro imT_is_subspace) from subs_im have vs_im: "vectorspace K (W.vs imT)" by (rule W.subspace_is_vs) from vs_im interpret im: vectorspace K "(W.vs imT)" by auto have imInC: "imT\carrier W" by (unfold im_def, auto) (* obvious fact *) have zero_same[simp]: "\\<^bsub>V.vs kerT\<^esub> = \\<^bsub>V\<^esub>" apply (unfold ker_def) by auto \ \Show ker T has a finite basis. This is not obvious. Show that any linearly independent set has size at most that of V. There exists a maximal linearly independent set, which is the basis.\ have every_li_small: "\A. (A \ kerT)\ ker.lin_indpt A \ finite A \ card A \ V.dim" proof - fix A assume eli_asm: "(A \ kerT)\ ker.lin_indpt A" (*annoying: I can't just use subst V.module.span_li_not_depend(2) in the show ?thesis statement because it doesn't appear in the conclusion.*) note V.module.span_li_not_depend(2)[where ?N="kerT" and ?S="A"] from this subs_ker fd eli_asm kerInC show "?thesis A" apply (intro conjI) by (auto intro!: V.li_le_dim) qed from every_li_small have exA: "\A. finite A \ maximal A (\S. S\carrier (V.vs kerT) \ ker.lin_indpt S)" apply (intro maximal_exists[where ?N="V.dim" and ?B="{}"]) apply auto by (unfold ker.lin_dep_def, auto) from exA obtain A where A:" finite A \ maximal A (\S. S\carrier (V.vs kerT) \ ker.lin_indpt S)" by blast hence finA: "finite A" and Ainker: "A\carrier (V.vs kerT)" and AinC: "A\carrier V" by (unfold maximal_def ker_def, auto) \ \We obtain the basis A of kerT. It is also linearly independent when considered in V rather than kerT\ from A have Abasis: "ker.basis A" by (intro ker.max_li_is_basis, auto) from subs_ker Abasis have spanA: "V.module.span A = kerT" apply (unfold ker.basis_def) by (subst sym[OF V.module.span_li_not_depend(1)[where ?N="kerT"]], auto) from Abasis have Akerli: "ker.lin_indpt A" apply (unfold ker.basis_def) by auto from subs_ker Ainker Akerli have Ali: "V.module.lin_indpt A" by (auto simp add: V.module.span_li_not_depend(2)) txt\Use the replacement theorem to find C such that $A\cup C$ is a basis of V.\ from fd obtain B where B: "finite B\ V.basis B" by (metis V.finite_basis_exists) from B have Bfin: "finite B" and Bbasis:"V.basis B" by auto from B have Bcard: "V.dim = card B" by (intro V.dim_basis, auto) from Bbasis have 62: "V.module.span B = carrier V" by (unfold V.basis_def, auto) from A Abasis Ali B vs_ker have "\C. finite C \ C\carrier V \ C\ V.module.span B \ C\A={} \ int (card C) \ (int (card B)) - (int (card A)) \ (V.module.span (A \ C) = V.module.span B)" apply (intro V.replacement) apply (unfold vectorspace.basis_def V.basis_def) by (unfold ker_def, auto) txt \From replacement we got $|C|\leq |B|-|A|$. Equality must actually hold, because no generating set can be smaller than $B$. Now $A\cup C$ is a maximal generating set, hence a basis; its cardinality equals the dimension.\ txt \We claim that $T(C)$ is basis for $\text{im}(T)$.\ then obtain C where C: "finite C \ C\carrier V \ C\ V.module.span B \ C\A={} \ int (card C) \ (int (card B)) - (int (card A)) \ (V.module.span (A \ C) = V.module.span B)" by auto hence Cfin: "finite C" and CinC: "C\carrier V" and CinspanB: " C\V.module.span B" and CAdis: "C\A={}" and Ccard: "int (card C) \ (int (card B)) - (int (card A))" and ACspanB: "(V.module.span (A \ C) = V.module.span B)" by auto from C have cardLe: "card A + card C \ card B" by auto from B C have ACgen: "V.module.gen_set (A\C)" apply (unfold V.basis_def) by auto from finA C ACgen AinC B have cardGe: "card (A\C) \ card B" by (intro V.li_smaller_than_gen, unfold V.basis_def, auto) from finA C have cardUn: "card (A\C)\ card A + card C" by (metis Int_commute card_Un_disjoint le_refl) from cardLe cardUn cardGe Bcard have cardEq: "card (A\C) = card A + card C" "card (A\C) = card B" "card (A\C) = V.dim" by auto from Abasis C cardEq have disj: "A\C={}" by auto from finA AinC C cardEq 62 have ACfin: "finite (A\C)" and ACbasis: "V.basis (A\C)" by (auto intro!: V.dim_gen_is_basis) have lm: "linear_map K V W T".. txt \Let $C'$ be the image of $C$ under $T$. We will show $C'$ is a basis for $\text{im}(T)$.\ let ?C' = "T`C" from Cfin have C'fin: "finite ?C'" by auto from AinC C have cim: "?C'\imT" by (unfold im_def, auto) txt \"There is a subtle detail: we first have to show $T$ is injective on $C$.\ txt \We establish that no nontrivial linear combination of $C$ can have image 0 under $T$, because that would mean it is a linear combination of $A$, giving that $A\cup C$ is linearly dependent, contradiction. We use this result in 2 ways: (1) if $T$ is not injective on $C$, then we obtain $v$, $w\in C$ such that $v-w$ is in the kernel, contradiction, (2) if $T(C)$ is linearly dependent, taking the inverse image of that linear combination gives a linear combination of $C$ in the kernel, contradiction. Hence $T$ is injective on $C$ and $T(C)$ is linearly independent.\ have lc_in_ker: "\d D v. \D\C; d\D\carrier K; T (V.module.lincomb d D) = \\<^bsub>W\<^esub>; v\D; d v \\\<^bsub>K\<^esub>\\False" proof - fix d D v assume D: "D\C" and d: "d\D\carrier K" and T0: "T (V.module.lincomb d D) = \\<^bsub>W\<^esub>" and v: "v\D" and dvnz: "d v \\\<^bsub>K\<^esub>" from D Cfin have Dfin: "finite D" by (auto intro: finite_subset) from D CinC have DinC: "D\carrier V" by auto from T0 d Dfin DinC have lc_d: "V.module.lincomb d D\kerT" by (unfold ker_def, auto) from lc_d spanA AinC have "\a' A'. A'\A \ a'\A'\carrier K \ V.module.lincomb a' A'= V.module.lincomb d D" by (intro V.module.in_span, auto) then obtain a' A' where a': "A'\A \ a'\A'\carrier K \ V.module.lincomb d D = V.module.lincomb a' A'" by metis hence A'sub: "A'\A" and a'fun: "a'\A'\carrier K" and a'_lc:"V.module.lincomb d D = V.module.lincomb a' A'" by auto from a' finA Dfin have A'fin: "finite (A')" by (auto intro: finite_subset) from AinC A'sub have A'inC: "A'\carrier V" by auto let ?e = "(\v. if v \ A' then a' v else \\<^bsub>K\<^esub>\\<^bsub>K\<^esub>\\<^bsub>K\<^esub> d v)" from a'fun d have e_fun: "?e \ A' \ D \ carrier K" apply (unfold Pi_def) by auto from A'fin Dfin (*finiteness*) A'inC DinC (*in carrier*) a'fun d e_fun (*coefficients valid*) disj D A'sub (*A and C disjoint*) have lccomp1: "V.module.lincomb a' A' \\<^bsub>V\<^esub> \\<^bsub>K\<^esub>\\<^bsub>K\<^esub>\\<^bsub>V\<^esub> V.module.lincomb d D = V.module.lincomb (\v. if v\A' then a' v else \\<^bsub>K\<^esub>\\<^bsub>K\<^esub>\\<^bsub>K\<^esub> d v) (A'\D)" apply (subst sym[OF V.module.lincomb_smult]) apply (simp_all) apply (subst V.module.lincomb_union2) by (auto) from A'fin (*finiteness*) A'inC (*in carrier*) a'fun (*coefficients valid*) have lccomp2: "V.module.lincomb a' A' \\<^bsub>V\<^esub> \\<^bsub>K\<^esub>\\<^bsub>K\<^esub>\\<^bsub>V\<^esub> V.module.lincomb d D = \\<^bsub>V\<^esub>" by (simp add: a'_lc V.module.smult_minus_1 V.module.M.r_neg) from lccomp1 lccomp2 have lc0: "V.module.lincomb (\v. if v\A' then a' v else \\<^bsub>K\<^esub>\\<^bsub>K\<^esub>\\<^bsub>K\<^esub> d v) (A'\D) =\\<^bsub>V\<^esub>" by auto from disj a' v D have v_nin: "v\A'" by auto from A'fin Dfin (*finiteness*) A'inC DinC (*in carrier*) e_fun d (*coefficients valid*) A'sub D disj (*A' D are disjoint subsets*) v dvnz (*d v is nonzero coefficient*) lc0 have AC_ld: "V.module.lin_dep (A\C)" apply (intro V.module.lin_dep_crit[where ?A="A'\D" and ?S="A\C" and ?a="\v. if v\A' then a' v else \\<^bsub>K\<^esub>\\<^bsub>K\<^esub>\\<^bsub>K\<^esub> d v" and ?v="v"]) by (auto dest: integral) from AC_ld ACbasis show False by (unfold V.basis_def, auto) qed have C'_card: "inj_on T C" "card C = card ?C'" proof - show "inj_on T C" proof (rule ccontr) assume "\inj_on T C" then obtain v w where "v\C" "w\C" "v\w" "T v = T w" by (unfold inj_on_def, auto) from this CinC show False - apply (intro lc_in_ker[where ?D="{v,w}" and ?d="\x. if x=v then \\<^bsub>K\<^esub> else \\<^bsub>K\<^esub>\\<^bsub>K\<^esub>" - and ?v="v"]) + apply (intro lc_in_ker[where ?D1="{v,w}" and ?d1="\x. if x=v then \\<^bsub>K\<^esub> else \\<^bsub>K\<^esub>\\<^bsub>K\<^esub>" + and ?v1="v"]) by (auto simp add: V.module.lincomb_def hom_sum ring_subset_carrier W.module.smult_minus_1 r_neg T_im) qed from this Cfin show "card C = card ?C'" by (metis card_image) qed let ?f="the_inv_into C T" have f: "\x. x\C \ ?f (T x) = x" "\y. y\?C' \ T (?f y) = y" apply (insert C'_card(1)) apply (metis the_inv_into_f_f) by (metis f_the_inv_into_f) (*We show C' is a basis for the image. First we show it is linearly independent.*) have C'_li: "im.lin_indpt ?C'" proof (rule ccontr) assume Cld: "\im.lin_indpt ?C'" from Cld cim subs_im have CldW: "W.module.lin_dep ?C'" apply (subst sym[OF W.module.span_li_not_depend(2)[where ?S="T`C" and ?N="imT"]]) by auto from C CldW have "\c' v'. (c'\ (?C'\carrier K)) \ (W.module.lincomb c' ?C' = \\<^bsub>W\<^esub>) \ (v'\?C') \ (c' v'\ \\<^bsub>K\<^esub>)" by (intro W.module.finite_lin_dep, auto) then obtain c' v' where c': "(c'\ (?C'\carrier K)) \ (W.module.lincomb c' ?C' = \\<^bsub>W\<^esub>) \ (v'\?C') \ (c' v'\ \\<^bsub>K\<^esub>)" by auto hence c'fun: "(c'\ (?C'\carrier K))" and c'lc: "(W.module.lincomb c' ?C' = \\<^bsub>W\<^esub>)" and v':"(v'\?C')" and cvnz: "(c' v'\ \\<^bsub>K\<^esub>)" by auto (*can't get c' directly with metis/auto with W.module.finite_lin_dep*) txt \We take the inverse image of $C'$ under $T$ to get a linear combination of $C$ that is in the kernel and hence a linear combination of $A$. This contradicts $A\cup C$ being linearly independent.\ let ?c="\v. c' (T v)" from c'fun have c_fun: "?c\ C\carrier K" by auto from Cfin (*C finite*) c_fun c'fun (*coefficients valid*) C'_card (*bijective*) CinC (*C in carrier*) f (*inverse to T*) c'lc (*lc c' = 0*) have "T (V.module.lincomb ?c C) = \\<^bsub>W\<^esub>" apply (unfold V.module.lincomb_def W.module.lincomb_def) apply (subst hom_sum, auto) apply (simp cong: finsum_cong add: ring_subset_carrier coeff_in_ring) apply (subst finsum_reindex[where ?f="\w. c' w \\<^bsub>W\<^esub> w" and ?h="T" and ?A="C", THEN sym]) by auto with f c'fun cvnz v' show False - by (intro lc_in_ker[where ?D="C" and ?d="?c" and ?v="?f v'"], auto) + by (intro lc_in_ker[where ?D1="C" and ?d1="?c" and ?v1="?f v'"], auto) qed have C'_gen: "im.gen_set ?C'" proof - have C'_span: "span ?C' = imT" proof (rule equalityI) from cim subs_im show "W.module.span ?C' \ imT" by (intro span_is_subset, unfold subspace_def, auto) next show "imT\W.module.span ?C'" proof (auto) fix w assume w: "w\imT" from this finA Cfin AinC CinC obtain v where v_inC: "v\carrier V" and w_eq_T_v: "w= T v" by (unfold im_def image_def, auto) from finA Cfin AinC CinC v_inC ACgen have "\a. a \ A\C \ carrier K\ V.module.lincomb a (A\C) = v" by (intro V.module.finite_in_span, auto) then obtain a where a_fun: "a \ A\C \ carrier K" and lc_a_v: "v= V.module.lincomb a (A\C)" by auto let ?a'="\v. a (?f v)" from finA Cfin AinC CinC a_fun disj Ainker f C'_card have Tv: "T v = W.module.lincomb ?a' ?C'" apply (subst lc_a_v) apply (subst V.module.lincomb_union, simp_all) (*Break up the union A\C*) apply (unfold lincomb_def V.module.lincomb_def) apply (subst hom_sum, auto) (*Take T inside the sum over A*) apply (simp add: subsetD coeff_in_ring hom_sum (*Take T inside the sum over C*) T_ker (*all terms become 0 because the vectors are in the kernel.*) ) apply (subst finsum_reindex[where ?h="T" and ?f="\v. ?a' v\\<^bsub>W\<^esub> v"], auto) by (auto cong: finsum_cong simp add: coeff_in_ring ring_subset_carrier) from a_fun f have a'_fun: "?a'\?C' \ carrier K" by auto from C'fin CinC this w_eq_T_v a'_fun Tv show "w \ LinearCombinations.module.span K W (T ` C)" by (subst finite_span, auto) qed qed from this subs_im CinC show ?thesis apply (subst span_li_not_depend(1)) by (unfold im_def subspace_def, auto) qed from C'_li C'_gen C cim have C'_basis: "im.basis (T`C)" by (unfold im.basis_def, auto) have C_card_im: "card C = (vectorspace.dim K (W.vs imT))" using C'_basis C'_card(2) C'fin im.dim_basis by auto from finA Abasis have "ker.dim = card A" by (rule ker.dim_basis) note * = this C_card_im cardEq show "(vectorspace.dim K (W.vs imT)) + (vectorspace.dim K (V.vs kerT)) = V.dim" using * by auto assume "T` (carrier V) = carrier W" from * surj_imp_imT_carrier[OF this] show "W.fin_dim" using C'_basis C'fin unfolding W.fin_dim_def im.basis_def by auto qed theorem (in linear_map) rank_nullity: assumes fd: "V.fin_dim" shows "(vectorspace.dim K (W.vs imT)) + (vectorspace.dim K (V.vs kerT)) = V.dim" by (rule rank_nullity_main[OF fd]) end