theory Employee_DesignModel_UMLPart_generated_generated imports "OCL.UML_Main" "FOCL.Static" "FOCL.Generator_dynamic_sequential" begin (* 1 ************************************ 0 + 0 *) (* term Floor1_infra.print_infra_enum_synonym *) (* 2 ************************************ 0 + 1 *) text \ \label{ex:Employee-DesignModel-UMLPart-generated-generatedemployee-design:uml} \ (* 3 ************************************ 1 + 1 *) text \\ (* 4 ************************************ 2 + 1 *) section \Class Model: Introduction\ (* 5 ************************************ 3 + 1 *) text \ For certain concepts like classes and class-types, only a generic definition for its resulting semantics can be given. Generic means, there is a function outside \HOL that ``compiles'' a concrete, closed-world class diagram into a ``theory'' of this data model, consisting of a bunch of definitions for classes, accessors, method, casts, and tests for actual types, as well as proofs for the fundamental properties of these operations in this concrete data model. \ (* 6 ************************************ 4 + 1 *) text \ Such generic function or ``compiler'' can be implemented in Isabelle on the \ML level. This has been done, for a semantics following the open-world assumption, for \UML 2.0 in~\cite{brucker.ea:extensible:2008-b, brucker:interactive:2007}. In this paper, we follow another approach for \UML 2.4: we define the concepts of the compilation informally, and present a concrete example which is verified in Isabelle/\HOL. \ (* 7 ************************************ 5 + 1 *) subsection \Outlining the Example\ (* 8 ************************************ 6 + 1 *) text \ We are presenting here a ``design-model'' of the (slightly modified) example Figure 7.3, page 20 of the \OCL standard~\cite{omg:ocl:2012}. To be precise, this theory contains the formalization of the data-part covered by the \UML class model (see \autoref{fig:Employee-DesignModel-UMLPart-generated-generatedperson}):\ (* 9 ************************************ 7 + 1 *) text \\ (* 10 ************************************ 8 + 1 *) text_raw \ \begin{figure} \centering\scalebox{.3}{\includegraphics{figures/person.png}}% \caption{A simple \UML class model drawn from Figure 7.3, page 20 of~\cite{omg:ocl:2012}. \label{fig:Employee-DesignModel-UMLPart-generated-generatedperson}} \end{figure} \ (* 11 ************************************ 9 + 1 *) text_raw \\ (* 12 ************************************ 10 + 1 *) text \ This means that the association (attached to the association class \inlineocl{EmployeeRanking}) with the association ends \inlineocl+boss+ and \inlineocl+employees+ is implemented by the attribute \inlineocl+boss+ and the operation \inlineocl+employees+ (to be discussed in the \OCL part captured by the subsequent theory). \ (* 13 ************************************ 11 + 1 *) section \Class Model: The Construction of the Object Universe\ (* 14 ************************************ 12 + 1 *) text \ Our data universe consists in the concrete class diagram just of node's, and implicitly of the class object. Each class implies the existence of a class type defined for the corresponding object representations as follows: \ (* 15 ************************************ 13 + 8 *) (* term Floor1_infra.print_infra_datatype_class_1 *) datatype ty\\\\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n = mk\\\\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n "oid" "nat option" "int option" "unit option" "bool option" "oid list list option" datatype ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n = mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n "ty\\\\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" "oid list option" "int option" datatype ty\\\\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t = mk\\\\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n "ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" | mk\\\\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t "oid" "unit option" "bool option" "oid list list option" datatype ty\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t = mk\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t "ty\\\\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t" "nat option" "int option" datatype ty\\\\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y = mk\\\\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t "ty\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t" | mk\\\\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n "ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" | mk\\\\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y "oid" datatype ty\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y = mk\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y "ty\\\\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y" "unit option" "bool option" "oid list list option" datatype ty\\\\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y = mk\\\\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y "ty\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y" | mk\\\\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t "ty\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t" | mk\\\\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n "ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" | mk\\\\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y "oid" datatype ty\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y = mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y "ty\\\\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y" (* 16 ************************************ 21 + 11 *) (* term Floor1_infra.print_infra_datatype_class_2 *) datatype ty2\\\\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n = mk2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n "oid list option" "int option" datatype ty2oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n = mk2oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n "oid" "nat option" "int option" "unit option" "bool option" "oid list list option" "ty2\\\\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" datatype ty2\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t = mk2\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n "ty2\\\\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" datatype ty2\\\\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t = mk2\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t "nat option" "int option" "ty2\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t option" datatype ty2oid\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t = mk2oid\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t "oid" "unit option" "bool option" "oid list list option" "ty2\\\\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t" datatype ty2\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y = mk2\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t "ty2\\\\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t" datatype ty2\\\\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y = mk2\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y "unit option" "bool option" "oid list list option" "ty2\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y option" datatype ty2oid\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y = mk2oid\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y "oid" "ty2\\\\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y" datatype ty2\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y = mk2\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y "ty2\\\\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y" datatype ty2\\\\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y = mk2\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y "ty2\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y option" datatype ty2oid\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y = mk2oid\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y "oid" "ty2\\\\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y" (* 17 ************************************ 32 + 8 *) (* term Floor1_infra.print_infra_datatype_equiv_2of1 *) definition "class_ty_ext_equiv_2of1_aux\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n = (\oid inh\<^sub>w\<^sub>o\<^sub>r\<^sub>m\<^sub>h\<^sub>o\<^sub>l\<^sub>e inh\<^sub>w\<^sub>e\<^sub>i\<^sub>g\<^sub>h\<^sub>t inh\<^sub>s\<^sub>o\<^sub>u\<^sub>n\<^sub>d inh\<^sub>m\<^sub>o\<^sub>v\<^sub>i\<^sub>n\<^sub>g inh\<^sub>o\<^sub>u\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>w\<^sub>o\<^sub>r\<^sub>l\<^sub>d. (\ (mk2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (own\<^sub>b\<^sub>o\<^sub>s\<^sub>s) (own\<^sub>s\<^sub>a\<^sub>l\<^sub>a\<^sub>r\<^sub>y)) \ (mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n ((mk\\\\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (oid) (inh\<^sub>w\<^sub>o\<^sub>r\<^sub>m\<^sub>h\<^sub>o\<^sub>l\<^sub>e) (inh\<^sub>w\<^sub>e\<^sub>i\<^sub>g\<^sub>h\<^sub>t) (inh\<^sub>s\<^sub>o\<^sub>u\<^sub>n\<^sub>d) (inh\<^sub>m\<^sub>o\<^sub>v\<^sub>i\<^sub>n\<^sub>g) (inh\<^sub>o\<^sub>u\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>w\<^sub>o\<^sub>r\<^sub>l\<^sub>d))) (own\<^sub>b\<^sub>o\<^sub>s\<^sub>s) (own\<^sub>s\<^sub>a\<^sub>l\<^sub>a\<^sub>r\<^sub>y))))" definition "class_ty_ext_equiv_2of1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n = (\ (mk2oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (oid) (inh\<^sub>w\<^sub>o\<^sub>r\<^sub>m\<^sub>h\<^sub>o\<^sub>l\<^sub>e) (inh\<^sub>w\<^sub>e\<^sub>i\<^sub>g\<^sub>h\<^sub>t) (inh\<^sub>s\<^sub>o\<^sub>u\<^sub>n\<^sub>d) (inh\<^sub>m\<^sub>o\<^sub>v\<^sub>i\<^sub>n\<^sub>g) (inh\<^sub>o\<^sub>u\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>w\<^sub>o\<^sub>r\<^sub>l\<^sub>d) (t)) \ (class_ty_ext_equiv_2of1_aux\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (oid) (inh\<^sub>w\<^sub>o\<^sub>r\<^sub>m\<^sub>h\<^sub>o\<^sub>l\<^sub>e) (inh\<^sub>w\<^sub>e\<^sub>i\<^sub>g\<^sub>h\<^sub>t) (inh\<^sub>s\<^sub>o\<^sub>u\<^sub>n\<^sub>d) (inh\<^sub>m\<^sub>o\<^sub>v\<^sub>i\<^sub>n\<^sub>g) (inh\<^sub>o\<^sub>u\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>w\<^sub>o\<^sub>r\<^sub>l\<^sub>d) (t)))" definition "class_ty_ext_equiv_2of1_aux\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t = (\oid inh\<^sub>s\<^sub>o\<^sub>u\<^sub>n\<^sub>d inh\<^sub>m\<^sub>o\<^sub>v\<^sub>i\<^sub>n\<^sub>g inh\<^sub>o\<^sub>u\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>w\<^sub>o\<^sub>r\<^sub>l\<^sub>d. (\ (mk2\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (own\<^sub>w\<^sub>o\<^sub>r\<^sub>m\<^sub>h\<^sub>o\<^sub>l\<^sub>e) (own\<^sub>w\<^sub>e\<^sub>i\<^sub>g\<^sub>h\<^sub>t) (t)) \ (mk\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t ((case t of None \ (mk\\\\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (oid) (inh\<^sub>s\<^sub>o\<^sub>u\<^sub>n\<^sub>d) (inh\<^sub>m\<^sub>o\<^sub>v\<^sub>i\<^sub>n\<^sub>g) (inh\<^sub>o\<^sub>u\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>w\<^sub>o\<^sub>r\<^sub>l\<^sub>d)) | \(mk2\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (t))\ \ (case (class_ty_ext_equiv_2of1_aux\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (oid) (own\<^sub>w\<^sub>o\<^sub>r\<^sub>m\<^sub>h\<^sub>o\<^sub>l\<^sub>e) (own\<^sub>w\<^sub>e\<^sub>i\<^sub>g\<^sub>h\<^sub>t) (inh\<^sub>s\<^sub>o\<^sub>u\<^sub>n\<^sub>d) (inh\<^sub>m\<^sub>o\<^sub>v\<^sub>i\<^sub>n\<^sub>g) (inh\<^sub>o\<^sub>u\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>w\<^sub>o\<^sub>r\<^sub>l\<^sub>d) (t)) of (mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n ((mk\\\\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (oid) (own\<^sub>w\<^sub>o\<^sub>r\<^sub>m\<^sub>h\<^sub>o\<^sub>l\<^sub>e) (own\<^sub>w\<^sub>e\<^sub>i\<^sub>g\<^sub>h\<^sub>t) (inh\<^sub>s\<^sub>o\<^sub>u\<^sub>n\<^sub>d) (inh\<^sub>m\<^sub>o\<^sub>v\<^sub>i\<^sub>n\<^sub>g) (inh\<^sub>o\<^sub>u\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>w\<^sub>o\<^sub>r\<^sub>l\<^sub>d))) (own\<^sub>b\<^sub>o\<^sub>s\<^sub>s) (own\<^sub>s\<^sub>a\<^sub>l\<^sub>a\<^sub>r\<^sub>y)) \ (mk\\\\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n ((mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n ((mk\\\\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (oid) (own\<^sub>w\<^sub>o\<^sub>r\<^sub>m\<^sub>h\<^sub>o\<^sub>l\<^sub>e) (own\<^sub>w\<^sub>e\<^sub>i\<^sub>g\<^sub>h\<^sub>t) (inh\<^sub>s\<^sub>o\<^sub>u\<^sub>n\<^sub>d) (inh\<^sub>m\<^sub>o\<^sub>v\<^sub>i\<^sub>n\<^sub>g) (inh\<^sub>o\<^sub>u\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>w\<^sub>o\<^sub>r\<^sub>l\<^sub>d))) (own\<^sub>b\<^sub>o\<^sub>s\<^sub>s) (own\<^sub>s\<^sub>a\<^sub>l\<^sub>a\<^sub>r\<^sub>y))))))) (own\<^sub>w\<^sub>o\<^sub>r\<^sub>m\<^sub>h\<^sub>o\<^sub>l\<^sub>e) (own\<^sub>w\<^sub>e\<^sub>i\<^sub>g\<^sub>h\<^sub>t))))" definition "class_ty_ext_equiv_2of1\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t = (\ (mk2oid\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (oid) (inh\<^sub>s\<^sub>o\<^sub>u\<^sub>n\<^sub>d) (inh\<^sub>m\<^sub>o\<^sub>v\<^sub>i\<^sub>n\<^sub>g) (inh\<^sub>o\<^sub>u\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>w\<^sub>o\<^sub>r\<^sub>l\<^sub>d) (t)) \ (class_ty_ext_equiv_2of1_aux\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (oid) (inh\<^sub>s\<^sub>o\<^sub>u\<^sub>n\<^sub>d) (inh\<^sub>m\<^sub>o\<^sub>v\<^sub>i\<^sub>n\<^sub>g) (inh\<^sub>o\<^sub>u\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>w\<^sub>o\<^sub>r\<^sub>l\<^sub>d) (t)))" definition "class_ty_ext_equiv_2of1_aux\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y = (\oid. (\ (mk2\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y (own\<^sub>s\<^sub>o\<^sub>u\<^sub>n\<^sub>d) (own\<^sub>m\<^sub>o\<^sub>v\<^sub>i\<^sub>n\<^sub>g) (own\<^sub>o\<^sub>u\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>w\<^sub>o\<^sub>r\<^sub>l\<^sub>d) (t)) \ (mk\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y ((case t of None \ (mk\\\\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y (oid)) | \(mk2\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (t))\ \ (case (class_ty_ext_equiv_2of1_aux\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (oid) (own\<^sub>s\<^sub>o\<^sub>u\<^sub>n\<^sub>d) (own\<^sub>m\<^sub>o\<^sub>v\<^sub>i\<^sub>n\<^sub>g) (own\<^sub>o\<^sub>u\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>w\<^sub>o\<^sub>r\<^sub>l\<^sub>d) (t)) of (mk\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t ((mk\\\\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (oid) (own\<^sub>s\<^sub>o\<^sub>u\<^sub>n\<^sub>d) (own\<^sub>m\<^sub>o\<^sub>v\<^sub>i\<^sub>n\<^sub>g) (own\<^sub>o\<^sub>u\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>w\<^sub>o\<^sub>r\<^sub>l\<^sub>d))) (own\<^sub>w\<^sub>o\<^sub>r\<^sub>m\<^sub>h\<^sub>o\<^sub>l\<^sub>e) (own\<^sub>w\<^sub>e\<^sub>i\<^sub>g\<^sub>h\<^sub>t)) \ (mk\\\\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t ((mk\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t ((mk\\\\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (oid) (own\<^sub>s\<^sub>o\<^sub>u\<^sub>n\<^sub>d) (own\<^sub>m\<^sub>o\<^sub>v\<^sub>i\<^sub>n\<^sub>g) (own\<^sub>o\<^sub>u\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>w\<^sub>o\<^sub>r\<^sub>l\<^sub>d))) (own\<^sub>w\<^sub>o\<^sub>r\<^sub>m\<^sub>h\<^sub>o\<^sub>l\<^sub>e) (own\<^sub>w\<^sub>e\<^sub>i\<^sub>g\<^sub>h\<^sub>t)))) | (mk\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t ((mk\\\\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (t))) (own\<^sub>w\<^sub>o\<^sub>r\<^sub>m\<^sub>h\<^sub>o\<^sub>l\<^sub>e) (own\<^sub>w\<^sub>e\<^sub>i\<^sub>g\<^sub>h\<^sub>t)) \ (mk\\\\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (t))))) (own\<^sub>s\<^sub>o\<^sub>u\<^sub>n\<^sub>d) (own\<^sub>m\<^sub>o\<^sub>v\<^sub>i\<^sub>n\<^sub>g) (own\<^sub>o\<^sub>u\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>w\<^sub>o\<^sub>r\<^sub>l\<^sub>d))))" definition "class_ty_ext_equiv_2of1\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y = (\ (mk2oid\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y (oid) (t)) \ (class_ty_ext_equiv_2of1_aux\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y (oid) (t)))" definition "class_ty_ext_equiv_2of1_aux\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y = (\oid. (\ (mk2\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y (t)) \ (mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y ((case t of None \ (mk\\\\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y (oid)) | \(mk2\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y (t))\ \ (case (class_ty_ext_equiv_2of1_aux\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y (oid) (t)) of (mk\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y ((mk\\\\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y (oid))) (own\<^sub>s\<^sub>o\<^sub>u\<^sub>n\<^sub>d) (own\<^sub>m\<^sub>o\<^sub>v\<^sub>i\<^sub>n\<^sub>g) (own\<^sub>o\<^sub>u\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>w\<^sub>o\<^sub>r\<^sub>l\<^sub>d)) \ (mk\\\\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y ((mk\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y ((mk\\\\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y (oid))) (own\<^sub>s\<^sub>o\<^sub>u\<^sub>n\<^sub>d) (own\<^sub>m\<^sub>o\<^sub>v\<^sub>i\<^sub>n\<^sub>g) (own\<^sub>o\<^sub>u\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>w\<^sub>o\<^sub>r\<^sub>l\<^sub>d)))) | (mk\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y ((mk\\\\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (t))) (own\<^sub>s\<^sub>o\<^sub>u\<^sub>n\<^sub>d) (own\<^sub>m\<^sub>o\<^sub>v\<^sub>i\<^sub>n\<^sub>g) (own\<^sub>o\<^sub>u\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>w\<^sub>o\<^sub>r\<^sub>l\<^sub>d)) \ (mk\\\\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (t)) | (mk\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y ((mk\\\\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (t))) (own\<^sub>s\<^sub>o\<^sub>u\<^sub>n\<^sub>d) (own\<^sub>m\<^sub>o\<^sub>v\<^sub>i\<^sub>n\<^sub>g) (own\<^sub>o\<^sub>u\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>w\<^sub>o\<^sub>r\<^sub>l\<^sub>d)) \ (mk\\\\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (t))))))))" definition "class_ty_ext_equiv_2of1\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y = (\ (mk2oid\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y (oid) (t)) \ (class_ty_ext_equiv_2of1_aux\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y (oid) (t)))" (* 18 ************************************ 40 + 12 *) (* term Floor1_infra.print_infra_datatype_equiv_1of2 *) definition "class_ty_ext_equiv_1of2_get_oid_inh\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n = (\ (mk\\\\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (oid) (inh\<^sub>w\<^sub>o\<^sub>r\<^sub>m\<^sub>h\<^sub>o\<^sub>l\<^sub>e) (inh\<^sub>w\<^sub>e\<^sub>i\<^sub>g\<^sub>h\<^sub>t) (inh\<^sub>s\<^sub>o\<^sub>u\<^sub>n\<^sub>d) (inh\<^sub>m\<^sub>o\<^sub>v\<^sub>i\<^sub>n\<^sub>g) (inh\<^sub>o\<^sub>u\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>w\<^sub>o\<^sub>r\<^sub>l\<^sub>d)) \ (oid , inh\<^sub>w\<^sub>o\<^sub>r\<^sub>m\<^sub>h\<^sub>o\<^sub>l\<^sub>e , inh\<^sub>w\<^sub>e\<^sub>i\<^sub>g\<^sub>h\<^sub>t , inh\<^sub>s\<^sub>o\<^sub>u\<^sub>n\<^sub>d , inh\<^sub>m\<^sub>o\<^sub>v\<^sub>i\<^sub>n\<^sub>g , inh\<^sub>o\<^sub>u\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>w\<^sub>o\<^sub>r\<^sub>l\<^sub>d))" definition "class_ty_ext_equiv_1of2_aux\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n = (\ (mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (t) (own\<^sub>b\<^sub>o\<^sub>s\<^sub>s) (own\<^sub>s\<^sub>a\<^sub>l\<^sub>a\<^sub>r\<^sub>y)) \ (mk2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (own\<^sub>b\<^sub>o\<^sub>s\<^sub>s) (own\<^sub>s\<^sub>a\<^sub>l\<^sub>a\<^sub>r\<^sub>y)))" definition "class_ty_ext_equiv_1of2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n = (\ (mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (t) (own\<^sub>b\<^sub>o\<^sub>s\<^sub>s) (own\<^sub>s\<^sub>a\<^sub>l\<^sub>a\<^sub>r\<^sub>y)) \ (case (class_ty_ext_equiv_1of2_get_oid_inh\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (t)) of (oid , inh\<^sub>w\<^sub>o\<^sub>r\<^sub>m\<^sub>h\<^sub>o\<^sub>l\<^sub>e , inh\<^sub>w\<^sub>e\<^sub>i\<^sub>g\<^sub>h\<^sub>t , inh\<^sub>s\<^sub>o\<^sub>u\<^sub>n\<^sub>d , inh\<^sub>m\<^sub>o\<^sub>v\<^sub>i\<^sub>n\<^sub>g , inh\<^sub>o\<^sub>u\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>w\<^sub>o\<^sub>r\<^sub>l\<^sub>d) \ (mk2oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (oid) (inh\<^sub>w\<^sub>o\<^sub>r\<^sub>m\<^sub>h\<^sub>o\<^sub>l\<^sub>e) (inh\<^sub>w\<^sub>e\<^sub>i\<^sub>g\<^sub>h\<^sub>t) (inh\<^sub>s\<^sub>o\<^sub>u\<^sub>n\<^sub>d) (inh\<^sub>m\<^sub>o\<^sub>v\<^sub>i\<^sub>n\<^sub>g) (inh\<^sub>o\<^sub>u\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>w\<^sub>o\<^sub>r\<^sub>l\<^sub>d) ((class_ty_ext_equiv_1of2_aux\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n ((mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (t) (own\<^sub>b\<^sub>o\<^sub>s\<^sub>s) (own\<^sub>s\<^sub>a\<^sub>l\<^sub>a\<^sub>r\<^sub>y))))))))" definition "class_ty_ext_equiv_1of2_get_oid_inh\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t = (\ (mk\\\\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (oid) (inh\<^sub>s\<^sub>o\<^sub>u\<^sub>n\<^sub>d) (inh\<^sub>m\<^sub>o\<^sub>v\<^sub>i\<^sub>n\<^sub>g) (inh\<^sub>o\<^sub>u\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>w\<^sub>o\<^sub>r\<^sub>l\<^sub>d)) \ (oid , inh\<^sub>s\<^sub>o\<^sub>u\<^sub>n\<^sub>d , inh\<^sub>m\<^sub>o\<^sub>v\<^sub>i\<^sub>n\<^sub>g , inh\<^sub>o\<^sub>u\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>w\<^sub>o\<^sub>r\<^sub>l\<^sub>d) | (mk\\\\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n ((mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (t) (var\<^sub>b\<^sub>o\<^sub>s\<^sub>s) (var\<^sub>s\<^sub>a\<^sub>l\<^sub>a\<^sub>r\<^sub>y)))) \ (case (class_ty_ext_equiv_1of2_get_oid_inh\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (t)) of (oid , var\<^sub>w\<^sub>o\<^sub>r\<^sub>m\<^sub>h\<^sub>o\<^sub>l\<^sub>e , var\<^sub>w\<^sub>e\<^sub>i\<^sub>g\<^sub>h\<^sub>t , var\<^sub>s\<^sub>o\<^sub>u\<^sub>n\<^sub>d , var\<^sub>m\<^sub>o\<^sub>v\<^sub>i\<^sub>n\<^sub>g , var\<^sub>o\<^sub>u\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>w\<^sub>o\<^sub>r\<^sub>l\<^sub>d) \ (oid , var\<^sub>s\<^sub>o\<^sub>u\<^sub>n\<^sub>d , var\<^sub>m\<^sub>o\<^sub>v\<^sub>i\<^sub>n\<^sub>g , var\<^sub>o\<^sub>u\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>w\<^sub>o\<^sub>r\<^sub>l\<^sub>d)))" definition "class_ty_ext_equiv_1of2_aux\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t = (\ (mk\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (t) (own\<^sub>w\<^sub>o\<^sub>r\<^sub>m\<^sub>h\<^sub>o\<^sub>l\<^sub>e) (own\<^sub>w\<^sub>e\<^sub>i\<^sub>g\<^sub>h\<^sub>t)) \ (mk2\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (own\<^sub>w\<^sub>o\<^sub>r\<^sub>m\<^sub>h\<^sub>o\<^sub>l\<^sub>e) (own\<^sub>w\<^sub>e\<^sub>i\<^sub>g\<^sub>h\<^sub>t) ((case t of (mk\\\\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (oid) (inh\<^sub>s\<^sub>o\<^sub>u\<^sub>n\<^sub>d) (inh\<^sub>m\<^sub>o\<^sub>v\<^sub>i\<^sub>n\<^sub>g) (inh\<^sub>o\<^sub>u\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>w\<^sub>o\<^sub>r\<^sub>l\<^sub>d)) \ None | (mk\\\\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (tt)) \ (case (case tt of (mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (t) (var\<^sub>b\<^sub>o\<^sub>s\<^sub>s) (var\<^sub>s\<^sub>a\<^sub>l\<^sub>a\<^sub>r\<^sub>y)) \ (class_ty_ext_equiv_1of2_get_oid_inh\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (t))) of (oid , var\<^sub>w\<^sub>o\<^sub>r\<^sub>m\<^sub>h\<^sub>o\<^sub>l\<^sub>e , var\<^sub>w\<^sub>e\<^sub>i\<^sub>g\<^sub>h\<^sub>t , var\<^sub>s\<^sub>o\<^sub>u\<^sub>n\<^sub>d , var\<^sub>m\<^sub>o\<^sub>v\<^sub>i\<^sub>n\<^sub>g , var\<^sub>o\<^sub>u\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>w\<^sub>o\<^sub>r\<^sub>l\<^sub>d) \ \(mk2\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n ((class_ty_ext_equiv_1of2_aux\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (tt))))\)))))" definition "class_ty_ext_equiv_1of2\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t = (\ (mk\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (t) (own\<^sub>w\<^sub>o\<^sub>r\<^sub>m\<^sub>h\<^sub>o\<^sub>l\<^sub>e) (own\<^sub>w\<^sub>e\<^sub>i\<^sub>g\<^sub>h\<^sub>t)) \ (case (class_ty_ext_equiv_1of2_get_oid_inh\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (t)) of (oid , inh\<^sub>s\<^sub>o\<^sub>u\<^sub>n\<^sub>d , inh\<^sub>m\<^sub>o\<^sub>v\<^sub>i\<^sub>n\<^sub>g , inh\<^sub>o\<^sub>u\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>w\<^sub>o\<^sub>r\<^sub>l\<^sub>d) \ (mk2oid\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (oid) (inh\<^sub>s\<^sub>o\<^sub>u\<^sub>n\<^sub>d) (inh\<^sub>m\<^sub>o\<^sub>v\<^sub>i\<^sub>n\<^sub>g) (inh\<^sub>o\<^sub>u\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>w\<^sub>o\<^sub>r\<^sub>l\<^sub>d) ((class_ty_ext_equiv_1of2_aux\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t ((mk\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (t) (own\<^sub>w\<^sub>o\<^sub>r\<^sub>m\<^sub>h\<^sub>o\<^sub>l\<^sub>e) (own\<^sub>w\<^sub>e\<^sub>i\<^sub>g\<^sub>h\<^sub>t))))))))" definition "class_ty_ext_equiv_1of2_get_oid_inh\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y = (\ (mk\\\\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y (oid)) \ (oid) | (mk\\\\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n ((mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (t) (var\<^sub>b\<^sub>o\<^sub>s\<^sub>s) (var\<^sub>s\<^sub>a\<^sub>l\<^sub>a\<^sub>r\<^sub>y)))) \ (case (class_ty_ext_equiv_1of2_get_oid_inh\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (t)) of (oid , var\<^sub>w\<^sub>o\<^sub>r\<^sub>m\<^sub>h\<^sub>o\<^sub>l\<^sub>e , var\<^sub>w\<^sub>e\<^sub>i\<^sub>g\<^sub>h\<^sub>t , var\<^sub>s\<^sub>o\<^sub>u\<^sub>n\<^sub>d , var\<^sub>m\<^sub>o\<^sub>v\<^sub>i\<^sub>n\<^sub>g , var\<^sub>o\<^sub>u\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>w\<^sub>o\<^sub>r\<^sub>l\<^sub>d) \ (oid)) | (mk\\\\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t ((mk\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (t) (var\<^sub>w\<^sub>o\<^sub>r\<^sub>m\<^sub>h\<^sub>o\<^sub>l\<^sub>e) (var\<^sub>w\<^sub>e\<^sub>i\<^sub>g\<^sub>h\<^sub>t)))) \ (case (class_ty_ext_equiv_1of2_get_oid_inh\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (t)) of (oid , var\<^sub>s\<^sub>o\<^sub>u\<^sub>n\<^sub>d , var\<^sub>m\<^sub>o\<^sub>v\<^sub>i\<^sub>n\<^sub>g , var\<^sub>o\<^sub>u\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>w\<^sub>o\<^sub>r\<^sub>l\<^sub>d) \ (oid)))" definition "class_ty_ext_equiv_1of2_aux\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y = (\ (mk\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y (t) (own\<^sub>s\<^sub>o\<^sub>u\<^sub>n\<^sub>d) (own\<^sub>m\<^sub>o\<^sub>v\<^sub>i\<^sub>n\<^sub>g) (own\<^sub>o\<^sub>u\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>w\<^sub>o\<^sub>r\<^sub>l\<^sub>d)) \ (mk2\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y (own\<^sub>s\<^sub>o\<^sub>u\<^sub>n\<^sub>d) (own\<^sub>m\<^sub>o\<^sub>v\<^sub>i\<^sub>n\<^sub>g) (own\<^sub>o\<^sub>u\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>w\<^sub>o\<^sub>r\<^sub>l\<^sub>d) ((case t of (mk\\\\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y (oid)) \ None | (mk\\\\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (tt)) \ (case (case tt of (mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (t) (var\<^sub>b\<^sub>o\<^sub>s\<^sub>s) (var\<^sub>s\<^sub>a\<^sub>l\<^sub>a\<^sub>r\<^sub>y)) \ (class_ty_ext_equiv_1of2_get_oid_inh\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (t))) of (oid , var\<^sub>w\<^sub>o\<^sub>r\<^sub>m\<^sub>h\<^sub>o\<^sub>l\<^sub>e , var\<^sub>w\<^sub>e\<^sub>i\<^sub>g\<^sub>h\<^sub>t , var\<^sub>s\<^sub>o\<^sub>u\<^sub>n\<^sub>d , var\<^sub>m\<^sub>o\<^sub>v\<^sub>i\<^sub>n\<^sub>g , var\<^sub>o\<^sub>u\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>w\<^sub>o\<^sub>r\<^sub>l\<^sub>d) \ \(mk2\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t ((mk2\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (var\<^sub>w\<^sub>o\<^sub>r\<^sub>m\<^sub>h\<^sub>o\<^sub>l\<^sub>e) (var\<^sub>w\<^sub>e\<^sub>i\<^sub>g\<^sub>h\<^sub>t) (\(mk2\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n ((class_ty_ext_equiv_1of2_aux\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (tt))))\))))\) | (mk\\\\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (tt)) \ (case (case tt of (mk\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (t) (var\<^sub>w\<^sub>o\<^sub>r\<^sub>m\<^sub>h\<^sub>o\<^sub>l\<^sub>e) (var\<^sub>w\<^sub>e\<^sub>i\<^sub>g\<^sub>h\<^sub>t)) \ (class_ty_ext_equiv_1of2_get_oid_inh\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (t))) of (oid , var\<^sub>s\<^sub>o\<^sub>u\<^sub>n\<^sub>d , var\<^sub>m\<^sub>o\<^sub>v\<^sub>i\<^sub>n\<^sub>g , var\<^sub>o\<^sub>u\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>w\<^sub>o\<^sub>r\<^sub>l\<^sub>d) \ \(mk2\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t ((class_ty_ext_equiv_1of2_aux\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (tt))))\)))))" definition "class_ty_ext_equiv_1of2\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y = (\ (mk\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y (t) (own\<^sub>s\<^sub>o\<^sub>u\<^sub>n\<^sub>d) (own\<^sub>m\<^sub>o\<^sub>v\<^sub>i\<^sub>n\<^sub>g) (own\<^sub>o\<^sub>u\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>w\<^sub>o\<^sub>r\<^sub>l\<^sub>d)) \ (case (class_ty_ext_equiv_1of2_get_oid_inh\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y (t)) of (oid) \ (mk2oid\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y (oid) ((class_ty_ext_equiv_1of2_aux\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y ((mk\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y (t) (own\<^sub>s\<^sub>o\<^sub>u\<^sub>n\<^sub>d) (own\<^sub>m\<^sub>o\<^sub>v\<^sub>i\<^sub>n\<^sub>g) (own\<^sub>o\<^sub>u\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>w\<^sub>o\<^sub>r\<^sub>l\<^sub>d))))))))" definition "class_ty_ext_equiv_1of2_get_oid_inh\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y = (\ (mk\\\\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y (oid)) \ (oid) | (mk\\\\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n ((mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (t) (var\<^sub>b\<^sub>o\<^sub>s\<^sub>s) (var\<^sub>s\<^sub>a\<^sub>l\<^sub>a\<^sub>r\<^sub>y)))) \ (case (class_ty_ext_equiv_1of2_get_oid_inh\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (t)) of (oid , var\<^sub>w\<^sub>o\<^sub>r\<^sub>m\<^sub>h\<^sub>o\<^sub>l\<^sub>e , var\<^sub>w\<^sub>e\<^sub>i\<^sub>g\<^sub>h\<^sub>t , var\<^sub>s\<^sub>o\<^sub>u\<^sub>n\<^sub>d , var\<^sub>m\<^sub>o\<^sub>v\<^sub>i\<^sub>n\<^sub>g , var\<^sub>o\<^sub>u\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>w\<^sub>o\<^sub>r\<^sub>l\<^sub>d) \ (oid)) | (mk\\\\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t ((mk\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (t) (var\<^sub>w\<^sub>o\<^sub>r\<^sub>m\<^sub>h\<^sub>o\<^sub>l\<^sub>e) (var\<^sub>w\<^sub>e\<^sub>i\<^sub>g\<^sub>h\<^sub>t)))) \ (case (class_ty_ext_equiv_1of2_get_oid_inh\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (t)) of (oid , var\<^sub>s\<^sub>o\<^sub>u\<^sub>n\<^sub>d , var\<^sub>m\<^sub>o\<^sub>v\<^sub>i\<^sub>n\<^sub>g , var\<^sub>o\<^sub>u\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>w\<^sub>o\<^sub>r\<^sub>l\<^sub>d) \ (oid)) | (mk\\\\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y ((mk\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y (t) (var\<^sub>s\<^sub>o\<^sub>u\<^sub>n\<^sub>d) (var\<^sub>m\<^sub>o\<^sub>v\<^sub>i\<^sub>n\<^sub>g) (var\<^sub>o\<^sub>u\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>w\<^sub>o\<^sub>r\<^sub>l\<^sub>d)))) \ (case (class_ty_ext_equiv_1of2_get_oid_inh\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y (t)) of (oid) \ (oid)))" definition "class_ty_ext_equiv_1of2_aux\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y = (\ (mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y (t)) \ (mk2\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y ((case t of (mk\\\\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y (oid)) \ None | (mk\\\\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (tt)) \ (case (case tt of (mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (t) (var\<^sub>b\<^sub>o\<^sub>s\<^sub>s) (var\<^sub>s\<^sub>a\<^sub>l\<^sub>a\<^sub>r\<^sub>y)) \ (class_ty_ext_equiv_1of2_get_oid_inh\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (t))) of (oid , var\<^sub>w\<^sub>o\<^sub>r\<^sub>m\<^sub>h\<^sub>o\<^sub>l\<^sub>e , var\<^sub>w\<^sub>e\<^sub>i\<^sub>g\<^sub>h\<^sub>t , var\<^sub>s\<^sub>o\<^sub>u\<^sub>n\<^sub>d , var\<^sub>m\<^sub>o\<^sub>v\<^sub>i\<^sub>n\<^sub>g , var\<^sub>o\<^sub>u\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>w\<^sub>o\<^sub>r\<^sub>l\<^sub>d) \ \(mk2\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y ((mk2\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y (var\<^sub>s\<^sub>o\<^sub>u\<^sub>n\<^sub>d) (var\<^sub>m\<^sub>o\<^sub>v\<^sub>i\<^sub>n\<^sub>g) (var\<^sub>o\<^sub>u\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>w\<^sub>o\<^sub>r\<^sub>l\<^sub>d) (\(mk2\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t ((mk2\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (var\<^sub>w\<^sub>o\<^sub>r\<^sub>m\<^sub>h\<^sub>o\<^sub>l\<^sub>e) (var\<^sub>w\<^sub>e\<^sub>i\<^sub>g\<^sub>h\<^sub>t) (\(mk2\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n ((class_ty_ext_equiv_1of2_aux\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (tt))))\))))\))))\) | (mk\\\\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (tt)) \ (case (case tt of (mk\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (t) (var\<^sub>w\<^sub>o\<^sub>r\<^sub>m\<^sub>h\<^sub>o\<^sub>l\<^sub>e) (var\<^sub>w\<^sub>e\<^sub>i\<^sub>g\<^sub>h\<^sub>t)) \ (class_ty_ext_equiv_1of2_get_oid_inh\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (t))) of (oid , var\<^sub>s\<^sub>o\<^sub>u\<^sub>n\<^sub>d , var\<^sub>m\<^sub>o\<^sub>v\<^sub>i\<^sub>n\<^sub>g , var\<^sub>o\<^sub>u\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>w\<^sub>o\<^sub>r\<^sub>l\<^sub>d) \ \(mk2\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y ((mk2\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y (var\<^sub>s\<^sub>o\<^sub>u\<^sub>n\<^sub>d) (var\<^sub>m\<^sub>o\<^sub>v\<^sub>i\<^sub>n\<^sub>g) (var\<^sub>o\<^sub>u\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>w\<^sub>o\<^sub>r\<^sub>l\<^sub>d) (\(mk2\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t ((class_ty_ext_equiv_1of2_aux\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (tt))))\))))\) | (mk\\\\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y (tt)) \ (case (case tt of (mk\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y (t) (var\<^sub>s\<^sub>o\<^sub>u\<^sub>n\<^sub>d) (var\<^sub>m\<^sub>o\<^sub>v\<^sub>i\<^sub>n\<^sub>g) (var\<^sub>o\<^sub>u\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>w\<^sub>o\<^sub>r\<^sub>l\<^sub>d)) \ (class_ty_ext_equiv_1of2_get_oid_inh\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y (t))) of (oid) \ \(mk2\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y ((class_ty_ext_equiv_1of2_aux\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y (tt))))\)))))" definition "class_ty_ext_equiv_1of2\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y = (\ (mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y (t)) \ (case (class_ty_ext_equiv_1of2_get_oid_inh\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y (t)) of (oid) \ (mk2oid\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y (oid) ((class_ty_ext_equiv_1of2_aux\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y ((mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y (t))))))))" (* 19 ************************************ 52 + 1 *) text \ Now, we construct a concrete ``universe of OclAny types'' by injection into a sum type containing the class types. This type of OclAny will be used as instance for all respective type-variables. \ (* 20 ************************************ 53 + 1 *) (* term Floor1_infra.print_infra_datatype_universe *) datatype \ = in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n "ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" | in\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t "ty\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t" | in\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y "ty\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y" | in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y "ty\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y" (* 21 ************************************ 54 + 1 *) text \ Having fixed the object universe, we can introduce type synonyms that exactly correspond to \OCL types. Again, we exploit that our representation of \OCL is a ``shallow embedding'' with a one-to-one correspondance of \OCL-types to types of the meta-language \HOL. \ (* 22 ************************************ 55 + 7 *) (* term Floor1_infra.print_infra_type_synonym_class *) type_synonym Void = "\ Void" type_synonym Boolean = "\ Boolean" type_synonym Integer = "\ Integer" type_synonym Real = "\ Real" type_synonym String = "\ String" type_synonym '\ val' = "(\, '\) val" type_notation val' ("\(_)") (* 23 ************************************ 62 + 4 *) (* term Floor1_infra.print_infra_type_synonym_class_higher *) type_synonym Person = "\\ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\<^sub>\\\<^sub>\" type_synonym Planet = "\\ty\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t\\<^sub>\\\<^sub>\" type_synonym Galaxy = "\\ty\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y\\<^sub>\\\<^sub>\" type_synonym OclAny = "\\ty\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y\\<^sub>\\\<^sub>\" (* 24 ************************************ 66 + 3 *) (* term Floor1_infra.print_infra_type_synonym_class_rec *) type_synonym Sequence_Person = "(\, ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n option option Sequence\<^sub>b\<^sub>a\<^sub>s\<^sub>e) val" type_synonym Set_Person = "(\, ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n option option Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e) val" type_synonym Set_Sequence_Planet = "(\, ty\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t option option Sequence\<^sub>b\<^sub>a\<^sub>s\<^sub>e Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e) val" (* 25 ************************************ 69 + 0 *) (* term Floor1_infra.print_infra_enum_syn *) (* 26 ************************************ 69 + 1 *) text \ To reuse key-elements of the library like referential equality, we have to show that the object universe belongs to the type class ``oclany,'' \ie, each class type has to provide a function @{term oid_of} yielding the Object ID (oid) of the object. \ (* 27 ************************************ 70 + 4 *) (* term Floor1_infra.print_infra_instantiation_class *) instantiation ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n :: object begin definition oid_of_ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_def : "oid_of = (\ mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n t _ _ \ (case t of (mk\\\\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (t) (_) (_) (_) (_) (_)) \ t))" instance .. end instantiation ty\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t :: object begin definition oid_of_ty\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_def : "oid_of = (\ mk\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t t _ _ \ (case t of (mk\\\\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (t) (_) (_) (_)) \ t | (mk\\\\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (t)) \ (oid_of (t))))" instance .. end instantiation ty\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y :: object begin definition oid_of_ty\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_def : "oid_of = (\ mk\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y t _ _ _ \ (case t of (mk\\\\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y (t)) \ t | (mk\\\\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (t)) \ (oid_of (t)) | (mk\\\\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (t)) \ (oid_of (t))))" instance .. end instantiation ty\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y :: object begin definition oid_of_ty\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_def : "oid_of = (\ mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y t \ (case t of (mk\\\\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y (t)) \ t | (mk\\\\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (t)) \ (oid_of (t)) | (mk\\\\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (t)) \ (oid_of (t)) | (mk\\\\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y (t)) \ (oid_of (t))))" instance .. end (* 28 ************************************ 74 + 1 *) (* term Floor1_infra.print_infra_instantiation_universe *) instantiation \ :: object begin definition oid_of_\_def : "oid_of = (\ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n Person \ oid_of Person | in\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t Planet \ oid_of Planet | in\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y Galaxy \ oid_of Galaxy | in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y OclAny \ oid_of OclAny)" instance .. end (* 29 ************************************ 75 + 1 *) section \Class Model: Instantiation of the Generic Strict Equality\ (* 30 ************************************ 76 + 1 *) text \ We instantiate the referential equality on @{text "Person"} and @{text "OclAny"} \ (* 31 ************************************ 77 + 4 *) (* term Floor1_infra.print_instantia_def_strictrefeq *) overloading StrictRefEq \ "(StrictRefEq::(\Person) \ _ \ _)" begin definition StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n : "(x::\Person) \ y \ StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t x y" end overloading StrictRefEq \ "(StrictRefEq::(\Planet) \ _ \ _)" begin definition StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t : "(x::\Planet) \ y \ StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t x y" end overloading StrictRefEq \ "(StrictRefEq::(\Galaxy) \ _ \ _)" begin definition StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y : "(x::\Galaxy) \ y \ StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t x y" end overloading StrictRefEq \ "(StrictRefEq::(\OclAny) \ _ \ _)" begin definition StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y : "(x::\OclAny) \ y \ StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t x y" end (* 32 ************************************ 81 + 1 *) (* term Floor1_infra.print_instantia_lemmas_strictrefeq *) lemmas[simp,code_unfold] = StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y (* 33 ************************************ 82 + 1 *) text \ For each Class \emph{C}, we will have a casting operation \inlineocl{.oclAsType($C$)}, a test on the actual type \inlineocl{.oclIsTypeOf($C$)} as well as its relaxed form \inlineocl{.oclIsKindOf($C$)} (corresponding exactly to Java's \verb+instanceof+-operator. \ (* 34 ************************************ 83 + 1 *) text \ Thus, since we have two class-types in our concrete class hierarchy, we have two operations to declare and to provide two overloading definitions for the two static types. \ (* 35 ************************************ 84 + 1 *) section \Class Model: OclAsType\ (* 36 ************************************ 85 + 1 *) subsection \Definition\ (* 37 ************************************ 86 + 4 *) (* term Floor1_astype.print_astype_consts *) consts OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n :: "'\ \ \Person" ("(_) .oclAsType'(Person')") consts OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t :: "'\ \ \Planet" ("(_) .oclAsType'(Planet')") consts OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y :: "'\ \ \Galaxy" ("(_) .oclAsType'(Galaxy')") consts OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y :: "'\ \ \OclAny" ("(_) .oclAsType'(OclAny')") (* 38 ************************************ 90 + 16 *) (* term Floor1_astype.print_astype_class *) overloading OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n \ "(OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n::(\Person) \ _)" begin definition OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person : "(x::\Person) .oclAsType(Person) \ x" end overloading OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n \ "(OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n::(\Planet) \ _)" begin definition OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet : "(x::\Planet) .oclAsType(Person) \ (\\. (case (x (\)) of \ \ (invalid (\)) | \\\ \ (null (\)) | \\(mk\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t ((mk\\\\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (Person))) (_) (_))\\ \ \\Person\\ | _ \ (invalid (\))))" end overloading OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n \ "(OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n::(\Galaxy) \ _)" begin definition OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy : "(x::\Galaxy) .oclAsType(Person) \ (\\. (case (x (\)) of \ \ (invalid (\)) | \\\ \ (null (\)) | \\(mk\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y ((mk\\\\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (Person))) (_) (_) (_))\\ \ \\Person\\ | _ \ (invalid (\))))" end overloading OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n \ "(OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n::(\OclAny) \ _)" begin definition OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny : "(x::\OclAny) .oclAsType(Person) \ (\\. (case (x (\)) of \ \ (invalid (\)) | \\\ \ (null (\)) | \\(mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y ((mk\\\\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (Person))))\\ \ \\Person\\ | _ \ (invalid (\))))" end overloading OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t \ "(OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t::(\Planet) \ _)" begin definition OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet : "(x::\Planet) .oclAsType(Planet) \ x" end overloading OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t \ "(OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t::(\Galaxy) \ _)" begin definition OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy : "(x::\Galaxy) .oclAsType(Planet) \ (\\. (case (x (\)) of \ \ (invalid (\)) | \\\ \ (null (\)) | \\(mk\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y ((mk\\\\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (Planet))) (_) (_) (_))\\ \ \\Planet\\ | _ \ (invalid (\))))" end overloading OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t \ "(OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t::(\OclAny) \ _)" begin definition OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny : "(x::\OclAny) .oclAsType(Planet) \ (\\. (case (x (\)) of \ \ (invalid (\)) | \\\ \ (null (\)) | \\(mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y ((mk\\\\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (Planet))))\\ \ \\Planet\\ | _ \ (invalid (\))))" end overloading OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t \ "(OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t::(\Person) \ _)" begin definition OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person : "(x::\Person) .oclAsType(Planet) \ (\\. (case (x (\)) of \ \ (invalid (\)) | \\\ \ (null (\)) | \\Person\\ \ \\(mk\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t ((mk\\\\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (Person))) (None) (None))\\))" end overloading OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y \ "(OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y::(\Galaxy) \ _)" begin definition OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy : "(x::\Galaxy) .oclAsType(Galaxy) \ x" end overloading OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y \ "(OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y::(\OclAny) \ _)" begin definition OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny : "(x::\OclAny) .oclAsType(Galaxy) \ (\\. (case (x (\)) of \ \ (invalid (\)) | \\\ \ (null (\)) | \\(mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y ((mk\\\\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y (Galaxy))))\\ \ \\Galaxy\\ | _ \ (invalid (\))))" end overloading OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y \ "(OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y::(\Person) \ _)" begin definition OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person : "(x::\Person) .oclAsType(Galaxy) \ (\\. (case (x (\)) of \ \ (invalid (\)) | \\\ \ (null (\)) | \\Person\\ \ \\(mk\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y ((mk\\\\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (Person))) (None) (None) (None))\\))" end overloading OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y \ "(OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y::(\Planet) \ _)" begin definition OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet : "(x::\Planet) .oclAsType(Galaxy) \ (\\. (case (x (\)) of \ \ (invalid (\)) | \\\ \ (null (\)) | \\Planet\\ \ \\(mk\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y ((mk\\\\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (Planet))) (None) (None) (None))\\))" end overloading OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y \ "(OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y::(\OclAny) \ _)" begin definition OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny : "(x::\OclAny) .oclAsType(OclAny) \ x" end overloading OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y \ "(OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y::(\Person) \ _)" begin definition OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person : "(x::\Person) .oclAsType(OclAny) \ (\\. (case (x (\)) of \ \ (invalid (\)) | \\\ \ (null (\)) | \\Person\\ \ \\(mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y ((mk\\\\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (Person))))\\))" end overloading OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y \ "(OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y::(\Planet) \ _)" begin definition OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet : "(x::\Planet) .oclAsType(OclAny) \ (\\. (case (x (\)) of \ \ (invalid (\)) | \\\ \ (null (\)) | \\Planet\\ \ \\(mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y ((mk\\\\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (Planet))))\\))" end overloading OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y \ "(OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y::(\Galaxy) \ _)" begin definition OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy : "(x::\Galaxy) .oclAsType(OclAny) \ (\\. (case (x (\)) of \ \ (invalid (\)) | \\\ \ (null (\)) | \\Galaxy\\ \ \\(mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y ((mk\\\\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y (Galaxy))))\\))" end (* 39 ************************************ 106 + 4 *) (* term Floor1_astype.print_astype_from_universe *) definition "OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_\ = (\ (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (Person)) \ \Person\ | (in\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t ((mk\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t ((mk\\\\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (Person))) (_) (_)))) \ \Person\ | (in\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y ((mk\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y ((mk\\\\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (Person))) (_) (_) (_)))) \ \Person\ | (in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y ((mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y ((mk\\\\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (Person)))))) \ \Person\ | _ \ None)" definition "OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_\ = (\ (in\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (Planet)) \ \Planet\ | (in\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y ((mk\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y ((mk\\\\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (Planet))) (_) (_) (_)))) \ \Planet\ | (in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y ((mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y ((mk\\\\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (Planet)))))) \ \Planet\ | (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (Person)) \ \(mk\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t ((mk\\\\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (Person))) (None) (None))\ | _ \ None)" definition "OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_\ = (\ (in\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y (Galaxy)) \ \Galaxy\ | (in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y ((mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y ((mk\\\\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y (Galaxy)))))) \ \Galaxy\ | (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (Person)) \ \(mk\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y ((mk\\\\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (Person))) (None) (None) (None))\ | (in\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (Planet)) \ \(mk\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y ((mk\\\\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (Planet))) (None) (None) (None))\ | _ \ None)" definition "OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\ = Some o (\ (in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y (OclAny)) \ OclAny | (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (Person)) \ (mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y ((mk\\\\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (Person)))) | (in\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (Planet)) \ (mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y ((mk\\\\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (Planet)))) | (in\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y (Galaxy)) \ (mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y ((mk\\\\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y (Galaxy)))))" (* 40 ************************************ 110 + 1 *) (* term Floor1_astype.print_astype_lemmas_id *) lemmas[simp,code_unfold] = OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny (* 41 ************************************ 111 + 1 *) subsection \Context Passing\ (* 42 ************************************ 112 + 64 *) (* term Floor1_astype.print_astype_lemma_cp *) lemma cp_OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_OclAny : "(cp (p)) \ (cp ((\x. (((p ((x::\OclAny)))::\OclAny) .oclAsType(OclAny)))))" by(rule cpI1, simp) lemma cp_OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy_OclAny : "(cp (p)) \ (cp ((\x. (((p ((x::\Galaxy)))::\OclAny) .oclAsType(OclAny)))))" by(rule cpI1, simp) lemma cp_OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet_OclAny : "(cp (p)) \ (cp ((\x. (((p ((x::\Planet)))::\OclAny) .oclAsType(OclAny)))))" by(rule cpI1, simp) lemma cp_OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_OclAny : "(cp (p)) \ (cp ((\x. (((p ((x::\Person)))::\OclAny) .oclAsType(OclAny)))))" by(rule cpI1, simp) lemma cp_OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_Galaxy : "(cp (p)) \ (cp ((\x. (((p ((x::\OclAny)))::\Galaxy) .oclAsType(OclAny)))))" by(rule cpI1, simp add: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy) lemma cp_OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy_Galaxy : "(cp (p)) \ (cp ((\x. (((p ((x::\Galaxy)))::\Galaxy) .oclAsType(OclAny)))))" by(rule cpI1, simp add: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy) lemma cp_OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet_Galaxy : "(cp (p)) \ (cp ((\x. (((p ((x::\Planet)))::\Galaxy) .oclAsType(OclAny)))))" by(rule cpI1, simp add: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy) lemma cp_OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_Galaxy : "(cp (p)) \ (cp ((\x. (((p ((x::\Person)))::\Galaxy) .oclAsType(OclAny)))))" by(rule cpI1, simp add: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy) lemma cp_OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_Planet : "(cp (p)) \ (cp ((\x. (((p ((x::\OclAny)))::\Planet) .oclAsType(OclAny)))))" by(rule cpI1, simp add: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet) lemma cp_OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy_Planet : "(cp (p)) \ (cp ((\x. (((p ((x::\Galaxy)))::\Planet) .oclAsType(OclAny)))))" by(rule cpI1, simp add: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet) lemma cp_OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet_Planet : "(cp (p)) \ (cp ((\x. (((p ((x::\Planet)))::\Planet) .oclAsType(OclAny)))))" by(rule cpI1, simp add: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet) lemma cp_OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_Planet : "(cp (p)) \ (cp ((\x. (((p ((x::\Person)))::\Planet) .oclAsType(OclAny)))))" by(rule cpI1, simp add: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet) lemma cp_OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_Person : "(cp (p)) \ (cp ((\x. (((p ((x::\OclAny)))::\Person) .oclAsType(OclAny)))))" by(rule cpI1, simp add: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person) lemma cp_OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy_Person : "(cp (p)) \ (cp ((\x. (((p ((x::\Galaxy)))::\Person) .oclAsType(OclAny)))))" by(rule cpI1, simp add: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person) lemma cp_OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet_Person : "(cp (p)) \ (cp ((\x. (((p ((x::\Planet)))::\Person) .oclAsType(OclAny)))))" by(rule cpI1, simp add: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person) lemma cp_OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_Person : "(cp (p)) \ (cp ((\x. (((p ((x::\Person)))::\Person) .oclAsType(OclAny)))))" by(rule cpI1, simp add: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person) lemma cp_OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny_OclAny : "(cp (p)) \ (cp ((\x. (((p ((x::\OclAny)))::\OclAny) .oclAsType(Galaxy)))))" by(rule cpI1, simp add: OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny) lemma cp_OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy_OclAny : "(cp (p)) \ (cp ((\x. (((p ((x::\Galaxy)))::\OclAny) .oclAsType(Galaxy)))))" by(rule cpI1, simp add: OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny) lemma cp_OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet_OclAny : "(cp (p)) \ (cp ((\x. (((p ((x::\Planet)))::\OclAny) .oclAsType(Galaxy)))))" by(rule cpI1, simp add: OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny) lemma cp_OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person_OclAny : "(cp (p)) \ (cp ((\x. (((p ((x::\Person)))::\OclAny) .oclAsType(Galaxy)))))" by(rule cpI1, simp add: OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny) lemma cp_OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny_Galaxy : "(cp (p)) \ (cp ((\x. (((p ((x::\OclAny)))::\Galaxy) .oclAsType(Galaxy)))))" by(rule cpI1, simp) lemma cp_OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy_Galaxy : "(cp (p)) \ (cp ((\x. (((p ((x::\Galaxy)))::\Galaxy) .oclAsType(Galaxy)))))" by(rule cpI1, simp) lemma cp_OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet_Galaxy : "(cp (p)) \ (cp ((\x. (((p ((x::\Planet)))::\Galaxy) .oclAsType(Galaxy)))))" by(rule cpI1, simp) lemma cp_OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person_Galaxy : "(cp (p)) \ (cp ((\x. (((p ((x::\Person)))::\Galaxy) .oclAsType(Galaxy)))))" by(rule cpI1, simp) lemma cp_OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny_Planet : "(cp (p)) \ (cp ((\x. (((p ((x::\OclAny)))::\Planet) .oclAsType(Galaxy)))))" by(rule cpI1, simp add: OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet) lemma cp_OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy_Planet : "(cp (p)) \ (cp ((\x. (((p ((x::\Galaxy)))::\Planet) .oclAsType(Galaxy)))))" by(rule cpI1, simp add: OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet) lemma cp_OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet_Planet : "(cp (p)) \ (cp ((\x. (((p ((x::\Planet)))::\Planet) .oclAsType(Galaxy)))))" by(rule cpI1, simp add: OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet) lemma cp_OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person_Planet : "(cp (p)) \ (cp ((\x. (((p ((x::\Person)))::\Planet) .oclAsType(Galaxy)))))" by(rule cpI1, simp add: OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet) lemma cp_OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny_Person : "(cp (p)) \ (cp ((\x. (((p ((x::\OclAny)))::\Person) .oclAsType(Galaxy)))))" by(rule cpI1, simp add: OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person) lemma cp_OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy_Person : "(cp (p)) \ (cp ((\x. (((p ((x::\Galaxy)))::\Person) .oclAsType(Galaxy)))))" by(rule cpI1, simp add: OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person) lemma cp_OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet_Person : "(cp (p)) \ (cp ((\x. (((p ((x::\Planet)))::\Person) .oclAsType(Galaxy)))))" by(rule cpI1, simp add: OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person) lemma cp_OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person_Person : "(cp (p)) \ (cp ((\x. (((p ((x::\Person)))::\Person) .oclAsType(Galaxy)))))" by(rule cpI1, simp add: OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person) lemma cp_OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny_OclAny : "(cp (p)) \ (cp ((\x. (((p ((x::\OclAny)))::\OclAny) .oclAsType(Planet)))))" by(rule cpI1, simp add: OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny) lemma cp_OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy_OclAny : "(cp (p)) \ (cp ((\x. (((p ((x::\Galaxy)))::\OclAny) .oclAsType(Planet)))))" by(rule cpI1, simp add: OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny) lemma cp_OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet_OclAny : "(cp (p)) \ (cp ((\x. (((p ((x::\Planet)))::\OclAny) .oclAsType(Planet)))))" by(rule cpI1, simp add: OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny) lemma cp_OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person_OclAny : "(cp (p)) \ (cp ((\x. (((p ((x::\Person)))::\OclAny) .oclAsType(Planet)))))" by(rule cpI1, simp add: OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny) lemma cp_OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny_Galaxy : "(cp (p)) \ (cp ((\x. (((p ((x::\OclAny)))::\Galaxy) .oclAsType(Planet)))))" by(rule cpI1, simp add: OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy) lemma cp_OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy_Galaxy : "(cp (p)) \ (cp ((\x. (((p ((x::\Galaxy)))::\Galaxy) .oclAsType(Planet)))))" by(rule cpI1, simp add: OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy) lemma cp_OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet_Galaxy : "(cp (p)) \ (cp ((\x. (((p ((x::\Planet)))::\Galaxy) .oclAsType(Planet)))))" by(rule cpI1, simp add: OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy) lemma cp_OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person_Galaxy : "(cp (p)) \ (cp ((\x. (((p ((x::\Person)))::\Galaxy) .oclAsType(Planet)))))" by(rule cpI1, simp add: OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy) lemma cp_OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny_Planet : "(cp (p)) \ (cp ((\x. (((p ((x::\OclAny)))::\Planet) .oclAsType(Planet)))))" by(rule cpI1, simp) lemma cp_OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy_Planet : "(cp (p)) \ (cp ((\x. (((p ((x::\Galaxy)))::\Planet) .oclAsType(Planet)))))" by(rule cpI1, simp) lemma cp_OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet_Planet : "(cp (p)) \ (cp ((\x. (((p ((x::\Planet)))::\Planet) .oclAsType(Planet)))))" by(rule cpI1, simp) lemma cp_OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person_Planet : "(cp (p)) \ (cp ((\x. (((p ((x::\Person)))::\Planet) .oclAsType(Planet)))))" by(rule cpI1, simp) lemma cp_OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny_Person : "(cp (p)) \ (cp ((\x. (((p ((x::\OclAny)))::\Person) .oclAsType(Planet)))))" by(rule cpI1, simp add: OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person) lemma cp_OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy_Person : "(cp (p)) \ (cp ((\x. (((p ((x::\Galaxy)))::\Person) .oclAsType(Planet)))))" by(rule cpI1, simp add: OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person) lemma cp_OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet_Person : "(cp (p)) \ (cp ((\x. (((p ((x::\Planet)))::\Person) .oclAsType(Planet)))))" by(rule cpI1, simp add: OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person) lemma cp_OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person_Person : "(cp (p)) \ (cp ((\x. (((p ((x::\Person)))::\Person) .oclAsType(Planet)))))" by(rule cpI1, simp add: OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person) lemma cp_OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_OclAny : "(cp (p)) \ (cp ((\x. (((p ((x::\OclAny)))::\OclAny) .oclAsType(Person)))))" by(rule cpI1, simp add: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny) lemma cp_OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy_OclAny : "(cp (p)) \ (cp ((\x. (((p ((x::\Galaxy)))::\OclAny) .oclAsType(Person)))))" by(rule cpI1, simp add: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny) lemma cp_OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet_OclAny : "(cp (p)) \ (cp ((\x. (((p ((x::\Planet)))::\OclAny) .oclAsType(Person)))))" by(rule cpI1, simp add: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny) lemma cp_OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_OclAny : "(cp (p)) \ (cp ((\x. (((p ((x::\Person)))::\OclAny) .oclAsType(Person)))))" by(rule cpI1, simp add: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny) lemma cp_OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_Galaxy : "(cp (p)) \ (cp ((\x. (((p ((x::\OclAny)))::\Galaxy) .oclAsType(Person)))))" by(rule cpI1, simp add: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy) lemma cp_OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy_Galaxy : "(cp (p)) \ (cp ((\x. (((p ((x::\Galaxy)))::\Galaxy) .oclAsType(Person)))))" by(rule cpI1, simp add: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy) lemma cp_OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet_Galaxy : "(cp (p)) \ (cp ((\x. (((p ((x::\Planet)))::\Galaxy) .oclAsType(Person)))))" by(rule cpI1, simp add: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy) lemma cp_OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_Galaxy : "(cp (p)) \ (cp ((\x. (((p ((x::\Person)))::\Galaxy) .oclAsType(Person)))))" by(rule cpI1, simp add: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy) lemma cp_OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_Planet : "(cp (p)) \ (cp ((\x. (((p ((x::\OclAny)))::\Planet) .oclAsType(Person)))))" by(rule cpI1, simp add: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet) lemma cp_OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy_Planet : "(cp (p)) \ (cp ((\x. (((p ((x::\Galaxy)))::\Planet) .oclAsType(Person)))))" by(rule cpI1, simp add: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet) lemma cp_OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet_Planet : "(cp (p)) \ (cp ((\x. (((p ((x::\Planet)))::\Planet) .oclAsType(Person)))))" by(rule cpI1, simp add: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet) lemma cp_OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_Planet : "(cp (p)) \ (cp ((\x. (((p ((x::\Person)))::\Planet) .oclAsType(Person)))))" by(rule cpI1, simp add: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet) lemma cp_OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_Person : "(cp (p)) \ (cp ((\x. (((p ((x::\OclAny)))::\Person) .oclAsType(Person)))))" by(rule cpI1, simp) lemma cp_OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy_Person : "(cp (p)) \ (cp ((\x. (((p ((x::\Galaxy)))::\Person) .oclAsType(Person)))))" by(rule cpI1, simp) lemma cp_OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet_Person : "(cp (p)) \ (cp ((\x. (((p ((x::\Planet)))::\Person) .oclAsType(Person)))))" by(rule cpI1, simp) lemma cp_OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_Person : "(cp (p)) \ (cp ((\x. (((p ((x::\Person)))::\Person) .oclAsType(Person)))))" by(rule cpI1, simp) (* 43 ************************************ 176 + 1 *) (* term Floor1_astype.print_astype_lemmas_cp *) lemmas[simp,code_unfold] = cp_OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_OclAny cp_OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy_OclAny cp_OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet_OclAny cp_OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_OclAny cp_OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_Galaxy cp_OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy_Galaxy cp_OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet_Galaxy cp_OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_Galaxy cp_OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_Planet cp_OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy_Planet cp_OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet_Planet cp_OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_Planet cp_OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_Person cp_OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy_Person cp_OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet_Person cp_OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_Person cp_OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny_OclAny cp_OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy_OclAny cp_OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet_OclAny cp_OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person_OclAny cp_OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny_Galaxy cp_OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy_Galaxy cp_OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet_Galaxy cp_OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person_Galaxy cp_OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny_Planet cp_OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy_Planet cp_OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet_Planet cp_OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person_Planet cp_OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny_Person cp_OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy_Person cp_OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet_Person cp_OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person_Person cp_OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny_OclAny cp_OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy_OclAny cp_OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet_OclAny cp_OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person_OclAny cp_OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny_Galaxy cp_OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy_Galaxy cp_OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet_Galaxy cp_OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person_Galaxy cp_OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny_Planet cp_OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy_Planet cp_OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet_Planet cp_OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person_Planet cp_OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny_Person cp_OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy_Person cp_OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet_Person cp_OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person_Person cp_OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_OclAny cp_OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy_OclAny cp_OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet_OclAny cp_OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_OclAny cp_OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_Galaxy cp_OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy_Galaxy cp_OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet_Galaxy cp_OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_Galaxy cp_OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_Planet cp_OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy_Planet cp_OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet_Planet cp_OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_Planet cp_OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_Person cp_OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy_Person cp_OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet_Person cp_OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_Person (* 44 ************************************ 177 + 1 *) subsection \Execution with Invalid or Null as Argument\ (* 45 ************************************ 178 + 32 *) (* term Floor1_astype.print_astype_lemma_strict *) lemma OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_invalid : "((invalid::\OclAny) .oclAsType(OclAny)) = invalid" by(simp) lemma OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy_invalid : "((invalid::\Galaxy) .oclAsType(OclAny)) = invalid" by(rule ext, simp add: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy bot_option_def invalid_def) lemma OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet_invalid : "((invalid::\Planet) .oclAsType(OclAny)) = invalid" by(rule ext, simp add: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet bot_option_def invalid_def) lemma OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_invalid : "((invalid::\Person) .oclAsType(OclAny)) = invalid" by(rule ext, simp add: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person bot_option_def invalid_def) lemma OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_null : "((null::\OclAny) .oclAsType(OclAny)) = null" by(simp) lemma OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy_null : "((null::\Galaxy) .oclAsType(OclAny)) = null" by(rule ext, simp add: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy bot_option_def null_fun_def null_option_def) lemma OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet_null : "((null::\Planet) .oclAsType(OclAny)) = null" by(rule ext, simp add: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet bot_option_def null_fun_def null_option_def) lemma OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_null : "((null::\Person) .oclAsType(OclAny)) = null" by(rule ext, simp add: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person bot_option_def null_fun_def null_option_def) lemma OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny_invalid : "((invalid::\OclAny) .oclAsType(Galaxy)) = invalid" by(rule ext, simp add: OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny bot_option_def invalid_def) lemma OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy_invalid : "((invalid::\Galaxy) .oclAsType(Galaxy)) = invalid" by(simp) lemma OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet_invalid : "((invalid::\Planet) .oclAsType(Galaxy)) = invalid" by(rule ext, simp add: OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet bot_option_def invalid_def) lemma OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person_invalid : "((invalid::\Person) .oclAsType(Galaxy)) = invalid" by(rule ext, simp add: OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person bot_option_def invalid_def) lemma OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny_null : "((null::\OclAny) .oclAsType(Galaxy)) = null" by(rule ext, simp add: OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny bot_option_def null_fun_def null_option_def) lemma OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy_null : "((null::\Galaxy) .oclAsType(Galaxy)) = null" by(simp) lemma OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet_null : "((null::\Planet) .oclAsType(Galaxy)) = null" by(rule ext, simp add: OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet bot_option_def null_fun_def null_option_def) lemma OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person_null : "((null::\Person) .oclAsType(Galaxy)) = null" by(rule ext, simp add: OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person bot_option_def null_fun_def null_option_def) lemma OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny_invalid : "((invalid::\OclAny) .oclAsType(Planet)) = invalid" by(rule ext, simp add: OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny bot_option_def invalid_def) lemma OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy_invalid : "((invalid::\Galaxy) .oclAsType(Planet)) = invalid" by(rule ext, simp add: OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy bot_option_def invalid_def) lemma OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet_invalid : "((invalid::\Planet) .oclAsType(Planet)) = invalid" by(simp) lemma OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person_invalid : "((invalid::\Person) .oclAsType(Planet)) = invalid" by(rule ext, simp add: OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person bot_option_def invalid_def) lemma OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny_null : "((null::\OclAny) .oclAsType(Planet)) = null" by(rule ext, simp add: OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny bot_option_def null_fun_def null_option_def) lemma OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy_null : "((null::\Galaxy) .oclAsType(Planet)) = null" by(rule ext, simp add: OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy bot_option_def null_fun_def null_option_def) lemma OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet_null : "((null::\Planet) .oclAsType(Planet)) = null" by(simp) lemma OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person_null : "((null::\Person) .oclAsType(Planet)) = null" by(rule ext, simp add: OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person bot_option_def null_fun_def null_option_def) lemma OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_invalid : "((invalid::\OclAny) .oclAsType(Person)) = invalid" by(rule ext, simp add: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny bot_option_def invalid_def) lemma OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy_invalid : "((invalid::\Galaxy) .oclAsType(Person)) = invalid" by(rule ext, simp add: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy bot_option_def invalid_def) lemma OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet_invalid : "((invalid::\Planet) .oclAsType(Person)) = invalid" by(rule ext, simp add: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet bot_option_def invalid_def) lemma OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_invalid : "((invalid::\Person) .oclAsType(Person)) = invalid" by(simp) lemma OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_null : "((null::\OclAny) .oclAsType(Person)) = null" by(rule ext, simp add: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny bot_option_def null_fun_def null_option_def) lemma OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy_null : "((null::\Galaxy) .oclAsType(Person)) = null" by(rule ext, simp add: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy bot_option_def null_fun_def null_option_def) lemma OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet_null : "((null::\Planet) .oclAsType(Person)) = null" by(rule ext, simp add: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet bot_option_def null_fun_def null_option_def) lemma OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_null : "((null::\Person) .oclAsType(Person)) = null" by(simp) (* 46 ************************************ 210 + 1 *) (* term Floor1_astype.print_astype_lemmas_strict *) lemmas[simp,code_unfold] = OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_invalid OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy_invalid OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet_invalid OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_invalid OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_null OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy_null OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet_null OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_null OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny_invalid OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy_invalid OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet_invalid OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person_invalid OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny_null OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy_null OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet_null OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person_null OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny_invalid OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy_invalid OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet_invalid OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person_invalid OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny_null OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy_null OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet_null OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person_null OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_invalid OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy_invalid OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet_invalid OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_invalid OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_null OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy_null OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet_null OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_null (* 47 ************************************ 211 + 1 *) subsection \Validity and Definedness Properties\ (* 48 ************************************ 212 + 6 *) (* term Floor1_astype.print_astype_defined *) lemma OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person_defined : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\Person) .oclAsType(Planet)))" using isdef by(auto simp: OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person foundation16 null_option_def bot_option_def) lemma OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person_defined : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\Person) .oclAsType(Galaxy)))" using isdef by(auto simp: OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person foundation16 null_option_def bot_option_def) lemma OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet_defined : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\Planet) .oclAsType(Galaxy)))" using isdef by(auto simp: OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet foundation16 null_option_def bot_option_def) lemma OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_defined : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\Person) .oclAsType(OclAny)))" using isdef by(auto simp: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person foundation16 null_option_def bot_option_def) lemma OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet_defined : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\Planet) .oclAsType(OclAny)))" using isdef by(auto simp: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet foundation16 null_option_def bot_option_def) lemma OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy_defined : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\Galaxy) .oclAsType(OclAny)))" using isdef by(auto simp: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy foundation16 null_option_def bot_option_def) (* 49 ************************************ 218 + 1 *) subsection \Up Down Casting\ (* 50 ************************************ 219 + 6 *) (* term Floor1_astype.print_astype_up_d_cast0 *) lemma up\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_down\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_cast0 : assumes isdef: "\ \ (\ (X))" shows "\ \ (((X::\Person) .oclAsType(Planet)) .oclAsType(Person)) \ X" using isdef by(auto simp: OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet foundation22 foundation16 null_option_def bot_option_def split: ty\\\\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n.split ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n.split) lemma up\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_down\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_cast0 : assumes isdef: "\ \ (\ (X))" shows "\ \ (((X::\Person) .oclAsType(Galaxy)) .oclAsType(Person)) \ X" using isdef by(auto simp: OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy foundation22 foundation16 null_option_def bot_option_def split: ty\\\\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n.split ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n.split) lemma up\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_down\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_cast0 : assumes isdef: "\ \ (\ (X))" shows "\ \ (((X::\Person) .oclAsType(OclAny)) .oclAsType(Person)) \ X" using isdef by(auto simp: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny foundation22 foundation16 null_option_def bot_option_def split: ty\\\\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n.split ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n.split) lemma up\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_down\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_cast0 : assumes isdef: "\ \ (\ (X))" shows "\ \ (((X::\Planet) .oclAsType(Galaxy)) .oclAsType(Planet)) \ X" using isdef by(auto simp: OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy foundation22 foundation16 null_option_def bot_option_def split: ty\\\\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t.split ty\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t.split) lemma up\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_down\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_cast0 : assumes isdef: "\ \ (\ (X))" shows "\ \ (((X::\Planet) .oclAsType(OclAny)) .oclAsType(Planet)) \ X" using isdef by(auto simp: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny foundation22 foundation16 null_option_def bot_option_def split: ty\\\\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t.split ty\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t.split) lemma up\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_down\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_cast0 : assumes isdef: "\ \ (\ (X))" shows "\ \ (((X::\Galaxy) .oclAsType(OclAny)) .oclAsType(Galaxy)) \ X" using isdef by(auto simp: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny foundation22 foundation16 null_option_def bot_option_def split: ty\\\\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y.split ty\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y.split) (* 51 ************************************ 225 + 6 *) (* term Floor1_astype.print_astype_up_d_cast *) lemma up\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_down\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_cast : shows "(((X::\Person) .oclAsType(Planet)) .oclAsType(Person)) = X" apply(rule ext, rename_tac \) apply(rule foundation22[THEN iffD1]) apply(case_tac "\ \ (\ (X))", simp add: up\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_down\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_cast0) apply(simp add: defined_split, elim disjE) apply((erule StrongEq_L_subst2_rev, simp, simp)+) done lemma up\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_down\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_cast : shows "(((X::\Person) .oclAsType(Galaxy)) .oclAsType(Person)) = X" apply(rule ext, rename_tac \) apply(rule foundation22[THEN iffD1]) apply(case_tac "\ \ (\ (X))", simp add: up\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_down\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_cast0) apply(simp add: defined_split, elim disjE) apply((erule StrongEq_L_subst2_rev, simp, simp)+) done lemma up\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_down\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_cast : shows "(((X::\Person) .oclAsType(OclAny)) .oclAsType(Person)) = X" apply(rule ext, rename_tac \) apply(rule foundation22[THEN iffD1]) apply(case_tac "\ \ (\ (X))", simp add: up\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_down\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_cast0) apply(simp add: defined_split, elim disjE) apply((erule StrongEq_L_subst2_rev, simp, simp)+) done lemma up\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_down\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_cast : shows "(((X::\Planet) .oclAsType(Galaxy)) .oclAsType(Planet)) = X" apply(rule ext, rename_tac \) apply(rule foundation22[THEN iffD1]) apply(case_tac "\ \ (\ (X))", simp add: up\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_down\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_cast0) apply(simp add: defined_split, elim disjE) apply((erule StrongEq_L_subst2_rev, simp, simp)+) done lemma up\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_down\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_cast : shows "(((X::\Planet) .oclAsType(OclAny)) .oclAsType(Planet)) = X" apply(rule ext, rename_tac \) apply(rule foundation22[THEN iffD1]) apply(case_tac "\ \ (\ (X))", simp add: up\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_down\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_cast0) apply(simp add: defined_split, elim disjE) apply((erule StrongEq_L_subst2_rev, simp, simp)+) done lemma up\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_down\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_cast : shows "(((X::\Galaxy) .oclAsType(OclAny)) .oclAsType(Galaxy)) = X" apply(rule ext, rename_tac \) apply(rule foundation22[THEN iffD1]) apply(case_tac "\ \ (\ (X))", simp add: up\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_down\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_cast0) apply(simp add: defined_split, elim disjE) apply((erule StrongEq_L_subst2_rev, simp, simp)+) done (* 52 ************************************ 231 + 6 *) (* term Floor1_astype.print_astype_d_up_cast *) lemma down\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_up\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_cast : assumes def_X: "X = ((Y::\Person) .oclAsType(Planet))" shows "(\ \ ((not ((\ (X)))) or ((X .oclAsType(Person)) .oclAsType(Planet)) \ X))" apply(case_tac "(\ \ ((not ((\ (X))))))", rule foundation25, simp) by(rule foundation25', simp add: def_X up\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_down\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_cast StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_sym) lemma down\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_up\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_cast : assumes def_X: "X = ((Y::\Person) .oclAsType(Galaxy))" shows "(\ \ ((not ((\ (X)))) or ((X .oclAsType(Person)) .oclAsType(Galaxy)) \ X))" apply(case_tac "(\ \ ((not ((\ (X))))))", rule foundation25, simp) by(rule foundation25', simp add: def_X up\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_down\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_cast StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_sym) lemma down\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_up\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_cast : assumes def_X: "X = ((Y::\Person) .oclAsType(OclAny))" shows "(\ \ ((not ((\ (X)))) or ((X .oclAsType(Person)) .oclAsType(OclAny)) \ X))" apply(case_tac "(\ \ ((not ((\ (X))))))", rule foundation25, simp) by(rule foundation25', simp add: def_X up\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_down\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_cast StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_sym) lemma down\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_up\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_cast : assumes def_X: "X = ((Y::\Planet) .oclAsType(Galaxy))" shows "(\ \ ((not ((\ (X)))) or ((X .oclAsType(Planet)) .oclAsType(Galaxy)) \ X))" apply(case_tac "(\ \ ((not ((\ (X))))))", rule foundation25, simp) by(rule foundation25', simp add: def_X up\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_down\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_cast StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_sym) lemma down\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_up\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_cast : assumes def_X: "X = ((Y::\Planet) .oclAsType(OclAny))" shows "(\ \ ((not ((\ (X)))) or ((X .oclAsType(Planet)) .oclAsType(OclAny)) \ X))" apply(case_tac "(\ \ ((not ((\ (X))))))", rule foundation25, simp) by(rule foundation25', simp add: def_X up\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_down\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_cast StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_sym) lemma down\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_up\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_cast : assumes def_X: "X = ((Y::\Galaxy) .oclAsType(OclAny))" shows "(\ \ ((not ((\ (X)))) or ((X .oclAsType(Galaxy)) .oclAsType(OclAny)) \ X))" apply(case_tac "(\ \ ((not ((\ (X))))))", rule foundation25, simp) by(rule foundation25', simp add: def_X up\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_down\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_cast StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_sym) (* 53 ************************************ 237 + 1 *) subsection \Const\ (* 54 ************************************ 238 + 16 *) (* term Floor1_astype.print_astype_lemma_const *) lemma OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_const : "(const ((X::\OclAny))) \ (const (X .oclAsType(OclAny)))" by(simp add: const_def, (metis (no_types) OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny prod.collapse bot_option_def invalid_def null_fun_def null_option_def)?) lemma OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy_const : "(const ((X::\Galaxy))) \ (const (X .oclAsType(OclAny)))" by(simp add: const_def, (metis (no_types) OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy prod.collapse bot_option_def invalid_def null_fun_def null_option_def)?) lemma OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet_const : "(const ((X::\Planet))) \ (const (X .oclAsType(OclAny)))" by(simp add: const_def, (metis (no_types) OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet prod.collapse bot_option_def invalid_def null_fun_def null_option_def)?) lemma OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_const : "(const ((X::\Person))) \ (const (X .oclAsType(OclAny)))" by(simp add: const_def, (metis (no_types) OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person prod.collapse bot_option_def invalid_def null_fun_def null_option_def)?) lemma OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny_const : "(const ((X::\OclAny))) \ (const (X .oclAsType(Galaxy)))" by(simp add: const_def, (metis (no_types) OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny prod.collapse bot_option_def invalid_def null_fun_def null_option_def)?) lemma OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy_const : "(const ((X::\Galaxy))) \ (const (X .oclAsType(Galaxy)))" by(simp add: const_def, (metis (no_types) OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy prod.collapse bot_option_def invalid_def null_fun_def null_option_def)?) lemma OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet_const : "(const ((X::\Planet))) \ (const (X .oclAsType(Galaxy)))" by(simp add: const_def, (metis (no_types) OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet prod.collapse bot_option_def invalid_def null_fun_def null_option_def)?) lemma OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person_const : "(const ((X::\Person))) \ (const (X .oclAsType(Galaxy)))" by(simp add: const_def, (metis (no_types) OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person prod.collapse bot_option_def invalid_def null_fun_def null_option_def)?) lemma OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny_const : "(const ((X::\OclAny))) \ (const (X .oclAsType(Planet)))" by(simp add: const_def, (metis (no_types) OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny prod.collapse bot_option_def invalid_def null_fun_def null_option_def)?) lemma OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy_const : "(const ((X::\Galaxy))) \ (const (X .oclAsType(Planet)))" by(simp add: const_def, (metis (no_types) OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy prod.collapse bot_option_def invalid_def null_fun_def null_option_def)?) lemma OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet_const : "(const ((X::\Planet))) \ (const (X .oclAsType(Planet)))" by(simp add: const_def, (metis (no_types) OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet prod.collapse bot_option_def invalid_def null_fun_def null_option_def)?) lemma OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person_const : "(const ((X::\Person))) \ (const (X .oclAsType(Planet)))" by(simp add: const_def, (metis (no_types) OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person prod.collapse bot_option_def invalid_def null_fun_def null_option_def)?) lemma OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_const : "(const ((X::\OclAny))) \ (const (X .oclAsType(Person)))" by(simp add: const_def, (metis (no_types) OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny prod.collapse bot_option_def invalid_def null_fun_def null_option_def)?) lemma OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy_const : "(const ((X::\Galaxy))) \ (const (X .oclAsType(Person)))" by(simp add: const_def, (metis (no_types) OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy prod.collapse bot_option_def invalid_def null_fun_def null_option_def)?) lemma OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet_const : "(const ((X::\Planet))) \ (const (X .oclAsType(Person)))" by(simp add: const_def, (metis (no_types) OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet prod.collapse bot_option_def invalid_def null_fun_def null_option_def)?) lemma OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_const : "(const ((X::\Person))) \ (const (X .oclAsType(Person)))" by(simp add: const_def, (metis (no_types) OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person prod.collapse bot_option_def invalid_def null_fun_def null_option_def)?) (* 55 ************************************ 254 + 1 *) (* term Floor1_astype.print_astype_lemmas_const *) lemmas[simp,code_unfold] = OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_const OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy_const OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet_const OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_const OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny_const OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy_const OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet_const OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person_const OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny_const OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy_const OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet_const OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person_const OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_const OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy_const OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet_const OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_const (* 56 ************************************ 255 + 1 *) section \Class Model: OclIsTypeOf\ (* 57 ************************************ 256 + 1 *) subsection \Definition\ (* 58 ************************************ 257 + 4 *) (* term Floor1_istypeof.print_istypeof_consts *) consts OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n :: "'\ \ Boolean" ("(_) .oclIsTypeOf'(Person')") consts OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t :: "'\ \ Boolean" ("(_) .oclIsTypeOf'(Planet')") consts OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y :: "'\ \ Boolean" ("(_) .oclIsTypeOf'(Galaxy')") consts OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y :: "'\ \ Boolean" ("(_) .oclIsTypeOf'(OclAny')") (* 59 ************************************ 261 + 16 *) (* term Floor1_istypeof.print_istypeof_class *) overloading OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n \ "(OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n::(\Person) \ _)" begin definition OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person : "(x::\Person) .oclIsTypeOf(Person) \ (\\. (case (x (\)) of \ \ (invalid (\)) | \\\ \ (true (\)) | \\(mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n ((mk\\\\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (_) (_) (_) (_) (_) (_))) (_) (_))\\ \ (true (\))))" end overloading OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n \ "(OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n::(\Planet) \ _)" begin definition OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet : "(x::\Planet) .oclIsTypeOf(Person) \ (\\. (case (x (\)) of \ \ (invalid (\)) | \\\ \ (true (\)) | \\(mk\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t ((mk\\\\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (_))) (_) (_))\\ \ (true (\)) | _ \ (false (\))))" end overloading OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n \ "(OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n::(\Galaxy) \ _)" begin definition OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy : "(x::\Galaxy) .oclIsTypeOf(Person) \ (\\. (case (x (\)) of \ \ (invalid (\)) | \\\ \ (true (\)) | \\(mk\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y ((mk\\\\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (_))) (_) (_) (_))\\ \ (true (\)) | _ \ (false (\))))" end overloading OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n \ "(OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n::(\OclAny) \ _)" begin definition OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny : "(x::\OclAny) .oclIsTypeOf(Person) \ (\\. (case (x (\)) of \ \ (invalid (\)) | \\\ \ (true (\)) | \\(mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y ((mk\\\\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (_))))\\ \ (true (\)) | _ \ (false (\))))" end overloading OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t \ "(OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t::(\Planet) \ _)" begin definition OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet : "(x::\Planet) .oclIsTypeOf(Planet) \ (\\. (case (x (\)) of \ \ (invalid (\)) | \\\ \ (true (\)) | \\(mk\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t ((mk\\\\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (_) (_) (_) (_))) (_) (_))\\ \ (true (\)) | _ \ (false (\))))" end overloading OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t \ "(OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t::(\Galaxy) \ _)" begin definition OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy : "(x::\Galaxy) .oclIsTypeOf(Planet) \ (\\. (case (x (\)) of \ \ (invalid (\)) | \\\ \ (true (\)) | \\(mk\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y ((mk\\\\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (_))) (_) (_) (_))\\ \ (true (\)) | _ \ (false (\))))" end overloading OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t \ "(OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t::(\OclAny) \ _)" begin definition OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny : "(x::\OclAny) .oclIsTypeOf(Planet) \ (\\. (case (x (\)) of \ \ (invalid (\)) | \\\ \ (true (\)) | \\(mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y ((mk\\\\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (_))))\\ \ (true (\)) | _ \ (false (\))))" end overloading OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t \ "(OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t::(\Person) \ _)" begin definition OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person : "(x::\Person) .oclIsTypeOf(Planet) \ (\\. (case (x (\)) of \ \ (invalid (\)) | \\\ \ (true (\)) | _ \ (false (\))))" end overloading OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y \ "(OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y::(\Galaxy) \ _)" begin definition OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy : "(x::\Galaxy) .oclIsTypeOf(Galaxy) \ (\\. (case (x (\)) of \ \ (invalid (\)) | \\\ \ (true (\)) | \\(mk\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y ((mk\\\\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y (_))) (_) (_) (_))\\ \ (true (\)) | _ \ (false (\))))" end overloading OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y \ "(OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y::(\OclAny) \ _)" begin definition OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny : "(x::\OclAny) .oclIsTypeOf(Galaxy) \ (\\. (case (x (\)) of \ \ (invalid (\)) | \\\ \ (true (\)) | \\(mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y ((mk\\\\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y (_))))\\ \ (true (\)) | _ \ (false (\))))" end overloading OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y \ "(OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y::(\Person) \ _)" begin definition OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person : "(x::\Person) .oclIsTypeOf(Galaxy) \ (\\. (case (x (\)) of \ \ (invalid (\)) | \\\ \ (true (\)) | _ \ (false (\))))" end overloading OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y \ "(OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y::(\Planet) \ _)" begin definition OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet : "(x::\Planet) .oclIsTypeOf(Galaxy) \ (\\. (case (x (\)) of \ \ (invalid (\)) | \\\ \ (true (\)) | _ \ (false (\))))" end overloading OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y \ "(OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y::(\OclAny) \ _)" begin definition OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny : "(x::\OclAny) .oclIsTypeOf(OclAny) \ (\\. (case (x (\)) of \ \ (invalid (\)) | \\\ \ (true (\)) | \\(mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y ((mk\\\\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y (_))))\\ \ (true (\)) | _ \ (false (\))))" end overloading OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y \ "(OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y::(\Person) \ _)" begin definition OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person : "(x::\Person) .oclIsTypeOf(OclAny) \ (\\. (case (x (\)) of \ \ (invalid (\)) | \\\ \ (true (\)) | _ \ (false (\))))" end overloading OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y \ "(OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y::(\Planet) \ _)" begin definition OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet : "(x::\Planet) .oclIsTypeOf(OclAny) \ (\\. (case (x (\)) of \ \ (invalid (\)) | \\\ \ (true (\)) | _ \ (false (\))))" end overloading OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y \ "(OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y::(\Galaxy) \ _)" begin definition OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy : "(x::\Galaxy) .oclIsTypeOf(OclAny) \ (\\. (case (x (\)) of \ \ (invalid (\)) | \\\ \ (true (\)) | _ \ (false (\))))" end (* 60 ************************************ 277 + 4 *) (* term Floor1_istypeof.print_istypeof_from_universe *) definition "OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_\ = (\ (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (Person)) \ (((((\x _. \\x\\)) (Person))::\Person) .oclIsTypeOf(Person)) | (in\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (Planet)) \ (((((\x _. \\x\\)) (Planet))::\Planet) .oclIsTypeOf(Person)) | (in\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y (Galaxy)) \ (((((\x _. \\x\\)) (Galaxy))::\Galaxy) .oclIsTypeOf(Person)) | (in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y (OclAny)) \ (((((\x _. \\x\\)) (OclAny))::\OclAny) .oclIsTypeOf(Person)))" definition "OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_\ = (\ (in\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (Planet)) \ (((((\x _. \\x\\)) (Planet))::\Planet) .oclIsTypeOf(Planet)) | (in\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y (Galaxy)) \ (((((\x _. \\x\\)) (Galaxy))::\Galaxy) .oclIsTypeOf(Planet)) | (in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y (OclAny)) \ (((((\x _. \\x\\)) (OclAny))::\OclAny) .oclIsTypeOf(Planet)) | (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (Person)) \ (((((\x _. \\x\\)) (Person))::\Person) .oclIsTypeOf(Planet)))" definition "OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_\ = (\ (in\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y (Galaxy)) \ (((((\x _. \\x\\)) (Galaxy))::\Galaxy) .oclIsTypeOf(Galaxy)) | (in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y (OclAny)) \ (((((\x _. \\x\\)) (OclAny))::\OclAny) .oclIsTypeOf(Galaxy)) | (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (Person)) \ (((((\x _. \\x\\)) (Person))::\Person) .oclIsTypeOf(Galaxy)) | (in\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (Planet)) \ (((((\x _. \\x\\)) (Planet))::\Planet) .oclIsTypeOf(Galaxy)))" definition "OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\ = (\ (in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y (OclAny)) \ (((((\x _. \\x\\)) (OclAny))::\OclAny) .oclIsTypeOf(OclAny)) | (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (Person)) \ (((((\x _. \\x\\)) (Person))::\Person) .oclIsTypeOf(OclAny)) | (in\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (Planet)) \ (((((\x _. \\x\\)) (Planet))::\Planet) .oclIsTypeOf(OclAny)) | (in\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y (Galaxy)) \ (((((\x _. \\x\\)) (Galaxy))::\Galaxy) .oclIsTypeOf(OclAny)))" (* 61 ************************************ 281 + 1 *) (* term Floor1_istypeof.print_istypeof_lemmas_id *) lemmas[simp,code_unfold] = OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny (* 62 ************************************ 282 + 1 *) subsection \Context Passing\ (* 63 ************************************ 283 + 64 *) (* term Floor1_istypeof.print_istypeof_lemma_cp *) lemma cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_OclAny : "(cp (p)) \ (cp ((\x. (((p ((x::\OclAny)))::\OclAny) .oclIsTypeOf(OclAny)))))" by(rule cpI1, simp) lemma cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy_OclAny : "(cp (p)) \ (cp ((\x. (((p ((x::\Galaxy)))::\OclAny) .oclIsTypeOf(OclAny)))))" by(rule cpI1, simp) lemma cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet_OclAny : "(cp (p)) \ (cp ((\x. (((p ((x::\Planet)))::\OclAny) .oclIsTypeOf(OclAny)))))" by(rule cpI1, simp) lemma cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_OclAny : "(cp (p)) \ (cp ((\x. (((p ((x::\Person)))::\OclAny) .oclIsTypeOf(OclAny)))))" by(rule cpI1, simp) lemma cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_Galaxy : "(cp (p)) \ (cp ((\x. (((p ((x::\OclAny)))::\Galaxy) .oclIsTypeOf(OclAny)))))" by(rule cpI1, simp add: OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy) lemma cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy_Galaxy : "(cp (p)) \ (cp ((\x. (((p ((x::\Galaxy)))::\Galaxy) .oclIsTypeOf(OclAny)))))" by(rule cpI1, simp add: OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy) lemma cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet_Galaxy : "(cp (p)) \ (cp ((\x. (((p ((x::\Planet)))::\Galaxy) .oclIsTypeOf(OclAny)))))" by(rule cpI1, simp add: OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy) lemma cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_Galaxy : "(cp (p)) \ (cp ((\x. (((p ((x::\Person)))::\Galaxy) .oclIsTypeOf(OclAny)))))" by(rule cpI1, simp add: OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy) lemma cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_Planet : "(cp (p)) \ (cp ((\x. (((p ((x::\OclAny)))::\Planet) .oclIsTypeOf(OclAny)))))" by(rule cpI1, simp add: OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet) lemma cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy_Planet : "(cp (p)) \ (cp ((\x. (((p ((x::\Galaxy)))::\Planet) .oclIsTypeOf(OclAny)))))" by(rule cpI1, simp add: OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet) lemma cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet_Planet : "(cp (p)) \ (cp ((\x. (((p ((x::\Planet)))::\Planet) .oclIsTypeOf(OclAny)))))" by(rule cpI1, simp add: OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet) lemma cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_Planet : "(cp (p)) \ (cp ((\x. (((p ((x::\Person)))::\Planet) .oclIsTypeOf(OclAny)))))" by(rule cpI1, simp add: OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet) lemma cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_Person : "(cp (p)) \ (cp ((\x. (((p ((x::\OclAny)))::\Person) .oclIsTypeOf(OclAny)))))" by(rule cpI1, simp add: OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person) lemma cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy_Person : "(cp (p)) \ (cp ((\x. (((p ((x::\Galaxy)))::\Person) .oclIsTypeOf(OclAny)))))" by(rule cpI1, simp add: OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person) lemma cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet_Person : "(cp (p)) \ (cp ((\x. (((p ((x::\Planet)))::\Person) .oclIsTypeOf(OclAny)))))" by(rule cpI1, simp add: OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person) lemma cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_Person : "(cp (p)) \ (cp ((\x. (((p ((x::\Person)))::\Person) .oclIsTypeOf(OclAny)))))" by(rule cpI1, simp add: OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person) lemma cp_OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny_OclAny : "(cp (p)) \ (cp ((\x. (((p ((x::\OclAny)))::\OclAny) .oclIsTypeOf(Galaxy)))))" by(rule cpI1, simp add: OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny) lemma cp_OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy_OclAny : "(cp (p)) \ (cp ((\x. (((p ((x::\Galaxy)))::\OclAny) .oclIsTypeOf(Galaxy)))))" by(rule cpI1, simp add: OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny) lemma cp_OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet_OclAny : "(cp (p)) \ (cp ((\x. (((p ((x::\Planet)))::\OclAny) .oclIsTypeOf(Galaxy)))))" by(rule cpI1, simp add: OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny) lemma cp_OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person_OclAny : "(cp (p)) \ (cp ((\x. (((p ((x::\Person)))::\OclAny) .oclIsTypeOf(Galaxy)))))" by(rule cpI1, simp add: OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny) lemma cp_OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny_Galaxy : "(cp (p)) \ (cp ((\x. (((p ((x::\OclAny)))::\Galaxy) .oclIsTypeOf(Galaxy)))))" by(rule cpI1, simp) lemma cp_OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy_Galaxy : "(cp (p)) \ (cp ((\x. (((p ((x::\Galaxy)))::\Galaxy) .oclIsTypeOf(Galaxy)))))" by(rule cpI1, simp) lemma cp_OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet_Galaxy : "(cp (p)) \ (cp ((\x. (((p ((x::\Planet)))::\Galaxy) .oclIsTypeOf(Galaxy)))))" by(rule cpI1, simp) lemma cp_OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person_Galaxy : "(cp (p)) \ (cp ((\x. (((p ((x::\Person)))::\Galaxy) .oclIsTypeOf(Galaxy)))))" by(rule cpI1, simp) lemma cp_OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny_Planet : "(cp (p)) \ (cp ((\x. (((p ((x::\OclAny)))::\Planet) .oclIsTypeOf(Galaxy)))))" by(rule cpI1, simp add: OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet) lemma cp_OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy_Planet : "(cp (p)) \ (cp ((\x. (((p ((x::\Galaxy)))::\Planet) .oclIsTypeOf(Galaxy)))))" by(rule cpI1, simp add: OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet) lemma cp_OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet_Planet : "(cp (p)) \ (cp ((\x. (((p ((x::\Planet)))::\Planet) .oclIsTypeOf(Galaxy)))))" by(rule cpI1, simp add: OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet) lemma cp_OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person_Planet : "(cp (p)) \ (cp ((\x. (((p ((x::\Person)))::\Planet) .oclIsTypeOf(Galaxy)))))" by(rule cpI1, simp add: OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet) lemma cp_OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny_Person : "(cp (p)) \ (cp ((\x. (((p ((x::\OclAny)))::\Person) .oclIsTypeOf(Galaxy)))))" by(rule cpI1, simp add: OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person) lemma cp_OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy_Person : "(cp (p)) \ (cp ((\x. (((p ((x::\Galaxy)))::\Person) .oclIsTypeOf(Galaxy)))))" by(rule cpI1, simp add: OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person) lemma cp_OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet_Person : "(cp (p)) \ (cp ((\x. (((p ((x::\Planet)))::\Person) .oclIsTypeOf(Galaxy)))))" by(rule cpI1, simp add: OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person) lemma cp_OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person_Person : "(cp (p)) \ (cp ((\x. (((p ((x::\Person)))::\Person) .oclIsTypeOf(Galaxy)))))" by(rule cpI1, simp add: OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person) lemma cp_OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny_OclAny : "(cp (p)) \ (cp ((\x. (((p ((x::\OclAny)))::\OclAny) .oclIsTypeOf(Planet)))))" by(rule cpI1, simp add: OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny) lemma cp_OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy_OclAny : "(cp (p)) \ (cp ((\x. (((p ((x::\Galaxy)))::\OclAny) .oclIsTypeOf(Planet)))))" by(rule cpI1, simp add: OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny) lemma cp_OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet_OclAny : "(cp (p)) \ (cp ((\x. (((p ((x::\Planet)))::\OclAny) .oclIsTypeOf(Planet)))))" by(rule cpI1, simp add: OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny) lemma cp_OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person_OclAny : "(cp (p)) \ (cp ((\x. (((p ((x::\Person)))::\OclAny) .oclIsTypeOf(Planet)))))" by(rule cpI1, simp add: OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny) lemma cp_OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny_Galaxy : "(cp (p)) \ (cp ((\x. (((p ((x::\OclAny)))::\Galaxy) .oclIsTypeOf(Planet)))))" by(rule cpI1, simp add: OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy) lemma cp_OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy_Galaxy : "(cp (p)) \ (cp ((\x. (((p ((x::\Galaxy)))::\Galaxy) .oclIsTypeOf(Planet)))))" by(rule cpI1, simp add: OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy) lemma cp_OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet_Galaxy : "(cp (p)) \ (cp ((\x. (((p ((x::\Planet)))::\Galaxy) .oclIsTypeOf(Planet)))))" by(rule cpI1, simp add: OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy) lemma cp_OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person_Galaxy : "(cp (p)) \ (cp ((\x. (((p ((x::\Person)))::\Galaxy) .oclIsTypeOf(Planet)))))" by(rule cpI1, simp add: OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy) lemma cp_OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny_Planet : "(cp (p)) \ (cp ((\x. (((p ((x::\OclAny)))::\Planet) .oclIsTypeOf(Planet)))))" by(rule cpI1, simp) lemma cp_OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy_Planet : "(cp (p)) \ (cp ((\x. (((p ((x::\Galaxy)))::\Planet) .oclIsTypeOf(Planet)))))" by(rule cpI1, simp) lemma cp_OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet_Planet : "(cp (p)) \ (cp ((\x. (((p ((x::\Planet)))::\Planet) .oclIsTypeOf(Planet)))))" by(rule cpI1, simp) lemma cp_OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person_Planet : "(cp (p)) \ (cp ((\x. (((p ((x::\Person)))::\Planet) .oclIsTypeOf(Planet)))))" by(rule cpI1, simp) lemma cp_OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny_Person : "(cp (p)) \ (cp ((\x. (((p ((x::\OclAny)))::\Person) .oclIsTypeOf(Planet)))))" by(rule cpI1, simp add: OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person) lemma cp_OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy_Person : "(cp (p)) \ (cp ((\x. (((p ((x::\Galaxy)))::\Person) .oclIsTypeOf(Planet)))))" by(rule cpI1, simp add: OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person) lemma cp_OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet_Person : "(cp (p)) \ (cp ((\x. (((p ((x::\Planet)))::\Person) .oclIsTypeOf(Planet)))))" by(rule cpI1, simp add: OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person) lemma cp_OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person_Person : "(cp (p)) \ (cp ((\x. (((p ((x::\Person)))::\Person) .oclIsTypeOf(Planet)))))" by(rule cpI1, simp add: OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person) lemma cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_OclAny : "(cp (p)) \ (cp ((\x. (((p ((x::\OclAny)))::\OclAny) .oclIsTypeOf(Person)))))" by(rule cpI1, simp add: OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny) lemma cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy_OclAny : "(cp (p)) \ (cp ((\x. (((p ((x::\Galaxy)))::\OclAny) .oclIsTypeOf(Person)))))" by(rule cpI1, simp add: OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny) lemma cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet_OclAny : "(cp (p)) \ (cp ((\x. (((p ((x::\Planet)))::\OclAny) .oclIsTypeOf(Person)))))" by(rule cpI1, simp add: OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny) lemma cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_OclAny : "(cp (p)) \ (cp ((\x. (((p ((x::\Person)))::\OclAny) .oclIsTypeOf(Person)))))" by(rule cpI1, simp add: OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny) lemma cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_Galaxy : "(cp (p)) \ (cp ((\x. (((p ((x::\OclAny)))::\Galaxy) .oclIsTypeOf(Person)))))" by(rule cpI1, simp add: OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy) lemma cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy_Galaxy : "(cp (p)) \ (cp ((\x. (((p ((x::\Galaxy)))::\Galaxy) .oclIsTypeOf(Person)))))" by(rule cpI1, simp add: OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy) lemma cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet_Galaxy : "(cp (p)) \ (cp ((\x. (((p ((x::\Planet)))::\Galaxy) .oclIsTypeOf(Person)))))" by(rule cpI1, simp add: OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy) lemma cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_Galaxy : "(cp (p)) \ (cp ((\x. (((p ((x::\Person)))::\Galaxy) .oclIsTypeOf(Person)))))" by(rule cpI1, simp add: OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy) lemma cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_Planet : "(cp (p)) \ (cp ((\x. (((p ((x::\OclAny)))::\Planet) .oclIsTypeOf(Person)))))" by(rule cpI1, simp add: OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet) lemma cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy_Planet : "(cp (p)) \ (cp ((\x. (((p ((x::\Galaxy)))::\Planet) .oclIsTypeOf(Person)))))" by(rule cpI1, simp add: OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet) lemma cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet_Planet : "(cp (p)) \ (cp ((\x. (((p ((x::\Planet)))::\Planet) .oclIsTypeOf(Person)))))" by(rule cpI1, simp add: OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet) lemma cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_Planet : "(cp (p)) \ (cp ((\x. (((p ((x::\Person)))::\Planet) .oclIsTypeOf(Person)))))" by(rule cpI1, simp add: OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet) lemma cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_Person : "(cp (p)) \ (cp ((\x. (((p ((x::\OclAny)))::\Person) .oclIsTypeOf(Person)))))" by(rule cpI1, simp) lemma cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy_Person : "(cp (p)) \ (cp ((\x. (((p ((x::\Galaxy)))::\Person) .oclIsTypeOf(Person)))))" by(rule cpI1, simp) lemma cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet_Person : "(cp (p)) \ (cp ((\x. (((p ((x::\Planet)))::\Person) .oclIsTypeOf(Person)))))" by(rule cpI1, simp) lemma cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_Person : "(cp (p)) \ (cp ((\x. (((p ((x::\Person)))::\Person) .oclIsTypeOf(Person)))))" by(rule cpI1, simp) (* 64 ************************************ 347 + 1 *) (* term Floor1_istypeof.print_istypeof_lemmas_cp *) lemmas[simp,code_unfold] = cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_OclAny cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy_OclAny cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet_OclAny cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_OclAny cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_Galaxy cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy_Galaxy cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet_Galaxy cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_Galaxy cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_Planet cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy_Planet cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet_Planet cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_Planet cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_Person cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy_Person cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet_Person cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_Person cp_OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny_OclAny cp_OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy_OclAny cp_OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet_OclAny cp_OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person_OclAny cp_OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny_Galaxy cp_OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy_Galaxy cp_OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet_Galaxy cp_OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person_Galaxy cp_OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny_Planet cp_OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy_Planet cp_OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet_Planet cp_OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person_Planet cp_OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny_Person cp_OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy_Person cp_OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet_Person cp_OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person_Person cp_OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny_OclAny cp_OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy_OclAny cp_OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet_OclAny cp_OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person_OclAny cp_OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny_Galaxy cp_OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy_Galaxy cp_OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet_Galaxy cp_OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person_Galaxy cp_OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny_Planet cp_OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy_Planet cp_OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet_Planet cp_OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person_Planet cp_OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny_Person cp_OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy_Person cp_OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet_Person cp_OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person_Person cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_OclAny cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy_OclAny cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet_OclAny cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_OclAny cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_Galaxy cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy_Galaxy cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet_Galaxy cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_Galaxy cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_Planet cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy_Planet cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet_Planet cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_Planet cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_Person cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy_Person cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet_Person cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_Person (* 65 ************************************ 348 + 1 *) subsection \Execution with Invalid or Null as Argument\ (* 66 ************************************ 349 + 32 *) (* term Floor1_istypeof.print_istypeof_lemma_strict *) lemma OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_invalid : "((invalid::\OclAny) .oclIsTypeOf(OclAny)) = invalid" by(rule ext, simp add: bot_option_def invalid_def) lemma OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy_invalid : "((invalid::\Galaxy) .oclIsTypeOf(OclAny)) = invalid" by(rule ext, simp add: OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy bot_option_def invalid_def) lemma OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet_invalid : "((invalid::\Planet) .oclIsTypeOf(OclAny)) = invalid" by(rule ext, simp add: OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet bot_option_def invalid_def) lemma OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_invalid : "((invalid::\Person) .oclIsTypeOf(OclAny)) = invalid" by(rule ext, simp add: OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person bot_option_def invalid_def) lemma OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_null : "((null::\OclAny) .oclIsTypeOf(OclAny)) = true" by(rule ext, simp add: bot_option_def null_fun_def null_option_def) lemma OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy_null : "((null::\Galaxy) .oclIsTypeOf(OclAny)) = true" by(rule ext, simp add: OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy bot_option_def null_fun_def null_option_def) lemma OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet_null : "((null::\Planet) .oclIsTypeOf(OclAny)) = true" by(rule ext, simp add: OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet bot_option_def null_fun_def null_option_def) lemma OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_null : "((null::\Person) .oclIsTypeOf(OclAny)) = true" by(rule ext, simp add: OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person bot_option_def null_fun_def null_option_def) lemma OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny_invalid : "((invalid::\OclAny) .oclIsTypeOf(Galaxy)) = invalid" by(rule ext, simp add: OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny bot_option_def invalid_def) lemma OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy_invalid : "((invalid::\Galaxy) .oclIsTypeOf(Galaxy)) = invalid" by(rule ext, simp add: bot_option_def invalid_def) lemma OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet_invalid : "((invalid::\Planet) .oclIsTypeOf(Galaxy)) = invalid" by(rule ext, simp add: OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet bot_option_def invalid_def) lemma OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person_invalid : "((invalid::\Person) .oclIsTypeOf(Galaxy)) = invalid" by(rule ext, simp add: OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person bot_option_def invalid_def) lemma OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny_null : "((null::\OclAny) .oclIsTypeOf(Galaxy)) = true" by(rule ext, simp add: OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny bot_option_def null_fun_def null_option_def) lemma OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy_null : "((null::\Galaxy) .oclIsTypeOf(Galaxy)) = true" by(rule ext, simp add: bot_option_def null_fun_def null_option_def) lemma OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet_null : "((null::\Planet) .oclIsTypeOf(Galaxy)) = true" by(rule ext, simp add: OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet bot_option_def null_fun_def null_option_def) lemma OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person_null : "((null::\Person) .oclIsTypeOf(Galaxy)) = true" by(rule ext, simp add: OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person bot_option_def null_fun_def null_option_def) lemma OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny_invalid : "((invalid::\OclAny) .oclIsTypeOf(Planet)) = invalid" by(rule ext, simp add: OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny bot_option_def invalid_def) lemma OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy_invalid : "((invalid::\Galaxy) .oclIsTypeOf(Planet)) = invalid" by(rule ext, simp add: OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy bot_option_def invalid_def) lemma OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet_invalid : "((invalid::\Planet) .oclIsTypeOf(Planet)) = invalid" by(rule ext, simp add: bot_option_def invalid_def) lemma OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person_invalid : "((invalid::\Person) .oclIsTypeOf(Planet)) = invalid" by(rule ext, simp add: OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person bot_option_def invalid_def) lemma OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny_null : "((null::\OclAny) .oclIsTypeOf(Planet)) = true" by(rule ext, simp add: OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny bot_option_def null_fun_def null_option_def) lemma OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy_null : "((null::\Galaxy) .oclIsTypeOf(Planet)) = true" by(rule ext, simp add: OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy bot_option_def null_fun_def null_option_def) lemma OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet_null : "((null::\Planet) .oclIsTypeOf(Planet)) = true" by(rule ext, simp add: bot_option_def null_fun_def null_option_def) lemma OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person_null : "((null::\Person) .oclIsTypeOf(Planet)) = true" by(rule ext, simp add: OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person bot_option_def null_fun_def null_option_def) lemma OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_invalid : "((invalid::\OclAny) .oclIsTypeOf(Person)) = invalid" by(rule ext, simp add: OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny bot_option_def invalid_def) lemma OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy_invalid : "((invalid::\Galaxy) .oclIsTypeOf(Person)) = invalid" by(rule ext, simp add: OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy bot_option_def invalid_def) lemma OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet_invalid : "((invalid::\Planet) .oclIsTypeOf(Person)) = invalid" by(rule ext, simp add: OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet bot_option_def invalid_def) lemma OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_invalid : "((invalid::\Person) .oclIsTypeOf(Person)) = invalid" by(rule ext, simp add: bot_option_def invalid_def) lemma OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_null : "((null::\OclAny) .oclIsTypeOf(Person)) = true" by(rule ext, simp add: OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny bot_option_def null_fun_def null_option_def) lemma OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy_null : "((null::\Galaxy) .oclIsTypeOf(Person)) = true" by(rule ext, simp add: OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy bot_option_def null_fun_def null_option_def) lemma OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet_null : "((null::\Planet) .oclIsTypeOf(Person)) = true" by(rule ext, simp add: OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet bot_option_def null_fun_def null_option_def) lemma OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_null : "((null::\Person) .oclIsTypeOf(Person)) = true" by(rule ext, simp add: bot_option_def null_fun_def null_option_def) (* 67 ************************************ 381 + 1 *) (* term Floor1_istypeof.print_istypeof_lemmas_strict *) lemmas[simp,code_unfold] = OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_invalid OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy_invalid OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet_invalid OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_invalid OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_null OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy_null OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet_null OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_null OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny_invalid OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy_invalid OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet_invalid OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person_invalid OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny_null OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy_null OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet_null OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person_null OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny_invalid OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy_invalid OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet_invalid OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person_invalid OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny_null OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy_null OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet_null OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person_null OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_invalid OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy_invalid OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet_invalid OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_invalid OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_null OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy_null OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet_null OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_null (* 68 ************************************ 382 + 1 *) subsection \Validity and Definedness Properties\ (* 69 ************************************ 383 + 16 *) (* term Floor1_istypeof.print_istypeof_defined *) lemma OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_defined : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\Person) .oclIsTypeOf(Person)))" apply(insert isdef[simplified foundation18'], simp only: OclValid_def, subst cp_defined) by(auto simp: cp_defined[symmetric ] bot_option_def OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person split: option.split ty\\\\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n.split ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n.split) lemma OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet_defined : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\Planet) .oclIsTypeOf(Person)))" apply(insert isdef[simplified foundation18'], simp only: OclValid_def, subst cp_defined) by(auto simp: cp_defined[symmetric ] bot_option_def OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet split: option.split ty\\\\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t.split ty\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t.split) lemma OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy_defined : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\Galaxy) .oclIsTypeOf(Person)))" apply(insert isdef[simplified foundation18'], simp only: OclValid_def, subst cp_defined) by(auto simp: cp_defined[symmetric ] bot_option_def OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy split: option.split ty\\\\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y.split ty\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y.split) lemma OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_defined : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\OclAny) .oclIsTypeOf(Person)))" apply(insert isdef[simplified foundation18'], simp only: OclValid_def, subst cp_defined) by(auto simp: cp_defined[symmetric ] bot_option_def OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny split: option.split ty\\\\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y.split ty\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y.split) lemma OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet_defined : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\Planet) .oclIsTypeOf(Planet)))" apply(insert isdef[simplified foundation18'], simp only: OclValid_def, subst cp_defined) by(auto simp: cp_defined[symmetric ] bot_option_def OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet split: option.split ty\\\\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t.split ty\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t.split) lemma OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy_defined : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\Galaxy) .oclIsTypeOf(Planet)))" apply(insert isdef[simplified foundation18'], simp only: OclValid_def, subst cp_defined) by(auto simp: cp_defined[symmetric ] bot_option_def OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy split: option.split ty\\\\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y.split ty\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y.split) lemma OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny_defined : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\OclAny) .oclIsTypeOf(Planet)))" apply(insert isdef[simplified foundation18'], simp only: OclValid_def, subst cp_defined) by(auto simp: cp_defined[symmetric ] bot_option_def OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny split: option.split ty\\\\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y.split ty\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y.split) lemma OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person_defined : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\Person) .oclIsTypeOf(Planet)))" apply(insert isdef[simplified foundation18'], simp only: OclValid_def, subst cp_defined) by(auto simp: cp_defined[symmetric ] bot_option_def OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person split: option.split ty\\\\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n.split ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n.split) lemma OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy_defined : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\Galaxy) .oclIsTypeOf(Galaxy)))" apply(insert isdef[simplified foundation18'], simp only: OclValid_def, subst cp_defined) by(auto simp: cp_defined[symmetric ] bot_option_def OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy split: option.split ty\\\\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y.split ty\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y.split) lemma OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny_defined : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\OclAny) .oclIsTypeOf(Galaxy)))" apply(insert isdef[simplified foundation18'], simp only: OclValid_def, subst cp_defined) by(auto simp: cp_defined[symmetric ] bot_option_def OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny split: option.split ty\\\\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y.split ty\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y.split) lemma OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person_defined : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\Person) .oclIsTypeOf(Galaxy)))" apply(insert isdef[simplified foundation18'], simp only: OclValid_def, subst cp_defined) by(auto simp: cp_defined[symmetric ] bot_option_def OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person split: option.split ty\\\\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n.split ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n.split) lemma OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet_defined : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\Planet) .oclIsTypeOf(Galaxy)))" apply(insert isdef[simplified foundation18'], simp only: OclValid_def, subst cp_defined) by(auto simp: cp_defined[symmetric ] bot_option_def OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet split: option.split ty\\\\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t.split ty\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t.split) lemma OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_defined : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\OclAny) .oclIsTypeOf(OclAny)))" apply(insert isdef[simplified foundation18'], simp only: OclValid_def, subst cp_defined) by(auto simp: cp_defined[symmetric ] bot_option_def OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny split: option.split ty\\\\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y.split ty\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y.split) lemma OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_defined : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\Person) .oclIsTypeOf(OclAny)))" apply(insert isdef[simplified foundation18'], simp only: OclValid_def, subst cp_defined) by(auto simp: cp_defined[symmetric ] bot_option_def OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person split: option.split ty\\\\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n.split ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n.split) lemma OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet_defined : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\Planet) .oclIsTypeOf(OclAny)))" apply(insert isdef[simplified foundation18'], simp only: OclValid_def, subst cp_defined) by(auto simp: cp_defined[symmetric ] bot_option_def OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet split: option.split ty\\\\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t.split ty\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t.split) lemma OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy_defined : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\Galaxy) .oclIsTypeOf(OclAny)))" apply(insert isdef[simplified foundation18'], simp only: OclValid_def, subst cp_defined) by(auto simp: cp_defined[symmetric ] bot_option_def OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy split: option.split ty\\\\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y.split ty\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y.split) (* 70 ************************************ 399 + 16 *) (* term Floor1_istypeof.print_istypeof_defined' *) lemma OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_defined' : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\Person) .oclIsTypeOf(Person)))" by(rule OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_defined[OF isdef[THEN foundation20]]) lemma OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet_defined' : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\Planet) .oclIsTypeOf(Person)))" by(rule OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet_defined[OF isdef[THEN foundation20]]) lemma OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy_defined' : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\Galaxy) .oclIsTypeOf(Person)))" by(rule OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy_defined[OF isdef[THEN foundation20]]) lemma OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_defined' : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\OclAny) .oclIsTypeOf(Person)))" by(rule OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_defined[OF isdef[THEN foundation20]]) lemma OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet_defined' : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\Planet) .oclIsTypeOf(Planet)))" by(rule OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet_defined[OF isdef[THEN foundation20]]) lemma OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy_defined' : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\Galaxy) .oclIsTypeOf(Planet)))" by(rule OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy_defined[OF isdef[THEN foundation20]]) lemma OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny_defined' : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\OclAny) .oclIsTypeOf(Planet)))" by(rule OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny_defined[OF isdef[THEN foundation20]]) lemma OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person_defined' : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\Person) .oclIsTypeOf(Planet)))" by(rule OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person_defined[OF isdef[THEN foundation20]]) lemma OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy_defined' : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\Galaxy) .oclIsTypeOf(Galaxy)))" by(rule OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy_defined[OF isdef[THEN foundation20]]) lemma OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny_defined' : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\OclAny) .oclIsTypeOf(Galaxy)))" by(rule OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny_defined[OF isdef[THEN foundation20]]) lemma OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person_defined' : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\Person) .oclIsTypeOf(Galaxy)))" by(rule OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person_defined[OF isdef[THEN foundation20]]) lemma OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet_defined' : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\Planet) .oclIsTypeOf(Galaxy)))" by(rule OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet_defined[OF isdef[THEN foundation20]]) lemma OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_defined' : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\OclAny) .oclIsTypeOf(OclAny)))" by(rule OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_defined[OF isdef[THEN foundation20]]) lemma OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_defined' : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\Person) .oclIsTypeOf(OclAny)))" by(rule OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_defined[OF isdef[THEN foundation20]]) lemma OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet_defined' : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\Planet) .oclIsTypeOf(OclAny)))" by(rule OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet_defined[OF isdef[THEN foundation20]]) lemma OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy_defined' : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\Galaxy) .oclIsTypeOf(OclAny)))" by(rule OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy_defined[OF isdef[THEN foundation20]]) (* 71 ************************************ 415 + 1 *) subsection \Up Down Casting\ (* 72 ************************************ 416 + 6 *) (* term Floor1_istypeof.print_istypeof_up_larger *) lemma actualType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_larger_staticType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t : assumes isdef: "\ \ (\ (X))" shows "\ \ ((X::\Person) .oclIsTypeOf(Planet)) \ false" using isdef by(auto simp: OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person foundation22 foundation16 null_option_def bot_option_def) lemma actualType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_larger_staticType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y : assumes isdef: "\ \ (\ (X))" shows "\ \ ((X::\Person) .oclIsTypeOf(Galaxy)) \ false" using isdef by(auto simp: OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person foundation22 foundation16 null_option_def bot_option_def) lemma actualType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_larger_staticType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y : assumes isdef: "\ \ (\ (X))" shows "\ \ ((X::\Person) .oclIsTypeOf(OclAny)) \ false" using isdef by(auto simp: OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person foundation22 foundation16 null_option_def bot_option_def) lemma actualType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_larger_staticType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y : assumes isdef: "\ \ (\ (X))" shows "\ \ ((X::\Planet) .oclIsTypeOf(Galaxy)) \ false" using isdef by(auto simp: OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet foundation22 foundation16 null_option_def bot_option_def) lemma actualType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_larger_staticType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y : assumes isdef: "\ \ (\ (X))" shows "\ \ ((X::\Planet) .oclIsTypeOf(OclAny)) \ false" using isdef by(auto simp: OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet foundation22 foundation16 null_option_def bot_option_def) lemma actualType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_larger_staticType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y : assumes isdef: "\ \ (\ (X))" shows "\ \ ((X::\Galaxy) .oclIsTypeOf(OclAny)) \ false" using isdef by(auto simp: OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy foundation22 foundation16 null_option_def bot_option_def) (* 73 ************************************ 422 + 10 *) (* term Floor1_istypeof.print_istypeof_up_d_cast *) lemma down_cast_type\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_from_Planet_to_Person : assumes istyp: "\ \ ((X::\Planet) .oclIsTypeOf(Planet))" assumes isdef: "\ \ (\ (X))" shows "\ \ (X .oclAsType(Person)) \ invalid" using istyp isdef apply(auto simp: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet foundation22 foundation16 null_option_def bot_option_def split: ty\\\\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t.split ty\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t.split) by(simp add: OclValid_def false_def true_def) lemma down_cast_type\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_from_Galaxy_to_Person : assumes istyp: "\ \ ((X::\Galaxy) .oclIsTypeOf(Galaxy))" assumes isdef: "\ \ (\ (X))" shows "\ \ (X .oclAsType(Person)) \ invalid" using istyp isdef apply(auto simp: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy foundation22 foundation16 null_option_def bot_option_def split: ty\\\\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y.split ty\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y.split) by(simp add: OclValid_def false_def true_def) lemma down_cast_type\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_from_OclAny_to_Person : assumes istyp: "\ \ ((X::\OclAny) .oclIsTypeOf(OclAny))" assumes isdef: "\ \ (\ (X))" shows "\ \ (X .oclAsType(Person)) \ invalid" using istyp isdef apply(auto simp: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny foundation22 foundation16 null_option_def bot_option_def split: ty\\\\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y.split ty\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y.split) by(simp add: OclValid_def false_def true_def) lemma down_cast_type\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_from_Galaxy_to_Person : assumes istyp: "\ \ ((X::\Galaxy) .oclIsTypeOf(Planet))" assumes isdef: "\ \ (\ (X))" shows "\ \ (X .oclAsType(Person)) \ invalid" using istyp isdef apply(auto simp: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy foundation22 foundation16 null_option_def bot_option_def split: ty\\\\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y.split ty\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y.split) by(simp add: OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy OclValid_def false_def true_def) lemma down_cast_type\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_from_OclAny_to_Person : assumes istyp: "\ \ ((X::\OclAny) .oclIsTypeOf(Planet))" assumes isdef: "\ \ (\ (X))" shows "\ \ (X .oclAsType(Person)) \ invalid" using istyp isdef apply(auto simp: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny foundation22 foundation16 null_option_def bot_option_def split: ty\\\\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y.split ty\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y.split) by(simp add: OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny OclValid_def false_def true_def) lemma down_cast_type\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_from_OclAny_to_Person : assumes istyp: "\ \ ((X::\OclAny) .oclIsTypeOf(Galaxy))" assumes isdef: "\ \ (\ (X))" shows "\ \ (X .oclAsType(Person)) \ invalid" using istyp isdef apply(auto simp: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny foundation22 foundation16 null_option_def bot_option_def split: ty\\\\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y.split ty\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y.split) by(simp add: OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny OclValid_def false_def true_def) lemma down_cast_type\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_from_Galaxy_to_Planet : assumes istyp: "\ \ ((X::\Galaxy) .oclIsTypeOf(Galaxy))" assumes isdef: "\ \ (\ (X))" shows "\ \ (X .oclAsType(Planet)) \ invalid" using istyp isdef apply(auto simp: OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy foundation22 foundation16 null_option_def bot_option_def split: ty\\\\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y.split ty\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y.split) by(simp add: OclValid_def false_def true_def) lemma down_cast_type\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_from_OclAny_to_Planet : assumes istyp: "\ \ ((X::\OclAny) .oclIsTypeOf(OclAny))" assumes isdef: "\ \ (\ (X))" shows "\ \ (X .oclAsType(Planet)) \ invalid" using istyp isdef apply(auto simp: OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny foundation22 foundation16 null_option_def bot_option_def split: ty\\\\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y.split ty\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y.split) by(simp add: OclValid_def false_def true_def) lemma down_cast_type\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_from_OclAny_to_Planet : assumes istyp: "\ \ ((X::\OclAny) .oclIsTypeOf(Galaxy))" assumes isdef: "\ \ (\ (X))" shows "\ \ (X .oclAsType(Planet)) \ invalid" using istyp isdef apply(auto simp: OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny foundation22 foundation16 null_option_def bot_option_def split: ty\\\\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y.split ty\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y.split) by(simp add: OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny OclValid_def false_def true_def) lemma down_cast_type\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_from_OclAny_to_Galaxy : assumes istyp: "\ \ ((X::\OclAny) .oclIsTypeOf(OclAny))" assumes isdef: "\ \ (\ (X))" shows "\ \ (X .oclAsType(Galaxy)) \ invalid" using istyp isdef apply(auto simp: OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny foundation22 foundation16 null_option_def bot_option_def split: ty\\\\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y.split ty\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y.split) by(simp add: OclValid_def false_def true_def) (* 74 ************************************ 432 + 1 *) subsection \Const\ (* 75 ************************************ 433 + 1 *) section \Class Model: OclIsKindOf\ (* 76 ************************************ 434 + 1 *) subsection \Definition\ (* 77 ************************************ 435 + 4 *) (* term Floor1_iskindof.print_iskindof_consts *) consts OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n :: "'\ \ Boolean" ("(_) .oclIsKindOf'(Person')") consts OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t :: "'\ \ Boolean" ("(_) .oclIsKindOf'(Planet')") consts OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y :: "'\ \ Boolean" ("(_) .oclIsKindOf'(Galaxy')") consts OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y :: "'\ \ Boolean" ("(_) .oclIsKindOf'(OclAny')") (* 78 ************************************ 439 + 16 *) (* term Floor1_iskindof.print_iskindof_class *) overloading OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n \ "(OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n::(\Person) \ _)" begin definition OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person : "(x::\Person) .oclIsKindOf(Person) \ (x .oclIsTypeOf(Person))" end overloading OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n \ "(OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n::(\Planet) \ _)" begin definition OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet : "(x::\Planet) .oclIsKindOf(Person) \ (x .oclIsTypeOf(Person))" end overloading OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n \ "(OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n::(\Galaxy) \ _)" begin definition OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy : "(x::\Galaxy) .oclIsKindOf(Person) \ (x .oclIsTypeOf(Person))" end overloading OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n \ "(OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n::(\OclAny) \ _)" begin definition OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny : "(x::\OclAny) .oclIsKindOf(Person) \ (x .oclIsTypeOf(Person))" end overloading OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t \ "(OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t::(\Planet) \ _)" begin definition OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet : "(x::\Planet) .oclIsKindOf(Planet) \ (x .oclIsTypeOf(Planet)) or (x .oclIsKindOf(Person))" end overloading OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t \ "(OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t::(\Galaxy) \ _)" begin definition OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy : "(x::\Galaxy) .oclIsKindOf(Planet) \ (x .oclIsTypeOf(Planet)) or (x .oclIsKindOf(Person))" end overloading OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t \ "(OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t::(\OclAny) \ _)" begin definition OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny : "(x::\OclAny) .oclIsKindOf(Planet) \ (x .oclIsTypeOf(Planet)) or (x .oclIsKindOf(Person))" end overloading OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t \ "(OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t::(\Person) \ _)" begin definition OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person : "(x::\Person) .oclIsKindOf(Planet) \ (x .oclIsTypeOf(Planet)) or (x .oclIsKindOf(Person))" end overloading OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y \ "(OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y::(\Galaxy) \ _)" begin definition OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy : "(x::\Galaxy) .oclIsKindOf(Galaxy) \ (x .oclIsTypeOf(Galaxy)) or (x .oclIsKindOf(Planet))" end overloading OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y \ "(OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y::(\OclAny) \ _)" begin definition OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny : "(x::\OclAny) .oclIsKindOf(Galaxy) \ (x .oclIsTypeOf(Galaxy)) or (x .oclIsKindOf(Planet))" end overloading OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y \ "(OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y::(\Person) \ _)" begin definition OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person : "(x::\Person) .oclIsKindOf(Galaxy) \ (x .oclIsTypeOf(Galaxy)) or (x .oclIsKindOf(Planet))" end overloading OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y \ "(OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y::(\Planet) \ _)" begin definition OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet : "(x::\Planet) .oclIsKindOf(Galaxy) \ (x .oclIsTypeOf(Galaxy)) or (x .oclIsKindOf(Planet))" end overloading OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y \ "(OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y::(\OclAny) \ _)" begin definition OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny : "(x::\OclAny) .oclIsKindOf(OclAny) \ (x .oclIsTypeOf(OclAny)) or (x .oclIsKindOf(Galaxy))" end overloading OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y \ "(OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y::(\Person) \ _)" begin definition OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person : "(x::\Person) .oclIsKindOf(OclAny) \ (x .oclIsTypeOf(OclAny)) or (x .oclIsKindOf(Galaxy))" end overloading OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y \ "(OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y::(\Planet) \ _)" begin definition OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet : "(x::\Planet) .oclIsKindOf(OclAny) \ (x .oclIsTypeOf(OclAny)) or (x .oclIsKindOf(Galaxy))" end overloading OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y \ "(OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y::(\Galaxy) \ _)" begin definition OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy : "(x::\Galaxy) .oclIsKindOf(OclAny) \ (x .oclIsTypeOf(OclAny)) or (x .oclIsKindOf(Galaxy))" end (* 79 ************************************ 455 + 4 *) (* term Floor1_iskindof.print_iskindof_from_universe *) definition "OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_\ = (\ (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (Person)) \ (((((\x _. \\x\\)) (Person))::\Person) .oclIsKindOf(Person)) | (in\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (Planet)) \ (((((\x _. \\x\\)) (Planet))::\Planet) .oclIsKindOf(Person)) | (in\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y (Galaxy)) \ (((((\x _. \\x\\)) (Galaxy))::\Galaxy) .oclIsKindOf(Person)) | (in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y (OclAny)) \ (((((\x _. \\x\\)) (OclAny))::\OclAny) .oclIsKindOf(Person)))" definition "OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_\ = (\ (in\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (Planet)) \ (((((\x _. \\x\\)) (Planet))::\Planet) .oclIsKindOf(Planet)) | (in\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y (Galaxy)) \ (((((\x _. \\x\\)) (Galaxy))::\Galaxy) .oclIsKindOf(Planet)) | (in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y (OclAny)) \ (((((\x _. \\x\\)) (OclAny))::\OclAny) .oclIsKindOf(Planet)) | (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (Person)) \ (((((\x _. \\x\\)) (Person))::\Person) .oclIsKindOf(Planet)))" definition "OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_\ = (\ (in\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y (Galaxy)) \ (((((\x _. \\x\\)) (Galaxy))::\Galaxy) .oclIsKindOf(Galaxy)) | (in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y (OclAny)) \ (((((\x _. \\x\\)) (OclAny))::\OclAny) .oclIsKindOf(Galaxy)) | (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (Person)) \ (((((\x _. \\x\\)) (Person))::\Person) .oclIsKindOf(Galaxy)) | (in\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (Planet)) \ (((((\x _. \\x\\)) (Planet))::\Planet) .oclIsKindOf(Galaxy)))" definition "OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\ = (\ (in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y (OclAny)) \ (((((\x _. \\x\\)) (OclAny))::\OclAny) .oclIsKindOf(OclAny)) | (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (Person)) \ (((((\x _. \\x\\)) (Person))::\Person) .oclIsKindOf(OclAny)) | (in\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (Planet)) \ (((((\x _. \\x\\)) (Planet))::\Planet) .oclIsKindOf(OclAny)) | (in\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y (Galaxy)) \ (((((\x _. \\x\\)) (Galaxy))::\Galaxy) .oclIsKindOf(OclAny)))" (* 80 ************************************ 459 + 1 *) (* term Floor1_iskindof.print_iskindof_lemmas_id *) lemmas[simp,code_unfold] = OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny (* 81 ************************************ 460 + 1 *) subsection \Context Passing\ (* 82 ************************************ 461 + 64 *) (* term Floor1_iskindof.print_iskindof_lemma_cp *) lemma cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_Person : "(cp (p)) \ (cp ((\x. (((p ((x::\Person)))::\Person) .oclIsKindOf(Person)))))" by(simp only: OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person, simp only: cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_Person) lemma cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet_Person : "(cp (p)) \ (cp ((\x. (((p ((x::\Planet)))::\Person) .oclIsKindOf(Person)))))" by(simp only: OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person, simp only: cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet_Person) lemma cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy_Person : "(cp (p)) \ (cp ((\x. (((p ((x::\Galaxy)))::\Person) .oclIsKindOf(Person)))))" by(simp only: OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person, simp only: cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy_Person) lemma cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_Person : "(cp (p)) \ (cp ((\x. (((p ((x::\OclAny)))::\Person) .oclIsKindOf(Person)))))" by(simp only: OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person, simp only: cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_Person) lemma cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_Planet : "(cp (p)) \ (cp ((\x. (((p ((x::\Person)))::\Planet) .oclIsKindOf(Person)))))" by(simp only: OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet, simp only: cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_Planet) lemma cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet_Planet : "(cp (p)) \ (cp ((\x. (((p ((x::\Planet)))::\Planet) .oclIsKindOf(Person)))))" by(simp only: OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet, simp only: cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet_Planet) lemma cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy_Planet : "(cp (p)) \ (cp ((\x. (((p ((x::\Galaxy)))::\Planet) .oclIsKindOf(Person)))))" by(simp only: OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet, simp only: cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy_Planet) lemma cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_Planet : "(cp (p)) \ (cp ((\x. (((p ((x::\OclAny)))::\Planet) .oclIsKindOf(Person)))))" by(simp only: OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet, simp only: cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_Planet) lemma cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_Galaxy : "(cp (p)) \ (cp ((\x. (((p ((x::\Person)))::\Galaxy) .oclIsKindOf(Person)))))" by(simp only: OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy, simp only: cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_Galaxy) lemma cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet_Galaxy : "(cp (p)) \ (cp ((\x. (((p ((x::\Planet)))::\Galaxy) .oclIsKindOf(Person)))))" by(simp only: OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy, simp only: cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet_Galaxy) lemma cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy_Galaxy : "(cp (p)) \ (cp ((\x. (((p ((x::\Galaxy)))::\Galaxy) .oclIsKindOf(Person)))))" by(simp only: OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy, simp only: cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy_Galaxy) lemma cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_Galaxy : "(cp (p)) \ (cp ((\x. (((p ((x::\OclAny)))::\Galaxy) .oclIsKindOf(Person)))))" by(simp only: OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy, simp only: cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_Galaxy) lemma cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_OclAny : "(cp (p)) \ (cp ((\x. (((p ((x::\Person)))::\OclAny) .oclIsKindOf(Person)))))" by(simp only: OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny, simp only: cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_OclAny) lemma cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet_OclAny : "(cp (p)) \ (cp ((\x. (((p ((x::\Planet)))::\OclAny) .oclIsKindOf(Person)))))" by(simp only: OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny, simp only: cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet_OclAny) lemma cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy_OclAny : "(cp (p)) \ (cp ((\x. (((p ((x::\Galaxy)))::\OclAny) .oclIsKindOf(Person)))))" by(simp only: OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny, simp only: cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy_OclAny) lemma cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_OclAny : "(cp (p)) \ (cp ((\x. (((p ((x::\OclAny)))::\OclAny) .oclIsKindOf(Person)))))" by(simp only: OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny, simp only: cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_OclAny) lemma cp_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet_Planet : "(cp (p)) \ (cp ((\x. (((p ((x::\Planet)))::\Planet) .oclIsKindOf(Planet)))))" apply(simp only: OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet) apply((rule cpI2[where f = "(or)"], (rule allI)+, rule cp_OclOr)+) apply(simp only: cp_OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet_Planet) by(simp only: cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet_Planet) lemma cp_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy_Planet : "(cp (p)) \ (cp ((\x. (((p ((x::\Galaxy)))::\Planet) .oclIsKindOf(Planet)))))" apply(simp only: OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet) apply((rule cpI2[where f = "(or)"], (rule allI)+, rule cp_OclOr)+) apply(simp only: cp_OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy_Planet) by(simp only: cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy_Planet) lemma cp_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny_Planet : "(cp (p)) \ (cp ((\x. (((p ((x::\OclAny)))::\Planet) .oclIsKindOf(Planet)))))" apply(simp only: OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet) apply((rule cpI2[where f = "(or)"], (rule allI)+, rule cp_OclOr)+) apply(simp only: cp_OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny_Planet) by(simp only: cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_Planet) lemma cp_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person_Planet : "(cp (p)) \ (cp ((\x. (((p ((x::\Person)))::\Planet) .oclIsKindOf(Planet)))))" apply(simp only: OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet) apply((rule cpI2[where f = "(or)"], (rule allI)+, rule cp_OclOr)+) apply(simp only: cp_OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person_Planet) by(simp only: cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_Planet) lemma cp_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet_Galaxy : "(cp (p)) \ (cp ((\x. (((p ((x::\Planet)))::\Galaxy) .oclIsKindOf(Planet)))))" apply(simp only: OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy) apply((rule cpI2[where f = "(or)"], (rule allI)+, rule cp_OclOr)+) apply(simp only: cp_OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet_Galaxy) by(simp only: cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet_Galaxy) lemma cp_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy_Galaxy : "(cp (p)) \ (cp ((\x. (((p ((x::\Galaxy)))::\Galaxy) .oclIsKindOf(Planet)))))" apply(simp only: OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy) apply((rule cpI2[where f = "(or)"], (rule allI)+, rule cp_OclOr)+) apply(simp only: cp_OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy_Galaxy) by(simp only: cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy_Galaxy) lemma cp_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny_Galaxy : "(cp (p)) \ (cp ((\x. (((p ((x::\OclAny)))::\Galaxy) .oclIsKindOf(Planet)))))" apply(simp only: OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy) apply((rule cpI2[where f = "(or)"], (rule allI)+, rule cp_OclOr)+) apply(simp only: cp_OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny_Galaxy) by(simp only: cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_Galaxy) lemma cp_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person_Galaxy : "(cp (p)) \ (cp ((\x. (((p ((x::\Person)))::\Galaxy) .oclIsKindOf(Planet)))))" apply(simp only: OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy) apply((rule cpI2[where f = "(or)"], (rule allI)+, rule cp_OclOr)+) apply(simp only: cp_OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person_Galaxy) by(simp only: cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_Galaxy) lemma cp_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet_OclAny : "(cp (p)) \ (cp ((\x. (((p ((x::\Planet)))::\OclAny) .oclIsKindOf(Planet)))))" apply(simp only: OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny) apply((rule cpI2[where f = "(or)"], (rule allI)+, rule cp_OclOr)+) apply(simp only: cp_OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet_OclAny) by(simp only: cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet_OclAny) lemma cp_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy_OclAny : "(cp (p)) \ (cp ((\x. (((p ((x::\Galaxy)))::\OclAny) .oclIsKindOf(Planet)))))" apply(simp only: OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny) apply((rule cpI2[where f = "(or)"], (rule allI)+, rule cp_OclOr)+) apply(simp only: cp_OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy_OclAny) by(simp only: cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy_OclAny) lemma cp_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny_OclAny : "(cp (p)) \ (cp ((\x. (((p ((x::\OclAny)))::\OclAny) .oclIsKindOf(Planet)))))" apply(simp only: OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny) apply((rule cpI2[where f = "(or)"], (rule allI)+, rule cp_OclOr)+) apply(simp only: cp_OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny_OclAny) by(simp only: cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_OclAny) lemma cp_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person_OclAny : "(cp (p)) \ (cp ((\x. (((p ((x::\Person)))::\OclAny) .oclIsKindOf(Planet)))))" apply(simp only: OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny) apply((rule cpI2[where f = "(or)"], (rule allI)+, rule cp_OclOr)+) apply(simp only: cp_OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person_OclAny) by(simp only: cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_OclAny) lemma cp_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet_Person : "(cp (p)) \ (cp ((\x. (((p ((x::\Planet)))::\Person) .oclIsKindOf(Planet)))))" apply(simp only: OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person) apply((rule cpI2[where f = "(or)"], (rule allI)+, rule cp_OclOr)+) apply(simp only: cp_OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet_Person) by(simp only: cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet_Person) lemma cp_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy_Person : "(cp (p)) \ (cp ((\x. (((p ((x::\Galaxy)))::\Person) .oclIsKindOf(Planet)))))" apply(simp only: OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person) apply((rule cpI2[where f = "(or)"], (rule allI)+, rule cp_OclOr)+) apply(simp only: cp_OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy_Person) by(simp only: cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy_Person) lemma cp_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny_Person : "(cp (p)) \ (cp ((\x. (((p ((x::\OclAny)))::\Person) .oclIsKindOf(Planet)))))" apply(simp only: OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person) apply((rule cpI2[where f = "(or)"], (rule allI)+, rule cp_OclOr)+) apply(simp only: cp_OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny_Person) by(simp only: cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_Person) lemma cp_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person_Person : "(cp (p)) \ (cp ((\x. (((p ((x::\Person)))::\Person) .oclIsKindOf(Planet)))))" apply(simp only: OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person) apply((rule cpI2[where f = "(or)"], (rule allI)+, rule cp_OclOr)+) apply(simp only: cp_OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person_Person) by(simp only: cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_Person) lemma cp_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy_Galaxy : "(cp (p)) \ (cp ((\x. (((p ((x::\Galaxy)))::\Galaxy) .oclIsKindOf(Galaxy)))))" apply(simp only: OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy) apply((rule cpI2[where f = "(or)"], (rule allI)+, rule cp_OclOr)+) apply(simp only: cp_OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy_Galaxy) by(simp only: cp_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy_Galaxy) lemma cp_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny_Galaxy : "(cp (p)) \ (cp ((\x. (((p ((x::\OclAny)))::\Galaxy) .oclIsKindOf(Galaxy)))))" apply(simp only: OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy) apply((rule cpI2[where f = "(or)"], (rule allI)+, rule cp_OclOr)+) apply(simp only: cp_OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny_Galaxy) by(simp only: cp_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny_Galaxy) lemma cp_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person_Galaxy : "(cp (p)) \ (cp ((\x. (((p ((x::\Person)))::\Galaxy) .oclIsKindOf(Galaxy)))))" apply(simp only: OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy) apply((rule cpI2[where f = "(or)"], (rule allI)+, rule cp_OclOr)+) apply(simp only: cp_OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person_Galaxy) by(simp only: cp_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person_Galaxy) lemma cp_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet_Galaxy : "(cp (p)) \ (cp ((\x. (((p ((x::\Planet)))::\Galaxy) .oclIsKindOf(Galaxy)))))" apply(simp only: OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy) apply((rule cpI2[where f = "(or)"], (rule allI)+, rule cp_OclOr)+) apply(simp only: cp_OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet_Galaxy) by(simp only: cp_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet_Galaxy) lemma cp_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy_OclAny : "(cp (p)) \ (cp ((\x. (((p ((x::\Galaxy)))::\OclAny) .oclIsKindOf(Galaxy)))))" apply(simp only: OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny) apply((rule cpI2[where f = "(or)"], (rule allI)+, rule cp_OclOr)+) apply(simp only: cp_OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy_OclAny) by(simp only: cp_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy_OclAny) lemma cp_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny_OclAny : "(cp (p)) \ (cp ((\x. (((p ((x::\OclAny)))::\OclAny) .oclIsKindOf(Galaxy)))))" apply(simp only: OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny) apply((rule cpI2[where f = "(or)"], (rule allI)+, rule cp_OclOr)+) apply(simp only: cp_OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny_OclAny) by(simp only: cp_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny_OclAny) lemma cp_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person_OclAny : "(cp (p)) \ (cp ((\x. (((p ((x::\Person)))::\OclAny) .oclIsKindOf(Galaxy)))))" apply(simp only: OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny) apply((rule cpI2[where f = "(or)"], (rule allI)+, rule cp_OclOr)+) apply(simp only: cp_OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person_OclAny) by(simp only: cp_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person_OclAny) lemma cp_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet_OclAny : "(cp (p)) \ (cp ((\x. (((p ((x::\Planet)))::\OclAny) .oclIsKindOf(Galaxy)))))" apply(simp only: OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny) apply((rule cpI2[where f = "(or)"], (rule allI)+, rule cp_OclOr)+) apply(simp only: cp_OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet_OclAny) by(simp only: cp_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet_OclAny) lemma cp_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy_Person : "(cp (p)) \ (cp ((\x. (((p ((x::\Galaxy)))::\Person) .oclIsKindOf(Galaxy)))))" apply(simp only: OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person) apply((rule cpI2[where f = "(or)"], (rule allI)+, rule cp_OclOr)+) apply(simp only: cp_OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy_Person) by(simp only: cp_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy_Person) lemma cp_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny_Person : "(cp (p)) \ (cp ((\x. (((p ((x::\OclAny)))::\Person) .oclIsKindOf(Galaxy)))))" apply(simp only: OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person) apply((rule cpI2[where f = "(or)"], (rule allI)+, rule cp_OclOr)+) apply(simp only: cp_OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny_Person) by(simp only: cp_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny_Person) lemma cp_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person_Person : "(cp (p)) \ (cp ((\x. (((p ((x::\Person)))::\Person) .oclIsKindOf(Galaxy)))))" apply(simp only: OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person) apply((rule cpI2[where f = "(or)"], (rule allI)+, rule cp_OclOr)+) apply(simp only: cp_OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person_Person) by(simp only: cp_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person_Person) lemma cp_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet_Person : "(cp (p)) \ (cp ((\x. (((p ((x::\Planet)))::\Person) .oclIsKindOf(Galaxy)))))" apply(simp only: OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person) apply((rule cpI2[where f = "(or)"], (rule allI)+, rule cp_OclOr)+) apply(simp only: cp_OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet_Person) by(simp only: cp_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet_Person) lemma cp_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy_Planet : "(cp (p)) \ (cp ((\x. (((p ((x::\Galaxy)))::\Planet) .oclIsKindOf(Galaxy)))))" apply(simp only: OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet) apply((rule cpI2[where f = "(or)"], (rule allI)+, rule cp_OclOr)+) apply(simp only: cp_OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy_Planet) by(simp only: cp_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy_Planet) lemma cp_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny_Planet : "(cp (p)) \ (cp ((\x. (((p ((x::\OclAny)))::\Planet) .oclIsKindOf(Galaxy)))))" apply(simp only: OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet) apply((rule cpI2[where f = "(or)"], (rule allI)+, rule cp_OclOr)+) apply(simp only: cp_OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny_Planet) by(simp only: cp_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny_Planet) lemma cp_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person_Planet : "(cp (p)) \ (cp ((\x. (((p ((x::\Person)))::\Planet) .oclIsKindOf(Galaxy)))))" apply(simp only: OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet) apply((rule cpI2[where f = "(or)"], (rule allI)+, rule cp_OclOr)+) apply(simp only: cp_OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person_Planet) by(simp only: cp_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person_Planet) lemma cp_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet_Planet : "(cp (p)) \ (cp ((\x. (((p ((x::\Planet)))::\Planet) .oclIsKindOf(Galaxy)))))" apply(simp only: OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet) apply((rule cpI2[where f = "(or)"], (rule allI)+, rule cp_OclOr)+) apply(simp only: cp_OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet_Planet) by(simp only: cp_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet_Planet) lemma cp_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_OclAny : "(cp (p)) \ (cp ((\x. (((p ((x::\OclAny)))::\OclAny) .oclIsKindOf(OclAny)))))" apply(simp only: OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny) apply((rule cpI2[where f = "(or)"], (rule allI)+, rule cp_OclOr)+) apply(simp only: cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_OclAny) by(simp only: cp_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny_OclAny) lemma cp_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_OclAny : "(cp (p)) \ (cp ((\x. (((p ((x::\Person)))::\OclAny) .oclIsKindOf(OclAny)))))" apply(simp only: OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny) apply((rule cpI2[where f = "(or)"], (rule allI)+, rule cp_OclOr)+) apply(simp only: cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_OclAny) by(simp only: cp_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person_OclAny) lemma cp_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet_OclAny : "(cp (p)) \ (cp ((\x. (((p ((x::\Planet)))::\OclAny) .oclIsKindOf(OclAny)))))" apply(simp only: OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny) apply((rule cpI2[where f = "(or)"], (rule allI)+, rule cp_OclOr)+) apply(simp only: cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet_OclAny) by(simp only: cp_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet_OclAny) lemma cp_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy_OclAny : "(cp (p)) \ (cp ((\x. (((p ((x::\Galaxy)))::\OclAny) .oclIsKindOf(OclAny)))))" apply(simp only: OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny) apply((rule cpI2[where f = "(or)"], (rule allI)+, rule cp_OclOr)+) apply(simp only: cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy_OclAny) by(simp only: cp_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy_OclAny) lemma cp_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_Person : "(cp (p)) \ (cp ((\x. (((p ((x::\OclAny)))::\Person) .oclIsKindOf(OclAny)))))" apply(simp only: OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person) apply((rule cpI2[where f = "(or)"], (rule allI)+, rule cp_OclOr)+) apply(simp only: cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_Person) by(simp only: cp_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny_Person) lemma cp_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_Person : "(cp (p)) \ (cp ((\x. (((p ((x::\Person)))::\Person) .oclIsKindOf(OclAny)))))" apply(simp only: OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person) apply((rule cpI2[where f = "(or)"], (rule allI)+, rule cp_OclOr)+) apply(simp only: cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_Person) by(simp only: cp_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person_Person) lemma cp_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet_Person : "(cp (p)) \ (cp ((\x. (((p ((x::\Planet)))::\Person) .oclIsKindOf(OclAny)))))" apply(simp only: OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person) apply((rule cpI2[where f = "(or)"], (rule allI)+, rule cp_OclOr)+) apply(simp only: cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet_Person) by(simp only: cp_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet_Person) lemma cp_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy_Person : "(cp (p)) \ (cp ((\x. (((p ((x::\Galaxy)))::\Person) .oclIsKindOf(OclAny)))))" apply(simp only: OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person) apply((rule cpI2[where f = "(or)"], (rule allI)+, rule cp_OclOr)+) apply(simp only: cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy_Person) by(simp only: cp_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy_Person) lemma cp_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_Planet : "(cp (p)) \ (cp ((\x. (((p ((x::\OclAny)))::\Planet) .oclIsKindOf(OclAny)))))" apply(simp only: OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet) apply((rule cpI2[where f = "(or)"], (rule allI)+, rule cp_OclOr)+) apply(simp only: cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_Planet) by(simp only: cp_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny_Planet) lemma cp_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_Planet : "(cp (p)) \ (cp ((\x. (((p ((x::\Person)))::\Planet) .oclIsKindOf(OclAny)))))" apply(simp only: OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet) apply((rule cpI2[where f = "(or)"], (rule allI)+, rule cp_OclOr)+) apply(simp only: cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_Planet) by(simp only: cp_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person_Planet) lemma cp_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet_Planet : "(cp (p)) \ (cp ((\x. (((p ((x::\Planet)))::\Planet) .oclIsKindOf(OclAny)))))" apply(simp only: OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet) apply((rule cpI2[where f = "(or)"], (rule allI)+, rule cp_OclOr)+) apply(simp only: cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet_Planet) by(simp only: cp_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet_Planet) lemma cp_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy_Planet : "(cp (p)) \ (cp ((\x. (((p ((x::\Galaxy)))::\Planet) .oclIsKindOf(OclAny)))))" apply(simp only: OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet) apply((rule cpI2[where f = "(or)"], (rule allI)+, rule cp_OclOr)+) apply(simp only: cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy_Planet) by(simp only: cp_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy_Planet) lemma cp_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_Galaxy : "(cp (p)) \ (cp ((\x. (((p ((x::\OclAny)))::\Galaxy) .oclIsKindOf(OclAny)))))" apply(simp only: OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy) apply((rule cpI2[where f = "(or)"], (rule allI)+, rule cp_OclOr)+) apply(simp only: cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_Galaxy) by(simp only: cp_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny_Galaxy) lemma cp_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_Galaxy : "(cp (p)) \ (cp ((\x. (((p ((x::\Person)))::\Galaxy) .oclIsKindOf(OclAny)))))" apply(simp only: OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy) apply((rule cpI2[where f = "(or)"], (rule allI)+, rule cp_OclOr)+) apply(simp only: cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_Galaxy) by(simp only: cp_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person_Galaxy) lemma cp_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet_Galaxy : "(cp (p)) \ (cp ((\x. (((p ((x::\Planet)))::\Galaxy) .oclIsKindOf(OclAny)))))" apply(simp only: OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy) apply((rule cpI2[where f = "(or)"], (rule allI)+, rule cp_OclOr)+) apply(simp only: cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet_Galaxy) by(simp only: cp_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet_Galaxy) lemma cp_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy_Galaxy : "(cp (p)) \ (cp ((\x. (((p ((x::\Galaxy)))::\Galaxy) .oclIsKindOf(OclAny)))))" apply(simp only: OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy) apply((rule cpI2[where f = "(or)"], (rule allI)+, rule cp_OclOr)+) apply(simp only: cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy_Galaxy) by(simp only: cp_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy_Galaxy) (* 83 ************************************ 525 + 1 *) (* term Floor1_iskindof.print_iskindof_lemmas_cp *) lemmas[simp,code_unfold] = cp_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_OclAny cp_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy_OclAny cp_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet_OclAny cp_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_OclAny cp_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_Galaxy cp_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy_Galaxy cp_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet_Galaxy cp_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_Galaxy cp_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_Planet cp_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy_Planet cp_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet_Planet cp_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_Planet cp_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_Person cp_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy_Person cp_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet_Person cp_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_Person cp_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny_OclAny cp_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy_OclAny cp_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet_OclAny cp_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person_OclAny cp_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny_Galaxy cp_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy_Galaxy cp_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet_Galaxy cp_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person_Galaxy cp_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny_Planet cp_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy_Planet cp_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet_Planet cp_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person_Planet cp_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny_Person cp_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy_Person cp_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet_Person cp_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person_Person cp_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny_OclAny cp_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy_OclAny cp_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet_OclAny cp_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person_OclAny cp_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny_Galaxy cp_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy_Galaxy cp_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet_Galaxy cp_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person_Galaxy cp_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny_Planet cp_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy_Planet cp_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet_Planet cp_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person_Planet cp_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny_Person cp_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy_Person cp_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet_Person cp_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person_Person cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_OclAny cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy_OclAny cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet_OclAny cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_OclAny cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_Galaxy cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy_Galaxy cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet_Galaxy cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_Galaxy cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_Planet cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy_Planet cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet_Planet cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_Planet cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_Person cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy_Person cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet_Person cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_Person (* 84 ************************************ 526 + 1 *) subsection \Execution with Invalid or Null as Argument\ (* 85 ************************************ 527 + 32 *) (* term Floor1_iskindof.print_iskindof_lemma_strict *) lemma OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_invalid : "((invalid::\Person) .oclIsKindOf(Person)) = invalid" by(simp only: OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_invalid) lemma OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_null : "((null::\Person) .oclIsKindOf(Person)) = true" by(simp only: OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_null) lemma OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet_invalid : "((invalid::\Planet) .oclIsKindOf(Person)) = invalid" by(simp only: OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet_invalid) lemma OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet_null : "((null::\Planet) .oclIsKindOf(Person)) = true" by(simp only: OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet_null) lemma OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy_invalid : "((invalid::\Galaxy) .oclIsKindOf(Person)) = invalid" by(simp only: OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy_invalid) lemma OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy_null : "((null::\Galaxy) .oclIsKindOf(Person)) = true" by(simp only: OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy_null) lemma OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_invalid : "((invalid::\OclAny) .oclIsKindOf(Person)) = invalid" by(simp only: OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_invalid) lemma OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_null : "((null::\OclAny) .oclIsKindOf(Person)) = true" by(simp only: OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_null) lemma OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet_invalid : "((invalid::\Planet) .oclIsKindOf(Planet)) = invalid" by(simp only: OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet_invalid OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet_invalid, simp) lemma OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet_null : "((null::\Planet) .oclIsKindOf(Planet)) = true" by(simp only: OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet_null OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet_null, simp) lemma OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy_invalid : "((invalid::\Galaxy) .oclIsKindOf(Planet)) = invalid" by(simp only: OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy_invalid OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy_invalid, simp) lemma OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy_null : "((null::\Galaxy) .oclIsKindOf(Planet)) = true" by(simp only: OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy_null OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy_null, simp) lemma OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny_invalid : "((invalid::\OclAny) .oclIsKindOf(Planet)) = invalid" by(simp only: OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny_invalid OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_invalid, simp) lemma OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny_null : "((null::\OclAny) .oclIsKindOf(Planet)) = true" by(simp only: OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny_null OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_null, simp) lemma OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person_invalid : "((invalid::\Person) .oclIsKindOf(Planet)) = invalid" by(simp only: OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person_invalid OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_invalid, simp) lemma OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person_null : "((null::\Person) .oclIsKindOf(Planet)) = true" by(simp only: OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person_null OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_null, simp) lemma OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy_invalid : "((invalid::\Galaxy) .oclIsKindOf(Galaxy)) = invalid" by(simp only: OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy_invalid OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy_invalid, simp) lemma OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy_null : "((null::\Galaxy) .oclIsKindOf(Galaxy)) = true" by(simp only: OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy_null OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy_null, simp) lemma OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny_invalid : "((invalid::\OclAny) .oclIsKindOf(Galaxy)) = invalid" by(simp only: OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny_invalid OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny_invalid, simp) lemma OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny_null : "((null::\OclAny) .oclIsKindOf(Galaxy)) = true" by(simp only: OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny_null OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny_null, simp) lemma OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person_invalid : "((invalid::\Person) .oclIsKindOf(Galaxy)) = invalid" by(simp only: OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person_invalid OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person_invalid, simp) lemma OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person_null : "((null::\Person) .oclIsKindOf(Galaxy)) = true" by(simp only: OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person_null OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person_null, simp) lemma OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet_invalid : "((invalid::\Planet) .oclIsKindOf(Galaxy)) = invalid" by(simp only: OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet_invalid OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet_invalid, simp) lemma OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet_null : "((null::\Planet) .oclIsKindOf(Galaxy)) = true" by(simp only: OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet_null OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet_null, simp) lemma OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_invalid : "((invalid::\OclAny) .oclIsKindOf(OclAny)) = invalid" by(simp only: OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_invalid OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny_invalid, simp) lemma OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_null : "((null::\OclAny) .oclIsKindOf(OclAny)) = true" by(simp only: OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_null OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny_null, simp) lemma OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_invalid : "((invalid::\Person) .oclIsKindOf(OclAny)) = invalid" by(simp only: OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_invalid OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person_invalid, simp) lemma OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_null : "((null::\Person) .oclIsKindOf(OclAny)) = true" by(simp only: OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_null OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person_null, simp) lemma OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet_invalid : "((invalid::\Planet) .oclIsKindOf(OclAny)) = invalid" by(simp only: OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet_invalid OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet_invalid, simp) lemma OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet_null : "((null::\Planet) .oclIsKindOf(OclAny)) = true" by(simp only: OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet_null OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet_null, simp) lemma OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy_invalid : "((invalid::\Galaxy) .oclIsKindOf(OclAny)) = invalid" by(simp only: OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy_invalid OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy_invalid, simp) lemma OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy_null : "((null::\Galaxy) .oclIsKindOf(OclAny)) = true" by(simp only: OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy_null OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy_null, simp) (* 86 ************************************ 559 + 1 *) (* term Floor1_iskindof.print_iskindof_lemmas_strict *) lemmas[simp,code_unfold] = OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_invalid OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy_invalid OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet_invalid OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_invalid OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_null OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy_null OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet_null OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_null OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny_invalid OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy_invalid OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet_invalid OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person_invalid OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny_null OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy_null OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet_null OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person_null OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny_invalid OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy_invalid OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet_invalid OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person_invalid OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny_null OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy_null OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet_null OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person_null OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_invalid OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy_invalid OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet_invalid OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_invalid OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_null OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy_null OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet_null OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_null (* 87 ************************************ 560 + 1 *) subsection \Validity and Definedness Properties\ (* 88 ************************************ 561 + 16 *) (* term Floor1_iskindof.print_iskindof_defined *) lemma OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_defined : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\Person) .oclIsKindOf(Person)))" by(simp only: OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person, rule OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_defined[OF isdef]) lemma OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet_defined : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\Planet) .oclIsKindOf(Person)))" by(simp only: OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet, rule OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet_defined[OF isdef]) lemma OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy_defined : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\Galaxy) .oclIsKindOf(Person)))" by(simp only: OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy, rule OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy_defined[OF isdef]) lemma OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_defined : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\OclAny) .oclIsKindOf(Person)))" by(simp only: OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny, rule OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_defined[OF isdef]) lemma OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet_defined : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\Planet) .oclIsKindOf(Planet)))" by(simp only: OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet, rule defined_or_I[OF OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet_defined[OF isdef], OF OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet_defined[OF isdef]]) lemma OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy_defined : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\Galaxy) .oclIsKindOf(Planet)))" by(simp only: OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy, rule defined_or_I[OF OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy_defined[OF isdef], OF OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy_defined[OF isdef]]) lemma OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny_defined : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\OclAny) .oclIsKindOf(Planet)))" by(simp only: OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny, rule defined_or_I[OF OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny_defined[OF isdef], OF OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_defined[OF isdef]]) lemma OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person_defined : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\Person) .oclIsKindOf(Planet)))" by(simp only: OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person, rule defined_or_I[OF OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person_defined[OF isdef], OF OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_defined[OF isdef]]) lemma OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy_defined : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\Galaxy) .oclIsKindOf(Galaxy)))" by(simp only: OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy, rule defined_or_I[OF OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy_defined[OF isdef], OF OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy_defined[OF isdef]]) lemma OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny_defined : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\OclAny) .oclIsKindOf(Galaxy)))" by(simp only: OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny, rule defined_or_I[OF OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny_defined[OF isdef], OF OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny_defined[OF isdef]]) lemma OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person_defined : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\Person) .oclIsKindOf(Galaxy)))" by(simp only: OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person, rule defined_or_I[OF OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person_defined[OF isdef], OF OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person_defined[OF isdef]]) lemma OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet_defined : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\Planet) .oclIsKindOf(Galaxy)))" by(simp only: OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet, rule defined_or_I[OF OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet_defined[OF isdef], OF OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet_defined[OF isdef]]) lemma OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_defined : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\OclAny) .oclIsKindOf(OclAny)))" by(simp only: OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny, rule defined_or_I[OF OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_defined[OF isdef], OF OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny_defined[OF isdef]]) lemma OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_defined : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\Person) .oclIsKindOf(OclAny)))" by(simp only: OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person, rule defined_or_I[OF OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_defined[OF isdef], OF OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person_defined[OF isdef]]) lemma OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet_defined : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\Planet) .oclIsKindOf(OclAny)))" by(simp only: OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet, rule defined_or_I[OF OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet_defined[OF isdef], OF OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet_defined[OF isdef]]) lemma OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy_defined : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\Galaxy) .oclIsKindOf(OclAny)))" by(simp only: OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy, rule defined_or_I[OF OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy_defined[OF isdef], OF OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy_defined[OF isdef]]) (* 89 ************************************ 577 + 16 *) (* term Floor1_iskindof.print_iskindof_defined' *) lemma OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_defined' : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\Person) .oclIsKindOf(Person)))" by(rule OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_defined[OF isdef[THEN foundation20]]) lemma OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet_defined' : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\Planet) .oclIsKindOf(Person)))" by(rule OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet_defined[OF isdef[THEN foundation20]]) lemma OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy_defined' : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\Galaxy) .oclIsKindOf(Person)))" by(rule OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy_defined[OF isdef[THEN foundation20]]) lemma OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_defined' : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\OclAny) .oclIsKindOf(Person)))" by(rule OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_defined[OF isdef[THEN foundation20]]) lemma OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet_defined' : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\Planet) .oclIsKindOf(Planet)))" by(rule OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet_defined[OF isdef[THEN foundation20]]) lemma OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy_defined' : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\Galaxy) .oclIsKindOf(Planet)))" by(rule OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy_defined[OF isdef[THEN foundation20]]) lemma OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny_defined' : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\OclAny) .oclIsKindOf(Planet)))" by(rule OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny_defined[OF isdef[THEN foundation20]]) lemma OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person_defined' : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\Person) .oclIsKindOf(Planet)))" by(rule OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person_defined[OF isdef[THEN foundation20]]) lemma OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy_defined' : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\Galaxy) .oclIsKindOf(Galaxy)))" by(rule OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy_defined[OF isdef[THEN foundation20]]) lemma OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny_defined' : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\OclAny) .oclIsKindOf(Galaxy)))" by(rule OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny_defined[OF isdef[THEN foundation20]]) lemma OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person_defined' : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\Person) .oclIsKindOf(Galaxy)))" by(rule OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person_defined[OF isdef[THEN foundation20]]) lemma OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet_defined' : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\Planet) .oclIsKindOf(Galaxy)))" by(rule OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet_defined[OF isdef[THEN foundation20]]) lemma OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_defined' : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\OclAny) .oclIsKindOf(OclAny)))" by(rule OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_defined[OF isdef[THEN foundation20]]) lemma OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_defined' : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\Person) .oclIsKindOf(OclAny)))" by(rule OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_defined[OF isdef[THEN foundation20]]) lemma OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet_defined' : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\Planet) .oclIsKindOf(OclAny)))" by(rule OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet_defined[OF isdef[THEN foundation20]]) lemma OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy_defined' : assumes isdef: "\ \ (\ (X))" shows "\ \ (\ ((X::\Galaxy) .oclIsKindOf(OclAny)))" by(rule OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy_defined[OF isdef[THEN foundation20]]) (* 90 ************************************ 593 + 1 *) subsection \Up Down Casting\ (* 91 ************************************ 594 + 4 *) (* term Floor1_iskindof.print_iskindof_up_eq_asty *) lemma actual_eq_static\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n : assumes isdef: "\ \ (\ (X))" shows "\ \ ((X::\Person) .oclIsKindOf(Person))" apply(simp only: OclValid_def, insert isdef) apply(simp only: OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person) apply(auto simp: foundation16 bot_option_def split: option.split ty\\\\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n.split ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n.split) by((simp_all add: false_def true_def OclOr_def OclAnd_def OclNot_def)?) lemma actual_eq_static\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t : assumes isdef: "\ \ (\ (X))" shows "\ \ ((X::\Planet) .oclIsKindOf(Planet))" apply(simp only: OclValid_def, insert isdef) apply(simp only: OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet, subst (1) cp_OclOr, simp only: OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet) apply(auto simp: cp_OclOr[symmetric ] foundation16 bot_option_def OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet split: option.split ty\\\\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t.split ty\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t.split ty\\\\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n.split ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n.split) by((simp_all add: false_def true_def OclOr_def OclAnd_def OclNot_def)?) lemma actual_eq_static\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y : assumes isdef: "\ \ (\ (X))" shows "\ \ ((X::\Galaxy) .oclIsKindOf(Galaxy))" apply(simp only: OclValid_def, insert isdef) apply(simp only: OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy, subst (1) cp_OclOr, simp only: OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy, subst (2 1) cp_OclOr, simp only: OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy) apply(auto simp: cp_OclOr[symmetric ] foundation16 bot_option_def OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy split: option.split ty\\\\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y.split ty\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y.split ty\\\\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n.split ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n.split ty\\\\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t.split ty\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t.split) by((simp_all add: false_def true_def OclOr_def OclAnd_def OclNot_def)?) lemma actual_eq_static\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y : assumes isdef: "\ \ (\ (X))" shows "\ \ ((X::\OclAny) .oclIsKindOf(OclAny))" apply(simp only: OclValid_def, insert isdef) apply(simp only: OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny, subst (1) cp_OclOr, simp only: OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny, subst (2 1) cp_OclOr, simp only: OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny, subst (3 2 1) cp_OclOr, simp only: OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny) apply(auto simp: cp_OclOr[symmetric ] foundation16 bot_option_def OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny split: option.split ty\\\\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y.split ty\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y.split ty\\\\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n.split ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n.split ty\\\\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t.split ty\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t.split ty\\\\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y.split ty\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y.split) by((simp_all add: false_def true_def OclOr_def OclAnd_def OclNot_def)?) (* 92 ************************************ 598 + 6 *) (* term Floor1_iskindof.print_iskindof_up_larger *) lemma actualKind\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_larger_staticKind\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t : assumes isdef: "\ \ (\ (X))" shows "\ \ ((X::\Person) .oclIsKindOf(Planet))" apply(simp only: OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person) by(rule foundation25', rule actual_eq_static\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n[OF isdef]) lemma actualKind\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_larger_staticKind\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y : assumes isdef: "\ \ (\ (X))" shows "\ \ ((X::\Person) .oclIsKindOf(Galaxy))" apply(simp only: OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person) by(rule foundation25', rule actualKind\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_larger_staticKind\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t[OF isdef]) lemma actualKind\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_larger_staticKind\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y : assumes isdef: "\ \ (\ (X))" shows "\ \ ((X::\Person) .oclIsKindOf(OclAny))" apply(simp only: OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person) by(rule foundation25', rule actualKind\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_larger_staticKind\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y[OF isdef]) lemma actualKind\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_larger_staticKind\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y : assumes isdef: "\ \ (\ (X))" shows "\ \ ((X::\Planet) .oclIsKindOf(Galaxy))" apply(simp only: OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet) by(rule foundation25', rule actual_eq_static\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t[OF isdef]) lemma actualKind\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_larger_staticKind\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y : assumes isdef: "\ \ (\ (X))" shows "\ \ ((X::\Planet) .oclIsKindOf(OclAny))" apply(simp only: OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet) by(rule foundation25', rule actualKind\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_larger_staticKind\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y[OF isdef]) lemma actualKind\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_larger_staticKind\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y : assumes isdef: "\ \ (\ (X))" shows "\ \ ((X::\Galaxy) .oclIsKindOf(OclAny))" apply(simp only: OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy) by(rule foundation25', rule actual_eq_static\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y[OF isdef]) (* 93 ************************************ 604 + 6 *) (* term Floor1_iskindof.print_iskindof_up_istypeof_unfold *) lemma not_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_then_Planet_OclIsTypeOf_others_unfold : assumes isdef: "(\ \ (\ (X)))" assumes iskin: "(\ \ ((X::\Planet) .oclIsKindOf(Person)))" shows "(\ \ ((X::\Planet) .oclIsTypeOf(Person)))" using iskin apply(simp only: OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet) done lemma not_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_then_Galaxy_OclIsTypeOf_others_unfold : assumes isdef: "(\ \ (\ (X)))" assumes iskin: "(\ \ ((X::\Galaxy) .oclIsKindOf(Person)))" shows "(\ \ ((X::\Galaxy) .oclIsTypeOf(Person)))" using iskin apply(simp only: OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy) done lemma not_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_then_OclAny_OclIsTypeOf_others_unfold : assumes isdef: "(\ \ (\ (X)))" assumes iskin: "(\ \ ((X::\OclAny) .oclIsKindOf(Person)))" shows "(\ \ ((X::\OclAny) .oclIsTypeOf(Person)))" using iskin apply(simp only: OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny) done lemma not_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_then_Galaxy_OclIsTypeOf_others_unfold : assumes isdef: "(\ \ (\ (X)))" assumes iskin: "(\ \ ((X::\Galaxy) .oclIsKindOf(Planet)))" shows "((\ \ ((X::\Galaxy) .oclIsTypeOf(Planet))) \ (\ \ ((X::\Galaxy) .oclIsTypeOf(Person))))" using iskin apply(simp only: OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy) apply(erule foundation26[OF OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy_defined'[OF isdef], OF OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy_defined'[OF isdef]]) apply(simp) apply(drule not_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_then_Galaxy_OclIsTypeOf_others_unfold[OF isdef], blast) done lemma not_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_then_OclAny_OclIsTypeOf_others_unfold : assumes isdef: "(\ \ (\ (X)))" assumes iskin: "(\ \ ((X::\OclAny) .oclIsKindOf(Planet)))" shows "((\ \ ((X::\OclAny) .oclIsTypeOf(Planet))) \ (\ \ ((X::\OclAny) .oclIsTypeOf(Person))))" using iskin apply(simp only: OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny) apply(erule foundation26[OF OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny_defined'[OF isdef], OF OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_defined'[OF isdef]]) apply(simp) apply(drule not_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_then_OclAny_OclIsTypeOf_others_unfold[OF isdef], blast) done lemma not_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_then_OclAny_OclIsTypeOf_others_unfold : assumes isdef: "(\ \ (\ (X)))" assumes iskin: "(\ \ ((X::\OclAny) .oclIsKindOf(Galaxy)))" shows "((\ \ ((X::\OclAny) .oclIsTypeOf(Galaxy))) \ ((\ \ ((X::\OclAny) .oclIsTypeOf(Person))) \ (\ \ ((X::\OclAny) .oclIsTypeOf(Planet)))))" using iskin apply(simp only: OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny) apply(erule foundation26[OF OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny_defined'[OF isdef], OF OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny_defined'[OF isdef]]) apply(simp) apply(drule not_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_then_OclAny_OclIsTypeOf_others_unfold[OF isdef], blast) done (* 94 ************************************ 610 + 6 *) (* term Floor1_iskindof.print_iskindof_up_istypeof *) lemma not_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_then_Planet_OclIsTypeOf_others : assumes iskin: "\ \ \ ((X::\Planet) .oclIsKindOf(Person))" assumes isdef: "\ \ (\ (X))" shows "\ \ ((X::\Planet) .oclIsTypeOf(Planet))" using actual_eq_static\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t[OF isdef] apply(simp only: OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet) apply(erule foundation26[OF OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet_defined'[OF isdef], OF OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet_defined'[OF isdef]]) apply(simp) apply(simp add: iskin) done lemma not_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_then_Galaxy_OclIsTypeOf_others : assumes iskin: "\ \ \ ((X::\Galaxy) .oclIsKindOf(Person))" assumes isdef: "\ \ (\ (X))" shows "(\ \ ((X::\Galaxy) .oclIsTypeOf(Galaxy)) \ \ \ ((X::\Galaxy) .oclIsTypeOf(Planet)))" using actual_eq_static\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y[OF isdef] apply(simp only: OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy) apply(erule foundation26[OF OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy_defined'[OF isdef], OF OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy_defined'[OF isdef]]) apply(simp) apply(simp only: OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy) apply(erule foundation26[OF OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy_defined'[OF isdef], OF OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy_defined'[OF isdef]]) apply(simp) apply(simp add: iskin) done lemma not_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_then_OclAny_OclIsTypeOf_others : assumes iskin: "\ \ \ ((X::\OclAny) .oclIsKindOf(Person))" assumes isdef: "\ \ (\ (X))" shows "(\ \ ((X::\OclAny) .oclIsTypeOf(OclAny)) \ (\ \ ((X::\OclAny) .oclIsTypeOf(Galaxy)) \ \ \ ((X::\OclAny) .oclIsTypeOf(Planet))))" using actual_eq_static\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y[OF isdef] apply(simp only: OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny) apply(erule foundation26[OF OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_defined'[OF isdef], OF OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny_defined'[OF isdef]]) apply(simp) apply(simp only: OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny) apply(erule foundation26[OF OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny_defined'[OF isdef], OF OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny_defined'[OF isdef]]) apply(simp) apply(simp only: OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny) apply(erule foundation26[OF OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny_defined'[OF isdef], OF OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_defined'[OF isdef]]) apply(simp) apply(simp add: iskin) done lemma not_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_then_Galaxy_OclIsTypeOf_others : assumes iskin: "\ \ \ ((X::\Galaxy) .oclIsKindOf(Planet))" assumes isdef: "\ \ (\ (X))" shows "\ \ ((X::\Galaxy) .oclIsTypeOf(Galaxy))" using actual_eq_static\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y[OF isdef] apply(simp only: OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy) apply(erule foundation26[OF OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy_defined'[OF isdef], OF OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy_defined'[OF isdef]]) apply(simp) apply(simp add: iskin) done lemma not_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_then_OclAny_OclIsTypeOf_others : assumes iskin: "\ \ \ ((X::\OclAny) .oclIsKindOf(Planet))" assumes isdef: "\ \ (\ (X))" shows "(\ \ ((X::\OclAny) .oclIsTypeOf(OclAny)) \ \ \ ((X::\OclAny) .oclIsTypeOf(Galaxy)))" using actual_eq_static\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y[OF isdef] apply(simp only: OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny) apply(erule foundation26[OF OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_defined'[OF isdef], OF OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny_defined'[OF isdef]]) apply(simp) apply(simp only: OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny) apply(erule foundation26[OF OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny_defined'[OF isdef], OF OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny_defined'[OF isdef]]) apply(simp) apply(simp add: iskin) done lemma not_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_then_OclAny_OclIsTypeOf_others : assumes iskin: "\ \ \ ((X::\OclAny) .oclIsKindOf(Galaxy))" assumes isdef: "\ \ (\ (X))" shows "\ \ ((X::\OclAny) .oclIsTypeOf(OclAny))" using actual_eq_static\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y[OF isdef] apply(simp only: OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny) apply(erule foundation26[OF OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_defined'[OF isdef], OF OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny_defined'[OF isdef]]) apply(simp) apply(simp add: iskin) done (* 95 ************************************ 616 + 10 *) (* term Floor1_iskindof.print_iskindof_up_d_cast *) lemma down_cast_kind\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_from_Planet_to_Person : assumes iskin: "\ \ \ ((X::\Planet) .oclIsKindOf(Person))" assumes isdef: "\ \ (\ (X))" shows "\ \ (X .oclAsType(Person)) \ invalid" apply(insert not_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_then_Planet_OclIsTypeOf_others[OF iskin, OF isdef]) apply(rule down_cast_type\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_from_Planet_to_Person, simp only: , simp only: isdef) done lemma down_cast_kind\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_from_Galaxy_to_Person : assumes iskin: "\ \ \ ((X::\Galaxy) .oclIsKindOf(Person))" assumes isdef: "\ \ (\ (X))" shows "\ \ (X .oclAsType(Person)) \ invalid" apply(insert not_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_then_Galaxy_OclIsTypeOf_others[OF iskin, OF isdef], elim disjE) apply(rule down_cast_type\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_from_Galaxy_to_Person, simp only: , simp only: isdef) apply(rule down_cast_type\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_from_Galaxy_to_Person, simp only: , simp only: isdef) done lemma down_cast_kind\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_from_OclAny_to_Person : assumes iskin: "\ \ \ ((X::\OclAny) .oclIsKindOf(Person))" assumes isdef: "\ \ (\ (X))" shows "\ \ (X .oclAsType(Person)) \ invalid" apply(insert not_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_then_OclAny_OclIsTypeOf_others[OF iskin, OF isdef], elim disjE) apply(rule down_cast_type\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_from_OclAny_to_Person, simp only: , simp only: isdef) apply(rule down_cast_type\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_from_OclAny_to_Person, simp only: , simp only: isdef) apply(rule down_cast_type\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_from_OclAny_to_Person, simp only: , simp only: isdef) done lemma down_cast_kind\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_from_Galaxy_to_Planet : assumes iskin: "\ \ \ ((X::\Galaxy) .oclIsKindOf(Planet))" assumes isdef: "\ \ (\ (X))" shows "\ \ (X .oclAsType(Planet)) \ invalid" apply(insert not_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_then_Galaxy_OclIsTypeOf_others[OF iskin, OF isdef]) apply(rule down_cast_type\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_from_Galaxy_to_Planet, simp only: , simp only: isdef) done lemma down_cast_kind\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_from_Galaxy_to_Person : assumes iskin: "\ \ \ ((X::\Galaxy) .oclIsKindOf(Planet))" assumes isdef: "\ \ (\ (X))" shows "\ \ (X .oclAsType(Person)) \ invalid" apply(insert not_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_then_Galaxy_OclIsTypeOf_others[OF iskin, OF isdef]) apply(rule down_cast_type\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_from_Galaxy_to_Person, simp only: , simp only: isdef) done lemma down_cast_kind\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_from_OclAny_to_Planet : assumes iskin: "\ \ \ ((X::\OclAny) .oclIsKindOf(Planet))" assumes isdef: "\ \ (\ (X))" shows "\ \ (X .oclAsType(Planet)) \ invalid" apply(insert not_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_then_OclAny_OclIsTypeOf_others[OF iskin, OF isdef], elim disjE) apply(rule down_cast_type\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_from_OclAny_to_Planet, simp only: , simp only: isdef) apply(rule down_cast_type\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_from_OclAny_to_Planet, simp only: , simp only: isdef) done lemma down_cast_kind\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_from_OclAny_to_Person : assumes iskin: "\ \ \ ((X::\OclAny) .oclIsKindOf(Planet))" assumes isdef: "\ \ (\ (X))" shows "\ \ (X .oclAsType(Person)) \ invalid" apply(insert not_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_then_OclAny_OclIsTypeOf_others[OF iskin, OF isdef], elim disjE) apply(rule down_cast_type\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_from_OclAny_to_Person, simp only: , simp only: isdef) apply(rule down_cast_type\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_from_OclAny_to_Person, simp only: , simp only: isdef) done lemma down_cast_kind\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_from_OclAny_to_Galaxy : assumes iskin: "\ \ \ ((X::\OclAny) .oclIsKindOf(Galaxy))" assumes isdef: "\ \ (\ (X))" shows "\ \ (X .oclAsType(Galaxy)) \ invalid" apply(insert not_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_then_OclAny_OclIsTypeOf_others[OF iskin, OF isdef]) apply(rule down_cast_type\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_from_OclAny_to_Galaxy, simp only: , simp only: isdef) done lemma down_cast_kind\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_from_OclAny_to_Person : assumes iskin: "\ \ \ ((X::\OclAny) .oclIsKindOf(Galaxy))" assumes isdef: "\ \ (\ (X))" shows "\ \ (X .oclAsType(Person)) \ invalid" apply(insert not_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_then_OclAny_OclIsTypeOf_others[OF iskin, OF isdef]) apply(rule down_cast_type\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_from_OclAny_to_Person, simp only: , simp only: isdef) done lemma down_cast_kind\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_from_OclAny_to_Planet : assumes iskin: "\ \ \ ((X::\OclAny) .oclIsKindOf(Galaxy))" assumes isdef: "\ \ (\ (X))" shows "\ \ (X .oclAsType(Planet)) \ invalid" apply(insert not_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_then_OclAny_OclIsTypeOf_others[OF iskin, OF isdef]) apply(rule down_cast_type\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_from_OclAny_to_Planet, simp only: , simp only: isdef) done (* 96 ************************************ 626 + 1 *) subsection \Const\ (* 97 ************************************ 627 + 1 *) section \Class Model: OclAllInstances\ (* 98 ************************************ 628 + 1 *) text \ To denote \OCL-types occurring in \OCL expressions syntactically---as, for example, as ``argument'' of \inlineisar{oclAllInstances()}---we use the inverses of the injection functions into the object universes; we show that this is sufficient ``characterization.'' \ (* 99 ************************************ 629 + 4 *) (* term Floor1_allinst.print_allinst_def_id *) definition "Person = OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_\" definition "Planet = OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_\" definition "Galaxy = OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_\" definition "OclAny = OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\" (* 100 ************************************ 633 + 1 *) (* term Floor1_allinst.print_allinst_lemmas_id *) lemmas[simp,code_unfold] = Person_def Planet_def Galaxy_def OclAny_def (* 101 ************************************ 634 + 1 *) (* term Floor1_allinst.print_allinst_astype *) lemma OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\_some : "(OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\ (x)) \ None" by(simp add: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\_def) (* 102 ************************************ 635 + 3 *) (* term Floor1_allinst.print_allinst_exec *) lemma OclAllInstances_generic\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_exec : shows "(OclAllInstances_generic (pre_post) (OclAny)) = (\\. (Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (\\Some ` OclAny ` (ran ((heap ((pre_post (\))))))\\)))" proof - let ?S1 = "(\\. OclAny ` (ran ((heap ((pre_post (\)))))))" show ?thesis proof - let ?S2 = "(\\. ((?S1) (\)) - {None})" show ?thesis proof - have B: "(\\. ((?S2) (\)) \ ((?S1) (\)))" by(auto) show ?thesis proof - have C: "(\\. ((?S1) (\)) \ ((?S2) (\)))" by(auto simp: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\_some) show ?thesis apply(simp add: OclValid_def del: OclAllInstances_generic_def OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny) by(insert equalityI[OF B, OF C], simp) qed qed qed qed lemma OclAllInstances_at_post\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_exec : shows "(OclAllInstances_at_post (OclAny)) = (\\. (Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (\\Some ` OclAny ` (ran ((heap ((snd (\))))))\\)))" unfolding OclAllInstances_at_post_def by(rule OclAllInstances_generic\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_exec) lemma OclAllInstances_at_pre\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_exec : shows "(OclAllInstances_at_pre (OclAny)) = (\\. (Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e (\\Some ` OclAny ` (ran ((heap ((fst (\))))))\\)))" unfolding OclAllInstances_at_pre_def by(rule OclAllInstances_generic\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_exec) (* 103 ************************************ 638 + 1 *) subsection \OclIsTypeOf\ (* 104 ************************************ 639 + 2 *) (* term Floor1_allinst.print_allinst_istypeof_pre *) lemma ex_ssubst : "(\x \ B. (s (x)) = (t (x))) \ (\x \ B. (P ((s (x))))) = (\x \ B. (P ((t (x)))))" by(simp) lemma ex_def : "x \ \\\\Some ` (X - {None})\\\\ \ (\y. x = \\y\\)" by(auto) (* 105 ************************************ 641 + 21 *) (* term Floor1_allinst.print_allinst_istypeof *) lemma Person_OclAllInstances_generic_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n : "\ \ (UML_Set.OclForall ((OclAllInstances_generic (pre_post) (Person))) (OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))" apply(simp add: OclValid_def del: OclAllInstances_generic_def) apply(simp only: UML_Set.OclForall_def refl if_True OclAllInstances_generic_defined[simplified OclValid_def]) apply(simp only: OclAllInstances_generic_def) apply(subst (1 2 3) Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse, simp add: bot_option_def) apply(subst (1 2 3) ex_ssubst[where s = "(\x. (((\_. x) .oclIsTypeOf(Person)) (\)))" and t = "(\_. (true (\)))"]) apply(intro ballI actual_eq_static\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n[simplified OclValid_def, simplified OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person]) apply(drule ex_def, erule exE, simp) by(simp) lemma Person_OclAllInstances_at_post_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n : shows "\ \ (UML_Set.OclForall ((OclAllInstances_at_post (Person))) (OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))" unfolding OclAllInstances_at_post_def by(rule Person_OclAllInstances_generic_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n) lemma Person_OclAllInstances_at_pre_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n : shows "\ \ (UML_Set.OclForall ((OclAllInstances_at_pre (Person))) (OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))" unfolding OclAllInstances_at_pre_def by(rule Person_OclAllInstances_generic_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n) lemma Planet_OclAllInstances_generic_OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t1 : assumes [simp]: "(\x. (pre_post ((x , x))) = x)" shows "(\\. \ \ (UML_Set.OclForall ((OclAllInstances_generic (pre_post) (Planet))) (OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t)))" apply(rule exI[where x = "\\<^sub>0"], simp add: \\<^sub>0_def OclValid_def del: OclAllInstances_generic_def) apply(simp only: UML_Set.OclForall_def refl if_True OclAllInstances_generic_defined[simplified OclValid_def]) apply(simp only: OclAllInstances_generic_def) apply(subst (1 2 3) Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse, simp add: bot_option_def) by(simp) lemma Planet_OclAllInstances_at_post_OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t1 : shows "(\\. \ \ (UML_Set.OclForall ((OclAllInstances_at_post (Planet))) (OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t)))" unfolding OclAllInstances_at_post_def by(rule Planet_OclAllInstances_generic_OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t1, simp) lemma Planet_OclAllInstances_at_pre_OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t1 : shows "(\\. \ \ (UML_Set.OclForall ((OclAllInstances_at_pre (Planet))) (OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t)))" unfolding OclAllInstances_at_pre_def by(rule Planet_OclAllInstances_generic_OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t1, simp) lemma Planet_OclAllInstances_generic_OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t2 : assumes [simp]: "(\x. (pre_post ((x , x))) = x)" shows "(\\. \ \ (not ((UML_Set.OclForall ((OclAllInstances_generic (pre_post) (Planet))) (OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t)))))" proof - fix oid a show ?thesis proof - let ?t0 = "(state.make ((Map.empty (oid \ (in\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t ((mk\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t ((mk\\\\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (a))) (None) (None))))))) (Map.empty))" show ?thesis apply(rule exI[where x = "(?t0 , ?t0)"], simp add: OclValid_def del: OclAllInstances_generic_def) apply(simp only: UML_Set.OclForall_def refl if_True OclAllInstances_generic_defined[simplified OclValid_def]) apply(simp only: OclAllInstances_generic_def OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_\_def) apply(subst (1 2 3) Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse, simp add: bot_option_def) by(simp add: state.make_def OclNot_def) qed qed lemma Planet_OclAllInstances_at_post_OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t2 : shows "(\\. \ \ (not ((UML_Set.OclForall ((OclAllInstances_at_post (Planet))) (OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t)))))" unfolding OclAllInstances_at_post_def by(rule Planet_OclAllInstances_generic_OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t2, simp) lemma Planet_OclAllInstances_at_pre_OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t2 : shows "(\\. \ \ (not ((UML_Set.OclForall ((OclAllInstances_at_pre (Planet))) (OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t)))))" unfolding OclAllInstances_at_pre_def by(rule Planet_OclAllInstances_generic_OclIsTypeOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t2, simp) lemma Galaxy_OclAllInstances_generic_OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y1 : assumes [simp]: "(\x. (pre_post ((x , x))) = x)" shows "(\\. \ \ (UML_Set.OclForall ((OclAllInstances_generic (pre_post) (Galaxy))) (OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y)))" apply(rule exI[where x = "\\<^sub>0"], simp add: \\<^sub>0_def OclValid_def del: OclAllInstances_generic_def) apply(simp only: UML_Set.OclForall_def refl if_True OclAllInstances_generic_defined[simplified OclValid_def]) apply(simp only: OclAllInstances_generic_def) apply(subst (1 2 3) Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse, simp add: bot_option_def) by(simp) lemma Galaxy_OclAllInstances_at_post_OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y1 : shows "(\\. \ \ (UML_Set.OclForall ((OclAllInstances_at_post (Galaxy))) (OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y)))" unfolding OclAllInstances_at_post_def by(rule Galaxy_OclAllInstances_generic_OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y1, simp) lemma Galaxy_OclAllInstances_at_pre_OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y1 : shows "(\\. \ \ (UML_Set.OclForall ((OclAllInstances_at_pre (Galaxy))) (OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y)))" unfolding OclAllInstances_at_pre_def by(rule Galaxy_OclAllInstances_generic_OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y1, simp) lemma Galaxy_OclAllInstances_generic_OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y2 : assumes [simp]: "(\x. (pre_post ((x , x))) = x)" shows "(\\. \ \ (not ((UML_Set.OclForall ((OclAllInstances_generic (pre_post) (Galaxy))) (OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y)))))" proof - fix oid a show ?thesis proof - let ?t0 = "(state.make ((Map.empty (oid \ (in\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y ((mk\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y ((mk\\\\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (a))) (None) (None) (None))))))) (Map.empty))" show ?thesis apply(rule exI[where x = "(?t0 , ?t0)"], simp add: OclValid_def del: OclAllInstances_generic_def) apply(simp only: UML_Set.OclForall_def refl if_True OclAllInstances_generic_defined[simplified OclValid_def]) apply(simp only: OclAllInstances_generic_def OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_\_def) apply(subst (1 2 3) Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse, simp add: bot_option_def) by(simp add: state.make_def OclNot_def) qed qed lemma Galaxy_OclAllInstances_at_post_OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y2 : shows "(\\. \ \ (not ((UML_Set.OclForall ((OclAllInstances_at_post (Galaxy))) (OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y)))))" unfolding OclAllInstances_at_post_def by(rule Galaxy_OclAllInstances_generic_OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y2, simp) lemma Galaxy_OclAllInstances_at_pre_OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y2 : shows "(\\. \ \ (not ((UML_Set.OclForall ((OclAllInstances_at_pre (Galaxy))) (OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y)))))" unfolding OclAllInstances_at_pre_def by(rule Galaxy_OclAllInstances_generic_OclIsTypeOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y2, simp) lemma OclAny_OclAllInstances_generic_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y1 : assumes [simp]: "(\x. (pre_post ((x , x))) = x)" shows "(\\. \ \ (UML_Set.OclForall ((OclAllInstances_generic (pre_post) (OclAny))) (OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y)))" apply(rule exI[where x = "\\<^sub>0"], simp add: \\<^sub>0_def OclValid_def del: OclAllInstances_generic_def) apply(simp only: UML_Set.OclForall_def refl if_True OclAllInstances_generic_defined[simplified OclValid_def]) apply(simp only: OclAllInstances_generic_def) apply(subst (1 2 3) Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse, simp add: bot_option_def) by(simp) lemma OclAny_OclAllInstances_at_post_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y1 : shows "(\\. \ \ (UML_Set.OclForall ((OclAllInstances_at_post (OclAny))) (OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y)))" unfolding OclAllInstances_at_post_def by(rule OclAny_OclAllInstances_generic_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y1, simp) lemma OclAny_OclAllInstances_at_pre_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y1 : shows "(\\. \ \ (UML_Set.OclForall ((OclAllInstances_at_pre (OclAny))) (OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y)))" unfolding OclAllInstances_at_pre_def by(rule OclAny_OclAllInstances_generic_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y1, simp) lemma OclAny_OclAllInstances_generic_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y2 : assumes [simp]: "(\x. (pre_post ((x , x))) = x)" shows "(\\. \ \ (not ((UML_Set.OclForall ((OclAllInstances_generic (pre_post) (OclAny))) (OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y)))))" proof - fix oid a show ?thesis proof - let ?t0 = "(state.make ((Map.empty (oid \ (in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y ((mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y ((mk\\\\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y (a))))))))) (Map.empty))" show ?thesis apply(rule exI[where x = "(?t0 , ?t0)"], simp add: OclValid_def del: OclAllInstances_generic_def) apply(simp only: UML_Set.OclForall_def refl if_True OclAllInstances_generic_defined[simplified OclValid_def]) apply(simp only: OclAllInstances_generic_def OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\_def) apply(subst (1 2 3) Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse, simp add: bot_option_def) by(simp add: state.make_def OclNot_def) qed qed lemma OclAny_OclAllInstances_at_post_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y2 : shows "(\\. \ \ (not ((UML_Set.OclForall ((OclAllInstances_at_post (OclAny))) (OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y)))))" unfolding OclAllInstances_at_post_def by(rule OclAny_OclAllInstances_generic_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y2, simp) lemma OclAny_OclAllInstances_at_pre_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y2 : shows "(\\. \ \ (not ((UML_Set.OclForall ((OclAllInstances_at_pre (OclAny))) (OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y)))))" unfolding OclAllInstances_at_pre_def by(rule OclAny_OclAllInstances_generic_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y2, simp) (* 106 ************************************ 662 + 1 *) subsection \OclIsKindOf\ (* 107 ************************************ 663 + 12 *) (* term Floor1_allinst.print_allinst_iskindof_eq *) lemma Person_OclAllInstances_generic_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n : "\ \ (UML_Set.OclForall ((OclAllInstances_generic (pre_post) (Person))) (OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))" apply(simp add: OclValid_def del: OclAllInstances_generic_def OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person) apply(simp only: UML_Set.OclForall_def refl if_True OclAllInstances_generic_defined[simplified OclValid_def]) apply(simp only: OclAllInstances_generic_def) apply(subst (1 2 3) Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse, simp add: bot_option_def) apply(subst (1 2 3) ex_ssubst[where s = "(\x. (((\_. x) .oclIsKindOf(Person)) (\)))" and t = "(\_. (true (\)))"]) apply(intro ballI actual_eq_static\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n[simplified OclValid_def]) apply(drule ex_def, erule exE, simp) by(simp) lemma Person_OclAllInstances_at_post_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n : shows "\ \ (UML_Set.OclForall ((OclAllInstances_at_post (Person))) (OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))" unfolding OclAllInstances_at_post_def by(rule Person_OclAllInstances_generic_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n) lemma Person_OclAllInstances_at_pre_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n : shows "\ \ (UML_Set.OclForall ((OclAllInstances_at_pre (Person))) (OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))" unfolding OclAllInstances_at_pre_def by(rule Person_OclAllInstances_generic_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n) lemma Planet_OclAllInstances_generic_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t : "\ \ (UML_Set.OclForall ((OclAllInstances_generic (pre_post) (Planet))) (OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t))" apply(simp add: OclValid_def del: OclAllInstances_generic_def OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Planet) apply(simp only: UML_Set.OclForall_def refl if_True OclAllInstances_generic_defined[simplified OclValid_def]) apply(simp only: OclAllInstances_generic_def) apply(subst (1 2 3) Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse, simp add: bot_option_def) apply(subst (1 2 3) ex_ssubst[where s = "(\x. (((\_. x) .oclIsKindOf(Planet)) (\)))" and t = "(\_. (true (\)))"]) apply(intro ballI actual_eq_static\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t[simplified OclValid_def]) apply(drule ex_def, erule exE, simp) by(simp) lemma Planet_OclAllInstances_at_post_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t : shows "\ \ (UML_Set.OclForall ((OclAllInstances_at_post (Planet))) (OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t))" unfolding OclAllInstances_at_post_def by(rule Planet_OclAllInstances_generic_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t) lemma Planet_OclAllInstances_at_pre_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t : shows "\ \ (UML_Set.OclForall ((OclAllInstances_at_pre (Planet))) (OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t))" unfolding OclAllInstances_at_pre_def by(rule Planet_OclAllInstances_generic_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t) lemma Galaxy_OclAllInstances_generic_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y : "\ \ (UML_Set.OclForall ((OclAllInstances_generic (pre_post) (Galaxy))) (OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y))" apply(simp add: OclValid_def del: OclAllInstances_generic_def OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Galaxy) apply(simp only: UML_Set.OclForall_def refl if_True OclAllInstances_generic_defined[simplified OclValid_def]) apply(simp only: OclAllInstances_generic_def) apply(subst (1 2 3) Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse, simp add: bot_option_def) apply(subst (1 2 3) ex_ssubst[where s = "(\x. (((\_. x) .oclIsKindOf(Galaxy)) (\)))" and t = "(\_. (true (\)))"]) apply(intro ballI actual_eq_static\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y[simplified OclValid_def]) apply(drule ex_def, erule exE, simp) by(simp) lemma Galaxy_OclAllInstances_at_post_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y : shows "\ \ (UML_Set.OclForall ((OclAllInstances_at_post (Galaxy))) (OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y))" unfolding OclAllInstances_at_post_def by(rule Galaxy_OclAllInstances_generic_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y) lemma Galaxy_OclAllInstances_at_pre_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y : shows "\ \ (UML_Set.OclForall ((OclAllInstances_at_pre (Galaxy))) (OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y))" unfolding OclAllInstances_at_pre_def by(rule Galaxy_OclAllInstances_generic_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y) lemma OclAny_OclAllInstances_generic_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y : "\ \ (UML_Set.OclForall ((OclAllInstances_generic (pre_post) (OclAny))) (OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y))" apply(simp add: OclValid_def del: OclAllInstances_generic_def OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny) apply(simp only: UML_Set.OclForall_def refl if_True OclAllInstances_generic_defined[simplified OclValid_def]) apply(simp only: OclAllInstances_generic_def) apply(subst (1 2 3) Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse, simp add: bot_option_def) apply(subst (1 2 3) ex_ssubst[where s = "(\x. (((\_. x) .oclIsKindOf(OclAny)) (\)))" and t = "(\_. (true (\)))"]) apply(intro ballI actual_eq_static\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y[simplified OclValid_def]) apply(drule ex_def, erule exE, simp) by(simp) lemma OclAny_OclAllInstances_at_post_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y : shows "\ \ (UML_Set.OclForall ((OclAllInstances_at_post (OclAny))) (OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y))" unfolding OclAllInstances_at_post_def by(rule OclAny_OclAllInstances_generic_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y) lemma OclAny_OclAllInstances_at_pre_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y : shows "\ \ (UML_Set.OclForall ((OclAllInstances_at_pre (OclAny))) (OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y))" unfolding OclAllInstances_at_pre_def by(rule OclAny_OclAllInstances_generic_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y) (* 108 ************************************ 675 + 18 *) (* term Floor1_allinst.print_allinst_iskindof_larger *) lemma Person_OclAllInstances_generic_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t : "\ \ (UML_Set.OclForall ((OclAllInstances_generic (pre_post) (Person))) (OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t))" apply(simp add: OclValid_def del: OclAllInstances_generic_def OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person) apply(simp only: UML_Set.OclForall_def refl if_True OclAllInstances_generic_defined[simplified OclValid_def]) apply(simp only: OclAllInstances_generic_def) apply(subst (1 2 3) Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse, simp add: bot_option_def) apply(subst (1 2 3) ex_ssubst[where s = "(\x. (((\_. x) .oclIsKindOf(Planet)) (\)))" and t = "(\_. (true (\)))"]) apply(intro ballI actualKind\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_larger_staticKind\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t[simplified OclValid_def]) apply(drule ex_def, erule exE, simp) by(simp) lemma Person_OclAllInstances_at_post_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t : shows "\ \ (UML_Set.OclForall ((OclAllInstances_at_post (Person))) (OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t))" unfolding OclAllInstances_at_post_def by(rule Person_OclAllInstances_generic_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t) lemma Person_OclAllInstances_at_pre_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t : shows "\ \ (UML_Set.OclForall ((OclAllInstances_at_pre (Person))) (OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t))" unfolding OclAllInstances_at_pre_def by(rule Person_OclAllInstances_generic_OclIsKindOf\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t) lemma Person_OclAllInstances_generic_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y : "\ \ (UML_Set.OclForall ((OclAllInstances_generic (pre_post) (Person))) (OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y))" apply(simp add: OclValid_def del: OclAllInstances_generic_def OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person) apply(simp only: UML_Set.OclForall_def refl if_True OclAllInstances_generic_defined[simplified OclValid_def]) apply(simp only: OclAllInstances_generic_def) apply(subst (1 2 3) Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse, simp add: bot_option_def) apply(subst (1 2 3) ex_ssubst[where s = "(\x. (((\_. x) .oclIsKindOf(Galaxy)) (\)))" and t = "(\_. (true (\)))"]) apply(intro ballI actualKind\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_larger_staticKind\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y[simplified OclValid_def]) apply(drule ex_def, erule exE, simp) by(simp) lemma Person_OclAllInstances_at_post_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y : shows "\ \ (UML_Set.OclForall ((OclAllInstances_at_post (Person))) (OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y))" unfolding OclAllInstances_at_post_def by(rule Person_OclAllInstances_generic_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y) lemma Person_OclAllInstances_at_pre_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y : shows "\ \ (UML_Set.OclForall ((OclAllInstances_at_pre (Person))) (OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y))" unfolding OclAllInstances_at_pre_def by(rule Person_OclAllInstances_generic_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y) lemma Person_OclAllInstances_generic_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y : "\ \ (UML_Set.OclForall ((OclAllInstances_generic (pre_post) (Person))) (OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y))" apply(simp add: OclValid_def del: OclAllInstances_generic_def OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person) apply(simp only: UML_Set.OclForall_def refl if_True OclAllInstances_generic_defined[simplified OclValid_def]) apply(simp only: OclAllInstances_generic_def) apply(subst (1 2 3) Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse, simp add: bot_option_def) apply(subst (1 2 3) ex_ssubst[where s = "(\x. (((\_. x) .oclIsKindOf(OclAny)) (\)))" and t = "(\_. (true (\)))"]) apply(intro ballI actualKind\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_larger_staticKind\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y[simplified OclValid_def]) apply(drule ex_def, erule exE, simp) by(simp) lemma Person_OclAllInstances_at_post_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y : shows "\ \ (UML_Set.OclForall ((OclAllInstances_at_post (Person))) (OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y))" unfolding OclAllInstances_at_post_def by(rule Person_OclAllInstances_generic_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y) lemma Person_OclAllInstances_at_pre_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y : shows "\ \ (UML_Set.OclForall ((OclAllInstances_at_pre (Person))) (OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y))" unfolding OclAllInstances_at_pre_def by(rule Person_OclAllInstances_generic_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y) lemma Planet_OclAllInstances_generic_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y : "\ \ (UML_Set.OclForall ((OclAllInstances_generic (pre_post) (Planet))) (OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y))" apply(simp add: OclValid_def del: OclAllInstances_generic_def OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet) apply(simp only: UML_Set.OclForall_def refl if_True OclAllInstances_generic_defined[simplified OclValid_def]) apply(simp only: OclAllInstances_generic_def) apply(subst (1 2 3) Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse, simp add: bot_option_def) apply(subst (1 2 3) ex_ssubst[where s = "(\x. (((\_. x) .oclIsKindOf(Galaxy)) (\)))" and t = "(\_. (true (\)))"]) apply(intro ballI actualKind\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_larger_staticKind\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y[simplified OclValid_def]) apply(drule ex_def, erule exE, simp) by(simp) lemma Planet_OclAllInstances_at_post_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y : shows "\ \ (UML_Set.OclForall ((OclAllInstances_at_post (Planet))) (OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y))" unfolding OclAllInstances_at_post_def by(rule Planet_OclAllInstances_generic_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y) lemma Planet_OclAllInstances_at_pre_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y : shows "\ \ (UML_Set.OclForall ((OclAllInstances_at_pre (Planet))) (OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y))" unfolding OclAllInstances_at_pre_def by(rule Planet_OclAllInstances_generic_OclIsKindOf\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y) lemma Planet_OclAllInstances_generic_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y : "\ \ (UML_Set.OclForall ((OclAllInstances_generic (pre_post) (Planet))) (OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y))" apply(simp add: OclValid_def del: OclAllInstances_generic_def OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet) apply(simp only: UML_Set.OclForall_def refl if_True OclAllInstances_generic_defined[simplified OclValid_def]) apply(simp only: OclAllInstances_generic_def) apply(subst (1 2 3) Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse, simp add: bot_option_def) apply(subst (1 2 3) ex_ssubst[where s = "(\x. (((\_. x) .oclIsKindOf(OclAny)) (\)))" and t = "(\_. (true (\)))"]) apply(intro ballI actualKind\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_larger_staticKind\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y[simplified OclValid_def]) apply(drule ex_def, erule exE, simp) by(simp) lemma Planet_OclAllInstances_at_post_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y : shows "\ \ (UML_Set.OclForall ((OclAllInstances_at_post (Planet))) (OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y))" unfolding OclAllInstances_at_post_def by(rule Planet_OclAllInstances_generic_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y) lemma Planet_OclAllInstances_at_pre_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y : shows "\ \ (UML_Set.OclForall ((OclAllInstances_at_pre (Planet))) (OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y))" unfolding OclAllInstances_at_pre_def by(rule Planet_OclAllInstances_generic_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y) lemma Galaxy_OclAllInstances_generic_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y : "\ \ (UML_Set.OclForall ((OclAllInstances_generic (pre_post) (Galaxy))) (OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y))" apply(simp add: OclValid_def del: OclAllInstances_generic_def OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy) apply(simp only: UML_Set.OclForall_def refl if_True OclAllInstances_generic_defined[simplified OclValid_def]) apply(simp only: OclAllInstances_generic_def) apply(subst (1 2 3) Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse, simp add: bot_option_def) apply(subst (1 2 3) ex_ssubst[where s = "(\x. (((\_. x) .oclIsKindOf(OclAny)) (\)))" and t = "(\_. (true (\)))"]) apply(intro ballI actualKind\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_larger_staticKind\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y[simplified OclValid_def]) apply(drule ex_def, erule exE, simp) by(simp) lemma Galaxy_OclAllInstances_at_post_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y : shows "\ \ (UML_Set.OclForall ((OclAllInstances_at_post (Galaxy))) (OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y))" unfolding OclAllInstances_at_post_def by(rule Galaxy_OclAllInstances_generic_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y) lemma Galaxy_OclAllInstances_at_pre_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y : shows "\ \ (UML_Set.OclForall ((OclAllInstances_at_pre (Galaxy))) (OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y))" unfolding OclAllInstances_at_pre_def by(rule Galaxy_OclAllInstances_generic_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y) (* 109 ************************************ 693 + 1 *) section \Class Model: The Accessors\ (* 110 ************************************ 694 + 1 *) text \ \label{sec:Employee-DesignModel-UMLPart-generated-generatededm-accessors}\ (* 111 ************************************ 695 + 1 *) text \\ (* 112 ************************************ 696 + 1 *) subsection \Definition\ (* 113 ************************************ 697 + 1 *) text \\ (* 114 ************************************ 698 + 1 *) (* term Floor1_access.print_access_oid_uniq_ml *) ML \val oidPerson_0_boss = 0\ (* 115 ************************************ 699 + 1 *) (* term Floor1_access.print_access_oid_uniq *) definition "oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_0___boss = 0" (* 116 ************************************ 700 + 1 *) text \\ (* 117 ************************************ 701 + 5 *) (* term Floor1_access.print_access_eval_extract *) definition "eval_extract x f = (\\. (case x \ of \\obj\\ \ (f ((oid_of (obj))) (\)) | _ \ invalid \))" definition "in_pre_state = fst" definition "in_post_state = snd" definition "reconst_basetype = (\x _. \\x\\)" definition "reconst_basetype\<^sub>V\<^sub>o\<^sub>i\<^sub>d x = Abs_Void\<^sub>b\<^sub>a\<^sub>s\<^sub>e o (reconst_basetype (x))" (* 118 ************************************ 706 + 1 *) text \\ (* 119 ************************************ 707 + 2 *) (* term Floor1_access.print_access_choose_ml *) ML \val switch2_01 = (fn [x0 , x1] => (x0 , x1))\ ML \val switch2_10 = (fn [x0 , x1] => (x1 , x0))\ (* 120 ************************************ 709 + 3 *) (* term Floor1_access.print_access_choose *) definition "switch\<^sub>2_01 = (\ [x0 , x1] \ (x0 , x1))" definition "switch\<^sub>2_10 = (\ [x0 , x1] \ (x1 , x0))" definition "deref_assocs pre_post to_from assoc_oid f oid = (\\. (case (assocs ((pre_post (\))) (assoc_oid)) of \S\ \ (f ((deref_assocs_list (to_from) (oid) (S))) (\)) | _ \ (invalid (\))))" (* 121 ************************************ 712 + 4 *) (* term Floor1_access.print_access_deref_oid *) definition "deref_oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n fst_snd f oid = (\\. (case (heap (fst_snd \) (oid)) of \in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n obj\ \ f obj \ | _ \ invalid \))" definition "deref_oid\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t fst_snd f oid = (\\. (case (heap (fst_snd \) (oid)) of \in\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t obj\ \ f obj \ | _ \ invalid \))" definition "deref_oid\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y fst_snd f oid = (\\. (case (heap (fst_snd \) (oid)) of \in\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y obj\ \ f obj \ | _ \ invalid \))" definition "deref_oid\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y fst_snd f oid = (\\. (case (heap (fst_snd \) (oid)) of \in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y obj\ \ f obj \ | _ \ invalid \))" (* 122 ************************************ 716 + 0 *) (* term Floor1_access.print_access_deref_assocs *) (* 123 ************************************ 716 + 1 *) text \ pointer undefined in state or not referencing a type conform object representation \ (* 124 ************************************ 717 + 15 *) (* term Floor1_access.print_access_select *) definition "select\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__boss f = (\ (mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (_) (\) (_)) \ null | (mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (_) (\x___boss\) (_)) \ (f (x___boss)))" definition "select\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__salary f = (\ (mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (_) (_) (\)) \ null | (mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (_) (_) (\x___salary\)) \ (f (x___salary)))" definition "select\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__wormhole f = (\ (mk\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (_) (\) (_)) \ null | (mk\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (_) (\x___wormhole\) (_)) \ (f (x___wormhole)))" definition "select\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__weight f = (\ (mk\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (_) (_) (\)) \ null | (mk\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (_) (_) (\x___weight\)) \ (f (x___weight)))" definition "select\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y__sound f = (\ (mk\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y (_) (\) (_) (_)) \ null | (mk\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y (_) (\x___sound\) (_) (_)) \ (f (x___sound)))" definition "select\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y__moving f = (\ (mk\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y (_) (_) (\) (_)) \ null | (mk\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y (_) (_) (\x___moving\) (_)) \ (f (x___moving)))" definition "select\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y__outer_world f = (\ (mk\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y (_) (_) (_) (\)) \ null | (mk\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y (_) (_) (_) (\x___outer_world\)) \ (f (x___outer_world)))" definition "select\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__wormhole f = (\ (mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n ((mk\\\\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (_) (\) (_) (_) (_) (_))) (_) (_)) \ null | (mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n ((mk\\\\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (_) (\x___wormhole\) (_) (_) (_) (_))) (_) (_)) \ (f (x___wormhole)))" definition "select\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__weight f = (\ (mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n ((mk\\\\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (_) (_) (\) (_) (_) (_))) (_) (_)) \ null | (mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n ((mk\\\\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (_) (_) (\x___weight\) (_) (_) (_))) (_) (_)) \ (f (x___weight)))" definition "select\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__sound f = (\ (mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n ((mk\\\\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (_) (_) (_) (\) (_) (_))) (_) (_)) \ null | (mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n ((mk\\\\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (_) (_) (_) (\x___sound\) (_) (_))) (_) (_)) \ (f (x___sound)))" definition "select\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__moving f = (\ (mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n ((mk\\\\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (_) (_) (_) (_) (\) (_))) (_) (_)) \ null | (mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n ((mk\\\\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (_) (_) (_) (_) (\x___moving\) (_))) (_) (_)) \ (f (x___moving)))" definition "select\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__outer_world f = (\ (mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n ((mk\\\\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (_) (_) (_) (_) (_) (\))) (_) (_)) \ null | (mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n ((mk\\\\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (_) (_) (_) (_) (_) (\x___outer_world\))) (_) (_)) \ (f (x___outer_world)))" definition "select\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__sound f = (\ (mk\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t ((mk\\\\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (_) (\) (_) (_))) (_) (_)) \ null | (mk\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t ((mk\\\\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (_) (\x___sound\) (_) (_))) (_) (_)) \ (f (x___sound)) | (mk\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t ((mk\\\\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (person))) (_) (_)) \ (select\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__sound (f) (person)))" definition "select\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__moving f = (\ (mk\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t ((mk\\\\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (_) (_) (\) (_))) (_) (_)) \ null | (mk\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t ((mk\\\\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (_) (_) (\x___moving\) (_))) (_) (_)) \ (f (x___moving)) | (mk\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t ((mk\\\\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (person))) (_) (_)) \ (select\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__moving (f) (person)))" definition "select\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__outer_world f = (\ (mk\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t ((mk\\\\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (_) (_) (_) (\))) (_) (_)) \ null | (mk\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t ((mk\\\\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (_) (_) (_) (\x___outer_world\))) (_) (_)) \ (f (x___outer_world)) | (mk\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t ((mk\\\\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (person))) (_) (_)) \ (select\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__outer_world (f) (person)))" (* 125 ************************************ 732 + 0 *) (* term Floor1_access.print_access_select_obj *) (* 126 ************************************ 732 + 14 *) (* term Floor1_access.print_access_dot_consts *) consts dot_0___boss :: "(\, '\) val \ \Person" ("(_) .boss") consts dot_0___bossat_pre :: "(\, '\) val \ \Person" ("(_) .boss@pre") consts dot__salary :: "(\, '\) val \ Integer" ("(_) .salary") consts dot__salaryat_pre :: "(\, '\) val \ Integer" ("(_) .salary@pre") consts dot__wormhole :: "(\, '\) val \ (\, nat option option) val" ("(_) .wormhole") consts dot__wormholeat_pre :: "(\, '\) val \ (\, nat option option) val" ("(_) .wormhole@pre") consts dot__weight :: "(\, '\) val \ Integer" ("(_) .weight") consts dot__weightat_pre :: "(\, '\) val \ Integer" ("(_) .weight@pre") consts dot__sound :: "(\, '\) val \ Void" ("(_) .sound") consts dot__soundat_pre :: "(\, '\) val \ Void" ("(_) .sound@pre") consts dot__moving :: "(\, '\) val \ Boolean" ("(_) .moving") consts dot__movingat_pre :: "(\, '\) val \ Boolean" ("(_) .moving@pre") consts dot__outer_world :: "(\, '\) val \ Set_Sequence_Planet" ("(_) .outer'_world") consts dot__outer_worldat_pre :: "(\, '\) val \ Set_Sequence_Planet" ("(_) .outer'_world@pre") (* 127 ************************************ 746 + 30 *) (* term Floor1_access.print_access_dot *) overloading dot_0___boss \ "(dot_0___boss::(\Person) \ _)" begin definition dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_0___boss : "(x::\Person) .boss \ (eval_extract (x) ((deref_oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (in_post_state) ((select\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__boss ((select_object_any\<^sub>S\<^sub>e\<^sub>t ((deref_oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (in_post_state) (reconst_basetype))))))))))" end overloading dot__salary \ "(dot__salary::(\Person) \ _)" begin definition dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__salary : "(x::\Person) .salary \ (eval_extract (x) ((deref_oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (in_post_state) ((select\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__salary (reconst_basetype))))))" end overloading dot_0___bossat_pre \ "(dot_0___bossat_pre::(\Person) \ _)" begin definition dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_0___bossat_pre : "(x::\Person) .boss@pre \ (eval_extract (x) ((deref_oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (in_pre_state) ((select\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__boss ((select_object_any\<^sub>S\<^sub>e\<^sub>t ((deref_oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (in_pre_state) (reconst_basetype))))))))))" end overloading dot__salaryat_pre \ "(dot__salaryat_pre::(\Person) \ _)" begin definition dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__salaryat_pre : "(x::\Person) .salary@pre \ (eval_extract (x) ((deref_oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (in_pre_state) ((select\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__salary (reconst_basetype))))))" end overloading dot__wormhole \ "(dot__wormhole::(\Planet) \ _)" begin definition dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__wormhole : "(x::\Planet) .wormhole \ (eval_extract (x) ((deref_oid\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (in_post_state) ((select\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__wormhole (reconst_basetype))))))" end overloading dot__weight \ "(dot__weight::(\Planet) \ _)" begin definition dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__weight : "(x::\Planet) .weight \ (eval_extract (x) ((deref_oid\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (in_post_state) ((select\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__weight (reconst_basetype))))))" end overloading dot__wormholeat_pre \ "(dot__wormholeat_pre::(\Planet) \ _)" begin definition dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__wormholeat_pre : "(x::\Planet) .wormhole@pre \ (eval_extract (x) ((deref_oid\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (in_pre_state) ((select\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__wormhole (reconst_basetype))))))" end overloading dot__weightat_pre \ "(dot__weightat_pre::(\Planet) \ _)" begin definition dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__weightat_pre : "(x::\Planet) .weight@pre \ (eval_extract (x) ((deref_oid\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (in_pre_state) ((select\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__weight (reconst_basetype))))))" end overloading dot__sound \ "(dot__sound::(\Galaxy) \ _)" begin definition dot\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y__sound : "(x::\Galaxy) .sound \ (eval_extract (x) ((deref_oid\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y (in_post_state) ((select\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y__sound (reconst_basetype\<^sub>V\<^sub>o\<^sub>i\<^sub>d))))))" end overloading dot__moving \ "(dot__moving::(\Galaxy) \ _)" begin definition dot\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y__moving : "(x::\Galaxy) .moving \ (eval_extract (x) ((deref_oid\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y (in_post_state) ((select\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y__moving (reconst_basetype))))))" end overloading dot__outer_world \ "(dot__outer_world::(\Galaxy) \ _)" begin definition dot\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y__outer_world : "(x::\Galaxy) .outer_world \ (eval_extract (x) ((deref_oid\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y (in_post_state) ((select\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y__outer_world ((select_object\<^sub>S\<^sub>e\<^sub>t ((select_object\<^sub>S\<^sub>e\<^sub>q ((deref_oid\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (in_post_state) (reconst_basetype))))))))))))" end overloading dot__soundat_pre \ "(dot__soundat_pre::(\Galaxy) \ _)" begin definition dot\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y__soundat_pre : "(x::\Galaxy) .sound@pre \ (eval_extract (x) ((deref_oid\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y (in_pre_state) ((select\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y__sound (reconst_basetype\<^sub>V\<^sub>o\<^sub>i\<^sub>d))))))" end overloading dot__movingat_pre \ "(dot__movingat_pre::(\Galaxy) \ _)" begin definition dot\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y__movingat_pre : "(x::\Galaxy) .moving@pre \ (eval_extract (x) ((deref_oid\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y (in_pre_state) ((select\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y__moving (reconst_basetype))))))" end overloading dot__outer_worldat_pre \ "(dot__outer_worldat_pre::(\Galaxy) \ _)" begin definition dot\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y__outer_worldat_pre : "(x::\Galaxy) .outer_world@pre \ (eval_extract (x) ((deref_oid\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y (in_pre_state) ((select\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y__outer_world ((select_object\<^sub>S\<^sub>e\<^sub>t ((select_object\<^sub>S\<^sub>e\<^sub>q ((deref_oid\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (in_pre_state) (reconst_basetype))))))))))))" end overloading dot__wormhole \ "(dot__wormhole::(\Person) \ _)" begin definition dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__wormhole : "(x::\Person) .wormhole \ (eval_extract (x) ((deref_oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (in_post_state) ((select\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__wormhole (reconst_basetype))))))" end overloading dot__weight \ "(dot__weight::(\Person) \ _)" begin definition dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__weight : "(x::\Person) .weight \ (eval_extract (x) ((deref_oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (in_post_state) ((select\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__weight (reconst_basetype))))))" end overloading dot__sound \ "(dot__sound::(\Person) \ _)" begin definition dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__sound : "(x::\Person) .sound \ (eval_extract (x) ((deref_oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (in_post_state) ((select\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__sound (reconst_basetype\<^sub>V\<^sub>o\<^sub>i\<^sub>d))))))" end overloading dot__moving \ "(dot__moving::(\Person) \ _)" begin definition dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__moving : "(x::\Person) .moving \ (eval_extract (x) ((deref_oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (in_post_state) ((select\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__moving (reconst_basetype))))))" end overloading dot__outer_world \ "(dot__outer_world::(\Person) \ _)" begin definition dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__outer_world : "(x::\Person) .outer_world \ (eval_extract (x) ((deref_oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (in_post_state) ((select\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__outer_world ((select_object\<^sub>S\<^sub>e\<^sub>t ((select_object\<^sub>S\<^sub>e\<^sub>q ((deref_oid\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (in_post_state) (reconst_basetype))))))))))))" end overloading dot__wormholeat_pre \ "(dot__wormholeat_pre::(\Person) \ _)" begin definition dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__wormholeat_pre : "(x::\Person) .wormhole@pre \ (eval_extract (x) ((deref_oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (in_pre_state) ((select\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__wormhole (reconst_basetype))))))" end overloading dot__weightat_pre \ "(dot__weightat_pre::(\Person) \ _)" begin definition dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__weightat_pre : "(x::\Person) .weight@pre \ (eval_extract (x) ((deref_oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (in_pre_state) ((select\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__weight (reconst_basetype))))))" end overloading dot__soundat_pre \ "(dot__soundat_pre::(\Person) \ _)" begin definition dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__soundat_pre : "(x::\Person) .sound@pre \ (eval_extract (x) ((deref_oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (in_pre_state) ((select\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__sound (reconst_basetype\<^sub>V\<^sub>o\<^sub>i\<^sub>d))))))" end overloading dot__movingat_pre \ "(dot__movingat_pre::(\Person) \ _)" begin definition dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__movingat_pre : "(x::\Person) .moving@pre \ (eval_extract (x) ((deref_oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (in_pre_state) ((select\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__moving (reconst_basetype))))))" end overloading dot__outer_worldat_pre \ "(dot__outer_worldat_pre::(\Person) \ _)" begin definition dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__outer_worldat_pre : "(x::\Person) .outer_world@pre \ (eval_extract (x) ((deref_oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (in_pre_state) ((select\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__outer_world ((select_object\<^sub>S\<^sub>e\<^sub>t ((select_object\<^sub>S\<^sub>e\<^sub>q ((deref_oid\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (in_pre_state) (reconst_basetype))))))))))))" end overloading dot__sound \ "(dot__sound::(\Planet) \ _)" begin definition dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__sound : "(x::\Planet) .sound \ (eval_extract (x) ((deref_oid\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (in_post_state) ((select\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__sound (reconst_basetype\<^sub>V\<^sub>o\<^sub>i\<^sub>d))))))" end overloading dot__moving \ "(dot__moving::(\Planet) \ _)" begin definition dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__moving : "(x::\Planet) .moving \ (eval_extract (x) ((deref_oid\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (in_post_state) ((select\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__moving (reconst_basetype))))))" end overloading dot__outer_world \ "(dot__outer_world::(\Planet) \ _)" begin definition dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__outer_world : "(x::\Planet) .outer_world \ (eval_extract (x) ((deref_oid\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (in_post_state) ((select\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__outer_world ((select_object\<^sub>S\<^sub>e\<^sub>t ((select_object\<^sub>S\<^sub>e\<^sub>q ((deref_oid\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (in_post_state) (reconst_basetype))))))))))))" end overloading dot__soundat_pre \ "(dot__soundat_pre::(\Planet) \ _)" begin definition dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__soundat_pre : "(x::\Planet) .sound@pre \ (eval_extract (x) ((deref_oid\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (in_pre_state) ((select\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__sound (reconst_basetype\<^sub>V\<^sub>o\<^sub>i\<^sub>d))))))" end overloading dot__movingat_pre \ "(dot__movingat_pre::(\Planet) \ _)" begin definition dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__movingat_pre : "(x::\Planet) .moving@pre \ (eval_extract (x) ((deref_oid\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (in_pre_state) ((select\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__moving (reconst_basetype))))))" end overloading dot__outer_worldat_pre \ "(dot__outer_worldat_pre::(\Planet) \ _)" begin definition dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__outer_worldat_pre : "(x::\Planet) .outer_world@pre \ (eval_extract (x) ((deref_oid\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (in_pre_state) ((select\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__outer_world ((select_object\<^sub>S\<^sub>e\<^sub>t ((select_object\<^sub>S\<^sub>e\<^sub>q ((deref_oid\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (in_pre_state) (reconst_basetype))))))))))))" end (* 128 ************************************ 776 + 1 *) (* term Floor1_access.print_access_dot_lemmas_id *) lemmas dot_accessor = dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_0___boss dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__salary dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_0___bossat_pre dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__salaryat_pre dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__wormhole dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__weight dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__wormholeat_pre dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__weightat_pre dot\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y__sound dot\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y__moving dot\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y__outer_world dot\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y__soundat_pre dot\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y__movingat_pre dot\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y__outer_worldat_pre dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__wormhole dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__weight dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__sound dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__moving dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__outer_world dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__wormholeat_pre dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__weightat_pre dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__soundat_pre dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__movingat_pre dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__outer_worldat_pre dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__sound dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__moving dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__outer_world dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__soundat_pre dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__movingat_pre dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__outer_worldat_pre (* 129 ************************************ 777 + 1 *) subsection \Context Passing\ (* 130 ************************************ 778 + 1 *) (* term Floor1_access.print_access_dot_cp_lemmas *) lemmas[simp,code_unfold] = eval_extract_def (* 131 ************************************ 779 + 30 *) (* term Floor1_access.print_access_dot_lemma_cp *) lemma cp_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_0___boss : "(cp ((\X. (X::\Person) .boss)))" by(auto simp: dot_accessor cp_def) lemma cp_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__salary : "(cp ((\X. (X::\Person) .salary)))" by(auto simp: dot_accessor cp_def) lemma cp_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_0___bossat_pre : "(cp ((\X. (X::\Person) .boss@pre)))" by(auto simp: dot_accessor cp_def) lemma cp_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__salaryat_pre : "(cp ((\X. (X::\Person) .salary@pre)))" by(auto simp: dot_accessor cp_def) lemma cp_dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__wormhole : "(cp ((\X. (X::\Planet) .wormhole)))" by(auto simp: dot_accessor cp_def) lemma cp_dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__weight : "(cp ((\X. (X::\Planet) .weight)))" by(auto simp: dot_accessor cp_def) lemma cp_dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__wormholeat_pre : "(cp ((\X. (X::\Planet) .wormhole@pre)))" by(auto simp: dot_accessor cp_def) lemma cp_dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__weightat_pre : "(cp ((\X. (X::\Planet) .weight@pre)))" by(auto simp: dot_accessor cp_def) lemma cp_dot\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y__sound : "(cp ((\X. (X::\Galaxy) .sound)))" by(auto simp: dot_accessor cp_def) lemma cp_dot\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y__moving : "(cp ((\X. (X::\Galaxy) .moving)))" by(auto simp: dot_accessor cp_def) lemma cp_dot\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y__outer_world : "(cp ((\X. (X::\Galaxy) .outer_world)))" by(auto simp: dot_accessor cp_def) lemma cp_dot\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y__soundat_pre : "(cp ((\X. (X::\Galaxy) .sound@pre)))" by(auto simp: dot_accessor cp_def) lemma cp_dot\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y__movingat_pre : "(cp ((\X. (X::\Galaxy) .moving@pre)))" by(auto simp: dot_accessor cp_def) lemma cp_dot\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y__outer_worldat_pre : "(cp ((\X. (X::\Galaxy) .outer_world@pre)))" by(auto simp: dot_accessor cp_def) lemma cp_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__wormhole : "(cp ((\X. (X::\Person) .wormhole)))" by(auto simp: dot_accessor cp_def) lemma cp_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__weight : "(cp ((\X. (X::\Person) .weight)))" by(auto simp: dot_accessor cp_def) lemma cp_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__sound : "(cp ((\X. (X::\Person) .sound)))" by(auto simp: dot_accessor cp_def) lemma cp_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__moving : "(cp ((\X. (X::\Person) .moving)))" by(auto simp: dot_accessor cp_def) lemma cp_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__outer_world : "(cp ((\X. (X::\Person) .outer_world)))" by(auto simp: dot_accessor cp_def) lemma cp_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__wormholeat_pre : "(cp ((\X. (X::\Person) .wormhole@pre)))" by(auto simp: dot_accessor cp_def) lemma cp_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__weightat_pre : "(cp ((\X. (X::\Person) .weight@pre)))" by(auto simp: dot_accessor cp_def) lemma cp_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__soundat_pre : "(cp ((\X. (X::\Person) .sound@pre)))" by(auto simp: dot_accessor cp_def) lemma cp_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__movingat_pre : "(cp ((\X. (X::\Person) .moving@pre)))" by(auto simp: dot_accessor cp_def) lemma cp_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__outer_worldat_pre : "(cp ((\X. (X::\Person) .outer_world@pre)))" by(auto simp: dot_accessor cp_def) lemma cp_dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__sound : "(cp ((\X. (X::\Planet) .sound)))" by(auto simp: dot_accessor cp_def) lemma cp_dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__moving : "(cp ((\X. (X::\Planet) .moving)))" by(auto simp: dot_accessor cp_def) lemma cp_dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__outer_world : "(cp ((\X. (X::\Planet) .outer_world)))" by(auto simp: dot_accessor cp_def) lemma cp_dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__soundat_pre : "(cp ((\X. (X::\Planet) .sound@pre)))" by(auto simp: dot_accessor cp_def) lemma cp_dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__movingat_pre : "(cp ((\X. (X::\Planet) .moving@pre)))" by(auto simp: dot_accessor cp_def) lemma cp_dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__outer_worldat_pre : "(cp ((\X. (X::\Planet) .outer_world@pre)))" by(auto simp: dot_accessor cp_def) (* 132 ************************************ 809 + 1 *) (* term Floor1_access.print_access_dot_lemmas_cp *) lemmas[simp,code_unfold] = cp_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_0___boss cp_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__salary cp_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_0___bossat_pre cp_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__salaryat_pre cp_dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__wormhole cp_dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__weight cp_dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__wormholeat_pre cp_dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__weightat_pre cp_dot\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y__sound cp_dot\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y__moving cp_dot\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y__outer_world cp_dot\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y__soundat_pre cp_dot\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y__movingat_pre cp_dot\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y__outer_worldat_pre cp_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__wormhole cp_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__weight cp_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__sound cp_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__moving cp_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__outer_world cp_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__wormholeat_pre cp_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__weightat_pre cp_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__soundat_pre cp_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__movingat_pre cp_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__outer_worldat_pre cp_dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__sound cp_dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__moving cp_dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__outer_world cp_dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__soundat_pre cp_dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__movingat_pre cp_dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__outer_worldat_pre (* 133 ************************************ 810 + 1 *) subsection \Execution with Invalid or Null as Argument\ (* 134 ************************************ 811 + 60 *) (* term Floor1_access.print_access_lemma_strict *) lemma dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_0___boss_invalid : "(invalid::\Person) .boss = invalid" by(rule ext, simp add: dot_accessor bot_option_def invalid_def) lemma dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_0___boss_null : "(null::\Person) .boss = invalid" by(rule ext, simp add: dot_accessor bot_option_def null_fun_def null_option_def) lemma dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__salary_invalid : "(invalid::\Person) .salary = invalid" by(rule ext, simp add: dot_accessor bot_option_def invalid_def) lemma dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__salary_null : "(null::\Person) .salary = invalid" by(rule ext, simp add: dot_accessor bot_option_def null_fun_def null_option_def) lemma dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_0___bossat_pre_invalid : "(invalid::\Person) .boss@pre = invalid" by(rule ext, simp add: dot_accessor bot_option_def invalid_def) lemma dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_0___bossat_pre_null : "(null::\Person) .boss@pre = invalid" by(rule ext, simp add: dot_accessor bot_option_def null_fun_def null_option_def) lemma dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__salaryat_pre_invalid : "(invalid::\Person) .salary@pre = invalid" by(rule ext, simp add: dot_accessor bot_option_def invalid_def) lemma dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__salaryat_pre_null : "(null::\Person) .salary@pre = invalid" by(rule ext, simp add: dot_accessor bot_option_def null_fun_def null_option_def) lemma dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__wormhole_invalid : "(invalid::\Planet) .wormhole = invalid" by(rule ext, simp add: dot_accessor bot_option_def invalid_def) lemma dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__wormhole_null : "(null::\Planet) .wormhole = invalid" by(rule ext, simp add: dot_accessor bot_option_def null_fun_def null_option_def) lemma dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__weight_invalid : "(invalid::\Planet) .weight = invalid" by(rule ext, simp add: dot_accessor bot_option_def invalid_def) lemma dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__weight_null : "(null::\Planet) .weight = invalid" by(rule ext, simp add: dot_accessor bot_option_def null_fun_def null_option_def) lemma dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__wormholeat_pre_invalid : "(invalid::\Planet) .wormhole@pre = invalid" by(rule ext, simp add: dot_accessor bot_option_def invalid_def) lemma dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__wormholeat_pre_null : "(null::\Planet) .wormhole@pre = invalid" by(rule ext, simp add: dot_accessor bot_option_def null_fun_def null_option_def) lemma dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__weightat_pre_invalid : "(invalid::\Planet) .weight@pre = invalid" by(rule ext, simp add: dot_accessor bot_option_def invalid_def) lemma dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__weightat_pre_null : "(null::\Planet) .weight@pre = invalid" by(rule ext, simp add: dot_accessor bot_option_def null_fun_def null_option_def) lemma dot\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y__sound_invalid : "(invalid::\Galaxy) .sound = invalid" by(rule ext, simp add: dot_accessor bot_option_def invalid_def) lemma dot\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y__sound_null : "(null::\Galaxy) .sound = invalid" by(rule ext, simp add: dot_accessor bot_option_def null_fun_def null_option_def) lemma dot\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y__moving_invalid : "(invalid::\Galaxy) .moving = invalid" by(rule ext, simp add: dot_accessor bot_option_def invalid_def) lemma dot\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y__moving_null : "(null::\Galaxy) .moving = invalid" by(rule ext, simp add: dot_accessor bot_option_def null_fun_def null_option_def) lemma dot\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y__outer_world_invalid : "(invalid::\Galaxy) .outer_world = invalid" by(rule ext, simp add: dot_accessor bot_option_def invalid_def) lemma dot\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y__outer_world_null : "(null::\Galaxy) .outer_world = invalid" by(rule ext, simp add: dot_accessor bot_option_def null_fun_def null_option_def) lemma dot\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y__soundat_pre_invalid : "(invalid::\Galaxy) .sound@pre = invalid" by(rule ext, simp add: dot_accessor bot_option_def invalid_def) lemma dot\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y__soundat_pre_null : "(null::\Galaxy) .sound@pre = invalid" by(rule ext, simp add: dot_accessor bot_option_def null_fun_def null_option_def) lemma dot\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y__movingat_pre_invalid : "(invalid::\Galaxy) .moving@pre = invalid" by(rule ext, simp add: dot_accessor bot_option_def invalid_def) lemma dot\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y__movingat_pre_null : "(null::\Galaxy) .moving@pre = invalid" by(rule ext, simp add: dot_accessor bot_option_def null_fun_def null_option_def) lemma dot\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y__outer_worldat_pre_invalid : "(invalid::\Galaxy) .outer_world@pre = invalid" by(rule ext, simp add: dot_accessor bot_option_def invalid_def) lemma dot\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y__outer_worldat_pre_null : "(null::\Galaxy) .outer_world@pre = invalid" by(rule ext, simp add: dot_accessor bot_option_def null_fun_def null_option_def) lemma dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__wormhole_invalid : "(invalid::\Person) .wormhole = invalid" by(rule ext, simp add: dot_accessor bot_option_def invalid_def) lemma dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__wormhole_null : "(null::\Person) .wormhole = invalid" by(rule ext, simp add: dot_accessor bot_option_def null_fun_def null_option_def) lemma dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__weight_invalid : "(invalid::\Person) .weight = invalid" by(rule ext, simp add: dot_accessor bot_option_def invalid_def) lemma dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__weight_null : "(null::\Person) .weight = invalid" by(rule ext, simp add: dot_accessor bot_option_def null_fun_def null_option_def) lemma dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__sound_invalid : "(invalid::\Person) .sound = invalid" by(rule ext, simp add: dot_accessor bot_option_def invalid_def) lemma dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__sound_null : "(null::\Person) .sound = invalid" by(rule ext, simp add: dot_accessor bot_option_def null_fun_def null_option_def) lemma dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__moving_invalid : "(invalid::\Person) .moving = invalid" by(rule ext, simp add: dot_accessor bot_option_def invalid_def) lemma dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__moving_null : "(null::\Person) .moving = invalid" by(rule ext, simp add: dot_accessor bot_option_def null_fun_def null_option_def) lemma dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__outer_world_invalid : "(invalid::\Person) .outer_world = invalid" by(rule ext, simp add: dot_accessor bot_option_def invalid_def) lemma dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__outer_world_null : "(null::\Person) .outer_world = invalid" by(rule ext, simp add: dot_accessor bot_option_def null_fun_def null_option_def) lemma dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__wormholeat_pre_invalid : "(invalid::\Person) .wormhole@pre = invalid" by(rule ext, simp add: dot_accessor bot_option_def invalid_def) lemma dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__wormholeat_pre_null : "(null::\Person) .wormhole@pre = invalid" by(rule ext, simp add: dot_accessor bot_option_def null_fun_def null_option_def) lemma dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__weightat_pre_invalid : "(invalid::\Person) .weight@pre = invalid" by(rule ext, simp add: dot_accessor bot_option_def invalid_def) lemma dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__weightat_pre_null : "(null::\Person) .weight@pre = invalid" by(rule ext, simp add: dot_accessor bot_option_def null_fun_def null_option_def) lemma dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__soundat_pre_invalid : "(invalid::\Person) .sound@pre = invalid" by(rule ext, simp add: dot_accessor bot_option_def invalid_def) lemma dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__soundat_pre_null : "(null::\Person) .sound@pre = invalid" by(rule ext, simp add: dot_accessor bot_option_def null_fun_def null_option_def) lemma dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__movingat_pre_invalid : "(invalid::\Person) .moving@pre = invalid" by(rule ext, simp add: dot_accessor bot_option_def invalid_def) lemma dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__movingat_pre_null : "(null::\Person) .moving@pre = invalid" by(rule ext, simp add: dot_accessor bot_option_def null_fun_def null_option_def) lemma dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__outer_worldat_pre_invalid : "(invalid::\Person) .outer_world@pre = invalid" by(rule ext, simp add: dot_accessor bot_option_def invalid_def) lemma dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__outer_worldat_pre_null : "(null::\Person) .outer_world@pre = invalid" by(rule ext, simp add: dot_accessor bot_option_def null_fun_def null_option_def) lemma dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__sound_invalid : "(invalid::\Planet) .sound = invalid" by(rule ext, simp add: dot_accessor bot_option_def invalid_def) lemma dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__sound_null : "(null::\Planet) .sound = invalid" by(rule ext, simp add: dot_accessor bot_option_def null_fun_def null_option_def) lemma dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__moving_invalid : "(invalid::\Planet) .moving = invalid" by(rule ext, simp add: dot_accessor bot_option_def invalid_def) lemma dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__moving_null : "(null::\Planet) .moving = invalid" by(rule ext, simp add: dot_accessor bot_option_def null_fun_def null_option_def) lemma dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__outer_world_invalid : "(invalid::\Planet) .outer_world = invalid" by(rule ext, simp add: dot_accessor bot_option_def invalid_def) lemma dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__outer_world_null : "(null::\Planet) .outer_world = invalid" by(rule ext, simp add: dot_accessor bot_option_def null_fun_def null_option_def) lemma dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__soundat_pre_invalid : "(invalid::\Planet) .sound@pre = invalid" by(rule ext, simp add: dot_accessor bot_option_def invalid_def) lemma dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__soundat_pre_null : "(null::\Planet) .sound@pre = invalid" by(rule ext, simp add: dot_accessor bot_option_def null_fun_def null_option_def) lemma dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__movingat_pre_invalid : "(invalid::\Planet) .moving@pre = invalid" by(rule ext, simp add: dot_accessor bot_option_def invalid_def) lemma dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__movingat_pre_null : "(null::\Planet) .moving@pre = invalid" by(rule ext, simp add: dot_accessor bot_option_def null_fun_def null_option_def) lemma dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__outer_worldat_pre_invalid : "(invalid::\Planet) .outer_world@pre = invalid" by(rule ext, simp add: dot_accessor bot_option_def invalid_def) lemma dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__outer_worldat_pre_null : "(null::\Planet) .outer_world@pre = invalid" by(rule ext, simp add: dot_accessor bot_option_def null_fun_def null_option_def) (* 135 ************************************ 871 + 1 *) subsection \Representation in States\ (* 136 ************************************ 872 + 30 *) (* term Floor1_access.print_access_def_mono *) lemma defined_mono_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_0___boss : "\ \ (\ ((X::\Person) .boss)) \ \ \ (\ (X))" apply(case_tac "\ \ (X \ invalid)", insert StrongEq_L_subst2[where P = "(\x. (\ (x .boss)))" and \ = "\" and x = "X" and y = "invalid"], simp add: foundation16' dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_0___boss_invalid) apply(case_tac "\ \ (X \ null)", insert StrongEq_L_subst2[where P = "(\x. (\ (x .boss)))" and \ = "\" and x = "X" and y = "null"], simp add: foundation16' dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_0___boss_null) by(simp add: defined_split) lemma defined_mono_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__salary : "\ \ (\ ((X::\Person) .salary)) \ \ \ (\ (X))" apply(case_tac "\ \ (X \ invalid)", insert StrongEq_L_subst2[where P = "(\x. (\ (x .salary)))" and \ = "\" and x = "X" and y = "invalid"], simp add: foundation16' dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__salary_invalid) apply(case_tac "\ \ (X \ null)", insert StrongEq_L_subst2[where P = "(\x. (\ (x .salary)))" and \ = "\" and x = "X" and y = "null"], simp add: foundation16' dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__salary_null) by(simp add: defined_split) lemma defined_mono_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_0___bossat_pre : "\ \ (\ ((X::\Person) .boss@pre)) \ \ \ (\ (X))" apply(case_tac "\ \ (X \ invalid)", insert StrongEq_L_subst2[where P = "(\x. (\ (x .boss@pre)))" and \ = "\" and x = "X" and y = "invalid"], simp add: foundation16' dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_0___bossat_pre_invalid) apply(case_tac "\ \ (X \ null)", insert StrongEq_L_subst2[where P = "(\x. (\ (x .boss@pre)))" and \ = "\" and x = "X" and y = "null"], simp add: foundation16' dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_0___bossat_pre_null) by(simp add: defined_split) lemma defined_mono_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__salaryat_pre : "\ \ (\ ((X::\Person) .salary@pre)) \ \ \ (\ (X))" apply(case_tac "\ \ (X \ invalid)", insert StrongEq_L_subst2[where P = "(\x. (\ (x .salary@pre)))" and \ = "\" and x = "X" and y = "invalid"], simp add: foundation16' dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__salaryat_pre_invalid) apply(case_tac "\ \ (X \ null)", insert StrongEq_L_subst2[where P = "(\x. (\ (x .salary@pre)))" and \ = "\" and x = "X" and y = "null"], simp add: foundation16' dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__salaryat_pre_null) by(simp add: defined_split) lemma defined_mono_dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__wormhole : "\ \ (\ ((X::\Planet) .wormhole)) \ \ \ (\ (X))" apply(case_tac "\ \ (X \ invalid)", insert StrongEq_L_subst2[where P = "(\x. (\ (x .wormhole)))" and \ = "\" and x = "X" and y = "invalid"], simp add: foundation16' dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__wormhole_invalid) apply(case_tac "\ \ (X \ null)", insert StrongEq_L_subst2[where P = "(\x. (\ (x .wormhole)))" and \ = "\" and x = "X" and y = "null"], simp add: foundation16' dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__wormhole_null) by(simp add: defined_split) lemma defined_mono_dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__weight : "\ \ (\ ((X::\Planet) .weight)) \ \ \ (\ (X))" apply(case_tac "\ \ (X \ invalid)", insert StrongEq_L_subst2[where P = "(\x. (\ (x .weight)))" and \ = "\" and x = "X" and y = "invalid"], simp add: foundation16' dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__weight_invalid) apply(case_tac "\ \ (X \ null)", insert StrongEq_L_subst2[where P = "(\x. (\ (x .weight)))" and \ = "\" and x = "X" and y = "null"], simp add: foundation16' dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__weight_null) by(simp add: defined_split) lemma defined_mono_dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__wormholeat_pre : "\ \ (\ ((X::\Planet) .wormhole@pre)) \ \ \ (\ (X))" apply(case_tac "\ \ (X \ invalid)", insert StrongEq_L_subst2[where P = "(\x. (\ (x .wormhole@pre)))" and \ = "\" and x = "X" and y = "invalid"], simp add: foundation16' dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__wormholeat_pre_invalid) apply(case_tac "\ \ (X \ null)", insert StrongEq_L_subst2[where P = "(\x. (\ (x .wormhole@pre)))" and \ = "\" and x = "X" and y = "null"], simp add: foundation16' dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__wormholeat_pre_null) by(simp add: defined_split) lemma defined_mono_dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__weightat_pre : "\ \ (\ ((X::\Planet) .weight@pre)) \ \ \ (\ (X))" apply(case_tac "\ \ (X \ invalid)", insert StrongEq_L_subst2[where P = "(\x. (\ (x .weight@pre)))" and \ = "\" and x = "X" and y = "invalid"], simp add: foundation16' dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__weightat_pre_invalid) apply(case_tac "\ \ (X \ null)", insert StrongEq_L_subst2[where P = "(\x. (\ (x .weight@pre)))" and \ = "\" and x = "X" and y = "null"], simp add: foundation16' dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__weightat_pre_null) by(simp add: defined_split) lemma defined_mono_dot\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y__sound : "\ \ (\ ((X::\Galaxy) .sound)) \ \ \ (\ (X))" apply(case_tac "\ \ (X \ invalid)", insert StrongEq_L_subst2[where P = "(\x. (\ (x .sound)))" and \ = "\" and x = "X" and y = "invalid"], simp add: foundation16' dot\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y__sound_invalid) apply(case_tac "\ \ (X \ null)", insert StrongEq_L_subst2[where P = "(\x. (\ (x .sound)))" and \ = "\" and x = "X" and y = "null"], simp add: foundation16' dot\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y__sound_null) by(simp add: defined_split) lemma defined_mono_dot\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y__moving : "\ \ (\ ((X::\Galaxy) .moving)) \ \ \ (\ (X))" apply(case_tac "\ \ (X \ invalid)", insert StrongEq_L_subst2[where P = "(\x. (\ (x .moving)))" and \ = "\" and x = "X" and y = "invalid"], simp add: foundation16' dot\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y__moving_invalid) apply(case_tac "\ \ (X \ null)", insert StrongEq_L_subst2[where P = "(\x. (\ (x .moving)))" and \ = "\" and x = "X" and y = "null"], simp add: foundation16' dot\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y__moving_null) by(simp add: defined_split) lemma defined_mono_dot\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y__outer_world : "\ \ (\ ((X::\Galaxy) .outer_world)) \ \ \ (\ (X))" apply(case_tac "\ \ (X \ invalid)", insert StrongEq_L_subst2[where P = "(\x. (\ (x .outer_world)))" and \ = "\" and x = "X" and y = "invalid"], simp add: foundation16' dot\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y__outer_world_invalid) apply(case_tac "\ \ (X \ null)", insert StrongEq_L_subst2[where P = "(\x. (\ (x .outer_world)))" and \ = "\" and x = "X" and y = "null"], simp add: foundation16' dot\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y__outer_world_null) by(simp add: defined_split) lemma defined_mono_dot\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y__soundat_pre : "\ \ (\ ((X::\Galaxy) .sound@pre)) \ \ \ (\ (X))" apply(case_tac "\ \ (X \ invalid)", insert StrongEq_L_subst2[where P = "(\x. (\ (x .sound@pre)))" and \ = "\" and x = "X" and y = "invalid"], simp add: foundation16' dot\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y__soundat_pre_invalid) apply(case_tac "\ \ (X \ null)", insert StrongEq_L_subst2[where P = "(\x. (\ (x .sound@pre)))" and \ = "\" and x = "X" and y = "null"], simp add: foundation16' dot\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y__soundat_pre_null) by(simp add: defined_split) lemma defined_mono_dot\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y__movingat_pre : "\ \ (\ ((X::\Galaxy) .moving@pre)) \ \ \ (\ (X))" apply(case_tac "\ \ (X \ invalid)", insert StrongEq_L_subst2[where P = "(\x. (\ (x .moving@pre)))" and \ = "\" and x = "X" and y = "invalid"], simp add: foundation16' dot\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y__movingat_pre_invalid) apply(case_tac "\ \ (X \ null)", insert StrongEq_L_subst2[where P = "(\x. (\ (x .moving@pre)))" and \ = "\" and x = "X" and y = "null"], simp add: foundation16' dot\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y__movingat_pre_null) by(simp add: defined_split) lemma defined_mono_dot\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y__outer_worldat_pre : "\ \ (\ ((X::\Galaxy) .outer_world@pre)) \ \ \ (\ (X))" apply(case_tac "\ \ (X \ invalid)", insert StrongEq_L_subst2[where P = "(\x. (\ (x .outer_world@pre)))" and \ = "\" and x = "X" and y = "invalid"], simp add: foundation16' dot\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y__outer_worldat_pre_invalid) apply(case_tac "\ \ (X \ null)", insert StrongEq_L_subst2[where P = "(\x. (\ (x .outer_world@pre)))" and \ = "\" and x = "X" and y = "null"], simp add: foundation16' dot\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y__outer_worldat_pre_null) by(simp add: defined_split) lemma defined_mono_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__wormhole : "\ \ (\ ((X::\Person) .wormhole)) \ \ \ (\ (X))" apply(case_tac "\ \ (X \ invalid)", insert StrongEq_L_subst2[where P = "(\x. (\ (x .wormhole)))" and \ = "\" and x = "X" and y = "invalid"], simp add: foundation16' dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__wormhole_invalid) apply(case_tac "\ \ (X \ null)", insert StrongEq_L_subst2[where P = "(\x. (\ (x .wormhole)))" and \ = "\" and x = "X" and y = "null"], simp add: foundation16' dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__wormhole_null) by(simp add: defined_split) lemma defined_mono_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__weight : "\ \ (\ ((X::\Person) .weight)) \ \ \ (\ (X))" apply(case_tac "\ \ (X \ invalid)", insert StrongEq_L_subst2[where P = "(\x. (\ (x .weight)))" and \ = "\" and x = "X" and y = "invalid"], simp add: foundation16' dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__weight_invalid) apply(case_tac "\ \ (X \ null)", insert StrongEq_L_subst2[where P = "(\x. (\ (x .weight)))" and \ = "\" and x = "X" and y = "null"], simp add: foundation16' dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__weight_null) by(simp add: defined_split) lemma defined_mono_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__sound : "\ \ (\ ((X::\Person) .sound)) \ \ \ (\ (X))" apply(case_tac "\ \ (X \ invalid)", insert StrongEq_L_subst2[where P = "(\x. (\ (x .sound)))" and \ = "\" and x = "X" and y = "invalid"], simp add: foundation16' dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__sound_invalid) apply(case_tac "\ \ (X \ null)", insert StrongEq_L_subst2[where P = "(\x. (\ (x .sound)))" and \ = "\" and x = "X" and y = "null"], simp add: foundation16' dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__sound_null) by(simp add: defined_split) lemma defined_mono_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__moving : "\ \ (\ ((X::\Person) .moving)) \ \ \ (\ (X))" apply(case_tac "\ \ (X \ invalid)", insert StrongEq_L_subst2[where P = "(\x. (\ (x .moving)))" and \ = "\" and x = "X" and y = "invalid"], simp add: foundation16' dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__moving_invalid) apply(case_tac "\ \ (X \ null)", insert StrongEq_L_subst2[where P = "(\x. (\ (x .moving)))" and \ = "\" and x = "X" and y = "null"], simp add: foundation16' dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__moving_null) by(simp add: defined_split) lemma defined_mono_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__outer_world : "\ \ (\ ((X::\Person) .outer_world)) \ \ \ (\ (X))" apply(case_tac "\ \ (X \ invalid)", insert StrongEq_L_subst2[where P = "(\x. (\ (x .outer_world)))" and \ = "\" and x = "X" and y = "invalid"], simp add: foundation16' dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__outer_world_invalid) apply(case_tac "\ \ (X \ null)", insert StrongEq_L_subst2[where P = "(\x. (\ (x .outer_world)))" and \ = "\" and x = "X" and y = "null"], simp add: foundation16' dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__outer_world_null) by(simp add: defined_split) lemma defined_mono_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__wormholeat_pre : "\ \ (\ ((X::\Person) .wormhole@pre)) \ \ \ (\ (X))" apply(case_tac "\ \ (X \ invalid)", insert StrongEq_L_subst2[where P = "(\x. (\ (x .wormhole@pre)))" and \ = "\" and x = "X" and y = "invalid"], simp add: foundation16' dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__wormholeat_pre_invalid) apply(case_tac "\ \ (X \ null)", insert StrongEq_L_subst2[where P = "(\x. (\ (x .wormhole@pre)))" and \ = "\" and x = "X" and y = "null"], simp add: foundation16' dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__wormholeat_pre_null) by(simp add: defined_split) lemma defined_mono_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__weightat_pre : "\ \ (\ ((X::\Person) .weight@pre)) \ \ \ (\ (X))" apply(case_tac "\ \ (X \ invalid)", insert StrongEq_L_subst2[where P = "(\x. (\ (x .weight@pre)))" and \ = "\" and x = "X" and y = "invalid"], simp add: foundation16' dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__weightat_pre_invalid) apply(case_tac "\ \ (X \ null)", insert StrongEq_L_subst2[where P = "(\x. (\ (x .weight@pre)))" and \ = "\" and x = "X" and y = "null"], simp add: foundation16' dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__weightat_pre_null) by(simp add: defined_split) lemma defined_mono_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__soundat_pre : "\ \ (\ ((X::\Person) .sound@pre)) \ \ \ (\ (X))" apply(case_tac "\ \ (X \ invalid)", insert StrongEq_L_subst2[where P = "(\x. (\ (x .sound@pre)))" and \ = "\" and x = "X" and y = "invalid"], simp add: foundation16' dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__soundat_pre_invalid) apply(case_tac "\ \ (X \ null)", insert StrongEq_L_subst2[where P = "(\x. (\ (x .sound@pre)))" and \ = "\" and x = "X" and y = "null"], simp add: foundation16' dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__soundat_pre_null) by(simp add: defined_split) lemma defined_mono_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__movingat_pre : "\ \ (\ ((X::\Person) .moving@pre)) \ \ \ (\ (X))" apply(case_tac "\ \ (X \ invalid)", insert StrongEq_L_subst2[where P = "(\x. (\ (x .moving@pre)))" and \ = "\" and x = "X" and y = "invalid"], simp add: foundation16' dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__movingat_pre_invalid) apply(case_tac "\ \ (X \ null)", insert StrongEq_L_subst2[where P = "(\x. (\ (x .moving@pre)))" and \ = "\" and x = "X" and y = "null"], simp add: foundation16' dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__movingat_pre_null) by(simp add: defined_split) lemma defined_mono_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__outer_worldat_pre : "\ \ (\ ((X::\Person) .outer_world@pre)) \ \ \ (\ (X))" apply(case_tac "\ \ (X \ invalid)", insert StrongEq_L_subst2[where P = "(\x. (\ (x .outer_world@pre)))" and \ = "\" and x = "X" and y = "invalid"], simp add: foundation16' dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__outer_worldat_pre_invalid) apply(case_tac "\ \ (X \ null)", insert StrongEq_L_subst2[where P = "(\x. (\ (x .outer_world@pre)))" and \ = "\" and x = "X" and y = "null"], simp add: foundation16' dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__outer_worldat_pre_null) by(simp add: defined_split) lemma defined_mono_dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__sound : "\ \ (\ ((X::\Planet) .sound)) \ \ \ (\ (X))" apply(case_tac "\ \ (X \ invalid)", insert StrongEq_L_subst2[where P = "(\x. (\ (x .sound)))" and \ = "\" and x = "X" and y = "invalid"], simp add: foundation16' dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__sound_invalid) apply(case_tac "\ \ (X \ null)", insert StrongEq_L_subst2[where P = "(\x. (\ (x .sound)))" and \ = "\" and x = "X" and y = "null"], simp add: foundation16' dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__sound_null) by(simp add: defined_split) lemma defined_mono_dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__moving : "\ \ (\ ((X::\Planet) .moving)) \ \ \ (\ (X))" apply(case_tac "\ \ (X \ invalid)", insert StrongEq_L_subst2[where P = "(\x. (\ (x .moving)))" and \ = "\" and x = "X" and y = "invalid"], simp add: foundation16' dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__moving_invalid) apply(case_tac "\ \ (X \ null)", insert StrongEq_L_subst2[where P = "(\x. (\ (x .moving)))" and \ = "\" and x = "X" and y = "null"], simp add: foundation16' dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__moving_null) by(simp add: defined_split) lemma defined_mono_dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__outer_world : "\ \ (\ ((X::\Planet) .outer_world)) \ \ \ (\ (X))" apply(case_tac "\ \ (X \ invalid)", insert StrongEq_L_subst2[where P = "(\x. (\ (x .outer_world)))" and \ = "\" and x = "X" and y = "invalid"], simp add: foundation16' dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__outer_world_invalid) apply(case_tac "\ \ (X \ null)", insert StrongEq_L_subst2[where P = "(\x. (\ (x .outer_world)))" and \ = "\" and x = "X" and y = "null"], simp add: foundation16' dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__outer_world_null) by(simp add: defined_split) lemma defined_mono_dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__soundat_pre : "\ \ (\ ((X::\Planet) .sound@pre)) \ \ \ (\ (X))" apply(case_tac "\ \ (X \ invalid)", insert StrongEq_L_subst2[where P = "(\x. (\ (x .sound@pre)))" and \ = "\" and x = "X" and y = "invalid"], simp add: foundation16' dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__soundat_pre_invalid) apply(case_tac "\ \ (X \ null)", insert StrongEq_L_subst2[where P = "(\x. (\ (x .sound@pre)))" and \ = "\" and x = "X" and y = "null"], simp add: foundation16' dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__soundat_pre_null) by(simp add: defined_split) lemma defined_mono_dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__movingat_pre : "\ \ (\ ((X::\Planet) .moving@pre)) \ \ \ (\ (X))" apply(case_tac "\ \ (X \ invalid)", insert StrongEq_L_subst2[where P = "(\x. (\ (x .moving@pre)))" and \ = "\" and x = "X" and y = "invalid"], simp add: foundation16' dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__movingat_pre_invalid) apply(case_tac "\ \ (X \ null)", insert StrongEq_L_subst2[where P = "(\x. (\ (x .moving@pre)))" and \ = "\" and x = "X" and y = "null"], simp add: foundation16' dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__movingat_pre_null) by(simp add: defined_split) lemma defined_mono_dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__outer_worldat_pre : "\ \ (\ ((X::\Planet) .outer_world@pre)) \ \ \ (\ (X))" apply(case_tac "\ \ (X \ invalid)", insert StrongEq_L_subst2[where P = "(\x. (\ (x .outer_world@pre)))" and \ = "\" and x = "X" and y = "invalid"], simp add: foundation16' dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__outer_worldat_pre_invalid) apply(case_tac "\ \ (X \ null)", insert StrongEq_L_subst2[where P = "(\x. (\ (x .outer_world@pre)))" and \ = "\" and x = "X" and y = "null"], simp add: foundation16' dot\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t__outer_worldat_pre_null) by(simp add: defined_split) (* 137 ************************************ 902 + 2 *) (* term Floor1_access.print_access_is_repr *) lemma is_repr_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_0___boss : assumes def_dot: "\ \ (\ ((X::\Person) .boss))" shows "(is_represented_in_state (in_post_state) (X .boss) (Person) (\))" apply(insert defined_mono_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_0___boss[OF def_dot, simplified foundation16]) apply(case_tac "(X (\))", simp add: bot_option_def) proof - fix a0 show "(X (\)) = (Some (a0)) \ ?thesis" when "(X (\)) \ null" apply(insert that, case_tac "a0", simp add: null_option_def bot_option_def, clarify) proof - fix a show "(X (\)) = (Some ((Some (a)))) \ ?thesis" apply(case_tac "(heap ((in_post_state (\))) ((oid_of (a))))", simp add: invalid_def bot_option_def) apply(insert def_dot, simp add: dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_0___boss is_represented_in_state_def select\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__boss_def deref_oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_def in_post_state_def defined_def OclValid_def false_def true_def invalid_def bot_fun_def split: if_split_asm) proof - fix b show "(X (\)) = (Some ((Some (a)))) \ (heap ((in_post_state (\))) ((oid_of (a)))) = (Some (b)) \ ?thesis" apply(insert def_dot[simplified foundation16], auto simp: dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_0___boss is_represented_in_state_def deref_oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_def bot_option_def null_option_def) apply(case_tac "b", simp_all add: invalid_def bot_option_def) proof - fix r typeoid let ?t = "(Some ((Some (r)))) \ (Some o OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_\) ` (ran ((heap ((in_post_state (\))))))" let ?sel_any = "(select_object_any\<^sub>S\<^sub>e\<^sub>t ((deref_oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (in_post_state) (reconst_basetype))))" show "(select\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__boss (?sel_any) (typeoid) (\)) = (Some ((Some (r)))) \ ?t" apply(case_tac "typeoid", simp add: select\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__boss_def) proof - fix opt show "(((case opt of None \ null | (Some (x)) \ ((?sel_any) (x)))) (\)) = (Some ((Some (r)))) \ ?t" apply(case_tac "opt", auto simp: null_fun_def null_option_def bot_option_def) proof - fix aa show "((?sel_any) (aa) (\)) = (Some ((Some (r)))) \ ?t" when "\ \ (\ (((?sel_any) (aa))))" apply(insert that, drule select_object_any_exec\<^sub>S\<^sub>e\<^sub>t[simplified foundation22], erule exE) proof - fix e show "?t" when "((?sel_any) (aa) (\)) = (Some ((Some (r))))" "((?sel_any) (aa) (\)) = (deref_oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (in_post_state) (reconst_basetype) (e) (\))" apply(insert that, simp add: deref_oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_def) apply(case_tac "(heap ((in_post_state (\))) (e))", simp add: invalid_def bot_option_def, simp) proof - fix aaa show "(case aaa of (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (obj)) \ (reconst_basetype (obj) (\)) | _ \ (invalid (\))) = (Some ((Some (r)))) \ (heap ((in_post_state (\))) (e)) = (Some (aaa)) \ ?t" apply(case_tac "aaa", auto simp: invalid_def bot_option_def image_def ran_def) apply(rule exI[where x = "(in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (r))"], simp add: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_\_def Let_def reconst_basetype_def split: if_split_asm) by(rule) qed apply_end((blast)+) qed apply_end(simp add: foundation16 bot_option_def null_option_def) qed qed qed qed qed apply_end(simp_all) qed lemma is_repr_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_0___bossat_pre : assumes def_dot: "\ \ (\ ((X::\Person) .boss@pre))" shows "(is_represented_in_state (in_pre_state) (X .boss@pre) (Person) (\))" apply(insert defined_mono_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_0___bossat_pre[OF def_dot, simplified foundation16]) apply(case_tac "(X (\))", simp add: bot_option_def) proof - fix a0 show "(X (\)) = (Some (a0)) \ ?thesis" when "(X (\)) \ null" apply(insert that, case_tac "a0", simp add: null_option_def bot_option_def, clarify) proof - fix a show "(X (\)) = (Some ((Some (a)))) \ ?thesis" apply(case_tac "(heap ((in_pre_state (\))) ((oid_of (a))))", simp add: invalid_def bot_option_def) apply(insert def_dot, simp add: dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_0___bossat_pre is_represented_in_state_def select\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__boss_def deref_oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_def in_pre_state_def defined_def OclValid_def false_def true_def invalid_def bot_fun_def split: if_split_asm) proof - fix b show "(X (\)) = (Some ((Some (a)))) \ (heap ((in_pre_state (\))) ((oid_of (a)))) = (Some (b)) \ ?thesis" apply(insert def_dot[simplified foundation16], auto simp: dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_0___bossat_pre is_represented_in_state_def deref_oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_def bot_option_def null_option_def) apply(case_tac "b", simp_all add: invalid_def bot_option_def) proof - fix r typeoid let ?t = "(Some ((Some (r)))) \ (Some o OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_\) ` (ran ((heap ((in_pre_state (\))))))" let ?sel_any = "(select_object_any\<^sub>S\<^sub>e\<^sub>t ((deref_oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (in_pre_state) (reconst_basetype))))" show "(select\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__boss (?sel_any) (typeoid) (\)) = (Some ((Some (r)))) \ ?t" apply(case_tac "typeoid", simp add: select\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n__boss_def) proof - fix opt show "(((case opt of None \ null | (Some (x)) \ ((?sel_any) (x)))) (\)) = (Some ((Some (r)))) \ ?t" apply(case_tac "opt", auto simp: null_fun_def null_option_def bot_option_def) proof - fix aa show "((?sel_any) (aa) (\)) = (Some ((Some (r)))) \ ?t" when "\ \ (\ (((?sel_any) (aa))))" apply(insert that, drule select_object_any_exec\<^sub>S\<^sub>e\<^sub>t[simplified foundation22], erule exE) proof - fix e show "?t" when "((?sel_any) (aa) (\)) = (Some ((Some (r))))" "((?sel_any) (aa) (\)) = (deref_oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (in_pre_state) (reconst_basetype) (e) (\))" apply(insert that, simp add: deref_oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_def) apply(case_tac "(heap ((in_pre_state (\))) (e))", simp add: invalid_def bot_option_def, simp) proof - fix aaa show "(case aaa of (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (obj)) \ (reconst_basetype (obj) (\)) | _ \ (invalid (\))) = (Some ((Some (r)))) \ (heap ((in_pre_state (\))) (e)) = (Some (aaa)) \ ?t" apply(case_tac "aaa", auto simp: invalid_def bot_option_def image_def ran_def) apply(rule exI[where x = "(in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (r))"], simp add: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_\_def Let_def reconst_basetype_def split: if_split_asm) by(rule) qed apply_end((blast)+) qed apply_end(simp add: foundation16 bot_option_def null_option_def) qed qed qed qed qed apply_end(simp_all) qed (* 138 ************************************ 904 + 0 *) (* term Floor1_access.print_access_repr_allinst *) (* 139 ************************************ 904 + 1 *) section \Class Model: Towards the Object Instances\ (* 140 ************************************ 905 + 1 *) text \ The example we are defining in this section comes from the \autoref{fig:Employee-DesignModel-UMLPart-generated-generatededm1_system-states}. \ (* 141 ************************************ 906 + 1 *) text_raw \ \begin{figure} \includegraphics[width=\textwidth]{figures/pre-post.pdf} \caption{(a) pre-state $\sigma_1$ and (b) post-state $\sigma_1'$.} \label{fig:Employee-DesignModel-UMLPart-generated-generatededm1_system-states} \end{figure} \ (* 142 ************************************ 907 + 1 *) text \\ (* 143 ************************************ 908 + 1 *) text_raw \\ (* 144 ************************************ 909 + 1 *) (* term Floor1_examp.print_examp_def_st_defs *) lemmas [simp,code_unfold] = state.defs const_ss (* 145 ************************************ 910 + 1 *) (* term Floor1_astype.print_astype_lemmas_id2 *) lemmas[simp,code_unfold] = OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Planet OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Galaxy OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Galaxy OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_OclAny OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_OclAny OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Planet OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Planet OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Galaxy (* 146 ************************************ 911 + 1 *) section \Instance\ (* 147 ************************************ 912 + 2 *) (* term Floor1_examp.print_examp_instance_defassoc_typecheck_var *) definition "(typecheck_instance_bad_head_on_lhs_P1_X0_X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9_X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8_X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7_X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6_X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5_X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4_X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3_X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2_X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 (P1) (X0) (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9) (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8) (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7) (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6) (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5) (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4) (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3) (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2) (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1)) = ()" definition "typecheck_instance_extra_variables_on_rhs_P1_X0_X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9_X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8_X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7_X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6_X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5_X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4_X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3_X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2_X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 = (\P1 X0 X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9 X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8 X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7 X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6 X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5 X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4 X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3 X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1. (P1 , P1 , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7 , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7 , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2))" (* 148 ************************************ 914 + 12 *) (* term Floor1_examp.print_examp_instance_defassoc *) definition "oid1 = 1" definition "oid2 = 2" definition "oid3 = 3" definition "oid4 = 4" definition "oid5 = 5" definition "oid6 = 6" definition "oid7 = 7" definition "oid8 = 8" definition "oid9 = 9" definition "oid10 = 10" definition "oid11 = 11" definition "inst_assoc1 = (\oid_class to_from oid. ((case (deref_assocs_list ((to_from::oid list list \ oid list \ oid list)) ((oid::oid)) ((the ((((map_of_list ([(oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_0___boss , (List.map ((\(x , y). [x , y]) o switch\<^sub>2_01) ([[[oid7] , [oid7]] , [[oid6] , [oid7]] , [[oid2] , [oid2]] , [[oid1] , [oid2]]])))]))) ((oid_class::oid))))))) of Nil \ None | l \ (Some (l)))::oid list option))" (* 149 ************************************ 926 + 22 *) (* term Floor1_examp.print_examp_instance *) definition "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n inst_assoc = (mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n ((mk\\\\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (oid1) (None) (None) (None) (None) (None))) (((inst_assoc) (oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_0___boss) (switch\<^sub>2_01) (oid1))) (\1300\))" definition "(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1::\Person) = ((\_. \\(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (inst_assoc1))\\))" definition "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n inst_assoc = (mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n ((mk\\\\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (oid2) (None) (None) (None) (None) (None))) (((inst_assoc) (oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_0___boss) (switch\<^sub>2_01) (oid2))) (\1800\))" definition "(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2::\Person) = ((\_. \\(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (inst_assoc1))\\))" definition "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n inst_assoc = (mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n ((mk\\\\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (oid3) (None) (None) (None) (None) (None))) (((inst_assoc) (oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_0___boss) (switch\<^sub>2_01) (oid3))) (None))" definition "(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3::\Person) = ((\_. \\(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (inst_assoc1))\\))" definition "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n inst_assoc = (mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n ((mk\\\\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (oid4) (None) (None) (None) (None) (None))) (((inst_assoc) (oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_0___boss) (switch\<^sub>2_01) (oid4))) (\2900\))" definition "(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4::\Person) = ((\_. \\(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (inst_assoc1))\\))" definition "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n inst_assoc = (mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n ((mk\\\\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (oid5) (None) (None) (None) (None) (None))) (((inst_assoc) (oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_0___boss) (switch\<^sub>2_01) (oid5))) (\3500\))" definition "(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5::\Person) = ((\_. \\(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (inst_assoc1))\\))" definition "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n inst_assoc = (mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n ((mk\\\\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (oid6) (None) (None) (None) (None) (None))) (((inst_assoc) (oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_0___boss) (switch\<^sub>2_01) (oid6))) (\2500\))" definition "(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6::\Person) = ((\_. \\(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (inst_assoc1))\\))" definition "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y inst_assoc = (mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n ((mk\\\\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (oid7) (None) (None) (None) (None) (None))) (((inst_assoc) (oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_0___boss) (switch\<^sub>2_01) (oid7))) (\3200\))" definition "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7 = ((((\_. \\(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y (inst_assoc1))\\)::\Person)) .oclAsType(OclAny))" definition "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y inst_assoc = (mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y ((mk\\\\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y (oid8))))" definition "(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8::\OclAny) = ((\_. \\(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y (inst_assoc1))\\))" definition "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n inst_assoc = (mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n ((mk\\\\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (oid9) (None) (None) (None) (None) (None))) (((inst_assoc) (oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_0___boss) (switch\<^sub>2_01) (oid9))) (\0\))" definition "(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9::\Person) = ((\_. \\(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (inst_assoc1))\\))" definition "X0\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n inst_assoc = (mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n ((mk\\\\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (oid10) (None) (None) (None) (None) (\[[oid11]]\))) (((inst_assoc) (oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_0___boss) (switch\<^sub>2_01) (oid10))) (None))" definition "(X0::\Person) = ((\_. \\(X0\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (inst_assoc1))\\))" definition "P1\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t inst_assoc = (mk\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t ((mk\\\\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (oid11) (None) (None) (\[[oid11] , [oid11]]\))) (None) (None))" definition "(P1::\Planet) = ((\_. \\(P1\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t (inst_assoc1))\\))" (* 150 ************************************ 948 + 1 *) (* term Floor1_examp.print_examp_instance_defassoc_typecheck *) ML \(Ty'.check ([(META.Writeln , "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .boss \ Set{ X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 }") , (META.Writeln , "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 /* unnamed attribute */ \ Set{}") , (META.Writeln , "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 .boss \ Set{ X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 }") , (META.Writeln , "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 /* unnamed attribute */ \ Set{ X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 }") , (META.Writeln , "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3 .boss \ Set{}") , (META.Writeln , "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3 /* unnamed attribute */ \ Set{}") , (META.Writeln , "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4 .boss \ Set{}") , (META.Writeln , "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4 /* unnamed attribute */ \ Set{}") , (META.Writeln , "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5 .boss \ Set{}") , (META.Writeln , "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5 /* unnamed attribute */ \ Set{}") , (META.Writeln , "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6 .boss \ Set{ X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7 }") , (META.Writeln , "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6 /* unnamed attribute */ \ Set{}") , (META.Writeln , "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7 .boss \ Set{ X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7 }") , (META.Writeln , "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7 /* unnamed attribute */ \ Set{ X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6 , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7 }") , (META.Writeln , "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9 .boss \ Set{}") , (META.Writeln , "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9 /* unnamed attribute */ \ Set{}") , (META.Writeln , "X0 .boss \ Set{}") , (META.Writeln , "X0 /* unnamed attribute */ \ Set{}")]) (" error(s)"))\ (* 151 ************************************ 949 + 1 *) section \Instance\ (* 152 ************************************ 950 + 2 *) (* term Floor1_examp.print_examp_instance_defassoc_typecheck_var *) definition "(typecheck_instance_bad_head_on_lhs_\\<^sub>1_object4_\\<^sub>1_object2_\\<^sub>1_object1_\\<^sub>1_object0 (\\<^sub>1_object4) (\\<^sub>1_object2) (\\<^sub>1_object1) (\\<^sub>1_object0)) = ()" definition "typecheck_instance_extra_variables_on_rhs_\\<^sub>1_object4_\\<^sub>1_object2_\\<^sub>1_object1_\\<^sub>1_object0 = (\\\<^sub>1_object4 \\<^sub>1_object2 \\<^sub>1_object1 \\<^sub>1_object0. (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6 , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4 , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5 , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1))" (* 153 ************************************ 952 + 1 *) (* term Floor1_examp.print_examp_instance_defassoc *) definition "inst_assoc12 = (\oid_class to_from oid. ((case (deref_assocs_list ((to_from::oid list list \ oid list \ oid list)) ((oid::oid)) ((the ((((map_of_list ([(oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_0___boss , (List.map ((\(x , y). [x , y]) o switch\<^sub>2_01) ([[[oid6] , [oid4]] , [[oid4] , [oid5]] , [[oid1] , [oid2]]])))]))) ((oid_class::oid))))))) of Nil \ None | l \ (Some (l)))::oid list option))" (* 154 ************************************ 953 + 8 *) (* term Floor1_examp.print_examp_instance *) definition "\\<^sub>1_object0\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n inst_assoc = (mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n ((mk\\\\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (oid1) (None) (None) (None) (None) (None))) (((inst_assoc) (oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_0___boss) (switch\<^sub>2_01) (oid1))) (\1000\))" definition "(\\<^sub>1_object0::\Person) = ((\_. \\(\\<^sub>1_object0\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (inst_assoc12))\\))" definition "\\<^sub>1_object1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n inst_assoc = (mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n ((mk\\\\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (oid2) (None) (None) (None) (None) (None))) (((inst_assoc) (oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_0___boss) (switch\<^sub>2_01) (oid2))) (\1200\))" definition "(\\<^sub>1_object1::\Person) = ((\_. \\(\\<^sub>1_object1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (inst_assoc12))\\))" definition "\\<^sub>1_object2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n inst_assoc = (mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n ((mk\\\\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (oid4) (None) (None) (None) (None) (None))) (((inst_assoc) (oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_0___boss) (switch\<^sub>2_01) (oid4))) (\2600\))" definition "(\\<^sub>1_object2::\Person) = ((\_. \\(\\<^sub>1_object2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (inst_assoc12))\\))" definition "\\<^sub>1_object4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n inst_assoc = (mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n ((mk\\\\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (oid6) (None) (None) (None) (None) (None))) (((inst_assoc) (oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_0___boss) (switch\<^sub>2_01) (oid6))) (\2300\))" definition "(\\<^sub>1_object4::\Person) = ((\_. \\(\\<^sub>1_object4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (inst_assoc12))\\))" (* 155 ************************************ 961 + 1 *) (* term Floor1_examp.print_examp_instance_defassoc_typecheck *) ML \(Ty'.check ([(META.Writeln , "\\<^sub>1_object0 .boss \ Set{ \\<^sub>1_object1 }") , (META.Writeln , "\\<^sub>1_object0 /* unnamed attribute */ \ Set{}") , (META.Writeln , "\\<^sub>1_object1 .boss \ Set{}") , (META.Writeln , "\\<^sub>1_object1 /* unnamed attribute */ \ Set{ \\<^sub>1_object0 }") , (META.Writeln , "\\<^sub>1_object2 .boss \ Set{ /*5*/ }") , (META.Writeln , "\\<^sub>1_object2 /* unnamed attribute */ \ Set{ \\<^sub>1_object4 }") , (META.Writeln , "\\<^sub>1_object4 .boss \ Set{ \\<^sub>1_object2 }") , (META.Writeln , "\\<^sub>1_object4 /* unnamed attribute */ \ Set{}")]) (" error(s)"))\ (* 156 ************************************ 962 + 1 *) section \State (Floor 2)\ (* 157 ************************************ 963 + 1 *) locale state_\\<^sub>1 = fixes "oid1" :: "nat" fixes "oid2" :: "nat" fixes "oid4" :: "nat" fixes "oid5" :: "nat" fixes "oid6" :: "nat" fixes "oid9" :: "nat" assumes distinct_oid: "(distinct ([oid1 , oid2 , oid4 , oid5 , oid6 , oid9]))" fixes "\\<^sub>1_object0\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" :: "ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" fixes "\\<^sub>1_object0" :: "\Person" assumes \\<^sub>1_object0_def: "\\<^sub>1_object0 = (\_. \\\\<^sub>1_object0\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)" fixes "\\<^sub>1_object1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" :: "ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" fixes "\\<^sub>1_object1" :: "\Person" assumes \\<^sub>1_object1_def: "\\<^sub>1_object1 = (\_. \\\\<^sub>1_object1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)" fixes "\\<^sub>1_object2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" :: "ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" fixes "\\<^sub>1_object2" :: "\Person" assumes \\<^sub>1_object2_def: "\\<^sub>1_object2 = (\_. \\\\<^sub>1_object2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)" fixes "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" :: "ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" fixes "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5" :: "\Person" assumes X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5_def: "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5 = (\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)" fixes "\\<^sub>1_object4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" :: "ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" fixes "\\<^sub>1_object4" :: "\Person" assumes \\<^sub>1_object4_def: "\\<^sub>1_object4 = (\_. \\\\<^sub>1_object4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)" fixes "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" :: "ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" fixes "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9" :: "\Person" assumes X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9_def: "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9 = (\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)" begin definition "\\<^sub>1 = (state.make ((Map.empty (oid1 \ (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object0\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))) (oid2 \ (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))) (oid4 \ (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))) (oid5 \ (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))) (oid6 \ (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))) (oid9 \ (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))) ((map_of_list ([(oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_0___boss , (List.map ((\(x , y). [x , y]) o switch\<^sub>2_01) ([[[oid1] , [oid2]] , [[oid4] , [oid5]] , [[oid6] , [oid4]]])))]))))" lemma dom_\\<^sub>1 : "(dom ((heap (\\<^sub>1)))) = {oid1 , oid2 , oid4 , oid5 , oid6 , oid9}" by(auto simp: \\<^sub>1_def) lemmas[simp,code_unfold] = dom_\\<^sub>1 lemma perm_\\<^sub>1 : "\\<^sub>1 = (state.make ((Map.empty (oid9 \ (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))) (oid6 \ (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))) (oid5 \ (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))) (oid4 \ (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))) (oid2 \ (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))) (oid1 \ (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object0\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))) ((assocs (\\<^sub>1))))" apply(simp add: \\<^sub>1_def) apply(subst (1) fun_upd_twist, metis distinct_oid distinct_length_2_or_more) apply(subst (2) fun_upd_twist, metis distinct_oid distinct_length_2_or_more) apply(subst (1) fun_upd_twist, metis distinct_oid distinct_length_2_or_more) apply(subst (3) fun_upd_twist, metis distinct_oid distinct_length_2_or_more) apply(subst (2) fun_upd_twist, metis distinct_oid distinct_length_2_or_more) apply(subst (1) fun_upd_twist, metis distinct_oid distinct_length_2_or_more) apply(subst (4) fun_upd_twist, metis distinct_oid distinct_length_2_or_more) apply(subst (3) fun_upd_twist, metis distinct_oid distinct_length_2_or_more) apply(subst (2) fun_upd_twist, metis distinct_oid distinct_length_2_or_more) apply(subst (1) fun_upd_twist, metis distinct_oid distinct_length_2_or_more) apply(subst (5) fun_upd_twist, metis distinct_oid distinct_length_2_or_more) apply(subst (4) fun_upd_twist, metis distinct_oid distinct_length_2_or_more) apply(subst (3) fun_upd_twist, metis distinct_oid distinct_length_2_or_more) apply(subst (2) fun_upd_twist, metis distinct_oid distinct_length_2_or_more) apply(subst (1) fun_upd_twist, metis distinct_oid distinct_length_2_or_more) by(simp) lemma \\<^sub>1_OclAllInstances_generic_exec_Person : assumes [simp]: "(Person ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object0\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Person ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Person ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Person ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Person ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Person ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(\a. (pre_post ((mk (a)))) = a)" shows "(mk (\\<^sub>1)) \ (OclAllInstances_generic (pre_post) (Person)) \ Set{\\<^sub>1_object0 , \\<^sub>1_object1 , \\<^sub>1_object2 , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5 , \\<^sub>1_object4 , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9}" apply(subst perm_\\<^sub>1) apply(simp only: state.make_def \\<^sub>1_object0_def \\<^sub>1_object1_def \\<^sub>1_object2_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5_def \\<^sub>1_object4_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9_def) apply(subst state_update_vs_allInstances_generic_tc, simp, simp, (metis distinct_oid distinct_length_2_or_more)?, simp only: assms, blast, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp, simp, simp, rule OclIncluding_cong, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def) apply(subst state_update_vs_allInstances_generic_tc, simp, simp, (metis distinct_oid distinct_length_2_or_more)?, simp only: assms, blast, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp, simp, simp, rule OclIncluding_cong, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def) apply(subst state_update_vs_allInstances_generic_tc, simp, simp, (metis distinct_oid distinct_length_2_or_more)?, simp only: assms, blast, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp, simp, simp, rule OclIncluding_cong, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def) apply(subst state_update_vs_allInstances_generic_tc, simp, simp, (metis distinct_oid distinct_length_2_or_more)?, simp only: assms, blast, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp, simp, simp, rule OclIncluding_cong, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def) apply(subst state_update_vs_allInstances_generic_tc, simp, simp, (metis distinct_oid distinct_length_2_or_more)?, simp only: assms, blast, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp, simp, simp, rule OclIncluding_cong, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def) apply(subst state_update_vs_allInstances_generic_tc, simp, simp, (metis distinct_oid distinct_length_2_or_more)?, simp only: assms, blast, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp, simp, simp, rule OclIncluding_cong, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def) apply(rule state_update_vs_allInstances_generic_empty) by(simp_all only: assms, (simp_all add: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_\_def)?) lemma \\<^sub>1_OclAllInstances_at_post_exec_Person : assumes [simp]: "(Person ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object0\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Person ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Person ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Person ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Person ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Person ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" shows "(st , \\<^sub>1) \ (OclAllInstances_at_post (Person)) \ Set{\\<^sub>1_object0 , \\<^sub>1_object1 , \\<^sub>1_object2 , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5 , \\<^sub>1_object4 , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9}" unfolding OclAllInstances_at_post_def by(rule \\<^sub>1_OclAllInstances_generic_exec_Person, simp_all only: assms, simp_all) lemma \\<^sub>1_OclAllInstances_at_pre_exec_Person : assumes [simp]: "(Person ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object0\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Person ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Person ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Person ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Person ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Person ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" shows "(\\<^sub>1 , st) \ (OclAllInstances_at_pre (Person)) \ Set{\\<^sub>1_object0 , \\<^sub>1_object1 , \\<^sub>1_object2 , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5 , \\<^sub>1_object4 , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9}" unfolding OclAllInstances_at_pre_def by(rule \\<^sub>1_OclAllInstances_generic_exec_Person, simp_all only: assms, simp_all) lemma \\<^sub>1_OclAllInstances_generic_exec_Planet : assumes [simp]: "(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object0\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(\_. \(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object0\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\\\<^sub>1_object0\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Planet))" assumes [simp]: "(\_. \(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\\\<^sub>1_object1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Planet))" assumes [simp]: "(\_. \(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\\\<^sub>1_object2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Planet))" assumes [simp]: "(\_. \(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Planet))" assumes [simp]: "(\_. \(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\\\<^sub>1_object4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Planet))" assumes [simp]: "(\_. \(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Planet))" assumes [simp]: "(\a. (pre_post ((mk (a)))) = a)" shows "(mk (\\<^sub>1)) \ (OclAllInstances_generic (pre_post) (Planet)) \ Set{\\<^sub>1_object0 .oclAsType(Planet) , \\<^sub>1_object1 .oclAsType(Planet) , \\<^sub>1_object2 .oclAsType(Planet) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5 .oclAsType(Planet) , \\<^sub>1_object4 .oclAsType(Planet) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9 .oclAsType(Planet)}" apply(subst perm_\\<^sub>1) apply(simp only: state.make_def \\<^sub>1_object0_def \\<^sub>1_object1_def \\<^sub>1_object2_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5_def \\<^sub>1_object4_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9_def) apply(subst state_update_vs_allInstances_generic_tc, simp, simp, (metis distinct_oid distinct_length_2_or_more)?, simp only: assms, blast, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp del: OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person, simp del: OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person, simp, rule OclIncluding_cong, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def) apply(subst state_update_vs_allInstances_generic_tc, simp, simp, (metis distinct_oid distinct_length_2_or_more)?, simp only: assms, blast, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp del: OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person, simp del: OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person, simp, rule OclIncluding_cong, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def) apply(subst state_update_vs_allInstances_generic_tc, simp, simp, (metis distinct_oid distinct_length_2_or_more)?, simp only: assms, blast, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp del: OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person, simp del: OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person, simp, rule OclIncluding_cong, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def) apply(subst state_update_vs_allInstances_generic_tc, simp, simp, (metis distinct_oid distinct_length_2_or_more)?, simp only: assms, blast, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp del: OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person, simp del: OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person, simp, rule OclIncluding_cong, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def) apply(subst state_update_vs_allInstances_generic_tc, simp, simp, (metis distinct_oid distinct_length_2_or_more)?, simp only: assms, blast, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp del: OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person, simp del: OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person, simp, rule OclIncluding_cong, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def) apply(subst state_update_vs_allInstances_generic_tc, simp, simp, (metis distinct_oid distinct_length_2_or_more)?, simp only: assms, blast, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp del: OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person, simp del: OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person, simp, rule OclIncluding_cong, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def) apply(rule state_update_vs_allInstances_generic_empty) by(simp_all only: assms, (simp_all add: OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_\_def)?) lemma \\<^sub>1_OclAllInstances_at_post_exec_Planet : assumes [simp]: "(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object0\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(\_. \(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object0\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\\\<^sub>1_object0\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Planet))" assumes [simp]: "(\_. \(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\\\<^sub>1_object1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Planet))" assumes [simp]: "(\_. \(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\\\<^sub>1_object2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Planet))" assumes [simp]: "(\_. \(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Planet))" assumes [simp]: "(\_. \(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\\\<^sub>1_object4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Planet))" assumes [simp]: "(\_. \(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Planet))" shows "(st , \\<^sub>1) \ (OclAllInstances_at_post (Planet)) \ Set{\\<^sub>1_object0 .oclAsType(Planet) , \\<^sub>1_object1 .oclAsType(Planet) , \\<^sub>1_object2 .oclAsType(Planet) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5 .oclAsType(Planet) , \\<^sub>1_object4 .oclAsType(Planet) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9 .oclAsType(Planet)}" unfolding OclAllInstances_at_post_def by(rule \\<^sub>1_OclAllInstances_generic_exec_Planet, simp_all only: assms, simp_all) lemma \\<^sub>1_OclAllInstances_at_pre_exec_Planet : assumes [simp]: "(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object0\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(\_. \(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object0\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\\\<^sub>1_object0\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Planet))" assumes [simp]: "(\_. \(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\\\<^sub>1_object1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Planet))" assumes [simp]: "(\_. \(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\\\<^sub>1_object2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Planet))" assumes [simp]: "(\_. \(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Planet))" assumes [simp]: "(\_. \(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\\\<^sub>1_object4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Planet))" assumes [simp]: "(\_. \(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Planet))" shows "(\\<^sub>1 , st) \ (OclAllInstances_at_pre (Planet)) \ Set{\\<^sub>1_object0 .oclAsType(Planet) , \\<^sub>1_object1 .oclAsType(Planet) , \\<^sub>1_object2 .oclAsType(Planet) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5 .oclAsType(Planet) , \\<^sub>1_object4 .oclAsType(Planet) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9 .oclAsType(Planet)}" unfolding OclAllInstances_at_pre_def by(rule \\<^sub>1_OclAllInstances_generic_exec_Planet, simp_all only: assms, simp_all) lemma \\<^sub>1_OclAllInstances_generic_exec_Galaxy : assumes [simp]: "(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object0\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(\_. \(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object0\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\\\<^sub>1_object0\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Galaxy))" assumes [simp]: "(\_. \(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\\\<^sub>1_object1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Galaxy))" assumes [simp]: "(\_. \(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\\\<^sub>1_object2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Galaxy))" assumes [simp]: "(\_. \(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Galaxy))" assumes [simp]: "(\_. \(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\\\<^sub>1_object4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Galaxy))" assumes [simp]: "(\_. \(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Galaxy))" assumes [simp]: "(\a. (pre_post ((mk (a)))) = a)" shows "(mk (\\<^sub>1)) \ (OclAllInstances_generic (pre_post) (Galaxy)) \ Set{\\<^sub>1_object0 .oclAsType(Galaxy) , \\<^sub>1_object1 .oclAsType(Galaxy) , \\<^sub>1_object2 .oclAsType(Galaxy) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5 .oclAsType(Galaxy) , \\<^sub>1_object4 .oclAsType(Galaxy) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9 .oclAsType(Galaxy)}" apply(subst perm_\\<^sub>1) apply(simp only: state.make_def \\<^sub>1_object0_def \\<^sub>1_object1_def \\<^sub>1_object2_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5_def \\<^sub>1_object4_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9_def) apply(subst state_update_vs_allInstances_generic_tc, simp, simp, (metis distinct_oid distinct_length_2_or_more)?, simp only: assms, blast, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp del: OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person, simp del: OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person, simp, rule OclIncluding_cong, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def) apply(subst state_update_vs_allInstances_generic_tc, simp, simp, (metis distinct_oid distinct_length_2_or_more)?, simp only: assms, blast, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp del: OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person, simp del: OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person, simp, rule OclIncluding_cong, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def) apply(subst state_update_vs_allInstances_generic_tc, simp, simp, (metis distinct_oid distinct_length_2_or_more)?, simp only: assms, blast, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp del: OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person, simp del: OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person, simp, rule OclIncluding_cong, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def) apply(subst state_update_vs_allInstances_generic_tc, simp, simp, (metis distinct_oid distinct_length_2_or_more)?, simp only: assms, blast, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp del: OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person, simp del: OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person, simp, rule OclIncluding_cong, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def) apply(subst state_update_vs_allInstances_generic_tc, simp, simp, (metis distinct_oid distinct_length_2_or_more)?, simp only: assms, blast, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp del: OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person, simp del: OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person, simp, rule OclIncluding_cong, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def) apply(subst state_update_vs_allInstances_generic_tc, simp, simp, (metis distinct_oid distinct_length_2_or_more)?, simp only: assms, blast, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp del: OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person, simp del: OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person, simp, rule OclIncluding_cong, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def) apply(rule state_update_vs_allInstances_generic_empty) by(simp_all only: assms, (simp_all add: OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_\_def)?) lemma \\<^sub>1_OclAllInstances_at_post_exec_Galaxy : assumes [simp]: "(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object0\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(\_. \(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object0\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\\\<^sub>1_object0\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Galaxy))" assumes [simp]: "(\_. \(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\\\<^sub>1_object1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Galaxy))" assumes [simp]: "(\_. \(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\\\<^sub>1_object2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Galaxy))" assumes [simp]: "(\_. \(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Galaxy))" assumes [simp]: "(\_. \(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\\\<^sub>1_object4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Galaxy))" assumes [simp]: "(\_. \(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Galaxy))" shows "(st , \\<^sub>1) \ (OclAllInstances_at_post (Galaxy)) \ Set{\\<^sub>1_object0 .oclAsType(Galaxy) , \\<^sub>1_object1 .oclAsType(Galaxy) , \\<^sub>1_object2 .oclAsType(Galaxy) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5 .oclAsType(Galaxy) , \\<^sub>1_object4 .oclAsType(Galaxy) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9 .oclAsType(Galaxy)}" unfolding OclAllInstances_at_post_def by(rule \\<^sub>1_OclAllInstances_generic_exec_Galaxy, simp_all only: assms, simp_all) lemma \\<^sub>1_OclAllInstances_at_pre_exec_Galaxy : assumes [simp]: "(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object0\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(\_. \(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object0\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\\\<^sub>1_object0\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Galaxy))" assumes [simp]: "(\_. \(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\\\<^sub>1_object1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Galaxy))" assumes [simp]: "(\_. \(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\\\<^sub>1_object2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Galaxy))" assumes [simp]: "(\_. \(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Galaxy))" assumes [simp]: "(\_. \(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\\\<^sub>1_object4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Galaxy))" assumes [simp]: "(\_. \(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Galaxy))" shows "(\\<^sub>1 , st) \ (OclAllInstances_at_pre (Galaxy)) \ Set{\\<^sub>1_object0 .oclAsType(Galaxy) , \\<^sub>1_object1 .oclAsType(Galaxy) , \\<^sub>1_object2 .oclAsType(Galaxy) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5 .oclAsType(Galaxy) , \\<^sub>1_object4 .oclAsType(Galaxy) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9 .oclAsType(Galaxy)}" unfolding OclAllInstances_at_pre_def by(rule \\<^sub>1_OclAllInstances_generic_exec_Galaxy, simp_all only: assms, simp_all) lemma \\<^sub>1_OclAllInstances_generic_exec_OclAny : assumes [simp]: "(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object0\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(\_. \(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object0\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\\\<^sub>1_object0\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(OclAny))" assumes [simp]: "(\_. \(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\\\<^sub>1_object1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(OclAny))" assumes [simp]: "(\_. \(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\\\<^sub>1_object2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(OclAny))" assumes [simp]: "(\_. \(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(OclAny))" assumes [simp]: "(\_. \(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\\\<^sub>1_object4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(OclAny))" assumes [simp]: "(\_. \(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(OclAny))" assumes [simp]: "(\a. (pre_post ((mk (a)))) = a)" shows "(mk (\\<^sub>1)) \ (OclAllInstances_generic (pre_post) (OclAny)) \ Set{\\<^sub>1_object0 .oclAsType(OclAny) , \\<^sub>1_object1 .oclAsType(OclAny) , \\<^sub>1_object2 .oclAsType(OclAny) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5 .oclAsType(OclAny) , \\<^sub>1_object4 .oclAsType(OclAny) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9 .oclAsType(OclAny)}" apply(subst perm_\\<^sub>1) apply(simp only: state.make_def \\<^sub>1_object0_def \\<^sub>1_object1_def \\<^sub>1_object2_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5_def \\<^sub>1_object4_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9_def) apply(subst state_update_vs_allInstances_generic_tc, simp, simp, (metis distinct_oid distinct_length_2_or_more)?, simp only: assms, blast, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp del: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person, simp del: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person, simp, rule OclIncluding_cong, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def) apply(subst state_update_vs_allInstances_generic_tc, simp, simp, (metis distinct_oid distinct_length_2_or_more)?, simp only: assms, blast, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp del: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person, simp del: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person, simp, rule OclIncluding_cong, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def) apply(subst state_update_vs_allInstances_generic_tc, simp, simp, (metis distinct_oid distinct_length_2_or_more)?, simp only: assms, blast, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp del: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person, simp del: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person, simp, rule OclIncluding_cong, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def) apply(subst state_update_vs_allInstances_generic_tc, simp, simp, (metis distinct_oid distinct_length_2_or_more)?, simp only: assms, blast, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp del: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person, simp del: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person, simp, rule OclIncluding_cong, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def) apply(subst state_update_vs_allInstances_generic_tc, simp, simp, (metis distinct_oid distinct_length_2_or_more)?, simp only: assms, blast, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp del: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person, simp del: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person, simp, rule OclIncluding_cong, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def) apply(subst state_update_vs_allInstances_generic_tc, simp, simp, (metis distinct_oid distinct_length_2_or_more)?, simp only: assms, blast, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp del: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person, simp del: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person, simp, rule OclIncluding_cong, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def) apply(rule state_update_vs_allInstances_generic_empty) by(simp_all only: assms, (simp_all add: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\_def)?) lemma \\<^sub>1_OclAllInstances_at_post_exec_OclAny : assumes [simp]: "(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object0\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(\_. \(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object0\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\\\<^sub>1_object0\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(OclAny))" assumes [simp]: "(\_. \(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\\\<^sub>1_object1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(OclAny))" assumes [simp]: "(\_. \(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\\\<^sub>1_object2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(OclAny))" assumes [simp]: "(\_. \(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(OclAny))" assumes [simp]: "(\_. \(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\\\<^sub>1_object4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(OclAny))" assumes [simp]: "(\_. \(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(OclAny))" shows "(st , \\<^sub>1) \ (OclAllInstances_at_post (OclAny)) \ Set{\\<^sub>1_object0 .oclAsType(OclAny) , \\<^sub>1_object1 .oclAsType(OclAny) , \\<^sub>1_object2 .oclAsType(OclAny) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5 .oclAsType(OclAny) , \\<^sub>1_object4 .oclAsType(OclAny) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9 .oclAsType(OclAny)}" unfolding OclAllInstances_at_post_def by(rule \\<^sub>1_OclAllInstances_generic_exec_OclAny, simp_all only: assms, simp_all) lemma \\<^sub>1_OclAllInstances_at_pre_exec_OclAny : assumes [simp]: "(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object0\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(\_. \(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object0\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\\\<^sub>1_object0\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(OclAny))" assumes [simp]: "(\_. \(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\\\<^sub>1_object1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(OclAny))" assumes [simp]: "(\_. \(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\\\<^sub>1_object2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(OclAny))" assumes [simp]: "(\_. \(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(OclAny))" assumes [simp]: "(\_. \(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\\\<^sub>1_object4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(OclAny))" assumes [simp]: "(\_. \(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(OclAny))" shows "(\\<^sub>1 , st) \ (OclAllInstances_at_pre (OclAny)) \ Set{\\<^sub>1_object0 .oclAsType(OclAny) , \\<^sub>1_object1 .oclAsType(OclAny) , \\<^sub>1_object2 .oclAsType(OclAny) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5 .oclAsType(OclAny) , \\<^sub>1_object4 .oclAsType(OclAny) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9 .oclAsType(OclAny)}" unfolding OclAllInstances_at_pre_def by(rule \\<^sub>1_OclAllInstances_generic_exec_OclAny, simp_all only: assms, simp_all) ML \(Ty'.check ([]) (" error(s)"))\ end (* 158 ************************************ 964 + 1 *) (* term Floor2_examp.print_examp_def_st_def_interp *) definition "(state_interpretation_\\<^sub>1 (\)) = (state_\\<^sub>1 (oid1) (oid2) (oid4) (oid5) (oid6) (oid9) (\\(\\<^sub>1_object0 (\))\\) (\\<^sub>1_object0) (\\(\\<^sub>1_object1 (\))\\) (\\<^sub>1_object1) (\\(\\<^sub>1_object2 (\))\\) (\\<^sub>1_object2) (\\(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5 (\))\\) (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5) (\\(\\<^sub>1_object4 (\))\\) (\\<^sub>1_object4) (\\(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9 (\))\\) (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9))" (* 159 ************************************ 965 + 1 *) section \State (Floor 2)\ (* 160 ************************************ 966 + 1 *) locale state_\\<^sub>1' = fixes "oid1" :: "nat" fixes "oid2" :: "nat" fixes "oid3" :: "nat" fixes "oid4" :: "nat" fixes "oid6" :: "nat" fixes "oid7" :: "nat" fixes "oid8" :: "nat" fixes "oid9" :: "nat" assumes distinct_oid: "(distinct ([oid1 , oid2 , oid3 , oid4 , oid6 , oid7 , oid8 , oid9]))" fixes "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" :: "ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" fixes "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1" :: "\Person" assumes X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1_def: "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 = (\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)" fixes "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" :: "ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" fixes "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2" :: "\Person" assumes X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2_def: "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 = (\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)" fixes "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" :: "ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" fixes "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3" :: "\Person" assumes X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3_def: "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3 = (\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)" fixes "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" :: "ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" fixes "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4" :: "\Person" assumes X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4_def: "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4 = (\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)" fixes "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" :: "ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" fixes "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6" :: "\Person" assumes X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6_def: "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6 = (\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)" fixes "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y" :: "ty\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y" fixes "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7" :: "\OclAny" assumes X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7_def: "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7 = (\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y\\)" fixes "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y" :: "ty\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y" fixes "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8" :: "\OclAny" assumes X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8_def: "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8 = (\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y\\)" fixes "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" :: "ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" fixes "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9" :: "\Person" assumes X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9_def: "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9 = (\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)" begin definition "\\<^sub>1' = (state.make ((Map.empty (oid1 \ (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))) (oid2 \ (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))) (oid3 \ (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))) (oid4 \ (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))) (oid6 \ (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))) (oid7 \ (in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y))) (oid8 \ (in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y))) (oid9 \ (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))) ((map_of_list ([(oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_0___boss , (List.map ((\(x , y). [x , y]) o switch\<^sub>2_01) ([[[oid1] , [oid2]] , [[oid2] , [oid2]] , [[oid6] , [oid7]] , [[oid7] , [oid7]]])))]))))" lemma dom_\\<^sub>1' : "(dom ((heap (\\<^sub>1')))) = {oid1 , oid2 , oid3 , oid4 , oid6 , oid7 , oid8 , oid9}" by(auto simp: \\<^sub>1'_def) lemmas[simp,code_unfold] = dom_\\<^sub>1' lemma perm_\\<^sub>1' : "\\<^sub>1' = (state.make ((Map.empty (oid9 \ (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))) (oid8 \ (in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y))) (oid7 \ (in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y))) (oid6 \ (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))) (oid4 \ (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))) (oid3 \ (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))) (oid2 \ (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))) (oid1 \ (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))) ((assocs (\\<^sub>1'))))" apply(simp add: \\<^sub>1'_def) apply(subst (1) fun_upd_twist, metis distinct_oid distinct_length_2_or_more) apply(subst (2) fun_upd_twist, metis distinct_oid distinct_length_2_or_more) apply(subst (1) fun_upd_twist, metis distinct_oid distinct_length_2_or_more) apply(subst (3) fun_upd_twist, metis distinct_oid distinct_length_2_or_more) apply(subst (2) fun_upd_twist, metis distinct_oid distinct_length_2_or_more) apply(subst (1) fun_upd_twist, metis distinct_oid distinct_length_2_or_more) apply(subst (4) fun_upd_twist, metis distinct_oid distinct_length_2_or_more) apply(subst (3) fun_upd_twist, metis distinct_oid distinct_length_2_or_more) apply(subst (2) fun_upd_twist, metis distinct_oid distinct_length_2_or_more) apply(subst (1) fun_upd_twist, metis distinct_oid distinct_length_2_or_more) apply(subst (5) fun_upd_twist, metis distinct_oid distinct_length_2_or_more) apply(subst (4) fun_upd_twist, metis distinct_oid distinct_length_2_or_more) apply(subst (3) fun_upd_twist, metis distinct_oid distinct_length_2_or_more) apply(subst (2) fun_upd_twist, metis distinct_oid distinct_length_2_or_more) apply(subst (1) fun_upd_twist, metis distinct_oid distinct_length_2_or_more) apply(subst (6) fun_upd_twist, metis distinct_oid distinct_length_2_or_more) apply(subst (5) fun_upd_twist, metis distinct_oid distinct_length_2_or_more) apply(subst (4) fun_upd_twist, metis distinct_oid distinct_length_2_or_more) apply(subst (3) fun_upd_twist, metis distinct_oid distinct_length_2_or_more) apply(subst (2) fun_upd_twist, metis distinct_oid distinct_length_2_or_more) apply(subst (1) fun_upd_twist, metis distinct_oid distinct_length_2_or_more) apply(subst (7) fun_upd_twist, metis distinct_oid distinct_length_2_or_more) apply(subst (6) fun_upd_twist, metis distinct_oid distinct_length_2_or_more) apply(subst (5) fun_upd_twist, metis distinct_oid distinct_length_2_or_more) apply(subst (4) fun_upd_twist, metis distinct_oid distinct_length_2_or_more) apply(subst (3) fun_upd_twist, metis distinct_oid distinct_length_2_or_more) apply(subst (2) fun_upd_twist, metis distinct_oid distinct_length_2_or_more) apply(subst (1) fun_upd_twist, metis distinct_oid distinct_length_2_or_more) by(simp) lemma \\<^sub>1'_OclAllInstances_generic_exec_Person : assumes [simp]: "(Person ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Person ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Person ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Person ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Person ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Person ((in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y)))) \ None" assumes [simp]: "(Person ((in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y)))) = None" assumes [simp]: "(Person ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(\_. \(Person ((in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y\\)::\OclAny)) .oclAsType(Person))" assumes [simp]: "(\a. (pre_post ((mk (a)))) = a)" shows "(mk (\\<^sub>1')) \ (OclAllInstances_generic (pre_post) (Person)) \ Set{X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3 , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4 , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6 , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7 .oclAsType(Person) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9}" apply(subst perm_\\<^sub>1') apply(simp only: state.make_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9_def) apply(subst state_update_vs_allInstances_generic_tc, simp, simp, (metis distinct_oid distinct_length_2_or_more)?, simp only: assms, blast, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp del: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny, simp del: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny, simp, rule OclIncluding_cong, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def) apply(subst state_update_vs_allInstances_generic_tc, simp, simp, (metis distinct_oid distinct_length_2_or_more)?, simp only: assms, blast, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp del: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny, simp del: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny, simp, rule OclIncluding_cong, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def) apply(subst state_update_vs_allInstances_generic_tc, simp, simp, (metis distinct_oid distinct_length_2_or_more)?, simp only: assms, blast, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp del: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny, simp del: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny, simp, rule OclIncluding_cong, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def) apply(subst state_update_vs_allInstances_generic_tc, simp, simp, (metis distinct_oid distinct_length_2_or_more)?, simp only: assms, blast, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp del: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny, simp del: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny, simp, rule OclIncluding_cong, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def) apply(subst state_update_vs_allInstances_generic_tc, simp, simp, (metis distinct_oid distinct_length_2_or_more)?, simp only: assms, blast, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp del: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny, simp del: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny, simp, rule OclIncluding_cong, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def) apply(subst state_update_vs_allInstances_generic_tc, simp, simp, (metis distinct_oid distinct_length_2_or_more)?, simp only: assms, blast, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp del: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny, simp del: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny, simp, rule OclIncluding_cong, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def) apply(subst state_update_vs_allInstances_generic_ntc, simp, simp, (metis distinct_oid distinct_length_2_or_more)?, simp only: assms, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp del: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny, simp del: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny, simp) apply(subst state_update_vs_allInstances_generic_tc, simp, simp, (metis distinct_oid distinct_length_2_or_more)?, simp only: assms, blast, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp del: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny, simp del: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny, simp, rule OclIncluding_cong, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def) apply(rule state_update_vs_allInstances_generic_empty) by(simp_all only: assms, (simp_all add: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_\_def)?) lemma \\<^sub>1'_OclAllInstances_at_post_exec_Person : assumes [simp]: "(Person ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Person ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Person ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Person ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Person ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Person ((in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y)))) \ None" assumes [simp]: "(Person ((in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y)))) = None" assumes [simp]: "(Person ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(\_. \(Person ((in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y\\)::\OclAny)) .oclAsType(Person))" shows "(st , \\<^sub>1') \ (OclAllInstances_at_post (Person)) \ Set{X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3 , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4 , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6 , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7 .oclAsType(Person) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9}" unfolding OclAllInstances_at_post_def by(rule \\<^sub>1'_OclAllInstances_generic_exec_Person, simp_all only: assms, simp_all) lemma \\<^sub>1'_OclAllInstances_at_pre_exec_Person : assumes [simp]: "(Person ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Person ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Person ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Person ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Person ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Person ((in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y)))) \ None" assumes [simp]: "(Person ((in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y)))) = None" assumes [simp]: "(Person ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(\_. \(Person ((in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y\\)::\OclAny)) .oclAsType(Person))" shows "(\\<^sub>1' , st) \ (OclAllInstances_at_pre (Person)) \ Set{X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3 , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4 , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6 , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7 .oclAsType(Person) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9}" unfolding OclAllInstances_at_pre_def by(rule \\<^sub>1'_OclAllInstances_generic_exec_Person, simp_all only: assms, simp_all) lemma \\<^sub>1'_OclAllInstances_generic_exec_Planet : assumes [simp]: "(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Planet ((in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y)))) = None" assumes [simp]: "(Planet ((in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y)))) = None" assumes [simp]: "(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(\_. \(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Planet))" assumes [simp]: "(\_. \(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Planet))" assumes [simp]: "(\_. \(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Planet))" assumes [simp]: "(\_. \(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Planet))" assumes [simp]: "(\_. \(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Planet))" assumes [simp]: "(\_. \(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Planet))" assumes [simp]: "(\a. (pre_post ((mk (a)))) = a)" shows "(mk (\\<^sub>1')) \ (OclAllInstances_generic (pre_post) (Planet)) \ Set{X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .oclAsType(Planet) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 .oclAsType(Planet) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3 .oclAsType(Planet) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4 .oclAsType(Planet) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6 .oclAsType(Planet) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9 .oclAsType(Planet)}" apply(subst perm_\\<^sub>1') apply(simp only: state.make_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9_def) apply(subst state_update_vs_allInstances_generic_tc, simp, simp, (metis distinct_oid distinct_length_2_or_more)?, simp only: assms, blast, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp del: OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person, simp del: OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person, simp, rule OclIncluding_cong, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def) apply(subst state_update_vs_allInstances_generic_tc, simp, simp, (metis distinct_oid distinct_length_2_or_more)?, simp only: assms, blast, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp del: OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person, simp del: OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person, simp, rule OclIncluding_cong, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def) apply(subst state_update_vs_allInstances_generic_tc, simp, simp, (metis distinct_oid distinct_length_2_or_more)?, simp only: assms, blast, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp del: OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person, simp del: OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person, simp, rule OclIncluding_cong, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def) apply(subst state_update_vs_allInstances_generic_tc, simp, simp, (metis distinct_oid distinct_length_2_or_more)?, simp only: assms, blast, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp del: OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person, simp del: OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person, simp, rule OclIncluding_cong, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def) apply(subst state_update_vs_allInstances_generic_tc, simp, simp, (metis distinct_oid distinct_length_2_or_more)?, simp only: assms, blast, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp del: OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person, simp del: OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person, simp, rule OclIncluding_cong, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def) apply(subst state_update_vs_allInstances_generic_ntc, simp, simp, (metis distinct_oid distinct_length_2_or_more)?, simp only: assms, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp del: OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person, simp del: OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person, simp) apply(subst state_update_vs_allInstances_generic_ntc, simp, simp, (metis distinct_oid distinct_length_2_or_more)?, simp only: assms, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp del: OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person, simp del: OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person, simp) apply(subst state_update_vs_allInstances_generic_tc, simp, simp, (metis distinct_oid distinct_length_2_or_more)?, simp only: assms, blast, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp del: OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person, simp del: OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_Person, simp, rule OclIncluding_cong, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def) apply(rule state_update_vs_allInstances_generic_empty) by(simp_all only: assms, (simp_all add: OclAsType\<^sub>P\<^sub>l\<^sub>a\<^sub>n\<^sub>e\<^sub>t_\_def)?) lemma \\<^sub>1'_OclAllInstances_at_post_exec_Planet : assumes [simp]: "(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Planet ((in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y)))) = None" assumes [simp]: "(Planet ((in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y)))) = None" assumes [simp]: "(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(\_. \(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Planet))" assumes [simp]: "(\_. \(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Planet))" assumes [simp]: "(\_. \(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Planet))" assumes [simp]: "(\_. \(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Planet))" assumes [simp]: "(\_. \(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Planet))" assumes [simp]: "(\_. \(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Planet))" shows "(st , \\<^sub>1') \ (OclAllInstances_at_post (Planet)) \ Set{X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .oclAsType(Planet) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 .oclAsType(Planet) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3 .oclAsType(Planet) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4 .oclAsType(Planet) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6 .oclAsType(Planet) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9 .oclAsType(Planet)}" unfolding OclAllInstances_at_post_def by(rule \\<^sub>1'_OclAllInstances_generic_exec_Planet, simp_all only: assms, simp_all) lemma \\<^sub>1'_OclAllInstances_at_pre_exec_Planet : assumes [simp]: "(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Planet ((in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y)))) = None" assumes [simp]: "(Planet ((in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y)))) = None" assumes [simp]: "(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(\_. \(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Planet))" assumes [simp]: "(\_. \(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Planet))" assumes [simp]: "(\_. \(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Planet))" assumes [simp]: "(\_. \(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Planet))" assumes [simp]: "(\_. \(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Planet))" assumes [simp]: "(\_. \(Planet ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Planet))" shows "(\\<^sub>1' , st) \ (OclAllInstances_at_pre (Planet)) \ Set{X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .oclAsType(Planet) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 .oclAsType(Planet) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3 .oclAsType(Planet) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4 .oclAsType(Planet) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6 .oclAsType(Planet) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9 .oclAsType(Planet)}" unfolding OclAllInstances_at_pre_def by(rule \\<^sub>1'_OclAllInstances_generic_exec_Planet, simp_all only: assms, simp_all) lemma \\<^sub>1'_OclAllInstances_generic_exec_Galaxy : assumes [simp]: "(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Galaxy ((in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y)))) = None" assumes [simp]: "(Galaxy ((in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y)))) = None" assumes [simp]: "(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(\_. \(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Galaxy))" assumes [simp]: "(\_. \(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Galaxy))" assumes [simp]: "(\_. \(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Galaxy))" assumes [simp]: "(\_. \(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Galaxy))" assumes [simp]: "(\_. \(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Galaxy))" assumes [simp]: "(\_. \(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Galaxy))" assumes [simp]: "(\a. (pre_post ((mk (a)))) = a)" shows "(mk (\\<^sub>1')) \ (OclAllInstances_generic (pre_post) (Galaxy)) \ Set{X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .oclAsType(Galaxy) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 .oclAsType(Galaxy) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3 .oclAsType(Galaxy) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4 .oclAsType(Galaxy) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6 .oclAsType(Galaxy) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9 .oclAsType(Galaxy)}" apply(subst perm_\\<^sub>1') apply(simp only: state.make_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9_def) apply(subst state_update_vs_allInstances_generic_tc, simp, simp, (metis distinct_oid distinct_length_2_or_more)?, simp only: assms, blast, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp del: OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person, simp del: OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person, simp, rule OclIncluding_cong, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def) apply(subst state_update_vs_allInstances_generic_tc, simp, simp, (metis distinct_oid distinct_length_2_or_more)?, simp only: assms, blast, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp del: OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person, simp del: OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person, simp, rule OclIncluding_cong, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def) apply(subst state_update_vs_allInstances_generic_tc, simp, simp, (metis distinct_oid distinct_length_2_or_more)?, simp only: assms, blast, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp del: OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person, simp del: OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person, simp, rule OclIncluding_cong, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def) apply(subst state_update_vs_allInstances_generic_tc, simp, simp, (metis distinct_oid distinct_length_2_or_more)?, simp only: assms, blast, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp del: OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person, simp del: OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person, simp, rule OclIncluding_cong, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def) apply(subst state_update_vs_allInstances_generic_tc, simp, simp, (metis distinct_oid distinct_length_2_or_more)?, simp only: assms, blast, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp del: OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person, simp del: OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person, simp, rule OclIncluding_cong, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def) apply(subst state_update_vs_allInstances_generic_ntc, simp, simp, (metis distinct_oid distinct_length_2_or_more)?, simp only: assms, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp del: OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person, simp del: OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person, simp) apply(subst state_update_vs_allInstances_generic_ntc, simp, simp, (metis distinct_oid distinct_length_2_or_more)?, simp only: assms, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp del: OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person, simp del: OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person, simp) apply(subst state_update_vs_allInstances_generic_tc, simp, simp, (metis distinct_oid distinct_length_2_or_more)?, simp only: assms, blast, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp del: OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person, simp del: OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_Person, simp, rule OclIncluding_cong, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def) apply(rule state_update_vs_allInstances_generic_empty) by(simp_all only: assms, (simp_all add: OclAsType\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_\_def)?) lemma \\<^sub>1'_OclAllInstances_at_post_exec_Galaxy : assumes [simp]: "(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Galaxy ((in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y)))) = None" assumes [simp]: "(Galaxy ((in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y)))) = None" assumes [simp]: "(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(\_. \(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Galaxy))" assumes [simp]: "(\_. \(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Galaxy))" assumes [simp]: "(\_. \(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Galaxy))" assumes [simp]: "(\_. \(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Galaxy))" assumes [simp]: "(\_. \(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Galaxy))" assumes [simp]: "(\_. \(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Galaxy))" shows "(st , \\<^sub>1') \ (OclAllInstances_at_post (Galaxy)) \ Set{X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .oclAsType(Galaxy) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 .oclAsType(Galaxy) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3 .oclAsType(Galaxy) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4 .oclAsType(Galaxy) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6 .oclAsType(Galaxy) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9 .oclAsType(Galaxy)}" unfolding OclAllInstances_at_post_def by(rule \\<^sub>1'_OclAllInstances_generic_exec_Galaxy, simp_all only: assms, simp_all) lemma \\<^sub>1'_OclAllInstances_at_pre_exec_Galaxy : assumes [simp]: "(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(Galaxy ((in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y)))) = None" assumes [simp]: "(Galaxy ((in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y)))) = None" assumes [simp]: "(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(\_. \(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Galaxy))" assumes [simp]: "(\_. \(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Galaxy))" assumes [simp]: "(\_. \(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Galaxy))" assumes [simp]: "(\_. \(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Galaxy))" assumes [simp]: "(\_. \(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Galaxy))" assumes [simp]: "(\_. \(Galaxy ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(Galaxy))" shows "(\\<^sub>1' , st) \ (OclAllInstances_at_pre (Galaxy)) \ Set{X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .oclAsType(Galaxy) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 .oclAsType(Galaxy) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3 .oclAsType(Galaxy) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4 .oclAsType(Galaxy) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6 .oclAsType(Galaxy) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9 .oclAsType(Galaxy)}" unfolding OclAllInstances_at_pre_def by(rule \\<^sub>1'_OclAllInstances_generic_exec_Galaxy, simp_all only: assms, simp_all) lemma \\<^sub>1'_OclAllInstances_generic_exec_OclAny : assumes [simp]: "(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(OclAny ((in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y)))) \ None" assumes [simp]: "(OclAny ((in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y)))) \ None" assumes [simp]: "(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(\_. \(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(OclAny))" assumes [simp]: "(\_. \(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(OclAny))" assumes [simp]: "(\_. \(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(OclAny))" assumes [simp]: "(\_. \(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(OclAny))" assumes [simp]: "(\_. \(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(OclAny))" assumes [simp]: "(\_. \(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(OclAny))" assumes [simp]: "(\a. (pre_post ((mk (a)))) = a)" shows "(mk (\\<^sub>1')) \ (OclAllInstances_generic (pre_post) (OclAny)) \ Set{X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .oclAsType(OclAny) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 .oclAsType(OclAny) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3 .oclAsType(OclAny) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4 .oclAsType(OclAny) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6 .oclAsType(OclAny) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7 , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8 , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9 .oclAsType(OclAny)}" apply(subst perm_\\<^sub>1') apply(simp only: state.make_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9_def) apply(subst state_update_vs_allInstances_generic_tc, simp, simp, (metis distinct_oid distinct_length_2_or_more)?, simp only: assms, blast, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp del: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person, simp del: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person, simp, rule OclIncluding_cong, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def) apply(subst state_update_vs_allInstances_generic_tc, simp, simp, (metis distinct_oid distinct_length_2_or_more)?, simp only: assms, blast, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp del: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person, simp del: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person, simp, rule OclIncluding_cong, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def) apply(subst state_update_vs_allInstances_generic_tc, simp, simp, (metis distinct_oid distinct_length_2_or_more)?, simp only: assms, blast, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp del: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person, simp del: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person, simp, rule OclIncluding_cong, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def) apply(subst state_update_vs_allInstances_generic_tc, simp, simp, (metis distinct_oid distinct_length_2_or_more)?, simp only: assms, blast, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp del: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person, simp del: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person, simp, rule OclIncluding_cong, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def) apply(subst state_update_vs_allInstances_generic_tc, simp, simp, (metis distinct_oid distinct_length_2_or_more)?, simp only: assms, blast, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp del: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person, simp del: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person, simp, rule OclIncluding_cong, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def) apply(subst state_update_vs_allInstances_generic_tc, simp, simp, (metis distinct_oid distinct_length_2_or_more)?, simp only: assms, blast, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp del: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person, simp del: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person, simp, rule OclIncluding_cong, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def) apply(subst state_update_vs_allInstances_generic_tc, simp, simp, (metis distinct_oid distinct_length_2_or_more)?, simp only: assms, blast, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp del: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person, simp del: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person, simp, rule OclIncluding_cong, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def) apply(subst state_update_vs_allInstances_generic_tc, simp, simp, (metis distinct_oid distinct_length_2_or_more)?, simp only: assms, blast, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp del: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person, simp del: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person, simp, rule OclIncluding_cong, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def, (simp only: assms[symmetric ])?, simp add: valid_def OclValid_def bot_fun_def bot_option_def) apply(rule state_update_vs_allInstances_generic_empty) by(simp_all only: assms, (simp_all add: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\_def)?) lemma \\<^sub>1'_OclAllInstances_at_post_exec_OclAny : assumes [simp]: "(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(OclAny ((in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y)))) \ None" assumes [simp]: "(OclAny ((in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y)))) \ None" assumes [simp]: "(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(\_. \(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(OclAny))" assumes [simp]: "(\_. \(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(OclAny))" assumes [simp]: "(\_. \(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(OclAny))" assumes [simp]: "(\_. \(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(OclAny))" assumes [simp]: "(\_. \(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(OclAny))" assumes [simp]: "(\_. \(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(OclAny))" shows "(st , \\<^sub>1') \ (OclAllInstances_at_post (OclAny)) \ Set{X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .oclAsType(OclAny) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 .oclAsType(OclAny) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3 .oclAsType(OclAny) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4 .oclAsType(OclAny) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6 .oclAsType(OclAny) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7 , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8 , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9 .oclAsType(OclAny)}" unfolding OclAllInstances_at_post_def by(rule \\<^sub>1'_OclAllInstances_generic_exec_OclAny, simp_all only: assms, simp_all) lemma \\<^sub>1'_OclAllInstances_at_pre_exec_OclAny : assumes [simp]: "(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(OclAny ((in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y)))) \ None" assumes [simp]: "(OclAny ((in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y)))) \ None" assumes [simp]: "(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) \ None" assumes [simp]: "(\_. \(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(OclAny))" assumes [simp]: "(\_. \(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(OclAny))" assumes [simp]: "(\_. \(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(OclAny))" assumes [simp]: "(\_. \(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(OclAny))" assumes [simp]: "(\_. \(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(OclAny))" assumes [simp]: "(\_. \(OclAny ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))\) = ((((\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)::\Person)) .oclAsType(OclAny))" shows "(\\<^sub>1' , st) \ (OclAllInstances_at_pre (OclAny)) \ Set{X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .oclAsType(OclAny) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 .oclAsType(OclAny) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3 .oclAsType(OclAny) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4 .oclAsType(OclAny) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6 .oclAsType(OclAny) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7 , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8 , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9 .oclAsType(OclAny)}" unfolding OclAllInstances_at_pre_def by(rule \\<^sub>1'_OclAllInstances_generic_exec_OclAny, simp_all only: assms, simp_all) ML \(Ty'.check ([]) (" error(s)"))\ end (* 161 ************************************ 967 + 1 *) (* term Floor2_examp.print_examp_def_st_def_interp *) definition "(state_interpretation_\\<^sub>1' (\)) = (state_\\<^sub>1' (oid1) (oid2) (oid3) (oid4) (oid6) (oid7) (oid8) (oid9) (\\(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 (\))\\) (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1) (\\(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 (\))\\) (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2) (\\(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3 (\))\\) (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3) (\\(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4 (\))\\) (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4) (\\(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6 (\))\\) (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6) (\\(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7 (\))\\) (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7) (\\(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8 (\))\\) (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8) (\\(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9 (\))\\) (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9))" (* 162 ************************************ 968 + 1 *) section \Transition (Floor 2)\ (* 163 ************************************ 969 + 1 *) locale transition_\\<^sub>1_\\<^sub>1' = fixes "oid1" :: "nat" fixes "oid2" :: "nat" fixes "oid3" :: "nat" fixes "oid4" :: "nat" fixes "oid5" :: "nat" fixes "oid6" :: "nat" fixes "oid7" :: "nat" fixes "oid8" :: "nat" fixes "oid9" :: "nat" assumes distinct_oid: "(distinct ([oid1 , oid2 , oid3 , oid4 , oid5 , oid6 , oid7 , oid8 , oid9]))" fixes "\\<^sub>1_object0\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" :: "ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" fixes "\\<^sub>1_object0" :: "\Person" assumes \\<^sub>1_object0_def: "\\<^sub>1_object0 = (\_. \\\\<^sub>1_object0\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)" fixes "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" :: "ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" fixes "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1" :: "\Person" assumes X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1_def: "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 = (\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)" fixes "\\<^sub>1_object1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" :: "ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" fixes "\\<^sub>1_object1" :: "\Person" assumes \\<^sub>1_object1_def: "\\<^sub>1_object1 = (\_. \\\\<^sub>1_object1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)" fixes "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" :: "ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" fixes "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2" :: "\Person" assumes X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2_def: "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 = (\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)" fixes "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" :: "ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" fixes "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3" :: "\Person" assumes X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3_def: "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3 = (\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)" fixes "\\<^sub>1_object2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" :: "ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" fixes "\\<^sub>1_object2" :: "\Person" assumes \\<^sub>1_object2_def: "\\<^sub>1_object2 = (\_. \\\\<^sub>1_object2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)" fixes "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" :: "ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" fixes "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4" :: "\Person" assumes X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4_def: "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4 = (\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)" fixes "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" :: "ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" fixes "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5" :: "\Person" assumes X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5_def: "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5 = (\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)" fixes "\\<^sub>1_object4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" :: "ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" fixes "\\<^sub>1_object4" :: "\Person" assumes \\<^sub>1_object4_def: "\\<^sub>1_object4 = (\_. \\\\<^sub>1_object4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)" fixes "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" :: "ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" fixes "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6" :: "\Person" assumes X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6_def: "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6 = (\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)" fixes "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y" :: "ty\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y" fixes "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7" :: "\OclAny" assumes X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7_def: "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7 = (\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y\\)" fixes "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y" :: "ty\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y" fixes "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8" :: "\OclAny" assumes X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8_def: "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8 = (\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y\\)" fixes "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" :: "ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" fixes "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9" :: "\Person" assumes X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9_def: "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9 = (\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)" assumes \\<^sub>1: "(state_\\<^sub>1 (oid1) (oid2) (oid4) (oid5) (oid6) (oid9) (\\<^sub>1_object0\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n) (\\<^sub>1_object0) (\\<^sub>1_object1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n) (\\<^sub>1_object1) (\\<^sub>1_object2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n) (\\<^sub>1_object2) (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n) (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5) (\\<^sub>1_object4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n) (\\<^sub>1_object4) (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n) (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9))" assumes \\<^sub>1': "(state_\\<^sub>1' (oid1) (oid2) (oid3) (oid4) (oid6) (oid7) (oid8) (oid9) (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n) (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1) (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n) (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2) (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n) (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3) (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n) (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4) (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n) (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6) (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y) (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7) (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y) (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8) (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n) (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9))" begin interpretation state_\\<^sub>1: state_\\<^sub>1 "oid1" "oid2" "oid4" "oid5" "oid6" "oid9" "\\<^sub>1_object0\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" "\\<^sub>1_object0" "\\<^sub>1_object1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" "\\<^sub>1_object1" "\\<^sub>1_object2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" "\\<^sub>1_object2" "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5" "\\<^sub>1_object4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" "\\<^sub>1_object4" "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9" by(rule \\<^sub>1) interpretation state_\\<^sub>1': state_\\<^sub>1' "oid1" "oid2" "oid3" "oid4" "oid6" "oid7" "oid8" "oid9" "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1" "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2" "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3" "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4" "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6" "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y" "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7" "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y" "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8" "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9" by(rule \\<^sub>1') definition "\\<^sub>1 = state_\\<^sub>1.\\<^sub>1" definition "\\<^sub>1' = state_\\<^sub>1'.\\<^sub>1'" lemma basic_\\<^sub>1_\\<^sub>1'_wff : assumes [simp]: "(oid_of ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object0\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) = oid1" assumes [simp]: "(oid_of ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) = oid1" assumes [simp]: "(oid_of ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) = oid2" assumes [simp]: "(oid_of ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) = oid2" assumes [simp]: "(oid_of ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) = oid3" assumes [simp]: "(oid_of ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) = oid4" assumes [simp]: "(oid_of ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) = oid4" assumes [simp]: "(oid_of ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) = oid5" assumes [simp]: "(oid_of ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) = oid6" assumes [simp]: "(oid_of ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) = oid6" assumes [simp]: "(oid_of ((in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y)))) = oid7" assumes [simp]: "(oid_of ((in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y)))) = oid8" assumes [simp]: "(oid_of ((in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)))) = oid9" shows "(WFF ((state_\\<^sub>1.\\<^sub>1 , state_\\<^sub>1'.\\<^sub>1')))" proof - have [simp]: "oid1 \ oid2" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid1 \ oid3" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid1 \ oid4" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid1 \ oid5" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid1 \ oid6" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid1 \ oid7" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid1 \ oid8" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid1 \ oid9" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid2 \ oid1" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid2 \ oid3" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid2 \ oid4" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid2 \ oid5" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid2 \ oid6" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid2 \ oid7" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid2 \ oid8" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid2 \ oid9" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid3 \ oid1" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid3 \ oid2" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid3 \ oid4" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid3 \ oid5" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid3 \ oid6" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid3 \ oid7" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid3 \ oid8" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid3 \ oid9" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid4 \ oid1" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid4 \ oid2" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid4 \ oid3" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid4 \ oid5" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid4 \ oid6" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid4 \ oid7" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid4 \ oid8" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid4 \ oid9" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid5 \ oid1" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid5 \ oid2" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid5 \ oid3" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid5 \ oid4" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid5 \ oid6" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid5 \ oid7" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid5 \ oid8" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid5 \ oid9" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid6 \ oid1" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid6 \ oid2" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid6 \ oid3" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid6 \ oid4" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid6 \ oid5" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid6 \ oid7" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid6 \ oid8" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid6 \ oid9" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid7 \ oid1" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid7 \ oid2" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid7 \ oid3" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid7 \ oid4" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid7 \ oid5" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid7 \ oid6" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid7 \ oid8" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid7 \ oid9" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid8 \ oid1" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid8 \ oid2" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid8 \ oid3" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid8 \ oid4" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid8 \ oid5" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid8 \ oid6" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid8 \ oid7" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid8 \ oid9" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid9 \ oid1" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid9 \ oid2" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid9 \ oid3" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid9 \ oid4" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid9 \ oid5" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid9 \ oid6" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid9 \ oid7" by(metis distinct_oid distinct_length_2_or_more) show ?thesis proof - have [simp]: "oid9 \ oid8" by(metis distinct_oid distinct_length_2_or_more) show ?thesis by(auto simp: WFF_def state_\\<^sub>1.\\<^sub>1_def state_\\<^sub>1'.\\<^sub>1'_def) qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed lemma oid1\\<^sub>1\\<^sub>1'_\\<^sub>1_OclIsMaintained : assumes [simp]: "(oid_of (\\<^sub>1_object0\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)) = oid1" shows "(state_\\<^sub>1.\\<^sub>1 , state_\\<^sub>1'.\\<^sub>1') \ (OclIsMaintained (\\<^sub>1_object0))" apply(simp add: state_\\<^sub>1.\\<^sub>1_def state_\\<^sub>1'.\\<^sub>1'_def \\<^sub>1_object0_def OclIsMaintained_def OclValid_def oid_of_option_def) by((metis distinct_oid distinct_length_2_or_more)?) lemma oid1\\<^sub>1\\<^sub>1'_\\<^sub>1'_OclIsMaintained : assumes [simp]: "(oid_of (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)) = oid1" shows "(state_\\<^sub>1.\\<^sub>1 , state_\\<^sub>1'.\\<^sub>1') \ (OclIsMaintained (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1))" apply(simp add: state_\\<^sub>1.\\<^sub>1_def state_\\<^sub>1'.\\<^sub>1'_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1_def OclIsMaintained_def OclValid_def oid_of_option_def) by((metis distinct_oid distinct_length_2_or_more)?) lemma oid2\\<^sub>1\\<^sub>1'_\\<^sub>1_OclIsMaintained : assumes [simp]: "(oid_of (\\<^sub>1_object1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)) = oid2" shows "(state_\\<^sub>1.\\<^sub>1 , state_\\<^sub>1'.\\<^sub>1') \ (OclIsMaintained (\\<^sub>1_object1))" apply(simp add: state_\\<^sub>1.\\<^sub>1_def state_\\<^sub>1'.\\<^sub>1'_def \\<^sub>1_object1_def OclIsMaintained_def OclValid_def oid_of_option_def) by((metis distinct_oid distinct_length_2_or_more)?) lemma oid2\\<^sub>1\\<^sub>1'_\\<^sub>1'_OclIsMaintained : assumes [simp]: "(oid_of (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)) = oid2" shows "(state_\\<^sub>1.\\<^sub>1 , state_\\<^sub>1'.\\<^sub>1') \ (OclIsMaintained (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2))" apply(simp add: state_\\<^sub>1.\\<^sub>1_def state_\\<^sub>1'.\\<^sub>1'_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2_def OclIsMaintained_def OclValid_def oid_of_option_def) by((metis distinct_oid distinct_length_2_or_more)?) lemma oid3\\<^sub>1\\<^sub>1'_\\<^sub>1'_OclIsNew : assumes [simp]: "(oid_of (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)) = oid3" shows "(state_\\<^sub>1.\\<^sub>1 , state_\\<^sub>1'.\\<^sub>1') \ (OclIsNew (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3))" apply(simp add: state_\\<^sub>1.\\<^sub>1_def state_\\<^sub>1'.\\<^sub>1'_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3_def OclIsNew_def OclValid_def oid_of_option_def) by((metis distinct_oid distinct_length_2_or_more)?) lemma oid4\\<^sub>1\\<^sub>1'_\\<^sub>1_OclIsMaintained : assumes [simp]: "(oid_of (\\<^sub>1_object2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)) = oid4" shows "(state_\\<^sub>1.\\<^sub>1 , state_\\<^sub>1'.\\<^sub>1') \ (OclIsMaintained (\\<^sub>1_object2))" apply(simp add: state_\\<^sub>1.\\<^sub>1_def state_\\<^sub>1'.\\<^sub>1'_def \\<^sub>1_object2_def OclIsMaintained_def OclValid_def oid_of_option_def) by((metis distinct_oid distinct_length_2_or_more)?) lemma oid4\\<^sub>1\\<^sub>1'_\\<^sub>1'_OclIsMaintained : assumes [simp]: "(oid_of (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)) = oid4" shows "(state_\\<^sub>1.\\<^sub>1 , state_\\<^sub>1'.\\<^sub>1') \ (OclIsMaintained (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4))" apply(simp add: state_\\<^sub>1.\\<^sub>1_def state_\\<^sub>1'.\\<^sub>1'_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4_def OclIsMaintained_def OclValid_def oid_of_option_def) by((metis distinct_oid distinct_length_2_or_more)?) lemma oid5\\<^sub>1\\<^sub>1'_\\<^sub>1_OclIsDeleted : assumes [simp]: "(oid_of (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)) = oid5" shows "(state_\\<^sub>1.\\<^sub>1 , state_\\<^sub>1'.\\<^sub>1') \ (OclIsDeleted (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5))" apply(simp add: state_\\<^sub>1.\\<^sub>1_def state_\\<^sub>1'.\\<^sub>1'_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5_def OclIsDeleted_def OclValid_def oid_of_option_def) by((metis distinct_oid distinct_length_2_or_more)?) lemma oid6\\<^sub>1\\<^sub>1'_\\<^sub>1_OclIsMaintained : assumes [simp]: "(oid_of (\\<^sub>1_object4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)) = oid6" shows "(state_\\<^sub>1.\\<^sub>1 , state_\\<^sub>1'.\\<^sub>1') \ (OclIsMaintained (\\<^sub>1_object4))" apply(simp add: state_\\<^sub>1.\\<^sub>1_def state_\\<^sub>1'.\\<^sub>1'_def \\<^sub>1_object4_def OclIsMaintained_def OclValid_def oid_of_option_def) by((metis distinct_oid distinct_length_2_or_more)?) lemma oid6\\<^sub>1\\<^sub>1'_\\<^sub>1'_OclIsMaintained : assumes [simp]: "(oid_of (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)) = oid6" shows "(state_\\<^sub>1.\\<^sub>1 , state_\\<^sub>1'.\\<^sub>1') \ (OclIsMaintained (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6))" apply(simp add: state_\\<^sub>1.\\<^sub>1_def state_\\<^sub>1'.\\<^sub>1'_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6_def OclIsMaintained_def OclValid_def oid_of_option_def) by((metis distinct_oid distinct_length_2_or_more)?) lemma oid7\\<^sub>1\\<^sub>1'_\\<^sub>1'_OclIsNew : assumes [simp]: "(oid_of (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y)) = oid7" shows "(state_\\<^sub>1.\\<^sub>1 , state_\\<^sub>1'.\\<^sub>1') \ (OclIsNew (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7))" apply(simp add: state_\\<^sub>1.\\<^sub>1_def state_\\<^sub>1'.\\<^sub>1'_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7_def OclIsNew_def OclValid_def oid_of_option_def) by((metis distinct_oid distinct_length_2_or_more)?) lemma oid8\\<^sub>1\\<^sub>1'_\\<^sub>1'_OclIsNew : assumes [simp]: "(oid_of (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y)) = oid8" shows "(state_\\<^sub>1.\\<^sub>1 , state_\\<^sub>1'.\\<^sub>1') \ (OclIsNew (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8))" apply(simp add: state_\\<^sub>1.\\<^sub>1_def state_\\<^sub>1'.\\<^sub>1'_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8_def OclIsNew_def OclValid_def oid_of_option_def) by((metis distinct_oid distinct_length_2_or_more)?) lemma oid9\\<^sub>1\\<^sub>1'_\\<^sub>1_OclIsMaintained : assumes [simp]: "(oid_of (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)) = oid9" shows "(state_\\<^sub>1.\\<^sub>1 , state_\\<^sub>1'.\\<^sub>1') \ (OclIsMaintained (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9))" apply(simp add: state_\\<^sub>1.\\<^sub>1_def state_\\<^sub>1'.\\<^sub>1'_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9_def OclIsMaintained_def OclValid_def oid_of_option_def) by((metis distinct_oid distinct_length_2_or_more)?) lemma oid9\\<^sub>1\\<^sub>1'_\\<^sub>1'_OclIsMaintained : assumes [simp]: "(oid_of (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)) = oid9" shows "(state_\\<^sub>1.\\<^sub>1 , state_\\<^sub>1'.\\<^sub>1') \ (OclIsMaintained (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9))" apply(simp add: state_\\<^sub>1.\\<^sub>1_def state_\\<^sub>1'.\\<^sub>1'_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9_def OclIsMaintained_def OclValid_def oid_of_option_def) by((metis distinct_oid distinct_length_2_or_more)?) end (* 164 ************************************ 970 + 1 *) (* term Floor2_examp.print_transition_def_interp *) definition "(pp_\\<^sub>1_\\<^sub>1' (\)) = (transition_\\<^sub>1_\\<^sub>1' (oid1) (oid2) (oid3) (oid4) (oid5) (oid6) (oid7) (oid8) (oid9) (\\(\\<^sub>1_object0 (\))\\) (\\<^sub>1_object0) (\\(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 (\))\\) (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1) (\\(\\<^sub>1_object1 (\))\\) (\\<^sub>1_object1) (\\(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 (\))\\) (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2) (\\(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3 (\))\\) (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3) (\\(\\<^sub>1_object2 (\))\\) (\\<^sub>1_object2) (\\(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4 (\))\\) (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4) (\\(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5 (\))\\) (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5) (\\(\\<^sub>1_object4 (\))\\) (\\<^sub>1_object4) (\\(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6 (\))\\) (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6) (\\(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7 (\))\\) (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7) (\\(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8 (\))\\) (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8) (\\(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9 (\))\\) (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9))" (* 165 ************************************ 971 + 3 *) (* term Floor2_examp.print_transition_lemmas_oid *) lemmas pp_oid_\\<^sub>1_\\<^sub>1' = oid1_def oid2_def oid3_def oid4_def oid5_def oid6_def oid7_def oid8_def oid9_def lemmas pp_object_\\<^sub>1_\\<^sub>1' = \\<^sub>1_object0_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1_def \\<^sub>1_object1_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3_def \\<^sub>1_object2_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5_def \\<^sub>1_object4_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9_def lemmas pp_object_ty_\\<^sub>1_\\<^sub>1' = \\<^sub>1_object0\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_def \\<^sub>1_object1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_def \\<^sub>1_object2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_def \\<^sub>1_object4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_def (* 166 ************************************ 974 + 1 *) section \Context (Floor 2)\ (* 167 ************************************ 975 + 6 *) (* term Floor2_ctxt.print_ctxt_pre_post *) axiomatization where dot__contents_Person_def: "(self::\Person) .contents() \ (\\. (Eps ((\result. (HOL.Let ((\_. result)) ((\result. (if ((\ \ ((\ (self))))) then (\ \ ((((UML_Logic.false :: (((_, Product_Type.unit) UML_Types.state.state_ext \ (_, Product_Type.unit) UML_Types.state.state_ext) \ ((HOL.bool) Option.option) Option.option)))))) \ (\ \ ((((((UML_Logic.StrongEq :: ((((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ (((Int.int) Option.option) Option.option) UML_Types.Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e) \ ((((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ (((Int.int) Option.option) Option.option) UML_Types.Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e) \ (((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ ((HOL.bool) Option.option) Option.option))))) (result)) (((((UML_Logic.OclIf :: ((((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ ((HOL.bool) Option.option) Option.option) \ ((((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ (((Int.int) Option.option) Option.option) UML_Types.Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e) \ ((((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ (((Int.int) Option.option) Option.option) UML_Types.Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e) \ (((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ (((Int.int) Option.option) Option.option) UML_Types.Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e)))))) ((((UML_Logic.StrictRefEq :: ((((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ ((Employee_DesignModel_UMLPart_generated.ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n) Option.option) Option.option) \ ((((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ ((Employee_DesignModel_UMLPart_generated.ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n) Option.option) Option.option) \ (((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ ((HOL.bool) Option.option) Option.option))))) (((Employee_DesignModel_UMLPart_generated.dot_0___boss :: ((((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ _) \ (((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ ((Employee_DesignModel_UMLPart_generated.ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n) Option.option) Option.option)))) (self))) ((UML_Types.null_class.null :: (((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ ((Employee_DesignModel_UMLPart_generated.ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n) Option.option) Option.option))))) ((((UML_Set.OclIncluding :: ((((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ (((Int.int) Option.option) Option.option) UML_Types.Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e) \ ((((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ ((Int.int) Option.option) Option.option) \ (((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ (((Int.int) Option.option) Option.option) UML_Types.Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e))))) ((UML_Set.mtSet :: (((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ (((Int.int) Option.option) Option.option) UML_Types.Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e)))) (((Employee_DesignModel_UMLPart_generated.dot__salary :: ((((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ _) \ (((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ ((Int.int) Option.option) Option.option)))) (self)))) ((((UML_Set.OclIncluding :: ((((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ (((Int.int) Option.option) Option.option) UML_Types.Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e) \ ((((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ ((Int.int) Option.option) Option.option) \ (((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ (((Int.int) Option.option) Option.option) UML_Types.Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e))))) (((Employee_DesignModel_UMLPart_generated.dot__contents :: ((((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ ((Employee_DesignModel_UMLPart_generated.ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n) Option.option) Option.option) \ (((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ (((Int.int) Option.option) Option.option) UML_Types.Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e)))) (((Employee_DesignModel_UMLPart_generated.dot_0___boss :: ((((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ _) \ (((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ ((Employee_DesignModel_UMLPart_generated.ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n) Option.option) Option.option)))) (self)))) (((Employee_DesignModel_UMLPart_generated.dot__salary :: ((((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ _) \ (((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ ((Int.int) Option.option) Option.option)))) (self)))) and (UML_Logic.true :: (((_, Product_Type.unit) UML_Types.state.state_ext \ (_, Product_Type.unit) UML_Types.state.state_ext) \ ((HOL.bool) Option.option) Option.option)))))) else (\ \ (result \ invalid))))))))))" thm dot__contents_Person_def overloading dot__contents \ "(dot__contents::(\Planet) \ _)" begin definition dot__contents_Planet : "(x::\Planet) .contents() \ x .oclAsType(Person) .contents()" end overloading dot__contents \ "(dot__contents::(\Galaxy) \ _)" begin definition dot__contents_Galaxy : "(x::\Galaxy) .contents() \ x .oclAsType(Person) .contents()" end overloading dot__contents \ "(dot__contents::(\OclAny) \ _)" begin definition dot__contents_OclAny : "(x::\OclAny) .contents() \ x .oclAsType(Person) .contents()" end ML \(Ty'.check ([]) (" error(s)"))\ (* 168 ************************************ 981 + 0 *) (* term Floor2_ctxt.print_ctxt_inv *) (* 169 ************************************ 981 + 0 *) (* term Floor2_ctxt.print_ctxt_thm *) (* 170 ************************************ 981 + 1 *) section \Context (Floor 2)\ (* 171 ************************************ 982 + 0 *) (* term Floor2_ctxt.print_ctxt_pre_post *) (* 172 ************************************ 982 + 3 *) (* term Floor2_ctxt.print_ctxt_inv *) definition "Person_aat_pre = (\\. (\ \ (UML_Set.OclForall ((OclAllInstances_at_pre (Person))) ((\self. (((UML_Logic.OclImplies :: ((((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ ((HOL.bool) Option.option) Option.option) \ ((((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ ((HOL.bool) Option.option) Option.option) \ (((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ ((HOL.bool) Option.option) Option.option))))) (((UML_Logic.OclNot :: ((((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ ((HOL.bool) Option.option) Option.option) \ (((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ ((HOL.bool) Option.option) Option.option)))) ((((UML_Logic.StrictRefEq :: ((((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ ((Employee_DesignModel_UMLPart_generated.ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n) Option.option) Option.option) \ ((((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ ((Employee_DesignModel_UMLPart_generated.ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n) Option.option) Option.option) \ (((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ ((HOL.bool) Option.option) Option.option))))) (((Employee_DesignModel_UMLPart_generated.dot_0___bossat_pre :: ((((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ _) \ (((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ ((Employee_DesignModel_UMLPart_generated.ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n) Option.option) Option.option)))) (self))) ((UML_Types.null_class.null :: (((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ ((Employee_DesignModel_UMLPart_generated.ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n) Option.option) Option.option)))))) ((((UML_Logic.StrongEq :: ((((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ ((Int.int) Option.option) Option.option) \ ((((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ ((Int.int) Option.option) Option.option) \ (((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ ((HOL.bool) Option.option) Option.option))))) (((Employee_DesignModel_UMLPart_generated.dot__salaryat_pre :: ((((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ _) \ (((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ ((Int.int) Option.option) Option.option)))) (self))) (((Employee_DesignModel_UMLPart_generated.dot__salaryat_pre :: ((((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ ((Employee_DesignModel_UMLPart_generated.ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n) Option.option) Option.option) \ (((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ ((Int.int) Option.option) Option.option)))) (((Employee_DesignModel_UMLPart_generated.dot_0___bossat_pre :: ((((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ _) \ (((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ ((Employee_DesignModel_UMLPart_generated.ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n) Option.option) Option.option)))) (self)))))))))" definition "Person_a = (\\. (\ \ (UML_Set.OclForall ((OclAllInstances_at_post (Person))) ((\self. (((UML_Logic.OclImplies :: ((((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ ((HOL.bool) Option.option) Option.option) \ ((((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ ((HOL.bool) Option.option) Option.option) \ (((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ ((HOL.bool) Option.option) Option.option))))) (((UML_Logic.OclNot :: ((((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ ((HOL.bool) Option.option) Option.option) \ (((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ ((HOL.bool) Option.option) Option.option)))) ((((UML_Logic.StrictRefEq :: ((((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ ((Employee_DesignModel_UMLPart_generated.ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n) Option.option) Option.option) \ ((((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ ((Employee_DesignModel_UMLPart_generated.ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n) Option.option) Option.option) \ (((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ ((HOL.bool) Option.option) Option.option))))) (((Employee_DesignModel_UMLPart_generated.dot_0___boss :: ((((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ _) \ (((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ ((Employee_DesignModel_UMLPart_generated.ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n) Option.option) Option.option)))) (self))) ((UML_Types.null_class.null :: (((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ ((Employee_DesignModel_UMLPart_generated.ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n) Option.option) Option.option)))))) ((((UML_Logic.StrongEq :: ((((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ ((Int.int) Option.option) Option.option) \ ((((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ ((Int.int) Option.option) Option.option) \ (((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ ((HOL.bool) Option.option) Option.option))))) (((Employee_DesignModel_UMLPart_generated.dot__salary :: ((((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ _) \ (((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ ((Int.int) Option.option) Option.option)))) (self))) (((Employee_DesignModel_UMLPart_generated.dot__salary :: ((((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ ((Employee_DesignModel_UMLPart_generated.ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n) Option.option) Option.option) \ (((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ ((Int.int) Option.option) Option.option)))) (((Employee_DesignModel_UMLPart_generated.dot_0___boss :: ((((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ _) \ (((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ ((Employee_DesignModel_UMLPart_generated.ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n) Option.option) Option.option)))) (self)))))))))" ML \(Ty'.check ([]) (" error(s)"))\ (* 173 ************************************ 985 + 1 *) (* term Floor2_ctxt.print_ctxt_thm *) thm Person_aat_pre_def Person_a_def (* 174 ************************************ 986 + 1 *) section \Context (Floor 2)\ (* 175 ************************************ 987 + 0 *) (* term Floor2_ctxt.print_ctxt_pre_post *) (* 176 ************************************ 987 + 3 *) (* term Floor2_ctxt.print_ctxt_inv *) definition "Planet_Aat_pre = (\\. (\ \ (UML_Set.OclForall ((OclAllInstances_at_pre (Planet))) ((\self. (((UML_Logic.OclAnd :: ((((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ ((HOL.bool) Option.option) Option.option) \ ((((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ ((HOL.bool) Option.option) Option.option) \ (((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ ((HOL.bool) Option.option) Option.option))))) ((UML_Logic.true :: (((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ ((HOL.bool) Option.option) Option.option)))) ((((UML_Integer.OclLe\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r :: ((((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ ((Int.int) Option.option) Option.option) \ ((((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ ((Int.int) Option.option) Option.option) \ (((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ ((HOL.bool) Option.option) Option.option))))) (((Employee_DesignModel_UMLPart_generated.dot__weightat_pre :: ((((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ _) \ (((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ ((Int.int) Option.option) Option.option)))) (self))) ((UML_Integer.OclInt0 :: (((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ ((Int.int) Option.option) Option.option)))))))))" definition "Planet_A = (\\. (\ \ (UML_Set.OclForall ((OclAllInstances_at_post (Planet))) ((\self. (((UML_Logic.OclAnd :: ((((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ ((HOL.bool) Option.option) Option.option) \ ((((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ ((HOL.bool) Option.option) Option.option) \ (((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ ((HOL.bool) Option.option) Option.option))))) ((UML_Logic.true :: (((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ ((HOL.bool) Option.option) Option.option)))) ((((UML_Integer.OclLe\<^sub>I\<^sub>n\<^sub>t\<^sub>e\<^sub>g\<^sub>e\<^sub>r :: ((((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ ((Int.int) Option.option) Option.option) \ ((((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ ((Int.int) Option.option) Option.option) \ (((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ ((HOL.bool) Option.option) Option.option))))) (((Employee_DesignModel_UMLPart_generated.dot__weight :: ((((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ _) \ (((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ ((Int.int) Option.option) Option.option)))) (self))) ((UML_Integer.OclInt0 :: (((Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext \ (Employee_DesignModel_UMLPart_generated.\, Product_Type.unit) UML_Types.state.state_ext) \ ((Int.int) Option.option) Option.option)))))))))" ML \(Ty'.check ([]) (" error(s)"))\ (* 177 ************************************ 990 + 1 *) (* term Floor2_ctxt.print_ctxt_thm *) thm Planet_Aat_pre_def Planet_A_def end